From f0407209dfb62aa238df40f6f7fda66105208605 Mon Sep 17 00:00:00 2001 From: Sebastian Ullrich Date: Sat, 18 Jun 2022 11:07:25 +0200 Subject: [PATCH 1/4] fix: auto-group `syntax` parsers where necessary --- RELEASES.md | 2 + src/Lean/Elab/Syntax.lean | 100 ++++++++++++++++++++------------- src/Lean/Parser.lean | 20 +++---- src/Lean/Parser/Extension.lean | 13 ++++- src/Lean/Parser/Extra.lean | 30 +++++----- 5 files changed, 100 insertions(+), 65 deletions(-) diff --git a/RELEASES.md b/RELEASES.md index 91e6693bdd6..1444d94e5a8 100644 --- a/RELEASES.md +++ b/RELEASES.md @@ -1,6 +1,8 @@ Unreleased --------- +* The `group(·)` `syntax` combinator is now introduced automatically where necessary, such as when using multiple parsers inside `(...)+`. + * Add ["Typed Macros"](https://github.com/leanprover/lean4/pull/1251): syntax trees produced and accepted by syntax antiquotations now remember their syntax kinds, preventing accidental production of ill-formed syntax trees and reducing the need for explicit `:kind` antiquotation annotations. See PR for details. * Aliases of protected definitions are protected too. Example: diff --git a/src/Lean/Elab/Syntax.lean b/src/Lean/Elab/Syntax.lean index 3d5a6acd40f..b668e8b9c6f 100644 --- a/src/Lean/Elab/Syntax.lean +++ b/src/Lean/Elab/Syntax.lean @@ -17,16 +17,17 @@ def expandOptPrecedence (stx : Syntax) : MacroM (Option Nat) := else return some (← evalPrec stx[0][1]) -private def mkParserSeq (ds : Array Term) : TermElabM Syntax := do +private def mkParserSeq (ds : Array (Term × Nat)) : TermElabM (Term × Nat) := do if ds.size == 0 then throwUnsupportedSyntax else if ds.size == 1 then pure ds[0] else - let mut r := ds[0] - for d in ds[1:ds.size] do + let mut (r, stackSum) := ds[0] + for (d, stackSz) in ds[1:ds.size] do r ← `(ParserDescr.binary `andthen $r $d) - return r + stackSum := stackSum + stackSz + return (r, stackSum) structure ToParserDescrContext where catName : Name @@ -36,12 +37,20 @@ structure ToParserDescrContext where behavior : Parser.LeadingIdentBehavior abbrev ToParserDescrM := ReaderT ToParserDescrContext (StateRefT (Option Nat) TermElabM) +abbrev ToParserDescr := ToParserDescrM (Term × Nat) private def markAsTrailingParser (lhsPrec : Nat) : ToParserDescrM Unit := set (some lhsPrec) @[inline] private def withNotFirst {α} (x : ToParserDescrM α) : ToParserDescrM α := withReader (fun ctx => { ctx with first := false }) x -@[inline] private def withNestedParser {α} (x : ToParserDescrM α) : ToParserDescrM α := +def ensureUnaryOutput (x : Term × Nat) : Term := + let (stx, stackSz) := x + if stackSz != 1 then + Unhygienic.run ``(ParserDescr.unary $(quote `group) $stx) + else + stx + +@[inline] private def withNestedParser (x : ToParserDescr) : ToParserDescr := do withReader (fun ctx => { ctx with leftRec := false, first := false }) x def checkLeftRec (stx : Syntax) : ToParserDescrM Bool := do @@ -83,12 +92,12 @@ open TSyntax.Compat in Given a `stx` of category `syntax`, return a pair `(newStx, lhsPrec?)`, where `newStx` is of category `term`. After elaboration, `newStx` should have type `TrailingParserDescr` if `lhsPrec?.isSome`, and `ParserDescr` otherwise. -/ -partial def toParserDescr (stx : Syntax) (catName : Name) : TermElabM (Term × Option Nat) := do +partial def toParserDescr (stx : Syntax) (catName : Name) : TermElabM ((Term × Nat) × Option Nat) := do let env ← getEnv let behavior := Parser.leadingIdentBehavior env catName (process stx { catName := catName, first := true, leftRec := true, behavior := behavior }).run none where - process (stx : Syntax) : ToParserDescrM Term := withRef stx do + process (stx : Syntax) : ToParserDescr := withRef stx do let kind := stx.getKind if kind == nullKind then processSeq stx @@ -99,9 +108,9 @@ where else if kind == ``Lean.Parser.Syntax.cat then processNullaryOrCat stx else if kind == ``Lean.Parser.Syntax.unary then - processUnary stx + processAlias stx[0] #[stx[2]] else if kind == ``Lean.Parser.Syntax.binary then - processBinary stx + processAlias stx[0] #[stx[2], stx[4]] else if kind == ``Lean.Parser.Syntax.sepBy then processSepBy stx else if kind == ``Lean.Parser.Syntax.sepBy1 then @@ -138,12 +147,39 @@ where throwErrorAt stx "invalid atomic left recursive syntax" let prec? ← liftMacroM <| expandOptPrecedence stx[1] let prec := prec?.getD 0 - `(ParserDescr.cat $(quote catName) $(quote prec)) - + return (← `(ParserDescr.cat $(quote catName) $(quote prec)), 1) + + processAlias (id : Syntax) (args : Array Syntax) := do + let aliasName := id.getId.eraseMacroScopes + let info ← Parser.getParserAliasInfo aliasName + let args ← args.mapM (withNestedParser ∘ process) + let (args, stackSz) := if let some stackSz := info.stackSz? then + if !info.autoGroupArgs then + (args.map (·.1), stackSz) + else + (args.map ensureUnaryOutput, stackSz) + else + let (args, stackSzs) := args.unzip + (args, stackSzs.foldl (· + ·) 0) + let stx ← match args with + | #[] => Parser.ensureConstantParserAlias aliasName; ``(ParserDescr.const $(quote aliasName)) + | #[p1] => Parser.ensureUnaryParserAlias aliasName; ``(ParserDescr.unary $(quote aliasName) $p1) + | #[p1, p2] => Parser.ensureBinaryParserAlias aliasName; ``(ParserDescr.binary $(quote aliasName) $p1 $p2) + | _ => unreachable! + return (stx, stackSz) + processNullaryOrCat (stx : Syntax) := do match (← resolveParserName stx[0]) with - | [(c, true)] => ensureNoPrec stx; return mkIdentFrom stx c - | [(c, false)] => ensureNoPrec stx; `(ParserDescr.parser $(quote c)) + | [(c, true)] => + ensureNoPrec stx + -- `syntax _ :=` at least enforces this + let stackSz := 1 + return (mkIdentFrom stx c, stackSz) + | [(c, false)] => + ensureNoPrec stx + -- as usual, we assume that people using `Parser` know what they are doing + let stackSz := 1 + return (← `(ParserDescr.parser $(quote c)), stackSz) | cs@(_ :: _ :: _) => throwError "ambiguous parser declaration {cs.map (·.1)}" | [] => let id := stx[0].getId.eraseMacroScopes @@ -151,37 +187,23 @@ where processParserCategory stx else if (← Parser.isParserAlias id) then ensureNoPrec stx - Parser.ensureConstantParserAlias id - `(ParserDescr.const $(quote id)) + processAlias stx[0] #[] else throwError "unknown parser declaration/category/alias '{id}'" - processUnary (stx : Syntax) := do - let aliasName := (stx[0].getId).eraseMacroScopes - Parser.ensureUnaryParserAlias aliasName - let d ← withNestedParser do process stx[2] - `(ParserDescr.unary $(quote aliasName) $d) - - processBinary (stx : Syntax) := do - let aliasName := (stx[0].getId).eraseMacroScopes - Parser.ensureBinaryParserAlias aliasName - let d₁ ← withNestedParser do process stx[2] - let d₂ ← withNestedParser do process stx[4] - `(ParserDescr.binary $(quote aliasName) $d₁ $d₂) - processSepBy (stx : Syntax) := do - let p ← withNestedParser $ process stx[1] + let p ← ensureUnaryOutput <$> withNestedParser do process stx[1] let sep := stx[3] - let psep ← if stx[4].isNone then `(ParserDescr.symbol $sep) else process stx[4][1] + let psep ← if stx[4].isNone then `(ParserDescr.symbol $sep) else ensureUnaryOutput <$> withNestedParser do process stx[4][1] let allowTrailingSep := !stx[5].isNone - `(ParserDescr.sepBy $p $sep $psep $(quote allowTrailingSep)) + return (← `(ParserDescr.sepBy $p $sep $psep $(quote allowTrailingSep)), 1) processSepBy1 (stx : Syntax) := do - let p ← withNestedParser do process stx[1] + let p ← ensureUnaryOutput <$> withNestedParser do process stx[1] let sep := stx[3] - let psep ← if stx[4].isNone then `(ParserDescr.symbol $sep) else process stx[4][1] + let psep ← if stx[4].isNone then `(ParserDescr.symbol $sep) else ensureUnaryOutput <$> withNestedParser do process stx[4][1] let allowTrailingSep := !stx[5].isNone - `(ParserDescr.sepBy1 $p $sep $psep $(quote allowTrailingSep)) + return (← `(ParserDescr.sepBy1 $p $sep $psep $(quote allowTrailingSep)), 1) isValidAtom (s : String) : Bool := !s.isEmpty && @@ -198,14 +220,14 @@ where /- For syntax categories where initialized with `LeadingIdentBehavior` different from default (e.g., `tactic`), we automatically mark the first symbol as nonReserved. -/ if (← read).behavior != Parser.LeadingIdentBehavior.default && (← read).first then - `(ParserDescr.nonReservedSymbol $(quote atom) false) + return (← `(ParserDescr.nonReservedSymbol $(quote atom) false), 1) else - `(ParserDescr.symbol $(quote atom)) + return (← `(ParserDescr.symbol $(quote atom)), 1) | none => throwUnsupportedSyntax processNonReserved (stx : Syntax) := do match stx[1].isStrLit? with - | some atom => `(ParserDescr.nonReservedSymbol $(quote atom) false) + | some atom => return (← `(ParserDescr.nonReservedSymbol $(quote atom) false), 1) | none => throwUnsupportedSyntax @@ -319,7 +341,7 @@ def resolveSyntaxKind (k : Name) : CommandElabM Name := do let prio ← liftMacroM <| evalOptPrio prio? let stxNodeKind := (← getCurrNamespace) ++ name let catParserId := mkIdentFrom stx (cat.appendAfter "Parser") - let (val, lhsPrec?) ← runTermElabM none fun _ => Term.toParserDescr syntaxParser cat + let ((val, _), lhsPrec?) ← runTermElabM none fun _ => Term.toParserDescr syntaxParser cat let declName := mkIdentFrom stx name let d ← if let some lhsPrec := lhsPrec? then `($[$doc?:docComment]? @[$attrKind:attrKind $catParserId:ident $(quote prio):num] def $declName:ident : Lean.TrailingParserDescr := @@ -333,7 +355,7 @@ def resolveSyntaxKind (k : Name) : CommandElabM Name := do @[builtinCommandElab «syntaxAbbrev»] def elabSyntaxAbbrev : CommandElab := fun stx => do let `($[$doc?:docComment]? syntax $declName:ident := $[$ps:stx]*) ← pure stx | throwUnsupportedSyntax -- TODO: nonatomic names - let (val, _) ← runTermElabM none fun _ => Term.toParserDescr (mkNullNode ps) Name.anonymous + let ((val, _), _) ← runTermElabM none fun _ => Term.toParserDescr (mkNullNode ps) Name.anonymous let stxNodeKind := (← getCurrNamespace) ++ declName.getId let stx' ← `($[$doc?:docComment]? def $declName:ident : Lean.ParserDescr := ParserDescr.nodeWithAntiquot $(quote (toString declName.getId)) $(quote stxNodeKind) $val) withMacroExpansion stx stx' <| elabCommand stx' diff --git a/src/Lean/Parser.lean b/src/Lean/Parser.lean index e9b849b1d37..fa19ed88911 100644 --- a/src/Lean/Parser.lean +++ b/src/Lean/Parser.lean @@ -19,28 +19,28 @@ open Lean.PrettyPrinter.Parenthesizer open Lean.PrettyPrinter.Formatter builtin_initialize - register_parser_alias "ws" checkWsBefore - register_parser_alias "noWs" checkNoWsBefore - register_parser_alias "linebreak" checkLinebreakBefore + register_parser_alias "ws" checkWsBefore { stackSz? := none } + register_parser_alias "noWs" checkNoWsBefore { stackSz? := none } + register_parser_alias "linebreak" checkLinebreakBefore { stackSz? := none } register_parser_alias (kind := numLitKind) "num" numLit register_parser_alias (kind := strLitKind) "str" strLit register_parser_alias (kind := charLitKind) "char" charLit register_parser_alias (kind := nameLitKind) "name" nameLit register_parser_alias (kind := scientificLitKind) "scientific" scientificLit register_parser_alias (kind := identKind) "ident" ident - register_parser_alias "colGt" checkColGt - register_parser_alias "colGe" checkColGe - register_parser_alias lookahead - register_parser_alias atomic + register_parser_alias "colGt" checkColGt { stackSz? := none } + register_parser_alias "colGe" checkColGe { stackSz? := none } + register_parser_alias lookahead { stackSz? := some 0 } + register_parser_alias atomic { stackSz? := none } register_parser_alias many register_parser_alias many1 register_parser_alias manyIndent register_parser_alias many1Indent - register_parser_alias optional - register_parser_alias withPosition + register_parser_alias optional { autoGroupArgs := false } + register_parser_alias withPosition { stackSz? := none } register_parser_alias (kind := interpolatedStrKind) interpolatedStr register_parser_alias orelse - register_parser_alias andthen + register_parser_alias andthen { stackSz? := none } registerAlias "notFollowedBy" (notFollowedBy · "element") Parenthesizer.registerAlias "notFollowedBy" notFollowedBy.parenthesizer diff --git a/src/Lean/Parser/Extension.lean b/src/Lean/Parser/Extension.lean index 64dfa210000..3d0e7e9d788 100644 --- a/src/Lean/Parser/Extension.lean +++ b/src/Lean/Parser/Extension.lean @@ -202,14 +202,25 @@ def getBinaryAlias {α} (mapRef : IO.Ref (AliasTable α)) (aliasName : Name) : I abbrev ParserAliasValue := AliasValue Parser +structure ParserAliasInfo where + /-- Number of syntax nodes produced by this parser. `none` means "sum of input sizes". -/ + stackSz? : Option Nat := some 1 + /-- Whether arguments should be wrapped in `group(·)` if they do not produce exactly one syntax node. -/ + autoGroupArgs : Bool := stackSz?.isSome + builtin_initialize parserAliasesRef : IO.Ref (NameMap ParserAliasValue) ← IO.mkRef {} builtin_initialize parserAlias2kindRef : IO.Ref (NameMap SyntaxNodeKind) ← IO.mkRef {} +builtin_initialize parserAliases2infoRef : IO.Ref (NameMap ParserAliasInfo) ← IO.mkRef {} + +def getParserAliasInfo (aliasName : Name) : IO ParserAliasInfo := do + return (← parserAliases2infoRef.get).findD aliasName {} -- Later, we define macro `register_parser_alias` which registers a parser, formatter and parenthesizer -def registerAlias (aliasName : Name) (p : ParserAliasValue) (kind? : Option SyntaxNodeKind := none) : IO Unit := do +def registerAlias (aliasName : Name) (p : ParserAliasValue) (kind? : Option SyntaxNodeKind := none) (info : ParserAliasInfo := {}) : IO Unit := do registerAliasCore parserAliasesRef aliasName p if let some kind := kind? then parserAlias2kindRef.modify (·.insert aliasName kind) + parserAliases2infoRef.modify (·.insert aliasName info) instance : Coe Parser ParserAliasValue := { coe := AliasValue.const } instance : Coe (Parser → Parser) ParserAliasValue := { coe := AliasValue.unary } diff --git a/src/Lean/Parser/Extra.lean b/src/Lean/Parser/Extra.lean index 18a3c3387e6..bf5e0c35d13 100644 --- a/src/Lean/Parser/Extra.lean +++ b/src/Lean/Parser/Extra.lean @@ -169,29 +169,29 @@ attribute [runBuiltinParserAttributeHooks] ppHardSpace ppSpace ppLine ppGroup ppRealGroup ppRealFill ppIndent ppDedent ppAllowUngrouped ppDedentIfGrouped ppHardLineUnlessUngrouped -syntax "register_parser_alias" group("(" &"kind" " := " term ")")? (strLit)? ident : term +syntax "register_parser_alias" group("(" &"kind" " := " term ")")? (strLit)? ident (colGt term)? : term macro_rules - | `(register_parser_alias $[(kind := $kind?)]? $(aliasName?)? $declName) => do + | `(register_parser_alias $[(kind := $kind?)]? $(aliasName?)? $declName $(info?)?) => do let [(fullDeclName, [])] ← Macro.resolveGlobalName declName.getId | Macro.throwError "expected non-overloaded constant name" let aliasName := aliasName?.getD (Syntax.mkStrLit declName.getId.toString) - `(do Parser.registerAlias $aliasName $declName (kind? := some $(kind?.getD (quote fullDeclName))) + `(do Parser.registerAlias $aliasName $declName $(info?.getD (Unhygienic.run `({}))) (kind? := some $(kind?.getD (quote fullDeclName))) PrettyPrinter.Formatter.registerAlias $aliasName $(mkIdentFrom declName (declName.getId ++ `formatter)) PrettyPrinter.Parenthesizer.registerAlias $aliasName $(mkIdentFrom declName (declName.getId ++ `parenthesizer))) builtin_initialize - register_parser_alias group - register_parser_alias ppHardSpace - register_parser_alias ppSpace - register_parser_alias ppLine - register_parser_alias ppGroup - register_parser_alias ppRealGroup - register_parser_alias ppRealFill - register_parser_alias ppIndent - register_parser_alias ppDedent - register_parser_alias ppAllowUngrouped - register_parser_alias ppDedentIfGrouped - register_parser_alias ppHardLineUnlessUngrouped + register_parser_alias group { autoGroupArgs := false } + register_parser_alias ppHardSpace { stackSz? := none } + register_parser_alias ppSpace { stackSz? := none } + register_parser_alias ppLine { stackSz? := none } + register_parser_alias ppGroup { stackSz? := none } + register_parser_alias ppRealGroup { stackSz? := none } + register_parser_alias ppRealFill { stackSz? := none } + register_parser_alias ppIndent { stackSz? := none } + register_parser_alias ppDedent { stackSz? := none } + register_parser_alias ppAllowUngrouped { stackSz? := none } + register_parser_alias ppDedentIfGrouped { stackSz? := none } + register_parser_alias ppHardLineUnlessUngrouped { stackSz? := none } end Parser From d6659e7802135efd19ea7f477801f4e871895915 Mon Sep 17 00:00:00 2001 From: Sebastian Ullrich Date: Tue, 28 Jun 2022 10:56:43 +0200 Subject: [PATCH 2/4] chore: update stage0 --- stage0/src/Init/Conv.lean | 4 +- stage0/src/Init/Data/Format/Syntax.lean | 3 + stage0/src/Init/Data/List/Basic.lean | 29 +- stage0/src/Init/Data/OfScientific.lean | 16 +- stage0/src/Init/Data/Option/Basic.lean | 3 - stage0/src/Init/Data/Random.lean | 4 +- stage0/src/Init/Data/String/Basic.lean | 2 +- stage0/src/Init/Meta.lean | 231 +- stage0/src/Init/Notation.lean | 29 +- stage0/src/Init/NotationExtra.lean | 48 +- stage0/src/Init/Prelude.lean | 47 +- stage0/src/Init/SimpLemmas.lean | 3 + stage0/src/Init/System/IO.lean | 7 +- stage0/src/Init/Tactics.lean | 32 +- stage0/src/Lean/Elab/App.lean | 31 +- stage0/src/Lean/Elab/AuxDef.lean | 2 +- stage0/src/Lean/Elab/AuxDiscr.lean | 5 +- stage0/src/Lean/Elab/Binders.lean | 3 +- stage0/src/Lean/Elab/BindersUtil.lean | 24 +- stage0/src/Lean/Elab/BuiltinCommand.lean | 11 +- stage0/src/Lean/Elab/BuiltinNotation.lean | 83 +- stage0/src/Lean/Elab/Command.lean | 4 +- stage0/src/Lean/Elab/Declaration.lean | 3 +- stage0/src/Lean/Elab/Deriving/BEq.lean | 12 +- stage0/src/Lean/Elab/Deriving/Basic.lean | 6 +- stage0/src/Lean/Elab/Deriving/DecEq.lean | 8 +- stage0/src/Lean/Elab/Deriving/FromToJson.lean | 58 +- stage0/src/Lean/Elab/Deriving/Hashable.lean | 10 +- stage0/src/Lean/Elab/Deriving/Inhabited.lean | 8 +- stage0/src/Lean/Elab/Deriving/Ord.lean | 16 +- stage0/src/Lean/Elab/Deriving/Repr.lean | 20 +- stage0/src/Lean/Elab/Deriving/Util.lean | 23 +- stage0/src/Lean/Elab/Do.lean | 9 +- stage0/src/Lean/Elab/ElabRules.lean | 13 +- stage0/src/Lean/Elab/Macro.lean | 9 +- stage0/src/Lean/Elab/MacroArgUtil.lean | 22 +- stage0/src/Lean/Elab/MacroRules.lean | 6 +- stage0/src/Lean/Elab/Match.lean | 10 +- stage0/src/Lean/Elab/Mixfix.lean | 2 +- stage0/src/Lean/Elab/MutualDef.lean | 4 +- stage0/src/Lean/Elab/Notation.lean | 95 +- stage0/src/Lean/Elab/PatternVar.lean | 9 +- stage0/src/Lean/Elab/PreDefinition/Main.lean | 2 +- stage0/src/Lean/Elab/Print.lean | 2 +- stage0/src/Lean/Elab/Quotation.lean | 31 +- stage0/src/Lean/Elab/Quotation/Precheck.lean | 4 +- stage0/src/Lean/Elab/Quotation/Util.lean | 4 +- stage0/src/Lean/Elab/StructInst.lean | 13 +- stage0/src/Lean/Elab/Structure.lean | 3 +- stage0/src/Lean/Elab/Syntax.lean | 109 +- .../src/Lean/Elab/Tactic/BuiltinTactic.lean | 9 +- stage0/src/Lean/Elab/Tactic/Conv/Congr.lean | 2 +- stage0/src/Lean/Elab/Tactic/Match.lean | 7 +- stage0/src/Lean/Elab/Tactic/Simp.lean | 19 +- stage0/src/Lean/Elab/Term.lean | 2 +- stage0/src/Lean/Exception.lean | 14 +- stage0/src/Lean/Hygiene.lean | 7 +- stage0/src/Lean/Level.lean | 10 +- stage0/src/Lean/Message.lean | 1 + stage0/src/Lean/Meta/Basic.lean | 18 +- stage0/src/Lean/Meta/DiscrTree.lean | 2 +- stage0/src/Lean/Parser.lean | 20 +- stage0/src/Lean/Parser/Basic.lean | 3 + stage0/src/Lean/Parser/Command.lean | 3 +- stage0/src/Lean/Parser/Extension.lean | 13 +- stage0/src/Lean/Parser/Extra.lean | 40 +- stage0/src/Lean/Parser/Module.lean | 2 +- stage0/src/Lean/Parser/Term.lean | 19 +- stage0/src/Lean/PrettyPrinter.lean | 2 +- .../Lean/PrettyPrinter/Delaborator/Basic.lean | 14 +- .../PrettyPrinter/Delaborator/Builtins.lean | 17 +- .../src/Lean/PrettyPrinter/Parenthesizer.lean | 6 +- stage0/src/Lean/ResolveName.lean | 17 +- stage0/src/Lean/Server/FileWorker.lean | 101 +- .../Server/FileWorker/RequestHandling.lean | 14 +- stage0/src/Lean/Server/FileWorker/Utils.lean | 7 +- stage0/src/Lean/Server/InfoUtils.lean | 2 - stage0/src/Lean/Server/Requests.lean | 4 +- stage0/src/Lean/Server/Rpc/Deriving.lean | 54 +- stage0/src/Lean/Server/Snapshots.lean | 5 - stage0/src/Lean/Util/Trace.lean | 2 +- stage0/src/bin/lean-gdb.py | 7 +- stage0/src/stdlib_flags.h | 2 +- stage0/stdlib/Init/Classical.c | 2440 +- stage0/stdlib/Init/Control/Basic.c | 360 +- stage0/stdlib/Init/Conv.c | 1717 +- stage0/stdlib/Init/Core.c | 836 +- stage0/stdlib/Init/Data/Format/Macro.c | 47 +- stage0/stdlib/Init/Data/Format/Syntax.c | 233 +- stage0/stdlib/Init/Data/List/Basic.c | 181 +- stage0/stdlib/Init/Data/List/BasicAux.c | 882 +- stage0/stdlib/Init/Data/OfScientific.c | 124 +- stage0/stdlib/Init/Data/Option/Basic.c | 80 +- stage0/stdlib/Init/Data/Random.c | 77 +- stage0/stdlib/Init/Data/String/Basic.c | 37 +- stage0/stdlib/Init/Data/String/Extra.c | 27 +- stage0/stdlib/Init/Data/ToString/Macro.c | 47 +- stage0/stdlib/Init/Meta.c | 17372 +++++----- stage0/stdlib/Init/Notation.c | 7787 ++--- stage0/stdlib/Init/NotationExtra.c | 13562 ++++---- stage0/stdlib/Init/Prelude.c | 135 +- stage0/stdlib/Init/System/IO.c | 1173 +- stage0/stdlib/Init/Tactics.c | 3610 +- stage0/stdlib/Init/WFTactics.c | 30 +- stage0/stdlib/Lean/Attributes.c | 46 +- stage0/stdlib/Lean/Compiler/ExternAttr.c | 48 +- stage0/stdlib/Lean/Compiler/IR/Basic.c | 12 +- stage0/stdlib/Lean/Compiler/IR/EmitC.c | 10 +- .../stdlib/Lean/Compiler/ImplementedByAttr.c | 4 +- stage0/stdlib/Lean/Compiler/InitAttr.c | 4 +- stage0/stdlib/Lean/Data/FuzzyMatching.c | 17 +- stage0/stdlib/Lean/Data/Json/FromToJson.c | 1 + stage0/stdlib/Lean/Data/Json/Parser.c | 8 +- stage0/stdlib/Lean/Data/JsonRpc.c | 4 +- stage0/stdlib/Lean/Data/Lsp/Capabilities.c | 8 +- stage0/stdlib/Lean/Data/Lsp/Diagnostics.c | 74 +- .../stdlib/Lean/Data/Lsp/LanguageFeatures.c | 2764 +- stage0/stdlib/Lean/Data/Name.c | 27 +- stage0/stdlib/Lean/Data/Options.c | 1947 +- stage0/stdlib/Lean/Elab/App.c | 5699 ++-- stage0/stdlib/Lean/Elab/Attributes.c | 41 +- stage0/stdlib/Lean/Elab/AuxDef.c | 925 +- stage0/stdlib/Lean/Elab/Binders.c | 1787 +- stage0/stdlib/Lean/Elab/BindersUtil.c | 1251 +- stage0/stdlib/Lean/Elab/BuiltinCommand.c | 1893 +- stage0/stdlib/Lean/Elab/BuiltinNotation.c | 7853 ++--- stage0/stdlib/Lean/Elab/BuiltinTerm.c | 216 +- stage0/stdlib/Lean/Elab/Command.c | 372 +- stage0/stdlib/Lean/Elab/DeclModifiers.c | 29 +- stage0/stdlib/Lean/Elab/Declaration.c | 233 +- stage0/stdlib/Lean/Elab/DefView.c | 299 +- stage0/stdlib/Lean/Elab/Deriving/BEq.c | 2374 +- stage0/stdlib/Lean/Elab/Deriving/Basic.c | 770 +- stage0/stdlib/Lean/Elab/Deriving/DecEq.c | 1256 +- stage0/stdlib/Lean/Elab/Deriving/FromToJson.c | 3525 +- stage0/stdlib/Lean/Elab/Deriving/Hashable.c | 761 +- stage0/stdlib/Lean/Elab/Deriving/Inhabited.c | 1574 +- stage0/stdlib/Lean/Elab/Deriving/Ord.c | 1721 +- stage0/stdlib/Lean/Elab/Deriving/Repr.c | 3861 +-- stage0/stdlib/Lean/Elab/Deriving/Util.c | 1515 +- stage0/stdlib/Lean/Elab/Do.c | 4079 ++- stage0/stdlib/Lean/Elab/ElabRules.c | 2073 +- stage0/stdlib/Lean/Elab/Extra.c | 226 +- stage0/stdlib/Lean/Elab/GenInjective.c | 4 +- stage0/stdlib/Lean/Elab/Inductive.c | 12 +- stage0/stdlib/Lean/Elab/Level.c | 118 +- stage0/stdlib/Lean/Elab/Macro.c | 744 +- stage0/stdlib/Lean/Elab/MacroArgUtil.c | 15928 ++++----- stage0/stdlib/Lean/Elab/MacroRules.c | 2163 +- stage0/stdlib/Lean/Elab/Match.c | 3520 +- stage0/stdlib/Lean/Elab/Mixfix.c | 2727 +- stage0/stdlib/Lean/Elab/MutualDef.c | 1598 +- stage0/stdlib/Lean/Elab/Notation.c | 7713 ++--- stage0/stdlib/Lean/Elab/Open.c | 4 +- stage0/stdlib/Lean/Elab/PatternVar.c | 2837 +- stage0/stdlib/Lean/Elab/PreDefinition/Main.c | 895 +- .../PreDefinition/Structural/FindRecArg.c | 4 +- .../stdlib/Lean/Elab/PreDefinition/WF/Rel.c | 4727 +-- stage0/stdlib/Lean/Elab/Print.c | 183 +- stage0/stdlib/Lean/Elab/Quotation.c | 27791 ++++++++-------- stage0/stdlib/Lean/Elab/Quotation/Precheck.c | 328 +- stage0/stdlib/Lean/Elab/Quotation/Util.c | 232 +- stage0/stdlib/Lean/Elab/SetOption.c | 184 +- stage0/stdlib/Lean/Elab/StructInst.c | 2845 +- stage0/stdlib/Lean/Elab/Structure.c | 356 +- stage0/stdlib/Lean/Elab/Syntax.c | 19093 ++++++----- stage0/stdlib/Lean/Elab/Tactic/Basic.c | 55 +- .../stdlib/Lean/Elab/Tactic/BuiltinTactic.c | 1394 +- stage0/stdlib/Lean/Elab/Tactic/Config.c | 5618 ++-- stage0/stdlib/Lean/Elab/Tactic/Conv/Basic.c | 188 +- stage0/stdlib/Lean/Elab/Tactic/Conv/Congr.c | 163 +- stage0/stdlib/Lean/Elab/Tactic/Delta.c | 4 +- stage0/stdlib/Lean/Elab/Tactic/Induction.c | 45 +- stage0/stdlib/Lean/Elab/Tactic/Match.c | 2070 +- stage0/stdlib/Lean/Elab/Tactic/Rewrite.c | 257 +- stage0/stdlib/Lean/Elab/Tactic/Simp.c | 572 +- stage0/stdlib/Lean/Elab/Tactic/Unfold.c | 4 +- stage0/stdlib/Lean/Elab/Term.c | 221 +- stage0/stdlib/Lean/Elab/Util.c | 112 +- stage0/stdlib/Lean/Exception.c | 1011 +- stage0/stdlib/Lean/Expr.c | 6 +- stage0/stdlib/Lean/Hygiene.c | 78 +- stage0/stdlib/Lean/KeyedDeclsAttribute.c | 24 +- stage0/stdlib/Lean/Level.c | 55 +- stage0/stdlib/Lean/Message.c | 76 +- stage0/stdlib/Lean/Meta/Basic.c | 24 +- stage0/stdlib/Lean/Meta/GetConst.c | 6 +- stage0/stdlib/Lean/Meta/IndPredBelow.c | 4 +- stage0/stdlib/Lean/Meta/Match/Basic.c | 14 +- stage0/stdlib/Lean/Meta/Match/Match.c | 62 +- stage0/stdlib/Lean/Meta/PPGoal.c | 18 +- stage0/stdlib/Lean/Meta/RecursorInfo.c | 96 +- stage0/stdlib/Lean/Meta/Reduce.c | 6 +- .../Lean/Meta/Tactic/LinearArith/Solver.c | 412 +- .../Lean/Meta/Tactic/Simp/SimpTheorems.c | 2343 +- stage0/stdlib/Lean/Meta/WHNF.c | 8 +- stage0/stdlib/Lean/MetavarContext.c | 4 +- stage0/stdlib/Lean/Parser.c | 2523 +- stage0/stdlib/Lean/Parser/Attr.c | 112 +- stage0/stdlib/Lean/Parser/Basic.c | 641 +- stage0/stdlib/Lean/Parser/Command.c | 2828 +- stage0/stdlib/Lean/Parser/Do.c | 956 +- stage0/stdlib/Lean/Parser/Extension.c | 1267 +- stage0/stdlib/Lean/Parser/Extra.c | 5389 +-- stage0/stdlib/Lean/Parser/Level.c | 70 +- stage0/stdlib/Lean/Parser/Module.c | 158 +- stage0/stdlib/Lean/Parser/StrInterpolation.c | 79 +- stage0/stdlib/Lean/Parser/Syntax.c | 402 +- stage0/stdlib/Lean/Parser/Tactic.c | 195 +- stage0/stdlib/Lean/Parser/Term.c | 5484 +-- stage0/stdlib/Lean/ParserCompiler.c | 85 +- .../Lean/PrettyPrinter/Delaborator/Basic.c | 18 +- .../Lean/PrettyPrinter/Delaborator/Builtins.c | 2300 +- .../Delaborator/TopDownAnalyze.c | 4 +- stage0/stdlib/Lean/PrettyPrinter/Formatter.c | 259 +- .../stdlib/Lean/PrettyPrinter/Parenthesizer.c | 123 +- stage0/stdlib/Lean/ResolveName.c | 398 +- stage0/stdlib/Lean/Server/Completion.c | 215 +- stage0/stdlib/Lean/Server/FileWorker.c | 3534 +- .../Lean/Server/FileWorker/RequestHandling.c | 7118 ++-- stage0/stdlib/Lean/Server/FileWorker/Utils.c | 25 - .../Lean/Server/FileWorker/WidgetRequests.c | 4 +- stage0/stdlib/Lean/Server/InfoUtils.c | 660 +- stage0/stdlib/Lean/Server/References.c | 12 +- stage0/stdlib/Lean/Server/Requests.c | 33 +- stage0/stdlib/Lean/Server/Rpc/Deriving.c | 6344 ++-- .../stdlib/Lean/Server/Rpc/RequestHandling.c | 344 +- stage0/stdlib/Lean/Server/Snapshots.c | 204 +- stage0/stdlib/Lean/Server/Watchdog.c | 173 +- stage0/stdlib/Lean/SubExpr.c | 4 +- stage0/stdlib/Lean/Syntax.c | 314 +- stage0/stdlib/Lean/Util/Path.c | 12 +- stage0/stdlib/Lean/Util/SCC.c | 14 +- stage0/stdlib/Lean/Util/Trace.c | 1766 +- stage0/stdlib/Lean/Widget/InteractiveGoal.c | 16 +- 235 files changed, 132499 insertions(+), 122181 deletions(-) diff --git a/stage0/src/Init/Conv.lean b/stage0/src/Init/Conv.lean index 6b650375d62..8636c6c97cb 100644 --- a/stage0/src/Init/Conv.lean +++ b/stage0/src/Init/Conv.lean @@ -54,7 +54,9 @@ macro "erw " s:rwRuleSeq : conv => `(rw (config := { transparency := Meta.Transp macro "args" : conv => `(congr) macro "left" : conv => `(lhs) macro "right" : conv => `(rhs) -macro "intro " xs:(colGt ident)* : conv => `(ext $xs*) +syntax "intro " (colGt ident)* : conv +macro_rules + | `(conv| intro $[$xs:ident]*) => `(conv| ext $xs*) syntax enterArg := ident <|> group("@"? num) syntax "enter " "[" (colGt enterArg),+ "]": conv diff --git a/stage0/src/Init/Data/Format/Syntax.lean b/stage0/src/Init/Data/Format/Syntax.lean index a970c912941..4a551e71858 100644 --- a/stage0/src/Init/Data/Format/Syntax.lean +++ b/stage0/src/Init/Data/Format/Syntax.lean @@ -44,4 +44,7 @@ def formatStx (stx : Syntax) (maxDepth : Option Nat := none) (showInfo := false) instance : ToFormat (Syntax) := ⟨formatStx⟩ instance : ToString (Syntax) := ⟨@toString Format _ ∘ format⟩ +instance : ToFormat (TSyntax k) := ⟨(format ·.raw)⟩ +instance : ToString (TSyntax k) := ⟨(toString ·.raw)⟩ + end Lean.Syntax diff --git a/stage0/src/Init/Data/List/Basic.lean b/stage0/src/Init/Data/List/Basic.lean index 1de7d3ad75c..e5fa0249223 100644 --- a/stage0/src/Init/Data/List/Basic.lean +++ b/stage0/src/Init/Data/List/Basic.lean @@ -162,14 +162,33 @@ def join : List (List α) → List α | none => filterMap f as | some b => b :: filterMap f as -@[specialize] def filterAux (p : α → Bool) : List α → List α → List α +def filter (p : α → Bool) : List α → List α + | [] => [] + | a::as => match p a with + | true => a :: filter p as + | false => filter p as + +@[specialize] def filterTRAux (p : α → Bool) : List α → List α → List α | [], rs => rs.reverse | a::as, rs => match p a with - | true => filterAux p as (a::rs) - | false => filterAux p as rs + | true => filterTRAux p as (a::rs) + | false => filterTRAux p as rs + +@[inline] def filterTR (p : α → Bool) (as : List α) : List α := + filterTRAux p as [] + +theorem filterTRAux_eq (p : α → Bool) (as bs : List α) : filterTRAux p as bs = bs.reverse ++ filter p as := by + induction as generalizing bs with + | nil => simp [filterTRAux, filter] + | cons a as ih => + simp [filterTRAux, filter] + split + next => rw [ih, reverse_cons, append_assoc]; simp + next => rw [ih] -@[inline] def filter (p : α → Bool) (as : List α) : List α := - filterAux p as [] +@[csimp] theorem filter_eq_filterTR : @filter = @filterTR := by + apply funext; intro α; apply funext; intro p; apply funext; intro as + simp [filterTR, filterTRAux_eq] @[specialize] def partitionAux (p : α → Bool) : List α → List α × List α → List α × List α | [], (bs, cs) => (bs.reverse, cs.reverse) diff --git a/stage0/src/Init/Data/OfScientific.lean b/stage0/src/Init/Data/OfScientific.lean index 91f409d75a9..f1105181075 100644 --- a/stage0/src/Init/Data/OfScientific.lean +++ b/stage0/src/Init/Data/OfScientific.lean @@ -23,6 +23,14 @@ def Float.ofBinaryScientific (m : Nat) (e : Int) : Float := let e := e + s m.toFloat.scaleB e +protected opaque Float.ofScientific (m : Nat) (s : Bool) (e : Nat) : Float := + if s then + let s := 64 - m.log2 -- ensure we have 64 bits of mantissa left after division + let m := (m <<< (3 * e + s)) / 5^e + Float.ofBinaryScientific m (-4 * e - s) + else + Float.ofBinaryScientific (m * 5^e) e + /- The `OfScientifi Float` must have priority higher than `mid` since the default instance `Neg Int` has `mid` priority. @@ -32,13 +40,7 @@ def Float.ofBinaryScientific (m : Nat) (e : Int) : Float := -/ @[defaultInstance mid+1] instance : OfScientific Float where - ofScientific m s e := - if s then - let s := 64 - m.log2 -- ensure we have 64 bits of mantissa left after division - let m := (m <<< (3 * e + s)) / 5^e - Float.ofBinaryScientific m (-4 * e - s) - else - Float.ofBinaryScientific (m * 5^e) e + ofScientific := Float.ofScientific @[export lean_float_of_nat] def Float.ofNat (n : Nat) : Float := diff --git a/stage0/src/Init/Data/Option/Basic.lean b/stage0/src/Init/Data/Option/Basic.lean index 8792e478306..895628da58a 100644 --- a/stage0/src/Init/Data/Option/Basic.lean +++ b/stage0/src/Init/Data/Option/Basic.lean @@ -34,9 +34,6 @@ def toMonad [Monad m] [Alternative m] : Option α → m α | none, _ => none | some a, b => b a -@[inline] protected def map (f : α → β) (o : Option α) : Option β := - Option.bind o (some ∘ f) - @[inline] protected def mapM [Monad m] (f : α → m β) (o : Option α) : m (Option β) := do if let some a := o then return some (← f a) diff --git a/stage0/src/Init/Data/Random.lean b/stage0/src/Init/Data/Random.lean index 42cec3efdd8..b00e3751635 100644 --- a/stage0/src/Init/Data/Random.lean +++ b/stage0/src/Init/Data/Random.lean @@ -112,7 +112,9 @@ def randBool {gen : Type u} [RandomGen gen] (g : gen) : Bool × gen := let (v, g') := randNat g 0 1 (v = 1, g') -initialize IO.stdGenRef : IO.Ref StdGen ← IO.mkRef mkStdGen +initialize IO.stdGenRef : IO.Ref StdGen ← + let seed := UInt64.toNat (ByteArray.toUInt64LE! (← IO.getRandomBytes 8)) + IO.mkRef (mkStdGen seed) def IO.setRandSeed (n : Nat) : IO Unit := IO.stdGenRef.set (mkStdGen n) diff --git a/stage0/src/Init/Data/String/Basic.lean b/stage0/src/Init/Data/String/Basic.lean index dbe5880441d..cf86c7a5e53 100644 --- a/stage0/src/Init/Data/String/Basic.lean +++ b/stage0/src/Init/Data/String/Basic.lean @@ -333,7 +333,7 @@ s.any (fun a => a == c) mapAux f 0 s def isNat (s : String) : Bool := - s.all fun c => c.isDigit + !s.isEmpty && s.all (·.isDigit) def toNat? (s : String) : Option Nat := if s.isNat then diff --git a/stage0/src/Init/Meta.lean b/stage0/src/Init/Meta.lean index 8b405562bd2..554089afe1c 100644 --- a/stage0/src/Init/Meta.lean +++ b/stage0/src/Init/Meta.lean @@ -7,6 +7,7 @@ Additional goodies for writing macros -/ prelude import Init.Data.Array.Basic +import Init.Data.Option.BasicAux namespace Lean @@ -169,6 +170,15 @@ def replacePrefix : Name → Name → Name → Name | n@(str p s _), queryP, newP => if n == queryP then newP else Name.mkStr (p.replacePrefix queryP newP) s | n@(num p s _), queryP, newP => if n == queryP then newP else Name.mkNum (p.replacePrefix queryP newP) s +/-- + `eraseSuffix? n s` return `n'` if `n` is of the form `n == n' ++ s`. +-/ +def eraseSuffix? : Name → Name → Option Name + | n, anonymous => some n + | str p s _, str p' s' _ => if s == s' then eraseSuffix? p p' else none + | num p s _, num p' s' _ => if s == s' then eraseSuffix? p p' else none + | _, _ => none + /-- Remove macros scopes, apply `f`, and put them back -/ @[inline] def modifyBase (n : Name) (f : Name → Name) : Name := if n.hasMacroScopes then @@ -236,6 +246,64 @@ instance monadNameGeneratorLift (m n : Type → Type) [MonadLift m n] [MonadName namespace Syntax +abbrev Term := TSyntax `term +abbrev Command := TSyntax `command +protected abbrev Level := TSyntax `level +abbrev Prec := TSyntax `prec +abbrev Prio := TSyntax `prio +abbrev Ident := TSyntax identKind +abbrev StrLit := TSyntax strLitKind +abbrev CharLit := TSyntax charLitKind +abbrev NameLit := TSyntax nameLitKind +abbrev NumLit := TSyntax numLitKind + +end Syntax + +export Syntax (Term Command Prec Prio Ident StrLit CharLit NameLit NumLit) + +namespace TSyntax + +instance : Coe (TSyntax [k]) (TSyntax (k :: ks)) where + coe stx := ⟨stx⟩ + +instance [Coe (TSyntax [k]) (TSyntax ks)] : Coe (TSyntax [k]) (TSyntax (k' :: ks)) where + coe stx := ⟨stx⟩ + +instance : Coe Ident Term where + coe s := ⟨s.raw⟩ + +instance : CoeDep Term ⟨Syntax.ident info ss n res⟩ Ident where + coe := ⟨Syntax.ident info ss n res⟩ + +instance : Coe StrLit Term where + coe s := ⟨s.raw⟩ + +instance : Coe NameLit Term where + coe s := ⟨s.raw⟩ + +instance : Coe NumLit Term where + coe s := ⟨s.raw⟩ + +instance : Coe CharLit Term where + coe s := ⟨s.raw⟩ + +instance : Coe NumLit Prec where + coe s := ⟨s.raw⟩ + +namespace Compat + +scoped instance : CoeTail Syntax (TSyntax k) where + coe s := ⟨s⟩ + +scoped instance : CoeTail (Array Syntax) (TSyntaxArray k) where + coe := .mk + +end Compat + +end TSyntax + +namespace Syntax + partial def structEq : Syntax → Syntax → Bool | Syntax.missing, Syntax.missing => true | Syntax.node _ k args, Syntax.node _ k' args' => k == k' && args.isEqv args' structEq @@ -244,6 +312,7 @@ partial def structEq : Syntax → Syntax → Bool | _, _ => false instance : BEq Lean.Syntax := ⟨structEq⟩ +instance : BEq (Lean.TSyntax k) := ⟨(·.raw == ·.raw)⟩ partial def getTailInfo? : Syntax → Option SourceInfo | atom info _ => info @@ -355,8 +424,8 @@ end Syntax | none => x | some ref => withRef ref x -@[inline] def mkNode (k : SyntaxNodeKind) (args : Array Syntax) : Syntax := - Syntax.node SourceInfo.none k args +@[inline] def mkNode (k : SyntaxNodeKind) (args : Array Syntax) : TSyntax k := + ⟨Syntax.node SourceInfo.none k args⟩ /- Syntax objects for a Lean module. -/ structure Module where @@ -378,30 +447,30 @@ partial def expandMacros : Syntax → MacroM Syntax /-- Create an identifier copying the position from `src`. To refer to a specific constant, use `mkCIdentFrom` instead. -/ -def mkIdentFrom (src : Syntax) (val : Name) : Syntax := - Syntax.ident (SourceInfo.fromRef src) (toString val).toSubstring val [] +def mkIdentFrom (src : Syntax) (val : Name) : Ident := + ⟨Syntax.ident (SourceInfo.fromRef src) (toString val).toSubstring val []⟩ -def mkIdentFromRef [Monad m] [MonadRef m] (val : Name) : m Syntax := do +def mkIdentFromRef [Monad m] [MonadRef m] (val : Name) : m Ident := do return mkIdentFrom (← getRef) val /-- Create an identifier referring to a constant `c` copying the position from `src`. This variant of `mkIdentFrom` makes sure that the identifier cannot accidentally be captured. -/ -def mkCIdentFrom (src : Syntax) (c : Name) : Syntax := +def mkCIdentFrom (src : Syntax) (c : Name) : Ident := -- Remark: We use the reserved macro scope to make sure there are no accidental collision with our frontend let id := addMacroScope `_internal c reservedMacroScope - Syntax.ident (SourceInfo.fromRef src) (toString id).toSubstring id [(c, [])] + ⟨Syntax.ident (SourceInfo.fromRef src) (toString id).toSubstring id [(c, [])]⟩ def mkCIdentFromRef [Monad m] [MonadRef m] (c : Name) : m Syntax := do return mkCIdentFrom (← getRef) c -def mkCIdent (c : Name) : Syntax := +def mkCIdent (c : Name) : Ident := mkCIdentFrom Syntax.missing c @[export lean_mk_syntax_ident] -def mkIdent (val : Name) : Syntax := - Syntax.ident SourceInfo.none (toString val).toSubstring val [] +def mkIdent (val : Name) : Ident := + ⟨Syntax.ident SourceInfo.none (toString val).toSubstring val []⟩ @[inline] def mkNullNode (args : Array Syntax := #[]) : Syntax := mkNode nullKind args @@ -440,31 +509,34 @@ def SepArray.ofElemsUsingRef [Monad m] [MonadRef m] {sep} (elems : Array Syntax) let ref ← getRef; return ⟨mkSepArray elems (if sep.isEmpty then mkNullNode else mkAtomFrom ref sep)⟩ -instance (sep) : Coe (Array Syntax) (SepArray sep) where +instance : Coe (Array Syntax) (SepArray sep) where coe := SepArray.ofElems +instance : Coe (TSyntaxArray k) (TSepArray k sep) where + coe a := ⟨mkSepArray a.raw (mkAtom sep)⟩ + /-- Create syntax representing a Lean term application, but avoid degenerate empty applications. -/ -def mkApp (fn : Syntax) : (args : Array Syntax) → Syntax +def mkApp (fn : Term) : (args : TSyntaxArray `term) → Term | #[] => fn - | args => mkNode `Lean.Parser.Term.app #[fn, mkNullNode args] + | args => ⟨mkNode `Lean.Parser.Term.app #[fn, mkNullNode args.raw]⟩ -def mkCApp (fn : Name) (args : Array Syntax) : Syntax := +def mkCApp (fn : Name) (args : TSyntaxArray `term) : Term := mkApp (mkCIdent fn) args -def mkLit (kind : SyntaxNodeKind) (val : String) (info := SourceInfo.none) : Syntax := +def mkLit (kind : SyntaxNodeKind) (val : String) (info := SourceInfo.none) : TSyntax kind := let atom : Syntax := Syntax.atom info val mkNode kind #[atom] -def mkStrLit (val : String) (info := SourceInfo.none) : Syntax := +def mkStrLit (val : String) (info := SourceInfo.none) : StrLit := mkLit strLitKind (String.quote val) info -def mkNumLit (val : String) (info := SourceInfo.none) : Syntax := +def mkNumLit (val : String) (info := SourceInfo.none) : NumLit := mkLit numLitKind val info -def mkScientificLit (val : String) (info := SourceInfo.none) : Syntax := +def mkScientificLit (val : String) (info := SourceInfo.none) : TSyntax scientificLitKind := mkLit scientificLitKind val info -def mkNameLit (val : String) (info := SourceInfo.none) : Syntax := +def mkNameLit (val : String) (info := SourceInfo.none) : NameLit := mkLit nameLitKind val info /- Recall that we don't have special Syntax constructors for storing numeric and string atoms. @@ -750,11 +822,6 @@ def isNone (stx : Syntax) : Bool := | Syntax.missing => true | _ => false -def getOptional? (stx : Syntax) : Option Syntax := - match stx with - | Syntax.node _ k args => if k == nullKind && args.size == 1 then some (args.get! 0) else none - | _ => none - def getOptionalIdent? (stx : Syntax) : Option Name := match stx.getOptional? with | some stx => some stx.getId @@ -769,16 +836,38 @@ def find? (stx : Syntax) (p : Syntax → Bool) : Option Syntax := end Syntax +namespace TSyntax + +def getNat (s : TSyntax numLitKind) : Nat := + s.raw.isNatLit?.get! + +def getId (s : Ident) : Name := + s.raw.getId + +def getString (s : TSyntax strLitKind) : String := + s.raw.isStrLit?.get! + +namespace Compat + +scoped instance : CoeTail (Array Syntax) (Syntax.TSepArray k sep) where + coe a := (a : TSyntaxArray k) + +end Compat + +end TSyntax + /-- Reflect a runtime datum back to surface syntax (best-effort). -/ -class Quote (α : Type) where - quote : α → Syntax +class Quote (α : Type) (k : SyntaxNodeKind := `term) where + quote : α → TSyntax k export Quote (quote) -instance : Quote Syntax := ⟨id⟩ +instance [Quote α k] [CoeHTCT (TSyntax k) (TSyntax [k'])]: Quote α k' := ⟨fun a => quote (k := k) a⟩ + +instance : Quote Term := ⟨id⟩ instance : Quote Bool := ⟨fun | true => mkCIdent `Bool.true | false => mkCIdent `Bool.false⟩ -instance : Quote String := ⟨Syntax.mkStrLit⟩ -instance : Quote Nat := ⟨fun n => Syntax.mkNumLit <| toString n⟩ +instance : Quote String strLitKind := ⟨Syntax.mkStrLit⟩ +instance : Quote Nat numLitKind := ⟨fun n => Syntax.mkNumLit <| toString n⟩ instance : Quote Substring := ⟨fun s => Syntax.mkCApp `String.toSubstring #[quote s.toString]⟩ -- in contrast to `Name.toString`, we can, and want to be, precise here @@ -789,43 +878,42 @@ private def getEscapedNameParts? (acc : List String) : Name → Option (List Str getEscapedNameParts? (s::acc) n | Name.num _ _ _ => none -private def quoteNameMk : Name → Syntax +def quoteNameMk : Name → Term | Name.anonymous => mkCIdent ``Name.anonymous | Name.str n s _ => Syntax.mkCApp ``Name.mkStr #[quoteNameMk n, quote s] | Name.num n i _ => Syntax.mkCApp ``Name.mkNum #[quoteNameMk n, quote i] -instance : Quote Name where +instance : Quote Name `term where quote n := match getEscapedNameParts? [] n with - | some ss => mkNode `Lean.Parser.Term.quotedName #[Syntax.mkNameLit ("`" ++ ".".intercalate ss)] - | none => quoteNameMk n + | some ss => ⟨mkNode `Lean.Parser.Term.quotedName #[Syntax.mkNameLit ("`" ++ ".".intercalate ss)]⟩ + | none => ⟨quoteNameMk n⟩ -instance {α β : Type} [Quote α] [Quote β] : Quote (α × β) where +instance [Quote α `term] [Quote β `term] : Quote (α × β) `term where quote | ⟨a, b⟩ => Syntax.mkCApp ``Prod.mk #[quote a, quote b] -private def quoteList {α : Type} [Quote α] : List α → Syntax +private def quoteList [Quote α `term] : List α → Term | [] => mkCIdent ``List.nil | (x::xs) => Syntax.mkCApp ``List.cons #[quote x, quoteList xs] -instance {α : Type} [Quote α] : Quote (List α) where +instance [Quote α `term] : Quote (List α) `term where quote := quoteList -instance {α : Type} [Quote α] : Quote (Array α) where +instance [Quote α `term] : Quote (Array α) `term where quote xs := Syntax.mkCApp ``List.toArray #[quote xs.toList] -private def quoteOption {α : Type} [Quote α] : Option α → Syntax - | none => mkIdent ``none - | (some x) => Syntax.mkCApp ``some #[quote x] +instance Option.hasQuote {α : Type} [Quote α `term] : Quote (Option α) `term where + quote + | none => mkIdent ``none + | (some x) => Syntax.mkCApp ``some #[quote x] -instance Option.hasQuote {α : Type} [Quote α] : Quote (Option α) where - quote := quoteOption /- Evaluator for `prec` DSL -/ def evalPrec (stx : Syntax) : MacroM Nat := Macro.withIncRecDepth stx do let stx ← expandMacros stx match stx with - | `(prec| $num:num) => return num.isNatLit?.getD 0 + | `(prec| $num:num) => return num.getNat | _ => Macro.throwErrorAt stx "unexpected precedence" macro_rules @@ -834,14 +922,14 @@ macro_rules macro_rules | `(prec| $a - $b) => do `(prec| $(quote <| (← evalPrec a) - (← evalPrec b)):num) -macro "eval_prec " p:prec:max : term => return quote (← evalPrec p) +macro "eval_prec " p:prec:max : term => return quote (k := `term) (← evalPrec p) /- Evaluator for `prio` DSL -/ def evalPrio (stx : Syntax) : MacroM Nat := Macro.withIncRecDepth stx do let stx ← expandMacros stx match stx with - | `(prio| $num:num) => return num.isNatLit?.getD 0 + | `(prio| $num:num) => return num.getNat | _ => Macro.throwErrorAt stx "unexpected priority" macro_rules @@ -850,9 +938,9 @@ macro_rules macro_rules | `(prio| $a - $b) => do `(prio| $(quote <| (← evalPrio a) - (← evalPrio b)):num) -macro "eval_prio " p:prio:max : term => return quote (← evalPrio p) +macro "eval_prio " p:prio:max : term => return quote (k := `term) (← evalPrio p) -def evalOptPrio : Option Syntax → MacroM Nat +def evalOptPrio : Option (TSyntax `prio) → MacroM Nat | some prio => evalPrio prio | none => return 1000 -- TODO: FIX back eval_prio default @@ -906,20 +994,38 @@ def mapSepElems (a : Array Syntax) (f : Syntax → Syntax) : Array Syntax := end Array -namespace Lean.Syntax.SepArray +namespace Lean.Syntax -def getElems {sep} (sa : SepArray sep) : Array Syntax := +def SepArray.getElems (sa : SepArray sep) : Array Syntax := sa.elemsAndSeps.getSepElems +def TSepArray.getElems (sa : TSepArray k sep) : TSyntaxArray k := + .mk sa.elemsAndSeps.getSepElems + /- We use `CoeTail` here instead of `Coe` to avoid a "loop" when computing `CoeTC`. The "loop" is interrupted using the maximum instance size threshold, but it is a performance bottleneck. The loop occurs because the predicate `isNewAnswer` is too imprecise. -/ -instance (sep) : CoeTail (SepArray sep) (Array Syntax) where - coe := getElems +instance : CoeTail (SepArray sep) (Array Syntax) where + coe := SepArray.getElems + +instance : Coe (TSepArray k sep) (TSyntaxArray k) where + coe := TSepArray.getElems + +instance [Coe (TSyntax k) (TSyntax k')] : Coe (TSyntaxArray k) (TSyntaxArray k') where + coe a := .mk a.raw -end Lean.Syntax.SepArray +instance : Coe (TSyntaxArray k) (Array Syntax) where + coe a := a.raw + +instance : Coe Ident (TSyntax `Lean.Parser.Command.declId) where + coe id := mkNode _ #[id, mkNullNode #[]] + +instance : Coe (Lean.Term) (Lean.TSyntax `Lean.Parser.Term.funBinder) where + coe stx := ⟨stx⟩ + +end Lean.Syntax set_option linter.unusedVariables.funArgs false in /-- @@ -961,6 +1067,13 @@ partial def isInterpolatedStrLit? (stx : Syntax) : Option String := | none => none | some val => decodeInterpStrLit val +def getSepArgs (stx : Syntax) : Array Syntax := + stx.getArgs.getSepElems + +end Syntax + +namespace TSyntax + def expandInterpolatedStrChunks (chunks : Array Syntax) (mkAppend : Syntax → Syntax → MacroM Syntax) (mkElem : Syntax → MacroM Syntax) : MacroM Syntax := do let mut i := 0 let mut result := Syntax.missing @@ -975,14 +1088,12 @@ def expandInterpolatedStrChunks (chunks : Array Syntax) (mkAppend : Syntax → S i := i+1 return result -def expandInterpolatedStr (interpStr : Syntax) (type : Syntax) (toTypeFn : Syntax) : MacroM Syntax := do - let r ← expandInterpolatedStrChunks interpStr.getArgs (fun a b => `($a ++ $b)) (fun a => `($toTypeFn $a)) +open TSyntax.Compat in +def expandInterpolatedStr (interpStr : TSyntax interpolatedStrKind) (type : Term) (toTypeFn : Term) : MacroM Term := do + let r ← expandInterpolatedStrChunks interpStr.raw.getArgs (fun a b => `($a ++ $b)) (fun a => `($toTypeFn $a)) `(($r : $type)) -def getSepArgs (stx : Syntax) : Array Syntax := - stx.getArgs.getSepElems - -end Syntax +end TSyntax namespace Meta @@ -1072,9 +1183,9 @@ syntax dsimpKind := atomic("(" &"dsimp") " := " &"true" ")" macro "declare_simp_like_tactic" opt:((simpAllKind <|> dsimpKind)?) tacName:ident tacToken:str updateCfg:term : command => do let (kind, tkn, stx) ← - if opt.isNone then + if opt.raw.isNone then pure (← `(``simp), ← `("simp "), ← `(syntax (name := $tacName:ident) $tacToken:str (config)? (discharger)? (&"only ")? ("[" (simpStar <|> simpErase <|> simpLemma),* "]")? (location)? : tactic)) - else if opt[0].getKind == ``simpAllKind then + else if opt.raw[0].getKind == ``simpAllKind then pure (← `(``simpAll), ← `("simp_all "), ← `(syntax (name := $tacName:ident) $tacToken:str (config)? (discharger)? (&"only ")? ("[" (simpErase <|> simpLemma),* "]")? : tactic)) else pure (← `(``dsimp), ← `("dsimp "), ← `(syntax (name := $tacName:ident) $tacToken:str (config)? (discharger)? (&"only ")? ("[" (simpErase <|> simpLemma),* "]")? (location)? : tactic)) diff --git a/stage0/src/Init/Notation.lean b/stage0/src/Init/Notation.lean index b9c4282f5e3..9ada407268c 100644 --- a/stage0/src/Init/Notation.lean +++ b/stage0/src/Init/Notation.lean @@ -22,8 +22,13 @@ syntax:65 (name := subPrio) prio " - " prio:66 : prio end Lean.Parser.Syntax namespace Lean -instance : Coe (TSyntax k) Syntax where + +instance : Coe (TSyntax ks) Syntax where coe stx := stx.raw + +instance : Coe SyntaxNodeKind SyntaxNodeKinds where + coe k := List.cons k List.nil + end Lean macro "max" : prec => `(1024) -- maximum precedence used in term parsers, in particular for terms in function position (`ident`, `paren`, ...) @@ -151,35 +156,35 @@ syntax (name := termDepIfThenElse) ppDedent(ppSpace) ppRealFill("else " term)) : term macro_rules - | `(if $h:ident : $c then $t:term else $e:term) => `(let_mvar% ?m := $c; wait_if_type_mvar% ?m; dite ?m (fun $h:ident => $t) (fun $h:ident => $e)) + | `(if $h : $c then $t else $e) => `(let_mvar% ?m := $c; wait_if_type_mvar% ?m; dite ?m (fun $h:ident => $t) (fun $h:ident => $e)) syntax (name := termIfThenElse) ppRealGroup(ppRealFill(ppIndent("if " term " then") ppSpace term) ppDedent(ppSpace) ppRealFill("else " term)) : term macro_rules - | `(if $c then $t:term else $e:term) => `(let_mvar% ?m := $c; wait_if_type_mvar% ?m; ite ?m $t $e) + | `(if $c then $t else $e) => `(let_mvar% ?m := $c; wait_if_type_mvar% ?m; ite ?m $t $e) macro "if " "let " pat:term " := " d:term " then " t:term " else " e:term : term => - `(match $d:term with | $pat:term => $t | _ => $e) + `(match $d:term with | $pat => $t | _ => $e) syntax (name := boolIfThenElse) ppRealGroup(ppRealFill(ppIndent("bif " term " then") ppSpace term) ppDedent(ppSpace) ppRealFill("else " term)) : term macro_rules - | `(bif $c then $t:term else $e:term) => `(cond $c $t $e) + | `(bif $c then $t else $e) => `(cond $c $t $e) syntax:min term " <| " term:min : term macro_rules - | `($f $args* <| $a) => let args := args.push a; `($f $args*) + | `($f $args* <| $a) => `($f $args* $a) | `($f <| $a) => `($f $a) syntax:min term " |> " term:min1 : term macro_rules - | `($a |> $f $args*) => let args := args.push a; `($f $args*) + | `($a |> $f $args*) => `($f $args* $a) | `($a |> $f) => `($f $a) -- Haskell-like pipe <| @@ -187,7 +192,7 @@ macro_rules syntax:min term atomic(" $" ws) term:min : term macro_rules - | `($f $args* $ $a) => let args := args.push a; `($f $args*) + | `($f $args* $ $a) => `($f $args* $a) | `($f $ $a) => `($f $a) syntax "{ " ident (" : " term)? " // " term " }" : term @@ -209,11 +214,12 @@ namespace Lean macro_rules | `([ $elems,* ]) => do - let rec expandListLit (i : Nat) (skip : Bool) (result : Syntax) : MacroM Syntax := do + -- NOTE: we do not have `TSepArray.getElems` yet at this point + let rec expandListLit (i : Nat) (skip : Bool) (result : TSyntax `term) : MacroM Syntax := do match i, skip with | 0, _ => pure result | i+1, true => expandListLit i false result - | i+1, false => expandListLit i true (← ``(List.cons $(elems.elemsAndSeps[i]) $result)) + | i+1, false => expandListLit i true (← ``(List.cons $(⟨elems.elemsAndSeps[i]⟩) $result)) if elems.elemsAndSeps.size < 64 then expandListLit elems.elemsAndSeps.size false (← ``(List.nil)) else @@ -228,3 +234,6 @@ macro tk:"this" : term => return Syntax.ident tk.getHeadInfo "this".toSubstring Category for carrying raw syntax trees between macros; any content is printed as is by the pretty printer. The only accepted parser for this category is an antiquotation. -/ declare_syntax_cat rawStx + +instance : Coe Syntax (TSyntax `rawStx) where + coe stx := ⟨stx⟩ diff --git a/stage0/src/Init/NotationExtra.lean b/stage0/src/Init/NotationExtra.lean index bdbb78b0a38..6b978e8012d 100644 --- a/stage0/src/Init/NotationExtra.lean +++ b/stage0/src/Init/NotationExtra.lean @@ -16,11 +16,11 @@ macro "Macro.trace[" id:ident "]" s:interpolatedStr(term) : term => -- Auxiliary parsers and functions for declaring notation with binders -syntax binderIdent := ident <|> "_" syntax unbracketedExplicitBinders := binderIdent+ (" : " term)? syntax bracketedExplicitBinders := "(" binderIdent+ " : " term ")" syntax explicitBinders := bracketedExplicitBinders+ <|> unbracketedExplicitBinders +open TSyntax.Compat in def expandExplicitBindersAux (combinator : Syntax) (idents : Array Syntax) (type? : Option Syntax) (body : Syntax) : MacroM Syntax := let rec loop (i : Nat) (acc : Syntax) := do match i with @@ -66,19 +66,13 @@ syntax unifConstraintElem := colGe unifConstraint ", "? syntax attrKind "unif_hint " (ident)? bracketedBinder* " where " withPosition(unifConstraintElem*) ("|-" <|> "⊢ ") unifConstraint : command -private def mkHintBody (cs : Array Syntax) (p : Syntax) : MacroM Syntax := do - let mut body ← `($(p[0]) = $(p[2])) - for c in cs.reverse do - body ← `($(c[0][0]) = $(c[0][2]) → $body) - return body - macro_rules - | `($kind:attrKind unif_hint $bs:explicitBinder* where $cs* |- $p) => do - let body ← mkHintBody cs p - `(@[$kind:attrKind unificationHint] def hint $bs:explicitBinder* : Sort _ := $body) - | `($kind:attrKind unif_hint $n:ident $bs* where $cs* |- $p) => do - let body ← mkHintBody cs p - `(@[$kind:attrKind unificationHint] def $n:ident $bs:explicitBinder* : Sort _ := $body) + | `($kind:attrKind unif_hint $(n)? $bs:bracketedBinder* where $[$cs₁:term ≟ $cs₂]* |- $t₁:term ≟ $t₂) => do + let mut body ← `($t₁ = $t₂) + for (c₁, c₂) in cs₁.zip cs₂ |>.reverse do + body ← `($c₁ = $c₂ → $body) + let hint : Ident ← `(hint) + `(@[$kind:attrKind unificationHint] def $(n.getD hint):ident $bs:bracketedBinder* : Sort _ := $body) end Lean open Lean @@ -94,7 +88,7 @@ macro:35 xs:bracketedExplicitBinders " ×' " b:term:35 : term => expandBrackedBi syntax calcStep := ppIndent(colGe term " := " withPosition(term)) syntax (name := calc) "calc" ppLine withPosition((calcStep ppLine)+) : term -macro "calc " steps:withPosition(calcStep+) : tactic => `(exact calc $(steps.getArgs)*) +macro "calc " steps:withPosition(calcStep+) : tactic => `(exact calc $steps*) @[appUnexpander Unit.unit] def unexpandUnit : Lean.PrettyPrinter.Unexpander | `($(_)) => `(()) @@ -154,6 +148,18 @@ macro "calc " steps:withPosition(calcStep+) : tactic => `(exact calc $(steps.get | `($(_) fun $x:ident => $p) => `({ $x // $p }) | _ => throw () +@[appUnexpander TSyntax] def unexpandTSyntax : Lean.PrettyPrinter.Unexpander + | `($f [$k]) => `($f $k) + | _ => throw () + +@[appUnexpander TSyntaxArray] def unexpandTSyntaxArray : Lean.PrettyPrinter.Unexpander + | `($f [$k]) => `($f $k) + | _ => throw () + +@[appUnexpander Syntax.TSepArray] def unexpandTSepArray : Lean.PrettyPrinter.Unexpander + | `($f [$k] $sep) => `($f $k $sep) + | _ => throw () + /-- Apply function extensionality and introduce new hypotheses. The tactic `funext` will keep applying new the `funext` lemma until the goal target is not reducible to @@ -170,16 +176,13 @@ Patterns can be used like in the `intro` tactic. Example, given a goal syntax "funext " (colGt term:max)+ : tactic macro_rules - | `(tactic|funext $xs*) => - if xs.size == 1 then - `(tactic| apply funext; intro $(xs[0]):term) - else - `(tactic| apply funext; intro $(xs[0]):term; funext $(xs[1:])*) + | `(tactic|funext $x) => `(tactic| apply funext; intro $x:term) + | `(tactic|funext $x $xs*) => `(tactic| apply funext; intro $x:term; funext $xs*) macro_rules | `(%[ $[$x],* | $k ]) => if x.size < 8 then - x.foldrM (init := k) fun x k => + x.foldrM (β := Term) (init := k) fun x k => `(List.cons $x $k) else let m := x.size / 2 @@ -203,9 +206,8 @@ syntax declModifiers "class " "abbrev " declId bracketedBinder* (":" term)? ":=" withPosition(group(colGe term ","?)*) : command macro_rules - | `($mods:declModifiers class abbrev $id $params* $[: $ty:term]? := $[ $parents:term $[,]? ]*) => - let name := id[0] - let ctor := mkIdentFrom name <| name.getId.modifyBase (. ++ `mk) + | `($mods:declModifiers class abbrev $id $params* $[: $ty]? := $[ $parents $[,]? ]*) => + let ctor := mkIdentFrom id <| id.raw[0].getId.modifyBase (. ++ `mk) `($mods:declModifiers class $id $params* extends $[$parents:term],* $[: $ty]? attribute [instance] $ctor) diff --git a/stage0/src/Init/Prelude.lean b/stage0/src/Init/Prelude.lean index 9f6b03973c8..7f46a72fda6 100644 --- a/stage0/src/Init/Prelude.lean +++ b/stage0/src/Init/Prelude.lean @@ -1065,6 +1065,10 @@ instance {α} : Inhabited (Option α) where | some x, _ => x | none, e => e +@[inline] protected def Option.map (f : α → β) : Option α → Option β + | some x => some (f x) + | none => none + inductive List (α : Type u) where | nil : List α | cons (head : α) (tail : List α) : List α @@ -1353,7 +1357,7 @@ def Array.sequenceMap {α : Type u} {β : Type v} {m : Type v → Type w} [Monad | 0 => pure bs | Nat.succ i' => Bind.bind (f (as.get ⟨j, hlt⟩)) fun b => loop i' (hAdd j 1) (bs.push b)) (fun _ => pure bs) - loop as.size 0 Array.empty + loop as.size 0 (Array.mkEmpty as.size) /-- A Function for lifting a computation from an inner Monad to an outer Monad. Like [MonadTrans](https://hackage.haskell.org/package/transformers-0.5.5.0/docs/Control-Monad-Trans-Class.html), @@ -1843,25 +1847,34 @@ inductive Syntax where def SyntaxNodeKinds := List SyntaxNodeKind +/-- + A `Syntax` value of one of the given syntax kinds. + Note that while syntax quotations produce/expect `TSyntax` values of the correct kinds, + this is not otherwise enforced and can easily be circumvented by direct use of the constructor. + The namespace `TSyntax.Compat` can be opened to expose a general coercion from `Syntax` to any + `TSyntax ks` for porting older code. -/ structure TSyntax (ks : SyntaxNodeKinds) where raw : Syntax instance : Inhabited Syntax where default := Syntax.missing +instance : Inhabited (TSyntax ks) where + default := ⟨default⟩ + /- Builtin kinds -/ -def choiceKind : SyntaxNodeKind := `choice -def nullKind : SyntaxNodeKind := `null -def groupKind : SyntaxNodeKind := `group -def identKind : SyntaxNodeKind := `ident -def strLitKind : SyntaxNodeKind := `str -def charLitKind : SyntaxNodeKind := `char -def numLitKind : SyntaxNodeKind := `num -def scientificLitKind : SyntaxNodeKind := `scientific -def nameLitKind : SyntaxNodeKind := `name -def fieldIdxKind : SyntaxNodeKind := `fieldIdx -def interpolatedStrLitKind : SyntaxNodeKind := `interpolatedStrLitKind -def interpolatedStrKind : SyntaxNodeKind := `interpolatedStrKind +abbrev choiceKind : SyntaxNodeKind := `choice +abbrev nullKind : SyntaxNodeKind := `null +abbrev groupKind : SyntaxNodeKind := `group +abbrev identKind : SyntaxNodeKind := `ident +abbrev strLitKind : SyntaxNodeKind := `str +abbrev charLitKind : SyntaxNodeKind := `char +abbrev numLitKind : SyntaxNodeKind := `num +abbrev scientificLitKind : SyntaxNodeKind := `scientific +abbrev nameLitKind : SyntaxNodeKind := `name +abbrev fieldIdxKind : SyntaxNodeKind := `fieldIdx +abbrev interpolatedStrLitKind : SyntaxNodeKind := `interpolatedStrLitKind +abbrev interpolatedStrKind : SyntaxNodeKind := `interpolatedStrKind namespace Syntax @@ -1902,6 +1915,13 @@ def getNumArgs (stx : Syntax) : Nat := | Syntax.node _ _ args => args.size | _ => 0 +def getOptional? (stx : Syntax) : Option Syntax := + match stx with + | Syntax.node _ k args => match and (beq k nullKind) (beq args.size 1) with + | true => some (args.get! 0) + | false => none + | _ => none + def isMissing : Syntax → Bool | Syntax.missing => true | _ => false @@ -1975,6 +1995,7 @@ partial def getTailPos? (stx : Syntax) (originalOnly := false) : Option String.P structure SepArray (sep : String) where elemsAndSeps : Array Syntax +/-- A typed version of `SepArray`. -/ structure TSepArray (ks : SyntaxNodeKinds) (sep : String) where elemsAndSeps : Array Syntax diff --git a/stage0/src/Init/SimpLemmas.lean b/stage0/src/Init/SimpLemmas.lean index 15d24bcaac6..3d1cecdf326 100644 --- a/stage0/src/Init/SimpLemmas.lean +++ b/stage0/src/Init/SimpLemmas.lean @@ -156,6 +156,9 @@ theorem dite_congr {s : Decidable b} [Decidable c] @[simp] theorem beq_self_eq_true [BEq α] [LawfulBEq α] (a : α) : (a == a) = true := LawfulBEq.rfl @[simp] theorem beq_self_eq_true' [DecidableEq α] (a : α) : (a == a) = true := by simp [BEq.beq] +@[simp] theorem bne_self_eq_false [BEq α] [LawfulBEq α] (a : α) : (a != a) = false := by simp [bne] +@[simp] theorem bne_self_eq_false' [DecidableEq α] (a : α) : (a != a) = false := by simp [bne] + @[simp] theorem Nat.le_zero_eq (a : Nat) : (a ≤ 0) = (a = 0) := propext <| Iff.intro (fun h => Nat.le_antisymm h (Nat.zero_le ..)) (fun h => by simp [h]) diff --git a/stage0/src/Init/System/IO.lean b/stage0/src/Init/System/IO.lean index b58810557b0..69c437031fb 100644 --- a/stage0/src/Init/System/IO.lean +++ b/stage0/src/Init/System/IO.lean @@ -695,8 +695,5 @@ end Lean syntax "println! " (interpolatedStr(term) <|> term) : term macro_rules - | `(println! $msg) => - if msg.getKind == Lean.interpolatedStrKind then - `((IO.println (s! $msg) : IO Unit)) - else - `((IO.println $msg : IO Unit)) + | `(println! $msg:interpolatedStr) => `((IO.println (s! $msg) : IO Unit)) + | `(println! $msg:term) => `((IO.println $msg : IO Unit)) diff --git a/stage0/src/Init/Tactics.lean b/stage0/src/Init/Tactics.lean index b959d68fcfe..074f1988f88 100644 --- a/stage0/src/Init/Tactics.lean +++ b/stage0/src/Init/Tactics.lean @@ -8,6 +8,8 @@ import Init.Notation namespace Lean +syntax binderIdent := ident <|> "_" + namespace Parser.Tactic /-- `with_annotate_state stx t` annotates the lexical range of `stx : Syntax` with the initial and final state of running tactic `t`. -/ scoped syntax (name := withAnnotateState) "with_annotate_state " rawStx ppSpace tactic : tactic @@ -103,18 +105,18 @@ syntax (name := constructor) "constructor" : tactic /-- `case tag => tac` focuses on the goal with case name `tag` and solves it using `tac`, or else fails. `case tag x₁ ... xₙ => tac` additionally renames the `n` most recent hypotheses with inaccessible names to the given names. -/ -syntax (name := case) "case " (ident <|> "_") (ident <|> "_")* " => " tacticSeq : tactic +syntax (name := case) "case " binderIdent binderIdent* " => " tacticSeq : tactic /-- `case'` is similar to the `case tag => tac` tactic, but does not ensure the goal has been solved after applying `tac`, nor admits the goal if `tac` failed. Recall that `case` closes the goal using `sorry` when `tac` fails, and the tactic execution is not interrupted. -/ -syntax (name := case') "case' " (ident <|> "_") (ident <|> "_")* " => " tacticSeq : tactic +syntax (name := case') "case' " binderIdent binderIdent* " => " tacticSeq : tactic /-- `next => tac` focuses on the next goal solves it using `tac`, or else fails. `next x₁ ... xₙ => tac` additionally renames the `n` most recent hypotheses with inaccessible names to the given names. -/ -macro "next " args:(ident <|> "_")* " => " tac:tacticSeq : tactic => `(tactic| case _ $args* => $tac) +macro "next " args:binderIdent* " => " tac:tacticSeq : tactic => `(tactic| case _ $args* => $tac) /-- `all_goals tac` runs `tac` on each goal, concatenating the resulting goals, if any. -/ syntax (name := allGoals) "all_goals " tacticSeq : tactic @@ -201,19 +203,12 @@ syntax (name := rewriteSeq) "rewrite " (config)? rwRuleSeq (location)? : tactic /-- `rw` is like `rewrite`, but also tries to close the goal by "cheap" (reducible) `rfl` afterwards. -/ -syntax (name := rwSeq) "rw " (config)? rwRuleSeq (location)? : tactic - -def rwWithRfl (kind : SyntaxNodeKind) (atom : String) (stx : Syntax) : MacroM Syntax := do - -- We show the `rfl` state on `]` - let seq := stx[2] - let rbrak := seq[2] - -- Replace `]` token with one without position information in the expanded tactic - let seq := seq.setArg 2 (mkAtom "]") - let tac := stx.setKind kind |>.setArg 0 (mkAtomFrom stx atom) |>.setArg 2 seq - `(tactic| $tac; try (with_reducible rfl%$rbrak)) - -@[macro rwSeq] def expandRwSeq : Macro := - rwWithRfl ``Lean.Parser.Tactic.rewriteSeq "rewrite" +macro (name := rwSeq) rw:"rw " c:(config)? s:rwRuleSeq l:(location)? : tactic => + match s with + | `(rwRuleSeq| [%$lbrak $rs:rwRule,* ]%$rbrak) => + -- We show the `rfl` state on `]` + `(tactic| rewrite%$rw $(c)? [%$lbrak $rs,*] $(l)?; try (with_reducible rfl%$rbrak)) + | _ => Macro.throwUnsupported /-- The `injection` tactic is based on the fact that constructors of inductive data types are injections. @@ -285,7 +280,8 @@ macro "have " d:haveDecl : tactic => `(refine_lift have $d:haveDecl; ?_) /-- `have h := e` adds the hypothesis `h : t` if `e : t`. -/ -macro (priority := high) "have" x:ident " := " p:term : tactic => `(have $x:ident : _ := $p) +-- TODO: `DecidableEq` derive handler depends on the old name, see if we can get rid of it in a future stage +macro (name := «tacticHave__:=_») (priority := high) "have" x:ident " := " p:term : tactic => `(have $x:ident : _ := $p) /-- Given a main goal `ctx |- t`, `suffices h : t' from e` replaces the main goal with `ctx |- t'`, `e` must have type `t` in the context `ctx, h : t'`. @@ -356,7 +352,7 @@ and one goal with hypothesis `h : P (Nat.succ a)` and target `Q (Nat.succ a)`. H syntax (name := cases) "cases " casesTarget,+ (" using " ident)? (inductionAlts)? : tactic /-- `rename_i x_1 ... x_n` renames the last `n` inaccessible names using the given names. -/ -syntax (name := renameI) "rename_i " (colGt (ident <|> "_"))+ : tactic +syntax (name := renameI) "rename_i " (colGt binderIdent)+ : tactic /-- `repeat tac` applies `tac` to main goal. If the application succeeds, diff --git a/stage0/src/Lean/Elab/App.lean b/stage0/src/Lean/Elab/App.lean index c5af25a8784..ca43debd920 100644 --- a/stage0/src/Lean/Elab/App.lean +++ b/stage0/src/Lean/Elab/App.lean @@ -401,7 +401,7 @@ mutual | Except.error err => throwError err | Except.ok tacticSyntax => -- TODO(Leo): does this work correctly for tactic sequences? - let tacticBlock ← `(by $tacticSyntax) + let tacticBlock ← `(by $(⟨tacticSyntax⟩)) let argNew := Arg.stx tacticBlock propagateExpectedType argNew elabAndAddNewArg argName argNew @@ -527,7 +527,8 @@ private def throwLValError (e : Expr) (eType : Expr) (msg : MessageData) : TermE /-- `findMethod? env S fName`. 1- If `env` contains `S ++ fName`, return `(S, S++fName)` 2- Otherwise if `env` contains private name `prv` for `S ++ fName`, return `(S, prv)`, o - 3- Otherwise for each parent structure `S'` of `S`, we try `findMethod? env S' fname` -/ + 3- Otherwise for each parent structure `S'` of `S`, we try `findMethod? env S' fname` +-/ private partial def findMethod? (env : Environment) (structName fieldName : Name) : Option (Name × Name) := let fullName := structName ++ fieldName match env.find? fullName with @@ -542,6 +543,24 @@ private partial def findMethod? (env : Environment) (structName fieldName : Name else none +/-- + Return `some (structName', fullName)` if `structName ++ fieldName` is an alias for `fullName`, and + `fullName` is of the form `structName' ++ fieldName`. + + TODO: if there is more than one applicable alias, it returns `none`. We should consider throwing an error or + warning. +-/ +private def findMethodAlias? (env : Environment) (structName fieldName : Name) : Option (Name × Name) := + let fullName := structName ++ fieldName + -- We never skip `protected` aliases when resolving dot-notation. + let aliasesCandidates := getAliases env fullName (skipProtected := false) |>.filterMap fun alias => + match alias.eraseSuffix? fieldName with + | none => none + | some structName' => some (structName', alias) + match aliasesCandidates with + | [r] => some r + | _ => none + private def throwInvalidFieldNotation (e eType : Expr) : TermElabM α := throwLValError e eType "invalid field notation, type is not of the form (C ...) where C is a constant" @@ -574,9 +593,11 @@ private def resolveLValAux (e : Expr) (eType : Expr) (lval : LVal) : TermElabM L | some structName, LVal.fieldName _ fieldName _ _ => let env ← getEnv let searchEnv : Unit → TermElabM LValResolution := fun _ => do - match findMethod? env structName (Name.mkSimple fieldName) with - | some (baseStructName, fullName) => return LValResolution.const baseStructName structName fullName - | none => + if let some (baseStructName, fullName) := findMethod? env structName fieldName then + return LValResolution.const baseStructName structName fullName + else if let some (structName', fullName) := findMethodAlias? env structName fieldName then + return LValResolution.const structName' structName' fullName + else throwLValError e eType m!"invalid field '{fieldName}', the environment does not contain '{Name.mkStr structName fieldName}'" -- search local context first, then environment diff --git a/stage0/src/Lean/Elab/AuxDef.lean b/stage0/src/Lean/Elab/AuxDef.lean index f26d391aa28..d662cdab6ad 100644 --- a/stage0/src/Lean/Elab/AuxDef.lean +++ b/stage0/src/Lean/Elab/AuxDef.lean @@ -18,7 +18,7 @@ scoped syntax (name := aux_def) docComment ? attributes ? "aux_def" ident+ ":" t @[builtinCommandElab «aux_def»] def elabAuxDef : CommandElab - | `($[$doc?:docComment]? $[$attrs?:attributes]? aux_def $[$suggestion:ident]* : $ty := $body) => do + | `($[$doc?:docComment]? $[$attrs?:attributes]? aux_def $suggestion* : $ty := $body) => do let id := suggestion.map (·.getId.eraseMacroScopes) |>.foldl (· ++ ·) Name.anonymous let id := `_aux ++ (← getMainModule) ++ `_ ++ id let id := String.intercalate "_" <| id.components.map (·.toString (escape := false)) diff --git a/stage0/src/Lean/Elab/AuxDiscr.lean b/stage0/src/Lean/Elab/AuxDiscr.lean index 0e9e1051c20..c07326fac38 100644 --- a/stage0/src/Lean/Elab/AuxDiscr.lean +++ b/stage0/src/Lean/Elab/AuxDiscr.lean @@ -5,18 +5,19 @@ Authors: Leonardo de Moura -/ namespace Lean + /-- Create an auxiliary identifier for storing non-atomic discriminants. This kind of free variable is cleared before elaborating a `match` alternative rhs. -/ -def mkAuxDiscr [Monad m] [MonadQuotation m] : m Syntax := +def mkAuxDiscr [Monad m] [MonadQuotation m] : m Ident := `(_discr) /-- Create an auxiliary identifier for expanding notation such as `fun (a, b) => ...`. This kind of free variable is cleared before elaborating a `match` alternative rhs. -/ -def mkAuxFunDiscr [Monad m] [MonadQuotation m] : m Syntax := +def mkAuxFunDiscr [Monad m] [MonadQuotation m] : m Ident := `(_fun_discr) /-- Return true iff `n` is an auxiliary variable created by `expandNonAtomicDiscrs?` -/ diff --git a/stage0/src/Lean/Elab/Binders.lean b/stage0/src/Lean/Elab/Binders.lean index f09d3834bb9..98a009eaab5 100644 --- a/stage0/src/Lean/Elab/Binders.lean +++ b/stage0/src/Lean/Elab/Binders.lean @@ -12,6 +12,7 @@ import Lean.Elab.AuxDiscr namespace Lean.Elab.Term open Meta open Lean.Parser.Term +open TSyntax.Compat /-- Given syntax of the forms @@ -244,7 +245,7 @@ in the literature. -/ Auxiliary functions for converting `id_1 ... id_n` application into `#[id_1, ..., id_m]` It is used at `expandFunBinders`. -/ private partial def getFunBinderIds? (stx : Syntax) : OptionT MacroM (Array Syntax) := - let convertElem (stx : Syntax) : OptionT MacroM Syntax := + let convertElem (stx : Term) : OptionT MacroM Syntax := match stx with | `(_) => do let ident ← mkFreshIdent stx; pure ident | `($id:ident) => return id diff --git a/stage0/src/Lean/Elab/BindersUtil.lean b/stage0/src/Lean/Elab/BindersUtil.lean index fe31dca2099..2c08ecac7a8 100644 --- a/stage0/src/Lean/Elab/BindersUtil.lean +++ b/stage0/src/Lean/Elab/BindersUtil.lean @@ -30,24 +30,24 @@ def getMatchAltsNumPatterns (matchAlts : Syntax) : Nat := /-- Expand a match alternative such as `| 0 | 1 => rhs` to an array containing `| 0 => rhs` and `| 1 => rhs`. -/ -def expandMatchAlt (matchAlt : Syntax) : Array Syntax := - let patss := matchAlt[1] - if patss.getArgs.size ≤ 1 then - #[matchAlt] - else - patss.getSepArgs.map fun pats => - let patss := mkNullNode #[pats] - matchAlt.setArg 1 patss +def expandMatchAlt (stx : TSyntax ``matchAlt) : MacroM (Array (TSyntax ``matchAlt)) := + match stx with + | `(matchAltExpr| | $[$patss,*]|* => $rhs) => + if patss.size ≤ 1 then + return #[stx] + else + patss.mapM fun pats => `(matchAltExpr| | $pats,* => $rhs) + | _ => return #[stx] -def shouldExpandMatchAlt (matchAlt : Syntax) : Bool := - let patss := matchAlt[1] - patss.getArgs.size > 1 +def shouldExpandMatchAlt : TSyntax ``matchAlt → Bool + | `(matchAltExpr| | $[$patss,*]|* => $_) => patss.size > 1 + | _ => false def expandMatchAlts? (stx : Syntax) : MacroM (Option Syntax) := do match stx with | `(match $[$gen]? $[$motive]? $discrs,* with $alts:matchAlt*) => if alts.any shouldExpandMatchAlt then - let alts := alts.foldl (init := #[]) fun alts alt => alts ++ expandMatchAlt alt + let alts ← alts.foldlM (init := #[]) fun alts alt => return alts ++ (← expandMatchAlt alt) `(match $[$gen]? $[$motive]? $discrs,* with $alts:matchAlt*) else return none diff --git a/stage0/src/Lean/Elab/BuiltinCommand.lean b/stage0/src/Lean/Elab/BuiltinCommand.lean index acaea509d34..6db18e41828 100644 --- a/stage0/src/Lean/Elab/BuiltinCommand.lean +++ b/stage0/src/Lean/Elab/BuiltinCommand.lean @@ -172,14 +172,14 @@ private def matchBinderNames (ids : Array Syntax) (binderNames : Array Name) : C variable {α} -- trying to update part of the binder block defined above. ``` -/ -private def replaceBinderAnnotation (binder : Syntax) : CommandElabM Bool := do +private def replaceBinderAnnotation (binder : TSyntax ``Parser.Term.bracketedBinder) : CommandElabM Bool := do if let some (binderNames, explicit) := typelessBinder? binder then let varDecls := (← getScope).varDecls let mut varDeclsNew := #[] let mut found := false for varDecl in varDecls do if let some (ids, ty?, annot?) := - match varDecl with + match (varDecl : TSyntax ``Parser.Term.bracketedBinder) with | `(bracketedBinder|($ids* $[: $ty?]? $(annot?)?)) => some (ids, ty?, annot?) | `(bracketedBinder|{$ids* $[: $ty?]?}) => some (ids, ty?, none) | `(bracketedBinder|[$id : $ty]) => some (#[id], some ty, none) @@ -373,9 +373,8 @@ opaque elabEval : CommandElab modify fun s => { s with maxRecDepth := maxRecDepth.get options } modifyScope fun scope => { scope with opts := options } -@[builtinMacro Lean.Parser.Command.«in»] def expandInCmd : Macro := fun stx => do - let cmd₁ := stx[0] - let cmd₂ := stx[2] - `(section $cmd₁:command $cmd₂:command end) +@[builtinMacro Lean.Parser.Command.«in»] def expandInCmd : Macro + | `($cmd₁ in $cmd₂) => `(section $cmd₁:command $cmd₂ end) + | _ => Macro.throwUnsupported end Lean.Elab.Command diff --git a/stage0/src/Lean/Elab/BuiltinNotation.lean b/stage0/src/Lean/Elab/BuiltinNotation.lean index dd33fde13e2..96717f897a7 100644 --- a/stage0/src/Lean/Elab/BuiltinNotation.lean +++ b/stage0/src/Lean/Elab/BuiltinNotation.lean @@ -95,7 +95,7 @@ are turned into a new anonymous constructor application. For example, | _ => Macro.throwUnsupported open Lean.Parser in -private def elabParserMacroAux (prec : Syntax) (e : Syntax) (withAnonymousAntiquot : Bool) : TermElabM Syntax := do +private def elabParserMacroAux (prec e : Term) (withAnonymousAntiquot : Bool) : TermElabM Syntax := do let (some declName) ← getDeclName? | throwError "invalid `leading_parser` macro, it must be used in definitions" match extractMacroScopes declName with @@ -108,10 +108,10 @@ private def elabParserMacroAux (prec : Syntax) (e : Syntax) (withAnonymousAntiqu @[builtinTermElab «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 (·.isOfKind ``Parser.Term.trueVal)) + elabParserMacroAux (prec?.getD (quote Parser.maxPrec)) e (anon?.all (·.raw.isOfKind ``Parser.Term.trueVal)) | _ => throwUnsupportedSyntax -private def elabTParserMacroAux (prec lhsPrec : Syntax) (e : Syntax) : TermElabM Syntax := do +private def elabTParserMacroAux (prec lhsPrec e : Term) : TermElabM Syntax := do let declName? ← getDeclName? match declName? with | some declName => let kind := quote declName; ``(Lean.Parser.trailingNode $kind $prec $lhsPrec $e) @@ -131,37 +131,36 @@ function `lean_set_panic_messages(false)` has been executed before. If the C function `lean_set_exit_on_panic(true)` has been executed before, the process is then aborted. -/ @[builtinTermElab Lean.Parser.Term.panic] def elabPanic : TermElab := fun stx expectedType? => do - let arg := stx[1] - let pos ← getRefPosition - let env ← getEnv - let stxNew ← match (← getDeclName?) with - | some declName => `(panicWithPosWithDecl $(quote (toString env.mainModule)) $(quote (toString declName)) $(quote pos.line) $(quote pos.column) $arg) - | none => `(panicWithPos $(quote (toString env.mainModule)) $(quote pos.line) $(quote pos.column) $arg) - withMacroExpansion stx stxNew $ elabTerm stxNew expectedType? + match stx with + | `(panic! $arg) => + let pos ← getRefPosition + let env ← getEnv + let stxNew ← match (← getDeclName?) with + | some declName => `(panicWithPosWithDecl $(quote (toString env.mainModule)) $(quote (toString declName)) $(quote pos.line) $(quote pos.column) $arg) + | none => `(panicWithPos $(quote (toString env.mainModule)) $(quote pos.line) $(quote pos.column) $arg) + withMacroExpansion stx stxNew $ elabTerm stxNew expectedType? + | _ => throwUnsupportedSyntax /-- A shorthand for `panic! "unreachable code has been reached"`. -/ @[builtinMacro Lean.Parser.Term.unreachable] def expandUnreachable : Macro := fun _ => `(panic! "unreachable code has been reached") /-- `assert! cond` panics if `cond` evaluates to `false`. -/ -@[builtinMacro Lean.Parser.Term.assert] def expandAssert : Macro := fun stx => - -- TODO: support for disabling runtime assertions - let cond := stx[1] - let body := stx[3] - match cond.reprint with - | some code => `(if $cond then $body else panic! ("assertion violation: " ++ $(quote code))) - | none => `(if $cond then $body else panic! ("assertion violation")) +@[builtinMacro Lean.Parser.Term.assert] def expandAssert : Macro + | `(assert! $cond; $body) => + -- TODO: support for disabling runtime assertions + match cond.raw.reprint with + | some code => `(if $cond then $body else panic! ("assertion violation: " ++ $(quote code))) + | none => `(if $cond then $body else panic! ("assertion violation")) + | _ => Macro.throwUnsupported /-- `dbg_trace e; body` evaluates to `body` and prints `e` (which can be an interpolated string literal) to stderr. It should only be used for debugging. -/ -@[builtinMacro Lean.Parser.Term.dbgTrace] def expandDbgTrace : Macro := fun stx => - let arg := stx[1] - let body := stx[3] - if arg.getKind == interpolatedStrKind then - `(dbgTrace (s! $arg) fun _ => $body) - else - `(dbgTrace (toString $arg) fun _ => $body) +@[builtinMacro Lean.Parser.Term.dbgTrace] def expandDbgTrace : Macro + | `(dbg_trace $arg:interpolatedStr; $body) => `(dbgTrace (s! $arg) fun _ => $body) + | `(dbg_trace $arg:term; $body) => `(dbgTrace (toString $arg) fun _ => $body) + | _ => Macro.throwUnsupported /-- A temporary placeholder for a missing proof or value. -/ @[builtinTermElab «sorry»] def elabSorry : TermElab := fun stx expectedType? => do @@ -169,8 +168,8 @@ interpolated string literal) to stderr. It should only be used for debugging. -/ withMacroExpansion stx stxNew <| elabTerm stxNew expectedType? /-- Return syntax `Prod.mk elems[0] (Prod.mk elems[1] ... (Prod.mk elems[elems.size - 2] elems[elems.size - 1])))` -/ -partial def mkPairs (elems : Array Syntax) : MacroM Syntax := - let rec loop (i : Nat) (acc : Syntax) := do +partial def mkPairs (elems : Array Term) : MacroM Term := + let rec loop (i : Nat) (acc : Term) := do if i > 0 then let i := i - 1 let elem := elems[i] @@ -193,10 +192,10 @@ private partial def hasCDot : Syntax → Bool Examples: - `· + 1` => `fun _a_1 => _a_1 + 1` - `f · · b` => `fun _a_1 _a_2 => f _a_1 _a_2 b` -/ -partial def expandCDot? (stx : Syntax) : MacroM (Option Syntax) := do +partial def expandCDot? (stx : Term) : MacroM (Option Term) := do if hasCDot stx then - let (newStx, binders) ← (go stx).run #[]; - `(fun $binders* => $newStx) + let (newStx, binders) ← (go stx).run #[] + `(fun $binders* => $(⟨newStx⟩)) else pure none where @@ -205,23 +204,23 @@ where The extra state `Array Syntax` contains the new binder names. If `stx` is a `·`, we create a fresh identifier, store in the extra state, and return it. Otherwise, we just return `stx`. -/ - go : Syntax → StateT (Array Syntax) MacroM Syntax - | stx@(Syntax.node i k args) => - if k == ``Lean.Parser.Term.paren then pure stx - else if k == ``Lean.Parser.Term.cdot then withFreshMacroScope do - let id ← `(a) - modify fun s => s.push id; - pure id - else do + go : Syntax → StateT (Array Ident) MacroM Syntax + | stx@`(($(_))) => pure stx + | stx@`(·) => withFreshMacroScope do + let id : Ident ← `(a) + modify fun s => s.push id + pure id + | stx => match stx with + | .node i k args => do let args ← args.mapM go pure $ Syntax.node i k args - | stx => pure stx + | _ => pure stx /-- Helper method for elaborating terms such as `(.+.)` where a constant name is expected. This method is usually used to implement tactics that function names as arguments (e.g., `simp`). -/ -def elabCDotFunctionAlias? (stx : Syntax) : TermElabM (Option Expr) := do +def elabCDotFunctionAlias? (stx : Term) : TermElabM (Option Expr) := do let some stx ← liftMacroM <| expandCDotArg? stx | pure none let stx ← liftMacroM <| expandMacros stx match stx with @@ -237,7 +236,7 @@ def elabCDotFunctionAlias? (stx : Syntax) : TermElabM (Option Expr) := do return none | _ => return none where - expandCDotArg? (stx : Syntax) : MacroM (Option Syntax) := + expandCDotArg? (stx : Term) : MacroM (Option Term) := match stx with | `(($e)) => Term.expandCDot? e | _ => Term.expandCDot? stx @@ -264,7 +263,7 @@ where | stx => if !stx[1][0].isMissing && stx[1][1].isMissing then -- parsed `(` and `term`, assume it's a basic parenthesis to get any elaboration output at all - `(($(stx[1][0]))) + `(($(⟨stx[1][0]⟩))) else throw <| Macro.Exception.error stx "unexpected parentheses notation" @@ -287,7 +286,7 @@ private def isSubstCandidate (lhs rhs : Expr) : MetaM Bool := Given an expression `e` that is the elaboration of `stx`, if `e` is a free variable, then return `k stx`. Otherwise, return `(fun x => k x) e` -/ -private def withLocalIdentFor (stx : Syntax) (e : Expr) (k : Syntax → TermElabM Expr) : TermElabM Expr := do +private def withLocalIdentFor (stx : Term) (e : Expr) (k : Term → TermElabM Expr) : TermElabM Expr := do if e.isFVar then k stx else diff --git a/stage0/src/Lean/Elab/Command.lean b/stage0/src/Lean/Elab/Command.lean index 1befba23f8d..65a29672e85 100644 --- a/stage0/src/Lean/Elab/Command.lean +++ b/stage0/src/Lean/Elab/Command.lean @@ -23,8 +23,8 @@ structure Scope where currNamespace : Name := Name.anonymous openDecls : List OpenDecl := [] levelNames : List Name := [] - /-- section variables as `bracketedBinder`s -/ - varDecls : Array Syntax := #[] + /-- section variables -/ + varDecls : Array (TSyntax ``Parser.Term.bracketedBinder) := #[] /-- Globally unique internal identifiers for the `varDecls` -/ varUIds : Array Name := #[] /-- noncomputable sections automatically add the `noncomputable` modifier to any declaration we cannot generate code for. -/ diff --git a/stage0/src/Lean/Elab/Declaration.lean b/stage0/src/Lean/Elab/Declaration.lean index c1deaf80325..13ff5c20b67 100644 --- a/stage0/src/Lean/Elab/Declaration.lean +++ b/stage0/src/Lean/Elab/Declaration.lean @@ -13,6 +13,7 @@ import Lean.Elab.DeclarationRange namespace Lean.Elab.Command open Meta +open TSyntax.Compat private def ensureValidNamespace (name : Name) : MacroM Unit := do match name with @@ -37,7 +38,7 @@ private def expandDeclIdNamespace? (declId : Syntax) : MacroM (Option (Name × S let nameNew := { scpView with name := Name.mkSimple s }.review -- preserve "original" info, if any, so that hover etc. on the namespaced -- name access the info tree node of the declaration name - let id := mkIdent nameNew |>.setInfo declId.getHeadInfo + let id := mkIdent nameNew |>.raw.setInfo declId.getHeadInfo if declId.isIdent then return some (pre, id) else diff --git a/stage0/src/Lean/Elab/Deriving/BEq.lean b/stage0/src/Lean/Elab/Deriving/BEq.lean index 56c02dd72c0..91d896570fe 100644 --- a/stage0/src/Lean/Elab/Deriving/BEq.lean +++ b/stage0/src/Lean/Elab/Deriving/BEq.lean @@ -14,12 +14,12 @@ open Meta def mkBEqHeader (indVal : InductiveVal) : TermElabM Header := do mkHeader `BEq 2 indVal -def mkMatch (header : Header) (indVal : InductiveVal) (auxFunName : Name) : TermElabM Syntax := do +def mkMatch (header : Header) (indVal : InductiveVal) (auxFunName : Name) : TermElabM Term := do let discrs ← mkDiscrs header indVal let alts ← mkAlts `(match $[$discrs],* with $alts:matchAlt*) where - mkElseAlt : TermElabM Syntax := do + mkElseAlt : TermElabM (TSyntax ``matchAltExpr) := do let mut patterns := #[] -- add `_` pattern for indices for _ in [:indVal.numIndices] do @@ -29,7 +29,7 @@ where let altRhs ← `(false) `(matchAltExpr| | $[$patterns:term],* => $altRhs:term) - mkAlts : TermElabM (Array Syntax) := do + mkAlts : TermElabM (Array (TSyntax ``matchAlt)) := do let mut alts := #[] for ctorName in indVal.ctors do let ctorInfo ← getConstInfoCtor ctorName @@ -68,7 +68,7 @@ where alts := alts.push (← mkElseAlt) return alts -def mkAuxFunction (ctx : Context) (i : Nat) : TermElabM Syntax := do +def mkAuxFunction (ctx : Context) (i : Nat) : TermElabM Command := do let auxFunName := ctx.auxFunNames[i] let indVal := ctx.typeInfos[i] let header ← mkBEqHeader indVal @@ -78,9 +78,9 @@ def mkAuxFunction (ctx : Context) (i : Nat) : TermElabM Syntax := do body ← mkLet letDecls body let binders := header.binders if ctx.usePartial then - `(private partial def $(mkIdent auxFunName):ident $binders:explicitBinder* : Bool := $body:term) + `(private partial def $(mkIdent auxFunName):ident $binders:bracketedBinder* : Bool := $body:term) else - `(private def $(mkIdent auxFunName):ident $binders:explicitBinder* : Bool := $body:term) + `(private def $(mkIdent auxFunName):ident $binders:bracketedBinder* : Bool := $body:term) def mkMutualBlock (ctx : Context) : TermElabM Syntax := do let mut auxDefs := #[] diff --git a/stage0/src/Lean/Elab/Deriving/Basic.lean b/stage0/src/Lean/Elab/Deriving/Basic.lean index 1522cc074f1..4a72b66b280 100644 --- a/stage0/src/Lean/Elab/Deriving/Basic.lean +++ b/stage0/src/Lean/Elab/Deriving/Basic.lean @@ -9,7 +9,7 @@ import Lean.Elab.MutualDef namespace Lean.Elab open Command -def DerivingHandler := (typeNames : Array Name) → (args? : Option Syntax) → CommandElabM Bool +def DerivingHandler := (typeNames : Array Name) → (args? : Option (TSyntax ``Parser.Term.structInst)) → CommandElabM Bool def DerivingHandlerNoArgs := (typeNames : Array Name) → CommandElabM Bool builtin_initialize derivingHandlersRef : IO.Ref (NameMap DerivingHandler) ← IO.mkRef {} @@ -33,7 +33,7 @@ def registerBuiltinDerivingHandler (className : Name) (handler : DerivingHandler def defaultHandler (className : Name) (typeNames : Array Name) : CommandElabM Unit := do throwError "default handlers have not been implemented yet, class: '{className}' types: {typeNames}" -def applyDerivingHandlers (className : Name) (typeNames : Array Name) (args? : Option Syntax) : CommandElabM Unit := do +def applyDerivingHandlers (className : Name) (typeNames : Array Name) (args? : Option (TSyntax ``Parser.Term.structInst)) : CommandElabM Unit := do match (← derivingHandlersRef.get).find? className with | some handler => unless (← handler typeNames args?) do @@ -62,7 +62,7 @@ private def tryApplyDefHandler (className : Name) (declName : Name) : CommandEla structure DerivingClassView where ref : Syntax className : Name - args? : Option Syntax + args? : Option (TSyntax ``Parser.Term.structInst) def getOptDerivingClasses [Monad m] [MonadEnv m] [MonadResolveName m] [MonadError m] [MonadInfoTree m] (optDeriving : Syntax) : m (Array DerivingClassView) := do match optDeriving with diff --git a/stage0/src/Lean/Elab/Deriving/DecEq.lean b/stage0/src/Lean/Elab/Deriving/DecEq.lean index 6b232a2da7a..4820dc89d4e 100644 --- a/stage0/src/Lean/Elab/Deriving/DecEq.lean +++ b/stage0/src/Lean/Elab/Deriving/DecEq.lean @@ -15,12 +15,12 @@ open Meta def mkDecEqHeader (indVal : InductiveVal) : TermElabM Header := do mkHeader `DecidableEq 2 indVal -def mkMatch (header : Header) (indVal : InductiveVal) (auxFunName : Name) : TermElabM Syntax := do +def mkMatch (header : Header) (indVal : InductiveVal) (auxFunName : Name) : TermElabM Term := do let discrs ← mkDiscrs header indVal let alts ← mkAlts `(match $[$discrs],* with $alts:matchAlt*) where - mkSameCtorRhs : List (Syntax × Syntax × Bool × Bool) → TermElabM Syntax + mkSameCtorRhs : List (Ident × Ident × Bool × Bool) → TermElabM Term | [] => ``(isTrue rfl) | (a, b, recField, isProof) :: todo => withFreshMacroScope do let rhs ← if isProof then @@ -36,7 +36,7 @@ where else return rhs - mkAlts : TermElabM (Array Syntax) := do + mkAlts : TermElabM (Array (TSyntax ``matchAlt)) := do let mut alts := #[] for ctorName₁ in indVal.ctors do let ctorInfo ← getConstInfoCtor ctorName₁ @@ -88,7 +88,7 @@ def mkAuxFunction (ctx : Context) : TermElabM Syntax := do let mut body ← mkMatch header indVal auxFunName let binders := header.binders let type ← `(Decidable ($(mkIdent header.targetNames[0]) = $(mkIdent header.targetNames[1]))) - `(private def $(mkIdent auxFunName):ident $binders:explicitBinder* : $type:term := $body:term) + `(private def $(mkIdent auxFunName):ident $binders:bracketedBinder* : $type:term := $body:term) def mkDecEqCmds (indVal : InductiveVal) : TermElabM (Array Syntax) := do let ctx ← mkContext "decEq" indVal.name diff --git a/stage0/src/Lean/Elab/Deriving/FromToJson.lean b/stage0/src/Lean/Elab/Deriving/FromToJson.lean index 54497424bfc..b4f63dcbddc 100644 --- a/stage0/src/Lean/Elab/Deriving/FromToJson.lean +++ b/stage0/src/Lean/Elab/Deriving/FromToJson.lean @@ -14,7 +14,7 @@ open Lean.Json open Lean.Parser.Term open Lean.Meta -def mkJsonField (n : Name) : Bool × Syntax := +def mkJsonField (n : Name) : Bool × Term := let s := n.toString let s₁ := s.dropRightWhile (· == '?') (s != s₁, Syntax.mkStrLit s₁) @@ -26,11 +26,11 @@ def mkToJsonInstanceHandler (declNames : Array Name) : CommandElabM Bool := do let ctx ← mkContext "toJson" declNames[0] let header ← mkHeader ``ToJson 1 ctx.typeInfos[0] let fields := getStructureFieldsFlattened (← getEnv) declNames[0] (includeSubobjectFields := false) - let fields : Array Syntax ← fields.mapM fun field => do + let fields ← fields.mapM fun field => do let (isOptField, nm) := mkJsonField field if isOptField then ``(opt $nm $(mkIdent <| header.targetNames[0] ++ field)) else ``([($nm, toJson $(mkIdent <| header.targetNames[0] ++ field))]) - let cmd ← `(private def $(mkIdent ctx.auxFunNames[0]):ident $header.binders:explicitBinder* := + let cmd ← `(private def $(mkIdent ctx.auxFunNames[0]):ident $header.binders:bracketedBinder* := mkObj <| List.join [$fields,*]) return #[cmd] ++ (← mkInstanceCmds ctx ``ToJson declNames) cmds.forM elabCommand @@ -42,7 +42,7 @@ def mkToJsonInstanceHandler (declNames : Array Name) : CommandElabM Bool := do let toJsonFuncId := mkIdent ctx.auxFunNames[0] -- Return syntax to JSONify `id`, either via `ToJson` or recursively -- if `id`'s type is the type we're deriving for. - let mkToJson (id : Syntax) (type : Expr) : TermElabM Syntax := do + let mkToJson (id : Ident) (type : Expr) : TermElabM Term := do if type.isAppOf indVal.name then `($toJsonFuncId:ident $id:ident) else ``(toJson $id:ident) let header ← mkHeader ``ToJson 1 ctx.typeInfos[0] @@ -58,13 +58,14 @@ def mkToJsonInstanceHandler (declNames : Array Name) : CommandElabM Bool := do let xs ← xs.mapIdxM fun idx (x, t) => do `(($(quote userNames[idx].getString!), $(← mkToJson x t))) ``(mkObj [($(quote ctor.name.getString!), mkObj [$[$xs:term],*])]) - let mut auxCmd ← `(match $[$discrs],* with $alts:matchAlt*) - if ctx.usePartial then - let letDecls ← mkLocalInstanceLetDecls ctx ``ToJson header.argNames - auxCmd ← mkLet letDecls auxCmd - auxCmd ← `(private partial def $toJsonFuncId:ident $header.binders:explicitBinder* := $auxCmd) - else - auxCmd ← `(private def $toJsonFuncId:ident $header.binders:explicitBinder* := $auxCmd) + let auxTerm ← `(match $[$discrs],* with $alts:matchAlt*) + let auxCmd ← + if ctx.usePartial then + let letDecls ← mkLocalInstanceLetDecls ctx ``ToJson header.argNames + let auxTerm ← mkLet letDecls auxTerm + `(private partial def $toJsonFuncId:ident $header.binders:bracketedBinder* := $auxTerm) + else + `(private def $toJsonFuncId:ident $header.binders:bracketedBinder* := $auxTerm) return #[auxCmd] ++ (← mkInstanceCmds ctx ``ToJson declNames) cmds.forM elabCommand return true @@ -73,7 +74,7 @@ def mkToJsonInstanceHandler (declNames : Array Name) : CommandElabM Bool := do where mkAlts (indVal : InductiveVal) - (rhs : ConstructorVal → Array (Syntax × Expr) → (Option $ Array Name) → TermElabM Syntax) : TermElabM (Array Syntax) := do + (rhs : ConstructorVal → Array (Ident × Expr) → Option (Array Name) → TermElabM Term) : TermElabM (Array (TSyntax ``matchAlt)) := do indVal.ctors.toArray.mapM fun ctor => do let ctorInfo ← getConstInfoCtor ctor forallTelescopeReducing ctorInfo.type fun xs _ => do @@ -109,7 +110,7 @@ def mkFromJsonInstanceHandler (declNames : Array Name) : CommandElabM Bool := do let fields := getStructureFieldsFlattened (← getEnv) declNames[0] (includeSubobjectFields := false) let jsonFields := fields.map (Prod.snd ∘ mkJsonField) let fields := fields.map mkIdent - let cmd ← `(private def $(mkIdent ctx.auxFunNames[0]):ident $header.binders:explicitBinder* (j : Json) + let cmd ← `(private def $(mkIdent ctx.auxFunNames[0]):ident $header.binders:bracketedBinder* (j : Json) : Except String $(← mkInductiveApp ctx.typeInfos[0] header.argNames) := do $[let $fields:ident ← getObjValAs? j _ $jsonFields]* return { $[$fields:ident := $(id fields)],* }) @@ -123,29 +124,28 @@ def mkFromJsonInstanceHandler (declNames : Array Name) : CommandElabM Bool := do let header ← mkHeader ``FromJson 0 ctx.typeInfos[0] let fromJsonFuncId := mkIdent ctx.auxFunNames[0] let alts ← mkAlts indVal fromJsonFuncId - let mut auxCmd ← alts.foldrM (fun xs x => `(Except.orElseLazy $xs (fun _ => $x))) (← `(Except.error "no inductive constructor matched")) + let mut auxTerm ← alts.foldrM (fun xs x => `(Except.orElseLazy $xs (fun _ => $x))) (← `(Except.error "no inductive constructor matched")) if ctx.usePartial then let letDecls ← mkLocalInstanceLetDecls ctx ``FromJson header.argNames - auxCmd ← mkLet letDecls auxCmd + auxTerm ← mkLet letDecls auxTerm -- FromJson is not structurally recursive even non-nested recursive inductives, -- so we also use `partial` then. - if ctx.usePartial || indVal.isRec then - auxCmd ← `( - private partial def $fromJsonFuncId:ident $header.binders:explicitBinder* (json : Json) - : Except String $(← mkInductiveApp ctx.typeInfos[0] header.argNames) := - $auxCmd) - else - auxCmd ← `( - private def $fromJsonFuncId:ident $header.binders:explicitBinder* (json : Json) - : Except String $(← mkInductiveApp ctx.typeInfos[0] header.argNames) := - $auxCmd) + let auxCmd ← + if ctx.usePartial || indVal.isRec then + `(private partial def $fromJsonFuncId:ident $header.binders:bracketedBinder* (json : Json) + : Except String $(← mkInductiveApp ctx.typeInfos[0] header.argNames) := + $auxTerm) + else + `(private def $fromJsonFuncId:ident $header.binders:bracketedBinder* (json : Json) + : Except String $(← mkInductiveApp ctx.typeInfos[0] header.argNames) := + $auxTerm) return #[auxCmd] ++ (← mkInstanceCmds ctx ``FromJson declNames) cmds.forM elabCommand return true else return false where - mkAlts (indVal : InductiveVal) (fromJsonFuncId : Syntax) : TermElabM (Array Syntax) := do + mkAlts (indVal : InductiveVal) (fromJsonFuncId : Ident) : TermElabM (Array Term) := do let alts ← indVal.ctors.toArray.mapM fun ctor => do let ctorInfo ← getConstInfoCtor ctor @@ -162,19 +162,19 @@ where -- Return syntax to parse `id`, either via `FromJson` or recursively -- if `id`'s type is the type we're deriving for. - let mkFromJson (idx : Nat) (type : Expr) : TermElabM Syntax := + let mkFromJson (idx : Nat) (type : Expr) : TermElabM (TSyntax ``doExpr) := if type.isAppOf indVal.name then `(Lean.Parser.Term.doExpr| $fromJsonFuncId:ident jsons[$(quote idx)]) else `(Lean.Parser.Term.doExpr| fromJson? jsons[$(quote idx)]) let identNames := binders.map Prod.fst let fromJsons ← binders.mapIdxM fun idx (_, type) => mkFromJson idx type let userNamesOpt ← if binders.size == userNames.size then - ``(some #[$[$(userNames.map quote):ident],*]) + ``(some #[$[$(userNames.map quote)],*]) else ``(none) let stx ← `((Json.parseTagged json $(quote ctor.getString!) $(quote ctorInfo.numFields) $(quote userNamesOpt)).bind (fun jsons => do - $[let $identNames:ident ← $fromJsons]* + $[let $identNames:ident ← $fromJsons:doExpr]* return $(mkIdent ctor):ident $identNames*)) pure (stx, ctorInfo.numFields) -- the smaller cases, especially the ones without fields are likely faster diff --git a/stage0/src/Lean/Elab/Deriving/Hashable.lean b/stage0/src/Lean/Elab/Deriving/Hashable.lean index 0c036e86841..94fbe62fc3e 100644 --- a/stage0/src/Lean/Elab/Deriving/Hashable.lean +++ b/stage0/src/Lean/Elab/Deriving/Hashable.lean @@ -15,13 +15,13 @@ open Meta def mkHashableHeader (indVal : InductiveVal) : TermElabM Header := do mkHeader `Hashable 1 indVal -def mkMatch (ctx : Context) (header : Header) (indVal : InductiveVal) : TermElabM Syntax := do +def mkMatch (ctx : Context) (header : Header) (indVal : InductiveVal) : TermElabM Term := do let discrs ← mkDiscrs header indVal let alts ← mkAlts `(match $[$discrs],* with $alts:matchAlt*) where - mkAlts : TermElabM (Array Syntax) := do + mkAlts : TermElabM (Array (TSyntax ``matchAlt)) := do let mut alts := #[] let mut ctorIdx := 0 let allIndVals := indVal.all.toArray @@ -54,7 +54,7 @@ where ctorIdx := ctorIdx + 1 return alts -def mkAuxFunction (ctx : Context) (i : Nat) : TermElabM Syntax := do +def mkAuxFunction (ctx : Context) (i : Nat) : TermElabM Command := do let auxFunName := ctx.auxFunNames[i] let indVal := ctx.typeInfos[i] let header ← mkHashableHeader indVal @@ -62,9 +62,9 @@ def mkAuxFunction (ctx : Context) (i : Nat) : TermElabM Syntax := do let binders := header.binders if ctx.usePartial then -- TODO(Dany): Get rid of this code branch altogether once we have well-founded recursion - `(private partial def $(mkIdent auxFunName):ident $binders:explicitBinder* : UInt64 := $body:term) + `(private partial def $(mkIdent auxFunName):ident $binders:bracketedBinder* : UInt64 := $body:term) else - `(private def $(mkIdent auxFunName):ident $binders:explicitBinder* : UInt64 := $body:term) + `(private def $(mkIdent auxFunName):ident $binders:bracketedBinder* : UInt64 := $body:term) def mkHashFuncs (ctx : Context) : TermElabM Syntax := do let mut auxDefs := #[] diff --git a/stage0/src/Lean/Elab/Deriving/Inhabited.lean b/stage0/src/Lean/Elab/Deriving/Inhabited.lean index b3f8e5d80db..48d80e69934 100644 --- a/stage0/src/Lean/Elab/Deriving/Inhabited.lean +++ b/stage0/src/Lean/Elab/Deriving/Inhabited.lean @@ -66,10 +66,10 @@ where for i in [:indVal.numParams + indVal.numIndices] do let arg := mkIdent (← mkFreshUserName `a) indArgs := indArgs.push arg - let binder ← `(implicitBinderF| { $arg:ident }) + let binder ← `(bracketedBinder| { $arg:ident }) binders := binders.push binder if assumingParamIdxs.contains i then - let binder ← `(instBinderF| [ Inhabited $arg:ident ]) + let binder ← `(bracketedBinder| [Inhabited $arg:ident ]) binders := binders.push binder let type ← `(Inhabited (@$(mkIdent inductiveTypeName):ident $indArgs:ident*)) let mut ctorArgs := #[] @@ -77,8 +77,8 @@ where ctorArgs := ctorArgs.push (← `(_)) for _ in [:ctorVal.numFields] do ctorArgs := ctorArgs.push (← ``(Inhabited.default)) - let val ← `(⟨@$(mkIdent ctorName):ident $ctorArgs:ident*⟩) - `(instance $binders:explicitBinder* : $type := $val) + let val ← `(⟨@$(mkIdent ctorName):ident $ctorArgs*⟩) + `(instance $binders:bracketedBinder* : $type := $val) mkInstanceCmd? : TermElabM (Option Syntax) := do let ctorVal ← getConstInfoCtor ctorName diff --git a/stage0/src/Lean/Elab/Deriving/Ord.lean b/stage0/src/Lean/Elab/Deriving/Ord.lean index 0e398468d91..5e092ea82a4 100644 --- a/stage0/src/Lean/Elab/Deriving/Ord.lean +++ b/stage0/src/Lean/Elab/Deriving/Ord.lean @@ -14,12 +14,12 @@ open Meta def mkOrdHeader (indVal : InductiveVal) : TermElabM Header := do mkHeader `Ord 2 indVal -def mkMatch (header : Header) (indVal : InductiveVal) : TermElabM Syntax := do +def mkMatch (header : Header) (indVal : InductiveVal) : TermElabM Term := do let discrs ← mkDiscrs header indVal let alts ← mkAlts `(match $[$discrs],* with $alts:matchAlt*) where - mkAlts : TermElabM (Array Syntax) := do + mkAlts : TermElabM (Array (TSyntax ``matchAlt)) := do let mut alts := #[] for ctorName in indVal.ctors do let ctorInfo ← getConstInfoCtor ctorName @@ -32,7 +32,7 @@ where let mut ctorArgs1 := #[] let mut ctorArgs2 := #[] -- construct RHS top-down as continuation over the remaining comparison - let mut rhsCont : Syntax → TermElabM Syntax := fun rhs => pure rhs + let mut rhsCont : Term → TermElabM Term := fun rhs => pure rhs -- add `_` for inductive parameters, they are inaccessible for _ in [:indVal.numParams] do ctorArgs1 := ctorArgs1.push (← `(_)) @@ -48,7 +48,7 @@ where let b := mkIdent (← mkFreshUserName `b) ctorArgs1 := ctorArgs1.push a ctorArgs2 := ctorArgs2.push b - rhsCont := fun rhs => `(match compare $a:ident $b:ident with + rhsCont := fun rhs => `(match compare $a $b with | Ordering.lt => Ordering.lt | Ordering.gt => Ordering.gt | Ordering.eq => $rhs) >>= rhsCont @@ -61,10 +61,10 @@ where pure #[←`(matchAltExpr| | $[$(patterns):term],* => $rhs:term), ←`(matchAltExpr| | $[$(ltPatterns):term],* => Ordering.lt), ←`(matchAltExpr| | $[$(gtPatterns):term],* => Ordering.gt)] - alts := alts ++ alt + alts := alts ++ (alt : Array (TSyntax ``matchAlt)) return alts.pop.pop -def mkAuxFunction (ctx : Context) (i : Nat) : TermElabM Syntax := do +def mkAuxFunction (ctx : Context) (i : Nat) : TermElabM Command := do let auxFunName := ctx.auxFunNames[i] let indVal := ctx.typeInfos[i] let header ← mkOrdHeader indVal @@ -74,9 +74,9 @@ def mkAuxFunction (ctx : Context) (i : Nat) : TermElabM Syntax := do body ← mkLet letDecls body let binders := header.binders if ctx.usePartial || indVal.isRec then - `(private partial def $(mkIdent auxFunName):ident $binders:explicitBinder* : Ordering := $body:term) + `(private partial def $(mkIdent auxFunName):ident $binders:bracketedBinder* : Ordering := $body:term) else - `(private def $(mkIdent auxFunName):ident $binders:explicitBinder* : Ordering := $body:term) + `(private def $(mkIdent auxFunName):ident $binders:bracketedBinder* : Ordering := $body:term) def mkMutualBlock (ctx : Context) : TermElabM Syntax := do let mut auxDefs := #[] diff --git a/stage0/src/Lean/Elab/Deriving/Repr.lean b/stage0/src/Lean/Elab/Deriving/Repr.lean index f57ad1ea30a..105f4a82a29 100644 --- a/stage0/src/Lean/Elab/Deriving/Repr.lean +++ b/stage0/src/Lean/Elab/Deriving/Repr.lean @@ -16,16 +16,16 @@ open Std def mkReprHeader (indVal : InductiveVal) : TermElabM Header := do let header ← mkHeader `Repr 1 indVal return { header with - binders := header.binders.push (← `(explicitBinderF| (prec : Nat))) + binders := header.binders.push (← `(bracketedBinder| (prec : Nat))) } -def mkBodyForStruct (header : Header) (indVal : InductiveVal) : TermElabM Syntax := do +def mkBodyForStruct (header : Header) (indVal : InductiveVal) : TermElabM Term := do let ctorVal ← getConstInfoCtor indVal.ctors.head! let fieldNames := getStructureFields (← getEnv) indVal.name let numParams := indVal.numParams let target := mkIdent header.targetNames[0] forallTelescopeReducing ctorVal.type fun xs _ => do - let mut fields : Syntax ← `(Format.nil) + let mut fields ← `(Format.nil) let mut first := true if xs.size != numParams + fieldNames.size then throwError "'deriving Repr' failed, unexpected number of fields in structure" @@ -43,12 +43,12 @@ def mkBodyForStruct (header : Header) (indVal : InductiveVal) : TermElabM Syntax fields ← `($fields ++ $fieldNameLit ++ " := " ++ repr ($target.$(mkIdent fieldName):ident)) `(Format.bracket "{ " $fields:term " }") -def mkBodyForInduct (header : Header) (indVal : InductiveVal) (auxFunName : Name) : TermElabM Syntax := do +def mkBodyForInduct (header : Header) (indVal : InductiveVal) (auxFunName : Name) : TermElabM Term := do let discrs ← mkDiscrs header indVal let alts ← mkAlts `(match $[$discrs],* with $alts:matchAlt*) where - mkAlts : TermElabM (Array Syntax) := do + mkAlts : TermElabM (Array (TSyntax ``matchAlt)) := do let mut alts := #[] for ctorName in indVal.ctors do let ctorInfo ← getConstInfoCtor ctorName @@ -58,7 +58,7 @@ where for _ in [:indVal.numIndices] do patterns := patterns.push (← `(_)) let mut ctorArgs := #[] - let mut rhs := Syntax.mkStrLit (toString ctorInfo.name) + let mut rhs : Term := Syntax.mkStrLit (toString ctorInfo.name) rhs ← `(Format.text $rhs) -- add `_` for inductive parameters, they are inaccessible for _ in [:indVal.numParams] do @@ -78,13 +78,13 @@ where alts := alts.push alt return alts -def mkBody (header : Header) (indVal : InductiveVal) (auxFunName : Name) : TermElabM Syntax := do +def mkBody (header : Header) (indVal : InductiveVal) (auxFunName : Name) : TermElabM Term := do if isStructure (← getEnv) indVal.name then mkBodyForStruct header indVal else mkBodyForInduct header indVal auxFunName -def mkAuxFunction (ctx : Context) (i : Nat) : TermElabM Syntax := do +def mkAuxFunction (ctx : Context) (i : Nat) : TermElabM Command := do let auxFunName := ctx.auxFunNames[i] let indVal := ctx.typeInfos[i] let header ← mkReprHeader indVal @@ -94,9 +94,9 @@ def mkAuxFunction (ctx : Context) (i : Nat) : TermElabM Syntax := do body ← mkLet letDecls body let binders := header.binders if ctx.usePartial then - `(private partial def $(mkIdent auxFunName):ident $binders:explicitBinder* : Format := $body:term) + `(private partial def $(mkIdent auxFunName):ident $binders:bracketedBinder* : Format := $body:term) else - `(private def $(mkIdent auxFunName):ident $binders:explicitBinder* : Format := $body:term) + `(private def $(mkIdent auxFunName):ident $binders:bracketedBinder* : Format := $body:term) def mkMutualBlock (ctx : Context) : TermElabM Syntax := do let mut auxDefs := #[] diff --git a/stage0/src/Lean/Elab/Deriving/Util.lean b/stage0/src/Lean/Elab/Deriving/Util.lean index d37d312e8d8..745d5acaf5e 100644 --- a/stage0/src/Lean/Elab/Deriving/Util.lean +++ b/stage0/src/Lean/Elab/Deriving/Util.lean @@ -26,15 +26,16 @@ def mkInductArgNames (indVal : InductiveVal) : TermElabM (Array Name) := do pure argNames /-- Return the inductive declaration's type applied to the arguments in `argNames`. -/ -def mkInductiveApp (indVal : InductiveVal) (argNames : Array Name) : TermElabM Syntax := +def mkInductiveApp (indVal : InductiveVal) (argNames : Array Name) : TermElabM Term := let f := mkIdent indVal.name let args := argNames.map mkIdent `(@$f $args*) +open TSyntax.Compat in /-- Return implicit binder syntaxes for the given `argNames`. The output matches `implicitBinder*`. For example, ``#[`foo,`bar]`` gives `` `({foo} {bar})``. -/ -def mkImplicitBinders (argNames : Array Name) : TermElabM (Array Syntax) := +def mkImplicitBinders (argNames : Array Name) : TermElabM (Array (TSyntax ``Parser.Term.implicitBinder)) := argNames.mapM fun argName => `(implicitBinderF| { $(mkIdent argName) }) @@ -53,7 +54,7 @@ def mkInstImplicitBinders (className : Name) (indVal : InductiveVal) (argNames : let c ← mkAppM className #[x] if (← isTypeCorrect c) then let argName := argNames[i] - let binder ← `(instBinderF| [ $(mkIdent className):ident $(mkIdent argName):ident ]) + let binder : Syntax ← `(instBinderF| [ $(mkIdent className):ident $(mkIdent argName):ident ]) binders := binders.push binder catch _ => pure () @@ -82,7 +83,7 @@ def mkContext (fnPrefix : String) (typeName : Name) : TermElabM Context := do usePartial := usePartial } -def mkLocalInstanceLetDecls (ctx : Context) (className : Name) (argNames : Array Name) : TermElabM (Array Syntax) := do +def mkLocalInstanceLetDecls (ctx : Context) (className : Name) (argNames : Array Name) : TermElabM (Array (TSyntax ``Parser.Term.letDecl)) := do let mut letDecls := #[] for i in [:ctx.typeInfos.size] do let indVal := ctx.typeInfos[i] @@ -100,11 +101,12 @@ def mkLocalInstanceLetDecls (ctx : Context) (className : Name) (argNames : Array letDecls := letDecls.push letDecl return letDecls -def mkLet (letDecls : Array Syntax) (body : Syntax) : TermElabM Syntax := +def mkLet (letDecls : Array (TSyntax ``Parser.Term.letDecl)) (body : Term) : TermElabM Term := letDecls.foldrM (init := body) fun letDecl body => `(let $letDecl:letDecl; $body) -def mkInstanceCmds (ctx : Context) (className : Name) (typeNames : Array Name) (useAnonCtor := true) : TermElabM (Array Syntax) := do +open TSyntax.Compat in +def mkInstanceCmds (ctx : Context) (className : Name) (typeNames : Array Name) (useAnonCtor := true) : TermElabM (Array Command) := do let mut instances := #[] for i in [:ctx.typeInfos.size] do let indVal := ctx.typeInfos[i] @@ -120,15 +122,16 @@ def mkInstanceCmds (ctx : Context) (className : Name) (typeNames : Array Name) ( instances := instances.push instCmd return instances -def mkDiscr (varName : Name) : TermElabM Syntax := +def mkDiscr (varName : Name) : TermElabM (TSyntax ``Parser.Term.matchDiscr) := `(Parser.Term.matchDiscr| $(mkIdent varName):term) structure Header where - binders : Array Syntax + binders : Array (TSyntax ``Parser.Term.bracketedBinder) argNames : Array Name targetNames : Array Name - targetType : Syntax + targetType : Term +open TSyntax.Compat in def mkHeader (className : Name) (arity : Nat) (indVal : InductiveVal) : TermElabM Header := do let argNames ← mkInductArgNames indVal let binders ← mkImplicitBinders argNames @@ -145,7 +148,7 @@ def mkHeader (className : Name) (arity : Nat) (indVal : InductiveVal) : TermElab targetType := targetType } -def mkDiscrs (header : Header) (indVal : InductiveVal) : TermElabM (Array Syntax) := do +def mkDiscrs (header : Header) (indVal : InductiveVal) : TermElabM (Array (TSyntax ``Parser.Term.matchDiscr)) := do let mut discrs := #[] -- add indices for argName in header.argNames[indVal.numParams:] do diff --git a/stage0/src/Lean/Elab/Do.lean b/stage0/src/Lean/Elab/Do.lean index bf0f6f86f4b..367bac9a98d 100644 --- a/stage0/src/Lean/Elab/Do.lean +++ b/stage0/src/Lean/Elab/Do.lean @@ -15,6 +15,7 @@ set_option compiler.reuse false namespace Lean.Elab.Term open Lean.Parser.Term open Meta +open TSyntax.Compat private def getDoSeqElems (doSeq : Syntax) : List Syntax := if doSeq.getKind == ``Lean.Parser.Term.doSeqBracketed then @@ -122,7 +123,7 @@ private partial def extractBind (expectedType? : Option Expr) : TermElabM Extrac namespace Do -abbrev Var := Syntax -- TODO: should be `TSyntax identKind` +abbrev Var := Syntax -- TODO: should be `Ident` /- A `doMatch` alternative. `vars` is the array of variables declared by `patterns`. -/ structure Alt (σ : Type) where @@ -305,7 +306,7 @@ def attachJPs (jpDecls : Array JPDecl) (k : Code) : Code := def mkFreshJP (ps : Array (Var × Bool)) (body : Code) : TermElabM JPDecl := do let ps ← if ps.isEmpty then let y ← `(y) - pure #[(y, false)] + pure #[(y.raw, false)] else pure ps -- Remark: the compiler frontend implemented in C++ currently detects jointpoints created by @@ -1171,7 +1172,7 @@ private partial def expandLiftMethodAux (inQuot : Bool) (inBinder : Bool) : Synt throwErrorAt stx "cannot lift `(<- ...)` over a binder, this error usually happens when you are trying to lift a method nested in a `fun`, `let`, or `match`-alternative, and it can often be fixed by adding a missing `do`" let term := args[1] let term ← expandLiftMethodAux inQuot inBinder term - let auxDoElem ← `(doElem| let a ← $term:term) + let auxDoElem : Syntax ← `(doElem| let a ← $term:term) modify fun s => s ++ [auxDoElem] `(a) else do @@ -1439,7 +1440,7 @@ mutual let optMotive := doMatch[2] let discrs := doMatch[3] let matchAlts := doMatch[5][0].getArgs -- Array of `doMatchAlt` - let matchAlts := matchAlts.foldl (init := #[]) fun result matchAlt => result ++ expandMatchAlt matchAlt + let matchAlts ← matchAlts.foldlM (init := #[]) fun result matchAlt => return result ++ (← liftMacroM <| expandMatchAlt matchAlt) let alts ← matchAlts.mapM fun matchAlt => do let patterns := matchAlt[1][0] let vars ← getPatternsVarsEx patterns.getSepArgs diff --git a/stage0/src/Lean/Elab/ElabRules.lean b/stage0/src/Lean/Elab/ElabRules.lean index 07062ac6986..2661a529602 100644 --- a/stage0/src/Lean/Elab/ElabRules.lean +++ b/stage0/src/Lean/Elab/ElabRules.lean @@ -9,6 +9,7 @@ import Lean.Elab.AuxDef namespace Lean.Elab.Command open Lean.Syntax open Lean.Parser.Term hiding macroArg +open Lean.Parser.Command def withExpectedType (expectedType? : Option Expr) (x : Expr → TermElabM Expr) : TermElabM Expr := do Term.tryPostponeIfNoneOrMVar expectedType? @@ -16,8 +17,8 @@ def withExpectedType (expectedType? : Option Expr) (x : Expr → TermElabM Expr) | throwError "expected type must be known" x expectedType -def elabElabRulesAux (doc? : Option Syntax) (attrKind : Syntax) (k : SyntaxNodeKind) (cat? expty? : Option Syntax) (alts : Array Syntax) : CommandElabM Syntax := do - let alts ← alts.mapM fun alt => match alt with +def elabElabRulesAux (doc? : Option (TSyntax ``docComment)) (attrKind : TSyntax ``attrKind) (k : SyntaxNodeKind) (cat? expty? : Option (Ident)) (alts : Array (TSyntax ``matchAlt)) : CommandElabM Syntax := do + let alts ← alts.mapM fun (alt : TSyntax ``matchAlt) => match alt with | `(matchAltExpr| | $pats,* => $rhs) => do let pat := pats.elemsAndSeps[0] if !pat.isQuot then @@ -31,7 +32,7 @@ def elabElabRulesAux (doc? : Option Syntax) (attrKind : Syntax) (k : SyntaxNodeK | none => throwErrorAt alt "invalid elab_rules alternative, expected syntax node kind '{k}'" | some quoted => let pat := pat.setArg 1 quoted - let pats := pats.elemsAndSeps.set! 0 pat + let pats := ⟨pats.elemsAndSeps.set! 0 pat⟩ `(matchAltExpr| | $pats,* => $rhs) else throwErrorAt alt "invalid elab_rules alternative, unexpected syntax node kind '{k'}'" @@ -87,9 +88,9 @@ def elabElab : CommandElab let name ← match name? with | some name => pure name.getId | none => liftMacroM <| mkNameFromParserSyntax cat.getId (mkNullNode stxParts) - let pat := mkNode ((← getCurrNamespace) ++ name) patArgs - `($[$doc?:docComment]? $attrKind:attrKind syntax$[:$prec?]? (name := $(← mkIdentFromRef name)) (priority := $(quote prio)) $[$stxParts]* : $cat - $[$doc?:docComment]? elab_rules : $cat $[<= $expectedType?]? | `($pat) => $rhs) >>= elabCommand + let pat := ⟨mkNode ((← getCurrNamespace) ++ name) patArgs⟩ + `($[$doc?:docComment]? $attrKind:attrKind syntax$[:$prec?]? (name := $(← mkIdentFromRef name)) (priority := $(quote prio):num) $[$stxParts]* : $cat + $[$doc?:docComment]? elab_rules : $cat $[<= $expectedType?]? | `($pat) => $rhs) >>= (elabCommand ·) | _ => throwUnsupportedSyntax end Lean.Elab.Command diff --git a/stage0/src/Lean/Elab/Macro.lean b/stage0/src/Lean/Elab/Macro.lean index 3878a60a0a1..3616d036ef2 100644 --- a/stage0/src/Lean/Elab/Macro.lean +++ b/stage0/src/Lean/Elab/Macro.lean @@ -22,16 +22,17 @@ open Lean.Parser.Command | none => liftMacroM <| mkNameFromParserSyntax cat.getId (mkNullNode stxParts) /- The command `syntax [] ...` adds the current namespace to the syntax node kind. So, we must include current namespace when we create a pattern for the following `macro_rules` commands. -/ - let pat := mkNode ((← getCurrNamespace) ++ name) patArgs + let pat := ⟨mkNode ((← getCurrNamespace) ++ name) patArgs⟩ let stxCmd ← `($[$doc?:docComment]? $attrKind:attrKind - syntax%$tk$[:$prec?]? (name := $(← mkIdentFromRef name)) (priority := $(quote prio)) $[$stxParts]* : $cat) + syntax%$tk$[:$prec?]? (name := $(← mkIdentFromRef name)) (priority := $(quote prio):num) $[$stxParts]* : $cat) + let rhs := rhs.raw let macroRulesCmd ← if rhs.getArgs.size == 1 then -- `rhs` is a `term` - let rhs := rhs[0] + let rhs := ⟨rhs[0]⟩ `($[$doc?:docComment]? macro_rules%$tk | `($pat) => $rhs) else -- `rhs` is of the form `` `( $body ) `` - let rhsBody := rhs[1] + let rhsBody := ⟨rhs[1]⟩ `($[$doc?:docComment]? macro_rules%$tk | `($pat) => `($rhsBody)) elabCommand <| mkNullNode #[stxCmd, macroRulesCmd] | _ => throwUnsupportedSyntax diff --git a/stage0/src/Lean/Elab/MacroArgUtil.lean b/stage0/src/Lean/Elab/MacroArgUtil.lean index b8aae5dfa06..1596aa2b26c 100644 --- a/stage0/src/Lean/Elab/MacroArgUtil.lean +++ b/stage0/src/Lean/Elab/MacroArgUtil.lean @@ -11,17 +11,17 @@ open Lean.Parser.Term hiding macroArg open Lean.Parser.Command /- Convert `macro` arg into a `syntax` command item and a pattern element -/ -partial def expandMacroArg (stx : Syntax) : CommandElabM (Syntax × Syntax) := do +partial def expandMacroArg (stx : TSyntax ``macroArg) : CommandElabM (TSyntax `stx × Term) := do let (id?, id, stx) ← match (← liftMacroM <| expandMacros stx) with - | `(macroArg| $id:ident:$stx) => pure (some id, id, stx) + | `(macroArg| $id:ident:$stx) => pure (some id, (id : Term), stx) | `(macroArg| $stx:stx) => pure (none, (← `(x)), stx) | _ => throwUnsupportedSyntax mkSyntaxAndPat id? id stx where mkSyntaxAndPat id? id stx := do let pat ← match stx with - | `(stx| $s:str) => pure <| mkNode `token_antiquot #[← liftMacroM <| strLitToPattern s, mkAtom "%", mkAtom "$", id] - | `(stx| &$s:str) => pure <| mkNode `token_antiquot #[← liftMacroM <| strLitToPattern s, mkAtom "%", mkAtom "$", id] + | `(stx| $s:str) => pure ⟨mkNode `token_antiquot #[← liftMacroM <| strLitToPattern s, mkAtom "%", mkAtom "$", id]⟩ + | `(stx| &$s:str) => pure ⟨mkNode `token_antiquot #[← liftMacroM <| strLitToPattern s, mkAtom "%", mkAtom "$", id]⟩ | `(stx| optional($stx)) => mkSplicePat `optional stx id "?" | `(stx| many($stx)) => mkSplicePat `many stx id "*" | `(stx| many1($stx)) => mkSplicePat `many stx id "*" @@ -30,7 +30,7 @@ where | `(stx| sepBy1($stx, $sep:str $[, $stxsep]? $[, allowTrailingSep]?)) => mkSplicePat `sepBy stx id ((isStrLit? sep).get! ++ "*") -- NOTE: all `interpolatedStr(·)` reuse the same node kind - | `(stx| interpolatedStr(term)) => pure <| Syntax.mkAntiquotNode interpolatedStrKind id + | `(stx| interpolatedStr(term)) => pure ⟨Syntax.mkAntiquotNode interpolatedStrKind id⟩ -- bind through withPosition | `(stx| withPosition($stx)) => return (← mkSyntaxAndPat id? id stx) @@ -40,9 +40,9 @@ where -- otherwise `group` the syntax to enforce arity 1, e.g. for `noWs` | none => return (← `(stx| group($stx)), (← mkAntiquotNode stx id)) pure (stx, pat) - mkSplicePat kind stx id suffix := - return mkNullNode #[mkAntiquotSuffixSpliceNode kind (← mkAntiquotNode stx id) suffix] - mkAntiquotNode + mkSplicePat kind stx id suffix : CommandElabM Term := + return ⟨mkNullNode #[mkAntiquotSuffixSpliceNode kind (← mkAntiquotNode stx id) suffix]⟩ + mkAntiquotNode : TSyntax `stx → Term → CommandElabM Term | `(stx| ($stx)), term => mkAntiquotNode stx term | `(stx| $id:ident$[:$p:prec]?), term => do let kind ← match (← Elab.Term.resolveParserName id) with @@ -55,13 +55,13 @@ where | [] => let id := id.getId.eraseMacroScopes if Parser.isParserCategory (← getEnv) id then - return Syntax.mkAntiquotNode id term (isPseudoKind := true) + return ⟨Syntax.mkAntiquotNode id term (isPseudoKind := true)⟩ else if (← Parser.isParserAlias id) then pure <| (← Parser.getSyntaxKindOfParserAlias? id).getD Name.anonymous else throwError "unknown parser declaration/category/alias '{id}'" - pure <| Syntax.mkAntiquotNode kind term + pure ⟨Syntax.mkAntiquotNode kind term⟩ | stx, term => - pure <| Syntax.mkAntiquotNode Name.anonymous term (isPseudoKind := true) + pure ⟨Syntax.mkAntiquotNode Name.anonymous term (isPseudoKind := true)⟩ end Lean.Elab.Command diff --git a/stage0/src/Lean/Elab/MacroRules.lean b/stage0/src/Lean/Elab/MacroRules.lean index a6f047debaa..406cb308f7f 100644 --- a/stage0/src/Lean/Elab/MacroRules.lean +++ b/stage0/src/Lean/Elab/MacroRules.lean @@ -15,8 +15,8 @@ open Lean.Parser.Command Remark: `k` is the user provided kind with the current namespace included. Recall that syntax node kinds contain the current namespace. -/ -def elabMacroRulesAux (doc? : Option Syntax) (attrKind tk : Syntax) (k : SyntaxNodeKind) (alts : Array Syntax) : CommandElabM Syntax := do - let alts ← alts.mapM fun alt => match alt with +def elabMacroRulesAux (doc? : Option (TSyntax ``docComment)) (attrKind : TSyntax ``attrKind) (tk : Syntax) (k : SyntaxNodeKind) (alts : Array (TSyntax ``matchAlt)) : CommandElabM Syntax := do + let alts ← alts.mapM fun (alt : TSyntax ``matchAlt) => match alt with | `(matchAltExpr| | $pats,* => $rhs) => do let pat := pats.elemsAndSeps[0] if !pat.isQuot then @@ -31,7 +31,7 @@ def elabMacroRulesAux (doc? : Option Syntax) (attrKind tk : Syntax) (k : SyntaxN | some quoted => let pat := pat.setArg 1 quoted let pats := pats.elemsAndSeps.set! 0 pat - `(matchAltExpr| | $pats,* => $rhs) + `(matchAltExpr| | $(⟨pats⟩),* => $rhs) else throwErrorAt alt "invalid macro_rules alternative, unexpected syntax node kind '{k'}'" | _ => throwUnsupportedSyntax diff --git a/stage0/src/Lean/Elab/Match.lean b/stage0/src/Lean/Elab/Match.lean index fb4e053f865..5e3fb8c2114 100644 --- a/stage0/src/Lean/Elab/Match.lean +++ b/stage0/src/Lean/Elab/Match.lean @@ -19,7 +19,7 @@ namespace Lean.Elab.Term open Meta open Lean.Parser.Term -private def expandSimpleMatch (stx discr lhsVar rhs : Syntax) (expectedType? : Option Expr) : TermElabM Expr := do +private def expandSimpleMatch (stx : Syntax) (discr : Term) (lhsVar : Ident) (rhs : Term) (expectedType? : Option Expr) : TermElabM Expr := do let newStx ← `(let $lhsVar := $discr; $rhs) withMacroExpansion stx newStx <| elabTerm newStx expectedType? @@ -658,12 +658,9 @@ where if inaccessible? p |>.isSome then return mkMData k (← withReader (fun _ => false) (go b)) else if let some (stx, p) := patternWithRef? p then - let p ← go p - if p.isFVar && !(← read) then + Elab.withInfoContext' (go p) fun p => do /- If `p` is a free variable and we are not inside of an "inaccessible" pattern, this `p` is a binder. -/ - addTermInfo stx p (isBinder := true) - else - addTermInfo stx p + mkTermInfo Name.anonymous stx p (isBinder := p.isFVar && !(← read)) else return mkMData k (← go b) | _ => return p @@ -1124,6 +1121,7 @@ private def getDiscrs (matchStx : Syntax) : Array Syntax := private def getMatchOptMotive (matchStx : Syntax) : Syntax := matchStx[2] +open TSyntax.Compat in private def expandNonAtomicDiscrs? (matchStx : Syntax) : TermElabM (Option Syntax) := let matchOptMotive := getMatchOptMotive matchStx if matchOptMotive.isNone then do diff --git a/stage0/src/Lean/Elab/Mixfix.lean b/stage0/src/Lean/Elab/Mixfix.lean index 722384823e9..f65d7bd4222 100644 --- a/stage0/src/Lean/Elab/Mixfix.lean +++ b/stage0/src/Lean/Elab/Mixfix.lean @@ -30,6 +30,6 @@ where let attrKind := stx[0] let stx := stx.setArg 0 mkAttrKindGlobal let stx ← f stx - return stx.setArg 0 attrKind + return stx.raw.setArg 0 attrKind end Lean.Elab.Command diff --git a/stage0/src/Lean/Elab/MutualDef.lean b/stage0/src/Lean/Elab/MutualDef.lean index fb64a63b7f4..49dba0ddb28 100644 --- a/stage0/src/Lean/Elab/MutualDef.lean +++ b/stage0/src/Lean/Elab/MutualDef.lean @@ -169,11 +169,11 @@ private def expandWhereStructInst : Macro | `(letDecl|$decl:letIdDecl) => pure decl | _ => Macro.throwUnsupported let structInstFields ← letIdDecls.mapM fun - | stx@`(letIdDecl|$id:ident $[$binders]* $[: $ty?]? := $val) => withRef stx do + | stx@`(letIdDecl|$id:ident $binders* $[: $ty?]? := $val) => withRef stx do let mut val := val if let some ty := ty? then val ← `(($val : $ty)) - val ← if binders.size > 0 then `(fun $[$binders]* => $val:term) else pure val + val ← if binders.size > 0 then `(fun $binders* => $val) else pure val `(structInstField|$id:ident := $val) | _ => Macro.throwUnsupported let body ← `({ $structInstFields,* }) diff --git a/stage0/src/Lean/Elab/Notation.lean b/stage0/src/Lean/Elab/Notation.lean index c4c517427ff..2bd12f9b4c8 100644 --- a/stage0/src/Lean/Elab/Notation.lean +++ b/stage0/src/Lean/Elab/Notation.lean @@ -25,14 +25,10 @@ private partial def antiquote (vars : Array Syntax) : Syntax → Syntax | stx => stx /- Convert `notation` command lhs item into a `syntax` command item -/ -def expandNotationItemIntoSyntaxItem (stx : Syntax) : MacroM Syntax := - let k := stx.getKind - if k == `Lean.Parser.Command.identPrec then - pure $ mkNode `Lean.Parser.Syntax.cat #[mkIdentFrom stx `term, stx[1]] - else if k == strLitKind then - pure $ mkNode `Lean.Parser.Syntax.atom #[stx] - else - Macro.throwUnsupported +def expandNotationItemIntoSyntaxItem : TSyntax ``notationItem → MacroM (TSyntax `stx) + | `(notationItem| $id:ident$[:$prec?]?) => `(stx| term $[:$prec?]?) + | `(notationItem| $s:str) => `(stx| $s:str) + | _ => Macro.throwUnsupported /- Convert `notation` command lhs item into a pattern element -/ def expandNotationItemIntoPattern (stx : Syntax) : MacroM Syntax := @@ -75,37 +71,33 @@ partial def hasDuplicateAntiquot (stxs : Array Syntax) : Bool := Id.run do The notation must be of the form `notation ... => c body` where `c` is a declaration in the current scope and `body` any syntax that contains each variable from the LHS at most once. -/ -def mkSimpleDelab (attrKind : Syntax) (pat qrhs : Syntax) : OptionT MacroM Syntax := do - match qrhs with - | `($c:ident $args*) => - let [(c, [])] ← Macro.resolveGlobalName c.getId | failure - /- - Try to remove all non semantic parenthesis. Since the parenthesizer - runs after appUnexpanders we should not match on parenthesis that the user - syntax inserted here for example the right hand side of: - notation "{" x "|" p "}" => setOf (fun x => p) - Should be matched as: setOf fun x => p - -/ - let args ← args.mapM (liftM ∘ removeParentheses) - /- - The user could mention the same antiquotation from the lhs multiple - times on the rhs, this heuristic does not support this. - -/ - let dup := hasDuplicateAntiquot args - guard !dup - -- replace head constant with (unused) antiquotation so we're not dependent on the exact pretty printing of the head - -- The reference is attached to the syntactic representation of the called function itself, not the entire function application - `(@[$attrKind:attrKind appUnexpander $(mkIdent c):ident] - aux_def unexpand $(mkIdent c) : Lean.PrettyPrinter.Unexpander := fun - | `($$f:ident $args*) => withRef f `($pat) - | _ => throw ()) - | `($c:ident) => - let [(c, [])] ← Macro.resolveGlobalName c.getId | failure - `(@[$attrKind:attrKind appUnexpander $(mkIdent c):ident] - aux_def unexpand $(mkIdent c) : Lean.PrettyPrinter.Unexpander := fun - | `($$f:ident) => withRef f `($pat) - | _ => throw ()) - | _ => failure +def mkSimpleDelab (attrKind : TSyntax ``attrKind) (pat qrhs : Term) : OptionT MacroM Syntax := do + let (c, args) ← match qrhs with + | `($c:ident $args*) => pure (c, args) + | `($c:ident) => pure (c, #[]) + | _ => failure + let [(c, [])] ← Macro.resolveGlobalName c.getId | failure + /- + Try to remove all non semantic parenthesis. Since the parenthesizer + runs after appUnexpanders we should not match on parenthesis that the user + syntax inserted here for example the right hand side of: + notation "{" x "|" p "}" => setOf (fun x => p) + Should be matched as: setOf fun x => p + -/ + let args ← liftM <| args.mapM removeParentheses + /- + The user could mention the same antiquotation from the lhs multiple + times on the rhs, this heuristic does not support this. + -/ + guard !hasDuplicateAntiquot args + -- replace head constant with antiquotation so we're not dependent on the exact pretty printing of the head + -- The reference is attached to the syntactic representation of the called function itself, not the entire function application + let lhs ← `($$f:ident) + let lhs := Syntax.mkApp lhs (.mk args) + `(@[$attrKind:attrKind appUnexpander $(mkIdent c):ident] + aux_def unexpand $(mkIdent c) : Lean.PrettyPrinter.Unexpander := fun + | `($lhs) => withRef f `($pat) + | _ => throw ()) private def isLocalAttrKind (attrKind : Syntax) : Bool := match attrKind with @@ -113,7 +105,7 @@ private def isLocalAttrKind (attrKind : Syntax) : Bool := | _ => false private def expandNotationAux (ref : Syntax) - (currNamespace : Name) (attrKind : Syntax) (prec? : Option Syntax) (name? : Option Syntax) (prio? : Option Syntax) (items : Array Syntax) (rhs : Syntax) : MacroM Syntax := do + (currNamespace : Name) (attrKind : TSyntax ``attrKind) (prec? : Option Prec) (name? : Option Ident) (prio? : Option Prio) (items : Array (TSyntax ``notationItem)) (rhs : Term) : MacroM Syntax := do let prio ← evalOptPrio prio? -- build parser let syntaxParts ← items.mapM expandNotationItemIntoSyntaxItem @@ -123,22 +115,25 @@ private def expandNotationAux (ref : Syntax) | some name => pure name.getId | none => mkNameFromParserSyntax `term (mkNullNode syntaxParts) -- build macro rules - let vars := items.filter fun item => item.getKind == `Lean.Parser.Command.identPrec - let vars := vars.map fun var => var[0] - let qrhs := antiquote vars rhs + let vars := items.filter fun item => item.raw.getKind == `Lean.Parser.Command.identPrec + let vars := vars.map fun var => var.raw[0] + let qrhs := ⟨antiquote vars rhs⟩ let patArgs ← items.mapM expandNotationItemIntoPattern /- The command `syntax [] ...` adds the current namespace to the syntax node kind. So, we must include current namespace when we create a pattern for the following `macro_rules` commands. -/ let fullName := currNamespace ++ name - let pat := mkNode fullName patArgs + let pat : Term := ⟨mkNode fullName patArgs⟩ let stxDecl ← `($attrKind:attrKind syntax $[: $prec?]? (name := $(mkIdent name)) (priority := $(quote prio):num) $[$syntaxParts]* : $cat) - let mut macroDecl ← `(macro_rules | `($pat) => ``($qrhs)) - if isLocalAttrKind attrKind then - -- Make sure the quotation pre-checker takes section variables into account for local notation. - macroDecl ← `(section set_option quotPrecheck.allowSectionVars true $macroDecl end) + let macroDecl ← `(macro_rules | `($pat) => ``($qrhs)) + let macroDecls ← + if isLocalAttrKind attrKind then + -- Make sure the quotation pre-checker takes section variables into account for local notation. + `(section set_option quotPrecheck.allowSectionVars true $macroDecl end) + else + pure ⟨mkNullNode #[macroDecl]⟩ match (← mkSimpleDelab attrKind pat qrhs |>.run) with - | some delabDecl => return mkNullNode #[stxDecl, macroDecl, delabDecl] - | none => return mkNullNode #[stxDecl, macroDecl] + | some delabDecl => return mkNullNode #[stxDecl, macroDecls, delabDecl] + | none => return mkNullNode #[stxDecl, macroDecls] @[builtinMacro Lean.Parser.Command.notation] def expandNotation : Macro | stx@`($attrKind:attrKind notation $[: $prec? ]? $[(name := $name?)]? $[(priority := $prio?)]? $items* => $rhs) => do diff --git a/stage0/src/Lean/Elab/PatternVar.lean b/stage0/src/Lean/Elab/PatternVar.lean index a5e399037e9..08daafdb4a0 100644 --- a/stage0/src/Lean/Elab/PatternVar.lean +++ b/stage0/src/Lean/Elab/PatternVar.lean @@ -11,7 +11,7 @@ namespace Lean.Elab.Term open Meta -abbrev PatternVar := Syntax -- TODO: should be `TSyntax identKind` +abbrev PatternVar := Syntax -- TODO: should be `Ident` /- Patterns define new local variables. @@ -60,7 +60,7 @@ An application in a pattern can be -/ structure Context where - funId : Syntax + funId : Ident ctorVal? : Option ConstructorVal -- It is `some`, if constructor application explicit : Bool ellipsis : Bool @@ -68,7 +68,7 @@ structure Context where paramDeclIdx : Nat := 0 namedArgs : Array NamedArg args : List Arg - newArgs : Array Syntax := #[] + newArgs : Array Term := #[] deriving Inhabited private def isDone (ctx : Context) : Bool := @@ -109,7 +109,7 @@ private def processVar (idStx : Syntax) : M Syntax := do modify fun s => { s with vars := s.vars.push idStx, found := s.found.insert id } return idStx -private def nameToPattern : Name → TermElabM Syntax +private def nameToPattern : Name → TermElabM Term | Name.anonymous => `(Name.anonymous) | Name.str p s _ => do let p ← nameToPattern p; `(Name.str $p $(quote s) _) | Name.num p n _ => do let p ← nameToPattern p; `(Name.num $p $(quote n) _) @@ -128,6 +128,7 @@ private def samePatternsVariables (startingAt : Nat) (s₁ s₂ : State) : Bool else false +open TSyntax.Compat in partial def collect (stx : Syntax) : M Syntax := withRef stx <| withFreshMacroScope do let k := stx.getKind if k == identKind then diff --git a/stage0/src/Lean/Elab/PreDefinition/Main.lean b/stage0/src/Lean/Elab/PreDefinition/Main.lean index 88c72934127..c4da4a4e5f3 100644 --- a/stage0/src/Lean/Elab/PreDefinition/Main.lean +++ b/stage0/src/Lean/Elab/PreDefinition/Main.lean @@ -124,7 +124,7 @@ def addPreDefinitions (preDefs : Array PreDefinition) (hints : TerminationHints) wf? := some wf terminationBy := terminationBy.markAsUsed (preDefs.map (·.declName)) if let some { ref, value := decrTactic } := decreasingBy.find? (preDefs.map (·.declName)) then - decrTactic? := some (← withRef ref `(by $decrTactic)) + decrTactic? := some (← withRef ref `(by $(⟨decrTactic⟩))) decreasingBy := decreasingBy.markAsUsed (preDefs.map (·.declName)) if wf?.isSome || decrTactic?.isSome then wfRecursion preDefs wf? decrTactic? diff --git a/stage0/src/Lean/Elab/Print.lean b/stage0/src/Lean/Elab/Print.lean index 1fb84deb744..addcb935594 100644 --- a/stage0/src/Lean/Elab/Print.lean +++ b/stage0/src/Lean/Elab/Print.lean @@ -83,7 +83,7 @@ private def printId (id : Syntax) : CommandElabM Unit := do @[builtinCommandElab «print»] def elabPrint : CommandElab | `(#print%$tk $id:ident) => withRef tk <| printId id - | `(#print%$tk $s:str) => logInfoAt tk s.isStrLit?.get! + | `(#print%$tk $s:str) => logInfoAt tk s.getString | _ => throwError "invalid #print command" namespace CollectAxioms diff --git a/stage0/src/Lean/Elab/Quotation.lean b/stage0/src/Lean/Elab/Quotation.lean index d6994d609ca..55a2849d8fb 100644 --- a/stage0/src/Lean/Elab/Quotation.lean +++ b/stage0/src/Lean/Elab/Quotation.lean @@ -16,6 +16,7 @@ namespace Lean.Elab.Term.Quotation open Lean.Parser.Term open Lean.Syntax open Meta +open TSyntax.Compat /-- `C[$(e)]` ~> `let a := e; C[$a]`. Used in the implementation of antiquot splices. -/ private partial def floatOutAntiquotTerms (stx : Syntax) : StateT (Syntax → TermElabM Syntax) TermElabM Syntax := @@ -46,7 +47,7 @@ private def getSepFromSplice (splice : Syntax) : String := private def getSepStxFromSplice (splice : Syntax) : Syntax := Unhygienic.run do match getSepFromSplice splice with | "" => `(mkNullNode) -- sepByIdent uses the null node for separator-less enumerations - | sep => `(mkAtom $(.mkStrLit sep)) + | sep => `(mkAtom $(Syntax.mkStrLit sep)) partial def mkTuple : Array Syntax → TermElabM Syntax | #[] => `(Unit.unit) @@ -69,13 +70,13 @@ def resolveSectionVariable (sectionVars : NameMap Name) (id : Name) : List (Name loop extractionResult.name [] /-- Transform sequence of pushes and appends into acceptable code -/ -def ArrayStxBuilder := Sum (Array Syntax) Syntax +def ArrayStxBuilder := Sum (Array Term) Term namespace ArrayStxBuilder def empty : ArrayStxBuilder := Sum.inl #[] -def build : ArrayStxBuilder → Syntax +def build : ArrayStxBuilder → Term | Sum.inl elems => quote elems | Sum.inr arr => arr @@ -90,8 +91,8 @@ def append (b : ArrayStxBuilder) (arr : Syntax) (appendName := ``Array.append) : end ArrayStxBuilder -- Elaborate the content of a syntax quotation term -private partial def quoteSyntax : Syntax → TermElabM Syntax - | Syntax.ident _ rawVal val preresolved => do +private partial def quoteSyntax : Syntax → TermElabM Term + | Syntax.ident _ rawVal val preresolved => do if !hygiene.get (← getOptions) then return ← `(Syntax.ident info $(quote rawVal) $(quote val) $(quote preresolved)) -- Add global scopes at compilation time (now), add macro scope at runtime (in the quotation). @@ -134,7 +135,9 @@ private partial def quoteSyntax : Syntax → TermElabM Syntax | some x => Array.empty.push x | none => Array.empty) | `many => `(@TSyntaxArray.raw $(quote ks) $val) - | `sepBy => `(@TSepArray.elemsAndSeps $(quote ks) $(quote <| getSepFromSplice arg) $val) + | `sepBy => + let sep := quote <| getSepFromSplice arg + `(@TSepArray.elemsAndSeps $(quote ks) $sep $val) | k => throwErrorAt arg "invalid antiquotation suffix splice kind '{k}'" else if k == nullKind && isAntiquotSplice arg then let k := antiquotSpliceKind? arg @@ -148,7 +151,7 @@ private partial def quoteSyntax : Syntax → TermElabM Syntax | $[some $ids:ident],* => $(quote inner) | $[_%$ids],* => Array.empty) | _ => - let arr ← ids[:ids.size-1].foldrM (fun id arr => `(Array.zip $id $arr)) ids.back + let arr ← ids[:ids.size-1].foldrM (fun id arr => `(Array.zip $id:ident $arr)) ids.back `(Array.map (fun $(← mkTuple ids) => $(inner[0])) $arr) let arr ← if k == `sepBy then `(mkSepArray $arr $(getSepStxFromSplice arg)) @@ -234,7 +237,7 @@ elab_stx_quot Parser.Command.quot /- match -/ -- an "alternative" of patterns plus right-hand side -private abbrev Alt := List Syntax × Syntax +private abbrev Alt := List Term × Term /-- In a single match step, we match the first discriminant against the "head" of the first pattern of the first @@ -280,7 +283,7 @@ structure HeadInfo where -- compute compatibility of pattern with given head check onMatch (taken : HeadCheck) : MatchResult -- actually run the specified head check, with the discriminant bound to `discr` - doMatch (yes : (newDiscrs : List Syntax) → TermElabM Syntax) (no : TermElabM Syntax) : TermElabM Syntax + doMatch (yes : (newDiscrs : List Term) → TermElabM Term) (no : TermElabM Term) : TermElabM Term /-- Adapt alternatives that do not introduce new discriminants in `doMatch`, but are covered by those that do so. -/ private def noOpMatchAdaptPats : HeadCheck → Alt → Alt @@ -288,7 +291,7 @@ private def noOpMatchAdaptPats : HeadCheck → Alt → Alt | slice p s, (pats, rhs) => (List.replicate (p + 1 + s) (Unhygienic.run `(_)) ++ pats, rhs) | _, alt => alt -private def adaptRhs (fn : Syntax → TermElabM Syntax) : Alt → TermElabM Alt +private def adaptRhs (fn : Term → TermElabM Term) : Alt → TermElabM Alt | (pats, rhs) => return (pats, ← fn rhs) private partial def getHeadInfo (alt : Alt) : TermElabM HeadInfo := @@ -489,7 +492,7 @@ private def deduplicate (floatedLetDecls : Array Syntax) : Alt → TermElabM (Ar let rhs' ← `(rhs $vars*) pure (floatedLetDecls.push (← `(letDecl|rhs $vars:ident* := $rhs)), (pats, rhs')) -private partial def compileStxMatch (discrs : List Syntax) (alts : List Alt) : TermElabM Syntax := do +private partial def compileStxMatch (discrs : List Term) (alts : List Alt) : TermElabM Syntax := do trace[Elab.match_syntax] "match {discrs} with {alts}" match discrs, alts with | [], ([], rhs)::_ => pure rhs -- nothing left to match @@ -536,7 +539,7 @@ private partial def compileStxMatch (discrs : List Syntax) (alts : List Alt) : T withFreshMacroScope $ compileStxMatch (newDiscrs ++ discrs) yesAlts.toList) (no := withFreshMacroScope $ compileStxMatch (discr::discrs) nonExhaustiveAlts.toList) for d in floatedLetDecls do - stx ← `(let_delayed $d:letDecl; $stx) + stx ← `(let_delayed $(⟨d⟩):letDecl; $stx) `(let discr := $discr; $stx) | _, _ => unreachable! @@ -544,8 +547,8 @@ def match_syntax.expand (stx : Syntax) : TermElabM Syntax := do match stx with | `(match $[$discrs:term],* with $[| $[$patss],* => $rhss]*) => do if !patss.any (·.any (fun - | `($_@$pat) => pat.isQuot - | pat => pat.isQuot)) then + | `($_@$pat) => pat.raw.isQuot + | pat => pat.raw.isQuot)) then -- no quotations => fall back to regular `match` throwUnsupportedSyntax let stx ← compileStxMatch discrs.toList (patss.map (·.toList) |>.zip rhss).toList diff --git a/stage0/src/Lean/Elab/Quotation/Precheck.lean b/stage0/src/Lean/Elab/Quotation/Precheck.lean index b701cf71a6b..c4e6c42a3e5 100644 --- a/stage0/src/Lean/Elab/Quotation/Precheck.lean +++ b/stage0/src/Lean/Elab/Quotation/Precheck.lean @@ -109,7 +109,7 @@ private def isSectionVariable (e : Expr) : TermElabM Bool := do @[builtinQuotPrecheck Lean.Parser.Term.app] def precheckApp : Precheck | `($f $args*) => do precheck f - for arg in args do + for arg in args.raw do match arg with | `(argument| ($_ := $e)) => precheck e | `(argument| $e:term) => precheck e @@ -125,7 +125,7 @@ private def isSectionVariable (e : Expr) : TermElabM Bool := do | `(($e)) => precheck e | `(($e, $es,*)) => do precheck e - es.getElems.forM precheck + es.getElems.raw.forM precheck | _ => throwUnsupportedSyntax @[builtinQuotPrecheck choice] def precheckChoice : Precheck := fun stx => do diff --git a/stage0/src/Lean/Elab/Quotation/Util.lean b/stage0/src/Lean/Elab/Quotation/Util.lean index d9870e416b2..dd52fa4fe2f 100644 --- a/stage0/src/Lean/Elab/Quotation/Util.lean +++ b/stage0/src/Lean/Elab/Quotation/Util.lean @@ -13,12 +13,12 @@ register_builtin_option hygiene : Bool := { descr := "Annotate identifiers in quotations such that they are resolved relative to the scope at their declaration, not that at their eventual use/expansion, to avoid accidental capturing. Note that quotations/notations already defined are unaffected." } -def getAntiquotationIds (stx : Syntax) : TermElabM (Array Syntax) := do +def getAntiquotationIds (stx : Syntax) : TermElabM (Array Ident) := do let mut ids := #[] for stx in stx.topDown (firstChoiceOnly := true) do if (isAntiquot stx || isTokenAntiquot stx) && !isEscapedAntiquot stx then let anti := getAntiquotTerm stx - if anti.isIdent then ids := ids.push anti + if anti.isIdent then ids := ids.push ⟨anti⟩ else if anti.isOfKind ``Parser.Term.hole then pure () else throwErrorAt stx "complex antiquotation not allowed here" return ids diff --git a/stage0/src/Lean/Elab/StructInst.lean b/stage0/src/Lean/Elab/StructInst.lean index a5062fd3202..6bc7837bc1e 100644 --- a/stage0/src/Lean/Elab/StructInst.lean +++ b/stage0/src/Lean/Elab/StructInst.lean @@ -13,6 +13,7 @@ namespace Lean.Elab.Term.StructInst open Std (HashMap) open Meta +open TSyntax.Compat /-- Structure instances are of the form: @@ -36,13 +37,11 @@ open Meta /-- Expand field abbreviations. Example: `{ x, y := 0 }` expands to `{ x := x, y := 0 }` -/ @[builtinMacro Lean.Parser.Term.structInst] def expandStructInstFieldAbbrev : Macro | `({ $[$srcs,* with]? $fields,* $[..%$ell]? $[: $ty]? }) => - if fields.getElems.any (·.getKind == ``Lean.Parser.Term.structInstFieldAbbrev) then do - let fieldsNew ← fields.getElems.mapM fun field => do - if field.getKind == ``Lean.Parser.Term.structInstFieldAbbrev then - let id := field[0] - `(Lean.Parser.Term.structInstField| $id:ident := $id:ident) - else - return field + if fields.getElems.raw.any (·.getKind == ``Lean.Parser.Term.structInstFieldAbbrev) then do + let fieldsNew ← fields.getElems.mapM fun + | `(Parser.Term.structInstFieldAbbrev| $id:ident) => + `(Parser.Term.structInstField| $id:ident := $id:ident) + | field => return field `({ $[$srcs,* with]? $fieldsNew,* $[..%$ell]? $[: $ty]? }) else Macro.throwUnsupported diff --git a/stage0/src/Lean/Elab/Structure.lean b/stage0/src/Lean/Elab/Structure.lean index ae78e1857ab..6f14bce042a 100644 --- a/stage0/src/Lean/Elab/Structure.lean +++ b/stage0/src/Lean/Elab/Structure.lean @@ -19,6 +19,7 @@ import Lean.Elab.Binders namespace Lean.Elab.Command open Meta +open TSyntax.Compat /- Recall that the `structure command syntax is ``` @@ -180,7 +181,7 @@ private def expandFields (structStx : Syntax) (structModifiers : Modifiers) (str -- It is safe to reset the binders to a "null" node since there is no value to be elaborated let type ← `(forall $(binders.getArgs):bracketedBinder*, $type) let type ← `(autoParam $type $(mkIdentFrom tac name)) - pure (mkNullNode, some type) + pure (mkNullNode, some type.raw) else let (binders, type) := expandDeclSig fieldBinder[3] pure (binders, some type) diff --git a/stage0/src/Lean/Elab/Syntax.lean b/stage0/src/Lean/Elab/Syntax.lean index 13cedcb80c6..b668e8b9c6f 100644 --- a/stage0/src/Lean/Elab/Syntax.lean +++ b/stage0/src/Lean/Elab/Syntax.lean @@ -17,16 +17,17 @@ def expandOptPrecedence (stx : Syntax) : MacroM (Option Nat) := else return some (← evalPrec stx[0][1]) -private def mkParserSeq (ds : Array Syntax) : TermElabM Syntax := do +private def mkParserSeq (ds : Array (Term × Nat)) : TermElabM (Term × Nat) := do if ds.size == 0 then throwUnsupportedSyntax else if ds.size == 1 then pure ds[0] else - let mut r := ds[0] - for d in ds[1:ds.size] do + let mut (r, stackSum) := ds[0] + for (d, stackSz) in ds[1:ds.size] do r ← `(ParserDescr.binary `andthen $r $d) - return r + stackSum := stackSum + stackSz + return (r, stackSum) structure ToParserDescrContext where catName : Name @@ -36,12 +37,20 @@ structure ToParserDescrContext where behavior : Parser.LeadingIdentBehavior abbrev ToParserDescrM := ReaderT ToParserDescrContext (StateRefT (Option Nat) TermElabM) +abbrev ToParserDescr := ToParserDescrM (Term × Nat) private def markAsTrailingParser (lhsPrec : Nat) : ToParserDescrM Unit := set (some lhsPrec) @[inline] private def withNotFirst {α} (x : ToParserDescrM α) : ToParserDescrM α := withReader (fun ctx => { ctx with first := false }) x -@[inline] private def withNestedParser {α} (x : ToParserDescrM α) : ToParserDescrM α := +def ensureUnaryOutput (x : Term × Nat) : Term := + let (stx, stackSz) := x + if stackSz != 1 then + Unhygienic.run ``(ParserDescr.unary $(quote `group) $stx) + else + stx + +@[inline] private def withNestedParser (x : ToParserDescr) : ToParserDescr := do withReader (fun ctx => { ctx with leftRec := false, first := false }) x def checkLeftRec (stx : Syntax) : ToParserDescrM Bool := do @@ -78,16 +87,17 @@ def resolveParserName [Monad m] [MonadInfoTree m] [MonadResolveName m] [MonadEnv | _ => none catch _ => return [] +open TSyntax.Compat in /-- Given a `stx` of category `syntax`, return a pair `(newStx, lhsPrec?)`, where `newStx` is of category `term`. After elaboration, `newStx` should have type `TrailingParserDescr` if `lhsPrec?.isSome`, and `ParserDescr` otherwise. -/ -partial def toParserDescr (stx : Syntax) (catName : Name) : TermElabM (Syntax × Option Nat) := do +partial def toParserDescr (stx : Syntax) (catName : Name) : TermElabM ((Term × Nat) × Option Nat) := do let env ← getEnv let behavior := Parser.leadingIdentBehavior env catName (process stx { catName := catName, first := true, leftRec := true, behavior := behavior }).run none where - process (stx : Syntax) : ToParserDescrM Syntax := withRef stx do + process (stx : Syntax) : ToParserDescr := withRef stx do let kind := stx.getKind if kind == nullKind then processSeq stx @@ -98,9 +108,9 @@ where else if kind == ``Lean.Parser.Syntax.cat then processNullaryOrCat stx else if kind == ``Lean.Parser.Syntax.unary then - processUnary stx + processAlias stx[0] #[stx[2]] else if kind == ``Lean.Parser.Syntax.binary then - processBinary stx + processAlias stx[0] #[stx[2], stx[4]] else if kind == ``Lean.Parser.Syntax.sepBy then processSepBy stx else if kind == ``Lean.Parser.Syntax.sepBy1 then @@ -137,12 +147,39 @@ where throwErrorAt stx "invalid atomic left recursive syntax" let prec? ← liftMacroM <| expandOptPrecedence stx[1] let prec := prec?.getD 0 - `(ParserDescr.cat $(quote catName) $(quote prec)) - + return (← `(ParserDescr.cat $(quote catName) $(quote prec)), 1) + + processAlias (id : Syntax) (args : Array Syntax) := do + let aliasName := id.getId.eraseMacroScopes + let info ← Parser.getParserAliasInfo aliasName + let args ← args.mapM (withNestedParser ∘ process) + let (args, stackSz) := if let some stackSz := info.stackSz? then + if !info.autoGroupArgs then + (args.map (·.1), stackSz) + else + (args.map ensureUnaryOutput, stackSz) + else + let (args, stackSzs) := args.unzip + (args, stackSzs.foldl (· + ·) 0) + let stx ← match args with + | #[] => Parser.ensureConstantParserAlias aliasName; ``(ParserDescr.const $(quote aliasName)) + | #[p1] => Parser.ensureUnaryParserAlias aliasName; ``(ParserDescr.unary $(quote aliasName) $p1) + | #[p1, p2] => Parser.ensureBinaryParserAlias aliasName; ``(ParserDescr.binary $(quote aliasName) $p1 $p2) + | _ => unreachable! + return (stx, stackSz) + processNullaryOrCat (stx : Syntax) := do match (← resolveParserName stx[0]) with - | [(c, true)] => ensureNoPrec stx; return mkIdentFrom stx c - | [(c, false)] => ensureNoPrec stx; `(ParserDescr.parser $(quote c)) + | [(c, true)] => + ensureNoPrec stx + -- `syntax _ :=` at least enforces this + let stackSz := 1 + return (mkIdentFrom stx c, stackSz) + | [(c, false)] => + ensureNoPrec stx + -- as usual, we assume that people using `Parser` know what they are doing + let stackSz := 1 + return (← `(ParserDescr.parser $(quote c)), stackSz) | cs@(_ :: _ :: _) => throwError "ambiguous parser declaration {cs.map (·.1)}" | [] => let id := stx[0].getId.eraseMacroScopes @@ -150,37 +187,23 @@ where processParserCategory stx else if (← Parser.isParserAlias id) then ensureNoPrec stx - Parser.ensureConstantParserAlias id - `(ParserDescr.const $(quote id)) + processAlias stx[0] #[] else throwError "unknown parser declaration/category/alias '{id}'" - processUnary (stx : Syntax) := do - let aliasName := (stx[0].getId).eraseMacroScopes - Parser.ensureUnaryParserAlias aliasName - let d ← withNestedParser do process stx[2] - `(ParserDescr.unary $(quote aliasName) $d) - - processBinary (stx : Syntax) := do - let aliasName := (stx[0].getId).eraseMacroScopes - Parser.ensureBinaryParserAlias aliasName - let d₁ ← withNestedParser do process stx[2] - let d₂ ← withNestedParser do process stx[4] - `(ParserDescr.binary $(quote aliasName) $d₁ $d₂) - processSepBy (stx : Syntax) := do - let p ← withNestedParser $ process stx[1] + let p ← ensureUnaryOutput <$> withNestedParser do process stx[1] let sep := stx[3] - let psep ← if stx[4].isNone then `(ParserDescr.symbol $sep) else process stx[4][1] + let psep ← if stx[4].isNone then `(ParserDescr.symbol $sep) else ensureUnaryOutput <$> withNestedParser do process stx[4][1] let allowTrailingSep := !stx[5].isNone - `(ParserDescr.sepBy $p $sep $psep $(quote allowTrailingSep)) + return (← `(ParserDescr.sepBy $p $sep $psep $(quote allowTrailingSep)), 1) processSepBy1 (stx : Syntax) := do - let p ← withNestedParser do process stx[1] + let p ← ensureUnaryOutput <$> withNestedParser do process stx[1] let sep := stx[3] - let psep ← if stx[4].isNone then `(ParserDescr.symbol $sep) else process stx[4][1] + let psep ← if stx[4].isNone then `(ParserDescr.symbol $sep) else ensureUnaryOutput <$> withNestedParser do process stx[4][1] let allowTrailingSep := !stx[5].isNone - `(ParserDescr.sepBy1 $p $sep $psep $(quote allowTrailingSep)) + return (← `(ParserDescr.sepBy1 $p $sep $psep $(quote allowTrailingSep)), 1) isValidAtom (s : String) : Bool := !s.isEmpty && @@ -197,14 +220,14 @@ where /- For syntax categories where initialized with `LeadingIdentBehavior` different from default (e.g., `tactic`), we automatically mark the first symbol as nonReserved. -/ if (← read).behavior != Parser.LeadingIdentBehavior.default && (← read).first then - `(ParserDescr.nonReservedSymbol $(quote atom) false) + return (← `(ParserDescr.nonReservedSymbol $(quote atom) false), 1) else - `(ParserDescr.symbol $(quote atom)) + return (← `(ParserDescr.symbol $(quote atom)), 1) | none => throwUnsupportedSyntax processNonReserved (stx : Syntax) := do match stx[1].isStrLit? with - | some atom => `(ParserDescr.nonReservedSymbol $(quote atom) false) + | some atom => return (← `(ParserDescr.nonReservedSymbol $(quote atom) false), 1) | none => throwUnsupportedSyntax @@ -318,7 +341,7 @@ def resolveSyntaxKind (k : Name) : CommandElabM Name := do let prio ← liftMacroM <| evalOptPrio prio? let stxNodeKind := (← getCurrNamespace) ++ name let catParserId := mkIdentFrom stx (cat.appendAfter "Parser") - let (val, lhsPrec?) ← runTermElabM none fun _ => Term.toParserDescr syntaxParser cat + let ((val, _), lhsPrec?) ← runTermElabM none fun _ => Term.toParserDescr syntaxParser cat let declName := mkIdentFrom stx name let d ← if let some lhsPrec := lhsPrec? then `($[$doc?:docComment]? @[$attrKind:attrKind $catParserId:ident $(quote prio):num] def $declName:ident : Lean.TrailingParserDescr := @@ -332,7 +355,7 @@ def resolveSyntaxKind (k : Name) : CommandElabM Name := do @[builtinCommandElab «syntaxAbbrev»] def elabSyntaxAbbrev : CommandElab := fun stx => do let `($[$doc?:docComment]? syntax $declName:ident := $[$ps:stx]*) ← pure stx | throwUnsupportedSyntax -- TODO: nonatomic names - let (val, _) ← runTermElabM none fun _ => Term.toParserDescr (mkNullNode ps) Name.anonymous + let ((val, _), _) ← runTermElabM none fun _ => Term.toParserDescr (mkNullNode ps) Name.anonymous let stxNodeKind := (← getCurrNamespace) ++ declName.getId let stx' ← `($[$doc?:docComment]? def $declName:ident : Lean.ParserDescr := ParserDescr.nodeWithAntiquot $(quote (toString declName.getId)) $(quote stxNodeKind) $val) withMacroExpansion stx stx' <| elabCommand stx' @@ -340,9 +363,9 @@ def resolveSyntaxKind (k : Name) : CommandElabM Name := do def checkRuleKind (given expected : SyntaxNodeKind) : Bool := given == expected || given == expected ++ `antiquot -def inferMacroRulesAltKind : Syntax → CommandElabM SyntaxNodeKind +def inferMacroRulesAltKind : TSyntax ``matchAlt → CommandElabM SyntaxNodeKind | `(matchAltExpr| | $pat:term => $_) => do - if !pat.isQuot then + if !pat.raw.isQuot then throwUnsupportedSyntax let quoted := getQuotContent pat pure quoted.getKind @@ -351,7 +374,7 @@ def inferMacroRulesAltKind : Syntax → CommandElabM SyntaxNodeKind /-- Infer syntax kind `k` from first pattern, put alternatives of same kind into new `macro/elab_rules (kind := k)` via `mkCmd (some k)`, leave remaining alternatives (via `mkCmd none`) to be recursively expanded. -/ -def expandNoKindMacroRulesAux (alts : Array Syntax) (cmdName : String) (mkCmd : Option Name → Array Syntax → CommandElabM Syntax) : CommandElabM Syntax := do +def expandNoKindMacroRulesAux (alts : Array (TSyntax ``matchAlt)) (cmdName : String) (mkCmd : Option Name → Array (TSyntax ``matchAlt) → CommandElabM Command) : CommandElabM Command := do let mut k ← inferMacroRulesAltKind alts[0] if k.isStr && k.getString! == "antiquot" then k := k.getPrefix @@ -364,7 +387,7 @@ def expandNoKindMacroRulesAux (alts : Array Syntax) (cmdName : String) (mkCmd : if altsNotK.isEmpty then mkCmd k altsK else - return mkNullNode #[← mkCmd k altsK, ← mkCmd none altsNotK] + `($(← mkCmd k altsK):command $(← mkCmd none altsNotK)) def strLitToPattern (stx: Syntax) : MacroM Syntax := match stx.isStrLit? with diff --git a/stage0/src/Lean/Elab/Tactic/BuiltinTactic.lean b/stage0/src/Lean/Elab/Tactic/BuiltinTactic.lean index 9e566b6ef58..9937e1fb8d7 100644 --- a/stage0/src/Lean/Elab/Tactic/BuiltinTactic.lean +++ b/stage0/src/Lean/Elab/Tactic/BuiltinTactic.lean @@ -222,7 +222,7 @@ where return (fvars, [mvarId]) withMainContext do for stx in ids, fvar in fvars do - Term.addLocalVarInfo stx (mkFVar fvar) + Term.addLocalVarInfo stx.raw (mkFVar fvar) | _ => throwUnsupportedSyntax @[builtinTactic Lean.Parser.Tactic.revert] def evalRevert : Tactic := fun stx => @@ -365,11 +365,8 @@ where let goals ← getGoals let goalsMsg := MessageData.joinSep (goals.map MessageData.ofGoal) m!"\n\n" match stx with - | `(tactic| fail) => throwError "tactic 'fail' failed\n{goalsMsg}" - | `(tactic| fail $msg) => - match msg.isStrLit? with - | none => throwIllFormedSyntax - | some msg => throwError "{msg}\n{goalsMsg}" + | `(tactic| fail) => throwError "tactic 'fail' failed\n{goalsMsg}" + | `(tactic| fail $msg:str) => throwError "{msg.getString}\n{goalsMsg}" | _ => throwUnsupportedSyntax @[builtinTactic dbgTrace] def evalDbgTrace : Tactic := fun stx => do diff --git a/stage0/src/Lean/Elab/Tactic/Conv/Congr.lean b/stage0/src/Lean/Elab/Tactic/Conv/Congr.lean index 4714ee0a81a..4c920dec196 100644 --- a/stage0/src/Lean/Elab/Tactic/Conv/Congr.lean +++ b/stage0/src/Lean/Elab/Tactic/Conv/Congr.lean @@ -98,7 +98,7 @@ private def selectIdx (tacticName : String) (mvarIds : List (Option MVarId)) (i @[builtinTactic Lean.Parser.Tactic.Conv.arg] def evalArg : Tactic := fun stx => do match stx with | `(conv| arg $[@%$tk?]? $i:num) => - let i := i.isNatLit?.getD 0 + let i := i.getNat if i == 0 then throwError "invalid 'arg' conv tactic, index must be greater than 0" let i := i - 1 diff --git a/stage0/src/Lean/Elab/Tactic/Match.lean b/stage0/src/Lean/Elab/Tactic/Match.lean index d724c8d71ec..30ca7a03da8 100644 --- a/stage0/src/Lean/Elab/Tactic/Match.lean +++ b/stage0/src/Lean/Elab/Tactic/Match.lean @@ -11,8 +11,9 @@ import Lean.Elab.Tactic.Induction namespace Lean.Elab.Tactic open Meta +open TSyntax.Compat in open Parser.Tactic in -private def mkAuxiliaryMatchTerm (parentTag : Name) (matchTac : Syntax) : MacroM (Syntax × Array Syntax) := do +private def mkAuxiliaryMatchTerm (parentTag : Name) (matchTac : Syntax) : MacroM (Term × Array Syntax) := do let matchAlts := matchTac[5] let alts := matchAlts[0].getArgs let mut newAlts := #[] @@ -35,9 +36,9 @@ private def mkAuxiliaryMatchTerm (parentTag : Name) (matchTac : Syntax) : MacroM alt := alt.setArg 3 newHole else let newHole ← withFreshMacroScope `(?rhs) - let newHoleId := newHole[1] + let newHoleId := newHole.raw[1] let newCase ← `(tactic| - case $newHoleId =>%$(alt[2]) + case $newHoleId:ident =>%$(alt[2]) -- annotate `| ... =>` with state after `case` with_annotate_state $(mkNullNode #[alt[0], alt[2]]) skip $holeOrTacticSeq) diff --git a/stage0/src/Lean/Elab/Tactic/Simp.lean b/stage0/src/Lean/Elab/Tactic/Simp.lean index a4b3299e42d..830071a0f27 100644 --- a/stage0/src/Lean/Elab/Tactic/Simp.lean +++ b/stage0/src/Lean/Elab/Tactic/Simp.lean @@ -13,6 +13,7 @@ import Lean.Elab.Tactic.Config namespace Lean.Elab.Tactic open Meta +open TSyntax.Compat declare_config_elab elabSimpConfigCore Meta.Simp.Config declare_config_elab elabSimpConfigCtxCore Meta.Simp.ConfigCtx @@ -174,24 +175,26 @@ private def elabSimpArgs (stx : Syntax) (ctx : Simp.Context) (eraseLocal : Bool) throwUnsupportedSyntax return { ctx := { ctx with simpTheorems := thmsArray.set! 0 thms }, starArg } where - resolveSimpIdTheorem? (simpArgTerm : Syntax) : TacticM ResolveSimpIdResult := do + resolveSimpIdTheorem? (simpArgTerm : Term) : TacticM ResolveSimpIdResult := do let resolveExt (n : Name) : TacticM ResolveSimpIdResult := do if let some ext ← getSimpExtension? n then return .ext ext else return .none - if simpArgTerm.isIdent then + match simpArgTerm with + | `($id:ident) => try if let some e ← Term.resolveId? simpArgTerm (withInfo := true) then return .expr e else - resolveExt simpArgTerm.getId.eraseMacroScopes + resolveExt id.getId.eraseMacroScopes catch _ => - resolveExt simpArgTerm.getId.eraseMacroScopes - else if let some e ← Term.elabCDotFunctionAlias? simpArgTerm then - return .expr e - else - return .none + resolveExt id.getId.eraseMacroScopes + | _ => + if let some e ← Term.elabCDotFunctionAlias? simpArgTerm then + return .expr e + else + return .none structure MkSimpContextResult where ctx : Simp.Context diff --git a/stage0/src/Lean/Elab/Term.lean b/stage0/src/Lean/Elab/Term.lean index b0ccbce943d..81d7bd0b7e8 100644 --- a/stage0/src/Lean/Elab/Term.lean +++ b/stage0/src/Lean/Elab/Term.lean @@ -1102,7 +1102,7 @@ private def isExplicitApp (stx : Syntax) : Bool := Example: `fun {α} (a : α) => a` -/ private def isLambdaWithImplicit (stx : Syntax) : Bool := match stx with - | `(fun $binders* => $_) => binders.any fun b => b.isOfKind ``Lean.Parser.Term.implicitBinder || b.isOfKind `Lean.Parser.Term.instBinder + | `(fun $binders* => $_) => binders.raw.any fun b => b.isOfKind ``Lean.Parser.Term.implicitBinder || b.isOfKind `Lean.Parser.Term.instBinder | _ => false private partial def dropTermParens : Syntax → Syntax := fun stx => diff --git a/stage0/src/Lean/Exception.lean b/stage0/src/Lean/Exception.lean index 843ad266e8a..bd9839db9c3 100644 --- a/stage0/src/Lean/Exception.lean +++ b/stage0/src/Lean/Exception.lean @@ -106,17 +106,11 @@ syntax "throwError " (interpolatedStr(term) <|> term) : term syntax "throwErrorAt " term:max (interpolatedStr(term) <|> term) : term macro_rules - | `(throwError $msg) => - if msg.getKind == interpolatedStrKind then - `(Lean.throwError (m! $msg)) - else - `(Lean.throwError $msg) + | `(throwError $msg:interpolatedStr) => `(Lean.throwError (m! $msg)) + | `(throwError $msg:term) => `(Lean.throwError $msg) macro_rules - | `(throwErrorAt $ref $msg) => - if msg.getKind == interpolatedStrKind then - `(Lean.throwErrorAt $ref (m! $msg)) - else - `(Lean.throwErrorAt $ref $msg) + | `(throwErrorAt $ref $msg:interpolatedStr) => `(Lean.throwErrorAt $ref (m! $msg)) + | `(throwErrorAt $ref $msg:term) => `(Lean.throwErrorAt $ref $msg) end Lean diff --git a/stage0/src/Lean/Hygiene.lean b/stage0/src/Lean/Hygiene.lean index 228e7a15dcd..5b7ff6f28db 100644 --- a/stage0/src/Lean/Hygiene.lean +++ b/stage0/src/Lean/Hygiene.lean @@ -95,9 +95,10 @@ def sanitizeName (userName : Name) : StateM NameSanitizerState Name := do private partial def sanitizeSyntaxAux : Syntax → StateM NameSanitizerState Syntax | stx@(Syntax.ident _ _ n _) => do - mkIdentFrom stx <$> match (← get).userName2Sanitized.find? n with - | some n' => pure n' - | none => if n.hasMacroScopes then sanitizeName n else pure n + let n ← match (← get).userName2Sanitized.find? n with + | some n' => pure n' + | none => if n.hasMacroScopes then sanitizeName n else pure n + return mkIdentFrom stx n | Syntax.node info k args => Syntax.node info k <$> args.mapM sanitizeSyntaxAux | stx => pure stx diff --git a/stage0/src/Lean/Level.lean b/stage0/src/Lean/Level.lean index 4ec9e052d24..cd39ec318db 100644 --- a/stage0/src/Lean/Level.lean +++ b/stage0/src/Lean/Level.lean @@ -447,8 +447,8 @@ mutual | Result.imaxNode fs, r => parenIfFalse (Format.group <| "imax" ++ formatLst fs) r end -protected partial def Result.quote (r : Result) (prec : Nat) : Syntax := - let addParen (s : Syntax) := +protected partial def Result.quote (r : Result) (prec : Nat) : Syntax.Level := + let addParen (s : Syntax.Level) := if prec > 0 then Unhygienic.run `(level| ( $s )) else s match r with | Result.leaf n => Unhygienic.run `(level| $(mkIdent n):ident) @@ -469,11 +469,11 @@ instance : ToFormat Level where instance : ToString Level where toString u := Format.pretty (Level.format u) -protected def quote (u : Level) (prec : Nat := 0) : Syntax := +protected def quote (u : Level) (prec : Nat := 0) : Syntax.Level := (PP.toResult u).quote prec -instance : Quote Level where - quote u := Level.quote u +instance : Quote Level `level where + quote := Level.quote end Level diff --git a/stage0/src/Lean/Message.lean b/stage0/src/Lean/Message.lean index c530afaf9ed..45d15420845 100644 --- a/stage0/src/Lean/Message.lean +++ b/stage0/src/Lean/Message.lean @@ -294,6 +294,7 @@ instance : ToMessageData Level := ⟨MessageData.ofLevel⟩ instance : ToMessageData Name := ⟨MessageData.ofName⟩ instance : ToMessageData String := ⟨stringToMessageData⟩ instance : ToMessageData Syntax := ⟨MessageData.ofSyntax⟩ +instance : ToMessageData (TSyntax k) := ⟨(MessageData.ofSyntax ·)⟩ instance : ToMessageData Format := ⟨MessageData.ofFormat⟩ instance : ToMessageData MVarId := ⟨MessageData.ofGoal⟩ instance : ToMessageData MessageData := ⟨id⟩ diff --git a/stage0/src/Lean/Meta/Basic.lean b/stage0/src/Lean/Meta/Basic.lean index bc4ec31f8bf..93150ed8570 100644 --- a/stage0/src/Lean/Meta/Basic.lean +++ b/stage0/src/Lean/Meta/Basic.lean @@ -1027,7 +1027,11 @@ private def withNewMCtxDepthImp (x : MetaM α) : MetaM α := do /-- Save cache and `MetavarContext`, bump the `MetavarContext` depth, execute `x`, - and restore saved data. -/ + and restore saved data. + + Metavariable context depths are used to control which metavariables may be assigned in `isDefEq`. + See the docstring of `isDefEq` for more information. + -/ def withNewMCtxDepth : n α → n α := mapMetaM withNewMCtxDepthImp @@ -1343,7 +1347,17 @@ def isExprDefEq (t s : Expr) : MetaM Bool := trace[Meta.isDefEq] "{t} =?= {s} ... {if b then "success" else "failure"}" return b -/-- Determines whether two expressions are definitionally equal to each other. -/ +/-- Determines whether two expressions are definitionally equal to each other. + + To control how metavariables are assigned and unified, metavariables and their context have a "depth". + Given a metavariable `?m` and a `MetavarContext` `mctx`, `?m` is not assigned if `?m.depth != mctx.depth`. + The combinator `withNewMCtxDepth x` will bump the depth while executing `x`. + So, `withNewMCtxDepth (isDefEq a b)` is `isDefEq` without any mvar assignment happening + whereas `isDefEq a b` will assign any metavariables of the current depth in `a` and `b` to unify them. + + For matching (where only mvars in `b` should be assigned), we create the term inside the `withNewMCtxDepth`. + For an example, see [Lean.Meta.Simp.tryTheoremWithExtraArgs?](https://github.com/leanprover/lean4/blob/master/src/Lean/Meta/Tactic/Simp/Rewrite.lean#L100-L106) + -/ abbrev isDefEq (t s : Expr) : MetaM Bool := isExprDefEq t s diff --git a/stage0/src/Lean/Meta/DiscrTree.lean b/stage0/src/Lean/Meta/DiscrTree.lean index c9ee301aa8d..52158e5ffc0 100644 --- a/stage0/src/Lean/Meta/DiscrTree.lean +++ b/stage0/src/Lean/Meta/DiscrTree.lean @@ -215,7 +215,7 @@ private def isOffset (fName : Name) (e : Expr) : MetaM Bool := do `Nat.succ ?m |-> v`, and we are trying to retrieve the matches for `Expr.lit (Literal.natVal 1) _`. In this scenario, we want to retrieve `Nat.succ ?m |-> v` -/ private def shouldAddAsStar (fName : Name) (e : Expr) : MetaM Bool := do - if fName == `Nat.zero then + if fName == ``Nat.zero then return true else isOffset fName e diff --git a/stage0/src/Lean/Parser.lean b/stage0/src/Lean/Parser.lean index e9b849b1d37..fa19ed88911 100644 --- a/stage0/src/Lean/Parser.lean +++ b/stage0/src/Lean/Parser.lean @@ -19,28 +19,28 @@ open Lean.PrettyPrinter.Parenthesizer open Lean.PrettyPrinter.Formatter builtin_initialize - register_parser_alias "ws" checkWsBefore - register_parser_alias "noWs" checkNoWsBefore - register_parser_alias "linebreak" checkLinebreakBefore + register_parser_alias "ws" checkWsBefore { stackSz? := none } + register_parser_alias "noWs" checkNoWsBefore { stackSz? := none } + register_parser_alias "linebreak" checkLinebreakBefore { stackSz? := none } register_parser_alias (kind := numLitKind) "num" numLit register_parser_alias (kind := strLitKind) "str" strLit register_parser_alias (kind := charLitKind) "char" charLit register_parser_alias (kind := nameLitKind) "name" nameLit register_parser_alias (kind := scientificLitKind) "scientific" scientificLit register_parser_alias (kind := identKind) "ident" ident - register_parser_alias "colGt" checkColGt - register_parser_alias "colGe" checkColGe - register_parser_alias lookahead - register_parser_alias atomic + register_parser_alias "colGt" checkColGt { stackSz? := none } + register_parser_alias "colGe" checkColGe { stackSz? := none } + register_parser_alias lookahead { stackSz? := some 0 } + register_parser_alias atomic { stackSz? := none } register_parser_alias many register_parser_alias many1 register_parser_alias manyIndent register_parser_alias many1Indent - register_parser_alias optional - register_parser_alias withPosition + register_parser_alias optional { autoGroupArgs := false } + register_parser_alias withPosition { stackSz? := none } register_parser_alias (kind := interpolatedStrKind) interpolatedStr register_parser_alias orelse - register_parser_alias andthen + register_parser_alias andthen { stackSz? := none } registerAlias "notFollowedBy" (notFollowedBy · "element") Parenthesizer.registerAlias "notFollowedBy" notFollowedBy.parenthesizer diff --git a/stage0/src/Lean/Parser/Basic.lean b/stage0/src/Lean/Parser/Basic.lean index 584b5f576b4..d595ee7ed40 100644 --- a/stage0/src/Lean/Parser/Basic.lean +++ b/stage0/src/Lean/Parser/Basic.lean @@ -485,6 +485,8 @@ def mergeOrElseErrors (s : ParserState) (error1 : Error) (iniPos : String.Pos) ( else s | other => other +-- When `mergeAntiquots` is true, if `p` parses an antiquotation, we try `q` as well and return a choice node if they +-- both return antiquotations def orelseFnCore (p q : ParserFn) (mergeAntiquots : Bool) : ParserFn := fun c s => Id.run do let s0 := s let iniSz := s.stackSize @@ -1745,6 +1747,7 @@ def mkAntiquot (name : String) (kind : SyntaxNodeKind) (anonymous := true) (isPs @[inline] def withAntiquotFn (antiquotP p : ParserFn) : ParserFn := fun c s => -- fast check that is false in most cases if c.input.get s.pos == '$' then + -- do not allow antiquotation choice nodes here as `antiquotP` is the strictly more general antiquotation than any in `p` orelseFnCore (mergeAntiquots := false) antiquotP p c s else p c s diff --git a/stage0/src/Lean/Parser/Command.lean b/stage0/src/Lean/Parser/Command.lean index d2e0595b58e..4b8fc0c8c53 100644 --- a/stage0/src/Lean/Parser/Command.lean +++ b/stage0/src/Lean/Parser/Command.lean @@ -65,7 +65,8 @@ def whereStructInst := leading_parser " where" >> sepBy1Indent (ppGroup whereSt Remark: we should not use `Term.whereDecls` at `declVal` because `Term.whereDecls` is defined using `Term.letRecDecl` which may contain attributes. Issue #753 showns an example that fails to be parsed when we used `Term.whereDecls`. -/ -def declVal := declValSimple <|> declValEqns <|> whereStructInst +def declVal := withAntiquot (mkAntiquot "declVal" `Lean.Parser.Command.declVal (isPseudoKind := true)) <| + declValSimple <|> declValEqns <|> whereStructInst def «abbrev» := leading_parser "abbrev " >> declId >> ppIndent optDeclSig >> declVal def optDefDeriving := optional (atomic ("deriving " >> notSymbol "instance") >> sepBy1 ident ", ") def «def» := leading_parser "def " >> declId >> ppIndent optDeclSig >> declVal >> optDefDeriving >> terminationSuffix diff --git a/stage0/src/Lean/Parser/Extension.lean b/stage0/src/Lean/Parser/Extension.lean index 64dfa210000..3d0e7e9d788 100644 --- a/stage0/src/Lean/Parser/Extension.lean +++ b/stage0/src/Lean/Parser/Extension.lean @@ -202,14 +202,25 @@ def getBinaryAlias {α} (mapRef : IO.Ref (AliasTable α)) (aliasName : Name) : I abbrev ParserAliasValue := AliasValue Parser +structure ParserAliasInfo where + /-- Number of syntax nodes produced by this parser. `none` means "sum of input sizes". -/ + stackSz? : Option Nat := some 1 + /-- Whether arguments should be wrapped in `group(·)` if they do not produce exactly one syntax node. -/ + autoGroupArgs : Bool := stackSz?.isSome + builtin_initialize parserAliasesRef : IO.Ref (NameMap ParserAliasValue) ← IO.mkRef {} builtin_initialize parserAlias2kindRef : IO.Ref (NameMap SyntaxNodeKind) ← IO.mkRef {} +builtin_initialize parserAliases2infoRef : IO.Ref (NameMap ParserAliasInfo) ← IO.mkRef {} + +def getParserAliasInfo (aliasName : Name) : IO ParserAliasInfo := do + return (← parserAliases2infoRef.get).findD aliasName {} -- Later, we define macro `register_parser_alias` which registers a parser, formatter and parenthesizer -def registerAlias (aliasName : Name) (p : ParserAliasValue) (kind? : Option SyntaxNodeKind := none) : IO Unit := do +def registerAlias (aliasName : Name) (p : ParserAliasValue) (kind? : Option SyntaxNodeKind := none) (info : ParserAliasInfo := {}) : IO Unit := do registerAliasCore parserAliasesRef aliasName p if let some kind := kind? then parserAlias2kindRef.modify (·.insert aliasName kind) + parserAliases2infoRef.modify (·.insert aliasName info) instance : Coe Parser ParserAliasValue := { coe := AliasValue.const } instance : Coe (Parser → Parser) ParserAliasValue := { coe := AliasValue.unary } diff --git a/stage0/src/Lean/Parser/Extra.lean b/stage0/src/Lean/Parser/Extra.lean index 35f027026e4..bf5e0c35d13 100644 --- a/stage0/src/Lean/Parser/Extra.lean +++ b/stage0/src/Lean/Parser/Extra.lean @@ -169,27 +169,29 @@ attribute [runBuiltinParserAttributeHooks] ppHardSpace ppSpace ppLine ppGroup ppRealGroup ppRealFill ppIndent ppDedent ppAllowUngrouped ppDedentIfGrouped ppHardLineUnlessUngrouped -macro "register_parser_alias" kind?:group("(" &"kind" " := " term ")")? aliasName?:optional(strLit) declName:ident : term => do - let [(fullDeclName, [])] ← Macro.resolveGlobalName declName.getId | - Macro.throwError "expected non-overloaded constant name" - let aliasName := aliasName?.getD (Syntax.mkStrLit declName.getId.toString) - `(do Parser.registerAlias $aliasName $declName (kind? := some $(kind?.map (·[3]) |>.getD (quote fullDeclName))) - PrettyPrinter.Formatter.registerAlias $aliasName $(mkIdentFrom declName (declName.getId ++ `formatter)) - PrettyPrinter.Parenthesizer.registerAlias $aliasName $(mkIdentFrom declName (declName.getId ++ `parenthesizer))) +syntax "register_parser_alias" group("(" &"kind" " := " term ")")? (strLit)? ident (colGt term)? : term +macro_rules + | `(register_parser_alias $[(kind := $kind?)]? $(aliasName?)? $declName $(info?)?) => do + let [(fullDeclName, [])] ← Macro.resolveGlobalName declName.getId | + Macro.throwError "expected non-overloaded constant name" + let aliasName := aliasName?.getD (Syntax.mkStrLit declName.getId.toString) + `(do Parser.registerAlias $aliasName $declName $(info?.getD (Unhygienic.run `({}))) (kind? := some $(kind?.getD (quote fullDeclName))) + PrettyPrinter.Formatter.registerAlias $aliasName $(mkIdentFrom declName (declName.getId ++ `formatter)) + PrettyPrinter.Parenthesizer.registerAlias $aliasName $(mkIdentFrom declName (declName.getId ++ `parenthesizer))) builtin_initialize - register_parser_alias group - register_parser_alias ppHardSpace - register_parser_alias ppSpace - register_parser_alias ppLine - register_parser_alias ppGroup - register_parser_alias ppRealGroup - register_parser_alias ppRealFill - register_parser_alias ppIndent - register_parser_alias ppDedent - register_parser_alias ppAllowUngrouped - register_parser_alias ppDedentIfGrouped - register_parser_alias ppHardLineUnlessUngrouped + register_parser_alias group { autoGroupArgs := false } + register_parser_alias ppHardSpace { stackSz? := none } + register_parser_alias ppSpace { stackSz? := none } + register_parser_alias ppLine { stackSz? := none } + register_parser_alias ppGroup { stackSz? := none } + register_parser_alias ppRealGroup { stackSz? := none } + register_parser_alias ppRealFill { stackSz? := none } + register_parser_alias ppIndent { stackSz? := none } + register_parser_alias ppDedent { stackSz? := none } + register_parser_alias ppAllowUngrouped { stackSz? := none } + register_parser_alias ppDedentIfGrouped { stackSz? := none } + register_parser_alias ppHardLineUnlessUngrouped { stackSz? := none } end Parser diff --git a/stage0/src/Lean/Parser/Module.lean b/stage0/src/Lean/Parser/Module.lean index b0f9e11739e..5fb1cd38296 100644 --- a/stage0/src/Lean/Parser/Module.lean +++ b/stage0/src/Lean/Parser/Module.lean @@ -134,7 +134,7 @@ def testParseModule (env : Environment) (fname contents : String) : IO Syntax := let (header, state, messages) ← parseHeader inputCtx let cmds ← testParseModuleAux env inputCtx state messages #[] let stx := mkNode `Lean.Parser.Module.module #[header, mkListNode cmds] - pure stx.updateLeading + pure stx.raw.updateLeading def testParseFile (env : Environment) (fname : System.FilePath) : IO Syntax := do let contents ← IO.FS.readFile fname diff --git a/stage0/src/Lean/Parser/Term.lean b/stage0/src/Lean/Parser/Term.lean index 7bdf08f362b..0fc37ce678f 100644 --- a/stage0/src/Lean/Parser/Term.lean +++ b/stage0/src/Lean/Parser/Term.lean @@ -111,7 +111,7 @@ def strictImplicitLeftBracket := atomic (group (symbol "{" >> "{")) <|> "⦃" def strictImplicitRightBracket := atomic (group (symbol "}" >> "}")) <|> "⦄" def strictImplicitBinder (requireType := false) := ppGroup $ leading_parser strictImplicitLeftBracket >> many1 binderIdent >> binderType requireType >> strictImplicitRightBracket def instBinder := ppGroup $ leading_parser "[" >> optIdent >> termParser >> "]" -def bracketedBinder (requireType := false) := withAntiquot (mkAntiquot "bracketedBinder" `Lean.Parser.Term.bracketedBinder (anonymous := false) (isPseudoKind := true)) <| +def bracketedBinder (requireType := false) := withAntiquot (mkAntiquot "bracketedBinder" `Lean.Parser.Term.bracketedBinder (isPseudoKind := true)) <| explicitBinder requireType <|> strictImplicitBinder requireType <|> implicitBinder requireType <|> instBinder /- @@ -143,6 +143,9 @@ def matchAlt (rhsParser : Parser := termParser) : Parser := work with other `rhsParser`s (of arity 1). -/ def matchAltExpr := matchAlt +instance : Coe (TSyntax ``matchAltExpr) (TSyntax ``matchAlt) where + coe stx := ⟨stx.raw⟩ + def matchAlts (rhsParser : Parser := termParser) : Parser := leading_parser withPosition $ many1Indent (ppLine >> matchAlt rhsParser) @@ -157,10 +160,10 @@ def motive := leading_parser atomic ("(" >> nonReservedSymbol "motive" >> " := " @[builtinTermParser] def «match» := leading_parser:leadPrec "match " >> optional generalizingParam >> optional motive >> sepBy1 matchDiscr ", " >> " with " >> ppDedent matchAlts @[builtinTermParser] def «nomatch» := leading_parser:leadPrec "nomatch " >> termParser -def funImplicitBinder := atomic (lookahead ("{" >> many1 binderIdent >> (symbol " : " <|> "}"))) >> implicitBinder +def funImplicitBinder := withAntiquot (mkAntiquot "implicitBinder" ``implicitBinder) <| atomic (lookahead ("{" >> many1 binderIdent >> (symbol " : " <|> "}"))) >> implicitBinder def funStrictImplicitBinder := atomic (lookahead (strictImplicitLeftBracket >> many1 binderIdent >> (symbol " : " <|> strictImplicitRightBracket))) >> strictImplicitBinder -def funSimpleBinder := atomic (lookahead (many1 binderIdent >> " : ")) >> simpleBinder -def funBinder : Parser := funStrictImplicitBinder <|> funImplicitBinder <|> instBinder <|> funSimpleBinder <|> termParser maxPrec +def funSimpleBinder := withAntiquot (mkAntiquot "simpleBinder" ``simpleBinder) <| atomic (lookahead (many1 binderIdent >> " : ")) >> simpleBinder +def funBinder : Parser := withAntiquot (mkAntiquot "funBinder" `Lean.Parser.Term.funBinder (isPseudoKind := true)) (funStrictImplicitBinder <|> funImplicitBinder <|> instBinder <|> funSimpleBinder <|> termParser maxPrec) -- NOTE: we disable anonymous antiquotations to ensure that `fun $b => ...` remains a `term` antiquotation def basicFun : Parser := leading_parser (withAnonymousAntiquot := false) ppGroup (many1 (ppSpace >> funBinder) >> " =>") >> ppSpace >> termParser @[builtinTermParser] def «fun» := leading_parser:maxPrec ppAllowUngrouped >> unicodeSymbol "λ" "fun" >> (basicFun <|> matchAlts) @@ -180,8 +183,9 @@ def withAnonymousAntiquot := leading_parser atomic ("(" >> nonReservedSymbol "wi def simpleBinderWithoutType := nodeWithAntiquot "simpleBinder" ``simpleBinder (anonymous := true) (many1 binderIdent >> pushNone) +def letIdBinder := withAntiquot (mkAntiquot "letIdBinder" `Lean.Parser.Term.letIdBinder (isPseudoKind := true)) (simpleBinderWithoutType <|> bracketedBinder) /- Remark: we use `checkWsBefore` to ensure `let x[i] := e; b` is not parsed as `let x [i] := e; b` where `[i]` is an `instBinder`. -/ -def letIdLhs : Parser := ident >> notFollowedBy (checkNoWsBefore "" >> "[") "space is required before instance '[...]' binders to distinguish them from array updates `let x[i] := e; ...`" >> many (ppSpace >> (simpleBinderWithoutType <|> bracketedBinder)) >> optType +def letIdLhs : Parser := ident >> notFollowedBy (checkNoWsBefore "" >> "[") "space is required before instance '[...]' binders to distinguish them from array updates `let x[i] := e; ...`" >> many (ppSpace >> letIdBinder) >> optType def letIdDecl := leading_parser (withAnonymousAntiquot := false) atomic (letIdLhs >> " := ") >> termParser def letPatDecl := leading_parser (withAnonymousAntiquot := false) atomic (termParser >> pushNone >> optType >> " := ") >> termParser /- @@ -205,8 +209,11 @@ def letDecl := leading_parser (withAnonymousAntiquot := false) notFollowedBy -- `let`-declaration that is only included in the elaborated term if variable is still there @[builtinTermParser] def «let_tmp» := leading_parser:leadPrec withPosition ("let_tmp " >> letDecl) >> optSemicolon termParser +instance : Coe (TSyntax ``letIdBinder) (TSyntax ``funBinder) where + coe stx := ⟨stx⟩ -- `simpleBinderWithoutType` prevents using a proper quotation for this + -- like `let_fun` but with optional name -def haveIdLhs := optional (ident >> many (ppSpace >> (simpleBinderWithoutType <|> bracketedBinder))) >> optType +def haveIdLhs := optional (ident >> many (ppSpace >> letIdBinder)) >> optType def haveIdDecl := leading_parser (withAnonymousAntiquot := false) atomic (haveIdLhs >> " := ") >> termParser def haveEqnsDecl := leading_parser (withAnonymousAntiquot := false) haveIdLhs >> matchAlts def haveDecl := leading_parser (withAnonymousAntiquot := false) haveIdDecl <|> letPatDecl <|> haveEqnsDecl diff --git a/stage0/src/Lean/PrettyPrinter.lean b/stage0/src/Lean/PrettyPrinter.lean index 77bb2eade3d..e31f729b8bf 100644 --- a/stage0/src/Lean/PrettyPrinter.lean +++ b/stage0/src/Lean/PrettyPrinter.lean @@ -25,7 +25,7 @@ def ppTerm (stx : Syntax) : CoreM Format := do let stx := (sanitizeSyntax stx).run' { options := opts } parenthesizeTerm stx >>= formatTerm -def ppUsing (e : Expr) (delab : Expr → MetaM Syntax) : MetaM Format := do +def ppUsing (e : Expr) (delab : Expr → MetaM Term) : MetaM Format := do let lctx := (← getLCtx).sanitizeNames.run' { options := (← getOptions) } Meta.withLCtx lctx #[] do ppTerm (← delab e) diff --git a/stage0/src/Lean/PrettyPrinter/Delaborator/Basic.lean b/stage0/src/Lean/PrettyPrinter/Delaborator/Basic.lean index b2ee14cabbe..8c46b15c46f 100644 --- a/stage0/src/Lean/PrettyPrinter/Delaborator/Basic.lean +++ b/stage0/src/Lean/PrettyPrinter/Delaborator/Basic.lean @@ -61,7 +61,7 @@ structure State where builtin_initialize delabFailureId : InternalExceptionId ← registerInternalExceptionId `delabFailure abbrev DelabM := ReaderT Context (StateRefT State MetaM) -abbrev Delab := DelabM Syntax +abbrev Delab := DelabM Term instance : Inhabited (DelabM α) where default := throw default @@ -166,10 +166,10 @@ def withOptionAtCurrPos (k : Name) (v : DataValue) (x : DelabM α) : DelabM α : { ctx with optionsPerPos := ctx.optionsPerPos.insert pos opts' }) x -def annotatePos (pos : Pos) (stx : Syntax) : Syntax := - stx.setInfo (SourceInfo.synthetic ⟨pos⟩ ⟨pos⟩) +def annotatePos (pos : Pos) (stx : Term) : Term := + ⟨stx.raw.setInfo (SourceInfo.synthetic ⟨pos⟩ ⟨pos⟩)⟩ -def annotateCurPos (stx : Syntax) : Delab := +def annotateCurPos (stx : Term) : Delab := return annotatePos (← getPos) stx def getUnusedName (suggestion : Name) (body : Expr) : DelabM Name := do @@ -248,7 +248,7 @@ partial def delab : Delab := do else return ← ``(_) let k ← getExprKind - let stx ← delabFor k <|> (liftM $ show MetaM Syntax from throwError "don't know how to delaborate '{k}'") + let stx ← delabFor k <|> (liftM $ show MetaM _ from throwError "don't know how to delaborate '{k}'") if ← getPPOption getPPAnalyzeTypeAscriptions <&&> getPPOption getPPAnalysisNeedsType <&&> pure !e.isMData then let typeStx ← withType delab `(($stx:term : $typeStx:term)) >>= annotateCurPos @@ -274,7 +274,7 @@ end Delaborator open SubExpr (Pos) open Delaborator (OptionsPerPos topDownAnalyze) -def delabCore (e : Expr) (optionsPerPos : OptionsPerPos := {}) (delab := Delaborator.delab) : MetaM (Syntax × Std.RBMap Pos Elab.Info compare) := do +def delabCore (e : Expr) (optionsPerPos : OptionsPerPos := {}) (delab := Delaborator.delab) : MetaM (Term × Std.RBMap Pos Elab.Info compare) := do /- Using `erasePatternAnnotations` here is a bit hackish, but we do it `Expr.mdata` affects the delaborator. TODO: should we fix that? -/ let e ← Meta.erasePatternRefAnnotations e @@ -302,7 +302,7 @@ def delabCore (e : Expr) (optionsPerPos : OptionsPerPos := {}) (delab := Delabor return (stx, infos) /-- "Delaborate" the given term into surface-level syntax using the default and given subterm-specific options. -/ -def delab (e : Expr) (optionsPerPos : OptionsPerPos := {}) : MetaM Syntax := do +def delab (e : Expr) (optionsPerPos : OptionsPerPos := {}) : MetaM Term := do let (stx, _) ← delabCore e optionsPerPos return stx diff --git a/stage0/src/Lean/PrettyPrinter/Delaborator/Builtins.lean b/stage0/src/Lean/PrettyPrinter/Delaborator/Builtins.lean index dcec957fbb2..dc7c7e3c582 100644 --- a/stage0/src/Lean/PrettyPrinter/Delaborator/Builtins.lean +++ b/stage0/src/Lean/PrettyPrinter/Delaborator/Builtins.lean @@ -13,6 +13,7 @@ namespace Lean.PrettyPrinter.Delaborator open Lean.Meta open Lean.Parser.Term open SubExpr +open TSyntax.Compat def maybeAddBlockImplicit (ident : Syntax) : DelabM Syntax := do if ← getPPOption getPPAnalysisBlockImplicit then `(@$ident:ident) else pure ident @@ -165,7 +166,7 @@ def delabAppExplicit : Delab := do let (fnStx, _, argStxs) ← withAppFnArgs (do let stx ← withOptionAtCurrPos `pp.tagAppFns tagAppFn delabAppFn - let needsExplicit := stx.getKind != ``Lean.Parser.Term.explicit + let needsExplicit := stx.raw.getKind != ``Lean.Parser.Term.explicit let stx ← if needsExplicit then `(@$stx) else pure stx pure (stx, paramKinds.toList, #[])) (fun ⟨fnStx, paramKinds, argStxs⟩ => do @@ -273,7 +274,7 @@ def delabAppImplicit : Delab := do let v := param.defVal.get! if !v.hasLooseBVars && v == arg then pure none else delab else if !param.isRegularExplicit && param.defVal.isNone then - if ← getPPOption getPPAnalysisNamedArg <||> (pure (param.name == `motive) <&&> shouldShowMotive arg opts) then mkNamedArg param.name (← delab) else pure none + if ← getPPOption getPPAnalysisNamedArg <||> (pure (param.name == `motive) <&&> shouldShowMotive arg opts) then some <$> mkNamedArg param.name (← delab) else pure none else delab let argStxs := match argStx? with | none => argStxs @@ -293,17 +294,17 @@ structure AppMatchState where info : MatcherInfo matcherTy : Expr params : Array Expr := #[] - motive : Option (Syntax × Expr) := none + motive : Option (Term × Expr) := none motiveNamed : Bool := false - discrs : Array Syntax := #[] + discrs : Array Term := #[] varNames : Array (Array Name) := #[] - rhss : Array Syntax := #[] + rhss : Array Term := #[] -- additional arguments applied to the result of the `match` expression - moreArgs : Array Syntax := #[] + moreArgs : Array Term := #[] /-- Extract arguments of motive applications from the matcher type. For the example below: `#[#[`([])], #[`(a::as)]]` -/ -private partial def delabPatterns (st : AppMatchState) : DelabM (Array (Array Syntax)) := +private partial def delabPatterns (st : AppMatchState) : DelabM (Array (Array Term)) := withReader (fun ctx => { ctx with inPattern := true, optionsPerPos := {} }) do let ty ← instantiateForall st.matcherTy st.params -- need to reduce `let`s that are lifted into the matcher type @@ -704,7 +705,7 @@ def delabNamedPattern : Delab := do let p ← withAppFn $ withAppArg delab -- TODO: we should hide `h` if it has an inaccessible name and is not used in the rhs let h ← withAppArg delab - guard x.isIdent + guard x.raw.isIdent `($x:ident@$h:ident:$p:term) -- Sigma and PSigma delaborators diff --git a/stage0/src/Lean/PrettyPrinter/Parenthesizer.lean b/stage0/src/Lean/PrettyPrinter/Parenthesizer.lean index 6c7c656c247..24ac87e4ff7 100644 --- a/stage0/src/Lean/PrettyPrinter/Parenthesizer.lean +++ b/stage0/src/Lean/PrettyPrinter/Parenthesizer.lean @@ -334,17 +334,17 @@ def parserOfStack.parenthesizer (offset : Nat) (_prec : Nat := 0) : Parenthesize @[builtinCategoryParenthesizer term] def term.parenthesizer : CategoryParenthesizer | prec => do - maybeParenthesize `term true (fun stx => Unhygienic.run `(($stx))) prec $ + maybeParenthesize `term true (fun stx => Unhygienic.run `(($(⟨stx⟩)))) prec $ parenthesizeCategoryCore `term prec @[builtinCategoryParenthesizer tactic] def tactic.parenthesizer : CategoryParenthesizer | prec => do - maybeParenthesize `tactic false (fun stx => Unhygienic.run `(tactic|($stx))) prec $ + maybeParenthesize `tactic false (fun stx => Unhygienic.run `(tactic|($(⟨stx⟩)))) prec $ parenthesizeCategoryCore `tactic prec @[builtinCategoryParenthesizer level] def level.parenthesizer : CategoryParenthesizer | prec => do - maybeParenthesize `level false (fun stx => Unhygienic.run `(level|($stx))) prec $ + maybeParenthesize `level false (fun stx => Unhygienic.run `(level|($(⟨stx⟩)))) prec $ parenthesizeCategoryCore `level prec @[builtinCategoryParenthesizer rawStx] diff --git a/stage0/src/Lean/ResolveName.lean b/stage0/src/Lean/ResolveName.lean index 84525a46cdf..a713d9c7c74 100644 --- a/stage0/src/Lean/ResolveName.lean +++ b/stage0/src/Lean/ResolveName.lean @@ -36,10 +36,18 @@ builtin_initialize aliasExtension : SimplePersistentEnvExtension AliasEntry Alia def getAliasState (env : Environment) : AliasState := aliasExtension.getState env -def getAliases (env : Environment) (a : Name) : List Name := +/- + Retrieve aliases for `a`. If `skipProtected` is `true`, then the resulting list only includes + declarations that are not marked as `proctected`. +-/ +def getAliases (env : Environment) (a : Name) (skipProtected : Bool) : List Name := match aliasExtension.getState env |>.find? a with | none => [] - | some es => es + | some es => + if skipProtected then + es.filter (!isProtected env ·) + else + es -- slower, but only used in the pretty printer def getRevAliases (env : Environment) (e : Name) : List Name := @@ -51,7 +59,8 @@ namespace ResolveName /- Check whether `ns ++ id` is a valid namepace name and/or there are aliases names `ns ++ id`. -/ private def resolveQualifiedName (env : Environment) (ns : Name) (id : Name) : List Name := let resolvedId := ns ++ id - let resolvedIds := getAliases env resolvedId + -- We ignore protected aliases if `id` is atomic. + let resolvedIds := getAliases env resolvedId (skipProtected := id.isAtomic) if env.contains resolvedId && (!id.isAtomic || !isProtected env resolvedId) then resolvedId :: resolvedIds else @@ -122,7 +131,7 @@ def resolveGlobalName (env : Environment) (ns : Name) (openDecls : List OpenDecl let idPrv := mkPrivateName env id let resolvedIds := if env.contains idPrv then [idPrv] ++ resolvedIds else resolvedIds let resolvedIds := resolveOpenDecls env id openDecls resolvedIds - let resolvedIds := getAliases env id ++ resolvedIds + let resolvedIds := getAliases env id (skipProtected := id.isAtomic) ++ resolvedIds match resolvedIds with | _ :: _ => resolvedIds.eraseDups.map fun id => (id, projs) | [] => loop p (s::projs) diff --git a/stage0/src/Lean/Server/FileWorker.lean b/stage0/src/Lean/Server/FileWorker.lean index 972920c3b09..ce9c819f3f2 100644 --- a/stage0/src/Lean/Server/FileWorker.lean +++ b/stage0/src/Lean/Server/FileWorker.lean @@ -61,19 +61,15 @@ structure WorkerContext where hIn : FS.Stream hOut : FS.Stream hLog : FS.Stream - srcSearchPath : SearchPath + headerTask : Task (Except Error (Snapshot × SearchPath)) initParams : InitializeParams clientHasWidgets : Bool /- Asynchronous snapshot elaboration. -/ section Elab structure AsyncElabState where - headerSnap : Snapshot snaps : Array Snapshot - private def AsyncElabState.lastSnap (s : AsyncElabState) : Snapshot := - s.snaps.getD (s.snaps.size - 1) s.headerSnap - abbrev AsyncElabM := StateT AsyncElabState <| EIO ElabTaskError -- Placed here instead of Lean.Server.Utils because of an import loop @@ -95,13 +91,13 @@ section Elab : AsyncElabM (Option Snapshot) := do cancelTk.check let s ← get - let lastSnap := s.lastSnap + let lastSnap := s.snaps.back if lastSnap.isAtEnd then publishDiagnostics m lastSnap.diagnostics.toArray ctx.hOut publishProgressDone m ctx.hOut -- This will overwrite existing ilean info for the file, in case something -- went wrong during the incremental updates. - publishIleanInfoFinal m ctx.hOut <| s.snaps.insertAt 0 s.headerSnap + publishIleanInfoFinal m ctx.hOut s.snaps return none publishProgressAtPos m lastSnap.endPos ctx.hOut let snap ← compileNextCmd m.mkInputContext lastSnap ctx.clientHasWidgets @@ -124,22 +120,22 @@ section Elab publishIleanInfoUpdate m ctx.hOut #[snap] return some snap - /-- Elaborates all commands after the last snap (using `headerSnap` if `snaps` - is empty), emitting the diagnostics into `hOut`. -/ - def unfoldCmdSnaps (m : DocumentMeta) (headerSnap : Snapshot) (snaps : Array Snapshot) (cancelTk : CancelToken) + /-- Elaborates all commands after the last snap (at least the header snap is assumed to exist), emitting the diagnostics into `hOut`. -/ + def unfoldCmdSnaps (m : DocumentMeta) (snaps : Array Snapshot) (cancelTk : CancelToken) : ReaderT WorkerContext IO (AsyncList ElabTaskError Snapshot) := do let ctx ← read - if snaps.isEmpty && headerSnap.msgLog.hasErrors then + let headerSnap := snaps[0] + if headerSnap.msgLog.hasErrors then -- Treat header processing errors as fatal so users aren't swamped with -- followup errors publishProgressAtPos m headerSnap.beginPos ctx.hOut (kind := LeanFileProgressKind.fatalError) publishIleanInfoFinal m ctx.hOut #[headerSnap] - pure AsyncList.nil + return AsyncList.ofList [headerSnap] else -- This will overwrite existing ilean info for the file since this has a -- higher version number. - publishIleanInfoUpdate m ctx.hOut <| snaps.insertAt 0 headerSnap - AsyncList.unfoldAsync (nextCmdSnap ctx m cancelTk) { headerSnap, snaps } + publishIleanInfoUpdate m ctx.hOut snaps + return AsyncList.ofList snaps.toList ++ (← AsyncList.unfoldAsync (nextCmdSnap ctx m cancelTk) { snaps }) end Elab -- Pending requests are tracked so they can be cancelled @@ -248,18 +244,20 @@ section Initialization def initializeWorker (meta : DocumentMeta) (i o e : FS.Stream) (initParams : InitializeParams) (opts : Options) : IO (WorkerContext × WorkerState) := do let clientHasWidgets := initParams.initializationOptions?.bind (·.hasWidgets?) |>.getD false - let (headerSnap, srcSearchPath) ← compileHeader meta o opts (hasWidgets := clientHasWidgets) + let headerTask ← EIO.asTask <| compileHeader meta o opts (hasWidgets := clientHasWidgets) let cancelTk ← CancelToken.new let ctx := { hIn := i hOut := o hLog := e - srcSearchPath + headerTask initParams clientHasWidgets } - let cmdSnaps ← unfoldCmdSnaps meta headerSnap #[] cancelTk ctx - let doc : EditableDocument := ⟨meta, headerSnap, cmdSnaps, cancelTk⟩ + let cmdSnaps ← EIO.mapTask (t := headerTask) (match · with + | Except.ok (s, _) => unfoldCmdSnaps meta #[s] cancelTk ctx + | Except.error e => throw (e : ElabTaskError)) + let doc : EditableDocument := ⟨meta, AsyncList.asyncTail cmdSnaps, cancelTk⟩ return (ctx, { doc := doc pendingRequests := RBMap.empty @@ -273,31 +271,28 @@ section Updates /-- Given the new document, updates editable doc state. -/ def updateDocument (newMeta : DocumentMeta) : WorkerM Unit := do - -- The watchdog only restarts the file worker when the syntax tree of the header changes. - -- If e.g. a newline is deleted, it will not restart this file worker, but we still - -- need to reparse the header so that the offsets are correct. let ctx ← read let oldDoc := (←get).doc - let newHeaderSnap ← reparseHeader newMeta.mkInputContext oldDoc.headerSnap - if newHeaderSnap.stx != oldDoc.headerSnap.stx then - throwServerError "Internal server error: header changed but worker wasn't restarted." let ⟨cmdSnaps, e?⟩ ← oldDoc.cmdSnaps.updateFinishedPrefix match e? with - -- This case should not be possible. only the main task aborts tasks and ensures that aborted tasks - -- do not show up in `snapshots` of an EditableDocument. | some ElabTaskError.aborted => + -- This case should not be possible. only the main task aborts tasks and ensures that aborted tasks + -- do not show up in `snapshots` of an EditableDocument. throwServerError "Internal server error: elab task was aborted while still in use." | some (ElabTaskError.ioError ioError) => throw ioError | _ => -- No error or EOF + -- The watchdog only restarts the file worker when the semantic content of the header changes. + -- If e.g. a newline is deleted, it will not restart this file worker, but we still + -- need to reparse the header so that the offsets are correct. + let (newHeaderStx, newMpState, _) ← Parser.parseHeader newMeta.mkInputContext + let newHeaderSnap := { cmdSnaps.finishedPrefix.head! with stx := newHeaderStx, mpState := newMpState } oldDoc.cancelTk.set let changePos := oldDoc.meta.text.source.firstDiffPos newMeta.text.source -- NOTE(WN): we invalidate eagerly as `endPos` consumes input greedily. To re-elaborate only -- when really necessary, we could do a whitespace-aware `Syntax` comparison instead. let mut validSnaps := cmdSnaps.finishedPrefix.takeWhile (fun s => s.endPos < changePos) - if validSnaps.length = 0 then - let cancelTk ← CancelToken.new - let newCmdSnaps ← unfoldCmdSnaps newMeta newHeaderSnap #[] cancelTk ctx - modify fun st => { st with doc := ⟨newMeta, newHeaderSnap, newCmdSnaps, cancelTk⟩ } + if validSnaps.length ≤ 1 then + validSnaps := [newHeaderSnap] else /- When at least one valid non-header snap exists, it may happen that a change does not fall within the syntactic range of that last snap but still modifies it by appending tokens. @@ -311,10 +306,9 @@ section Updates let newLastStx ← parseNextCmd newMeta.mkInputContext preLastSnap if newLastStx != lastSnap.stx then validSnaps := validSnaps.dropLast - let cancelTk ← CancelToken.new - let newSnaps ← unfoldCmdSnaps newMeta newHeaderSnap validSnaps.toArray cancelTk ctx - let newCmdSnaps := AsyncList.ofList validSnaps ++ newSnaps - modify fun st => { st with doc := ⟨newMeta, newHeaderSnap, newCmdSnaps, cancelTk⟩ } + let cancelTk ← CancelToken.new + let newSnaps ← unfoldCmdSnaps newMeta validSnaps.toArray cancelTk ctx + modify fun st => { st with doc := ⟨newMeta, newSnaps, cancelTk⟩ } end Updates /- Notifications are handled in the main thread. They may change global worker state @@ -405,23 +399,26 @@ section MessageHandling message := toString e } return - let rc : RequestContext := - { rpcSessions := st.rpcSessions - srcSearchPath := ctx.srcSearchPath - doc := st.doc - hLog := ctx.hLog - initParams := ctx.initParams } - let t? ← EIO.toIO' <| handleLspRequest method params rc - let t₁ ← match t? with - | Except.error e => - IO.asTask do - ctx.hOut.writeLspResponseError <| e.toLspResponseError id - | Except.ok t => (IO.mapTask · t) fun - | Except.ok resp => - ctx.hOut.writeLspResponse ⟨id, resp⟩ - | Except.error e => - ctx.hOut.writeLspResponseError <| e.toLspResponseError id - queueRequest id t₁ + -- we assume that every request requires at least the header snapshot or the search path + let t ← IO.bindTask ctx.headerTask fun x => do + let (_, srcSearchPath) ← IO.ofExcept x + let rc : RequestContext := + { rpcSessions := st.rpcSessions + srcSearchPath + doc := st.doc + hLog := ctx.hLog + initParams := ctx.initParams } + let t? ← EIO.toIO' <| handleLspRequest method params rc + let t₁ ← match t? with + | Except.error e => + IO.asTask do + ctx.hOut.writeLspResponseError <| e.toLspResponseError id + | Except.ok t => (IO.mapTask · t) fun + | Except.ok resp => + ctx.hOut.writeLspResponse ⟨id, resp⟩ + | Except.error e => + ctx.hOut.writeLspResponseError <| e.toLspResponseError id + queueRequest id t end MessageHandling section MainLoop @@ -491,8 +488,6 @@ def workerMain (opts : Options) : IO UInt32 := do let o ← IO.getStdout let e ← IO.getStderr try - let seed ← (UInt64.toNat ∘ ByteArray.toUInt64LE!) <$> IO.getRandomBytes 8 - IO.setRandSeed seed let exitCode ← initAndRunWorker i o e opts -- HACK: all `Task`s are currently "foreground", i.e. we join on them on main thread exit, but we definitely don't -- want to do that in the case of the worker processes, which can produce non-terminating tasks evaluating user code diff --git a/stage0/src/Lean/Server/FileWorker/RequestHandling.lean b/stage0/src/Lean/Server/FileWorker/RequestHandling.lean index 505c189fc15..e7fb5101443 100644 --- a/stage0/src/Lean/Server/FileWorker/RequestHandling.lean +++ b/stage0/src/Lean/Server/FileWorker/RequestHandling.lean @@ -226,7 +226,7 @@ partial def handleDocumentHighlight (p : DocumentHighlightParams) withWaitFindSnap doc (fun s => s.endPos > pos) (notFoundX := pure #[]) fun snap => do - let (snaps, _) ← doc.allSnaps.updateFinishedPrefix + let (snaps, _) ← doc.cmdSnaps.updateFinishedPrefix if let some his := highlightRefs? snaps.finishedPrefix.toArray then return his if let some hi := highlightReturn? none snap.stx then @@ -247,7 +247,7 @@ partial def handleDocumentSymbol (_ : DocumentSymbolParams) throw (e : RequestError) | _ => pure () - let lastSnap := cmdSnaps.finishedPrefix.getLastD doc.headerSnap + let lastSnap := cmdSnaps.finishedPrefix.getLast! stxs := stxs ++ (← parseAhead doc.meta.mkInputContext lastSnap).toList let (syms, _) := toDocumentSymbols doc.meta.text stxs return { syms := syms.toArray } @@ -256,7 +256,7 @@ partial def handleDocumentSymbol (_ : DocumentSymbolParams) | [] => ([], []) | stx::stxs => match stx with | `(namespace $id) => sectionLikeToDocumentSymbols text stx stxs (id.getId.toString) SymbolKind.namespace id - | `(section $(id)?) => sectionLikeToDocumentSymbols text stx stxs ((·.getId.toString) <$> id |>.getD "
") SymbolKind.namespace (id.getD stx) + | `(section $(id)?) => sectionLikeToDocumentSymbols text stx stxs ((·.getId.toString) <$> id |>.getD "
") SymbolKind.namespace (id.map (·.raw) |>.getD stx) | `(end $(_id)?) => ([], stx::stxs) | _ => Id.run do let (syms, stxs') := toDocumentSymbols text stxs @@ -265,7 +265,7 @@ partial def handleDocumentSymbol (_ : DocumentSymbolParams) if let some stxRange := stx.getRange? then let (name, selection) := match stx with | `($_:declModifiers $_:attrKind instance $[$np:namedPrio]? $[$id:ident$[.{$ls,*}]?]? $sig:declSig $_) => - ((·.getId.toString) <$> id |>.getD s!"instance {sig.reprint.getD ""}", id.getD sig) + ((·.getId.toString) <$> id |>.getD s!"instance {sig.raw.reprint.getD ""}", id.map (·.raw) |>.getD sig) | _ => match stx[1][1] with | `(declId|$id:ident$[.{$ls,*}]?) => (id.getId.toString, id) | _ => (stx[1][0].isIdOrAtom?.getD "", stx[1][0]) @@ -402,7 +402,7 @@ def handleSemanticTokensRange (p : SemanticTokensRangeParams) partial def handleFoldingRange (_ : FoldingRangeParams) : RequestM (RequestTask (Array FoldingRange)) := do let doc ← readDoc - let t ← doc.allSnaps.waitAll + let t ← doc.cmdSnaps.waitAll mapTask t fun (snaps, _) => do let stxs := snaps.map (·.stx) let (_, ranges) ← StateT.run (addRanges doc.meta.text [] stxs) #[] @@ -423,7 +423,7 @@ partial def handleFoldingRange (_ : FoldingRangeParams) addRanges text sections stxs | `(mutual $body* end) => do addRangeFromSyntax text FoldingRangeKind.region stx - addRanges text [] body.toList + addRanges text [] body.raw.toList addRanges text sections stxs | _ => do if isImport stx then @@ -444,7 +444,7 @@ partial def handleFoldingRange (_ : FoldingRangeParams) -- separately to the main definition. -- We never fold other modifiers, such as annotations. if let `($dm:declModifiers $decl) := stx then - if let some comment := dm[0].getOptional? then + if let some comment := dm.raw[0].getOptional? then addRangeFromSyntax text FoldingRangeKind.comment comment addRangeFromSyntax text FoldingRangeKind.region decl diff --git a/stage0/src/Lean/Server/FileWorker/Utils.lean b/stage0/src/Lean/Server/FileWorker/Utils.lean index bde91f58bb9..ef5996e9187 100644 --- a/stage0/src/Lean/Server/FileWorker/Utils.lean +++ b/stage0/src/Lean/Server/FileWorker/Utils.lean @@ -50,17 +50,12 @@ and parser state after each command so that edits can be applied without recompiling code appearing earlier in the file. -/ structure EditableDocument where meta : DocumentMeta - /- The first snapshot is that after the header. -/ - headerSnap : Snapshot - /- Subsequent snapshots occur after each command. -/ + /- State snapshots after header and each command. -/ cmdSnaps : AsyncList ElabTaskError Snapshot cancelTk : CancelToken namespace EditableDocument -def allSnaps (doc : EditableDocument) : AsyncList ElabTaskError Snapshot := - AsyncList.cons doc.headerSnap doc.cmdSnaps - end EditableDocument structure RpcSession where diff --git a/stage0/src/Lean/Server/InfoUtils.lean b/stage0/src/Lean/Server/InfoUtils.lean index 0252bf99562..944613a61c1 100644 --- a/stage0/src/Lean/Server/InfoUtils.lean +++ b/stage0/src/Lean/Server/InfoUtils.lean @@ -154,8 +154,6 @@ partial def InfoTree.hoverableInfoAt? (t : InfoTree) (hoverPos : String.Pos) (in else if i matches .ofTermInfo { expr := .fvar .., .. } then 0 -- prefer results for constants over variables (which overlap at declaration names) else 1 - -- must prioritize smaller nodes because pattern infos are not properly nested - let priority := (-(r.stop - r.start |>.byteIdx : Int), priority) [(priority, ctx, i)] else results) |>.getD [] diff --git a/stage0/src/Lean/Server/Requests.lean b/stage0/src/Lean/Server/Requests.lean index e4588cf6c59..dd905307b86 100644 --- a/stage0/src/Lean/Server/Requests.lean +++ b/stage0/src/Lean/Server/Requests.lean @@ -108,7 +108,7 @@ def withWaitFindSnap (doc : EditableDocument) (p : Snapshot → Bool) (notFoundX : RequestM β) (x : Snapshot → RequestM β) : RequestM (RequestTask β) := do - let findTask ← doc.allSnaps.waitFind? p + let findTask ← doc.cmdSnaps.waitFind? p mapTask findTask <| waitFindSnapAux notFoundX x /-- See `withWaitFindSnap`. -/ @@ -116,7 +116,7 @@ def bindWaitFindSnap (doc : EditableDocument) (p : Snapshot → Bool) (notFoundX : RequestM (RequestTask β)) (x : Snapshot → RequestM (RequestTask β)) : RequestM (RequestTask β) := do - let findTask ← doc.allSnaps.waitFind? p + let findTask ← doc.cmdSnaps.waitFind? p bindTask findTask <| waitFindSnapAux notFoundX x end RequestM diff --git a/stage0/src/Lean/Server/Rpc/Deriving.lean b/stage0/src/Lean/Server/Rpc/Deriving.lean index e0e4af003cc..d4173b4d04e 100644 --- a/stage0/src/Lean/Server/Rpc/Deriving.lean +++ b/stage0/src/Lean/Server/Rpc/Deriving.lean @@ -88,7 +88,7 @@ end def isOptField (n : Name) : Bool := n.toString.endsWith "?" -private def deriveStructureInstance (indVal : InductiveVal) (params : Array Expr) : TermElabM Syntax := +private def deriveStructureInstance (indVal : InductiveVal) (params : Array Expr) : TermElabM Command := withFields indVal params fun fields => do trace[Elab.Deriving.RpcEncoding] "for structure {indVal.name} with params {params}" -- Postulate that every field have a rpc encoding, storing the encoding type ident @@ -96,9 +96,9 @@ private def deriveStructureInstance (indVal : InductiveVal) (params : Array Expr -- as otherwise typeclass synthesis fails. let mut binders := #[] let mut fieldIds := #[] - let mut fieldEncIds : Array Syntax := #[] - let mut uniqFieldEncIds : Array Syntax := #[] - let mut fieldEncIds' : DiscrTree Syntax := {} + let mut fieldEncIds : Array Term := #[] + let mut uniqFieldEncIds : Array Ident := #[] + let mut fieldEncIds' : DiscrTree Ident := {} for (fieldName, fieldTp) in fields do let mut fieldTp := fieldTp if isOptField fieldName then @@ -108,14 +108,14 @@ private def deriveStructureInstance (indVal : InductiveVal) (params : Array Expr -- postulate that the field has an encoding and remember the encoding's binder name fieldIds := fieldIds.push <| mkIdent fieldName - let mut fieldEncId := Syntax.missing + let mut fieldEncId : Ident := ⟨Syntax.missing⟩ match (← fieldEncIds'.getMatch fieldTp).back? with | none => fieldEncId ← mkIdent <$> mkFreshUserName fieldName - binders := binders.push <| ← `(Deriving.explicitBinderF| ( $fieldEncId:ident )) + binders := binders.push (← `(bracketedBinder| ( $fieldEncId:ident ))) let stx ← PrettyPrinter.delab fieldTp - binders := binders.push <| - ← `(Deriving.instBinderF| [ $(mkIdent ``Lean.Server.RpcEncoding) $stx $fieldEncId:ident ]) + binders := binders.push + (← `(bracketedBinder| [ $(mkIdent ``Lean.Server.RpcEncoding) $stx $fieldEncId:ident ])) fieldEncIds' ← fieldEncIds'.insert fieldTp fieldEncId uniqFieldEncIds := uniqFieldEncIds.push fieldEncId | some fid => fieldEncId := fid @@ -143,7 +143,7 @@ private def deriveStructureInstance (indVal : InductiveVal) (params : Array Expr $[($fieldIds : $fieldEncIds)]* deriving $(mkIdent ``FromJson), $(mkIdent ``ToJson) - instance : $(mkIdent ``RpcEncoding) $typeId $packetAppliedId:ident where + instance : $(mkIdent ``RpcEncoding) $typeId $packetAppliedId where rpcEncode a := return { $[$encInits],* } @@ -157,17 +157,17 @@ private structure CtorState where encArgTypes : DiscrTree Name := {} uniqEncArgTypes : Array Name := #[] -- binders for `encArgTypes` as well as the relevant `RpcEncoding`s - binders : Array Syntax := #[] + binders : Array (TSyntax ``Parser.Term.bracketedBinder) := #[] -- the syntax of each constructor in the packet - ctors : Array Syntax := #[] + ctors : Array (TSyntax ``Parser.Command.ctor) := #[] -- syntax of each arm of the `rpcEncode` pattern-match - encodes : Array Syntax := #[] + encodes : Array (TSyntax ``Parser.Term.matchAlt) := #[] -- syntax of each arm of the `rpcDecode` pattern-match - decodes : Array Syntax := #[] + decodes : Array (TSyntax ``Parser.Term.matchAlt) := #[] deriving Inhabited private def matchF := Lean.Parser.Term.matchAlt (rhsParser := Lean.Parser.termParser) -private def deriveInductiveInstance (indVal : InductiveVal) (params : Array Expr) : TermElabM Syntax := do +private def deriveInductiveInstance (indVal : InductiveVal) (params : Array Expr) : TermElabM Command := do trace[Elab.Deriving.RpcEncoding] "for inductive {indVal.name} with params {params}" -- produce all encoding types and binders for them @@ -186,8 +186,8 @@ private def deriveInductiveInstance (indVal : InductiveVal) (params : Array Expr acc := { acc with encArgTypes := ← acc.encArgTypes.insert argTp tid uniqEncArgTypes := acc.uniqEncArgTypes.push tid binders := acc.binders.append #[ - ← `(Deriving.explicitBinderF| ( $(mkIdent tid):ident )), - ← `(Deriving.instBinderF| [ $(mkIdent ``Lean.Server.RpcEncoding) $argTpStx $(mkIdent tid):ident ]) + (← `(bracketedBinder| ( $(mkIdent tid):ident ))), + (← `(bracketedBinder| [ $(mkIdent ``Lean.Server.RpcEncoding) $argTpStx $(mkIdent tid):ident ])) ] } return acc @@ -214,7 +214,7 @@ private def deriveInductiveInstance (indVal : InductiveVal) (params : Array Expr -- create encoder and decoder match arms let nms ← argVars.mapM fun _ => mkIdent <$> mkFreshBinderName let mkPattern (src : Name) := Syntax.mkApp (mkIdent <| Name.mkStr src ctor.getString!) nms - let mkBody (tgt : Name) (func : Name) : TermElabM Syntax := do + let mkBody (tgt : Name) (func : Name) : TermElabM Term := do let items ← nms.mapM fun nm => `(← $(mkIdent func) $nm) let tm := Syntax.mkApp (mkIdent <| Name.mkStr tgt ctor.getString!) items `(return $tm:term) @@ -223,13 +223,13 @@ private def deriveInductiveInstance (indVal : InductiveVal) (params : Array Expr let decArm ← `(matchF| | $(mkPattern packetNm):term => $(← mkBody indVal.name ``rpcDecode)) return { acc with ctors := acc.ctors.push pktCtor - encodes := acc.encodes.push encArm - decodes := acc.decodes.push decArm } + encodes := acc.encodes.push ⟨encArm⟩ + decodes := acc.decodes.push ⟨decArm⟩ } -- helpers for type name syntax let paramIds ← params.mapM fun p => return mkIdent (← getFVarLocalDecl p).userName let typeId := Syntax.mkApp (← `(@$(mkIdent indVal.name))) paramIds - let packetAppliedId := Syntax.mkApp (mkIdent packetNm) (st.uniqEncArgTypes.map mkIdent) + let packetAppliedId := Syntax.mkApp (mkIdent packetNm) (st.uniqEncArgTypes.map (mkIdent ·)) `(variable $st.binders* @@ -237,7 +237,7 @@ private def deriveInductiveInstance (indVal : InductiveVal) (params : Array Expr $[$(st.ctors):ctor]* deriving $(mkIdent ``FromJson), $(mkIdent ``ToJson) - instance : $(mkIdent ``RpcEncoding) $typeId:ident $packetAppliedId:ident where + instance : $(mkIdent ``RpcEncoding) $typeId $packetAppliedId where rpcEncode := fun x => match x with $[$(st.encodes):matchAlt]* rpcDecode := fun x => match x with @@ -264,11 +264,11 @@ private def deriveInstance (typeName : Name) : CommandElabM Bool := do let mut binders := #[] for param in params do let paramNm := (←getFVarLocalDecl param).userName - binders := binders.push <| ← `(Deriving.explicitBinderF| ( $(mkIdent paramNm) )) + binders := binders.push (← `(bracketedBinder| ( $(mkIdent paramNm) ))) -- only look for encodings for `Type` parameters if !(← inferType param).isType then continue - binders := binders.push <| - ← `(Deriving.instBinderF| [ $(mkIdent ``Lean.Server.RpcEncoding) $(mkIdent paramNm) _ ]) + binders := binders.push + (← `(bracketedBinder| [ $(mkIdent ``Lean.Server.RpcEncoding) $(mkIdent paramNm) _ ])) return #[ ← `(section), @@ -281,10 +281,10 @@ private def deriveInstance (typeName : Name) : CommandElabM Bool := do ] for cmd in cmds do - elabCommand cmd + elabCommand cmd.raw return true -private unsafe def dispatchDeriveInstanceUnsafe (declNames : Array Name) (args? : Option Syntax) : CommandElabM Bool := do +private unsafe def dispatchDeriveInstanceUnsafe (declNames : Array Name) (args? : Option (TSyntax ``Parser.Term.structInst)) : CommandElabM Bool := do if declNames.size ≠ 1 then return false let args ← @@ -301,7 +301,7 @@ private unsafe def dispatchDeriveInstanceUnsafe (declNames : Array Name) (args? deriveInstance declNames[0] @[implementedBy dispatchDeriveInstanceUnsafe] -private opaque dispatchDeriveInstance (declNames : Array Name) (args? : Option Syntax) : CommandElabM Bool +private opaque dispatchDeriveInstance (declNames : Array Name) (args? : Option (TSyntax ``Parser.Term.structInst)) : CommandElabM Bool builtin_initialize Elab.registerBuiltinDerivingHandlerWithArgs ``RpcEncoding dispatchDeriveInstance diff --git a/stage0/src/Lean/Server/Snapshots.lean b/stage0/src/Lean/Server/Snapshots.lean index 79f16b8630e..713b3953e66 100644 --- a/stage0/src/Lean/Server/Snapshots.lean +++ b/stage0/src/Lean/Server/Snapshots.lean @@ -73,11 +73,6 @@ def isAtEnd (s : Snapshot) : Bool := end Snapshot -/-- Reparses the header syntax but does not re-elaborate it. Used to ignore whitespace-only changes. -/ -def reparseHeader (inputCtx : Parser.InputContext) (header : Snapshot) : IO Snapshot := do - let (newStx, newMpState, _) ← Parser.parseHeader inputCtx - pure { header with stx := newStx, mpState := newMpState } - /-- Parses the next command occurring after the given snapshot without elaborating it. -/ def parseNextCmd (inputCtx : Parser.InputContext) (snap : Snapshot) : IO Syntax := do diff --git a/stage0/src/Lean/Util/Trace.lean b/stage0/src/Lean/Util/Trace.lean index 15ede7a70fa..7a7daec3ea9 100644 --- a/stage0/src/Lean/Util/Trace.lean +++ b/stage0/src/Lean/Util/Trace.lean @@ -147,7 +147,7 @@ def registerTraceClass (traceClassName : Name) : IO Unit := registerOption (`trace ++ traceClassName) { group := "trace", defValue := false, descr := "enable/disable tracing for the given module and submodules" } macro "trace[" id:ident "]" s:(interpolatedStr(term) <|> term) : doElem => do - let msg ← if s.getKind == interpolatedStrKind then `(m! $s) else `(($s : MessageData)) + let msg ← if s.raw.getKind == interpolatedStrKind then `(m! $(⟨s⟩)) else `(($(⟨s⟩) : MessageData)) `(doElem| do let cls := $(quote id.getId.eraseMacroScopes) if (← Lean.isTracingEnabledFor cls) then diff --git a/stage0/src/bin/lean-gdb.py b/stage0/src/bin/lean-gdb.py index b47d682494c..be1d84b8479 100644 --- a/stage0/src/bin/lean-gdb.py +++ b/stage0/src/bin/lean-gdb.py @@ -73,7 +73,7 @@ def to_string(self): if is_scalar(self.val): return unbox(self.val) else: - return "{} ({})".format(LeanObjectPrinter.kinds[self.kind][0], get_rc(self.val)) + return "{} (RC {})".format(LeanObjectPrinter.kinds[self.kind][0], get_rc(self.val)) def children(self): if is_scalar(self.val): @@ -82,11 +82,14 @@ def children(self): return typ, fields = LeanObjectPrinter.kinds[self.kind] val = self.val.cast(gdb.lookup_type("lean_" + typ + "_object").pointer()).dereference() + if self.kind == 0: + yield ('tag', get_tag(self.val)) for f in fields: yield (f, val[f]) if typ == 'ctor': for i in range(get_num_objs(self.val)): - yield ('', val['m_objs'][i].cast(gdb.lookup_type('lean_object').pointer())) + p = val['m_objs'][i].cast(gdb.lookup_type('lean_object').pointer()) + yield ('', p if is_scalar(p) else p.dereference()) elif typ == 'array': for i in range(val['m_size']): yield ('', val['m_data'][i].cast(gdb.lookup_type('lean_object').pointer())) diff --git a/stage0/src/stdlib_flags.h b/stage0/src/stdlib_flags.h index d1703c2f0f3..e3a2e51437d 100644 --- a/stage0/src/stdlib_flags.h +++ b/stage0/src/stdlib_flags.h @@ -8,7 +8,7 @@ options get_default_options() { // switch to `true` for ABI-breaking changes affecting meta code opts = opts.update({"interpreter", "prefer_native"}, false); // switch to `true` for changing built-in parsers used in quotations - opts = opts.update({"internal", "parseQuotWithCurrentStage"}, true); + opts = opts.update({"internal", "parseQuotWithCurrentStage"}, false); opts = opts.update({"pp", "rawOnError"}, true); #endif return opts; diff --git a/stage0/stdlib/Init/Classical.c b/stage0/stdlib/Init/Classical.c index f976b83bf61..d807b1f458f 100644 --- a/stage0/stdlib/Init/Classical.c +++ b/stage0/stdlib/Init/Classical.c @@ -13,101 +13,100 @@ #ifdef __cplusplus extern "C" { #endif +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__45; +static lean_object* l_Classical_tacticBy__cases___x3a_____closed__9; lean_object* lean_mk_empty_array_with_capacity(lean_object*); -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__5; -extern lean_object* l_Lean_nullKind; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__16; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__51; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__9; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__40; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__27; lean_object* lean_name_mk_string(lean_object*, lean_object*); -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__24; -static lean_object* l_Classical_tacticBy__cases_____x3a_____closed__9; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__49; -static lean_object* l_Classical_tacticBy__cases_____x3a_____closed__22; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__7; -static lean_object* l_Classical_tacticBy__cases_____x3a_____closed__25; -static lean_object* l_Classical_tacticBy__cases_____x3a_____closed__23; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__39; +static lean_object* l_Classical_tacticBy__cases___x3a_____closed__26; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__19; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__35; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__31; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__21; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__38; +static lean_object* l_Classical_tacticBy__cases___x3a_____closed__2; +static lean_object* l_Classical_tacticBy__cases___x3a_____closed__18; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__2; +static lean_object* l_Classical_tacticBy__cases___x3a_____closed__5; +static lean_object* l_Classical_tacticBy__cases___x3a_____closed__14; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__4; lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(lean_object*, lean_object*); -static lean_object* l_Classical_tacticBy__cases_____x3a_____closed__13; lean_object* lean_array_push(lean_object*, lean_object*); -static lean_object* l_Classical_tacticBy__cases_____x3a_____closed__14; -static lean_object* l_Classical_tacticBy__cases_____x3a_____closed__5; -static lean_object* l_Classical_tacticBy__cases_____x3a_____closed__18; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__20; -static lean_object* l_Classical_tacticBy__cases_____x3a_____closed__15; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__43; +static lean_object* l_Classical_tacticBy__cases___x3a_____closed__13; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__44; +static lean_object* l_Classical_tacticBy__cases___x3a_____closed__23; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__17; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__8; lean_object* lean_string_utf8_byte_size(lean_object*); -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__32; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__10; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__26; -static lean_object* l_Classical_tacticBy__cases_____x3a_____closed__6; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__12; -static lean_object* l_Classical_tacticBy__cases_____x3a_____closed__20; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__36; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__11; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__19; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__51; -static lean_object* l_Classical_tacticBy__cases_____x3a_____closed__16; -static lean_object* l_Classical_tacticBy__cases_____x3a_____closed__3; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__39; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__44; -static lean_object* l_Classical_tacticBy__cases_____x3a_____closed__10; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__31; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__35; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__1; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__17; -static lean_object* l_Classical_tacticBy__cases_____x3a_____closed__26; -LEAN_EXPORT lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__8; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__45; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__9; -static lean_object* l_Classical_tacticBy__cases_____x3a_____closed__7; -static lean_object* l_Classical_tacticBy__cases_____x3a_____closed__24; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__34; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__36; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__1; +static lean_object* l_Classical_tacticBy__cases___x3a_____closed__6; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__6; +static lean_object* l_Classical_tacticBy__cases___x3a_____closed__12; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__23; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__46; +static lean_object* l_Classical_tacticBy__cases___x3a_____closed__16; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__34; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__48; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__13; +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__41; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__47; +static lean_object* l_Classical_tacticBy__cases___x3a_____closed__20; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__18; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__14; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__28; +static lean_object* l_Classical_tacticBy__cases___x3a_____closed__21; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__29; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__25; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__5; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__30; +static lean_object* l_Classical_tacticBy__cases___x3a_____closed__19; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__52; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__50; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__54; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__49; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__24; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__7; +static lean_object* l_Classical_tacticBy__cases___x3a_____closed__25; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__15; +static lean_object* l_Classical_tacticBy__cases___x3a_____closed__24; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__3; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__42; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__53; lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__4; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__27; -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__48; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__13; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__23; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__40; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__41; -static lean_object* l_Classical_tacticBy__cases_____x3a_____closed__11; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__2; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__38; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__21; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__47; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__33; -static lean_object* l_Classical_tacticBy__cases_____x3a_____closed__8; -static lean_object* l_Classical_tacticBy__cases_____x3a_____closed__17; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__52; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__29; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__33; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__55; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__22; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__37; +static lean_object* l_Classical_tacticBy__cases___x3a_____closed__10; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__43; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__20; +static lean_object* l_Classical_tacticBy__cases___x3a_____closed__11; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__16; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__56; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__12; +static lean_object* l_Classical_tacticBy__cases___x3a_____closed__4; +LEAN_EXPORT lean_object* l_Classical_tacticBy__cases___x3a__; +static lean_object* l_Classical_tacticBy__cases___x3a_____closed__8; +LEAN_EXPORT lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__53; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__42; -static lean_object* l_Classical_tacticBy__cases_____x3a_____closed__12; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__6; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__25; +static lean_object* l_Classical_tacticBy__cases___x3a_____closed__17; +static lean_object* l_Classical_tacticBy__cases___x3a_____closed__1; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__32; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__10; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); -static lean_object* l_Classical_tacticBy__cases_____x3a_____closed__2; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__54; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__50; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__55; -static lean_object* l_Classical_tacticBy__cases_____x3a_____closed__1; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__3; -static lean_object* l_Classical_tacticBy__cases_____x3a_____closed__4; -LEAN_EXPORT lean_object* l_Classical_tacticBy__cases_____x3a__; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__46; -static lean_object* l_Classical_tacticBy__cases_____x3a_____closed__19; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__14; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__15; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__56; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__18; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__30; -static lean_object* l_Classical_tacticBy__cases_____x3a_____closed__21; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__28; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__22; -static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__37; -static lean_object* _init_l_Classical_tacticBy__cases_____x3a_____closed__1() { +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__11; +static lean_object* l_Classical_tacticBy__cases___x3a_____closed__22; +static lean_object* l_Classical_tacticBy__cases___x3a_____closed__15; +static lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__26; +static lean_object* l_Classical_tacticBy__cases___x3a_____closed__3; +static lean_object* l_Classical_tacticBy__cases___x3a_____closed__7; +static lean_object* _init_l_Classical_tacticBy__cases___x3a_____closed__1() { _start: { lean_object* x_1; @@ -115,35 +114,35 @@ x_1 = lean_mk_string_from_bytes("Classical", 9); return x_1; } } -static lean_object* _init_l_Classical_tacticBy__cases_____x3a_____closed__2() { +static lean_object* _init_l_Classical_tacticBy__cases___x3a_____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Classical_tacticBy__cases_____x3a_____closed__1; +x_2 = l_Classical_tacticBy__cases___x3a_____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Classical_tacticBy__cases_____x3a_____closed__3() { +static lean_object* _init_l_Classical_tacticBy__cases___x3a_____closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("tacticBy_cases__:_", 18); +x_1 = lean_mk_string_from_bytes("tacticBy_cases_:_", 17); return x_1; } } -static lean_object* _init_l_Classical_tacticBy__cases_____x3a_____closed__4() { +static lean_object* _init_l_Classical_tacticBy__cases___x3a_____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Classical_tacticBy__cases_____x3a_____closed__2; -x_2 = l_Classical_tacticBy__cases_____x3a_____closed__3; +x_1 = l_Classical_tacticBy__cases___x3a_____closed__2; +x_2 = l_Classical_tacticBy__cases___x3a_____closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Classical_tacticBy__cases_____x3a_____closed__5() { +static lean_object* _init_l_Classical_tacticBy__cases___x3a_____closed__5() { _start: { lean_object* x_1; @@ -151,17 +150,17 @@ x_1 = lean_mk_string_from_bytes("andthen", 7); return x_1; } } -static lean_object* _init_l_Classical_tacticBy__cases_____x3a_____closed__6() { +static lean_object* _init_l_Classical_tacticBy__cases___x3a_____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Classical_tacticBy__cases_____x3a_____closed__5; +x_2 = l_Classical_tacticBy__cases___x3a_____closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Classical_tacticBy__cases_____x3a_____closed__7() { +static lean_object* _init_l_Classical_tacticBy__cases___x3a_____closed__7() { _start: { lean_object* x_1; @@ -169,11 +168,11 @@ x_1 = lean_mk_string_from_bytes("by_cases", 8); return x_1; } } -static lean_object* _init_l_Classical_tacticBy__cases_____x3a_____closed__8() { +static lean_object* _init_l_Classical_tacticBy__cases___x3a_____closed__8() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Classical_tacticBy__cases_____x3a_____closed__7; +x_1 = l_Classical_tacticBy__cases___x3a_____closed__7; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -181,7 +180,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l_Classical_tacticBy__cases_____x3a_____closed__9() { +static lean_object* _init_l_Classical_tacticBy__cases___x3a_____closed__9() { _start: { lean_object* x_1; @@ -189,17 +188,17 @@ x_1 = lean_mk_string_from_bytes("optional", 8); return x_1; } } -static lean_object* _init_l_Classical_tacticBy__cases_____x3a_____closed__10() { +static lean_object* _init_l_Classical_tacticBy__cases___x3a_____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Classical_tacticBy__cases_____x3a_____closed__9; +x_2 = l_Classical_tacticBy__cases___x3a_____closed__9; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Classical_tacticBy__cases_____x3a_____closed__11() { +static lean_object* _init_l_Classical_tacticBy__cases___x3a_____closed__11() { _start: { lean_object* x_1; @@ -207,17 +206,17 @@ x_1 = lean_mk_string_from_bytes("atomic", 6); return x_1; } } -static lean_object* _init_l_Classical_tacticBy__cases_____x3a_____closed__12() { +static lean_object* _init_l_Classical_tacticBy__cases___x3a_____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Classical_tacticBy__cases_____x3a_____closed__11; +x_2 = l_Classical_tacticBy__cases___x3a_____closed__11; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Classical_tacticBy__cases_____x3a_____closed__13() { +static lean_object* _init_l_Classical_tacticBy__cases___x3a_____closed__13() { _start: { lean_object* x_1; @@ -225,27 +224,27 @@ x_1 = lean_mk_string_from_bytes("ident", 5); return x_1; } } -static lean_object* _init_l_Classical_tacticBy__cases_____x3a_____closed__14() { +static lean_object* _init_l_Classical_tacticBy__cases___x3a_____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Classical_tacticBy__cases_____x3a_____closed__13; +x_2 = l_Classical_tacticBy__cases___x3a_____closed__13; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Classical_tacticBy__cases_____x3a_____closed__15() { +static lean_object* _init_l_Classical_tacticBy__cases___x3a_____closed__15() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Classical_tacticBy__cases_____x3a_____closed__14; +x_1 = l_Classical_tacticBy__cases___x3a_____closed__14; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Classical_tacticBy__cases_____x3a_____closed__16() { +static lean_object* _init_l_Classical_tacticBy__cases___x3a_____closed__16() { _start: { lean_object* x_1; @@ -253,23 +252,23 @@ x_1 = lean_mk_string_from_bytes(":", 1); return x_1; } } -static lean_object* _init_l_Classical_tacticBy__cases_____x3a_____closed__17() { +static lean_object* _init_l_Classical_tacticBy__cases___x3a_____closed__17() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Classical_tacticBy__cases_____x3a_____closed__16; +x_1 = l_Classical_tacticBy__cases___x3a_____closed__16; x_2 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Classical_tacticBy__cases_____x3a_____closed__18() { +static lean_object* _init_l_Classical_tacticBy__cases___x3a_____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Classical_tacticBy__cases_____x3a_____closed__6; -x_2 = l_Classical_tacticBy__cases_____x3a_____closed__15; -x_3 = l_Classical_tacticBy__cases_____x3a_____closed__17; +x_1 = l_Classical_tacticBy__cases___x3a_____closed__6; +x_2 = l_Classical_tacticBy__cases___x3a_____closed__15; +x_3 = l_Classical_tacticBy__cases___x3a_____closed__17; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -277,37 +276,37 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Classical_tacticBy__cases_____x3a_____closed__19() { +static lean_object* _init_l_Classical_tacticBy__cases___x3a_____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Classical_tacticBy__cases_____x3a_____closed__12; -x_2 = l_Classical_tacticBy__cases_____x3a_____closed__18; +x_1 = l_Classical_tacticBy__cases___x3a_____closed__12; +x_2 = l_Classical_tacticBy__cases___x3a_____closed__18; x_3 = lean_alloc_ctor(1, 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_Classical_tacticBy__cases_____x3a_____closed__20() { +static lean_object* _init_l_Classical_tacticBy__cases___x3a_____closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Classical_tacticBy__cases_____x3a_____closed__10; -x_2 = l_Classical_tacticBy__cases_____x3a_____closed__19; +x_1 = l_Classical_tacticBy__cases___x3a_____closed__10; +x_2 = l_Classical_tacticBy__cases___x3a_____closed__19; x_3 = lean_alloc_ctor(1, 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_Classical_tacticBy__cases_____x3a_____closed__21() { +static lean_object* _init_l_Classical_tacticBy__cases___x3a_____closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Classical_tacticBy__cases_____x3a_____closed__6; -x_2 = l_Classical_tacticBy__cases_____x3a_____closed__8; -x_3 = l_Classical_tacticBy__cases_____x3a_____closed__20; +x_1 = l_Classical_tacticBy__cases___x3a_____closed__6; +x_2 = l_Classical_tacticBy__cases___x3a_____closed__8; +x_3 = l_Classical_tacticBy__cases___x3a_____closed__20; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -315,7 +314,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Classical_tacticBy__cases_____x3a_____closed__22() { +static lean_object* _init_l_Classical_tacticBy__cases___x3a_____closed__22() { _start: { lean_object* x_1; @@ -323,21 +322,21 @@ x_1 = lean_mk_string_from_bytes("term", 4); return x_1; } } -static lean_object* _init_l_Classical_tacticBy__cases_____x3a_____closed__23() { +static lean_object* _init_l_Classical_tacticBy__cases___x3a_____closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Classical_tacticBy__cases_____x3a_____closed__22; +x_2 = l_Classical_tacticBy__cases___x3a_____closed__22; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Classical_tacticBy__cases_____x3a_____closed__24() { +static lean_object* _init_l_Classical_tacticBy__cases___x3a_____closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Classical_tacticBy__cases_____x3a_____closed__23; +x_1 = l_Classical_tacticBy__cases___x3a_____closed__23; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -345,13 +344,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Classical_tacticBy__cases_____x3a_____closed__25() { +static lean_object* _init_l_Classical_tacticBy__cases___x3a_____closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Classical_tacticBy__cases_____x3a_____closed__6; -x_2 = l_Classical_tacticBy__cases_____x3a_____closed__21; -x_3 = l_Classical_tacticBy__cases_____x3a_____closed__24; +x_1 = l_Classical_tacticBy__cases___x3a_____closed__6; +x_2 = l_Classical_tacticBy__cases___x3a_____closed__21; +x_3 = l_Classical_tacticBy__cases___x3a_____closed__24; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -359,13 +358,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Classical_tacticBy__cases_____x3a_____closed__26() { +static lean_object* _init_l_Classical_tacticBy__cases___x3a_____closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Classical_tacticBy__cases_____x3a_____closed__4; +x_1 = l_Classical_tacticBy__cases___x3a_____closed__4; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Classical_tacticBy__cases_____x3a_____closed__25; +x_3 = l_Classical_tacticBy__cases___x3a_____closed__25; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -373,15 +372,15 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Classical_tacticBy__cases_____x3a__() { +static lean_object* _init_l_Classical_tacticBy__cases___x3a__() { _start: { lean_object* x_1; -x_1 = l_Classical_tacticBy__cases_____x3a_____closed__26; +x_1 = l_Classical_tacticBy__cases___x3a_____closed__26; return x_1; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__1() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__1() { _start: { lean_object* x_1; @@ -389,17 +388,17 @@ x_1 = lean_mk_string_from_bytes("Lean", 4); return x_1; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__2() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__1; +x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__3() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__3() { _start: { lean_object* x_1; @@ -407,17 +406,17 @@ x_1 = lean_mk_string_from_bytes("Parser", 6); return x_1; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__4() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__2; -x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__3; +x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__2; +x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__5() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__5() { _start: { lean_object* x_1; @@ -425,17 +424,17 @@ x_1 = lean_mk_string_from_bytes("Tactic", 6); return x_1; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__6() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__4; -x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__5; +x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__4; +x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__7() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__7() { _start: { lean_object* x_1; @@ -443,17 +442,17 @@ x_1 = lean_mk_string_from_bytes("cases", 5); return x_1; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__8() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__6; -x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__7; +x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__6; +x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__7; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__9() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__9() { _start: { lean_object* x_1; @@ -461,17 +460,17 @@ x_1 = lean_mk_string_from_bytes("null", 4); return x_1; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__10() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__9; +x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__9; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__11() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__11() { _start: { lean_object* x_1; @@ -479,17 +478,17 @@ x_1 = lean_mk_string_from_bytes("casesTarget", 11); return x_1; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__12() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__6; -x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__11; +x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__6; +x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__11; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__13() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__13() { _start: { lean_object* x_1; lean_object* x_2; @@ -498,13 +497,13 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__14() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__10; -x_3 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__13; +x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__10; +x_3 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__13; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -512,7 +511,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__15() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__15() { _start: { lean_object* x_1; @@ -520,17 +519,17 @@ x_1 = lean_mk_string_from_bytes("Term", 4); return x_1; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__16() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__4; -x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__15; +x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__4; +x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__15; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__17() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__17() { _start: { lean_object* x_1; @@ -538,17 +537,17 @@ x_1 = lean_mk_string_from_bytes("app", 3); return x_1; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__18() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__16; -x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__17; +x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__16; +x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__17; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__19() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__19() { _start: { lean_object* x_1; @@ -556,22 +555,22 @@ x_1 = lean_mk_string_from_bytes("em", 2); return x_1; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__20() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__20() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__19; +x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__19; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__21() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__19; +x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__19; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__20; +x_3 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__20; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -579,51 +578,51 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__22() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__19; +x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__19; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__23() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Classical_tacticBy__cases_____x3a_____closed__2; -x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__19; +x_1 = l_Classical_tacticBy__cases___x3a_____closed__2; +x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__19; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__24() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__23; +x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__23; 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_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__25() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__24; +x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__24; x_3 = lean_alloc_ctor(1, 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_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__26() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__26() { _start: { lean_object* x_1; lean_object* x_2; @@ -632,7 +631,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__27() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__27() { _start: { lean_object* x_1; lean_object* x_2; @@ -641,17 +640,17 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__28() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__27; -x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__14; +x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__27; +x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__14; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__29() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__29() { _start: { lean_object* x_1; @@ -659,17 +658,17 @@ x_1 = lean_mk_string_from_bytes("inductionAlts", 13); return x_1; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__30() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__30() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__6; -x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__29; +x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__6; +x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__29; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__31() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__31() { _start: { lean_object* x_1; @@ -677,7 +676,7 @@ x_1 = lean_mk_string_from_bytes("with", 4); return x_1; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__32() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__32() { _start: { lean_object* x_1; @@ -685,17 +684,17 @@ x_1 = lean_mk_string_from_bytes("inductionAlt", 12); return x_1; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__33() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__33() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__6; -x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__32; +x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__6; +x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__32; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__34() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__34() { _start: { lean_object* x_1; @@ -703,17 +702,17 @@ x_1 = lean_mk_string_from_bytes("inductionAltLHS", 15); return x_1; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__35() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__35() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__6; -x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__34; +x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__6; +x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__34; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__36() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__36() { _start: { lean_object* x_1; @@ -721,7 +720,7 @@ x_1 = lean_mk_string_from_bytes("|", 1); return x_1; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__37() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__37() { _start: { lean_object* x_1; @@ -729,17 +728,17 @@ x_1 = lean_mk_string_from_bytes("group", 5); return x_1; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__38() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__38() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__37; +x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__37; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__39() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__39() { _start: { lean_object* x_1; @@ -747,22 +746,22 @@ x_1 = lean_mk_string_from_bytes("inl", 3); return x_1; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__40() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__40() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__39; +x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__39; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__41() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__41() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__39; +x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__39; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__40; +x_3 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__40; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -770,17 +769,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__42() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__42() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__39; +x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__39; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__43() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__43() { _start: { lean_object* x_1; @@ -788,22 +787,22 @@ x_1 = lean_mk_string_from_bytes("h", 1); return x_1; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__44() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__44() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__43; +x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__43; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__45() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__45() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__43; +x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__43; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__44; +x_3 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__44; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -811,17 +810,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__46() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__46() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__43; +x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__43; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__47() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__47() { _start: { lean_object* x_1; lean_object* x_2; @@ -830,7 +829,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__48() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__48() { _start: { lean_object* x_1; @@ -838,7 +837,7 @@ x_1 = lean_mk_string_from_bytes("=>", 2); return x_1; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__49() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__49() { _start: { lean_object* x_1; @@ -846,17 +845,17 @@ x_1 = lean_mk_string_from_bytes("hole", 4); return x_1; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__50() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__50() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__16; -x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__49; +x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__16; +x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__49; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__51() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__51() { _start: { lean_object* x_1; @@ -864,7 +863,7 @@ x_1 = lean_mk_string_from_bytes("_", 1); return x_1; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__52() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__52() { _start: { lean_object* x_1; @@ -872,22 +871,22 @@ x_1 = lean_mk_string_from_bytes("inr", 3); return x_1; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__53() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__53() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__52; +x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__52; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__54() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__54() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__52; +x_1 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__52; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__53; +x_3 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__53; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -895,17 +894,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__55() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__55() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__52; +x_2 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__52; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__56() { +static lean_object* _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__56() { _start: { lean_object* x_1; lean_object* x_2; @@ -914,11 +913,11 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; uint8_t x_5; -x_4 = l_Classical_tacticBy__cases_____x3a_____closed__4; +x_4 = l_Classical_tacticBy__cases___x3a_____closed__4; lean_inc(x_1); x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); if (x_5 == 0) @@ -934,873 +933,872 @@ return x_7; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; x_8 = lean_unsigned_to_nat(1u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); -x_10 = l_Lean_nullKind; -x_11 = lean_unsigned_to_nat(2u); +x_10 = lean_unsigned_to_nat(2u); lean_inc(x_9); -x_12 = l_Lean_Syntax_isNodeOf(x_9, x_10, x_11); -if (x_12 == 0) +x_11 = l_Lean_Syntax_matchesNull(x_9, x_10); +if (x_11 == 0) { -lean_object* x_13; uint8_t x_14; -x_13 = lean_unsigned_to_nat(0u); -x_14 = l_Lean_Syntax_isNodeOf(x_9, x_10, x_13); -if (x_14 == 0) +lean_object* x_12; uint8_t x_13; +x_12 = lean_unsigned_to_nat(0u); +x_13 = l_Lean_Syntax_matchesNull(x_9, x_12); +if (x_13 == 0) { -lean_object* x_15; lean_object* x_16; +lean_object* x_14; lean_object* x_15; lean_dec(x_2); lean_dec(x_1); -x_15 = lean_box(1); -x_16 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_3); -return x_16; +x_14 = lean_box(1); +x_15 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_3); +return x_15; } else { -lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_17 = l_Lean_Syntax_getArg(x_1, x_11); +lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_16 = l_Lean_Syntax_getArg(x_1, x_10); lean_dec(x_1); lean_inc(x_2); -x_18 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_19 = !lean_is_exclusive(x_18); -if (x_19 == 0) +x_17 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) { -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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; -x_20 = lean_ctor_get(x_18, 0); -x_21 = lean_ctor_get(x_2, 2); +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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; +x_19 = lean_ctor_get(x_17, 0); +x_20 = lean_ctor_get(x_2, 2); +lean_inc(x_20); +x_21 = lean_ctor_get(x_2, 1); lean_inc(x_21); -x_22 = lean_ctor_get(x_2, 1); -lean_inc(x_22); lean_dec(x_2); -x_23 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__7; +x_22 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__7; +lean_inc(x_19); +x_23 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_23, 0, x_19); +lean_ctor_set(x_23, 1, x_22); +x_24 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__22; lean_inc(x_20); -x_24 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_24, 0, x_20); -lean_ctor_set(x_24, 1, x_23); -x_25 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__22; lean_inc(x_21); -lean_inc(x_22); -x_26 = l_Lean_addMacroScope(x_22, x_25, x_21); -x_27 = lean_box(0); -x_28 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__21; -x_29 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__25; -lean_inc(x_20); -x_30 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_30, 0, x_20); -lean_ctor_set(x_30, 1, x_28); -lean_ctor_set(x_30, 2, x_26); -lean_ctor_set(x_30, 3, x_29); -x_31 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__26; -x_32 = lean_array_push(x_31, x_17); -x_33 = lean_box(2); -x_34 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__10; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -x_36 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__27; -x_37 = lean_array_push(x_36, x_30); -x_38 = lean_array_push(x_37, x_35); -x_39 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__18; -x_40 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_40, 0, x_33); -lean_ctor_set(x_40, 1, x_39); -lean_ctor_set(x_40, 2, x_38); -x_41 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__28; -x_42 = lean_array_push(x_41, x_40); -x_43 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__12; -x_44 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_44, 0, x_33); -lean_ctor_set(x_44, 1, x_43); -lean_ctor_set(x_44, 2, x_42); -x_45 = lean_array_push(x_31, x_44); -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_33); -lean_ctor_set(x_46, 1, x_34); -lean_ctor_set(x_46, 2, x_45); -x_47 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__31; -lean_inc(x_20); -x_48 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_48, 0, x_20); -lean_ctor_set(x_48, 1, x_47); -x_49 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__36; +x_25 = l_Lean_addMacroScope(x_21, x_24, x_20); +x_26 = lean_box(0); +x_27 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__21; +x_28 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__25; +lean_inc(x_19); +x_29 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_29, 0, x_19); +lean_ctor_set(x_29, 1, x_27); +lean_ctor_set(x_29, 2, x_25); +lean_ctor_set(x_29, 3, x_28); +x_30 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__26; +x_31 = lean_array_push(x_30, x_16); +x_32 = lean_box(2); +x_33 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__10; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +x_35 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__27; +x_36 = lean_array_push(x_35, x_29); +x_37 = lean_array_push(x_36, x_34); +x_38 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__18; +x_39 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_39, 0, x_32); +lean_ctor_set(x_39, 1, x_38); +lean_ctor_set(x_39, 2, x_37); +x_40 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__28; +x_41 = lean_array_push(x_40, x_39); +x_42 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__12; +x_43 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_43, 0, x_32); +lean_ctor_set(x_43, 1, x_42); +lean_ctor_set(x_43, 2, x_41); +x_44 = lean_array_push(x_30, x_43); +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_32); +lean_ctor_set(x_45, 1, x_33); +lean_ctor_set(x_45, 2, x_44); +x_46 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__31; +lean_inc(x_19); +x_47 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_47, 0, x_19); +lean_ctor_set(x_47, 1, x_46); +x_48 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__36; +lean_inc(x_19); +x_49 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_49, 0, x_19); +lean_ctor_set(x_49, 1, x_48); +x_50 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__42; lean_inc(x_20); -x_50 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_50, 0, x_20); -lean_ctor_set(x_50, 1, x_49); -x_51 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__42; lean_inc(x_21); -lean_inc(x_22); -x_52 = l_Lean_addMacroScope(x_22, x_51, x_21); -x_53 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__41; +x_51 = l_Lean_addMacroScope(x_21, x_50, x_20); +x_52 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__41; +lean_inc(x_19); +x_53 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_53, 0, x_19); +lean_ctor_set(x_53, 1, x_52); +lean_ctor_set(x_53, 2, x_51); +lean_ctor_set(x_53, 3, x_26); +x_54 = lean_array_push(x_40, x_53); +x_55 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__38; +x_56 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_56, 0, x_32); +lean_ctor_set(x_56, 1, x_55); +lean_ctor_set(x_56, 2, x_54); +x_57 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__46; lean_inc(x_20); -x_54 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_54, 0, x_20); -lean_ctor_set(x_54, 1, x_53); -lean_ctor_set(x_54, 2, x_52); -lean_ctor_set(x_54, 3, x_27); -x_55 = lean_array_push(x_41, x_54); -x_56 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__38; -x_57 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_57, 0, x_33); -lean_ctor_set(x_57, 1, x_56); -lean_ctor_set(x_57, 2, x_55); -x_58 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__46; lean_inc(x_21); -lean_inc(x_22); -x_59 = l_Lean_addMacroScope(x_22, x_58, x_21); -x_60 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__45; -lean_inc(x_20); -x_61 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_61, 0, x_20); -lean_ctor_set(x_61, 1, x_60); -lean_ctor_set(x_61, 2, x_59); -lean_ctor_set(x_61, 3, x_27); -x_62 = lean_array_push(x_31, x_61); -x_63 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_63, 0, x_33); -lean_ctor_set(x_63, 1, x_34); -lean_ctor_set(x_63, 2, x_62); -x_64 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__47; -x_65 = lean_array_push(x_64, x_50); -lean_inc(x_65); -x_66 = lean_array_push(x_65, x_57); -lean_inc(x_63); -x_67 = lean_array_push(x_66, x_63); -x_68 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__35; -x_69 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_69, 0, x_33); -lean_ctor_set(x_69, 1, x_68); -lean_ctor_set(x_69, 2, x_67); -x_70 = lean_array_push(x_31, x_69); -x_71 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_71, 0, x_33); -lean_ctor_set(x_71, 1, x_34); -lean_ctor_set(x_71, 2, x_70); -x_72 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__48; -lean_inc(x_20); -x_73 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_73, 0, x_20); -lean_ctor_set(x_73, 1, x_72); -x_74 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__51; -lean_inc(x_20); -x_75 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_75, 0, x_20); -lean_ctor_set(x_75, 1, x_74); -x_76 = lean_array_push(x_31, x_75); -x_77 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__50; -x_78 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_78, 0, x_33); -lean_ctor_set(x_78, 1, x_77); -lean_ctor_set(x_78, 2, x_76); -x_79 = lean_array_push(x_64, x_71); -lean_inc(x_73); -x_80 = lean_array_push(x_79, x_73); -lean_inc(x_78); -x_81 = lean_array_push(x_80, x_78); -x_82 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__33; -x_83 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_83, 0, x_33); -lean_ctor_set(x_83, 1, x_82); -lean_ctor_set(x_83, 2, x_81); -x_84 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__55; -x_85 = l_Lean_addMacroScope(x_22, x_84, x_21); -x_86 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__54; -x_87 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_87, 0, x_20); -lean_ctor_set(x_87, 1, x_86); -lean_ctor_set(x_87, 2, x_85); -lean_ctor_set(x_87, 3, x_27); -x_88 = lean_array_push(x_41, x_87); -x_89 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_89, 0, x_33); -lean_ctor_set(x_89, 1, x_56); -lean_ctor_set(x_89, 2, x_88); -x_90 = lean_array_push(x_65, x_89); -x_91 = lean_array_push(x_90, x_63); -x_92 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_92, 0, x_33); -lean_ctor_set(x_92, 1, x_68); -lean_ctor_set(x_92, 2, x_91); -x_93 = lean_array_push(x_31, x_92); -x_94 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_94, 0, x_33); -lean_ctor_set(x_94, 1, x_34); -lean_ctor_set(x_94, 2, x_93); -x_95 = lean_array_push(x_64, x_94); -x_96 = lean_array_push(x_95, x_73); -x_97 = lean_array_push(x_96, x_78); -x_98 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_98, 0, x_33); -lean_ctor_set(x_98, 1, x_82); -lean_ctor_set(x_98, 2, x_97); -x_99 = lean_array_push(x_36, x_83); -x_100 = lean_array_push(x_99, x_98); -x_101 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_101, 0, x_33); -lean_ctor_set(x_101, 1, x_34); -lean_ctor_set(x_101, 2, x_100); -x_102 = lean_array_push(x_64, x_48); -x_103 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__14; -x_104 = lean_array_push(x_102, x_103); -x_105 = lean_array_push(x_104, x_101); -x_106 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__30; -x_107 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_107, 0, x_33); -lean_ctor_set(x_107, 1, x_106); -lean_ctor_set(x_107, 2, x_105); -x_108 = lean_array_push(x_31, x_107); -x_109 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_109, 0, x_33); -lean_ctor_set(x_109, 1, x_34); -lean_ctor_set(x_109, 2, x_108); -x_110 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__56; -x_111 = lean_array_push(x_110, x_24); -x_112 = lean_array_push(x_111, x_46); -x_113 = lean_array_push(x_112, x_103); -x_114 = lean_array_push(x_113, x_109); -x_115 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__8; -x_116 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_116, 0, x_33); -lean_ctor_set(x_116, 1, x_115); -lean_ctor_set(x_116, 2, x_114); -lean_ctor_set(x_18, 0, x_116); -return x_18; +x_58 = l_Lean_addMacroScope(x_21, x_57, x_20); +x_59 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__45; +lean_inc(x_19); +x_60 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_60, 0, x_19); +lean_ctor_set(x_60, 1, x_59); +lean_ctor_set(x_60, 2, x_58); +lean_ctor_set(x_60, 3, x_26); +x_61 = lean_array_push(x_30, x_60); +x_62 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_62, 0, x_32); +lean_ctor_set(x_62, 1, x_33); +lean_ctor_set(x_62, 2, x_61); +x_63 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__47; +x_64 = lean_array_push(x_63, x_49); +lean_inc(x_64); +x_65 = lean_array_push(x_64, x_56); +lean_inc(x_62); +x_66 = lean_array_push(x_65, x_62); +x_67 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__35; +x_68 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_68, 0, x_32); +lean_ctor_set(x_68, 1, x_67); +lean_ctor_set(x_68, 2, x_66); +x_69 = lean_array_push(x_30, x_68); +x_70 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_70, 0, x_32); +lean_ctor_set(x_70, 1, x_33); +lean_ctor_set(x_70, 2, x_69); +x_71 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__48; +lean_inc(x_19); +x_72 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_72, 0, x_19); +lean_ctor_set(x_72, 1, x_71); +x_73 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__51; +lean_inc(x_19); +x_74 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_74, 0, x_19); +lean_ctor_set(x_74, 1, x_73); +x_75 = lean_array_push(x_30, x_74); +x_76 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__50; +x_77 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_77, 0, x_32); +lean_ctor_set(x_77, 1, x_76); +lean_ctor_set(x_77, 2, x_75); +x_78 = lean_array_push(x_63, x_70); +lean_inc(x_72); +x_79 = lean_array_push(x_78, x_72); +lean_inc(x_77); +x_80 = lean_array_push(x_79, x_77); +x_81 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__33; +x_82 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_82, 0, x_32); +lean_ctor_set(x_82, 1, x_81); +lean_ctor_set(x_82, 2, x_80); +x_83 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__55; +x_84 = l_Lean_addMacroScope(x_21, x_83, x_20); +x_85 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__54; +x_86 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_86, 0, x_19); +lean_ctor_set(x_86, 1, x_85); +lean_ctor_set(x_86, 2, x_84); +lean_ctor_set(x_86, 3, x_26); +x_87 = lean_array_push(x_40, x_86); +x_88 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_88, 0, x_32); +lean_ctor_set(x_88, 1, x_55); +lean_ctor_set(x_88, 2, x_87); +x_89 = lean_array_push(x_64, x_88); +x_90 = lean_array_push(x_89, x_62); +x_91 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_91, 0, x_32); +lean_ctor_set(x_91, 1, x_67); +lean_ctor_set(x_91, 2, x_90); +x_92 = lean_array_push(x_30, x_91); +x_93 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_93, 0, x_32); +lean_ctor_set(x_93, 1, x_33); +lean_ctor_set(x_93, 2, x_92); +x_94 = lean_array_push(x_63, x_93); +x_95 = lean_array_push(x_94, x_72); +x_96 = lean_array_push(x_95, x_77); +x_97 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_97, 0, x_32); +lean_ctor_set(x_97, 1, x_81); +lean_ctor_set(x_97, 2, x_96); +x_98 = lean_array_push(x_35, x_82); +x_99 = lean_array_push(x_98, x_97); +x_100 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_100, 0, x_32); +lean_ctor_set(x_100, 1, x_33); +lean_ctor_set(x_100, 2, x_99); +x_101 = lean_array_push(x_63, x_47); +x_102 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__14; +x_103 = lean_array_push(x_101, x_102); +x_104 = lean_array_push(x_103, x_100); +x_105 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__30; +x_106 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_106, 0, x_32); +lean_ctor_set(x_106, 1, x_105); +lean_ctor_set(x_106, 2, x_104); +x_107 = lean_array_push(x_30, x_106); +x_108 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_108, 0, x_32); +lean_ctor_set(x_108, 1, x_33); +lean_ctor_set(x_108, 2, x_107); +x_109 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__56; +x_110 = lean_array_push(x_109, x_23); +x_111 = lean_array_push(x_110, x_45); +x_112 = lean_array_push(x_111, x_102); +x_113 = lean_array_push(x_112, x_108); +x_114 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__8; +x_115 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_115, 0, x_32); +lean_ctor_set(x_115, 1, x_114); +lean_ctor_set(x_115, 2, x_113); +lean_ctor_set(x_17, 0, x_115); +return x_17; } else { -lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; -x_117 = lean_ctor_get(x_18, 0); -x_118 = lean_ctor_get(x_18, 1); -lean_inc(x_118); +lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; +x_116 = lean_ctor_get(x_17, 0); +x_117 = lean_ctor_get(x_17, 1); lean_inc(x_117); -lean_dec(x_18); -x_119 = lean_ctor_get(x_2, 2); +lean_inc(x_116); +lean_dec(x_17); +x_118 = lean_ctor_get(x_2, 2); +lean_inc(x_118); +x_119 = lean_ctor_get(x_2, 1); lean_inc(x_119); -x_120 = lean_ctor_get(x_2, 1); -lean_inc(x_120); lean_dec(x_2); -x_121 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__7; -lean_inc(x_117); -x_122 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_122, 0, x_117); -lean_ctor_set(x_122, 1, x_121); -x_123 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__22; +x_120 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__7; +lean_inc(x_116); +x_121 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_121, 0, x_116); +lean_ctor_set(x_121, 1, x_120); +x_122 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__22; +lean_inc(x_118); lean_inc(x_119); -lean_inc(x_120); -x_124 = l_Lean_addMacroScope(x_120, x_123, x_119); -x_125 = lean_box(0); -x_126 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__21; -x_127 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__25; -lean_inc(x_117); -x_128 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_128, 0, x_117); -lean_ctor_set(x_128, 1, x_126); -lean_ctor_set(x_128, 2, x_124); -lean_ctor_set(x_128, 3, x_127); -x_129 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__26; -x_130 = lean_array_push(x_129, x_17); -x_131 = lean_box(2); -x_132 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__10; -x_133 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_133, 0, x_131); -lean_ctor_set(x_133, 1, x_132); -lean_ctor_set(x_133, 2, x_130); -x_134 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__27; -x_135 = lean_array_push(x_134, x_128); -x_136 = lean_array_push(x_135, x_133); -x_137 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__18; -x_138 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_138, 0, x_131); -lean_ctor_set(x_138, 1, x_137); -lean_ctor_set(x_138, 2, x_136); -x_139 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__28; -x_140 = lean_array_push(x_139, x_138); -x_141 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__12; -x_142 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_142, 0, x_131); -lean_ctor_set(x_142, 1, x_141); -lean_ctor_set(x_142, 2, x_140); -x_143 = lean_array_push(x_129, x_142); -x_144 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_144, 0, x_131); -lean_ctor_set(x_144, 1, x_132); -lean_ctor_set(x_144, 2, x_143); -x_145 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__31; -lean_inc(x_117); -x_146 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_146, 0, x_117); -lean_ctor_set(x_146, 1, x_145); -x_147 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__36; -lean_inc(x_117); -x_148 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_148, 0, x_117); -lean_ctor_set(x_148, 1, x_147); -x_149 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__42; +x_123 = l_Lean_addMacroScope(x_119, x_122, x_118); +x_124 = lean_box(0); +x_125 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__21; +x_126 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__25; +lean_inc(x_116); +x_127 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_127, 0, x_116); +lean_ctor_set(x_127, 1, x_125); +lean_ctor_set(x_127, 2, x_123); +lean_ctor_set(x_127, 3, x_126); +x_128 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__26; +x_129 = lean_array_push(x_128, x_16); +x_130 = lean_box(2); +x_131 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__10; +x_132 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_132, 0, x_130); +lean_ctor_set(x_132, 1, x_131); +lean_ctor_set(x_132, 2, x_129); +x_133 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__27; +x_134 = lean_array_push(x_133, x_127); +x_135 = lean_array_push(x_134, x_132); +x_136 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__18; +x_137 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_137, 0, x_130); +lean_ctor_set(x_137, 1, x_136); +lean_ctor_set(x_137, 2, x_135); +x_138 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__28; +x_139 = lean_array_push(x_138, x_137); +x_140 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__12; +x_141 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_141, 0, x_130); +lean_ctor_set(x_141, 1, x_140); +lean_ctor_set(x_141, 2, x_139); +x_142 = lean_array_push(x_128, x_141); +x_143 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_143, 0, x_130); +lean_ctor_set(x_143, 1, x_131); +lean_ctor_set(x_143, 2, x_142); +x_144 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__31; +lean_inc(x_116); +x_145 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_145, 0, x_116); +lean_ctor_set(x_145, 1, x_144); +x_146 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__36; +lean_inc(x_116); +x_147 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_147, 0, x_116); +lean_ctor_set(x_147, 1, x_146); +x_148 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__42; +lean_inc(x_118); lean_inc(x_119); -lean_inc(x_120); -x_150 = l_Lean_addMacroScope(x_120, x_149, x_119); -x_151 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__41; -lean_inc(x_117); -x_152 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_152, 0, x_117); -lean_ctor_set(x_152, 1, x_151); -lean_ctor_set(x_152, 2, x_150); -lean_ctor_set(x_152, 3, x_125); -x_153 = lean_array_push(x_139, x_152); -x_154 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__38; -x_155 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_155, 0, x_131); -lean_ctor_set(x_155, 1, x_154); -lean_ctor_set(x_155, 2, x_153); -x_156 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__46; +x_149 = l_Lean_addMacroScope(x_119, x_148, x_118); +x_150 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__41; +lean_inc(x_116); +x_151 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_151, 0, x_116); +lean_ctor_set(x_151, 1, x_150); +lean_ctor_set(x_151, 2, x_149); +lean_ctor_set(x_151, 3, x_124); +x_152 = lean_array_push(x_138, x_151); +x_153 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__38; +x_154 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_154, 0, x_130); +lean_ctor_set(x_154, 1, x_153); +lean_ctor_set(x_154, 2, x_152); +x_155 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__46; +lean_inc(x_118); lean_inc(x_119); -lean_inc(x_120); -x_157 = l_Lean_addMacroScope(x_120, x_156, x_119); -x_158 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__45; -lean_inc(x_117); -x_159 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_159, 0, x_117); -lean_ctor_set(x_159, 1, x_158); -lean_ctor_set(x_159, 2, x_157); -lean_ctor_set(x_159, 3, x_125); -x_160 = lean_array_push(x_129, x_159); -x_161 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_161, 0, x_131); -lean_ctor_set(x_161, 1, x_132); -lean_ctor_set(x_161, 2, x_160); -x_162 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__47; -x_163 = lean_array_push(x_162, x_148); -lean_inc(x_163); -x_164 = lean_array_push(x_163, x_155); -lean_inc(x_161); -x_165 = lean_array_push(x_164, x_161); -x_166 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__35; -x_167 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_167, 0, x_131); -lean_ctor_set(x_167, 1, x_166); -lean_ctor_set(x_167, 2, x_165); -x_168 = lean_array_push(x_129, x_167); -x_169 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_169, 0, x_131); -lean_ctor_set(x_169, 1, x_132); -lean_ctor_set(x_169, 2, x_168); -x_170 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__48; -lean_inc(x_117); -x_171 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_171, 0, x_117); -lean_ctor_set(x_171, 1, x_170); -x_172 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__51; -lean_inc(x_117); -x_173 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_173, 0, x_117); -lean_ctor_set(x_173, 1, x_172); -x_174 = lean_array_push(x_129, x_173); -x_175 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__50; -x_176 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_176, 0, x_131); -lean_ctor_set(x_176, 1, x_175); -lean_ctor_set(x_176, 2, x_174); -x_177 = lean_array_push(x_162, x_169); -lean_inc(x_171); -x_178 = lean_array_push(x_177, x_171); -lean_inc(x_176); -x_179 = lean_array_push(x_178, x_176); -x_180 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__33; -x_181 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_181, 0, x_131); -lean_ctor_set(x_181, 1, x_180); -lean_ctor_set(x_181, 2, x_179); -x_182 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__55; -x_183 = l_Lean_addMacroScope(x_120, x_182, x_119); -x_184 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__54; -x_185 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_185, 0, x_117); -lean_ctor_set(x_185, 1, x_184); -lean_ctor_set(x_185, 2, x_183); -lean_ctor_set(x_185, 3, x_125); -x_186 = lean_array_push(x_139, x_185); -x_187 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_187, 0, x_131); -lean_ctor_set(x_187, 1, x_154); -lean_ctor_set(x_187, 2, x_186); -x_188 = lean_array_push(x_163, x_187); -x_189 = lean_array_push(x_188, x_161); -x_190 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_190, 0, x_131); -lean_ctor_set(x_190, 1, x_166); -lean_ctor_set(x_190, 2, x_189); -x_191 = lean_array_push(x_129, x_190); -x_192 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_192, 0, x_131); -lean_ctor_set(x_192, 1, x_132); -lean_ctor_set(x_192, 2, x_191); -x_193 = lean_array_push(x_162, x_192); -x_194 = lean_array_push(x_193, x_171); -x_195 = lean_array_push(x_194, x_176); -x_196 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_196, 0, x_131); -lean_ctor_set(x_196, 1, x_180); -lean_ctor_set(x_196, 2, x_195); -x_197 = lean_array_push(x_134, x_181); -x_198 = lean_array_push(x_197, x_196); -x_199 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_199, 0, x_131); -lean_ctor_set(x_199, 1, x_132); -lean_ctor_set(x_199, 2, x_198); -x_200 = lean_array_push(x_162, x_146); -x_201 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__14; -x_202 = lean_array_push(x_200, x_201); -x_203 = lean_array_push(x_202, x_199); -x_204 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__30; -x_205 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_205, 0, x_131); -lean_ctor_set(x_205, 1, x_204); -lean_ctor_set(x_205, 2, x_203); -x_206 = lean_array_push(x_129, x_205); -x_207 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_207, 0, x_131); -lean_ctor_set(x_207, 1, x_132); -lean_ctor_set(x_207, 2, x_206); -x_208 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__56; -x_209 = lean_array_push(x_208, x_122); -x_210 = lean_array_push(x_209, x_144); -x_211 = lean_array_push(x_210, x_201); -x_212 = lean_array_push(x_211, x_207); -x_213 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__8; -x_214 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_214, 0, x_131); -lean_ctor_set(x_214, 1, x_213); -lean_ctor_set(x_214, 2, x_212); -x_215 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_215, 0, x_214); -lean_ctor_set(x_215, 1, x_118); -return x_215; +x_156 = l_Lean_addMacroScope(x_119, x_155, x_118); +x_157 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__45; +lean_inc(x_116); +x_158 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_158, 0, x_116); +lean_ctor_set(x_158, 1, x_157); +lean_ctor_set(x_158, 2, x_156); +lean_ctor_set(x_158, 3, x_124); +x_159 = lean_array_push(x_128, x_158); +x_160 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_160, 0, x_130); +lean_ctor_set(x_160, 1, x_131); +lean_ctor_set(x_160, 2, x_159); +x_161 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__47; +x_162 = lean_array_push(x_161, x_147); +lean_inc(x_162); +x_163 = lean_array_push(x_162, x_154); +lean_inc(x_160); +x_164 = lean_array_push(x_163, x_160); +x_165 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__35; +x_166 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_166, 0, x_130); +lean_ctor_set(x_166, 1, x_165); +lean_ctor_set(x_166, 2, x_164); +x_167 = lean_array_push(x_128, x_166); +x_168 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_168, 0, x_130); +lean_ctor_set(x_168, 1, x_131); +lean_ctor_set(x_168, 2, x_167); +x_169 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__48; +lean_inc(x_116); +x_170 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_170, 0, x_116); +lean_ctor_set(x_170, 1, x_169); +x_171 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__51; +lean_inc(x_116); +x_172 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_172, 0, x_116); +lean_ctor_set(x_172, 1, x_171); +x_173 = lean_array_push(x_128, x_172); +x_174 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__50; +x_175 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_175, 0, x_130); +lean_ctor_set(x_175, 1, x_174); +lean_ctor_set(x_175, 2, x_173); +x_176 = lean_array_push(x_161, x_168); +lean_inc(x_170); +x_177 = lean_array_push(x_176, x_170); +lean_inc(x_175); +x_178 = lean_array_push(x_177, x_175); +x_179 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__33; +x_180 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_180, 0, x_130); +lean_ctor_set(x_180, 1, x_179); +lean_ctor_set(x_180, 2, x_178); +x_181 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__55; +x_182 = l_Lean_addMacroScope(x_119, x_181, x_118); +x_183 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__54; +x_184 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_184, 0, x_116); +lean_ctor_set(x_184, 1, x_183); +lean_ctor_set(x_184, 2, x_182); +lean_ctor_set(x_184, 3, x_124); +x_185 = lean_array_push(x_138, x_184); +x_186 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_186, 0, x_130); +lean_ctor_set(x_186, 1, x_153); +lean_ctor_set(x_186, 2, x_185); +x_187 = lean_array_push(x_162, x_186); +x_188 = lean_array_push(x_187, x_160); +x_189 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_189, 0, x_130); +lean_ctor_set(x_189, 1, x_165); +lean_ctor_set(x_189, 2, x_188); +x_190 = lean_array_push(x_128, x_189); +x_191 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_191, 0, x_130); +lean_ctor_set(x_191, 1, x_131); +lean_ctor_set(x_191, 2, x_190); +x_192 = lean_array_push(x_161, x_191); +x_193 = lean_array_push(x_192, x_170); +x_194 = lean_array_push(x_193, x_175); +x_195 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_195, 0, x_130); +lean_ctor_set(x_195, 1, x_179); +lean_ctor_set(x_195, 2, x_194); +x_196 = lean_array_push(x_133, x_180); +x_197 = lean_array_push(x_196, x_195); +x_198 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_198, 0, x_130); +lean_ctor_set(x_198, 1, x_131); +lean_ctor_set(x_198, 2, x_197); +x_199 = lean_array_push(x_161, x_145); +x_200 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__14; +x_201 = lean_array_push(x_199, x_200); +x_202 = lean_array_push(x_201, x_198); +x_203 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__30; +x_204 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_204, 0, x_130); +lean_ctor_set(x_204, 1, x_203); +lean_ctor_set(x_204, 2, x_202); +x_205 = lean_array_push(x_128, x_204); +x_206 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_206, 0, x_130); +lean_ctor_set(x_206, 1, x_131); +lean_ctor_set(x_206, 2, x_205); +x_207 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__56; +x_208 = lean_array_push(x_207, x_121); +x_209 = lean_array_push(x_208, x_143); +x_210 = lean_array_push(x_209, x_200); +x_211 = lean_array_push(x_210, x_206); +x_212 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__8; +x_213 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_213, 0, x_130); +lean_ctor_set(x_213, 1, x_212); +lean_ctor_set(x_213, 2, x_211); +x_214 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_214, 0, x_213); +lean_ctor_set(x_214, 1, x_117); +return x_214; } } } else { -lean_object* x_216; lean_object* x_217; lean_object* x_218; uint8_t x_219; -x_216 = lean_unsigned_to_nat(0u); -x_217 = l_Lean_Syntax_getArg(x_9, x_216); +lean_object* x_215; lean_object* x_216; lean_object* x_217; uint8_t x_218; +x_215 = lean_unsigned_to_nat(0u); +x_216 = l_Lean_Syntax_getArg(x_9, x_215); lean_dec(x_9); -x_218 = l_Classical_tacticBy__cases_____x3a_____closed__14; -lean_inc(x_217); -x_219 = l_Lean_Syntax_isOfKind(x_217, x_218); -if (x_219 == 0) +x_217 = l_Classical_tacticBy__cases___x3a_____closed__14; +lean_inc(x_216); +x_218 = l_Lean_Syntax_isOfKind(x_216, x_217); +if (x_218 == 0) { -lean_object* x_220; lean_object* x_221; -lean_dec(x_217); +lean_object* x_219; lean_object* x_220; +lean_dec(x_216); lean_dec(x_2); lean_dec(x_1); -x_220 = lean_box(1); -x_221 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_221, 0, x_220); -lean_ctor_set(x_221, 1, x_3); -return x_221; +x_219 = lean_box(1); +x_220 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_220, 0, x_219); +lean_ctor_set(x_220, 1, x_3); +return x_220; } else { -lean_object* x_222; lean_object* x_223; uint8_t x_224; -x_222 = l_Lean_Syntax_getArg(x_1, x_11); +lean_object* x_221; lean_object* x_222; uint8_t x_223; +x_221 = l_Lean_Syntax_getArg(x_1, x_10); lean_dec(x_1); lean_inc(x_2); -x_223 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_224 = !lean_is_exclusive(x_223); -if (x_224 == 0) +x_222 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_223 = !lean_is_exclusive(x_222); +if (x_223 == 0) { -lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; -x_225 = lean_ctor_get(x_223, 0); -x_226 = lean_ctor_get(x_2, 2); +lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; +x_224 = lean_ctor_get(x_222, 0); +x_225 = lean_ctor_get(x_2, 2); +lean_inc(x_225); +x_226 = lean_ctor_get(x_2, 1); lean_inc(x_226); -x_227 = lean_ctor_get(x_2, 1); -lean_inc(x_227); lean_dec(x_2); -x_228 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__7; +x_227 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__7; +lean_inc(x_224); +x_228 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_228, 0, x_224); +lean_ctor_set(x_228, 1, x_227); +x_229 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__22; lean_inc(x_225); -x_229 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_229, 0, x_225); -lean_ctor_set(x_229, 1, x_228); -x_230 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__22; lean_inc(x_226); -lean_inc(x_227); -x_231 = l_Lean_addMacroScope(x_227, x_230, x_226); -x_232 = lean_box(0); -x_233 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__21; -x_234 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__25; -lean_inc(x_225); -x_235 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_235, 0, x_225); -lean_ctor_set(x_235, 1, x_233); -lean_ctor_set(x_235, 2, x_231); -lean_ctor_set(x_235, 3, x_234); -x_236 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__26; -x_237 = lean_array_push(x_236, x_222); -x_238 = lean_box(2); -x_239 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__10; -x_240 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_240, 0, x_238); -lean_ctor_set(x_240, 1, x_239); -lean_ctor_set(x_240, 2, x_237); -x_241 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__27; -x_242 = lean_array_push(x_241, x_235); -x_243 = lean_array_push(x_242, x_240); -x_244 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__18; -x_245 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_245, 0, x_238); -lean_ctor_set(x_245, 1, x_244); -lean_ctor_set(x_245, 2, x_243); -x_246 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__28; -x_247 = lean_array_push(x_246, x_245); -x_248 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__12; -x_249 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_249, 0, x_238); -lean_ctor_set(x_249, 1, x_248); -lean_ctor_set(x_249, 2, x_247); -x_250 = lean_array_push(x_236, x_249); -x_251 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_251, 0, x_238); -lean_ctor_set(x_251, 1, x_239); -lean_ctor_set(x_251, 2, x_250); -x_252 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__31; +x_230 = l_Lean_addMacroScope(x_226, x_229, x_225); +x_231 = lean_box(0); +x_232 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__21; +x_233 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__25; +lean_inc(x_224); +x_234 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_234, 0, x_224); +lean_ctor_set(x_234, 1, x_232); +lean_ctor_set(x_234, 2, x_230); +lean_ctor_set(x_234, 3, x_233); +x_235 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__26; +x_236 = lean_array_push(x_235, x_221); +x_237 = lean_box(2); +x_238 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__10; +x_239 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_239, 0, x_237); +lean_ctor_set(x_239, 1, x_238); +lean_ctor_set(x_239, 2, x_236); +x_240 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__27; +x_241 = lean_array_push(x_240, x_234); +x_242 = lean_array_push(x_241, x_239); +x_243 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__18; +x_244 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_244, 0, x_237); +lean_ctor_set(x_244, 1, x_243); +lean_ctor_set(x_244, 2, x_242); +x_245 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__28; +x_246 = lean_array_push(x_245, x_244); +x_247 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__12; +x_248 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_248, 0, x_237); +lean_ctor_set(x_248, 1, x_247); +lean_ctor_set(x_248, 2, x_246); +x_249 = lean_array_push(x_235, x_248); +x_250 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_250, 0, x_237); +lean_ctor_set(x_250, 1, x_238); +lean_ctor_set(x_250, 2, x_249); +x_251 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__31; +lean_inc(x_224); +x_252 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_252, 0, x_224); +lean_ctor_set(x_252, 1, x_251); +x_253 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__36; +lean_inc(x_224); +x_254 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_254, 0, x_224); +lean_ctor_set(x_254, 1, x_253); +x_255 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__42; lean_inc(x_225); -x_253 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_253, 0, x_225); -lean_ctor_set(x_253, 1, x_252); -x_254 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__36; -lean_inc(x_225); -x_255 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_255, 0, x_225); -lean_ctor_set(x_255, 1, x_254); -x_256 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__42; lean_inc(x_226); -lean_inc(x_227); -x_257 = l_Lean_addMacroScope(x_227, x_256, x_226); -x_258 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__41; -lean_inc(x_225); -x_259 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_259, 0, x_225); -lean_ctor_set(x_259, 1, x_258); -lean_ctor_set(x_259, 2, x_257); -lean_ctor_set(x_259, 3, x_232); -x_260 = lean_array_push(x_246, x_259); -x_261 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__38; -x_262 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_262, 0, x_238); -lean_ctor_set(x_262, 1, x_261); -lean_ctor_set(x_262, 2, x_260); -x_263 = lean_array_push(x_236, x_217); -x_264 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_264, 0, x_238); -lean_ctor_set(x_264, 1, x_239); -lean_ctor_set(x_264, 2, x_263); -x_265 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__47; -x_266 = lean_array_push(x_265, x_255); -lean_inc(x_266); -x_267 = lean_array_push(x_266, x_262); -lean_inc(x_264); -x_268 = lean_array_push(x_267, x_264); -x_269 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__35; -x_270 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_270, 0, x_238); -lean_ctor_set(x_270, 1, x_269); -lean_ctor_set(x_270, 2, x_268); -x_271 = lean_array_push(x_236, x_270); -x_272 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_272, 0, x_238); -lean_ctor_set(x_272, 1, x_239); -lean_ctor_set(x_272, 2, x_271); -x_273 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__48; -lean_inc(x_225); -x_274 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_274, 0, x_225); -lean_ctor_set(x_274, 1, x_273); -x_275 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__51; -lean_inc(x_225); -x_276 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_276, 0, x_225); -lean_ctor_set(x_276, 1, x_275); -x_277 = lean_array_push(x_236, x_276); -x_278 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__50; -x_279 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_279, 0, x_238); -lean_ctor_set(x_279, 1, x_278); -lean_ctor_set(x_279, 2, x_277); -x_280 = lean_array_push(x_265, x_272); -lean_inc(x_274); -x_281 = lean_array_push(x_280, x_274); -lean_inc(x_279); -x_282 = lean_array_push(x_281, x_279); -x_283 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__33; -x_284 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_284, 0, x_238); -lean_ctor_set(x_284, 1, x_283); -lean_ctor_set(x_284, 2, x_282); -x_285 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__55; -x_286 = l_Lean_addMacroScope(x_227, x_285, x_226); -x_287 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__54; -x_288 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_288, 0, x_225); -lean_ctor_set(x_288, 1, x_287); -lean_ctor_set(x_288, 2, x_286); -lean_ctor_set(x_288, 3, x_232); -x_289 = lean_array_push(x_246, x_288); -x_290 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_290, 0, x_238); -lean_ctor_set(x_290, 1, x_261); -lean_ctor_set(x_290, 2, x_289); -x_291 = lean_array_push(x_266, x_290); -x_292 = lean_array_push(x_291, x_264); -x_293 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_293, 0, x_238); -lean_ctor_set(x_293, 1, x_269); -lean_ctor_set(x_293, 2, x_292); -x_294 = lean_array_push(x_236, x_293); -x_295 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_295, 0, x_238); -lean_ctor_set(x_295, 1, x_239); -lean_ctor_set(x_295, 2, x_294); -x_296 = lean_array_push(x_265, x_295); -x_297 = lean_array_push(x_296, x_274); -x_298 = lean_array_push(x_297, x_279); -x_299 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_299, 0, x_238); -lean_ctor_set(x_299, 1, x_283); -lean_ctor_set(x_299, 2, x_298); -x_300 = lean_array_push(x_241, x_284); -x_301 = lean_array_push(x_300, x_299); -x_302 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_302, 0, x_238); -lean_ctor_set(x_302, 1, x_239); -lean_ctor_set(x_302, 2, x_301); -x_303 = lean_array_push(x_265, x_253); -x_304 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__14; -x_305 = lean_array_push(x_303, x_304); -x_306 = lean_array_push(x_305, x_302); -x_307 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__30; -x_308 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_308, 0, x_238); -lean_ctor_set(x_308, 1, x_307); -lean_ctor_set(x_308, 2, x_306); -x_309 = lean_array_push(x_236, x_308); -x_310 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_310, 0, x_238); -lean_ctor_set(x_310, 1, x_239); -lean_ctor_set(x_310, 2, x_309); -x_311 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__56; -x_312 = lean_array_push(x_311, x_229); -x_313 = lean_array_push(x_312, x_251); -x_314 = lean_array_push(x_313, x_304); -x_315 = lean_array_push(x_314, x_310); -x_316 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__8; -x_317 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_317, 0, x_238); -lean_ctor_set(x_317, 1, x_316); -lean_ctor_set(x_317, 2, x_315); -lean_ctor_set(x_223, 0, x_317); -return x_223; +x_256 = l_Lean_addMacroScope(x_226, x_255, x_225); +x_257 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__41; +lean_inc(x_224); +x_258 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_258, 0, x_224); +lean_ctor_set(x_258, 1, x_257); +lean_ctor_set(x_258, 2, x_256); +lean_ctor_set(x_258, 3, x_231); +x_259 = lean_array_push(x_245, x_258); +x_260 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__38; +x_261 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_261, 0, x_237); +lean_ctor_set(x_261, 1, x_260); +lean_ctor_set(x_261, 2, x_259); +x_262 = lean_array_push(x_235, x_216); +x_263 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_263, 0, x_237); +lean_ctor_set(x_263, 1, x_238); +lean_ctor_set(x_263, 2, x_262); +x_264 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__47; +x_265 = lean_array_push(x_264, x_254); +lean_inc(x_265); +x_266 = lean_array_push(x_265, x_261); +lean_inc(x_263); +x_267 = lean_array_push(x_266, x_263); +x_268 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__35; +x_269 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_269, 0, x_237); +lean_ctor_set(x_269, 1, x_268); +lean_ctor_set(x_269, 2, x_267); +x_270 = lean_array_push(x_235, x_269); +x_271 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_271, 0, x_237); +lean_ctor_set(x_271, 1, x_238); +lean_ctor_set(x_271, 2, x_270); +x_272 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__48; +lean_inc(x_224); +x_273 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_273, 0, x_224); +lean_ctor_set(x_273, 1, x_272); +x_274 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__51; +lean_inc(x_224); +x_275 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_275, 0, x_224); +lean_ctor_set(x_275, 1, x_274); +x_276 = lean_array_push(x_235, x_275); +x_277 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__50; +x_278 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_278, 0, x_237); +lean_ctor_set(x_278, 1, x_277); +lean_ctor_set(x_278, 2, x_276); +x_279 = lean_array_push(x_264, x_271); +lean_inc(x_273); +x_280 = lean_array_push(x_279, x_273); +lean_inc(x_278); +x_281 = lean_array_push(x_280, x_278); +x_282 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__33; +x_283 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_283, 0, x_237); +lean_ctor_set(x_283, 1, x_282); +lean_ctor_set(x_283, 2, x_281); +x_284 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__55; +x_285 = l_Lean_addMacroScope(x_226, x_284, x_225); +x_286 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__54; +x_287 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_287, 0, x_224); +lean_ctor_set(x_287, 1, x_286); +lean_ctor_set(x_287, 2, x_285); +lean_ctor_set(x_287, 3, x_231); +x_288 = lean_array_push(x_245, x_287); +x_289 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_289, 0, x_237); +lean_ctor_set(x_289, 1, x_260); +lean_ctor_set(x_289, 2, x_288); +x_290 = lean_array_push(x_265, x_289); +x_291 = lean_array_push(x_290, x_263); +x_292 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_292, 0, x_237); +lean_ctor_set(x_292, 1, x_268); +lean_ctor_set(x_292, 2, x_291); +x_293 = lean_array_push(x_235, x_292); +x_294 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_294, 0, x_237); +lean_ctor_set(x_294, 1, x_238); +lean_ctor_set(x_294, 2, x_293); +x_295 = lean_array_push(x_264, x_294); +x_296 = lean_array_push(x_295, x_273); +x_297 = lean_array_push(x_296, x_278); +x_298 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_298, 0, x_237); +lean_ctor_set(x_298, 1, x_282); +lean_ctor_set(x_298, 2, x_297); +x_299 = lean_array_push(x_240, x_283); +x_300 = lean_array_push(x_299, x_298); +x_301 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_301, 0, x_237); +lean_ctor_set(x_301, 1, x_238); +lean_ctor_set(x_301, 2, x_300); +x_302 = lean_array_push(x_264, x_252); +x_303 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__14; +x_304 = lean_array_push(x_302, x_303); +x_305 = lean_array_push(x_304, x_301); +x_306 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__30; +x_307 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_307, 0, x_237); +lean_ctor_set(x_307, 1, x_306); +lean_ctor_set(x_307, 2, x_305); +x_308 = lean_array_push(x_235, x_307); +x_309 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_309, 0, x_237); +lean_ctor_set(x_309, 1, x_238); +lean_ctor_set(x_309, 2, x_308); +x_310 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__56; +x_311 = lean_array_push(x_310, x_228); +x_312 = lean_array_push(x_311, x_250); +x_313 = lean_array_push(x_312, x_303); +x_314 = lean_array_push(x_313, x_309); +x_315 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__8; +x_316 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_316, 0, x_237); +lean_ctor_set(x_316, 1, x_315); +lean_ctor_set(x_316, 2, x_314); +lean_ctor_set(x_222, 0, x_316); +return x_222; } else { -lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; -x_318 = lean_ctor_get(x_223, 0); -x_319 = lean_ctor_get(x_223, 1); -lean_inc(x_319); +lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; +x_317 = lean_ctor_get(x_222, 0); +x_318 = lean_ctor_get(x_222, 1); lean_inc(x_318); -lean_dec(x_223); -x_320 = lean_ctor_get(x_2, 2); +lean_inc(x_317); +lean_dec(x_222); +x_319 = lean_ctor_get(x_2, 2); +lean_inc(x_319); +x_320 = lean_ctor_get(x_2, 1); lean_inc(x_320); -x_321 = lean_ctor_get(x_2, 1); -lean_inc(x_321); lean_dec(x_2); -x_322 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__7; -lean_inc(x_318); -x_323 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_323, 0, x_318); -lean_ctor_set(x_323, 1, x_322); -x_324 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__22; +x_321 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__7; +lean_inc(x_317); +x_322 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_322, 0, x_317); +lean_ctor_set(x_322, 1, x_321); +x_323 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__22; +lean_inc(x_319); lean_inc(x_320); -lean_inc(x_321); -x_325 = l_Lean_addMacroScope(x_321, x_324, x_320); -x_326 = lean_box(0); -x_327 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__21; -x_328 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__25; -lean_inc(x_318); -x_329 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_329, 0, x_318); -lean_ctor_set(x_329, 1, x_327); -lean_ctor_set(x_329, 2, x_325); -lean_ctor_set(x_329, 3, x_328); -x_330 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__26; -x_331 = lean_array_push(x_330, x_222); -x_332 = lean_box(2); -x_333 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__10; -x_334 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_334, 0, x_332); -lean_ctor_set(x_334, 1, x_333); -lean_ctor_set(x_334, 2, x_331); -x_335 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__27; -x_336 = lean_array_push(x_335, x_329); -x_337 = lean_array_push(x_336, x_334); -x_338 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__18; -x_339 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_339, 0, x_332); -lean_ctor_set(x_339, 1, x_338); -lean_ctor_set(x_339, 2, x_337); -x_340 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__28; -x_341 = lean_array_push(x_340, x_339); -x_342 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__12; -x_343 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_343, 0, x_332); -lean_ctor_set(x_343, 1, x_342); -lean_ctor_set(x_343, 2, x_341); -x_344 = lean_array_push(x_330, x_343); -x_345 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_345, 0, x_332); -lean_ctor_set(x_345, 1, x_333); -lean_ctor_set(x_345, 2, x_344); -x_346 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__31; -lean_inc(x_318); -x_347 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_347, 0, x_318); -lean_ctor_set(x_347, 1, x_346); -x_348 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__36; -lean_inc(x_318); -x_349 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_349, 0, x_318); -lean_ctor_set(x_349, 1, x_348); -x_350 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__42; +x_324 = l_Lean_addMacroScope(x_320, x_323, x_319); +x_325 = lean_box(0); +x_326 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__21; +x_327 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__25; +lean_inc(x_317); +x_328 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_328, 0, x_317); +lean_ctor_set(x_328, 1, x_326); +lean_ctor_set(x_328, 2, x_324); +lean_ctor_set(x_328, 3, x_327); +x_329 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__26; +x_330 = lean_array_push(x_329, x_221); +x_331 = lean_box(2); +x_332 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__10; +x_333 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_333, 0, x_331); +lean_ctor_set(x_333, 1, x_332); +lean_ctor_set(x_333, 2, x_330); +x_334 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__27; +x_335 = lean_array_push(x_334, x_328); +x_336 = lean_array_push(x_335, x_333); +x_337 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__18; +x_338 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_338, 0, x_331); +lean_ctor_set(x_338, 1, x_337); +lean_ctor_set(x_338, 2, x_336); +x_339 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__28; +x_340 = lean_array_push(x_339, x_338); +x_341 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__12; +x_342 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_342, 0, x_331); +lean_ctor_set(x_342, 1, x_341); +lean_ctor_set(x_342, 2, x_340); +x_343 = lean_array_push(x_329, x_342); +x_344 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_344, 0, x_331); +lean_ctor_set(x_344, 1, x_332); +lean_ctor_set(x_344, 2, x_343); +x_345 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__31; +lean_inc(x_317); +x_346 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_346, 0, x_317); +lean_ctor_set(x_346, 1, x_345); +x_347 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__36; +lean_inc(x_317); +x_348 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_348, 0, x_317); +lean_ctor_set(x_348, 1, x_347); +x_349 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__42; +lean_inc(x_319); lean_inc(x_320); -lean_inc(x_321); -x_351 = l_Lean_addMacroScope(x_321, x_350, x_320); -x_352 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__41; -lean_inc(x_318); -x_353 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_353, 0, x_318); -lean_ctor_set(x_353, 1, x_352); -lean_ctor_set(x_353, 2, x_351); -lean_ctor_set(x_353, 3, x_326); -x_354 = lean_array_push(x_340, x_353); -x_355 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__38; -x_356 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_356, 0, x_332); -lean_ctor_set(x_356, 1, x_355); -lean_ctor_set(x_356, 2, x_354); -x_357 = lean_array_push(x_330, x_217); -x_358 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_358, 0, x_332); -lean_ctor_set(x_358, 1, x_333); -lean_ctor_set(x_358, 2, x_357); -x_359 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__47; -x_360 = lean_array_push(x_359, x_349); -lean_inc(x_360); -x_361 = lean_array_push(x_360, x_356); -lean_inc(x_358); -x_362 = lean_array_push(x_361, x_358); -x_363 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__35; -x_364 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_364, 0, x_332); -lean_ctor_set(x_364, 1, x_363); -lean_ctor_set(x_364, 2, x_362); -x_365 = lean_array_push(x_330, x_364); -x_366 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_366, 0, x_332); -lean_ctor_set(x_366, 1, x_333); -lean_ctor_set(x_366, 2, x_365); -x_367 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__48; -lean_inc(x_318); -x_368 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_368, 0, x_318); -lean_ctor_set(x_368, 1, x_367); -x_369 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__51; -lean_inc(x_318); -x_370 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_370, 0, x_318); -lean_ctor_set(x_370, 1, x_369); -x_371 = lean_array_push(x_330, x_370); -x_372 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__50; -x_373 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_373, 0, x_332); -lean_ctor_set(x_373, 1, x_372); -lean_ctor_set(x_373, 2, x_371); -x_374 = lean_array_push(x_359, x_366); -lean_inc(x_368); -x_375 = lean_array_push(x_374, x_368); -lean_inc(x_373); -x_376 = lean_array_push(x_375, x_373); -x_377 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__33; -x_378 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_378, 0, x_332); -lean_ctor_set(x_378, 1, x_377); -lean_ctor_set(x_378, 2, x_376); -x_379 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__55; -x_380 = l_Lean_addMacroScope(x_321, x_379, x_320); -x_381 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__54; -x_382 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_382, 0, x_318); -lean_ctor_set(x_382, 1, x_381); -lean_ctor_set(x_382, 2, x_380); -lean_ctor_set(x_382, 3, x_326); -x_383 = lean_array_push(x_340, x_382); -x_384 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_384, 0, x_332); -lean_ctor_set(x_384, 1, x_355); -lean_ctor_set(x_384, 2, x_383); -x_385 = lean_array_push(x_360, x_384); -x_386 = lean_array_push(x_385, x_358); -x_387 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_387, 0, x_332); -lean_ctor_set(x_387, 1, x_363); -lean_ctor_set(x_387, 2, x_386); -x_388 = lean_array_push(x_330, x_387); -x_389 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_389, 0, x_332); -lean_ctor_set(x_389, 1, x_333); -lean_ctor_set(x_389, 2, x_388); -x_390 = lean_array_push(x_359, x_389); -x_391 = lean_array_push(x_390, x_368); -x_392 = lean_array_push(x_391, x_373); -x_393 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_393, 0, x_332); -lean_ctor_set(x_393, 1, x_377); -lean_ctor_set(x_393, 2, x_392); -x_394 = lean_array_push(x_335, x_378); -x_395 = lean_array_push(x_394, x_393); -x_396 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_396, 0, x_332); -lean_ctor_set(x_396, 1, x_333); -lean_ctor_set(x_396, 2, x_395); -x_397 = lean_array_push(x_359, x_347); -x_398 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__14; -x_399 = lean_array_push(x_397, x_398); -x_400 = lean_array_push(x_399, x_396); -x_401 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__30; -x_402 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_402, 0, x_332); -lean_ctor_set(x_402, 1, x_401); -lean_ctor_set(x_402, 2, x_400); -x_403 = lean_array_push(x_330, x_402); -x_404 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_404, 0, x_332); -lean_ctor_set(x_404, 1, x_333); -lean_ctor_set(x_404, 2, x_403); -x_405 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__56; -x_406 = lean_array_push(x_405, x_323); -x_407 = lean_array_push(x_406, x_345); -x_408 = lean_array_push(x_407, x_398); -x_409 = lean_array_push(x_408, x_404); -x_410 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__8; -x_411 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_411, 0, x_332); -lean_ctor_set(x_411, 1, x_410); -lean_ctor_set(x_411, 2, x_409); -x_412 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_412, 0, x_411); -lean_ctor_set(x_412, 1, x_319); -return x_412; +x_350 = l_Lean_addMacroScope(x_320, x_349, x_319); +x_351 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__41; +lean_inc(x_317); +x_352 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_352, 0, x_317); +lean_ctor_set(x_352, 1, x_351); +lean_ctor_set(x_352, 2, x_350); +lean_ctor_set(x_352, 3, x_325); +x_353 = lean_array_push(x_339, x_352); +x_354 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__38; +x_355 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_355, 0, x_331); +lean_ctor_set(x_355, 1, x_354); +lean_ctor_set(x_355, 2, x_353); +x_356 = lean_array_push(x_329, x_216); +x_357 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_357, 0, x_331); +lean_ctor_set(x_357, 1, x_332); +lean_ctor_set(x_357, 2, x_356); +x_358 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__47; +x_359 = lean_array_push(x_358, x_348); +lean_inc(x_359); +x_360 = lean_array_push(x_359, x_355); +lean_inc(x_357); +x_361 = lean_array_push(x_360, x_357); +x_362 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__35; +x_363 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_363, 0, x_331); +lean_ctor_set(x_363, 1, x_362); +lean_ctor_set(x_363, 2, x_361); +x_364 = lean_array_push(x_329, x_363); +x_365 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_365, 0, x_331); +lean_ctor_set(x_365, 1, x_332); +lean_ctor_set(x_365, 2, x_364); +x_366 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__48; +lean_inc(x_317); +x_367 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_367, 0, x_317); +lean_ctor_set(x_367, 1, x_366); +x_368 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__51; +lean_inc(x_317); +x_369 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_369, 0, x_317); +lean_ctor_set(x_369, 1, x_368); +x_370 = lean_array_push(x_329, x_369); +x_371 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__50; +x_372 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_372, 0, x_331); +lean_ctor_set(x_372, 1, x_371); +lean_ctor_set(x_372, 2, x_370); +x_373 = lean_array_push(x_358, x_365); +lean_inc(x_367); +x_374 = lean_array_push(x_373, x_367); +lean_inc(x_372); +x_375 = lean_array_push(x_374, x_372); +x_376 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__33; +x_377 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_377, 0, x_331); +lean_ctor_set(x_377, 1, x_376); +lean_ctor_set(x_377, 2, x_375); +x_378 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__55; +x_379 = l_Lean_addMacroScope(x_320, x_378, x_319); +x_380 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__54; +x_381 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_381, 0, x_317); +lean_ctor_set(x_381, 1, x_380); +lean_ctor_set(x_381, 2, x_379); +lean_ctor_set(x_381, 3, x_325); +x_382 = lean_array_push(x_339, x_381); +x_383 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_383, 0, x_331); +lean_ctor_set(x_383, 1, x_354); +lean_ctor_set(x_383, 2, x_382); +x_384 = lean_array_push(x_359, x_383); +x_385 = lean_array_push(x_384, x_357); +x_386 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_386, 0, x_331); +lean_ctor_set(x_386, 1, x_362); +lean_ctor_set(x_386, 2, x_385); +x_387 = lean_array_push(x_329, x_386); +x_388 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_388, 0, x_331); +lean_ctor_set(x_388, 1, x_332); +lean_ctor_set(x_388, 2, x_387); +x_389 = lean_array_push(x_358, x_388); +x_390 = lean_array_push(x_389, x_367); +x_391 = lean_array_push(x_390, x_372); +x_392 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_392, 0, x_331); +lean_ctor_set(x_392, 1, x_376); +lean_ctor_set(x_392, 2, x_391); +x_393 = lean_array_push(x_334, x_377); +x_394 = lean_array_push(x_393, x_392); +x_395 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_395, 0, x_331); +lean_ctor_set(x_395, 1, x_332); +lean_ctor_set(x_395, 2, x_394); +x_396 = lean_array_push(x_358, x_346); +x_397 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__14; +x_398 = lean_array_push(x_396, x_397); +x_399 = lean_array_push(x_398, x_395); +x_400 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__30; +x_401 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_401, 0, x_331); +lean_ctor_set(x_401, 1, x_400); +lean_ctor_set(x_401, 2, x_399); +x_402 = lean_array_push(x_329, x_401); +x_403 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_403, 0, x_331); +lean_ctor_set(x_403, 1, x_332); +lean_ctor_set(x_403, 2, x_402); +x_404 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__56; +x_405 = lean_array_push(x_404, x_322); +x_406 = lean_array_push(x_405, x_344); +x_407 = lean_array_push(x_406, x_397); +x_408 = lean_array_push(x_407, x_403); +x_409 = l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__8; +x_410 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_410, 0, x_331); +lean_ctor_set(x_410, 1, x_409); +lean_ctor_set(x_410, 2, x_408); +x_411 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_411, 0, x_410); +lean_ctor_set(x_411, 1, x_318); +return x_411; } } } @@ -1820,172 +1818,172 @@ lean_dec_ref(res); res = initialize_Init_NotationExtra(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Classical_tacticBy__cases_____x3a_____closed__1 = _init_l_Classical_tacticBy__cases_____x3a_____closed__1(); -lean_mark_persistent(l_Classical_tacticBy__cases_____x3a_____closed__1); -l_Classical_tacticBy__cases_____x3a_____closed__2 = _init_l_Classical_tacticBy__cases_____x3a_____closed__2(); -lean_mark_persistent(l_Classical_tacticBy__cases_____x3a_____closed__2); -l_Classical_tacticBy__cases_____x3a_____closed__3 = _init_l_Classical_tacticBy__cases_____x3a_____closed__3(); -lean_mark_persistent(l_Classical_tacticBy__cases_____x3a_____closed__3); -l_Classical_tacticBy__cases_____x3a_____closed__4 = _init_l_Classical_tacticBy__cases_____x3a_____closed__4(); -lean_mark_persistent(l_Classical_tacticBy__cases_____x3a_____closed__4); -l_Classical_tacticBy__cases_____x3a_____closed__5 = _init_l_Classical_tacticBy__cases_____x3a_____closed__5(); -lean_mark_persistent(l_Classical_tacticBy__cases_____x3a_____closed__5); -l_Classical_tacticBy__cases_____x3a_____closed__6 = _init_l_Classical_tacticBy__cases_____x3a_____closed__6(); -lean_mark_persistent(l_Classical_tacticBy__cases_____x3a_____closed__6); -l_Classical_tacticBy__cases_____x3a_____closed__7 = _init_l_Classical_tacticBy__cases_____x3a_____closed__7(); -lean_mark_persistent(l_Classical_tacticBy__cases_____x3a_____closed__7); -l_Classical_tacticBy__cases_____x3a_____closed__8 = _init_l_Classical_tacticBy__cases_____x3a_____closed__8(); -lean_mark_persistent(l_Classical_tacticBy__cases_____x3a_____closed__8); -l_Classical_tacticBy__cases_____x3a_____closed__9 = _init_l_Classical_tacticBy__cases_____x3a_____closed__9(); -lean_mark_persistent(l_Classical_tacticBy__cases_____x3a_____closed__9); -l_Classical_tacticBy__cases_____x3a_____closed__10 = _init_l_Classical_tacticBy__cases_____x3a_____closed__10(); -lean_mark_persistent(l_Classical_tacticBy__cases_____x3a_____closed__10); -l_Classical_tacticBy__cases_____x3a_____closed__11 = _init_l_Classical_tacticBy__cases_____x3a_____closed__11(); -lean_mark_persistent(l_Classical_tacticBy__cases_____x3a_____closed__11); -l_Classical_tacticBy__cases_____x3a_____closed__12 = _init_l_Classical_tacticBy__cases_____x3a_____closed__12(); -lean_mark_persistent(l_Classical_tacticBy__cases_____x3a_____closed__12); -l_Classical_tacticBy__cases_____x3a_____closed__13 = _init_l_Classical_tacticBy__cases_____x3a_____closed__13(); -lean_mark_persistent(l_Classical_tacticBy__cases_____x3a_____closed__13); -l_Classical_tacticBy__cases_____x3a_____closed__14 = _init_l_Classical_tacticBy__cases_____x3a_____closed__14(); -lean_mark_persistent(l_Classical_tacticBy__cases_____x3a_____closed__14); -l_Classical_tacticBy__cases_____x3a_____closed__15 = _init_l_Classical_tacticBy__cases_____x3a_____closed__15(); -lean_mark_persistent(l_Classical_tacticBy__cases_____x3a_____closed__15); -l_Classical_tacticBy__cases_____x3a_____closed__16 = _init_l_Classical_tacticBy__cases_____x3a_____closed__16(); -lean_mark_persistent(l_Classical_tacticBy__cases_____x3a_____closed__16); -l_Classical_tacticBy__cases_____x3a_____closed__17 = _init_l_Classical_tacticBy__cases_____x3a_____closed__17(); -lean_mark_persistent(l_Classical_tacticBy__cases_____x3a_____closed__17); -l_Classical_tacticBy__cases_____x3a_____closed__18 = _init_l_Classical_tacticBy__cases_____x3a_____closed__18(); -lean_mark_persistent(l_Classical_tacticBy__cases_____x3a_____closed__18); -l_Classical_tacticBy__cases_____x3a_____closed__19 = _init_l_Classical_tacticBy__cases_____x3a_____closed__19(); -lean_mark_persistent(l_Classical_tacticBy__cases_____x3a_____closed__19); -l_Classical_tacticBy__cases_____x3a_____closed__20 = _init_l_Classical_tacticBy__cases_____x3a_____closed__20(); -lean_mark_persistent(l_Classical_tacticBy__cases_____x3a_____closed__20); -l_Classical_tacticBy__cases_____x3a_____closed__21 = _init_l_Classical_tacticBy__cases_____x3a_____closed__21(); -lean_mark_persistent(l_Classical_tacticBy__cases_____x3a_____closed__21); -l_Classical_tacticBy__cases_____x3a_____closed__22 = _init_l_Classical_tacticBy__cases_____x3a_____closed__22(); -lean_mark_persistent(l_Classical_tacticBy__cases_____x3a_____closed__22); -l_Classical_tacticBy__cases_____x3a_____closed__23 = _init_l_Classical_tacticBy__cases_____x3a_____closed__23(); -lean_mark_persistent(l_Classical_tacticBy__cases_____x3a_____closed__23); -l_Classical_tacticBy__cases_____x3a_____closed__24 = _init_l_Classical_tacticBy__cases_____x3a_____closed__24(); -lean_mark_persistent(l_Classical_tacticBy__cases_____x3a_____closed__24); -l_Classical_tacticBy__cases_____x3a_____closed__25 = _init_l_Classical_tacticBy__cases_____x3a_____closed__25(); -lean_mark_persistent(l_Classical_tacticBy__cases_____x3a_____closed__25); -l_Classical_tacticBy__cases_____x3a_____closed__26 = _init_l_Classical_tacticBy__cases_____x3a_____closed__26(); -lean_mark_persistent(l_Classical_tacticBy__cases_____x3a_____closed__26); -l_Classical_tacticBy__cases_____x3a__ = _init_l_Classical_tacticBy__cases_____x3a__(); -lean_mark_persistent(l_Classical_tacticBy__cases_____x3a__); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__1 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__1(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__1); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__2 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__2(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__2); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__3 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__3(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__3); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__4 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__4(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__4); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__5 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__5(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__5); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__6 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__6(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__6); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__7 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__7(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__7); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__8 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__8(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__8); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__9 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__9(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__9); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__10 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__10(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__10); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__11 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__11(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__11); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__12 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__12(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__12); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__13 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__13(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__13); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__14 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__14(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__14); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__15 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__15(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__15); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__16 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__16(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__16); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__17 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__17(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__17); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__18 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__18(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__18); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__19 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__19(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__19); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__20 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__20(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__20); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__21 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__21(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__21); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__22 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__22(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__22); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__23 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__23(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__23); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__24 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__24(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__24); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__25 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__25(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__25); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__26 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__26(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__26); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__27 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__27(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__27); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__28 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__28(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__28); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__29 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__29(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__29); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__30 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__30(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__30); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__31 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__31(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__31); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__32 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__32(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__32); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__33 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__33(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__33); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__34 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__34(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__34); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__35 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__35(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__35); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__36 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__36(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__36); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__37 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__37(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__37); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__38 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__38(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__38); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__39 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__39(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__39); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__40 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__40(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__40); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__41 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__41(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__41); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__42 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__42(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__42); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__43 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__43(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__43); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__44 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__44(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__44); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__45 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__45(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__45); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__46 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__46(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__46); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__47 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__47(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__47); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__48 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__48(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__48); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__49 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__49(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__49); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__50 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__50(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__50); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__51 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__51(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__51); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__52 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__52(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__52); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__53 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__53(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__53); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__54 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__54(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__54); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__55 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__55(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__55); -l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__56 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__56(); -lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases_____x3a____1___closed__56); +l_Classical_tacticBy__cases___x3a_____closed__1 = _init_l_Classical_tacticBy__cases___x3a_____closed__1(); +lean_mark_persistent(l_Classical_tacticBy__cases___x3a_____closed__1); +l_Classical_tacticBy__cases___x3a_____closed__2 = _init_l_Classical_tacticBy__cases___x3a_____closed__2(); +lean_mark_persistent(l_Classical_tacticBy__cases___x3a_____closed__2); +l_Classical_tacticBy__cases___x3a_____closed__3 = _init_l_Classical_tacticBy__cases___x3a_____closed__3(); +lean_mark_persistent(l_Classical_tacticBy__cases___x3a_____closed__3); +l_Classical_tacticBy__cases___x3a_____closed__4 = _init_l_Classical_tacticBy__cases___x3a_____closed__4(); +lean_mark_persistent(l_Classical_tacticBy__cases___x3a_____closed__4); +l_Classical_tacticBy__cases___x3a_____closed__5 = _init_l_Classical_tacticBy__cases___x3a_____closed__5(); +lean_mark_persistent(l_Classical_tacticBy__cases___x3a_____closed__5); +l_Classical_tacticBy__cases___x3a_____closed__6 = _init_l_Classical_tacticBy__cases___x3a_____closed__6(); +lean_mark_persistent(l_Classical_tacticBy__cases___x3a_____closed__6); +l_Classical_tacticBy__cases___x3a_____closed__7 = _init_l_Classical_tacticBy__cases___x3a_____closed__7(); +lean_mark_persistent(l_Classical_tacticBy__cases___x3a_____closed__7); +l_Classical_tacticBy__cases___x3a_____closed__8 = _init_l_Classical_tacticBy__cases___x3a_____closed__8(); +lean_mark_persistent(l_Classical_tacticBy__cases___x3a_____closed__8); +l_Classical_tacticBy__cases___x3a_____closed__9 = _init_l_Classical_tacticBy__cases___x3a_____closed__9(); +lean_mark_persistent(l_Classical_tacticBy__cases___x3a_____closed__9); +l_Classical_tacticBy__cases___x3a_____closed__10 = _init_l_Classical_tacticBy__cases___x3a_____closed__10(); +lean_mark_persistent(l_Classical_tacticBy__cases___x3a_____closed__10); +l_Classical_tacticBy__cases___x3a_____closed__11 = _init_l_Classical_tacticBy__cases___x3a_____closed__11(); +lean_mark_persistent(l_Classical_tacticBy__cases___x3a_____closed__11); +l_Classical_tacticBy__cases___x3a_____closed__12 = _init_l_Classical_tacticBy__cases___x3a_____closed__12(); +lean_mark_persistent(l_Classical_tacticBy__cases___x3a_____closed__12); +l_Classical_tacticBy__cases___x3a_____closed__13 = _init_l_Classical_tacticBy__cases___x3a_____closed__13(); +lean_mark_persistent(l_Classical_tacticBy__cases___x3a_____closed__13); +l_Classical_tacticBy__cases___x3a_____closed__14 = _init_l_Classical_tacticBy__cases___x3a_____closed__14(); +lean_mark_persistent(l_Classical_tacticBy__cases___x3a_____closed__14); +l_Classical_tacticBy__cases___x3a_____closed__15 = _init_l_Classical_tacticBy__cases___x3a_____closed__15(); +lean_mark_persistent(l_Classical_tacticBy__cases___x3a_____closed__15); +l_Classical_tacticBy__cases___x3a_____closed__16 = _init_l_Classical_tacticBy__cases___x3a_____closed__16(); +lean_mark_persistent(l_Classical_tacticBy__cases___x3a_____closed__16); +l_Classical_tacticBy__cases___x3a_____closed__17 = _init_l_Classical_tacticBy__cases___x3a_____closed__17(); +lean_mark_persistent(l_Classical_tacticBy__cases___x3a_____closed__17); +l_Classical_tacticBy__cases___x3a_____closed__18 = _init_l_Classical_tacticBy__cases___x3a_____closed__18(); +lean_mark_persistent(l_Classical_tacticBy__cases___x3a_____closed__18); +l_Classical_tacticBy__cases___x3a_____closed__19 = _init_l_Classical_tacticBy__cases___x3a_____closed__19(); +lean_mark_persistent(l_Classical_tacticBy__cases___x3a_____closed__19); +l_Classical_tacticBy__cases___x3a_____closed__20 = _init_l_Classical_tacticBy__cases___x3a_____closed__20(); +lean_mark_persistent(l_Classical_tacticBy__cases___x3a_____closed__20); +l_Classical_tacticBy__cases___x3a_____closed__21 = _init_l_Classical_tacticBy__cases___x3a_____closed__21(); +lean_mark_persistent(l_Classical_tacticBy__cases___x3a_____closed__21); +l_Classical_tacticBy__cases___x3a_____closed__22 = _init_l_Classical_tacticBy__cases___x3a_____closed__22(); +lean_mark_persistent(l_Classical_tacticBy__cases___x3a_____closed__22); +l_Classical_tacticBy__cases___x3a_____closed__23 = _init_l_Classical_tacticBy__cases___x3a_____closed__23(); +lean_mark_persistent(l_Classical_tacticBy__cases___x3a_____closed__23); +l_Classical_tacticBy__cases___x3a_____closed__24 = _init_l_Classical_tacticBy__cases___x3a_____closed__24(); +lean_mark_persistent(l_Classical_tacticBy__cases___x3a_____closed__24); +l_Classical_tacticBy__cases___x3a_____closed__25 = _init_l_Classical_tacticBy__cases___x3a_____closed__25(); +lean_mark_persistent(l_Classical_tacticBy__cases___x3a_____closed__25); +l_Classical_tacticBy__cases___x3a_____closed__26 = _init_l_Classical_tacticBy__cases___x3a_____closed__26(); +lean_mark_persistent(l_Classical_tacticBy__cases___x3a_____closed__26); +l_Classical_tacticBy__cases___x3a__ = _init_l_Classical_tacticBy__cases___x3a__(); +lean_mark_persistent(l_Classical_tacticBy__cases___x3a__); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__1 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__1(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__1); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__2 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__2(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__2); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__3 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__3(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__3); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__4 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__4(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__4); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__5 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__5(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__5); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__6 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__6(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__6); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__7 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__7(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__7); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__8 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__8(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__8); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__9 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__9(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__9); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__10 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__10(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__10); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__11 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__11(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__11); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__12 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__12(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__12); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__13 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__13(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__13); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__14 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__14(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__14); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__15 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__15(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__15); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__16 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__16(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__16); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__17 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__17(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__17); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__18 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__18(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__18); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__19 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__19(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__19); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__20 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__20(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__20); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__21 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__21(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__21); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__22 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__22(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__22); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__23 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__23(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__23); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__24 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__24(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__24); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__25 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__25(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__25); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__26 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__26(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__26); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__27 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__27(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__27); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__28 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__28(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__28); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__29 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__29(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__29); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__30 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__30(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__30); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__31 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__31(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__31); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__32 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__32(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__32); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__33 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__33(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__33); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__34 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__34(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__34); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__35 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__35(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__35); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__36 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__36(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__36); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__37 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__37(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__37); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__38 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__38(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__38); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__39 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__39(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__39); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__40 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__40(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__40); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__41 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__41(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__41); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__42 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__42(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__42); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__43 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__43(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__43); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__44 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__44(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__44); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__45 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__45(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__45); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__46 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__46(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__46); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__47 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__47(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__47); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__48 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__48(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__48); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__49 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__49(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__49); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__50 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__50(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__50); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__51 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__51(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__51); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__52 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__52(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__52); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__53 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__53(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__53); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__54 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__54(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__54); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__55 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__55(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__55); +l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__56 = _init_l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__56(); +lean_mark_persistent(l_Classical___aux__Init__Classical______macroRules__Classical__tacticBy__cases___x3a____1___closed__56); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Init/Control/Basic.c b/stage0/stdlib/Init/Control/Basic.c index 7fd0c19f2cc..763308e74db 100644 --- a/stage0/stdlib/Init/Control/Basic.c +++ b/stage0/stdlib/Init/Control/Basic.c @@ -20,7 +20,6 @@ lean_object* lean_mk_empty_array_with_capacity(lean_object*); static lean_object* l___aux__Init__Control__Basic______macroRules__term___x3c_x26_x3e____1___closed__14; LEAN_EXPORT lean_object* l___aux__Init__Control__Basic______macroRules__term___x3c_x26_x26_x3e____1(lean_object*, lean_object*, lean_object*); static lean_object* l_term___x3c_x26_x26_x3e_____closed__7; -extern lean_object* l_Lean_nullKind; lean_object* lean_name_mk_string(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_term___x3c_x26_x3e__; LEAN_EXPORT lean_object* l_term___x3c_x7c_x7c_x3e__; @@ -64,6 +63,7 @@ static lean_object* l_optional___rarg___closed__1; static lean_object* l_term___x3c_x26_x3e_____closed__5; LEAN_EXPORT lean_object* l_instMonadControlT___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_instMonadControlT__1(lean_object*); +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); lean_object* l_not___boxed(lean_object*); lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_andM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -97,7 +97,6 @@ lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_instMonadControlT___rarg___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_notM(lean_object*); static lean_object* l___aux__Init__Control__Basic______macroRules__term___x3c_x26_x3e____1___closed__18; -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_instMonadControlT___rarg___lambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_bool(lean_object*, lean_object*); static lean_object* l_term___x3c_x26_x26_x3e_____closed__2; @@ -656,82 +655,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) -{ -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___x3c_x26_x3e_____closed__5; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Control__Basic______unexpand__Functor__mapRev__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___x3c_x26_x3e_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return x_24; +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(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; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___x3c_x26_x3e_____closed__5; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Control__Basic______unexpand__Functor__mapRev__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___x3c_x26_x3e_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___x3c_x26_x3e_____closed__5; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Control__Basic______unexpand__Functor__mapRev__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___x3c_x26_x3e_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___x3c_x26_x3e_____closed__5; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Control__Basic______unexpand__Functor__mapRev__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___x3c_x26_x3e_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -1314,82 +1312,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) -{ -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___x3c_x7c_x7c_x3e_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Control__Basic______unexpand__Functor__mapRev__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___x3c_x7c_x7c_x3e_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return x_24; +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(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; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___x3c_x7c_x7c_x3e_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Control__Basic______unexpand__Functor__mapRev__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___x3c_x7c_x7c_x3e_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___x3c_x7c_x7c_x3e_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Control__Basic______unexpand__Functor__mapRev__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___x3c_x7c_x7c_x3e_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___x3c_x7c_x7c_x3e_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Control__Basic______unexpand__Functor__mapRev__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___x3c_x7c_x7c_x3e_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -1766,82 +1763,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) -{ -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___x3c_x26_x26_x3e_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Control__Basic______unexpand__Functor__mapRev__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___x3c_x26_x26_x3e_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return x_24; +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(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; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___x3c_x26_x26_x3e_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Control__Basic______unexpand__Functor__mapRev__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___x3c_x26_x26_x3e_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___x3c_x26_x26_x3e_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Control__Basic______unexpand__Functor__mapRev__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___x3c_x26_x26_x3e_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___x3c_x26_x26_x3e_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Control__Basic______unexpand__Functor__mapRev__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___x3c_x26_x26_x3e_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } diff --git a/stage0/stdlib/Init/Conv.c b/stage0/stdlib/Init/Conv.c index 291b88858a7..c5cd6503513 100644 --- a/stage0/stdlib/Init/Conv.c +++ b/stage0/stdlib/Init/Conv.c @@ -30,6 +30,7 @@ static lean_object* l_Lean_Parser_Tactic_Conv_nestedTactic___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_convSeq___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_convSeq; static lean_object* l_Lean_Parser_Tactic_Conv_conv_xb7_x2e_____closed__3; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convIntro______1___lambda__1(lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_zeta___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convRepeat____1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_change___closed__5; @@ -41,7 +42,6 @@ static lean_object* l_Lean_Parser_Tactic_Conv_unfold___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_simp___closed__7; static lean_object* l_Lean_Parser_Tactic_Conv_rewrite___closed__5; static lean_object* l_Lean_Parser_Tactic_Conv_simp___closed__24; -extern lean_object* l_Lean_nullKind; static lean_object* l_Lean_Parser_Tactic_Conv_ext___closed__6; static lean_object* l_Lean_Parser_Tactic_Conv_conv___closed__6; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convRw______1___closed__2; @@ -53,7 +53,6 @@ static lean_object* l_Lean_Parser_Tactic_Conv_arg___closed__6; static lean_object* l_Lean_Parser_Tactic_Conv_delta___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__11; static lean_object* l_Lean_Parser_Tactic_Conv_convConvSeq___closed__1; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_convIntro______; static lean_object* l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__1; lean_object* l_Array_append___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_convTrace__state___closed__1; @@ -78,6 +77,7 @@ static lean_object* l_Lean_Parser_Tactic_Conv_change___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_conv_quot___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_enterArg___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_convApply_____closed__4; +static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convIntro______1___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_simp___closed__10; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_nestedConv; static lean_object* l_Lean_Parser_Tactic_Conv_lhs___closed__4; @@ -89,6 +89,7 @@ static lean_object* l_Lean_Parser_Tactic_Conv_zeta___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convApply____1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_lhs___closed__1; +static lean_object* l_Lean_Parser_Tactic_Conv_convIntro_______closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convErw____1___closed__31; static lean_object* l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__7; @@ -104,6 +105,7 @@ static lean_object* l_Lean_Parser_Tactic_Conv_paren___closed__7; static lean_object* l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__5; static lean_object* l_Lean_Parser_Tactic_Conv_change___closed__6; lean_object* lean_array_push(lean_object*, lean_object*); +lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convDone__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_arg___closed__11; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_convRepeat__; @@ -128,6 +130,7 @@ static lean_object* l_Lean_Parser_Tactic_Conv_rewrite___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_convSeqBracketed___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_nestedTactic___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_congr___closed__2; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convIntro______1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_enterArg; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convErw____1___closed__19; static lean_object* l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__3; @@ -137,6 +140,7 @@ static lean_object* l_Lean_Parser_Tactic_Conv_convSeq___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_conv___closed__7; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convRepeat____1___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_convRepeat_____closed__6; +lean_object* lean_nat_add(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_rewrite___closed__7; static lean_object* l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convErw____1___closed__11; @@ -154,13 +158,13 @@ static lean_object* l_Lean_Parser_Tactic_Conv_arg___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convTrace__state__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_unfold___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__6; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convIntro________1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_reduce___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_nestedTactic___closed__5; static lean_object* l_Lean_Parser_Tactic_Conv_conv_quot___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_nestedTactic___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_conv_xb7_x2e_____closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_simp___closed__12; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_convIntro____; static lean_object* l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_conv_xb7_x2e_____closed__2; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convApply____1___closed__1; @@ -170,44 +174,48 @@ static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules_ static lean_object* l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__12; static lean_object* l_Lean_Parser_Tactic_Conv_lhs___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__4; -static lean_object* l_Lean_Parser_Tactic_Conv_convIntro_________closed__4; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_change___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_conv_quot___closed__2; +lean_object* lean_array_fget(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_change___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_convDone___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_rhs___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__2; +uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_conv_xb7_x2e_____closed__8; static lean_object* l_Lean_Parser_Tactic_Conv_simp___closed__26; static lean_object* l_Lean_Parser_Tactic_Conv_enterArg___closed__6; static lean_object* l_Lean_Parser_Tactic_Conv_conv_quot___closed__6; static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__1; +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_conv___closed__20; -static lean_object* l_Lean_Parser_Tactic_Conv_convIntro_________closed__6; static lean_object* l_Lean_Parser_Tactic_Conv_ext___closed__10; static lean_object* l_Lean_Parser_Tactic_Conv_unfold___closed__5; +static lean_object* l_Lean_Parser_Tactic_Conv_convIntro_______closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_convRight; static lean_object* l_Lean_Parser_Tactic_Conv_arg___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_pattern___closed__3; +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convIntro______1___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__9; static lean_object* l_Lean_Parser_Tactic_Conv_convConvSeq___closed__5; static lean_object* l_Lean_Parser_Tactic_Conv_unfold___closed__6; static lean_object* l_Lean_Parser_Tactic_Conv_convSkip___closed__2; +static lean_object* l_Lean_Parser_Tactic_Conv_convIntro_______closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_convConvSeq; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convErw____1___closed__17; static lean_object* l_Lean_Parser_Tactic_Conv_arg___closed__8; +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convIntro______1___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_Conv_convIntro_______closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_convSkip___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convErw____1___closed__24; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_zeta; static lean_object* l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__17; static lean_object* l_Lean_Parser_Tactic_Conv_conv_xb7_x2e_____closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_simp___closed__21; -static lean_object* l_Lean_Parser_Tactic_Conv_convIntro_________closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_convErw_____closed__6; static lean_object* l_Lean_Parser_Tactic_Conv_ext___closed__13; -static lean_object* l_Lean_Parser_Tactic_Conv_convIntro_________closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__13; static lean_object* l_Lean_Parser_Tactic_Conv_delta___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_conv___closed__22; @@ -229,12 +237,12 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroR static lean_object* l_Lean_Parser_Tactic_Conv_conv___closed__14; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_whnf; static lean_object* l_Lean_Parser_Tactic_Conv_convSeqBracketed___closed__2; +static lean_object* l_Lean_Parser_Tactic_Conv_convIntro_______closed__2; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convErw____1___closed__7; static lean_object* l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__21; static lean_object* l_Lean_Parser_Tactic_Conv_convConvSeq___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_conv___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_convSkip; -static lean_object* l_Lean_Parser_Tactic_Conv_convIntro_________closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_simp___closed__23; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_rewrite; extern lean_object* l_Lean_Parser_Tactic_rwRuleSeq; @@ -245,7 +253,6 @@ static lean_object* l_Lean_Parser_Tactic_Conv_conv___closed__16; static lean_object* l_Lean_Parser_Tactic_Conv_enterArg___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convRight__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__23; -static lean_object* l_Lean_Parser_Tactic_Conv_convIntro_________closed__5; static lean_object* l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__13; static lean_object* l_Lean_Parser_Tactic_Conv_convSeqBracketed___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_change___closed__4; @@ -264,6 +271,8 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_first; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_conv_quot; static lean_object* l_Lean_Parser_Tactic_Conv_simp___closed__8; static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__7; +LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convIntro______1___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_Conv_conv_quot___closed__22; static lean_object* l_Lean_Parser_Tactic_Conv_whnf___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_rhs___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_convSkip___closed__4; @@ -291,7 +300,6 @@ extern lean_object* l_Lean_Parser_Tactic_simpLemma; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_convLeft; static lean_object* l_Lean_Parser_Tactic_Conv_delta___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_convRepeat_____closed__4; -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_convErw_____closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_convSeqBracketed___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_simp___closed__1; @@ -347,6 +355,7 @@ static lean_object* l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__11; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convTrace__state__1___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_simp___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convSkip__1___closed__1; +static lean_object* l_Lean_Parser_Tactic_Conv_conv_quot___closed__21; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convErw____1___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_convLeft___closed__1; @@ -368,6 +377,7 @@ static lean_object* l_Lean_Parser_Tactic_Conv_simpMatch___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convErw____1___closed__18; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convErw____1___closed__30; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convErw____1___closed__5; +static lean_object* l_Lean_Parser_Tactic_Conv_convIntro_______closed__5; static lean_object* l_Lean_Parser_Tactic_Conv_convSkip___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_simpMatch___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_lhs; @@ -429,6 +439,7 @@ static lean_object* l_Lean_Parser_Tactic_Conv_whnf___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_delta___closed__6; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convErw____1___closed__9; static lean_object* l_Lean_Parser_Tactic_Conv_convApply_____closed__6; +LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convIntro______1___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convTrace__state__1___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_convRepeat_____closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__6; @@ -476,6 +487,7 @@ static lean_object* l_Lean_Parser_Tactic_Conv_conv___closed__9; static lean_object* l_Lean_Parser_Tactic_Conv_ext___closed__12; static lean_object* l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__9; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_convDone; +uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_conv_quot___closed__15; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convErw____1___closed__23; static lean_object* l_Lean_Parser_Tactic_Conv_arg___closed__13; @@ -556,7 +568,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv_quot___closed__9() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("andthen", 7); +x_1 = lean_mk_string_from_bytes("conv", 4); return x_1; } } @@ -573,44 +585,54 @@ return x_3; static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv_quot___closed__11() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("`(conv|", 7); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_2 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__7; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; } } static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv_quot___closed__12() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__11; -x_2 = lean_alloc_ctor(5, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv_quot___closed__13() { -_start: -{ lean_object* x_1; -x_1 = lean_mk_string_from_bytes("conv", 4); +x_1 = lean_mk_string_from_bytes("andthen", 7); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv_quot___closed__14() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv_quot___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; +x_2 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__12; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } +static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv_quot___closed__14() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("`(conv|", 7); +return x_1; +} +} static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv_quot___closed__15() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; +lean_object* x_1; lean_object* x_2; x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__14; +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv_quot___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -618,7 +640,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv_quot___closed__16() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv_quot___closed__17() { _start: { lean_object* x_1; @@ -626,23 +648,23 @@ x_1 = lean_mk_string_from_bytes(")", 1); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv_quot___closed__17() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv_quot___closed__18() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__16; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__17; x_2 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv_quot___closed__18() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv_quot___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; -x_2 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__15; -x_3 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__17; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; +x_2 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__16; +x_3 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__18; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -650,13 +672,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv_quot___closed__19() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv_quot___closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; -x_2 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__12; -x_3 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__18; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; +x_2 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__15; +x_3 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__19; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -664,13 +686,27 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv_quot___closed__20() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv_quot___closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__11; +x_2 = lean_unsigned_to_nat(1024u); +x_3 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__20; +x_4 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv_quot___closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__8; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__19; +x_3 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__21; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -682,7 +718,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv_quot() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__20; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__22; return x_1; } } @@ -826,9 +862,9 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__1 _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__15; -x_3 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__15; +x_3 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__16; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -888,7 +924,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__2 _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__16; x_3 = l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__21; x_4 = lean_alloc_ctor(2, 3, 0); @@ -996,8 +1032,8 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_convSeqBracketed___closed__5 _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; -x_2 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__15; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; +x_2 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__16; x_3 = l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__21; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -1034,7 +1070,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_convSeqBracketed___closed__8 _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_convSeqBracketed___closed__4; x_3 = l_Lean_Parser_Tactic_Conv_convSeqBracketed___closed__7; x_4 = lean_alloc_ctor(2, 3, 0); @@ -1066,7 +1102,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_convSeqBracketed___closed__1 _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_convSeqBracketed___closed__8; x_3 = l_Lean_Parser_Tactic_Conv_convSeqBracketed___closed__10; x_4 = lean_alloc_ctor(2, 3, 0); @@ -1175,7 +1211,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv___closed__1() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__4; -x_2 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; +x_2 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__9; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } @@ -1250,7 +1286,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_conv___closed__5; x_3 = l_Lean_Parser_Tactic_Conv_conv___closed__8; x_4 = lean_alloc_ctor(2, 3, 0); @@ -1276,7 +1312,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_conv___closed__3; x_3 = l_Lean_Parser_Tactic_Conv_conv___closed__10; x_4 = lean_alloc_ctor(2, 3, 0); @@ -1338,7 +1374,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_conv___closed__13; x_3 = l_Lean_Parser_Tactic_Conv_conv___closed__16; x_4 = lean_alloc_ctor(2, 3, 0); @@ -1364,7 +1400,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_conv___closed__11; x_3 = l_Lean_Parser_Tactic_Conv_conv___closed__18; x_4 = lean_alloc_ctor(2, 3, 0); @@ -1396,7 +1432,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv___closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_conv___closed__19; x_3 = l_Lean_Parser_Tactic_Conv_conv___closed__21; x_4 = lean_alloc_ctor(2, 3, 0); @@ -1410,7 +1446,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv___closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_conv___closed__22; x_3 = l_Lean_Parser_Tactic_Conv_convSeq; x_4 = lean_alloc_ctor(2, 3, 0); @@ -1826,7 +1862,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_arg___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_arg___closed__4; x_3 = l_Lean_Parser_Tactic_Conv_arg___closed__7; x_4 = lean_alloc_ctor(2, 3, 0); @@ -1868,7 +1904,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_arg___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_arg___closed__8; x_3 = l_Lean_Parser_Tactic_Conv_arg___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); @@ -1988,7 +2024,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_ext___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_ext___closed__9; x_3 = l_Lean_Parser_Tactic_Conv_conv___closed__8; x_4 = lean_alloc_ctor(2, 3, 0); @@ -2014,7 +2050,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_ext___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_ext___closed__4; x_3 = l_Lean_Parser_Tactic_Conv_ext___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); @@ -2088,7 +2124,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_change___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_change___closed__4; x_3 = l_Lean_Parser_Tactic_Conv_conv___closed__16; x_4 = lean_alloc_ctor(2, 3, 0); @@ -2162,7 +2198,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_delta___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_delta___closed__4; x_3 = l_Lean_Parser_Tactic_Conv_conv___closed__8; x_4 = lean_alloc_ctor(2, 3, 0); @@ -2236,7 +2272,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_unfold___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_unfold___closed__4; x_3 = l_Lean_Parser_Tactic_Conv_conv___closed__8; x_4 = lean_alloc_ctor(2, 3, 0); @@ -2310,7 +2346,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_pattern___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_pattern___closed__4; x_3 = l_Lean_Parser_Tactic_Conv_conv___closed__16; x_4 = lean_alloc_ctor(2, 3, 0); @@ -2396,7 +2432,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_rewrite___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_rewrite___closed__4; x_3 = l_Lean_Parser_Tactic_Conv_rewrite___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -2410,7 +2446,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_rewrite___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_rewrite___closed__6; x_3 = l_Lean_Parser_Tactic_rwRuleSeq; x_4 = lean_alloc_ctor(2, 3, 0); @@ -2484,7 +2520,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_simp___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_simp___closed__4; x_3 = l_Lean_Parser_Tactic_Conv_rewrite___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -2510,7 +2546,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_simp___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_simp___closed__5; x_3 = l_Lean_Parser_Tactic_Conv_simp___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); @@ -2556,7 +2592,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_simp___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_simp___closed__7; x_3 = l_Lean_Parser_Tactic_Conv_simp___closed__10; x_4 = lean_alloc_ctor(2, 3, 0); @@ -2658,7 +2694,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_simp___closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_simp___closed__13; x_3 = l_Lean_Parser_Tactic_Conv_simp___closed__19; x_4 = lean_alloc_ctor(2, 3, 0); @@ -2690,7 +2726,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_simp___closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_simp___closed__20; x_3 = l_Lean_Parser_Tactic_Conv_simp___closed__22; x_4 = lean_alloc_ctor(2, 3, 0); @@ -2716,7 +2752,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_simp___closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_simp___closed__11; x_3 = l_Lean_Parser_Tactic_Conv_simp___closed__24; x_4 = lean_alloc_ctor(2, 3, 0); @@ -2850,7 +2886,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_nestedTacticCore___closed__5 _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_nestedTacticCore___closed__4; x_3 = l_Lean_Parser_Tactic_Conv_conv___closed__21; x_4 = lean_alloc_ctor(2, 3, 0); @@ -2892,7 +2928,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_nestedTacticCore___closed__9 _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_nestedTacticCore___closed__5; x_3 = l_Lean_Parser_Tactic_Conv_nestedTacticCore___closed__8; x_4 = lean_alloc_ctor(2, 3, 0); @@ -2966,7 +3002,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_nestedTactic___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_nestedTactic___closed__4; x_3 = l_Lean_Parser_Tactic_Conv_conv___closed__21; x_4 = lean_alloc_ctor(2, 3, 0); @@ -2980,7 +3016,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_nestedTactic___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_nestedTactic___closed__5; x_3 = l_Lean_Parser_Tactic_Conv_nestedTacticCore___closed__8; x_4 = lean_alloc_ctor(2, 3, 0); @@ -3094,7 +3130,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_paren___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_paren___closed__4; x_3 = l_Lean_Parser_Tactic_Conv_convSeq; x_4 = lean_alloc_ctor(2, 3, 0); @@ -3108,9 +3144,9 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_paren___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_paren___closed__5; -x_3 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__17; +x_3 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__18; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3162,7 +3198,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_convConvSeq___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_conv___closed__3; x_3 = l_Lean_Parser_Tactic_Conv_conv___closed__21; x_4 = lean_alloc_ctor(2, 3, 0); @@ -3176,7 +3212,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_convConvSeq___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_convConvSeq___closed__3; x_3 = l_Lean_Parser_Tactic_Conv_convSeq; x_4 = lean_alloc_ctor(2, 3, 0); @@ -3280,7 +3316,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv_xb7_x2e_____closed__8() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_conv_xb7_x2e_____closed__7; x_3 = l_Lean_Parser_Tactic_Conv_convSeq; x_4 = lean_alloc_ctor(2, 3, 0); @@ -3419,7 +3455,7 @@ lean_inc(x_14); x_17 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_17, 0, x_14); lean_ctor_set(x_17, 1, x_16); -x_18 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__16; +x_18 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__17; lean_inc(x_14); x_19 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_19, 0, x_14); @@ -3523,7 +3559,7 @@ lean_inc(x_60); x_64 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_64, 0, x_60); lean_ctor_set(x_64, 1, x_63); -x_65 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__16; +x_65 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__17; lean_inc(x_60); x_66 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_66, 0, x_60); @@ -3661,7 +3697,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_convRw_______closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_convRw_______closed__4; x_3 = l_Lean_Parser_Tactic_Conv_rewrite___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -3675,7 +3711,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_convRw_______closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_convRw_______closed__5; x_3 = l_Lean_Parser_Tactic_rwRuleSeq; x_4 = lean_alloc_ctor(2, 3, 0); @@ -3750,44 +3786,94 @@ return x_7; } 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; uint8_t x_14; +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; lean_object* x_40; lean_object* x_41; x_8 = lean_unsigned_to_nat(1u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); x_10 = lean_unsigned_to_nat(2u); x_11 = l_Lean_Syntax_getArg(x_1, x_10); lean_dec(x_1); -x_12 = l_Lean_Syntax_getOptional_x3f(x_9); +x_40 = l_Lean_Syntax_getOptional_x3f(x_9); lean_dec(x_9); -x_13 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) +x_41 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +if (lean_obj_tag(x_40) == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_15 = lean_ctor_get(x_13, 0); -x_16 = l_Lean_Parser_Tactic_Conv_rewrite___closed__1; -x_17 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_17, 0, x_15); -lean_ctor_set(x_17, 1, x_16); -x_18 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__3; -x_19 = lean_array_push(x_18, x_17); +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_41, 1); +lean_inc(x_43); +lean_dec(x_41); +x_44 = lean_box(0); +x_12 = x_44; +x_13 = x_42; +x_14 = x_43; +goto block_39; +} +else +{ +uint8_t x_45; +x_45 = !lean_is_exclusive(x_40); +if (x_45 == 0) +{ +lean_object* x_46; lean_object* x_47; +x_46 = lean_ctor_get(x_41, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_41, 1); +lean_inc(x_47); +lean_dec(x_41); +x_12 = x_40; +x_13 = x_46; +x_14 = x_47; +goto block_39; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_48 = lean_ctor_get(x_40, 0); +lean_inc(x_48); +lean_dec(x_40); +x_49 = lean_ctor_get(x_41, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_41, 1); +lean_inc(x_50); +lean_dec(x_41); +x_51 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_51, 0, x_48); +x_12 = x_51; +x_13 = x_49; +x_14 = x_50; +goto block_39; +} +} +block_39: +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_15 = l_Lean_Parser_Tactic_Conv_rewrite___closed__1; +x_16 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_16, 0, x_13); +lean_ctor_set(x_16, 1, x_15); +x_17 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__3; +x_18 = lean_array_push(x_17, x_16); if (lean_obj_tag(x_12) == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_20 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convRw______1___closed__2; -x_21 = lean_array_push(x_19, x_20); -x_22 = lean_array_push(x_21, x_11); -x_23 = lean_box(2); -x_24 = l_Lean_Parser_Tactic_Conv_rewrite___closed__2; -x_25 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_25, 0, x_23); -lean_ctor_set(x_25, 1, x_24); -lean_ctor_set(x_25, 2, x_22); -lean_ctor_set(x_13, 0, x_25); -return x_13; +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; +x_19 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convRw______1___closed__2; +x_20 = lean_array_push(x_18, x_19); +x_21 = lean_array_push(x_20, x_11); +x_22 = lean_box(2); +x_23 = l_Lean_Parser_Tactic_Conv_rewrite___closed__2; +x_24 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +lean_ctor_set(x_24, 2, x_21); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_14); +return x_25; } else { -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_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_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; 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; x_26 = lean_ctor_get(x_12, 0); lean_inc(x_26); lean_dec(x_12); @@ -3801,75 +3887,17 @@ x_33 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_33, 0, x_31); lean_ctor_set(x_33, 1, x_32); lean_ctor_set(x_33, 2, x_30); -x_34 = lean_array_push(x_19, x_33); +x_34 = lean_array_push(x_18, x_33); x_35 = lean_array_push(x_34, x_11); x_36 = l_Lean_Parser_Tactic_Conv_rewrite___closed__2; x_37 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_37, 0, x_31); lean_ctor_set(x_37, 1, x_36); lean_ctor_set(x_37, 2, x_35); -lean_ctor_set(x_13, 0, x_37); -return x_13; -} -} -else -{ -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_38 = lean_ctor_get(x_13, 0); -x_39 = lean_ctor_get(x_13, 1); -lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_13); -x_40 = l_Lean_Parser_Tactic_Conv_rewrite___closed__1; -x_41 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_41, 0, x_38); -lean_ctor_set(x_41, 1, x_40); -x_42 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__3; -x_43 = lean_array_push(x_42, x_41); -if (lean_obj_tag(x_12) == 0) -{ -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_44 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convRw______1___closed__2; -x_45 = lean_array_push(x_43, x_44); -x_46 = lean_array_push(x_45, x_11); -x_47 = lean_box(2); -x_48 = l_Lean_Parser_Tactic_Conv_rewrite___closed__2; -x_49 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_49, 0, x_47); -lean_ctor_set(x_49, 1, x_48); -lean_ctor_set(x_49, 2, x_46); -x_50 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_50, 0, x_49); -lean_ctor_set(x_50, 1, x_39); -return x_50; -} -else -{ -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_51 = lean_ctor_get(x_12, 0); -lean_inc(x_51); -lean_dec(x_12); -x_52 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__7; -x_53 = lean_array_push(x_52, x_51); -x_54 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__4; -x_55 = l_Array_append___rarg(x_54, x_53); -x_56 = lean_box(2); -x_57 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__2; -x_58 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_58, 0, x_56); -lean_ctor_set(x_58, 1, x_57); -lean_ctor_set(x_58, 2, x_55); -x_59 = lean_array_push(x_43, x_58); -x_60 = lean_array_push(x_59, x_11); -x_61 = l_Lean_Parser_Tactic_Conv_rewrite___closed__2; -x_62 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_62, 0, x_56); -lean_ctor_set(x_62, 1, x_61); -lean_ctor_set(x_62, 2, x_60); -x_63 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_63, 0, x_62); -lean_ctor_set(x_63, 1, x_39); -return x_63; +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_14); +return x_38; } } } @@ -3917,7 +3945,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_convErw_____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_convErw_____closed__4; x_3 = l_Lean_Parser_Tactic_rwRuleSeq; x_4 = lean_alloc_ctor(2, 3, 0); @@ -4405,7 +4433,7 @@ x_63 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_63, 0, x_34); lean_ctor_set(x_63, 1, x_62); lean_ctor_set(x_63, 2, x_61); -x_64 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__16; +x_64 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__17; x_65 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_65, 0, x_12); lean_ctor_set(x_65, 1, x_64); @@ -4541,7 +4569,7 @@ x_133 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_133, 0, x_104); lean_ctor_set(x_133, 1, x_132); lean_ctor_set(x_133, 2, x_131); -x_134 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__16; +x_134 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__17; x_135 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_135, 0, x_81); lean_ctor_set(x_135, 1, x_134); @@ -4961,25 +4989,25 @@ return x_27; } } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_convIntro_________closed__1() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_convIntro_______closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("convIntro___", 12); +x_1 = lean_mk_string_from_bytes("convIntro__", 11); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_convIntro_________closed__2() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_convIntro_______closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__4; -x_2 = l_Lean_Parser_Tactic_Conv_convIntro_________closed__1; +x_2 = l_Lean_Parser_Tactic_Conv_convIntro_______closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_convIntro_________closed__3() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_convIntro_______closed__3() { _start: { lean_object* x_1; @@ -4987,11 +5015,11 @@ x_1 = lean_mk_string_from_bytes("intro ", 6); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_convIntro_________closed__4() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_convIntro_______closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_Conv_convIntro_________closed__3; +x_1 = l_Lean_Parser_Tactic_Conv_convIntro_______closed__3; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -4999,12 +5027,12 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_convIntro_________closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_convIntro_______closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; -x_2 = l_Lean_Parser_Tactic_Conv_convIntro_________closed__4; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; +x_2 = l_Lean_Parser_Tactic_Conv_convIntro_______closed__4; x_3 = l_Lean_Parser_Tactic_Conv_ext___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -5013,13 +5041,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_convIntro_________closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_convIntro_______closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_convIntro_________closed__2; +x_1 = l_Lean_Parser_Tactic_Conv_convIntro_______closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Parser_Tactic_Conv_convIntro_________closed__5; +x_3 = l_Lean_Parser_Tactic_Conv_convIntro_______closed__5; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -5027,19 +5055,130 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_convIntro______() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_convIntro____() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Parser_Tactic_Conv_convIntro_______closed__6; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convIntro______1___spec__2(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; uint8_t x_7; +x_6 = lean_array_get_size(x_1); +x_7 = lean_nat_dec_lt(x_4, x_6); +lean_dec(x_6); +if (x_7 == 0) +{ +lean_object* x_8; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_5); +return x_8; +} +else +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_unsigned_to_nat(0u); +x_10 = lean_nat_dec_eq(x_3, x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_sub(x_3, x_11); +lean_dec(x_3); +x_13 = lean_array_fget(x_1, x_4); +lean_inc(x_2); +x_14 = lean_apply_1(x_2, x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; +lean_dec(x_12); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_15 = lean_box(0); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_14, 0); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_nat_add(x_4, x_11); +lean_dec(x_4); +x_18 = lean_array_push(x_5, x_16); +x_3 = x_12; +x_4 = x_17; +x_5 = x_18; +goto _start; +} +} +else +{ +lean_object* x_20; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_20 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_20, 0, x_5); +return x_20; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convIntro______1___spec__1(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; +x_3 = lean_array_get_size(x_1); +x_4 = lean_mk_empty_array_with_capacity(x_3); +x_5 = lean_unsigned_to_nat(0u); +x_6 = l_Array_sequenceMap_loop___at_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convIntro______1___spec__2(x_1, x_2, x_3, x_5, x_4); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convIntro______1___lambda__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; +x_2 = l_Lean_Parser_Tactic_Conv_conv___closed__7; +lean_inc(x_1); +x_3 = l_Lean_Syntax_isOfKind(x_1, x_2); +if (x_3 == 0) +{ +lean_object* x_4; +lean_dec(x_1); +x_4 = lean_box(0); +return x_4; +} +else +{ +lean_object* x_5; +x_5 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_5, 0, x_1); +return x_5; +} +} +} +static lean_object* _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convIntro______1___closed__1() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_Conv_convIntro_________closed__6; +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convIntro______1___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convIntro________1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convIntro______1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; uint8_t x_5; -x_4 = l_Lean_Parser_Tactic_Conv_convIntro_________closed__2; +x_4 = l_Lean_Parser_Tactic_Conv_convIntro_______closed__2; lean_inc(x_1); x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); if (x_5 == 0) @@ -5055,77 +5194,115 @@ return x_7; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; x_8 = lean_unsigned_to_nat(1u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); lean_dec(x_1); x_10 = l_Lean_Syntax_getArgs(x_9); lean_dec(x_9); -x_11 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) -{ -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; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_13 = lean_ctor_get(x_11, 0); -x_14 = l_Lean_Parser_Tactic_Conv_ext___closed__1; -x_15 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_15, 0, x_13); -lean_ctor_set(x_15, 1, x_14); -x_16 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__4; -x_17 = l_Array_append___rarg(x_16, x_10); -x_18 = lean_box(2); -x_19 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__2; -x_20 = lean_alloc_ctor(1, 3, 0); +x_11 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convIntro______1___closed__1; +x_12 = l_Array_sequenceMap___at_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convIntro______1___spec__1(x_10, x_11); +lean_dec(x_10); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_2); +x_13 = lean_box(1); +x_14 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_3); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = lean_ctor_get(x_12, 0); +lean_inc(x_15); +lean_dec(x_12); +x_16 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +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; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_18 = lean_ctor_get(x_16, 0); +x_19 = l_Lean_Parser_Tactic_Conv_ext___closed__1; +x_20 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_20, 0, x_18); lean_ctor_set(x_20, 1, x_19); -lean_ctor_set(x_20, 2, x_17); -x_21 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__6; -x_22 = lean_array_push(x_21, x_15); -x_23 = lean_array_push(x_22, x_20); -x_24 = l_Lean_Parser_Tactic_Conv_ext___closed__2; +x_21 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__4; +x_22 = l_Array_append___rarg(x_21, x_15); +x_23 = lean_box(2); +x_24 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__2; x_25 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_25, 0, x_18); +lean_ctor_set(x_25, 0, x_23); lean_ctor_set(x_25, 1, x_24); -lean_ctor_set(x_25, 2, x_23); -lean_ctor_set(x_11, 0, x_25); -return x_11; +lean_ctor_set(x_25, 2, x_22); +x_26 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__6; +x_27 = lean_array_push(x_26, x_20); +x_28 = lean_array_push(x_27, x_25); +x_29 = l_Lean_Parser_Tactic_Conv_ext___closed__2; +x_30 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_30, 0, x_23); +lean_ctor_set(x_30, 1, x_29); +lean_ctor_set(x_30, 2, x_28); +lean_ctor_set(x_16, 0, x_30); +return x_16; } else { -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_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; -x_26 = lean_ctor_get(x_11, 0); -x_27 = lean_ctor_get(x_11, 1); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_11); -x_28 = l_Lean_Parser_Tactic_Conv_ext___closed__1; -x_29 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_29, 0, x_26); -lean_ctor_set(x_29, 1, x_28); -x_30 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__4; -x_31 = l_Array_append___rarg(x_30, x_10); -x_32 = lean_box(2); -x_33 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__2; -x_34 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_34, 0, x_32); +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; lean_object* x_44; lean_object* x_45; +x_31 = lean_ctor_get(x_16, 0); +x_32 = lean_ctor_get(x_16, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_16); +x_33 = l_Lean_Parser_Tactic_Conv_ext___closed__1; +x_34 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_34, 0, x_31); lean_ctor_set(x_34, 1, x_33); -lean_ctor_set(x_34, 2, x_31); -x_35 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__6; -x_36 = lean_array_push(x_35, x_29); -x_37 = lean_array_push(x_36, x_34); -x_38 = l_Lean_Parser_Tactic_Conv_ext___closed__2; +x_35 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__4; +x_36 = l_Array_append___rarg(x_35, x_15); +x_37 = lean_box(2); +x_38 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__2; x_39 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_39, 0, x_32); +lean_ctor_set(x_39, 0, x_37); lean_ctor_set(x_39, 1, x_38); -lean_ctor_set(x_39, 2, x_37); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_27); -return x_40; +lean_ctor_set(x_39, 2, x_36); +x_40 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__6; +x_41 = lean_array_push(x_40, x_34); +x_42 = lean_array_push(x_41, x_39); +x_43 = l_Lean_Parser_Tactic_Conv_ext___closed__2; +x_44 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_44, 0, x_37); +lean_ctor_set(x_44, 1, x_43); +lean_ctor_set(x_44, 2, x_42); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_32); +return x_45; } } } } +} +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convIntro______1___spec__2___boxed(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; +x_6 = l_Array_sequenceMap_loop___at_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convIntro______1___spec__2(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convIntro______1___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Array_sequenceMap___at_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convIntro______1___spec__1(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} static lean_object* _init_l_Lean_Parser_Tactic_Conv_enterArg___closed__1() { _start: { @@ -5148,7 +5325,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_enterArg___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_arg___closed__7; x_3 = l_Lean_Parser_Tactic_Conv_arg___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); @@ -5248,7 +5425,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__4; x_3 = l_Lean_Parser_Tactic_Conv_simp___closed__13; x_4 = lean_alloc_ctor(2, 3, 0); @@ -5262,7 +5439,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_ext___closed__9; x_3 = l_Lean_Parser_Tactic_Conv_enterArg; x_4 = lean_alloc_ctor(2, 3, 0); @@ -5292,7 +5469,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__5; x_3 = l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__7; x_4 = lean_alloc_ctor(2, 3, 0); @@ -5306,7 +5483,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__8; x_3 = l_Lean_Parser_Tactic_Conv_simp___closed__22; x_4 = lean_alloc_ctor(2, 3, 0); @@ -5375,72 +5552,72 @@ return x_7; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; x_8 = lean_unsigned_to_nat(2u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); lean_dec(x_1); -x_10 = l_Lean_nullKind; -x_11 = lean_unsigned_to_nat(1u); +x_10 = lean_unsigned_to_nat(1u); lean_inc(x_9); -x_12 = l_Lean_Syntax_isNodeOf(x_9, x_10, x_11); -if (x_12 == 0) +x_11 = l_Lean_Syntax_matchesNull(x_9, x_10); +if (x_11 == 0) { -lean_object* x_13; uint8_t x_14; -x_13 = l_Lean_Syntax_getNumArgs(x_9); -x_14 = lean_nat_dec_le(x_8, x_13); -if (x_14 == 0) +lean_object* x_12; uint8_t x_13; +x_12 = l_Lean_Syntax_getNumArgs(x_9); +x_13 = lean_nat_dec_le(x_8, x_12); +if (x_13 == 0) { -lean_object* x_15; lean_object* x_16; -lean_dec(x_13); +lean_object* x_14; lean_object* x_15; +lean_dec(x_12); lean_dec(x_9); lean_dec(x_2); -x_15 = lean_box(1); -x_16 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_3); -return x_16; +x_14 = lean_box(1); +x_15 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_3); +return x_15; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_17 = lean_unsigned_to_nat(0u); -x_18 = l_Lean_Syntax_getArg(x_9, x_17); -x_19 = l_Lean_Parser_Tactic_Conv_enterArg___closed__2; -lean_inc(x_18); -x_20 = l_Lean_Syntax_isOfKind(x_18, x_19); -if (x_20 == 0) +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_16 = lean_unsigned_to_nat(0u); +x_17 = l_Lean_Syntax_getArg(x_9, x_16); +x_18 = l_Lean_Parser_Tactic_Conv_enterArg___closed__2; +lean_inc(x_17); +x_19 = l_Lean_Syntax_isOfKind(x_17, x_18); +if (x_19 == 0) { -lean_object* x_21; lean_object* x_22; -lean_dec(x_18); -lean_dec(x_13); +lean_object* x_20; lean_object* x_21; +lean_dec(x_17); +lean_dec(x_12); lean_dec(x_9); lean_dec(x_2); -x_21 = lean_box(1); -x_22 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_3); -return x_22; +x_20 = lean_box(1); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_3); +return x_21; } else { -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; uint8_t x_30; -x_23 = l_Lean_Syntax_getArgs(x_9); +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; uint8_t x_30; +x_22 = l_Lean_Syntax_getArgs(x_9); lean_dec(x_9); -x_24 = lean_nat_sub(x_13, x_17); -lean_dec(x_13); -x_25 = l_Array_extract___rarg(x_23, x_8, x_24); -x_26 = lean_box(2); +x_23 = lean_nat_sub(x_12, x_16); +lean_dec(x_12); +x_24 = l_Array_extract___rarg(x_22, x_8, x_23); +x_25 = lean_box(2); +x_26 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__2; x_27 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_10); -lean_ctor_set(x_27, 2, x_25); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +lean_ctor_set(x_27, 2, x_24); x_28 = l_Lean_Syntax_getArgs(x_27); lean_dec(x_27); x_29 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); x_30 = !lean_is_exclusive(x_29); if (x_30 == 0) { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; x_31 = lean_ctor_get(x_29, 0); x_32 = l_Lean_Parser_Tactic_Conv_paren___closed__3; lean_inc(x_31); @@ -5458,527 +5635,525 @@ x_37 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_37, 0, x_31); lean_ctor_set(x_37, 1, x_36); x_38 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__7; -x_39 = lean_array_push(x_38, x_18); -x_40 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__2; -x_41 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_41, 0, x_26); -lean_ctor_set(x_41, 1, x_40); -lean_ctor_set(x_41, 2, x_39); -x_42 = l_Lean_Parser_Tactic_Conv_simp___closed__21; +x_39 = lean_array_push(x_38, x_17); +x_40 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_40, 0, x_25); +lean_ctor_set(x_40, 1, x_26); +lean_ctor_set(x_40, 2, x_39); +x_41 = l_Lean_Parser_Tactic_Conv_simp___closed__21; lean_inc(x_31); -x_43 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_43, 0, x_31); -lean_ctor_set(x_43, 1, x_42); -x_44 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b_____x5d__1___closed__2; -x_45 = lean_array_push(x_44, x_35); -x_46 = lean_array_push(x_45, x_37); -lean_inc(x_46); -x_47 = lean_array_push(x_46, x_41); -lean_inc(x_43); -x_48 = lean_array_push(x_47, x_43); -x_49 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_49, 0, x_26); -lean_ctor_set(x_49, 1, x_4); -lean_ctor_set(x_49, 2, x_48); -x_50 = l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__19; +x_42 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_42, 0, x_31); +lean_ctor_set(x_42, 1, x_41); +x_43 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b_____x5d__1___closed__2; +x_44 = lean_array_push(x_43, x_35); +x_45 = lean_array_push(x_44, x_37); +lean_inc(x_45); +x_46 = lean_array_push(x_45, x_40); +lean_inc(x_42); +x_47 = lean_array_push(x_46, x_42); +x_48 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_48, 0, x_25); +lean_ctor_set(x_48, 1, x_4); +lean_ctor_set(x_48, 2, x_47); +x_49 = l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__19; lean_inc(x_31); -x_51 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_51, 0, x_31); -lean_ctor_set(x_51, 1, x_50); -x_52 = lean_array_push(x_38, x_51); -x_53 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_53, 0, x_26); -lean_ctor_set(x_53, 1, x_40); -lean_ctor_set(x_53, 2, x_52); -x_54 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__6; -x_55 = lean_array_push(x_54, x_49); -x_56 = lean_array_push(x_55, x_53); -x_57 = l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__12; -x_58 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_58, 0, x_26); -lean_ctor_set(x_58, 1, x_57); -lean_ctor_set(x_58, 2, x_56); -x_59 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__4; -x_60 = l_Array_append___rarg(x_59, x_28); -x_61 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_61, 0, x_26); -lean_ctor_set(x_61, 1, x_40); -lean_ctor_set(x_61, 2, x_60); -x_62 = lean_array_push(x_46, x_61); -x_63 = lean_array_push(x_62, x_43); -x_64 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_64, 0, x_26); -lean_ctor_set(x_64, 1, x_4); -lean_ctor_set(x_64, 2, x_63); -x_65 = lean_array_push(x_54, x_64); -x_66 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__5; -x_67 = lean_array_push(x_65, x_66); -x_68 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_68, 0, x_26); -lean_ctor_set(x_68, 1, x_57); -lean_ctor_set(x_68, 2, x_67); -x_69 = lean_array_push(x_54, x_58); -x_70 = lean_array_push(x_69, x_68); -x_71 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_71, 0, x_26); -lean_ctor_set(x_71, 1, x_40); -lean_ctor_set(x_71, 2, x_70); -x_72 = lean_array_push(x_38, x_71); -x_73 = l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__6; -x_74 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_74, 0, x_26); -lean_ctor_set(x_74, 1, x_73); -lean_ctor_set(x_74, 2, x_72); -x_75 = lean_array_push(x_38, x_74); -x_76 = l_Lean_Parser_Tactic_Conv_convSeq___closed__2; -x_77 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_77, 0, x_26); -lean_ctor_set(x_77, 1, x_76); -lean_ctor_set(x_77, 2, x_75); -x_78 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__16; -x_79 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_79, 0, x_31); -lean_ctor_set(x_79, 1, x_78); -x_80 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__3; -x_81 = lean_array_push(x_80, x_33); -x_82 = lean_array_push(x_81, x_77); -x_83 = lean_array_push(x_82, x_79); -x_84 = l_Lean_Parser_Tactic_Conv_paren___closed__2; -x_85 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_85, 0, x_26); -lean_ctor_set(x_85, 1, x_84); -lean_ctor_set(x_85, 2, x_83); -lean_ctor_set(x_29, 0, x_85); +x_50 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_50, 0, x_31); +lean_ctor_set(x_50, 1, x_49); +x_51 = lean_array_push(x_38, x_50); +x_52 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_52, 0, x_25); +lean_ctor_set(x_52, 1, x_26); +lean_ctor_set(x_52, 2, x_51); +x_53 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__6; +x_54 = lean_array_push(x_53, x_48); +x_55 = lean_array_push(x_54, x_52); +x_56 = l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__12; +x_57 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_57, 0, x_25); +lean_ctor_set(x_57, 1, x_56); +lean_ctor_set(x_57, 2, x_55); +x_58 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__4; +x_59 = l_Array_append___rarg(x_58, x_28); +x_60 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_60, 0, x_25); +lean_ctor_set(x_60, 1, x_26); +lean_ctor_set(x_60, 2, x_59); +x_61 = lean_array_push(x_45, x_60); +x_62 = lean_array_push(x_61, x_42); +x_63 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_63, 0, x_25); +lean_ctor_set(x_63, 1, x_4); +lean_ctor_set(x_63, 2, x_62); +x_64 = lean_array_push(x_53, x_63); +x_65 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__5; +x_66 = lean_array_push(x_64, x_65); +x_67 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_67, 0, x_25); +lean_ctor_set(x_67, 1, x_56); +lean_ctor_set(x_67, 2, x_66); +x_68 = lean_array_push(x_53, x_57); +x_69 = lean_array_push(x_68, x_67); +x_70 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_70, 0, x_25); +lean_ctor_set(x_70, 1, x_26); +lean_ctor_set(x_70, 2, x_69); +x_71 = lean_array_push(x_38, x_70); +x_72 = l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__6; +x_73 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_73, 0, x_25); +lean_ctor_set(x_73, 1, x_72); +lean_ctor_set(x_73, 2, x_71); +x_74 = lean_array_push(x_38, x_73); +x_75 = l_Lean_Parser_Tactic_Conv_convSeq___closed__2; +x_76 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_76, 0, x_25); +lean_ctor_set(x_76, 1, x_75); +lean_ctor_set(x_76, 2, x_74); +x_77 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__17; +x_78 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_78, 0, x_31); +lean_ctor_set(x_78, 1, x_77); +x_79 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__3; +x_80 = lean_array_push(x_79, x_33); +x_81 = lean_array_push(x_80, x_76); +x_82 = lean_array_push(x_81, x_78); +x_83 = l_Lean_Parser_Tactic_Conv_paren___closed__2; +x_84 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_84, 0, x_25); +lean_ctor_set(x_84, 1, x_83); +lean_ctor_set(x_84, 2, x_82); +lean_ctor_set(x_29, 0, x_84); return x_29; } else { -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; -x_86 = lean_ctor_get(x_29, 0); -x_87 = lean_ctor_get(x_29, 1); -lean_inc(x_87); +lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; +x_85 = lean_ctor_get(x_29, 0); +x_86 = lean_ctor_get(x_29, 1); lean_inc(x_86); +lean_inc(x_85); lean_dec(x_29); -x_88 = l_Lean_Parser_Tactic_Conv_paren___closed__3; -lean_inc(x_86); -x_89 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_89, 0, x_86); -lean_ctor_set(x_89, 1, x_88); -x_90 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b_____x5d__1___closed__1; -lean_inc(x_86); -x_91 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_91, 0, x_86); -lean_ctor_set(x_91, 1, x_90); -x_92 = l_Lean_Parser_Tactic_Conv_simp___closed__12; -lean_inc(x_86); -x_93 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_93, 0, x_86); -lean_ctor_set(x_93, 1, x_92); -x_94 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__7; -x_95 = lean_array_push(x_94, x_18); -x_96 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__2; -x_97 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_97, 0, x_26); +x_87 = l_Lean_Parser_Tactic_Conv_paren___closed__3; +lean_inc(x_85); +x_88 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_88, 0, x_85); +lean_ctor_set(x_88, 1, x_87); +x_89 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b_____x5d__1___closed__1; +lean_inc(x_85); +x_90 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_90, 0, x_85); +lean_ctor_set(x_90, 1, x_89); +x_91 = l_Lean_Parser_Tactic_Conv_simp___closed__12; +lean_inc(x_85); +x_92 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_92, 0, x_85); +lean_ctor_set(x_92, 1, x_91); +x_93 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__7; +x_94 = lean_array_push(x_93, x_17); +x_95 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_95, 0, x_25); +lean_ctor_set(x_95, 1, x_26); +lean_ctor_set(x_95, 2, x_94); +x_96 = l_Lean_Parser_Tactic_Conv_simp___closed__21; +lean_inc(x_85); +x_97 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_97, 0, x_85); lean_ctor_set(x_97, 1, x_96); -lean_ctor_set(x_97, 2, x_95); -x_98 = l_Lean_Parser_Tactic_Conv_simp___closed__21; -lean_inc(x_86); -x_99 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_99, 0, x_86); -lean_ctor_set(x_99, 1, x_98); -x_100 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b_____x5d__1___closed__2; -x_101 = lean_array_push(x_100, x_91); -x_102 = lean_array_push(x_101, x_93); -lean_inc(x_102); -x_103 = lean_array_push(x_102, x_97); -lean_inc(x_99); -x_104 = lean_array_push(x_103, x_99); -x_105 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_105, 0, x_26); -lean_ctor_set(x_105, 1, x_4); -lean_ctor_set(x_105, 2, x_104); -x_106 = l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__19; -lean_inc(x_86); -x_107 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_107, 0, x_86); -lean_ctor_set(x_107, 1, x_106); -x_108 = lean_array_push(x_94, x_107); -x_109 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_109, 0, x_26); -lean_ctor_set(x_109, 1, x_96); -lean_ctor_set(x_109, 2, x_108); -x_110 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__6; -x_111 = lean_array_push(x_110, x_105); -x_112 = lean_array_push(x_111, x_109); -x_113 = l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__12; -x_114 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_114, 0, x_26); -lean_ctor_set(x_114, 1, x_113); -lean_ctor_set(x_114, 2, x_112); -x_115 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__4; -x_116 = l_Array_append___rarg(x_115, x_28); -x_117 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_117, 0, x_26); -lean_ctor_set(x_117, 1, x_96); -lean_ctor_set(x_117, 2, x_116); -x_118 = lean_array_push(x_102, x_117); -x_119 = lean_array_push(x_118, x_99); -x_120 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_120, 0, x_26); -lean_ctor_set(x_120, 1, x_4); -lean_ctor_set(x_120, 2, x_119); -x_121 = lean_array_push(x_110, x_120); -x_122 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__5; -x_123 = lean_array_push(x_121, x_122); -x_124 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_124, 0, x_26); -lean_ctor_set(x_124, 1, x_113); -lean_ctor_set(x_124, 2, x_123); -x_125 = lean_array_push(x_110, x_114); -x_126 = lean_array_push(x_125, x_124); -x_127 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_127, 0, x_26); -lean_ctor_set(x_127, 1, x_96); -lean_ctor_set(x_127, 2, x_126); -x_128 = lean_array_push(x_94, x_127); -x_129 = l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__6; -x_130 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_130, 0, x_26); -lean_ctor_set(x_130, 1, x_129); -lean_ctor_set(x_130, 2, x_128); -x_131 = lean_array_push(x_94, x_130); -x_132 = l_Lean_Parser_Tactic_Conv_convSeq___closed__2; -x_133 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_133, 0, x_26); +x_98 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b_____x5d__1___closed__2; +x_99 = lean_array_push(x_98, x_90); +x_100 = lean_array_push(x_99, x_92); +lean_inc(x_100); +x_101 = lean_array_push(x_100, x_95); +lean_inc(x_97); +x_102 = lean_array_push(x_101, x_97); +x_103 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_103, 0, x_25); +lean_ctor_set(x_103, 1, x_4); +lean_ctor_set(x_103, 2, x_102); +x_104 = l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__19; +lean_inc(x_85); +x_105 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_105, 0, x_85); +lean_ctor_set(x_105, 1, x_104); +x_106 = lean_array_push(x_93, x_105); +x_107 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_107, 0, x_25); +lean_ctor_set(x_107, 1, x_26); +lean_ctor_set(x_107, 2, x_106); +x_108 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__6; +x_109 = lean_array_push(x_108, x_103); +x_110 = lean_array_push(x_109, x_107); +x_111 = l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__12; +x_112 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_112, 0, x_25); +lean_ctor_set(x_112, 1, x_111); +lean_ctor_set(x_112, 2, x_110); +x_113 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__4; +x_114 = l_Array_append___rarg(x_113, x_28); +x_115 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_115, 0, x_25); +lean_ctor_set(x_115, 1, x_26); +lean_ctor_set(x_115, 2, x_114); +x_116 = lean_array_push(x_100, x_115); +x_117 = lean_array_push(x_116, x_97); +x_118 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_118, 0, x_25); +lean_ctor_set(x_118, 1, x_4); +lean_ctor_set(x_118, 2, x_117); +x_119 = lean_array_push(x_108, x_118); +x_120 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__5; +x_121 = lean_array_push(x_119, x_120); +x_122 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_122, 0, x_25); +lean_ctor_set(x_122, 1, x_111); +lean_ctor_set(x_122, 2, x_121); +x_123 = lean_array_push(x_108, x_112); +x_124 = lean_array_push(x_123, x_122); +x_125 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_125, 0, x_25); +lean_ctor_set(x_125, 1, x_26); +lean_ctor_set(x_125, 2, x_124); +x_126 = lean_array_push(x_93, x_125); +x_127 = l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__6; +x_128 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_128, 0, x_25); +lean_ctor_set(x_128, 1, x_127); +lean_ctor_set(x_128, 2, x_126); +x_129 = lean_array_push(x_93, x_128); +x_130 = l_Lean_Parser_Tactic_Conv_convSeq___closed__2; +x_131 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_131, 0, x_25); +lean_ctor_set(x_131, 1, x_130); +lean_ctor_set(x_131, 2, x_129); +x_132 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__17; +x_133 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_133, 0, x_85); lean_ctor_set(x_133, 1, x_132); -lean_ctor_set(x_133, 2, x_131); -x_134 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__16; -x_135 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_135, 0, x_86); -lean_ctor_set(x_135, 1, x_134); -x_136 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__3; -x_137 = lean_array_push(x_136, x_89); -x_138 = lean_array_push(x_137, x_133); -x_139 = lean_array_push(x_138, x_135); -x_140 = l_Lean_Parser_Tactic_Conv_paren___closed__2; -x_141 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_141, 0, x_26); -lean_ctor_set(x_141, 1, x_140); -lean_ctor_set(x_141, 2, x_139); -x_142 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_142, 0, x_141); -lean_ctor_set(x_142, 1, x_87); -return x_142; +x_134 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__3; +x_135 = lean_array_push(x_134, x_88); +x_136 = lean_array_push(x_135, x_131); +x_137 = lean_array_push(x_136, x_133); +x_138 = l_Lean_Parser_Tactic_Conv_paren___closed__2; +x_139 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_139, 0, x_25); +lean_ctor_set(x_139, 1, x_138); +lean_ctor_set(x_139, 2, x_137); +x_140 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_140, 0, x_139); +lean_ctor_set(x_140, 1, x_86); +return x_140; } } } } else { -lean_object* x_143; lean_object* x_144; lean_object* x_145; uint8_t x_146; -x_143 = lean_unsigned_to_nat(0u); -x_144 = l_Lean_Syntax_getArg(x_9, x_143); +lean_object* x_141; lean_object* x_142; lean_object* x_143; uint8_t x_144; +x_141 = lean_unsigned_to_nat(0u); +x_142 = l_Lean_Syntax_getArg(x_9, x_141); lean_dec(x_9); -x_145 = l_Lean_Parser_Tactic_Conv_enterArg___closed__2; -lean_inc(x_144); -x_146 = l_Lean_Syntax_isOfKind(x_144, x_145); -if (x_146 == 0) +x_143 = l_Lean_Parser_Tactic_Conv_enterArg___closed__2; +lean_inc(x_142); +x_144 = l_Lean_Syntax_isOfKind(x_142, x_143); +if (x_144 == 0) { -lean_object* x_147; lean_object* x_148; -lean_dec(x_144); +lean_object* x_145; lean_object* x_146; +lean_dec(x_142); lean_dec(x_2); -x_147 = lean_box(1); -x_148 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_148, 0, x_147); -lean_ctor_set(x_148, 1, x_3); -return x_148; +x_145 = lean_box(1); +x_146 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_146, 0, x_145); +lean_ctor_set(x_146, 1, x_3); +return x_146; } else { -lean_object* x_149; lean_object* x_150; uint8_t x_151; -x_149 = l_Lean_Syntax_getArg(x_144, x_143); -lean_dec(x_144); -x_150 = l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__12; -lean_inc(x_149); -x_151 = l_Lean_Syntax_isOfKind(x_149, x_150); +lean_object* x_147; lean_object* x_148; uint8_t x_149; +x_147 = l_Lean_Syntax_getArg(x_142, x_141); +lean_dec(x_142); +x_148 = l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__12; +lean_inc(x_147); +x_149 = l_Lean_Syntax_isOfKind(x_147, x_148); +if (x_149 == 0) +{ +lean_object* x_150; uint8_t x_151; +x_150 = l_Lean_Parser_Tactic_Conv_conv___closed__7; +lean_inc(x_147); +x_151 = l_Lean_Syntax_isOfKind(x_147, x_150); if (x_151 == 0) { -lean_object* x_152; uint8_t x_153; -x_152 = l_Lean_Parser_Tactic_Conv_conv___closed__7; -lean_inc(x_149); -x_153 = l_Lean_Syntax_isOfKind(x_149, x_152); -if (x_153 == 0) -{ -lean_object* x_154; lean_object* x_155; -lean_dec(x_149); +lean_object* x_152; lean_object* x_153; +lean_dec(x_147); lean_dec(x_2); -x_154 = lean_box(1); -x_155 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_155, 0, x_154); -lean_ctor_set(x_155, 1, x_3); -return x_155; +x_152 = lean_box(1); +x_153 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_153, 0, x_152); +lean_ctor_set(x_153, 1, x_3); +return x_153; } else { -lean_object* x_156; uint8_t x_157; -x_156 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_157 = !lean_is_exclusive(x_156); -if (x_157 == 0) -{ -lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; -x_158 = lean_ctor_get(x_156, 0); -x_159 = l_Lean_Parser_Tactic_Conv_ext___closed__1; -x_160 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_160, 0, x_158); -lean_ctor_set(x_160, 1, x_159); -x_161 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__7; -x_162 = lean_array_push(x_161, x_149); -x_163 = lean_box(2); -x_164 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__2; -x_165 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_165, 0, x_163); -lean_ctor_set(x_165, 1, x_164); -lean_ctor_set(x_165, 2, x_162); -x_166 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__6; -x_167 = lean_array_push(x_166, x_160); -x_168 = lean_array_push(x_167, x_165); -x_169 = l_Lean_Parser_Tactic_Conv_ext___closed__2; -x_170 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_170, 0, x_163); -lean_ctor_set(x_170, 1, x_169); -lean_ctor_set(x_170, 2, x_168); -lean_ctor_set(x_156, 0, x_170); -return x_156; +lean_object* x_154; uint8_t x_155; +x_154 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_155 = !lean_is_exclusive(x_154); +if (x_155 == 0) +{ +lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; +x_156 = lean_ctor_get(x_154, 0); +x_157 = l_Lean_Parser_Tactic_Conv_ext___closed__1; +x_158 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_158, 0, x_156); +lean_ctor_set(x_158, 1, x_157); +x_159 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__7; +x_160 = lean_array_push(x_159, x_147); +x_161 = lean_box(2); +x_162 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__2; +x_163 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_163, 0, x_161); +lean_ctor_set(x_163, 1, x_162); +lean_ctor_set(x_163, 2, x_160); +x_164 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__6; +x_165 = lean_array_push(x_164, x_158); +x_166 = lean_array_push(x_165, x_163); +x_167 = l_Lean_Parser_Tactic_Conv_ext___closed__2; +x_168 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_168, 0, x_161); +lean_ctor_set(x_168, 1, x_167); +lean_ctor_set(x_168, 2, x_166); +lean_ctor_set(x_154, 0, x_168); +return x_154; } else { -lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; -x_171 = lean_ctor_get(x_156, 0); -x_172 = lean_ctor_get(x_156, 1); -lean_inc(x_172); -lean_inc(x_171); -lean_dec(x_156); -x_173 = l_Lean_Parser_Tactic_Conv_ext___closed__1; -x_174 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_174, 0, x_171); -lean_ctor_set(x_174, 1, x_173); -x_175 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__7; -x_176 = lean_array_push(x_175, x_149); -x_177 = lean_box(2); -x_178 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__2; -x_179 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_179, 0, x_177); -lean_ctor_set(x_179, 1, x_178); -lean_ctor_set(x_179, 2, x_176); -x_180 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__6; -x_181 = lean_array_push(x_180, x_174); -x_182 = lean_array_push(x_181, x_179); -x_183 = l_Lean_Parser_Tactic_Conv_ext___closed__2; -x_184 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_184, 0, x_177); -lean_ctor_set(x_184, 1, x_183); -lean_ctor_set(x_184, 2, x_182); -x_185 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_185, 0, x_184); -lean_ctor_set(x_185, 1, x_172); -return x_185; +lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; +x_169 = lean_ctor_get(x_154, 0); +x_170 = lean_ctor_get(x_154, 1); +lean_inc(x_170); +lean_inc(x_169); +lean_dec(x_154); +x_171 = l_Lean_Parser_Tactic_Conv_ext___closed__1; +x_172 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_172, 0, x_169); +lean_ctor_set(x_172, 1, x_171); +x_173 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__7; +x_174 = lean_array_push(x_173, x_147); +x_175 = lean_box(2); +x_176 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__2; +x_177 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_177, 0, x_175); +lean_ctor_set(x_177, 1, x_176); +lean_ctor_set(x_177, 2, x_174); +x_178 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__6; +x_179 = lean_array_push(x_178, x_172); +x_180 = lean_array_push(x_179, x_177); +x_181 = l_Lean_Parser_Tactic_Conv_ext___closed__2; +x_182 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_182, 0, x_175); +lean_ctor_set(x_182, 1, x_181); +lean_ctor_set(x_182, 2, x_180); +x_183 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_183, 0, x_182); +lean_ctor_set(x_183, 1, x_170); +return x_183; } } } else { -lean_object* x_186; uint8_t x_187; -x_186 = l_Lean_Syntax_getArg(x_149, x_143); -lean_inc(x_186); -x_187 = l_Lean_Syntax_isNodeOf(x_186, x_10, x_143); -if (x_187 == 0) +lean_object* x_184; uint8_t x_185; +x_184 = l_Lean_Syntax_getArg(x_147, x_141); +lean_inc(x_184); +x_185 = l_Lean_Syntax_matchesNull(x_184, x_141); +if (x_185 == 0) { -uint8_t x_188; -x_188 = l_Lean_Syntax_isNodeOf(x_186, x_10, x_11); -if (x_188 == 0) +uint8_t x_186; +x_186 = l_Lean_Syntax_matchesNull(x_184, x_10); +if (x_186 == 0) { -lean_object* x_189; lean_object* x_190; -lean_dec(x_149); +lean_object* x_187; lean_object* x_188; +lean_dec(x_147); lean_dec(x_2); -x_189 = lean_box(1); -x_190 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_190, 0, x_189); -lean_ctor_set(x_190, 1, x_3); -return x_190; +x_187 = lean_box(1); +x_188 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_188, 0, x_187); +lean_ctor_set(x_188, 1, x_3); +return x_188; } else { -lean_object* x_191; lean_object* x_192; uint8_t x_193; -x_191 = l_Lean_Syntax_getArg(x_149, x_11); -lean_dec(x_149); -x_192 = l_Lean_Parser_Tactic_Conv_arg___closed__10; -lean_inc(x_191); -x_193 = l_Lean_Syntax_isOfKind(x_191, x_192); -if (x_193 == 0) +lean_object* x_189; lean_object* x_190; uint8_t x_191; +x_189 = l_Lean_Syntax_getArg(x_147, x_10); +lean_dec(x_147); +x_190 = l_Lean_Parser_Tactic_Conv_arg___closed__10; +lean_inc(x_189); +x_191 = l_Lean_Syntax_isOfKind(x_189, x_190); +if (x_191 == 0) { -lean_object* x_194; lean_object* x_195; -lean_dec(x_191); +lean_object* x_192; lean_object* x_193; +lean_dec(x_189); lean_dec(x_2); -x_194 = lean_box(1); -x_195 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_195, 0, x_194); -lean_ctor_set(x_195, 1, x_3); -return x_195; +x_192 = lean_box(1); +x_193 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_193, 0, x_192); +lean_ctor_set(x_193, 1, x_3); +return x_193; } else { -lean_object* x_196; uint8_t x_197; -x_196 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_197 = !lean_is_exclusive(x_196); -if (x_197 == 0) -{ -lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; -x_198 = lean_ctor_get(x_196, 0); -x_199 = l_Lean_Parser_Tactic_Conv_arg___closed__1; -lean_inc(x_198); +lean_object* x_194; uint8_t x_195; +x_194 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_195 = !lean_is_exclusive(x_194); +if (x_195 == 0) +{ +lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; +x_196 = lean_ctor_get(x_194, 0); +x_197 = l_Lean_Parser_Tactic_Conv_arg___closed__1; +lean_inc(x_196); +x_198 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_198, 0, x_196); +lean_ctor_set(x_198, 1, x_197); +x_199 = l_Lean_Parser_Tactic_Conv_arg___closed__5; x_200 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_200, 0, x_198); +lean_ctor_set(x_200, 0, x_196); lean_ctor_set(x_200, 1, x_199); -x_201 = l_Lean_Parser_Tactic_Conv_arg___closed__5; -x_202 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_202, 0, x_198); -lean_ctor_set(x_202, 1, x_201); -x_203 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__7; -x_204 = lean_array_push(x_203, x_202); -x_205 = lean_box(2); -x_206 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__2; -x_207 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_207, 0, x_205); -lean_ctor_set(x_207, 1, x_206); -lean_ctor_set(x_207, 2, x_204); -x_208 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__3; -x_209 = lean_array_push(x_208, x_200); -x_210 = lean_array_push(x_209, x_207); -x_211 = lean_array_push(x_210, x_191); -x_212 = l_Lean_Parser_Tactic_Conv_arg___closed__2; -x_213 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_213, 0, x_205); -lean_ctor_set(x_213, 1, x_212); -lean_ctor_set(x_213, 2, x_211); -lean_ctor_set(x_196, 0, x_213); -return x_196; +x_201 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__7; +x_202 = lean_array_push(x_201, x_200); +x_203 = lean_box(2); +x_204 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__2; +x_205 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_205, 0, x_203); +lean_ctor_set(x_205, 1, x_204); +lean_ctor_set(x_205, 2, x_202); +x_206 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__3; +x_207 = lean_array_push(x_206, x_198); +x_208 = lean_array_push(x_207, x_205); +x_209 = lean_array_push(x_208, x_189); +x_210 = l_Lean_Parser_Tactic_Conv_arg___closed__2; +x_211 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_211, 0, x_203); +lean_ctor_set(x_211, 1, x_210); +lean_ctor_set(x_211, 2, x_209); +lean_ctor_set(x_194, 0, x_211); +return x_194; } else { -lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; -x_214 = lean_ctor_get(x_196, 0); -x_215 = lean_ctor_get(x_196, 1); -lean_inc(x_215); -lean_inc(x_214); -lean_dec(x_196); -x_216 = l_Lean_Parser_Tactic_Conv_arg___closed__1; -lean_inc(x_214); +lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; +x_212 = lean_ctor_get(x_194, 0); +x_213 = lean_ctor_get(x_194, 1); +lean_inc(x_213); +lean_inc(x_212); +lean_dec(x_194); +x_214 = l_Lean_Parser_Tactic_Conv_arg___closed__1; +lean_inc(x_212); +x_215 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_215, 0, x_212); +lean_ctor_set(x_215, 1, x_214); +x_216 = l_Lean_Parser_Tactic_Conv_arg___closed__5; x_217 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_217, 0, x_214); +lean_ctor_set(x_217, 0, x_212); lean_ctor_set(x_217, 1, x_216); -x_218 = l_Lean_Parser_Tactic_Conv_arg___closed__5; -x_219 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_219, 0, x_214); -lean_ctor_set(x_219, 1, x_218); -x_220 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__7; -x_221 = lean_array_push(x_220, x_219); -x_222 = lean_box(2); -x_223 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__2; -x_224 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_224, 0, x_222); -lean_ctor_set(x_224, 1, x_223); -lean_ctor_set(x_224, 2, x_221); -x_225 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__3; -x_226 = lean_array_push(x_225, x_217); -x_227 = lean_array_push(x_226, x_224); -x_228 = lean_array_push(x_227, x_191); -x_229 = l_Lean_Parser_Tactic_Conv_arg___closed__2; -x_230 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_230, 0, x_222); -lean_ctor_set(x_230, 1, x_229); -lean_ctor_set(x_230, 2, x_228); -x_231 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_231, 0, x_230); -lean_ctor_set(x_231, 1, x_215); -return x_231; +x_218 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__7; +x_219 = lean_array_push(x_218, x_217); +x_220 = lean_box(2); +x_221 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__2; +x_222 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_222, 0, x_220); +lean_ctor_set(x_222, 1, x_221); +lean_ctor_set(x_222, 2, x_219); +x_223 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__3; +x_224 = lean_array_push(x_223, x_215); +x_225 = lean_array_push(x_224, x_222); +x_226 = lean_array_push(x_225, x_189); +x_227 = l_Lean_Parser_Tactic_Conv_arg___closed__2; +x_228 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_228, 0, x_220); +lean_ctor_set(x_228, 1, x_227); +lean_ctor_set(x_228, 2, x_226); +x_229 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_229, 0, x_228); +lean_ctor_set(x_229, 1, x_213); +return x_229; } } } } else { -lean_object* x_232; lean_object* x_233; uint8_t x_234; -lean_dec(x_186); -x_232 = l_Lean_Syntax_getArg(x_149, x_11); -lean_dec(x_149); -x_233 = l_Lean_Parser_Tactic_Conv_arg___closed__10; -lean_inc(x_232); -x_234 = l_Lean_Syntax_isOfKind(x_232, x_233); -if (x_234 == 0) -{ -lean_object* x_235; lean_object* x_236; -lean_dec(x_232); +lean_object* x_230; lean_object* x_231; uint8_t x_232; +lean_dec(x_184); +x_230 = l_Lean_Syntax_getArg(x_147, x_10); +lean_dec(x_147); +x_231 = l_Lean_Parser_Tactic_Conv_arg___closed__10; +lean_inc(x_230); +x_232 = l_Lean_Syntax_isOfKind(x_230, x_231); +if (x_232 == 0) +{ +lean_object* x_233; lean_object* x_234; +lean_dec(x_230); lean_dec(x_2); -x_235 = lean_box(1); -x_236 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_236, 0, x_235); -lean_ctor_set(x_236, 1, x_3); -return x_236; +x_233 = lean_box(1); +x_234 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_234, 0, x_233); +lean_ctor_set(x_234, 1, x_3); +return x_234; } else { -lean_object* x_237; uint8_t x_238; -x_237 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_238 = !lean_is_exclusive(x_237); -if (x_238 == 0) -{ -lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; -x_239 = lean_ctor_get(x_237, 0); -x_240 = l_Lean_Parser_Tactic_Conv_arg___closed__1; -x_241 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_241, 0, x_239); -lean_ctor_set(x_241, 1, x_240); -x_242 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__3; -x_243 = lean_array_push(x_242, x_241); -x_244 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__5; -x_245 = lean_array_push(x_243, x_244); -x_246 = lean_array_push(x_245, x_232); -x_247 = lean_box(2); -x_248 = l_Lean_Parser_Tactic_Conv_arg___closed__2; -x_249 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_249, 0, x_247); -lean_ctor_set(x_249, 1, x_248); -lean_ctor_set(x_249, 2, x_246); -lean_ctor_set(x_237, 0, x_249); -return x_237; +lean_object* x_235; uint8_t x_236; +x_235 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_236 = !lean_is_exclusive(x_235); +if (x_236 == 0) +{ +lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; +x_237 = lean_ctor_get(x_235, 0); +x_238 = l_Lean_Parser_Tactic_Conv_arg___closed__1; +x_239 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_239, 0, x_237); +lean_ctor_set(x_239, 1, x_238); +x_240 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__3; +x_241 = lean_array_push(x_240, x_239); +x_242 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__5; +x_243 = lean_array_push(x_241, x_242); +x_244 = lean_array_push(x_243, x_230); +x_245 = lean_box(2); +x_246 = l_Lean_Parser_Tactic_Conv_arg___closed__2; +x_247 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_247, 0, x_245); +lean_ctor_set(x_247, 1, x_246); +lean_ctor_set(x_247, 2, x_244); +lean_ctor_set(x_235, 0, x_247); +return x_235; } else { -lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; -x_250 = lean_ctor_get(x_237, 0); -x_251 = lean_ctor_get(x_237, 1); -lean_inc(x_251); -lean_inc(x_250); -lean_dec(x_237); -x_252 = l_Lean_Parser_Tactic_Conv_arg___closed__1; -x_253 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_253, 0, x_250); -lean_ctor_set(x_253, 1, x_252); -x_254 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__3; -x_255 = lean_array_push(x_254, x_253); -x_256 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__5; -x_257 = lean_array_push(x_255, x_256); -x_258 = lean_array_push(x_257, x_232); -x_259 = lean_box(2); -x_260 = l_Lean_Parser_Tactic_Conv_arg___closed__2; -x_261 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_261, 0, x_259); -lean_ctor_set(x_261, 1, x_260); -lean_ctor_set(x_261, 2, x_258); -x_262 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_262, 0, x_261); -lean_ctor_set(x_262, 1, x_251); -return x_262; +lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; +x_248 = lean_ctor_get(x_235, 0); +x_249 = lean_ctor_get(x_235, 1); +lean_inc(x_249); +lean_inc(x_248); +lean_dec(x_235); +x_250 = l_Lean_Parser_Tactic_Conv_arg___closed__1; +x_251 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_251, 0, x_248); +lean_ctor_set(x_251, 1, x_250); +x_252 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__3; +x_253 = lean_array_push(x_252, x_251); +x_254 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv_xb7_x2e____1___closed__5; +x_255 = lean_array_push(x_253, x_254); +x_256 = lean_array_push(x_255, x_230); +x_257 = lean_box(2); +x_258 = l_Lean_Parser_Tactic_Conv_arg___closed__2; +x_259 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_259, 0, x_257); +lean_ctor_set(x_259, 1, x_258); +lean_ctor_set(x_259, 2, x_256); +x_260 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_260, 0, x_259); +lean_ctor_set(x_260, 1, x_249); +return x_260; } } } @@ -6780,7 +6955,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_convApply_____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_convApply_____closed__4; x_3 = l_Lean_Parser_Tactic_Conv_conv___closed__16; x_4 = lean_alloc_ctor(2, 3, 0); @@ -7057,7 +7232,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_first___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__15; x_3 = l_Lean_Parser_Tactic_Conv_first___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); @@ -7071,7 +7246,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_first___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_first___closed__7; x_3 = l_Lean_Parser_Tactic_Conv_convSeq; x_4 = lean_alloc_ctor(2, 3, 0); @@ -7121,7 +7296,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_first___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_first___closed__4; x_3 = l_Lean_Parser_Tactic_Conv_first___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); @@ -7195,7 +7370,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_convRepeat_____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__13; x_2 = l_Lean_Parser_Tactic_Conv_convRepeat_____closed__4; x_3 = l_Lean_Parser_Tactic_Conv_convSeq; x_4 = lean_alloc_ctor(2, 3, 0); @@ -7280,7 +7455,7 @@ lean_inc(x_12); x_18 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_18, 0, x_12); lean_ctor_set(x_18, 1, x_17); -x_19 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__16; +x_19 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__17; lean_inc(x_12); x_20 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_20, 0, x_12); @@ -7434,7 +7609,7 @@ lean_inc(x_83); x_90 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_90, 0, x_83); lean_ctor_set(x_90, 1, x_89); -x_91 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__16; +x_91 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__17; lean_inc(x_83); x_92 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_92, 0, x_83); @@ -7619,6 +7794,10 @@ l_Lean_Parser_Tactic_Conv_conv_quot___closed__19 = _init_l_Lean_Parser_Tactic_Co lean_mark_persistent(l_Lean_Parser_Tactic_Conv_conv_quot___closed__19); l_Lean_Parser_Tactic_Conv_conv_quot___closed__20 = _init_l_Lean_Parser_Tactic_Conv_conv_quot___closed__20(); lean_mark_persistent(l_Lean_Parser_Tactic_Conv_conv_quot___closed__20); +l_Lean_Parser_Tactic_Conv_conv_quot___closed__21 = _init_l_Lean_Parser_Tactic_Conv_conv_quot___closed__21(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_conv_quot___closed__21); +l_Lean_Parser_Tactic_Conv_conv_quot___closed__22 = _init_l_Lean_Parser_Tactic_Conv_conv_quot___closed__22(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_conv_quot___closed__22); l_Lean_Parser_Tactic_Conv_conv_quot = _init_l_Lean_Parser_Tactic_Conv_conv_quot(); lean_mark_persistent(l_Lean_Parser_Tactic_Conv_conv_quot); l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__1 = _init_l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__1(); @@ -8267,20 +8446,22 @@ l_Lean_Parser_Tactic_Conv_convRight___closed__5 = _init_l_Lean_Parser_Tactic_Con lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convRight___closed__5); l_Lean_Parser_Tactic_Conv_convRight = _init_l_Lean_Parser_Tactic_Conv_convRight(); lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convRight); -l_Lean_Parser_Tactic_Conv_convIntro_________closed__1 = _init_l_Lean_Parser_Tactic_Conv_convIntro_________closed__1(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convIntro_________closed__1); -l_Lean_Parser_Tactic_Conv_convIntro_________closed__2 = _init_l_Lean_Parser_Tactic_Conv_convIntro_________closed__2(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convIntro_________closed__2); -l_Lean_Parser_Tactic_Conv_convIntro_________closed__3 = _init_l_Lean_Parser_Tactic_Conv_convIntro_________closed__3(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convIntro_________closed__3); -l_Lean_Parser_Tactic_Conv_convIntro_________closed__4 = _init_l_Lean_Parser_Tactic_Conv_convIntro_________closed__4(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convIntro_________closed__4); -l_Lean_Parser_Tactic_Conv_convIntro_________closed__5 = _init_l_Lean_Parser_Tactic_Conv_convIntro_________closed__5(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convIntro_________closed__5); -l_Lean_Parser_Tactic_Conv_convIntro_________closed__6 = _init_l_Lean_Parser_Tactic_Conv_convIntro_________closed__6(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convIntro_________closed__6); -l_Lean_Parser_Tactic_Conv_convIntro______ = _init_l_Lean_Parser_Tactic_Conv_convIntro______(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convIntro______); +l_Lean_Parser_Tactic_Conv_convIntro_______closed__1 = _init_l_Lean_Parser_Tactic_Conv_convIntro_______closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convIntro_______closed__1); +l_Lean_Parser_Tactic_Conv_convIntro_______closed__2 = _init_l_Lean_Parser_Tactic_Conv_convIntro_______closed__2(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convIntro_______closed__2); +l_Lean_Parser_Tactic_Conv_convIntro_______closed__3 = _init_l_Lean_Parser_Tactic_Conv_convIntro_______closed__3(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convIntro_______closed__3); +l_Lean_Parser_Tactic_Conv_convIntro_______closed__4 = _init_l_Lean_Parser_Tactic_Conv_convIntro_______closed__4(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convIntro_______closed__4); +l_Lean_Parser_Tactic_Conv_convIntro_______closed__5 = _init_l_Lean_Parser_Tactic_Conv_convIntro_______closed__5(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convIntro_______closed__5); +l_Lean_Parser_Tactic_Conv_convIntro_______closed__6 = _init_l_Lean_Parser_Tactic_Conv_convIntro_______closed__6(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convIntro_______closed__6); +l_Lean_Parser_Tactic_Conv_convIntro____ = _init_l_Lean_Parser_Tactic_Conv_convIntro____(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convIntro____); +l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convIntro______1___closed__1 = _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convIntro______1___closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convIntro______1___closed__1); l_Lean_Parser_Tactic_Conv_enterArg___closed__1 = _init_l_Lean_Parser_Tactic_Conv_enterArg___closed__1(); lean_mark_persistent(l_Lean_Parser_Tactic_Conv_enterArg___closed__1); l_Lean_Parser_Tactic_Conv_enterArg___closed__2 = _init_l_Lean_Parser_Tactic_Conv_enterArg___closed__2(); diff --git a/stage0/stdlib/Init/Core.c b/stage0/stdlib/Init/Core.c index 5c60a887398..83c3161cc27 100644 --- a/stage0/stdlib/Init/Core.c +++ b/stage0/stdlib/Init/Core.c @@ -30,7 +30,6 @@ LEAN_EXPORT lean_object* l_inline(lean_object*); LEAN_EXPORT lean_object* l_instLTProd___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Quotient_lift_u2082___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_nullKind; static lean_object* l___aux__Init__Core______macroRules__term___u2248____1___closed__4; LEAN_EXPORT lean_object* l_Quotient_rec___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___aux__Init__Core______macroRules__term_u2205__1(lean_object*, lean_object*, lean_object*); @@ -150,6 +149,7 @@ static lean_object* l_term___u2194_____closed__4; LEAN_EXPORT lean_object* l_Sum_inhabitedLeft___rarg(lean_object*); static lean_object* l_term___u2295_____closed__4; LEAN_EXPORT lean_object* l_instDecidableDitePropNot___rarg___boxed(lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); static lean_object* l_term___u2295_x27_____closed__2; static lean_object* l_term___u2295_x27_____closed__1; LEAN_EXPORT lean_object* l_Quot_hrecOn(lean_object*, lean_object*, lean_object*); @@ -240,7 +240,6 @@ LEAN_EXPORT lean_object* l_Thunk_map___rarg(lean_object*, lean_object*); static lean_object* l_term___x3c_x2d_x3e_____closed__3; LEAN_EXPORT lean_object* l_instDecidableIteProp(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Eq_mp(lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_decidable__of__decidable__of__eq___rarg(uint8_t, lean_object*); static lean_object* l_term___u2295_____closed__6; static lean_object* l___aux__Init__Core______macroRules__term___x3c_x2d_x3e____1___closed__16; @@ -1009,82 +1008,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) -{ -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___x3c_x2d_x3e_____closed__5; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Core______unexpand__Iff__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___x3c_x2d_x3e_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return x_24; +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(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; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___x3c_x2d_x3e_____closed__5; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Core______unexpand__Iff__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___x3c_x2d_x3e_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___x3c_x2d_x3e_____closed__5; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Core______unexpand__Iff__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___x3c_x2d_x3e_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___x3c_x2d_x3e_____closed__5; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Core______unexpand__Iff__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___x3c_x2d_x3e_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -1324,82 +1322,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) -{ -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___u2194_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Core______unexpand__Iff__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___u2194_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return x_24; +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(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; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___u2194_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Core______unexpand__Iff__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___u2194_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___u2194_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Core______unexpand__Iff__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___u2194_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___u2194_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Core______unexpand__Iff__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___u2194_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -1716,82 +1713,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) -{ -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___u2295_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Core______unexpand__Iff__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___u2295_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return x_24; +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(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; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___u2295_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Core______unexpand__Iff__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___u2295_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___u2295_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Core______unexpand__Iff__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___u2295_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___u2295_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Core______unexpand__Iff__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___u2295_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -2096,82 +2092,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) -{ -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___u2295_x27_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Core______unexpand__Iff__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___u2295_x27_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return x_24; +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(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; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___u2295_x27_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Core______unexpand__Iff__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___u2295_x27_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___u2295_x27_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Core______unexpand__Iff__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___u2295_x27_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___u2295_x27_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Core______unexpand__Iff__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___u2295_x27_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -2514,82 +2509,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) -{ -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___u2248_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Core______unexpand__Iff__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___u2248_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return x_24; +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(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; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___u2248_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Core______unexpand__Iff__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___u2248_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___u2248_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Core______unexpand__Iff__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___u2248_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___u2248_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Core______unexpand__Iff__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___u2248_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -3605,82 +3599,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) -{ -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___x21_x3d_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Core______unexpand__Iff__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___x21_x3d_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return x_24; +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(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; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___x21_x3d_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Core______unexpand__Iff__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___x21_x3d_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___x21_x3d_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Core______unexpand__Iff__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___x21_x3d_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___x21_x3d_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Core______unexpand__Iff__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___x21_x3d_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -4033,82 +4026,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) -{ -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___u2260_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Core______unexpand__Iff__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___u2260_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return x_24; +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(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; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___u2260_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Core______unexpand__Iff__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___u2260_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___u2260_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Core______unexpand__Iff__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___u2260_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___u2260_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Core______unexpand__Iff__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___u2260_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } diff --git a/stage0/stdlib/Init/Data/Format/Macro.c b/stage0/stdlib/Init/Data/Format/Macro.c index 84c8dbb29ca..9cebb78e1da 100644 --- a/stage0/stdlib/Init/Data/Format/Macro.c +++ b/stage0/stdlib/Init/Data/Format/Macro.c @@ -30,8 +30,8 @@ static lean_object* l_Std___aux__Init__Data__Format__Macro______macroRules__Std_ static lean_object* l_Std___aux__Init__Data__Format__Macro______macroRules__Std__termF_x21____1___closed__4; LEAN_EXPORT lean_object* l_Std_termF_x21__; static lean_object* l_Std___aux__Init__Data__Format__Macro______macroRules__Std__termF_x21____1___closed__10; -lean_object* l_Lean_Syntax_expandInterpolatedStr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_termF_x21_____closed__10; +lean_object* l_Lean_TSyntax_expandInterpolatedStr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_termF_x21_____closed__14; static lean_object* l_Std_termF_x21_____closed__5; static lean_object* l_Std_termF_x21_____closed__3; @@ -461,10 +461,53 @@ lean_ctor_set(x_27, 0, x_21); lean_ctor_set(x_27, 1, x_25); lean_ctor_set(x_27, 2, x_24); lean_ctor_set(x_27, 3, x_26); -x_28 = l_Lean_Syntax_expandInterpolatedStr(x_9, x_19, x_27, x_2, x_22); +x_28 = l_Lean_TSyntax_expandInterpolatedStr(x_9, x_19, x_27, x_2, x_22); lean_dec(x_9); +if (lean_obj_tag(x_28) == 0) +{ +uint8_t x_29; +x_29 = !lean_is_exclusive(x_28); +if (x_29 == 0) +{ +return x_28; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_28, 0); +x_31 = lean_ctor_get(x_28, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_28); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +return x_32; +} +} +else +{ +uint8_t x_33; +x_33 = !lean_is_exclusive(x_28); +if (x_33 == 0) +{ return x_28; } +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_28, 0); +x_35 = lean_ctor_get(x_28, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_28); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +} +} } } lean_object* initialize_Init_Data_Format_Basic(uint8_t builtin, lean_object*); diff --git a/stage0/stdlib/Init/Data/Format/Syntax.c b/stage0/stdlib/Init/Data/Format/Syntax.c index 35132a44b9d..2e54fc3e57e 100644 --- a/stage0/stdlib/Init/Data/Format/Syntax.c +++ b/stage0/stdlib/Init/Data/Format/Syntax.c @@ -19,8 +19,9 @@ static lean_object* l_Lean_Syntax_formatStxAux___closed__1; LEAN_EXPORT lean_object* l_List_mapTRAux___at_Lean_Syntax_formatStxAux___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Syntax_formatStxAux___closed__13; static lean_object* l_Lean_Syntax_formatStxAux___closed__23; -extern lean_object* l_Lean_nullKind; +LEAN_EXPORT lean_object* l_Lean_Syntax_instToFormatTSyntax(lean_object*); lean_object* lean_name_mk_string(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Syntax_instToStringTSyntax(lean_object*); extern lean_object* l_Std_Format_defWidth; static lean_object* l_Lean_Syntax_formatStxAux___closed__21; static lean_object* l_Lean_Syntax_formatStxAux___closed__25; @@ -29,6 +30,7 @@ LEAN_EXPORT lean_object* l___private_Init_Data_Format_Syntax_0__Lean_Syntax_form static lean_object* l_Lean_Syntax_formatStxAux___closed__7; uint8_t lean_name_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTRAux___at_Lean_Syntax_formatStxAux___spec__4(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Syntax_instToStringTSyntax___boxed(lean_object*); static lean_object* l_Lean_Syntax_formatStxAux___closed__17; static lean_object* l_Lean_Syntax_formatStxAux___closed__8; lean_object* lean_array_get_size(lean_object*); @@ -48,21 +50,25 @@ lean_object* l_Lean_Name_toString(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Syntax_instToFormatSyntax(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_formatStx___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTRAux___at_Lean_Syntax_formatStxAux___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Syntax_formatStxAux___closed__29; LEAN_EXPORT lean_object* l_List_mapTRAux___at_Lean_Syntax_formatStxAux___spec__5(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); lean_object* l_Nat_repr(lean_object*); static lean_object* l_Lean_Syntax_formatStxAux___closed__12; lean_object* lean_format_pretty(lean_object*, lean_object*); +static lean_object* l_Lean_Syntax_formatStxAux___closed__28; LEAN_EXPORT lean_object* l_Lean_Syntax_formatStx(lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_Syntax_formatStxAux___closed__22; static lean_object* l___private_Init_Data_Format_Syntax_0__Lean_Syntax_formatInfo___closed__1; lean_object* l_Function_comp___rarg(lean_object*, lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Syntax_instToFormatTSyntax___boxed(lean_object*); static lean_object* l_Lean_Syntax_formatStxAux___closed__10; LEAN_EXPORT lean_object* l_List_mapTRAux___at_Lean_Syntax_formatStxAux___spec__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_instToStringSyntax___lambda__2(lean_object*); static lean_object* l___private_Init_Data_Format_Syntax_0__Lean_Syntax_formatInfo___closed__4; LEAN_EXPORT lean_object* l_List_mapTRAux___at_Lean_Syntax_formatStxAux___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_instToStringSyntax; +LEAN_EXPORT lean_object* l_Lean_Syntax_instToFormatTSyntax___rarg(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Format_Syntax_0__Lean_Syntax_formatInfo___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTRAux___at_Lean_Syntax_formatStxAux___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTRAux___at_Lean_Syntax_formatStxAux___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -84,6 +90,7 @@ static lean_object* l_Lean_Syntax_formatStxAux___closed__6; static lean_object* l___private_Init_Data_Format_Syntax_0__Lean_Syntax_formatInfo___closed__3; static lean_object* l_Lean_Syntax_formatStxAux___closed__26; LEAN_EXPORT lean_object* l_List_mapTRAux___at_Lean_Syntax_formatStxAux___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Syntax_instToStringTSyntax___rarg(lean_object*); lean_object* l_Lean_Name_replacePrefix(lean_object*, lean_object*, lean_object*); lean_object* lean_nat_to_int(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_formatStxAux(lean_object*, uint8_t, lean_object*, lean_object*); @@ -577,7 +584,7 @@ static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Lean", 4); +x_1 = lean_mk_string_from_bytes("null", 4); return x_1; } } @@ -595,7 +602,7 @@ static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__5() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Parser", 6); +x_1 = lean_mk_string_from_bytes("Lean", 4); return x_1; } } @@ -603,7 +610,7 @@ static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Syntax_formatStxAux___closed__4; +x_1 = lean_box(0); x_2 = l_Lean_Syntax_formatStxAux___closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -613,39 +620,57 @@ static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__7() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("(", 1); +x_1 = lean_mk_string_from_bytes("Parser", 6); return x_1; } } static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__8() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Syntax_formatStxAux___closed__6; +x_2 = l_Lean_Syntax_formatStxAux___closed__7; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("(", 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__10() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Syntax_formatStxAux___closed__7; +x_1 = l_Lean_Syntax_formatStxAux___closed__9; x_2 = lean_string_length(x_1); return x_2; } } -static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__9() { +static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Syntax_formatStxAux___closed__8; +x_1 = l_Lean_Syntax_formatStxAux___closed__10; x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__10() { +static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Syntax_formatStxAux___closed__7; +x_1 = l_Lean_Syntax_formatStxAux___closed__9; x_2 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__11() { +static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__13() { _start: { lean_object* x_1; @@ -653,17 +678,17 @@ x_1 = lean_mk_string_from_bytes(")", 1); return x_1; } } -static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__12() { +static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__14() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Syntax_formatStxAux___closed__11; +x_1 = l_Lean_Syntax_formatStxAux___closed__13; x_2 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__13() { +static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__15() { _start: { lean_object* x_1; @@ -671,29 +696,29 @@ x_1 = lean_mk_string_from_bytes("..", 2); return x_1; } } -static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__14() { +static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__16() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Syntax_formatStxAux___closed__13; +x_1 = l_Lean_Syntax_formatStxAux___closed__15; x_2 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__15() { +static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Syntax_formatStxAux___closed__14; +x_2 = l_Lean_Syntax_formatStxAux___closed__16; x_3 = lean_alloc_ctor(1, 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_Syntax_formatStxAux___closed__16() { +static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__18() { _start: { lean_object* x_1; @@ -701,35 +726,35 @@ x_1 = lean_mk_string_from_bytes("[", 1); return x_1; } } -static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__17() { +static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__19() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Syntax_formatStxAux___closed__16; +x_1 = l_Lean_Syntax_formatStxAux___closed__18; x_2 = lean_string_length(x_1); return x_2; } } -static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__18() { +static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__20() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Syntax_formatStxAux___closed__17; +x_1 = l_Lean_Syntax_formatStxAux___closed__19; x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__19() { +static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__21() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Syntax_formatStxAux___closed__16; +x_1 = l_Lean_Syntax_formatStxAux___closed__18; x_2 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__20() { +static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__22() { _start: { lean_object* x_1; @@ -737,57 +762,57 @@ x_1 = lean_mk_string_from_bytes("]", 1); return x_1; } } -static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__21() { +static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__23() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Syntax_formatStxAux___closed__20; +x_1 = l_Lean_Syntax_formatStxAux___closed__22; x_2 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__22() { +static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Syntax_formatStxAux___closed__19; -x_2 = l_Lean_Syntax_formatStxAux___closed__14; +x_1 = l_Lean_Syntax_formatStxAux___closed__21; +x_2 = l_Lean_Syntax_formatStxAux___closed__16; x_3 = lean_alloc_ctor(4, 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_Lean_Syntax_formatStxAux___closed__23() { +static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Syntax_formatStxAux___closed__22; -x_2 = l_Lean_Syntax_formatStxAux___closed__21; +x_1 = l_Lean_Syntax_formatStxAux___closed__24; +x_2 = l_Lean_Syntax_formatStxAux___closed__23; x_3 = lean_alloc_ctor(4, 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_Lean_Syntax_formatStxAux___closed__24() { +static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Syntax_formatStxAux___closed__18; -x_2 = l_Lean_Syntax_formatStxAux___closed__23; +x_1 = l_Lean_Syntax_formatStxAux___closed__20; +x_2 = l_Lean_Syntax_formatStxAux___closed__25; 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_Lean_Syntax_formatStxAux___closed__25() { +static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__27() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_Syntax_formatStxAux___closed__24; +x_1 = l_Lean_Syntax_formatStxAux___closed__26; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -795,7 +820,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__26() { +static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__28() { _start: { lean_object* x_1; @@ -803,11 +828,11 @@ x_1 = lean_mk_string_from_bytes("`", 1); return x_1; } } -static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__27() { +static lean_object* _init_l_Lean_Syntax_formatStxAux___closed__29() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Syntax_formatStxAux___closed__26; +x_1 = l_Lean_Syntax_formatStxAux___closed__28; x_2 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; @@ -833,12 +858,12 @@ lean_inc(x_7); lean_dec(x_4); x_8 = lean_unsigned_to_nat(1u); x_9 = lean_nat_add(x_3, x_8); -x_10 = l_Lean_nullKind; +x_10 = l_Lean_Syntax_formatStxAux___closed__4; x_11 = lean_name_eq(x_6, x_10); if (x_11 == 0) { lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_12 = l_Lean_Syntax_formatStxAux___closed__6; +x_12 = l_Lean_Syntax_formatStxAux___closed__8; x_13 = lean_box(0); x_14 = l_Lean_Name_replacePrefix(x_6, x_12, x_13); x_15 = 1; @@ -862,15 +887,15 @@ lean_ctor_set(x_24, 1, x_23); x_25 = lean_box(1); x_26 = l_Std_Format_joinSep___at_instReprProd___spec__1(x_24, x_25); lean_dec(x_24); -x_27 = l_Lean_Syntax_formatStxAux___closed__10; +x_27 = l_Lean_Syntax_formatStxAux___closed__12; x_28 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_28, 0, x_27); lean_ctor_set(x_28, 1, x_26); -x_29 = l_Lean_Syntax_formatStxAux___closed__12; +x_29 = l_Lean_Syntax_formatStxAux___closed__14; x_30 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_30, 0, x_28); lean_ctor_set(x_30, 1, x_29); -x_31 = l_Lean_Syntax_formatStxAux___closed__9; +x_31 = l_Lean_Syntax_formatStxAux___closed__11; x_32 = lean_alloc_ctor(3, 2, 0); lean_ctor_set(x_32, 0, x_31); lean_ctor_set(x_32, 1, x_30); @@ -899,15 +924,15 @@ lean_ctor_set(x_39, 1, x_38); x_40 = lean_box(1); x_41 = l_Std_Format_joinSep___at_instReprProd___spec__1(x_39, x_40); lean_dec(x_39); -x_42 = l_Lean_Syntax_formatStxAux___closed__10; +x_42 = l_Lean_Syntax_formatStxAux___closed__12; x_43 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_43, 0, x_42); lean_ctor_set(x_43, 1, x_41); -x_44 = l_Lean_Syntax_formatStxAux___closed__12; +x_44 = l_Lean_Syntax_formatStxAux___closed__14; x_45 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_45, 0, x_43); lean_ctor_set(x_45, 1, x_44); -x_46 = l_Lean_Syntax_formatStxAux___closed__9; +x_46 = l_Lean_Syntax_formatStxAux___closed__11; x_47 = lean_alloc_ctor(3, 2, 0); lean_ctor_set(x_47, 0, x_46); lean_ctor_set(x_47, 1, x_45); @@ -922,22 +947,22 @@ else lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; uint8_t x_60; lean_object* x_61; lean_dec(x_9); lean_dec(x_7); -x_50 = l_Lean_Syntax_formatStxAux___closed__15; +x_50 = l_Lean_Syntax_formatStxAux___closed__17; x_51 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_51, 0, x_17); lean_ctor_set(x_51, 1, x_50); x_52 = lean_box(1); x_53 = l_Std_Format_joinSep___at_instReprProd___spec__1(x_51, x_52); lean_dec(x_51); -x_54 = l_Lean_Syntax_formatStxAux___closed__10; +x_54 = l_Lean_Syntax_formatStxAux___closed__12; x_55 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_55, 0, x_54); lean_ctor_set(x_55, 1, x_53); -x_56 = l_Lean_Syntax_formatStxAux___closed__12; +x_56 = l_Lean_Syntax_formatStxAux___closed__14; x_57 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_57, 0, x_55); lean_ctor_set(x_57, 1, x_56); -x_58 = l_Lean_Syntax_formatStxAux___closed__9; +x_58 = l_Lean_Syntax_formatStxAux___closed__11; x_59 = lean_alloc_ctor(3, 2, 0); lean_ctor_set(x_59, 0, x_58); lean_ctor_set(x_59, 1, x_57); @@ -966,15 +991,15 @@ lean_ctor_set(x_67, 1, x_66); x_68 = lean_box(1); x_69 = l_Std_Format_joinSep___at_instReprProd___spec__1(x_67, x_68); lean_dec(x_67); -x_70 = l_Lean_Syntax_formatStxAux___closed__10; +x_70 = l_Lean_Syntax_formatStxAux___closed__12; x_71 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_71, 0, x_70); lean_ctor_set(x_71, 1, x_69); -x_72 = l_Lean_Syntax_formatStxAux___closed__12; +x_72 = l_Lean_Syntax_formatStxAux___closed__14; x_73 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_73, 0, x_71); lean_ctor_set(x_73, 1, x_72); -x_74 = l_Lean_Syntax_formatStxAux___closed__9; +x_74 = l_Lean_Syntax_formatStxAux___closed__11; x_75 = lean_alloc_ctor(3, 2, 0); lean_ctor_set(x_75, 0, x_74); lean_ctor_set(x_75, 1, x_73); @@ -989,22 +1014,22 @@ else lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; uint8_t x_88; lean_object* x_89; lean_dec(x_9); lean_dec(x_7); -x_78 = l_Lean_Syntax_formatStxAux___closed__15; +x_78 = l_Lean_Syntax_formatStxAux___closed__17; x_79 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_79, 0, x_17); lean_ctor_set(x_79, 1, x_78); x_80 = lean_box(1); x_81 = l_Std_Format_joinSep___at_instReprProd___spec__1(x_79, x_80); lean_dec(x_79); -x_82 = l_Lean_Syntax_formatStxAux___closed__10; +x_82 = l_Lean_Syntax_formatStxAux___closed__12; x_83 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_83, 0, x_82); lean_ctor_set(x_83, 1, x_81); -x_84 = l_Lean_Syntax_formatStxAux___closed__12; +x_84 = l_Lean_Syntax_formatStxAux___closed__14; x_85 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_85, 0, x_83); lean_ctor_set(x_85, 1, x_84); -x_86 = l_Lean_Syntax_formatStxAux___closed__9; +x_86 = l_Lean_Syntax_formatStxAux___closed__11; x_87 = lean_alloc_ctor(3, 2, 0); lean_ctor_set(x_87, 0, x_86); lean_ctor_set(x_87, 1, x_85); @@ -1035,15 +1060,15 @@ lean_dec(x_9); x_96 = lean_box(1); x_97 = l_Std_Format_joinSep___at_instReprProd___spec__1(x_95, x_96); lean_dec(x_95); -x_98 = l_Lean_Syntax_formatStxAux___closed__19; +x_98 = l_Lean_Syntax_formatStxAux___closed__21; x_99 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_99, 0, x_98); lean_ctor_set(x_99, 1, x_97); -x_100 = l_Lean_Syntax_formatStxAux___closed__21; +x_100 = l_Lean_Syntax_formatStxAux___closed__23; x_101 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_101, 0, x_99); lean_ctor_set(x_101, 1, x_100); -x_102 = l_Lean_Syntax_formatStxAux___closed__18; +x_102 = l_Lean_Syntax_formatStxAux___closed__20; x_103 = lean_alloc_ctor(3, 2, 0); lean_ctor_set(x_103, 0, x_102); lean_ctor_set(x_103, 1, x_101); @@ -1069,15 +1094,15 @@ lean_dec(x_9); x_110 = lean_box(1); x_111 = l_Std_Format_joinSep___at_instReprProd___spec__1(x_109, x_110); lean_dec(x_109); -x_112 = l_Lean_Syntax_formatStxAux___closed__19; +x_112 = l_Lean_Syntax_formatStxAux___closed__21; x_113 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_113, 0, x_112); lean_ctor_set(x_113, 1, x_111); -x_114 = l_Lean_Syntax_formatStxAux___closed__21; +x_114 = l_Lean_Syntax_formatStxAux___closed__23; x_115 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_115, 0, x_113); lean_ctor_set(x_115, 1, x_114); -x_116 = l_Lean_Syntax_formatStxAux___closed__18; +x_116 = l_Lean_Syntax_formatStxAux___closed__20; x_117 = lean_alloc_ctor(3, 2, 0); lean_ctor_set(x_117, 0, x_116); lean_ctor_set(x_117, 1, x_115); @@ -1092,7 +1117,7 @@ else lean_object* x_120; lean_dec(x_9); lean_dec(x_7); -x_120 = l_Lean_Syntax_formatStxAux___closed__25; +x_120 = l_Lean_Syntax_formatStxAux___closed__27; return x_120; } } @@ -1111,15 +1136,15 @@ lean_dec(x_9); x_126 = lean_box(1); x_127 = l_Std_Format_joinSep___at_instReprProd___spec__1(x_125, x_126); lean_dec(x_125); -x_128 = l_Lean_Syntax_formatStxAux___closed__19; +x_128 = l_Lean_Syntax_formatStxAux___closed__21; x_129 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_129, 0, x_128); lean_ctor_set(x_129, 1, x_127); -x_130 = l_Lean_Syntax_formatStxAux___closed__21; +x_130 = l_Lean_Syntax_formatStxAux___closed__23; x_131 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_131, 0, x_129); lean_ctor_set(x_131, 1, x_130); -x_132 = l_Lean_Syntax_formatStxAux___closed__18; +x_132 = l_Lean_Syntax_formatStxAux___closed__20; x_133 = lean_alloc_ctor(3, 2, 0); lean_ctor_set(x_133, 0, x_132); lean_ctor_set(x_133, 1, x_131); @@ -1134,7 +1159,7 @@ else lean_object* x_136; lean_dec(x_9); lean_dec(x_7); -x_136 = l_Lean_Syntax_formatStxAux___closed__25; +x_136 = l_Lean_Syntax_formatStxAux___closed__27; return x_136; } } @@ -1168,7 +1193,7 @@ x_144 = 1; x_145 = l_Lean_Name_toString(x_143, x_144); x_146 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_146, 0, x_145); -x_147 = l_Lean_Syntax_formatStxAux___closed__27; +x_147 = l_Lean_Syntax_formatStxAux___closed__29; x_148 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_148, 0, x_147); lean_ctor_set(x_148, 1, x_146); @@ -1349,6 +1374,64 @@ x_1 = l_Lean_Syntax_instToStringSyntax___closed__3; return x_1; } } +LEAN_EXPORT lean_object* l_Lean_Syntax_instToFormatTSyntax___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; +x_2 = lean_box(0); +x_3 = 0; +x_4 = lean_unsigned_to_nat(0u); +x_5 = l_Lean_Syntax_formatStxAux(x_2, x_3, x_4, x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Syntax_instToFormatTSyntax(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Syntax_instToFormatTSyntax___rarg), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Syntax_instToFormatTSyntax___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Syntax_instToFormatTSyntax(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Syntax_instToStringTSyntax___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_2 = lean_box(0); +x_3 = 0; +x_4 = lean_unsigned_to_nat(0u); +x_5 = l_Lean_Syntax_formatStxAux(x_2, x_3, x_4, x_1); +x_6 = l_Std_Format_defWidth; +x_7 = lean_format_pretty(x_5, x_6); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Syntax_instToStringTSyntax(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Syntax_instToStringTSyntax___rarg), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Syntax_instToStringTSyntax___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Syntax_instToStringTSyntax(x_1); +lean_dec(x_1); +return x_2; +} +} lean_object* initialize_Init_Data_Format_Macro(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_Format_Instances(uint8_t builtin, lean_object*); lean_object* initialize_Init_Meta(uint8_t builtin, lean_object*); @@ -1428,6 +1511,10 @@ l_Lean_Syntax_formatStxAux___closed__26 = _init_l_Lean_Syntax_formatStxAux___clo lean_mark_persistent(l_Lean_Syntax_formatStxAux___closed__26); l_Lean_Syntax_formatStxAux___closed__27 = _init_l_Lean_Syntax_formatStxAux___closed__27(); lean_mark_persistent(l_Lean_Syntax_formatStxAux___closed__27); +l_Lean_Syntax_formatStxAux___closed__28 = _init_l_Lean_Syntax_formatStxAux___closed__28(); +lean_mark_persistent(l_Lean_Syntax_formatStxAux___closed__28); +l_Lean_Syntax_formatStxAux___closed__29 = _init_l_Lean_Syntax_formatStxAux___closed__29(); +lean_mark_persistent(l_Lean_Syntax_formatStxAux___closed__29); l_Lean_Syntax_instToStringSyntax___closed__1 = _init_l_Lean_Syntax_instToStringSyntax___closed__1(); lean_mark_persistent(l_Lean_Syntax_instToStringSyntax___closed__1); l_Lean_Syntax_instToStringSyntax___closed__2 = _init_l_Lean_Syntax_instToStringSyntax___closed__2(); diff --git a/stage0/stdlib/Init/Data/List/Basic.c b/stage0/stdlib/Init/Data/List/Basic.c index 5f2014ad844..3fa4056f841 100644 --- a/stage0/stdlib/Init/Data/List/Basic.c +++ b/stage0/stdlib/Init/Data/List/Basic.c @@ -46,6 +46,7 @@ LEAN_EXPORT lean_object* l_List_reverseAux(lean_object*); LEAN_EXPORT lean_object* l_List_replace___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_List_isSuffixOf___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_List_isPrefixOf___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_filterTRAux___at_List_removeAll___spec__1(lean_object*); LEAN_EXPORT lean_object* l_List_isPrefixOf___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_all___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_append___rarg(lean_object*, lean_object*); @@ -71,6 +72,7 @@ LEAN_EXPORT lean_object* l_List_groupBy___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Basic_0__List_concat_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_instMembershipList(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Basic_0__List_init_match__1_splitter(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_filterTRAux___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_join___rarg(lean_object*); LEAN_EXPORT lean_object* l_List_elem___at_List_instDecidableMemListInstMembershipList___spec__1(lean_object*); LEAN_EXPORT lean_object* l_max___at_List_maximum_x3f___spec__1(lean_object*, lean_object*); @@ -85,10 +87,10 @@ LEAN_EXPORT lean_object* l_List_or___boxed(lean_object*); LEAN_EXPORT lean_object* l_List_replicate___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_isEmpty___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Basic_0__List_beq_match__1_splitter___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_filterTR___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_any(lean_object*); LEAN_EXPORT lean_object* l_List_instLTList___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_bind___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_filterAux___at_List_removeAll___spec__1(lean_object*); LEAN_EXPORT lean_object* l_List_reverse(lean_object*); LEAN_EXPORT lean_object* l_List_minimum_x3f___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_minimum_x3f(lean_object*); @@ -104,7 +106,6 @@ LEAN_EXPORT lean_object* l_List_eraseIdx(lean_object*); LEAN_EXPORT lean_object* l_List_groupByAux___rarg(lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_List_all___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_filterAux___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_List_instDecidableMemListInstMembershipList___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_zip___rarg___lambda__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_List_foldr___at_List_or___spec__1(uint8_t, lean_object*); @@ -115,6 +116,7 @@ LEAN_EXPORT lean_object* l_List_init___rarg(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Basic_0__List_length_match__1_splitter___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_isSuffixOf(lean_object*); LEAN_EXPORT uint8_t l_List_or(lean_object*); +LEAN_EXPORT lean_object* l_List_filterTR(lean_object*); LEAN_EXPORT lean_object* l_List_replicate(lean_object*); LEAN_EXPORT lean_object* l_List_append(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Basic_0__List_beq_match__1_splitter(lean_object*, lean_object*); @@ -154,6 +156,8 @@ LEAN_EXPORT lean_object* l_List_removeAll(lean_object*); LEAN_EXPORT lean_object* l_List_replicateTR_loop(lean_object*); LEAN_EXPORT lean_object* l_List_enumFrom___rarg(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_List_foldr___at_List_and___spec__1(uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_List_filterTRAux___at_List_removeAll___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_filterTRAux(lean_object*); LEAN_EXPORT uint8_t l_List_hasDecidableLt___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Basic_0__List_replicateTR_loop_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_partitionAux(lean_object*); @@ -161,7 +165,6 @@ LEAN_EXPORT lean_object* l_List_isEmpty(lean_object*); LEAN_EXPORT lean_object* l_List_drop___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_take___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_span___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_filterAux___at_List_removeAll___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_erase(lean_object*); LEAN_EXPORT uint8_t l_List_beq___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_drop___rarg___boxed(lean_object*, lean_object*); @@ -206,7 +209,6 @@ LEAN_EXPORT lean_object* l_List_zip___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_replicate___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_and___boxed(lean_object*); LEAN_EXPORT lean_object* l_List_erase___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_filterAux(lean_object*); LEAN_EXPORT lean_object* l_List_foldr(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_iotaTR_go(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_List_elem___at_List_instDecidableMemListInstMembershipList___spec__1___rarg(lean_object*, lean_object*, lean_object*); @@ -1098,7 +1100,86 @@ x_3 = lean_alloc_closure((void*)(l_List_filterMap___rarg), 2, 0); return x_3; } } -LEAN_EXPORT lean_object* l_List_filterAux___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_filter___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; +lean_dec(x_1); +x_3 = lean_box(0); +return x_3; +} +else +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_2); +if (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); +x_6 = lean_ctor_get(x_2, 1); +lean_inc(x_1); +lean_inc(x_5); +x_7 = lean_apply_1(x_1, x_5); +x_8 = lean_unbox(x_7); +lean_dec(x_7); +if (x_8 == 0) +{ +lean_free_object(x_2); +lean_dec(x_5); +x_2 = x_6; +goto _start; +} +else +{ +lean_object* x_10; +x_10 = l_List_filter___rarg(x_1, x_6); +lean_ctor_set(x_2, 1, x_10); +return x_2; +} +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_ctor_get(x_2, 0); +x_12 = lean_ctor_get(x_2, 1); +lean_inc(x_12); +lean_inc(x_11); +lean_dec(x_2); +lean_inc(x_1); +lean_inc(x_11); +x_13 = lean_apply_1(x_1, x_11); +x_14 = lean_unbox(x_13); +lean_dec(x_13); +if (x_14 == 0) +{ +lean_dec(x_11); +x_2 = x_12; +goto _start; +} +else +{ +lean_object* x_16; lean_object* x_17; +x_16 = l_List_filter___rarg(x_1, x_12); +x_17 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_17, 0, x_11); +lean_ctor_set(x_17, 1, x_16); +return x_17; +} +} +} +} +} +LEAN_EXPORT lean_object* l_List_filter(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_List_filter___rarg), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_List_filterTRAux___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_2) == 0) @@ -1174,31 +1255,66 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_List_filterAux(lean_object* x_1) { +LEAN_EXPORT lean_object* l_List_filterTRAux(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_List_filterAux___rarg), 3, 0); +x_2 = lean_alloc_closure((void*)(l_List_filterTRAux___rarg), 3, 0); return x_2; } } -LEAN_EXPORT lean_object* l_List_filter___rarg(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_List_filterTR___rarg(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; x_3 = lean_box(0); -x_4 = l_List_filterAux___rarg(x_1, x_2, x_3); +x_4 = l_List_filterTRAux___rarg(x_1, x_2, x_3); return x_4; } } -LEAN_EXPORT lean_object* l_List_filter(lean_object* x_1) { +LEAN_EXPORT lean_object* l_List_filterTR(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_List_filter___rarg), 2, 0); +x_2 = lean_alloc_closure((void*)(l_List_filterTR___rarg), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_List_Basic_0__List_erase_match__1_splitter___rarg(uint8_t x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (x_1 == 0) +{ +lean_inc(x_3); +return x_3; +} +else +{ +lean_inc(x_2); return x_2; } } +} +LEAN_EXPORT lean_object* l___private_Init_Data_List_Basic_0__List_erase_match__1_splitter(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Init_Data_List_Basic_0__List_erase_match__1_splitter___rarg___boxed), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_List_Basic_0__List_erase_match__1_splitter___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = lean_unbox(x_1); +lean_dec(x_1); +x_5 = l___private_Init_Data_List_Basic_0__List_erase_match__1_splitter___rarg(x_4, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +return x_5; +} +} LEAN_EXPORT lean_object* l_List_partitionAux___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { @@ -1806,41 +1922,6 @@ x_2 = lean_box(0); return x_2; } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Basic_0__List_erase_match__1_splitter___rarg(uint8_t x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -if (x_1 == 0) -{ -lean_inc(x_3); -return x_3; -} -else -{ -lean_inc(x_2); -return x_2; -} -} -} -LEAN_EXPORT lean_object* l___private_Init_Data_List_Basic_0__List_erase_match__1_splitter(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Init_Data_List_Basic_0__List_erase_match__1_splitter___rarg___boxed), 3, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l___private_Init_Data_List_Basic_0__List_erase_match__1_splitter___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -uint8_t x_4; lean_object* x_5; -x_4 = lean_unbox(x_1); -lean_dec(x_1); -x_5 = l___private_Init_Data_List_Basic_0__List_erase_match__1_splitter___rarg(x_4, x_2, x_3); -lean_dec(x_3); -lean_dec(x_2); -return x_5; -} -} LEAN_EXPORT uint8_t l_List_elem___at_List_instDecidableMemListInstMembershipList___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { @@ -2699,7 +2780,7 @@ x_3 = lean_alloc_closure((void*)(l_List_lookup___rarg), 3, 0); return x_3; } } -LEAN_EXPORT lean_object* l_List_filterAux___at_List_removeAll___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_List_filterTRAux___at_List_removeAll___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { if (lean_obj_tag(x_3) == 0) @@ -2774,11 +2855,11 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_List_filterAux___at_List_removeAll___spec__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_List_filterTRAux___at_List_removeAll___spec__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_List_filterAux___at_List_removeAll___spec__1___rarg), 4, 0); +x_2 = lean_alloc_closure((void*)(l_List_filterTRAux___at_List_removeAll___spec__1___rarg), 4, 0); return x_2; } } @@ -2787,7 +2868,7 @@ LEAN_EXPORT lean_object* l_List_removeAll___rarg(lean_object* x_1, lean_object* { lean_object* x_4; lean_object* x_5; x_4 = lean_box(0); -x_5 = l_List_filterAux___at_List_removeAll___spec__1___rarg(x_1, x_3, x_2, x_4); +x_5 = l_List_filterTRAux___at_List_removeAll___spec__1___rarg(x_1, x_3, x_2, x_4); return x_5; } } diff --git a/stage0/stdlib/Init/Data/List/BasicAux.c b/stage0/stdlib/Init/Data/List/BasicAux.c index 9c5e4a00429..d55680616f4 100644 --- a/stage0/stdlib/Init/Data/List/BasicAux.c +++ b/stage0/stdlib/Init/Data/List/BasicAux.c @@ -15,6 +15,7 @@ extern "C" { #endif static lean_object* l_List_getLast_x21___rarg___closed__1; static lean_object* l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__7; +static lean_object* l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__71; static lean_object* l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__16; static lean_object* l_List___aux__Init__Data__List__BasicAux______macroRules__tacticDecreasing__trivial__1___closed__2; LEAN_EXPORT lean_object* l_List_tail_x21___rarg(lean_object*); @@ -154,6 +155,7 @@ lean_object* l_List_take___rarg(lean_object*, lean_object*); static lean_object* l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__54; static lean_object* l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__62; LEAN_EXPORT lean_object* l_List_head_x21(lean_object*); +static lean_object* l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__72; static lean_object* l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__47; static lean_object* l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__60; static lean_object* l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__41; @@ -1702,20 +1704,38 @@ static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRul _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("=>", 2); +x_1 = lean_mk_string_from_bytes("binderIdent", 11); return x_1; } } static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__66() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__2; +x_2 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__65; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__67() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("=>", 2); +return x_1; +} +} +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__68() { +_start: +{ lean_object* x_1; lean_object* x_2; x_1 = lean_unsigned_to_nat(5u); x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__67() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__69() { _start: { lean_object* x_1; @@ -1723,17 +1743,17 @@ x_1 = lean_mk_string_from_bytes("simpArith", 9); return x_1; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__68() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__70() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__6; -x_2 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__67; +x_2 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__69; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__69() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__71() { _start: { lean_object* x_1; @@ -1741,7 +1761,7 @@ x_1 = lean_mk_string_from_bytes("simp_arith", 10); return x_1; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__70() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__72() { _start: { lean_object* x_1; lean_object* x_2; @@ -1774,7 +1794,7 @@ x_8 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRule 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; 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; +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; 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; x_10 = lean_ctor_get(x_8, 0); x_11 = lean_ctor_get(x_2, 2); lean_inc(x_11); @@ -2000,465 +2020,477 @@ lean_inc(x_10); x_116 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_116, 0, x_10); lean_ctor_set(x_116, 1, x_115); -x_117 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__65; +x_117 = lean_array_push(x_33, x_84); +x_118 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__66; +x_119 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_119, 0, x_28); +lean_ctor_set(x_119, 1, x_118); +lean_ctor_set(x_119, 2, x_117); +x_120 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__67; lean_inc(x_10); -x_118 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_118, 0, x_10); -lean_ctor_set(x_118, 1, x_117); -x_119 = lean_array_push(x_46, x_55); -x_120 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_120, 0, x_28); -lean_ctor_set(x_120, 1, x_39); -lean_ctor_set(x_120, 2, x_119); -x_121 = lean_array_push(x_33, x_120); -x_122 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_122, 0, x_28); -lean_ctor_set(x_122, 1, x_35); -lean_ctor_set(x_122, 2, x_121); -x_123 = lean_array_push(x_33, x_122); -x_124 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_124, 0, x_28); -lean_ctor_set(x_124, 1, x_64); -lean_ctor_set(x_124, 2, x_123); -x_125 = lean_array_push(x_33, x_124); -x_126 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_126, 0, x_28); -lean_ctor_set(x_126, 1, x_67); -lean_ctor_set(x_126, 2, x_125); -x_127 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__66; -x_128 = lean_array_push(x_127, x_116); -x_129 = lean_array_push(x_128, x_84); -x_130 = lean_array_push(x_129, x_55); -x_131 = lean_array_push(x_130, x_118); -x_132 = lean_array_push(x_131, x_126); -x_133 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__64; -x_134 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_134, 0, x_28); -lean_ctor_set(x_134, 1, x_133); -lean_ctor_set(x_134, 2, x_132); -x_135 = lean_array_push(x_25, x_134); -x_136 = lean_array_push(x_135, x_55); +x_121 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_121, 0, x_10); +lean_ctor_set(x_121, 1, x_120); +x_122 = lean_array_push(x_46, x_55); +x_123 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_123, 0, x_28); +lean_ctor_set(x_123, 1, x_39); +lean_ctor_set(x_123, 2, x_122); +x_124 = lean_array_push(x_33, x_123); +x_125 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_125, 0, x_28); +lean_ctor_set(x_125, 1, x_35); +lean_ctor_set(x_125, 2, x_124); +x_126 = lean_array_push(x_33, x_125); +x_127 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_127, 0, x_28); +lean_ctor_set(x_127, 1, x_64); +lean_ctor_set(x_127, 2, x_126); +x_128 = lean_array_push(x_33, x_127); +x_129 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_129, 0, x_28); +lean_ctor_set(x_129, 1, x_67); +lean_ctor_set(x_129, 2, x_128); +x_130 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__68; +x_131 = lean_array_push(x_130, x_116); +x_132 = lean_array_push(x_131, x_119); +x_133 = lean_array_push(x_132, x_55); +x_134 = lean_array_push(x_133, x_121); +x_135 = lean_array_push(x_134, x_129); +x_136 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__64; x_137 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_137, 0, x_28); -lean_ctor_set(x_137, 1, x_39); -lean_ctor_set(x_137, 2, x_136); -x_138 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__69; -x_139 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_139, 0, x_10); -lean_ctor_set(x_139, 1, x_138); -x_140 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__70; -x_141 = lean_array_push(x_140, x_139); -x_142 = lean_array_push(x_141, x_55); -x_143 = lean_array_push(x_142, x_55); -x_144 = lean_array_push(x_143, x_55); +lean_ctor_set(x_137, 1, x_136); +lean_ctor_set(x_137, 2, x_135); +x_138 = lean_array_push(x_25, x_137); +x_139 = lean_array_push(x_138, x_55); +x_140 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_140, 0, x_28); +lean_ctor_set(x_140, 1, x_39); +lean_ctor_set(x_140, 2, x_139); +x_141 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__71; +x_142 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_142, 0, x_10); +lean_ctor_set(x_142, 1, x_141); +x_143 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__72; +x_144 = lean_array_push(x_143, x_142); x_145 = lean_array_push(x_144, x_55); x_146 = lean_array_push(x_145, x_55); -x_147 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__68; -x_148 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_148, 0, x_28); -lean_ctor_set(x_148, 1, x_147); -lean_ctor_set(x_148, 2, x_146); -x_149 = lean_array_push(x_25, x_148); -x_150 = lean_array_push(x_149, x_55); +x_147 = lean_array_push(x_146, x_55); +x_148 = lean_array_push(x_147, x_55); +x_149 = lean_array_push(x_148, x_55); +x_150 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__70; x_151 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_151, 0, x_28); -lean_ctor_set(x_151, 1, x_39); -lean_ctor_set(x_151, 2, x_150); -x_152 = lean_array_push(x_58, x_114); -x_153 = lean_array_push(x_152, x_137); -x_154 = lean_array_push(x_153, x_151); -x_155 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_155, 0, x_28); -lean_ctor_set(x_155, 1, x_35); -lean_ctor_set(x_155, 2, x_154); -x_156 = lean_array_push(x_33, x_155); -x_157 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_157, 0, x_28); -lean_ctor_set(x_157, 1, x_64); -lean_ctor_set(x_157, 2, x_156); -x_158 = lean_array_push(x_33, x_157); -x_159 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_159, 0, x_28); -lean_ctor_set(x_159, 1, x_67); -lean_ctor_set(x_159, 2, x_158); -x_160 = lean_array_push(x_69, x_159); -x_161 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_161, 0, x_28); -lean_ctor_set(x_161, 1, x_39); -lean_ctor_set(x_161, 2, x_160); -x_162 = lean_array_push(x_25, x_71); -x_163 = lean_array_push(x_162, x_161); +lean_ctor_set(x_151, 1, x_150); +lean_ctor_set(x_151, 2, x_149); +x_152 = lean_array_push(x_25, x_151); +x_153 = lean_array_push(x_152, x_55); +x_154 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_154, 0, x_28); +lean_ctor_set(x_154, 1, x_39); +lean_ctor_set(x_154, 2, x_153); +x_155 = lean_array_push(x_58, x_114); +x_156 = lean_array_push(x_155, x_140); +x_157 = lean_array_push(x_156, x_154); +x_158 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_158, 0, x_28); +lean_ctor_set(x_158, 1, x_35); +lean_ctor_set(x_158, 2, x_157); +x_159 = lean_array_push(x_33, x_158); +x_160 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_160, 0, x_28); +lean_ctor_set(x_160, 1, x_64); +lean_ctor_set(x_160, 2, x_159); +x_161 = lean_array_push(x_33, x_160); +x_162 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_162, 0, x_28); +lean_ctor_set(x_162, 1, x_67); +lean_ctor_set(x_162, 2, x_161); +x_163 = lean_array_push(x_69, x_162); x_164 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_164, 0, x_28); -lean_ctor_set(x_164, 1, x_35); +lean_ctor_set(x_164, 1, x_39); lean_ctor_set(x_164, 2, x_163); -x_165 = lean_array_push(x_25, x_14); +x_165 = lean_array_push(x_25, x_71); x_166 = lean_array_push(x_165, x_164); -x_167 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__12; -x_168 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_168, 0, x_28); -lean_ctor_set(x_168, 1, x_167); -lean_ctor_set(x_168, 2, x_166); -x_169 = lean_array_push(x_33, x_168); -x_170 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_170, 0, x_28); -lean_ctor_set(x_170, 1, x_35); -lean_ctor_set(x_170, 2, x_169); -x_171 = lean_array_push(x_33, x_170); -x_172 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__8; +x_167 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_167, 0, x_28); +lean_ctor_set(x_167, 1, x_35); +lean_ctor_set(x_167, 2, x_166); +x_168 = lean_array_push(x_25, x_14); +x_169 = lean_array_push(x_168, x_167); +x_170 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__12; +x_171 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_171, 0, x_28); +lean_ctor_set(x_171, 1, x_170); +lean_ctor_set(x_171, 2, x_169); +x_172 = lean_array_push(x_33, x_171); x_173 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_173, 0, x_28); -lean_ctor_set(x_173, 1, x_172); -lean_ctor_set(x_173, 2, x_171); -lean_ctor_set(x_8, 0, x_173); +lean_ctor_set(x_173, 1, x_35); +lean_ctor_set(x_173, 2, x_172); +x_174 = lean_array_push(x_33, x_173); +x_175 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__8; +x_176 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_176, 0, x_28); +lean_ctor_set(x_176, 1, x_175); +lean_ctor_set(x_176, 2, x_174); +lean_ctor_set(x_8, 0, x_176); return x_8; } else { -lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; -x_174 = lean_ctor_get(x_8, 0); -x_175 = lean_ctor_get(x_8, 1); -lean_inc(x_175); -lean_inc(x_174); -lean_dec(x_8); -x_176 = lean_ctor_get(x_2, 2); -lean_inc(x_176); -x_177 = lean_ctor_get(x_2, 1); +lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; +x_177 = lean_ctor_get(x_8, 0); +x_178 = lean_ctor_get(x_8, 1); +lean_inc(x_178); lean_inc(x_177); +lean_dec(x_8); +x_179 = lean_ctor_get(x_2, 2); +lean_inc(x_179); +x_180 = lean_ctor_get(x_2, 1); +lean_inc(x_180); lean_dec(x_2); -x_178 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__11; -lean_inc(x_174); -x_179 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_179, 0, x_174); -lean_ctor_set(x_179, 1, x_178); -x_180 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__15; -lean_inc(x_174); -x_181 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_181, 0, x_174); -lean_ctor_set(x_181, 1, x_180); -x_182 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__20; -lean_inc(x_174); -x_183 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_183, 0, x_174); -lean_ctor_set(x_183, 1, x_182); -x_184 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__25; -lean_inc(x_176); +x_181 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__11; +lean_inc(x_177); +x_182 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_182, 0, x_177); +lean_ctor_set(x_182, 1, x_181); +x_183 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__15; +lean_inc(x_177); +x_184 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_184, 0, x_177); +lean_ctor_set(x_184, 1, x_183); +x_185 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__20; +lean_inc(x_177); +x_186 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_186, 0, x_177); +lean_ctor_set(x_186, 1, x_185); +x_187 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__25; +lean_inc(x_179); +lean_inc(x_180); +x_188 = l_Lean_addMacroScope(x_180, x_187, x_179); +x_189 = lean_box(0); +x_190 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__24; +x_191 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__28; lean_inc(x_177); -x_185 = l_Lean_addMacroScope(x_177, x_184, x_176); -x_186 = lean_box(0); -x_187 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__24; -x_188 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__28; -lean_inc(x_174); -x_189 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_189, 0, x_174); -lean_ctor_set(x_189, 1, x_187); -lean_ctor_set(x_189, 2, x_185); -lean_ctor_set(x_189, 3, x_188); -x_190 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__29; -x_191 = lean_array_push(x_190, x_183); -lean_inc(x_189); -lean_inc(x_191); -x_192 = lean_array_push(x_191, x_189); -x_193 = lean_box(2); -x_194 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__21; -x_195 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_195, 0, x_193); -lean_ctor_set(x_195, 1, x_194); -lean_ctor_set(x_195, 2, x_192); -x_196 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__30; -lean_inc(x_174); -x_197 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_197, 0, x_174); -lean_ctor_set(x_197, 1, x_196); -x_198 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__31; -x_199 = lean_array_push(x_198, x_197); -x_200 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__10; -x_201 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_201, 0, x_193); -lean_ctor_set(x_201, 1, x_200); -lean_ctor_set(x_201, 2, x_199); -x_202 = lean_array_push(x_190, x_195); -lean_inc(x_201); -x_203 = lean_array_push(x_202, x_201); -x_204 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__14; -x_205 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_205, 0, x_193); -lean_ctor_set(x_205, 1, x_204); -lean_ctor_set(x_205, 2, x_203); -x_206 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__32; -lean_inc(x_174); -x_207 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_207, 0, x_174); -lean_ctor_set(x_207, 1, x_206); -x_208 = lean_array_push(x_198, x_207); -x_209 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__33; -x_210 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_210, 0, x_193); +x_192 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_192, 0, x_177); +lean_ctor_set(x_192, 1, x_190); +lean_ctor_set(x_192, 2, x_188); +lean_ctor_set(x_192, 3, x_191); +x_193 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__29; +x_194 = lean_array_push(x_193, x_186); +lean_inc(x_192); +lean_inc(x_194); +x_195 = lean_array_push(x_194, x_192); +x_196 = lean_box(2); +x_197 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__21; +x_198 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_198, 0, x_196); +lean_ctor_set(x_198, 1, x_197); +lean_ctor_set(x_198, 2, x_195); +x_199 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__30; +lean_inc(x_177); +x_200 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_200, 0, x_177); +lean_ctor_set(x_200, 1, x_199); +x_201 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__31; +x_202 = lean_array_push(x_201, x_200); +x_203 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__10; +x_204 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_204, 0, x_196); +lean_ctor_set(x_204, 1, x_203); +lean_ctor_set(x_204, 2, x_202); +x_205 = lean_array_push(x_193, x_198); +lean_inc(x_204); +x_206 = lean_array_push(x_205, x_204); +x_207 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__14; +x_208 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_208, 0, x_196); +lean_ctor_set(x_208, 1, x_207); +lean_ctor_set(x_208, 2, x_206); +x_209 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__32; +lean_inc(x_177); +x_210 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_210, 0, x_177); lean_ctor_set(x_210, 1, x_209); -lean_ctor_set(x_210, 2, x_208); -x_211 = lean_array_push(x_190, x_210); -lean_inc(x_211); -x_212 = lean_array_push(x_211, x_201); +x_211 = lean_array_push(x_201, x_210); +x_212 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__33; x_213 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_213, 0, x_193); -lean_ctor_set(x_213, 1, x_204); -lean_ctor_set(x_213, 2, x_212); -x_214 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__34; -lean_inc(x_174); -x_215 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_215, 0, x_174); -lean_ctor_set(x_215, 1, x_214); -x_216 = lean_array_push(x_198, x_215); -x_217 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__35; -x_218 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_218, 0, x_193); +lean_ctor_set(x_213, 0, x_196); +lean_ctor_set(x_213, 1, x_212); +lean_ctor_set(x_213, 2, x_211); +x_214 = lean_array_push(x_193, x_213); +lean_inc(x_214); +x_215 = lean_array_push(x_214, x_204); +x_216 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_216, 0, x_196); +lean_ctor_set(x_216, 1, x_207); +lean_ctor_set(x_216, 2, x_215); +x_217 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__34; +lean_inc(x_177); +x_218 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_218, 0, x_177); lean_ctor_set(x_218, 1, x_217); -lean_ctor_set(x_218, 2, x_216); -x_219 = lean_array_push(x_190, x_218); -x_220 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__37; -x_221 = lean_array_push(x_219, x_220); -x_222 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_222, 0, x_193); -lean_ctor_set(x_222, 1, x_204); -lean_ctor_set(x_222, 2, x_221); -x_223 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__38; -x_224 = lean_array_push(x_223, x_205); -x_225 = lean_array_push(x_224, x_213); -x_226 = lean_array_push(x_225, x_222); -x_227 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_227, 0, x_193); -lean_ctor_set(x_227, 1, x_200); -lean_ctor_set(x_227, 2, x_226); -x_228 = lean_array_push(x_198, x_227); -x_229 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__19; +x_219 = lean_array_push(x_201, x_218); +x_220 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__35; +x_221 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_221, 0, x_196); +lean_ctor_set(x_221, 1, x_220); +lean_ctor_set(x_221, 2, x_219); +x_222 = lean_array_push(x_193, x_221); +x_223 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__37; +x_224 = lean_array_push(x_222, x_223); +x_225 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_225, 0, x_196); +lean_ctor_set(x_225, 1, x_207); +lean_ctor_set(x_225, 2, x_224); +x_226 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__38; +x_227 = lean_array_push(x_226, x_208); +x_228 = lean_array_push(x_227, x_216); +x_229 = lean_array_push(x_228, x_225); x_230 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_230, 0, x_193); -lean_ctor_set(x_230, 1, x_229); -lean_ctor_set(x_230, 2, x_228); -x_231 = lean_array_push(x_198, x_230); -x_232 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__17; +lean_ctor_set(x_230, 0, x_196); +lean_ctor_set(x_230, 1, x_203); +lean_ctor_set(x_230, 2, x_229); +x_231 = lean_array_push(x_201, x_230); +x_232 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__19; x_233 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_233, 0, x_193); +lean_ctor_set(x_233, 0, x_196); lean_ctor_set(x_233, 1, x_232); lean_ctor_set(x_233, 2, x_231); -x_234 = lean_array_push(x_190, x_181); -lean_inc(x_234); -x_235 = lean_array_push(x_234, x_233); +x_234 = lean_array_push(x_201, x_233); +x_235 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__17; x_236 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_236, 0, x_193); -lean_ctor_set(x_236, 1, x_204); -lean_ctor_set(x_236, 2, x_235); -x_237 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__49; -lean_inc(x_176); +lean_ctor_set(x_236, 0, x_196); +lean_ctor_set(x_236, 1, x_235); +lean_ctor_set(x_236, 2, x_234); +x_237 = lean_array_push(x_193, x_184); +lean_inc(x_237); +x_238 = lean_array_push(x_237, x_236); +x_239 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_239, 0, x_196); +lean_ctor_set(x_239, 1, x_207); +lean_ctor_set(x_239, 2, x_238); +x_240 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__49; +lean_inc(x_179); +lean_inc(x_180); +x_241 = l_Lean_addMacroScope(x_180, x_240, x_179); +x_242 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__45; +x_243 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__51; +lean_inc(x_177); +x_244 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_244, 0, x_177); +lean_ctor_set(x_244, 1, x_242); +lean_ctor_set(x_244, 2, x_241); +lean_ctor_set(x_244, 3, x_243); +x_245 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__54; +lean_inc(x_177); +x_246 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_246, 0, x_177); +lean_ctor_set(x_246, 1, x_245); +x_247 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__57; lean_inc(x_177); -x_238 = l_Lean_addMacroScope(x_177, x_237, x_176); -x_239 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__45; -x_240 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__51; -lean_inc(x_174); -x_241 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_241, 0, x_174); -lean_ctor_set(x_241, 1, x_239); -lean_ctor_set(x_241, 2, x_238); -lean_ctor_set(x_241, 3, x_240); -x_242 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__54; -lean_inc(x_174); -x_243 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_243, 0, x_174); -lean_ctor_set(x_243, 1, x_242); -x_244 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__57; -lean_inc(x_174); -x_245 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_245, 0, x_174); -lean_ctor_set(x_245, 1, x_244); -x_246 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__61; -x_247 = l_Lean_addMacroScope(x_177, x_246, x_176); -x_248 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__60; -lean_inc(x_174); -x_249 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_249, 0, x_174); -lean_ctor_set(x_249, 1, x_248); -lean_ctor_set(x_249, 2, x_247); -lean_ctor_set(x_249, 3, x_186); -x_250 = lean_array_push(x_190, x_245); -lean_inc(x_249); -x_251 = lean_array_push(x_250, x_249); -x_252 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__56; -x_253 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_253, 0, x_193); -lean_ctor_set(x_253, 1, x_252); -lean_ctor_set(x_253, 2, x_251); -x_254 = lean_array_push(x_198, x_253); -x_255 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_255, 0, x_193); -lean_ctor_set(x_255, 1, x_200); -lean_ctor_set(x_255, 2, x_254); -x_256 = lean_array_push(x_190, x_189); -x_257 = lean_array_push(x_256, x_255); -x_258 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__42; -x_259 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_259, 0, x_193); -lean_ctor_set(x_259, 1, x_258); -lean_ctor_set(x_259, 2, x_257); -x_260 = lean_array_push(x_190, x_259); -x_261 = lean_array_push(x_260, x_220); +x_248 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_248, 0, x_177); +lean_ctor_set(x_248, 1, x_247); +x_249 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__61; +x_250 = l_Lean_addMacroScope(x_180, x_249, x_179); +x_251 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__60; +lean_inc(x_177); +x_252 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_252, 0, x_177); +lean_ctor_set(x_252, 1, x_251); +lean_ctor_set(x_252, 2, x_250); +lean_ctor_set(x_252, 3, x_189); +x_253 = lean_array_push(x_193, x_248); +lean_inc(x_252); +x_254 = lean_array_push(x_253, x_252); +x_255 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__56; +x_256 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_256, 0, x_196); +lean_ctor_set(x_256, 1, x_255); +lean_ctor_set(x_256, 2, x_254); +x_257 = lean_array_push(x_201, x_256); +x_258 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_258, 0, x_196); +lean_ctor_set(x_258, 1, x_203); +lean_ctor_set(x_258, 2, x_257); +x_259 = lean_array_push(x_193, x_192); +x_260 = lean_array_push(x_259, x_258); +x_261 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__42; x_262 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_262, 0, x_193); -lean_ctor_set(x_262, 1, x_200); -lean_ctor_set(x_262, 2, x_261); -x_263 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__62; -lean_inc(x_174); -x_264 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_264, 0, x_174); -lean_ctor_set(x_264, 1, x_263); -x_265 = lean_array_push(x_223, x_243); -x_266 = lean_array_push(x_265, x_262); -x_267 = lean_array_push(x_266, x_264); -x_268 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__53; -x_269 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_269, 0, x_193); -lean_ctor_set(x_269, 1, x_268); -lean_ctor_set(x_269, 2, x_267); -x_270 = lean_array_push(x_198, x_269); -x_271 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_271, 0, x_193); -lean_ctor_set(x_271, 1, x_200); -lean_ctor_set(x_271, 2, x_270); -x_272 = lean_array_push(x_190, x_241); -x_273 = lean_array_push(x_272, x_271); +lean_ctor_set(x_262, 0, x_196); +lean_ctor_set(x_262, 1, x_261); +lean_ctor_set(x_262, 2, x_260); +x_263 = lean_array_push(x_193, x_262); +x_264 = lean_array_push(x_263, x_223); +x_265 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_265, 0, x_196); +lean_ctor_set(x_265, 1, x_203); +lean_ctor_set(x_265, 2, x_264); +x_266 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__62; +lean_inc(x_177); +x_267 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_267, 0, x_177); +lean_ctor_set(x_267, 1, x_266); +x_268 = lean_array_push(x_226, x_246); +x_269 = lean_array_push(x_268, x_265); +x_270 = lean_array_push(x_269, x_267); +x_271 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__53; +x_272 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_272, 0, x_196); +lean_ctor_set(x_272, 1, x_271); +lean_ctor_set(x_272, 2, x_270); +x_273 = lean_array_push(x_201, x_272); x_274 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_274, 0, x_193); -lean_ctor_set(x_274, 1, x_258); +lean_ctor_set(x_274, 0, x_196); +lean_ctor_set(x_274, 1, x_203); lean_ctor_set(x_274, 2, x_273); -x_275 = lean_array_push(x_191, x_274); -x_276 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_276, 0, x_193); -lean_ctor_set(x_276, 1, x_194); -lean_ctor_set(x_276, 2, x_275); -x_277 = lean_array_push(x_190, x_276); -x_278 = lean_array_push(x_277, x_220); +x_275 = lean_array_push(x_193, x_244); +x_276 = lean_array_push(x_275, x_274); +x_277 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_277, 0, x_196); +lean_ctor_set(x_277, 1, x_261); +lean_ctor_set(x_277, 2, x_276); +x_278 = lean_array_push(x_194, x_277); x_279 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_279, 0, x_193); -lean_ctor_set(x_279, 1, x_204); +lean_ctor_set(x_279, 0, x_196); +lean_ctor_set(x_279, 1, x_197); lean_ctor_set(x_279, 2, x_278); -x_280 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__63; -lean_inc(x_174); -x_281 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_281, 0, x_174); -lean_ctor_set(x_281, 1, x_280); -x_282 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__65; -lean_inc(x_174); -x_283 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_283, 0, x_174); -lean_ctor_set(x_283, 1, x_282); -x_284 = lean_array_push(x_211, x_220); -x_285 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_285, 0, x_193); -lean_ctor_set(x_285, 1, x_204); -lean_ctor_set(x_285, 2, x_284); -x_286 = lean_array_push(x_198, x_285); +x_280 = lean_array_push(x_193, x_279); +x_281 = lean_array_push(x_280, x_223); +x_282 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_282, 0, x_196); +lean_ctor_set(x_282, 1, x_207); +lean_ctor_set(x_282, 2, x_281); +x_283 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__63; +lean_inc(x_177); +x_284 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_284, 0, x_177); +lean_ctor_set(x_284, 1, x_283); +x_285 = lean_array_push(x_201, x_252); +x_286 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__66; x_287 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_287, 0, x_193); -lean_ctor_set(x_287, 1, x_200); -lean_ctor_set(x_287, 2, x_286); -x_288 = lean_array_push(x_198, x_287); -x_289 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_289, 0, x_193); -lean_ctor_set(x_289, 1, x_229); -lean_ctor_set(x_289, 2, x_288); -x_290 = lean_array_push(x_198, x_289); +lean_ctor_set(x_287, 0, x_196); +lean_ctor_set(x_287, 1, x_286); +lean_ctor_set(x_287, 2, x_285); +x_288 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__67; +lean_inc(x_177); +x_289 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_289, 0, x_177); +lean_ctor_set(x_289, 1, x_288); +x_290 = lean_array_push(x_214, x_223); x_291 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_291, 0, x_193); -lean_ctor_set(x_291, 1, x_232); +lean_ctor_set(x_291, 0, x_196); +lean_ctor_set(x_291, 1, x_207); lean_ctor_set(x_291, 2, x_290); -x_292 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__66; -x_293 = lean_array_push(x_292, x_281); -x_294 = lean_array_push(x_293, x_249); -x_295 = lean_array_push(x_294, x_220); -x_296 = lean_array_push(x_295, x_283); -x_297 = lean_array_push(x_296, x_291); -x_298 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__64; -x_299 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_299, 0, x_193); -lean_ctor_set(x_299, 1, x_298); -lean_ctor_set(x_299, 2, x_297); -x_300 = lean_array_push(x_190, x_299); -x_301 = lean_array_push(x_300, x_220); -x_302 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_302, 0, x_193); -lean_ctor_set(x_302, 1, x_204); -lean_ctor_set(x_302, 2, x_301); -x_303 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__69; -x_304 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_304, 0, x_174); -lean_ctor_set(x_304, 1, x_303); -x_305 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__70; -x_306 = lean_array_push(x_305, x_304); -x_307 = lean_array_push(x_306, x_220); -x_308 = lean_array_push(x_307, x_220); -x_309 = lean_array_push(x_308, x_220); -x_310 = lean_array_push(x_309, x_220); -x_311 = lean_array_push(x_310, x_220); -x_312 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__68; -x_313 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_313, 0, x_193); -lean_ctor_set(x_313, 1, x_312); -lean_ctor_set(x_313, 2, x_311); -x_314 = lean_array_push(x_190, x_313); -x_315 = lean_array_push(x_314, x_220); -x_316 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_316, 0, x_193); -lean_ctor_set(x_316, 1, x_204); -lean_ctor_set(x_316, 2, x_315); -x_317 = lean_array_push(x_223, x_279); -x_318 = lean_array_push(x_317, x_302); -x_319 = lean_array_push(x_318, x_316); -x_320 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_320, 0, x_193); -lean_ctor_set(x_320, 1, x_200); -lean_ctor_set(x_320, 2, x_319); -x_321 = lean_array_push(x_198, x_320); +x_292 = lean_array_push(x_201, x_291); +x_293 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_293, 0, x_196); +lean_ctor_set(x_293, 1, x_203); +lean_ctor_set(x_293, 2, x_292); +x_294 = lean_array_push(x_201, x_293); +x_295 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_295, 0, x_196); +lean_ctor_set(x_295, 1, x_232); +lean_ctor_set(x_295, 2, x_294); +x_296 = lean_array_push(x_201, x_295); +x_297 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_297, 0, x_196); +lean_ctor_set(x_297, 1, x_235); +lean_ctor_set(x_297, 2, x_296); +x_298 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__68; +x_299 = lean_array_push(x_298, x_284); +x_300 = lean_array_push(x_299, x_287); +x_301 = lean_array_push(x_300, x_223); +x_302 = lean_array_push(x_301, x_289); +x_303 = lean_array_push(x_302, x_297); +x_304 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__64; +x_305 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_305, 0, x_196); +lean_ctor_set(x_305, 1, x_304); +lean_ctor_set(x_305, 2, x_303); +x_306 = lean_array_push(x_193, x_305); +x_307 = lean_array_push(x_306, x_223); +x_308 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_308, 0, x_196); +lean_ctor_set(x_308, 1, x_207); +lean_ctor_set(x_308, 2, x_307); +x_309 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__71; +x_310 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_310, 0, x_177); +lean_ctor_set(x_310, 1, x_309); +x_311 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__72; +x_312 = lean_array_push(x_311, x_310); +x_313 = lean_array_push(x_312, x_223); +x_314 = lean_array_push(x_313, x_223); +x_315 = lean_array_push(x_314, x_223); +x_316 = lean_array_push(x_315, x_223); +x_317 = lean_array_push(x_316, x_223); +x_318 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__70; +x_319 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_319, 0, x_196); +lean_ctor_set(x_319, 1, x_318); +lean_ctor_set(x_319, 2, x_317); +x_320 = lean_array_push(x_193, x_319); +x_321 = lean_array_push(x_320, x_223); x_322 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_322, 0, x_193); -lean_ctor_set(x_322, 1, x_229); +lean_ctor_set(x_322, 0, x_196); +lean_ctor_set(x_322, 1, x_207); lean_ctor_set(x_322, 2, x_321); -x_323 = lean_array_push(x_198, x_322); -x_324 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_324, 0, x_193); -lean_ctor_set(x_324, 1, x_232); -lean_ctor_set(x_324, 2, x_323); -x_325 = lean_array_push(x_234, x_324); +x_323 = lean_array_push(x_226, x_282); +x_324 = lean_array_push(x_323, x_308); +x_325 = lean_array_push(x_324, x_322); x_326 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_326, 0, x_193); -lean_ctor_set(x_326, 1, x_204); +lean_ctor_set(x_326, 0, x_196); +lean_ctor_set(x_326, 1, x_203); lean_ctor_set(x_326, 2, x_325); -x_327 = lean_array_push(x_190, x_236); -x_328 = lean_array_push(x_327, x_326); -x_329 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_329, 0, x_193); -lean_ctor_set(x_329, 1, x_200); -lean_ctor_set(x_329, 2, x_328); -x_330 = lean_array_push(x_190, x_179); -x_331 = lean_array_push(x_330, x_329); -x_332 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__12; -x_333 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_333, 0, x_193); -lean_ctor_set(x_333, 1, x_332); -lean_ctor_set(x_333, 2, x_331); -x_334 = lean_array_push(x_198, x_333); +x_327 = lean_array_push(x_201, x_326); +x_328 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_328, 0, x_196); +lean_ctor_set(x_328, 1, x_232); +lean_ctor_set(x_328, 2, x_327); +x_329 = lean_array_push(x_201, x_328); +x_330 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_330, 0, x_196); +lean_ctor_set(x_330, 1, x_235); +lean_ctor_set(x_330, 2, x_329); +x_331 = lean_array_push(x_237, x_330); +x_332 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_332, 0, x_196); +lean_ctor_set(x_332, 1, x_207); +lean_ctor_set(x_332, 2, x_331); +x_333 = lean_array_push(x_193, x_239); +x_334 = lean_array_push(x_333, x_332); x_335 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_335, 0, x_193); -lean_ctor_set(x_335, 1, x_200); +lean_ctor_set(x_335, 0, x_196); +lean_ctor_set(x_335, 1, x_203); lean_ctor_set(x_335, 2, x_334); -x_336 = lean_array_push(x_198, x_335); -x_337 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__8; -x_338 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_338, 0, x_193); -lean_ctor_set(x_338, 1, x_337); -lean_ctor_set(x_338, 2, x_336); -x_339 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_339, 0, x_338); -lean_ctor_set(x_339, 1, x_175); -return x_339; +x_336 = lean_array_push(x_193, x_182); +x_337 = lean_array_push(x_336, x_335); +x_338 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__12; +x_339 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_339, 0, x_196); +lean_ctor_set(x_339, 1, x_338); +lean_ctor_set(x_339, 2, x_337); +x_340 = lean_array_push(x_201, x_339); +x_341 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_341, 0, x_196); +lean_ctor_set(x_341, 1, x_203); +lean_ctor_set(x_341, 2, x_340); +x_342 = lean_array_push(x_201, x_341); +x_343 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__8; +x_344 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_344, 0, x_196); +lean_ctor_set(x_344, 1, x_343); +lean_ctor_set(x_344, 2, x_342); +x_345 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_345, 0, x_344); +lean_ctor_set(x_345, 1, x_178); +return x_345; } } } @@ -2762,6 +2794,10 @@ l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__li lean_mark_persistent(l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__69); l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__70 = _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__70(); lean_mark_persistent(l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__70); +l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__71 = _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__71(); +lean_mark_persistent(l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__71); +l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__72 = _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__72(); +lean_mark_persistent(l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__72); l_List___aux__Init__Data__List__BasicAux______macroRules__tacticDecreasing__trivial__1___closed__1 = _init_l_List___aux__Init__Data__List__BasicAux______macroRules__tacticDecreasing__trivial__1___closed__1(); lean_mark_persistent(l_List___aux__Init__Data__List__BasicAux______macroRules__tacticDecreasing__trivial__1___closed__1); l_List___aux__Init__Data__List__BasicAux______macroRules__tacticDecreasing__trivial__1___closed__2 = _init_l_List___aux__Init__Data__List__BasicAux______macroRules__tacticDecreasing__trivial__1___closed__2(); diff --git a/stage0/stdlib/Init/Data/OfScientific.c b/stage0/stdlib/Init/Data/OfScientific.c index 7c1c30eb55e..517f4edcb99 100644 --- a/stage0/stdlib/Init/Data/OfScientific.c +++ b/stage0/stdlib/Init/Data/OfScientific.c @@ -15,7 +15,7 @@ extern "C" { #endif LEAN_EXPORT lean_object* l_Float_ofBinaryScientific___boxed(lean_object*, lean_object*); lean_object* lean_nat_div(lean_object*, lean_object*); -static lean_object* l_instOfScientificFloat___closed__2; +LEAN_EXPORT double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); lean_object* lean_nat_log2(lean_object*); LEAN_EXPORT double lean_float_of_nat(lean_object*); LEAN_EXPORT double l_Nat_toFloat(lean_object*); @@ -23,19 +23,21 @@ static lean_object* l_instOfScientificFloat___closed__1; double lean_uint64_to_float(uint64_t); LEAN_EXPORT double l_instOfNatFloat(lean_object*); lean_object* lean_nat_pow(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_instOfScientificFloat___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT double l_Float_ofBinaryScientific(lean_object*, lean_object*); double lean_float_scaleb(double, lean_object*); lean_object* lean_int_mul(lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Float_ofNat___boxed(lean_object*); -static lean_object* l_Float_ofNat___closed__1; +static lean_object* l_Float_ofScientific___closed__1; +LEAN_EXPORT lean_object* l_Float_ofScientific___boxed(lean_object*, lean_object*, lean_object*); uint64_t lean_uint64_of_nat(lean_object*); -LEAN_EXPORT double l_instOfScientificFloat(lean_object*, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_instOfScientificFloat; +static lean_object* l_Float_ofInt___closed__1; LEAN_EXPORT lean_object* l_Float_ofInt___boxed(lean_object*); lean_object* lean_int_neg(lean_object*); LEAN_EXPORT lean_object* l_Nat_toFloat___boxed(lean_object*); +static lean_object* l_Float_ofScientific___closed__2; uint8_t lean_int_dec_lt(lean_object*, lean_object*); double lean_float_negate(double); lean_object* lean_nat_shiftl(lean_object*, lean_object*); @@ -78,7 +80,7 @@ x_4 = lean_box_float(x_3); return x_4; } } -static lean_object* _init_l_instOfScientificFloat___closed__1() { +static lean_object* _init_l_Float_ofScientific___closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -87,16 +89,16 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l_instOfScientificFloat___closed__2() { +static lean_object* _init_l_Float_ofScientific___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_instOfScientificFloat___closed__1; +x_1 = l_Float_ofScientific___closed__1; x_2 = lean_int_neg(x_1); return x_2; } } -LEAN_EXPORT double l_instOfScientificFloat(lean_object* x_1, uint8_t x_2, lean_object* x_3) { +LEAN_EXPORT double l_Float_ofScientific(lean_object* x_1, uint8_t x_2, lean_object* x_3) { _start: { if (x_2 == 0) @@ -131,7 +133,7 @@ x_18 = lean_nat_div(x_15, x_17); lean_dec(x_17); lean_dec(x_15); x_19 = lean_nat_to_int(x_3); -x_20 = l_instOfScientificFloat___closed__2; +x_20 = l_Float_ofScientific___closed__2; x_21 = lean_int_mul(x_20, x_19); lean_dec(x_19); x_22 = lean_nat_to_int(x_11); @@ -145,38 +147,43 @@ return x_24; } } } -LEAN_EXPORT lean_object* l_instOfScientificFloat___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Float_ofScientific___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { uint8_t x_4; double x_5; lean_object* x_6; x_4 = lean_unbox(x_2); lean_dec(x_2); -x_5 = l_instOfScientificFloat(x_1, x_4, x_3); +x_5 = l_Float_ofScientific(x_1, x_4, x_3); lean_dec(x_1); x_6 = lean_box_float(x_5); return x_6; } } -static lean_object* _init_l_Float_ofNat___closed__1() { +static lean_object* _init_l_instOfScientificFloat___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(0u); -x_2 = lean_nat_to_int(x_1); -return x_2; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Float_ofScientific___boxed), 3, 0); +return x_1; +} +} +static lean_object* _init_l_instOfScientificFloat() { +_start: +{ +lean_object* x_1; +x_1 = l_instOfScientificFloat___closed__1; +return x_1; } } LEAN_EXPORT double lean_float_of_nat(lean_object* x_1) { _start: { -lean_object* x_2; lean_object* x_3; lean_object* x_4; double x_5; -x_2 = lean_unsigned_to_nat(1u); -x_3 = lean_nat_mul(x_1, x_2); +uint8_t x_2; lean_object* x_3; double x_4; +x_2 = 0; +x_3 = lean_unsigned_to_nat(0u); +x_4 = l_Float_ofScientific(x_1, x_2, x_3); lean_dec(x_1); -x_4 = l_Float_ofNat___closed__1; -x_5 = l_Float_ofBinaryScientific(x_3, x_4); -lean_dec(x_3); -return x_5; +return x_4; } } LEAN_EXPORT lean_object* l_Float_ofNat___boxed(lean_object* x_1) { @@ -188,31 +195,46 @@ x_3 = lean_box_float(x_2); return x_3; } } +static lean_object* _init_l_Float_ofInt___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} LEAN_EXPORT double l_Float_ofInt(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; -x_2 = l_Float_ofNat___closed__1; +x_2 = l_Float_ofInt___closed__1; x_3 = lean_int_dec_lt(x_1, x_2); if (x_3 == 0) { -lean_object* x_4; double x_5; +lean_object* x_4; uint8_t x_5; lean_object* x_6; double x_7; x_4 = lean_nat_abs(x_1); -x_5 = lean_float_of_nat(x_4); -return x_5; +x_5 = 0; +x_6 = lean_unsigned_to_nat(0u); +x_7 = l_Float_ofScientific(x_4, x_5, x_6); +lean_dec(x_4); +return x_7; } else { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; double x_10; double x_11; -x_6 = lean_nat_abs(x_1); -x_7 = lean_unsigned_to_nat(1u); -x_8 = lean_nat_sub(x_6, x_7); -lean_dec(x_6); -x_9 = lean_nat_add(x_8, x_7); +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; double x_14; double x_15; +x_8 = lean_nat_abs(x_1); +x_9 = lean_unsigned_to_nat(1u); +x_10 = lean_nat_sub(x_8, x_9); lean_dec(x_8); -x_10 = lean_float_of_nat(x_9); -x_11 = lean_float_negate(x_10); -return x_11; +x_11 = lean_nat_add(x_10, x_9); +lean_dec(x_10); +x_12 = 0; +x_13 = lean_unsigned_to_nat(0u); +x_14 = l_Float_ofScientific(x_11, x_12, x_13); +lean_dec(x_11); +x_15 = lean_float_negate(x_14); +return x_15; } } } @@ -229,9 +251,11 @@ return x_3; LEAN_EXPORT double l_instOfNatFloat(lean_object* x_1) { _start: { -double x_2; -x_2 = lean_float_of_nat(x_1); -return x_2; +uint8_t x_2; lean_object* x_3; double x_4; +x_2 = 0; +x_3 = lean_unsigned_to_nat(0u); +x_4 = l_Float_ofScientific(x_1, x_2, x_3); +return x_4; } } LEAN_EXPORT lean_object* l_instOfNatFloat___boxed(lean_object* x_1) { @@ -239,6 +263,7 @@ LEAN_EXPORT lean_object* l_instOfNatFloat___boxed(lean_object* x_1) { { double x_2; lean_object* x_3; x_2 = l_instOfNatFloat(x_1); +lean_dec(x_1); x_3 = lean_box_float(x_2); return x_3; } @@ -246,9 +271,11 @@ return x_3; LEAN_EXPORT double l_Nat_toFloat(lean_object* x_1) { _start: { -double x_2; -x_2 = lean_float_of_nat(x_1); -return x_2; +uint8_t x_2; lean_object* x_3; double x_4; +x_2 = 0; +x_3 = lean_unsigned_to_nat(0u); +x_4 = l_Float_ofScientific(x_1, x_2, x_3); +return x_4; } } LEAN_EXPORT lean_object* l_Nat_toFloat___boxed(lean_object* x_1) { @@ -256,6 +283,7 @@ LEAN_EXPORT lean_object* l_Nat_toFloat___boxed(lean_object* x_1) { { double x_2; lean_object* x_3; x_2 = l_Nat_toFloat(x_1); +lean_dec(x_1); x_3 = lean_box_float(x_2); return x_3; } @@ -277,12 +305,16 @@ lean_dec_ref(res); res = initialize_Init_Data_Nat(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +l_Float_ofScientific___closed__1 = _init_l_Float_ofScientific___closed__1(); +lean_mark_persistent(l_Float_ofScientific___closed__1); +l_Float_ofScientific___closed__2 = _init_l_Float_ofScientific___closed__2(); +lean_mark_persistent(l_Float_ofScientific___closed__2); l_instOfScientificFloat___closed__1 = _init_l_instOfScientificFloat___closed__1(); lean_mark_persistent(l_instOfScientificFloat___closed__1); -l_instOfScientificFloat___closed__2 = _init_l_instOfScientificFloat___closed__2(); -lean_mark_persistent(l_instOfScientificFloat___closed__2); -l_Float_ofNat___closed__1 = _init_l_Float_ofNat___closed__1(); -lean_mark_persistent(l_Float_ofNat___closed__1); +l_instOfScientificFloat = _init_l_instOfScientificFloat(); +lean_mark_persistent(l_instOfScientificFloat); +l_Float_ofInt___closed__1 = _init_l_Float_ofInt___closed__1(); +lean_mark_persistent(l_Float_ofInt___closed__1); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Init/Data/Option/Basic.c b/stage0/stdlib/Init/Data/Option/Basic.c index f9a2e552c39..5f17ef001a0 100644 --- a/stage0/stdlib/Init/Data/Option/Basic.c +++ b/stage0/stdlib/Init/Data/Option/Basic.c @@ -23,6 +23,7 @@ LEAN_EXPORT lean_object* l_instMonadOption___lambda__4(lean_object*, lean_object LEAN_EXPORT lean_object* l_Option_bind(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_instMonadOption___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Option_orElse(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__decEqOption____x40_Init_Data_Option_Basic___hyg_661____at_instDecidableEqOption___spec__1___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_instAlternativeOption; LEAN_EXPORT lean_object* l_Option_tryCatch___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Option_tryCatch___rarg(lean_object*, lean_object*); @@ -32,13 +33,14 @@ LEAN_EXPORT lean_object* l_instLTOption(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Option_filter___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Option_bind___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_instMonadExceptOfUnitOption___lambda__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__decEqOption____x40_Init_Data_Option_Basic___hyg_661____rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Option_toBool(lean_object*); static lean_object* l_instMonadOption___closed__4; LEAN_EXPORT lean_object* l_instBEqOption___rarg(lean_object*); LEAN_EXPORT lean_object* l_Option_instFunctorOption; LEAN_EXPORT lean_object* l_instMonadOption___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Option_all(lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822_(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800_(lean_object*); static lean_object* l_instMonadOption___closed__2; LEAN_EXPORT uint8_t l_Option_isSome___rarg(lean_object*); LEAN_EXPORT lean_object* l_Option_isSome(lean_object*); @@ -60,7 +62,6 @@ LEAN_EXPORT lean_object* l_instAlternativeOption___lambda__2(lean_object*, lean_ static lean_object* l_instMonadExceptOfUnitOption___closed__2; static lean_object* l_instMonadOption___closed__6; LEAN_EXPORT lean_object* l_Option_toMonad___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_instMonadExceptOfUnitOption; static lean_object* l_Option_instFunctorOption___closed__1; LEAN_EXPORT lean_object* l_instMonadExceptOfUnitOption___lambda__2___boxed(lean_object*, lean_object*, lean_object*); @@ -68,7 +69,6 @@ static lean_object* l_instMonadExceptOfUnitOption___closed__1; LEAN_EXPORT lean_object* l_Option_mapM___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_instMonadExceptOfUnitOption___lambda__1(lean_object*, lean_object*); static lean_object* l_instMonadOption___closed__1; -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__decEqOption____x40_Init_Data_Option_Basic___hyg_683____at_instDecidableEqOption___spec__1___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_instBEqOption(lean_object*); LEAN_EXPORT lean_object* l_Option_any(lean_object*); static lean_object* l_Option_instFunctorOption___closed__2; @@ -81,23 +81,21 @@ LEAN_EXPORT lean_object* l_instAlternativeOption___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Option_instDecidableRelOptionLt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_instDecidableEqOption___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_instAlternativeOption___closed__2; -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__decEqOption____x40_Init_Data_Option_Basic___hyg_683_(lean_object*); -LEAN_EXPORT lean_object* l_Option_map___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Option_map(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__decEqOption____x40_Init_Data_Option_Basic___hyg_661_(lean_object*); static lean_object* l_instMonadOption___closed__7; LEAN_EXPORT lean_object* l_liftOption___rarg(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Option_isNone___rarg(lean_object*); LEAN_EXPORT lean_object* l_instLTOption___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__decEqOption____x40_Init_Data_Option_Basic___hyg_683____at_instDecidableEqOption___spec__1(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__decEqOption____x40_Init_Data_Option_Basic___hyg_661____at_instDecidableEqOption___spec__1(lean_object*); LEAN_EXPORT uint8_t l_Option_toBool___rarg(lean_object*); static lean_object* l_Option_instFunctorOption___closed__3; LEAN_EXPORT lean_object* l_instMonadExceptOfUnitOption___lambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_instFunctorOption; -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__decEqOption____x40_Init_Data_Option_Basic___hyg_683____rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_instMonadOption___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_instMonadOption___closed__3; LEAN_EXPORT lean_object* l_Option_tryCatch(lean_object*); LEAN_EXPORT lean_object* l_Option_filter(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Option_toBool___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Option_toMonad___rarg(lean_object* x_1, lean_object* x_2) { _start: @@ -310,50 +308,6 @@ x_3 = lean_alloc_closure((void*)(l_Option_bind___rarg), 2, 0); return x_3; } } -LEAN_EXPORT lean_object* l_Option_map___rarg(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_2) == 0) -{ -lean_object* x_3; -lean_dec(x_1); -x_3 = lean_box(0); -return x_3; -} -else -{ -uint8_t x_4; -x_4 = !lean_is_exclusive(x_2); -if (x_4 == 0) -{ -lean_object* x_5; lean_object* x_6; -x_5 = lean_ctor_get(x_2, 0); -x_6 = lean_apply_1(x_1, x_5); -lean_ctor_set(x_2, 0, x_6); -return x_2; -} -else -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_7 = lean_ctor_get(x_2, 0); -lean_inc(x_7); -lean_dec(x_2); -x_8 = lean_apply_1(x_1, x_7); -x_9 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_9, 0, x_8); -return x_9; -} -} -} -} -LEAN_EXPORT lean_object* l_Option_map(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Option_map___rarg), 2, 0); -return x_3; -} -} LEAN_EXPORT lean_object* l_Option_mapM___rarg___lambda__1(lean_object* x_1, lean_object* x_2) { _start: { @@ -783,7 +737,7 @@ x_3 = lean_alloc_closure((void*)(l_Option_instDecidableRelOptionLt___rarg), 3, 0 return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__decEqOption____x40_Init_Data_Option_Basic___hyg_683____rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__decEqOption____x40_Init_Data_Option_Basic___hyg_661____rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_2) == 0) @@ -831,15 +785,15 @@ return x_12; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__decEqOption____x40_Init_Data_Option_Basic___hyg_683_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__decEqOption____x40_Init_Data_Option_Basic___hyg_661_(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Init_Data_Option_Basic_0__decEqOption____x40_Init_Data_Option_Basic___hyg_683____rarg), 3, 0); +x_2 = lean_alloc_closure((void*)(l___private_Init_Data_Option_Basic_0__decEqOption____x40_Init_Data_Option_Basic___hyg_661____rarg), 3, 0); return x_2; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__decEqOption____x40_Init_Data_Option_Basic___hyg_683____at_instDecidableEqOption___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__decEqOption____x40_Init_Data_Option_Basic___hyg_661____at_instDecidableEqOption___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_2) == 0) @@ -887,11 +841,11 @@ return x_12; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__decEqOption____x40_Init_Data_Option_Basic___hyg_683____at_instDecidableEqOption___spec__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__decEqOption____x40_Init_Data_Option_Basic___hyg_661____at_instDecidableEqOption___spec__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Init_Data_Option_Basic_0__decEqOption____x40_Init_Data_Option_Basic___hyg_683____at_instDecidableEqOption___spec__1___rarg), 3, 0); +x_2 = lean_alloc_closure((void*)(l___private_Init_Data_Option_Basic_0__decEqOption____x40_Init_Data_Option_Basic___hyg_661____at_instDecidableEqOption___spec__1___rarg), 3, 0); return x_2; } } @@ -899,7 +853,7 @@ LEAN_EXPORT lean_object* l_instDecidableEqOption___rarg(lean_object* x_1, lean_o _start: { lean_object* x_4; -x_4 = l___private_Init_Data_Option_Basic_0__decEqOption____x40_Init_Data_Option_Basic___hyg_683____at_instDecidableEqOption___spec__1___rarg(x_1, x_2, x_3); +x_4 = l___private_Init_Data_Option_Basic_0__decEqOption____x40_Init_Data_Option_Basic___hyg_661____at_instDecidableEqOption___spec__1___rarg(x_1, x_2, x_3); return x_4; } } @@ -911,7 +865,7 @@ x_2 = lean_alloc_closure((void*)(l_instDecidableEqOption___rarg), 3, 0); return x_2; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_2) == 0) @@ -959,11 +913,11 @@ return x_12; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800_(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____rarg), 3, 0); +x_2 = lean_alloc_closure((void*)(l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____rarg), 3, 0); return x_2; } } @@ -971,7 +925,7 @@ LEAN_EXPORT lean_object* l_instBEqOption___rarg(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____rarg), 3, 1); +x_2 = lean_alloc_closure((void*)(l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____rarg), 3, 1); lean_closure_set(x_2, 0, x_1); return x_2; } diff --git a/stage0/stdlib/Init/Data/Random.c b/stage0/stdlib/Init/Data/Random.c index 319e4bed542..4d5608f0597 100644 --- a/stage0/stdlib/Init/Data/Random.c +++ b/stage0/stdlib/Init/Data/Random.c @@ -50,7 +50,6 @@ lean_object* lean_st_mk_ref(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_stdGenRef; static lean_object* l_stdRange___closed__1; LEAN_EXPORT lean_object* l___private_Init_Data_Random_0__randNatAux(lean_object*); -static lean_object* l_initFn____x40_Init_Data_Random___hyg_712____closed__1; static lean_object* l_instReprStdGen___closed__2; LEAN_EXPORT lean_object* l_IO_rand(lean_object*, lean_object*, lean_object*); static lean_object* l_instRandomGenStdGen___closed__1; @@ -66,7 +65,9 @@ lean_object* lean_int_div(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_randBool___rarg(lean_object*, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_randNat___at_IO_rand___spec__1___boxed(lean_object*, lean_object*, lean_object*); +lean_object* lean_io_get_random_bytes(size_t, lean_object*); static lean_object* l_stdNext___closed__5; +lean_object* lean_uint64_to_nat(uint64_t); static lean_object* l_instRandomGenStdGen___closed__2; LEAN_EXPORT lean_object* l_randNat___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_int_sub(lean_object*, lean_object*); @@ -77,6 +78,7 @@ static lean_object* l_stdNext___closed__7; lean_object* lean_int_add(lean_object*, lean_object*); lean_object* lean_nat_mod(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_initFn____x40_Init_Data_Random___hyg_712_(lean_object*); +lean_object* l_ByteArray_toUInt64LE_x21(lean_object*); LEAN_EXPORT lean_object* l_instReprStdGen(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_instInhabitedStdGen; static lean_object* l_instRandomGenStdGen___closed__4; @@ -1289,38 +1291,67 @@ x_2 = lean_alloc_closure((void*)(l_randBool___rarg), 2, 0); return x_2; } } -static lean_object* _init_l_initFn____x40_Init_Data_Random___hyg_712____closed__1() { +LEAN_EXPORT lean_object* l_initFn____x40_Init_Data_Random___hyg_712_(lean_object* x_1) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(0u); -x_2 = l_mkStdGen(x_1); -return x_2; +size_t x_2; lean_object* x_3; +x_2 = 8; +x_3 = lean_io_get_random_bytes(x_2, x_1); +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint64_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_3, 1); +lean_inc(x_5); +lean_dec(x_3); +x_6 = l_ByteArray_toUInt64LE_x21(x_4); +x_7 = lean_unbox_uint64(x_6); +lean_dec(x_6); +x_8 = lean_uint64_to_nat(x_7); +x_9 = l_mkStdGen(x_8); +lean_dec(x_8); +x_10 = lean_st_mk_ref(x_9, x_5); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +return x_10; } +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_10); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +return x_14; } -LEAN_EXPORT lean_object* l_initFn____x40_Init_Data_Random___hyg_712_(lean_object* x_1) { -_start: +} +else { -lean_object* x_2; lean_object* x_3; uint8_t x_4; -x_2 = l_initFn____x40_Init_Data_Random___hyg_712____closed__1; -x_3 = lean_st_mk_ref(x_2, x_1); -x_4 = !lean_is_exclusive(x_3); -if (x_4 == 0) +uint8_t x_15; +x_15 = !lean_is_exclusive(x_3); +if (x_15 == 0) { return x_3; } else { -lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_5 = lean_ctor_get(x_3, 0); -x_6 = lean_ctor_get(x_3, 1); -lean_inc(x_6); -lean_inc(x_5); +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_3, 0); +x_17 = lean_ctor_get(x_3, 1); +lean_inc(x_17); +lean_inc(x_16); lean_dec(x_3); -x_7 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_7, 0, x_5); -lean_ctor_set(x_7, 1, x_6); -return x_7; +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +return x_18; +} } } } @@ -1824,8 +1855,6 @@ l_instRandomGenStdGen___closed__4 = _init_l_instRandomGenStdGen___closed__4(); lean_mark_persistent(l_instRandomGenStdGen___closed__4); l_instRandomGenStdGen = _init_l_instRandomGenStdGen(); lean_mark_persistent(l_instRandomGenStdGen); -l_initFn____x40_Init_Data_Random___hyg_712____closed__1 = _init_l_initFn____x40_Init_Data_Random___hyg_712____closed__1(); -lean_mark_persistent(l_initFn____x40_Init_Data_Random___hyg_712____closed__1); res = l_initFn____x40_Init_Data_Random___hyg_712_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_IO_stdGenRef = lean_io_result_get_value(res); diff --git a/stage0/stdlib/Init/Data/String/Basic.c b/stage0/stdlib/Init/Data/String/Basic.c index f67b771eef8..ac561e72d72 100644 --- a/stage0/stdlib/Init/Data/String/Basic.c +++ b/stage0/stdlib/Init/Data/String/Basic.c @@ -2763,24 +2763,35 @@ return x_1; LEAN_EXPORT uint8_t l_String_isNat(lean_object* x_1) { _start: { -lean_object* x_2; lean_object* x_3; lean_object* x_4; uint8_t x_5; -x_2 = lean_string_utf8_byte_size(x_1); -x_3 = l_String_isNat___closed__1; -x_4 = lean_unsigned_to_nat(0u); -x_5 = l_String_anyAux_loop(x_1, x_2, x_3, x_4); -lean_dec(x_2); -if (x_5 == 0) +uint8_t x_2; +x_2 = l_String_isEmpty(x_1); +if (x_2 == 0) { -uint8_t x_6; -x_6 = 1; -return x_6; -} -else +lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_3 = lean_string_utf8_byte_size(x_1); +x_4 = l_String_isNat___closed__1; +x_5 = lean_unsigned_to_nat(0u); +x_6 = l_String_anyAux_loop(x_1, x_3, x_4, x_5); +lean_dec(x_3); +if (x_6 == 0) { uint8_t x_7; -x_7 = 0; +x_7 = 1; return x_7; } +else +{ +uint8_t x_8; +x_8 = 0; +return x_8; +} +} +else +{ +uint8_t x_9; +x_9 = 0; +return x_9; +} } } LEAN_EXPORT lean_object* l_String_isNat___lambda__1___boxed(lean_object* x_1) { diff --git a/stage0/stdlib/Init/Data/String/Extra.c b/stage0/stdlib/Init/Data/String/Extra.c index 40a9b96b393..910e5c3ce00 100644 --- a/stage0/stdlib/Init/Data/String/Extra.c +++ b/stage0/stdlib/Init/Data/String/Extra.c @@ -22,7 +22,6 @@ lean_object* lean_name_mk_string(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1(lean_object*, lean_object*, lean_object*); static lean_object* l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__2___closed__2; LEAN_EXPORT lean_object* l_String_toNat_x21___lambda__1(lean_object*, uint32_t); -extern lean_object* l_instInhabitedNat; lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); static lean_object* l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__2___closed__1; @@ -44,7 +43,6 @@ static lean_object* l_String___aux__Init__Data__String__Extra______macroRules__t static lean_object* l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__22; static lean_object* l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__28; static lean_object* l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__14; -LEAN_EXPORT lean_object* l_panic___at_String_toNat_x21___spec__1(lean_object*); uint8_t l_String_isNat(lean_object*); static lean_object* l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__29; static lean_object* l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__25; @@ -67,10 +65,11 @@ static lean_object* l_String___aux__Init__Data__String__Extra______macroRules__t static lean_object* l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__9; static lean_object* l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__8; static lean_object* l_String_toNat_x21___closed__4; -lean_object* lean_panic_fn(lean_object*, lean_object*); +lean_object* l_panic___at_Lean_TSyntax_getNat___spec__1(lean_object*); static lean_object* l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__27; uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_fromUTF8Unchecked___boxed(lean_object*); +LEAN_EXPORT lean_object* l_String_toNat_x21___boxed(lean_object*); static lean_object* l_String_toNat_x21___closed__1; static lean_object* l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__2___closed__5; static lean_object* l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__17; @@ -85,15 +84,6 @@ LEAN_EXPORT lean_object* l___private_Init_Data_String_Extra_0__String_utf8ByteSi static lean_object* l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__2___closed__7; static lean_object* l_String_toNat_x21___closed__5; static lean_object* l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__19; -LEAN_EXPORT lean_object* l_panic___at_String_toNat_x21___spec__1(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; -x_2 = l_instInhabitedNat; -x_3 = lean_panic_fn(x_2, x_1); -return x_3; -} -} LEAN_EXPORT lean_object* l_String_toNat_x21___lambda__1(lean_object* x_1, uint32_t x_2) { _start: { @@ -163,9 +153,8 @@ x_2 = l_String_isNat(x_1); if (x_2 == 0) { lean_object* x_3; lean_object* x_4; -lean_dec(x_1); x_3 = l_String_toNat_x21___closed__4; -x_4 = l_panic___at_String_toNat_x21___spec__1(x_3); +x_4 = l_panic___at_Lean_TSyntax_getNat___spec__1(x_3); return x_4; } else @@ -176,7 +165,6 @@ x_6 = l_String_toNat_x21___closed__5; x_7 = lean_unsigned_to_nat(0u); x_8 = l_String_foldlAux_loop___rarg(x_6, x_1, x_5, x_7, x_7); lean_dec(x_5); -lean_dec(x_1); return x_8; } } @@ -192,6 +180,15 @@ lean_dec(x_1); return x_4; } } +LEAN_EXPORT lean_object* l_String_toNat_x21___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_String_toNat_x21(x_1); +lean_dec(x_1); +return x_2; +} +} LEAN_EXPORT lean_object* l_String_fromUTF8Unchecked___boxed(lean_object* x_1) { _start: { diff --git a/stage0/stdlib/Init/Data/ToString/Macro.c b/stage0/stdlib/Init/Data/ToString/Macro.c index be031f0a60d..ec5d2e0023d 100644 --- a/stage0/stdlib/Init/Data/ToString/Macro.c +++ b/stage0/stdlib/Init/Data/ToString/Macro.c @@ -25,8 +25,8 @@ static lean_object* l_termS_x21_____closed__2; static lean_object* l___aux__Init__Data__ToString__Macro______macroRules__termS_x21____1___closed__15; static lean_object* l_termS_x21_____closed__13; static lean_object* l___aux__Init__Data__ToString__Macro______macroRules__termS_x21____1___closed__6; -lean_object* l_Lean_Syntax_expandInterpolatedStr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_termS_x21_____closed__6; +lean_object* l_Lean_TSyntax_expandInterpolatedStr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__Data__ToString__Macro______macroRules__termS_x21____1___closed__3; static lean_object* l___aux__Init__Data__ToString__Macro______macroRules__termS_x21____1___closed__5; static lean_object* l___aux__Init__Data__ToString__Macro______macroRules__termS_x21____1___closed__14; @@ -421,10 +421,53 @@ lean_ctor_set(x_27, 0, x_21); lean_ctor_set(x_27, 1, x_25); lean_ctor_set(x_27, 2, x_24); lean_ctor_set(x_27, 3, x_26); -x_28 = l_Lean_Syntax_expandInterpolatedStr(x_9, x_19, x_27, x_2, x_22); +x_28 = l_Lean_TSyntax_expandInterpolatedStr(x_9, x_19, x_27, x_2, x_22); lean_dec(x_9); +if (lean_obj_tag(x_28) == 0) +{ +uint8_t x_29; +x_29 = !lean_is_exclusive(x_28); +if (x_29 == 0) +{ +return x_28; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_28, 0); +x_31 = lean_ctor_get(x_28, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_28); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +return x_32; +} +} +else +{ +uint8_t x_33; +x_33 = !lean_is_exclusive(x_28); +if (x_33 == 0) +{ return x_28; } +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_28, 0); +x_35 = lean_ctor_get(x_28, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_28); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +} +} } } lean_object* initialize_Init_Meta(uint8_t builtin, lean_object*); diff --git a/stage0/stdlib/Init/Meta.c b/stage0/stdlib/Init/Meta.c index 53e99724f46..bd239aad449 100644 --- a/stage0/stdlib/Init/Meta.c +++ b/stage0/stdlib/Init/Meta.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Init.Meta -// Imports: Init.Data.Array.Basic +// Imports: Init.Data.Array.Basic Init.Data.Option.BasicAux #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -13,1299 +13,1374 @@ #ifdef __cplusplus extern "C" { #endif -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__7; lean_object* l_List_reverse___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_getSepElems___rarg___boxed(lean_object*); static uint8_t l_Lean_versionString___closed__2; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__10; +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__73; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__32; static lean_object* l_Lean_mkHole___closed__3; static lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__10; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f_decode___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkCIdentFrom(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__3; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__14; +static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__7; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_version_getMajor___boxed(lean_object*); lean_object* lean_string_push(lean_object*, uint32_t); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeOctalLitAux___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__17; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__12; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__81; LEAN_EXPORT lean_object* l_Lean_Syntax_isIdOrAtom_x3f___boxed(lean_object*); lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_dsimpKind; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__16; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__13; +static lean_object* l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___closed__2; static lean_object* l_Lean_Parser_Tactic_simpAllArith___closed__3; LEAN_EXPORT lean_object* l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__1(lean_object*); lean_object* l_Lean_extractMacroScopes(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkIdentFromRef(lean_object*); static lean_object* l_Lean_Parser_Tactic_dsimpKind___closed__2; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__32; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__10; +static lean_object* l_Lean_instQuoteNameMkStrAnonymous___closed__1; size_t lean_usize_add(size_t, size_t); -extern lean_object* l_Lean_fieldIdxKind; LEAN_EXPORT lean_object* l_Lean_Syntax_isNatLit_x3f___boxed(lean_object*); static lean_object* l_Lean_Syntax_unsetTrailing___closed__1; static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__6; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__19; LEAN_EXPORT lean_object* l_Lean_Syntax_setTailInfoAux(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTailSepArrayArraySyntax___boxed(lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__20; static lean_object* l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__23; +static lean_object* l_Lean_Syntax_mkNumLit___closed__2; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_updateFirst___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__89; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____lambda__1___closed__1; LEAN_EXPORT uint8_t l_Lean_Meta_DSimp_Config_decide___default; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeDecimalLitAux___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__1; LEAN_EXPORT uint8_t l_Lean_Meta_DSimp_Config_zeta___default; -LEAN_EXPORT lean_object* l_Lean_instQuoteBool(uint8_t); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__20; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f_decodeAfterExp___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__12; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_mkSepArray___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__6; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__5; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__8; LEAN_EXPORT uint8_t lean_is_inaccessible_user_name(lean_object*); +extern lean_object* l_String_instInhabitedString; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__23; static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__2; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__92; +static lean_object* l_Lean_instQuoteTermMkStrAnonymous___closed__1; static lean_object* l_Lean_versionString___closed__1; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__35; +static lean_object* l_Lean_quoteNameMk___closed__6; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__4; lean_object* lean_mk_empty_array_with_capacity(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__8; LEAN_EXPORT lean_object* l_Lean_Syntax_getSubstring_x3f___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__81; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__15; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__10; static lean_object* l_Lean_Meta_instReprTransparencyMode___closed__1; static lean_object* l_Lean_Name_toString_maybePseudoSyntax___closed__2; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__11; +static lean_object* l_Lean_Option_hasQuote___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_Syntax_getTrailingSize(lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__8; -LEAN_EXPORT lean_object* l_Lean_Syntax_expandInterpolatedStr___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTermTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_instQuoteProdMkStrAnonymous___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__90; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__6; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__69; static lean_object* l_Lean_Name_reprPrec___closed__7; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__17; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__94; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__5; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__12; LEAN_EXPORT uint8_t l_Lean_Meta_DSimp_Config_autoUnfold___default; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_updateFirst___at_Lean_Syntax_setHeadInfoAux___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__5; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__21; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__14; -extern lean_object* l_Lean_nullKind; +LEAN_EXPORT lean_object* l_Lean_instQuoteListMkStrAnonymous___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_getSepArgs___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__38; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Array_filterSepElemsMAux___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpArith___closed__8; -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__9; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__7; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__15; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__2; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__2; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeHexDigit(lean_object*, lean_object*); +static lean_object* l_Lean_instQuoteBoolMkStrAnonymous___closed__8; static lean_object* l_Lean_Name_escapePart___closed__1; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__95; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__24; lean_object* lean_name_mk_string(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Syntax_getOptional_x3f(lean_object*); +lean_object* l_Lean_Syntax_getOptional_x3f(lean_object*); static lean_object* l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__3; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__18; +LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSepArray___boxed(lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); -static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__20; lean_object* lean_array_uget(lean_object*, size_t); static lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__9; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__44; static lean_object* l_Lean_termEval__prio_____closed__9; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__7; LEAN_EXPORT lean_object* l_Lean_versionStringCore; -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__13; static lean_object* l_Lean_versionString___closed__8; LEAN_EXPORT uint32_t l_Lean_idBeginEscape; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__43; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__1; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__26; +LEAN_EXPORT lean_object* l_Lean_instQuote___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__75; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailSyntaxTSyntax(lean_object*); lean_object* l_Array_append___rarg(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_instInhabitedEtaStructMode; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__87; LEAN_EXPORT lean_object* l_Lean_termEval__prio__; -LEAN_EXPORT lean_object* l_Lean_Syntax_expandInterpolatedStrChunks(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_TSyntax_expandInterpolatedStrChunks___spec__1___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeQuotedChar(lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__3; +static lean_object* l_Lean_quoteNameMk___closed__3; static lean_object* l_Lean_mkSepArray___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_simpAutoUnfold; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__1; static lean_object* l_Lean_versionString___closed__4; LEAN_EXPORT lean_object* l_Lean_Syntax_mkSep(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__56; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_simpAllArith; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__66; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__34; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__11; static lean_object* l_Lean_mkHole___closed__7; -static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__10; +static lean_object* l_Lean_quoteNameMk___closed__5; LEAN_EXPORT lean_object* l_Array_filterSepElemsM___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_instQuoteSubstring___closed__3; static lean_object* l_Lean_toolchain___closed__3; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__58; lean_object* l_Lean_SourceInfo_fromRef(lean_object*); uint8_t l_String_anyAux_loop(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Name_capitalize(lean_object*); static lean_object* l_Lean_versionString___closed__9; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__21; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__38; LEAN_EXPORT lean_object* l_Lean_Syntax_isAtom___boxed(lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__23; static lean_object* l_Lean_termEval__prec_____closed__6; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__96; static lean_object* l_Lean_version_specialDesc___closed__1; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__68; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__35; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_updateFirst(lean_object*); +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeIdentTerm(lean_object*); LEAN_EXPORT lean_object* l_Lean_monadNameGeneratorLift(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__72; LEAN_EXPORT lean_object* l_Array_mapSepElemsM___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_toolchain___closed__9; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__73; extern lean_object* l_Lean_maxRecDepthErrorMessage; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__13; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__48; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f_decodeAfterDot(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_toNat___boxed(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__5; static lean_object* l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__15; static lean_object* l_Lean_Parser_Tactic_simpArith___closed__6; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__21; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15670____closed__1; +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__13; LEAN_EXPORT uint8_t l_Lean_Meta_Simp_Config_memoize___default; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__13; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__2; static lean_object* l_Lean_toolchain___closed__4; +LEAN_EXPORT lean_object* l_Lean_TSyntax_getNat(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic__________; LEAN_EXPORT lean_object* l_Lean_Syntax_getHead_x3f(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_dsimpKind___closed__8; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__64; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_DSimp_Config_eta___default; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__98; LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__5(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__11; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__97; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__25; LEAN_EXPORT uint8_t l_Lean_Syntax_structEq(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__6; +LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTailSepArrayArraySyntax(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_EtaStructMode_noConfusion___rarg(uint8_t, uint8_t, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeInterpStrQuotedChar(lean_object*, lean_object*); static lean_object* l_Lean_githash___closed__1; +static lean_object* l_Lean_instQuoteSubstringMkStrAnonymous___closed__3; LEAN_EXPORT lean_object* l_Lean_termEval__prec__; +static lean_object* l_Lean_quoteNameMk___closed__8; LEAN_EXPORT lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__termEval__prio____1(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__7; -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Syntax_expandInterpolatedStrChunks___spec__1___lambda__2___closed__1; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Syntax_expandInterpolatedStrChunks___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__21; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__42; -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__1; +static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__14; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__68; static lean_object* l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__4; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__63; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__7; lean_object* lean_string_utf8_prev(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__71; -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__85; static lean_object* l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__1; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__6; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_SepArray_ofElemsUsingRef___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__40; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__12; -static lean_object* l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__5; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__87; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____boxed(lean_object*, lean_object*); lean_object* l_id___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeOctalLitAux(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_instQuote(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__17; -static lean_object* l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__3; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Syntax_findAux___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__49; static lean_object* l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__8; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__87; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Syntax_getHead_x3f___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__12; uint8_t l_Char_isDigit(uint32_t); -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__2; LEAN_EXPORT lean_object* l_Lean_Name_reprPrec(lean_object*, lean_object*); -static lean_object* l_Lean_instQuoteBool___closed__7; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__75; LEAN_EXPORT lean_object* l_Lean_Name_escapePart(lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__20; +LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailSyntaxTSyntax___rarg___boxed(lean_object*); +static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__5; LEAN_EXPORT lean_object* l_Lean_withHeadRefOnly___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__26; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f_decodeExp(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__68; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__2; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__13; LEAN_EXPORT lean_object* l_Lean_isGreek___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__16; static lean_object* l_Lean_Parser_Tactic_dsimpKind___closed__1; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Syntax_expandInterpolatedStrChunks___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeBinLitAux(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__24; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__58; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeNumLitTerm(lean_object*); LEAN_EXPORT uint32_t l_Lean_idEndEscape; -LEAN_EXPORT lean_object* l_Lean_Syntax_SepArray_instCoeTailSepArrayArraySyntax(lean_object*); LEAN_EXPORT lean_object* l_Lean_Name_escapePart___lambda__1___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__11; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeQuotedChar___boxed__const__5; static lean_object* l_Lean_Name_escapePart___closed__2; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__22; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__69; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__23; LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__Substring_takeWhileAux___at___private_Init_Meta_0__Lean_Syntax_splitNameLitAux___spec__3___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isIdRest___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_instQuoteProd(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__49; +LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkIdentFrom(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__84; +static lean_object* l_Lean_quoteNameMk___closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__1; LEAN_EXPORT uint8_t l_Lean_isIdBeginEscape(uint32_t); static lean_object* l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___closed__3; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__16; LEAN_EXPORT lean_object* l_Lean_Meta_EtaStructMode_toCtorIdx(uint8_t); +static lean_object* l_Lean_Syntax_mkNumLit___closed__1; LEAN_EXPORT uint8_t l_Lean_Name_escapePart___lambda__1(uint32_t); -static lean_object* l_Lean_instQuoteProd___rarg___closed__3; +static lean_object* l_Lean_instQuoteBoolMkStrAnonymous___closed__4; LEAN_EXPORT lean_object* l_Lean_Syntax_mkScientificLit(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__54; +extern lean_object* l_instInhabitedNat; lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(lean_object*, lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__8; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__16; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__62; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__63; extern lean_object* l_Lean_Parser_Tactic_config; -static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__12; static lean_object* l_Lean_termEval__prec_____closed__3; +static lean_object* l_Lean_Option_hasQuote___rarg___closed__7; +static lean_object* l_Lean_TSyntax_expandInterpolatedStr___closed__3; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__24; static lean_object* l_Lean_Parser_Tactic_dsimpKind___closed__9; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__79; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__9; LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__Substring_takeWhileAux___at___private_Init_Meta_0__Lean_Syntax_splitNameLitAux___spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllArith___closed__6; LEAN_EXPORT lean_object* l_Lean_isIdFirst___boxed(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__13; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Array_filterSepElemsMAux___at_Array_filterSepElems___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Name_toString___boxed(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__16; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__27; +static lean_object* l_Lean_Syntax_isInterpolatedStrLit_x3f___closed__1; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__7; lean_object* lean_string_append(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_beqConfig____x40_Init_Meta___hyg_8947____boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_dsimpKind___closed__4; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Syntax_expandInterpolatedStrChunks___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__47; LEAN_EXPORT lean_object* l_Lean_version_getSpecialDesc___boxed(lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__90; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeHexDigit___boxed(lean_object*, lean_object*); lean_object* lean_get_githash(lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__19; LEAN_EXPORT lean_object* l_Lean_Syntax_setTailInfo(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__6; -LEAN_EXPORT lean_object* l_Lean_instQuoteArray(lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__1; lean_object* lean_string_utf8_extract(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__65; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__39; -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__4; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__2; -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSyntaxArrayArraySyntax(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__7; +static lean_object* l_Lean_quoteNameMk___closed__4; LEAN_EXPORT uint8_t l_Lean_Meta_Simp_Config_decide___default; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__51; static lean_object* l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__21; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__70; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__31; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__59; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__2; -static lean_object* l_Lean_instQuoteBool___closed__8; +lean_object* l_Lean_TSyntaxArray_rawImpl___rarg___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__1; static lean_object* l_Lean_Meta_instBEqEtaStructMode___closed__1; static lean_object* l___private_Init_Meta_0__Lean_quoteList___rarg___closed__3; lean_object* l_Std_Format_joinSep___at_instReprProd___spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__2; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__65; static lean_object* l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__13; lean_object* l___private_Init_Data_Repr_0__reprSourceInfo____x40_Init_Data_Repr___hyg_1755_(lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__21; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__39; +LEAN_EXPORT lean_object* l_Lean_instQuoteSubstringMkStrAnonymous(lean_object*); LEAN_EXPORT lean_object* l_Array_filterSepElemsM___at_Array_filterSepElems___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_quoteNameMk___closed__9; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__1; LEAN_EXPORT uint8_t l_Lean_Meta_Simp_Config_contextual___default; LEAN_EXPORT lean_object* l_Lean_Meta_TransparencyMode_toCtorIdx(uint8_t); -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__8; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__23; -LEAN_EXPORT lean_object* l_Lean_instQuoteList___rarg(lean_object*); lean_object* lean_string_utf8_byte_size(lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_instInhabitedTransparencyMode; LEAN_EXPORT lean_object* l_Lean_NameGenerator_namePrefix___default; -LEAN_EXPORT lean_object* l_Lean_Syntax_expandInterpolatedStr___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__12; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__16; +static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__16; LEAN_EXPORT lean_object* l_Array_filterSepElems___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__9; -LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_8567_(uint8_t, uint8_t); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__82; +LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_9542_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_Syntax_setHeadInfo(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__7; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__16; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__61; static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__8; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__12; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__4; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__10; +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeTSyntaxConsSyntaxNodeKindNil(lean_object*, lean_object*); static lean_object* l_Lean_toolchain___closed__5; -static lean_object* l_Lean_instQuoteBool___closed__1; uint8_t lean_usize_dec_lt(size_t, size_t); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__74; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__60; +static lean_object* l_Lean_Option_hasQuote___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_version_patch; -extern lean_object* l_Lean_nameLitKind; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__77; static lean_object* l___private_Init_Meta_0__Lean_quoteList___rarg___closed__5; static lean_object* l_Lean_versionString___closed__11; LEAN_EXPORT lean_object* l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_isNatLitAux___boxed(lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__22; LEAN_EXPORT lean_object* l_Lean_withHeadRefOnly___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_instQuoteBool___boxed(lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__36; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__2; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__61; LEAN_EXPORT lean_object* l_Lean_version_specialDesc; -static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__3; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__34; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__19; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__84; lean_object* lean_nat_add(lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__6; LEAN_EXPORT lean_object* l_Lean_Syntax_mkSep___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__3; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeQuotedChar___boxed__const__6; LEAN_EXPORT lean_object* l_Lean_isSubScriptAlnum___boxed(lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__84; LEAN_EXPORT lean_object* l_Array_mapSepElems(lean_object*, lean_object*); static lean_object* l_Lean_Meta_DSimp_instBEqConfig___closed__1; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__14; static lean_object* l_Lean_Parser_Tactic_tacticErw_______closed__5; static lean_object* l_Lean_versionStringCore___closed__2; static lean_object* l_Lean_Parser_Tactic_simpAllArith___closed__2; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__58; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__33; +static lean_object* l_Lean_TSyntax_getNat___closed__3; +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_beqConfig____x40_Init_Meta___hyg_9922____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Name_instReprName; static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__5; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__4; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__8; static lean_object* l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__7; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__59; LEAN_EXPORT lean_object* l_Lean_Name_toStringWithSep(lean_object*, uint8_t, lean_object*); static lean_object* l_Lean_Parser_Tactic_tacticErw_______closed__7; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__31; LEAN_EXPORT lean_object* l_Lean_Syntax_SepArray_ofElems(lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__15; static lean_object* l_Lean_Name_escapePart___closed__4; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__19; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__7; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__11; static lean_object* l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__16; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__82; +LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSepArrayTSyntaxArray(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__7; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__24; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__69; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__19; static lean_object* l_Lean_termEval__prio_____closed__4; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__22; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__82; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__93; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__51; static lean_object* l_Lean_version_major___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_EtaStructMode_toCtorIdx___boxed(lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__33; +LEAN_EXPORT lean_object* l_Lean_TSyntax_expandInterpolatedStr___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Option_hasQuote___rarg___closed__4; LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__2(lean_object*, lean_object*); -static lean_object* l_Lean_instQuoteSubstring___closed__1; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__5; static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__7; LEAN_EXPORT uint8_t l_Lean_Meta_Simp_Config_etaStruct___default; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__87; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__45; lean_object* lean_string_utf8_next(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__2; static lean_object* l_Lean_Name_isInaccessibleUserName___closed__1; +static lean_object* l_Lean_Syntax_mkStrLit___closed__2; static lean_object* l_Lean_versionStringCore___closed__1; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__78; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__91; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__7; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__9; static lean_object* l_Lean_termEval__prio_____closed__2; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__1; +LEAN_EXPORT lean_object* l_Lean_TSyntax_getId(lean_object*); +LEAN_EXPORT lean_object* l_Lean_TSyntax_expandInterpolatedStrChunks___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Name_toStringWithSep_maybeEscape(uint8_t, lean_object*); -static lean_object* l_Lean_instQuoteSubstring___closed__2; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__83; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__26; LEAN_EXPORT lean_object* l_Lean_Syntax_isLit_x3f___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_tacticErw_______closed__3; LEAN_EXPORT uint8_t l_Lean_Meta_DSimp_Config_proj___default; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__39; LEAN_EXPORT lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__addPrio__1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__51; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__16; LEAN_EXPORT lean_object* l_Lean_Syntax_hasArgs___boxed(lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__4; static lean_object* l_Lean_Name_reprPrec___closed__6; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__3; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeInterpStrLit_loop___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__11; LEAN_EXPORT lean_object* l_Lean_Syntax_getSubstring_x3f(lean_object*, uint8_t, uint8_t); +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__10; LEAN_EXPORT lean_object* l_Lean_Syntax_SepArray_getElems___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_findSomeRevM_x3f_find___at_Lean_Syntax_getTailInfo_x3f___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Simp_Config_singlePass___default; LEAN_EXPORT lean_object* l_Lean_Syntax_mkApp(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapSepElems___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__40; +LEAN_EXPORT lean_object* l_Lean_Syntax_instBEqTSyntax___boxed(lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__67; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_Config_maxSteps___default; -static lean_object* l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__6; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Syntax_expandInterpolatedStrChunks___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Syntax_TSepArray_getElems___rarg___boxed(lean_object*); +static lean_object* l_Lean_TSyntax_expandInterpolatedStrChunks___closed__1; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_version_getMinor___boxed(lean_object*); static lean_object* l_Lean_evalPrio___closed__1; uint8_t l_Lean_Name_hasMacroScopes(lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__27; static uint8_t l_Lean_version_isRelease___closed__1; -static lean_object* l_Lean_instQuoteBool___closed__5; +LEAN_EXPORT lean_object* l_Lean_Syntax_instBEqTSyntax(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__39; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__10; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__5; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__57; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__19; static lean_object* l_Lean_termEval__prec_____closed__10; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__69; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f_decodeExp___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_isLit_x3f(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_instQuoteList(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_quoteNameMk___closed__6; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__45; static lean_object* l_Lean_Name_reprPrec___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_instInhabitedConfig; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__17; LEAN_EXPORT lean_object* l_Lean_Syntax_isScientificLit_x3f___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_isCharLit_x3f(lean_object*); +static lean_object* l_Lean_instQuoteProdMkStrAnonymous___rarg___closed__4; lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_getTrailingSize___boxed(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__11; +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__14; LEAN_EXPORT uint8_t l_Lean_Meta_DSimp_Config_iota___default; LEAN_EXPORT lean_object* l_Lean_Syntax_isNatLit_x3f(lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__29; +LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSyntaxArray(lean_object*); static lean_object* l_Lean_Parser_Tactic_dsimpKind___closed__5; LEAN_EXPORT lean_object* l_Lean_version_major; -static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__18; -static lean_object* l_Lean_instQuoteProd___rarg___closed__4; +LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTermTSyntaxConsSyntaxNodeKindMkStrAnonymousNil(lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__94; LEAN_EXPORT lean_object* l_Lean_Syntax_instBEqSyntax; LEAN_EXPORT uint8_t l_Lean_Meta_Simp_Config_zeta___default; -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_8567____boxed(lean_object*, lean_object*); static lean_object* l_Lean_versionStringCore___closed__3; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__17; -static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__15; -static lean_object* l_Lean_instQuoteSyntax___closed__1; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__44; LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Syntax_structEq___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__20; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__8; +static lean_object* l_Lean_Syntax_mkStrLit___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_SepArray_ofElemsUsingRef___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Name_appendIndexAfter___closed__1; uint8_t l_instDecidableNot___rarg(uint8_t); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_neutralConfig; uint8_t l_String_contains(lean_object*, uint32_t); static lean_object* l_Lean_Parser_Tactic_tacticErw_______closed__1; +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__1; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__18; LEAN_EXPORT lean_object* l_Lean_Syntax_mkStrLit___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__61; -static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__22; static lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__subPrec__1___closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__4; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__20; -LEAN_EXPORT lean_object* l_Lean_Syntax_expandInterpolatedStr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__22; static lean_object* l_Lean_termEval__prec_____closed__7; static lean_object* l_Lean_Parser_Tactic_simpAllArith___closed__9; -extern lean_object* l_Lean_numLitKind; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__83; +LEAN_EXPORT lean_object* l_Lean_instQuoteNatNumLitKind(lean_object*); static lean_object* l_Lean_Parser_Tactic_tacticErw_______closed__2; -LEAN_EXPORT lean_object* l_Lean_instQuoteString(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__10; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeDecimalLitAux(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapSepElemsM(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkGroupNode(lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); static lean_object* l_Lean_termEval__prio_____closed__1; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__7; -static lean_object* l_Lean_instQuoteProd___rarg___closed__2; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__31; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__11; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__54; +static lean_object* l_Lean_mkGroupNode___closed__1; +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__10; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__35; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__15; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__62; LEAN_EXPORT uint8_t l_List_beq___at_Lean_Syntax_structEq___spec__3(lean_object*, lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__9; LEAN_EXPORT lean_object* l_Lean_mkFreshId___rarg___lambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__65; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__36; LEAN_EXPORT lean_object* l_Lean_Name_reprPrec___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f_decode(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__3; +static lean_object* l_Lean_mkGroupNode___closed__2; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Array_mapSepElemsMAux(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_mkCApp(lean_object*, lean_object*); -extern lean_object* l_Lean_strLitKind; +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__3; +static lean_object* l_Lean_Syntax_mkNameLit___closed__2; static lean_object* l_Lean_termEval__prio_____closed__6; static lean_object* l_Lean_versionStringCore___closed__5; LEAN_EXPORT lean_object* l_Lean_Syntax_isStrLit_x3f(lean_object*); static lean_object* l_Lean_Name_reprPrec___closed__1; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__35; static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__3; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__70; LEAN_EXPORT lean_object* l_Array_getSepElems(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__Substring_takeWhileAux___at___private_Init_Meta_0__Lean_Syntax_splitNameLitAux___spec__3(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__48; +static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__19; LEAN_EXPORT lean_object* l_Lean_monadNameGeneratorLift___rarg(lean_object*, lean_object*); static lean_object* l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__13; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__7; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__18; LEAN_EXPORT lean_object* l_Lean_Syntax_SepArray_getElems(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__17; LEAN_EXPORT lean_object* l_Lean_Syntax_structEq___boxed(lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__8; static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__5; static lean_object* l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__5; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__52; -LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Meta_Simp_beqConfig____x40_Init_Meta___hyg_9383_(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Meta_Simp_beqConfig____x40_Init_Meta___hyg_10358_(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Simp_ConfigCtx_contextual___default; LEAN_EXPORT lean_object* l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3(lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__41; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___rarg(lean_object*, lean_object*); static lean_object* l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__14; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__25; LEAN_EXPORT lean_object* l_Lean_Syntax_isFieldIdx_x3f(lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__13; -static lean_object* l_Lean_instQuoteProd___rarg___closed__1; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__88; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__36; +LEAN_EXPORT lean_object* l_Lean_TSyntax_expandInterpolatedStr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSyntaxArrayTSepArray___boxed(lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__6; +LEAN_EXPORT lean_object* l_Lean_instQuoteStringStrLitKind(lean_object*); lean_object* l_Lean_Syntax_getHeadInfo(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_SepArray_ofElems___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__36; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeHexLitAux(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_isNumericSubscript(uint32_t); -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__11; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_mkSepArray___spec__1___lambda__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_isScientificLit_x3f(lean_object*); static lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__1; LEAN_EXPORT lean_object* l_Lean_Name_toString(lean_object*, uint8_t); +LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSyntaxArray___rarg(lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__57; static lean_object* l_Lean_toolchain___closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_DSimp_instInhabitedConfig; -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteOption___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__6; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__17; lean_object* l_Lean_replaceRef(lean_object*, lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_instReprTransparencyMode; -static lean_object* l_Lean_instQuoteBool___closed__6; static lean_object* l_Lean_Parser_Tactic_simpAllArith___closed__4; LEAN_EXPORT lean_object* l_Lean_mkIdentFromRef___rarg___lambda__1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__16; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__14; LEAN_EXPORT lean_object* l_Lean_mkFreshId___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__80; lean_object* l_String_capitalize(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__5; +LEAN_EXPORT lean_object* l_panic___at_Lean_TSyntax_getString___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_NameGenerator_next(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_decodeCharLit___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_decodeNatLitVal_x3f(lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748_(uint8_t, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723_(uint8_t, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__9; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__3; LEAN_EXPORT lean_object* l_Lean_Syntax_SepArray_getElems___boxed(lean_object*); +static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__10; LEAN_EXPORT lean_object* l_Array_findSomeRevM_x3f_find___at_Lean_Syntax_getTailInfo_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailSyntaxTSyntax___boxed(lean_object*); +static lean_object* l_Lean_mkNullNode___closed__2; static lean_object* l_Lean_Parser_Tactic_dsimpKind___closed__6; +static lean_object* l_Lean_Option_hasQuote___rarg___closed__6; extern lean_object* l_Lean_Parser_Tactic_simpStar; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__8; lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Internal_isStage0(lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__94; static lean_object* l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__11; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__25; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__31; +static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__20; static lean_object* l_Lean_mkHole___closed__5; -static lean_object* l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_withHeadRefOnly___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_TransparencyMode_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__12; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__86; +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeStrLitTerm(lean_object*); +LEAN_EXPORT lean_object* l_Lean_TSyntax_expandInterpolatedStrChunks(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_setKind(lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_name_append_index_after(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Name_eraseSuffix_x3f(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_instReprConfig___closed__1; LEAN_EXPORT lean_object* l_Array_getSepElems___rarg(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__3; extern lean_object* l_Lean_reservedMacroScope; static lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__6; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__46; static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__10; static lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__8; +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeDepTermMkConsSyntaxNodeKindMkStrAnonymousNilIdentIdent(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Syntax_unsetTrailing___closed__2; -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__11; static lean_object* l_Lean_Parser_Tactic_simpArith___closed__10; static lean_object* l_Lean_NameGenerator_namePrefix___default___closed__1; static lean_object* l_Lean_Meta_DSimp_instInhabitedConfig___closed__1; +static lean_object* l_Lean_Syntax_instCoeTailSepArrayArraySyntax___closed__1; LEAN_EXPORT lean_object* l_Lean_mkNullNode(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__15; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__12; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__25; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__15; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__4; LEAN_EXPORT lean_object* l_Lean_Name_instToStringName(lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__28; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__5; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__90; +static lean_object* l_Lean_quoteNameMk___closed__2; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__21; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__38; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__18; LEAN_EXPORT lean_object* l_Lean_Name_toStringWithSep_maybeEscape___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Name_modifyBase(lean_object*, lean_object*); static lean_object* l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___closed__2; -static lean_object* l_Lean_instQuoteBool___closed__3; static lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__3; -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__12; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__2; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__27; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__72; +LEAN_EXPORT lean_object* l_Lean_instQuote___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__28; LEAN_EXPORT uint8_t l_Lean_Meta_Rewrite_Config_transparency___default; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__71; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__11; lean_object* l_Nat_repr(lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__7; +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeTSyntaxConsSyntaxNodeKindNil__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapSepElemsM___at_Array_mapSepElems___spec__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_splitNameLitAux___closed__1; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__11; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__12; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__30; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__42; -LEAN_EXPORT lean_object* l_Lean_instQuoteSubstring___boxed(lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__63; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__4; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__13; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__41; +LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailSyntaxTSyntax___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__termEval__prec____1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____lambda__1___closed__1; +static lean_object* l_Lean_mkNullNode___closed__1; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__23; static lean_object* l_Lean_versionStringCore___closed__8; -LEAN_EXPORT lean_object* l_Lean_instQuoteSubstring(lean_object*); -static lean_object* l_Lean_Syntax_expandInterpolatedStr___closed__3; static lean_object* l_Lean_Name_instReprSyntax___closed__1; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__64; LEAN_EXPORT lean_object* l_Lean_Internal_isStage0___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__10; static lean_object* l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__19; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__6; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__53; -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__7; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__34; +LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSyntaxArrayTSepArray(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__14; static lean_object* l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__10; lean_object* l_Lean_Syntax_getId(lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteNameMk(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__12; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_splitNameLitAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_mkSepArray___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__2; +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeNameLitTerm(lean_object*); static lean_object* l_Lean_versionStringCore___closed__4; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeInterpStrQuotedChar___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__92; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__9; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__3; extern lean_object* l_Lean_Parser_Tactic_rwRuleSeq; -extern lean_object* l_Lean_charLitKind; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__86; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__88; +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__6; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__74; +static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__12; static lean_object* l_Lean_Syntax_instBEqSyntax___closed__1; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__15; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__17; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__1; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__89; +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_9542____boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__47; LEAN_EXPORT lean_object* l_Lean_NameGenerator_idx___default; static lean_object* l_Lean_Parser_Tactic_simpAllArith___closed__5; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__46; +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeTSyntaxConsSyntaxNodeKindNil__1___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_EtaStructMode_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__6; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____lambda__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__60; static lean_object* l_Lean_Meta_instBEqTransparencyMode___closed__1; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__37; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_TSyntax_expandInterpolatedStrChunks___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_instQuoteProdMkStrAnonymous___rarg___closed__1; +static lean_object* l_Lean_instQuoteBoolMkStrAnonymous___closed__3; lean_object* lean_array_to_list(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshId(lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__44; lean_object* l_Lean_Macro_throwErrorAt___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Syntax_expandInterpolatedStr___closed__5; extern lean_object* l_Lean_Parser_Tactic_location; LEAN_EXPORT lean_object* l_Lean_Meta_TransparencyMode_toCtorIdx___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkCIdent(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__16; LEAN_EXPORT lean_object* l_Lean_version_getIsRelease___boxed(lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__55; static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__3; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__1; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__16; uint32_t lean_string_utf8_get(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__3; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Name_hasNum___boxed(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__11; LEAN_EXPORT lean_object* l_Lean_Syntax_mkStrLit(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__77; LEAN_EXPORT lean_object* l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__4(lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__39; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__77; static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__4; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__18; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Array_filterSepElemsMAux___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); +static lean_object* l_Lean_Syntax_isCharLit_x3f___closed__2; static lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__subPrio__1___closed__2; -static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__13; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f_decodeAfterExp(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__60; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__2; +static lean_object* l_Lean_instQuoteBoolMkStrAnonymous___closed__7; +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeTSyntaxConsSyntaxNodeKindNil___rarg___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__1; LEAN_EXPORT lean_object* l_Lean_mkOptionalNode(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_version_getPatch___boxed(lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__74; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__70; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079_(lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_quoteNameMk___closed__3; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__81; static lean_object* l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__6; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__79; LEAN_EXPORT lean_object* l_Lean_Syntax_copyHeadTailInfoFrom(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__89; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__9; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__6; +LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSyntaxArray(lean_object*, lean_object*, lean_object*); lean_object* l_Nat_pred(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeHexLitAux___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Array_filterSepElemsMAux(lean_object*); static lean_object* l_Lean_Name_reprPrec___closed__4; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__2; -static lean_object* l___private_Init_Meta_0__Lean_quoteNameMk___closed__5; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__72; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__11; lean_object* l_panic___at___private_Init_Prelude_0__Lean_assembleParts___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__6(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__9; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__23; lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_versionString___closed__10; -static lean_object* l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__4; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__89; +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeCharLitTerm___boxed(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__17; LEAN_EXPORT lean_object* l_Lean_versionString; static lean_object* l_Lean_version_patch___closed__1; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__43; uint8_t l_Array_isEmpty___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15345_(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759_(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14552_(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17812_(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703_(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246_(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17029_(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__47; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__60; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861_(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901_(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18195_(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18994_(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396_(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16479_(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15670_(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__88; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__47; uint8_t l_Substring_beq(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__74; LEAN_EXPORT lean_object* l_Lean_Name_toStringWithSep___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_filterSepElemsM(lean_object*); static lean_object* l_Lean_Name_reprPrec___closed__8; LEAN_EXPORT lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__subPrio__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__13; static lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__2; extern lean_object* l_Lean_instInhabitedSyntax; LEAN_EXPORT lean_object* l_Lean_version_getSpecialDesc(lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__97; static lean_object* l_Lean_toolchain___closed__6; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__41; static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__6; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__4; +static lean_object* l_Lean_instQuoteProdMkStrAnonymous___rarg___closed__3; static lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__addPrio__1___closed__2; static lean_object* l_Lean_version_minor___closed__1; static lean_object* l_Lean_Name_reprPrec___closed__9; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__6; LEAN_EXPORT lean_object* l_Lean_mkCIdentFromRef___rarg___lambda__1(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__6; LEAN_EXPORT lean_object* l_Lean_mkSepArray(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__76; LEAN_EXPORT lean_object* l_Lean_Syntax_SepArray_ofElemsUsingRef___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_Config_maxDischargeDepth___default; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__37; +LEAN_EXPORT lean_object* l_Lean_Syntax_TSepArray_getElems(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_dsimpKind___closed__10; LEAN_EXPORT lean_object* l_Lean_isNumericSubscript___boxed(lean_object*); lean_object* l_Substring_nextn(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__5; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__14; LEAN_EXPORT lean_object* l_Lean_evalOptPrio(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__5; static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__10; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__23; -static lean_object* l___private_Init_Meta_0__Lean_quoteNameMk___closed__2; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__28; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__25; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_TSyntax_expandInterpolatedStrChunks___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Syntax_mkApp___closed__2; +static lean_object* l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___closed__3; LEAN_EXPORT lean_object* l_Lean_NameGenerator_mkChild(lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__24; -static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__6; LEAN_EXPORT lean_object* l_Lean_getGithash___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_instReprEtaStructMode; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__13; +static lean_object* l_Lean_Syntax_isCharLit_x3f___closed__1; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__86; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_updateFirst___at_Lean_Syntax_setHeadInfoAux___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeNumLitTerm___boxed(lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__46; lean_object* l_String_dropRight(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__14; LEAN_EXPORT lean_object* l_Lean_Syntax_isCharLit_x3f___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_splitNameLit(lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__18; static lean_object* l___private_Init_Meta_0__Lean_quoteList___rarg___closed__6; -static lean_object* l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__1; -static lean_object* l___private_Init_Meta_0__Lean_quoteNameMk___closed__7; -static lean_object* l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__4; -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__5; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__38; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__21; +static lean_object* l_Lean_instQuoteSubstringMkStrAnonymous___closed__2; size_t lean_usize_of_nat(lean_object*); -static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__2; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__56; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__94; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__8; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeInterpStrLit(lean_object*); static lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__4; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14552____closed__1; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__68; static lean_object* l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__12; LEAN_EXPORT lean_object* l_Lean_version_minor; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeStrLit___boxed(lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Array_mapSepElemsMAux___at_Array_mapSepElems___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_NameGenerator_namePrefix___default___closed__2; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__24; lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__21; +static lean_object* l_Lean_instQuoteArrayMkStrAnonymous___rarg___closed__1; LEAN_EXPORT uint8_t l_Lean_isIdEndEscape(uint32_t); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____lambda__1___closed__1; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__52; static lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__4; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__50; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_version_getMajor(lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__17; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__66; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Syntax_expandInterpolatedStrChunks___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__44; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__7; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_updateLast___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__13; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeNatLitVal_x3f___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__14; +static lean_object* l_Lean_instQuoteBoolMkStrAnonymous___closed__5; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeCharLit(lean_object*); LEAN_EXPORT uint8_t l_Lean_version_isRelease; -LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Meta_DSimp_beqConfig____x40_Init_Meta___hyg_8947_(lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__3; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__85; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__57; +LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Meta_DSimp_beqConfig____x40_Init_Meta___hyg_9922_(lean_object*, lean_object*); static lean_object* l_Lean_Syntax_decodeNatLitVal_x3f___closed__1; LEAN_EXPORT uint8_t l_Lean_Meta_Simp_Config_arith___default; static lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__5; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__10; LEAN_EXPORT lean_object* l_Lean_Meta_TransparencyMode_noConfusion___rarg(uint8_t, uint8_t, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__1; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__29; static lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__addPrio__1___closed__1; +static lean_object* l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___closed__4; uint8_t l_Char_isAlpha(uint32_t); static lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__8; LEAN_EXPORT lean_object* l_Lean_Option_hasQuote(lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__11; LEAN_EXPORT uint8_t l_Lean_Syntax_isAtom(lean_object*); +LEAN_EXPORT lean_object* l_Lean_instQuoteTermMkStrAnonymous; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Array_getSepElems___spec__1(lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__4; -static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__5; +LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSepArray___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_instQuoteBoolMkStrAnonymous(uint8_t); static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__15; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__37; LEAN_EXPORT uint8_t l_Lean_isLetterLike(uint32_t); LEAN_EXPORT uint8_t l_Lean_Meta_DSimp_Config_beta___default; -static lean_object* l_Lean_instQuoteName___closed__2; -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_beqEtaStructMode____x40_Init_Meta___hyg_8732____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_evalPrec(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_isStrLit_x3f___boxed(lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__80; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__6; lean_object* l_Lean_Macro_expandMacro_x3f(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_simpArith; extern lean_object* l_Lean_Parser_Tactic_simpLemma; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__95; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__2; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__6; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__20; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__34; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__55; lean_object* l_Lean_mkAtomFrom(lean_object*, lean_object*); +lean_object* l_Lean_TSyntaxArray_mkImpl___rarg___boxed(lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__48; static lean_object* l_Lean_termEval__prio_____closed__8; LEAN_EXPORT lean_object* l_Lean_Syntax_mkNumLit(lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedNameGenerator___closed__1; static lean_object* l_Lean_Meta_TransparencyMode_noConfusion___rarg___closed__1; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__64; +LEAN_EXPORT lean_object* l_Lean_Syntax_TSepArray_getElems___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_SepArray_ofElemsUsingRef(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__5; static lean_object* l_Lean_versionString___closed__7; static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__13; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__1; -extern lean_object* l_Lean_groupKind; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__13; LEAN_EXPORT lean_object* l_Lean_Syntax_unsetTrailing(lean_object*); static lean_object* l_Lean_termEval__prio_____closed__7; LEAN_EXPORT lean_object* l_Lean_isLetterLike___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__26; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_isIdOrAtom_x3f(lean_object*); +static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__13; +LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSyntaxArray___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeNumLitPrec___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkIdentFromRef___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__81; -static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__14; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Array_mapSepElemsMAux___at_Array_mapSepElems___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_decodeQuotedChar___boxed__const__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__2; LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__5___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__7; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__50; static lean_object* l_Lean_Parser_Tactic_tacticErw_______closed__4; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__4; static lean_object* l_Lean_Syntax_mkApp___closed__1; uint8_t lean_uint32_dec_eq(uint32_t, uint32_t); +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__24; +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeCharLitTerm(lean_object*); static lean_object* l_Lean_Name_instReprName___closed__1; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__19; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__8; +static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__11; lean_object* l_String_intercalate(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_EtaStructMode_noConfusion(lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__3; static lean_object* l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__20; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__7; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__25; static lean_object* l_Lean_mkHole___closed__8; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__27; +static lean_object* l_Lean_Syntax_instCoeTSepArrayTSyntaxArray___closed__1; +LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSepArrayTSyntaxArray___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Syntax_mkNameLit___closed__1; LEAN_EXPORT uint8_t l_Lean_version_getIsRelease(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__10; LEAN_EXPORT lean_object* lean_name_append_after(lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_quoteNameMk___closed__9; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_version_getMinor(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_getSepArgs(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__1; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__80; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__6; uint8_t l_String_isEmpty(lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__9; static lean_object* l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__16; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__50; +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeTSyntaxConsSyntaxNodeKindNil__1___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_TransparencyMode_noConfusion(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_quoteNameMk___closed__4; +static lean_object* l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__4; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__4; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__4; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__83; static lean_object* l_Lean_Parser_Tactic_tacticErw_______closed__9; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__14; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__28; static lean_object* l_Lean_Parser_Tactic_dsimpKind___closed__7; LEAN_EXPORT lean_object* l_Lean_mkHole(lean_object*); -extern lean_object* l_Lean_scientificLitKind; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__12; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Array_getSepElems___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__59; LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Syntax_structEq___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__20; lean_object* l_Lean_Syntax_setArg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__7; -LEAN_EXPORT lean_object* l_Lean_instQuoteString___boxed(lean_object*); -static lean_object* l_Lean_Syntax_expandInterpolatedStrChunks___closed__1; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__27; -static lean_object* l_Lean_instQuoteBool___closed__2; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__8; +static lean_object* l_Lean_instQuoteSubstringMkStrAnonymous___closed__1; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__29; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__52; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__9; -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Syntax_getHead_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Simp_Config_eta___default; LEAN_EXPORT lean_object* l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__4___boxed(lean_object*); -static lean_object* l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__2; -LEAN_EXPORT lean_object* l_Lean_Syntax_expandInterpolatedStrChunks___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_tacticErw_______closed__8; LEAN_EXPORT uint8_t l_Lean_Syntax_hasArgs(lean_object*); static lean_object* l_Lean_toolchain___closed__1; static lean_object* l_Lean_Parser_Tactic_simpAllArith___closed__1; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__11; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__67; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__40; +static lean_object* l_Lean_TSyntax_getNat___closed__4; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__45; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__30; +LEAN_EXPORT lean_object* l_Lean_instQuoteSubstringMkStrAnonymous___boxed(lean_object*); uint8_t l_String_isPrefixOf(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__9; LEAN_EXPORT lean_object* l_Lean_githash; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList(lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__15; +static lean_object* l_Lean_Syntax_isFieldIdx_x3f___closed__2; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__9; LEAN_EXPORT lean_object* l_Lean_withHeadRefOnly(lean_object*); lean_object* l_String_quote(lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__5; uint8_t l_Char_isAlphanum(uint32_t); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__4; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__4; LEAN_EXPORT lean_object* l_Lean_Syntax_copyHeadTailInfoFrom___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__76; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__32; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__22; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__18; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_instReprConfig; +static lean_object* l_Lean_instQuoteProdMkStrAnonymous___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_instInhabitedNameGenerator; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__17; +LEAN_EXPORT lean_object* l_Lean_instQuoteNameMkStrAnonymous(lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__40; +static lean_object* l_Lean_instQuoteBoolMkStrAnonymous___closed__2; +LEAN_EXPORT uint8_t l_Lean_Syntax_instBEqTSyntax___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__67; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__9; static lean_object* l_Lean_mkHole___closed__4; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__10; -static lean_object* l_Lean_evalPrec___closed__2; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__80; +static lean_object* l_Lean_TSyntax_getNat___closed__2; lean_object* l_Lean_Syntax_getArgs(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__8; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__86; +LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSepArray___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f(lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__6; LEAN_EXPORT uint8_t l_Lean_isGreek(uint32_t); LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_tacticErw____; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__17; lean_object* l_Lean_Syntax_getKind(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllArith___closed__8; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__53; LEAN_EXPORT lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__subPrec__1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__1; lean_object* l_Lean_MacroScopesView_review(lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__20; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__3; +LEAN_EXPORT lean_object* l_Lean_quoteNameMk(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeInterpStrQuotedChar___boxed__const__1; -static lean_object* l___private_Init_Meta_0__Lean_quoteNameMk___closed__8; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__88; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__8; static lean_object* l_Lean_toolchain___closed__2; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__13; -static lean_object* l___private_Init_Meta_0__Lean_quoteNameMk___closed__1; LEAN_EXPORT lean_object* l_Array_filterSepElems(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapSepElemsM___at_Array_mapSepElems___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_instQuoteArrayMkStrAnonymous___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__53; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__42; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__22; +static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__1; +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621_(lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Name_hasNum(lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__55; +static lean_object* l_Lean_Syntax_mkScientificLit___closed__1; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__93; LEAN_EXPORT lean_object* l_Lean_Meta_instBEqEtaStructMode; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__17; +static lean_object* l_Lean_Syntax_mkScientificLit___closed__2; static lean_object* l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__10; -static lean_object* l_Lean_instQuoteSubstring___closed__4; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__72; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__96; static lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__9; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__23; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__9; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Name_isInaccessibleUserName___boxed(lean_object*); static uint8_t l_Lean_versionString___closed__3; +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__7; LEAN_EXPORT lean_object* l_Lean_Name_replacePrefix___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__91; static lean_object* l_Lean_versionString___closed__5; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__78; static lean_object* l_Lean_termEval__prec_____closed__11; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__57; LEAN_EXPORT lean_object* l_Lean_Syntax_getOptionalIdent_x3f(lean_object*); static lean_object* l_Lean_evalPrec___closed__1; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__41; +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeTSyntaxConsSyntaxNodeKindNil___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Name_toStringWithSep___closed__1; LEAN_EXPORT lean_object* l_Lean_NameGenerator_curr(lean_object*); static lean_object* l_Lean_Parser_Tactic_dsimpKind___closed__3; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__4; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__4; LEAN_EXPORT lean_object* l_Lean_Syntax_SepArray_getElems___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_DSimp_instBEqConfig; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__7; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__67; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__9; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__40; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__10; -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583_(uint8_t, lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__1; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_TSyntax_expandInterpolatedStrChunks___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__17; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__12; +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558_(uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_isNameLit_x3f___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Array_mapSepElemsMAux___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__1; static lean_object* l_Lean_mkHole___closed__1; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__20; -static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__7; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__64; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15670____closed__2; static lean_object* l_Lean_Syntax_mkApp___closed__3; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__19; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__8; static lean_object* l_Lean_versionStringCore___closed__6; LEAN_EXPORT lean_object* l_Lean_isIdBeginEscape___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic_tacticErw_______closed__6; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__54; static lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__2; LEAN_EXPORT lean_object* l_Lean_mkFreshId___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Rewrite_Config_offsetCnstrs___default; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Array_filterSepElemsMAux___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__25; LEAN_EXPORT uint8_t l_Lean_Meta_DSimp_Config_etaStruct___default; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__62; static lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__subPrio__1___closed__1; LEAN_EXPORT lean_object* l_List_beq___at_Lean_Syntax_structEq___spec__2___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__1; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__23; lean_object* lean_nat_mul(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__12; static lean_object* l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__15; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__91; LEAN_EXPORT lean_object* l_Lean_Syntax_mkNameLit(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__19; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__27; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__11; static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__9; static lean_object* l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__8; LEAN_EXPORT lean_object* l_Lean_mkCIdentFromRef___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_decodeQuotedChar___boxed__const__4; +static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__6; static lean_object* l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___closed__1; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__78; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__93; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeQuotedChar___boxed__const__1; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__12; +static lean_object* l_Lean_TSyntax_expandInterpolatedStr___closed__1; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__22; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__13; +static lean_object* l_Lean_instQuoteBoolMkStrAnonymous___closed__6; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__5; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__29; +LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSyntaxArrayTSepArray___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Meta_instReprEtaStructMode___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_getTailInfo_x3f(lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__92; +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__2; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_instBEqConfig; static lean_object* l_Lean_termEval__prec_____closed__2; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__21; static lean_object* l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__4; +LEAN_EXPORT lean_object* l_Lean_Syntax_instBEqTSyntax___rarg___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__3; LEAN_EXPORT uint8_t l_Lean_Meta_Simp_Config_iota___default; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__31; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__76; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__16; +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_setInfo(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__13; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__30; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__59; -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteOption(lean_object*); +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeTSyntaxConsSyntaxNodeKindNil__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__addPrec__1(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_instQuoteNat(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__9; +static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__15; +LEAN_EXPORT lean_object* l_panic___at_Lean_TSyntax_getNat___spec__1(lean_object*); +static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__18; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_TSyntax_expandInterpolatedStrChunks___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_neutralConfig___closed__1; static lean_object* l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__3; LEAN_EXPORT uint8_t l_Lean_Syntax_isNone(lean_object*); lean_object* lean_name_mk_numeral(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_TSyntax_getId___boxed(lean_object*); static lean_object* l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__17; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__28; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__10; LEAN_EXPORT lean_object* l_List_beq___at_Lean_Syntax_structEq___spec__3___boxed(lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__4; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f_decodeAfterDot___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__9; static lean_object* l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__7; LEAN_EXPORT lean_object* l_Lean_Name_toString_maybePseudoSyntax___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_expandMacros(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__50; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_updateLast(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__3; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeNameLit(lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__20; -static lean_object* l_Lean_instQuoteArray___rarg___closed__1; static lean_object* l_Lean_Name_toString_maybePseudoSyntax___closed__1; +LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSepArray(lean_object*); +static lean_object* l_Lean_TSyntax_expandInterpolatedStr___closed__4; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__15; static lean_object* l_Lean_mkOptionalNode___closed__1; static lean_object* l_Lean_Meta_DSimp_instReprConfig___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_isNameLit_x3f(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__25; +static lean_object* l_Lean_quoteNameMk___closed__7; LEAN_EXPORT lean_object* l_Lean_Syntax_isInterpolatedStrLit_x3f(lean_object*); +lean_object* lean_panic_fn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_getHead_x3f___lambda__1___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__21; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__38; +static lean_object* l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__2; lean_object* l_String_trim(lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__24; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeQuotedChar___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isIdEndEscape___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_mkSepArray___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__4; LEAN_EXPORT lean_object* l_Lean_Syntax_getTailInfo_x3f___boxed(lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__3; LEAN_EXPORT lean_object* l_Lean_Syntax_getTailInfo(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_getOptionalIdent_x3f___boxed(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__5; +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeStrLitTerm___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__18; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__14; -static lean_object* l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__3; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__79; +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__8; LEAN_EXPORT lean_object* l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_termEval__prec_____closed__5; -static lean_object* l_Lean_instQuoteBool___closed__4; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic____________; static lean_object* l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__24; static lean_object* l___private_Init_Meta_0__Lean_quoteList___rarg___closed__7; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__15; +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__30; LEAN_EXPORT lean_object* l_Lean_mkCIdentFromRef(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__19; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__49; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__95; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeStrLit(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__1; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__18; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__76; LEAN_EXPORT uint8_t l_Lean_isIdFirst(uint32_t); -LEAN_EXPORT lean_object* l_Lean_instQuoteName(lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__16; uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__21; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__24; LEAN_EXPORT lean_object* l_Lean_Syntax_isNone___boxed(lean_object*); -static lean_object* l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__2; +LEAN_EXPORT lean_object* l_Lean_TSyntax_getString(lean_object*); +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeNameLitTerm___boxed(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__15; +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_beqEtaStructMode____x40_Init_Meta___hyg_9707____boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__18; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__1; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__52; LEAN_EXPORT lean_object* l_Lean_Syntax_isToken___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__82; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__6; +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__11; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__3; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__15; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__14; static lean_object* l_Lean_mkHole___closed__2; extern lean_object* l_Lean_Parser_Tactic_simpErase; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__6; static lean_object* l_Lean_mkCIdentFrom___closed__1; static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__1; +LEAN_EXPORT lean_object* l_Lean_Syntax_TSepArray_getElems___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_getTailInfo___boxed(lean_object*); +static lean_object* l_Lean_TSyntax_expandInterpolatedStr___closed__5; +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeIdentTerm___boxed(lean_object*); static lean_object* l_Lean_termEval__prec_____closed__8; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeQuotedChar___boxed__const__3; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__3; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__85; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__2; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__23; static lean_object* l_Lean_toolchain___closed__7; static lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__addPrec__1___closed__3; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__56; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__66; LEAN_EXPORT lean_object* l_Lean_Syntax_isInterpolatedStrLit_x3f___boxed(lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__75; -LEAN_EXPORT lean_object* l_Lean_Syntax_SepArray_instCoeTailSepArrayArraySyntax___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_origin; +static lean_object* l_Lean_Syntax_isInterpolatedStrLit_x3f___closed__2; LEAN_EXPORT lean_object* l_Lean_Name_instReprSyntax; -static lean_object* l_Lean_instQuoteArray___rarg___closed__2; +LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSyntaxArray___rarg___boxed(lean_object*); lean_object* lean_string_length(lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__58; -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_beqConfig____x40_Init_Meta___hyg_9383____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Syntax_findAux___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -static lean_object* l_Lean_Syntax_SepArray_instCoeTailSepArrayArraySyntax___closed__1; static lean_object* l_Lean_mkHole___closed__6; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__18; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__26; LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__Substring_takeWhileAux___at___private_Init_Meta_0__Lean_Syntax_splitNameLitAux___spec__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_tacticErw_______closed__10; LEAN_EXPORT uint8_t l_Lean_isSubScriptAlnum(uint32_t); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__19; static lean_object* l_Lean_Name_reprPrec___closed__10; LEAN_EXPORT uint8_t l_Lean_Meta_Simp_Config_beta___default; -static lean_object* l_Lean_Syntax_expandInterpolatedStr___closed__4; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_getHead_x3f___lambda__1(lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__8; +static lean_object* l_Lean_Option_hasQuote___rarg___closed__3; LEAN_EXPORT uint8_t l_List_beq___at_Lean_Syntax_structEq___spec__2(lean_object*, lean_object*); +static lean_object* l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_withHeadRefOnly___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__5; LEAN_EXPORT lean_object* lean_mk_syntax_ident(lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__70; extern lean_object* l_Lean_Parser_Tactic_discharger; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_TransparencyMode_noConfusion___rarg___lambda__1(lean_object*); -static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__1; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__71; static lean_object* l_Lean_Syntax_getHead_x3f___closed__1; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_updateFirst___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__42; static lean_object* l_Lean_Parser_Tactic_simpArith___closed__7; +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__9; LEAN_EXPORT lean_object* l_Lean_Syntax_mkSynthetic(lean_object*); +static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__3; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__93; static lean_object* l_Lean_termEval__prec_____closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_toNat(lean_object*); static lean_object* l_Lean_versionString___closed__6; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__37; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__53; static lean_object* l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__11; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__1; LEAN_EXPORT lean_object* lean_name_append_before(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_version_getPatch(lean_object*); LEAN_EXPORT lean_object* l_Lean_toolchain; static lean_object* l_Lean_mkOptionalNode___closed__2; -LEAN_EXPORT lean_object* l_Lean_Syntax_getOptional_x3f___boxed(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__17; +static lean_object* l_Lean_instQuoteSubstringMkStrAnonymous___closed__4; static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__4; static lean_object* l_Lean_mkCIdentFrom___closed__2; static lean_object* l_Lean_origin___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold; -LEAN_EXPORT lean_object* l_Lean_instQuoteArray___rarg(lean_object*, lean_object*); lean_object* lean_nat_mod(lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__14; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__33; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__37; +static lean_object* l_Lean_TSyntax_getNat___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__18; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__61; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__55; +LEAN_EXPORT lean_object* l_Lean_instQuoteListMkStrAnonymous(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__6; lean_object* l_Lean_SourceInfo_getPos_x3f(lean_object*, uint8_t); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Array_filterSepElemsMAux___at_Array_filterSepElems___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__5; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__1; static lean_object* l_Lean_Meta_Simp_instBEqConfig___closed__1; -LEAN_EXPORT lean_object* l_Lean_instQuoteProd___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshId___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__17; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__4; static lean_object* l_Lean_termEval__prec_____closed__4; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__45; LEAN_EXPORT uint8_t l_Lean_Meta_Simp_Config_proj___default; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_instBEqTransparencyMode; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__18; LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeArraySyntaxSepArray(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__13; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__28; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__16; static lean_object* l_Lean_Parser_Tactic_tacticErw_______closed__11; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__5; +static lean_object* l_Lean_Option_hasQuote___rarg___closed__5; +LEAN_EXPORT lean_object* l_Lean_instQuoteStringStrLitKind___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_find_x3f(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__83; static lean_object* l_Lean_Parser_Tactic_tacticErw_______closed__12; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__90; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__5; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__63; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeStrLitAux___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__14; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__11; uint8_t l_List_isEmpty___rarg(lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__12; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__8; lean_object* l_String_drop(lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__22; -LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Meta_beqEtaStructMode____x40_Init_Meta___hyg_8732_(uint8_t, uint8_t); -static lean_object* l_Lean_Syntax_expandInterpolatedStr___closed__2; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__7; +LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Meta_beqEtaStructMode____x40_Init_Meta___hyg_9707_(uint8_t, uint8_t); +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeTSyntaxConsSyntaxNodeKindNil___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_filterSepElemsM___at_Array_filterSepElems___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__9; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__17; LEAN_EXPORT lean_object* l_Lean_monadNameGeneratorLift___rarg___lambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__71; +static lean_object* l_Lean_TSyntax_expandInterpolatedStr___closed__2; static lean_object* l_Lean_Parser_Tactic_simpArith___closed__1; lean_object* l_Nat_min(lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__6; -static lean_object* l_Lean_Syntax_expandInterpolatedStr___closed__1; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__19; static lean_object* l___private_Init_Meta_0__Lean_quoteList___rarg___closed__4; -LEAN_EXPORT lean_object* l_Lean_Syntax_expandInterpolatedStr___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_instQuoteArrayMkStrAnonymous(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__33; static lean_object* l_Lean_mkSepArray___closed__1; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__3; static lean_object* l_Lean_Parser_Tactic_simpArith___closed__5; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__8; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__36; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__15; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__10; static lean_object* l_Lean_Parser_Tactic_simpArith___closed__3; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__32; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__73; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_TSyntax_expandInterpolatedStrChunks___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__1; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__12; +LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSyntaxArrayArraySyntax___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Name_getRoot___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeBinLitAux___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Name_getRoot(lean_object*); uint8_t lean_uint32_dec_le(uint32_t, uint32_t); static lean_object* l_Lean_Name_reprPrec___closed__5; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__22; -static lean_object* l_Lean_evalPrec___closed__3; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__85; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__79; LEAN_EXPORT lean_object* l_Lean_mkSepArray___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14552____closed__2; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__43; +static lean_object* l_Lean_Syntax_instCoeTSyntaxArrayArraySyntax___closed__1; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__5; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15670____closed__3; static lean_object* l_Lean_Name_reprPrec___closed__3; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__46; static lean_object* l___private_Init_Meta_0__Lean_quoteList___rarg___closed__2; LEAN_EXPORT uint8_t l_Lean_Name_toString_maybePseudoSyntax(lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__26; +LEAN_EXPORT lean_object* l_Lean_TSyntax_expandInterpolatedStr___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_beqConfig____x40_Init_Meta___hyg_10358____boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__2; -static lean_object* l_Lean_instQuoteName___closed__1; static lean_object* l_Lean_termEval__prec_____closed__9; -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__14; LEAN_EXPORT lean_object* l_Lean_Meta_TransparencyMode_noConfusion___rarg___lambda__1___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeNumLitPrec(lean_object*); static lean_object* l_Array_getSepElems___rarg___closed__1; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__26; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Array_getSepElems___spec__1___rarg(lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__2; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__95; static lean_object* l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__14; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__16; static lean_object* l_Lean_Parser_Tactic_simpAllArith___closed__7; LEAN_EXPORT lean_object* l_Lean_Syntax_findAux(lean_object*, lean_object*); static lean_object* l_Lean_Name_escapePart___closed__3; +LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSyntaxArrayTSepArray___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_decodeStrLitAux(lean_object*, lean_object*, lean_object*); static lean_object* l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__2; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__49; uint8_t lean_string_utf8_at_end(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_instQuoteSyntax; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__28; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__2; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__7; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__14; LEAN_EXPORT lean_object* l_Lean_mkNode(lean_object*, lean_object*); static lean_object* l___private_Init_Meta_0__Lean_quoteList___rarg___closed__1; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__78; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__91; lean_object* lean_uint32_to_nat(uint32_t); static lean_object* l_Array_forInUnsafe_loop___at_Lean_mkSepArray___spec__1___closed__1; +static lean_object* l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___closed__1; LEAN_EXPORT lean_object* l_Lean_Name_replacePrefix(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_termEval__prio_____closed__5; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__2; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__6; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__73; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeInterpStrLit_loop(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__subPrec__1___closed__2; static lean_object* l_Lean_Parser_Tactic_simpArith___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_simpAllKind; +LEAN_EXPORT lean_object* l_Lean_TSyntax_expandInterpolatedStr___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__18; -static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__8; static lean_object* l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___closed__4; lean_object* lean_nat_to_int(lean_object*); LEAN_EXPORT uint8_t l_Lean_Syntax_isToken(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__48; LEAN_EXPORT lean_object* l_Lean_Syntax_isFieldIdx_x3f___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__24; LEAN_EXPORT uint8_t l_Lean_Meta_Simp_Config_autoUnfold___default; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__92; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeInterpStrLit___boxed(lean_object*); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__19; +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__12; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__18; LEAN_EXPORT lean_object* l_Lean_evalPrio(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__32; static lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_defaultMaxSteps; -static lean_object* l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__4; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__65; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__16; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14552____closed__3; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__10; +static lean_object* l_Lean_instQuoteArrayMkStrAnonymous___rarg___closed__2; static lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__addPrec__1___closed__4; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__34; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__5; +LEAN_EXPORT lean_object* l_Lean_instQuoteProdMkStrAnonymous(lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__14; LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__Substring_takeWhileAux___at___private_Init_Meta_0__Lean_Syntax_splitNameLitAux___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(lean_object*, size_t, size_t, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__8; +LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSyntaxArray___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_setHeadInfoAux(lean_object*, lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__22; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__20; static lean_object* l_Lean_Parser_Tactic_simpArith___closed__9; static lean_object* l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__9; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__30; -LEAN_EXPORT lean_object* l_Lean_Option_hasQuote___rarg(lean_object*); +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__77; +LEAN_EXPORT lean_object* l_Lean_Option_hasQuote___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__8; -extern lean_object* l_Lean_interpolatedStrLitKind; static lean_object* l_Lean_termEval__prio_____closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__75; +static lean_object* l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSyntaxArray___closed__1; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__4; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__84; static lean_object* l_Lean_Meta_Simp_instInhabitedConfig___closed__1; +static lean_object* l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__3; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__54; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__3; +static lean_object* l_Lean_instQuoteBoolMkStrAnonymous___closed__1; static lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__addPrec__1___closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__29; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__3; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_updateLast___at_Lean_Syntax_setTailInfoAux___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Syntax_isFieldIdx_x3f___closed__1; LEAN_EXPORT uint8_t l_Lean_isIdRest(uint32_t); -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__51; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__35; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__43; static lean_object* l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__5; static lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_DSimp_instReprConfig; +LEAN_EXPORT lean_object* l_Lean_instQuoteBoolMkStrAnonymous___boxed(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__27; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__66; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__10; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__56; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__33; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__9; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__22; uint8_t lean_string_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_mkLit(lean_object*, lean_object*, lean_object*); lean_object* l_Char_ofNat(lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Array_mapSepElemsMAux___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_instQuoteNameMkStrAnonymous___closed__2; static lean_object* l_Lean_versionStringCore___closed__7; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__20; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__62; lean_object* l_Repr_addAppParen(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpArith___closed__2; LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__Substring_takeWhileAux___at___private_Init_Meta_0__Lean_Syntax_splitNameLitAux___spec__1(lean_object*, lean_object*, lean_object*); @@ -4356,6 +4431,132 @@ lean_dec(x_2); return x_4; } } +LEAN_EXPORT lean_object* l_Lean_Name_eraseSuffix_x3f(lean_object* x_1, lean_object* x_2) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 0: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +else +{ +lean_object* x_4; +lean_dec(x_2); +x_4 = lean_box(0); +return x_4; +} +} +case 1: +{ +switch (lean_obj_tag(x_2)) { +case 0: +{ +lean_object* x_5; +x_5 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_5, 0, x_1); +return x_5; +} +case 1: +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_6 = lean_ctor_get(x_1, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_1, 1); +lean_inc(x_7); +lean_dec(x_1); +x_8 = lean_ctor_get(x_2, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_2, 1); +lean_inc(x_9); +lean_dec(x_2); +x_10 = lean_string_dec_eq(x_7, x_9); +lean_dec(x_9); +lean_dec(x_7); +if (x_10 == 0) +{ +lean_object* x_11; +lean_dec(x_8); +lean_dec(x_6); +x_11 = lean_box(0); +return x_11; +} +else +{ +x_1 = x_6; +x_2 = x_8; +goto _start; +} +} +default: +{ +lean_object* x_13; +lean_dec(x_2); +lean_dec(x_1); +x_13 = lean_box(0); +return x_13; +} +} +} +default: +{ +switch (lean_obj_tag(x_2)) { +case 0: +{ +lean_object* x_14; +x_14 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, x_1); +return x_14; +} +case 1: +{ +lean_object* x_15; +lean_dec(x_2); +lean_dec(x_1); +x_15 = lean_box(0); +return x_15; +} +default: +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_16 = lean_ctor_get(x_1, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_1, 1); +lean_inc(x_17); +lean_dec(x_1); +x_18 = lean_ctor_get(x_2, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_2, 1); +lean_inc(x_19); +lean_dec(x_2); +x_20 = lean_nat_dec_eq(x_17, x_19); +lean_dec(x_19); +lean_dec(x_17); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_18); +lean_dec(x_16); +x_21 = lean_box(0); +return x_21; +} +else +{ +x_1 = x_16; +x_2 = x_18; +goto _start; +} +} +} +} +} +} +} LEAN_EXPORT lean_object* l_Lean_Name_modifyBase(lean_object* x_1, lean_object* x_2) { _start: { @@ -5147,126 +5348,362 @@ x_3 = lean_alloc_closure((void*)(l_Lean_monadNameGeneratorLift___rarg), 2, 0); return x_3; } } -LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Syntax_structEq___spec__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_EXPORT lean_object* l_Lean_TSyntax_instCoeTSyntaxConsSyntaxNodeKindNil___rarg(lean_object* x_1) { _start: { -lean_object* x_7; uint8_t x_8; -x_7 = lean_array_get_size(x_4); -x_8 = lean_nat_dec_lt(x_6, x_7); -lean_dec(x_7); -if (x_8 == 0) -{ -uint8_t x_9; -lean_dec(x_6); -x_9 = 1; -return x_9; +lean_inc(x_1); +return x_1; } -else -{ -lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = lean_array_fget(x_4, x_6); -x_11 = lean_array_fget(x_5, x_6); -x_12 = l_Lean_Syntax_structEq(x_10, x_11); -if (x_12 == 0) -{ -uint8_t x_13; -lean_dec(x_6); -x_13 = 0; -return x_13; } -else +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeTSyntaxConsSyntaxNodeKindNil(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_14; lean_object* x_15; -x_14 = lean_unsigned_to_nat(1u); -x_15 = lean_nat_add(x_6, x_14); -lean_dec(x_6); -x_3 = lean_box(0); -x_6 = x_15; -goto _start; +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_TSyntax_instCoeTSyntaxConsSyntaxNodeKindNil___rarg___boxed), 1, 0); +return x_3; } } +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeTSyntaxConsSyntaxNodeKindNil___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_TSyntax_instCoeTSyntaxConsSyntaxNodeKindNil___rarg(x_1); +lean_dec(x_1); +return x_2; } } -LEAN_EXPORT uint8_t l_List_beq___at_Lean_Syntax_structEq___spec__3(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeTSyntaxConsSyntaxNodeKindNil___boxed(lean_object* x_1, lean_object* x_2) { _start: { -if (lean_obj_tag(x_1) == 0) -{ -if (lean_obj_tag(x_2) == 0) -{ -uint8_t x_3; -x_3 = 1; +lean_object* x_3; +x_3 = l_Lean_TSyntax_instCoeTSyntaxConsSyntaxNodeKindNil(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); return x_3; } -else +} +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeTSyntaxConsSyntaxNodeKindNil__1___rarg(lean_object* x_1) { +_start: { -uint8_t x_4; -x_4 = 0; -return x_4; +lean_inc(x_1); +return x_1; } } -else -{ -if (lean_obj_tag(x_2) == 0) +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeTSyntaxConsSyntaxNodeKindNil__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -uint8_t x_5; -x_5 = 0; +lean_object* x_5; +x_5 = lean_alloc_closure((void*)(l_Lean_TSyntax_instCoeTSyntaxConsSyntaxNodeKindNil__1___rarg___boxed), 1, 0); return x_5; } -else -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; -x_6 = lean_ctor_get(x_1, 0); -x_7 = lean_ctor_get(x_1, 1); -x_8 = lean_ctor_get(x_2, 0); -x_9 = lean_ctor_get(x_2, 1); -x_10 = lean_string_dec_eq(x_6, x_8); -if (x_10 == 0) -{ -uint8_t x_11; -x_11 = 0; -return x_11; } -else +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeTSyntaxConsSyntaxNodeKindNil__1___rarg___boxed(lean_object* x_1) { +_start: { -x_1 = x_7; -x_2 = x_9; -goto _start; -} +lean_object* x_2; +x_2 = l_Lean_TSyntax_instCoeTSyntaxConsSyntaxNodeKindNil__1___rarg(x_1); +lean_dec(x_1); +return x_2; } } +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeTSyntaxConsSyntaxNodeKindNil__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_TSyntax_instCoeTSyntaxConsSyntaxNodeKindNil__1(x_1, x_2, x_3, x_4); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_5; } } -LEAN_EXPORT uint8_t l_List_beq___at_Lean_Syntax_structEq___spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeIdentTerm(lean_object* x_1) { _start: { -if (lean_obj_tag(x_1) == 0) -{ -if (lean_obj_tag(x_2) == 0) +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeIdentTerm___boxed(lean_object* x_1) { +_start: { -uint8_t x_3; -x_3 = 1; -return x_3; +lean_object* x_2; +x_2 = l_Lean_TSyntax_instCoeIdentTerm(x_1); +lean_dec(x_1); +return x_2; } -else +} +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeDepTermMkConsSyntaxNodeKindMkStrAnonymousNilIdentIdent(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -uint8_t x_4; -x_4 = 0; -return x_4; +lean_object* x_5; +x_5 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_2); +lean_ctor_set(x_5, 2, x_3); +lean_ctor_set(x_5, 3, x_4); +return x_5; } } -else +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeStrLitTerm(lean_object* x_1) { +_start: { -if (lean_obj_tag(x_2) == 0) +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeStrLitTerm___boxed(lean_object* x_1) { +_start: { -uint8_t x_5; -x_5 = 0; -return x_5; +lean_object* x_2; +x_2 = l_Lean_TSyntax_instCoeStrLitTerm(x_1); +lean_dec(x_1); +return x_2; } -else +} +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeNameLitTerm(lean_object* x_1) { +_start: { -lean_object* x_6; lean_object* x_7; lean_object* x_8; 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_6 = lean_ctor_get(x_1, 0); +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeNameLitTerm___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_TSyntax_instCoeNameLitTerm(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeNumLitTerm(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeNumLitTerm___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_TSyntax_instCoeNumLitTerm(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeCharLitTerm(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeCharLitTerm___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_TSyntax_instCoeCharLitTerm(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeNumLitPrec(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeNumLitPrec___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_TSyntax_instCoeNumLitPrec(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailSyntaxTSyntax___rarg(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailSyntaxTSyntax(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_TSyntax_Compat_instCoeTailSyntaxTSyntax___rarg___boxed), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailSyntaxTSyntax___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_TSyntax_Compat_instCoeTailSyntaxTSyntax___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailSyntaxTSyntax___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_TSyntax_Compat_instCoeTailSyntaxTSyntax(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSyntaxArray___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_TSyntaxArray_mkImpl___rarg___boxed), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSyntaxArray(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSyntaxArray___closed__1; +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSyntaxArray___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSyntaxArray(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Syntax_structEq___spec__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: +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_array_get_size(x_4); +x_8 = lean_nat_dec_lt(x_6, x_7); +lean_dec(x_7); +if (x_8 == 0) +{ +uint8_t x_9; +lean_dec(x_6); +x_9 = 1; +return x_9; +} +else +{ +lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_10 = lean_array_fget(x_4, x_6); +x_11 = lean_array_fget(x_5, x_6); +x_12 = l_Lean_Syntax_structEq(x_10, x_11); +if (x_12 == 0) +{ +uint8_t x_13; +lean_dec(x_6); +x_13 = 0; +return x_13; +} +else +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_unsigned_to_nat(1u); +x_15 = lean_nat_add(x_6, x_14); +lean_dec(x_6); +x_3 = lean_box(0); +x_6 = x_15; +goto _start; +} +} +} +} +LEAN_EXPORT uint8_t l_List_beq___at_Lean_Syntax_structEq___spec__3(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_3; +x_3 = 1; +return x_3; +} +else +{ +uint8_t x_4; +x_4 = 0; +return x_4; +} +} +else +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_5; +x_5 = 0; +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_6 = lean_ctor_get(x_1, 0); +x_7 = lean_ctor_get(x_1, 1); +x_8 = lean_ctor_get(x_2, 0); +x_9 = lean_ctor_get(x_2, 1); +x_10 = lean_string_dec_eq(x_6, x_8); +if (x_10 == 0) +{ +uint8_t x_11; +x_11 = 0; +return x_11; +} +else +{ +x_1 = x_7; +x_2 = x_9; +goto _start; +} +} +} +} +} +LEAN_EXPORT uint8_t l_List_beq___at_Lean_Syntax_structEq___spec__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_3; +x_3 = 1; +return x_3; +} +else +{ +uint8_t x_4; +x_4 = 0; +return x_4; +} +} +else +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_5; +x_5 = 0; +return x_5; +} +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; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_6 = lean_ctor_get(x_1, 0); x_7 = lean_ctor_get(x_2, 0); x_8 = lean_ctor_get(x_1, 1); x_9 = lean_ctor_get(x_2, 1); @@ -5535,6 +5972,40 @@ x_1 = l_Lean_Syntax_instBEqSyntax___closed__1; return x_1; } } +LEAN_EXPORT uint8_t l_Lean_Syntax_instBEqTSyntax___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = l_Lean_Syntax_structEq(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Syntax_instBEqTSyntax(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Syntax_instBEqTSyntax___rarg___boxed), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Syntax_instBEqTSyntax___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_Syntax_instBEqTSyntax___rarg(x_1, x_2); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Syntax_instBEqTSyntax___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Syntax_instBEqTSyntax(x_1); +lean_dec(x_1); +return x_2; +} +} LEAN_EXPORT lean_object* l_Array_findSomeRevM_x3f_find___at_Lean_Syntax_getTailInfo_x3f___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { @@ -7867,25 +8338,61 @@ lean_ctor_set(x_9, 3, x_7); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_mkNullNode(lean_object* x_1) { +static lean_object* _init_l_Lean_mkNullNode___closed__1() { _start: { -lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = lean_box(2); -x_3 = l_Lean_nullKind; -x_4 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_4, 0, x_2); -lean_ctor_set(x_4, 1, x_3); -lean_ctor_set(x_4, 2, x_1); -return x_4; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("null", 4); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_mkGroupNode(lean_object* x_1) { +static lean_object* _init_l_Lean_mkNullNode___closed__2() { _start: { -lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = lean_box(2); -x_3 = l_Lean_groupKind; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_mkNullNode___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_mkNullNode(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = lean_box(2); +x_3 = l_Lean_mkNullNode___closed__2; +x_4 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_4, 0, x_2); +lean_ctor_set(x_4, 1, x_3); +lean_ctor_set(x_4, 2, x_1); +return x_4; +} +} +static lean_object* _init_l_Lean_mkGroupNode___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("group", 5); +return x_1; +} +} +static lean_object* _init_l_Lean_mkGroupNode___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_mkGroupNode___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_mkGroupNode(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = lean_box(2); +x_3 = l_Lean_mkGroupNode___closed__2; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_2); lean_ctor_set(x_4, 1, x_3); @@ -8072,7 +8579,7 @@ static lean_object* _init_l_Lean_mkOptionalNode___closed__1() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean_nullKind; +x_2 = l_Lean_mkNullNode___closed__2; x_3 = l_Lean_mkSepArray___closed__1; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -8108,7 +8615,7 @@ lean_dec(x_1); x_4 = l_Lean_mkOptionalNode___closed__2; x_5 = lean_array_push(x_4, x_3); x_6 = lean_box(2); -x_7 = l_Lean_nullKind; +x_7 = l_Lean_mkNullNode___closed__2; x_8 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_8, 0, x_6); lean_ctor_set(x_8, 1, x_7); @@ -8212,7 +8719,7 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_mkSep(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_3 = l_Lean_mkSepArray(x_1, x_2); x_4 = lean_box(2); -x_5 = l_Lean_nullKind; +x_5 = l_Lean_mkNullNode___closed__2; x_6 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_6, 0, x_4); lean_ctor_set(x_6, 1, x_5); @@ -8337,6 +8844,44 @@ lean_closure_set(x_2, 0, x_1); return x_2; } } +LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSyntaxArrayTSepArray___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_box(2); +x_4 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_1); +x_5 = l_Lean_mkSepArray(x_2, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSyntaxArrayTSepArray(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Syntax_instCoeTSyntaxArrayTSepArray___rarg___boxed), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSyntaxArrayTSepArray___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Syntax_instCoeTSyntaxArrayTSepArray___rarg(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSyntaxArrayTSepArray___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Syntax_instCoeTSyntaxArrayTSepArray(x_1); +lean_dec(x_1); +return x_2; +} +} static lean_object* _init_l_Lean_Syntax_mkApp___closed__1() { _start: { @@ -8376,7 +8921,7 @@ if (x_5 == 0) { lean_object* x_6; lean_object* x_7; 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_6 = lean_box(2); -x_7 = l_Lean_nullKind; +x_7 = l_Lean_mkNullNode___closed__2; x_8 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_8, 0, x_6); lean_ctor_set(x_8, 1, x_7); @@ -8425,12 +8970,30 @@ lean_ctor_set(x_8, 2, x_6); return x_8; } } +static lean_object* _init_l_Lean_Syntax_mkStrLit___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("str", 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Syntax_mkStrLit___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Syntax_mkStrLit___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Syntax_mkStrLit(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; x_3 = l_String_quote(x_1); -x_4 = l_Lean_strLitKind; +x_4 = l_Lean_Syntax_mkStrLit___closed__2; x_5 = l_Lean_Syntax_mkLit(x_4, x_3, x_2); return x_5; } @@ -8444,29 +9007,83 @@ lean_dec(x_1); return x_3; } } +static lean_object* _init_l_Lean_Syntax_mkNumLit___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("num", 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Syntax_mkNumLit___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Syntax_mkNumLit___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Syntax_mkNumLit(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; -x_3 = l_Lean_numLitKind; +x_3 = l_Lean_Syntax_mkNumLit___closed__2; x_4 = l_Lean_Syntax_mkLit(x_3, x_1, x_2); return x_4; } } +static lean_object* _init_l_Lean_Syntax_mkScientificLit___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("scientific", 10); +return x_1; +} +} +static lean_object* _init_l_Lean_Syntax_mkScientificLit___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Syntax_mkScientificLit___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Syntax_mkScientificLit(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; -x_3 = l_Lean_scientificLitKind; +x_3 = l_Lean_Syntax_mkScientificLit___closed__2; x_4 = l_Lean_Syntax_mkLit(x_3, x_1, x_2); return x_4; } } +static lean_object* _init_l_Lean_Syntax_mkNameLit___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("name", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Syntax_mkNameLit___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Syntax_mkNameLit___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Syntax_mkNameLit(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; -x_3 = l_Lean_nameLitKind; +x_3 = l_Lean_Syntax_mkNameLit___closed__2; x_4 = l_Lean_Syntax_mkLit(x_3, x_1, x_2); return x_4; } @@ -9193,7 +9810,7 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_isNatLit_x3f(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_numLitKind; +x_2 = l_Lean_Syntax_mkNumLit___closed__2; x_3 = l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(x_2, x_1); return x_3; } @@ -9207,11 +9824,29 @@ lean_dec(x_1); return x_2; } } +static lean_object* _init_l_Lean_Syntax_isFieldIdx_x3f___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("fieldIdx", 8); +return x_1; +} +} +static lean_object* _init_l_Lean_Syntax_isFieldIdx_x3f___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Syntax_isFieldIdx_x3f___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Syntax_isFieldIdx_x3f(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_fieldIdxKind; +x_2 = l_Lean_Syntax_isFieldIdx_x3f___closed__2; x_3 = l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(x_2, x_1); return x_3; } @@ -9744,7 +10379,7 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_isScientificLit_x3f(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_scientificLitKind; +x_2 = l_Lean_Syntax_mkScientificLit___closed__2; x_3 = l_Lean_Syntax_isLit_x3f(x_2, x_1); if (lean_obj_tag(x_3) == 0) { @@ -9819,22 +10454,21 @@ return x_2; LEAN_EXPORT lean_object* l_Lean_Syntax_toNat(lean_object* x_1) { _start: { -lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_numLitKind; -x_3 = l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(x_2, x_1); -if (lean_obj_tag(x_3) == 0) +lean_object* x_2; +x_2 = l_Lean_Syntax_isNatLit_x3f(x_1); +if (lean_obj_tag(x_2) == 0) { -lean_object* x_4; -x_4 = lean_unsigned_to_nat(0u); -return x_4; +lean_object* x_3; +x_3 = lean_unsigned_to_nat(0u); +return x_3; } else { -lean_object* x_5; -x_5 = lean_ctor_get(x_3, 0); -lean_inc(x_5); -lean_dec(x_3); -return x_5; +lean_object* x_4; +x_4 = lean_ctor_get(x_2, 0); +lean_inc(x_4); +lean_dec(x_2); +return x_4; } } } @@ -10474,7 +11108,7 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_isStrLit_x3f(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_strLitKind; +x_2 = l_Lean_Syntax_mkStrLit___closed__2; x_3 = l_Lean_Syntax_isLit_x3f(x_2, x_1); if (lean_obj_tag(x_3) == 0) { @@ -10572,11 +11206,29 @@ lean_dec(x_1); return x_2; } } +static lean_object* _init_l_Lean_Syntax_isCharLit_x3f___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("char", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Syntax_isCharLit_x3f___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Syntax_isCharLit_x3f___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Syntax_isCharLit_x3f(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_charLitKind; +x_2 = l_Lean_Syntax_isCharLit_x3f___closed__2; x_3 = l_Lean_Syntax_isLit_x3f(x_2, x_1); if (lean_obj_tag(x_3) == 0) { @@ -11423,7 +12075,7 @@ static lean_object* _init_l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__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; x_1 = l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___closed__1; x_2 = l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___closed__2; -x_3 = lean_unsigned_to_nat(723u); +x_3 = lean_unsigned_to_nat(795u); x_4 = lean_unsigned_to_nat(12u); x_5 = l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -11567,7 +12219,7 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_isNameLit_x3f(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_nameLitKind; +x_2 = l_Lean_Syntax_mkNameLit___closed__2; x_3 = l_Lean_Syntax_isLit_x3f(x_2, x_1); if (lean_obj_tag(x_3) == 0) { @@ -11701,7 +12353,7 @@ case 1: lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; x_3 = lean_ctor_get(x_1, 1); x_4 = lean_ctor_get(x_1, 2); -x_5 = l_Lean_nullKind; +x_5 = l_Lean_mkNullNode___closed__2; x_6 = lean_name_eq(x_3, x_5); if (x_6 == 0) { @@ -11738,64 +12390,6 @@ x_3 = lean_box(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Syntax_getOptional_x3f(lean_object* x_1) { -_start: -{ -if (lean_obj_tag(x_1) == 1) -{ -lean_object* x_2; lean_object* x_3; lean_object* x_4; uint8_t x_5; -x_2 = lean_ctor_get(x_1, 1); -x_3 = lean_ctor_get(x_1, 2); -x_4 = l_Lean_nullKind; -x_5 = lean_name_eq(x_2, x_4); -if (x_5 == 0) -{ -lean_object* x_6; -x_6 = lean_box(0); -return x_6; -} -else -{ -lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_7 = lean_array_get_size(x_3); -x_8 = lean_unsigned_to_nat(1u); -x_9 = lean_nat_dec_eq(x_7, x_8); -lean_dec(x_7); -if (x_9 == 0) -{ -lean_object* x_10; -x_10 = lean_box(0); -return x_10; -} -else -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_11 = l_Lean_instInhabitedSyntax; -x_12 = lean_unsigned_to_nat(0u); -x_13 = lean_array_get(x_11, x_3, x_12); -x_14 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_14, 0, x_13); -return x_14; -} -} -} -else -{ -lean_object* x_15; -x_15 = lean_box(0); -return x_15; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Syntax_getOptional_x3f___boxed(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = l_Lean_Syntax_getOptional_x3f(x_1); -lean_dec(x_1); -return x_2; -} -} LEAN_EXPORT lean_object* l_Lean_Syntax_getOptionalIdent_x3f(lean_object* x_1) { _start: { @@ -12015,124 +12609,307 @@ x_3 = l_Lean_Syntax_findAux(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_instQuoteSyntax___closed__1() { +LEAN_EXPORT lean_object* l_panic___at_Lean_TSyntax_getNat___spec__1(lean_object* x_1) { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_id___rarg___boxed), 1, 0); -return x_1; +lean_object* x_2; lean_object* x_3; +x_2 = l_instInhabitedNat; +x_3 = lean_panic_fn(x_2, x_1); +return x_3; } } -static lean_object* _init_l_Lean_instQuoteSyntax() { +static lean_object* _init_l_Lean_TSyntax_getNat___closed__1() { _start: { lean_object* x_1; -x_1 = l_Lean_instQuoteSyntax___closed__1; +x_1 = lean_mk_string_from_bytes("Init.Data.Option.BasicAux", 25); return x_1; } } -static lean_object* _init_l_Lean_instQuoteBool___closed__1() { +static lean_object* _init_l_Lean_TSyntax_getNat___closed__2() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Bool", 4); +x_1 = lean_mk_string_from_bytes("Option.get!", 11); return x_1; } } -static lean_object* _init_l_Lean_instQuoteBool___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_instQuoteBool___closed__1; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_instQuoteBool___closed__3() { +static lean_object* _init_l_Lean_TSyntax_getNat___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("false", 5); +x_1 = lean_mk_string_from_bytes("value is none", 13); return x_1; } } -static lean_object* _init_l_Lean_instQuoteBool___closed__4() { +static lean_object* _init_l_Lean_TSyntax_getNat___closed__4() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_instQuoteBool___closed__2; -x_2 = l_Lean_instQuoteBool___closed__3; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; +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_TSyntax_getNat___closed__1; +x_2 = l_Lean_TSyntax_getNat___closed__2; +x_3 = lean_unsigned_to_nat(16u); +x_4 = lean_unsigned_to_nat(14u); +x_5 = l_Lean_TSyntax_getNat___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; } } -static lean_object* _init_l_Lean_instQuoteBool___closed__5() { +LEAN_EXPORT lean_object* l_Lean_TSyntax_getNat(lean_object* x_1) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_instQuoteBool___closed__4; -x_3 = l_Lean_mkCIdentFrom(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_instQuoteBool___closed__6() { -_start: +lean_object* x_2; +x_2 = l_Lean_Syntax_isNatLit_x3f(x_1); +lean_dec(x_1); +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; lean_object* x_4; +x_3 = l_Lean_TSyntax_getNat___closed__4; +x_4 = l_panic___at_Lean_TSyntax_getNat___spec__1(x_3); +return x_4; +} +else +{ +lean_object* x_5; +x_5 = lean_ctor_get(x_2, 0); +lean_inc(x_5); +lean_dec(x_2); +return x_5; +} +} +} +LEAN_EXPORT lean_object* l_Lean_TSyntax_getId(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Syntax_getId(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_TSyntax_getId___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_TSyntax_getId(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_panic___at_Lean_TSyntax_getString___spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_String_instInhabitedString; +x_3 = lean_panic_fn(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_TSyntax_getString(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Syntax_isStrLit_x3f(x_1); +lean_dec(x_1); +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; lean_object* x_4; +x_3 = l_Lean_TSyntax_getNat___closed__4; +x_4 = l_panic___at_Lean_TSyntax_getString___spec__1(x_3); +return x_4; +} +else +{ +lean_object* x_5; +x_5 = lean_ctor_get(x_2, 0); +lean_inc(x_5); +lean_dec(x_2); +return x_5; +} +} +} +LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSepArray___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_box(2); +x_4 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_1); +x_5 = l_Lean_mkSepArray(x_2, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSepArray(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSepArray___rarg___boxed), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSepArray___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSepArray___rarg(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSepArray___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSepArray(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_instQuote___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_apply_1(x_1, x_3); +x_5 = lean_apply_1(x_2, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_instQuote(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Lean_instQuote___rarg), 3, 0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_instQuote___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_instQuote(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +return x_4; +} +} +static lean_object* _init_l_Lean_instQuoteTermMkStrAnonymous___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_id___rarg___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_instQuoteTermMkStrAnonymous() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_instQuoteTermMkStrAnonymous___closed__1; +return x_1; +} +} +static lean_object* _init_l_Lean_instQuoteBoolMkStrAnonymous___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Bool", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_instQuoteBoolMkStrAnonymous___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_instQuoteBoolMkStrAnonymous___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_instQuoteBoolMkStrAnonymous___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("false", 5); +return x_1; +} +} +static lean_object* _init_l_Lean_instQuoteBoolMkStrAnonymous___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_instQuoteBoolMkStrAnonymous___closed__2; +x_2 = l_Lean_instQuoteBoolMkStrAnonymous___closed__3; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_instQuoteBoolMkStrAnonymous___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_instQuoteBoolMkStrAnonymous___closed__4; +x_3 = l_Lean_mkCIdentFrom(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_instQuoteBoolMkStrAnonymous___closed__6() { +_start: { lean_object* x_1; x_1 = lean_mk_string_from_bytes("true", 4); return x_1; } } -static lean_object* _init_l_Lean_instQuoteBool___closed__7() { +static lean_object* _init_l_Lean_instQuoteBoolMkStrAnonymous___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_instQuoteBool___closed__2; -x_2 = l_Lean_instQuoteBool___closed__6; +x_1 = l_Lean_instQuoteBoolMkStrAnonymous___closed__2; +x_2 = l_Lean_instQuoteBoolMkStrAnonymous___closed__6; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_instQuoteBool___closed__8() { +static lean_object* _init_l_Lean_instQuoteBoolMkStrAnonymous___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_instQuoteBool___closed__7; +x_2 = l_Lean_instQuoteBoolMkStrAnonymous___closed__7; x_3 = l_Lean_mkCIdentFrom(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_instQuoteBool(uint8_t x_1) { +LEAN_EXPORT lean_object* l_Lean_instQuoteBoolMkStrAnonymous(uint8_t x_1) { _start: { if (x_1 == 0) { lean_object* x_2; -x_2 = l_Lean_instQuoteBool___closed__5; +x_2 = l_Lean_instQuoteBoolMkStrAnonymous___closed__5; return x_2; } else { lean_object* x_3; -x_3 = l_Lean_instQuoteBool___closed__8; +x_3 = l_Lean_instQuoteBoolMkStrAnonymous___closed__8; return x_3; } } } -LEAN_EXPORT lean_object* l_Lean_instQuoteBool___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_instQuoteBoolMkStrAnonymous___boxed(lean_object* x_1) { _start: { uint8_t x_2; lean_object* x_3; x_2 = lean_unbox(x_1); lean_dec(x_1); -x_3 = l_Lean_instQuoteBool(x_2); +x_3 = l_Lean_instQuoteBoolMkStrAnonymous(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_instQuoteString(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_instQuoteStringStrLitKind(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; @@ -12141,27 +12918,26 @@ x_3 = l_Lean_Syntax_mkStrLit(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_instQuoteString___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_instQuoteStringStrLitKind___boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_instQuoteString(x_1); +x_2 = l_Lean_instQuoteStringStrLitKind(x_1); lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_instQuoteNat(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_instQuoteNatNumLitKind(lean_object* x_1) { _start: { -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +lean_object* x_2; lean_object* x_3; lean_object* x_4; x_2 = l_Nat_repr(x_1); -x_3 = l_Lean_numLitKind; -x_4 = lean_box(2); -x_5 = l_Lean_Syntax_mkLit(x_3, x_2, x_4); -return x_5; +x_3 = lean_box(2); +x_4 = l_Lean_Syntax_mkNumLit(x_2, x_3); +return x_4; } } -static lean_object* _init_l_Lean_instQuoteSubstring___closed__1() { +static lean_object* _init_l_Lean_instQuoteSubstringMkStrAnonymous___closed__1() { _start: { lean_object* x_1; @@ -12169,17 +12945,17 @@ x_1 = lean_mk_string_from_bytes("String", 6); return x_1; } } -static lean_object* _init_l_Lean_instQuoteSubstring___closed__2() { +static lean_object* _init_l_Lean_instQuoteSubstringMkStrAnonymous___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_instQuoteSubstring___closed__1; +x_2 = l_Lean_instQuoteSubstringMkStrAnonymous___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_instQuoteSubstring___closed__3() { +static lean_object* _init_l_Lean_instQuoteSubstringMkStrAnonymous___closed__3() { _start: { lean_object* x_1; @@ -12187,17 +12963,17 @@ x_1 = lean_mk_string_from_bytes("toSubstring", 11); return x_1; } } -static lean_object* _init_l_Lean_instQuoteSubstring___closed__4() { +static lean_object* _init_l_Lean_instQuoteSubstringMkStrAnonymous___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_instQuoteSubstring___closed__2; -x_2 = l_Lean_instQuoteSubstring___closed__3; +x_1 = l_Lean_instQuoteSubstringMkStrAnonymous___closed__2; +x_2 = l_Lean_instQuoteSubstringMkStrAnonymous___closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_instQuoteSubstring(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_instQuoteSubstringMkStrAnonymous(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_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; @@ -12210,16 +12986,16 @@ x_7 = l_Lean_Syntax_mkStrLit(x_5, x_6); lean_dec(x_5); x_8 = l_Lean_mkOptionalNode___closed__2; x_9 = lean_array_push(x_8, x_7); -x_10 = l_Lean_instQuoteSubstring___closed__4; +x_10 = l_Lean_instQuoteSubstringMkStrAnonymous___closed__4; x_11 = l_Lean_Syntax_mkCApp(x_10, x_9); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_instQuoteSubstring___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_instQuoteSubstringMkStrAnonymous___boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_instQuoteSubstring(x_1); +x_2 = l_Lean_instQuoteSubstringMkStrAnonymous(x_1); lean_dec(x_1); return x_2; } @@ -12289,7 +13065,7 @@ return x_13; } } } -static lean_object* _init_l___private_Init_Meta_0__Lean_quoteNameMk___closed__1() { +static lean_object* _init_l_Lean_quoteNameMk___closed__1() { _start: { lean_object* x_1; @@ -12297,17 +13073,17 @@ x_1 = lean_mk_string_from_bytes("Name", 4); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_quoteNameMk___closed__2() { +static lean_object* _init_l_Lean_quoteNameMk___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_mkHole___closed__2; -x_2 = l___private_Init_Meta_0__Lean_quoteNameMk___closed__1; +x_2 = l_Lean_quoteNameMk___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_quoteNameMk___closed__3() { +static lean_object* _init_l_Lean_quoteNameMk___closed__3() { _start: { lean_object* x_1; @@ -12315,27 +13091,27 @@ x_1 = lean_mk_string_from_bytes("anonymous", 9); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_quoteNameMk___closed__4() { +static lean_object* _init_l_Lean_quoteNameMk___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_quoteNameMk___closed__2; -x_2 = l___private_Init_Meta_0__Lean_quoteNameMk___closed__3; +x_1 = l_Lean_quoteNameMk___closed__2; +x_2 = l_Lean_quoteNameMk___closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_quoteNameMk___closed__5() { +static lean_object* _init_l_Lean_quoteNameMk___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Init_Meta_0__Lean_quoteNameMk___closed__4; +x_2 = l_Lean_quoteNameMk___closed__4; x_3 = l_Lean_mkCIdentFrom(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_quoteNameMk___closed__6() { +static lean_object* _init_l_Lean_quoteNameMk___closed__6() { _start: { lean_object* x_1; @@ -12343,17 +13119,17 @@ x_1 = lean_mk_string_from_bytes("mkStr", 5); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_quoteNameMk___closed__7() { +static lean_object* _init_l_Lean_quoteNameMk___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_quoteNameMk___closed__2; -x_2 = l___private_Init_Meta_0__Lean_quoteNameMk___closed__6; +x_1 = l_Lean_quoteNameMk___closed__2; +x_2 = l_Lean_quoteNameMk___closed__6; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_quoteNameMk___closed__8() { +static lean_object* _init_l_Lean_quoteNameMk___closed__8() { _start: { lean_object* x_1; @@ -12361,24 +13137,24 @@ x_1 = lean_mk_string_from_bytes("mkNum", 5); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_quoteNameMk___closed__9() { +static lean_object* _init_l_Lean_quoteNameMk___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_quoteNameMk___closed__2; -x_2 = l___private_Init_Meta_0__Lean_quoteNameMk___closed__8; +x_1 = l_Lean_quoteNameMk___closed__2; +x_2 = l_Lean_quoteNameMk___closed__8; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteNameMk(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_quoteNameMk(lean_object* x_1) { _start: { switch (lean_obj_tag(x_1)) { case 0: { lean_object* x_2; -x_2 = l___private_Init_Meta_0__Lean_quoteNameMk___closed__5; +x_2 = l_Lean_quoteNameMk___closed__5; return x_2; } case 1: @@ -12389,41 +13165,40 @@ lean_inc(x_3); x_4 = lean_ctor_get(x_1, 1); lean_inc(x_4); lean_dec(x_1); -x_5 = l___private_Init_Meta_0__Lean_quoteNameMk(x_3); +x_5 = l_Lean_quoteNameMk(x_3); x_6 = lean_box(2); x_7 = l_Lean_Syntax_mkStrLit(x_4, x_6); lean_dec(x_4); x_8 = l_Lean_Syntax_mkApp___closed__3; x_9 = lean_array_push(x_8, x_5); x_10 = lean_array_push(x_9, x_7); -x_11 = l___private_Init_Meta_0__Lean_quoteNameMk___closed__7; +x_11 = l_Lean_quoteNameMk___closed__7; x_12 = l_Lean_Syntax_mkCApp(x_11, x_10); return x_12; } default: { -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; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +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; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; x_13 = lean_ctor_get(x_1, 0); lean_inc(x_13); x_14 = lean_ctor_get(x_1, 1); lean_inc(x_14); lean_dec(x_1); -x_15 = l___private_Init_Meta_0__Lean_quoteNameMk(x_13); +x_15 = l_Lean_quoteNameMk(x_13); x_16 = l_Nat_repr(x_14); -x_17 = l_Lean_numLitKind; -x_18 = lean_box(2); -x_19 = l_Lean_Syntax_mkLit(x_17, x_16, x_18); -x_20 = l_Lean_Syntax_mkApp___closed__3; -x_21 = lean_array_push(x_20, x_15); -x_22 = lean_array_push(x_21, x_19); -x_23 = l___private_Init_Meta_0__Lean_quoteNameMk___closed__9; -x_24 = l_Lean_Syntax_mkCApp(x_23, x_22); -return x_24; +x_17 = lean_box(2); +x_18 = l_Lean_Syntax_mkNumLit(x_16, x_17); +x_19 = l_Lean_Syntax_mkApp___closed__3; +x_20 = lean_array_push(x_19, x_15); +x_21 = lean_array_push(x_20, x_18); +x_22 = l_Lean_quoteNameMk___closed__9; +x_23 = l_Lean_Syntax_mkCApp(x_22, x_21); +return x_23; } } } } -static lean_object* _init_l_Lean_instQuoteName___closed__1() { +static lean_object* _init_l_Lean_instQuoteNameMkStrAnonymous___closed__1() { _start: { lean_object* x_1; @@ -12431,17 +13206,17 @@ x_1 = lean_mk_string_from_bytes("quotedName", 10); return x_1; } } -static lean_object* _init_l_Lean_instQuoteName___closed__2() { +static lean_object* _init_l_Lean_instQuoteNameMkStrAnonymous___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_mkHole___closed__6; -x_2 = l_Lean_instQuoteName___closed__1; +x_2 = l_Lean_instQuoteNameMkStrAnonymous___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_instQuoteName(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_instQuoteNameMkStrAnonymous(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; @@ -12451,12 +13226,12 @@ x_3 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_2, x_1); if (lean_obj_tag(x_3) == 0) { lean_object* x_4; -x_4 = l___private_Init_Meta_0__Lean_quoteNameMk(x_1); +x_4 = l_Lean_quoteNameMk(x_1); return x_4; } else { -lean_object* x_5; lean_object* x_6; lean_object* x_7; 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; lean_object* x_15; lean_object* x_16; +lean_object* x_5; lean_object* x_6; lean_object* x_7; 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; lean_object* x_15; lean_dec(x_1); x_5 = lean_ctor_get(x_3, 0); lean_inc(x_5); @@ -12466,21 +13241,20 @@ x_7 = l_String_intercalate(x_6, x_5); x_8 = l_Lean_Name_reprPrec___closed__3; x_9 = lean_string_append(x_8, x_7); lean_dec(x_7); -x_10 = l_Lean_nameLitKind; -x_11 = lean_box(2); -x_12 = l_Lean_Syntax_mkLit(x_10, x_9, x_11); -x_13 = l_Lean_mkOptionalNode___closed__2; -x_14 = lean_array_push(x_13, x_12); -x_15 = l_Lean_instQuoteName___closed__2; -x_16 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_16, 0, x_11); -lean_ctor_set(x_16, 1, x_15); -lean_ctor_set(x_16, 2, x_14); -return x_16; +x_10 = lean_box(2); +x_11 = l_Lean_Syntax_mkNameLit(x_9, x_10); +x_12 = l_Lean_mkOptionalNode___closed__2; +x_13 = lean_array_push(x_12, x_11); +x_14 = l_Lean_instQuoteNameMkStrAnonymous___closed__2; +x_15 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_15, 0, x_10); +lean_ctor_set(x_15, 1, x_14); +lean_ctor_set(x_15, 2, x_13); +return x_15; } } } -static lean_object* _init_l_Lean_instQuoteProd___rarg___closed__1() { +static lean_object* _init_l_Lean_instQuoteProdMkStrAnonymous___rarg___closed__1() { _start: { lean_object* x_1; @@ -12488,17 +13262,17 @@ x_1 = lean_mk_string_from_bytes("Prod", 4); return x_1; } } -static lean_object* _init_l_Lean_instQuoteProd___rarg___closed__2() { +static lean_object* _init_l_Lean_instQuoteProdMkStrAnonymous___rarg___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_instQuoteProd___rarg___closed__1; +x_2 = l_Lean_instQuoteProdMkStrAnonymous___rarg___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_instQuoteProd___rarg___closed__3() { +static lean_object* _init_l_Lean_instQuoteProdMkStrAnonymous___rarg___closed__3() { _start: { lean_object* x_1; @@ -12506,17 +13280,17 @@ x_1 = lean_mk_string_from_bytes("mk", 2); return x_1; } } -static lean_object* _init_l_Lean_instQuoteProd___rarg___closed__4() { +static lean_object* _init_l_Lean_instQuoteProdMkStrAnonymous___rarg___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_instQuoteProd___rarg___closed__2; -x_2 = l_Lean_instQuoteProd___rarg___closed__3; +x_1 = l_Lean_instQuoteProdMkStrAnonymous___rarg___closed__2; +x_2 = l_Lean_instQuoteProdMkStrAnonymous___rarg___closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_instQuoteProd___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_instQuoteProdMkStrAnonymous___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { 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; lean_object* x_11; lean_object* x_12; @@ -12530,16 +13304,16 @@ x_7 = lean_apply_1(x_2, x_5); x_8 = l_Lean_Syntax_mkApp___closed__3; x_9 = lean_array_push(x_8, x_6); x_10 = lean_array_push(x_9, x_7); -x_11 = l_Lean_instQuoteProd___rarg___closed__4; +x_11 = l_Lean_instQuoteProdMkStrAnonymous___rarg___closed__4; x_12 = l_Lean_Syntax_mkCApp(x_11, x_10); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_instQuoteProd(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_instQuoteProdMkStrAnonymous(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_instQuoteProd___rarg), 3, 0); +x_3 = lean_alloc_closure((void*)(l_Lean_instQuoteProdMkStrAnonymous___rarg), 3, 0); return x_3; } } @@ -12645,7 +13419,7 @@ x_2 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_quoteList___rarg) return x_2; } } -LEAN_EXPORT lean_object* l_Lean_instQuoteList___rarg(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_instQuoteListMkStrAnonymous___rarg(lean_object* x_1) { _start: { lean_object* x_2; @@ -12654,15 +13428,15 @@ lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_instQuoteList(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_instQuoteListMkStrAnonymous(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_instQuoteList___rarg), 1, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_instQuoteListMkStrAnonymous___rarg), 1, 0); return x_2; } } -static lean_object* _init_l_Lean_instQuoteArray___rarg___closed__1() { +static lean_object* _init_l_Lean_instQuoteArrayMkStrAnonymous___rarg___closed__1() { _start: { lean_object* x_1; @@ -12670,17 +13444,17 @@ x_1 = lean_mk_string_from_bytes("toArray", 7); return x_1; } } -static lean_object* _init_l_Lean_instQuoteArray___rarg___closed__2() { +static lean_object* _init_l_Lean_instQuoteArrayMkStrAnonymous___rarg___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_quoteList___rarg___closed__2; -x_2 = l_Lean_instQuoteArray___rarg___closed__1; +x_2 = l_Lean_instQuoteArrayMkStrAnonymous___rarg___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_instQuoteArray___rarg(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_instQuoteArrayMkStrAnonymous___rarg(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; @@ -12688,20 +13462,20 @@ x_3 = lean_array_to_list(lean_box(0), x_2); x_4 = l___private_Init_Meta_0__Lean_quoteList___rarg(x_1, x_3); x_5 = l_Lean_mkOptionalNode___closed__2; x_6 = lean_array_push(x_5, x_4); -x_7 = l_Lean_instQuoteArray___rarg___closed__2; +x_7 = l_Lean_instQuoteArrayMkStrAnonymous___rarg___closed__2; x_8 = l_Lean_Syntax_mkCApp(x_7, x_6); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_instQuoteArray(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_instQuoteArrayMkStrAnonymous(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_instQuoteArray___rarg), 2, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_instQuoteArrayMkStrAnonymous___rarg), 2, 0); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__1() { +static lean_object* _init_l_Lean_Option_hasQuote___rarg___closed__1() { _start: { lean_object* x_1; @@ -12709,17 +13483,17 @@ x_1 = lean_mk_string_from_bytes("Option", 6); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__2() { +static lean_object* _init_l_Lean_Option_hasQuote___rarg___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__1; +x_2 = l_Lean_Option_hasQuote___rarg___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__3() { +static lean_object* _init_l_Lean_Option_hasQuote___rarg___closed__3() { _start: { lean_object* x_1; @@ -12727,26 +13501,26 @@ x_1 = lean_mk_string_from_bytes("none", 4); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__4() { +static lean_object* _init_l_Lean_Option_hasQuote___rarg___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__2; -x_2 = l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__3; +x_1 = l_Lean_Option_hasQuote___rarg___closed__2; +x_2 = l_Lean_Option_hasQuote___rarg___closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__5() { +static lean_object* _init_l_Lean_Option_hasQuote___rarg___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__4; +x_1 = l_Lean_Option_hasQuote___rarg___closed__4; x_2 = lean_mk_syntax_ident(x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__6() { +static lean_object* _init_l_Lean_Option_hasQuote___rarg___closed__6() { _start: { lean_object* x_1; @@ -12754,24 +13528,24 @@ x_1 = lean_mk_string_from_bytes("some", 4); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__7() { +static lean_object* _init_l_Lean_Option_hasQuote___rarg___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__2; -x_2 = l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__6; +x_1 = l_Lean_Option_hasQuote___rarg___closed__2; +x_2 = l_Lean_Option_hasQuote___rarg___closed__6; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteOption___rarg(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Option_hasQuote___rarg(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) { lean_object* x_3; lean_dec(x_1); -x_3 = l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__5; +x_3 = l_Lean_Option_hasQuote___rarg___closed__5; return x_3; } else @@ -12783,34 +13557,17 @@ lean_dec(x_2); x_5 = lean_apply_1(x_1, x_4); x_6 = l_Lean_mkOptionalNode___closed__2; x_7 = lean_array_push(x_6, x_5); -x_8 = l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__7; +x_8 = l_Lean_Option_hasQuote___rarg___closed__7; x_9 = l_Lean_Syntax_mkCApp(x_8, x_7); return x_9; } } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteOption(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_quoteOption___rarg), 2, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Option_hasQuote___rarg(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_quoteOption___rarg), 2, 1); -lean_closure_set(x_2, 0, x_1); -return x_2; -} -} LEAN_EXPORT lean_object* l_Lean_Option_hasQuote(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Option_hasQuote___rarg), 1, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Option_hasQuote___rarg), 2, 0); return x_2; } } @@ -12818,24 +13575,6 @@ static lean_object* _init_l_Lean_evalPrec___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("num", 3); -return x_1; -} -} -static lean_object* _init_l_Lean_evalPrec___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_evalPrec___closed__1; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_evalPrec___closed__3() { -_start: -{ -lean_object* x_1; x_1 = lean_mk_string_from_bytes("unexpected precedence", 21); return x_1; } @@ -12873,119 +13612,85 @@ if (x_15 == 0) lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; x_16 = lean_ctor_get(x_14, 0); x_17 = lean_ctor_get(x_14, 1); -x_18 = l_Lean_evalPrec___closed__2; +x_18 = l_Lean_Syntax_mkNumLit___closed__2; lean_inc(x_16); x_19 = l_Lean_Syntax_isOfKind(x_16, x_18); if (x_19 == 0) { lean_object* x_20; lean_object* x_21; lean_free_object(x_14); -x_20 = l_Lean_evalPrec___closed__3; +x_20 = l_Lean_evalPrec___closed__1; x_21 = l_Lean_Macro_throwErrorAt___rarg(x_16, x_20, x_2, x_17); lean_dec(x_16); return x_21; } else { -lean_object* x_22; lean_object* x_23; +lean_object* x_22; lean_dec(x_2); -x_22 = l_Lean_numLitKind; -x_23 = l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(x_22, x_16); -lean_dec(x_16); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; -x_24 = lean_unsigned_to_nat(0u); -lean_ctor_set(x_14, 0, x_24); -return x_14; -} -else -{ -lean_object* x_25; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -lean_ctor_set(x_14, 0, x_25); +x_22 = l_Lean_TSyntax_getNat(x_16); +lean_ctor_set(x_14, 0, x_22); return x_14; } } -} else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_26 = lean_ctor_get(x_14, 0); -x_27 = lean_ctor_get(x_14, 1); -lean_inc(x_27); -lean_inc(x_26); +lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_23 = lean_ctor_get(x_14, 0); +x_24 = lean_ctor_get(x_14, 1); +lean_inc(x_24); +lean_inc(x_23); lean_dec(x_14); -x_28 = l_Lean_evalPrec___closed__2; -lean_inc(x_26); -x_29 = l_Lean_Syntax_isOfKind(x_26, x_28); -if (x_29 == 0) +x_25 = l_Lean_Syntax_mkNumLit___closed__2; +lean_inc(x_23); +x_26 = l_Lean_Syntax_isOfKind(x_23, x_25); +if (x_26 == 0) { -lean_object* x_30; lean_object* x_31; -x_30 = l_Lean_evalPrec___closed__3; -x_31 = l_Lean_Macro_throwErrorAt___rarg(x_26, x_30, x_2, x_27); -lean_dec(x_26); -return x_31; +lean_object* x_27; lean_object* x_28; +x_27 = l_Lean_evalPrec___closed__1; +x_28 = l_Lean_Macro_throwErrorAt___rarg(x_23, x_27, x_2, x_24); +lean_dec(x_23); +return x_28; } else { -lean_object* x_32; lean_object* x_33; +lean_object* x_29; lean_object* x_30; lean_dec(x_2); -x_32 = l_Lean_numLitKind; -x_33 = l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(x_32, x_26); -lean_dec(x_26); -if (lean_obj_tag(x_33) == 0) -{ -lean_object* x_34; lean_object* x_35; -x_34 = lean_unsigned_to_nat(0u); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_27); -return x_35; -} -else -{ -lean_object* x_36; lean_object* x_37; -x_36 = lean_ctor_get(x_33, 0); -lean_inc(x_36); -lean_dec(x_33); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_27); -return x_37; -} +x_29 = l_Lean_TSyntax_getNat(x_23); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_24); +return x_30; } } } else { -uint8_t x_38; +uint8_t x_31; lean_dec(x_2); -x_38 = !lean_is_exclusive(x_14); -if (x_38 == 0) +x_31 = !lean_is_exclusive(x_14); +if (x_31 == 0) { return x_14; } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_39 = lean_ctor_get(x_14, 0); -x_40 = lean_ctor_get(x_14, 1); -lean_inc(x_40); -lean_inc(x_39); -lean_dec(x_14); -x_41 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_40); -return x_41; +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_14, 0); +x_33 = lean_ctor_get(x_14, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_14); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; } } } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_free_object(x_2); lean_dec(x_10); lean_dec(x_9); @@ -12993,155 +13698,133 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_42 = l_Lean_maxRecDepthErrorMessage; -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_1); -lean_ctor_set(x_43, 1, x_42); -x_44 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_44, 1, x_3); -return x_44; +x_35 = l_Lean_maxRecDepthErrorMessage; +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_1); +lean_ctor_set(x_36, 1, x_35); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_3); +return x_37; } } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; -x_45 = lean_ctor_get(x_2, 0); -x_46 = lean_ctor_get(x_2, 1); -x_47 = lean_ctor_get(x_2, 2); -x_48 = lean_ctor_get(x_2, 3); -x_49 = lean_ctor_get(x_2, 4); -x_50 = lean_ctor_get(x_2, 5); -lean_inc(x_50); -lean_inc(x_49); -lean_inc(x_48); -lean_inc(x_47); -lean_inc(x_46); -lean_inc(x_45); +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; +x_38 = lean_ctor_get(x_2, 0); +x_39 = lean_ctor_get(x_2, 1); +x_40 = lean_ctor_get(x_2, 2); +x_41 = lean_ctor_get(x_2, 3); +x_42 = lean_ctor_get(x_2, 4); +x_43 = lean_ctor_get(x_2, 5); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_inc(x_39); +lean_inc(x_38); lean_dec(x_2); -x_51 = lean_nat_dec_eq(x_48, x_49); -if (x_51 == 0) +x_44 = lean_nat_dec_eq(x_41, x_42); +if (x_44 == 0) { -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_52 = lean_unsigned_to_nat(1u); -x_53 = lean_nat_add(x_48, x_52); -lean_dec(x_48); -x_54 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_54, 0, x_45); -lean_ctor_set(x_54, 1, x_46); -lean_ctor_set(x_54, 2, x_47); -lean_ctor_set(x_54, 3, x_53); -lean_ctor_set(x_54, 4, x_49); -lean_ctor_set(x_54, 5, x_50); -lean_inc(x_54); -x_55 = l_Lean_expandMacros(x_1, x_54, x_3); -if (lean_obj_tag(x_55) == 0) +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_45 = lean_unsigned_to_nat(1u); +x_46 = lean_nat_add(x_41, x_45); +lean_dec(x_41); +x_47 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_47, 0, x_38); +lean_ctor_set(x_47, 1, x_39); +lean_ctor_set(x_47, 2, x_40); +lean_ctor_set(x_47, 3, x_46); +lean_ctor_set(x_47, 4, x_42); +lean_ctor_set(x_47, 5, x_43); +lean_inc(x_47); +x_48 = l_Lean_expandMacros(x_1, x_47, x_3); +if (lean_obj_tag(x_48) == 0) { -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; uint8_t x_60; -x_56 = lean_ctor_get(x_55, 0); -lean_inc(x_56); -x_57 = lean_ctor_get(x_55, 1); -lean_inc(x_57); -if (lean_is_exclusive(x_55)) { - lean_ctor_release(x_55, 0); - lean_ctor_release(x_55, 1); - x_58 = x_55; +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; +x_49 = lean_ctor_get(x_48, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_48, 1); +lean_inc(x_50); +if (lean_is_exclusive(x_48)) { + lean_ctor_release(x_48, 0); + lean_ctor_release(x_48, 1); + x_51 = x_48; } else { - lean_dec_ref(x_55); - x_58 = lean_box(0); -} -x_59 = l_Lean_evalPrec___closed__2; -lean_inc(x_56); -x_60 = l_Lean_Syntax_isOfKind(x_56, x_59); -if (x_60 == 0) -{ -lean_object* x_61; lean_object* x_62; -lean_dec(x_58); -x_61 = l_Lean_evalPrec___closed__3; -x_62 = l_Lean_Macro_throwErrorAt___rarg(x_56, x_61, x_54, x_57); -lean_dec(x_56); -return x_62; + lean_dec_ref(x_48); + x_51 = lean_box(0); } -else -{ -lean_object* x_63; lean_object* x_64; -lean_dec(x_54); -x_63 = l_Lean_numLitKind; -x_64 = l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(x_63, x_56); -lean_dec(x_56); -if (lean_obj_tag(x_64) == 0) +x_52 = l_Lean_Syntax_mkNumLit___closed__2; +lean_inc(x_49); +x_53 = l_Lean_Syntax_isOfKind(x_49, x_52); +if (x_53 == 0) { -lean_object* x_65; lean_object* x_66; -x_65 = lean_unsigned_to_nat(0u); -if (lean_is_scalar(x_58)) { - x_66 = lean_alloc_ctor(0, 2, 0); -} else { - x_66 = x_58; -} -lean_ctor_set(x_66, 0, x_65); -lean_ctor_set(x_66, 1, x_57); -return x_66; +lean_object* x_54; lean_object* x_55; +lean_dec(x_51); +x_54 = l_Lean_evalPrec___closed__1; +x_55 = l_Lean_Macro_throwErrorAt___rarg(x_49, x_54, x_47, x_50); +lean_dec(x_49); +return x_55; } else { -lean_object* x_67; lean_object* x_68; -x_67 = lean_ctor_get(x_64, 0); -lean_inc(x_67); -lean_dec(x_64); -if (lean_is_scalar(x_58)) { - x_68 = lean_alloc_ctor(0, 2, 0); +lean_object* x_56; lean_object* x_57; +lean_dec(x_47); +x_56 = l_Lean_TSyntax_getNat(x_49); +if (lean_is_scalar(x_51)) { + x_57 = lean_alloc_ctor(0, 2, 0); } else { - x_68 = x_58; -} -lean_ctor_set(x_68, 0, x_67); -lean_ctor_set(x_68, 1, x_57); -return x_68; + x_57 = x_51; } +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_50); +return x_57; } } else { -lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; -lean_dec(x_54); -x_69 = lean_ctor_get(x_55, 0); -lean_inc(x_69); -x_70 = lean_ctor_get(x_55, 1); -lean_inc(x_70); -if (lean_is_exclusive(x_55)) { - lean_ctor_release(x_55, 0); - lean_ctor_release(x_55, 1); - x_71 = x_55; +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +lean_dec(x_47); +x_58 = lean_ctor_get(x_48, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_48, 1); +lean_inc(x_59); +if (lean_is_exclusive(x_48)) { + lean_ctor_release(x_48, 0); + lean_ctor_release(x_48, 1); + x_60 = x_48; } else { - lean_dec_ref(x_55); - x_71 = lean_box(0); + lean_dec_ref(x_48); + x_60 = lean_box(0); } -if (lean_is_scalar(x_71)) { - x_72 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_60)) { + x_61 = lean_alloc_ctor(1, 2, 0); } else { - x_72 = x_71; + x_61 = x_60; } -lean_ctor_set(x_72, 0, x_69); -lean_ctor_set(x_72, 1, x_70); -return x_72; +lean_ctor_set(x_61, 0, x_58); +lean_ctor_set(x_61, 1, x_59); +return x_61; } } else { -lean_object* x_73; lean_object* x_74; lean_object* x_75; -lean_dec(x_50); -lean_dec(x_49); -lean_dec(x_48); -lean_dec(x_47); -lean_dec(x_46); -lean_dec(x_45); -x_73 = l_Lean_maxRecDepthErrorMessage; -x_74 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_74, 0, x_1); -lean_ctor_set(x_74, 1, x_73); -x_75 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_75, 0, x_74); -lean_ctor_set(x_75, 1, x_3); -return x_75; +lean_object* x_62; lean_object* x_63; lean_object* x_64; +lean_dec(x_43); +lean_dec(x_42); +lean_dec(x_41); +lean_dec(x_40); +lean_dec(x_39); +lean_dec(x_38); +x_62 = l_Lean_maxRecDepthErrorMessage; +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_1); +lean_ctor_set(x_63, 1, x_62); +x_64 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set(x_64, 1, x_3); +return x_64; } } } @@ -13232,85 +13915,83 @@ x_18 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRul x_19 = !lean_is_exclusive(x_18); if (x_19 == 0) { -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_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; x_20 = lean_ctor_get(x_18, 0); lean_dec(x_20); x_21 = lean_nat_add(x_13, x_16); lean_dec(x_16); lean_dec(x_13); x_22 = l_Nat_repr(x_21); -x_23 = l_Lean_numLitKind; -x_24 = lean_box(2); -x_25 = l_Lean_Syntax_mkLit(x_23, x_22, x_24); -lean_ctor_set(x_18, 0, x_25); +x_23 = lean_box(2); +x_24 = l_Lean_Syntax_mkNumLit(x_22, x_23); +lean_ctor_set(x_18, 0, x_24); return x_18; } else { -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_object* x_32; -x_26 = lean_ctor_get(x_18, 1); -lean_inc(x_26); +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_25 = lean_ctor_get(x_18, 1); +lean_inc(x_25); lean_dec(x_18); -x_27 = lean_nat_add(x_13, x_16); +x_26 = lean_nat_add(x_13, x_16); lean_dec(x_16); lean_dec(x_13); -x_28 = l_Nat_repr(x_27); -x_29 = l_Lean_numLitKind; -x_30 = lean_box(2); -x_31 = l_Lean_Syntax_mkLit(x_29, x_28, x_30); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_26); -return x_32; +x_27 = l_Nat_repr(x_26); +x_28 = lean_box(2); +x_29 = l_Lean_Syntax_mkNumLit(x_27, x_28); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_25); +return x_30; } } else { -uint8_t x_33; +uint8_t x_31; lean_dec(x_13); lean_dec(x_2); -x_33 = !lean_is_exclusive(x_15); -if (x_33 == 0) +x_31 = !lean_is_exclusive(x_15); +if (x_31 == 0) { return x_15; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_15, 0); -x_35 = lean_ctor_get(x_15, 1); -lean_inc(x_35); -lean_inc(x_34); +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_15, 0); +x_33 = lean_ctor_get(x_15, 1); +lean_inc(x_33); +lean_inc(x_32); lean_dec(x_15); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -return x_36; +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; } } } else { -uint8_t x_37; +uint8_t x_35; lean_dec(x_11); lean_dec(x_2); -x_37 = !lean_is_exclusive(x_12); -if (x_37 == 0) +x_35 = !lean_is_exclusive(x_12); +if (x_35 == 0) { return x_12; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_12, 0); -x_39 = lean_ctor_get(x_12, 1); -lean_inc(x_39); -lean_inc(x_38); +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_12, 0); +x_37 = lean_ctor_get(x_12, 1); +lean_inc(x_37); +lean_inc(x_36); lean_dec(x_12); -x_40 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -return x_40; +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +return x_38; } } } @@ -13384,85 +14065,83 @@ x_18 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRul x_19 = !lean_is_exclusive(x_18); if (x_19 == 0) { -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_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; x_20 = lean_ctor_get(x_18, 0); lean_dec(x_20); x_21 = lean_nat_sub(x_13, x_16); lean_dec(x_16); lean_dec(x_13); x_22 = l_Nat_repr(x_21); -x_23 = l_Lean_numLitKind; -x_24 = lean_box(2); -x_25 = l_Lean_Syntax_mkLit(x_23, x_22, x_24); -lean_ctor_set(x_18, 0, x_25); +x_23 = lean_box(2); +x_24 = l_Lean_Syntax_mkNumLit(x_22, x_23); +lean_ctor_set(x_18, 0, x_24); return x_18; } else { -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_object* x_32; -x_26 = lean_ctor_get(x_18, 1); -lean_inc(x_26); +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_25 = lean_ctor_get(x_18, 1); +lean_inc(x_25); lean_dec(x_18); -x_27 = lean_nat_sub(x_13, x_16); +x_26 = lean_nat_sub(x_13, x_16); lean_dec(x_16); lean_dec(x_13); -x_28 = l_Nat_repr(x_27); -x_29 = l_Lean_numLitKind; -x_30 = lean_box(2); -x_31 = l_Lean_Syntax_mkLit(x_29, x_28, x_30); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_26); -return x_32; +x_27 = l_Nat_repr(x_26); +x_28 = lean_box(2); +x_29 = l_Lean_Syntax_mkNumLit(x_27, x_28); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_25); +return x_30; } } else { -uint8_t x_33; +uint8_t x_31; lean_dec(x_13); lean_dec(x_2); -x_33 = !lean_is_exclusive(x_15); -if (x_33 == 0) +x_31 = !lean_is_exclusive(x_15); +if (x_31 == 0) { return x_15; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_15, 0); -x_35 = lean_ctor_get(x_15, 1); -lean_inc(x_35); -lean_inc(x_34); +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_15, 0); +x_33 = lean_ctor_get(x_15, 1); +lean_inc(x_33); +lean_inc(x_32); lean_dec(x_15); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -return x_36; +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; } } } else { -uint8_t x_37; +uint8_t x_35; lean_dec(x_11); lean_dec(x_2); -x_37 = !lean_is_exclusive(x_12); -if (x_37 == 0) +x_35 = !lean_is_exclusive(x_12); +if (x_35 == 0) { return x_12; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_12, 0); -x_39 = lean_ctor_get(x_12, 1); -lean_inc(x_39); -lean_inc(x_38); +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_12, 0); +x_37 = lean_ctor_get(x_12, 1); +lean_inc(x_37); +lean_inc(x_36); lean_dec(x_12); -x_40 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -return x_40; +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +return x_38; } } } @@ -13619,53 +14298,51 @@ uint8_t x_11; x_11 = !lean_is_exclusive(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; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; x_12 = lean_ctor_get(x_10, 0); x_13 = l_Nat_repr(x_12); -x_14 = l_Lean_numLitKind; -x_15 = lean_box(2); -x_16 = l_Lean_Syntax_mkLit(x_14, x_13, x_15); -lean_ctor_set(x_10, 0, x_16); +x_14 = lean_box(2); +x_15 = l_Lean_Syntax_mkNumLit(x_13, x_14); +lean_ctor_set(x_10, 0, x_15); return x_10; } else { -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_17 = lean_ctor_get(x_10, 0); -x_18 = lean_ctor_get(x_10, 1); -lean_inc(x_18); +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_16 = lean_ctor_get(x_10, 0); +x_17 = lean_ctor_get(x_10, 1); lean_inc(x_17); +lean_inc(x_16); lean_dec(x_10); -x_19 = l_Nat_repr(x_17); -x_20 = l_Lean_numLitKind; -x_21 = lean_box(2); -x_22 = l_Lean_Syntax_mkLit(x_20, x_19, x_21); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_18); -return x_23; +x_18 = l_Nat_repr(x_16); +x_19 = lean_box(2); +x_20 = l_Lean_Syntax_mkNumLit(x_18, x_19); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_17); +return x_21; } } else { -uint8_t x_24; -x_24 = !lean_is_exclusive(x_10); -if (x_24 == 0) +uint8_t x_22; +x_22 = !lean_is_exclusive(x_10); +if (x_22 == 0) { return x_10; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_10, 0); -x_26 = lean_ctor_get(x_10, 1); -lean_inc(x_26); -lean_inc(x_25); +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_10, 0); +x_24 = lean_ctor_get(x_10, 1); +lean_inc(x_24); +lean_inc(x_23); lean_dec(x_10); -x_27 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_27, 0, x_25); -lean_ctor_set(x_27, 1, x_26); -return x_27; +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; } } } @@ -13712,7 +14389,7 @@ if (x_15 == 0) lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; x_16 = lean_ctor_get(x_14, 0); x_17 = lean_ctor_get(x_14, 1); -x_18 = l_Lean_evalPrec___closed__2; +x_18 = l_Lean_Syntax_mkNumLit___closed__2; lean_inc(x_16); x_19 = l_Lean_Syntax_isOfKind(x_16, x_18); if (x_19 == 0) @@ -13726,105 +14403,71 @@ return x_21; } else { -lean_object* x_22; lean_object* x_23; +lean_object* x_22; lean_dec(x_2); -x_22 = l_Lean_numLitKind; -x_23 = l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(x_22, x_16); -lean_dec(x_16); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; -x_24 = lean_unsigned_to_nat(0u); -lean_ctor_set(x_14, 0, x_24); -return x_14; -} -else -{ -lean_object* x_25; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -lean_ctor_set(x_14, 0, x_25); +x_22 = l_Lean_TSyntax_getNat(x_16); +lean_ctor_set(x_14, 0, x_22); return x_14; } } -} else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_26 = lean_ctor_get(x_14, 0); -x_27 = lean_ctor_get(x_14, 1); -lean_inc(x_27); -lean_inc(x_26); +lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_23 = lean_ctor_get(x_14, 0); +x_24 = lean_ctor_get(x_14, 1); +lean_inc(x_24); +lean_inc(x_23); lean_dec(x_14); -x_28 = l_Lean_evalPrec___closed__2; -lean_inc(x_26); -x_29 = l_Lean_Syntax_isOfKind(x_26, x_28); -if (x_29 == 0) +x_25 = l_Lean_Syntax_mkNumLit___closed__2; +lean_inc(x_23); +x_26 = l_Lean_Syntax_isOfKind(x_23, x_25); +if (x_26 == 0) { -lean_object* x_30; lean_object* x_31; -x_30 = l_Lean_evalPrio___closed__1; -x_31 = l_Lean_Macro_throwErrorAt___rarg(x_26, x_30, x_2, x_27); -lean_dec(x_26); -return x_31; +lean_object* x_27; lean_object* x_28; +x_27 = l_Lean_evalPrio___closed__1; +x_28 = l_Lean_Macro_throwErrorAt___rarg(x_23, x_27, x_2, x_24); +lean_dec(x_23); +return x_28; } else { -lean_object* x_32; lean_object* x_33; +lean_object* x_29; lean_object* x_30; lean_dec(x_2); -x_32 = l_Lean_numLitKind; -x_33 = l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(x_32, x_26); -lean_dec(x_26); -if (lean_obj_tag(x_33) == 0) -{ -lean_object* x_34; lean_object* x_35; -x_34 = lean_unsigned_to_nat(0u); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_27); -return x_35; -} -else -{ -lean_object* x_36; lean_object* x_37; -x_36 = lean_ctor_get(x_33, 0); -lean_inc(x_36); -lean_dec(x_33); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_27); -return x_37; -} +x_29 = l_Lean_TSyntax_getNat(x_23); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_24); +return x_30; } } } else { -uint8_t x_38; +uint8_t x_31; lean_dec(x_2); -x_38 = !lean_is_exclusive(x_14); -if (x_38 == 0) +x_31 = !lean_is_exclusive(x_14); +if (x_31 == 0) { return x_14; } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_39 = lean_ctor_get(x_14, 0); -x_40 = lean_ctor_get(x_14, 1); -lean_inc(x_40); -lean_inc(x_39); +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_14, 0); +x_33 = lean_ctor_get(x_14, 1); +lean_inc(x_33); +lean_inc(x_32); lean_dec(x_14); -x_41 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_40); -return x_41; -} +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; +} } } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_free_object(x_2); lean_dec(x_10); lean_dec(x_9); @@ -13832,155 +14475,133 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_42 = l_Lean_maxRecDepthErrorMessage; -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_1); -lean_ctor_set(x_43, 1, x_42); -x_44 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_44, 1, x_3); -return x_44; +x_35 = l_Lean_maxRecDepthErrorMessage; +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_1); +lean_ctor_set(x_36, 1, x_35); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_3); +return x_37; } } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; -x_45 = lean_ctor_get(x_2, 0); -x_46 = lean_ctor_get(x_2, 1); -x_47 = lean_ctor_get(x_2, 2); -x_48 = lean_ctor_get(x_2, 3); -x_49 = lean_ctor_get(x_2, 4); -x_50 = lean_ctor_get(x_2, 5); -lean_inc(x_50); -lean_inc(x_49); -lean_inc(x_48); -lean_inc(x_47); -lean_inc(x_46); -lean_inc(x_45); +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; +x_38 = lean_ctor_get(x_2, 0); +x_39 = lean_ctor_get(x_2, 1); +x_40 = lean_ctor_get(x_2, 2); +x_41 = lean_ctor_get(x_2, 3); +x_42 = lean_ctor_get(x_2, 4); +x_43 = lean_ctor_get(x_2, 5); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_inc(x_39); +lean_inc(x_38); lean_dec(x_2); -x_51 = lean_nat_dec_eq(x_48, x_49); -if (x_51 == 0) +x_44 = lean_nat_dec_eq(x_41, x_42); +if (x_44 == 0) { -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_52 = lean_unsigned_to_nat(1u); -x_53 = lean_nat_add(x_48, x_52); -lean_dec(x_48); -x_54 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_54, 0, x_45); -lean_ctor_set(x_54, 1, x_46); -lean_ctor_set(x_54, 2, x_47); -lean_ctor_set(x_54, 3, x_53); -lean_ctor_set(x_54, 4, x_49); -lean_ctor_set(x_54, 5, x_50); -lean_inc(x_54); -x_55 = l_Lean_expandMacros(x_1, x_54, x_3); -if (lean_obj_tag(x_55) == 0) +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_45 = lean_unsigned_to_nat(1u); +x_46 = lean_nat_add(x_41, x_45); +lean_dec(x_41); +x_47 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_47, 0, x_38); +lean_ctor_set(x_47, 1, x_39); +lean_ctor_set(x_47, 2, x_40); +lean_ctor_set(x_47, 3, x_46); +lean_ctor_set(x_47, 4, x_42); +lean_ctor_set(x_47, 5, x_43); +lean_inc(x_47); +x_48 = l_Lean_expandMacros(x_1, x_47, x_3); +if (lean_obj_tag(x_48) == 0) { -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; uint8_t x_60; -x_56 = lean_ctor_get(x_55, 0); -lean_inc(x_56); -x_57 = lean_ctor_get(x_55, 1); -lean_inc(x_57); -if (lean_is_exclusive(x_55)) { - lean_ctor_release(x_55, 0); - lean_ctor_release(x_55, 1); - x_58 = x_55; +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; +x_49 = lean_ctor_get(x_48, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_48, 1); +lean_inc(x_50); +if (lean_is_exclusive(x_48)) { + lean_ctor_release(x_48, 0); + lean_ctor_release(x_48, 1); + x_51 = x_48; } else { - lean_dec_ref(x_55); - x_58 = lean_box(0); -} -x_59 = l_Lean_evalPrec___closed__2; -lean_inc(x_56); -x_60 = l_Lean_Syntax_isOfKind(x_56, x_59); -if (x_60 == 0) -{ -lean_object* x_61; lean_object* x_62; -lean_dec(x_58); -x_61 = l_Lean_evalPrio___closed__1; -x_62 = l_Lean_Macro_throwErrorAt___rarg(x_56, x_61, x_54, x_57); -lean_dec(x_56); -return x_62; + lean_dec_ref(x_48); + x_51 = lean_box(0); } -else -{ -lean_object* x_63; lean_object* x_64; -lean_dec(x_54); -x_63 = l_Lean_numLitKind; -x_64 = l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(x_63, x_56); -lean_dec(x_56); -if (lean_obj_tag(x_64) == 0) +x_52 = l_Lean_Syntax_mkNumLit___closed__2; +lean_inc(x_49); +x_53 = l_Lean_Syntax_isOfKind(x_49, x_52); +if (x_53 == 0) { -lean_object* x_65; lean_object* x_66; -x_65 = lean_unsigned_to_nat(0u); -if (lean_is_scalar(x_58)) { - x_66 = lean_alloc_ctor(0, 2, 0); -} else { - x_66 = x_58; -} -lean_ctor_set(x_66, 0, x_65); -lean_ctor_set(x_66, 1, x_57); -return x_66; +lean_object* x_54; lean_object* x_55; +lean_dec(x_51); +x_54 = l_Lean_evalPrio___closed__1; +x_55 = l_Lean_Macro_throwErrorAt___rarg(x_49, x_54, x_47, x_50); +lean_dec(x_49); +return x_55; } else { -lean_object* x_67; lean_object* x_68; -x_67 = lean_ctor_get(x_64, 0); -lean_inc(x_67); -lean_dec(x_64); -if (lean_is_scalar(x_58)) { - x_68 = lean_alloc_ctor(0, 2, 0); +lean_object* x_56; lean_object* x_57; +lean_dec(x_47); +x_56 = l_Lean_TSyntax_getNat(x_49); +if (lean_is_scalar(x_51)) { + x_57 = lean_alloc_ctor(0, 2, 0); } else { - x_68 = x_58; -} -lean_ctor_set(x_68, 0, x_67); -lean_ctor_set(x_68, 1, x_57); -return x_68; + x_57 = x_51; } +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_50); +return x_57; } } else { -lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; -lean_dec(x_54); -x_69 = lean_ctor_get(x_55, 0); -lean_inc(x_69); -x_70 = lean_ctor_get(x_55, 1); -lean_inc(x_70); -if (lean_is_exclusive(x_55)) { - lean_ctor_release(x_55, 0); - lean_ctor_release(x_55, 1); - x_71 = x_55; +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +lean_dec(x_47); +x_58 = lean_ctor_get(x_48, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_48, 1); +lean_inc(x_59); +if (lean_is_exclusive(x_48)) { + lean_ctor_release(x_48, 0); + lean_ctor_release(x_48, 1); + x_60 = x_48; } else { - lean_dec_ref(x_55); - x_71 = lean_box(0); + lean_dec_ref(x_48); + x_60 = lean_box(0); } -if (lean_is_scalar(x_71)) { - x_72 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_60)) { + x_61 = lean_alloc_ctor(1, 2, 0); } else { - x_72 = x_71; + x_61 = x_60; } -lean_ctor_set(x_72, 0, x_69); -lean_ctor_set(x_72, 1, x_70); -return x_72; +lean_ctor_set(x_61, 0, x_58); +lean_ctor_set(x_61, 1, x_59); +return x_61; } } else { -lean_object* x_73; lean_object* x_74; lean_object* x_75; -lean_dec(x_50); -lean_dec(x_49); -lean_dec(x_48); -lean_dec(x_47); -lean_dec(x_46); -lean_dec(x_45); -x_73 = l_Lean_maxRecDepthErrorMessage; -x_74 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_74, 0, x_1); -lean_ctor_set(x_74, 1, x_73); -x_75 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_75, 0, x_74); -lean_ctor_set(x_75, 1, x_3); -return x_75; +lean_object* x_62; lean_object* x_63; lean_object* x_64; +lean_dec(x_43); +lean_dec(x_42); +lean_dec(x_41); +lean_dec(x_40); +lean_dec(x_39); +lean_dec(x_38); +x_62 = l_Lean_maxRecDepthErrorMessage; +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_1); +lean_ctor_set(x_63, 1, x_62); +x_64 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set(x_64, 1, x_3); +return x_64; } } } @@ -14053,85 +14674,83 @@ x_18 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRul x_19 = !lean_is_exclusive(x_18); if (x_19 == 0) { -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_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; x_20 = lean_ctor_get(x_18, 0); lean_dec(x_20); x_21 = lean_nat_add(x_13, x_16); lean_dec(x_16); lean_dec(x_13); x_22 = l_Nat_repr(x_21); -x_23 = l_Lean_numLitKind; -x_24 = lean_box(2); -x_25 = l_Lean_Syntax_mkLit(x_23, x_22, x_24); -lean_ctor_set(x_18, 0, x_25); +x_23 = lean_box(2); +x_24 = l_Lean_Syntax_mkNumLit(x_22, x_23); +lean_ctor_set(x_18, 0, x_24); return x_18; } else { -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_object* x_32; -x_26 = lean_ctor_get(x_18, 1); -lean_inc(x_26); +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_25 = lean_ctor_get(x_18, 1); +lean_inc(x_25); lean_dec(x_18); -x_27 = lean_nat_add(x_13, x_16); +x_26 = lean_nat_add(x_13, x_16); lean_dec(x_16); lean_dec(x_13); -x_28 = l_Nat_repr(x_27); -x_29 = l_Lean_numLitKind; -x_30 = lean_box(2); -x_31 = l_Lean_Syntax_mkLit(x_29, x_28, x_30); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_26); -return x_32; +x_27 = l_Nat_repr(x_26); +x_28 = lean_box(2); +x_29 = l_Lean_Syntax_mkNumLit(x_27, x_28); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_25); +return x_30; } } else { -uint8_t x_33; +uint8_t x_31; lean_dec(x_13); lean_dec(x_2); -x_33 = !lean_is_exclusive(x_15); -if (x_33 == 0) +x_31 = !lean_is_exclusive(x_15); +if (x_31 == 0) { return x_15; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_15, 0); -x_35 = lean_ctor_get(x_15, 1); -lean_inc(x_35); -lean_inc(x_34); +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_15, 0); +x_33 = lean_ctor_get(x_15, 1); +lean_inc(x_33); +lean_inc(x_32); lean_dec(x_15); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -return x_36; +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; } } } else { -uint8_t x_37; +uint8_t x_35; lean_dec(x_11); lean_dec(x_2); -x_37 = !lean_is_exclusive(x_12); -if (x_37 == 0) +x_35 = !lean_is_exclusive(x_12); +if (x_35 == 0) { return x_12; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_12, 0); -x_39 = lean_ctor_get(x_12, 1); -lean_inc(x_39); -lean_inc(x_38); +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_12, 0); +x_37 = lean_ctor_get(x_12, 1); +lean_inc(x_37); +lean_inc(x_36); lean_dec(x_12); -x_40 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -return x_40; +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +return x_38; } } } @@ -14205,85 +14824,83 @@ x_18 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRul x_19 = !lean_is_exclusive(x_18); if (x_19 == 0) { -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_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; x_20 = lean_ctor_get(x_18, 0); lean_dec(x_20); x_21 = lean_nat_sub(x_13, x_16); lean_dec(x_16); lean_dec(x_13); x_22 = l_Nat_repr(x_21); -x_23 = l_Lean_numLitKind; -x_24 = lean_box(2); -x_25 = l_Lean_Syntax_mkLit(x_23, x_22, x_24); -lean_ctor_set(x_18, 0, x_25); +x_23 = lean_box(2); +x_24 = l_Lean_Syntax_mkNumLit(x_22, x_23); +lean_ctor_set(x_18, 0, x_24); return x_18; } else { -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_object* x_32; -x_26 = lean_ctor_get(x_18, 1); -lean_inc(x_26); +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_25 = lean_ctor_get(x_18, 1); +lean_inc(x_25); lean_dec(x_18); -x_27 = lean_nat_sub(x_13, x_16); +x_26 = lean_nat_sub(x_13, x_16); lean_dec(x_16); lean_dec(x_13); -x_28 = l_Nat_repr(x_27); -x_29 = l_Lean_numLitKind; -x_30 = lean_box(2); -x_31 = l_Lean_Syntax_mkLit(x_29, x_28, x_30); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_26); -return x_32; +x_27 = l_Nat_repr(x_26); +x_28 = lean_box(2); +x_29 = l_Lean_Syntax_mkNumLit(x_27, x_28); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_25); +return x_30; } } else { -uint8_t x_33; +uint8_t x_31; lean_dec(x_13); lean_dec(x_2); -x_33 = !lean_is_exclusive(x_15); -if (x_33 == 0) +x_31 = !lean_is_exclusive(x_15); +if (x_31 == 0) { return x_15; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_15, 0); -x_35 = lean_ctor_get(x_15, 1); -lean_inc(x_35); -lean_inc(x_34); +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_15, 0); +x_33 = lean_ctor_get(x_15, 1); +lean_inc(x_33); +lean_inc(x_32); lean_dec(x_15); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -return x_36; +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; } } } else { -uint8_t x_37; +uint8_t x_35; lean_dec(x_11); lean_dec(x_2); -x_37 = !lean_is_exclusive(x_12); -if (x_37 == 0) +x_35 = !lean_is_exclusive(x_12); +if (x_35 == 0) { return x_12; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_12, 0); -x_39 = lean_ctor_get(x_12, 1); -lean_inc(x_39); -lean_inc(x_38); +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_12, 0); +x_37 = lean_ctor_get(x_12, 1); +lean_inc(x_37); +lean_inc(x_36); lean_dec(x_12); -x_40 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -return x_40; +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +return x_38; } } } @@ -14422,53 +15039,51 @@ uint8_t x_11; x_11 = !lean_is_exclusive(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; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; x_12 = lean_ctor_get(x_10, 0); x_13 = l_Nat_repr(x_12); -x_14 = l_Lean_numLitKind; -x_15 = lean_box(2); -x_16 = l_Lean_Syntax_mkLit(x_14, x_13, x_15); -lean_ctor_set(x_10, 0, x_16); +x_14 = lean_box(2); +x_15 = l_Lean_Syntax_mkNumLit(x_13, x_14); +lean_ctor_set(x_10, 0, x_15); return x_10; } else { -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_17 = lean_ctor_get(x_10, 0); -x_18 = lean_ctor_get(x_10, 1); -lean_inc(x_18); +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_16 = lean_ctor_get(x_10, 0); +x_17 = lean_ctor_get(x_10, 1); lean_inc(x_17); +lean_inc(x_16); lean_dec(x_10); -x_19 = l_Nat_repr(x_17); -x_20 = l_Lean_numLitKind; -x_21 = lean_box(2); -x_22 = l_Lean_Syntax_mkLit(x_20, x_19, x_21); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_18); -return x_23; +x_18 = l_Nat_repr(x_16); +x_19 = lean_box(2); +x_20 = l_Lean_Syntax_mkNumLit(x_18, x_19); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_17); +return x_21; } } else { -uint8_t x_24; -x_24 = !lean_is_exclusive(x_10); -if (x_24 == 0) +uint8_t x_22; +x_22 = !lean_is_exclusive(x_10); +if (x_22 == 0) { return x_10; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_10, 0); -x_26 = lean_ctor_get(x_10, 1); -lean_inc(x_26); -lean_inc(x_25); +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_10, 0); +x_24 = lean_ctor_get(x_10, 1); +lean_inc(x_24); +lean_inc(x_23); lean_dec(x_10); -x_27 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_27, 0, x_25); -lean_ctor_set(x_27, 1, x_26); -return x_27; +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; } } } @@ -15310,50 +15925,273 @@ lean_dec(x_1); return x_2; } } -static lean_object* _init_l_Lean_Syntax_SepArray_instCoeTailSepArrayArraySyntax___closed__1() { +LEAN_EXPORT lean_object* l_Lean_Syntax_TSepArray_getElems___rarg(lean_object* x_1) { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Syntax_SepArray_getElems___rarg___boxed), 1, 0); -return x_1; +lean_object* x_2; lean_object* x_3; uint8_t x_4; +x_2 = lean_array_get_size(x_1); +x_3 = lean_unsigned_to_nat(0u); +x_4 = lean_nat_dec_lt(x_3, x_2); +if (x_4 == 0) +{ +lean_object* x_5; +lean_dec(x_2); +x_5 = l_Lean_mkSepArray___closed__1; +return x_5; +} +else +{ +uint8_t x_6; +x_6 = lean_nat_dec_le(x_2, x_2); +if (x_6 == 0) +{ +lean_object* x_7; +lean_dec(x_2); +x_7 = l_Lean_mkSepArray___closed__1; +return x_7; +} +else +{ +size_t x_8; size_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_8 = 0; +x_9 = lean_usize_of_nat(x_2); +lean_dec(x_2); +x_10 = l_Array_getSepElems___rarg___closed__1; +x_11 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(x_1, x_8, x_9, x_10); +x_12 = lean_ctor_get(x_11, 1); +lean_inc(x_12); +lean_dec(x_11); +return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Syntax_SepArray_instCoeTailSepArrayArraySyntax(lean_object* x_1) { +} +} +LEAN_EXPORT lean_object* l_Lean_Syntax_TSepArray_getElems(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_2; -x_2 = l_Lean_Syntax_SepArray_instCoeTailSepArrayArraySyntax___closed__1; -return x_2; +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_Syntax_TSepArray_getElems___rarg___boxed), 1, 0); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Syntax_SepArray_instCoeTailSepArrayArraySyntax___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Syntax_TSepArray_getElems___rarg___boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Syntax_SepArray_instCoeTailSepArrayArraySyntax(x_1); +x_2 = l_Lean_Syntax_TSepArray_getElems___rarg(x_1); lean_dec(x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Syntax_decodeInterpStrQuotedChar___boxed__const__1() { +LEAN_EXPORT lean_object* l_Lean_Syntax_TSepArray_getElems___boxed(lean_object* x_1, lean_object* x_2) { _start: { -uint32_t x_1; lean_object* x_2; -x_1 = 123; -x_2 = lean_box_uint32(x_1); -return x_2; +lean_object* x_3; +x_3 = l_Lean_Syntax_TSepArray_getElems(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeInterpStrQuotedChar(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_Syntax_instCoeTailSepArrayArraySyntax___closed__1() { _start: { -lean_object* x_3; -x_3 = l_Lean_Syntax_decodeQuotedChar(x_1, x_2); -if (lean_obj_tag(x_3) == 0) +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Syntax_SepArray_getElems___rarg___boxed), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTailSepArrayArraySyntax(lean_object* x_1) { +_start: { -uint32_t x_4; lean_object* x_5; uint32_t x_6; uint8_t x_7; -x_4 = lean_string_utf8_get(x_1, x_2); -x_5 = lean_string_utf8_next(x_1, x_2); +lean_object* x_2; +x_2 = l_Lean_Syntax_instCoeTailSepArrayArraySyntax___closed__1; +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTailSepArrayArraySyntax___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Syntax_instCoeTailSepArrayArraySyntax(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Syntax_instCoeTSepArrayTSyntaxArray___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Syntax_TSepArray_getElems___rarg___boxed), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSepArrayTSyntaxArray(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Syntax_instCoeTSepArrayTSyntaxArray___closed__1; +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSepArrayTSyntaxArray___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Syntax_instCoeTSepArrayTSyntaxArray(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSyntaxArray___rarg(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSyntaxArray(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Lean_Syntax_instCoeTSyntaxArray___rarg___boxed), 1, 0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSyntaxArray___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Syntax_instCoeTSyntaxArray___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSyntaxArray___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_Syntax_instCoeTSyntaxArray(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +static lean_object* _init_l_Lean_Syntax_instCoeTSyntaxArrayArraySyntax___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_TSyntaxArray_rawImpl___rarg___boxed), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSyntaxArrayArraySyntax(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Syntax_instCoeTSyntaxArrayArraySyntax___closed__1; +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSyntaxArrayArraySyntax___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Syntax_instCoeTSyntaxArrayArraySyntax(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Command", 7); +return x_1; +} +} +static lean_object* _init_l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_mkHole___closed__4; +x_2 = l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("declId", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___closed__2; +x_2 = l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___closed__3; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil(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_8; +x_2 = l_Lean_Syntax_mkApp___closed__3; +x_3 = lean_array_push(x_2, x_1); +x_4 = l_Lean_mkOptionalNode___closed__1; +x_5 = lean_array_push(x_3, x_4); +x_6 = lean_box(2); +x_7 = l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___closed__4; +x_8 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_8, 0, x_6); +lean_ctor_set(x_8, 1, x_7); +lean_ctor_set(x_8, 2, x_5); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTermTSyntaxConsSyntaxNodeKindMkStrAnonymousNil(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTermTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Syntax_instCoeTermTSyntaxConsSyntaxNodeKindMkStrAnonymousNil(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Init_Meta_0__Lean_Syntax_decodeInterpStrQuotedChar___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 123; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeInterpStrQuotedChar(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Syntax_decodeQuotedChar(x_1, x_2); +if (lean_obj_tag(x_3) == 0) +{ +uint32_t x_4; lean_object* x_5; uint32_t x_6; uint8_t x_7; +x_4 = lean_string_utf8_get(x_1, x_2); +x_5 = lean_string_utf8_next(x_1, x_2); x_6 = 123; x_7 = lean_uint32_dec_eq(x_4, x_6); if (x_7 == 0) @@ -15525,11 +16363,29 @@ lean_dec(x_1); return x_2; } } +static lean_object* _init_l_Lean_Syntax_isInterpolatedStrLit_x3f___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("interpolatedStrLitKind", 22); +return x_1; +} +} +static lean_object* _init_l_Lean_Syntax_isInterpolatedStrLit_x3f___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Syntax_isInterpolatedStrLit_x3f___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Syntax_isInterpolatedStrLit_x3f(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_interpolatedStrLitKind; +x_2 = l_Lean_Syntax_isInterpolatedStrLit_x3f___closed__2; x_3 = l_Lean_Syntax_isLit_x3f(x_2, x_1); if (lean_obj_tag(x_3) == 0) { @@ -15560,7 +16416,61 @@ lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Syntax_expandInterpolatedStrChunks___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_EXPORT lean_object* l_Lean_Syntax_getSepArgs(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_2 = l_Lean_Syntax_getArgs(x_1); +x_3 = lean_array_get_size(x_2); +x_4 = lean_unsigned_to_nat(0u); +x_5 = lean_nat_dec_lt(x_4, x_3); +if (x_5 == 0) +{ +lean_object* x_6; +lean_dec(x_3); +lean_dec(x_2); +x_6 = l_Lean_mkSepArray___closed__1; +return x_6; +} +else +{ +uint8_t x_7; +x_7 = lean_nat_dec_le(x_3, x_3); +if (x_7 == 0) +{ +lean_object* x_8; +lean_dec(x_3); +lean_dec(x_2); +x_8 = l_Lean_mkSepArray___closed__1; +return x_8; +} +else +{ +size_t x_9; size_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_9 = 0; +x_10 = lean_usize_of_nat(x_3); +lean_dec(x_3); +x_11 = l_Array_getSepElems___rarg___closed__1; +x_12 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(x_2, x_9, x_10, x_11); +lean_dec(x_2); +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +lean_dec(x_12); +return x_13; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Syntax_getSepArgs___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Syntax_getSepArgs(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_TSyntax_expandInterpolatedStrChunks___spec__1___lambda__1(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; lean_object* x_9; lean_object* x_10; @@ -15577,19 +16487,19 @@ lean_ctor_set(x_10, 1, x_5); return x_10; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Syntax_expandInterpolatedStrChunks___spec__1___lambda__2___closed__1() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_TSyntax_expandInterpolatedStrChunks___spec__1___lambda__2___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Syntax_expandInterpolatedStrChunks___spec__1___lambda__1___boxed), 5, 0); +x_1 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_TSyntax_expandInterpolatedStrChunks___spec__1___lambda__1___boxed), 5, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Syntax_expandInterpolatedStrChunks___spec__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) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_TSyntax_expandInterpolatedStrChunks___spec__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; lean_object* x_8; uint8_t x_9; -x_7 = l_Array_forInUnsafe_loop___at_Lean_Syntax_expandInterpolatedStrChunks___spec__1___lambda__2___closed__1; +x_7 = l_Array_forInUnsafe_loop___at_Lean_TSyntax_expandInterpolatedStrChunks___spec__1___lambda__2___closed__1; x_8 = lean_unsigned_to_nat(0u); x_9 = lean_nat_dec_eq(x_2, x_8); if (x_9 == 0) @@ -15645,7 +16555,7 @@ return x_20; } } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Syntax_expandInterpolatedStrChunks___spec__1(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_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_TSyntax_expandInterpolatedStrChunks___spec__1(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_object* x_7, lean_object* x_8) { _start: { uint8_t x_9; @@ -15687,7 +16597,7 @@ lean_inc(x_17); lean_dec(x_15); lean_inc(x_7); lean_inc(x_1); -x_18 = l_Array_forInUnsafe_loop___at_Lean_Syntax_expandInterpolatedStrChunks___spec__1___lambda__2(x_1, x_12, x_13, x_16, x_7, x_17); +x_18 = l_Array_forInUnsafe_loop___at_Lean_TSyntax_expandInterpolatedStrChunks___spec__1___lambda__2(x_1, x_12, x_13, x_16, x_7, x_17); if (lean_obj_tag(x_18) == 0) { lean_object* x_19; @@ -15820,7 +16730,7 @@ lean_inc(x_44); lean_dec(x_42); lean_inc(x_7); lean_inc(x_1); -x_45 = l_Array_forInUnsafe_loop___at_Lean_Syntax_expandInterpolatedStrChunks___spec__1___lambda__2(x_1, x_12, x_13, x_43, x_7, x_44); +x_45 = l_Array_forInUnsafe_loop___at_Lean_TSyntax_expandInterpolatedStrChunks___spec__1___lambda__2(x_1, x_12, x_13, x_43, x_7, x_44); if (lean_obj_tag(x_45) == 0) { lean_object* x_46; @@ -15933,7 +16843,7 @@ return x_65; } } } -static lean_object* _init_l_Lean_Syntax_expandInterpolatedStrChunks___closed__1() { +static lean_object* _init_l_Lean_TSyntax_expandInterpolatedStrChunks___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -15945,7 +16855,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Syntax_expandInterpolatedStrChunks(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_TSyntax_expandInterpolatedStrChunks(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; size_t x_7; size_t x_8; lean_object* x_9; lean_object* x_10; @@ -15953,8 +16863,8 @@ x_6 = lean_array_get_size(x_1); x_7 = lean_usize_of_nat(x_6); lean_dec(x_6); x_8 = 0; -x_9 = l_Lean_Syntax_expandInterpolatedStrChunks___closed__1; -x_10 = l_Array_forInUnsafe_loop___at_Lean_Syntax_expandInterpolatedStrChunks___spec__1(x_2, x_3, x_1, x_7, x_8, x_9, x_4, x_5); +x_9 = l_Lean_TSyntax_expandInterpolatedStrChunks___closed__1; +x_10 = l_Array_forInUnsafe_loop___at_Lean_TSyntax_expandInterpolatedStrChunks___spec__1(x_2, x_3, x_1, x_7, x_8, x_9, x_4, x_5); if (lean_obj_tag(x_10) == 0) { uint8_t x_11; @@ -16010,18 +16920,18 @@ return x_21; } } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Syntax_expandInterpolatedStrChunks___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_TSyntax_expandInterpolatedStrChunks___spec__1___lambda__1___boxed(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; -x_6 = l_Array_forInUnsafe_loop___at_Lean_Syntax_expandInterpolatedStrChunks___spec__1___lambda__1(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Array_forInUnsafe_loop___at_Lean_TSyntax_expandInterpolatedStrChunks___spec__1___lambda__1(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); return x_6; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Syntax_expandInterpolatedStrChunks___spec__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) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_TSyntax_expandInterpolatedStrChunks___spec__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: { size_t x_9; size_t x_10; lean_object* x_11; @@ -16029,21 +16939,21 @@ x_9 = lean_unbox_usize(x_4); lean_dec(x_4); x_10 = lean_unbox_usize(x_5); lean_dec(x_5); -x_11 = l_Array_forInUnsafe_loop___at_Lean_Syntax_expandInterpolatedStrChunks___spec__1(x_1, x_2, x_3, x_9, x_10, x_6, x_7, x_8); +x_11 = l_Array_forInUnsafe_loop___at_Lean_TSyntax_expandInterpolatedStrChunks___spec__1(x_1, x_2, x_3, x_9, x_10, x_6, x_7, x_8); lean_dec(x_3); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Syntax_expandInterpolatedStrChunks___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_TSyntax_expandInterpolatedStrChunks___boxed(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; -x_6 = l_Lean_Syntax_expandInterpolatedStrChunks(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_TSyntax_expandInterpolatedStrChunks(x_1, x_2, x_3, x_4, x_5); lean_dec(x_1); return x_6; } } -static lean_object* _init_l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__1() { +static lean_object* _init_l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__1() { _start: { lean_object* x_1; @@ -16051,17 +16961,17 @@ x_1 = lean_mk_string_from_bytes("term_++_", 8); return x_1; } } -static lean_object* _init_l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__2() { +static lean_object* _init_l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__1; +x_2 = l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__3() { +static lean_object* _init_l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__3() { _start: { lean_object* x_1; @@ -16069,7 +16979,7 @@ x_1 = lean_mk_string_from_bytes("++", 2); return x_1; } } -static lean_object* _init_l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__4() { +static lean_object* _init_l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; @@ -16078,7 +16988,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Syntax_expandInterpolatedStr___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_TSyntax_expandInterpolatedStr___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; @@ -16088,16 +16998,16 @@ if (x_6 == 0) { lean_object* x_7; 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; lean_object* x_15; lean_object* x_16; x_7 = lean_ctor_get(x_5, 0); -x_8 = l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__3; +x_8 = l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__3; x_9 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_9, 0, x_7); lean_ctor_set(x_9, 1, x_8); -x_10 = l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__4; +x_10 = l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__4; x_11 = lean_array_push(x_10, x_1); x_12 = lean_array_push(x_11, x_9); x_13 = lean_array_push(x_12, x_2); x_14 = lean_box(2); -x_15 = l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__2; +x_15 = l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__2; x_16 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_16, 0, x_14); lean_ctor_set(x_16, 1, x_15); @@ -16113,16 +17023,16 @@ x_18 = lean_ctor_get(x_5, 1); lean_inc(x_18); lean_inc(x_17); lean_dec(x_5); -x_19 = l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__3; +x_19 = l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__3; x_20 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_20, 0, x_17); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__4; +x_21 = l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__4; x_22 = lean_array_push(x_21, x_1); x_23 = lean_array_push(x_22, x_20); x_24 = lean_array_push(x_23, x_2); x_25 = lean_box(2); -x_26 = l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__2; +x_26 = l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__2; x_27 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_27, 0, x_25); lean_ctor_set(x_27, 1, x_26); @@ -16134,25 +17044,7 @@ return x_28; } } } -static lean_object* _init_l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("null", 4); -return x_1; -} -} -static lean_object* _init_l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__1; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_Syntax_expandInterpolatedStr___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_TSyntax_expandInterpolatedStr___lambda__2(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; @@ -16166,7 +17058,7 @@ lean_dec(x_7); x_8 = l_Lean_mkOptionalNode___closed__2; x_9 = lean_array_push(x_8, x_2); x_10 = lean_box(2); -x_11 = l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__2; +x_11 = l_Lean_mkNullNode___closed__2; x_12 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_12, 0, x_10); lean_ctor_set(x_12, 1, x_11); @@ -16191,7 +17083,7 @@ lean_dec(x_5); x_19 = l_Lean_mkOptionalNode___closed__2; x_20 = lean_array_push(x_19, x_2); x_21 = lean_box(2); -x_22 = l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__2; +x_22 = l_Lean_mkNullNode___closed__2; x_23 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_23, 0, x_21); lean_ctor_set(x_23, 1, x_22); @@ -16211,15 +17103,15 @@ return x_29; } } } -static lean_object* _init_l_Lean_Syntax_expandInterpolatedStr___closed__1() { +static lean_object* _init_l_Lean_TSyntax_expandInterpolatedStr___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Syntax_expandInterpolatedStr___lambda__1), 4, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_TSyntax_expandInterpolatedStr___lambda__1), 4, 0); return x_1; } } -static lean_object* _init_l_Lean_Syntax_expandInterpolatedStr___closed__2() { +static lean_object* _init_l_Lean_TSyntax_expandInterpolatedStr___closed__2() { _start: { lean_object* x_1; @@ -16227,17 +17119,17 @@ x_1 = lean_mk_string_from_bytes("paren", 5); return x_1; } } -static lean_object* _init_l_Lean_Syntax_expandInterpolatedStr___closed__3() { +static lean_object* _init_l_Lean_TSyntax_expandInterpolatedStr___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_mkHole___closed__6; -x_2 = l_Lean_Syntax_expandInterpolatedStr___closed__2; +x_2 = l_Lean_TSyntax_expandInterpolatedStr___closed__2; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Syntax_expandInterpolatedStr___closed__4() { +static lean_object* _init_l_Lean_TSyntax_expandInterpolatedStr___closed__4() { _start: { lean_object* x_1; @@ -16245,26 +17137,26 @@ x_1 = lean_mk_string_from_bytes("typeAscription", 14); return x_1; } } -static lean_object* _init_l_Lean_Syntax_expandInterpolatedStr___closed__5() { +static lean_object* _init_l_Lean_TSyntax_expandInterpolatedStr___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_mkHole___closed__6; -x_2 = l_Lean_Syntax_expandInterpolatedStr___closed__4; +x_2 = l_Lean_TSyntax_expandInterpolatedStr___closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Syntax_expandInterpolatedStr(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_TSyntax_expandInterpolatedStr(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; lean_object* x_9; x_6 = l_Lean_Syntax_getArgs(x_1); -x_7 = lean_alloc_closure((void*)(l_Lean_Syntax_expandInterpolatedStr___lambda__2), 4, 1); +x_7 = lean_alloc_closure((void*)(l_Lean_TSyntax_expandInterpolatedStr___lambda__2), 4, 1); lean_closure_set(x_7, 0, x_3); -x_8 = l_Lean_Syntax_expandInterpolatedStr___closed__1; +x_8 = l_Lean_TSyntax_expandInterpolatedStr___closed__1; lean_inc(x_4); -x_9 = l_Lean_Syntax_expandInterpolatedStrChunks(x_6, x_8, x_7, x_4, x_5); +x_9 = l_Lean_TSyntax_expandInterpolatedStrChunks(x_6, x_8, x_7, x_4, x_5); lean_dec(x_6); if (lean_obj_tag(x_9) == 0) { @@ -16294,14 +17186,14 @@ x_19 = l_Lean_Syntax_mkApp___closed__3; x_20 = lean_array_push(x_19, x_18); x_21 = lean_array_push(x_20, x_2); x_22 = lean_box(2); -x_23 = l_Lean_Syntax_expandInterpolatedStr___closed__5; +x_23 = l_Lean_TSyntax_expandInterpolatedStr___closed__5; x_24 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_24, 0, x_22); lean_ctor_set(x_24, 1, x_23); lean_ctor_set(x_24, 2, x_21); x_25 = l_Lean_mkOptionalNode___closed__2; x_26 = lean_array_push(x_25, x_24); -x_27 = l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__2; +x_27 = l_Lean_mkNullNode___closed__2; x_28 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_28, 0, x_22); lean_ctor_set(x_28, 1, x_27); @@ -16316,11 +17208,11 @@ x_32 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta_ x_33 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_33, 0, x_14); lean_ctor_set(x_33, 1, x_32); -x_34 = l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__4; +x_34 = l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__4; x_35 = lean_array_push(x_34, x_16); x_36 = lean_array_push(x_35, x_31); x_37 = lean_array_push(x_36, x_33); -x_38 = l_Lean_Syntax_expandInterpolatedStr___closed__3; +x_38 = l_Lean_TSyntax_expandInterpolatedStr___closed__3; x_39 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_39, 0, x_22); lean_ctor_set(x_39, 1, x_38); @@ -16350,14 +17242,14 @@ x_46 = l_Lean_Syntax_mkApp___closed__3; x_47 = lean_array_push(x_46, x_45); x_48 = lean_array_push(x_47, x_2); x_49 = lean_box(2); -x_50 = l_Lean_Syntax_expandInterpolatedStr___closed__5; +x_50 = l_Lean_TSyntax_expandInterpolatedStr___closed__5; x_51 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_51, 0, x_49); lean_ctor_set(x_51, 1, x_50); lean_ctor_set(x_51, 2, x_48); x_52 = l_Lean_mkOptionalNode___closed__2; x_53 = lean_array_push(x_52, x_51); -x_54 = l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__2; +x_54 = l_Lean_mkNullNode___closed__2; x_55 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_55, 0, x_49); lean_ctor_set(x_55, 1, x_54); @@ -16372,11 +17264,11 @@ x_59 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta_ x_60 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_60, 0, x_40); lean_ctor_set(x_60, 1, x_59); -x_61 = l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__4; +x_61 = l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__4; x_62 = lean_array_push(x_61, x_43); x_63 = lean_array_push(x_62, x_58); x_64 = lean_array_push(x_63, x_60); -x_65 = l_Lean_Syntax_expandInterpolatedStr___closed__3; +x_65 = l_Lean_TSyntax_expandInterpolatedStr___closed__3; x_66 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_66, 0, x_49); lean_ctor_set(x_66, 1, x_65); @@ -16413,80 +17305,26 @@ return x_71; } } } -LEAN_EXPORT lean_object* l_Lean_Syntax_expandInterpolatedStr___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_TSyntax_expandInterpolatedStr___boxed(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; -x_6 = l_Lean_Syntax_expandInterpolatedStr(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_TSyntax_expandInterpolatedStr(x_1, x_2, x_3, x_4, x_5); lean_dec(x_1); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Syntax_getSepArgs(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_TransparencyMode_toCtorIdx(uint8_t x_1) { _start: { -lean_object* x_2; lean_object* x_3; lean_object* x_4; uint8_t x_5; -x_2 = l_Lean_Syntax_getArgs(x_1); -x_3 = lean_array_get_size(x_2); -x_4 = lean_unsigned_to_nat(0u); -x_5 = lean_nat_dec_lt(x_4, x_3); -if (x_5 == 0) +switch (x_1) { +case 0: { -lean_object* x_6; -lean_dec(x_3); -lean_dec(x_2); -x_6 = l_Lean_mkSepArray___closed__1; -return x_6; +lean_object* x_2; +x_2 = lean_unsigned_to_nat(0u); +return x_2; } -else -{ -uint8_t x_7; -x_7 = lean_nat_dec_le(x_3, x_3); -if (x_7 == 0) -{ -lean_object* x_8; -lean_dec(x_3); -lean_dec(x_2); -x_8 = l_Lean_mkSepArray___closed__1; -return x_8; -} -else -{ -size_t x_9; size_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_9 = 0; -x_10 = lean_usize_of_nat(x_3); -lean_dec(x_3); -x_11 = l_Array_getSepElems___rarg___closed__1; -x_12 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(x_2, x_9, x_10, x_11); -lean_dec(x_2); -x_13 = lean_ctor_get(x_12, 1); -lean_inc(x_13); -lean_dec(x_12); -return x_13; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Syntax_getSepArgs___boxed(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = l_Lean_Syntax_getSepArgs(x_1); -lean_dec(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_TransparencyMode_toCtorIdx(uint8_t x_1) { -_start: -{ -switch (x_1) { -case 0: -{ -lean_object* x_2; -x_2 = lean_unsigned_to_nat(0u); -return x_2; -} -case 1: +case 1: { lean_object* x_3; x_3 = lean_unsigned_to_nat(1u); @@ -16577,7 +17415,7 @@ x_1 = 0; return x_1; } } -LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_8567_(uint8_t x_1, uint8_t x_2) { +LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_9542_(uint8_t x_1, uint8_t x_2) { _start: { lean_object* x_3; lean_object* x_4; uint8_t x_5; @@ -16589,7 +17427,7 @@ lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_8567____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_9542____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; @@ -16597,7 +17435,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_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_8567_(x_3, x_4); +x_5 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_9542_(x_3, x_4); x_6 = lean_box(x_5); return x_6; } @@ -16606,7 +17444,7 @@ static lean_object* _init_l_Lean_Meta_instBEqTransparencyMode___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_8567____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_9542____boxed), 2, 0); return x_1; } } @@ -16618,7 +17456,7 @@ x_1 = l_Lean_Meta_instBEqTransparencyMode___closed__1; return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__1() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__1() { _start: { lean_object* x_1; @@ -16626,33 +17464,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Meta.TransparencyMode.all", 30); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__2() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__1; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____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_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__3() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__3; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__2; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____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_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__4() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__3; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__3; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -16660,23 +17498,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__5() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__6; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__2; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____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_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__6() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__5; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__5; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -16684,7 +17522,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__7() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__7() { _start: { lean_object* x_1; @@ -16692,33 +17530,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Meta.TransparencyMode.default", 34); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__8() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__7; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__7; 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_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__9() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__3; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__8; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__8; 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_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__10() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__10() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__9; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__9; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -16726,23 +17564,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__11() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__6; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__8; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__8; 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_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__12() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__12() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__11; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__11; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -16750,7 +17588,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__13() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__13() { _start: { lean_object* x_1; @@ -16758,33 +17596,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Meta.TransparencyMode.reducible", 36); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__14() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__14() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__13; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__13; 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_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__15() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__3; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__14; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__14; 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_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__16() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__16() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__15; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__15; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -16792,23 +17630,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__17() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__6; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__14; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__14; 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_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__18() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__18() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__17; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__17; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -16816,7 +17654,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__19() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__19() { _start: { lean_object* x_1; @@ -16824,33 +17662,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Meta.TransparencyMode.instances", 36); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__20() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__20() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__19; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__19; 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_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__21() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__3; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__20; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__20; 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_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__22() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__22() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__21; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__21; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -16858,23 +17696,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__23() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__6; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__20; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__20; 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_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__24() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__24() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__23; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__23; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -16882,7 +17720,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583_(uint8_t x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558_(uint8_t x_1, lean_object* x_2) { _start: { switch (x_1) { @@ -16894,14 +17732,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_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__4; +x_5 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__4; 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_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__6; +x_7 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__6; x_8 = l_Repr_addAppParen(x_7, x_2); return x_8; } @@ -16914,14 +17752,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_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__10; +x_11 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__10; 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_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__12; +x_13 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__12; x_14 = l_Repr_addAppParen(x_13, x_2); return x_14; } @@ -16934,14 +17772,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_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__16; +x_17 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__16; 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_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__18; +x_19 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__18; x_20 = l_Repr_addAppParen(x_19, x_2); return x_20; } @@ -16954,14 +17792,14 @@ x_22 = lean_nat_dec_le(x_21, x_2); if (x_22 == 0) { lean_object* x_23; lean_object* x_24; -x_23 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__22; +x_23 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__22; x_24 = l_Repr_addAppParen(x_23, x_2); return x_24; } else { lean_object* x_25; lean_object* x_26; -x_25 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__24; +x_25 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__24; x_26 = l_Repr_addAppParen(x_25, x_2); return x_26; } @@ -16969,13 +17807,13 @@ return x_26; } } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____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_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583_(x_3, x_2); +x_4 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558_(x_3, x_2); lean_dec(x_2); return x_4; } @@ -16984,7 +17822,7 @@ static lean_object* _init_l_Lean_Meta_instReprTransparencyMode___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____boxed), 2, 0); return x_1; } } @@ -17067,7 +17905,7 @@ x_1 = 0; return x_1; } } -LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Meta_beqEtaStructMode____x40_Init_Meta___hyg_8732_(uint8_t x_1, uint8_t x_2) { +LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Meta_beqEtaStructMode____x40_Init_Meta___hyg_9707_(uint8_t x_1, uint8_t x_2) { _start: { lean_object* x_3; lean_object* x_4; uint8_t x_5; @@ -17079,7 +17917,7 @@ lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_beqEtaStructMode____x40_Init_Meta___hyg_8732____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_beqEtaStructMode____x40_Init_Meta___hyg_9707____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; @@ -17087,7 +17925,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_Init_Meta_0__Lean_Meta_beqEtaStructMode____x40_Init_Meta___hyg_8732_(x_3, x_4); +x_5 = l___private_Init_Meta_0__Lean_Meta_beqEtaStructMode____x40_Init_Meta___hyg_9707_(x_3, x_4); x_6 = lean_box(x_5); return x_6; } @@ -17096,7 +17934,7 @@ static lean_object* _init_l_Lean_Meta_instBEqEtaStructMode___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_beqEtaStructMode____x40_Init_Meta___hyg_8732____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_beqEtaStructMode____x40_Init_Meta___hyg_9707____boxed), 2, 0); return x_1; } } @@ -17108,7 +17946,7 @@ x_1 = l_Lean_Meta_instBEqEtaStructMode___closed__1; return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__1() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__1() { _start: { lean_object* x_1; @@ -17116,33 +17954,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Meta.EtaStructMode.all", 27); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__2() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__1; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____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_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__3() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__3; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__2; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____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_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__4() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__3; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__3; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -17150,23 +17988,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__5() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__6; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__2; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____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_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__6() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__5; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__5; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -17174,7 +18012,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__7() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__7() { _start: { lean_object* x_1; @@ -17182,33 +18020,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Meta.EtaStructMode.notClasses", 34); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__8() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__7; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__7; 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_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__9() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__3; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__8; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__8; 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_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__10() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__10() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__9; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__9; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -17216,23 +18054,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__11() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__6; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__8; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__8; 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_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__12() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__12() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__11; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__11; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -17240,7 +18078,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__13() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__13() { _start: { lean_object* x_1; @@ -17248,33 +18086,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Meta.EtaStructMode.none", 28); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__14() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__14() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__13; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__13; 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_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__15() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__3; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__14; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__14; 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_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__16() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__16() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__15; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__15; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -17282,23 +18120,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__17() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__6; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__14; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__14; 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_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__18() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__18() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__17; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__17; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -17306,7 +18144,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748_(uint8_t x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723_(uint8_t x_1, lean_object* x_2) { _start: { switch (x_1) { @@ -17318,14 +18156,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_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__4; +x_5 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__4; 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_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__6; +x_7 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__6; x_8 = l_Repr_addAppParen(x_7, x_2); return x_8; } @@ -17338,14 +18176,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_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__10; +x_11 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__10; 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_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__12; +x_13 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__12; x_14 = l_Repr_addAppParen(x_13, x_2); return x_14; } @@ -17358,14 +18196,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_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__16; +x_17 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__16; 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_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__18; +x_19 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__18; x_20 = l_Repr_addAppParen(x_19, x_2); return x_20; } @@ -17373,13 +18211,13 @@ return x_20; } } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____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_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748_(x_3, x_2); +x_4 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723_(x_3, x_2); lean_dec(x_2); return x_4; } @@ -17388,7 +18226,7 @@ static lean_object* _init_l_Lean_Meta_instReprEtaStructMode___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____boxed), 2, 0); return x_1; } } @@ -17490,7 +18328,7 @@ x_1 = l_Lean_Meta_DSimp_instInhabitedConfig___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Meta_DSimp_beqConfig____x40_Init_Meta___hyg_8947_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Meta_DSimp_beqConfig____x40_Init_Meta___hyg_9922_(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; uint8_t x_4; uint8_t x_5; uint8_t x_6; uint8_t x_7; uint8_t x_8; uint8_t x_9; uint8_t x_10; uint8_t x_11; uint8_t x_12; uint8_t x_13; uint8_t x_14; uint8_t x_15; uint8_t x_16; uint8_t x_17; uint8_t x_18; lean_object* x_19; lean_object* x_23; lean_object* x_29; lean_object* x_35; lean_object* x_43; uint8_t x_49; @@ -17643,7 +18481,7 @@ goto block_28; { uint8_t x_36; lean_dec(x_35); -x_36 = l___private_Init_Meta_0__Lean_Meta_beqEtaStructMode____x40_Init_Meta___hyg_8732_(x_6, x_14); +x_36 = l___private_Init_Meta_0__Lean_Meta_beqEtaStructMode____x40_Init_Meta___hyg_9707_(x_6, x_14); if (x_36 == 0) { uint8_t x_37; @@ -17768,11 +18606,11 @@ goto block_48; } } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_beqConfig____x40_Init_Meta___hyg_8947____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_beqConfig____x40_Init_Meta___hyg_9922____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_Meta_0__Lean_Meta_DSimp_beqConfig____x40_Init_Meta___hyg_8947_(x_1, x_2); +x_3 = l___private_Init_Meta_0__Lean_Meta_DSimp_beqConfig____x40_Init_Meta___hyg_9922_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -17783,7 +18621,7 @@ static lean_object* _init_l_Lean_Meta_DSimp_instBEqConfig___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_DSimp_beqConfig____x40_Init_Meta___hyg_8947____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_DSimp_beqConfig____x40_Init_Meta___hyg_9922____boxed), 2, 0); return x_1; } } @@ -17795,7 +18633,7 @@ x_1 = l_Lean_Meta_DSimp_instBEqConfig___closed__1; return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__1() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__1() { _start: { lean_object* x_1; @@ -17803,29 +18641,29 @@ x_1 = lean_mk_string_from_bytes("zeta", 4); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__2() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__1; +x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____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_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__3() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__2; +x_2 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__2; x_3 = lean_alloc_ctor(4, 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_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__4() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__4() { _start: { lean_object* x_1; @@ -17833,29 +18671,29 @@ x_1 = lean_mk_string_from_bytes(" := ", 4); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__5() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__4; +x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__4; 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_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__6() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__3; -x_2 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__5; +x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__3; +x_2 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__5; x_3 = lean_alloc_ctor(4, 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_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__7() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__7() { _start: { lean_object* x_1; @@ -17863,17 +18701,17 @@ x_1 = lean_mk_string_from_bytes("beta", 4); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__8() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__7; +x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__7; 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_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__9() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__9() { _start: { lean_object* x_1; @@ -17881,17 +18719,17 @@ x_1 = lean_mk_string_from_bytes("eta", 3); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__10() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__9; +x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__9; 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_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__11() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__11() { _start: { lean_object* x_1; @@ -17899,17 +18737,17 @@ x_1 = lean_mk_string_from_bytes("etaStruct", 9); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__12() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__11; +x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__11; 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_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__13() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__13() { _start: { lean_object* x_1; @@ -17917,17 +18755,17 @@ x_1 = lean_mk_string_from_bytes("iota", 4); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__14() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__14() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__13; +x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__13; 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_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__15() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__15() { _start: { lean_object* x_1; @@ -17935,17 +18773,17 @@ x_1 = lean_mk_string_from_bytes("proj", 4); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__16() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__16() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__15; +x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__15; 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_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__17() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__17() { _start: { lean_object* x_1; @@ -17953,17 +18791,17 @@ x_1 = lean_mk_string_from_bytes("decide", 6); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__18() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__18() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__17; +x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__17; 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_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__19() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__19() { _start: { lean_object* x_1; @@ -17971,17 +18809,17 @@ x_1 = lean_mk_string_from_bytes("autoUnfold", 10); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__20() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__20() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__19; +x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__19; 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_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__21() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__21() { _start: { lean_object* x_1; @@ -17989,35 +18827,35 @@ x_1 = lean_mk_string_from_bytes("{ ", 2); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__22() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__22() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__21; +x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__21; x_2 = lean_string_length(x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__23() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__23() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__22; +x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__22; x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__24() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__24() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__21; +x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__21; 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_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__25() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__25() { _start: { lean_object* x_1; @@ -18025,37 +18863,37 @@ x_1 = lean_mk_string_from_bytes(" }", 2); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__26() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__26() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__25; +x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__25; 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_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__27() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__27() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_instQuoteBool___closed__3; +x_1 = l_Lean_instQuoteBoolMkStrAnonymous___closed__3; 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_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__28() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__28() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_instQuoteBool___closed__6; +x_1 = l_Lean_instQuoteBoolMkStrAnonymous___closed__6; x_2 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089_(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; uint8_t x_4; uint8_t x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; uint8_t x_10; uint8_t x_11; uint8_t x_12; lean_object* x_13; @@ -18064,7 +18902,7 @@ x_4 = lean_ctor_get_uint8(x_1, 1); x_5 = lean_ctor_get_uint8(x_1, 2); x_6 = lean_ctor_get_uint8(x_1, 3); x_7 = lean_unsigned_to_nat(0u); -x_8 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748_(x_6, x_7); +x_8 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723_(x_6, x_7); x_9 = lean_ctor_get_uint8(x_1, 4); x_10 = lean_ctor_get_uint8(x_1, 5); x_11 = lean_ctor_get_uint8(x_1, 6); @@ -18072,21 +18910,21 @@ x_12 = lean_ctor_get_uint8(x_1, 7); if (x_3 == 0) { lean_object* x_123; -x_123 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__27; +x_123 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__27; x_13 = x_123; goto block_122; } else { lean_object* x_124; -x_124 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__28; +x_124 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__28; x_13 = x_124; goto block_122; } block_122: { 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 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__6; +x_14 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__6; x_15 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_15, 0, x_14); lean_ctor_set(x_15, 1, x_13); @@ -18098,25 +18936,25 @@ x_18 = lean_box(1); x_19 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_19, 0, x_17); lean_ctor_set(x_19, 1, x_18); -x_20 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__8; +x_20 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__8; x_21 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_21, 0, x_19); lean_ctor_set(x_21, 1, x_20); -x_22 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__5; +x_22 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__5; x_23 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_23, 0, x_21); lean_ctor_set(x_23, 1, x_22); if (x_4 == 0) { lean_object* x_120; -x_120 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__27; +x_120 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__27; x_24 = x_120; goto block_119; } else { lean_object* x_121; -x_121 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__28; +x_121 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__28; x_24 = x_121; goto block_119; } @@ -18132,7 +18970,7 @@ lean_ctor_set(x_26, 1, x_16); x_27 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_27, 0, x_26); lean_ctor_set(x_27, 1, x_18); -x_28 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__10; +x_28 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__10; x_29 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_29, 0, x_27); lean_ctor_set(x_29, 1, x_28); @@ -18142,14 +18980,14 @@ lean_ctor_set(x_30, 1, x_22); if (x_5 == 0) { lean_object* x_117; -x_117 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__27; +x_117 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__27; x_31 = x_117; goto block_116; } else { lean_object* x_118; -x_118 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__28; +x_118 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__28; x_31 = x_118; goto block_116; } @@ -18165,7 +19003,7 @@ lean_ctor_set(x_33, 1, x_16); x_34 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_34, 0, x_33); lean_ctor_set(x_34, 1, x_18); -x_35 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__12; +x_35 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__12; x_36 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_36, 0, x_34); lean_ctor_set(x_36, 1, x_35); @@ -18181,7 +19019,7 @@ lean_ctor_set(x_39, 1, x_16); x_40 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_40, 0, x_39); lean_ctor_set(x_40, 1, x_18); -x_41 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__14; +x_41 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__14; x_42 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_42, 0, x_40); lean_ctor_set(x_42, 1, x_41); @@ -18191,14 +19029,14 @@ lean_ctor_set(x_43, 1, x_22); if (x_9 == 0) { lean_object* x_114; -x_114 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__27; +x_114 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__27; x_44 = x_114; goto block_113; } else { lean_object* x_115; -x_115 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__28; +x_115 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__28; x_44 = x_115; goto block_113; } @@ -18214,7 +19052,7 @@ lean_ctor_set(x_46, 1, x_16); x_47 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_47, 0, x_46); lean_ctor_set(x_47, 1, x_18); -x_48 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__16; +x_48 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__16; x_49 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_49, 0, x_47); lean_ctor_set(x_49, 1, x_48); @@ -18224,14 +19062,14 @@ lean_ctor_set(x_50, 1, x_22); if (x_10 == 0) { lean_object* x_111; -x_111 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__27; +x_111 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__27; x_51 = x_111; goto block_110; } else { lean_object* x_112; -x_112 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__28; +x_112 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__28; x_51 = x_112; goto block_110; } @@ -18247,7 +19085,7 @@ lean_ctor_set(x_53, 1, x_16); x_54 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_54, 0, x_53); lean_ctor_set(x_54, 1, x_18); -x_55 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__18; +x_55 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__18; x_56 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_56, 0, x_54); lean_ctor_set(x_56, 1, x_55); @@ -18257,7 +19095,7 @@ lean_ctor_set(x_57, 1, x_22); if (x_11 == 0) { lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_58 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__27; +x_58 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__27; x_59 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_59, 0, x_57); lean_ctor_set(x_59, 1, x_58); @@ -18267,7 +19105,7 @@ lean_ctor_set(x_60, 1, x_16); x_61 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_61, 0, x_60); lean_ctor_set(x_61, 1, x_18); -x_62 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__20; +x_62 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__20; x_63 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_63, 0, x_61); lean_ctor_set(x_63, 1, x_62); @@ -18280,15 +19118,15 @@ lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean x_65 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_65, 0, x_64); lean_ctor_set(x_65, 1, x_58); -x_66 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__24; +x_66 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__24; x_67 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_67, 0, x_66); lean_ctor_set(x_67, 1, x_65); -x_68 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__26; +x_68 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__26; x_69 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_69, 0, x_67); lean_ctor_set(x_69, 1, x_68); -x_70 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__23; +x_70 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__23; x_71 = lean_alloc_ctor(3, 2, 0); lean_ctor_set(x_71, 0, x_70); lean_ctor_set(x_71, 1, x_69); @@ -18301,19 +19139,19 @@ return x_73; else { lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; uint8_t x_82; lean_object* x_83; -x_74 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__28; +x_74 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__28; x_75 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_75, 0, x_64); lean_ctor_set(x_75, 1, x_74); -x_76 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__24; +x_76 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__24; x_77 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_77, 0, x_76); lean_ctor_set(x_77, 1, x_75); -x_78 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__26; +x_78 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__26; x_79 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_79, 0, x_77); lean_ctor_set(x_79, 1, x_78); -x_80 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__23; +x_80 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__23; x_81 = lean_alloc_ctor(3, 2, 0); lean_ctor_set(x_81, 0, x_80); lean_ctor_set(x_81, 1, x_79); @@ -18327,7 +19165,7 @@ return x_83; else { lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; -x_84 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__28; +x_84 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__28; x_85 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_85, 0, x_57); lean_ctor_set(x_85, 1, x_84); @@ -18337,7 +19175,7 @@ lean_ctor_set(x_86, 1, x_16); x_87 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_87, 0, x_86); lean_ctor_set(x_87, 1, x_18); -x_88 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__20; +x_88 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__20; x_89 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_89, 0, x_87); lean_ctor_set(x_89, 1, x_88); @@ -18347,19 +19185,19 @@ lean_ctor_set(x_90, 1, x_22); if (x_12 == 0) { lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; uint8_t x_99; lean_object* x_100; -x_91 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__27; +x_91 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__27; x_92 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_92, 0, x_90); lean_ctor_set(x_92, 1, x_91); -x_93 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__24; +x_93 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__24; x_94 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_94, 0, x_93); lean_ctor_set(x_94, 1, x_92); -x_95 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__26; +x_95 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__26; x_96 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_96, 0, x_94); lean_ctor_set(x_96, 1, x_95); -x_97 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__23; +x_97 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__23; x_98 = lean_alloc_ctor(3, 2, 0); lean_ctor_set(x_98, 0, x_97); lean_ctor_set(x_98, 1, x_96); @@ -18375,15 +19213,15 @@ lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; x_101 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_101, 0, x_90); lean_ctor_set(x_101, 1, x_84); -x_102 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__24; +x_102 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__24; x_103 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_103, 0, x_102); lean_ctor_set(x_103, 1, x_101); -x_104 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__26; +x_104 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__26; x_105 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_105, 0, x_103); lean_ctor_set(x_105, 1, x_104); -x_106 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__23; +x_106 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__23; x_107 = lean_alloc_ctor(3, 2, 0); lean_ctor_set(x_107, 0, x_106); lean_ctor_set(x_107, 1, x_105); @@ -18401,11 +19239,11 @@ return x_109; } } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114_(x_1, x_2); +x_3 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; @@ -18415,7 +19253,7 @@ static lean_object* _init_l_Lean_Meta_DSimp_instReprConfig___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____boxed), 2, 0); return x_1; } } @@ -18580,7 +19418,7 @@ x_1 = l_Lean_Meta_Simp_instInhabitedConfig___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Meta_Simp_beqConfig____x40_Init_Meta___hyg_9383_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Meta_Simp_beqConfig____x40_Init_Meta___hyg_10358_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; uint8_t x_5; uint8_t x_6; uint8_t x_7; uint8_t x_8; uint8_t x_9; uint8_t x_10; uint8_t x_11; uint8_t x_12; uint8_t x_13; uint8_t x_14; uint8_t x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; uint8_t x_20; uint8_t x_21; uint8_t x_22; uint8_t x_23; uint8_t x_24; uint8_t x_25; uint8_t x_26; uint8_t x_27; uint8_t x_28; uint8_t x_29; uint8_t x_30; lean_object* x_31; lean_object* x_35; lean_object* x_41; lean_object* x_47; lean_object* x_53; lean_object* x_61; lean_object* x_67; lean_object* x_73; lean_object* x_79; lean_object* x_85; uint8_t x_91; @@ -18800,7 +19638,7 @@ goto block_46; { uint8_t x_54; lean_dec(x_53); -x_54 = l___private_Init_Meta_0__Lean_Meta_beqEtaStructMode____x40_Init_Meta___hyg_8732_(x_11, x_25); +x_54 = l___private_Init_Meta_0__Lean_Meta_beqEtaStructMode____x40_Init_Meta___hyg_9707_(x_11, x_25); if (x_54 == 0) { uint8_t x_55; @@ -19025,11 +19863,11 @@ goto block_84; } } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_beqConfig____x40_Init_Meta___hyg_9383____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_beqConfig____x40_Init_Meta___hyg_10358____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_Meta_0__Lean_Meta_Simp_beqConfig____x40_Init_Meta___hyg_9383_(x_1, x_2); +x_3 = l___private_Init_Meta_0__Lean_Meta_Simp_beqConfig____x40_Init_Meta___hyg_10358_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -19040,7 +19878,7 @@ static lean_object* _init_l_Lean_Meta_Simp_instBEqConfig___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_Simp_beqConfig____x40_Init_Meta___hyg_9383____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_Simp_beqConfig____x40_Init_Meta___hyg_10358____boxed), 2, 0); return x_1; } } @@ -19052,7 +19890,7 @@ x_1 = l_Lean_Meta_Simp_instBEqConfig___closed__1; return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__1() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__1() { _start: { lean_object* x_1; @@ -19060,41 +19898,41 @@ x_1 = lean_mk_string_from_bytes("maxSteps", 8); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__2() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__1; +x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____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_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__3() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__2; +x_2 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__2; x_3 = lean_alloc_ctor(4, 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_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__4() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__3; -x_2 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__5; +x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__3; +x_2 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__5; x_3 = lean_alloc_ctor(4, 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_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__5() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__5() { _start: { lean_object* x_1; @@ -19102,17 +19940,17 @@ x_1 = lean_mk_string_from_bytes("maxDischargeDepth", 17); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__6() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__5; +x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__5; 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_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__7() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__7() { _start: { lean_object* x_1; @@ -19120,17 +19958,17 @@ x_1 = lean_mk_string_from_bytes("contextual", 10); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__8() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__7; +x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__7; 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_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__9() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__9() { _start: { lean_object* x_1; @@ -19138,17 +19976,17 @@ x_1 = lean_mk_string_from_bytes("memoize", 7); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__10() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__9; +x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__9; 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_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__11() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__11() { _start: { lean_object* x_1; @@ -19156,17 +19994,17 @@ x_1 = lean_mk_string_from_bytes("singlePass", 10); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__12() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__11; +x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__11; 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_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__13() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__13() { _start: { lean_object* x_1; @@ -19174,17 +20012,17 @@ x_1 = lean_mk_string_from_bytes("arith", 5); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__14() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__14() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__13; +x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__13; x_2 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621_(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; 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; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; uint8_t x_26; uint8_t x_27; uint8_t x_28; uint8_t x_29; uint8_t x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; uint8_t x_35; uint8_t x_36; uint8_t x_37; uint8_t x_38; lean_object* x_39; @@ -19193,7 +20031,7 @@ lean_inc(x_3); x_4 = l_Nat_repr(x_3); x_5 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_5, 0, x_4); -x_6 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__4; +x_6 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__4; x_7 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); @@ -19205,11 +20043,11 @@ x_10 = lean_box(1); x_11 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_11, 0, x_9); lean_ctor_set(x_11, 1, x_10); -x_12 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__6; +x_12 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__6; x_13 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_13, 0, x_11); lean_ctor_set(x_13, 1, x_12); -x_14 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__5; +x_14 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__5; x_15 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_15, 0, x_13); lean_ctor_set(x_15, 1, x_14); @@ -19227,7 +20065,7 @@ lean_ctor_set(x_20, 1, x_8); x_21 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_21, 0, x_20); lean_ctor_set(x_21, 1, x_10); -x_22 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__8; +x_22 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__8; x_23 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_23, 0, x_21); lean_ctor_set(x_23, 1, x_22); @@ -19242,7 +20080,7 @@ x_29 = lean_ctor_get_uint8(x_1, sizeof(void*)*2 + 4); x_30 = lean_ctor_get_uint8(x_1, sizeof(void*)*2 + 5); x_31 = lean_ctor_get_uint8(x_1, sizeof(void*)*2 + 6); x_32 = lean_unsigned_to_nat(0u); -x_33 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748_(x_31, x_32); +x_33 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723_(x_31, x_32); x_34 = lean_ctor_get_uint8(x_1, sizeof(void*)*2 + 7); x_35 = lean_ctor_get_uint8(x_1, sizeof(void*)*2 + 8); x_36 = lean_ctor_get_uint8(x_1, sizeof(void*)*2 + 9); @@ -19252,14 +20090,14 @@ lean_dec(x_1); if (x_25 == 0) { lean_object* x_185; -x_185 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__27; +x_185 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__27; x_39 = x_185; goto block_184; } else { lean_object* x_186; -x_186 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__28; +x_186 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__28; x_39 = x_186; goto block_184; } @@ -19275,7 +20113,7 @@ lean_ctor_set(x_41, 1, x_8); x_42 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_42, 0, x_41); lean_ctor_set(x_42, 1, x_10); -x_43 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__10; +x_43 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__10; x_44 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_44, 0, x_42); lean_ctor_set(x_44, 1, x_43); @@ -19285,14 +20123,14 @@ lean_ctor_set(x_45, 1, x_14); if (x_26 == 0) { lean_object* x_182; -x_182 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__27; +x_182 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__27; x_46 = x_182; goto block_181; } else { lean_object* x_183; -x_183 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__28; +x_183 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__28; x_46 = x_183; goto block_181; } @@ -19308,7 +20146,7 @@ lean_ctor_set(x_48, 1, x_8); x_49 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_49, 0, x_48); lean_ctor_set(x_49, 1, x_10); -x_50 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__12; +x_50 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__12; x_51 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_51, 0, x_49); lean_ctor_set(x_51, 1, x_50); @@ -19318,14 +20156,14 @@ lean_ctor_set(x_52, 1, x_14); if (x_27 == 0) { lean_object* x_179; -x_179 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__27; +x_179 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__27; x_53 = x_179; goto block_178; } else { lean_object* x_180; -x_180 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__28; +x_180 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__28; x_53 = x_180; goto block_178; } @@ -19341,7 +20179,7 @@ lean_ctor_set(x_55, 1, x_8); x_56 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_56, 0, x_55); lean_ctor_set(x_56, 1, x_10); -x_57 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__2; +x_57 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__2; x_58 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_58, 0, x_56); lean_ctor_set(x_58, 1, x_57); @@ -19351,14 +20189,14 @@ lean_ctor_set(x_59, 1, x_14); if (x_28 == 0) { lean_object* x_176; -x_176 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__27; +x_176 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__27; x_60 = x_176; goto block_175; } else { lean_object* x_177; -x_177 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__28; +x_177 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__28; x_60 = x_177; goto block_175; } @@ -19374,7 +20212,7 @@ lean_ctor_set(x_62, 1, x_8); x_63 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_63, 0, x_62); lean_ctor_set(x_63, 1, x_10); -x_64 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__8; +x_64 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__8; x_65 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_65, 0, x_63); lean_ctor_set(x_65, 1, x_64); @@ -19384,14 +20222,14 @@ lean_ctor_set(x_66, 1, x_14); if (x_29 == 0) { lean_object* x_173; -x_173 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__27; +x_173 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__27; x_67 = x_173; goto block_172; } else { lean_object* x_174; -x_174 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__28; +x_174 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__28; x_67 = x_174; goto block_172; } @@ -19407,7 +20245,7 @@ lean_ctor_set(x_69, 1, x_8); x_70 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_70, 0, x_69); lean_ctor_set(x_70, 1, x_10); -x_71 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__10; +x_71 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__10; x_72 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_72, 0, x_70); lean_ctor_set(x_72, 1, x_71); @@ -19417,14 +20255,14 @@ lean_ctor_set(x_73, 1, x_14); if (x_30 == 0) { lean_object* x_170; -x_170 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__27; +x_170 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__27; x_74 = x_170; goto block_169; } else { lean_object* x_171; -x_171 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__28; +x_171 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__28; x_74 = x_171; goto block_169; } @@ -19440,7 +20278,7 @@ lean_ctor_set(x_76, 1, x_8); x_77 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_77, 0, x_76); lean_ctor_set(x_77, 1, x_10); -x_78 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__12; +x_78 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__12; x_79 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_79, 0, x_77); lean_ctor_set(x_79, 1, x_78); @@ -19456,7 +20294,7 @@ lean_ctor_set(x_82, 1, x_8); x_83 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_83, 0, x_82); lean_ctor_set(x_83, 1, x_10); -x_84 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__14; +x_84 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__14; x_85 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_85, 0, x_83); lean_ctor_set(x_85, 1, x_84); @@ -19466,14 +20304,14 @@ lean_ctor_set(x_86, 1, x_14); if (x_34 == 0) { lean_object* x_167; -x_167 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__27; +x_167 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__27; x_87 = x_167; goto block_166; } else { lean_object* x_168; -x_168 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__28; +x_168 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__28; x_87 = x_168; goto block_166; } @@ -19489,7 +20327,7 @@ lean_ctor_set(x_89, 1, x_8); x_90 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_90, 0, x_89); lean_ctor_set(x_90, 1, x_10); -x_91 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__16; +x_91 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__16; x_92 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_92, 0, x_90); lean_ctor_set(x_92, 1, x_91); @@ -19499,14 +20337,14 @@ lean_ctor_set(x_93, 1, x_14); if (x_35 == 0) { lean_object* x_164; -x_164 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__27; +x_164 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__27; x_94 = x_164; goto block_163; } else { lean_object* x_165; -x_165 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__28; +x_165 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__28; x_94 = x_165; goto block_163; } @@ -19522,7 +20360,7 @@ lean_ctor_set(x_96, 1, x_8); x_97 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_97, 0, x_96); lean_ctor_set(x_97, 1, x_10); -x_98 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__18; +x_98 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__18; x_99 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_99, 0, x_97); lean_ctor_set(x_99, 1, x_98); @@ -19532,14 +20370,14 @@ lean_ctor_set(x_100, 1, x_14); if (x_36 == 0) { lean_object* x_161; -x_161 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__27; +x_161 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__27; x_101 = x_161; goto block_160; } else { lean_object* x_162; -x_162 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__28; +x_162 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__28; x_101 = x_162; goto block_160; } @@ -19555,7 +20393,7 @@ lean_ctor_set(x_103, 1, x_8); x_104 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_104, 0, x_103); lean_ctor_set(x_104, 1, x_10); -x_105 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__14; +x_105 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__14; x_106 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_106, 0, x_104); lean_ctor_set(x_106, 1, x_105); @@ -19565,7 +20403,7 @@ lean_ctor_set(x_107, 1, x_14); if (x_37 == 0) { lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; -x_108 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__27; +x_108 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__27; x_109 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_109, 0, x_107); lean_ctor_set(x_109, 1, x_108); @@ -19575,7 +20413,7 @@ lean_ctor_set(x_110, 1, x_8); x_111 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_111, 0, x_110); lean_ctor_set(x_111, 1, x_10); -x_112 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__20; +x_112 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__20; x_113 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_113, 0, x_111); lean_ctor_set(x_113, 1, x_112); @@ -19588,15 +20426,15 @@ lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; x_115 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_115, 0, x_114); lean_ctor_set(x_115, 1, x_108); -x_116 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__24; +x_116 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__24; x_117 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_117, 0, x_116); lean_ctor_set(x_117, 1, x_115); -x_118 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__26; +x_118 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__26; x_119 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_119, 0, x_117); lean_ctor_set(x_119, 1, x_118); -x_120 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__23; +x_120 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__23; x_121 = lean_alloc_ctor(3, 2, 0); lean_ctor_set(x_121, 0, x_120); lean_ctor_set(x_121, 1, x_119); @@ -19609,19 +20447,19 @@ return x_123; else { lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; uint8_t x_132; lean_object* x_133; -x_124 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__28; +x_124 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__28; x_125 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_125, 0, x_114); lean_ctor_set(x_125, 1, x_124); -x_126 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__24; +x_126 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__24; x_127 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_127, 0, x_126); lean_ctor_set(x_127, 1, x_125); -x_128 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__26; +x_128 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__26; x_129 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_129, 0, x_127); lean_ctor_set(x_129, 1, x_128); -x_130 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__23; +x_130 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__23; x_131 = lean_alloc_ctor(3, 2, 0); lean_ctor_set(x_131, 0, x_130); lean_ctor_set(x_131, 1, x_129); @@ -19635,7 +20473,7 @@ return x_133; else { lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; -x_134 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__28; +x_134 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__28; x_135 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_135, 0, x_107); lean_ctor_set(x_135, 1, x_134); @@ -19645,7 +20483,7 @@ lean_ctor_set(x_136, 1, x_8); x_137 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_137, 0, x_136); lean_ctor_set(x_137, 1, x_10); -x_138 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__20; +x_138 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__20; x_139 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_139, 0, x_137); lean_ctor_set(x_139, 1, x_138); @@ -19655,19 +20493,19 @@ lean_ctor_set(x_140, 1, x_14); if (x_38 == 0) { lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; uint8_t x_149; lean_object* x_150; -x_141 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__27; +x_141 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__27; x_142 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_142, 0, x_140); lean_ctor_set(x_142, 1, x_141); -x_143 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__24; +x_143 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__24; x_144 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_144, 0, x_143); lean_ctor_set(x_144, 1, x_142); -x_145 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__26; +x_145 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__26; x_146 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_146, 0, x_144); lean_ctor_set(x_146, 1, x_145); -x_147 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__23; +x_147 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__23; x_148 = lean_alloc_ctor(3, 2, 0); lean_ctor_set(x_148, 0, x_147); lean_ctor_set(x_148, 1, x_146); @@ -19683,15 +20521,15 @@ lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; x_151 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_151, 0, x_140); lean_ctor_set(x_151, 1, x_134); -x_152 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__24; +x_152 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__24; x_153 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_153, 0, x_152); lean_ctor_set(x_153, 1, x_151); -x_154 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__26; +x_154 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__26; x_155 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_155, 0, x_153); lean_ctor_set(x_155, 1, x_154); -x_156 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__23; +x_156 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__23; x_157 = lean_alloc_ctor(3, 2, 0); lean_ctor_set(x_157, 0, x_156); lean_ctor_set(x_157, 1, x_155); @@ -19713,11 +20551,11 @@ return x_159; } } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646_(x_1, x_2); +x_3 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621_(x_1, x_2); lean_dec(x_2); return x_3; } @@ -19726,7 +20564,7 @@ static lean_object* _init_l_Lean_Meta_Simp_instReprConfig___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____boxed), 2, 0); return x_1; } } @@ -20034,7 +20872,7 @@ static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__2; +x_2 = l_Lean_mkNullNode___closed__2; x_3 = l_Lean_mkSepArray___closed__1; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -20124,7 +20962,7 @@ static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__7; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__7; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } @@ -20133,7 +20971,7 @@ static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__7; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__7; x_2 = lean_unsigned_to_nat(0u); x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__21; x_4 = lean_alloc_ctor(0, 3, 0); @@ -20312,7 +21150,7 @@ static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__2; +x_2 = l_Lean_mkNullNode___closed__2; x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__39; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -20341,389 +21179,250 @@ return x_7; } 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; uint8_t x_14; +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; x_8 = lean_unsigned_to_nat(1u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); -x_10 = lean_unsigned_to_nat(2u); -x_11 = l_Lean_Syntax_getArg(x_1, x_10); +x_105 = lean_unsigned_to_nat(2u); +x_106 = l_Lean_Syntax_getArg(x_1, x_105); lean_dec(x_1); -x_12 = l_Lean_Syntax_getOptional_x3f(x_11); -lean_dec(x_11); +x_107 = l_Lean_Syntax_getOptional_x3f(x_106); +lean_dec(x_106); lean_inc(x_2); -x_13 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) +x_108 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +if (lean_obj_tag(x_107) == 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; lean_object* 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; -x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_ctor_get(x_2, 2); -lean_inc(x_16); -x_17 = lean_ctor_get(x_2, 1); -lean_inc(x_17); -lean_dec(x_2); -x_18 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__5; -lean_inc(x_15); -x_19 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_19, 0, x_15); -lean_ctor_set(x_19, 1, x_18); -x_20 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__4; -lean_inc(x_15); -x_21 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_21, 0, x_15); -lean_ctor_set(x_21, 1, x_20); -x_22 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__6; -lean_inc(x_15); -x_23 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_23, 0, x_15); -lean_ctor_set(x_23, 1, x_22); -x_24 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__8; -lean_inc(x_15); -x_25 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_25, 0, x_15); -lean_ctor_set(x_25, 1, x_24); -x_26 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__11; -lean_inc(x_15); -x_27 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_27, 0, x_15); -lean_ctor_set(x_27, 1, x_26); -x_28 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__20; -lean_inc(x_16); -lean_inc(x_17); -x_29 = l_Lean_addMacroScope(x_17, x_28, x_16); -x_30 = lean_box(0); -x_31 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__19; -lean_inc(x_15); -x_32 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_32, 0, x_15); -lean_ctor_set(x_32, 1, x_31); -lean_ctor_set(x_32, 2, x_29); -lean_ctor_set(x_32, 3, x_30); -x_33 = l_Lean_Syntax_mkApp___closed__3; -x_34 = lean_array_push(x_33, x_32); -x_35 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__12; -x_36 = lean_array_push(x_34, x_35); -x_37 = lean_box(2); -x_38 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__16; -x_39 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_39, 0, x_37); -lean_ctor_set(x_39, 1, x_38); -lean_ctor_set(x_39, 2, x_36); -x_40 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__28; -x_41 = l_Lean_addMacroScope(x_17, x_40, x_16); -x_42 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__22; -x_43 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__30; -lean_inc(x_15); -x_44 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_44, 0, x_15); -lean_ctor_set(x_44, 1, x_42); -lean_ctor_set(x_44, 2, x_41); -lean_ctor_set(x_44, 3, x_43); -x_45 = l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__4; -x_46 = lean_array_push(x_45, x_39); -lean_inc(x_25); -x_47 = lean_array_push(x_46, x_25); -x_48 = lean_array_push(x_47, x_44); -x_49 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__14; -x_50 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_50, 0, x_37); -lean_ctor_set(x_50, 1, x_49); -lean_ctor_set(x_50, 2, x_48); -x_51 = l_Lean_mkOptionalNode___closed__2; -x_52 = lean_array_push(x_51, x_50); -x_53 = l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__2; -x_54 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_54, 0, x_37); -lean_ctor_set(x_54, 1, x_53); -lean_ctor_set(x_54, 2, x_52); -x_55 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__35; -lean_inc(x_15); -x_56 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_56, 0, x_15); -lean_ctor_set(x_56, 1, x_55); -x_57 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__36; -x_58 = lean_array_push(x_57, x_27); -x_59 = lean_array_push(x_58, x_35); -x_60 = lean_array_push(x_59, x_54); -x_61 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__34; -x_62 = lean_array_push(x_60, x_61); -x_63 = lean_array_push(x_62, x_35); -x_64 = lean_array_push(x_63, x_56); -x_65 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__10; -x_66 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_66, 0, x_37); -lean_ctor_set(x_66, 1, x_65); -lean_ctor_set(x_66, 2, x_64); -x_67 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__8; -x_68 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_68, 0, x_15); -lean_ctor_set(x_68, 1, x_67); -x_69 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__37; -x_70 = lean_array_push(x_69, x_21); -x_71 = lean_array_push(x_70, x_23); -x_72 = lean_array_push(x_71, x_25); -x_73 = lean_array_push(x_72, x_66); -x_74 = lean_array_push(x_73, x_68); -x_75 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__7; -x_76 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_76, 0, x_37); -lean_ctor_set(x_76, 1, x_75); -lean_ctor_set(x_76, 2, x_74); -x_77 = lean_array_push(x_51, x_76); -x_78 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_78, 0, x_37); -lean_ctor_set(x_78, 1, x_53); -lean_ctor_set(x_78, 2, x_77); -x_79 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__38; -x_80 = lean_array_push(x_79, x_19); -x_81 = lean_array_push(x_80, x_78); -x_82 = lean_array_push(x_81, x_9); -if (lean_obj_tag(x_12) == 0) +lean_object* x_109; lean_object* x_110; lean_object* x_111; +x_109 = lean_ctor_get(x_108, 0); +lean_inc(x_109); +x_110 = lean_ctor_get(x_108, 1); +lean_inc(x_110); +lean_dec(x_108); +x_111 = lean_box(0); +x_10 = x_111; +x_11 = x_109; +x_12 = x_110; +goto block_104; +} +else { -lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; -x_83 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__40; -x_84 = lean_array_push(x_82, x_83); -x_85 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__4; -x_86 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_86, 0, x_37); -lean_ctor_set(x_86, 1, x_85); -lean_ctor_set(x_86, 2, x_84); -x_87 = lean_array_push(x_51, x_86); -x_88 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_88, 0, x_37); -lean_ctor_set(x_88, 1, x_53); -lean_ctor_set(x_88, 2, x_87); -x_89 = lean_array_push(x_51, x_88); -x_90 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__2; -x_91 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_91, 0, x_37); -lean_ctor_set(x_91, 1, x_90); -lean_ctor_set(x_91, 2, x_89); -lean_ctor_set(x_13, 0, x_91); -return x_13; +uint8_t x_112; +x_112 = !lean_is_exclusive(x_107); +if (x_112 == 0) +{ +lean_object* x_113; lean_object* x_114; +x_113 = lean_ctor_get(x_108, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_108, 1); +lean_inc(x_114); +lean_dec(x_108); +x_10 = x_107; +x_11 = x_113; +x_12 = x_114; +goto block_104; } else { -lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; -x_92 = lean_ctor_get(x_12, 0); -lean_inc(x_92); -lean_dec(x_12); -x_93 = lean_array_push(x_51, x_92); -x_94 = l_Lean_mkSepArray___closed__1; -x_95 = l_Array_append___rarg(x_94, x_93); -x_96 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_96, 0, x_37); -lean_ctor_set(x_96, 1, x_53); -lean_ctor_set(x_96, 2, x_95); -x_97 = lean_array_push(x_82, x_96); -x_98 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__4; -x_99 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_99, 0, x_37); -lean_ctor_set(x_99, 1, x_98); -lean_ctor_set(x_99, 2, x_97); -x_100 = lean_array_push(x_51, x_99); -x_101 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_101, 0, x_37); -lean_ctor_set(x_101, 1, x_53); -lean_ctor_set(x_101, 2, x_100); -x_102 = lean_array_push(x_51, x_101); -x_103 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__2; -x_104 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_104, 0, x_37); -lean_ctor_set(x_104, 1, x_103); -lean_ctor_set(x_104, 2, x_102); -lean_ctor_set(x_13, 0, x_104); -return x_13; +lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; +x_115 = lean_ctor_get(x_107, 0); +lean_inc(x_115); +lean_dec(x_107); +x_116 = lean_ctor_get(x_108, 0); +lean_inc(x_116); +x_117 = lean_ctor_get(x_108, 1); +lean_inc(x_117); +lean_dec(x_108); +x_118 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_118, 0, x_115); +x_10 = x_118; +x_11 = x_116; +x_12 = x_117; +goto block_104; } } -else +block_104: { -lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; -x_105 = lean_ctor_get(x_13, 0); -x_106 = lean_ctor_get(x_13, 1); -lean_inc(x_106); -lean_inc(x_105); -lean_dec(x_13); -x_107 = lean_ctor_get(x_2, 2); -lean_inc(x_107); -x_108 = lean_ctor_get(x_2, 1); -lean_inc(x_108); +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; 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; +x_13 = lean_ctor_get(x_2, 2); +lean_inc(x_13); +x_14 = lean_ctor_get(x_2, 1); +lean_inc(x_14); lean_dec(x_2); -x_109 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__5; -lean_inc(x_105); -x_110 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_110, 0, x_105); -lean_ctor_set(x_110, 1, x_109); -x_111 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__4; -lean_inc(x_105); -x_112 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_112, 0, x_105); -lean_ctor_set(x_112, 1, x_111); -x_113 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__6; -lean_inc(x_105); -x_114 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_114, 0, x_105); -lean_ctor_set(x_114, 1, x_113); -x_115 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__8; -lean_inc(x_105); -x_116 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_116, 0, x_105); -lean_ctor_set(x_116, 1, x_115); -x_117 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__11; -lean_inc(x_105); -x_118 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_118, 0, x_105); -lean_ctor_set(x_118, 1, x_117); -x_119 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__20; -lean_inc(x_107); -lean_inc(x_108); -x_120 = l_Lean_addMacroScope(x_108, x_119, x_107); -x_121 = lean_box(0); -x_122 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__19; -lean_inc(x_105); -x_123 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_123, 0, x_105); -lean_ctor_set(x_123, 1, x_122); -lean_ctor_set(x_123, 2, x_120); -lean_ctor_set(x_123, 3, x_121); -x_124 = l_Lean_Syntax_mkApp___closed__3; -x_125 = lean_array_push(x_124, x_123); -x_126 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__12; -x_127 = lean_array_push(x_125, x_126); -x_128 = lean_box(2); -x_129 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__16; -x_130 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_130, 0, x_128); -lean_ctor_set(x_130, 1, x_129); -lean_ctor_set(x_130, 2, x_127); -x_131 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__28; -x_132 = l_Lean_addMacroScope(x_108, x_131, x_107); -x_133 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__22; -x_134 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__30; -lean_inc(x_105); -x_135 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_135, 0, x_105); -lean_ctor_set(x_135, 1, x_133); -lean_ctor_set(x_135, 2, x_132); -lean_ctor_set(x_135, 3, x_134); -x_136 = l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__4; -x_137 = lean_array_push(x_136, x_130); -lean_inc(x_116); -x_138 = lean_array_push(x_137, x_116); -x_139 = lean_array_push(x_138, x_135); -x_140 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__14; -x_141 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_141, 0, x_128); -lean_ctor_set(x_141, 1, x_140); -lean_ctor_set(x_141, 2, x_139); -x_142 = l_Lean_mkOptionalNode___closed__2; -x_143 = lean_array_push(x_142, x_141); -x_144 = l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__2; -x_145 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_145, 0, x_128); -lean_ctor_set(x_145, 1, x_144); -lean_ctor_set(x_145, 2, x_143); -x_146 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__35; -lean_inc(x_105); -x_147 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_147, 0, x_105); -lean_ctor_set(x_147, 1, x_146); -x_148 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__36; -x_149 = lean_array_push(x_148, x_118); -x_150 = lean_array_push(x_149, x_126); -x_151 = lean_array_push(x_150, x_145); -x_152 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__34; -x_153 = lean_array_push(x_151, x_152); -x_154 = lean_array_push(x_153, x_126); -x_155 = lean_array_push(x_154, x_147); -x_156 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__10; -x_157 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_157, 0, x_128); -lean_ctor_set(x_157, 1, x_156); -lean_ctor_set(x_157, 2, x_155); -x_158 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__8; -x_159 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_159, 0, x_105); -lean_ctor_set(x_159, 1, x_158); -x_160 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__37; -x_161 = lean_array_push(x_160, x_112); -x_162 = lean_array_push(x_161, x_114); -x_163 = lean_array_push(x_162, x_116); -x_164 = lean_array_push(x_163, x_157); -x_165 = lean_array_push(x_164, x_159); -x_166 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__7; -x_167 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_167, 0, x_128); -lean_ctor_set(x_167, 1, x_166); -lean_ctor_set(x_167, 2, x_165); -x_168 = lean_array_push(x_142, x_167); -x_169 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_169, 0, x_128); -lean_ctor_set(x_169, 1, x_144); -lean_ctor_set(x_169, 2, x_168); -x_170 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__38; -x_171 = lean_array_push(x_170, x_110); -x_172 = lean_array_push(x_171, x_169); -x_173 = lean_array_push(x_172, x_9); -if (lean_obj_tag(x_12) == 0) +x_15 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__5; +lean_inc(x_11); +x_16 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_16, 0, x_11); +lean_ctor_set(x_16, 1, x_15); +x_17 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__4; +lean_inc(x_11); +x_18 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_18, 0, x_11); +lean_ctor_set(x_18, 1, x_17); +x_19 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__6; +lean_inc(x_11); +x_20 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_20, 0, x_11); +lean_ctor_set(x_20, 1, x_19); +x_21 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__8; +lean_inc(x_11); +x_22 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_22, 0, x_11); +lean_ctor_set(x_22, 1, x_21); +x_23 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__11; +lean_inc(x_11); +x_24 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_24, 0, x_11); +lean_ctor_set(x_24, 1, x_23); +x_25 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__20; +lean_inc(x_13); +lean_inc(x_14); +x_26 = l_Lean_addMacroScope(x_14, x_25, x_13); +x_27 = lean_box(0); +x_28 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__19; +lean_inc(x_11); +x_29 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_29, 0, x_11); +lean_ctor_set(x_29, 1, x_28); +lean_ctor_set(x_29, 2, x_26); +lean_ctor_set(x_29, 3, x_27); +x_30 = l_Lean_Syntax_mkApp___closed__3; +x_31 = lean_array_push(x_30, x_29); +x_32 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__12; +x_33 = lean_array_push(x_31, x_32); +x_34 = lean_box(2); +x_35 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__16; +x_36 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +lean_ctor_set(x_36, 2, x_33); +x_37 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__28; +x_38 = l_Lean_addMacroScope(x_14, x_37, x_13); +x_39 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__22; +x_40 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__30; +lean_inc(x_11); +x_41 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_41, 0, x_11); +lean_ctor_set(x_41, 1, x_39); +lean_ctor_set(x_41, 2, x_38); +lean_ctor_set(x_41, 3, x_40); +x_42 = l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__4; +x_43 = lean_array_push(x_42, x_36); +lean_inc(x_22); +x_44 = lean_array_push(x_43, x_22); +x_45 = lean_array_push(x_44, x_41); +x_46 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__14; +x_47 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_47, 0, x_34); +lean_ctor_set(x_47, 1, x_46); +lean_ctor_set(x_47, 2, x_45); +x_48 = l_Lean_mkOptionalNode___closed__2; +x_49 = lean_array_push(x_48, x_47); +x_50 = l_Lean_mkNullNode___closed__2; +x_51 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_51, 0, x_34); +lean_ctor_set(x_51, 1, x_50); +lean_ctor_set(x_51, 2, x_49); +x_52 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__35; +lean_inc(x_11); +x_53 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_53, 0, x_11); +lean_ctor_set(x_53, 1, x_52); +x_54 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__36; +x_55 = lean_array_push(x_54, x_24); +x_56 = lean_array_push(x_55, x_32); +x_57 = lean_array_push(x_56, x_51); +x_58 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__34; +x_59 = lean_array_push(x_57, x_58); +x_60 = lean_array_push(x_59, x_32); +x_61 = lean_array_push(x_60, x_53); +x_62 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__10; +x_63 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_63, 0, x_34); +lean_ctor_set(x_63, 1, x_62); +lean_ctor_set(x_63, 2, x_61); +x_64 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__8; +x_65 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_65, 0, x_11); +lean_ctor_set(x_65, 1, x_64); +x_66 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__37; +x_67 = lean_array_push(x_66, x_18); +x_68 = lean_array_push(x_67, x_20); +x_69 = lean_array_push(x_68, x_22); +x_70 = lean_array_push(x_69, x_63); +x_71 = lean_array_push(x_70, x_65); +x_72 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__7; +x_73 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_73, 0, x_34); +lean_ctor_set(x_73, 1, x_72); +lean_ctor_set(x_73, 2, x_71); +x_74 = lean_array_push(x_48, x_73); +x_75 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_75, 0, x_34); +lean_ctor_set(x_75, 1, x_50); +lean_ctor_set(x_75, 2, x_74); +x_76 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__38; +x_77 = lean_array_push(x_76, x_16); +x_78 = lean_array_push(x_77, x_75); +x_79 = lean_array_push(x_78, x_9); +if (lean_obj_tag(x_10) == 0) { -lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; -x_174 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__40; -x_175 = lean_array_push(x_173, x_174); -x_176 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__4; -x_177 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_177, 0, x_128); -lean_ctor_set(x_177, 1, x_176); -lean_ctor_set(x_177, 2, x_175); -x_178 = lean_array_push(x_142, x_177); -x_179 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_179, 0, x_128); -lean_ctor_set(x_179, 1, x_144); -lean_ctor_set(x_179, 2, x_178); -x_180 = lean_array_push(x_142, x_179); -x_181 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__2; -x_182 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_182, 0, x_128); -lean_ctor_set(x_182, 1, x_181); -lean_ctor_set(x_182, 2, x_180); -x_183 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_183, 0, x_182); -lean_ctor_set(x_183, 1, x_106); -return x_183; +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_80 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__40; +x_81 = lean_array_push(x_79, x_80); +x_82 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__4; +x_83 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_83, 0, x_34); +lean_ctor_set(x_83, 1, x_82); +lean_ctor_set(x_83, 2, x_81); +x_84 = lean_array_push(x_48, x_83); +x_85 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_85, 0, x_34); +lean_ctor_set(x_85, 1, x_50); +lean_ctor_set(x_85, 2, x_84); +x_86 = lean_array_push(x_48, x_85); +x_87 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__2; +x_88 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_88, 0, x_34); +lean_ctor_set(x_88, 1, x_87); +lean_ctor_set(x_88, 2, x_86); +x_89 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_89, 0, x_88); +lean_ctor_set(x_89, 1, x_12); +return x_89; } else { -lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; -x_184 = lean_ctor_get(x_12, 0); -lean_inc(x_184); -lean_dec(x_12); -x_185 = lean_array_push(x_142, x_184); -x_186 = l_Lean_mkSepArray___closed__1; -x_187 = l_Array_append___rarg(x_186, x_185); -x_188 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_188, 0, x_128); -lean_ctor_set(x_188, 1, x_144); -lean_ctor_set(x_188, 2, x_187); -x_189 = lean_array_push(x_173, x_188); -x_190 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__4; -x_191 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_191, 0, x_128); -lean_ctor_set(x_191, 1, x_190); -lean_ctor_set(x_191, 2, x_189); -x_192 = lean_array_push(x_142, x_191); -x_193 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_193, 0, x_128); -lean_ctor_set(x_193, 1, x_144); -lean_ctor_set(x_193, 2, x_192); -x_194 = lean_array_push(x_142, x_193); -x_195 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__2; -x_196 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_196, 0, x_128); -lean_ctor_set(x_196, 1, x_195); -lean_ctor_set(x_196, 2, x_194); -x_197 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_197, 0, x_196); -lean_ctor_set(x_197, 1, x_106); -return x_197; +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; +x_90 = lean_ctor_get(x_10, 0); +lean_inc(x_90); +lean_dec(x_10); +x_91 = lean_array_push(x_48, x_90); +x_92 = l_Lean_mkSepArray___closed__1; +x_93 = l_Array_append___rarg(x_92, x_91); +x_94 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_94, 0, x_34); +lean_ctor_set(x_94, 1, x_50); +lean_ctor_set(x_94, 2, x_93); +x_95 = lean_array_push(x_79, x_94); +x_96 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__4; +x_97 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_97, 0, x_34); +lean_ctor_set(x_97, 1, x_96); +lean_ctor_set(x_97, 2, x_95); +x_98 = lean_array_push(x_48, x_97); +x_99 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_99, 0, x_34); +lean_ctor_set(x_99, 1, x_50); +lean_ctor_set(x_99, 2, x_98); +x_100 = lean_array_push(x_48, x_99); +x_101 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__2; +x_102 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_102, 0, x_34); +lean_ctor_set(x_102, 1, x_101); +lean_ctor_set(x_102, 2, x_100); +x_103 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_103, 0, x_102); +lean_ctor_set(x_103, 1, x_12); +return x_103; } } } @@ -20825,7 +21524,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_simpAllKind___closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__4; +x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__4; x_2 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; @@ -20849,7 +21548,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_simpAllKind___closed__12() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_instQuoteBool___closed__6; +x_1 = l_Lean_instQuoteBoolMkStrAnonymous___closed__6; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -21045,25 +21744,25 @@ x_1 = l_Lean_Parser_Tactic_dsimpKind___closed__10; return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__1() { +static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("commandDeclare_simp_like_tactic______", 37); +x_1 = lean_mk_string_from_bytes("commandDeclare_simp_like_tactic_____", 36); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__2() { +static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_tacticErw_______closed__2; -x_2 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__1; +x_2 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__3() { +static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__3() { _start: { lean_object* x_1; @@ -21071,17 +21770,17 @@ x_1 = lean_mk_string_from_bytes("declare_simp_like_tactic", 24); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__4() { +static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__3; +x_1 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__3; x_2 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__5() { _start: { lean_object* x_1; @@ -21089,21 +21788,21 @@ x_1 = lean_mk_string_from_bytes("orelse", 6); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__5; +x_2 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__7() { +static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__6; +x_1 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__6; x_2 = l_Lean_Parser_Tactic_simpAllKind; x_3 = l_Lean_Parser_Tactic_dsimpKind; x_4 = lean_alloc_ctor(2, 3, 0); @@ -21113,25 +21812,25 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__8() { +static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_tacticErw_______closed__9; -x_2 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__7; +x_2 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__7; x_3 = lean_alloc_ctor(1, 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_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__9() { +static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_termEval__prec_____closed__4; -x_2 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__4; -x_3 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__8; +x_2 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__4; +x_3 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__8; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -21139,7 +21838,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__10() { +static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__10() { _start: { lean_object* x_1; @@ -21147,33 +21846,33 @@ x_1 = lean_mk_string_from_bytes("ident", 5); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__11() { +static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__10; +x_2 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__10; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__12() { +static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__11; +x_1 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__11; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__13() { +static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_termEval__prec_____closed__4; -x_2 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__9; -x_3 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__12; +x_2 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__9; +x_3 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__12; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -21181,41 +21880,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__14() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("str", 3); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__15() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__14; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__16() { +static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__14() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__15; +x_1 = l_Lean_Syntax_mkStrLit___closed__2; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__17() { +static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_termEval__prec_____closed__4; -x_2 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__13; -x_3 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__16; +x_2 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__13; +x_3 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__14; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -21223,7 +21904,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__18() { +static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__16() { _start: { lean_object* x_1; @@ -21231,21 +21912,21 @@ x_1 = lean_mk_string_from_bytes("term", 4); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__19() { +static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__18; +x_2 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__16; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__20() { +static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__19; +x_1 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__17; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -21253,13 +21934,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__21() { +static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_termEval__prec_____closed__4; -x_2 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__17; -x_3 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__20; +x_2 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__15; +x_3 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__18; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -21267,13 +21948,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__22() { +static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__2; +x_1 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__21; +x_3 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__19; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -21281,23 +21962,15 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic____________() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__22; -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__1() { +static lean_object* _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic__________() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Command", 7); +x_1 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__20; return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__2() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__1() { _start: { lean_object* x_1; @@ -21305,7 +21978,7 @@ x_1 = lean_mk_string_from_bytes("declaration", 11); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__3() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__2() { _start: { lean_object* x_1; @@ -21313,7 +21986,7 @@ x_1 = lean_mk_string_from_bytes("declModifiers", 13); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__4() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__3() { _start: { lean_object* x_1; @@ -21321,7 +21994,7 @@ x_1 = lean_mk_string_from_bytes("attributes", 10); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__4() { _start: { lean_object* x_1; @@ -21329,7 +22002,7 @@ x_1 = lean_mk_string_from_bytes("@[", 2); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__5() { _start: { lean_object* x_1; @@ -21337,7 +22010,7 @@ x_1 = lean_mk_string_from_bytes("attrInstance", 12); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__7() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__6() { _start: { lean_object* x_1; @@ -21345,7 +22018,7 @@ x_1 = lean_mk_string_from_bytes("attrKind", 8); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__8() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__7() { _start: { lean_object* x_1; @@ -21353,7 +22026,7 @@ x_1 = lean_mk_string_from_bytes("Attr", 4); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__9() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__8() { _start: { lean_object* x_1; @@ -21361,7 +22034,7 @@ x_1 = lean_mk_string_from_bytes("macro", 5); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__10() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -21371,7 +22044,7 @@ x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__11() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__10() { _start: { lean_object* x_1; @@ -21379,15 +22052,7 @@ x_1 = lean_mk_string_from_bytes("def", 3); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__12() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("declId", 6); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__13() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__11() { _start: { lean_object* x_1; @@ -21395,22 +22060,22 @@ x_1 = lean_mk_string_from_bytes("expandSimp", 10); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__14() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__13; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__11; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__15() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__13; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__11; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__14; +x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__12; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -21418,17 +22083,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__16() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__13; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__11; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__17() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__15() { _start: { lean_object* x_1; @@ -21436,7 +22101,7 @@ x_1 = lean_mk_string_from_bytes("optDeclSig", 10); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__18() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__16() { _start: { lean_object* x_1; @@ -21444,7 +22109,7 @@ x_1 = lean_mk_string_from_bytes("typeSpec", 8); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__19() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__17() { _start: { lean_object* x_1; @@ -21452,22 +22117,22 @@ x_1 = lean_mk_string_from_bytes("Macro", 5); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__20() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__18() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__19; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__17; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__21() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__19; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__17; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__20; +x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__18; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -21475,17 +22140,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__22() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__19; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__17; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__23() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -21495,7 +22160,7 @@ x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__24() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__22() { _start: { lean_object* x_1; @@ -21503,7 +22168,7 @@ x_1 = lean_mk_string_from_bytes("declValSimple", 13); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__25() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__23() { _start: { lean_object* x_1; @@ -21511,7 +22176,7 @@ x_1 = lean_mk_string_from_bytes("fun", 3); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__26() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__24() { _start: { lean_object* x_1; @@ -21519,7 +22184,7 @@ x_1 = lean_mk_string_from_bytes("basicFun", 8); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__27() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__25() { _start: { lean_object* x_1; @@ -21527,22 +22192,22 @@ x_1 = lean_mk_string_from_bytes("s", 1); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__28() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__26() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__27; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__25; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__29() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__27() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__27; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__25; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__28; +x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__26; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -21550,17 +22215,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__30() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__27; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__25; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__31() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__29() { _start: { lean_object* x_1; @@ -21568,7 +22233,7 @@ x_1 = lean_mk_string_from_bytes("=>", 2); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__32() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__30() { _start: { lean_object* x_1; @@ -21576,7 +22241,7 @@ x_1 = lean_mk_string_from_bytes("do", 2); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__33() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__31() { _start: { lean_object* x_1; @@ -21584,7 +22249,7 @@ x_1 = lean_mk_string_from_bytes("doSeqIndent", 11); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__34() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__32() { _start: { lean_object* x_1; @@ -21592,7 +22257,7 @@ x_1 = lean_mk_string_from_bytes("doSeqItem", 9); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__35() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__33() { _start: { lean_object* x_1; @@ -21600,7 +22265,7 @@ x_1 = lean_mk_string_from_bytes("doLetArrow", 10); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__36() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__34() { _start: { lean_object* x_1; @@ -21608,7 +22273,7 @@ x_1 = lean_mk_string_from_bytes("let", 3); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__37() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__35() { _start: { lean_object* x_1; @@ -21616,7 +22281,7 @@ x_1 = lean_mk_string_from_bytes("doIdDecl", 8); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__38() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__36() { _start: { lean_object* x_1; @@ -21624,22 +22289,22 @@ x_1 = lean_mk_string_from_bytes("c", 1); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__39() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__37() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__38; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__36; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__40() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__38() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__38; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__36; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__39; +x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__37; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -21647,17 +22312,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__41() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__39() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__38; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__36; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__42() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__40() { _start: { lean_object* x_1; @@ -21665,7 +22330,7 @@ x_1 = lean_mk_string_from_bytes("←", 3); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__43() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__41() { _start: { lean_object* x_1; @@ -21673,7 +22338,7 @@ x_1 = lean_mk_string_from_bytes("doMatch", 7); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__44() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__42() { _start: { lean_object* x_1; @@ -21681,7 +22346,7 @@ x_1 = lean_mk_string_from_bytes("match", 5); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__45() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__43() { _start: { lean_object* x_1; @@ -21689,7 +22354,7 @@ x_1 = lean_mk_string_from_bytes("matchDiscr", 10); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__46() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__44() { _start: { lean_object* x_1; @@ -21697,7 +22362,7 @@ x_1 = lean_mk_string_from_bytes("arrayRef", 8); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__47() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__45() { _start: { lean_object* x_1; @@ -21705,7 +22370,7 @@ x_1 = lean_mk_string_from_bytes("1", 1); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__48() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__46() { _start: { lean_object* x_1; @@ -21713,7 +22378,7 @@ x_1 = lean_mk_string_from_bytes("0", 1); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__49() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__47() { _start: { lean_object* x_1; @@ -21721,7 +22386,7 @@ x_1 = lean_mk_string_from_bytes("with", 4); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__50() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__48() { _start: { lean_object* x_1; @@ -21729,7 +22394,7 @@ x_1 = lean_mk_string_from_bytes("matchAlts", 9); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__51() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__49() { _start: { lean_object* x_1; @@ -21737,7 +22402,7 @@ x_1 = lean_mk_string_from_bytes("matchAlt", 8); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__52() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__50() { _start: { lean_object* x_1; @@ -21745,7 +22410,7 @@ x_1 = lean_mk_string_from_bytes("|", 1); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__53() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__51() { _start: { lean_object* x_1; @@ -21753,7 +22418,7 @@ x_1 = lean_mk_string_from_bytes("dynamicQuot", 11); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__54() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__52() { _start: { lean_object* x_1; @@ -21761,7 +22426,7 @@ x_1 = lean_mk_string_from_bytes("`(", 2); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__55() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__53() { _start: { lean_object* x_1; lean_object* x_2; @@ -21770,13 +22435,13 @@ x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__56() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__54() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__6; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__55; +x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__53; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -21784,7 +22449,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__57() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__55() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -21794,25 +22459,23 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__58() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__56() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("antiquot", 8); +x_1 = lean_mk_string_from_bytes("pseudo", 6); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__59() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__57() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__58; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("antiquot", 8); +return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__60() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__58() { _start: { lean_object* x_1; @@ -21820,7 +22483,7 @@ x_1 = lean_mk_string_from_bytes("$", 1); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__61() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__59() { _start: { lean_object* x_1; @@ -21828,17 +22491,17 @@ x_1 = lean_mk_string_from_bytes("antiquotName", 12); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__62() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__60() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__61; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__59; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__63() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__61() { _start: { lean_object* x_1; @@ -21846,15 +22509,7 @@ x_1 = lean_mk_string_from_bytes("doExpr", 6); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__64() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("namedArgument", 13); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__65() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__62() { _start: { lean_object* x_1; @@ -21862,17 +22517,17 @@ x_1 = lean_mk_string_from_bytes("choice", 6); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__66() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__63() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__65; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__62; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__67() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__64() { _start: { lean_object* x_1; @@ -21880,17 +22535,17 @@ x_1 = lean_mk_string_from_bytes("term{}", 6); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__68() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__65() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__67; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__64; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__69() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__66() { _start: { lean_object* x_1; @@ -21898,7 +22553,7 @@ x_1 = lean_mk_string_from_bytes("doLet", 5); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__70() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__67() { _start: { lean_object* x_1; @@ -21906,7 +22561,7 @@ x_1 = lean_mk_string_from_bytes("letDecl", 7); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__71() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__68() { _start: { lean_object* x_1; @@ -21914,7 +22569,7 @@ x_1 = lean_mk_string_from_bytes("letIdDecl", 9); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__72() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__69() { _start: { lean_object* x_1; @@ -21922,22 +22577,22 @@ x_1 = lean_mk_string_from_bytes("s.setKind", 9); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__73() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__70() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__72; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__69; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__74() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__71() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__72; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__69; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__73; +x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__70; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -21945,7 +22600,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__75() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__72() { _start: { lean_object* x_1; @@ -21953,17 +22608,17 @@ x_1 = lean_mk_string_from_bytes("setKind", 7); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__76() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__73() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__30; -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__75; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__28; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__72; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__77() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__74() { _start: { lean_object* x_1; @@ -21971,22 +22626,22 @@ x_1 = lean_mk_string_from_bytes("s.setArg", 8); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__78() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__75() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__77; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__74; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__79() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__76() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__77; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__74; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__78; +x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__75; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -21994,7 +22649,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__80() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__77() { _start: { lean_object* x_1; @@ -22002,17 +22657,17 @@ x_1 = lean_mk_string_from_bytes("setArg", 6); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__81() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__78() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__30; -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__80; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__28; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__77; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__82() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__79() { _start: { lean_object* x_1; @@ -22020,22 +22675,22 @@ x_1 = lean_mk_string_from_bytes("mkAtomFrom", 10); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__83() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__80() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__82; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__79; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__84() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__81() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__82; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__79; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__83; +x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__80; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -22043,17 +22698,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__85() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__82() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__82; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__79; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__86() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__83() { _start: { lean_object* x_1; @@ -22061,22 +22716,22 @@ x_1 = lean_mk_string_from_bytes("r", 1); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__87() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__84() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__86; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__83; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__88() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__85() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__86; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__83; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__87; +x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__84; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -22084,17 +22739,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__89() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__86() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__86; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__83; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__90() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__87() { _start: { lean_object* x_1; @@ -22102,22 +22757,22 @@ x_1 = lean_mk_string_from_bytes("mkNullNode", 10); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__91() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__88() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__90; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__87; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__92() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__89() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__90; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__87; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__91; +x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__88; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -22125,17 +22780,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__93() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__90() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__90; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__87; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__94() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__91() { _start: { lean_object* x_1; @@ -22143,17 +22798,17 @@ x_1 = lean_mk_string_from_bytes("term#[_,]", 9); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__95() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__92() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__94; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__91; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__96() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__93() { _start: { lean_object* x_1; @@ -22161,7 +22816,7 @@ x_1 = lean_mk_string_from_bytes("doReturn", 8); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__97() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__94() { _start: { lean_object* x_1; @@ -22169,7 +22824,7 @@ x_1 = lean_mk_string_from_bytes("return", 6); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__98() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__95() { _start: { lean_object* x_1; lean_object* x_2; @@ -22178,1141 +22833,1140 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________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, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________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, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -uint8_t x_9; -x_9 = !lean_is_exclusive(x_6); -if (x_9 == 0) +uint8_t x_10; +x_10 = !lean_is_exclusive(x_7); +if (x_10 == 0) { -lean_object* x_10; uint8_t x_11; -x_10 = lean_ctor_get(x_6, 1); -x_11 = !lean_is_exclusive(x_10); -if (x_11 == 0) +lean_object* x_11; uint8_t x_12; +x_11 = lean_ctor_get(x_7, 1); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_12 = lean_ctor_get(x_6, 0); -x_13 = lean_ctor_get(x_10, 0); -x_14 = lean_ctor_get(x_10, 1); -lean_inc(x_7); -x_15 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_7, x_8); -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_13 = lean_ctor_get(x_7, 0); +x_14 = lean_ctor_get(x_11, 0); +x_15 = lean_ctor_get(x_11, 1); +lean_inc(x_8); +x_16 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_8, x_9); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) { -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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; -x_17 = lean_ctor_get(x_15, 0); -x_18 = lean_ctor_get(x_7, 2); -lean_inc(x_18); -x_19 = lean_ctor_get(x_7, 1); +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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; +x_18 = lean_ctor_get(x_16, 0); +x_19 = lean_ctor_get(x_8, 2); lean_inc(x_19); -lean_dec(x_7); -x_20 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__1; +x_20 = lean_ctor_get(x_8, 1); +lean_inc(x_20); +lean_dec(x_8); +x_21 = l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___closed__1; lean_inc(x_1); -x_21 = lean_name_mk_string(x_1, x_20); -x_22 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__2; -lean_inc(x_21); -x_23 = lean_name_mk_string(x_21, x_22); -x_24 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__3; -lean_inc(x_21); -x_25 = lean_name_mk_string(x_21, x_24); -x_26 = l_Lean_mkHole___closed__5; +x_22 = lean_name_mk_string(x_1, x_21); +x_23 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__1; +lean_inc(x_22); +x_24 = lean_name_mk_string(x_22, x_23); +x_25 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__2; +lean_inc(x_22); +x_26 = lean_name_mk_string(x_22, x_25); +x_27 = l_Lean_mkHole___closed__5; lean_inc(x_1); -x_27 = lean_name_mk_string(x_1, x_26); -x_28 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__4; -lean_inc(x_27); -x_29 = lean_name_mk_string(x_27, x_28); -x_30 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__5; -lean_inc(x_17); -x_31 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_31, 0, x_17); -lean_ctor_set(x_31, 1, x_30); -x_32 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__6; -lean_inc(x_27); -x_33 = lean_name_mk_string(x_27, x_32); -x_34 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__7; -lean_inc(x_27); -x_35 = lean_name_mk_string(x_27, x_34); -x_36 = lean_box(2); -x_37 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__33; -x_38 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_35); -lean_ctor_set(x_38, 2, x_37); -x_39 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__8; -x_40 = lean_name_mk_string(x_1, x_39); -x_41 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__9; -x_42 = lean_name_mk_string(x_40, x_41); -lean_inc(x_17); -x_43 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_43, 0, x_17); -lean_ctor_set(x_43, 1, x_41); -x_44 = l_Lean_Syntax_mkApp___closed__3; -x_45 = lean_array_push(x_44, x_43); -x_46 = lean_array_push(x_45, x_2); -x_47 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_47, 0, x_36); -lean_ctor_set(x_47, 1, x_42); -lean_ctor_set(x_47, 2, x_46); -x_48 = lean_array_push(x_44, x_38); -x_49 = lean_array_push(x_48, x_47); -x_50 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_50, 0, x_36); -lean_ctor_set(x_50, 1, x_33); -lean_ctor_set(x_50, 2, x_49); -x_51 = l_Lean_mkOptionalNode___closed__2; -x_52 = lean_array_push(x_51, x_50); -x_53 = l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__2; -x_54 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_54, 0, x_36); -lean_ctor_set(x_54, 1, x_53); -lean_ctor_set(x_54, 2, x_52); -x_55 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__16; -lean_inc(x_17); -x_56 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_56, 0, x_17); -lean_ctor_set(x_56, 1, x_55); -x_57 = l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__4; -x_58 = lean_array_push(x_57, x_31); -x_59 = lean_array_push(x_58, x_54); -lean_inc(x_56); -x_60 = lean_array_push(x_59, x_56); -x_61 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_61, 0, x_36); -lean_ctor_set(x_61, 1, x_29); -lean_ctor_set(x_61, 2, x_60); -x_62 = lean_array_push(x_51, x_61); -x_63 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_63, 0, x_36); -lean_ctor_set(x_63, 1, x_53); -lean_ctor_set(x_63, 2, x_62); -x_64 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__10; -x_65 = lean_array_push(x_64, x_63); -x_66 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__12; -x_67 = lean_array_push(x_65, x_66); -x_68 = lean_array_push(x_67, x_66); -x_69 = lean_array_push(x_68, x_66); -x_70 = lean_array_push(x_69, x_66); -x_71 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_71, 0, x_36); -lean_ctor_set(x_71, 1, x_25); -lean_ctor_set(x_71, 2, x_70); -x_72 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__11; -lean_inc(x_21); -x_73 = lean_name_mk_string(x_21, x_72); -lean_inc(x_17); -x_74 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_74, 0, x_17); -lean_ctor_set(x_74, 1, x_72); -x_75 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__12; -lean_inc(x_21); -x_76 = lean_name_mk_string(x_21, x_75); -x_77 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__16; +x_28 = lean_name_mk_string(x_1, x_27); +x_29 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__3; +lean_inc(x_28); +x_30 = lean_name_mk_string(x_28, x_29); +x_31 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__4; lean_inc(x_18); +x_32 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_32, 0, x_18); +lean_ctor_set(x_32, 1, x_31); +x_33 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__5; +lean_inc(x_28); +x_34 = lean_name_mk_string(x_28, x_33); +x_35 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__6; +lean_inc(x_28); +x_36 = lean_name_mk_string(x_28, x_35); +x_37 = lean_box(2); +x_38 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__33; +x_39 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_36); +lean_ctor_set(x_39, 2, x_38); +x_40 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__7; +x_41 = lean_name_mk_string(x_1, x_40); +x_42 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__8; +x_43 = lean_name_mk_string(x_41, x_42); +lean_inc(x_18); +x_44 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_44, 0, x_18); +lean_ctor_set(x_44, 1, x_42); +x_45 = l_Lean_Syntax_mkApp___closed__3; +x_46 = lean_array_push(x_45, x_44); +x_47 = lean_array_push(x_46, x_2); +x_48 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_48, 0, x_37); +lean_ctor_set(x_48, 1, x_43); +lean_ctor_set(x_48, 2, x_47); +x_49 = lean_array_push(x_45, x_39); +x_50 = lean_array_push(x_49, x_48); +x_51 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_51, 0, x_37); +lean_ctor_set(x_51, 1, x_34); +lean_ctor_set(x_51, 2, x_50); +x_52 = l_Lean_mkOptionalNode___closed__2; +x_53 = lean_array_push(x_52, x_51); +x_54 = l_Lean_mkNullNode___closed__2; +x_55 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_55, 0, x_37); +lean_ctor_set(x_55, 1, x_54); +lean_ctor_set(x_55, 2, x_53); +x_56 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__16; +lean_inc(x_18); +x_57 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_57, 0, x_18); +lean_ctor_set(x_57, 1, x_56); +x_58 = l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__4; +x_59 = lean_array_push(x_58, x_32); +x_60 = lean_array_push(x_59, x_55); +lean_inc(x_57); +x_61 = lean_array_push(x_60, x_57); +x_62 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_62, 0, x_37); +lean_ctor_set(x_62, 1, x_30); +lean_ctor_set(x_62, 2, x_61); +x_63 = lean_array_push(x_52, x_62); +x_64 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_64, 0, x_37); +lean_ctor_set(x_64, 1, x_54); +lean_ctor_set(x_64, 2, x_63); +x_65 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__9; +x_66 = lean_array_push(x_65, x_64); +x_67 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__12; +x_68 = lean_array_push(x_66, x_67); +x_69 = lean_array_push(x_68, x_67); +x_70 = lean_array_push(x_69, x_67); +x_71 = lean_array_push(x_70, x_67); +x_72 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_72, 0, x_37); +lean_ctor_set(x_72, 1, x_26); +lean_ctor_set(x_72, 2, x_71); +x_73 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__10; +lean_inc(x_22); +x_74 = lean_name_mk_string(x_22, x_73); +lean_inc(x_18); +x_75 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_75, 0, x_18); +lean_ctor_set(x_75, 1, x_73); +x_76 = l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___closed__3; +lean_inc(x_22); +x_77 = lean_name_mk_string(x_22, x_76); +x_78 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__14; lean_inc(x_19); -x_78 = l_Lean_addMacroScope(x_19, x_77, x_18); -x_79 = lean_box(0); -x_80 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__15; -lean_inc(x_17); -x_81 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_81, 0, x_17); -lean_ctor_set(x_81, 1, x_80); -lean_ctor_set(x_81, 2, x_78); -lean_ctor_set(x_81, 3, x_79); -x_82 = lean_array_push(x_44, x_81); -x_83 = lean_array_push(x_82, x_66); -x_84 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_84, 0, x_36); -lean_ctor_set(x_84, 1, x_76); -lean_ctor_set(x_84, 2, x_83); -x_85 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__17; -lean_inc(x_21); -x_86 = lean_name_mk_string(x_21, x_85); -x_87 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__18; -lean_inc(x_27); -x_88 = lean_name_mk_string(x_27, x_87); -x_89 = l_Lean_toolchain___closed__1; -lean_inc(x_17); -x_90 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_90, 0, x_17); -lean_ctor_set(x_90, 1, x_89); -x_91 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__22; +lean_inc(x_20); +x_79 = l_Lean_addMacroScope(x_20, x_78, x_19); +x_80 = lean_box(0); +x_81 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__13; lean_inc(x_18); +x_82 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_82, 0, x_18); +lean_ctor_set(x_82, 1, x_81); +lean_ctor_set(x_82, 2, x_79); +lean_ctor_set(x_82, 3, x_80); +x_83 = lean_array_push(x_45, x_82); +x_84 = lean_array_push(x_83, x_67); +x_85 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_85, 0, x_37); +lean_ctor_set(x_85, 1, x_77); +lean_ctor_set(x_85, 2, x_84); +x_86 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__15; +lean_inc(x_22); +x_87 = lean_name_mk_string(x_22, x_86); +x_88 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__16; +lean_inc(x_28); +x_89 = lean_name_mk_string(x_28, x_88); +x_90 = l_Lean_toolchain___closed__1; +lean_inc(x_18); +x_91 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_91, 0, x_18); +lean_ctor_set(x_91, 1, x_90); +x_92 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__20; lean_inc(x_19); -x_92 = l_Lean_addMacroScope(x_19, x_91, x_18); -x_93 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__19; +lean_inc(x_20); +x_93 = l_Lean_addMacroScope(x_20, x_92, x_19); +x_94 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__17; lean_inc(x_3); -x_94 = lean_name_mk_string(x_3, x_93); -lean_ctor_set(x_10, 1, x_79); -lean_ctor_set(x_10, 0, x_94); -x_95 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_95, 0, x_10); -lean_ctor_set(x_95, 1, x_79); -x_96 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__21; -lean_inc(x_17); -x_97 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_97, 0, x_17); -lean_ctor_set(x_97, 1, x_96); -lean_ctor_set(x_97, 2, x_92); -lean_ctor_set(x_97, 3, x_95); -x_98 = lean_array_push(x_44, x_90); -lean_inc(x_98); -x_99 = lean_array_push(x_98, x_97); -x_100 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_100, 0, x_36); -lean_ctor_set(x_100, 1, x_88); -lean_ctor_set(x_100, 2, x_99); -x_101 = lean_array_push(x_51, x_100); -x_102 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_102, 0, x_36); -lean_ctor_set(x_102, 1, x_53); -lean_ctor_set(x_102, 2, x_101); -x_103 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__23; -x_104 = lean_array_push(x_103, x_102); -x_105 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_105, 0, x_36); -lean_ctor_set(x_105, 1, x_86); -lean_ctor_set(x_105, 2, x_104); -x_106 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__24; -x_107 = lean_name_mk_string(x_21, x_106); -x_108 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__8; -lean_inc(x_17); -x_109 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_109, 0, x_17); -lean_ctor_set(x_109, 1, x_108); -x_110 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__25; -lean_inc(x_27); -x_111 = lean_name_mk_string(x_27, x_110); -lean_inc(x_17); -x_112 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_112, 0, x_17); -lean_ctor_set(x_112, 1, x_110); -x_113 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__26; -lean_inc(x_27); -x_114 = lean_name_mk_string(x_27, x_113); -x_115 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__30; +x_95 = lean_name_mk_string(x_3, x_94); +lean_ctor_set(x_11, 1, x_80); +lean_ctor_set(x_11, 0, x_95); +x_96 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_96, 0, x_11); +lean_ctor_set(x_96, 1, x_80); +x_97 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__19; lean_inc(x_18); +x_98 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_98, 0, x_18); +lean_ctor_set(x_98, 1, x_97); +lean_ctor_set(x_98, 2, x_93); +lean_ctor_set(x_98, 3, x_96); +x_99 = lean_array_push(x_45, x_91); +lean_inc(x_99); +x_100 = lean_array_push(x_99, x_98); +x_101 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_101, 0, x_37); +lean_ctor_set(x_101, 1, x_89); +lean_ctor_set(x_101, 2, x_100); +x_102 = lean_array_push(x_52, x_101); +x_103 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_103, 0, x_37); +lean_ctor_set(x_103, 1, x_54); +lean_ctor_set(x_103, 2, x_102); +x_104 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__21; +x_105 = lean_array_push(x_104, x_103); +x_106 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_106, 0, x_37); +lean_ctor_set(x_106, 1, x_87); +lean_ctor_set(x_106, 2, x_105); +x_107 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__22; +x_108 = lean_name_mk_string(x_22, x_107); +x_109 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__8; +lean_inc(x_18); +x_110 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_110, 0, x_18); +lean_ctor_set(x_110, 1, x_109); +x_111 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__23; +lean_inc(x_28); +x_112 = lean_name_mk_string(x_28, x_111); +lean_inc(x_18); +x_113 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_113, 0, x_18); +lean_ctor_set(x_113, 1, x_111); +x_114 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__24; +lean_inc(x_28); +x_115 = lean_name_mk_string(x_28, x_114); +x_116 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__28; lean_inc(x_19); -x_116 = l_Lean_addMacroScope(x_19, x_115, x_18); -x_117 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__29; -lean_inc(x_17); -x_118 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_118, 0, x_17); -lean_ctor_set(x_118, 1, x_117); -lean_ctor_set(x_118, 2, x_116); -lean_ctor_set(x_118, 3, x_79); -lean_inc(x_118); -x_119 = lean_array_push(x_51, x_118); -x_120 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_120, 0, x_36); -lean_ctor_set(x_120, 1, x_53); -lean_ctor_set(x_120, 2, x_119); -x_121 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__31; -lean_inc(x_17); -x_122 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_122, 0, x_17); -lean_ctor_set(x_122, 1, x_121); -x_123 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__32; -lean_inc(x_27); -x_124 = lean_name_mk_string(x_27, x_123); -lean_inc(x_17); -x_125 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_125, 0, x_17); -lean_ctor_set(x_125, 1, x_123); -x_126 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__33; -lean_inc(x_27); -x_127 = lean_name_mk_string(x_27, x_126); -x_128 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__34; -lean_inc(x_27); -x_129 = lean_name_mk_string(x_27, x_128); -x_130 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__35; -lean_inc(x_27); -x_131 = lean_name_mk_string(x_27, x_130); -x_132 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__36; -lean_inc(x_17); -x_133 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_133, 0, x_17); -lean_ctor_set(x_133, 1, x_132); -x_134 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__37; -lean_inc(x_27); -x_135 = lean_name_mk_string(x_27, x_134); -x_136 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__41; +lean_inc(x_20); +x_117 = l_Lean_addMacroScope(x_20, x_116, x_19); +x_118 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__27; lean_inc(x_18); +x_119 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_119, 0, x_18); +lean_ctor_set(x_119, 1, x_118); +lean_ctor_set(x_119, 2, x_117); +lean_ctor_set(x_119, 3, x_80); +lean_inc(x_119); +x_120 = lean_array_push(x_52, x_119); +x_121 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_121, 0, x_37); +lean_ctor_set(x_121, 1, x_54); +lean_ctor_set(x_121, 2, x_120); +x_122 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__29; +lean_inc(x_18); +x_123 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_123, 0, x_18); +lean_ctor_set(x_123, 1, x_122); +x_124 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__30; +lean_inc(x_28); +x_125 = lean_name_mk_string(x_28, x_124); +lean_inc(x_18); +x_126 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_126, 0, x_18); +lean_ctor_set(x_126, 1, x_124); +x_127 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__31; +lean_inc(x_28); +x_128 = lean_name_mk_string(x_28, x_127); +x_129 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__32; +lean_inc(x_28); +x_130 = lean_name_mk_string(x_28, x_129); +x_131 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__33; +lean_inc(x_28); +x_132 = lean_name_mk_string(x_28, x_131); +x_133 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__34; +lean_inc(x_18); +x_134 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_134, 0, x_18); +lean_ctor_set(x_134, 1, x_133); +x_135 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__35; +lean_inc(x_28); +x_136 = lean_name_mk_string(x_28, x_135); +x_137 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__39; lean_inc(x_19); -x_137 = l_Lean_addMacroScope(x_19, x_136, x_18); -x_138 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__40; -lean_inc(x_17); -x_139 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_139, 0, x_17); -lean_ctor_set(x_139, 1, x_138); -lean_ctor_set(x_139, 2, x_137); -lean_ctor_set(x_139, 3, x_79); -x_140 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__42; -lean_inc(x_17); -x_141 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_141, 0, x_17); -lean_ctor_set(x_141, 1, x_140); -x_142 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__43; -lean_inc(x_27); -x_143 = lean_name_mk_string(x_27, x_142); -x_144 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__44; -lean_inc(x_17); -x_145 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_145, 0, x_17); -lean_ctor_set(x_145, 1, x_144); -x_146 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__45; -lean_inc(x_27); -x_147 = lean_name_mk_string(x_27, x_146); -x_148 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__46; -lean_inc(x_27); -x_149 = lean_name_mk_string(x_27, x_148); -x_150 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__12; -lean_inc(x_17); -x_151 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_151, 0, x_17); -lean_ctor_set(x_151, 1, x_150); -x_152 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__47; -lean_inc(x_17); -x_153 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_153, 0, x_17); -lean_ctor_set(x_153, 1, x_152); -x_154 = lean_array_push(x_51, x_153); -x_155 = l_Lean_evalPrec___closed__2; -x_156 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_156, 0, x_36); -lean_ctor_set(x_156, 1, x_155); -lean_ctor_set(x_156, 2, x_154); -x_157 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__38; -lean_inc(x_118); -x_158 = lean_array_push(x_157, x_118); -lean_inc(x_151); -x_159 = lean_array_push(x_158, x_151); -lean_inc(x_156); -lean_inc(x_159); -x_160 = lean_array_push(x_159, x_156); -lean_inc(x_56); -x_161 = lean_array_push(x_160, x_56); -lean_inc(x_149); -x_162 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_162, 0, x_36); -lean_ctor_set(x_162, 1, x_149); -lean_ctor_set(x_162, 2, x_161); -x_163 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__48; -lean_inc(x_17); -x_164 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_164, 0, x_17); -lean_ctor_set(x_164, 1, x_163); -x_165 = lean_array_push(x_51, x_164); -x_166 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_166, 0, x_36); -lean_ctor_set(x_166, 1, x_155); -lean_ctor_set(x_166, 2, x_165); -x_167 = lean_array_push(x_157, x_162); -x_168 = lean_array_push(x_167, x_151); -lean_inc(x_166); -x_169 = lean_array_push(x_168, x_166); -lean_inc(x_56); -x_170 = lean_array_push(x_169, x_56); -lean_inc(x_149); -x_171 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_171, 0, x_36); -lean_ctor_set(x_171, 1, x_149); -lean_ctor_set(x_171, 2, x_170); -x_172 = lean_array_push(x_103, x_171); -x_173 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_173, 0, x_36); -lean_ctor_set(x_173, 1, x_147); -lean_ctor_set(x_173, 2, x_172); -x_174 = lean_array_push(x_51, x_173); -x_175 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_175, 0, x_36); -lean_ctor_set(x_175, 1, x_53); -lean_ctor_set(x_175, 2, x_174); -x_176 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__49; -lean_inc(x_17); -x_177 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_177, 0, x_17); -lean_ctor_set(x_177, 1, x_176); -x_178 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__50; -lean_inc(x_27); -x_179 = lean_name_mk_string(x_27, x_178); -x_180 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__51; -lean_inc(x_27); -x_181 = lean_name_mk_string(x_27, x_180); -x_182 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__52; -lean_inc(x_17); -x_183 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_183, 0, x_17); -lean_ctor_set(x_183, 1, x_182); -x_184 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__53; -lean_inc(x_27); -x_185 = lean_name_mk_string(x_27, x_184); -x_186 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__54; -lean_inc(x_17); -x_187 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_187, 0, x_17); -lean_ctor_set(x_187, 1, x_186); -x_188 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__57; +lean_inc(x_20); +x_138 = l_Lean_addMacroScope(x_20, x_137, x_19); +x_139 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__38; +lean_inc(x_18); +x_140 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_140, 0, x_18); +lean_ctor_set(x_140, 1, x_139); +lean_ctor_set(x_140, 2, x_138); +lean_ctor_set(x_140, 3, x_80); +x_141 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__40; +lean_inc(x_18); +x_142 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_142, 0, x_18); +lean_ctor_set(x_142, 1, x_141); +x_143 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__41; +lean_inc(x_28); +x_144 = lean_name_mk_string(x_28, x_143); +x_145 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__42; +lean_inc(x_18); +x_146 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_146, 0, x_18); +lean_ctor_set(x_146, 1, x_145); +x_147 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__43; +lean_inc(x_28); +x_148 = lean_name_mk_string(x_28, x_147); +x_149 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__44; +lean_inc(x_28); +x_150 = lean_name_mk_string(x_28, x_149); +x_151 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__12; lean_inc(x_18); +x_152 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_152, 0, x_18); +lean_ctor_set(x_152, 1, x_151); +x_153 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__45; +lean_inc(x_18); +x_154 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_154, 0, x_18); +lean_ctor_set(x_154, 1, x_153); +x_155 = lean_array_push(x_52, x_154); +x_156 = l_Lean_Syntax_mkNumLit___closed__2; +x_157 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_157, 0, x_37); +lean_ctor_set(x_157, 1, x_156); +lean_ctor_set(x_157, 2, x_155); +x_158 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__38; +lean_inc(x_119); +x_159 = lean_array_push(x_158, x_119); +lean_inc(x_152); +x_160 = lean_array_push(x_159, x_152); +lean_inc(x_157); +lean_inc(x_160); +x_161 = lean_array_push(x_160, x_157); +lean_inc(x_57); +x_162 = lean_array_push(x_161, x_57); +lean_inc(x_150); +x_163 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_163, 0, x_37); +lean_ctor_set(x_163, 1, x_150); +lean_ctor_set(x_163, 2, x_162); +x_164 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__46; +lean_inc(x_18); +x_165 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_165, 0, x_18); +lean_ctor_set(x_165, 1, x_164); +x_166 = lean_array_push(x_52, x_165); +x_167 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_167, 0, x_37); +lean_ctor_set(x_167, 1, x_156); +lean_ctor_set(x_167, 2, x_166); +x_168 = lean_array_push(x_158, x_163); +x_169 = lean_array_push(x_168, x_152); +lean_inc(x_167); +x_170 = lean_array_push(x_169, x_167); +lean_inc(x_57); +x_171 = lean_array_push(x_170, x_57); +lean_inc(x_150); +x_172 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_172, 0, x_37); +lean_ctor_set(x_172, 1, x_150); +lean_ctor_set(x_172, 2, x_171); +x_173 = lean_array_push(x_104, x_172); +x_174 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_174, 0, x_37); +lean_ctor_set(x_174, 1, x_148); +lean_ctor_set(x_174, 2, x_173); +x_175 = lean_array_push(x_52, x_174); +x_176 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_176, 0, x_37); +lean_ctor_set(x_176, 1, x_54); +lean_ctor_set(x_176, 2, x_175); +x_177 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__47; +lean_inc(x_18); +x_178 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_178, 0, x_18); +lean_ctor_set(x_178, 1, x_177); +x_179 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__48; +lean_inc(x_28); +x_180 = lean_name_mk_string(x_28, x_179); +x_181 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__49; +lean_inc(x_28); +x_182 = lean_name_mk_string(x_28, x_181); +x_183 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__50; +lean_inc(x_18); +x_184 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_184, 0, x_18); +lean_ctor_set(x_184, 1, x_183); +x_185 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__51; +lean_inc(x_28); +x_186 = lean_name_mk_string(x_28, x_185); +x_187 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__52; +lean_inc(x_18); +x_188 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_188, 0, x_18); +lean_ctor_set(x_188, 1, x_187); +x_189 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__55; lean_inc(x_19); -x_189 = l_Lean_addMacroScope(x_19, x_188, x_18); -x_190 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__6; -x_191 = lean_name_mk_string(x_4, x_190); -lean_inc(x_191); -lean_ctor_set(x_6, 1, x_79); -lean_ctor_set(x_6, 0, x_191); -x_192 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_192, 0, x_6); -lean_ctor_set(x_192, 1, x_79); -x_193 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__56; -lean_inc(x_17); -x_194 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_194, 0, x_17); -lean_ctor_set(x_194, 1, x_193); -lean_ctor_set(x_194, 2, x_189); -lean_ctor_set(x_194, 3, x_192); -x_195 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__4; -lean_inc(x_17); -x_196 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_196, 0, x_17); -lean_ctor_set(x_196, 1, x_195); -lean_inc(x_17); +lean_inc(x_20); +x_190 = l_Lean_addMacroScope(x_20, x_189, x_19); +x_191 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__6; +x_192 = lean_name_mk_string(x_4, x_191); +lean_inc(x_192); +lean_ctor_set(x_7, 1, x_80); +lean_ctor_set(x_7, 0, x_192); +x_193 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_193, 0, x_7); +lean_ctor_set(x_193, 1, x_80); +x_194 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__54; +lean_inc(x_18); +x_195 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_195, 0, x_18); +lean_ctor_set(x_195, 1, x_194); +lean_ctor_set(x_195, 2, x_190); +lean_ctor_set(x_195, 3, x_193); +x_196 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__4; +lean_inc(x_18); x_197 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_197, 0, x_17); -lean_ctor_set(x_197, 1, x_190); -x_198 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__60; -lean_inc(x_17); -x_199 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_199, 0, x_17); -lean_ctor_set(x_199, 1, x_198); -x_200 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__18; -lean_inc(x_17); -x_201 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_201, 0, x_17); -lean_ctor_set(x_201, 1, x_200); -x_202 = lean_array_push(x_98, x_201); -x_203 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__62; -x_204 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_204, 0, x_36); +lean_ctor_set(x_197, 0, x_18); +lean_ctor_set(x_197, 1, x_196); +lean_inc(x_18); +x_198 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_198, 0, x_18); +lean_ctor_set(x_198, 1, x_191); +x_199 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__56; +x_200 = lean_name_mk_string(x_5, x_199); +x_201 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__57; +x_202 = lean_name_mk_string(x_200, x_201); +x_203 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__58; +lean_inc(x_18); +x_204 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_204, 0, x_18); lean_ctor_set(x_204, 1, x_203); -lean_ctor_set(x_204, 2, x_202); -x_205 = lean_array_push(x_157, x_199); -x_206 = lean_array_push(x_205, x_66); -lean_inc(x_139); -x_207 = lean_array_push(x_206, x_139); -lean_inc(x_207); -x_208 = lean_array_push(x_207, x_204); -x_209 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__59; -x_210 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_210, 0, x_36); -lean_ctor_set(x_210, 1, x_209); -lean_ctor_set(x_210, 2, x_208); -x_211 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__8; -lean_inc(x_17); -x_212 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_212, 0, x_17); -lean_ctor_set(x_212, 1, x_211); -x_213 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__37; -lean_inc(x_196); -x_214 = lean_array_push(x_213, x_196); -x_215 = lean_array_push(x_214, x_197); -lean_inc(x_109); -x_216 = lean_array_push(x_215, x_109); -lean_inc(x_216); -x_217 = lean_array_push(x_216, x_210); -lean_inc(x_212); -x_218 = lean_array_push(x_217, x_212); -lean_inc(x_191); -x_219 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_219, 0, x_36); -lean_ctor_set(x_219, 1, x_191); -lean_ctor_set(x_219, 2, x_218); -x_220 = lean_array_push(x_213, x_187); -x_221 = lean_array_push(x_220, x_194); -lean_inc(x_183); -x_222 = lean_array_push(x_221, x_183); -lean_inc(x_222); -x_223 = lean_array_push(x_222, x_219); +x_205 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__16; +lean_inc(x_18); +x_206 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_206, 0, x_18); +lean_ctor_set(x_206, 1, x_205); +x_207 = lean_array_push(x_99, x_206); +x_208 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__60; +x_209 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_209, 0, x_37); +lean_ctor_set(x_209, 1, x_208); +lean_ctor_set(x_209, 2, x_207); +x_210 = lean_array_push(x_158, x_204); +x_211 = lean_array_push(x_210, x_67); +lean_inc(x_140); +x_212 = lean_array_push(x_211, x_140); lean_inc(x_212); -x_224 = lean_array_push(x_223, x_212); -lean_inc(x_185); -x_225 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_225, 0, x_36); -lean_ctor_set(x_225, 1, x_185); -lean_ctor_set(x_225, 2, x_224); -x_226 = lean_array_push(x_51, x_225); -x_227 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_227, 0, x_36); -lean_ctor_set(x_227, 1, x_53); -lean_ctor_set(x_227, 2, x_226); -x_228 = lean_array_push(x_51, x_227); +x_213 = lean_array_push(x_212, x_209); +lean_inc(x_202); +x_214 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_214, 0, x_37); +lean_ctor_set(x_214, 1, x_202); +lean_ctor_set(x_214, 2, x_213); +x_215 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__8; +lean_inc(x_18); +x_216 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_216, 0, x_18); +lean_ctor_set(x_216, 1, x_215); +x_217 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__37; +lean_inc(x_197); +x_218 = lean_array_push(x_217, x_197); +x_219 = lean_array_push(x_218, x_198); +lean_inc(x_110); +x_220 = lean_array_push(x_219, x_110); +lean_inc(x_220); +x_221 = lean_array_push(x_220, x_214); +lean_inc(x_216); +x_222 = lean_array_push(x_221, x_216); +lean_inc(x_192); +x_223 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_223, 0, x_37); +lean_ctor_set(x_223, 1, x_192); +lean_ctor_set(x_223, 2, x_222); +x_224 = lean_array_push(x_217, x_188); +x_225 = lean_array_push(x_224, x_195); +lean_inc(x_184); +x_226 = lean_array_push(x_225, x_184); +lean_inc(x_226); +x_227 = lean_array_push(x_226, x_223); +lean_inc(x_216); +x_228 = lean_array_push(x_227, x_216); +lean_inc(x_186); x_229 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_229, 0, x_36); -lean_ctor_set(x_229, 1, x_53); +lean_ctor_set(x_229, 0, x_37); +lean_ctor_set(x_229, 1, x_186); lean_ctor_set(x_229, 2, x_228); -x_230 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__63; -lean_inc(x_27); -x_231 = lean_name_mk_string(x_27, x_230); -x_232 = l_Lean_Syntax_mkApp___closed__1; -lean_inc(x_27); -x_233 = lean_name_mk_string(x_27, x_232); -x_234 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__64; -lean_inc(x_27); -x_235 = lean_name_mk_string(x_27, x_234); -x_236 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__58; -x_237 = lean_name_mk_string(x_235, x_236); -x_238 = lean_array_push(x_207, x_66); +x_230 = lean_array_push(x_52, x_229); +x_231 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_231, 0, x_37); +lean_ctor_set(x_231, 1, x_54); +lean_ctor_set(x_231, 2, x_230); +x_232 = lean_array_push(x_52, x_231); +x_233 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_233, 0, x_37); +lean_ctor_set(x_233, 1, x_54); +lean_ctor_set(x_233, 2, x_232); +x_234 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__61; +lean_inc(x_28); +x_235 = lean_name_mk_string(x_28, x_234); +x_236 = l_Lean_Syntax_mkApp___closed__1; +lean_inc(x_28); +x_237 = lean_name_mk_string(x_28, x_236); +x_238 = lean_array_push(x_212, x_67); x_239 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_239, 0, x_36); -lean_ctor_set(x_239, 1, x_237); +lean_ctor_set(x_239, 0, x_37); +lean_ctor_set(x_239, 1, x_202); lean_ctor_set(x_239, 2, x_238); -x_240 = lean_array_push(x_51, x_239); +x_240 = lean_array_push(x_52, x_239); x_241 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_241, 0, x_36); -lean_ctor_set(x_241, 1, x_53); +lean_ctor_set(x_241, 0, x_37); +lean_ctor_set(x_241, 1, x_54); lean_ctor_set(x_241, 2, x_240); -x_242 = lean_array_push(x_44, x_5); +x_242 = lean_array_push(x_45, x_6); lean_inc(x_242); x_243 = lean_array_push(x_242, x_241); -lean_inc(x_233); +lean_inc(x_237); x_244 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_244, 0, x_36); -lean_ctor_set(x_244, 1, x_233); +lean_ctor_set(x_244, 0, x_37); +lean_ctor_set(x_244, 1, x_237); lean_ctor_set(x_244, 2, x_243); +lean_inc(x_220); +x_245 = lean_array_push(x_220, x_244); lean_inc(x_216); -x_245 = lean_array_push(x_216, x_244); -lean_inc(x_212); -x_246 = lean_array_push(x_245, x_212); -lean_inc(x_191); +x_246 = lean_array_push(x_245, x_216); +lean_inc(x_192); x_247 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_247, 0, x_36); -lean_ctor_set(x_247, 1, x_191); +lean_ctor_set(x_247, 0, x_37); +lean_ctor_set(x_247, 1, x_192); lean_ctor_set(x_247, 2, x_246); -lean_inc(x_222); -x_248 = lean_array_push(x_222, x_247); -lean_inc(x_212); -x_249 = lean_array_push(x_248, x_212); -lean_inc(x_185); +lean_inc(x_226); +x_248 = lean_array_push(x_226, x_247); +lean_inc(x_216); +x_249 = lean_array_push(x_248, x_216); +lean_inc(x_186); x_250 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_250, 0, x_36); -lean_ctor_set(x_250, 1, x_185); +lean_ctor_set(x_250, 0, x_37); +lean_ctor_set(x_250, 1, x_186); lean_ctor_set(x_250, 2, x_249); -x_251 = lean_array_push(x_51, x_250); -lean_inc(x_231); +x_251 = lean_array_push(x_52, x_250); +lean_inc(x_235); x_252 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_252, 0, x_36); -lean_ctor_set(x_252, 1, x_231); +lean_ctor_set(x_252, 0, x_37); +lean_ctor_set(x_252, 1, x_235); lean_ctor_set(x_252, 2, x_251); -x_253 = lean_array_push(x_44, x_252); -x_254 = lean_array_push(x_253, x_66); -lean_inc(x_129); +x_253 = lean_array_push(x_45, x_252); +x_254 = lean_array_push(x_253, x_67); +lean_inc(x_130); x_255 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_255, 0, x_36); -lean_ctor_set(x_255, 1, x_129); +lean_ctor_set(x_255, 0, x_37); +lean_ctor_set(x_255, 1, x_130); lean_ctor_set(x_255, 2, x_254); -x_256 = lean_array_push(x_51, x_255); +x_256 = lean_array_push(x_52, x_255); x_257 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_257, 0, x_36); -lean_ctor_set(x_257, 1, x_53); +lean_ctor_set(x_257, 0, x_37); +lean_ctor_set(x_257, 1, x_54); lean_ctor_set(x_257, 2, x_256); -x_258 = lean_array_push(x_51, x_257); -lean_inc(x_127); +x_258 = lean_array_push(x_52, x_257); +lean_inc(x_128); x_259 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_259, 0, x_36); -lean_ctor_set(x_259, 1, x_127); +lean_ctor_set(x_259, 0, x_37); +lean_ctor_set(x_259, 1, x_128); lean_ctor_set(x_259, 2, x_258); -x_260 = lean_array_push(x_157, x_183); +x_260 = lean_array_push(x_158, x_184); lean_inc(x_260); -x_261 = lean_array_push(x_260, x_229); -lean_inc(x_122); -x_262 = lean_array_push(x_261, x_122); +x_261 = lean_array_push(x_260, x_233); +lean_inc(x_123); +x_262 = lean_array_push(x_261, x_123); x_263 = lean_array_push(x_262, x_259); -lean_inc(x_181); +lean_inc(x_182); x_264 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_264, 0, x_36); -lean_ctor_set(x_264, 1, x_181); +lean_ctor_set(x_264, 0, x_37); +lean_ctor_set(x_264, 1, x_182); lean_ctor_set(x_264, 2, x_263); x_265 = l_Lean_mkHole___closed__7; -lean_inc(x_27); -x_266 = lean_name_mk_string(x_27, x_265); +lean_inc(x_28); +x_266 = lean_name_mk_string(x_28, x_265); x_267 = l_Lean_Name_appendIndexAfter___closed__1; -lean_inc(x_17); +lean_inc(x_18); x_268 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_268, 0, x_17); +lean_ctor_set(x_268, 0, x_18); lean_ctor_set(x_268, 1, x_267); -x_269 = lean_array_push(x_51, x_268); +x_269 = lean_array_push(x_52, x_268); x_270 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_270, 0, x_36); +lean_ctor_set(x_270, 0, x_37); lean_ctor_set(x_270, 1, x_266); lean_ctor_set(x_270, 2, x_269); -x_271 = lean_array_push(x_51, x_270); +x_271 = lean_array_push(x_52, x_270); x_272 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_272, 0, x_36); -lean_ctor_set(x_272, 1, x_53); +lean_ctor_set(x_272, 0, x_37); +lean_ctor_set(x_272, 1, x_54); lean_ctor_set(x_272, 2, x_271); -x_273 = lean_array_push(x_51, x_272); +x_273 = lean_array_push(x_52, x_272); x_274 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_274, 0, x_36); -lean_ctor_set(x_274, 1, x_53); +lean_ctor_set(x_274, 0, x_37); +lean_ctor_set(x_274, 1, x_54); lean_ctor_set(x_274, 2, x_273); x_275 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__11; -lean_inc(x_17); +lean_inc(x_18); x_276 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_276, 0, x_17); +lean_ctor_set(x_276, 0, x_18); lean_ctor_set(x_276, 1, x_275); x_277 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__35; -lean_inc(x_17); +lean_inc(x_18); x_278 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_278, 0, x_17); +lean_ctor_set(x_278, 0, x_18); lean_ctor_set(x_278, 1, x_277); lean_inc(x_276); -x_279 = lean_array_push(x_44, x_276); +x_279 = lean_array_push(x_45, x_276); lean_inc(x_278); x_280 = lean_array_push(x_279, x_278); -x_281 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__68; +x_281 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__65; x_282 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_282, 0, x_36); +lean_ctor_set(x_282, 0, x_37); lean_ctor_set(x_282, 1, x_281); lean_ctor_set(x_282, 2, x_280); x_283 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__9; -lean_inc(x_27); -x_284 = lean_name_mk_string(x_27, x_283); +lean_inc(x_28); +x_284 = lean_name_mk_string(x_28, x_283); x_285 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__31; -lean_inc(x_27); -x_286 = lean_name_mk_string(x_27, x_285); +lean_inc(x_28); +x_286 = lean_name_mk_string(x_28, x_285); x_287 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_287, 0, x_36); +lean_ctor_set(x_287, 0, x_37); lean_ctor_set(x_287, 1, x_286); -lean_ctor_set(x_287, 2, x_37); +lean_ctor_set(x_287, 2, x_38); x_288 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__36; x_289 = lean_array_push(x_288, x_276); -x_290 = lean_array_push(x_289, x_66); -x_291 = lean_array_push(x_290, x_66); +x_290 = lean_array_push(x_289, x_67); +x_291 = lean_array_push(x_290, x_67); x_292 = lean_array_push(x_291, x_287); -x_293 = lean_array_push(x_292, x_66); +x_293 = lean_array_push(x_292, x_67); x_294 = lean_array_push(x_293, x_278); x_295 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_295, 0, x_36); +lean_ctor_set(x_295, 0, x_37); lean_ctor_set(x_295, 1, x_284); lean_ctor_set(x_295, 2, x_294); -x_296 = lean_array_push(x_44, x_282); +x_296 = lean_array_push(x_45, x_282); x_297 = lean_array_push(x_296, x_295); -x_298 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__66; +x_298 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__63; x_299 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_299, 0, x_36); +lean_ctor_set(x_299, 0, x_37); lean_ctor_set(x_299, 1, x_298); lean_ctor_set(x_299, 2, x_297); -x_300 = lean_array_push(x_51, x_299); +x_300 = lean_array_push(x_52, x_299); x_301 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_301, 0, x_36); -lean_ctor_set(x_301, 1, x_53); +lean_ctor_set(x_301, 0, x_37); +lean_ctor_set(x_301, 1, x_54); lean_ctor_set(x_301, 2, x_300); x_302 = lean_array_push(x_242, x_301); -lean_inc(x_233); +lean_inc(x_237); x_303 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_303, 0, x_36); -lean_ctor_set(x_303, 1, x_233); +lean_ctor_set(x_303, 0, x_37); +lean_ctor_set(x_303, 1, x_237); lean_ctor_set(x_303, 2, x_302); -x_304 = lean_array_push(x_216, x_303); -lean_inc(x_212); -x_305 = lean_array_push(x_304, x_212); +x_304 = lean_array_push(x_220, x_303); +lean_inc(x_216); +x_305 = lean_array_push(x_304, x_216); x_306 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_306, 0, x_36); -lean_ctor_set(x_306, 1, x_191); +lean_ctor_set(x_306, 0, x_37); +lean_ctor_set(x_306, 1, x_192); lean_ctor_set(x_306, 2, x_305); -x_307 = lean_array_push(x_222, x_306); -lean_inc(x_212); -x_308 = lean_array_push(x_307, x_212); +x_307 = lean_array_push(x_226, x_306); +lean_inc(x_216); +x_308 = lean_array_push(x_307, x_216); x_309 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_309, 0, x_36); -lean_ctor_set(x_309, 1, x_185); +lean_ctor_set(x_309, 0, x_37); +lean_ctor_set(x_309, 1, x_186); lean_ctor_set(x_309, 2, x_308); -x_310 = lean_array_push(x_51, x_309); +x_310 = lean_array_push(x_52, x_309); x_311 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_311, 0, x_36); -lean_ctor_set(x_311, 1, x_231); +lean_ctor_set(x_311, 0, x_37); +lean_ctor_set(x_311, 1, x_235); lean_ctor_set(x_311, 2, x_310); -x_312 = lean_array_push(x_44, x_311); -x_313 = lean_array_push(x_312, x_66); -lean_inc(x_129); +x_312 = lean_array_push(x_45, x_311); +x_313 = lean_array_push(x_312, x_67); +lean_inc(x_130); x_314 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_314, 0, x_36); -lean_ctor_set(x_314, 1, x_129); +lean_ctor_set(x_314, 0, x_37); +lean_ctor_set(x_314, 1, x_130); lean_ctor_set(x_314, 2, x_313); -x_315 = lean_array_push(x_51, x_314); +x_315 = lean_array_push(x_52, x_314); x_316 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_316, 0, x_36); -lean_ctor_set(x_316, 1, x_53); +lean_ctor_set(x_316, 0, x_37); +lean_ctor_set(x_316, 1, x_54); lean_ctor_set(x_316, 2, x_315); -x_317 = lean_array_push(x_51, x_316); -lean_inc(x_127); +x_317 = lean_array_push(x_52, x_316); +lean_inc(x_128); x_318 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_318, 0, x_36); -lean_ctor_set(x_318, 1, x_127); +lean_ctor_set(x_318, 0, x_37); +lean_ctor_set(x_318, 1, x_128); lean_ctor_set(x_318, 2, x_317); x_319 = lean_array_push(x_260, x_274); -lean_inc(x_122); -x_320 = lean_array_push(x_319, x_122); +lean_inc(x_123); +x_320 = lean_array_push(x_319, x_123); x_321 = lean_array_push(x_320, x_318); x_322 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_322, 0, x_36); -lean_ctor_set(x_322, 1, x_181); +lean_ctor_set(x_322, 0, x_37); +lean_ctor_set(x_322, 1, x_182); lean_ctor_set(x_322, 2, x_321); -x_323 = lean_array_push(x_44, x_264); +x_323 = lean_array_push(x_45, x_264); x_324 = lean_array_push(x_323, x_322); x_325 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_325, 0, x_36); -lean_ctor_set(x_325, 1, x_53); +lean_ctor_set(x_325, 0, x_37); +lean_ctor_set(x_325, 1, x_54); lean_ctor_set(x_325, 2, x_324); -x_326 = lean_array_push(x_51, x_325); +x_326 = lean_array_push(x_52, x_325); x_327 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_327, 0, x_36); -lean_ctor_set(x_327, 1, x_179); +lean_ctor_set(x_327, 0, x_37); +lean_ctor_set(x_327, 1, x_180); lean_ctor_set(x_327, 2, x_326); -x_328 = lean_array_push(x_288, x_145); -x_329 = lean_array_push(x_328, x_66); -x_330 = lean_array_push(x_329, x_66); -x_331 = lean_array_push(x_330, x_175); -x_332 = lean_array_push(x_331, x_177); +x_328 = lean_array_push(x_288, x_146); +x_329 = lean_array_push(x_328, x_67); +x_330 = lean_array_push(x_329, x_67); +x_331 = lean_array_push(x_330, x_176); +x_332 = lean_array_push(x_331, x_178); x_333 = lean_array_push(x_332, x_327); x_334 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_334, 0, x_36); -lean_ctor_set(x_334, 1, x_143); +lean_ctor_set(x_334, 0, x_37); +lean_ctor_set(x_334, 1, x_144); lean_ctor_set(x_334, 2, x_333); -lean_inc(x_139); -x_335 = lean_array_push(x_157, x_139); -x_336 = lean_array_push(x_335, x_66); -x_337 = lean_array_push(x_336, x_141); +lean_inc(x_140); +x_335 = lean_array_push(x_158, x_140); +x_336 = lean_array_push(x_335, x_67); +x_337 = lean_array_push(x_336, x_142); x_338 = lean_array_push(x_337, x_334); x_339 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_339, 0, x_36); -lean_ctor_set(x_339, 1, x_135); +lean_ctor_set(x_339, 0, x_37); +lean_ctor_set(x_339, 1, x_136); lean_ctor_set(x_339, 2, x_338); -x_340 = lean_array_push(x_57, x_133); -x_341 = lean_array_push(x_340, x_66); +x_340 = lean_array_push(x_58, x_134); +x_341 = lean_array_push(x_340, x_67); lean_inc(x_341); x_342 = lean_array_push(x_341, x_339); x_343 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_343, 0, x_36); -lean_ctor_set(x_343, 1, x_131); +lean_ctor_set(x_343, 0, x_37); +lean_ctor_set(x_343, 1, x_132); lean_ctor_set(x_343, 2, x_342); -x_344 = lean_array_push(x_44, x_343); -x_345 = lean_array_push(x_344, x_66); -lean_inc(x_129); +x_344 = lean_array_push(x_45, x_343); +x_345 = lean_array_push(x_344, x_67); +lean_inc(x_130); x_346 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_346, 0, x_36); -lean_ctor_set(x_346, 1, x_129); -lean_ctor_set(x_346, 2, x_345); -x_347 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__69; -lean_inc(x_27); -x_348 = lean_name_mk_string(x_27, x_347); -x_349 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__70; -lean_inc(x_27); -x_350 = lean_name_mk_string(x_27, x_349); -x_351 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__71; -lean_inc(x_27); -x_352 = lean_name_mk_string(x_27, x_351); -x_353 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__76; -lean_inc(x_18); +lean_ctor_set(x_346, 0, x_37); +lean_ctor_set(x_346, 1, x_130); +lean_ctor_set(x_346, 2, x_345); +x_347 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__66; +lean_inc(x_28); +x_348 = lean_name_mk_string(x_28, x_347); +x_349 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__67; +lean_inc(x_28); +x_350 = lean_name_mk_string(x_28, x_349); +x_351 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__68; +lean_inc(x_28); +x_352 = lean_name_mk_string(x_28, x_351); +x_353 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__73; lean_inc(x_19); -x_354 = l_Lean_addMacroScope(x_19, x_353, x_18); -x_355 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__74; -lean_inc(x_17); +lean_inc(x_20); +x_354 = l_Lean_addMacroScope(x_20, x_353, x_19); +x_355 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__71; +lean_inc(x_18); x_356 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_356, 0, x_17); +lean_ctor_set(x_356, 0, x_18); lean_ctor_set(x_356, 1, x_355); lean_ctor_set(x_356, 2, x_354); -lean_ctor_set(x_356, 3, x_79); -x_357 = lean_array_push(x_51, x_12); +lean_ctor_set(x_356, 3, x_80); +x_357 = lean_array_push(x_52, x_13); x_358 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_358, 0, x_36); -lean_ctor_set(x_358, 1, x_53); +lean_ctor_set(x_358, 0, x_37); +lean_ctor_set(x_358, 1, x_54); lean_ctor_set(x_358, 2, x_357); -x_359 = lean_array_push(x_44, x_356); +x_359 = lean_array_push(x_45, x_356); x_360 = lean_array_push(x_359, x_358); -lean_inc(x_233); +lean_inc(x_237); x_361 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_361, 0, x_36); -lean_ctor_set(x_361, 1, x_233); +lean_ctor_set(x_361, 0, x_37); +lean_ctor_set(x_361, 1, x_237); lean_ctor_set(x_361, 2, x_360); -x_362 = lean_array_push(x_213, x_118); -x_363 = lean_array_push(x_362, x_66); -x_364 = lean_array_push(x_363, x_66); -lean_inc(x_109); -x_365 = lean_array_push(x_364, x_109); +x_362 = lean_array_push(x_217, x_119); +x_363 = lean_array_push(x_362, x_67); +x_364 = lean_array_push(x_363, x_67); +lean_inc(x_110); +x_365 = lean_array_push(x_364, x_110); lean_inc(x_365); x_366 = lean_array_push(x_365, x_361); lean_inc(x_352); x_367 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_367, 0, x_36); +lean_ctor_set(x_367, 0, x_37); lean_ctor_set(x_367, 1, x_352); lean_ctor_set(x_367, 2, x_366); -x_368 = lean_array_push(x_51, x_367); +x_368 = lean_array_push(x_52, x_367); lean_inc(x_350); x_369 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_369, 0, x_36); +lean_ctor_set(x_369, 0, x_37); lean_ctor_set(x_369, 1, x_350); lean_ctor_set(x_369, 2, x_368); lean_inc(x_341); x_370 = lean_array_push(x_341, x_369); lean_inc(x_348); x_371 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_371, 0, x_36); +lean_ctor_set(x_371, 0, x_37); lean_ctor_set(x_371, 1, x_348); lean_ctor_set(x_371, 2, x_370); -x_372 = lean_array_push(x_44, x_371); -x_373 = lean_array_push(x_372, x_66); -lean_inc(x_129); +x_372 = lean_array_push(x_45, x_371); +x_373 = lean_array_push(x_372, x_67); +lean_inc(x_130); x_374 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_374, 0, x_36); -lean_ctor_set(x_374, 1, x_129); +lean_ctor_set(x_374, 0, x_37); +lean_ctor_set(x_374, 1, x_130); lean_ctor_set(x_374, 2, x_373); -x_375 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__81; -lean_inc(x_18); +x_375 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__78; lean_inc(x_19); -x_376 = l_Lean_addMacroScope(x_19, x_375, x_18); -x_377 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__79; -lean_inc(x_17); +lean_inc(x_20); +x_376 = l_Lean_addMacroScope(x_20, x_375, x_19); +x_377 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__76; +lean_inc(x_18); x_378 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_378, 0, x_17); +lean_ctor_set(x_378, 0, x_18); lean_ctor_set(x_378, 1, x_377); lean_ctor_set(x_378, 2, x_376); -lean_ctor_set(x_378, 3, x_79); -x_379 = l_Lean_Syntax_expandInterpolatedStr___closed__2; -lean_inc(x_27); -x_380 = lean_name_mk_string(x_27, x_379); -x_381 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__85; -lean_inc(x_18); +lean_ctor_set(x_378, 3, x_80); +x_379 = l_Lean_TSyntax_expandInterpolatedStr___closed__2; +lean_inc(x_28); +x_380 = lean_name_mk_string(x_28, x_379); +x_381 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__82; lean_inc(x_19); -x_382 = l_Lean_addMacroScope(x_19, x_381, x_18); -x_383 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__82; +lean_inc(x_20); +x_382 = l_Lean_addMacroScope(x_20, x_381, x_19); +x_383 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__79; lean_inc(x_3); x_384 = lean_name_mk_string(x_3, x_383); x_385 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_385, 0, x_384); -lean_ctor_set(x_385, 1, x_79); +lean_ctor_set(x_385, 1, x_80); x_386 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_386, 0, x_385); -lean_ctor_set(x_386, 1, x_79); -x_387 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__84; -lean_inc(x_17); +lean_ctor_set(x_386, 1, x_80); +x_387 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__81; +lean_inc(x_18); x_388 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_388, 0, x_17); +lean_ctor_set(x_388, 0, x_18); lean_ctor_set(x_388, 1, x_387); lean_ctor_set(x_388, 2, x_382); lean_ctor_set(x_388, 3, x_386); -lean_inc(x_166); -x_389 = lean_array_push(x_159, x_166); -lean_inc(x_56); -x_390 = lean_array_push(x_389, x_56); +lean_inc(x_167); +x_389 = lean_array_push(x_160, x_167); +lean_inc(x_57); +x_390 = lean_array_push(x_389, x_57); x_391 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_391, 0, x_36); -lean_ctor_set(x_391, 1, x_149); +lean_ctor_set(x_391, 0, x_37); +lean_ctor_set(x_391, 1, x_150); lean_ctor_set(x_391, 2, x_390); -x_392 = lean_array_push(x_44, x_391); -x_393 = lean_array_push(x_392, x_13); +x_392 = lean_array_push(x_45, x_391); +x_393 = lean_array_push(x_392, x_14); x_394 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_394, 0, x_36); -lean_ctor_set(x_394, 1, x_53); +lean_ctor_set(x_394, 0, x_37); +lean_ctor_set(x_394, 1, x_54); lean_ctor_set(x_394, 2, x_393); -x_395 = lean_array_push(x_44, x_388); +x_395 = lean_array_push(x_45, x_388); x_396 = lean_array_push(x_395, x_394); -lean_inc(x_233); +lean_inc(x_237); x_397 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_397, 0, x_36); -lean_ctor_set(x_397, 1, x_233); +lean_ctor_set(x_397, 0, x_37); +lean_ctor_set(x_397, 1, x_237); lean_ctor_set(x_397, 2, x_396); -x_398 = lean_array_push(x_44, x_397); -x_399 = lean_array_push(x_398, x_66); +x_398 = lean_array_push(x_45, x_397); +x_399 = lean_array_push(x_398, x_67); x_400 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_400, 0, x_36); -lean_ctor_set(x_400, 1, x_53); +lean_ctor_set(x_400, 0, x_37); +lean_ctor_set(x_400, 1, x_54); lean_ctor_set(x_400, 2, x_399); -x_401 = lean_array_push(x_57, x_196); +x_401 = lean_array_push(x_58, x_197); lean_inc(x_401); x_402 = lean_array_push(x_401, x_400); -lean_inc(x_212); -x_403 = lean_array_push(x_402, x_212); +lean_inc(x_216); +x_403 = lean_array_push(x_402, x_216); lean_inc(x_380); x_404 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_404, 0, x_36); +lean_ctor_set(x_404, 0, x_37); lean_ctor_set(x_404, 1, x_380); lean_ctor_set(x_404, 2, x_403); -x_405 = lean_array_push(x_44, x_166); +x_405 = lean_array_push(x_45, x_167); x_406 = lean_array_push(x_405, x_404); x_407 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_407, 0, x_36); -lean_ctor_set(x_407, 1, x_53); +lean_ctor_set(x_407, 0, x_37); +lean_ctor_set(x_407, 1, x_54); lean_ctor_set(x_407, 2, x_406); -x_408 = lean_array_push(x_44, x_378); +x_408 = lean_array_push(x_45, x_378); lean_inc(x_408); x_409 = lean_array_push(x_408, x_407); -lean_inc(x_233); +lean_inc(x_237); x_410 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_410, 0, x_36); -lean_ctor_set(x_410, 1, x_233); +lean_ctor_set(x_410, 0, x_37); +lean_ctor_set(x_410, 1, x_237); lean_ctor_set(x_410, 2, x_409); x_411 = lean_array_push(x_365, x_410); lean_inc(x_352); x_412 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_412, 0, x_36); +lean_ctor_set(x_412, 0, x_37); lean_ctor_set(x_412, 1, x_352); lean_ctor_set(x_412, 2, x_411); -x_413 = lean_array_push(x_51, x_412); +x_413 = lean_array_push(x_52, x_412); lean_inc(x_350); x_414 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_414, 0, x_36); +lean_ctor_set(x_414, 0, x_37); lean_ctor_set(x_414, 1, x_350); lean_ctor_set(x_414, 2, x_413); lean_inc(x_341); x_415 = lean_array_push(x_341, x_414); lean_inc(x_348); x_416 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_416, 0, x_36); +lean_ctor_set(x_416, 0, x_37); lean_ctor_set(x_416, 1, x_348); lean_ctor_set(x_416, 2, x_415); -x_417 = lean_array_push(x_44, x_416); -x_418 = lean_array_push(x_417, x_66); -lean_inc(x_129); +x_417 = lean_array_push(x_45, x_416); +x_418 = lean_array_push(x_417, x_67); +lean_inc(x_130); x_419 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_419, 0, x_36); -lean_ctor_set(x_419, 1, x_129); +lean_ctor_set(x_419, 0, x_37); +lean_ctor_set(x_419, 1, x_130); lean_ctor_set(x_419, 2, x_418); -x_420 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__89; -lean_inc(x_18); +x_420 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__86; lean_inc(x_19); -x_421 = l_Lean_addMacroScope(x_19, x_420, x_18); -x_422 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__88; -lean_inc(x_17); +lean_inc(x_20); +x_421 = l_Lean_addMacroScope(x_20, x_420, x_19); +x_422 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__85; +lean_inc(x_18); x_423 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_423, 0, x_17); +lean_ctor_set(x_423, 0, x_18); lean_ctor_set(x_423, 1, x_422); lean_ctor_set(x_423, 2, x_421); -lean_ctor_set(x_423, 3, x_79); -x_424 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__93; -x_425 = l_Lean_addMacroScope(x_19, x_424, x_18); -x_426 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__90; +lean_ctor_set(x_423, 3, x_80); +x_424 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__90; +x_425 = l_Lean_addMacroScope(x_20, x_424, x_19); +x_426 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__87; x_427 = lean_name_mk_string(x_3, x_426); x_428 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_428, 0, x_427); -lean_ctor_set(x_428, 1, x_79); +lean_ctor_set(x_428, 1, x_80); x_429 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_429, 0, x_428); -lean_ctor_set(x_429, 1, x_79); -x_430 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__92; -lean_inc(x_17); +lean_ctor_set(x_429, 1, x_80); +x_430 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__89; +lean_inc(x_18); x_431 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_431, 0, x_17); +lean_ctor_set(x_431, 0, x_18); lean_ctor_set(x_431, 1, x_430); lean_ctor_set(x_431, 2, x_425); lean_ctor_set(x_431, 3, x_429); x_432 = l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__12; -lean_inc(x_17); +lean_inc(x_18); x_433 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_433, 0, x_17); +lean_ctor_set(x_433, 0, x_18); lean_ctor_set(x_433, 1, x_432); -x_434 = lean_array_push(x_51, x_139); +x_434 = lean_array_push(x_52, x_140); x_435 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_435, 0, x_36); -lean_ctor_set(x_435, 1, x_53); +lean_ctor_set(x_435, 0, x_37); +lean_ctor_set(x_435, 1, x_54); lean_ctor_set(x_435, 2, x_434); -x_436 = lean_array_push(x_57, x_433); +x_436 = lean_array_push(x_58, x_433); x_437 = lean_array_push(x_436, x_435); -x_438 = lean_array_push(x_437, x_56); -x_439 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__95; +x_438 = lean_array_push(x_437, x_57); +x_439 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__92; x_440 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_440, 0, x_36); +lean_ctor_set(x_440, 0, x_37); lean_ctor_set(x_440, 1, x_439); lean_ctor_set(x_440, 2, x_438); -x_441 = lean_array_push(x_51, x_440); +x_441 = lean_array_push(x_52, x_440); x_442 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_442, 0, x_36); -lean_ctor_set(x_442, 1, x_53); +lean_ctor_set(x_442, 0, x_37); +lean_ctor_set(x_442, 1, x_54); lean_ctor_set(x_442, 2, x_441); -x_443 = lean_array_push(x_44, x_431); +x_443 = lean_array_push(x_45, x_431); x_444 = lean_array_push(x_443, x_442); -lean_inc(x_233); +lean_inc(x_237); x_445 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_445, 0, x_36); -lean_ctor_set(x_445, 1, x_233); +lean_ctor_set(x_445, 0, x_37); +lean_ctor_set(x_445, 1, x_237); lean_ctor_set(x_445, 2, x_444); -x_446 = lean_array_push(x_44, x_445); -x_447 = lean_array_push(x_446, x_66); +x_446 = lean_array_push(x_45, x_445); +x_447 = lean_array_push(x_446, x_67); x_448 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_448, 0, x_36); -lean_ctor_set(x_448, 1, x_53); +lean_ctor_set(x_448, 0, x_37); +lean_ctor_set(x_448, 1, x_54); lean_ctor_set(x_448, 2, x_447); x_449 = lean_array_push(x_401, x_448); -x_450 = lean_array_push(x_449, x_212); +x_450 = lean_array_push(x_449, x_216); x_451 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_451, 0, x_36); +lean_ctor_set(x_451, 0, x_37); lean_ctor_set(x_451, 1, x_380); lean_ctor_set(x_451, 2, x_450); -x_452 = lean_array_push(x_44, x_156); +x_452 = lean_array_push(x_45, x_157); x_453 = lean_array_push(x_452, x_451); x_454 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_454, 0, x_36); -lean_ctor_set(x_454, 1, x_53); +lean_ctor_set(x_454, 0, x_37); +lean_ctor_set(x_454, 1, x_54); lean_ctor_set(x_454, 2, x_453); x_455 = lean_array_push(x_408, x_454); x_456 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_456, 0, x_36); -lean_ctor_set(x_456, 1, x_233); +lean_ctor_set(x_456, 0, x_37); +lean_ctor_set(x_456, 1, x_237); lean_ctor_set(x_456, 2, x_455); lean_inc(x_423); -x_457 = lean_array_push(x_213, x_423); -x_458 = lean_array_push(x_457, x_66); -x_459 = lean_array_push(x_458, x_66); -lean_inc(x_109); -x_460 = lean_array_push(x_459, x_109); +x_457 = lean_array_push(x_217, x_423); +x_458 = lean_array_push(x_457, x_67); +x_459 = lean_array_push(x_458, x_67); +lean_inc(x_110); +x_460 = lean_array_push(x_459, x_110); x_461 = lean_array_push(x_460, x_456); x_462 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_462, 0, x_36); +lean_ctor_set(x_462, 0, x_37); lean_ctor_set(x_462, 1, x_352); lean_ctor_set(x_462, 2, x_461); -x_463 = lean_array_push(x_51, x_462); +x_463 = lean_array_push(x_52, x_462); x_464 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_464, 0, x_36); +lean_ctor_set(x_464, 0, x_37); lean_ctor_set(x_464, 1, x_350); lean_ctor_set(x_464, 2, x_463); x_465 = lean_array_push(x_341, x_464); x_466 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_466, 0, x_36); +lean_ctor_set(x_466, 0, x_37); lean_ctor_set(x_466, 1, x_348); lean_ctor_set(x_466, 2, x_465); -x_467 = lean_array_push(x_44, x_466); -x_468 = lean_array_push(x_467, x_66); -lean_inc(x_129); +x_467 = lean_array_push(x_45, x_466); +x_468 = lean_array_push(x_467, x_67); +lean_inc(x_130); x_469 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_469, 0, x_36); -lean_ctor_set(x_469, 1, x_129); +lean_ctor_set(x_469, 0, x_37); +lean_ctor_set(x_469, 1, x_130); lean_ctor_set(x_469, 2, x_468); -x_470 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__96; -x_471 = lean_name_mk_string(x_27, x_470); -x_472 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__97; +x_470 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__93; +x_471 = lean_name_mk_string(x_28, x_470); +x_472 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__94; x_473 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_473, 0, x_17); +lean_ctor_set(x_473, 0, x_18); lean_ctor_set(x_473, 1, x_472); -x_474 = lean_array_push(x_51, x_423); +x_474 = lean_array_push(x_52, x_423); x_475 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_475, 0, x_36); -lean_ctor_set(x_475, 1, x_53); +lean_ctor_set(x_475, 0, x_37); +lean_ctor_set(x_475, 1, x_54); lean_ctor_set(x_475, 2, x_474); -x_476 = lean_array_push(x_44, x_473); +x_476 = lean_array_push(x_45, x_473); x_477 = lean_array_push(x_476, x_475); x_478 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_478, 0, x_36); +lean_ctor_set(x_478, 0, x_37); lean_ctor_set(x_478, 1, x_471); lean_ctor_set(x_478, 2, x_477); -x_479 = lean_array_push(x_44, x_478); -x_480 = lean_array_push(x_479, x_66); +x_479 = lean_array_push(x_45, x_478); +x_480 = lean_array_push(x_479, x_67); x_481 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_481, 0, x_36); -lean_ctor_set(x_481, 1, x_129); +lean_ctor_set(x_481, 0, x_37); +lean_ctor_set(x_481, 1, x_130); lean_ctor_set(x_481, 2, x_480); -x_482 = lean_array_push(x_213, x_346); +x_482 = lean_array_push(x_217, x_346); x_483 = lean_array_push(x_482, x_374); x_484 = lean_array_push(x_483, x_419); x_485 = lean_array_push(x_484, x_469); x_486 = lean_array_push(x_485, x_481); x_487 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_487, 0, x_36); -lean_ctor_set(x_487, 1, x_53); +lean_ctor_set(x_487, 0, x_37); +lean_ctor_set(x_487, 1, x_54); lean_ctor_set(x_487, 2, x_486); -x_488 = lean_array_push(x_51, x_487); +x_488 = lean_array_push(x_52, x_487); x_489 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_489, 0, x_36); -lean_ctor_set(x_489, 1, x_127); +lean_ctor_set(x_489, 0, x_37); +lean_ctor_set(x_489, 1, x_128); lean_ctor_set(x_489, 2, x_488); -x_490 = lean_array_push(x_44, x_125); +x_490 = lean_array_push(x_45, x_126); x_491 = lean_array_push(x_490, x_489); x_492 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_492, 0, x_36); -lean_ctor_set(x_492, 1, x_124); +lean_ctor_set(x_492, 0, x_37); +lean_ctor_set(x_492, 1, x_125); lean_ctor_set(x_492, 2, x_491); -x_493 = lean_array_push(x_57, x_120); -x_494 = lean_array_push(x_493, x_122); +x_493 = lean_array_push(x_58, x_121); +x_494 = lean_array_push(x_493, x_123); x_495 = lean_array_push(x_494, x_492); x_496 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_496, 0, x_36); -lean_ctor_set(x_496, 1, x_114); +lean_ctor_set(x_496, 0, x_37); +lean_ctor_set(x_496, 1, x_115); lean_ctor_set(x_496, 2, x_495); -x_497 = lean_array_push(x_44, x_112); +x_497 = lean_array_push(x_45, x_113); x_498 = lean_array_push(x_497, x_496); x_499 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_499, 0, x_36); -lean_ctor_set(x_499, 1, x_111); +lean_ctor_set(x_499, 0, x_37); +lean_ctor_set(x_499, 1, x_112); lean_ctor_set(x_499, 2, x_498); -x_500 = lean_array_push(x_57, x_109); +x_500 = lean_array_push(x_58, x_110); x_501 = lean_array_push(x_500, x_499); -x_502 = lean_array_push(x_501, x_66); +x_502 = lean_array_push(x_501, x_67); x_503 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_503, 0, x_36); -lean_ctor_set(x_503, 1, x_107); +lean_ctor_set(x_503, 0, x_37); +lean_ctor_set(x_503, 1, x_108); lean_ctor_set(x_503, 2, x_502); -x_504 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__98; -x_505 = lean_array_push(x_504, x_74); -x_506 = lean_array_push(x_505, x_84); -x_507 = lean_array_push(x_506, x_105); +x_504 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__95; +x_505 = lean_array_push(x_504, x_75); +x_506 = lean_array_push(x_505, x_85); +x_507 = lean_array_push(x_506, x_106); x_508 = lean_array_push(x_507, x_503); -x_509 = lean_array_push(x_508, x_66); -x_510 = lean_array_push(x_509, x_66); -x_511 = lean_array_push(x_510, x_66); +x_509 = lean_array_push(x_508, x_67); +x_510 = lean_array_push(x_509, x_67); +x_511 = lean_array_push(x_510, x_67); x_512 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_512, 0, x_36); -lean_ctor_set(x_512, 1, x_73); +lean_ctor_set(x_512, 0, x_37); +lean_ctor_set(x_512, 1, x_74); lean_ctor_set(x_512, 2, x_511); -x_513 = lean_array_push(x_44, x_71); +x_513 = lean_array_push(x_45, x_72); x_514 = lean_array_push(x_513, x_512); x_515 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_515, 0, x_36); -lean_ctor_set(x_515, 1, x_23); +lean_ctor_set(x_515, 0, x_37); +lean_ctor_set(x_515, 1, x_24); lean_ctor_set(x_515, 2, x_514); -x_516 = lean_array_push(x_44, x_14); +x_516 = lean_array_push(x_45, x_15); x_517 = lean_array_push(x_516, x_515); x_518 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_518, 0, x_36); -lean_ctor_set(x_518, 1, x_53); +lean_ctor_set(x_518, 0, x_37); +lean_ctor_set(x_518, 1, x_54); lean_ctor_set(x_518, 2, x_517); -lean_ctor_set(x_15, 0, x_518); -return x_15; +lean_ctor_set(x_16, 0, x_518); +return x_16; } else { -lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; lean_object* x_565; lean_object* x_566; lean_object* x_567; lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; lean_object* x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; lean_object* x_678; lean_object* x_679; lean_object* x_680; lean_object* x_681; lean_object* x_682; lean_object* x_683; lean_object* x_684; lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; lean_object* x_690; lean_object* x_691; lean_object* x_692; lean_object* x_693; lean_object* x_694; lean_object* x_695; lean_object* x_696; lean_object* x_697; lean_object* x_698; lean_object* x_699; lean_object* x_700; lean_object* x_701; lean_object* x_702; lean_object* x_703; lean_object* x_704; lean_object* x_705; lean_object* x_706; lean_object* x_707; lean_object* x_708; lean_object* x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; lean_object* x_718; lean_object* x_719; lean_object* x_720; lean_object* x_721; lean_object* x_722; lean_object* x_723; lean_object* x_724; lean_object* x_725; lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; lean_object* x_730; lean_object* x_731; lean_object* x_732; lean_object* x_733; lean_object* x_734; lean_object* x_735; lean_object* x_736; lean_object* x_737; lean_object* x_738; lean_object* x_739; lean_object* x_740; lean_object* x_741; lean_object* x_742; lean_object* x_743; lean_object* x_744; lean_object* x_745; lean_object* x_746; lean_object* x_747; lean_object* x_748; lean_object* x_749; lean_object* x_750; lean_object* x_751; lean_object* x_752; lean_object* x_753; lean_object* x_754; lean_object* x_755; lean_object* x_756; lean_object* x_757; lean_object* x_758; lean_object* x_759; lean_object* x_760; lean_object* x_761; lean_object* x_762; lean_object* x_763; lean_object* x_764; lean_object* x_765; lean_object* x_766; lean_object* x_767; lean_object* x_768; lean_object* x_769; lean_object* x_770; lean_object* x_771; lean_object* x_772; lean_object* x_773; lean_object* x_774; lean_object* x_775; lean_object* x_776; lean_object* x_777; lean_object* x_778; lean_object* x_779; lean_object* x_780; lean_object* x_781; lean_object* x_782; lean_object* x_783; lean_object* x_784; lean_object* x_785; lean_object* x_786; lean_object* x_787; lean_object* x_788; lean_object* x_789; lean_object* x_790; lean_object* x_791; lean_object* x_792; lean_object* x_793; lean_object* x_794; lean_object* x_795; lean_object* x_796; lean_object* x_797; lean_object* x_798; lean_object* x_799; lean_object* x_800; lean_object* x_801; lean_object* x_802; lean_object* x_803; lean_object* x_804; lean_object* x_805; lean_object* x_806; lean_object* x_807; lean_object* x_808; lean_object* x_809; lean_object* x_810; lean_object* x_811; lean_object* x_812; lean_object* x_813; lean_object* x_814; lean_object* x_815; lean_object* x_816; lean_object* x_817; lean_object* x_818; lean_object* x_819; lean_object* x_820; lean_object* x_821; lean_object* x_822; lean_object* x_823; lean_object* x_824; lean_object* x_825; lean_object* x_826; lean_object* x_827; lean_object* x_828; lean_object* x_829; lean_object* x_830; lean_object* x_831; lean_object* x_832; lean_object* x_833; lean_object* x_834; lean_object* x_835; lean_object* x_836; lean_object* x_837; lean_object* x_838; lean_object* x_839; lean_object* x_840; lean_object* x_841; lean_object* x_842; lean_object* x_843; lean_object* x_844; lean_object* x_845; lean_object* x_846; lean_object* x_847; lean_object* x_848; lean_object* x_849; lean_object* x_850; lean_object* x_851; lean_object* x_852; lean_object* x_853; lean_object* x_854; lean_object* x_855; lean_object* x_856; lean_object* x_857; lean_object* x_858; lean_object* x_859; lean_object* x_860; lean_object* x_861; lean_object* x_862; lean_object* x_863; lean_object* x_864; lean_object* x_865; lean_object* x_866; lean_object* x_867; lean_object* x_868; lean_object* x_869; lean_object* x_870; lean_object* x_871; lean_object* x_872; lean_object* x_873; lean_object* x_874; lean_object* x_875; lean_object* x_876; lean_object* x_877; lean_object* x_878; lean_object* x_879; lean_object* x_880; lean_object* x_881; lean_object* x_882; lean_object* x_883; lean_object* x_884; lean_object* x_885; lean_object* x_886; lean_object* x_887; lean_object* x_888; lean_object* x_889; lean_object* x_890; lean_object* x_891; lean_object* x_892; lean_object* x_893; lean_object* x_894; lean_object* x_895; lean_object* x_896; lean_object* x_897; lean_object* x_898; lean_object* x_899; lean_object* x_900; lean_object* x_901; lean_object* x_902; lean_object* x_903; lean_object* x_904; lean_object* x_905; lean_object* x_906; lean_object* x_907; lean_object* x_908; lean_object* x_909; lean_object* x_910; lean_object* x_911; lean_object* x_912; lean_object* x_913; lean_object* x_914; lean_object* x_915; lean_object* x_916; lean_object* x_917; lean_object* x_918; lean_object* x_919; lean_object* x_920; lean_object* x_921; lean_object* x_922; lean_object* x_923; lean_object* x_924; lean_object* x_925; lean_object* x_926; lean_object* x_927; lean_object* x_928; lean_object* x_929; lean_object* x_930; lean_object* x_931; lean_object* x_932; lean_object* x_933; lean_object* x_934; lean_object* x_935; lean_object* x_936; lean_object* x_937; lean_object* x_938; lean_object* x_939; lean_object* x_940; lean_object* x_941; lean_object* x_942; lean_object* x_943; lean_object* x_944; lean_object* x_945; lean_object* x_946; lean_object* x_947; lean_object* x_948; lean_object* x_949; lean_object* x_950; lean_object* x_951; lean_object* x_952; lean_object* x_953; lean_object* x_954; lean_object* x_955; lean_object* x_956; lean_object* x_957; lean_object* x_958; lean_object* x_959; lean_object* x_960; lean_object* x_961; lean_object* x_962; lean_object* x_963; lean_object* x_964; lean_object* x_965; lean_object* x_966; lean_object* x_967; lean_object* x_968; lean_object* x_969; lean_object* x_970; lean_object* x_971; lean_object* x_972; lean_object* x_973; lean_object* x_974; lean_object* x_975; lean_object* x_976; lean_object* x_977; lean_object* x_978; lean_object* x_979; lean_object* x_980; lean_object* x_981; lean_object* x_982; lean_object* x_983; lean_object* x_984; lean_object* x_985; lean_object* x_986; lean_object* x_987; lean_object* x_988; lean_object* x_989; lean_object* x_990; lean_object* x_991; lean_object* x_992; lean_object* x_993; lean_object* x_994; lean_object* x_995; lean_object* x_996; lean_object* x_997; lean_object* x_998; lean_object* x_999; lean_object* x_1000; lean_object* x_1001; lean_object* x_1002; lean_object* x_1003; lean_object* x_1004; lean_object* x_1005; lean_object* x_1006; lean_object* x_1007; lean_object* x_1008; lean_object* x_1009; lean_object* x_1010; lean_object* x_1011; lean_object* x_1012; lean_object* x_1013; lean_object* x_1014; lean_object* x_1015; lean_object* x_1016; lean_object* x_1017; lean_object* x_1018; lean_object* x_1019; lean_object* x_1020; lean_object* x_1021; lean_object* x_1022; -x_519 = lean_ctor_get(x_15, 0); -x_520 = lean_ctor_get(x_15, 1); +lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; lean_object* x_565; lean_object* x_566; lean_object* x_567; lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; lean_object* x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; lean_object* x_678; lean_object* x_679; lean_object* x_680; lean_object* x_681; lean_object* x_682; lean_object* x_683; lean_object* x_684; lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; lean_object* x_690; lean_object* x_691; lean_object* x_692; lean_object* x_693; lean_object* x_694; lean_object* x_695; lean_object* x_696; lean_object* x_697; lean_object* x_698; lean_object* x_699; lean_object* x_700; lean_object* x_701; lean_object* x_702; lean_object* x_703; lean_object* x_704; lean_object* x_705; lean_object* x_706; lean_object* x_707; lean_object* x_708; lean_object* x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; lean_object* x_718; lean_object* x_719; lean_object* x_720; lean_object* x_721; lean_object* x_722; lean_object* x_723; lean_object* x_724; lean_object* x_725; lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; lean_object* x_730; lean_object* x_731; lean_object* x_732; lean_object* x_733; lean_object* x_734; lean_object* x_735; lean_object* x_736; lean_object* x_737; lean_object* x_738; lean_object* x_739; lean_object* x_740; lean_object* x_741; lean_object* x_742; lean_object* x_743; lean_object* x_744; lean_object* x_745; lean_object* x_746; lean_object* x_747; lean_object* x_748; lean_object* x_749; lean_object* x_750; lean_object* x_751; lean_object* x_752; lean_object* x_753; lean_object* x_754; lean_object* x_755; lean_object* x_756; lean_object* x_757; lean_object* x_758; lean_object* x_759; lean_object* x_760; lean_object* x_761; lean_object* x_762; lean_object* x_763; lean_object* x_764; lean_object* x_765; lean_object* x_766; lean_object* x_767; lean_object* x_768; lean_object* x_769; lean_object* x_770; lean_object* x_771; lean_object* x_772; lean_object* x_773; lean_object* x_774; lean_object* x_775; lean_object* x_776; lean_object* x_777; lean_object* x_778; lean_object* x_779; lean_object* x_780; lean_object* x_781; lean_object* x_782; lean_object* x_783; lean_object* x_784; lean_object* x_785; lean_object* x_786; lean_object* x_787; lean_object* x_788; lean_object* x_789; lean_object* x_790; lean_object* x_791; lean_object* x_792; lean_object* x_793; lean_object* x_794; lean_object* x_795; lean_object* x_796; lean_object* x_797; lean_object* x_798; lean_object* x_799; lean_object* x_800; lean_object* x_801; lean_object* x_802; lean_object* x_803; lean_object* x_804; lean_object* x_805; lean_object* x_806; lean_object* x_807; lean_object* x_808; lean_object* x_809; lean_object* x_810; lean_object* x_811; lean_object* x_812; lean_object* x_813; lean_object* x_814; lean_object* x_815; lean_object* x_816; lean_object* x_817; lean_object* x_818; lean_object* x_819; lean_object* x_820; lean_object* x_821; lean_object* x_822; lean_object* x_823; lean_object* x_824; lean_object* x_825; lean_object* x_826; lean_object* x_827; lean_object* x_828; lean_object* x_829; lean_object* x_830; lean_object* x_831; lean_object* x_832; lean_object* x_833; lean_object* x_834; lean_object* x_835; lean_object* x_836; lean_object* x_837; lean_object* x_838; lean_object* x_839; lean_object* x_840; lean_object* x_841; lean_object* x_842; lean_object* x_843; lean_object* x_844; lean_object* x_845; lean_object* x_846; lean_object* x_847; lean_object* x_848; lean_object* x_849; lean_object* x_850; lean_object* x_851; lean_object* x_852; lean_object* x_853; lean_object* x_854; lean_object* x_855; lean_object* x_856; lean_object* x_857; lean_object* x_858; lean_object* x_859; lean_object* x_860; lean_object* x_861; lean_object* x_862; lean_object* x_863; lean_object* x_864; lean_object* x_865; lean_object* x_866; lean_object* x_867; lean_object* x_868; lean_object* x_869; lean_object* x_870; lean_object* x_871; lean_object* x_872; lean_object* x_873; lean_object* x_874; lean_object* x_875; lean_object* x_876; lean_object* x_877; lean_object* x_878; lean_object* x_879; lean_object* x_880; lean_object* x_881; lean_object* x_882; lean_object* x_883; lean_object* x_884; lean_object* x_885; lean_object* x_886; lean_object* x_887; lean_object* x_888; lean_object* x_889; lean_object* x_890; lean_object* x_891; lean_object* x_892; lean_object* x_893; lean_object* x_894; lean_object* x_895; lean_object* x_896; lean_object* x_897; lean_object* x_898; lean_object* x_899; lean_object* x_900; lean_object* x_901; lean_object* x_902; lean_object* x_903; lean_object* x_904; lean_object* x_905; lean_object* x_906; lean_object* x_907; lean_object* x_908; lean_object* x_909; lean_object* x_910; lean_object* x_911; lean_object* x_912; lean_object* x_913; lean_object* x_914; lean_object* x_915; lean_object* x_916; lean_object* x_917; lean_object* x_918; lean_object* x_919; lean_object* x_920; lean_object* x_921; lean_object* x_922; lean_object* x_923; lean_object* x_924; lean_object* x_925; lean_object* x_926; lean_object* x_927; lean_object* x_928; lean_object* x_929; lean_object* x_930; lean_object* x_931; lean_object* x_932; lean_object* x_933; lean_object* x_934; lean_object* x_935; lean_object* x_936; lean_object* x_937; lean_object* x_938; lean_object* x_939; lean_object* x_940; lean_object* x_941; lean_object* x_942; lean_object* x_943; lean_object* x_944; lean_object* x_945; lean_object* x_946; lean_object* x_947; lean_object* x_948; lean_object* x_949; lean_object* x_950; lean_object* x_951; lean_object* x_952; lean_object* x_953; lean_object* x_954; lean_object* x_955; lean_object* x_956; lean_object* x_957; lean_object* x_958; lean_object* x_959; lean_object* x_960; lean_object* x_961; lean_object* x_962; lean_object* x_963; lean_object* x_964; lean_object* x_965; lean_object* x_966; lean_object* x_967; lean_object* x_968; lean_object* x_969; lean_object* x_970; lean_object* x_971; lean_object* x_972; lean_object* x_973; lean_object* x_974; lean_object* x_975; lean_object* x_976; lean_object* x_977; lean_object* x_978; lean_object* x_979; lean_object* x_980; lean_object* x_981; lean_object* x_982; lean_object* x_983; lean_object* x_984; lean_object* x_985; lean_object* x_986; lean_object* x_987; lean_object* x_988; lean_object* x_989; lean_object* x_990; lean_object* x_991; lean_object* x_992; lean_object* x_993; lean_object* x_994; lean_object* x_995; lean_object* x_996; lean_object* x_997; lean_object* x_998; lean_object* x_999; lean_object* x_1000; lean_object* x_1001; lean_object* x_1002; lean_object* x_1003; lean_object* x_1004; lean_object* x_1005; lean_object* x_1006; lean_object* x_1007; lean_object* x_1008; lean_object* x_1009; lean_object* x_1010; lean_object* x_1011; lean_object* x_1012; lean_object* x_1013; lean_object* x_1014; lean_object* x_1015; lean_object* x_1016; lean_object* x_1017; lean_object* x_1018; lean_object* x_1019; lean_object* x_1020; lean_object* x_1021; +x_519 = lean_ctor_get(x_16, 0); +x_520 = lean_ctor_get(x_16, 1); lean_inc(x_520); lean_inc(x_519); -lean_dec(x_15); -x_521 = lean_ctor_get(x_7, 2); +lean_dec(x_16); +x_521 = lean_ctor_get(x_8, 2); lean_inc(x_521); -x_522 = lean_ctor_get(x_7, 1); +x_522 = lean_ctor_get(x_8, 1); lean_inc(x_522); -lean_dec(x_7); -x_523 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__1; +lean_dec(x_8); +x_523 = l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___closed__1; lean_inc(x_1); x_524 = lean_name_mk_string(x_1, x_523); -x_525 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__2; +x_525 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__1; lean_inc(x_524); x_526 = lean_name_mk_string(x_524, x_525); -x_527 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__3; +x_527 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__2; lean_inc(x_524); x_528 = lean_name_mk_string(x_524, x_527); x_529 = l_Lean_mkHole___closed__5; lean_inc(x_1); x_530 = lean_name_mk_string(x_1, x_529); -x_531 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__4; +x_531 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__3; lean_inc(x_530); x_532 = lean_name_mk_string(x_530, x_531); -x_533 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__5; +x_533 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__4; lean_inc(x_519); x_534 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_534, 0, x_519); lean_ctor_set(x_534, 1, x_533); -x_535 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__6; +x_535 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__5; lean_inc(x_530); x_536 = lean_name_mk_string(x_530, x_535); -x_537 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__7; +x_537 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__6; lean_inc(x_530); x_538 = lean_name_mk_string(x_530, x_537); x_539 = lean_box(2); @@ -23321,9 +23975,9 @@ x_541 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_541, 0, x_539); lean_ctor_set(x_541, 1, x_538); lean_ctor_set(x_541, 2, x_540); -x_542 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__8; +x_542 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__7; x_543 = lean_name_mk_string(x_1, x_542); -x_544 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__9; +x_544 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__8; x_545 = lean_name_mk_string(x_543, x_544); lean_inc(x_519); x_546 = lean_alloc_ctor(2, 2, 0); @@ -23344,7 +23998,7 @@ lean_ctor_set(x_553, 1, x_536); lean_ctor_set(x_553, 2, x_552); x_554 = l_Lean_mkOptionalNode___closed__2; x_555 = lean_array_push(x_554, x_553); -x_556 = l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__2; +x_556 = l_Lean_mkNullNode___closed__2; x_557 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_557, 0, x_539); lean_ctor_set(x_557, 1, x_556); @@ -23354,7 +24008,7 @@ lean_inc(x_519); x_559 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_559, 0, x_519); lean_ctor_set(x_559, 1, x_558); -x_560 = l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__4; +x_560 = l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__4; x_561 = lean_array_push(x_560, x_534); x_562 = lean_array_push(x_561, x_557); lean_inc(x_559); @@ -23368,7 +24022,7 @@ x_566 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_566, 0, x_539); lean_ctor_set(x_566, 1, x_556); lean_ctor_set(x_566, 2, x_565); -x_567 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__10; +x_567 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__9; x_568 = lean_array_push(x_567, x_566); x_569 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__12; x_570 = lean_array_push(x_568, x_569); @@ -23379,22 +24033,22 @@ x_574 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_574, 0, x_539); lean_ctor_set(x_574, 1, x_528); lean_ctor_set(x_574, 2, x_573); -x_575 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__11; +x_575 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__10; lean_inc(x_524); x_576 = lean_name_mk_string(x_524, x_575); lean_inc(x_519); x_577 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_577, 0, x_519); lean_ctor_set(x_577, 1, x_575); -x_578 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__12; +x_578 = l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___closed__3; lean_inc(x_524); x_579 = lean_name_mk_string(x_524, x_578); -x_580 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__16; +x_580 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__14; lean_inc(x_521); lean_inc(x_522); x_581 = l_Lean_addMacroScope(x_522, x_580, x_521); x_582 = lean_box(0); -x_583 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__15; +x_583 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__13; lean_inc(x_519); x_584 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_584, 0, x_519); @@ -23407,10 +24061,10 @@ x_587 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_587, 0, x_539); lean_ctor_set(x_587, 1, x_579); lean_ctor_set(x_587, 2, x_586); -x_588 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__17; +x_588 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__15; lean_inc(x_524); x_589 = lean_name_mk_string(x_524, x_588); -x_590 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__18; +x_590 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__16; lean_inc(x_530); x_591 = lean_name_mk_string(x_530, x_590); x_592 = l_Lean_toolchain___closed__1; @@ -23418,19 +24072,19 @@ lean_inc(x_519); x_593 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_593, 0, x_519); lean_ctor_set(x_593, 1, x_592); -x_594 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__22; +x_594 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__20; lean_inc(x_521); lean_inc(x_522); x_595 = l_Lean_addMacroScope(x_522, x_594, x_521); -x_596 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__19; +x_596 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__17; lean_inc(x_3); x_597 = lean_name_mk_string(x_3, x_596); -lean_ctor_set(x_10, 1, x_582); -lean_ctor_set(x_10, 0, x_597); +lean_ctor_set(x_11, 1, x_582); +lean_ctor_set(x_11, 0, x_597); x_598 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_598, 0, x_10); +lean_ctor_set(x_598, 0, x_11); lean_ctor_set(x_598, 1, x_582); -x_599 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__21; +x_599 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__19; lean_inc(x_519); x_600 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_600, 0, x_519); @@ -23449,34 +24103,34 @@ x_605 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_605, 0, x_539); lean_ctor_set(x_605, 1, x_556); lean_ctor_set(x_605, 2, x_604); -x_606 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__23; +x_606 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__21; x_607 = lean_array_push(x_606, x_605); x_608 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_608, 0, x_539); lean_ctor_set(x_608, 1, x_589); lean_ctor_set(x_608, 2, x_607); -x_609 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__24; +x_609 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__22; x_610 = lean_name_mk_string(x_524, x_609); x_611 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__8; lean_inc(x_519); x_612 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_612, 0, x_519); lean_ctor_set(x_612, 1, x_611); -x_613 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__25; +x_613 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__23; lean_inc(x_530); x_614 = lean_name_mk_string(x_530, x_613); lean_inc(x_519); x_615 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_615, 0, x_519); lean_ctor_set(x_615, 1, x_613); -x_616 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__26; +x_616 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__24; lean_inc(x_530); x_617 = lean_name_mk_string(x_530, x_616); -x_618 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__30; +x_618 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__28; lean_inc(x_521); lean_inc(x_522); x_619 = l_Lean_addMacroScope(x_522, x_618, x_521); -x_620 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__29; +x_620 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__27; lean_inc(x_519); x_621 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_621, 0, x_519); @@ -23489,63 +24143,63 @@ x_623 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_623, 0, x_539); lean_ctor_set(x_623, 1, x_556); lean_ctor_set(x_623, 2, x_622); -x_624 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__31; +x_624 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__29; lean_inc(x_519); x_625 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_625, 0, x_519); lean_ctor_set(x_625, 1, x_624); -x_626 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__32; +x_626 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__30; lean_inc(x_530); x_627 = lean_name_mk_string(x_530, x_626); lean_inc(x_519); x_628 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_628, 0, x_519); lean_ctor_set(x_628, 1, x_626); -x_629 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__33; +x_629 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__31; lean_inc(x_530); x_630 = lean_name_mk_string(x_530, x_629); -x_631 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__34; +x_631 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__32; lean_inc(x_530); x_632 = lean_name_mk_string(x_530, x_631); -x_633 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__35; +x_633 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__33; lean_inc(x_530); x_634 = lean_name_mk_string(x_530, x_633); -x_635 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__36; +x_635 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__34; lean_inc(x_519); x_636 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_636, 0, x_519); lean_ctor_set(x_636, 1, x_635); -x_637 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__37; +x_637 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__35; lean_inc(x_530); x_638 = lean_name_mk_string(x_530, x_637); -x_639 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__41; +x_639 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__39; lean_inc(x_521); lean_inc(x_522); x_640 = l_Lean_addMacroScope(x_522, x_639, x_521); -x_641 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__40; +x_641 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__38; lean_inc(x_519); x_642 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_642, 0, x_519); lean_ctor_set(x_642, 1, x_641); lean_ctor_set(x_642, 2, x_640); lean_ctor_set(x_642, 3, x_582); -x_643 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__42; +x_643 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__40; lean_inc(x_519); x_644 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_644, 0, x_519); lean_ctor_set(x_644, 1, x_643); -x_645 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__43; +x_645 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__41; lean_inc(x_530); x_646 = lean_name_mk_string(x_530, x_645); -x_647 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__44; +x_647 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__42; lean_inc(x_519); x_648 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_648, 0, x_519); lean_ctor_set(x_648, 1, x_647); -x_649 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__45; +x_649 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__43; lean_inc(x_530); x_650 = lean_name_mk_string(x_530, x_649); -x_651 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__46; +x_651 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__44; lean_inc(x_530); x_652 = lean_name_mk_string(x_530, x_651); x_653 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__12; @@ -23553,13 +24207,13 @@ lean_inc(x_519); x_654 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_654, 0, x_519); lean_ctor_set(x_654, 1, x_653); -x_655 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__47; +x_655 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__45; lean_inc(x_519); x_656 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_656, 0, x_519); lean_ctor_set(x_656, 1, x_655); x_657 = lean_array_push(x_554, x_656); -x_658 = l_Lean_evalPrec___closed__2; +x_658 = l_Lean_Syntax_mkNumLit___closed__2; x_659 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_659, 0, x_539); lean_ctor_set(x_659, 1, x_658); @@ -23579,7 +24233,7 @@ x_665 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_665, 0, x_539); lean_ctor_set(x_665, 1, x_652); lean_ctor_set(x_665, 2, x_664); -x_666 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__48; +x_666 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__46; lean_inc(x_519); x_667 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_667, 0, x_519); @@ -23610,43 +24264,43 @@ x_678 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_678, 0, x_539); lean_ctor_set(x_678, 1, x_556); lean_ctor_set(x_678, 2, x_677); -x_679 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__49; +x_679 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__47; lean_inc(x_519); x_680 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_680, 0, x_519); lean_ctor_set(x_680, 1, x_679); -x_681 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__50; +x_681 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__48; lean_inc(x_530); x_682 = lean_name_mk_string(x_530, x_681); -x_683 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__51; +x_683 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__49; lean_inc(x_530); x_684 = lean_name_mk_string(x_530, x_683); -x_685 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__52; +x_685 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__50; lean_inc(x_519); x_686 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_686, 0, x_519); lean_ctor_set(x_686, 1, x_685); -x_687 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__53; +x_687 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__51; lean_inc(x_530); x_688 = lean_name_mk_string(x_530, x_687); -x_689 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__54; +x_689 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__52; lean_inc(x_519); x_690 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_690, 0, x_519); lean_ctor_set(x_690, 1, x_689); -x_691 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__57; +x_691 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__55; lean_inc(x_521); lean_inc(x_522); x_692 = l_Lean_addMacroScope(x_522, x_691, x_521); x_693 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__6; x_694 = lean_name_mk_string(x_4, x_693); lean_inc(x_694); -lean_ctor_set(x_6, 1, x_582); -lean_ctor_set(x_6, 0, x_694); +lean_ctor_set(x_7, 1, x_582); +lean_ctor_set(x_7, 0, x_694); x_695 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_695, 0, x_6); +lean_ctor_set(x_695, 0, x_7); lean_ctor_set(x_695, 1, x_582); -x_696 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__56; +x_696 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__54; lean_inc(x_519); x_697 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_697, 0, x_519); @@ -23662,2936 +24316,2933 @@ lean_inc(x_519); x_700 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_700, 0, x_519); lean_ctor_set(x_700, 1, x_693); -x_701 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__60; +x_701 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__56; +x_702 = lean_name_mk_string(x_5, x_701); +x_703 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__57; +x_704 = lean_name_mk_string(x_702, x_703); +x_705 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__58; lean_inc(x_519); -x_702 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_702, 0, x_519); -lean_ctor_set(x_702, 1, x_701); -x_703 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__18; +x_706 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_706, 0, x_519); +lean_ctor_set(x_706, 1, x_705); +x_707 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__16; lean_inc(x_519); -x_704 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_704, 0, x_519); -lean_ctor_set(x_704, 1, x_703); -x_705 = lean_array_push(x_601, x_704); -x_706 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__62; -x_707 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_707, 0, x_539); -lean_ctor_set(x_707, 1, x_706); -lean_ctor_set(x_707, 2, x_705); -x_708 = lean_array_push(x_660, x_702); -x_709 = lean_array_push(x_708, x_569); +x_708 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_708, 0, x_519); +lean_ctor_set(x_708, 1, x_707); +x_709 = lean_array_push(x_601, x_708); +x_710 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__60; +x_711 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_711, 0, x_539); +lean_ctor_set(x_711, 1, x_710); +lean_ctor_set(x_711, 2, x_709); +x_712 = lean_array_push(x_660, x_706); +x_713 = lean_array_push(x_712, x_569); lean_inc(x_642); -x_710 = lean_array_push(x_709, x_642); -lean_inc(x_710); -x_711 = lean_array_push(x_710, x_707); -x_712 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__59; -x_713 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_713, 0, x_539); -lean_ctor_set(x_713, 1, x_712); -lean_ctor_set(x_713, 2, x_711); -x_714 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__8; +x_714 = lean_array_push(x_713, x_642); +lean_inc(x_714); +x_715 = lean_array_push(x_714, x_711); +lean_inc(x_704); +x_716 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_716, 0, x_539); +lean_ctor_set(x_716, 1, x_704); +lean_ctor_set(x_716, 2, x_715); +x_717 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__8; lean_inc(x_519); -x_715 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_715, 0, x_519); -lean_ctor_set(x_715, 1, x_714); -x_716 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__37; +x_718 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_718, 0, x_519); +lean_ctor_set(x_718, 1, x_717); +x_719 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__37; lean_inc(x_699); -x_717 = lean_array_push(x_716, x_699); -x_718 = lean_array_push(x_717, x_700); +x_720 = lean_array_push(x_719, x_699); +x_721 = lean_array_push(x_720, x_700); lean_inc(x_612); -x_719 = lean_array_push(x_718, x_612); -lean_inc(x_719); -x_720 = lean_array_push(x_719, x_713); -lean_inc(x_715); -x_721 = lean_array_push(x_720, x_715); +x_722 = lean_array_push(x_721, x_612); +lean_inc(x_722); +x_723 = lean_array_push(x_722, x_716); +lean_inc(x_718); +x_724 = lean_array_push(x_723, x_718); lean_inc(x_694); -x_722 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_722, 0, x_539); -lean_ctor_set(x_722, 1, x_694); -lean_ctor_set(x_722, 2, x_721); -x_723 = lean_array_push(x_716, x_690); -x_724 = lean_array_push(x_723, x_697); +x_725 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_725, 0, x_539); +lean_ctor_set(x_725, 1, x_694); +lean_ctor_set(x_725, 2, x_724); +x_726 = lean_array_push(x_719, x_690); +x_727 = lean_array_push(x_726, x_697); lean_inc(x_686); -x_725 = lean_array_push(x_724, x_686); -lean_inc(x_725); -x_726 = lean_array_push(x_725, x_722); -lean_inc(x_715); -x_727 = lean_array_push(x_726, x_715); +x_728 = lean_array_push(x_727, x_686); +lean_inc(x_728); +x_729 = lean_array_push(x_728, x_725); +lean_inc(x_718); +x_730 = lean_array_push(x_729, x_718); lean_inc(x_688); -x_728 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_728, 0, x_539); -lean_ctor_set(x_728, 1, x_688); -lean_ctor_set(x_728, 2, x_727); -x_729 = lean_array_push(x_554, x_728); -x_730 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_730, 0, x_539); -lean_ctor_set(x_730, 1, x_556); -lean_ctor_set(x_730, 2, x_729); -x_731 = lean_array_push(x_554, x_730); -x_732 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_732, 0, x_539); -lean_ctor_set(x_732, 1, x_556); -lean_ctor_set(x_732, 2, x_731); -x_733 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__63; -lean_inc(x_530); -x_734 = lean_name_mk_string(x_530, x_733); -x_735 = l_Lean_Syntax_mkApp___closed__1; +x_731 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_731, 0, x_539); +lean_ctor_set(x_731, 1, x_688); +lean_ctor_set(x_731, 2, x_730); +x_732 = lean_array_push(x_554, x_731); +x_733 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_733, 0, x_539); +lean_ctor_set(x_733, 1, x_556); +lean_ctor_set(x_733, 2, x_732); +x_734 = lean_array_push(x_554, x_733); +x_735 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_735, 0, x_539); +lean_ctor_set(x_735, 1, x_556); +lean_ctor_set(x_735, 2, x_734); +x_736 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__61; lean_inc(x_530); -x_736 = lean_name_mk_string(x_530, x_735); -x_737 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__64; +x_737 = lean_name_mk_string(x_530, x_736); +x_738 = l_Lean_Syntax_mkApp___closed__1; lean_inc(x_530); -x_738 = lean_name_mk_string(x_530, x_737); -x_739 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__58; -x_740 = lean_name_mk_string(x_738, x_739); -x_741 = lean_array_push(x_710, x_569); -x_742 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_742, 0, x_539); -lean_ctor_set(x_742, 1, x_740); -lean_ctor_set(x_742, 2, x_741); -x_743 = lean_array_push(x_554, x_742); -x_744 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_744, 0, x_539); -lean_ctor_set(x_744, 1, x_556); -lean_ctor_set(x_744, 2, x_743); -x_745 = lean_array_push(x_547, x_5); -lean_inc(x_745); -x_746 = lean_array_push(x_745, x_744); -lean_inc(x_736); -x_747 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_747, 0, x_539); -lean_ctor_set(x_747, 1, x_736); -lean_ctor_set(x_747, 2, x_746); -lean_inc(x_719); -x_748 = lean_array_push(x_719, x_747); -lean_inc(x_715); -x_749 = lean_array_push(x_748, x_715); +x_739 = lean_name_mk_string(x_530, x_738); +x_740 = lean_array_push(x_714, x_569); +x_741 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_741, 0, x_539); +lean_ctor_set(x_741, 1, x_704); +lean_ctor_set(x_741, 2, x_740); +x_742 = lean_array_push(x_554, x_741); +x_743 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_743, 0, x_539); +lean_ctor_set(x_743, 1, x_556); +lean_ctor_set(x_743, 2, x_742); +x_744 = lean_array_push(x_547, x_6); +lean_inc(x_744); +x_745 = lean_array_push(x_744, x_743); +lean_inc(x_739); +x_746 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_746, 0, x_539); +lean_ctor_set(x_746, 1, x_739); +lean_ctor_set(x_746, 2, x_745); +lean_inc(x_722); +x_747 = lean_array_push(x_722, x_746); +lean_inc(x_718); +x_748 = lean_array_push(x_747, x_718); lean_inc(x_694); -x_750 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_750, 0, x_539); -lean_ctor_set(x_750, 1, x_694); -lean_ctor_set(x_750, 2, x_749); -lean_inc(x_725); -x_751 = lean_array_push(x_725, x_750); -lean_inc(x_715); -x_752 = lean_array_push(x_751, x_715); +x_749 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_749, 0, x_539); +lean_ctor_set(x_749, 1, x_694); +lean_ctor_set(x_749, 2, x_748); +lean_inc(x_728); +x_750 = lean_array_push(x_728, x_749); +lean_inc(x_718); +x_751 = lean_array_push(x_750, x_718); lean_inc(x_688); -x_753 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_753, 0, x_539); -lean_ctor_set(x_753, 1, x_688); -lean_ctor_set(x_753, 2, x_752); -x_754 = lean_array_push(x_554, x_753); -lean_inc(x_734); -x_755 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_755, 0, x_539); -lean_ctor_set(x_755, 1, x_734); -lean_ctor_set(x_755, 2, x_754); -x_756 = lean_array_push(x_547, x_755); -x_757 = lean_array_push(x_756, x_569); +x_752 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_752, 0, x_539); +lean_ctor_set(x_752, 1, x_688); +lean_ctor_set(x_752, 2, x_751); +x_753 = lean_array_push(x_554, x_752); +lean_inc(x_737); +x_754 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_754, 0, x_539); +lean_ctor_set(x_754, 1, x_737); +lean_ctor_set(x_754, 2, x_753); +x_755 = lean_array_push(x_547, x_754); +x_756 = lean_array_push(x_755, x_569); lean_inc(x_632); -x_758 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_758, 0, x_539); -lean_ctor_set(x_758, 1, x_632); -lean_ctor_set(x_758, 2, x_757); -x_759 = lean_array_push(x_554, x_758); -x_760 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_760, 0, x_539); -lean_ctor_set(x_760, 1, x_556); -lean_ctor_set(x_760, 2, x_759); -x_761 = lean_array_push(x_554, x_760); +x_757 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_757, 0, x_539); +lean_ctor_set(x_757, 1, x_632); +lean_ctor_set(x_757, 2, x_756); +x_758 = lean_array_push(x_554, x_757); +x_759 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_759, 0, x_539); +lean_ctor_set(x_759, 1, x_556); +lean_ctor_set(x_759, 2, x_758); +x_760 = lean_array_push(x_554, x_759); lean_inc(x_630); -x_762 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_762, 0, x_539); -lean_ctor_set(x_762, 1, x_630); -lean_ctor_set(x_762, 2, x_761); -x_763 = lean_array_push(x_660, x_686); -lean_inc(x_763); -x_764 = lean_array_push(x_763, x_732); +x_761 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_761, 0, x_539); +lean_ctor_set(x_761, 1, x_630); +lean_ctor_set(x_761, 2, x_760); +x_762 = lean_array_push(x_660, x_686); +lean_inc(x_762); +x_763 = lean_array_push(x_762, x_735); lean_inc(x_625); -x_765 = lean_array_push(x_764, x_625); -x_766 = lean_array_push(x_765, x_762); +x_764 = lean_array_push(x_763, x_625); +x_765 = lean_array_push(x_764, x_761); lean_inc(x_684); -x_767 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_767, 0, x_539); -lean_ctor_set(x_767, 1, x_684); -lean_ctor_set(x_767, 2, x_766); -x_768 = l_Lean_mkHole___closed__7; +x_766 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_766, 0, x_539); +lean_ctor_set(x_766, 1, x_684); +lean_ctor_set(x_766, 2, x_765); +x_767 = l_Lean_mkHole___closed__7; lean_inc(x_530); -x_769 = lean_name_mk_string(x_530, x_768); -x_770 = l_Lean_Name_appendIndexAfter___closed__1; +x_768 = lean_name_mk_string(x_530, x_767); +x_769 = l_Lean_Name_appendIndexAfter___closed__1; lean_inc(x_519); -x_771 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_771, 0, x_519); -lean_ctor_set(x_771, 1, x_770); -x_772 = lean_array_push(x_554, x_771); -x_773 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_773, 0, x_539); -lean_ctor_set(x_773, 1, x_769); -lean_ctor_set(x_773, 2, x_772); -x_774 = lean_array_push(x_554, x_773); -x_775 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_775, 0, x_539); -lean_ctor_set(x_775, 1, x_556); -lean_ctor_set(x_775, 2, x_774); -x_776 = lean_array_push(x_554, x_775); -x_777 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_777, 0, x_539); -lean_ctor_set(x_777, 1, x_556); -lean_ctor_set(x_777, 2, x_776); -x_778 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__11; +x_770 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_770, 0, x_519); +lean_ctor_set(x_770, 1, x_769); +x_771 = lean_array_push(x_554, x_770); +x_772 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_772, 0, x_539); +lean_ctor_set(x_772, 1, x_768); +lean_ctor_set(x_772, 2, x_771); +x_773 = lean_array_push(x_554, x_772); +x_774 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_774, 0, x_539); +lean_ctor_set(x_774, 1, x_556); +lean_ctor_set(x_774, 2, x_773); +x_775 = lean_array_push(x_554, x_774); +x_776 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_776, 0, x_539); +lean_ctor_set(x_776, 1, x_556); +lean_ctor_set(x_776, 2, x_775); +x_777 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__11; lean_inc(x_519); -x_779 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_779, 0, x_519); -lean_ctor_set(x_779, 1, x_778); -x_780 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__35; +x_778 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_778, 0, x_519); +lean_ctor_set(x_778, 1, x_777); +x_779 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__35; lean_inc(x_519); -x_781 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_781, 0, x_519); -lean_ctor_set(x_781, 1, x_780); -lean_inc(x_779); -x_782 = lean_array_push(x_547, x_779); -lean_inc(x_781); -x_783 = lean_array_push(x_782, x_781); -x_784 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__68; -x_785 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_785, 0, x_539); -lean_ctor_set(x_785, 1, x_784); -lean_ctor_set(x_785, 2, x_783); -x_786 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__9; +x_780 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_780, 0, x_519); +lean_ctor_set(x_780, 1, x_779); +lean_inc(x_778); +x_781 = lean_array_push(x_547, x_778); +lean_inc(x_780); +x_782 = lean_array_push(x_781, x_780); +x_783 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__65; +x_784 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_784, 0, x_539); +lean_ctor_set(x_784, 1, x_783); +lean_ctor_set(x_784, 2, x_782); +x_785 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__9; lean_inc(x_530); -x_787 = lean_name_mk_string(x_530, x_786); -x_788 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__31; +x_786 = lean_name_mk_string(x_530, x_785); +x_787 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__31; lean_inc(x_530); -x_789 = lean_name_mk_string(x_530, x_788); -x_790 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_790, 0, x_539); -lean_ctor_set(x_790, 1, x_789); -lean_ctor_set(x_790, 2, x_540); -x_791 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__36; -x_792 = lean_array_push(x_791, x_779); +x_788 = lean_name_mk_string(x_530, x_787); +x_789 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_789, 0, x_539); +lean_ctor_set(x_789, 1, x_788); +lean_ctor_set(x_789, 2, x_540); +x_790 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__36; +x_791 = lean_array_push(x_790, x_778); +x_792 = lean_array_push(x_791, x_569); x_793 = lean_array_push(x_792, x_569); -x_794 = lean_array_push(x_793, x_569); -x_795 = lean_array_push(x_794, x_790); -x_796 = lean_array_push(x_795, x_569); -x_797 = lean_array_push(x_796, x_781); -x_798 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_798, 0, x_539); -lean_ctor_set(x_798, 1, x_787); -lean_ctor_set(x_798, 2, x_797); -x_799 = lean_array_push(x_547, x_785); -x_800 = lean_array_push(x_799, x_798); -x_801 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__66; -x_802 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_802, 0, x_539); -lean_ctor_set(x_802, 1, x_801); -lean_ctor_set(x_802, 2, x_800); -x_803 = lean_array_push(x_554, x_802); -x_804 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_804, 0, x_539); -lean_ctor_set(x_804, 1, x_556); -lean_ctor_set(x_804, 2, x_803); -x_805 = lean_array_push(x_745, x_804); -lean_inc(x_736); -x_806 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_806, 0, x_539); -lean_ctor_set(x_806, 1, x_736); -lean_ctor_set(x_806, 2, x_805); -x_807 = lean_array_push(x_719, x_806); -lean_inc(x_715); -x_808 = lean_array_push(x_807, x_715); -x_809 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_809, 0, x_539); -lean_ctor_set(x_809, 1, x_694); -lean_ctor_set(x_809, 2, x_808); -x_810 = lean_array_push(x_725, x_809); -lean_inc(x_715); -x_811 = lean_array_push(x_810, x_715); -x_812 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_812, 0, x_539); -lean_ctor_set(x_812, 1, x_688); -lean_ctor_set(x_812, 2, x_811); -x_813 = lean_array_push(x_554, x_812); -x_814 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_814, 0, x_539); -lean_ctor_set(x_814, 1, x_734); -lean_ctor_set(x_814, 2, x_813); -x_815 = lean_array_push(x_547, x_814); -x_816 = lean_array_push(x_815, x_569); +x_794 = lean_array_push(x_793, x_789); +x_795 = lean_array_push(x_794, x_569); +x_796 = lean_array_push(x_795, x_780); +x_797 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_797, 0, x_539); +lean_ctor_set(x_797, 1, x_786); +lean_ctor_set(x_797, 2, x_796); +x_798 = lean_array_push(x_547, x_784); +x_799 = lean_array_push(x_798, x_797); +x_800 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__63; +x_801 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_801, 0, x_539); +lean_ctor_set(x_801, 1, x_800); +lean_ctor_set(x_801, 2, x_799); +x_802 = lean_array_push(x_554, x_801); +x_803 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_803, 0, x_539); +lean_ctor_set(x_803, 1, x_556); +lean_ctor_set(x_803, 2, x_802); +x_804 = lean_array_push(x_744, x_803); +lean_inc(x_739); +x_805 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_805, 0, x_539); +lean_ctor_set(x_805, 1, x_739); +lean_ctor_set(x_805, 2, x_804); +x_806 = lean_array_push(x_722, x_805); +lean_inc(x_718); +x_807 = lean_array_push(x_806, x_718); +x_808 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_808, 0, x_539); +lean_ctor_set(x_808, 1, x_694); +lean_ctor_set(x_808, 2, x_807); +x_809 = lean_array_push(x_728, x_808); +lean_inc(x_718); +x_810 = lean_array_push(x_809, x_718); +x_811 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_811, 0, x_539); +lean_ctor_set(x_811, 1, x_688); +lean_ctor_set(x_811, 2, x_810); +x_812 = lean_array_push(x_554, x_811); +x_813 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_813, 0, x_539); +lean_ctor_set(x_813, 1, x_737); +lean_ctor_set(x_813, 2, x_812); +x_814 = lean_array_push(x_547, x_813); +x_815 = lean_array_push(x_814, x_569); lean_inc(x_632); -x_817 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_817, 0, x_539); -lean_ctor_set(x_817, 1, x_632); -lean_ctor_set(x_817, 2, x_816); -x_818 = lean_array_push(x_554, x_817); -x_819 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_819, 0, x_539); -lean_ctor_set(x_819, 1, x_556); -lean_ctor_set(x_819, 2, x_818); -x_820 = lean_array_push(x_554, x_819); +x_816 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_816, 0, x_539); +lean_ctor_set(x_816, 1, x_632); +lean_ctor_set(x_816, 2, x_815); +x_817 = lean_array_push(x_554, x_816); +x_818 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_818, 0, x_539); +lean_ctor_set(x_818, 1, x_556); +lean_ctor_set(x_818, 2, x_817); +x_819 = lean_array_push(x_554, x_818); lean_inc(x_630); -x_821 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_821, 0, x_539); -lean_ctor_set(x_821, 1, x_630); -lean_ctor_set(x_821, 2, x_820); -x_822 = lean_array_push(x_763, x_777); +x_820 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_820, 0, x_539); +lean_ctor_set(x_820, 1, x_630); +lean_ctor_set(x_820, 2, x_819); +x_821 = lean_array_push(x_762, x_776); lean_inc(x_625); -x_823 = lean_array_push(x_822, x_625); -x_824 = lean_array_push(x_823, x_821); -x_825 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_825, 0, x_539); -lean_ctor_set(x_825, 1, x_684); -lean_ctor_set(x_825, 2, x_824); -x_826 = lean_array_push(x_547, x_767); -x_827 = lean_array_push(x_826, x_825); -x_828 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_828, 0, x_539); -lean_ctor_set(x_828, 1, x_556); -lean_ctor_set(x_828, 2, x_827); -x_829 = lean_array_push(x_554, x_828); -x_830 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_830, 0, x_539); -lean_ctor_set(x_830, 1, x_682); -lean_ctor_set(x_830, 2, x_829); -x_831 = lean_array_push(x_791, x_648); +x_822 = lean_array_push(x_821, x_625); +x_823 = lean_array_push(x_822, x_820); +x_824 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_824, 0, x_539); +lean_ctor_set(x_824, 1, x_684); +lean_ctor_set(x_824, 2, x_823); +x_825 = lean_array_push(x_547, x_766); +x_826 = lean_array_push(x_825, x_824); +x_827 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_827, 0, x_539); +lean_ctor_set(x_827, 1, x_556); +lean_ctor_set(x_827, 2, x_826); +x_828 = lean_array_push(x_554, x_827); +x_829 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_829, 0, x_539); +lean_ctor_set(x_829, 1, x_682); +lean_ctor_set(x_829, 2, x_828); +x_830 = lean_array_push(x_790, x_648); +x_831 = lean_array_push(x_830, x_569); x_832 = lean_array_push(x_831, x_569); -x_833 = lean_array_push(x_832, x_569); -x_834 = lean_array_push(x_833, x_678); -x_835 = lean_array_push(x_834, x_680); -x_836 = lean_array_push(x_835, x_830); -x_837 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_837, 0, x_539); -lean_ctor_set(x_837, 1, x_646); -lean_ctor_set(x_837, 2, x_836); +x_833 = lean_array_push(x_832, x_678); +x_834 = lean_array_push(x_833, x_680); +x_835 = lean_array_push(x_834, x_829); +x_836 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_836, 0, x_539); +lean_ctor_set(x_836, 1, x_646); +lean_ctor_set(x_836, 2, x_835); lean_inc(x_642); -x_838 = lean_array_push(x_660, x_642); -x_839 = lean_array_push(x_838, x_569); -x_840 = lean_array_push(x_839, x_644); -x_841 = lean_array_push(x_840, x_837); -x_842 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_842, 0, x_539); -lean_ctor_set(x_842, 1, x_638); -lean_ctor_set(x_842, 2, x_841); -x_843 = lean_array_push(x_560, x_636); -x_844 = lean_array_push(x_843, x_569); -lean_inc(x_844); -x_845 = lean_array_push(x_844, x_842); -x_846 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_846, 0, x_539); -lean_ctor_set(x_846, 1, x_634); -lean_ctor_set(x_846, 2, x_845); -x_847 = lean_array_push(x_547, x_846); -x_848 = lean_array_push(x_847, x_569); +x_837 = lean_array_push(x_660, x_642); +x_838 = lean_array_push(x_837, x_569); +x_839 = lean_array_push(x_838, x_644); +x_840 = lean_array_push(x_839, x_836); +x_841 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_841, 0, x_539); +lean_ctor_set(x_841, 1, x_638); +lean_ctor_set(x_841, 2, x_840); +x_842 = lean_array_push(x_560, x_636); +x_843 = lean_array_push(x_842, x_569); +lean_inc(x_843); +x_844 = lean_array_push(x_843, x_841); +x_845 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_845, 0, x_539); +lean_ctor_set(x_845, 1, x_634); +lean_ctor_set(x_845, 2, x_844); +x_846 = lean_array_push(x_547, x_845); +x_847 = lean_array_push(x_846, x_569); lean_inc(x_632); -x_849 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_849, 0, x_539); -lean_ctor_set(x_849, 1, x_632); -lean_ctor_set(x_849, 2, x_848); -x_850 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__69; +x_848 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_848, 0, x_539); +lean_ctor_set(x_848, 1, x_632); +lean_ctor_set(x_848, 2, x_847); +x_849 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__66; lean_inc(x_530); -x_851 = lean_name_mk_string(x_530, x_850); -x_852 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__70; +x_850 = lean_name_mk_string(x_530, x_849); +x_851 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__67; lean_inc(x_530); -x_853 = lean_name_mk_string(x_530, x_852); -x_854 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__71; +x_852 = lean_name_mk_string(x_530, x_851); +x_853 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__68; lean_inc(x_530); -x_855 = lean_name_mk_string(x_530, x_854); -x_856 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__76; +x_854 = lean_name_mk_string(x_530, x_853); +x_855 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__73; lean_inc(x_521); lean_inc(x_522); -x_857 = l_Lean_addMacroScope(x_522, x_856, x_521); -x_858 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__74; +x_856 = l_Lean_addMacroScope(x_522, x_855, x_521); +x_857 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__71; lean_inc(x_519); -x_859 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_859, 0, x_519); -lean_ctor_set(x_859, 1, x_858); -lean_ctor_set(x_859, 2, x_857); -lean_ctor_set(x_859, 3, x_582); -x_860 = lean_array_push(x_554, x_12); -x_861 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_861, 0, x_539); -lean_ctor_set(x_861, 1, x_556); -lean_ctor_set(x_861, 2, x_860); -x_862 = lean_array_push(x_547, x_859); -x_863 = lean_array_push(x_862, x_861); -lean_inc(x_736); -x_864 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_864, 0, x_539); -lean_ctor_set(x_864, 1, x_736); -lean_ctor_set(x_864, 2, x_863); -x_865 = lean_array_push(x_716, x_621); +x_858 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_858, 0, x_519); +lean_ctor_set(x_858, 1, x_857); +lean_ctor_set(x_858, 2, x_856); +lean_ctor_set(x_858, 3, x_582); +x_859 = lean_array_push(x_554, x_13); +x_860 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_860, 0, x_539); +lean_ctor_set(x_860, 1, x_556); +lean_ctor_set(x_860, 2, x_859); +x_861 = lean_array_push(x_547, x_858); +x_862 = lean_array_push(x_861, x_860); +lean_inc(x_739); +x_863 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_863, 0, x_539); +lean_ctor_set(x_863, 1, x_739); +lean_ctor_set(x_863, 2, x_862); +x_864 = lean_array_push(x_719, x_621); +x_865 = lean_array_push(x_864, x_569); x_866 = lean_array_push(x_865, x_569); -x_867 = lean_array_push(x_866, x_569); lean_inc(x_612); -x_868 = lean_array_push(x_867, x_612); -lean_inc(x_868); -x_869 = lean_array_push(x_868, x_864); -lean_inc(x_855); -x_870 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_870, 0, x_539); -lean_ctor_set(x_870, 1, x_855); -lean_ctor_set(x_870, 2, x_869); -x_871 = lean_array_push(x_554, x_870); -lean_inc(x_853); -x_872 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_872, 0, x_539); -lean_ctor_set(x_872, 1, x_853); -lean_ctor_set(x_872, 2, x_871); -lean_inc(x_844); -x_873 = lean_array_push(x_844, x_872); -lean_inc(x_851); -x_874 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_874, 0, x_539); -lean_ctor_set(x_874, 1, x_851); -lean_ctor_set(x_874, 2, x_873); -x_875 = lean_array_push(x_547, x_874); -x_876 = lean_array_push(x_875, x_569); +x_867 = lean_array_push(x_866, x_612); +lean_inc(x_867); +x_868 = lean_array_push(x_867, x_863); +lean_inc(x_854); +x_869 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_869, 0, x_539); +lean_ctor_set(x_869, 1, x_854); +lean_ctor_set(x_869, 2, x_868); +x_870 = lean_array_push(x_554, x_869); +lean_inc(x_852); +x_871 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_871, 0, x_539); +lean_ctor_set(x_871, 1, x_852); +lean_ctor_set(x_871, 2, x_870); +lean_inc(x_843); +x_872 = lean_array_push(x_843, x_871); +lean_inc(x_850); +x_873 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_873, 0, x_539); +lean_ctor_set(x_873, 1, x_850); +lean_ctor_set(x_873, 2, x_872); +x_874 = lean_array_push(x_547, x_873); +x_875 = lean_array_push(x_874, x_569); lean_inc(x_632); -x_877 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_877, 0, x_539); -lean_ctor_set(x_877, 1, x_632); -lean_ctor_set(x_877, 2, x_876); -x_878 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__81; +x_876 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_876, 0, x_539); +lean_ctor_set(x_876, 1, x_632); +lean_ctor_set(x_876, 2, x_875); +x_877 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__78; lean_inc(x_521); lean_inc(x_522); -x_879 = l_Lean_addMacroScope(x_522, x_878, x_521); -x_880 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__79; +x_878 = l_Lean_addMacroScope(x_522, x_877, x_521); +x_879 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__76; lean_inc(x_519); -x_881 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_881, 0, x_519); -lean_ctor_set(x_881, 1, x_880); -lean_ctor_set(x_881, 2, x_879); -lean_ctor_set(x_881, 3, x_582); -x_882 = l_Lean_Syntax_expandInterpolatedStr___closed__2; +x_880 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_880, 0, x_519); +lean_ctor_set(x_880, 1, x_879); +lean_ctor_set(x_880, 2, x_878); +lean_ctor_set(x_880, 3, x_582); +x_881 = l_Lean_TSyntax_expandInterpolatedStr___closed__2; lean_inc(x_530); -x_883 = lean_name_mk_string(x_530, x_882); -x_884 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__85; +x_882 = lean_name_mk_string(x_530, x_881); +x_883 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__82; lean_inc(x_521); lean_inc(x_522); -x_885 = l_Lean_addMacroScope(x_522, x_884, x_521); -x_886 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__82; +x_884 = l_Lean_addMacroScope(x_522, x_883, x_521); +x_885 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__79; lean_inc(x_3); -x_887 = lean_name_mk_string(x_3, x_886); -x_888 = lean_alloc_ctor(0, 2, 0); +x_886 = lean_name_mk_string(x_3, x_885); +x_887 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_887, 0, x_886); +lean_ctor_set(x_887, 1, x_582); +x_888 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_888, 0, x_887); lean_ctor_set(x_888, 1, x_582); -x_889 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_889, 0, x_888); -lean_ctor_set(x_889, 1, x_582); -x_890 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__84; +x_889 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__81; lean_inc(x_519); -x_891 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_891, 0, x_519); -lean_ctor_set(x_891, 1, x_890); -lean_ctor_set(x_891, 2, x_885); -lean_ctor_set(x_891, 3, x_889); +x_890 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_890, 0, x_519); +lean_ctor_set(x_890, 1, x_889); +lean_ctor_set(x_890, 2, x_884); +lean_ctor_set(x_890, 3, x_888); lean_inc(x_669); -x_892 = lean_array_push(x_662, x_669); +x_891 = lean_array_push(x_662, x_669); lean_inc(x_559); -x_893 = lean_array_push(x_892, x_559); -x_894 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_894, 0, x_539); -lean_ctor_set(x_894, 1, x_652); -lean_ctor_set(x_894, 2, x_893); -x_895 = lean_array_push(x_547, x_894); -x_896 = lean_array_push(x_895, x_13); -x_897 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_897, 0, x_539); -lean_ctor_set(x_897, 1, x_556); -lean_ctor_set(x_897, 2, x_896); -x_898 = lean_array_push(x_547, x_891); -x_899 = lean_array_push(x_898, x_897); -lean_inc(x_736); -x_900 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_900, 0, x_539); -lean_ctor_set(x_900, 1, x_736); -lean_ctor_set(x_900, 2, x_899); -x_901 = lean_array_push(x_547, x_900); -x_902 = lean_array_push(x_901, x_569); -x_903 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_903, 0, x_539); -lean_ctor_set(x_903, 1, x_556); -lean_ctor_set(x_903, 2, x_902); -x_904 = lean_array_push(x_560, x_699); -lean_inc(x_904); -x_905 = lean_array_push(x_904, x_903); -lean_inc(x_715); -x_906 = lean_array_push(x_905, x_715); -lean_inc(x_883); -x_907 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_907, 0, x_539); -lean_ctor_set(x_907, 1, x_883); -lean_ctor_set(x_907, 2, x_906); -x_908 = lean_array_push(x_547, x_669); -x_909 = lean_array_push(x_908, x_907); -x_910 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_910, 0, x_539); -lean_ctor_set(x_910, 1, x_556); -lean_ctor_set(x_910, 2, x_909); -x_911 = lean_array_push(x_547, x_881); -lean_inc(x_911); -x_912 = lean_array_push(x_911, x_910); -lean_inc(x_736); -x_913 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_913, 0, x_539); -lean_ctor_set(x_913, 1, x_736); -lean_ctor_set(x_913, 2, x_912); -x_914 = lean_array_push(x_868, x_913); -lean_inc(x_855); -x_915 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_915, 0, x_539); -lean_ctor_set(x_915, 1, x_855); -lean_ctor_set(x_915, 2, x_914); -x_916 = lean_array_push(x_554, x_915); -lean_inc(x_853); -x_917 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_917, 0, x_539); -lean_ctor_set(x_917, 1, x_853); -lean_ctor_set(x_917, 2, x_916); -lean_inc(x_844); -x_918 = lean_array_push(x_844, x_917); -lean_inc(x_851); -x_919 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_919, 0, x_539); -lean_ctor_set(x_919, 1, x_851); -lean_ctor_set(x_919, 2, x_918); -x_920 = lean_array_push(x_547, x_919); -x_921 = lean_array_push(x_920, x_569); +x_892 = lean_array_push(x_891, x_559); +x_893 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_893, 0, x_539); +lean_ctor_set(x_893, 1, x_652); +lean_ctor_set(x_893, 2, x_892); +x_894 = lean_array_push(x_547, x_893); +x_895 = lean_array_push(x_894, x_14); +x_896 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_896, 0, x_539); +lean_ctor_set(x_896, 1, x_556); +lean_ctor_set(x_896, 2, x_895); +x_897 = lean_array_push(x_547, x_890); +x_898 = lean_array_push(x_897, x_896); +lean_inc(x_739); +x_899 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_899, 0, x_539); +lean_ctor_set(x_899, 1, x_739); +lean_ctor_set(x_899, 2, x_898); +x_900 = lean_array_push(x_547, x_899); +x_901 = lean_array_push(x_900, x_569); +x_902 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_902, 0, x_539); +lean_ctor_set(x_902, 1, x_556); +lean_ctor_set(x_902, 2, x_901); +x_903 = lean_array_push(x_560, x_699); +lean_inc(x_903); +x_904 = lean_array_push(x_903, x_902); +lean_inc(x_718); +x_905 = lean_array_push(x_904, x_718); +lean_inc(x_882); +x_906 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_906, 0, x_539); +lean_ctor_set(x_906, 1, x_882); +lean_ctor_set(x_906, 2, x_905); +x_907 = lean_array_push(x_547, x_669); +x_908 = lean_array_push(x_907, x_906); +x_909 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_909, 0, x_539); +lean_ctor_set(x_909, 1, x_556); +lean_ctor_set(x_909, 2, x_908); +x_910 = lean_array_push(x_547, x_880); +lean_inc(x_910); +x_911 = lean_array_push(x_910, x_909); +lean_inc(x_739); +x_912 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_912, 0, x_539); +lean_ctor_set(x_912, 1, x_739); +lean_ctor_set(x_912, 2, x_911); +x_913 = lean_array_push(x_867, x_912); +lean_inc(x_854); +x_914 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_914, 0, x_539); +lean_ctor_set(x_914, 1, x_854); +lean_ctor_set(x_914, 2, x_913); +x_915 = lean_array_push(x_554, x_914); +lean_inc(x_852); +x_916 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_916, 0, x_539); +lean_ctor_set(x_916, 1, x_852); +lean_ctor_set(x_916, 2, x_915); +lean_inc(x_843); +x_917 = lean_array_push(x_843, x_916); +lean_inc(x_850); +x_918 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_918, 0, x_539); +lean_ctor_set(x_918, 1, x_850); +lean_ctor_set(x_918, 2, x_917); +x_919 = lean_array_push(x_547, x_918); +x_920 = lean_array_push(x_919, x_569); lean_inc(x_632); -x_922 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_922, 0, x_539); -lean_ctor_set(x_922, 1, x_632); -lean_ctor_set(x_922, 2, x_921); -x_923 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__89; +x_921 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_921, 0, x_539); +lean_ctor_set(x_921, 1, x_632); +lean_ctor_set(x_921, 2, x_920); +x_922 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__86; lean_inc(x_521); lean_inc(x_522); -x_924 = l_Lean_addMacroScope(x_522, x_923, x_521); -x_925 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__88; +x_923 = l_Lean_addMacroScope(x_522, x_922, x_521); +x_924 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__85; lean_inc(x_519); -x_926 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_926, 0, x_519); -lean_ctor_set(x_926, 1, x_925); -lean_ctor_set(x_926, 2, x_924); -lean_ctor_set(x_926, 3, x_582); -x_927 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__93; -x_928 = l_Lean_addMacroScope(x_522, x_927, x_521); -x_929 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__90; -x_930 = lean_name_mk_string(x_3, x_929); -x_931 = lean_alloc_ctor(0, 2, 0); +x_925 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_925, 0, x_519); +lean_ctor_set(x_925, 1, x_924); +lean_ctor_set(x_925, 2, x_923); +lean_ctor_set(x_925, 3, x_582); +x_926 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__90; +x_927 = l_Lean_addMacroScope(x_522, x_926, x_521); +x_928 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__87; +x_929 = lean_name_mk_string(x_3, x_928); +x_930 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_930, 0, x_929); +lean_ctor_set(x_930, 1, x_582); +x_931 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_931, 0, x_930); lean_ctor_set(x_931, 1, x_582); -x_932 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_932, 0, x_931); -lean_ctor_set(x_932, 1, x_582); -x_933 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__92; +x_932 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__89; lean_inc(x_519); -x_934 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_934, 0, x_519); -lean_ctor_set(x_934, 1, x_933); -lean_ctor_set(x_934, 2, x_928); -lean_ctor_set(x_934, 3, x_932); -x_935 = l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__12; +x_933 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_933, 0, x_519); +lean_ctor_set(x_933, 1, x_932); +lean_ctor_set(x_933, 2, x_927); +lean_ctor_set(x_933, 3, x_931); +x_934 = l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__12; lean_inc(x_519); -x_936 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_936, 0, x_519); -lean_ctor_set(x_936, 1, x_935); -x_937 = lean_array_push(x_554, x_642); -x_938 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_938, 0, x_539); -lean_ctor_set(x_938, 1, x_556); -lean_ctor_set(x_938, 2, x_937); -x_939 = lean_array_push(x_560, x_936); -x_940 = lean_array_push(x_939, x_938); -x_941 = lean_array_push(x_940, x_559); -x_942 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__95; -x_943 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_943, 0, x_539); -lean_ctor_set(x_943, 1, x_942); -lean_ctor_set(x_943, 2, x_941); -x_944 = lean_array_push(x_554, x_943); -x_945 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_945, 0, x_539); -lean_ctor_set(x_945, 1, x_556); -lean_ctor_set(x_945, 2, x_944); -x_946 = lean_array_push(x_547, x_934); -x_947 = lean_array_push(x_946, x_945); -lean_inc(x_736); -x_948 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_948, 0, x_539); -lean_ctor_set(x_948, 1, x_736); -lean_ctor_set(x_948, 2, x_947); -x_949 = lean_array_push(x_547, x_948); -x_950 = lean_array_push(x_949, x_569); -x_951 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_951, 0, x_539); -lean_ctor_set(x_951, 1, x_556); -lean_ctor_set(x_951, 2, x_950); -x_952 = lean_array_push(x_904, x_951); -x_953 = lean_array_push(x_952, x_715); -x_954 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_954, 0, x_539); -lean_ctor_set(x_954, 1, x_883); -lean_ctor_set(x_954, 2, x_953); -x_955 = lean_array_push(x_547, x_659); -x_956 = lean_array_push(x_955, x_954); -x_957 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_957, 0, x_539); -lean_ctor_set(x_957, 1, x_556); -lean_ctor_set(x_957, 2, x_956); -x_958 = lean_array_push(x_911, x_957); -x_959 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_959, 0, x_539); -lean_ctor_set(x_959, 1, x_736); -lean_ctor_set(x_959, 2, x_958); -lean_inc(x_926); -x_960 = lean_array_push(x_716, x_926); +x_935 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_935, 0, x_519); +lean_ctor_set(x_935, 1, x_934); +x_936 = lean_array_push(x_554, x_642); +x_937 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_937, 0, x_539); +lean_ctor_set(x_937, 1, x_556); +lean_ctor_set(x_937, 2, x_936); +x_938 = lean_array_push(x_560, x_935); +x_939 = lean_array_push(x_938, x_937); +x_940 = lean_array_push(x_939, x_559); +x_941 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__92; +x_942 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_942, 0, x_539); +lean_ctor_set(x_942, 1, x_941); +lean_ctor_set(x_942, 2, x_940); +x_943 = lean_array_push(x_554, x_942); +x_944 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_944, 0, x_539); +lean_ctor_set(x_944, 1, x_556); +lean_ctor_set(x_944, 2, x_943); +x_945 = lean_array_push(x_547, x_933); +x_946 = lean_array_push(x_945, x_944); +lean_inc(x_739); +x_947 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_947, 0, x_539); +lean_ctor_set(x_947, 1, x_739); +lean_ctor_set(x_947, 2, x_946); +x_948 = lean_array_push(x_547, x_947); +x_949 = lean_array_push(x_948, x_569); +x_950 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_950, 0, x_539); +lean_ctor_set(x_950, 1, x_556); +lean_ctor_set(x_950, 2, x_949); +x_951 = lean_array_push(x_903, x_950); +x_952 = lean_array_push(x_951, x_718); +x_953 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_953, 0, x_539); +lean_ctor_set(x_953, 1, x_882); +lean_ctor_set(x_953, 2, x_952); +x_954 = lean_array_push(x_547, x_659); +x_955 = lean_array_push(x_954, x_953); +x_956 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_956, 0, x_539); +lean_ctor_set(x_956, 1, x_556); +lean_ctor_set(x_956, 2, x_955); +x_957 = lean_array_push(x_910, x_956); +x_958 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_958, 0, x_539); +lean_ctor_set(x_958, 1, x_739); +lean_ctor_set(x_958, 2, x_957); +lean_inc(x_925); +x_959 = lean_array_push(x_719, x_925); +x_960 = lean_array_push(x_959, x_569); x_961 = lean_array_push(x_960, x_569); -x_962 = lean_array_push(x_961, x_569); lean_inc(x_612); -x_963 = lean_array_push(x_962, x_612); -x_964 = lean_array_push(x_963, x_959); -x_965 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_965, 0, x_539); -lean_ctor_set(x_965, 1, x_855); -lean_ctor_set(x_965, 2, x_964); -x_966 = lean_array_push(x_554, x_965); -x_967 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_967, 0, x_539); -lean_ctor_set(x_967, 1, x_853); -lean_ctor_set(x_967, 2, x_966); -x_968 = lean_array_push(x_844, x_967); -x_969 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_969, 0, x_539); -lean_ctor_set(x_969, 1, x_851); -lean_ctor_set(x_969, 2, x_968); -x_970 = lean_array_push(x_547, x_969); -x_971 = lean_array_push(x_970, x_569); +x_962 = lean_array_push(x_961, x_612); +x_963 = lean_array_push(x_962, x_958); +x_964 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_964, 0, x_539); +lean_ctor_set(x_964, 1, x_854); +lean_ctor_set(x_964, 2, x_963); +x_965 = lean_array_push(x_554, x_964); +x_966 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_966, 0, x_539); +lean_ctor_set(x_966, 1, x_852); +lean_ctor_set(x_966, 2, x_965); +x_967 = lean_array_push(x_843, x_966); +x_968 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_968, 0, x_539); +lean_ctor_set(x_968, 1, x_850); +lean_ctor_set(x_968, 2, x_967); +x_969 = lean_array_push(x_547, x_968); +x_970 = lean_array_push(x_969, x_569); lean_inc(x_632); -x_972 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_972, 0, x_539); -lean_ctor_set(x_972, 1, x_632); -lean_ctor_set(x_972, 2, x_971); -x_973 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__96; -x_974 = lean_name_mk_string(x_530, x_973); -x_975 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__97; -x_976 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_976, 0, x_519); -lean_ctor_set(x_976, 1, x_975); -x_977 = lean_array_push(x_554, x_926); -x_978 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_978, 0, x_539); -lean_ctor_set(x_978, 1, x_556); -lean_ctor_set(x_978, 2, x_977); -x_979 = lean_array_push(x_547, x_976); -x_980 = lean_array_push(x_979, x_978); -x_981 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_981, 0, x_539); -lean_ctor_set(x_981, 1, x_974); -lean_ctor_set(x_981, 2, x_980); -x_982 = lean_array_push(x_547, x_981); -x_983 = lean_array_push(x_982, x_569); -x_984 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_984, 0, x_539); -lean_ctor_set(x_984, 1, x_632); -lean_ctor_set(x_984, 2, x_983); -x_985 = lean_array_push(x_716, x_849); -x_986 = lean_array_push(x_985, x_877); -x_987 = lean_array_push(x_986, x_922); -x_988 = lean_array_push(x_987, x_972); -x_989 = lean_array_push(x_988, x_984); -x_990 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_990, 0, x_539); -lean_ctor_set(x_990, 1, x_556); -lean_ctor_set(x_990, 2, x_989); -x_991 = lean_array_push(x_554, x_990); -x_992 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_992, 0, x_539); -lean_ctor_set(x_992, 1, x_630); -lean_ctor_set(x_992, 2, x_991); -x_993 = lean_array_push(x_547, x_628); -x_994 = lean_array_push(x_993, x_992); -x_995 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_995, 0, x_539); -lean_ctor_set(x_995, 1, x_627); -lean_ctor_set(x_995, 2, x_994); -x_996 = lean_array_push(x_560, x_623); -x_997 = lean_array_push(x_996, x_625); -x_998 = lean_array_push(x_997, x_995); -x_999 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_999, 0, x_539); -lean_ctor_set(x_999, 1, x_617); -lean_ctor_set(x_999, 2, x_998); -x_1000 = lean_array_push(x_547, x_615); -x_1001 = lean_array_push(x_1000, x_999); -x_1002 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1002, 0, x_539); -lean_ctor_set(x_1002, 1, x_614); -lean_ctor_set(x_1002, 2, x_1001); -x_1003 = lean_array_push(x_560, x_612); -x_1004 = lean_array_push(x_1003, x_1002); -x_1005 = lean_array_push(x_1004, x_569); -x_1006 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1006, 0, x_539); -lean_ctor_set(x_1006, 1, x_610); -lean_ctor_set(x_1006, 2, x_1005); -x_1007 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__98; -x_1008 = lean_array_push(x_1007, x_577); -x_1009 = lean_array_push(x_1008, x_587); -x_1010 = lean_array_push(x_1009, x_608); -x_1011 = lean_array_push(x_1010, x_1006); +x_971 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_971, 0, x_539); +lean_ctor_set(x_971, 1, x_632); +lean_ctor_set(x_971, 2, x_970); +x_972 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__93; +x_973 = lean_name_mk_string(x_530, x_972); +x_974 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__94; +x_975 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_975, 0, x_519); +lean_ctor_set(x_975, 1, x_974); +x_976 = lean_array_push(x_554, x_925); +x_977 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_977, 0, x_539); +lean_ctor_set(x_977, 1, x_556); +lean_ctor_set(x_977, 2, x_976); +x_978 = lean_array_push(x_547, x_975); +x_979 = lean_array_push(x_978, x_977); +x_980 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_980, 0, x_539); +lean_ctor_set(x_980, 1, x_973); +lean_ctor_set(x_980, 2, x_979); +x_981 = lean_array_push(x_547, x_980); +x_982 = lean_array_push(x_981, x_569); +x_983 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_983, 0, x_539); +lean_ctor_set(x_983, 1, x_632); +lean_ctor_set(x_983, 2, x_982); +x_984 = lean_array_push(x_719, x_848); +x_985 = lean_array_push(x_984, x_876); +x_986 = lean_array_push(x_985, x_921); +x_987 = lean_array_push(x_986, x_971); +x_988 = lean_array_push(x_987, x_983); +x_989 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_989, 0, x_539); +lean_ctor_set(x_989, 1, x_556); +lean_ctor_set(x_989, 2, x_988); +x_990 = lean_array_push(x_554, x_989); +x_991 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_991, 0, x_539); +lean_ctor_set(x_991, 1, x_630); +lean_ctor_set(x_991, 2, x_990); +x_992 = lean_array_push(x_547, x_628); +x_993 = lean_array_push(x_992, x_991); +x_994 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_994, 0, x_539); +lean_ctor_set(x_994, 1, x_627); +lean_ctor_set(x_994, 2, x_993); +x_995 = lean_array_push(x_560, x_623); +x_996 = lean_array_push(x_995, x_625); +x_997 = lean_array_push(x_996, x_994); +x_998 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_998, 0, x_539); +lean_ctor_set(x_998, 1, x_617); +lean_ctor_set(x_998, 2, x_997); +x_999 = lean_array_push(x_547, x_615); +x_1000 = lean_array_push(x_999, x_998); +x_1001 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1001, 0, x_539); +lean_ctor_set(x_1001, 1, x_614); +lean_ctor_set(x_1001, 2, x_1000); +x_1002 = lean_array_push(x_560, x_612); +x_1003 = lean_array_push(x_1002, x_1001); +x_1004 = lean_array_push(x_1003, x_569); +x_1005 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1005, 0, x_539); +lean_ctor_set(x_1005, 1, x_610); +lean_ctor_set(x_1005, 2, x_1004); +x_1006 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__95; +x_1007 = lean_array_push(x_1006, x_577); +x_1008 = lean_array_push(x_1007, x_587); +x_1009 = lean_array_push(x_1008, x_608); +x_1010 = lean_array_push(x_1009, x_1005); +x_1011 = lean_array_push(x_1010, x_569); x_1012 = lean_array_push(x_1011, x_569); x_1013 = lean_array_push(x_1012, x_569); -x_1014 = lean_array_push(x_1013, x_569); -x_1015 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1015, 0, x_539); -lean_ctor_set(x_1015, 1, x_576); -lean_ctor_set(x_1015, 2, x_1014); -x_1016 = lean_array_push(x_547, x_574); -x_1017 = lean_array_push(x_1016, x_1015); -x_1018 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1018, 0, x_539); -lean_ctor_set(x_1018, 1, x_526); -lean_ctor_set(x_1018, 2, x_1017); -x_1019 = lean_array_push(x_547, x_14); -x_1020 = lean_array_push(x_1019, x_1018); -x_1021 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1021, 0, x_539); -lean_ctor_set(x_1021, 1, x_556); -lean_ctor_set(x_1021, 2, x_1020); -x_1022 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1022, 0, x_1021); -lean_ctor_set(x_1022, 1, x_520); -return x_1022; -} -} -else -{ -lean_object* x_1023; lean_object* x_1024; lean_object* x_1025; lean_object* x_1026; lean_object* x_1027; lean_object* x_1028; lean_object* x_1029; lean_object* x_1030; lean_object* x_1031; lean_object* x_1032; lean_object* x_1033; lean_object* x_1034; lean_object* x_1035; lean_object* x_1036; lean_object* x_1037; lean_object* x_1038; lean_object* x_1039; lean_object* x_1040; lean_object* x_1041; lean_object* x_1042; lean_object* x_1043; lean_object* x_1044; lean_object* x_1045; lean_object* x_1046; lean_object* x_1047; lean_object* x_1048; lean_object* x_1049; lean_object* x_1050; lean_object* x_1051; lean_object* x_1052; lean_object* x_1053; lean_object* x_1054; lean_object* x_1055; lean_object* x_1056; lean_object* x_1057; lean_object* x_1058; lean_object* x_1059; lean_object* x_1060; lean_object* x_1061; lean_object* x_1062; lean_object* x_1063; lean_object* x_1064; lean_object* x_1065; lean_object* x_1066; lean_object* x_1067; lean_object* x_1068; lean_object* x_1069; lean_object* x_1070; lean_object* x_1071; lean_object* x_1072; lean_object* x_1073; lean_object* x_1074; lean_object* x_1075; lean_object* x_1076; lean_object* x_1077; lean_object* x_1078; lean_object* x_1079; lean_object* x_1080; lean_object* x_1081; lean_object* x_1082; lean_object* x_1083; lean_object* x_1084; lean_object* x_1085; lean_object* x_1086; lean_object* x_1087; lean_object* x_1088; lean_object* x_1089; lean_object* x_1090; lean_object* x_1091; lean_object* x_1092; lean_object* x_1093; lean_object* x_1094; lean_object* x_1095; lean_object* x_1096; lean_object* x_1097; lean_object* x_1098; lean_object* x_1099; lean_object* x_1100; lean_object* x_1101; lean_object* x_1102; lean_object* x_1103; lean_object* x_1104; lean_object* x_1105; lean_object* x_1106; lean_object* x_1107; lean_object* x_1108; lean_object* x_1109; lean_object* x_1110; lean_object* x_1111; lean_object* x_1112; lean_object* x_1113; lean_object* x_1114; lean_object* x_1115; lean_object* x_1116; lean_object* x_1117; lean_object* x_1118; lean_object* x_1119; lean_object* x_1120; lean_object* x_1121; lean_object* x_1122; lean_object* x_1123; lean_object* x_1124; lean_object* x_1125; lean_object* x_1126; lean_object* x_1127; lean_object* x_1128; lean_object* x_1129; lean_object* x_1130; lean_object* x_1131; lean_object* x_1132; lean_object* x_1133; lean_object* x_1134; lean_object* x_1135; lean_object* x_1136; lean_object* x_1137; lean_object* x_1138; lean_object* x_1139; lean_object* x_1140; lean_object* x_1141; lean_object* x_1142; lean_object* x_1143; lean_object* x_1144; lean_object* x_1145; lean_object* x_1146; lean_object* x_1147; lean_object* x_1148; lean_object* x_1149; lean_object* x_1150; lean_object* x_1151; lean_object* x_1152; lean_object* x_1153; lean_object* x_1154; lean_object* x_1155; lean_object* x_1156; lean_object* x_1157; lean_object* x_1158; lean_object* x_1159; lean_object* x_1160; lean_object* x_1161; lean_object* x_1162; lean_object* x_1163; lean_object* x_1164; lean_object* x_1165; lean_object* x_1166; lean_object* x_1167; lean_object* x_1168; lean_object* x_1169; lean_object* x_1170; lean_object* x_1171; lean_object* x_1172; lean_object* x_1173; lean_object* x_1174; lean_object* x_1175; lean_object* x_1176; lean_object* x_1177; lean_object* x_1178; lean_object* x_1179; lean_object* x_1180; lean_object* x_1181; lean_object* x_1182; lean_object* x_1183; lean_object* x_1184; lean_object* x_1185; lean_object* x_1186; lean_object* x_1187; lean_object* x_1188; lean_object* x_1189; lean_object* x_1190; lean_object* x_1191; lean_object* x_1192; lean_object* x_1193; lean_object* x_1194; lean_object* x_1195; lean_object* x_1196; lean_object* x_1197; lean_object* x_1198; lean_object* x_1199; lean_object* x_1200; lean_object* x_1201; lean_object* x_1202; lean_object* x_1203; lean_object* x_1204; lean_object* x_1205; lean_object* x_1206; lean_object* x_1207; lean_object* x_1208; lean_object* x_1209; lean_object* x_1210; lean_object* x_1211; lean_object* x_1212; lean_object* x_1213; lean_object* x_1214; lean_object* x_1215; lean_object* x_1216; lean_object* x_1217; lean_object* x_1218; lean_object* x_1219; lean_object* x_1220; lean_object* x_1221; lean_object* x_1222; lean_object* x_1223; lean_object* x_1224; lean_object* x_1225; lean_object* x_1226; lean_object* x_1227; lean_object* x_1228; lean_object* x_1229; lean_object* x_1230; lean_object* x_1231; lean_object* x_1232; lean_object* x_1233; lean_object* x_1234; lean_object* x_1235; lean_object* x_1236; lean_object* x_1237; lean_object* x_1238; lean_object* x_1239; lean_object* x_1240; lean_object* x_1241; lean_object* x_1242; lean_object* x_1243; lean_object* x_1244; lean_object* x_1245; lean_object* x_1246; lean_object* x_1247; lean_object* x_1248; lean_object* x_1249; lean_object* x_1250; lean_object* x_1251; lean_object* x_1252; lean_object* x_1253; lean_object* x_1254; lean_object* x_1255; lean_object* x_1256; lean_object* x_1257; lean_object* x_1258; lean_object* x_1259; lean_object* x_1260; lean_object* x_1261; lean_object* x_1262; lean_object* x_1263; lean_object* x_1264; lean_object* x_1265; lean_object* x_1266; lean_object* x_1267; lean_object* x_1268; lean_object* x_1269; lean_object* x_1270; lean_object* x_1271; lean_object* x_1272; lean_object* x_1273; lean_object* x_1274; lean_object* x_1275; lean_object* x_1276; lean_object* x_1277; lean_object* x_1278; lean_object* x_1279; lean_object* x_1280; lean_object* x_1281; lean_object* x_1282; lean_object* x_1283; lean_object* x_1284; lean_object* x_1285; lean_object* x_1286; lean_object* x_1287; lean_object* x_1288; lean_object* x_1289; lean_object* x_1290; lean_object* x_1291; lean_object* x_1292; lean_object* x_1293; lean_object* x_1294; lean_object* x_1295; lean_object* x_1296; lean_object* x_1297; lean_object* x_1298; lean_object* x_1299; lean_object* x_1300; lean_object* x_1301; lean_object* x_1302; lean_object* x_1303; lean_object* x_1304; lean_object* x_1305; lean_object* x_1306; lean_object* x_1307; lean_object* x_1308; lean_object* x_1309; lean_object* x_1310; lean_object* x_1311; lean_object* x_1312; lean_object* x_1313; lean_object* x_1314; lean_object* x_1315; lean_object* x_1316; lean_object* x_1317; lean_object* x_1318; lean_object* x_1319; lean_object* x_1320; lean_object* x_1321; lean_object* x_1322; lean_object* x_1323; lean_object* x_1324; lean_object* x_1325; lean_object* x_1326; lean_object* x_1327; lean_object* x_1328; lean_object* x_1329; lean_object* x_1330; lean_object* x_1331; lean_object* x_1332; lean_object* x_1333; lean_object* x_1334; lean_object* x_1335; lean_object* x_1336; lean_object* x_1337; lean_object* x_1338; lean_object* x_1339; lean_object* x_1340; lean_object* x_1341; lean_object* x_1342; lean_object* x_1343; lean_object* x_1344; lean_object* x_1345; lean_object* x_1346; lean_object* x_1347; lean_object* x_1348; lean_object* x_1349; lean_object* x_1350; lean_object* x_1351; lean_object* x_1352; lean_object* x_1353; lean_object* x_1354; lean_object* x_1355; lean_object* x_1356; lean_object* x_1357; lean_object* x_1358; lean_object* x_1359; lean_object* x_1360; lean_object* x_1361; lean_object* x_1362; lean_object* x_1363; lean_object* x_1364; lean_object* x_1365; lean_object* x_1366; lean_object* x_1367; lean_object* x_1368; lean_object* x_1369; lean_object* x_1370; lean_object* x_1371; lean_object* x_1372; lean_object* x_1373; lean_object* x_1374; lean_object* x_1375; lean_object* x_1376; lean_object* x_1377; lean_object* x_1378; lean_object* x_1379; lean_object* x_1380; lean_object* x_1381; lean_object* x_1382; lean_object* x_1383; lean_object* x_1384; lean_object* x_1385; lean_object* x_1386; lean_object* x_1387; lean_object* x_1388; lean_object* x_1389; lean_object* x_1390; lean_object* x_1391; lean_object* x_1392; lean_object* x_1393; lean_object* x_1394; lean_object* x_1395; lean_object* x_1396; lean_object* x_1397; lean_object* x_1398; lean_object* x_1399; lean_object* x_1400; lean_object* x_1401; lean_object* x_1402; lean_object* x_1403; lean_object* x_1404; lean_object* x_1405; lean_object* x_1406; lean_object* x_1407; lean_object* x_1408; lean_object* x_1409; lean_object* x_1410; lean_object* x_1411; lean_object* x_1412; lean_object* x_1413; lean_object* x_1414; lean_object* x_1415; lean_object* x_1416; lean_object* x_1417; lean_object* x_1418; lean_object* x_1419; lean_object* x_1420; lean_object* x_1421; lean_object* x_1422; lean_object* x_1423; lean_object* x_1424; lean_object* x_1425; lean_object* x_1426; lean_object* x_1427; lean_object* x_1428; lean_object* x_1429; lean_object* x_1430; lean_object* x_1431; lean_object* x_1432; lean_object* x_1433; lean_object* x_1434; lean_object* x_1435; lean_object* x_1436; lean_object* x_1437; lean_object* x_1438; lean_object* x_1439; lean_object* x_1440; lean_object* x_1441; lean_object* x_1442; lean_object* x_1443; lean_object* x_1444; lean_object* x_1445; lean_object* x_1446; lean_object* x_1447; lean_object* x_1448; lean_object* x_1449; lean_object* x_1450; lean_object* x_1451; lean_object* x_1452; lean_object* x_1453; lean_object* x_1454; lean_object* x_1455; lean_object* x_1456; lean_object* x_1457; lean_object* x_1458; lean_object* x_1459; lean_object* x_1460; lean_object* x_1461; lean_object* x_1462; lean_object* x_1463; lean_object* x_1464; lean_object* x_1465; lean_object* x_1466; lean_object* x_1467; lean_object* x_1468; lean_object* x_1469; lean_object* x_1470; lean_object* x_1471; lean_object* x_1472; lean_object* x_1473; lean_object* x_1474; lean_object* x_1475; lean_object* x_1476; lean_object* x_1477; lean_object* x_1478; lean_object* x_1479; lean_object* x_1480; lean_object* x_1481; lean_object* x_1482; lean_object* x_1483; lean_object* x_1484; lean_object* x_1485; lean_object* x_1486; lean_object* x_1487; lean_object* x_1488; lean_object* x_1489; lean_object* x_1490; lean_object* x_1491; lean_object* x_1492; lean_object* x_1493; lean_object* x_1494; lean_object* x_1495; lean_object* x_1496; lean_object* x_1497; lean_object* x_1498; lean_object* x_1499; lean_object* x_1500; lean_object* x_1501; lean_object* x_1502; lean_object* x_1503; lean_object* x_1504; lean_object* x_1505; lean_object* x_1506; lean_object* x_1507; lean_object* x_1508; lean_object* x_1509; lean_object* x_1510; lean_object* x_1511; lean_object* x_1512; lean_object* x_1513; lean_object* x_1514; lean_object* x_1515; lean_object* x_1516; lean_object* x_1517; lean_object* x_1518; lean_object* x_1519; lean_object* x_1520; lean_object* x_1521; lean_object* x_1522; lean_object* x_1523; lean_object* x_1524; lean_object* x_1525; lean_object* x_1526; lean_object* x_1527; lean_object* x_1528; lean_object* x_1529; lean_object* x_1530; lean_object* x_1531; lean_object* x_1532; -x_1023 = lean_ctor_get(x_6, 0); -x_1024 = lean_ctor_get(x_10, 0); -x_1025 = lean_ctor_get(x_10, 1); -lean_inc(x_1025); +x_1014 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1014, 0, x_539); +lean_ctor_set(x_1014, 1, x_576); +lean_ctor_set(x_1014, 2, x_1013); +x_1015 = lean_array_push(x_547, x_574); +x_1016 = lean_array_push(x_1015, x_1014); +x_1017 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1017, 0, x_539); +lean_ctor_set(x_1017, 1, x_526); +lean_ctor_set(x_1017, 2, x_1016); +x_1018 = lean_array_push(x_547, x_15); +x_1019 = lean_array_push(x_1018, x_1017); +x_1020 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1020, 0, x_539); +lean_ctor_set(x_1020, 1, x_556); +lean_ctor_set(x_1020, 2, x_1019); +x_1021 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1021, 0, x_1020); +lean_ctor_set(x_1021, 1, x_520); +return x_1021; +} +} +else +{ +lean_object* x_1022; lean_object* x_1023; lean_object* x_1024; lean_object* x_1025; lean_object* x_1026; lean_object* x_1027; lean_object* x_1028; lean_object* x_1029; lean_object* x_1030; lean_object* x_1031; lean_object* x_1032; lean_object* x_1033; lean_object* x_1034; lean_object* x_1035; lean_object* x_1036; lean_object* x_1037; lean_object* x_1038; lean_object* x_1039; lean_object* x_1040; lean_object* x_1041; lean_object* x_1042; lean_object* x_1043; lean_object* x_1044; lean_object* x_1045; lean_object* x_1046; lean_object* x_1047; lean_object* x_1048; lean_object* x_1049; lean_object* x_1050; lean_object* x_1051; lean_object* x_1052; lean_object* x_1053; lean_object* x_1054; lean_object* x_1055; lean_object* x_1056; lean_object* x_1057; lean_object* x_1058; lean_object* x_1059; lean_object* x_1060; lean_object* x_1061; lean_object* x_1062; lean_object* x_1063; lean_object* x_1064; lean_object* x_1065; lean_object* x_1066; lean_object* x_1067; lean_object* x_1068; lean_object* x_1069; lean_object* x_1070; lean_object* x_1071; lean_object* x_1072; lean_object* x_1073; lean_object* x_1074; lean_object* x_1075; lean_object* x_1076; lean_object* x_1077; lean_object* x_1078; lean_object* x_1079; lean_object* x_1080; lean_object* x_1081; lean_object* x_1082; lean_object* x_1083; lean_object* x_1084; lean_object* x_1085; lean_object* x_1086; lean_object* x_1087; lean_object* x_1088; lean_object* x_1089; lean_object* x_1090; lean_object* x_1091; lean_object* x_1092; lean_object* x_1093; lean_object* x_1094; lean_object* x_1095; lean_object* x_1096; lean_object* x_1097; lean_object* x_1098; lean_object* x_1099; lean_object* x_1100; lean_object* x_1101; lean_object* x_1102; lean_object* x_1103; lean_object* x_1104; lean_object* x_1105; lean_object* x_1106; lean_object* x_1107; lean_object* x_1108; lean_object* x_1109; lean_object* x_1110; lean_object* x_1111; lean_object* x_1112; lean_object* x_1113; lean_object* x_1114; lean_object* x_1115; lean_object* x_1116; lean_object* x_1117; lean_object* x_1118; lean_object* x_1119; lean_object* x_1120; lean_object* x_1121; lean_object* x_1122; lean_object* x_1123; lean_object* x_1124; lean_object* x_1125; lean_object* x_1126; lean_object* x_1127; lean_object* x_1128; lean_object* x_1129; lean_object* x_1130; lean_object* x_1131; lean_object* x_1132; lean_object* x_1133; lean_object* x_1134; lean_object* x_1135; lean_object* x_1136; lean_object* x_1137; lean_object* x_1138; lean_object* x_1139; lean_object* x_1140; lean_object* x_1141; lean_object* x_1142; lean_object* x_1143; lean_object* x_1144; lean_object* x_1145; lean_object* x_1146; lean_object* x_1147; lean_object* x_1148; lean_object* x_1149; lean_object* x_1150; lean_object* x_1151; lean_object* x_1152; lean_object* x_1153; lean_object* x_1154; lean_object* x_1155; lean_object* x_1156; lean_object* x_1157; lean_object* x_1158; lean_object* x_1159; lean_object* x_1160; lean_object* x_1161; lean_object* x_1162; lean_object* x_1163; lean_object* x_1164; lean_object* x_1165; lean_object* x_1166; lean_object* x_1167; lean_object* x_1168; lean_object* x_1169; lean_object* x_1170; lean_object* x_1171; lean_object* x_1172; lean_object* x_1173; lean_object* x_1174; lean_object* x_1175; lean_object* x_1176; lean_object* x_1177; lean_object* x_1178; lean_object* x_1179; lean_object* x_1180; lean_object* x_1181; lean_object* x_1182; lean_object* x_1183; lean_object* x_1184; lean_object* x_1185; lean_object* x_1186; lean_object* x_1187; lean_object* x_1188; lean_object* x_1189; lean_object* x_1190; lean_object* x_1191; lean_object* x_1192; lean_object* x_1193; lean_object* x_1194; lean_object* x_1195; lean_object* x_1196; lean_object* x_1197; lean_object* x_1198; lean_object* x_1199; lean_object* x_1200; lean_object* x_1201; lean_object* x_1202; lean_object* x_1203; lean_object* x_1204; lean_object* x_1205; lean_object* x_1206; lean_object* x_1207; lean_object* x_1208; lean_object* x_1209; lean_object* x_1210; lean_object* x_1211; lean_object* x_1212; lean_object* x_1213; lean_object* x_1214; lean_object* x_1215; lean_object* x_1216; lean_object* x_1217; lean_object* x_1218; lean_object* x_1219; lean_object* x_1220; lean_object* x_1221; lean_object* x_1222; lean_object* x_1223; lean_object* x_1224; lean_object* x_1225; lean_object* x_1226; lean_object* x_1227; lean_object* x_1228; lean_object* x_1229; lean_object* x_1230; lean_object* x_1231; lean_object* x_1232; lean_object* x_1233; lean_object* x_1234; lean_object* x_1235; lean_object* x_1236; lean_object* x_1237; lean_object* x_1238; lean_object* x_1239; lean_object* x_1240; lean_object* x_1241; lean_object* x_1242; lean_object* x_1243; lean_object* x_1244; lean_object* x_1245; lean_object* x_1246; lean_object* x_1247; lean_object* x_1248; lean_object* x_1249; lean_object* x_1250; lean_object* x_1251; lean_object* x_1252; lean_object* x_1253; lean_object* x_1254; lean_object* x_1255; lean_object* x_1256; lean_object* x_1257; lean_object* x_1258; lean_object* x_1259; lean_object* x_1260; lean_object* x_1261; lean_object* x_1262; lean_object* x_1263; lean_object* x_1264; lean_object* x_1265; lean_object* x_1266; lean_object* x_1267; lean_object* x_1268; lean_object* x_1269; lean_object* x_1270; lean_object* x_1271; lean_object* x_1272; lean_object* x_1273; lean_object* x_1274; lean_object* x_1275; lean_object* x_1276; lean_object* x_1277; lean_object* x_1278; lean_object* x_1279; lean_object* x_1280; lean_object* x_1281; lean_object* x_1282; lean_object* x_1283; lean_object* x_1284; lean_object* x_1285; lean_object* x_1286; lean_object* x_1287; lean_object* x_1288; lean_object* x_1289; lean_object* x_1290; lean_object* x_1291; lean_object* x_1292; lean_object* x_1293; lean_object* x_1294; lean_object* x_1295; lean_object* x_1296; lean_object* x_1297; lean_object* x_1298; lean_object* x_1299; lean_object* x_1300; lean_object* x_1301; lean_object* x_1302; lean_object* x_1303; lean_object* x_1304; lean_object* x_1305; lean_object* x_1306; lean_object* x_1307; lean_object* x_1308; lean_object* x_1309; lean_object* x_1310; lean_object* x_1311; lean_object* x_1312; lean_object* x_1313; lean_object* x_1314; lean_object* x_1315; lean_object* x_1316; lean_object* x_1317; lean_object* x_1318; lean_object* x_1319; lean_object* x_1320; lean_object* x_1321; lean_object* x_1322; lean_object* x_1323; lean_object* x_1324; lean_object* x_1325; lean_object* x_1326; lean_object* x_1327; lean_object* x_1328; lean_object* x_1329; lean_object* x_1330; lean_object* x_1331; lean_object* x_1332; lean_object* x_1333; lean_object* x_1334; lean_object* x_1335; lean_object* x_1336; lean_object* x_1337; lean_object* x_1338; lean_object* x_1339; lean_object* x_1340; lean_object* x_1341; lean_object* x_1342; lean_object* x_1343; lean_object* x_1344; lean_object* x_1345; lean_object* x_1346; lean_object* x_1347; lean_object* x_1348; lean_object* x_1349; lean_object* x_1350; lean_object* x_1351; lean_object* x_1352; lean_object* x_1353; lean_object* x_1354; lean_object* x_1355; lean_object* x_1356; lean_object* x_1357; lean_object* x_1358; lean_object* x_1359; lean_object* x_1360; lean_object* x_1361; lean_object* x_1362; lean_object* x_1363; lean_object* x_1364; lean_object* x_1365; lean_object* x_1366; lean_object* x_1367; lean_object* x_1368; lean_object* x_1369; lean_object* x_1370; lean_object* x_1371; lean_object* x_1372; lean_object* x_1373; lean_object* x_1374; lean_object* x_1375; lean_object* x_1376; lean_object* x_1377; lean_object* x_1378; lean_object* x_1379; lean_object* x_1380; lean_object* x_1381; lean_object* x_1382; lean_object* x_1383; lean_object* x_1384; lean_object* x_1385; lean_object* x_1386; lean_object* x_1387; lean_object* x_1388; lean_object* x_1389; lean_object* x_1390; lean_object* x_1391; lean_object* x_1392; lean_object* x_1393; lean_object* x_1394; lean_object* x_1395; lean_object* x_1396; lean_object* x_1397; lean_object* x_1398; lean_object* x_1399; lean_object* x_1400; lean_object* x_1401; lean_object* x_1402; lean_object* x_1403; lean_object* x_1404; lean_object* x_1405; lean_object* x_1406; lean_object* x_1407; lean_object* x_1408; lean_object* x_1409; lean_object* x_1410; lean_object* x_1411; lean_object* x_1412; lean_object* x_1413; lean_object* x_1414; lean_object* x_1415; lean_object* x_1416; lean_object* x_1417; lean_object* x_1418; lean_object* x_1419; lean_object* x_1420; lean_object* x_1421; lean_object* x_1422; lean_object* x_1423; lean_object* x_1424; lean_object* x_1425; lean_object* x_1426; lean_object* x_1427; lean_object* x_1428; lean_object* x_1429; lean_object* x_1430; lean_object* x_1431; lean_object* x_1432; lean_object* x_1433; lean_object* x_1434; lean_object* x_1435; lean_object* x_1436; lean_object* x_1437; lean_object* x_1438; lean_object* x_1439; lean_object* x_1440; lean_object* x_1441; lean_object* x_1442; lean_object* x_1443; lean_object* x_1444; lean_object* x_1445; lean_object* x_1446; lean_object* x_1447; lean_object* x_1448; lean_object* x_1449; lean_object* x_1450; lean_object* x_1451; lean_object* x_1452; lean_object* x_1453; lean_object* x_1454; lean_object* x_1455; lean_object* x_1456; lean_object* x_1457; lean_object* x_1458; lean_object* x_1459; lean_object* x_1460; lean_object* x_1461; lean_object* x_1462; lean_object* x_1463; lean_object* x_1464; lean_object* x_1465; lean_object* x_1466; lean_object* x_1467; lean_object* x_1468; lean_object* x_1469; lean_object* x_1470; lean_object* x_1471; lean_object* x_1472; lean_object* x_1473; lean_object* x_1474; lean_object* x_1475; lean_object* x_1476; lean_object* x_1477; lean_object* x_1478; lean_object* x_1479; lean_object* x_1480; lean_object* x_1481; lean_object* x_1482; lean_object* x_1483; lean_object* x_1484; lean_object* x_1485; lean_object* x_1486; lean_object* x_1487; lean_object* x_1488; lean_object* x_1489; lean_object* x_1490; lean_object* x_1491; lean_object* x_1492; lean_object* x_1493; lean_object* x_1494; lean_object* x_1495; lean_object* x_1496; lean_object* x_1497; lean_object* x_1498; lean_object* x_1499; lean_object* x_1500; lean_object* x_1501; lean_object* x_1502; lean_object* x_1503; lean_object* x_1504; lean_object* x_1505; lean_object* x_1506; lean_object* x_1507; lean_object* x_1508; lean_object* x_1509; lean_object* x_1510; lean_object* x_1511; lean_object* x_1512; lean_object* x_1513; lean_object* x_1514; lean_object* x_1515; lean_object* x_1516; lean_object* x_1517; lean_object* x_1518; lean_object* x_1519; lean_object* x_1520; lean_object* x_1521; lean_object* x_1522; lean_object* x_1523; lean_object* x_1524; lean_object* x_1525; lean_object* x_1526; lean_object* x_1527; lean_object* x_1528; lean_object* x_1529; lean_object* x_1530; +x_1022 = lean_ctor_get(x_7, 0); +x_1023 = lean_ctor_get(x_11, 0); +x_1024 = lean_ctor_get(x_11, 1); lean_inc(x_1024); -lean_dec(x_10); -lean_inc(x_7); -x_1026 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_7, x_8); -x_1027 = lean_ctor_get(x_1026, 0); +lean_inc(x_1023); +lean_dec(x_11); +lean_inc(x_8); +x_1025 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_8, x_9); +x_1026 = lean_ctor_get(x_1025, 0); +lean_inc(x_1026); +x_1027 = lean_ctor_get(x_1025, 1); lean_inc(x_1027); -x_1028 = lean_ctor_get(x_1026, 1); -lean_inc(x_1028); -if (lean_is_exclusive(x_1026)) { - lean_ctor_release(x_1026, 0); - lean_ctor_release(x_1026, 1); - x_1029 = x_1026; +if (lean_is_exclusive(x_1025)) { + lean_ctor_release(x_1025, 0); + lean_ctor_release(x_1025, 1); + x_1028 = x_1025; } else { - lean_dec_ref(x_1026); - x_1029 = lean_box(0); + lean_dec_ref(x_1025); + x_1028 = lean_box(0); } -x_1030 = lean_ctor_get(x_7, 2); +x_1029 = lean_ctor_get(x_8, 2); +lean_inc(x_1029); +x_1030 = lean_ctor_get(x_8, 1); lean_inc(x_1030); -x_1031 = lean_ctor_get(x_7, 1); -lean_inc(x_1031); -lean_dec(x_7); -x_1032 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__1; +lean_dec(x_8); +x_1031 = l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___closed__1; lean_inc(x_1); -x_1033 = lean_name_mk_string(x_1, x_1032); -x_1034 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__2; -lean_inc(x_1033); -x_1035 = lean_name_mk_string(x_1033, x_1034); -x_1036 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__3; -lean_inc(x_1033); -x_1037 = lean_name_mk_string(x_1033, x_1036); -x_1038 = l_Lean_mkHole___closed__5; +x_1032 = lean_name_mk_string(x_1, x_1031); +x_1033 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__1; +lean_inc(x_1032); +x_1034 = lean_name_mk_string(x_1032, x_1033); +x_1035 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__2; +lean_inc(x_1032); +x_1036 = lean_name_mk_string(x_1032, x_1035); +x_1037 = l_Lean_mkHole___closed__5; lean_inc(x_1); -x_1039 = lean_name_mk_string(x_1, x_1038); -x_1040 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__4; -lean_inc(x_1039); -x_1041 = lean_name_mk_string(x_1039, x_1040); -x_1042 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__5; -lean_inc(x_1027); -x_1043 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1043, 0, x_1027); -lean_ctor_set(x_1043, 1, x_1042); -x_1044 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__6; -lean_inc(x_1039); -x_1045 = lean_name_mk_string(x_1039, x_1044); -x_1046 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__7; -lean_inc(x_1039); -x_1047 = lean_name_mk_string(x_1039, x_1046); -x_1048 = lean_box(2); -x_1049 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__33; -x_1050 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1050, 0, x_1048); -lean_ctor_set(x_1050, 1, x_1047); -lean_ctor_set(x_1050, 2, x_1049); -x_1051 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__8; -x_1052 = lean_name_mk_string(x_1, x_1051); -x_1053 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__9; -x_1054 = lean_name_mk_string(x_1052, x_1053); -lean_inc(x_1027); -x_1055 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1055, 0, x_1027); -lean_ctor_set(x_1055, 1, x_1053); -x_1056 = l_Lean_Syntax_mkApp___closed__3; -x_1057 = lean_array_push(x_1056, x_1055); -x_1058 = lean_array_push(x_1057, x_2); -x_1059 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1059, 0, x_1048); -lean_ctor_set(x_1059, 1, x_1054); -lean_ctor_set(x_1059, 2, x_1058); -x_1060 = lean_array_push(x_1056, x_1050); -x_1061 = lean_array_push(x_1060, x_1059); -x_1062 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1062, 0, x_1048); -lean_ctor_set(x_1062, 1, x_1045); -lean_ctor_set(x_1062, 2, x_1061); -x_1063 = l_Lean_mkOptionalNode___closed__2; -x_1064 = lean_array_push(x_1063, x_1062); -x_1065 = l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__2; -x_1066 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1066, 0, x_1048); -lean_ctor_set(x_1066, 1, x_1065); -lean_ctor_set(x_1066, 2, x_1064); -x_1067 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__16; -lean_inc(x_1027); -x_1068 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1068, 0, x_1027); -lean_ctor_set(x_1068, 1, x_1067); -x_1069 = l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__4; -x_1070 = lean_array_push(x_1069, x_1043); -x_1071 = lean_array_push(x_1070, x_1066); -lean_inc(x_1068); -x_1072 = lean_array_push(x_1071, x_1068); -x_1073 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1073, 0, x_1048); -lean_ctor_set(x_1073, 1, x_1041); -lean_ctor_set(x_1073, 2, x_1072); -x_1074 = lean_array_push(x_1063, x_1073); -x_1075 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1075, 0, x_1048); -lean_ctor_set(x_1075, 1, x_1065); -lean_ctor_set(x_1075, 2, x_1074); -x_1076 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__10; -x_1077 = lean_array_push(x_1076, x_1075); -x_1078 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__12; -x_1079 = lean_array_push(x_1077, x_1078); -x_1080 = lean_array_push(x_1079, x_1078); -x_1081 = lean_array_push(x_1080, x_1078); -x_1082 = lean_array_push(x_1081, x_1078); -x_1083 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1083, 0, x_1048); -lean_ctor_set(x_1083, 1, x_1037); -lean_ctor_set(x_1083, 2, x_1082); -x_1084 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__11; -lean_inc(x_1033); -x_1085 = lean_name_mk_string(x_1033, x_1084); -lean_inc(x_1027); -x_1086 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1086, 0, x_1027); -lean_ctor_set(x_1086, 1, x_1084); -x_1087 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__12; -lean_inc(x_1033); -x_1088 = lean_name_mk_string(x_1033, x_1087); -x_1089 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__16; +x_1038 = lean_name_mk_string(x_1, x_1037); +x_1039 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__3; +lean_inc(x_1038); +x_1040 = lean_name_mk_string(x_1038, x_1039); +x_1041 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__4; +lean_inc(x_1026); +x_1042 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1042, 0, x_1026); +lean_ctor_set(x_1042, 1, x_1041); +x_1043 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__5; +lean_inc(x_1038); +x_1044 = lean_name_mk_string(x_1038, x_1043); +x_1045 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__6; +lean_inc(x_1038); +x_1046 = lean_name_mk_string(x_1038, x_1045); +x_1047 = lean_box(2); +x_1048 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__33; +x_1049 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1049, 0, x_1047); +lean_ctor_set(x_1049, 1, x_1046); +lean_ctor_set(x_1049, 2, x_1048); +x_1050 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__7; +x_1051 = lean_name_mk_string(x_1, x_1050); +x_1052 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__8; +x_1053 = lean_name_mk_string(x_1051, x_1052); +lean_inc(x_1026); +x_1054 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1054, 0, x_1026); +lean_ctor_set(x_1054, 1, x_1052); +x_1055 = l_Lean_Syntax_mkApp___closed__3; +x_1056 = lean_array_push(x_1055, x_1054); +x_1057 = lean_array_push(x_1056, x_2); +x_1058 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1058, 0, x_1047); +lean_ctor_set(x_1058, 1, x_1053); +lean_ctor_set(x_1058, 2, x_1057); +x_1059 = lean_array_push(x_1055, x_1049); +x_1060 = lean_array_push(x_1059, x_1058); +x_1061 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1061, 0, x_1047); +lean_ctor_set(x_1061, 1, x_1044); +lean_ctor_set(x_1061, 2, x_1060); +x_1062 = l_Lean_mkOptionalNode___closed__2; +x_1063 = lean_array_push(x_1062, x_1061); +x_1064 = l_Lean_mkNullNode___closed__2; +x_1065 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1065, 0, x_1047); +lean_ctor_set(x_1065, 1, x_1064); +lean_ctor_set(x_1065, 2, x_1063); +x_1066 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__16; +lean_inc(x_1026); +x_1067 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1067, 0, x_1026); +lean_ctor_set(x_1067, 1, x_1066); +x_1068 = l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__4; +x_1069 = lean_array_push(x_1068, x_1042); +x_1070 = lean_array_push(x_1069, x_1065); +lean_inc(x_1067); +x_1071 = lean_array_push(x_1070, x_1067); +x_1072 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1072, 0, x_1047); +lean_ctor_set(x_1072, 1, x_1040); +lean_ctor_set(x_1072, 2, x_1071); +x_1073 = lean_array_push(x_1062, x_1072); +x_1074 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1074, 0, x_1047); +lean_ctor_set(x_1074, 1, x_1064); +lean_ctor_set(x_1074, 2, x_1073); +x_1075 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__9; +x_1076 = lean_array_push(x_1075, x_1074); +x_1077 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__12; +x_1078 = lean_array_push(x_1076, x_1077); +x_1079 = lean_array_push(x_1078, x_1077); +x_1080 = lean_array_push(x_1079, x_1077); +x_1081 = lean_array_push(x_1080, x_1077); +x_1082 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1082, 0, x_1047); +lean_ctor_set(x_1082, 1, x_1036); +lean_ctor_set(x_1082, 2, x_1081); +x_1083 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__10; +lean_inc(x_1032); +x_1084 = lean_name_mk_string(x_1032, x_1083); +lean_inc(x_1026); +x_1085 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1085, 0, x_1026); +lean_ctor_set(x_1085, 1, x_1083); +x_1086 = l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___closed__3; +lean_inc(x_1032); +x_1087 = lean_name_mk_string(x_1032, x_1086); +x_1088 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__14; +lean_inc(x_1029); lean_inc(x_1030); -lean_inc(x_1031); -x_1090 = l_Lean_addMacroScope(x_1031, x_1089, x_1030); -x_1091 = lean_box(0); -x_1092 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__15; -lean_inc(x_1027); -x_1093 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1093, 0, x_1027); -lean_ctor_set(x_1093, 1, x_1092); -lean_ctor_set(x_1093, 2, x_1090); -lean_ctor_set(x_1093, 3, x_1091); -x_1094 = lean_array_push(x_1056, x_1093); -x_1095 = lean_array_push(x_1094, x_1078); -x_1096 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1096, 0, x_1048); -lean_ctor_set(x_1096, 1, x_1088); -lean_ctor_set(x_1096, 2, x_1095); -x_1097 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__17; -lean_inc(x_1033); -x_1098 = lean_name_mk_string(x_1033, x_1097); -x_1099 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__18; -lean_inc(x_1039); -x_1100 = lean_name_mk_string(x_1039, x_1099); -x_1101 = l_Lean_toolchain___closed__1; -lean_inc(x_1027); -x_1102 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1102, 0, x_1027); -lean_ctor_set(x_1102, 1, x_1101); -x_1103 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__22; +x_1089 = l_Lean_addMacroScope(x_1030, x_1088, x_1029); +x_1090 = lean_box(0); +x_1091 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__13; +lean_inc(x_1026); +x_1092 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1092, 0, x_1026); +lean_ctor_set(x_1092, 1, x_1091); +lean_ctor_set(x_1092, 2, x_1089); +lean_ctor_set(x_1092, 3, x_1090); +x_1093 = lean_array_push(x_1055, x_1092); +x_1094 = lean_array_push(x_1093, x_1077); +x_1095 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1095, 0, x_1047); +lean_ctor_set(x_1095, 1, x_1087); +lean_ctor_set(x_1095, 2, x_1094); +x_1096 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__15; +lean_inc(x_1032); +x_1097 = lean_name_mk_string(x_1032, x_1096); +x_1098 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__16; +lean_inc(x_1038); +x_1099 = lean_name_mk_string(x_1038, x_1098); +x_1100 = l_Lean_toolchain___closed__1; +lean_inc(x_1026); +x_1101 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1101, 0, x_1026); +lean_ctor_set(x_1101, 1, x_1100); +x_1102 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__20; +lean_inc(x_1029); lean_inc(x_1030); -lean_inc(x_1031); -x_1104 = l_Lean_addMacroScope(x_1031, x_1103, x_1030); -x_1105 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__19; +x_1103 = l_Lean_addMacroScope(x_1030, x_1102, x_1029); +x_1104 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__17; lean_inc(x_3); -x_1106 = lean_name_mk_string(x_3, x_1105); -x_1107 = lean_alloc_ctor(0, 2, 0); +x_1105 = lean_name_mk_string(x_3, x_1104); +x_1106 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1106, 0, x_1105); +lean_ctor_set(x_1106, 1, x_1090); +x_1107 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_1107, 0, x_1106); -lean_ctor_set(x_1107, 1, x_1091); -x_1108 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1108, 0, x_1107); -lean_ctor_set(x_1108, 1, x_1091); -x_1109 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__21; -lean_inc(x_1027); -x_1110 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1110, 0, x_1027); -lean_ctor_set(x_1110, 1, x_1109); -lean_ctor_set(x_1110, 2, x_1104); -lean_ctor_set(x_1110, 3, x_1108); -x_1111 = lean_array_push(x_1056, x_1102); -lean_inc(x_1111); -x_1112 = lean_array_push(x_1111, x_1110); -x_1113 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1113, 0, x_1048); -lean_ctor_set(x_1113, 1, x_1100); -lean_ctor_set(x_1113, 2, x_1112); -x_1114 = lean_array_push(x_1063, x_1113); -x_1115 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1115, 0, x_1048); -lean_ctor_set(x_1115, 1, x_1065); -lean_ctor_set(x_1115, 2, x_1114); -x_1116 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__23; -x_1117 = lean_array_push(x_1116, x_1115); -x_1118 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1118, 0, x_1048); -lean_ctor_set(x_1118, 1, x_1098); -lean_ctor_set(x_1118, 2, x_1117); -x_1119 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__24; -x_1120 = lean_name_mk_string(x_1033, x_1119); -x_1121 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__8; -lean_inc(x_1027); -x_1122 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1122, 0, x_1027); -lean_ctor_set(x_1122, 1, x_1121); -x_1123 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__25; -lean_inc(x_1039); -x_1124 = lean_name_mk_string(x_1039, x_1123); -lean_inc(x_1027); -x_1125 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1125, 0, x_1027); -lean_ctor_set(x_1125, 1, x_1123); -x_1126 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__26; -lean_inc(x_1039); -x_1127 = lean_name_mk_string(x_1039, x_1126); -x_1128 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__30; +lean_ctor_set(x_1107, 1, x_1090); +x_1108 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__19; +lean_inc(x_1026); +x_1109 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1109, 0, x_1026); +lean_ctor_set(x_1109, 1, x_1108); +lean_ctor_set(x_1109, 2, x_1103); +lean_ctor_set(x_1109, 3, x_1107); +x_1110 = lean_array_push(x_1055, x_1101); +lean_inc(x_1110); +x_1111 = lean_array_push(x_1110, x_1109); +x_1112 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1112, 0, x_1047); +lean_ctor_set(x_1112, 1, x_1099); +lean_ctor_set(x_1112, 2, x_1111); +x_1113 = lean_array_push(x_1062, x_1112); +x_1114 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1114, 0, x_1047); +lean_ctor_set(x_1114, 1, x_1064); +lean_ctor_set(x_1114, 2, x_1113); +x_1115 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__21; +x_1116 = lean_array_push(x_1115, x_1114); +x_1117 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1117, 0, x_1047); +lean_ctor_set(x_1117, 1, x_1097); +lean_ctor_set(x_1117, 2, x_1116); +x_1118 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__22; +x_1119 = lean_name_mk_string(x_1032, x_1118); +x_1120 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__8; +lean_inc(x_1026); +x_1121 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1121, 0, x_1026); +lean_ctor_set(x_1121, 1, x_1120); +x_1122 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__23; +lean_inc(x_1038); +x_1123 = lean_name_mk_string(x_1038, x_1122); +lean_inc(x_1026); +x_1124 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1124, 0, x_1026); +lean_ctor_set(x_1124, 1, x_1122); +x_1125 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__24; +lean_inc(x_1038); +x_1126 = lean_name_mk_string(x_1038, x_1125); +x_1127 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__28; +lean_inc(x_1029); lean_inc(x_1030); -lean_inc(x_1031); -x_1129 = l_Lean_addMacroScope(x_1031, x_1128, x_1030); -x_1130 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__29; -lean_inc(x_1027); -x_1131 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1131, 0, x_1027); -lean_ctor_set(x_1131, 1, x_1130); -lean_ctor_set(x_1131, 2, x_1129); -lean_ctor_set(x_1131, 3, x_1091); -lean_inc(x_1131); -x_1132 = lean_array_push(x_1063, x_1131); -x_1133 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1133, 0, x_1048); -lean_ctor_set(x_1133, 1, x_1065); -lean_ctor_set(x_1133, 2, x_1132); -x_1134 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__31; -lean_inc(x_1027); -x_1135 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1135, 0, x_1027); -lean_ctor_set(x_1135, 1, x_1134); -x_1136 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__32; -lean_inc(x_1039); -x_1137 = lean_name_mk_string(x_1039, x_1136); -lean_inc(x_1027); -x_1138 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1138, 0, x_1027); -lean_ctor_set(x_1138, 1, x_1136); -x_1139 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__33; -lean_inc(x_1039); -x_1140 = lean_name_mk_string(x_1039, x_1139); -x_1141 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__34; -lean_inc(x_1039); -x_1142 = lean_name_mk_string(x_1039, x_1141); -x_1143 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__35; -lean_inc(x_1039); -x_1144 = lean_name_mk_string(x_1039, x_1143); -x_1145 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__36; -lean_inc(x_1027); -x_1146 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1146, 0, x_1027); -lean_ctor_set(x_1146, 1, x_1145); -x_1147 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__37; -lean_inc(x_1039); -x_1148 = lean_name_mk_string(x_1039, x_1147); -x_1149 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__41; +x_1128 = l_Lean_addMacroScope(x_1030, x_1127, x_1029); +x_1129 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__27; +lean_inc(x_1026); +x_1130 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1130, 0, x_1026); +lean_ctor_set(x_1130, 1, x_1129); +lean_ctor_set(x_1130, 2, x_1128); +lean_ctor_set(x_1130, 3, x_1090); +lean_inc(x_1130); +x_1131 = lean_array_push(x_1062, x_1130); +x_1132 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1132, 0, x_1047); +lean_ctor_set(x_1132, 1, x_1064); +lean_ctor_set(x_1132, 2, x_1131); +x_1133 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__29; +lean_inc(x_1026); +x_1134 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1134, 0, x_1026); +lean_ctor_set(x_1134, 1, x_1133); +x_1135 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__30; +lean_inc(x_1038); +x_1136 = lean_name_mk_string(x_1038, x_1135); +lean_inc(x_1026); +x_1137 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1137, 0, x_1026); +lean_ctor_set(x_1137, 1, x_1135); +x_1138 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__31; +lean_inc(x_1038); +x_1139 = lean_name_mk_string(x_1038, x_1138); +x_1140 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__32; +lean_inc(x_1038); +x_1141 = lean_name_mk_string(x_1038, x_1140); +x_1142 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__33; +lean_inc(x_1038); +x_1143 = lean_name_mk_string(x_1038, x_1142); +x_1144 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__34; +lean_inc(x_1026); +x_1145 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1145, 0, x_1026); +lean_ctor_set(x_1145, 1, x_1144); +x_1146 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__35; +lean_inc(x_1038); +x_1147 = lean_name_mk_string(x_1038, x_1146); +x_1148 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__39; +lean_inc(x_1029); lean_inc(x_1030); -lean_inc(x_1031); -x_1150 = l_Lean_addMacroScope(x_1031, x_1149, x_1030); -x_1151 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__40; -lean_inc(x_1027); -x_1152 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1152, 0, x_1027); -lean_ctor_set(x_1152, 1, x_1151); -lean_ctor_set(x_1152, 2, x_1150); -lean_ctor_set(x_1152, 3, x_1091); -x_1153 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__42; -lean_inc(x_1027); -x_1154 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1154, 0, x_1027); -lean_ctor_set(x_1154, 1, x_1153); -x_1155 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__43; -lean_inc(x_1039); -x_1156 = lean_name_mk_string(x_1039, x_1155); -x_1157 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__44; -lean_inc(x_1027); -x_1158 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1158, 0, x_1027); -lean_ctor_set(x_1158, 1, x_1157); -x_1159 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__45; -lean_inc(x_1039); -x_1160 = lean_name_mk_string(x_1039, x_1159); -x_1161 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__46; -lean_inc(x_1039); -x_1162 = lean_name_mk_string(x_1039, x_1161); -x_1163 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__12; -lean_inc(x_1027); -x_1164 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1164, 0, x_1027); -lean_ctor_set(x_1164, 1, x_1163); -x_1165 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__47; -lean_inc(x_1027); -x_1166 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1166, 0, x_1027); -lean_ctor_set(x_1166, 1, x_1165); -x_1167 = lean_array_push(x_1063, x_1166); -x_1168 = l_Lean_evalPrec___closed__2; -x_1169 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1169, 0, x_1048); -lean_ctor_set(x_1169, 1, x_1168); -lean_ctor_set(x_1169, 2, x_1167); -x_1170 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__38; -lean_inc(x_1131); -x_1171 = lean_array_push(x_1170, x_1131); -lean_inc(x_1164); -x_1172 = lean_array_push(x_1171, x_1164); -lean_inc(x_1169); -lean_inc(x_1172); -x_1173 = lean_array_push(x_1172, x_1169); -lean_inc(x_1068); -x_1174 = lean_array_push(x_1173, x_1068); -lean_inc(x_1162); -x_1175 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1175, 0, x_1048); -lean_ctor_set(x_1175, 1, x_1162); -lean_ctor_set(x_1175, 2, x_1174); -x_1176 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__48; -lean_inc(x_1027); -x_1177 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1177, 0, x_1027); -lean_ctor_set(x_1177, 1, x_1176); -x_1178 = lean_array_push(x_1063, x_1177); -x_1179 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1179, 0, x_1048); -lean_ctor_set(x_1179, 1, x_1168); -lean_ctor_set(x_1179, 2, x_1178); -x_1180 = lean_array_push(x_1170, x_1175); -x_1181 = lean_array_push(x_1180, x_1164); -lean_inc(x_1179); -x_1182 = lean_array_push(x_1181, x_1179); -lean_inc(x_1068); -x_1183 = lean_array_push(x_1182, x_1068); -lean_inc(x_1162); -x_1184 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1184, 0, x_1048); -lean_ctor_set(x_1184, 1, x_1162); -lean_ctor_set(x_1184, 2, x_1183); -x_1185 = lean_array_push(x_1116, x_1184); -x_1186 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1186, 0, x_1048); -lean_ctor_set(x_1186, 1, x_1160); -lean_ctor_set(x_1186, 2, x_1185); -x_1187 = lean_array_push(x_1063, x_1186); -x_1188 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1188, 0, x_1048); -lean_ctor_set(x_1188, 1, x_1065); -lean_ctor_set(x_1188, 2, x_1187); -x_1189 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__49; -lean_inc(x_1027); -x_1190 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1190, 0, x_1027); -lean_ctor_set(x_1190, 1, x_1189); -x_1191 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__50; -lean_inc(x_1039); -x_1192 = lean_name_mk_string(x_1039, x_1191); -x_1193 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__51; -lean_inc(x_1039); -x_1194 = lean_name_mk_string(x_1039, x_1193); -x_1195 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__52; -lean_inc(x_1027); -x_1196 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1196, 0, x_1027); -lean_ctor_set(x_1196, 1, x_1195); -x_1197 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__53; -lean_inc(x_1039); -x_1198 = lean_name_mk_string(x_1039, x_1197); -x_1199 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__54; -lean_inc(x_1027); -x_1200 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1200, 0, x_1027); -lean_ctor_set(x_1200, 1, x_1199); -x_1201 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__57; +x_1149 = l_Lean_addMacroScope(x_1030, x_1148, x_1029); +x_1150 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__38; +lean_inc(x_1026); +x_1151 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1151, 0, x_1026); +lean_ctor_set(x_1151, 1, x_1150); +lean_ctor_set(x_1151, 2, x_1149); +lean_ctor_set(x_1151, 3, x_1090); +x_1152 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__40; +lean_inc(x_1026); +x_1153 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1153, 0, x_1026); +lean_ctor_set(x_1153, 1, x_1152); +x_1154 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__41; +lean_inc(x_1038); +x_1155 = lean_name_mk_string(x_1038, x_1154); +x_1156 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__42; +lean_inc(x_1026); +x_1157 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1157, 0, x_1026); +lean_ctor_set(x_1157, 1, x_1156); +x_1158 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__43; +lean_inc(x_1038); +x_1159 = lean_name_mk_string(x_1038, x_1158); +x_1160 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__44; +lean_inc(x_1038); +x_1161 = lean_name_mk_string(x_1038, x_1160); +x_1162 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__12; +lean_inc(x_1026); +x_1163 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1163, 0, x_1026); +lean_ctor_set(x_1163, 1, x_1162); +x_1164 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__45; +lean_inc(x_1026); +x_1165 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1165, 0, x_1026); +lean_ctor_set(x_1165, 1, x_1164); +x_1166 = lean_array_push(x_1062, x_1165); +x_1167 = l_Lean_Syntax_mkNumLit___closed__2; +x_1168 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1168, 0, x_1047); +lean_ctor_set(x_1168, 1, x_1167); +lean_ctor_set(x_1168, 2, x_1166); +x_1169 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__38; +lean_inc(x_1130); +x_1170 = lean_array_push(x_1169, x_1130); +lean_inc(x_1163); +x_1171 = lean_array_push(x_1170, x_1163); +lean_inc(x_1168); +lean_inc(x_1171); +x_1172 = lean_array_push(x_1171, x_1168); +lean_inc(x_1067); +x_1173 = lean_array_push(x_1172, x_1067); +lean_inc(x_1161); +x_1174 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1174, 0, x_1047); +lean_ctor_set(x_1174, 1, x_1161); +lean_ctor_set(x_1174, 2, x_1173); +x_1175 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__46; +lean_inc(x_1026); +x_1176 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1176, 0, x_1026); +lean_ctor_set(x_1176, 1, x_1175); +x_1177 = lean_array_push(x_1062, x_1176); +x_1178 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1178, 0, x_1047); +lean_ctor_set(x_1178, 1, x_1167); +lean_ctor_set(x_1178, 2, x_1177); +x_1179 = lean_array_push(x_1169, x_1174); +x_1180 = lean_array_push(x_1179, x_1163); +lean_inc(x_1178); +x_1181 = lean_array_push(x_1180, x_1178); +lean_inc(x_1067); +x_1182 = lean_array_push(x_1181, x_1067); +lean_inc(x_1161); +x_1183 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1183, 0, x_1047); +lean_ctor_set(x_1183, 1, x_1161); +lean_ctor_set(x_1183, 2, x_1182); +x_1184 = lean_array_push(x_1115, x_1183); +x_1185 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1185, 0, x_1047); +lean_ctor_set(x_1185, 1, x_1159); +lean_ctor_set(x_1185, 2, x_1184); +x_1186 = lean_array_push(x_1062, x_1185); +x_1187 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1187, 0, x_1047); +lean_ctor_set(x_1187, 1, x_1064); +lean_ctor_set(x_1187, 2, x_1186); +x_1188 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__47; +lean_inc(x_1026); +x_1189 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1189, 0, x_1026); +lean_ctor_set(x_1189, 1, x_1188); +x_1190 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__48; +lean_inc(x_1038); +x_1191 = lean_name_mk_string(x_1038, x_1190); +x_1192 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__49; +lean_inc(x_1038); +x_1193 = lean_name_mk_string(x_1038, x_1192); +x_1194 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__50; +lean_inc(x_1026); +x_1195 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1195, 0, x_1026); +lean_ctor_set(x_1195, 1, x_1194); +x_1196 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__51; +lean_inc(x_1038); +x_1197 = lean_name_mk_string(x_1038, x_1196); +x_1198 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__52; +lean_inc(x_1026); +x_1199 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1199, 0, x_1026); +lean_ctor_set(x_1199, 1, x_1198); +x_1200 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__55; +lean_inc(x_1029); lean_inc(x_1030); -lean_inc(x_1031); -x_1202 = l_Lean_addMacroScope(x_1031, x_1201, x_1030); -x_1203 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__6; -x_1204 = lean_name_mk_string(x_4, x_1203); -lean_inc(x_1204); -lean_ctor_set(x_6, 1, x_1091); -lean_ctor_set(x_6, 0, x_1204); -x_1205 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1205, 0, x_6); -lean_ctor_set(x_1205, 1, x_1091); -x_1206 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__56; -lean_inc(x_1027); -x_1207 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1207, 0, x_1027); -lean_ctor_set(x_1207, 1, x_1206); -lean_ctor_set(x_1207, 2, x_1202); -lean_ctor_set(x_1207, 3, x_1205); -x_1208 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__4; -lean_inc(x_1027); +x_1201 = l_Lean_addMacroScope(x_1030, x_1200, x_1029); +x_1202 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__6; +x_1203 = lean_name_mk_string(x_4, x_1202); +lean_inc(x_1203); +lean_ctor_set(x_7, 1, x_1090); +lean_ctor_set(x_7, 0, x_1203); +x_1204 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1204, 0, x_7); +lean_ctor_set(x_1204, 1, x_1090); +x_1205 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__54; +lean_inc(x_1026); +x_1206 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1206, 0, x_1026); +lean_ctor_set(x_1206, 1, x_1205); +lean_ctor_set(x_1206, 2, x_1201); +lean_ctor_set(x_1206, 3, x_1204); +x_1207 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__4; +lean_inc(x_1026); +x_1208 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1208, 0, x_1026); +lean_ctor_set(x_1208, 1, x_1207); +lean_inc(x_1026); x_1209 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1209, 0, x_1027); -lean_ctor_set(x_1209, 1, x_1208); -lean_inc(x_1027); -x_1210 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1210, 0, x_1027); -lean_ctor_set(x_1210, 1, x_1203); -x_1211 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__60; -lean_inc(x_1027); -x_1212 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1212, 0, x_1027); -lean_ctor_set(x_1212, 1, x_1211); -x_1213 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__18; -lean_inc(x_1027); -x_1214 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1214, 0, x_1027); -lean_ctor_set(x_1214, 1, x_1213); -x_1215 = lean_array_push(x_1111, x_1214); -x_1216 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__62; -x_1217 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1217, 0, x_1048); +lean_ctor_set(x_1209, 0, x_1026); +lean_ctor_set(x_1209, 1, x_1202); +x_1210 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__56; +x_1211 = lean_name_mk_string(x_5, x_1210); +x_1212 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__57; +x_1213 = lean_name_mk_string(x_1211, x_1212); +x_1214 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__58; +lean_inc(x_1026); +x_1215 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1215, 0, x_1026); +lean_ctor_set(x_1215, 1, x_1214); +x_1216 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__16; +lean_inc(x_1026); +x_1217 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1217, 0, x_1026); lean_ctor_set(x_1217, 1, x_1216); -lean_ctor_set(x_1217, 2, x_1215); -x_1218 = lean_array_push(x_1170, x_1212); -x_1219 = lean_array_push(x_1218, x_1078); -lean_inc(x_1152); -x_1220 = lean_array_push(x_1219, x_1152); -lean_inc(x_1220); -x_1221 = lean_array_push(x_1220, x_1217); -x_1222 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__59; -x_1223 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1223, 0, x_1048); -lean_ctor_set(x_1223, 1, x_1222); -lean_ctor_set(x_1223, 2, x_1221); -x_1224 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__8; -lean_inc(x_1027); -x_1225 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1225, 0, x_1027); -lean_ctor_set(x_1225, 1, x_1224); -x_1226 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__37; -lean_inc(x_1209); -x_1227 = lean_array_push(x_1226, x_1209); -x_1228 = lean_array_push(x_1227, x_1210); -lean_inc(x_1122); -x_1229 = lean_array_push(x_1228, x_1122); -lean_inc(x_1229); -x_1230 = lean_array_push(x_1229, x_1223); -lean_inc(x_1225); -x_1231 = lean_array_push(x_1230, x_1225); -lean_inc(x_1204); -x_1232 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1232, 0, x_1048); -lean_ctor_set(x_1232, 1, x_1204); -lean_ctor_set(x_1232, 2, x_1231); -x_1233 = lean_array_push(x_1226, x_1200); -x_1234 = lean_array_push(x_1233, x_1207); -lean_inc(x_1196); -x_1235 = lean_array_push(x_1234, x_1196); -lean_inc(x_1235); -x_1236 = lean_array_push(x_1235, x_1232); -lean_inc(x_1225); -x_1237 = lean_array_push(x_1236, x_1225); -lean_inc(x_1198); -x_1238 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1238, 0, x_1048); -lean_ctor_set(x_1238, 1, x_1198); -lean_ctor_set(x_1238, 2, x_1237); -x_1239 = lean_array_push(x_1063, x_1238); +x_1218 = lean_array_push(x_1110, x_1217); +x_1219 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__60; +x_1220 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1220, 0, x_1047); +lean_ctor_set(x_1220, 1, x_1219); +lean_ctor_set(x_1220, 2, x_1218); +x_1221 = lean_array_push(x_1169, x_1215); +x_1222 = lean_array_push(x_1221, x_1077); +lean_inc(x_1151); +x_1223 = lean_array_push(x_1222, x_1151); +lean_inc(x_1223); +x_1224 = lean_array_push(x_1223, x_1220); +lean_inc(x_1213); +x_1225 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1225, 0, x_1047); +lean_ctor_set(x_1225, 1, x_1213); +lean_ctor_set(x_1225, 2, x_1224); +x_1226 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__8; +lean_inc(x_1026); +x_1227 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1227, 0, x_1026); +lean_ctor_set(x_1227, 1, x_1226); +x_1228 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__37; +lean_inc(x_1208); +x_1229 = lean_array_push(x_1228, x_1208); +x_1230 = lean_array_push(x_1229, x_1209); +lean_inc(x_1121); +x_1231 = lean_array_push(x_1230, x_1121); +lean_inc(x_1231); +x_1232 = lean_array_push(x_1231, x_1225); +lean_inc(x_1227); +x_1233 = lean_array_push(x_1232, x_1227); +lean_inc(x_1203); +x_1234 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1234, 0, x_1047); +lean_ctor_set(x_1234, 1, x_1203); +lean_ctor_set(x_1234, 2, x_1233); +x_1235 = lean_array_push(x_1228, x_1199); +x_1236 = lean_array_push(x_1235, x_1206); +lean_inc(x_1195); +x_1237 = lean_array_push(x_1236, x_1195); +lean_inc(x_1237); +x_1238 = lean_array_push(x_1237, x_1234); +lean_inc(x_1227); +x_1239 = lean_array_push(x_1238, x_1227); +lean_inc(x_1197); x_1240 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1240, 0, x_1048); -lean_ctor_set(x_1240, 1, x_1065); +lean_ctor_set(x_1240, 0, x_1047); +lean_ctor_set(x_1240, 1, x_1197); lean_ctor_set(x_1240, 2, x_1239); -x_1241 = lean_array_push(x_1063, x_1240); +x_1241 = lean_array_push(x_1062, x_1240); x_1242 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1242, 0, x_1048); -lean_ctor_set(x_1242, 1, x_1065); +lean_ctor_set(x_1242, 0, x_1047); +lean_ctor_set(x_1242, 1, x_1064); lean_ctor_set(x_1242, 2, x_1241); -x_1243 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__63; -lean_inc(x_1039); -x_1244 = lean_name_mk_string(x_1039, x_1243); -x_1245 = l_Lean_Syntax_mkApp___closed__1; -lean_inc(x_1039); -x_1246 = lean_name_mk_string(x_1039, x_1245); -x_1247 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__64; -lean_inc(x_1039); -x_1248 = lean_name_mk_string(x_1039, x_1247); -x_1249 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__58; -x_1250 = lean_name_mk_string(x_1248, x_1249); -x_1251 = lean_array_push(x_1220, x_1078); +x_1243 = lean_array_push(x_1062, x_1242); +x_1244 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1244, 0, x_1047); +lean_ctor_set(x_1244, 1, x_1064); +lean_ctor_set(x_1244, 2, x_1243); +x_1245 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__61; +lean_inc(x_1038); +x_1246 = lean_name_mk_string(x_1038, x_1245); +x_1247 = l_Lean_Syntax_mkApp___closed__1; +lean_inc(x_1038); +x_1248 = lean_name_mk_string(x_1038, x_1247); +x_1249 = lean_array_push(x_1223, x_1077); +x_1250 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1250, 0, x_1047); +lean_ctor_set(x_1250, 1, x_1213); +lean_ctor_set(x_1250, 2, x_1249); +x_1251 = lean_array_push(x_1062, x_1250); x_1252 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1252, 0, x_1048); -lean_ctor_set(x_1252, 1, x_1250); +lean_ctor_set(x_1252, 0, x_1047); +lean_ctor_set(x_1252, 1, x_1064); lean_ctor_set(x_1252, 2, x_1251); -x_1253 = lean_array_push(x_1063, x_1252); -x_1254 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1254, 0, x_1048); -lean_ctor_set(x_1254, 1, x_1065); -lean_ctor_set(x_1254, 2, x_1253); -x_1255 = lean_array_push(x_1056, x_5); -lean_inc(x_1255); -x_1256 = lean_array_push(x_1255, x_1254); +x_1253 = lean_array_push(x_1055, x_6); +lean_inc(x_1253); +x_1254 = lean_array_push(x_1253, x_1252); +lean_inc(x_1248); +x_1255 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1255, 0, x_1047); +lean_ctor_set(x_1255, 1, x_1248); +lean_ctor_set(x_1255, 2, x_1254); +lean_inc(x_1231); +x_1256 = lean_array_push(x_1231, x_1255); +lean_inc(x_1227); +x_1257 = lean_array_push(x_1256, x_1227); +lean_inc(x_1203); +x_1258 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1258, 0, x_1047); +lean_ctor_set(x_1258, 1, x_1203); +lean_ctor_set(x_1258, 2, x_1257); +lean_inc(x_1237); +x_1259 = lean_array_push(x_1237, x_1258); +lean_inc(x_1227); +x_1260 = lean_array_push(x_1259, x_1227); +lean_inc(x_1197); +x_1261 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1261, 0, x_1047); +lean_ctor_set(x_1261, 1, x_1197); +lean_ctor_set(x_1261, 2, x_1260); +x_1262 = lean_array_push(x_1062, x_1261); lean_inc(x_1246); -x_1257 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1257, 0, x_1048); -lean_ctor_set(x_1257, 1, x_1246); -lean_ctor_set(x_1257, 2, x_1256); -lean_inc(x_1229); -x_1258 = lean_array_push(x_1229, x_1257); -lean_inc(x_1225); -x_1259 = lean_array_push(x_1258, x_1225); -lean_inc(x_1204); -x_1260 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1260, 0, x_1048); -lean_ctor_set(x_1260, 1, x_1204); -lean_ctor_set(x_1260, 2, x_1259); -lean_inc(x_1235); -x_1261 = lean_array_push(x_1235, x_1260); -lean_inc(x_1225); -x_1262 = lean_array_push(x_1261, x_1225); -lean_inc(x_1198); x_1263 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1263, 0, x_1048); -lean_ctor_set(x_1263, 1, x_1198); +lean_ctor_set(x_1263, 0, x_1047); +lean_ctor_set(x_1263, 1, x_1246); lean_ctor_set(x_1263, 2, x_1262); -x_1264 = lean_array_push(x_1063, x_1263); -lean_inc(x_1244); -x_1265 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1265, 0, x_1048); -lean_ctor_set(x_1265, 1, x_1244); -lean_ctor_set(x_1265, 2, x_1264); -x_1266 = lean_array_push(x_1056, x_1265); -x_1267 = lean_array_push(x_1266, x_1078); -lean_inc(x_1142); +x_1264 = lean_array_push(x_1055, x_1263); +x_1265 = lean_array_push(x_1264, x_1077); +lean_inc(x_1141); +x_1266 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1266, 0, x_1047); +lean_ctor_set(x_1266, 1, x_1141); +lean_ctor_set(x_1266, 2, x_1265); +x_1267 = lean_array_push(x_1062, x_1266); x_1268 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1268, 0, x_1048); -lean_ctor_set(x_1268, 1, x_1142); +lean_ctor_set(x_1268, 0, x_1047); +lean_ctor_set(x_1268, 1, x_1064); lean_ctor_set(x_1268, 2, x_1267); -x_1269 = lean_array_push(x_1063, x_1268); +x_1269 = lean_array_push(x_1062, x_1268); +lean_inc(x_1139); x_1270 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1270, 0, x_1048); -lean_ctor_set(x_1270, 1, x_1065); +lean_ctor_set(x_1270, 0, x_1047); +lean_ctor_set(x_1270, 1, x_1139); lean_ctor_set(x_1270, 2, x_1269); -x_1271 = lean_array_push(x_1063, x_1270); -lean_inc(x_1140); -x_1272 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1272, 0, x_1048); -lean_ctor_set(x_1272, 1, x_1140); -lean_ctor_set(x_1272, 2, x_1271); -x_1273 = lean_array_push(x_1170, x_1196); -lean_inc(x_1273); -x_1274 = lean_array_push(x_1273, x_1242); -lean_inc(x_1135); -x_1275 = lean_array_push(x_1274, x_1135); -x_1276 = lean_array_push(x_1275, x_1272); -lean_inc(x_1194); -x_1277 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1277, 0, x_1048); -lean_ctor_set(x_1277, 1, x_1194); -lean_ctor_set(x_1277, 2, x_1276); -x_1278 = l_Lean_mkHole___closed__7; -lean_inc(x_1039); -x_1279 = lean_name_mk_string(x_1039, x_1278); -x_1280 = l_Lean_Name_appendIndexAfter___closed__1; -lean_inc(x_1027); -x_1281 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1281, 0, x_1027); -lean_ctor_set(x_1281, 1, x_1280); -x_1282 = lean_array_push(x_1063, x_1281); +x_1271 = lean_array_push(x_1169, x_1195); +lean_inc(x_1271); +x_1272 = lean_array_push(x_1271, x_1244); +lean_inc(x_1134); +x_1273 = lean_array_push(x_1272, x_1134); +x_1274 = lean_array_push(x_1273, x_1270); +lean_inc(x_1193); +x_1275 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1275, 0, x_1047); +lean_ctor_set(x_1275, 1, x_1193); +lean_ctor_set(x_1275, 2, x_1274); +x_1276 = l_Lean_mkHole___closed__7; +lean_inc(x_1038); +x_1277 = lean_name_mk_string(x_1038, x_1276); +x_1278 = l_Lean_Name_appendIndexAfter___closed__1; +lean_inc(x_1026); +x_1279 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1279, 0, x_1026); +lean_ctor_set(x_1279, 1, x_1278); +x_1280 = lean_array_push(x_1062, x_1279); +x_1281 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1281, 0, x_1047); +lean_ctor_set(x_1281, 1, x_1277); +lean_ctor_set(x_1281, 2, x_1280); +x_1282 = lean_array_push(x_1062, x_1281); x_1283 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1283, 0, x_1048); -lean_ctor_set(x_1283, 1, x_1279); +lean_ctor_set(x_1283, 0, x_1047); +lean_ctor_set(x_1283, 1, x_1064); lean_ctor_set(x_1283, 2, x_1282); -x_1284 = lean_array_push(x_1063, x_1283); +x_1284 = lean_array_push(x_1062, x_1283); x_1285 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1285, 0, x_1048); -lean_ctor_set(x_1285, 1, x_1065); +lean_ctor_set(x_1285, 0, x_1047); +lean_ctor_set(x_1285, 1, x_1064); lean_ctor_set(x_1285, 2, x_1284); -x_1286 = lean_array_push(x_1063, x_1285); -x_1287 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1287, 0, x_1048); -lean_ctor_set(x_1287, 1, x_1065); -lean_ctor_set(x_1287, 2, x_1286); -x_1288 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__11; -lean_inc(x_1027); +x_1286 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__11; +lean_inc(x_1026); +x_1287 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1287, 0, x_1026); +lean_ctor_set(x_1287, 1, x_1286); +x_1288 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__35; +lean_inc(x_1026); x_1289 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1289, 0, x_1027); +lean_ctor_set(x_1289, 0, x_1026); lean_ctor_set(x_1289, 1, x_1288); -x_1290 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__35; -lean_inc(x_1027); -x_1291 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1291, 0, x_1027); -lean_ctor_set(x_1291, 1, x_1290); +lean_inc(x_1287); +x_1290 = lean_array_push(x_1055, x_1287); lean_inc(x_1289); -x_1292 = lean_array_push(x_1056, x_1289); -lean_inc(x_1291); -x_1293 = lean_array_push(x_1292, x_1291); -x_1294 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__68; -x_1295 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1295, 0, x_1048); -lean_ctor_set(x_1295, 1, x_1294); -lean_ctor_set(x_1295, 2, x_1293); -x_1296 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__9; -lean_inc(x_1039); -x_1297 = lean_name_mk_string(x_1039, x_1296); -x_1298 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__31; -lean_inc(x_1039); -x_1299 = lean_name_mk_string(x_1039, x_1298); -x_1300 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1300, 0, x_1048); -lean_ctor_set(x_1300, 1, x_1299); -lean_ctor_set(x_1300, 2, x_1049); -x_1301 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__36; -x_1302 = lean_array_push(x_1301, x_1289); -x_1303 = lean_array_push(x_1302, x_1078); -x_1304 = lean_array_push(x_1303, x_1078); -x_1305 = lean_array_push(x_1304, x_1300); -x_1306 = lean_array_push(x_1305, x_1078); -x_1307 = lean_array_push(x_1306, x_1291); -x_1308 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1308, 0, x_1048); -lean_ctor_set(x_1308, 1, x_1297); -lean_ctor_set(x_1308, 2, x_1307); -x_1309 = lean_array_push(x_1056, x_1295); -x_1310 = lean_array_push(x_1309, x_1308); -x_1311 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__66; +x_1291 = lean_array_push(x_1290, x_1289); +x_1292 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__65; +x_1293 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1293, 0, x_1047); +lean_ctor_set(x_1293, 1, x_1292); +lean_ctor_set(x_1293, 2, x_1291); +x_1294 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__9; +lean_inc(x_1038); +x_1295 = lean_name_mk_string(x_1038, x_1294); +x_1296 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__31; +lean_inc(x_1038); +x_1297 = lean_name_mk_string(x_1038, x_1296); +x_1298 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1298, 0, x_1047); +lean_ctor_set(x_1298, 1, x_1297); +lean_ctor_set(x_1298, 2, x_1048); +x_1299 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__36; +x_1300 = lean_array_push(x_1299, x_1287); +x_1301 = lean_array_push(x_1300, x_1077); +x_1302 = lean_array_push(x_1301, x_1077); +x_1303 = lean_array_push(x_1302, x_1298); +x_1304 = lean_array_push(x_1303, x_1077); +x_1305 = lean_array_push(x_1304, x_1289); +x_1306 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1306, 0, x_1047); +lean_ctor_set(x_1306, 1, x_1295); +lean_ctor_set(x_1306, 2, x_1305); +x_1307 = lean_array_push(x_1055, x_1293); +x_1308 = lean_array_push(x_1307, x_1306); +x_1309 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__63; +x_1310 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1310, 0, x_1047); +lean_ctor_set(x_1310, 1, x_1309); +lean_ctor_set(x_1310, 2, x_1308); +x_1311 = lean_array_push(x_1062, x_1310); x_1312 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1312, 0, x_1048); -lean_ctor_set(x_1312, 1, x_1311); -lean_ctor_set(x_1312, 2, x_1310); -x_1313 = lean_array_push(x_1063, x_1312); +lean_ctor_set(x_1312, 0, x_1047); +lean_ctor_set(x_1312, 1, x_1064); +lean_ctor_set(x_1312, 2, x_1311); +x_1313 = lean_array_push(x_1253, x_1312); +lean_inc(x_1248); x_1314 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1314, 0, x_1048); -lean_ctor_set(x_1314, 1, x_1065); +lean_ctor_set(x_1314, 0, x_1047); +lean_ctor_set(x_1314, 1, x_1248); lean_ctor_set(x_1314, 2, x_1313); -x_1315 = lean_array_push(x_1255, x_1314); -lean_inc(x_1246); -x_1316 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1316, 0, x_1048); -lean_ctor_set(x_1316, 1, x_1246); -lean_ctor_set(x_1316, 2, x_1315); -x_1317 = lean_array_push(x_1229, x_1316); -lean_inc(x_1225); -x_1318 = lean_array_push(x_1317, x_1225); -x_1319 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1319, 0, x_1048); -lean_ctor_set(x_1319, 1, x_1204); -lean_ctor_set(x_1319, 2, x_1318); -x_1320 = lean_array_push(x_1235, x_1319); -lean_inc(x_1225); -x_1321 = lean_array_push(x_1320, x_1225); +x_1315 = lean_array_push(x_1231, x_1314); +lean_inc(x_1227); +x_1316 = lean_array_push(x_1315, x_1227); +x_1317 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1317, 0, x_1047); +lean_ctor_set(x_1317, 1, x_1203); +lean_ctor_set(x_1317, 2, x_1316); +x_1318 = lean_array_push(x_1237, x_1317); +lean_inc(x_1227); +x_1319 = lean_array_push(x_1318, x_1227); +x_1320 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1320, 0, x_1047); +lean_ctor_set(x_1320, 1, x_1197); +lean_ctor_set(x_1320, 2, x_1319); +x_1321 = lean_array_push(x_1062, x_1320); x_1322 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1322, 0, x_1048); -lean_ctor_set(x_1322, 1, x_1198); +lean_ctor_set(x_1322, 0, x_1047); +lean_ctor_set(x_1322, 1, x_1246); lean_ctor_set(x_1322, 2, x_1321); -x_1323 = lean_array_push(x_1063, x_1322); -x_1324 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1324, 0, x_1048); -lean_ctor_set(x_1324, 1, x_1244); -lean_ctor_set(x_1324, 2, x_1323); -x_1325 = lean_array_push(x_1056, x_1324); -x_1326 = lean_array_push(x_1325, x_1078); -lean_inc(x_1142); +x_1323 = lean_array_push(x_1055, x_1322); +x_1324 = lean_array_push(x_1323, x_1077); +lean_inc(x_1141); +x_1325 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1325, 0, x_1047); +lean_ctor_set(x_1325, 1, x_1141); +lean_ctor_set(x_1325, 2, x_1324); +x_1326 = lean_array_push(x_1062, x_1325); x_1327 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1327, 0, x_1048); -lean_ctor_set(x_1327, 1, x_1142); +lean_ctor_set(x_1327, 0, x_1047); +lean_ctor_set(x_1327, 1, x_1064); lean_ctor_set(x_1327, 2, x_1326); -x_1328 = lean_array_push(x_1063, x_1327); +x_1328 = lean_array_push(x_1062, x_1327); +lean_inc(x_1139); x_1329 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1329, 0, x_1048); -lean_ctor_set(x_1329, 1, x_1065); +lean_ctor_set(x_1329, 0, x_1047); +lean_ctor_set(x_1329, 1, x_1139); lean_ctor_set(x_1329, 2, x_1328); -x_1330 = lean_array_push(x_1063, x_1329); -lean_inc(x_1140); -x_1331 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1331, 0, x_1048); -lean_ctor_set(x_1331, 1, x_1140); -lean_ctor_set(x_1331, 2, x_1330); -x_1332 = lean_array_push(x_1273, x_1287); -lean_inc(x_1135); -x_1333 = lean_array_push(x_1332, x_1135); -x_1334 = lean_array_push(x_1333, x_1331); -x_1335 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1335, 0, x_1048); -lean_ctor_set(x_1335, 1, x_1194); -lean_ctor_set(x_1335, 2, x_1334); -x_1336 = lean_array_push(x_1056, x_1277); -x_1337 = lean_array_push(x_1336, x_1335); +x_1330 = lean_array_push(x_1271, x_1285); +lean_inc(x_1134); +x_1331 = lean_array_push(x_1330, x_1134); +x_1332 = lean_array_push(x_1331, x_1329); +x_1333 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1333, 0, x_1047); +lean_ctor_set(x_1333, 1, x_1193); +lean_ctor_set(x_1333, 2, x_1332); +x_1334 = lean_array_push(x_1055, x_1275); +x_1335 = lean_array_push(x_1334, x_1333); +x_1336 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1336, 0, x_1047); +lean_ctor_set(x_1336, 1, x_1064); +lean_ctor_set(x_1336, 2, x_1335); +x_1337 = lean_array_push(x_1062, x_1336); x_1338 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1338, 0, x_1048); -lean_ctor_set(x_1338, 1, x_1065); +lean_ctor_set(x_1338, 0, x_1047); +lean_ctor_set(x_1338, 1, x_1191); lean_ctor_set(x_1338, 2, x_1337); -x_1339 = lean_array_push(x_1063, x_1338); -x_1340 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1340, 0, x_1048); -lean_ctor_set(x_1340, 1, x_1192); -lean_ctor_set(x_1340, 2, x_1339); -x_1341 = lean_array_push(x_1301, x_1158); -x_1342 = lean_array_push(x_1341, x_1078); -x_1343 = lean_array_push(x_1342, x_1078); -x_1344 = lean_array_push(x_1343, x_1188); -x_1345 = lean_array_push(x_1344, x_1190); -x_1346 = lean_array_push(x_1345, x_1340); -x_1347 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1347, 0, x_1048); -lean_ctor_set(x_1347, 1, x_1156); -lean_ctor_set(x_1347, 2, x_1346); -lean_inc(x_1152); -x_1348 = lean_array_push(x_1170, x_1152); -x_1349 = lean_array_push(x_1348, x_1078); -x_1350 = lean_array_push(x_1349, x_1154); -x_1351 = lean_array_push(x_1350, x_1347); -x_1352 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1352, 0, x_1048); -lean_ctor_set(x_1352, 1, x_1148); -lean_ctor_set(x_1352, 2, x_1351); -x_1353 = lean_array_push(x_1069, x_1146); -x_1354 = lean_array_push(x_1353, x_1078); -lean_inc(x_1354); -x_1355 = lean_array_push(x_1354, x_1352); -x_1356 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1356, 0, x_1048); -lean_ctor_set(x_1356, 1, x_1144); -lean_ctor_set(x_1356, 2, x_1355); -x_1357 = lean_array_push(x_1056, x_1356); -x_1358 = lean_array_push(x_1357, x_1078); -lean_inc(x_1142); -x_1359 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1359, 0, x_1048); -lean_ctor_set(x_1359, 1, x_1142); -lean_ctor_set(x_1359, 2, x_1358); -x_1360 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__69; -lean_inc(x_1039); -x_1361 = lean_name_mk_string(x_1039, x_1360); -x_1362 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__70; -lean_inc(x_1039); -x_1363 = lean_name_mk_string(x_1039, x_1362); -x_1364 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__71; -lean_inc(x_1039); -x_1365 = lean_name_mk_string(x_1039, x_1364); -x_1366 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__76; +x_1339 = lean_array_push(x_1299, x_1157); +x_1340 = lean_array_push(x_1339, x_1077); +x_1341 = lean_array_push(x_1340, x_1077); +x_1342 = lean_array_push(x_1341, x_1187); +x_1343 = lean_array_push(x_1342, x_1189); +x_1344 = lean_array_push(x_1343, x_1338); +x_1345 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1345, 0, x_1047); +lean_ctor_set(x_1345, 1, x_1155); +lean_ctor_set(x_1345, 2, x_1344); +lean_inc(x_1151); +x_1346 = lean_array_push(x_1169, x_1151); +x_1347 = lean_array_push(x_1346, x_1077); +x_1348 = lean_array_push(x_1347, x_1153); +x_1349 = lean_array_push(x_1348, x_1345); +x_1350 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1350, 0, x_1047); +lean_ctor_set(x_1350, 1, x_1147); +lean_ctor_set(x_1350, 2, x_1349); +x_1351 = lean_array_push(x_1068, x_1145); +x_1352 = lean_array_push(x_1351, x_1077); +lean_inc(x_1352); +x_1353 = lean_array_push(x_1352, x_1350); +x_1354 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1354, 0, x_1047); +lean_ctor_set(x_1354, 1, x_1143); +lean_ctor_set(x_1354, 2, x_1353); +x_1355 = lean_array_push(x_1055, x_1354); +x_1356 = lean_array_push(x_1355, x_1077); +lean_inc(x_1141); +x_1357 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1357, 0, x_1047); +lean_ctor_set(x_1357, 1, x_1141); +lean_ctor_set(x_1357, 2, x_1356); +x_1358 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__66; +lean_inc(x_1038); +x_1359 = lean_name_mk_string(x_1038, x_1358); +x_1360 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__67; +lean_inc(x_1038); +x_1361 = lean_name_mk_string(x_1038, x_1360); +x_1362 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__68; +lean_inc(x_1038); +x_1363 = lean_name_mk_string(x_1038, x_1362); +x_1364 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__73; +lean_inc(x_1029); lean_inc(x_1030); -lean_inc(x_1031); -x_1367 = l_Lean_addMacroScope(x_1031, x_1366, x_1030); -x_1368 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__74; -lean_inc(x_1027); -x_1369 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1369, 0, x_1027); -lean_ctor_set(x_1369, 1, x_1368); -lean_ctor_set(x_1369, 2, x_1367); -lean_ctor_set(x_1369, 3, x_1091); -x_1370 = lean_array_push(x_1063, x_1023); -x_1371 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1371, 0, x_1048); -lean_ctor_set(x_1371, 1, x_1065); -lean_ctor_set(x_1371, 2, x_1370); -x_1372 = lean_array_push(x_1056, x_1369); -x_1373 = lean_array_push(x_1372, x_1371); -lean_inc(x_1246); -x_1374 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1374, 0, x_1048); -lean_ctor_set(x_1374, 1, x_1246); -lean_ctor_set(x_1374, 2, x_1373); -x_1375 = lean_array_push(x_1226, x_1131); -x_1376 = lean_array_push(x_1375, x_1078); -x_1377 = lean_array_push(x_1376, x_1078); -lean_inc(x_1122); -x_1378 = lean_array_push(x_1377, x_1122); -lean_inc(x_1378); -x_1379 = lean_array_push(x_1378, x_1374); -lean_inc(x_1365); +x_1365 = l_Lean_addMacroScope(x_1030, x_1364, x_1029); +x_1366 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__71; +lean_inc(x_1026); +x_1367 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1367, 0, x_1026); +lean_ctor_set(x_1367, 1, x_1366); +lean_ctor_set(x_1367, 2, x_1365); +lean_ctor_set(x_1367, 3, x_1090); +x_1368 = lean_array_push(x_1062, x_1022); +x_1369 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1369, 0, x_1047); +lean_ctor_set(x_1369, 1, x_1064); +lean_ctor_set(x_1369, 2, x_1368); +x_1370 = lean_array_push(x_1055, x_1367); +x_1371 = lean_array_push(x_1370, x_1369); +lean_inc(x_1248); +x_1372 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1372, 0, x_1047); +lean_ctor_set(x_1372, 1, x_1248); +lean_ctor_set(x_1372, 2, x_1371); +x_1373 = lean_array_push(x_1228, x_1130); +x_1374 = lean_array_push(x_1373, x_1077); +x_1375 = lean_array_push(x_1374, x_1077); +lean_inc(x_1121); +x_1376 = lean_array_push(x_1375, x_1121); +lean_inc(x_1376); +x_1377 = lean_array_push(x_1376, x_1372); +lean_inc(x_1363); +x_1378 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1378, 0, x_1047); +lean_ctor_set(x_1378, 1, x_1363); +lean_ctor_set(x_1378, 2, x_1377); +x_1379 = lean_array_push(x_1062, x_1378); +lean_inc(x_1361); x_1380 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1380, 0, x_1048); -lean_ctor_set(x_1380, 1, x_1365); +lean_ctor_set(x_1380, 0, x_1047); +lean_ctor_set(x_1380, 1, x_1361); lean_ctor_set(x_1380, 2, x_1379); -x_1381 = lean_array_push(x_1063, x_1380); -lean_inc(x_1363); +lean_inc(x_1352); +x_1381 = lean_array_push(x_1352, x_1380); +lean_inc(x_1359); x_1382 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1382, 0, x_1048); -lean_ctor_set(x_1382, 1, x_1363); +lean_ctor_set(x_1382, 0, x_1047); +lean_ctor_set(x_1382, 1, x_1359); lean_ctor_set(x_1382, 2, x_1381); -lean_inc(x_1354); -x_1383 = lean_array_push(x_1354, x_1382); -lean_inc(x_1361); -x_1384 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1384, 0, x_1048); -lean_ctor_set(x_1384, 1, x_1361); -lean_ctor_set(x_1384, 2, x_1383); -x_1385 = lean_array_push(x_1056, x_1384); -x_1386 = lean_array_push(x_1385, x_1078); -lean_inc(x_1142); -x_1387 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1387, 0, x_1048); -lean_ctor_set(x_1387, 1, x_1142); -lean_ctor_set(x_1387, 2, x_1386); -x_1388 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__81; +x_1383 = lean_array_push(x_1055, x_1382); +x_1384 = lean_array_push(x_1383, x_1077); +lean_inc(x_1141); +x_1385 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1385, 0, x_1047); +lean_ctor_set(x_1385, 1, x_1141); +lean_ctor_set(x_1385, 2, x_1384); +x_1386 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__78; +lean_inc(x_1029); lean_inc(x_1030); -lean_inc(x_1031); -x_1389 = l_Lean_addMacroScope(x_1031, x_1388, x_1030); -x_1390 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__79; -lean_inc(x_1027); -x_1391 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1391, 0, x_1027); -lean_ctor_set(x_1391, 1, x_1390); -lean_ctor_set(x_1391, 2, x_1389); -lean_ctor_set(x_1391, 3, x_1091); -x_1392 = l_Lean_Syntax_expandInterpolatedStr___closed__2; -lean_inc(x_1039); -x_1393 = lean_name_mk_string(x_1039, x_1392); -x_1394 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__85; +x_1387 = l_Lean_addMacroScope(x_1030, x_1386, x_1029); +x_1388 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__76; +lean_inc(x_1026); +x_1389 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1389, 0, x_1026); +lean_ctor_set(x_1389, 1, x_1388); +lean_ctor_set(x_1389, 2, x_1387); +lean_ctor_set(x_1389, 3, x_1090); +x_1390 = l_Lean_TSyntax_expandInterpolatedStr___closed__2; +lean_inc(x_1038); +x_1391 = lean_name_mk_string(x_1038, x_1390); +x_1392 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__82; +lean_inc(x_1029); lean_inc(x_1030); -lean_inc(x_1031); -x_1395 = l_Lean_addMacroScope(x_1031, x_1394, x_1030); -x_1396 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__82; +x_1393 = l_Lean_addMacroScope(x_1030, x_1392, x_1029); +x_1394 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__79; lean_inc(x_3); -x_1397 = lean_name_mk_string(x_3, x_1396); -x_1398 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1398, 0, x_1397); -lean_ctor_set(x_1398, 1, x_1091); -x_1399 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1399, 0, x_1398); -lean_ctor_set(x_1399, 1, x_1091); -x_1400 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__84; -lean_inc(x_1027); -x_1401 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1401, 0, x_1027); -lean_ctor_set(x_1401, 1, x_1400); -lean_ctor_set(x_1401, 2, x_1395); -lean_ctor_set(x_1401, 3, x_1399); -lean_inc(x_1179); -x_1402 = lean_array_push(x_1172, x_1179); -lean_inc(x_1068); -x_1403 = lean_array_push(x_1402, x_1068); -x_1404 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1404, 0, x_1048); -lean_ctor_set(x_1404, 1, x_1162); -lean_ctor_set(x_1404, 2, x_1403); -x_1405 = lean_array_push(x_1056, x_1404); -x_1406 = lean_array_push(x_1405, x_1024); -x_1407 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1407, 0, x_1048); -lean_ctor_set(x_1407, 1, x_1065); -lean_ctor_set(x_1407, 2, x_1406); -x_1408 = lean_array_push(x_1056, x_1401); -x_1409 = lean_array_push(x_1408, x_1407); -lean_inc(x_1246); -x_1410 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1410, 0, x_1048); -lean_ctor_set(x_1410, 1, x_1246); -lean_ctor_set(x_1410, 2, x_1409); -x_1411 = lean_array_push(x_1056, x_1410); -x_1412 = lean_array_push(x_1411, x_1078); -x_1413 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1413, 0, x_1048); -lean_ctor_set(x_1413, 1, x_1065); -lean_ctor_set(x_1413, 2, x_1412); -x_1414 = lean_array_push(x_1069, x_1209); -lean_inc(x_1414); -x_1415 = lean_array_push(x_1414, x_1413); -lean_inc(x_1225); -x_1416 = lean_array_push(x_1415, x_1225); -lean_inc(x_1393); -x_1417 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1417, 0, x_1048); -lean_ctor_set(x_1417, 1, x_1393); -lean_ctor_set(x_1417, 2, x_1416); -x_1418 = lean_array_push(x_1056, x_1179); -x_1419 = lean_array_push(x_1418, x_1417); -x_1420 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1420, 0, x_1048); -lean_ctor_set(x_1420, 1, x_1065); -lean_ctor_set(x_1420, 2, x_1419); -x_1421 = lean_array_push(x_1056, x_1391); -lean_inc(x_1421); -x_1422 = lean_array_push(x_1421, x_1420); -lean_inc(x_1246); +x_1395 = lean_name_mk_string(x_3, x_1394); +x_1396 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1396, 0, x_1395); +lean_ctor_set(x_1396, 1, x_1090); +x_1397 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1397, 0, x_1396); +lean_ctor_set(x_1397, 1, x_1090); +x_1398 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__81; +lean_inc(x_1026); +x_1399 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1399, 0, x_1026); +lean_ctor_set(x_1399, 1, x_1398); +lean_ctor_set(x_1399, 2, x_1393); +lean_ctor_set(x_1399, 3, x_1397); +lean_inc(x_1178); +x_1400 = lean_array_push(x_1171, x_1178); +lean_inc(x_1067); +x_1401 = lean_array_push(x_1400, x_1067); +x_1402 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1402, 0, x_1047); +lean_ctor_set(x_1402, 1, x_1161); +lean_ctor_set(x_1402, 2, x_1401); +x_1403 = lean_array_push(x_1055, x_1402); +x_1404 = lean_array_push(x_1403, x_1023); +x_1405 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1405, 0, x_1047); +lean_ctor_set(x_1405, 1, x_1064); +lean_ctor_set(x_1405, 2, x_1404); +x_1406 = lean_array_push(x_1055, x_1399); +x_1407 = lean_array_push(x_1406, x_1405); +lean_inc(x_1248); +x_1408 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1408, 0, x_1047); +lean_ctor_set(x_1408, 1, x_1248); +lean_ctor_set(x_1408, 2, x_1407); +x_1409 = lean_array_push(x_1055, x_1408); +x_1410 = lean_array_push(x_1409, x_1077); +x_1411 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1411, 0, x_1047); +lean_ctor_set(x_1411, 1, x_1064); +lean_ctor_set(x_1411, 2, x_1410); +x_1412 = lean_array_push(x_1068, x_1208); +lean_inc(x_1412); +x_1413 = lean_array_push(x_1412, x_1411); +lean_inc(x_1227); +x_1414 = lean_array_push(x_1413, x_1227); +lean_inc(x_1391); +x_1415 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1415, 0, x_1047); +lean_ctor_set(x_1415, 1, x_1391); +lean_ctor_set(x_1415, 2, x_1414); +x_1416 = lean_array_push(x_1055, x_1178); +x_1417 = lean_array_push(x_1416, x_1415); +x_1418 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1418, 0, x_1047); +lean_ctor_set(x_1418, 1, x_1064); +lean_ctor_set(x_1418, 2, x_1417); +x_1419 = lean_array_push(x_1055, x_1389); +lean_inc(x_1419); +x_1420 = lean_array_push(x_1419, x_1418); +lean_inc(x_1248); +x_1421 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1421, 0, x_1047); +lean_ctor_set(x_1421, 1, x_1248); +lean_ctor_set(x_1421, 2, x_1420); +x_1422 = lean_array_push(x_1376, x_1421); +lean_inc(x_1363); x_1423 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1423, 0, x_1048); -lean_ctor_set(x_1423, 1, x_1246); +lean_ctor_set(x_1423, 0, x_1047); +lean_ctor_set(x_1423, 1, x_1363); lean_ctor_set(x_1423, 2, x_1422); -x_1424 = lean_array_push(x_1378, x_1423); -lean_inc(x_1365); +x_1424 = lean_array_push(x_1062, x_1423); +lean_inc(x_1361); x_1425 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1425, 0, x_1048); -lean_ctor_set(x_1425, 1, x_1365); +lean_ctor_set(x_1425, 0, x_1047); +lean_ctor_set(x_1425, 1, x_1361); lean_ctor_set(x_1425, 2, x_1424); -x_1426 = lean_array_push(x_1063, x_1425); -lean_inc(x_1363); +lean_inc(x_1352); +x_1426 = lean_array_push(x_1352, x_1425); +lean_inc(x_1359); x_1427 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1427, 0, x_1048); -lean_ctor_set(x_1427, 1, x_1363); +lean_ctor_set(x_1427, 0, x_1047); +lean_ctor_set(x_1427, 1, x_1359); lean_ctor_set(x_1427, 2, x_1426); -lean_inc(x_1354); -x_1428 = lean_array_push(x_1354, x_1427); -lean_inc(x_1361); -x_1429 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1429, 0, x_1048); -lean_ctor_set(x_1429, 1, x_1361); -lean_ctor_set(x_1429, 2, x_1428); -x_1430 = lean_array_push(x_1056, x_1429); -x_1431 = lean_array_push(x_1430, x_1078); -lean_inc(x_1142); -x_1432 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1432, 0, x_1048); -lean_ctor_set(x_1432, 1, x_1142); -lean_ctor_set(x_1432, 2, x_1431); -x_1433 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__89; +x_1428 = lean_array_push(x_1055, x_1427); +x_1429 = lean_array_push(x_1428, x_1077); +lean_inc(x_1141); +x_1430 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1430, 0, x_1047); +lean_ctor_set(x_1430, 1, x_1141); +lean_ctor_set(x_1430, 2, x_1429); +x_1431 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__86; +lean_inc(x_1029); lean_inc(x_1030); -lean_inc(x_1031); -x_1434 = l_Lean_addMacroScope(x_1031, x_1433, x_1030); -x_1435 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__88; -lean_inc(x_1027); -x_1436 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1436, 0, x_1027); -lean_ctor_set(x_1436, 1, x_1435); -lean_ctor_set(x_1436, 2, x_1434); -lean_ctor_set(x_1436, 3, x_1091); -x_1437 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__93; -x_1438 = l_Lean_addMacroScope(x_1031, x_1437, x_1030); -x_1439 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__90; -x_1440 = lean_name_mk_string(x_3, x_1439); -x_1441 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1441, 0, x_1440); -lean_ctor_set(x_1441, 1, x_1091); -x_1442 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1442, 0, x_1441); -lean_ctor_set(x_1442, 1, x_1091); -x_1443 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__92; -lean_inc(x_1027); -x_1444 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1444, 0, x_1027); +x_1432 = l_Lean_addMacroScope(x_1030, x_1431, x_1029); +x_1433 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__85; +lean_inc(x_1026); +x_1434 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1434, 0, x_1026); +lean_ctor_set(x_1434, 1, x_1433); +lean_ctor_set(x_1434, 2, x_1432); +lean_ctor_set(x_1434, 3, x_1090); +x_1435 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__90; +x_1436 = l_Lean_addMacroScope(x_1030, x_1435, x_1029); +x_1437 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__87; +x_1438 = lean_name_mk_string(x_3, x_1437); +x_1439 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1439, 0, x_1438); +lean_ctor_set(x_1439, 1, x_1090); +x_1440 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1440, 0, x_1439); +lean_ctor_set(x_1440, 1, x_1090); +x_1441 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__89; +lean_inc(x_1026); +x_1442 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1442, 0, x_1026); +lean_ctor_set(x_1442, 1, x_1441); +lean_ctor_set(x_1442, 2, x_1436); +lean_ctor_set(x_1442, 3, x_1440); +x_1443 = l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__12; +lean_inc(x_1026); +x_1444 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1444, 0, x_1026); lean_ctor_set(x_1444, 1, x_1443); -lean_ctor_set(x_1444, 2, x_1438); -lean_ctor_set(x_1444, 3, x_1442); -x_1445 = l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__12; -lean_inc(x_1027); -x_1446 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1446, 0, x_1027); -lean_ctor_set(x_1446, 1, x_1445); -x_1447 = lean_array_push(x_1063, x_1152); -x_1448 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1448, 0, x_1048); -lean_ctor_set(x_1448, 1, x_1065); -lean_ctor_set(x_1448, 2, x_1447); -x_1449 = lean_array_push(x_1069, x_1446); -x_1450 = lean_array_push(x_1449, x_1448); -x_1451 = lean_array_push(x_1450, x_1068); -x_1452 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__95; +x_1445 = lean_array_push(x_1062, x_1151); +x_1446 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1446, 0, x_1047); +lean_ctor_set(x_1446, 1, x_1064); +lean_ctor_set(x_1446, 2, x_1445); +x_1447 = lean_array_push(x_1068, x_1444); +x_1448 = lean_array_push(x_1447, x_1446); +x_1449 = lean_array_push(x_1448, x_1067); +x_1450 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__92; +x_1451 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1451, 0, x_1047); +lean_ctor_set(x_1451, 1, x_1450); +lean_ctor_set(x_1451, 2, x_1449); +x_1452 = lean_array_push(x_1062, x_1451); x_1453 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1453, 0, x_1048); -lean_ctor_set(x_1453, 1, x_1452); -lean_ctor_set(x_1453, 2, x_1451); -x_1454 = lean_array_push(x_1063, x_1453); -x_1455 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1455, 0, x_1048); -lean_ctor_set(x_1455, 1, x_1065); -lean_ctor_set(x_1455, 2, x_1454); -x_1456 = lean_array_push(x_1056, x_1444); -x_1457 = lean_array_push(x_1456, x_1455); -lean_inc(x_1246); -x_1458 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1458, 0, x_1048); -lean_ctor_set(x_1458, 1, x_1246); -lean_ctor_set(x_1458, 2, x_1457); -x_1459 = lean_array_push(x_1056, x_1458); -x_1460 = lean_array_push(x_1459, x_1078); -x_1461 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1461, 0, x_1048); -lean_ctor_set(x_1461, 1, x_1065); -lean_ctor_set(x_1461, 2, x_1460); -x_1462 = lean_array_push(x_1414, x_1461); -x_1463 = lean_array_push(x_1462, x_1225); -x_1464 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1464, 0, x_1048); -lean_ctor_set(x_1464, 1, x_1393); -lean_ctor_set(x_1464, 2, x_1463); -x_1465 = lean_array_push(x_1056, x_1169); -x_1466 = lean_array_push(x_1465, x_1464); +lean_ctor_set(x_1453, 0, x_1047); +lean_ctor_set(x_1453, 1, x_1064); +lean_ctor_set(x_1453, 2, x_1452); +x_1454 = lean_array_push(x_1055, x_1442); +x_1455 = lean_array_push(x_1454, x_1453); +lean_inc(x_1248); +x_1456 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1456, 0, x_1047); +lean_ctor_set(x_1456, 1, x_1248); +lean_ctor_set(x_1456, 2, x_1455); +x_1457 = lean_array_push(x_1055, x_1456); +x_1458 = lean_array_push(x_1457, x_1077); +x_1459 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1459, 0, x_1047); +lean_ctor_set(x_1459, 1, x_1064); +lean_ctor_set(x_1459, 2, x_1458); +x_1460 = lean_array_push(x_1412, x_1459); +x_1461 = lean_array_push(x_1460, x_1227); +x_1462 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1462, 0, x_1047); +lean_ctor_set(x_1462, 1, x_1391); +lean_ctor_set(x_1462, 2, x_1461); +x_1463 = lean_array_push(x_1055, x_1168); +x_1464 = lean_array_push(x_1463, x_1462); +x_1465 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1465, 0, x_1047); +lean_ctor_set(x_1465, 1, x_1064); +lean_ctor_set(x_1465, 2, x_1464); +x_1466 = lean_array_push(x_1419, x_1465); x_1467 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1467, 0, x_1048); -lean_ctor_set(x_1467, 1, x_1065); +lean_ctor_set(x_1467, 0, x_1047); +lean_ctor_set(x_1467, 1, x_1248); lean_ctor_set(x_1467, 2, x_1466); -x_1468 = lean_array_push(x_1421, x_1467); -x_1469 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1469, 0, x_1048); -lean_ctor_set(x_1469, 1, x_1246); -lean_ctor_set(x_1469, 2, x_1468); -lean_inc(x_1436); -x_1470 = lean_array_push(x_1226, x_1436); -x_1471 = lean_array_push(x_1470, x_1078); -x_1472 = lean_array_push(x_1471, x_1078); -lean_inc(x_1122); -x_1473 = lean_array_push(x_1472, x_1122); -x_1474 = lean_array_push(x_1473, x_1469); +lean_inc(x_1434); +x_1468 = lean_array_push(x_1228, x_1434); +x_1469 = lean_array_push(x_1468, x_1077); +x_1470 = lean_array_push(x_1469, x_1077); +lean_inc(x_1121); +x_1471 = lean_array_push(x_1470, x_1121); +x_1472 = lean_array_push(x_1471, x_1467); +x_1473 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1473, 0, x_1047); +lean_ctor_set(x_1473, 1, x_1363); +lean_ctor_set(x_1473, 2, x_1472); +x_1474 = lean_array_push(x_1062, x_1473); x_1475 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1475, 0, x_1048); -lean_ctor_set(x_1475, 1, x_1365); +lean_ctor_set(x_1475, 0, x_1047); +lean_ctor_set(x_1475, 1, x_1361); lean_ctor_set(x_1475, 2, x_1474); -x_1476 = lean_array_push(x_1063, x_1475); +x_1476 = lean_array_push(x_1352, x_1475); x_1477 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1477, 0, x_1048); -lean_ctor_set(x_1477, 1, x_1363); +lean_ctor_set(x_1477, 0, x_1047); +lean_ctor_set(x_1477, 1, x_1359); lean_ctor_set(x_1477, 2, x_1476); -x_1478 = lean_array_push(x_1354, x_1477); -x_1479 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1479, 0, x_1048); -lean_ctor_set(x_1479, 1, x_1361); -lean_ctor_set(x_1479, 2, x_1478); -x_1480 = lean_array_push(x_1056, x_1479); -x_1481 = lean_array_push(x_1480, x_1078); -lean_inc(x_1142); -x_1482 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1482, 0, x_1048); -lean_ctor_set(x_1482, 1, x_1142); -lean_ctor_set(x_1482, 2, x_1481); -x_1483 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__96; -x_1484 = lean_name_mk_string(x_1039, x_1483); -x_1485 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__97; -x_1486 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1486, 0, x_1027); -lean_ctor_set(x_1486, 1, x_1485); -x_1487 = lean_array_push(x_1063, x_1436); -x_1488 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1488, 0, x_1048); -lean_ctor_set(x_1488, 1, x_1065); -lean_ctor_set(x_1488, 2, x_1487); -x_1489 = lean_array_push(x_1056, x_1486); -x_1490 = lean_array_push(x_1489, x_1488); -x_1491 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1491, 0, x_1048); -lean_ctor_set(x_1491, 1, x_1484); -lean_ctor_set(x_1491, 2, x_1490); -x_1492 = lean_array_push(x_1056, x_1491); -x_1493 = lean_array_push(x_1492, x_1078); -x_1494 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1494, 0, x_1048); -lean_ctor_set(x_1494, 1, x_1142); -lean_ctor_set(x_1494, 2, x_1493); -x_1495 = lean_array_push(x_1226, x_1359); -x_1496 = lean_array_push(x_1495, x_1387); -x_1497 = lean_array_push(x_1496, x_1432); -x_1498 = lean_array_push(x_1497, x_1482); -x_1499 = lean_array_push(x_1498, x_1494); +x_1478 = lean_array_push(x_1055, x_1477); +x_1479 = lean_array_push(x_1478, x_1077); +lean_inc(x_1141); +x_1480 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1480, 0, x_1047); +lean_ctor_set(x_1480, 1, x_1141); +lean_ctor_set(x_1480, 2, x_1479); +x_1481 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__93; +x_1482 = lean_name_mk_string(x_1038, x_1481); +x_1483 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__94; +x_1484 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1484, 0, x_1026); +lean_ctor_set(x_1484, 1, x_1483); +x_1485 = lean_array_push(x_1062, x_1434); +x_1486 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1486, 0, x_1047); +lean_ctor_set(x_1486, 1, x_1064); +lean_ctor_set(x_1486, 2, x_1485); +x_1487 = lean_array_push(x_1055, x_1484); +x_1488 = lean_array_push(x_1487, x_1486); +x_1489 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1489, 0, x_1047); +lean_ctor_set(x_1489, 1, x_1482); +lean_ctor_set(x_1489, 2, x_1488); +x_1490 = lean_array_push(x_1055, x_1489); +x_1491 = lean_array_push(x_1490, x_1077); +x_1492 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1492, 0, x_1047); +lean_ctor_set(x_1492, 1, x_1141); +lean_ctor_set(x_1492, 2, x_1491); +x_1493 = lean_array_push(x_1228, x_1357); +x_1494 = lean_array_push(x_1493, x_1385); +x_1495 = lean_array_push(x_1494, x_1430); +x_1496 = lean_array_push(x_1495, x_1480); +x_1497 = lean_array_push(x_1496, x_1492); +x_1498 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1498, 0, x_1047); +lean_ctor_set(x_1498, 1, x_1064); +lean_ctor_set(x_1498, 2, x_1497); +x_1499 = lean_array_push(x_1062, x_1498); x_1500 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1500, 0, x_1048); -lean_ctor_set(x_1500, 1, x_1065); +lean_ctor_set(x_1500, 0, x_1047); +lean_ctor_set(x_1500, 1, x_1139); lean_ctor_set(x_1500, 2, x_1499); -x_1501 = lean_array_push(x_1063, x_1500); -x_1502 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1502, 0, x_1048); -lean_ctor_set(x_1502, 1, x_1140); -lean_ctor_set(x_1502, 2, x_1501); -x_1503 = lean_array_push(x_1056, x_1138); -x_1504 = lean_array_push(x_1503, x_1502); -x_1505 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1505, 0, x_1048); -lean_ctor_set(x_1505, 1, x_1137); -lean_ctor_set(x_1505, 2, x_1504); -x_1506 = lean_array_push(x_1069, x_1133); -x_1507 = lean_array_push(x_1506, x_1135); -x_1508 = lean_array_push(x_1507, x_1505); -x_1509 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1509, 0, x_1048); -lean_ctor_set(x_1509, 1, x_1127); -lean_ctor_set(x_1509, 2, x_1508); -x_1510 = lean_array_push(x_1056, x_1125); -x_1511 = lean_array_push(x_1510, x_1509); -x_1512 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1512, 0, x_1048); -lean_ctor_set(x_1512, 1, x_1124); -lean_ctor_set(x_1512, 2, x_1511); -x_1513 = lean_array_push(x_1069, x_1122); -x_1514 = lean_array_push(x_1513, x_1512); -x_1515 = lean_array_push(x_1514, x_1078); -x_1516 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1516, 0, x_1048); -lean_ctor_set(x_1516, 1, x_1120); -lean_ctor_set(x_1516, 2, x_1515); -x_1517 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__98; -x_1518 = lean_array_push(x_1517, x_1086); -x_1519 = lean_array_push(x_1518, x_1096); -x_1520 = lean_array_push(x_1519, x_1118); -x_1521 = lean_array_push(x_1520, x_1516); -x_1522 = lean_array_push(x_1521, x_1078); -x_1523 = lean_array_push(x_1522, x_1078); -x_1524 = lean_array_push(x_1523, x_1078); -x_1525 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1525, 0, x_1048); -lean_ctor_set(x_1525, 1, x_1085); -lean_ctor_set(x_1525, 2, x_1524); -x_1526 = lean_array_push(x_1056, x_1083); -x_1527 = lean_array_push(x_1526, x_1525); -x_1528 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1528, 0, x_1048); -lean_ctor_set(x_1528, 1, x_1035); -lean_ctor_set(x_1528, 2, x_1527); -x_1529 = lean_array_push(x_1056, x_1025); -x_1530 = lean_array_push(x_1529, x_1528); -x_1531 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1531, 0, x_1048); -lean_ctor_set(x_1531, 1, x_1065); -lean_ctor_set(x_1531, 2, x_1530); -if (lean_is_scalar(x_1029)) { - x_1532 = lean_alloc_ctor(0, 2, 0); +x_1501 = lean_array_push(x_1055, x_1137); +x_1502 = lean_array_push(x_1501, x_1500); +x_1503 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1503, 0, x_1047); +lean_ctor_set(x_1503, 1, x_1136); +lean_ctor_set(x_1503, 2, x_1502); +x_1504 = lean_array_push(x_1068, x_1132); +x_1505 = lean_array_push(x_1504, x_1134); +x_1506 = lean_array_push(x_1505, x_1503); +x_1507 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1507, 0, x_1047); +lean_ctor_set(x_1507, 1, x_1126); +lean_ctor_set(x_1507, 2, x_1506); +x_1508 = lean_array_push(x_1055, x_1124); +x_1509 = lean_array_push(x_1508, x_1507); +x_1510 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1510, 0, x_1047); +lean_ctor_set(x_1510, 1, x_1123); +lean_ctor_set(x_1510, 2, x_1509); +x_1511 = lean_array_push(x_1068, x_1121); +x_1512 = lean_array_push(x_1511, x_1510); +x_1513 = lean_array_push(x_1512, x_1077); +x_1514 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1514, 0, x_1047); +lean_ctor_set(x_1514, 1, x_1119); +lean_ctor_set(x_1514, 2, x_1513); +x_1515 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__95; +x_1516 = lean_array_push(x_1515, x_1085); +x_1517 = lean_array_push(x_1516, x_1095); +x_1518 = lean_array_push(x_1517, x_1117); +x_1519 = lean_array_push(x_1518, x_1514); +x_1520 = lean_array_push(x_1519, x_1077); +x_1521 = lean_array_push(x_1520, x_1077); +x_1522 = lean_array_push(x_1521, x_1077); +x_1523 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1523, 0, x_1047); +lean_ctor_set(x_1523, 1, x_1084); +lean_ctor_set(x_1523, 2, x_1522); +x_1524 = lean_array_push(x_1055, x_1082); +x_1525 = lean_array_push(x_1524, x_1523); +x_1526 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1526, 0, x_1047); +lean_ctor_set(x_1526, 1, x_1034); +lean_ctor_set(x_1526, 2, x_1525); +x_1527 = lean_array_push(x_1055, x_1024); +x_1528 = lean_array_push(x_1527, x_1526); +x_1529 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1529, 0, x_1047); +lean_ctor_set(x_1529, 1, x_1064); +lean_ctor_set(x_1529, 2, x_1528); +if (lean_is_scalar(x_1028)) { + x_1530 = lean_alloc_ctor(0, 2, 0); } else { - x_1532 = x_1029; + x_1530 = x_1028; } -lean_ctor_set(x_1532, 0, x_1531); -lean_ctor_set(x_1532, 1, x_1028); -return x_1532; +lean_ctor_set(x_1530, 0, x_1529); +lean_ctor_set(x_1530, 1, x_1027); +return x_1530; } } else { -lean_object* x_1533; lean_object* x_1534; lean_object* x_1535; lean_object* x_1536; lean_object* x_1537; lean_object* x_1538; lean_object* x_1539; lean_object* x_1540; lean_object* x_1541; lean_object* x_1542; lean_object* x_1543; lean_object* x_1544; lean_object* x_1545; lean_object* x_1546; lean_object* x_1547; lean_object* x_1548; lean_object* x_1549; lean_object* x_1550; lean_object* x_1551; lean_object* x_1552; lean_object* x_1553; lean_object* x_1554; lean_object* x_1555; lean_object* x_1556; lean_object* x_1557; lean_object* x_1558; lean_object* x_1559; lean_object* x_1560; lean_object* x_1561; lean_object* x_1562; lean_object* x_1563; lean_object* x_1564; lean_object* x_1565; lean_object* x_1566; lean_object* x_1567; lean_object* x_1568; lean_object* x_1569; lean_object* x_1570; lean_object* x_1571; lean_object* x_1572; lean_object* x_1573; lean_object* x_1574; lean_object* x_1575; lean_object* x_1576; lean_object* x_1577; lean_object* x_1578; lean_object* x_1579; lean_object* x_1580; lean_object* x_1581; lean_object* x_1582; lean_object* x_1583; lean_object* x_1584; lean_object* x_1585; lean_object* x_1586; lean_object* x_1587; lean_object* x_1588; lean_object* x_1589; lean_object* x_1590; lean_object* x_1591; lean_object* x_1592; lean_object* x_1593; lean_object* x_1594; lean_object* x_1595; lean_object* x_1596; lean_object* x_1597; lean_object* x_1598; lean_object* x_1599; lean_object* x_1600; lean_object* x_1601; lean_object* x_1602; lean_object* x_1603; lean_object* x_1604; lean_object* x_1605; lean_object* x_1606; lean_object* x_1607; lean_object* x_1608; lean_object* x_1609; lean_object* x_1610; lean_object* x_1611; lean_object* x_1612; lean_object* x_1613; lean_object* x_1614; lean_object* x_1615; lean_object* x_1616; lean_object* x_1617; lean_object* x_1618; lean_object* x_1619; lean_object* x_1620; lean_object* x_1621; lean_object* x_1622; lean_object* x_1623; lean_object* x_1624; lean_object* x_1625; lean_object* x_1626; lean_object* x_1627; lean_object* x_1628; lean_object* x_1629; lean_object* x_1630; lean_object* x_1631; lean_object* x_1632; lean_object* x_1633; lean_object* x_1634; lean_object* x_1635; lean_object* x_1636; lean_object* x_1637; lean_object* x_1638; lean_object* x_1639; lean_object* x_1640; lean_object* x_1641; lean_object* x_1642; lean_object* x_1643; lean_object* x_1644; lean_object* x_1645; lean_object* x_1646; lean_object* x_1647; lean_object* x_1648; lean_object* x_1649; lean_object* x_1650; lean_object* x_1651; lean_object* x_1652; lean_object* x_1653; lean_object* x_1654; lean_object* x_1655; lean_object* x_1656; lean_object* x_1657; lean_object* x_1658; lean_object* x_1659; lean_object* x_1660; lean_object* x_1661; lean_object* x_1662; lean_object* x_1663; lean_object* x_1664; lean_object* x_1665; lean_object* x_1666; lean_object* x_1667; lean_object* x_1668; lean_object* x_1669; lean_object* x_1670; lean_object* x_1671; lean_object* x_1672; lean_object* x_1673; lean_object* x_1674; lean_object* x_1675; lean_object* x_1676; lean_object* x_1677; lean_object* x_1678; lean_object* x_1679; lean_object* x_1680; lean_object* x_1681; lean_object* x_1682; lean_object* x_1683; lean_object* x_1684; lean_object* x_1685; lean_object* x_1686; lean_object* x_1687; lean_object* x_1688; lean_object* x_1689; lean_object* x_1690; lean_object* x_1691; lean_object* x_1692; lean_object* x_1693; lean_object* x_1694; lean_object* x_1695; lean_object* x_1696; lean_object* x_1697; lean_object* x_1698; lean_object* x_1699; lean_object* x_1700; lean_object* x_1701; lean_object* x_1702; lean_object* x_1703; lean_object* x_1704; lean_object* x_1705; lean_object* x_1706; lean_object* x_1707; lean_object* x_1708; lean_object* x_1709; lean_object* x_1710; lean_object* x_1711; lean_object* x_1712; lean_object* x_1713; lean_object* x_1714; lean_object* x_1715; lean_object* x_1716; lean_object* x_1717; lean_object* x_1718; lean_object* x_1719; lean_object* x_1720; lean_object* x_1721; lean_object* x_1722; lean_object* x_1723; lean_object* x_1724; lean_object* x_1725; lean_object* x_1726; lean_object* x_1727; lean_object* x_1728; lean_object* x_1729; lean_object* x_1730; lean_object* x_1731; lean_object* x_1732; lean_object* x_1733; lean_object* x_1734; lean_object* x_1735; lean_object* x_1736; lean_object* x_1737; lean_object* x_1738; lean_object* x_1739; lean_object* x_1740; lean_object* x_1741; lean_object* x_1742; lean_object* x_1743; lean_object* x_1744; lean_object* x_1745; lean_object* x_1746; lean_object* x_1747; lean_object* x_1748; lean_object* x_1749; lean_object* x_1750; lean_object* x_1751; lean_object* x_1752; lean_object* x_1753; lean_object* x_1754; lean_object* x_1755; lean_object* x_1756; lean_object* x_1757; lean_object* x_1758; lean_object* x_1759; lean_object* x_1760; lean_object* x_1761; lean_object* x_1762; lean_object* x_1763; lean_object* x_1764; lean_object* x_1765; lean_object* x_1766; lean_object* x_1767; lean_object* x_1768; lean_object* x_1769; lean_object* x_1770; lean_object* x_1771; lean_object* x_1772; lean_object* x_1773; lean_object* x_1774; lean_object* x_1775; lean_object* x_1776; lean_object* x_1777; lean_object* x_1778; lean_object* x_1779; lean_object* x_1780; lean_object* x_1781; lean_object* x_1782; lean_object* x_1783; lean_object* x_1784; lean_object* x_1785; lean_object* x_1786; lean_object* x_1787; lean_object* x_1788; lean_object* x_1789; lean_object* x_1790; lean_object* x_1791; lean_object* x_1792; lean_object* x_1793; lean_object* x_1794; lean_object* x_1795; lean_object* x_1796; lean_object* x_1797; lean_object* x_1798; lean_object* x_1799; lean_object* x_1800; lean_object* x_1801; lean_object* x_1802; lean_object* x_1803; lean_object* x_1804; lean_object* x_1805; lean_object* x_1806; lean_object* x_1807; lean_object* x_1808; lean_object* x_1809; lean_object* x_1810; lean_object* x_1811; lean_object* x_1812; lean_object* x_1813; lean_object* x_1814; lean_object* x_1815; lean_object* x_1816; lean_object* x_1817; lean_object* x_1818; lean_object* x_1819; lean_object* x_1820; lean_object* x_1821; lean_object* x_1822; lean_object* x_1823; lean_object* x_1824; lean_object* x_1825; lean_object* x_1826; lean_object* x_1827; lean_object* x_1828; lean_object* x_1829; lean_object* x_1830; lean_object* x_1831; lean_object* x_1832; lean_object* x_1833; lean_object* x_1834; lean_object* x_1835; lean_object* x_1836; lean_object* x_1837; lean_object* x_1838; lean_object* x_1839; lean_object* x_1840; lean_object* x_1841; lean_object* x_1842; lean_object* x_1843; lean_object* x_1844; lean_object* x_1845; lean_object* x_1846; lean_object* x_1847; lean_object* x_1848; lean_object* x_1849; lean_object* x_1850; lean_object* x_1851; lean_object* x_1852; lean_object* x_1853; lean_object* x_1854; lean_object* x_1855; lean_object* x_1856; lean_object* x_1857; lean_object* x_1858; lean_object* x_1859; lean_object* x_1860; lean_object* x_1861; lean_object* x_1862; lean_object* x_1863; lean_object* x_1864; lean_object* x_1865; lean_object* x_1866; lean_object* x_1867; lean_object* x_1868; lean_object* x_1869; lean_object* x_1870; lean_object* x_1871; lean_object* x_1872; lean_object* x_1873; lean_object* x_1874; lean_object* x_1875; lean_object* x_1876; lean_object* x_1877; lean_object* x_1878; lean_object* x_1879; lean_object* x_1880; lean_object* x_1881; lean_object* x_1882; lean_object* x_1883; lean_object* x_1884; lean_object* x_1885; lean_object* x_1886; lean_object* x_1887; lean_object* x_1888; lean_object* x_1889; lean_object* x_1890; lean_object* x_1891; lean_object* x_1892; lean_object* x_1893; lean_object* x_1894; lean_object* x_1895; lean_object* x_1896; lean_object* x_1897; lean_object* x_1898; lean_object* x_1899; lean_object* x_1900; lean_object* x_1901; lean_object* x_1902; lean_object* x_1903; lean_object* x_1904; lean_object* x_1905; lean_object* x_1906; lean_object* x_1907; lean_object* x_1908; lean_object* x_1909; lean_object* x_1910; lean_object* x_1911; lean_object* x_1912; lean_object* x_1913; lean_object* x_1914; lean_object* x_1915; lean_object* x_1916; lean_object* x_1917; lean_object* x_1918; lean_object* x_1919; lean_object* x_1920; lean_object* x_1921; lean_object* x_1922; lean_object* x_1923; lean_object* x_1924; lean_object* x_1925; lean_object* x_1926; lean_object* x_1927; lean_object* x_1928; lean_object* x_1929; lean_object* x_1930; lean_object* x_1931; lean_object* x_1932; lean_object* x_1933; lean_object* x_1934; lean_object* x_1935; lean_object* x_1936; lean_object* x_1937; lean_object* x_1938; lean_object* x_1939; lean_object* x_1940; lean_object* x_1941; lean_object* x_1942; lean_object* x_1943; lean_object* x_1944; lean_object* x_1945; lean_object* x_1946; lean_object* x_1947; lean_object* x_1948; lean_object* x_1949; lean_object* x_1950; lean_object* x_1951; lean_object* x_1952; lean_object* x_1953; lean_object* x_1954; lean_object* x_1955; lean_object* x_1956; lean_object* x_1957; lean_object* x_1958; lean_object* x_1959; lean_object* x_1960; lean_object* x_1961; lean_object* x_1962; lean_object* x_1963; lean_object* x_1964; lean_object* x_1965; lean_object* x_1966; lean_object* x_1967; lean_object* x_1968; lean_object* x_1969; lean_object* x_1970; lean_object* x_1971; lean_object* x_1972; lean_object* x_1973; lean_object* x_1974; lean_object* x_1975; lean_object* x_1976; lean_object* x_1977; lean_object* x_1978; lean_object* x_1979; lean_object* x_1980; lean_object* x_1981; lean_object* x_1982; lean_object* x_1983; lean_object* x_1984; lean_object* x_1985; lean_object* x_1986; lean_object* x_1987; lean_object* x_1988; lean_object* x_1989; lean_object* x_1990; lean_object* x_1991; lean_object* x_1992; lean_object* x_1993; lean_object* x_1994; lean_object* x_1995; lean_object* x_1996; lean_object* x_1997; lean_object* x_1998; lean_object* x_1999; lean_object* x_2000; lean_object* x_2001; lean_object* x_2002; lean_object* x_2003; lean_object* x_2004; lean_object* x_2005; lean_object* x_2006; lean_object* x_2007; lean_object* x_2008; lean_object* x_2009; lean_object* x_2010; lean_object* x_2011; lean_object* x_2012; lean_object* x_2013; lean_object* x_2014; lean_object* x_2015; lean_object* x_2016; lean_object* x_2017; lean_object* x_2018; lean_object* x_2019; lean_object* x_2020; lean_object* x_2021; lean_object* x_2022; lean_object* x_2023; lean_object* x_2024; lean_object* x_2025; lean_object* x_2026; lean_object* x_2027; lean_object* x_2028; lean_object* x_2029; lean_object* x_2030; lean_object* x_2031; lean_object* x_2032; lean_object* x_2033; lean_object* x_2034; lean_object* x_2035; lean_object* x_2036; lean_object* x_2037; lean_object* x_2038; lean_object* x_2039; lean_object* x_2040; lean_object* x_2041; lean_object* x_2042; lean_object* x_2043; lean_object* x_2044; lean_object* x_2045; -x_1533 = lean_ctor_get(x_6, 1); -x_1534 = lean_ctor_get(x_6, 0); +lean_object* x_1531; lean_object* x_1532; lean_object* x_1533; lean_object* x_1534; lean_object* x_1535; lean_object* x_1536; lean_object* x_1537; lean_object* x_1538; lean_object* x_1539; lean_object* x_1540; lean_object* x_1541; lean_object* x_1542; lean_object* x_1543; lean_object* x_1544; lean_object* x_1545; lean_object* x_1546; lean_object* x_1547; lean_object* x_1548; lean_object* x_1549; lean_object* x_1550; lean_object* x_1551; lean_object* x_1552; lean_object* x_1553; lean_object* x_1554; lean_object* x_1555; lean_object* x_1556; lean_object* x_1557; lean_object* x_1558; lean_object* x_1559; lean_object* x_1560; lean_object* x_1561; lean_object* x_1562; lean_object* x_1563; lean_object* x_1564; lean_object* x_1565; lean_object* x_1566; lean_object* x_1567; lean_object* x_1568; lean_object* x_1569; lean_object* x_1570; lean_object* x_1571; lean_object* x_1572; lean_object* x_1573; lean_object* x_1574; lean_object* x_1575; lean_object* x_1576; lean_object* x_1577; lean_object* x_1578; lean_object* x_1579; lean_object* x_1580; lean_object* x_1581; lean_object* x_1582; lean_object* x_1583; lean_object* x_1584; lean_object* x_1585; lean_object* x_1586; lean_object* x_1587; lean_object* x_1588; lean_object* x_1589; lean_object* x_1590; lean_object* x_1591; lean_object* x_1592; lean_object* x_1593; lean_object* x_1594; lean_object* x_1595; lean_object* x_1596; lean_object* x_1597; lean_object* x_1598; lean_object* x_1599; lean_object* x_1600; lean_object* x_1601; lean_object* x_1602; lean_object* x_1603; lean_object* x_1604; lean_object* x_1605; lean_object* x_1606; lean_object* x_1607; lean_object* x_1608; lean_object* x_1609; lean_object* x_1610; lean_object* x_1611; lean_object* x_1612; lean_object* x_1613; lean_object* x_1614; lean_object* x_1615; lean_object* x_1616; lean_object* x_1617; lean_object* x_1618; lean_object* x_1619; lean_object* x_1620; lean_object* x_1621; lean_object* x_1622; lean_object* x_1623; lean_object* x_1624; lean_object* x_1625; lean_object* x_1626; lean_object* x_1627; lean_object* x_1628; lean_object* x_1629; lean_object* x_1630; lean_object* x_1631; lean_object* x_1632; lean_object* x_1633; lean_object* x_1634; lean_object* x_1635; lean_object* x_1636; lean_object* x_1637; lean_object* x_1638; lean_object* x_1639; lean_object* x_1640; lean_object* x_1641; lean_object* x_1642; lean_object* x_1643; lean_object* x_1644; lean_object* x_1645; lean_object* x_1646; lean_object* x_1647; lean_object* x_1648; lean_object* x_1649; lean_object* x_1650; lean_object* x_1651; lean_object* x_1652; lean_object* x_1653; lean_object* x_1654; lean_object* x_1655; lean_object* x_1656; lean_object* x_1657; lean_object* x_1658; lean_object* x_1659; lean_object* x_1660; lean_object* x_1661; lean_object* x_1662; lean_object* x_1663; lean_object* x_1664; lean_object* x_1665; lean_object* x_1666; lean_object* x_1667; lean_object* x_1668; lean_object* x_1669; lean_object* x_1670; lean_object* x_1671; lean_object* x_1672; lean_object* x_1673; lean_object* x_1674; lean_object* x_1675; lean_object* x_1676; lean_object* x_1677; lean_object* x_1678; lean_object* x_1679; lean_object* x_1680; lean_object* x_1681; lean_object* x_1682; lean_object* x_1683; lean_object* x_1684; lean_object* x_1685; lean_object* x_1686; lean_object* x_1687; lean_object* x_1688; lean_object* x_1689; lean_object* x_1690; lean_object* x_1691; lean_object* x_1692; lean_object* x_1693; lean_object* x_1694; lean_object* x_1695; lean_object* x_1696; lean_object* x_1697; lean_object* x_1698; lean_object* x_1699; lean_object* x_1700; lean_object* x_1701; lean_object* x_1702; lean_object* x_1703; lean_object* x_1704; lean_object* x_1705; lean_object* x_1706; lean_object* x_1707; lean_object* x_1708; lean_object* x_1709; lean_object* x_1710; lean_object* x_1711; lean_object* x_1712; lean_object* x_1713; lean_object* x_1714; lean_object* x_1715; lean_object* x_1716; lean_object* x_1717; lean_object* x_1718; lean_object* x_1719; lean_object* x_1720; lean_object* x_1721; lean_object* x_1722; lean_object* x_1723; lean_object* x_1724; lean_object* x_1725; lean_object* x_1726; lean_object* x_1727; lean_object* x_1728; lean_object* x_1729; lean_object* x_1730; lean_object* x_1731; lean_object* x_1732; lean_object* x_1733; lean_object* x_1734; lean_object* x_1735; lean_object* x_1736; lean_object* x_1737; lean_object* x_1738; lean_object* x_1739; lean_object* x_1740; lean_object* x_1741; lean_object* x_1742; lean_object* x_1743; lean_object* x_1744; lean_object* x_1745; lean_object* x_1746; lean_object* x_1747; lean_object* x_1748; lean_object* x_1749; lean_object* x_1750; lean_object* x_1751; lean_object* x_1752; lean_object* x_1753; lean_object* x_1754; lean_object* x_1755; lean_object* x_1756; lean_object* x_1757; lean_object* x_1758; lean_object* x_1759; lean_object* x_1760; lean_object* x_1761; lean_object* x_1762; lean_object* x_1763; lean_object* x_1764; lean_object* x_1765; lean_object* x_1766; lean_object* x_1767; lean_object* x_1768; lean_object* x_1769; lean_object* x_1770; lean_object* x_1771; lean_object* x_1772; lean_object* x_1773; lean_object* x_1774; lean_object* x_1775; lean_object* x_1776; lean_object* x_1777; lean_object* x_1778; lean_object* x_1779; lean_object* x_1780; lean_object* x_1781; lean_object* x_1782; lean_object* x_1783; lean_object* x_1784; lean_object* x_1785; lean_object* x_1786; lean_object* x_1787; lean_object* x_1788; lean_object* x_1789; lean_object* x_1790; lean_object* x_1791; lean_object* x_1792; lean_object* x_1793; lean_object* x_1794; lean_object* x_1795; lean_object* x_1796; lean_object* x_1797; lean_object* x_1798; lean_object* x_1799; lean_object* x_1800; lean_object* x_1801; lean_object* x_1802; lean_object* x_1803; lean_object* x_1804; lean_object* x_1805; lean_object* x_1806; lean_object* x_1807; lean_object* x_1808; lean_object* x_1809; lean_object* x_1810; lean_object* x_1811; lean_object* x_1812; lean_object* x_1813; lean_object* x_1814; lean_object* x_1815; lean_object* x_1816; lean_object* x_1817; lean_object* x_1818; lean_object* x_1819; lean_object* x_1820; lean_object* x_1821; lean_object* x_1822; lean_object* x_1823; lean_object* x_1824; lean_object* x_1825; lean_object* x_1826; lean_object* x_1827; lean_object* x_1828; lean_object* x_1829; lean_object* x_1830; lean_object* x_1831; lean_object* x_1832; lean_object* x_1833; lean_object* x_1834; lean_object* x_1835; lean_object* x_1836; lean_object* x_1837; lean_object* x_1838; lean_object* x_1839; lean_object* x_1840; lean_object* x_1841; lean_object* x_1842; lean_object* x_1843; lean_object* x_1844; lean_object* x_1845; lean_object* x_1846; lean_object* x_1847; lean_object* x_1848; lean_object* x_1849; lean_object* x_1850; lean_object* x_1851; lean_object* x_1852; lean_object* x_1853; lean_object* x_1854; lean_object* x_1855; lean_object* x_1856; lean_object* x_1857; lean_object* x_1858; lean_object* x_1859; lean_object* x_1860; lean_object* x_1861; lean_object* x_1862; lean_object* x_1863; lean_object* x_1864; lean_object* x_1865; lean_object* x_1866; lean_object* x_1867; lean_object* x_1868; lean_object* x_1869; lean_object* x_1870; lean_object* x_1871; lean_object* x_1872; lean_object* x_1873; lean_object* x_1874; lean_object* x_1875; lean_object* x_1876; lean_object* x_1877; lean_object* x_1878; lean_object* x_1879; lean_object* x_1880; lean_object* x_1881; lean_object* x_1882; lean_object* x_1883; lean_object* x_1884; lean_object* x_1885; lean_object* x_1886; lean_object* x_1887; lean_object* x_1888; lean_object* x_1889; lean_object* x_1890; lean_object* x_1891; lean_object* x_1892; lean_object* x_1893; lean_object* x_1894; lean_object* x_1895; lean_object* x_1896; lean_object* x_1897; lean_object* x_1898; lean_object* x_1899; lean_object* x_1900; lean_object* x_1901; lean_object* x_1902; lean_object* x_1903; lean_object* x_1904; lean_object* x_1905; lean_object* x_1906; lean_object* x_1907; lean_object* x_1908; lean_object* x_1909; lean_object* x_1910; lean_object* x_1911; lean_object* x_1912; lean_object* x_1913; lean_object* x_1914; lean_object* x_1915; lean_object* x_1916; lean_object* x_1917; lean_object* x_1918; lean_object* x_1919; lean_object* x_1920; lean_object* x_1921; lean_object* x_1922; lean_object* x_1923; lean_object* x_1924; lean_object* x_1925; lean_object* x_1926; lean_object* x_1927; lean_object* x_1928; lean_object* x_1929; lean_object* x_1930; lean_object* x_1931; lean_object* x_1932; lean_object* x_1933; lean_object* x_1934; lean_object* x_1935; lean_object* x_1936; lean_object* x_1937; lean_object* x_1938; lean_object* x_1939; lean_object* x_1940; lean_object* x_1941; lean_object* x_1942; lean_object* x_1943; lean_object* x_1944; lean_object* x_1945; lean_object* x_1946; lean_object* x_1947; lean_object* x_1948; lean_object* x_1949; lean_object* x_1950; lean_object* x_1951; lean_object* x_1952; lean_object* x_1953; lean_object* x_1954; lean_object* x_1955; lean_object* x_1956; lean_object* x_1957; lean_object* x_1958; lean_object* x_1959; lean_object* x_1960; lean_object* x_1961; lean_object* x_1962; lean_object* x_1963; lean_object* x_1964; lean_object* x_1965; lean_object* x_1966; lean_object* x_1967; lean_object* x_1968; lean_object* x_1969; lean_object* x_1970; lean_object* x_1971; lean_object* x_1972; lean_object* x_1973; lean_object* x_1974; lean_object* x_1975; lean_object* x_1976; lean_object* x_1977; lean_object* x_1978; lean_object* x_1979; lean_object* x_1980; lean_object* x_1981; lean_object* x_1982; lean_object* x_1983; lean_object* x_1984; lean_object* x_1985; lean_object* x_1986; lean_object* x_1987; lean_object* x_1988; lean_object* x_1989; lean_object* x_1990; lean_object* x_1991; lean_object* x_1992; lean_object* x_1993; lean_object* x_1994; lean_object* x_1995; lean_object* x_1996; lean_object* x_1997; lean_object* x_1998; lean_object* x_1999; lean_object* x_2000; lean_object* x_2001; lean_object* x_2002; lean_object* x_2003; lean_object* x_2004; lean_object* x_2005; lean_object* x_2006; lean_object* x_2007; lean_object* x_2008; lean_object* x_2009; lean_object* x_2010; lean_object* x_2011; lean_object* x_2012; lean_object* x_2013; lean_object* x_2014; lean_object* x_2015; lean_object* x_2016; lean_object* x_2017; lean_object* x_2018; lean_object* x_2019; lean_object* x_2020; lean_object* x_2021; lean_object* x_2022; lean_object* x_2023; lean_object* x_2024; lean_object* x_2025; lean_object* x_2026; lean_object* x_2027; lean_object* x_2028; lean_object* x_2029; lean_object* x_2030; lean_object* x_2031; lean_object* x_2032; lean_object* x_2033; lean_object* x_2034; lean_object* x_2035; lean_object* x_2036; lean_object* x_2037; lean_object* x_2038; lean_object* x_2039; lean_object* x_2040; lean_object* x_2041; lean_object* x_2042; +x_1531 = lean_ctor_get(x_7, 1); +x_1532 = lean_ctor_get(x_7, 0); +lean_inc(x_1531); +lean_inc(x_1532); +lean_dec(x_7); +x_1533 = lean_ctor_get(x_1531, 0); lean_inc(x_1533); +x_1534 = lean_ctor_get(x_1531, 1); lean_inc(x_1534); -lean_dec(x_6); -x_1535 = lean_ctor_get(x_1533, 0); -lean_inc(x_1535); -x_1536 = lean_ctor_get(x_1533, 1); -lean_inc(x_1536); -if (lean_is_exclusive(x_1533)) { - lean_ctor_release(x_1533, 0); - lean_ctor_release(x_1533, 1); - x_1537 = x_1533; +if (lean_is_exclusive(x_1531)) { + lean_ctor_release(x_1531, 0); + lean_ctor_release(x_1531, 1); + x_1535 = x_1531; } else { - lean_dec_ref(x_1533); - x_1537 = lean_box(0); + lean_dec_ref(x_1531); + x_1535 = lean_box(0); } -lean_inc(x_7); -x_1538 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_7, x_8); -x_1539 = lean_ctor_get(x_1538, 0); -lean_inc(x_1539); -x_1540 = lean_ctor_get(x_1538, 1); -lean_inc(x_1540); -if (lean_is_exclusive(x_1538)) { - lean_ctor_release(x_1538, 0); - lean_ctor_release(x_1538, 1); - x_1541 = x_1538; +lean_inc(x_8); +x_1536 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_8, x_9); +x_1537 = lean_ctor_get(x_1536, 0); +lean_inc(x_1537); +x_1538 = lean_ctor_get(x_1536, 1); +lean_inc(x_1538); +if (lean_is_exclusive(x_1536)) { + lean_ctor_release(x_1536, 0); + lean_ctor_release(x_1536, 1); + x_1539 = x_1536; } else { - lean_dec_ref(x_1538); - x_1541 = lean_box(0); + lean_dec_ref(x_1536); + x_1539 = lean_box(0); } -x_1542 = lean_ctor_get(x_7, 2); -lean_inc(x_1542); -x_1543 = lean_ctor_get(x_7, 1); -lean_inc(x_1543); -lean_dec(x_7); -x_1544 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__1; +x_1540 = lean_ctor_get(x_8, 2); +lean_inc(x_1540); +x_1541 = lean_ctor_get(x_8, 1); +lean_inc(x_1541); +lean_dec(x_8); +x_1542 = l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___closed__1; lean_inc(x_1); -x_1545 = lean_name_mk_string(x_1, x_1544); -x_1546 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__2; -lean_inc(x_1545); -x_1547 = lean_name_mk_string(x_1545, x_1546); -x_1548 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__3; -lean_inc(x_1545); -x_1549 = lean_name_mk_string(x_1545, x_1548); -x_1550 = l_Lean_mkHole___closed__5; +x_1543 = lean_name_mk_string(x_1, x_1542); +x_1544 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__1; +lean_inc(x_1543); +x_1545 = lean_name_mk_string(x_1543, x_1544); +x_1546 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__2; +lean_inc(x_1543); +x_1547 = lean_name_mk_string(x_1543, x_1546); +x_1548 = l_Lean_mkHole___closed__5; lean_inc(x_1); -x_1551 = lean_name_mk_string(x_1, x_1550); -x_1552 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__4; -lean_inc(x_1551); -x_1553 = lean_name_mk_string(x_1551, x_1552); -x_1554 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__5; -lean_inc(x_1539); -x_1555 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1555, 0, x_1539); -lean_ctor_set(x_1555, 1, x_1554); -x_1556 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__6; -lean_inc(x_1551); -x_1557 = lean_name_mk_string(x_1551, x_1556); -x_1558 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__7; -lean_inc(x_1551); -x_1559 = lean_name_mk_string(x_1551, x_1558); -x_1560 = lean_box(2); -x_1561 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__33; -x_1562 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1562, 0, x_1560); -lean_ctor_set(x_1562, 1, x_1559); -lean_ctor_set(x_1562, 2, x_1561); -x_1563 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__8; -x_1564 = lean_name_mk_string(x_1, x_1563); -x_1565 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__9; -x_1566 = lean_name_mk_string(x_1564, x_1565); -lean_inc(x_1539); -x_1567 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1567, 0, x_1539); -lean_ctor_set(x_1567, 1, x_1565); -x_1568 = l_Lean_Syntax_mkApp___closed__3; -x_1569 = lean_array_push(x_1568, x_1567); -x_1570 = lean_array_push(x_1569, x_2); -x_1571 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1571, 0, x_1560); -lean_ctor_set(x_1571, 1, x_1566); -lean_ctor_set(x_1571, 2, x_1570); -x_1572 = lean_array_push(x_1568, x_1562); -x_1573 = lean_array_push(x_1572, x_1571); -x_1574 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1574, 0, x_1560); -lean_ctor_set(x_1574, 1, x_1557); -lean_ctor_set(x_1574, 2, x_1573); -x_1575 = l_Lean_mkOptionalNode___closed__2; -x_1576 = lean_array_push(x_1575, x_1574); -x_1577 = l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__2; -x_1578 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1578, 0, x_1560); +x_1549 = lean_name_mk_string(x_1, x_1548); +x_1550 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__3; +lean_inc(x_1549); +x_1551 = lean_name_mk_string(x_1549, x_1550); +x_1552 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__4; +lean_inc(x_1537); +x_1553 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1553, 0, x_1537); +lean_ctor_set(x_1553, 1, x_1552); +x_1554 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__5; +lean_inc(x_1549); +x_1555 = lean_name_mk_string(x_1549, x_1554); +x_1556 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__6; +lean_inc(x_1549); +x_1557 = lean_name_mk_string(x_1549, x_1556); +x_1558 = lean_box(2); +x_1559 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__33; +x_1560 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1560, 0, x_1558); +lean_ctor_set(x_1560, 1, x_1557); +lean_ctor_set(x_1560, 2, x_1559); +x_1561 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__7; +x_1562 = lean_name_mk_string(x_1, x_1561); +x_1563 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__8; +x_1564 = lean_name_mk_string(x_1562, x_1563); +lean_inc(x_1537); +x_1565 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1565, 0, x_1537); +lean_ctor_set(x_1565, 1, x_1563); +x_1566 = l_Lean_Syntax_mkApp___closed__3; +x_1567 = lean_array_push(x_1566, x_1565); +x_1568 = lean_array_push(x_1567, x_2); +x_1569 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1569, 0, x_1558); +lean_ctor_set(x_1569, 1, x_1564); +lean_ctor_set(x_1569, 2, x_1568); +x_1570 = lean_array_push(x_1566, x_1560); +x_1571 = lean_array_push(x_1570, x_1569); +x_1572 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1572, 0, x_1558); +lean_ctor_set(x_1572, 1, x_1555); +lean_ctor_set(x_1572, 2, x_1571); +x_1573 = l_Lean_mkOptionalNode___closed__2; +x_1574 = lean_array_push(x_1573, x_1572); +x_1575 = l_Lean_mkNullNode___closed__2; +x_1576 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1576, 0, x_1558); +lean_ctor_set(x_1576, 1, x_1575); +lean_ctor_set(x_1576, 2, x_1574); +x_1577 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__16; +lean_inc(x_1537); +x_1578 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1578, 0, x_1537); lean_ctor_set(x_1578, 1, x_1577); -lean_ctor_set(x_1578, 2, x_1576); -x_1579 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__16; -lean_inc(x_1539); -x_1580 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1580, 0, x_1539); -lean_ctor_set(x_1580, 1, x_1579); -x_1581 = l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__4; -x_1582 = lean_array_push(x_1581, x_1555); -x_1583 = lean_array_push(x_1582, x_1578); -lean_inc(x_1580); -x_1584 = lean_array_push(x_1583, x_1580); +x_1579 = l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__4; +x_1580 = lean_array_push(x_1579, x_1553); +x_1581 = lean_array_push(x_1580, x_1576); +lean_inc(x_1578); +x_1582 = lean_array_push(x_1581, x_1578); +x_1583 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1583, 0, x_1558); +lean_ctor_set(x_1583, 1, x_1551); +lean_ctor_set(x_1583, 2, x_1582); +x_1584 = lean_array_push(x_1573, x_1583); x_1585 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1585, 0, x_1560); -lean_ctor_set(x_1585, 1, x_1553); +lean_ctor_set(x_1585, 0, x_1558); +lean_ctor_set(x_1585, 1, x_1575); lean_ctor_set(x_1585, 2, x_1584); -x_1586 = lean_array_push(x_1575, x_1585); -x_1587 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1587, 0, x_1560); -lean_ctor_set(x_1587, 1, x_1577); -lean_ctor_set(x_1587, 2, x_1586); -x_1588 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__10; -x_1589 = lean_array_push(x_1588, x_1587); -x_1590 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__12; -x_1591 = lean_array_push(x_1589, x_1590); -x_1592 = lean_array_push(x_1591, x_1590); -x_1593 = lean_array_push(x_1592, x_1590); -x_1594 = lean_array_push(x_1593, x_1590); -x_1595 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1595, 0, x_1560); -lean_ctor_set(x_1595, 1, x_1549); -lean_ctor_set(x_1595, 2, x_1594); -x_1596 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__11; -lean_inc(x_1545); -x_1597 = lean_name_mk_string(x_1545, x_1596); -lean_inc(x_1539); -x_1598 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1598, 0, x_1539); -lean_ctor_set(x_1598, 1, x_1596); -x_1599 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__12; -lean_inc(x_1545); -x_1600 = lean_name_mk_string(x_1545, x_1599); -x_1601 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__16; -lean_inc(x_1542); +x_1586 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__9; +x_1587 = lean_array_push(x_1586, x_1585); +x_1588 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__12; +x_1589 = lean_array_push(x_1587, x_1588); +x_1590 = lean_array_push(x_1589, x_1588); +x_1591 = lean_array_push(x_1590, x_1588); +x_1592 = lean_array_push(x_1591, x_1588); +x_1593 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1593, 0, x_1558); +lean_ctor_set(x_1593, 1, x_1547); +lean_ctor_set(x_1593, 2, x_1592); +x_1594 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__10; +lean_inc(x_1543); +x_1595 = lean_name_mk_string(x_1543, x_1594); +lean_inc(x_1537); +x_1596 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1596, 0, x_1537); +lean_ctor_set(x_1596, 1, x_1594); +x_1597 = l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___closed__3; lean_inc(x_1543); -x_1602 = l_Lean_addMacroScope(x_1543, x_1601, x_1542); -x_1603 = lean_box(0); -x_1604 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__15; -lean_inc(x_1539); -x_1605 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1605, 0, x_1539); -lean_ctor_set(x_1605, 1, x_1604); -lean_ctor_set(x_1605, 2, x_1602); -lean_ctor_set(x_1605, 3, x_1603); -x_1606 = lean_array_push(x_1568, x_1605); -x_1607 = lean_array_push(x_1606, x_1590); -x_1608 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1608, 0, x_1560); -lean_ctor_set(x_1608, 1, x_1600); -lean_ctor_set(x_1608, 2, x_1607); -x_1609 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__17; -lean_inc(x_1545); -x_1610 = lean_name_mk_string(x_1545, x_1609); -x_1611 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__18; -lean_inc(x_1551); -x_1612 = lean_name_mk_string(x_1551, x_1611); -x_1613 = l_Lean_toolchain___closed__1; -lean_inc(x_1539); -x_1614 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1614, 0, x_1539); -lean_ctor_set(x_1614, 1, x_1613); -x_1615 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__22; -lean_inc(x_1542); +x_1598 = lean_name_mk_string(x_1543, x_1597); +x_1599 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__14; +lean_inc(x_1540); +lean_inc(x_1541); +x_1600 = l_Lean_addMacroScope(x_1541, x_1599, x_1540); +x_1601 = lean_box(0); +x_1602 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__13; +lean_inc(x_1537); +x_1603 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1603, 0, x_1537); +lean_ctor_set(x_1603, 1, x_1602); +lean_ctor_set(x_1603, 2, x_1600); +lean_ctor_set(x_1603, 3, x_1601); +x_1604 = lean_array_push(x_1566, x_1603); +x_1605 = lean_array_push(x_1604, x_1588); +x_1606 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1606, 0, x_1558); +lean_ctor_set(x_1606, 1, x_1598); +lean_ctor_set(x_1606, 2, x_1605); +x_1607 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__15; lean_inc(x_1543); -x_1616 = l_Lean_addMacroScope(x_1543, x_1615, x_1542); -x_1617 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__19; +x_1608 = lean_name_mk_string(x_1543, x_1607); +x_1609 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__16; +lean_inc(x_1549); +x_1610 = lean_name_mk_string(x_1549, x_1609); +x_1611 = l_Lean_toolchain___closed__1; +lean_inc(x_1537); +x_1612 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1612, 0, x_1537); +lean_ctor_set(x_1612, 1, x_1611); +x_1613 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__20; +lean_inc(x_1540); +lean_inc(x_1541); +x_1614 = l_Lean_addMacroScope(x_1541, x_1613, x_1540); +x_1615 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__17; lean_inc(x_3); -x_1618 = lean_name_mk_string(x_3, x_1617); -if (lean_is_scalar(x_1537)) { - x_1619 = lean_alloc_ctor(0, 2, 0); +x_1616 = lean_name_mk_string(x_3, x_1615); +if (lean_is_scalar(x_1535)) { + x_1617 = lean_alloc_ctor(0, 2, 0); } else { - x_1619 = x_1537; -} -lean_ctor_set(x_1619, 0, x_1618); -lean_ctor_set(x_1619, 1, x_1603); -x_1620 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1620, 0, x_1619); -lean_ctor_set(x_1620, 1, x_1603); -x_1621 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__21; -lean_inc(x_1539); -x_1622 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1622, 0, x_1539); -lean_ctor_set(x_1622, 1, x_1621); -lean_ctor_set(x_1622, 2, x_1616); -lean_ctor_set(x_1622, 3, x_1620); -x_1623 = lean_array_push(x_1568, x_1614); -lean_inc(x_1623); -x_1624 = lean_array_push(x_1623, x_1622); + x_1617 = x_1535; +} +lean_ctor_set(x_1617, 0, x_1616); +lean_ctor_set(x_1617, 1, x_1601); +x_1618 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1618, 0, x_1617); +lean_ctor_set(x_1618, 1, x_1601); +x_1619 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__19; +lean_inc(x_1537); +x_1620 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1620, 0, x_1537); +lean_ctor_set(x_1620, 1, x_1619); +lean_ctor_set(x_1620, 2, x_1614); +lean_ctor_set(x_1620, 3, x_1618); +x_1621 = lean_array_push(x_1566, x_1612); +lean_inc(x_1621); +x_1622 = lean_array_push(x_1621, x_1620); +x_1623 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1623, 0, x_1558); +lean_ctor_set(x_1623, 1, x_1610); +lean_ctor_set(x_1623, 2, x_1622); +x_1624 = lean_array_push(x_1573, x_1623); x_1625 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1625, 0, x_1560); -lean_ctor_set(x_1625, 1, x_1612); +lean_ctor_set(x_1625, 0, x_1558); +lean_ctor_set(x_1625, 1, x_1575); lean_ctor_set(x_1625, 2, x_1624); -x_1626 = lean_array_push(x_1575, x_1625); -x_1627 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1627, 0, x_1560); -lean_ctor_set(x_1627, 1, x_1577); -lean_ctor_set(x_1627, 2, x_1626); -x_1628 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__23; -x_1629 = lean_array_push(x_1628, x_1627); -x_1630 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1630, 0, x_1560); -lean_ctor_set(x_1630, 1, x_1610); -lean_ctor_set(x_1630, 2, x_1629); -x_1631 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__24; -x_1632 = lean_name_mk_string(x_1545, x_1631); -x_1633 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__8; -lean_inc(x_1539); -x_1634 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1634, 0, x_1539); -lean_ctor_set(x_1634, 1, x_1633); -x_1635 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__25; -lean_inc(x_1551); -x_1636 = lean_name_mk_string(x_1551, x_1635); -lean_inc(x_1539); -x_1637 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1637, 0, x_1539); -lean_ctor_set(x_1637, 1, x_1635); -x_1638 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__26; -lean_inc(x_1551); -x_1639 = lean_name_mk_string(x_1551, x_1638); -x_1640 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__30; -lean_inc(x_1542); -lean_inc(x_1543); -x_1641 = l_Lean_addMacroScope(x_1543, x_1640, x_1542); -x_1642 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__29; -lean_inc(x_1539); -x_1643 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1643, 0, x_1539); -lean_ctor_set(x_1643, 1, x_1642); -lean_ctor_set(x_1643, 2, x_1641); -lean_ctor_set(x_1643, 3, x_1603); -lean_inc(x_1643); -x_1644 = lean_array_push(x_1575, x_1643); -x_1645 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1645, 0, x_1560); -lean_ctor_set(x_1645, 1, x_1577); -lean_ctor_set(x_1645, 2, x_1644); -x_1646 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__31; -lean_inc(x_1539); -x_1647 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1647, 0, x_1539); -lean_ctor_set(x_1647, 1, x_1646); -x_1648 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__32; -lean_inc(x_1551); -x_1649 = lean_name_mk_string(x_1551, x_1648); -lean_inc(x_1539); -x_1650 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1650, 0, x_1539); -lean_ctor_set(x_1650, 1, x_1648); -x_1651 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__33; -lean_inc(x_1551); -x_1652 = lean_name_mk_string(x_1551, x_1651); -x_1653 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__34; -lean_inc(x_1551); -x_1654 = lean_name_mk_string(x_1551, x_1653); -x_1655 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__35; -lean_inc(x_1551); -x_1656 = lean_name_mk_string(x_1551, x_1655); -x_1657 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__36; -lean_inc(x_1539); -x_1658 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1658, 0, x_1539); -lean_ctor_set(x_1658, 1, x_1657); -x_1659 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__37; -lean_inc(x_1551); -x_1660 = lean_name_mk_string(x_1551, x_1659); -x_1661 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__41; -lean_inc(x_1542); -lean_inc(x_1543); -x_1662 = l_Lean_addMacroScope(x_1543, x_1661, x_1542); -x_1663 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__40; -lean_inc(x_1539); -x_1664 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1664, 0, x_1539); +x_1626 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__21; +x_1627 = lean_array_push(x_1626, x_1625); +x_1628 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1628, 0, x_1558); +lean_ctor_set(x_1628, 1, x_1608); +lean_ctor_set(x_1628, 2, x_1627); +x_1629 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__22; +x_1630 = lean_name_mk_string(x_1543, x_1629); +x_1631 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__8; +lean_inc(x_1537); +x_1632 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1632, 0, x_1537); +lean_ctor_set(x_1632, 1, x_1631); +x_1633 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__23; +lean_inc(x_1549); +x_1634 = lean_name_mk_string(x_1549, x_1633); +lean_inc(x_1537); +x_1635 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1635, 0, x_1537); +lean_ctor_set(x_1635, 1, x_1633); +x_1636 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__24; +lean_inc(x_1549); +x_1637 = lean_name_mk_string(x_1549, x_1636); +x_1638 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__28; +lean_inc(x_1540); +lean_inc(x_1541); +x_1639 = l_Lean_addMacroScope(x_1541, x_1638, x_1540); +x_1640 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__27; +lean_inc(x_1537); +x_1641 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1641, 0, x_1537); +lean_ctor_set(x_1641, 1, x_1640); +lean_ctor_set(x_1641, 2, x_1639); +lean_ctor_set(x_1641, 3, x_1601); +lean_inc(x_1641); +x_1642 = lean_array_push(x_1573, x_1641); +x_1643 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1643, 0, x_1558); +lean_ctor_set(x_1643, 1, x_1575); +lean_ctor_set(x_1643, 2, x_1642); +x_1644 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__29; +lean_inc(x_1537); +x_1645 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1645, 0, x_1537); +lean_ctor_set(x_1645, 1, x_1644); +x_1646 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__30; +lean_inc(x_1549); +x_1647 = lean_name_mk_string(x_1549, x_1646); +lean_inc(x_1537); +x_1648 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1648, 0, x_1537); +lean_ctor_set(x_1648, 1, x_1646); +x_1649 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__31; +lean_inc(x_1549); +x_1650 = lean_name_mk_string(x_1549, x_1649); +x_1651 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__32; +lean_inc(x_1549); +x_1652 = lean_name_mk_string(x_1549, x_1651); +x_1653 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__33; +lean_inc(x_1549); +x_1654 = lean_name_mk_string(x_1549, x_1653); +x_1655 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__34; +lean_inc(x_1537); +x_1656 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1656, 0, x_1537); +lean_ctor_set(x_1656, 1, x_1655); +x_1657 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__35; +lean_inc(x_1549); +x_1658 = lean_name_mk_string(x_1549, x_1657); +x_1659 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__39; +lean_inc(x_1540); +lean_inc(x_1541); +x_1660 = l_Lean_addMacroScope(x_1541, x_1659, x_1540); +x_1661 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__38; +lean_inc(x_1537); +x_1662 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1662, 0, x_1537); +lean_ctor_set(x_1662, 1, x_1661); +lean_ctor_set(x_1662, 2, x_1660); +lean_ctor_set(x_1662, 3, x_1601); +x_1663 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__40; +lean_inc(x_1537); +x_1664 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1664, 0, x_1537); lean_ctor_set(x_1664, 1, x_1663); -lean_ctor_set(x_1664, 2, x_1662); -lean_ctor_set(x_1664, 3, x_1603); -x_1665 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__42; -lean_inc(x_1539); -x_1666 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1666, 0, x_1539); -lean_ctor_set(x_1666, 1, x_1665); -x_1667 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__43; -lean_inc(x_1551); -x_1668 = lean_name_mk_string(x_1551, x_1667); -x_1669 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__44; -lean_inc(x_1539); -x_1670 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1670, 0, x_1539); -lean_ctor_set(x_1670, 1, x_1669); -x_1671 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__45; -lean_inc(x_1551); -x_1672 = lean_name_mk_string(x_1551, x_1671); -x_1673 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__46; -lean_inc(x_1551); -x_1674 = lean_name_mk_string(x_1551, x_1673); -x_1675 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__12; -lean_inc(x_1539); +x_1665 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__41; +lean_inc(x_1549); +x_1666 = lean_name_mk_string(x_1549, x_1665); +x_1667 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__42; +lean_inc(x_1537); +x_1668 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1668, 0, x_1537); +lean_ctor_set(x_1668, 1, x_1667); +x_1669 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__43; +lean_inc(x_1549); +x_1670 = lean_name_mk_string(x_1549, x_1669); +x_1671 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__44; +lean_inc(x_1549); +x_1672 = lean_name_mk_string(x_1549, x_1671); +x_1673 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__12; +lean_inc(x_1537); +x_1674 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1674, 0, x_1537); +lean_ctor_set(x_1674, 1, x_1673); +x_1675 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__45; +lean_inc(x_1537); x_1676 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1676, 0, x_1539); +lean_ctor_set(x_1676, 0, x_1537); lean_ctor_set(x_1676, 1, x_1675); -x_1677 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__47; -lean_inc(x_1539); -x_1678 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1678, 0, x_1539); -lean_ctor_set(x_1678, 1, x_1677); -x_1679 = lean_array_push(x_1575, x_1678); -x_1680 = l_Lean_evalPrec___closed__2; -x_1681 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1681, 0, x_1560); -lean_ctor_set(x_1681, 1, x_1680); -lean_ctor_set(x_1681, 2, x_1679); -x_1682 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__38; -lean_inc(x_1643); -x_1683 = lean_array_push(x_1682, x_1643); -lean_inc(x_1676); -x_1684 = lean_array_push(x_1683, x_1676); -lean_inc(x_1681); -lean_inc(x_1684); -x_1685 = lean_array_push(x_1684, x_1681); -lean_inc(x_1580); -x_1686 = lean_array_push(x_1685, x_1580); -lean_inc(x_1674); -x_1687 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1687, 0, x_1560); -lean_ctor_set(x_1687, 1, x_1674); -lean_ctor_set(x_1687, 2, x_1686); -x_1688 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__48; -lean_inc(x_1539); -x_1689 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1689, 0, x_1539); -lean_ctor_set(x_1689, 1, x_1688); -x_1690 = lean_array_push(x_1575, x_1689); -x_1691 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1691, 0, x_1560); -lean_ctor_set(x_1691, 1, x_1680); -lean_ctor_set(x_1691, 2, x_1690); -x_1692 = lean_array_push(x_1682, x_1687); -x_1693 = lean_array_push(x_1692, x_1676); -lean_inc(x_1691); -x_1694 = lean_array_push(x_1693, x_1691); -lean_inc(x_1580); -x_1695 = lean_array_push(x_1694, x_1580); +x_1677 = lean_array_push(x_1573, x_1676); +x_1678 = l_Lean_Syntax_mkNumLit___closed__2; +x_1679 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1679, 0, x_1558); +lean_ctor_set(x_1679, 1, x_1678); +lean_ctor_set(x_1679, 2, x_1677); +x_1680 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__38; +lean_inc(x_1641); +x_1681 = lean_array_push(x_1680, x_1641); lean_inc(x_1674); +x_1682 = lean_array_push(x_1681, x_1674); +lean_inc(x_1679); +lean_inc(x_1682); +x_1683 = lean_array_push(x_1682, x_1679); +lean_inc(x_1578); +x_1684 = lean_array_push(x_1683, x_1578); +lean_inc(x_1672); +x_1685 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1685, 0, x_1558); +lean_ctor_set(x_1685, 1, x_1672); +lean_ctor_set(x_1685, 2, x_1684); +x_1686 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__46; +lean_inc(x_1537); +x_1687 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1687, 0, x_1537); +lean_ctor_set(x_1687, 1, x_1686); +x_1688 = lean_array_push(x_1573, x_1687); +x_1689 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1689, 0, x_1558); +lean_ctor_set(x_1689, 1, x_1678); +lean_ctor_set(x_1689, 2, x_1688); +x_1690 = lean_array_push(x_1680, x_1685); +x_1691 = lean_array_push(x_1690, x_1674); +lean_inc(x_1689); +x_1692 = lean_array_push(x_1691, x_1689); +lean_inc(x_1578); +x_1693 = lean_array_push(x_1692, x_1578); +lean_inc(x_1672); +x_1694 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1694, 0, x_1558); +lean_ctor_set(x_1694, 1, x_1672); +lean_ctor_set(x_1694, 2, x_1693); +x_1695 = lean_array_push(x_1626, x_1694); x_1696 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1696, 0, x_1560); -lean_ctor_set(x_1696, 1, x_1674); +lean_ctor_set(x_1696, 0, x_1558); +lean_ctor_set(x_1696, 1, x_1670); lean_ctor_set(x_1696, 2, x_1695); -x_1697 = lean_array_push(x_1628, x_1696); +x_1697 = lean_array_push(x_1573, x_1696); x_1698 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1698, 0, x_1560); -lean_ctor_set(x_1698, 1, x_1672); +lean_ctor_set(x_1698, 0, x_1558); +lean_ctor_set(x_1698, 1, x_1575); lean_ctor_set(x_1698, 2, x_1697); -x_1699 = lean_array_push(x_1575, x_1698); -x_1700 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1700, 0, x_1560); -lean_ctor_set(x_1700, 1, x_1577); -lean_ctor_set(x_1700, 2, x_1699); -x_1701 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__49; -lean_inc(x_1539); -x_1702 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1702, 0, x_1539); -lean_ctor_set(x_1702, 1, x_1701); -x_1703 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__50; -lean_inc(x_1551); -x_1704 = lean_name_mk_string(x_1551, x_1703); -x_1705 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__51; -lean_inc(x_1551); -x_1706 = lean_name_mk_string(x_1551, x_1705); -x_1707 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__52; -lean_inc(x_1539); -x_1708 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1708, 0, x_1539); -lean_ctor_set(x_1708, 1, x_1707); -x_1709 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__53; -lean_inc(x_1551); -x_1710 = lean_name_mk_string(x_1551, x_1709); -x_1711 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__54; -lean_inc(x_1539); -x_1712 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1712, 0, x_1539); -lean_ctor_set(x_1712, 1, x_1711); -x_1713 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__57; -lean_inc(x_1542); -lean_inc(x_1543); -x_1714 = l_Lean_addMacroScope(x_1543, x_1713, x_1542); -x_1715 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__6; -x_1716 = lean_name_mk_string(x_4, x_1715); -lean_inc(x_1716); -x_1717 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1717, 0, x_1716); -lean_ctor_set(x_1717, 1, x_1603); -x_1718 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1718, 0, x_1717); -lean_ctor_set(x_1718, 1, x_1603); -x_1719 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__56; -lean_inc(x_1539); -x_1720 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1720, 0, x_1539); +x_1699 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__47; +lean_inc(x_1537); +x_1700 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1700, 0, x_1537); +lean_ctor_set(x_1700, 1, x_1699); +x_1701 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__48; +lean_inc(x_1549); +x_1702 = lean_name_mk_string(x_1549, x_1701); +x_1703 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__49; +lean_inc(x_1549); +x_1704 = lean_name_mk_string(x_1549, x_1703); +x_1705 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__50; +lean_inc(x_1537); +x_1706 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1706, 0, x_1537); +lean_ctor_set(x_1706, 1, x_1705); +x_1707 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__51; +lean_inc(x_1549); +x_1708 = lean_name_mk_string(x_1549, x_1707); +x_1709 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__52; +lean_inc(x_1537); +x_1710 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1710, 0, x_1537); +lean_ctor_set(x_1710, 1, x_1709); +x_1711 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__55; +lean_inc(x_1540); +lean_inc(x_1541); +x_1712 = l_Lean_addMacroScope(x_1541, x_1711, x_1540); +x_1713 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__6; +x_1714 = lean_name_mk_string(x_4, x_1713); +lean_inc(x_1714); +x_1715 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1715, 0, x_1714); +lean_ctor_set(x_1715, 1, x_1601); +x_1716 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1716, 0, x_1715); +lean_ctor_set(x_1716, 1, x_1601); +x_1717 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__54; +lean_inc(x_1537); +x_1718 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1718, 0, x_1537); +lean_ctor_set(x_1718, 1, x_1717); +lean_ctor_set(x_1718, 2, x_1712); +lean_ctor_set(x_1718, 3, x_1716); +x_1719 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__4; +lean_inc(x_1537); +x_1720 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1720, 0, x_1537); lean_ctor_set(x_1720, 1, x_1719); -lean_ctor_set(x_1720, 2, x_1714); -lean_ctor_set(x_1720, 3, x_1718); -x_1721 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__4; -lean_inc(x_1539); -x_1722 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1722, 0, x_1539); -lean_ctor_set(x_1722, 1, x_1721); -lean_inc(x_1539); -x_1723 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1723, 0, x_1539); -lean_ctor_set(x_1723, 1, x_1715); -x_1724 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__60; -lean_inc(x_1539); -x_1725 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1725, 0, x_1539); -lean_ctor_set(x_1725, 1, x_1724); -x_1726 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__18; -lean_inc(x_1539); +lean_inc(x_1537); +x_1721 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1721, 0, x_1537); +lean_ctor_set(x_1721, 1, x_1713); +x_1722 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__56; +x_1723 = lean_name_mk_string(x_5, x_1722); +x_1724 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__57; +x_1725 = lean_name_mk_string(x_1723, x_1724); +x_1726 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__58; +lean_inc(x_1537); x_1727 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1727, 0, x_1539); +lean_ctor_set(x_1727, 0, x_1537); lean_ctor_set(x_1727, 1, x_1726); -x_1728 = lean_array_push(x_1623, x_1727); -x_1729 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__62; -x_1730 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1730, 0, x_1560); -lean_ctor_set(x_1730, 1, x_1729); -lean_ctor_set(x_1730, 2, x_1728); -x_1731 = lean_array_push(x_1682, x_1725); -x_1732 = lean_array_push(x_1731, x_1590); -lean_inc(x_1664); -x_1733 = lean_array_push(x_1732, x_1664); -lean_inc(x_1733); -x_1734 = lean_array_push(x_1733, x_1730); -x_1735 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__59; -x_1736 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1736, 0, x_1560); -lean_ctor_set(x_1736, 1, x_1735); -lean_ctor_set(x_1736, 2, x_1734); -x_1737 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__8; -lean_inc(x_1539); -x_1738 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1738, 0, x_1539); -lean_ctor_set(x_1738, 1, x_1737); -x_1739 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__37; -lean_inc(x_1722); -x_1740 = lean_array_push(x_1739, x_1722); -x_1741 = lean_array_push(x_1740, x_1723); -lean_inc(x_1634); -x_1742 = lean_array_push(x_1741, x_1634); -lean_inc(x_1742); -x_1743 = lean_array_push(x_1742, x_1736); -lean_inc(x_1738); -x_1744 = lean_array_push(x_1743, x_1738); -lean_inc(x_1716); -x_1745 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1745, 0, x_1560); -lean_ctor_set(x_1745, 1, x_1716); -lean_ctor_set(x_1745, 2, x_1744); -x_1746 = lean_array_push(x_1739, x_1712); -x_1747 = lean_array_push(x_1746, x_1720); +x_1728 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__16; +lean_inc(x_1537); +x_1729 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1729, 0, x_1537); +lean_ctor_set(x_1729, 1, x_1728); +x_1730 = lean_array_push(x_1621, x_1729); +x_1731 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__60; +x_1732 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1732, 0, x_1558); +lean_ctor_set(x_1732, 1, x_1731); +lean_ctor_set(x_1732, 2, x_1730); +x_1733 = lean_array_push(x_1680, x_1727); +x_1734 = lean_array_push(x_1733, x_1588); +lean_inc(x_1662); +x_1735 = lean_array_push(x_1734, x_1662); +lean_inc(x_1735); +x_1736 = lean_array_push(x_1735, x_1732); +lean_inc(x_1725); +x_1737 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1737, 0, x_1558); +lean_ctor_set(x_1737, 1, x_1725); +lean_ctor_set(x_1737, 2, x_1736); +x_1738 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__8; +lean_inc(x_1537); +x_1739 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1739, 0, x_1537); +lean_ctor_set(x_1739, 1, x_1738); +x_1740 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__37; +lean_inc(x_1720); +x_1741 = lean_array_push(x_1740, x_1720); +x_1742 = lean_array_push(x_1741, x_1721); +lean_inc(x_1632); +x_1743 = lean_array_push(x_1742, x_1632); +lean_inc(x_1743); +x_1744 = lean_array_push(x_1743, x_1737); +lean_inc(x_1739); +x_1745 = lean_array_push(x_1744, x_1739); +lean_inc(x_1714); +x_1746 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1746, 0, x_1558); +lean_ctor_set(x_1746, 1, x_1714); +lean_ctor_set(x_1746, 2, x_1745); +x_1747 = lean_array_push(x_1740, x_1710); +x_1748 = lean_array_push(x_1747, x_1718); +lean_inc(x_1706); +x_1749 = lean_array_push(x_1748, x_1706); +lean_inc(x_1749); +x_1750 = lean_array_push(x_1749, x_1746); +lean_inc(x_1739); +x_1751 = lean_array_push(x_1750, x_1739); lean_inc(x_1708); -x_1748 = lean_array_push(x_1747, x_1708); -lean_inc(x_1748); -x_1749 = lean_array_push(x_1748, x_1745); -lean_inc(x_1738); -x_1750 = lean_array_push(x_1749, x_1738); -lean_inc(x_1710); -x_1751 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1751, 0, x_1560); -lean_ctor_set(x_1751, 1, x_1710); -lean_ctor_set(x_1751, 2, x_1750); -x_1752 = lean_array_push(x_1575, x_1751); -x_1753 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1753, 0, x_1560); -lean_ctor_set(x_1753, 1, x_1577); -lean_ctor_set(x_1753, 2, x_1752); -x_1754 = lean_array_push(x_1575, x_1753); -x_1755 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1755, 0, x_1560); -lean_ctor_set(x_1755, 1, x_1577); -lean_ctor_set(x_1755, 2, x_1754); -x_1756 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__63; -lean_inc(x_1551); -x_1757 = lean_name_mk_string(x_1551, x_1756); -x_1758 = l_Lean_Syntax_mkApp___closed__1; -lean_inc(x_1551); -x_1759 = lean_name_mk_string(x_1551, x_1758); -x_1760 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__64; -lean_inc(x_1551); -x_1761 = lean_name_mk_string(x_1551, x_1760); -x_1762 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__58; -x_1763 = lean_name_mk_string(x_1761, x_1762); -x_1764 = lean_array_push(x_1733, x_1590); -x_1765 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1765, 0, x_1560); -lean_ctor_set(x_1765, 1, x_1763); -lean_ctor_set(x_1765, 2, x_1764); -x_1766 = lean_array_push(x_1575, x_1765); +x_1752 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1752, 0, x_1558); +lean_ctor_set(x_1752, 1, x_1708); +lean_ctor_set(x_1752, 2, x_1751); +x_1753 = lean_array_push(x_1573, x_1752); +x_1754 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1754, 0, x_1558); +lean_ctor_set(x_1754, 1, x_1575); +lean_ctor_set(x_1754, 2, x_1753); +x_1755 = lean_array_push(x_1573, x_1754); +x_1756 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1756, 0, x_1558); +lean_ctor_set(x_1756, 1, x_1575); +lean_ctor_set(x_1756, 2, x_1755); +x_1757 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__61; +lean_inc(x_1549); +x_1758 = lean_name_mk_string(x_1549, x_1757); +x_1759 = l_Lean_Syntax_mkApp___closed__1; +lean_inc(x_1549); +x_1760 = lean_name_mk_string(x_1549, x_1759); +x_1761 = lean_array_push(x_1735, x_1588); +x_1762 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1762, 0, x_1558); +lean_ctor_set(x_1762, 1, x_1725); +lean_ctor_set(x_1762, 2, x_1761); +x_1763 = lean_array_push(x_1573, x_1762); +x_1764 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1764, 0, x_1558); +lean_ctor_set(x_1764, 1, x_1575); +lean_ctor_set(x_1764, 2, x_1763); +x_1765 = lean_array_push(x_1566, x_6); +lean_inc(x_1765); +x_1766 = lean_array_push(x_1765, x_1764); +lean_inc(x_1760); x_1767 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1767, 0, x_1560); -lean_ctor_set(x_1767, 1, x_1577); +lean_ctor_set(x_1767, 0, x_1558); +lean_ctor_set(x_1767, 1, x_1760); lean_ctor_set(x_1767, 2, x_1766); -x_1768 = lean_array_push(x_1568, x_5); -lean_inc(x_1768); -x_1769 = lean_array_push(x_1768, x_1767); -lean_inc(x_1759); +lean_inc(x_1743); +x_1768 = lean_array_push(x_1743, x_1767); +lean_inc(x_1739); +x_1769 = lean_array_push(x_1768, x_1739); +lean_inc(x_1714); x_1770 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1770, 0, x_1560); -lean_ctor_set(x_1770, 1, x_1759); +lean_ctor_set(x_1770, 0, x_1558); +lean_ctor_set(x_1770, 1, x_1714); lean_ctor_set(x_1770, 2, x_1769); -lean_inc(x_1742); -x_1771 = lean_array_push(x_1742, x_1770); -lean_inc(x_1738); -x_1772 = lean_array_push(x_1771, x_1738); -lean_inc(x_1716); +lean_inc(x_1749); +x_1771 = lean_array_push(x_1749, x_1770); +lean_inc(x_1739); +x_1772 = lean_array_push(x_1771, x_1739); +lean_inc(x_1708); x_1773 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1773, 0, x_1560); -lean_ctor_set(x_1773, 1, x_1716); +lean_ctor_set(x_1773, 0, x_1558); +lean_ctor_set(x_1773, 1, x_1708); lean_ctor_set(x_1773, 2, x_1772); -lean_inc(x_1748); -x_1774 = lean_array_push(x_1748, x_1773); -lean_inc(x_1738); -x_1775 = lean_array_push(x_1774, x_1738); -lean_inc(x_1710); -x_1776 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1776, 0, x_1560); -lean_ctor_set(x_1776, 1, x_1710); -lean_ctor_set(x_1776, 2, x_1775); -x_1777 = lean_array_push(x_1575, x_1776); -lean_inc(x_1757); +x_1774 = lean_array_push(x_1573, x_1773); +lean_inc(x_1758); +x_1775 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1775, 0, x_1558); +lean_ctor_set(x_1775, 1, x_1758); +lean_ctor_set(x_1775, 2, x_1774); +x_1776 = lean_array_push(x_1566, x_1775); +x_1777 = lean_array_push(x_1776, x_1588); +lean_inc(x_1652); x_1778 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1778, 0, x_1560); -lean_ctor_set(x_1778, 1, x_1757); +lean_ctor_set(x_1778, 0, x_1558); +lean_ctor_set(x_1778, 1, x_1652); lean_ctor_set(x_1778, 2, x_1777); -x_1779 = lean_array_push(x_1568, x_1778); -x_1780 = lean_array_push(x_1779, x_1590); -lean_inc(x_1654); -x_1781 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1781, 0, x_1560); -lean_ctor_set(x_1781, 1, x_1654); -lean_ctor_set(x_1781, 2, x_1780); -x_1782 = lean_array_push(x_1575, x_1781); -x_1783 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1783, 0, x_1560); -lean_ctor_set(x_1783, 1, x_1577); -lean_ctor_set(x_1783, 2, x_1782); -x_1784 = lean_array_push(x_1575, x_1783); -lean_inc(x_1652); -x_1785 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1785, 0, x_1560); -lean_ctor_set(x_1785, 1, x_1652); -lean_ctor_set(x_1785, 2, x_1784); -x_1786 = lean_array_push(x_1682, x_1708); -lean_inc(x_1786); -x_1787 = lean_array_push(x_1786, x_1755); -lean_inc(x_1647); -x_1788 = lean_array_push(x_1787, x_1647); -x_1789 = lean_array_push(x_1788, x_1785); -lean_inc(x_1706); -x_1790 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1790, 0, x_1560); -lean_ctor_set(x_1790, 1, x_1706); -lean_ctor_set(x_1790, 2, x_1789); -x_1791 = l_Lean_mkHole___closed__7; -lean_inc(x_1551); -x_1792 = lean_name_mk_string(x_1551, x_1791); -x_1793 = l_Lean_Name_appendIndexAfter___closed__1; -lean_inc(x_1539); -x_1794 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1794, 0, x_1539); -lean_ctor_set(x_1794, 1, x_1793); -x_1795 = lean_array_push(x_1575, x_1794); -x_1796 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1796, 0, x_1560); -lean_ctor_set(x_1796, 1, x_1792); -lean_ctor_set(x_1796, 2, x_1795); -x_1797 = lean_array_push(x_1575, x_1796); -x_1798 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1798, 0, x_1560); -lean_ctor_set(x_1798, 1, x_1577); -lean_ctor_set(x_1798, 2, x_1797); -x_1799 = lean_array_push(x_1575, x_1798); -x_1800 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1800, 0, x_1560); -lean_ctor_set(x_1800, 1, x_1577); -lean_ctor_set(x_1800, 2, x_1799); -x_1801 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__11; -lean_inc(x_1539); -x_1802 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1802, 0, x_1539); -lean_ctor_set(x_1802, 1, x_1801); -x_1803 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__35; -lean_inc(x_1539); -x_1804 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1804, 0, x_1539); -lean_ctor_set(x_1804, 1, x_1803); -lean_inc(x_1802); -x_1805 = lean_array_push(x_1568, x_1802); -lean_inc(x_1804); -x_1806 = lean_array_push(x_1805, x_1804); -x_1807 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__68; -x_1808 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1808, 0, x_1560); -lean_ctor_set(x_1808, 1, x_1807); -lean_ctor_set(x_1808, 2, x_1806); -x_1809 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__9; -lean_inc(x_1551); -x_1810 = lean_name_mk_string(x_1551, x_1809); -x_1811 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__31; -lean_inc(x_1551); -x_1812 = lean_name_mk_string(x_1551, x_1811); -x_1813 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1813, 0, x_1560); -lean_ctor_set(x_1813, 1, x_1812); -lean_ctor_set(x_1813, 2, x_1561); -x_1814 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__36; -x_1815 = lean_array_push(x_1814, x_1802); -x_1816 = lean_array_push(x_1815, x_1590); -x_1817 = lean_array_push(x_1816, x_1590); -x_1818 = lean_array_push(x_1817, x_1813); -x_1819 = lean_array_push(x_1818, x_1590); -x_1820 = lean_array_push(x_1819, x_1804); -x_1821 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1821, 0, x_1560); -lean_ctor_set(x_1821, 1, x_1810); -lean_ctor_set(x_1821, 2, x_1820); -x_1822 = lean_array_push(x_1568, x_1808); -x_1823 = lean_array_push(x_1822, x_1821); -x_1824 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__66; -x_1825 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1825, 0, x_1560); -lean_ctor_set(x_1825, 1, x_1824); -lean_ctor_set(x_1825, 2, x_1823); -x_1826 = lean_array_push(x_1575, x_1825); -x_1827 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1827, 0, x_1560); -lean_ctor_set(x_1827, 1, x_1577); -lean_ctor_set(x_1827, 2, x_1826); -x_1828 = lean_array_push(x_1768, x_1827); -lean_inc(x_1759); +x_1779 = lean_array_push(x_1573, x_1778); +x_1780 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1780, 0, x_1558); +lean_ctor_set(x_1780, 1, x_1575); +lean_ctor_set(x_1780, 2, x_1779); +x_1781 = lean_array_push(x_1573, x_1780); +lean_inc(x_1650); +x_1782 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1782, 0, x_1558); +lean_ctor_set(x_1782, 1, x_1650); +lean_ctor_set(x_1782, 2, x_1781); +x_1783 = lean_array_push(x_1680, x_1706); +lean_inc(x_1783); +x_1784 = lean_array_push(x_1783, x_1756); +lean_inc(x_1645); +x_1785 = lean_array_push(x_1784, x_1645); +x_1786 = lean_array_push(x_1785, x_1782); +lean_inc(x_1704); +x_1787 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1787, 0, x_1558); +lean_ctor_set(x_1787, 1, x_1704); +lean_ctor_set(x_1787, 2, x_1786); +x_1788 = l_Lean_mkHole___closed__7; +lean_inc(x_1549); +x_1789 = lean_name_mk_string(x_1549, x_1788); +x_1790 = l_Lean_Name_appendIndexAfter___closed__1; +lean_inc(x_1537); +x_1791 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1791, 0, x_1537); +lean_ctor_set(x_1791, 1, x_1790); +x_1792 = lean_array_push(x_1573, x_1791); +x_1793 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1793, 0, x_1558); +lean_ctor_set(x_1793, 1, x_1789); +lean_ctor_set(x_1793, 2, x_1792); +x_1794 = lean_array_push(x_1573, x_1793); +x_1795 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1795, 0, x_1558); +lean_ctor_set(x_1795, 1, x_1575); +lean_ctor_set(x_1795, 2, x_1794); +x_1796 = lean_array_push(x_1573, x_1795); +x_1797 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1797, 0, x_1558); +lean_ctor_set(x_1797, 1, x_1575); +lean_ctor_set(x_1797, 2, x_1796); +x_1798 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__11; +lean_inc(x_1537); +x_1799 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1799, 0, x_1537); +lean_ctor_set(x_1799, 1, x_1798); +x_1800 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__35; +lean_inc(x_1537); +x_1801 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1801, 0, x_1537); +lean_ctor_set(x_1801, 1, x_1800); +lean_inc(x_1799); +x_1802 = lean_array_push(x_1566, x_1799); +lean_inc(x_1801); +x_1803 = lean_array_push(x_1802, x_1801); +x_1804 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__65; +x_1805 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1805, 0, x_1558); +lean_ctor_set(x_1805, 1, x_1804); +lean_ctor_set(x_1805, 2, x_1803); +x_1806 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__9; +lean_inc(x_1549); +x_1807 = lean_name_mk_string(x_1549, x_1806); +x_1808 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__31; +lean_inc(x_1549); +x_1809 = lean_name_mk_string(x_1549, x_1808); +x_1810 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1810, 0, x_1558); +lean_ctor_set(x_1810, 1, x_1809); +lean_ctor_set(x_1810, 2, x_1559); +x_1811 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__36; +x_1812 = lean_array_push(x_1811, x_1799); +x_1813 = lean_array_push(x_1812, x_1588); +x_1814 = lean_array_push(x_1813, x_1588); +x_1815 = lean_array_push(x_1814, x_1810); +x_1816 = lean_array_push(x_1815, x_1588); +x_1817 = lean_array_push(x_1816, x_1801); +x_1818 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1818, 0, x_1558); +lean_ctor_set(x_1818, 1, x_1807); +lean_ctor_set(x_1818, 2, x_1817); +x_1819 = lean_array_push(x_1566, x_1805); +x_1820 = lean_array_push(x_1819, x_1818); +x_1821 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__63; +x_1822 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1822, 0, x_1558); +lean_ctor_set(x_1822, 1, x_1821); +lean_ctor_set(x_1822, 2, x_1820); +x_1823 = lean_array_push(x_1573, x_1822); +x_1824 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1824, 0, x_1558); +lean_ctor_set(x_1824, 1, x_1575); +lean_ctor_set(x_1824, 2, x_1823); +x_1825 = lean_array_push(x_1765, x_1824); +lean_inc(x_1760); +x_1826 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1826, 0, x_1558); +lean_ctor_set(x_1826, 1, x_1760); +lean_ctor_set(x_1826, 2, x_1825); +x_1827 = lean_array_push(x_1743, x_1826); +lean_inc(x_1739); +x_1828 = lean_array_push(x_1827, x_1739); x_1829 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1829, 0, x_1560); -lean_ctor_set(x_1829, 1, x_1759); +lean_ctor_set(x_1829, 0, x_1558); +lean_ctor_set(x_1829, 1, x_1714); lean_ctor_set(x_1829, 2, x_1828); -x_1830 = lean_array_push(x_1742, x_1829); -lean_inc(x_1738); -x_1831 = lean_array_push(x_1830, x_1738); +x_1830 = lean_array_push(x_1749, x_1829); +lean_inc(x_1739); +x_1831 = lean_array_push(x_1830, x_1739); x_1832 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1832, 0, x_1560); -lean_ctor_set(x_1832, 1, x_1716); +lean_ctor_set(x_1832, 0, x_1558); +lean_ctor_set(x_1832, 1, x_1708); lean_ctor_set(x_1832, 2, x_1831); -x_1833 = lean_array_push(x_1748, x_1832); -lean_inc(x_1738); -x_1834 = lean_array_push(x_1833, x_1738); -x_1835 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1835, 0, x_1560); -lean_ctor_set(x_1835, 1, x_1710); -lean_ctor_set(x_1835, 2, x_1834); -x_1836 = lean_array_push(x_1575, x_1835); +x_1833 = lean_array_push(x_1573, x_1832); +x_1834 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1834, 0, x_1558); +lean_ctor_set(x_1834, 1, x_1758); +lean_ctor_set(x_1834, 2, x_1833); +x_1835 = lean_array_push(x_1566, x_1834); +x_1836 = lean_array_push(x_1835, x_1588); +lean_inc(x_1652); x_1837 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1837, 0, x_1560); -lean_ctor_set(x_1837, 1, x_1757); +lean_ctor_set(x_1837, 0, x_1558); +lean_ctor_set(x_1837, 1, x_1652); lean_ctor_set(x_1837, 2, x_1836); -x_1838 = lean_array_push(x_1568, x_1837); -x_1839 = lean_array_push(x_1838, x_1590); -lean_inc(x_1654); -x_1840 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1840, 0, x_1560); -lean_ctor_set(x_1840, 1, x_1654); -lean_ctor_set(x_1840, 2, x_1839); -x_1841 = lean_array_push(x_1575, x_1840); -x_1842 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1842, 0, x_1560); -lean_ctor_set(x_1842, 1, x_1577); -lean_ctor_set(x_1842, 2, x_1841); -x_1843 = lean_array_push(x_1575, x_1842); -lean_inc(x_1652); -x_1844 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1844, 0, x_1560); -lean_ctor_set(x_1844, 1, x_1652); -lean_ctor_set(x_1844, 2, x_1843); -x_1845 = lean_array_push(x_1786, x_1800); -lean_inc(x_1647); -x_1846 = lean_array_push(x_1845, x_1647); -x_1847 = lean_array_push(x_1846, x_1844); +x_1838 = lean_array_push(x_1573, x_1837); +x_1839 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1839, 0, x_1558); +lean_ctor_set(x_1839, 1, x_1575); +lean_ctor_set(x_1839, 2, x_1838); +x_1840 = lean_array_push(x_1573, x_1839); +lean_inc(x_1650); +x_1841 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1841, 0, x_1558); +lean_ctor_set(x_1841, 1, x_1650); +lean_ctor_set(x_1841, 2, x_1840); +x_1842 = lean_array_push(x_1783, x_1797); +lean_inc(x_1645); +x_1843 = lean_array_push(x_1842, x_1645); +x_1844 = lean_array_push(x_1843, x_1841); +x_1845 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1845, 0, x_1558); +lean_ctor_set(x_1845, 1, x_1704); +lean_ctor_set(x_1845, 2, x_1844); +x_1846 = lean_array_push(x_1566, x_1787); +x_1847 = lean_array_push(x_1846, x_1845); x_1848 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1848, 0, x_1560); -lean_ctor_set(x_1848, 1, x_1706); +lean_ctor_set(x_1848, 0, x_1558); +lean_ctor_set(x_1848, 1, x_1575); lean_ctor_set(x_1848, 2, x_1847); -x_1849 = lean_array_push(x_1568, x_1790); -x_1850 = lean_array_push(x_1849, x_1848); -x_1851 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1851, 0, x_1560); -lean_ctor_set(x_1851, 1, x_1577); -lean_ctor_set(x_1851, 2, x_1850); -x_1852 = lean_array_push(x_1575, x_1851); -x_1853 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1853, 0, x_1560); -lean_ctor_set(x_1853, 1, x_1704); -lean_ctor_set(x_1853, 2, x_1852); -x_1854 = lean_array_push(x_1814, x_1670); -x_1855 = lean_array_push(x_1854, x_1590); -x_1856 = lean_array_push(x_1855, x_1590); -x_1857 = lean_array_push(x_1856, x_1700); -x_1858 = lean_array_push(x_1857, x_1702); -x_1859 = lean_array_push(x_1858, x_1853); -x_1860 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1860, 0, x_1560); -lean_ctor_set(x_1860, 1, x_1668); -lean_ctor_set(x_1860, 2, x_1859); -lean_inc(x_1664); -x_1861 = lean_array_push(x_1682, x_1664); -x_1862 = lean_array_push(x_1861, x_1590); -x_1863 = lean_array_push(x_1862, x_1666); -x_1864 = lean_array_push(x_1863, x_1860); -x_1865 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1865, 0, x_1560); -lean_ctor_set(x_1865, 1, x_1660); -lean_ctor_set(x_1865, 2, x_1864); -x_1866 = lean_array_push(x_1581, x_1658); -x_1867 = lean_array_push(x_1866, x_1590); -lean_inc(x_1867); -x_1868 = lean_array_push(x_1867, x_1865); +x_1849 = lean_array_push(x_1573, x_1848); +x_1850 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1850, 0, x_1558); +lean_ctor_set(x_1850, 1, x_1702); +lean_ctor_set(x_1850, 2, x_1849); +x_1851 = lean_array_push(x_1811, x_1668); +x_1852 = lean_array_push(x_1851, x_1588); +x_1853 = lean_array_push(x_1852, x_1588); +x_1854 = lean_array_push(x_1853, x_1698); +x_1855 = lean_array_push(x_1854, x_1700); +x_1856 = lean_array_push(x_1855, x_1850); +x_1857 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1857, 0, x_1558); +lean_ctor_set(x_1857, 1, x_1666); +lean_ctor_set(x_1857, 2, x_1856); +lean_inc(x_1662); +x_1858 = lean_array_push(x_1680, x_1662); +x_1859 = lean_array_push(x_1858, x_1588); +x_1860 = lean_array_push(x_1859, x_1664); +x_1861 = lean_array_push(x_1860, x_1857); +x_1862 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1862, 0, x_1558); +lean_ctor_set(x_1862, 1, x_1658); +lean_ctor_set(x_1862, 2, x_1861); +x_1863 = lean_array_push(x_1579, x_1656); +x_1864 = lean_array_push(x_1863, x_1588); +lean_inc(x_1864); +x_1865 = lean_array_push(x_1864, x_1862); +x_1866 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1866, 0, x_1558); +lean_ctor_set(x_1866, 1, x_1654); +lean_ctor_set(x_1866, 2, x_1865); +x_1867 = lean_array_push(x_1566, x_1866); +x_1868 = lean_array_push(x_1867, x_1588); +lean_inc(x_1652); x_1869 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1869, 0, x_1560); -lean_ctor_set(x_1869, 1, x_1656); +lean_ctor_set(x_1869, 0, x_1558); +lean_ctor_set(x_1869, 1, x_1652); lean_ctor_set(x_1869, 2, x_1868); -x_1870 = lean_array_push(x_1568, x_1869); -x_1871 = lean_array_push(x_1870, x_1590); -lean_inc(x_1654); -x_1872 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1872, 0, x_1560); -lean_ctor_set(x_1872, 1, x_1654); -lean_ctor_set(x_1872, 2, x_1871); -x_1873 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__69; -lean_inc(x_1551); -x_1874 = lean_name_mk_string(x_1551, x_1873); -x_1875 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__70; -lean_inc(x_1551); -x_1876 = lean_name_mk_string(x_1551, x_1875); -x_1877 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__71; -lean_inc(x_1551); -x_1878 = lean_name_mk_string(x_1551, x_1877); -x_1879 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__76; -lean_inc(x_1542); -lean_inc(x_1543); -x_1880 = l_Lean_addMacroScope(x_1543, x_1879, x_1542); -x_1881 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__74; -lean_inc(x_1539); -x_1882 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1882, 0, x_1539); -lean_ctor_set(x_1882, 1, x_1881); -lean_ctor_set(x_1882, 2, x_1880); -lean_ctor_set(x_1882, 3, x_1603); -x_1883 = lean_array_push(x_1575, x_1534); +x_1870 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__66; +lean_inc(x_1549); +x_1871 = lean_name_mk_string(x_1549, x_1870); +x_1872 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__67; +lean_inc(x_1549); +x_1873 = lean_name_mk_string(x_1549, x_1872); +x_1874 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__68; +lean_inc(x_1549); +x_1875 = lean_name_mk_string(x_1549, x_1874); +x_1876 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__73; +lean_inc(x_1540); +lean_inc(x_1541); +x_1877 = l_Lean_addMacroScope(x_1541, x_1876, x_1540); +x_1878 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__71; +lean_inc(x_1537); +x_1879 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1879, 0, x_1537); +lean_ctor_set(x_1879, 1, x_1878); +lean_ctor_set(x_1879, 2, x_1877); +lean_ctor_set(x_1879, 3, x_1601); +x_1880 = lean_array_push(x_1573, x_1532); +x_1881 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1881, 0, x_1558); +lean_ctor_set(x_1881, 1, x_1575); +lean_ctor_set(x_1881, 2, x_1880); +x_1882 = lean_array_push(x_1566, x_1879); +x_1883 = lean_array_push(x_1882, x_1881); +lean_inc(x_1760); x_1884 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1884, 0, x_1560); -lean_ctor_set(x_1884, 1, x_1577); +lean_ctor_set(x_1884, 0, x_1558); +lean_ctor_set(x_1884, 1, x_1760); lean_ctor_set(x_1884, 2, x_1883); -x_1885 = lean_array_push(x_1568, x_1882); -x_1886 = lean_array_push(x_1885, x_1884); -lean_inc(x_1759); -x_1887 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1887, 0, x_1560); -lean_ctor_set(x_1887, 1, x_1759); -lean_ctor_set(x_1887, 2, x_1886); -x_1888 = lean_array_push(x_1739, x_1643); -x_1889 = lean_array_push(x_1888, x_1590); -x_1890 = lean_array_push(x_1889, x_1590); -lean_inc(x_1634); -x_1891 = lean_array_push(x_1890, x_1634); -lean_inc(x_1891); -x_1892 = lean_array_push(x_1891, x_1887); -lean_inc(x_1878); -x_1893 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1893, 0, x_1560); -lean_ctor_set(x_1893, 1, x_1878); -lean_ctor_set(x_1893, 2, x_1892); -x_1894 = lean_array_push(x_1575, x_1893); -lean_inc(x_1876); -x_1895 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1895, 0, x_1560); -lean_ctor_set(x_1895, 1, x_1876); -lean_ctor_set(x_1895, 2, x_1894); -lean_inc(x_1867); -x_1896 = lean_array_push(x_1867, x_1895); -lean_inc(x_1874); +x_1885 = lean_array_push(x_1740, x_1641); +x_1886 = lean_array_push(x_1885, x_1588); +x_1887 = lean_array_push(x_1886, x_1588); +lean_inc(x_1632); +x_1888 = lean_array_push(x_1887, x_1632); +lean_inc(x_1888); +x_1889 = lean_array_push(x_1888, x_1884); +lean_inc(x_1875); +x_1890 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1890, 0, x_1558); +lean_ctor_set(x_1890, 1, x_1875); +lean_ctor_set(x_1890, 2, x_1889); +x_1891 = lean_array_push(x_1573, x_1890); +lean_inc(x_1873); +x_1892 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1892, 0, x_1558); +lean_ctor_set(x_1892, 1, x_1873); +lean_ctor_set(x_1892, 2, x_1891); +lean_inc(x_1864); +x_1893 = lean_array_push(x_1864, x_1892); +lean_inc(x_1871); +x_1894 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1894, 0, x_1558); +lean_ctor_set(x_1894, 1, x_1871); +lean_ctor_set(x_1894, 2, x_1893); +x_1895 = lean_array_push(x_1566, x_1894); +x_1896 = lean_array_push(x_1895, x_1588); +lean_inc(x_1652); x_1897 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1897, 0, x_1560); -lean_ctor_set(x_1897, 1, x_1874); +lean_ctor_set(x_1897, 0, x_1558); +lean_ctor_set(x_1897, 1, x_1652); lean_ctor_set(x_1897, 2, x_1896); -x_1898 = lean_array_push(x_1568, x_1897); -x_1899 = lean_array_push(x_1898, x_1590); -lean_inc(x_1654); -x_1900 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1900, 0, x_1560); -lean_ctor_set(x_1900, 1, x_1654); -lean_ctor_set(x_1900, 2, x_1899); -x_1901 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__81; -lean_inc(x_1542); -lean_inc(x_1543); -x_1902 = l_Lean_addMacroScope(x_1543, x_1901, x_1542); -x_1903 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__79; -lean_inc(x_1539); -x_1904 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1904, 0, x_1539); -lean_ctor_set(x_1904, 1, x_1903); -lean_ctor_set(x_1904, 2, x_1902); -lean_ctor_set(x_1904, 3, x_1603); -x_1905 = l_Lean_Syntax_expandInterpolatedStr___closed__2; -lean_inc(x_1551); -x_1906 = lean_name_mk_string(x_1551, x_1905); -x_1907 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__85; -lean_inc(x_1542); -lean_inc(x_1543); -x_1908 = l_Lean_addMacroScope(x_1543, x_1907, x_1542); -x_1909 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__82; +x_1898 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__78; +lean_inc(x_1540); +lean_inc(x_1541); +x_1899 = l_Lean_addMacroScope(x_1541, x_1898, x_1540); +x_1900 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__76; +lean_inc(x_1537); +x_1901 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1901, 0, x_1537); +lean_ctor_set(x_1901, 1, x_1900); +lean_ctor_set(x_1901, 2, x_1899); +lean_ctor_set(x_1901, 3, x_1601); +x_1902 = l_Lean_TSyntax_expandInterpolatedStr___closed__2; +lean_inc(x_1549); +x_1903 = lean_name_mk_string(x_1549, x_1902); +x_1904 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__82; +lean_inc(x_1540); +lean_inc(x_1541); +x_1905 = l_Lean_addMacroScope(x_1541, x_1904, x_1540); +x_1906 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__79; lean_inc(x_3); -x_1910 = lean_name_mk_string(x_3, x_1909); -x_1911 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1911, 0, x_1910); -lean_ctor_set(x_1911, 1, x_1603); -x_1912 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1912, 0, x_1911); -lean_ctor_set(x_1912, 1, x_1603); -x_1913 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__84; -lean_inc(x_1539); -x_1914 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1914, 0, x_1539); -lean_ctor_set(x_1914, 1, x_1913); -lean_ctor_set(x_1914, 2, x_1908); -lean_ctor_set(x_1914, 3, x_1912); -lean_inc(x_1691); -x_1915 = lean_array_push(x_1684, x_1691); -lean_inc(x_1580); -x_1916 = lean_array_push(x_1915, x_1580); +x_1907 = lean_name_mk_string(x_3, x_1906); +x_1908 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1908, 0, x_1907); +lean_ctor_set(x_1908, 1, x_1601); +x_1909 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1909, 0, x_1908); +lean_ctor_set(x_1909, 1, x_1601); +x_1910 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__81; +lean_inc(x_1537); +x_1911 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1911, 0, x_1537); +lean_ctor_set(x_1911, 1, x_1910); +lean_ctor_set(x_1911, 2, x_1905); +lean_ctor_set(x_1911, 3, x_1909); +lean_inc(x_1689); +x_1912 = lean_array_push(x_1682, x_1689); +lean_inc(x_1578); +x_1913 = lean_array_push(x_1912, x_1578); +x_1914 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1914, 0, x_1558); +lean_ctor_set(x_1914, 1, x_1672); +lean_ctor_set(x_1914, 2, x_1913); +x_1915 = lean_array_push(x_1566, x_1914); +x_1916 = lean_array_push(x_1915, x_1533); x_1917 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1917, 0, x_1560); -lean_ctor_set(x_1917, 1, x_1674); +lean_ctor_set(x_1917, 0, x_1558); +lean_ctor_set(x_1917, 1, x_1575); lean_ctor_set(x_1917, 2, x_1916); -x_1918 = lean_array_push(x_1568, x_1917); -x_1919 = lean_array_push(x_1918, x_1535); +x_1918 = lean_array_push(x_1566, x_1911); +x_1919 = lean_array_push(x_1918, x_1917); +lean_inc(x_1760); x_1920 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1920, 0, x_1560); -lean_ctor_set(x_1920, 1, x_1577); +lean_ctor_set(x_1920, 0, x_1558); +lean_ctor_set(x_1920, 1, x_1760); lean_ctor_set(x_1920, 2, x_1919); -x_1921 = lean_array_push(x_1568, x_1914); -x_1922 = lean_array_push(x_1921, x_1920); -lean_inc(x_1759); +x_1921 = lean_array_push(x_1566, x_1920); +x_1922 = lean_array_push(x_1921, x_1588); x_1923 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1923, 0, x_1560); -lean_ctor_set(x_1923, 1, x_1759); +lean_ctor_set(x_1923, 0, x_1558); +lean_ctor_set(x_1923, 1, x_1575); lean_ctor_set(x_1923, 2, x_1922); -x_1924 = lean_array_push(x_1568, x_1923); -x_1925 = lean_array_push(x_1924, x_1590); -x_1926 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1926, 0, x_1560); -lean_ctor_set(x_1926, 1, x_1577); -lean_ctor_set(x_1926, 2, x_1925); -x_1927 = lean_array_push(x_1581, x_1722); -lean_inc(x_1927); -x_1928 = lean_array_push(x_1927, x_1926); -lean_inc(x_1738); -x_1929 = lean_array_push(x_1928, x_1738); -lean_inc(x_1906); +x_1924 = lean_array_push(x_1579, x_1720); +lean_inc(x_1924); +x_1925 = lean_array_push(x_1924, x_1923); +lean_inc(x_1739); +x_1926 = lean_array_push(x_1925, x_1739); +lean_inc(x_1903); +x_1927 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1927, 0, x_1558); +lean_ctor_set(x_1927, 1, x_1903); +lean_ctor_set(x_1927, 2, x_1926); +x_1928 = lean_array_push(x_1566, x_1689); +x_1929 = lean_array_push(x_1928, x_1927); x_1930 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1930, 0, x_1560); -lean_ctor_set(x_1930, 1, x_1906); +lean_ctor_set(x_1930, 0, x_1558); +lean_ctor_set(x_1930, 1, x_1575); lean_ctor_set(x_1930, 2, x_1929); -x_1931 = lean_array_push(x_1568, x_1691); +x_1931 = lean_array_push(x_1566, x_1901); +lean_inc(x_1931); x_1932 = lean_array_push(x_1931, x_1930); +lean_inc(x_1760); x_1933 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1933, 0, x_1560); -lean_ctor_set(x_1933, 1, x_1577); +lean_ctor_set(x_1933, 0, x_1558); +lean_ctor_set(x_1933, 1, x_1760); lean_ctor_set(x_1933, 2, x_1932); -x_1934 = lean_array_push(x_1568, x_1904); -lean_inc(x_1934); -x_1935 = lean_array_push(x_1934, x_1933); -lean_inc(x_1759); -x_1936 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1936, 0, x_1560); -lean_ctor_set(x_1936, 1, x_1759); -lean_ctor_set(x_1936, 2, x_1935); -x_1937 = lean_array_push(x_1891, x_1936); -lean_inc(x_1878); -x_1938 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1938, 0, x_1560); -lean_ctor_set(x_1938, 1, x_1878); -lean_ctor_set(x_1938, 2, x_1937); -x_1939 = lean_array_push(x_1575, x_1938); -lean_inc(x_1876); -x_1940 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1940, 0, x_1560); -lean_ctor_set(x_1940, 1, x_1876); -lean_ctor_set(x_1940, 2, x_1939); -lean_inc(x_1867); -x_1941 = lean_array_push(x_1867, x_1940); -lean_inc(x_1874); +x_1934 = lean_array_push(x_1888, x_1933); +lean_inc(x_1875); +x_1935 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1935, 0, x_1558); +lean_ctor_set(x_1935, 1, x_1875); +lean_ctor_set(x_1935, 2, x_1934); +x_1936 = lean_array_push(x_1573, x_1935); +lean_inc(x_1873); +x_1937 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1937, 0, x_1558); +lean_ctor_set(x_1937, 1, x_1873); +lean_ctor_set(x_1937, 2, x_1936); +lean_inc(x_1864); +x_1938 = lean_array_push(x_1864, x_1937); +lean_inc(x_1871); +x_1939 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1939, 0, x_1558); +lean_ctor_set(x_1939, 1, x_1871); +lean_ctor_set(x_1939, 2, x_1938); +x_1940 = lean_array_push(x_1566, x_1939); +x_1941 = lean_array_push(x_1940, x_1588); +lean_inc(x_1652); x_1942 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1942, 0, x_1560); -lean_ctor_set(x_1942, 1, x_1874); +lean_ctor_set(x_1942, 0, x_1558); +lean_ctor_set(x_1942, 1, x_1652); lean_ctor_set(x_1942, 2, x_1941); -x_1943 = lean_array_push(x_1568, x_1942); -x_1944 = lean_array_push(x_1943, x_1590); -lean_inc(x_1654); -x_1945 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1945, 0, x_1560); -lean_ctor_set(x_1945, 1, x_1654); -lean_ctor_set(x_1945, 2, x_1944); -x_1946 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__89; -lean_inc(x_1542); -lean_inc(x_1543); -x_1947 = l_Lean_addMacroScope(x_1543, x_1946, x_1542); -x_1948 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__88; -lean_inc(x_1539); -x_1949 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1949, 0, x_1539); -lean_ctor_set(x_1949, 1, x_1948); -lean_ctor_set(x_1949, 2, x_1947); -lean_ctor_set(x_1949, 3, x_1603); -x_1950 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__93; -x_1951 = l_Lean_addMacroScope(x_1543, x_1950, x_1542); -x_1952 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__90; -x_1953 = lean_name_mk_string(x_3, x_1952); -x_1954 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1954, 0, x_1953); -lean_ctor_set(x_1954, 1, x_1603); -x_1955 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1955, 0, x_1954); -lean_ctor_set(x_1955, 1, x_1603); -x_1956 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__92; -lean_inc(x_1539); -x_1957 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1957, 0, x_1539); -lean_ctor_set(x_1957, 1, x_1956); -lean_ctor_set(x_1957, 2, x_1951); -lean_ctor_set(x_1957, 3, x_1955); -x_1958 = l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__12; -lean_inc(x_1539); -x_1959 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1959, 0, x_1539); -lean_ctor_set(x_1959, 1, x_1958); -x_1960 = lean_array_push(x_1575, x_1664); -x_1961 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1961, 0, x_1560); -lean_ctor_set(x_1961, 1, x_1577); -lean_ctor_set(x_1961, 2, x_1960); -x_1962 = lean_array_push(x_1581, x_1959); -x_1963 = lean_array_push(x_1962, x_1961); -x_1964 = lean_array_push(x_1963, x_1580); -x_1965 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__95; -x_1966 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1966, 0, x_1560); -lean_ctor_set(x_1966, 1, x_1965); -lean_ctor_set(x_1966, 2, x_1964); -x_1967 = lean_array_push(x_1575, x_1966); +x_1943 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__86; +lean_inc(x_1540); +lean_inc(x_1541); +x_1944 = l_Lean_addMacroScope(x_1541, x_1943, x_1540); +x_1945 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__85; +lean_inc(x_1537); +x_1946 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1946, 0, x_1537); +lean_ctor_set(x_1946, 1, x_1945); +lean_ctor_set(x_1946, 2, x_1944); +lean_ctor_set(x_1946, 3, x_1601); +x_1947 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__90; +x_1948 = l_Lean_addMacroScope(x_1541, x_1947, x_1540); +x_1949 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__87; +x_1950 = lean_name_mk_string(x_3, x_1949); +x_1951 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1951, 0, x_1950); +lean_ctor_set(x_1951, 1, x_1601); +x_1952 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1952, 0, x_1951); +lean_ctor_set(x_1952, 1, x_1601); +x_1953 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__89; +lean_inc(x_1537); +x_1954 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1954, 0, x_1537); +lean_ctor_set(x_1954, 1, x_1953); +lean_ctor_set(x_1954, 2, x_1948); +lean_ctor_set(x_1954, 3, x_1952); +x_1955 = l___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____closed__12; +lean_inc(x_1537); +x_1956 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1956, 0, x_1537); +lean_ctor_set(x_1956, 1, x_1955); +x_1957 = lean_array_push(x_1573, x_1662); +x_1958 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1958, 0, x_1558); +lean_ctor_set(x_1958, 1, x_1575); +lean_ctor_set(x_1958, 2, x_1957); +x_1959 = lean_array_push(x_1579, x_1956); +x_1960 = lean_array_push(x_1959, x_1958); +x_1961 = lean_array_push(x_1960, x_1578); +x_1962 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__92; +x_1963 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1963, 0, x_1558); +lean_ctor_set(x_1963, 1, x_1962); +lean_ctor_set(x_1963, 2, x_1961); +x_1964 = lean_array_push(x_1573, x_1963); +x_1965 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1965, 0, x_1558); +lean_ctor_set(x_1965, 1, x_1575); +lean_ctor_set(x_1965, 2, x_1964); +x_1966 = lean_array_push(x_1566, x_1954); +x_1967 = lean_array_push(x_1966, x_1965); +lean_inc(x_1760); x_1968 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1968, 0, x_1560); -lean_ctor_set(x_1968, 1, x_1577); +lean_ctor_set(x_1968, 0, x_1558); +lean_ctor_set(x_1968, 1, x_1760); lean_ctor_set(x_1968, 2, x_1967); -x_1969 = lean_array_push(x_1568, x_1957); -x_1970 = lean_array_push(x_1969, x_1968); -lean_inc(x_1759); +x_1969 = lean_array_push(x_1566, x_1968); +x_1970 = lean_array_push(x_1969, x_1588); x_1971 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1971, 0, x_1560); -lean_ctor_set(x_1971, 1, x_1759); +lean_ctor_set(x_1971, 0, x_1558); +lean_ctor_set(x_1971, 1, x_1575); lean_ctor_set(x_1971, 2, x_1970); -x_1972 = lean_array_push(x_1568, x_1971); -x_1973 = lean_array_push(x_1972, x_1590); +x_1972 = lean_array_push(x_1924, x_1971); +x_1973 = lean_array_push(x_1972, x_1739); x_1974 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1974, 0, x_1560); -lean_ctor_set(x_1974, 1, x_1577); +lean_ctor_set(x_1974, 0, x_1558); +lean_ctor_set(x_1974, 1, x_1903); lean_ctor_set(x_1974, 2, x_1973); -x_1975 = lean_array_push(x_1927, x_1974); -x_1976 = lean_array_push(x_1975, x_1738); +x_1975 = lean_array_push(x_1566, x_1679); +x_1976 = lean_array_push(x_1975, x_1974); x_1977 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1977, 0, x_1560); -lean_ctor_set(x_1977, 1, x_1906); +lean_ctor_set(x_1977, 0, x_1558); +lean_ctor_set(x_1977, 1, x_1575); lean_ctor_set(x_1977, 2, x_1976); -x_1978 = lean_array_push(x_1568, x_1681); -x_1979 = lean_array_push(x_1978, x_1977); -x_1980 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1980, 0, x_1560); -lean_ctor_set(x_1980, 1, x_1577); -lean_ctor_set(x_1980, 2, x_1979); -x_1981 = lean_array_push(x_1934, x_1980); -x_1982 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1982, 0, x_1560); -lean_ctor_set(x_1982, 1, x_1759); -lean_ctor_set(x_1982, 2, x_1981); -lean_inc(x_1949); -x_1983 = lean_array_push(x_1739, x_1949); -x_1984 = lean_array_push(x_1983, x_1590); -x_1985 = lean_array_push(x_1984, x_1590); -lean_inc(x_1634); -x_1986 = lean_array_push(x_1985, x_1634); -x_1987 = lean_array_push(x_1986, x_1982); -x_1988 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1988, 0, x_1560); -lean_ctor_set(x_1988, 1, x_1878); -lean_ctor_set(x_1988, 2, x_1987); -x_1989 = lean_array_push(x_1575, x_1988); -x_1990 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1990, 0, x_1560); -lean_ctor_set(x_1990, 1, x_1876); -lean_ctor_set(x_1990, 2, x_1989); -x_1991 = lean_array_push(x_1867, x_1990); +x_1978 = lean_array_push(x_1931, x_1977); +x_1979 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1979, 0, x_1558); +lean_ctor_set(x_1979, 1, x_1760); +lean_ctor_set(x_1979, 2, x_1978); +lean_inc(x_1946); +x_1980 = lean_array_push(x_1740, x_1946); +x_1981 = lean_array_push(x_1980, x_1588); +x_1982 = lean_array_push(x_1981, x_1588); +lean_inc(x_1632); +x_1983 = lean_array_push(x_1982, x_1632); +x_1984 = lean_array_push(x_1983, x_1979); +x_1985 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1985, 0, x_1558); +lean_ctor_set(x_1985, 1, x_1875); +lean_ctor_set(x_1985, 2, x_1984); +x_1986 = lean_array_push(x_1573, x_1985); +x_1987 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1987, 0, x_1558); +lean_ctor_set(x_1987, 1, x_1873); +lean_ctor_set(x_1987, 2, x_1986); +x_1988 = lean_array_push(x_1864, x_1987); +x_1989 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1989, 0, x_1558); +lean_ctor_set(x_1989, 1, x_1871); +lean_ctor_set(x_1989, 2, x_1988); +x_1990 = lean_array_push(x_1566, x_1989); +x_1991 = lean_array_push(x_1990, x_1588); +lean_inc(x_1652); x_1992 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1992, 0, x_1560); -lean_ctor_set(x_1992, 1, x_1874); +lean_ctor_set(x_1992, 0, x_1558); +lean_ctor_set(x_1992, 1, x_1652); lean_ctor_set(x_1992, 2, x_1991); -x_1993 = lean_array_push(x_1568, x_1992); -x_1994 = lean_array_push(x_1993, x_1590); -lean_inc(x_1654); -x_1995 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1995, 0, x_1560); -lean_ctor_set(x_1995, 1, x_1654); -lean_ctor_set(x_1995, 2, x_1994); -x_1996 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__96; -x_1997 = lean_name_mk_string(x_1551, x_1996); -x_1998 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__97; -x_1999 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1999, 0, x_1539); -lean_ctor_set(x_1999, 1, x_1998); -x_2000 = lean_array_push(x_1575, x_1949); +x_1993 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__93; +x_1994 = lean_name_mk_string(x_1549, x_1993); +x_1995 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__94; +x_1996 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1996, 0, x_1537); +lean_ctor_set(x_1996, 1, x_1995); +x_1997 = lean_array_push(x_1573, x_1946); +x_1998 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1998, 0, x_1558); +lean_ctor_set(x_1998, 1, x_1575); +lean_ctor_set(x_1998, 2, x_1997); +x_1999 = lean_array_push(x_1566, x_1996); +x_2000 = lean_array_push(x_1999, x_1998); x_2001 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_2001, 0, x_1560); -lean_ctor_set(x_2001, 1, x_1577); +lean_ctor_set(x_2001, 0, x_1558); +lean_ctor_set(x_2001, 1, x_1994); lean_ctor_set(x_2001, 2, x_2000); -x_2002 = lean_array_push(x_1568, x_1999); -x_2003 = lean_array_push(x_2002, x_2001); +x_2002 = lean_array_push(x_1566, x_2001); +x_2003 = lean_array_push(x_2002, x_1588); x_2004 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_2004, 0, x_1560); -lean_ctor_set(x_2004, 1, x_1997); +lean_ctor_set(x_2004, 0, x_1558); +lean_ctor_set(x_2004, 1, x_1652); lean_ctor_set(x_2004, 2, x_2003); -x_2005 = lean_array_push(x_1568, x_2004); -x_2006 = lean_array_push(x_2005, x_1590); -x_2007 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_2007, 0, x_1560); -lean_ctor_set(x_2007, 1, x_1654); -lean_ctor_set(x_2007, 2, x_2006); -x_2008 = lean_array_push(x_1739, x_1872); -x_2009 = lean_array_push(x_2008, x_1900); -x_2010 = lean_array_push(x_2009, x_1945); -x_2011 = lean_array_push(x_2010, x_1995); -x_2012 = lean_array_push(x_2011, x_2007); -x_2013 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_2013, 0, x_1560); -lean_ctor_set(x_2013, 1, x_1577); -lean_ctor_set(x_2013, 2, x_2012); -x_2014 = lean_array_push(x_1575, x_2013); +x_2005 = lean_array_push(x_1740, x_1869); +x_2006 = lean_array_push(x_2005, x_1897); +x_2007 = lean_array_push(x_2006, x_1942); +x_2008 = lean_array_push(x_2007, x_1992); +x_2009 = lean_array_push(x_2008, x_2004); +x_2010 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_2010, 0, x_1558); +lean_ctor_set(x_2010, 1, x_1575); +lean_ctor_set(x_2010, 2, x_2009); +x_2011 = lean_array_push(x_1573, x_2010); +x_2012 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_2012, 0, x_1558); +lean_ctor_set(x_2012, 1, x_1650); +lean_ctor_set(x_2012, 2, x_2011); +x_2013 = lean_array_push(x_1566, x_1648); +x_2014 = lean_array_push(x_2013, x_2012); x_2015 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_2015, 0, x_1560); -lean_ctor_set(x_2015, 1, x_1652); +lean_ctor_set(x_2015, 0, x_1558); +lean_ctor_set(x_2015, 1, x_1647); lean_ctor_set(x_2015, 2, x_2014); -x_2016 = lean_array_push(x_1568, x_1650); -x_2017 = lean_array_push(x_2016, x_2015); -x_2018 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_2018, 0, x_1560); -lean_ctor_set(x_2018, 1, x_1649); -lean_ctor_set(x_2018, 2, x_2017); -x_2019 = lean_array_push(x_1581, x_1645); -x_2020 = lean_array_push(x_2019, x_1647); -x_2021 = lean_array_push(x_2020, x_2018); +x_2016 = lean_array_push(x_1579, x_1643); +x_2017 = lean_array_push(x_2016, x_1645); +x_2018 = lean_array_push(x_2017, x_2015); +x_2019 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_2019, 0, x_1558); +lean_ctor_set(x_2019, 1, x_1637); +lean_ctor_set(x_2019, 2, x_2018); +x_2020 = lean_array_push(x_1566, x_1635); +x_2021 = lean_array_push(x_2020, x_2019); x_2022 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_2022, 0, x_1560); -lean_ctor_set(x_2022, 1, x_1639); +lean_ctor_set(x_2022, 0, x_1558); +lean_ctor_set(x_2022, 1, x_1634); lean_ctor_set(x_2022, 2, x_2021); -x_2023 = lean_array_push(x_1568, x_1637); +x_2023 = lean_array_push(x_1579, x_1632); x_2024 = lean_array_push(x_2023, x_2022); -x_2025 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_2025, 0, x_1560); -lean_ctor_set(x_2025, 1, x_1636); -lean_ctor_set(x_2025, 2, x_2024); -x_2026 = lean_array_push(x_1581, x_1634); -x_2027 = lean_array_push(x_2026, x_2025); -x_2028 = lean_array_push(x_2027, x_1590); -x_2029 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_2029, 0, x_1560); -lean_ctor_set(x_2029, 1, x_1632); -lean_ctor_set(x_2029, 2, x_2028); -x_2030 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__98; -x_2031 = lean_array_push(x_2030, x_1598); -x_2032 = lean_array_push(x_2031, x_1608); -x_2033 = lean_array_push(x_2032, x_1630); -x_2034 = lean_array_push(x_2033, x_2029); -x_2035 = lean_array_push(x_2034, x_1590); -x_2036 = lean_array_push(x_2035, x_1590); -x_2037 = lean_array_push(x_2036, x_1590); +x_2025 = lean_array_push(x_2024, x_1588); +x_2026 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_2026, 0, x_1558); +lean_ctor_set(x_2026, 1, x_1630); +lean_ctor_set(x_2026, 2, x_2025); +x_2027 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__95; +x_2028 = lean_array_push(x_2027, x_1596); +x_2029 = lean_array_push(x_2028, x_1606); +x_2030 = lean_array_push(x_2029, x_1628); +x_2031 = lean_array_push(x_2030, x_2026); +x_2032 = lean_array_push(x_2031, x_1588); +x_2033 = lean_array_push(x_2032, x_1588); +x_2034 = lean_array_push(x_2033, x_1588); +x_2035 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_2035, 0, x_1558); +lean_ctor_set(x_2035, 1, x_1595); +lean_ctor_set(x_2035, 2, x_2034); +x_2036 = lean_array_push(x_1566, x_1593); +x_2037 = lean_array_push(x_2036, x_2035); x_2038 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_2038, 0, x_1560); -lean_ctor_set(x_2038, 1, x_1597); +lean_ctor_set(x_2038, 0, x_1558); +lean_ctor_set(x_2038, 1, x_1545); lean_ctor_set(x_2038, 2, x_2037); -x_2039 = lean_array_push(x_1568, x_1595); +x_2039 = lean_array_push(x_1566, x_1534); x_2040 = lean_array_push(x_2039, x_2038); x_2041 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_2041, 0, x_1560); -lean_ctor_set(x_2041, 1, x_1547); +lean_ctor_set(x_2041, 0, x_1558); +lean_ctor_set(x_2041, 1, x_1575); lean_ctor_set(x_2041, 2, x_2040); -x_2042 = lean_array_push(x_1568, x_1536); -x_2043 = lean_array_push(x_2042, x_2041); -x_2044 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_2044, 0, x_1560); -lean_ctor_set(x_2044, 1, x_1577); -lean_ctor_set(x_2044, 2, x_2043); -if (lean_is_scalar(x_1541)) { - x_2045 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_1539)) { + x_2042 = lean_alloc_ctor(0, 2, 0); } else { - x_2045 = x_1541; + x_2042 = x_1539; } -lean_ctor_set(x_2045, 0, x_2044); -lean_ctor_set(x_2045, 1, x_1540); -return x_2045; +lean_ctor_set(x_2042, 0, x_2041); +lean_ctor_set(x_2042, 1, x_1538); +return x_2042; } } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__1() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__1() { _start: { lean_object* x_1; @@ -26599,17 +27250,17 @@ x_1 = lean_mk_string_from_bytes("doubleQuotedName", 16); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__2() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_mkHole___closed__6; -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__1; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__3() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__3() { _start: { lean_object* x_1; lean_object* x_2; @@ -26618,13 +27269,13 @@ x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__4() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Tactic_dsimpKind___closed__3; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__3; +x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__3; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -26632,7 +27283,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -26642,7 +27293,7 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -26652,31 +27303,31 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__7() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__6; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__6; 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_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__8() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__7; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__7; x_3 = lean_alloc_ctor(1, 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_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__9() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__9() { _start: { lean_object* x_1; @@ -26684,17 +27335,7 @@ x_1 = lean_mk_string_from_bytes("\"dsimp \"", 8); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__10() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_mkHole___closed__4; -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__1; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__11() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__10() { _start: { lean_object* x_1; @@ -26702,27 +27343,27 @@ x_1 = lean_mk_string_from_bytes("syntax", 6); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__12() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__10; -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__11; +x_1 = l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___closed__2; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__10; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__13() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_mkHole___closed__6; -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__7; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__6; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__14() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -26732,13 +27373,13 @@ x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__15() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__13; -x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__14; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__12; +x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__13; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -26746,7 +27387,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__16() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__15() { _start: { lean_object* x_1; @@ -26754,25 +27395,17 @@ x_1 = lean_mk_string_from_bytes("namedName", 9); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__17() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__10; -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__16; +x_1 = l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___closed__2; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__15; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__18() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("name", 4); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__19() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__17() { _start: { lean_object* x_1; @@ -26780,17 +27413,17 @@ x_1 = lean_mk_string_from_bytes("atom", 4); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__20() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__addPrec__1___closed__2; -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__19; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__17; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__21() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__19() { _start: { lean_object* x_1; @@ -26798,27 +27431,27 @@ x_1 = lean_mk_string_from_bytes("stx_?", 5); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__22() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__21; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__19; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__23() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__addPrec__1___closed__2; -x_2 = l_Lean_Syntax_expandInterpolatedStr___closed__2; +x_2 = l_Lean_TSyntax_expandInterpolatedStr___closed__2; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__24() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__22() { _start: { lean_object* x_1; @@ -26826,17 +27459,17 @@ x_1 = lean_mk_string_from_bytes("cat", 3); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__25() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__addPrec__1___closed__2; -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__24; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__22; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__26() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -26848,19 +27481,19 @@ lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__27() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__26; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__24; x_3 = lean_alloc_ctor(1, 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_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__28() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__26() { _start: { lean_object* x_1; @@ -26868,22 +27501,22 @@ x_1 = lean_mk_string_from_bytes("discharger", 10); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__29() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__27() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__28; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__26; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__30() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__28; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__26; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__29; +x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__27; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -26891,51 +27524,51 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__31() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__29() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__28; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__26; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__32() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__30() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_tacticErw_______closed__2; -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__28; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__26; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__33() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__31() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__32; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__30; 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_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__34() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__32() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__33; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__31; x_3 = lean_alloc_ctor(1, 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_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__35() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__33() { _start: { lean_object* x_1; @@ -26943,17 +27576,17 @@ x_1 = lean_mk_string_from_bytes("nonReserved", 11); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__36() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__34() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__addPrec__1___closed__2; -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__35; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__33; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__37() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__35() { _start: { lean_object* x_1; @@ -26961,7 +27594,7 @@ x_1 = lean_mk_string_from_bytes("&", 1); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__38() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__36() { _start: { lean_object* x_1; @@ -26969,7 +27602,7 @@ x_1 = lean_mk_string_from_bytes("\"only \"", 7); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__39() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__37() { _start: { lean_object* x_1; @@ -26977,7 +27610,7 @@ x_1 = lean_mk_string_from_bytes("\"[\"", 3); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__40() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__38() { _start: { lean_object* x_1; @@ -26985,17 +27618,17 @@ x_1 = lean_mk_string_from_bytes("stx_,*", 6); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__41() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__39() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__40; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__38; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__42() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__40() { _start: { lean_object* x_1; @@ -27003,17 +27636,17 @@ x_1 = lean_mk_string_from_bytes("stx_<|>_", 8); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__43() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__41() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__42; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__40; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__44() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__42() { _start: { lean_object* x_1; @@ -27021,22 +27654,22 @@ x_1 = lean_mk_string_from_bytes("simpErase", 9); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__45() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__43() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__44; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__42; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__46() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__44() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__44; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__42; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__45; +x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__43; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -27044,51 +27677,51 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__47() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__45() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__44; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__42; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__48() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__46() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_tacticErw_______closed__2; -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__44; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__42; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__49() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__47() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__48; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__46; 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_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__50() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__48() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__49; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__47; x_3 = lean_alloc_ctor(1, 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_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__51() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__49() { _start: { lean_object* x_1; @@ -27096,7 +27729,7 @@ x_1 = lean_mk_string_from_bytes("<|>", 3); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__52() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__50() { _start: { lean_object* x_1; @@ -27104,22 +27737,22 @@ x_1 = lean_mk_string_from_bytes("simpLemma", 9); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__53() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__51() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__52; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__50; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__54() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__52() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__52; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__50; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__53; +x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__51; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -27127,51 +27760,51 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__55() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__53() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__52; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__50; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__56() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__54() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_tacticErw_______closed__2; -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__52; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__50; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__57() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__55() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__56; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__54; 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_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__58() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__56() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__57; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__55; x_3 = lean_alloc_ctor(1, 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_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__59() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__57() { _start: { lean_object* x_1; @@ -27179,7 +27812,7 @@ x_1 = lean_mk_string_from_bytes(",*", 2); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__60() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__58() { _start: { lean_object* x_1; @@ -27187,7 +27820,7 @@ x_1 = lean_mk_string_from_bytes("\"]\"", 3); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__61() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__59() { _start: { lean_object* x_1; @@ -27195,22 +27828,22 @@ x_1 = lean_mk_string_from_bytes("location", 8); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__62() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__60() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__61; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__59; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__63() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__61() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__61; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__59; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__62; +x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__60; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -27218,51 +27851,51 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__64() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__62() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__61; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__59; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__65() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__63() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_tacticErw_______closed__2; -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__61; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__59; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__66() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__64() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__65; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__63; 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_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__67() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__65() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__66; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__64; x_3 = lean_alloc_ctor(1, 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_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__68() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__66() { _start: { lean_object* x_1; @@ -27270,22 +27903,22 @@ x_1 = lean_mk_string_from_bytes("tactic", 6); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__69() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__67() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__68; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__66; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__70() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__68() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__68; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__66; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__69; +x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__67; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -27293,17 +27926,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__71() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__69() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__68; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__66; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__72() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__70() { _start: { lean_object* x_1; lean_object* x_2; @@ -27312,27 +27945,27 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__73() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__71() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__72; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__70; x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__12; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__74() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__72() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__73; -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__15; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__71; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__14; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__75() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__73() { _start: { lean_object* x_1; @@ -27340,22 +27973,22 @@ x_1 = lean_mk_string_from_bytes("simpAll", 7); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__76() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__74() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__75; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__73; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__77() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__75() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__75; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__73; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__76; +x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__74; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -27363,51 +27996,51 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__78() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__76() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__75; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__73; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__79() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__77() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_tacticErw_______closed__2; -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__75; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__73; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__80() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__78() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__79; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__77; 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_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__81() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__79() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__80; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__78; x_3 = lean_alloc_ctor(1, 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_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__82() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__80() { _start: { lean_object* x_1; @@ -27415,7 +28048,7 @@ x_1 = lean_mk_string_from_bytes("\"simp_all \"", 11); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__83() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__81() { _start: { lean_object* x_1; @@ -27423,22 +28056,22 @@ x_1 = lean_mk_string_from_bytes("simp", 4); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__84() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__82() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__83; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__81; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__85() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__83() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__83; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__81; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__84; +x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__82; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -27446,51 +28079,51 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__86() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__84() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__83; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__81; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__87() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__85() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_tacticErw_______closed__2; -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__83; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__81; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__88() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__86() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__87; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__85; 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_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__89() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__87() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__88; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__86; x_3 = lean_alloc_ctor(1, 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_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__90() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__88() { _start: { lean_object* x_1; @@ -27498,7 +28131,7 @@ x_1 = lean_mk_string_from_bytes("\"simp \"", 7); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__91() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__89() { _start: { lean_object* x_1; @@ -27506,22 +28139,22 @@ x_1 = lean_mk_string_from_bytes("simpStar", 8); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__92() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__90() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__91; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__89; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__93() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__91() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__91; +x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__89; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__92; +x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__90; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -27529,55 +28162,55 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__94() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__92() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__91; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__89; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__95() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__93() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_tacticErw_______closed__2; -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__91; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__89; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__96() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__94() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__95; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__93; 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_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__97() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__95() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__96; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__94; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; uint8_t x_5; -x_4 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__2; +x_4 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__2; lean_inc(x_1); x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); if (x_5 == 0) @@ -27616,7 +28249,7 @@ x_21 = lean_name_eq(x_19, x_20); lean_dec(x_19); if (x_21 == 0) { -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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; +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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_inc(x_2); x_22 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); x_23 = lean_ctor_get(x_22, 0); @@ -27633,25 +28266,25 @@ lean_inc(x_23); x_28 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_28, 0, x_23); lean_ctor_set(x_28, 1, x_27); -x_29 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__5; +x_29 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__5; lean_inc(x_25); lean_inc(x_26); x_30 = l_Lean_addMacroScope(x_26, x_29, x_25); x_31 = lean_box(0); -x_32 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__4; -x_33 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__8; +x_32 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__4; +x_33 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__8; x_34 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_34, 0, x_23); lean_ctor_set(x_34, 1, x_32); lean_ctor_set(x_34, 2, x_30); lean_ctor_set(x_34, 3, x_33); -x_35 = l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__4; +x_35 = l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__4; lean_inc(x_28); x_36 = lean_array_push(x_35, x_28); x_37 = lean_array_push(x_36, x_28); x_38 = lean_array_push(x_37, x_34); x_39 = lean_box(2); -x_40 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__2; +x_40 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__2; x_41 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_41, 0, x_39); lean_ctor_set(x_41, 1, x_40); @@ -27663,13 +28296,13 @@ lean_inc(x_43); x_44 = lean_ctor_get(x_42, 1); lean_inc(x_44); lean_dec(x_42); -x_45 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__9; +x_45 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__9; x_46 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_46, 0, x_43); lean_ctor_set(x_46, 1, x_45); x_47 = l_Lean_mkOptionalNode___closed__2; x_48 = lean_array_push(x_47, x_46); -x_49 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__15; +x_49 = l_Lean_Syntax_mkStrLit___closed__2; x_50 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_50, 0, x_39); lean_ctor_set(x_50, 1, x_49); @@ -27681,7 +28314,7 @@ lean_inc(x_52); x_53 = lean_ctor_get(x_51, 1); lean_inc(x_53); lean_dec(x_51); -x_54 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__11; +x_54 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__10; lean_inc(x_52); x_55 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_55, 0, x_52); @@ -27691,7 +28324,7 @@ lean_inc(x_52); x_57 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_57, 0, x_52); lean_ctor_set(x_57, 1, x_56); -x_58 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__18; +x_58 = l_Lean_Syntax_mkNameLit___closed__1; lean_inc(x_52); x_59 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_59, 0, x_52); @@ -27715,29 +28348,29 @@ lean_inc(x_11); x_68 = lean_array_push(x_67, x_11); lean_inc(x_63); x_69 = lean_array_push(x_68, x_63); -x_70 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__17; +x_70 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__16; x_71 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_71, 0, x_39); lean_ctor_set(x_71, 1, x_70); lean_ctor_set(x_71, 2, x_69); x_72 = lean_array_push(x_47, x_71); -x_73 = l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__2; +x_73 = l_Lean_mkNullNode___closed__2; x_74 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_74, 0, x_39); lean_ctor_set(x_74, 1, x_73); lean_ctor_set(x_74, 2, x_72); x_75 = lean_array_push(x_47, x_13); -x_76 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__20; +x_76 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__18; x_77 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_77, 0, x_39); lean_ctor_set(x_77, 1, x_76); lean_ctor_set(x_77, 2, x_75); -x_78 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__57; +x_78 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__55; lean_inc(x_25); lean_inc(x_26); x_79 = l_Lean_addMacroScope(x_26, x_78, x_25); -x_80 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__56; -x_81 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__27; +x_80 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__54; +x_81 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__25; lean_inc(x_52); x_82 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_82, 0, x_52); @@ -27748,7 +28381,7 @@ x_83 = l_Lean_Syntax_mkApp___closed__3; x_84 = lean_array_push(x_83, x_82); x_85 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__12; x_86 = lean_array_push(x_84, x_85); -x_87 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__25; +x_87 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__23; x_88 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_88, 0, x_39); lean_ctor_set(x_88, 1, x_87); @@ -27763,7 +28396,7 @@ lean_inc(x_91); x_92 = lean_array_push(x_91, x_90); lean_inc(x_63); x_93 = lean_array_push(x_92, x_63); -x_94 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__23; +x_94 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__21; x_95 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_95, 0, x_39); lean_ctor_set(x_95, 1, x_94); @@ -27776,17 +28409,17 @@ lean_ctor_set(x_97, 1, x_96); x_98 = lean_array_push(x_83, x_95); lean_inc(x_97); x_99 = lean_array_push(x_98, x_97); -x_100 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__22; +x_100 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__20; x_101 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_101, 0, x_39); lean_ctor_set(x_101, 1, x_100); lean_ctor_set(x_101, 2, x_99); -x_102 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__31; +x_102 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__29; lean_inc(x_25); lean_inc(x_26); x_103 = l_Lean_addMacroScope(x_26, x_102, x_25); -x_104 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__30; -x_105 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__34; +x_104 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__28; +x_105 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__32; lean_inc(x_52); x_106 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_106, 0, x_52); @@ -27819,12 +28452,12 @@ x_117 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_117, 0, x_39); lean_ctor_set(x_117, 1, x_100); lean_ctor_set(x_117, 2, x_116); -x_118 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__37; +x_118 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__35; lean_inc(x_52); x_119 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_119, 0, x_52); lean_ctor_set(x_119, 1, x_118); -x_120 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__38; +x_120 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__36; lean_inc(x_52); x_121 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_121, 0, x_52); @@ -27836,7 +28469,7 @@ lean_ctor_set(x_123, 1, x_49); lean_ctor_set(x_123, 2, x_122); x_124 = lean_array_push(x_83, x_119); x_125 = lean_array_push(x_124, x_123); -x_126 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__36; +x_126 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__34; x_127 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_127, 0, x_39); lean_ctor_set(x_127, 1, x_126); @@ -27861,7 +28494,7 @@ x_135 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_135, 0, x_39); lean_ctor_set(x_135, 1, x_100); lean_ctor_set(x_135, 2, x_134); -x_136 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__39; +x_136 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__37; lean_inc(x_52); x_137 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_137, 0, x_52); @@ -27876,12 +28509,12 @@ x_141 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_141, 0, x_39); lean_ctor_set(x_141, 1, x_76); lean_ctor_set(x_141, 2, x_140); -x_142 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__47; +x_142 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__45; lean_inc(x_25); lean_inc(x_26); x_143 = l_Lean_addMacroScope(x_26, x_142, x_25); -x_144 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__46; -x_145 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__50; +x_144 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__44; +x_145 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__48; lean_inc(x_52); x_146 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_146, 0, x_52); @@ -27894,17 +28527,17 @@ x_149 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_149, 0, x_39); lean_ctor_set(x_149, 1, x_87); lean_ctor_set(x_149, 2, x_148); -x_150 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__51; +x_150 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__49; lean_inc(x_52); x_151 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_151, 0, x_52); lean_ctor_set(x_151, 1, x_150); -x_152 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__55; +x_152 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__53; lean_inc(x_25); lean_inc(x_26); x_153 = l_Lean_addMacroScope(x_26, x_152, x_25); -x_154 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__54; -x_155 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__58; +x_154 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__52; +x_155 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__56; lean_inc(x_52); x_156 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_156, 0, x_52); @@ -27920,7 +28553,7 @@ lean_ctor_set(x_159, 2, x_158); x_160 = lean_array_push(x_35, x_149); x_161 = lean_array_push(x_160, x_151); x_162 = lean_array_push(x_161, x_159); -x_163 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__43; +x_163 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__41; x_164 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_164, 0, x_39); lean_ctor_set(x_164, 1, x_163); @@ -27938,19 +28571,19 @@ x_169 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_169, 0, x_39); lean_ctor_set(x_169, 1, x_94); lean_ctor_set(x_169, 2, x_168); -x_170 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__59; +x_170 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__57; lean_inc(x_52); x_171 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_171, 0, x_52); lean_ctor_set(x_171, 1, x_170); x_172 = lean_array_push(x_83, x_169); x_173 = lean_array_push(x_172, x_171); -x_174 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__41; +x_174 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__39; x_175 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_175, 0, x_39); lean_ctor_set(x_175, 1, x_174); lean_ctor_set(x_175, 2, x_173); -x_176 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__60; +x_176 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__58; lean_inc(x_52); x_177 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_177, 0, x_52); @@ -27987,12 +28620,12 @@ x_191 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_191, 0, x_39); lean_ctor_set(x_191, 1, x_100); lean_ctor_set(x_191, 2, x_190); -x_192 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__64; +x_192 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__62; lean_inc(x_25); lean_inc(x_26); x_193 = l_Lean_addMacroScope(x_26, x_192, x_25); -x_194 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__63; -x_195 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__67; +x_194 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__61; +x_195 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__65; lean_inc(x_52); x_196 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_196, 0, x_52); @@ -28038,15 +28671,15 @@ lean_inc(x_52); x_217 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_217, 0, x_52); lean_ctor_set(x_217, 1, x_216); -x_218 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__71; +x_218 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__69; x_219 = l_Lean_addMacroScope(x_26, x_218, x_25); -x_220 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__70; +x_220 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__68; x_221 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_221, 0, x_52); lean_ctor_set(x_221, 1, x_220); lean_ctor_set(x_221, 2, x_219); lean_ctor_set(x_221, 3, x_31); -x_222 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__74; +x_222 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__72; x_223 = lean_array_push(x_222, x_55); x_224 = lean_array_push(x_223, x_85); x_225 = lean_array_push(x_224, x_74); @@ -28054,7 +28687,7 @@ x_226 = lean_array_push(x_225, x_85); x_227 = lean_array_push(x_226, x_215); x_228 = lean_array_push(x_227, x_217); x_229 = lean_array_push(x_228, x_221); -x_230 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__12; +x_230 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__11; x_231 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_231, 0, x_39); lean_ctor_set(x_231, 1, x_230); @@ -28068,910 +28701,913 @@ lean_ctor_set(x_233, 1, x_232); x_234 = l_Lean_mkHole___closed__4; x_235 = l_Lean_mkHole___closed__2; x_236 = l_Lean_Parser_Tactic_tacticErw_______closed__2; -x_237 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1(x_234, x_11, x_235, x_236, x_15, x_233, x_2, x_53); -return x_237; +x_237 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__17; +x_238 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1(x_234, x_11, x_235, x_236, x_237, x_15, x_233, x_2, x_53); +return x_238; } else { -lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; +lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_inc(x_2); -x_238 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_239 = lean_ctor_get(x_238, 0); -lean_inc(x_239); -x_240 = lean_ctor_get(x_238, 1); +x_239 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_240 = lean_ctor_get(x_239, 0); lean_inc(x_240); -lean_dec(x_238); -x_241 = lean_ctor_get(x_2, 2); +x_241 = lean_ctor_get(x_239, 1); lean_inc(x_241); -x_242 = lean_ctor_get(x_2, 1); +lean_dec(x_239); +x_242 = lean_ctor_get(x_2, 2); lean_inc(x_242); -x_243 = l_Lean_Name_reprPrec___closed__3; -lean_inc(x_239); -x_244 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_244, 0, x_239); -lean_ctor_set(x_244, 1, x_243); -x_245 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__78; -lean_inc(x_241); +x_243 = lean_ctor_get(x_2, 1); +lean_inc(x_243); +x_244 = l_Lean_Name_reprPrec___closed__3; +lean_inc(x_240); +x_245 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_245, 0, x_240); +lean_ctor_set(x_245, 1, x_244); +x_246 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__76; lean_inc(x_242); -x_246 = l_Lean_addMacroScope(x_242, x_245, x_241); -x_247 = lean_box(0); -x_248 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__77; -x_249 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__81; -x_250 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_250, 0, x_239); -lean_ctor_set(x_250, 1, x_248); -lean_ctor_set(x_250, 2, x_246); -lean_ctor_set(x_250, 3, x_249); -x_251 = l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__4; -lean_inc(x_244); -x_252 = lean_array_push(x_251, x_244); -x_253 = lean_array_push(x_252, x_244); -x_254 = lean_array_push(x_253, x_250); -x_255 = lean_box(2); -x_256 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__2; -x_257 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_257, 0, x_255); -lean_ctor_set(x_257, 1, x_256); -lean_ctor_set(x_257, 2, x_254); +lean_inc(x_243); +x_247 = l_Lean_addMacroScope(x_243, x_246, x_242); +x_248 = lean_box(0); +x_249 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__75; +x_250 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__79; +x_251 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_251, 0, x_240); +lean_ctor_set(x_251, 1, x_249); +lean_ctor_set(x_251, 2, x_247); +lean_ctor_set(x_251, 3, x_250); +x_252 = l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__4; +lean_inc(x_245); +x_253 = lean_array_push(x_252, x_245); +x_254 = lean_array_push(x_253, x_245); +x_255 = lean_array_push(x_254, x_251); +x_256 = lean_box(2); +x_257 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__2; +x_258 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_258, 0, x_256); +lean_ctor_set(x_258, 1, x_257); +lean_ctor_set(x_258, 2, x_255); lean_inc(x_2); -x_258 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_240); -x_259 = lean_ctor_get(x_258, 0); -lean_inc(x_259); -x_260 = lean_ctor_get(x_258, 1); +x_259 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_241); +x_260 = lean_ctor_get(x_259, 0); lean_inc(x_260); -lean_dec(x_258); -x_261 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__82; -x_262 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_262, 0, x_259); -lean_ctor_set(x_262, 1, x_261); -x_263 = l_Lean_mkOptionalNode___closed__2; -x_264 = lean_array_push(x_263, x_262); -x_265 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__15; -x_266 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_266, 0, x_255); -lean_ctor_set(x_266, 1, x_265); -lean_ctor_set(x_266, 2, x_264); +x_261 = lean_ctor_get(x_259, 1); +lean_inc(x_261); +lean_dec(x_259); +x_262 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__80; +x_263 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_263, 0, x_260); +lean_ctor_set(x_263, 1, x_262); +x_264 = l_Lean_mkOptionalNode___closed__2; +x_265 = lean_array_push(x_264, x_263); +x_266 = l_Lean_Syntax_mkStrLit___closed__2; +x_267 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_267, 0, x_256); +lean_ctor_set(x_267, 1, x_266); +lean_ctor_set(x_267, 2, x_265); lean_inc(x_2); -x_267 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_260); -x_268 = lean_ctor_get(x_267, 0); -lean_inc(x_268); -x_269 = lean_ctor_get(x_267, 1); +x_268 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_261); +x_269 = lean_ctor_get(x_268, 0); lean_inc(x_269); -lean_dec(x_267); -x_270 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__11; -lean_inc(x_268); -x_271 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_271, 0, x_268); -lean_ctor_set(x_271, 1, x_270); -x_272 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__4; -lean_inc(x_268); -x_273 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_273, 0, x_268); -lean_ctor_set(x_273, 1, x_272); -x_274 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__18; -lean_inc(x_268); -x_275 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_275, 0, x_268); -lean_ctor_set(x_275, 1, x_274); -x_276 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__8; -lean_inc(x_268); -x_277 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_277, 0, x_268); -lean_ctor_set(x_277, 1, x_276); -x_278 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__8; -lean_inc(x_268); -x_279 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_279, 0, x_268); -lean_ctor_set(x_279, 1, x_278); -x_280 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__37; -lean_inc(x_273); -x_281 = lean_array_push(x_280, x_273); -x_282 = lean_array_push(x_281, x_275); -x_283 = lean_array_push(x_282, x_277); +x_270 = lean_ctor_get(x_268, 1); +lean_inc(x_270); +lean_dec(x_268); +x_271 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__10; +lean_inc(x_269); +x_272 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_272, 0, x_269); +lean_ctor_set(x_272, 1, x_271); +x_273 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__4; +lean_inc(x_269); +x_274 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_274, 0, x_269); +lean_ctor_set(x_274, 1, x_273); +x_275 = l_Lean_Syntax_mkNameLit___closed__1; +lean_inc(x_269); +x_276 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_276, 0, x_269); +lean_ctor_set(x_276, 1, x_275); +x_277 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__8; +lean_inc(x_269); +x_278 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_278, 0, x_269); +lean_ctor_set(x_278, 1, x_277); +x_279 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__8; +lean_inc(x_269); +x_280 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_280, 0, x_269); +lean_ctor_set(x_280, 1, x_279); +x_281 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__37; +lean_inc(x_274); +x_282 = lean_array_push(x_281, x_274); +x_283 = lean_array_push(x_282, x_276); +x_284 = lean_array_push(x_283, x_278); lean_inc(x_11); -x_284 = lean_array_push(x_283, x_11); -lean_inc(x_279); -x_285 = lean_array_push(x_284, x_279); -x_286 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__17; -x_287 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_287, 0, x_255); -lean_ctor_set(x_287, 1, x_286); -lean_ctor_set(x_287, 2, x_285); -x_288 = lean_array_push(x_263, x_287); -x_289 = l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__2; -x_290 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_290, 0, x_255); -lean_ctor_set(x_290, 1, x_289); -lean_ctor_set(x_290, 2, x_288); -x_291 = lean_array_push(x_263, x_13); -x_292 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__20; -x_293 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_293, 0, x_255); -lean_ctor_set(x_293, 1, x_292); -lean_ctor_set(x_293, 2, x_291); -x_294 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__57; -lean_inc(x_241); +x_285 = lean_array_push(x_284, x_11); +lean_inc(x_280); +x_286 = lean_array_push(x_285, x_280); +x_287 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__16; +x_288 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_288, 0, x_256); +lean_ctor_set(x_288, 1, x_287); +lean_ctor_set(x_288, 2, x_286); +x_289 = lean_array_push(x_264, x_288); +x_290 = l_Lean_mkNullNode___closed__2; +x_291 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_291, 0, x_256); +lean_ctor_set(x_291, 1, x_290); +lean_ctor_set(x_291, 2, x_289); +x_292 = lean_array_push(x_264, x_13); +x_293 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__18; +x_294 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_294, 0, x_256); +lean_ctor_set(x_294, 1, x_293); +lean_ctor_set(x_294, 2, x_292); +x_295 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__55; lean_inc(x_242); -x_295 = l_Lean_addMacroScope(x_242, x_294, x_241); -x_296 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__56; -x_297 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__27; -lean_inc(x_268); -x_298 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_298, 0, x_268); -lean_ctor_set(x_298, 1, x_296); -lean_ctor_set(x_298, 2, x_295); -lean_ctor_set(x_298, 3, x_297); -x_299 = l_Lean_Syntax_mkApp___closed__3; -x_300 = lean_array_push(x_299, x_298); -x_301 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__12; -x_302 = lean_array_push(x_300, x_301); -x_303 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__25; -x_304 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_304, 0, x_255); -lean_ctor_set(x_304, 1, x_303); -lean_ctor_set(x_304, 2, x_302); -x_305 = lean_array_push(x_263, x_304); -x_306 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_306, 0, x_255); -lean_ctor_set(x_306, 1, x_289); -lean_ctor_set(x_306, 2, x_305); -x_307 = lean_array_push(x_251, x_273); -lean_inc(x_307); -x_308 = lean_array_push(x_307, x_306); -lean_inc(x_279); -x_309 = lean_array_push(x_308, x_279); -x_310 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__23; -x_311 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_311, 0, x_255); -lean_ctor_set(x_311, 1, x_310); -lean_ctor_set(x_311, 2, x_309); -x_312 = l_Lean_Name_toString_maybePseudoSyntax___closed__2; -lean_inc(x_268); -x_313 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_313, 0, x_268); -lean_ctor_set(x_313, 1, x_312); -x_314 = lean_array_push(x_299, x_311); -lean_inc(x_313); -x_315 = lean_array_push(x_314, x_313); -x_316 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__22; -x_317 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_317, 0, x_255); -lean_ctor_set(x_317, 1, x_316); -lean_ctor_set(x_317, 2, x_315); -x_318 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__31; -lean_inc(x_241); +lean_inc(x_243); +x_296 = l_Lean_addMacroScope(x_243, x_295, x_242); +x_297 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__54; +x_298 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__25; +lean_inc(x_269); +x_299 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_299, 0, x_269); +lean_ctor_set(x_299, 1, x_297); +lean_ctor_set(x_299, 2, x_296); +lean_ctor_set(x_299, 3, x_298); +x_300 = l_Lean_Syntax_mkApp___closed__3; +x_301 = lean_array_push(x_300, x_299); +x_302 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__12; +x_303 = lean_array_push(x_301, x_302); +x_304 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__23; +x_305 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_305, 0, x_256); +lean_ctor_set(x_305, 1, x_304); +lean_ctor_set(x_305, 2, x_303); +x_306 = lean_array_push(x_264, x_305); +x_307 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_307, 0, x_256); +lean_ctor_set(x_307, 1, x_290); +lean_ctor_set(x_307, 2, x_306); +x_308 = lean_array_push(x_252, x_274); +lean_inc(x_308); +x_309 = lean_array_push(x_308, x_307); +lean_inc(x_280); +x_310 = lean_array_push(x_309, x_280); +x_311 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__21; +x_312 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_312, 0, x_256); +lean_ctor_set(x_312, 1, x_311); +lean_ctor_set(x_312, 2, x_310); +x_313 = l_Lean_Name_toString_maybePseudoSyntax___closed__2; +lean_inc(x_269); +x_314 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_314, 0, x_269); +lean_ctor_set(x_314, 1, x_313); +x_315 = lean_array_push(x_300, x_312); +lean_inc(x_314); +x_316 = lean_array_push(x_315, x_314); +x_317 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__20; +x_318 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_318, 0, x_256); +lean_ctor_set(x_318, 1, x_317); +lean_ctor_set(x_318, 2, x_316); +x_319 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__29; lean_inc(x_242); -x_319 = l_Lean_addMacroScope(x_242, x_318, x_241); -x_320 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__30; -x_321 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__34; -lean_inc(x_268); -x_322 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_322, 0, x_268); -lean_ctor_set(x_322, 1, x_320); -lean_ctor_set(x_322, 2, x_319); -lean_ctor_set(x_322, 3, x_321); -x_323 = lean_array_push(x_299, x_322); -x_324 = lean_array_push(x_323, x_301); -x_325 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_325, 0, x_255); -lean_ctor_set(x_325, 1, x_303); -lean_ctor_set(x_325, 2, x_324); -x_326 = lean_array_push(x_263, x_325); -x_327 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_327, 0, x_255); -lean_ctor_set(x_327, 1, x_289); -lean_ctor_set(x_327, 2, x_326); -lean_inc(x_307); -x_328 = lean_array_push(x_307, x_327); -lean_inc(x_279); -x_329 = lean_array_push(x_328, x_279); -x_330 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_330, 0, x_255); -lean_ctor_set(x_330, 1, x_310); -lean_ctor_set(x_330, 2, x_329); -x_331 = lean_array_push(x_299, x_330); -lean_inc(x_313); -x_332 = lean_array_push(x_331, x_313); -x_333 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_333, 0, x_255); -lean_ctor_set(x_333, 1, x_316); -lean_ctor_set(x_333, 2, x_332); -x_334 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__37; -lean_inc(x_268); -x_335 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_335, 0, x_268); -lean_ctor_set(x_335, 1, x_334); -x_336 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__38; -lean_inc(x_268); -x_337 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_337, 0, x_268); -lean_ctor_set(x_337, 1, x_336); -x_338 = lean_array_push(x_263, x_337); -x_339 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_339, 0, x_255); -lean_ctor_set(x_339, 1, x_265); -lean_ctor_set(x_339, 2, x_338); -x_340 = lean_array_push(x_299, x_335); -x_341 = lean_array_push(x_340, x_339); -x_342 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__36; -x_343 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_343, 0, x_255); -lean_ctor_set(x_343, 1, x_342); -lean_ctor_set(x_343, 2, x_341); -x_344 = lean_array_push(x_263, x_343); -x_345 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_345, 0, x_255); -lean_ctor_set(x_345, 1, x_289); -lean_ctor_set(x_345, 2, x_344); -lean_inc(x_307); -x_346 = lean_array_push(x_307, x_345); -lean_inc(x_279); -x_347 = lean_array_push(x_346, x_279); -x_348 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_348, 0, x_255); -lean_ctor_set(x_348, 1, x_310); -lean_ctor_set(x_348, 2, x_347); -x_349 = lean_array_push(x_299, x_348); -lean_inc(x_313); -x_350 = lean_array_push(x_349, x_313); -x_351 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_351, 0, x_255); -lean_ctor_set(x_351, 1, x_316); -lean_ctor_set(x_351, 2, x_350); -x_352 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__39; -lean_inc(x_268); -x_353 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_353, 0, x_268); -lean_ctor_set(x_353, 1, x_352); -x_354 = lean_array_push(x_263, x_353); -x_355 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_355, 0, x_255); -lean_ctor_set(x_355, 1, x_265); -lean_ctor_set(x_355, 2, x_354); -x_356 = lean_array_push(x_263, x_355); -x_357 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_357, 0, x_255); -lean_ctor_set(x_357, 1, x_292); -lean_ctor_set(x_357, 2, x_356); -x_358 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__47; -lean_inc(x_241); +lean_inc(x_243); +x_320 = l_Lean_addMacroScope(x_243, x_319, x_242); +x_321 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__28; +x_322 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__32; +lean_inc(x_269); +x_323 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_323, 0, x_269); +lean_ctor_set(x_323, 1, x_321); +lean_ctor_set(x_323, 2, x_320); +lean_ctor_set(x_323, 3, x_322); +x_324 = lean_array_push(x_300, x_323); +x_325 = lean_array_push(x_324, x_302); +x_326 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_326, 0, x_256); +lean_ctor_set(x_326, 1, x_304); +lean_ctor_set(x_326, 2, x_325); +x_327 = lean_array_push(x_264, x_326); +x_328 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_328, 0, x_256); +lean_ctor_set(x_328, 1, x_290); +lean_ctor_set(x_328, 2, x_327); +lean_inc(x_308); +x_329 = lean_array_push(x_308, x_328); +lean_inc(x_280); +x_330 = lean_array_push(x_329, x_280); +x_331 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_331, 0, x_256); +lean_ctor_set(x_331, 1, x_311); +lean_ctor_set(x_331, 2, x_330); +x_332 = lean_array_push(x_300, x_331); +lean_inc(x_314); +x_333 = lean_array_push(x_332, x_314); +x_334 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_334, 0, x_256); +lean_ctor_set(x_334, 1, x_317); +lean_ctor_set(x_334, 2, x_333); +x_335 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__35; +lean_inc(x_269); +x_336 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_336, 0, x_269); +lean_ctor_set(x_336, 1, x_335); +x_337 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__36; +lean_inc(x_269); +x_338 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_338, 0, x_269); +lean_ctor_set(x_338, 1, x_337); +x_339 = lean_array_push(x_264, x_338); +x_340 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_340, 0, x_256); +lean_ctor_set(x_340, 1, x_266); +lean_ctor_set(x_340, 2, x_339); +x_341 = lean_array_push(x_300, x_336); +x_342 = lean_array_push(x_341, x_340); +x_343 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__34; +x_344 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_344, 0, x_256); +lean_ctor_set(x_344, 1, x_343); +lean_ctor_set(x_344, 2, x_342); +x_345 = lean_array_push(x_264, x_344); +x_346 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_346, 0, x_256); +lean_ctor_set(x_346, 1, x_290); +lean_ctor_set(x_346, 2, x_345); +lean_inc(x_308); +x_347 = lean_array_push(x_308, x_346); +lean_inc(x_280); +x_348 = lean_array_push(x_347, x_280); +x_349 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_349, 0, x_256); +lean_ctor_set(x_349, 1, x_311); +lean_ctor_set(x_349, 2, x_348); +x_350 = lean_array_push(x_300, x_349); +lean_inc(x_314); +x_351 = lean_array_push(x_350, x_314); +x_352 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_352, 0, x_256); +lean_ctor_set(x_352, 1, x_317); +lean_ctor_set(x_352, 2, x_351); +x_353 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__37; +lean_inc(x_269); +x_354 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_354, 0, x_269); +lean_ctor_set(x_354, 1, x_353); +x_355 = lean_array_push(x_264, x_354); +x_356 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_356, 0, x_256); +lean_ctor_set(x_356, 1, x_266); +lean_ctor_set(x_356, 2, x_355); +x_357 = lean_array_push(x_264, x_356); +x_358 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_358, 0, x_256); +lean_ctor_set(x_358, 1, x_293); +lean_ctor_set(x_358, 2, x_357); +x_359 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__45; lean_inc(x_242); -x_359 = l_Lean_addMacroScope(x_242, x_358, x_241); -x_360 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__46; -x_361 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__50; -lean_inc(x_268); -x_362 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_362, 0, x_268); -lean_ctor_set(x_362, 1, x_360); -lean_ctor_set(x_362, 2, x_359); -lean_ctor_set(x_362, 3, x_361); -x_363 = lean_array_push(x_299, x_362); -x_364 = lean_array_push(x_363, x_301); -x_365 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_365, 0, x_255); -lean_ctor_set(x_365, 1, x_303); -lean_ctor_set(x_365, 2, x_364); -x_366 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__51; -lean_inc(x_268); -x_367 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_367, 0, x_268); -lean_ctor_set(x_367, 1, x_366); -x_368 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__55; -lean_inc(x_241); +lean_inc(x_243); +x_360 = l_Lean_addMacroScope(x_243, x_359, x_242); +x_361 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__44; +x_362 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__48; +lean_inc(x_269); +x_363 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_363, 0, x_269); +lean_ctor_set(x_363, 1, x_361); +lean_ctor_set(x_363, 2, x_360); +lean_ctor_set(x_363, 3, x_362); +x_364 = lean_array_push(x_300, x_363); +x_365 = lean_array_push(x_364, x_302); +x_366 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_366, 0, x_256); +lean_ctor_set(x_366, 1, x_304); +lean_ctor_set(x_366, 2, x_365); +x_367 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__49; +lean_inc(x_269); +x_368 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_368, 0, x_269); +lean_ctor_set(x_368, 1, x_367); +x_369 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__53; lean_inc(x_242); -x_369 = l_Lean_addMacroScope(x_242, x_368, x_241); -x_370 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__54; -x_371 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__58; -lean_inc(x_268); -x_372 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_372, 0, x_268); -lean_ctor_set(x_372, 1, x_370); -lean_ctor_set(x_372, 2, x_369); -lean_ctor_set(x_372, 3, x_371); -x_373 = lean_array_push(x_299, x_372); -x_374 = lean_array_push(x_373, x_301); -x_375 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_375, 0, x_255); -lean_ctor_set(x_375, 1, x_303); -lean_ctor_set(x_375, 2, x_374); -x_376 = lean_array_push(x_251, x_365); -x_377 = lean_array_push(x_376, x_367); -x_378 = lean_array_push(x_377, x_375); -x_379 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__43; -x_380 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_380, 0, x_255); -lean_ctor_set(x_380, 1, x_379); -lean_ctor_set(x_380, 2, x_378); -x_381 = lean_array_push(x_263, x_380); -x_382 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_382, 0, x_255); -lean_ctor_set(x_382, 1, x_289); -lean_ctor_set(x_382, 2, x_381); -lean_inc(x_307); -x_383 = lean_array_push(x_307, x_382); -lean_inc(x_279); -x_384 = lean_array_push(x_383, x_279); -x_385 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_385, 0, x_255); -lean_ctor_set(x_385, 1, x_310); -lean_ctor_set(x_385, 2, x_384); -x_386 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__59; -lean_inc(x_268); -x_387 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_387, 0, x_268); -lean_ctor_set(x_387, 1, x_386); -x_388 = lean_array_push(x_299, x_385); -x_389 = lean_array_push(x_388, x_387); -x_390 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__41; -x_391 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_391, 0, x_255); -lean_ctor_set(x_391, 1, x_390); -lean_ctor_set(x_391, 2, x_389); -x_392 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__60; -lean_inc(x_268); -x_393 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_393, 0, x_268); -lean_ctor_set(x_393, 1, x_392); -x_394 = lean_array_push(x_263, x_393); -x_395 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_395, 0, x_255); -lean_ctor_set(x_395, 1, x_265); -lean_ctor_set(x_395, 2, x_394); -x_396 = lean_array_push(x_263, x_395); -x_397 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_397, 0, x_255); -lean_ctor_set(x_397, 1, x_292); -lean_ctor_set(x_397, 2, x_396); -x_398 = lean_array_push(x_251, x_357); -x_399 = lean_array_push(x_398, x_391); -x_400 = lean_array_push(x_399, x_397); -x_401 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_401, 0, x_255); -lean_ctor_set(x_401, 1, x_289); -lean_ctor_set(x_401, 2, x_400); -x_402 = lean_array_push(x_307, x_401); -x_403 = lean_array_push(x_402, x_279); -x_404 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_404, 0, x_255); -lean_ctor_set(x_404, 1, x_310); -lean_ctor_set(x_404, 2, x_403); -x_405 = lean_array_push(x_299, x_404); -x_406 = lean_array_push(x_405, x_313); -x_407 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_407, 0, x_255); -lean_ctor_set(x_407, 1, x_316); -lean_ctor_set(x_407, 2, x_406); -x_408 = lean_array_push(x_280, x_293); -x_409 = lean_array_push(x_408, x_317); -x_410 = lean_array_push(x_409, x_333); -x_411 = lean_array_push(x_410, x_351); -x_412 = lean_array_push(x_411, x_407); -x_413 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_413, 0, x_255); -lean_ctor_set(x_413, 1, x_289); -lean_ctor_set(x_413, 2, x_412); -x_414 = l_Lean_toolchain___closed__1; -lean_inc(x_268); -x_415 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_415, 0, x_268); -lean_ctor_set(x_415, 1, x_414); -x_416 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__71; -x_417 = l_Lean_addMacroScope(x_242, x_416, x_241); -x_418 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__70; -x_419 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_419, 0, x_268); -lean_ctor_set(x_419, 1, x_418); -lean_ctor_set(x_419, 2, x_417); -lean_ctor_set(x_419, 3, x_247); -x_420 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__74; -x_421 = lean_array_push(x_420, x_271); -x_422 = lean_array_push(x_421, x_301); -x_423 = lean_array_push(x_422, x_290); -x_424 = lean_array_push(x_423, x_301); -x_425 = lean_array_push(x_424, x_413); -x_426 = lean_array_push(x_425, x_415); -x_427 = lean_array_push(x_426, x_419); -x_428 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__12; -x_429 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_429, 0, x_255); -lean_ctor_set(x_429, 1, x_428); -lean_ctor_set(x_429, 2, x_427); -x_430 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_430, 0, x_266); +lean_inc(x_243); +x_370 = l_Lean_addMacroScope(x_243, x_369, x_242); +x_371 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__52; +x_372 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__56; +lean_inc(x_269); +x_373 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_373, 0, x_269); +lean_ctor_set(x_373, 1, x_371); +lean_ctor_set(x_373, 2, x_370); +lean_ctor_set(x_373, 3, x_372); +x_374 = lean_array_push(x_300, x_373); +x_375 = lean_array_push(x_374, x_302); +x_376 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_376, 0, x_256); +lean_ctor_set(x_376, 1, x_304); +lean_ctor_set(x_376, 2, x_375); +x_377 = lean_array_push(x_252, x_366); +x_378 = lean_array_push(x_377, x_368); +x_379 = lean_array_push(x_378, x_376); +x_380 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__41; +x_381 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_381, 0, x_256); +lean_ctor_set(x_381, 1, x_380); +lean_ctor_set(x_381, 2, x_379); +x_382 = lean_array_push(x_264, x_381); +x_383 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_383, 0, x_256); +lean_ctor_set(x_383, 1, x_290); +lean_ctor_set(x_383, 2, x_382); +lean_inc(x_308); +x_384 = lean_array_push(x_308, x_383); +lean_inc(x_280); +x_385 = lean_array_push(x_384, x_280); +x_386 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_386, 0, x_256); +lean_ctor_set(x_386, 1, x_311); +lean_ctor_set(x_386, 2, x_385); +x_387 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__57; +lean_inc(x_269); +x_388 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_388, 0, x_269); +lean_ctor_set(x_388, 1, x_387); +x_389 = lean_array_push(x_300, x_386); +x_390 = lean_array_push(x_389, x_388); +x_391 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__39; +x_392 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_392, 0, x_256); +lean_ctor_set(x_392, 1, x_391); +lean_ctor_set(x_392, 2, x_390); +x_393 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__58; +lean_inc(x_269); +x_394 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_394, 0, x_269); +lean_ctor_set(x_394, 1, x_393); +x_395 = lean_array_push(x_264, x_394); +x_396 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_396, 0, x_256); +lean_ctor_set(x_396, 1, x_266); +lean_ctor_set(x_396, 2, x_395); +x_397 = lean_array_push(x_264, x_396); +x_398 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_398, 0, x_256); +lean_ctor_set(x_398, 1, x_293); +lean_ctor_set(x_398, 2, x_397); +x_399 = lean_array_push(x_252, x_358); +x_400 = lean_array_push(x_399, x_392); +x_401 = lean_array_push(x_400, x_398); +x_402 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_402, 0, x_256); +lean_ctor_set(x_402, 1, x_290); +lean_ctor_set(x_402, 2, x_401); +x_403 = lean_array_push(x_308, x_402); +x_404 = lean_array_push(x_403, x_280); +x_405 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_405, 0, x_256); +lean_ctor_set(x_405, 1, x_311); +lean_ctor_set(x_405, 2, x_404); +x_406 = lean_array_push(x_300, x_405); +x_407 = lean_array_push(x_406, x_314); +x_408 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_408, 0, x_256); +lean_ctor_set(x_408, 1, x_317); +lean_ctor_set(x_408, 2, x_407); +x_409 = lean_array_push(x_281, x_294); +x_410 = lean_array_push(x_409, x_318); +x_411 = lean_array_push(x_410, x_334); +x_412 = lean_array_push(x_411, x_352); +x_413 = lean_array_push(x_412, x_408); +x_414 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_414, 0, x_256); +lean_ctor_set(x_414, 1, x_290); +lean_ctor_set(x_414, 2, x_413); +x_415 = l_Lean_toolchain___closed__1; +lean_inc(x_269); +x_416 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_416, 0, x_269); +lean_ctor_set(x_416, 1, x_415); +x_417 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__69; +x_418 = l_Lean_addMacroScope(x_243, x_417, x_242); +x_419 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__68; +x_420 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_420, 0, x_269); +lean_ctor_set(x_420, 1, x_419); +lean_ctor_set(x_420, 2, x_418); +lean_ctor_set(x_420, 3, x_248); +x_421 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__72; +x_422 = lean_array_push(x_421, x_272); +x_423 = lean_array_push(x_422, x_302); +x_424 = lean_array_push(x_423, x_291); +x_425 = lean_array_push(x_424, x_302); +x_426 = lean_array_push(x_425, x_414); +x_427 = lean_array_push(x_426, x_416); +x_428 = lean_array_push(x_427, x_420); +x_429 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__11; +x_430 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_430, 0, x_256); lean_ctor_set(x_430, 1, x_429); +lean_ctor_set(x_430, 2, x_428); x_431 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_431, 0, x_257); +lean_ctor_set(x_431, 0, x_267); lean_ctor_set(x_431, 1, x_430); -x_432 = l_Lean_mkHole___closed__4; -x_433 = l_Lean_mkHole___closed__2; -x_434 = l_Lean_Parser_Tactic_tacticErw_______closed__2; -x_435 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1(x_432, x_11, x_433, x_434, x_15, x_431, x_2, x_269); -return x_435; +x_432 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_432, 0, x_258); +lean_ctor_set(x_432, 1, x_431); +x_433 = l_Lean_mkHole___closed__4; +x_434 = l_Lean_mkHole___closed__2; +x_435 = l_Lean_Parser_Tactic_tacticErw_______closed__2; +x_436 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__17; +x_437 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1(x_433, x_11, x_434, x_435, x_436, x_15, x_432, x_2, x_270); +return x_437; } } else { -lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; lean_object* x_565; lean_object* x_566; lean_object* x_567; lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; +lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; lean_object* x_565; lean_object* x_566; lean_object* x_567; lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; lean_dec(x_9); lean_inc(x_2); -x_436 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_437 = lean_ctor_get(x_436, 0); -lean_inc(x_437); -x_438 = lean_ctor_get(x_436, 1); -lean_inc(x_438); -lean_dec(x_436); -x_439 = lean_ctor_get(x_2, 2); +x_438 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_439 = lean_ctor_get(x_438, 0); lean_inc(x_439); -x_440 = lean_ctor_get(x_2, 1); +x_440 = lean_ctor_get(x_438, 1); lean_inc(x_440); -x_441 = l_Lean_Name_reprPrec___closed__3; -lean_inc(x_437); -x_442 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_442, 0, x_437); -lean_ctor_set(x_442, 1, x_441); -x_443 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__86; +lean_dec(x_438); +x_441 = lean_ctor_get(x_2, 2); +lean_inc(x_441); +x_442 = lean_ctor_get(x_2, 1); +lean_inc(x_442); +x_443 = l_Lean_Name_reprPrec___closed__3; lean_inc(x_439); -lean_inc(x_440); -x_444 = l_Lean_addMacroScope(x_440, x_443, x_439); -x_445 = lean_box(0); -x_446 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__85; -x_447 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__89; -x_448 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_448, 0, x_437); -lean_ctor_set(x_448, 1, x_446); -lean_ctor_set(x_448, 2, x_444); -lean_ctor_set(x_448, 3, x_447); -x_449 = l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__4; +x_444 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_444, 0, x_439); +lean_ctor_set(x_444, 1, x_443); +x_445 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__84; +lean_inc(x_441); lean_inc(x_442); -x_450 = lean_array_push(x_449, x_442); -x_451 = lean_array_push(x_450, x_442); -x_452 = lean_array_push(x_451, x_448); -x_453 = lean_box(2); -x_454 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__2; -x_455 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_455, 0, x_453); -lean_ctor_set(x_455, 1, x_454); -lean_ctor_set(x_455, 2, x_452); +x_446 = l_Lean_addMacroScope(x_442, x_445, x_441); +x_447 = lean_box(0); +x_448 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__83; +x_449 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__87; +x_450 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_450, 0, x_439); +lean_ctor_set(x_450, 1, x_448); +lean_ctor_set(x_450, 2, x_446); +lean_ctor_set(x_450, 3, x_449); +x_451 = l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__4; +lean_inc(x_444); +x_452 = lean_array_push(x_451, x_444); +x_453 = lean_array_push(x_452, x_444); +x_454 = lean_array_push(x_453, x_450); +x_455 = lean_box(2); +x_456 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__2; +x_457 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_457, 0, x_455); +lean_ctor_set(x_457, 1, x_456); +lean_ctor_set(x_457, 2, x_454); lean_inc(x_2); -x_456 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_438); -x_457 = lean_ctor_get(x_456, 0); -lean_inc(x_457); -x_458 = lean_ctor_get(x_456, 1); -lean_inc(x_458); -lean_dec(x_456); -x_459 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__90; -x_460 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_460, 0, x_457); -lean_ctor_set(x_460, 1, x_459); -x_461 = l_Lean_mkOptionalNode___closed__2; -x_462 = lean_array_push(x_461, x_460); -x_463 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__15; -x_464 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_464, 0, x_453); -lean_ctor_set(x_464, 1, x_463); -lean_ctor_set(x_464, 2, x_462); +x_458 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_440); +x_459 = lean_ctor_get(x_458, 0); +lean_inc(x_459); +x_460 = lean_ctor_get(x_458, 1); +lean_inc(x_460); +lean_dec(x_458); +x_461 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__88; +x_462 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_462, 0, x_459); +lean_ctor_set(x_462, 1, x_461); +x_463 = l_Lean_mkOptionalNode___closed__2; +x_464 = lean_array_push(x_463, x_462); +x_465 = l_Lean_Syntax_mkStrLit___closed__2; +x_466 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_466, 0, x_455); +lean_ctor_set(x_466, 1, x_465); +lean_ctor_set(x_466, 2, x_464); lean_inc(x_2); -x_465 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_458); -x_466 = lean_ctor_get(x_465, 0); -lean_inc(x_466); -x_467 = lean_ctor_get(x_465, 1); -lean_inc(x_467); -lean_dec(x_465); -x_468 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__11; -lean_inc(x_466); -x_469 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_469, 0, x_466); -lean_ctor_set(x_469, 1, x_468); -x_470 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__4; -lean_inc(x_466); +x_467 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_460); +x_468 = lean_ctor_get(x_467, 0); +lean_inc(x_468); +x_469 = lean_ctor_get(x_467, 1); +lean_inc(x_469); +lean_dec(x_467); +x_470 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__10; +lean_inc(x_468); x_471 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_471, 0, x_466); +lean_ctor_set(x_471, 0, x_468); lean_ctor_set(x_471, 1, x_470); -x_472 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__18; -lean_inc(x_466); +x_472 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__4; +lean_inc(x_468); x_473 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_473, 0, x_466); +lean_ctor_set(x_473, 0, x_468); lean_ctor_set(x_473, 1, x_472); -x_474 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__8; -lean_inc(x_466); +x_474 = l_Lean_Syntax_mkNameLit___closed__1; +lean_inc(x_468); x_475 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_475, 0, x_466); +lean_ctor_set(x_475, 0, x_468); lean_ctor_set(x_475, 1, x_474); -x_476 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__8; -lean_inc(x_466); +x_476 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__8; +lean_inc(x_468); x_477 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_477, 0, x_466); +lean_ctor_set(x_477, 0, x_468); lean_ctor_set(x_477, 1, x_476); -x_478 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__37; -lean_inc(x_471); -x_479 = lean_array_push(x_478, x_471); -x_480 = lean_array_push(x_479, x_473); -x_481 = lean_array_push(x_480, x_475); -lean_inc(x_11); -x_482 = lean_array_push(x_481, x_11); -lean_inc(x_477); +x_478 = l_repr___at___private_Init_Meta_0__Lean_Name_reprSyntax____x40_Init_Meta___hyg_1079____spec__3___closed__8; +lean_inc(x_468); +x_479 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_479, 0, x_468); +lean_ctor_set(x_479, 1, x_478); +x_480 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__37; +lean_inc(x_473); +x_481 = lean_array_push(x_480, x_473); +x_482 = lean_array_push(x_481, x_475); x_483 = lean_array_push(x_482, x_477); -x_484 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__17; -x_485 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_485, 0, x_453); -lean_ctor_set(x_485, 1, x_484); -lean_ctor_set(x_485, 2, x_483); -x_486 = lean_array_push(x_461, x_485); -x_487 = l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__2; -x_488 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_488, 0, x_453); -lean_ctor_set(x_488, 1, x_487); -lean_ctor_set(x_488, 2, x_486); -x_489 = lean_array_push(x_461, x_13); -x_490 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__20; -x_491 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_491, 0, x_453); -lean_ctor_set(x_491, 1, x_490); -lean_ctor_set(x_491, 2, x_489); -x_492 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__57; -lean_inc(x_439); -lean_inc(x_440); -x_493 = l_Lean_addMacroScope(x_440, x_492, x_439); -x_494 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__56; -x_495 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__27; -lean_inc(x_466); -x_496 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_496, 0, x_466); -lean_ctor_set(x_496, 1, x_494); -lean_ctor_set(x_496, 2, x_493); -lean_ctor_set(x_496, 3, x_495); -x_497 = l_Lean_Syntax_mkApp___closed__3; -x_498 = lean_array_push(x_497, x_496); -x_499 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__12; -x_500 = lean_array_push(x_498, x_499); -x_501 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__25; -x_502 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_502, 0, x_453); -lean_ctor_set(x_502, 1, x_501); -lean_ctor_set(x_502, 2, x_500); -x_503 = lean_array_push(x_461, x_502); +lean_inc(x_11); +x_484 = lean_array_push(x_483, x_11); +lean_inc(x_479); +x_485 = lean_array_push(x_484, x_479); +x_486 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__16; +x_487 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_487, 0, x_455); +lean_ctor_set(x_487, 1, x_486); +lean_ctor_set(x_487, 2, x_485); +x_488 = lean_array_push(x_463, x_487); +x_489 = l_Lean_mkNullNode___closed__2; +x_490 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_490, 0, x_455); +lean_ctor_set(x_490, 1, x_489); +lean_ctor_set(x_490, 2, x_488); +x_491 = lean_array_push(x_463, x_13); +x_492 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__18; +x_493 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_493, 0, x_455); +lean_ctor_set(x_493, 1, x_492); +lean_ctor_set(x_493, 2, x_491); +x_494 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__55; +lean_inc(x_441); +lean_inc(x_442); +x_495 = l_Lean_addMacroScope(x_442, x_494, x_441); +x_496 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__54; +x_497 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__25; +lean_inc(x_468); +x_498 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_498, 0, x_468); +lean_ctor_set(x_498, 1, x_496); +lean_ctor_set(x_498, 2, x_495); +lean_ctor_set(x_498, 3, x_497); +x_499 = l_Lean_Syntax_mkApp___closed__3; +x_500 = lean_array_push(x_499, x_498); +x_501 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__12; +x_502 = lean_array_push(x_500, x_501); +x_503 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__23; x_504 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_504, 0, x_453); -lean_ctor_set(x_504, 1, x_487); -lean_ctor_set(x_504, 2, x_503); -x_505 = lean_array_push(x_449, x_471); -lean_inc(x_505); -x_506 = lean_array_push(x_505, x_504); -lean_inc(x_477); -x_507 = lean_array_push(x_506, x_477); -x_508 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__23; -x_509 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_509, 0, x_453); -lean_ctor_set(x_509, 1, x_508); -lean_ctor_set(x_509, 2, x_507); -x_510 = l_Lean_Name_toString_maybePseudoSyntax___closed__2; -lean_inc(x_466); -x_511 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_511, 0, x_466); +lean_ctor_set(x_504, 0, x_455); +lean_ctor_set(x_504, 1, x_503); +lean_ctor_set(x_504, 2, x_502); +x_505 = lean_array_push(x_463, x_504); +x_506 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_506, 0, x_455); +lean_ctor_set(x_506, 1, x_489); +lean_ctor_set(x_506, 2, x_505); +x_507 = lean_array_push(x_451, x_473); +lean_inc(x_507); +x_508 = lean_array_push(x_507, x_506); +lean_inc(x_479); +x_509 = lean_array_push(x_508, x_479); +x_510 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__21; +x_511 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_511, 0, x_455); lean_ctor_set(x_511, 1, x_510); -x_512 = lean_array_push(x_497, x_509); -lean_inc(x_511); -x_513 = lean_array_push(x_512, x_511); -x_514 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__22; -x_515 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_515, 0, x_453); -lean_ctor_set(x_515, 1, x_514); -lean_ctor_set(x_515, 2, x_513); -x_516 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__31; -lean_inc(x_439); -lean_inc(x_440); -x_517 = l_Lean_addMacroScope(x_440, x_516, x_439); -x_518 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__30; -x_519 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__34; -lean_inc(x_466); -x_520 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_520, 0, x_466); -lean_ctor_set(x_520, 1, x_518); -lean_ctor_set(x_520, 2, x_517); -lean_ctor_set(x_520, 3, x_519); -x_521 = lean_array_push(x_497, x_520); -x_522 = lean_array_push(x_521, x_499); -x_523 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_523, 0, x_453); -lean_ctor_set(x_523, 1, x_501); -lean_ctor_set(x_523, 2, x_522); -x_524 = lean_array_push(x_461, x_523); +lean_ctor_set(x_511, 2, x_509); +x_512 = l_Lean_Name_toString_maybePseudoSyntax___closed__2; +lean_inc(x_468); +x_513 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_513, 0, x_468); +lean_ctor_set(x_513, 1, x_512); +x_514 = lean_array_push(x_499, x_511); +lean_inc(x_513); +x_515 = lean_array_push(x_514, x_513); +x_516 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__20; +x_517 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_517, 0, x_455); +lean_ctor_set(x_517, 1, x_516); +lean_ctor_set(x_517, 2, x_515); +x_518 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__29; +lean_inc(x_441); +lean_inc(x_442); +x_519 = l_Lean_addMacroScope(x_442, x_518, x_441); +x_520 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__28; +x_521 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__32; +lean_inc(x_468); +x_522 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_522, 0, x_468); +lean_ctor_set(x_522, 1, x_520); +lean_ctor_set(x_522, 2, x_519); +lean_ctor_set(x_522, 3, x_521); +x_523 = lean_array_push(x_499, x_522); +x_524 = lean_array_push(x_523, x_501); x_525 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_525, 0, x_453); -lean_ctor_set(x_525, 1, x_487); +lean_ctor_set(x_525, 0, x_455); +lean_ctor_set(x_525, 1, x_503); lean_ctor_set(x_525, 2, x_524); -lean_inc(x_505); -x_526 = lean_array_push(x_505, x_525); -lean_inc(x_477); -x_527 = lean_array_push(x_526, x_477); -x_528 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_528, 0, x_453); -lean_ctor_set(x_528, 1, x_508); -lean_ctor_set(x_528, 2, x_527); -x_529 = lean_array_push(x_497, x_528); -lean_inc(x_511); -x_530 = lean_array_push(x_529, x_511); -x_531 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_531, 0, x_453); -lean_ctor_set(x_531, 1, x_514); -lean_ctor_set(x_531, 2, x_530); -x_532 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__37; -lean_inc(x_466); -x_533 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_533, 0, x_466); -lean_ctor_set(x_533, 1, x_532); -x_534 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__38; -lean_inc(x_466); +x_526 = lean_array_push(x_463, x_525); +x_527 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_527, 0, x_455); +lean_ctor_set(x_527, 1, x_489); +lean_ctor_set(x_527, 2, x_526); +lean_inc(x_507); +x_528 = lean_array_push(x_507, x_527); +lean_inc(x_479); +x_529 = lean_array_push(x_528, x_479); +x_530 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_530, 0, x_455); +lean_ctor_set(x_530, 1, x_510); +lean_ctor_set(x_530, 2, x_529); +x_531 = lean_array_push(x_499, x_530); +lean_inc(x_513); +x_532 = lean_array_push(x_531, x_513); +x_533 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_533, 0, x_455); +lean_ctor_set(x_533, 1, x_516); +lean_ctor_set(x_533, 2, x_532); +x_534 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__35; +lean_inc(x_468); x_535 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_535, 0, x_466); +lean_ctor_set(x_535, 0, x_468); lean_ctor_set(x_535, 1, x_534); -x_536 = lean_array_push(x_461, x_535); -x_537 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_537, 0, x_453); -lean_ctor_set(x_537, 1, x_463); -lean_ctor_set(x_537, 2, x_536); -x_538 = lean_array_push(x_497, x_533); -x_539 = lean_array_push(x_538, x_537); -x_540 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__36; -x_541 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_541, 0, x_453); -lean_ctor_set(x_541, 1, x_540); -lean_ctor_set(x_541, 2, x_539); -x_542 = lean_array_push(x_461, x_541); +x_536 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__36; +lean_inc(x_468); +x_537 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_537, 0, x_468); +lean_ctor_set(x_537, 1, x_536); +x_538 = lean_array_push(x_463, x_537); +x_539 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_539, 0, x_455); +lean_ctor_set(x_539, 1, x_465); +lean_ctor_set(x_539, 2, x_538); +x_540 = lean_array_push(x_499, x_535); +x_541 = lean_array_push(x_540, x_539); +x_542 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__34; x_543 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_543, 0, x_453); -lean_ctor_set(x_543, 1, x_487); -lean_ctor_set(x_543, 2, x_542); -lean_inc(x_505); -x_544 = lean_array_push(x_505, x_543); -lean_inc(x_477); -x_545 = lean_array_push(x_544, x_477); -x_546 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_546, 0, x_453); -lean_ctor_set(x_546, 1, x_508); -lean_ctor_set(x_546, 2, x_545); -x_547 = lean_array_push(x_497, x_546); -lean_inc(x_511); -x_548 = lean_array_push(x_547, x_511); -x_549 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_549, 0, x_453); -lean_ctor_set(x_549, 1, x_514); -lean_ctor_set(x_549, 2, x_548); -x_550 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__39; -lean_inc(x_466); -x_551 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_551, 0, x_466); -lean_ctor_set(x_551, 1, x_550); -x_552 = lean_array_push(x_461, x_551); -x_553 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_553, 0, x_453); -lean_ctor_set(x_553, 1, x_463); -lean_ctor_set(x_553, 2, x_552); -x_554 = lean_array_push(x_461, x_553); +lean_ctor_set(x_543, 0, x_455); +lean_ctor_set(x_543, 1, x_542); +lean_ctor_set(x_543, 2, x_541); +x_544 = lean_array_push(x_463, x_543); +x_545 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_545, 0, x_455); +lean_ctor_set(x_545, 1, x_489); +lean_ctor_set(x_545, 2, x_544); +lean_inc(x_507); +x_546 = lean_array_push(x_507, x_545); +lean_inc(x_479); +x_547 = lean_array_push(x_546, x_479); +x_548 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_548, 0, x_455); +lean_ctor_set(x_548, 1, x_510); +lean_ctor_set(x_548, 2, x_547); +x_549 = lean_array_push(x_499, x_548); +lean_inc(x_513); +x_550 = lean_array_push(x_549, x_513); +x_551 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_551, 0, x_455); +lean_ctor_set(x_551, 1, x_516); +lean_ctor_set(x_551, 2, x_550); +x_552 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__37; +lean_inc(x_468); +x_553 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_553, 0, x_468); +lean_ctor_set(x_553, 1, x_552); +x_554 = lean_array_push(x_463, x_553); x_555 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_555, 0, x_453); -lean_ctor_set(x_555, 1, x_490); +lean_ctor_set(x_555, 0, x_455); +lean_ctor_set(x_555, 1, x_465); lean_ctor_set(x_555, 2, x_554); -x_556 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__94; -lean_inc(x_439); -lean_inc(x_440); -x_557 = l_Lean_addMacroScope(x_440, x_556, x_439); -x_558 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__93; -x_559 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__97; -lean_inc(x_466); -x_560 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_560, 0, x_466); -lean_ctor_set(x_560, 1, x_558); -lean_ctor_set(x_560, 2, x_557); -lean_ctor_set(x_560, 3, x_559); -x_561 = lean_array_push(x_497, x_560); -x_562 = lean_array_push(x_561, x_499); -x_563 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_563, 0, x_453); -lean_ctor_set(x_563, 1, x_501); -lean_ctor_set(x_563, 2, x_562); -x_564 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__51; -lean_inc(x_466); -x_565 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_565, 0, x_466); -lean_ctor_set(x_565, 1, x_564); -x_566 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__47; -lean_inc(x_439); -lean_inc(x_440); -x_567 = l_Lean_addMacroScope(x_440, x_566, x_439); -x_568 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__46; -x_569 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__50; -lean_inc(x_466); -x_570 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_570, 0, x_466); -lean_ctor_set(x_570, 1, x_568); -lean_ctor_set(x_570, 2, x_567); -lean_ctor_set(x_570, 3, x_569); -x_571 = lean_array_push(x_497, x_570); -x_572 = lean_array_push(x_571, x_499); -x_573 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_573, 0, x_453); -lean_ctor_set(x_573, 1, x_501); -lean_ctor_set(x_573, 2, x_572); -x_574 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__55; -lean_inc(x_439); -lean_inc(x_440); -x_575 = l_Lean_addMacroScope(x_440, x_574, x_439); -x_576 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__54; -x_577 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__58; -lean_inc(x_466); -x_578 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_578, 0, x_466); -lean_ctor_set(x_578, 1, x_576); -lean_ctor_set(x_578, 2, x_575); -lean_ctor_set(x_578, 3, x_577); -x_579 = lean_array_push(x_497, x_578); -x_580 = lean_array_push(x_579, x_499); -x_581 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_581, 0, x_453); -lean_ctor_set(x_581, 1, x_501); -lean_ctor_set(x_581, 2, x_580); -x_582 = lean_array_push(x_449, x_573); -lean_inc(x_565); -x_583 = lean_array_push(x_582, x_565); -x_584 = lean_array_push(x_583, x_581); -x_585 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__43; -x_586 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_586, 0, x_453); -lean_ctor_set(x_586, 1, x_585); -lean_ctor_set(x_586, 2, x_584); -x_587 = lean_array_push(x_449, x_563); -x_588 = lean_array_push(x_587, x_565); -x_589 = lean_array_push(x_588, x_586); -x_590 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_590, 0, x_453); -lean_ctor_set(x_590, 1, x_585); -lean_ctor_set(x_590, 2, x_589); -x_591 = lean_array_push(x_461, x_590); +x_556 = lean_array_push(x_463, x_555); +x_557 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_557, 0, x_455); +lean_ctor_set(x_557, 1, x_492); +lean_ctor_set(x_557, 2, x_556); +x_558 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__92; +lean_inc(x_441); +lean_inc(x_442); +x_559 = l_Lean_addMacroScope(x_442, x_558, x_441); +x_560 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__91; +x_561 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__95; +lean_inc(x_468); +x_562 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_562, 0, x_468); +lean_ctor_set(x_562, 1, x_560); +lean_ctor_set(x_562, 2, x_559); +lean_ctor_set(x_562, 3, x_561); +x_563 = lean_array_push(x_499, x_562); +x_564 = lean_array_push(x_563, x_501); +x_565 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_565, 0, x_455); +lean_ctor_set(x_565, 1, x_503); +lean_ctor_set(x_565, 2, x_564); +x_566 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__49; +lean_inc(x_468); +x_567 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_567, 0, x_468); +lean_ctor_set(x_567, 1, x_566); +x_568 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__45; +lean_inc(x_441); +lean_inc(x_442); +x_569 = l_Lean_addMacroScope(x_442, x_568, x_441); +x_570 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__44; +x_571 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__48; +lean_inc(x_468); +x_572 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_572, 0, x_468); +lean_ctor_set(x_572, 1, x_570); +lean_ctor_set(x_572, 2, x_569); +lean_ctor_set(x_572, 3, x_571); +x_573 = lean_array_push(x_499, x_572); +x_574 = lean_array_push(x_573, x_501); +x_575 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_575, 0, x_455); +lean_ctor_set(x_575, 1, x_503); +lean_ctor_set(x_575, 2, x_574); +x_576 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__53; +lean_inc(x_441); +lean_inc(x_442); +x_577 = l_Lean_addMacroScope(x_442, x_576, x_441); +x_578 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__52; +x_579 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__56; +lean_inc(x_468); +x_580 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_580, 0, x_468); +lean_ctor_set(x_580, 1, x_578); +lean_ctor_set(x_580, 2, x_577); +lean_ctor_set(x_580, 3, x_579); +x_581 = lean_array_push(x_499, x_580); +x_582 = lean_array_push(x_581, x_501); +x_583 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_583, 0, x_455); +lean_ctor_set(x_583, 1, x_503); +lean_ctor_set(x_583, 2, x_582); +x_584 = lean_array_push(x_451, x_575); +lean_inc(x_567); +x_585 = lean_array_push(x_584, x_567); +x_586 = lean_array_push(x_585, x_583); +x_587 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__41; +x_588 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_588, 0, x_455); +lean_ctor_set(x_588, 1, x_587); +lean_ctor_set(x_588, 2, x_586); +x_589 = lean_array_push(x_451, x_565); +x_590 = lean_array_push(x_589, x_567); +x_591 = lean_array_push(x_590, x_588); x_592 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_592, 0, x_453); -lean_ctor_set(x_592, 1, x_487); +lean_ctor_set(x_592, 0, x_455); +lean_ctor_set(x_592, 1, x_587); lean_ctor_set(x_592, 2, x_591); -lean_inc(x_505); -x_593 = lean_array_push(x_505, x_592); -lean_inc(x_477); -x_594 = lean_array_push(x_593, x_477); -x_595 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_595, 0, x_453); -lean_ctor_set(x_595, 1, x_508); -lean_ctor_set(x_595, 2, x_594); -x_596 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__59; -lean_inc(x_466); -x_597 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_597, 0, x_466); -lean_ctor_set(x_597, 1, x_596); -x_598 = lean_array_push(x_497, x_595); -x_599 = lean_array_push(x_598, x_597); -x_600 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__41; -x_601 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_601, 0, x_453); -lean_ctor_set(x_601, 1, x_600); -lean_ctor_set(x_601, 2, x_599); -x_602 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__60; -lean_inc(x_466); -x_603 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_603, 0, x_466); +x_593 = lean_array_push(x_463, x_592); +x_594 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_594, 0, x_455); +lean_ctor_set(x_594, 1, x_489); +lean_ctor_set(x_594, 2, x_593); +lean_inc(x_507); +x_595 = lean_array_push(x_507, x_594); +lean_inc(x_479); +x_596 = lean_array_push(x_595, x_479); +x_597 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_597, 0, x_455); +lean_ctor_set(x_597, 1, x_510); +lean_ctor_set(x_597, 2, x_596); +x_598 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__57; +lean_inc(x_468); +x_599 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_599, 0, x_468); +lean_ctor_set(x_599, 1, x_598); +x_600 = lean_array_push(x_499, x_597); +x_601 = lean_array_push(x_600, x_599); +x_602 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__39; +x_603 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_603, 0, x_455); lean_ctor_set(x_603, 1, x_602); -x_604 = lean_array_push(x_461, x_603); -x_605 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_605, 0, x_453); -lean_ctor_set(x_605, 1, x_463); -lean_ctor_set(x_605, 2, x_604); -x_606 = lean_array_push(x_461, x_605); +lean_ctor_set(x_603, 2, x_601); +x_604 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__58; +lean_inc(x_468); +x_605 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_605, 0, x_468); +lean_ctor_set(x_605, 1, x_604); +x_606 = lean_array_push(x_463, x_605); x_607 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_607, 0, x_453); -lean_ctor_set(x_607, 1, x_490); +lean_ctor_set(x_607, 0, x_455); +lean_ctor_set(x_607, 1, x_465); lean_ctor_set(x_607, 2, x_606); -x_608 = lean_array_push(x_449, x_555); -x_609 = lean_array_push(x_608, x_601); -x_610 = lean_array_push(x_609, x_607); -x_611 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_611, 0, x_453); -lean_ctor_set(x_611, 1, x_487); -lean_ctor_set(x_611, 2, x_610); -lean_inc(x_505); -x_612 = lean_array_push(x_505, x_611); -lean_inc(x_477); -x_613 = lean_array_push(x_612, x_477); -x_614 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_614, 0, x_453); -lean_ctor_set(x_614, 1, x_508); -lean_ctor_set(x_614, 2, x_613); -x_615 = lean_array_push(x_497, x_614); -lean_inc(x_511); -x_616 = lean_array_push(x_615, x_511); -x_617 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_617, 0, x_453); -lean_ctor_set(x_617, 1, x_514); -lean_ctor_set(x_617, 2, x_616); -x_618 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__64; -lean_inc(x_439); -lean_inc(x_440); -x_619 = l_Lean_addMacroScope(x_440, x_618, x_439); -x_620 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__63; -x_621 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__67; -lean_inc(x_466); -x_622 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_622, 0, x_466); -lean_ctor_set(x_622, 1, x_620); -lean_ctor_set(x_622, 2, x_619); -lean_ctor_set(x_622, 3, x_621); -x_623 = lean_array_push(x_497, x_622); -x_624 = lean_array_push(x_623, x_499); -x_625 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_625, 0, x_453); -lean_ctor_set(x_625, 1, x_501); -lean_ctor_set(x_625, 2, x_624); -x_626 = lean_array_push(x_461, x_625); +x_608 = lean_array_push(x_463, x_607); +x_609 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_609, 0, x_455); +lean_ctor_set(x_609, 1, x_492); +lean_ctor_set(x_609, 2, x_608); +x_610 = lean_array_push(x_451, x_557); +x_611 = lean_array_push(x_610, x_603); +x_612 = lean_array_push(x_611, x_609); +x_613 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_613, 0, x_455); +lean_ctor_set(x_613, 1, x_489); +lean_ctor_set(x_613, 2, x_612); +lean_inc(x_507); +x_614 = lean_array_push(x_507, x_613); +lean_inc(x_479); +x_615 = lean_array_push(x_614, x_479); +x_616 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_616, 0, x_455); +lean_ctor_set(x_616, 1, x_510); +lean_ctor_set(x_616, 2, x_615); +x_617 = lean_array_push(x_499, x_616); +lean_inc(x_513); +x_618 = lean_array_push(x_617, x_513); +x_619 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_619, 0, x_455); +lean_ctor_set(x_619, 1, x_516); +lean_ctor_set(x_619, 2, x_618); +x_620 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__62; +lean_inc(x_441); +lean_inc(x_442); +x_621 = l_Lean_addMacroScope(x_442, x_620, x_441); +x_622 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__61; +x_623 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__65; +lean_inc(x_468); +x_624 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_624, 0, x_468); +lean_ctor_set(x_624, 1, x_622); +lean_ctor_set(x_624, 2, x_621); +lean_ctor_set(x_624, 3, x_623); +x_625 = lean_array_push(x_499, x_624); +x_626 = lean_array_push(x_625, x_501); x_627 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_627, 0, x_453); -lean_ctor_set(x_627, 1, x_487); +lean_ctor_set(x_627, 0, x_455); +lean_ctor_set(x_627, 1, x_503); lean_ctor_set(x_627, 2, x_626); -x_628 = lean_array_push(x_505, x_627); -x_629 = lean_array_push(x_628, x_477); -x_630 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_630, 0, x_453); -lean_ctor_set(x_630, 1, x_508); -lean_ctor_set(x_630, 2, x_629); -x_631 = lean_array_push(x_497, x_630); -x_632 = lean_array_push(x_631, x_511); -x_633 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_633, 0, x_453); -lean_ctor_set(x_633, 1, x_514); -lean_ctor_set(x_633, 2, x_632); -x_634 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__36; -x_635 = lean_array_push(x_634, x_491); -x_636 = lean_array_push(x_635, x_515); -x_637 = lean_array_push(x_636, x_531); -x_638 = lean_array_push(x_637, x_549); -x_639 = lean_array_push(x_638, x_617); -x_640 = lean_array_push(x_639, x_633); -x_641 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_641, 0, x_453); -lean_ctor_set(x_641, 1, x_487); -lean_ctor_set(x_641, 2, x_640); -x_642 = l_Lean_toolchain___closed__1; -lean_inc(x_466); -x_643 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_643, 0, x_466); -lean_ctor_set(x_643, 1, x_642); -x_644 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__71; -x_645 = l_Lean_addMacroScope(x_440, x_644, x_439); -x_646 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__70; -x_647 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_647, 0, x_466); -lean_ctor_set(x_647, 1, x_646); -lean_ctor_set(x_647, 2, x_645); -lean_ctor_set(x_647, 3, x_445); -x_648 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__74; -x_649 = lean_array_push(x_648, x_469); -x_650 = lean_array_push(x_649, x_499); -x_651 = lean_array_push(x_650, x_488); -x_652 = lean_array_push(x_651, x_499); -x_653 = lean_array_push(x_652, x_641); -x_654 = lean_array_push(x_653, x_643); -x_655 = lean_array_push(x_654, x_647); -x_656 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__12; -x_657 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_657, 0, x_453); -lean_ctor_set(x_657, 1, x_656); -lean_ctor_set(x_657, 2, x_655); -x_658 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_658, 0, x_464); -lean_ctor_set(x_658, 1, x_657); -x_659 = lean_alloc_ctor(0, 2, 0); +x_628 = lean_array_push(x_463, x_627); +x_629 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_629, 0, x_455); +lean_ctor_set(x_629, 1, x_489); +lean_ctor_set(x_629, 2, x_628); +x_630 = lean_array_push(x_507, x_629); +x_631 = lean_array_push(x_630, x_479); +x_632 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_632, 0, x_455); +lean_ctor_set(x_632, 1, x_510); +lean_ctor_set(x_632, 2, x_631); +x_633 = lean_array_push(x_499, x_632); +x_634 = lean_array_push(x_633, x_513); +x_635 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_635, 0, x_455); +lean_ctor_set(x_635, 1, x_516); +lean_ctor_set(x_635, 2, x_634); +x_636 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__36; +x_637 = lean_array_push(x_636, x_493); +x_638 = lean_array_push(x_637, x_517); +x_639 = lean_array_push(x_638, x_533); +x_640 = lean_array_push(x_639, x_551); +x_641 = lean_array_push(x_640, x_619); +x_642 = lean_array_push(x_641, x_635); +x_643 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_643, 0, x_455); +lean_ctor_set(x_643, 1, x_489); +lean_ctor_set(x_643, 2, x_642); +x_644 = l_Lean_toolchain___closed__1; +lean_inc(x_468); +x_645 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_645, 0, x_468); +lean_ctor_set(x_645, 1, x_644); +x_646 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__69; +x_647 = l_Lean_addMacroScope(x_442, x_646, x_441); +x_648 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__68; +x_649 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_649, 0, x_468); +lean_ctor_set(x_649, 1, x_648); +lean_ctor_set(x_649, 2, x_647); +lean_ctor_set(x_649, 3, x_447); +x_650 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__72; +x_651 = lean_array_push(x_650, x_471); +x_652 = lean_array_push(x_651, x_501); +x_653 = lean_array_push(x_652, x_490); +x_654 = lean_array_push(x_653, x_501); +x_655 = lean_array_push(x_654, x_643); +x_656 = lean_array_push(x_655, x_645); +x_657 = lean_array_push(x_656, x_649); +x_658 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__11; +x_659 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_659, 0, x_455); lean_ctor_set(x_659, 1, x_658); -x_660 = l_Lean_mkHole___closed__4; -x_661 = l_Lean_mkHole___closed__2; -x_662 = l_Lean_Parser_Tactic_tacticErw_______closed__2; -x_663 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1(x_660, x_11, x_661, x_662, x_15, x_659, x_2, x_467); -return x_663; +lean_ctor_set(x_659, 2, x_657); +x_660 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_660, 0, x_466); +lean_ctor_set(x_660, 1, x_659); +x_661 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_661, 0, x_457); +lean_ctor_set(x_661, 1, x_660); +x_662 = l_Lean_mkHole___closed__4; +x_663 = l_Lean_mkHole___closed__2; +x_664 = l_Lean_Parser_Tactic_tacticErw_______closed__2; +x_665 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__17; +x_666 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1(x_662, x_11, x_663, x_664, x_665, x_15, x_661, x_2, x_469); +return x_666; } } } @@ -29126,7 +29762,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_simpAutoUnfold___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__6; +x_1 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__6; x_2 = l_Lean_Parser_Tactic_simpErase; x_3 = l_Lean_Parser_Tactic_simpLemma; x_4 = lean_alloc_ctor(2, 3, 0); @@ -29140,7 +29776,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_simpAutoUnfold___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__6; +x_1 = l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__6; x_2 = l_Lean_Parser_Tactic_simpStar; x_3 = l_Lean_Parser_Tactic_simpAutoUnfold___closed__14; x_4 = lean_alloc_ctor(2, 3, 0); @@ -29284,7 +29920,7 @@ x_1 = l_Lean_Parser_Tactic_simpAutoUnfold___closed__25; return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____lambda__1___closed__1() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____lambda__1___closed__1() { _start: { lean_object* x_1; @@ -29292,21 +29928,21 @@ x_1 = lean_mk_string_from_bytes("simp ", 5); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____lambda__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; 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_5 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__87; +x_5 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__85; x_6 = l_Lean_Syntax_setKind(x_1, x_5); x_7 = lean_unsigned_to_nat(0u); x_8 = l_Lean_Syntax_getArg(x_6, x_7); -x_9 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____lambda__1___closed__1; +x_9 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____lambda__1___closed__1; x_10 = l_Lean_mkAtomFrom(x_8, x_9); x_11 = l_Lean_Syntax_setArg(x_6, x_7, x_10); x_12 = l_Lean_mkOptionalNode___closed__2; x_13 = lean_array_push(x_12, x_2); x_14 = lean_box(2); -x_15 = l_Lean_nullKind; +x_15 = l_Lean_mkNullNode___closed__2; x_16 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_16, 0, x_14); lean_ctor_set(x_16, 1, x_15); @@ -29319,27 +29955,27 @@ lean_ctor_set(x_19, 1, x_4); return x_19; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__1() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_mkHole___closed__6; -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__25; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__23; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__2() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_mkHole___closed__6; -x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__26; +x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__24; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__3() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__3() { _start: { lean_object* x_1; @@ -29347,22 +29983,22 @@ x_1 = lean_mk_string_from_bytes("Lean.Meta.Simp.Config", 21); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__4() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__3; +x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__3; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__3; +x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__3; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__4; +x_3 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__4; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -29370,7 +30006,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__6() { _start: { lean_object* x_1; @@ -29378,17 +30014,17 @@ x_1 = lean_mk_string_from_bytes("Simp", 4); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__7() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__24; -x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__6; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__6; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__8() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__8() { _start: { lean_object* x_1; @@ -29396,56 +30032,56 @@ x_1 = lean_mk_string_from_bytes("Config", 6); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__9() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__7; -x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__8; +x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__7; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__8; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__10() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__9; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__9; 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_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__11() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__10; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__10; x_3 = lean_alloc_ctor(1, 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_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__12() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__19; +x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__19; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__13() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__19; +x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__19; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__12; +x_3 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__12; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -29453,32 +30089,32 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__14() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__19; +x_2 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__19; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__15() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__15() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_instQuoteBool___closed__6; +x_1 = l_Lean_instQuoteBoolMkStrAnonymous___closed__6; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__16() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_instQuoteBool___closed__6; +x_1 = l_Lean_instQuoteBoolMkStrAnonymous___closed__6; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__15; +x_3 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__15; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -29486,47 +30122,47 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__17() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_instQuoteBool___closed__6; +x_2 = l_Lean_instQuoteBoolMkStrAnonymous___closed__6; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__18() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_instQuoteBool___closed__7; +x_2 = l_Lean_instQuoteBoolMkStrAnonymous___closed__7; 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_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__19() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__18; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__18; x_3 = lean_alloc_ctor(1, 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_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__20() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__32; -x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__14; +x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__13; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -29534,7 +30170,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; @@ -29576,17 +30212,17 @@ lean_inc(x_11); x_20 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_20, 0, x_11); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__25; +x_21 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__23; lean_inc(x_11); x_22 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_22, 0, x_11); lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__41; +x_23 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__39; lean_inc(x_13); lean_inc(x_14); x_24 = l_Lean_addMacroScope(x_14, x_23, x_13); x_25 = lean_box(0); -x_26 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__40; +x_26 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__38; lean_inc(x_11); x_27 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_27, 0, x_11); @@ -29598,12 +30234,12 @@ lean_inc(x_11); x_29 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_29, 0, x_11); lean_ctor_set(x_29, 1, x_28); -x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__9; +x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__9; lean_inc(x_13); lean_inc(x_14); x_31 = l_Lean_addMacroScope(x_14, x_30, x_13); -x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__5; -x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__11; +x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__5; +x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__11; lean_inc(x_11); x_34 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_34, 0, x_11); @@ -29614,14 +30250,14 @@ x_35 = l_Lean_Syntax_mkApp___closed__3; x_36 = lean_array_push(x_35, x_29); x_37 = lean_array_push(x_36, x_34); x_38 = lean_box(2); -x_39 = l_Lean_Syntax_expandInterpolatedStr___closed__5; +x_39 = l_Lean_TSyntax_expandInterpolatedStr___closed__5; x_40 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_40, 0, x_38); lean_ctor_set(x_40, 1, x_39); lean_ctor_set(x_40, 2, x_37); x_41 = l_Lean_mkOptionalNode___closed__2; x_42 = lean_array_push(x_41, x_40); -x_43 = l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__2; +x_43 = l_Lean_mkNullNode___closed__2; x_44 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_44, 0, x_38); lean_ctor_set(x_44, 1, x_43); @@ -29638,13 +30274,13 @@ lean_inc(x_11); x_49 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_49, 0, x_11); lean_ctor_set(x_49, 1, x_48); -x_50 = l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__4; +x_50 = l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__4; lean_inc(x_16); x_51 = lean_array_push(x_50, x_16); x_52 = lean_array_push(x_51, x_47); lean_inc(x_49); x_53 = lean_array_push(x_52, x_49); -x_54 = l_Lean_Syntax_expandInterpolatedStr___closed__3; +x_54 = l_Lean_TSyntax_expandInterpolatedStr___closed__3; x_55 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_55, 0, x_38); lean_ctor_set(x_55, 1, x_54); @@ -29654,7 +30290,7 @@ x_57 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_57, 0, x_38); lean_ctor_set(x_57, 1, x_43); lean_ctor_set(x_57, 2, x_56); -x_58 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__31; +x_58 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__29; lean_inc(x_11); x_59 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_59, 0, x_11); @@ -29669,7 +30305,7 @@ x_63 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_63, 0, x_38); lean_ctor_set(x_63, 1, x_43); lean_ctor_set(x_63, 2, x_62); -x_64 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__49; +x_64 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__47; lean_inc(x_11); x_65 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_65, 0, x_11); @@ -29680,11 +30316,11 @@ x_68 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_68, 0, x_38); lean_ctor_set(x_68, 1, x_43); lean_ctor_set(x_68, 2, x_67); -x_69 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__14; +x_69 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__14; lean_inc(x_13); lean_inc(x_14); x_70 = l_Lean_addMacroScope(x_14, x_69, x_13); -x_71 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__13; +x_71 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__13; lean_inc(x_11); x_72 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_72, 0, x_11); @@ -29699,10 +30335,10 @@ x_77 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_77, 0, x_38); lean_ctor_set(x_77, 1, x_76); lean_ctor_set(x_77, 2, x_75); -x_78 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__17; +x_78 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__17; x_79 = l_Lean_addMacroScope(x_14, x_78, x_13); -x_80 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__16; -x_81 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__19; +x_80 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__16; +x_81 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__19; lean_inc(x_11); x_82 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_82, 0, x_11); @@ -29733,7 +30369,7 @@ x_93 = lean_array_push(x_92, x_61); lean_inc(x_93); x_94 = lean_array_push(x_93, x_68); x_95 = lean_array_push(x_94, x_89); -x_96 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__20; +x_96 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__20; x_97 = lean_array_push(x_95, x_96); x_98 = lean_array_push(x_97, x_74); lean_inc(x_91); @@ -29746,14 +30382,14 @@ lean_ctor_set(x_101, 2, x_99); x_102 = lean_array_push(x_50, x_57); x_103 = lean_array_push(x_102, x_59); x_104 = lean_array_push(x_103, x_101); -x_105 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__2; +x_105 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__2; x_106 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_106, 0, x_38); lean_ctor_set(x_106, 1, x_105); lean_ctor_set(x_106, 2, x_104); x_107 = lean_array_push(x_35, x_22); x_108 = lean_array_push(x_107, x_106); -x_109 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__1; +x_109 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__1; x_110 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_110, 0, x_38); lean_ctor_set(x_110, 1, x_109); @@ -29761,7 +30397,7 @@ lean_ctor_set(x_110, 2, x_108); x_111 = lean_array_push(x_35, x_61); lean_inc(x_91); x_112 = lean_array_push(x_111, x_91); -x_113 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__68; +x_113 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__65; x_114 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_114, 0, x_38); lean_ctor_set(x_114, 1, x_113); @@ -29777,7 +30413,7 @@ lean_ctor_set(x_120, 1, x_100); lean_ctor_set(x_120, 2, x_119); x_121 = lean_array_push(x_35, x_114); x_122 = lean_array_push(x_121, x_120); -x_123 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__66; +x_123 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__63; x_124 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_124, 0, x_38); lean_ctor_set(x_124, 1, x_123); @@ -29804,7 +30440,7 @@ x_137 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_137, 0, x_38); lean_ctor_set(x_137, 1, x_8); lean_ctor_set(x_137, 2, x_136); -x_138 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____lambda__1(x_1, x_137, x_2, x_12); +x_138 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____lambda__1(x_1, x_137, x_2, x_12); lean_dec(x_2); return x_138; } @@ -29840,17 +30476,17 @@ lean_inc(x_142); x_151 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_151, 0, x_142); lean_ctor_set(x_151, 1, x_150); -x_152 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__25; +x_152 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__23; lean_inc(x_142); x_153 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_153, 0, x_142); lean_ctor_set(x_153, 1, x_152); -x_154 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__41; +x_154 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__39; lean_inc(x_144); lean_inc(x_145); x_155 = l_Lean_addMacroScope(x_145, x_154, x_144); x_156 = lean_box(0); -x_157 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__40; +x_157 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__38; lean_inc(x_142); x_158 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_158, 0, x_142); @@ -29862,12 +30498,12 @@ lean_inc(x_142); x_160 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_160, 0, x_142); lean_ctor_set(x_160, 1, x_159); -x_161 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__9; +x_161 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__9; lean_inc(x_144); lean_inc(x_145); x_162 = l_Lean_addMacroScope(x_145, x_161, x_144); -x_163 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__5; -x_164 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__11; +x_163 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__5; +x_164 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__11; lean_inc(x_142); x_165 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_165, 0, x_142); @@ -29878,14 +30514,14 @@ x_166 = l_Lean_Syntax_mkApp___closed__3; x_167 = lean_array_push(x_166, x_160); x_168 = lean_array_push(x_167, x_165); x_169 = lean_box(2); -x_170 = l_Lean_Syntax_expandInterpolatedStr___closed__5; +x_170 = l_Lean_TSyntax_expandInterpolatedStr___closed__5; x_171 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_171, 0, x_169); lean_ctor_set(x_171, 1, x_170); lean_ctor_set(x_171, 2, x_168); x_172 = l_Lean_mkOptionalNode___closed__2; x_173 = lean_array_push(x_172, x_171); -x_174 = l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__2; +x_174 = l_Lean_mkNullNode___closed__2; x_175 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_175, 0, x_169); lean_ctor_set(x_175, 1, x_174); @@ -29902,13 +30538,13 @@ lean_inc(x_142); x_180 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_180, 0, x_142); lean_ctor_set(x_180, 1, x_179); -x_181 = l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__4; +x_181 = l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__4; lean_inc(x_147); x_182 = lean_array_push(x_181, x_147); x_183 = lean_array_push(x_182, x_178); lean_inc(x_180); x_184 = lean_array_push(x_183, x_180); -x_185 = l_Lean_Syntax_expandInterpolatedStr___closed__3; +x_185 = l_Lean_TSyntax_expandInterpolatedStr___closed__3; x_186 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_186, 0, x_169); lean_ctor_set(x_186, 1, x_185); @@ -29918,7 +30554,7 @@ x_188 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_188, 0, x_169); lean_ctor_set(x_188, 1, x_174); lean_ctor_set(x_188, 2, x_187); -x_189 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__31; +x_189 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__29; lean_inc(x_142); x_190 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_190, 0, x_142); @@ -29933,7 +30569,7 @@ x_194 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_194, 0, x_169); lean_ctor_set(x_194, 1, x_174); lean_ctor_set(x_194, 2, x_193); -x_195 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__49; +x_195 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__47; lean_inc(x_142); x_196 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_196, 0, x_142); @@ -29944,11 +30580,11 @@ x_199 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_199, 0, x_169); lean_ctor_set(x_199, 1, x_174); lean_ctor_set(x_199, 2, x_198); -x_200 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__14; +x_200 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__14; lean_inc(x_144); lean_inc(x_145); x_201 = l_Lean_addMacroScope(x_145, x_200, x_144); -x_202 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__13; +x_202 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__13; lean_inc(x_142); x_203 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_203, 0, x_142); @@ -29963,10 +30599,10 @@ x_208 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_208, 0, x_169); lean_ctor_set(x_208, 1, x_207); lean_ctor_set(x_208, 2, x_206); -x_209 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__17; +x_209 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__17; x_210 = l_Lean_addMacroScope(x_145, x_209, x_144); -x_211 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__16; -x_212 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__19; +x_211 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__16; +x_212 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__19; lean_inc(x_142); x_213 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_213, 0, x_142); @@ -29995,7 +30631,7 @@ x_223 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Ta x_224 = lean_array_push(x_223, x_192); x_225 = lean_array_push(x_224, x_199); x_226 = lean_array_push(x_225, x_220); -x_227 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__20; +x_227 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__20; x_228 = lean_array_push(x_226, x_227); x_229 = lean_array_push(x_228, x_205); x_230 = lean_array_push(x_229, x_222); @@ -30007,14 +30643,14 @@ lean_ctor_set(x_232, 2, x_230); x_233 = lean_array_push(x_181, x_188); x_234 = lean_array_push(x_233, x_190); x_235 = lean_array_push(x_234, x_232); -x_236 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__2; +x_236 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__2; x_237 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_237, 0, x_169); lean_ctor_set(x_237, 1, x_236); lean_ctor_set(x_237, 2, x_235); x_238 = lean_array_push(x_166, x_153); x_239 = lean_array_push(x_238, x_237); -x_240 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__1; +x_240 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__1; x_241 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_241, 0, x_169); lean_ctor_set(x_241, 1, x_240); @@ -30041,17 +30677,17 @@ x_254 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_254, 0, x_169); lean_ctor_set(x_254, 1, x_8); lean_ctor_set(x_254, 2, x_253); -x_255 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____lambda__1(x_1, x_254, x_2, x_143); +x_255 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____lambda__1(x_1, x_254, x_2, x_143); lean_dec(x_2); return x_255; } } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____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_Tactic_expandSimp____x40_Init_Meta___hyg_13759____lambda__1(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____lambda__1(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } @@ -30186,22 +30822,22 @@ x_1 = l_Lean_Parser_Tactic_simpArith___closed__10; return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14552____closed__1() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15670____closed__1() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__13; +x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__13; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14552____closed__2() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15670____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__13; +x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__13; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14552____closed__1; +x_3 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15670____closed__1; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -30209,17 +30845,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14552____closed__3() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15670____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__13; +x_2 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__13; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14552_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15670_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; @@ -30261,17 +30897,17 @@ lean_inc(x_11); x_20 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_20, 0, x_11); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__25; +x_21 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__23; lean_inc(x_11); x_22 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_22, 0, x_11); lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__41; +x_23 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__39; lean_inc(x_13); lean_inc(x_14); x_24 = l_Lean_addMacroScope(x_14, x_23, x_13); x_25 = lean_box(0); -x_26 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__40; +x_26 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__38; lean_inc(x_11); x_27 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_27, 0, x_11); @@ -30283,12 +30919,12 @@ lean_inc(x_11); x_29 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_29, 0, x_11); lean_ctor_set(x_29, 1, x_28); -x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__9; +x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__9; lean_inc(x_13); lean_inc(x_14); x_31 = l_Lean_addMacroScope(x_14, x_30, x_13); -x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__5; -x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__11; +x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__5; +x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__11; lean_inc(x_11); x_34 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_34, 0, x_11); @@ -30299,14 +30935,14 @@ x_35 = l_Lean_Syntax_mkApp___closed__3; x_36 = lean_array_push(x_35, x_29); x_37 = lean_array_push(x_36, x_34); x_38 = lean_box(2); -x_39 = l_Lean_Syntax_expandInterpolatedStr___closed__5; +x_39 = l_Lean_TSyntax_expandInterpolatedStr___closed__5; x_40 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_40, 0, x_38); lean_ctor_set(x_40, 1, x_39); lean_ctor_set(x_40, 2, x_37); x_41 = l_Lean_mkOptionalNode___closed__2; x_42 = lean_array_push(x_41, x_40); -x_43 = l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__2; +x_43 = l_Lean_mkNullNode___closed__2; x_44 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_44, 0, x_38); lean_ctor_set(x_44, 1, x_43); @@ -30323,13 +30959,13 @@ lean_inc(x_11); x_49 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_49, 0, x_11); lean_ctor_set(x_49, 1, x_48); -x_50 = l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__4; +x_50 = l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__4; lean_inc(x_16); x_51 = lean_array_push(x_50, x_16); x_52 = lean_array_push(x_51, x_47); lean_inc(x_49); x_53 = lean_array_push(x_52, x_49); -x_54 = l_Lean_Syntax_expandInterpolatedStr___closed__3; +x_54 = l_Lean_TSyntax_expandInterpolatedStr___closed__3; x_55 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_55, 0, x_38); lean_ctor_set(x_55, 1, x_54); @@ -30339,7 +30975,7 @@ x_57 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_57, 0, x_38); lean_ctor_set(x_57, 1, x_43); lean_ctor_set(x_57, 2, x_56); -x_58 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__31; +x_58 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__29; lean_inc(x_11); x_59 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_59, 0, x_11); @@ -30354,7 +30990,7 @@ x_63 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_63, 0, x_38); lean_ctor_set(x_63, 1, x_43); lean_ctor_set(x_63, 2, x_62); -x_64 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__49; +x_64 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__47; lean_inc(x_11); x_65 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_65, 0, x_11); @@ -30365,11 +31001,11 @@ x_68 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_68, 0, x_38); lean_ctor_set(x_68, 1, x_43); lean_ctor_set(x_68, 2, x_67); -x_69 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14552____closed__3; +x_69 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15670____closed__3; lean_inc(x_13); lean_inc(x_14); x_70 = l_Lean_addMacroScope(x_14, x_69, x_13); -x_71 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14552____closed__2; +x_71 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15670____closed__2; lean_inc(x_11); x_72 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_72, 0, x_11); @@ -30384,10 +31020,10 @@ x_77 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_77, 0, x_38); lean_ctor_set(x_77, 1, x_76); lean_ctor_set(x_77, 2, x_75); -x_78 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__17; +x_78 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__17; x_79 = l_Lean_addMacroScope(x_14, x_78, x_13); -x_80 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__16; -x_81 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__19; +x_80 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__16; +x_81 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__19; lean_inc(x_11); x_82 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_82, 0, x_11); @@ -30418,7 +31054,7 @@ x_93 = lean_array_push(x_92, x_61); lean_inc(x_93); x_94 = lean_array_push(x_93, x_68); x_95 = lean_array_push(x_94, x_89); -x_96 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__20; +x_96 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__20; x_97 = lean_array_push(x_95, x_96); x_98 = lean_array_push(x_97, x_74); lean_inc(x_91); @@ -30431,14 +31067,14 @@ lean_ctor_set(x_101, 2, x_99); x_102 = lean_array_push(x_50, x_57); x_103 = lean_array_push(x_102, x_59); x_104 = lean_array_push(x_103, x_101); -x_105 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__2; +x_105 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__2; x_106 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_106, 0, x_38); lean_ctor_set(x_106, 1, x_105); lean_ctor_set(x_106, 2, x_104); x_107 = lean_array_push(x_35, x_22); x_108 = lean_array_push(x_107, x_106); -x_109 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__1; +x_109 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__1; x_110 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_110, 0, x_38); lean_ctor_set(x_110, 1, x_109); @@ -30446,7 +31082,7 @@ lean_ctor_set(x_110, 2, x_108); x_111 = lean_array_push(x_35, x_61); lean_inc(x_91); x_112 = lean_array_push(x_111, x_91); -x_113 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__68; +x_113 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__65; x_114 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_114, 0, x_38); lean_ctor_set(x_114, 1, x_113); @@ -30462,7 +31098,7 @@ lean_ctor_set(x_120, 1, x_100); lean_ctor_set(x_120, 2, x_119); x_121 = lean_array_push(x_35, x_114); x_122 = lean_array_push(x_121, x_120); -x_123 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__66; +x_123 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__63; x_124 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_124, 0, x_38); lean_ctor_set(x_124, 1, x_123); @@ -30489,7 +31125,7 @@ x_137 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_137, 0, x_38); lean_ctor_set(x_137, 1, x_8); lean_ctor_set(x_137, 2, x_136); -x_138 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____lambda__1(x_1, x_137, x_2, x_12); +x_138 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____lambda__1(x_1, x_137, x_2, x_12); lean_dec(x_2); return x_138; } @@ -30525,17 +31161,17 @@ lean_inc(x_142); x_151 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_151, 0, x_142); lean_ctor_set(x_151, 1, x_150); -x_152 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__25; +x_152 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__23; lean_inc(x_142); x_153 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_153, 0, x_142); lean_ctor_set(x_153, 1, x_152); -x_154 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__41; +x_154 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__39; lean_inc(x_144); lean_inc(x_145); x_155 = l_Lean_addMacroScope(x_145, x_154, x_144); x_156 = lean_box(0); -x_157 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__40; +x_157 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__38; lean_inc(x_142); x_158 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_158, 0, x_142); @@ -30547,12 +31183,12 @@ lean_inc(x_142); x_160 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_160, 0, x_142); lean_ctor_set(x_160, 1, x_159); -x_161 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__9; +x_161 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__9; lean_inc(x_144); lean_inc(x_145); x_162 = l_Lean_addMacroScope(x_145, x_161, x_144); -x_163 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__5; -x_164 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__11; +x_163 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__5; +x_164 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__11; lean_inc(x_142); x_165 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_165, 0, x_142); @@ -30563,14 +31199,14 @@ x_166 = l_Lean_Syntax_mkApp___closed__3; x_167 = lean_array_push(x_166, x_160); x_168 = lean_array_push(x_167, x_165); x_169 = lean_box(2); -x_170 = l_Lean_Syntax_expandInterpolatedStr___closed__5; +x_170 = l_Lean_TSyntax_expandInterpolatedStr___closed__5; x_171 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_171, 0, x_169); lean_ctor_set(x_171, 1, x_170); lean_ctor_set(x_171, 2, x_168); x_172 = l_Lean_mkOptionalNode___closed__2; x_173 = lean_array_push(x_172, x_171); -x_174 = l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__2; +x_174 = l_Lean_mkNullNode___closed__2; x_175 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_175, 0, x_169); lean_ctor_set(x_175, 1, x_174); @@ -30587,13 +31223,13 @@ lean_inc(x_142); x_180 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_180, 0, x_142); lean_ctor_set(x_180, 1, x_179); -x_181 = l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__4; +x_181 = l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__4; lean_inc(x_147); x_182 = lean_array_push(x_181, x_147); x_183 = lean_array_push(x_182, x_178); lean_inc(x_180); x_184 = lean_array_push(x_183, x_180); -x_185 = l_Lean_Syntax_expandInterpolatedStr___closed__3; +x_185 = l_Lean_TSyntax_expandInterpolatedStr___closed__3; x_186 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_186, 0, x_169); lean_ctor_set(x_186, 1, x_185); @@ -30603,7 +31239,7 @@ x_188 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_188, 0, x_169); lean_ctor_set(x_188, 1, x_174); lean_ctor_set(x_188, 2, x_187); -x_189 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__31; +x_189 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__29; lean_inc(x_142); x_190 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_190, 0, x_142); @@ -30618,7 +31254,7 @@ x_194 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_194, 0, x_169); lean_ctor_set(x_194, 1, x_174); lean_ctor_set(x_194, 2, x_193); -x_195 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__49; +x_195 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__47; lean_inc(x_142); x_196 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_196, 0, x_142); @@ -30629,11 +31265,11 @@ x_199 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_199, 0, x_169); lean_ctor_set(x_199, 1, x_174); lean_ctor_set(x_199, 2, x_198); -x_200 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14552____closed__3; +x_200 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15670____closed__3; lean_inc(x_144); lean_inc(x_145); x_201 = l_Lean_addMacroScope(x_145, x_200, x_144); -x_202 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14552____closed__2; +x_202 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15670____closed__2; lean_inc(x_142); x_203 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_203, 0, x_142); @@ -30648,10 +31284,10 @@ x_208 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_208, 0, x_169); lean_ctor_set(x_208, 1, x_207); lean_ctor_set(x_208, 2, x_206); -x_209 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__17; +x_209 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__17; x_210 = l_Lean_addMacroScope(x_145, x_209, x_144); -x_211 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__16; -x_212 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__19; +x_211 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__16; +x_212 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__19; lean_inc(x_142); x_213 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_213, 0, x_142); @@ -30680,7 +31316,7 @@ x_223 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Ta x_224 = lean_array_push(x_223, x_192); x_225 = lean_array_push(x_224, x_199); x_226 = lean_array_push(x_225, x_220); -x_227 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__20; +x_227 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__20; x_228 = lean_array_push(x_226, x_227); x_229 = lean_array_push(x_228, x_205); x_230 = lean_array_push(x_229, x_222); @@ -30692,14 +31328,14 @@ lean_ctor_set(x_232, 2, x_230); x_233 = lean_array_push(x_181, x_188); x_234 = lean_array_push(x_233, x_190); x_235 = lean_array_push(x_234, x_232); -x_236 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__2; +x_236 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__2; x_237 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_237, 0, x_169); lean_ctor_set(x_237, 1, x_236); lean_ctor_set(x_237, 2, x_235); x_238 = lean_array_push(x_166, x_153); x_239 = lean_array_push(x_238, x_237); -x_240 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__1; +x_240 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__1; x_241 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_241, 0, x_169); lean_ctor_set(x_241, 1, x_240); @@ -30726,7 +31362,7 @@ x_254 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_254, 0, x_169); lean_ctor_set(x_254, 1, x_8); lean_ctor_set(x_254, 2, x_253); -x_255 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____lambda__1(x_1, x_254, x_2, x_143); +x_255 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____lambda__1(x_1, x_254, x_2, x_143); lean_dec(x_2); return x_255; } @@ -30862,7 +31498,7 @@ x_1 = l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__10; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15345_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16479_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; @@ -30904,17 +31540,17 @@ lean_inc(x_11); x_20 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_20, 0, x_11); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__25; +x_21 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__23; lean_inc(x_11); x_22 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_22, 0, x_11); lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__41; +x_23 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__39; lean_inc(x_13); lean_inc(x_14); x_24 = l_Lean_addMacroScope(x_14, x_23, x_13); x_25 = lean_box(0); -x_26 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__40; +x_26 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__38; lean_inc(x_11); x_27 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_27, 0, x_11); @@ -30926,12 +31562,12 @@ lean_inc(x_11); x_29 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_29, 0, x_11); lean_ctor_set(x_29, 1, x_28); -x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__9; +x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__9; lean_inc(x_13); lean_inc(x_14); x_31 = l_Lean_addMacroScope(x_14, x_30, x_13); -x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__5; -x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__11; +x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__5; +x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__11; lean_inc(x_11); x_34 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_34, 0, x_11); @@ -30942,14 +31578,14 @@ x_35 = l_Lean_Syntax_mkApp___closed__3; x_36 = lean_array_push(x_35, x_29); x_37 = lean_array_push(x_36, x_34); x_38 = lean_box(2); -x_39 = l_Lean_Syntax_expandInterpolatedStr___closed__5; +x_39 = l_Lean_TSyntax_expandInterpolatedStr___closed__5; x_40 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_40, 0, x_38); lean_ctor_set(x_40, 1, x_39); lean_ctor_set(x_40, 2, x_37); x_41 = l_Lean_mkOptionalNode___closed__2; x_42 = lean_array_push(x_41, x_40); -x_43 = l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__2; +x_43 = l_Lean_mkNullNode___closed__2; x_44 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_44, 0, x_38); lean_ctor_set(x_44, 1, x_43); @@ -30966,13 +31602,13 @@ lean_inc(x_11); x_49 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_49, 0, x_11); lean_ctor_set(x_49, 1, x_48); -x_50 = l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__4; +x_50 = l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__4; lean_inc(x_16); x_51 = lean_array_push(x_50, x_16); x_52 = lean_array_push(x_51, x_47); lean_inc(x_49); x_53 = lean_array_push(x_52, x_49); -x_54 = l_Lean_Syntax_expandInterpolatedStr___closed__3; +x_54 = l_Lean_TSyntax_expandInterpolatedStr___closed__3; x_55 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_55, 0, x_38); lean_ctor_set(x_55, 1, x_54); @@ -30982,7 +31618,7 @@ x_57 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_57, 0, x_38); lean_ctor_set(x_57, 1, x_43); lean_ctor_set(x_57, 2, x_56); -x_58 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__31; +x_58 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__29; lean_inc(x_11); x_59 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_59, 0, x_11); @@ -30997,7 +31633,7 @@ x_63 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_63, 0, x_38); lean_ctor_set(x_63, 1, x_43); lean_ctor_set(x_63, 2, x_62); -x_64 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__49; +x_64 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__47; lean_inc(x_11); x_65 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_65, 0, x_11); @@ -31008,11 +31644,11 @@ x_68 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_68, 0, x_38); lean_ctor_set(x_68, 1, x_43); lean_ctor_set(x_68, 2, x_67); -x_69 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14552____closed__3; +x_69 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15670____closed__3; lean_inc(x_13); lean_inc(x_14); x_70 = l_Lean_addMacroScope(x_14, x_69, x_13); -x_71 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14552____closed__2; +x_71 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15670____closed__2; lean_inc(x_11); x_72 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_72, 0, x_11); @@ -31027,12 +31663,12 @@ x_77 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_77, 0, x_38); lean_ctor_set(x_77, 1, x_76); lean_ctor_set(x_77, 2, x_75); -x_78 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__17; +x_78 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__17; lean_inc(x_13); lean_inc(x_14); x_79 = l_Lean_addMacroScope(x_14, x_78, x_13); -x_80 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__16; -x_81 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__19; +x_80 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__16; +x_81 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__19; lean_inc(x_11); x_82 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_82, 0, x_11); @@ -31054,9 +31690,9 @@ lean_inc(x_11); x_89 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_89, 0, x_11); lean_ctor_set(x_89, 1, x_88); -x_90 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__14; +x_90 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__14; x_91 = l_Lean_addMacroScope(x_14, x_90, x_13); -x_92 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__13; +x_92 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__13; lean_inc(x_11); x_93 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_93, 0, x_11); @@ -31094,7 +31730,7 @@ x_108 = lean_array_push(x_107, x_61); lean_inc(x_108); x_109 = lean_array_push(x_108, x_68); x_110 = lean_array_push(x_109, x_104); -x_111 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__20; +x_111 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__20; x_112 = lean_array_push(x_110, x_111); x_113 = lean_array_push(x_112, x_74); lean_inc(x_106); @@ -31107,14 +31743,14 @@ lean_ctor_set(x_116, 2, x_114); x_117 = lean_array_push(x_50, x_57); x_118 = lean_array_push(x_117, x_59); x_119 = lean_array_push(x_118, x_116); -x_120 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__2; +x_120 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__2; x_121 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_121, 0, x_38); lean_ctor_set(x_121, 1, x_120); lean_ctor_set(x_121, 2, x_119); x_122 = lean_array_push(x_35, x_22); x_123 = lean_array_push(x_122, x_121); -x_124 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__1; +x_124 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__1; x_125 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_125, 0, x_38); lean_ctor_set(x_125, 1, x_124); @@ -31122,7 +31758,7 @@ lean_ctor_set(x_125, 2, x_123); x_126 = lean_array_push(x_35, x_61); lean_inc(x_106); x_127 = lean_array_push(x_126, x_106); -x_128 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__68; +x_128 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__65; x_129 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_129, 0, x_38); lean_ctor_set(x_129, 1, x_128); @@ -31138,7 +31774,7 @@ lean_ctor_set(x_135, 1, x_115); lean_ctor_set(x_135, 2, x_134); x_136 = lean_array_push(x_35, x_129); x_137 = lean_array_push(x_136, x_135); -x_138 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__66; +x_138 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__63; x_139 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_139, 0, x_38); lean_ctor_set(x_139, 1, x_138); @@ -31165,7 +31801,7 @@ x_152 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_152, 0, x_38); lean_ctor_set(x_152, 1, x_8); lean_ctor_set(x_152, 2, x_151); -x_153 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____lambda__1(x_1, x_152, x_2, x_12); +x_153 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____lambda__1(x_1, x_152, x_2, x_12); lean_dec(x_2); return x_153; } @@ -31201,17 +31837,17 @@ lean_inc(x_157); x_166 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_166, 0, x_157); lean_ctor_set(x_166, 1, x_165); -x_167 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__25; +x_167 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__23; lean_inc(x_157); x_168 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_168, 0, x_157); lean_ctor_set(x_168, 1, x_167); -x_169 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__41; +x_169 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__39; lean_inc(x_159); lean_inc(x_160); x_170 = l_Lean_addMacroScope(x_160, x_169, x_159); x_171 = lean_box(0); -x_172 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__40; +x_172 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__38; lean_inc(x_157); x_173 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_173, 0, x_157); @@ -31223,12 +31859,12 @@ lean_inc(x_157); x_175 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_175, 0, x_157); lean_ctor_set(x_175, 1, x_174); -x_176 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__9; +x_176 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__9; lean_inc(x_159); lean_inc(x_160); x_177 = l_Lean_addMacroScope(x_160, x_176, x_159); -x_178 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__5; -x_179 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__11; +x_178 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__5; +x_179 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__11; lean_inc(x_157); x_180 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_180, 0, x_157); @@ -31239,14 +31875,14 @@ x_181 = l_Lean_Syntax_mkApp___closed__3; x_182 = lean_array_push(x_181, x_175); x_183 = lean_array_push(x_182, x_180); x_184 = lean_box(2); -x_185 = l_Lean_Syntax_expandInterpolatedStr___closed__5; +x_185 = l_Lean_TSyntax_expandInterpolatedStr___closed__5; x_186 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_186, 0, x_184); lean_ctor_set(x_186, 1, x_185); lean_ctor_set(x_186, 2, x_183); x_187 = l_Lean_mkOptionalNode___closed__2; x_188 = lean_array_push(x_187, x_186); -x_189 = l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__2; +x_189 = l_Lean_mkNullNode___closed__2; x_190 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_190, 0, x_184); lean_ctor_set(x_190, 1, x_189); @@ -31263,13 +31899,13 @@ lean_inc(x_157); x_195 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_195, 0, x_157); lean_ctor_set(x_195, 1, x_194); -x_196 = l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__4; +x_196 = l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__4; lean_inc(x_162); x_197 = lean_array_push(x_196, x_162); x_198 = lean_array_push(x_197, x_193); lean_inc(x_195); x_199 = lean_array_push(x_198, x_195); -x_200 = l_Lean_Syntax_expandInterpolatedStr___closed__3; +x_200 = l_Lean_TSyntax_expandInterpolatedStr___closed__3; x_201 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_201, 0, x_184); lean_ctor_set(x_201, 1, x_200); @@ -31279,7 +31915,7 @@ x_203 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_203, 0, x_184); lean_ctor_set(x_203, 1, x_189); lean_ctor_set(x_203, 2, x_202); -x_204 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__31; +x_204 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__29; lean_inc(x_157); x_205 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_205, 0, x_157); @@ -31294,7 +31930,7 @@ x_209 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_209, 0, x_184); lean_ctor_set(x_209, 1, x_189); lean_ctor_set(x_209, 2, x_208); -x_210 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__49; +x_210 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__47; lean_inc(x_157); x_211 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_211, 0, x_157); @@ -31305,11 +31941,11 @@ x_214 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_214, 0, x_184); lean_ctor_set(x_214, 1, x_189); lean_ctor_set(x_214, 2, x_213); -x_215 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14552____closed__3; +x_215 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15670____closed__3; lean_inc(x_159); lean_inc(x_160); x_216 = l_Lean_addMacroScope(x_160, x_215, x_159); -x_217 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14552____closed__2; +x_217 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15670____closed__2; lean_inc(x_157); x_218 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_218, 0, x_157); @@ -31324,12 +31960,12 @@ x_223 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_223, 0, x_184); lean_ctor_set(x_223, 1, x_222); lean_ctor_set(x_223, 2, x_221); -x_224 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__17; +x_224 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__17; lean_inc(x_159); lean_inc(x_160); x_225 = l_Lean_addMacroScope(x_160, x_224, x_159); -x_226 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__16; -x_227 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__19; +x_226 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__16; +x_227 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__19; lean_inc(x_157); x_228 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_228, 0, x_157); @@ -31351,9 +31987,9 @@ lean_inc(x_157); x_235 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_235, 0, x_157); lean_ctor_set(x_235, 1, x_234); -x_236 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__14; +x_236 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__14; x_237 = l_Lean_addMacroScope(x_160, x_236, x_159); -x_238 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__13; +x_238 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__13; lean_inc(x_157); x_239 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_239, 0, x_157); @@ -31389,7 +32025,7 @@ x_253 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Ta x_254 = lean_array_push(x_253, x_207); x_255 = lean_array_push(x_254, x_214); x_256 = lean_array_push(x_255, x_250); -x_257 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__20; +x_257 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__20; x_258 = lean_array_push(x_256, x_257); x_259 = lean_array_push(x_258, x_220); x_260 = lean_array_push(x_259, x_252); @@ -31401,14 +32037,14 @@ lean_ctor_set(x_262, 2, x_260); x_263 = lean_array_push(x_196, x_203); x_264 = lean_array_push(x_263, x_205); x_265 = lean_array_push(x_264, x_262); -x_266 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__2; +x_266 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__2; x_267 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_267, 0, x_184); lean_ctor_set(x_267, 1, x_266); lean_ctor_set(x_267, 2, x_265); x_268 = lean_array_push(x_181, x_168); x_269 = lean_array_push(x_268, x_267); -x_270 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__1; +x_270 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__1; x_271 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_271, 0, x_184); lean_ctor_set(x_271, 1, x_270); @@ -31435,7 +32071,7 @@ x_284 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_284, 0, x_184); lean_ctor_set(x_284, 1, x_8); lean_ctor_set(x_284, 2, x_283); -x_285 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____lambda__1(x_1, x_284, x_2, x_158); +x_285 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____lambda__1(x_1, x_284, x_2, x_158); lean_dec(x_2); return x_285; } @@ -31613,7 +32249,7 @@ x_1 = l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__13; return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____lambda__1___closed__1() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____lambda__1___closed__1() { _start: { lean_object* x_1; @@ -31621,21 +32257,21 @@ x_1 = lean_mk_string_from_bytes("simp_all ", 9); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____lambda__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; 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_5 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__79; +x_5 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__77; x_6 = l_Lean_Syntax_setKind(x_1, x_5); x_7 = lean_unsigned_to_nat(0u); x_8 = l_Lean_Syntax_getArg(x_6, x_7); -x_9 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____lambda__1___closed__1; +x_9 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____lambda__1___closed__1; x_10 = l_Lean_mkAtomFrom(x_8, x_9); x_11 = l_Lean_Syntax_setArg(x_6, x_7, x_10); x_12 = l_Lean_mkOptionalNode___closed__2; x_13 = lean_array_push(x_12, x_2); x_14 = lean_box(2); -x_15 = l_Lean_nullKind; +x_15 = l_Lean_mkNullNode___closed__2; x_16 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_16, 0, x_14); lean_ctor_set(x_16, 1, x_15); @@ -31648,7 +32284,7 @@ lean_ctor_set(x_19, 1, x_4); return x_19; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__1() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__1() { _start: { lean_object* x_1; @@ -31656,22 +32292,22 @@ x_1 = lean_mk_string_from_bytes("Lean.Meta.Simp.ConfigCtx", 24); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__2() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__1; +x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__1; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__3() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__1; +x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__2; +x_3 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__2; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -31679,7 +32315,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__4() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__4() { _start: { lean_object* x_1; @@ -31687,41 +32323,41 @@ x_1 = lean_mk_string_from_bytes("ConfigCtx", 9); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__7; -x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__4; +x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__7; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__5; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__5; 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_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__7() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__6; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__6; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; @@ -31763,17 +32399,17 @@ lean_inc(x_11); x_20 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_20, 0, x_11); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__25; +x_21 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__23; lean_inc(x_11); x_22 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_22, 0, x_11); lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__41; +x_23 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__39; lean_inc(x_13); lean_inc(x_14); x_24 = l_Lean_addMacroScope(x_14, x_23, x_13); x_25 = lean_box(0); -x_26 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__40; +x_26 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__38; lean_inc(x_11); x_27 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_27, 0, x_11); @@ -31785,12 +32421,12 @@ lean_inc(x_11); x_29 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_29, 0, x_11); lean_ctor_set(x_29, 1, x_28); -x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__5; +x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__5; lean_inc(x_13); lean_inc(x_14); x_31 = l_Lean_addMacroScope(x_14, x_30, x_13); -x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__3; -x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__7; +x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__3; +x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__7; lean_inc(x_11); x_34 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_34, 0, x_11); @@ -31801,14 +32437,14 @@ x_35 = l_Lean_Syntax_mkApp___closed__3; x_36 = lean_array_push(x_35, x_29); x_37 = lean_array_push(x_36, x_34); x_38 = lean_box(2); -x_39 = l_Lean_Syntax_expandInterpolatedStr___closed__5; +x_39 = l_Lean_TSyntax_expandInterpolatedStr___closed__5; x_40 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_40, 0, x_38); lean_ctor_set(x_40, 1, x_39); lean_ctor_set(x_40, 2, x_37); x_41 = l_Lean_mkOptionalNode___closed__2; x_42 = lean_array_push(x_41, x_40); -x_43 = l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__2; +x_43 = l_Lean_mkNullNode___closed__2; x_44 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_44, 0, x_38); lean_ctor_set(x_44, 1, x_43); @@ -31825,13 +32461,13 @@ lean_inc(x_11); x_49 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_49, 0, x_11); lean_ctor_set(x_49, 1, x_48); -x_50 = l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__4; +x_50 = l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__4; lean_inc(x_16); x_51 = lean_array_push(x_50, x_16); x_52 = lean_array_push(x_51, x_47); lean_inc(x_49); x_53 = lean_array_push(x_52, x_49); -x_54 = l_Lean_Syntax_expandInterpolatedStr___closed__3; +x_54 = l_Lean_TSyntax_expandInterpolatedStr___closed__3; x_55 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_55, 0, x_38); lean_ctor_set(x_55, 1, x_54); @@ -31841,7 +32477,7 @@ x_57 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_57, 0, x_38); lean_ctor_set(x_57, 1, x_43); lean_ctor_set(x_57, 2, x_56); -x_58 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__31; +x_58 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__29; lean_inc(x_11); x_59 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_59, 0, x_11); @@ -31856,7 +32492,7 @@ x_63 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_63, 0, x_38); lean_ctor_set(x_63, 1, x_43); lean_ctor_set(x_63, 2, x_62); -x_64 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__49; +x_64 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__47; lean_inc(x_11); x_65 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_65, 0, x_11); @@ -31867,11 +32503,11 @@ x_68 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_68, 0, x_38); lean_ctor_set(x_68, 1, x_43); lean_ctor_set(x_68, 2, x_67); -x_69 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__14; +x_69 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__14; lean_inc(x_13); lean_inc(x_14); x_70 = l_Lean_addMacroScope(x_14, x_69, x_13); -x_71 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__13; +x_71 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__13; lean_inc(x_11); x_72 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_72, 0, x_11); @@ -31886,10 +32522,10 @@ x_77 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_77, 0, x_38); lean_ctor_set(x_77, 1, x_76); lean_ctor_set(x_77, 2, x_75); -x_78 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__17; +x_78 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__17; x_79 = l_Lean_addMacroScope(x_14, x_78, x_13); -x_80 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__16; -x_81 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__19; +x_80 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__16; +x_81 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__19; lean_inc(x_11); x_82 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_82, 0, x_11); @@ -31920,7 +32556,7 @@ x_93 = lean_array_push(x_92, x_61); lean_inc(x_93); x_94 = lean_array_push(x_93, x_68); x_95 = lean_array_push(x_94, x_89); -x_96 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__20; +x_96 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__20; x_97 = lean_array_push(x_95, x_96); x_98 = lean_array_push(x_97, x_74); lean_inc(x_91); @@ -31933,14 +32569,14 @@ lean_ctor_set(x_101, 2, x_99); x_102 = lean_array_push(x_50, x_57); x_103 = lean_array_push(x_102, x_59); x_104 = lean_array_push(x_103, x_101); -x_105 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__2; +x_105 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__2; x_106 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_106, 0, x_38); lean_ctor_set(x_106, 1, x_105); lean_ctor_set(x_106, 2, x_104); x_107 = lean_array_push(x_35, x_22); x_108 = lean_array_push(x_107, x_106); -x_109 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__1; +x_109 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__1; x_110 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_110, 0, x_38); lean_ctor_set(x_110, 1, x_109); @@ -31948,7 +32584,7 @@ lean_ctor_set(x_110, 2, x_108); x_111 = lean_array_push(x_35, x_61); lean_inc(x_91); x_112 = lean_array_push(x_111, x_91); -x_113 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__68; +x_113 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__65; x_114 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_114, 0, x_38); lean_ctor_set(x_114, 1, x_113); @@ -31964,7 +32600,7 @@ lean_ctor_set(x_120, 1, x_100); lean_ctor_set(x_120, 2, x_119); x_121 = lean_array_push(x_35, x_114); x_122 = lean_array_push(x_121, x_120); -x_123 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__66; +x_123 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__63; x_124 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_124, 0, x_38); lean_ctor_set(x_124, 1, x_123); @@ -31991,7 +32627,7 @@ x_137 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_137, 0, x_38); lean_ctor_set(x_137, 1, x_8); lean_ctor_set(x_137, 2, x_136); -x_138 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____lambda__1(x_1, x_137, x_2, x_12); +x_138 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____lambda__1(x_1, x_137, x_2, x_12); lean_dec(x_2); return x_138; } @@ -32027,17 +32663,17 @@ lean_inc(x_142); x_151 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_151, 0, x_142); lean_ctor_set(x_151, 1, x_150); -x_152 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__25; +x_152 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__23; lean_inc(x_142); x_153 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_153, 0, x_142); lean_ctor_set(x_153, 1, x_152); -x_154 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__41; +x_154 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__39; lean_inc(x_144); lean_inc(x_145); x_155 = l_Lean_addMacroScope(x_145, x_154, x_144); x_156 = lean_box(0); -x_157 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__40; +x_157 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__38; lean_inc(x_142); x_158 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_158, 0, x_142); @@ -32049,12 +32685,12 @@ lean_inc(x_142); x_160 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_160, 0, x_142); lean_ctor_set(x_160, 1, x_159); -x_161 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__5; +x_161 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__5; lean_inc(x_144); lean_inc(x_145); x_162 = l_Lean_addMacroScope(x_145, x_161, x_144); -x_163 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__3; -x_164 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__7; +x_163 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__3; +x_164 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__7; lean_inc(x_142); x_165 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_165, 0, x_142); @@ -32065,14 +32701,14 @@ x_166 = l_Lean_Syntax_mkApp___closed__3; x_167 = lean_array_push(x_166, x_160); x_168 = lean_array_push(x_167, x_165); x_169 = lean_box(2); -x_170 = l_Lean_Syntax_expandInterpolatedStr___closed__5; +x_170 = l_Lean_TSyntax_expandInterpolatedStr___closed__5; x_171 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_171, 0, x_169); lean_ctor_set(x_171, 1, x_170); lean_ctor_set(x_171, 2, x_168); x_172 = l_Lean_mkOptionalNode___closed__2; x_173 = lean_array_push(x_172, x_171); -x_174 = l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__2; +x_174 = l_Lean_mkNullNode___closed__2; x_175 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_175, 0, x_169); lean_ctor_set(x_175, 1, x_174); @@ -32089,13 +32725,13 @@ lean_inc(x_142); x_180 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_180, 0, x_142); lean_ctor_set(x_180, 1, x_179); -x_181 = l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__4; +x_181 = l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__4; lean_inc(x_147); x_182 = lean_array_push(x_181, x_147); x_183 = lean_array_push(x_182, x_178); lean_inc(x_180); x_184 = lean_array_push(x_183, x_180); -x_185 = l_Lean_Syntax_expandInterpolatedStr___closed__3; +x_185 = l_Lean_TSyntax_expandInterpolatedStr___closed__3; x_186 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_186, 0, x_169); lean_ctor_set(x_186, 1, x_185); @@ -32105,7 +32741,7 @@ x_188 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_188, 0, x_169); lean_ctor_set(x_188, 1, x_174); lean_ctor_set(x_188, 2, x_187); -x_189 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__31; +x_189 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__29; lean_inc(x_142); x_190 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_190, 0, x_142); @@ -32120,7 +32756,7 @@ x_194 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_194, 0, x_169); lean_ctor_set(x_194, 1, x_174); lean_ctor_set(x_194, 2, x_193); -x_195 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__49; +x_195 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__47; lean_inc(x_142); x_196 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_196, 0, x_142); @@ -32131,11 +32767,11 @@ x_199 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_199, 0, x_169); lean_ctor_set(x_199, 1, x_174); lean_ctor_set(x_199, 2, x_198); -x_200 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__14; +x_200 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__14; lean_inc(x_144); lean_inc(x_145); x_201 = l_Lean_addMacroScope(x_145, x_200, x_144); -x_202 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__13; +x_202 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__13; lean_inc(x_142); x_203 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_203, 0, x_142); @@ -32150,10 +32786,10 @@ x_208 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_208, 0, x_169); lean_ctor_set(x_208, 1, x_207); lean_ctor_set(x_208, 2, x_206); -x_209 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__17; +x_209 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__17; x_210 = l_Lean_addMacroScope(x_145, x_209, x_144); -x_211 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__16; -x_212 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__19; +x_211 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__16; +x_212 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__19; lean_inc(x_142); x_213 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_213, 0, x_142); @@ -32182,7 +32818,7 @@ x_223 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Ta x_224 = lean_array_push(x_223, x_192); x_225 = lean_array_push(x_224, x_199); x_226 = lean_array_push(x_225, x_220); -x_227 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__20; +x_227 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__20; x_228 = lean_array_push(x_226, x_227); x_229 = lean_array_push(x_228, x_205); x_230 = lean_array_push(x_229, x_222); @@ -32194,14 +32830,14 @@ lean_ctor_set(x_232, 2, x_230); x_233 = lean_array_push(x_181, x_188); x_234 = lean_array_push(x_233, x_190); x_235 = lean_array_push(x_234, x_232); -x_236 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__2; +x_236 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__2; x_237 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_237, 0, x_169); lean_ctor_set(x_237, 1, x_236); lean_ctor_set(x_237, 2, x_235); x_238 = lean_array_push(x_166, x_153); x_239 = lean_array_push(x_238, x_237); -x_240 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__1; +x_240 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__1; x_241 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_241, 0, x_169); lean_ctor_set(x_241, 1, x_240); @@ -32228,17 +32864,17 @@ x_254 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_254, 0, x_169); lean_ctor_set(x_254, 1, x_8); lean_ctor_set(x_254, 2, x_253); -x_255 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____lambda__1(x_1, x_254, x_2, x_143); +x_255 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____lambda__1(x_1, x_254, x_2, x_143); lean_dec(x_2); return x_255; } } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____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_Tactic_expandSimp____x40_Init_Meta___hyg_16246____lambda__1(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____lambda__1(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } @@ -32359,7 +32995,7 @@ x_1 = l_Lean_Parser_Tactic_simpAllArith___closed__9; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17029_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18195_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; @@ -32401,17 +33037,17 @@ lean_inc(x_11); x_20 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_20, 0, x_11); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__25; +x_21 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__23; lean_inc(x_11); x_22 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_22, 0, x_11); lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__41; +x_23 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__39; lean_inc(x_13); lean_inc(x_14); x_24 = l_Lean_addMacroScope(x_14, x_23, x_13); x_25 = lean_box(0); -x_26 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__40; +x_26 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__38; lean_inc(x_11); x_27 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_27, 0, x_11); @@ -32423,12 +33059,12 @@ lean_inc(x_11); x_29 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_29, 0, x_11); lean_ctor_set(x_29, 1, x_28); -x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__5; +x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__5; lean_inc(x_13); lean_inc(x_14); x_31 = l_Lean_addMacroScope(x_14, x_30, x_13); -x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__3; -x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__7; +x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__3; +x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__7; lean_inc(x_11); x_34 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_34, 0, x_11); @@ -32439,14 +33075,14 @@ x_35 = l_Lean_Syntax_mkApp___closed__3; x_36 = lean_array_push(x_35, x_29); x_37 = lean_array_push(x_36, x_34); x_38 = lean_box(2); -x_39 = l_Lean_Syntax_expandInterpolatedStr___closed__5; +x_39 = l_Lean_TSyntax_expandInterpolatedStr___closed__5; x_40 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_40, 0, x_38); lean_ctor_set(x_40, 1, x_39); lean_ctor_set(x_40, 2, x_37); x_41 = l_Lean_mkOptionalNode___closed__2; x_42 = lean_array_push(x_41, x_40); -x_43 = l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__2; +x_43 = l_Lean_mkNullNode___closed__2; x_44 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_44, 0, x_38); lean_ctor_set(x_44, 1, x_43); @@ -32463,13 +33099,13 @@ lean_inc(x_11); x_49 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_49, 0, x_11); lean_ctor_set(x_49, 1, x_48); -x_50 = l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__4; +x_50 = l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__4; lean_inc(x_16); x_51 = lean_array_push(x_50, x_16); x_52 = lean_array_push(x_51, x_47); lean_inc(x_49); x_53 = lean_array_push(x_52, x_49); -x_54 = l_Lean_Syntax_expandInterpolatedStr___closed__3; +x_54 = l_Lean_TSyntax_expandInterpolatedStr___closed__3; x_55 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_55, 0, x_38); lean_ctor_set(x_55, 1, x_54); @@ -32479,7 +33115,7 @@ x_57 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_57, 0, x_38); lean_ctor_set(x_57, 1, x_43); lean_ctor_set(x_57, 2, x_56); -x_58 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__31; +x_58 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__29; lean_inc(x_11); x_59 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_59, 0, x_11); @@ -32494,7 +33130,7 @@ x_63 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_63, 0, x_38); lean_ctor_set(x_63, 1, x_43); lean_ctor_set(x_63, 2, x_62); -x_64 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__49; +x_64 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__47; lean_inc(x_11); x_65 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_65, 0, x_11); @@ -32505,11 +33141,11 @@ x_68 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_68, 0, x_38); lean_ctor_set(x_68, 1, x_43); lean_ctor_set(x_68, 2, x_67); -x_69 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14552____closed__3; +x_69 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15670____closed__3; lean_inc(x_13); lean_inc(x_14); x_70 = l_Lean_addMacroScope(x_14, x_69, x_13); -x_71 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14552____closed__2; +x_71 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15670____closed__2; lean_inc(x_11); x_72 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_72, 0, x_11); @@ -32524,10 +33160,10 @@ x_77 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_77, 0, x_38); lean_ctor_set(x_77, 1, x_76); lean_ctor_set(x_77, 2, x_75); -x_78 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__17; +x_78 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__17; x_79 = l_Lean_addMacroScope(x_14, x_78, x_13); -x_80 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__16; -x_81 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__19; +x_80 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__16; +x_81 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__19; lean_inc(x_11); x_82 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_82, 0, x_11); @@ -32558,7 +33194,7 @@ x_93 = lean_array_push(x_92, x_61); lean_inc(x_93); x_94 = lean_array_push(x_93, x_68); x_95 = lean_array_push(x_94, x_89); -x_96 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__20; +x_96 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__20; x_97 = lean_array_push(x_95, x_96); x_98 = lean_array_push(x_97, x_74); lean_inc(x_91); @@ -32571,14 +33207,14 @@ lean_ctor_set(x_101, 2, x_99); x_102 = lean_array_push(x_50, x_57); x_103 = lean_array_push(x_102, x_59); x_104 = lean_array_push(x_103, x_101); -x_105 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__2; +x_105 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__2; x_106 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_106, 0, x_38); lean_ctor_set(x_106, 1, x_105); lean_ctor_set(x_106, 2, x_104); x_107 = lean_array_push(x_35, x_22); x_108 = lean_array_push(x_107, x_106); -x_109 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__1; +x_109 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__1; x_110 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_110, 0, x_38); lean_ctor_set(x_110, 1, x_109); @@ -32586,7 +33222,7 @@ lean_ctor_set(x_110, 2, x_108); x_111 = lean_array_push(x_35, x_61); lean_inc(x_91); x_112 = lean_array_push(x_111, x_91); -x_113 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__68; +x_113 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__65; x_114 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_114, 0, x_38); lean_ctor_set(x_114, 1, x_113); @@ -32602,7 +33238,7 @@ lean_ctor_set(x_120, 1, x_100); lean_ctor_set(x_120, 2, x_119); x_121 = lean_array_push(x_35, x_114); x_122 = lean_array_push(x_121, x_120); -x_123 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__66; +x_123 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__63; x_124 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_124, 0, x_38); lean_ctor_set(x_124, 1, x_123); @@ -32629,7 +33265,7 @@ x_137 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_137, 0, x_38); lean_ctor_set(x_137, 1, x_8); lean_ctor_set(x_137, 2, x_136); -x_138 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____lambda__1(x_1, x_137, x_2, x_12); +x_138 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____lambda__1(x_1, x_137, x_2, x_12); lean_dec(x_2); return x_138; } @@ -32665,17 +33301,17 @@ lean_inc(x_142); x_151 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_151, 0, x_142); lean_ctor_set(x_151, 1, x_150); -x_152 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__25; +x_152 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__23; lean_inc(x_142); x_153 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_153, 0, x_142); lean_ctor_set(x_153, 1, x_152); -x_154 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__41; +x_154 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__39; lean_inc(x_144); lean_inc(x_145); x_155 = l_Lean_addMacroScope(x_145, x_154, x_144); x_156 = lean_box(0); -x_157 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__40; +x_157 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__38; lean_inc(x_142); x_158 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_158, 0, x_142); @@ -32687,12 +33323,12 @@ lean_inc(x_142); x_160 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_160, 0, x_142); lean_ctor_set(x_160, 1, x_159); -x_161 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__5; +x_161 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__5; lean_inc(x_144); lean_inc(x_145); x_162 = l_Lean_addMacroScope(x_145, x_161, x_144); -x_163 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__3; -x_164 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__7; +x_163 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__3; +x_164 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__7; lean_inc(x_142); x_165 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_165, 0, x_142); @@ -32703,14 +33339,14 @@ x_166 = l_Lean_Syntax_mkApp___closed__3; x_167 = lean_array_push(x_166, x_160); x_168 = lean_array_push(x_167, x_165); x_169 = lean_box(2); -x_170 = l_Lean_Syntax_expandInterpolatedStr___closed__5; +x_170 = l_Lean_TSyntax_expandInterpolatedStr___closed__5; x_171 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_171, 0, x_169); lean_ctor_set(x_171, 1, x_170); lean_ctor_set(x_171, 2, x_168); x_172 = l_Lean_mkOptionalNode___closed__2; x_173 = lean_array_push(x_172, x_171); -x_174 = l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__2; +x_174 = l_Lean_mkNullNode___closed__2; x_175 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_175, 0, x_169); lean_ctor_set(x_175, 1, x_174); @@ -32727,13 +33363,13 @@ lean_inc(x_142); x_180 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_180, 0, x_142); lean_ctor_set(x_180, 1, x_179); -x_181 = l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__4; +x_181 = l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__4; lean_inc(x_147); x_182 = lean_array_push(x_181, x_147); x_183 = lean_array_push(x_182, x_178); lean_inc(x_180); x_184 = lean_array_push(x_183, x_180); -x_185 = l_Lean_Syntax_expandInterpolatedStr___closed__3; +x_185 = l_Lean_TSyntax_expandInterpolatedStr___closed__3; x_186 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_186, 0, x_169); lean_ctor_set(x_186, 1, x_185); @@ -32743,7 +33379,7 @@ x_188 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_188, 0, x_169); lean_ctor_set(x_188, 1, x_174); lean_ctor_set(x_188, 2, x_187); -x_189 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__31; +x_189 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__29; lean_inc(x_142); x_190 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_190, 0, x_142); @@ -32758,7 +33394,7 @@ x_194 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_194, 0, x_169); lean_ctor_set(x_194, 1, x_174); lean_ctor_set(x_194, 2, x_193); -x_195 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__49; +x_195 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__47; lean_inc(x_142); x_196 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_196, 0, x_142); @@ -32769,11 +33405,11 @@ x_199 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_199, 0, x_169); lean_ctor_set(x_199, 1, x_174); lean_ctor_set(x_199, 2, x_198); -x_200 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14552____closed__3; +x_200 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15670____closed__3; lean_inc(x_144); lean_inc(x_145); x_201 = l_Lean_addMacroScope(x_145, x_200, x_144); -x_202 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14552____closed__2; +x_202 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15670____closed__2; lean_inc(x_142); x_203 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_203, 0, x_142); @@ -32788,10 +33424,10 @@ x_208 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_208, 0, x_169); lean_ctor_set(x_208, 1, x_207); lean_ctor_set(x_208, 2, x_206); -x_209 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__17; +x_209 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__17; x_210 = l_Lean_addMacroScope(x_145, x_209, x_144); -x_211 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__16; -x_212 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__19; +x_211 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__16; +x_212 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__19; lean_inc(x_142); x_213 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_213, 0, x_142); @@ -32820,7 +33456,7 @@ x_223 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Ta x_224 = lean_array_push(x_223, x_192); x_225 = lean_array_push(x_224, x_199); x_226 = lean_array_push(x_225, x_220); -x_227 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__20; +x_227 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__20; x_228 = lean_array_push(x_226, x_227); x_229 = lean_array_push(x_228, x_205); x_230 = lean_array_push(x_229, x_222); @@ -32832,14 +33468,14 @@ lean_ctor_set(x_232, 2, x_230); x_233 = lean_array_push(x_181, x_188); x_234 = lean_array_push(x_233, x_190); x_235 = lean_array_push(x_234, x_232); -x_236 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__2; +x_236 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__2; x_237 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_237, 0, x_169); lean_ctor_set(x_237, 1, x_236); lean_ctor_set(x_237, 2, x_235); x_238 = lean_array_push(x_166, x_153); x_239 = lean_array_push(x_238, x_237); -x_240 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__1; +x_240 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__1; x_241 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_241, 0, x_169); lean_ctor_set(x_241, 1, x_240); @@ -32866,7 +33502,7 @@ x_254 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_254, 0, x_169); lean_ctor_set(x_254, 1, x_8); lean_ctor_set(x_254, 2, x_253); -x_255 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____lambda__1(x_1, x_254, x_2, x_143); +x_255 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____lambda__1(x_1, x_254, x_2, x_143); lean_dec(x_2); return x_255; } @@ -32988,7 +33624,7 @@ x_1 = l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__9; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17812_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18994_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; @@ -33030,17 +33666,17 @@ lean_inc(x_11); x_20 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_20, 0, x_11); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__25; +x_21 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__23; lean_inc(x_11); x_22 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_22, 0, x_11); lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__41; +x_23 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__39; lean_inc(x_13); lean_inc(x_14); x_24 = l_Lean_addMacroScope(x_14, x_23, x_13); x_25 = lean_box(0); -x_26 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__40; +x_26 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__38; lean_inc(x_11); x_27 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_27, 0, x_11); @@ -33052,12 +33688,12 @@ lean_inc(x_11); x_29 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_29, 0, x_11); lean_ctor_set(x_29, 1, x_28); -x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__5; +x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__5; lean_inc(x_13); lean_inc(x_14); x_31 = l_Lean_addMacroScope(x_14, x_30, x_13); -x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__3; -x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__7; +x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__3; +x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__7; lean_inc(x_11); x_34 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_34, 0, x_11); @@ -33068,14 +33704,14 @@ x_35 = l_Lean_Syntax_mkApp___closed__3; x_36 = lean_array_push(x_35, x_29); x_37 = lean_array_push(x_36, x_34); x_38 = lean_box(2); -x_39 = l_Lean_Syntax_expandInterpolatedStr___closed__5; +x_39 = l_Lean_TSyntax_expandInterpolatedStr___closed__5; x_40 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_40, 0, x_38); lean_ctor_set(x_40, 1, x_39); lean_ctor_set(x_40, 2, x_37); x_41 = l_Lean_mkOptionalNode___closed__2; x_42 = lean_array_push(x_41, x_40); -x_43 = l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__2; +x_43 = l_Lean_mkNullNode___closed__2; x_44 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_44, 0, x_38); lean_ctor_set(x_44, 1, x_43); @@ -33092,13 +33728,13 @@ lean_inc(x_11); x_49 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_49, 0, x_11); lean_ctor_set(x_49, 1, x_48); -x_50 = l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__4; +x_50 = l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__4; lean_inc(x_16); x_51 = lean_array_push(x_50, x_16); x_52 = lean_array_push(x_51, x_47); lean_inc(x_49); x_53 = lean_array_push(x_52, x_49); -x_54 = l_Lean_Syntax_expandInterpolatedStr___closed__3; +x_54 = l_Lean_TSyntax_expandInterpolatedStr___closed__3; x_55 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_55, 0, x_38); lean_ctor_set(x_55, 1, x_54); @@ -33108,7 +33744,7 @@ x_57 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_57, 0, x_38); lean_ctor_set(x_57, 1, x_43); lean_ctor_set(x_57, 2, x_56); -x_58 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__31; +x_58 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__29; lean_inc(x_11); x_59 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_59, 0, x_11); @@ -33123,7 +33759,7 @@ x_63 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_63, 0, x_38); lean_ctor_set(x_63, 1, x_43); lean_ctor_set(x_63, 2, x_62); -x_64 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__49; +x_64 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__47; lean_inc(x_11); x_65 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_65, 0, x_11); @@ -33134,11 +33770,11 @@ x_68 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_68, 0, x_38); lean_ctor_set(x_68, 1, x_43); lean_ctor_set(x_68, 2, x_67); -x_69 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14552____closed__3; +x_69 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15670____closed__3; lean_inc(x_13); lean_inc(x_14); x_70 = l_Lean_addMacroScope(x_14, x_69, x_13); -x_71 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14552____closed__2; +x_71 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15670____closed__2; lean_inc(x_11); x_72 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_72, 0, x_11); @@ -33153,12 +33789,12 @@ x_77 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_77, 0, x_38); lean_ctor_set(x_77, 1, x_76); lean_ctor_set(x_77, 2, x_75); -x_78 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__17; +x_78 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__17; lean_inc(x_13); lean_inc(x_14); x_79 = l_Lean_addMacroScope(x_14, x_78, x_13); -x_80 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__16; -x_81 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__19; +x_80 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__16; +x_81 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__19; lean_inc(x_11); x_82 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_82, 0, x_11); @@ -33180,9 +33816,9 @@ lean_inc(x_11); x_89 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_89, 0, x_11); lean_ctor_set(x_89, 1, x_88); -x_90 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__14; +x_90 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__14; x_91 = l_Lean_addMacroScope(x_14, x_90, x_13); -x_92 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__13; +x_92 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__13; lean_inc(x_11); x_93 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_93, 0, x_11); @@ -33220,7 +33856,7 @@ x_108 = lean_array_push(x_107, x_61); lean_inc(x_108); x_109 = lean_array_push(x_108, x_68); x_110 = lean_array_push(x_109, x_104); -x_111 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__20; +x_111 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__20; x_112 = lean_array_push(x_110, x_111); x_113 = lean_array_push(x_112, x_74); lean_inc(x_106); @@ -33233,14 +33869,14 @@ lean_ctor_set(x_116, 2, x_114); x_117 = lean_array_push(x_50, x_57); x_118 = lean_array_push(x_117, x_59); x_119 = lean_array_push(x_118, x_116); -x_120 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__2; +x_120 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__2; x_121 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_121, 0, x_38); lean_ctor_set(x_121, 1, x_120); lean_ctor_set(x_121, 2, x_119); x_122 = lean_array_push(x_35, x_22); x_123 = lean_array_push(x_122, x_121); -x_124 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__1; +x_124 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__1; x_125 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_125, 0, x_38); lean_ctor_set(x_125, 1, x_124); @@ -33248,7 +33884,7 @@ lean_ctor_set(x_125, 2, x_123); x_126 = lean_array_push(x_35, x_61); lean_inc(x_106); x_127 = lean_array_push(x_126, x_106); -x_128 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__68; +x_128 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__65; x_129 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_129, 0, x_38); lean_ctor_set(x_129, 1, x_128); @@ -33264,7 +33900,7 @@ lean_ctor_set(x_135, 1, x_115); lean_ctor_set(x_135, 2, x_134); x_136 = lean_array_push(x_35, x_129); x_137 = lean_array_push(x_136, x_135); -x_138 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__66; +x_138 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__63; x_139 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_139, 0, x_38); lean_ctor_set(x_139, 1, x_138); @@ -33291,7 +33927,7 @@ x_152 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_152, 0, x_38); lean_ctor_set(x_152, 1, x_8); lean_ctor_set(x_152, 2, x_151); -x_153 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____lambda__1(x_1, x_152, x_2, x_12); +x_153 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____lambda__1(x_1, x_152, x_2, x_12); lean_dec(x_2); return x_153; } @@ -33327,17 +33963,17 @@ lean_inc(x_157); x_166 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_166, 0, x_157); lean_ctor_set(x_166, 1, x_165); -x_167 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__25; +x_167 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__23; lean_inc(x_157); x_168 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_168, 0, x_157); lean_ctor_set(x_168, 1, x_167); -x_169 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__41; +x_169 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__39; lean_inc(x_159); lean_inc(x_160); x_170 = l_Lean_addMacroScope(x_160, x_169, x_159); x_171 = lean_box(0); -x_172 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__40; +x_172 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__38; lean_inc(x_157); x_173 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_173, 0, x_157); @@ -33349,12 +33985,12 @@ lean_inc(x_157); x_175 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_175, 0, x_157); lean_ctor_set(x_175, 1, x_174); -x_176 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__5; +x_176 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__5; lean_inc(x_159); lean_inc(x_160); x_177 = l_Lean_addMacroScope(x_160, x_176, x_159); -x_178 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__3; -x_179 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__7; +x_178 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__3; +x_179 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__7; lean_inc(x_157); x_180 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_180, 0, x_157); @@ -33365,14 +34001,14 @@ x_181 = l_Lean_Syntax_mkApp___closed__3; x_182 = lean_array_push(x_181, x_175); x_183 = lean_array_push(x_182, x_180); x_184 = lean_box(2); -x_185 = l_Lean_Syntax_expandInterpolatedStr___closed__5; +x_185 = l_Lean_TSyntax_expandInterpolatedStr___closed__5; x_186 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_186, 0, x_184); lean_ctor_set(x_186, 1, x_185); lean_ctor_set(x_186, 2, x_183); x_187 = l_Lean_mkOptionalNode___closed__2; x_188 = lean_array_push(x_187, x_186); -x_189 = l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__2; +x_189 = l_Lean_mkNullNode___closed__2; x_190 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_190, 0, x_184); lean_ctor_set(x_190, 1, x_189); @@ -33389,13 +34025,13 @@ lean_inc(x_157); x_195 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_195, 0, x_157); lean_ctor_set(x_195, 1, x_194); -x_196 = l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__4; +x_196 = l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__4; lean_inc(x_162); x_197 = lean_array_push(x_196, x_162); x_198 = lean_array_push(x_197, x_193); lean_inc(x_195); x_199 = lean_array_push(x_198, x_195); -x_200 = l_Lean_Syntax_expandInterpolatedStr___closed__3; +x_200 = l_Lean_TSyntax_expandInterpolatedStr___closed__3; x_201 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_201, 0, x_184); lean_ctor_set(x_201, 1, x_200); @@ -33405,7 +34041,7 @@ x_203 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_203, 0, x_184); lean_ctor_set(x_203, 1, x_189); lean_ctor_set(x_203, 2, x_202); -x_204 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__31; +x_204 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__29; lean_inc(x_157); x_205 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_205, 0, x_157); @@ -33420,7 +34056,7 @@ x_209 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_209, 0, x_184); lean_ctor_set(x_209, 1, x_189); lean_ctor_set(x_209, 2, x_208); -x_210 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__49; +x_210 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__47; lean_inc(x_157); x_211 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_211, 0, x_157); @@ -33431,11 +34067,11 @@ x_214 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_214, 0, x_184); lean_ctor_set(x_214, 1, x_189); lean_ctor_set(x_214, 2, x_213); -x_215 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14552____closed__3; +x_215 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15670____closed__3; lean_inc(x_159); lean_inc(x_160); x_216 = l_Lean_addMacroScope(x_160, x_215, x_159); -x_217 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14552____closed__2; +x_217 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15670____closed__2; lean_inc(x_157); x_218 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_218, 0, x_157); @@ -33450,12 +34086,12 @@ x_223 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_223, 0, x_184); lean_ctor_set(x_223, 1, x_222); lean_ctor_set(x_223, 2, x_221); -x_224 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__17; +x_224 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__17; lean_inc(x_159); lean_inc(x_160); x_225 = l_Lean_addMacroScope(x_160, x_224, x_159); -x_226 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__16; -x_227 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__19; +x_226 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__16; +x_227 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__19; lean_inc(x_157); x_228 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_228, 0, x_157); @@ -33477,9 +34113,9 @@ lean_inc(x_157); x_235 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_235, 0, x_157); lean_ctor_set(x_235, 1, x_234); -x_236 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__14; +x_236 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__14; x_237 = l_Lean_addMacroScope(x_160, x_236, x_159); -x_238 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__13; +x_238 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__13; lean_inc(x_157); x_239 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_239, 0, x_157); @@ -33515,7 +34151,7 @@ x_253 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Ta x_254 = lean_array_push(x_253, x_207); x_255 = lean_array_push(x_254, x_214); x_256 = lean_array_push(x_255, x_250); -x_257 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__20; +x_257 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__20; x_258 = lean_array_push(x_256, x_257); x_259 = lean_array_push(x_258, x_220); x_260 = lean_array_push(x_259, x_252); @@ -33527,14 +34163,14 @@ lean_ctor_set(x_262, 2, x_260); x_263 = lean_array_push(x_196, x_203); x_264 = lean_array_push(x_263, x_205); x_265 = lean_array_push(x_264, x_262); -x_266 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__2; +x_266 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__2; x_267 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_267, 0, x_184); lean_ctor_set(x_267, 1, x_266); lean_ctor_set(x_267, 2, x_265); x_268 = lean_array_push(x_181, x_168); x_269 = lean_array_push(x_268, x_267); -x_270 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__1; +x_270 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__1; x_271 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_271, 0, x_184); lean_ctor_set(x_271, 1, x_270); @@ -33561,7 +34197,7 @@ x_284 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_284, 0, x_184); lean_ctor_set(x_284, 1, x_8); lean_ctor_set(x_284, 2, x_283); -x_285 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____lambda__1(x_1, x_284, x_2, x_158); +x_285 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____lambda__1(x_1, x_284, x_2, x_158); lean_dec(x_2); return x_285; } @@ -33697,7 +34333,7 @@ x_1 = l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__10; return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____lambda__1___closed__1() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____lambda__1___closed__1() { _start: { lean_object* x_1; @@ -33705,21 +34341,21 @@ x_1 = lean_mk_string_from_bytes("dsimp ", 6); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____lambda__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; 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_5 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__6; +x_5 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__6; x_6 = l_Lean_Syntax_setKind(x_1, x_5); x_7 = lean_unsigned_to_nat(0u); x_8 = l_Lean_Syntax_getArg(x_6, x_7); -x_9 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____lambda__1___closed__1; +x_9 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____lambda__1___closed__1; x_10 = l_Lean_mkAtomFrom(x_8, x_9); x_11 = l_Lean_Syntax_setArg(x_6, x_7, x_10); x_12 = l_Lean_mkOptionalNode___closed__2; x_13 = lean_array_push(x_12, x_2); x_14 = lean_box(2); -x_15 = l_Lean_nullKind; +x_15 = l_Lean_mkNullNode___closed__2; x_16 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_16, 0, x_14); lean_ctor_set(x_16, 1, x_15); @@ -33732,7 +34368,7 @@ lean_ctor_set(x_19, 1, x_4); return x_19; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__1() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__1() { _start: { lean_object* x_1; @@ -33740,22 +34376,22 @@ x_1 = lean_mk_string_from_bytes("Lean.Meta.DSimp.Config", 22); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__2() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__1; +x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__1; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__3() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__1; +x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__2; +x_3 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__2; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -33763,7 +34399,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__4() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__4() { _start: { lean_object* x_1; @@ -33771,51 +34407,51 @@ x_1 = lean_mk_string_from_bytes("DSimp", 5); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__24; -x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__4; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__5; -x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__8; +x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__5; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__8; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__7() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__6; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__6; 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_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__8() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__7; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__7; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; @@ -33857,17 +34493,17 @@ lean_inc(x_11); x_20 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_20, 0, x_11); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__25; +x_21 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__23; lean_inc(x_11); x_22 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_22, 0, x_11); lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__41; +x_23 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__39; lean_inc(x_13); lean_inc(x_14); x_24 = l_Lean_addMacroScope(x_14, x_23, x_13); x_25 = lean_box(0); -x_26 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__40; +x_26 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__38; lean_inc(x_11); x_27 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_27, 0, x_11); @@ -33879,12 +34515,12 @@ lean_inc(x_11); x_29 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_29, 0, x_11); lean_ctor_set(x_29, 1, x_28); -x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__6; +x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__6; lean_inc(x_13); lean_inc(x_14); x_31 = l_Lean_addMacroScope(x_14, x_30, x_13); -x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__3; -x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__8; +x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__3; +x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__8; lean_inc(x_11); x_34 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_34, 0, x_11); @@ -33895,14 +34531,14 @@ x_35 = l_Lean_Syntax_mkApp___closed__3; x_36 = lean_array_push(x_35, x_29); x_37 = lean_array_push(x_36, x_34); x_38 = lean_box(2); -x_39 = l_Lean_Syntax_expandInterpolatedStr___closed__5; +x_39 = l_Lean_TSyntax_expandInterpolatedStr___closed__5; x_40 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_40, 0, x_38); lean_ctor_set(x_40, 1, x_39); lean_ctor_set(x_40, 2, x_37); x_41 = l_Lean_mkOptionalNode___closed__2; x_42 = lean_array_push(x_41, x_40); -x_43 = l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__2; +x_43 = l_Lean_mkNullNode___closed__2; x_44 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_44, 0, x_38); lean_ctor_set(x_44, 1, x_43); @@ -33919,13 +34555,13 @@ lean_inc(x_11); x_49 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_49, 0, x_11); lean_ctor_set(x_49, 1, x_48); -x_50 = l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__4; +x_50 = l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__4; lean_inc(x_16); x_51 = lean_array_push(x_50, x_16); x_52 = lean_array_push(x_51, x_47); lean_inc(x_49); x_53 = lean_array_push(x_52, x_49); -x_54 = l_Lean_Syntax_expandInterpolatedStr___closed__3; +x_54 = l_Lean_TSyntax_expandInterpolatedStr___closed__3; x_55 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_55, 0, x_38); lean_ctor_set(x_55, 1, x_54); @@ -33935,7 +34571,7 @@ x_57 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_57, 0, x_38); lean_ctor_set(x_57, 1, x_43); lean_ctor_set(x_57, 2, x_56); -x_58 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__31; +x_58 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__29; lean_inc(x_11); x_59 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_59, 0, x_11); @@ -33950,7 +34586,7 @@ x_63 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_63, 0, x_38); lean_ctor_set(x_63, 1, x_43); lean_ctor_set(x_63, 2, x_62); -x_64 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__49; +x_64 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__47; lean_inc(x_11); x_65 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_65, 0, x_11); @@ -33961,11 +34597,11 @@ x_68 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_68, 0, x_38); lean_ctor_set(x_68, 1, x_43); lean_ctor_set(x_68, 2, x_67); -x_69 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__14; +x_69 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__14; lean_inc(x_13); lean_inc(x_14); x_70 = l_Lean_addMacroScope(x_14, x_69, x_13); -x_71 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__13; +x_71 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__13; lean_inc(x_11); x_72 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_72, 0, x_11); @@ -33980,10 +34616,10 @@ x_77 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_77, 0, x_38); lean_ctor_set(x_77, 1, x_76); lean_ctor_set(x_77, 2, x_75); -x_78 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__17; +x_78 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__17; x_79 = l_Lean_addMacroScope(x_14, x_78, x_13); -x_80 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__16; -x_81 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__19; +x_80 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__16; +x_81 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__19; lean_inc(x_11); x_82 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_82, 0, x_11); @@ -34014,7 +34650,7 @@ x_93 = lean_array_push(x_92, x_61); lean_inc(x_93); x_94 = lean_array_push(x_93, x_68); x_95 = lean_array_push(x_94, x_89); -x_96 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__20; +x_96 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__20; x_97 = lean_array_push(x_95, x_96); x_98 = lean_array_push(x_97, x_74); lean_inc(x_91); @@ -34027,14 +34663,14 @@ lean_ctor_set(x_101, 2, x_99); x_102 = lean_array_push(x_50, x_57); x_103 = lean_array_push(x_102, x_59); x_104 = lean_array_push(x_103, x_101); -x_105 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__2; +x_105 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__2; x_106 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_106, 0, x_38); lean_ctor_set(x_106, 1, x_105); lean_ctor_set(x_106, 2, x_104); x_107 = lean_array_push(x_35, x_22); x_108 = lean_array_push(x_107, x_106); -x_109 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__1; +x_109 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__1; x_110 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_110, 0, x_38); lean_ctor_set(x_110, 1, x_109); @@ -34042,7 +34678,7 @@ lean_ctor_set(x_110, 2, x_108); x_111 = lean_array_push(x_35, x_61); lean_inc(x_91); x_112 = lean_array_push(x_111, x_91); -x_113 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__68; +x_113 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__65; x_114 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_114, 0, x_38); lean_ctor_set(x_114, 1, x_113); @@ -34058,7 +34694,7 @@ lean_ctor_set(x_120, 1, x_100); lean_ctor_set(x_120, 2, x_119); x_121 = lean_array_push(x_35, x_114); x_122 = lean_array_push(x_121, x_120); -x_123 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__66; +x_123 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__63; x_124 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_124, 0, x_38); lean_ctor_set(x_124, 1, x_123); @@ -34085,7 +34721,7 @@ x_137 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_137, 0, x_38); lean_ctor_set(x_137, 1, x_8); lean_ctor_set(x_137, 2, x_136); -x_138 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____lambda__1(x_1, x_137, x_2, x_12); +x_138 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____lambda__1(x_1, x_137, x_2, x_12); lean_dec(x_2); return x_138; } @@ -34121,17 +34757,17 @@ lean_inc(x_142); x_151 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_151, 0, x_142); lean_ctor_set(x_151, 1, x_150); -x_152 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__25; +x_152 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__23; lean_inc(x_142); x_153 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_153, 0, x_142); lean_ctor_set(x_153, 1, x_152); -x_154 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__41; +x_154 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__39; lean_inc(x_144); lean_inc(x_145); x_155 = l_Lean_addMacroScope(x_145, x_154, x_144); x_156 = lean_box(0); -x_157 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__40; +x_157 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__38; lean_inc(x_142); x_158 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_158, 0, x_142); @@ -34143,12 +34779,12 @@ lean_inc(x_142); x_160 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_160, 0, x_142); lean_ctor_set(x_160, 1, x_159); -x_161 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__6; +x_161 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__6; lean_inc(x_144); lean_inc(x_145); x_162 = l_Lean_addMacroScope(x_145, x_161, x_144); -x_163 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__3; -x_164 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__8; +x_163 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__3; +x_164 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__8; lean_inc(x_142); x_165 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_165, 0, x_142); @@ -34159,14 +34795,14 @@ x_166 = l_Lean_Syntax_mkApp___closed__3; x_167 = lean_array_push(x_166, x_160); x_168 = lean_array_push(x_167, x_165); x_169 = lean_box(2); -x_170 = l_Lean_Syntax_expandInterpolatedStr___closed__5; +x_170 = l_Lean_TSyntax_expandInterpolatedStr___closed__5; x_171 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_171, 0, x_169); lean_ctor_set(x_171, 1, x_170); lean_ctor_set(x_171, 2, x_168); x_172 = l_Lean_mkOptionalNode___closed__2; x_173 = lean_array_push(x_172, x_171); -x_174 = l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__2; +x_174 = l_Lean_mkNullNode___closed__2; x_175 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_175, 0, x_169); lean_ctor_set(x_175, 1, x_174); @@ -34183,13 +34819,13 @@ lean_inc(x_142); x_180 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_180, 0, x_142); lean_ctor_set(x_180, 1, x_179); -x_181 = l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__4; +x_181 = l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__4; lean_inc(x_147); x_182 = lean_array_push(x_181, x_147); x_183 = lean_array_push(x_182, x_178); lean_inc(x_180); x_184 = lean_array_push(x_183, x_180); -x_185 = l_Lean_Syntax_expandInterpolatedStr___closed__3; +x_185 = l_Lean_TSyntax_expandInterpolatedStr___closed__3; x_186 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_186, 0, x_169); lean_ctor_set(x_186, 1, x_185); @@ -34199,7 +34835,7 @@ x_188 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_188, 0, x_169); lean_ctor_set(x_188, 1, x_174); lean_ctor_set(x_188, 2, x_187); -x_189 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__31; +x_189 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__29; lean_inc(x_142); x_190 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_190, 0, x_142); @@ -34214,7 +34850,7 @@ x_194 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_194, 0, x_169); lean_ctor_set(x_194, 1, x_174); lean_ctor_set(x_194, 2, x_193); -x_195 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__49; +x_195 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__47; lean_inc(x_142); x_196 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_196, 0, x_142); @@ -34225,11 +34861,11 @@ x_199 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_199, 0, x_169); lean_ctor_set(x_199, 1, x_174); lean_ctor_set(x_199, 2, x_198); -x_200 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__14; +x_200 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__14; lean_inc(x_144); lean_inc(x_145); x_201 = l_Lean_addMacroScope(x_145, x_200, x_144); -x_202 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__13; +x_202 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__13; lean_inc(x_142); x_203 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_203, 0, x_142); @@ -34244,10 +34880,10 @@ x_208 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_208, 0, x_169); lean_ctor_set(x_208, 1, x_207); lean_ctor_set(x_208, 2, x_206); -x_209 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__17; +x_209 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__17; x_210 = l_Lean_addMacroScope(x_145, x_209, x_144); -x_211 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__16; -x_212 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__19; +x_211 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__16; +x_212 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__19; lean_inc(x_142); x_213 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_213, 0, x_142); @@ -34276,7 +34912,7 @@ x_223 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Ta x_224 = lean_array_push(x_223, x_192); x_225 = lean_array_push(x_224, x_199); x_226 = lean_array_push(x_225, x_220); -x_227 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__20; +x_227 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__20; x_228 = lean_array_push(x_226, x_227); x_229 = lean_array_push(x_228, x_205); x_230 = lean_array_push(x_229, x_222); @@ -34288,14 +34924,14 @@ lean_ctor_set(x_232, 2, x_230); x_233 = lean_array_push(x_181, x_188); x_234 = lean_array_push(x_233, x_190); x_235 = lean_array_push(x_234, x_232); -x_236 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__2; +x_236 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__2; x_237 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_237, 0, x_169); lean_ctor_set(x_237, 1, x_236); lean_ctor_set(x_237, 2, x_235); x_238 = lean_array_push(x_166, x_153); x_239 = lean_array_push(x_238, x_237); -x_240 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__1; +x_240 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__1; x_241 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_241, 0, x_169); lean_ctor_set(x_241, 1, x_240); @@ -34322,22 +34958,23 @@ x_254 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_254, 0, x_169); lean_ctor_set(x_254, 1, x_8); lean_ctor_set(x_254, 2, x_253); -x_255 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____lambda__1(x_1, x_254, x_2, x_143); +x_255 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____lambda__1(x_1, x_254, x_2, x_143); lean_dec(x_2); return x_255; } } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____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_Tactic_expandSimp____x40_Init_Meta___hyg_18703____lambda__1(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____lambda__1(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } lean_object* initialize_Init_Data_Array_Basic(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Data_Option_BasicAux(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Init_Meta(uint8_t builtin, lean_object* w) { lean_object * res; @@ -34346,6 +34983,9 @@ _G_initialized = true; res = initialize_Init_Data_Array_Basic(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Init_Data_Option_BasicAux(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); l_Lean_version_major___closed__1 = _init_l_Lean_version_major___closed__1(); lean_mark_persistent(l_Lean_version_major___closed__1); l_Lean_version_major = _init_l_Lean_version_major(); @@ -34574,6 +35214,8 @@ l_Lean_instInhabitedNameGenerator___closed__1 = _init_l_Lean_instInhabitedNameGe lean_mark_persistent(l_Lean_instInhabitedNameGenerator___closed__1); l_Lean_instInhabitedNameGenerator = _init_l_Lean_instInhabitedNameGenerator(); lean_mark_persistent(l_Lean_instInhabitedNameGenerator); +l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSyntaxArray___closed__1 = _init_l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSyntaxArray___closed__1(); +lean_mark_persistent(l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSyntaxArray___closed__1); l_Lean_Syntax_instBEqSyntax___closed__1 = _init_l_Lean_Syntax_instBEqSyntax___closed__1(); lean_mark_persistent(l_Lean_Syntax_instBEqSyntax___closed__1); l_Lean_Syntax_instBEqSyntax = _init_l_Lean_Syntax_instBEqSyntax(); @@ -34588,6 +35230,14 @@ l_Lean_mkCIdentFrom___closed__1 = _init_l_Lean_mkCIdentFrom___closed__1(); lean_mark_persistent(l_Lean_mkCIdentFrom___closed__1); l_Lean_mkCIdentFrom___closed__2 = _init_l_Lean_mkCIdentFrom___closed__2(); lean_mark_persistent(l_Lean_mkCIdentFrom___closed__2); +l_Lean_mkNullNode___closed__1 = _init_l_Lean_mkNullNode___closed__1(); +lean_mark_persistent(l_Lean_mkNullNode___closed__1); +l_Lean_mkNullNode___closed__2 = _init_l_Lean_mkNullNode___closed__2(); +lean_mark_persistent(l_Lean_mkNullNode___closed__2); +l_Lean_mkGroupNode___closed__1 = _init_l_Lean_mkGroupNode___closed__1(); +lean_mark_persistent(l_Lean_mkGroupNode___closed__1); +l_Lean_mkGroupNode___closed__2 = _init_l_Lean_mkGroupNode___closed__2(); +lean_mark_persistent(l_Lean_mkGroupNode___closed__2); l_Array_forInUnsafe_loop___at_Lean_mkSepArray___spec__1___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_mkSepArray___spec__1___closed__1(); lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_mkSepArray___spec__1___closed__1); l_Lean_mkSepArray___closed__1 = _init_l_Lean_mkSepArray___closed__1(); @@ -34620,8 +35270,28 @@ l_Lean_Syntax_mkApp___closed__2 = _init_l_Lean_Syntax_mkApp___closed__2(); lean_mark_persistent(l_Lean_Syntax_mkApp___closed__2); l_Lean_Syntax_mkApp___closed__3 = _init_l_Lean_Syntax_mkApp___closed__3(); lean_mark_persistent(l_Lean_Syntax_mkApp___closed__3); +l_Lean_Syntax_mkStrLit___closed__1 = _init_l_Lean_Syntax_mkStrLit___closed__1(); +lean_mark_persistent(l_Lean_Syntax_mkStrLit___closed__1); +l_Lean_Syntax_mkStrLit___closed__2 = _init_l_Lean_Syntax_mkStrLit___closed__2(); +lean_mark_persistent(l_Lean_Syntax_mkStrLit___closed__2); +l_Lean_Syntax_mkNumLit___closed__1 = _init_l_Lean_Syntax_mkNumLit___closed__1(); +lean_mark_persistent(l_Lean_Syntax_mkNumLit___closed__1); +l_Lean_Syntax_mkNumLit___closed__2 = _init_l_Lean_Syntax_mkNumLit___closed__2(); +lean_mark_persistent(l_Lean_Syntax_mkNumLit___closed__2); +l_Lean_Syntax_mkScientificLit___closed__1 = _init_l_Lean_Syntax_mkScientificLit___closed__1(); +lean_mark_persistent(l_Lean_Syntax_mkScientificLit___closed__1); +l_Lean_Syntax_mkScientificLit___closed__2 = _init_l_Lean_Syntax_mkScientificLit___closed__2(); +lean_mark_persistent(l_Lean_Syntax_mkScientificLit___closed__2); +l_Lean_Syntax_mkNameLit___closed__1 = _init_l_Lean_Syntax_mkNameLit___closed__1(); +lean_mark_persistent(l_Lean_Syntax_mkNameLit___closed__1); +l_Lean_Syntax_mkNameLit___closed__2 = _init_l_Lean_Syntax_mkNameLit___closed__2(); +lean_mark_persistent(l_Lean_Syntax_mkNameLit___closed__2); l_Lean_Syntax_decodeNatLitVal_x3f___closed__1 = _init_l_Lean_Syntax_decodeNatLitVal_x3f___closed__1(); lean_mark_persistent(l_Lean_Syntax_decodeNatLitVal_x3f___closed__1); +l_Lean_Syntax_isFieldIdx_x3f___closed__1 = _init_l_Lean_Syntax_isFieldIdx_x3f___closed__1(); +lean_mark_persistent(l_Lean_Syntax_isFieldIdx_x3f___closed__1); +l_Lean_Syntax_isFieldIdx_x3f___closed__2 = _init_l_Lean_Syntax_isFieldIdx_x3f___closed__2(); +lean_mark_persistent(l_Lean_Syntax_isFieldIdx_x3f___closed__2); l_Lean_Syntax_decodeQuotedChar___boxed__const__1 = _init_l_Lean_Syntax_decodeQuotedChar___boxed__const__1(); lean_mark_persistent(l_Lean_Syntax_decodeQuotedChar___boxed__const__1); l_Lean_Syntax_decodeQuotedChar___boxed__const__2 = _init_l_Lean_Syntax_decodeQuotedChar___boxed__const__2(); @@ -34634,6 +35304,10 @@ l_Lean_Syntax_decodeQuotedChar___boxed__const__5 = _init_l_Lean_Syntax_decodeQuo lean_mark_persistent(l_Lean_Syntax_decodeQuotedChar___boxed__const__5); l_Lean_Syntax_decodeQuotedChar___boxed__const__6 = _init_l_Lean_Syntax_decodeQuotedChar___boxed__const__6(); lean_mark_persistent(l_Lean_Syntax_decodeQuotedChar___boxed__const__6); +l_Lean_Syntax_isCharLit_x3f___closed__1 = _init_l_Lean_Syntax_isCharLit_x3f___closed__1(); +lean_mark_persistent(l_Lean_Syntax_isCharLit_x3f___closed__1); +l_Lean_Syntax_isCharLit_x3f___closed__2 = _init_l_Lean_Syntax_isCharLit_x3f___closed__2(); +lean_mark_persistent(l_Lean_Syntax_isCharLit_x3f___closed__2); l___private_Init_Meta_0__Lean_Syntax_splitNameLitAux___closed__1 = _init_l___private_Init_Meta_0__Lean_Syntax_splitNameLitAux___closed__1(); lean_mark_persistent(l___private_Init_Meta_0__Lean_Syntax_splitNameLitAux___closed__1); l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___closed__1 = _init_l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___closed__1(); @@ -34644,64 +35318,72 @@ l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___closed__3 = _init_l_List lean_mark_persistent(l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___closed__3); l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___closed__4 = _init_l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___closed__4(); lean_mark_persistent(l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___closed__4); -l_Lean_instQuoteSyntax___closed__1 = _init_l_Lean_instQuoteSyntax___closed__1(); -lean_mark_persistent(l_Lean_instQuoteSyntax___closed__1); -l_Lean_instQuoteSyntax = _init_l_Lean_instQuoteSyntax(); -lean_mark_persistent(l_Lean_instQuoteSyntax); -l_Lean_instQuoteBool___closed__1 = _init_l_Lean_instQuoteBool___closed__1(); -lean_mark_persistent(l_Lean_instQuoteBool___closed__1); -l_Lean_instQuoteBool___closed__2 = _init_l_Lean_instQuoteBool___closed__2(); -lean_mark_persistent(l_Lean_instQuoteBool___closed__2); -l_Lean_instQuoteBool___closed__3 = _init_l_Lean_instQuoteBool___closed__3(); -lean_mark_persistent(l_Lean_instQuoteBool___closed__3); -l_Lean_instQuoteBool___closed__4 = _init_l_Lean_instQuoteBool___closed__4(); -lean_mark_persistent(l_Lean_instQuoteBool___closed__4); -l_Lean_instQuoteBool___closed__5 = _init_l_Lean_instQuoteBool___closed__5(); -lean_mark_persistent(l_Lean_instQuoteBool___closed__5); -l_Lean_instQuoteBool___closed__6 = _init_l_Lean_instQuoteBool___closed__6(); -lean_mark_persistent(l_Lean_instQuoteBool___closed__6); -l_Lean_instQuoteBool___closed__7 = _init_l_Lean_instQuoteBool___closed__7(); -lean_mark_persistent(l_Lean_instQuoteBool___closed__7); -l_Lean_instQuoteBool___closed__8 = _init_l_Lean_instQuoteBool___closed__8(); -lean_mark_persistent(l_Lean_instQuoteBool___closed__8); -l_Lean_instQuoteSubstring___closed__1 = _init_l_Lean_instQuoteSubstring___closed__1(); -lean_mark_persistent(l_Lean_instQuoteSubstring___closed__1); -l_Lean_instQuoteSubstring___closed__2 = _init_l_Lean_instQuoteSubstring___closed__2(); -lean_mark_persistent(l_Lean_instQuoteSubstring___closed__2); -l_Lean_instQuoteSubstring___closed__3 = _init_l_Lean_instQuoteSubstring___closed__3(); -lean_mark_persistent(l_Lean_instQuoteSubstring___closed__3); -l_Lean_instQuoteSubstring___closed__4 = _init_l_Lean_instQuoteSubstring___closed__4(); -lean_mark_persistent(l_Lean_instQuoteSubstring___closed__4); -l___private_Init_Meta_0__Lean_quoteNameMk___closed__1 = _init_l___private_Init_Meta_0__Lean_quoteNameMk___closed__1(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_quoteNameMk___closed__1); -l___private_Init_Meta_0__Lean_quoteNameMk___closed__2 = _init_l___private_Init_Meta_0__Lean_quoteNameMk___closed__2(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_quoteNameMk___closed__2); -l___private_Init_Meta_0__Lean_quoteNameMk___closed__3 = _init_l___private_Init_Meta_0__Lean_quoteNameMk___closed__3(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_quoteNameMk___closed__3); -l___private_Init_Meta_0__Lean_quoteNameMk___closed__4 = _init_l___private_Init_Meta_0__Lean_quoteNameMk___closed__4(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_quoteNameMk___closed__4); -l___private_Init_Meta_0__Lean_quoteNameMk___closed__5 = _init_l___private_Init_Meta_0__Lean_quoteNameMk___closed__5(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_quoteNameMk___closed__5); -l___private_Init_Meta_0__Lean_quoteNameMk___closed__6 = _init_l___private_Init_Meta_0__Lean_quoteNameMk___closed__6(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_quoteNameMk___closed__6); -l___private_Init_Meta_0__Lean_quoteNameMk___closed__7 = _init_l___private_Init_Meta_0__Lean_quoteNameMk___closed__7(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_quoteNameMk___closed__7); -l___private_Init_Meta_0__Lean_quoteNameMk___closed__8 = _init_l___private_Init_Meta_0__Lean_quoteNameMk___closed__8(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_quoteNameMk___closed__8); -l___private_Init_Meta_0__Lean_quoteNameMk___closed__9 = _init_l___private_Init_Meta_0__Lean_quoteNameMk___closed__9(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_quoteNameMk___closed__9); -l_Lean_instQuoteName___closed__1 = _init_l_Lean_instQuoteName___closed__1(); -lean_mark_persistent(l_Lean_instQuoteName___closed__1); -l_Lean_instQuoteName___closed__2 = _init_l_Lean_instQuoteName___closed__2(); -lean_mark_persistent(l_Lean_instQuoteName___closed__2); -l_Lean_instQuoteProd___rarg___closed__1 = _init_l_Lean_instQuoteProd___rarg___closed__1(); -lean_mark_persistent(l_Lean_instQuoteProd___rarg___closed__1); -l_Lean_instQuoteProd___rarg___closed__2 = _init_l_Lean_instQuoteProd___rarg___closed__2(); -lean_mark_persistent(l_Lean_instQuoteProd___rarg___closed__2); -l_Lean_instQuoteProd___rarg___closed__3 = _init_l_Lean_instQuoteProd___rarg___closed__3(); -lean_mark_persistent(l_Lean_instQuoteProd___rarg___closed__3); -l_Lean_instQuoteProd___rarg___closed__4 = _init_l_Lean_instQuoteProd___rarg___closed__4(); -lean_mark_persistent(l_Lean_instQuoteProd___rarg___closed__4); +l_Lean_TSyntax_getNat___closed__1 = _init_l_Lean_TSyntax_getNat___closed__1(); +lean_mark_persistent(l_Lean_TSyntax_getNat___closed__1); +l_Lean_TSyntax_getNat___closed__2 = _init_l_Lean_TSyntax_getNat___closed__2(); +lean_mark_persistent(l_Lean_TSyntax_getNat___closed__2); +l_Lean_TSyntax_getNat___closed__3 = _init_l_Lean_TSyntax_getNat___closed__3(); +lean_mark_persistent(l_Lean_TSyntax_getNat___closed__3); +l_Lean_TSyntax_getNat___closed__4 = _init_l_Lean_TSyntax_getNat___closed__4(); +lean_mark_persistent(l_Lean_TSyntax_getNat___closed__4); +l_Lean_instQuoteTermMkStrAnonymous___closed__1 = _init_l_Lean_instQuoteTermMkStrAnonymous___closed__1(); +lean_mark_persistent(l_Lean_instQuoteTermMkStrAnonymous___closed__1); +l_Lean_instQuoteTermMkStrAnonymous = _init_l_Lean_instQuoteTermMkStrAnonymous(); +lean_mark_persistent(l_Lean_instQuoteTermMkStrAnonymous); +l_Lean_instQuoteBoolMkStrAnonymous___closed__1 = _init_l_Lean_instQuoteBoolMkStrAnonymous___closed__1(); +lean_mark_persistent(l_Lean_instQuoteBoolMkStrAnonymous___closed__1); +l_Lean_instQuoteBoolMkStrAnonymous___closed__2 = _init_l_Lean_instQuoteBoolMkStrAnonymous___closed__2(); +lean_mark_persistent(l_Lean_instQuoteBoolMkStrAnonymous___closed__2); +l_Lean_instQuoteBoolMkStrAnonymous___closed__3 = _init_l_Lean_instQuoteBoolMkStrAnonymous___closed__3(); +lean_mark_persistent(l_Lean_instQuoteBoolMkStrAnonymous___closed__3); +l_Lean_instQuoteBoolMkStrAnonymous___closed__4 = _init_l_Lean_instQuoteBoolMkStrAnonymous___closed__4(); +lean_mark_persistent(l_Lean_instQuoteBoolMkStrAnonymous___closed__4); +l_Lean_instQuoteBoolMkStrAnonymous___closed__5 = _init_l_Lean_instQuoteBoolMkStrAnonymous___closed__5(); +lean_mark_persistent(l_Lean_instQuoteBoolMkStrAnonymous___closed__5); +l_Lean_instQuoteBoolMkStrAnonymous___closed__6 = _init_l_Lean_instQuoteBoolMkStrAnonymous___closed__6(); +lean_mark_persistent(l_Lean_instQuoteBoolMkStrAnonymous___closed__6); +l_Lean_instQuoteBoolMkStrAnonymous___closed__7 = _init_l_Lean_instQuoteBoolMkStrAnonymous___closed__7(); +lean_mark_persistent(l_Lean_instQuoteBoolMkStrAnonymous___closed__7); +l_Lean_instQuoteBoolMkStrAnonymous___closed__8 = _init_l_Lean_instQuoteBoolMkStrAnonymous___closed__8(); +lean_mark_persistent(l_Lean_instQuoteBoolMkStrAnonymous___closed__8); +l_Lean_instQuoteSubstringMkStrAnonymous___closed__1 = _init_l_Lean_instQuoteSubstringMkStrAnonymous___closed__1(); +lean_mark_persistent(l_Lean_instQuoteSubstringMkStrAnonymous___closed__1); +l_Lean_instQuoteSubstringMkStrAnonymous___closed__2 = _init_l_Lean_instQuoteSubstringMkStrAnonymous___closed__2(); +lean_mark_persistent(l_Lean_instQuoteSubstringMkStrAnonymous___closed__2); +l_Lean_instQuoteSubstringMkStrAnonymous___closed__3 = _init_l_Lean_instQuoteSubstringMkStrAnonymous___closed__3(); +lean_mark_persistent(l_Lean_instQuoteSubstringMkStrAnonymous___closed__3); +l_Lean_instQuoteSubstringMkStrAnonymous___closed__4 = _init_l_Lean_instQuoteSubstringMkStrAnonymous___closed__4(); +lean_mark_persistent(l_Lean_instQuoteSubstringMkStrAnonymous___closed__4); +l_Lean_quoteNameMk___closed__1 = _init_l_Lean_quoteNameMk___closed__1(); +lean_mark_persistent(l_Lean_quoteNameMk___closed__1); +l_Lean_quoteNameMk___closed__2 = _init_l_Lean_quoteNameMk___closed__2(); +lean_mark_persistent(l_Lean_quoteNameMk___closed__2); +l_Lean_quoteNameMk___closed__3 = _init_l_Lean_quoteNameMk___closed__3(); +lean_mark_persistent(l_Lean_quoteNameMk___closed__3); +l_Lean_quoteNameMk___closed__4 = _init_l_Lean_quoteNameMk___closed__4(); +lean_mark_persistent(l_Lean_quoteNameMk___closed__4); +l_Lean_quoteNameMk___closed__5 = _init_l_Lean_quoteNameMk___closed__5(); +lean_mark_persistent(l_Lean_quoteNameMk___closed__5); +l_Lean_quoteNameMk___closed__6 = _init_l_Lean_quoteNameMk___closed__6(); +lean_mark_persistent(l_Lean_quoteNameMk___closed__6); +l_Lean_quoteNameMk___closed__7 = _init_l_Lean_quoteNameMk___closed__7(); +lean_mark_persistent(l_Lean_quoteNameMk___closed__7); +l_Lean_quoteNameMk___closed__8 = _init_l_Lean_quoteNameMk___closed__8(); +lean_mark_persistent(l_Lean_quoteNameMk___closed__8); +l_Lean_quoteNameMk___closed__9 = _init_l_Lean_quoteNameMk___closed__9(); +lean_mark_persistent(l_Lean_quoteNameMk___closed__9); +l_Lean_instQuoteNameMkStrAnonymous___closed__1 = _init_l_Lean_instQuoteNameMkStrAnonymous___closed__1(); +lean_mark_persistent(l_Lean_instQuoteNameMkStrAnonymous___closed__1); +l_Lean_instQuoteNameMkStrAnonymous___closed__2 = _init_l_Lean_instQuoteNameMkStrAnonymous___closed__2(); +lean_mark_persistent(l_Lean_instQuoteNameMkStrAnonymous___closed__2); +l_Lean_instQuoteProdMkStrAnonymous___rarg___closed__1 = _init_l_Lean_instQuoteProdMkStrAnonymous___rarg___closed__1(); +lean_mark_persistent(l_Lean_instQuoteProdMkStrAnonymous___rarg___closed__1); +l_Lean_instQuoteProdMkStrAnonymous___rarg___closed__2 = _init_l_Lean_instQuoteProdMkStrAnonymous___rarg___closed__2(); +lean_mark_persistent(l_Lean_instQuoteProdMkStrAnonymous___rarg___closed__2); +l_Lean_instQuoteProdMkStrAnonymous___rarg___closed__3 = _init_l_Lean_instQuoteProdMkStrAnonymous___rarg___closed__3(); +lean_mark_persistent(l_Lean_instQuoteProdMkStrAnonymous___rarg___closed__3); +l_Lean_instQuoteProdMkStrAnonymous___rarg___closed__4 = _init_l_Lean_instQuoteProdMkStrAnonymous___rarg___closed__4(); +lean_mark_persistent(l_Lean_instQuoteProdMkStrAnonymous___rarg___closed__4); l___private_Init_Meta_0__Lean_quoteList___rarg___closed__1 = _init_l___private_Init_Meta_0__Lean_quoteList___rarg___closed__1(); lean_mark_persistent(l___private_Init_Meta_0__Lean_quoteList___rarg___closed__1); l___private_Init_Meta_0__Lean_quoteList___rarg___closed__2 = _init_l___private_Init_Meta_0__Lean_quoteList___rarg___closed__2(); @@ -34716,30 +35398,26 @@ l___private_Init_Meta_0__Lean_quoteList___rarg___closed__6 = _init_l___private_I lean_mark_persistent(l___private_Init_Meta_0__Lean_quoteList___rarg___closed__6); l___private_Init_Meta_0__Lean_quoteList___rarg___closed__7 = _init_l___private_Init_Meta_0__Lean_quoteList___rarg___closed__7(); lean_mark_persistent(l___private_Init_Meta_0__Lean_quoteList___rarg___closed__7); -l_Lean_instQuoteArray___rarg___closed__1 = _init_l_Lean_instQuoteArray___rarg___closed__1(); -lean_mark_persistent(l_Lean_instQuoteArray___rarg___closed__1); -l_Lean_instQuoteArray___rarg___closed__2 = _init_l_Lean_instQuoteArray___rarg___closed__2(); -lean_mark_persistent(l_Lean_instQuoteArray___rarg___closed__2); -l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__1 = _init_l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__1(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__1); -l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__2 = _init_l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__2(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__2); -l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__3 = _init_l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__3(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__3); -l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__4 = _init_l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__4(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__4); -l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__5 = _init_l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__5(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__5); -l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__6 = _init_l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__6(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__6); -l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__7 = _init_l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__7(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_quoteOption___rarg___closed__7); +l_Lean_instQuoteArrayMkStrAnonymous___rarg___closed__1 = _init_l_Lean_instQuoteArrayMkStrAnonymous___rarg___closed__1(); +lean_mark_persistent(l_Lean_instQuoteArrayMkStrAnonymous___rarg___closed__1); +l_Lean_instQuoteArrayMkStrAnonymous___rarg___closed__2 = _init_l_Lean_instQuoteArrayMkStrAnonymous___rarg___closed__2(); +lean_mark_persistent(l_Lean_instQuoteArrayMkStrAnonymous___rarg___closed__2); +l_Lean_Option_hasQuote___rarg___closed__1 = _init_l_Lean_Option_hasQuote___rarg___closed__1(); +lean_mark_persistent(l_Lean_Option_hasQuote___rarg___closed__1); +l_Lean_Option_hasQuote___rarg___closed__2 = _init_l_Lean_Option_hasQuote___rarg___closed__2(); +lean_mark_persistent(l_Lean_Option_hasQuote___rarg___closed__2); +l_Lean_Option_hasQuote___rarg___closed__3 = _init_l_Lean_Option_hasQuote___rarg___closed__3(); +lean_mark_persistent(l_Lean_Option_hasQuote___rarg___closed__3); +l_Lean_Option_hasQuote___rarg___closed__4 = _init_l_Lean_Option_hasQuote___rarg___closed__4(); +lean_mark_persistent(l_Lean_Option_hasQuote___rarg___closed__4); +l_Lean_Option_hasQuote___rarg___closed__5 = _init_l_Lean_Option_hasQuote___rarg___closed__5(); +lean_mark_persistent(l_Lean_Option_hasQuote___rarg___closed__5); +l_Lean_Option_hasQuote___rarg___closed__6 = _init_l_Lean_Option_hasQuote___rarg___closed__6(); +lean_mark_persistent(l_Lean_Option_hasQuote___rarg___closed__6); +l_Lean_Option_hasQuote___rarg___closed__7 = _init_l_Lean_Option_hasQuote___rarg___closed__7(); +lean_mark_persistent(l_Lean_Option_hasQuote___rarg___closed__7); l_Lean_evalPrec___closed__1 = _init_l_Lean_evalPrec___closed__1(); lean_mark_persistent(l_Lean_evalPrec___closed__1); -l_Lean_evalPrec___closed__2 = _init_l_Lean_evalPrec___closed__2(); -lean_mark_persistent(l_Lean_evalPrec___closed__2); -l_Lean_evalPrec___closed__3 = _init_l_Lean_evalPrec___closed__3(); -lean_mark_persistent(l_Lean_evalPrec___closed__3); l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__addPrec__1___closed__1 = _init_l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__addPrec__1___closed__1(); lean_mark_persistent(l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__addPrec__1___closed__1); l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__addPrec__1___closed__2 = _init_l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__addPrec__1___closed__2(); @@ -34808,36 +35486,48 @@ l_Lean_termEval__prio__ = _init_l_Lean_termEval__prio__(); lean_mark_persistent(l_Lean_termEval__prio__); l_Array_getSepElems___rarg___closed__1 = _init_l_Array_getSepElems___rarg___closed__1(); lean_mark_persistent(l_Array_getSepElems___rarg___closed__1); -l_Lean_Syntax_SepArray_instCoeTailSepArrayArraySyntax___closed__1 = _init_l_Lean_Syntax_SepArray_instCoeTailSepArrayArraySyntax___closed__1(); -lean_mark_persistent(l_Lean_Syntax_SepArray_instCoeTailSepArrayArraySyntax___closed__1); +l_Lean_Syntax_instCoeTailSepArrayArraySyntax___closed__1 = _init_l_Lean_Syntax_instCoeTailSepArrayArraySyntax___closed__1(); +lean_mark_persistent(l_Lean_Syntax_instCoeTailSepArrayArraySyntax___closed__1); +l_Lean_Syntax_instCoeTSepArrayTSyntaxArray___closed__1 = _init_l_Lean_Syntax_instCoeTSepArrayTSyntaxArray___closed__1(); +lean_mark_persistent(l_Lean_Syntax_instCoeTSepArrayTSyntaxArray___closed__1); +l_Lean_Syntax_instCoeTSyntaxArrayArraySyntax___closed__1 = _init_l_Lean_Syntax_instCoeTSyntaxArrayArraySyntax___closed__1(); +lean_mark_persistent(l_Lean_Syntax_instCoeTSyntaxArrayArraySyntax___closed__1); +l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___closed__1 = _init_l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___closed__1(); +lean_mark_persistent(l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___closed__1); +l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___closed__2 = _init_l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___closed__2(); +lean_mark_persistent(l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___closed__2); +l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___closed__3 = _init_l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___closed__3(); +lean_mark_persistent(l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___closed__3); +l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___closed__4 = _init_l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___closed__4(); +lean_mark_persistent(l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___closed__4); l___private_Init_Meta_0__Lean_Syntax_decodeInterpStrQuotedChar___boxed__const__1 = _init_l___private_Init_Meta_0__Lean_Syntax_decodeInterpStrQuotedChar___boxed__const__1(); lean_mark_persistent(l___private_Init_Meta_0__Lean_Syntax_decodeInterpStrQuotedChar___boxed__const__1); -l_Array_forInUnsafe_loop___at_Lean_Syntax_expandInterpolatedStrChunks___spec__1___lambda__2___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Syntax_expandInterpolatedStrChunks___spec__1___lambda__2___closed__1(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Syntax_expandInterpolatedStrChunks___spec__1___lambda__2___closed__1); -l_Lean_Syntax_expandInterpolatedStrChunks___closed__1 = _init_l_Lean_Syntax_expandInterpolatedStrChunks___closed__1(); -lean_mark_persistent(l_Lean_Syntax_expandInterpolatedStrChunks___closed__1); -l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__1 = _init_l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__1); -l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__2 = _init_l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__2(); -lean_mark_persistent(l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__2); -l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__3 = _init_l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__3(); -lean_mark_persistent(l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__3); -l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__4 = _init_l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__4(); -lean_mark_persistent(l_Lean_Syntax_expandInterpolatedStr___lambda__1___closed__4); -l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__1 = _init_l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__1); -l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__2 = _init_l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__2(); -lean_mark_persistent(l_Lean_Syntax_expandInterpolatedStr___lambda__2___closed__2); -l_Lean_Syntax_expandInterpolatedStr___closed__1 = _init_l_Lean_Syntax_expandInterpolatedStr___closed__1(); -lean_mark_persistent(l_Lean_Syntax_expandInterpolatedStr___closed__1); -l_Lean_Syntax_expandInterpolatedStr___closed__2 = _init_l_Lean_Syntax_expandInterpolatedStr___closed__2(); -lean_mark_persistent(l_Lean_Syntax_expandInterpolatedStr___closed__2); -l_Lean_Syntax_expandInterpolatedStr___closed__3 = _init_l_Lean_Syntax_expandInterpolatedStr___closed__3(); -lean_mark_persistent(l_Lean_Syntax_expandInterpolatedStr___closed__3); -l_Lean_Syntax_expandInterpolatedStr___closed__4 = _init_l_Lean_Syntax_expandInterpolatedStr___closed__4(); -lean_mark_persistent(l_Lean_Syntax_expandInterpolatedStr___closed__4); -l_Lean_Syntax_expandInterpolatedStr___closed__5 = _init_l_Lean_Syntax_expandInterpolatedStr___closed__5(); -lean_mark_persistent(l_Lean_Syntax_expandInterpolatedStr___closed__5); +l_Lean_Syntax_isInterpolatedStrLit_x3f___closed__1 = _init_l_Lean_Syntax_isInterpolatedStrLit_x3f___closed__1(); +lean_mark_persistent(l_Lean_Syntax_isInterpolatedStrLit_x3f___closed__1); +l_Lean_Syntax_isInterpolatedStrLit_x3f___closed__2 = _init_l_Lean_Syntax_isInterpolatedStrLit_x3f___closed__2(); +lean_mark_persistent(l_Lean_Syntax_isInterpolatedStrLit_x3f___closed__2); +l_Array_forInUnsafe_loop___at_Lean_TSyntax_expandInterpolatedStrChunks___spec__1___lambda__2___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_TSyntax_expandInterpolatedStrChunks___spec__1___lambda__2___closed__1(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_TSyntax_expandInterpolatedStrChunks___spec__1___lambda__2___closed__1); +l_Lean_TSyntax_expandInterpolatedStrChunks___closed__1 = _init_l_Lean_TSyntax_expandInterpolatedStrChunks___closed__1(); +lean_mark_persistent(l_Lean_TSyntax_expandInterpolatedStrChunks___closed__1); +l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__1 = _init_l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__1); +l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__2 = _init_l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__2); +l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__3 = _init_l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__3(); +lean_mark_persistent(l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__3); +l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__4 = _init_l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__4(); +lean_mark_persistent(l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__4); +l_Lean_TSyntax_expandInterpolatedStr___closed__1 = _init_l_Lean_TSyntax_expandInterpolatedStr___closed__1(); +lean_mark_persistent(l_Lean_TSyntax_expandInterpolatedStr___closed__1); +l_Lean_TSyntax_expandInterpolatedStr___closed__2 = _init_l_Lean_TSyntax_expandInterpolatedStr___closed__2(); +lean_mark_persistent(l_Lean_TSyntax_expandInterpolatedStr___closed__2); +l_Lean_TSyntax_expandInterpolatedStr___closed__3 = _init_l_Lean_TSyntax_expandInterpolatedStr___closed__3(); +lean_mark_persistent(l_Lean_TSyntax_expandInterpolatedStr___closed__3); +l_Lean_TSyntax_expandInterpolatedStr___closed__4 = _init_l_Lean_TSyntax_expandInterpolatedStr___closed__4(); +lean_mark_persistent(l_Lean_TSyntax_expandInterpolatedStr___closed__4); +l_Lean_TSyntax_expandInterpolatedStr___closed__5 = _init_l_Lean_TSyntax_expandInterpolatedStr___closed__5(); +lean_mark_persistent(l_Lean_TSyntax_expandInterpolatedStr___closed__5); l_Lean_Meta_TransparencyMode_noConfusion___rarg___closed__1 = _init_l_Lean_Meta_TransparencyMode_noConfusion___rarg___closed__1(); lean_mark_persistent(l_Lean_Meta_TransparencyMode_noConfusion___rarg___closed__1); l_Lean_Meta_instInhabitedTransparencyMode = _init_l_Lean_Meta_instInhabitedTransparencyMode(); @@ -34845,54 +35535,54 @@ l_Lean_Meta_instBEqTransparencyMode___closed__1 = _init_l_Lean_Meta_instBEqTrans lean_mark_persistent(l_Lean_Meta_instBEqTransparencyMode___closed__1); l_Lean_Meta_instBEqTransparencyMode = _init_l_Lean_Meta_instBEqTransparencyMode(); lean_mark_persistent(l_Lean_Meta_instBEqTransparencyMode); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__1 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__1(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__1); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__2 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__2(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__2); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__3 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__3(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__3); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__4 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__4(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__4); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__5 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__5(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__5); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__6 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__6(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__6); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__7 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__7(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__7); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__8 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__8(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__8); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__9 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__9(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__9); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__10 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__10(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__10); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__11 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__11(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__11); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__12 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__12(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__12); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__13 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__13(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__13); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__14 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__14(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__14); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__15 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__15(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__15); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__16 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__16(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__16); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__17 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__17(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__17); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__18 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__18(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__18); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__19 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__19(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__19); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__20 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__20(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__20); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__21 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__21(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__21); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__22 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__22(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__22); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__23 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__23(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__23); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__24 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__24(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_8583____closed__24); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__1 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__1(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__1); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__2 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__2(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__2); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__3 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__3(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__3); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__4 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__4(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__4); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__5 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__5(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__5); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__6 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__6(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__6); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__7 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__7(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__7); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__8 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__8(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__8); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__9 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__9(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__9); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__10 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__10(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__10); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__11 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__11(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__11); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__12 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__12(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__12); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__13 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__13(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__13); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__14 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__14(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__14); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__15 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__15(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__15); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__16 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__16(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__16); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__17 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__17(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__17); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__18 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__18(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__18); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__19 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__19(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__19); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__20 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__20(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__20); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__21 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__21(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__21); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__22 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__22(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__22); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__23 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__23(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__23); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__24 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__24(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_9558____closed__24); l_Lean_Meta_instReprTransparencyMode___closed__1 = _init_l_Lean_Meta_instReprTransparencyMode___closed__1(); lean_mark_persistent(l_Lean_Meta_instReprTransparencyMode___closed__1); l_Lean_Meta_instReprTransparencyMode = _init_l_Lean_Meta_instReprTransparencyMode(); @@ -34902,42 +35592,42 @@ l_Lean_Meta_instBEqEtaStructMode___closed__1 = _init_l_Lean_Meta_instBEqEtaStruc lean_mark_persistent(l_Lean_Meta_instBEqEtaStructMode___closed__1); l_Lean_Meta_instBEqEtaStructMode = _init_l_Lean_Meta_instBEqEtaStructMode(); lean_mark_persistent(l_Lean_Meta_instBEqEtaStructMode); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__1 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__1(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__1); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__2 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__2(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__2); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__3 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__3(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__3); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__4 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__4(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__4); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__5 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__5(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__5); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__6 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__6(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__6); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__7 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__7(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__7); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__8 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__8(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__8); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__9 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__9(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__9); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__10 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__10(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__10); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__11 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__11(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__11); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__12 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__12(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__12); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__13 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__13(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__13); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__14 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__14(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__14); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__15 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__15(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__15); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__16 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__16(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__16); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__17 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__17(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__17); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__18 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__18(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_8748____closed__18); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__1 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__1(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__1); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__2 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__2(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__2); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__3 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__3(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__3); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__4 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__4(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__4); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__5 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__5(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__5); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__6 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__6(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__6); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__7 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__7(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__7); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__8 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__8(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__8); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__9 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__9(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__9); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__10 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__10(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__10); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__11 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__11(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__11); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__12 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__12(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__12); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__13 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__13(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__13); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__14 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__14(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__14); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__15 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__15(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__15); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__16 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__16(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__16); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__17 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__17(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__17); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__18 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__18(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_9723____closed__18); l_Lean_Meta_instReprEtaStructMode___closed__1 = _init_l_Lean_Meta_instReprEtaStructMode___closed__1(); lean_mark_persistent(l_Lean_Meta_instReprEtaStructMode___closed__1); l_Lean_Meta_instReprEtaStructMode = _init_l_Lean_Meta_instReprEtaStructMode(); @@ -34958,62 +35648,62 @@ l_Lean_Meta_DSimp_instBEqConfig___closed__1 = _init_l_Lean_Meta_DSimp_instBEqCon lean_mark_persistent(l_Lean_Meta_DSimp_instBEqConfig___closed__1); l_Lean_Meta_DSimp_instBEqConfig = _init_l_Lean_Meta_DSimp_instBEqConfig(); lean_mark_persistent(l_Lean_Meta_DSimp_instBEqConfig); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__1 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__1(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__1); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__2 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__2(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__2); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__3 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__3(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__3); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__4 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__4(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__4); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__5 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__5(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__5); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__6 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__6(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__6); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__7 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__7(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__7); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__8 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__8(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__8); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__9 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__9(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__9); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__10 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__10(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__10); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__11 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__11(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__11); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__12 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__12(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__12); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__13 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__13(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__13); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__14 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__14(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__14); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__15 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__15(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__15); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__16 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__16(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__16); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__17 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__17(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__17); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__18 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__18(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__18); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__19 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__19(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__19); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__20 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__20(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__20); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__21 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__21(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__21); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__22 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__22(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__22); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__23 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__23(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__23); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__24 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__24(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__24); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__25 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__25(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__25); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__26 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__26(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__26); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__27 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__27(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__27); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__28 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__28(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_9114____closed__28); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__1 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__1(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__1); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__2 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__2(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__2); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__3 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__3(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__3); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__4 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__4(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__4); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__5 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__5(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__5); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__6 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__6(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__6); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__7 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__7(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__7); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__8 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__8(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__8); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__9 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__9(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__9); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__10 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__10(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__10); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__11 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__11(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__11); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__12 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__12(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__12); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__13 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__13(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__13); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__14 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__14(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__14); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__15 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__15(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__15); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__16 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__16(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__16); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__17 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__17(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__17); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__18 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__18(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__18); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__19 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__19(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__19); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__20 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__20(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__20); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__21 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__21(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__21); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__22 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__22(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__22); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__23 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__23(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__23); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__24 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__24(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__24); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__25 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__25(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__25); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__26 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__26(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__26); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__27 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__27(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__27); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__28 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__28(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_10089____closed__28); l_Lean_Meta_DSimp_instReprConfig___closed__1 = _init_l_Lean_Meta_DSimp_instReprConfig___closed__1(); lean_mark_persistent(l_Lean_Meta_DSimp_instReprConfig___closed__1); l_Lean_Meta_DSimp_instReprConfig = _init_l_Lean_Meta_DSimp_instReprConfig(); @@ -35044,34 +35734,34 @@ l_Lean_Meta_Simp_instBEqConfig___closed__1 = _init_l_Lean_Meta_Simp_instBEqConfi lean_mark_persistent(l_Lean_Meta_Simp_instBEqConfig___closed__1); l_Lean_Meta_Simp_instBEqConfig = _init_l_Lean_Meta_Simp_instBEqConfig(); lean_mark_persistent(l_Lean_Meta_Simp_instBEqConfig); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__1 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__1(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__1); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__2 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__2(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__2); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__3 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__3(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__3); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__4 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__4(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__4); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__5 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__5(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__5); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__6 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__6(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__6); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__7 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__7(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__7); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__8 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__8(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__8); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__9 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__9(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__9); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__10 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__10(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__10); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__11 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__11(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__11); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__12 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__12(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__12); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__13 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__13(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__13); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__14 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__14(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_9646____closed__14); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__1 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__1(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__1); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__2 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__2(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__2); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__3 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__3(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__3); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__4 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__4(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__4); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__5 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__5(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__5); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__6 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__6(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__6); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__7 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__7(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__7); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__8 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__8(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__8); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__9 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__9(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__9); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__10 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__10(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__10); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__11 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__11(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__11); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__12 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__12(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__12); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__13 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__13(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__13); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__14 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__14(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_10621____closed__14); l_Lean_Meta_Simp_instReprConfig___closed__1 = _init_l_Lean_Meta_Simp_instReprConfig___closed__1(); lean_mark_persistent(l_Lean_Meta_Simp_instReprConfig___closed__1); l_Lean_Meta_Simp_instReprConfig = _init_l_Lean_Meta_Simp_instReprConfig(); @@ -35245,442 +35935,428 @@ l_Lean_Parser_Tactic_dsimpKind___closed__10 = _init_l_Lean_Parser_Tactic_dsimpKi lean_mark_persistent(l_Lean_Parser_Tactic_dsimpKind___closed__10); l_Lean_Parser_Tactic_dsimpKind = _init_l_Lean_Parser_Tactic_dsimpKind(); lean_mark_persistent(l_Lean_Parser_Tactic_dsimpKind); -l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__1 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__1(); -lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__1); -l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__2 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__2(); -lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__2); -l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__3 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__3(); -lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__3); -l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__4 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__4(); -lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__4); -l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__5 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__5(); -lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__5); -l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__6 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__6(); -lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__6); -l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__7 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__7(); -lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__7); -l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__8 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__8(); -lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__8); -l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__9 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__9(); -lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__9); -l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__10 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__10(); -lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__10); -l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__11 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__11(); -lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__11); -l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__12 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__12(); -lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__12); -l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__13 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__13(); -lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__13); -l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__14 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__14(); -lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__14); -l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__15 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__15(); -lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__15); -l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__16 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__16(); -lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__16); -l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__17 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__17(); -lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__17); -l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__18 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__18(); -lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__18); -l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__19 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__19(); -lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__19); -l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__20 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__20(); -lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__20); -l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__21 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__21(); -lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__21); -l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__22 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__22(); -lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_______________closed__22); -l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic____________ = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic____________(); -lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic____________); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__1 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__1); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__2 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__2(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__2); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__3 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__3(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__3); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__4 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__4(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__4); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__5 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__5(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__5); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__6 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__6(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__6); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__7 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__7(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__7); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__8 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__8(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__8); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__9 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__9(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__9); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__10 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__10(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__10); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__11 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__11(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__11); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__12 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__12(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__12); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__13 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__13(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__13); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__14 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__14(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__14); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__15 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__15(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__15); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__16 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__16(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__16); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__17 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__17(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__17); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__18 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__18(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__18); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__19 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__19(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__19); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__20 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__20(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__20); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__21 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__21(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__21); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__22 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__22(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__22); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__23 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__23(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__23); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__24 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__24(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__24); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__25 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__25(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__25); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__26 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__26(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__26); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__27 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__27(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__27); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__28 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__28(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__28); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__29 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__29(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__29); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__30 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__30(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__30); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__31 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__31(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__31); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__32 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__32(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__32); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__33 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__33(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__33); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__34 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__34(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__34); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__35 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__35(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__35); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__36 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__36(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__36); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__37 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__37(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__37); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__38 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__38(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__38); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__39 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__39(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__39); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__40 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__40(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__40); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__41 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__41(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__41); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__42 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__42(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__42); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__43 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__43(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__43); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__44 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__44(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__44); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__45 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__45(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__45); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__46 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__46(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__46); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__47 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__47(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__47); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__48 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__48(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__48); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__49 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__49(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__49); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__50 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__50(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__50); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__51 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__51(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__51); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__52 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__52(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__52); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__53 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__53(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__53); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__54 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__54(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__54); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__55 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__55(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__55); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__56 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__56(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__56); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__57 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__57(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__57); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__58 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__58(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__58); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__59 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__59(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__59); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__60 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__60(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__60); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__61 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__61(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__61); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__62 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__62(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__62); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__63 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__63(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__63); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__64 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__64(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__64); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__65 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__65(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__65); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__66 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__66(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__66); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__67 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__67(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__67); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__68 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__68(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__68); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__69 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__69(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__69); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__70 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__70(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__70); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__71 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__71(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__71); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__72 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__72(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__72); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__73 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__73(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__73); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__74 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__74(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__74); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__75 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__75(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__75); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__76 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__76(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__76); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__77 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__77(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__77); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__78 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__78(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__78); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__79 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__79(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__79); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__80 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__80(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__80); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__81 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__81(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__81); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__82 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__82(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__82); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__83 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__83(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__83); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__84 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__84(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__84); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__85 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__85(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__85); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__86 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__86(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__86); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__87 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__87(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__87); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__88 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__88(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__88); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__89 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__89(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__89); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__90 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__90(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__90); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__91 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__91(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__91); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__92 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__92(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__92); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__93 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__93(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__93); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__94 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__94(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__94); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__95 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__95(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__95); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__96 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__96(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__96); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__97 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__97(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__97); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__98 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__98(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___lambda__1___closed__98); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__1 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__1(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__1); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__2 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__2(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__2); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__3 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__3(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__3); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__4 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__4(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__4); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__5 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__5(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__5); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__6 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__6(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__6); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__7 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__7(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__7); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__8 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__8(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__8); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__9 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__9(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__9); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__10 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__10(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__10); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__11 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__11(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__11); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__12 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__12(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__12); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__13 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__13(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__13); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__14 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__14(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__14); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__15 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__15(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__15); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__16 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__16(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__16); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__17 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__17(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__17); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__18 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__18(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__18); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__19 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__19(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__19); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__20 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__20(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__20); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__21 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__21(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__21); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__22 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__22(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__22); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__23 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__23(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__23); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__24 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__24(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__24); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__25 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__25(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__25); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__26 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__26(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__26); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__27 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__27(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__27); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__28 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__28(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__28); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__29 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__29(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__29); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__30 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__30(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__30); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__31 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__31(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__31); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__32 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__32(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__32); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__33 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__33(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__33); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__34 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__34(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__34); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__35 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__35(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__35); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__36 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__36(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__36); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__37 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__37(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__37); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__38 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__38(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__38); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__39 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__39(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__39); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__40 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__40(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__40); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__41 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__41(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__41); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__42 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__42(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__42); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__43 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__43(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__43); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__44 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__44(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__44); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__45 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__45(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__45); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__46 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__46(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__46); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__47 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__47(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__47); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__48 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__48(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__48); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__49 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__49(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__49); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__50 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__50(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__50); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__51 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__51(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__51); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__52 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__52(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__52); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__53 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__53(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__53); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__54 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__54(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__54); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__55 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__55(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__55); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__56 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__56(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__56); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__57 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__57(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__57); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__58 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__58(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__58); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__59 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__59(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__59); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__60 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__60(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__60); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__61 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__61(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__61); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__62 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__62(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__62); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__63 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__63(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__63); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__64 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__64(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__64); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__65 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__65(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__65); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__66 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__66(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__66); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__67 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__67(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__67); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__68 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__68(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__68); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__69 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__69(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__69); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__70 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__70(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__70); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__71 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__71(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__71); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__72 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__72(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__72); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__73 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__73(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__73); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__74 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__74(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__74); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__75 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__75(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__75); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__76 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__76(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__76); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__77 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__77(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__77); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__78 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__78(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__78); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__79 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__79(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__79); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__80 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__80(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__80); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__81 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__81(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__81); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__82 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__82(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__82); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__83 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__83(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__83); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__84 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__84(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__84); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__85 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__85(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__85); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__86 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__86(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__86); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__87 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__87(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__87); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__88 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__88(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__88); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__89 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__89(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__89); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__90 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__90(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__90); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__91 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__91(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__91); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__92 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__92(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__92); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__93 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__93(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__93); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__94 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__94(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__94); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__95 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__95(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__95); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__96 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__96(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__96); -l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__97 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__97(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic______________1___closed__97); +l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__1 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__1); +l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__2 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__2(); +lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__2); +l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__3 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__3(); +lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__3); +l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__4 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__4(); +lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__4); +l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__5 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__5(); +lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__5); +l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__6 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__6(); +lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__6); +l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__7 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__7(); +lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__7); +l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__8 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__8(); +lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__8); +l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__9 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__9(); +lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__9); +l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__10 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__10(); +lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__10); +l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__11 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__11(); +lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__11); +l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__12 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__12(); +lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__12); +l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__13 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__13(); +lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__13); +l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__14 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__14(); +lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__14); +l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__15 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__15(); +lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__15); +l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__16 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__16(); +lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__16); +l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__17 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__17(); +lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__17); +l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__18 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__18(); +lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__18); +l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__19 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__19(); +lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__19); +l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__20 = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__20(); +lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic_____________closed__20); +l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic__________ = _init_l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic__________(); +lean_mark_persistent(l_Lean_Parser_Tactic_commandDeclare__simp__like__tactic__________); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__1 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__1); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__2 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__2); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__3 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__3(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__3); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__4 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__4(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__4); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__5 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__5(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__5); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__6 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__6(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__6); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__7 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__7(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__7); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__8 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__8(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__8); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__9 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__9(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__9); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__10 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__10(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__10); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__11 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__11(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__11); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__12 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__12(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__12); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__13 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__13(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__13); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__14 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__14(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__14); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__15 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__15(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__15); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__16 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__16(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__16); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__17 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__17(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__17); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__18 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__18(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__18); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__19 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__19(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__19); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__20 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__20(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__20); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__21 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__21(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__21); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__22 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__22(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__22); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__23 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__23(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__23); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__24 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__24(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__24); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__25 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__25(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__25); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__26 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__26(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__26); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__27 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__27(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__27); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__28 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__28(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__28); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__29 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__29(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__29); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__30 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__30(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__30); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__31 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__31(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__31); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__32 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__32(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__32); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__33 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__33(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__33); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__34 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__34(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__34); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__35 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__35(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__35); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__36 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__36(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__36); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__37 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__37(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__37); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__38 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__38(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__38); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__39 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__39(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__39); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__40 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__40(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__40); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__41 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__41(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__41); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__42 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__42(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__42); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__43 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__43(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__43); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__44 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__44(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__44); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__45 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__45(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__45); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__46 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__46(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__46); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__47 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__47(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__47); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__48 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__48(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__48); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__49 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__49(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__49); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__50 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__50(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__50); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__51 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__51(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__51); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__52 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__52(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__52); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__53 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__53(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__53); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__54 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__54(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__54); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__55 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__55(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__55); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__56 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__56(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__56); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__57 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__57(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__57); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__58 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__58(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__58); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__59 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__59(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__59); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__60 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__60(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__60); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__61 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__61(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__61); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__62 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__62(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__62); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__63 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__63(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__63); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__64 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__64(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__64); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__65 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__65(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__65); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__66 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__66(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__66); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__67 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__67(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__67); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__68 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__68(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__68); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__69 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__69(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__69); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__70 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__70(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__70); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__71 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__71(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__71); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__72 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__72(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__72); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__73 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__73(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__73); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__74 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__74(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__74); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__75 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__75(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__75); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__76 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__76(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__76); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__77 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__77(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__77); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__78 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__78(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__78); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__79 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__79(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__79); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__80 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__80(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__80); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__81 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__81(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__81); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__82 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__82(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__82); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__83 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__83(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__83); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__84 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__84(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__84); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__85 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__85(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__85); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__86 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__86(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__86); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__87 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__87(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__87); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__88 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__88(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__88); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__89 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__89(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__89); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__90 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__90(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__90); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__91 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__91(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__91); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__92 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__92(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__92); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__93 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__93(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__93); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__94 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__94(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__94); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__95 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__95(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___lambda__1___closed__95); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__1 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__1); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__2 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__2(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__2); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__3 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__3(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__3); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__4 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__4(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__4); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__5 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__5(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__5); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__6 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__6(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__6); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__7 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__7(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__7); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__8 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__8(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__8); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__9 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__9(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__9); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__10 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__10(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__10); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__11 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__11(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__11); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__12 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__12(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__12); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__13 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__13(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__13); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__14 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__14(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__14); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__15 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__15(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__15); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__16 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__16(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__16); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__17 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__17(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__17); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__18 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__18(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__18); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__19 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__19(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__19); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__20 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__20(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__20); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__21 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__21(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__21); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__22 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__22(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__22); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__23 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__23(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__23); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__24 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__24(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__24); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__25 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__25(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__25); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__26 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__26(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__26); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__27 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__27(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__27); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__28 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__28(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__28); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__29 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__29(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__29); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__30 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__30(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__30); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__31 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__31(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__31); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__32 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__32(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__32); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__33 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__33(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__33); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__34 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__34(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__34); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__35 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__35(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__35); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__36 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__36(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__36); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__37 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__37(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__37); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__38 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__38(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__38); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__39 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__39(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__39); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__40 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__40(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__40); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__41 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__41(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__41); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__42 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__42(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__42); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__43 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__43(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__43); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__44 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__44(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__44); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__45 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__45(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__45); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__46 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__46(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__46); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__47 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__47(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__47); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__48 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__48(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__48); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__49 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__49(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__49); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__50 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__50(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__50); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__51 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__51(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__51); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__52 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__52(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__52); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__53 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__53(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__53); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__54 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__54(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__54); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__55 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__55(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__55); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__56 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__56(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__56); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__57 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__57(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__57); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__58 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__58(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__58); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__59 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__59(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__59); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__60 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__60(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__60); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__61 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__61(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__61); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__62 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__62(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__62); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__63 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__63(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__63); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__64 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__64(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__64); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__65 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__65(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__65); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__66 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__66(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__66); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__67 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__67(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__67); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__68 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__68(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__68); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__69 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__69(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__69); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__70 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__70(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__70); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__71 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__71(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__71); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__72 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__72(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__72); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__73 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__73(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__73); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__74 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__74(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__74); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__75 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__75(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__75); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__76 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__76(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__76); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__77 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__77(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__77); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__78 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__78(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__78); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__79 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__79(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__79); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__80 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__80(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__80); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__81 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__81(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__81); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__82 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__82(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__82); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__83 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__83(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__83); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__84 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__84(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__84); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__85 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__85(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__85); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__86 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__86(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__86); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__87 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__87(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__87); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__88 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__88(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__88); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__89 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__89(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__89); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__90 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__90(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__90); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__91 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__91(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__91); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__92 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__92(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__92); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__93 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__93(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__93); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__94 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__94(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__94); +l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__95 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__95(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__commandDeclare__simp__like__tactic____________1___closed__95); l_Lean_Parser_Tactic_simpAutoUnfold___closed__1 = _init_l_Lean_Parser_Tactic_simpAutoUnfold___closed__1(); lean_mark_persistent(l_Lean_Parser_Tactic_simpAutoUnfold___closed__1); l_Lean_Parser_Tactic_simpAutoUnfold___closed__2 = _init_l_Lean_Parser_Tactic_simpAutoUnfold___closed__2(); @@ -35733,48 +36409,48 @@ l_Lean_Parser_Tactic_simpAutoUnfold___closed__25 = _init_l_Lean_Parser_Tactic_si lean_mark_persistent(l_Lean_Parser_Tactic_simpAutoUnfold___closed__25); l_Lean_Parser_Tactic_simpAutoUnfold = _init_l_Lean_Parser_Tactic_simpAutoUnfold(); lean_mark_persistent(l_Lean_Parser_Tactic_simpAutoUnfold); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____lambda__1___closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____lambda__1___closed__1); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__1(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__1); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__2 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__2(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__2); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__3 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__3(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__3); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__4 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__4(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__4); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__5 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__5(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__5); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__6 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__6(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__6); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__7 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__7(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__7); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__8 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__8(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__8); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__9 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__9(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__9); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__10 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__10(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__10); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__11 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__11(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__11); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__12 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__12(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__12); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__13 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__13(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__13); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__14 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__14(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__14); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__15 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__15(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__15); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__16 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__16(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__16); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__17 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__17(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__17); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__18 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__18(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__18); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__19 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__19(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__19); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__20 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__20(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_13759____closed__20); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____lambda__1___closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____lambda__1___closed__1); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__1); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__2 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__2(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__2); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__3 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__3(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__3); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__4 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__4(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__4); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__5 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__5(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__5); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__6 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__6(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__6); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__7 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__7(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__7); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__8 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__8(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__8); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__9 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__9(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__9); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__10 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__10(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__10); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__11 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__11(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__11); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__12 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__12(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__12); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__13 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__13(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__13); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__14 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__14(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__14); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__15 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__15(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__15); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__16 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__16(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__16); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__17 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__17(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__17); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__18 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__18(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__18); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__19 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__19(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__19); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__20 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__20(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14861____closed__20); l_Lean_Parser_Tactic_simpArith___closed__1 = _init_l_Lean_Parser_Tactic_simpArith___closed__1(); lean_mark_persistent(l_Lean_Parser_Tactic_simpArith___closed__1); l_Lean_Parser_Tactic_simpArith___closed__2 = _init_l_Lean_Parser_Tactic_simpArith___closed__2(); @@ -35797,12 +36473,12 @@ l_Lean_Parser_Tactic_simpArith___closed__10 = _init_l_Lean_Parser_Tactic_simpAri lean_mark_persistent(l_Lean_Parser_Tactic_simpArith___closed__10); l_Lean_Parser_Tactic_simpArith = _init_l_Lean_Parser_Tactic_simpArith(); lean_mark_persistent(l_Lean_Parser_Tactic_simpArith); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14552____closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14552____closed__1(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14552____closed__1); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14552____closed__2 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14552____closed__2(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14552____closed__2); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14552____closed__3 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14552____closed__3(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_14552____closed__3); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15670____closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15670____closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15670____closed__1); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15670____closed__2 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15670____closed__2(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15670____closed__2); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15670____closed__3 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15670____closed__3(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15670____closed__3); l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__1 = _init_l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__1(); lean_mark_persistent(l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__1); l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__2 = _init_l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__2(); @@ -35853,22 +36529,22 @@ l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__13 = _init_l_Lean_Parser_Tactic lean_mark_persistent(l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__13); l_Lean_Parser_Tactic_simpAllAutoUnfold = _init_l_Lean_Parser_Tactic_simpAllAutoUnfold(); lean_mark_persistent(l_Lean_Parser_Tactic_simpAllAutoUnfold); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____lambda__1___closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____lambda__1___closed__1); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__1(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__1); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__2 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__2(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__2); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__3 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__3(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__3); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__4 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__4(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__4); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__5 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__5(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__5); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__6 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__6(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__6); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__7 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__7(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16246____closed__7); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____lambda__1___closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____lambda__1___closed__1); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__1); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__2 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__2(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__2); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__3 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__3(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__3); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__4 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__4(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__4); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__5 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__5(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__5); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__6 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__6(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__6); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__7 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__7(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17396____closed__7); l_Lean_Parser_Tactic_simpAllArith___closed__1 = _init_l_Lean_Parser_Tactic_simpAllArith___closed__1(); lean_mark_persistent(l_Lean_Parser_Tactic_simpAllArith___closed__1); l_Lean_Parser_Tactic_simpAllArith___closed__2 = _init_l_Lean_Parser_Tactic_simpAllArith___closed__2(); @@ -35931,24 +36607,24 @@ l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__10 = _init_l_Lean_Parser_Tactic_d lean_mark_persistent(l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__10); l_Lean_Parser_Tactic_dsimpAutoUnfold = _init_l_Lean_Parser_Tactic_dsimpAutoUnfold(); lean_mark_persistent(l_Lean_Parser_Tactic_dsimpAutoUnfold); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____lambda__1___closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____lambda__1___closed__1); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__1(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__1); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__2 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__2(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__2); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__3 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__3(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__3); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__4 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__4(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__4); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__5 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__5(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__5); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__6 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__6(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__6); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__7 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__7(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__7); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__8 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__8(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18703____closed__8); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____lambda__1___closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____lambda__1___closed__1); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__1); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__2 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__2(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__2); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__3 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__3(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__3); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__4 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__4(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__4); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__5 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__5(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__5); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__6 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__6(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__6); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__7 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__7(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__7); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__8 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__8(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19901____closed__8); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Init/Notation.c b/stage0/stdlib/Init/Notation.c index cd11047c0cd..b997b1e4be0 100644 --- a/stage0/stdlib/Init/Notation.c +++ b/stage0/stdlib/Init/Notation.c @@ -37,6 +37,7 @@ static lean_object* l_term___x25_____closed__1; LEAN_EXPORT lean_object* l_term___x3d__; static lean_object* l_stx___x3c_x7c_x3e_____closed__9; static lean_object* l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__14; +static lean_object* l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__5; LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__Complement__complement__1(lean_object*, lean_object*, lean_object*); static lean_object* l_term___x3c_x24_x3e_____closed__2; static lean_object* l___aux__Init__Notation______macroRules__term___x3e_x3e____1___closed__5; @@ -74,6 +75,7 @@ LEAN_EXPORT lean_object* l_term_x2d__; lean_object* lean_mk_empty_array_with_capacity(lean_object*); static lean_object* l_term___u2264_____closed__5; static lean_object* l_term___x3e_x3d_____closed__1; +static lean_object* l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__7; static lean_object* l___aux__Init__Notation______macroRules__stx___x2b__1___closed__9; static lean_object* l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__13; static lean_object* l_precMax___closed__3; @@ -88,7 +90,6 @@ static lean_object* l___aux__Init__Notation______macroRules__term___x3e____1___c static lean_object* l___aux__Init__Notation______macroRules__term___x2a____1___closed__3; static lean_object* l___aux__Init__Notation______macroRules__term___x3d_x3d____1___closed__1; LEAN_EXPORT lean_object* l_Lean___aux__Init__Notation______macroRules__Lean__termThis__1(lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_nullKind; static lean_object* l___aux__Init__Notation______macroRules__term___u2218____1___closed__11; LEAN_EXPORT lean_object* l_term___x3c_x3c_x3c__; static lean_object* l_term___x3d_x3d_____closed__6; @@ -113,7 +114,7 @@ LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__prioHigh__1(le static lean_object* l_term___x26_x26_x26_____closed__3; static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x3d____2___closed__1; static lean_object* l_Lean_Parser_Syntax_addPrec___closed__1; -static lean_object* l_term_x7b_____x3a___x2f_x2f___x7d___closed__6; +static lean_object* l_term_x7b___x3a___x2f_x2f___x7d___closed__9; static lean_object* l_term___x7c_x7c_____closed__3; static lean_object* l_Lean_Parser_Syntax_addPrio___closed__4; static lean_object* l_term___x2d_____closed__2; @@ -124,6 +125,7 @@ static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x3c_x3c_ static lean_object* l___aux__Init__Notation______macroRules__term_xac____1___closed__2; static lean_object* l_term___x3e_x3e_x3e_____closed__2; static lean_object* l_term___x3e_x3d_____closed__2; +LEAN_EXPORT lean_object* l_Lean_instCoeSyntaxNodeKindSyntaxNodeKinds(lean_object*); LEAN_EXPORT lean_object* l_term___x3e_x3e_x3d__; static lean_object* l___aux__Init__Notation______macroRules__term___x2b_x2b____1___closed__1; static lean_object* l_term___x7c_x7c_x7c_____closed__1; @@ -229,7 +231,6 @@ static lean_object* l_term___x3e_x3d_____closed__6; static lean_object* l_termIfLet___x3a_x3d__Then__Else_____closed__12; LEAN_EXPORT lean_object* l_Lean___aux__Init__Notation______macroRules__term_x5b___x5d__1(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__8; -static lean_object* l_term_x7b_____x3a___x2f_x2f___x7d___closed__16; static lean_object* l___aux__Init__Notation______macroRules__term___x3e_x3d____1___closed__6; static lean_object* l___aux__Init__Notation______macroRules__term___x2b____1___closed__2; static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__26; @@ -240,6 +241,7 @@ static lean_object* l___aux__Init__Notation______macroRules__term___x3e_x3e_x3d_ static lean_object* l_term___x3e_x3e_____closed__3; static lean_object* l_rawNatLit___closed__2; static lean_object* l_Lean_Parser_Syntax_addPrec___closed__6; +static lean_object* l_term_x7b___x3a___x2f_x2f___x7d___closed__3; static lean_object* l_term___xd7_____closed__7; static lean_object* l___aux__Init__Notation______macroRules__term___x3e_x3e_x3d____1___closed__8; static lean_object* l___aux__Init__Notation______macroRules__term___x2d____1___closed__7; @@ -252,6 +254,7 @@ LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__Seq__seq__1___bo static lean_object* l___aux__Init__Notation______macroRules__term___x3a_x3a____1___closed__3; static lean_object* l_term___u2209_____closed__6; lean_object* lean_array_push(lean_object*, lean_object*); +static lean_object* l_term_x7b___x3a___x2f_x2f___x7d___closed__15; static lean_object* l___aux__Init__Notation______macroRules__term___x3a_x3a____1___closed__5; lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_termDepIfThenElse; @@ -319,6 +322,7 @@ static lean_object* l___aux__Init__Notation______macroRules__term_x7e_x7e_x7e___ static lean_object* l_term_x2d_____closed__2; static lean_object* l_term_x21_____closed__3; static lean_object* l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__11; +static lean_object* l_Lean_rawStx_quot___closed__11; static lean_object* l_term___u2265_____closed__1; static lean_object* l___aux__Init__Notation______macroRules__termWithout__expected__type____1___closed__6; LEAN_EXPORT lean_object* l_term___x26_x26__; @@ -359,7 +363,6 @@ static lean_object* l_stx___x2c_x2a___closed__1; static lean_object* l_term___x24_______closed__2; static lean_object* l_term___x2f_x5c_____closed__5; static lean_object* l_term___x7c_x7c_____closed__2; -static lean_object* l_term_x7b_____x3a___x2f_x2f___x7d___closed__8; static lean_object* l___aux__Init__Notation______macroRules__term___x2b_x2b____1___closed__9; static lean_object* l_term_x21_____closed__4; static lean_object* l___aux__Init__Notation______macroRules__term___x26_x26_x26____1___closed__7; @@ -375,6 +378,7 @@ static lean_object* l___aux__Init__Notation______macroRules__termIfThenElse__1__ static lean_object* l___aux__Init__Notation______macroRules__stx___x3c_x7c_x3e____1___closed__2; static lean_object* l___aux__Init__Notation______macroRules__term___x2f____1___closed__8; static lean_object* l_term___u2218_____closed__8; +static lean_object* l_term_x7b___x3a___x2f_x2f___x7d___closed__16; static lean_object* l_termIfLet___x3a_x3d__Then__Else_____closed__18; LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__term___x3c_x3d____2(lean_object*, lean_object*, lean_object*); static lean_object* l_term_x21_____closed__5; @@ -384,8 +388,8 @@ static lean_object* l_term___x3c_x3d_____closed__1; static lean_object* l___aux__Init__Notation______macroRules__term___x5e_x5e_x5e____1___closed__9; static lean_object* l_stx___x3f___closed__1; static lean_object* l_term___x24_______closed__8; +static lean_object* l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__6; static lean_object* l_Lean_Parser_Syntax_addPrio___closed__7; -static lean_object* l_term_x7b_____x3a___x2f_x2f___x7d___closed__15; static lean_object* l_term___x3a_x3a_____closed__4; static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x24_x3e____1___closed__6; static lean_object* l___aux__Init__Notation______macroRules__term___x3e_x3e____1___closed__1; @@ -443,11 +447,12 @@ static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x3d____2 static lean_object* l_term___x3c_x7c_____closed__2; static lean_object* l_termWithout__expected__type_____closed__4; static lean_object* l___aux__Init__Notation______macroRules__term___x2f____1___closed__4; -static lean_object* l_term_x7b_____x3a___x2f_x2f___x7d___closed__9; LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__term___x3c_x2a____1(lean_object*, lean_object*, lean_object*); static lean_object* l_term___x3c_x7c_x3e_____closed__4; static lean_object* l_term_x21_____closed__1; +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); static lean_object* l___aux__Init__Notation______macroRules__stx___x3c_x7c_x3e____1___closed__1; +static lean_object* l_term_x7b___x3a___x2f_x2f___x7d___closed__6; static lean_object* l_term___u2218_____closed__2; static lean_object* l_Lean_rawStx_quot___closed__3; LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__HXor__hXor__1___boxed(lean_object*, lean_object*, lean_object*); @@ -482,13 +487,13 @@ LEAN_EXPORT lean_object* l_term___x7c_x7c__; static lean_object* l_stx___x2c_x2a___closed__2; static lean_object* l_term___x25_____closed__5; static lean_object* l___aux__Init__Notation______macroRules__term___x5e____1___closed__9; +static lean_object* l_term_x7b___x3a___x2f_x2f___x7d___closed__10; static lean_object* l_stx___x3c_x7c_x3e_____closed__3; static lean_object* l___aux__Init__Notation______macroRules__term___x3e____1___closed__4; static lean_object* l_term___x3c_x24_x3e_____closed__5; static lean_object* l_Lean___aux__Init__Notation______macroRules__Lean__termThis__1___closed__3; LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__termMax__prec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_term___x3e_x3d_____closed__3; -static lean_object* l_term_x7b_____x3a___x2f_x2f___x7d___closed__2; LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__Or__2(lean_object*, lean_object*, lean_object*); static lean_object* l_term___x2a_____closed__7; static lean_object* l_term___x24_______closed__9; @@ -553,7 +558,6 @@ static lean_object* l_stx_x21_____closed__5; static lean_object* l_prioHigh___closed__3; static lean_object* l_termIfLet___x3a_x3d__Then__Else_____closed__15; static lean_object* l___aux__Init__Notation______macroRules__term___x7c_x7c____1___closed__5; -static lean_object* l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__7; static lean_object* l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__4; static lean_object* l_stx___x2c_x2a_x2c_x3f___closed__1; static lean_object* l___aux__Init__Notation______macroRules__term___x3e_x3d____1___closed__7; @@ -602,6 +606,7 @@ static lean_object* l___aux__Init__Notation______macroRules__term___u2209____1__ LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__Eq__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__Notation______macroRules__term___x2b_x2b____1___closed__2; static lean_object* l_Lean_Parser_Syntax_addPrec___closed__7; +static lean_object* l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__3; static lean_object* l_precMin1___closed__3; LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__Functor__map__1(lean_object*, lean_object*, lean_object*); static lean_object* l_term___u2218_____closed__9; @@ -615,12 +620,12 @@ static lean_object* l_term___x3e_x3d_____closed__4; static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x7c_x3e____1___closed__2; static lean_object* l_term_xac_____closed__3; LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__term___x3d____2(lean_object*, lean_object*, lean_object*); -static lean_object* l_term_x7b_____x3a___x2f_x2f___x7d___closed__3; static lean_object* l___aux__Init__Notation______macroRules__stx___x2b__1___closed__3; static lean_object* l_term___x7c_x7c_____closed__4; static lean_object* l_prio_x28___x29___closed__5; LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__GE__ge__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__16; +static lean_object* l_term_x7b___x3a___x2f_x2f___x7d___closed__11; static lean_object* l___aux__Init__Notation______macroRules__term___x3d_x3d____1___closed__6; static lean_object* l_term___x2a_____closed__5; static lean_object* l_prioDefault___closed__2; @@ -631,6 +636,7 @@ static lean_object* l_prioMid___closed__4; static lean_object* l_term___x3a_x3a_____closed__1; static lean_object* l_precLead___closed__1; static lean_object* l_Lean_rawStx_quot___closed__6; +LEAN_EXPORT lean_object* l_Lean_instCoeSyntaxTSyntaxConsSyntaxNodeKindMkStrAnonymousNil(lean_object*); LEAN_EXPORT lean_object* l_term___x3e_x3d__; LEAN_EXPORT lean_object* l_prioDefault; static lean_object* l_term___x3c_x7c_x3e_____closed__3; @@ -640,15 +646,14 @@ static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x7c_x3e_ static lean_object* l_Lean___aux__Init__Notation______macroRules__term_x5b___x5d__1___closed__6; static lean_object* l_termDepIfThenElse___closed__6; static lean_object* l___aux__Init__Notation______macroRules__term___x2a_x3e____1___closed__6; -static lean_object* l_term_x7b_____x3a___x2f_x2f___x7d___closed__14; static lean_object* l_term___x24_______closed__13; static lean_object* l___aux__Init__Notation______macroRules__term___x3d____1___closed__4; static lean_object* l_termIfThenElse___closed__10; static lean_object* l___aux__Init__Notation______macroRules__term_x2d____1___closed__6; static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__25; LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__term___x5e____1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_instCoeSyntaxTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___boxed(lean_object*); static lean_object* l_Lean_Parser_Syntax_addPrio___closed__5; -static lean_object* l_term_x7b_____x3a___x2f_x2f___x7d___closed__5; static lean_object* l___aux__Init__Notation______macroRules__term___x3a_x3a____1___closed__8; static lean_object* l___aux__Init__Notation______macroRules__precMax__1___closed__2; static lean_object* l_prioDefault___closed__1; @@ -767,7 +772,6 @@ static lean_object* l_termDepIfThenElse___closed__30; static lean_object* l___aux__Init__Notation______macroRules__term___x26_x26____1___closed__6; static lean_object* l___aux__Init__Notation______macroRules__term___x25____1___closed__4; static lean_object* l_Lean_Parser_Syntax_addPrec___closed__10; -static lean_object* l_term_x7b_____x3a___x2f_x2f___x7d___closed__13; static lean_object* l_term___x7c_x7c_x7c_____closed__3; LEAN_EXPORT lean_object* l_term_x25_x5b___x7c___x5d; static lean_object* l_stx___x2c_x2a___closed__5; @@ -793,9 +797,11 @@ static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x24_x3e_ LEAN_EXPORT lean_object* l_term___x2f_x5c__; static lean_object* l_term___x2b_x2b_____closed__2; static lean_object* l_term___x3e_____closed__3; +static lean_object* l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__1; static lean_object* l_term___x25_____closed__4; static lean_object* l___aux__Init__Notation______macroRules__term___x26_x26_x26____1___closed__9; LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__GE__ge__1(lean_object*, lean_object*, lean_object*); +static lean_object* l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__4; LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__term___x7c_x7c_x7c____2(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__Notation______macroRules__term___x3c____1___closed__8; lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); @@ -808,6 +814,7 @@ static lean_object* l___aux__Init__Notation______macroRules__term___x3d____1___c static lean_object* l___aux__Init__Notation______macroRules__precMax__1___closed__1; static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__5; LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__HAdd__hAdd__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_term_x7b___x3a___x2f_x2f___x7d___closed__7; static lean_object* l___aux__Init__Notation______macroRules__term___x26_x26_x26____1___closed__8; LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__precMin__1(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x24_x3e____1___closed__9; @@ -842,12 +849,10 @@ static lean_object* l___aux__Init__Notation______macroRules__term___x5c_x2f____1 static lean_object* l___aux__Init__Notation______macroRules__term___x2f____1___closed__5; static lean_object* l___aux__Init__Notation______macroRules__term___x26_x26_x26____1___closed__2; static lean_object* l___aux__Init__Notation______macroRules__term_xac____1___closed__5; -static lean_object* l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__5; LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__HShiftLeft__hShiftLeft__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__20; LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__term___x3c_x7c_x3e____1(lean_object*, lean_object*, lean_object*); static lean_object* l_term___x5e_____closed__4; -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__Notation______macroRules__term___x3c____1___closed__2; static lean_object* l_Lean_Parser_Syntax_subPrio___closed__4; LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__HMod__hMod__1(lean_object*, lean_object*, lean_object*); @@ -903,8 +908,6 @@ static lean_object* l_boolIfThenElse___closed__12; static lean_object* l___aux__Init__Notation______macroRules__term___x3d_x3d____2___closed__3; static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x2a____1___closed__5; LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__stx___x3c_x7c_x3e____1(lean_object*, lean_object*, lean_object*); -static lean_object* l_term_x7b_____x3a___x2f_x2f___x7d___closed__1; -static lean_object* l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__6; static lean_object* l___aux__Init__Notation______macroRules__term___x3d_x3d____1___closed__9; static lean_object* l_stx___x2c_x2b___closed__2; static lean_object* l___aux__Init__Notation______macroRules__term___x3e_x3e_x3d____1___closed__6; @@ -921,7 +924,6 @@ LEAN_EXPORT lean_object* l_term_xac__; static lean_object* l_term___u2218_____closed__6; static lean_object* l_term___x3c_x7c_____closed__1; static lean_object* l_stx_x21_____closed__6; -static lean_object* l_term_x7b_____x3a___x2f_x2f___x7d___closed__4; static lean_object* l_stx___x2b___closed__4; static lean_object* l___aux__Init__Notation______macroRules__precMin__1___closed__1; LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__Seq__seq__1(lean_object*, lean_object*, lean_object*); @@ -995,7 +997,6 @@ static lean_object* l___aux__Init__Notation______macroRules__term___x3e_x3e____1 static lean_object* l___aux__Init__Notation______macroRules__term___x2f_x5c____1___closed__3; static lean_object* l___aux__Init__Notation______macroRules__term___x3d_x3d____1___closed__5; static lean_object* l___aux__Init__Notation______macroRules__stx_x21____1___closed__3; -LEAN_EXPORT lean_object* l_term_x7b_____x3a___x2f_x2f___x7d; LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__term___x7c_x3e____1(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__Notation______macroRules__term___x3c____1___closed__7; static lean_object* l_term___x7c_x7c_____closed__6; @@ -1022,6 +1023,7 @@ LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__Prod__1(lean_obj LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__HSub__hSub__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Syntax_subPrec___closed__6; static lean_object* l___aux__Init__Notation______macroRules__term___x2a_x3e____1___closed__1; +static lean_object* l_term_x7b___x3a___x2f_x2f___x7d___closed__1; LEAN_EXPORT lean_object* l_Lean___aux__Init__Notation______macroRules__term_x5b___x5d__1_expandListLit(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x3c_x3c____1___closed__6; LEAN_EXPORT lean_object* l_term___x3c_x24_x3e__; @@ -1035,6 +1037,7 @@ static lean_object* l_Lean_rawStx_quot___closed__5; static lean_object* l_termIfLet___x3a_x3d__Then__Else_____closed__6; static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__6; static lean_object* l_termDepIfThenElse___closed__29; +static lean_object* l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__2; static lean_object* l_term___u2264_____closed__4; static lean_object* l_stx___x2b___closed__2; static lean_object* l_term___x2f_x5c_____closed__3; @@ -1067,6 +1070,7 @@ static lean_object* l_term___x3c_____closed__6; static lean_object* l___aux__Init__Notation______macroRules__term___x3c____1___closed__5; static lean_object* l___aux__Init__Notation______macroRules__term___u2218____1___closed__12; LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__SeqLeft__seqLeft__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_term_x7b___x3a___x2f_x2f___x7d___closed__13; static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x7c_x3e____1___closed__10; static lean_object* l_term___x3c_x3c_x3c_____closed__1; static lean_object* l_term___u2228_____closed__3; @@ -1079,6 +1083,7 @@ static lean_object* l_boolIfThenElse___closed__11; static lean_object* l___aux__Init__Notation______macroRules__term___x3e_x3e_x3e____1___closed__8; LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__prec_x28___x29__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x7c_x3e____1___closed__5; +static lean_object* l_term_x7b___x3a___x2f_x2f___x7d___closed__8; static lean_object* l_term_x5b___x5d___closed__12; static lean_object* l___aux__Init__Notation______macroRules__term___x2d____1___closed__2; static lean_object* l___aux__Init__Notation______macroRules__term___x3d____1___closed__5; @@ -1092,7 +1097,6 @@ static lean_object* l_Lean_termThis___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Syntax_addPrio; static lean_object* l___aux__Init__Notation______macroRules__term_x21____1___closed__4; static lean_object* l_term___x2a_____closed__2; -static lean_object* l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__1; static lean_object* l___aux__Init__Notation______unexpand__Function__comp__1___closed__1; static lean_object* l_term___xd7_____closed__1; static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__22; @@ -1108,13 +1112,12 @@ LEAN_EXPORT lean_object* l_term___x2f__; static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__19; static lean_object* l_term___u2265_____closed__6; static lean_object* l_termDepIfThenElse___closed__4; -static lean_object* l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__8; static lean_object* l_termIfThenElse___closed__2; static lean_object* l_term___u2228_____closed__5; static lean_object* l_stx___x2c_x2b_x2c_x3f___closed__5; +static lean_object* l_term_x7b___x3a___x2f_x2f___x7d___closed__4; static lean_object* l_stx_x21_____closed__3; static lean_object* l___aux__Init__Notation______macroRules__term___x2f____1___closed__7; -static lean_object* l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__4; static lean_object* l___aux__Init__Notation______macroRules__stx___x2b__1___closed__6; static lean_object* l___aux__Init__Notation______macroRules__term___x3e_x3e____1___closed__3; LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__termIfLet___x3a_x3d__Then__Else____1(lean_object*, lean_object*, lean_object*); @@ -1129,13 +1132,12 @@ static lean_object* l_precArg___closed__2; uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); static lean_object* l_term___u2208_____closed__5; static lean_object* l_term_x7e_x7e_x7e_____closed__3; +static lean_object* l_term_x7b___x3a___x2f_x2f___x7d___closed__12; static lean_object* l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__2; static lean_object* l___aux__Init__Notation______macroRules__stx___x3f__1___closed__2; -static lean_object* l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__2; static lean_object* l_stx_x21_____closed__4; static lean_object* l_term___x3d_x3d_____closed__2; LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__stx___x3f__1(lean_object*, lean_object*, lean_object*); -static lean_object* l_term_x7b_____x3a___x2f_x2f___x7d___closed__7; static lean_object* l_precMin1___closed__2; LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__HXor__hXor__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__term_x2d____1(lean_object*, lean_object*, lean_object*); @@ -1161,6 +1163,7 @@ static lean_object* l_term___x3e_x3e_x3d_____closed__6; LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__term_x21____1(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x7c_x3e____1___closed__3; static lean_object* l_term___x26_x26_x26_____closed__5; +LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x3d____1___closed__7; static lean_object* l___aux__Init__Notation______macroRules__termIfLet___x3a_x3d__Then__Else____1___closed__3; static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x2a____1___closed__7; @@ -1180,10 +1183,10 @@ static lean_object* l___aux__Init__Notation______macroRules__term___x3e_x3d____1 static lean_object* l_termDepIfThenElse___closed__2; static lean_object* l_termIfThenElse___closed__6; static lean_object* l_term___x3e_x3e_____closed__7; -static lean_object* l_term_x7b_____x3a___x2f_x2f___x7d___closed__10; static lean_object* l_term___x3c_x2a_x3e_____closed__2; static lean_object* l_prec_x28___x29___closed__3; static lean_object* l_term_x5b___x5d___closed__1; +static lean_object* l_term_x7b___x3a___x2f_x2f___x7d___closed__2; static lean_object* l_stx___x2c_x2b_x2c_x3f___closed__1; static lean_object* l_term___x24_______closed__7; LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__term___x3c_x7c____1(lean_object*, lean_object*, lean_object*); @@ -1207,7 +1210,6 @@ static lean_object* l___aux__Init__Notation______macroRules__term___x5e____1___c static lean_object* l_termIfLet___x3a_x3d__Then__Else_____closed__11; static lean_object* l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__21; static lean_object* l___aux__Init__Notation______macroRules__termIfLet___x3a_x3d__Then__Else____1___closed__5; -static lean_object* l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__3; static lean_object* l_term_x2d_____closed__7; static lean_object* l_term___u2208_____closed__1; static lean_object* l_term___u2265_____closed__2; @@ -1245,7 +1247,6 @@ static lean_object* l_term___x2a_____closed__4; static lean_object* l_Lean_Parser_Syntax_addPrec___closed__11; LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__Or__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_term_x7e_x7e_x7e_____closed__2; -LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1(lean_object*, lean_object*, lean_object*); static lean_object* l_term___u2227_____closed__6; static lean_object* l_stx___x2c_x2b_x2c_x3f___closed__2; static lean_object* l_term___x7c_x3e_____closed__4; @@ -1279,7 +1280,6 @@ static lean_object* l___aux__Init__Notation______macroRules__term___x2d____1___c static lean_object* l___aux__Init__Notation______macroRules__term___x26_x26_x26____1___closed__1; static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__17; LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__BEq__beq__1___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_term_x7b_____x3a___x2f_x2f___x7d___closed__12; static lean_object* l_term___x2f_x5c_____closed__2; static lean_object* l___aux__Init__Notation______macroRules__stx_x21____1___closed__1; static lean_object* l_term___x26_x26_____closed__3; @@ -1292,8 +1292,10 @@ static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e_ LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__HAdd__hAdd__1(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__Notation______macroRules__term___x3d____1___closed__6; static lean_object* l_term___x3d_____closed__3; +static lean_object* l_Lean_rawStx_quot___closed__12; static lean_object* l_term___x7c_x3e_____closed__6; static lean_object* l_term___x3c_x7c_____closed__4; +static lean_object* l_term_x7b___x3a___x2f_x2f___x7d___closed__14; static lean_object* l___aux__Init__Notation______macroRules__term___u2218____1___closed__10; static lean_object* l_termDepIfThenElse___closed__5; static lean_object* l_term___x2f_x5c_____closed__4; @@ -1307,6 +1309,7 @@ static lean_object* l_term___x3c_x3c_x3c_____closed__4; static lean_object* l_boolIfThenElse___closed__1; static lean_object* l___aux__Init__Notation______macroRules__term___x2b____1___closed__1; static lean_object* l_term_x7e_x7e_x7e_____closed__4; +static lean_object* l_term_x7b___x3a___x2f_x2f___x7d___closed__5; LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__term___x3c____2(lean_object*, lean_object*, lean_object*); static lean_object* l_precArg___closed__4; static lean_object* l_termWithout__expected__type_____closed__3; @@ -1314,8 +1317,8 @@ static lean_object* l___aux__Init__Notation______macroRules__term___x3e____1___c static lean_object* l_stx___x2c_x2b___closed__5; static lean_object* l_term___u2209_____closed__1; LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__term___x25____1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_term_x7b___x3a___x2f_x2f___x7d; static lean_object* l_term___u2265_____closed__4; -static lean_object* l_term_x7b_____x3a___x2f_x2f___x7d___closed__11; LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__term___x2a_x3e____1(lean_object*, lean_object*, lean_object*); static lean_object* l_term___x3e_x3e_____closed__2; static lean_object* l_Lean_Parser_Syntax_subPrec___closed__1; @@ -1332,6 +1335,7 @@ LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__prioMid__1(lea static lean_object* l___aux__Init__Notation______macroRules__stx___x3f__1___closed__3; static lean_object* l_term___x2a_x3e_____closed__2; LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__prio_x28___x29__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__8; static lean_object* l___aux__Init__Notation______macroRules__term___x2b_x2b____1___closed__7; static lean_object* l_Lean_termThis___closed__2; LEAN_EXPORT lean_object* l_term___x5c_x2f__; @@ -1791,6 +1795,17 @@ lean_dec(x_1); return x_2; } } +LEAN_EXPORT lean_object* l_Lean_instCoeSyntaxNodeKindSyntaxNodeKinds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_box(0); +x_3 = lean_alloc_ctor(1, 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_precMax___closed__1() { _start: { @@ -6538,82 +6553,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___u2218_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___u2218_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___u2218_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___u2218_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___u2218_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___u2218_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___u2218_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___u2218_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -6930,82 +6944,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___xd7_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___xd7_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___xd7_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___xd7_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___xd7_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___xd7_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___xd7_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___xd7_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -7347,82 +7360,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___x7c_x7c_x7c_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___x7c_x7c_x7c_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___x7c_x7c_x7c_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___x7c_x7c_x7c_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___x7c_x7c_x7c_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___x7c_x7c_x7c_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___x7c_x7c_x7c_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___x7c_x7c_x7c_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -7764,82 +7776,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___x5e_x5e_x5e_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___x5e_x5e_x5e_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___x5e_x5e_x5e_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___x5e_x5e_x5e_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___x5e_x5e_x5e_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___x5e_x5e_x5e_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___x5e_x5e_x5e_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___x5e_x5e_x5e_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -8181,82 +8192,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___x26_x26_x26_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___x26_x26_x26_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___x26_x26_x26_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___x26_x26_x26_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___x26_x26_x26_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___x26_x26_x26_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___x26_x26_x26_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___x26_x26_x26_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -8580,82 +8590,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_Lean_Parser_Syntax_addPrec___closed__11; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___x2b_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_Lean_Parser_Syntax_addPrec___closed__11; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___x2b_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_Lean_Parser_Syntax_addPrec___closed__11; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___x2b_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_Lean_Parser_Syntax_addPrec___closed__11; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___x2b_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -8967,82 +8976,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_Lean_Parser_Syntax_subPrec___closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___x2d_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_Lean_Parser_Syntax_subPrec___closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___x2d_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_Lean_Parser_Syntax_subPrec___closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___x2d_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_Lean_Parser_Syntax_subPrec___closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___x2d_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -9384,82 +9392,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___x2a_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___x2a_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___x2a_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___x2a_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; +} else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___x2a_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___x2a_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___x2a_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___x2a_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -9789,82 +9796,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___x2f_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___x2f_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___x2f_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___x2f_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___x2f_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___x2f_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___x2f_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___x2f_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -10194,82 +10200,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___x25_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___x25_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___x25_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___x25_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___x25_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___x25_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___x25_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___x25_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -10611,82 +10616,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___x3c_x3c_x3c_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___x3c_x3c_x3c_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___x3c_x3c_x3c_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___x3c_x3c_x3c_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___x3c_x3c_x3c_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___x3c_x3c_x3c_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___x3c_x3c_x3c_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___x3c_x3c_x3c_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -11016,82 +11020,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___x3e_x3e_x3e_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___x3e_x3e_x3e_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___x3e_x3e_x3e_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___x3e_x3e_x3e_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___x3e_x3e_x3e_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___x3e_x3e_x3e_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___x3e_x3e_x3e_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___x3e_x3e_x3e_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -11434,82 +11437,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___x5e_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___x5e_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___x5e_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___x5e_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___x5e_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___x5e_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___x5e_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___x5e_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -11839,82 +11841,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___x2b_x2b_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___x2b_x2b_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___x2b_x2b_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___x2b_x2b_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___x2b_x2b_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___x2b_x2b_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___x2b_x2b_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___x2b_x2b_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -12253,81 +12254,80 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +lean_object* x_14; lean_object* x_15; uint8_t x_16; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; lean_inc(x_15); -x_17 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_14); -if (x_17 == 0) +x_16 = l_Lean_Syntax_matchesNull(x_15, x_14); +if (x_16 == 0) { -lean_object* x_18; lean_object* x_19; +lean_object* x_17; lean_object* x_18; lean_dec(x_15); lean_dec(x_9); -x_18 = lean_box(0); -x_19 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_3); -return x_19; +x_17 = lean_box(0); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_3); +return x_18; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_20 = l_Lean_Syntax_getArg(x_15, x_8); +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_19 = l_Lean_Syntax_getArg(x_15, x_8); lean_dec(x_15); -x_21 = l_Lean_replaceRef(x_9, x_2); +x_20 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_22 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_21, x_3); -x_23 = !lean_is_exclusive(x_22); -if (x_23 == 0) +x_21 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_20, x_3); +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) { -lean_object* 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; lean_object* x_32; -x_24 = lean_ctor_get(x_22, 0); -x_25 = l_term_x2d_____closed__3; -x_26 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -x_27 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; -x_28 = lean_array_push(x_27, x_26); -x_29 = lean_array_push(x_28, x_20); -x_30 = lean_box(2); -x_31 = l_term_x2d_____closed__2; -x_32 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -lean_ctor_set(x_32, 2, x_29); -lean_ctor_set(x_22, 0, x_32); -return 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; lean_object* x_30; lean_object* x_31; +x_23 = lean_ctor_get(x_21, 0); +x_24 = l_term_x2d_____closed__3; +x_25 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +x_26 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; +x_27 = lean_array_push(x_26, x_25); +x_28 = lean_array_push(x_27, x_19); +x_29 = lean_box(2); +x_30 = l_term_x2d_____closed__2; +x_31 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +lean_ctor_set(x_31, 2, x_28); +lean_ctor_set(x_21, 0, x_31); +return x_21; } else { -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_33 = lean_ctor_get(x_22, 0); -x_34 = lean_ctor_get(x_22, 1); -lean_inc(x_34); +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; +x_32 = lean_ctor_get(x_21, 0); +x_33 = lean_ctor_get(x_21, 1); lean_inc(x_33); -lean_dec(x_22); -x_35 = l_term_x2d_____closed__3; -x_36 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_36, 0, x_33); -lean_ctor_set(x_36, 1, x_35); -x_37 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; -x_38 = lean_array_push(x_37, x_36); -x_39 = lean_array_push(x_38, x_20); -x_40 = lean_box(2); -x_41 = l_term_x2d_____closed__2; -x_42 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_42, 0, x_40); -lean_ctor_set(x_42, 1, x_41); -lean_ctor_set(x_42, 2, x_39); -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_43, 1, x_34); -return x_43; -} -} -} +lean_inc(x_32); +lean_dec(x_21); +x_34 = l_term_x2d_____closed__3; +x_35 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_35, 0, x_32); +lean_ctor_set(x_35, 1, x_34); +x_36 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; +x_37 = lean_array_push(x_36, x_35); +x_38 = lean_array_push(x_37, x_19); +x_39 = lean_box(2); +x_40 = l_term_x2d_____closed__2; +x_41 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +lean_ctor_set(x_41, 2, x_38); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_33); +return x_42; +} +} +} } } } @@ -12651,78 +12651,77 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +lean_object* x_14; lean_object* x_15; uint8_t x_16; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; lean_inc(x_15); -x_17 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_14); -if (x_17 == 0) +x_16 = l_Lean_Syntax_matchesNull(x_15, x_14); +if (x_16 == 0) { -lean_object* x_18; lean_object* x_19; +lean_object* x_17; lean_object* x_18; lean_dec(x_15); lean_dec(x_9); -x_18 = lean_box(0); -x_19 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_3); -return x_19; +x_17 = lean_box(0); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_3); +return x_18; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_20 = l_Lean_Syntax_getArg(x_15, x_8); +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_19 = l_Lean_Syntax_getArg(x_15, x_8); lean_dec(x_15); -x_21 = l_Lean_replaceRef(x_9, x_2); +x_20 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_22 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_21, x_3); -x_23 = !lean_is_exclusive(x_22); -if (x_23 == 0) +x_21 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_20, x_3); +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) { -lean_object* 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; lean_object* x_32; -x_24 = lean_ctor_get(x_22, 0); -x_25 = l_term_x7e_x7e_x7e_____closed__3; -x_26 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -x_27 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; -x_28 = lean_array_push(x_27, x_26); -x_29 = lean_array_push(x_28, x_20); -x_30 = lean_box(2); -x_31 = l_term_x7e_x7e_x7e_____closed__2; -x_32 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -lean_ctor_set(x_32, 2, x_29); -lean_ctor_set(x_22, 0, x_32); -return 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; lean_object* x_30; lean_object* x_31; +x_23 = lean_ctor_get(x_21, 0); +x_24 = l_term_x7e_x7e_x7e_____closed__3; +x_25 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +x_26 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; +x_27 = lean_array_push(x_26, x_25); +x_28 = lean_array_push(x_27, x_19); +x_29 = lean_box(2); +x_30 = l_term_x7e_x7e_x7e_____closed__2; +x_31 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +lean_ctor_set(x_31, 2, x_28); +lean_ctor_set(x_21, 0, x_31); +return x_21; } else { -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_33 = lean_ctor_get(x_22, 0); -x_34 = lean_ctor_get(x_22, 1); -lean_inc(x_34); +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; +x_32 = lean_ctor_get(x_21, 0); +x_33 = lean_ctor_get(x_21, 1); lean_inc(x_33); -lean_dec(x_22); -x_35 = l_term_x7e_x7e_x7e_____closed__3; -x_36 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_36, 0, x_33); -lean_ctor_set(x_36, 1, x_35); -x_37 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; -x_38 = lean_array_push(x_37, x_36); -x_39 = lean_array_push(x_38, x_20); -x_40 = lean_box(2); -x_41 = l_term_x7e_x7e_x7e_____closed__2; -x_42 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_42, 0, x_40); -lean_ctor_set(x_42, 1, x_41); -lean_ctor_set(x_42, 2, x_39); -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_43, 1, x_34); -return x_43; +lean_inc(x_32); +lean_dec(x_21); +x_34 = l_term_x7e_x7e_x7e_____closed__3; +x_35 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_35, 0, x_32); +lean_ctor_set(x_35, 1, x_34); +x_36 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; +x_37 = lean_array_push(x_36, x_35); +x_38 = lean_array_push(x_37, x_19); +x_39 = lean_box(2); +x_40 = l_term_x7e_x7e_x7e_____closed__2; +x_41 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +lean_ctor_set(x_41, 2, x_38); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_33); +return x_42; } } } @@ -13987,82 +13986,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___x3c_x3d_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___x3c_x3d_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___x3c_x3d_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___x3c_x3d_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___x3c_x3d_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___x3c_x3d_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___x3c_x3d_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___x3c_x3d_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -14302,82 +14300,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___u2264_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___u2264_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___u2264_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___u2264_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___u2264_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___u2264_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___u2264_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___u2264_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -14708,82 +14705,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___x3c_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___x3c_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___x3c_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___x3c_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___x3c_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___x3c_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___x3c_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___x3c_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -15114,82 +15110,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___x3e_x3d_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___x3e_x3d_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___x3e_x3d_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___x3e_x3d_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___x3e_x3d_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___x3e_x3d_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___x3e_x3d_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___x3e_x3d_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -15429,82 +15424,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___u2265_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___u2265_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___u2265_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___u2265_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___u2265_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___u2265_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___u2265_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___u2265_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -15835,82 +15829,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___x3e_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___x3e_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___x3e_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___x3e_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___x3e_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___x3e_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___x3e_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___x3e_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -16215,82 +16208,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___x3d_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___x3d_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___x3d_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___x3d_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___x3d_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___x3d_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___x3d_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___x3d_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -16621,82 +16613,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___x3d_x3d_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___x3d_x3d_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___x3d_x3d_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___x3d_x3d_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___x3d_x3d_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___x3d_x3d_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___x3d_x3d_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___x3d_x3d_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -17949,82 +17940,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___x2f_x5c_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___x2f_x5c_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___x2f_x5c_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___x2f_x5c_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___x2f_x5c_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___x2f_x5c_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___x2f_x5c_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___x2f_x5c_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -18264,82 +18254,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___u2227_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___u2227_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___u2227_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___u2227_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___u2227_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___u2227_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___u2227_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___u2227_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -18656,82 +18645,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___x5c_x2f_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___x5c_x2f_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___x5c_x2f_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___x5c_x2f_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___x5c_x2f_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___x5c_x2f_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___x5c_x2f_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___x5c_x2f_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -18971,82 +18959,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___u2228_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___u2228_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___u2228_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___u2228_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___u2228_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___u2228_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___u2228_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___u2228_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -19359,78 +19346,77 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +lean_object* x_14; lean_object* x_15; uint8_t x_16; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; lean_inc(x_15); -x_17 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_14); -if (x_17 == 0) +x_16 = l_Lean_Syntax_matchesNull(x_15, x_14); +if (x_16 == 0) { -lean_object* x_18; lean_object* x_19; +lean_object* x_17; lean_object* x_18; lean_dec(x_15); lean_dec(x_9); -x_18 = lean_box(0); -x_19 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_3); -return x_19; +x_17 = lean_box(0); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_3); +return x_18; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_20 = l_Lean_Syntax_getArg(x_15, x_8); +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_19 = l_Lean_Syntax_getArg(x_15, x_8); lean_dec(x_15); -x_21 = l_Lean_replaceRef(x_9, x_2); +x_20 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_22 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_21, x_3); -x_23 = !lean_is_exclusive(x_22); -if (x_23 == 0) +x_21 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_20, x_3); +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) { -lean_object* 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; lean_object* x_32; -x_24 = lean_ctor_get(x_22, 0); -x_25 = l_term_xac_____closed__3; -x_26 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -x_27 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; -x_28 = lean_array_push(x_27, x_26); -x_29 = lean_array_push(x_28, x_20); -x_30 = lean_box(2); -x_31 = l_term_xac_____closed__2; -x_32 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -lean_ctor_set(x_32, 2, x_29); -lean_ctor_set(x_22, 0, x_32); -return 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; lean_object* x_30; lean_object* x_31; +x_23 = lean_ctor_get(x_21, 0); +x_24 = l_term_xac_____closed__3; +x_25 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +x_26 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; +x_27 = lean_array_push(x_26, x_25); +x_28 = lean_array_push(x_27, x_19); +x_29 = lean_box(2); +x_30 = l_term_xac_____closed__2; +x_31 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +lean_ctor_set(x_31, 2, x_28); +lean_ctor_set(x_21, 0, x_31); +return x_21; } else { -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_33 = lean_ctor_get(x_22, 0); -x_34 = lean_ctor_get(x_22, 1); -lean_inc(x_34); +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; +x_32 = lean_ctor_get(x_21, 0); +x_33 = lean_ctor_get(x_21, 1); lean_inc(x_33); -lean_dec(x_22); -x_35 = l_term_xac_____closed__3; -x_36 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_36, 0, x_33); -lean_ctor_set(x_36, 1, x_35); -x_37 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; -x_38 = lean_array_push(x_37, x_36); -x_39 = lean_array_push(x_38, x_20); -x_40 = lean_box(2); -x_41 = l_term_xac_____closed__2; -x_42 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_42, 0, x_40); -lean_ctor_set(x_42, 1, x_41); -lean_ctor_set(x_42, 2, x_39); -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_43, 1, x_34); -return x_43; +lean_inc(x_32); +lean_dec(x_21); +x_34 = l_term_xac_____closed__3; +x_35 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_35, 0, x_32); +lean_ctor_set(x_35, 1, x_34); +x_36 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; +x_37 = lean_array_push(x_36, x_35); +x_38 = lean_array_push(x_37, x_19); +x_39 = lean_box(2); +x_40 = l_term_xac_____closed__2; +x_41 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +lean_ctor_set(x_41, 2, x_38); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_33); +return x_42; } } } @@ -19746,82 +19732,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___x26_x26_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___x26_x26_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return x_24; -} -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_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___x26_x26_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___x26_x26_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; +} +else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___x26_x26_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___x26_x26_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___x26_x26_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___x26_x26_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -20137,82 +20122,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___x7c_x7c_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___x7c_x7c_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___x7c_x7c_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___x7c_x7c_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___x7c_x7c_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___x7c_x7c_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___x7c_x7c_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___x7c_x7c_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -20505,78 +20489,77 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +lean_object* x_14; lean_object* x_15; uint8_t x_16; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; lean_inc(x_15); -x_17 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_14); -if (x_17 == 0) +x_16 = l_Lean_Syntax_matchesNull(x_15, x_14); +if (x_16 == 0) { -lean_object* x_18; lean_object* x_19; +lean_object* x_17; lean_object* x_18; lean_dec(x_15); lean_dec(x_9); -x_18 = lean_box(0); -x_19 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_3); -return x_19; +x_17 = lean_box(0); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_3); +return x_18; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_20 = l_Lean_Syntax_getArg(x_15, x_8); +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_19 = l_Lean_Syntax_getArg(x_15, x_8); lean_dec(x_15); -x_21 = l_Lean_replaceRef(x_9, x_2); +x_20 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_22 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_21, x_3); -x_23 = !lean_is_exclusive(x_22); -if (x_23 == 0) +x_21 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_20, x_3); +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) { -lean_object* 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; lean_object* x_32; -x_24 = lean_ctor_get(x_22, 0); -x_25 = l_stx_x21_____closed__3; -x_26 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -x_27 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; -x_28 = lean_array_push(x_27, x_26); -x_29 = lean_array_push(x_28, x_20); -x_30 = lean_box(2); -x_31 = l_term_x21_____closed__2; -x_32 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -lean_ctor_set(x_32, 2, x_29); -lean_ctor_set(x_22, 0, x_32); -return 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; lean_object* x_30; lean_object* x_31; +x_23 = lean_ctor_get(x_21, 0); +x_24 = l_stx_x21_____closed__3; +x_25 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +x_26 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; +x_27 = lean_array_push(x_26, x_25); +x_28 = lean_array_push(x_27, x_19); +x_29 = lean_box(2); +x_30 = l_term_x21_____closed__2; +x_31 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +lean_ctor_set(x_31, 2, x_28); +lean_ctor_set(x_21, 0, x_31); +return x_21; } else { -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_33 = lean_ctor_get(x_22, 0); -x_34 = lean_ctor_get(x_22, 1); -lean_inc(x_34); +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; +x_32 = lean_ctor_get(x_21, 0); +x_33 = lean_ctor_get(x_21, 1); lean_inc(x_33); -lean_dec(x_22); -x_35 = l_stx_x21_____closed__3; -x_36 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_36, 0, x_33); -lean_ctor_set(x_36, 1, x_35); -x_37 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; -x_38 = lean_array_push(x_37, x_36); -x_39 = lean_array_push(x_38, x_20); -x_40 = lean_box(2); -x_41 = l_term_x21_____closed__2; -x_42 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_42, 0, x_40); -lean_ctor_set(x_42, 1, x_41); -lean_ctor_set(x_42, 2, x_39); -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_43, 1, x_34); -return x_43; +lean_inc(x_32); +lean_dec(x_21); +x_34 = l_stx_x21_____closed__3; +x_35 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_35, 0, x_32); +lean_ctor_set(x_35, 1, x_34); +x_36 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; +x_37 = lean_array_push(x_36, x_35); +x_38 = lean_array_push(x_37, x_19); +x_39 = lean_box(2); +x_40 = l_term_x21_____closed__2; +x_41 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +lean_ctor_set(x_41, 2, x_38); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_33); +return x_42; } } } @@ -20907,82 +20890,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___u2208_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___u2208_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___u2208_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___u2208_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___u2208_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___u2208_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___u2208_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___u2208_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -21593,82 +21575,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___x3a_x3a_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___x3a_x3a_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___x3a_x3a_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___x3a_x3a_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___x3a_x3a_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___x3a_x3a_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___x3a_x3a_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___x3a_x3a_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -22152,82 +22133,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___x3e_x3e_x3d_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___x3e_x3e_x3d_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___x3e_x3e_x3d_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___x3e_x3e_x3d_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___x3e_x3e_x3d_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___x3e_x3e_x3d_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___x3e_x3e_x3d_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___x3e_x3e_x3d_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -22904,263 +22884,262 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__11; -lean_inc(x_22); -x_24 = l_Lean_Syntax_isOfKind(x_22, x_23); -if (x_24 == 0) +x_22 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__11; +lean_inc(x_21); +x_23 = l_Lean_Syntax_isOfKind(x_21, x_22); +if (x_23 == 0) { -lean_object* x_25; lean_object* x_26; -lean_dec(x_22); +lean_object* x_24; lean_object* x_25; lean_dec(x_21); +lean_dec(x_20); lean_dec(x_9); -x_25 = lean_box(0); -x_26 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_3); -return x_26; +x_24 = lean_box(0); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_3); +return x_25; } else { -lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_27 = l_Lean_Syntax_getArg(x_22, x_14); -lean_dec(x_22); -x_28 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__13; -lean_inc(x_27); -x_29 = l_Lean_Syntax_isOfKind(x_27, x_28); -if (x_29 == 0) -{ -lean_object* x_30; lean_object* x_31; -lean_dec(x_27); +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = l_Lean_Syntax_getArg(x_21, x_14); lean_dec(x_21); +x_27 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__13; +lean_inc(x_26); +x_28 = l_Lean_Syntax_isOfKind(x_26, x_27); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; +lean_dec(x_26); +lean_dec(x_20); lean_dec(x_9); -x_30 = lean_box(0); -x_31 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_3); -return x_31; +x_29 = lean_box(0); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_3); +return x_30; } else { -lean_object* x_32; uint8_t x_33; -x_32 = l_Lean_Syntax_getArg(x_27, x_8); -lean_inc(x_32); -x_33 = l_Lean_Syntax_isNodeOf(x_32, x_16, x_14); -if (x_33 == 0) +lean_object* x_31; uint8_t x_32; +x_31 = l_Lean_Syntax_getArg(x_26, x_8); +lean_inc(x_31); +x_32 = l_Lean_Syntax_matchesNull(x_31, x_14); +if (x_32 == 0) { -lean_object* x_34; lean_object* x_35; -lean_dec(x_32); -lean_dec(x_27); -lean_dec(x_21); +lean_object* x_33; lean_object* x_34; +lean_dec(x_31); +lean_dec(x_26); +lean_dec(x_20); lean_dec(x_9); -x_34 = lean_box(0); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_3); -return x_35; +x_33 = lean_box(0); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_3); +return x_34; } else { -lean_object* x_36; lean_object* x_37; uint8_t x_38; -x_36 = l_Lean_Syntax_getArg(x_32, x_8); -lean_dec(x_32); -x_37 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__15; -lean_inc(x_36); -x_38 = l_Lean_Syntax_isOfKind(x_36, x_37); -if (x_38 == 0) +lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_35 = l_Lean_Syntax_getArg(x_31, x_8); +lean_dec(x_31); +x_36 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__15; +lean_inc(x_35); +x_37 = l_Lean_Syntax_isOfKind(x_35, x_36); +if (x_37 == 0) { -lean_object* x_39; lean_object* x_40; -lean_dec(x_36); -lean_dec(x_27); -lean_dec(x_21); +lean_object* x_38; lean_object* x_39; +lean_dec(x_35); +lean_dec(x_26); +lean_dec(x_20); lean_dec(x_9); -x_39 = lean_box(0); -x_40 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_3); -return x_40; +x_38 = lean_box(0); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_3); +return x_39; } else { -lean_object* x_41; uint8_t x_42; -x_41 = l_Lean_Syntax_getArg(x_36, x_8); -lean_inc(x_41); -x_42 = l_Lean_Syntax_isNodeOf(x_41, x_16, x_14); -if (x_42 == 0) +lean_object* x_40; uint8_t x_41; +x_40 = l_Lean_Syntax_getArg(x_35, x_8); +lean_inc(x_40); +x_41 = l_Lean_Syntax_matchesNull(x_40, x_14); +if (x_41 == 0) { -lean_object* x_43; lean_object* x_44; -lean_dec(x_41); -lean_dec(x_36); -lean_dec(x_27); -lean_dec(x_21); +lean_object* x_42; lean_object* x_43; +lean_dec(x_40); +lean_dec(x_35); +lean_dec(x_26); +lean_dec(x_20); lean_dec(x_9); -x_43 = lean_box(0); -x_44 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_44, 1, x_3); -return x_44; +x_42 = lean_box(0); +x_43 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_3); +return x_43; } else { -lean_object* x_45; lean_object* x_46; uint8_t x_47; -x_45 = l_Lean_Syntax_getArg(x_41, x_8); -lean_dec(x_41); -x_46 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__17; -x_47 = l_Lean_Syntax_isOfKind(x_45, x_46); -if (x_47 == 0) +lean_object* x_44; lean_object* x_45; uint8_t x_46; +x_44 = l_Lean_Syntax_getArg(x_40, x_8); +lean_dec(x_40); +x_45 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__17; +x_46 = l_Lean_Syntax_isOfKind(x_44, x_45); +if (x_46 == 0) { -lean_object* x_48; lean_object* x_49; -lean_dec(x_36); -lean_dec(x_27); -lean_dec(x_21); +lean_object* x_47; lean_object* x_48; +lean_dec(x_35); +lean_dec(x_26); +lean_dec(x_20); lean_dec(x_9); -x_48 = lean_box(0); -x_49 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set(x_49, 1, x_3); -return x_49; +x_47 = lean_box(0); +x_48 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_3); +return x_48; } else { -lean_object* x_50; uint8_t x_51; -x_50 = l_Lean_Syntax_getArg(x_36, x_14); -lean_dec(x_36); -lean_inc(x_50); -x_51 = l_Lean_Syntax_isNodeOf(x_50, x_16, x_14); -if (x_51 == 0) +lean_object* x_49; uint8_t x_50; +x_49 = l_Lean_Syntax_getArg(x_35, x_14); +lean_dec(x_35); +lean_inc(x_49); +x_50 = l_Lean_Syntax_matchesNull(x_49, x_14); +if (x_50 == 0) { -lean_object* x_52; lean_object* x_53; -lean_dec(x_50); -lean_dec(x_27); -lean_dec(x_21); +lean_object* x_51; lean_object* x_52; +lean_dec(x_49); +lean_dec(x_26); +lean_dec(x_20); lean_dec(x_9); -x_52 = lean_box(0); -x_53 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_53, 0, x_52); -lean_ctor_set(x_53, 1, x_3); -return x_53; +x_51 = lean_box(0); +x_52 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_3); +return x_52; } else { -lean_object* x_54; lean_object* x_55; uint8_t x_56; -x_54 = l_Lean_Syntax_getArg(x_50, x_8); -lean_dec(x_50); -x_55 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__20; -lean_inc(x_54); -x_56 = l_Lean_Syntax_isOfKind(x_54, x_55); -if (x_56 == 0) +lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_53 = l_Lean_Syntax_getArg(x_49, x_8); +lean_dec(x_49); +x_54 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__20; +lean_inc(x_53); +x_55 = l_Lean_Syntax_isOfKind(x_53, x_54); +if (x_55 == 0) { -lean_object* x_57; lean_object* x_58; -lean_dec(x_54); -lean_dec(x_27); -lean_dec(x_21); +lean_object* x_56; lean_object* x_57; +lean_dec(x_53); +lean_dec(x_26); +lean_dec(x_20); lean_dec(x_9); -x_57 = lean_box(0); -x_58 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_58, 0, x_57); -lean_ctor_set(x_58, 1, x_3); -return x_58; +x_56 = lean_box(0); +x_57 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_3); +return x_57; } else { -lean_object* x_59; lean_object* x_60; uint8_t x_61; -x_59 = l_Lean_Syntax_getArg(x_54, x_14); -lean_dec(x_54); -x_60 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__25; -x_61 = l_Lean_Syntax_matchesIdent(x_59, x_60); -lean_dec(x_59); -if (x_61 == 0) +lean_object* x_58; lean_object* x_59; uint8_t x_60; +x_58 = l_Lean_Syntax_getArg(x_53, x_14); +lean_dec(x_53); +x_59 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__25; +x_60 = l_Lean_Syntax_matchesIdent(x_58, x_59); +lean_dec(x_58); +if (x_60 == 0) { -lean_object* x_62; lean_object* x_63; -lean_dec(x_27); -lean_dec(x_21); +lean_object* x_61; lean_object* x_62; +lean_dec(x_26); +lean_dec(x_20); lean_dec(x_9); -x_62 = lean_box(0); -x_63 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_63, 0, x_62); -lean_ctor_set(x_63, 1, x_3); -return x_63; +x_61 = lean_box(0); +x_62 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_62, 0, x_61); +lean_ctor_set(x_62, 1, x_3); +return x_62; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; -x_64 = l_Lean_Syntax_getArg(x_27, x_17); -lean_dec(x_27); -x_65 = l_Lean_replaceRef(x_9, x_2); +lean_object* x_63; lean_object* x_64; lean_object* x_65; uint8_t x_66; +x_63 = l_Lean_Syntax_getArg(x_26, x_16); +lean_dec(x_26); +x_64 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_66 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_65, x_3); -x_67 = !lean_is_exclusive(x_66); -if (x_67 == 0) +x_65 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_64, x_3); +x_66 = !lean_is_exclusive(x_65); +if (x_66 == 0) +{ +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; +x_67 = lean_ctor_get(x_65, 0); +x_68 = l_term___x3c_x2a_x3e_____closed__3; +x_69 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_69, 0, x_67); +lean_ctor_set(x_69, 1, x_68); +x_70 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_71 = lean_array_push(x_70, x_20); +x_72 = lean_array_push(x_71, x_69); +x_73 = lean_array_push(x_72, x_63); +x_74 = lean_box(2); +x_75 = l_term___x3c_x2a_x3e_____closed__2; +x_76 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_76, 0, x_74); +lean_ctor_set(x_76, 1, x_75); +lean_ctor_set(x_76, 2, x_73); +lean_ctor_set(x_65, 0, x_76); +return x_65; +} +else { -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; -x_68 = lean_ctor_get(x_66, 0); -x_69 = l_term___x3c_x2a_x3e_____closed__3; -x_70 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_69); -x_71 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_72 = lean_array_push(x_71, x_21); -x_73 = lean_array_push(x_72, x_70); -x_74 = lean_array_push(x_73, x_64); -x_75 = lean_box(2); -x_76 = l_term___x3c_x2a_x3e_____closed__2; -x_77 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_77, 0, x_75); -lean_ctor_set(x_77, 1, x_76); -lean_ctor_set(x_77, 2, x_74); -lean_ctor_set(x_66, 0, x_77); -return x_66; -} -else -{ -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_78 = lean_ctor_get(x_66, 0); -x_79 = lean_ctor_get(x_66, 1); -lean_inc(x_79); +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_77 = lean_ctor_get(x_65, 0); +x_78 = lean_ctor_get(x_65, 1); lean_inc(x_78); -lean_dec(x_66); -x_80 = l_term___x3c_x2a_x3e_____closed__3; -x_81 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_81, 0, x_78); -lean_ctor_set(x_81, 1, x_80); -x_82 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_83 = lean_array_push(x_82, x_21); -x_84 = lean_array_push(x_83, x_81); -x_85 = lean_array_push(x_84, x_64); -x_86 = lean_box(2); -x_87 = l_term___x3c_x2a_x3e_____closed__2; -x_88 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_88, 0, x_86); -lean_ctor_set(x_88, 1, x_87); -lean_ctor_set(x_88, 2, x_85); -x_89 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_89, 0, x_88); -lean_ctor_set(x_89, 1, x_79); -return x_89; +lean_inc(x_77); +lean_dec(x_65); +x_79 = l_term___x3c_x2a_x3e_____closed__3; +x_80 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_80, 0, x_77); +lean_ctor_set(x_80, 1, x_79); +x_81 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_82 = lean_array_push(x_81, x_20); +x_83 = lean_array_push(x_82, x_80); +x_84 = lean_array_push(x_83, x_63); +x_85 = lean_box(2); +x_86 = l_term___x3c_x2a_x3e_____closed__2; +x_87 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_87, 0, x_85); +lean_ctor_set(x_87, 1, x_86); +lean_ctor_set(x_87, 2, x_84); +x_88 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_88, 0, x_87); +lean_ctor_set(x_88, 1, x_78); +return x_88; } } } @@ -23667,263 +23646,262 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__11; -lean_inc(x_22); -x_24 = l_Lean_Syntax_isOfKind(x_22, x_23); -if (x_24 == 0) +x_22 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__11; +lean_inc(x_21); +x_23 = l_Lean_Syntax_isOfKind(x_21, x_22); +if (x_23 == 0) { -lean_object* x_25; lean_object* x_26; -lean_dec(x_22); +lean_object* x_24; lean_object* x_25; lean_dec(x_21); +lean_dec(x_20); lean_dec(x_9); -x_25 = lean_box(0); -x_26 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_3); -return x_26; +x_24 = lean_box(0); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_3); +return x_25; } else { -lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_27 = l_Lean_Syntax_getArg(x_22, x_14); -lean_dec(x_22); -x_28 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__13; -lean_inc(x_27); -x_29 = l_Lean_Syntax_isOfKind(x_27, x_28); -if (x_29 == 0) -{ -lean_object* x_30; lean_object* x_31; -lean_dec(x_27); +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = l_Lean_Syntax_getArg(x_21, x_14); lean_dec(x_21); +x_27 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__13; +lean_inc(x_26); +x_28 = l_Lean_Syntax_isOfKind(x_26, x_27); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; +lean_dec(x_26); +lean_dec(x_20); lean_dec(x_9); -x_30 = lean_box(0); -x_31 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_3); -return x_31; +x_29 = lean_box(0); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_3); +return x_30; } else { -lean_object* x_32; uint8_t x_33; -x_32 = l_Lean_Syntax_getArg(x_27, x_8); -lean_inc(x_32); -x_33 = l_Lean_Syntax_isNodeOf(x_32, x_16, x_14); -if (x_33 == 0) +lean_object* x_31; uint8_t x_32; +x_31 = l_Lean_Syntax_getArg(x_26, x_8); +lean_inc(x_31); +x_32 = l_Lean_Syntax_matchesNull(x_31, x_14); +if (x_32 == 0) { -lean_object* x_34; lean_object* x_35; -lean_dec(x_32); -lean_dec(x_27); -lean_dec(x_21); +lean_object* x_33; lean_object* x_34; +lean_dec(x_31); +lean_dec(x_26); +lean_dec(x_20); lean_dec(x_9); -x_34 = lean_box(0); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_3); -return x_35; +x_33 = lean_box(0); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_3); +return x_34; } else { -lean_object* x_36; lean_object* x_37; uint8_t x_38; -x_36 = l_Lean_Syntax_getArg(x_32, x_8); -lean_dec(x_32); -x_37 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__15; -lean_inc(x_36); -x_38 = l_Lean_Syntax_isOfKind(x_36, x_37); -if (x_38 == 0) +lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_35 = l_Lean_Syntax_getArg(x_31, x_8); +lean_dec(x_31); +x_36 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__15; +lean_inc(x_35); +x_37 = l_Lean_Syntax_isOfKind(x_35, x_36); +if (x_37 == 0) { -lean_object* x_39; lean_object* x_40; -lean_dec(x_36); -lean_dec(x_27); -lean_dec(x_21); +lean_object* x_38; lean_object* x_39; +lean_dec(x_35); +lean_dec(x_26); +lean_dec(x_20); lean_dec(x_9); -x_39 = lean_box(0); -x_40 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_3); -return x_40; +x_38 = lean_box(0); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_3); +return x_39; } else { -lean_object* x_41; uint8_t x_42; -x_41 = l_Lean_Syntax_getArg(x_36, x_8); -lean_inc(x_41); -x_42 = l_Lean_Syntax_isNodeOf(x_41, x_16, x_14); -if (x_42 == 0) +lean_object* x_40; uint8_t x_41; +x_40 = l_Lean_Syntax_getArg(x_35, x_8); +lean_inc(x_40); +x_41 = l_Lean_Syntax_matchesNull(x_40, x_14); +if (x_41 == 0) { -lean_object* x_43; lean_object* x_44; -lean_dec(x_41); -lean_dec(x_36); -lean_dec(x_27); -lean_dec(x_21); +lean_object* x_42; lean_object* x_43; +lean_dec(x_40); +lean_dec(x_35); +lean_dec(x_26); +lean_dec(x_20); lean_dec(x_9); -x_43 = lean_box(0); -x_44 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_44, 1, x_3); -return x_44; +x_42 = lean_box(0); +x_43 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_3); +return x_43; } else { -lean_object* x_45; lean_object* x_46; uint8_t x_47; -x_45 = l_Lean_Syntax_getArg(x_41, x_8); -lean_dec(x_41); -x_46 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__17; -x_47 = l_Lean_Syntax_isOfKind(x_45, x_46); -if (x_47 == 0) +lean_object* x_44; lean_object* x_45; uint8_t x_46; +x_44 = l_Lean_Syntax_getArg(x_40, x_8); +lean_dec(x_40); +x_45 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__17; +x_46 = l_Lean_Syntax_isOfKind(x_44, x_45); +if (x_46 == 0) { -lean_object* x_48; lean_object* x_49; -lean_dec(x_36); -lean_dec(x_27); -lean_dec(x_21); +lean_object* x_47; lean_object* x_48; +lean_dec(x_35); +lean_dec(x_26); +lean_dec(x_20); lean_dec(x_9); -x_48 = lean_box(0); -x_49 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set(x_49, 1, x_3); -return x_49; +x_47 = lean_box(0); +x_48 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_3); +return x_48; } else { -lean_object* x_50; uint8_t x_51; -x_50 = l_Lean_Syntax_getArg(x_36, x_14); -lean_dec(x_36); -lean_inc(x_50); -x_51 = l_Lean_Syntax_isNodeOf(x_50, x_16, x_14); -if (x_51 == 0) +lean_object* x_49; uint8_t x_50; +x_49 = l_Lean_Syntax_getArg(x_35, x_14); +lean_dec(x_35); +lean_inc(x_49); +x_50 = l_Lean_Syntax_matchesNull(x_49, x_14); +if (x_50 == 0) { -lean_object* x_52; lean_object* x_53; -lean_dec(x_50); -lean_dec(x_27); -lean_dec(x_21); +lean_object* x_51; lean_object* x_52; +lean_dec(x_49); +lean_dec(x_26); +lean_dec(x_20); lean_dec(x_9); -x_52 = lean_box(0); -x_53 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_53, 0, x_52); -lean_ctor_set(x_53, 1, x_3); -return x_53; +x_51 = lean_box(0); +x_52 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_3); +return x_52; } else { -lean_object* x_54; lean_object* x_55; uint8_t x_56; -x_54 = l_Lean_Syntax_getArg(x_50, x_8); -lean_dec(x_50); -x_55 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__20; -lean_inc(x_54); -x_56 = l_Lean_Syntax_isOfKind(x_54, x_55); -if (x_56 == 0) +lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_53 = l_Lean_Syntax_getArg(x_49, x_8); +lean_dec(x_49); +x_54 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__20; +lean_inc(x_53); +x_55 = l_Lean_Syntax_isOfKind(x_53, x_54); +if (x_55 == 0) { -lean_object* x_57; lean_object* x_58; -lean_dec(x_54); -lean_dec(x_27); -lean_dec(x_21); +lean_object* x_56; lean_object* x_57; +lean_dec(x_53); +lean_dec(x_26); +lean_dec(x_20); lean_dec(x_9); -x_57 = lean_box(0); -x_58 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_58, 0, x_57); -lean_ctor_set(x_58, 1, x_3); -return x_58; +x_56 = lean_box(0); +x_57 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_3); +return x_57; } else { -lean_object* x_59; lean_object* x_60; uint8_t x_61; -x_59 = l_Lean_Syntax_getArg(x_54, x_14); -lean_dec(x_54); -x_60 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__25; -x_61 = l_Lean_Syntax_matchesIdent(x_59, x_60); -lean_dec(x_59); -if (x_61 == 0) +lean_object* x_58; lean_object* x_59; uint8_t x_60; +x_58 = l_Lean_Syntax_getArg(x_53, x_14); +lean_dec(x_53); +x_59 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__25; +x_60 = l_Lean_Syntax_matchesIdent(x_58, x_59); +lean_dec(x_58); +if (x_60 == 0) { -lean_object* x_62; lean_object* x_63; -lean_dec(x_27); -lean_dec(x_21); +lean_object* x_61; lean_object* x_62; +lean_dec(x_26); +lean_dec(x_20); lean_dec(x_9); -x_62 = lean_box(0); -x_63 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_63, 0, x_62); -lean_ctor_set(x_63, 1, x_3); -return x_63; +x_61 = lean_box(0); +x_62 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_62, 0, x_61); +lean_ctor_set(x_62, 1, x_3); +return x_62; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; -x_64 = l_Lean_Syntax_getArg(x_27, x_17); -lean_dec(x_27); -x_65 = l_Lean_replaceRef(x_9, x_2); +lean_object* x_63; lean_object* x_64; lean_object* x_65; uint8_t x_66; +x_63 = l_Lean_Syntax_getArg(x_26, x_16); +lean_dec(x_26); +x_64 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_66 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_65, x_3); -x_67 = !lean_is_exclusive(x_66); -if (x_67 == 0) -{ -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; -x_68 = lean_ctor_get(x_66, 0); -x_69 = l_term___x3c_x2a_____closed__3; -x_70 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_69); -x_71 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_72 = lean_array_push(x_71, x_21); -x_73 = lean_array_push(x_72, x_70); -x_74 = lean_array_push(x_73, x_64); -x_75 = lean_box(2); -x_76 = l_term___x3c_x2a_____closed__2; -x_77 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_77, 0, x_75); -lean_ctor_set(x_77, 1, x_76); -lean_ctor_set(x_77, 2, x_74); -lean_ctor_set(x_66, 0, x_77); -return x_66; +x_65 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_64, x_3); +x_66 = !lean_is_exclusive(x_65); +if (x_66 == 0) +{ +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; +x_67 = lean_ctor_get(x_65, 0); +x_68 = l_term___x3c_x2a_____closed__3; +x_69 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_69, 0, x_67); +lean_ctor_set(x_69, 1, x_68); +x_70 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_71 = lean_array_push(x_70, x_20); +x_72 = lean_array_push(x_71, x_69); +x_73 = lean_array_push(x_72, x_63); +x_74 = lean_box(2); +x_75 = l_term___x3c_x2a_____closed__2; +x_76 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_76, 0, x_74); +lean_ctor_set(x_76, 1, x_75); +lean_ctor_set(x_76, 2, x_73); +lean_ctor_set(x_65, 0, x_76); +return x_65; } else { -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_78 = lean_ctor_get(x_66, 0); -x_79 = lean_ctor_get(x_66, 1); -lean_inc(x_79); +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_77 = lean_ctor_get(x_65, 0); +x_78 = lean_ctor_get(x_65, 1); lean_inc(x_78); -lean_dec(x_66); -x_80 = l_term___x3c_x2a_____closed__3; -x_81 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_81, 0, x_78); -lean_ctor_set(x_81, 1, x_80); -x_82 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_83 = lean_array_push(x_82, x_21); -x_84 = lean_array_push(x_83, x_81); -x_85 = lean_array_push(x_84, x_64); -x_86 = lean_box(2); -x_87 = l_term___x3c_x2a_____closed__2; -x_88 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_88, 0, x_86); -lean_ctor_set(x_88, 1, x_87); -lean_ctor_set(x_88, 2, x_85); -x_89 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_89, 0, x_88); -lean_ctor_set(x_89, 1, x_79); -return x_89; +lean_inc(x_77); +lean_dec(x_65); +x_79 = l_term___x3c_x2a_____closed__3; +x_80 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_80, 0, x_77); +lean_ctor_set(x_80, 1, x_79); +x_81 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_82 = lean_array_push(x_81, x_20); +x_83 = lean_array_push(x_82, x_80); +x_84 = lean_array_push(x_83, x_63); +x_85 = lean_box(2); +x_86 = l_term___x3c_x2a_____closed__2; +x_87 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_87, 0, x_85); +lean_ctor_set(x_87, 1, x_86); +lean_ctor_set(x_87, 2, x_84); +x_88 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_88, 0, x_87); +lean_ctor_set(x_88, 1, x_78); +return x_88; } } } @@ -24430,263 +24408,262 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__11; -lean_inc(x_22); -x_24 = l_Lean_Syntax_isOfKind(x_22, x_23); -if (x_24 == 0) +x_22 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__11; +lean_inc(x_21); +x_23 = l_Lean_Syntax_isOfKind(x_21, x_22); +if (x_23 == 0) { -lean_object* x_25; lean_object* x_26; -lean_dec(x_22); +lean_object* x_24; lean_object* x_25; lean_dec(x_21); +lean_dec(x_20); lean_dec(x_9); -x_25 = lean_box(0); -x_26 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_3); -return x_26; +x_24 = lean_box(0); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_3); +return x_25; } else { -lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_27 = l_Lean_Syntax_getArg(x_22, x_14); -lean_dec(x_22); -x_28 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__13; -lean_inc(x_27); -x_29 = l_Lean_Syntax_isOfKind(x_27, x_28); -if (x_29 == 0) -{ -lean_object* x_30; lean_object* x_31; -lean_dec(x_27); +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = l_Lean_Syntax_getArg(x_21, x_14); lean_dec(x_21); +x_27 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__13; +lean_inc(x_26); +x_28 = l_Lean_Syntax_isOfKind(x_26, x_27); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; +lean_dec(x_26); +lean_dec(x_20); lean_dec(x_9); -x_30 = lean_box(0); -x_31 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_3); -return x_31; +x_29 = lean_box(0); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_3); +return x_30; } else { -lean_object* x_32; uint8_t x_33; -x_32 = l_Lean_Syntax_getArg(x_27, x_8); -lean_inc(x_32); -x_33 = l_Lean_Syntax_isNodeOf(x_32, x_16, x_14); -if (x_33 == 0) +lean_object* x_31; uint8_t x_32; +x_31 = l_Lean_Syntax_getArg(x_26, x_8); +lean_inc(x_31); +x_32 = l_Lean_Syntax_matchesNull(x_31, x_14); +if (x_32 == 0) { -lean_object* x_34; lean_object* x_35; -lean_dec(x_32); -lean_dec(x_27); -lean_dec(x_21); +lean_object* x_33; lean_object* x_34; +lean_dec(x_31); +lean_dec(x_26); +lean_dec(x_20); lean_dec(x_9); -x_34 = lean_box(0); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_3); -return x_35; +x_33 = lean_box(0); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_3); +return x_34; } else { -lean_object* x_36; lean_object* x_37; uint8_t x_38; -x_36 = l_Lean_Syntax_getArg(x_32, x_8); -lean_dec(x_32); -x_37 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__15; -lean_inc(x_36); -x_38 = l_Lean_Syntax_isOfKind(x_36, x_37); -if (x_38 == 0) +lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_35 = l_Lean_Syntax_getArg(x_31, x_8); +lean_dec(x_31); +x_36 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__15; +lean_inc(x_35); +x_37 = l_Lean_Syntax_isOfKind(x_35, x_36); +if (x_37 == 0) { -lean_object* x_39; lean_object* x_40; -lean_dec(x_36); -lean_dec(x_27); -lean_dec(x_21); +lean_object* x_38; lean_object* x_39; +lean_dec(x_35); +lean_dec(x_26); +lean_dec(x_20); lean_dec(x_9); -x_39 = lean_box(0); -x_40 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_3); -return x_40; +x_38 = lean_box(0); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_3); +return x_39; } else { -lean_object* x_41; uint8_t x_42; -x_41 = l_Lean_Syntax_getArg(x_36, x_8); -lean_inc(x_41); -x_42 = l_Lean_Syntax_isNodeOf(x_41, x_16, x_14); -if (x_42 == 0) +lean_object* x_40; uint8_t x_41; +x_40 = l_Lean_Syntax_getArg(x_35, x_8); +lean_inc(x_40); +x_41 = l_Lean_Syntax_matchesNull(x_40, x_14); +if (x_41 == 0) { -lean_object* x_43; lean_object* x_44; -lean_dec(x_41); -lean_dec(x_36); -lean_dec(x_27); -lean_dec(x_21); +lean_object* x_42; lean_object* x_43; +lean_dec(x_40); +lean_dec(x_35); +lean_dec(x_26); +lean_dec(x_20); lean_dec(x_9); -x_43 = lean_box(0); -x_44 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_44, 1, x_3); -return x_44; +x_42 = lean_box(0); +x_43 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_3); +return x_43; } else { -lean_object* x_45; lean_object* x_46; uint8_t x_47; -x_45 = l_Lean_Syntax_getArg(x_41, x_8); -lean_dec(x_41); -x_46 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__17; -x_47 = l_Lean_Syntax_isOfKind(x_45, x_46); -if (x_47 == 0) +lean_object* x_44; lean_object* x_45; uint8_t x_46; +x_44 = l_Lean_Syntax_getArg(x_40, x_8); +lean_dec(x_40); +x_45 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__17; +x_46 = l_Lean_Syntax_isOfKind(x_44, x_45); +if (x_46 == 0) { -lean_object* x_48; lean_object* x_49; -lean_dec(x_36); -lean_dec(x_27); -lean_dec(x_21); +lean_object* x_47; lean_object* x_48; +lean_dec(x_35); +lean_dec(x_26); +lean_dec(x_20); lean_dec(x_9); -x_48 = lean_box(0); -x_49 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set(x_49, 1, x_3); -return x_49; +x_47 = lean_box(0); +x_48 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_3); +return x_48; } else { -lean_object* x_50; uint8_t x_51; -x_50 = l_Lean_Syntax_getArg(x_36, x_14); -lean_dec(x_36); -lean_inc(x_50); -x_51 = l_Lean_Syntax_isNodeOf(x_50, x_16, x_14); -if (x_51 == 0) +lean_object* x_49; uint8_t x_50; +x_49 = l_Lean_Syntax_getArg(x_35, x_14); +lean_dec(x_35); +lean_inc(x_49); +x_50 = l_Lean_Syntax_matchesNull(x_49, x_14); +if (x_50 == 0) { -lean_object* x_52; lean_object* x_53; -lean_dec(x_50); -lean_dec(x_27); -lean_dec(x_21); +lean_object* x_51; lean_object* x_52; +lean_dec(x_49); +lean_dec(x_26); +lean_dec(x_20); lean_dec(x_9); -x_52 = lean_box(0); -x_53 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_53, 0, x_52); -lean_ctor_set(x_53, 1, x_3); -return x_53; +x_51 = lean_box(0); +x_52 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_3); +return x_52; } else { -lean_object* x_54; lean_object* x_55; uint8_t x_56; -x_54 = l_Lean_Syntax_getArg(x_50, x_8); -lean_dec(x_50); -x_55 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__20; -lean_inc(x_54); -x_56 = l_Lean_Syntax_isOfKind(x_54, x_55); -if (x_56 == 0) +lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_53 = l_Lean_Syntax_getArg(x_49, x_8); +lean_dec(x_49); +x_54 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__20; +lean_inc(x_53); +x_55 = l_Lean_Syntax_isOfKind(x_53, x_54); +if (x_55 == 0) { -lean_object* x_57; lean_object* x_58; -lean_dec(x_54); -lean_dec(x_27); -lean_dec(x_21); +lean_object* x_56; lean_object* x_57; +lean_dec(x_53); +lean_dec(x_26); +lean_dec(x_20); lean_dec(x_9); -x_57 = lean_box(0); -x_58 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_58, 0, x_57); -lean_ctor_set(x_58, 1, x_3); -return x_58; +x_56 = lean_box(0); +x_57 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_3); +return x_57; } else { -lean_object* x_59; lean_object* x_60; uint8_t x_61; -x_59 = l_Lean_Syntax_getArg(x_54, x_14); -lean_dec(x_54); -x_60 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__25; -x_61 = l_Lean_Syntax_matchesIdent(x_59, x_60); -lean_dec(x_59); -if (x_61 == 0) +lean_object* x_58; lean_object* x_59; uint8_t x_60; +x_58 = l_Lean_Syntax_getArg(x_53, x_14); +lean_dec(x_53); +x_59 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__25; +x_60 = l_Lean_Syntax_matchesIdent(x_58, x_59); +lean_dec(x_58); +if (x_60 == 0) { -lean_object* x_62; lean_object* x_63; -lean_dec(x_27); -lean_dec(x_21); +lean_object* x_61; lean_object* x_62; +lean_dec(x_26); +lean_dec(x_20); lean_dec(x_9); -x_62 = lean_box(0); -x_63 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_63, 0, x_62); -lean_ctor_set(x_63, 1, x_3); -return x_63; +x_61 = lean_box(0); +x_62 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_62, 0, x_61); +lean_ctor_set(x_62, 1, x_3); +return x_62; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; -x_64 = l_Lean_Syntax_getArg(x_27, x_17); -lean_dec(x_27); -x_65 = l_Lean_replaceRef(x_9, x_2); +lean_object* x_63; lean_object* x_64; lean_object* x_65; uint8_t x_66; +x_63 = l_Lean_Syntax_getArg(x_26, x_16); +lean_dec(x_26); +x_64 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_66 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_65, x_3); -x_67 = !lean_is_exclusive(x_66); -if (x_67 == 0) -{ -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; -x_68 = lean_ctor_get(x_66, 0); -x_69 = l_term___x2a_x3e_____closed__3; -x_70 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_69); -x_71 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_72 = lean_array_push(x_71, x_21); -x_73 = lean_array_push(x_72, x_70); -x_74 = lean_array_push(x_73, x_64); -x_75 = lean_box(2); -x_76 = l_term___x2a_x3e_____closed__2; -x_77 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_77, 0, x_75); -lean_ctor_set(x_77, 1, x_76); -lean_ctor_set(x_77, 2, x_74); -lean_ctor_set(x_66, 0, x_77); -return x_66; +x_65 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_64, x_3); +x_66 = !lean_is_exclusive(x_65); +if (x_66 == 0) +{ +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; +x_67 = lean_ctor_get(x_65, 0); +x_68 = l_term___x2a_x3e_____closed__3; +x_69 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_69, 0, x_67); +lean_ctor_set(x_69, 1, x_68); +x_70 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_71 = lean_array_push(x_70, x_20); +x_72 = lean_array_push(x_71, x_69); +x_73 = lean_array_push(x_72, x_63); +x_74 = lean_box(2); +x_75 = l_term___x2a_x3e_____closed__2; +x_76 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_76, 0, x_74); +lean_ctor_set(x_76, 1, x_75); +lean_ctor_set(x_76, 2, x_73); +lean_ctor_set(x_65, 0, x_76); +return x_65; } else { -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_78 = lean_ctor_get(x_66, 0); -x_79 = lean_ctor_get(x_66, 1); -lean_inc(x_79); +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_77 = lean_ctor_get(x_65, 0); +x_78 = lean_ctor_get(x_65, 1); lean_inc(x_78); -lean_dec(x_66); -x_80 = l_term___x2a_x3e_____closed__3; -x_81 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_81, 0, x_78); -lean_ctor_set(x_81, 1, x_80); -x_82 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_83 = lean_array_push(x_82, x_21); -x_84 = lean_array_push(x_83, x_81); -x_85 = lean_array_push(x_84, x_64); -x_86 = lean_box(2); -x_87 = l_term___x2a_x3e_____closed__2; -x_88 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_88, 0, x_86); -lean_ctor_set(x_88, 1, x_87); -lean_ctor_set(x_88, 2, x_85); -x_89 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_89, 0, x_88); -lean_ctor_set(x_89, 1, x_79); -return x_89; +lean_inc(x_77); +lean_dec(x_65); +x_79 = l_term___x2a_x3e_____closed__3; +x_80 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_80, 0, x_77); +lean_ctor_set(x_80, 1, x_79); +x_81 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_82 = lean_array_push(x_81, x_20); +x_83 = lean_array_push(x_82, x_80); +x_84 = lean_array_push(x_83, x_63); +x_85 = lean_box(2); +x_86 = l_term___x2a_x3e_____closed__2; +x_87 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_87, 0, x_85); +lean_ctor_set(x_87, 1, x_86); +lean_ctor_set(x_87, 2, x_84); +x_88 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_88, 0, x_87); +lean_ctor_set(x_88, 1, x_78); +return x_88; } } } @@ -25026,82 +25003,81 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(2u); +x_16 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; +lean_object* x_18; lean_object* x_19; lean_dec(x_15); lean_dec(x_9); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = l_Lean_Syntax_getArg(x_15, x_8); -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = l_Lean_Syntax_getArg(x_15, x_8); +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = l_Lean_replaceRef(x_9, x_2); +x_22 = l_Lean_replaceRef(x_9, x_2); lean_dec(x_9); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_23, x_3); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_22, x_3); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_26 = lean_ctor_get(x_24, 0); -x_27 = l_term___x3c_x24_x3e_____closed__3; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_30 = lean_array_push(x_29, x_21); -x_31 = lean_array_push(x_30, x_28); -x_32 = lean_array_push(x_31, x_22); -x_33 = lean_box(2); -x_34 = l_term___x3c_x24_x3e_____closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_24, 0, x_35); -return 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_23, 0); +x_26 = l_term___x3c_x24_x3e_____closed__3; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_29 = lean_array_push(x_28, x_20); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_array_push(x_30, x_21); +x_32 = lean_box(2); +x_33 = l_term___x3c_x24_x3e_____closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_36 = lean_ctor_get(x_24, 0); -x_37 = lean_ctor_get(x_24, 1); -lean_inc(x_37); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 1); lean_inc(x_36); -lean_dec(x_24); -x_38 = l_term___x3c_x24_x3e_____closed__3; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_41 = lean_array_push(x_40, x_21); -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_42, x_22); -x_44 = lean_box(2); -x_45 = l_term___x3c_x24_x3e_____closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_37); -return x_47; +lean_inc(x_35); +lean_dec(x_23); +x_37 = l_term___x3c_x24_x3e_____closed__3; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_40 = lean_array_push(x_39, x_20); +x_41 = lean_array_push(x_40, x_38); +x_42 = lean_array_push(x_41, x_21); +x_43 = lean_box(2); +x_44 = l_term___x3c_x24_x3e_____closed__2; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_36); +return x_46; } } } @@ -26190,431 +26166,412 @@ return x_7; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +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; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_8 = lean_unsigned_to_nat(1u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); -x_10 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__2; -lean_inc(x_9); -x_11 = l_Lean_Syntax_isOfKind(x_9, x_10); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; -lean_dec(x_9); -lean_dec(x_2); -lean_dec(x_1); -x_12 = lean_box(1); -x_13 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_3); -return x_13; -} -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; lean_object* x_20; uint8_t x_21; -x_14 = lean_unsigned_to_nat(3u); +x_10 = lean_unsigned_to_nat(3u); +x_11 = l_Lean_Syntax_getArg(x_1, x_10); +x_12 = lean_unsigned_to_nat(5u); +x_13 = l_Lean_Syntax_getArg(x_1, x_12); +x_14 = lean_unsigned_to_nat(7u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); -x_16 = lean_unsigned_to_nat(5u); -x_17 = l_Lean_Syntax_getArg(x_1, x_16); -x_18 = lean_unsigned_to_nat(7u); -x_19 = l_Lean_Syntax_getArg(x_1, x_18); lean_dec(x_1); lean_inc(x_2); -x_20 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_21 = !lean_is_exclusive(x_20); -if (x_21 == 0) +x_16 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) { -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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; -x_22 = lean_ctor_get(x_20, 0); -x_23 = lean_ctor_get(x_2, 2); -lean_inc(x_23); -x_24 = lean_ctor_get(x_2, 1); -lean_inc(x_24); +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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; +x_18 = lean_ctor_get(x_16, 0); +x_19 = lean_ctor_get(x_2, 2); +lean_inc(x_19); +x_20 = lean_ctor_get(x_2, 1); +lean_inc(x_20); lean_dec(x_2); -x_25 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__3; -lean_inc(x_22); -x_26 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_26, 0, x_22); -lean_ctor_set(x_26, 1, x_25); -x_27 = l_stx___x3f___closed__3; -lean_inc(x_22); -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_22); -lean_ctor_set(x_28, 1, x_27); -x_29 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__7; -lean_inc(x_23); -lean_inc(x_24); -x_30 = l_Lean_addMacroScope(x_24, x_29, x_23); -x_31 = lean_box(0); -x_32 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__6; -lean_inc(x_22); -x_33 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_33, 0, x_22); +x_21 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__3; +lean_inc(x_18); +x_22 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_22, 0, x_18); +lean_ctor_set(x_22, 1, x_21); +x_23 = l_stx___x3f___closed__3; +lean_inc(x_18); +x_24 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_24, 0, x_18); +lean_ctor_set(x_24, 1, x_23); +x_25 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__7; +lean_inc(x_19); +lean_inc(x_20); +x_26 = l_Lean_addMacroScope(x_20, x_25, x_19); +x_27 = lean_box(0); +x_28 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__6; +lean_inc(x_18); +x_29 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_29, 0, x_18); +lean_ctor_set(x_29, 1, x_28); +lean_ctor_set(x_29, 2, x_26); +lean_ctor_set(x_29, 3, x_27); +x_30 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__8; +lean_inc(x_18); +x_31 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_31, 0, x_18); +lean_ctor_set(x_31, 1, x_30); +x_32 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__9; +lean_inc(x_18); +x_33 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_33, 0, x_18); lean_ctor_set(x_33, 1, x_32); -lean_ctor_set(x_33, 2, x_30); -lean_ctor_set(x_33, 3, x_31); -x_34 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__8; -lean_inc(x_22); +x_34 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__12; +lean_inc(x_18); x_35 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_35, 0, x_22); +lean_ctor_set(x_35, 0, x_18); lean_ctor_set(x_35, 1, x_34); -x_36 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__9; -lean_inc(x_22); -x_37 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_37, 0, x_22); -lean_ctor_set(x_37, 1, x_36); -x_38 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__12; -lean_inc(x_22); -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_22); -lean_ctor_set(x_39, 1, x_38); -x_40 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__16; -x_41 = l_Lean_addMacroScope(x_24, x_40, x_23); -x_42 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__15; -x_43 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__18; -lean_inc(x_22); -x_44 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_44, 0, x_22); -lean_ctor_set(x_44, 1, x_42); -lean_ctor_set(x_44, 2, x_41); -lean_ctor_set(x_44, 3, x_43); -x_45 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; -lean_inc(x_28); -x_46 = lean_array_push(x_45, x_28); -lean_inc(x_33); -x_47 = lean_array_push(x_46, x_33); -x_48 = lean_box(2); -x_49 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__20; -x_50 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_50, 0, x_48); +x_36 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__16; +x_37 = l_Lean_addMacroScope(x_20, x_36, x_19); +x_38 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__15; +x_39 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__18; +lean_inc(x_18); +x_40 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_40, 0, x_18); +lean_ctor_set(x_40, 1, x_38); +lean_ctor_set(x_40, 2, x_37); +lean_ctor_set(x_40, 3, x_39); +x_41 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; +lean_inc(x_24); +x_42 = lean_array_push(x_41, x_24); +lean_inc(x_29); +x_43 = lean_array_push(x_42, x_29); +x_44 = lean_box(2); +x_45 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__20; +x_46 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +lean_ctor_set(x_46, 2, x_43); +x_47 = l_prec_x28___x29___closed__3; +lean_inc(x_18); +x_48 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_48, 0, x_18); +lean_ctor_set(x_48, 1, x_47); +x_49 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__10; +lean_inc(x_18); +x_50 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_50, 0, x_18); lean_ctor_set(x_50, 1, x_49); -lean_ctor_set(x_50, 2, x_47); -x_51 = l_prec_x28___x29___closed__3; -lean_inc(x_22); -x_52 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_52, 0, x_22); -lean_ctor_set(x_52, 1, x_51); -x_53 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__10; -lean_inc(x_22); -x_54 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_54, 0, x_22); +x_51 = l___aux__Init__Notation______macroRules__precMax__1___closed__4; +x_52 = lean_array_push(x_51, x_9); +x_53 = l___aux__Init__Notation______macroRules__stx___x2b__1___closed__8; +x_54 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_54, 0, x_44); lean_ctor_set(x_54, 1, x_53); -x_55 = l___aux__Init__Notation______macroRules__precMax__1___closed__4; -x_56 = lean_array_push(x_55, x_9); -x_57 = l___aux__Init__Notation______macroRules__stx___x2b__1___closed__8; -x_58 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_58, 0, x_48); -lean_ctor_set(x_58, 1, x_57); -lean_ctor_set(x_58, 2, x_56); -x_59 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__28; -lean_inc(x_22); -x_60 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_60, 0, x_22); -lean_ctor_set(x_60, 1, x_59); -x_61 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_62 = lean_array_push(x_61, x_58); -x_63 = lean_array_push(x_62, x_60); +lean_ctor_set(x_54, 2, x_52); +x_55 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__28; +lean_inc(x_18); +x_56 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_56, 0, x_18); +lean_ctor_set(x_56, 1, x_55); +x_57 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_58 = lean_array_push(x_57, x_54); +x_59 = lean_array_push(x_58, x_56); +lean_inc(x_59); +x_60 = lean_array_push(x_59, x_13); +x_61 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__13; +x_62 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_62, 0, x_44); +lean_ctor_set(x_62, 1, x_61); +lean_ctor_set(x_62, 2, x_60); +x_63 = lean_array_push(x_41, x_50); lean_inc(x_63); -x_64 = lean_array_push(x_63, x_17); -x_65 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__13; +x_64 = lean_array_push(x_63, x_62); +x_65 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__11; x_66 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_66, 0, x_48); +lean_ctor_set(x_66, 0, x_44); lean_ctor_set(x_66, 1, x_65); lean_ctor_set(x_66, 2, x_64); -x_67 = lean_array_push(x_45, x_54); -lean_inc(x_67); -x_68 = lean_array_push(x_67, x_66); -x_69 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__11; +x_67 = lean_array_push(x_41, x_66); +x_68 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__12; +x_69 = lean_array_push(x_67, x_68); x_70 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_70, 0, x_48); -lean_ctor_set(x_70, 1, x_69); -lean_ctor_set(x_70, 2, x_68); -x_71 = lean_array_push(x_45, x_70); -x_72 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__12; -x_73 = lean_array_push(x_71, x_72); -x_74 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_74, 0, x_48); -lean_ctor_set(x_74, 1, x_57); -lean_ctor_set(x_74, 2, x_73); -x_75 = l_prec_x28___x29___closed__7; -x_76 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_76, 0, x_22); -lean_ctor_set(x_76, 1, x_75); -x_77 = lean_array_push(x_61, x_52); -lean_inc(x_77); -x_78 = lean_array_push(x_77, x_74); -lean_inc(x_76); -x_79 = lean_array_push(x_78, x_76); -x_80 = l___aux__Init__Notation______macroRules__term___u2209____1___closed__2; +lean_ctor_set(x_70, 0, x_44); +lean_ctor_set(x_70, 1, x_53); +lean_ctor_set(x_70, 2, x_69); +x_71 = l_prec_x28___x29___closed__7; +x_72 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_72, 0, x_18); +lean_ctor_set(x_72, 1, x_71); +x_73 = lean_array_push(x_57, x_48); +lean_inc(x_73); +x_74 = lean_array_push(x_73, x_70); +lean_inc(x_72); +x_75 = lean_array_push(x_74, x_72); +x_76 = l___aux__Init__Notation______macroRules__term___u2209____1___closed__2; +x_77 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_77, 0, x_44); +lean_ctor_set(x_77, 1, x_76); +lean_ctor_set(x_77, 2, x_75); +x_78 = lean_array_push(x_59, x_15); +x_79 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_79, 0, x_44); +lean_ctor_set(x_79, 1, x_61); +lean_ctor_set(x_79, 2, x_78); +x_80 = lean_array_push(x_63, x_79); x_81 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_81, 0, x_48); -lean_ctor_set(x_81, 1, x_80); -lean_ctor_set(x_81, 2, x_79); -x_82 = lean_array_push(x_63, x_19); -x_83 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_83, 0, x_48); -lean_ctor_set(x_83, 1, x_65); -lean_ctor_set(x_83, 2, x_82); -x_84 = lean_array_push(x_67, x_83); -x_85 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_85, 0, x_48); -lean_ctor_set(x_85, 1, x_69); -lean_ctor_set(x_85, 2, x_84); -x_86 = lean_array_push(x_45, x_85); -x_87 = lean_array_push(x_86, x_72); -x_88 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_88, 0, x_48); -lean_ctor_set(x_88, 1, x_57); -lean_ctor_set(x_88, 2, x_87); -x_89 = lean_array_push(x_77, x_88); -x_90 = lean_array_push(x_89, x_76); +lean_ctor_set(x_81, 0, x_44); +lean_ctor_set(x_81, 1, x_65); +lean_ctor_set(x_81, 2, x_80); +x_82 = lean_array_push(x_41, x_81); +x_83 = lean_array_push(x_82, x_68); +x_84 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_84, 0, x_44); +lean_ctor_set(x_84, 1, x_53); +lean_ctor_set(x_84, 2, x_83); +x_85 = lean_array_push(x_73, x_84); +x_86 = lean_array_push(x_85, x_72); +x_87 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_87, 0, x_44); +lean_ctor_set(x_87, 1, x_76); +lean_ctor_set(x_87, 2, x_86); +x_88 = lean_array_push(x_57, x_46); +x_89 = lean_array_push(x_88, x_77); +x_90 = lean_array_push(x_89, x_87); x_91 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_91, 0, x_48); -lean_ctor_set(x_91, 1, x_80); +lean_ctor_set(x_91, 0, x_44); +lean_ctor_set(x_91, 1, x_53); lean_ctor_set(x_91, 2, x_90); -x_92 = lean_array_push(x_61, x_50); -x_93 = lean_array_push(x_92, x_81); -x_94 = lean_array_push(x_93, x_91); +x_92 = lean_array_push(x_41, x_40); +x_93 = lean_array_push(x_92, x_91); +x_94 = l___aux__Init__Notation______macroRules__term___u2218____1___closed__4; x_95 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_95, 0, x_48); -lean_ctor_set(x_95, 1, x_57); -lean_ctor_set(x_95, 2, x_94); -x_96 = lean_array_push(x_45, x_44); -x_97 = lean_array_push(x_96, x_95); -x_98 = l___aux__Init__Notation______macroRules__term___u2218____1___closed__4; -x_99 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_99, 0, x_48); -lean_ctor_set(x_99, 1, x_98); -lean_ctor_set(x_99, 2, x_97); -x_100 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__21; -x_101 = lean_array_push(x_100, x_39); -lean_inc(x_28); -x_102 = lean_array_push(x_101, x_28); +lean_ctor_set(x_95, 0, x_44); +lean_ctor_set(x_95, 1, x_94); +lean_ctor_set(x_95, 2, x_93); +x_96 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__21; +x_97 = lean_array_push(x_96, x_35); +lean_inc(x_24); +x_98 = lean_array_push(x_97, x_24); +lean_inc(x_29); +x_99 = lean_array_push(x_98, x_29); lean_inc(x_33); -x_103 = lean_array_push(x_102, x_33); -lean_inc(x_37); -x_104 = lean_array_push(x_103, x_37); -x_105 = lean_array_push(x_104, x_99); -x_106 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__11; -x_107 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_107, 0, x_48); -lean_ctor_set(x_107, 1, x_106); -lean_ctor_set(x_107, 2, x_105); -x_108 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__13; -x_109 = lean_array_push(x_108, x_26); -x_110 = lean_array_push(x_109, x_28); -x_111 = lean_array_push(x_110, x_33); -x_112 = lean_array_push(x_111, x_35); -x_113 = lean_array_push(x_112, x_15); -x_114 = lean_array_push(x_113, x_37); -x_115 = lean_array_push(x_114, x_107); -x_116 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__2; -x_117 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_117, 0, x_48); -lean_ctor_set(x_117, 1, x_116); -lean_ctor_set(x_117, 2, x_115); -lean_ctor_set(x_20, 0, x_117); -return x_20; -} -else -{ -lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; -x_118 = lean_ctor_get(x_20, 0); -x_119 = lean_ctor_get(x_20, 1); -lean_inc(x_119); -lean_inc(x_118); -lean_dec(x_20); -x_120 = lean_ctor_get(x_2, 2); -lean_inc(x_120); -x_121 = lean_ctor_get(x_2, 1); -lean_inc(x_121); -lean_dec(x_2); -x_122 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__3; -lean_inc(x_118); -x_123 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_123, 0, x_118); -lean_ctor_set(x_123, 1, x_122); -x_124 = l_stx___x3f___closed__3; -lean_inc(x_118); -x_125 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_125, 0, x_118); -lean_ctor_set(x_125, 1, x_124); -x_126 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__7; -lean_inc(x_120); -lean_inc(x_121); -x_127 = l_Lean_addMacroScope(x_121, x_126, x_120); -x_128 = lean_box(0); -x_129 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__6; -lean_inc(x_118); -x_130 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_130, 0, x_118); +x_100 = lean_array_push(x_99, x_33); +x_101 = lean_array_push(x_100, x_95); +x_102 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__11; +x_103 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_103, 0, x_44); +lean_ctor_set(x_103, 1, x_102); +lean_ctor_set(x_103, 2, x_101); +x_104 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__13; +x_105 = lean_array_push(x_104, x_22); +x_106 = lean_array_push(x_105, x_24); +x_107 = lean_array_push(x_106, x_29); +x_108 = lean_array_push(x_107, x_31); +x_109 = lean_array_push(x_108, x_11); +x_110 = lean_array_push(x_109, x_33); +x_111 = lean_array_push(x_110, x_103); +x_112 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__2; +x_113 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_113, 0, x_44); +lean_ctor_set(x_113, 1, x_112); +lean_ctor_set(x_113, 2, x_111); +lean_ctor_set(x_16, 0, x_113); +return x_16; +} +else +{ +lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; +x_114 = lean_ctor_get(x_16, 0); +x_115 = lean_ctor_get(x_16, 1); +lean_inc(x_115); +lean_inc(x_114); +lean_dec(x_16); +x_116 = lean_ctor_get(x_2, 2); +lean_inc(x_116); +x_117 = lean_ctor_get(x_2, 1); +lean_inc(x_117); +lean_dec(x_2); +x_118 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__3; +lean_inc(x_114); +x_119 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_119, 0, x_114); +lean_ctor_set(x_119, 1, x_118); +x_120 = l_stx___x3f___closed__3; +lean_inc(x_114); +x_121 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_121, 0, x_114); +lean_ctor_set(x_121, 1, x_120); +x_122 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__7; +lean_inc(x_116); +lean_inc(x_117); +x_123 = l_Lean_addMacroScope(x_117, x_122, x_116); +x_124 = lean_box(0); +x_125 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__6; +lean_inc(x_114); +x_126 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_126, 0, x_114); +lean_ctor_set(x_126, 1, x_125); +lean_ctor_set(x_126, 2, x_123); +lean_ctor_set(x_126, 3, x_124); +x_127 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__8; +lean_inc(x_114); +x_128 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_128, 0, x_114); +lean_ctor_set(x_128, 1, x_127); +x_129 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__9; +lean_inc(x_114); +x_130 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_130, 0, x_114); lean_ctor_set(x_130, 1, x_129); -lean_ctor_set(x_130, 2, x_127); -lean_ctor_set(x_130, 3, x_128); -x_131 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__8; -lean_inc(x_118); +x_131 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__12; +lean_inc(x_114); x_132 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_132, 0, x_118); +lean_ctor_set(x_132, 0, x_114); lean_ctor_set(x_132, 1, x_131); -x_133 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__9; -lean_inc(x_118); -x_134 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_134, 0, x_118); -lean_ctor_set(x_134, 1, x_133); -x_135 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__12; -lean_inc(x_118); -x_136 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_136, 0, x_118); -lean_ctor_set(x_136, 1, x_135); -x_137 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__16; -x_138 = l_Lean_addMacroScope(x_121, x_137, x_120); -x_139 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__15; -x_140 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__18; -lean_inc(x_118); -x_141 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_141, 0, x_118); -lean_ctor_set(x_141, 1, x_139); -lean_ctor_set(x_141, 2, x_138); -lean_ctor_set(x_141, 3, x_140); -x_142 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; -lean_inc(x_125); -x_143 = lean_array_push(x_142, x_125); -lean_inc(x_130); -x_144 = lean_array_push(x_143, x_130); -x_145 = lean_box(2); -x_146 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__20; -x_147 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_147, 0, x_145); +x_133 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__16; +x_134 = l_Lean_addMacroScope(x_117, x_133, x_116); +x_135 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__15; +x_136 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__18; +lean_inc(x_114); +x_137 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_137, 0, x_114); +lean_ctor_set(x_137, 1, x_135); +lean_ctor_set(x_137, 2, x_134); +lean_ctor_set(x_137, 3, x_136); +x_138 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; +lean_inc(x_121); +x_139 = lean_array_push(x_138, x_121); +lean_inc(x_126); +x_140 = lean_array_push(x_139, x_126); +x_141 = lean_box(2); +x_142 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__20; +x_143 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_143, 0, x_141); +lean_ctor_set(x_143, 1, x_142); +lean_ctor_set(x_143, 2, x_140); +x_144 = l_prec_x28___x29___closed__3; +lean_inc(x_114); +x_145 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_145, 0, x_114); +lean_ctor_set(x_145, 1, x_144); +x_146 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__10; +lean_inc(x_114); +x_147 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_147, 0, x_114); lean_ctor_set(x_147, 1, x_146); -lean_ctor_set(x_147, 2, x_144); -x_148 = l_prec_x28___x29___closed__3; -lean_inc(x_118); -x_149 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_149, 0, x_118); -lean_ctor_set(x_149, 1, x_148); -x_150 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__10; -lean_inc(x_118); -x_151 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_151, 0, x_118); +x_148 = l___aux__Init__Notation______macroRules__precMax__1___closed__4; +x_149 = lean_array_push(x_148, x_9); +x_150 = l___aux__Init__Notation______macroRules__stx___x2b__1___closed__8; +x_151 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_151, 0, x_141); lean_ctor_set(x_151, 1, x_150); -x_152 = l___aux__Init__Notation______macroRules__precMax__1___closed__4; -x_153 = lean_array_push(x_152, x_9); -x_154 = l___aux__Init__Notation______macroRules__stx___x2b__1___closed__8; -x_155 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_155, 0, x_145); -lean_ctor_set(x_155, 1, x_154); -lean_ctor_set(x_155, 2, x_153); -x_156 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__28; -lean_inc(x_118); -x_157 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_157, 0, x_118); -lean_ctor_set(x_157, 1, x_156); -x_158 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_159 = lean_array_push(x_158, x_155); -x_160 = lean_array_push(x_159, x_157); +lean_ctor_set(x_151, 2, x_149); +x_152 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__28; +lean_inc(x_114); +x_153 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_153, 0, x_114); +lean_ctor_set(x_153, 1, x_152); +x_154 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_155 = lean_array_push(x_154, x_151); +x_156 = lean_array_push(x_155, x_153); +lean_inc(x_156); +x_157 = lean_array_push(x_156, x_13); +x_158 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__13; +x_159 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_159, 0, x_141); +lean_ctor_set(x_159, 1, x_158); +lean_ctor_set(x_159, 2, x_157); +x_160 = lean_array_push(x_138, x_147); lean_inc(x_160); -x_161 = lean_array_push(x_160, x_17); -x_162 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__13; +x_161 = lean_array_push(x_160, x_159); +x_162 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__11; x_163 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_163, 0, x_145); +lean_ctor_set(x_163, 0, x_141); lean_ctor_set(x_163, 1, x_162); lean_ctor_set(x_163, 2, x_161); -x_164 = lean_array_push(x_142, x_151); -lean_inc(x_164); -x_165 = lean_array_push(x_164, x_163); -x_166 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__11; +x_164 = lean_array_push(x_138, x_163); +x_165 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__12; +x_166 = lean_array_push(x_164, x_165); x_167 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_167, 0, x_145); -lean_ctor_set(x_167, 1, x_166); -lean_ctor_set(x_167, 2, x_165); -x_168 = lean_array_push(x_142, x_167); -x_169 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__12; -x_170 = lean_array_push(x_168, x_169); -x_171 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_171, 0, x_145); -lean_ctor_set(x_171, 1, x_154); -lean_ctor_set(x_171, 2, x_170); -x_172 = l_prec_x28___x29___closed__7; -x_173 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_173, 0, x_118); -lean_ctor_set(x_173, 1, x_172); -x_174 = lean_array_push(x_158, x_149); -lean_inc(x_174); -x_175 = lean_array_push(x_174, x_171); -lean_inc(x_173); -x_176 = lean_array_push(x_175, x_173); -x_177 = l___aux__Init__Notation______macroRules__term___u2209____1___closed__2; +lean_ctor_set(x_167, 0, x_141); +lean_ctor_set(x_167, 1, x_150); +lean_ctor_set(x_167, 2, x_166); +x_168 = l_prec_x28___x29___closed__7; +x_169 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_169, 0, x_114); +lean_ctor_set(x_169, 1, x_168); +x_170 = lean_array_push(x_154, x_145); +lean_inc(x_170); +x_171 = lean_array_push(x_170, x_167); +lean_inc(x_169); +x_172 = lean_array_push(x_171, x_169); +x_173 = l___aux__Init__Notation______macroRules__term___u2209____1___closed__2; +x_174 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_174, 0, x_141); +lean_ctor_set(x_174, 1, x_173); +lean_ctor_set(x_174, 2, x_172); +x_175 = lean_array_push(x_156, x_15); +x_176 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_176, 0, x_141); +lean_ctor_set(x_176, 1, x_158); +lean_ctor_set(x_176, 2, x_175); +x_177 = lean_array_push(x_160, x_176); x_178 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_178, 0, x_145); -lean_ctor_set(x_178, 1, x_177); -lean_ctor_set(x_178, 2, x_176); -x_179 = lean_array_push(x_160, x_19); -x_180 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_180, 0, x_145); -lean_ctor_set(x_180, 1, x_162); -lean_ctor_set(x_180, 2, x_179); -x_181 = lean_array_push(x_164, x_180); -x_182 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_182, 0, x_145); -lean_ctor_set(x_182, 1, x_166); -lean_ctor_set(x_182, 2, x_181); -x_183 = lean_array_push(x_142, x_182); -x_184 = lean_array_push(x_183, x_169); -x_185 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_185, 0, x_145); -lean_ctor_set(x_185, 1, x_154); -lean_ctor_set(x_185, 2, x_184); -x_186 = lean_array_push(x_174, x_185); -x_187 = lean_array_push(x_186, x_173); +lean_ctor_set(x_178, 0, x_141); +lean_ctor_set(x_178, 1, x_162); +lean_ctor_set(x_178, 2, x_177); +x_179 = lean_array_push(x_138, x_178); +x_180 = lean_array_push(x_179, x_165); +x_181 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_181, 0, x_141); +lean_ctor_set(x_181, 1, x_150); +lean_ctor_set(x_181, 2, x_180); +x_182 = lean_array_push(x_170, x_181); +x_183 = lean_array_push(x_182, x_169); +x_184 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_184, 0, x_141); +lean_ctor_set(x_184, 1, x_173); +lean_ctor_set(x_184, 2, x_183); +x_185 = lean_array_push(x_154, x_143); +x_186 = lean_array_push(x_185, x_174); +x_187 = lean_array_push(x_186, x_184); x_188 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_188, 0, x_145); -lean_ctor_set(x_188, 1, x_177); +lean_ctor_set(x_188, 0, x_141); +lean_ctor_set(x_188, 1, x_150); lean_ctor_set(x_188, 2, x_187); -x_189 = lean_array_push(x_158, x_147); -x_190 = lean_array_push(x_189, x_178); -x_191 = lean_array_push(x_190, x_188); +x_189 = lean_array_push(x_138, x_137); +x_190 = lean_array_push(x_189, x_188); +x_191 = l___aux__Init__Notation______macroRules__term___u2218____1___closed__4; x_192 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_192, 0, x_145); -lean_ctor_set(x_192, 1, x_154); -lean_ctor_set(x_192, 2, x_191); -x_193 = lean_array_push(x_142, x_141); -x_194 = lean_array_push(x_193, x_192); -x_195 = l___aux__Init__Notation______macroRules__term___u2218____1___closed__4; -x_196 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_196, 0, x_145); -lean_ctor_set(x_196, 1, x_195); -lean_ctor_set(x_196, 2, x_194); -x_197 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__21; -x_198 = lean_array_push(x_197, x_136); -lean_inc(x_125); -x_199 = lean_array_push(x_198, x_125); +lean_ctor_set(x_192, 0, x_141); +lean_ctor_set(x_192, 1, x_191); +lean_ctor_set(x_192, 2, x_190); +x_193 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__21; +x_194 = lean_array_push(x_193, x_132); +lean_inc(x_121); +x_195 = lean_array_push(x_194, x_121); +lean_inc(x_126); +x_196 = lean_array_push(x_195, x_126); lean_inc(x_130); -x_200 = lean_array_push(x_199, x_130); -lean_inc(x_134); -x_201 = lean_array_push(x_200, x_134); -x_202 = lean_array_push(x_201, x_196); -x_203 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__11; -x_204 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_204, 0, x_145); -lean_ctor_set(x_204, 1, x_203); -lean_ctor_set(x_204, 2, x_202); -x_205 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__13; -x_206 = lean_array_push(x_205, x_123); -x_207 = lean_array_push(x_206, x_125); -x_208 = lean_array_push(x_207, x_130); -x_209 = lean_array_push(x_208, x_132); -x_210 = lean_array_push(x_209, x_15); -x_211 = lean_array_push(x_210, x_134); -x_212 = lean_array_push(x_211, x_204); -x_213 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__2; -x_214 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_214, 0, x_145); -lean_ctor_set(x_214, 1, x_213); -lean_ctor_set(x_214, 2, x_212); -x_215 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_215, 0, x_214); -lean_ctor_set(x_215, 1, x_119); -return x_215; -} +x_197 = lean_array_push(x_196, x_130); +x_198 = lean_array_push(x_197, x_192); +x_199 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__11; +x_200 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_200, 0, x_141); +lean_ctor_set(x_200, 1, x_199); +lean_ctor_set(x_200, 2, x_198); +x_201 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__13; +x_202 = lean_array_push(x_201, x_119); +x_203 = lean_array_push(x_202, x_121); +x_204 = lean_array_push(x_203, x_126); +x_205 = lean_array_push(x_204, x_128); +x_206 = lean_array_push(x_205, x_11); +x_207 = lean_array_push(x_206, x_130); +x_208 = lean_array_push(x_207, x_200); +x_209 = l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__2; +x_210 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_210, 0, x_141); +lean_ctor_set(x_210, 1, x_209); +lean_ctor_set(x_210, 2, x_208); +x_211 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_211, 0, x_210); +lean_ctor_set(x_211, 1, x_115); +return x_211; } } } @@ -28238,7 +28195,7 @@ return x_36; } else { -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; lean_object* x_44; uint8_t x_45; +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; uint8_t x_44; x_37 = l_Lean_Syntax_getArg(x_9, x_8); x_38 = lean_unsigned_to_nat(1u); x_39 = l_Lean_Syntax_getArg(x_9, x_38); @@ -28248,17 +28205,17 @@ x_41 = l_Lean_Syntax_getArg(x_1, x_40); lean_dec(x_1); x_42 = l_Lean_Syntax_getArgs(x_39); lean_dec(x_39); -x_43 = lean_array_push(x_42, x_41); -x_44 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_45 = !lean_is_exclusive(x_44); -if (x_45 == 0) -{ -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_46 = lean_ctor_get(x_44, 0); -lean_dec(x_46); -x_47 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__11; -x_48 = l_Array_appendCore___rarg(x_47, x_43); -lean_dec(x_43); +x_43 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_44 = !lean_is_exclusive(x_43); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_45 = lean_ctor_get(x_43, 0); +lean_dec(x_45); +x_46 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__11; +x_47 = l_Array_appendCore___rarg(x_46, x_42); +lean_dec(x_42); +x_48 = lean_array_push(x_47, x_41); x_49 = lean_box(2); x_50 = l___aux__Init__Notation______macroRules__stx___x2b__1___closed__8; x_51 = lean_alloc_ctor(1, 3, 0); @@ -28272,35 +28229,36 @@ x_55 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_55, 0, x_49); lean_ctor_set(x_55, 1, x_10); lean_ctor_set(x_55, 2, x_54); -lean_ctor_set(x_44, 0, x_55); -return x_44; +lean_ctor_set(x_43, 0, x_55); +return x_43; } else { -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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_56 = lean_ctor_get(x_44, 1); +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; lean_object* x_67; +x_56 = lean_ctor_get(x_43, 1); lean_inc(x_56); -lean_dec(x_44); -x_57 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__11; -x_58 = l_Array_appendCore___rarg(x_57, x_43); lean_dec(x_43); -x_59 = lean_box(2); -x_60 = l___aux__Init__Notation______macroRules__stx___x2b__1___closed__8; -x_61 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_61, 0, x_59); -lean_ctor_set(x_61, 1, x_60); -lean_ctor_set(x_61, 2, x_58); -x_62 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; -x_63 = lean_array_push(x_62, x_37); -x_64 = lean_array_push(x_63, x_61); -x_65 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_65, 0, x_59); -lean_ctor_set(x_65, 1, x_10); -lean_ctor_set(x_65, 2, x_64); -x_66 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_66, 0, x_65); -lean_ctor_set(x_66, 1, x_56); -return x_66; +x_57 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__11; +x_58 = l_Array_appendCore___rarg(x_57, x_42); +lean_dec(x_42); +x_59 = lean_array_push(x_58, x_41); +x_60 = lean_box(2); +x_61 = l___aux__Init__Notation______macroRules__stx___x2b__1___closed__8; +x_62 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_62, 0, x_60); +lean_ctor_set(x_62, 1, x_61); +lean_ctor_set(x_62, 2, x_59); +x_63 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; +x_64 = lean_array_push(x_63, x_37); +x_65 = lean_array_push(x_64, x_62); +x_66 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_66, 0, x_60); +lean_ctor_set(x_66, 1, x_10); +lean_ctor_set(x_66, 2, x_65); +x_67 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_67, 1, x_56); +return x_67; } } } @@ -28478,24 +28436,24 @@ return x_36; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; x_37 = l_Lean_Syntax_getArg(x_11, x_8); x_38 = lean_unsigned_to_nat(1u); x_39 = l_Lean_Syntax_getArg(x_11, x_38); lean_dec(x_11); x_40 = l_Lean_Syntax_getArgs(x_39); lean_dec(x_39); -x_41 = lean_array_push(x_40, x_9); -x_42 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_43 = !lean_is_exclusive(x_42); -if (x_43 == 0) -{ -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; -x_44 = lean_ctor_get(x_42, 0); -lean_dec(x_44); -x_45 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__11; -x_46 = l_Array_appendCore___rarg(x_45, x_41); -lean_dec(x_41); +x_41 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_42 = !lean_is_exclusive(x_41); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_43 = lean_ctor_get(x_41, 0); +lean_dec(x_43); +x_44 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__11; +x_45 = l_Array_appendCore___rarg(x_44, x_40); +lean_dec(x_40); +x_46 = lean_array_push(x_45, x_9); x_47 = lean_box(2); x_48 = l___aux__Init__Notation______macroRules__stx___x2b__1___closed__8; x_49 = lean_alloc_ctor(1, 3, 0); @@ -28509,35 +28467,36 @@ x_53 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_53, 0, x_47); lean_ctor_set(x_53, 1, x_12); lean_ctor_set(x_53, 2, x_52); -lean_ctor_set(x_42, 0, x_53); -return x_42; +lean_ctor_set(x_41, 0, x_53); +return x_41; } else { -lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_54 = lean_ctor_get(x_42, 1); +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_54 = lean_ctor_get(x_41, 1); lean_inc(x_54); -lean_dec(x_42); -x_55 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__11; -x_56 = l_Array_appendCore___rarg(x_55, x_41); lean_dec(x_41); -x_57 = lean_box(2); -x_58 = l___aux__Init__Notation______macroRules__stx___x2b__1___closed__8; -x_59 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_59, 0, x_57); -lean_ctor_set(x_59, 1, x_58); -lean_ctor_set(x_59, 2, x_56); -x_60 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; -x_61 = lean_array_push(x_60, x_37); -x_62 = lean_array_push(x_61, x_59); -x_63 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_63, 0, x_57); -lean_ctor_set(x_63, 1, x_12); -lean_ctor_set(x_63, 2, x_62); -x_64 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_64, 0, x_63); -lean_ctor_set(x_64, 1, x_54); -return x_64; +x_55 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__11; +x_56 = l_Array_appendCore___rarg(x_55, x_40); +lean_dec(x_40); +x_57 = lean_array_push(x_56, x_9); +x_58 = lean_box(2); +x_59 = l___aux__Init__Notation______macroRules__stx___x2b__1___closed__8; +x_60 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +lean_ctor_set(x_60, 2, x_57); +x_61 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; +x_62 = lean_array_push(x_61, x_37); +x_63 = lean_array_push(x_62, x_60); +x_64 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_64, 0, x_58); +lean_ctor_set(x_64, 1, x_12); +lean_ctor_set(x_64, 2, x_63); +x_65 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_65, 0, x_64); +lean_ctor_set(x_65, 1, x_54); +return x_65; } } } @@ -28775,7 +28734,7 @@ return x_36; } else { -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; lean_object* x_44; uint8_t x_45; +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; uint8_t x_44; x_37 = l_Lean_Syntax_getArg(x_9, x_8); x_38 = lean_unsigned_to_nat(1u); x_39 = l_Lean_Syntax_getArg(x_9, x_38); @@ -28785,17 +28744,17 @@ x_41 = l_Lean_Syntax_getArg(x_1, x_40); lean_dec(x_1); x_42 = l_Lean_Syntax_getArgs(x_39); lean_dec(x_39); -x_43 = lean_array_push(x_42, x_41); -x_44 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_45 = !lean_is_exclusive(x_44); -if (x_45 == 0) -{ -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_46 = lean_ctor_get(x_44, 0); -lean_dec(x_46); -x_47 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__11; -x_48 = l_Array_appendCore___rarg(x_47, x_43); -lean_dec(x_43); +x_43 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_44 = !lean_is_exclusive(x_43); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_45 = lean_ctor_get(x_43, 0); +lean_dec(x_45); +x_46 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__11; +x_47 = l_Array_appendCore___rarg(x_46, x_42); +lean_dec(x_42); +x_48 = lean_array_push(x_47, x_41); x_49 = lean_box(2); x_50 = l___aux__Init__Notation______macroRules__stx___x2b__1___closed__8; x_51 = lean_alloc_ctor(1, 3, 0); @@ -28809,59 +28768,60 @@ x_55 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_55, 0, x_49); lean_ctor_set(x_55, 1, x_10); lean_ctor_set(x_55, 2, x_54); -lean_ctor_set(x_44, 0, x_55); -return x_44; +lean_ctor_set(x_43, 0, x_55); +return x_43; } else { -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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_56 = lean_ctor_get(x_44, 1); +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; lean_object* x_67; +x_56 = lean_ctor_get(x_43, 1); lean_inc(x_56); -lean_dec(x_44); -x_57 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__11; -x_58 = l_Array_appendCore___rarg(x_57, x_43); lean_dec(x_43); -x_59 = lean_box(2); -x_60 = l___aux__Init__Notation______macroRules__stx___x2b__1___closed__8; -x_61 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_61, 0, x_59); -lean_ctor_set(x_61, 1, x_60); -lean_ctor_set(x_61, 2, x_58); -x_62 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; -x_63 = lean_array_push(x_62, x_37); -x_64 = lean_array_push(x_63, x_61); -x_65 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_65, 0, x_59); -lean_ctor_set(x_65, 1, x_10); -lean_ctor_set(x_65, 2, x_64); -x_66 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_66, 0, x_65); -lean_ctor_set(x_66, 1, x_56); -return x_66; +x_57 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__11; +x_58 = l_Array_appendCore___rarg(x_57, x_42); +lean_dec(x_42); +x_59 = lean_array_push(x_58, x_41); +x_60 = lean_box(2); +x_61 = l___aux__Init__Notation______macroRules__stx___x2b__1___closed__8; +x_62 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_62, 0, x_60); +lean_ctor_set(x_62, 1, x_61); +lean_ctor_set(x_62, 2, x_59); +x_63 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; +x_64 = lean_array_push(x_63, x_37); +x_65 = lean_array_push(x_64, x_62); +x_66 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_66, 0, x_60); +lean_ctor_set(x_66, 1, x_10); +lean_ctor_set(x_66, 2, x_65); +x_67 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_67, 1, x_56); +return x_67; } } } } } -static lean_object* _init_l_term_x7b_____x3a___x2f_x2f___x7d___closed__1() { +static lean_object* _init_l_term_x7b___x3a___x2f_x2f___x7d___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("term{__:_//_}", 13); +x_1 = lean_mk_string_from_bytes("term{_:_//_}", 12); return x_1; } } -static lean_object* _init_l_term_x7b_____x3a___x2f_x2f___x7d___closed__2() { +static lean_object* _init_l_term_x7b___x3a___x2f_x2f___x7d___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_term_x7b_____x3a___x2f_x2f___x7d___closed__1; +x_2 = l_term_x7b___x3a___x2f_x2f___x7d___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_term_x7b_____x3a___x2f_x2f___x7d___closed__3() { +static lean_object* _init_l_term_x7b___x3a___x2f_x2f___x7d___closed__3() { _start: { lean_object* x_1; @@ -28869,22 +28829,22 @@ x_1 = lean_mk_string_from_bytes("{ ", 2); return x_1; } } -static lean_object* _init_l_term_x7b_____x3a___x2f_x2f___x7d___closed__4() { +static lean_object* _init_l_term_x7b___x3a___x2f_x2f___x7d___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_term_x7b_____x3a___x2f_x2f___x7d___closed__3; +x_1 = l_term_x7b___x3a___x2f_x2f___x7d___closed__3; x_2 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_term_x7b_____x3a___x2f_x2f___x7d___closed__5() { +static lean_object* _init_l_term_x7b___x3a___x2f_x2f___x7d___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Syntax_addPrec___closed__10; -x_2 = l_term_x7b_____x3a___x2f_x2f___x7d___closed__4; +x_2 = l_term_x7b___x3a___x2f_x2f___x7d___closed__4; x_3 = l_termDepIfThenElse___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -28893,7 +28853,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_term_x7b_____x3a___x2f_x2f___x7d___closed__6() { +static lean_object* _init_l_term_x7b___x3a___x2f_x2f___x7d___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; @@ -28907,25 +28867,25 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_term_x7b_____x3a___x2f_x2f___x7d___closed__7() { +static lean_object* _init_l_term_x7b___x3a___x2f_x2f___x7d___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___aux__Init__Notation______macroRules__stx___x3f__1___closed__4; -x_2 = l_term_x7b_____x3a___x2f_x2f___x7d___closed__6; +x_2 = l_term_x7b___x3a___x2f_x2f___x7d___closed__6; x_3 = lean_alloc_ctor(1, 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_term_x7b_____x3a___x2f_x2f___x7d___closed__8() { +static lean_object* _init_l_term_x7b___x3a___x2f_x2f___x7d___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Syntax_addPrec___closed__10; -x_2 = l_term_x7b_____x3a___x2f_x2f___x7d___closed__5; -x_3 = l_term_x7b_____x3a___x2f_x2f___x7d___closed__7; +x_2 = l_term_x7b___x3a___x2f_x2f___x7d___closed__5; +x_3 = l_term_x7b___x3a___x2f_x2f___x7d___closed__7; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -28933,7 +28893,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_term_x7b_____x3a___x2f_x2f___x7d___closed__9() { +static lean_object* _init_l_term_x7b___x3a___x2f_x2f___x7d___closed__9() { _start: { lean_object* x_1; @@ -28941,23 +28901,23 @@ x_1 = lean_mk_string_from_bytes(" // ", 4); return x_1; } } -static lean_object* _init_l_term_x7b_____x3a___x2f_x2f___x7d___closed__10() { +static lean_object* _init_l_term_x7b___x3a___x2f_x2f___x7d___closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_term_x7b_____x3a___x2f_x2f___x7d___closed__9; +x_1 = l_term_x7b___x3a___x2f_x2f___x7d___closed__9; x_2 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_term_x7b_____x3a___x2f_x2f___x7d___closed__11() { +static lean_object* _init_l_term_x7b___x3a___x2f_x2f___x7d___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Syntax_addPrec___closed__10; -x_2 = l_term_x7b_____x3a___x2f_x2f___x7d___closed__8; -x_3 = l_term_x7b_____x3a___x2f_x2f___x7d___closed__10; +x_2 = l_term_x7b___x3a___x2f_x2f___x7d___closed__8; +x_3 = l_term_x7b___x3a___x2f_x2f___x7d___closed__10; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -28965,12 +28925,12 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_term_x7b_____x3a___x2f_x2f___x7d___closed__12() { +static lean_object* _init_l_term_x7b___x3a___x2f_x2f___x7d___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Syntax_addPrec___closed__10; -x_2 = l_term_x7b_____x3a___x2f_x2f___x7d___closed__11; +x_2 = l_term_x7b___x3a___x2f_x2f___x7d___closed__11; x_3 = l_termDepIfThenElse___closed__16; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -28979,7 +28939,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_term_x7b_____x3a___x2f_x2f___x7d___closed__13() { +static lean_object* _init_l_term_x7b___x3a___x2f_x2f___x7d___closed__13() { _start: { lean_object* x_1; @@ -28987,23 +28947,23 @@ x_1 = lean_mk_string_from_bytes(" }", 2); return x_1; } } -static lean_object* _init_l_term_x7b_____x3a___x2f_x2f___x7d___closed__14() { +static lean_object* _init_l_term_x7b___x3a___x2f_x2f___x7d___closed__14() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_term_x7b_____x3a___x2f_x2f___x7d___closed__13; +x_1 = l_term_x7b___x3a___x2f_x2f___x7d___closed__13; x_2 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_term_x7b_____x3a___x2f_x2f___x7d___closed__15() { +static lean_object* _init_l_term_x7b___x3a___x2f_x2f___x7d___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Syntax_addPrec___closed__10; -x_2 = l_term_x7b_____x3a___x2f_x2f___x7d___closed__12; -x_3 = l_term_x7b_____x3a___x2f_x2f___x7d___closed__14; +x_2 = l_term_x7b___x3a___x2f_x2f___x7d___closed__12; +x_3 = l_term_x7b___x3a___x2f_x2f___x7d___closed__14; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -29011,13 +28971,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_term_x7b_____x3a___x2f_x2f___x7d___closed__16() { +static lean_object* _init_l_term_x7b___x3a___x2f_x2f___x7d___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_term_x7b_____x3a___x2f_x2f___x7d___closed__2; +x_1 = l_term_x7b___x3a___x2f_x2f___x7d___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_term_x7b_____x3a___x2f_x2f___x7d___closed__15; +x_3 = l_term_x7b___x3a___x2f_x2f___x7d___closed__15; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -29025,15 +28985,15 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_term_x7b_____x3a___x2f_x2f___x7d() { +static lean_object* _init_l_term_x7b___x3a___x2f_x2f___x7d() { _start: { lean_object* x_1; -x_1 = l_term_x7b_____x3a___x2f_x2f___x7d___closed__16; +x_1 = l_term_x7b___x3a___x2f_x2f___x7d___closed__16; return x_1; } } -static lean_object* _init_l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__1() { +static lean_object* _init_l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__1() { _start: { lean_object* x_1; @@ -29041,22 +29001,22 @@ x_1 = lean_mk_string_from_bytes("Subtype", 7); return x_1; } } -static lean_object* _init_l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__2() { +static lean_object* _init_l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__1; +x_1 = l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__1; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__3() { +static lean_object* _init_l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__1; +x_1 = l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__2; +x_3 = l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__2; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -29064,41 +29024,41 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__4() { +static lean_object* _init_l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__1; +x_2 = l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__5() { +static lean_object* _init_l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__4; +x_2 = l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__4; 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___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__6() { +static lean_object* _init_l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__5; +x_2 = l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__5; x_3 = lean_alloc_ctor(1, 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___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__7() { +static lean_object* _init_l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__7() { _start: { lean_object* x_1; @@ -29106,21 +29066,21 @@ x_1 = lean_mk_string_from_bytes("typeAscription", 14); return x_1; } } -static lean_object* _init_l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__8() { +static lean_object* _init_l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___aux__Init__Notation______macroRules__term___u2218____1___closed__2; -x_2 = l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__7; +x_2 = l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__7; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; uint8_t x_5; -x_4 = l_term_x7b_____x3a___x2f_x2f___x7d___closed__2; +x_4 = l_term_x7b___x3a___x2f_x2f___x7d___closed__2; lean_inc(x_1); x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); if (x_5 == 0) @@ -29136,582 +29096,581 @@ return x_7; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; x_8 = lean_unsigned_to_nat(1u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); x_10 = lean_unsigned_to_nat(2u); x_11 = l_Lean_Syntax_getArg(x_1, x_10); -x_12 = l_Lean_nullKind; lean_inc(x_11); -x_13 = l_Lean_Syntax_isNodeOf(x_11, x_12, x_10); -if (x_13 == 0) +x_12 = l_Lean_Syntax_matchesNull(x_11, x_10); +if (x_12 == 0) { -lean_object* x_14; uint8_t x_15; -x_14 = lean_unsigned_to_nat(0u); -x_15 = l_Lean_Syntax_isNodeOf(x_11, x_12, x_14); -if (x_15 == 0) +lean_object* x_13; uint8_t x_14; +x_13 = lean_unsigned_to_nat(0u); +x_14 = l_Lean_Syntax_matchesNull(x_11, x_13); +if (x_14 == 0) { -lean_object* x_16; lean_object* x_17; +lean_object* x_15; lean_object* x_16; lean_dec(x_9); lean_dec(x_2); lean_dec(x_1); -x_16 = lean_box(1); -x_17 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_3); -return x_17; +x_15 = lean_box(1); +x_16 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_3); +return x_16; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_18 = lean_unsigned_to_nat(4u); -x_19 = l_Lean_Syntax_getArg(x_1, x_18); +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_17 = lean_unsigned_to_nat(4u); +x_18 = l_Lean_Syntax_getArg(x_1, x_17); lean_dec(x_1); lean_inc(x_2); -x_20 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_21 = !lean_is_exclusive(x_20); -if (x_21 == 0) +x_19 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) { -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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; -x_22 = lean_ctor_get(x_20, 0); -x_23 = lean_ctor_get(x_2, 2); +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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_21 = lean_ctor_get(x_19, 0); +x_22 = lean_ctor_get(x_2, 2); +lean_inc(x_22); +x_23 = lean_ctor_get(x_2, 1); lean_inc(x_23); -x_24 = lean_ctor_get(x_2, 1); -lean_inc(x_24); lean_dec(x_2); -x_25 = l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__4; -x_26 = l_Lean_addMacroScope(x_24, x_25, x_23); -x_27 = l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__3; -x_28 = l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__6; -lean_inc(x_22); -x_29 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_29, 0, x_22); -lean_ctor_set(x_29, 1, x_27); -lean_ctor_set(x_29, 2, x_26); -lean_ctor_set(x_29, 3, x_28); -x_30 = l_prec_x28___x29___closed__3; -lean_inc(x_22); -x_31 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_31, 0, x_22); -lean_ctor_set(x_31, 1, x_30); -x_32 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__10; -lean_inc(x_22); -x_33 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_33, 0, x_22); -lean_ctor_set(x_33, 1, x_32); -x_34 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__21; -lean_inc(x_22); -x_35 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_35, 0, x_22); -lean_ctor_set(x_35, 1, x_34); -x_36 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__18; -lean_inc(x_22); -x_37 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_37, 0, x_22); -lean_ctor_set(x_37, 1, x_36); -x_38 = l___aux__Init__Notation______macroRules__precMax__1___closed__4; -x_39 = lean_array_push(x_38, x_37); -x_40 = lean_box(2); -x_41 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__17; -x_42 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_42, 0, x_40); -lean_ctor_set(x_42, 1, x_41); -lean_ctor_set(x_42, 2, x_39); -x_43 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; -x_44 = lean_array_push(x_43, x_35); -x_45 = lean_array_push(x_44, x_42); -x_46 = l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__8; -x_47 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_47, 0, x_40); -lean_ctor_set(x_47, 1, x_46); -lean_ctor_set(x_47, 2, x_45); -x_48 = lean_array_push(x_38, x_47); -x_49 = l___aux__Init__Notation______macroRules__stx___x2b__1___closed__8; -x_50 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_50, 0, x_40); -lean_ctor_set(x_50, 1, x_49); -lean_ctor_set(x_50, 2, x_48); -x_51 = lean_array_push(x_43, x_9); -x_52 = lean_array_push(x_51, x_50); -x_53 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_53, 0, x_40); -lean_ctor_set(x_53, 1, x_49); -lean_ctor_set(x_53, 2, x_52); -x_54 = l_prec_x28___x29___closed__7; -lean_inc(x_22); -x_55 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_55, 0, x_22); -lean_ctor_set(x_55, 1, x_54); -x_56 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_57 = lean_array_push(x_56, x_31); -lean_inc(x_57); -x_58 = lean_array_push(x_57, x_53); -lean_inc(x_55); -x_59 = lean_array_push(x_58, x_55); -x_60 = l___aux__Init__Notation______macroRules__term___u2209____1___closed__2; -x_61 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_61, 0, x_40); -lean_ctor_set(x_61, 1, x_60); -lean_ctor_set(x_61, 2, x_59); -x_62 = lean_array_push(x_38, x_61); -x_63 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_63, 0, x_40); -lean_ctor_set(x_63, 1, x_49); -lean_ctor_set(x_63, 2, x_62); -x_64 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__28; -x_65 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_65, 0, x_22); -lean_ctor_set(x_65, 1, x_64); -x_66 = lean_array_push(x_56, x_63); -x_67 = lean_array_push(x_66, x_65); -x_68 = lean_array_push(x_67, x_19); -x_69 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__13; -x_70 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_70, 0, x_40); -lean_ctor_set(x_70, 1, x_69); -lean_ctor_set(x_70, 2, x_68); -x_71 = lean_array_push(x_43, x_33); -x_72 = lean_array_push(x_71, x_70); -x_73 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__11; -x_74 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_74, 0, x_40); -lean_ctor_set(x_74, 1, x_73); -lean_ctor_set(x_74, 2, x_72); -x_75 = lean_array_push(x_43, x_74); -x_76 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__12; -x_77 = lean_array_push(x_75, x_76); -x_78 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_78, 0, x_40); -lean_ctor_set(x_78, 1, x_49); -lean_ctor_set(x_78, 2, x_77); -x_79 = lean_array_push(x_57, x_78); -x_80 = lean_array_push(x_79, x_55); -x_81 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_81, 0, x_40); -lean_ctor_set(x_81, 1, x_60); -lean_ctor_set(x_81, 2, x_80); -x_82 = lean_array_push(x_38, x_81); -x_83 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_83, 0, x_40); -lean_ctor_set(x_83, 1, x_49); -lean_ctor_set(x_83, 2, x_82); -x_84 = lean_array_push(x_43, x_29); -x_85 = lean_array_push(x_84, x_83); -x_86 = l___aux__Init__Notation______macroRules__term___u2218____1___closed__4; -x_87 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_87, 0, x_40); -lean_ctor_set(x_87, 1, x_86); -lean_ctor_set(x_87, 2, x_85); -lean_ctor_set(x_20, 0, x_87); -return x_20; +x_24 = l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__4; +x_25 = l_Lean_addMacroScope(x_23, x_24, x_22); +x_26 = l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__3; +x_27 = l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__6; +lean_inc(x_21); +x_28 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_28, 0, x_21); +lean_ctor_set(x_28, 1, x_26); +lean_ctor_set(x_28, 2, x_25); +lean_ctor_set(x_28, 3, x_27); +x_29 = l_prec_x28___x29___closed__3; +lean_inc(x_21); +x_30 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_30, 0, x_21); +lean_ctor_set(x_30, 1, x_29); +x_31 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__10; +lean_inc(x_21); +x_32 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_32, 0, x_21); +lean_ctor_set(x_32, 1, x_31); +x_33 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__21; +lean_inc(x_21); +x_34 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_34, 0, x_21); +lean_ctor_set(x_34, 1, x_33); +x_35 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__18; +lean_inc(x_21); +x_36 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_36, 0, x_21); +lean_ctor_set(x_36, 1, x_35); +x_37 = l___aux__Init__Notation______macroRules__precMax__1___closed__4; +x_38 = lean_array_push(x_37, x_36); +x_39 = lean_box(2); +x_40 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__17; +x_41 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +lean_ctor_set(x_41, 2, x_38); +x_42 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; +x_43 = lean_array_push(x_42, x_34); +x_44 = lean_array_push(x_43, x_41); +x_45 = l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__8; +x_46 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_46, 0, x_39); +lean_ctor_set(x_46, 1, x_45); +lean_ctor_set(x_46, 2, x_44); +x_47 = lean_array_push(x_37, x_46); +x_48 = l___aux__Init__Notation______macroRules__stx___x2b__1___closed__8; +x_49 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_49, 0, x_39); +lean_ctor_set(x_49, 1, x_48); +lean_ctor_set(x_49, 2, x_47); +x_50 = lean_array_push(x_42, x_9); +x_51 = lean_array_push(x_50, x_49); +x_52 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_52, 0, x_39); +lean_ctor_set(x_52, 1, x_48); +lean_ctor_set(x_52, 2, x_51); +x_53 = l_prec_x28___x29___closed__7; +lean_inc(x_21); +x_54 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_54, 0, x_21); +lean_ctor_set(x_54, 1, x_53); +x_55 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_56 = lean_array_push(x_55, x_30); +lean_inc(x_56); +x_57 = lean_array_push(x_56, x_52); +lean_inc(x_54); +x_58 = lean_array_push(x_57, x_54); +x_59 = l___aux__Init__Notation______macroRules__term___u2209____1___closed__2; +x_60 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_60, 0, x_39); +lean_ctor_set(x_60, 1, x_59); +lean_ctor_set(x_60, 2, x_58); +x_61 = lean_array_push(x_37, x_60); +x_62 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_62, 0, x_39); +lean_ctor_set(x_62, 1, x_48); +lean_ctor_set(x_62, 2, x_61); +x_63 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__28; +x_64 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_64, 0, x_21); +lean_ctor_set(x_64, 1, x_63); +x_65 = lean_array_push(x_55, x_62); +x_66 = lean_array_push(x_65, x_64); +x_67 = lean_array_push(x_66, x_18); +x_68 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__13; +x_69 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_69, 0, x_39); +lean_ctor_set(x_69, 1, x_68); +lean_ctor_set(x_69, 2, x_67); +x_70 = lean_array_push(x_42, x_32); +x_71 = lean_array_push(x_70, x_69); +x_72 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__11; +x_73 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_73, 0, x_39); +lean_ctor_set(x_73, 1, x_72); +lean_ctor_set(x_73, 2, x_71); +x_74 = lean_array_push(x_42, x_73); +x_75 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__12; +x_76 = lean_array_push(x_74, x_75); +x_77 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_77, 0, x_39); +lean_ctor_set(x_77, 1, x_48); +lean_ctor_set(x_77, 2, x_76); +x_78 = lean_array_push(x_56, x_77); +x_79 = lean_array_push(x_78, x_54); +x_80 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_80, 0, x_39); +lean_ctor_set(x_80, 1, x_59); +lean_ctor_set(x_80, 2, x_79); +x_81 = lean_array_push(x_37, x_80); +x_82 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_82, 0, x_39); +lean_ctor_set(x_82, 1, x_48); +lean_ctor_set(x_82, 2, x_81); +x_83 = lean_array_push(x_42, x_28); +x_84 = lean_array_push(x_83, x_82); +x_85 = l___aux__Init__Notation______macroRules__term___u2218____1___closed__4; +x_86 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_86, 0, x_39); +lean_ctor_set(x_86, 1, x_85); +lean_ctor_set(x_86, 2, x_84); +lean_ctor_set(x_19, 0, x_86); +return x_19; } else { -lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; -x_88 = lean_ctor_get(x_20, 0); -x_89 = lean_ctor_get(x_20, 1); -lean_inc(x_89); +lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; +x_87 = lean_ctor_get(x_19, 0); +x_88 = lean_ctor_get(x_19, 1); lean_inc(x_88); -lean_dec(x_20); -x_90 = lean_ctor_get(x_2, 2); +lean_inc(x_87); +lean_dec(x_19); +x_89 = lean_ctor_get(x_2, 2); +lean_inc(x_89); +x_90 = lean_ctor_get(x_2, 1); lean_inc(x_90); -x_91 = lean_ctor_get(x_2, 1); -lean_inc(x_91); lean_dec(x_2); -x_92 = l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__4; -x_93 = l_Lean_addMacroScope(x_91, x_92, x_90); -x_94 = l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__3; -x_95 = l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__6; -lean_inc(x_88); -x_96 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_96, 0, x_88); -lean_ctor_set(x_96, 1, x_94); -lean_ctor_set(x_96, 2, x_93); -lean_ctor_set(x_96, 3, x_95); -x_97 = l_prec_x28___x29___closed__3; -lean_inc(x_88); -x_98 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_98, 0, x_88); -lean_ctor_set(x_98, 1, x_97); -x_99 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__10; -lean_inc(x_88); -x_100 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_100, 0, x_88); -lean_ctor_set(x_100, 1, x_99); -x_101 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__21; -lean_inc(x_88); -x_102 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_102, 0, x_88); -lean_ctor_set(x_102, 1, x_101); -x_103 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__18; -lean_inc(x_88); -x_104 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_104, 0, x_88); -lean_ctor_set(x_104, 1, x_103); -x_105 = l___aux__Init__Notation______macroRules__precMax__1___closed__4; -x_106 = lean_array_push(x_105, x_104); -x_107 = lean_box(2); -x_108 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__17; -x_109 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_109, 0, x_107); -lean_ctor_set(x_109, 1, x_108); -lean_ctor_set(x_109, 2, x_106); -x_110 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; -x_111 = lean_array_push(x_110, x_102); -x_112 = lean_array_push(x_111, x_109); -x_113 = l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__8; -x_114 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_114, 0, x_107); -lean_ctor_set(x_114, 1, x_113); -lean_ctor_set(x_114, 2, x_112); -x_115 = lean_array_push(x_105, x_114); -x_116 = l___aux__Init__Notation______macroRules__stx___x2b__1___closed__8; -x_117 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_117, 0, x_107); -lean_ctor_set(x_117, 1, x_116); -lean_ctor_set(x_117, 2, x_115); -x_118 = lean_array_push(x_110, x_9); -x_119 = lean_array_push(x_118, x_117); -x_120 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_120, 0, x_107); -lean_ctor_set(x_120, 1, x_116); -lean_ctor_set(x_120, 2, x_119); -x_121 = l_prec_x28___x29___closed__7; -lean_inc(x_88); -x_122 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_122, 0, x_88); -lean_ctor_set(x_122, 1, x_121); -x_123 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_124 = lean_array_push(x_123, x_98); -lean_inc(x_124); -x_125 = lean_array_push(x_124, x_120); -lean_inc(x_122); -x_126 = lean_array_push(x_125, x_122); -x_127 = l___aux__Init__Notation______macroRules__term___u2209____1___closed__2; -x_128 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_128, 0, x_107); -lean_ctor_set(x_128, 1, x_127); -lean_ctor_set(x_128, 2, x_126); -x_129 = lean_array_push(x_105, x_128); -x_130 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_130, 0, x_107); -lean_ctor_set(x_130, 1, x_116); -lean_ctor_set(x_130, 2, x_129); -x_131 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__28; -x_132 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_132, 0, x_88); -lean_ctor_set(x_132, 1, x_131); -x_133 = lean_array_push(x_123, x_130); -x_134 = lean_array_push(x_133, x_132); -x_135 = lean_array_push(x_134, x_19); -x_136 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__13; -x_137 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_137, 0, x_107); -lean_ctor_set(x_137, 1, x_136); -lean_ctor_set(x_137, 2, x_135); -x_138 = lean_array_push(x_110, x_100); -x_139 = lean_array_push(x_138, x_137); -x_140 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__11; -x_141 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_141, 0, x_107); -lean_ctor_set(x_141, 1, x_140); -lean_ctor_set(x_141, 2, x_139); -x_142 = lean_array_push(x_110, x_141); -x_143 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__12; -x_144 = lean_array_push(x_142, x_143); -x_145 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_145, 0, x_107); -lean_ctor_set(x_145, 1, x_116); -lean_ctor_set(x_145, 2, x_144); -x_146 = lean_array_push(x_124, x_145); -x_147 = lean_array_push(x_146, x_122); -x_148 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_148, 0, x_107); -lean_ctor_set(x_148, 1, x_127); -lean_ctor_set(x_148, 2, x_147); -x_149 = lean_array_push(x_105, x_148); -x_150 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_150, 0, x_107); -lean_ctor_set(x_150, 1, x_116); -lean_ctor_set(x_150, 2, x_149); -x_151 = lean_array_push(x_110, x_96); -x_152 = lean_array_push(x_151, x_150); -x_153 = l___aux__Init__Notation______macroRules__term___u2218____1___closed__4; -x_154 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_154, 0, x_107); -lean_ctor_set(x_154, 1, x_153); -lean_ctor_set(x_154, 2, x_152); -x_155 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_155, 0, x_154); -lean_ctor_set(x_155, 1, x_89); -return x_155; -} -} -} -else -{ -lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; uint8_t x_160; -x_156 = l_Lean_Syntax_getArg(x_11, x_8); +x_91 = l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__4; +x_92 = l_Lean_addMacroScope(x_90, x_91, x_89); +x_93 = l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__3; +x_94 = l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__6; +lean_inc(x_87); +x_95 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_95, 0, x_87); +lean_ctor_set(x_95, 1, x_93); +lean_ctor_set(x_95, 2, x_92); +lean_ctor_set(x_95, 3, x_94); +x_96 = l_prec_x28___x29___closed__3; +lean_inc(x_87); +x_97 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_97, 0, x_87); +lean_ctor_set(x_97, 1, x_96); +x_98 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__10; +lean_inc(x_87); +x_99 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_99, 0, x_87); +lean_ctor_set(x_99, 1, x_98); +x_100 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__21; +lean_inc(x_87); +x_101 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_101, 0, x_87); +lean_ctor_set(x_101, 1, x_100); +x_102 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__18; +lean_inc(x_87); +x_103 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_103, 0, x_87); +lean_ctor_set(x_103, 1, x_102); +x_104 = l___aux__Init__Notation______macroRules__precMax__1___closed__4; +x_105 = lean_array_push(x_104, x_103); +x_106 = lean_box(2); +x_107 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__17; +x_108 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_108, 0, x_106); +lean_ctor_set(x_108, 1, x_107); +lean_ctor_set(x_108, 2, x_105); +x_109 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; +x_110 = lean_array_push(x_109, x_101); +x_111 = lean_array_push(x_110, x_108); +x_112 = l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__8; +x_113 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_113, 0, x_106); +lean_ctor_set(x_113, 1, x_112); +lean_ctor_set(x_113, 2, x_111); +x_114 = lean_array_push(x_104, x_113); +x_115 = l___aux__Init__Notation______macroRules__stx___x2b__1___closed__8; +x_116 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_116, 0, x_106); +lean_ctor_set(x_116, 1, x_115); +lean_ctor_set(x_116, 2, x_114); +x_117 = lean_array_push(x_109, x_9); +x_118 = lean_array_push(x_117, x_116); +x_119 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_119, 0, x_106); +lean_ctor_set(x_119, 1, x_115); +lean_ctor_set(x_119, 2, x_118); +x_120 = l_prec_x28___x29___closed__7; +lean_inc(x_87); +x_121 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_121, 0, x_87); +lean_ctor_set(x_121, 1, x_120); +x_122 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_123 = lean_array_push(x_122, x_97); +lean_inc(x_123); +x_124 = lean_array_push(x_123, x_119); +lean_inc(x_121); +x_125 = lean_array_push(x_124, x_121); +x_126 = l___aux__Init__Notation______macroRules__term___u2209____1___closed__2; +x_127 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_127, 0, x_106); +lean_ctor_set(x_127, 1, x_126); +lean_ctor_set(x_127, 2, x_125); +x_128 = lean_array_push(x_104, x_127); +x_129 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_129, 0, x_106); +lean_ctor_set(x_129, 1, x_115); +lean_ctor_set(x_129, 2, x_128); +x_130 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__28; +x_131 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_131, 0, x_87); +lean_ctor_set(x_131, 1, x_130); +x_132 = lean_array_push(x_122, x_129); +x_133 = lean_array_push(x_132, x_131); +x_134 = lean_array_push(x_133, x_18); +x_135 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__13; +x_136 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_136, 0, x_106); +lean_ctor_set(x_136, 1, x_135); +lean_ctor_set(x_136, 2, x_134); +x_137 = lean_array_push(x_109, x_99); +x_138 = lean_array_push(x_137, x_136); +x_139 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__11; +x_140 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_140, 0, x_106); +lean_ctor_set(x_140, 1, x_139); +lean_ctor_set(x_140, 2, x_138); +x_141 = lean_array_push(x_109, x_140); +x_142 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__12; +x_143 = lean_array_push(x_141, x_142); +x_144 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_144, 0, x_106); +lean_ctor_set(x_144, 1, x_115); +lean_ctor_set(x_144, 2, x_143); +x_145 = lean_array_push(x_123, x_144); +x_146 = lean_array_push(x_145, x_121); +x_147 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_147, 0, x_106); +lean_ctor_set(x_147, 1, x_126); +lean_ctor_set(x_147, 2, x_146); +x_148 = lean_array_push(x_104, x_147); +x_149 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_149, 0, x_106); +lean_ctor_set(x_149, 1, x_115); +lean_ctor_set(x_149, 2, x_148); +x_150 = lean_array_push(x_109, x_95); +x_151 = lean_array_push(x_150, x_149); +x_152 = l___aux__Init__Notation______macroRules__term___u2218____1___closed__4; +x_153 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_153, 0, x_106); +lean_ctor_set(x_153, 1, x_152); +lean_ctor_set(x_153, 2, x_151); +x_154 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_154, 0, x_153); +lean_ctor_set(x_154, 1, x_88); +return x_154; +} +} +} +else +{ +lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; uint8_t x_159; +x_155 = l_Lean_Syntax_getArg(x_11, x_8); lean_dec(x_11); -x_157 = lean_unsigned_to_nat(4u); -x_158 = l_Lean_Syntax_getArg(x_1, x_157); +x_156 = lean_unsigned_to_nat(4u); +x_157 = l_Lean_Syntax_getArg(x_1, x_156); lean_dec(x_1); lean_inc(x_2); -x_159 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_160 = !lean_is_exclusive(x_159); -if (x_160 == 0) +x_158 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_159 = !lean_is_exclusive(x_158); +if (x_159 == 0) { -lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; -x_161 = lean_ctor_get(x_159, 0); -x_162 = lean_ctor_get(x_2, 2); +lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; +x_160 = lean_ctor_get(x_158, 0); +x_161 = lean_ctor_get(x_2, 2); +lean_inc(x_161); +x_162 = lean_ctor_get(x_2, 1); lean_inc(x_162); -x_163 = lean_ctor_get(x_2, 1); -lean_inc(x_163); lean_dec(x_2); -x_164 = l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__4; -x_165 = l_Lean_addMacroScope(x_163, x_164, x_162); -x_166 = l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__3; -x_167 = l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__6; -lean_inc(x_161); -x_168 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_168, 0, x_161); -lean_ctor_set(x_168, 1, x_166); -lean_ctor_set(x_168, 2, x_165); -lean_ctor_set(x_168, 3, x_167); -x_169 = l_prec_x28___x29___closed__3; -lean_inc(x_161); -x_170 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_170, 0, x_161); -lean_ctor_set(x_170, 1, x_169); -x_171 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__10; -lean_inc(x_161); -x_172 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_172, 0, x_161); -lean_ctor_set(x_172, 1, x_171); -x_173 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__21; -lean_inc(x_161); -x_174 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_174, 0, x_161); -lean_ctor_set(x_174, 1, x_173); -x_175 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; -x_176 = lean_array_push(x_175, x_174); -x_177 = lean_array_push(x_176, x_156); -x_178 = lean_box(2); -x_179 = l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__8; -x_180 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_180, 0, x_178); -lean_ctor_set(x_180, 1, x_179); -lean_ctor_set(x_180, 2, x_177); -x_181 = l___aux__Init__Notation______macroRules__precMax__1___closed__4; -x_182 = lean_array_push(x_181, x_180); -x_183 = l___aux__Init__Notation______macroRules__stx___x2b__1___closed__8; -x_184 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_184, 0, x_178); -lean_ctor_set(x_184, 1, x_183); -lean_ctor_set(x_184, 2, x_182); -x_185 = lean_array_push(x_175, x_9); -x_186 = lean_array_push(x_185, x_184); -x_187 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_187, 0, x_178); -lean_ctor_set(x_187, 1, x_183); -lean_ctor_set(x_187, 2, x_186); -x_188 = l_prec_x28___x29___closed__7; -lean_inc(x_161); -x_189 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_189, 0, x_161); -lean_ctor_set(x_189, 1, x_188); -x_190 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_191 = lean_array_push(x_190, x_170); -lean_inc(x_191); -x_192 = lean_array_push(x_191, x_187); -lean_inc(x_189); -x_193 = lean_array_push(x_192, x_189); -x_194 = l___aux__Init__Notation______macroRules__term___u2209____1___closed__2; -x_195 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_195, 0, x_178); -lean_ctor_set(x_195, 1, x_194); -lean_ctor_set(x_195, 2, x_193); -x_196 = lean_array_push(x_181, x_195); -x_197 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_197, 0, x_178); -lean_ctor_set(x_197, 1, x_183); -lean_ctor_set(x_197, 2, x_196); -x_198 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__28; -x_199 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_199, 0, x_161); -lean_ctor_set(x_199, 1, x_198); -x_200 = lean_array_push(x_190, x_197); -x_201 = lean_array_push(x_200, x_199); -x_202 = lean_array_push(x_201, x_158); -x_203 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__13; -x_204 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_204, 0, x_178); -lean_ctor_set(x_204, 1, x_203); -lean_ctor_set(x_204, 2, x_202); -x_205 = lean_array_push(x_175, x_172); -x_206 = lean_array_push(x_205, x_204); -x_207 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__11; -x_208 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_208, 0, x_178); -lean_ctor_set(x_208, 1, x_207); -lean_ctor_set(x_208, 2, x_206); -x_209 = lean_array_push(x_175, x_208); -x_210 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__12; -x_211 = lean_array_push(x_209, x_210); -x_212 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_212, 0, x_178); -lean_ctor_set(x_212, 1, x_183); -lean_ctor_set(x_212, 2, x_211); -x_213 = lean_array_push(x_191, x_212); -x_214 = lean_array_push(x_213, x_189); -x_215 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_215, 0, x_178); -lean_ctor_set(x_215, 1, x_194); -lean_ctor_set(x_215, 2, x_214); -x_216 = lean_array_push(x_181, x_215); -x_217 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_217, 0, x_178); -lean_ctor_set(x_217, 1, x_183); -lean_ctor_set(x_217, 2, x_216); -x_218 = lean_array_push(x_175, x_168); -x_219 = lean_array_push(x_218, x_217); -x_220 = l___aux__Init__Notation______macroRules__term___u2218____1___closed__4; -x_221 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_221, 0, x_178); -lean_ctor_set(x_221, 1, x_220); -lean_ctor_set(x_221, 2, x_219); -lean_ctor_set(x_159, 0, x_221); -return x_159; -} -else -{ -lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; -x_222 = lean_ctor_get(x_159, 0); -x_223 = lean_ctor_get(x_159, 1); -lean_inc(x_223); +x_163 = l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__4; +x_164 = l_Lean_addMacroScope(x_162, x_163, x_161); +x_165 = l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__3; +x_166 = l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__6; +lean_inc(x_160); +x_167 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_167, 0, x_160); +lean_ctor_set(x_167, 1, x_165); +lean_ctor_set(x_167, 2, x_164); +lean_ctor_set(x_167, 3, x_166); +x_168 = l_prec_x28___x29___closed__3; +lean_inc(x_160); +x_169 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_169, 0, x_160); +lean_ctor_set(x_169, 1, x_168); +x_170 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__10; +lean_inc(x_160); +x_171 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_171, 0, x_160); +lean_ctor_set(x_171, 1, x_170); +x_172 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__21; +lean_inc(x_160); +x_173 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_173, 0, x_160); +lean_ctor_set(x_173, 1, x_172); +x_174 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; +x_175 = lean_array_push(x_174, x_173); +x_176 = lean_array_push(x_175, x_155); +x_177 = lean_box(2); +x_178 = l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__8; +x_179 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_179, 0, x_177); +lean_ctor_set(x_179, 1, x_178); +lean_ctor_set(x_179, 2, x_176); +x_180 = l___aux__Init__Notation______macroRules__precMax__1___closed__4; +x_181 = lean_array_push(x_180, x_179); +x_182 = l___aux__Init__Notation______macroRules__stx___x2b__1___closed__8; +x_183 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_183, 0, x_177); +lean_ctor_set(x_183, 1, x_182); +lean_ctor_set(x_183, 2, x_181); +x_184 = lean_array_push(x_174, x_9); +x_185 = lean_array_push(x_184, x_183); +x_186 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_186, 0, x_177); +lean_ctor_set(x_186, 1, x_182); +lean_ctor_set(x_186, 2, x_185); +x_187 = l_prec_x28___x29___closed__7; +lean_inc(x_160); +x_188 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_188, 0, x_160); +lean_ctor_set(x_188, 1, x_187); +x_189 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_190 = lean_array_push(x_189, x_169); +lean_inc(x_190); +x_191 = lean_array_push(x_190, x_186); +lean_inc(x_188); +x_192 = lean_array_push(x_191, x_188); +x_193 = l___aux__Init__Notation______macroRules__term___u2209____1___closed__2; +x_194 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_194, 0, x_177); +lean_ctor_set(x_194, 1, x_193); +lean_ctor_set(x_194, 2, x_192); +x_195 = lean_array_push(x_180, x_194); +x_196 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_196, 0, x_177); +lean_ctor_set(x_196, 1, x_182); +lean_ctor_set(x_196, 2, x_195); +x_197 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__28; +x_198 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_198, 0, x_160); +lean_ctor_set(x_198, 1, x_197); +x_199 = lean_array_push(x_189, x_196); +x_200 = lean_array_push(x_199, x_198); +x_201 = lean_array_push(x_200, x_157); +x_202 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__13; +x_203 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_203, 0, x_177); +lean_ctor_set(x_203, 1, x_202); +lean_ctor_set(x_203, 2, x_201); +x_204 = lean_array_push(x_174, x_171); +x_205 = lean_array_push(x_204, x_203); +x_206 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__11; +x_207 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_207, 0, x_177); +lean_ctor_set(x_207, 1, x_206); +lean_ctor_set(x_207, 2, x_205); +x_208 = lean_array_push(x_174, x_207); +x_209 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__12; +x_210 = lean_array_push(x_208, x_209); +x_211 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_211, 0, x_177); +lean_ctor_set(x_211, 1, x_182); +lean_ctor_set(x_211, 2, x_210); +x_212 = lean_array_push(x_190, x_211); +x_213 = lean_array_push(x_212, x_188); +x_214 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_214, 0, x_177); +lean_ctor_set(x_214, 1, x_193); +lean_ctor_set(x_214, 2, x_213); +x_215 = lean_array_push(x_180, x_214); +x_216 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_216, 0, x_177); +lean_ctor_set(x_216, 1, x_182); +lean_ctor_set(x_216, 2, x_215); +x_217 = lean_array_push(x_174, x_167); +x_218 = lean_array_push(x_217, x_216); +x_219 = l___aux__Init__Notation______macroRules__term___u2218____1___closed__4; +x_220 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_220, 0, x_177); +lean_ctor_set(x_220, 1, x_219); +lean_ctor_set(x_220, 2, x_218); +lean_ctor_set(x_158, 0, x_220); +return x_158; +} +else +{ +lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; +x_221 = lean_ctor_get(x_158, 0); +x_222 = lean_ctor_get(x_158, 1); lean_inc(x_222); -lean_dec(x_159); -x_224 = lean_ctor_get(x_2, 2); +lean_inc(x_221); +lean_dec(x_158); +x_223 = lean_ctor_get(x_2, 2); +lean_inc(x_223); +x_224 = lean_ctor_get(x_2, 1); lean_inc(x_224); -x_225 = lean_ctor_get(x_2, 1); -lean_inc(x_225); lean_dec(x_2); -x_226 = l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__4; -x_227 = l_Lean_addMacroScope(x_225, x_226, x_224); -x_228 = l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__3; -x_229 = l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__6; -lean_inc(x_222); -x_230 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_230, 0, x_222); -lean_ctor_set(x_230, 1, x_228); -lean_ctor_set(x_230, 2, x_227); -lean_ctor_set(x_230, 3, x_229); -x_231 = l_prec_x28___x29___closed__3; -lean_inc(x_222); -x_232 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_232, 0, x_222); -lean_ctor_set(x_232, 1, x_231); -x_233 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__10; -lean_inc(x_222); -x_234 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_234, 0, x_222); -lean_ctor_set(x_234, 1, x_233); -x_235 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__21; -lean_inc(x_222); -x_236 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_236, 0, x_222); -lean_ctor_set(x_236, 1, x_235); -x_237 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; -x_238 = lean_array_push(x_237, x_236); -x_239 = lean_array_push(x_238, x_156); -x_240 = lean_box(2); -x_241 = l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__8; -x_242 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_242, 0, x_240); -lean_ctor_set(x_242, 1, x_241); -lean_ctor_set(x_242, 2, x_239); -x_243 = l___aux__Init__Notation______macroRules__precMax__1___closed__4; -x_244 = lean_array_push(x_243, x_242); -x_245 = l___aux__Init__Notation______macroRules__stx___x2b__1___closed__8; -x_246 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_246, 0, x_240); -lean_ctor_set(x_246, 1, x_245); -lean_ctor_set(x_246, 2, x_244); -x_247 = lean_array_push(x_237, x_9); -x_248 = lean_array_push(x_247, x_246); -x_249 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_249, 0, x_240); -lean_ctor_set(x_249, 1, x_245); -lean_ctor_set(x_249, 2, x_248); -x_250 = l_prec_x28___x29___closed__7; -lean_inc(x_222); -x_251 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_251, 0, x_222); -lean_ctor_set(x_251, 1, x_250); -x_252 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; -x_253 = lean_array_push(x_252, x_232); -lean_inc(x_253); -x_254 = lean_array_push(x_253, x_249); -lean_inc(x_251); -x_255 = lean_array_push(x_254, x_251); -x_256 = l___aux__Init__Notation______macroRules__term___u2209____1___closed__2; -x_257 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_257, 0, x_240); -lean_ctor_set(x_257, 1, x_256); -lean_ctor_set(x_257, 2, x_255); -x_258 = lean_array_push(x_243, x_257); -x_259 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_259, 0, x_240); -lean_ctor_set(x_259, 1, x_245); -lean_ctor_set(x_259, 2, x_258); -x_260 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__28; -x_261 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_261, 0, x_222); -lean_ctor_set(x_261, 1, x_260); -x_262 = lean_array_push(x_252, x_259); -x_263 = lean_array_push(x_262, x_261); -x_264 = lean_array_push(x_263, x_158); -x_265 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__13; -x_266 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_266, 0, x_240); -lean_ctor_set(x_266, 1, x_265); -lean_ctor_set(x_266, 2, x_264); -x_267 = lean_array_push(x_237, x_234); -x_268 = lean_array_push(x_267, x_266); -x_269 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__11; -x_270 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_270, 0, x_240); -lean_ctor_set(x_270, 1, x_269); -lean_ctor_set(x_270, 2, x_268); -x_271 = lean_array_push(x_237, x_270); -x_272 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__12; -x_273 = lean_array_push(x_271, x_272); -x_274 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_274, 0, x_240); -lean_ctor_set(x_274, 1, x_245); -lean_ctor_set(x_274, 2, x_273); -x_275 = lean_array_push(x_253, x_274); -x_276 = lean_array_push(x_275, x_251); -x_277 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_277, 0, x_240); -lean_ctor_set(x_277, 1, x_256); -lean_ctor_set(x_277, 2, x_276); -x_278 = lean_array_push(x_243, x_277); -x_279 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_279, 0, x_240); -lean_ctor_set(x_279, 1, x_245); -lean_ctor_set(x_279, 2, x_278); -x_280 = lean_array_push(x_237, x_230); -x_281 = lean_array_push(x_280, x_279); -x_282 = l___aux__Init__Notation______macroRules__term___u2218____1___closed__4; -x_283 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_283, 0, x_240); -lean_ctor_set(x_283, 1, x_282); -lean_ctor_set(x_283, 2, x_281); -x_284 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_284, 0, x_283); -lean_ctor_set(x_284, 1, x_223); -return x_284; +x_225 = l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__4; +x_226 = l_Lean_addMacroScope(x_224, x_225, x_223); +x_227 = l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__3; +x_228 = l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__6; +lean_inc(x_221); +x_229 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_229, 0, x_221); +lean_ctor_set(x_229, 1, x_227); +lean_ctor_set(x_229, 2, x_226); +lean_ctor_set(x_229, 3, x_228); +x_230 = l_prec_x28___x29___closed__3; +lean_inc(x_221); +x_231 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_231, 0, x_221); +lean_ctor_set(x_231, 1, x_230); +x_232 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__10; +lean_inc(x_221); +x_233 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_233, 0, x_221); +lean_ctor_set(x_233, 1, x_232); +x_234 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__21; +lean_inc(x_221); +x_235 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_235, 0, x_221); +lean_ctor_set(x_235, 1, x_234); +x_236 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; +x_237 = lean_array_push(x_236, x_235); +x_238 = lean_array_push(x_237, x_155); +x_239 = lean_box(2); +x_240 = l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__8; +x_241 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_241, 0, x_239); +lean_ctor_set(x_241, 1, x_240); +lean_ctor_set(x_241, 2, x_238); +x_242 = l___aux__Init__Notation______macroRules__precMax__1___closed__4; +x_243 = lean_array_push(x_242, x_241); +x_244 = l___aux__Init__Notation______macroRules__stx___x2b__1___closed__8; +x_245 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_245, 0, x_239); +lean_ctor_set(x_245, 1, x_244); +lean_ctor_set(x_245, 2, x_243); +x_246 = lean_array_push(x_236, x_9); +x_247 = lean_array_push(x_246, x_245); +x_248 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_248, 0, x_239); +lean_ctor_set(x_248, 1, x_244); +lean_ctor_set(x_248, 2, x_247); +x_249 = l_prec_x28___x29___closed__7; +lean_inc(x_221); +x_250 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_250, 0, x_221); +lean_ctor_set(x_250, 1, x_249); +x_251 = l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; +x_252 = lean_array_push(x_251, x_231); +lean_inc(x_252); +x_253 = lean_array_push(x_252, x_248); +lean_inc(x_250); +x_254 = lean_array_push(x_253, x_250); +x_255 = l___aux__Init__Notation______macroRules__term___u2209____1___closed__2; +x_256 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_256, 0, x_239); +lean_ctor_set(x_256, 1, x_255); +lean_ctor_set(x_256, 2, x_254); +x_257 = lean_array_push(x_242, x_256); +x_258 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_258, 0, x_239); +lean_ctor_set(x_258, 1, x_244); +lean_ctor_set(x_258, 2, x_257); +x_259 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__28; +x_260 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_260, 0, x_221); +lean_ctor_set(x_260, 1, x_259); +x_261 = lean_array_push(x_251, x_258); +x_262 = lean_array_push(x_261, x_260); +x_263 = lean_array_push(x_262, x_157); +x_264 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__13; +x_265 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_265, 0, x_239); +lean_ctor_set(x_265, 1, x_264); +lean_ctor_set(x_265, 2, x_263); +x_266 = lean_array_push(x_236, x_233); +x_267 = lean_array_push(x_266, x_265); +x_268 = l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__11; +x_269 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_269, 0, x_239); +lean_ctor_set(x_269, 1, x_268); +lean_ctor_set(x_269, 2, x_267); +x_270 = lean_array_push(x_236, x_269); +x_271 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__12; +x_272 = lean_array_push(x_270, x_271); +x_273 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_273, 0, x_239); +lean_ctor_set(x_273, 1, x_244); +lean_ctor_set(x_273, 2, x_272); +x_274 = lean_array_push(x_252, x_273); +x_275 = lean_array_push(x_274, x_250); +x_276 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_276, 0, x_239); +lean_ctor_set(x_276, 1, x_255); +lean_ctor_set(x_276, 2, x_275); +x_277 = lean_array_push(x_242, x_276); +x_278 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_278, 0, x_239); +lean_ctor_set(x_278, 1, x_244); +lean_ctor_set(x_278, 2, x_277); +x_279 = lean_array_push(x_236, x_229); +x_280 = lean_array_push(x_279, x_278); +x_281 = l___aux__Init__Notation______macroRules__term___u2218____1___closed__4; +x_282 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_282, 0, x_239); +lean_ctor_set(x_282, 1, x_281); +lean_ctor_set(x_282, 2, x_280); +x_283 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_283, 0, x_282); +lean_ctor_set(x_283, 1, x_222); +return x_283; } } } @@ -30829,43 +30788,53 @@ static lean_object* _init_l_Lean_rawStx_quot___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("`(rawStx|", 9); +x_1 = lean_mk_string_from_bytes("rawStx", 6); return x_1; } } static lean_object* _init_l_Lean_rawStx_quot___closed__4() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_rawStx_quot___closed__3; -x_2 = lean_alloc_ctor(5, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_rawStx_quot___closed__3; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; } } static lean_object* _init_l_Lean_rawStx_quot___closed__5() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("rawStx", 6); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_rawStx_quot___closed__4; +x_2 = l_Lean_rawStx_quot___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; } } static lean_object* _init_l_Lean_rawStx_quot___closed__6() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_rawStx_quot___closed__5; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("`(rawStx|", 9); +return x_1; } } static lean_object* _init_l_Lean_rawStx_quot___closed__7() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; +lean_object* x_1; lean_object* x_2; x_1 = l_Lean_rawStx_quot___closed__6; +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_rawStx_quot___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_rawStx_quot___closed__4; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30873,12 +30842,12 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_rawStx_quot___closed__8() { +static lean_object* _init_l_Lean_rawStx_quot___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Syntax_addPrec___closed__10; -x_2 = l_Lean_rawStx_quot___closed__7; +x_2 = l_Lean_rawStx_quot___closed__8; x_3 = l_prec_x28___x29___closed__8; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -30887,13 +30856,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_rawStx_quot___closed__9() { +static lean_object* _init_l_Lean_rawStx_quot___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Syntax_addPrec___closed__10; -x_2 = l_Lean_rawStx_quot___closed__4; -x_3 = l_Lean_rawStx_quot___closed__8; +x_2 = l_Lean_rawStx_quot___closed__7; +x_3 = l_Lean_rawStx_quot___closed__9; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -30901,13 +30870,27 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_rawStx_quot___closed__10() { +static lean_object* _init_l_Lean_rawStx_quot___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_rawStx_quot___closed__5; +x_2 = lean_unsigned_to_nat(1024u); +x_3 = l_Lean_rawStx_quot___closed__10; +x_4 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_rawStx_quot___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_rawStx_quot___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_rawStx_quot___closed__9; +x_3 = l_Lean_rawStx_quot___closed__11; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -30919,10 +30902,26 @@ static lean_object* _init_l_Lean_rawStx_quot() { _start: { lean_object* x_1; -x_1 = l_Lean_rawStx_quot___closed__10; +x_1 = l_Lean_rawStx_quot___closed__12; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_instCoeSyntaxTSyntaxConsSyntaxNodeKindMkStrAnonymousNil(lean_object* x_1) { +_start: +{ +lean_inc(x_1); return x_1; } } +LEAN_EXPORT lean_object* l_Lean_instCoeSyntaxTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_instCoeSyntaxTSyntaxConsSyntaxNodeKindMkStrAnonymousNil(x_1); +lean_dec(x_1); +return x_2; +} +} lean_object* initialize_Init_Prelude(uint8_t builtin, lean_object*); lean_object* initialize_Init_Coe(uint8_t builtin, lean_object*); static bool _G_initialized = false; @@ -33056,56 +33055,56 @@ l_term___x24_______closed__13 = _init_l_term___x24_______closed__13(); lean_mark_persistent(l_term___x24_______closed__13); l_term___x24____ = _init_l_term___x24____(); lean_mark_persistent(l_term___x24____); -l_term_x7b_____x3a___x2f_x2f___x7d___closed__1 = _init_l_term_x7b_____x3a___x2f_x2f___x7d___closed__1(); -lean_mark_persistent(l_term_x7b_____x3a___x2f_x2f___x7d___closed__1); -l_term_x7b_____x3a___x2f_x2f___x7d___closed__2 = _init_l_term_x7b_____x3a___x2f_x2f___x7d___closed__2(); -lean_mark_persistent(l_term_x7b_____x3a___x2f_x2f___x7d___closed__2); -l_term_x7b_____x3a___x2f_x2f___x7d___closed__3 = _init_l_term_x7b_____x3a___x2f_x2f___x7d___closed__3(); -lean_mark_persistent(l_term_x7b_____x3a___x2f_x2f___x7d___closed__3); -l_term_x7b_____x3a___x2f_x2f___x7d___closed__4 = _init_l_term_x7b_____x3a___x2f_x2f___x7d___closed__4(); -lean_mark_persistent(l_term_x7b_____x3a___x2f_x2f___x7d___closed__4); -l_term_x7b_____x3a___x2f_x2f___x7d___closed__5 = _init_l_term_x7b_____x3a___x2f_x2f___x7d___closed__5(); -lean_mark_persistent(l_term_x7b_____x3a___x2f_x2f___x7d___closed__5); -l_term_x7b_____x3a___x2f_x2f___x7d___closed__6 = _init_l_term_x7b_____x3a___x2f_x2f___x7d___closed__6(); -lean_mark_persistent(l_term_x7b_____x3a___x2f_x2f___x7d___closed__6); -l_term_x7b_____x3a___x2f_x2f___x7d___closed__7 = _init_l_term_x7b_____x3a___x2f_x2f___x7d___closed__7(); -lean_mark_persistent(l_term_x7b_____x3a___x2f_x2f___x7d___closed__7); -l_term_x7b_____x3a___x2f_x2f___x7d___closed__8 = _init_l_term_x7b_____x3a___x2f_x2f___x7d___closed__8(); -lean_mark_persistent(l_term_x7b_____x3a___x2f_x2f___x7d___closed__8); -l_term_x7b_____x3a___x2f_x2f___x7d___closed__9 = _init_l_term_x7b_____x3a___x2f_x2f___x7d___closed__9(); -lean_mark_persistent(l_term_x7b_____x3a___x2f_x2f___x7d___closed__9); -l_term_x7b_____x3a___x2f_x2f___x7d___closed__10 = _init_l_term_x7b_____x3a___x2f_x2f___x7d___closed__10(); -lean_mark_persistent(l_term_x7b_____x3a___x2f_x2f___x7d___closed__10); -l_term_x7b_____x3a___x2f_x2f___x7d___closed__11 = _init_l_term_x7b_____x3a___x2f_x2f___x7d___closed__11(); -lean_mark_persistent(l_term_x7b_____x3a___x2f_x2f___x7d___closed__11); -l_term_x7b_____x3a___x2f_x2f___x7d___closed__12 = _init_l_term_x7b_____x3a___x2f_x2f___x7d___closed__12(); -lean_mark_persistent(l_term_x7b_____x3a___x2f_x2f___x7d___closed__12); -l_term_x7b_____x3a___x2f_x2f___x7d___closed__13 = _init_l_term_x7b_____x3a___x2f_x2f___x7d___closed__13(); -lean_mark_persistent(l_term_x7b_____x3a___x2f_x2f___x7d___closed__13); -l_term_x7b_____x3a___x2f_x2f___x7d___closed__14 = _init_l_term_x7b_____x3a___x2f_x2f___x7d___closed__14(); -lean_mark_persistent(l_term_x7b_____x3a___x2f_x2f___x7d___closed__14); -l_term_x7b_____x3a___x2f_x2f___x7d___closed__15 = _init_l_term_x7b_____x3a___x2f_x2f___x7d___closed__15(); -lean_mark_persistent(l_term_x7b_____x3a___x2f_x2f___x7d___closed__15); -l_term_x7b_____x3a___x2f_x2f___x7d___closed__16 = _init_l_term_x7b_____x3a___x2f_x2f___x7d___closed__16(); -lean_mark_persistent(l_term_x7b_____x3a___x2f_x2f___x7d___closed__16); -l_term_x7b_____x3a___x2f_x2f___x7d = _init_l_term_x7b_____x3a___x2f_x2f___x7d(); -lean_mark_persistent(l_term_x7b_____x3a___x2f_x2f___x7d); -l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__1 = _init_l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__1(); -lean_mark_persistent(l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__1); -l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__2 = _init_l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__2(); -lean_mark_persistent(l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__2); -l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__3 = _init_l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__3(); -lean_mark_persistent(l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__3); -l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__4 = _init_l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__4(); -lean_mark_persistent(l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__4); -l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__5 = _init_l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__5(); -lean_mark_persistent(l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__5); -l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__6 = _init_l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__6(); -lean_mark_persistent(l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__6); -l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__7 = _init_l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__7(); -lean_mark_persistent(l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__7); -l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__8 = _init_l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__8(); -lean_mark_persistent(l___aux__Init__Notation______macroRules__term_x7b_____x3a___x2f_x2f___x7d__1___closed__8); +l_term_x7b___x3a___x2f_x2f___x7d___closed__1 = _init_l_term_x7b___x3a___x2f_x2f___x7d___closed__1(); +lean_mark_persistent(l_term_x7b___x3a___x2f_x2f___x7d___closed__1); +l_term_x7b___x3a___x2f_x2f___x7d___closed__2 = _init_l_term_x7b___x3a___x2f_x2f___x7d___closed__2(); +lean_mark_persistent(l_term_x7b___x3a___x2f_x2f___x7d___closed__2); +l_term_x7b___x3a___x2f_x2f___x7d___closed__3 = _init_l_term_x7b___x3a___x2f_x2f___x7d___closed__3(); +lean_mark_persistent(l_term_x7b___x3a___x2f_x2f___x7d___closed__3); +l_term_x7b___x3a___x2f_x2f___x7d___closed__4 = _init_l_term_x7b___x3a___x2f_x2f___x7d___closed__4(); +lean_mark_persistent(l_term_x7b___x3a___x2f_x2f___x7d___closed__4); +l_term_x7b___x3a___x2f_x2f___x7d___closed__5 = _init_l_term_x7b___x3a___x2f_x2f___x7d___closed__5(); +lean_mark_persistent(l_term_x7b___x3a___x2f_x2f___x7d___closed__5); +l_term_x7b___x3a___x2f_x2f___x7d___closed__6 = _init_l_term_x7b___x3a___x2f_x2f___x7d___closed__6(); +lean_mark_persistent(l_term_x7b___x3a___x2f_x2f___x7d___closed__6); +l_term_x7b___x3a___x2f_x2f___x7d___closed__7 = _init_l_term_x7b___x3a___x2f_x2f___x7d___closed__7(); +lean_mark_persistent(l_term_x7b___x3a___x2f_x2f___x7d___closed__7); +l_term_x7b___x3a___x2f_x2f___x7d___closed__8 = _init_l_term_x7b___x3a___x2f_x2f___x7d___closed__8(); +lean_mark_persistent(l_term_x7b___x3a___x2f_x2f___x7d___closed__8); +l_term_x7b___x3a___x2f_x2f___x7d___closed__9 = _init_l_term_x7b___x3a___x2f_x2f___x7d___closed__9(); +lean_mark_persistent(l_term_x7b___x3a___x2f_x2f___x7d___closed__9); +l_term_x7b___x3a___x2f_x2f___x7d___closed__10 = _init_l_term_x7b___x3a___x2f_x2f___x7d___closed__10(); +lean_mark_persistent(l_term_x7b___x3a___x2f_x2f___x7d___closed__10); +l_term_x7b___x3a___x2f_x2f___x7d___closed__11 = _init_l_term_x7b___x3a___x2f_x2f___x7d___closed__11(); +lean_mark_persistent(l_term_x7b___x3a___x2f_x2f___x7d___closed__11); +l_term_x7b___x3a___x2f_x2f___x7d___closed__12 = _init_l_term_x7b___x3a___x2f_x2f___x7d___closed__12(); +lean_mark_persistent(l_term_x7b___x3a___x2f_x2f___x7d___closed__12); +l_term_x7b___x3a___x2f_x2f___x7d___closed__13 = _init_l_term_x7b___x3a___x2f_x2f___x7d___closed__13(); +lean_mark_persistent(l_term_x7b___x3a___x2f_x2f___x7d___closed__13); +l_term_x7b___x3a___x2f_x2f___x7d___closed__14 = _init_l_term_x7b___x3a___x2f_x2f___x7d___closed__14(); +lean_mark_persistent(l_term_x7b___x3a___x2f_x2f___x7d___closed__14); +l_term_x7b___x3a___x2f_x2f___x7d___closed__15 = _init_l_term_x7b___x3a___x2f_x2f___x7d___closed__15(); +lean_mark_persistent(l_term_x7b___x3a___x2f_x2f___x7d___closed__15); +l_term_x7b___x3a___x2f_x2f___x7d___closed__16 = _init_l_term_x7b___x3a___x2f_x2f___x7d___closed__16(); +lean_mark_persistent(l_term_x7b___x3a___x2f_x2f___x7d___closed__16); +l_term_x7b___x3a___x2f_x2f___x7d = _init_l_term_x7b___x3a___x2f_x2f___x7d(); +lean_mark_persistent(l_term_x7b___x3a___x2f_x2f___x7d); +l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__1 = _init_l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__1(); +lean_mark_persistent(l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__1); +l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__2 = _init_l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__2(); +lean_mark_persistent(l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__2); +l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__3 = _init_l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__3(); +lean_mark_persistent(l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__3); +l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__4 = _init_l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__4(); +lean_mark_persistent(l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__4); +l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__5 = _init_l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__5(); +lean_mark_persistent(l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__5); +l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__6 = _init_l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__6(); +lean_mark_persistent(l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__6); +l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__7 = _init_l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__7(); +lean_mark_persistent(l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__7); +l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__8 = _init_l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__8(); +lean_mark_persistent(l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__8); l_termWithout__expected__type_____closed__1 = _init_l_termWithout__expected__type_____closed__1(); lean_mark_persistent(l_termWithout__expected__type_____closed__1); l_termWithout__expected__type_____closed__2 = _init_l_termWithout__expected__type_____closed__2(); @@ -33240,6 +33239,10 @@ l_Lean_rawStx_quot___closed__9 = _init_l_Lean_rawStx_quot___closed__9(); lean_mark_persistent(l_Lean_rawStx_quot___closed__9); l_Lean_rawStx_quot___closed__10 = _init_l_Lean_rawStx_quot___closed__10(); lean_mark_persistent(l_Lean_rawStx_quot___closed__10); +l_Lean_rawStx_quot___closed__11 = _init_l_Lean_rawStx_quot___closed__11(); +lean_mark_persistent(l_Lean_rawStx_quot___closed__11); +l_Lean_rawStx_quot___closed__12 = _init_l_Lean_rawStx_quot___closed__12(); +lean_mark_persistent(l_Lean_rawStx_quot___closed__12); l_Lean_rawStx_quot = _init_l_Lean_rawStx_quot(); lean_mark_persistent(l_Lean_rawStx_quot); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Init/NotationExtra.c b/stage0/stdlib/Init/NotationExtra.c index 8e3fe5c9141..278601c3922 100644 --- a/stage0/stdlib/Init/NotationExtra.c +++ b/stage0/stdlib/Init/NotationExtra.c @@ -13,31 +13,29 @@ #ifdef __cplusplus extern "C" { #endif -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__15; static lean_object* l_calcStep___closed__7; +static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__3; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__28; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile__Do____1___closed__13; -static lean_object* l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__27; static lean_object* l_Lean_unbracketedExplicitBinders___closed__4; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__18; -static lean_object* l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__4; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__14; LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__termExists___x2c____1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__5; LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__2(lean_object*); static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__25; static lean_object* l_unexpandListCons___closed__1; +static lean_object* l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__31; static lean_object* l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__12; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__13; LEAN_EXPORT lean_object* l_term_u03a3_x27___x2c__; static lean_object* l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__7; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__3(size_t, size_t, lean_object*); static lean_object* l_termExists___x2c_____closed__2; LEAN_EXPORT lean_object* l_unexpandSigma(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__27; lean_object* l_Lean_extractMacroScopes(lean_object*); -static lean_object* l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__17; size_t lean_usize_add(size_t, size_t); static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__29; -static lean_object* l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__8; LEAN_EXPORT lean_object* l_calcStep; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__9; LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___rarg___lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_term_u03a3___x2c_____closed__6; static lean_object* l_Lean_unbracketedExplicitBinders___closed__1; @@ -46,100 +44,98 @@ static lean_object* l_Lean_unifConstraint___closed__1; static lean_object* l_tactic_xb7_x2e_____x3b_____closed__22; static lean_object* l_Lean_expandExplicitBindersAux_loop___closed__5; static lean_object* l_term_u03a3_x27___x2c_____closed__4; +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at___aux__Init__NotationExtra______macroRules__solve__1___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_erase_macro_scopes(lean_object*); static lean_object* l_unexpandIte___closed__4; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__24; LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__19; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__32; static lean_object* l_unexpandListNil___rarg___closed__2; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__18; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__10; static lean_object* l_Lean_doElemWhile__Do_____closed__5; LEAN_EXPORT lean_object* l_term_u03a3___x2c__; static lean_object* l_Lean_doElemWhile__Do_____closed__3; lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__solve__1___spec__1___boxed(lean_object*, lean_object*); static lean_object* l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__13; static lean_object* l_term___xd7____1___closed__7; lean_object* lean_nat_div(lean_object*, lean_object*); lean_object* l_Lean_Syntax_getHeadInfo_x3f(lean_object*); static lean_object* l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__6; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__5; LEAN_EXPORT lean_object* l_unexpandSorryAx(lean_object*, lean_object*, lean_object*); static lean_object* l_termExists___x2c_____closed__5; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__14; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__30; static lean_object* l_Lean_explicitBinders___closed__4; -static lean_object* l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__9; -static lean_object* l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__9; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__22; static lean_object* l_unexpandSubtype___closed__4; static lean_object* l_Lean_unbracketedExplicitBinders___closed__2; -extern lean_object* l_Lean_nullKind; -LEAN_EXPORT lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_unifConstraint___closed__8; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4; static lean_object* l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__11; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__10; static lean_object* l_calc___closed__12; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__32; static lean_object* l_Lean_unifConstraint___closed__2; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__33; -static lean_object* l_Lean_binderIdent___closed__6; +static lean_object* l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__12; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__4___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_tactic_xb7_x2e_____x3b_____closed__5; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__22; static lean_object* l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__5; lean_object* lean_name_mk_string(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_expandExplicitBinders___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_getOptional_x3f(lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); LEAN_EXPORT lean_object* l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__1___boxed(lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); static lean_object* l_Lean_explicitBinders___closed__3; static lean_object* l_termExists___x2c_____closed__1; +static lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__solve__1___spec__3___closed__1; static lean_object* l_Lean_expandExplicitBinders___closed__1; LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__solve__1___lambda__1(lean_object*); static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile__Do____1___closed__2; +static lean_object* l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__23; +static lean_object* l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__7; static lean_object* l_Lean_unifConstraint___closed__3; static lean_object* l_tactic_xb7_x2e_____x3b_____closed__18; static lean_object* l___aux__Init__NotationExtra______macroRules__term_u2203___x2c____1___closed__2; lean_object* l_Array_append___rarg(lean_object*, lean_object*); static lean_object* l_tactic_xb7_x2e_____x3b_____closed__15; static lean_object* l_term_u2203___x2c_____closed__7; +static lean_object* l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__13; static lean_object* l_tacticFunext_______closed__7; LEAN_EXPORT lean_object* l_Lean_expandExplicitBindersAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_term_u03a3_x27___x2c_____closed__1; static lean_object* l___aux__Init__NotationExtra______macroRules__tacticCalc____1___closed__2; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__28; +static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__7; static lean_object* l_tactic_xb7_x2e_____x3b_____closed__3; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__7; +lean_object* l_Lean_SourceInfo_fromRef(lean_object*); static lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__6; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__14; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__1; -static lean_object* l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__16; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__16; static lean_object* l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__20; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__8; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__16; static lean_object* l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__3; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile__Do____1___closed__1; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__28; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__solve__1___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_unexpandIte___closed__1; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__14; static lean_object* l_calc___closed__10; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__17; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__37; +static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__16; +lean_object* l_Array_extract___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_tacticCalc__; static lean_object* l_tactic_xb7_x2e_____x3b_____closed__14; +LEAN_EXPORT lean_object* l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__solve__1___spec__1(lean_object*, lean_object*); static lean_object* l_calcStep___closed__10; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__3; static lean_object* l_calcStep___closed__11; static lean_object* l_unexpandSubtype___closed__2; size_t lean_usize_sub(size_t, size_t); LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__term_u03a3___x2c____1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_toCtorIdx___boxed(lean_object*); static lean_object* l_Lean_expandExplicitBindersAux_loop___closed__12; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__16; static lean_object* l_tacticFunext_______closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___spec__4(size_t, size_t, lean_object*); static lean_object* l_Lean_explicitBinders___closed__2; -static lean_object* l___aux__Init__NotationExtra______macroRules__solve__1___lambda__1___closed__1; static lean_object* l___aux__Init__NotationExtra______macroRules__tacticCalc____1___closed__1; static lean_object* l_term_u03a3_x27___x2c_____closed__2; static lean_object* l_Lean_term__Matches___x7c___closed__3; @@ -154,39 +150,40 @@ static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__ter uint8_t lean_name_eq(lean_object*, lean_object*); static lean_object* l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__31; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile__Do____1___closed__4; -static lean_object* l_Lean_binderIdent___closed__3; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__25; static lean_object* l_term_u03a3___x2c_____closed__5; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__6; static lean_object* l_Lean_doElemRepeat__Until_____closed__6; static lean_object* l_tacticFunext_______closed__1; static lean_object* l_Lean_doElemWhile__Do_____closed__6; -static lean_object* l_Lean_binderIdent___closed__5; -LEAN_EXPORT lean_object* l_Lean_termMacro_x2etrace_x5b_____x5d__; LEAN_EXPORT lean_object* l_tactic_xb7_x2e_____x3b__; static lean_object* l___aux__Init__NotationExtra______macroRules__term_u2203___x2c____1___closed__1; lean_object* l_Lean_mkIdentFrom(lean_object*, lean_object*); static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__10; static lean_object* l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__18; LEAN_EXPORT lean_object* l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___spec__1___boxed(lean_object*, lean_object*); +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__35; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__20; lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(lean_object*, lean_object*); static lean_object* l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__15; -static lean_object* l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__31; static lean_object* l___aux__Init__NotationExtra______macroRules__term_u03a3___x2c____1___closed__1; +static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__10; LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__11; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__19; lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__11; LEAN_EXPORT lean_object* l_Lean_unifConstraintElem; LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__36; +static lean_object* l_unexpandExists___closed__6; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__1; static lean_object* l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__2; lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_expandExplicitBindersAux_loop(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__32; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__10; LEAN_EXPORT lean_object* l_unexpandListToArray(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__26; static lean_object* l_tacticFunext_______closed__3; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__12; static lean_object* l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__9; @@ -196,99 +193,100 @@ static lean_object* l_Lean_doElemWhile__Do_____closed__2; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__21; static lean_object* l_Lean_explicitBinders___closed__1; static lean_object* l___aux__Init__NotationExtra______macroRules__term_u03a3_x27___x2c____1___closed__1; -static lean_object* l_Lean_binderIdent___closed__7; static lean_object* l_tactic_xb7_x2e_____x3b_____closed__16; static lean_object* l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__26; LEAN_EXPORT lean_object* l_Lean_expandExplicitBindersAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__24; +static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__20; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__39; static lean_object* l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__9; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__23; static lean_object* l_term_u03a3___x2c_____closed__3; lean_object* lean_string_utf8_byte_size(lean_object*); static lean_object* l_Lean_bracketedExplicitBinders___closed__5; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__6; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__7; static lean_object* l_Lean_unbracketedExplicitBinders___closed__7; static lean_object* l_unexpandExists___closed__5; -LEAN_EXPORT lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__19; LEAN_EXPORT lean_object* l_calc; static lean_object* l_tacticFunext_______closed__5; -LEAN_EXPORT lean_object* l___private_Init_NotationExtra_0__Lean_mkHintBody(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_unexpandListToArray___closed__2; static lean_object* l___aux__Init__NotationExtra______macroRules__solve__1___closed__2; +LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__1(lean_object*, lean_object*); static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile__Do____1___closed__9; static lean_object* l_Lean_unifConstraint___closed__6; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; static lean_object* l_Lean_doElemRepeat__Until_____closed__2; static lean_object* l___aux__Init__NotationExtra______macroRules__term_u03a3_x27___x2c____1___closed__2; static lean_object* l_Lean_doElemRepeat_____closed__1; uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l_tactic_xb7_x2e_____x3b_____closed__7; -extern lean_object* l_Lean_nameLitKind; static lean_object* l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__8; static lean_object* l_term_u03a3_x27___x2c_____closed__8; LEAN_EXPORT lean_object* l_unexpandListNil___boxed(lean_object*); static lean_object* l_Lean_unbracketedExplicitBinders___closed__3; LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_expandExplicitBindersAux_loop___closed__13; static lean_object* l_termExists___x2c_____closed__7; static lean_object* l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__17; static lean_object* l_Lean_doElemRepeat__Until_____closed__4; +static lean_object* l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__32; +static lean_object* l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__10; static lean_object* l_term___xd7_x27_____closed__1; +static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__23; lean_object* lean_nat_add(lean_object*, lean_object*); -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__21; LEAN_EXPORT lean_object* l_Lean_Loop_noConfusion(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__10; static lean_object* l_calcStep___closed__6; LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__solve__1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__12; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___spec__3(size_t, size_t, lean_object*); +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__25; static lean_object* l_tactic_xb7_x2e_____x3b_____closed__1; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__20; +static lean_object* l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__20; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile__Do____1___closed__3; static lean_object* l_solve___closed__11; +static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__13; static lean_object* l_calc___closed__11; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__16; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__20; +static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at___aux__Init__NotationExtra______macroRules__solve__1___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_explicitBinders___closed__6; lean_object* l_Array_zip___rarg(lean_object*, lean_object*); static lean_object* l___aux__Init__NotationExtra______macroRules__solve__1___closed__7; LEAN_EXPORT lean_object* l_Lean_expandBrackedBindersAux_loop(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_term_u03a3___x2c_____closed__7; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat__Until____1___closed__2; static lean_object* l_tactic_xb7_x2e_____x3b_____closed__17; -static lean_object* l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__14; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__43; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__31; LEAN_EXPORT lean_object* l_Lean_expandBrackedBindersAux_loop___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__7; LEAN_EXPORT lean_object* l_Lean_expandBrackedBindersAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_unexpandSorryAx___closed__1; static lean_object* l_term___xd7_x27_____closed__2; static lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__8; -static lean_object* l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__22; static lean_object* l_tactic_xb7_x2e_____x3b_____closed__8; static lean_object* l_Lean_expandExplicitBindersAux_loop___closed__10; static lean_object* l_solve___closed__2; static lean_object* l_Lean_doElemRepeat__Until_____closed__3; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__2; +static lean_object* l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__24; static lean_object* l_unexpandEqNDRec___closed__3; static lean_object* l_calc___closed__4; -LEAN_EXPORT lean_object* l_Lean_binderIdent; +extern lean_object* l_Lean_binderIdent; static lean_object* l_unexpandProdMk___closed__1; LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop(lean_object*, lean_object*); static lean_object* l_Lean_doElemRepeat__Until_____closed__1; -lean_object* l_Lean_Syntax_SepArray_getElems___rarg(lean_object*); static lean_object* l_Lean_term__Matches___x7c___closed__4; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(size_t, size_t, lean_object*); -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__31; static lean_object* l_tacticCalc_____closed__5; static lean_object* l_Lean_doElemRepeat_____closed__2; uint8_t l_Lean_Name_hasMacroScopes(lean_object*); LEAN_EXPORT lean_object* l_Lean_expandBrackedBinders(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__11; static lean_object* l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__11; LEAN_EXPORT lean_object* l_Lean_instForInLoopUnit___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__1(lean_object*, lean_object*); static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__20; static lean_object* l_term___xd7_x27_____closed__4; -static lean_object* l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__25; static lean_object* l_solve___closed__1; -static lean_object* l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__9; static lean_object* l___aux__Init__NotationExtra______macroRules__solve__1___closed__5; static lean_object* l_unexpandSorryAx___closed__2; static lean_object* l_tactic_xb7_x2e_____x3b_____closed__21; @@ -297,20 +295,22 @@ static lean_object* l_Lean_doElemRepeat__Until_____closed__5; static lean_object* l_Lean_doElemRepeat_____closed__4; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__10; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__11; +static lean_object* l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__6; static lean_object* l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__10; static lean_object* l_Lean_doElemRepeat_____closed__8; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_term__Matches___x7c; +static lean_object* l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__19; static lean_object* l_Lean_unbracketedExplicitBinders___closed__9; LEAN_EXPORT lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1(lean_object*, lean_object*, lean_object*); static lean_object* l_tacticFunext_______closed__6; -static lean_object* l_Array_forInUnsafe_loop___at___private_Init_NotationExtra_0__Lean_mkHintBody___spec__1___closed__2; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__solve__1___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_unexpandExists___closed__4; static lean_object* l_Lean_unifConstraintElem___closed__9; -static lean_object* l_Array_forInUnsafe_loop___at___private_Init_NotationExtra_0__Lean_mkHintBody___spec__1___closed__1; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__12; lean_object* l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(lean_object*, lean_object*); +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); static lean_object* l_Lean_doElemRepeat_____closed__9; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__26; static lean_object* l_solve___closed__8; static lean_object* l_tacticCalc_____closed__3; static lean_object* l_tactic_xb7_x2e_____x3b_____closed__2; @@ -320,29 +320,26 @@ static lean_object* l_Lean_unbracketedExplicitBinders___closed__11; static lean_object* l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__29; static lean_object* l_Lean_unifConstraintElem___closed__4; lean_object* lean_nat_sub(lean_object*, lean_object*); +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(lean_object*, lean_object*); static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat__Until____1___closed__1; LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__term_u2203___x2c____1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__29; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; +LEAN_EXPORT lean_object* l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2__; LEAN_EXPORT lean_object* l_unexpandIte(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_solve; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile__Do____1___closed__5; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4; static lean_object* l_termExists___x2c_____closed__4; static lean_object* l_calc___closed__9; static lean_object* l_term___xd7____1___closed__1; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__15; -static lean_object* l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__4; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__45; static lean_object* l_Lean_doElemRepeat__Until_____closed__7; static lean_object* l___aux__Init__NotationExtra______macroRules__solve__1___closed__1; LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___lambda__2(lean_object*); -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__27; static lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__1; +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_unifConstraintElem___closed__8; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__9; -static lean_object* l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__19; static lean_object* l_Lean_bracketedExplicitBinders___closed__6; static lean_object* l_tacticCalc_____closed__1; static lean_object* l_tacticCalc_____closed__2; @@ -352,44 +349,49 @@ static lean_object* l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__1 static lean_object* l_tactic_xb7_x2e_____x3b_____closed__23; LEAN_EXPORT lean_object* l_term___xd7____1; static lean_object* l_Lean_doElemWhile__Do_____closed__7; -static lean_object* l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__32; static lean_object* l_Lean_explicitBinders___closed__5; -static lean_object* l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__10; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__39; static lean_object* l_tactic_xb7_x2e_____x3b_____closed__13; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_expandExplicitBinders___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_termExists___x2c_____closed__8; -static lean_object* l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__20; static lean_object* l_Lean_doElemRepeat_____closed__6; static lean_object* l_Lean_expandExplicitBindersAux_loop___closed__7; static lean_object* l_Lean_unbracketedExplicitBinders___closed__8; static lean_object* l_Lean_unifConstraintElem___closed__1; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__12; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__31; static lean_object* l_term_u2203___x2c_____closed__4; static lean_object* l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__7; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__26; LEAN_EXPORT lean_object* l_Lean_expandExplicitBinders(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile__Do____1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__18; static lean_object* l_Lean_unifConstraintElem___closed__2; +static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__22; static lean_object* l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__24; static lean_object* l_term___xd7____1___closed__2; +LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4(lean_object*, lean_object*); +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__41; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__9; -static lean_object* l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__10; static lean_object* l___aux__Init__NotationExtra______macroRules__term_u03a3___x2c____1___closed__2; +static lean_object* l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__16; LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_tacticCalc_____closed__4; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; static lean_object* l_unexpandExists___closed__2; +static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__15; static lean_object* l_Lean_unifConstraint___closed__4; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__21; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__38; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__4; static lean_object* l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__16; static lean_object* l_Lean_unifConstraintElem___closed__11; lean_object* l_Lean_Syntax_getId(lean_object*); -lean_object* l___private_Init_Meta_0__Lean_quoteNameMk(lean_object*); +static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__14; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__34; LEAN_EXPORT lean_object* l_Lean_Loop_noConfusion___rarg(lean_object*); +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__11; static lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__3; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__24; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__23; static lean_object* l_unexpandProdMk___closed__2; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__3; static lean_object* l_termExists___x2c_____closed__6; @@ -399,276 +401,286 @@ static lean_object* l___aux__Init__NotationExtra______macroRules__command__Class static lean_object* l_solve___closed__13; LEAN_EXPORT lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__18; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__2; static lean_object* l_Lean_unbracketedExplicitBinders___closed__13; LEAN_EXPORT lean_object* l_Lean_bracketedExplicitBinders; static lean_object* l_tactic_xb7_x2e_____x3b_____closed__6; -static lean_object* l_tacticCalc_____closed__8; LEAN_EXPORT lean_object* l_unexpandEqRec(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__13; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__30; +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__17; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__5; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile__Do____1___closed__6; +LEAN_EXPORT lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___lambda__1(lean_object*, lean_object*); static lean_object* l_Lean_doElemWhile__Do_____closed__10; static lean_object* l___aux__Init__NotationExtra______macroRules__solve__1___closed__3; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__solve__1___spec__1___boxed(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___aux__Init__NotationExtra______macroRules__term_u03a3_x27___x2c____1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_unifConstraint___closed__10; -static lean_object* l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__3; -static lean_object* l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__16; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_bracketedExplicitBinders___closed__2; +static lean_object* l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__33; +static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__19; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__26; -static lean_object* l_Array_forInUnsafe_loop___at___private_Init_NotationExtra_0__Lean_mkHintBody___spec__1___closed__3; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__11; +static lean_object* l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__22; +static lean_object* l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__5___closed__1; static lean_object* l_calcStep___closed__9; +static lean_object* l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__30; LEAN_EXPORT lean_object* l_term___xd7_x27__; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__1; static lean_object* l_Lean_expandExplicitBindersAux_loop___closed__11; LEAN_EXPORT lean_object* l_unexpandEqNDRec(lean_object*, lean_object*, lean_object*); static lean_object* l_termExists___x2c_____closed__3; -static lean_object* l___private_Init_NotationExtra_0__Lean_mkHintBody___closed__2; +static lean_object* l_Lean_explicitBinders___closed__7; lean_object* l_Array_reverse___rarg(lean_object*); static lean_object* l_Lean_term__Matches___x7c___closed__8; LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__tacticCalc____1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_expandExplicitBinders___spec__1(lean_object*, lean_object*, size_t, size_t); extern lean_object* l_Lean_instInhabitedSyntax; +static lean_object* l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__6; +static lean_object* l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__3; static lean_object* l_Lean_unifConstraintElem___closed__5; LEAN_EXPORT lean_object* l_unexpandUnit(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___spec__4___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_tacticCalc_____closed__6; lean_object* l_Lean_mkSepArray(lean_object*, lean_object*); static lean_object* l_calcStep___closed__1; -static lean_object* l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__20; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__11; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__31; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__17; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__6; -static lean_object* l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__24; static lean_object* l___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___closed__2; -static lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__solve__1___spec__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2__; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__8; +static lean_object* l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__5___closed__3; static lean_object* l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__14; LEAN_EXPORT lean_object* l_Lean_Loop_forIn(lean_object*, lean_object*); static lean_object* l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__28; static lean_object* l_tacticFunext_______closed__8; static lean_object* l_term___xd7____1___closed__8; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__36; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__44; static lean_object* l_Lean_term__Matches___x7c___closed__6; -static lean_object* l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__4; size_t lean_usize_of_nat(lean_object*); static lean_object* l_Lean_unifConstraintElem___closed__6; static lean_object* l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__15; static lean_object* l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__18; static lean_object* l_term_u03a3___x2c_____closed__2; static lean_object* l_term___xd7____1___closed__4; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__19; lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); static lean_object* l_term_u2203___x2c_____closed__8; -static lean_object* l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__1; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__12; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__1; static lean_object* l_Lean_unbracketedExplicitBinders___closed__12; LEAN_EXPORT lean_object* l_tacticFunext____; static lean_object* l_term___xd7_x27_____closed__6; static lean_object* l_Lean_bracketedExplicitBinders___closed__1; static lean_object* l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__3; +static lean_object* l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__9; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__27; static lean_object* l_Lean_bracketedExplicitBinders___closed__8; static lean_object* l_calc___closed__8; static lean_object* l_unexpandListNil___rarg___closed__3; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile__Do____1___closed__10; LEAN_EXPORT lean_object* l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__4; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__13; static lean_object* l_calc___closed__2; static lean_object* l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__2; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__11; static lean_object* l_Lean_bracketedExplicitBinders___closed__9; static lean_object* l_tacticFunext_______closed__9; -static lean_object* l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__8; +LEAN_EXPORT lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_unexpandPSigma___closed__1; -static lean_object* l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__17; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__9; LEAN_EXPORT lean_object* l_unexpandPSigma(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__11; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__40; static lean_object* l_solve___closed__12; +static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__9; static lean_object* l_unexpandSigma___closed__1; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__17; +static lean_object* l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__5___closed__2; static lean_object* l_tactic_xb7_x2e_____x3b_____closed__11; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__16; -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__8; static lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__5; +LEAN_EXPORT lean_object* l_unexpandTSyntax(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_TSepArray_getElems___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat__Until____1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__35; static lean_object* l_Lean_bracketedExplicitBinders___closed__4; -static lean_object* l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__7; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__12; +static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__4; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__6; LEAN_EXPORT lean_object* l_unexpandUnit___rarg(lean_object*, lean_object*); -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__4; +static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__8; static lean_object* l_unexpandListNil___rarg___closed__1; -static lean_object* l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__30; +static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__17; static lean_object* l_calcStep___closed__3; lean_object* l_String_intercalate(lean_object*, lean_object*); static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__8; static lean_object* l_calc___closed__6; -LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Macro_throwError___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__23; -static lean_object* l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__18; LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__13; LEAN_EXPORT lean_object* l_unexpandListNil___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_unexpandTSyntaxArray(lean_object*, lean_object*, lean_object*); static lean_object* l_calcStep___closed__5; static lean_object* l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__22; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__1; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__34; static lean_object* l_tactic_xb7_x2e_____x3b_____closed__19; -static lean_object* l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__21; -static lean_object* l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__1; LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_solve___closed__7; -static lean_object* l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__5; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Init_NotationExtra_0__Lean_mkHintBody___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_getNumArgs(lean_object*); LEAN_EXPORT lean_object* l_unexpandSubtype(lean_object*, lean_object*, lean_object*); static lean_object* l_calcStep___closed__4; -static lean_object* l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__6; -static lean_object* l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__1; static lean_object* l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__10; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__2; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__37; static lean_object* l_unexpandListToArray___closed__3; static lean_object* l_Lean_doElemRepeat_____closed__7; -static lean_object* l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__33; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__27; static lean_object* l_Lean_expandExplicitBindersAux_loop___closed__6; -static lean_object* l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__22; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile__Do____1___closed__12; +static lean_object* l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__28; static lean_object* l_Lean_doElemWhile__Do_____closed__12; static lean_object* l_term_u2203___x2c_____closed__1; static lean_object* l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__6; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__9; -static lean_object* l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__14; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__3___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__21; uint8_t lean_nat_dec_le(lean_object*, lean_object*); -static lean_object* l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__28; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__40; +static lean_object* l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__5; static lean_object* l_tacticFunext_______closed__4; LEAN_EXPORT lean_object* l_Lean_Loop_noConfusion___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__18; static lean_object* l_term_u03a3_x27___x2c_____closed__7; -static lean_object* l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__8; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__21; LEAN_EXPORT lean_object* l_Lean_doElemWhile__Do__; static lean_object* l_term_u2203___x2c_____closed__2; lean_object* l_Lean_Syntax_getArgs(lean_object*); lean_object* l_Lean_Name_append(lean_object*, lean_object*); static lean_object* l_term_u03a3_x27___x2c_____closed__3; static lean_object* l_tactic_xb7_x2e_____x3b_____closed__9; +static lean_object* l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__25; static lean_object* l_Lean_unbracketedExplicitBinders___closed__10; +static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__2; +LEAN_EXPORT lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getKind(lean_object*); -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__8; lean_object* l_Lean_MacroScopesView_review(lean_object*); +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__42; +lean_object* l_Lean_quoteNameMk(lean_object*); LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_tactic_xb7_x2e_____x3b_____closed__4; static lean_object* l_tacticFunext_______closed__11; +LEAN_EXPORT lean_object* l_unexpandTSepArray(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__5; static lean_object* l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__25; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__17; LEAN_EXPORT lean_object* l_unexpandProdMk(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__29; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__30; LEAN_EXPORT lean_object* l_Lean_doElemRepeat__Until__; static lean_object* l_Lean_doElemRepeat_____closed__5; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__18; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__3; static lean_object* l_Lean_doElemRepeat_____closed__3; -static lean_object* l_Array_forInUnsafe_loop___at___private_Init_NotationExtra_0__Lean_mkHintBody___spec__1___closed__4; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__15; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__34; +static lean_object* l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__5; static lean_object* l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__16; -static lean_object* l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__15; LEAN_EXPORT lean_object* l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__3; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__1; +LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__1___boxed(lean_object*, lean_object*); +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__23; LEAN_EXPORT lean_object* l_Lean_doElemRepeat__; static lean_object* l_solve___closed__3; static lean_object* l_tacticFunext_______closed__10; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__3; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__13; static lean_object* l_Lean_term__Matches___x7c___closed__2; -static lean_object* l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__2; static lean_object* l_Lean_term__Matches___x7c___closed__1; lean_object* l_Array_ofSubarray___rarg(lean_object*); static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile__Do____1___closed__7; static lean_object* l_solve___closed__15; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile__Do____1___closed__11; static lean_object* l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__9; +lean_object* l_Lean_Syntax_mkNameLit(lean_object*, lean_object*); static lean_object* l_calc___closed__5; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__4; static lean_object* l_solve___closed__14; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__8; static lean_object* l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__6; -uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__2; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__12; LEAN_EXPORT lean_object* l_unexpandUnit___boxed(lean_object*); static lean_object* l_solve___closed__5; +static lean_object* l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__2; static lean_object* l_Lean_doElemWhile__Do_____closed__13; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__2; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__17; static lean_object* l_calc___closed__3; static lean_object* l_tactic_xb7_x2e_____x3b_____closed__10; LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_doElemWhile__Do_____closed__1; uint8_t l_Lean_Syntax_isNone(lean_object*); -static lean_object* l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__11; static lean_object* l_term___xd7_x27_____closed__7; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__solve__1___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_instForInLoopUnit(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_expandBrackedBindersAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__13; static lean_object* l_tacticCalc_____closed__7; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__20; static lean_object* l_term_u2203___x2c_____closed__5; -static lean_object* l_Lean_binderIdent___closed__8; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__22; static lean_object* l_term___xd7____1___closed__6; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__38; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__21; static lean_object* l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__4; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__10; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__32; LEAN_EXPORT lean_object* l_unexpandExists(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__23; +static lean_object* l___aux__Init__NotationExtra______macroRules__solve__1___closed__8; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__7; +static lean_object* l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__9; static lean_object* l_unexpandIte___closed__5; -static lean_object* l_Lean_binderIdent___closed__4; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Init_NotationExtra_0__Lean_mkHintBody___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); static lean_object* l_solve___closed__10; -static lean_object* l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__12; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__33; LEAN_EXPORT lean_object* l_unexpandListNil(lean_object*); static lean_object* l_term_u2203___x2c_____closed__3; static lean_object* l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__7; +static lean_object* l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__27; static lean_object* l_unexpandExists___closed__1; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__22; +static lean_object* l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__14; static lean_object* l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__14; uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); static lean_object* l_unexpandSubtype___closed__3; static lean_object* l___aux__Init__NotationExtra______macroRules__tacticCalc____1___closed__4; +static lean_object* l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__15; static lean_object* l_Lean_expandExplicitBindersAux_loop___closed__1; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__5; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__29; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__14; LEAN_EXPORT lean_object* l_termExists___x2c__; static lean_object* l_Lean_unifConstraintElem___closed__7; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__41; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile__Do____1___closed__8; LEAN_EXPORT lean_object* l_Lean_expandExplicitBindersAux_loop___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__21; LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1(lean_object*, lean_object*, lean_object*); static lean_object* l_term_u2203___x2c_____closed__6; +static lean_object* l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__3; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__2; static lean_object* l_Lean_term__Matches___x7c___closed__5; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__13; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__15; static lean_object* l_Lean_term__Matches___x7c___closed__7; static lean_object* l_Lean_expandExplicitBindersAux_loop___closed__9; static lean_object* l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__23; static lean_object* l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__5; LEAN_EXPORT lean_object* l_Lean_unbracketedExplicitBinders; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__6; static lean_object* l_term_u03a3___x2c_____closed__1; static lean_object* l_tactic_xb7_x2e_____x3b_____closed__20; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_tacticFunext_______closed__12; static lean_object* l_Lean_unbracketedExplicitBinders___closed__5; -static lean_object* l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__26; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__29; +static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__12; +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__7; static lean_object* l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__3; LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__term___xd7____1__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__25; LEAN_EXPORT lean_object* l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___spec__1(lean_object*, lean_object*); static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__27; static lean_object* l_Lean_expandExplicitBindersAux_loop___closed__8; static lean_object* l_term_u03a3_x27___x2c_____closed__5; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__18; -static lean_object* l_Lean_binderIdent___closed__2; +static lean_object* l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__1; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); -static lean_object* l_Lean_binderIdent___closed__1; LEAN_EXPORT lean_object* l_unexpandListCons(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_unifConstraint; LEAN_EXPORT lean_object* l_term_u2203___x2c__; @@ -676,82 +688,84 @@ static lean_object* l___aux__Init__NotationExtra______macroRules__term_x25_x5b__ static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__9; static lean_object* l_term_u03a3___x2c_____closed__8; static lean_object* l_solve___closed__4; -static lean_object* l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__6; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__33; static lean_object* l_Lean_expandExplicitBindersAux_loop___closed__2; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__3; static lean_object* l_term_u03a3___x2c_____closed__4; static lean_object* l___aux__Init__NotationExtra______macroRules__solve__1___closed__4; static lean_object* l_Lean_doElemWhile__Do_____closed__4; static lean_object* l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__1; -static lean_object* l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__18; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__25; static lean_object* l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__5; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__5; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; static lean_object* l_Lean_expandExplicitBindersAux_loop___closed__4; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__24; static lean_object* l_solve___closed__9; static lean_object* l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__30; -static lean_object* l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__5; static lean_object* l_Lean_doElemWhile__Do_____closed__8; LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext______1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__24; +static lean_object* l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__1; static lean_object* l_calc___closed__7; -static lean_object* l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__13; +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__19; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__7; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___spec__5(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__29; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___spec__3___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__7; -static lean_object* l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__21; -static lean_object* l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__15; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; +static lean_object* l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__4; static lean_object* l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__8; static lean_object* l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__17; static lean_object* l___aux__Init__NotationExtra______macroRules__tacticCalc____1___closed__3; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__42; static lean_object* l_unexpandSubtype___closed__5; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__2; -static lean_object* l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__7; -static lean_object* l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__7; -static lean_object* l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__2; LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___boxed(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___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___lambda__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_doElemWhile__Do_____closed__9; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__15; +static lean_object* l_unexpandExists___closed__7; static lean_object* l_Lean_unifConstraintElem___closed__10; static lean_object* l___aux__Init__NotationExtra______macroRules__tacticCalc____1___closed__5; +static lean_object* l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__21; static lean_object* l_term_u03a3_x27___x2c_____closed__6; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__18; +LEAN_EXPORT lean_object* l_Lean_termMacro_x2etrace_x5b___x5d__; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__4(size_t, size_t, lean_object*); static lean_object* l_calcStep___closed__12; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__33; +static lean_object* l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__8; static lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__4; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__29; static lean_object* l_term___xd7_x27_____closed__5; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat__Until____1___closed__3; +static lean_object* l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__4; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__22; static lean_object* l_term___xd7_x27_____closed__3; static lean_object* l_tactic_xb7_x2e_____x3b_____closed__12; static lean_object* l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__27; -static lean_object* l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__5; static lean_object* l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__8; static lean_object* l_Lean_doElemWhile__Do_____closed__11; LEAN_EXPORT lean_object* l_Lean_Loop_toCtorIdx(lean_object*); static lean_object* l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__1; -LEAN_EXPORT lean_object* l___private_Init_NotationExtra_0__Lean_mkHintBody___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__4; -static lean_object* l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__2; +static lean_object* l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__17; static lean_object* l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__12; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__16; static lean_object* l_unexpandIte___closed__3; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__14; +static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__11; static lean_object* l_term___xd7____1___closed__5; lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__8; static lean_object* l_calcStep___closed__8; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__34; -static lean_object* l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__23; LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__term___xd7_x27____1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_unifConstraint___closed__5; static lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__11; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__6; static lean_object* l___aux__Init__NotationExtra______macroRules__tacticCalc____1___closed__6; static lean_object* l___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___closed__1; -static lean_object* l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__3; -static lean_object* l___private_Init_NotationExtra_0__Lean_mkHintBody___closed__1; -static lean_object* l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__1; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__3; +static lean_object* l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__2; static lean_object* l_unexpandSubtype___closed__1; static lean_object* l_Lean_unifConstraintElem___closed__3; static lean_object* l_unexpandEqNDRec___closed__2; @@ -760,16 +774,14 @@ static lean_object* l_term___xd7____1___closed__3; static lean_object* l_Lean_unifConstraint___closed__7; static lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__9; static lean_object* l_calcStep___closed__2; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__23; -lean_object* l_Lean_Syntax_mkLit(lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_unexpandEqNDRec___closed__1; +static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__5; static lean_object* l___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___closed__3; uint8_t l_Lean_Syntax_isIdent(lean_object*); LEAN_EXPORT lean_object* l_Lean_explicitBinders; static lean_object* l_calc___closed__1; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__15; -static lean_object* _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__1() { +static lean_object* _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__1() { _start: { lean_object* x_1; @@ -777,35 +789,35 @@ x_1 = lean_mk_string_from_bytes("Lean", 4); return x_1; } } -static lean_object* _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__2() { +static lean_object* _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__1; +x_2 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__3() { +static lean_object* _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("termMacro.trace[__]_", 20); +x_1 = lean_mk_string_from_bytes("termMacro.trace[_]_", 19); return x_1; } } -static lean_object* _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__4() { +static lean_object* _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__2; -x_2 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__3; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__2; +x_2 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__5() { +static lean_object* _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__5() { _start: { lean_object* x_1; @@ -813,17 +825,17 @@ x_1 = lean_mk_string_from_bytes("andthen", 7); return x_1; } } -static lean_object* _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6() { +static lean_object* _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__5; +x_2 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__7() { +static lean_object* _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__7() { _start: { lean_object* x_1; @@ -831,17 +843,17 @@ x_1 = lean_mk_string_from_bytes("Macro.trace[", 12); return x_1; } } -static lean_object* _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__8() { +static lean_object* _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__7; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__7; x_2 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__9() { +static lean_object* _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__9() { _start: { lean_object* x_1; @@ -849,33 +861,33 @@ x_1 = lean_mk_string_from_bytes("ident", 5); return x_1; } } -static lean_object* _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__10() { +static lean_object* _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__9; +x_2 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__9; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__11() { +static lean_object* _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__10; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__10; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__12() { +static lean_object* _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; -x_2 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__8; -x_3 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__11; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; +x_2 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__8; +x_3 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__11; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -883,7 +895,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__13() { +static lean_object* _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__13() { _start: { lean_object* x_1; @@ -891,23 +903,23 @@ x_1 = lean_mk_string_from_bytes("]", 1); return x_1; } } -static lean_object* _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__14() { +static lean_object* _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__14() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__13; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__13; x_2 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__15() { +static lean_object* _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; -x_2 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__12; -x_3 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__14; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; +x_2 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__12; +x_3 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__14; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -915,7 +927,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__16() { +static lean_object* _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__16() { _start: { lean_object* x_1; @@ -923,17 +935,17 @@ x_1 = lean_mk_string_from_bytes("interpolatedStr", 15); return x_1; } } -static lean_object* _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__17() { +static lean_object* _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__16; +x_2 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__16; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__18() { +static lean_object* _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__18() { _start: { lean_object* x_1; @@ -941,21 +953,21 @@ x_1 = lean_mk_string_from_bytes("term", 4); return x_1; } } -static lean_object* _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__19() { +static lean_object* _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__18; +x_2 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__18; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__20() { +static lean_object* _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__19; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__19; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -963,25 +975,25 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__21() { +static lean_object* _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__17; -x_2 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__20; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__17; +x_2 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__20; x_3 = lean_alloc_ctor(1, 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_Lean_termMacro_x2etrace_x5b_____x5d_____closed__22() { +static lean_object* _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; -x_2 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__15; -x_3 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__21; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; +x_2 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__15; +x_3 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__21; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -989,13 +1001,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__23() { +static lean_object* _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__4; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__4; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__22; +x_3 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__22; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1003,15 +1015,15 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_termMacro_x2etrace_x5b_____x5d__() { +static lean_object* _init_l_Lean_termMacro_x2etrace_x5b___x5d__() { _start: { lean_object* x_1; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__23; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__23; return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__1() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__1() { _start: { lean_object* x_1; @@ -1019,17 +1031,17 @@ x_1 = lean_mk_string_from_bytes("Parser", 6); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__2() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__2; -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__1; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__2; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__3() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__3() { _start: { lean_object* x_1; @@ -1037,17 +1049,17 @@ x_1 = lean_mk_string_from_bytes("Term", 4); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__2; -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__3; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__2; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__5() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__5() { _start: { lean_object* x_1; @@ -1055,17 +1067,17 @@ x_1 = lean_mk_string_from_bytes("app", 3); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__6() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4; -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__5; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__7() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__7() { _start: { lean_object* x_1; @@ -1073,22 +1085,22 @@ x_1 = lean_mk_string_from_bytes("Macro.trace", 11); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__8() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__7; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__7; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__9() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__7; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__7; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__8; +x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__8; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1096,7 +1108,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__10() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__10() { _start: { lean_object* x_1; @@ -1104,17 +1116,17 @@ x_1 = lean_mk_string_from_bytes("Macro", 5); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__11() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__10; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__10; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__12() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__12() { _start: { lean_object* x_1; @@ -1122,61 +1134,61 @@ x_1 = lean_mk_string_from_bytes("trace", 5); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__13() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__11; -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__12; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__11; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__12; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__14() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__2; -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__10; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__2; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__10; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__15() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__14; -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__12; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__14; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__12; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__16() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__15; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__15; 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___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__17() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__16; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__16; x_3 = lean_alloc_ctor(1, 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___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__18() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__18() { _start: { lean_object* x_1; @@ -1184,17 +1196,17 @@ x_1 = lean_mk_string_from_bytes("null", 4); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__18; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__18; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__20() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__20() { _start: { lean_object* x_1; @@ -1202,17 +1214,17 @@ x_1 = lean_mk_string_from_bytes("paren", 5); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__21() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4; -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__20; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__20; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__22() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__22() { _start: { lean_object* x_1; @@ -1220,7 +1232,7 @@ x_1 = lean_mk_string_from_bytes("(", 1); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__23() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__23() { _start: { lean_object* x_1; @@ -1228,17 +1240,17 @@ x_1 = lean_mk_string_from_bytes("termS!_", 7); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__24() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__23; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__23; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__25() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__25() { _start: { lean_object* x_1; @@ -1246,7 +1258,7 @@ x_1 = lean_mk_string_from_bytes("s!", 2); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26() { _start: { lean_object* x_1; lean_object* x_2; @@ -1255,7 +1267,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__27() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__27() { _start: { lean_object* x_1; lean_object* x_2; @@ -1264,13 +1276,13 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__27; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__27; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1278,7 +1290,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__29() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__29() { _start: { lean_object* x_1; @@ -1286,7 +1298,7 @@ x_1 = lean_mk_string_from_bytes(")", 1); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30() { _start: { lean_object* x_1; lean_object* x_2; @@ -1295,7 +1307,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__31() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__31() { _start: { lean_object* x_1; @@ -1303,17 +1315,17 @@ x_1 = lean_mk_string_from_bytes("quotedName", 10); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__32() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__32() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4; -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__31; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__31; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__33() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__33() { _start: { lean_object* x_1; @@ -1321,7 +1333,7 @@ x_1 = lean_mk_string_from_bytes(".", 1); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__34() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__34() { _start: { lean_object* x_1; @@ -1329,7 +1341,7 @@ x_1 = lean_mk_string_from_bytes("`", 1); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35() { _start: { lean_object* x_1; lean_object* x_2; @@ -1338,11 +1350,11 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; uint8_t x_5; -x_4 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__4; +x_4 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__4; lean_inc(x_1); x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); if (x_5 == 0) @@ -1376,11 +1388,11 @@ lean_inc(x_15); x_16 = lean_ctor_get(x_2, 1); lean_inc(x_16); lean_dec(x_2); -x_17 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__13; +x_17 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__13; x_18 = l_Lean_addMacroScope(x_16, x_17, x_15); x_19 = lean_box(0); -x_20 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__9; -x_21 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__17; +x_20 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__9; +x_21 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__17; lean_inc(x_14); x_22 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_22, 0, x_14); @@ -1392,42 +1404,42 @@ lean_dec(x_9); x_24 = lean_erase_macro_scopes(x_23); lean_inc(x_24); x_25 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_19, x_24); -x_26 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__22; +x_26 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__22; lean_inc(x_14); x_27 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_27, 0, x_14); lean_ctor_set(x_27, 1, x_26); -x_28 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__25; +x_28 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__25; lean_inc(x_14); x_29 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_29, 0, x_14); lean_ctor_set(x_29, 1, x_28); -x_30 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; +x_30 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; x_31 = lean_array_push(x_30, x_29); x_32 = lean_array_push(x_31, x_11); x_33 = lean_box(2); -x_34 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__24; +x_34 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__24; x_35 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_35, 0, x_33); lean_ctor_set(x_35, 1, x_34); lean_ctor_set(x_35, 2, x_32); x_36 = lean_array_push(x_30, x_35); -x_37 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; +x_37 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; x_38 = lean_array_push(x_36, x_37); -x_39 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; +x_39 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; x_40 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_40, 0, x_33); lean_ctor_set(x_40, 1, x_39); lean_ctor_set(x_40, 2, x_38); -x_41 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__29; +x_41 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__29; x_42 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_42, 0, x_14); lean_ctor_set(x_42, 1, x_41); -x_43 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; +x_43 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; x_44 = lean_array_push(x_43, x_27); x_45 = lean_array_push(x_44, x_40); x_46 = lean_array_push(x_45, x_42); -x_47 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__21; +x_47 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__21; x_48 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_48, 0, x_33); lean_ctor_set(x_48, 1, x_47); @@ -1436,7 +1448,7 @@ x_49 = lean_array_push(x_30, x_22); if (lean_obj_tag(x_25) == 0) { lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_50 = l___private_Init_Meta_0__Lean_quoteNameMk(x_24); +x_50 = l_Lean_quoteNameMk(x_24); x_51 = lean_array_push(x_30, x_50); x_52 = lean_array_push(x_51, x_48); x_53 = lean_alloc_ctor(1, 3, 0); @@ -1444,7 +1456,7 @@ lean_ctor_set(x_53, 0, x_33); lean_ctor_set(x_53, 1, x_39); lean_ctor_set(x_53, 2, x_52); x_54 = lean_array_push(x_49, x_53); -x_55 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__6; +x_55 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__6; x_56 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_56, 0, x_33); lean_ctor_set(x_56, 1, x_55); @@ -1454,262 +1466,170 @@ return x_12; } else { -lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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_57; lean_object* x_58; lean_object* x_59; 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; 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_dec(x_24); x_57 = lean_ctor_get(x_25, 0); lean_inc(x_57); lean_dec(x_25); -x_58 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__33; +x_58 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__33; x_59 = l_String_intercalate(x_58, x_57); -x_60 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__34; +x_60 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__34; x_61 = lean_string_append(x_60, x_59); lean_dec(x_59); -x_62 = l_Lean_nameLitKind; -x_63 = l_Lean_Syntax_mkLit(x_62, x_61, x_33); -x_64 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_65 = lean_array_push(x_64, x_63); -x_66 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__32; -x_67 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_67, 0, x_33); -lean_ctor_set(x_67, 1, x_66); -lean_ctor_set(x_67, 2, x_65); -x_68 = lean_array_push(x_30, x_67); -x_69 = lean_array_push(x_68, x_48); -x_70 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_70, 0, x_33); -lean_ctor_set(x_70, 1, x_39); -lean_ctor_set(x_70, 2, x_69); -x_71 = lean_array_push(x_49, x_70); -x_72 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__6; -x_73 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_73, 0, x_33); -lean_ctor_set(x_73, 1, x_72); -lean_ctor_set(x_73, 2, x_71); -lean_ctor_set(x_12, 0, x_73); +x_62 = l_Lean_Syntax_mkNameLit(x_61, x_33); +x_63 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_64 = lean_array_push(x_63, x_62); +x_65 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__32; +x_66 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_66, 0, x_33); +lean_ctor_set(x_66, 1, x_65); +lean_ctor_set(x_66, 2, x_64); +x_67 = lean_array_push(x_30, x_66); +x_68 = lean_array_push(x_67, x_48); +x_69 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_69, 0, x_33); +lean_ctor_set(x_69, 1, x_39); +lean_ctor_set(x_69, 2, x_68); +x_70 = lean_array_push(x_49, x_69); +x_71 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__6; +x_72 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_72, 0, x_33); +lean_ctor_set(x_72, 1, x_71); +lean_ctor_set(x_72, 2, x_70); +lean_ctor_set(x_12, 0, x_72); return x_12; } } else { -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_74 = lean_ctor_get(x_12, 0); -x_75 = lean_ctor_get(x_12, 1); -lean_inc(x_75); +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; +x_73 = lean_ctor_get(x_12, 0); +x_74 = lean_ctor_get(x_12, 1); lean_inc(x_74); +lean_inc(x_73); lean_dec(x_12); -x_76 = lean_ctor_get(x_2, 2); +x_75 = lean_ctor_get(x_2, 2); +lean_inc(x_75); +x_76 = lean_ctor_get(x_2, 1); lean_inc(x_76); -x_77 = lean_ctor_get(x_2, 1); -lean_inc(x_77); lean_dec(x_2); -x_78 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__13; -x_79 = l_Lean_addMacroScope(x_77, x_78, x_76); -x_80 = lean_box(0); -x_81 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__9; -x_82 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__17; -lean_inc(x_74); -x_83 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_83, 0, x_74); -lean_ctor_set(x_83, 1, x_81); -lean_ctor_set(x_83, 2, x_79); -lean_ctor_set(x_83, 3, x_82); -x_84 = l_Lean_Syntax_getId(x_9); +x_77 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__13; +x_78 = l_Lean_addMacroScope(x_76, x_77, x_75); +x_79 = lean_box(0); +x_80 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__9; +x_81 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__17; +lean_inc(x_73); +x_82 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_82, 0, x_73); +lean_ctor_set(x_82, 1, x_80); +lean_ctor_set(x_82, 2, x_78); +lean_ctor_set(x_82, 3, x_81); +x_83 = l_Lean_Syntax_getId(x_9); lean_dec(x_9); -x_85 = lean_erase_macro_scopes(x_84); -lean_inc(x_85); -x_86 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_80, x_85); -x_87 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__22; -lean_inc(x_74); -x_88 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_88, 0, x_74); -lean_ctor_set(x_88, 1, x_87); -x_89 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__25; -lean_inc(x_74); -x_90 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_90, 0, x_74); -lean_ctor_set(x_90, 1, x_89); -x_91 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_92 = lean_array_push(x_91, x_90); -x_93 = lean_array_push(x_92, x_11); -x_94 = lean_box(2); -x_95 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__24; -x_96 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_96, 0, x_94); -lean_ctor_set(x_96, 1, x_95); -lean_ctor_set(x_96, 2, x_93); -x_97 = lean_array_push(x_91, x_96); -x_98 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; -x_99 = lean_array_push(x_97, x_98); -x_100 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_101 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_101, 0, x_94); -lean_ctor_set(x_101, 1, x_100); -lean_ctor_set(x_101, 2, x_99); -x_102 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__29; -x_103 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_103, 0, x_74); -lean_ctor_set(x_103, 1, x_102); -x_104 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; -x_105 = lean_array_push(x_104, x_88); -x_106 = lean_array_push(x_105, x_101); -x_107 = lean_array_push(x_106, x_103); -x_108 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__21; -x_109 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_109, 0, x_94); -lean_ctor_set(x_109, 1, x_108); -lean_ctor_set(x_109, 2, x_107); -x_110 = lean_array_push(x_91, x_83); -if (lean_obj_tag(x_86) == 0) -{ -lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; -x_111 = l___private_Init_Meta_0__Lean_quoteNameMk(x_85); -x_112 = lean_array_push(x_91, x_111); -x_113 = lean_array_push(x_112, x_109); -x_114 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_114, 0, x_94); -lean_ctor_set(x_114, 1, x_100); -lean_ctor_set(x_114, 2, x_113); -x_115 = lean_array_push(x_110, x_114); -x_116 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__6; -x_117 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_117, 0, x_94); -lean_ctor_set(x_117, 1, x_116); -lean_ctor_set(x_117, 2, x_115); -x_118 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_118, 0, x_117); -lean_ctor_set(x_118, 1, x_75); -return x_118; +x_84 = lean_erase_macro_scopes(x_83); +lean_inc(x_84); +x_85 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_79, x_84); +x_86 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__22; +lean_inc(x_73); +x_87 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_87, 0, x_73); +lean_ctor_set(x_87, 1, x_86); +x_88 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__25; +lean_inc(x_73); +x_89 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_89, 0, x_73); +lean_ctor_set(x_89, 1, x_88); +x_90 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_91 = lean_array_push(x_90, x_89); +x_92 = lean_array_push(x_91, x_11); +x_93 = lean_box(2); +x_94 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__24; +x_95 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_95, 0, x_93); +lean_ctor_set(x_95, 1, x_94); +lean_ctor_set(x_95, 2, x_92); +x_96 = lean_array_push(x_90, x_95); +x_97 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; +x_98 = lean_array_push(x_96, x_97); +x_99 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_100 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_100, 0, x_93); +lean_ctor_set(x_100, 1, x_99); +lean_ctor_set(x_100, 2, x_98); +x_101 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__29; +x_102 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_102, 0, x_73); +lean_ctor_set(x_102, 1, x_101); +x_103 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; +x_104 = lean_array_push(x_103, x_87); +x_105 = lean_array_push(x_104, x_100); +x_106 = lean_array_push(x_105, x_102); +x_107 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__21; +x_108 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_108, 0, x_93); +lean_ctor_set(x_108, 1, x_107); +lean_ctor_set(x_108, 2, x_106); +x_109 = lean_array_push(x_90, x_82); +if (lean_obj_tag(x_85) == 0) +{ +lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; +x_110 = l_Lean_quoteNameMk(x_84); +x_111 = lean_array_push(x_90, x_110); +x_112 = lean_array_push(x_111, x_108); +x_113 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_113, 0, x_93); +lean_ctor_set(x_113, 1, x_99); +lean_ctor_set(x_113, 2, x_112); +x_114 = lean_array_push(x_109, x_113); +x_115 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__6; +x_116 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_116, 0, x_93); +lean_ctor_set(x_116, 1, x_115); +lean_ctor_set(x_116, 2, x_114); +x_117 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_117, 0, x_116); +lean_ctor_set(x_117, 1, x_74); +return x_117; } else { -lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; +lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; +lean_dec(x_84); +x_118 = lean_ctor_get(x_85, 0); +lean_inc(x_118); lean_dec(x_85); -x_119 = lean_ctor_get(x_86, 0); -lean_inc(x_119); -lean_dec(x_86); -x_120 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__33; -x_121 = l_String_intercalate(x_120, x_119); -x_122 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__34; -x_123 = lean_string_append(x_122, x_121); -lean_dec(x_121); -x_124 = l_Lean_nameLitKind; -x_125 = l_Lean_Syntax_mkLit(x_124, x_123, x_94); -x_126 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_127 = lean_array_push(x_126, x_125); -x_128 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__32; -x_129 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_129, 0, x_94); -lean_ctor_set(x_129, 1, x_128); -lean_ctor_set(x_129, 2, x_127); -x_130 = lean_array_push(x_91, x_129); -x_131 = lean_array_push(x_130, x_109); -x_132 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_132, 0, x_94); -lean_ctor_set(x_132, 1, x_100); -lean_ctor_set(x_132, 2, x_131); -x_133 = lean_array_push(x_110, x_132); -x_134 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__6; -x_135 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_135, 0, x_94); -lean_ctor_set(x_135, 1, x_134); -lean_ctor_set(x_135, 2, x_133); -x_136 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_136, 0, x_135); -lean_ctor_set(x_136, 1, x_75); -return x_136; -} -} -} -} -} -static lean_object* _init_l_Lean_binderIdent___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("binderIdent", 11); -return x_1; -} -} -static lean_object* _init_l_Lean_binderIdent___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__2; -x_2 = l_Lean_binderIdent___closed__1; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_binderIdent___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("orelse", 6); -return x_1; -} -} -static lean_object* _init_l_Lean_binderIdent___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_binderIdent___closed__3; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_binderIdent___closed__5() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("_", 1); -return x_1; -} -} -static lean_object* _init_l_Lean_binderIdent___closed__6() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_binderIdent___closed__5; -x_2 = lean_alloc_ctor(5, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_binderIdent___closed__7() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_binderIdent___closed__4; -x_2 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__11; -x_3 = l_Lean_binderIdent___closed__6; -x_4 = lean_alloc_ctor(2, 3, 0); -lean_ctor_set(x_4, 0, x_1); -lean_ctor_set(x_4, 1, x_2); -lean_ctor_set(x_4, 2, x_3); -return x_4; -} +x_119 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__33; +x_120 = l_String_intercalate(x_119, x_118); +x_121 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__34; +x_122 = lean_string_append(x_121, x_120); +lean_dec(x_120); +x_123 = l_Lean_Syntax_mkNameLit(x_122, x_93); +x_124 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_125 = lean_array_push(x_124, x_123); +x_126 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__32; +x_127 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_127, 0, x_93); +lean_ctor_set(x_127, 1, x_126); +lean_ctor_set(x_127, 2, x_125); +x_128 = lean_array_push(x_90, x_127); +x_129 = lean_array_push(x_128, x_108); +x_130 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_130, 0, x_93); +lean_ctor_set(x_130, 1, x_99); +lean_ctor_set(x_130, 2, x_129); +x_131 = lean_array_push(x_109, x_130); +x_132 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__6; +x_133 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_133, 0, x_93); +lean_ctor_set(x_133, 1, x_132); +lean_ctor_set(x_133, 2, x_131); +x_134 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_134, 0, x_133); +lean_ctor_set(x_134, 1, x_74); +return x_134; } -static lean_object* _init_l_Lean_binderIdent___closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_binderIdent___closed__1; -x_2 = l_Lean_binderIdent___closed__2; -x_3 = l_Lean_binderIdent___closed__7; -x_4 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_4, 0, x_1); -lean_ctor_set(x_4, 1, x_2); -lean_ctor_set(x_4, 2, x_3); -return x_4; } } -static lean_object* _init_l_Lean_binderIdent() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_binderIdent___closed__8; -return x_1; } } static lean_object* _init_l_Lean_unbracketedExplicitBinders___closed__1() { @@ -1724,7 +1644,7 @@ static lean_object* _init_l_Lean_unbracketedExplicitBinders___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__2; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__2; x_2 = l_Lean_unbracketedExplicitBinders___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -1800,9 +1720,9 @@ static lean_object* _init_l_Lean_unbracketedExplicitBinders___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_Lean_unbracketedExplicitBinders___closed__9; -x_3 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__20; +x_3 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__20; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1826,7 +1746,7 @@ static lean_object* _init_l_Lean_unbracketedExplicitBinders___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_Lean_unbracketedExplicitBinders___closed__5; x_3 = l_Lean_unbracketedExplicitBinders___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); @@ -1870,7 +1790,7 @@ static lean_object* _init_l_Lean_bracketedExplicitBinders___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__2; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__2; x_2 = l_Lean_bracketedExplicitBinders___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -1880,7 +1800,7 @@ static lean_object* _init_l_Lean_bracketedExplicitBinders___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__22; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__22; x_2 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; @@ -1890,7 +1810,7 @@ static lean_object* _init_l_Lean_bracketedExplicitBinders___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_Lean_bracketedExplicitBinders___closed__3; x_3 = l_Lean_unbracketedExplicitBinders___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -1904,7 +1824,7 @@ static lean_object* _init_l_Lean_bracketedExplicitBinders___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_Lean_bracketedExplicitBinders___closed__4; x_3 = l_Lean_unbracketedExplicitBinders___closed__9; x_4 = lean_alloc_ctor(2, 3, 0); @@ -1918,9 +1838,9 @@ static lean_object* _init_l_Lean_bracketedExplicitBinders___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_Lean_bracketedExplicitBinders___closed__5; -x_3 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__20; +x_3 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__20; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1932,7 +1852,7 @@ static lean_object* _init_l_Lean_bracketedExplicitBinders___closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__29; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__29; x_2 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; @@ -1942,7 +1862,7 @@ static lean_object* _init_l_Lean_bracketedExplicitBinders___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_Lean_bracketedExplicitBinders___closed__6; x_3 = l_Lean_bracketedExplicitBinders___closed__7; x_4 = lean_alloc_ctor(2, 3, 0); @@ -1986,7 +1906,7 @@ static lean_object* _init_l_Lean_explicitBinders___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__2; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__2; x_2 = l_Lean_explicitBinders___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -1995,6 +1915,24 @@ return x_3; static lean_object* _init_l_Lean_explicitBinders___closed__3() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("orelse", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_explicitBinders___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_explicitBinders___closed__3; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_explicitBinders___closed__5() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_unbracketedExplicitBinders___closed__4; x_2 = l_Lean_bracketedExplicitBinders; @@ -2004,12 +1942,12 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_explicitBinders___closed__4() { +static lean_object* _init_l_Lean_explicitBinders___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_binderIdent___closed__4; -x_2 = l_Lean_explicitBinders___closed__3; +x_1 = l_Lean_explicitBinders___closed__4; +x_2 = l_Lean_explicitBinders___closed__5; x_3 = l_Lean_unbracketedExplicitBinders; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -2018,13 +1956,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_explicitBinders___closed__5() { +static lean_object* _init_l_Lean_explicitBinders___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_explicitBinders___closed__1; x_2 = l_Lean_explicitBinders___closed__2; -x_3 = l_Lean_explicitBinders___closed__4; +x_3 = l_Lean_explicitBinders___closed__6; x_4 = lean_alloc_ctor(9, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2036,7 +1974,7 @@ static lean_object* _init_l_Lean_explicitBinders() { _start: { lean_object* x_1; -x_1 = l_Lean_explicitBinders___closed__5; +x_1 = l_Lean_explicitBinders___closed__7; return x_1; } } @@ -2052,7 +1990,7 @@ static lean_object* _init_l_Lean_expandExplicitBindersAux_loop___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4; x_2 = l_Lean_expandExplicitBindersAux_loop___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -2070,7 +2008,7 @@ static lean_object* _init_l_Lean_expandExplicitBindersAux_loop___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4; x_2 = l_Lean_expandExplicitBindersAux_loop___closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -2088,7 +2026,7 @@ static lean_object* _init_l_Lean_expandExplicitBindersAux_loop___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4; x_2 = l_Lean_expandExplicitBindersAux_loop___closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -2098,7 +2036,7 @@ static lean_object* _init_l_Lean_expandExplicitBindersAux_loop___closed__7() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("=>", 2); +x_1 = lean_mk_string_from_bytes("_", 1); return x_1; } } @@ -2106,21 +2044,29 @@ static lean_object* _init_l_Lean_expandExplicitBindersAux_loop___closed__8() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("simpleBinder", 12); +x_1 = lean_mk_string_from_bytes("=>", 2); return x_1; } } static lean_object* _init_l_Lean_expandExplicitBindersAux_loop___closed__9() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("simpleBinder", 12); +return x_1; +} +} +static lean_object* _init_l_Lean_expandExplicitBindersAux_loop___closed__10() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4; -x_2 = l_Lean_expandExplicitBindersAux_loop___closed__8; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4; +x_2 = l_Lean_expandExplicitBindersAux_loop___closed__9; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_expandExplicitBindersAux_loop___closed__10() { +static lean_object* _init_l_Lean_expandExplicitBindersAux_loop___closed__11() { _start: { lean_object* x_1; @@ -2128,17 +2074,17 @@ x_1 = lean_mk_string_from_bytes("typeSpec", 8); return x_1; } } -static lean_object* _init_l_Lean_expandExplicitBindersAux_loop___closed__11() { +static lean_object* _init_l_Lean_expandExplicitBindersAux_loop___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4; -x_2 = l_Lean_expandExplicitBindersAux_loop___closed__10; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4; +x_2 = l_Lean_expandExplicitBindersAux_loop___closed__11; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_expandExplicitBindersAux_loop___closed__12() { +static lean_object* _init_l_Lean_expandExplicitBindersAux_loop___closed__13() { _start: { lean_object* x_1; @@ -2181,12 +2127,12 @@ lean_inc(x_17); x_20 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_20, 0, x_17); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_binderIdent___closed__5; +x_21 = l_Lean_expandExplicitBindersAux_loop___closed__7; lean_inc(x_17); x_22 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_22, 0, x_17); lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; +x_23 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; x_24 = lean_array_push(x_23, x_22); x_25 = lean_box(2); x_26 = l_Lean_expandExplicitBindersAux_loop___closed__6; @@ -2195,16 +2141,16 @@ lean_ctor_set(x_27, 0, x_25); lean_ctor_set(x_27, 1, x_26); lean_ctor_set(x_27, 2, x_24); x_28 = lean_array_push(x_23, x_27); -x_29 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; +x_29 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; x_30 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_30, 0, x_25); lean_ctor_set(x_30, 1, x_29); lean_ctor_set(x_30, 2, x_28); -x_31 = l_Lean_expandExplicitBindersAux_loop___closed__7; +x_31 = l_Lean_expandExplicitBindersAux_loop___closed__8; x_32 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_32, 0, x_17); lean_ctor_set(x_32, 1, x_31); -x_33 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; +x_33 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; x_34 = lean_array_push(x_33, x_30); x_35 = lean_array_push(x_34, x_32); x_36 = lean_array_push(x_35, x_5); @@ -2213,7 +2159,7 @@ x_38 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_38, 0, x_25); lean_ctor_set(x_38, 1, x_37); lean_ctor_set(x_38, 2, x_36); -x_39 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; +x_39 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; x_40 = lean_array_push(x_39, x_20); x_41 = lean_array_push(x_40, x_38); x_42 = l_Lean_expandExplicitBindersAux_loop___closed__2; @@ -2229,7 +2175,7 @@ lean_ctor_set(x_45, 2, x_44); lean_inc(x_1); x_46 = lean_array_push(x_39, x_1); x_47 = lean_array_push(x_46, x_45); -x_48 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__6; +x_48 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__6; x_49 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_49, 0, x_25); lean_ctor_set(x_49, 1, x_48); @@ -2256,12 +2202,12 @@ lean_inc(x_53); x_56 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_56, 0, x_53); lean_ctor_set(x_56, 1, x_55); -x_57 = l_Lean_binderIdent___closed__5; +x_57 = l_Lean_expandExplicitBindersAux_loop___closed__7; lean_inc(x_53); x_58 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_58, 0, x_53); lean_ctor_set(x_58, 1, x_57); -x_59 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; +x_59 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; x_60 = lean_array_push(x_59, x_58); x_61 = lean_box(2); x_62 = l_Lean_expandExplicitBindersAux_loop___closed__6; @@ -2270,20 +2216,20 @@ lean_ctor_set(x_63, 0, x_61); lean_ctor_set(x_63, 1, x_62); lean_ctor_set(x_63, 2, x_60); x_64 = lean_array_push(x_59, x_63); -x_65 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; +x_65 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; x_66 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_66, 0, x_61); lean_ctor_set(x_66, 1, x_65); lean_ctor_set(x_66, 2, x_64); -x_67 = l_Lean_expandExplicitBindersAux_loop___closed__12; +x_67 = l_Lean_expandExplicitBindersAux_loop___closed__13; lean_inc(x_53); x_68 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_68, 0, x_53); lean_ctor_set(x_68, 1, x_67); -x_69 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; +x_69 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; x_70 = lean_array_push(x_69, x_68); x_71 = lean_array_push(x_70, x_51); -x_72 = l_Lean_expandExplicitBindersAux_loop___closed__11; +x_72 = l_Lean_expandExplicitBindersAux_loop___closed__12; x_73 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_73, 0, x_61); lean_ctor_set(x_73, 1, x_72); @@ -2295,7 +2241,7 @@ lean_ctor_set(x_75, 1, x_65); lean_ctor_set(x_75, 2, x_74); x_76 = lean_array_push(x_69, x_66); x_77 = lean_array_push(x_76, x_75); -x_78 = l_Lean_expandExplicitBindersAux_loop___closed__9; +x_78 = l_Lean_expandExplicitBindersAux_loop___closed__10; x_79 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_79, 0, x_61); lean_ctor_set(x_79, 1, x_78); @@ -2305,11 +2251,11 @@ x_81 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_81, 0, x_61); lean_ctor_set(x_81, 1, x_65); lean_ctor_set(x_81, 2, x_80); -x_82 = l_Lean_expandExplicitBindersAux_loop___closed__7; +x_82 = l_Lean_expandExplicitBindersAux_loop___closed__8; x_83 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_83, 0, x_53); lean_ctor_set(x_83, 1, x_82); -x_84 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; +x_84 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; x_85 = lean_array_push(x_84, x_81); x_86 = lean_array_push(x_85, x_83); x_87 = lean_array_push(x_86, x_5); @@ -2333,7 +2279,7 @@ lean_ctor_set(x_95, 2, x_94); lean_inc(x_1); x_96 = lean_array_push(x_69, x_1); x_97 = lean_array_push(x_96, x_95); -x_98 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__6; +x_98 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__6; x_99 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_99, 0, x_61); lean_ctor_set(x_99, 1, x_98); @@ -2361,19 +2307,19 @@ lean_inc(x_102); x_105 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_105, 0, x_102); lean_ctor_set(x_105, 1, x_104); -x_106 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; +x_106 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; x_107 = lean_array_push(x_106, x_14); x_108 = lean_box(2); -x_109 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; +x_109 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; x_110 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_110, 0, x_108); lean_ctor_set(x_110, 1, x_109); lean_ctor_set(x_110, 2, x_107); -x_111 = l_Lean_expandExplicitBindersAux_loop___closed__7; +x_111 = l_Lean_expandExplicitBindersAux_loop___closed__8; x_112 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_112, 0, x_102); lean_ctor_set(x_112, 1, x_111); -x_113 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; +x_113 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; x_114 = lean_array_push(x_113, x_110); x_115 = lean_array_push(x_114, x_112); x_116 = lean_array_push(x_115, x_5); @@ -2382,7 +2328,7 @@ x_118 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_118, 0, x_108); lean_ctor_set(x_118, 1, x_117); lean_ctor_set(x_118, 2, x_116); -x_119 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; +x_119 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; x_120 = lean_array_push(x_119, x_105); x_121 = lean_array_push(x_120, x_118); x_122 = l_Lean_expandExplicitBindersAux_loop___closed__2; @@ -2398,7 +2344,7 @@ lean_ctor_set(x_125, 2, x_124); lean_inc(x_1); x_126 = lean_array_push(x_119, x_1); x_127 = lean_array_push(x_126, x_125); -x_128 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__6; +x_128 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__6; x_129 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_129, 0, x_108); lean_ctor_set(x_129, 1, x_128); @@ -2425,23 +2371,23 @@ lean_inc(x_133); x_136 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_136, 0, x_133); lean_ctor_set(x_136, 1, x_135); -x_137 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; +x_137 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; x_138 = lean_array_push(x_137, x_14); x_139 = lean_box(2); -x_140 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; +x_140 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; x_141 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_141, 0, x_139); lean_ctor_set(x_141, 1, x_140); lean_ctor_set(x_141, 2, x_138); -x_142 = l_Lean_expandExplicitBindersAux_loop___closed__12; +x_142 = l_Lean_expandExplicitBindersAux_loop___closed__13; lean_inc(x_133); x_143 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_143, 0, x_133); lean_ctor_set(x_143, 1, x_142); -x_144 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; +x_144 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; x_145 = lean_array_push(x_144, x_143); x_146 = lean_array_push(x_145, x_131); -x_147 = l_Lean_expandExplicitBindersAux_loop___closed__11; +x_147 = l_Lean_expandExplicitBindersAux_loop___closed__12; x_148 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_148, 0, x_139); lean_ctor_set(x_148, 1, x_147); @@ -2453,7 +2399,7 @@ lean_ctor_set(x_150, 1, x_140); lean_ctor_set(x_150, 2, x_149); x_151 = lean_array_push(x_144, x_141); x_152 = lean_array_push(x_151, x_150); -x_153 = l_Lean_expandExplicitBindersAux_loop___closed__9; +x_153 = l_Lean_expandExplicitBindersAux_loop___closed__10; x_154 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_154, 0, x_139); lean_ctor_set(x_154, 1, x_153); @@ -2463,11 +2409,11 @@ x_156 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_156, 0, x_139); lean_ctor_set(x_156, 1, x_140); lean_ctor_set(x_156, 2, x_155); -x_157 = l_Lean_expandExplicitBindersAux_loop___closed__7; +x_157 = l_Lean_expandExplicitBindersAux_loop___closed__8; x_158 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_158, 0, x_133); lean_ctor_set(x_158, 1, x_157); -x_159 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; +x_159 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; x_160 = lean_array_push(x_159, x_156); x_161 = lean_array_push(x_160, x_158); x_162 = lean_array_push(x_161, x_5); @@ -2491,7 +2437,7 @@ lean_ctor_set(x_170, 2, x_169); lean_inc(x_1); x_171 = lean_array_push(x_144, x_1); x_172 = lean_array_push(x_171, x_170); -x_173 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__6; +x_173 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__6; x_174 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_174, 0, x_139); lean_ctor_set(x_174, 1, x_173); @@ -2716,7 +2662,7 @@ size_t x_19; size_t x_20; lean_object* x_21; uint8_t x_22; x_19 = 0; x_20 = lean_usize_of_nat(x_14); lean_dec(x_14); -x_21 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__2; +x_21 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__2; x_22 = l_Array_anyMUnsafe_any___at_Lean_expandExplicitBinders___spec__1(x_21, x_13, x_19, x_20); if (x_22 == 0) { @@ -2802,7 +2748,7 @@ lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_obj x_6 = lean_ctor_get(x_4, 5); lean_inc(x_6); x_7 = l_Lean_mkIdentFrom(x_6, x_1); -x_8 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; +x_8 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; x_9 = lean_array_push(x_8, x_2); x_10 = l_Lean_expandBrackedBindersAux(x_7, x_9, x_3, x_4, x_5); lean_dec(x_9); @@ -2821,7 +2767,7 @@ static lean_object* _init_l_Lean_unifConstraint___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__2; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__2; x_2 = l_Lean_unifConstraint___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -2867,7 +2813,7 @@ static lean_object* _init_l_Lean_unifConstraint___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_binderIdent___closed__4; +x_1 = l_Lean_explicitBinders___closed__4; x_2 = l_Lean_unifConstraint___closed__4; x_3 = l_Lean_unifConstraint___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); @@ -2881,8 +2827,8 @@ static lean_object* _init_l_Lean_unifConstraint___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; -x_2 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__20; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; +x_2 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__20; x_3 = l_Lean_unifConstraint___closed__7; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -2895,9 +2841,9 @@ static lean_object* _init_l_Lean_unifConstraint___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_Lean_unifConstraint___closed__8; -x_3 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__20; +x_3 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__20; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2939,7 +2885,7 @@ static lean_object* _init_l_Lean_unifConstraintElem___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__2; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__2; x_2 = l_Lean_unifConstraintElem___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -2977,7 +2923,7 @@ static lean_object* _init_l_Lean_unifConstraintElem___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_Lean_unifConstraintElem___closed__5; x_3 = l_Lean_unifConstraint; x_4 = lean_alloc_ctor(2, 3, 0); @@ -3021,7 +2967,7 @@ static lean_object* _init_l_Lean_unifConstraintElem___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_Lean_unifConstraintElem___closed__6; x_3 = l_Lean_unifConstraintElem___closed__9; x_4 = lean_alloc_ctor(2, 3, 0); @@ -3053,25 +2999,25 @@ x_1 = l_Lean_unifConstraintElem___closed__11; return x_1; } } -static lean_object* _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__1() { +static lean_object* _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("command_Unif_hint___Where_|-⊢_", 32); +x_1 = lean_mk_string_from_bytes("command_Unif_hint__Where_|-⊢_", 31); return x_1; } } -static lean_object* _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__2() { +static lean_object* _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__2; -x_2 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__1; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__2; +x_2 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__3() { +static lean_object* _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__3() { _start: { lean_object* x_1; @@ -3079,27 +3025,27 @@ x_1 = lean_mk_string_from_bytes("attrKind", 8); return x_1; } } -static lean_object* _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__4() { +static lean_object* _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__3; +x_2 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__5() { +static lean_object* _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__4; +x_1 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__4; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__6() { +static lean_object* _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__6() { _start: { lean_object* x_1; @@ -3107,23 +3053,23 @@ x_1 = lean_mk_string_from_bytes("unif_hint ", 10); return x_1; } } -static lean_object* _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__7() { +static lean_object* _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__6; +x_1 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__6; x_2 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__8() { +static lean_object* _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; -x_2 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__5; -x_3 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__7; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; +x_2 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__5; +x_3 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__7; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3131,25 +3077,25 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__9() { +static lean_object* _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_unbracketedExplicitBinders___closed__7; -x_2 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__11; +x_2 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__11; x_3 = lean_alloc_ctor(1, 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_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__10() { +static lean_object* _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; -x_2 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__8; -x_3 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__9; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; +x_2 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__8; +x_3 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__9; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3157,7 +3103,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__11() { +static lean_object* _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__11() { _start: { lean_object* x_1; @@ -3165,17 +3111,17 @@ x_1 = lean_mk_string_from_bytes("many", 4); return x_1; } } -static lean_object* _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__12() { +static lean_object* _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__11; +x_2 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__11; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__13() { +static lean_object* _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__13() { _start: { lean_object* x_1; @@ -3183,45 +3129,45 @@ x_1 = lean_mk_string_from_bytes("bracketedBinder", 15); return x_1; } } -static lean_object* _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__14() { +static lean_object* _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__13; +x_2 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__13; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__15() { +static lean_object* _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__15() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__14; +x_1 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__14; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__16() { +static lean_object* _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__12; -x_2 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__15; +x_1 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__12; +x_2 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__15; x_3 = lean_alloc_ctor(1, 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_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__17() { +static lean_object* _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; -x_2 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__10; -x_3 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__16; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; +x_2 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__10; +x_3 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__16; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3229,7 +3175,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__18() { +static lean_object* _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__18() { _start: { lean_object* x_1; @@ -3237,23 +3183,23 @@ x_1 = lean_mk_string_from_bytes(" where ", 7); return x_1; } } -static lean_object* _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__19() { +static lean_object* _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__19() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__18; +x_1 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__18; x_2 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__20() { +static lean_object* _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; -x_2 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__17; -x_3 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__19; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; +x_2 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__17; +x_3 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__19; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3261,7 +3207,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__21() { +static lean_object* _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__21() { _start: { lean_object* x_1; @@ -3269,21 +3215,21 @@ x_1 = lean_mk_string_from_bytes("withPosition", 12); return x_1; } } -static lean_object* _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__22() { +static lean_object* _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__21; +x_2 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__21; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__23() { +static lean_object* _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__12; +x_1 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__12; x_2 = l_Lean_unifConstraintElem; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -3291,25 +3237,25 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__24() { +static lean_object* _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__22; -x_2 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__23; +x_1 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__22; +x_2 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__23; x_3 = lean_alloc_ctor(1, 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_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__25() { +static lean_object* _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; -x_2 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__20; -x_3 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__24; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; +x_2 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__20; +x_3 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__24; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3317,7 +3263,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__26() { +static lean_object* _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__26() { _start: { lean_object* x_1; @@ -3325,17 +3271,17 @@ x_1 = lean_mk_string_from_bytes("|-", 2); return x_1; } } -static lean_object* _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__27() { +static lean_object* _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__27() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__26; +x_1 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__26; x_2 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__28() { +static lean_object* _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__28() { _start: { lean_object* x_1; @@ -3343,23 +3289,23 @@ x_1 = lean_mk_string_from_bytes("⊢ ", 4); return x_1; } } -static lean_object* _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__29() { +static lean_object* _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__29() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__28; +x_1 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__28; x_2 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__30() { +static lean_object* _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__30() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_binderIdent___closed__4; -x_2 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__27; -x_3 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__29; +x_1 = l_Lean_explicitBinders___closed__4; +x_2 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__27; +x_3 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__29; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3367,13 +3313,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__31() { +static lean_object* _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__31() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; -x_2 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__25; -x_3 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__30; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; +x_2 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__25; +x_3 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__30; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3381,12 +3327,12 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__32() { +static lean_object* _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__32() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; -x_2 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__31; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; +x_2 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__31; x_3 = l_Lean_unifConstraint; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -3395,13 +3341,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__33() { +static lean_object* _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__33() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__2; +x_1 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__32; +x_3 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__32; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3409,226 +3355,395 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2__() { +static lean_object* _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2__() { _start: { lean_object* x_1; -x_1 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__33; +x_1 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__33; return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Init_NotationExtra_0__Lean_mkHintBody___spec__1___closed__1() { +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("arrow", 5); -return x_1; +lean_object* x_6; uint8_t x_7; +x_6 = lean_array_get_size(x_1); +x_7 = lean_nat_dec_lt(x_4, x_6); +lean_dec(x_6); +if (x_7 == 0) +{ +lean_object* x_8; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_5); +return x_8; } +else +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_unsigned_to_nat(0u); +x_10 = lean_nat_dec_eq(x_3, x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_sub(x_3, x_11); +lean_dec(x_3); +x_13 = lean_array_fget(x_1, x_4); +lean_inc(x_2); +x_14 = lean_apply_1(x_2, x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; +lean_dec(x_12); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_15 = lean_box(0); +return x_15; } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Init_NotationExtra_0__Lean_mkHintBody___spec__1___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4; -x_2 = l_Array_forInUnsafe_loop___at___private_Init_NotationExtra_0__Lean_mkHintBody___spec__1___closed__1; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_14, 0); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_nat_add(x_4, x_11); +lean_dec(x_4); +x_18 = lean_array_push(x_5, x_16); +x_3 = x_12; +x_4 = x_17; +x_5 = x_18; +goto _start; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Init_NotationExtra_0__Lean_mkHintBody___spec__1___closed__3() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("=", 1); -return x_1; +lean_object* x_20; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_20 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_20, 0, x_5); +return x_20; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Init_NotationExtra_0__Lean_mkHintBody___spec__1___closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("→", 3); -return x_1; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Init_NotationExtra_0__Lean_mkHintBody___spec__1(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_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__1(lean_object* x_1, lean_object* x_2) { _start: { -uint8_t x_9; -x_9 = lean_usize_dec_lt(x_5, x_4); -if (x_9 == 0) +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_3 = lean_array_get_size(x_1); +x_4 = lean_mk_empty_array_with_capacity(x_3); +x_5 = lean_unsigned_to_nat(0u); +x_6 = l_Array_sequenceMap_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__2(x_1, x_2, x_3, x_5, x_4); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__3(size_t x_1, size_t x_2, lean_object* x_3) { +_start: { -lean_object* x_10; -lean_dec(x_7); -lean_dec(x_2); +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; lean_object* x_11; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = 1; +x_9 = lean_usize_add(x_2, x_8); +x_10 = lean_ctor_get(x_5, 1); +lean_inc(x_10); +lean_dec(x_5); +x_11 = lean_array_uset(x_7, x_2, x_10); +x_2 = x_9; +x_3 = x_11; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__4(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; lean_object* x_11; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = 1; +x_9 = lean_usize_add(x_2, x_8); +x_10 = lean_ctor_get(x_5, 0); +lean_inc(x_10); +lean_dec(x_5); +x_11 = lean_array_uset(x_7, x_2, x_10); +x_2 = x_9; +x_3 = x_11; +goto _start; +} +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__5___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("arrow", 5); +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__5___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("=", 1); +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__5___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("→", 3); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, size_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; +x_11 = lean_usize_dec_lt(x_7, x_6); +if (x_11 == 0) +{ +lean_object* x_12; +lean_dec(x_9); +lean_dec(x_4); +lean_dec(x_3); lean_dec(x_1); -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_6); -lean_ctor_set(x_10, 1, x_8); -return x_10; +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_8); +lean_ctor_set(x_12, 1, x_10); +return x_12; } else { -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; 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; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; size_t x_34; size_t x_35; -x_11 = lean_array_uget(x_3, x_5); -lean_inc(x_7); -x_12 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_7, x_8); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); +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; 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; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; size_t x_34; size_t x_35; +x_13 = lean_array_uget(x_5, x_7); +x_14 = lean_ctor_get(x_13, 0); lean_inc(x_14); -lean_dec(x_12); -x_15 = lean_unsigned_to_nat(0u); -x_16 = l_Lean_Syntax_getArg(x_11, x_15); -lean_dec(x_11); -x_17 = l_Lean_Syntax_getArg(x_16, x_15); -x_18 = l_Array_forInUnsafe_loop___at___private_Init_NotationExtra_0__Lean_mkHintBody___spec__1___closed__3; -lean_inc(x_13); -x_19 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_19, 0, x_13); -lean_ctor_set(x_19, 1, x_18); -x_20 = lean_unsigned_to_nat(2u); -x_21 = l_Lean_Syntax_getArg(x_16, x_20); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +lean_inc(x_9); +x_16 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_9, x_10); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); lean_dec(x_16); -lean_inc(x_2); -x_22 = lean_array_push(x_2, x_17); -x_23 = lean_array_push(x_22, x_19); -x_24 = lean_array_push(x_23, x_21); -x_25 = lean_box(2); +x_19 = l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__5___closed__1; lean_inc(x_1); -x_26 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_1); -lean_ctor_set(x_26, 2, x_24); -x_27 = l_Array_forInUnsafe_loop___at___private_Init_NotationExtra_0__Lean_mkHintBody___spec__1___closed__4; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_13); -lean_ctor_set(x_28, 1, x_27); -lean_inc(x_2); -x_29 = lean_array_push(x_2, x_26); -x_30 = lean_array_push(x_29, x_28); -x_31 = lean_array_push(x_30, x_6); -x_32 = l_Array_forInUnsafe_loop___at___private_Init_NotationExtra_0__Lean_mkHintBody___spec__1___closed__2; +x_20 = lean_name_mk_string(x_1, x_19); +x_21 = l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__5___closed__2; +lean_inc(x_17); +x_22 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_22, 0, x_17); +lean_ctor_set(x_22, 1, x_21); +lean_inc(x_4); +x_23 = lean_array_push(x_4, x_14); +x_24 = lean_array_push(x_23, x_22); +x_25 = lean_array_push(x_24, x_15); +x_26 = lean_box(2); +lean_inc(x_3); +x_27 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_3); +lean_ctor_set(x_27, 2, x_25); +x_28 = l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__5___closed__3; +x_29 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_29, 0, x_17); +lean_ctor_set(x_29, 1, x_28); +lean_inc(x_4); +x_30 = lean_array_push(x_4, x_27); +x_31 = lean_array_push(x_30, x_29); +x_32 = lean_array_push(x_31, x_8); x_33 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_33, 0, x_25); -lean_ctor_set(x_33, 1, x_32); -lean_ctor_set(x_33, 2, x_31); +lean_ctor_set(x_33, 0, x_26); +lean_ctor_set(x_33, 1, x_20); +lean_ctor_set(x_33, 2, x_32); x_34 = 1; -x_35 = lean_usize_add(x_5, x_34); -x_5 = x_35; -x_6 = x_33; -x_8 = x_14; +x_35 = lean_usize_add(x_7, x_34); +x_7 = x_35; +x_8 = x_33; +x_10 = x_18; goto _start; } } } -static lean_object* _init_l___private_Init_NotationExtra_0__Lean_mkHintBody___closed__1() { +LEAN_EXPORT lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___lambda__1(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("term_=_", 7); -return x_1; +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = l_Lean_unifConstraintElem___closed__1; +lean_inc(x_1); +x_4 = lean_name_mk_string(x_1, x_3); +lean_inc(x_2); +x_5 = l_Lean_Syntax_isOfKind(x_2, x_4); +lean_dec(x_4); +if (x_5 == 0) +{ +lean_object* x_6; +lean_dec(x_2); +lean_dec(x_1); +x_6 = lean_box(0); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_7 = lean_unsigned_to_nat(0u); +x_8 = l_Lean_Syntax_getArg(x_2, x_7); +x_9 = l_Lean_unifConstraint___closed__1; +x_10 = lean_name_mk_string(x_1, x_9); +lean_inc(x_8); +x_11 = l_Lean_Syntax_isOfKind(x_8, x_10); +lean_dec(x_10); +if (x_11 == 0) +{ +lean_object* x_12; +lean_dec(x_8); +lean_dec(x_2); +x_12 = lean_box(0); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_13 = l_Lean_Syntax_getArg(x_8, x_7); +x_14 = lean_unsigned_to_nat(2u); +x_15 = l_Lean_Syntax_getArg(x_8, x_14); +lean_dec(x_8); +x_16 = lean_unsigned_to_nat(1u); +x_17 = l_Lean_Syntax_getArg(x_2, x_16); +lean_dec(x_2); +x_18 = l_Lean_Syntax_matchesNull(x_17, x_7); +if (x_18 == 0) +{ +lean_object* x_19; +lean_dec(x_15); +lean_dec(x_13); +x_19 = lean_box(0); +return x_19; +} +else +{ +lean_object* x_20; lean_object* x_21; +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_13); +lean_ctor_set(x_20, 1, x_15); +x_21 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_21, 0, x_20); +return x_21; +} +} } } -static lean_object* _init_l___private_Init_NotationExtra_0__Lean_mkHintBody___closed__2() { +} +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Init_NotationExtra_0__Lean_mkHintBody___closed__1; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4; +x_2 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_NotationExtra_0__Lean_mkHintBody(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__2() { _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; 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; lean_object* x_20; lean_object* x_21; lean_object* x_22; size_t x_23; size_t x_24; lean_object* x_25; uint8_t x_26; -lean_inc(x_3); -x_5 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_3, x_4); -x_6 = lean_ctor_get(x_5, 0); -lean_inc(x_6); -x_7 = lean_ctor_get(x_5, 1); -lean_inc(x_7); -lean_dec(x_5); -x_8 = lean_unsigned_to_nat(0u); -x_9 = l_Lean_Syntax_getArg(x_2, x_8); -x_10 = l_Array_forInUnsafe_loop___at___private_Init_NotationExtra_0__Lean_mkHintBody___spec__1___closed__3; -x_11 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_11, 0, x_6); -lean_ctor_set(x_11, 1, x_10); -x_12 = lean_unsigned_to_nat(2u); -x_13 = l_Lean_Syntax_getArg(x_2, x_12); -x_14 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; -x_15 = lean_array_push(x_14, x_9); -x_16 = lean_array_push(x_15, x_11); -x_17 = lean_array_push(x_16, x_13); -x_18 = lean_box(2); -x_19 = l___private_Init_NotationExtra_0__Lean_mkHintBody___closed__2; -x_20 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_20, 0, x_18); -lean_ctor_set(x_20, 1, x_19); -lean_ctor_set(x_20, 2, x_17); -x_21 = l_Array_reverse___rarg(x_1); -x_22 = lean_array_get_size(x_21); -x_23 = lean_usize_of_nat(x_22); -lean_dec(x_22); -x_24 = 0; -x_25 = l_Array_forInUnsafe_loop___at___private_Init_NotationExtra_0__Lean_mkHintBody___spec__1(x_19, x_14, x_21, x_23, x_24, x_20, x_3, x_7); -lean_dec(x_21); -x_26 = !lean_is_exclusive(x_25); -if (x_26 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__19; +x_3 = lean_alloc_ctor(1, 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___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__3() { +_start: { -return x_25; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("term_=_", 7); +return x_1; } -else +} +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__4() { +_start: { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_25, 0); -x_28 = lean_ctor_get(x_25, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_25); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -return x_29; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__3; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; } } +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("hint", 4); +return x_1; +} } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Init_NotationExtra_0__Lean_mkHintBody___spec__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) { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__6() { _start: { -size_t x_9; size_t x_10; lean_object* x_11; -x_9 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_10 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_11 = l_Array_forInUnsafe_loop___at___private_Init_NotationExtra_0__Lean_mkHintBody___spec__1(x_1, x_2, x_3, x_9, x_10, x_6, x_7, x_8); -lean_dec(x_3); -return x_11; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__5; +x_2 = lean_string_utf8_byte_size(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l___private_Init_NotationExtra_0__Lean_mkHintBody___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__7() { _start: { -lean_object* x_5; -x_5 = l___private_Init_NotationExtra_0__Lean_mkHintBody(x_1, x_2, x_3, x_4); -lean_dec(x_2); -return x_5; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__5; +x_2 = lean_unsigned_to_nat(0u); +x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__6; +x_4 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__1() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4; -x_2 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__3; +x_1 = lean_box(0); +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__2() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__9() { _start: { lean_object* x_1; @@ -3636,17 +3751,17 @@ x_1 = lean_mk_string_from_bytes("Command", 7); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__3() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__2; -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__2; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__2; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__9; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__4() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__11() { _start: { lean_object* x_1; @@ -3654,17 +3769,17 @@ x_1 = lean_mk_string_from_bytes("declaration", 11); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__5() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__3; -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__4; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__10; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__11; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__6() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__13() { _start: { lean_object* x_1; @@ -3672,17 +3787,17 @@ x_1 = lean_mk_string_from_bytes("declModifiers", 13); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__7() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__3; -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__6; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__10; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__13; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__8() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__15() { _start: { lean_object* x_1; @@ -3690,17 +3805,17 @@ x_1 = lean_mk_string_from_bytes("attributes", 10); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__9() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4; -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__8; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__15; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__10() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__17() { _start: { lean_object* x_1; @@ -3708,7 +3823,7 @@ x_1 = lean_mk_string_from_bytes("@[", 2); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__11() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__18() { _start: { lean_object* x_1; @@ -3716,17 +3831,17 @@ x_1 = lean_mk_string_from_bytes("attrInstance", 12); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__12() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4; -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__11; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__18; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__13() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__20() { _start: { lean_object* x_1; @@ -3734,17 +3849,17 @@ x_1 = lean_mk_string_from_bytes("Attr", 4); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__14() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__2; -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__13; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__2; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__20; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__15() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__22() { _start: { lean_object* x_1; @@ -3752,17 +3867,17 @@ x_1 = lean_mk_string_from_bytes("simple", 6); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__16() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__14; -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__15; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__21; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__22; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__17() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__24() { _start: { lean_object* x_1; @@ -3770,22 +3885,22 @@ x_1 = lean_mk_string_from_bytes("unificationHint", 15); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__18() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__25() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__17; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__24; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__19() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__17; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__24; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__18; +x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__25; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3793,17 +3908,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__20() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__27() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__17; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__24; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__21() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__28() { _start: { lean_object* x_1; lean_object* x_2; @@ -3812,17 +3927,17 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__22() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__29() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__21; -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__28; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__23() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__30() { _start: { lean_object* x_1; @@ -3830,17 +3945,17 @@ x_1 = lean_mk_string_from_bytes("def", 3); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__24() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__31() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__3; -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__23; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__10; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__30; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__25() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__32() { _start: { lean_object* x_1; @@ -3848,17 +3963,17 @@ x_1 = lean_mk_string_from_bytes("declId", 6); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__26() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__33() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__3; -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__25; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__10; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__32; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__27() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__34() { _start: { lean_object* x_1; @@ -3866,17 +3981,17 @@ x_1 = lean_mk_string_from_bytes("optDeclSig", 10); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__28() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__35() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__3; -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__27; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__10; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__34; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__29() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__36() { _start: { lean_object* x_1; @@ -3884,17 +3999,17 @@ x_1 = lean_mk_string_from_bytes("sort", 4); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__30() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__37() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4; -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__29; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__36; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__31() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__38() { _start: { lean_object* x_1; @@ -3902,7 +4017,7 @@ x_1 = lean_mk_string_from_bytes("Sort", 4); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__32() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__39() { _start: { lean_object* x_1; @@ -3910,27 +4025,27 @@ x_1 = lean_mk_string_from_bytes("Level", 5); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__33() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__40() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__2; -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__32; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__2; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__39; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__34() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__41() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__33; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__40; x_2 = l_Lean_expandExplicitBindersAux_loop___closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__35() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__42() { _start: { lean_object* x_1; @@ -3938,17 +4053,17 @@ x_1 = lean_mk_string_from_bytes("declValSimple", 13); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__36() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__43() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__3; -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__35; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__10; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__42; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__37() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__44() { _start: { lean_object* x_1; @@ -3956,7 +4071,7 @@ x_1 = lean_mk_string_from_bytes(":=", 2); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__38() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__45() { _start: { lean_object* x_1; lean_object* x_2; @@ -3965,55 +4080,14 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__39() { +LEAN_EXPORT lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("hint", 4); -return x_1; -} -} -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__40() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__39; -x_2 = lean_string_utf8_byte_size(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__41() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__39; -x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__40; -x_4 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_4, 0, x_1); -lean_ctor_set(x_4, 1, x_2); -lean_ctor_set(x_4, 2, x_3); -return x_4; -} -} -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__42() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__39; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; uint8_t x_5; -x_4 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__2; -lean_inc(x_1); -x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); -if (x_5 == 0) +lean_object* x_4; uint8_t x_5; +x_4 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__2; +lean_inc(x_1); +x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); +if (x_5 == 0) { lean_object* x_6; lean_object* x_7; lean_dec(x_2); @@ -4029,7 +4103,7 @@ else lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; x_8 = lean_unsigned_to_nat(0u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); -x_10 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__1; +x_10 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__1; lean_inc(x_9); x_11 = l_Lean_Syntax_isOfKind(x_9, x_10); if (x_11 == 0) @@ -4046,875 +4120,667 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +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; x_14 = lean_unsigned_to_nat(2u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); -x_16 = l_Lean_nullKind; -lean_inc(x_15); -x_17 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_8); -if (x_17 == 0) -{ -lean_object* x_18; uint8_t x_19; -x_18 = lean_unsigned_to_nat(1u); -lean_inc(x_15); -x_19 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_18); -if (x_19 == 0) +x_16 = lean_unsigned_to_nat(3u); +x_17 = l_Lean_Syntax_getArg(x_1, x_16); +x_18 = lean_unsigned_to_nat(5u); +x_19 = l_Lean_Syntax_getArg(x_1, x_18); +x_20 = l_Lean_Syntax_getArgs(x_19); +lean_dec(x_19); +x_21 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__2; +x_22 = lean_alloc_closure((void*)(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___lambda__1), 2, 1); +lean_closure_set(x_22, 0, x_21); +x_23 = l_Array_sequenceMap___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__1(x_20, x_22); +lean_dec(x_20); +if (lean_obj_tag(x_23) == 0) { -lean_object* x_20; lean_object* x_21; +lean_object* x_24; lean_object* x_25; +lean_dec(x_17); lean_dec(x_15); lean_dec(x_9); lean_dec(x_2); lean_dec(x_1); -x_20 = lean_box(1); -x_21 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_3); -return x_21; +x_24 = lean_box(1); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_3); +return x_25; } else { -lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_22 = l_Lean_Syntax_getArg(x_15, x_8); -lean_dec(x_15); -x_23 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__10; -lean_inc(x_22); -x_24 = l_Lean_Syntax_isOfKind(x_22, x_23); -if (x_24 == 0) +lean_object* x_26; lean_object* x_27; size_t x_28; size_t x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_26 = lean_ctor_get(x_23, 0); +lean_inc(x_26); +lean_dec(x_23); +x_27 = lean_array_get_size(x_26); +x_28 = lean_usize_of_nat(x_27); +lean_dec(x_27); +x_29 = 0; +lean_inc(x_26); +x_30 = l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__3(x_28, x_29, x_26); +x_31 = l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__4(x_28, x_29, x_26); +x_32 = lean_unsigned_to_nat(7u); +x_33 = l_Lean_Syntax_getArg(x_1, x_32); +lean_dec(x_1); +x_34 = l_Lean_unifConstraint___closed__2; +lean_inc(x_33); +x_35 = l_Lean_Syntax_isOfKind(x_33, x_34); +if (x_35 == 0) { -lean_object* x_25; lean_object* x_26; -lean_dec(x_22); +lean_object* x_36; lean_object* x_37; +lean_dec(x_33); +lean_dec(x_31); +lean_dec(x_30); +lean_dec(x_17); +lean_dec(x_15); lean_dec(x_9); lean_dec(x_2); -lean_dec(x_1); -x_25 = lean_box(1); -x_26 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_3); -return x_26; +x_36 = lean_box(1); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_3); +return x_37; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; 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; uint8_t x_39; -x_27 = lean_unsigned_to_nat(3u); -x_28 = l_Lean_Syntax_getArg(x_1, x_27); -x_29 = lean_unsigned_to_nat(5u); -x_30 = l_Lean_Syntax_getArg(x_1, x_29); -x_31 = lean_unsigned_to_nat(7u); -x_32 = l_Lean_Syntax_getArg(x_1, x_31); -lean_dec(x_1); -x_33 = l_Lean_Syntax_getArgs(x_30); +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_38 = l_Lean_Syntax_getArg(x_33, x_8); +x_39 = l_Lean_Syntax_getArg(x_33, x_14); +lean_dec(x_33); +x_40 = lean_box(0); +x_41 = l_Lean_Syntax_getArgs(x_17); +lean_dec(x_17); +x_42 = l_Lean_Syntax_getOptional_x3f(x_15); +lean_dec(x_15); +if (lean_obj_tag(x_42) == 0) +{ +lean_object* x_298; +x_298 = lean_box(0); +x_43 = x_298; +goto block_297; +} +else +{ +uint8_t x_299; +x_299 = !lean_is_exclusive(x_42); +if (x_299 == 0) +{ +x_43 = x_42; +goto block_297; +} +else +{ +lean_object* x_300; lean_object* x_301; +x_300 = lean_ctor_get(x_42, 0); +lean_inc(x_300); +lean_dec(x_42); +x_301 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_301, 0, x_300); +x_43 = x_301; +goto block_297; +} +} +block_297: +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; size_t x_59; 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; 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; uint8_t x_75; +lean_inc(x_2); +x_44 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_45 = lean_ctor_get(x_44, 0); +lean_inc(x_45); +x_46 = lean_ctor_get(x_44, 1); +lean_inc(x_46); +lean_dec(x_44); +x_47 = l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__5___closed__2; +x_48 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_48, 0, x_45); +lean_ctor_set(x_48, 1, x_47); +x_49 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; +x_50 = lean_array_push(x_49, x_38); +x_51 = lean_array_push(x_50, x_48); +x_52 = lean_array_push(x_51, x_39); +x_53 = lean_box(2); +x_54 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__4; +x_55 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +lean_ctor_set(x_55, 2, x_52); +x_56 = l_Array_zip___rarg(x_31, x_30); lean_dec(x_30); -x_34 = l_Lean_Syntax_getArgs(x_28); -lean_dec(x_28); +lean_dec(x_31); +x_57 = l_Array_reverse___rarg(x_56); +x_58 = lean_array_get_size(x_57); +x_59 = lean_usize_of_nat(x_58); +lean_dec(x_58); +x_60 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4; +x_61 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__2; lean_inc(x_2); -x_35 = l___private_Init_NotationExtra_0__Lean_mkHintBody(x_33, x_32, x_2, x_3); -lean_dec(x_32); -x_36 = lean_ctor_get(x_35, 0); -lean_inc(x_36); -x_37 = lean_ctor_get(x_35, 1); -lean_inc(x_37); -lean_dec(x_35); +x_62 = l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__5(x_60, x_61, x_54, x_49, x_57, x_59, x_29, x_55, x_2, x_46); +lean_dec(x_57); +x_63 = lean_ctor_get(x_62, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_62, 1); +lean_inc(x_64); +lean_dec(x_62); lean_inc(x_2); -x_38 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_37); -x_39 = !lean_is_exclusive(x_38); -if (x_39 == 0) -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; -x_40 = lean_ctor_get(x_38, 0); -x_41 = lean_ctor_get(x_2, 2); -lean_inc(x_41); -x_42 = lean_ctor_get(x_2, 1); -lean_inc(x_42); -lean_dec(x_2); -x_43 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__10; -lean_inc(x_40); -x_44 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_44, 0, x_40); -lean_ctor_set(x_44, 1, x_43); -x_45 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__20; -x_46 = l_Lean_addMacroScope(x_42, x_45, x_41); -x_47 = lean_box(0); -x_48 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__19; -lean_inc(x_40); -x_49 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_49, 0, x_40); -lean_ctor_set(x_49, 1, x_48); -lean_ctor_set(x_49, 2, x_46); -lean_ctor_set(x_49, 3, x_47); -x_50 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_51 = lean_array_push(x_50, x_49); -x_52 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; -x_53 = lean_array_push(x_51, x_52); -x_54 = lean_box(2); -x_55 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__16; -x_56 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_56, 0, x_54); -lean_ctor_set(x_56, 1, x_55); -lean_ctor_set(x_56, 2, x_53); -x_57 = lean_array_push(x_50, x_9); -x_58 = lean_array_push(x_57, x_56); -x_59 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__12; -x_60 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_60, 0, x_54); -lean_ctor_set(x_60, 1, x_59); -lean_ctor_set(x_60, 2, x_58); -x_61 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_62 = lean_array_push(x_61, x_60); -x_63 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_64 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_64, 0, x_54); -lean_ctor_set(x_64, 1, x_63); -lean_ctor_set(x_64, 2, x_62); -x_65 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__13; -lean_inc(x_40); -x_66 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_66, 0, x_40); -lean_ctor_set(x_66, 1, x_65); -x_67 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; -x_68 = lean_array_push(x_67, x_44); -x_69 = lean_array_push(x_68, x_64); -x_70 = lean_array_push(x_69, x_66); -x_71 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__9; -x_72 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_72, 0, x_54); -lean_ctor_set(x_72, 1, x_71); -lean_ctor_set(x_72, 2, x_70); -x_73 = lean_array_push(x_61, x_72); -x_74 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_74, 0, x_54); -lean_ctor_set(x_74, 1, x_63); -lean_ctor_set(x_74, 2, x_73); -x_75 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__22; -x_76 = lean_array_push(x_75, x_74); -x_77 = lean_array_push(x_76, x_52); -x_78 = lean_array_push(x_77, x_52); -x_79 = lean_array_push(x_78, x_52); -x_80 = lean_array_push(x_79, x_52); -x_81 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__7; -x_82 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_82, 0, x_54); +x_65 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_64); +x_66 = lean_ctor_get(x_65, 0); +lean_inc(x_66); +x_67 = lean_ctor_get(x_65, 1); +lean_inc(x_67); +lean_dec(x_65); +x_68 = lean_ctor_get(x_2, 2); +lean_inc(x_68); +x_69 = lean_ctor_get(x_2, 1); +lean_inc(x_69); +x_70 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__8; +lean_inc(x_68); +lean_inc(x_69); +x_71 = l_Lean_addMacroScope(x_69, x_70, x_68); +x_72 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__7; +x_73 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_73, 0, x_66); +lean_ctor_set(x_73, 1, x_72); +lean_ctor_set(x_73, 2, x_71); +lean_ctor_set(x_73, 3, x_40); +x_74 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_67); +x_75 = !lean_is_exclusive(x_74); +if (x_75 == 0) +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; +x_76 = lean_ctor_get(x_74, 0); +x_77 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__17; +lean_inc(x_76); +x_78 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_78, 0, x_76); +lean_ctor_set(x_78, 1, x_77); +x_79 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__27; +x_80 = l_Lean_addMacroScope(x_69, x_79, x_68); +x_81 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__26; +lean_inc(x_76); +x_82 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_82, 0, x_76); lean_ctor_set(x_82, 1, x_81); lean_ctor_set(x_82, 2, x_80); -x_83 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__23; -lean_inc(x_40); -x_84 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_84, 0, x_40); -lean_ctor_set(x_84, 1, x_83); -x_85 = lean_array_push(x_50, x_22); -x_86 = lean_array_push(x_85, x_52); -x_87 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__26; +lean_ctor_set(x_82, 3, x_40); +x_83 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_84 = lean_array_push(x_83, x_82); +x_85 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; +x_86 = lean_array_push(x_84, x_85); +x_87 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__23; x_88 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_88, 0, x_54); +lean_ctor_set(x_88, 0, x_53); lean_ctor_set(x_88, 1, x_87); lean_ctor_set(x_88, 2, x_86); -x_89 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__27; -x_90 = l_Array_append___rarg(x_89, x_34); -x_91 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_91, 0, x_54); -lean_ctor_set(x_91, 1, x_63); -lean_ctor_set(x_91, 2, x_90); -x_92 = l_Lean_expandExplicitBindersAux_loop___closed__12; -lean_inc(x_40); -x_93 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_93, 0, x_40); -lean_ctor_set(x_93, 1, x_92); -x_94 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__31; -lean_inc(x_40); -x_95 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_95, 0, x_40); -lean_ctor_set(x_95, 1, x_94); -x_96 = l_Lean_binderIdent___closed__5; -lean_inc(x_40); -x_97 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_97, 0, x_40); -lean_ctor_set(x_97, 1, x_96); -x_98 = lean_array_push(x_61, x_97); -x_99 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__34; -x_100 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_100, 0, x_54); -lean_ctor_set(x_100, 1, x_99); -lean_ctor_set(x_100, 2, x_98); -x_101 = lean_array_push(x_61, x_100); -x_102 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_102, 0, x_54); -lean_ctor_set(x_102, 1, x_63); -lean_ctor_set(x_102, 2, x_101); -x_103 = lean_array_push(x_50, x_95); -x_104 = lean_array_push(x_103, x_102); -x_105 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__30; -x_106 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_106, 0, x_54); -lean_ctor_set(x_106, 1, x_105); -lean_ctor_set(x_106, 2, x_104); -x_107 = lean_array_push(x_50, x_93); -x_108 = lean_array_push(x_107, x_106); -x_109 = l_Lean_expandExplicitBindersAux_loop___closed__11; -x_110 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_110, 0, x_54); -lean_ctor_set(x_110, 1, x_109); -lean_ctor_set(x_110, 2, x_108); -x_111 = lean_array_push(x_61, x_110); -x_112 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_112, 0, x_54); -lean_ctor_set(x_112, 1, x_63); -lean_ctor_set(x_112, 2, x_111); -x_113 = lean_array_push(x_50, x_91); -x_114 = lean_array_push(x_113, x_112); -x_115 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__28; -x_116 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_116, 0, x_54); -lean_ctor_set(x_116, 1, x_115); -lean_ctor_set(x_116, 2, x_114); -x_117 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__37; -x_118 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_118, 0, x_40); -lean_ctor_set(x_118, 1, x_117); -x_119 = lean_array_push(x_67, x_118); -x_120 = lean_array_push(x_119, x_36); -x_121 = lean_array_push(x_120, x_52); -x_122 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__36; -x_123 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_123, 0, x_54); -lean_ctor_set(x_123, 1, x_122); -lean_ctor_set(x_123, 2, x_121); -x_124 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__38; -x_125 = lean_array_push(x_124, x_84); -x_126 = lean_array_push(x_125, x_88); -x_127 = lean_array_push(x_126, x_116); -x_128 = lean_array_push(x_127, x_123); -x_129 = lean_array_push(x_128, x_52); -x_130 = lean_array_push(x_129, x_52); -x_131 = lean_array_push(x_130, x_52); -x_132 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__24; +x_89 = lean_array_push(x_83, x_9); +x_90 = lean_array_push(x_89, x_88); +x_91 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__19; +x_92 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_92, 0, x_53); +lean_ctor_set(x_92, 1, x_91); +lean_ctor_set(x_92, 2, x_90); +x_93 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_94 = lean_array_push(x_93, x_92); +x_95 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_96 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_96, 0, x_53); +lean_ctor_set(x_96, 1, x_95); +lean_ctor_set(x_96, 2, x_94); +x_97 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__13; +lean_inc(x_76); +x_98 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_98, 0, x_76); +lean_ctor_set(x_98, 1, x_97); +x_99 = lean_array_push(x_49, x_78); +x_100 = lean_array_push(x_99, x_96); +x_101 = lean_array_push(x_100, x_98); +x_102 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__16; +x_103 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_103, 0, x_53); +lean_ctor_set(x_103, 1, x_102); +lean_ctor_set(x_103, 2, x_101); +x_104 = lean_array_push(x_93, x_103); +x_105 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_105, 0, x_53); +lean_ctor_set(x_105, 1, x_95); +lean_ctor_set(x_105, 2, x_104); +x_106 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__29; +x_107 = lean_array_push(x_106, x_105); +x_108 = lean_array_push(x_107, x_85); +x_109 = lean_array_push(x_108, x_85); +x_110 = lean_array_push(x_109, x_85); +x_111 = lean_array_push(x_110, x_85); +x_112 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__14; +x_113 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_113, 0, x_53); +lean_ctor_set(x_113, 1, x_112); +lean_ctor_set(x_113, 2, x_111); +x_114 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__30; +lean_inc(x_76); +x_115 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_115, 0, x_76); +lean_ctor_set(x_115, 1, x_114); +x_116 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__27; +x_117 = l_Array_append___rarg(x_116, x_41); +x_118 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_118, 0, x_53); +lean_ctor_set(x_118, 1, x_95); +lean_ctor_set(x_118, 2, x_117); +x_119 = l_Lean_expandExplicitBindersAux_loop___closed__13; +lean_inc(x_76); +x_120 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_120, 0, x_76); +lean_ctor_set(x_120, 1, x_119); +x_121 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__38; +lean_inc(x_76); +x_122 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_122, 0, x_76); +lean_ctor_set(x_122, 1, x_121); +x_123 = l_Lean_expandExplicitBindersAux_loop___closed__7; +lean_inc(x_76); +x_124 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_124, 0, x_76); +lean_ctor_set(x_124, 1, x_123); +x_125 = lean_array_push(x_93, x_124); +x_126 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__41; +x_127 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_127, 0, x_53); +lean_ctor_set(x_127, 1, x_126); +lean_ctor_set(x_127, 2, x_125); +x_128 = lean_array_push(x_93, x_127); +x_129 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_129, 0, x_53); +lean_ctor_set(x_129, 1, x_95); +lean_ctor_set(x_129, 2, x_128); +x_130 = lean_array_push(x_83, x_122); +x_131 = lean_array_push(x_130, x_129); +x_132 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__37; x_133 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_133, 0, x_54); +lean_ctor_set(x_133, 0, x_53); lean_ctor_set(x_133, 1, x_132); lean_ctor_set(x_133, 2, x_131); -x_134 = lean_array_push(x_50, x_82); +x_134 = lean_array_push(x_83, x_120); x_135 = lean_array_push(x_134, x_133); -x_136 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__5; +x_136 = l_Lean_expandExplicitBindersAux_loop___closed__12; x_137 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_137, 0, x_54); +lean_ctor_set(x_137, 0, x_53); lean_ctor_set(x_137, 1, x_136); lean_ctor_set(x_137, 2, x_135); -lean_ctor_set(x_38, 0, x_137); -return x_38; +x_138 = lean_array_push(x_93, x_137); +x_139 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_139, 0, x_53); +lean_ctor_set(x_139, 1, x_95); +lean_ctor_set(x_139, 2, x_138); +x_140 = lean_array_push(x_83, x_118); +x_141 = lean_array_push(x_140, x_139); +x_142 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__35; +x_143 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_143, 0, x_53); +lean_ctor_set(x_143, 1, x_142); +lean_ctor_set(x_143, 2, x_141); +x_144 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__44; +x_145 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_145, 0, x_76); +lean_ctor_set(x_145, 1, x_144); +x_146 = lean_array_push(x_49, x_145); +x_147 = lean_array_push(x_146, x_63); +x_148 = lean_array_push(x_147, x_85); +x_149 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__43; +x_150 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_150, 0, x_53); +lean_ctor_set(x_150, 1, x_149); +lean_ctor_set(x_150, 2, x_148); +x_151 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__45; +x_152 = lean_array_push(x_151, x_115); +x_153 = lean_array_push(x_83, x_113); +if (lean_obj_tag(x_43) == 0) +{ +lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; +x_154 = lean_array_push(x_83, x_73); +x_155 = lean_array_push(x_154, x_85); +x_156 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__33; +x_157 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_157, 0, x_53); +lean_ctor_set(x_157, 1, x_156); +lean_ctor_set(x_157, 2, x_155); +x_158 = lean_array_push(x_152, x_157); +x_159 = lean_array_push(x_158, x_143); +x_160 = lean_array_push(x_159, x_150); +x_161 = lean_array_push(x_160, x_85); +x_162 = lean_array_push(x_161, x_85); +x_163 = lean_array_push(x_162, x_85); +x_164 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__31; +x_165 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_165, 0, x_53); +lean_ctor_set(x_165, 1, x_164); +lean_ctor_set(x_165, 2, x_163); +x_166 = lean_array_push(x_153, x_165); +x_167 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__12; +x_168 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_168, 0, x_53); +lean_ctor_set(x_168, 1, x_167); +lean_ctor_set(x_168, 2, x_166); +lean_ctor_set(x_74, 0, x_168); +return x_74; } else { -lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; -x_138 = lean_ctor_get(x_38, 0); -x_139 = lean_ctor_get(x_38, 1); -lean_inc(x_139); -lean_inc(x_138); -lean_dec(x_38); -x_140 = lean_ctor_get(x_2, 2); -lean_inc(x_140); -x_141 = lean_ctor_get(x_2, 1); -lean_inc(x_141); -lean_dec(x_2); -x_142 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__10; -lean_inc(x_138); -x_143 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_143, 0, x_138); -lean_ctor_set(x_143, 1, x_142); -x_144 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__20; -x_145 = l_Lean_addMacroScope(x_141, x_144, x_140); -x_146 = lean_box(0); -x_147 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__19; -lean_inc(x_138); -x_148 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_148, 0, x_138); -lean_ctor_set(x_148, 1, x_147); -lean_ctor_set(x_148, 2, x_145); -lean_ctor_set(x_148, 3, x_146); -x_149 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_150 = lean_array_push(x_149, x_148); -x_151 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; -x_152 = lean_array_push(x_150, x_151); -x_153 = lean_box(2); -x_154 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__16; -x_155 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_155, 0, x_153); -lean_ctor_set(x_155, 1, x_154); -lean_ctor_set(x_155, 2, x_152); -x_156 = lean_array_push(x_149, x_9); -x_157 = lean_array_push(x_156, x_155); -x_158 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__12; -x_159 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_159, 0, x_153); -lean_ctor_set(x_159, 1, x_158); -lean_ctor_set(x_159, 2, x_157); -x_160 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_161 = lean_array_push(x_160, x_159); -x_162 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_163 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_163, 0, x_153); -lean_ctor_set(x_163, 1, x_162); -lean_ctor_set(x_163, 2, x_161); -x_164 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__13; -lean_inc(x_138); -x_165 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_165, 0, x_138); -lean_ctor_set(x_165, 1, x_164); -x_166 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; -x_167 = lean_array_push(x_166, x_143); -x_168 = lean_array_push(x_167, x_163); -x_169 = lean_array_push(x_168, x_165); -x_170 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__9; -x_171 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_171, 0, x_153); -lean_ctor_set(x_171, 1, x_170); -lean_ctor_set(x_171, 2, x_169); -x_172 = lean_array_push(x_160, x_171); +lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; +lean_dec(x_73); +x_169 = lean_ctor_get(x_43, 0); +lean_inc(x_169); +lean_dec(x_43); +x_170 = lean_array_push(x_83, x_169); +x_171 = lean_array_push(x_170, x_85); +x_172 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__33; x_173 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_173, 0, x_153); -lean_ctor_set(x_173, 1, x_162); -lean_ctor_set(x_173, 2, x_172); -x_174 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__22; -x_175 = lean_array_push(x_174, x_173); -x_176 = lean_array_push(x_175, x_151); -x_177 = lean_array_push(x_176, x_151); -x_178 = lean_array_push(x_177, x_151); -x_179 = lean_array_push(x_178, x_151); -x_180 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__7; +lean_ctor_set(x_173, 0, x_53); +lean_ctor_set(x_173, 1, x_172); +lean_ctor_set(x_173, 2, x_171); +x_174 = lean_array_push(x_152, x_173); +x_175 = lean_array_push(x_174, x_143); +x_176 = lean_array_push(x_175, x_150); +x_177 = lean_array_push(x_176, x_85); +x_178 = lean_array_push(x_177, x_85); +x_179 = lean_array_push(x_178, x_85); +x_180 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__31; x_181 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_181, 0, x_153); +lean_ctor_set(x_181, 0, x_53); lean_ctor_set(x_181, 1, x_180); lean_ctor_set(x_181, 2, x_179); -x_182 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__23; -lean_inc(x_138); -x_183 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_183, 0, x_138); -lean_ctor_set(x_183, 1, x_182); -x_184 = lean_array_push(x_149, x_22); -x_185 = lean_array_push(x_184, x_151); -x_186 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__26; -x_187 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_187, 0, x_153); -lean_ctor_set(x_187, 1, x_186); -lean_ctor_set(x_187, 2, x_185); -x_188 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__27; -x_189 = l_Array_append___rarg(x_188, x_34); -x_190 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_190, 0, x_153); -lean_ctor_set(x_190, 1, x_162); -lean_ctor_set(x_190, 2, x_189); -x_191 = l_Lean_expandExplicitBindersAux_loop___closed__12; -lean_inc(x_138); -x_192 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_192, 0, x_138); -lean_ctor_set(x_192, 1, x_191); -x_193 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__31; -lean_inc(x_138); -x_194 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_194, 0, x_138); -lean_ctor_set(x_194, 1, x_193); -x_195 = l_Lean_binderIdent___closed__5; -lean_inc(x_138); -x_196 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_196, 0, x_138); -lean_ctor_set(x_196, 1, x_195); -x_197 = lean_array_push(x_160, x_196); -x_198 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__34; -x_199 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_199, 0, x_153); -lean_ctor_set(x_199, 1, x_198); -lean_ctor_set(x_199, 2, x_197); -x_200 = lean_array_push(x_160, x_199); -x_201 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_201, 0, x_153); -lean_ctor_set(x_201, 1, x_162); -lean_ctor_set(x_201, 2, x_200); -x_202 = lean_array_push(x_149, x_194); -x_203 = lean_array_push(x_202, x_201); -x_204 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__30; -x_205 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_205, 0, x_153); -lean_ctor_set(x_205, 1, x_204); -lean_ctor_set(x_205, 2, x_203); -x_206 = lean_array_push(x_149, x_192); -x_207 = lean_array_push(x_206, x_205); -x_208 = l_Lean_expandExplicitBindersAux_loop___closed__11; -x_209 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_209, 0, x_153); -lean_ctor_set(x_209, 1, x_208); -lean_ctor_set(x_209, 2, x_207); -x_210 = lean_array_push(x_160, x_209); -x_211 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_211, 0, x_153); -lean_ctor_set(x_211, 1, x_162); -lean_ctor_set(x_211, 2, x_210); -x_212 = lean_array_push(x_149, x_190); -x_213 = lean_array_push(x_212, x_211); -x_214 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__28; -x_215 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_215, 0, x_153); -lean_ctor_set(x_215, 1, x_214); -lean_ctor_set(x_215, 2, x_213); -x_216 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__37; -x_217 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_217, 0, x_138); -lean_ctor_set(x_217, 1, x_216); -x_218 = lean_array_push(x_166, x_217); -x_219 = lean_array_push(x_218, x_36); -x_220 = lean_array_push(x_219, x_151); -x_221 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__36; -x_222 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_222, 0, x_153); -lean_ctor_set(x_222, 1, x_221); -lean_ctor_set(x_222, 2, x_220); -x_223 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__38; -x_224 = lean_array_push(x_223, x_183); -x_225 = lean_array_push(x_224, x_187); -x_226 = lean_array_push(x_225, x_215); -x_227 = lean_array_push(x_226, x_222); -x_228 = lean_array_push(x_227, x_151); -x_229 = lean_array_push(x_228, x_151); -x_230 = lean_array_push(x_229, x_151); -x_231 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__24; -x_232 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_232, 0, x_153); -lean_ctor_set(x_232, 1, x_231); -lean_ctor_set(x_232, 2, x_230); -x_233 = lean_array_push(x_149, x_181); -x_234 = lean_array_push(x_233, x_232); -x_235 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__5; -x_236 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_236, 0, x_153); -lean_ctor_set(x_236, 1, x_235); -lean_ctor_set(x_236, 2, x_234); -x_237 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_237, 0, x_236); -lean_ctor_set(x_237, 1, x_139); -return x_237; -} -} +x_182 = lean_array_push(x_153, x_181); +x_183 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__12; +x_184 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_184, 0, x_53); +lean_ctor_set(x_184, 1, x_183); +lean_ctor_set(x_184, 2, x_182); +lean_ctor_set(x_74, 0, x_184); +return x_74; } } else { -lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; uint8_t x_250; -lean_dec(x_15); -x_238 = lean_unsigned_to_nat(3u); -x_239 = l_Lean_Syntax_getArg(x_1, x_238); -x_240 = lean_unsigned_to_nat(5u); -x_241 = l_Lean_Syntax_getArg(x_1, x_240); -x_242 = lean_unsigned_to_nat(7u); -x_243 = l_Lean_Syntax_getArg(x_1, x_242); -lean_dec(x_1); -x_244 = l_Lean_Syntax_getArgs(x_241); -lean_dec(x_241); -x_245 = l_Lean_Syntax_getArgs(x_239); -lean_dec(x_239); -lean_inc(x_2); -x_246 = l___private_Init_NotationExtra_0__Lean_mkHintBody(x_244, x_243, x_2, x_3); -lean_dec(x_243); -x_247 = lean_ctor_get(x_246, 0); -lean_inc(x_247); -x_248 = lean_ctor_get(x_246, 1); -lean_inc(x_248); -lean_dec(x_246); -lean_inc(x_2); -x_249 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_248); -x_250 = !lean_is_exclusive(x_249); -if (x_250 == 0) -{ -lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; -x_251 = lean_ctor_get(x_249, 0); -x_252 = lean_ctor_get(x_2, 2); -lean_inc(x_252); -x_253 = lean_ctor_get(x_2, 1); -lean_inc(x_253); -lean_dec(x_2); -x_254 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__10; -lean_inc(x_251); -x_255 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_255, 0, x_251); -lean_ctor_set(x_255, 1, x_254); -x_256 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__20; -lean_inc(x_252); -lean_inc(x_253); -x_257 = l_Lean_addMacroScope(x_253, x_256, x_252); -x_258 = lean_box(0); -x_259 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__19; -lean_inc(x_251); -x_260 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_260, 0, x_251); -lean_ctor_set(x_260, 1, x_259); -lean_ctor_set(x_260, 2, x_257); -lean_ctor_set(x_260, 3, x_258); -x_261 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_262 = lean_array_push(x_261, x_260); -x_263 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; -x_264 = lean_array_push(x_262, x_263); -x_265 = lean_box(2); -x_266 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__16; -x_267 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_267, 0, x_265); -lean_ctor_set(x_267, 1, x_266); -lean_ctor_set(x_267, 2, x_264); -x_268 = lean_array_push(x_261, x_9); -x_269 = lean_array_push(x_268, x_267); -x_270 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__12; -x_271 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_271, 0, x_265); -lean_ctor_set(x_271, 1, x_270); -lean_ctor_set(x_271, 2, x_269); -x_272 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_273 = lean_array_push(x_272, x_271); -x_274 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_275 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_275, 0, x_265); -lean_ctor_set(x_275, 1, x_274); -lean_ctor_set(x_275, 2, x_273); -x_276 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__13; -lean_inc(x_251); -x_277 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_277, 0, x_251); -lean_ctor_set(x_277, 1, x_276); -x_278 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; -x_279 = lean_array_push(x_278, x_255); -x_280 = lean_array_push(x_279, x_275); -x_281 = lean_array_push(x_280, x_277); -x_282 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__9; -x_283 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_283, 0, x_265); -lean_ctor_set(x_283, 1, x_282); -lean_ctor_set(x_283, 2, x_281); -x_284 = lean_array_push(x_272, x_283); -x_285 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_285, 0, x_265); -lean_ctor_set(x_285, 1, x_274); -lean_ctor_set(x_285, 2, x_284); -x_286 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__22; -x_287 = lean_array_push(x_286, x_285); -x_288 = lean_array_push(x_287, x_263); -x_289 = lean_array_push(x_288, x_263); -x_290 = lean_array_push(x_289, x_263); -x_291 = lean_array_push(x_290, x_263); -x_292 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__7; -x_293 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_293, 0, x_265); -lean_ctor_set(x_293, 1, x_292); -lean_ctor_set(x_293, 2, x_291); -x_294 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__23; -lean_inc(x_251); -x_295 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_295, 0, x_251); -lean_ctor_set(x_295, 1, x_294); -x_296 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__42; -x_297 = l_Lean_addMacroScope(x_253, x_296, x_252); -x_298 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__41; -lean_inc(x_251); -x_299 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_299, 0, x_251); -lean_ctor_set(x_299, 1, x_298); -lean_ctor_set(x_299, 2, x_297); -lean_ctor_set(x_299, 3, x_258); -x_300 = lean_array_push(x_261, x_299); -x_301 = lean_array_push(x_300, x_263); -x_302 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__26; -x_303 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_303, 0, x_265); -lean_ctor_set(x_303, 1, x_302); -lean_ctor_set(x_303, 2, x_301); -x_304 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__27; -x_305 = l_Array_append___rarg(x_304, x_245); -x_306 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_306, 0, x_265); -lean_ctor_set(x_306, 1, x_274); -lean_ctor_set(x_306, 2, x_305); -x_307 = l_Lean_expandExplicitBindersAux_loop___closed__12; -lean_inc(x_251); -x_308 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_308, 0, x_251); -lean_ctor_set(x_308, 1, x_307); -x_309 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__31; -lean_inc(x_251); -x_310 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_310, 0, x_251); -lean_ctor_set(x_310, 1, x_309); -x_311 = l_Lean_binderIdent___closed__5; -lean_inc(x_251); -x_312 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_312, 0, x_251); -lean_ctor_set(x_312, 1, x_311); -x_313 = lean_array_push(x_272, x_312); -x_314 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__34; -x_315 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_315, 0, x_265); -lean_ctor_set(x_315, 1, x_314); -lean_ctor_set(x_315, 2, x_313); -x_316 = lean_array_push(x_272, x_315); -x_317 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_317, 0, x_265); -lean_ctor_set(x_317, 1, x_274); -lean_ctor_set(x_317, 2, x_316); -x_318 = lean_array_push(x_261, x_310); -x_319 = lean_array_push(x_318, x_317); -x_320 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__30; -x_321 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_321, 0, x_265); -lean_ctor_set(x_321, 1, x_320); -lean_ctor_set(x_321, 2, x_319); -x_322 = lean_array_push(x_261, x_308); -x_323 = lean_array_push(x_322, x_321); -x_324 = l_Lean_expandExplicitBindersAux_loop___closed__11; -x_325 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_325, 0, x_265); -lean_ctor_set(x_325, 1, x_324); -lean_ctor_set(x_325, 2, x_323); -x_326 = lean_array_push(x_272, x_325); -x_327 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_327, 0, x_265); -lean_ctor_set(x_327, 1, x_274); -lean_ctor_set(x_327, 2, x_326); -x_328 = lean_array_push(x_261, x_306); -x_329 = lean_array_push(x_328, x_327); -x_330 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__28; -x_331 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_331, 0, x_265); -lean_ctor_set(x_331, 1, x_330); -lean_ctor_set(x_331, 2, x_329); -x_332 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__37; -x_333 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_333, 0, x_251); -lean_ctor_set(x_333, 1, x_332); -x_334 = lean_array_push(x_278, x_333); -x_335 = lean_array_push(x_334, x_247); -x_336 = lean_array_push(x_335, x_263); -x_337 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__36; -x_338 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_338, 0, x_265); -lean_ctor_set(x_338, 1, x_337); -lean_ctor_set(x_338, 2, x_336); -x_339 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__38; -x_340 = lean_array_push(x_339, x_295); -x_341 = lean_array_push(x_340, x_303); -x_342 = lean_array_push(x_341, x_331); -x_343 = lean_array_push(x_342, x_338); -x_344 = lean_array_push(x_343, x_263); -x_345 = lean_array_push(x_344, x_263); -x_346 = lean_array_push(x_345, x_263); -x_347 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__24; -x_348 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_348, 0, x_265); -lean_ctor_set(x_348, 1, x_347); -lean_ctor_set(x_348, 2, x_346); -x_349 = lean_array_push(x_261, x_293); -x_350 = lean_array_push(x_349, x_348); -x_351 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__5; -x_352 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_352, 0, x_265); -lean_ctor_set(x_352, 1, x_351); -lean_ctor_set(x_352, 2, x_350); -lean_ctor_set(x_249, 0, x_352); -return x_249; +lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; +x_185 = lean_ctor_get(x_74, 0); +x_186 = lean_ctor_get(x_74, 1); +lean_inc(x_186); +lean_inc(x_185); +lean_dec(x_74); +x_187 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__17; +lean_inc(x_185); +x_188 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_188, 0, x_185); +lean_ctor_set(x_188, 1, x_187); +x_189 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__27; +x_190 = l_Lean_addMacroScope(x_69, x_189, x_68); +x_191 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__26; +lean_inc(x_185); +x_192 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_192, 0, x_185); +lean_ctor_set(x_192, 1, x_191); +lean_ctor_set(x_192, 2, x_190); +lean_ctor_set(x_192, 3, x_40); +x_193 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_194 = lean_array_push(x_193, x_192); +x_195 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; +x_196 = lean_array_push(x_194, x_195); +x_197 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__23; +x_198 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_198, 0, x_53); +lean_ctor_set(x_198, 1, x_197); +lean_ctor_set(x_198, 2, x_196); +x_199 = lean_array_push(x_193, x_9); +x_200 = lean_array_push(x_199, x_198); +x_201 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__19; +x_202 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_202, 0, x_53); +lean_ctor_set(x_202, 1, x_201); +lean_ctor_set(x_202, 2, x_200); +x_203 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_204 = lean_array_push(x_203, x_202); +x_205 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_206 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_206, 0, x_53); +lean_ctor_set(x_206, 1, x_205); +lean_ctor_set(x_206, 2, x_204); +x_207 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__13; +lean_inc(x_185); +x_208 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_208, 0, x_185); +lean_ctor_set(x_208, 1, x_207); +x_209 = lean_array_push(x_49, x_188); +x_210 = lean_array_push(x_209, x_206); +x_211 = lean_array_push(x_210, x_208); +x_212 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__16; +x_213 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_213, 0, x_53); +lean_ctor_set(x_213, 1, x_212); +lean_ctor_set(x_213, 2, x_211); +x_214 = lean_array_push(x_203, x_213); +x_215 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_215, 0, x_53); +lean_ctor_set(x_215, 1, x_205); +lean_ctor_set(x_215, 2, x_214); +x_216 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__29; +x_217 = lean_array_push(x_216, x_215); +x_218 = lean_array_push(x_217, x_195); +x_219 = lean_array_push(x_218, x_195); +x_220 = lean_array_push(x_219, x_195); +x_221 = lean_array_push(x_220, x_195); +x_222 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__14; +x_223 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_223, 0, x_53); +lean_ctor_set(x_223, 1, x_222); +lean_ctor_set(x_223, 2, x_221); +x_224 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__30; +lean_inc(x_185); +x_225 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_225, 0, x_185); +lean_ctor_set(x_225, 1, x_224); +x_226 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__27; +x_227 = l_Array_append___rarg(x_226, x_41); +x_228 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_228, 0, x_53); +lean_ctor_set(x_228, 1, x_205); +lean_ctor_set(x_228, 2, x_227); +x_229 = l_Lean_expandExplicitBindersAux_loop___closed__13; +lean_inc(x_185); +x_230 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_230, 0, x_185); +lean_ctor_set(x_230, 1, x_229); +x_231 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__38; +lean_inc(x_185); +x_232 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_232, 0, x_185); +lean_ctor_set(x_232, 1, x_231); +x_233 = l_Lean_expandExplicitBindersAux_loop___closed__7; +lean_inc(x_185); +x_234 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_234, 0, x_185); +lean_ctor_set(x_234, 1, x_233); +x_235 = lean_array_push(x_203, x_234); +x_236 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__41; +x_237 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_237, 0, x_53); +lean_ctor_set(x_237, 1, x_236); +lean_ctor_set(x_237, 2, x_235); +x_238 = lean_array_push(x_203, x_237); +x_239 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_239, 0, x_53); +lean_ctor_set(x_239, 1, x_205); +lean_ctor_set(x_239, 2, x_238); +x_240 = lean_array_push(x_193, x_232); +x_241 = lean_array_push(x_240, x_239); +x_242 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__37; +x_243 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_243, 0, x_53); +lean_ctor_set(x_243, 1, x_242); +lean_ctor_set(x_243, 2, x_241); +x_244 = lean_array_push(x_193, x_230); +x_245 = lean_array_push(x_244, x_243); +x_246 = l_Lean_expandExplicitBindersAux_loop___closed__12; +x_247 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_247, 0, x_53); +lean_ctor_set(x_247, 1, x_246); +lean_ctor_set(x_247, 2, x_245); +x_248 = lean_array_push(x_203, x_247); +x_249 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_249, 0, x_53); +lean_ctor_set(x_249, 1, x_205); +lean_ctor_set(x_249, 2, x_248); +x_250 = lean_array_push(x_193, x_228); +x_251 = lean_array_push(x_250, x_249); +x_252 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__35; +x_253 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_253, 0, x_53); +lean_ctor_set(x_253, 1, x_252); +lean_ctor_set(x_253, 2, x_251); +x_254 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__44; +x_255 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_255, 0, x_185); +lean_ctor_set(x_255, 1, x_254); +x_256 = lean_array_push(x_49, x_255); +x_257 = lean_array_push(x_256, x_63); +x_258 = lean_array_push(x_257, x_195); +x_259 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__43; +x_260 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_260, 0, x_53); +lean_ctor_set(x_260, 1, x_259); +lean_ctor_set(x_260, 2, x_258); +x_261 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__45; +x_262 = lean_array_push(x_261, x_225); +x_263 = lean_array_push(x_193, x_223); +if (lean_obj_tag(x_43) == 0) +{ +lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; +x_264 = lean_array_push(x_193, x_73); +x_265 = lean_array_push(x_264, x_195); +x_266 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__33; +x_267 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_267, 0, x_53); +lean_ctor_set(x_267, 1, x_266); +lean_ctor_set(x_267, 2, x_265); +x_268 = lean_array_push(x_262, x_267); +x_269 = lean_array_push(x_268, x_253); +x_270 = lean_array_push(x_269, x_260); +x_271 = lean_array_push(x_270, x_195); +x_272 = lean_array_push(x_271, x_195); +x_273 = lean_array_push(x_272, x_195); +x_274 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__31; +x_275 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_275, 0, x_53); +lean_ctor_set(x_275, 1, x_274); +lean_ctor_set(x_275, 2, x_273); +x_276 = lean_array_push(x_263, x_275); +x_277 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__12; +x_278 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_278, 0, x_53); +lean_ctor_set(x_278, 1, x_277); +lean_ctor_set(x_278, 2, x_276); +x_279 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_279, 0, x_278); +lean_ctor_set(x_279, 1, x_186); +return x_279; } else { -lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; -x_353 = lean_ctor_get(x_249, 0); -x_354 = lean_ctor_get(x_249, 1); -lean_inc(x_354); -lean_inc(x_353); -lean_dec(x_249); -x_355 = lean_ctor_get(x_2, 2); -lean_inc(x_355); -x_356 = lean_ctor_get(x_2, 1); -lean_inc(x_356); +lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; +lean_dec(x_73); +x_280 = lean_ctor_get(x_43, 0); +lean_inc(x_280); +lean_dec(x_43); +x_281 = lean_array_push(x_193, x_280); +x_282 = lean_array_push(x_281, x_195); +x_283 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__33; +x_284 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_284, 0, x_53); +lean_ctor_set(x_284, 1, x_283); +lean_ctor_set(x_284, 2, x_282); +x_285 = lean_array_push(x_262, x_284); +x_286 = lean_array_push(x_285, x_253); +x_287 = lean_array_push(x_286, x_260); +x_288 = lean_array_push(x_287, x_195); +x_289 = lean_array_push(x_288, x_195); +x_290 = lean_array_push(x_289, x_195); +x_291 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__31; +x_292 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_292, 0, x_53); +lean_ctor_set(x_292, 1, x_291); +lean_ctor_set(x_292, 2, x_290); +x_293 = lean_array_push(x_263, x_292); +x_294 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__12; +x_295 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_295, 0, x_53); +lean_ctor_set(x_295, 1, x_294); +lean_ctor_set(x_295, 2, x_293); +x_296 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_296, 0, x_295); +lean_ctor_set(x_296, 1, x_186); +return x_296; +} +} +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__2___boxed(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; +x_6 = l_Array_sequenceMap_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__2(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Array_sequenceMap___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__1(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_357 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__10; -lean_inc(x_353); -x_358 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_358, 0, x_353); -lean_ctor_set(x_358, 1, x_357); -x_359 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__20; -lean_inc(x_355); -lean_inc(x_356); -x_360 = l_Lean_addMacroScope(x_356, x_359, x_355); -x_361 = lean_box(0); -x_362 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__19; -lean_inc(x_353); -x_363 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_363, 0, x_353); -lean_ctor_set(x_363, 1, x_362); -lean_ctor_set(x_363, 2, x_360); -lean_ctor_set(x_363, 3, x_361); -x_364 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_365 = lean_array_push(x_364, x_363); -x_366 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; -x_367 = lean_array_push(x_365, x_366); -x_368 = lean_box(2); -x_369 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__16; -x_370 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_370, 0, x_368); -lean_ctor_set(x_370, 1, x_369); -lean_ctor_set(x_370, 2, x_367); -x_371 = lean_array_push(x_364, x_9); -x_372 = lean_array_push(x_371, x_370); -x_373 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__12; -x_374 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_374, 0, x_368); -lean_ctor_set(x_374, 1, x_373); -lean_ctor_set(x_374, 2, x_372); -x_375 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_376 = lean_array_push(x_375, x_374); -x_377 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_378 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_378, 0, x_368); -lean_ctor_set(x_378, 1, x_377); -lean_ctor_set(x_378, 2, x_376); -x_379 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__13; -lean_inc(x_353); -x_380 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_380, 0, x_353); -lean_ctor_set(x_380, 1, x_379); -x_381 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; -x_382 = lean_array_push(x_381, x_358); -x_383 = lean_array_push(x_382, x_378); -x_384 = lean_array_push(x_383, x_380); -x_385 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__9; -x_386 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_386, 0, x_368); -lean_ctor_set(x_386, 1, x_385); -lean_ctor_set(x_386, 2, x_384); -x_387 = lean_array_push(x_375, x_386); -x_388 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_388, 0, x_368); -lean_ctor_set(x_388, 1, x_377); -lean_ctor_set(x_388, 2, x_387); -x_389 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__22; -x_390 = lean_array_push(x_389, x_388); -x_391 = lean_array_push(x_390, x_366); -x_392 = lean_array_push(x_391, x_366); -x_393 = lean_array_push(x_392, x_366); -x_394 = lean_array_push(x_393, x_366); -x_395 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__7; -x_396 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_396, 0, x_368); -lean_ctor_set(x_396, 1, x_395); -lean_ctor_set(x_396, 2, x_394); -x_397 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__23; -lean_inc(x_353); -x_398 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_398, 0, x_353); -lean_ctor_set(x_398, 1, x_397); -x_399 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__42; -x_400 = l_Lean_addMacroScope(x_356, x_399, x_355); -x_401 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__41; -lean_inc(x_353); -x_402 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_402, 0, x_353); -lean_ctor_set(x_402, 1, x_401); -lean_ctor_set(x_402, 2, x_400); -lean_ctor_set(x_402, 3, x_361); -x_403 = lean_array_push(x_364, x_402); -x_404 = lean_array_push(x_403, x_366); -x_405 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__26; -x_406 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_406, 0, x_368); -lean_ctor_set(x_406, 1, x_405); -lean_ctor_set(x_406, 2, x_404); -x_407 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__27; -x_408 = l_Array_append___rarg(x_407, x_245); -x_409 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_409, 0, x_368); -lean_ctor_set(x_409, 1, x_377); -lean_ctor_set(x_409, 2, x_408); -x_410 = l_Lean_expandExplicitBindersAux_loop___closed__12; -lean_inc(x_353); -x_411 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_411, 0, x_353); -lean_ctor_set(x_411, 1, x_410); -x_412 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__31; -lean_inc(x_353); -x_413 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_413, 0, x_353); -lean_ctor_set(x_413, 1, x_412); -x_414 = l_Lean_binderIdent___closed__5; -lean_inc(x_353); -x_415 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_415, 0, x_353); -lean_ctor_set(x_415, 1, x_414); -x_416 = lean_array_push(x_375, x_415); -x_417 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__34; -x_418 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_418, 0, x_368); -lean_ctor_set(x_418, 1, x_417); -lean_ctor_set(x_418, 2, x_416); -x_419 = lean_array_push(x_375, x_418); -x_420 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_420, 0, x_368); -lean_ctor_set(x_420, 1, x_377); -lean_ctor_set(x_420, 2, x_419); -x_421 = lean_array_push(x_364, x_413); -x_422 = lean_array_push(x_421, x_420); -x_423 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__30; -x_424 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_424, 0, x_368); -lean_ctor_set(x_424, 1, x_423); -lean_ctor_set(x_424, 2, x_422); -x_425 = lean_array_push(x_364, x_411); -x_426 = lean_array_push(x_425, x_424); -x_427 = l_Lean_expandExplicitBindersAux_loop___closed__11; -x_428 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_428, 0, x_368); -lean_ctor_set(x_428, 1, x_427); -lean_ctor_set(x_428, 2, x_426); -x_429 = lean_array_push(x_375, x_428); -x_430 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_430, 0, x_368); -lean_ctor_set(x_430, 1, x_377); -lean_ctor_set(x_430, 2, x_429); -x_431 = lean_array_push(x_364, x_409); -x_432 = lean_array_push(x_431, x_430); -x_433 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__28; -x_434 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_434, 0, x_368); -lean_ctor_set(x_434, 1, x_433); -lean_ctor_set(x_434, 2, x_432); -x_435 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__37; -x_436 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_436, 0, x_353); -lean_ctor_set(x_436, 1, x_435); -x_437 = lean_array_push(x_381, x_436); -x_438 = lean_array_push(x_437, x_247); -x_439 = lean_array_push(x_438, x_366); -x_440 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__36; -x_441 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_441, 0, x_368); -lean_ctor_set(x_441, 1, x_440); -lean_ctor_set(x_441, 2, x_439); -x_442 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__38; -x_443 = lean_array_push(x_442, x_398); -x_444 = lean_array_push(x_443, x_406); -x_445 = lean_array_push(x_444, x_434); -x_446 = lean_array_push(x_445, x_441); -x_447 = lean_array_push(x_446, x_366); -x_448 = lean_array_push(x_447, x_366); -x_449 = lean_array_push(x_448, x_366); -x_450 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__24; -x_451 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_451, 0, x_368); -lean_ctor_set(x_451, 1, x_450); -lean_ctor_set(x_451, 2, x_449); -x_452 = lean_array_push(x_364, x_396); -x_453 = lean_array_push(x_452, x_451); -x_454 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__5; -x_455 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_455, 0, x_368); -lean_ctor_set(x_455, 1, x_454); -lean_ctor_set(x_455, 2, x_453); -x_456 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_456, 0, x_455); -lean_ctor_set(x_456, 1, x_354); -return x_456; +x_6 = l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__3(x_4, x_5, x_3); +return x_6; } } +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__4(x_4, x_5, x_3); +return x_6; } } +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__5___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_object* x_9, lean_object* x_10) { +_start: +{ +size_t x_11; size_t x_12; lean_object* x_13; +x_11 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_12 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_13 = l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__5(x_1, x_2, x_3, x_4, x_5, x_11, x_12, x_8, x_9, x_10); +lean_dec(x_5); +lean_dec(x_2); +return x_13; } } static lean_object* _init_l_term_u2203___x2c_____closed__1() { @@ -4957,7 +4823,7 @@ static lean_object* _init_l_term_u2203___x2c_____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_term_u2203___x2c_____closed__4; x_3 = l_Lean_explicitBinders; x_4 = lean_alloc_ctor(2, 3, 0); @@ -4971,7 +4837,7 @@ static lean_object* _init_l_term_u2203___x2c_____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_term_u2203___x2c_____closed__5; x_3 = l_Lean_unifConstraintElem___closed__8; x_4 = lean_alloc_ctor(2, 3, 0); @@ -4985,9 +4851,9 @@ static lean_object* _init_l_term_u2203___x2c_____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_term_u2203___x2c_____closed__6; -x_3 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__20; +x_3 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__20; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -5108,7 +4974,7 @@ static lean_object* _init_l_termExists___x2c_____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_termExists___x2c_____closed__4; x_3 = l_Lean_explicitBinders; x_4 = lean_alloc_ctor(2, 3, 0); @@ -5122,7 +4988,7 @@ static lean_object* _init_l_termExists___x2c_____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_termExists___x2c_____closed__5; x_3 = l_Lean_unifConstraintElem___closed__8; x_4 = lean_alloc_ctor(2, 3, 0); @@ -5136,9 +5002,9 @@ static lean_object* _init_l_termExists___x2c_____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_termExists___x2c_____closed__6; -x_3 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__20; +x_3 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__20; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -5241,7 +5107,7 @@ static lean_object* _init_l_term_u03a3___x2c_____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_term_u03a3___x2c_____closed__4; x_3 = l_Lean_explicitBinders; x_4 = lean_alloc_ctor(2, 3, 0); @@ -5255,7 +5121,7 @@ static lean_object* _init_l_term_u03a3___x2c_____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_term_u03a3___x2c_____closed__5; x_3 = l_Lean_unifConstraintElem___closed__8; x_4 = lean_alloc_ctor(2, 3, 0); @@ -5269,9 +5135,9 @@ static lean_object* _init_l_term_u03a3___x2c_____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_term_u03a3___x2c_____closed__6; -x_3 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__20; +x_3 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__20; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -5392,7 +5258,7 @@ static lean_object* _init_l_term_u03a3_x27___x2c_____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_term_u03a3_x27___x2c_____closed__4; x_3 = l_Lean_explicitBinders; x_4 = lean_alloc_ctor(2, 3, 0); @@ -5406,7 +5272,7 @@ static lean_object* _init_l_term_u03a3_x27___x2c_____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_term_u03a3_x27___x2c_____closed__5; x_3 = l_Lean_unifConstraintElem___closed__8; x_4 = lean_alloc_ctor(2, 3, 0); @@ -5420,9 +5286,9 @@ static lean_object* _init_l_term_u03a3_x27___x2c_____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_term_u03a3_x27___x2c_____closed__6; -x_3 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__20; +x_3 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__20; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -5543,7 +5409,7 @@ static lean_object* _init_l_term___xd7____1___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_Lean_bracketedExplicitBinders; x_3 = l_term___xd7____1___closed__4; x_4 = lean_alloc_ctor(2, 3, 0); @@ -5557,7 +5423,7 @@ static lean_object* _init_l_term___xd7____1___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__19; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__19; x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -5569,7 +5435,7 @@ static lean_object* _init_l_term___xd7____1___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_term___xd7____1___closed__5; x_3 = l_term___xd7____1___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); @@ -5673,7 +5539,7 @@ static lean_object* _init_l_term___xd7_x27_____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_Lean_bracketedExplicitBinders; x_3 = l_term___xd7_x27_____closed__4; x_4 = lean_alloc_ctor(2, 3, 0); @@ -5687,7 +5553,7 @@ static lean_object* _init_l_term___xd7_x27_____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_term___xd7_x27_____closed__5; x_3 = l_term___xd7____1___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); @@ -5791,9 +5657,9 @@ static lean_object* _init_l_calcStep___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_Lean_unifConstraintElem___closed__5; -x_3 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__20; +x_3 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__20; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -5823,7 +5689,7 @@ static lean_object* _init_l_calcStep___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_calcStep___closed__5; x_3 = l_calcStep___closed__7; x_4 = lean_alloc_ctor(2, 3, 0); @@ -5837,8 +5703,8 @@ static lean_object* _init_l_calcStep___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__22; -x_2 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__20; +x_1 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__22; +x_2 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__20; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -5849,7 +5715,7 @@ static lean_object* _init_l_calcStep___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_calcStep___closed__8; x_3 = l_calcStep___closed__9; x_4 = lean_alloc_ctor(2, 3, 0); @@ -5953,7 +5819,7 @@ static lean_object* _init_l_calc___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_calc___closed__3; x_3 = l_calc___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); @@ -5967,7 +5833,7 @@ static lean_object* _init_l_calc___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_calcStep; x_3 = l_calc___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); @@ -5993,7 +5859,7 @@ static lean_object* _init_l_calc___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__22; +x_1 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__22; x_2 = l_calc___closed__9; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6005,7 +5871,7 @@ static lean_object* _init_l_calc___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_calc___closed__7; x_3 = l_calc___closed__10; x_4 = lean_alloc_ctor(2, 3, 0); @@ -6090,22 +5956,10 @@ return x_3; static lean_object* _init_l_tacticCalc_____closed__6() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__22; -x_2 = l_tacticCalc_____closed__5; -x_3 = lean_alloc_ctor(1, 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_tacticCalc_____closed__7() { -_start: -{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_tacticCalc_____closed__4; -x_3 = l_tacticCalc_____closed__6; +x_3 = l_tacticCalc_____closed__5; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -6113,13 +5967,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_tacticCalc_____closed__8() { +static lean_object* _init_l_tacticCalc_____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_tacticCalc_____closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_tacticCalc_____closed__7; +x_3 = l_tacticCalc_____closed__6; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -6131,7 +5985,7 @@ static lean_object* _init_l_tacticCalc__() { _start: { lean_object* x_1; -x_1 = l_tacticCalc_____closed__8; +x_1 = l_tacticCalc_____closed__7; return x_1; } } @@ -6147,7 +6001,7 @@ static lean_object* _init_l___aux__Init__NotationExtra______macroRules__tacticCa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__2; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__2; x_2 = l___aux__Init__NotationExtra______macroRules__tacticCalc____1___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -6209,51 +6063,51 @@ return x_7; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; x_8 = lean_unsigned_to_nat(1u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); lean_dec(x_1); -x_10 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_11 = !lean_is_exclusive(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; 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; 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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_12 = lean_ctor_get(x_10, 0); -x_13 = l___aux__Init__NotationExtra______macroRules__tacticCalc____1___closed__5; -lean_inc(x_12); -x_14 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_14, 0, x_12); -lean_ctor_set(x_14, 1, x_13); -x_15 = l_calc___closed__1; -x_16 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_16, 0, x_12); -lean_ctor_set(x_16, 1, x_15); -x_17 = l_Lean_Syntax_getArgs(x_9); +x_10 = l_Lean_Syntax_getArgs(x_9); lean_dec(x_9); -x_18 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__27; -x_19 = l_Array_append___rarg(x_18, x_17); +x_11 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +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; 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; 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; +x_13 = lean_ctor_get(x_11, 0); +x_14 = l___aux__Init__NotationExtra______macroRules__tacticCalc____1___closed__5; +lean_inc(x_13); +x_15 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_14); +x_16 = l_calc___closed__1; +x_17 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_17, 0, x_13); +lean_ctor_set(x_17, 1, x_16); +x_18 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__27; +x_19 = l_Array_append___rarg(x_18, x_10); x_20 = lean_box(2); -x_21 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; +x_21 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; x_22 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_22, 0, x_20); lean_ctor_set(x_22, 1, x_21); lean_ctor_set(x_22, 2, x_19); -x_23 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_24 = lean_array_push(x_23, x_16); +x_23 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_24 = lean_array_push(x_23, x_17); x_25 = lean_array_push(x_24, x_22); x_26 = l_calc___closed__2; x_27 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_27, 0, x_20); lean_ctor_set(x_27, 1, x_26); lean_ctor_set(x_27, 2, x_25); -x_28 = lean_array_push(x_23, x_14); +x_28 = lean_array_push(x_23, x_15); x_29 = lean_array_push(x_28, x_27); x_30 = l___aux__Init__NotationExtra______macroRules__tacticCalc____1___closed__6; x_31 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_31, 0, x_20); lean_ctor_set(x_31, 1, x_30); lean_ctor_set(x_31, 2, x_29); -x_32 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; +x_32 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; x_33 = lean_array_push(x_32, x_31); x_34 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_34, 0, x_20); @@ -6265,17 +6119,17 @@ x_37 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_37, 0, x_20); lean_ctor_set(x_37, 1, x_36); lean_ctor_set(x_37, 2, x_35); -lean_ctor_set(x_10, 0, x_37); -return x_10; +lean_ctor_set(x_11, 0, x_37); +return x_11; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_38 = lean_ctor_get(x_10, 0); -x_39 = lean_ctor_get(x_10, 1); +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_38 = lean_ctor_get(x_11, 0); +x_39 = lean_ctor_get(x_11, 1); lean_inc(x_39); lean_inc(x_38); -lean_dec(x_10); +lean_dec(x_11); x_40 = l___aux__Init__NotationExtra______macroRules__tacticCalc____1___closed__5; lean_inc(x_38); x_41 = lean_alloc_ctor(2, 2, 0); @@ -6285,47 +6139,45 @@ x_42 = l_calc___closed__1; x_43 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_43, 0, x_38); lean_ctor_set(x_43, 1, x_42); -x_44 = l_Lean_Syntax_getArgs(x_9); -lean_dec(x_9); -x_45 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__27; -x_46 = l_Array_append___rarg(x_45, x_44); -x_47 = lean_box(2); -x_48 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_49 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_49, 0, x_47); -lean_ctor_set(x_49, 1, x_48); -lean_ctor_set(x_49, 2, x_46); -x_50 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_51 = lean_array_push(x_50, x_43); -x_52 = lean_array_push(x_51, x_49); -x_53 = l_calc___closed__2; -x_54 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_54, 0, x_47); -lean_ctor_set(x_54, 1, x_53); -lean_ctor_set(x_54, 2, x_52); -x_55 = lean_array_push(x_50, x_41); -x_56 = lean_array_push(x_55, x_54); -x_57 = l___aux__Init__NotationExtra______macroRules__tacticCalc____1___closed__6; -x_58 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_58, 0, x_47); -lean_ctor_set(x_58, 1, x_57); -lean_ctor_set(x_58, 2, x_56); -x_59 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_60 = lean_array_push(x_59, x_58); -x_61 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_61, 0, x_47); -lean_ctor_set(x_61, 1, x_48); -lean_ctor_set(x_61, 2, x_60); -x_62 = lean_array_push(x_59, x_61); -x_63 = l___aux__Init__NotationExtra______macroRules__tacticCalc____1___closed__4; -x_64 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_64, 0, x_47); -lean_ctor_set(x_64, 1, x_63); -lean_ctor_set(x_64, 2, x_62); -x_65 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_65, 0, x_64); -lean_ctor_set(x_65, 1, x_39); -return x_65; +x_44 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__27; +x_45 = l_Array_append___rarg(x_44, x_10); +x_46 = lean_box(2); +x_47 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_48 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_48, 0, x_46); +lean_ctor_set(x_48, 1, x_47); +lean_ctor_set(x_48, 2, x_45); +x_49 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_50 = lean_array_push(x_49, x_43); +x_51 = lean_array_push(x_50, x_48); +x_52 = l_calc___closed__2; +x_53 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_53, 0, x_46); +lean_ctor_set(x_53, 1, x_52); +lean_ctor_set(x_53, 2, x_51); +x_54 = lean_array_push(x_49, x_41); +x_55 = lean_array_push(x_54, x_53); +x_56 = l___aux__Init__NotationExtra______macroRules__tacticCalc____1___closed__6; +x_57 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_57, 0, x_46); +lean_ctor_set(x_57, 1, x_56); +lean_ctor_set(x_57, 2, x_55); +x_58 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_59 = lean_array_push(x_58, x_57); +x_60 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_60, 0, x_46); +lean_ctor_set(x_60, 1, x_47); +lean_ctor_set(x_60, 2, x_59); +x_61 = lean_array_push(x_58, x_60); +x_62 = l___aux__Init__NotationExtra______macroRules__tacticCalc____1___closed__4; +x_63 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_63, 0, x_46); +lean_ctor_set(x_63, 1, x_62); +lean_ctor_set(x_63, 2, x_61); +x_64 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set(x_64, 1, x_39); +return x_64; } } } @@ -6340,22 +6192,22 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; 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; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_5 = lean_ctor_get(x_3, 0); -x_6 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__22; +x_6 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__22; lean_inc(x_5); x_7 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_7, 0, x_5); lean_ctor_set(x_7, 1, x_6); -x_8 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__29; +x_8 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__29; x_9 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_9, 0, x_5); lean_ctor_set(x_9, 1, x_8); -x_10 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; +x_10 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; x_11 = lean_array_push(x_10, x_7); -x_12 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; +x_12 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; x_13 = lean_array_push(x_11, x_12); x_14 = lean_array_push(x_13, x_9); x_15 = lean_box(2); -x_16 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__21; +x_16 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__21; x_17 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_17, 0, x_15); lean_ctor_set(x_17, 1, x_16); @@ -6371,22 +6223,22 @@ x_19 = lean_ctor_get(x_3, 1); lean_inc(x_19); lean_inc(x_18); lean_dec(x_3); -x_20 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__22; +x_20 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__22; lean_inc(x_18); x_21 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_21, 0, x_18); lean_ctor_set(x_21, 1, x_20); -x_22 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__29; +x_22 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__29; x_23 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_23, 0, x_18); lean_ctor_set(x_23, 1, x_22); -x_24 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; +x_24 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; x_25 = lean_array_push(x_24, x_21); -x_26 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; +x_26 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; x_27 = lean_array_push(x_25, x_26); x_28 = lean_array_push(x_27, x_23); x_29 = lean_box(2); -x_30 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__21; +x_30 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__21; x_31 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_31, 0, x_29); lean_ctor_set(x_31, 1, x_30); @@ -6456,13 +6308,13 @@ lean_inc(x_5); x_7 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_7, 0, x_5); lean_ctor_set(x_7, 1, x_6); -x_8 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__13; +x_8 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__13; x_9 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_9, 0, x_5); lean_ctor_set(x_9, 1, x_8); -x_10 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; +x_10 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; x_11 = lean_array_push(x_10, x_7); -x_12 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; +x_12 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; x_13 = lean_array_push(x_11, x_12); x_14 = lean_array_push(x_13, x_9); x_15 = lean_box(2); @@ -6487,13 +6339,13 @@ lean_inc(x_18); x_21 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_21, 0, x_18); lean_ctor_set(x_21, 1, x_20); -x_22 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__13; +x_22 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__13; x_23 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_23, 0, x_18); lean_ctor_set(x_23, 1, x_22); -x_24 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; +x_24 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; x_25 = lean_array_push(x_24, x_21); -x_26 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; +x_26 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; x_27 = lean_array_push(x_25, x_26); x_28 = lean_array_push(x_27, x_23); x_29 = lean_box(2); @@ -6538,7 +6390,7 @@ LEAN_EXPORT lean_object* l_unexpandListCons(lean_object* x_1, lean_object* x_2, _start: { lean_object* x_4; uint8_t x_5; -x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__6; +x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__6; lean_inc(x_1); x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); if (x_5 == 0) @@ -6554,221 +6406,220 @@ return x_7; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; x_8 = lean_unsigned_to_nat(1u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); lean_dec(x_1); -x_10 = l_Lean_nullKind; -x_11 = lean_unsigned_to_nat(2u); +x_10 = lean_unsigned_to_nat(2u); lean_inc(x_9); -x_12 = l_Lean_Syntax_isNodeOf(x_9, x_10, x_11); -if (x_12 == 0) +x_11 = l_Lean_Syntax_matchesNull(x_9, x_10); +if (x_11 == 0) { -lean_object* x_13; lean_object* x_14; +lean_object* x_12; lean_object* x_13; lean_dec(x_9); lean_dec(x_2); -x_13 = lean_box(0); -x_14 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_3); -return x_14; +x_12 = lean_box(0); +x_13 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_3); +return x_13; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_15 = lean_unsigned_to_nat(0u); -x_16 = l_Lean_Syntax_getArg(x_9, x_15); -x_17 = l_Lean_Syntax_getArg(x_9, x_8); +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_14 = lean_unsigned_to_nat(0u); +x_15 = l_Lean_Syntax_getArg(x_9, x_14); +x_16 = l_Lean_Syntax_getArg(x_9, x_8); lean_dec(x_9); -x_18 = l_unexpandListNil___rarg___closed__2; -lean_inc(x_17); -x_19 = l_Lean_Syntax_isOfKind(x_17, x_18); -if (x_19 == 0) +x_17 = l_unexpandListNil___rarg___closed__2; +lean_inc(x_16); +x_18 = l_Lean_Syntax_isOfKind(x_16, x_17); +if (x_18 == 0) { -lean_object* x_20; lean_object* x_21; -lean_dec(x_17); +lean_object* x_19; lean_object* x_20; lean_dec(x_16); +lean_dec(x_15); lean_dec(x_2); -x_20 = lean_box(0); -x_21 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_3); -return x_21; +x_19 = lean_box(0); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_3); +return x_20; } else { -lean_object* x_22; uint8_t x_23; -x_22 = l_Lean_Syntax_getArg(x_17, x_8); -lean_dec(x_17); -lean_inc(x_22); -x_23 = l_Lean_Syntax_isNodeOf(x_22, x_10, x_15); -if (x_23 == 0) +lean_object* x_21; uint8_t x_22; +x_21 = l_Lean_Syntax_getArg(x_16, x_8); +lean_dec(x_16); +lean_inc(x_21); +x_22 = l_Lean_Syntax_matchesNull(x_21, x_14); +if (x_22 == 0) { -lean_object* x_24; lean_object* x_25; uint8_t x_26; -x_24 = l_Lean_Syntax_getArgs(x_22); -lean_dec(x_22); -x_25 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); -x_26 = !lean_is_exclusive(x_25); -if (x_26 == 0) +lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_23 = l_Lean_Syntax_getArgs(x_21); +lean_dec(x_21); +x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); +x_25 = !lean_is_exclusive(x_24); +if (x_25 == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; 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; lean_object* x_44; lean_object* x_45; -x_27 = lean_ctor_get(x_25, 0); -x_28 = l_unexpandListNil___rarg___closed__3; -lean_inc(x_27); -x_29 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -x_30 = l_unexpandListCons___closed__1; -lean_inc(x_27); -x_31 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_31, 0, x_27); -lean_ctor_set(x_31, 1, x_30); -x_32 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_33 = lean_array_push(x_32, x_16); -x_34 = lean_array_push(x_33, x_31); -x_35 = l_Array_append___rarg(x_34, x_24); -x_36 = lean_box(2); -x_37 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_38 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -lean_ctor_set(x_38, 2, x_35); -x_39 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__13; -x_40 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_40, 0, x_27); -lean_ctor_set(x_40, 1, x_39); -x_41 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; -x_42 = lean_array_push(x_41, x_29); -x_43 = lean_array_push(x_42, x_38); -x_44 = lean_array_push(x_43, x_40); -x_45 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_45, 0, x_36); -lean_ctor_set(x_45, 1, x_18); -lean_ctor_set(x_45, 2, x_44); -lean_ctor_set(x_25, 0, x_45); -return 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_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; lean_object* x_44; +x_26 = lean_ctor_get(x_24, 0); +x_27 = l_unexpandListNil___rarg___closed__3; +lean_inc(x_26); +x_28 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +x_29 = l_unexpandListCons___closed__1; +lean_inc(x_26); +x_30 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_30, 0, x_26); +lean_ctor_set(x_30, 1, x_29); +x_31 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_32 = lean_array_push(x_31, x_15); +x_33 = lean_array_push(x_32, x_30); +x_34 = l_Array_append___rarg(x_33, x_23); +x_35 = lean_box(2); +x_36 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_37 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +lean_ctor_set(x_37, 2, x_34); +x_38 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__13; +x_39 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_39, 0, x_26); +lean_ctor_set(x_39, 1, x_38); +x_40 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; +x_41 = lean_array_push(x_40, x_28); +x_42 = lean_array_push(x_41, x_37); +x_43 = lean_array_push(x_42, x_39); +x_44 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_44, 0, x_35); +lean_ctor_set(x_44, 1, x_17); +lean_ctor_set(x_44, 2, x_43); +lean_ctor_set(x_24, 0, x_44); +return x_24; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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_46 = lean_ctor_get(x_25, 0); -x_47 = lean_ctor_get(x_25, 1); -lean_inc(x_47); -lean_inc(x_46); -lean_dec(x_25); -x_48 = l_unexpandListNil___rarg___closed__3; -lean_inc(x_46); -x_49 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_49, 0, x_46); -lean_ctor_set(x_49, 1, x_48); -x_50 = l_unexpandListCons___closed__1; +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_45 = lean_ctor_get(x_24, 0); +x_46 = lean_ctor_get(x_24, 1); lean_inc(x_46); -x_51 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_51, 0, x_46); -lean_ctor_set(x_51, 1, x_50); -x_52 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_53 = lean_array_push(x_52, x_16); -x_54 = lean_array_push(x_53, x_51); -x_55 = l_Array_append___rarg(x_54, x_24); -x_56 = lean_box(2); -x_57 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_58 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_58, 0, x_56); -lean_ctor_set(x_58, 1, x_57); -lean_ctor_set(x_58, 2, x_55); -x_59 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__13; -x_60 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_60, 0, x_46); -lean_ctor_set(x_60, 1, x_59); -x_61 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; -x_62 = lean_array_push(x_61, x_49); -x_63 = lean_array_push(x_62, x_58); -x_64 = lean_array_push(x_63, x_60); -x_65 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_65, 0, x_56); -lean_ctor_set(x_65, 1, x_18); -lean_ctor_set(x_65, 2, x_64); -x_66 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_66, 0, x_65); -lean_ctor_set(x_66, 1, x_47); -return x_66; +lean_inc(x_45); +lean_dec(x_24); +x_47 = l_unexpandListNil___rarg___closed__3; +lean_inc(x_45); +x_48 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_48, 0, x_45); +lean_ctor_set(x_48, 1, x_47); +x_49 = l_unexpandListCons___closed__1; +lean_inc(x_45); +x_50 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_50, 0, x_45); +lean_ctor_set(x_50, 1, x_49); +x_51 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_52 = lean_array_push(x_51, x_15); +x_53 = lean_array_push(x_52, x_50); +x_54 = l_Array_append___rarg(x_53, x_23); +x_55 = lean_box(2); +x_56 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_57 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_56); +lean_ctor_set(x_57, 2, x_54); +x_58 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__13; +x_59 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_59, 0, x_45); +lean_ctor_set(x_59, 1, x_58); +x_60 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; +x_61 = lean_array_push(x_60, x_48); +x_62 = lean_array_push(x_61, x_57); +x_63 = lean_array_push(x_62, x_59); +x_64 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_64, 0, x_55); +lean_ctor_set(x_64, 1, x_17); +lean_ctor_set(x_64, 2, x_63); +x_65 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_65, 0, x_64); +lean_ctor_set(x_65, 1, x_46); +return x_65; } } else { -lean_object* x_67; uint8_t x_68; -lean_dec(x_22); -x_67 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); -x_68 = !lean_is_exclusive(x_67); -if (x_68 == 0) +lean_object* x_66; uint8_t x_67; +lean_dec(x_21); +x_66 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); +x_67 = !lean_is_exclusive(x_66); +if (x_67 == 0) { -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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; -x_69 = lean_ctor_get(x_67, 0); -x_70 = l_unexpandListNil___rarg___closed__3; -lean_inc(x_69); -x_71 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_71, 0, x_69); -lean_ctor_set(x_71, 1, x_70); -x_72 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_73 = lean_array_push(x_72, x_16); -x_74 = lean_box(2); -x_75 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_76 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_76, 0, x_74); -lean_ctor_set(x_76, 1, x_75); -lean_ctor_set(x_76, 2, x_73); -x_77 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__13; -x_78 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_78, 0, x_69); -lean_ctor_set(x_78, 1, x_77); -x_79 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; -x_80 = lean_array_push(x_79, x_71); -x_81 = lean_array_push(x_80, x_76); -x_82 = lean_array_push(x_81, x_78); -x_83 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_83, 0, x_74); -lean_ctor_set(x_83, 1, x_18); -lean_ctor_set(x_83, 2, x_82); -lean_ctor_set(x_67, 0, x_83); -return 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_68 = lean_ctor_get(x_66, 0); +x_69 = l_unexpandListNil___rarg___closed__3; +lean_inc(x_68); +x_70 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_70, 0, x_68); +lean_ctor_set(x_70, 1, x_69); +x_71 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_72 = lean_array_push(x_71, x_15); +x_73 = lean_box(2); +x_74 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_75 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set(x_75, 1, x_74); +lean_ctor_set(x_75, 2, x_72); +x_76 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__13; +x_77 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_77, 0, x_68); +lean_ctor_set(x_77, 1, x_76); +x_78 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; +x_79 = lean_array_push(x_78, x_70); +x_80 = lean_array_push(x_79, x_75); +x_81 = lean_array_push(x_80, x_77); +x_82 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_82, 0, x_73); +lean_ctor_set(x_82, 1, x_17); +lean_ctor_set(x_82, 2, x_81); +lean_ctor_set(x_66, 0, x_82); +return x_66; } else { -lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_84 = lean_ctor_get(x_67, 0); -x_85 = lean_ctor_get(x_67, 1); -lean_inc(x_85); -lean_inc(x_84); -lean_dec(x_67); -x_86 = l_unexpandListNil___rarg___closed__3; +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_83 = lean_ctor_get(x_66, 0); +x_84 = lean_ctor_get(x_66, 1); lean_inc(x_84); -x_87 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_87, 0, x_84); -lean_ctor_set(x_87, 1, x_86); -x_88 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_89 = lean_array_push(x_88, x_16); -x_90 = lean_box(2); -x_91 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_92 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_92, 0, x_90); -lean_ctor_set(x_92, 1, x_91); -lean_ctor_set(x_92, 2, x_89); -x_93 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__13; -x_94 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_94, 0, x_84); -lean_ctor_set(x_94, 1, x_93); -x_95 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; -x_96 = lean_array_push(x_95, x_87); -x_97 = lean_array_push(x_96, x_92); -x_98 = lean_array_push(x_97, x_94); -x_99 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_99, 0, x_90); -lean_ctor_set(x_99, 1, x_18); -lean_ctor_set(x_99, 2, x_98); -x_100 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_100, 0, x_99); -lean_ctor_set(x_100, 1, x_85); -return x_100; +lean_inc(x_83); +lean_dec(x_66); +x_85 = l_unexpandListNil___rarg___closed__3; +lean_inc(x_83); +x_86 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_86, 0, x_83); +lean_ctor_set(x_86, 1, x_85); +x_87 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_88 = lean_array_push(x_87, x_15); +x_89 = lean_box(2); +x_90 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_91 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_91, 0, x_89); +lean_ctor_set(x_91, 1, x_90); +lean_ctor_set(x_91, 2, x_88); +x_92 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__13; +x_93 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_93, 0, x_83); +lean_ctor_set(x_93, 1, x_92); +x_94 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; +x_95 = lean_array_push(x_94, x_86); +x_96 = lean_array_push(x_95, x_91); +x_97 = lean_array_push(x_96, x_93); +x_98 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_98, 0, x_89); +lean_ctor_set(x_98, 1, x_17); +lean_ctor_set(x_98, 2, x_97); +x_99 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_99, 0, x_98); +lean_ctor_set(x_99, 1, x_84); +return x_99; } } } @@ -6806,7 +6657,7 @@ LEAN_EXPORT lean_object* l_unexpandListToArray(lean_object* x_1, lean_object* x_ _start: { lean_object* x_4; uint8_t x_5; -x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__6; +x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__6; lean_inc(x_1); x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); if (x_5 == 0) @@ -6822,124 +6673,123 @@ return x_7; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +lean_object* x_8; lean_object* x_9; uint8_t x_10; x_8 = lean_unsigned_to_nat(1u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); lean_dec(x_1); -x_10 = l_Lean_nullKind; lean_inc(x_9); -x_11 = l_Lean_Syntax_isNodeOf(x_9, x_10, x_8); -if (x_11 == 0) +x_10 = l_Lean_Syntax_matchesNull(x_9, x_8); +if (x_10 == 0) { -lean_object* x_12; lean_object* x_13; +lean_object* x_11; lean_object* x_12; lean_dec(x_9); lean_dec(x_2); -x_12 = lean_box(0); -x_13 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_3); -return x_13; +x_11 = lean_box(0); +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_3); +return x_12; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_14 = lean_unsigned_to_nat(0u); -x_15 = l_Lean_Syntax_getArg(x_9, x_14); +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_13 = lean_unsigned_to_nat(0u); +x_14 = l_Lean_Syntax_getArg(x_9, x_13); lean_dec(x_9); -x_16 = l_unexpandListNil___rarg___closed__2; -lean_inc(x_15); -x_17 = l_Lean_Syntax_isOfKind(x_15, x_16); -if (x_17 == 0) +x_15 = l_unexpandListNil___rarg___closed__2; +lean_inc(x_14); +x_16 = l_Lean_Syntax_isOfKind(x_14, x_15); +if (x_16 == 0) { -lean_object* x_18; lean_object* x_19; -lean_dec(x_15); +lean_object* x_17; lean_object* x_18; +lean_dec(x_14); lean_dec(x_2); -x_18 = lean_box(0); -x_19 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_3); -return x_19; +x_17 = lean_box(0); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_3); +return x_18; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_20 = l_Lean_Syntax_getArg(x_15, x_8); -lean_dec(x_15); -x_21 = l_Lean_Syntax_getArgs(x_20); -lean_dec(x_20); -x_22 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); -x_23 = !lean_is_exclusive(x_22); -if (x_23 == 0) +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_19 = l_Lean_Syntax_getArg(x_14, x_8); +lean_dec(x_14); +x_20 = l_Lean_Syntax_getArgs(x_19); +lean_dec(x_19); +x_21 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) { -lean_object* 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; 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; -x_24 = lean_ctor_get(x_22, 0); -x_25 = l_unexpandListToArray___closed__3; -lean_inc(x_24); -x_26 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -x_27 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__27; -x_28 = l_Array_append___rarg(x_27, x_21); -x_29 = lean_box(2); -x_30 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_31 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_30); -lean_ctor_set(x_31, 2, x_28); -x_32 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__13; -x_33 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_33, 0, x_24); -lean_ctor_set(x_33, 1, x_32); -x_34 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; -x_35 = lean_array_push(x_34, x_26); -x_36 = lean_array_push(x_35, x_31); -x_37 = lean_array_push(x_36, x_33); -x_38 = l_unexpandListToArray___closed__2; -x_39 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_39, 0, x_29); -lean_ctor_set(x_39, 1, x_38); -lean_ctor_set(x_39, 2, x_37); -lean_ctor_set(x_22, 0, x_39); -return 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; 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; +x_23 = lean_ctor_get(x_21, 0); +x_24 = l_unexpandListToArray___closed__3; +lean_inc(x_23); +x_25 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +x_26 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__27; +x_27 = l_Array_append___rarg(x_26, x_20); +x_28 = lean_box(2); +x_29 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_30 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +lean_ctor_set(x_30, 2, x_27); +x_31 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__13; +x_32 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_32, 0, x_23); +lean_ctor_set(x_32, 1, x_31); +x_33 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; +x_34 = lean_array_push(x_33, x_25); +x_35 = lean_array_push(x_34, x_30); +x_36 = lean_array_push(x_35, x_32); +x_37 = l_unexpandListToArray___closed__2; +x_38 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_38, 0, x_28); +lean_ctor_set(x_38, 1, x_37); +lean_ctor_set(x_38, 2, x_36); +lean_ctor_set(x_21, 0, x_38); +return x_21; } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_40 = lean_ctor_get(x_22, 0); -x_41 = lean_ctor_get(x_22, 1); -lean_inc(x_41); -lean_inc(x_40); -lean_dec(x_22); -x_42 = l_unexpandListToArray___closed__3; +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_39 = lean_ctor_get(x_21, 0); +x_40 = lean_ctor_get(x_21, 1); lean_inc(x_40); -x_43 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_43, 0, x_40); -lean_ctor_set(x_43, 1, x_42); -x_44 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__27; -x_45 = l_Array_append___rarg(x_44, x_21); -x_46 = lean_box(2); -x_47 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_48 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_48, 0, x_46); -lean_ctor_set(x_48, 1, x_47); -lean_ctor_set(x_48, 2, x_45); -x_49 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__13; -x_50 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_50, 0, x_40); -lean_ctor_set(x_50, 1, x_49); -x_51 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; -x_52 = lean_array_push(x_51, x_43); -x_53 = lean_array_push(x_52, x_48); -x_54 = lean_array_push(x_53, x_50); -x_55 = l_unexpandListToArray___closed__2; -x_56 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_56, 0, x_46); -lean_ctor_set(x_56, 1, x_55); -lean_ctor_set(x_56, 2, x_54); -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_41); -return x_57; +lean_inc(x_39); +lean_dec(x_21); +x_41 = l_unexpandListToArray___closed__3; +lean_inc(x_39); +x_42 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_42, 0, x_39); +lean_ctor_set(x_42, 1, x_41); +x_43 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__27; +x_44 = l_Array_append___rarg(x_43, x_20); +x_45 = lean_box(2); +x_46 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_47 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +lean_ctor_set(x_47, 2, x_44); +x_48 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__13; +x_49 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_49, 0, x_39); +lean_ctor_set(x_49, 1, x_48); +x_50 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; +x_51 = lean_array_push(x_50, x_42); +x_52 = lean_array_push(x_51, x_47); +x_53 = lean_array_push(x_52, x_49); +x_54 = l_unexpandListToArray___closed__2; +x_55 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_55, 0, x_45); +lean_ctor_set(x_55, 1, x_54); +lean_ctor_set(x_55, 2, x_53); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_40); +return x_56; } } } @@ -6958,7 +6808,7 @@ static lean_object* _init_l_unexpandProdMk___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4; x_2 = l_unexpandProdMk___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -6968,7 +6818,7 @@ LEAN_EXPORT lean_object* l_unexpandProdMk(lean_object* x_1, lean_object* x_2, le _start: { lean_object* x_4; uint8_t x_5; -x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__6; +x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__6; lean_inc(x_1); x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); if (x_5 == 0) @@ -6984,687 +6834,686 @@ return x_7; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; x_8 = lean_unsigned_to_nat(1u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); lean_dec(x_1); -x_10 = l_Lean_nullKind; -x_11 = lean_unsigned_to_nat(2u); +x_10 = lean_unsigned_to_nat(2u); lean_inc(x_9); -x_12 = l_Lean_Syntax_isNodeOf(x_9, x_10, x_11); -if (x_12 == 0) +x_11 = l_Lean_Syntax_matchesNull(x_9, x_10); +if (x_11 == 0) { -lean_object* x_13; lean_object* x_14; +lean_object* x_12; lean_object* x_13; lean_dec(x_9); lean_dec(x_2); -x_13 = lean_box(0); -x_14 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_3); -return x_14; +x_12 = lean_box(0); +x_13 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_3); +return x_13; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_15 = lean_unsigned_to_nat(0u); -x_16 = l_Lean_Syntax_getArg(x_9, x_15); -x_17 = l_Lean_Syntax_getArg(x_9, x_8); +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_14 = lean_unsigned_to_nat(0u); +x_15 = l_Lean_Syntax_getArg(x_9, x_14); +x_16 = l_Lean_Syntax_getArg(x_9, x_8); lean_dec(x_9); -x_18 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__21; -lean_inc(x_17); -x_19 = l_Lean_Syntax_isOfKind(x_17, x_18); -if (x_19 == 0) +x_17 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__21; +lean_inc(x_16); +x_18 = l_Lean_Syntax_isOfKind(x_16, x_17); +if (x_18 == 0) { -lean_object* x_20; uint8_t x_21; -x_20 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); -x_21 = !lean_is_exclusive(x_20); -if (x_21 == 0) +lean_object* x_19; uint8_t x_20; +x_19 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) { -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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_22 = lean_ctor_get(x_20, 0); -x_23 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__22; -lean_inc(x_22); -x_24 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_23); -x_25 = l_unexpandListCons___closed__1; -lean_inc(x_22); -x_26 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_26, 0, x_22); -lean_ctor_set(x_26, 1, x_25); -x_27 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_28 = lean_array_push(x_27, x_17); -x_29 = lean_box(2); -x_30 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_31 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_30); -lean_ctor_set(x_31, 2, x_28); -x_32 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_33 = lean_array_push(x_32, x_26); -x_34 = lean_array_push(x_33, x_31); -x_35 = l_unexpandProdMk___closed__2; -x_36 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_36, 0, x_29); -lean_ctor_set(x_36, 1, x_35); -lean_ctor_set(x_36, 2, x_34); -x_37 = lean_array_push(x_27, x_36); -x_38 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_38, 0, x_29); -lean_ctor_set(x_38, 1, x_30); -lean_ctor_set(x_38, 2, x_37); -x_39 = lean_array_push(x_32, x_16); -x_40 = lean_array_push(x_39, x_38); -x_41 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_41, 0, x_29); -lean_ctor_set(x_41, 1, x_30); -lean_ctor_set(x_41, 2, x_40); -x_42 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__29; -x_43 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_43, 0, x_22); -lean_ctor_set(x_43, 1, x_42); -x_44 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; -x_45 = lean_array_push(x_44, x_24); -x_46 = lean_array_push(x_45, x_41); -x_47 = lean_array_push(x_46, x_43); -x_48 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_48, 0, x_29); -lean_ctor_set(x_48, 1, x_18); -lean_ctor_set(x_48, 2, x_47); -lean_ctor_set(x_20, 0, x_48); -return 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_21 = lean_ctor_get(x_19, 0); +x_22 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__22; +lean_inc(x_21); +x_23 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +x_24 = l_unexpandListCons___closed__1; +lean_inc(x_21); +x_25 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_25, 0, x_21); +lean_ctor_set(x_25, 1, x_24); +x_26 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_27 = lean_array_push(x_26, x_16); +x_28 = lean_box(2); +x_29 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_30 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +lean_ctor_set(x_30, 2, x_27); +x_31 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_32 = lean_array_push(x_31, x_25); +x_33 = lean_array_push(x_32, x_30); +x_34 = l_unexpandProdMk___closed__2; +x_35 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_35, 0, x_28); +lean_ctor_set(x_35, 1, x_34); +lean_ctor_set(x_35, 2, x_33); +x_36 = lean_array_push(x_26, x_35); +x_37 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_37, 0, x_28); +lean_ctor_set(x_37, 1, x_29); +lean_ctor_set(x_37, 2, x_36); +x_38 = lean_array_push(x_31, x_15); +x_39 = lean_array_push(x_38, x_37); +x_40 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_40, 0, x_28); +lean_ctor_set(x_40, 1, x_29); +lean_ctor_set(x_40, 2, x_39); +x_41 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__29; +x_42 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_42, 0, x_21); +lean_ctor_set(x_42, 1, x_41); +x_43 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; +x_44 = lean_array_push(x_43, x_23); +x_45 = lean_array_push(x_44, x_40); +x_46 = lean_array_push(x_45, x_42); +x_47 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_47, 0, x_28); +lean_ctor_set(x_47, 1, x_17); +lean_ctor_set(x_47, 2, x_46); +lean_ctor_set(x_19, 0, x_47); +return x_19; } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; -x_49 = lean_ctor_get(x_20, 0); -x_50 = lean_ctor_get(x_20, 1); -lean_inc(x_50); -lean_inc(x_49); -lean_dec(x_20); -x_51 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__22; +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; +x_48 = lean_ctor_get(x_19, 0); +x_49 = lean_ctor_get(x_19, 1); lean_inc(x_49); -x_52 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_52, 0, x_49); -lean_ctor_set(x_52, 1, x_51); -x_53 = l_unexpandListCons___closed__1; -lean_inc(x_49); -x_54 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_54, 0, x_49); -lean_ctor_set(x_54, 1, x_53); -x_55 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_56 = lean_array_push(x_55, x_17); -x_57 = lean_box(2); -x_58 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_59 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_59, 0, x_57); -lean_ctor_set(x_59, 1, x_58); -lean_ctor_set(x_59, 2, x_56); -x_60 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_61 = lean_array_push(x_60, x_54); -x_62 = lean_array_push(x_61, x_59); -x_63 = l_unexpandProdMk___closed__2; -x_64 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_64, 0, x_57); -lean_ctor_set(x_64, 1, x_63); -lean_ctor_set(x_64, 2, x_62); -x_65 = lean_array_push(x_55, x_64); -x_66 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_66, 0, x_57); -lean_ctor_set(x_66, 1, x_58); -lean_ctor_set(x_66, 2, x_65); -x_67 = lean_array_push(x_60, x_16); -x_68 = lean_array_push(x_67, x_66); -x_69 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_69, 0, x_57); -lean_ctor_set(x_69, 1, x_58); -lean_ctor_set(x_69, 2, x_68); -x_70 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__29; -x_71 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_71, 0, x_49); -lean_ctor_set(x_71, 1, x_70); -x_72 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; -x_73 = lean_array_push(x_72, x_52); -x_74 = lean_array_push(x_73, x_69); -x_75 = lean_array_push(x_74, x_71); -x_76 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_76, 0, x_57); -lean_ctor_set(x_76, 1, x_18); -lean_ctor_set(x_76, 2, x_75); -x_77 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_77, 0, x_76); -lean_ctor_set(x_77, 1, x_50); -return x_77; +lean_inc(x_48); +lean_dec(x_19); +x_50 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__22; +lean_inc(x_48); +x_51 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_51, 0, x_48); +lean_ctor_set(x_51, 1, x_50); +x_52 = l_unexpandListCons___closed__1; +lean_inc(x_48); +x_53 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_53, 0, x_48); +lean_ctor_set(x_53, 1, x_52); +x_54 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_55 = lean_array_push(x_54, x_16); +x_56 = lean_box(2); +x_57 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_58 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_58, 0, x_56); +lean_ctor_set(x_58, 1, x_57); +lean_ctor_set(x_58, 2, x_55); +x_59 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_60 = lean_array_push(x_59, x_53); +x_61 = lean_array_push(x_60, x_58); +x_62 = l_unexpandProdMk___closed__2; +x_63 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_63, 0, x_56); +lean_ctor_set(x_63, 1, x_62); +lean_ctor_set(x_63, 2, x_61); +x_64 = lean_array_push(x_54, x_63); +x_65 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_65, 0, x_56); +lean_ctor_set(x_65, 1, x_57); +lean_ctor_set(x_65, 2, x_64); +x_66 = lean_array_push(x_59, x_15); +x_67 = lean_array_push(x_66, x_65); +x_68 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_68, 0, x_56); +lean_ctor_set(x_68, 1, x_57); +lean_ctor_set(x_68, 2, x_67); +x_69 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__29; +x_70 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_70, 0, x_48); +lean_ctor_set(x_70, 1, x_69); +x_71 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; +x_72 = lean_array_push(x_71, x_51); +x_73 = lean_array_push(x_72, x_68); +x_74 = lean_array_push(x_73, x_70); +x_75 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_75, 0, x_56); +lean_ctor_set(x_75, 1, x_17); +lean_ctor_set(x_75, 2, x_74); +x_76 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_76, 0, x_75); +lean_ctor_set(x_76, 1, x_49); +return x_76; } } else { -lean_object* x_78; uint8_t x_79; -x_78 = l_Lean_Syntax_getArg(x_17, x_8); -lean_inc(x_78); -x_79 = l_Lean_Syntax_isNodeOf(x_78, x_10, x_11); -if (x_79 == 0) +lean_object* x_77; uint8_t x_78; +x_77 = l_Lean_Syntax_getArg(x_16, x_8); +lean_inc(x_77); +x_78 = l_Lean_Syntax_matchesNull(x_77, x_10); +if (x_78 == 0) { -lean_object* x_80; uint8_t x_81; -lean_dec(x_78); -x_80 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); -x_81 = !lean_is_exclusive(x_80); -if (x_81 == 0) +lean_object* x_79; uint8_t x_80; +lean_dec(x_77); +x_79 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); +x_80 = !lean_is_exclusive(x_79); +if (x_80 == 0) { -lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; -x_82 = lean_ctor_get(x_80, 0); -x_83 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__22; -lean_inc(x_82); -x_84 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_84, 0, x_82); -lean_ctor_set(x_84, 1, x_83); -x_85 = l_unexpandListCons___closed__1; -lean_inc(x_82); -x_86 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_86, 0, x_82); -lean_ctor_set(x_86, 1, x_85); -x_87 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_88 = lean_array_push(x_87, x_17); -x_89 = lean_box(2); -x_90 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_91 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_91, 0, x_89); -lean_ctor_set(x_91, 1, x_90); -lean_ctor_set(x_91, 2, x_88); -x_92 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_93 = lean_array_push(x_92, x_86); -x_94 = lean_array_push(x_93, x_91); -x_95 = l_unexpandProdMk___closed__2; -x_96 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_96, 0, x_89); -lean_ctor_set(x_96, 1, x_95); -lean_ctor_set(x_96, 2, x_94); -x_97 = lean_array_push(x_87, x_96); -x_98 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_98, 0, x_89); -lean_ctor_set(x_98, 1, x_90); -lean_ctor_set(x_98, 2, x_97); -x_99 = lean_array_push(x_92, x_16); -x_100 = lean_array_push(x_99, x_98); -x_101 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_101, 0, x_89); -lean_ctor_set(x_101, 1, x_90); -lean_ctor_set(x_101, 2, x_100); -x_102 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__29; -x_103 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_103, 0, x_82); -lean_ctor_set(x_103, 1, x_102); -x_104 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; -x_105 = lean_array_push(x_104, x_84); -x_106 = lean_array_push(x_105, x_101); -x_107 = lean_array_push(x_106, x_103); -x_108 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_108, 0, x_89); -lean_ctor_set(x_108, 1, x_18); -lean_ctor_set(x_108, 2, x_107); -lean_ctor_set(x_80, 0, x_108); -return x_80; +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_81 = lean_ctor_get(x_79, 0); +x_82 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__22; +lean_inc(x_81); +x_83 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_83, 0, x_81); +lean_ctor_set(x_83, 1, x_82); +x_84 = l_unexpandListCons___closed__1; +lean_inc(x_81); +x_85 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_85, 0, x_81); +lean_ctor_set(x_85, 1, x_84); +x_86 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_87 = lean_array_push(x_86, x_16); +x_88 = lean_box(2); +x_89 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_90 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_90, 0, x_88); +lean_ctor_set(x_90, 1, x_89); +lean_ctor_set(x_90, 2, x_87); +x_91 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_92 = lean_array_push(x_91, x_85); +x_93 = lean_array_push(x_92, x_90); +x_94 = l_unexpandProdMk___closed__2; +x_95 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_95, 0, x_88); +lean_ctor_set(x_95, 1, x_94); +lean_ctor_set(x_95, 2, x_93); +x_96 = lean_array_push(x_86, x_95); +x_97 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_97, 0, x_88); +lean_ctor_set(x_97, 1, x_89); +lean_ctor_set(x_97, 2, x_96); +x_98 = lean_array_push(x_91, x_15); +x_99 = lean_array_push(x_98, x_97); +x_100 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_100, 0, x_88); +lean_ctor_set(x_100, 1, x_89); +lean_ctor_set(x_100, 2, x_99); +x_101 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__29; +x_102 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_102, 0, x_81); +lean_ctor_set(x_102, 1, x_101); +x_103 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; +x_104 = lean_array_push(x_103, x_83); +x_105 = lean_array_push(x_104, x_100); +x_106 = lean_array_push(x_105, x_102); +x_107 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_107, 0, x_88); +lean_ctor_set(x_107, 1, x_17); +lean_ctor_set(x_107, 2, x_106); +lean_ctor_set(x_79, 0, x_107); +return x_79; } else { -lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; -x_109 = lean_ctor_get(x_80, 0); -x_110 = lean_ctor_get(x_80, 1); -lean_inc(x_110); -lean_inc(x_109); -lean_dec(x_80); -x_111 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__22; -lean_inc(x_109); -x_112 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_112, 0, x_109); -lean_ctor_set(x_112, 1, x_111); -x_113 = l_unexpandListCons___closed__1; +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; +x_108 = lean_ctor_get(x_79, 0); +x_109 = lean_ctor_get(x_79, 1); lean_inc(x_109); -x_114 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_114, 0, x_109); -lean_ctor_set(x_114, 1, x_113); -x_115 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_116 = lean_array_push(x_115, x_17); -x_117 = lean_box(2); -x_118 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_119 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_119, 0, x_117); -lean_ctor_set(x_119, 1, x_118); -lean_ctor_set(x_119, 2, x_116); -x_120 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_121 = lean_array_push(x_120, x_114); -x_122 = lean_array_push(x_121, x_119); -x_123 = l_unexpandProdMk___closed__2; -x_124 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_124, 0, x_117); -lean_ctor_set(x_124, 1, x_123); -lean_ctor_set(x_124, 2, x_122); -x_125 = lean_array_push(x_115, x_124); -x_126 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_126, 0, x_117); -lean_ctor_set(x_126, 1, x_118); -lean_ctor_set(x_126, 2, x_125); -x_127 = lean_array_push(x_120, x_16); -x_128 = lean_array_push(x_127, x_126); -x_129 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_129, 0, x_117); -lean_ctor_set(x_129, 1, x_118); -lean_ctor_set(x_129, 2, x_128); -x_130 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__29; -x_131 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_131, 0, x_109); -lean_ctor_set(x_131, 1, x_130); -x_132 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; -x_133 = lean_array_push(x_132, x_112); -x_134 = lean_array_push(x_133, x_129); -x_135 = lean_array_push(x_134, x_131); -x_136 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_136, 0, x_117); -lean_ctor_set(x_136, 1, x_18); -lean_ctor_set(x_136, 2, x_135); -x_137 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_137, 0, x_136); -lean_ctor_set(x_137, 1, x_110); -return x_137; +lean_inc(x_108); +lean_dec(x_79); +x_110 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__22; +lean_inc(x_108); +x_111 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_110); +x_112 = l_unexpandListCons___closed__1; +lean_inc(x_108); +x_113 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_113, 0, x_108); +lean_ctor_set(x_113, 1, x_112); +x_114 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_115 = lean_array_push(x_114, x_16); +x_116 = lean_box(2); +x_117 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_118 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_118, 0, x_116); +lean_ctor_set(x_118, 1, x_117); +lean_ctor_set(x_118, 2, x_115); +x_119 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_120 = lean_array_push(x_119, x_113); +x_121 = lean_array_push(x_120, x_118); +x_122 = l_unexpandProdMk___closed__2; +x_123 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_123, 0, x_116); +lean_ctor_set(x_123, 1, x_122); +lean_ctor_set(x_123, 2, x_121); +x_124 = lean_array_push(x_114, x_123); +x_125 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_125, 0, x_116); +lean_ctor_set(x_125, 1, x_117); +lean_ctor_set(x_125, 2, x_124); +x_126 = lean_array_push(x_119, x_15); +x_127 = lean_array_push(x_126, x_125); +x_128 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_128, 0, x_116); +lean_ctor_set(x_128, 1, x_117); +lean_ctor_set(x_128, 2, x_127); +x_129 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__29; +x_130 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_130, 0, x_108); +lean_ctor_set(x_130, 1, x_129); +x_131 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; +x_132 = lean_array_push(x_131, x_111); +x_133 = lean_array_push(x_132, x_128); +x_134 = lean_array_push(x_133, x_130); +x_135 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_135, 0, x_116); +lean_ctor_set(x_135, 1, x_17); +lean_ctor_set(x_135, 2, x_134); +x_136 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_136, 0, x_135); +lean_ctor_set(x_136, 1, x_109); +return x_136; } } else { -lean_object* x_138; lean_object* x_139; uint8_t x_140; -x_138 = l_Lean_Syntax_getArg(x_78, x_15); -x_139 = l_Lean_Syntax_getArg(x_78, x_8); -lean_dec(x_78); -lean_inc(x_139); -x_140 = l_Lean_Syntax_isNodeOf(x_139, x_10, x_8); -if (x_140 == 0) +lean_object* x_137; lean_object* x_138; uint8_t x_139; +x_137 = l_Lean_Syntax_getArg(x_77, x_14); +x_138 = l_Lean_Syntax_getArg(x_77, x_8); +lean_dec(x_77); +lean_inc(x_138); +x_139 = l_Lean_Syntax_matchesNull(x_138, x_8); +if (x_139 == 0) { -lean_object* x_141; uint8_t x_142; -lean_dec(x_139); +lean_object* x_140; uint8_t x_141; lean_dec(x_138); -x_141 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); -x_142 = !lean_is_exclusive(x_141); -if (x_142 == 0) +lean_dec(x_137); +x_140 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); +x_141 = !lean_is_exclusive(x_140); +if (x_141 == 0) { -lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; -x_143 = lean_ctor_get(x_141, 0); -x_144 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__22; -lean_inc(x_143); -x_145 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_145, 0, x_143); -lean_ctor_set(x_145, 1, x_144); -x_146 = l_unexpandListCons___closed__1; -lean_inc(x_143); -x_147 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_147, 0, x_143); -lean_ctor_set(x_147, 1, x_146); -x_148 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_149 = lean_array_push(x_148, x_17); -x_150 = lean_box(2); -x_151 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_152 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_152, 0, x_150); -lean_ctor_set(x_152, 1, x_151); -lean_ctor_set(x_152, 2, x_149); -x_153 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_154 = lean_array_push(x_153, x_147); -x_155 = lean_array_push(x_154, x_152); -x_156 = l_unexpandProdMk___closed__2; -x_157 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_157, 0, x_150); -lean_ctor_set(x_157, 1, x_156); -lean_ctor_set(x_157, 2, x_155); -x_158 = lean_array_push(x_148, x_157); -x_159 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_159, 0, x_150); -lean_ctor_set(x_159, 1, x_151); -lean_ctor_set(x_159, 2, x_158); -x_160 = lean_array_push(x_153, x_16); -x_161 = lean_array_push(x_160, x_159); -x_162 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_162, 0, x_150); -lean_ctor_set(x_162, 1, x_151); -lean_ctor_set(x_162, 2, x_161); -x_163 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__29; -x_164 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_164, 0, x_143); -lean_ctor_set(x_164, 1, x_163); -x_165 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; -x_166 = lean_array_push(x_165, x_145); -x_167 = lean_array_push(x_166, x_162); -x_168 = lean_array_push(x_167, x_164); -x_169 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_169, 0, x_150); -lean_ctor_set(x_169, 1, x_18); -lean_ctor_set(x_169, 2, x_168); -lean_ctor_set(x_141, 0, x_169); -return x_141; +lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; +x_142 = lean_ctor_get(x_140, 0); +x_143 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__22; +lean_inc(x_142); +x_144 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_144, 0, x_142); +lean_ctor_set(x_144, 1, x_143); +x_145 = l_unexpandListCons___closed__1; +lean_inc(x_142); +x_146 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_146, 0, x_142); +lean_ctor_set(x_146, 1, x_145); +x_147 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_148 = lean_array_push(x_147, x_16); +x_149 = lean_box(2); +x_150 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_151 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_151, 0, x_149); +lean_ctor_set(x_151, 1, x_150); +lean_ctor_set(x_151, 2, x_148); +x_152 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_153 = lean_array_push(x_152, x_146); +x_154 = lean_array_push(x_153, x_151); +x_155 = l_unexpandProdMk___closed__2; +x_156 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_156, 0, x_149); +lean_ctor_set(x_156, 1, x_155); +lean_ctor_set(x_156, 2, x_154); +x_157 = lean_array_push(x_147, x_156); +x_158 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_158, 0, x_149); +lean_ctor_set(x_158, 1, x_150); +lean_ctor_set(x_158, 2, x_157); +x_159 = lean_array_push(x_152, x_15); +x_160 = lean_array_push(x_159, x_158); +x_161 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_161, 0, x_149); +lean_ctor_set(x_161, 1, x_150); +lean_ctor_set(x_161, 2, x_160); +x_162 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__29; +x_163 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_163, 0, x_142); +lean_ctor_set(x_163, 1, x_162); +x_164 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; +x_165 = lean_array_push(x_164, x_144); +x_166 = lean_array_push(x_165, x_161); +x_167 = lean_array_push(x_166, x_163); +x_168 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_168, 0, x_149); +lean_ctor_set(x_168, 1, x_17); +lean_ctor_set(x_168, 2, x_167); +lean_ctor_set(x_140, 0, x_168); +return x_140; } else { -lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; -x_170 = lean_ctor_get(x_141, 0); -x_171 = lean_ctor_get(x_141, 1); -lean_inc(x_171); -lean_inc(x_170); -lean_dec(x_141); -x_172 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__22; -lean_inc(x_170); -x_173 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_173, 0, x_170); -lean_ctor_set(x_173, 1, x_172); -x_174 = l_unexpandListCons___closed__1; +lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; +x_169 = lean_ctor_get(x_140, 0); +x_170 = lean_ctor_get(x_140, 1); lean_inc(x_170); -x_175 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_175, 0, x_170); -lean_ctor_set(x_175, 1, x_174); -x_176 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_177 = lean_array_push(x_176, x_17); -x_178 = lean_box(2); -x_179 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_180 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_180, 0, x_178); -lean_ctor_set(x_180, 1, x_179); -lean_ctor_set(x_180, 2, x_177); -x_181 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_182 = lean_array_push(x_181, x_175); -x_183 = lean_array_push(x_182, x_180); -x_184 = l_unexpandProdMk___closed__2; -x_185 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_185, 0, x_178); -lean_ctor_set(x_185, 1, x_184); -lean_ctor_set(x_185, 2, x_183); -x_186 = lean_array_push(x_176, x_185); -x_187 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_187, 0, x_178); -lean_ctor_set(x_187, 1, x_179); -lean_ctor_set(x_187, 2, x_186); -x_188 = lean_array_push(x_181, x_16); -x_189 = lean_array_push(x_188, x_187); -x_190 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_190, 0, x_178); -lean_ctor_set(x_190, 1, x_179); -lean_ctor_set(x_190, 2, x_189); -x_191 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__29; -x_192 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_192, 0, x_170); -lean_ctor_set(x_192, 1, x_191); -x_193 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; -x_194 = lean_array_push(x_193, x_173); -x_195 = lean_array_push(x_194, x_190); -x_196 = lean_array_push(x_195, x_192); -x_197 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_197, 0, x_178); -lean_ctor_set(x_197, 1, x_18); -lean_ctor_set(x_197, 2, x_196); -x_198 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_198, 0, x_197); -lean_ctor_set(x_198, 1, x_171); -return x_198; +lean_inc(x_169); +lean_dec(x_140); +x_171 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__22; +lean_inc(x_169); +x_172 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_172, 0, x_169); +lean_ctor_set(x_172, 1, x_171); +x_173 = l_unexpandListCons___closed__1; +lean_inc(x_169); +x_174 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_174, 0, x_169); +lean_ctor_set(x_174, 1, x_173); +x_175 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_176 = lean_array_push(x_175, x_16); +x_177 = lean_box(2); +x_178 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_179 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_179, 0, x_177); +lean_ctor_set(x_179, 1, x_178); +lean_ctor_set(x_179, 2, x_176); +x_180 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_181 = lean_array_push(x_180, x_174); +x_182 = lean_array_push(x_181, x_179); +x_183 = l_unexpandProdMk___closed__2; +x_184 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_184, 0, x_177); +lean_ctor_set(x_184, 1, x_183); +lean_ctor_set(x_184, 2, x_182); +x_185 = lean_array_push(x_175, x_184); +x_186 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_186, 0, x_177); +lean_ctor_set(x_186, 1, x_178); +lean_ctor_set(x_186, 2, x_185); +x_187 = lean_array_push(x_180, x_15); +x_188 = lean_array_push(x_187, x_186); +x_189 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_189, 0, x_177); +lean_ctor_set(x_189, 1, x_178); +lean_ctor_set(x_189, 2, x_188); +x_190 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__29; +x_191 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_191, 0, x_169); +lean_ctor_set(x_191, 1, x_190); +x_192 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; +x_193 = lean_array_push(x_192, x_172); +x_194 = lean_array_push(x_193, x_189); +x_195 = lean_array_push(x_194, x_191); +x_196 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_196, 0, x_177); +lean_ctor_set(x_196, 1, x_17); +lean_ctor_set(x_196, 2, x_195); +x_197 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_197, 0, x_196); +lean_ctor_set(x_197, 1, x_170); +return x_197; } } else { -lean_object* x_199; lean_object* x_200; uint8_t x_201; -x_199 = l_Lean_Syntax_getArg(x_139, x_15); -lean_dec(x_139); -x_200 = l_unexpandProdMk___closed__2; -lean_inc(x_199); -x_201 = l_Lean_Syntax_isOfKind(x_199, x_200); -if (x_201 == 0) -{ -lean_object* x_202; uint8_t x_203; -lean_dec(x_199); +lean_object* x_198; lean_object* x_199; uint8_t x_200; +x_198 = l_Lean_Syntax_getArg(x_138, x_14); lean_dec(x_138); -x_202 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); -x_203 = !lean_is_exclusive(x_202); -if (x_203 == 0) -{ -lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; -x_204 = lean_ctor_get(x_202, 0); -x_205 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__22; -lean_inc(x_204); -x_206 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_206, 0, x_204); -lean_ctor_set(x_206, 1, x_205); -x_207 = l_unexpandListCons___closed__1; -lean_inc(x_204); -x_208 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_208, 0, x_204); -lean_ctor_set(x_208, 1, x_207); -x_209 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_210 = lean_array_push(x_209, x_17); -x_211 = lean_box(2); -x_212 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_213 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_213, 0, x_211); -lean_ctor_set(x_213, 1, x_212); -lean_ctor_set(x_213, 2, x_210); -x_214 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_215 = lean_array_push(x_214, x_208); -x_216 = lean_array_push(x_215, x_213); -x_217 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_217, 0, x_211); -lean_ctor_set(x_217, 1, x_200); -lean_ctor_set(x_217, 2, x_216); -x_218 = lean_array_push(x_209, x_217); -x_219 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_219, 0, x_211); -lean_ctor_set(x_219, 1, x_212); -lean_ctor_set(x_219, 2, x_218); -x_220 = lean_array_push(x_214, x_16); -x_221 = lean_array_push(x_220, x_219); -x_222 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_222, 0, x_211); -lean_ctor_set(x_222, 1, x_212); -lean_ctor_set(x_222, 2, x_221); -x_223 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__29; -x_224 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_224, 0, x_204); -lean_ctor_set(x_224, 1, x_223); -x_225 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; -x_226 = lean_array_push(x_225, x_206); -x_227 = lean_array_push(x_226, x_222); -x_228 = lean_array_push(x_227, x_224); -x_229 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_229, 0, x_211); -lean_ctor_set(x_229, 1, x_18); -lean_ctor_set(x_229, 2, x_228); -lean_ctor_set(x_202, 0, x_229); -return x_202; +x_199 = l_unexpandProdMk___closed__2; +lean_inc(x_198); +x_200 = l_Lean_Syntax_isOfKind(x_198, x_199); +if (x_200 == 0) +{ +lean_object* x_201; uint8_t x_202; +lean_dec(x_198); +lean_dec(x_137); +x_201 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); +x_202 = !lean_is_exclusive(x_201); +if (x_202 == 0) +{ +lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; +x_203 = lean_ctor_get(x_201, 0); +x_204 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__22; +lean_inc(x_203); +x_205 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_205, 0, x_203); +lean_ctor_set(x_205, 1, x_204); +x_206 = l_unexpandListCons___closed__1; +lean_inc(x_203); +x_207 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_207, 0, x_203); +lean_ctor_set(x_207, 1, x_206); +x_208 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_209 = lean_array_push(x_208, x_16); +x_210 = lean_box(2); +x_211 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_212 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_212, 0, x_210); +lean_ctor_set(x_212, 1, x_211); +lean_ctor_set(x_212, 2, x_209); +x_213 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_214 = lean_array_push(x_213, x_207); +x_215 = lean_array_push(x_214, x_212); +x_216 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_216, 0, x_210); +lean_ctor_set(x_216, 1, x_199); +lean_ctor_set(x_216, 2, x_215); +x_217 = lean_array_push(x_208, x_216); +x_218 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_218, 0, x_210); +lean_ctor_set(x_218, 1, x_211); +lean_ctor_set(x_218, 2, x_217); +x_219 = lean_array_push(x_213, x_15); +x_220 = lean_array_push(x_219, x_218); +x_221 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_221, 0, x_210); +lean_ctor_set(x_221, 1, x_211); +lean_ctor_set(x_221, 2, x_220); +x_222 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__29; +x_223 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_223, 0, x_203); +lean_ctor_set(x_223, 1, x_222); +x_224 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; +x_225 = lean_array_push(x_224, x_205); +x_226 = lean_array_push(x_225, x_221); +x_227 = lean_array_push(x_226, x_223); +x_228 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_228, 0, x_210); +lean_ctor_set(x_228, 1, x_17); +lean_ctor_set(x_228, 2, x_227); +lean_ctor_set(x_201, 0, x_228); +return x_201; } else { -lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; -x_230 = lean_ctor_get(x_202, 0); -x_231 = lean_ctor_get(x_202, 1); -lean_inc(x_231); +lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; +x_229 = lean_ctor_get(x_201, 0); +x_230 = lean_ctor_get(x_201, 1); lean_inc(x_230); -lean_dec(x_202); -x_232 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__22; -lean_inc(x_230); -x_233 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_233, 0, x_230); -lean_ctor_set(x_233, 1, x_232); -x_234 = l_unexpandListCons___closed__1; -lean_inc(x_230); -x_235 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_235, 0, x_230); -lean_ctor_set(x_235, 1, x_234); -x_236 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_237 = lean_array_push(x_236, x_17); -x_238 = lean_box(2); -x_239 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_240 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_240, 0, x_238); -lean_ctor_set(x_240, 1, x_239); -lean_ctor_set(x_240, 2, x_237); -x_241 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_242 = lean_array_push(x_241, x_235); -x_243 = lean_array_push(x_242, x_240); -x_244 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_244, 0, x_238); -lean_ctor_set(x_244, 1, x_200); -lean_ctor_set(x_244, 2, x_243); -x_245 = lean_array_push(x_236, x_244); -x_246 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_246, 0, x_238); -lean_ctor_set(x_246, 1, x_239); -lean_ctor_set(x_246, 2, x_245); -x_247 = lean_array_push(x_241, x_16); -x_248 = lean_array_push(x_247, x_246); -x_249 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_249, 0, x_238); -lean_ctor_set(x_249, 1, x_239); -lean_ctor_set(x_249, 2, x_248); -x_250 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__29; -x_251 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_251, 0, x_230); -lean_ctor_set(x_251, 1, x_250); -x_252 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; -x_253 = lean_array_push(x_252, x_233); -x_254 = lean_array_push(x_253, x_249); -x_255 = lean_array_push(x_254, x_251); -x_256 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_256, 0, x_238); -lean_ctor_set(x_256, 1, x_18); -lean_ctor_set(x_256, 2, x_255); -x_257 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_257, 0, x_256); -lean_ctor_set(x_257, 1, x_231); -return x_257; +lean_inc(x_229); +lean_dec(x_201); +x_231 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__22; +lean_inc(x_229); +x_232 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_232, 0, x_229); +lean_ctor_set(x_232, 1, x_231); +x_233 = l_unexpandListCons___closed__1; +lean_inc(x_229); +x_234 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_234, 0, x_229); +lean_ctor_set(x_234, 1, x_233); +x_235 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_236 = lean_array_push(x_235, x_16); +x_237 = lean_box(2); +x_238 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_239 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_239, 0, x_237); +lean_ctor_set(x_239, 1, x_238); +lean_ctor_set(x_239, 2, x_236); +x_240 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_241 = lean_array_push(x_240, x_234); +x_242 = lean_array_push(x_241, x_239); +x_243 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_243, 0, x_237); +lean_ctor_set(x_243, 1, x_199); +lean_ctor_set(x_243, 2, x_242); +x_244 = lean_array_push(x_235, x_243); +x_245 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_245, 0, x_237); +lean_ctor_set(x_245, 1, x_238); +lean_ctor_set(x_245, 2, x_244); +x_246 = lean_array_push(x_240, x_15); +x_247 = lean_array_push(x_246, x_245); +x_248 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_248, 0, x_237); +lean_ctor_set(x_248, 1, x_238); +lean_ctor_set(x_248, 2, x_247); +x_249 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__29; +x_250 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_250, 0, x_229); +lean_ctor_set(x_250, 1, x_249); +x_251 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; +x_252 = lean_array_push(x_251, x_232); +x_253 = lean_array_push(x_252, x_248); +x_254 = lean_array_push(x_253, x_250); +x_255 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_255, 0, x_237); +lean_ctor_set(x_255, 1, x_17); +lean_ctor_set(x_255, 2, x_254); +x_256 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_256, 0, x_255); +lean_ctor_set(x_256, 1, x_230); +return x_256; } } else { -lean_object* x_258; lean_object* x_259; lean_object* x_260; uint8_t x_261; -lean_dec(x_17); -x_258 = l_Lean_Syntax_getArg(x_199, x_8); -lean_dec(x_199); -x_259 = l_Lean_Syntax_getArgs(x_258); -lean_dec(x_258); -x_260 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); -x_261 = !lean_is_exclusive(x_260); -if (x_261 == 0) -{ -lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; -x_262 = lean_ctor_get(x_260, 0); -x_263 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__22; -lean_inc(x_262); -x_264 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_264, 0, x_262); -lean_ctor_set(x_264, 1, x_263); -x_265 = l_unexpandListCons___closed__1; -lean_inc(x_262); -x_266 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_266, 0, x_262); -lean_ctor_set(x_266, 1, x_265); -x_267 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_268 = lean_array_push(x_267, x_138); -lean_inc(x_266); -x_269 = lean_array_push(x_268, x_266); -x_270 = l_Array_append___rarg(x_269, x_259); -x_271 = lean_box(2); -x_272 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_273 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_273, 0, x_271); -lean_ctor_set(x_273, 1, x_272); -lean_ctor_set(x_273, 2, x_270); -x_274 = lean_array_push(x_267, x_266); -x_275 = lean_array_push(x_274, x_273); -x_276 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_276, 0, x_271); -lean_ctor_set(x_276, 1, x_200); -lean_ctor_set(x_276, 2, x_275); -x_277 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_278 = lean_array_push(x_277, x_276); -x_279 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_279, 0, x_271); -lean_ctor_set(x_279, 1, x_272); -lean_ctor_set(x_279, 2, x_278); -x_280 = lean_array_push(x_267, x_16); -x_281 = lean_array_push(x_280, x_279); -x_282 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_282, 0, x_271); -lean_ctor_set(x_282, 1, x_272); -lean_ctor_set(x_282, 2, x_281); -x_283 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__29; -x_284 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_284, 0, x_262); -lean_ctor_set(x_284, 1, x_283); -x_285 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; -x_286 = lean_array_push(x_285, x_264); -x_287 = lean_array_push(x_286, x_282); -x_288 = lean_array_push(x_287, x_284); -x_289 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_289, 0, x_271); -lean_ctor_set(x_289, 1, x_18); -lean_ctor_set(x_289, 2, x_288); -lean_ctor_set(x_260, 0, x_289); -return x_260; +lean_object* x_257; lean_object* x_258; lean_object* x_259; uint8_t x_260; +lean_dec(x_16); +x_257 = l_Lean_Syntax_getArg(x_198, x_8); +lean_dec(x_198); +x_258 = l_Lean_Syntax_getArgs(x_257); +lean_dec(x_257); +x_259 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); +x_260 = !lean_is_exclusive(x_259); +if (x_260 == 0) +{ +lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; +x_261 = lean_ctor_get(x_259, 0); +x_262 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__22; +lean_inc(x_261); +x_263 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_263, 0, x_261); +lean_ctor_set(x_263, 1, x_262); +x_264 = l_unexpandListCons___closed__1; +lean_inc(x_261); +x_265 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_265, 0, x_261); +lean_ctor_set(x_265, 1, x_264); +x_266 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_267 = lean_array_push(x_266, x_137); +lean_inc(x_265); +x_268 = lean_array_push(x_267, x_265); +x_269 = l_Array_append___rarg(x_268, x_258); +x_270 = lean_box(2); +x_271 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_272 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_272, 0, x_270); +lean_ctor_set(x_272, 1, x_271); +lean_ctor_set(x_272, 2, x_269); +x_273 = lean_array_push(x_266, x_265); +x_274 = lean_array_push(x_273, x_272); +x_275 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_275, 0, x_270); +lean_ctor_set(x_275, 1, x_199); +lean_ctor_set(x_275, 2, x_274); +x_276 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_277 = lean_array_push(x_276, x_275); +x_278 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_278, 0, x_270); +lean_ctor_set(x_278, 1, x_271); +lean_ctor_set(x_278, 2, x_277); +x_279 = lean_array_push(x_266, x_15); +x_280 = lean_array_push(x_279, x_278); +x_281 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_281, 0, x_270); +lean_ctor_set(x_281, 1, x_271); +lean_ctor_set(x_281, 2, x_280); +x_282 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__29; +x_283 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_283, 0, x_261); +lean_ctor_set(x_283, 1, x_282); +x_284 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; +x_285 = lean_array_push(x_284, x_263); +x_286 = lean_array_push(x_285, x_281); +x_287 = lean_array_push(x_286, x_283); +x_288 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_288, 0, x_270); +lean_ctor_set(x_288, 1, x_17); +lean_ctor_set(x_288, 2, x_287); +lean_ctor_set(x_259, 0, x_288); +return x_259; } else { -lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; -x_290 = lean_ctor_get(x_260, 0); -x_291 = lean_ctor_get(x_260, 1); -lean_inc(x_291); -lean_inc(x_290); -lean_dec(x_260); -x_292 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__22; -lean_inc(x_290); -x_293 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_293, 0, x_290); -lean_ctor_set(x_293, 1, x_292); -x_294 = l_unexpandListCons___closed__1; +lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; +x_289 = lean_ctor_get(x_259, 0); +x_290 = lean_ctor_get(x_259, 1); lean_inc(x_290); -x_295 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_295, 0, x_290); -lean_ctor_set(x_295, 1, x_294); -x_296 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_297 = lean_array_push(x_296, x_138); -lean_inc(x_295); -x_298 = lean_array_push(x_297, x_295); -x_299 = l_Array_append___rarg(x_298, x_259); -x_300 = lean_box(2); -x_301 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_302 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_302, 0, x_300); -lean_ctor_set(x_302, 1, x_301); -lean_ctor_set(x_302, 2, x_299); -x_303 = lean_array_push(x_296, x_295); -x_304 = lean_array_push(x_303, x_302); -x_305 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_305, 0, x_300); -lean_ctor_set(x_305, 1, x_200); -lean_ctor_set(x_305, 2, x_304); -x_306 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_307 = lean_array_push(x_306, x_305); -x_308 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_308, 0, x_300); -lean_ctor_set(x_308, 1, x_301); -lean_ctor_set(x_308, 2, x_307); -x_309 = lean_array_push(x_296, x_16); -x_310 = lean_array_push(x_309, x_308); -x_311 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_311, 0, x_300); -lean_ctor_set(x_311, 1, x_301); -lean_ctor_set(x_311, 2, x_310); -x_312 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__29; -x_313 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_313, 0, x_290); -lean_ctor_set(x_313, 1, x_312); -x_314 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; -x_315 = lean_array_push(x_314, x_293); -x_316 = lean_array_push(x_315, x_311); -x_317 = lean_array_push(x_316, x_313); -x_318 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_318, 0, x_300); -lean_ctor_set(x_318, 1, x_18); -lean_ctor_set(x_318, 2, x_317); -x_319 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_319, 0, x_318); -lean_ctor_set(x_319, 1, x_291); -return x_319; +lean_inc(x_289); +lean_dec(x_259); +x_291 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__22; +lean_inc(x_289); +x_292 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_292, 0, x_289); +lean_ctor_set(x_292, 1, x_291); +x_293 = l_unexpandListCons___closed__1; +lean_inc(x_289); +x_294 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_294, 0, x_289); +lean_ctor_set(x_294, 1, x_293); +x_295 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_296 = lean_array_push(x_295, x_137); +lean_inc(x_294); +x_297 = lean_array_push(x_296, x_294); +x_298 = l_Array_append___rarg(x_297, x_258); +x_299 = lean_box(2); +x_300 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_301 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_301, 0, x_299); +lean_ctor_set(x_301, 1, x_300); +lean_ctor_set(x_301, 2, x_298); +x_302 = lean_array_push(x_295, x_294); +x_303 = lean_array_push(x_302, x_301); +x_304 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_304, 0, x_299); +lean_ctor_set(x_304, 1, x_199); +lean_ctor_set(x_304, 2, x_303); +x_305 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_306 = lean_array_push(x_305, x_304); +x_307 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_307, 0, x_299); +lean_ctor_set(x_307, 1, x_300); +lean_ctor_set(x_307, 2, x_306); +x_308 = lean_array_push(x_295, x_15); +x_309 = lean_array_push(x_308, x_307); +x_310 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_310, 0, x_299); +lean_ctor_set(x_310, 1, x_300); +lean_ctor_set(x_310, 2, x_309); +x_311 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__29; +x_312 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_312, 0, x_289); +lean_ctor_set(x_312, 1, x_311); +x_313 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; +x_314 = lean_array_push(x_313, x_292); +x_315 = lean_array_push(x_314, x_310); +x_316 = lean_array_push(x_315, x_312); +x_317 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_317, 0, x_299); +lean_ctor_set(x_317, 1, x_17); +lean_ctor_set(x_317, 2, x_316); +x_318 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_318, 0, x_317); +lean_ctor_set(x_318, 1, x_290); +return x_318; } } } @@ -7720,7 +7569,7 @@ LEAN_EXPORT lean_object* l_unexpandIte(lean_object* x_1, lean_object* x_2, lean_ _start: { lean_object* x_4; uint8_t x_5; -x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__6; +x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__6; lean_inc(x_1); x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); if (x_5 == 0) @@ -7736,109 +7585,108 @@ return x_7; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; x_8 = lean_unsigned_to_nat(1u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); lean_dec(x_1); -x_10 = l_Lean_nullKind; -x_11 = lean_unsigned_to_nat(3u); +x_10 = lean_unsigned_to_nat(3u); lean_inc(x_9); -x_12 = l_Lean_Syntax_isNodeOf(x_9, x_10, x_11); -if (x_12 == 0) +x_11 = l_Lean_Syntax_matchesNull(x_9, x_10); +if (x_11 == 0) { -lean_object* x_13; lean_object* x_14; +lean_object* x_12; lean_object* x_13; lean_dec(x_9); lean_dec(x_2); -x_13 = lean_box(0); -x_14 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_3); -return x_14; -} -else -{ -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 = lean_unsigned_to_nat(0u); -x_16 = l_Lean_Syntax_getArg(x_9, x_15); -x_17 = l_Lean_Syntax_getArg(x_9, x_8); -x_18 = lean_unsigned_to_nat(2u); -x_19 = l_Lean_Syntax_getArg(x_9, x_18); +x_12 = lean_box(0); +x_13 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_3); +return x_13; +} +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 = lean_unsigned_to_nat(0u); +x_15 = l_Lean_Syntax_getArg(x_9, x_14); +x_16 = l_Lean_Syntax_getArg(x_9, x_8); +x_17 = lean_unsigned_to_nat(2u); +x_18 = l_Lean_Syntax_getArg(x_9, x_17); lean_dec(x_9); -x_20 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); -x_21 = !lean_is_exclusive(x_20); -if (x_21 == 0) +x_19 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) { -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; 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; -x_22 = lean_ctor_get(x_20, 0); -x_23 = l_unexpandIte___closed__3; -lean_inc(x_22); -x_24 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_23); -x_25 = l_unexpandIte___closed__4; -lean_inc(x_22); -x_26 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_26, 0, x_22); -lean_ctor_set(x_26, 1, x_25); -x_27 = l_unexpandIte___closed__5; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_22); -lean_ctor_set(x_28, 1, x_27); -x_29 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__21; -x_30 = lean_array_push(x_29, x_24); -x_31 = lean_array_push(x_30, x_16); -x_32 = lean_array_push(x_31, x_26); -x_33 = lean_array_push(x_32, x_17); -x_34 = lean_array_push(x_33, x_28); -x_35 = lean_array_push(x_34, x_19); -x_36 = lean_box(2); -x_37 = l_unexpandIte___closed__2; -x_38 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -lean_ctor_set(x_38, 2, x_35); -lean_ctor_set(x_20, 0, x_38); -return 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; 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; +x_21 = lean_ctor_get(x_19, 0); +x_22 = l_unexpandIte___closed__3; +lean_inc(x_21); +x_23 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +x_24 = l_unexpandIte___closed__4; +lean_inc(x_21); +x_25 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_25, 0, x_21); +lean_ctor_set(x_25, 1, x_24); +x_26 = l_unexpandIte___closed__5; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_21); +lean_ctor_set(x_27, 1, x_26); +x_28 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__28; +x_29 = lean_array_push(x_28, x_23); +x_30 = lean_array_push(x_29, x_15); +x_31 = lean_array_push(x_30, x_25); +x_32 = lean_array_push(x_31, x_16); +x_33 = lean_array_push(x_32, x_27); +x_34 = lean_array_push(x_33, x_18); +x_35 = lean_box(2); +x_36 = l_unexpandIte___closed__2; +x_37 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +lean_ctor_set(x_37, 2, x_34); +lean_ctor_set(x_19, 0, x_37); +return x_19; } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_39 = lean_ctor_get(x_20, 0); -x_40 = lean_ctor_get(x_20, 1); -lean_inc(x_40); -lean_inc(x_39); -lean_dec(x_20); -x_41 = l_unexpandIte___closed__3; -lean_inc(x_39); -x_42 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_42, 0, x_39); -lean_ctor_set(x_42, 1, x_41); -x_43 = l_unexpandIte___closed__4; +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_38 = lean_ctor_get(x_19, 0); +x_39 = lean_ctor_get(x_19, 1); lean_inc(x_39); -x_44 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_44, 0, x_39); -lean_ctor_set(x_44, 1, x_43); -x_45 = l_unexpandIte___closed__5; -x_46 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_46, 0, x_39); -lean_ctor_set(x_46, 1, x_45); -x_47 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__21; -x_48 = lean_array_push(x_47, x_42); -x_49 = lean_array_push(x_48, x_16); -x_50 = lean_array_push(x_49, x_44); -x_51 = lean_array_push(x_50, x_17); -x_52 = lean_array_push(x_51, x_46); -x_53 = lean_array_push(x_52, x_19); -x_54 = lean_box(2); -x_55 = l_unexpandIte___closed__2; -x_56 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_56, 0, x_54); -lean_ctor_set(x_56, 1, x_55); -lean_ctor_set(x_56, 2, x_53); -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_40); -return x_57; +lean_inc(x_38); +lean_dec(x_19); +x_40 = l_unexpandIte___closed__3; +lean_inc(x_38); +x_41 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_41, 0, x_38); +lean_ctor_set(x_41, 1, x_40); +x_42 = l_unexpandIte___closed__4; +lean_inc(x_38); +x_43 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_43, 0, x_38); +lean_ctor_set(x_43, 1, x_42); +x_44 = l_unexpandIte___closed__5; +x_45 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_45, 0, x_38); +lean_ctor_set(x_45, 1, x_44); +x_46 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__28; +x_47 = lean_array_push(x_46, x_41); +x_48 = lean_array_push(x_47, x_15); +x_49 = lean_array_push(x_48, x_43); +x_50 = lean_array_push(x_49, x_16); +x_51 = lean_array_push(x_50, x_45); +x_52 = lean_array_push(x_51, x_18); +x_53 = lean_box(2); +x_54 = l_unexpandIte___closed__2; +x_55 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +lean_ctor_set(x_55, 2, x_52); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_39); +return x_56; } } } @@ -7856,7 +7704,7 @@ static lean_object* _init_l_unexpandSorryAx___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4; x_2 = l_unexpandSorryAx___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -7866,7 +7714,7 @@ LEAN_EXPORT lean_object* l_unexpandSorryAx(lean_object* x_1, lean_object* x_2, l _start: { lean_object* x_4; uint8_t x_5; -x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__6; +x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__6; lean_inc(x_1); x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); if (x_5 == 0) @@ -7882,112 +7730,111 @@ return x_7; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +lean_object* x_8; lean_object* x_9; uint8_t x_10; x_8 = lean_unsigned_to_nat(1u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); lean_dec(x_1); -x_10 = l_Lean_nullKind; lean_inc(x_9); -x_11 = l_Lean_Syntax_isNodeOf(x_9, x_10, x_8); -if (x_11 == 0) +x_10 = l_Lean_Syntax_matchesNull(x_9, x_8); +if (x_10 == 0) { -lean_object* x_12; uint8_t x_13; -x_12 = lean_unsigned_to_nat(2u); +lean_object* x_11; uint8_t x_12; +x_11 = lean_unsigned_to_nat(2u); lean_inc(x_9); -x_13 = l_Lean_Syntax_isNodeOf(x_9, x_10, x_12); -if (x_13 == 0) +x_12 = l_Lean_Syntax_matchesNull(x_9, x_11); +if (x_12 == 0) { -lean_object* x_14; lean_object* x_15; +lean_object* x_13; lean_object* x_14; lean_dec(x_9); lean_dec(x_2); -x_14 = lean_box(0); -x_15 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_15, 1, x_3); -return x_15; +x_13 = lean_box(0); +x_14 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_3); +return x_14; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_16 = lean_unsigned_to_nat(0u); -x_17 = l_Lean_Syntax_getArg(x_9, x_16); -x_18 = l_Lean_expandExplicitBindersAux_loop___closed__6; -x_19 = l_Lean_Syntax_isOfKind(x_17, x_18); -if (x_19 == 0) +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_15 = lean_unsigned_to_nat(0u); +x_16 = l_Lean_Syntax_getArg(x_9, x_15); +x_17 = l_Lean_expandExplicitBindersAux_loop___closed__6; +x_18 = l_Lean_Syntax_isOfKind(x_16, x_17); +if (x_18 == 0) { -lean_object* x_20; lean_object* x_21; +lean_object* x_19; lean_object* x_20; lean_dec(x_9); lean_dec(x_2); -x_20 = lean_box(0); -x_21 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_3); -return x_21; +x_19 = lean_box(0); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_3); +return x_20; } else { -lean_object* x_22; uint8_t x_23; -x_22 = l_Lean_Syntax_getArg(x_9, x_8); +lean_object* x_21; uint8_t x_22; +x_21 = l_Lean_Syntax_getArg(x_9, x_8); lean_dec(x_9); -x_23 = l_Lean_Syntax_isOfKind(x_22, x_18); -if (x_23 == 0) +x_22 = l_Lean_Syntax_isOfKind(x_21, x_17); +if (x_22 == 0) { -lean_object* x_24; lean_object* x_25; +lean_object* x_23; lean_object* x_24; lean_dec(x_2); -x_24 = lean_box(0); -x_25 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_3); -return x_25; +x_23 = lean_box(0); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_3); +return x_24; } else { -lean_object* x_26; uint8_t x_27; -x_26 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); -x_27 = !lean_is_exclusive(x_26); -if (x_27 == 0) +lean_object* x_25; uint8_t x_26; +x_25 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); +x_26 = !lean_is_exclusive(x_25); +if (x_26 == 0) { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_28 = lean_ctor_get(x_26, 0); -x_29 = l_unexpandSorryAx___closed__1; -x_30 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -x_31 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_32 = lean_array_push(x_31, x_30); -x_33 = lean_box(2); -x_34 = l_unexpandSorryAx___closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_26, 0, x_35); -return x_26; +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_27 = lean_ctor_get(x_25, 0); +x_28 = l_unexpandSorryAx___closed__1; +x_29 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +x_30 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_31 = lean_array_push(x_30, x_29); +x_32 = lean_box(2); +x_33 = l_unexpandSorryAx___closed__2; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_25, 0, x_34); +return x_25; } else { -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; lean_object* x_44; lean_object* x_45; -x_36 = lean_ctor_get(x_26, 0); -x_37 = lean_ctor_get(x_26, 1); -lean_inc(x_37); +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; lean_object* x_44; +x_35 = lean_ctor_get(x_25, 0); +x_36 = lean_ctor_get(x_25, 1); lean_inc(x_36); -lean_dec(x_26); -x_38 = l_unexpandSorryAx___closed__1; -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -x_40 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_41 = lean_array_push(x_40, x_39); -x_42 = lean_box(2); -x_43 = l_unexpandSorryAx___closed__2; -x_44 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_44, 1, x_43); -lean_ctor_set(x_44, 2, x_41); -x_45 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_45, 0, x_44); -lean_ctor_set(x_45, 1, x_37); -return x_45; +lean_inc(x_35); +lean_dec(x_25); +x_37 = l_unexpandSorryAx___closed__1; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +x_39 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_40 = lean_array_push(x_39, x_38); +x_41 = lean_box(2); +x_42 = l_unexpandSorryAx___closed__2; +x_43 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +lean_ctor_set(x_43, 2, x_40); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_36); +return x_44; } } } @@ -7995,70 +7842,70 @@ return x_45; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; -x_46 = lean_unsigned_to_nat(0u); -x_47 = l_Lean_Syntax_getArg(x_9, x_46); +lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; +x_45 = lean_unsigned_to_nat(0u); +x_46 = l_Lean_Syntax_getArg(x_9, x_45); lean_dec(x_9); -x_48 = l_Lean_expandExplicitBindersAux_loop___closed__6; -x_49 = l_Lean_Syntax_isOfKind(x_47, x_48); -if (x_49 == 0) +x_47 = l_Lean_expandExplicitBindersAux_loop___closed__6; +x_48 = l_Lean_Syntax_isOfKind(x_46, x_47); +if (x_48 == 0) { -lean_object* x_50; lean_object* x_51; +lean_object* x_49; lean_object* x_50; lean_dec(x_2); -x_50 = lean_box(0); -x_51 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_3); -return x_51; +x_49 = lean_box(0); +x_50 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_50, 0, x_49); +lean_ctor_set(x_50, 1, x_3); +return x_50; } else { -lean_object* x_52; uint8_t x_53; -x_52 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); -x_53 = !lean_is_exclusive(x_52); -if (x_53 == 0) +lean_object* x_51; uint8_t x_52; +x_51 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); +x_52 = !lean_is_exclusive(x_51); +if (x_52 == 0) { -lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_54 = lean_ctor_get(x_52, 0); -x_55 = l_unexpandSorryAx___closed__1; -x_56 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_56, 0, x_54); -lean_ctor_set(x_56, 1, x_55); -x_57 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_58 = lean_array_push(x_57, x_56); -x_59 = lean_box(2); -x_60 = l_unexpandSorryAx___closed__2; -x_61 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_61, 0, x_59); -lean_ctor_set(x_61, 1, x_60); -lean_ctor_set(x_61, 2, x_58); -lean_ctor_set(x_52, 0, x_61); -return x_52; +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_53 = lean_ctor_get(x_51, 0); +x_54 = l_unexpandSorryAx___closed__1; +x_55 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +x_56 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_57 = lean_array_push(x_56, x_55); +x_58 = lean_box(2); +x_59 = l_unexpandSorryAx___closed__2; +x_60 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +lean_ctor_set(x_60, 2, x_57); +lean_ctor_set(x_51, 0, x_60); +return x_51; } else { -lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_62 = lean_ctor_get(x_52, 0); -x_63 = lean_ctor_get(x_52, 1); -lean_inc(x_63); +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_61 = lean_ctor_get(x_51, 0); +x_62 = lean_ctor_get(x_51, 1); lean_inc(x_62); -lean_dec(x_52); -x_64 = l_unexpandSorryAx___closed__1; -x_65 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_65, 0, x_62); -lean_ctor_set(x_65, 1, x_64); -x_66 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_67 = lean_array_push(x_66, x_65); -x_68 = lean_box(2); -x_69 = l_unexpandSorryAx___closed__2; -x_70 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_69); -lean_ctor_set(x_70, 2, x_67); -x_71 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_63); -return x_71; +lean_inc(x_61); +lean_dec(x_51); +x_63 = l_unexpandSorryAx___closed__1; +x_64 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_64, 0, x_61); +lean_ctor_set(x_64, 1, x_63); +x_65 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_66 = lean_array_push(x_65, x_64); +x_67 = lean_box(2); +x_68 = l_unexpandSorryAx___closed__2; +x_69 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_69, 0, x_67); +lean_ctor_set(x_69, 1, x_68); +lean_ctor_set(x_69, 2, x_66); +x_70 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set(x_70, 1, x_62); +return x_70; } } } @@ -8077,7 +7924,7 @@ static lean_object* _init_l_unexpandEqNDRec___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4; x_2 = l_unexpandEqNDRec___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -8095,7 +7942,7 @@ LEAN_EXPORT lean_object* l_unexpandEqNDRec(lean_object* x_1, lean_object* x_2, l _start: { lean_object* x_4; uint8_t x_5; -x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__6; +x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__6; lean_inc(x_1); x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); if (x_5 == 0) @@ -8111,95 +7958,94 @@ return x_7; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; x_8 = lean_unsigned_to_nat(1u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); lean_dec(x_1); -x_10 = l_Lean_nullKind; -x_11 = lean_unsigned_to_nat(2u); +x_10 = lean_unsigned_to_nat(2u); lean_inc(x_9); -x_12 = l_Lean_Syntax_isNodeOf(x_9, x_10, x_11); -if (x_12 == 0) +x_11 = l_Lean_Syntax_matchesNull(x_9, x_10); +if (x_11 == 0) { -lean_object* x_13; lean_object* x_14; +lean_object* x_12; lean_object* x_13; lean_dec(x_9); lean_dec(x_2); -x_13 = lean_box(0); -x_14 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_3); -return x_14; +x_12 = lean_box(0); +x_13 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_3); +return x_13; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_15 = lean_unsigned_to_nat(0u); -x_16 = l_Lean_Syntax_getArg(x_9, x_15); -x_17 = l_Lean_Syntax_getArg(x_9, x_8); +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_14 = lean_unsigned_to_nat(0u); +x_15 = l_Lean_Syntax_getArg(x_9, x_14); +x_16 = l_Lean_Syntax_getArg(x_9, x_8); lean_dec(x_9); -x_18 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); -x_19 = !lean_is_exclusive(x_18); -if (x_19 == 0) +x_17 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) { -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; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_20 = lean_ctor_get(x_18, 0); -x_21 = l_unexpandEqNDRec___closed__3; -x_22 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_22, 0, x_20); -lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_24 = lean_array_push(x_23, x_16); -x_25 = lean_box(2); -x_26 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_27 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_27, 0, x_25); -lean_ctor_set(x_27, 1, x_26); -lean_ctor_set(x_27, 2, x_24); -x_28 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; -x_29 = lean_array_push(x_28, x_17); -x_30 = lean_array_push(x_29, x_22); -x_31 = lean_array_push(x_30, x_27); -x_32 = l_unexpandEqNDRec___closed__2; -x_33 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_33, 0, x_25); -lean_ctor_set(x_33, 1, x_32); -lean_ctor_set(x_33, 2, x_31); -lean_ctor_set(x_18, 0, x_33); -return 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; 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_object* x_32; +x_19 = lean_ctor_get(x_17, 0); +x_20 = l_unexpandEqNDRec___closed__3; +x_21 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +x_22 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_23 = lean_array_push(x_22, x_15); +x_24 = lean_box(2); +x_25 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_26 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +lean_ctor_set(x_26, 2, x_23); +x_27 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; +x_28 = lean_array_push(x_27, x_16); +x_29 = lean_array_push(x_28, x_21); +x_30 = lean_array_push(x_29, x_26); +x_31 = l_unexpandEqNDRec___closed__2; +x_32 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_32, 0, x_24); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set(x_32, 2, x_30); +lean_ctor_set(x_17, 0, x_32); +return x_17; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_34 = lean_ctor_get(x_18, 0); -x_35 = lean_ctor_get(x_18, 1); -lean_inc(x_35); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_33 = lean_ctor_get(x_17, 0); +x_34 = lean_ctor_get(x_17, 1); lean_inc(x_34); -lean_dec(x_18); -x_36 = l_unexpandEqNDRec___closed__3; -x_37 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_37, 0, x_34); -lean_ctor_set(x_37, 1, x_36); -x_38 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_39 = lean_array_push(x_38, x_16); -x_40 = lean_box(2); -x_41 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_42 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_42, 0, x_40); -lean_ctor_set(x_42, 1, x_41); -lean_ctor_set(x_42, 2, x_39); -x_43 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; -x_44 = lean_array_push(x_43, x_17); -x_45 = lean_array_push(x_44, x_37); -x_46 = lean_array_push(x_45, x_42); -x_47 = l_unexpandEqNDRec___closed__2; -x_48 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_48, 0, x_40); -lean_ctor_set(x_48, 1, x_47); -lean_ctor_set(x_48, 2, x_46); -x_49 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set(x_49, 1, x_35); -return x_49; +lean_inc(x_33); +lean_dec(x_17); +x_35 = l_unexpandEqNDRec___closed__3; +x_36 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_36, 0, x_33); +lean_ctor_set(x_36, 1, x_35); +x_37 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_38 = lean_array_push(x_37, x_15); +x_39 = lean_box(2); +x_40 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_41 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +lean_ctor_set(x_41, 2, x_38); +x_42 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; +x_43 = lean_array_push(x_42, x_16); +x_44 = lean_array_push(x_43, x_36); +x_45 = lean_array_push(x_44, x_41); +x_46 = l_unexpandEqNDRec___closed__2; +x_47 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_47, 0, x_39); +lean_ctor_set(x_47, 1, x_46); +lean_ctor_set(x_47, 2, x_45); +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_34); +return x_48; } } } @@ -8209,7 +8055,7 @@ LEAN_EXPORT lean_object* l_unexpandEqRec(lean_object* x_1, lean_object* x_2, lea _start: { lean_object* x_4; uint8_t x_5; -x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__6; +x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__6; lean_inc(x_1); x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); if (x_5 == 0) @@ -8225,95 +8071,94 @@ return x_7; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; x_8 = lean_unsigned_to_nat(1u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); lean_dec(x_1); -x_10 = l_Lean_nullKind; -x_11 = lean_unsigned_to_nat(2u); +x_10 = lean_unsigned_to_nat(2u); lean_inc(x_9); -x_12 = l_Lean_Syntax_isNodeOf(x_9, x_10, x_11); -if (x_12 == 0) +x_11 = l_Lean_Syntax_matchesNull(x_9, x_10); +if (x_11 == 0) { -lean_object* x_13; lean_object* x_14; +lean_object* x_12; lean_object* x_13; lean_dec(x_9); lean_dec(x_2); -x_13 = lean_box(0); -x_14 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_3); -return x_14; +x_12 = lean_box(0); +x_13 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_3); +return x_13; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_15 = lean_unsigned_to_nat(0u); -x_16 = l_Lean_Syntax_getArg(x_9, x_15); -x_17 = l_Lean_Syntax_getArg(x_9, x_8); +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_14 = lean_unsigned_to_nat(0u); +x_15 = l_Lean_Syntax_getArg(x_9, x_14); +x_16 = l_Lean_Syntax_getArg(x_9, x_8); lean_dec(x_9); -x_18 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); -x_19 = !lean_is_exclusive(x_18); -if (x_19 == 0) +x_17 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) { -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; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_20 = lean_ctor_get(x_18, 0); -x_21 = l_unexpandEqNDRec___closed__3; -x_22 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_22, 0, x_20); -lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_24 = lean_array_push(x_23, x_16); -x_25 = lean_box(2); -x_26 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_27 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_27, 0, x_25); -lean_ctor_set(x_27, 1, x_26); -lean_ctor_set(x_27, 2, x_24); -x_28 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; -x_29 = lean_array_push(x_28, x_17); -x_30 = lean_array_push(x_29, x_22); -x_31 = lean_array_push(x_30, x_27); -x_32 = l_unexpandEqNDRec___closed__2; -x_33 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_33, 0, x_25); -lean_ctor_set(x_33, 1, x_32); -lean_ctor_set(x_33, 2, x_31); -lean_ctor_set(x_18, 0, x_33); -return 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; 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_object* x_32; +x_19 = lean_ctor_get(x_17, 0); +x_20 = l_unexpandEqNDRec___closed__3; +x_21 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +x_22 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_23 = lean_array_push(x_22, x_15); +x_24 = lean_box(2); +x_25 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_26 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +lean_ctor_set(x_26, 2, x_23); +x_27 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; +x_28 = lean_array_push(x_27, x_16); +x_29 = lean_array_push(x_28, x_21); +x_30 = lean_array_push(x_29, x_26); +x_31 = l_unexpandEqNDRec___closed__2; +x_32 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_32, 0, x_24); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set(x_32, 2, x_30); +lean_ctor_set(x_17, 0, x_32); +return x_17; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_34 = lean_ctor_get(x_18, 0); -x_35 = lean_ctor_get(x_18, 1); -lean_inc(x_35); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_33 = lean_ctor_get(x_17, 0); +x_34 = lean_ctor_get(x_17, 1); lean_inc(x_34); -lean_dec(x_18); -x_36 = l_unexpandEqNDRec___closed__3; -x_37 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_37, 0, x_34); -lean_ctor_set(x_37, 1, x_36); -x_38 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_39 = lean_array_push(x_38, x_16); -x_40 = lean_box(2); -x_41 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_42 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_42, 0, x_40); -lean_ctor_set(x_42, 1, x_41); -lean_ctor_set(x_42, 2, x_39); -x_43 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; -x_44 = lean_array_push(x_43, x_17); -x_45 = lean_array_push(x_44, x_37); -x_46 = lean_array_push(x_45, x_42); -x_47 = l_unexpandEqNDRec___closed__2; -x_48 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_48, 0, x_40); -lean_ctor_set(x_48, 1, x_47); -lean_ctor_set(x_48, 2, x_46); -x_49 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set(x_49, 1, x_35); -return x_49; +lean_inc(x_33); +lean_dec(x_17); +x_35 = l_unexpandEqNDRec___closed__3; +x_36 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_36, 0, x_33); +lean_ctor_set(x_36, 1, x_35); +x_37 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_38 = lean_array_push(x_37, x_15); +x_39 = lean_box(2); +x_40 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_41 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +lean_ctor_set(x_41, 2, x_38); +x_42 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; +x_43 = lean_array_push(x_42, x_16); +x_44 = lean_array_push(x_43, x_36); +x_45 = lean_array_push(x_44, x_41); +x_46 = l_unexpandEqNDRec___closed__2; +x_47 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_47, 0, x_39); +lean_ctor_set(x_47, 1, x_46); +lean_ctor_set(x_47, 2, x_45); +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_34); +return x_48; } } } @@ -8331,7 +8176,7 @@ static lean_object* _init_l_unexpandExists___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4; x_2 = l_unexpandExists___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -8348,13 +8193,31 @@ return x_1; static lean_object* _init_l_unexpandExists___closed__4() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("binderIdent", 11); +return x_1; +} +} +static lean_object* _init_l_unexpandExists___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__2; +x_2 = l_unexpandExists___closed__4; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_unexpandExists___closed__6() { +_start: +{ lean_object* x_1; lean_object* x_2; x_1 = lean_unsigned_to_nat(5u); x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_unexpandExists___closed__5() { +static lean_object* _init_l_unexpandExists___closed__7() { _start: { lean_object* x_1; lean_object* x_2; @@ -8367,7 +8230,7 @@ LEAN_EXPORT lean_object* l_unexpandExists(lean_object* x_1, lean_object* x_2, le _start: { lean_object* x_4; uint8_t x_5; -x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__6; +x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__6; lean_inc(x_1); x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); if (x_5 == 0) @@ -8383,355 +8246,354 @@ return x_7; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +lean_object* x_8; lean_object* x_9; uint8_t x_10; x_8 = lean_unsigned_to_nat(1u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); lean_dec(x_1); -x_10 = l_Lean_nullKind; lean_inc(x_9); -x_11 = l_Lean_Syntax_isNodeOf(x_9, x_10, x_8); -if (x_11 == 0) +x_10 = l_Lean_Syntax_matchesNull(x_9, x_8); +if (x_10 == 0) { -lean_object* x_12; lean_object* x_13; +lean_object* x_11; lean_object* x_12; lean_dec(x_9); lean_dec(x_2); -x_12 = lean_box(0); -x_13 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_3); -return x_13; +x_11 = lean_box(0); +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_3); +return x_12; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_14 = lean_unsigned_to_nat(0u); -x_15 = l_Lean_Syntax_getArg(x_9, x_14); +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_13 = lean_unsigned_to_nat(0u); +x_14 = l_Lean_Syntax_getArg(x_9, x_13); lean_dec(x_9); -x_16 = l_Lean_expandExplicitBindersAux_loop___closed__2; -lean_inc(x_15); -x_17 = l_Lean_Syntax_isOfKind(x_15, x_16); -if (x_17 == 0) +x_15 = l_Lean_expandExplicitBindersAux_loop___closed__2; +lean_inc(x_14); +x_16 = l_Lean_Syntax_isOfKind(x_14, x_15); +if (x_16 == 0) { -lean_object* x_18; lean_object* x_19; -lean_dec(x_15); +lean_object* x_17; lean_object* x_18; +lean_dec(x_14); lean_dec(x_2); -x_18 = lean_box(0); -x_19 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_3); -return x_19; +x_17 = lean_box(0); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_3); +return x_18; } else { -lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_20 = l_Lean_Syntax_getArg(x_15, x_8); -lean_dec(x_15); -x_21 = l_Lean_expandExplicitBindersAux_loop___closed__4; -lean_inc(x_20); -x_22 = l_Lean_Syntax_isOfKind(x_20, x_21); -if (x_22 == 0) +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = l_Lean_Syntax_getArg(x_14, x_8); +lean_dec(x_14); +x_20 = l_Lean_expandExplicitBindersAux_loop___closed__4; +lean_inc(x_19); +x_21 = l_Lean_Syntax_isOfKind(x_19, x_20); +if (x_21 == 0) { -lean_object* x_23; lean_object* x_24; -lean_dec(x_20); +lean_object* x_22; lean_object* x_23; +lean_dec(x_19); lean_dec(x_2); -x_23 = lean_box(0); -x_24 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_3); -return x_24; +x_22 = lean_box(0); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_3); +return x_23; } else { -lean_object* x_25; uint8_t x_26; -x_25 = l_Lean_Syntax_getArg(x_20, x_14); -lean_inc(x_25); -x_26 = l_Lean_Syntax_isNodeOf(x_25, x_10, x_8); -if (x_26 == 0) +lean_object* x_24; uint8_t x_25; +x_24 = l_Lean_Syntax_getArg(x_19, x_13); +lean_inc(x_24); +x_25 = l_Lean_Syntax_matchesNull(x_24, x_8); +if (x_25 == 0) { -lean_object* x_27; lean_object* x_28; -lean_dec(x_25); -lean_dec(x_20); +lean_object* x_26; lean_object* x_27; +lean_dec(x_24); +lean_dec(x_19); lean_dec(x_2); -x_27 = lean_box(0); -x_28 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_28, 0, x_27); -lean_ctor_set(x_28, 1, x_3); -return x_28; +x_26 = lean_box(0); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_3); +return x_27; } else { -lean_object* x_29; lean_object* x_30; uint8_t x_31; -x_29 = l_Lean_Syntax_getArg(x_25, x_14); -lean_dec(x_25); -x_30 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__10; -lean_inc(x_29); -x_31 = l_Lean_Syntax_isOfKind(x_29, x_30); -if (x_31 == 0) +lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_28 = l_Lean_Syntax_getArg(x_24, x_13); +lean_dec(x_24); +x_29 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__10; +lean_inc(x_28); +x_30 = l_Lean_Syntax_isOfKind(x_28, x_29); +if (x_30 == 0) { -lean_object* x_32; uint8_t x_33; -x_32 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__21; -lean_inc(x_29); -x_33 = l_Lean_Syntax_isOfKind(x_29, x_32); -if (x_33 == 0) +lean_object* x_31; uint8_t x_32; +x_31 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__21; +lean_inc(x_28); +x_32 = l_Lean_Syntax_isOfKind(x_28, x_31); +if (x_32 == 0) { -lean_object* x_34; lean_object* x_35; -lean_dec(x_29); -lean_dec(x_20); +lean_object* x_33; lean_object* x_34; +lean_dec(x_28); +lean_dec(x_19); lean_dec(x_2); -x_34 = lean_box(0); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_3); -return x_35; +x_33 = lean_box(0); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_3); +return x_34; } else { -lean_object* x_36; lean_object* x_37; uint8_t x_38; -x_36 = l_Lean_Syntax_getArg(x_29, x_8); -lean_dec(x_29); -x_37 = lean_unsigned_to_nat(2u); -lean_inc(x_36); -x_38 = l_Lean_Syntax_isNodeOf(x_36, x_10, x_37); -if (x_38 == 0) +lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_35 = l_Lean_Syntax_getArg(x_28, x_8); +lean_dec(x_28); +x_36 = lean_unsigned_to_nat(2u); +lean_inc(x_35); +x_37 = l_Lean_Syntax_matchesNull(x_35, x_36); +if (x_37 == 0) { -lean_object* x_39; lean_object* x_40; -lean_dec(x_36); -lean_dec(x_20); +lean_object* x_38; lean_object* x_39; +lean_dec(x_35); +lean_dec(x_19); lean_dec(x_2); -x_39 = lean_box(0); -x_40 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_3); -return x_40; +x_38 = lean_box(0); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_3); +return x_39; } else { -lean_object* x_41; uint8_t x_42; -x_41 = l_Lean_Syntax_getArg(x_36, x_14); -lean_inc(x_41); -x_42 = l_Lean_Syntax_isOfKind(x_41, x_30); -if (x_42 == 0) +lean_object* x_40; uint8_t x_41; +x_40 = l_Lean_Syntax_getArg(x_35, x_13); +lean_inc(x_40); +x_41 = l_Lean_Syntax_isOfKind(x_40, x_29); +if (x_41 == 0) { -lean_object* x_43; lean_object* x_44; -lean_dec(x_41); -lean_dec(x_36); -lean_dec(x_20); +lean_object* x_42; lean_object* x_43; +lean_dec(x_40); +lean_dec(x_35); +lean_dec(x_19); lean_dec(x_2); -x_43 = lean_box(0); -x_44 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_44, 1, x_3); -return x_44; +x_42 = lean_box(0); +x_43 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_3); +return x_43; } else { -lean_object* x_45; uint8_t x_46; -x_45 = l_Lean_Syntax_getArg(x_36, x_8); -lean_dec(x_36); -lean_inc(x_45); -x_46 = l_Lean_Syntax_isNodeOf(x_45, x_10, x_8); -if (x_46 == 0) +lean_object* x_44; uint8_t x_45; +x_44 = l_Lean_Syntax_getArg(x_35, x_8); +lean_dec(x_35); +lean_inc(x_44); +x_45 = l_Lean_Syntax_matchesNull(x_44, x_8); +if (x_45 == 0) { -lean_object* x_47; lean_object* x_48; -lean_dec(x_45); -lean_dec(x_41); -lean_dec(x_20); +lean_object* x_46; lean_object* x_47; +lean_dec(x_44); +lean_dec(x_40); +lean_dec(x_19); lean_dec(x_2); -x_47 = lean_box(0); -x_48 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_3); -return x_48; +x_46 = lean_box(0); +x_47 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_3); +return x_47; } else { -lean_object* x_49; lean_object* x_50; uint8_t x_51; -x_49 = l_Lean_Syntax_getArg(x_45, x_14); -lean_dec(x_45); -x_50 = l_unexpandExists___closed__2; -lean_inc(x_49); -x_51 = l_Lean_Syntax_isOfKind(x_49, x_50); -if (x_51 == 0) +lean_object* x_48; lean_object* x_49; uint8_t x_50; +x_48 = l_Lean_Syntax_getArg(x_44, x_13); +lean_dec(x_44); +x_49 = l_unexpandExists___closed__2; +lean_inc(x_48); +x_50 = l_Lean_Syntax_isOfKind(x_48, x_49); +if (x_50 == 0) { -lean_object* x_52; lean_object* x_53; -lean_dec(x_49); -lean_dec(x_41); -lean_dec(x_20); +lean_object* x_51; lean_object* x_52; +lean_dec(x_48); +lean_dec(x_40); +lean_dec(x_19); lean_dec(x_2); -x_52 = lean_box(0); -x_53 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_53, 0, x_52); -lean_ctor_set(x_53, 1, x_3); -return x_53; +x_51 = lean_box(0); +x_52 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_3); +return x_52; } else { -lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; -x_54 = l_Lean_Syntax_getArg(x_49, x_8); -lean_dec(x_49); -x_55 = l_Lean_Syntax_getArg(x_20, x_37); -lean_dec(x_20); -x_56 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); -x_57 = !lean_is_exclusive(x_56); -if (x_57 == 0) +lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; +x_53 = l_Lean_Syntax_getArg(x_48, x_8); +lean_dec(x_48); +x_54 = l_Lean_Syntax_getArg(x_19, x_36); +lean_dec(x_19); +x_55 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); +x_56 = !lean_is_exclusive(x_55); +if (x_56 == 0) { -lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; -x_58 = lean_ctor_get(x_56, 0); -x_59 = l_unexpandExists___closed__3; -lean_inc(x_58); -x_60 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_60, 0, x_58); -lean_ctor_set(x_60, 1, x_59); -x_61 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__22; -lean_inc(x_58); -x_62 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_62, 0, x_58); -lean_ctor_set(x_62, 1, x_61); -x_63 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_64 = lean_array_push(x_63, x_41); -x_65 = lean_box(2); -x_66 = l_Lean_binderIdent___closed__2; -x_67 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_67, 0, x_65); -lean_ctor_set(x_67, 1, x_66); -lean_ctor_set(x_67, 2, x_64); -x_68 = lean_array_push(x_63, x_67); -x_69 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_70 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_70, 0, x_65); -lean_ctor_set(x_70, 1, x_69); -lean_ctor_set(x_70, 2, x_68); -x_71 = l_Lean_expandExplicitBindersAux_loop___closed__12; -lean_inc(x_58); -x_72 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_72, 0, x_58); -lean_ctor_set(x_72, 1, x_71); -x_73 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__29; -lean_inc(x_58); -x_74 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_74, 0, x_58); -lean_ctor_set(x_74, 1, x_73); -x_75 = l_unexpandExists___closed__4; -x_76 = lean_array_push(x_75, x_62); -x_77 = lean_array_push(x_76, x_70); -x_78 = lean_array_push(x_77, x_72); -x_79 = lean_array_push(x_78, x_54); -x_80 = lean_array_push(x_79, x_74); -x_81 = l_Lean_bracketedExplicitBinders___closed__2; -x_82 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_82, 0, x_65); -lean_ctor_set(x_82, 1, x_81); -lean_ctor_set(x_82, 2, x_80); -x_83 = lean_array_push(x_63, x_82); -x_84 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_84, 0, x_65); -lean_ctor_set(x_84, 1, x_69); -lean_ctor_set(x_84, 2, x_83); -x_85 = lean_array_push(x_63, x_84); -x_86 = l_Lean_explicitBinders___closed__2; -x_87 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_87, 0, x_65); -lean_ctor_set(x_87, 1, x_86); -lean_ctor_set(x_87, 2, x_85); -x_88 = l_unexpandListCons___closed__1; -x_89 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_89, 0, x_58); -lean_ctor_set(x_89, 1, x_88); -x_90 = l_unexpandExists___closed__5; -x_91 = lean_array_push(x_90, x_60); -x_92 = lean_array_push(x_91, x_87); -x_93 = lean_array_push(x_92, x_89); -x_94 = lean_array_push(x_93, x_55); -x_95 = l_term_u2203___x2c_____closed__2; -x_96 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_96, 0, x_65); -lean_ctor_set(x_96, 1, x_95); -lean_ctor_set(x_96, 2, x_94); -lean_ctor_set(x_56, 0, x_96); -return x_56; -} +lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_57 = lean_ctor_get(x_55, 0); +x_58 = l_unexpandExists___closed__3; +lean_inc(x_57); +x_59 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_58); +x_60 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__22; +lean_inc(x_57); +x_61 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_61, 0, x_57); +lean_ctor_set(x_61, 1, x_60); +x_62 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_63 = lean_array_push(x_62, x_40); +x_64 = lean_box(2); +x_65 = l_unexpandExists___closed__5; +x_66 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +lean_ctor_set(x_66, 2, x_63); +x_67 = lean_array_push(x_62, x_66); +x_68 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_69 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_69, 0, x_64); +lean_ctor_set(x_69, 1, x_68); +lean_ctor_set(x_69, 2, x_67); +x_70 = l_Lean_expandExplicitBindersAux_loop___closed__13; +lean_inc(x_57); +x_71 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_71, 0, x_57); +lean_ctor_set(x_71, 1, x_70); +x_72 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__29; +lean_inc(x_57); +x_73 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_73, 0, x_57); +lean_ctor_set(x_73, 1, x_72); +x_74 = l_unexpandExists___closed__6; +x_75 = lean_array_push(x_74, x_61); +x_76 = lean_array_push(x_75, x_69); +x_77 = lean_array_push(x_76, x_71); +x_78 = lean_array_push(x_77, x_53); +x_79 = lean_array_push(x_78, x_73); +x_80 = l_Lean_bracketedExplicitBinders___closed__2; +x_81 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_81, 0, x_64); +lean_ctor_set(x_81, 1, x_80); +lean_ctor_set(x_81, 2, x_79); +x_82 = lean_array_push(x_62, x_81); +x_83 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_83, 0, x_64); +lean_ctor_set(x_83, 1, x_68); +lean_ctor_set(x_83, 2, x_82); +x_84 = lean_array_push(x_62, x_83); +x_85 = l_Lean_explicitBinders___closed__2; +x_86 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_86, 0, x_64); +lean_ctor_set(x_86, 1, x_85); +lean_ctor_set(x_86, 2, x_84); +x_87 = l_unexpandListCons___closed__1; +x_88 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_88, 0, x_57); +lean_ctor_set(x_88, 1, x_87); +x_89 = l_unexpandExists___closed__7; +x_90 = lean_array_push(x_89, x_59); +x_91 = lean_array_push(x_90, x_86); +x_92 = lean_array_push(x_91, x_88); +x_93 = lean_array_push(x_92, x_54); +x_94 = l_term_u2203___x2c_____closed__2; +x_95 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_95, 0, x_64); +lean_ctor_set(x_95, 1, x_94); +lean_ctor_set(x_95, 2, x_93); +lean_ctor_set(x_55, 0, x_95); +return x_55; +} else { -lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; -x_97 = lean_ctor_get(x_56, 0); -x_98 = lean_ctor_get(x_56, 1); -lean_inc(x_98); -lean_inc(x_97); -lean_dec(x_56); -x_99 = l_unexpandExists___closed__3; -lean_inc(x_97); -x_100 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_100, 0, x_97); -lean_ctor_set(x_100, 1, x_99); -x_101 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__22; -lean_inc(x_97); -x_102 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_102, 0, x_97); -lean_ctor_set(x_102, 1, x_101); -x_103 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_104 = lean_array_push(x_103, x_41); -x_105 = lean_box(2); -x_106 = l_Lean_binderIdent___closed__2; -x_107 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_107, 0, x_105); -lean_ctor_set(x_107, 1, x_106); -lean_ctor_set(x_107, 2, x_104); -x_108 = lean_array_push(x_103, x_107); -x_109 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_110 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_110, 0, x_105); -lean_ctor_set(x_110, 1, x_109); -lean_ctor_set(x_110, 2, x_108); -x_111 = l_Lean_expandExplicitBindersAux_loop___closed__12; -lean_inc(x_97); -x_112 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_112, 0, x_97); -lean_ctor_set(x_112, 1, x_111); -x_113 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__29; +lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; +x_96 = lean_ctor_get(x_55, 0); +x_97 = lean_ctor_get(x_55, 1); lean_inc(x_97); -x_114 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_114, 0, x_97); -lean_ctor_set(x_114, 1, x_113); -x_115 = l_unexpandExists___closed__4; -x_116 = lean_array_push(x_115, x_102); -x_117 = lean_array_push(x_116, x_110); -x_118 = lean_array_push(x_117, x_112); -x_119 = lean_array_push(x_118, x_54); -x_120 = lean_array_push(x_119, x_114); -x_121 = l_Lean_bracketedExplicitBinders___closed__2; -x_122 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_122, 0, x_105); -lean_ctor_set(x_122, 1, x_121); -lean_ctor_set(x_122, 2, x_120); -x_123 = lean_array_push(x_103, x_122); -x_124 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_124, 0, x_105); -lean_ctor_set(x_124, 1, x_109); -lean_ctor_set(x_124, 2, x_123); -x_125 = lean_array_push(x_103, x_124); -x_126 = l_Lean_explicitBinders___closed__2; -x_127 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_127, 0, x_105); -lean_ctor_set(x_127, 1, x_126); -lean_ctor_set(x_127, 2, x_125); -x_128 = l_unexpandListCons___closed__1; -x_129 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_129, 0, x_97); -lean_ctor_set(x_129, 1, x_128); -x_130 = l_unexpandExists___closed__5; -x_131 = lean_array_push(x_130, x_100); -x_132 = lean_array_push(x_131, x_127); -x_133 = lean_array_push(x_132, x_129); -x_134 = lean_array_push(x_133, x_55); -x_135 = l_term_u2203___x2c_____closed__2; -x_136 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_136, 0, x_105); -lean_ctor_set(x_136, 1, x_135); -lean_ctor_set(x_136, 2, x_134); -x_137 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_137, 0, x_136); -lean_ctor_set(x_137, 1, x_98); -return x_137; +lean_inc(x_96); +lean_dec(x_55); +x_98 = l_unexpandExists___closed__3; +lean_inc(x_96); +x_99 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_99, 0, x_96); +lean_ctor_set(x_99, 1, x_98); +x_100 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__22; +lean_inc(x_96); +x_101 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_101, 0, x_96); +lean_ctor_set(x_101, 1, x_100); +x_102 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_103 = lean_array_push(x_102, x_40); +x_104 = lean_box(2); +x_105 = l_unexpandExists___closed__5; +x_106 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_106, 0, x_104); +lean_ctor_set(x_106, 1, x_105); +lean_ctor_set(x_106, 2, x_103); +x_107 = lean_array_push(x_102, x_106); +x_108 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_109 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_109, 0, x_104); +lean_ctor_set(x_109, 1, x_108); +lean_ctor_set(x_109, 2, x_107); +x_110 = l_Lean_expandExplicitBindersAux_loop___closed__13; +lean_inc(x_96); +x_111 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_111, 0, x_96); +lean_ctor_set(x_111, 1, x_110); +x_112 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__29; +lean_inc(x_96); +x_113 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_113, 0, x_96); +lean_ctor_set(x_113, 1, x_112); +x_114 = l_unexpandExists___closed__6; +x_115 = lean_array_push(x_114, x_101); +x_116 = lean_array_push(x_115, x_109); +x_117 = lean_array_push(x_116, x_111); +x_118 = lean_array_push(x_117, x_53); +x_119 = lean_array_push(x_118, x_113); +x_120 = l_Lean_bracketedExplicitBinders___closed__2; +x_121 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_121, 0, x_104); +lean_ctor_set(x_121, 1, x_120); +lean_ctor_set(x_121, 2, x_119); +x_122 = lean_array_push(x_102, x_121); +x_123 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_123, 0, x_104); +lean_ctor_set(x_123, 1, x_108); +lean_ctor_set(x_123, 2, x_122); +x_124 = lean_array_push(x_102, x_123); +x_125 = l_Lean_explicitBinders___closed__2; +x_126 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_126, 0, x_104); +lean_ctor_set(x_126, 1, x_125); +lean_ctor_set(x_126, 2, x_124); +x_127 = l_unexpandListCons___closed__1; +x_128 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_128, 0, x_96); +lean_ctor_set(x_128, 1, x_127); +x_129 = l_unexpandExists___closed__7; +x_130 = lean_array_push(x_129, x_99); +x_131 = lean_array_push(x_130, x_126); +x_132 = lean_array_push(x_131, x_128); +x_133 = lean_array_push(x_132, x_54); +x_134 = l_term_u2203___x2c_____closed__2; +x_135 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_135, 0, x_104); +lean_ctor_set(x_135, 1, x_134); +lean_ctor_set(x_135, 2, x_133); +x_136 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_136, 0, x_135); +lean_ctor_set(x_136, 1, x_97); +return x_136; } } } @@ -8741,629 +8603,629 @@ return x_137; } else { -lean_object* x_138; lean_object* x_139; lean_object* x_140; uint8_t x_141; -x_138 = lean_unsigned_to_nat(2u); -x_139 = l_Lean_Syntax_getArg(x_20, x_138); -lean_dec(x_20); -x_140 = l_term_u2203___x2c_____closed__2; -lean_inc(x_139); -x_141 = l_Lean_Syntax_isOfKind(x_139, x_140); -if (x_141 == 0) +lean_object* x_137; lean_object* x_138; lean_object* x_139; uint8_t x_140; +x_137 = lean_unsigned_to_nat(2u); +x_138 = l_Lean_Syntax_getArg(x_19, x_137); +lean_dec(x_19); +x_139 = l_term_u2203___x2c_____closed__2; +lean_inc(x_138); +x_140 = l_Lean_Syntax_isOfKind(x_138, x_139); +if (x_140 == 0) { -lean_object* x_142; uint8_t x_143; -x_142 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); -x_143 = !lean_is_exclusive(x_142); -if (x_143 == 0) +lean_object* x_141; uint8_t x_142; +x_141 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); +x_142 = !lean_is_exclusive(x_141); +if (x_142 == 0) { -lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; -x_144 = lean_ctor_get(x_142, 0); -x_145 = l_unexpandExists___closed__3; -lean_inc(x_144); -x_146 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_146, 0, x_144); -lean_ctor_set(x_146, 1, x_145); -x_147 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_148 = lean_array_push(x_147, x_29); -x_149 = lean_box(2); -x_150 = l_Lean_binderIdent___closed__2; -x_151 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_151, 0, x_149); -lean_ctor_set(x_151, 1, x_150); -lean_ctor_set(x_151, 2, x_148); -x_152 = lean_array_push(x_147, x_151); -x_153 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_154 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_154, 0, x_149); -lean_ctor_set(x_154, 1, x_153); -lean_ctor_set(x_154, 2, x_152); -x_155 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_156 = lean_array_push(x_155, x_154); -x_157 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; -x_158 = lean_array_push(x_156, x_157); -x_159 = l_Lean_unbracketedExplicitBinders___closed__2; -x_160 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_160, 0, x_149); -lean_ctor_set(x_160, 1, x_159); -lean_ctor_set(x_160, 2, x_158); -x_161 = lean_array_push(x_147, x_160); -x_162 = l_Lean_explicitBinders___closed__2; -x_163 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_163, 0, x_149); -lean_ctor_set(x_163, 1, x_162); -lean_ctor_set(x_163, 2, x_161); -x_164 = l_unexpandListCons___closed__1; -x_165 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_165, 0, x_144); -lean_ctor_set(x_165, 1, x_164); -x_166 = l_unexpandExists___closed__5; -x_167 = lean_array_push(x_166, x_146); -x_168 = lean_array_push(x_167, x_163); -x_169 = lean_array_push(x_168, x_165); -x_170 = lean_array_push(x_169, x_139); -x_171 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_171, 0, x_149); -lean_ctor_set(x_171, 1, x_140); -lean_ctor_set(x_171, 2, x_170); -lean_ctor_set(x_142, 0, x_171); -return x_142; +lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; +x_143 = lean_ctor_get(x_141, 0); +x_144 = l_unexpandExists___closed__3; +lean_inc(x_143); +x_145 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_145, 0, x_143); +lean_ctor_set(x_145, 1, x_144); +x_146 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_147 = lean_array_push(x_146, x_28); +x_148 = lean_box(2); +x_149 = l_unexpandExists___closed__5; +x_150 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_150, 0, x_148); +lean_ctor_set(x_150, 1, x_149); +lean_ctor_set(x_150, 2, x_147); +x_151 = lean_array_push(x_146, x_150); +x_152 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_153 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_153, 0, x_148); +lean_ctor_set(x_153, 1, x_152); +lean_ctor_set(x_153, 2, x_151); +x_154 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_155 = lean_array_push(x_154, x_153); +x_156 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; +x_157 = lean_array_push(x_155, x_156); +x_158 = l_Lean_unbracketedExplicitBinders___closed__2; +x_159 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_159, 0, x_148); +lean_ctor_set(x_159, 1, x_158); +lean_ctor_set(x_159, 2, x_157); +x_160 = lean_array_push(x_146, x_159); +x_161 = l_Lean_explicitBinders___closed__2; +x_162 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_162, 0, x_148); +lean_ctor_set(x_162, 1, x_161); +lean_ctor_set(x_162, 2, x_160); +x_163 = l_unexpandListCons___closed__1; +x_164 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_164, 0, x_143); +lean_ctor_set(x_164, 1, x_163); +x_165 = l_unexpandExists___closed__7; +x_166 = lean_array_push(x_165, x_145); +x_167 = lean_array_push(x_166, x_162); +x_168 = lean_array_push(x_167, x_164); +x_169 = lean_array_push(x_168, x_138); +x_170 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_170, 0, x_148); +lean_ctor_set(x_170, 1, x_139); +lean_ctor_set(x_170, 2, x_169); +lean_ctor_set(x_141, 0, x_170); +return x_141; } else { -lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; -x_172 = lean_ctor_get(x_142, 0); -x_173 = lean_ctor_get(x_142, 1); -lean_inc(x_173); +lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; +x_171 = lean_ctor_get(x_141, 0); +x_172 = lean_ctor_get(x_141, 1); lean_inc(x_172); -lean_dec(x_142); -x_174 = l_unexpandExists___closed__3; -lean_inc(x_172); -x_175 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_175, 0, x_172); -lean_ctor_set(x_175, 1, x_174); -x_176 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_177 = lean_array_push(x_176, x_29); -x_178 = lean_box(2); -x_179 = l_Lean_binderIdent___closed__2; -x_180 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_180, 0, x_178); -lean_ctor_set(x_180, 1, x_179); -lean_ctor_set(x_180, 2, x_177); -x_181 = lean_array_push(x_176, x_180); -x_182 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_183 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_183, 0, x_178); -lean_ctor_set(x_183, 1, x_182); -lean_ctor_set(x_183, 2, x_181); -x_184 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_185 = lean_array_push(x_184, x_183); -x_186 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; -x_187 = lean_array_push(x_185, x_186); -x_188 = l_Lean_unbracketedExplicitBinders___closed__2; -x_189 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_189, 0, x_178); -lean_ctor_set(x_189, 1, x_188); -lean_ctor_set(x_189, 2, x_187); -x_190 = lean_array_push(x_176, x_189); -x_191 = l_Lean_explicitBinders___closed__2; -x_192 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_192, 0, x_178); -lean_ctor_set(x_192, 1, x_191); -lean_ctor_set(x_192, 2, x_190); -x_193 = l_unexpandListCons___closed__1; -x_194 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_194, 0, x_172); -lean_ctor_set(x_194, 1, x_193); -x_195 = l_unexpandExists___closed__5; -x_196 = lean_array_push(x_195, x_175); -x_197 = lean_array_push(x_196, x_192); -x_198 = lean_array_push(x_197, x_194); -x_199 = lean_array_push(x_198, x_139); -x_200 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_200, 0, x_178); -lean_ctor_set(x_200, 1, x_140); -lean_ctor_set(x_200, 2, x_199); -x_201 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_201, 0, x_200); -lean_ctor_set(x_201, 1, x_173); -return x_201; +lean_inc(x_171); +lean_dec(x_141); +x_173 = l_unexpandExists___closed__3; +lean_inc(x_171); +x_174 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_174, 0, x_171); +lean_ctor_set(x_174, 1, x_173); +x_175 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_176 = lean_array_push(x_175, x_28); +x_177 = lean_box(2); +x_178 = l_unexpandExists___closed__5; +x_179 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_179, 0, x_177); +lean_ctor_set(x_179, 1, x_178); +lean_ctor_set(x_179, 2, x_176); +x_180 = lean_array_push(x_175, x_179); +x_181 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_182 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_182, 0, x_177); +lean_ctor_set(x_182, 1, x_181); +lean_ctor_set(x_182, 2, x_180); +x_183 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_184 = lean_array_push(x_183, x_182); +x_185 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; +x_186 = lean_array_push(x_184, x_185); +x_187 = l_Lean_unbracketedExplicitBinders___closed__2; +x_188 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_188, 0, x_177); +lean_ctor_set(x_188, 1, x_187); +lean_ctor_set(x_188, 2, x_186); +x_189 = lean_array_push(x_175, x_188); +x_190 = l_Lean_explicitBinders___closed__2; +x_191 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_191, 0, x_177); +lean_ctor_set(x_191, 1, x_190); +lean_ctor_set(x_191, 2, x_189); +x_192 = l_unexpandListCons___closed__1; +x_193 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_193, 0, x_171); +lean_ctor_set(x_193, 1, x_192); +x_194 = l_unexpandExists___closed__7; +x_195 = lean_array_push(x_194, x_174); +x_196 = lean_array_push(x_195, x_191); +x_197 = lean_array_push(x_196, x_193); +x_198 = lean_array_push(x_197, x_138); +x_199 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_199, 0, x_177); +lean_ctor_set(x_199, 1, x_139); +lean_ctor_set(x_199, 2, x_198); +x_200 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_200, 0, x_199); +lean_ctor_set(x_200, 1, x_172); +return x_200; } } else { -lean_object* x_202; lean_object* x_203; uint8_t x_204; -x_202 = l_Lean_Syntax_getArg(x_139, x_8); -x_203 = l_Lean_explicitBinders___closed__2; -lean_inc(x_202); -x_204 = l_Lean_Syntax_isOfKind(x_202, x_203); -if (x_204 == 0) -{ -lean_object* x_205; uint8_t x_206; -lean_dec(x_202); -x_205 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); -x_206 = !lean_is_exclusive(x_205); -if (x_206 == 0) -{ -lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; -x_207 = lean_ctor_get(x_205, 0); -x_208 = l_unexpandExists___closed__3; -lean_inc(x_207); -x_209 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_209, 0, x_207); -lean_ctor_set(x_209, 1, x_208); -x_210 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_211 = lean_array_push(x_210, x_29); -x_212 = lean_box(2); -x_213 = l_Lean_binderIdent___closed__2; -x_214 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_214, 0, x_212); -lean_ctor_set(x_214, 1, x_213); -lean_ctor_set(x_214, 2, x_211); -x_215 = lean_array_push(x_210, x_214); -x_216 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_217 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_217, 0, x_212); -lean_ctor_set(x_217, 1, x_216); -lean_ctor_set(x_217, 2, x_215); -x_218 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_219 = lean_array_push(x_218, x_217); -x_220 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; -x_221 = lean_array_push(x_219, x_220); -x_222 = l_Lean_unbracketedExplicitBinders___closed__2; -x_223 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_223, 0, x_212); -lean_ctor_set(x_223, 1, x_222); -lean_ctor_set(x_223, 2, x_221); -x_224 = lean_array_push(x_210, x_223); -x_225 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_225, 0, x_212); -lean_ctor_set(x_225, 1, x_203); -lean_ctor_set(x_225, 2, x_224); -x_226 = l_unexpandListCons___closed__1; -x_227 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_227, 0, x_207); -lean_ctor_set(x_227, 1, x_226); -x_228 = l_unexpandExists___closed__5; -x_229 = lean_array_push(x_228, x_209); -x_230 = lean_array_push(x_229, x_225); -x_231 = lean_array_push(x_230, x_227); -x_232 = lean_array_push(x_231, x_139); -x_233 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_233, 0, x_212); -lean_ctor_set(x_233, 1, x_140); -lean_ctor_set(x_233, 2, x_232); -lean_ctor_set(x_205, 0, x_233); -return x_205; +lean_object* x_201; lean_object* x_202; uint8_t x_203; +x_201 = l_Lean_Syntax_getArg(x_138, x_8); +x_202 = l_Lean_explicitBinders___closed__2; +lean_inc(x_201); +x_203 = l_Lean_Syntax_isOfKind(x_201, x_202); +if (x_203 == 0) +{ +lean_object* x_204; uint8_t x_205; +lean_dec(x_201); +x_204 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); +x_205 = !lean_is_exclusive(x_204); +if (x_205 == 0) +{ +lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; +x_206 = lean_ctor_get(x_204, 0); +x_207 = l_unexpandExists___closed__3; +lean_inc(x_206); +x_208 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_208, 0, x_206); +lean_ctor_set(x_208, 1, x_207); +x_209 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_210 = lean_array_push(x_209, x_28); +x_211 = lean_box(2); +x_212 = l_unexpandExists___closed__5; +x_213 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_213, 0, x_211); +lean_ctor_set(x_213, 1, x_212); +lean_ctor_set(x_213, 2, x_210); +x_214 = lean_array_push(x_209, x_213); +x_215 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_216 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_216, 0, x_211); +lean_ctor_set(x_216, 1, x_215); +lean_ctor_set(x_216, 2, x_214); +x_217 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_218 = lean_array_push(x_217, x_216); +x_219 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; +x_220 = lean_array_push(x_218, x_219); +x_221 = l_Lean_unbracketedExplicitBinders___closed__2; +x_222 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_222, 0, x_211); +lean_ctor_set(x_222, 1, x_221); +lean_ctor_set(x_222, 2, x_220); +x_223 = lean_array_push(x_209, x_222); +x_224 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_224, 0, x_211); +lean_ctor_set(x_224, 1, x_202); +lean_ctor_set(x_224, 2, x_223); +x_225 = l_unexpandListCons___closed__1; +x_226 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_226, 0, x_206); +lean_ctor_set(x_226, 1, x_225); +x_227 = l_unexpandExists___closed__7; +x_228 = lean_array_push(x_227, x_208); +x_229 = lean_array_push(x_228, x_224); +x_230 = lean_array_push(x_229, x_226); +x_231 = lean_array_push(x_230, x_138); +x_232 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_232, 0, x_211); +lean_ctor_set(x_232, 1, x_139); +lean_ctor_set(x_232, 2, x_231); +lean_ctor_set(x_204, 0, x_232); +return x_204; } else { -lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; -x_234 = lean_ctor_get(x_205, 0); -x_235 = lean_ctor_get(x_205, 1); -lean_inc(x_235); +lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; +x_233 = lean_ctor_get(x_204, 0); +x_234 = lean_ctor_get(x_204, 1); lean_inc(x_234); -lean_dec(x_205); -x_236 = l_unexpandExists___closed__3; -lean_inc(x_234); -x_237 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_237, 0, x_234); -lean_ctor_set(x_237, 1, x_236); -x_238 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_239 = lean_array_push(x_238, x_29); -x_240 = lean_box(2); -x_241 = l_Lean_binderIdent___closed__2; -x_242 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_242, 0, x_240); -lean_ctor_set(x_242, 1, x_241); -lean_ctor_set(x_242, 2, x_239); -x_243 = lean_array_push(x_238, x_242); -x_244 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_245 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_245, 0, x_240); -lean_ctor_set(x_245, 1, x_244); -lean_ctor_set(x_245, 2, x_243); -x_246 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_247 = lean_array_push(x_246, x_245); -x_248 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; -x_249 = lean_array_push(x_247, x_248); -x_250 = l_Lean_unbracketedExplicitBinders___closed__2; -x_251 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_251, 0, x_240); -lean_ctor_set(x_251, 1, x_250); -lean_ctor_set(x_251, 2, x_249); -x_252 = lean_array_push(x_238, x_251); -x_253 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_253, 0, x_240); -lean_ctor_set(x_253, 1, x_203); -lean_ctor_set(x_253, 2, x_252); -x_254 = l_unexpandListCons___closed__1; -x_255 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_255, 0, x_234); -lean_ctor_set(x_255, 1, x_254); -x_256 = l_unexpandExists___closed__5; -x_257 = lean_array_push(x_256, x_237); -x_258 = lean_array_push(x_257, x_253); -x_259 = lean_array_push(x_258, x_255); -x_260 = lean_array_push(x_259, x_139); -x_261 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_261, 0, x_240); -lean_ctor_set(x_261, 1, x_140); -lean_ctor_set(x_261, 2, x_260); -x_262 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_262, 0, x_261); -lean_ctor_set(x_262, 1, x_235); -return x_262; +lean_inc(x_233); +lean_dec(x_204); +x_235 = l_unexpandExists___closed__3; +lean_inc(x_233); +x_236 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_236, 0, x_233); +lean_ctor_set(x_236, 1, x_235); +x_237 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_238 = lean_array_push(x_237, x_28); +x_239 = lean_box(2); +x_240 = l_unexpandExists___closed__5; +x_241 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_241, 0, x_239); +lean_ctor_set(x_241, 1, x_240); +lean_ctor_set(x_241, 2, x_238); +x_242 = lean_array_push(x_237, x_241); +x_243 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_244 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_244, 0, x_239); +lean_ctor_set(x_244, 1, x_243); +lean_ctor_set(x_244, 2, x_242); +x_245 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_246 = lean_array_push(x_245, x_244); +x_247 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; +x_248 = lean_array_push(x_246, x_247); +x_249 = l_Lean_unbracketedExplicitBinders___closed__2; +x_250 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_250, 0, x_239); +lean_ctor_set(x_250, 1, x_249); +lean_ctor_set(x_250, 2, x_248); +x_251 = lean_array_push(x_237, x_250); +x_252 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_252, 0, x_239); +lean_ctor_set(x_252, 1, x_202); +lean_ctor_set(x_252, 2, x_251); +x_253 = l_unexpandListCons___closed__1; +x_254 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_254, 0, x_233); +lean_ctor_set(x_254, 1, x_253); +x_255 = l_unexpandExists___closed__7; +x_256 = lean_array_push(x_255, x_236); +x_257 = lean_array_push(x_256, x_252); +x_258 = lean_array_push(x_257, x_254); +x_259 = lean_array_push(x_258, x_138); +x_260 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_260, 0, x_239); +lean_ctor_set(x_260, 1, x_139); +lean_ctor_set(x_260, 2, x_259); +x_261 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_261, 0, x_260); +lean_ctor_set(x_261, 1, x_234); +return x_261; } } else { -lean_object* x_263; lean_object* x_264; uint8_t x_265; -x_263 = l_Lean_Syntax_getArg(x_202, x_14); -lean_dec(x_202); -x_264 = l_Lean_unbracketedExplicitBinders___closed__2; -lean_inc(x_263); -x_265 = l_Lean_Syntax_isOfKind(x_263, x_264); -if (x_265 == 0) -{ -lean_object* x_266; uint8_t x_267; -lean_dec(x_263); -x_266 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); -x_267 = !lean_is_exclusive(x_266); -if (x_267 == 0) -{ -lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; -x_268 = lean_ctor_get(x_266, 0); -x_269 = l_unexpandExists___closed__3; -lean_inc(x_268); -x_270 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_270, 0, x_268); -lean_ctor_set(x_270, 1, x_269); -x_271 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_272 = lean_array_push(x_271, x_29); -x_273 = lean_box(2); -x_274 = l_Lean_binderIdent___closed__2; -x_275 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_275, 0, x_273); -lean_ctor_set(x_275, 1, x_274); -lean_ctor_set(x_275, 2, x_272); -x_276 = lean_array_push(x_271, x_275); -x_277 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_278 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_278, 0, x_273); -lean_ctor_set(x_278, 1, x_277); -lean_ctor_set(x_278, 2, x_276); -x_279 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_280 = lean_array_push(x_279, x_278); -x_281 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; -x_282 = lean_array_push(x_280, x_281); -x_283 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_283, 0, x_273); -lean_ctor_set(x_283, 1, x_264); -lean_ctor_set(x_283, 2, x_282); -x_284 = lean_array_push(x_271, x_283); -x_285 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_285, 0, x_273); -lean_ctor_set(x_285, 1, x_203); -lean_ctor_set(x_285, 2, x_284); -x_286 = l_unexpandListCons___closed__1; -x_287 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_287, 0, x_268); -lean_ctor_set(x_287, 1, x_286); -x_288 = l_unexpandExists___closed__5; -x_289 = lean_array_push(x_288, x_270); -x_290 = lean_array_push(x_289, x_285); -x_291 = lean_array_push(x_290, x_287); -x_292 = lean_array_push(x_291, x_139); -x_293 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_293, 0, x_273); -lean_ctor_set(x_293, 1, x_140); -lean_ctor_set(x_293, 2, x_292); -lean_ctor_set(x_266, 0, x_293); -return x_266; +lean_object* x_262; lean_object* x_263; uint8_t x_264; +x_262 = l_Lean_Syntax_getArg(x_201, x_13); +lean_dec(x_201); +x_263 = l_Lean_unbracketedExplicitBinders___closed__2; +lean_inc(x_262); +x_264 = l_Lean_Syntax_isOfKind(x_262, x_263); +if (x_264 == 0) +{ +lean_object* x_265; uint8_t x_266; +lean_dec(x_262); +x_265 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); +x_266 = !lean_is_exclusive(x_265); +if (x_266 == 0) +{ +lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; +x_267 = lean_ctor_get(x_265, 0); +x_268 = l_unexpandExists___closed__3; +lean_inc(x_267); +x_269 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_269, 0, x_267); +lean_ctor_set(x_269, 1, x_268); +x_270 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_271 = lean_array_push(x_270, x_28); +x_272 = lean_box(2); +x_273 = l_unexpandExists___closed__5; +x_274 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_274, 0, x_272); +lean_ctor_set(x_274, 1, x_273); +lean_ctor_set(x_274, 2, x_271); +x_275 = lean_array_push(x_270, x_274); +x_276 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_277 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_277, 0, x_272); +lean_ctor_set(x_277, 1, x_276); +lean_ctor_set(x_277, 2, x_275); +x_278 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_279 = lean_array_push(x_278, x_277); +x_280 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; +x_281 = lean_array_push(x_279, x_280); +x_282 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_282, 0, x_272); +lean_ctor_set(x_282, 1, x_263); +lean_ctor_set(x_282, 2, x_281); +x_283 = lean_array_push(x_270, x_282); +x_284 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_284, 0, x_272); +lean_ctor_set(x_284, 1, x_202); +lean_ctor_set(x_284, 2, x_283); +x_285 = l_unexpandListCons___closed__1; +x_286 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_286, 0, x_267); +lean_ctor_set(x_286, 1, x_285); +x_287 = l_unexpandExists___closed__7; +x_288 = lean_array_push(x_287, x_269); +x_289 = lean_array_push(x_288, x_284); +x_290 = lean_array_push(x_289, x_286); +x_291 = lean_array_push(x_290, x_138); +x_292 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_292, 0, x_272); +lean_ctor_set(x_292, 1, x_139); +lean_ctor_set(x_292, 2, x_291); +lean_ctor_set(x_265, 0, x_292); +return x_265; } else { -lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; -x_294 = lean_ctor_get(x_266, 0); -x_295 = lean_ctor_get(x_266, 1); -lean_inc(x_295); -lean_inc(x_294); -lean_dec(x_266); -x_296 = l_unexpandExists___closed__3; +lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; +x_293 = lean_ctor_get(x_265, 0); +x_294 = lean_ctor_get(x_265, 1); lean_inc(x_294); -x_297 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_297, 0, x_294); -lean_ctor_set(x_297, 1, x_296); -x_298 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_299 = lean_array_push(x_298, x_29); -x_300 = lean_box(2); -x_301 = l_Lean_binderIdent___closed__2; -x_302 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_302, 0, x_300); -lean_ctor_set(x_302, 1, x_301); -lean_ctor_set(x_302, 2, x_299); -x_303 = lean_array_push(x_298, x_302); -x_304 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_305 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_305, 0, x_300); -lean_ctor_set(x_305, 1, x_304); -lean_ctor_set(x_305, 2, x_303); -x_306 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_307 = lean_array_push(x_306, x_305); -x_308 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; -x_309 = lean_array_push(x_307, x_308); -x_310 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_310, 0, x_300); -lean_ctor_set(x_310, 1, x_264); -lean_ctor_set(x_310, 2, x_309); -x_311 = lean_array_push(x_298, x_310); -x_312 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_312, 0, x_300); -lean_ctor_set(x_312, 1, x_203); -lean_ctor_set(x_312, 2, x_311); -x_313 = l_unexpandListCons___closed__1; -x_314 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_314, 0, x_294); -lean_ctor_set(x_314, 1, x_313); -x_315 = l_unexpandExists___closed__5; -x_316 = lean_array_push(x_315, x_297); -x_317 = lean_array_push(x_316, x_312); -x_318 = lean_array_push(x_317, x_314); -x_319 = lean_array_push(x_318, x_139); -x_320 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_320, 0, x_300); -lean_ctor_set(x_320, 1, x_140); -lean_ctor_set(x_320, 2, x_319); -x_321 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_321, 0, x_320); -lean_ctor_set(x_321, 1, x_295); -return x_321; +lean_inc(x_293); +lean_dec(x_265); +x_295 = l_unexpandExists___closed__3; +lean_inc(x_293); +x_296 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_296, 0, x_293); +lean_ctor_set(x_296, 1, x_295); +x_297 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_298 = lean_array_push(x_297, x_28); +x_299 = lean_box(2); +x_300 = l_unexpandExists___closed__5; +x_301 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_301, 0, x_299); +lean_ctor_set(x_301, 1, x_300); +lean_ctor_set(x_301, 2, x_298); +x_302 = lean_array_push(x_297, x_301); +x_303 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_304 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_304, 0, x_299); +lean_ctor_set(x_304, 1, x_303); +lean_ctor_set(x_304, 2, x_302); +x_305 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_306 = lean_array_push(x_305, x_304); +x_307 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; +x_308 = lean_array_push(x_306, x_307); +x_309 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_309, 0, x_299); +lean_ctor_set(x_309, 1, x_263); +lean_ctor_set(x_309, 2, x_308); +x_310 = lean_array_push(x_297, x_309); +x_311 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_311, 0, x_299); +lean_ctor_set(x_311, 1, x_202); +lean_ctor_set(x_311, 2, x_310); +x_312 = l_unexpandListCons___closed__1; +x_313 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_313, 0, x_293); +lean_ctor_set(x_313, 1, x_312); +x_314 = l_unexpandExists___closed__7; +x_315 = lean_array_push(x_314, x_296); +x_316 = lean_array_push(x_315, x_311); +x_317 = lean_array_push(x_316, x_313); +x_318 = lean_array_push(x_317, x_138); +x_319 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_319, 0, x_299); +lean_ctor_set(x_319, 1, x_139); +lean_ctor_set(x_319, 2, x_318); +x_320 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_320, 0, x_319); +lean_ctor_set(x_320, 1, x_294); +return x_320; } } else { -lean_object* x_322; lean_object* x_323; uint8_t x_324; -x_322 = l_Lean_Syntax_getArg(x_263, x_14); -x_323 = l_Lean_Syntax_getArg(x_263, x_8); -lean_dec(x_263); -x_324 = l_Lean_Syntax_isNodeOf(x_323, x_10, x_14); -if (x_324 == 0) -{ -lean_object* x_325; uint8_t x_326; -lean_dec(x_322); -x_325 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); -x_326 = !lean_is_exclusive(x_325); -if (x_326 == 0) -{ -lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; -x_327 = lean_ctor_get(x_325, 0); -x_328 = l_unexpandExists___closed__3; -lean_inc(x_327); -x_329 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_329, 0, x_327); -lean_ctor_set(x_329, 1, x_328); -x_330 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_331 = lean_array_push(x_330, x_29); -x_332 = lean_box(2); -x_333 = l_Lean_binderIdent___closed__2; -x_334 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_334, 0, x_332); -lean_ctor_set(x_334, 1, x_333); -lean_ctor_set(x_334, 2, x_331); -x_335 = lean_array_push(x_330, x_334); -x_336 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_337 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_337, 0, x_332); -lean_ctor_set(x_337, 1, x_336); -lean_ctor_set(x_337, 2, x_335); -x_338 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_339 = lean_array_push(x_338, x_337); -x_340 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; -x_341 = lean_array_push(x_339, x_340); -x_342 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_342, 0, x_332); -lean_ctor_set(x_342, 1, x_264); -lean_ctor_set(x_342, 2, x_341); -x_343 = lean_array_push(x_330, x_342); -x_344 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_344, 0, x_332); -lean_ctor_set(x_344, 1, x_203); -lean_ctor_set(x_344, 2, x_343); -x_345 = l_unexpandListCons___closed__1; -x_346 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_346, 0, x_327); -lean_ctor_set(x_346, 1, x_345); -x_347 = l_unexpandExists___closed__5; -x_348 = lean_array_push(x_347, x_329); -x_349 = lean_array_push(x_348, x_344); -x_350 = lean_array_push(x_349, x_346); -x_351 = lean_array_push(x_350, x_139); -x_352 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_352, 0, x_332); -lean_ctor_set(x_352, 1, x_140); -lean_ctor_set(x_352, 2, x_351); -lean_ctor_set(x_325, 0, x_352); -return x_325; +lean_object* x_321; lean_object* x_322; uint8_t x_323; +x_321 = l_Lean_Syntax_getArg(x_262, x_13); +x_322 = l_Lean_Syntax_getArg(x_262, x_8); +lean_dec(x_262); +x_323 = l_Lean_Syntax_matchesNull(x_322, x_13); +if (x_323 == 0) +{ +lean_object* x_324; uint8_t x_325; +lean_dec(x_321); +x_324 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); +x_325 = !lean_is_exclusive(x_324); +if (x_325 == 0) +{ +lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; +x_326 = lean_ctor_get(x_324, 0); +x_327 = l_unexpandExists___closed__3; +lean_inc(x_326); +x_328 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_328, 0, x_326); +lean_ctor_set(x_328, 1, x_327); +x_329 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_330 = lean_array_push(x_329, x_28); +x_331 = lean_box(2); +x_332 = l_unexpandExists___closed__5; +x_333 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_333, 0, x_331); +lean_ctor_set(x_333, 1, x_332); +lean_ctor_set(x_333, 2, x_330); +x_334 = lean_array_push(x_329, x_333); +x_335 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_336 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_336, 0, x_331); +lean_ctor_set(x_336, 1, x_335); +lean_ctor_set(x_336, 2, x_334); +x_337 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_338 = lean_array_push(x_337, x_336); +x_339 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; +x_340 = lean_array_push(x_338, x_339); +x_341 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_341, 0, x_331); +lean_ctor_set(x_341, 1, x_263); +lean_ctor_set(x_341, 2, x_340); +x_342 = lean_array_push(x_329, x_341); +x_343 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_343, 0, x_331); +lean_ctor_set(x_343, 1, x_202); +lean_ctor_set(x_343, 2, x_342); +x_344 = l_unexpandListCons___closed__1; +x_345 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_345, 0, x_326); +lean_ctor_set(x_345, 1, x_344); +x_346 = l_unexpandExists___closed__7; +x_347 = lean_array_push(x_346, x_328); +x_348 = lean_array_push(x_347, x_343); +x_349 = lean_array_push(x_348, x_345); +x_350 = lean_array_push(x_349, x_138); +x_351 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_351, 0, x_331); +lean_ctor_set(x_351, 1, x_139); +lean_ctor_set(x_351, 2, x_350); +lean_ctor_set(x_324, 0, x_351); +return x_324; } else { -lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; -x_353 = lean_ctor_get(x_325, 0); -x_354 = lean_ctor_get(x_325, 1); -lean_inc(x_354); +lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; +x_352 = lean_ctor_get(x_324, 0); +x_353 = lean_ctor_get(x_324, 1); lean_inc(x_353); -lean_dec(x_325); -x_355 = l_unexpandExists___closed__3; -lean_inc(x_353); -x_356 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_356, 0, x_353); -lean_ctor_set(x_356, 1, x_355); -x_357 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_358 = lean_array_push(x_357, x_29); -x_359 = lean_box(2); -x_360 = l_Lean_binderIdent___closed__2; -x_361 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_361, 0, x_359); -lean_ctor_set(x_361, 1, x_360); -lean_ctor_set(x_361, 2, x_358); -x_362 = lean_array_push(x_357, x_361); -x_363 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_364 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_364, 0, x_359); -lean_ctor_set(x_364, 1, x_363); -lean_ctor_set(x_364, 2, x_362); -x_365 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_366 = lean_array_push(x_365, x_364); -x_367 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; -x_368 = lean_array_push(x_366, x_367); -x_369 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_369, 0, x_359); -lean_ctor_set(x_369, 1, x_264); -lean_ctor_set(x_369, 2, x_368); -x_370 = lean_array_push(x_357, x_369); -x_371 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_371, 0, x_359); -lean_ctor_set(x_371, 1, x_203); -lean_ctor_set(x_371, 2, x_370); -x_372 = l_unexpandListCons___closed__1; -x_373 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_373, 0, x_353); -lean_ctor_set(x_373, 1, x_372); -x_374 = l_unexpandExists___closed__5; -x_375 = lean_array_push(x_374, x_356); -x_376 = lean_array_push(x_375, x_371); -x_377 = lean_array_push(x_376, x_373); -x_378 = lean_array_push(x_377, x_139); -x_379 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_379, 0, x_359); -lean_ctor_set(x_379, 1, x_140); -lean_ctor_set(x_379, 2, x_378); -x_380 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_380, 0, x_379); -lean_ctor_set(x_380, 1, x_354); -return x_380; +lean_inc(x_352); +lean_dec(x_324); +x_354 = l_unexpandExists___closed__3; +lean_inc(x_352); +x_355 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_355, 0, x_352); +lean_ctor_set(x_355, 1, x_354); +x_356 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_357 = lean_array_push(x_356, x_28); +x_358 = lean_box(2); +x_359 = l_unexpandExists___closed__5; +x_360 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_360, 0, x_358); +lean_ctor_set(x_360, 1, x_359); +lean_ctor_set(x_360, 2, x_357); +x_361 = lean_array_push(x_356, x_360); +x_362 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_363 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_363, 0, x_358); +lean_ctor_set(x_363, 1, x_362); +lean_ctor_set(x_363, 2, x_361); +x_364 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_365 = lean_array_push(x_364, x_363); +x_366 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; +x_367 = lean_array_push(x_365, x_366); +x_368 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_368, 0, x_358); +lean_ctor_set(x_368, 1, x_263); +lean_ctor_set(x_368, 2, x_367); +x_369 = lean_array_push(x_356, x_368); +x_370 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_370, 0, x_358); +lean_ctor_set(x_370, 1, x_202); +lean_ctor_set(x_370, 2, x_369); +x_371 = l_unexpandListCons___closed__1; +x_372 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_372, 0, x_352); +lean_ctor_set(x_372, 1, x_371); +x_373 = l_unexpandExists___closed__7; +x_374 = lean_array_push(x_373, x_355); +x_375 = lean_array_push(x_374, x_370); +x_376 = lean_array_push(x_375, x_372); +x_377 = lean_array_push(x_376, x_138); +x_378 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_378, 0, x_358); +lean_ctor_set(x_378, 1, x_139); +lean_ctor_set(x_378, 2, x_377); +x_379 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_379, 0, x_378); +lean_ctor_set(x_379, 1, x_353); +return x_379; } } else { -lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; uint8_t x_385; -x_381 = lean_unsigned_to_nat(3u); -x_382 = l_Lean_Syntax_getArg(x_139, x_381); -lean_dec(x_139); -x_383 = l_Lean_Syntax_getArgs(x_322); -lean_dec(x_322); -x_384 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); -x_385 = !lean_is_exclusive(x_384); -if (x_385 == 0) -{ -lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; -x_386 = lean_ctor_get(x_384, 0); -x_387 = l_unexpandExists___closed__3; -lean_inc(x_386); -x_388 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_388, 0, x_386); -lean_ctor_set(x_388, 1, x_387); -x_389 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_390 = lean_array_push(x_389, x_29); -x_391 = lean_box(2); -x_392 = l_Lean_binderIdent___closed__2; -x_393 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_393, 0, x_391); -lean_ctor_set(x_393, 1, x_392); -lean_ctor_set(x_393, 2, x_390); -x_394 = lean_array_push(x_389, x_393); -x_395 = l_Array_append___rarg(x_394, x_383); -x_396 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_397 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_397, 0, x_391); -lean_ctor_set(x_397, 1, x_396); -lean_ctor_set(x_397, 2, x_395); -x_398 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_399 = lean_array_push(x_398, x_397); -x_400 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; -x_401 = lean_array_push(x_399, x_400); -x_402 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_402, 0, x_391); -lean_ctor_set(x_402, 1, x_264); -lean_ctor_set(x_402, 2, x_401); -x_403 = lean_array_push(x_389, x_402); -x_404 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_404, 0, x_391); -lean_ctor_set(x_404, 1, x_203); -lean_ctor_set(x_404, 2, x_403); -x_405 = l_unexpandListCons___closed__1; -x_406 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_406, 0, x_386); -lean_ctor_set(x_406, 1, x_405); -x_407 = l_unexpandExists___closed__5; -x_408 = lean_array_push(x_407, x_388); -x_409 = lean_array_push(x_408, x_404); -x_410 = lean_array_push(x_409, x_406); -x_411 = lean_array_push(x_410, x_382); -x_412 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_412, 0, x_391); -lean_ctor_set(x_412, 1, x_140); -lean_ctor_set(x_412, 2, x_411); -lean_ctor_set(x_384, 0, x_412); -return x_384; +lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; uint8_t x_384; +x_380 = lean_unsigned_to_nat(3u); +x_381 = l_Lean_Syntax_getArg(x_138, x_380); +lean_dec(x_138); +x_382 = l_Lean_Syntax_getArgs(x_321); +lean_dec(x_321); +x_383 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); +x_384 = !lean_is_exclusive(x_383); +if (x_384 == 0) +{ +lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; +x_385 = lean_ctor_get(x_383, 0); +x_386 = l_unexpandExists___closed__3; +lean_inc(x_385); +x_387 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_387, 0, x_385); +lean_ctor_set(x_387, 1, x_386); +x_388 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_389 = lean_array_push(x_388, x_28); +x_390 = lean_box(2); +x_391 = l_unexpandExists___closed__5; +x_392 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_392, 0, x_390); +lean_ctor_set(x_392, 1, x_391); +lean_ctor_set(x_392, 2, x_389); +x_393 = lean_array_push(x_388, x_392); +x_394 = l_Array_append___rarg(x_393, x_382); +x_395 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_396 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_396, 0, x_390); +lean_ctor_set(x_396, 1, x_395); +lean_ctor_set(x_396, 2, x_394); +x_397 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_398 = lean_array_push(x_397, x_396); +x_399 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; +x_400 = lean_array_push(x_398, x_399); +x_401 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_401, 0, x_390); +lean_ctor_set(x_401, 1, x_263); +lean_ctor_set(x_401, 2, x_400); +x_402 = lean_array_push(x_388, x_401); +x_403 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_403, 0, x_390); +lean_ctor_set(x_403, 1, x_202); +lean_ctor_set(x_403, 2, x_402); +x_404 = l_unexpandListCons___closed__1; +x_405 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_405, 0, x_385); +lean_ctor_set(x_405, 1, x_404); +x_406 = l_unexpandExists___closed__7; +x_407 = lean_array_push(x_406, x_387); +x_408 = lean_array_push(x_407, x_403); +x_409 = lean_array_push(x_408, x_405); +x_410 = lean_array_push(x_409, x_381); +x_411 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_411, 0, x_390); +lean_ctor_set(x_411, 1, x_139); +lean_ctor_set(x_411, 2, x_410); +lean_ctor_set(x_383, 0, x_411); +return x_383; } else { -lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; -x_413 = lean_ctor_get(x_384, 0); -x_414 = lean_ctor_get(x_384, 1); -lean_inc(x_414); -lean_inc(x_413); -lean_dec(x_384); -x_415 = l_unexpandExists___closed__3; +lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; +x_412 = lean_ctor_get(x_383, 0); +x_413 = lean_ctor_get(x_383, 1); lean_inc(x_413); -x_416 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_416, 0, x_413); -lean_ctor_set(x_416, 1, x_415); -x_417 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_418 = lean_array_push(x_417, x_29); -x_419 = lean_box(2); -x_420 = l_Lean_binderIdent___closed__2; -x_421 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_421, 0, x_419); -lean_ctor_set(x_421, 1, x_420); -lean_ctor_set(x_421, 2, x_418); -x_422 = lean_array_push(x_417, x_421); -x_423 = l_Array_append___rarg(x_422, x_383); -x_424 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_425 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_425, 0, x_419); -lean_ctor_set(x_425, 1, x_424); -lean_ctor_set(x_425, 2, x_423); -x_426 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_427 = lean_array_push(x_426, x_425); -x_428 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; -x_429 = lean_array_push(x_427, x_428); -x_430 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_430, 0, x_419); -lean_ctor_set(x_430, 1, x_264); -lean_ctor_set(x_430, 2, x_429); -x_431 = lean_array_push(x_417, x_430); -x_432 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_432, 0, x_419); -lean_ctor_set(x_432, 1, x_203); -lean_ctor_set(x_432, 2, x_431); -x_433 = l_unexpandListCons___closed__1; -x_434 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_434, 0, x_413); -lean_ctor_set(x_434, 1, x_433); -x_435 = l_unexpandExists___closed__5; -x_436 = lean_array_push(x_435, x_416); -x_437 = lean_array_push(x_436, x_432); -x_438 = lean_array_push(x_437, x_434); -x_439 = lean_array_push(x_438, x_382); -x_440 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_440, 0, x_419); -lean_ctor_set(x_440, 1, x_140); -lean_ctor_set(x_440, 2, x_439); -x_441 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_441, 0, x_440); -lean_ctor_set(x_441, 1, x_414); -return x_441; +lean_inc(x_412); +lean_dec(x_383); +x_414 = l_unexpandExists___closed__3; +lean_inc(x_412); +x_415 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_415, 0, x_412); +lean_ctor_set(x_415, 1, x_414); +x_416 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_417 = lean_array_push(x_416, x_28); +x_418 = lean_box(2); +x_419 = l_unexpandExists___closed__5; +x_420 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_420, 0, x_418); +lean_ctor_set(x_420, 1, x_419); +lean_ctor_set(x_420, 2, x_417); +x_421 = lean_array_push(x_416, x_420); +x_422 = l_Array_append___rarg(x_421, x_382); +x_423 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_424 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_424, 0, x_418); +lean_ctor_set(x_424, 1, x_423); +lean_ctor_set(x_424, 2, x_422); +x_425 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_426 = lean_array_push(x_425, x_424); +x_427 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; +x_428 = lean_array_push(x_426, x_427); +x_429 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_429, 0, x_418); +lean_ctor_set(x_429, 1, x_263); +lean_ctor_set(x_429, 2, x_428); +x_430 = lean_array_push(x_416, x_429); +x_431 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_431, 0, x_418); +lean_ctor_set(x_431, 1, x_202); +lean_ctor_set(x_431, 2, x_430); +x_432 = l_unexpandListCons___closed__1; +x_433 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_433, 0, x_412); +lean_ctor_set(x_433, 1, x_432); +x_434 = l_unexpandExists___closed__7; +x_435 = lean_array_push(x_434, x_415); +x_436 = lean_array_push(x_435, x_431); +x_437 = lean_array_push(x_436, x_433); +x_438 = lean_array_push(x_437, x_381); +x_439 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_439, 0, x_418); +lean_ctor_set(x_439, 1, x_139); +lean_ctor_set(x_439, 2, x_438); +x_440 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_440, 0, x_439); +lean_ctor_set(x_440, 1, x_413); +return x_440; } } } @@ -9389,7 +9251,7 @@ LEAN_EXPORT lean_object* l_unexpandSigma(lean_object* x_1, lean_object* x_2, lea _start: { lean_object* x_4; uint8_t x_5; -x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__6; +x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__6; lean_inc(x_1); x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); if (x_5 == 0) @@ -9405,316 +9267,315 @@ return x_7; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +lean_object* x_8; lean_object* x_9; uint8_t x_10; x_8 = lean_unsigned_to_nat(1u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); lean_dec(x_1); -x_10 = l_Lean_nullKind; lean_inc(x_9); -x_11 = l_Lean_Syntax_isNodeOf(x_9, x_10, x_8); -if (x_11 == 0) +x_10 = l_Lean_Syntax_matchesNull(x_9, x_8); +if (x_10 == 0) { -lean_object* x_12; lean_object* x_13; +lean_object* x_11; lean_object* x_12; lean_dec(x_9); lean_dec(x_2); -x_12 = lean_box(0); -x_13 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_3); -return x_13; +x_11 = lean_box(0); +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_3); +return x_12; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_14 = lean_unsigned_to_nat(0u); -x_15 = l_Lean_Syntax_getArg(x_9, x_14); +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_13 = lean_unsigned_to_nat(0u); +x_14 = l_Lean_Syntax_getArg(x_9, x_13); lean_dec(x_9); -x_16 = l_Lean_expandExplicitBindersAux_loop___closed__2; -lean_inc(x_15); -x_17 = l_Lean_Syntax_isOfKind(x_15, x_16); -if (x_17 == 0) +x_15 = l_Lean_expandExplicitBindersAux_loop___closed__2; +lean_inc(x_14); +x_16 = l_Lean_Syntax_isOfKind(x_14, x_15); +if (x_16 == 0) { -lean_object* x_18; lean_object* x_19; -lean_dec(x_15); +lean_object* x_17; lean_object* x_18; +lean_dec(x_14); lean_dec(x_2); -x_18 = lean_box(0); -x_19 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_3); -return x_19; +x_17 = lean_box(0); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_3); +return x_18; } else { -lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_20 = l_Lean_Syntax_getArg(x_15, x_8); -lean_dec(x_15); -x_21 = l_Lean_expandExplicitBindersAux_loop___closed__4; -lean_inc(x_20); -x_22 = l_Lean_Syntax_isOfKind(x_20, x_21); -if (x_22 == 0) +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = l_Lean_Syntax_getArg(x_14, x_8); +lean_dec(x_14); +x_20 = l_Lean_expandExplicitBindersAux_loop___closed__4; +lean_inc(x_19); +x_21 = l_Lean_Syntax_isOfKind(x_19, x_20); +if (x_21 == 0) { -lean_object* x_23; lean_object* x_24; -lean_dec(x_20); +lean_object* x_22; lean_object* x_23; +lean_dec(x_19); lean_dec(x_2); -x_23 = lean_box(0); -x_24 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_3); -return x_24; +x_22 = lean_box(0); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_3); +return x_23; } else { -lean_object* x_25; uint8_t x_26; -x_25 = l_Lean_Syntax_getArg(x_20, x_14); -lean_inc(x_25); -x_26 = l_Lean_Syntax_isNodeOf(x_25, x_10, x_8); -if (x_26 == 0) +lean_object* x_24; uint8_t x_25; +x_24 = l_Lean_Syntax_getArg(x_19, x_13); +lean_inc(x_24); +x_25 = l_Lean_Syntax_matchesNull(x_24, x_8); +if (x_25 == 0) { -lean_object* x_27; lean_object* x_28; -lean_dec(x_25); -lean_dec(x_20); +lean_object* x_26; lean_object* x_27; +lean_dec(x_24); +lean_dec(x_19); lean_dec(x_2); -x_27 = lean_box(0); -x_28 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_28, 0, x_27); -lean_ctor_set(x_28, 1, x_3); -return x_28; +x_26 = lean_box(0); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_3); +return x_27; } else { -lean_object* x_29; lean_object* x_30; uint8_t x_31; -x_29 = l_Lean_Syntax_getArg(x_25, x_14); -lean_dec(x_25); -x_30 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__21; -lean_inc(x_29); -x_31 = l_Lean_Syntax_isOfKind(x_29, x_30); -if (x_31 == 0) +lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_28 = l_Lean_Syntax_getArg(x_24, x_13); +lean_dec(x_24); +x_29 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__21; +lean_inc(x_28); +x_30 = l_Lean_Syntax_isOfKind(x_28, x_29); +if (x_30 == 0) { -lean_object* x_32; lean_object* x_33; -lean_dec(x_29); -lean_dec(x_20); +lean_object* x_31; lean_object* x_32; +lean_dec(x_28); +lean_dec(x_19); lean_dec(x_2); -x_32 = lean_box(0); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_3); -return x_33; +x_31 = lean_box(0); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_3); +return x_32; } else { -lean_object* x_34; lean_object* x_35; uint8_t x_36; -x_34 = l_Lean_Syntax_getArg(x_29, x_8); -lean_dec(x_29); -x_35 = lean_unsigned_to_nat(2u); -lean_inc(x_34); -x_36 = l_Lean_Syntax_isNodeOf(x_34, x_10, x_35); -if (x_36 == 0) +lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_33 = l_Lean_Syntax_getArg(x_28, x_8); +lean_dec(x_28); +x_34 = lean_unsigned_to_nat(2u); +lean_inc(x_33); +x_35 = l_Lean_Syntax_matchesNull(x_33, x_34); +if (x_35 == 0) { -lean_object* x_37; lean_object* x_38; -lean_dec(x_34); -lean_dec(x_20); +lean_object* x_36; lean_object* x_37; +lean_dec(x_33); +lean_dec(x_19); lean_dec(x_2); -x_37 = lean_box(0); -x_38 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_38, 0, x_37); -lean_ctor_set(x_38, 1, x_3); -return x_38; +x_36 = lean_box(0); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_3); +return x_37; } else { -lean_object* x_39; lean_object* x_40; uint8_t x_41; -x_39 = l_Lean_Syntax_getArg(x_34, x_14); -x_40 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__10; -lean_inc(x_39); -x_41 = l_Lean_Syntax_isOfKind(x_39, x_40); -if (x_41 == 0) +lean_object* x_38; lean_object* x_39; uint8_t x_40; +x_38 = l_Lean_Syntax_getArg(x_33, x_13); +x_39 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__10; +lean_inc(x_38); +x_40 = l_Lean_Syntax_isOfKind(x_38, x_39); +if (x_40 == 0) { -lean_object* x_42; lean_object* x_43; -lean_dec(x_39); -lean_dec(x_34); -lean_dec(x_20); +lean_object* x_41; lean_object* x_42; +lean_dec(x_38); +lean_dec(x_33); +lean_dec(x_19); lean_dec(x_2); -x_42 = lean_box(0); -x_43 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_43, 1, x_3); -return x_43; +x_41 = lean_box(0); +x_42 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_3); +return x_42; } else { -lean_object* x_44; uint8_t x_45; -x_44 = l_Lean_Syntax_getArg(x_34, x_8); -lean_dec(x_34); -lean_inc(x_44); -x_45 = l_Lean_Syntax_isNodeOf(x_44, x_10, x_8); -if (x_45 == 0) +lean_object* x_43; uint8_t x_44; +x_43 = l_Lean_Syntax_getArg(x_33, x_8); +lean_dec(x_33); +lean_inc(x_43); +x_44 = l_Lean_Syntax_matchesNull(x_43, x_8); +if (x_44 == 0) { -lean_object* x_46; lean_object* x_47; -lean_dec(x_44); -lean_dec(x_39); -lean_dec(x_20); +lean_object* x_45; lean_object* x_46; +lean_dec(x_43); +lean_dec(x_38); +lean_dec(x_19); lean_dec(x_2); -x_46 = lean_box(0); -x_47 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_3); -return x_47; +x_45 = lean_box(0); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_3); +return x_46; } else { -lean_object* x_48; lean_object* x_49; uint8_t x_50; -x_48 = l_Lean_Syntax_getArg(x_44, x_14); -lean_dec(x_44); -x_49 = l_unexpandExists___closed__2; -lean_inc(x_48); -x_50 = l_Lean_Syntax_isOfKind(x_48, x_49); -if (x_50 == 0) +lean_object* x_47; lean_object* x_48; uint8_t x_49; +x_47 = l_Lean_Syntax_getArg(x_43, x_13); +lean_dec(x_43); +x_48 = l_unexpandExists___closed__2; +lean_inc(x_47); +x_49 = l_Lean_Syntax_isOfKind(x_47, x_48); +if (x_49 == 0) { -lean_object* x_51; lean_object* x_52; -lean_dec(x_48); -lean_dec(x_39); -lean_dec(x_20); +lean_object* x_50; lean_object* x_51; +lean_dec(x_47); +lean_dec(x_38); +lean_dec(x_19); lean_dec(x_2); -x_51 = lean_box(0); -x_52 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_52, 0, x_51); -lean_ctor_set(x_52, 1, x_3); -return x_52; +x_50 = lean_box(0); +x_51 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_3); +return x_51; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; -x_53 = l_Lean_Syntax_getArg(x_48, x_8); -lean_dec(x_48); -x_54 = l_Lean_Syntax_getArg(x_20, x_35); -lean_dec(x_20); -x_55 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); -x_56 = !lean_is_exclusive(x_55); -if (x_56 == 0) +lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_52 = l_Lean_Syntax_getArg(x_47, x_8); +lean_dec(x_47); +x_53 = l_Lean_Syntax_getArg(x_19, x_34); +lean_dec(x_19); +x_54 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); +x_55 = !lean_is_exclusive(x_54); +if (x_55 == 0) { -lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; -x_57 = lean_ctor_get(x_55, 0); -x_58 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__22; -lean_inc(x_57); -x_59 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_59, 0, x_57); -lean_ctor_set(x_59, 1, x_58); -x_60 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_61 = lean_array_push(x_60, x_39); -x_62 = lean_box(2); -x_63 = l_Lean_binderIdent___closed__2; -x_64 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_64, 0, x_62); -lean_ctor_set(x_64, 1, x_63); -lean_ctor_set(x_64, 2, x_61); -x_65 = lean_array_push(x_60, x_64); -x_66 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_67 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_67, 0, x_62); -lean_ctor_set(x_67, 1, x_66); -lean_ctor_set(x_67, 2, x_65); -x_68 = l_Lean_expandExplicitBindersAux_loop___closed__12; -lean_inc(x_57); -x_69 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_69, 0, x_57); -lean_ctor_set(x_69, 1, x_68); -x_70 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__29; -lean_inc(x_57); -x_71 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_71, 0, x_57); -lean_ctor_set(x_71, 1, x_70); -x_72 = l_unexpandExists___closed__4; -x_73 = lean_array_push(x_72, x_59); -x_74 = lean_array_push(x_73, x_67); -x_75 = lean_array_push(x_74, x_69); -x_76 = lean_array_push(x_75, x_53); -x_77 = lean_array_push(x_76, x_71); -x_78 = l_Lean_bracketedExplicitBinders___closed__2; -x_79 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_79, 0, x_62); -lean_ctor_set(x_79, 1, x_78); -lean_ctor_set(x_79, 2, x_77); -x_80 = l_unexpandSigma___closed__1; -x_81 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_81, 0, x_57); -lean_ctor_set(x_81, 1, x_80); -x_82 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; -x_83 = lean_array_push(x_82, x_79); -x_84 = lean_array_push(x_83, x_81); -x_85 = lean_array_push(x_84, x_54); -x_86 = l_term___xd7____1___closed__2; -x_87 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_87, 0, x_62); -lean_ctor_set(x_87, 1, x_86); -lean_ctor_set(x_87, 2, x_85); -lean_ctor_set(x_55, 0, x_87); -return x_55; +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_56 = lean_ctor_get(x_54, 0); +x_57 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__22; +lean_inc(x_56); +x_58 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_58, 0, x_56); +lean_ctor_set(x_58, 1, x_57); +x_59 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_60 = lean_array_push(x_59, x_38); +x_61 = lean_box(2); +x_62 = l_unexpandExists___closed__5; +x_63 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_63, 0, x_61); +lean_ctor_set(x_63, 1, x_62); +lean_ctor_set(x_63, 2, x_60); +x_64 = lean_array_push(x_59, x_63); +x_65 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_66 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_66, 0, x_61); +lean_ctor_set(x_66, 1, x_65); +lean_ctor_set(x_66, 2, x_64); +x_67 = l_Lean_expandExplicitBindersAux_loop___closed__13; +lean_inc(x_56); +x_68 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_68, 0, x_56); +lean_ctor_set(x_68, 1, x_67); +x_69 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__29; +lean_inc(x_56); +x_70 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_70, 0, x_56); +lean_ctor_set(x_70, 1, x_69); +x_71 = l_unexpandExists___closed__6; +x_72 = lean_array_push(x_71, x_58); +x_73 = lean_array_push(x_72, x_66); +x_74 = lean_array_push(x_73, x_68); +x_75 = lean_array_push(x_74, x_52); +x_76 = lean_array_push(x_75, x_70); +x_77 = l_Lean_bracketedExplicitBinders___closed__2; +x_78 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_78, 0, x_61); +lean_ctor_set(x_78, 1, x_77); +lean_ctor_set(x_78, 2, x_76); +x_79 = l_unexpandSigma___closed__1; +x_80 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_80, 0, x_56); +lean_ctor_set(x_80, 1, x_79); +x_81 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; +x_82 = lean_array_push(x_81, x_78); +x_83 = lean_array_push(x_82, x_80); +x_84 = lean_array_push(x_83, x_53); +x_85 = l_term___xd7____1___closed__2; +x_86 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_86, 0, x_61); +lean_ctor_set(x_86, 1, x_85); +lean_ctor_set(x_86, 2, x_84); +lean_ctor_set(x_54, 0, x_86); +return x_54; } else { -lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; -x_88 = lean_ctor_get(x_55, 0); -x_89 = lean_ctor_get(x_55, 1); -lean_inc(x_89); -lean_inc(x_88); -lean_dec(x_55); -x_90 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__22; +lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; +x_87 = lean_ctor_get(x_54, 0); +x_88 = lean_ctor_get(x_54, 1); lean_inc(x_88); -x_91 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_91, 0, x_88); -lean_ctor_set(x_91, 1, x_90); -x_92 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_93 = lean_array_push(x_92, x_39); -x_94 = lean_box(2); -x_95 = l_Lean_binderIdent___closed__2; -x_96 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_96, 0, x_94); -lean_ctor_set(x_96, 1, x_95); -lean_ctor_set(x_96, 2, x_93); -x_97 = lean_array_push(x_92, x_96); -x_98 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_99 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_99, 0, x_94); -lean_ctor_set(x_99, 1, x_98); -lean_ctor_set(x_99, 2, x_97); -x_100 = l_Lean_expandExplicitBindersAux_loop___closed__12; -lean_inc(x_88); -x_101 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_101, 0, x_88); -lean_ctor_set(x_101, 1, x_100); -x_102 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__29; -lean_inc(x_88); -x_103 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_103, 0, x_88); -lean_ctor_set(x_103, 1, x_102); -x_104 = l_unexpandExists___closed__4; -x_105 = lean_array_push(x_104, x_91); -x_106 = lean_array_push(x_105, x_99); -x_107 = lean_array_push(x_106, x_101); -x_108 = lean_array_push(x_107, x_53); -x_109 = lean_array_push(x_108, x_103); -x_110 = l_Lean_bracketedExplicitBinders___closed__2; -x_111 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_111, 0, x_94); -lean_ctor_set(x_111, 1, x_110); -lean_ctor_set(x_111, 2, x_109); -x_112 = l_unexpandSigma___closed__1; -x_113 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_113, 0, x_88); -lean_ctor_set(x_113, 1, x_112); -x_114 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; -x_115 = lean_array_push(x_114, x_111); -x_116 = lean_array_push(x_115, x_113); -x_117 = lean_array_push(x_116, x_54); -x_118 = l_term___xd7____1___closed__2; -x_119 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_119, 0, x_94); -lean_ctor_set(x_119, 1, x_118); -lean_ctor_set(x_119, 2, x_117); -x_120 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_120, 0, x_119); -lean_ctor_set(x_120, 1, x_89); -return x_120; +lean_inc(x_87); +lean_dec(x_54); +x_89 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__22; +lean_inc(x_87); +x_90 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_90, 0, x_87); +lean_ctor_set(x_90, 1, x_89); +x_91 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_92 = lean_array_push(x_91, x_38); +x_93 = lean_box(2); +x_94 = l_unexpandExists___closed__5; +x_95 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_95, 0, x_93); +lean_ctor_set(x_95, 1, x_94); +lean_ctor_set(x_95, 2, x_92); +x_96 = lean_array_push(x_91, x_95); +x_97 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_98 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_98, 0, x_93); +lean_ctor_set(x_98, 1, x_97); +lean_ctor_set(x_98, 2, x_96); +x_99 = l_Lean_expandExplicitBindersAux_loop___closed__13; +lean_inc(x_87); +x_100 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_100, 0, x_87); +lean_ctor_set(x_100, 1, x_99); +x_101 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__29; +lean_inc(x_87); +x_102 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_102, 0, x_87); +lean_ctor_set(x_102, 1, x_101); +x_103 = l_unexpandExists___closed__6; +x_104 = lean_array_push(x_103, x_90); +x_105 = lean_array_push(x_104, x_98); +x_106 = lean_array_push(x_105, x_100); +x_107 = lean_array_push(x_106, x_52); +x_108 = lean_array_push(x_107, x_102); +x_109 = l_Lean_bracketedExplicitBinders___closed__2; +x_110 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_110, 0, x_93); +lean_ctor_set(x_110, 1, x_109); +lean_ctor_set(x_110, 2, x_108); +x_111 = l_unexpandSigma___closed__1; +x_112 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_112, 0, x_87); +lean_ctor_set(x_112, 1, x_111); +x_113 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; +x_114 = lean_array_push(x_113, x_110); +x_115 = lean_array_push(x_114, x_112); +x_116 = lean_array_push(x_115, x_53); +x_117 = l_term___xd7____1___closed__2; +x_118 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_118, 0, x_93); +lean_ctor_set(x_118, 1, x_117); +lean_ctor_set(x_118, 2, x_116); +x_119 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_119, 0, x_118); +lean_ctor_set(x_119, 1, x_88); +return x_119; } } } @@ -9740,7 +9601,7 @@ LEAN_EXPORT lean_object* l_unexpandPSigma(lean_object* x_1, lean_object* x_2, le _start: { lean_object* x_4; uint8_t x_5; -x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__6; +x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__6; lean_inc(x_1); x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); if (x_5 == 0) @@ -9756,316 +9617,766 @@ return x_7; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +lean_object* x_8; lean_object* x_9; uint8_t x_10; x_8 = lean_unsigned_to_nat(1u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); lean_dec(x_1); -x_10 = l_Lean_nullKind; lean_inc(x_9); -x_11 = l_Lean_Syntax_isNodeOf(x_9, x_10, x_8); -if (x_11 == 0) +x_10 = l_Lean_Syntax_matchesNull(x_9, x_8); +if (x_10 == 0) { -lean_object* x_12; lean_object* x_13; +lean_object* x_11; lean_object* x_12; lean_dec(x_9); lean_dec(x_2); -x_12 = lean_box(0); -x_13 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_3); -return x_13; +x_11 = lean_box(0); +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_3); +return x_12; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_14 = lean_unsigned_to_nat(0u); -x_15 = l_Lean_Syntax_getArg(x_9, x_14); +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_13 = lean_unsigned_to_nat(0u); +x_14 = l_Lean_Syntax_getArg(x_9, x_13); lean_dec(x_9); -x_16 = l_Lean_expandExplicitBindersAux_loop___closed__2; -lean_inc(x_15); -x_17 = l_Lean_Syntax_isOfKind(x_15, x_16); -if (x_17 == 0) +x_15 = l_Lean_expandExplicitBindersAux_loop___closed__2; +lean_inc(x_14); +x_16 = l_Lean_Syntax_isOfKind(x_14, x_15); +if (x_16 == 0) { -lean_object* x_18; lean_object* x_19; -lean_dec(x_15); +lean_object* x_17; lean_object* x_18; +lean_dec(x_14); lean_dec(x_2); -x_18 = lean_box(0); -x_19 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_3); -return x_19; +x_17 = lean_box(0); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_3); +return x_18; } else { -lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_20 = l_Lean_Syntax_getArg(x_15, x_8); -lean_dec(x_15); -x_21 = l_Lean_expandExplicitBindersAux_loop___closed__4; -lean_inc(x_20); -x_22 = l_Lean_Syntax_isOfKind(x_20, x_21); -if (x_22 == 0) +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = l_Lean_Syntax_getArg(x_14, x_8); +lean_dec(x_14); +x_20 = l_Lean_expandExplicitBindersAux_loop___closed__4; +lean_inc(x_19); +x_21 = l_Lean_Syntax_isOfKind(x_19, x_20); +if (x_21 == 0) { -lean_object* x_23; lean_object* x_24; -lean_dec(x_20); +lean_object* x_22; lean_object* x_23; +lean_dec(x_19); lean_dec(x_2); -x_23 = lean_box(0); -x_24 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_3); -return x_24; +x_22 = lean_box(0); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_3); +return x_23; } else { -lean_object* x_25; uint8_t x_26; -x_25 = l_Lean_Syntax_getArg(x_20, x_14); -lean_inc(x_25); -x_26 = l_Lean_Syntax_isNodeOf(x_25, x_10, x_8); -if (x_26 == 0) +lean_object* x_24; uint8_t x_25; +x_24 = l_Lean_Syntax_getArg(x_19, x_13); +lean_inc(x_24); +x_25 = l_Lean_Syntax_matchesNull(x_24, x_8); +if (x_25 == 0) { -lean_object* x_27; lean_object* x_28; -lean_dec(x_25); -lean_dec(x_20); +lean_object* x_26; lean_object* x_27; +lean_dec(x_24); +lean_dec(x_19); lean_dec(x_2); -x_27 = lean_box(0); -x_28 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_28, 0, x_27); -lean_ctor_set(x_28, 1, x_3); -return x_28; +x_26 = lean_box(0); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_3); +return x_27; } else { -lean_object* x_29; lean_object* x_30; uint8_t x_31; -x_29 = l_Lean_Syntax_getArg(x_25, x_14); -lean_dec(x_25); -x_30 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__21; -lean_inc(x_29); -x_31 = l_Lean_Syntax_isOfKind(x_29, x_30); -if (x_31 == 0) +lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_28 = l_Lean_Syntax_getArg(x_24, x_13); +lean_dec(x_24); +x_29 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__21; +lean_inc(x_28); +x_30 = l_Lean_Syntax_isOfKind(x_28, x_29); +if (x_30 == 0) { -lean_object* x_32; lean_object* x_33; -lean_dec(x_29); -lean_dec(x_20); +lean_object* x_31; lean_object* x_32; +lean_dec(x_28); +lean_dec(x_19); lean_dec(x_2); -x_32 = lean_box(0); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_3); -return x_33; +x_31 = lean_box(0); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_3); +return x_32; } else { -lean_object* x_34; lean_object* x_35; uint8_t x_36; -x_34 = l_Lean_Syntax_getArg(x_29, x_8); -lean_dec(x_29); -x_35 = lean_unsigned_to_nat(2u); -lean_inc(x_34); -x_36 = l_Lean_Syntax_isNodeOf(x_34, x_10, x_35); -if (x_36 == 0) +lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_33 = l_Lean_Syntax_getArg(x_28, x_8); +lean_dec(x_28); +x_34 = lean_unsigned_to_nat(2u); +lean_inc(x_33); +x_35 = l_Lean_Syntax_matchesNull(x_33, x_34); +if (x_35 == 0) { -lean_object* x_37; lean_object* x_38; -lean_dec(x_34); -lean_dec(x_20); +lean_object* x_36; lean_object* x_37; +lean_dec(x_33); +lean_dec(x_19); lean_dec(x_2); -x_37 = lean_box(0); -x_38 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_38, 0, x_37); -lean_ctor_set(x_38, 1, x_3); -return x_38; +x_36 = lean_box(0); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_3); +return x_37; } else { -lean_object* x_39; lean_object* x_40; uint8_t x_41; -x_39 = l_Lean_Syntax_getArg(x_34, x_14); -x_40 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__10; -lean_inc(x_39); -x_41 = l_Lean_Syntax_isOfKind(x_39, x_40); -if (x_41 == 0) +lean_object* x_38; lean_object* x_39; uint8_t x_40; +x_38 = l_Lean_Syntax_getArg(x_33, x_13); +x_39 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__10; +lean_inc(x_38); +x_40 = l_Lean_Syntax_isOfKind(x_38, x_39); +if (x_40 == 0) { -lean_object* x_42; lean_object* x_43; -lean_dec(x_39); -lean_dec(x_34); -lean_dec(x_20); +lean_object* x_41; lean_object* x_42; +lean_dec(x_38); +lean_dec(x_33); +lean_dec(x_19); lean_dec(x_2); -x_42 = lean_box(0); -x_43 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_43, 1, x_3); -return x_43; +x_41 = lean_box(0); +x_42 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_3); +return x_42; } else { -lean_object* x_44; uint8_t x_45; -x_44 = l_Lean_Syntax_getArg(x_34, x_8); -lean_dec(x_34); -lean_inc(x_44); -x_45 = l_Lean_Syntax_isNodeOf(x_44, x_10, x_8); -if (x_45 == 0) +lean_object* x_43; uint8_t x_44; +x_43 = l_Lean_Syntax_getArg(x_33, x_8); +lean_dec(x_33); +lean_inc(x_43); +x_44 = l_Lean_Syntax_matchesNull(x_43, x_8); +if (x_44 == 0) { -lean_object* x_46; lean_object* x_47; -lean_dec(x_44); -lean_dec(x_39); -lean_dec(x_20); +lean_object* x_45; lean_object* x_46; +lean_dec(x_43); +lean_dec(x_38); +lean_dec(x_19); lean_dec(x_2); -x_46 = lean_box(0); -x_47 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_3); -return x_47; +x_45 = lean_box(0); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_3); +return x_46; } else { -lean_object* x_48; lean_object* x_49; uint8_t x_50; -x_48 = l_Lean_Syntax_getArg(x_44, x_14); -lean_dec(x_44); -x_49 = l_unexpandExists___closed__2; -lean_inc(x_48); -x_50 = l_Lean_Syntax_isOfKind(x_48, x_49); -if (x_50 == 0) +lean_object* x_47; lean_object* x_48; uint8_t x_49; +x_47 = l_Lean_Syntax_getArg(x_43, x_13); +lean_dec(x_43); +x_48 = l_unexpandExists___closed__2; +lean_inc(x_47); +x_49 = l_Lean_Syntax_isOfKind(x_47, x_48); +if (x_49 == 0) { -lean_object* x_51; lean_object* x_52; -lean_dec(x_48); -lean_dec(x_39); -lean_dec(x_20); +lean_object* x_50; lean_object* x_51; +lean_dec(x_47); +lean_dec(x_38); +lean_dec(x_19); lean_dec(x_2); -x_51 = lean_box(0); -x_52 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_52, 0, x_51); -lean_ctor_set(x_52, 1, x_3); -return x_52; +x_50 = lean_box(0); +x_51 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_3); +return x_51; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; -x_53 = l_Lean_Syntax_getArg(x_48, x_8); -lean_dec(x_48); -x_54 = l_Lean_Syntax_getArg(x_20, x_35); -lean_dec(x_20); -x_55 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); -x_56 = !lean_is_exclusive(x_55); -if (x_56 == 0) +lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_52 = l_Lean_Syntax_getArg(x_47, x_8); +lean_dec(x_47); +x_53 = l_Lean_Syntax_getArg(x_19, x_34); +lean_dec(x_19); +x_54 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); +x_55 = !lean_is_exclusive(x_54); +if (x_55 == 0) { -lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; -x_57 = lean_ctor_get(x_55, 0); -x_58 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__22; +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_56 = lean_ctor_get(x_54, 0); +x_57 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__22; +lean_inc(x_56); +x_58 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_58, 0, x_56); +lean_ctor_set(x_58, 1, x_57); +x_59 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_60 = lean_array_push(x_59, x_38); +x_61 = lean_box(2); +x_62 = l_unexpandExists___closed__5; +x_63 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_63, 0, x_61); +lean_ctor_set(x_63, 1, x_62); +lean_ctor_set(x_63, 2, x_60); +x_64 = lean_array_push(x_59, x_63); +x_65 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_66 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_66, 0, x_61); +lean_ctor_set(x_66, 1, x_65); +lean_ctor_set(x_66, 2, x_64); +x_67 = l_Lean_expandExplicitBindersAux_loop___closed__13; +lean_inc(x_56); +x_68 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_68, 0, x_56); +lean_ctor_set(x_68, 1, x_67); +x_69 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__29; +lean_inc(x_56); +x_70 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_70, 0, x_56); +lean_ctor_set(x_70, 1, x_69); +x_71 = l_unexpandExists___closed__6; +x_72 = lean_array_push(x_71, x_58); +x_73 = lean_array_push(x_72, x_66); +x_74 = lean_array_push(x_73, x_68); +x_75 = lean_array_push(x_74, x_52); +x_76 = lean_array_push(x_75, x_70); +x_77 = l_Lean_bracketedExplicitBinders___closed__2; +x_78 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_78, 0, x_61); +lean_ctor_set(x_78, 1, x_77); +lean_ctor_set(x_78, 2, x_76); +x_79 = l_unexpandPSigma___closed__1; +x_80 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_80, 0, x_56); +lean_ctor_set(x_80, 1, x_79); +x_81 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; +x_82 = lean_array_push(x_81, x_78); +x_83 = lean_array_push(x_82, x_80); +x_84 = lean_array_push(x_83, x_53); +x_85 = l_term___xd7_x27_____closed__2; +x_86 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_86, 0, x_61); +lean_ctor_set(x_86, 1, x_85); +lean_ctor_set(x_86, 2, x_84); +lean_ctor_set(x_54, 0, x_86); +return x_54; +} +else +{ +lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; +x_87 = lean_ctor_get(x_54, 0); +x_88 = lean_ctor_get(x_54, 1); +lean_inc(x_88); +lean_inc(x_87); +lean_dec(x_54); +x_89 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__22; +lean_inc(x_87); +x_90 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_90, 0, x_87); +lean_ctor_set(x_90, 1, x_89); +x_91 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_92 = lean_array_push(x_91, x_38); +x_93 = lean_box(2); +x_94 = l_unexpandExists___closed__5; +x_95 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_95, 0, x_93); +lean_ctor_set(x_95, 1, x_94); +lean_ctor_set(x_95, 2, x_92); +x_96 = lean_array_push(x_91, x_95); +x_97 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_98 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_98, 0, x_93); +lean_ctor_set(x_98, 1, x_97); +lean_ctor_set(x_98, 2, x_96); +x_99 = l_Lean_expandExplicitBindersAux_loop___closed__13; +lean_inc(x_87); +x_100 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_100, 0, x_87); +lean_ctor_set(x_100, 1, x_99); +x_101 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__29; +lean_inc(x_87); +x_102 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_102, 0, x_87); +lean_ctor_set(x_102, 1, x_101); +x_103 = l_unexpandExists___closed__6; +x_104 = lean_array_push(x_103, x_90); +x_105 = lean_array_push(x_104, x_98); +x_106 = lean_array_push(x_105, x_100); +x_107 = lean_array_push(x_106, x_52); +x_108 = lean_array_push(x_107, x_102); +x_109 = l_Lean_bracketedExplicitBinders___closed__2; +x_110 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_110, 0, x_93); +lean_ctor_set(x_110, 1, x_109); +lean_ctor_set(x_110, 2, x_108); +x_111 = l_unexpandPSigma___closed__1; +x_112 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_112, 0, x_87); +lean_ctor_set(x_112, 1, x_111); +x_113 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; +x_114 = lean_array_push(x_113, x_110); +x_115 = lean_array_push(x_114, x_112); +x_116 = lean_array_push(x_115, x_53); +x_117 = l_term___xd7_x27_____closed__2; +x_118 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_118, 0, x_93); +lean_ctor_set(x_118, 1, x_117); +lean_ctor_set(x_118, 2, x_116); +x_119 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_119, 0, x_118); +lean_ctor_set(x_119, 1, x_88); +return x_119; +} +} +} +} +} +} +} +} +} +} +} +} +} +static lean_object* _init_l_unexpandSubtype___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("term{_:_//_}", 12); +return x_1; +} +} +static lean_object* _init_l_unexpandSubtype___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_unexpandSubtype___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_unexpandSubtype___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("{", 1); +return x_1; +} +} +static lean_object* _init_l_unexpandSubtype___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("//", 2); +return x_1; +} +} +static lean_object* _init_l_unexpandSubtype___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("}", 1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_unexpandSubtype(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__6; +lean_inc(x_1); +x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; +lean_dec(x_2); +lean_dec(x_1); +x_6 = lean_box(0); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_3); +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_8 = lean_unsigned_to_nat(1u); +x_9 = l_Lean_Syntax_getArg(x_1, x_8); +lean_dec(x_1); +lean_inc(x_9); +x_10 = l_Lean_Syntax_matchesNull(x_9, x_8); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +lean_dec(x_9); +lean_dec(x_2); +x_11 = lean_box(0); +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_3); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_13 = lean_unsigned_to_nat(0u); +x_14 = l_Lean_Syntax_getArg(x_9, x_13); +lean_dec(x_9); +x_15 = l_Lean_expandExplicitBindersAux_loop___closed__2; +lean_inc(x_14); +x_16 = l_Lean_Syntax_isOfKind(x_14, x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +lean_dec(x_14); +lean_dec(x_2); +x_17 = lean_box(0); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_3); +return x_18; +} +else +{ +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = l_Lean_Syntax_getArg(x_14, x_8); +lean_dec(x_14); +x_20 = l_Lean_expandExplicitBindersAux_loop___closed__4; +lean_inc(x_19); +x_21 = l_Lean_Syntax_isOfKind(x_19, x_20); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; +lean_dec(x_19); +lean_dec(x_2); +x_22 = lean_box(0); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_3); +return x_23; +} +else +{ +lean_object* x_24; uint8_t x_25; +x_24 = l_Lean_Syntax_getArg(x_19, x_13); +lean_inc(x_24); +x_25 = l_Lean_Syntax_matchesNull(x_24, x_8); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; +lean_dec(x_24); +lean_dec(x_19); +lean_dec(x_2); +x_26 = lean_box(0); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_3); +return x_27; +} +else +{ +lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_28 = l_Lean_Syntax_getArg(x_24, x_13); +lean_dec(x_24); +x_29 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__21; +lean_inc(x_28); +x_30 = l_Lean_Syntax_isOfKind(x_28, x_29); +if (x_30 == 0) +{ +lean_object* x_31; uint8_t x_32; +x_31 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__10; +lean_inc(x_28); +x_32 = l_Lean_Syntax_isOfKind(x_28, x_31); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; +lean_dec(x_28); +lean_dec(x_19); +lean_dec(x_2); +x_33 = lean_box(0); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_3); +return x_34; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; +x_35 = lean_unsigned_to_nat(2u); +x_36 = l_Lean_Syntax_getArg(x_19, x_35); +lean_dec(x_19); +x_37 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); +x_38 = !lean_is_exclusive(x_37); +if (x_38 == 0) +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_39 = lean_ctor_get(x_37, 0); +x_40 = l_unexpandSubtype___closed__3; +lean_inc(x_39); +x_41 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +x_42 = l_unexpandSubtype___closed__4; +lean_inc(x_39); +x_43 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_43, 0, x_39); +lean_ctor_set(x_43, 1, x_42); +x_44 = l_unexpandSubtype___closed__5; +x_45 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_45, 0, x_39); +lean_ctor_set(x_45, 1, x_44); +x_46 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__28; +x_47 = lean_array_push(x_46, x_41); +x_48 = lean_array_push(x_47, x_28); +x_49 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; +x_50 = lean_array_push(x_48, x_49); +x_51 = lean_array_push(x_50, x_43); +x_52 = lean_array_push(x_51, x_36); +x_53 = lean_array_push(x_52, x_45); +x_54 = lean_box(2); +x_55 = l_unexpandSubtype___closed__2; +x_56 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_56, 0, x_54); +lean_ctor_set(x_56, 1, x_55); +lean_ctor_set(x_56, 2, x_53); +lean_ctor_set(x_37, 0, x_56); +return x_37; +} +else +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; +x_57 = lean_ctor_get(x_37, 0); +x_58 = lean_ctor_get(x_37, 1); +lean_inc(x_58); lean_inc(x_57); -x_59 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_59, 0, x_57); -lean_ctor_set(x_59, 1, x_58); -x_60 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_61 = lean_array_push(x_60, x_39); -x_62 = lean_box(2); -x_63 = l_Lean_binderIdent___closed__2; -x_64 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_64, 0, x_62); -lean_ctor_set(x_64, 1, x_63); -lean_ctor_set(x_64, 2, x_61); -x_65 = lean_array_push(x_60, x_64); -x_66 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_67 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_67, 0, x_62); -lean_ctor_set(x_67, 1, x_66); -lean_ctor_set(x_67, 2, x_65); -x_68 = l_Lean_expandExplicitBindersAux_loop___closed__12; +lean_dec(x_37); +x_59 = l_unexpandSubtype___closed__3; lean_inc(x_57); -x_69 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_69, 0, x_57); -lean_ctor_set(x_69, 1, x_68); -x_70 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__29; +x_60 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_60, 0, x_57); +lean_ctor_set(x_60, 1, x_59); +x_61 = l_unexpandSubtype___closed__4; lean_inc(x_57); -x_71 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_71, 0, x_57); -lean_ctor_set(x_71, 1, x_70); -x_72 = l_unexpandExists___closed__4; -x_73 = lean_array_push(x_72, x_59); -x_74 = lean_array_push(x_73, x_67); -x_75 = lean_array_push(x_74, x_69); -x_76 = lean_array_push(x_75, x_53); -x_77 = lean_array_push(x_76, x_71); -x_78 = l_Lean_bracketedExplicitBinders___closed__2; -x_79 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_79, 0, x_62); -lean_ctor_set(x_79, 1, x_78); -lean_ctor_set(x_79, 2, x_77); -x_80 = l_unexpandPSigma___closed__1; -x_81 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_81, 0, x_57); -lean_ctor_set(x_81, 1, x_80); -x_82 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; -x_83 = lean_array_push(x_82, x_79); -x_84 = lean_array_push(x_83, x_81); -x_85 = lean_array_push(x_84, x_54); -x_86 = l_term___xd7_x27_____closed__2; -x_87 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_87, 0, x_62); -lean_ctor_set(x_87, 1, x_86); -lean_ctor_set(x_87, 2, x_85); -lean_ctor_set(x_55, 0, x_87); -return x_55; +x_62 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_62, 0, x_57); +lean_ctor_set(x_62, 1, x_61); +x_63 = l_unexpandSubtype___closed__5; +x_64 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_64, 0, x_57); +lean_ctor_set(x_64, 1, x_63); +x_65 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__28; +x_66 = lean_array_push(x_65, x_60); +x_67 = lean_array_push(x_66, x_28); +x_68 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; +x_69 = lean_array_push(x_67, x_68); +x_70 = lean_array_push(x_69, x_62); +x_71 = lean_array_push(x_70, x_36); +x_72 = lean_array_push(x_71, x_64); +x_73 = lean_box(2); +x_74 = l_unexpandSubtype___closed__2; +x_75 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set(x_75, 1, x_74); +lean_ctor_set(x_75, 2, x_72); +x_76 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_76, 0, x_75); +lean_ctor_set(x_76, 1, x_58); +return x_76; +} +} } else { -lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; -x_88 = lean_ctor_get(x_55, 0); -x_89 = lean_ctor_get(x_55, 1); -lean_inc(x_89); -lean_inc(x_88); -lean_dec(x_55); -x_90 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__22; -lean_inc(x_88); -x_91 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_91, 0, x_88); -lean_ctor_set(x_91, 1, x_90); -x_92 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_93 = lean_array_push(x_92, x_39); -x_94 = lean_box(2); -x_95 = l_Lean_binderIdent___closed__2; -x_96 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_96, 0, x_94); -lean_ctor_set(x_96, 1, x_95); -lean_ctor_set(x_96, 2, x_93); -x_97 = lean_array_push(x_92, x_96); -x_98 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_99 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_99, 0, x_94); -lean_ctor_set(x_99, 1, x_98); -lean_ctor_set(x_99, 2, x_97); -x_100 = l_Lean_expandExplicitBindersAux_loop___closed__12; -lean_inc(x_88); -x_101 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_101, 0, x_88); -lean_ctor_set(x_101, 1, x_100); -x_102 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__29; -lean_inc(x_88); -x_103 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_103, 0, x_88); -lean_ctor_set(x_103, 1, x_102); -x_104 = l_unexpandExists___closed__4; -x_105 = lean_array_push(x_104, x_91); -x_106 = lean_array_push(x_105, x_99); -x_107 = lean_array_push(x_106, x_101); -x_108 = lean_array_push(x_107, x_53); -x_109 = lean_array_push(x_108, x_103); -x_110 = l_Lean_bracketedExplicitBinders___closed__2; -x_111 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_111, 0, x_94); -lean_ctor_set(x_111, 1, x_110); -lean_ctor_set(x_111, 2, x_109); -x_112 = l_unexpandPSigma___closed__1; -x_113 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_113, 0, x_88); -lean_ctor_set(x_113, 1, x_112); -x_114 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; -x_115 = lean_array_push(x_114, x_111); -x_116 = lean_array_push(x_115, x_113); -x_117 = lean_array_push(x_116, x_54); -x_118 = l_term___xd7_x27_____closed__2; -x_119 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_119, 0, x_94); -lean_ctor_set(x_119, 1, x_118); -lean_ctor_set(x_119, 2, x_117); -x_120 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_120, 0, x_119); -lean_ctor_set(x_120, 1, x_89); -return x_120; +lean_object* x_77; lean_object* x_78; uint8_t x_79; +x_77 = l_Lean_Syntax_getArg(x_28, x_8); +lean_dec(x_28); +x_78 = lean_unsigned_to_nat(2u); +lean_inc(x_77); +x_79 = l_Lean_Syntax_matchesNull(x_77, x_78); +if (x_79 == 0) +{ +lean_object* x_80; lean_object* x_81; +lean_dec(x_77); +lean_dec(x_19); +lean_dec(x_2); +x_80 = lean_box(0); +x_81 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_81, 0, x_80); +lean_ctor_set(x_81, 1, x_3); +return x_81; +} +else +{ +lean_object* x_82; lean_object* x_83; uint8_t x_84; +x_82 = l_Lean_Syntax_getArg(x_77, x_13); +x_83 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__10; +lean_inc(x_82); +x_84 = l_Lean_Syntax_isOfKind(x_82, x_83); +if (x_84 == 0) +{ +lean_object* x_85; lean_object* x_86; +lean_dec(x_82); +lean_dec(x_77); +lean_dec(x_19); +lean_dec(x_2); +x_85 = lean_box(0); +x_86 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_86, 0, x_85); +lean_ctor_set(x_86, 1, x_3); +return x_86; +} +else +{ +lean_object* x_87; uint8_t x_88; +x_87 = l_Lean_Syntax_getArg(x_77, x_8); +lean_dec(x_77); +lean_inc(x_87); +x_88 = l_Lean_Syntax_matchesNull(x_87, x_8); +if (x_88 == 0) +{ +lean_object* x_89; lean_object* x_90; +lean_dec(x_87); +lean_dec(x_82); +lean_dec(x_19); +lean_dec(x_2); +x_89 = lean_box(0); +x_90 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_90, 0, x_89); +lean_ctor_set(x_90, 1, x_3); +return x_90; +} +else +{ +lean_object* x_91; lean_object* x_92; uint8_t x_93; +x_91 = l_Lean_Syntax_getArg(x_87, x_13); +lean_dec(x_87); +x_92 = l_unexpandExists___closed__2; +lean_inc(x_91); +x_93 = l_Lean_Syntax_isOfKind(x_91, x_92); +if (x_93 == 0) +{ +lean_object* x_94; lean_object* x_95; +lean_dec(x_91); +lean_dec(x_82); +lean_dec(x_19); +lean_dec(x_2); +x_94 = lean_box(0); +x_95 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_95, 0, x_94); +lean_ctor_set(x_95, 1, x_3); +return x_95; +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; uint8_t x_99; +x_96 = l_Lean_Syntax_getArg(x_91, x_8); +lean_dec(x_91); +x_97 = l_Lean_Syntax_getArg(x_19, x_78); +lean_dec(x_19); +x_98 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); +x_99 = !lean_is_exclusive(x_98); +if (x_99 == 0) +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_100 = lean_ctor_get(x_98, 0); +x_101 = l_unexpandSubtype___closed__3; +lean_inc(x_100); +x_102 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_102, 0, x_100); +lean_ctor_set(x_102, 1, x_101); +x_103 = l_Lean_expandExplicitBindersAux_loop___closed__13; +lean_inc(x_100); +x_104 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_104, 0, x_100); +lean_ctor_set(x_104, 1, x_103); +x_105 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_106 = lean_array_push(x_105, x_104); +x_107 = lean_array_push(x_106, x_96); +x_108 = lean_box(2); +x_109 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_110 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_110, 0, x_108); +lean_ctor_set(x_110, 1, x_109); +lean_ctor_set(x_110, 2, x_107); +x_111 = l_unexpandSubtype___closed__4; +lean_inc(x_100); +x_112 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_112, 0, x_100); +lean_ctor_set(x_112, 1, x_111); +x_113 = l_unexpandSubtype___closed__5; +x_114 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_114, 0, x_100); +lean_ctor_set(x_114, 1, x_113); +x_115 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__28; +x_116 = lean_array_push(x_115, x_102); +x_117 = lean_array_push(x_116, x_82); +x_118 = lean_array_push(x_117, x_110); +x_119 = lean_array_push(x_118, x_112); +x_120 = lean_array_push(x_119, x_97); +x_121 = lean_array_push(x_120, x_114); +x_122 = l_unexpandSubtype___closed__2; +x_123 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_123, 0, x_108); +lean_ctor_set(x_123, 1, x_122); +lean_ctor_set(x_123, 2, x_121); +lean_ctor_set(x_98, 0, x_123); +return x_98; +} +else +{ +lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; +x_124 = lean_ctor_get(x_98, 0); +x_125 = lean_ctor_get(x_98, 1); +lean_inc(x_125); +lean_inc(x_124); +lean_dec(x_98); +x_126 = l_unexpandSubtype___closed__3; +lean_inc(x_124); +x_127 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_127, 0, x_124); +lean_ctor_set(x_127, 1, x_126); +x_128 = l_Lean_expandExplicitBindersAux_loop___closed__13; +lean_inc(x_124); +x_129 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_129, 0, x_124); +lean_ctor_set(x_129, 1, x_128); +x_130 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_131 = lean_array_push(x_130, x_129); +x_132 = lean_array_push(x_131, x_96); +x_133 = lean_box(2); +x_134 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_135 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_135, 0, x_133); +lean_ctor_set(x_135, 1, x_134); +lean_ctor_set(x_135, 2, x_132); +x_136 = l_unexpandSubtype___closed__4; +lean_inc(x_124); +x_137 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_137, 0, x_124); +lean_ctor_set(x_137, 1, x_136); +x_138 = l_unexpandSubtype___closed__5; +x_139 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_139, 0, x_124); +lean_ctor_set(x_139, 1, x_138); +x_140 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__28; +x_141 = lean_array_push(x_140, x_127); +x_142 = lean_array_push(x_141, x_82); +x_143 = lean_array_push(x_142, x_135); +x_144 = lean_array_push(x_143, x_137); +x_145 = lean_array_push(x_144, x_97); +x_146 = lean_array_push(x_145, x_139); +x_147 = l_unexpandSubtype___closed__2; +x_148 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_148, 0, x_133); +lean_ctor_set(x_148, 1, x_147); +lean_ctor_set(x_148, 2, x_146); +x_149 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_149, 0, x_148); +lean_ctor_set(x_149, 1, x_125); +return x_149; +} } } } @@ -10078,54 +10389,152 @@ return x_120; } } } +LEAN_EXPORT lean_object* l_unexpandTSyntax(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__6; +lean_inc(x_1); +x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; +lean_dec(x_2); +lean_dec(x_1); +x_6 = lean_box(0); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_3); +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_8 = lean_unsigned_to_nat(0u); +x_9 = l_Lean_Syntax_getArg(x_1, x_8); +x_10 = lean_unsigned_to_nat(1u); +x_11 = l_Lean_Syntax_getArg(x_1, x_10); +lean_dec(x_1); +lean_inc(x_11); +x_12 = l_Lean_Syntax_matchesNull(x_11, x_10); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_2); +x_13 = lean_box(0); +x_14 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_3); +return x_14; } -static lean_object* _init_l_unexpandSubtype___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("term{__:_//_}", 13); -return x_1; -} +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = l_Lean_Syntax_getArg(x_11, x_8); +lean_dec(x_11); +x_16 = l_unexpandListNil___rarg___closed__2; +lean_inc(x_15); +x_17 = l_Lean_Syntax_isOfKind(x_15, x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +lean_dec(x_15); +lean_dec(x_9); +lean_dec(x_2); +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } -static lean_object* _init_l_unexpandSubtype___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_unexpandSubtype___closed__1; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; +lean_object* x_20; uint8_t x_21; +x_20 = l_Lean_Syntax_getArg(x_15, x_10); +lean_dec(x_15); +lean_inc(x_20); +x_21 = l_Lean_Syntax_matchesNull(x_20, x_10); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; +lean_dec(x_20); +lean_dec(x_9); +lean_dec(x_2); +x_22 = lean_box(0); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_3); +return x_23; } +else +{ +lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_24 = l_Lean_Syntax_getArg(x_20, x_8); +lean_dec(x_20); +x_25 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); +x_26 = !lean_is_exclusive(x_25); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; 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; +x_27 = lean_ctor_get(x_25, 0); +lean_dec(x_27); +x_28 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_29 = lean_array_push(x_28, x_24); +x_30 = lean_box(2); +x_31 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_32 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set(x_32, 2, x_29); +x_33 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_34 = lean_array_push(x_33, x_9); +x_35 = lean_array_push(x_34, x_32); +x_36 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_36, 0, x_30); +lean_ctor_set(x_36, 1, x_4); +lean_ctor_set(x_36, 2, x_35); +lean_ctor_set(x_25, 0, x_36); +return x_25; } -static lean_object* _init_l_unexpandSubtype___closed__3() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("{", 1); -return x_1; +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_37 = lean_ctor_get(x_25, 1); +lean_inc(x_37); +lean_dec(x_25); +x_38 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_39 = lean_array_push(x_38, x_24); +x_40 = lean_box(2); +x_41 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_42 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_41); +lean_ctor_set(x_42, 2, x_39); +x_43 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_44 = lean_array_push(x_43, x_9); +x_45 = lean_array_push(x_44, x_42); +x_46 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_46, 0, x_40); +lean_ctor_set(x_46, 1, x_4); +lean_ctor_set(x_46, 2, x_45); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_37); +return x_47; } } -static lean_object* _init_l_unexpandSubtype___closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("//", 2); -return x_1; } } -static lean_object* _init_l_unexpandSubtype___closed__5() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("}", 1); -return x_1; } } -LEAN_EXPORT lean_object* l_unexpandSubtype(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +} +LEAN_EXPORT lean_object* l_unexpandTSyntaxArray(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; uint8_t x_5; -x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__6; +x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__6; lean_inc(x_1); x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); if (x_5 == 0) @@ -10141,37 +10550,39 @@ return x_7; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_unsigned_to_nat(1u); +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_8 = lean_unsigned_to_nat(0u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); +x_10 = lean_unsigned_to_nat(1u); +x_11 = l_Lean_Syntax_getArg(x_1, x_10); lean_dec(x_1); -x_10 = l_Lean_nullKind; -lean_inc(x_9); -x_11 = l_Lean_Syntax_isNodeOf(x_9, x_10, x_8); -if (x_11 == 0) +lean_inc(x_11); +x_12 = l_Lean_Syntax_matchesNull(x_11, x_10); +if (x_12 == 0) { -lean_object* x_12; lean_object* x_13; +lean_object* x_13; lean_object* x_14; +lean_dec(x_11); lean_dec(x_9); lean_dec(x_2); -x_12 = lean_box(0); -x_13 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_3); -return x_13; +x_13 = lean_box(0); +x_14 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_3); +return x_14; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_14 = lean_unsigned_to_nat(0u); -x_15 = l_Lean_Syntax_getArg(x_9, x_14); -lean_dec(x_9); -x_16 = l_Lean_expandExplicitBindersAux_loop___closed__2; +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = l_Lean_Syntax_getArg(x_11, x_8); +lean_dec(x_11); +x_16 = l_unexpandListNil___rarg___closed__2; lean_inc(x_15); x_17 = l_Lean_Syntax_isOfKind(x_15, x_16); if (x_17 == 0) { lean_object* x_18; lean_object* x_19; lean_dec(x_15); +lean_dec(x_9); lean_dec(x_2); x_18 = lean_box(0); x_19 = lean_alloc_ctor(1, 2, 0); @@ -10181,348 +10592,223 @@ return x_19; } else { -lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_20 = l_Lean_Syntax_getArg(x_15, x_8); +lean_object* x_20; uint8_t x_21; +x_20 = l_Lean_Syntax_getArg(x_15, x_10); lean_dec(x_15); -x_21 = l_Lean_expandExplicitBindersAux_loop___closed__4; lean_inc(x_20); -x_22 = l_Lean_Syntax_isOfKind(x_20, x_21); -if (x_22 == 0) +x_21 = l_Lean_Syntax_matchesNull(x_20, x_10); +if (x_21 == 0) { -lean_object* x_23; lean_object* x_24; +lean_object* x_22; lean_object* x_23; lean_dec(x_20); +lean_dec(x_9); lean_dec(x_2); -x_23 = lean_box(0); -x_24 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_3); -return x_24; +x_22 = lean_box(0); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_3); +return x_23; } else { -lean_object* x_25; uint8_t x_26; -x_25 = l_Lean_Syntax_getArg(x_20, x_14); -lean_inc(x_25); -x_26 = l_Lean_Syntax_isNodeOf(x_25, x_10, x_8); +lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_24 = l_Lean_Syntax_getArg(x_20, x_8); +lean_dec(x_20); +x_25 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); +x_26 = !lean_is_exclusive(x_25); if (x_26 == 0) { -lean_object* x_27; lean_object* x_28; -lean_dec(x_25); -lean_dec(x_20); -lean_dec(x_2); -x_27 = lean_box(0); -x_28 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_28, 0, x_27); -lean_ctor_set(x_28, 1, x_3); -return x_28; +lean_object* x_27; lean_object* x_28; lean_object* x_29; 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; +x_27 = lean_ctor_get(x_25, 0); +lean_dec(x_27); +x_28 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_29 = lean_array_push(x_28, x_24); +x_30 = lean_box(2); +x_31 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_32 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set(x_32, 2, x_29); +x_33 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_34 = lean_array_push(x_33, x_9); +x_35 = lean_array_push(x_34, x_32); +x_36 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_36, 0, x_30); +lean_ctor_set(x_36, 1, x_4); +lean_ctor_set(x_36, 2, x_35); +lean_ctor_set(x_25, 0, x_36); +return x_25; } else { -lean_object* x_29; lean_object* x_30; uint8_t x_31; -x_29 = l_Lean_Syntax_getArg(x_25, x_14); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_37 = lean_ctor_get(x_25, 1); +lean_inc(x_37); lean_dec(x_25); -x_30 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__21; -lean_inc(x_29); -x_31 = l_Lean_Syntax_isOfKind(x_29, x_30); -if (x_31 == 0) -{ -lean_object* x_32; uint8_t x_33; -x_32 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__10; -lean_inc(x_29); -x_33 = l_Lean_Syntax_isOfKind(x_29, x_32); -if (x_33 == 0) -{ -lean_object* x_34; lean_object* x_35; -lean_dec(x_29); -lean_dec(x_20); -lean_dec(x_2); -x_34 = lean_box(0); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_3); -return x_35; -} -else -{ -lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; -x_36 = lean_unsigned_to_nat(2u); -x_37 = l_Lean_Syntax_getArg(x_20, x_36); -lean_dec(x_20); -x_38 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); -x_39 = !lean_is_exclusive(x_38); -if (x_39 == 0) -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_40 = lean_ctor_get(x_38, 0); -x_41 = l_unexpandSubtype___closed__3; -lean_inc(x_40); -x_42 = lean_alloc_ctor(2, 2, 0); +x_38 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_39 = lean_array_push(x_38, x_24); +x_40 = lean_box(2); +x_41 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_42 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_42, 0, x_40); lean_ctor_set(x_42, 1, x_41); -x_43 = l_unexpandSubtype___closed__4; -lean_inc(x_40); -x_44 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_44, 0, x_40); -lean_ctor_set(x_44, 1, x_43); -x_45 = l_unexpandSubtype___closed__5; -x_46 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_42, 2, x_39); +x_43 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_44 = lean_array_push(x_43, x_9); +x_45 = lean_array_push(x_44, x_42); +x_46 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_46, 0, x_40); -lean_ctor_set(x_46, 1, x_45); -x_47 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__21; -x_48 = lean_array_push(x_47, x_42); -x_49 = lean_array_push(x_48, x_29); -x_50 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; -x_51 = lean_array_push(x_49, x_50); -x_52 = lean_array_push(x_51, x_44); -x_53 = lean_array_push(x_52, x_37); -x_54 = lean_array_push(x_53, x_46); -x_55 = lean_box(2); -x_56 = l_unexpandSubtype___closed__2; -x_57 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_57, 0, x_55); -lean_ctor_set(x_57, 1, x_56); -lean_ctor_set(x_57, 2, x_54); -lean_ctor_set(x_38, 0, x_57); -return x_38; +lean_ctor_set(x_46, 1, x_4); +lean_ctor_set(x_46, 2, x_45); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_37); +return x_47; } -else -{ -lean_object* x_58; lean_object* x_59; 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; 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; -x_58 = lean_ctor_get(x_38, 0); -x_59 = lean_ctor_get(x_38, 1); -lean_inc(x_59); -lean_inc(x_58); -lean_dec(x_38); -x_60 = l_unexpandSubtype___closed__3; -lean_inc(x_58); -x_61 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_61, 0, x_58); -lean_ctor_set(x_61, 1, x_60); -x_62 = l_unexpandSubtype___closed__4; -lean_inc(x_58); -x_63 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_63, 0, x_58); -lean_ctor_set(x_63, 1, x_62); -x_64 = l_unexpandSubtype___closed__5; -x_65 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_65, 0, x_58); -lean_ctor_set(x_65, 1, x_64); -x_66 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__21; -x_67 = lean_array_push(x_66, x_61); -x_68 = lean_array_push(x_67, x_29); -x_69 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; -x_70 = lean_array_push(x_68, x_69); -x_71 = lean_array_push(x_70, x_63); -x_72 = lean_array_push(x_71, x_37); -x_73 = lean_array_push(x_72, x_65); -x_74 = lean_box(2); -x_75 = l_unexpandSubtype___closed__2; -x_76 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_76, 0, x_74); -lean_ctor_set(x_76, 1, x_75); -lean_ctor_set(x_76, 2, x_73); -x_77 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_77, 0, x_76); -lean_ctor_set(x_77, 1, x_59); -return x_77; } } } -else +} +} +} +LEAN_EXPORT lean_object* l_unexpandTSepArray(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_78; lean_object* x_79; uint8_t x_80; -x_78 = l_Lean_Syntax_getArg(x_29, x_8); -lean_dec(x_29); -x_79 = lean_unsigned_to_nat(2u); -lean_inc(x_78); -x_80 = l_Lean_Syntax_isNodeOf(x_78, x_10, x_79); -if (x_80 == 0) +lean_object* x_4; uint8_t x_5; +x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__6; +lean_inc(x_1); +x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); +if (x_5 == 0) { -lean_object* x_81; lean_object* x_82; -lean_dec(x_78); -lean_dec(x_20); +lean_object* x_6; lean_object* x_7; lean_dec(x_2); -x_81 = lean_box(0); -x_82 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_82, 0, x_81); -lean_ctor_set(x_82, 1, x_3); -return x_82; +lean_dec(x_1); +x_6 = lean_box(0); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_3); +return x_7; } else { -lean_object* x_83; lean_object* x_84; uint8_t x_85; -x_83 = l_Lean_Syntax_getArg(x_78, x_14); -x_84 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__10; -lean_inc(x_83); -x_85 = l_Lean_Syntax_isOfKind(x_83, x_84); -if (x_85 == 0) +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_8 = lean_unsigned_to_nat(0u); +x_9 = l_Lean_Syntax_getArg(x_1, x_8); +x_10 = lean_unsigned_to_nat(1u); +x_11 = l_Lean_Syntax_getArg(x_1, x_10); +lean_dec(x_1); +x_12 = lean_unsigned_to_nat(2u); +lean_inc(x_11); +x_13 = l_Lean_Syntax_matchesNull(x_11, x_12); +if (x_13 == 0) { -lean_object* x_86; lean_object* x_87; -lean_dec(x_83); -lean_dec(x_78); -lean_dec(x_20); +lean_object* x_14; lean_object* x_15; +lean_dec(x_11); +lean_dec(x_9); lean_dec(x_2); -x_86 = lean_box(0); -x_87 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_87, 0, x_86); -lean_ctor_set(x_87, 1, x_3); -return x_87; +x_14 = lean_box(0); +x_15 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_3); +return x_15; } else { -lean_object* x_88; uint8_t x_89; -x_88 = l_Lean_Syntax_getArg(x_78, x_8); -lean_dec(x_78); -lean_inc(x_88); -x_89 = l_Lean_Syntax_isNodeOf(x_88, x_10, x_8); -if (x_89 == 0) +lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_16 = l_Lean_Syntax_getArg(x_11, x_8); +x_17 = l_unexpandListNil___rarg___closed__2; +lean_inc(x_16); +x_18 = l_Lean_Syntax_isOfKind(x_16, x_17); +if (x_18 == 0) { -lean_object* x_90; lean_object* x_91; -lean_dec(x_88); -lean_dec(x_83); -lean_dec(x_20); +lean_object* x_19; lean_object* x_20; +lean_dec(x_16); +lean_dec(x_11); +lean_dec(x_9); lean_dec(x_2); -x_90 = lean_box(0); -x_91 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_91, 0, x_90); -lean_ctor_set(x_91, 1, x_3); -return x_91; +x_19 = lean_box(0); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_3); +return x_20; } else { -lean_object* x_92; lean_object* x_93; uint8_t x_94; -x_92 = l_Lean_Syntax_getArg(x_88, x_14); -lean_dec(x_88); -x_93 = l_unexpandExists___closed__2; -lean_inc(x_92); -x_94 = l_Lean_Syntax_isOfKind(x_92, x_93); -if (x_94 == 0) -{ -lean_object* x_95; lean_object* x_96; -lean_dec(x_92); -lean_dec(x_83); -lean_dec(x_20); +lean_object* x_21; uint8_t x_22; +x_21 = l_Lean_Syntax_getArg(x_16, x_10); +lean_dec(x_16); +lean_inc(x_21); +x_22 = l_Lean_Syntax_matchesNull(x_21, x_10); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; +lean_dec(x_21); +lean_dec(x_11); +lean_dec(x_9); lean_dec(x_2); -x_95 = lean_box(0); -x_96 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_96, 0, x_95); -lean_ctor_set(x_96, 1, x_3); -return x_96; +x_23 = lean_box(0); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_3); +return x_24; } else { -lean_object* x_97; lean_object* x_98; lean_object* x_99; uint8_t x_100; -x_97 = l_Lean_Syntax_getArg(x_92, x_8); -lean_dec(x_92); -x_98 = l_Lean_Syntax_getArg(x_20, x_79); -lean_dec(x_20); -x_99 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); -x_100 = !lean_is_exclusive(x_99); -if (x_100 == 0) +lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_25 = l_Lean_Syntax_getArg(x_21, x_8); +lean_dec(x_21); +x_26 = l_Lean_Syntax_getArg(x_11, x_10); +lean_dec(x_11); +x_27 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______unexpand__Function__comp__1___spec__1(x_2, x_3); +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) { -lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; -x_101 = lean_ctor_get(x_99, 0); -x_102 = l_unexpandSubtype___closed__3; -lean_inc(x_101); -x_103 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_103, 0, x_101); -lean_ctor_set(x_103, 1, x_102); -x_104 = l_Lean_expandExplicitBindersAux_loop___closed__12; -lean_inc(x_101); -x_105 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_105, 0, x_101); -lean_ctor_set(x_105, 1, x_104); -x_106 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_107 = lean_array_push(x_106, x_105); -x_108 = lean_array_push(x_107, x_97); -x_109 = lean_box(2); -x_110 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_111 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_111, 0, x_109); -lean_ctor_set(x_111, 1, x_110); -lean_ctor_set(x_111, 2, x_108); -x_112 = l_unexpandSubtype___closed__4; -lean_inc(x_101); -x_113 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_113, 0, x_101); -lean_ctor_set(x_113, 1, x_112); -x_114 = l_unexpandSubtype___closed__5; -x_115 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_115, 0, x_101); -lean_ctor_set(x_115, 1, x_114); -x_116 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__21; -x_117 = lean_array_push(x_116, x_103); -x_118 = lean_array_push(x_117, x_83); -x_119 = lean_array_push(x_118, x_111); -x_120 = lean_array_push(x_119, x_113); -x_121 = lean_array_push(x_120, x_98); -x_122 = lean_array_push(x_121, x_115); -x_123 = l_unexpandSubtype___closed__2; -x_124 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_124, 0, x_109); -lean_ctor_set(x_124, 1, x_123); -lean_ctor_set(x_124, 2, x_122); -lean_ctor_set(x_99, 0, x_124); -return x_99; +lean_object* x_29; 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; +x_29 = lean_ctor_get(x_27, 0); +lean_dec(x_29); +x_30 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_31 = lean_array_push(x_30, x_25); +x_32 = lean_array_push(x_31, x_26); +x_33 = lean_box(2); +x_34 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_35 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +lean_ctor_set(x_35, 2, x_32); +x_36 = lean_array_push(x_30, x_9); +x_37 = lean_array_push(x_36, x_35); +x_38 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_38, 0, x_33); +lean_ctor_set(x_38, 1, x_4); +lean_ctor_set(x_38, 2, x_37); +lean_ctor_set(x_27, 0, x_38); +return x_27; } else { -lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; -x_125 = lean_ctor_get(x_99, 0); -x_126 = lean_ctor_get(x_99, 1); -lean_inc(x_126); -lean_inc(x_125); -lean_dec(x_99); -x_127 = l_unexpandSubtype___closed__3; -lean_inc(x_125); -x_128 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_128, 0, x_125); -lean_ctor_set(x_128, 1, x_127); -x_129 = l_Lean_expandExplicitBindersAux_loop___closed__12; -lean_inc(x_125); -x_130 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_130, 0, x_125); -lean_ctor_set(x_130, 1, x_129); -x_131 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_132 = lean_array_push(x_131, x_130); -x_133 = lean_array_push(x_132, x_97); -x_134 = lean_box(2); -x_135 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_136 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_136, 0, x_134); -lean_ctor_set(x_136, 1, x_135); -lean_ctor_set(x_136, 2, x_133); -x_137 = l_unexpandSubtype___closed__4; -lean_inc(x_125); -x_138 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_138, 0, x_125); -lean_ctor_set(x_138, 1, x_137); -x_139 = l_unexpandSubtype___closed__5; -x_140 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_140, 0, x_125); -lean_ctor_set(x_140, 1, x_139); -x_141 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__21; -x_142 = lean_array_push(x_141, x_128); -x_143 = lean_array_push(x_142, x_83); -x_144 = lean_array_push(x_143, x_136); -x_145 = lean_array_push(x_144, x_138); -x_146 = lean_array_push(x_145, x_98); -x_147 = lean_array_push(x_146, x_140); -x_148 = l_unexpandSubtype___closed__2; -x_149 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_149, 0, x_134); -lean_ctor_set(x_149, 1, x_148); -lean_ctor_set(x_149, 2, x_147); -x_150 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_150, 0, x_149); -lean_ctor_set(x_150, 1, x_126); -return x_150; -} -} -} -} -} -} +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_39 = lean_ctor_get(x_27, 1); +lean_inc(x_39); +lean_dec(x_27); +x_40 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_41 = lean_array_push(x_40, x_25); +x_42 = lean_array_push(x_41, x_26); +x_43 = lean_box(2); +x_44 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = lean_array_push(x_40, x_9); +x_47 = lean_array_push(x_46, x_45); +x_48 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_48, 0, x_43); +lean_ctor_set(x_48, 1, x_4); +lean_ctor_set(x_48, 2, x_47); +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_39); +return x_49; } } } @@ -10600,7 +10886,7 @@ static lean_object* _init_l_tacticFunext_______closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__19; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__19; x_2 = lean_unsigned_to_nat(1024u); x_3 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10612,7 +10898,7 @@ static lean_object* _init_l_tacticFunext_______closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_tacticFunext_______closed__7; x_3 = l_tacticFunext_______closed__8; x_4 = lean_alloc_ctor(2, 3, 0); @@ -10638,7 +10924,7 @@ static lean_object* _init_l_tacticFunext_______closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_tacticFunext_______closed__4; x_3 = l_tacticFunext_______closed__10; x_4 = lean_alloc_ctor(2, 3, 0); @@ -10799,391 +11085,400 @@ return x_7; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +lean_object* x_8; lean_object* x_9; uint8_t x_10; x_8 = lean_unsigned_to_nat(1u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); lean_dec(x_1); -x_10 = l_Lean_Syntax_getArgs(x_9); -lean_dec(x_9); -x_11 = lean_array_get_size(x_10); -x_12 = lean_nat_dec_eq(x_11, x_8); -lean_dec(x_11); +lean_inc(x_9); +x_10 = l_Lean_Syntax_matchesNull(x_9, x_8); +if (x_10 == 0) +{ +lean_object* x_11; uint8_t x_12; +x_11 = l_Lean_Syntax_getNumArgs(x_9); +x_12 = lean_nat_dec_le(x_8, x_11); if (x_12 == 0) { -lean_object* x_13; uint8_t x_14; +lean_object* x_13; lean_object* x_14; +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_2); +x_13 = lean_box(1); +x_14 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_3); +return x_14; +} +else +{ +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; uint8_t x_25; +x_15 = lean_unsigned_to_nat(0u); +x_16 = l_Lean_Syntax_getArg(x_9, x_15); +x_17 = l_Lean_Syntax_getArgs(x_9); +lean_dec(x_9); +x_18 = lean_nat_sub(x_11, x_15); +lean_dec(x_11); +x_19 = l_Array_extract___rarg(x_17, x_8, x_18); +x_20 = lean_box(2); +x_21 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_22 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +lean_ctor_set(x_22, 2, x_19); +x_23 = l_Lean_Syntax_getArgs(x_22); +lean_dec(x_22); lean_inc(x_2); -x_13 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) +x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_25 = !lean_is_exclusive(x_24); +if (x_25 == 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; lean_object* 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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_15 = lean_ctor_get(x_13, 0); -x_16 = lean_ctor_get(x_2, 2); -lean_inc(x_16); -x_17 = lean_ctor_get(x_2, 1); -lean_inc(x_17); +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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_26 = lean_ctor_get(x_24, 0); +x_27 = lean_ctor_get(x_2, 2); +lean_inc(x_27); +x_28 = lean_ctor_get(x_2, 1); +lean_inc(x_28); lean_dec(x_2); -x_18 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__1; -lean_inc(x_15); -x_19 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_19, 0, x_15); -lean_ctor_set(x_19, 1, x_18); -x_20 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__6; -x_21 = l_Lean_addMacroScope(x_17, x_20, x_16); -x_22 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__5; -x_23 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__8; -lean_inc(x_15); -x_24 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_24, 0, x_15); -lean_ctor_set(x_24, 1, x_22); -lean_ctor_set(x_24, 2, x_21); -lean_ctor_set(x_24, 3, x_23); -x_25 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_26 = lean_array_push(x_25, x_19); -x_27 = lean_array_push(x_26, x_24); -x_28 = lean_box(2); -x_29 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__2; -x_30 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_30, 0, x_28); +x_29 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__1; +lean_inc(x_26); +x_30 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_30, 0, x_26); lean_ctor_set(x_30, 1, x_29); -lean_ctor_set(x_30, 2, x_27); -x_31 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__9; -lean_inc(x_15); -x_32 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_32, 0, x_15); -lean_ctor_set(x_32, 1, x_31); -x_33 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__10; -lean_inc(x_15); -x_34 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_34, 0, x_15); -lean_ctor_set(x_34, 1, x_33); -x_35 = l_Lean_instInhabitedSyntax; -x_36 = lean_unsigned_to_nat(0u); -x_37 = lean_array_get(x_35, x_10, x_36); -x_38 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_39 = lean_array_push(x_38, x_37); -x_40 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_41 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_41, 0, x_28); -lean_ctor_set(x_41, 1, x_40); -lean_ctor_set(x_41, 2, x_39); -x_42 = lean_array_push(x_25, x_34); -x_43 = lean_array_push(x_42, x_41); -x_44 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__11; -x_45 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_45, 0, x_28); -lean_ctor_set(x_45, 1, x_44); -lean_ctor_set(x_45, 2, x_43); -x_46 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__3; -x_47 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_47, 0, x_15); -lean_ctor_set(x_47, 1, x_46); -x_48 = lean_array_get_size(x_10); -x_49 = l_Array_toSubarray___rarg(x_10, x_8, x_48); -x_50 = l_Array_ofSubarray___rarg(x_49); -x_51 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__27; -x_52 = l_Array_append___rarg(x_51, x_50); -x_53 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_53, 0, x_28); -lean_ctor_set(x_53, 1, x_40); -lean_ctor_set(x_53, 2, x_52); -x_54 = lean_array_push(x_25, x_47); -x_55 = lean_array_push(x_54, x_53); +x_31 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__6; +x_32 = l_Lean_addMacroScope(x_28, x_31, x_27); +x_33 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__5; +x_34 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__8; +lean_inc(x_26); +x_35 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_35, 0, x_26); +lean_ctor_set(x_35, 1, x_33); +lean_ctor_set(x_35, 2, x_32); +lean_ctor_set(x_35, 3, x_34); +x_36 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_37 = lean_array_push(x_36, x_30); +x_38 = lean_array_push(x_37, x_35); +x_39 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__2; +x_40 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_40, 0, x_20); +lean_ctor_set(x_40, 1, x_39); +lean_ctor_set(x_40, 2, x_38); +x_41 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__9; +lean_inc(x_26); +x_42 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_42, 0, x_26); +lean_ctor_set(x_42, 1, x_41); +x_43 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__10; +lean_inc(x_26); +x_44 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_44, 0, x_26); +lean_ctor_set(x_44, 1, x_43); +x_45 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_46 = lean_array_push(x_45, x_16); +x_47 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_47, 0, x_20); +lean_ctor_set(x_47, 1, x_21); +lean_ctor_set(x_47, 2, x_46); +x_48 = lean_array_push(x_36, x_44); +x_49 = lean_array_push(x_48, x_47); +x_50 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__11; +x_51 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_51, 0, x_20); +lean_ctor_set(x_51, 1, x_50); +lean_ctor_set(x_51, 2, x_49); +x_52 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__3; +x_53 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_53, 0, x_26); +lean_ctor_set(x_53, 1, x_52); +x_54 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__27; +x_55 = l_Array_append___rarg(x_54, x_23); x_56 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_56, 0, x_28); -lean_ctor_set(x_56, 1, x_4); +lean_ctor_set(x_56, 0, x_20); +lean_ctor_set(x_56, 1, x_21); lean_ctor_set(x_56, 2, x_55); -x_57 = l_unexpandExists___closed__4; -x_58 = lean_array_push(x_57, x_30); -lean_inc(x_32); -x_59 = lean_array_push(x_58, x_32); -x_60 = lean_array_push(x_59, x_45); -x_61 = lean_array_push(x_60, x_32); -x_62 = lean_array_push(x_61, x_56); -x_63 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_63, 0, x_28); -lean_ctor_set(x_63, 1, x_40); -lean_ctor_set(x_63, 2, x_62); -x_64 = lean_array_push(x_38, x_63); -x_65 = l___aux__Init__NotationExtra______macroRules__tacticCalc____1___closed__4; +x_57 = lean_array_push(x_36, x_53); +x_58 = lean_array_push(x_57, x_56); +x_59 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_59, 0, x_20); +lean_ctor_set(x_59, 1, x_4); +lean_ctor_set(x_59, 2, x_58); +x_60 = l_unexpandExists___closed__6; +x_61 = lean_array_push(x_60, x_40); +lean_inc(x_42); +x_62 = lean_array_push(x_61, x_42); +x_63 = lean_array_push(x_62, x_51); +x_64 = lean_array_push(x_63, x_42); +x_65 = lean_array_push(x_64, x_59); x_66 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_66, 0, x_28); -lean_ctor_set(x_66, 1, x_65); -lean_ctor_set(x_66, 2, x_64); -lean_ctor_set(x_13, 0, x_66); -return x_13; +lean_ctor_set(x_66, 0, x_20); +lean_ctor_set(x_66, 1, x_21); +lean_ctor_set(x_66, 2, x_65); +x_67 = lean_array_push(x_45, x_66); +x_68 = l___aux__Init__NotationExtra______macroRules__tacticCalc____1___closed__4; +x_69 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_69, 0, x_20); +lean_ctor_set(x_69, 1, x_68); +lean_ctor_set(x_69, 2, x_67); +lean_ctor_set(x_24, 0, x_69); +return x_24; } 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; -x_67 = lean_ctor_get(x_13, 0); -x_68 = lean_ctor_get(x_13, 1); -lean_inc(x_68); -lean_inc(x_67); -lean_dec(x_13); -x_69 = lean_ctor_get(x_2, 2); -lean_inc(x_69); -x_70 = lean_ctor_get(x_2, 1); +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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; +x_70 = lean_ctor_get(x_24, 0); +x_71 = lean_ctor_get(x_24, 1); +lean_inc(x_71); lean_inc(x_70); +lean_dec(x_24); +x_72 = lean_ctor_get(x_2, 2); +lean_inc(x_72); +x_73 = lean_ctor_get(x_2, 1); +lean_inc(x_73); lean_dec(x_2); -x_71 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__1; -lean_inc(x_67); -x_72 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_72, 0, x_67); -lean_ctor_set(x_72, 1, x_71); -x_73 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__6; -x_74 = l_Lean_addMacroScope(x_70, x_73, x_69); -x_75 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__5; -x_76 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__8; -lean_inc(x_67); -x_77 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_77, 0, x_67); -lean_ctor_set(x_77, 1, x_75); -lean_ctor_set(x_77, 2, x_74); -lean_ctor_set(x_77, 3, x_76); -x_78 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_79 = lean_array_push(x_78, x_72); -x_80 = lean_array_push(x_79, x_77); -x_81 = lean_box(2); -x_82 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__2; -x_83 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_83, 0, x_81); -lean_ctor_set(x_83, 1, x_82); -lean_ctor_set(x_83, 2, x_80); -x_84 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__9; -lean_inc(x_67); -x_85 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_85, 0, x_67); +x_74 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__1; +lean_inc(x_70); +x_75 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_75, 0, x_70); +lean_ctor_set(x_75, 1, x_74); +x_76 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__6; +x_77 = l_Lean_addMacroScope(x_73, x_76, x_72); +x_78 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__5; +x_79 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__8; +lean_inc(x_70); +x_80 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_80, 0, x_70); +lean_ctor_set(x_80, 1, x_78); +lean_ctor_set(x_80, 2, x_77); +lean_ctor_set(x_80, 3, x_79); +x_81 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_82 = lean_array_push(x_81, x_75); +x_83 = lean_array_push(x_82, x_80); +x_84 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__2; +x_85 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_85, 0, x_20); lean_ctor_set(x_85, 1, x_84); -x_86 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__10; -lean_inc(x_67); +lean_ctor_set(x_85, 2, x_83); +x_86 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__9; +lean_inc(x_70); x_87 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_87, 0, x_67); +lean_ctor_set(x_87, 0, x_70); lean_ctor_set(x_87, 1, x_86); -x_88 = l_Lean_instInhabitedSyntax; -x_89 = lean_unsigned_to_nat(0u); -x_90 = lean_array_get(x_88, x_10, x_89); -x_91 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_92 = lean_array_push(x_91, x_90); -x_93 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_94 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_94, 0, x_81); -lean_ctor_set(x_94, 1, x_93); -lean_ctor_set(x_94, 2, x_92); -x_95 = lean_array_push(x_78, x_87); -x_96 = lean_array_push(x_95, x_94); -x_97 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__11; -x_98 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_98, 0, x_81); +x_88 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__10; +lean_inc(x_70); +x_89 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_89, 0, x_70); +lean_ctor_set(x_89, 1, x_88); +x_90 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_91 = lean_array_push(x_90, x_16); +x_92 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_92, 0, x_20); +lean_ctor_set(x_92, 1, x_21); +lean_ctor_set(x_92, 2, x_91); +x_93 = lean_array_push(x_81, x_89); +x_94 = lean_array_push(x_93, x_92); +x_95 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__11; +x_96 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_96, 0, x_20); +lean_ctor_set(x_96, 1, x_95); +lean_ctor_set(x_96, 2, x_94); +x_97 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__3; +x_98 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_98, 0, x_70); lean_ctor_set(x_98, 1, x_97); -lean_ctor_set(x_98, 2, x_96); -x_99 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__3; -x_100 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_100, 0, x_67); -lean_ctor_set(x_100, 1, x_99); -x_101 = lean_array_get_size(x_10); -x_102 = l_Array_toSubarray___rarg(x_10, x_8, x_101); -x_103 = l_Array_ofSubarray___rarg(x_102); -x_104 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__27; -x_105 = l_Array_append___rarg(x_104, x_103); -x_106 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_106, 0, x_81); -lean_ctor_set(x_106, 1, x_93); -lean_ctor_set(x_106, 2, x_105); -x_107 = lean_array_push(x_78, x_100); -x_108 = lean_array_push(x_107, x_106); -x_109 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_109, 0, x_81); -lean_ctor_set(x_109, 1, x_4); -lean_ctor_set(x_109, 2, x_108); -x_110 = l_unexpandExists___closed__4; -x_111 = lean_array_push(x_110, x_83); -lean_inc(x_85); -x_112 = lean_array_push(x_111, x_85); -x_113 = lean_array_push(x_112, x_98); -x_114 = lean_array_push(x_113, x_85); -x_115 = lean_array_push(x_114, x_109); -x_116 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_116, 0, x_81); -lean_ctor_set(x_116, 1, x_93); -lean_ctor_set(x_116, 2, x_115); -x_117 = lean_array_push(x_91, x_116); -x_118 = l___aux__Init__NotationExtra______macroRules__tacticCalc____1___closed__4; -x_119 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_119, 0, x_81); -lean_ctor_set(x_119, 1, x_118); -lean_ctor_set(x_119, 2, x_117); -x_120 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_120, 0, x_119); -lean_ctor_set(x_120, 1, x_68); -return x_120; +x_99 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__27; +x_100 = l_Array_append___rarg(x_99, x_23); +x_101 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_101, 0, x_20); +lean_ctor_set(x_101, 1, x_21); +lean_ctor_set(x_101, 2, x_100); +x_102 = lean_array_push(x_81, x_98); +x_103 = lean_array_push(x_102, x_101); +x_104 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_104, 0, x_20); +lean_ctor_set(x_104, 1, x_4); +lean_ctor_set(x_104, 2, x_103); +x_105 = l_unexpandExists___closed__6; +x_106 = lean_array_push(x_105, x_85); +lean_inc(x_87); +x_107 = lean_array_push(x_106, x_87); +x_108 = lean_array_push(x_107, x_96); +x_109 = lean_array_push(x_108, x_87); +x_110 = lean_array_push(x_109, x_104); +x_111 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_111, 0, x_20); +lean_ctor_set(x_111, 1, x_21); +lean_ctor_set(x_111, 2, x_110); +x_112 = lean_array_push(x_90, x_111); +x_113 = l___aux__Init__NotationExtra______macroRules__tacticCalc____1___closed__4; +x_114 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_114, 0, x_20); +lean_ctor_set(x_114, 1, x_113); +lean_ctor_set(x_114, 2, x_112); +x_115 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_115, 0, x_114); +lean_ctor_set(x_115, 1, x_71); +return x_115; +} } } else { -lean_object* x_121; uint8_t x_122; +lean_object* x_116; lean_object* x_117; lean_object* x_118; uint8_t x_119; +x_116 = lean_unsigned_to_nat(0u); +x_117 = l_Lean_Syntax_getArg(x_9, x_116); +lean_dec(x_9); lean_inc(x_2); -x_121 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_122 = !lean_is_exclusive(x_121); -if (x_122 == 0) -{ -lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; -x_123 = lean_ctor_get(x_121, 0); -x_124 = lean_ctor_get(x_2, 2); -lean_inc(x_124); -x_125 = lean_ctor_get(x_2, 1); -lean_inc(x_125); +x_118 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_119 = !lean_is_exclusive(x_118); +if (x_119 == 0) +{ +lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; +x_120 = lean_ctor_get(x_118, 0); +x_121 = lean_ctor_get(x_2, 2); +lean_inc(x_121); +x_122 = lean_ctor_get(x_2, 1); +lean_inc(x_122); lean_dec(x_2); -x_126 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__1; -lean_inc(x_123); -x_127 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_127, 0, x_123); -lean_ctor_set(x_127, 1, x_126); -x_128 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__6; -x_129 = l_Lean_addMacroScope(x_125, x_128, x_124); -x_130 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__5; -x_131 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__8; -lean_inc(x_123); -x_132 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_132, 0, x_123); -lean_ctor_set(x_132, 1, x_130); -lean_ctor_set(x_132, 2, x_129); -lean_ctor_set(x_132, 3, x_131); -x_133 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_134 = lean_array_push(x_133, x_127); -x_135 = lean_array_push(x_134, x_132); -x_136 = lean_box(2); -x_137 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__2; -x_138 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_138, 0, x_136); -lean_ctor_set(x_138, 1, x_137); -lean_ctor_set(x_138, 2, x_135); -x_139 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__9; -lean_inc(x_123); -x_140 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_140, 0, x_123); -lean_ctor_set(x_140, 1, x_139); -x_141 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__10; -x_142 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_142, 0, x_123); -lean_ctor_set(x_142, 1, x_141); -x_143 = l_Lean_instInhabitedSyntax; -x_144 = lean_unsigned_to_nat(0u); -x_145 = lean_array_get(x_143, x_10, x_144); -lean_dec(x_10); -x_146 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_147 = lean_array_push(x_146, x_145); -x_148 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_149 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_149, 0, x_136); -lean_ctor_set(x_149, 1, x_148); -lean_ctor_set(x_149, 2, x_147); -x_150 = lean_array_push(x_133, x_142); -x_151 = lean_array_push(x_150, x_149); -x_152 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__11; -x_153 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_153, 0, x_136); -lean_ctor_set(x_153, 1, x_152); -lean_ctor_set(x_153, 2, x_151); -x_154 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; -x_155 = lean_array_push(x_154, x_138); -x_156 = lean_array_push(x_155, x_140); -x_157 = lean_array_push(x_156, x_153); -x_158 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_158, 0, x_136); -lean_ctor_set(x_158, 1, x_148); -lean_ctor_set(x_158, 2, x_157); -x_159 = lean_array_push(x_146, x_158); -x_160 = l___aux__Init__NotationExtra______macroRules__tacticCalc____1___closed__4; -x_161 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_161, 0, x_136); -lean_ctor_set(x_161, 1, x_160); -lean_ctor_set(x_161, 2, x_159); -lean_ctor_set(x_121, 0, x_161); -return x_121; +x_123 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__1; +lean_inc(x_120); +x_124 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_124, 0, x_120); +lean_ctor_set(x_124, 1, x_123); +x_125 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__6; +x_126 = l_Lean_addMacroScope(x_122, x_125, x_121); +x_127 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__5; +x_128 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__8; +lean_inc(x_120); +x_129 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_129, 0, x_120); +lean_ctor_set(x_129, 1, x_127); +lean_ctor_set(x_129, 2, x_126); +lean_ctor_set(x_129, 3, x_128); +x_130 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_131 = lean_array_push(x_130, x_124); +x_132 = lean_array_push(x_131, x_129); +x_133 = lean_box(2); +x_134 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__2; +x_135 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_135, 0, x_133); +lean_ctor_set(x_135, 1, x_134); +lean_ctor_set(x_135, 2, x_132); +x_136 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__9; +lean_inc(x_120); +x_137 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_137, 0, x_120); +lean_ctor_set(x_137, 1, x_136); +x_138 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__10; +x_139 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_139, 0, x_120); +lean_ctor_set(x_139, 1, x_138); +x_140 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_141 = lean_array_push(x_140, x_117); +x_142 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_143 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_143, 0, x_133); +lean_ctor_set(x_143, 1, x_142); +lean_ctor_set(x_143, 2, x_141); +x_144 = lean_array_push(x_130, x_139); +x_145 = lean_array_push(x_144, x_143); +x_146 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__11; +x_147 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_147, 0, x_133); +lean_ctor_set(x_147, 1, x_146); +lean_ctor_set(x_147, 2, x_145); +x_148 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; +x_149 = lean_array_push(x_148, x_135); +x_150 = lean_array_push(x_149, x_137); +x_151 = lean_array_push(x_150, x_147); +x_152 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_152, 0, x_133); +lean_ctor_set(x_152, 1, x_142); +lean_ctor_set(x_152, 2, x_151); +x_153 = lean_array_push(x_140, x_152); +x_154 = l___aux__Init__NotationExtra______macroRules__tacticCalc____1___closed__4; +x_155 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_155, 0, x_133); +lean_ctor_set(x_155, 1, x_154); +lean_ctor_set(x_155, 2, x_153); +lean_ctor_set(x_118, 0, x_155); +return x_118; } else { -lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; -x_162 = lean_ctor_get(x_121, 0); -x_163 = lean_ctor_get(x_121, 1); -lean_inc(x_163); -lean_inc(x_162); -lean_dec(x_121); -x_164 = lean_ctor_get(x_2, 2); -lean_inc(x_164); -x_165 = lean_ctor_get(x_2, 1); -lean_inc(x_165); +lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; +x_156 = lean_ctor_get(x_118, 0); +x_157 = lean_ctor_get(x_118, 1); +lean_inc(x_157); +lean_inc(x_156); +lean_dec(x_118); +x_158 = lean_ctor_get(x_2, 2); +lean_inc(x_158); +x_159 = lean_ctor_get(x_2, 1); +lean_inc(x_159); lean_dec(x_2); -x_166 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__1; -lean_inc(x_162); -x_167 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_167, 0, x_162); -lean_ctor_set(x_167, 1, x_166); -x_168 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__6; -x_169 = l_Lean_addMacroScope(x_165, x_168, x_164); -x_170 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__5; -x_171 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__8; -lean_inc(x_162); -x_172 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_172, 0, x_162); -lean_ctor_set(x_172, 1, x_170); -lean_ctor_set(x_172, 2, x_169); -lean_ctor_set(x_172, 3, x_171); -x_173 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_174 = lean_array_push(x_173, x_167); -x_175 = lean_array_push(x_174, x_172); -x_176 = lean_box(2); -x_177 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__2; -x_178 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_178, 0, x_176); -lean_ctor_set(x_178, 1, x_177); -lean_ctor_set(x_178, 2, x_175); -x_179 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__9; -lean_inc(x_162); -x_180 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_180, 0, x_162); +x_160 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__1; +lean_inc(x_156); +x_161 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_161, 0, x_156); +lean_ctor_set(x_161, 1, x_160); +x_162 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__6; +x_163 = l_Lean_addMacroScope(x_159, x_162, x_158); +x_164 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__5; +x_165 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__8; +lean_inc(x_156); +x_166 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_166, 0, x_156); +lean_ctor_set(x_166, 1, x_164); +lean_ctor_set(x_166, 2, x_163); +lean_ctor_set(x_166, 3, x_165); +x_167 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_168 = lean_array_push(x_167, x_161); +x_169 = lean_array_push(x_168, x_166); +x_170 = lean_box(2); +x_171 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__2; +x_172 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_172, 0, x_170); +lean_ctor_set(x_172, 1, x_171); +lean_ctor_set(x_172, 2, x_169); +x_173 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__9; +lean_inc(x_156); +x_174 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_174, 0, x_156); +lean_ctor_set(x_174, 1, x_173); +x_175 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__10; +x_176 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_176, 0, x_156); +lean_ctor_set(x_176, 1, x_175); +x_177 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_178 = lean_array_push(x_177, x_117); +x_179 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_180 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_180, 0, x_170); lean_ctor_set(x_180, 1, x_179); -x_181 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__10; -x_182 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_182, 0, x_162); -lean_ctor_set(x_182, 1, x_181); -x_183 = l_Lean_instInhabitedSyntax; -x_184 = lean_unsigned_to_nat(0u); -x_185 = lean_array_get(x_183, x_10, x_184); -lean_dec(x_10); -x_186 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_187 = lean_array_push(x_186, x_185); -x_188 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; +lean_ctor_set(x_180, 2, x_178); +x_181 = lean_array_push(x_167, x_176); +x_182 = lean_array_push(x_181, x_180); +x_183 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__11; +x_184 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_184, 0, x_170); +lean_ctor_set(x_184, 1, x_183); +lean_ctor_set(x_184, 2, x_182); +x_185 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; +x_186 = lean_array_push(x_185, x_172); +x_187 = lean_array_push(x_186, x_174); +x_188 = lean_array_push(x_187, x_184); x_189 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_189, 0, x_176); -lean_ctor_set(x_189, 1, x_188); -lean_ctor_set(x_189, 2, x_187); -x_190 = lean_array_push(x_173, x_182); -x_191 = lean_array_push(x_190, x_189); -x_192 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__11; -x_193 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_193, 0, x_176); -lean_ctor_set(x_193, 1, x_192); -lean_ctor_set(x_193, 2, x_191); -x_194 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; -x_195 = lean_array_push(x_194, x_178); -x_196 = lean_array_push(x_195, x_180); -x_197 = lean_array_push(x_196, x_193); -x_198 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_198, 0, x_176); -lean_ctor_set(x_198, 1, x_188); -lean_ctor_set(x_198, 2, x_197); -x_199 = lean_array_push(x_186, x_198); -x_200 = l___aux__Init__NotationExtra______macroRules__tacticCalc____1___closed__4; -x_201 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_201, 0, x_176); -lean_ctor_set(x_201, 1, x_200); -lean_ctor_set(x_201, 2, x_199); -x_202 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_202, 0, x_201); -lean_ctor_set(x_202, 1, x_163); -return x_202; +lean_ctor_set(x_189, 0, x_170); +lean_ctor_set(x_189, 1, x_179); +lean_ctor_set(x_189, 2, x_188); +x_190 = lean_array_push(x_177, x_189); +x_191 = l___aux__Init__NotationExtra______macroRules__tacticCalc____1___closed__4; +x_192 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_192, 0, x_170); +lean_ctor_set(x_192, 1, x_191); +lean_ctor_set(x_192, 2, x_190); +x_193 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_193, 0, x_192); +lean_ctor_set(x_193, 1, x_157); +return x_193; } } } @@ -11263,9 +11558,9 @@ LEAN_EXPORT lean_object* l_Array_sequenceMap___at___aux__Init__NotationExtra____ { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_3 = lean_array_get_size(x_1); -x_4 = lean_unsigned_to_nat(0u); -x_5 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__27; -x_6 = l_Array_sequenceMap_loop___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_1, x_2, x_3, x_4, x_5); +x_4 = lean_mk_empty_array_with_capacity(x_3); +x_5 = lean_unsigned_to_nat(0u); +x_6 = l_Array_sequenceMap_loop___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__2(x_1, x_2, x_3, x_5, x_4); return x_6; } } @@ -11293,7 +11588,21 @@ goto _start; } } } -static lean_object* _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__1() { +LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_ctor_get(x_1, 5); +lean_inc(x_3); +lean_dec(x_1); +x_4 = l_Lean_SourceInfo_fromRef(x_3); +x_5 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_5, 0, x_4); +lean_ctor_set(x_5, 1, x_2); +return x_5; +} +} +static lean_object* _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__1() { _start: { lean_object* x_1; @@ -11301,22 +11610,22 @@ x_1 = lean_mk_string_from_bytes("List.cons", 9); return x_1; } } -static lean_object* _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__2() { +static lean_object* _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__1; +x_1 = l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__1; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__3() { +static lean_object* _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__1; +x_1 = l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__2; +x_3 = l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__2; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -11324,7 +11633,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__4() { +static lean_object* _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__4() { _start: { lean_object* x_1; @@ -11332,17 +11641,17 @@ x_1 = lean_mk_string_from_bytes("List", 4); return x_1; } } -static lean_object* _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__5() { +static lean_object* _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__4; +x_2 = l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__6() { +static lean_object* _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__6() { _start: { lean_object* x_1; @@ -11350,100 +11659,100 @@ x_1 = lean_mk_string_from_bytes("cons", 4); return x_1; } } -static lean_object* _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__7() { +static lean_object* _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__5; -x_2 = l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__6; +x_1 = l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__5; +x_2 = l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__6; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__8() { +static lean_object* _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__7; +x_2 = l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__7; 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_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__9() { +static lean_object* _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__8; +x_2 = l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__8; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5(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) { _start: { -uint8_t x_7; -x_7 = lean_usize_dec_eq(x_2, x_3); -if (x_7 == 0) +uint8_t x_8; +x_8 = lean_usize_dec_eq(x_3, x_4); +if (x_8 == 0) { -size_t x_8; size_t 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; 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; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_8 = 1; -x_9 = lean_usize_sub(x_2, x_8); -x_10 = lean_array_uget(x_1, x_9); -lean_inc(x_5); -x_11 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_5, x_6); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); +size_t x_9; size_t 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; 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; lean_object* x_30; lean_object* x_31; +x_9 = 1; +x_10 = lean_usize_sub(x_3, x_9); +x_11 = lean_array_uget(x_2, x_10); +lean_inc(x_6); +x_12 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4(x_6, x_7); +x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); -lean_dec(x_11); -x_14 = lean_ctor_get(x_5, 2); +x_14 = lean_ctor_get(x_12, 1); lean_inc(x_14); -x_15 = lean_ctor_get(x_5, 1); +lean_dec(x_12); +x_15 = lean_ctor_get(x_6, 2); lean_inc(x_15); -x_16 = l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__7; -x_17 = l_Lean_addMacroScope(x_15, x_16, x_14); -x_18 = l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__3; -x_19 = l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__9; -x_20 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_20, 0, x_12); -lean_ctor_set(x_20, 1, x_18); -lean_ctor_set(x_20, 2, x_17); -lean_ctor_set(x_20, 3, x_19); -x_21 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_22 = lean_array_push(x_21, x_10); -x_23 = lean_array_push(x_22, x_4); -x_24 = lean_box(2); -x_25 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_26 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -lean_ctor_set(x_26, 2, x_23); -x_27 = lean_array_push(x_21, x_20); -x_28 = lean_array_push(x_27, x_26); -x_29 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__6; -x_30 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_30, 0, x_24); -lean_ctor_set(x_30, 1, x_29); -lean_ctor_set(x_30, 2, x_28); -x_2 = x_9; -x_4 = x_30; -x_6 = x_13; +x_16 = lean_ctor_get(x_6, 1); +lean_inc(x_16); +x_17 = l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__7; +x_18 = l_Lean_addMacroScope(x_16, x_17, x_15); +x_19 = l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__3; +x_20 = l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__9; +x_21 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_21, 0, x_13); +lean_ctor_set(x_21, 1, x_19); +lean_ctor_set(x_21, 2, x_18); +lean_ctor_set(x_21, 3, x_20); +x_22 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_23 = lean_array_push(x_22, x_11); +x_24 = lean_array_push(x_23, x_5); +x_25 = lean_box(2); +x_26 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_27 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +lean_ctor_set(x_27, 2, x_24); +x_28 = lean_array_push(x_22, x_21); +x_29 = lean_array_push(x_28, x_27); +x_30 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__6; +x_31 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_31, 0, x_25); +lean_ctor_set(x_31, 1, x_30); +lean_ctor_set(x_31, 2, x_29); +x_3 = x_10; +x_5 = x_31; +x_7 = x_14; goto _start; } else { -lean_object* x_32; -lean_dec(x_5); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_4); -lean_ctor_set(x_32, 1, x_6); -return x_32; +lean_object* x_33; +lean_dec(x_6); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_5); +lean_ctor_set(x_33, 1, x_7); +return x_33; } } } @@ -11479,7 +11788,7 @@ static lean_object* _init_l___aux__Init__NotationExtra______macroRules__term_x25 { uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = 1; -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__27; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__27; x_3 = lean_box(x_1); x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); @@ -11507,7 +11816,7 @@ static lean_object* _init_l___aux__Init__NotationExtra______macroRules__term_x25 _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4; x_2 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -11525,7 +11834,7 @@ static lean_object* _init_l___aux__Init__NotationExtra______macroRules__term_x25 _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4; x_2 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__7; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -11543,7 +11852,7 @@ static lean_object* _init_l___aux__Init__NotationExtra______macroRules__term_x25 _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4; x_2 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__9; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -11595,8 +11904,8 @@ static lean_object* _init_l___aux__Init__NotationExtra______macroRules__term_x25 { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__27; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__27; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -11662,43 +11971,43 @@ x_12 = lean_unsigned_to_nat(0u); x_13 = lean_nat_dec_lt(x_12, x_11); if (x_13 == 0) { -lean_object* x_176; +lean_object* x_180; lean_dec(x_11); lean_dec(x_10); -x_176 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__27; -x_14 = x_176; -goto block_175; +x_180 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__27; +x_14 = x_180; +goto block_179; } else { -uint8_t x_177; -x_177 = lean_nat_dec_le(x_11, x_11); -if (x_177 == 0) +uint8_t x_181; +x_181 = lean_nat_dec_le(x_11, x_11); +if (x_181 == 0) { -lean_object* x_178; +lean_object* x_182; lean_dec(x_11); lean_dec(x_10); -x_178 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__27; -x_14 = x_178; -goto block_175; +x_182 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__27; +x_14 = x_182; +goto block_179; } else { -size_t x_179; size_t x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; -x_179 = 0; -x_180 = lean_usize_of_nat(x_11); +size_t x_183; size_t x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; +x_183 = 0; +x_184 = lean_usize_of_nat(x_11); lean_dec(x_11); -x_181 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__3; -x_182 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(x_10, x_179, x_180, x_181); +x_185 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__3; +x_186 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(x_10, x_183, x_184, x_185); lean_dec(x_10); -x_183 = lean_ctor_get(x_182, 1); -lean_inc(x_183); -lean_dec(x_182); -x_14 = x_183; -goto block_175; +x_187 = lean_ctor_get(x_186, 1); +lean_inc(x_187); +lean_dec(x_186); +x_14 = x_187; +goto block_179; } } -block_175: +block_179: { lean_object* x_15; lean_object* x_16; x_15 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__4; @@ -11717,65 +12026,65 @@ return x_18; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; x_19 = lean_ctor_get(x_16, 0); lean_inc(x_19); lean_dec(x_16); x_20 = lean_unsigned_to_nat(3u); x_21 = l_Lean_Syntax_getArg(x_1, x_20); lean_dec(x_1); -x_22 = lean_array_get_size(x_19); -x_23 = lean_unsigned_to_nat(8u); -x_24 = lean_nat_dec_lt(x_22, 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; lean_object* x_30; uint8_t x_31; -x_25 = lean_unsigned_to_nat(2u); -x_26 = lean_nat_div(x_22, x_25); -lean_dec(x_22); -x_27 = lean_array_get_size(x_19); -lean_inc(x_26); +x_22 = lean_box(0); +x_23 = lean_array_get_size(x_19); +x_24 = lean_unsigned_to_nat(8u); +x_25 = lean_nat_dec_lt(x_23, x_24); +if (x_25 == 0) +{ +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_26 = lean_unsigned_to_nat(2u); +x_27 = lean_nat_div(x_23, x_26); +lean_dec(x_23); +x_28 = lean_array_get_size(x_19); +lean_inc(x_27); lean_inc(x_19); -x_28 = l_Array_toSubarray___rarg(x_19, x_26, x_27); -x_29 = l_Array_toSubarray___rarg(x_19, x_12, x_26); +x_29 = l_Array_toSubarray___rarg(x_19, x_27, x_28); +x_30 = l_Array_toSubarray___rarg(x_19, x_12, x_27); lean_inc(x_2); -x_30 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_31 = !lean_is_exclusive(x_30); -if (x_31 == 0) +x_31 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_32 = !lean_is_exclusive(x_31); +if (x_32 == 0) { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; size_t x_48; size_t x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; size_t x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; -x_32 = lean_ctor_get(x_30, 0); -x_33 = lean_ctor_get(x_2, 2); -lean_inc(x_33); -x_34 = lean_ctor_get(x_2, 1); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; size_t x_48; size_t x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; size_t x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_33 = lean_ctor_get(x_31, 0); +x_34 = lean_ctor_get(x_2, 2); lean_inc(x_34); +x_35 = lean_ctor_get(x_2, 1); +lean_inc(x_35); lean_dec(x_2); -x_35 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__5; -lean_inc(x_32); -x_36 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_36, 0, x_32); -lean_ctor_set(x_36, 1, x_35); -x_37 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__14; -x_38 = l_Lean_addMacroScope(x_34, x_37, x_33); -x_39 = lean_box(0); +x_36 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__5; +lean_inc(x_33); +x_37 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_37, 0, x_33); +lean_ctor_set(x_37, 1, x_36); +x_38 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__14; +x_39 = l_Lean_addMacroScope(x_35, x_38, x_34); x_40 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__13; -lean_inc(x_32); +lean_inc(x_33); x_41 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_41, 0, x_32); +lean_ctor_set(x_41, 0, x_33); lean_ctor_set(x_41, 1, x_40); -lean_ctor_set(x_41, 2, x_38); -lean_ctor_set(x_41, 3, x_39); -x_42 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__37; -lean_inc(x_32); +lean_ctor_set(x_41, 2, x_39); +lean_ctor_set(x_41, 3, x_22); +x_42 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__44; +lean_inc(x_33); x_43 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_43, 0, x_32); +lean_ctor_set(x_43, 0, x_33); lean_ctor_set(x_43, 1, x_42); x_44 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__16; -lean_inc(x_32); +lean_inc(x_33); x_45 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_45, 0, x_32); +lean_ctor_set(x_45, 0, x_33); lean_ctor_set(x_45, 1, x_44); -x_46 = l_Array_ofSubarray___rarg(x_28); +x_46 = l_Array_ofSubarray___rarg(x_29); x_47 = lean_array_get_size(x_46); x_48 = lean_usize_of_nat(x_47); lean_dec(x_47); @@ -11784,24 +12093,24 @@ x_50 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__t x_51 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__17; x_52 = l_Lean_mkSepArray(x_50, x_51); lean_dec(x_50); -x_53 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__27; +x_53 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__27; x_54 = l_Array_append___rarg(x_53, x_52); x_55 = lean_box(2); -x_56 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; +x_56 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; x_57 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_57, 0, x_55); lean_ctor_set(x_57, 1, x_56); lean_ctor_set(x_57, 2, x_54); x_58 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__18; -lean_inc(x_32); +lean_inc(x_33); x_59 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_59, 0, x_32); +lean_ctor_set(x_59, 0, x_33); lean_ctor_set(x_59, 1, x_58); -x_60 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__13; +x_60 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__13; x_61 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_61, 0, x_32); +lean_ctor_set(x_61, 0, x_33); lean_ctor_set(x_61, 1, x_60); -x_62 = l_unexpandExists___closed__4; +x_62 = l_unexpandExists___closed__6; x_63 = lean_array_push(x_62, x_45); lean_inc(x_63); x_64 = lean_array_push(x_63, x_57); @@ -11826,14 +12135,14 @@ x_76 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_76, 0, x_55); lean_ctor_set(x_76, 1, x_75); lean_ctor_set(x_76, 2, x_74); -x_77 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; +x_77 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; x_78 = lean_array_push(x_77, x_76); x_79 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__8; x_80 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_80, 0, x_55); lean_ctor_set(x_80, 1, x_79); lean_ctor_set(x_80, 2, x_78); -x_81 = l_Array_ofSubarray___rarg(x_29); +x_81 = l_Array_ofSubarray___rarg(x_30); x_82 = lean_array_get_size(x_81); x_83 = lean_usize_of_nat(x_82); lean_dec(x_82); @@ -11853,8 +12162,8 @@ x_92 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_92, 0, x_55); lean_ctor_set(x_92, 1, x_4); lean_ctor_set(x_92, 2, x_91); -x_93 = l_unexpandExists___closed__5; -x_94 = lean_array_push(x_93, x_36); +x_93 = l_unexpandExists___closed__7; +x_94 = lean_array_push(x_93, x_37); x_95 = lean_array_push(x_94, x_80); x_96 = lean_array_push(x_95, x_70); x_97 = lean_array_push(x_96, x_92); @@ -11863,17 +12172,17 @@ x_99 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_99, 0, x_55); lean_ctor_set(x_99, 1, x_98); lean_ctor_set(x_99, 2, x_97); -lean_ctor_set(x_30, 0, x_99); -return x_30; +lean_ctor_set(x_31, 0, x_99); +return x_31; } else { -lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; size_t x_117; size_t x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; size_t x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; -x_100 = lean_ctor_get(x_30, 0); -x_101 = lean_ctor_get(x_30, 1); +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; size_t x_116; size_t x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; size_t x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; +x_100 = lean_ctor_get(x_31, 0); +x_101 = lean_ctor_get(x_31, 1); lean_inc(x_101); lean_inc(x_100); -lean_dec(x_30); +lean_dec(x_31); x_102 = lean_ctor_get(x_2, 2); lean_inc(x_102); x_103 = lean_ctor_get(x_2, 1); @@ -11886,143 +12195,160 @@ lean_ctor_set(x_105, 0, x_100); lean_ctor_set(x_105, 1, x_104); x_106 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__14; x_107 = l_Lean_addMacroScope(x_103, x_106, x_102); -x_108 = lean_box(0); -x_109 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__13; +x_108 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__13; lean_inc(x_100); -x_110 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_110, 0, x_100); -lean_ctor_set(x_110, 1, x_109); -lean_ctor_set(x_110, 2, x_107); -lean_ctor_set(x_110, 3, x_108); -x_111 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__37; +x_109 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_109, 0, x_100); +lean_ctor_set(x_109, 1, x_108); +lean_ctor_set(x_109, 2, x_107); +lean_ctor_set(x_109, 3, x_22); +x_110 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__44; lean_inc(x_100); -x_112 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_112, 0, x_100); -lean_ctor_set(x_112, 1, x_111); -x_113 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__16; +x_111 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_111, 0, x_100); +lean_ctor_set(x_111, 1, x_110); +x_112 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__16; lean_inc(x_100); -x_114 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_114, 0, x_100); -lean_ctor_set(x_114, 1, x_113); -x_115 = l_Array_ofSubarray___rarg(x_28); -x_116 = lean_array_get_size(x_115); -x_117 = lean_usize_of_nat(x_116); -lean_dec(x_116); -x_118 = 0; -x_119 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_117, x_118, x_115); -x_120 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__17; -x_121 = l_Lean_mkSepArray(x_119, x_120); -lean_dec(x_119); -x_122 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__27; -x_123 = l_Array_append___rarg(x_122, x_121); -x_124 = lean_box(2); -x_125 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_126 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_126, 0, x_124); -lean_ctor_set(x_126, 1, x_125); -lean_ctor_set(x_126, 2, x_123); -x_127 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__18; +x_113 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_113, 0, x_100); +lean_ctor_set(x_113, 1, x_112); +x_114 = l_Array_ofSubarray___rarg(x_29); +x_115 = lean_array_get_size(x_114); +x_116 = lean_usize_of_nat(x_115); +lean_dec(x_115); +x_117 = 0; +x_118 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_116, x_117, x_114); +x_119 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__17; +x_120 = l_Lean_mkSepArray(x_118, x_119); +lean_dec(x_118); +x_121 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__27; +x_122 = l_Array_append___rarg(x_121, x_120); +x_123 = lean_box(2); +x_124 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_125 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_125, 0, x_123); +lean_ctor_set(x_125, 1, x_124); +lean_ctor_set(x_125, 2, x_122); +x_126 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__18; lean_inc(x_100); -x_128 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_128, 0, x_100); -lean_ctor_set(x_128, 1, x_127); -x_129 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__13; -x_130 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_130, 0, x_100); -lean_ctor_set(x_130, 1, x_129); -x_131 = l_unexpandExists___closed__4; -x_132 = lean_array_push(x_131, x_114); -lean_inc(x_132); -x_133 = lean_array_push(x_132, x_126); -lean_inc(x_128); -x_134 = lean_array_push(x_133, x_128); -x_135 = lean_array_push(x_134, x_21); -lean_inc(x_130); -x_136 = lean_array_push(x_135, x_130); -x_137 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_137, 0, x_124); -lean_ctor_set(x_137, 1, x_4); -lean_ctor_set(x_137, 2, x_136); -lean_inc(x_110); -x_138 = lean_array_push(x_131, x_110); -x_139 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__15; -x_140 = lean_array_push(x_138, x_139); -x_141 = lean_array_push(x_140, x_139); -x_142 = lean_array_push(x_141, x_112); -x_143 = lean_array_push(x_142, x_137); -x_144 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__10; -x_145 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_145, 0, x_124); -lean_ctor_set(x_145, 1, x_144); -lean_ctor_set(x_145, 2, x_143); -x_146 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_147 = lean_array_push(x_146, x_145); -x_148 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__8; -x_149 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_149, 0, x_124); -lean_ctor_set(x_149, 1, x_148); -lean_ctor_set(x_149, 2, x_147); -x_150 = l_Array_ofSubarray___rarg(x_29); -x_151 = lean_array_get_size(x_150); -x_152 = lean_usize_of_nat(x_151); -lean_dec(x_151); -x_153 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_152, x_118, x_150); -x_154 = l_Lean_mkSepArray(x_153, x_120); -lean_dec(x_153); -x_155 = l_Array_append___rarg(x_122, x_154); -x_156 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_156, 0, x_124); -lean_ctor_set(x_156, 1, x_125); -lean_ctor_set(x_156, 2, x_155); -x_157 = lean_array_push(x_132, x_156); -x_158 = lean_array_push(x_157, x_128); -x_159 = lean_array_push(x_158, x_110); -x_160 = lean_array_push(x_159, x_130); -x_161 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_161, 0, x_124); -lean_ctor_set(x_161, 1, x_4); -lean_ctor_set(x_161, 2, x_160); -x_162 = l_unexpandExists___closed__5; -x_163 = lean_array_push(x_162, x_105); -x_164 = lean_array_push(x_163, x_149); -x_165 = lean_array_push(x_164, x_139); -x_166 = lean_array_push(x_165, x_161); -x_167 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__6; -x_168 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_168, 0, x_124); -lean_ctor_set(x_168, 1, x_167); -lean_ctor_set(x_168, 2, x_166); -x_169 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_169, 0, x_168); -lean_ctor_set(x_169, 1, x_101); -return x_169; +x_127 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_127, 0, x_100); +lean_ctor_set(x_127, 1, x_126); +x_128 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__13; +x_129 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_129, 0, x_100); +lean_ctor_set(x_129, 1, x_128); +x_130 = l_unexpandExists___closed__6; +x_131 = lean_array_push(x_130, x_113); +lean_inc(x_131); +x_132 = lean_array_push(x_131, x_125); +lean_inc(x_127); +x_133 = lean_array_push(x_132, x_127); +x_134 = lean_array_push(x_133, x_21); +lean_inc(x_129); +x_135 = lean_array_push(x_134, x_129); +x_136 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_136, 0, x_123); +lean_ctor_set(x_136, 1, x_4); +lean_ctor_set(x_136, 2, x_135); +lean_inc(x_109); +x_137 = lean_array_push(x_130, x_109); +x_138 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__15; +x_139 = lean_array_push(x_137, x_138); +x_140 = lean_array_push(x_139, x_138); +x_141 = lean_array_push(x_140, x_111); +x_142 = lean_array_push(x_141, x_136); +x_143 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__10; +x_144 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_144, 0, x_123); +lean_ctor_set(x_144, 1, x_143); +lean_ctor_set(x_144, 2, x_142); +x_145 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_146 = lean_array_push(x_145, x_144); +x_147 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__8; +x_148 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_148, 0, x_123); +lean_ctor_set(x_148, 1, x_147); +lean_ctor_set(x_148, 2, x_146); +x_149 = l_Array_ofSubarray___rarg(x_30); +x_150 = lean_array_get_size(x_149); +x_151 = lean_usize_of_nat(x_150); +lean_dec(x_150); +x_152 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_151, x_117, x_149); +x_153 = l_Lean_mkSepArray(x_152, x_119); +lean_dec(x_152); +x_154 = l_Array_append___rarg(x_121, x_153); +x_155 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_155, 0, x_123); +lean_ctor_set(x_155, 1, x_124); +lean_ctor_set(x_155, 2, x_154); +x_156 = lean_array_push(x_131, x_155); +x_157 = lean_array_push(x_156, x_127); +x_158 = lean_array_push(x_157, x_109); +x_159 = lean_array_push(x_158, x_129); +x_160 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_160, 0, x_123); +lean_ctor_set(x_160, 1, x_4); +lean_ctor_set(x_160, 2, x_159); +x_161 = l_unexpandExists___closed__7; +x_162 = lean_array_push(x_161, x_105); +x_163 = lean_array_push(x_162, x_148); +x_164 = lean_array_push(x_163, x_138); +x_165 = lean_array_push(x_164, x_160); +x_166 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__6; +x_167 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_167, 0, x_123); +lean_ctor_set(x_167, 1, x_166); +lean_ctor_set(x_167, 2, x_165); +x_168 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_168, 0, x_167); +lean_ctor_set(x_168, 1, x_101); +return x_168; } } else { -uint8_t x_170; -x_170 = lean_nat_dec_lt(x_12, x_22); -if (x_170 == 0) +uint8_t x_169; +x_169 = lean_nat_dec_lt(x_12, x_23); +if (x_169 == 0) { -lean_object* x_171; -lean_dec(x_22); +lean_object* x_170; +lean_dec(x_23); lean_dec(x_19); lean_dec(x_2); -x_171 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_171, 0, x_21); -lean_ctor_set(x_171, 1, x_3); -return x_171; +x_170 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_170, 0, x_21); +lean_ctor_set(x_170, 1, x_3); +return x_170; } else { -size_t x_172; size_t x_173; lean_object* x_174; -x_172 = lean_usize_of_nat(x_22); -lean_dec(x_22); -x_173 = 0; -x_174 = l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4(x_19, x_172, x_173, x_21, x_2, x_3); +size_t x_171; size_t x_172; lean_object* x_173; lean_object* x_174; uint8_t x_175; +x_171 = lean_usize_of_nat(x_23); +lean_dec(x_23); +x_172 = 0; +x_173 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__2; +x_174 = l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5(x_173, x_19, x_171, x_172, x_21, x_2, x_3); lean_dec(x_19); +x_175 = !lean_is_exclusive(x_174); +if (x_175 == 0) +{ return x_174; } +else +{ +lean_object* x_176; lean_object* x_177; lean_object* x_178; +x_176 = lean_ctor_get(x_174, 0); +x_177 = lean_ctor_get(x_174, 1); +lean_inc(x_177); +lean_inc(x_176); +lean_dec(x_174); +x_178 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_178, 0, x_176); +lean_ctor_set(x_178, 1, x_177); +return x_178; +} +} } } } @@ -12059,17 +12385,18 @@ x_6 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__te return x_6; } } -LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___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_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___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_2); -lean_dec(x_2); +size_t x_8; size_t x_9; lean_object* x_10; x_8 = lean_unbox_usize(x_3); lean_dec(x_3); -x_9 = l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4(x_1, x_7, x_8, x_4, x_5, x_6); +x_9 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_10 = l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5(x_1, x_2, x_8, x_9, x_5, x_6, x_7); +lean_dec(x_2); lean_dec(x_1); -return x_9; +return x_10; } } static lean_object* _init_l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__1() { @@ -12095,7 +12422,7 @@ static lean_object* _init_l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___clo { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__6; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__13; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } @@ -12132,7 +12459,7 @@ static lean_object* _init_l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__4; x_3 = l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); @@ -12164,7 +12491,7 @@ static lean_object* _init_l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__7; x_3 = l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__9; x_4 = lean_alloc_ctor(2, 3, 0); @@ -12179,7 +12506,7 @@ static lean_object* _init_l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___clo { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__25; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__32; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } @@ -12198,7 +12525,7 @@ static lean_object* _init_l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__10; x_3 = l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__12; x_4 = lean_alloc_ctor(2, 3, 0); @@ -12212,9 +12539,9 @@ static lean_object* _init_l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__13; -x_3 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__16; +x_3 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__16; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -12226,7 +12553,7 @@ static lean_object* _init_l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___clo _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_expandExplicitBindersAux_loop___closed__12; +x_1 = l_Lean_expandExplicitBindersAux_loop___closed__13; x_2 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; @@ -12236,9 +12563,9 @@ static lean_object* _init_l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__15; -x_3 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__20; +x_3 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__20; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -12262,7 +12589,7 @@ static lean_object* _init_l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__14; x_3 = l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__17; x_4 = lean_alloc_ctor(2, 3, 0); @@ -12276,7 +12603,7 @@ static lean_object* _init_l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___clo _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__37; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__44; x_2 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; @@ -12286,7 +12613,7 @@ static lean_object* _init_l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__18; x_3 = l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__19; x_4 = lean_alloc_ctor(2, 3, 0); @@ -12318,9 +12645,9 @@ static lean_object* _init_l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_Lean_unifConstraintElem___closed__5; -x_3 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__20; +x_3 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__20; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -12354,7 +12681,7 @@ static lean_object* _init_l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__23; x_3 = l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__25; x_4 = lean_alloc_ctor(2, 3, 0); @@ -12380,7 +12707,7 @@ static lean_object* _init_l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__12; +x_1 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__12; x_2 = l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__27; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -12392,7 +12719,7 @@ static lean_object* _init_l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__22; +x_1 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__22; x_2 = l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__28; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -12404,7 +12731,7 @@ static lean_object* _init_l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__20; x_3 = l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__29; x_4 = lean_alloc_ctor(2, 3, 0); @@ -12470,31 +12797,30 @@ lean_dec(x_1); x_9 = l_Lean_Syntax_isNone(x_8); if (x_9 == 0) { -lean_object* x_10; uint8_t x_11; -x_10 = l_Lean_nullKind; -x_11 = l_Lean_Syntax_isNodeOf(x_8, x_10, x_7); -if (x_11 == 0) +uint8_t x_10; +x_10 = l_Lean_Syntax_matchesNull(x_8, x_7); +if (x_10 == 0) { -lean_object* x_12; +lean_object* x_11; lean_dec(x_6); -x_12 = lean_box(0); -return x_12; +x_11 = lean_box(0); +return x_11; } else { -lean_object* x_13; -x_13 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_13, 0, x_6); -return x_13; +lean_object* x_12; +x_12 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_12, 0, x_6); +return x_12; } } else { -lean_object* x_14; +lean_object* x_13; lean_dec(x_8); -x_14 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_14, 0, x_6); -return x_14; +x_13 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_13, 0, x_6); +return x_13; } } } @@ -12551,8 +12877,8 @@ static lean_object* _init_l___aux__Init__NotationExtra______macroRules__command_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -12577,7 +12903,7 @@ static lean_object* _init_l___aux__Init__NotationExtra______macroRules__command_ _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__27; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__27; x_2 = l_Array_append___rarg(x_1, x_1); return x_2; } @@ -12587,7 +12913,7 @@ static lean_object* _init_l___aux__Init__NotationExtra______macroRules__command_ { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; x_3 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__10; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -12633,7 +12959,7 @@ lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); +lean_dec(x_2); x_16 = lean_box(1); x_17 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_17, 0, x_16); @@ -12642,246 +12968,248 @@ return x_17; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; +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; uint8_t x_25; lean_object* x_26; x_18 = lean_ctor_get(x_15, 0); lean_inc(x_18); lean_dec(x_15); -x_19 = l_Lean_Syntax_getArgs(x_2); -x_20 = lean_unsigned_to_nat(0u); -x_21 = l_Lean_Syntax_getArg(x_3, x_20); -x_22 = l_Lean_Syntax_getId(x_21); -x_23 = l_Lean_Name_hasMacroScopes(x_22); -if (x_23 == 0) -{ -lean_object* x_254; lean_object* x_255; -x_254 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__13; -x_255 = l_Lean_Name_append(x_22, x_254); -lean_dec(x_22); -x_24 = x_255; -goto block_253; +x_19 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__3; +lean_inc(x_2); +x_20 = lean_name_mk_string(x_2, x_19); +x_21 = l_Lean_Syntax_getArgs(x_3); +x_22 = lean_unsigned_to_nat(0u); +x_23 = l_Lean_Syntax_getArg(x_4, x_22); +x_24 = l_Lean_Syntax_getId(x_23); +lean_dec(x_23); +x_25 = l_Lean_Name_hasMacroScopes(x_24); +if (x_25 == 0) +{ +lean_object* x_252; lean_object* x_253; +x_252 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__13; +x_253 = l_Lean_Name_append(x_24, x_252); +lean_dec(x_24); +x_26 = x_253; +goto block_251; } else { -lean_object* x_256; uint8_t x_257; -x_256 = l_Lean_extractMacroScopes(x_22); -x_257 = !lean_is_exclusive(x_256); -if (x_257 == 0) -{ -lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; -x_258 = lean_ctor_get(x_256, 0); -x_259 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__13; -x_260 = l_Lean_Name_append(x_258, x_259); -lean_dec(x_258); -lean_ctor_set(x_256, 0, x_260); -x_261 = l_Lean_MacroScopesView_review(x_256); -x_24 = x_261; -goto block_253; +lean_object* x_254; uint8_t x_255; +x_254 = l_Lean_extractMacroScopes(x_24); +x_255 = !lean_is_exclusive(x_254); +if (x_255 == 0) +{ +lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; +x_256 = lean_ctor_get(x_254, 0); +x_257 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__13; +x_258 = l_Lean_Name_append(x_256, x_257); +lean_dec(x_256); +lean_ctor_set(x_254, 0, x_258); +x_259 = l_Lean_MacroScopesView_review(x_254); +x_26 = x_259; +goto block_251; } else { -lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; -x_262 = lean_ctor_get(x_256, 0); -x_263 = lean_ctor_get(x_256, 1); -x_264 = lean_ctor_get(x_256, 2); -x_265 = lean_ctor_get(x_256, 3); -lean_inc(x_265); -lean_inc(x_264); +lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; +x_260 = lean_ctor_get(x_254, 0); +x_261 = lean_ctor_get(x_254, 1); +x_262 = lean_ctor_get(x_254, 2); +x_263 = lean_ctor_get(x_254, 3); lean_inc(x_263); lean_inc(x_262); -lean_dec(x_256); -x_266 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__13; -x_267 = l_Lean_Name_append(x_262, x_266); -lean_dec(x_262); -x_268 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_268, 0, x_267); -lean_ctor_set(x_268, 1, x_263); -lean_ctor_set(x_268, 2, x_264); -lean_ctor_set(x_268, 3, x_265); -x_269 = l_Lean_MacroScopesView_review(x_268); -x_24 = x_269; -goto block_253; -} -} -block_253: -{ -lean_object* x_25; lean_object* x_26; uint8_t x_27; -x_25 = l_Lean_mkIdentFrom(x_21, x_24); -x_26 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_9, x_10); -x_27 = !lean_is_exclusive(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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; size_t x_49; size_t x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; -x_28 = lean_ctor_get(x_26, 0); -x_29 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__4; -lean_inc(x_4); -x_30 = lean_name_mk_string(x_4, x_29); -x_31 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__2; -lean_inc(x_4); -x_32 = lean_name_mk_string(x_4, x_31); -x_33 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__3; -lean_inc(x_4); -x_34 = lean_name_mk_string(x_4, x_33); -x_35 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__4; -lean_inc(x_28); -x_36 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_36, 0, x_28); -lean_ctor_set(x_36, 1, x_35); -x_37 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_38 = lean_array_push(x_37, x_36); -x_39 = lean_box(2); -x_40 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_34); -lean_ctor_set(x_40, 2, x_38); -x_41 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__27; -x_42 = l_Array_append___rarg(x_41, x_19); -x_43 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_44 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_44, 0, x_39); -lean_ctor_set(x_44, 1, x_43); -lean_ctor_set(x_44, 2, x_42); -x_45 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__5; +lean_inc(x_261); +lean_inc(x_260); +lean_dec(x_254); +x_264 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__13; +x_265 = l_Lean_Name_append(x_260, x_264); +lean_dec(x_260); +x_266 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_266, 0, x_265); +lean_ctor_set(x_266, 1, x_261); +lean_ctor_set(x_266, 2, x_262); +lean_ctor_set(x_266, 3, x_263); +x_267 = l_Lean_MacroScopesView_review(x_266); +x_26 = x_267; +goto block_251; +} +} +block_251: +{ +lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_inc(x_4); -x_46 = lean_name_mk_string(x_4, x_45); -lean_inc(x_28); -x_47 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_47, 0, x_28); -lean_ctor_set(x_47, 1, x_45); -x_48 = lean_array_get_size(x_18); -x_49 = lean_usize_of_nat(x_48); -lean_dec(x_48); -x_50 = 0; -x_51 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_49, x_50, x_18); -x_52 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__17; -x_53 = l_Lean_mkSepArray(x_51, x_52); -lean_dec(x_51); -x_54 = l_Array_append___rarg(x_41, x_53); -x_55 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_55, 0, x_39); -lean_ctor_set(x_55, 1, x_43); -lean_ctor_set(x_55, 2, x_54); -x_56 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_57 = lean_array_push(x_56, x_47); -x_58 = lean_array_push(x_57, x_55); -x_59 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_59, 0, x_39); -lean_ctor_set(x_59, 1, x_46); -lean_ctor_set(x_59, 2, x_58); -x_60 = lean_array_push(x_37, x_59); +x_27 = l_Lean_mkIdentFrom(x_4, x_26); +x_28 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_9, x_10); +x_29 = !lean_is_exclusive(x_28); +if (x_29 == 0) +{ +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; size_t x_51; size_t x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; +x_30 = lean_ctor_get(x_28, 0); +x_31 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__11; +lean_inc(x_5); +x_32 = lean_name_mk_string(x_5, x_31); +x_33 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__2; +lean_inc(x_5); +x_34 = lean_name_mk_string(x_5, x_33); +x_35 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__3; +lean_inc(x_5); +x_36 = lean_name_mk_string(x_5, x_35); +x_37 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__4; +lean_inc(x_30); +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_30); +lean_ctor_set(x_38, 1, x_37); +x_39 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_40 = lean_array_push(x_39, x_38); +x_41 = lean_box(2); +x_42 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_36); +lean_ctor_set(x_42, 2, x_40); +x_43 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__27; +x_44 = l_Array_append___rarg(x_43, x_21); +x_45 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_46 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_46, 0, x_41); +lean_ctor_set(x_46, 1, x_45); +lean_ctor_set(x_46, 2, x_44); +x_47 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__5; +lean_inc(x_5); +x_48 = lean_name_mk_string(x_5, x_47); +lean_inc(x_30); +x_49 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_49, 0, x_30); +lean_ctor_set(x_49, 1, x_47); +x_50 = lean_array_get_size(x_18); +x_51 = lean_usize_of_nat(x_50); +lean_dec(x_50); +x_52 = 0; +x_53 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_51, x_52, x_18); +x_54 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__17; +x_55 = l_Lean_mkSepArray(x_53, x_54); +lean_dec(x_53); +x_56 = l_Array_append___rarg(x_43, x_55); +x_57 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_57, 0, x_41); +lean_ctor_set(x_57, 1, x_45); +lean_ctor_set(x_57, 2, x_56); +x_58 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_59 = lean_array_push(x_58, x_49); +x_60 = lean_array_push(x_59, x_57); x_61 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_61, 0, x_39); -lean_ctor_set(x_61, 1, x_43); +lean_ctor_set(x_61, 0, x_41); +lean_ctor_set(x_61, 1, x_48); lean_ctor_set(x_61, 2, x_60); -x_62 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__6; -lean_inc(x_4); -x_63 = lean_name_mk_string(x_4, x_62); -x_64 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__7; -x_65 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_65, 0, x_39); -lean_ctor_set(x_65, 1, x_63); -lean_ctor_set(x_65, 2, x_64); -x_66 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__38; -x_67 = lean_array_push(x_66, x_40); -x_68 = lean_array_push(x_67, x_3); -x_69 = lean_array_push(x_68, x_44); -x_70 = lean_array_push(x_69, x_61); -x_71 = lean_array_push(x_56, x_5); -x_72 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__8; -x_73 = lean_name_mk_string(x_4, x_72); -lean_inc(x_28); -x_74 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_74, 0, x_28); -lean_ctor_set(x_74, 1, x_72); -x_75 = l_unexpandListNil___rarg___closed__3; -lean_inc(x_28); +x_62 = lean_array_push(x_39, x_61); +x_63 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_63, 0, x_41); +lean_ctor_set(x_63, 1, x_45); +lean_ctor_set(x_63, 2, x_62); +x_64 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__6; +lean_inc(x_5); +x_65 = lean_name_mk_string(x_5, x_64); +x_66 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__7; +x_67 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_67, 0, x_41); +lean_ctor_set(x_67, 1, x_65); +lean_ctor_set(x_67, 2, x_66); +x_68 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__45; +x_69 = lean_array_push(x_68, x_42); +x_70 = lean_array_push(x_69, x_4); +x_71 = lean_array_push(x_70, x_46); +x_72 = lean_array_push(x_71, x_63); +x_73 = lean_array_push(x_58, x_6); +x_74 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__8; +x_75 = lean_name_mk_string(x_5, x_74); +lean_inc(x_30); x_76 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_76, 0, x_28); -lean_ctor_set(x_76, 1, x_75); -x_77 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__3; -lean_inc(x_6); -x_78 = lean_name_mk_string(x_6, x_77); -x_79 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__11; -lean_inc(x_78); -x_80 = lean_name_mk_string(x_78, x_79); -x_81 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__3; -lean_inc(x_78); -x_82 = lean_name_mk_string(x_78, x_81); +lean_ctor_set(x_76, 0, x_30); +lean_ctor_set(x_76, 1, x_74); +x_77 = l_unexpandListNil___rarg___closed__3; +lean_inc(x_30); +x_78 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_78, 0, x_30); +lean_ctor_set(x_78, 1, x_77); +x_79 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__18; +lean_inc(x_20); +x_80 = lean_name_mk_string(x_20, x_79); +x_81 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__3; +lean_inc(x_20); +x_82 = lean_name_mk_string(x_20, x_81); x_83 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_83, 0, x_39); +lean_ctor_set(x_83, 0, x_41); lean_ctor_set(x_83, 1, x_82); -lean_ctor_set(x_83, 2, x_64); -x_84 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__13; -x_85 = lean_name_mk_string(x_6, x_84); +lean_ctor_set(x_83, 2, x_66); +x_84 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__20; +x_85 = lean_name_mk_string(x_2, x_84); x_86 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__9; x_87 = lean_name_mk_string(x_85, x_86); -lean_inc(x_28); +lean_inc(x_30); x_88 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_88, 0, x_28); +lean_ctor_set(x_88, 0, x_30); lean_ctor_set(x_88, 1, x_86); -x_89 = lean_array_push(x_56, x_88); -x_90 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; +x_89 = lean_array_push(x_58, x_88); +x_90 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; x_91 = lean_array_push(x_89, x_90); x_92 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_92, 0, x_39); +lean_ctor_set(x_92, 0, x_41); lean_ctor_set(x_92, 1, x_87); lean_ctor_set(x_92, 2, x_91); -x_93 = lean_array_push(x_56, x_83); +x_93 = lean_array_push(x_58, x_83); x_94 = lean_array_push(x_93, x_92); x_95 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_95, 0, x_39); +lean_ctor_set(x_95, 0, x_41); lean_ctor_set(x_95, 1, x_80); lean_ctor_set(x_95, 2, x_94); -x_96 = lean_array_push(x_37, x_95); +x_96 = lean_array_push(x_39, x_95); x_97 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_97, 0, x_39); -lean_ctor_set(x_97, 1, x_43); +lean_ctor_set(x_97, 0, x_41); +lean_ctor_set(x_97, 1, x_45); lean_ctor_set(x_97, 2, x_96); -x_98 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__13; -lean_inc(x_28); +x_98 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__13; +lean_inc(x_30); x_99 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_99, 0, x_28); +lean_ctor_set(x_99, 0, x_30); lean_ctor_set(x_99, 1, x_98); -x_100 = lean_array_push(x_37, x_25); +x_100 = lean_array_push(x_39, x_27); x_101 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_101, 0, x_39); -lean_ctor_set(x_101, 1, x_43); +lean_ctor_set(x_101, 0, x_41); +lean_ctor_set(x_101, 1, x_45); lean_ctor_set(x_101, 2, x_100); -x_102 = l_unexpandExists___closed__4; -x_103 = lean_array_push(x_102, x_74); -x_104 = lean_array_push(x_103, x_76); +x_102 = l_unexpandExists___closed__6; +x_103 = lean_array_push(x_102, x_76); +x_104 = lean_array_push(x_103, x_78); x_105 = lean_array_push(x_104, x_97); x_106 = lean_array_push(x_105, x_99); x_107 = lean_array_push(x_106, x_101); x_108 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_108, 0, x_39); -lean_ctor_set(x_108, 1, x_73); +lean_ctor_set(x_108, 0, x_41); +lean_ctor_set(x_108, 1, x_75); lean_ctor_set(x_108, 2, x_107); if (lean_obj_tag(x_8) == 0) { lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; -lean_dec(x_78); -lean_dec(x_28); +lean_dec(x_30); +lean_dec(x_20); x_109 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__11; -x_110 = lean_array_push(x_70, x_109); +x_110 = lean_array_push(x_72, x_109); x_111 = lean_array_push(x_110, x_90); -x_112 = lean_array_push(x_111, x_65); +x_112 = lean_array_push(x_111, x_67); x_113 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_113, 0, x_39); -lean_ctor_set(x_113, 1, x_32); +lean_ctor_set(x_113, 0, x_41); +lean_ctor_set(x_113, 1, x_34); lean_ctor_set(x_113, 2, x_112); -x_114 = lean_array_push(x_71, x_113); +x_114 = lean_array_push(x_73, x_113); x_115 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_115, 0, x_39); -lean_ctor_set(x_115, 1, x_30); +lean_ctor_set(x_115, 0, x_41); +lean_ctor_set(x_115, 1, x_32); lean_ctor_set(x_115, 2, x_114); -x_116 = lean_array_push(x_56, x_115); +x_116 = lean_array_push(x_58, x_115); x_117 = lean_array_push(x_116, x_108); x_118 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_118, 0, x_39); -lean_ctor_set(x_118, 1, x_43); +lean_ctor_set(x_118, 0, x_41); +lean_ctor_set(x_118, 1, x_45); lean_ctor_set(x_118, 2, x_117); -lean_ctor_set(x_26, 0, x_118); -return x_26; +lean_ctor_set(x_28, 0, x_118); +return x_28; } else { @@ -12889,85 +13217,85 @@ lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; x_119 = lean_ctor_get(x_8, 0); lean_inc(x_119); lean_dec(x_8); -x_120 = l_Lean_expandExplicitBindersAux_loop___closed__10; -x_121 = lean_name_mk_string(x_78, x_120); -x_122 = l_Lean_expandExplicitBindersAux_loop___closed__12; +x_120 = l_Lean_expandExplicitBindersAux_loop___closed__11; +x_121 = lean_name_mk_string(x_20, x_120); +x_122 = l_Lean_expandExplicitBindersAux_loop___closed__13; x_123 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_123, 0, x_28); +lean_ctor_set(x_123, 0, x_30); lean_ctor_set(x_123, 1, x_122); -x_124 = lean_array_push(x_56, x_123); +x_124 = lean_array_push(x_58, x_123); x_125 = lean_array_push(x_124, x_119); x_126 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_126, 0, x_39); +lean_ctor_set(x_126, 0, x_41); lean_ctor_set(x_126, 1, x_121); lean_ctor_set(x_126, 2, x_125); -x_127 = lean_array_push(x_37, x_126); -x_128 = l_Array_append___rarg(x_41, x_127); +x_127 = lean_array_push(x_39, x_126); +x_128 = l_Array_append___rarg(x_43, x_127); x_129 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_129, 0, x_39); -lean_ctor_set(x_129, 1, x_43); +lean_ctor_set(x_129, 0, x_41); +lean_ctor_set(x_129, 1, x_45); lean_ctor_set(x_129, 2, x_128); -x_130 = lean_array_push(x_70, x_129); +x_130 = lean_array_push(x_72, x_129); x_131 = lean_array_push(x_130, x_90); -x_132 = lean_array_push(x_131, x_65); +x_132 = lean_array_push(x_131, x_67); x_133 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_133, 0, x_39); -lean_ctor_set(x_133, 1, x_32); +lean_ctor_set(x_133, 0, x_41); +lean_ctor_set(x_133, 1, x_34); lean_ctor_set(x_133, 2, x_132); -x_134 = lean_array_push(x_71, x_133); +x_134 = lean_array_push(x_73, x_133); x_135 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_135, 0, x_39); -lean_ctor_set(x_135, 1, x_30); +lean_ctor_set(x_135, 0, x_41); +lean_ctor_set(x_135, 1, x_32); lean_ctor_set(x_135, 2, x_134); -x_136 = lean_array_push(x_56, x_135); +x_136 = lean_array_push(x_58, x_135); x_137 = lean_array_push(x_136, x_108); x_138 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_138, 0, x_39); -lean_ctor_set(x_138, 1, x_43); +lean_ctor_set(x_138, 0, x_41); +lean_ctor_set(x_138, 1, x_45); lean_ctor_set(x_138, 2, x_137); -lean_ctor_set(x_26, 0, x_138); -return x_26; +lean_ctor_set(x_28, 0, x_138); +return x_28; } } else { -lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; size_t x_161; size_t x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; -x_139 = lean_ctor_get(x_26, 0); -x_140 = lean_ctor_get(x_26, 1); +lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; size_t x_161; size_t x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; +x_139 = lean_ctor_get(x_28, 0); +x_140 = lean_ctor_get(x_28, 1); lean_inc(x_140); -lean_inc(x_139); -lean_dec(x_26); -x_141 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__4; -lean_inc(x_4); -x_142 = lean_name_mk_string(x_4, x_141); +lean_inc(x_139); +lean_dec(x_28); +x_141 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__11; +lean_inc(x_5); +x_142 = lean_name_mk_string(x_5, x_141); x_143 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__2; -lean_inc(x_4); -x_144 = lean_name_mk_string(x_4, x_143); +lean_inc(x_5); +x_144 = lean_name_mk_string(x_5, x_143); x_145 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__3; -lean_inc(x_4); -x_146 = lean_name_mk_string(x_4, x_145); +lean_inc(x_5); +x_146 = lean_name_mk_string(x_5, x_145); x_147 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__4; lean_inc(x_139); x_148 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_148, 0, x_139); lean_ctor_set(x_148, 1, x_147); -x_149 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; +x_149 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; x_150 = lean_array_push(x_149, x_148); x_151 = lean_box(2); x_152 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_152, 0, x_151); lean_ctor_set(x_152, 1, x_146); lean_ctor_set(x_152, 2, x_150); -x_153 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__27; -x_154 = l_Array_append___rarg(x_153, x_19); -x_155 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; +x_153 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__27; +x_154 = l_Array_append___rarg(x_153, x_21); +x_155 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; x_156 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_156, 0, x_151); lean_ctor_set(x_156, 1, x_155); lean_ctor_set(x_156, 2, x_154); x_157 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__5; -lean_inc(x_4); -x_158 = lean_name_mk_string(x_4, x_157); +lean_inc(x_5); +x_158 = lean_name_mk_string(x_5, x_157); lean_inc(x_139); x_159 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_159, 0, x_139); @@ -12985,7 +13313,7 @@ x_167 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_167, 0, x_151); lean_ctor_set(x_167, 1, x_155); lean_ctor_set(x_167, 2, x_166); -x_168 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; +x_168 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; x_169 = lean_array_push(x_168, x_159); x_170 = lean_array_push(x_169, x_167); x_171 = lean_alloc_ctor(1, 3, 0); @@ -12998,21 +13326,21 @@ lean_ctor_set(x_173, 0, x_151); lean_ctor_set(x_173, 1, x_155); lean_ctor_set(x_173, 2, x_172); x_174 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__6; -lean_inc(x_4); -x_175 = lean_name_mk_string(x_4, x_174); +lean_inc(x_5); +x_175 = lean_name_mk_string(x_5, x_174); x_176 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__7; x_177 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_177, 0, x_151); lean_ctor_set(x_177, 1, x_175); lean_ctor_set(x_177, 2, x_176); -x_178 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__38; +x_178 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__45; x_179 = lean_array_push(x_178, x_152); -x_180 = lean_array_push(x_179, x_3); +x_180 = lean_array_push(x_179, x_4); x_181 = lean_array_push(x_180, x_156); x_182 = lean_array_push(x_181, x_173); -x_183 = lean_array_push(x_168, x_5); +x_183 = lean_array_push(x_168, x_6); x_184 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__8; -x_185 = lean_name_mk_string(x_4, x_184); +x_185 = lean_name_mk_string(x_5, x_184); lean_inc(x_139); x_186 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_186, 0, x_139); @@ -13022,140 +13350,137 @@ lean_inc(x_139); x_188 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_188, 0, x_139); lean_ctor_set(x_188, 1, x_187); -x_189 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__3; -lean_inc(x_6); -x_190 = lean_name_mk_string(x_6, x_189); -x_191 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__11; -lean_inc(x_190); -x_192 = lean_name_mk_string(x_190, x_191); -x_193 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__3; -lean_inc(x_190); -x_194 = lean_name_mk_string(x_190, x_193); -x_195 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_195, 0, x_151); -lean_ctor_set(x_195, 1, x_194); -lean_ctor_set(x_195, 2, x_176); -x_196 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__13; -x_197 = lean_name_mk_string(x_6, x_196); -x_198 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__9; -x_199 = lean_name_mk_string(x_197, x_198); +x_189 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__18; +lean_inc(x_20); +x_190 = lean_name_mk_string(x_20, x_189); +x_191 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__3; +lean_inc(x_20); +x_192 = lean_name_mk_string(x_20, x_191); +x_193 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_193, 0, x_151); +lean_ctor_set(x_193, 1, x_192); +lean_ctor_set(x_193, 2, x_176); +x_194 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__20; +x_195 = lean_name_mk_string(x_2, x_194); +x_196 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__9; +x_197 = lean_name_mk_string(x_195, x_196); lean_inc(x_139); -x_200 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_200, 0, x_139); -lean_ctor_set(x_200, 1, x_198); -x_201 = lean_array_push(x_168, x_200); -x_202 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; -x_203 = lean_array_push(x_201, x_202); -x_204 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_204, 0, x_151); -lean_ctor_set(x_204, 1, x_199); -lean_ctor_set(x_204, 2, x_203); -x_205 = lean_array_push(x_168, x_195); -x_206 = lean_array_push(x_205, x_204); +x_198 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_198, 0, x_139); +lean_ctor_set(x_198, 1, x_196); +x_199 = lean_array_push(x_168, x_198); +x_200 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; +x_201 = lean_array_push(x_199, x_200); +x_202 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_202, 0, x_151); +lean_ctor_set(x_202, 1, x_197); +lean_ctor_set(x_202, 2, x_201); +x_203 = lean_array_push(x_168, x_193); +x_204 = lean_array_push(x_203, x_202); +x_205 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_205, 0, x_151); +lean_ctor_set(x_205, 1, x_190); +lean_ctor_set(x_205, 2, x_204); +x_206 = lean_array_push(x_149, x_205); x_207 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_207, 0, x_151); -lean_ctor_set(x_207, 1, x_192); +lean_ctor_set(x_207, 1, x_155); lean_ctor_set(x_207, 2, x_206); -x_208 = lean_array_push(x_149, x_207); -x_209 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_209, 0, x_151); -lean_ctor_set(x_209, 1, x_155); -lean_ctor_set(x_209, 2, x_208); -x_210 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__13; +x_208 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__13; lean_inc(x_139); -x_211 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_211, 0, x_139); -lean_ctor_set(x_211, 1, x_210); -x_212 = lean_array_push(x_149, x_25); -x_213 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_213, 0, x_151); -lean_ctor_set(x_213, 1, x_155); -lean_ctor_set(x_213, 2, x_212); -x_214 = l_unexpandExists___closed__4; -x_215 = lean_array_push(x_214, x_186); -x_216 = lean_array_push(x_215, x_188); -x_217 = lean_array_push(x_216, x_209); -x_218 = lean_array_push(x_217, x_211); -x_219 = lean_array_push(x_218, x_213); -x_220 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_220, 0, x_151); -lean_ctor_set(x_220, 1, x_185); -lean_ctor_set(x_220, 2, x_219); +x_209 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_209, 0, x_139); +lean_ctor_set(x_209, 1, x_208); +x_210 = lean_array_push(x_149, x_27); +x_211 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_211, 0, x_151); +lean_ctor_set(x_211, 1, x_155); +lean_ctor_set(x_211, 2, x_210); +x_212 = l_unexpandExists___closed__6; +x_213 = lean_array_push(x_212, x_186); +x_214 = lean_array_push(x_213, x_188); +x_215 = lean_array_push(x_214, x_207); +x_216 = lean_array_push(x_215, x_209); +x_217 = lean_array_push(x_216, x_211); +x_218 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_218, 0, x_151); +lean_ctor_set(x_218, 1, x_185); +lean_ctor_set(x_218, 2, x_217); if (lean_obj_tag(x_8) == 0) { -lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; -lean_dec(x_190); +lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_dec(x_139); -x_221 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__11; -x_222 = lean_array_push(x_182, x_221); -x_223 = lean_array_push(x_222, x_202); -x_224 = lean_array_push(x_223, x_177); +lean_dec(x_20); +x_219 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3___closed__11; +x_220 = lean_array_push(x_182, x_219); +x_221 = lean_array_push(x_220, x_200); +x_222 = lean_array_push(x_221, x_177); +x_223 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_223, 0, x_151); +lean_ctor_set(x_223, 1, x_144); +lean_ctor_set(x_223, 2, x_222); +x_224 = lean_array_push(x_183, x_223); x_225 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_225, 0, x_151); -lean_ctor_set(x_225, 1, x_144); +lean_ctor_set(x_225, 1, x_142); lean_ctor_set(x_225, 2, x_224); -x_226 = lean_array_push(x_183, x_225); -x_227 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_227, 0, x_151); -lean_ctor_set(x_227, 1, x_142); -lean_ctor_set(x_227, 2, x_226); -x_228 = lean_array_push(x_168, x_227); -x_229 = lean_array_push(x_228, x_220); -x_230 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_230, 0, x_151); -lean_ctor_set(x_230, 1, x_155); -lean_ctor_set(x_230, 2, x_229); -x_231 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_231, 0, x_230); -lean_ctor_set(x_231, 1, x_140); -return x_231; +x_226 = lean_array_push(x_168, x_225); +x_227 = lean_array_push(x_226, x_218); +x_228 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_228, 0, x_151); +lean_ctor_set(x_228, 1, x_155); +lean_ctor_set(x_228, 2, x_227); +x_229 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_229, 0, x_228); +lean_ctor_set(x_229, 1, x_140); +return x_229; } else { -lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; -x_232 = lean_ctor_get(x_8, 0); -lean_inc(x_232); +lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; +x_230 = lean_ctor_get(x_8, 0); +lean_inc(x_230); lean_dec(x_8); -x_233 = l_Lean_expandExplicitBindersAux_loop___closed__10; -x_234 = lean_name_mk_string(x_190, x_233); -x_235 = l_Lean_expandExplicitBindersAux_loop___closed__12; -x_236 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_236, 0, x_139); -lean_ctor_set(x_236, 1, x_235); -x_237 = lean_array_push(x_168, x_236); -x_238 = lean_array_push(x_237, x_232); -x_239 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_239, 0, x_151); -lean_ctor_set(x_239, 1, x_234); -lean_ctor_set(x_239, 2, x_238); -x_240 = lean_array_push(x_149, x_239); -x_241 = l_Array_append___rarg(x_153, x_240); -x_242 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_242, 0, x_151); -lean_ctor_set(x_242, 1, x_155); -lean_ctor_set(x_242, 2, x_241); -x_243 = lean_array_push(x_182, x_242); -x_244 = lean_array_push(x_243, x_202); -x_245 = lean_array_push(x_244, x_177); +x_231 = l_Lean_expandExplicitBindersAux_loop___closed__11; +x_232 = lean_name_mk_string(x_20, x_231); +x_233 = l_Lean_expandExplicitBindersAux_loop___closed__13; +x_234 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_234, 0, x_139); +lean_ctor_set(x_234, 1, x_233); +x_235 = lean_array_push(x_168, x_234); +x_236 = lean_array_push(x_235, x_230); +x_237 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_237, 0, x_151); +lean_ctor_set(x_237, 1, x_232); +lean_ctor_set(x_237, 2, x_236); +x_238 = lean_array_push(x_149, x_237); +x_239 = l_Array_append___rarg(x_153, x_238); +x_240 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_240, 0, x_151); +lean_ctor_set(x_240, 1, x_155); +lean_ctor_set(x_240, 2, x_239); +x_241 = lean_array_push(x_182, x_240); +x_242 = lean_array_push(x_241, x_200); +x_243 = lean_array_push(x_242, x_177); +x_244 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_244, 0, x_151); +lean_ctor_set(x_244, 1, x_144); +lean_ctor_set(x_244, 2, x_243); +x_245 = lean_array_push(x_183, x_244); x_246 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_246, 0, x_151); -lean_ctor_set(x_246, 1, x_144); +lean_ctor_set(x_246, 1, x_142); lean_ctor_set(x_246, 2, x_245); -x_247 = lean_array_push(x_183, x_246); -x_248 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_248, 0, x_151); -lean_ctor_set(x_248, 1, x_142); -lean_ctor_set(x_248, 2, x_247); -x_249 = lean_array_push(x_168, x_248); -x_250 = lean_array_push(x_249, x_220); -x_251 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_251, 0, x_151); -lean_ctor_set(x_251, 1, x_155); -lean_ctor_set(x_251, 2, x_250); -x_252 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_252, 0, x_251); -lean_ctor_set(x_252, 1, x_140); -return x_252; +x_247 = lean_array_push(x_168, x_246); +x_248 = lean_array_push(x_247, x_218); +x_249 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_249, 0, x_151); +lean_ctor_set(x_249, 1, x_155); +lean_ctor_set(x_249, 2, x_248); +x_250 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_250, 0, x_249); +lean_ctor_set(x_250, 1, x_140); +return x_250; } } } @@ -13185,7 +13510,7 @@ else lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; x_8 = lean_unsigned_to_nat(0u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); -x_10 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__7; +x_10 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__14; lean_inc(x_9); x_11 = l_Lean_Syntax_isOfKind(x_9, x_10); if (x_11 == 0) @@ -13212,55 +13537,54 @@ x_19 = l_Lean_Syntax_getArg(x_1, x_18); x_20 = l_Lean_Syntax_isNone(x_19); if (x_20 == 0) { -lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_21 = l_Lean_nullKind; -x_22 = lean_unsigned_to_nat(2u); +lean_object* x_21; uint8_t x_22; +x_21 = lean_unsigned_to_nat(2u); lean_inc(x_19); -x_23 = l_Lean_Syntax_isNodeOf(x_19, x_21, x_22); -if (x_23 == 0) +x_22 = l_Lean_Syntax_matchesNull(x_19, x_21); +if (x_22 == 0) { -lean_object* x_24; lean_object* x_25; +lean_object* x_23; lean_object* x_24; lean_dec(x_19); lean_dec(x_17); lean_dec(x_15); lean_dec(x_9); lean_dec(x_2); lean_dec(x_1); -x_24 = lean_box(1); -x_25 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_3); -return x_25; +x_23 = lean_box(1); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_3); +return x_24; } else { -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_object* x_32; -x_26 = lean_unsigned_to_nat(1u); -x_27 = l_Lean_Syntax_getArg(x_19, x_26); +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; +x_25 = lean_unsigned_to_nat(1u); +x_26 = l_Lean_Syntax_getArg(x_19, x_25); lean_dec(x_19); -x_28 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_28, 0, x_27); -x_29 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__3; -x_30 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__2; -x_31 = lean_box(0); -x_32 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3(x_1, x_17, x_15, x_29, x_9, x_30, x_31, x_28, x_2, x_3); +x_27 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_27, 0, x_26); +x_28 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__2; +x_29 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__10; +x_30 = lean_box(0); +x_31 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3(x_1, x_28, x_17, x_15, x_29, x_9, x_30, x_27, x_2, x_3); lean_dec(x_17); lean_dec(x_1); -return x_32; +return x_31; } } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_dec(x_19); -x_33 = lean_box(0); -x_34 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__3; -x_35 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__2; -x_36 = lean_box(0); -x_37 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3(x_1, x_17, x_15, x_34, x_9, x_35, x_36, x_33, x_2, x_3); +x_32 = lean_box(0); +x_33 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__2; +x_34 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__10; +x_35 = lean_box(0); +x_36 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3(x_1, x_33, x_17, x_15, x_34, x_9, x_35, x_32, x_2, x_3); lean_dec(x_17); lean_dec(x_1); -return x_37; +return x_36; } } } @@ -13281,7 +13605,7 @@ LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__command__ lean_object* x_11; x_11 = l___aux__Init__NotationExtra______macroRules__command__ClassAbbrev_____x3a___x3a_x3d_____x2c__1___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_7); -lean_dec(x_2); +lean_dec(x_3); lean_dec(x_1); return x_11; } @@ -13326,7 +13650,7 @@ static lean_object* _init_l_tactic_xb7_x2e_____x3b_____closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__33; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__33; x_2 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; @@ -13336,7 +13660,7 @@ static lean_object* _init_l_tactic_xb7_x2e_____x3b_____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_binderIdent___closed__4; +x_1 = l_Lean_explicitBinders___closed__4; x_2 = l_tactic_xb7_x2e_____x3b_____closed__4; x_3 = l_tactic_xb7_x2e_____x3b_____closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -13378,7 +13702,7 @@ static lean_object* _init_l_tactic_xb7_x2e_____x3b_____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_tactic_xb7_x2e_____x3b_____closed__6; x_3 = l_tactic_xb7_x2e_____x3b_____closed__9; x_4 = lean_alloc_ctor(2, 3, 0); @@ -13462,7 +13786,7 @@ static lean_object* _init_l_tactic_xb7_x2e_____x3b_____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_tactic_xb7_x2e_____x3b_____closed__15; x_3 = l_tactic_xb7_x2e_____x3b_____closed__17; x_4 = lean_alloc_ctor(2, 3, 0); @@ -13476,7 +13800,7 @@ static lean_object* _init_l_tactic_xb7_x2e_____x3b_____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_tactic_xb7_x2e_____x3b_____closed__18; x_3 = l_calc___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); @@ -13514,7 +13838,7 @@ static lean_object* _init_l_tactic_xb7_x2e_____x3b_____closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_tactic_xb7_x2e_____x3b_____closed__10; x_3 = l_tactic_xb7_x2e_____x3b_____closed__21; x_4 = lean_alloc_ctor(2, 3, 0); @@ -13620,9 +13944,9 @@ LEAN_EXPORT lean_object* l_Array_sequenceMap___at___aux__Init__NotationExtra____ { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_3 = lean_array_get_size(x_1); -x_4 = lean_unsigned_to_nat(0u); -x_5 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__27; -x_6 = l_Array_sequenceMap_loop___at___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___spec__2(x_1, x_2, x_3, x_4, x_5); +x_4 = lean_mk_empty_array_with_capacity(x_3); +x_5 = lean_unsigned_to_nat(0u); +x_6 = l_Array_sequenceMap_loop___at___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___spec__2(x_1, x_2, x_3, x_5, x_4); return x_6; } } @@ -13705,7 +14029,7 @@ lean_inc(x_13); x_14 = lean_ctor_get(x_8, 1); lean_inc(x_14); lean_dec(x_8); -x_15 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; +x_15 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; x_16 = lean_array_push(x_15, x_13); if (lean_obj_tag(x_14) == 0) { @@ -13745,7 +14069,7 @@ lean_inc(x_1); x_28 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_28, 0, x_1); lean_ctor_set(x_28, 1, x_27); -x_29 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; +x_29 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; x_30 = lean_array_push(x_29, x_28); lean_inc(x_3); x_31 = l_Array_append___rarg(x_3, x_30); @@ -13776,7 +14100,7 @@ x_40 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___close x_41 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_41, 0, x_39); lean_ctor_set(x_41, 1, x_40); -x_42 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; +x_42 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; x_43 = lean_array_push(x_42, x_41); lean_inc(x_3); x_44 = l_Array_append___rarg(x_3, x_43); @@ -13838,38 +14162,37 @@ lean_dec(x_1); x_9 = l_Lean_Syntax_isNone(x_8); if (x_9 == 0) { -lean_object* x_10; uint8_t x_11; -x_10 = l_Lean_nullKind; +uint8_t x_10; lean_inc(x_8); -x_11 = l_Lean_Syntax_isNodeOf(x_8, x_10, x_7); -if (x_11 == 0) +x_10 = l_Lean_Syntax_matchesNull(x_8, x_7); +if (x_10 == 0) { -lean_object* x_12; +lean_object* x_11; lean_dec(x_8); lean_dec(x_6); -x_12 = lean_box(0); -return x_12; +x_11 = lean_box(0); +return x_11; } else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_13 = l_Lean_Syntax_getArg(x_8, x_5); +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = l_Lean_Syntax_getArg(x_8, x_5); lean_dec(x_8); -x_14 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_14, 0, x_13); -x_15 = lean_box(0); -x_16 = l___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___lambda__1(x_6, x_15, x_14); -return x_16; +x_13 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_13, 0, x_12); +x_14 = lean_box(0); +x_15 = l___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___lambda__1(x_6, x_14, x_13); +return x_15; } } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_dec(x_8); +x_16 = lean_box(0); x_17 = lean_box(0); -x_18 = lean_box(0); -x_19 = l___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___lambda__1(x_6, x_18, x_17); -return x_19; +x_18 = l___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___lambda__1(x_6, x_17, x_16); +return x_18; } } } @@ -13969,8 +14292,8 @@ lean_dec(x_22); x_28 = lean_array_get_size(x_27); x_29 = lean_usize_of_nat(x_28); lean_dec(x_28); -x_30 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_31 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__27; +x_30 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_31 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__27; lean_inc(x_25); x_32 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___spec__5(x_25, x_30, x_31, x_29, x_20, x_27); x_33 = l_Array_append___rarg(x_31, x_32); @@ -13991,7 +14314,7 @@ x_38 = l_unexpandSubtype___closed__3; x_39 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_39, 0, x_25); lean_ctor_set(x_39, 1, x_38); -x_40 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; +x_40 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; x_41 = lean_array_push(x_40, x_39); x_42 = lean_array_push(x_41, x_35); x_43 = lean_array_push(x_42, x_37); @@ -14014,7 +14337,7 @@ x_47 = l_unexpandSubtype___closed__3; x_48 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_48, 0, x_46); lean_ctor_set(x_48, 1, x_47); -x_49 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; +x_49 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; x_50 = lean_array_push(x_49, x_48); x_51 = lean_array_push(x_50, x_35); x_52 = lean_array_push(x_51, x_37); @@ -14043,8 +14366,8 @@ lean_dec(x_22); x_59 = lean_array_get_size(x_58); x_60 = lean_usize_of_nat(x_59); lean_dec(x_59); -x_61 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; -x_62 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__27; +x_61 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; +x_62 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__27; lean_inc(x_55); x_63 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___spec__5(x_55, x_61, x_62, x_60, x_20, x_58); x_64 = l_Array_append___rarg(x_62, x_63); @@ -14065,7 +14388,7 @@ x_69 = l_unexpandSubtype___closed__3; x_70 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_70, 0, x_55); lean_ctor_set(x_70, 1, x_69); -x_71 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; +x_71 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; x_72 = lean_array_push(x_71, x_70); x_73 = lean_array_push(x_72, x_66); x_74 = lean_array_push(x_73, x_68); @@ -14090,7 +14413,7 @@ x_79 = l_unexpandSubtype___closed__3; x_80 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_80, 0, x_78); lean_ctor_set(x_80, 1, x_79); -x_81 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; +x_81 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; x_82 = lean_array_push(x_81, x_80); x_83 = lean_array_push(x_82, x_66); x_84 = lean_array_push(x_83, x_68); @@ -14224,7 +14547,7 @@ static lean_object* _init_l_solve___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_Lean_unifConstraintElem___closed__5; x_3 = l_solve___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -14266,7 +14589,7 @@ static lean_object* _init_l_solve___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_solve___closed__6; x_3 = l_solve___closed__9; x_4 = lean_alloc_ctor(2, 3, 0); @@ -14304,7 +14627,7 @@ static lean_object* _init_l_solve___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__22; +x_1 = l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__22; x_2 = l_solve___closed__12; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -14316,7 +14639,7 @@ static lean_object* _init_l_solve___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_solve___closed__4; x_3 = l_solve___closed__13; x_4 = lean_alloc_ctor(2, 3, 0); @@ -14348,7 +14671,87 @@ x_1 = l_solve___closed__15; return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__solve__1___spec__1___closed__1() { +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at___aux__Init__NotationExtra______macroRules__solve__1___spec__2(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; uint8_t x_7; +x_6 = lean_array_get_size(x_1); +x_7 = lean_nat_dec_lt(x_4, x_6); +lean_dec(x_6); +if (x_7 == 0) +{ +lean_object* x_8; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_5); +return x_8; +} +else +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_unsigned_to_nat(0u); +x_10 = lean_nat_dec_eq(x_3, x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_sub(x_3, x_11); +lean_dec(x_3); +x_13 = lean_array_fget(x_1, x_4); +lean_inc(x_2); +x_14 = lean_apply_1(x_2, x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; +lean_dec(x_12); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_15 = lean_box(0); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_14, 0); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_nat_add(x_4, x_11); +lean_dec(x_4); +x_18 = lean_array_push(x_5, x_16); +x_3 = x_12; +x_4 = x_17; +x_5 = x_18; +goto _start; +} +} +else +{ +lean_object* x_20; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_20 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_20, 0, x_5); +return x_20; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__solve__1___spec__1(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; +x_3 = lean_array_get_size(x_1); +x_4 = lean_mk_empty_array_with_capacity(x_3); +x_5 = lean_unsigned_to_nat(0u); +x_6 = l_Array_sequenceMap_loop___at___aux__Init__NotationExtra______macroRules__solve__1___spec__2(x_1, x_2, x_3, x_5, x_4); +return x_6; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__solve__1___spec__3___closed__1() { _start: { lean_object* x_1; @@ -14356,7 +14759,7 @@ x_1 = lean_mk_string_from_bytes("done", 4); return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__solve__1___spec__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, size_t x_8, size_t x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__solve__1___spec__3(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, size_t x_8, size_t x_9, lean_object* x_10) { _start: { uint8_t x_11; @@ -14374,7 +14777,7 @@ return x_10; } else { -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; 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; size_t x_59; size_t x_60; lean_object* x_61; +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; 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; size_t x_57; size_t x_58; lean_object* x_59; x_12 = lean_array_uget(x_10, x_9); x_13 = lean_unsigned_to_nat(0u); x_14 = lean_array_uset(x_10, x_9, x_13); @@ -14383,124 +14786,108 @@ lean_inc(x_1); x_16 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_16, 0, x_1); lean_ctor_set(x_16, 1, x_15); -x_17 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__20; +x_17 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__20; lean_inc(x_2); x_18 = lean_name_mk_string(x_2, x_17); -x_19 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__22; +x_19 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__22; lean_inc(x_1); x_20 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_20, 0, x_1); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; -x_22 = lean_array_push(x_21, x_12); -x_23 = lean_box(2); -lean_inc(x_3); -x_24 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_3); -lean_ctor_set(x_24, 2, x_22); -x_25 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__29; +x_21 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__29; lean_inc(x_1); -x_26 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_26, 0, x_1); -lean_ctor_set(x_26, 1, x_25); -x_27 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; -x_28 = lean_array_push(x_27, x_20); -x_29 = lean_array_push(x_28, x_24); -x_30 = lean_array_push(x_29, x_26); -x_31 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_31, 0, x_23); -lean_ctor_set(x_31, 1, x_18); -lean_ctor_set(x_31, 2, x_30); -x_32 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__9; +x_22 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_22, 0, x_1); +lean_ctor_set(x_22, 1, x_21); +x_23 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; +x_24 = lean_array_push(x_23, x_20); +x_25 = lean_array_push(x_24, x_12); +x_26 = lean_array_push(x_25, x_22); +x_27 = lean_box(2); +x_28 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_18); +lean_ctor_set(x_28, 2, x_26); +x_29 = l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__9; lean_inc(x_1); -x_33 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_33, 0, x_1); -lean_ctor_set(x_33, 1, x_32); -x_34 = lean_array_push(x_21, x_33); +x_30 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_30, 0, x_1); +lean_ctor_set(x_30, 1, x_29); +x_31 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; +x_32 = lean_array_push(x_31, x_30); lean_inc(x_5); -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_23); -lean_ctor_set(x_35, 1, x_5); -lean_ctor_set(x_35, 2, x_34); -x_36 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_37 = lean_array_push(x_36, x_31); -x_38 = lean_array_push(x_37, x_35); +x_33 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_33, 0, x_27); +lean_ctor_set(x_33, 1, x_5); +lean_ctor_set(x_33, 2, x_32); +x_34 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_35 = lean_array_push(x_34, x_28); +x_36 = lean_array_push(x_35, x_33); lean_inc(x_6); -x_39 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_39, 0, x_23); -lean_ctor_set(x_39, 1, x_6); -lean_ctor_set(x_39, 2, x_38); -x_40 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__solve__1___spec__1___closed__1; +x_37 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_37, 0, x_27); +lean_ctor_set(x_37, 1, x_6); +lean_ctor_set(x_37, 2, x_36); +x_38 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__solve__1___spec__3___closed__1; lean_inc(x_2); -x_41 = lean_name_mk_string(x_2, x_40); +x_39 = lean_name_mk_string(x_2, x_38); lean_inc(x_1); -x_42 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_42, 0, x_1); -lean_ctor_set(x_42, 1, x_40); -x_43 = lean_array_push(x_21, x_42); -x_44 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_44, 0, x_23); -lean_ctor_set(x_44, 1, x_41); -lean_ctor_set(x_44, 2, x_43); +x_40 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_40, 0, x_1); +lean_ctor_set(x_40, 1, x_38); +x_41 = lean_array_push(x_31, x_40); +x_42 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_42, 0, x_27); +lean_ctor_set(x_42, 1, x_39); +lean_ctor_set(x_42, 2, x_41); lean_inc(x_7); lean_inc(x_5); -x_45 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_45, 0, x_23); -lean_ctor_set(x_45, 1, x_5); -lean_ctor_set(x_45, 2, x_7); -x_46 = lean_array_push(x_36, x_44); -x_47 = lean_array_push(x_46, x_45); +x_43 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_43, 0, x_27); +lean_ctor_set(x_43, 1, x_5); +lean_ctor_set(x_43, 2, x_7); +x_44 = lean_array_push(x_34, x_42); +x_45 = lean_array_push(x_44, x_43); lean_inc(x_6); -x_48 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_48, 0, x_23); -lean_ctor_set(x_48, 1, x_6); -lean_ctor_set(x_48, 2, x_47); -x_49 = lean_array_push(x_36, x_39); -x_50 = lean_array_push(x_49, x_48); +x_46 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_46, 0, x_27); +lean_ctor_set(x_46, 1, x_6); +lean_ctor_set(x_46, 2, x_45); +x_47 = lean_array_push(x_34, x_37); +x_48 = lean_array_push(x_47, x_46); lean_inc(x_5); +x_49 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_49, 0, x_27); +lean_ctor_set(x_49, 1, x_5); +lean_ctor_set(x_49, 2, x_48); +x_50 = lean_array_push(x_31, x_49); +lean_inc(x_4); x_51 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_51, 0, x_23); -lean_ctor_set(x_51, 1, x_5); +lean_ctor_set(x_51, 0, x_27); +lean_ctor_set(x_51, 1, x_4); lean_ctor_set(x_51, 2, x_50); -x_52 = lean_array_push(x_21, x_51); -lean_inc(x_4); +x_52 = lean_array_push(x_31, x_51); +lean_inc(x_3); x_53 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_53, 0, x_23); -lean_ctor_set(x_53, 1, x_4); +lean_ctor_set(x_53, 0, x_27); +lean_ctor_set(x_53, 1, x_3); lean_ctor_set(x_53, 2, x_52); -x_54 = lean_array_push(x_21, x_53); -lean_inc(x_3); -x_55 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_55, 0, x_23); -lean_ctor_set(x_55, 1, x_3); -lean_ctor_set(x_55, 2, x_54); -x_56 = lean_array_push(x_36, x_16); -x_57 = lean_array_push(x_56, x_55); +x_54 = lean_array_push(x_34, x_16); +x_55 = lean_array_push(x_54, x_53); lean_inc(x_6); -x_58 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_58, 0, x_23); -lean_ctor_set(x_58, 1, x_6); -lean_ctor_set(x_58, 2, x_57); -x_59 = 1; -x_60 = lean_usize_add(x_9, x_59); -x_61 = lean_array_uset(x_14, x_9, x_58); -x_9 = x_60; -x_10 = x_61; +x_56 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_56, 0, x_27); +lean_ctor_set(x_56, 1, x_6); +lean_ctor_set(x_56, 2, x_55); +x_57 = 1; +x_58 = lean_usize_add(x_9, x_57); +x_59 = lean_array_uset(x_14, x_9, x_56); +x_9 = x_58; +x_10 = x_59; goto _start; } } } -static lean_object* _init_l___aux__Init__NotationExtra______macroRules__solve__1___lambda__1___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Init__NotationExtra______macroRules__tacticCalc____1___closed__2; -x_2 = l_solve___closed__7; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__solve__1___lambda__1(lean_object* x_1) { _start: { @@ -14517,30 +14904,13 @@ return x_4; } else { -lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_unsigned_to_nat(1u); x_6 = l_Lean_Syntax_getArg(x_1, x_5); lean_dec(x_1); -x_7 = l___aux__Init__NotationExtra______macroRules__solve__1___lambda__1___closed__1; -lean_inc(x_6); -x_8 = l_Lean_Syntax_isOfKind(x_6, x_7); -if (x_8 == 0) -{ -lean_object* x_9; -lean_dec(x_6); -x_9 = lean_box(0); -return x_9; -} -else -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_unsigned_to_nat(0u); -x_11 = l_Lean_Syntax_getArg(x_6, x_10); -lean_dec(x_6); -x_12 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_12, 0, x_11); -return x_12; -} +x_7 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_7, 0, x_6); +return x_7; } } } @@ -14573,22 +14943,32 @@ return x_3; static lean_object* _init_l___aux__Init__NotationExtra______macroRules__solve__1___closed__4() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___aux__Init__NotationExtra______macroRules__tacticCalc____1___closed__2; +x_2 = l_solve___closed__7; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___aux__Init__NotationExtra______macroRules__solve__1___closed__5() { +_start: +{ lean_object* x_1; x_1 = lean_mk_string_from_bytes("tacticSeq1Indented", 18); return x_1; } } -static lean_object* _init_l___aux__Init__NotationExtra______macroRules__solve__1___closed__5() { +static lean_object* _init_l___aux__Init__NotationExtra______macroRules__solve__1___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___aux__Init__NotationExtra______macroRules__tacticCalc____1___closed__2; -x_2 = l___aux__Init__NotationExtra______macroRules__solve__1___closed__4; +x_2 = l___aux__Init__NotationExtra______macroRules__solve__1___closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Init__NotationExtra______macroRules__solve__1___closed__6() { +static lean_object* _init_l___aux__Init__NotationExtra______macroRules__solve__1___closed__7() { _start: { lean_object* x_1; @@ -14596,12 +14976,12 @@ x_1 = lean_mk_string_from_bytes("first", 5); return x_1; } } -static lean_object* _init_l___aux__Init__NotationExtra______macroRules__solve__1___closed__7() { +static lean_object* _init_l___aux__Init__NotationExtra______macroRules__solve__1___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___aux__Init__NotationExtra______macroRules__tacticCalc____1___closed__2; -x_2 = l___aux__Init__NotationExtra______macroRules__solve__1___closed__6; +x_2 = l___aux__Init__NotationExtra______macroRules__solve__1___closed__7; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } @@ -14633,7 +15013,7 @@ lean_dec(x_1); x_10 = l_Lean_Syntax_getArgs(x_9); lean_dec(x_9); x_11 = l___aux__Init__NotationExtra______macroRules__solve__1___closed__1; -x_12 = l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__1(x_10, x_11); +x_12 = l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__solve__1___spec__1(x_10, x_11); lean_dec(x_10); if (lean_obj_tag(x_12) == 0) { @@ -14662,7 +15042,7 @@ lean_inc(x_18); x_20 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_20, 0, x_18); lean_ctor_set(x_20, 1, x_19); -x_21 = l___aux__Init__NotationExtra______macroRules__solve__1___closed__6; +x_21 = l___aux__Init__NotationExtra______macroRules__solve__1___closed__7; lean_inc(x_18); x_22 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_22, 0, x_18); @@ -14672,34 +15052,34 @@ x_24 = lean_usize_of_nat(x_23); lean_dec(x_23); x_25 = 0; x_26 = l___aux__Init__NotationExtra______macroRules__tacticCalc____1___closed__2; -x_27 = l___aux__Init__NotationExtra______macroRules__solve__1___lambda__1___closed__1; -x_28 = l___aux__Init__NotationExtra______macroRules__solve__1___closed__5; -x_29 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; +x_27 = l___aux__Init__NotationExtra______macroRules__solve__1___closed__4; +x_28 = l___aux__Init__NotationExtra______macroRules__solve__1___closed__6; +x_29 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; x_30 = l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__22; -x_31 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__27; -x_32 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__solve__1___spec__1(x_18, x_26, x_27, x_28, x_29, x_30, x_31, x_24, x_25, x_15); +x_31 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__27; +x_32 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__solve__1___spec__3(x_18, x_26, x_27, x_28, x_29, x_30, x_31, x_24, x_25, x_15); x_33 = l_Array_append___rarg(x_31, x_32); x_34 = lean_box(2); x_35 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_35, 0, x_34); lean_ctor_set(x_35, 1, x_29); lean_ctor_set(x_35, 2, x_33); -x_36 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; +x_36 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; x_37 = lean_array_push(x_36, x_22); x_38 = lean_array_push(x_37, x_35); -x_39 = l___aux__Init__NotationExtra______macroRules__solve__1___closed__7; +x_39 = l___aux__Init__NotationExtra______macroRules__solve__1___closed__8; x_40 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_40, 0, x_34); lean_ctor_set(x_40, 1, x_39); lean_ctor_set(x_40, 2, x_38); x_41 = lean_array_push(x_36, x_40); -x_42 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; +x_42 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; x_43 = lean_array_push(x_41, x_42); x_44 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_44, 0, x_34); lean_ctor_set(x_44, 1, x_30); lean_ctor_set(x_44, 2, x_43); -x_45 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; +x_45 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; x_46 = lean_array_push(x_45, x_44); x_47 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_47, 0, x_34); @@ -14738,7 +15118,7 @@ lean_inc(x_56); x_59 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_59, 0, x_56); lean_ctor_set(x_59, 1, x_58); -x_60 = l___aux__Init__NotationExtra______macroRules__solve__1___closed__6; +x_60 = l___aux__Init__NotationExtra______macroRules__solve__1___closed__7; lean_inc(x_56); x_61 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_61, 0, x_56); @@ -14748,34 +15128,34 @@ x_63 = lean_usize_of_nat(x_62); lean_dec(x_62); x_64 = 0; x_65 = l___aux__Init__NotationExtra______macroRules__tacticCalc____1___closed__2; -x_66 = l___aux__Init__NotationExtra______macroRules__solve__1___lambda__1___closed__1; -x_67 = l___aux__Init__NotationExtra______macroRules__solve__1___closed__5; -x_68 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; +x_66 = l___aux__Init__NotationExtra______macroRules__solve__1___closed__4; +x_67 = l___aux__Init__NotationExtra______macroRules__solve__1___closed__6; +x_68 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; x_69 = l_command__ClassAbbrev_____x3a___x3a_x3d_____x2c___closed__22; -x_70 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__27; -x_71 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__solve__1___spec__1(x_56, x_65, x_66, x_67, x_68, x_69, x_70, x_63, x_64, x_15); +x_70 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__27; +x_71 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__solve__1___spec__3(x_56, x_65, x_66, x_67, x_68, x_69, x_70, x_63, x_64, x_15); x_72 = l_Array_append___rarg(x_70, x_71); x_73 = lean_box(2); x_74 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_74, 0, x_73); lean_ctor_set(x_74, 1, x_68); lean_ctor_set(x_74, 2, x_72); -x_75 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; +x_75 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; x_76 = lean_array_push(x_75, x_61); x_77 = lean_array_push(x_76, x_74); -x_78 = l___aux__Init__NotationExtra______macroRules__solve__1___closed__7; +x_78 = l___aux__Init__NotationExtra______macroRules__solve__1___closed__8; x_79 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_79, 0, x_73); lean_ctor_set(x_79, 1, x_78); lean_ctor_set(x_79, 2, x_77); x_80 = lean_array_push(x_75, x_79); -x_81 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; +x_81 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; x_82 = lean_array_push(x_80, x_81); x_83 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_83, 0, x_73); lean_ctor_set(x_83, 1, x_69); lean_ctor_set(x_83, 2, x_82); -x_84 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; +x_84 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; x_85 = lean_array_push(x_84, x_83); x_86 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_86, 0, x_73); @@ -14807,7 +15187,25 @@ return x_95; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__solve__1___spec__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, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at___aux__Init__NotationExtra______macroRules__solve__1___spec__2___boxed(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; +x_6 = l_Array_sequenceMap_loop___at___aux__Init__NotationExtra______macroRules__solve__1___spec__2(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__solve__1___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__solve__1___spec__1(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__solve__1___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_object* x_9, lean_object* x_10) { _start: { size_t x_11; size_t x_12; lean_object* x_13; @@ -14815,7 +15213,7 @@ x_11 = lean_unbox_usize(x_8); lean_dec(x_8); x_12 = lean_unbox_usize(x_9); lean_dec(x_9); -x_13 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__solve__1___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_10); +x_13 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__solve__1___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_10); return x_13; } } @@ -14986,7 +15384,7 @@ static lean_object* _init_l_Lean_doElemRepeat_____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__2; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__2; x_2 = l_Lean_doElemRepeat_____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -15042,7 +15440,7 @@ static lean_object* _init_l_Lean_doElemRepeat_____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_Lean_doElemRepeat_____closed__4; x_3 = l_Lean_doElemRepeat_____closed__7; x_4 = lean_alloc_ctor(2, 3, 0); @@ -15086,7 +15484,7 @@ static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lea _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -15112,7 +15510,7 @@ static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lea _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -15189,7 +15587,7 @@ static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lea _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__2; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__2; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__10; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -15233,8 +15631,8 @@ static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lea _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_unexpandExists___closed__5; -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; +x_1 = l_unexpandExists___closed__7; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -15288,12 +15686,12 @@ lean_inc(x_12); x_16 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_16, 0, x_12); lean_ctor_set(x_16, 1, x_15); -x_17 = l_Lean_binderIdent___closed__5; +x_17 = l_Lean_expandExplicitBindersAux_loop___closed__7; lean_inc(x_12); x_18 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_18, 0, x_12); lean_ctor_set(x_18, 1, x_17); -x_19 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; +x_19 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; x_20 = lean_array_push(x_19, x_18); x_21 = lean_box(2); x_22 = l_Lean_expandExplicitBindersAux_loop___closed__6; @@ -15326,7 +15724,7 @@ lean_ctor_set(x_36, 0, x_21); lean_ctor_set(x_36, 1, x_35); lean_ctor_set(x_36, 2, x_34); x_37 = lean_array_push(x_19, x_36); -x_38 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; +x_38 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; x_39 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_39, 0, x_21); lean_ctor_set(x_39, 1, x_38); @@ -15335,7 +15733,7 @@ x_40 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____ x_41 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_41, 0, x_12); lean_ctor_set(x_41, 1, x_40); -x_42 = l_unexpandExists___closed__5; +x_42 = l_unexpandExists___closed__7; x_43 = lean_array_push(x_42, x_16); x_44 = lean_array_push(x_43, x_39); x_45 = lean_array_push(x_44, x_41); @@ -15366,12 +15764,12 @@ lean_inc(x_49); x_54 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_54, 0, x_49); lean_ctor_set(x_54, 1, x_53); -x_55 = l_Lean_binderIdent___closed__5; +x_55 = l_Lean_expandExplicitBindersAux_loop___closed__7; lean_inc(x_49); x_56 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_56, 0, x_49); lean_ctor_set(x_56, 1, x_55); -x_57 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; +x_57 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; x_58 = lean_array_push(x_57, x_56); x_59 = lean_box(2); x_60 = l_Lean_expandExplicitBindersAux_loop___closed__6; @@ -15404,7 +15802,7 @@ lean_ctor_set(x_74, 0, x_59); lean_ctor_set(x_74, 1, x_73); lean_ctor_set(x_74, 2, x_72); x_75 = lean_array_push(x_57, x_74); -x_76 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; +x_76 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; x_77 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_77, 0, x_59); lean_ctor_set(x_77, 1, x_76); @@ -15413,7 +15811,7 @@ x_78 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____ x_79 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_79, 0, x_49); lean_ctor_set(x_79, 1, x_78); -x_80 = l_unexpandExists___closed__5; +x_80 = l_unexpandExists___closed__7; x_81 = lean_array_push(x_80, x_54); x_82 = lean_array_push(x_81, x_77); x_83 = lean_array_push(x_82, x_79); @@ -15443,7 +15841,7 @@ static lean_object* _init_l_Lean_doElemWhile__Do_____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__2; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__2; x_2 = l_Lean_doElemWhile__Do_____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -15499,7 +15897,7 @@ static lean_object* _init_l_Lean_doElemWhile__Do_____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_Lean_doElemWhile__Do_____closed__4; x_3 = l_Lean_doElemWhile__Do_____closed__7; x_4 = lean_alloc_ctor(2, 3, 0); @@ -15531,7 +15929,7 @@ static lean_object* _init_l_Lean_doElemWhile__Do_____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_Lean_doElemWhile__Do_____closed__8; x_3 = l_Lean_doElemWhile__Do_____closed__10; x_4 = lean_alloc_ctor(2, 3, 0); @@ -15545,7 +15943,7 @@ static lean_object* _init_l_Lean_doElemWhile__Do_____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_Lean_doElemWhile__Do_____closed__11; x_3 = l_Lean_doElemRepeat_____closed__7; x_4 = lean_alloc_ctor(2, 3, 0); @@ -15597,7 +15995,7 @@ static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lea _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile__Do____1___closed__2; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -15615,7 +16013,7 @@ static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lea _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile__Do____1___closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -15633,7 +16031,7 @@ static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lea _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile__Do____1___closed__6; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -15651,7 +16049,7 @@ static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lea _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile__Do____1___closed__8; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -15661,8 +16059,8 @@ static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lea _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -15679,7 +16077,7 @@ static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lea _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile__Do____1___closed__11; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -15757,16 +16155,16 @@ x_28 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile__Do_ x_29 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_29, 0, x_14); lean_ctor_set(x_29, 1, x_28); -x_30 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; +x_30 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; x_31 = lean_array_push(x_30, x_29); x_32 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile__Do____1___closed__12; x_33 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_33, 0, x_21); lean_ctor_set(x_33, 1, x_32); lean_ctor_set(x_33, 2, x_31); -x_34 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; +x_34 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; x_35 = lean_array_push(x_34, x_33); -x_36 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; +x_36 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; x_37 = lean_array_push(x_35, x_36); x_38 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile__Do____1___closed__5; x_39 = lean_alloc_ctor(1, 3, 0); @@ -15774,7 +16172,7 @@ lean_ctor_set(x_39, 0, x_21); lean_ctor_set(x_39, 1, x_38); lean_ctor_set(x_39, 2, x_37); x_40 = lean_array_push(x_30, x_39); -x_41 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; +x_41 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; x_42 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_42, 0, x_21); lean_ctor_set(x_42, 1, x_41); @@ -15791,7 +16189,7 @@ x_48 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_48, 0, x_21); lean_ctor_set(x_48, 1, x_41); lean_ctor_set(x_48, 2, x_47); -x_49 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__21; +x_49 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__28; x_50 = lean_array_push(x_49, x_18); x_51 = lean_array_push(x_50, x_23); x_52 = lean_array_push(x_51, x_25); @@ -15869,16 +16267,16 @@ x_84 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile__Do_ x_85 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_85, 0, x_69); lean_ctor_set(x_85, 1, x_84); -x_86 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; +x_86 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; x_87 = lean_array_push(x_86, x_85); x_88 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile__Do____1___closed__12; x_89 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_89, 0, x_77); lean_ctor_set(x_89, 1, x_88); lean_ctor_set(x_89, 2, x_87); -x_90 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; +x_90 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; x_91 = lean_array_push(x_90, x_89); -x_92 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; +x_92 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; x_93 = lean_array_push(x_91, x_92); x_94 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile__Do____1___closed__5; x_95 = lean_alloc_ctor(1, 3, 0); @@ -15886,7 +16284,7 @@ lean_ctor_set(x_95, 0, x_77); lean_ctor_set(x_95, 1, x_94); lean_ctor_set(x_95, 2, x_93); x_96 = lean_array_push(x_86, x_95); -x_97 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; +x_97 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; x_98 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_98, 0, x_77); lean_ctor_set(x_98, 1, x_97); @@ -15903,7 +16301,7 @@ x_104 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_104, 0, x_77); lean_ctor_set(x_104, 1, x_97); lean_ctor_set(x_104, 2, x_103); -x_105 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__21; +x_105 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__28; x_106 = lean_array_push(x_105, x_74); x_107 = lean_array_push(x_106, x_79); x_108 = lean_array_push(x_107, x_81); @@ -15958,7 +16356,7 @@ static lean_object* _init_l_Lean_doElemRepeat__Until_____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__2; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__2; x_2 = l_Lean_doElemRepeat__Until_____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -15986,7 +16384,7 @@ static lean_object* _init_l_Lean_doElemRepeat__Until_____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_Lean_doElemRepeat_____closed__8; x_3 = l_Lean_doElemRepeat__Until_____closed__4; x_4 = lean_alloc_ctor(2, 3, 0); @@ -16000,9 +16398,9 @@ static lean_object* _init_l_Lean_doElemRepeat__Until_____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_Lean_doElemRepeat__Until_____closed__5; -x_3 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__20; +x_3 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__20; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -16044,7 +16442,7 @@ static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lea _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat__Until____1___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -16054,8 +16452,8 @@ static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lea _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -16102,7 +16500,7 @@ lean_inc(x_14); x_18 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_18, 0, x_14); lean_ctor_set(x_18, 1, x_17); -x_19 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; +x_19 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; x_20 = lean_array_push(x_19, x_18); x_21 = lean_array_push(x_20, x_9); x_22 = lean_box(2); @@ -16116,9 +16514,9 @@ lean_inc(x_14); x_26 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_26, 0, x_14); lean_ctor_set(x_26, 1, x_25); -x_27 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; +x_27 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; x_28 = lean_array_push(x_27, x_26); -x_29 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; +x_29 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; x_30 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_30, 0, x_22); lean_ctor_set(x_30, 1, x_29); @@ -16158,7 +16556,7 @@ lean_ctor_set(x_47, 0, x_22); lean_ctor_set(x_47, 1, x_46); lean_ctor_set(x_47, 2, x_45); x_48 = lean_array_push(x_19, x_47); -x_49 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; +x_49 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; x_50 = lean_array_push(x_48, x_49); x_51 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_51, 0, x_22); @@ -16175,7 +16573,7 @@ x_56 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_56, 0, x_22); lean_ctor_set(x_56, 1, x_55); lean_ctor_set(x_56, 2, x_54); -x_57 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__21; +x_57 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__28; x_58 = lean_array_push(x_57, x_36); x_59 = lean_array_push(x_58, x_40); x_60 = lean_array_push(x_59, x_42); @@ -16232,7 +16630,7 @@ lean_inc(x_78); x_83 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_83, 0, x_78); lean_ctor_set(x_83, 1, x_82); -x_84 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; +x_84 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; x_85 = lean_array_push(x_84, x_83); x_86 = lean_array_push(x_85, x_9); x_87 = lean_box(2); @@ -16246,9 +16644,9 @@ lean_inc(x_78); x_91 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_91, 0, x_78); lean_ctor_set(x_91, 1, x_90); -x_92 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; +x_92 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; x_93 = lean_array_push(x_92, x_91); -x_94 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; +x_94 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; x_95 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_95, 0, x_87); lean_ctor_set(x_95, 1, x_94); @@ -16288,7 +16686,7 @@ lean_ctor_set(x_112, 0, x_87); lean_ctor_set(x_112, 1, x_111); lean_ctor_set(x_112, 2, x_110); x_113 = lean_array_push(x_84, x_112); -x_114 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; +x_114 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; x_115 = lean_array_push(x_113, x_114); x_116 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_116, 0, x_87); @@ -16305,7 +16703,7 @@ x_121 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_121, 0, x_87); lean_ctor_set(x_121, 1, x_120); lean_ctor_set(x_121, 2, x_119); -x_122 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__21; +x_122 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__28; x_123 = lean_array_push(x_122, x_101); x_124 = lean_array_push(x_123, x_105); x_125 = lean_array_push(x_124, x_107); @@ -16361,7 +16759,7 @@ static lean_object* _init_l_Lean_term__Matches___x7c___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__2; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__2; x_2 = l_Lean_term__Matches___x7c___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -16389,7 +16787,7 @@ static lean_object* _init_l_Lean_term__Matches___x7c___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__19; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__19; x_2 = lean_unsigned_to_nat(51u); x_3 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -16417,7 +16815,7 @@ static lean_object* _init_l_Lean_term__Matches___x7c___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; x_2 = l_Lean_term__Matches___x7c___closed__4; x_3 = l_Lean_term__Matches___x7c___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); @@ -16497,7 +16895,7 @@ static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lea _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -16515,7 +16913,7 @@ static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lea _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -16541,7 +16939,7 @@ static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lea _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__6; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -16559,7 +16957,7 @@ static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lea _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__8; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -16832,7 +17230,7 @@ lean_inc(x_16); x_17 = lean_ctor_get(x_2, 1); lean_inc(x_17); lean_dec(x_2); -x_18 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__22; +x_18 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__22; lean_inc(x_15); x_19 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_19, 0, x_15); @@ -16850,9 +17248,9 @@ x_26 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_26, 0, x_24); lean_ctor_set(x_26, 1, x_25); lean_ctor_set(x_26, 2, x_23); -x_27 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; +x_27 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; x_28 = lean_array_push(x_27, x_26); -x_29 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; +x_29 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; x_30 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_30, 0, x_24); lean_ctor_set(x_30, 1, x_29); @@ -16867,7 +17265,7 @@ lean_inc(x_15); x_34 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_34, 0, x_15); lean_ctor_set(x_34, 1, x_33); -x_35 = l_Lean_Syntax_SepArray_getElems___rarg(x_12); +x_35 = l_Lean_Syntax_TSepArray_getElems___rarg(x_12); lean_dec(x_12); x_36 = lean_array_get_size(x_35); x_37 = lean_usize_of_nat(x_36); @@ -16877,13 +17275,13 @@ x_39 = l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRul x_40 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__10; x_41 = l_Lean_mkSepArray(x_39, x_40); lean_dec(x_39); -x_42 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__27; +x_42 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__27; x_43 = l_Array_append___rarg(x_42, x_41); x_44 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_44, 0, x_24); lean_ctor_set(x_44, 1, x_29); lean_ctor_set(x_44, 2, x_43); -x_45 = l_Lean_expandExplicitBindersAux_loop___closed__7; +x_45 = l_Lean_expandExplicitBindersAux_loop___closed__8; lean_inc(x_15); x_46 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_46, 0, x_15); @@ -16900,7 +17298,7 @@ lean_ctor_set(x_51, 0, x_15); lean_ctor_set(x_51, 1, x_49); lean_ctor_set(x_51, 2, x_48); lean_ctor_set(x_51, 3, x_50); -x_52 = l_unexpandExists___closed__5; +x_52 = l_unexpandExists___closed__7; x_53 = lean_array_push(x_52, x_34); lean_inc(x_53); x_54 = lean_array_push(x_53, x_44); @@ -16912,7 +17310,7 @@ x_58 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_58, 0, x_24); lean_ctor_set(x_58, 1, x_57); lean_ctor_set(x_58, 2, x_56); -x_59 = l_Lean_binderIdent___closed__5; +x_59 = l_Lean_expandExplicitBindersAux_loop___closed__7; lean_inc(x_15); x_60 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_60, 0, x_15); @@ -16952,7 +17350,7 @@ x_76 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_76, 0, x_24); lean_ctor_set(x_76, 1, x_57); lean_ctor_set(x_76, 2, x_75); -x_77 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; +x_77 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; x_78 = lean_array_push(x_77, x_58); x_79 = lean_array_push(x_78, x_76); x_80 = lean_alloc_ctor(1, 3, 0); @@ -16965,9 +17363,9 @@ x_83 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_83, 0, x_24); lean_ctor_set(x_83, 1, x_82); lean_ctor_set(x_83, 2, x_81); -x_84 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__21; +x_84 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__28; x_85 = lean_array_push(x_84, x_21); -x_86 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; +x_86 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; x_87 = lean_array_push(x_85, x_86); x_88 = lean_array_push(x_87, x_86); x_89 = lean_array_push(x_88, x_30); @@ -16984,23 +17382,23 @@ x_96 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_96, 0, x_24); lean_ctor_set(x_96, 1, x_29); lean_ctor_set(x_96, 2, x_95); -x_97 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__29; +x_97 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__29; lean_inc(x_15); x_98 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_98, 0, x_15); lean_ctor_set(x_98, 1, x_97); -x_99 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; +x_99 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; x_100 = lean_array_push(x_99, x_19); lean_inc(x_100); x_101 = lean_array_push(x_100, x_96); lean_inc(x_98); x_102 = lean_array_push(x_101, x_98); -x_103 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__21; +x_103 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__21; x_104 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_104, 0, x_24); lean_ctor_set(x_104, 1, x_103); lean_ctor_set(x_104, 2, x_102); -x_105 = l_Lean_expandExplicitBindersAux_loop___closed__12; +x_105 = l_Lean_expandExplicitBindersAux_loop___closed__13; lean_inc(x_15); x_106 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_106, 0, x_15); @@ -17054,7 +17452,7 @@ lean_inc(x_126); x_127 = lean_ctor_get(x_2, 1); lean_inc(x_127); lean_dec(x_2); -x_128 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__22; +x_128 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__22; lean_inc(x_124); x_129 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_129, 0, x_124); @@ -17072,9 +17470,9 @@ x_136 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_136, 0, x_134); lean_ctor_set(x_136, 1, x_135); lean_ctor_set(x_136, 2, x_133); -x_137 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35; +x_137 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35; x_138 = lean_array_push(x_137, x_136); -x_139 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19; +x_139 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; x_140 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_140, 0, x_134); lean_ctor_set(x_140, 1, x_139); @@ -17089,7 +17487,7 @@ lean_inc(x_124); x_144 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_144, 0, x_124); lean_ctor_set(x_144, 1, x_143); -x_145 = l_Lean_Syntax_SepArray_getElems___rarg(x_12); +x_145 = l_Lean_Syntax_TSepArray_getElems___rarg(x_12); lean_dec(x_12); x_146 = lean_array_get_size(x_145); x_147 = lean_usize_of_nat(x_146); @@ -17099,13 +17497,13 @@ x_149 = l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRu x_150 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__10; x_151 = l_Lean_mkSepArray(x_149, x_150); lean_dec(x_149); -x_152 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__27; +x_152 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__27; x_153 = l_Array_append___rarg(x_152, x_151); x_154 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_154, 0, x_134); lean_ctor_set(x_154, 1, x_139); lean_ctor_set(x_154, 2, x_153); -x_155 = l_Lean_expandExplicitBindersAux_loop___closed__7; +x_155 = l_Lean_expandExplicitBindersAux_loop___closed__8; lean_inc(x_124); x_156 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_156, 0, x_124); @@ -17122,7 +17520,7 @@ lean_ctor_set(x_161, 0, x_124); lean_ctor_set(x_161, 1, x_159); lean_ctor_set(x_161, 2, x_158); lean_ctor_set(x_161, 3, x_160); -x_162 = l_unexpandExists___closed__5; +x_162 = l_unexpandExists___closed__7; x_163 = lean_array_push(x_162, x_144); lean_inc(x_163); x_164 = lean_array_push(x_163, x_154); @@ -17134,7 +17532,7 @@ x_168 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_168, 0, x_134); lean_ctor_set(x_168, 1, x_167); lean_ctor_set(x_168, 2, x_166); -x_169 = l_Lean_binderIdent___closed__5; +x_169 = l_Lean_expandExplicitBindersAux_loop___closed__7; lean_inc(x_124); x_170 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_170, 0, x_124); @@ -17174,7 +17572,7 @@ x_186 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_186, 0, x_134); lean_ctor_set(x_186, 1, x_167); lean_ctor_set(x_186, 2, x_185); -x_187 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26; +x_187 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; x_188 = lean_array_push(x_187, x_168); x_189 = lean_array_push(x_188, x_186); x_190 = lean_alloc_ctor(1, 3, 0); @@ -17187,9 +17585,9 @@ x_193 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_193, 0, x_134); lean_ctor_set(x_193, 1, x_192); lean_ctor_set(x_193, 2, x_191); -x_194 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__21; +x_194 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__28; x_195 = lean_array_push(x_194, x_131); -x_196 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28; +x_196 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; x_197 = lean_array_push(x_195, x_196); x_198 = lean_array_push(x_197, x_196); x_199 = lean_array_push(x_198, x_140); @@ -17206,23 +17604,23 @@ x_206 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_206, 0, x_134); lean_ctor_set(x_206, 1, x_139); lean_ctor_set(x_206, 2, x_205); -x_207 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__29; +x_207 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__29; lean_inc(x_124); x_208 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_208, 0, x_124); lean_ctor_set(x_208, 1, x_207); -x_209 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30; +x_209 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30; x_210 = lean_array_push(x_209, x_129); lean_inc(x_210); x_211 = lean_array_push(x_210, x_206); lean_inc(x_208); x_212 = lean_array_push(x_211, x_208); -x_213 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__21; +x_213 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__21; x_214 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_214, 0, x_134); lean_ctor_set(x_214, 1, x_213); lean_ctor_set(x_214, 2, x_212); -x_215 = l_Lean_expandExplicitBindersAux_loop___closed__12; +x_215 = l_Lean_expandExplicitBindersAux_loop___closed__13; lean_inc(x_124); x_216 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_216, 0, x_124); @@ -17297,142 +17695,124 @@ lean_dec_ref(res); res = initialize_Init_Data_ToString(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__1 = _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__1(); -lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__1); -l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__2 = _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__2(); -lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__2); -l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__3 = _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__3(); -lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__3); -l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__4 = _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__4(); -lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__4); -l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__5 = _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__5(); -lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__5); -l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6 = _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6(); -lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__6); -l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__7 = _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__7(); -lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__7); -l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__8 = _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__8(); -lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__8); -l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__9 = _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__9(); -lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__9); -l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__10 = _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__10(); -lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__10); -l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__11 = _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__11(); -lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__11); -l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__12 = _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__12(); -lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__12); -l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__13 = _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__13(); -lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__13); -l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__14 = _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__14(); -lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__14); -l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__15 = _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__15(); -lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__15); -l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__16 = _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__16(); -lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__16); -l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__17 = _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__17(); -lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__17); -l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__18 = _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__18(); -lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__18); -l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__19 = _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__19(); -lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__19); -l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__20 = _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__20(); -lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__20); -l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__21 = _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__21(); -lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__21); -l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__22 = _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__22(); -lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__22); -l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__23 = _init_l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__23(); -lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b_____x5d_____closed__23); -l_Lean_termMacro_x2etrace_x5b_____x5d__ = _init_l_Lean_termMacro_x2etrace_x5b_____x5d__(); -lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b_____x5d__); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__1 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__1(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__1); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__2 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__2(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__2); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__3 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__3(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__3); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__4); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__5 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__5(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__5); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__6 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__6(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__6); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__7 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__7(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__7); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__8 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__8(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__8); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__9 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__9(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__9); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__10 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__10(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__10); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__11 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__11(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__11); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__12 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__12(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__12); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__13 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__13(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__13); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__14 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__14(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__14); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__15 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__15(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__15); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__16 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__16(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__16); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__17 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__17(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__17); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__18 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__18(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__18); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__19); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__20 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__20(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__20); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__21 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__21(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__21); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__22 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__22(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__22); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__23 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__23(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__23); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__24 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__24(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__24); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__25 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__25(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__25); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__26); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__27 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__27(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__27); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__28); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__29 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__29(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__29); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__30); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__31 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__31(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__31); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__32 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__32(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__32); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__33 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__33(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__33); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__34 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__34(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__34); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b_____x5d____1___closed__35); -l_Lean_binderIdent___closed__1 = _init_l_Lean_binderIdent___closed__1(); -lean_mark_persistent(l_Lean_binderIdent___closed__1); -l_Lean_binderIdent___closed__2 = _init_l_Lean_binderIdent___closed__2(); -lean_mark_persistent(l_Lean_binderIdent___closed__2); -l_Lean_binderIdent___closed__3 = _init_l_Lean_binderIdent___closed__3(); -lean_mark_persistent(l_Lean_binderIdent___closed__3); -l_Lean_binderIdent___closed__4 = _init_l_Lean_binderIdent___closed__4(); -lean_mark_persistent(l_Lean_binderIdent___closed__4); -l_Lean_binderIdent___closed__5 = _init_l_Lean_binderIdent___closed__5(); -lean_mark_persistent(l_Lean_binderIdent___closed__5); -l_Lean_binderIdent___closed__6 = _init_l_Lean_binderIdent___closed__6(); -lean_mark_persistent(l_Lean_binderIdent___closed__6); -l_Lean_binderIdent___closed__7 = _init_l_Lean_binderIdent___closed__7(); -lean_mark_persistent(l_Lean_binderIdent___closed__7); -l_Lean_binderIdent___closed__8 = _init_l_Lean_binderIdent___closed__8(); -lean_mark_persistent(l_Lean_binderIdent___closed__8); -l_Lean_binderIdent = _init_l_Lean_binderIdent(); -lean_mark_persistent(l_Lean_binderIdent); +l_Lean_termMacro_x2etrace_x5b___x5d_____closed__1 = _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__1(); +lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b___x5d_____closed__1); +l_Lean_termMacro_x2etrace_x5b___x5d_____closed__2 = _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__2(); +lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b___x5d_____closed__2); +l_Lean_termMacro_x2etrace_x5b___x5d_____closed__3 = _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__3(); +lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b___x5d_____closed__3); +l_Lean_termMacro_x2etrace_x5b___x5d_____closed__4 = _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__4(); +lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b___x5d_____closed__4); +l_Lean_termMacro_x2etrace_x5b___x5d_____closed__5 = _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__5(); +lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b___x5d_____closed__5); +l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6 = _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6(); +lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6); +l_Lean_termMacro_x2etrace_x5b___x5d_____closed__7 = _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__7(); +lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b___x5d_____closed__7); +l_Lean_termMacro_x2etrace_x5b___x5d_____closed__8 = _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__8(); +lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b___x5d_____closed__8); +l_Lean_termMacro_x2etrace_x5b___x5d_____closed__9 = _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__9(); +lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b___x5d_____closed__9); +l_Lean_termMacro_x2etrace_x5b___x5d_____closed__10 = _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__10(); +lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b___x5d_____closed__10); +l_Lean_termMacro_x2etrace_x5b___x5d_____closed__11 = _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__11(); +lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b___x5d_____closed__11); +l_Lean_termMacro_x2etrace_x5b___x5d_____closed__12 = _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__12(); +lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b___x5d_____closed__12); +l_Lean_termMacro_x2etrace_x5b___x5d_____closed__13 = _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__13(); +lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b___x5d_____closed__13); +l_Lean_termMacro_x2etrace_x5b___x5d_____closed__14 = _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__14(); +lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b___x5d_____closed__14); +l_Lean_termMacro_x2etrace_x5b___x5d_____closed__15 = _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__15(); +lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b___x5d_____closed__15); +l_Lean_termMacro_x2etrace_x5b___x5d_____closed__16 = _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__16(); +lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b___x5d_____closed__16); +l_Lean_termMacro_x2etrace_x5b___x5d_____closed__17 = _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__17(); +lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b___x5d_____closed__17); +l_Lean_termMacro_x2etrace_x5b___x5d_____closed__18 = _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__18(); +lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b___x5d_____closed__18); +l_Lean_termMacro_x2etrace_x5b___x5d_____closed__19 = _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__19(); +lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b___x5d_____closed__19); +l_Lean_termMacro_x2etrace_x5b___x5d_____closed__20 = _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__20(); +lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b___x5d_____closed__20); +l_Lean_termMacro_x2etrace_x5b___x5d_____closed__21 = _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__21(); +lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b___x5d_____closed__21); +l_Lean_termMacro_x2etrace_x5b___x5d_____closed__22 = _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__22(); +lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b___x5d_____closed__22); +l_Lean_termMacro_x2etrace_x5b___x5d_____closed__23 = _init_l_Lean_termMacro_x2etrace_x5b___x5d_____closed__23(); +lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b___x5d_____closed__23); +l_Lean_termMacro_x2etrace_x5b___x5d__ = _init_l_Lean_termMacro_x2etrace_x5b___x5d__(); +lean_mark_persistent(l_Lean_termMacro_x2etrace_x5b___x5d__); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__1 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__1(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__1); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__2 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__2(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__2); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__3 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__3(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__3); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__4); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__5 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__5(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__5); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__6 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__6(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__6); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__7 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__7(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__7); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__8 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__8(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__8); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__9 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__9(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__9); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__10 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__10(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__10); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__11 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__11(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__11); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__12 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__12(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__12); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__13 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__13(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__13); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__14 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__14(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__14); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__15 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__15(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__15); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__16 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__16(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__16); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__17 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__17(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__17); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__18 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__18(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__18); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__20 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__20(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__20); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__21 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__21(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__21); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__22 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__22(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__22); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__23 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__23(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__23); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__24 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__24(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__24); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__25 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__25(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__25); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__27 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__27(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__27); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__29 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__29(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__29); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__30); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__31 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__31(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__31); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__32 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__32(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__32); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__33 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__33(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__33); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__34 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__34(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__34); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__35); l_Lean_unbracketedExplicitBinders___closed__1 = _init_l_Lean_unbracketedExplicitBinders___closed__1(); lean_mark_persistent(l_Lean_unbracketedExplicitBinders___closed__1); l_Lean_unbracketedExplicitBinders___closed__2 = _init_l_Lean_unbracketedExplicitBinders___closed__2(); @@ -17491,6 +17871,10 @@ l_Lean_explicitBinders___closed__4 = _init_l_Lean_explicitBinders___closed__4(); lean_mark_persistent(l_Lean_explicitBinders___closed__4); l_Lean_explicitBinders___closed__5 = _init_l_Lean_explicitBinders___closed__5(); lean_mark_persistent(l_Lean_explicitBinders___closed__5); +l_Lean_explicitBinders___closed__6 = _init_l_Lean_explicitBinders___closed__6(); +lean_mark_persistent(l_Lean_explicitBinders___closed__6); +l_Lean_explicitBinders___closed__7 = _init_l_Lean_explicitBinders___closed__7(); +lean_mark_persistent(l_Lean_explicitBinders___closed__7); l_Lean_explicitBinders = _init_l_Lean_explicitBinders(); lean_mark_persistent(l_Lean_explicitBinders); l_Lean_expandExplicitBindersAux_loop___closed__1 = _init_l_Lean_expandExplicitBindersAux_loop___closed__1(); @@ -17517,6 +17901,8 @@ l_Lean_expandExplicitBindersAux_loop___closed__11 = _init_l_Lean_expandExplicitB lean_mark_persistent(l_Lean_expandExplicitBindersAux_loop___closed__11); l_Lean_expandExplicitBindersAux_loop___closed__12 = _init_l_Lean_expandExplicitBindersAux_loop___closed__12(); lean_mark_persistent(l_Lean_expandExplicitBindersAux_loop___closed__12); +l_Lean_expandExplicitBindersAux_loop___closed__13 = _init_l_Lean_expandExplicitBindersAux_loop___closed__13(); +lean_mark_persistent(l_Lean_expandExplicitBindersAux_loop___closed__13); l_Lean_expandExplicitBinders___closed__1 = _init_l_Lean_expandExplicitBinders___closed__1(); lean_mark_persistent(l_Lean_expandExplicitBinders___closed__1); l_Lean_unifConstraint___closed__1 = _init_l_Lean_unifConstraint___closed__1(); @@ -17565,170 +17951,170 @@ l_Lean_unifConstraintElem___closed__11 = _init_l_Lean_unifConstraintElem___close lean_mark_persistent(l_Lean_unifConstraintElem___closed__11); l_Lean_unifConstraintElem = _init_l_Lean_unifConstraintElem(); lean_mark_persistent(l_Lean_unifConstraintElem); -l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__1 = _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__1(); -lean_mark_persistent(l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__1); -l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__2 = _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__2(); -lean_mark_persistent(l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__2); -l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__3 = _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__3(); -lean_mark_persistent(l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__3); -l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__4 = _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__4(); -lean_mark_persistent(l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__4); -l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__5 = _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__5(); -lean_mark_persistent(l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__5); -l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__6 = _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__6(); -lean_mark_persistent(l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__6); -l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__7 = _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__7(); -lean_mark_persistent(l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__7); -l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__8 = _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__8(); -lean_mark_persistent(l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__8); -l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__9 = _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__9(); -lean_mark_persistent(l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__9); -l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__10 = _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__10(); -lean_mark_persistent(l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__10); -l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__11 = _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__11(); -lean_mark_persistent(l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__11); -l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__12 = _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__12(); -lean_mark_persistent(l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__12); -l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__13 = _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__13(); -lean_mark_persistent(l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__13); -l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__14 = _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__14(); -lean_mark_persistent(l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__14); -l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__15 = _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__15(); -lean_mark_persistent(l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__15); -l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__16 = _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__16(); -lean_mark_persistent(l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__16); -l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__17 = _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__17(); -lean_mark_persistent(l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__17); -l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__18 = _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__18(); -lean_mark_persistent(l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__18); -l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__19 = _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__19(); -lean_mark_persistent(l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__19); -l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__20 = _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__20(); -lean_mark_persistent(l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__20); -l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__21 = _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__21(); -lean_mark_persistent(l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__21); -l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__22 = _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__22(); -lean_mark_persistent(l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__22); -l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__23 = _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__23(); -lean_mark_persistent(l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__23); -l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__24 = _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__24(); -lean_mark_persistent(l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__24); -l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__25 = _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__25(); -lean_mark_persistent(l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__25); -l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__26 = _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__26(); -lean_mark_persistent(l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__26); -l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__27 = _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__27(); -lean_mark_persistent(l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__27); -l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__28 = _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__28(); -lean_mark_persistent(l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__28); -l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__29 = _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__29(); -lean_mark_persistent(l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__29); -l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__30 = _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__30(); -lean_mark_persistent(l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__30); -l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__31 = _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__31(); -lean_mark_persistent(l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__31); -l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__32 = _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__32(); -lean_mark_persistent(l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__32); -l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__33 = _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__33(); -lean_mark_persistent(l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2_____closed__33); -l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2__ = _init_l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2__(); -lean_mark_persistent(l_Lean_command__Unif__hint______Where___x7c_x2d_u22a2__); -l_Array_forInUnsafe_loop___at___private_Init_NotationExtra_0__Lean_mkHintBody___spec__1___closed__1 = _init_l_Array_forInUnsafe_loop___at___private_Init_NotationExtra_0__Lean_mkHintBody___spec__1___closed__1(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Init_NotationExtra_0__Lean_mkHintBody___spec__1___closed__1); -l_Array_forInUnsafe_loop___at___private_Init_NotationExtra_0__Lean_mkHintBody___spec__1___closed__2 = _init_l_Array_forInUnsafe_loop___at___private_Init_NotationExtra_0__Lean_mkHintBody___spec__1___closed__2(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Init_NotationExtra_0__Lean_mkHintBody___spec__1___closed__2); -l_Array_forInUnsafe_loop___at___private_Init_NotationExtra_0__Lean_mkHintBody___spec__1___closed__3 = _init_l_Array_forInUnsafe_loop___at___private_Init_NotationExtra_0__Lean_mkHintBody___spec__1___closed__3(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Init_NotationExtra_0__Lean_mkHintBody___spec__1___closed__3); -l_Array_forInUnsafe_loop___at___private_Init_NotationExtra_0__Lean_mkHintBody___spec__1___closed__4 = _init_l_Array_forInUnsafe_loop___at___private_Init_NotationExtra_0__Lean_mkHintBody___spec__1___closed__4(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Init_NotationExtra_0__Lean_mkHintBody___spec__1___closed__4); -l___private_Init_NotationExtra_0__Lean_mkHintBody___closed__1 = _init_l___private_Init_NotationExtra_0__Lean_mkHintBody___closed__1(); -lean_mark_persistent(l___private_Init_NotationExtra_0__Lean_mkHintBody___closed__1); -l___private_Init_NotationExtra_0__Lean_mkHintBody___closed__2 = _init_l___private_Init_NotationExtra_0__Lean_mkHintBody___closed__2(); -lean_mark_persistent(l___private_Init_NotationExtra_0__Lean_mkHintBody___closed__2); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__1 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__1(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__1); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__2 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__2(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__2); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__3 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__3(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__3); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__4 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__4(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__4); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__5 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__5(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__5); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__6 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__6(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__6); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__7 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__7(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__7); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__8 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__8(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__8); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__9 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__9(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__9); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__10 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__10(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__10); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__11 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__11(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__11); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__12 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__12(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__12); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__13 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__13(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__13); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__14 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__14(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__14); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__15 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__15(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__15); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__16 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__16(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__16); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__17 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__17(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__17); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__18 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__18(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__18); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__19 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__19(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__19); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__20 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__20(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__20); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__21 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__21(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__21); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__22 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__22(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__22); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__23 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__23(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__23); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__24 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__24(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__24); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__25 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__25(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__25); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__26 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__26(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__26); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__27 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__27(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__27); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__28 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__28(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__28); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__29 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__29(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__29); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__30 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__30(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__30); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__31 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__31(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__31); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__32 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__32(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__32); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__33 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__33(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__33); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__34 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__34(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__34); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__35 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__35(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__35); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__36 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__36(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__36); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__37 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__37(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__37); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__38 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__38(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__38); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__39 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__39(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__39); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__40 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__40(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__40); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__41 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__41(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__41); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__42 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__42(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint______Where___x7c_x2d_u22a2____1___closed__42); +l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__1 = _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__1(); +lean_mark_persistent(l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__1); +l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__2 = _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__2(); +lean_mark_persistent(l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__2); +l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__3 = _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__3(); +lean_mark_persistent(l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__3); +l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__4 = _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__4(); +lean_mark_persistent(l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__4); +l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__5 = _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__5(); +lean_mark_persistent(l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__5); +l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__6 = _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__6(); +lean_mark_persistent(l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__6); +l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__7 = _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__7(); +lean_mark_persistent(l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__7); +l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__8 = _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__8(); +lean_mark_persistent(l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__8); +l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__9 = _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__9(); +lean_mark_persistent(l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__9); +l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__10 = _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__10(); +lean_mark_persistent(l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__10); +l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__11 = _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__11(); +lean_mark_persistent(l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__11); +l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__12 = _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__12(); +lean_mark_persistent(l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__12); +l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__13 = _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__13(); +lean_mark_persistent(l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__13); +l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__14 = _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__14(); +lean_mark_persistent(l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__14); +l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__15 = _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__15(); +lean_mark_persistent(l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__15); +l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__16 = _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__16(); +lean_mark_persistent(l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__16); +l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__17 = _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__17(); +lean_mark_persistent(l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__17); +l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__18 = _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__18(); +lean_mark_persistent(l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__18); +l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__19 = _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__19(); +lean_mark_persistent(l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__19); +l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__20 = _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__20(); +lean_mark_persistent(l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__20); +l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__21 = _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__21(); +lean_mark_persistent(l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__21); +l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__22 = _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__22(); +lean_mark_persistent(l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__22); +l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__23 = _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__23(); +lean_mark_persistent(l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__23); +l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__24 = _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__24(); +lean_mark_persistent(l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__24); +l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__25 = _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__25(); +lean_mark_persistent(l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__25); +l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__26 = _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__26(); +lean_mark_persistent(l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__26); +l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__27 = _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__27(); +lean_mark_persistent(l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__27); +l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__28 = _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__28(); +lean_mark_persistent(l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__28); +l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__29 = _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__29(); +lean_mark_persistent(l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__29); +l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__30 = _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__30(); +lean_mark_persistent(l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__30); +l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__31 = _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__31(); +lean_mark_persistent(l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__31); +l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__32 = _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__32(); +lean_mark_persistent(l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__32); +l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__33 = _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__33(); +lean_mark_persistent(l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2_____closed__33); +l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2__ = _init_l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2__(); +lean_mark_persistent(l_Lean_command__Unif__hint____Where___x7c_x2d_u22a2__); +l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__5___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__5___closed__1(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__5___closed__1); +l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__5___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__5___closed__2(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__5___closed__2); +l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__5___closed__3 = _init_l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__5___closed__3(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___spec__5___closed__3); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__1 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__1(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__1); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__2 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__2(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__2); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__3 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__3(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__3); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__4 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__4(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__4); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__5 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__5(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__5); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__6 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__6(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__6); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__7 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__7(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__7); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__8 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__8(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__8); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__9 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__9(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__9); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__10 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__10(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__10); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__11 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__11(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__11); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__12 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__12(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__12); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__13 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__13(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__13); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__14 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__14(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__14); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__15 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__15(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__15); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__16 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__16(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__16); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__17 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__17(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__17); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__18 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__18(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__18); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__19 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__19(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__19); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__20 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__20(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__20); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__21 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__21(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__21); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__22 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__22(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__22); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__23 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__23(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__23); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__24 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__24(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__24); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__25 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__25(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__25); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__26 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__26(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__26); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__27 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__27(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__27); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__28 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__28(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__28); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__29 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__29(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__29); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__30 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__30(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__30); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__31 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__31(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__31); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__32 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__32(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__32); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__33 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__33(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__33); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__34 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__34(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__34); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__35 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__35(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__35); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__36 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__36(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__36); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__37 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__37(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__37); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__38 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__38(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__38); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__39 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__39(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__39); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__40 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__40(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__40); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__41 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__41(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__41); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__42 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__42(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__42); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__43 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__43(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__43); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__44 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__44(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__44); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__45 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__45(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command__Unif__hint____Where___x7c_x2d_u22a2____1___closed__45); l_term_u2203___x2c_____closed__1 = _init_l_term_u2203___x2c_____closed__1(); lean_mark_persistent(l_term_u2203___x2c_____closed__1); l_term_u2203___x2c_____closed__2 = _init_l_term_u2203___x2c_____closed__2(); @@ -17913,8 +18299,6 @@ l_tacticCalc_____closed__6 = _init_l_tacticCalc_____closed__6(); lean_mark_persistent(l_tacticCalc_____closed__6); l_tacticCalc_____closed__7 = _init_l_tacticCalc_____closed__7(); lean_mark_persistent(l_tacticCalc_____closed__7); -l_tacticCalc_____closed__8 = _init_l_tacticCalc_____closed__8(); -lean_mark_persistent(l_tacticCalc_____closed__8); l_tacticCalc__ = _init_l_tacticCalc__(); lean_mark_persistent(l_tacticCalc__); l___aux__Init__NotationExtra______macroRules__tacticCalc____1___closed__1 = _init_l___aux__Init__NotationExtra______macroRules__tacticCalc____1___closed__1(); @@ -17977,6 +18361,10 @@ l_unexpandExists___closed__4 = _init_l_unexpandExists___closed__4(); lean_mark_persistent(l_unexpandExists___closed__4); l_unexpandExists___closed__5 = _init_l_unexpandExists___closed__5(); lean_mark_persistent(l_unexpandExists___closed__5); +l_unexpandExists___closed__6 = _init_l_unexpandExists___closed__6(); +lean_mark_persistent(l_unexpandExists___closed__6); +l_unexpandExists___closed__7 = _init_l_unexpandExists___closed__7(); +lean_mark_persistent(l_unexpandExists___closed__7); l_unexpandSigma___closed__1 = _init_l_unexpandSigma___closed__1(); lean_mark_persistent(l_unexpandSigma___closed__1); l_unexpandPSigma___closed__1 = _init_l_unexpandPSigma___closed__1(); @@ -18039,24 +18427,24 @@ l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__10 = lean_mark_persistent(l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__10); l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__11 = _init_l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__11(); lean_mark_persistent(l___aux__Init__NotationExtra______macroRules__tacticFunext______1___closed__11); -l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__1 = _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__1(); -lean_mark_persistent(l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__1); -l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__2 = _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__2(); -lean_mark_persistent(l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__2); -l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__3 = _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__3(); -lean_mark_persistent(l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__3); -l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__4 = _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__4(); -lean_mark_persistent(l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__4); -l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__5 = _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__5(); -lean_mark_persistent(l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__5); -l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__6 = _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__6(); -lean_mark_persistent(l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__6); -l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__7 = _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__7(); -lean_mark_persistent(l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__7); -l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__8 = _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__8(); -lean_mark_persistent(l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__8); -l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__9 = _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__9(); -lean_mark_persistent(l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__9); +l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__1 = _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__1(); +lean_mark_persistent(l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__1); +l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__2 = _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__2(); +lean_mark_persistent(l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__2); +l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__3 = _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__3(); +lean_mark_persistent(l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__3); +l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__4 = _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__4(); +lean_mark_persistent(l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__4); +l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__5 = _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__5(); +lean_mark_persistent(l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__5); +l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__6 = _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__6(); +lean_mark_persistent(l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__6); +l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__7 = _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__7(); +lean_mark_persistent(l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__7); +l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__8 = _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__8(); +lean_mark_persistent(l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__8); +l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__9 = _init_l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__9(); +lean_mark_persistent(l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5___closed__9); l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__1 = _init_l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__1(); lean_mark_persistent(l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__1); l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__2 = _init_l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__2(); @@ -18269,10 +18657,8 @@ l_solve___closed__15 = _init_l_solve___closed__15(); lean_mark_persistent(l_solve___closed__15); l_solve = _init_l_solve(); lean_mark_persistent(l_solve); -l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__solve__1___spec__1___closed__1 = _init_l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__solve__1___spec__1___closed__1(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__solve__1___spec__1___closed__1); -l___aux__Init__NotationExtra______macroRules__solve__1___lambda__1___closed__1 = _init_l___aux__Init__NotationExtra______macroRules__solve__1___lambda__1___closed__1(); -lean_mark_persistent(l___aux__Init__NotationExtra______macroRules__solve__1___lambda__1___closed__1); +l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__solve__1___spec__3___closed__1 = _init_l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__solve__1___spec__3___closed__1(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__solve__1___spec__3___closed__1); l___aux__Init__NotationExtra______macroRules__solve__1___closed__1 = _init_l___aux__Init__NotationExtra______macroRules__solve__1___closed__1(); lean_mark_persistent(l___aux__Init__NotationExtra______macroRules__solve__1___closed__1); l___aux__Init__NotationExtra______macroRules__solve__1___closed__2 = _init_l___aux__Init__NotationExtra______macroRules__solve__1___closed__2(); @@ -18287,6 +18673,8 @@ l___aux__Init__NotationExtra______macroRules__solve__1___closed__6 = _init_l___a lean_mark_persistent(l___aux__Init__NotationExtra______macroRules__solve__1___closed__6); l___aux__Init__NotationExtra______macroRules__solve__1___closed__7 = _init_l___aux__Init__NotationExtra______macroRules__solve__1___closed__7(); lean_mark_persistent(l___aux__Init__NotationExtra______macroRules__solve__1___closed__7); +l___aux__Init__NotationExtra______macroRules__solve__1___closed__8 = _init_l___aux__Init__NotationExtra______macroRules__solve__1___closed__8(); +lean_mark_persistent(l___aux__Init__NotationExtra______macroRules__solve__1___closed__8); l_Lean_doElemRepeat_____closed__1 = _init_l_Lean_doElemRepeat_____closed__1(); lean_mark_persistent(l_Lean_doElemRepeat_____closed__1); l_Lean_doElemRepeat_____closed__2 = _init_l_Lean_doElemRepeat_____closed__2(); diff --git a/stage0/stdlib/Init/Prelude.c b/stage0/stdlib/Init/Prelude.c index ffa1f1d3014..5c648fd9fb3 100644 --- a/stage0/stdlib/Init/Prelude.c +++ b/stage0/stdlib/Init/Prelude.c @@ -74,6 +74,7 @@ static lean_object* l_Lean_identKind___closed__1; LEAN_EXPORT lean_object* l_UInt64_decEq___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_name_mk_string(lean_object*, lean_object*); static lean_object* l_Lean_fieldIdxKind___closed__2; +LEAN_EXPORT lean_object* l_Lean_Syntax_getOptional_x3f(lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); LEAN_EXPORT lean_object* l_instHXor(lean_object*); LEAN_EXPORT lean_object* l_dite(lean_object*, lean_object*); @@ -171,6 +172,7 @@ LEAN_EXPORT lean_object* l_ReaderT_instMonadReaderT___rarg___lambda__4(lean_obje LEAN_EXPORT lean_object* l_Lean_instMonadQuotation(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_interpolatedStrKind; LEAN_EXPORT lean_object* l_Monad_seqRight___default___rarg___lambda__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_instInhabitedTSyntax(lean_object*); lean_object* lean_nat_pow(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_instMonadLiftT__1___rarg(lean_object*); static lean_object* l_EStateM_instMonadEStateM___closed__4; @@ -284,6 +286,7 @@ LEAN_EXPORT uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); static lean_object* l_Lean_interpolatedStrLitKind___closed__1; static lean_object* l_Lean_Macro_instInhabitedMethods___closed__1; LEAN_EXPORT lean_object* l_ite(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_instInhabitedTSyntax___boxed(lean_object*); static lean_object* l_Lean_groupKind___closed__2; LEAN_EXPORT lean_object* l_Lean_numLitKind; LEAN_EXPORT lean_object* l_EStateM_run_x27___rarg(lean_object*, lean_object*); @@ -668,6 +671,7 @@ LEAN_EXPORT lean_object* l_instBEq___rarg(lean_object*, lean_object*, lean_objec LEAN_EXPORT lean_object* l_UInt64_toUSize___boxed(lean_object*); static lean_object* l_Char_ofNat___closed__1; LEAN_EXPORT lean_object* l_instHAddPosString(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Option_map___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_EStateM_instInhabitedEStateM___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Function_const(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Char_ofNat___boxed(lean_object*); @@ -678,6 +682,7 @@ lean_object* lean_panic_fn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_instHOr(lean_object*); LEAN_EXPORT lean_object* l_EStateM_instInhabitedResult(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_instLEFin___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Option_map(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_instLENat; LEAN_EXPORT lean_object* l_Array_getOp___rarg(lean_object*, lean_object*, lean_object*); uint8_t lean_uint8_of_nat(lean_object*); @@ -735,6 +740,7 @@ LEAN_EXPORT uint8_t l_instDecidableEqBool(uint8_t, uint8_t); static lean_object* l___private_Init_Prelude_0__Lean_assembleParts___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_getOp___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Macro_instInhabitedMethods; +LEAN_EXPORT lean_object* l_Lean_Syntax_getOptional_x3f___boxed(lean_object*); LEAN_EXPORT lean_object* l_instHOrElse___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Prelude_0__Lean_Macro_MethodsRefPointed; LEAN_EXPORT lean_object* l_dite___rarg___boxed(lean_object*, lean_object*, lean_object*); @@ -3127,6 +3133,50 @@ lean_dec(x_1); return x_3; } } +LEAN_EXPORT lean_object* l_Option_map___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; +lean_dec(x_1); +x_3 = lean_box(0); +return x_3; +} +else +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_2); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_ctor_get(x_2, 0); +x_6 = lean_apply_1(x_1, x_5); +lean_ctor_set(x_2, 0, x_6); +return x_2; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_ctor_get(x_2, 0); +lean_inc(x_7); +lean_dec(x_2); +x_8 = lean_apply_1(x_1, x_7); +x_9 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_9, 0, x_8); +return x_9; +} +} +} +} +LEAN_EXPORT lean_object* l_Option_map(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Option_map___rarg), 2, 0); +return x_3; +} +} LEAN_EXPORT lean_object* l_instInhabitedList(lean_object* x_1) { _start: { @@ -4730,9 +4780,9 @@ LEAN_EXPORT lean_object* l_Array_sequenceMap___rarg(lean_object* x_1, lean_objec { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; x_4 = lean_array_get_size(x_2); -x_5 = lean_unsigned_to_nat(0u); -x_6 = l_Array_empty___closed__1; -x_7 = l_Array_sequenceMap_loop___rarg(x_1, x_2, x_3, x_4, x_5, x_6); +x_5 = lean_mk_empty_array_with_capacity(x_4); +x_6 = lean_unsigned_to_nat(0u); +x_7 = l_Array_sequenceMap_loop___rarg(x_1, x_2, x_3, x_4, x_6, x_5); lean_dec(x_4); return x_7; } @@ -7209,6 +7259,23 @@ x_1 = lean_box(0); return x_1; } } +LEAN_EXPORT lean_object* l_Lean_instInhabitedTSyntax(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_box(0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_instInhabitedTSyntax___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_instInhabitedTSyntax(x_1); +lean_dec(x_1); +return x_2; +} +} static lean_object* _init_l_Lean_choiceKind___closed__1() { _start: { @@ -7571,7 +7638,7 @@ return x_6; { lean_object* x_7; lean_dec(x_1); -x_7 = l_Lean_identKind; +x_7 = l_Lean_identKind___closed__2; return x_7; } } @@ -7747,6 +7814,64 @@ lean_dec(x_1); return x_2; } } +LEAN_EXPORT lean_object* l_Lean_Syntax_getOptional_x3f(lean_object* x_1) { +_start: +{ +if (lean_obj_tag(x_1) == 1) +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_2 = lean_ctor_get(x_1, 1); +x_3 = lean_ctor_get(x_1, 2); +x_4 = l_Lean_nullKind___closed__2; +x_5 = lean_name_eq(x_2, x_4); +if (x_5 == 0) +{ +lean_object* x_6; +x_6 = lean_box(0); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_7 = lean_array_get_size(x_3); +x_8 = lean_unsigned_to_nat(1u); +x_9 = lean_nat_dec_eq(x_7, x_8); +lean_dec(x_7); +if (x_9 == 0) +{ +lean_object* x_10; +x_10 = lean_box(0); +return x_10; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = l_Lean_instInhabitedSyntax; +x_12 = lean_unsigned_to_nat(0u); +x_13 = lean_array_get(x_11, x_3, x_12); +x_14 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, x_13); +return x_14; +} +} +} +else +{ +lean_object* x_15; +x_15 = lean_box(0); +return x_15; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Syntax_getOptional_x3f___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Syntax_getOptional_x3f(x_1); +lean_dec(x_1); +return x_2; +} +} LEAN_EXPORT uint8_t l_Lean_Syntax_isMissing(lean_object* x_1) { _start: { @@ -9450,7 +9575,7 @@ LEAN_EXPORT uint8_t l_Lean_Syntax_matchesNull(lean_object* x_1, lean_object* x_2 _start: { lean_object* x_3; uint8_t x_4; -x_3 = l_Lean_nullKind; +x_3 = l_Lean_nullKind___closed__2; x_4 = l_Lean_Syntax_isNodeOf(x_1, x_3, x_2); return x_4; } diff --git a/stage0/stdlib/Init/System/IO.c b/stage0/stdlib/Init/System/IO.c index c82525bd5a8..3f841800131 100644 --- a/stage0/stdlib/Init/System/IO.c +++ b/stage0/stdlib/Init/System/IO.c @@ -99,7 +99,6 @@ static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init LEAN_EXPORT lean_object* l_EIO_toIO_x27(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_withStdin___at_Lean_runEval___spec__3(lean_object*, lean_object*, lean_object*); uint32_t lean_uint32_shift_left(uint32_t, uint32_t); -uint8_t lean_name_eq(lean_object*, lean_object*); lean_object* lean_io_prim_handle_write(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_System_IO_0__IO_FS_Handle_fopenFlags___closed__3; static lean_object* l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__11; @@ -124,7 +123,6 @@ static lean_object* l_termPrintln_x21_______closed__6; static lean_object* l___private_Init_System_IO_0__IO_FS_Handle_fopenFlags___closed__5; lean_object* lean_string_append(lean_object*, lean_object*); static lean_object* l_IO_FS_instReprSystemTime___closed__1; -extern lean_object* l_Lean_interpolatedStrKind; LEAN_EXPORT lean_object* l_IO_Process_SpawnArgs_args___default; static uint32_t l_IO_AccessRight_flags___closed__6; static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_1948____closed__21; @@ -430,7 +428,6 @@ LEAN_EXPORT lean_object* l_IO_FS_Mode_toCtorIdx(uint8_t); LEAN_EXPORT lean_object* l_IO_wait___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_BaseIO_mapTasks_go___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_IO_FS_instBEqFileType___closed__1; -lean_object* l_Lean_Syntax_getKind(lean_object*); LEAN_EXPORT lean_object* l_unsafeBaseIO(lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Stream_ofBuffer___elambda__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_unsafeBaseIO___rarg(lean_object*); @@ -487,6 +484,7 @@ static lean_object* l_IO_FS_instInhabitedStream___lambda__1___closed__3; LEAN_EXPORT lean_object* l_IO_FS_withIsolatedStreams(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_BaseIO_mapTasks_go___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_1948____closed__17; +static lean_object* l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__45; LEAN_EXPORT lean_object* l_IO_println___at_Lean_instEval__1___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_instInhabitedEIO___rarg(lean_object*); static lean_object* l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2120____closed__4; @@ -517,6 +515,7 @@ static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init static lean_object* l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2359____closed__7; static uint32_t l_IO_FS_instInhabitedStream___lambda__1___closed__1; static lean_object* l_instMonadEIO___closed__1; +static lean_object* l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__44; uint32_t lean_uint32_lor(uint32_t, uint32_t); lean_object* lean_io_app_path(lean_object*); static lean_object* l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__1; @@ -9677,7 +9676,7 @@ static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Lean", 4); +x_1 = lean_mk_string_from_bytes("interpolatedStrKind", 19); return x_1; } } @@ -9695,7 +9694,7 @@ static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Parser", 6); +x_1 = lean_mk_string_from_bytes("Lean", 4); return x_1; } } @@ -9703,7 +9702,7 @@ static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__2; +x_1 = lean_box(0); x_2 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -9713,7 +9712,7 @@ static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Term", 4); +x_1 = lean_mk_string_from_bytes("Parser", 6); return x_1; } } @@ -9731,7 +9730,7 @@ static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("paren", 5); +x_1 = lean_mk_string_from_bytes("Term", 4); return x_1; } } @@ -9749,29 +9748,47 @@ static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("(", 1); +x_1 = lean_mk_string_from_bytes("paren", 5); return x_1; } } static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__10() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__8; +x_2 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__9; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("(", 1); +return x_1; +} +} +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__12() { +_start: +{ lean_object* x_1; x_1 = lean_mk_string_from_bytes("null", 4); return x_1; } } -static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__11() { +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__10; +x_2 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__12; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__12() { +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__14() { _start: { lean_object* x_1; @@ -9779,17 +9796,17 @@ x_1 = lean_mk_string_from_bytes("app", 3); return x_1; } } -static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__13() { +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__6; -x_2 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__12; +x_1 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__8; +x_2 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__14; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__14() { +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__16() { _start: { lean_object* x_1; @@ -9797,22 +9814,22 @@ x_1 = lean_mk_string_from_bytes("IO.println", 10); return x_1; } } -static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__15() { +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__17() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__14; +x_1 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__16; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__16() { +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__14; +x_1 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__16; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__15; +x_3 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__17; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -9820,7 +9837,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__17() { +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__19() { _start: { lean_object* x_1; @@ -9828,17 +9845,17 @@ x_1 = lean_mk_string_from_bytes("IO", 2); return x_1; } } -static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__18() { +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__17; +x_2 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__19; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__19() { +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__21() { _start: { lean_object* x_1; @@ -9846,41 +9863,41 @@ x_1 = lean_mk_string_from_bytes("println", 7); return x_1; } } -static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__20() { +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__18; -x_2 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__19; +x_1 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__20; +x_2 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__21; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__21() { +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__20; +x_2 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__22; 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___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__22() { +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__21; +x_2 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__23; x_3 = lean_alloc_ctor(1, 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___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__23() { +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__25() { _start: { lean_object* x_1; lean_object* x_2; @@ -9889,7 +9906,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__24() { +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__26() { _start: { lean_object* x_1; lean_object* x_2; @@ -9898,7 +9915,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__25() { +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__27() { _start: { lean_object* x_1; @@ -9906,17 +9923,17 @@ x_1 = lean_mk_string_from_bytes("typeAscription", 14); return x_1; } } -static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__26() { +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__6; -x_2 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__25; +x_1 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__8; +x_2 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__27; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__27() { +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__29() { _start: { lean_object* x_1; @@ -9924,22 +9941,22 @@ x_1 = lean_mk_string_from_bytes(":", 1); return x_1; } } -static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__28() { +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__30() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__17; +x_1 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__19; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__29() { +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__31() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__17; +x_1 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__19; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__28; +x_3 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__30; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -9947,31 +9964,31 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__30() { +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__32() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__18; +x_2 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__20; 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___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__31() { +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__33() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__30; +x_2 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__32; x_3 = lean_alloc_ctor(1, 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___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__32() { +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__34() { _start: { lean_object* x_1; @@ -9979,22 +9996,22 @@ x_1 = lean_mk_string_from_bytes("Unit", 4); return x_1; } } -static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__33() { +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__35() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__32; +x_1 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__34; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__34() { +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__36() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__32; +x_1 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__34; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__33; +x_3 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__35; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -10002,41 +10019,41 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__35() { +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__37() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__32; +x_2 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__34; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__36() { +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__38() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__35; +x_2 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__37; 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___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__37() { +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__39() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__36; +x_2 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__38; x_3 = lean_alloc_ctor(1, 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___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__38() { +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__40() { _start: { lean_object* x_1; @@ -10044,7 +10061,7 @@ x_1 = lean_mk_string_from_bytes(")", 1); return x_1; } } -static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__39() { +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__41() { _start: { lean_object* x_1; lean_object* x_2; @@ -10053,7 +10070,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__40() { +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__42() { _start: { lean_object* x_1; @@ -10061,17 +10078,17 @@ x_1 = lean_mk_string_from_bytes("termS!_", 7); return x_1; } } -static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__41() { +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__43() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__40; +x_2 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__42; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__42() { +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__44() { _start: { lean_object* x_1; @@ -10079,12 +10096,12 @@ x_1 = lean_mk_string_from_bytes("s!", 2); return x_1; } } -static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__43() { +static lean_object* _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__45() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__11; +x_2 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__13; x_3 = l_IO_FS_lines___closed__1; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -10113,549 +10130,547 @@ return x_7; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; x_8 = lean_unsigned_to_nat(1u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); lean_dec(x_1); +x_10 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__2; lean_inc(x_9); -x_10 = l_Lean_Syntax_getKind(x_9); -x_11 = l_Lean_interpolatedStrKind; -x_12 = lean_name_eq(x_10, x_11); -lean_dec(x_10); -if (x_12 == 0) +x_11 = l_Lean_Syntax_isOfKind(x_9, x_10); +if (x_11 == 0) { -lean_object* x_13; uint8_t x_14; +lean_object* x_12; uint8_t x_13; lean_inc(x_2); -x_13 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) +x_12 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 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; lean_object* 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; lean_object* x_67; lean_object* x_68; -x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_ctor_get(x_2, 2); +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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; lean_object* x_67; +x_14 = lean_ctor_get(x_12, 0); +x_15 = lean_ctor_get(x_2, 2); +lean_inc(x_15); +x_16 = lean_ctor_get(x_2, 1); lean_inc(x_16); -x_17 = lean_ctor_get(x_2, 1); -lean_inc(x_17); lean_dec(x_2); -x_18 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__9; +x_17 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__11; +lean_inc(x_14); +x_18 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_18, 0, x_14); +lean_ctor_set(x_18, 1, x_17); +x_19 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__22; lean_inc(x_15); -x_19 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_19, 0, x_15); -lean_ctor_set(x_19, 1, x_18); -x_20 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__20; lean_inc(x_16); -lean_inc(x_17); -x_21 = l_Lean_addMacroScope(x_17, x_20, x_16); -x_22 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__16; -x_23 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__22; -lean_inc(x_15); -x_24 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_24, 0, x_15); -lean_ctor_set(x_24, 1, x_22); -lean_ctor_set(x_24, 2, x_21); -lean_ctor_set(x_24, 3, x_23); -x_25 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__23; -x_26 = lean_array_push(x_25, x_9); -x_27 = lean_box(2); -x_28 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__11; -x_29 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -lean_ctor_set(x_29, 2, x_26); -x_30 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__24; -x_31 = lean_array_push(x_30, x_24); -x_32 = lean_array_push(x_31, x_29); -x_33 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__13; -x_34 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_34, 0, x_27); -lean_ctor_set(x_34, 1, x_33); -lean_ctor_set(x_34, 2, x_32); -x_35 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__27; +x_20 = l_Lean_addMacroScope(x_16, x_19, x_15); +x_21 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__18; +x_22 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__24; +lean_inc(x_14); +x_23 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_23, 0, x_14); +lean_ctor_set(x_23, 1, x_21); +lean_ctor_set(x_23, 2, x_20); +lean_ctor_set(x_23, 3, x_22); +x_24 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__25; +x_25 = lean_array_push(x_24, x_9); +x_26 = lean_box(2); +x_27 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__13; +x_28 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +lean_ctor_set(x_28, 2, x_25); +x_29 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__26; +x_30 = lean_array_push(x_29, x_23); +x_31 = lean_array_push(x_30, x_28); +x_32 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__15; +x_33 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_33, 0, x_26); +lean_ctor_set(x_33, 1, x_32); +lean_ctor_set(x_33, 2, x_31); +x_34 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__29; +lean_inc(x_14); +x_35 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_35, 0, x_14); +lean_ctor_set(x_35, 1, x_34); +x_36 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__20; lean_inc(x_15); -x_36 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_36, 0, x_15); -lean_ctor_set(x_36, 1, x_35); -x_37 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__18; lean_inc(x_16); -lean_inc(x_17); -x_38 = l_Lean_addMacroScope(x_17, x_37, x_16); -x_39 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__29; -x_40 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__31; -lean_inc(x_15); -x_41 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_41, 0, x_15); -lean_ctor_set(x_41, 1, x_39); -lean_ctor_set(x_41, 2, x_38); -lean_ctor_set(x_41, 3, x_40); -x_42 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__35; -x_43 = l_Lean_addMacroScope(x_17, x_42, x_16); -x_44 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__34; -x_45 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__37; -lean_inc(x_15); -x_46 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_46, 0, x_15); -lean_ctor_set(x_46, 1, x_44); -lean_ctor_set(x_46, 2, x_43); -lean_ctor_set(x_46, 3, x_45); -x_47 = lean_array_push(x_25, x_46); -x_48 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_48, 0, x_27); -lean_ctor_set(x_48, 1, x_28); -lean_ctor_set(x_48, 2, x_47); -x_49 = lean_array_push(x_30, x_41); -x_50 = lean_array_push(x_49, x_48); -x_51 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_51, 0, x_27); -lean_ctor_set(x_51, 1, x_33); -lean_ctor_set(x_51, 2, x_50); -x_52 = lean_array_push(x_30, x_36); -x_53 = lean_array_push(x_52, x_51); -x_54 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__26; -x_55 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_55, 0, x_27); -lean_ctor_set(x_55, 1, x_54); -lean_ctor_set(x_55, 2, x_53); -x_56 = lean_array_push(x_25, x_55); -x_57 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_57, 0, x_27); -lean_ctor_set(x_57, 1, x_28); -lean_ctor_set(x_57, 2, x_56); -x_58 = lean_array_push(x_30, x_34); -x_59 = lean_array_push(x_58, x_57); -x_60 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_60, 0, x_27); -lean_ctor_set(x_60, 1, x_28); -lean_ctor_set(x_60, 2, x_59); -x_61 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__38; -x_62 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_62, 0, x_15); -lean_ctor_set(x_62, 1, x_61); -x_63 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__39; -x_64 = lean_array_push(x_63, x_19); -x_65 = lean_array_push(x_64, x_60); -x_66 = lean_array_push(x_65, x_62); -x_67 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__8; -x_68 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_68, 0, x_27); -lean_ctor_set(x_68, 1, x_67); -lean_ctor_set(x_68, 2, x_66); -lean_ctor_set(x_13, 0, x_68); -return x_13; +x_37 = l_Lean_addMacroScope(x_16, x_36, x_15); +x_38 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__31; +x_39 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__33; +lean_inc(x_14); +x_40 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_40, 0, x_14); +lean_ctor_set(x_40, 1, x_38); +lean_ctor_set(x_40, 2, x_37); +lean_ctor_set(x_40, 3, x_39); +x_41 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__37; +x_42 = l_Lean_addMacroScope(x_16, x_41, x_15); +x_43 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__36; +x_44 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__39; +lean_inc(x_14); +x_45 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_45, 0, x_14); +lean_ctor_set(x_45, 1, x_43); +lean_ctor_set(x_45, 2, x_42); +lean_ctor_set(x_45, 3, x_44); +x_46 = lean_array_push(x_24, x_45); +x_47 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_47, 0, x_26); +lean_ctor_set(x_47, 1, x_27); +lean_ctor_set(x_47, 2, x_46); +x_48 = lean_array_push(x_29, x_40); +x_49 = lean_array_push(x_48, x_47); +x_50 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_50, 0, x_26); +lean_ctor_set(x_50, 1, x_32); +lean_ctor_set(x_50, 2, x_49); +x_51 = lean_array_push(x_29, x_35); +x_52 = lean_array_push(x_51, x_50); +x_53 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__28; +x_54 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_54, 0, x_26); +lean_ctor_set(x_54, 1, x_53); +lean_ctor_set(x_54, 2, x_52); +x_55 = lean_array_push(x_24, x_54); +x_56 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_56, 0, x_26); +lean_ctor_set(x_56, 1, x_27); +lean_ctor_set(x_56, 2, x_55); +x_57 = lean_array_push(x_29, x_33); +x_58 = lean_array_push(x_57, x_56); +x_59 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_59, 0, x_26); +lean_ctor_set(x_59, 1, x_27); +lean_ctor_set(x_59, 2, x_58); +x_60 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__40; +x_61 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_61, 0, x_14); +lean_ctor_set(x_61, 1, x_60); +x_62 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__41; +x_63 = lean_array_push(x_62, x_18); +x_64 = lean_array_push(x_63, x_59); +x_65 = lean_array_push(x_64, x_61); +x_66 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__10; +x_67 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_67, 0, x_26); +lean_ctor_set(x_67, 1, x_66); +lean_ctor_set(x_67, 2, x_65); +lean_ctor_set(x_12, 0, x_67); +return x_12; } else { -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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; -x_69 = lean_ctor_get(x_13, 0); -x_70 = lean_ctor_get(x_13, 1); -lean_inc(x_70); +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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_68 = lean_ctor_get(x_12, 0); +x_69 = lean_ctor_get(x_12, 1); lean_inc(x_69); -lean_dec(x_13); -x_71 = lean_ctor_get(x_2, 2); +lean_inc(x_68); +lean_dec(x_12); +x_70 = lean_ctor_get(x_2, 2); +lean_inc(x_70); +x_71 = lean_ctor_get(x_2, 1); lean_inc(x_71); -x_72 = lean_ctor_get(x_2, 1); -lean_inc(x_72); lean_dec(x_2); -x_73 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__9; -lean_inc(x_69); -x_74 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_74, 0, x_69); -lean_ctor_set(x_74, 1, x_73); -x_75 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__20; +x_72 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__11; +lean_inc(x_68); +x_73 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_73, 0, x_68); +lean_ctor_set(x_73, 1, x_72); +x_74 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__22; +lean_inc(x_70); lean_inc(x_71); -lean_inc(x_72); -x_76 = l_Lean_addMacroScope(x_72, x_75, x_71); -x_77 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__16; -x_78 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__22; -lean_inc(x_69); -x_79 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_79, 0, x_69); -lean_ctor_set(x_79, 1, x_77); -lean_ctor_set(x_79, 2, x_76); -lean_ctor_set(x_79, 3, x_78); -x_80 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__23; -x_81 = lean_array_push(x_80, x_9); -x_82 = lean_box(2); -x_83 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__11; -x_84 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_84, 0, x_82); -lean_ctor_set(x_84, 1, x_83); -lean_ctor_set(x_84, 2, x_81); -x_85 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__24; -x_86 = lean_array_push(x_85, x_79); -x_87 = lean_array_push(x_86, x_84); -x_88 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__13; -x_89 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_89, 0, x_82); -lean_ctor_set(x_89, 1, x_88); -lean_ctor_set(x_89, 2, x_87); -x_90 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__27; -lean_inc(x_69); -x_91 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_91, 0, x_69); -lean_ctor_set(x_91, 1, x_90); -x_92 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__18; +x_75 = l_Lean_addMacroScope(x_71, x_74, x_70); +x_76 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__18; +x_77 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__24; +lean_inc(x_68); +x_78 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_78, 0, x_68); +lean_ctor_set(x_78, 1, x_76); +lean_ctor_set(x_78, 2, x_75); +lean_ctor_set(x_78, 3, x_77); +x_79 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__25; +x_80 = lean_array_push(x_79, x_9); +x_81 = lean_box(2); +x_82 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__13; +x_83 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_83, 0, x_81); +lean_ctor_set(x_83, 1, x_82); +lean_ctor_set(x_83, 2, x_80); +x_84 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__26; +x_85 = lean_array_push(x_84, x_78); +x_86 = lean_array_push(x_85, x_83); +x_87 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__15; +x_88 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_88, 0, x_81); +lean_ctor_set(x_88, 1, x_87); +lean_ctor_set(x_88, 2, x_86); +x_89 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__29; +lean_inc(x_68); +x_90 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_90, 0, x_68); +lean_ctor_set(x_90, 1, x_89); +x_91 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__20; +lean_inc(x_70); lean_inc(x_71); -lean_inc(x_72); -x_93 = l_Lean_addMacroScope(x_72, x_92, x_71); -x_94 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__29; -x_95 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__31; -lean_inc(x_69); -x_96 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_96, 0, x_69); -lean_ctor_set(x_96, 1, x_94); -lean_ctor_set(x_96, 2, x_93); -lean_ctor_set(x_96, 3, x_95); -x_97 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__35; -x_98 = l_Lean_addMacroScope(x_72, x_97, x_71); -x_99 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__34; -x_100 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__37; -lean_inc(x_69); -x_101 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_101, 0, x_69); -lean_ctor_set(x_101, 1, x_99); -lean_ctor_set(x_101, 2, x_98); -lean_ctor_set(x_101, 3, x_100); -x_102 = lean_array_push(x_80, x_101); -x_103 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_103, 0, x_82); -lean_ctor_set(x_103, 1, x_83); -lean_ctor_set(x_103, 2, x_102); -x_104 = lean_array_push(x_85, x_96); -x_105 = lean_array_push(x_104, x_103); -x_106 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_106, 0, x_82); -lean_ctor_set(x_106, 1, x_88); -lean_ctor_set(x_106, 2, x_105); -x_107 = lean_array_push(x_85, x_91); -x_108 = lean_array_push(x_107, x_106); -x_109 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__26; -x_110 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_110, 0, x_82); -lean_ctor_set(x_110, 1, x_109); -lean_ctor_set(x_110, 2, x_108); -x_111 = lean_array_push(x_80, x_110); -x_112 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_112, 0, x_82); -lean_ctor_set(x_112, 1, x_83); -lean_ctor_set(x_112, 2, x_111); -x_113 = lean_array_push(x_85, x_89); -x_114 = lean_array_push(x_113, x_112); -x_115 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_115, 0, x_82); -lean_ctor_set(x_115, 1, x_83); -lean_ctor_set(x_115, 2, x_114); -x_116 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__38; -x_117 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_117, 0, x_69); -lean_ctor_set(x_117, 1, x_116); -x_118 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__39; -x_119 = lean_array_push(x_118, x_74); -x_120 = lean_array_push(x_119, x_115); -x_121 = lean_array_push(x_120, x_117); -x_122 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__8; -x_123 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_123, 0, x_82); -lean_ctor_set(x_123, 1, x_122); -lean_ctor_set(x_123, 2, x_121); -x_124 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_124, 0, x_123); -lean_ctor_set(x_124, 1, x_70); -return x_124; -} -} -else -{ -lean_object* x_125; uint8_t x_126; +x_92 = l_Lean_addMacroScope(x_71, x_91, x_70); +x_93 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__31; +x_94 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__33; +lean_inc(x_68); +x_95 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_95, 0, x_68); +lean_ctor_set(x_95, 1, x_93); +lean_ctor_set(x_95, 2, x_92); +lean_ctor_set(x_95, 3, x_94); +x_96 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__37; +x_97 = l_Lean_addMacroScope(x_71, x_96, x_70); +x_98 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__36; +x_99 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__39; +lean_inc(x_68); +x_100 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_100, 0, x_68); +lean_ctor_set(x_100, 1, x_98); +lean_ctor_set(x_100, 2, x_97); +lean_ctor_set(x_100, 3, x_99); +x_101 = lean_array_push(x_79, x_100); +x_102 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_102, 0, x_81); +lean_ctor_set(x_102, 1, x_82); +lean_ctor_set(x_102, 2, x_101); +x_103 = lean_array_push(x_84, x_95); +x_104 = lean_array_push(x_103, x_102); +x_105 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_105, 0, x_81); +lean_ctor_set(x_105, 1, x_87); +lean_ctor_set(x_105, 2, x_104); +x_106 = lean_array_push(x_84, x_90); +x_107 = lean_array_push(x_106, x_105); +x_108 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__28; +x_109 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_109, 0, x_81); +lean_ctor_set(x_109, 1, x_108); +lean_ctor_set(x_109, 2, x_107); +x_110 = lean_array_push(x_79, x_109); +x_111 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_111, 0, x_81); +lean_ctor_set(x_111, 1, x_82); +lean_ctor_set(x_111, 2, x_110); +x_112 = lean_array_push(x_84, x_88); +x_113 = lean_array_push(x_112, x_111); +x_114 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_114, 0, x_81); +lean_ctor_set(x_114, 1, x_82); +lean_ctor_set(x_114, 2, x_113); +x_115 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__40; +x_116 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_116, 0, x_68); +lean_ctor_set(x_116, 1, x_115); +x_117 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__41; +x_118 = lean_array_push(x_117, x_73); +x_119 = lean_array_push(x_118, x_114); +x_120 = lean_array_push(x_119, x_116); +x_121 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__10; +x_122 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_122, 0, x_81); +lean_ctor_set(x_122, 1, x_121); +lean_ctor_set(x_122, 2, x_120); +x_123 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_123, 0, x_122); +lean_ctor_set(x_123, 1, x_69); +return x_123; +} +} +else +{ +lean_object* x_124; uint8_t x_125; lean_inc(x_2); -x_125 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_126 = !lean_is_exclusive(x_125); -if (x_126 == 0) +x_124 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_125 = !lean_is_exclusive(x_124); +if (x_125 == 0) { -lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; -x_127 = lean_ctor_get(x_125, 0); -x_128 = lean_ctor_get(x_2, 2); +lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; +x_126 = lean_ctor_get(x_124, 0); +x_127 = lean_ctor_get(x_2, 2); +lean_inc(x_127); +x_128 = lean_ctor_get(x_2, 1); lean_inc(x_128); -x_129 = lean_ctor_get(x_2, 1); -lean_inc(x_129); lean_dec(x_2); -x_130 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__9; +x_129 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__11; +lean_inc(x_126); +x_130 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_130, 0, x_126); +lean_ctor_set(x_130, 1, x_129); +x_131 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__22; lean_inc(x_127); -x_131 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_131, 0, x_127); -lean_ctor_set(x_131, 1, x_130); -x_132 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__20; lean_inc(x_128); -lean_inc(x_129); -x_133 = l_Lean_addMacroScope(x_129, x_132, x_128); -x_134 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__16; -x_135 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__22; -lean_inc(x_127); -x_136 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_136, 0, x_127); -lean_ctor_set(x_136, 1, x_134); -lean_ctor_set(x_136, 2, x_133); -lean_ctor_set(x_136, 3, x_135); -x_137 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__42; -lean_inc(x_127); -x_138 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_138, 0, x_127); -lean_ctor_set(x_138, 1, x_137); -x_139 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__24; -x_140 = lean_array_push(x_139, x_138); -x_141 = lean_array_push(x_140, x_9); -x_142 = lean_box(2); -x_143 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__41; -x_144 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_144, 0, x_142); -lean_ctor_set(x_144, 1, x_143); -lean_ctor_set(x_144, 2, x_141); -x_145 = lean_array_push(x_139, x_144); -x_146 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__43; -x_147 = lean_array_push(x_145, x_146); -x_148 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__11; -x_149 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_149, 0, x_142); -lean_ctor_set(x_149, 1, x_148); -lean_ctor_set(x_149, 2, x_147); -x_150 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__38; +x_132 = l_Lean_addMacroScope(x_128, x_131, x_127); +x_133 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__18; +x_134 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__24; +lean_inc(x_126); +x_135 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_135, 0, x_126); +lean_ctor_set(x_135, 1, x_133); +lean_ctor_set(x_135, 2, x_132); +lean_ctor_set(x_135, 3, x_134); +x_136 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__44; +lean_inc(x_126); +x_137 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_137, 0, x_126); +lean_ctor_set(x_137, 1, x_136); +x_138 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__26; +x_139 = lean_array_push(x_138, x_137); +x_140 = lean_array_push(x_139, x_9); +x_141 = lean_box(2); +x_142 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__43; +x_143 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_143, 0, x_141); +lean_ctor_set(x_143, 1, x_142); +lean_ctor_set(x_143, 2, x_140); +x_144 = lean_array_push(x_138, x_143); +x_145 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__45; +x_146 = lean_array_push(x_144, x_145); +x_147 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__13; +x_148 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_148, 0, x_141); +lean_ctor_set(x_148, 1, x_147); +lean_ctor_set(x_148, 2, x_146); +x_149 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__40; +lean_inc(x_126); +x_150 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_150, 0, x_126); +lean_ctor_set(x_150, 1, x_149); +x_151 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__41; +x_152 = lean_array_push(x_151, x_130); +lean_inc(x_152); +x_153 = lean_array_push(x_152, x_148); +lean_inc(x_150); +x_154 = lean_array_push(x_153, x_150); +x_155 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__10; +x_156 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_156, 0, x_141); +lean_ctor_set(x_156, 1, x_155); +lean_ctor_set(x_156, 2, x_154); +x_157 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__25; +x_158 = lean_array_push(x_157, x_156); +x_159 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_159, 0, x_141); +lean_ctor_set(x_159, 1, x_147); +lean_ctor_set(x_159, 2, x_158); +x_160 = lean_array_push(x_138, x_135); +x_161 = lean_array_push(x_160, x_159); +x_162 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__15; +x_163 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_163, 0, x_141); +lean_ctor_set(x_163, 1, x_162); +lean_ctor_set(x_163, 2, x_161); +x_164 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__29; +lean_inc(x_126); +x_165 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_165, 0, x_126); +lean_ctor_set(x_165, 1, x_164); +x_166 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__20; lean_inc(x_127); -x_151 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_151, 0, x_127); -lean_ctor_set(x_151, 1, x_150); -x_152 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__39; -x_153 = lean_array_push(x_152, x_131); -lean_inc(x_153); -x_154 = lean_array_push(x_153, x_149); -lean_inc(x_151); -x_155 = lean_array_push(x_154, x_151); -x_156 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__8; -x_157 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_157, 0, x_142); -lean_ctor_set(x_157, 1, x_156); -lean_ctor_set(x_157, 2, x_155); -x_158 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__23; -x_159 = lean_array_push(x_158, x_157); -x_160 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_160, 0, x_142); -lean_ctor_set(x_160, 1, x_148); -lean_ctor_set(x_160, 2, x_159); -x_161 = lean_array_push(x_139, x_136); -x_162 = lean_array_push(x_161, x_160); -x_163 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__13; -x_164 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_164, 0, x_142); -lean_ctor_set(x_164, 1, x_163); -lean_ctor_set(x_164, 2, x_162); -x_165 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__27; -lean_inc(x_127); -x_166 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_166, 0, x_127); -lean_ctor_set(x_166, 1, x_165); -x_167 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__18; lean_inc(x_128); -lean_inc(x_129); -x_168 = l_Lean_addMacroScope(x_129, x_167, x_128); -x_169 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__29; -x_170 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__31; -lean_inc(x_127); -x_171 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_171, 0, x_127); -lean_ctor_set(x_171, 1, x_169); -lean_ctor_set(x_171, 2, x_168); -lean_ctor_set(x_171, 3, x_170); -x_172 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__35; -x_173 = l_Lean_addMacroScope(x_129, x_172, x_128); -x_174 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__34; -x_175 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__37; -x_176 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_176, 0, x_127); -lean_ctor_set(x_176, 1, x_174); -lean_ctor_set(x_176, 2, x_173); -lean_ctor_set(x_176, 3, x_175); -x_177 = lean_array_push(x_158, x_176); -x_178 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_178, 0, x_142); -lean_ctor_set(x_178, 1, x_148); -lean_ctor_set(x_178, 2, x_177); -x_179 = lean_array_push(x_139, x_171); -x_180 = lean_array_push(x_179, x_178); -x_181 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_181, 0, x_142); -lean_ctor_set(x_181, 1, x_163); -lean_ctor_set(x_181, 2, x_180); -x_182 = lean_array_push(x_139, x_166); -x_183 = lean_array_push(x_182, x_181); -x_184 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__26; -x_185 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_185, 0, x_142); -lean_ctor_set(x_185, 1, x_184); -lean_ctor_set(x_185, 2, x_183); -x_186 = lean_array_push(x_158, x_185); -x_187 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_187, 0, x_142); -lean_ctor_set(x_187, 1, x_148); -lean_ctor_set(x_187, 2, x_186); -x_188 = lean_array_push(x_139, x_164); -x_189 = lean_array_push(x_188, x_187); -x_190 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_190, 0, x_142); -lean_ctor_set(x_190, 1, x_148); -lean_ctor_set(x_190, 2, x_189); -x_191 = lean_array_push(x_153, x_190); -x_192 = lean_array_push(x_191, x_151); -x_193 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_193, 0, x_142); -lean_ctor_set(x_193, 1, x_156); -lean_ctor_set(x_193, 2, x_192); -lean_ctor_set(x_125, 0, x_193); -return x_125; -} -else -{ -lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; -x_194 = lean_ctor_get(x_125, 0); -x_195 = lean_ctor_get(x_125, 1); -lean_inc(x_195); +x_167 = l_Lean_addMacroScope(x_128, x_166, x_127); +x_168 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__31; +x_169 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__33; +lean_inc(x_126); +x_170 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_170, 0, x_126); +lean_ctor_set(x_170, 1, x_168); +lean_ctor_set(x_170, 2, x_167); +lean_ctor_set(x_170, 3, x_169); +x_171 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__37; +x_172 = l_Lean_addMacroScope(x_128, x_171, x_127); +x_173 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__36; +x_174 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__39; +x_175 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_175, 0, x_126); +lean_ctor_set(x_175, 1, x_173); +lean_ctor_set(x_175, 2, x_172); +lean_ctor_set(x_175, 3, x_174); +x_176 = lean_array_push(x_157, x_175); +x_177 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_177, 0, x_141); +lean_ctor_set(x_177, 1, x_147); +lean_ctor_set(x_177, 2, x_176); +x_178 = lean_array_push(x_138, x_170); +x_179 = lean_array_push(x_178, x_177); +x_180 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_180, 0, x_141); +lean_ctor_set(x_180, 1, x_162); +lean_ctor_set(x_180, 2, x_179); +x_181 = lean_array_push(x_138, x_165); +x_182 = lean_array_push(x_181, x_180); +x_183 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__28; +x_184 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_184, 0, x_141); +lean_ctor_set(x_184, 1, x_183); +lean_ctor_set(x_184, 2, x_182); +x_185 = lean_array_push(x_157, x_184); +x_186 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_186, 0, x_141); +lean_ctor_set(x_186, 1, x_147); +lean_ctor_set(x_186, 2, x_185); +x_187 = lean_array_push(x_138, x_163); +x_188 = lean_array_push(x_187, x_186); +x_189 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_189, 0, x_141); +lean_ctor_set(x_189, 1, x_147); +lean_ctor_set(x_189, 2, x_188); +x_190 = lean_array_push(x_152, x_189); +x_191 = lean_array_push(x_190, x_150); +x_192 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_192, 0, x_141); +lean_ctor_set(x_192, 1, x_155); +lean_ctor_set(x_192, 2, x_191); +lean_ctor_set(x_124, 0, x_192); +return x_124; +} +else +{ +lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; +x_193 = lean_ctor_get(x_124, 0); +x_194 = lean_ctor_get(x_124, 1); lean_inc(x_194); -lean_dec(x_125); -x_196 = lean_ctor_get(x_2, 2); +lean_inc(x_193); +lean_dec(x_124); +x_195 = lean_ctor_get(x_2, 2); +lean_inc(x_195); +x_196 = lean_ctor_get(x_2, 1); lean_inc(x_196); -x_197 = lean_ctor_get(x_2, 1); -lean_inc(x_197); lean_dec(x_2); -x_198 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__9; -lean_inc(x_194); -x_199 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_199, 0, x_194); -lean_ctor_set(x_199, 1, x_198); -x_200 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__20; +x_197 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__11; +lean_inc(x_193); +x_198 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_198, 0, x_193); +lean_ctor_set(x_198, 1, x_197); +x_199 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__22; +lean_inc(x_195); lean_inc(x_196); -lean_inc(x_197); -x_201 = l_Lean_addMacroScope(x_197, x_200, x_196); -x_202 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__16; -x_203 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__22; -lean_inc(x_194); -x_204 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_204, 0, x_194); -lean_ctor_set(x_204, 1, x_202); -lean_ctor_set(x_204, 2, x_201); -lean_ctor_set(x_204, 3, x_203); -x_205 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__42; -lean_inc(x_194); -x_206 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_206, 0, x_194); -lean_ctor_set(x_206, 1, x_205); -x_207 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__24; -x_208 = lean_array_push(x_207, x_206); -x_209 = lean_array_push(x_208, x_9); -x_210 = lean_box(2); -x_211 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__41; -x_212 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_212, 0, x_210); -lean_ctor_set(x_212, 1, x_211); -lean_ctor_set(x_212, 2, x_209); -x_213 = lean_array_push(x_207, x_212); -x_214 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__43; -x_215 = lean_array_push(x_213, x_214); -x_216 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__11; -x_217 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_217, 0, x_210); -lean_ctor_set(x_217, 1, x_216); -lean_ctor_set(x_217, 2, x_215); -x_218 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__38; -lean_inc(x_194); -x_219 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_219, 0, x_194); -lean_ctor_set(x_219, 1, x_218); -x_220 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__39; -x_221 = lean_array_push(x_220, x_199); -lean_inc(x_221); -x_222 = lean_array_push(x_221, x_217); -lean_inc(x_219); -x_223 = lean_array_push(x_222, x_219); -x_224 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__8; -x_225 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_225, 0, x_210); -lean_ctor_set(x_225, 1, x_224); -lean_ctor_set(x_225, 2, x_223); -x_226 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__23; -x_227 = lean_array_push(x_226, x_225); -x_228 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_228, 0, x_210); -lean_ctor_set(x_228, 1, x_216); -lean_ctor_set(x_228, 2, x_227); -x_229 = lean_array_push(x_207, x_204); -x_230 = lean_array_push(x_229, x_228); -x_231 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__13; -x_232 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_232, 0, x_210); -lean_ctor_set(x_232, 1, x_231); -lean_ctor_set(x_232, 2, x_230); -x_233 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__27; -lean_inc(x_194); -x_234 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_234, 0, x_194); -lean_ctor_set(x_234, 1, x_233); -x_235 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__18; +x_200 = l_Lean_addMacroScope(x_196, x_199, x_195); +x_201 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__18; +x_202 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__24; +lean_inc(x_193); +x_203 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_203, 0, x_193); +lean_ctor_set(x_203, 1, x_201); +lean_ctor_set(x_203, 2, x_200); +lean_ctor_set(x_203, 3, x_202); +x_204 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__44; +lean_inc(x_193); +x_205 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_205, 0, x_193); +lean_ctor_set(x_205, 1, x_204); +x_206 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__26; +x_207 = lean_array_push(x_206, x_205); +x_208 = lean_array_push(x_207, x_9); +x_209 = lean_box(2); +x_210 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__43; +x_211 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_211, 0, x_209); +lean_ctor_set(x_211, 1, x_210); +lean_ctor_set(x_211, 2, x_208); +x_212 = lean_array_push(x_206, x_211); +x_213 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__45; +x_214 = lean_array_push(x_212, x_213); +x_215 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__13; +x_216 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_216, 0, x_209); +lean_ctor_set(x_216, 1, x_215); +lean_ctor_set(x_216, 2, x_214); +x_217 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__40; +lean_inc(x_193); +x_218 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_218, 0, x_193); +lean_ctor_set(x_218, 1, x_217); +x_219 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__41; +x_220 = lean_array_push(x_219, x_198); +lean_inc(x_220); +x_221 = lean_array_push(x_220, x_216); +lean_inc(x_218); +x_222 = lean_array_push(x_221, x_218); +x_223 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__10; +x_224 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_224, 0, x_209); +lean_ctor_set(x_224, 1, x_223); +lean_ctor_set(x_224, 2, x_222); +x_225 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__25; +x_226 = lean_array_push(x_225, x_224); +x_227 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_227, 0, x_209); +lean_ctor_set(x_227, 1, x_215); +lean_ctor_set(x_227, 2, x_226); +x_228 = lean_array_push(x_206, x_203); +x_229 = lean_array_push(x_228, x_227); +x_230 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__15; +x_231 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_231, 0, x_209); +lean_ctor_set(x_231, 1, x_230); +lean_ctor_set(x_231, 2, x_229); +x_232 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__29; +lean_inc(x_193); +x_233 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_233, 0, x_193); +lean_ctor_set(x_233, 1, x_232); +x_234 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__20; +lean_inc(x_195); lean_inc(x_196); -lean_inc(x_197); -x_236 = l_Lean_addMacroScope(x_197, x_235, x_196); -x_237 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__29; -x_238 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__31; -lean_inc(x_194); -x_239 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_239, 0, x_194); -lean_ctor_set(x_239, 1, x_237); -lean_ctor_set(x_239, 2, x_236); -lean_ctor_set(x_239, 3, x_238); -x_240 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__35; -x_241 = l_Lean_addMacroScope(x_197, x_240, x_196); -x_242 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__34; -x_243 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__37; -x_244 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_244, 0, x_194); -lean_ctor_set(x_244, 1, x_242); -lean_ctor_set(x_244, 2, x_241); -lean_ctor_set(x_244, 3, x_243); -x_245 = lean_array_push(x_226, x_244); -x_246 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_246, 0, x_210); -lean_ctor_set(x_246, 1, x_216); -lean_ctor_set(x_246, 2, x_245); -x_247 = lean_array_push(x_207, x_239); -x_248 = lean_array_push(x_247, x_246); -x_249 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_249, 0, x_210); -lean_ctor_set(x_249, 1, x_231); -lean_ctor_set(x_249, 2, x_248); -x_250 = lean_array_push(x_207, x_234); -x_251 = lean_array_push(x_250, x_249); -x_252 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__26; -x_253 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_253, 0, x_210); -lean_ctor_set(x_253, 1, x_252); -lean_ctor_set(x_253, 2, x_251); -x_254 = lean_array_push(x_226, x_253); -x_255 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_255, 0, x_210); -lean_ctor_set(x_255, 1, x_216); -lean_ctor_set(x_255, 2, x_254); -x_256 = lean_array_push(x_207, x_232); -x_257 = lean_array_push(x_256, x_255); -x_258 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_258, 0, x_210); -lean_ctor_set(x_258, 1, x_216); -lean_ctor_set(x_258, 2, x_257); -x_259 = lean_array_push(x_221, x_258); -x_260 = lean_array_push(x_259, x_219); -x_261 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_261, 0, x_210); -lean_ctor_set(x_261, 1, x_224); -lean_ctor_set(x_261, 2, x_260); -x_262 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_262, 0, x_261); -lean_ctor_set(x_262, 1, x_195); -return x_262; +x_235 = l_Lean_addMacroScope(x_196, x_234, x_195); +x_236 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__31; +x_237 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__33; +lean_inc(x_193); +x_238 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_238, 0, x_193); +lean_ctor_set(x_238, 1, x_236); +lean_ctor_set(x_238, 2, x_235); +lean_ctor_set(x_238, 3, x_237); +x_239 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__37; +x_240 = l_Lean_addMacroScope(x_196, x_239, x_195); +x_241 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__36; +x_242 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__39; +x_243 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_243, 0, x_193); +lean_ctor_set(x_243, 1, x_241); +lean_ctor_set(x_243, 2, x_240); +lean_ctor_set(x_243, 3, x_242); +x_244 = lean_array_push(x_225, x_243); +x_245 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_245, 0, x_209); +lean_ctor_set(x_245, 1, x_215); +lean_ctor_set(x_245, 2, x_244); +x_246 = lean_array_push(x_206, x_238); +x_247 = lean_array_push(x_246, x_245); +x_248 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_248, 0, x_209); +lean_ctor_set(x_248, 1, x_230); +lean_ctor_set(x_248, 2, x_247); +x_249 = lean_array_push(x_206, x_233); +x_250 = lean_array_push(x_249, x_248); +x_251 = l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__28; +x_252 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_252, 0, x_209); +lean_ctor_set(x_252, 1, x_251); +lean_ctor_set(x_252, 2, x_250); +x_253 = lean_array_push(x_225, x_252); +x_254 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_254, 0, x_209); +lean_ctor_set(x_254, 1, x_215); +lean_ctor_set(x_254, 2, x_253); +x_255 = lean_array_push(x_206, x_231); +x_256 = lean_array_push(x_255, x_254); +x_257 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_257, 0, x_209); +lean_ctor_set(x_257, 1, x_215); +lean_ctor_set(x_257, 2, x_256); +x_258 = lean_array_push(x_220, x_257); +x_259 = lean_array_push(x_258, x_218); +x_260 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_260, 0, x_209); +lean_ctor_set(x_260, 1, x_223); +lean_ctor_set(x_260, 2, x_259); +x_261 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_261, 0, x_260); +lean_ctor_set(x_261, 1, x_194); +return x_261; } } } @@ -11100,6 +11115,10 @@ l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__42 = lean_mark_persistent(l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__42); l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__43 = _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__43(); lean_mark_persistent(l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__43); +l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__44 = _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__44(); +lean_mark_persistent(l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__44); +l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__45 = _init_l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__45(); +lean_mark_persistent(l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__45); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Init/Tactics.c b/stage0/stdlib/Init/Tactics.c index 0c4c0e2fd9e..175ef2dd9b1 100644 --- a/stage0/stdlib/Init/Tactics.c +++ b/stage0/stdlib/Init/Tactics.c @@ -29,6 +29,7 @@ static lean_object* l_Lean_Parser_Tactic_cases___closed__6; static lean_object* l_Lean_Parser_Tactic_first___closed__17; static lean_object* l_Lean_Parser_Tactic_withAnnotateState___closed__12; static lean_object* l_Lean_Parser_Tactic_tacticRepeat_____closed__3; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__4; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticUnhygienic____1___closed__2; static lean_object* l_Lean_Parser_Tactic_save___closed__4; static lean_object* l_Lean_Parser_Tactic_first___closed__8; @@ -38,6 +39,7 @@ static lean_object* l_Lean_Parser_Tactic_withUnfoldingAll___closed__6; static lean_object* l_Lean_Parser_Tactic_case_x27___closed__7; static lean_object* l___aux__Init__Tactics______macroRules__term_u2039___u203a__1___closed__6; static lean_object* l_Lean_Parser_Tactic_letrec___closed__12; +static lean_object* l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__2; static lean_object* l_Lean_Parser_Tactic_revert___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_dsimp; static lean_object* l_Lean_Parser_Tactic_induction___closed__13; @@ -51,6 +53,7 @@ static lean_object* l_Lean_Parser_Tactic_inductionAlt___closed__11; static lean_object* l_Lean_Parser_Tactic_case___closed__6; static lean_object* l_Lean_Parser_Tactic_apply___closed__2; static lean_object* l_Lean_Parser_Tactic_dbgTrace___closed__5; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__5; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__6; static lean_object* l_Lean_Parser_Tactic_refl___closed__3; lean_object* lean_mk_empty_array_with_capacity(lean_object*); @@ -69,7 +72,6 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__L static lean_object* l_Lean_Parser_Tactic_simpErase___closed__6; static lean_object* l_Lean_Parser_Tactic_intro___closed__9; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTrivial__6___closed__7; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__5; static lean_object* l_Lean_Parser_Tactic_revert___closed__9; static lean_object* l_Lean_Parser_Tactic_inductionAlt___closed__19; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave_x27____1(lean_object*, lean_object*, lean_object*); @@ -77,11 +79,14 @@ static lean_object* l_Lean_Parser_Tactic_failIfSuccess___closed__6; static lean_object* l_Lean_Parser_Tactic_delta___closed__1; static lean_object* l_Lean_Parser_Tactic_dbgTrace___closed__3; static lean_object* l_Lean_Parser_Tactic_simpAll___closed__10; +static lean_object* l_Lean_binderIdent___closed__6; static lean_object* l_Lean_Parser_Tactic_location___closed__7; static lean_object* l_Lean_Parser_Tactic_generalize___closed__6; static lean_object* l_Lean_Parser_Tactic_save___closed__1; lean_object* lean_name_mk_string(lean_object*, lean_object*); +lean_object* l_Lean_Syntax_getOptional_x3f(lean_object*); static lean_object* l_Lean_Parser_Tactic_save___closed__2; +static lean_object* l_Lean_Parser_Tactic_renameI___closed__7; static lean_object* l_Lean_Parser_Attr_simp___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Attr_simp; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticUnhygienic____1___closed__4; @@ -90,6 +95,7 @@ static lean_object* l_Lean_Parser_Tactic_traceMessage___closed__1; static lean_object* l_Lean_Parser_Tactic_tacticRefine__lift_____closed__6; static lean_object* l_Lean_Parser_Tactic_injection___closed__2; static lean_object* l_Lean_Parser_Tactic_generalizeArg___closed__12; +static lean_object* l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__7; static lean_object* l_Lean_Parser_Tactic_first___closed__2; static lean_object* l_Lean_Parser_Tactic_rwRule___closed__4; static lean_object* l_Lean_Parser_Tactic_rwSeq___closed__4; @@ -108,9 +114,9 @@ static lean_object* l_Lean_Parser_Tactic_tacticSuffices_____closed__1; static lean_object* l_Lean_Parser_Tactic_induction___closed__2; static lean_object* l_Lean_Parser_Tactic_simpLemma___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_save; -static lean_object* l_Lean_Parser_Tactic_rwWithRfl___closed__2; static lean_object* l_Lean_Parser_Tactic_rotateLeft___closed__12; static lean_object* l_Lean_Parser_Tactic_refl___closed__5; +static lean_object* l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__3; static lean_object* l_Lean_Parser_Tactic_simpAll___closed__8; static lean_object* l_Lean_Parser_Tactic_discharger___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_simpAll; @@ -173,6 +179,7 @@ static lean_object* l_Lean_Parser_Tactic_withAnnotateState___closed__19; static lean_object* l_Lean_Parser_Tactic_tacticStop_____closed__7; static lean_object* l_Lean_Parser_Tactic_injection___closed__8; static lean_object* l_Lean_Parser_Tactic_changeWith___closed__5; +static lean_object* l_Lean_Parser_Tactic_inductionAltLHS___closed__13; static lean_object* l_Lean_Parser_Tactic_rwSeq___closed__1; static lean_object* l_Lean_Parser_Tactic_tacticInfer__instance___closed__3; static lean_object* l_Lean_Parser_Tactic_apply___closed__4; @@ -184,6 +191,7 @@ static lean_object* l_Lean_Parser_Tactic_change___closed__1; static lean_object* l_Lean_Parser_Tactic_tacticRefine__lift_____closed__5; static lean_object* l_Lean_Parser_Tactic_intro___closed__13; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticExists___x2c_x2c__1___closed__3; +static lean_object* l_Lean_binderIdent___closed__3; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticAdmit__1___closed__17; static lean_object* l_Lean_Parser_Tactic_rename___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_tacticUnhygienic__; @@ -194,15 +202,14 @@ static lean_object* l_Lean_Parser_Tactic_dsimp___closed__9; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticExists___x2c_x2c__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_tacticTrivial___closed__4; static lean_object* l_Lean_Parser_Tactic_case___closed__11; +static lean_object* l_Lean_binderIdent___closed__5; static lean_object* l_Lean_Parser_Tactic_tacticLet_____closed__4; static lean_object* l_Lean_Parser_Tactic_unfold___closed__8; static lean_object* l_Lean_Parser_Tactic_refl___closed__4; static lean_object* l_Lean_Parser_Tactic_rwRule___closed__8; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave_x27_____x3a_x3d____1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTrivial__6(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_tacticStop_____closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticAdmit__1___closed__8; -static lean_object* l_Lean_Parser_Tactic_intros___closed__16; static lean_object* l___aux__Init__Tactics______macroRules__term_u2039___u203a__1___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_tacticRepeat__; static lean_object* l_Lean_Parser_Tactic_injection___closed__4; @@ -215,9 +222,11 @@ static lean_object* l_Lean_Parser_Tactic_generalizeArg___closed__14; static lean_object* l_Lean_Parser_Tactic_config___closed__13; static lean_object* l_Lean_Parser_Tactic_withAnnotateState___closed__9; static lean_object* l_Lean_Parser_Tactic_cases___closed__5; +static lean_object* l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__8; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__letrec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_first___closed__11; static lean_object* l_Lean_Parser_Tactic_refine___closed__5; +static lean_object* l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__1; static lean_object* l_Lean_Parser_Tactic_tacticTrivial___closed__1; static lean_object* l_Lean_Parser_Tactic_inductionAlts___closed__10; static lean_object* l_Lean_Parser_Tactic_exact___closed__1; @@ -227,6 +236,7 @@ lean_object* lean_array_push(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_injection___closed__9; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticExists___x2c_x2c__1___closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tactic___x3c_x3b_x3e____1___closed__1; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__5; static lean_object* l_Lean_Parser_Tactic_contradiction___closed__1; static lean_object* l_Lean_Parser_Tactic_simp___closed__20; static lean_object* l___aux__Init__Tactics______macroRules__term_u2039___u203a__1___closed__3; @@ -236,14 +246,17 @@ static lean_object* l_Lean_Parser_Tactic_dsimp___closed__10; static lean_object* l_Lean_Parser_Tactic_rotateLeft___closed__7; static lean_object* l_Lean_Parser_Tactic_ac__refl___closed__3; static lean_object* l_Lean_Parser_Tactic_subst___closed__4; -static lean_object* l_Lean_Parser_Tactic_intros___closed__14; +static lean_object* l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_case_x27; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__2; static lean_object* l_Lean_Parser_Tactic_rotateRight___closed__1; static lean_object* l_Lean_Parser_Tactic_intros___closed__5; static lean_object* l_Lean_Parser_Tactic_renameI___closed__2; +static lean_object* l_Lean_binderIdent___closed__7; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave_____x3a_x3d____1___closed__2; static lean_object* l_Lean_Parser_Tactic_cases___closed__3; static lean_object* l_Lean_Parser_Tactic_generalizeArg___closed__5; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; static lean_object* l_Lean_Parser_Tactic_clear___closed__4; static lean_object* l_Lean_Parser_Tactic_tacticHave_____x3a_x3d_____closed__5; static lean_object* l_Lean_Parser_Tactic_tacticTrivial___closed__5; @@ -269,13 +282,10 @@ static lean_object* l_Lean_Parser_Tactic_case___closed__10; static lean_object* l_Lean_Parser_Tactic_locationHyp___closed__7; static lean_object* l_Lean_Parser_Tactic_done___closed__3; static lean_object* l_Lean_Parser_Tactic_generalizeArg___closed__4; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__3; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d__; static lean_object* l_Lean_Parser_Tactic_tacticInfer__instance___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_rename; static lean_object* l_Lean_Parser_Tactic_cases___closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticRepeat____1___closed__1; -static lean_object* l_Lean_Parser_Tactic_intros___closed__15; static lean_object* l_Lean_Parser_Tactic_tacticHave_____x3a_x3d_____closed__1; static lean_object* l_Lean_Parser_Tactic_tacticUnhygienic_____closed__2; static lean_object* l_Lean_Parser_Tactic_tacticTrivial___closed__2; @@ -304,7 +314,6 @@ static lean_object* l_Lean_Parser_Tactic_case___closed__4; static lean_object* l_Lean_Parser_Tactic_rotateLeft___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticAdmit__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simp___closed__17; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_rwWithRfl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticShow____1___closed__1; static lean_object* l_Lean_Parser_Tactic_tacticStop_____closed__5; static lean_object* l_Lean_Parser_Tactic_subst___closed__1; @@ -315,12 +324,10 @@ static lean_object* l_Lean_Parser_Tactic_tacticRefine__lift_x27_____closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_induction; static lean_object* l_Lean_Parser_Tactic_induction___closed__19; static lean_object* l_Lean_Parser_Tactic_case_x27___closed__6; -static lean_object* l_Lean_Parser_Tactic_rwWithRfl___closed__3; static lean_object* l_Lean_Parser_Tactic_letrec___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_simp; static lean_object* l_Lean_Parser_Tactic_simp___closed__9; static lean_object* l_Lean_Parser_Tactic_locationWildcard___closed__4; -static lean_object* l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__5; static lean_object* l_Lean_Parser_Tactic_withAnnotateState___closed__10; static lean_object* l_Lean_Parser_Tactic_discharger___closed__2; static lean_object* l_Lean_Parser_Tactic_tacticHave_____closed__7; @@ -331,8 +338,8 @@ static lean_object* l_Lean_Parser_Tactic_tacticAdmit___closed__3; static lean_object* l_Lean_Parser_Tactic_tacticInfer__instance___closed__1; static lean_object* l_Lean_Parser_Tactic_cases___closed__4; static lean_object* l_Lean_Parser_Tactic_dsimp___closed__5; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_assumption; +LEAN_EXPORT lean_object* l_Lean_binderIdent; static lean_object* l_Lean_Parser_Tactic_subst___closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticUnhygienic____1___closed__5; static lean_object* l_Lean_Parser_Tactic_inductionAlts___closed__7; @@ -379,7 +386,6 @@ static lean_object* l_Lean_Parser_Tactic_location___closed__2; static lean_object* l_Lean_Parser_Tactic_failIfSuccess___closed__5; static lean_object* l_Lean_Parser_Tactic_generalizeArg___closed__9; static lean_object* l_Lean_Parser_Tactic_delta___closed__6; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__1; static lean_object* l_Lean_Parser_Tactic_traceState___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_allGoals; static lean_object* l_Lean_Parser_Tactic_checkpoint___closed__5; @@ -389,6 +395,7 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__L static lean_object* l_Lean_Parser_Tactic_withReducibleAndInstances___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_cases; static lean_object* l_Lean_Parser_Tactic_traceMessage___closed__3; +static lean_object* l_Lean_binderIdent___closed__10; static lean_object* l_Lean_Parser_Tactic_dsimp___closed__4; static lean_object* l_Lean_Parser_Tactic_unfold___closed__3; static lean_object* l_Lean_Parser_Tactic_withAnnotateState___closed__16; @@ -396,13 +403,13 @@ static lean_object* l_Lean_Parser_Tactic_allGoals___closed__3; static lean_object* l_Lean_Parser_Tactic_subst___closed__6; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tactic___x3c_x3b_x3e____1___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_tacticRefine__lift__; -static lean_object* l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__8; static lean_object* l_Lean_Parser_Attr_simp___closed__5; static lean_object* l_Lean_Parser_Tactic_assumption___closed__1; static lean_object* l_Lean_Parser_Tactic_letrec___closed__13; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticRefine__lift_x27____1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTrivial__6___closed__2; static lean_object* l_Lean_Parser_Tactic_withReducible___closed__1; +static lean_object* l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__6; static lean_object* l_Lean_Parser_Tactic_tacticTry_____closed__5; static lean_object* l_Lean_Parser_Tactic_apply___closed__3; static lean_object* l_Lean_Parser_Tactic_clear___closed__5; @@ -419,6 +426,7 @@ static lean_object* l_Lean_Parser_Tactic_changeWith___closed__2; static lean_object* l_Lean_Parser_Tactic_withReducible___closed__2; static lean_object* l_Lean_Parser_Tactic_simpErase___closed__1; static lean_object* l_Lean_Parser_Tactic_generalizeArg___closed__13; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d__; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticAdmit__1___closed__3; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticAdmit__1___closed__16; static lean_object* l_Lean_Parser_Tactic_dsimp___closed__8; @@ -435,11 +443,13 @@ static lean_object* l_Lean_Parser_Tactic_discharger___closed__9; static lean_object* l_Lean_Parser_Tactic_tacticRfl___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_tacticAdmit; static lean_object* l_Lean_Parser_Tactic_location___closed__1; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_rewriteSeq; static lean_object* l_Lean_Parser_Tactic_traceMessage___closed__9; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_simpPre; static lean_object* l_Lean_Parser_Tactic_withReducibleAndInstances___closed__1; static lean_object* l_Lean_Parser_Tactic_checkpoint___closed__2; +static lean_object* l_Lean_Parser_Tactic_renameI___closed__8; static lean_object* l_Lean_Parser_Tactic_case_x27___closed__5; static lean_object* l_Lean_Parser_Attr_simp___closed__10; static lean_object* l_Lean_Parser_Tactic_fail___closed__2; @@ -457,6 +467,7 @@ static lean_object* l_Lean_Parser_Tactic_tacticShow_____closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_rotateLeft; static lean_object* l_Lean_Parser_Tactic_focus___closed__6; static lean_object* l_Lean_Parser_Tactic_changeWith___closed__1; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; static lean_object* l_Lean_Parser_Tactic_split___closed__2; static lean_object* l_Lean_Parser_Tactic_cases___closed__2; static lean_object* l_Lean_Parser_Tactic_tacticAdmit___closed__1; @@ -472,7 +483,6 @@ static lean_object* l_Lean_Parser_Tactic_tacticRfl___closed__3; static lean_object* l_Lean_Parser_Tactic_simpAll___closed__6; static lean_object* l_Lean_Parser_Tactic_split___closed__6; static lean_object* l_Lean_Parser_Tactic_simp___closed__4; -lean_object* l_Lean_Syntax_setKind(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_traceState___closed__4; static lean_object* l_Lean_Parser_Tactic_intro___closed__10; static lean_object* l_Lean_Parser_Tactic_injections___closed__2; @@ -490,7 +500,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_tacticRefine__lift_x27__; static lean_object* l_Lean_Parser_Tactic_simp___closed__13; static lean_object* l_Lean_Parser_Tactic_injection___closed__6; static lean_object* l_Lean_Parser_Tactic_fail___closed__1; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__7; static lean_object* l_Lean_Parser_Tactic_checkpoint___closed__1; static lean_object* l_Lean_Parser_Tactic_tacticRefine__lift_x27_____closed__3; static lean_object* l_Lean_Parser_Tactic_simpPre___closed__1; @@ -514,11 +523,10 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_delta; static lean_object* l_Lean_Parser_Tactic_casesTarget___closed__1; static lean_object* l_Lean_Parser_Tactic_injections___closed__4; static lean_object* l_Lean_Parser_Tactic_changeWith___closed__4; -static lean_object* l_Lean_Parser_Tactic_intros___closed__13; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTrivial__1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__8; static lean_object* l_Lean_Parser_Tactic_letrec___closed__1; static lean_object* l_Lean_Parser_Tactic_traceMessage___closed__4; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__2; static lean_object* l_Lean_Parser_Tactic_delta___closed__3; static lean_object* l_Lean_Parser_Tactic_tacticInfer__instance___closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTrivial__6___closed__4; @@ -537,15 +545,14 @@ static lean_object* l_Lean_Parser_Tactic_case_x27___closed__2; static lean_object* l_Lean_Parser_Tactic_withReducibleAndInstances___closed__3; static lean_object* l_Lean_Parser_Tactic_intro___closed__19; static lean_object* l_Lean_Parser_Tactic_renameI___closed__1; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_traceMessage___closed__8; -static lean_object* l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__6; static lean_object* l_Lean_Parser_Tactic_location___closed__8; static lean_object* l_Lean_Parser_Tactic_exact___closed__3; static lean_object* l_Lean_Parser_Tactic_inductionAlts___closed__3; static lean_object* l_Lean_Parser_Tactic_paren___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__save__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_discharger___closed__4; -static lean_object* l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__3; static lean_object* l_Lean_Parser_Tactic_tacticLet_x27_____closed__4; static lean_object* l_Lean_Parser_Tactic_assumption___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_location; @@ -553,14 +560,15 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRul static lean_object* l_Lean_Parser_Tactic_locationHyp___closed__6; static lean_object* l_Lean_Parser_Tactic_withReducibleAndInstances___closed__5; static lean_object* l_Lean_Parser_Tactic_checkpoint___closed__4; +static lean_object* l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__5; static lean_object* l_Lean_Parser_Tactic_tactic___x3c_x3b_x3e_____closed__4; static lean_object* l_Lean_Parser_Tactic_simpPre___closed__3; static lean_object* l_Lean_Parser_Tactic_exact___closed__5; static lean_object* l_Lean_Parser_Tactic_rwRuleSeq___closed__4; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__1; static lean_object* l_term_u2039___u203a___closed__6; static lean_object* l_Lean_Parser_Tactic_tacticHave_x27_____closed__5; static lean_object* l_Lean_Parser_Tactic_induction___closed__10; -static lean_object* l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__1; static lean_object* l_Lean_Parser_Tactic_case___closed__12; static lean_object* l_Lean_Parser_Tactic_paren___closed__3; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__letrec__1___closed__1; @@ -574,7 +582,6 @@ static lean_object* l_Lean_Parser_Tactic_inductionAlt___closed__15; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_checkpoint; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTrivial__6___closed__9; static lean_object* l_Lean_Parser_Tactic_simp___closed__6; -static lean_object* l_Lean_Parser_Tactic_withAnnotateState___closed__25; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticInfer__instance__1___closed__1; static lean_object* l_Lean_Parser_Tactic_rename___closed__2; static lean_object* l_Lean_Parser_Tactic_tacticHave_____closed__9; @@ -600,8 +607,10 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__L static lean_object* l_Lean_Parser_Tactic_allGoals___closed__2; static lean_object* l_Lean_Parser_Tactic_injection___closed__5; static lean_object* l_Lean_Parser_Tactic_tacticShow_____closed__5; +static lean_object* l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__8; static lean_object* l_Lean_Parser_Tactic_case_x27___closed__1; static lean_object* l_Lean_Parser_Tactic_withReducibleAndInstances___closed__4; +static lean_object* l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__4; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticShow____1___closed__3; static lean_object* l_Lean_Parser_Tactic_tactic___x3c_x3b_x3e_____closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_tacticLet_x27__; @@ -637,12 +646,12 @@ static lean_object* l_Lean_Parser_Tactic_simpAll___closed__1; static lean_object* l_Lean_Parser_Tactic_sleep___closed__2; static lean_object* l_Lean_Parser_Tactic_withReducible___closed__5; static lean_object* l_Lean_Parser_Tactic_letrec___closed__14; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_injection___closed__7; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__4; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_tacticNext_______x3d_x3e__; +static lean_object* l_Lean_binderIdent___closed__12; static lean_object* l_Lean_Parser_Tactic_tacticShow_____closed__4; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTrivial__6___closed__1; +static lean_object* l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__2; static lean_object* l_Lean_Parser_Tactic_inductionAlts___closed__6; lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_intro___closed__22; @@ -653,7 +662,7 @@ static lean_object* l_Lean_Parser_Tactic_generalizeArg___closed__6; static lean_object* l_Lean_Parser_Tactic_tacticHave_x27_____closed__3; static lean_object* l_Lean_Parser_Tactic_allGoals___closed__6; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTrivial__5___closed__1; -static lean_object* l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__5; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_tacticNext___x3d_x3e__; static lean_object* l_Lean_Parser_Tactic_fail___closed__4; static lean_object* l_Lean_Parser_Tactic_subst___closed__5; static lean_object* l_Lean_Parser_Tactic_inductionAlt___closed__5; @@ -662,27 +671,25 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_tacticExists___x2c_x2c; static lean_object* l_Lean_Parser_Tactic_tacticLet_x27_____closed__1; static lean_object* l_Lean_Parser_Tactic_tacticShow_____closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticRefine__lift____1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__6; static lean_object* l_Lean_Parser_Tactic_tacticRfl_x27___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_tacticLet__; static lean_object* l_Lean_Parser_Tactic_tacticRfl___closed__1; static lean_object* l_Lean_Parser_Tactic_intros___closed__4; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticAdmit__1___closed__23; static lean_object* l_Lean_Parser_Tactic_change___closed__6; +static lean_object* l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__7; static lean_object* l_Lean_Parser_Tactic_tacticRefine__lift_x27_____closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_simpLemma; static lean_object* l_Lean_Parser_Tactic_casesTarget___closed__3; static lean_object* l_Lean_Parser_Tactic_rotateRight___closed__5; static lean_object* l_Lean_Parser_Tactic_intro___closed__16; -lean_object* l_Lean_mkAtomFrom(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_delta___closed__4; static lean_object* l_Lean_Parser_Tactic_rewriteSeq___closed__6; static lean_object* l_Lean_Parser_Tactic_first___closed__10; static lean_object* l_Lean_Parser_Tactic_ac__refl___closed__4; -static lean_object* l_Lean_Parser_Tactic_intros___closed__9; +static lean_object* l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticSorry__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_sleep___closed__4; -static lean_object* l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__2; static lean_object* l_Lean_Parser_Tactic_tacticLet_x27_____closed__3; static lean_object* l_Lean_Parser_Tactic_generalizeArg___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_specialize; @@ -704,7 +711,6 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__L static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticInfer__instance__1___closed__2; LEAN_EXPORT lean_object* l___aux__Init__Tactics______macroRules__term_u2039___u203a__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_rwRuleSeq___closed__11; -static lean_object* l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__9; static lean_object* l_Lean_Parser_Tactic_tacticLet_x27_____closed__6; static lean_object* l_Lean_Parser_Tactic_refine___closed__2; static lean_object* l_Lean_Parser_Tactic_withAnnotateState___closed__18; @@ -715,12 +721,13 @@ static lean_object* l_Lean_Parser_Tactic_case___closed__3; static lean_object* l_Lean_Parser_Tactic_withAnnotateState___closed__3; static lean_object* l_Lean_Parser_Tactic_refine_x27___closed__3; static lean_object* l_Lean_Parser_Tactic_locationWildcard___closed__5; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_first___closed__15; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave_____x3a_x3d____1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTrivial__4___closed__1; +static lean_object* l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__5; static lean_object* l_Lean_Parser_Tactic_withAnnotateState___closed__7; static lean_object* l_Lean_Parser_Tactic_inductionAlts___closed__2; -static lean_object* l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__3; static lean_object* l_Lean_Parser_Tactic_changeWith___closed__6; static lean_object* l_Lean_Parser_Tactic_simpPre___closed__2; static lean_object* l_Lean_Parser_Tactic_generalizeArg___closed__11; @@ -729,7 +736,6 @@ static lean_object* l_Lean_Parser_Tactic_ac__refl___closed__1; static lean_object* l_Lean_Parser_Tactic_rwRuleSeq___closed__13; static lean_object* l_Lean_Parser_Tactic_sleep___closed__1; static lean_object* l_Lean_Parser_Tactic_focus___closed__5; -lean_object* l_Lean_Syntax_setArg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpErase___closed__2; static lean_object* l_Lean_Parser_Tactic_first___closed__18; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_generalizeArg; @@ -751,17 +757,14 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__L static lean_object* l_Lean_Parser_Tactic_tacticExists___x2c_x2c___closed__2; static lean_object* l_Lean_Parser_Tactic_refine___closed__4; static lean_object* l_Lean_Parser_Tactic_tacticTry_____closed__3; -static lean_object* l_Lean_Parser_Tactic_withAnnotateState___closed__24; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tactic___x3c_x3b_x3e____1___closed__3; +static lean_object* l_Lean_binderIdent___closed__11; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_exact; static lean_object* l_Lean_Parser_Tactic_tacticSorry___closed__3; static lean_object* l_Lean_Parser_Tactic_unfold___closed__5; static lean_object* l_Lean_Parser_Tactic_inductionAlt___closed__17; -static lean_object* l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__2; static lean_object* l_Lean_Parser_Tactic_first___closed__5; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; static lean_object* l_Lean_Parser_Tactic_substVars___closed__3; -static lean_object* l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__4; static lean_object* l_Lean_Parser_Tactic_case___closed__1; static lean_object* l_Lean_Parser_Tactic_tacticHave_____closed__1; lean_object* l_Lean_Syntax_getArgs(lean_object*); @@ -801,7 +804,6 @@ static lean_object* l_Lean_Parser_Tactic_locationHyp___closed__9; static lean_object* l_Lean_Parser_Tactic_injection___closed__10; static lean_object* l_Lean_Parser_Tactic_assumption___closed__3; static lean_object* l_Lean_Parser_Tactic_done___closed__1; -static lean_object* l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticStop____1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_rename___closed__9; static lean_object* l_Lean_Parser_Tactic_tacticHave_____closed__2; @@ -812,14 +814,12 @@ static lean_object* l_Lean_Parser_Tactic_rwRuleSeq___closed__12; static lean_object* l_Lean_Parser_Tactic_letrec___closed__8; static lean_object* l_Lean_Parser_Tactic_traceMessage___closed__6; static lean_object* l_Lean_Parser_Tactic_rwRule___closed__10; -static lean_object* l_Lean_Parser_Tactic_intros___closed__11; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_withReducibleAndInstances; static lean_object* l_Lean_Parser_Tactic_unfold___closed__6; static lean_object* l_Lean_Parser_Tactic_ac__refl___closed__5; static lean_object* l_Lean_Parser_Tactic_induction___closed__5; static lean_object* l_Lean_Parser_Tactic_case___closed__2; static lean_object* l_Lean_Parser_Tactic_simp___closed__2; -static lean_object* l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__1; static lean_object* l_Lean_Parser_Tactic_rename___closed__4; static lean_object* l_Lean_Parser_Tactic_generalizeArg___closed__10; static lean_object* l_Lean_Parser_Tactic_withAnnotateState___closed__4; @@ -867,27 +867,31 @@ static lean_object* l_Lean_Parser_Tactic_generalizeArg___closed__7; static lean_object* l_Lean_Parser_Tactic_constructor___closed__2; static lean_object* l_Lean_Parser_Tactic_tacticTrivial___closed__3; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave_x27____1___closed__1; +static lean_object* l_Lean_binderIdent___closed__9; static lean_object* l_Lean_Parser_Tactic_tacticExists___x2c_x2c___closed__4; static lean_object* l_Lean_Parser_Tactic_exact___closed__4; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__7; static lean_object* l_Lean_Parser_Tactic_tacticRefine__lift_____closed__4; +static lean_object* l_Lean_binderIdent___closed__8; static lean_object* l_Lean_Parser_Tactic_done___closed__4; -static lean_object* l_Lean_Parser_Tactic_rwWithRfl___closed__4; static lean_object* l_Lean_Parser_Tactic_tacticRfl_x27___closed__3; static lean_object* l_Lean_Parser_Tactic_simpLemma___closed__7; static lean_object* l_Lean_Parser_Tactic_apply___closed__1; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__4; static lean_object* l_Lean_Parser_Tactic_rwSeq___closed__7; static lean_object* l_Lean_Parser_Tactic_renameI___closed__5; +static lean_object* l_Lean_binderIdent___closed__4; static lean_object* l_Lean_Parser_Tactic_specialize___closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticRfl_x27__1___closed__8; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTrivial__5(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_letrec___closed__7; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__6; static lean_object* l_Lean_Parser_Tactic_inductionAlt___closed__13; static lean_object* l_Lean_Parser_Tactic_first___closed__6; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTrivial__5___closed__4; static lean_object* l_Lean_Parser_Tactic_generalize___closed__7; uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_renameI___closed__6; +static lean_object* l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_tacticStop__; static lean_object* l_Lean_Parser_Tactic_induction___closed__4; static lean_object* l_Lean_Parser_Tactic_inductionAltLHS___closed__2; @@ -923,8 +927,6 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__L static lean_object* l_Lean_Parser_Tactic_simpLemma___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_constructor; static lean_object* l_Lean_Parser_Tactic_tactic___x3c_x3b_x3e_____closed__5; -static lean_object* l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__7; -static lean_object* l_Lean_Parser_Tactic_intros___closed__12; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave_____x3a_x3d____1___closed__7; static lean_object* l_Lean_Parser_Tactic_substVars___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_sleep; @@ -934,19 +936,23 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__L static lean_object* l_Lean_Parser_Tactic_fail___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_discharger; static lean_object* l_Lean_Parser_Tactic_withAnnotateState___closed__8; +static lean_object* l_Lean_binderIdent___closed__2; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); +static lean_object* l_Lean_binderIdent___closed__1; static lean_object* l_Lean_Parser_Tactic_inductionAltLHS___closed__9; LEAN_EXPORT lean_object* l_term_u2039___u203a; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_withUnfoldingAll; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_dbgTrace___closed__6; static lean_object* l_Lean_Parser_Tactic_induction___closed__9; +static lean_object* l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__3; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTrivial__5___closed__8; static lean_object* l_Lean_Parser_Tactic_specialize___closed__3; static lean_object* l_Lean_Parser_Tactic_tacticSorry___closed__1; static lean_object* l_term_u2039___u203a___closed__8; static lean_object* l_Lean_Parser_Tactic_fail___closed__6; static lean_object* l_Lean_Parser_Tactic_rotateLeft___closed__1; +static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__3; static lean_object* l_Lean_Parser_Tactic_allGoals___closed__4; static lean_object* l_Lean_Parser_Tactic_tacticSorry___closed__2; static lean_object* l_Lean_Parser_Tactic_intro___closed__1; @@ -959,6 +965,7 @@ static lean_object* l_Lean_Parser_Tactic_inductionAlts___closed__4; static lean_object* l_Lean_Parser_Tactic_simpAll___closed__12; static lean_object* l_Lean_Parser_Tactic_intro___closed__18; static lean_object* l_Lean_Parser_Tactic_first___closed__13; +static lean_object* l_Lean_binderIdent___closed__13; static lean_object* l_Lean_Parser_Tactic_tacticLet_____closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_split; static lean_object* l_Lean_Parser_Tactic_dbgTrace___closed__2; @@ -989,7 +996,6 @@ static lean_object* l_Lean_Parser_Tactic_simpAll___closed__5; static lean_object* l_Lean_Parser_Tactic_inductionAlts___closed__8; static lean_object* l_Lean_Parser_Tactic_induction___closed__17; static lean_object* l_Lean_Parser_Tactic_withAnnotateState___closed__13; -static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_casesTarget; static lean_object* l_Lean_Parser_Tactic_letrec___closed__3; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave_____x3a_x3d____1___closed__3; @@ -1011,8 +1017,6 @@ static lean_object* l_Lean_Parser_Tactic_casesTarget___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_tacticShow__; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTrivial__5___closed__2; static lean_object* l_Lean_Parser_Tactic_rename___closed__6; -static lean_object* l_Lean_Parser_Tactic_expandRwSeq___closed__1; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandRwSeq(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_tacticRefine__lift_____closed__1; static lean_object* l_Lean_Parser_Tactic_rotateLeft___closed__9; static lean_object* l_Lean_Parser_Tactic_allGoals___closed__1; @@ -1034,9 +1038,9 @@ static lean_object* l_Lean_Parser_Tactic_tacticRefine__lift_x27_____closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_paren; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_intro; static lean_object* l_Lean_Parser_Tactic_tacticExists___x2c_x2c___closed__6; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave_x27___x3a_x3d____1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_constructor___closed__1; static lean_object* l_Lean_Parser_Tactic_tacticHave_x27_____closed__1; -static lean_object* l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_locationHyp; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_refine; static lean_object* l_Lean_Parser_Tactic_rwSeq___closed__6; @@ -1055,7 +1059,6 @@ static lean_object* l_Lean_Parser_Tactic_skip___closed__4; static lean_object* l_Lean_Parser_Tactic_rwRule___closed__7; static lean_object* l_Lean_Parser_Tactic_rotateRight___closed__6; static lean_object* l_term_u2039___u203a___closed__4; -static lean_object* l_Lean_Parser_Tactic_rwWithRfl___closed__1; static lean_object* l_Lean_Parser_Tactic_rewriteSeq___closed__2; static lean_object* l_Lean_Parser_Tactic_dbgTrace___closed__1; static lean_object* l_Lean_Parser_Tactic_inductionAltLHS___closed__7; @@ -1068,7 +1071,6 @@ static lean_object* l_Lean_Parser_Tactic_refine_x27___closed__1; static lean_object* l_Lean_Parser_Tactic_specialize___closed__4; static lean_object* l_Lean_Parser_Attr_simp___closed__4; static lean_object* l_Lean_Parser_Tactic_inductionAltLHS___closed__5; -static lean_object* l_Lean_Parser_Tactic_intros___closed__10; static lean_object* l_Lean_Parser_Tactic_refine_x27___closed__2; static lean_object* l_Lean_Parser_Tactic_tactic___x3c_x3b_x3e_____closed__3; static lean_object* l_Lean_Parser_Tactic_simpPost___closed__1; @@ -1089,7 +1091,7 @@ static lean_object* l_Lean_Parser_Tactic_config___closed__3; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticAdmit__1___closed__5; static lean_object* l_Lean_Parser_Tactic_config___closed__5; static lean_object* l_Lean_Parser_Attr_simp___closed__8; -static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__1() { +static lean_object* _init_l_Lean_binderIdent___closed__1() { _start: { lean_object* x_1; @@ -1097,17 +1099,135 @@ x_1 = lean_mk_string_from_bytes("Lean", 4); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__2() { +static lean_object* _init_l_Lean_binderIdent___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic_withAnnotateState___closed__1; +x_2 = l_Lean_binderIdent___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__3() { +static lean_object* _init_l_Lean_binderIdent___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("binderIdent", 11); +return x_1; +} +} +static lean_object* _init_l_Lean_binderIdent___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_binderIdent___closed__2; +x_2 = l_Lean_binderIdent___closed__3; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_binderIdent___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("orelse", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_binderIdent___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_binderIdent___closed__5; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_binderIdent___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("ident", 5); +return x_1; +} +} +static lean_object* _init_l_Lean_binderIdent___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_binderIdent___closed__7; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_binderIdent___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_binderIdent___closed__8; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_binderIdent___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("_", 1); +return x_1; +} +} +static lean_object* _init_l_Lean_binderIdent___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_binderIdent___closed__10; +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_binderIdent___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_binderIdent___closed__6; +x_2 = l_Lean_binderIdent___closed__9; +x_3 = l_Lean_binderIdent___closed__11; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_binderIdent___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_binderIdent___closed__3; +x_2 = l_Lean_binderIdent___closed__4; +x_3 = l_Lean_binderIdent___closed__12; +x_4 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_binderIdent() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_binderIdent___closed__13; +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__1() { _start: { lean_object* x_1; @@ -1115,17 +1235,17 @@ x_1 = lean_mk_string_from_bytes("Parser", 6); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__4() { +static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__2; -x_2 = l_Lean_Parser_Tactic_withAnnotateState___closed__3; +x_1 = l_Lean_binderIdent___closed__2; +x_2 = l_Lean_Parser_Tactic_withAnnotateState___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__3() { _start: { lean_object* x_1; @@ -1133,17 +1253,17 @@ x_1 = lean_mk_string_from_bytes("Tactic", 6); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; -x_2 = l_Lean_Parser_Tactic_withAnnotateState___closed__5; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__2; +x_2 = l_Lean_Parser_Tactic_withAnnotateState___closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__7() { +static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__5() { _start: { lean_object* x_1; @@ -1151,17 +1271,17 @@ x_1 = lean_mk_string_from_bytes("withAnnotateState", 17); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__8() { +static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; -x_2 = l_Lean_Parser_Tactic_withAnnotateState___closed__7; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; +x_2 = l_Lean_Parser_Tactic_withAnnotateState___closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__9() { +static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__7() { _start: { lean_object* x_1; @@ -1169,17 +1289,17 @@ x_1 = lean_mk_string_from_bytes("andthen", 7); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__10() { +static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic_withAnnotateState___closed__9; +x_2 = l_Lean_Parser_Tactic_withAnnotateState___closed__7; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__11() { +static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__9() { _start: { lean_object* x_1; @@ -1187,11 +1307,11 @@ x_1 = lean_mk_string_from_bytes("with_annotate_state ", 20); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__12() { +static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__10() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__11; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__9; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -1199,7 +1319,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__13() { +static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__11() { _start: { lean_object* x_1; @@ -1207,21 +1327,21 @@ x_1 = lean_mk_string_from_bytes("rawStx", 6); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__14() { +static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic_withAnnotateState___closed__13; +x_2 = l_Lean_Parser_Tactic_withAnnotateState___closed__11; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__15() { +static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__14; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__12; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1229,13 +1349,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__16() { +static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; -x_2 = l_Lean_Parser_Tactic_withAnnotateState___closed__12; -x_3 = l_Lean_Parser_Tactic_withAnnotateState___closed__15; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; +x_2 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_3 = l_Lean_Parser_Tactic_withAnnotateState___closed__13; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1243,7 +1363,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__17() { +static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__15() { _start: { lean_object* x_1; @@ -1251,33 +1371,33 @@ x_1 = lean_mk_string_from_bytes("ppSpace", 7); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__18() { +static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic_withAnnotateState___closed__17; +x_2 = l_Lean_Parser_Tactic_withAnnotateState___closed__15; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__19() { +static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__17() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__18; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__16; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__20() { +static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; -x_2 = l_Lean_Parser_Tactic_withAnnotateState___closed__16; -x_3 = l_Lean_Parser_Tactic_withAnnotateState___closed__19; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; +x_2 = l_Lean_Parser_Tactic_withAnnotateState___closed__14; +x_3 = l_Lean_Parser_Tactic_withAnnotateState___closed__17; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1285,7 +1405,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__21() { +static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__19() { _start: { lean_object* x_1; @@ -1293,21 +1413,21 @@ x_1 = lean_mk_string_from_bytes("tactic", 6); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__22() { +static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic_withAnnotateState___closed__21; +x_2 = l_Lean_Parser_Tactic_withAnnotateState___closed__19; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__23() { +static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__22; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__20; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1315,13 +1435,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__24() { +static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; -x_2 = l_Lean_Parser_Tactic_withAnnotateState___closed__20; -x_3 = l_Lean_Parser_Tactic_withAnnotateState___closed__23; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; +x_2 = l_Lean_Parser_Tactic_withAnnotateState___closed__18; +x_3 = l_Lean_Parser_Tactic_withAnnotateState___closed__21; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1329,13 +1449,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__25() { +static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState___closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Parser_Tactic_withAnnotateState___closed__24; +x_3 = l_Lean_Parser_Tactic_withAnnotateState___closed__22; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1347,7 +1467,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_withAnnotateState() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__25; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__23; return x_1; } } @@ -1363,7 +1483,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_intro___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_intro___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -1441,7 +1561,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_intro___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_intro___closed__4; x_3 = l_Lean_Parser_Tactic_intro___closed__9; x_4 = lean_alloc_ctor(2, 3, 0); @@ -1531,7 +1651,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_intro___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_intro___closed__15; x_3 = l_Lean_Parser_Tactic_intro___closed__18; x_4 = lean_alloc_ctor(2, 3, 0); @@ -1557,7 +1677,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_intro___closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_intro___closed__10; x_3 = l_Lean_Parser_Tactic_intro___closed__20; x_4 = lean_alloc_ctor(2, 3, 0); @@ -1601,7 +1721,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_intros___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_intros___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -1630,88 +1750,10 @@ return x_3; static lean_object* _init_l_Lean_Parser_Tactic_intros___closed__5() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("orelse", 6); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Tactic_intros___closed__6() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic_intros___closed__5; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Tactic_intros___closed__7() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("ident", 5); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Tactic_intros___closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic_intros___closed__7; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Tactic_intros___closed__9() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_intros___closed__8; -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Parser_Tactic_intros___closed__10() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("_", 1); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Tactic_intros___closed__11() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_intros___closed__10; -x_2 = lean_alloc_ctor(5, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Parser_Tactic_intros___closed__12() { -_start: -{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_intros___closed__6; -x_2 = l_Lean_Parser_Tactic_intros___closed__9; -x_3 = l_Lean_Parser_Tactic_intros___closed__11; -x_4 = lean_alloc_ctor(2, 3, 0); -lean_ctor_set(x_4, 0, x_1); -lean_ctor_set(x_4, 1, x_2); -lean_ctor_set(x_4, 2, x_3); -return x_4; -} -} -static lean_object* _init_l_Lean_Parser_Tactic_intros___closed__13() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_intro___closed__15; -x_3 = l_Lean_Parser_Tactic_intros___closed__12; +x_3 = l_Lean_binderIdent___closed__12; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1719,25 +1761,25 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_intros___closed__14() { +static lean_object* _init_l_Lean_Parser_Tactic_intros___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_intro___closed__12; -x_2 = l_Lean_Parser_Tactic_intros___closed__13; +x_2 = l_Lean_Parser_Tactic_intros___closed__5; x_3 = lean_alloc_ctor(1, 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_Lean_Parser_Tactic_intros___closed__15() { +static lean_object* _init_l_Lean_Parser_Tactic_intros___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_intros___closed__4; -x_3 = l_Lean_Parser_Tactic_intros___closed__14; +x_3 = l_Lean_Parser_Tactic_intros___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1745,13 +1787,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_intros___closed__16() { +static lean_object* _init_l_Lean_Parser_Tactic_intros___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Tactic_intros___closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Parser_Tactic_intros___closed__15; +x_3 = l_Lean_Parser_Tactic_intros___closed__7; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1763,7 +1805,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_intros() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_intros___closed__16; +x_1 = l_Lean_Parser_Tactic_intros___closed__8; return x_1; } } @@ -1779,7 +1821,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_rename___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_rename___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -1821,7 +1863,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_rename___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_rename___closed__4; x_3 = l_Lean_Parser_Tactic_rename___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -1853,7 +1895,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_rename___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_rename___closed__6; x_3 = l_Lean_Parser_Tactic_rename___closed__8; x_4 = lean_alloc_ctor(2, 3, 0); @@ -1867,9 +1909,9 @@ static lean_object* _init_l_Lean_Parser_Tactic_rename___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_rename___closed__9; -x_3 = l_Lean_Parser_Tactic_intros___closed__9; +x_3 = l_Lean_binderIdent___closed__9; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1911,7 +1953,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_revert___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_revert___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -1971,7 +2013,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_revert___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_revert___closed__4; x_3 = l_Lean_Parser_Tactic_revert___closed__7; x_4 = lean_alloc_ctor(2, 3, 0); @@ -2015,7 +2057,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_clear___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_clear___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -2045,7 +2087,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_clear___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_clear___closed__4; x_3 = l_Lean_Parser_Tactic_revert___closed__7; x_4 = lean_alloc_ctor(2, 3, 0); @@ -2089,7 +2131,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_subst___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_subst___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -2119,7 +2161,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_subst___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_subst___closed__4; x_3 = l_Lean_Parser_Tactic_revert___closed__7; x_4 = lean_alloc_ctor(2, 3, 0); @@ -2163,7 +2205,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_substVars___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_substVars___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -2223,7 +2265,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_assumption___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_assumption___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -2275,7 +2317,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_contradiction___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_contradiction___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -2327,7 +2369,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_apply___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_apply___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -2357,7 +2399,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_apply___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_apply___closed__4; x_3 = l_Lean_Parser_Tactic_rename___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -2401,7 +2443,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_exact___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_exact___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -2431,7 +2473,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_exact___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_exact___closed__4; x_3 = l_Lean_Parser_Tactic_rename___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -2475,7 +2517,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_refine___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_refine___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -2505,7 +2547,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_refine___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_refine___closed__4; x_3 = l_Lean_Parser_Tactic_rename___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -2549,7 +2591,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_refine_x27___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_refine_x27___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -2579,7 +2621,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_refine_x27___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_refine_x27___closed__4; x_3 = l_Lean_Parser_Tactic_rename___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -2623,7 +2665,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_constructor___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_constructor___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -2675,7 +2717,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_case___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_case___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -2705,9 +2747,9 @@ static lean_object* _init_l_Lean_Parser_Tactic_case___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_case___closed__4; -x_3 = l_Lean_Parser_Tactic_intros___closed__12; +x_3 = l_Lean_binderIdent; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2720,7 +2762,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_case___closed__6() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_intro___closed__12; -x_2 = l_Lean_Parser_Tactic_intros___closed__12; +x_2 = l_Lean_binderIdent; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -2731,7 +2773,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_case___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_case___closed__5; x_3 = l_Lean_Parser_Tactic_case___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); @@ -2745,7 +2787,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_case___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_case___closed__7; x_3 = l_Lean_Parser_Tactic_rename___closed__8; x_4 = lean_alloc_ctor(2, 3, 0); @@ -2787,7 +2829,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_case___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_case___closed__8; x_3 = l_Lean_Parser_Tactic_case___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); @@ -2831,7 +2873,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_case_x27___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_case_x27___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -2861,9 +2903,9 @@ static lean_object* _init_l_Lean_Parser_Tactic_case_x27___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_case_x27___closed__4; -x_3 = l_Lean_Parser_Tactic_intros___closed__12; +x_3 = l_Lean_binderIdent; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2875,7 +2917,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_case_x27___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_case_x27___closed__5; x_3 = l_Lean_Parser_Tactic_case___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); @@ -2889,7 +2931,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_case_x27___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_case_x27___closed__6; x_3 = l_Lean_Parser_Tactic_rename___closed__8; x_4 = lean_alloc_ctor(2, 3, 0); @@ -2903,7 +2945,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_case_x27___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_case_x27___closed__7; x_3 = l_Lean_Parser_Tactic_case___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); @@ -2935,25 +2977,25 @@ x_1 = l_Lean_Parser_Tactic_case_x27___closed__9; return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__1() { +static lean_object* _init_l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("tacticNext___=>_", 16); +x_1 = lean_mk_string_from_bytes("tacticNext_=>_", 14); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__2() { +static lean_object* _init_l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; -x_2 = l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__1; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; +x_2 = l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__3() { +static lean_object* _init_l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__3() { _start: { lean_object* x_1; @@ -2961,11 +3003,11 @@ x_1 = lean_mk_string_from_bytes("next ", 5); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__4() { +static lean_object* _init_l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__3; +x_1 = l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__3; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2973,25 +3015,13 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_intro___closed__12; -x_2 = l_Lean_Parser_Tactic_intros___closed__12; -x_3 = lean_alloc_ctor(1, 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_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; -x_2 = l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__4; -x_3 = l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__5; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; +x_2 = l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__4; +x_3 = l_Lean_Parser_Tactic_case___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2999,12 +3029,12 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__7() { +static lean_object* _init_l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; -x_2 = l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; +x_2 = l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__5; x_3 = l_Lean_Parser_Tactic_rename___closed__8; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -3013,12 +3043,12 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__8() { +static lean_object* _init_l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; -x_2 = l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__7; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; +x_2 = l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__6; x_3 = l_Lean_Parser_Tactic_case___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -3027,13 +3057,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__9() { +static lean_object* _init_l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__2; +x_1 = l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__8; +x_3 = l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__7; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3041,15 +3071,24 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_tacticNext_______x3d_x3e__() { +static lean_object* _init_l_Lean_Parser_Tactic_tacticNext___x3d_x3e__() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__9; +x_1 = l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__8; return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__1() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__2() { _start: { lean_object* x_1; @@ -3057,17 +3096,17 @@ x_1 = lean_mk_string_from_bytes("null", 4); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__1; +x_2 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__2; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__3() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; @@ -3076,7 +3115,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__4() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__5() { _start: { lean_object* x_1; @@ -3084,26 +3123,7 @@ x_1 = lean_mk_string_from_bytes("=>", 2); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; -x_2 = l_Lean_Parser_Tactic_case___closed__9; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(1u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__7() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__6() { _start: { lean_object* x_1; lean_object* x_2; @@ -3112,11 +3132,11 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; uint8_t x_5; -x_4 = l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__2; +x_4 = l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__2; lean_inc(x_1); x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); if (x_5 == 0) @@ -3151,37 +3171,37 @@ lean_inc(x_15); x_17 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_17, 0, x_15); lean_ctor_set(x_17, 1, x_16); -x_18 = l_Lean_Parser_Tactic_intros___closed__10; +x_18 = l_Lean_binderIdent___closed__10; lean_inc(x_15); x_19 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_19, 0, x_15); lean_ctor_set(x_19, 1, x_18); -x_20 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__3; -x_21 = l_Array_appendCore___rarg(x_20, x_12); -lean_dec(x_12); +x_20 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; +x_21 = lean_array_push(x_20, x_19); x_22 = lean_box(2); -x_23 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_23 = l_Lean_binderIdent___closed__4; x_24 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_24, 0, x_22); lean_ctor_set(x_24, 1, x_23); lean_ctor_set(x_24, 2, x_21); -x_25 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__4; -x_26 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_26, 0, x_15); -lean_ctor_set(x_26, 1, x_25); -x_27 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; -x_28 = lean_array_push(x_27, x_11); -x_29 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__5; -x_30 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_30, 0, x_22); +x_25 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__4; +x_26 = l_Array_appendCore___rarg(x_25, x_12); +lean_dec(x_12); +x_27 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; +x_28 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_28, 0, x_22); +lean_ctor_set(x_28, 1, x_27); +lean_ctor_set(x_28, 2, x_26); +x_29 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__5; +x_30 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_30, 0, x_15); lean_ctor_set(x_30, 1, x_29); -lean_ctor_set(x_30, 2, x_28); -x_31 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__7; +x_31 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__6; x_32 = lean_array_push(x_31, x_17); -x_33 = lean_array_push(x_32, x_19); -x_34 = lean_array_push(x_33, x_24); -x_35 = lean_array_push(x_34, x_26); -x_36 = lean_array_push(x_35, x_30); +x_33 = lean_array_push(x_32, x_24); +x_34 = lean_array_push(x_33, x_28); +x_35 = lean_array_push(x_34, x_30); +x_36 = lean_array_push(x_35, x_11); x_37 = l_Lean_Parser_Tactic_case___closed__2; x_38 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_38, 0, x_22); @@ -3203,37 +3223,37 @@ lean_inc(x_39); x_42 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_42, 0, x_39); lean_ctor_set(x_42, 1, x_41); -x_43 = l_Lean_Parser_Tactic_intros___closed__10; +x_43 = l_Lean_binderIdent___closed__10; lean_inc(x_39); x_44 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_44, 0, x_39); lean_ctor_set(x_44, 1, x_43); -x_45 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__3; -x_46 = l_Array_appendCore___rarg(x_45, x_12); -lean_dec(x_12); +x_45 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; +x_46 = lean_array_push(x_45, x_44); x_47 = lean_box(2); -x_48 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_48 = l_Lean_binderIdent___closed__4; x_49 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_49, 0, x_47); lean_ctor_set(x_49, 1, x_48); lean_ctor_set(x_49, 2, x_46); -x_50 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__4; -x_51 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_51, 0, x_39); -lean_ctor_set(x_51, 1, x_50); -x_52 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; -x_53 = lean_array_push(x_52, x_11); -x_54 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__5; -x_55 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_55, 0, x_47); +x_50 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__4; +x_51 = l_Array_appendCore___rarg(x_50, x_12); +lean_dec(x_12); +x_52 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; +x_53 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_53, 0, x_47); +lean_ctor_set(x_53, 1, x_52); +lean_ctor_set(x_53, 2, x_51); +x_54 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__5; +x_55 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_55, 0, x_39); lean_ctor_set(x_55, 1, x_54); -lean_ctor_set(x_55, 2, x_53); -x_56 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__7; +x_56 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__6; x_57 = lean_array_push(x_56, x_42); -x_58 = lean_array_push(x_57, x_44); -x_59 = lean_array_push(x_58, x_49); -x_60 = lean_array_push(x_59, x_51); -x_61 = lean_array_push(x_60, x_55); +x_58 = lean_array_push(x_57, x_49); +x_59 = lean_array_push(x_58, x_53); +x_60 = lean_array_push(x_59, x_55); +x_61 = lean_array_push(x_60, x_11); x_62 = l_Lean_Parser_Tactic_case___closed__2; x_63 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_63, 0, x_47); @@ -3259,7 +3279,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_allGoals___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_allGoals___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -3289,7 +3309,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_allGoals___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_allGoals___closed__4; x_3 = l_Lean_Parser_Tactic_case___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); @@ -3333,7 +3353,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_anyGoals___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_anyGoals___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -3363,7 +3383,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_anyGoals___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_anyGoals___closed__4; x_3 = l_Lean_Parser_Tactic_case___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); @@ -3407,7 +3427,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_focus___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_focus___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -3437,7 +3457,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_focus___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_focus___closed__4; x_3 = l_Lean_Parser_Tactic_case___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); @@ -3481,7 +3501,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_skip___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_skip___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -3533,7 +3553,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_done___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_done___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -3585,7 +3605,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_traceState___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_traceState___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -3645,7 +3665,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_traceMessage___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_traceMessage___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -3703,7 +3723,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_traceMessage___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_traceMessage___closed__4; x_3 = l_Lean_Parser_Tactic_traceMessage___closed__7; x_4 = lean_alloc_ctor(2, 3, 0); @@ -3747,7 +3767,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_failIfSuccess___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_failIfSuccess___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -3777,7 +3797,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_failIfSuccess___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_failIfSuccess___closed__4; x_3 = l_Lean_Parser_Tactic_case___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); @@ -3821,7 +3841,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_paren___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_paren___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -3851,7 +3871,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_paren___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_paren___closed__4; x_3 = l_Lean_Parser_Tactic_case___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); @@ -3883,7 +3903,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_paren___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_paren___closed__5; x_3 = l_Lean_Parser_Tactic_paren___closed__7; x_4 = lean_alloc_ctor(2, 3, 0); @@ -3927,7 +3947,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_withReducible___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_withReducible___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -3957,7 +3977,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_withReducible___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_withReducible___closed__4; x_3 = l_Lean_Parser_Tactic_case___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); @@ -4001,7 +4021,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_withReducibleAndInstances___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_withReducibleAndInstances___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -4031,7 +4051,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_withReducibleAndInstances___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_withReducibleAndInstances___closed__4; x_3 = l_Lean_Parser_Tactic_case___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); @@ -4075,7 +4095,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_withUnfoldingAll___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_withUnfoldingAll___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -4105,7 +4125,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_withUnfoldingAll___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_withUnfoldingAll___closed__4; x_3 = l_Lean_Parser_Tactic_case___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); @@ -4149,7 +4169,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_first___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_first___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -4243,7 +4263,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_first___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_first___closed__11; x_3 = l_Lean_Parser_Tactic_intro___closed__8; x_4 = lean_alloc_ctor(2, 3, 0); @@ -4257,7 +4277,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_first___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_first___closed__12; x_3 = l_Lean_Parser_Tactic_case___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); @@ -4307,7 +4327,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_first___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_first___closed__4; x_3 = l_Lean_Parser_Tactic_first___closed__16; x_4 = lean_alloc_ctor(2, 3, 0); @@ -4351,7 +4371,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_rotateLeft___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_rotateLeft___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -4439,7 +4459,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_rotateLeft___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_rotateLeft___closed__4; x_3 = l_Lean_Parser_Tactic_rotateLeft___closed__10; x_4 = lean_alloc_ctor(2, 3, 0); @@ -4483,7 +4503,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_rotateRight___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_rotateRight___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -4513,7 +4533,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_rotateRight___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_rotateRight___closed__4; x_3 = l_Lean_Parser_Tactic_rotateLeft___closed__10; x_4 = lean_alloc_ctor(2, 3, 0); @@ -4557,7 +4577,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticTry_____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_tacticTry_____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -4587,7 +4607,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticTry_____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_tacticTry_____closed__4; x_3 = l_Lean_Parser_Tactic_case___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); @@ -4622,22 +4642,32 @@ return x_1; static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__1() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; +x_2 = l_Lean_Parser_Tactic_case___closed__9; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2() { +_start: +{ lean_object* x_1; x_1 = lean_mk_string_from_bytes("seq1", 4); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; -x_2 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__1; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; +x_2 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; @@ -4646,7 +4676,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__5() { _start: { lean_object* x_1; @@ -4654,23 +4684,23 @@ x_1 = lean_mk_string_from_bytes("tacticSeq1Indented", 18); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; -x_2 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; +x_2 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; -x_3 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__3; +x_2 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; +x_3 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__4; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -4706,7 +4736,7 @@ x_10 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRul x_11 = !lean_is_exclusive(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; 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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +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; 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; x_12 = lean_ctor_get(x_10, 0); x_13 = l_Lean_Parser_Tactic_first___closed__1; lean_inc(x_12); @@ -4718,191 +4748,181 @@ lean_inc(x_12); x_16 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_16, 0, x_12); lean_ctor_set(x_16, 1, x_15); -x_17 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; -x_18 = lean_array_push(x_17, x_9); -x_19 = lean_box(2); -x_20 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__5; -x_21 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -lean_ctor_set(x_21, 2, x_18); -x_22 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; -x_23 = lean_array_push(x_22, x_16); -lean_inc(x_23); -x_24 = lean_array_push(x_23, x_21); -x_25 = l_Lean_Parser_Tactic_first___closed__8; -x_26 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_26, 0, x_19); -lean_ctor_set(x_26, 1, x_25); -lean_ctor_set(x_26, 2, x_24); -x_27 = l_Lean_Parser_Tactic_skip___closed__1; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_12); +x_17 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; +x_18 = lean_array_push(x_17, x_16); +lean_inc(x_18); +x_19 = lean_array_push(x_18, x_9); +x_20 = lean_box(2); +x_21 = l_Lean_Parser_Tactic_first___closed__8; +x_22 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +lean_ctor_set(x_22, 2, x_19); +x_23 = l_Lean_Parser_Tactic_skip___closed__1; +x_24 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_24, 0, x_12); +lean_ctor_set(x_24, 1, x_23); +x_25 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; +x_26 = lean_array_push(x_25, x_24); +x_27 = l_Lean_Parser_Tactic_skip___closed__2; +x_28 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_28, 0, x_20); lean_ctor_set(x_28, 1, x_27); +lean_ctor_set(x_28, 2, x_26); x_29 = lean_array_push(x_17, x_28); -x_30 = l_Lean_Parser_Tactic_skip___closed__2; -x_31 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_31, 0, x_19); -lean_ctor_set(x_31, 1, x_30); -lean_ctor_set(x_31, 2, x_29); -x_32 = lean_array_push(x_22, x_31); -x_33 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; -x_34 = lean_array_push(x_32, x_33); +x_30 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__7; +x_31 = lean_array_push(x_29, x_30); +x_32 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_32, 0, x_20); +lean_ctor_set(x_32, 1, x_21); +lean_ctor_set(x_32, 2, x_31); +x_33 = lean_array_push(x_25, x_32); +x_34 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_19); -lean_ctor_set(x_35, 1, x_25); -lean_ctor_set(x_35, 2, x_34); -x_36 = lean_array_push(x_17, x_35); -x_37 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +lean_ctor_set(x_35, 0, x_20); +lean_ctor_set(x_35, 1, x_34); +lean_ctor_set(x_35, 2, x_33); +x_36 = lean_array_push(x_25, x_35); +x_37 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; x_38 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_38, 0, x_19); +lean_ctor_set(x_38, 0, x_20); lean_ctor_set(x_38, 1, x_37); lean_ctor_set(x_38, 2, x_36); -x_39 = lean_array_push(x_17, x_38); -x_40 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__5; +x_39 = lean_array_push(x_25, x_38); +x_40 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__1; x_41 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_41, 0, x_19); +lean_ctor_set(x_41, 0, x_20); lean_ctor_set(x_41, 1, x_40); lean_ctor_set(x_41, 2, x_39); -x_42 = lean_array_push(x_17, x_41); +x_42 = lean_array_push(x_18, x_41); x_43 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_43, 0, x_19); -lean_ctor_set(x_43, 1, x_20); +lean_ctor_set(x_43, 0, x_20); +lean_ctor_set(x_43, 1, x_21); lean_ctor_set(x_43, 2, x_42); -x_44 = lean_array_push(x_23, x_43); -x_45 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_45, 0, x_19); -lean_ctor_set(x_45, 1, x_25); -lean_ctor_set(x_45, 2, x_44); -x_46 = lean_array_push(x_22, x_26); -x_47 = lean_array_push(x_46, x_45); -x_48 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_48, 0, x_19); -lean_ctor_set(x_48, 1, x_37); -lean_ctor_set(x_48, 2, x_47); -x_49 = lean_array_push(x_22, x_14); -x_50 = lean_array_push(x_49, x_48); -x_51 = l_Lean_Parser_Tactic_first___closed__2; +x_44 = lean_array_push(x_17, x_22); +x_45 = lean_array_push(x_44, x_43); +x_46 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_46, 0, x_20); +lean_ctor_set(x_46, 1, x_34); +lean_ctor_set(x_46, 2, x_45); +x_47 = lean_array_push(x_17, x_14); +x_48 = lean_array_push(x_47, x_46); +x_49 = l_Lean_Parser_Tactic_first___closed__2; +x_50 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_50, 0, x_20); +lean_ctor_set(x_50, 1, x_49); +lean_ctor_set(x_50, 2, x_48); +x_51 = lean_array_push(x_25, x_50); x_52 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_52, 0, x_19); -lean_ctor_set(x_52, 1, x_51); -lean_ctor_set(x_52, 2, x_50); -x_53 = lean_array_push(x_17, x_52); -x_54 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_54, 0, x_19); -lean_ctor_set(x_54, 1, x_37); -lean_ctor_set(x_54, 2, x_53); -x_55 = lean_array_push(x_17, x_54); -x_56 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; -x_57 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_57, 0, x_19); -lean_ctor_set(x_57, 1, x_56); -lean_ctor_set(x_57, 2, x_55); -lean_ctor_set(x_10, 0, x_57); +lean_ctor_set(x_52, 0, x_20); +lean_ctor_set(x_52, 1, x_34); +lean_ctor_set(x_52, 2, x_51); +x_53 = lean_array_push(x_25, x_52); +x_54 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_55 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_55, 0, x_20); +lean_ctor_set(x_55, 1, x_54); +lean_ctor_set(x_55, 2, x_53); +lean_ctor_set(x_10, 0, x_55); return x_10; } else { -lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; -x_58 = lean_ctor_get(x_10, 0); -x_59 = lean_ctor_get(x_10, 1); -lean_inc(x_59); -lean_inc(x_58); +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_56 = lean_ctor_get(x_10, 0); +x_57 = lean_ctor_get(x_10, 1); +lean_inc(x_57); +lean_inc(x_56); lean_dec(x_10); -x_60 = l_Lean_Parser_Tactic_first___closed__1; -lean_inc(x_58); +x_58 = l_Lean_Parser_Tactic_first___closed__1; +lean_inc(x_56); +x_59 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_59, 0, x_56); +lean_ctor_set(x_59, 1, x_58); +x_60 = l_Lean_Parser_Tactic_intro___closed__7; +lean_inc(x_56); x_61 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_61, 0, x_58); +lean_ctor_set(x_61, 0, x_56); lean_ctor_set(x_61, 1, x_60); -x_62 = l_Lean_Parser_Tactic_intro___closed__7; -lean_inc(x_58); -x_63 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_63, 0, x_58); -lean_ctor_set(x_63, 1, x_62); -x_64 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; -x_65 = lean_array_push(x_64, x_9); -x_66 = lean_box(2); -x_67 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__5; -x_68 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_68, 0, x_66); -lean_ctor_set(x_68, 1, x_67); -lean_ctor_set(x_68, 2, x_65); -x_69 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; -x_70 = lean_array_push(x_69, x_63); -lean_inc(x_70); -x_71 = lean_array_push(x_70, x_68); -x_72 = l_Lean_Parser_Tactic_first___closed__8; +x_62 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; +x_63 = lean_array_push(x_62, x_61); +lean_inc(x_63); +x_64 = lean_array_push(x_63, x_9); +x_65 = lean_box(2); +x_66 = l_Lean_Parser_Tactic_first___closed__8; +x_67 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_67, 0, x_65); +lean_ctor_set(x_67, 1, x_66); +lean_ctor_set(x_67, 2, x_64); +x_68 = l_Lean_Parser_Tactic_skip___closed__1; +x_69 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_69, 0, x_56); +lean_ctor_set(x_69, 1, x_68); +x_70 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; +x_71 = lean_array_push(x_70, x_69); +x_72 = l_Lean_Parser_Tactic_skip___closed__2; x_73 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_73, 0, x_66); +lean_ctor_set(x_73, 0, x_65); lean_ctor_set(x_73, 1, x_72); lean_ctor_set(x_73, 2, x_71); -x_74 = l_Lean_Parser_Tactic_skip___closed__1; -x_75 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_75, 0, x_58); -lean_ctor_set(x_75, 1, x_74); -x_76 = lean_array_push(x_64, x_75); -x_77 = l_Lean_Parser_Tactic_skip___closed__2; -x_78 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_78, 0, x_66); -lean_ctor_set(x_78, 1, x_77); -lean_ctor_set(x_78, 2, x_76); -x_79 = lean_array_push(x_69, x_78); -x_80 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; -x_81 = lean_array_push(x_79, x_80); -x_82 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_82, 0, x_66); -lean_ctor_set(x_82, 1, x_72); -lean_ctor_set(x_82, 2, x_81); -x_83 = lean_array_push(x_64, x_82); -x_84 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; -x_85 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_85, 0, x_66); -lean_ctor_set(x_85, 1, x_84); -lean_ctor_set(x_85, 2, x_83); -x_86 = lean_array_push(x_64, x_85); -x_87 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__5; -x_88 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_88, 0, x_66); -lean_ctor_set(x_88, 1, x_87); -lean_ctor_set(x_88, 2, x_86); -x_89 = lean_array_push(x_64, x_88); -x_90 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_90, 0, x_66); -lean_ctor_set(x_90, 1, x_67); -lean_ctor_set(x_90, 2, x_89); -x_91 = lean_array_push(x_70, x_90); -x_92 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_92, 0, x_66); -lean_ctor_set(x_92, 1, x_72); -lean_ctor_set(x_92, 2, x_91); -x_93 = lean_array_push(x_69, x_73); -x_94 = lean_array_push(x_93, x_92); -x_95 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_95, 0, x_66); -lean_ctor_set(x_95, 1, x_84); -lean_ctor_set(x_95, 2, x_94); -x_96 = lean_array_push(x_69, x_61); -x_97 = lean_array_push(x_96, x_95); -x_98 = l_Lean_Parser_Tactic_first___closed__2; -x_99 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_99, 0, x_66); -lean_ctor_set(x_99, 1, x_98); -lean_ctor_set(x_99, 2, x_97); -x_100 = lean_array_push(x_64, x_99); -x_101 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_101, 0, x_66); -lean_ctor_set(x_101, 1, x_84); -lean_ctor_set(x_101, 2, x_100); -x_102 = lean_array_push(x_64, x_101); -x_103 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; -x_104 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_104, 0, x_66); -lean_ctor_set(x_104, 1, x_103); -lean_ctor_set(x_104, 2, x_102); -x_105 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_105, 0, x_104); -lean_ctor_set(x_105, 1, x_59); -return x_105; +x_74 = lean_array_push(x_62, x_73); +x_75 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__7; +x_76 = lean_array_push(x_74, x_75); +x_77 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_77, 0, x_65); +lean_ctor_set(x_77, 1, x_66); +lean_ctor_set(x_77, 2, x_76); +x_78 = lean_array_push(x_70, x_77); +x_79 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; +x_80 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_80, 0, x_65); +lean_ctor_set(x_80, 1, x_79); +lean_ctor_set(x_80, 2, x_78); +x_81 = lean_array_push(x_70, x_80); +x_82 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; +x_83 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_83, 0, x_65); +lean_ctor_set(x_83, 1, x_82); +lean_ctor_set(x_83, 2, x_81); +x_84 = lean_array_push(x_70, x_83); +x_85 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__1; +x_86 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_86, 0, x_65); +lean_ctor_set(x_86, 1, x_85); +lean_ctor_set(x_86, 2, x_84); +x_87 = lean_array_push(x_63, x_86); +x_88 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_88, 0, x_65); +lean_ctor_set(x_88, 1, x_66); +lean_ctor_set(x_88, 2, x_87); +x_89 = lean_array_push(x_62, x_67); +x_90 = lean_array_push(x_89, x_88); +x_91 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_91, 0, x_65); +lean_ctor_set(x_91, 1, x_79); +lean_ctor_set(x_91, 2, x_90); +x_92 = lean_array_push(x_62, x_59); +x_93 = lean_array_push(x_92, x_91); +x_94 = l_Lean_Parser_Tactic_first___closed__2; +x_95 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_95, 0, x_65); +lean_ctor_set(x_95, 1, x_94); +lean_ctor_set(x_95, 2, x_93); +x_96 = lean_array_push(x_70, x_95); +x_97 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_97, 0, x_65); +lean_ctor_set(x_97, 1, x_79); +lean_ctor_set(x_97, 2, x_96); +x_98 = lean_array_push(x_70, x_97); +x_99 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_100 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_100, 0, x_65); +lean_ctor_set(x_100, 1, x_99); +lean_ctor_set(x_100, 2, x_98); +x_101 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_101, 0, x_100); +lean_ctor_set(x_101, 1, x_57); +return x_101; } } } @@ -4919,7 +4939,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tactic___x3c_x3b_x3e_____closed__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_tactic___x3c_x3b_x3e_____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -4947,9 +4967,9 @@ static lean_object* _init_l_Lean_Parser_Tactic_tactic___x3c_x3b_x3e_____closed__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_tactic___x3c_x3b_x3e_____closed__4; -x_3 = l_Lean_Parser_Tactic_withAnnotateState___closed__23; +x_3 = l_Lean_Parser_Tactic_withAnnotateState___closed__21; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -5045,9 +5065,9 @@ lean_inc(x_16); x_18 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_18, 0, x_16); lean_ctor_set(x_18, 1, x_17); -x_19 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_19 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_20 = lean_array_push(x_19, x_9); -x_21 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; +x_21 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__7; x_22 = lean_array_push(x_20, x_21); x_23 = lean_box(2); x_24 = l_Lean_Parser_Tactic_first___closed__8; @@ -5065,7 +5085,7 @@ lean_inc(x_16); x_29 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_29, 0, x_16); lean_ctor_set(x_29, 1, x_28); -x_30 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_30 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_31 = lean_array_push(x_30, x_29); x_32 = l_Lean_Parser_Tactic_skip___closed__2; x_33 = lean_alloc_ctor(1, 3, 0); @@ -5076,7 +5096,7 @@ x_34 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__ x_35 = lean_array_push(x_34, x_27); x_36 = lean_array_push(x_35, x_11); x_37 = lean_array_push(x_36, x_33); -x_38 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; +x_38 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; x_39 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_39, 0, x_23); lean_ctor_set(x_39, 1, x_38); @@ -5098,19 +5118,19 @@ lean_ctor_set(x_47, 0, x_23); lean_ctor_set(x_47, 1, x_24); lean_ctor_set(x_47, 2, x_46); x_48 = lean_array_push(x_30, x_47); -x_49 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_49 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_50 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_50, 0, x_23); lean_ctor_set(x_50, 1, x_49); lean_ctor_set(x_50, 2, x_48); x_51 = lean_array_push(x_30, x_50); -x_52 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__5; +x_52 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; x_53 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_53, 0, x_23); lean_ctor_set(x_53, 1, x_52); lean_ctor_set(x_53, 2, x_51); x_54 = lean_array_push(x_30, x_53); -x_55 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__5; +x_55 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__1; x_56 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_56, 0, x_23); lean_ctor_set(x_56, 1, x_55); @@ -5168,9 +5188,9 @@ lean_inc(x_76); x_79 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_79, 0, x_76); lean_ctor_set(x_79, 1, x_78); -x_80 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_80 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_81 = lean_array_push(x_80, x_9); -x_82 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; +x_82 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__7; x_83 = lean_array_push(x_81, x_82); x_84 = lean_box(2); x_85 = l_Lean_Parser_Tactic_first___closed__8; @@ -5188,7 +5208,7 @@ lean_inc(x_76); x_90 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_90, 0, x_76); lean_ctor_set(x_90, 1, x_89); -x_91 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_91 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_92 = lean_array_push(x_91, x_90); x_93 = l_Lean_Parser_Tactic_skip___closed__2; x_94 = lean_alloc_ctor(1, 3, 0); @@ -5199,7 +5219,7 @@ x_95 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__ x_96 = lean_array_push(x_95, x_88); x_97 = lean_array_push(x_96, x_11); x_98 = lean_array_push(x_97, x_94); -x_99 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; +x_99 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; x_100 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_100, 0, x_84); lean_ctor_set(x_100, 1, x_99); @@ -5221,19 +5241,19 @@ lean_ctor_set(x_108, 0, x_84); lean_ctor_set(x_108, 1, x_85); lean_ctor_set(x_108, 2, x_107); x_109 = lean_array_push(x_91, x_108); -x_110 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_110 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_111 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_111, 0, x_84); lean_ctor_set(x_111, 1, x_110); lean_ctor_set(x_111, 2, x_109); x_112 = lean_array_push(x_91, x_111); -x_113 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__5; +x_113 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; x_114 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_114, 0, x_84); lean_ctor_set(x_114, 1, x_113); lean_ctor_set(x_114, 2, x_112); x_115 = lean_array_push(x_91, x_114); -x_116 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__5; +x_116 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__1; x_117 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_117, 0, x_84); lean_ctor_set(x_117, 1, x_116); @@ -5295,7 +5315,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_refl___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_refl___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -5355,7 +5375,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticRfl___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_tacticRfl___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -5432,7 +5452,7 @@ x_11 = l_Lean_Parser_Tactic_refl___closed__3; x_12 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_12, 0, x_10); lean_ctor_set(x_12, 1, x_11); -x_13 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_13 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_14 = lean_array_push(x_13, x_12); x_15 = lean_box(2); x_16 = l_Lean_Parser_Tactic_refl___closed__2; @@ -5441,13 +5461,13 @@ lean_ctor_set(x_17, 0, x_15); lean_ctor_set(x_17, 1, x_16); lean_ctor_set(x_17, 2, x_14); x_18 = lean_array_push(x_13, x_17); -x_19 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_19 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_20 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_20, 0, x_15); lean_ctor_set(x_20, 1, x_19); lean_ctor_set(x_20, 2, x_18); x_21 = lean_array_push(x_13, x_20); -x_22 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_22 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_23 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_23, 0, x_15); lean_ctor_set(x_23, 1, x_22); @@ -5467,7 +5487,7 @@ x_26 = l_Lean_Parser_Tactic_refl___closed__3; x_27 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_27, 0, x_24); lean_ctor_set(x_27, 1, x_26); -x_28 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_28 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_29 = lean_array_push(x_28, x_27); x_30 = lean_box(2); x_31 = l_Lean_Parser_Tactic_refl___closed__2; @@ -5476,13 +5496,13 @@ lean_ctor_set(x_32, 0, x_30); lean_ctor_set(x_32, 1, x_31); lean_ctor_set(x_32, 2, x_29); x_33 = lean_array_push(x_28, x_32); -x_34 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_34 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_35 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_35, 0, x_30); lean_ctor_set(x_35, 1, x_34); lean_ctor_set(x_35, 2, x_33); x_36 = lean_array_push(x_28, x_35); -x_37 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_37 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_38 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_38, 0, x_30); lean_ctor_set(x_38, 1, x_37); @@ -5507,7 +5527,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticRfl_x27___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_tacticRfl_x27___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -5567,7 +5587,7 @@ static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRu _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticRfl_x27__1___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -5703,7 +5723,7 @@ x_26 = l_Lean_Parser_Tactic_tacticRfl___closed__3; x_27 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_27, 0, x_10); lean_ctor_set(x_27, 1, x_26); -x_28 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_28 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_29 = lean_array_push(x_28, x_27); x_30 = lean_box(2); x_31 = l_Lean_Parser_Tactic_tacticRfl___closed__2; @@ -5711,9 +5731,9 @@ x_32 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_32, 0, x_30); lean_ctor_set(x_32, 1, x_31); lean_ctor_set(x_32, 2, x_29); -x_33 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_33 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_34 = lean_array_push(x_33, x_32); -x_35 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; +x_35 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__7; x_36 = lean_array_push(x_34, x_35); x_37 = l_Lean_Parser_Tactic_first___closed__8; x_38 = lean_alloc_ctor(1, 3, 0); @@ -5721,19 +5741,19 @@ lean_ctor_set(x_38, 0, x_30); lean_ctor_set(x_38, 1, x_37); lean_ctor_set(x_38, 2, x_36); x_39 = lean_array_push(x_28, x_38); -x_40 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_40 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_41 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_41, 0, x_30); lean_ctor_set(x_41, 1, x_40); lean_ctor_set(x_41, 2, x_39); x_42 = lean_array_push(x_28, x_41); -x_43 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__5; +x_43 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; x_44 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_44, 0, x_30); lean_ctor_set(x_44, 1, x_43); lean_ctor_set(x_44, 2, x_42); x_45 = lean_array_push(x_28, x_44); -x_46 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__5; +x_46 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__1; x_47 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_47, 0, x_30); lean_ctor_set(x_47, 1, x_46); @@ -5766,7 +5786,7 @@ x_60 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_60, 0, x_30); lean_ctor_set(x_60, 1, x_46); lean_ctor_set(x_60, 2, x_59); -x_61 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__7; +x_61 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__6; x_62 = lean_array_push(x_61, x_14); x_63 = lean_array_push(x_62, x_19); x_64 = lean_array_push(x_63, x_21); @@ -5783,7 +5803,7 @@ lean_ctor_set(x_70, 0, x_30); lean_ctor_set(x_70, 1, x_40); lean_ctor_set(x_70, 2, x_69); x_71 = lean_array_push(x_28, x_70); -x_72 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_72 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_73 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_73, 0, x_30); lean_ctor_set(x_73, 1, x_72); @@ -5838,7 +5858,7 @@ x_91 = l_Lean_Parser_Tactic_tacticRfl___closed__3; x_92 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_92, 0, x_74); lean_ctor_set(x_92, 1, x_91); -x_93 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_93 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_94 = lean_array_push(x_93, x_92); x_95 = lean_box(2); x_96 = l_Lean_Parser_Tactic_tacticRfl___closed__2; @@ -5846,9 +5866,9 @@ x_97 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_97, 0, x_95); lean_ctor_set(x_97, 1, x_96); lean_ctor_set(x_97, 2, x_94); -x_98 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_98 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_99 = lean_array_push(x_98, x_97); -x_100 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; +x_100 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__7; x_101 = lean_array_push(x_99, x_100); x_102 = l_Lean_Parser_Tactic_first___closed__8; x_103 = lean_alloc_ctor(1, 3, 0); @@ -5856,19 +5876,19 @@ lean_ctor_set(x_103, 0, x_95); lean_ctor_set(x_103, 1, x_102); lean_ctor_set(x_103, 2, x_101); x_104 = lean_array_push(x_93, x_103); -x_105 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_105 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_106 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_106, 0, x_95); lean_ctor_set(x_106, 1, x_105); lean_ctor_set(x_106, 2, x_104); x_107 = lean_array_push(x_93, x_106); -x_108 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__5; +x_108 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; x_109 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_109, 0, x_95); lean_ctor_set(x_109, 1, x_108); lean_ctor_set(x_109, 2, x_107); x_110 = lean_array_push(x_93, x_109); -x_111 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__5; +x_111 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__1; x_112 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_112, 0, x_95); lean_ctor_set(x_112, 1, x_111); @@ -5901,7 +5921,7 @@ x_125 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_125, 0, x_95); lean_ctor_set(x_125, 1, x_111); lean_ctor_set(x_125, 2, x_124); -x_126 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__7; +x_126 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__6; x_127 = lean_array_push(x_126, x_79); x_128 = lean_array_push(x_127, x_84); x_129 = lean_array_push(x_128, x_86); @@ -5918,7 +5938,7 @@ lean_ctor_set(x_135, 0, x_95); lean_ctor_set(x_135, 1, x_105); lean_ctor_set(x_135, 2, x_134); x_136 = lean_array_push(x_93, x_135); -x_137 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_137 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_138 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_138, 0, x_95); lean_ctor_set(x_138, 1, x_137); @@ -5943,7 +5963,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_ac__refl___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_ac__refl___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -6003,7 +6023,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticAdmit___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_tacticAdmit___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -6063,7 +6083,7 @@ static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRu _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__2; x_2 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticAdmit__1___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -6334,7 +6354,7 @@ lean_ctor_set(x_21, 0, x_10); lean_ctor_set(x_21, 1, x_19); lean_ctor_set(x_21, 2, x_18); lean_ctor_set(x_21, 3, x_20); -x_22 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_22 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_23 = lean_array_push(x_22, x_16); x_24 = lean_array_push(x_23, x_21); x_25 = lean_box(2); @@ -6343,12 +6363,12 @@ x_27 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_27, 0, x_25); lean_ctor_set(x_27, 1, x_26); lean_ctor_set(x_27, 2, x_24); -x_28 = l_Lean_Parser_Tactic_intros___closed__10; +x_28 = l_Lean_binderIdent___closed__10; lean_inc(x_10); x_29 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_29, 0, x_10); lean_ctor_set(x_29, 1, x_28); -x_30 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_30 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_31 = lean_array_push(x_30, x_29); x_32 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticAdmit__1___closed__15; x_33 = lean_alloc_ctor(1, 3, 0); @@ -6366,7 +6386,7 @@ lean_ctor_set(x_38, 2, x_35); lean_ctor_set(x_38, 3, x_37); x_39 = lean_array_push(x_22, x_33); x_40 = lean_array_push(x_39, x_38); -x_41 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_41 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_42 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_42, 0, x_25); lean_ctor_set(x_42, 1, x_41); @@ -6391,7 +6411,7 @@ lean_ctor_set(x_52, 0, x_25); lean_ctor_set(x_52, 1, x_41); lean_ctor_set(x_52, 2, x_51); x_53 = lean_array_push(x_30, x_52); -x_54 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_54 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_55 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_55, 0, x_25); lean_ctor_set(x_55, 1, x_54); @@ -6434,7 +6454,7 @@ lean_ctor_set(x_68, 0, x_56); lean_ctor_set(x_68, 1, x_66); lean_ctor_set(x_68, 2, x_65); lean_ctor_set(x_68, 3, x_67); -x_69 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_69 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_70 = lean_array_push(x_69, x_63); x_71 = lean_array_push(x_70, x_68); x_72 = lean_box(2); @@ -6443,12 +6463,12 @@ x_74 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_74, 0, x_72); lean_ctor_set(x_74, 1, x_73); lean_ctor_set(x_74, 2, x_71); -x_75 = l_Lean_Parser_Tactic_intros___closed__10; +x_75 = l_Lean_binderIdent___closed__10; lean_inc(x_56); x_76 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_76, 0, x_56); lean_ctor_set(x_76, 1, x_75); -x_77 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_77 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_78 = lean_array_push(x_77, x_76); x_79 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticAdmit__1___closed__15; x_80 = lean_alloc_ctor(1, 3, 0); @@ -6466,7 +6486,7 @@ lean_ctor_set(x_85, 2, x_82); lean_ctor_set(x_85, 3, x_84); x_86 = lean_array_push(x_69, x_80); x_87 = lean_array_push(x_86, x_85); -x_88 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_88 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_89 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_89, 0, x_72); lean_ctor_set(x_89, 1, x_88); @@ -6491,7 +6511,7 @@ lean_ctor_set(x_99, 0, x_72); lean_ctor_set(x_99, 1, x_88); lean_ctor_set(x_99, 2, x_98); x_100 = lean_array_push(x_77, x_99); -x_101 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_101 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_102 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_102, 0, x_72); lean_ctor_set(x_102, 1, x_101); @@ -6516,7 +6536,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticSorry___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_tacticSorry___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -6617,7 +6637,7 @@ lean_ctor_set(x_21, 0, x_10); lean_ctor_set(x_21, 1, x_19); lean_ctor_set(x_21, 2, x_18); lean_ctor_set(x_21, 3, x_20); -x_22 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_22 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_23 = lean_array_push(x_22, x_16); x_24 = lean_array_push(x_23, x_21); x_25 = lean_box(2); @@ -6626,12 +6646,12 @@ x_27 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_27, 0, x_25); lean_ctor_set(x_27, 1, x_26); lean_ctor_set(x_27, 2, x_24); -x_28 = l_Lean_Parser_Tactic_intros___closed__10; +x_28 = l_Lean_binderIdent___closed__10; lean_inc(x_10); x_29 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_29, 0, x_10); lean_ctor_set(x_29, 1, x_28); -x_30 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_30 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_31 = lean_array_push(x_30, x_29); x_32 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticAdmit__1___closed__15; x_33 = lean_alloc_ctor(1, 3, 0); @@ -6649,7 +6669,7 @@ lean_ctor_set(x_38, 2, x_35); lean_ctor_set(x_38, 3, x_37); x_39 = lean_array_push(x_22, x_33); x_40 = lean_array_push(x_39, x_38); -x_41 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_41 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_42 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_42, 0, x_25); lean_ctor_set(x_42, 1, x_41); @@ -6674,7 +6694,7 @@ lean_ctor_set(x_52, 0, x_25); lean_ctor_set(x_52, 1, x_41); lean_ctor_set(x_52, 2, x_51); x_53 = lean_array_push(x_30, x_52); -x_54 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_54 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_55 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_55, 0, x_25); lean_ctor_set(x_55, 1, x_54); @@ -6717,7 +6737,7 @@ lean_ctor_set(x_68, 0, x_56); lean_ctor_set(x_68, 1, x_66); lean_ctor_set(x_68, 2, x_65); lean_ctor_set(x_68, 3, x_67); -x_69 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_69 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_70 = lean_array_push(x_69, x_63); x_71 = lean_array_push(x_70, x_68); x_72 = lean_box(2); @@ -6726,12 +6746,12 @@ x_74 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_74, 0, x_72); lean_ctor_set(x_74, 1, x_73); lean_ctor_set(x_74, 2, x_71); -x_75 = l_Lean_Parser_Tactic_intros___closed__10; +x_75 = l_Lean_binderIdent___closed__10; lean_inc(x_56); x_76 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_76, 0, x_56); lean_ctor_set(x_76, 1, x_75); -x_77 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_77 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_78 = lean_array_push(x_77, x_76); x_79 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticAdmit__1___closed__15; x_80 = lean_alloc_ctor(1, 3, 0); @@ -6749,7 +6769,7 @@ lean_ctor_set(x_85, 2, x_82); lean_ctor_set(x_85, 3, x_84); x_86 = lean_array_push(x_69, x_80); x_87 = lean_array_push(x_86, x_85); -x_88 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_88 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_89 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_89, 0, x_72); lean_ctor_set(x_89, 1, x_88); @@ -6774,7 +6794,7 @@ lean_ctor_set(x_99, 0, x_72); lean_ctor_set(x_99, 1, x_88); lean_ctor_set(x_99, 2, x_98); x_100 = lean_array_push(x_77, x_99); -x_101 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_101 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_102 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_102, 0, x_72); lean_ctor_set(x_102, 1, x_101); @@ -6799,7 +6819,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticInfer__instance___closed__2 _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_tacticInfer__instance___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -6957,7 +6977,7 @@ lean_ctor_set(x_19, 0, x_10); lean_ctor_set(x_19, 1, x_17); lean_ctor_set(x_19, 2, x_16); lean_ctor_set(x_19, 3, x_18); -x_20 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_20 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_21 = lean_array_push(x_20, x_14); x_22 = lean_array_push(x_21, x_19); x_23 = lean_box(2); @@ -6966,15 +6986,15 @@ x_25 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_25, 0, x_23); lean_ctor_set(x_25, 1, x_24); lean_ctor_set(x_25, 2, x_22); -x_26 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_26 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_27 = lean_array_push(x_26, x_25); -x_28 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_28 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_29 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_29, 0, x_23); lean_ctor_set(x_29, 1, x_28); lean_ctor_set(x_29, 2, x_27); x_30 = lean_array_push(x_26, x_29); -x_31 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_31 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_32 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_32, 0, x_23); lean_ctor_set(x_32, 1, x_31); @@ -7009,7 +7029,7 @@ lean_ctor_set(x_43, 0, x_33); lean_ctor_set(x_43, 1, x_41); lean_ctor_set(x_43, 2, x_40); lean_ctor_set(x_43, 3, x_42); -x_44 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_44 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_45 = lean_array_push(x_44, x_38); x_46 = lean_array_push(x_45, x_43); x_47 = lean_box(2); @@ -7018,15 +7038,15 @@ x_49 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_49, 0, x_47); lean_ctor_set(x_49, 1, x_48); lean_ctor_set(x_49, 2, x_46); -x_50 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_50 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_51 = lean_array_push(x_50, x_49); -x_52 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_52 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_53 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_53, 0, x_47); lean_ctor_set(x_53, 1, x_52); lean_ctor_set(x_53, 2, x_51); x_54 = lean_array_push(x_50, x_53); -x_55 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_55 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_56 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_56, 0, x_47); lean_ctor_set(x_56, 1, x_55); @@ -7051,7 +7071,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_config___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_config___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -7101,7 +7121,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_config___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_config___closed__5; x_3 = l_Lean_Parser_Tactic_config___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); @@ -7145,7 +7165,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_config___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_config___closed__8; x_3 = l_Lean_Parser_Tactic_config___closed__10; x_4 = lean_alloc_ctor(2, 3, 0); @@ -7159,7 +7179,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_config___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_config___closed__11; x_3 = l_Lean_Parser_Tactic_rename___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -7173,7 +7193,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_config___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_config___closed__12; x_3 = l_Lean_Parser_Tactic_paren___closed__7; x_4 = lean_alloc_ctor(2, 3, 0); @@ -7217,7 +7237,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_locationWildcard___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_locationWildcard___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -7275,7 +7295,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_locationHyp___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_locationHyp___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -7321,7 +7341,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_locationHyp___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_intros___closed__6; +x_1 = l_Lean_binderIdent___closed__6; x_2 = l_Lean_Parser_Tactic_locationHyp___closed__4; x_3 = l_Lean_Parser_Tactic_locationHyp___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); @@ -7347,7 +7367,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_locationHyp___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_revert___closed__7; x_3 = l_Lean_Parser_Tactic_locationHyp___closed__8; x_4 = lean_alloc_ctor(2, 3, 0); @@ -7391,7 +7411,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_location___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_location___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -7419,7 +7439,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_location___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_intros___closed__6; +x_1 = l_Lean_binderIdent___closed__6; x_2 = l_Lean_Parser_Tactic_locationWildcard; x_3 = l_Lean_Parser_Tactic_locationHyp; x_4 = lean_alloc_ctor(2, 3, 0); @@ -7433,7 +7453,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_location___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_location___closed__4; x_3 = l_Lean_Parser_Tactic_location___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -7489,7 +7509,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_change___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_change___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -7519,7 +7539,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_change___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_change___closed__4; x_3 = l_Lean_Parser_Tactic_rename___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -7545,7 +7565,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_change___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_change___closed__5; x_3 = l_Lean_Parser_Tactic_change___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); @@ -7589,7 +7609,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_changeWith___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_changeWith___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -7617,7 +7637,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_changeWith___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_change___closed__5; x_3 = l_Lean_Parser_Tactic_changeWith___closed__4; x_4 = lean_alloc_ctor(2, 3, 0); @@ -7631,7 +7651,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_changeWith___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_changeWith___closed__5; x_3 = l_Lean_Parser_Tactic_rename___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -7645,7 +7665,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_changeWith___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_changeWith___closed__6; x_3 = l_Lean_Parser_Tactic_change___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); @@ -7689,7 +7709,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_rwRule___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_rwRule___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -7735,7 +7755,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_rwRule___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_intros___closed__6; +x_1 = l_Lean_binderIdent___closed__6; x_2 = l_Lean_Parser_Tactic_rwRule___closed__4; x_3 = l_Lean_Parser_Tactic_rwRule___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); @@ -7761,7 +7781,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_rwRule___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_rwRule___closed__8; x_3 = l_Lean_Parser_Tactic_rename___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -7805,7 +7825,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_rwRuleSeq___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_rwRuleSeq___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -7875,7 +7895,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_rwRuleSeq___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_rwRuleSeq___closed__4; x_3 = l_Lean_Parser_Tactic_rwRuleSeq___closed__8; x_4 = lean_alloc_ctor(2, 3, 0); @@ -7907,7 +7927,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_rwRuleSeq___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_rwRuleSeq___closed__9; x_3 = l_Lean_Parser_Tactic_rwRuleSeq___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); @@ -7951,7 +7971,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_rewriteSeq___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_rewriteSeq___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -7993,7 +8013,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_rewriteSeq___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_rewriteSeq___closed__4; x_3 = l_Lean_Parser_Tactic_rewriteSeq___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -8007,7 +8027,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_rewriteSeq___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_rewriteSeq___closed__6; x_3 = l_Lean_Parser_Tactic_rwRuleSeq; x_4 = lean_alloc_ctor(2, 3, 0); @@ -8021,7 +8041,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_rewriteSeq___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_rewriteSeq___closed__7; x_3 = l_Lean_Parser_Tactic_change___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); @@ -8065,7 +8085,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_rwSeq___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_rwSeq___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -8095,7 +8115,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_rwSeq___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_rwSeq___closed__4; x_3 = l_Lean_Parser_Tactic_rewriteSeq___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -8109,7 +8129,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_rwSeq___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_rwSeq___closed__5; x_3 = l_Lean_Parser_Tactic_rwRuleSeq; x_4 = lean_alloc_ctor(2, 3, 0); @@ -8123,7 +8143,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_rwSeq___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_rwSeq___closed__6; x_3 = l_Lean_Parser_Tactic_change___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); @@ -8155,19 +8175,16 @@ x_1 = l_Lean_Parser_Tactic_rwSeq___closed__8; return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_rwWithRfl___closed__1() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(2); -x_2 = l_Lean_Parser_Tactic_rwRuleSeq___closed__10; -x_3 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(4u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_rwWithRfl___closed__2() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__2() { _start: { lean_object* x_1; @@ -8175,7 +8192,7 @@ x_1 = lean_mk_string_from_bytes(";", 1); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_rwWithRfl___closed__3() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__3() { _start: { lean_object* x_1; @@ -8183,7 +8200,7 @@ x_1 = lean_mk_string_from_bytes("try", 3); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_rwWithRfl___closed__4() { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__4() { _start: { lean_object* x_1; @@ -8191,223 +8208,452 @@ x_1 = lean_mk_string_from_bytes("with_reducible", 14); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_rwWithRfl(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__5() { _start: { -lean_object* x_6; lean_object* x_7; 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; 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; 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_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; -x_6 = lean_unsigned_to_nat(2u); -x_7 = l_Lean_Syntax_getArg(x_3, x_6); -x_8 = l_Lean_Syntax_getArg(x_7, x_6); -x_9 = l_Lean_Parser_Tactic_rwWithRfl___closed__1; -x_10 = l_Lean_Syntax_setArg(x_7, x_6, x_9); -lean_inc(x_3); -x_11 = l_Lean_Syntax_setKind(x_3, x_1); -x_12 = l_Lean_mkAtomFrom(x_3, x_2); -x_13 = lean_unsigned_to_nat(0u); -x_14 = l_Lean_Syntax_setArg(x_11, x_13, x_12); -x_15 = l_Lean_Syntax_setArg(x_14, x_6, x_10); -x_16 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_4, x_5); -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -if (lean_is_exclusive(x_16)) { - lean_ctor_release(x_16, 0); - lean_ctor_release(x_16, 1); - x_19 = x_16; -} else { - lean_dec_ref(x_16); - x_19 = lean_box(0); +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("rewrite", 7); +return x_1; } -x_20 = l_Lean_Parser_Tactic_rwWithRfl___closed__2; -lean_inc(x_17); -x_21 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_21, 0, x_17); -lean_ctor_set(x_21, 1, x_20); -x_22 = l_Lean_Parser_Tactic_rwWithRfl___closed__3; -lean_inc(x_17); -x_23 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_23, 0, x_17); -lean_ctor_set(x_23, 1, x_22); -x_24 = l_Lean_Parser_Tactic_paren___closed__3; -lean_inc(x_17); -x_25 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_25, 0, x_17); -lean_ctor_set(x_25, 1, x_24); -x_26 = l_Lean_Parser_Tactic_rwWithRfl___closed__4; -lean_inc(x_17); -x_27 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_27, 0, x_17); -lean_ctor_set(x_27, 1, x_26); -x_28 = l_Lean_Syntax_getHeadInfo_x3f(x_8); -lean_dec(x_8); -x_29 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; -x_30 = lean_array_push(x_29, x_27); -x_31 = l_Lean_Parser_Tactic_paren___closed__6; -lean_inc(x_17); -x_32 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_32, 0, x_17); -lean_ctor_set(x_32, 1, x_31); -x_33 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tactic___x3c_x3b_x3e____1___closed__2; -x_34 = lean_array_push(x_33, x_25); -x_35 = lean_array_push(x_29, x_23); -x_36 = lean_array_push(x_33, x_15); -x_37 = lean_array_push(x_36, x_21); -if (lean_obj_tag(x_28) == 0) +} +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = l_Lean_Parser_Tactic_rwSeq___closed__2; +lean_inc(x_1); +x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; +lean_dec(x_2); +lean_dec(x_1); +x_6 = lean_box(1); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_3); +return x_7; +} +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; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_object* x_20; +x_8 = lean_unsigned_to_nat(0u); +x_9 = l_Lean_Syntax_getArg(x_1, x_8); +x_10 = lean_unsigned_to_nat(1u); +x_11 = l_Lean_Syntax_getArg(x_1, x_10); +x_12 = lean_unsigned_to_nat(2u); +x_13 = l_Lean_Syntax_getArg(x_1, x_12); +x_14 = lean_unsigned_to_nat(3u); +x_15 = l_Lean_Syntax_getArg(x_1, x_14); +lean_dec(x_1); +x_16 = l_Lean_Syntax_getOptional_x3f(x_15); +lean_dec(x_15); +x_17 = l_Lean_Syntax_getOptional_x3f(x_11); +lean_dec(x_11); +x_18 = l_Lean_Parser_Tactic_rwRuleSeq___closed__2; +lean_inc(x_13); +x_19 = l_Lean_Syntax_isOfKind(x_13, x_18); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_154; +x_154 = lean_box(0); +x_20 = x_154; +goto block_153; +} +else +{ +uint8_t x_155; +x_155 = !lean_is_exclusive(x_16); +if (x_155 == 0) +{ +x_20 = x_16; +goto block_153; +} +else +{ +lean_object* x_156; lean_object* x_157; +x_156 = lean_ctor_get(x_16, 0); +lean_inc(x_156); +lean_dec(x_16); +x_157 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_157, 0, x_156); +x_20 = x_157; +goto block_153; +} +} +block_153: +{ +lean_object* x_21; +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_149; +x_149 = lean_box(0); +x_21 = x_149; +goto block_148; +} +else +{ +uint8_t x_150; +x_150 = !lean_is_exclusive(x_17); +if (x_150 == 0) { -x_38 = x_17; -goto block_94; +x_21 = x_17; +goto block_148; } else { -lean_object* x_95; +lean_object* x_151; lean_object* x_152; +x_151 = lean_ctor_get(x_17, 0); +lean_inc(x_151); lean_dec(x_17); -x_95 = lean_ctor_get(x_28, 0); -lean_inc(x_95); -lean_dec(x_28); -x_38 = x_95; -goto block_94; +x_152 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_152, 0, x_151); +x_21 = x_152; +goto block_148; +} +} +block_148: +{ +if (x_19 == 0) +{ +lean_object* x_22; lean_object* x_23; +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_13); +lean_dec(x_9); +lean_dec(x_2); +x_22 = lean_box(1); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_3); +return x_23; } -block_94: +else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_39 = l_Lean_Parser_Tactic_tacticRfl___closed__3; +lean_object* 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_24 = l_Lean_Syntax_getArg(x_13, x_8); +x_25 = l_Lean_Syntax_getArg(x_13, x_10); +x_26 = l_Lean_Syntax_getArg(x_13, x_12); +lean_dec(x_13); +x_27 = l_Lean_Syntax_getArgs(x_25); +lean_dec(x_25); +x_28 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +if (lean_is_exclusive(x_28)) { + lean_ctor_release(x_28, 0); + lean_ctor_release(x_28, 1); + x_31 = x_28; +} else { + lean_dec_ref(x_28); + x_31 = lean_box(0); +} +x_32 = l_Lean_Syntax_getHeadInfo_x3f(x_9); +lean_dec(x_9); +x_33 = l_Lean_Syntax_getHeadInfo_x3f(x_24); +lean_dec(x_24); +x_34 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__4; +x_35 = l_Array_appendCore___rarg(x_34, x_27); +lean_dec(x_27); +x_36 = lean_box(2); +x_37 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; +x_38 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +lean_ctor_set(x_38, 2, x_35); +x_39 = l_Lean_Parser_Tactic_rwRuleSeq___closed__10; +lean_inc(x_29); x_40 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_40, 0, x_38); +lean_ctor_set(x_40, 0, x_29); lean_ctor_set(x_40, 1, x_39); -x_41 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; -x_42 = lean_array_push(x_41, x_40); -x_43 = lean_box(2); -x_44 = l_Lean_Parser_Tactic_tacticRfl___closed__2; -x_45 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_45, 0, x_43); -lean_ctor_set(x_45, 1, x_44); -lean_ctor_set(x_45, 2, x_42); -x_46 = lean_array_push(x_29, x_45); -x_47 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; -x_48 = lean_array_push(x_46, x_47); -x_49 = l_Lean_Parser_Tactic_first___closed__8; -x_50 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_50, 0, x_43); -lean_ctor_set(x_50, 1, x_49); -lean_ctor_set(x_50, 2, x_48); -x_51 = lean_array_push(x_41, x_50); -x_52 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; -x_53 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_53, 0, x_43); +x_41 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__2; +lean_inc(x_29); +x_42 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_42, 0, x_29); +lean_ctor_set(x_42, 1, x_41); +x_43 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__3; +lean_inc(x_29); +x_44 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_44, 0, x_29); +lean_ctor_set(x_44, 1, x_43); +x_45 = l_Lean_Parser_Tactic_paren___closed__3; +lean_inc(x_29); +x_46 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_46, 0, x_29); +lean_ctor_set(x_46, 1, x_45); +x_47 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__4; +lean_inc(x_29); +x_48 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_48, 0, x_29); +lean_ctor_set(x_48, 1, x_47); +x_49 = l_Lean_Syntax_getHeadInfo_x3f(x_26); +lean_dec(x_26); +x_50 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; +x_51 = lean_array_push(x_50, x_48); +x_52 = l_Lean_Parser_Tactic_paren___closed__6; +lean_inc(x_29); +x_53 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_53, 0, x_29); lean_ctor_set(x_53, 1, x_52); -lean_ctor_set(x_53, 2, x_51); -x_54 = lean_array_push(x_41, x_53); -x_55 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__5; -x_56 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_56, 0, x_43); -lean_ctor_set(x_56, 1, x_55); -lean_ctor_set(x_56, 2, x_54); -x_57 = lean_array_push(x_41, x_56); -x_58 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__5; -x_59 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_59, 0, x_43); +x_54 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tactic___x3c_x3b_x3e____1___closed__2; +x_55 = lean_array_push(x_54, x_46); +x_56 = lean_array_push(x_50, x_44); +if (lean_obj_tag(x_32) == 0) +{ +lean_inc(x_29); +x_57 = x_29; +goto block_146; +} +else +{ +lean_object* x_147; +x_147 = lean_ctor_get(x_32, 0); +lean_inc(x_147); +lean_dec(x_32); +x_57 = x_147; +goto block_146; +} +block_146: +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_58 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__5; +x_59 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_59, 0, x_57); lean_ctor_set(x_59, 1, x_58); -lean_ctor_set(x_59, 2, x_57); -x_60 = lean_array_push(x_30, x_59); -x_61 = l_Lean_Parser_Tactic_withReducible___closed__2; -x_62 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_62, 0, x_43); -lean_ctor_set(x_62, 1, x_61); -lean_ctor_set(x_62, 2, x_60); -x_63 = lean_array_push(x_29, x_62); -x_64 = lean_array_push(x_63, x_47); -x_65 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_65, 0, x_43); -lean_ctor_set(x_65, 1, x_49); -lean_ctor_set(x_65, 2, x_64); -x_66 = lean_array_push(x_41, x_65); -x_67 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_67, 0, x_43); -lean_ctor_set(x_67, 1, x_52); -lean_ctor_set(x_67, 2, x_66); -x_68 = lean_array_push(x_41, x_67); -x_69 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_69, 0, x_43); -lean_ctor_set(x_69, 1, x_55); -lean_ctor_set(x_69, 2, x_68); -x_70 = lean_array_push(x_41, x_69); -x_71 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_71, 0, x_43); -lean_ctor_set(x_71, 1, x_58); -lean_ctor_set(x_71, 2, x_70); -x_72 = lean_array_push(x_34, x_71); -x_73 = lean_array_push(x_72, x_32); -x_74 = l_Lean_Parser_Tactic_paren___closed__2; -x_75 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_75, 0, x_43); -lean_ctor_set(x_75, 1, x_74); -lean_ctor_set(x_75, 2, x_73); -x_76 = lean_array_push(x_29, x_75); -x_77 = lean_array_push(x_76, x_47); -x_78 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_78, 0, x_43); -lean_ctor_set(x_78, 1, x_49); -lean_ctor_set(x_78, 2, x_77); -x_79 = lean_array_push(x_41, x_78); -x_80 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_80, 0, x_43); -lean_ctor_set(x_80, 1, x_52); -lean_ctor_set(x_80, 2, x_79); -x_81 = lean_array_push(x_41, x_80); -x_82 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_82, 0, x_43); -lean_ctor_set(x_82, 1, x_55); -lean_ctor_set(x_82, 2, x_81); -x_83 = lean_array_push(x_41, x_82); -x_84 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_84, 0, x_43); -lean_ctor_set(x_84, 1, x_58); -lean_ctor_set(x_84, 2, x_83); -x_85 = lean_array_push(x_35, x_84); -x_86 = l_Lean_Parser_Tactic_tacticTry_____closed__2; -x_87 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_87, 0, x_43); -lean_ctor_set(x_87, 1, x_86); -lean_ctor_set(x_87, 2, x_85); -x_88 = lean_array_push(x_37, x_87); -x_89 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_89, 0, x_43); -lean_ctor_set(x_89, 1, x_52); -lean_ctor_set(x_89, 2, x_88); -x_90 = lean_array_push(x_41, x_89); -x_91 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; -x_92 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_92, 0, x_43); -lean_ctor_set(x_92, 1, x_91); -lean_ctor_set(x_92, 2, x_90); -if (lean_is_scalar(x_19)) { - x_93 = lean_alloc_ctor(0, 2, 0); +x_60 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__1; +x_61 = lean_array_push(x_60, x_59); +if (lean_obj_tag(x_21) == 0) +{ +x_62 = x_34; +goto block_143; +} +else +{ +lean_object* x_144; lean_object* x_145; +x_144 = lean_ctor_get(x_21, 0); +lean_inc(x_144); +lean_dec(x_21); +x_145 = lean_array_push(x_34, x_144); +x_62 = x_145; +goto block_143; +} +block_143: +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_63 = l_Array_appendCore___rarg(x_34, x_62); +lean_dec(x_62); +x_64 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_64, 0, x_36); +lean_ctor_set(x_64, 1, x_37); +lean_ctor_set(x_64, 2, x_63); +x_65 = lean_array_push(x_61, x_64); +if (lean_obj_tag(x_33) == 0) +{ +lean_inc(x_29); +x_66 = x_29; +goto block_141; +} +else +{ +lean_object* x_142; +x_142 = lean_ctor_get(x_33, 0); +lean_inc(x_142); +lean_dec(x_33); +x_66 = x_142; +goto block_141; +} +block_141: +{ +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; +x_67 = l_Lean_Parser_Tactic_rwRuleSeq___closed__3; +x_68 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_68, 0, x_66); +lean_ctor_set(x_68, 1, x_67); +x_69 = lean_array_push(x_54, x_68); +x_70 = lean_array_push(x_69, x_38); +x_71 = lean_array_push(x_70, x_40); +x_72 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_72, 0, x_36); +lean_ctor_set(x_72, 1, x_18); +lean_ctor_set(x_72, 2, x_71); +x_73 = lean_array_push(x_65, x_72); +if (lean_obj_tag(x_20) == 0) +{ +x_74 = x_34; +goto block_138; +} +else +{ +lean_object* x_139; lean_object* x_140; +x_139 = lean_ctor_get(x_20, 0); +lean_inc(x_139); +lean_dec(x_20); +x_140 = lean_array_push(x_34, x_139); +x_74 = x_140; +goto block_138; +} +block_138: +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_75 = l_Array_appendCore___rarg(x_34, x_74); +lean_dec(x_74); +x_76 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_76, 0, x_36); +lean_ctor_set(x_76, 1, x_37); +lean_ctor_set(x_76, 2, x_75); +x_77 = lean_array_push(x_73, x_76); +x_78 = l_Lean_Parser_Tactic_rewriteSeq___closed__2; +x_79 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_79, 0, x_36); +lean_ctor_set(x_79, 1, x_78); +lean_ctor_set(x_79, 2, x_77); +x_80 = lean_array_push(x_54, x_79); +x_81 = lean_array_push(x_80, x_42); +if (lean_obj_tag(x_49) == 0) +{ +x_82 = x_29; +goto block_136; +} +else +{ +lean_object* x_137; +lean_dec(x_29); +x_137 = lean_ctor_get(x_49, 0); +lean_inc(x_137); +lean_dec(x_49); +x_82 = x_137; +goto block_136; +} +block_136: +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; +x_83 = l_Lean_Parser_Tactic_tacticRfl___closed__3; +x_84 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_84, 0, x_82); +lean_ctor_set(x_84, 1, x_83); +x_85 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; +x_86 = lean_array_push(x_85, x_84); +x_87 = l_Lean_Parser_Tactic_tacticRfl___closed__2; +x_88 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_88, 0, x_36); +lean_ctor_set(x_88, 1, x_87); +lean_ctor_set(x_88, 2, x_86); +x_89 = lean_array_push(x_50, x_88); +x_90 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__7; +x_91 = lean_array_push(x_89, x_90); +x_92 = l_Lean_Parser_Tactic_first___closed__8; +x_93 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_93, 0, x_36); +lean_ctor_set(x_93, 1, x_92); +lean_ctor_set(x_93, 2, x_91); +x_94 = lean_array_push(x_85, x_93); +x_95 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_95, 0, x_36); +lean_ctor_set(x_95, 1, x_37); +lean_ctor_set(x_95, 2, x_94); +x_96 = lean_array_push(x_85, x_95); +x_97 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; +x_98 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_98, 0, x_36); +lean_ctor_set(x_98, 1, x_97); +lean_ctor_set(x_98, 2, x_96); +x_99 = lean_array_push(x_85, x_98); +x_100 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__1; +x_101 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_101, 0, x_36); +lean_ctor_set(x_101, 1, x_100); +lean_ctor_set(x_101, 2, x_99); +x_102 = lean_array_push(x_51, x_101); +x_103 = l_Lean_Parser_Tactic_withReducible___closed__2; +x_104 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_104, 0, x_36); +lean_ctor_set(x_104, 1, x_103); +lean_ctor_set(x_104, 2, x_102); +x_105 = lean_array_push(x_50, x_104); +x_106 = lean_array_push(x_105, x_90); +x_107 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_107, 0, x_36); +lean_ctor_set(x_107, 1, x_92); +lean_ctor_set(x_107, 2, x_106); +x_108 = lean_array_push(x_85, x_107); +x_109 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_109, 0, x_36); +lean_ctor_set(x_109, 1, x_37); +lean_ctor_set(x_109, 2, x_108); +x_110 = lean_array_push(x_85, x_109); +x_111 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_111, 0, x_36); +lean_ctor_set(x_111, 1, x_97); +lean_ctor_set(x_111, 2, x_110); +x_112 = lean_array_push(x_85, x_111); +x_113 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_113, 0, x_36); +lean_ctor_set(x_113, 1, x_100); +lean_ctor_set(x_113, 2, x_112); +x_114 = lean_array_push(x_55, x_113); +x_115 = lean_array_push(x_114, x_53); +x_116 = l_Lean_Parser_Tactic_paren___closed__2; +x_117 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_117, 0, x_36); +lean_ctor_set(x_117, 1, x_116); +lean_ctor_set(x_117, 2, x_115); +x_118 = lean_array_push(x_50, x_117); +x_119 = lean_array_push(x_118, x_90); +x_120 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_120, 0, x_36); +lean_ctor_set(x_120, 1, x_92); +lean_ctor_set(x_120, 2, x_119); +x_121 = lean_array_push(x_85, x_120); +x_122 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_122, 0, x_36); +lean_ctor_set(x_122, 1, x_37); +lean_ctor_set(x_122, 2, x_121); +x_123 = lean_array_push(x_85, x_122); +x_124 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_124, 0, x_36); +lean_ctor_set(x_124, 1, x_97); +lean_ctor_set(x_124, 2, x_123); +x_125 = lean_array_push(x_85, x_124); +x_126 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_126, 0, x_36); +lean_ctor_set(x_126, 1, x_100); +lean_ctor_set(x_126, 2, x_125); +x_127 = lean_array_push(x_56, x_126); +x_128 = l_Lean_Parser_Tactic_tacticTry_____closed__2; +x_129 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_129, 0, x_36); +lean_ctor_set(x_129, 1, x_128); +lean_ctor_set(x_129, 2, x_127); +x_130 = lean_array_push(x_81, x_129); +x_131 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_131, 0, x_36); +lean_ctor_set(x_131, 1, x_37); +lean_ctor_set(x_131, 2, x_130); +x_132 = lean_array_push(x_85, x_131); +x_133 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_134 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_134, 0, x_36); +lean_ctor_set(x_134, 1, x_133); +lean_ctor_set(x_134, 2, x_132); +if (lean_is_scalar(x_31)) { + x_135 = lean_alloc_ctor(0, 2, 0); } else { - x_93 = x_19; + x_135 = x_31; +} +lean_ctor_set(x_135, 0, x_134); +lean_ctor_set(x_135, 1, x_30); +return x_135; +} +} +} } -lean_ctor_set(x_93, 0, x_92); -lean_ctor_set(x_93, 1, x_18); -return x_93; } } } -static lean_object* _init_l_Lean_Parser_Tactic_expandRwSeq___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("rewrite", 7); -return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandRwSeq(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_4 = l_Lean_Parser_Tactic_rewriteSeq___closed__2; -x_5 = l_Lean_Parser_Tactic_expandRwSeq___closed__1; -x_6 = l_Lean_Parser_Tactic_rwWithRfl(x_4, x_5, x_1, x_2, x_3); -return x_6; } } static lean_object* _init_l_Lean_Parser_Tactic_injection___closed__1() { @@ -8422,7 +8668,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_injection___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_injection___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -8452,7 +8698,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_injection___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_injection___closed__4; x_3 = l_Lean_Parser_Tactic_rename___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -8467,7 +8713,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_injection___closed__6() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_revert___closed__6; -x_2 = l_Lean_Parser_Tactic_intros___closed__13; +x_2 = l_Lean_Parser_Tactic_intros___closed__5; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -8478,7 +8724,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_injection___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_changeWith___closed__4; x_3 = l_Lean_Parser_Tactic_injection___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); @@ -8504,7 +8750,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_injection___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_injection___closed__5; x_3 = l_Lean_Parser_Tactic_injection___closed__8; x_4 = lean_alloc_ctor(2, 3, 0); @@ -8548,7 +8794,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_injections___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_injections___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -8600,7 +8846,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_discharger___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_discharger___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -8642,7 +8888,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_discharger___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_intros___closed__6; +x_1 = l_Lean_binderIdent___closed__6; x_2 = l_Lean_Parser_Tactic_discharger___closed__3; x_3 = l_Lean_Parser_Tactic_discharger___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -8656,7 +8902,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_discharger___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_config___closed__5; x_3 = l_Lean_Parser_Tactic_discharger___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); @@ -8682,7 +8928,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_discharger___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_discharger___closed__8; x_3 = l_Lean_Parser_Tactic_config___closed__10; x_4 = lean_alloc_ctor(2, 3, 0); @@ -8696,7 +8942,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_discharger___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_discharger___closed__9; x_3 = l_Lean_Parser_Tactic_case___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); @@ -8710,7 +8956,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_discharger___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_discharger___closed__10; x_3 = l_Lean_Parser_Tactic_paren___closed__7; x_4 = lean_alloc_ctor(2, 3, 0); @@ -8754,7 +9000,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_simpPre___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_simpPre___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -8812,7 +9058,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_simpPost___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_simpPost___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -8870,7 +9116,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_simpLemma___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_simpLemma___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -8880,7 +9126,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_simpLemma___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_intros___closed__6; +x_1 = l_Lean_binderIdent___closed__6; x_2 = l_Lean_Parser_Tactic_simpPre; x_3 = l_Lean_Parser_Tactic_simpPost; x_4 = lean_alloc_ctor(2, 3, 0); @@ -8906,7 +9152,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_simpLemma___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_simpLemma___closed__4; x_3 = l_Lean_Parser_Tactic_rwRule___closed__8; x_4 = lean_alloc_ctor(2, 3, 0); @@ -8920,7 +9166,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_simpLemma___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_simpLemma___closed__5; x_3 = l_Lean_Parser_Tactic_rename___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -8964,7 +9210,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_simpErase___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_simpErase___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -8992,7 +9238,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_simpErase___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_simpErase___closed__4; x_3 = l_Lean_Parser_Tactic_intro___closed__18; x_4 = lean_alloc_ctor(2, 3, 0); @@ -9036,7 +9282,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_simpStar___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_simpStar___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -9076,7 +9322,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_simp___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_simp___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -9106,7 +9352,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_simp___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_simp___closed__4; x_3 = l_Lean_Parser_Tactic_rewriteSeq___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -9132,7 +9378,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_simp___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_simp___closed__5; x_3 = l_Lean_Parser_Tactic_simp___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); @@ -9178,7 +9424,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_simp___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_simp___closed__7; x_3 = l_Lean_Parser_Tactic_simp___closed__10; x_4 = lean_alloc_ctor(2, 3, 0); @@ -9192,7 +9438,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_simp___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_intros___closed__6; +x_1 = l_Lean_binderIdent___closed__6; x_2 = l_Lean_Parser_Tactic_simpErase; x_3 = l_Lean_Parser_Tactic_simpLemma; x_4 = lean_alloc_ctor(2, 3, 0); @@ -9206,7 +9452,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_simp___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_intros___closed__6; +x_1 = l_Lean_binderIdent___closed__6; x_2 = l_Lean_Parser_Tactic_simpStar; x_3 = l_Lean_Parser_Tactic_simp___closed__12; x_4 = lean_alloc_ctor(2, 3, 0); @@ -9236,7 +9482,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_simp___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_rwRuleSeq___closed__4; x_3 = l_Lean_Parser_Tactic_simp___closed__14; x_4 = lean_alloc_ctor(2, 3, 0); @@ -9250,7 +9496,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_simp___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_simp___closed__15; x_3 = l_Lean_Parser_Tactic_rwRuleSeq___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); @@ -9276,7 +9522,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_simp___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_simp___closed__11; x_3 = l_Lean_Parser_Tactic_simp___closed__17; x_4 = lean_alloc_ctor(2, 3, 0); @@ -9290,7 +9536,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_simp___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_simp___closed__18; x_3 = l_Lean_Parser_Tactic_change___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); @@ -9334,7 +9580,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_simpAll___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_simpAll___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -9364,7 +9610,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_simpAll___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_simpAll___closed__4; x_3 = l_Lean_Parser_Tactic_rewriteSeq___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -9378,7 +9624,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_simpAll___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_simpAll___closed__5; x_3 = l_Lean_Parser_Tactic_simp___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); @@ -9392,7 +9638,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_simpAll___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_simpAll___closed__6; x_3 = l_Lean_Parser_Tactic_simp___closed__10; x_4 = lean_alloc_ctor(2, 3, 0); @@ -9422,7 +9668,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_simpAll___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_rwRuleSeq___closed__4; x_3 = l_Lean_Parser_Tactic_simpAll___closed__8; x_4 = lean_alloc_ctor(2, 3, 0); @@ -9436,7 +9682,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_simpAll___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_simpAll___closed__9; x_3 = l_Lean_Parser_Tactic_rwRuleSeq___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); @@ -9462,7 +9708,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_simpAll___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_simpAll___closed__7; x_3 = l_Lean_Parser_Tactic_simpAll___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); @@ -9506,7 +9752,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_dsimp___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_dsimp___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -9536,7 +9782,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_dsimp___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_dsimp___closed__4; x_3 = l_Lean_Parser_Tactic_rewriteSeq___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -9550,7 +9796,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_dsimp___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_dsimp___closed__5; x_3 = l_Lean_Parser_Tactic_simp___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); @@ -9564,7 +9810,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_dsimp___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_dsimp___closed__6; x_3 = l_Lean_Parser_Tactic_simp___closed__10; x_4 = lean_alloc_ctor(2, 3, 0); @@ -9578,7 +9824,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_dsimp___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_dsimp___closed__7; x_3 = l_Lean_Parser_Tactic_simpAll___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); @@ -9592,7 +9838,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_dsimp___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_dsimp___closed__8; x_3 = l_Lean_Parser_Tactic_change___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); @@ -9636,7 +9882,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_delta___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_delta___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -9666,9 +9912,9 @@ static lean_object* _init_l_Lean_Parser_Tactic_delta___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_delta___closed__4; -x_3 = l_Lean_Parser_Tactic_intros___closed__9; +x_3 = l_Lean_binderIdent___closed__9; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -9680,7 +9926,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_delta___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_delta___closed__5; x_3 = l_Lean_Parser_Tactic_change___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); @@ -9724,7 +9970,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_unfold___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_unfold___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -9754,7 +10000,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_unfold___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; uint8_t x_4; lean_object* x_5; -x_1 = l_Lean_Parser_Tactic_intros___closed__9; +x_1 = l_Lean_binderIdent___closed__9; x_2 = l_Lean_Parser_Tactic_rwRuleSeq___closed__7; x_3 = l_Lean_Parser_Tactic_rwRuleSeq___closed__6; x_4 = 0; @@ -9770,7 +10016,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_unfold___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_unfold___closed__4; x_3 = l_Lean_Parser_Tactic_unfold___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -9784,7 +10030,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_unfold___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_unfold___closed__6; x_3 = l_Lean_Parser_Tactic_change___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); @@ -9828,7 +10074,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticRefine__lift_____closed__2( _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_tacticRefine__lift_____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -9858,7 +10104,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticRefine__lift_____closed__5( _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_tacticRefine__lift_____closed__4; x_3 = l_Lean_Parser_Tactic_rename___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -9966,7 +10212,7 @@ lean_inc(x_12); x_20 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_20, 0, x_12); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_21 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_22 = lean_array_push(x_21, x_20); x_23 = lean_array_push(x_22, x_9); x_24 = lean_box(2); @@ -9982,14 +10228,14 @@ x_30 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_30, 0, x_24); lean_ctor_set(x_30, 1, x_29); lean_ctor_set(x_30, 2, x_28); -x_31 = l_Lean_Parser_Tactic_rwWithRfl___closed__2; +x_31 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__2; lean_inc(x_12); x_32 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_32, 0, x_12); lean_ctor_set(x_32, 1, x_31); -x_33 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_33 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_34 = lean_array_push(x_33, x_32); -x_35 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_35 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_36 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_36, 0, x_24); lean_ctor_set(x_36, 1, x_35); @@ -10007,7 +10253,7 @@ x_42 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_42, 0, x_12); lean_ctor_set(x_42, 1, x_41); x_43 = lean_array_push(x_21, x_42); -x_44 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; +x_44 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__7; x_45 = lean_array_push(x_43, x_44); x_46 = l_Lean_Parser_Tactic_rotateRight___closed__2; x_47 = lean_alloc_ctor(1, 3, 0); @@ -10027,13 +10273,13 @@ lean_ctor_set(x_53, 0, x_24); lean_ctor_set(x_53, 1, x_35); lean_ctor_set(x_53, 2, x_52); x_54 = lean_array_push(x_33, x_53); -x_55 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__5; +x_55 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; x_56 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_56, 0, x_24); lean_ctor_set(x_56, 1, x_55); lean_ctor_set(x_56, 2, x_54); x_57 = lean_array_push(x_33, x_56); -x_58 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__5; +x_58 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__1; x_59 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_59, 0, x_24); lean_ctor_set(x_59, 1, x_58); @@ -10085,7 +10331,7 @@ lean_ctor_set(x_82, 0, x_24); lean_ctor_set(x_82, 1, x_35); lean_ctor_set(x_82, 2, x_81); x_83 = lean_array_push(x_33, x_82); -x_84 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_84 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_85 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_85, 0, x_24); lean_ctor_set(x_85, 1, x_84); @@ -10121,7 +10367,7 @@ lean_inc(x_86); x_95 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_95, 0, x_86); lean_ctor_set(x_95, 1, x_94); -x_96 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_96 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_97 = lean_array_push(x_96, x_95); x_98 = lean_array_push(x_97, x_9); x_99 = lean_box(2); @@ -10137,14 +10383,14 @@ x_105 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_105, 0, x_99); lean_ctor_set(x_105, 1, x_104); lean_ctor_set(x_105, 2, x_103); -x_106 = l_Lean_Parser_Tactic_rwWithRfl___closed__2; +x_106 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__2; lean_inc(x_86); x_107 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_107, 0, x_86); lean_ctor_set(x_107, 1, x_106); -x_108 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_108 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_109 = lean_array_push(x_108, x_107); -x_110 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_110 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_111 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_111, 0, x_99); lean_ctor_set(x_111, 1, x_110); @@ -10162,7 +10408,7 @@ x_117 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_117, 0, x_86); lean_ctor_set(x_117, 1, x_116); x_118 = lean_array_push(x_96, x_117); -x_119 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; +x_119 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__7; x_120 = lean_array_push(x_118, x_119); x_121 = l_Lean_Parser_Tactic_rotateRight___closed__2; x_122 = lean_alloc_ctor(1, 3, 0); @@ -10182,13 +10428,13 @@ lean_ctor_set(x_128, 0, x_99); lean_ctor_set(x_128, 1, x_110); lean_ctor_set(x_128, 2, x_127); x_129 = lean_array_push(x_108, x_128); -x_130 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__5; +x_130 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; x_131 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_131, 0, x_99); lean_ctor_set(x_131, 1, x_130); lean_ctor_set(x_131, 2, x_129); x_132 = lean_array_push(x_108, x_131); -x_133 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__5; +x_133 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__1; x_134 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_134, 0, x_99); lean_ctor_set(x_134, 1, x_133); @@ -10240,7 +10486,7 @@ lean_ctor_set(x_157, 0, x_99); lean_ctor_set(x_157, 1, x_110); lean_ctor_set(x_157, 2, x_156); x_158 = lean_array_push(x_108, x_157); -x_159 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_159 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_160 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_160, 0, x_99); lean_ctor_set(x_160, 1, x_159); @@ -10265,7 +10511,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticHave_____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_tacticHave_____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -10323,7 +10569,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticHave_____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_tacticHave_____closed__4; x_3 = l_Lean_Parser_Tactic_tacticHave_____closed__7; x_4 = lean_alloc_ctor(2, 3, 0); @@ -10407,15 +10653,6 @@ x_1 = lean_mk_string_from_bytes("?", 1); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__7() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(4u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; -} -} LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { @@ -10456,7 +10693,7 @@ lean_inc(x_12); x_16 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_16, 0, x_12); lean_ctor_set(x_16, 1, x_15); -x_17 = l_Lean_Parser_Tactic_rwWithRfl___closed__2; +x_17 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__2; lean_inc(x_12); x_18 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_18, 0, x_12); @@ -10466,11 +10703,11 @@ lean_inc(x_12); x_20 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_20, 0, x_12); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_Parser_Tactic_intros___closed__10; +x_21 = l_Lean_binderIdent___closed__10; x_22 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_22, 0, x_12); lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_23 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_24 = lean_array_push(x_23, x_22); x_25 = lean_box(2); x_26 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticAdmit__1___closed__15; @@ -10478,7 +10715,7 @@ x_27 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_27, 0, x_25); lean_ctor_set(x_27, 1, x_26); lean_ctor_set(x_27, 2, x_24); -x_28 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_28 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_29 = lean_array_push(x_28, x_20); x_30 = lean_array_push(x_29, x_27); x_31 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__5; @@ -10486,7 +10723,7 @@ x_32 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_32, 0, x_25); lean_ctor_set(x_32, 1, x_31); lean_ctor_set(x_32, 2, x_30); -x_33 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__7; +x_33 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__1; x_34 = lean_array_push(x_33, x_16); x_35 = lean_array_push(x_34, x_9); x_36 = lean_array_push(x_35, x_18); @@ -10504,13 +10741,13 @@ lean_ctor_set(x_43, 0, x_25); lean_ctor_set(x_43, 1, x_42); lean_ctor_set(x_43, 2, x_41); x_44 = lean_array_push(x_23, x_43); -x_45 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_45 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_46 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_46, 0, x_25); lean_ctor_set(x_46, 1, x_45); lean_ctor_set(x_46, 2, x_44); x_47 = lean_array_push(x_23, x_46); -x_48 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_48 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_49 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_49, 0, x_25); lean_ctor_set(x_49, 1, x_48); @@ -10536,7 +10773,7 @@ lean_inc(x_50); x_55 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_55, 0, x_50); lean_ctor_set(x_55, 1, x_54); -x_56 = l_Lean_Parser_Tactic_rwWithRfl___closed__2; +x_56 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__2; lean_inc(x_50); x_57 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_57, 0, x_50); @@ -10546,11 +10783,11 @@ lean_inc(x_50); x_59 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_59, 0, x_50); lean_ctor_set(x_59, 1, x_58); -x_60 = l_Lean_Parser_Tactic_intros___closed__10; +x_60 = l_Lean_binderIdent___closed__10; x_61 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_61, 0, x_50); lean_ctor_set(x_61, 1, x_60); -x_62 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_62 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_63 = lean_array_push(x_62, x_61); x_64 = lean_box(2); x_65 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticAdmit__1___closed__15; @@ -10558,7 +10795,7 @@ x_66 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_66, 0, x_64); lean_ctor_set(x_66, 1, x_65); lean_ctor_set(x_66, 2, x_63); -x_67 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_67 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_68 = lean_array_push(x_67, x_59); x_69 = lean_array_push(x_68, x_66); x_70 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__5; @@ -10566,7 +10803,7 @@ x_71 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_71, 0, x_64); lean_ctor_set(x_71, 1, x_70); lean_ctor_set(x_71, 2, x_69); -x_72 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__7; +x_72 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__1; x_73 = lean_array_push(x_72, x_55); x_74 = lean_array_push(x_73, x_9); x_75 = lean_array_push(x_74, x_57); @@ -10584,13 +10821,13 @@ lean_ctor_set(x_82, 0, x_64); lean_ctor_set(x_82, 1, x_81); lean_ctor_set(x_82, 2, x_80); x_83 = lean_array_push(x_62, x_82); -x_84 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_84 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_85 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_85, 0, x_64); lean_ctor_set(x_85, 1, x_84); lean_ctor_set(x_85, 2, x_83); x_86 = lean_array_push(x_62, x_85); -x_87 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_87 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_88 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_88, 0, x_64); lean_ctor_set(x_88, 1, x_87); @@ -10615,7 +10852,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticHave_____x3a_x3d_____closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_tacticHave_____x3a_x3d_____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -10637,9 +10874,9 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticHave_____x3a_x3d_____closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_tacticHave_____x3a_x3d_____closed__3; -x_3 = l_Lean_Parser_Tactic_intros___closed__9; +x_3 = l_Lean_binderIdent___closed__9; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -10651,7 +10888,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticHave_____x3a_x3d_____closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_tacticHave_____x3a_x3d_____closed__4; x_3 = l_Lean_Parser_Tactic_config___closed__10; x_4 = lean_alloc_ctor(2, 3, 0); @@ -10665,7 +10902,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticHave_____x3a_x3d_____closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_tacticHave_____x3a_x3d_____closed__5; x_3 = l_Lean_Parser_Tactic_rename___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -10796,12 +11033,12 @@ lean_inc(x_14); x_16 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_16, 0, x_14); lean_ctor_set(x_16, 1, x_15); -x_17 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_17 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_18 = lean_array_push(x_17, x_9); -x_19 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; +x_19 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__7; x_20 = lean_array_push(x_18, x_19); x_21 = lean_box(2); -x_22 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_22 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_23 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_23, 0, x_21); lean_ctor_set(x_23, 1, x_22); @@ -10811,12 +11048,12 @@ lean_inc(x_14); x_25 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_25, 0, x_14); lean_ctor_set(x_25, 1, x_24); -x_26 = l_Lean_Parser_Tactic_intros___closed__10; +x_26 = l_Lean_binderIdent___closed__10; lean_inc(x_14); x_27 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_27, 0, x_14); lean_ctor_set(x_27, 1, x_26); -x_28 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_28 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_29 = lean_array_push(x_28, x_27); x_30 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticAdmit__1___closed__15; x_31 = lean_alloc_ctor(1, 3, 0); @@ -10839,7 +11076,7 @@ x_38 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__ x_39 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_39, 0, x_14); lean_ctor_set(x_39, 1, x_38); -x_40 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__7; +x_40 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__1; x_41 = lean_array_push(x_40, x_23); x_42 = lean_array_push(x_41, x_37); x_43 = lean_array_push(x_42, x_39); @@ -10868,7 +11105,7 @@ lean_ctor_set(x_55, 0, x_21); lean_ctor_set(x_55, 1, x_22); lean_ctor_set(x_55, 2, x_54); x_56 = lean_array_push(x_28, x_55); -x_57 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_57 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_58 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_58, 0, x_21); lean_ctor_set(x_58, 1, x_57); @@ -10889,12 +11126,12 @@ lean_inc(x_59); x_62 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_62, 0, x_59); lean_ctor_set(x_62, 1, x_61); -x_63 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_63 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_64 = lean_array_push(x_63, x_9); -x_65 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; +x_65 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__7; x_66 = lean_array_push(x_64, x_65); x_67 = lean_box(2); -x_68 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_68 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_69 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_69, 0, x_67); lean_ctor_set(x_69, 1, x_68); @@ -10904,12 +11141,12 @@ lean_inc(x_59); x_71 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_71, 0, x_59); lean_ctor_set(x_71, 1, x_70); -x_72 = l_Lean_Parser_Tactic_intros___closed__10; +x_72 = l_Lean_binderIdent___closed__10; lean_inc(x_59); x_73 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_73, 0, x_59); lean_ctor_set(x_73, 1, x_72); -x_74 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_74 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_75 = lean_array_push(x_74, x_73); x_76 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticAdmit__1___closed__15; x_77 = lean_alloc_ctor(1, 3, 0); @@ -10932,7 +11169,7 @@ x_84 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__ x_85 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_85, 0, x_59); lean_ctor_set(x_85, 1, x_84); -x_86 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__7; +x_86 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__1; x_87 = lean_array_push(x_86, x_69); x_88 = lean_array_push(x_87, x_83); x_89 = lean_array_push(x_88, x_85); @@ -10961,7 +11198,7 @@ lean_ctor_set(x_101, 0, x_67); lean_ctor_set(x_101, 1, x_68); lean_ctor_set(x_101, 2, x_100); x_102 = lean_array_push(x_74, x_101); -x_103 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_103 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_104 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_104, 0, x_67); lean_ctor_set(x_104, 1, x_103); @@ -10986,7 +11223,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticSuffices_____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_tacticSuffices_____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -11044,7 +11281,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticSuffices_____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_tacticSuffices_____closed__4; x_3 = l_Lean_Parser_Tactic_tacticSuffices_____closed__7; x_4 = lean_alloc_ctor(2, 3, 0); @@ -11134,7 +11371,7 @@ lean_inc(x_12); x_16 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_16, 0, x_12); lean_ctor_set(x_16, 1, x_15); -x_17 = l_Lean_Parser_Tactic_rwWithRfl___closed__2; +x_17 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__2; lean_inc(x_12); x_18 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_18, 0, x_12); @@ -11144,11 +11381,11 @@ lean_inc(x_12); x_20 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_20, 0, x_12); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_Parser_Tactic_intros___closed__10; +x_21 = l_Lean_binderIdent___closed__10; x_22 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_22, 0, x_12); lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_23 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_24 = lean_array_push(x_23, x_22); x_25 = lean_box(2); x_26 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticAdmit__1___closed__15; @@ -11156,7 +11393,7 @@ x_27 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_27, 0, x_25); lean_ctor_set(x_27, 1, x_26); lean_ctor_set(x_27, 2, x_24); -x_28 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_28 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_29 = lean_array_push(x_28, x_20); x_30 = lean_array_push(x_29, x_27); x_31 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__5; @@ -11164,7 +11401,7 @@ x_32 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_32, 0, x_25); lean_ctor_set(x_32, 1, x_31); lean_ctor_set(x_32, 2, x_30); -x_33 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__7; +x_33 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__1; x_34 = lean_array_push(x_33, x_16); x_35 = lean_array_push(x_34, x_9); x_36 = lean_array_push(x_35, x_18); @@ -11182,13 +11419,13 @@ lean_ctor_set(x_43, 0, x_25); lean_ctor_set(x_43, 1, x_42); lean_ctor_set(x_43, 2, x_41); x_44 = lean_array_push(x_23, x_43); -x_45 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_45 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_46 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_46, 0, x_25); lean_ctor_set(x_46, 1, x_45); lean_ctor_set(x_46, 2, x_44); x_47 = lean_array_push(x_23, x_46); -x_48 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_48 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_49 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_49, 0, x_25); lean_ctor_set(x_49, 1, x_48); @@ -11214,7 +11451,7 @@ lean_inc(x_50); x_55 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_55, 0, x_50); lean_ctor_set(x_55, 1, x_54); -x_56 = l_Lean_Parser_Tactic_rwWithRfl___closed__2; +x_56 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__2; lean_inc(x_50); x_57 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_57, 0, x_50); @@ -11224,11 +11461,11 @@ lean_inc(x_50); x_59 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_59, 0, x_50); lean_ctor_set(x_59, 1, x_58); -x_60 = l_Lean_Parser_Tactic_intros___closed__10; +x_60 = l_Lean_binderIdent___closed__10; x_61 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_61, 0, x_50); lean_ctor_set(x_61, 1, x_60); -x_62 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_62 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_63 = lean_array_push(x_62, x_61); x_64 = lean_box(2); x_65 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticAdmit__1___closed__15; @@ -11236,7 +11473,7 @@ x_66 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_66, 0, x_64); lean_ctor_set(x_66, 1, x_65); lean_ctor_set(x_66, 2, x_63); -x_67 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_67 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_68 = lean_array_push(x_67, x_59); x_69 = lean_array_push(x_68, x_66); x_70 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__5; @@ -11244,7 +11481,7 @@ x_71 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_71, 0, x_64); lean_ctor_set(x_71, 1, x_70); lean_ctor_set(x_71, 2, x_69); -x_72 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__7; +x_72 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__1; x_73 = lean_array_push(x_72, x_55); x_74 = lean_array_push(x_73, x_9); x_75 = lean_array_push(x_74, x_57); @@ -11262,13 +11499,13 @@ lean_ctor_set(x_82, 0, x_64); lean_ctor_set(x_82, 1, x_81); lean_ctor_set(x_82, 2, x_80); x_83 = lean_array_push(x_62, x_82); -x_84 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_84 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_85 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_85, 0, x_64); lean_ctor_set(x_85, 1, x_84); lean_ctor_set(x_85, 2, x_83); x_86 = lean_array_push(x_62, x_85); -x_87 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_87 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_88 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_88, 0, x_64); lean_ctor_set(x_88, 1, x_87); @@ -11293,7 +11530,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticLet_____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_tacticLet_____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -11351,7 +11588,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticLet_____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_tacticLet_____closed__4; x_3 = l_Lean_Parser_Tactic_tacticLet_____closed__7; x_4 = lean_alloc_ctor(2, 3, 0); @@ -11441,7 +11678,7 @@ lean_inc(x_12); x_16 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_16, 0, x_12); lean_ctor_set(x_16, 1, x_15); -x_17 = l_Lean_Parser_Tactic_rwWithRfl___closed__2; +x_17 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__2; lean_inc(x_12); x_18 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_18, 0, x_12); @@ -11451,11 +11688,11 @@ lean_inc(x_12); x_20 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_20, 0, x_12); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_Parser_Tactic_intros___closed__10; +x_21 = l_Lean_binderIdent___closed__10; x_22 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_22, 0, x_12); lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_23 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_24 = lean_array_push(x_23, x_22); x_25 = lean_box(2); x_26 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticAdmit__1___closed__15; @@ -11463,7 +11700,7 @@ x_27 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_27, 0, x_25); lean_ctor_set(x_27, 1, x_26); lean_ctor_set(x_27, 2, x_24); -x_28 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_28 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_29 = lean_array_push(x_28, x_20); x_30 = lean_array_push(x_29, x_27); x_31 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__5; @@ -11471,7 +11708,7 @@ x_32 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_32, 0, x_25); lean_ctor_set(x_32, 1, x_31); lean_ctor_set(x_32, 2, x_30); -x_33 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__7; +x_33 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__1; x_34 = lean_array_push(x_33, x_16); x_35 = lean_array_push(x_34, x_9); x_36 = lean_array_push(x_35, x_18); @@ -11489,13 +11726,13 @@ lean_ctor_set(x_43, 0, x_25); lean_ctor_set(x_43, 1, x_42); lean_ctor_set(x_43, 2, x_41); x_44 = lean_array_push(x_23, x_43); -x_45 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_45 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_46 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_46, 0, x_25); lean_ctor_set(x_46, 1, x_45); lean_ctor_set(x_46, 2, x_44); x_47 = lean_array_push(x_23, x_46); -x_48 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_48 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_49 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_49, 0, x_25); lean_ctor_set(x_49, 1, x_48); @@ -11521,7 +11758,7 @@ lean_inc(x_50); x_55 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_55, 0, x_50); lean_ctor_set(x_55, 1, x_54); -x_56 = l_Lean_Parser_Tactic_rwWithRfl___closed__2; +x_56 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__2; lean_inc(x_50); x_57 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_57, 0, x_50); @@ -11531,11 +11768,11 @@ lean_inc(x_50); x_59 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_59, 0, x_50); lean_ctor_set(x_59, 1, x_58); -x_60 = l_Lean_Parser_Tactic_intros___closed__10; +x_60 = l_Lean_binderIdent___closed__10; x_61 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_61, 0, x_50); lean_ctor_set(x_61, 1, x_60); -x_62 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_62 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_63 = lean_array_push(x_62, x_61); x_64 = lean_box(2); x_65 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticAdmit__1___closed__15; @@ -11543,7 +11780,7 @@ x_66 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_66, 0, x_64); lean_ctor_set(x_66, 1, x_65); lean_ctor_set(x_66, 2, x_63); -x_67 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_67 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_68 = lean_array_push(x_67, x_59); x_69 = lean_array_push(x_68, x_66); x_70 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__5; @@ -11551,7 +11788,7 @@ x_71 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_71, 0, x_64); lean_ctor_set(x_71, 1, x_70); lean_ctor_set(x_71, 2, x_69); -x_72 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__7; +x_72 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__1; x_73 = lean_array_push(x_72, x_55); x_74 = lean_array_push(x_73, x_9); x_75 = lean_array_push(x_74, x_57); @@ -11569,13 +11806,13 @@ lean_ctor_set(x_82, 0, x_64); lean_ctor_set(x_82, 1, x_81); lean_ctor_set(x_82, 2, x_80); x_83 = lean_array_push(x_62, x_82); -x_84 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_84 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_85 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_85, 0, x_64); lean_ctor_set(x_85, 1, x_84); lean_ctor_set(x_85, 2, x_83); x_86 = lean_array_push(x_62, x_85); -x_87 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_87 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_88 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_88, 0, x_64); lean_ctor_set(x_88, 1, x_87); @@ -11600,7 +11837,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticShow_____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_tacticShow_____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -11630,7 +11867,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticShow_____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_tacticShow_____closed__4; x_3 = l_Lean_Parser_Tactic_rename___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -11756,11 +11993,11 @@ lean_inc(x_12); x_20 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_20, 0, x_12); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_Parser_Tactic_intros___closed__10; +x_21 = l_Lean_binderIdent___closed__10; x_22 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_22, 0, x_12); lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_23 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_24 = lean_array_push(x_23, x_22); x_25 = lean_box(2); x_26 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticAdmit__1___closed__15; @@ -11768,7 +12005,7 @@ x_27 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_27, 0, x_25); lean_ctor_set(x_27, 1, x_26); lean_ctor_set(x_27, 2, x_24); -x_28 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_28 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_29 = lean_array_push(x_28, x_20); x_30 = lean_array_push(x_29, x_27); x_31 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__5; @@ -11800,13 +12037,13 @@ lean_ctor_set(x_46, 0, x_25); lean_ctor_set(x_46, 1, x_45); lean_ctor_set(x_46, 2, x_44); x_47 = lean_array_push(x_23, x_46); -x_48 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_48 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_49 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_49, 0, x_25); lean_ctor_set(x_49, 1, x_48); lean_ctor_set(x_49, 2, x_47); x_50 = lean_array_push(x_23, x_49); -x_51 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_51 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_52 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_52, 0, x_25); lean_ctor_set(x_52, 1, x_51); @@ -11842,11 +12079,11 @@ lean_inc(x_53); x_62 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_62, 0, x_53); lean_ctor_set(x_62, 1, x_61); -x_63 = l_Lean_Parser_Tactic_intros___closed__10; +x_63 = l_Lean_binderIdent___closed__10; x_64 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_64, 0, x_53); lean_ctor_set(x_64, 1, x_63); -x_65 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_65 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_66 = lean_array_push(x_65, x_64); x_67 = lean_box(2); x_68 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticAdmit__1___closed__15; @@ -11854,7 +12091,7 @@ x_69 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_69, 0, x_67); lean_ctor_set(x_69, 1, x_68); lean_ctor_set(x_69, 2, x_66); -x_70 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_70 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_71 = lean_array_push(x_70, x_62); x_72 = lean_array_push(x_71, x_69); x_73 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__5; @@ -11886,13 +12123,13 @@ lean_ctor_set(x_88, 0, x_67); lean_ctor_set(x_88, 1, x_87); lean_ctor_set(x_88, 2, x_86); x_89 = lean_array_push(x_65, x_88); -x_90 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_90 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_91 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_91, 0, x_67); lean_ctor_set(x_91, 1, x_90); lean_ctor_set(x_91, 2, x_89); x_92 = lean_array_push(x_65, x_91); -x_93 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_93 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_94 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_94, 0, x_67); lean_ctor_set(x_94, 1, x_93); @@ -11917,7 +12154,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_letrec___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_letrec___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -11957,7 +12194,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_letrec___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_letrec___closed__3; x_3 = l_Lean_Parser_Tactic_letrec___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -12023,7 +12260,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_letrec___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_letrec___closed__8; x_3 = l_Lean_Parser_Tactic_letrec___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); @@ -12175,7 +12412,7 @@ lean_inc(x_22); x_28 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_28, 0, x_22); lean_ctor_set(x_28, 1, x_27); -x_29 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_29 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_30 = lean_array_push(x_29, x_26); x_31 = lean_array_push(x_30, x_28); x_32 = lean_box(2); @@ -12183,7 +12420,7 @@ x_33 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_33, 0, x_32); lean_ctor_set(x_33, 1, x_10); lean_ctor_set(x_33, 2, x_31); -x_34 = l_Lean_Parser_Tactic_rwWithRfl___closed__2; +x_34 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__2; lean_inc(x_22); x_35 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_35, 0, x_22); @@ -12193,11 +12430,11 @@ lean_inc(x_22); x_37 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_37, 0, x_22); lean_ctor_set(x_37, 1, x_36); -x_38 = l_Lean_Parser_Tactic_intros___closed__10; +x_38 = l_Lean_binderIdent___closed__10; x_39 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_39, 0, x_22); lean_ctor_set(x_39, 1, x_38); -x_40 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_40 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_41 = lean_array_push(x_40, x_39); x_42 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticAdmit__1___closed__15; x_43 = lean_alloc_ctor(1, 3, 0); @@ -12211,7 +12448,7 @@ x_47 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_47, 0, x_32); lean_ctor_set(x_47, 1, x_46); lean_ctor_set(x_47, 2, x_45); -x_48 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__7; +x_48 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__1; x_49 = lean_array_push(x_48, x_33); x_50 = lean_array_push(x_49, x_15); x_51 = lean_array_push(x_50, x_35); @@ -12254,7 +12491,7 @@ lean_inc(x_59); x_66 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_66, 0, x_59); lean_ctor_set(x_66, 1, x_65); -x_67 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_67 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_68 = lean_array_push(x_67, x_64); x_69 = lean_array_push(x_68, x_66); x_70 = lean_box(2); @@ -12262,7 +12499,7 @@ x_71 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_71, 0, x_70); lean_ctor_set(x_71, 1, x_10); lean_ctor_set(x_71, 2, x_69); -x_72 = l_Lean_Parser_Tactic_rwWithRfl___closed__2; +x_72 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__2; lean_inc(x_59); x_73 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_73, 0, x_59); @@ -12272,11 +12509,11 @@ lean_inc(x_59); x_75 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_75, 0, x_59); lean_ctor_set(x_75, 1, x_74); -x_76 = l_Lean_Parser_Tactic_intros___closed__10; +x_76 = l_Lean_binderIdent___closed__10; x_77 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_77, 0, x_59); lean_ctor_set(x_77, 1, x_76); -x_78 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_78 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_79 = lean_array_push(x_78, x_77); x_80 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticAdmit__1___closed__15; x_81 = lean_alloc_ctor(1, 3, 0); @@ -12290,7 +12527,7 @@ x_85 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_85, 0, x_70); lean_ctor_set(x_85, 1, x_84); lean_ctor_set(x_85, 2, x_83); -x_86 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__7; +x_86 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__1; x_87 = lean_array_push(x_86, x_71); x_88 = lean_array_push(x_87, x_15); x_89 = lean_array_push(x_88, x_73); @@ -12329,7 +12566,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticRefine__lift_x27_____closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_tacticRefine__lift_x27_____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -12359,7 +12596,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticRefine__lift_x27_____closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_tacticRefine__lift_x27_____closed__4; x_3 = l_Lean_Parser_Tactic_rename___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -12441,7 +12678,7 @@ lean_inc(x_12); x_20 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_20, 0, x_12); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_21 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_22 = lean_array_push(x_21, x_20); x_23 = lean_array_push(x_22, x_9); x_24 = lean_box(2); @@ -12457,14 +12694,14 @@ x_30 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_30, 0, x_24); lean_ctor_set(x_30, 1, x_29); lean_ctor_set(x_30, 2, x_28); -x_31 = l_Lean_Parser_Tactic_rwWithRfl___closed__2; +x_31 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__2; lean_inc(x_12); x_32 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_32, 0, x_12); lean_ctor_set(x_32, 1, x_31); -x_33 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_33 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_34 = lean_array_push(x_33, x_32); -x_35 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_35 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_36 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_36, 0, x_24); lean_ctor_set(x_36, 1, x_35); @@ -12482,7 +12719,7 @@ x_42 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_42, 0, x_12); lean_ctor_set(x_42, 1, x_41); x_43 = lean_array_push(x_21, x_42); -x_44 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; +x_44 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__7; x_45 = lean_array_push(x_43, x_44); x_46 = l_Lean_Parser_Tactic_rotateRight___closed__2; x_47 = lean_alloc_ctor(1, 3, 0); @@ -12502,13 +12739,13 @@ lean_ctor_set(x_53, 0, x_24); lean_ctor_set(x_53, 1, x_35); lean_ctor_set(x_53, 2, x_52); x_54 = lean_array_push(x_33, x_53); -x_55 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__5; +x_55 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; x_56 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_56, 0, x_24); lean_ctor_set(x_56, 1, x_55); lean_ctor_set(x_56, 2, x_54); x_57 = lean_array_push(x_33, x_56); -x_58 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__5; +x_58 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__1; x_59 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_59, 0, x_24); lean_ctor_set(x_59, 1, x_58); @@ -12560,7 +12797,7 @@ lean_ctor_set(x_82, 0, x_24); lean_ctor_set(x_82, 1, x_35); lean_ctor_set(x_82, 2, x_81); x_83 = lean_array_push(x_33, x_82); -x_84 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_84 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_85 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_85, 0, x_24); lean_ctor_set(x_85, 1, x_84); @@ -12596,7 +12833,7 @@ lean_inc(x_86); x_95 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_95, 0, x_86); lean_ctor_set(x_95, 1, x_94); -x_96 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_96 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_97 = lean_array_push(x_96, x_95); x_98 = lean_array_push(x_97, x_9); x_99 = lean_box(2); @@ -12612,14 +12849,14 @@ x_105 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_105, 0, x_99); lean_ctor_set(x_105, 1, x_104); lean_ctor_set(x_105, 2, x_103); -x_106 = l_Lean_Parser_Tactic_rwWithRfl___closed__2; +x_106 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__2; lean_inc(x_86); x_107 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_107, 0, x_86); lean_ctor_set(x_107, 1, x_106); -x_108 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_108 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_109 = lean_array_push(x_108, x_107); -x_110 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_110 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_111 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_111, 0, x_99); lean_ctor_set(x_111, 1, x_110); @@ -12637,7 +12874,7 @@ x_117 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_117, 0, x_86); lean_ctor_set(x_117, 1, x_116); x_118 = lean_array_push(x_96, x_117); -x_119 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; +x_119 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__7; x_120 = lean_array_push(x_118, x_119); x_121 = l_Lean_Parser_Tactic_rotateRight___closed__2; x_122 = lean_alloc_ctor(1, 3, 0); @@ -12657,13 +12894,13 @@ lean_ctor_set(x_128, 0, x_99); lean_ctor_set(x_128, 1, x_110); lean_ctor_set(x_128, 2, x_127); x_129 = lean_array_push(x_108, x_128); -x_130 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__5; +x_130 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; x_131 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_131, 0, x_99); lean_ctor_set(x_131, 1, x_130); lean_ctor_set(x_131, 2, x_129); x_132 = lean_array_push(x_108, x_131); -x_133 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__5; +x_133 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__1; x_134 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_134, 0, x_99); lean_ctor_set(x_134, 1, x_133); @@ -12715,7 +12952,7 @@ lean_ctor_set(x_157, 0, x_99); lean_ctor_set(x_157, 1, x_110); lean_ctor_set(x_157, 2, x_156); x_158 = lean_array_push(x_108, x_157); -x_159 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_159 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_160 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_160, 0, x_99); lean_ctor_set(x_160, 1, x_159); @@ -12740,7 +12977,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticHave_x27_____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_tacticHave_x27_____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -12770,7 +13007,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticHave_x27_____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_tacticHave_x27_____closed__4; x_3 = l_Lean_Parser_Tactic_tacticHave_____closed__7; x_4 = lean_alloc_ctor(2, 3, 0); @@ -12850,7 +13087,7 @@ lean_inc(x_12); x_16 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_16, 0, x_12); lean_ctor_set(x_16, 1, x_15); -x_17 = l_Lean_Parser_Tactic_rwWithRfl___closed__2; +x_17 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__2; lean_inc(x_12); x_18 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_18, 0, x_12); @@ -12860,11 +13097,11 @@ lean_inc(x_12); x_20 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_20, 0, x_12); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_Parser_Tactic_intros___closed__10; +x_21 = l_Lean_binderIdent___closed__10; x_22 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_22, 0, x_12); lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_23 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_24 = lean_array_push(x_23, x_22); x_25 = lean_box(2); x_26 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticAdmit__1___closed__15; @@ -12872,7 +13109,7 @@ x_27 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_27, 0, x_25); lean_ctor_set(x_27, 1, x_26); lean_ctor_set(x_27, 2, x_24); -x_28 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_28 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_29 = lean_array_push(x_28, x_20); x_30 = lean_array_push(x_29, x_27); x_31 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__5; @@ -12880,7 +13117,7 @@ x_32 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_32, 0, x_25); lean_ctor_set(x_32, 1, x_31); lean_ctor_set(x_32, 2, x_30); -x_33 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__7; +x_33 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__1; x_34 = lean_array_push(x_33, x_16); x_35 = lean_array_push(x_34, x_9); x_36 = lean_array_push(x_35, x_18); @@ -12898,13 +13135,13 @@ lean_ctor_set(x_43, 0, x_25); lean_ctor_set(x_43, 1, x_42); lean_ctor_set(x_43, 2, x_41); x_44 = lean_array_push(x_23, x_43); -x_45 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_45 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_46 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_46, 0, x_25); lean_ctor_set(x_46, 1, x_45); lean_ctor_set(x_46, 2, x_44); x_47 = lean_array_push(x_23, x_46); -x_48 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_48 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_49 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_49, 0, x_25); lean_ctor_set(x_49, 1, x_48); @@ -12930,7 +13167,7 @@ lean_inc(x_50); x_55 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_55, 0, x_50); lean_ctor_set(x_55, 1, x_54); -x_56 = l_Lean_Parser_Tactic_rwWithRfl___closed__2; +x_56 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__2; lean_inc(x_50); x_57 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_57, 0, x_50); @@ -12940,11 +13177,11 @@ lean_inc(x_50); x_59 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_59, 0, x_50); lean_ctor_set(x_59, 1, x_58); -x_60 = l_Lean_Parser_Tactic_intros___closed__10; +x_60 = l_Lean_binderIdent___closed__10; x_61 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_61, 0, x_50); lean_ctor_set(x_61, 1, x_60); -x_62 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_62 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_63 = lean_array_push(x_62, x_61); x_64 = lean_box(2); x_65 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticAdmit__1___closed__15; @@ -12952,7 +13189,7 @@ x_66 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_66, 0, x_64); lean_ctor_set(x_66, 1, x_65); lean_ctor_set(x_66, 2, x_63); -x_67 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_67 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_68 = lean_array_push(x_67, x_59); x_69 = lean_array_push(x_68, x_66); x_70 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__5; @@ -12960,7 +13197,7 @@ x_71 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_71, 0, x_64); lean_ctor_set(x_71, 1, x_70); lean_ctor_set(x_71, 2, x_69); -x_72 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__7; +x_72 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__1; x_73 = lean_array_push(x_72, x_55); x_74 = lean_array_push(x_73, x_9); x_75 = lean_array_push(x_74, x_57); @@ -12978,13 +13215,13 @@ lean_ctor_set(x_82, 0, x_64); lean_ctor_set(x_82, 1, x_81); lean_ctor_set(x_82, 2, x_80); x_83 = lean_array_push(x_62, x_82); -x_84 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_84 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_85 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_85, 0, x_64); lean_ctor_set(x_85, 1, x_84); lean_ctor_set(x_85, 2, x_83); x_86 = lean_array_push(x_62, x_85); -x_87 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_87 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_88 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_88, 0, x_64); lean_ctor_set(x_88, 1, x_87); @@ -12997,25 +13234,25 @@ return x_89; } } } -static lean_object* _init_l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__1() { +static lean_object* _init_l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("tacticHave'__:=_", 16); +x_1 = lean_mk_string_from_bytes("tacticHave'_:=_", 15); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__2() { +static lean_object* _init_l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; -x_2 = l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__1; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; +x_2 = l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__3() { +static lean_object* _init_l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__3() { _start: { lean_object* x_1; @@ -13023,11 +13260,11 @@ x_1 = lean_mk_string_from_bytes("have'", 5); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__4() { +static lean_object* _init_l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__3; +x_1 = l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__3; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -13035,13 +13272,13 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; -x_2 = l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__4; -x_3 = l_Lean_Parser_Tactic_intros___closed__9; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; +x_2 = l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__4; +x_3 = l_Lean_binderIdent___closed__9; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -13049,12 +13286,12 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; -x_2 = l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__5; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; +x_2 = l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__5; x_3 = l_Lean_Parser_Tactic_config___closed__10; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -13063,12 +13300,12 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__7() { +static lean_object* _init_l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; -x_2 = l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; +x_2 = l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__6; x_3 = l_Lean_Parser_Tactic_rename___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -13077,13 +13314,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__8() { +static lean_object* _init_l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__2; +x_1 = l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__7; +x_3 = l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__7; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -13091,19 +13328,19 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d__() { +static lean_object* _init_l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d__() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__8; +x_1 = l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__8; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave_x27_____x3a_x3d____1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave_x27___x3a_x3d____1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; uint8_t x_5; -x_4 = l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__2; +x_4 = l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__2; lean_inc(x_1); x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); if (x_5 == 0) @@ -13131,17 +13368,17 @@ 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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; x_14 = lean_ctor_get(x_12, 0); -x_15 = l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__3; +x_15 = l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__3; lean_inc(x_14); x_16 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_16, 0, x_14); lean_ctor_set(x_16, 1, x_15); -x_17 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_17 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_18 = lean_array_push(x_17, x_9); -x_19 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; +x_19 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__7; x_20 = lean_array_push(x_18, x_19); x_21 = lean_box(2); -x_22 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_22 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_23 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_23, 0, x_21); lean_ctor_set(x_23, 1, x_22); @@ -13151,12 +13388,12 @@ lean_inc(x_14); x_25 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_25, 0, x_14); lean_ctor_set(x_25, 1, x_24); -x_26 = l_Lean_Parser_Tactic_intros___closed__10; +x_26 = l_Lean_binderIdent___closed__10; lean_inc(x_14); x_27 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_27, 0, x_14); lean_ctor_set(x_27, 1, x_26); -x_28 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_28 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_29 = lean_array_push(x_28, x_27); x_30 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticAdmit__1___closed__15; x_31 = lean_alloc_ctor(1, 3, 0); @@ -13179,7 +13416,7 @@ x_38 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__ x_39 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_39, 0, x_14); lean_ctor_set(x_39, 1, x_38); -x_40 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__7; +x_40 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__1; x_41 = lean_array_push(x_40, x_23); x_42 = lean_array_push(x_41, x_37); x_43 = lean_array_push(x_42, x_39); @@ -13208,7 +13445,7 @@ lean_ctor_set(x_55, 0, x_21); lean_ctor_set(x_55, 1, x_22); lean_ctor_set(x_55, 2, x_54); x_56 = lean_array_push(x_28, x_55); -x_57 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_57 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_58 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_58, 0, x_21); lean_ctor_set(x_58, 1, x_57); @@ -13224,17 +13461,17 @@ x_60 = lean_ctor_get(x_12, 1); lean_inc(x_60); lean_inc(x_59); lean_dec(x_12); -x_61 = l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__3; +x_61 = l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__3; lean_inc(x_59); x_62 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_62, 0, x_59); lean_ctor_set(x_62, 1, x_61); -x_63 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_63 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_64 = lean_array_push(x_63, x_9); -x_65 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; +x_65 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__7; x_66 = lean_array_push(x_64, x_65); x_67 = lean_box(2); -x_68 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_68 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_69 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_69, 0, x_67); lean_ctor_set(x_69, 1, x_68); @@ -13244,12 +13481,12 @@ lean_inc(x_59); x_71 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_71, 0, x_59); lean_ctor_set(x_71, 1, x_70); -x_72 = l_Lean_Parser_Tactic_intros___closed__10; +x_72 = l_Lean_binderIdent___closed__10; lean_inc(x_59); x_73 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_73, 0, x_59); lean_ctor_set(x_73, 1, x_72); -x_74 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_74 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_75 = lean_array_push(x_74, x_73); x_76 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticAdmit__1___closed__15; x_77 = lean_alloc_ctor(1, 3, 0); @@ -13272,7 +13509,7 @@ x_84 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__ x_85 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_85, 0, x_59); lean_ctor_set(x_85, 1, x_84); -x_86 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__7; +x_86 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__1; x_87 = lean_array_push(x_86, x_69); x_88 = lean_array_push(x_87, x_83); x_89 = lean_array_push(x_88, x_85); @@ -13301,7 +13538,7 @@ lean_ctor_set(x_101, 0, x_67); lean_ctor_set(x_101, 1, x_68); lean_ctor_set(x_101, 2, x_100); x_102 = lean_array_push(x_74, x_101); -x_103 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_103 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_104 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_104, 0, x_67); lean_ctor_set(x_104, 1, x_103); @@ -13326,7 +13563,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticLet_x27_____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_tacticLet_x27_____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -13356,7 +13593,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticLet_x27_____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_tacticLet_x27_____closed__4; x_3 = l_Lean_Parser_Tactic_tacticLet_____closed__7; x_4 = lean_alloc_ctor(2, 3, 0); @@ -13428,7 +13665,7 @@ lean_inc(x_12); x_16 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_16, 0, x_12); lean_ctor_set(x_16, 1, x_15); -x_17 = l_Lean_Parser_Tactic_rwWithRfl___closed__2; +x_17 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__2; lean_inc(x_12); x_18 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_18, 0, x_12); @@ -13438,11 +13675,11 @@ lean_inc(x_12); x_20 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_20, 0, x_12); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_Parser_Tactic_intros___closed__10; +x_21 = l_Lean_binderIdent___closed__10; x_22 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_22, 0, x_12); lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_23 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_24 = lean_array_push(x_23, x_22); x_25 = lean_box(2); x_26 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticAdmit__1___closed__15; @@ -13450,7 +13687,7 @@ x_27 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_27, 0, x_25); lean_ctor_set(x_27, 1, x_26); lean_ctor_set(x_27, 2, x_24); -x_28 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_28 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_29 = lean_array_push(x_28, x_20); x_30 = lean_array_push(x_29, x_27); x_31 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__5; @@ -13458,7 +13695,7 @@ x_32 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_32, 0, x_25); lean_ctor_set(x_32, 1, x_31); lean_ctor_set(x_32, 2, x_30); -x_33 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__7; +x_33 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__1; x_34 = lean_array_push(x_33, x_16); x_35 = lean_array_push(x_34, x_9); x_36 = lean_array_push(x_35, x_18); @@ -13476,13 +13713,13 @@ lean_ctor_set(x_43, 0, x_25); lean_ctor_set(x_43, 1, x_42); lean_ctor_set(x_43, 2, x_41); x_44 = lean_array_push(x_23, x_43); -x_45 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_45 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_46 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_46, 0, x_25); lean_ctor_set(x_46, 1, x_45); lean_ctor_set(x_46, 2, x_44); x_47 = lean_array_push(x_23, x_46); -x_48 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_48 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_49 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_49, 0, x_25); lean_ctor_set(x_49, 1, x_48); @@ -13508,7 +13745,7 @@ lean_inc(x_50); x_55 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_55, 0, x_50); lean_ctor_set(x_55, 1, x_54); -x_56 = l_Lean_Parser_Tactic_rwWithRfl___closed__2; +x_56 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__2; lean_inc(x_50); x_57 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_57, 0, x_50); @@ -13518,11 +13755,11 @@ lean_inc(x_50); x_59 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_59, 0, x_50); lean_ctor_set(x_59, 1, x_58); -x_60 = l_Lean_Parser_Tactic_intros___closed__10; +x_60 = l_Lean_binderIdent___closed__10; x_61 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_61, 0, x_50); lean_ctor_set(x_61, 1, x_60); -x_62 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_62 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_63 = lean_array_push(x_62, x_61); x_64 = lean_box(2); x_65 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticAdmit__1___closed__15; @@ -13530,7 +13767,7 @@ x_66 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_66, 0, x_64); lean_ctor_set(x_66, 1, x_65); lean_ctor_set(x_66, 2, x_63); -x_67 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_67 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_68 = lean_array_push(x_67, x_59); x_69 = lean_array_push(x_68, x_66); x_70 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__5; @@ -13538,7 +13775,7 @@ x_71 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_71, 0, x_64); lean_ctor_set(x_71, 1, x_70); lean_ctor_set(x_71, 2, x_69); -x_72 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__7; +x_72 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__1; x_73 = lean_array_push(x_72, x_55); x_74 = lean_array_push(x_73, x_9); x_75 = lean_array_push(x_74, x_57); @@ -13556,13 +13793,13 @@ lean_ctor_set(x_82, 0, x_64); lean_ctor_set(x_82, 1, x_81); lean_ctor_set(x_82, 2, x_80); x_83 = lean_array_push(x_62, x_82); -x_84 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_84 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_85 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_85, 0, x_64); lean_ctor_set(x_85, 1, x_84); lean_ctor_set(x_85, 2, x_83); x_86 = lean_array_push(x_62, x_85); -x_87 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_87 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_88 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_88, 0, x_64); lean_ctor_set(x_88, 1, x_87); @@ -13587,7 +13824,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_inductionAltLHS___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_inductionAltLHS___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -13637,9 +13874,9 @@ static lean_object* _init_l_Lean_Parser_Tactic_inductionAltLHS___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_inductionAltLHS___closed__6; -x_3 = l_Lean_Parser_Tactic_intros___closed__9; +x_3 = l_Lean_binderIdent___closed__9; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -13663,9 +13900,9 @@ static lean_object* _init_l_Lean_Parser_Tactic_inductionAltLHS___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_intros___closed__6; +x_1 = l_Lean_binderIdent___closed__6; x_2 = l_Lean_Parser_Tactic_inductionAltLHS___closed__8; -x_3 = l_Lean_Parser_Tactic_intros___closed__11; +x_3 = l_Lean_binderIdent___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -13677,7 +13914,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_inductionAltLHS___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_inductionAltLHS___closed__4; x_3 = l_Lean_Parser_Tactic_inductionAltLHS___closed__9; x_4 = lean_alloc_ctor(2, 3, 0); @@ -13690,10 +13927,22 @@ return x_4; static lean_object* _init_l_Lean_Parser_Tactic_inductionAltLHS___closed__11() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Tactic_intro___closed__12; +x_2 = l_Lean_binderIdent___closed__12; +x_3 = lean_alloc_ctor(1, 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_Lean_Parser_Tactic_inductionAltLHS___closed__12() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_inductionAltLHS___closed__10; -x_3 = l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__5; +x_3 = l_Lean_Parser_Tactic_inductionAltLHS___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -13701,13 +13950,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_inductionAltLHS___closed__12() { +static lean_object* _init_l_Lean_Parser_Tactic_inductionAltLHS___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Tactic_inductionAltLHS___closed__1; x_2 = l_Lean_Parser_Tactic_inductionAltLHS___closed__2; -x_3 = l_Lean_Parser_Tactic_inductionAltLHS___closed__11; +x_3 = l_Lean_Parser_Tactic_inductionAltLHS___closed__12; x_4 = lean_alloc_ctor(9, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -13719,7 +13968,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_inductionAltLHS() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_inductionAltLHS___closed__12; +x_1 = l_Lean_Parser_Tactic_inductionAltLHS___closed__13; return x_1; } } @@ -13735,7 +13984,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_inductionAlt___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_inductionAlt___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -13815,7 +14064,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_inductionAlt___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_inductionAlt___closed__8; x_3 = l_Lean_Parser_Tactic_inductionAlt___closed__9; x_4 = lean_alloc_ctor(2, 3, 0); @@ -13829,7 +14078,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_inductionAlt___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_inductionAlt___closed__10; x_3 = l_Lean_Parser_Tactic_rename___closed__8; x_4 = lean_alloc_ctor(2, 3, 0); @@ -13883,7 +14132,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_inductionAlt___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_intros___closed__6; +x_1 = l_Lean_binderIdent___closed__6; x_2 = l_Lean_Parser_Tactic_inductionAlt___closed__15; x_3 = l_Lean_Parser_Tactic_case___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); @@ -13897,7 +14146,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_inductionAlt___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_intros___closed__6; +x_1 = l_Lean_binderIdent___closed__6; x_2 = l_Lean_Parser_Tactic_inductionAlt___closed__13; x_3 = l_Lean_Parser_Tactic_inductionAlt___closed__16; x_4 = lean_alloc_ctor(2, 3, 0); @@ -13911,7 +14160,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_inductionAlt___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_inductionAlt___closed__11; x_3 = l_Lean_Parser_Tactic_inductionAlt___closed__17; x_4 = lean_alloc_ctor(2, 3, 0); @@ -13955,7 +14204,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_inductionAlts___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_inductionAlts___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -13984,7 +14233,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_inductionAlts___closed__5() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_rotateLeft___closed__6; -x_2 = l_Lean_Parser_Tactic_withAnnotateState___closed__23; +x_2 = l_Lean_Parser_Tactic_withAnnotateState___closed__21; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -13995,7 +14244,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_inductionAlts___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_inductionAlts___closed__4; x_3 = l_Lean_Parser_Tactic_inductionAlts___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -14009,7 +14258,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_inductionAlts___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_first___closed__11; x_3 = l_Lean_Parser_Tactic_inductionAlt; x_4 = lean_alloc_ctor(2, 3, 0); @@ -14047,7 +14296,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_inductionAlts___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_inductionAlts___closed__6; x_3 = l_Lean_Parser_Tactic_inductionAlts___closed__9; x_4 = lean_alloc_ctor(2, 3, 0); @@ -14091,7 +14340,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_induction___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_induction___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -14137,7 +14386,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_induction___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_induction___closed__4; x_3 = l_Lean_Parser_Tactic_induction___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -14169,9 +14418,9 @@ static lean_object* _init_l_Lean_Parser_Tactic_induction___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_induction___closed__8; -x_3 = l_Lean_Parser_Tactic_intros___closed__9; +x_3 = l_Lean_binderIdent___closed__9; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -14195,7 +14444,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_induction___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_induction___closed__6; x_3 = l_Lean_Parser_Tactic_induction___closed__10; x_4 = lean_alloc_ctor(2, 3, 0); @@ -14227,7 +14476,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_induction___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_induction___closed__13; x_3 = l_Lean_Parser_Tactic_revert___closed__7; x_4 = lean_alloc_ctor(2, 3, 0); @@ -14253,7 +14502,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_induction___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_induction___closed__11; x_3 = l_Lean_Parser_Tactic_induction___closed__15; x_4 = lean_alloc_ctor(2, 3, 0); @@ -14279,7 +14528,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_induction___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_induction___closed__16; x_3 = l_Lean_Parser_Tactic_induction___closed__17; x_4 = lean_alloc_ctor(2, 3, 0); @@ -14323,7 +14572,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_generalizeArg___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_generalizeArg___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -14351,8 +14600,8 @@ static lean_object* _init_l_Lean_Parser_Tactic_generalizeArg___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; -x_2 = l_Lean_Parser_Tactic_intros___closed__9; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; +x_2 = l_Lean_binderIdent___closed__9; x_3 = l_Lean_Parser_Tactic_generalizeArg___closed__4; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -14401,7 +14650,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_generalizeArg___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_generalizeArg___closed__7; x_3 = l_Lean_Parser_Tactic_generalizeArg___closed__8; x_4 = lean_alloc_ctor(2, 3, 0); @@ -14433,7 +14682,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_generalizeArg___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_generalizeArg___closed__9; x_3 = l_Lean_Parser_Tactic_generalizeArg___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); @@ -14447,9 +14696,9 @@ static lean_object* _init_l_Lean_Parser_Tactic_generalizeArg___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_generalizeArg___closed__12; -x_3 = l_Lean_Parser_Tactic_intros___closed__9; +x_3 = l_Lean_binderIdent___closed__9; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -14491,7 +14740,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_generalize___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_generalize___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -14537,7 +14786,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_generalize___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_generalize___closed__4; x_3 = l_Lean_Parser_Tactic_generalize___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -14581,7 +14830,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_casesTarget___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_casesTarget___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -14591,7 +14840,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_casesTarget___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_generalizeArg___closed__7; x_3 = l_Lean_Parser_Tactic_rename___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -14635,7 +14884,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_cases___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_cases___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -14681,7 +14930,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_cases___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_cases___closed__4; x_3 = l_Lean_Parser_Tactic_cases___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -14695,7 +14944,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_cases___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_cases___closed__6; x_3 = l_Lean_Parser_Tactic_induction___closed__10; x_4 = lean_alloc_ctor(2, 3, 0); @@ -14709,7 +14958,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_cases___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_cases___closed__7; x_3 = l_Lean_Parser_Tactic_induction___closed__17; x_4 = lean_alloc_ctor(2, 3, 0); @@ -14753,7 +15002,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_renameI___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_renameI___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -14783,9 +15032,9 @@ static lean_object* _init_l_Lean_Parser_Tactic_renameI___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; -x_2 = l_Lean_Parser_Tactic_renameI___closed__4; -x_3 = l_Lean_Parser_Tactic_injection___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; +x_2 = l_Lean_Parser_Tactic_intro___closed__15; +x_3 = l_Lean_binderIdent; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -14796,10 +15045,36 @@ return x_4; static lean_object* _init_l_Lean_Parser_Tactic_renameI___closed__6() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Tactic_revert___closed__6; +x_2 = l_Lean_Parser_Tactic_renameI___closed__5; +x_3 = lean_alloc_ctor(1, 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_Lean_Parser_Tactic_renameI___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; +x_2 = l_Lean_Parser_Tactic_renameI___closed__4; +x_3 = l_Lean_Parser_Tactic_renameI___closed__6; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_renameI___closed__8() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Tactic_renameI___closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Parser_Tactic_renameI___closed__5; +x_3 = l_Lean_Parser_Tactic_renameI___closed__7; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -14811,7 +15086,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_renameI() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_renameI___closed__6; +x_1 = l_Lean_Parser_Tactic_renameI___closed__8; return x_1; } } @@ -14827,7 +15102,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticRepeat_____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_tacticRepeat_____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -14857,7 +15132,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticRepeat_____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_tacticRepeat_____closed__4; x_3 = l_Lean_Parser_Tactic_case___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); @@ -14912,351 +15187,322 @@ lean_dec(x_1); x_6 = lean_box(1); x_7 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_7, 0, x_6); -lean_ctor_set(x_7, 1, x_3); -return x_7; -} -else -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_unsigned_to_nat(1u); -x_9 = l_Lean_Syntax_getArg(x_1, x_8); -lean_dec(x_1); -x_10 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__5; -lean_inc(x_9); -x_11 = l_Lean_Syntax_isOfKind(x_9, x_10); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; -lean_dec(x_9); -lean_dec(x_2); -x_12 = lean_box(1); -x_13 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_3); -return x_13; +lean_ctor_set(x_7, 1, x_3); +return x_7; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_14 = lean_unsigned_to_nat(0u); -x_15 = l_Lean_Syntax_getArg(x_9, x_14); -lean_dec(x_9); -x_16 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_17 = !lean_is_exclusive(x_16); -if (x_17 == 0) +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_unsigned_to_nat(1u); +x_9 = l_Lean_Syntax_getArg(x_1, x_8); +lean_dec(x_1); +x_10 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) { -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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_18 = lean_ctor_get(x_16, 0); -x_19 = l_Lean_Parser_Tactic_first___closed__1; -lean_inc(x_18); +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; 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_12 = lean_ctor_get(x_10, 0); +x_13 = l_Lean_Parser_Tactic_first___closed__1; +lean_inc(x_12); +x_14 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +x_15 = l_Lean_Parser_Tactic_intro___closed__7; +lean_inc(x_12); +x_16 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_16, 0, x_12); +lean_ctor_set(x_16, 1, x_15); +x_17 = l_Lean_Parser_Tactic_paren___closed__3; +lean_inc(x_12); +x_18 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_18, 0, x_12); +lean_ctor_set(x_18, 1, x_17); +x_19 = l_Lean_Parser_Tactic_paren___closed__6; +lean_inc(x_12); x_20 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 0, x_12); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_Parser_Tactic_intro___closed__7; -lean_inc(x_18); -x_22 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_22, 0, x_18); -lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean_Parser_Tactic_paren___closed__3; -lean_inc(x_18); -x_24 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_24, 0, x_18); -lean_ctor_set(x_24, 1, x_23); -x_25 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; -x_26 = lean_array_push(x_25, x_15); -x_27 = lean_box(2); -x_28 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_28, 0, x_27); -lean_ctor_set(x_28, 1, x_10); -lean_ctor_set(x_28, 2, x_26); -x_29 = l_Lean_Parser_Tactic_paren___closed__6; -lean_inc(x_18); -x_30 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_30, 0, x_18); -lean_ctor_set(x_30, 1, x_29); -x_31 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tactic___x3c_x3b_x3e____1___closed__2; -x_32 = lean_array_push(x_31, x_24); -lean_inc(x_28); -x_33 = lean_array_push(x_32, x_28); -x_34 = lean_array_push(x_33, x_30); -x_35 = l_Lean_Parser_Tactic_paren___closed__2; -x_36 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_36, 0, x_27); -lean_ctor_set(x_36, 1, x_35); -lean_ctor_set(x_36, 2, x_34); -x_37 = l_Lean_Parser_Tactic_rwWithRfl___closed__2; -lean_inc(x_18); -x_38 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_38, 0, x_18); +x_21 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tactic___x3c_x3b_x3e____1___closed__2; +x_22 = lean_array_push(x_21, x_18); +lean_inc(x_9); +x_23 = lean_array_push(x_22, x_9); +x_24 = lean_array_push(x_23, x_20); +x_25 = lean_box(2); +x_26 = l_Lean_Parser_Tactic_paren___closed__2; +x_27 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +lean_ctor_set(x_27, 2, x_24); +x_28 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__2; +lean_inc(x_12); +x_29 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_29, 0, x_12); +lean_ctor_set(x_29, 1, x_28); +x_30 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; +x_31 = lean_array_push(x_30, x_29); +x_32 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; +x_33 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_33, 0, x_25); +lean_ctor_set(x_33, 1, x_32); +lean_ctor_set(x_33, 2, x_31); +x_34 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; +x_35 = lean_array_push(x_34, x_27); +x_36 = lean_array_push(x_35, x_33); +x_37 = l_Lean_Parser_Tactic_first___closed__8; +x_38 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_38, 0, x_25); lean_ctor_set(x_38, 1, x_37); -x_39 = lean_array_push(x_25, x_38); -x_40 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; -x_41 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_41, 0, x_27); -lean_ctor_set(x_41, 1, x_40); -lean_ctor_set(x_41, 2, x_39); -x_42 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; -x_43 = lean_array_push(x_42, x_36); -x_44 = lean_array_push(x_43, x_41); -x_45 = l_Lean_Parser_Tactic_first___closed__8; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_27); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_44); -x_47 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticRepeat____1___closed__1; -lean_inc(x_18); -x_48 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_48, 0, x_18); -lean_ctor_set(x_48, 1, x_47); -x_49 = lean_array_push(x_42, x_48); -x_50 = lean_array_push(x_49, x_28); -x_51 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_51, 0, x_27); -lean_ctor_set(x_51, 1, x_4); -lean_ctor_set(x_51, 2, x_50); -x_52 = lean_array_push(x_42, x_51); -x_53 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; -x_54 = lean_array_push(x_52, x_53); -x_55 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_55, 0, x_27); -lean_ctor_set(x_55, 1, x_45); -lean_ctor_set(x_55, 2, x_54); -x_56 = lean_array_push(x_42, x_46); -x_57 = lean_array_push(x_56, x_55); -x_58 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_58, 0, x_27); -lean_ctor_set(x_58, 1, x_40); -lean_ctor_set(x_58, 2, x_57); -x_59 = lean_array_push(x_25, x_58); -x_60 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__5; -x_61 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_61, 0, x_27); +lean_ctor_set(x_38, 2, x_36); +x_39 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticRepeat____1___closed__1; +lean_inc(x_12); +x_40 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_40, 0, x_12); +lean_ctor_set(x_40, 1, x_39); +x_41 = lean_array_push(x_34, x_40); +x_42 = lean_array_push(x_41, x_9); +x_43 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_43, 0, x_25); +lean_ctor_set(x_43, 1, x_4); +lean_ctor_set(x_43, 2, x_42); +x_44 = lean_array_push(x_34, x_43); +x_45 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__7; +x_46 = lean_array_push(x_44, x_45); +x_47 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_47, 0, x_25); +lean_ctor_set(x_47, 1, x_37); +lean_ctor_set(x_47, 2, x_46); +x_48 = lean_array_push(x_34, x_38); +x_49 = lean_array_push(x_48, x_47); +x_50 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_50, 0, x_25); +lean_ctor_set(x_50, 1, x_32); +lean_ctor_set(x_50, 2, x_49); +x_51 = lean_array_push(x_30, x_50); +x_52 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; +x_53 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_53, 0, x_25); +lean_ctor_set(x_53, 1, x_52); +lean_ctor_set(x_53, 2, x_51); +x_54 = lean_array_push(x_30, x_53); +x_55 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__1; +x_56 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_56, 0, x_25); +lean_ctor_set(x_56, 1, x_55); +lean_ctor_set(x_56, 2, x_54); +x_57 = lean_array_push(x_34, x_16); +lean_inc(x_57); +x_58 = lean_array_push(x_57, x_56); +x_59 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_59, 0, x_25); +lean_ctor_set(x_59, 1, x_37); +lean_ctor_set(x_59, 2, x_58); +x_60 = l_Lean_Parser_Tactic_skip___closed__1; +x_61 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_61, 0, x_12); lean_ctor_set(x_61, 1, x_60); -lean_ctor_set(x_61, 2, x_59); -x_62 = lean_array_push(x_25, x_61); -x_63 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_63, 0, x_27); -lean_ctor_set(x_63, 1, x_10); -lean_ctor_set(x_63, 2, x_62); -x_64 = lean_array_push(x_42, x_22); -lean_inc(x_64); -x_65 = lean_array_push(x_64, x_63); -x_66 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_66, 0, x_27); -lean_ctor_set(x_66, 1, x_45); -lean_ctor_set(x_66, 2, x_65); -x_67 = l_Lean_Parser_Tactic_skip___closed__1; -x_68 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_68, 0, x_18); -lean_ctor_set(x_68, 1, x_67); -x_69 = lean_array_push(x_25, x_68); -x_70 = l_Lean_Parser_Tactic_skip___closed__2; +x_62 = lean_array_push(x_30, x_61); +x_63 = l_Lean_Parser_Tactic_skip___closed__2; +x_64 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_64, 0, x_25); +lean_ctor_set(x_64, 1, x_63); +lean_ctor_set(x_64, 2, x_62); +x_65 = lean_array_push(x_34, x_64); +x_66 = lean_array_push(x_65, x_45); +x_67 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_67, 0, x_25); +lean_ctor_set(x_67, 1, x_37); +lean_ctor_set(x_67, 2, x_66); +x_68 = lean_array_push(x_30, x_67); +x_69 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_69, 0, x_25); +lean_ctor_set(x_69, 1, x_32); +lean_ctor_set(x_69, 2, x_68); +x_70 = lean_array_push(x_30, x_69); x_71 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_71, 0, x_27); -lean_ctor_set(x_71, 1, x_70); -lean_ctor_set(x_71, 2, x_69); -x_72 = lean_array_push(x_42, x_71); -x_73 = lean_array_push(x_72, x_53); -x_74 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_74, 0, x_27); -lean_ctor_set(x_74, 1, x_45); -lean_ctor_set(x_74, 2, x_73); -x_75 = lean_array_push(x_25, x_74); -x_76 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_76, 0, x_27); -lean_ctor_set(x_76, 1, x_40); -lean_ctor_set(x_76, 2, x_75); -x_77 = lean_array_push(x_25, x_76); +lean_ctor_set(x_71, 0, x_25); +lean_ctor_set(x_71, 1, x_52); +lean_ctor_set(x_71, 2, x_70); +x_72 = lean_array_push(x_30, x_71); +x_73 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_73, 0, x_25); +lean_ctor_set(x_73, 1, x_55); +lean_ctor_set(x_73, 2, x_72); +x_74 = lean_array_push(x_57, x_73); +x_75 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_75, 0, x_25); +lean_ctor_set(x_75, 1, x_37); +lean_ctor_set(x_75, 2, x_74); +x_76 = lean_array_push(x_34, x_59); +x_77 = lean_array_push(x_76, x_75); x_78 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_78, 0, x_27); -lean_ctor_set(x_78, 1, x_60); +lean_ctor_set(x_78, 0, x_25); +lean_ctor_set(x_78, 1, x_32); lean_ctor_set(x_78, 2, x_77); -x_79 = lean_array_push(x_25, x_78); -x_80 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_80, 0, x_27); -lean_ctor_set(x_80, 1, x_10); -lean_ctor_set(x_80, 2, x_79); -x_81 = lean_array_push(x_64, x_80); +x_79 = lean_array_push(x_34, x_14); +x_80 = lean_array_push(x_79, x_78); +x_81 = l_Lean_Parser_Tactic_first___closed__2; x_82 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_82, 0, x_27); -lean_ctor_set(x_82, 1, x_45); -lean_ctor_set(x_82, 2, x_81); -x_83 = lean_array_push(x_42, x_66); -x_84 = lean_array_push(x_83, x_82); -x_85 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_85, 0, x_27); -lean_ctor_set(x_85, 1, x_40); -lean_ctor_set(x_85, 2, x_84); -x_86 = lean_array_push(x_42, x_20); -x_87 = lean_array_push(x_86, x_85); -x_88 = l_Lean_Parser_Tactic_first___closed__2; -x_89 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_89, 0, x_27); -lean_ctor_set(x_89, 1, x_88); -lean_ctor_set(x_89, 2, x_87); -lean_ctor_set(x_16, 0, x_89); -return x_16; +lean_ctor_set(x_82, 0, x_25); +lean_ctor_set(x_82, 1, x_81); +lean_ctor_set(x_82, 2, x_80); +lean_ctor_set(x_10, 0, x_82); +return x_10; } else { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; -x_90 = lean_ctor_get(x_16, 0); -x_91 = lean_ctor_get(x_16, 1); -lean_inc(x_91); -lean_inc(x_90); -lean_dec(x_16); -x_92 = l_Lean_Parser_Tactic_first___closed__1; -lean_inc(x_90); -x_93 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_93, 0, x_90); -lean_ctor_set(x_93, 1, x_92); -x_94 = l_Lean_Parser_Tactic_intro___closed__7; -lean_inc(x_90); -x_95 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_95, 0, x_90); -lean_ctor_set(x_95, 1, x_94); -x_96 = l_Lean_Parser_Tactic_paren___closed__3; -lean_inc(x_90); -x_97 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_97, 0, x_90); -lean_ctor_set(x_97, 1, x_96); -x_98 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; -x_99 = lean_array_push(x_98, x_15); -x_100 = lean_box(2); -x_101 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_101, 0, x_100); -lean_ctor_set(x_101, 1, x_10); -lean_ctor_set(x_101, 2, x_99); -x_102 = l_Lean_Parser_Tactic_paren___closed__6; -lean_inc(x_90); -x_103 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_103, 0, x_90); -lean_ctor_set(x_103, 1, x_102); -x_104 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tactic___x3c_x3b_x3e____1___closed__2; -x_105 = lean_array_push(x_104, x_97); -lean_inc(x_101); -x_106 = lean_array_push(x_105, x_101); -x_107 = lean_array_push(x_106, x_103); -x_108 = l_Lean_Parser_Tactic_paren___closed__2; -x_109 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_109, 0, x_100); -lean_ctor_set(x_109, 1, x_108); -lean_ctor_set(x_109, 2, x_107); -x_110 = l_Lean_Parser_Tactic_rwWithRfl___closed__2; -lean_inc(x_90); -x_111 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_111, 0, x_90); -lean_ctor_set(x_111, 1, x_110); -x_112 = lean_array_push(x_98, x_111); -x_113 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; -x_114 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_114, 0, x_100); -lean_ctor_set(x_114, 1, x_113); -lean_ctor_set(x_114, 2, x_112); -x_115 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; -x_116 = lean_array_push(x_115, x_109); -x_117 = lean_array_push(x_116, x_114); -x_118 = l_Lean_Parser_Tactic_first___closed__8; +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; +x_83 = lean_ctor_get(x_10, 0); +x_84 = lean_ctor_get(x_10, 1); +lean_inc(x_84); +lean_inc(x_83); +lean_dec(x_10); +x_85 = l_Lean_Parser_Tactic_first___closed__1; +lean_inc(x_83); +x_86 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_86, 0, x_83); +lean_ctor_set(x_86, 1, x_85); +x_87 = l_Lean_Parser_Tactic_intro___closed__7; +lean_inc(x_83); +x_88 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_88, 0, x_83); +lean_ctor_set(x_88, 1, x_87); +x_89 = l_Lean_Parser_Tactic_paren___closed__3; +lean_inc(x_83); +x_90 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_90, 0, x_83); +lean_ctor_set(x_90, 1, x_89); +x_91 = l_Lean_Parser_Tactic_paren___closed__6; +lean_inc(x_83); +x_92 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_92, 0, x_83); +lean_ctor_set(x_92, 1, x_91); +x_93 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tactic___x3c_x3b_x3e____1___closed__2; +x_94 = lean_array_push(x_93, x_90); +lean_inc(x_9); +x_95 = lean_array_push(x_94, x_9); +x_96 = lean_array_push(x_95, x_92); +x_97 = lean_box(2); +x_98 = l_Lean_Parser_Tactic_paren___closed__2; +x_99 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_99, 0, x_97); +lean_ctor_set(x_99, 1, x_98); +lean_ctor_set(x_99, 2, x_96); +x_100 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__2; +lean_inc(x_83); +x_101 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_101, 0, x_83); +lean_ctor_set(x_101, 1, x_100); +x_102 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; +x_103 = lean_array_push(x_102, x_101); +x_104 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; +x_105 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_105, 0, x_97); +lean_ctor_set(x_105, 1, x_104); +lean_ctor_set(x_105, 2, x_103); +x_106 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; +x_107 = lean_array_push(x_106, x_99); +x_108 = lean_array_push(x_107, x_105); +x_109 = l_Lean_Parser_Tactic_first___closed__8; +x_110 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_110, 0, x_97); +lean_ctor_set(x_110, 1, x_109); +lean_ctor_set(x_110, 2, x_108); +x_111 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticRepeat____1___closed__1; +lean_inc(x_83); +x_112 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_112, 0, x_83); +lean_ctor_set(x_112, 1, x_111); +x_113 = lean_array_push(x_106, x_112); +x_114 = lean_array_push(x_113, x_9); +x_115 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_115, 0, x_97); +lean_ctor_set(x_115, 1, x_4); +lean_ctor_set(x_115, 2, x_114); +x_116 = lean_array_push(x_106, x_115); +x_117 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__7; +x_118 = lean_array_push(x_116, x_117); x_119 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_119, 0, x_100); -lean_ctor_set(x_119, 1, x_118); -lean_ctor_set(x_119, 2, x_117); -x_120 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticRepeat____1___closed__1; -lean_inc(x_90); -x_121 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_121, 0, x_90); -lean_ctor_set(x_121, 1, x_120); -x_122 = lean_array_push(x_115, x_121); -x_123 = lean_array_push(x_122, x_101); -x_124 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_124, 0, x_100); -lean_ctor_set(x_124, 1, x_4); -lean_ctor_set(x_124, 2, x_123); -x_125 = lean_array_push(x_115, x_124); -x_126 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; -x_127 = lean_array_push(x_125, x_126); +lean_ctor_set(x_119, 0, x_97); +lean_ctor_set(x_119, 1, x_109); +lean_ctor_set(x_119, 2, x_118); +x_120 = lean_array_push(x_106, x_110); +x_121 = lean_array_push(x_120, x_119); +x_122 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_122, 0, x_97); +lean_ctor_set(x_122, 1, x_104); +lean_ctor_set(x_122, 2, x_121); +x_123 = lean_array_push(x_102, x_122); +x_124 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; +x_125 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_125, 0, x_97); +lean_ctor_set(x_125, 1, x_124); +lean_ctor_set(x_125, 2, x_123); +x_126 = lean_array_push(x_102, x_125); +x_127 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__1; x_128 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_128, 0, x_100); -lean_ctor_set(x_128, 1, x_118); -lean_ctor_set(x_128, 2, x_127); -x_129 = lean_array_push(x_115, x_119); +lean_ctor_set(x_128, 0, x_97); +lean_ctor_set(x_128, 1, x_127); +lean_ctor_set(x_128, 2, x_126); +x_129 = lean_array_push(x_106, x_88); +lean_inc(x_129); x_130 = lean_array_push(x_129, x_128); x_131 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_131, 0, x_100); -lean_ctor_set(x_131, 1, x_113); +lean_ctor_set(x_131, 0, x_97); +lean_ctor_set(x_131, 1, x_109); lean_ctor_set(x_131, 2, x_130); -x_132 = lean_array_push(x_98, x_131); -x_133 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__5; -x_134 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_134, 0, x_100); -lean_ctor_set(x_134, 1, x_133); -lean_ctor_set(x_134, 2, x_132); -x_135 = lean_array_push(x_98, x_134); +x_132 = l_Lean_Parser_Tactic_skip___closed__1; +x_133 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_133, 0, x_83); +lean_ctor_set(x_133, 1, x_132); +x_134 = lean_array_push(x_102, x_133); +x_135 = l_Lean_Parser_Tactic_skip___closed__2; x_136 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_136, 0, x_100); -lean_ctor_set(x_136, 1, x_10); -lean_ctor_set(x_136, 2, x_135); -x_137 = lean_array_push(x_115, x_95); -lean_inc(x_137); -x_138 = lean_array_push(x_137, x_136); +lean_ctor_set(x_136, 0, x_97); +lean_ctor_set(x_136, 1, x_135); +lean_ctor_set(x_136, 2, x_134); +x_137 = lean_array_push(x_106, x_136); +x_138 = lean_array_push(x_137, x_117); x_139 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_139, 0, x_100); -lean_ctor_set(x_139, 1, x_118); +lean_ctor_set(x_139, 0, x_97); +lean_ctor_set(x_139, 1, x_109); lean_ctor_set(x_139, 2, x_138); -x_140 = l_Lean_Parser_Tactic_skip___closed__1; -x_141 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_141, 0, x_90); -lean_ctor_set(x_141, 1, x_140); -x_142 = lean_array_push(x_98, x_141); -x_143 = l_Lean_Parser_Tactic_skip___closed__2; -x_144 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_144, 0, x_100); -lean_ctor_set(x_144, 1, x_143); -lean_ctor_set(x_144, 2, x_142); -x_145 = lean_array_push(x_115, x_144); -x_146 = lean_array_push(x_145, x_126); +x_140 = lean_array_push(x_102, x_139); +x_141 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_141, 0, x_97); +lean_ctor_set(x_141, 1, x_104); +lean_ctor_set(x_141, 2, x_140); +x_142 = lean_array_push(x_102, x_141); +x_143 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_143, 0, x_97); +lean_ctor_set(x_143, 1, x_124); +lean_ctor_set(x_143, 2, x_142); +x_144 = lean_array_push(x_102, x_143); +x_145 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_145, 0, x_97); +lean_ctor_set(x_145, 1, x_127); +lean_ctor_set(x_145, 2, x_144); +x_146 = lean_array_push(x_129, x_145); x_147 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_147, 0, x_100); -lean_ctor_set(x_147, 1, x_118); +lean_ctor_set(x_147, 0, x_97); +lean_ctor_set(x_147, 1, x_109); lean_ctor_set(x_147, 2, x_146); -x_148 = lean_array_push(x_98, x_147); -x_149 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_149, 0, x_100); -lean_ctor_set(x_149, 1, x_113); -lean_ctor_set(x_149, 2, x_148); -x_150 = lean_array_push(x_98, x_149); -x_151 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_151, 0, x_100); -lean_ctor_set(x_151, 1, x_133); -lean_ctor_set(x_151, 2, x_150); -x_152 = lean_array_push(x_98, x_151); -x_153 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_153, 0, x_100); -lean_ctor_set(x_153, 1, x_10); -lean_ctor_set(x_153, 2, x_152); -x_154 = lean_array_push(x_137, x_153); -x_155 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_155, 0, x_100); -lean_ctor_set(x_155, 1, x_118); -lean_ctor_set(x_155, 2, x_154); -x_156 = lean_array_push(x_115, x_139); -x_157 = lean_array_push(x_156, x_155); -x_158 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_158, 0, x_100); -lean_ctor_set(x_158, 1, x_113); -lean_ctor_set(x_158, 2, x_157); -x_159 = lean_array_push(x_115, x_93); -x_160 = lean_array_push(x_159, x_158); -x_161 = l_Lean_Parser_Tactic_first___closed__2; -x_162 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_162, 0, x_100); -lean_ctor_set(x_162, 1, x_161); -lean_ctor_set(x_162, 2, x_160); -x_163 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_163, 0, x_162); -lean_ctor_set(x_163, 1, x_91); -return x_163; -} +x_148 = lean_array_push(x_106, x_131); +x_149 = lean_array_push(x_148, x_147); +x_150 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_150, 0, x_97); +lean_ctor_set(x_150, 1, x_104); +lean_ctor_set(x_150, 2, x_149); +x_151 = lean_array_push(x_106, x_86); +x_152 = lean_array_push(x_151, x_150); +x_153 = l_Lean_Parser_Tactic_first___closed__2; +x_154 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_154, 0, x_97); +lean_ctor_set(x_154, 1, x_153); +lean_ctor_set(x_154, 2, x_152); +x_155 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_155, 0, x_154); +lean_ctor_set(x_155, 1, x_84); +return x_155; } } } @@ -15273,7 +15519,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticTrivial___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_tacticTrivial___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -15333,7 +15579,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_split___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_split___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -15363,7 +15609,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_split___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_intro___closed__15; x_3 = l_Lean_Parser_Tactic_rename___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -15389,7 +15635,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_split___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_split___closed__4; x_3 = l_Lean_Parser_Tactic_split___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); @@ -15403,7 +15649,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_split___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_split___closed__7; x_3 = l_Lean_Parser_Tactic_change___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); @@ -15447,7 +15693,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_dbgTrace___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_dbgTrace___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -15477,7 +15723,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_dbgTrace___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_dbgTrace___closed__4; x_3 = l_Lean_Parser_Tactic_traceMessage___closed__7; x_4 = lean_alloc_ctor(2, 3, 0); @@ -15521,7 +15767,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticStop_____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_tacticStop_____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -15563,7 +15809,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticStop_____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_tacticStop_____closed__4; x_3 = l_Lean_Parser_Tactic_tacticStop_____closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -15629,7 +15875,7 @@ x_13 = l_Lean_Parser_Tactic_tacticSorry___closed__3; x_14 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_14, 0, x_10); lean_ctor_set(x_14, 1, x_13); -x_15 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_15 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_16 = lean_array_push(x_15, x_14); x_17 = lean_box(2); x_18 = l_Lean_Parser_Tactic_tacticSorry___closed__2; @@ -15637,9 +15883,9 @@ x_19 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_19, 0, x_17); lean_ctor_set(x_19, 1, x_18); lean_ctor_set(x_19, 2, x_16); -x_20 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_20 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_21 = lean_array_push(x_20, x_19); -x_22 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; +x_22 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__7; x_23 = lean_array_push(x_21, x_22); x_24 = l_Lean_Parser_Tactic_first___closed__8; x_25 = lean_alloc_ctor(1, 3, 0); @@ -15647,19 +15893,19 @@ lean_ctor_set(x_25, 0, x_17); lean_ctor_set(x_25, 1, x_24); lean_ctor_set(x_25, 2, x_23); x_26 = lean_array_push(x_15, x_25); -x_27 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_27 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_28 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_28, 0, x_17); lean_ctor_set(x_28, 1, x_27); lean_ctor_set(x_28, 2, x_26); x_29 = lean_array_push(x_15, x_28); -x_30 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__5; +x_30 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; x_31 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_31, 0, x_17); lean_ctor_set(x_31, 1, x_30); lean_ctor_set(x_31, 2, x_29); x_32 = lean_array_push(x_15, x_31); -x_33 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__5; +x_33 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__1; x_34 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_34, 0, x_17); lean_ctor_set(x_34, 1, x_33); @@ -15677,7 +15923,7 @@ lean_ctor_set(x_40, 0, x_17); lean_ctor_set(x_40, 1, x_27); lean_ctor_set(x_40, 2, x_39); x_41 = lean_array_push(x_15, x_40); -x_42 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_42 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_43 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_43, 0, x_17); lean_ctor_set(x_43, 1, x_42); @@ -15702,7 +15948,7 @@ x_48 = l_Lean_Parser_Tactic_tacticSorry___closed__3; x_49 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_49, 0, x_44); lean_ctor_set(x_49, 1, x_48); -x_50 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_50 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_51 = lean_array_push(x_50, x_49); x_52 = lean_box(2); x_53 = l_Lean_Parser_Tactic_tacticSorry___closed__2; @@ -15710,9 +15956,9 @@ x_54 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_54, 0, x_52); lean_ctor_set(x_54, 1, x_53); lean_ctor_set(x_54, 2, x_51); -x_55 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_55 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_56 = lean_array_push(x_55, x_54); -x_57 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; +x_57 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__7; x_58 = lean_array_push(x_56, x_57); x_59 = l_Lean_Parser_Tactic_first___closed__8; x_60 = lean_alloc_ctor(1, 3, 0); @@ -15720,19 +15966,19 @@ lean_ctor_set(x_60, 0, x_52); lean_ctor_set(x_60, 1, x_59); lean_ctor_set(x_60, 2, x_58); x_61 = lean_array_push(x_50, x_60); -x_62 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_62 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_63 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_63, 0, x_52); lean_ctor_set(x_63, 1, x_62); lean_ctor_set(x_63, 2, x_61); x_64 = lean_array_push(x_50, x_63); -x_65 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__5; +x_65 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; x_66 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_66, 0, x_52); lean_ctor_set(x_66, 1, x_65); lean_ctor_set(x_66, 2, x_64); x_67 = lean_array_push(x_50, x_66); -x_68 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__5; +x_68 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__1; x_69 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_69, 0, x_52); lean_ctor_set(x_69, 1, x_68); @@ -15750,7 +15996,7 @@ lean_ctor_set(x_75, 0, x_52); lean_ctor_set(x_75, 1, x_62); lean_ctor_set(x_75, 2, x_74); x_76 = lean_array_push(x_50, x_75); -x_77 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_77 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_78 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_78, 0, x_52); lean_ctor_set(x_78, 1, x_77); @@ -15775,7 +16021,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_specialize___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_specialize___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -15805,7 +16051,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_specialize___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_specialize___closed__4; x_3 = l_Lean_Parser_Tactic_rename___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -15866,7 +16112,7 @@ x_11 = l_Lean_Parser_Tactic_assumption___closed__1; x_12 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_12, 0, x_10); lean_ctor_set(x_12, 1, x_11); -x_13 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_13 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_14 = lean_array_push(x_13, x_12); x_15 = lean_box(2); x_16 = l_Lean_Parser_Tactic_assumption___closed__2; @@ -15889,7 +16135,7 @@ x_20 = l_Lean_Parser_Tactic_assumption___closed__1; x_21 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_21, 0, x_18); lean_ctor_set(x_21, 1, x_20); -x_22 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_22 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_23 = lean_array_push(x_22, x_21); x_24 = lean_box(2); x_25 = l_Lean_Parser_Tactic_assumption___closed__2; @@ -15934,7 +16180,7 @@ x_11 = l_Lean_Parser_Tactic_tacticRfl___closed__3; x_12 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_12, 0, x_10); lean_ctor_set(x_12, 1, x_11); -x_13 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_13 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_14 = lean_array_push(x_13, x_12); x_15 = lean_box(2); x_16 = l_Lean_Parser_Tactic_tacticRfl___closed__2; @@ -15957,7 +16203,7 @@ x_20 = l_Lean_Parser_Tactic_tacticRfl___closed__3; x_21 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_21, 0, x_18); lean_ctor_set(x_21, 1, x_20); -x_22 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_22 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_23 = lean_array_push(x_22, x_21); x_24 = lean_box(2); x_25 = l_Lean_Parser_Tactic_tacticRfl___closed__2; @@ -16002,7 +16248,7 @@ x_11 = l_Lean_Parser_Tactic_contradiction___closed__1; x_12 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_12, 0, x_10); lean_ctor_set(x_12, 1, x_11); -x_13 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_13 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_14 = lean_array_push(x_13, x_12); x_15 = lean_box(2); x_16 = l_Lean_Parser_Tactic_contradiction___closed__2; @@ -16025,7 +16271,7 @@ x_20 = l_Lean_Parser_Tactic_contradiction___closed__1; x_21 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_21, 0, x_18); lean_ctor_set(x_21, 1, x_20); -x_22 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_22 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_23 = lean_array_push(x_22, x_21); x_24 = lean_box(2); x_25 = l_Lean_Parser_Tactic_contradiction___closed__2; @@ -16053,7 +16299,7 @@ static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRu _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTrivial__4___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -16088,7 +16334,7 @@ x_11 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__ x_12 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_12, 0, x_10); lean_ctor_set(x_12, 1, x_11); -x_13 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_13 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_14 = lean_array_push(x_13, x_12); x_15 = lean_box(2); x_16 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTrivial__4___closed__2; @@ -16111,7 +16357,7 @@ x_20 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__ x_21 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_21, 0, x_18); lean_ctor_set(x_21, 1, x_20); -x_22 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_22 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_23 = lean_array_push(x_22, x_21); x_24 = lean_box(2); x_25 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTrivial__4___closed__2; @@ -16255,7 +16501,7 @@ lean_ctor_set(x_19, 0, x_10); lean_ctor_set(x_19, 1, x_17); lean_ctor_set(x_19, 2, x_16); lean_ctor_set(x_19, 3, x_18); -x_20 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_20 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_21 = lean_array_push(x_20, x_14); x_22 = lean_array_push(x_21, x_19); x_23 = lean_box(2); @@ -16294,7 +16540,7 @@ lean_ctor_set(x_36, 0, x_26); lean_ctor_set(x_36, 1, x_34); lean_ctor_set(x_36, 2, x_33); lean_ctor_set(x_36, 3, x_35); -x_37 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_37 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_38 = lean_array_push(x_37, x_31); x_39 = lean_array_push(x_38, x_36); x_40 = lean_box(2); @@ -16448,7 +16694,7 @@ lean_ctor_set(x_19, 0, x_10); lean_ctor_set(x_19, 1, x_17); lean_ctor_set(x_19, 2, x_16); lean_ctor_set(x_19, 3, x_18); -x_20 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_20 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_21 = lean_array_push(x_20, x_14); x_22 = lean_array_push(x_21, x_19); x_23 = lean_box(2); @@ -16466,7 +16712,7 @@ x_28 = l_Lean_Parser_Tactic_tacticTrivial___closed__3; x_29 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_29, 0, x_10); lean_ctor_set(x_29, 1, x_28); -x_30 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_30 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_31 = lean_array_push(x_30, x_29); x_32 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_32, 0, x_23); @@ -16512,7 +16758,7 @@ lean_ctor_set(x_49, 0, x_39); lean_ctor_set(x_49, 1, x_47); lean_ctor_set(x_49, 2, x_46); lean_ctor_set(x_49, 3, x_48); -x_50 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_50 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_51 = lean_array_push(x_50, x_44); x_52 = lean_array_push(x_51, x_49); x_53 = lean_box(2); @@ -16530,7 +16776,7 @@ x_58 = l_Lean_Parser_Tactic_tacticTrivial___closed__3; x_59 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_59, 0, x_39); lean_ctor_set(x_59, 1, x_58); -x_60 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_60 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_61 = lean_array_push(x_60, x_59); x_62 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_62, 0, x_53); @@ -16565,7 +16811,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticUnhygienic_____closed__2() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_tacticUnhygienic_____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -16595,7 +16841,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticUnhygienic_____closed__5() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_tacticUnhygienic_____closed__4; x_3 = l_Lean_Parser_Tactic_case___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); @@ -16670,7 +16916,7 @@ static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRu _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__22; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__20; x_2 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticUnhygienic____1___closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -16736,7 +16982,7 @@ x_24 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__ x_25 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_25, 0, x_12); lean_ctor_set(x_25, 1, x_24); -x_26 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__7; +x_26 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__6; x_27 = lean_array_push(x_26, x_16); x_28 = lean_array_push(x_27, x_21); x_29 = lean_array_push(x_28, x_23); @@ -16748,15 +16994,15 @@ x_34 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_34, 0, x_32); lean_ctor_set(x_34, 1, x_33); lean_ctor_set(x_34, 2, x_31); -x_35 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_35 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_36 = lean_array_push(x_35, x_34); -x_37 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_37 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_38 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_38, 0, x_32); lean_ctor_set(x_38, 1, x_37); lean_ctor_set(x_38, 2, x_36); x_39 = lean_array_push(x_35, x_38); -x_40 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_40 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_41 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_41, 0, x_32); lean_ctor_set(x_41, 1, x_40); @@ -16801,7 +17047,7 @@ x_55 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__ x_56 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_56, 0, x_42); lean_ctor_set(x_56, 1, x_55); -x_57 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__7; +x_57 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__6; x_58 = lean_array_push(x_57, x_47); x_59 = lean_array_push(x_58, x_52); x_60 = lean_array_push(x_59, x_54); @@ -16813,15 +17059,15 @@ x_65 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_65, 0, x_63); lean_ctor_set(x_65, 1, x_64); lean_ctor_set(x_65, 2, x_62); -x_66 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_66 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_67 = lean_array_push(x_66, x_65); -x_68 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_68 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_69 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_69, 0, x_63); lean_ctor_set(x_69, 1, x_68); lean_ctor_set(x_69, 2, x_67); x_70 = lean_array_push(x_66, x_69); -x_71 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_71 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_72 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_72, 0, x_63); lean_ctor_set(x_72, 1, x_71); @@ -16846,7 +17092,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_fail___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_fail___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -16888,7 +17134,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_fail___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_fail___closed__4; x_3 = l_Lean_Parser_Tactic_fail___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -16932,7 +17178,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_checkpoint___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_checkpoint___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -16962,7 +17208,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_checkpoint___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_checkpoint___closed__4; x_3 = l_Lean_Parser_Tactic_case___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); @@ -17006,7 +17252,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_save___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_save___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -17075,7 +17321,7 @@ x_11 = l_Lean_Parser_Tactic_skip___closed__1; x_12 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_12, 0, x_10); lean_ctor_set(x_12, 1, x_11); -x_13 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_13 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_14 = lean_array_push(x_13, x_12); x_15 = lean_box(2); x_16 = l_Lean_Parser_Tactic_skip___closed__2; @@ -17084,13 +17330,13 @@ lean_ctor_set(x_17, 0, x_15); lean_ctor_set(x_17, 1, x_16); lean_ctor_set(x_17, 2, x_14); x_18 = lean_array_push(x_13, x_17); -x_19 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_19 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_20 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_20, 0, x_15); lean_ctor_set(x_20, 1, x_19); lean_ctor_set(x_20, 2, x_18); x_21 = lean_array_push(x_13, x_20); -x_22 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_22 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_23 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_23, 0, x_15); lean_ctor_set(x_23, 1, x_22); @@ -17110,7 +17356,7 @@ x_26 = l_Lean_Parser_Tactic_skip___closed__1; x_27 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_27, 0, x_24); lean_ctor_set(x_27, 1, x_26); -x_28 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_28 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_29 = lean_array_push(x_28, x_27); x_30 = lean_box(2); x_31 = l_Lean_Parser_Tactic_skip___closed__2; @@ -17119,13 +17365,13 @@ lean_ctor_set(x_32, 0, x_30); lean_ctor_set(x_32, 1, x_31); lean_ctor_set(x_32, 2, x_29); x_33 = lean_array_push(x_28, x_32); -x_34 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_34 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_35 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_35, 0, x_30); lean_ctor_set(x_35, 1, x_34); lean_ctor_set(x_35, 2, x_33); x_36 = lean_array_push(x_28, x_35); -x_37 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__2; +x_37 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; x_38 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_38, 0, x_30); lean_ctor_set(x_38, 1, x_37); @@ -17150,7 +17396,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_sleep___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_sleep___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -17172,7 +17418,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_sleep___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_sleep___closed__3; x_3 = l_Lean_Parser_Tactic_rotateLeft___closed__9; x_4 = lean_alloc_ctor(2, 3, 0); @@ -17216,7 +17462,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticExists___x2c_x2c___closed__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__6; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; x_2 = l_Lean_Parser_Tactic_tacticExists___x2c_x2c___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -17246,7 +17492,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticExists___x2c_x2c___closed__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Tactic_tacticExists___x2c_x2c___closed__4; x_3 = l_Lean_Parser_Tactic_induction___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -17359,7 +17605,7 @@ lean_inc(x_13); x_19 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_19, 0, x_13); lean_ctor_set(x_19, 1, x_18); -x_20 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__3; +x_20 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__4; x_21 = l_Array_appendCore___rarg(x_20, x_10); lean_dec(x_10); x_22 = l_Lean_Parser_Tactic_rwRuleSeq___closed__7; @@ -17373,12 +17619,12 @@ lean_inc(x_13); x_26 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_26, 0, x_13); lean_ctor_set(x_26, 1, x_25); -x_27 = l_Lean_Parser_Tactic_intros___closed__10; +x_27 = l_Lean_binderIdent___closed__10; lean_inc(x_13); x_28 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_28, 0, x_13); lean_ctor_set(x_28, 1, x_27); -x_29 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_29 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_30 = lean_array_push(x_29, x_28); x_31 = lean_box(2); x_32 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticAdmit__1___closed__15; @@ -17386,7 +17632,7 @@ x_33 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_33, 0, x_31); lean_ctor_set(x_33, 1, x_32); lean_ctor_set(x_33, 2, x_30); -x_34 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_34 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_35 = lean_array_push(x_34, x_26); x_36 = lean_array_push(x_35, x_33); x_37 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__5; @@ -17395,7 +17641,7 @@ lean_ctor_set(x_38, 0, x_31); lean_ctor_set(x_38, 1, x_37); lean_ctor_set(x_38, 2, x_36); x_39 = lean_array_push(x_24, x_38); -x_40 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_40 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_41 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_41, 0, x_31); lean_ctor_set(x_41, 1, x_40); @@ -17421,7 +17667,7 @@ x_53 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_53, 0, x_31); lean_ctor_set(x_53, 1, x_52); lean_ctor_set(x_53, 2, x_51); -x_54 = l_Lean_Parser_Tactic_rwWithRfl___closed__2; +x_54 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__2; lean_inc(x_13); x_55 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_55, 0, x_13); @@ -17438,7 +17684,7 @@ x_61 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_61, 0, x_31); lean_ctor_set(x_61, 1, x_60); lean_ctor_set(x_61, 2, x_59); -x_62 = l_Lean_Parser_Tactic_rwWithRfl___closed__3; +x_62 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__3; lean_inc(x_13); x_63 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_63, 0, x_13); @@ -17455,7 +17701,7 @@ lean_ctor_set(x_68, 0, x_31); lean_ctor_set(x_68, 1, x_67); lean_ctor_set(x_68, 2, x_66); x_69 = lean_array_push(x_34, x_68); -x_70 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; +x_70 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__7; x_71 = lean_array_push(x_69, x_70); x_72 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_72, 0, x_31); @@ -17467,13 +17713,13 @@ lean_ctor_set(x_74, 0, x_31); lean_ctor_set(x_74, 1, x_40); lean_ctor_set(x_74, 2, x_73); x_75 = lean_array_push(x_29, x_74); -x_76 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__5; +x_76 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; x_77 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_77, 0, x_31); lean_ctor_set(x_77, 1, x_76); lean_ctor_set(x_77, 2, x_75); x_78 = lean_array_push(x_29, x_77); -x_79 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__5; +x_79 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__1; x_80 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_80, 0, x_31); lean_ctor_set(x_80, 1, x_79); @@ -17545,7 +17791,7 @@ lean_inc(x_102); x_109 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_109, 0, x_102); lean_ctor_set(x_109, 1, x_108); -x_110 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__3; +x_110 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__4; x_111 = l_Array_appendCore___rarg(x_110, x_10); lean_dec(x_10); x_112 = l_Lean_Parser_Tactic_rwRuleSeq___closed__7; @@ -17559,12 +17805,12 @@ lean_inc(x_102); x_116 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_116, 0, x_102); lean_ctor_set(x_116, 1, x_115); -x_117 = l_Lean_Parser_Tactic_intros___closed__10; +x_117 = l_Lean_binderIdent___closed__10; lean_inc(x_102); x_118 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_118, 0, x_102); lean_ctor_set(x_118, 1, x_117); -x_119 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_119 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_120 = lean_array_push(x_119, x_118); x_121 = lean_box(2); x_122 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticAdmit__1___closed__15; @@ -17572,7 +17818,7 @@ x_123 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_123, 0, x_121); lean_ctor_set(x_123, 1, x_122); lean_ctor_set(x_123, 2, x_120); -x_124 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_124 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_125 = lean_array_push(x_124, x_116); x_126 = lean_array_push(x_125, x_123); x_127 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__5; @@ -17581,7 +17827,7 @@ lean_ctor_set(x_128, 0, x_121); lean_ctor_set(x_128, 1, x_127); lean_ctor_set(x_128, 2, x_126); x_129 = lean_array_push(x_114, x_128); -x_130 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_130 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_131 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_131, 0, x_121); lean_ctor_set(x_131, 1, x_130); @@ -17607,7 +17853,7 @@ x_143 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_143, 0, x_121); lean_ctor_set(x_143, 1, x_142); lean_ctor_set(x_143, 2, x_141); -x_144 = l_Lean_Parser_Tactic_rwWithRfl___closed__2; +x_144 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__2; lean_inc(x_102); x_145 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_145, 0, x_102); @@ -17624,7 +17870,7 @@ x_151 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_151, 0, x_121); lean_ctor_set(x_151, 1, x_150); lean_ctor_set(x_151, 2, x_149); -x_152 = l_Lean_Parser_Tactic_rwWithRfl___closed__3; +x_152 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__3; lean_inc(x_102); x_153 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_153, 0, x_102); @@ -17641,7 +17887,7 @@ lean_ctor_set(x_158, 0, x_121); lean_ctor_set(x_158, 1, x_157); lean_ctor_set(x_158, 2, x_156); x_159 = lean_array_push(x_124, x_158); -x_160 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; +x_160 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__7; x_161 = lean_array_push(x_159, x_160); x_162 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_162, 0, x_121); @@ -17653,13 +17899,13 @@ lean_ctor_set(x_164, 0, x_121); lean_ctor_set(x_164, 1, x_130); lean_ctor_set(x_164, 2, x_163); x_165 = lean_array_push(x_119, x_164); -x_166 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__5; +x_166 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; x_167 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_167, 0, x_121); lean_ctor_set(x_167, 1, x_166); lean_ctor_set(x_167, 2, x_165); x_168 = lean_array_push(x_119, x_167); -x_169 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__5; +x_169 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__1; x_170 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_170, 0, x_121); lean_ctor_set(x_170, 1, x_169); @@ -17725,7 +17971,7 @@ static lean_object* _init_l_Lean_Parser_Attr_simp___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__4; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__2; x_2 = l_Lean_Parser_Attr_simp___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -17757,7 +18003,7 @@ static lean_object* _init_l_Lean_Parser_Attr_simp___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Attr_simp___closed__4; x_3 = l_Lean_Parser_Tactic_simpLemma___closed__4; x_4 = lean_alloc_ctor(2, 3, 0); @@ -17813,7 +18059,7 @@ static lean_object* _init_l_Lean_Parser_Attr_simp___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_Lean_Parser_Attr_simp___closed__5; x_3 = l_Lean_Parser_Attr_simp___closed__9; x_4 = lean_alloc_ctor(2, 3, 0); @@ -17885,7 +18131,7 @@ static lean_object* _init_l_term_u2039___u203a___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_term_u2039___u203a___closed__4; x_3 = l_Lean_Parser_Tactic_rename___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); @@ -17917,7 +18163,7 @@ static lean_object* _init_l_term_u2039___u203a___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__10; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__8; x_2 = l_term_u2039___u203a___closed__5; x_3 = l_term_u2039___u203a___closed__7; x_4 = lean_alloc_ctor(2, 3, 0); @@ -18048,7 +18294,7 @@ lean_inc(x_12); x_18 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_18, 0, x_12); lean_ctor_set(x_18, 1, x_17); -x_19 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_19 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_20 = lean_array_push(x_19, x_18); x_21 = lean_box(2); x_22 = l_Lean_Parser_Tactic_assumption___closed__2; @@ -18056,9 +18302,9 @@ x_23 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_23, 0, x_21); lean_ctor_set(x_23, 1, x_22); lean_ctor_set(x_23, 2, x_20); -x_24 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_24 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_25 = lean_array_push(x_24, x_23); -x_26 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; +x_26 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__7; x_27 = lean_array_push(x_25, x_26); x_28 = l_Lean_Parser_Tactic_first___closed__8; x_29 = lean_alloc_ctor(1, 3, 0); @@ -18066,19 +18312,19 @@ lean_ctor_set(x_29, 0, x_21); lean_ctor_set(x_29, 1, x_28); lean_ctor_set(x_29, 2, x_27); x_30 = lean_array_push(x_19, x_29); -x_31 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_31 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_32 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_32, 0, x_21); lean_ctor_set(x_32, 1, x_31); lean_ctor_set(x_32, 2, x_30); x_33 = lean_array_push(x_19, x_32); -x_34 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__5; +x_34 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; x_35 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_35, 0, x_21); lean_ctor_set(x_35, 1, x_34); lean_ctor_set(x_35, 2, x_33); x_36 = lean_array_push(x_19, x_35); -x_37 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__5; +x_37 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__1; x_38 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_38, 0, x_21); lean_ctor_set(x_38, 1, x_37); @@ -18152,7 +18398,7 @@ lean_inc(x_62); x_69 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_69, 0, x_62); lean_ctor_set(x_69, 1, x_68); -x_70 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6; +x_70 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1; x_71 = lean_array_push(x_70, x_69); x_72 = lean_box(2); x_73 = l_Lean_Parser_Tactic_assumption___closed__2; @@ -18160,9 +18406,9 @@ x_74 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_74, 0, x_72); lean_ctor_set(x_74, 1, x_73); lean_ctor_set(x_74, 2, x_71); -x_75 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +x_75 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__4; x_76 = lean_array_push(x_75, x_74); -x_77 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; +x_77 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__7; x_78 = lean_array_push(x_76, x_77); x_79 = l_Lean_Parser_Tactic_first___closed__8; x_80 = lean_alloc_ctor(1, 3, 0); @@ -18170,19 +18416,19 @@ lean_ctor_set(x_80, 0, x_72); lean_ctor_set(x_80, 1, x_79); lean_ctor_set(x_80, 2, x_78); x_81 = lean_array_push(x_70, x_80); -x_82 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2; +x_82 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3; x_83 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_83, 0, x_72); lean_ctor_set(x_83, 1, x_82); lean_ctor_set(x_83, 2, x_81); x_84 = lean_array_push(x_70, x_83); -x_85 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__5; +x_85 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6; x_86 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_86, 0, x_72); lean_ctor_set(x_86, 1, x_85); lean_ctor_set(x_86, 2, x_84); x_87 = lean_array_push(x_70, x_86); -x_88 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__5; +x_88 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__1; x_89 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_89, 0, x_72); lean_ctor_set(x_89, 1, x_88); @@ -18247,6 +18493,34 @@ _G_initialized = true; res = initialize_Init_Notation(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +l_Lean_binderIdent___closed__1 = _init_l_Lean_binderIdent___closed__1(); +lean_mark_persistent(l_Lean_binderIdent___closed__1); +l_Lean_binderIdent___closed__2 = _init_l_Lean_binderIdent___closed__2(); +lean_mark_persistent(l_Lean_binderIdent___closed__2); +l_Lean_binderIdent___closed__3 = _init_l_Lean_binderIdent___closed__3(); +lean_mark_persistent(l_Lean_binderIdent___closed__3); +l_Lean_binderIdent___closed__4 = _init_l_Lean_binderIdent___closed__4(); +lean_mark_persistent(l_Lean_binderIdent___closed__4); +l_Lean_binderIdent___closed__5 = _init_l_Lean_binderIdent___closed__5(); +lean_mark_persistent(l_Lean_binderIdent___closed__5); +l_Lean_binderIdent___closed__6 = _init_l_Lean_binderIdent___closed__6(); +lean_mark_persistent(l_Lean_binderIdent___closed__6); +l_Lean_binderIdent___closed__7 = _init_l_Lean_binderIdent___closed__7(); +lean_mark_persistent(l_Lean_binderIdent___closed__7); +l_Lean_binderIdent___closed__8 = _init_l_Lean_binderIdent___closed__8(); +lean_mark_persistent(l_Lean_binderIdent___closed__8); +l_Lean_binderIdent___closed__9 = _init_l_Lean_binderIdent___closed__9(); +lean_mark_persistent(l_Lean_binderIdent___closed__9); +l_Lean_binderIdent___closed__10 = _init_l_Lean_binderIdent___closed__10(); +lean_mark_persistent(l_Lean_binderIdent___closed__10); +l_Lean_binderIdent___closed__11 = _init_l_Lean_binderIdent___closed__11(); +lean_mark_persistent(l_Lean_binderIdent___closed__11); +l_Lean_binderIdent___closed__12 = _init_l_Lean_binderIdent___closed__12(); +lean_mark_persistent(l_Lean_binderIdent___closed__12); +l_Lean_binderIdent___closed__13 = _init_l_Lean_binderIdent___closed__13(); +lean_mark_persistent(l_Lean_binderIdent___closed__13); +l_Lean_binderIdent = _init_l_Lean_binderIdent(); +lean_mark_persistent(l_Lean_binderIdent); l_Lean_Parser_Tactic_withAnnotateState___closed__1 = _init_l_Lean_Parser_Tactic_withAnnotateState___closed__1(); lean_mark_persistent(l_Lean_Parser_Tactic_withAnnotateState___closed__1); l_Lean_Parser_Tactic_withAnnotateState___closed__2 = _init_l_Lean_Parser_Tactic_withAnnotateState___closed__2(); @@ -18293,10 +18567,6 @@ l_Lean_Parser_Tactic_withAnnotateState___closed__22 = _init_l_Lean_Parser_Tactic lean_mark_persistent(l_Lean_Parser_Tactic_withAnnotateState___closed__22); l_Lean_Parser_Tactic_withAnnotateState___closed__23 = _init_l_Lean_Parser_Tactic_withAnnotateState___closed__23(); lean_mark_persistent(l_Lean_Parser_Tactic_withAnnotateState___closed__23); -l_Lean_Parser_Tactic_withAnnotateState___closed__24 = _init_l_Lean_Parser_Tactic_withAnnotateState___closed__24(); -lean_mark_persistent(l_Lean_Parser_Tactic_withAnnotateState___closed__24); -l_Lean_Parser_Tactic_withAnnotateState___closed__25 = _init_l_Lean_Parser_Tactic_withAnnotateState___closed__25(); -lean_mark_persistent(l_Lean_Parser_Tactic_withAnnotateState___closed__25); l_Lean_Parser_Tactic_withAnnotateState = _init_l_Lean_Parser_Tactic_withAnnotateState(); lean_mark_persistent(l_Lean_Parser_Tactic_withAnnotateState); l_Lean_Parser_Tactic_intro___closed__1 = _init_l_Lean_Parser_Tactic_intro___closed__1(); @@ -18361,22 +18631,6 @@ l_Lean_Parser_Tactic_intros___closed__7 = _init_l_Lean_Parser_Tactic_intros___cl lean_mark_persistent(l_Lean_Parser_Tactic_intros___closed__7); l_Lean_Parser_Tactic_intros___closed__8 = _init_l_Lean_Parser_Tactic_intros___closed__8(); lean_mark_persistent(l_Lean_Parser_Tactic_intros___closed__8); -l_Lean_Parser_Tactic_intros___closed__9 = _init_l_Lean_Parser_Tactic_intros___closed__9(); -lean_mark_persistent(l_Lean_Parser_Tactic_intros___closed__9); -l_Lean_Parser_Tactic_intros___closed__10 = _init_l_Lean_Parser_Tactic_intros___closed__10(); -lean_mark_persistent(l_Lean_Parser_Tactic_intros___closed__10); -l_Lean_Parser_Tactic_intros___closed__11 = _init_l_Lean_Parser_Tactic_intros___closed__11(); -lean_mark_persistent(l_Lean_Parser_Tactic_intros___closed__11); -l_Lean_Parser_Tactic_intros___closed__12 = _init_l_Lean_Parser_Tactic_intros___closed__12(); -lean_mark_persistent(l_Lean_Parser_Tactic_intros___closed__12); -l_Lean_Parser_Tactic_intros___closed__13 = _init_l_Lean_Parser_Tactic_intros___closed__13(); -lean_mark_persistent(l_Lean_Parser_Tactic_intros___closed__13); -l_Lean_Parser_Tactic_intros___closed__14 = _init_l_Lean_Parser_Tactic_intros___closed__14(); -lean_mark_persistent(l_Lean_Parser_Tactic_intros___closed__14); -l_Lean_Parser_Tactic_intros___closed__15 = _init_l_Lean_Parser_Tactic_intros___closed__15(); -lean_mark_persistent(l_Lean_Parser_Tactic_intros___closed__15); -l_Lean_Parser_Tactic_intros___closed__16 = _init_l_Lean_Parser_Tactic_intros___closed__16(); -lean_mark_persistent(l_Lean_Parser_Tactic_intros___closed__16); l_Lean_Parser_Tactic_intros = _init_l_Lean_Parser_Tactic_intros(); lean_mark_persistent(l_Lean_Parser_Tactic_intros); l_Lean_Parser_Tactic_rename___closed__1 = _init_l_Lean_Parser_Tactic_rename___closed__1(); @@ -18597,40 +18851,36 @@ l_Lean_Parser_Tactic_case_x27___closed__9 = _init_l_Lean_Parser_Tactic_case_x27_ lean_mark_persistent(l_Lean_Parser_Tactic_case_x27___closed__9); l_Lean_Parser_Tactic_case_x27 = _init_l_Lean_Parser_Tactic_case_x27(); lean_mark_persistent(l_Lean_Parser_Tactic_case_x27); -l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__1 = _init_l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__1(); -lean_mark_persistent(l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__1); -l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__2 = _init_l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__2(); -lean_mark_persistent(l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__2); -l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__3 = _init_l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__3(); -lean_mark_persistent(l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__3); -l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__4 = _init_l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__4(); -lean_mark_persistent(l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__4); -l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__5 = _init_l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__5(); -lean_mark_persistent(l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__5); -l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__6 = _init_l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__6(); -lean_mark_persistent(l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__6); -l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__7 = _init_l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__7(); -lean_mark_persistent(l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__7); -l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__8 = _init_l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__8(); -lean_mark_persistent(l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__8); -l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__9 = _init_l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__9(); -lean_mark_persistent(l_Lean_Parser_Tactic_tacticNext_______x3d_x3e_____closed__9); -l_Lean_Parser_Tactic_tacticNext_______x3d_x3e__ = _init_l_Lean_Parser_Tactic_tacticNext_______x3d_x3e__(); -lean_mark_persistent(l_Lean_Parser_Tactic_tacticNext_______x3d_x3e__); -l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__1 = _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__1(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__1); -l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2 = _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__2); -l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__3 = _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__3(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__3); -l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__4 = _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__4(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__4); -l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__5 = _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__5(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__5); -l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6 = _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__6); -l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__7 = _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__7(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext_______x3d_x3e____1___closed__7); +l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__1 = _init_l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__1); +l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__2 = _init_l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__2(); +lean_mark_persistent(l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__2); +l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__3 = _init_l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__3(); +lean_mark_persistent(l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__3); +l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__4 = _init_l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__4(); +lean_mark_persistent(l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__4); +l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__5 = _init_l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__5(); +lean_mark_persistent(l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__5); +l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__6 = _init_l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__6(); +lean_mark_persistent(l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__6); +l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__7 = _init_l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__7(); +lean_mark_persistent(l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__7); +l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__8 = _init_l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__8(); +lean_mark_persistent(l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__8); +l_Lean_Parser_Tactic_tacticNext___x3d_x3e__ = _init_l_Lean_Parser_Tactic_tacticNext___x3d_x3e__(); +lean_mark_persistent(l_Lean_Parser_Tactic_tacticNext___x3d_x3e__); +l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1 = _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__1); +l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__2 = _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__2(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__2); +l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3 = _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__3); +l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__4 = _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__4(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__4); +l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__5 = _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__5(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__5); +l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__6 = _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__6(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__6); l_Lean_Parser_Tactic_allGoals___closed__1 = _init_l_Lean_Parser_Tactic_allGoals___closed__1(); lean_mark_persistent(l_Lean_Parser_Tactic_allGoals___closed__1); l_Lean_Parser_Tactic_allGoals___closed__2 = _init_l_Lean_Parser_Tactic_allGoals___closed__2(); @@ -18905,6 +19155,8 @@ l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic_ lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__5); l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6 = _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6(); lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__6); +l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__7 = _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__7(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__7); l_Lean_Parser_Tactic_tactic___x3c_x3b_x3e_____closed__1 = _init_l_Lean_Parser_Tactic_tactic___x3c_x3b_x3e_____closed__1(); lean_mark_persistent(l_Lean_Parser_Tactic_tactic___x3c_x3b_x3e_____closed__1); l_Lean_Parser_Tactic_tactic___x3c_x3b_x3e_____closed__2 = _init_l_Lean_Parser_Tactic_tactic___x3c_x3b_x3e_____closed__2(); @@ -19291,16 +19543,16 @@ l_Lean_Parser_Tactic_rwSeq___closed__8 = _init_l_Lean_Parser_Tactic_rwSeq___clos lean_mark_persistent(l_Lean_Parser_Tactic_rwSeq___closed__8); l_Lean_Parser_Tactic_rwSeq = _init_l_Lean_Parser_Tactic_rwSeq(); lean_mark_persistent(l_Lean_Parser_Tactic_rwSeq); -l_Lean_Parser_Tactic_rwWithRfl___closed__1 = _init_l_Lean_Parser_Tactic_rwWithRfl___closed__1(); -lean_mark_persistent(l_Lean_Parser_Tactic_rwWithRfl___closed__1); -l_Lean_Parser_Tactic_rwWithRfl___closed__2 = _init_l_Lean_Parser_Tactic_rwWithRfl___closed__2(); -lean_mark_persistent(l_Lean_Parser_Tactic_rwWithRfl___closed__2); -l_Lean_Parser_Tactic_rwWithRfl___closed__3 = _init_l_Lean_Parser_Tactic_rwWithRfl___closed__3(); -lean_mark_persistent(l_Lean_Parser_Tactic_rwWithRfl___closed__3); -l_Lean_Parser_Tactic_rwWithRfl___closed__4 = _init_l_Lean_Parser_Tactic_rwWithRfl___closed__4(); -lean_mark_persistent(l_Lean_Parser_Tactic_rwWithRfl___closed__4); -l_Lean_Parser_Tactic_expandRwSeq___closed__1 = _init_l_Lean_Parser_Tactic_expandRwSeq___closed__1(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandRwSeq___closed__1); +l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__1 = _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__1); +l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__2 = _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__2(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__2); +l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__3 = _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__3(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__3); +l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__4 = _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__4(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__4); +l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__5 = _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__5(); +lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__5); l_Lean_Parser_Tactic_injection___closed__1 = _init_l_Lean_Parser_Tactic_injection___closed__1(); lean_mark_persistent(l_Lean_Parser_Tactic_injection___closed__1); l_Lean_Parser_Tactic_injection___closed__2 = _init_l_Lean_Parser_Tactic_injection___closed__2(); @@ -19599,8 +19851,6 @@ l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic_ lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__5); l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__6 = _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__6(); lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__6); -l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__7 = _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__7(); -lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave____1___closed__7); l_Lean_Parser_Tactic_tacticHave_____x3a_x3d_____closed__1 = _init_l_Lean_Parser_Tactic_tacticHave_____x3a_x3d_____closed__1(); lean_mark_persistent(l_Lean_Parser_Tactic_tacticHave_____x3a_x3d_____closed__1); l_Lean_Parser_Tactic_tacticHave_____x3a_x3d_____closed__2 = _init_l_Lean_Parser_Tactic_tacticHave_____x3a_x3d_____closed__2(); @@ -19769,24 +20019,24 @@ l_Lean_Parser_Tactic_tacticHave_x27__ = _init_l_Lean_Parser_Tactic_tacticHave_x2 lean_mark_persistent(l_Lean_Parser_Tactic_tacticHave_x27__); l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave_x27____1___closed__1 = _init_l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave_x27____1___closed__1(); lean_mark_persistent(l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave_x27____1___closed__1); -l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__1 = _init_l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__1(); -lean_mark_persistent(l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__1); -l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__2 = _init_l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__2(); -lean_mark_persistent(l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__2); -l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__3 = _init_l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__3(); -lean_mark_persistent(l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__3); -l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__4 = _init_l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__4(); -lean_mark_persistent(l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__4); -l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__5 = _init_l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__5(); -lean_mark_persistent(l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__5); -l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__6 = _init_l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__6(); -lean_mark_persistent(l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__6); -l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__7 = _init_l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__7(); -lean_mark_persistent(l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__7); -l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__8 = _init_l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__8(); -lean_mark_persistent(l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d_____closed__8); -l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d__ = _init_l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d__(); -lean_mark_persistent(l_Lean_Parser_Tactic_tacticHave_x27_____x3a_x3d__); +l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__1 = _init_l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__1); +l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__2 = _init_l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__2(); +lean_mark_persistent(l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__2); +l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__3 = _init_l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__3(); +lean_mark_persistent(l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__3); +l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__4 = _init_l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__4(); +lean_mark_persistent(l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__4); +l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__5 = _init_l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__5(); +lean_mark_persistent(l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__5); +l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__6 = _init_l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__6(); +lean_mark_persistent(l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__6); +l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__7 = _init_l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__7(); +lean_mark_persistent(l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__7); +l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__8 = _init_l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__8(); +lean_mark_persistent(l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__8); +l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d__ = _init_l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d__(); +lean_mark_persistent(l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d__); l_Lean_Parser_Tactic_tacticLet_x27_____closed__1 = _init_l_Lean_Parser_Tactic_tacticLet_x27_____closed__1(); lean_mark_persistent(l_Lean_Parser_Tactic_tacticLet_x27_____closed__1); l_Lean_Parser_Tactic_tacticLet_x27_____closed__2 = _init_l_Lean_Parser_Tactic_tacticLet_x27_____closed__2(); @@ -19825,6 +20075,8 @@ l_Lean_Parser_Tactic_inductionAltLHS___closed__11 = _init_l_Lean_Parser_Tactic_i lean_mark_persistent(l_Lean_Parser_Tactic_inductionAltLHS___closed__11); l_Lean_Parser_Tactic_inductionAltLHS___closed__12 = _init_l_Lean_Parser_Tactic_inductionAltLHS___closed__12(); lean_mark_persistent(l_Lean_Parser_Tactic_inductionAltLHS___closed__12); +l_Lean_Parser_Tactic_inductionAltLHS___closed__13 = _init_l_Lean_Parser_Tactic_inductionAltLHS___closed__13(); +lean_mark_persistent(l_Lean_Parser_Tactic_inductionAltLHS___closed__13); l_Lean_Parser_Tactic_inductionAltLHS = _init_l_Lean_Parser_Tactic_inductionAltLHS(); lean_mark_persistent(l_Lean_Parser_Tactic_inductionAltLHS); l_Lean_Parser_Tactic_inductionAlt___closed__1 = _init_l_Lean_Parser_Tactic_inductionAlt___closed__1(); @@ -20019,6 +20271,10 @@ l_Lean_Parser_Tactic_renameI___closed__5 = _init_l_Lean_Parser_Tactic_renameI___ lean_mark_persistent(l_Lean_Parser_Tactic_renameI___closed__5); l_Lean_Parser_Tactic_renameI___closed__6 = _init_l_Lean_Parser_Tactic_renameI___closed__6(); lean_mark_persistent(l_Lean_Parser_Tactic_renameI___closed__6); +l_Lean_Parser_Tactic_renameI___closed__7 = _init_l_Lean_Parser_Tactic_renameI___closed__7(); +lean_mark_persistent(l_Lean_Parser_Tactic_renameI___closed__7); +l_Lean_Parser_Tactic_renameI___closed__8 = _init_l_Lean_Parser_Tactic_renameI___closed__8(); +lean_mark_persistent(l_Lean_Parser_Tactic_renameI___closed__8); l_Lean_Parser_Tactic_renameI = _init_l_Lean_Parser_Tactic_renameI(); lean_mark_persistent(l_Lean_Parser_Tactic_renameI); l_Lean_Parser_Tactic_tacticRepeat_____closed__1 = _init_l_Lean_Parser_Tactic_tacticRepeat_____closed__1(); diff --git a/stage0/stdlib/Init/WFTactics.c b/stage0/stdlib/Init/WFTactics.c index e8ea8a3ad9e..45a9a9be15f 100644 --- a/stage0/stdlib/Init/WFTactics.c +++ b/stage0/stdlib/Init/WFTactics.c @@ -3230,19 +3230,19 @@ return x_1; static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__1() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("paren", 5); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__6; +x_2 = l_tacticDecreasing__with_____closed__7; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; } } static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__6; -x_2 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__1; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("paren", 5); +return x_1; } } static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__3() { @@ -3250,7 +3250,7 @@ static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecrea { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__6; -x_2 = l_tacticDecreasing__with_____closed__7; +x_2 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__2; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } @@ -3840,7 +3840,7 @@ lean_ctor_set(x_55, 0, x_21); lean_ctor_set(x_55, 1, x_54); lean_ctor_set(x_55, 2, x_53); x_56 = lean_array_push(x_19, x_55); -x_57 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__3; +x_57 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__1; x_58 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_58, 0, x_21); lean_ctor_set(x_58, 1, x_57); @@ -3943,7 +3943,7 @@ lean_inc(x_99); x_100 = lean_array_push(x_99, x_95); lean_inc(x_97); x_101 = lean_array_push(x_100, x_97); -x_102 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__2; +x_102 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__3; x_103 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_103, 0, x_21); lean_ctor_set(x_103, 1, x_102); @@ -4357,7 +4357,7 @@ lean_ctor_set(x_283, 0, x_249); lean_ctor_set(x_283, 1, x_282); lean_ctor_set(x_283, 2, x_281); x_284 = lean_array_push(x_247, x_283); -x_285 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__3; +x_285 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__1; x_286 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_286, 0, x_249); lean_ctor_set(x_286, 1, x_285); @@ -4460,7 +4460,7 @@ lean_inc(x_327); x_328 = lean_array_push(x_327, x_323); lean_inc(x_325); x_329 = lean_array_push(x_328, x_325); -x_330 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__2; +x_330 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__3; x_331 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_331, 0, x_249); lean_ctor_set(x_331, 1, x_330); @@ -4950,7 +4950,7 @@ lean_ctor_set(x_35, 0, x_21); lean_ctor_set(x_35, 1, x_34); lean_ctor_set(x_35, 2, x_33); x_36 = lean_array_push(x_19, x_35); -x_37 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__3; +x_37 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__1; x_38 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_38, 0, x_21); lean_ctor_set(x_38, 1, x_37); @@ -5123,7 +5123,7 @@ lean_ctor_set(x_114, 0, x_100); lean_ctor_set(x_114, 1, x_113); lean_ctor_set(x_114, 2, x_112); x_115 = lean_array_push(x_98, x_114); -x_116 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__3; +x_116 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__1; x_117 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_117, 0, x_100); lean_ctor_set(x_117, 1, x_116); diff --git a/stage0/stdlib/Lean/Attributes.c b/stage0/stdlib/Lean/Attributes.c index c465fdc103a..3927a7f2120 100644 --- a/stage0/stdlib/Lean/Attributes.c +++ b/stage0/stdlib/Lean/Attributes.c @@ -85,7 +85,6 @@ LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3220____lamb lean_object* lean_environment_find(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Std_HashMapImp_contains___at_Lean_registerAttributeImplBuilder___spec__7(lean_object*, lean_object*); static lean_object* l_Lean_registerTagAttribute___closed__1; -lean_object* l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Attribute_erase(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l_Lean_isAttribute___closed__1; @@ -159,6 +158,7 @@ static size_t l_Std_PersistentHashMap_insertAux___at_Lean_registerBuiltinAttribu lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedAttributeExtensionState; static lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3220____closed__1; +lean_object* l_Lean_Syntax_isNatLit_x3f(lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_find_x3f___at_Lean_getBuiltinAttributeImpl___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3220____closed__8; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_updateEnvAttributesImpl___spec__3(lean_object*, size_t, size_t, lean_object*); @@ -170,7 +170,6 @@ LEAN_EXPORT lean_object* l_Std_PersistentHashMap_foldlMAux___at_Lean_getBuiltinA static lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3220____closed__2; lean_object* lean_st_ref_take(lean_object*, lean_object*); static lean_object* l_Lean_registerAttributeImplBuilder___lambda__1___closed__1; -extern lean_object* l_Lean_numLitKind; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_updateEnvAttributesImpl___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Attribute_Builtin_ensureNoArgs___closed__9; static lean_object* l_Lean_instBEqAttributeKind___closed__1; @@ -3254,45 +3253,44 @@ uint8_t x_5; x_5 = l_Lean_Syntax_isNone(x_1); if (x_5 == 0) { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = lean_unsigned_to_nat(0u); x_7 = l_Lean_Syntax_getArg(x_1, x_6); -x_8 = l_Lean_numLitKind; -x_9 = l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(x_8, x_7); +x_8 = l_Lean_Syntax_isNatLit_x3f(x_7); lean_dec(x_7); -if (lean_obj_tag(x_9) == 0) +if (lean_obj_tag(x_8) == 0) { -lean_object* x_10; lean_object* x_11; -x_10 = l_Lean_getAttrParamOptPrio___closed__2; -x_11 = l_Lean_throwErrorAt___at_Lean_getAttrParamOptPrio___spec__1(x_1, x_10, x_2, x_3, x_4); -return x_11; +lean_object* x_9; lean_object* x_10; +x_9 = l_Lean_getAttrParamOptPrio___closed__2; +x_10 = l_Lean_throwErrorAt___at_Lean_getAttrParamOptPrio___spec__1(x_1, x_9, x_2, x_3, x_4); +return x_10; } else { -lean_object* x_12; lean_object* x_13; +lean_object* x_11; lean_object* x_12; lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_12 = lean_ctor_get(x_9, 0); -lean_inc(x_12); -lean_dec(x_9); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_4); -return x_13; +x_11 = lean_ctor_get(x_8, 0); +lean_inc(x_11); +lean_dec(x_8); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_4); +return x_12; } } else { -lean_object* x_14; lean_object* x_15; +lean_object* x_13; lean_object* x_14; lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_14 = lean_unsigned_to_nat(1000u); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_15, 1, x_4); -return x_15; +x_13 = lean_unsigned_to_nat(1000u); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_4); +return x_14; } } } diff --git a/stage0/stdlib/Lean/Compiler/ExternAttr.c b/stage0/stdlib/Lean/Compiler/ExternAttr.c index acecfdf731e..94a1b70e5b0 100644 --- a/stage0/stdlib/Lean/Compiler/ExternAttr.c +++ b/stage0/stdlib/Lean/Compiler/ExternAttr.c @@ -42,7 +42,6 @@ static lean_object* l___private_Lean_Compiler_ExternAttr_0__Lean_getExternConstA lean_object* lean_environment_find(lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedExternAttrData___closed__1; LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Compiler_ExternAttr___hyg_395____lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_expandExternPattern(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_ExternAttr_0__Lean_getExternConstArity___closed__6; lean_object* lean_st_ref_get(lean_object*, lean_object*); @@ -71,6 +70,7 @@ LEAN_EXPORT lean_object* l_Lean_expandExternPattern___boxed(lean_object*, lean_o LEAN_EXPORT lean_object* l_Lean_ExternEntry_backend(lean_object*); LEAN_EXPORT uint8_t l_Lean_isExternC(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_ExternAttr_0__Lean_getExternConstArity___closed__13; +lean_object* l_Lean_Syntax_isNatLit_x3f(lean_object*); static lean_object* l_Lean_getConstInfo___at___private_Lean_Compiler_ExternAttr_0__Lean_getExternConstArity___spec__1___closed__1; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); uint8_t l_instDecidableNot___rarg(uint8_t); @@ -79,8 +79,8 @@ LEAN_EXPORT lean_object* l_Lean_getConstInfo___at___private_Lean_Compiler_Extern lean_object* lean_add_extern(lean_object*, lean_object*); lean_object* l_Lean_MessageData_toString(lean_object*, lean_object*); static lean_object* l_Lean_getExternConstArityExport___closed__5; -extern lean_object* l_Lean_numLitKind; lean_object* lean_nat_sub(lean_object*, lean_object*); +uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ExternAttrData_arity_x3f___default; LEAN_EXPORT lean_object* l_Lean_expandExternPatternAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getExternConstArityExport___closed__3; @@ -147,7 +147,6 @@ LEAN_EXPORT lean_object* l_Lean_addExtern___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getExternEntryFor(lean_object*, lean_object*); static lean_object* l_Lean_getExternConstArityExport___closed__6; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Compiler_ExternAttr_0__Lean_syntaxToExternAttrData___spec__3(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Compiler_ExternAttr_0__Lean_syntaxToExternAttrData___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_mul(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Compiler_ExternAttr_0__Lean_getExternConstArity___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); @@ -629,47 +628,46 @@ x_12 = lean_unsigned_to_nat(0u); x_13 = lean_nat_dec_eq(x_11, x_12); if (x_7 == 0) { -lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_object* x_24; lean_object* x_25; x_24 = l_Lean_Syntax_getArg(x_6, x_12); lean_dec(x_6); -x_25 = l_Lean_numLitKind; -x_26 = l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(x_25, x_24); +x_25 = l_Lean_Syntax_isNatLit_x3f(x_24); lean_dec(x_24); -if (lean_obj_tag(x_26) == 0) +if (lean_obj_tag(x_25) == 0) { -lean_object* x_27; -x_27 = l___private_Lean_Compiler_ExternAttr_0__Lean_syntaxToExternAttrData___closed__4; -x_14 = x_27; +lean_object* x_26; +x_26 = l___private_Lean_Compiler_ExternAttr_0__Lean_syntaxToExternAttrData___closed__4; +x_14 = x_26; goto block_23; } else { -uint8_t x_28; -x_28 = !lean_is_exclusive(x_26); -if (x_28 == 0) +uint8_t x_27; +x_27 = !lean_is_exclusive(x_25); +if (x_27 == 0) { -x_14 = x_26; +x_14 = x_25; goto block_23; } else { -lean_object* x_29; lean_object* x_30; -x_29 = lean_ctor_get(x_26, 0); -lean_inc(x_29); -lean_dec(x_26); -x_30 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_30, 0, x_29); -x_14 = x_30; +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_25, 0); +lean_inc(x_28); +lean_dec(x_25); +x_29 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_29, 0, x_28); +x_14 = x_29; goto block_23; } } } else { -lean_object* x_31; +lean_object* x_30; lean_dec(x_6); -x_31 = lean_box(0); -x_14 = x_31; +x_30 = lean_box(0); +x_14 = x_30; goto block_23; } block_23: @@ -687,7 +685,7 @@ else { lean_object* x_17; uint8_t x_18; x_17 = lean_box(0); -x_18 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1(x_14, x_17); +x_18 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1(x_14, x_17); if (x_18 == 0) { lean_object* x_19; lean_object* x_20; diff --git a/stage0/stdlib/Lean/Compiler/IR/Basic.c b/stage0/stdlib/Lean/Compiler/IR/Basic.c index 83b725244d7..47ca78f3e6e 100644 --- a/stage0/stdlib/Lean/Compiler/IR/Basic.c +++ b/stage0/stdlib/Lean/Compiler/IR/Basic.c @@ -71,7 +71,6 @@ lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* lean_ir_mk_str_expr(lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_IR_addParamsRename___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_IR_IRType_beq___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_IR_Decl_updateBody_x21___closed__2; static lean_object* l_Lean_IR_instAlphaEqvExpr___closed__1; LEAN_EXPORT lean_object* l_Lean_IR_IRType_isIrrelevant___boxed(lean_object*); @@ -110,6 +109,7 @@ LEAN_EXPORT lean_object* l_Std_RBNode_find___at_Lean_IR_LocalContext_isJP___spec LEAN_EXPORT lean_object* l_Std_RBNode_del___at_Lean_IR_LocalContext_eraseJoinPointDecl___spec__2(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_IR_args_alphaEqv(lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_IR_IRType_beq___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_reprCtorInfo____x40_Lean_Compiler_IR_Basic___hyg_763____boxed(lean_object*, lean_object*); static lean_object* l_Lean_IR_reshapeAux___closed__4; LEAN_EXPORT uint8_t l_Lean_IR_VarId_alphaEqv(lean_object*, lean_object*, lean_object*); @@ -122,6 +122,7 @@ LEAN_EXPORT uint8_t l_Lean_IR_CtorInfo_isRef(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_CtorInfo_isRef___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_VarId_alphaEqv___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_addParamRename(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_IR_IRType_beq___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT uint64_t l_Lean_IR_instHashableVarId(lean_object*); static lean_object* l_Lean_IR_IRType_instBEqIRType___closed__1; LEAN_EXPORT lean_object* l_Lean_IR_FnBody_isTerminal___boxed(lean_object*); @@ -282,7 +283,6 @@ LEAN_EXPORT lean_object* l_Lean_IR_Arg_beq___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_reprCtorInfo____x40_Lean_Compiler_IR_Basic___hyg_763____closed__20; LEAN_EXPORT lean_object* l_Lean_IR_mkParam___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_IR_Basic_0__Lean_IR_reprCtorInfo____x40_Lean_Compiler_IR_Basic___hyg_763____closed__19; -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_IR_IRType_beq___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_RBNode_ins___at_Lean_IR_mkIndexSet___spec__2(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_KVMap_eqv(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_IR_Alt_isDefault(lean_object*); @@ -495,7 +495,7 @@ x_1 = lean_box(0); return x_1; } } -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_IR_IRType_beq___spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_IR_IRType_beq___spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -764,7 +764,7 @@ x_21 = lean_ctor_get(x_1, 0); x_22 = lean_ctor_get(x_1, 1); x_23 = lean_ctor_get(x_2, 0); x_24 = lean_ctor_get(x_2, 1); -x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_IR_IRType_beq___spec__1(x_21, x_23); +x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_IR_IRType_beq___spec__1(x_21, x_23); if (x_25 == 0) { uint8_t x_26; @@ -850,11 +850,11 @@ return x_46; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_IR_IRType_beq___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_IR_IRType_beq___spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_IR_IRType_beq___spec__1(x_1, x_2); +x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_IR_IRType_beq___spec__1(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); diff --git a/stage0/stdlib/Lean/Compiler/IR/EmitC.c b/stage0/stdlib/Lean/Compiler/IR/EmitC.c index 31874ca7a75..c03f45c0b15 100644 --- a/stage0/stdlib/Lean/Compiler/IR/EmitC.c +++ b/stage0/stdlib/Lean/Compiler/IR/EmitC.c @@ -149,7 +149,6 @@ static lean_object* l_Lean_IR_EmitC_emitMainFn___lambda__2___closed__16; static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_IR_EmitC_emitCase___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitFnDecl(lean_object*, uint8_t, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_quoteString___closed__2; -uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_IR_IRType_beq___spec__1(lean_object*, lean_object*); lean_object* l_Lean_getExternNameFor(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_quoteString___closed__1; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_quoteString___boxed(lean_object*); @@ -230,6 +229,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitDeclAux___lambda__2(lean_object*, l LEAN_EXPORT lean_object* l_Lean_IR_EmitC_toCInitName(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_overwriteParam___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitTailCall___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_IR_IRType_beq___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitTag___closed__1; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitTailCall(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_IR_getDecls(lean_object*); @@ -263,6 +263,7 @@ LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitCtor___boxed(lean_object*, lean_obj lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitSimpleExternalCall___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitBoxFn___closed__4; +lean_object* l_panic___at_Lean_TSyntax_getString___spec__1(lean_object*); static lean_object* l_Lean_IR_EmitC_emitMainFn___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitJmp(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitMainFn___lambda__1___closed__6; @@ -553,7 +554,6 @@ LEAN_EXPORT lean_object* l_Std_RBNode_revFold___at_Lean_IR_EmitC_emitFnDecls___s LEAN_EXPORT lean_object* l_Lean_IR_EmitC_emitOffset(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_IR_EmitC_emitInitFn___spec__3(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitSSet___closed__4; -lean_object* l_panic___at_Lean_Name_getString_x21___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Nat_forM_loop___at_Lean_IR_EmitC_emitTailCall___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_IR_EmitC_emitInitFn___closed__2; LEAN_EXPORT lean_object* l_List_forM___at_Lean_IR_EmitC_emitMainFn___spec__3(lean_object*, lean_object*, lean_object*); @@ -1177,14 +1177,14 @@ case 9: { lean_object* x_8; lean_object* x_9; x_8 = l_Lean_IR_EmitC_toCType___closed__11; -x_9 = l_panic___at_Lean_Name_getString_x21___spec__1(x_8); +x_9 = l_panic___at_Lean_TSyntax_getString___spec__1(x_8); return x_9; } case 10: { lean_object* x_10; lean_object* x_11; x_10 = l_Lean_IR_EmitC_toCType___closed__12; -x_11 = l_panic___at_Lean_Name_getString_x21___spec__1(x_10); +x_11 = l_panic___at_Lean_TSyntax_getString___spec__1(x_10); return x_11; } default: @@ -3034,7 +3034,7 @@ lean_dec(x_26); x_28 = l_Lean_Expr_constName_x3f(x_27); lean_dec(x_27); x_29 = l_Lean_IR_EmitC_emitMainFn___lambda__1___closed__4; -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_IR_IRType_beq___spec__1(x_28, x_29); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_IR_IRType_beq___spec__1(x_28, x_29); lean_dec(x_28); if (x_30 == 0) { diff --git a/stage0/stdlib/Lean/Compiler/ImplementedByAttr.c b/stage0/stdlib/Lean/Compiler/ImplementedByAttr.c index 9cd4af4a7f9..d3fb0103cd6 100644 --- a/stage0/stdlib/Lean/Compiler/ImplementedByAttr.c +++ b/stage0/stdlib/Lean/Compiler/ImplementedByAttr.c @@ -19,6 +19,7 @@ LEAN_EXPORT lean_object* l_Lean_resolveGlobalName___at_Lean_Compiler_initFn____x lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Compiler_initFn____x40_Lean_Compiler_ImplementedByAttr___hyg_4____spec__2(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_registerTagAttribute___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); lean_object* l_Lean_setEnv___rarg(lean_object*, lean_object*); lean_object* lean_name_mk_string(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_ImplementedByAttr___hyg_4____lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -72,7 +73,6 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Implemented static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_ImplementedByAttr___hyg_4____lambda__2___closed__3; static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_ImplementedByAttr___hyg_4____lambda__4___closed__2; lean_object* l_Lean_ParametricAttribute_setParam___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at_Lean_Compiler_initFn____x40_Lean_Compiler_ImplementedByAttr___hyg_4____spec__7___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_resolveGlobalConstNoOverload___at_Lean_Compiler_initFn____x40_Lean_Compiler_ImplementedByAttr___hyg_4____spec__3___closed__1; uint8_t lean_expr_eqv(lean_object*, lean_object*); @@ -478,7 +478,7 @@ x_7 = lean_ctor_get(x_5, 1); lean_inc(x_7); lean_dec(x_5); x_8 = lean_box(0); -x_9 = l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(x_6, x_8); +x_9 = l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(x_6, x_8); x_10 = l_List_isEmpty___rarg(x_9); if (x_10 == 0) { diff --git a/stage0/stdlib/Lean/Compiler/InitAttr.c b/stage0/stdlib/Lean/Compiler/InitAttr.c index a2ba0c4e8b2..f815088f5e2 100644 --- a/stage0/stdlib/Lean/Compiler/InitAttr.c +++ b/stage0/stdlib/Lean/Compiler/InitAttr.c @@ -28,6 +28,7 @@ static lean_object* l_Lean_declareBuiltin___closed__6; static lean_object* l_Lean_getBuiltinInitFnNameFor_x3f___closed__1; LEAN_EXPORT lean_object* l_Lean_builtinInitAttr; lean_object* l_Lean_throwError___at_Lean_registerTagAttribute___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at_Lean_registerInitAttrUnsafe___spec__10(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_name_mk_string(lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); @@ -124,7 +125,6 @@ LEAN_EXPORT uint8_t l_Lean_hasInitAttr(lean_object*, lean_object*); lean_object* l_Lean_ParametricAttribute_setParam___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_declareBuiltin___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_registerInitAttrUnsafe___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at_Lean_registerInitAttrUnsafe___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfo___at_Lean_registerInitAttrUnsafe___spec__1___closed__3; LEAN_EXPORT uint8_t l___private_Lean_Compiler_InitAttr_0__Lean_isIOUnit(lean_object*); @@ -765,7 +765,7 @@ x_7 = lean_ctor_get(x_5, 1); lean_inc(x_7); lean_dec(x_5); x_8 = lean_box(0); -x_9 = l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(x_6, x_8); +x_9 = l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(x_6, x_8); x_10 = l_List_isEmpty___rarg(x_9); if (x_10 == 0) { diff --git a/stage0/stdlib/Lean/Data/FuzzyMatching.c b/stage0/stdlib/Lean/Data/FuzzyMatching.c index fea6070f3d9..73ec2c94c27 100644 --- a/stage0/stdlib/Lean/Data/FuzzyMatching.c +++ b/stage0/stdlib/Lean/Data/FuzzyMatching.c @@ -31,8 +31,8 @@ LEAN_EXPORT lean_object* l_Lean_FuzzyMatching_CharType_toCtorIdx(uint8_t); LEAN_EXPORT uint8_t l_Lean_FuzzyMatching_charRole(lean_object*, uint8_t, lean_object*); static lean_object* l___private_Lean_Data_FuzzyMatching_0__Lean_FuzzyMatching_iterateLookaround___rarg___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Data_FuzzyMatching_0__Lean_FuzzyMatching_iterateLookaround___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Data_FuzzyMatching_0__Lean_FuzzyMatching_fuzzyMatchCore___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -double lean_float_of_nat(lean_object*); static double l_Lean_FuzzyMatching_fuzzyMatchScore_x3f___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Lean_FuzzyMatching_CharRole_toCtorIdx(uint8_t); static lean_object* l___private_Lean_Data_FuzzyMatching_0__Lean_FuzzyMatching_iterateLookaround___rarg___closed__1; @@ -2549,19 +2549,22 @@ return x_2; static double _init_l_Lean_FuzzyMatching_fuzzyMatchScore_x3f___lambda__1___closed__2() { _start: { -lean_object* x_1; double x_2; +lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; x_1 = lean_unsigned_to_nat(1u); -x_2 = lean_float_of_nat(x_1); -return x_2; +x_2 = 0; +x_3 = lean_unsigned_to_nat(0u); +x_4 = l_Float_ofScientific(x_1, x_2, x_3); +return x_4; } } static double _init_l_Lean_FuzzyMatching_fuzzyMatchScore_x3f___lambda__1___closed__3() { _start: { -lean_object* x_1; double x_2; +lean_object* x_1; uint8_t x_2; double x_3; x_1 = lean_unsigned_to_nat(0u); -x_2 = lean_float_of_nat(x_1); -return x_2; +x_2 = 0; +x_3 = l_Float_ofScientific(x_1, x_2, x_1); +return x_3; } } static lean_object* _init_l_Lean_FuzzyMatching_fuzzyMatchScore_x3f___lambda__1___closed__4___boxed__const__1() { diff --git a/stage0/stdlib/Lean/Data/Json/FromToJson.c b/stage0/stdlib/Lean/Data/Json/FromToJson.c index bf5f777c518..f7bd8ee064d 100644 --- a/stage0/stdlib/Lean/Data/Json/FromToJson.c +++ b/stage0/stdlib/Lean/Data/Json/FromToJson.c @@ -1833,6 +1833,7 @@ else lean_object* x_8; lean_object* x_9; lean_object* x_10; x_8 = lean_array_uget(x_2, x_4); x_9 = l_Lean_Name_getString_x21(x_8); +lean_dec(x_8); x_10 = l_Lean_Json_getObjVal_x3f(x_1, x_9); lean_dec(x_9); if (lean_obj_tag(x_10) == 0) diff --git a/stage0/stdlib/Lean/Data/Json/Parser.c b/stage0/stdlib/Lean/Data/Json/Parser.c index c50c309ef77..62364366c8d 100644 --- a/stage0/stdlib/Lean/Data/Json/Parser.c +++ b/stage0/stdlib/Lean/Data/Json/Parser.c @@ -30,6 +30,7 @@ static lean_object* l_Lean_Json_Parser_anyCore___closed__2; LEAN_EXPORT lean_object* l_Lean_Json_Parser_num___lambda__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_Parser_num___lambda__5(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Json_Parser_anyCore___closed__1; +lean_object* l___private_Init_Data_Option_Basic_0__decEqOption____x40_Init_Data_Option_Basic___hyg_661____at_instDecidableEqOption___spec__1___rarg(lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); lean_object* lean_nat_pow(lean_object*, lean_object*); @@ -68,7 +69,6 @@ uint8_t l_String_Iterator_hasNext(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_Parser_escapedChar___boxed__const__7; lean_object* lean_int_neg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_Parser_arrayCore(lean_object*, lean_object*, lean_object*); -lean_object* l___private_Init_Data_Option_Basic_0__decEqOption____x40_Init_Data_Option_Basic___hyg_683____at_instDecidableEqOption___spec__1___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_Parser_strCore(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_Parser_num___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_uint32_dec_eq(uint32_t, uint32_t); @@ -1994,14 +1994,14 @@ lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; x_41 = l_Lean_Json_Parser_num___lambda__3___closed__2; x_42 = l_Lean_Json_Parser_num___lambda__3___closed__1; lean_inc(x_40); -x_43 = l___private_Init_Data_Option_Basic_0__decEqOption____x40_Init_Data_Option_Basic___hyg_683____at_instDecidableEqOption___spec__1___rarg(x_41, x_40, x_42); +x_43 = l___private_Init_Data_Option_Basic_0__decEqOption____x40_Init_Data_Option_Basic___hyg_661____at_instDecidableEqOption___spec__1___rarg(x_41, x_40, x_42); x_44 = lean_unbox(x_43); lean_dec(x_43); if (x_44 == 0) { lean_object* x_45; lean_object* x_46; uint8_t x_47; x_45 = l_Lean_Json_Parser_num___lambda__3___closed__3; -x_46 = l___private_Init_Data_Option_Basic_0__decEqOption____x40_Init_Data_Option_Basic___hyg_683____at_instDecidableEqOption___spec__1___rarg(x_41, x_40, x_45); +x_46 = l___private_Init_Data_Option_Basic_0__decEqOption____x40_Init_Data_Option_Basic___hyg_661____at_instDecidableEqOption___spec__1___rarg(x_41, x_40, x_45); x_47 = lean_unbox(x_46); lean_dec(x_46); if (x_47 == 0) @@ -2154,7 +2154,7 @@ lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_6 = l_Lean_Json_Parser_num___lambda__5___closed__1; x_7 = l_Lean_Json_Parser_num___lambda__3___closed__2; x_8 = l_Lean_Json_Parser_num___lambda__5___closed__2; -x_9 = l___private_Init_Data_Option_Basic_0__decEqOption____x40_Init_Data_Option_Basic___hyg_683____at_instDecidableEqOption___spec__1___rarg(x_7, x_5, x_8); +x_9 = l___private_Init_Data_Option_Basic_0__decEqOption____x40_Init_Data_Option_Basic___hyg_661____at_instDecidableEqOption___spec__1___rarg(x_7, x_5, x_8); x_10 = lean_unbox(x_9); lean_dec(x_9); if (x_10 == 0) diff --git a/stage0/stdlib/Lean/Data/JsonRpc.c b/stage0/stdlib/Lean/Data/JsonRpc.c index 1bc6b5f802a..11e7185ed10 100644 --- a/stage0/stdlib/Lean/Data/JsonRpc.c +++ b/stage0/stdlib/Lean/Data/JsonRpc.c @@ -123,7 +123,6 @@ static lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMe static lean_object* l_Lean_JsonRpc_instToJsonErrorCode___closed__44; static lean_object* l_Lean_JsonRpc_instInhabitedRequestID___closed__2; static lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__4___closed__2; -lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readResponseAs___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Data_JsonRpc_0__Lean_JsonRpc_beqNotification____x40_Lean_Data_JsonRpc___hyg_1020____rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_JsonRpc_ErrorCode_toCtorIdx___boxed(lean_object*); @@ -237,6 +236,7 @@ LEAN_EXPORT lean_object* l_IO_FS_Stream_readRequestAs(lean_object*, lean_object* lean_object* l_List_appendTR___rarg(lean_object*, lean_object*); static lean_object* l_Lean_JsonRpc_instBEqRequestID___closed__1; LEAN_EXPORT lean_object* l_IO_FS_Stream_writeRequest___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_JsonRpc_instFromJsonErrorCode___closed__4; LEAN_EXPORT uint8_t l___private_Lean_Data_JsonRpc_0__Lean_JsonRpc_RequestID_lt(lean_object*, lean_object*); static lean_object* l_Lean_JsonRpc_instFromJsonRequestID___closed__2; @@ -2313,7 +2313,7 @@ return x_20; else { lean_object* x_21; -x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____rarg(x_1, x_7, x_11); +x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____rarg(x_1, x_7, x_11); return x_21; } } diff --git a/stage0/stdlib/Lean/Data/Lsp/Capabilities.c b/stage0/stdlib/Lean/Data/Lsp/Capabilities.c index 968dcfa2e35..4b859561744 100644 --- a/stage0/stdlib/Lean/Data/Lsp/Capabilities.c +++ b/stage0/stdlib/Lean/Data/Lsp/Capabilities.c @@ -44,7 +44,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_from LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonServerCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_705____spec__3___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonServerCapabilities; static lean_object* l_Lean_Lsp_instToJsonCompletionClientCapabilities___closed__1; -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3415_(lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3414_(lean_object*); LEAN_EXPORT uint8_t l_Lean_Lsp_ServerCapabilities_workspaceSymbolProvider___default; LEAN_EXPORT uint8_t l_Lean_Lsp_ServerCapabilities_hoverProvider___default; lean_object* l_List_join___rarg(lean_object*); @@ -57,7 +57,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCompletionClientCapabilities; LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_toJsonWindowClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_304____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_toJsonServerCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_576____closed__6; lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_112_(lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3346_(lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3345_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_ClientCapabilities_textDocument_x3f___default; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonTextDocumentClientCapabilities; static lean_object* l_Lean_Lsp_instToJsonShowDocumentClientCapabilities___closed__1; @@ -1777,7 +1777,7 @@ lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_obj x_4 = lean_ctor_get(x_2, 0); lean_inc(x_4); lean_dec(x_2); -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3415_(x_4); +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3414_(x_4); x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_1); lean_ctor_set(x_6, 1, x_5); @@ -2192,7 +2192,7 @@ return x_4; else { lean_object* x_5; -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3346_(x_3); +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3345_(x_3); lean_dec(x_3); if (lean_obj_tag(x_5) == 0) { diff --git a/stage0/stdlib/Lean/Data/Lsp/Diagnostics.c b/stage0/stdlib/Lean/Data/Lsp/Diagnostics.c index 450efe44c22..35ca447ffc5 100644 --- a/stage0/stdlib/Lean/Data/Lsp/Diagnostics.c +++ b/stage0/stdlib/Lean/Data/Lsp/Diagnostics.c @@ -35,6 +35,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDiagnosticWith(lean_object*); lean_object* l_Lean_Json_getObjValAs_x3f(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonDiagnosticSeverity___closed__3; static lean_object* l_Lean_Lsp_instFromJsonDiagnosticSeverity___closed__4; +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__3(lean_object*, lean_object*); lean_object* l_Lean_Json_getNat_x3f(lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_1351_(lean_object*); @@ -43,22 +44,20 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instInhabitedDiagnosticRelatedInformation; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_926____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_1307____spec__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqDiagnosticCode; static lean_object* l_Lean_Lsp_instFromJsonDiagnosticSeverity___closed__3; -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_1220____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1009____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_1351____spec__3___boxed(lean_object*); static lean_object* l_Lean_Lsp_instInhabitedDiagnosticRelatedInformation___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_1351____spec__4___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_DiagnosticSeverity_noConfusion(lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_469_(lean_object*, lean_object*); static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1009____spec__1___closed__3; -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__3___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1009____rarg(lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__1(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_926____spec__4___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDiagnosticCode(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDiagnosticWith(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instInhabitedDiagnosticCode; +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_1220____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1009____spec__6___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonDiagnosticSeverity___closed__2; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1009____spec__1(lean_object*, lean_object*); @@ -68,11 +67,10 @@ static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Dia static lean_object* l_Lean_Lsp_instInhabitedDiagnosticRelatedInformation___closed__4; lean_object* l_List_join___rarg(lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__5(lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1009____spec__1___closed__2; LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_1220____spec__2(lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__3(lean_object*, lean_object*); lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_931_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_DiagnosticSeverity_toCtorIdx___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDiagnosticSeverity(lean_object*); @@ -83,7 +81,6 @@ LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Data_Lsp_Diagnosti LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_1351____spec__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_540____closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_1307____spec__3___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_1220____spec__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_540____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_PublishDiagnosticsParams_version_x3f___default; lean_object* lean_array_fget(lean_object*, lean_object*); @@ -93,22 +90,20 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDiagnosticRelatedInformation; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_926____spec__6___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_926____spec__5(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__3___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1009____spec__5(lean_object*, lean_object*); static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1009____spec__1___closed__6; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_1351____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonPublishDiagnosticsParams___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_926_(lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__5___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_580____spec__1(lean_object*, lean_object*); lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_2125____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_580____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1009____spec__2___boxed(lean_object*, lean_object*); lean_object* l_Lean_Json_getInt_x3f(lean_object*); static lean_object* l_Lean_Lsp_instToJsonDiagnosticRelatedInformation___closed__1; -uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_SearchPath_findAllWithExt___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1009____spec__2(lean_object*, lean_object*); uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_509_(lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__5(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_DiagnosticSeverity_noConfusion___rarg___lambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqDiagnosticRelatedInformation; LEAN_EXPORT lean_object* l_Lean_Lsp_DiagnosticTag_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); @@ -120,6 +115,7 @@ LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostic static lean_object* l_Lean_Lsp_instFromJsonPublishDiagnosticsParams___closed__1; static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_1307____closed__2; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_580_(lean_object*); +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_1220____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticSeverity____x40_Lean_Data_Lsp_Diagnostics___hyg_18____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_1307_(lean_object*); lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_931____spec__1(lean_object*, lean_object*); @@ -137,6 +133,7 @@ static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiag static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1009____spec__1___closed__8; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDiagnosticTag___boxed(lean_object*); static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1009____spec__1___closed__4; +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__2(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instBEqDiagnosticSeverity___closed__1; uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqLocation____x40_Lean_Data_Lsp_Basic___hyg_820_(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_1220____spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -193,6 +190,7 @@ static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonPubl static lean_object* l_Lean_Lsp_instToJsonDiagnosticSeverity___closed__6; static lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_926____rarg___closed__6; LEAN_EXPORT lean_object* l_Lean_Lsp_DiagnosticTag_noConfusion(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__5___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instBEqDiagnosticRelatedInformation___closed__1; static lean_object* l_Lean_Lsp_instToJsonDiagnosticSeverity___closed__8; LEAN_EXPORT lean_object* l_Lean_Lsp_DiagnosticWith_fullRange___default___boxed(lean_object*); @@ -203,20 +201,21 @@ static lean_object* l_Lean_Lsp_instInhabitedDiagnosticCode___closed__1; lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_931____spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticRelatedInformation____x40_Lean_Data_Lsp_Diagnostics___hyg_580____boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDiagnosticTag(lean_object*); +uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_SearchPath_findAllWithExt___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDiagnosticSeverity___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_926____spec__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_DiagnosticTag_toCtorIdx___boxed(lean_object*); static lean_object* l_Lean_Lsp_instBEqPublishDiagnosticsParams___closed__1; static lean_object* l_Lean_Lsp_instInhabitedDiagnosticCode___closed__2; static lean_object* l_Lean_Lsp_instFromJsonDiagnosticSeverity___closed__5; -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDiagnosticWith___rarg(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__2___boxed(lean_object*, lean_object*); lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_617_(lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonDiagnosticSeverity___closed__5; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_926____spec__4(size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__2___boxed(lean_object*, lean_object*); uint8_t lean_int_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_DiagnosticWith_tags_x3f___default; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_1351____spec__2(lean_object*, lean_object*); @@ -231,6 +230,7 @@ lean_object* lean_nat_to_int(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqPublishDiagnosticsParams; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1009_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1009____spec__3(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_1009____rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_1220_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqDiagnosticWith___rarg(lean_object*); @@ -1582,7 +1582,7 @@ x_2 = lean_alloc_closure((void*)(l_Lean_Lsp_instInhabitedDiagnosticWith___rarg), return x_2; } } -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -1629,7 +1629,7 @@ return x_10; } } } -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__2(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -1710,7 +1710,7 @@ goto _start; } } } -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__3(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__3(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -1810,7 +1810,7 @@ goto _start; } } } -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__5(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__5(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -1958,7 +1958,7 @@ return x_23; else { uint8_t x_24; -x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__1(x_6, x_14); +x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__1(x_6, x_14); if (x_24 == 0) { uint8_t x_25; @@ -1979,7 +1979,7 @@ return x_25; else { uint8_t x_26; -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__2(x_7, x_15); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__2(x_7, x_15); lean_dec(x_15); lean_dec(x_7); if (x_26 == 0) @@ -2000,7 +2000,7 @@ return x_27; else { uint8_t x_28; -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_SearchPath_findAllWithExt___spec__1(x_8, x_16); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_SearchPath_findAllWithExt___spec__1(x_8, x_16); lean_dec(x_16); lean_dec(x_8); if (x_28 == 0) @@ -2035,7 +2035,7 @@ return x_32; else { uint8_t x_33; -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__3(x_10, x_18); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__3(x_10, x_18); lean_dec(x_18); if (x_33 == 0) { @@ -2048,7 +2048,7 @@ return x_34; else { uint8_t x_35; -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__5(x_11, x_19); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__5(x_11, x_19); lean_dec(x_19); return x_35; } @@ -2068,20 +2068,20 @@ x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Diagnostics_0__Lean_L return x_2; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__1(x_1, x_2); +x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__1(x_1, x_2); x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__2___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__2___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__2(x_1, x_2); +x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__2(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -2101,11 +2101,11 @@ x_8 = lean_box(x_7); return x_8; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__3___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__3___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__3(x_1, x_2); +x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__3(x_1, x_2); lean_dec(x_2); x_4 = lean_box(x_3); return x_4; @@ -2124,11 +2124,11 @@ x_8 = lean_box(x_7); return x_8; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__5___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__5___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__5(x_1, x_2); +x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__5(x_1, x_2); lean_dec(x_2); x_4 = lean_box(x_3); return x_4; @@ -3757,7 +3757,7 @@ x_1 = l_Lean_Lsp_instInhabitedPublishDiagnosticsParams___closed__2; return x_1; } } -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_1220____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_1220____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -3882,7 +3882,7 @@ return x_22; else { uint8_t x_23; -x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__1(x_5, x_13); +x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__1(x_5, x_13); if (x_23 == 0) { uint8_t x_24; @@ -3902,7 +3902,7 @@ return x_24; else { uint8_t x_25; -x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__2(x_6, x_14); +x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__2(x_6, x_14); lean_dec(x_14); lean_dec(x_6); if (x_25 == 0) @@ -3922,7 +3922,7 @@ return x_26; else { uint8_t x_27; -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_SearchPath_findAllWithExt___spec__1(x_7, x_15); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_SearchPath_findAllWithExt___spec__1(x_7, x_15); lean_dec(x_15); lean_dec(x_7); if (x_27 == 0) @@ -3956,7 +3956,7 @@ return x_30; else { uint8_t x_31; -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__3(x_9, x_17); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__3(x_9, x_17); lean_dec(x_17); if (x_31 == 0) { @@ -3969,7 +3969,7 @@ return x_32; else { uint8_t x_33; -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__5(x_10, x_18); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqDiagnosticWith____x40_Lean_Data_Lsp_Diagnostics___hyg_744____spec__5(x_10, x_18); lean_dec(x_18); return x_33; } @@ -4055,7 +4055,7 @@ return x_10; else { uint8_t x_11; -x_11 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_1220____spec__1(x_4, x_7); +x_11 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_1220____spec__1(x_4, x_7); lean_dec(x_7); lean_dec(x_4); if (x_11 == 0) @@ -4095,11 +4095,11 @@ return x_18; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_1220____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_1220____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_1220____spec__1(x_1, x_2); +x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_beqPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_1220____spec__1(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); diff --git a/stage0/stdlib/Lean/Data/Lsp/LanguageFeatures.c b/stage0/stdlib/Lean/Data/Lsp/LanguageFeatures.c index 1022f1c5a51..68b5906da4f 100644 --- a/stage0/stdlib/Lean/Data/Lsp/LanguageFeatures.c +++ b/stage0/stdlib/Lean/Data/Lsp/LanguageFeatures.c @@ -16,54 +16,47 @@ extern "C" { LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSemanticTokens; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSemanticTokensParams; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCompletionOptions; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__138; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__90; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__102; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1371____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCompletionItemKind(uint8_t); static lean_object* l_Lean_Lsp_instToJsonReferenceContext___closed__1; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2256____spec__1(lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__143; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2255____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonSemanticTokens___closed__1; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____spec__3(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3664____spec__2(size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1665____closed__1; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3663____spec__2(size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSymbolTag___boxed(lean_object*); lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_891_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2208_(uint8_t); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1664____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2207_(uint8_t); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__31; lean_object* lean_mk_empty_array_with_capacity(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____spec__2___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__138; LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolTag_noConfusion___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instDecidableEqCompletionItemKind___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3088____spec__1(size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__90; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonHoverParams; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonDocumentHighlight___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__143; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__95; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__3; lean_object* lean_array_uget(lean_object*, size_t); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonInsertReplaceEdit; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3088____closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1761_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1760_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3087____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__134; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__3; LEAN_EXPORT uint8_t l_Lean_Lsp_CompletionItemKind_ofNat(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1697____boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSemanticTokensRangeParams; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemKind_ofNat___boxed(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1698____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2325_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__5; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__74; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__134; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__92; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2324_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__59; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__31; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemKind_noConfusion___rarg___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__124; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemKind_noConfusion___rarg(uint8_t, uint8_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__5; lean_object* l_Lean_Json_getNat_x3f(lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__31; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__58; static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCompletionItemKind___boxed(lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__26; @@ -71,530 +64,537 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenType_toCtorIdx(uint8_t); static lean_object* l_Lean_Lsp_instToJsonInsertReplaceEdit___closed__1; static lean_object* l_Lean_Lsp_instFromJsonHoverParams___closed__1; static lean_object* l_Lean_Lsp_instToJsonDefinitionParams___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__79; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__95; static lean_object* l_Lean_Lsp_instFromJsonSemanticTokensLegend___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____rarg___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__124; static lean_object* l_Lean_Lsp_instToJsonSemanticTokensRangeParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__92; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__144; LEAN_EXPORT lean_object* l_Lean_Lsp_FoldingRangeKind_toCtorIdx(uint8_t); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__64; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3227____spec__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3346____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__144; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1449____spec__2(size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3664____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1665____closed__2; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1448____spec__2(size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3345____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__89; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__47; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1500____spec__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__74; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__60; LEAN_EXPORT uint8_t l_Lean_Lsp_CompletionOptions_resolveProvider___default; LEAN_EXPORT lean_object* l_Lean_Lsp_DocumentSymbolAux_children_x3f___default(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__94; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__8; LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolTag_noConfusion___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentSymbolAux___rarg(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__150; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__12; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__97; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1664____closed__2; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__6; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemKind_noConfusion___rarg___lambda__1___boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__57; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__150; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__17; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1761____boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentSymbol; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__122; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1100_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1099_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__64; static lean_object* l_Lean_Lsp_CompletionItemKind_noConfusion___rarg___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__97; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3087____closed__2; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3663____spec__1(lean_object*, lean_object*); lean_object* l_Lean_Json_getStr_x3f(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSemanticTokens; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentHighlightKind___boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3088____closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1760____boxed(lean_object*); static lean_object* l_Lean_Lsp_instToJsonDocumentHighlightKind___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__1; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3663____spec__2___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1697____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____spec__1(lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__47; lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDeclarationParams; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__89; static lean_object* l_Lean_Lsp_instReprCompletionItemKind___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__60; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3664____spec__2___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDefinitionParams; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDefinitionParams; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__40; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__67; static lean_object* l_Lean_Lsp_instFromJsonDocumentHighlightParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__131; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____spec__2___rarg(lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3698____spec__1(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instInhabitedCompletionItem; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__4; lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_657____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__7; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__9; LEAN_EXPORT lean_object* l_Lean_Json_opt___at_Lean_Lsp_instToJsonDocumentSymbol_go___spec__2(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__42; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1864_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____rarg___closed__2; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2600____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__40; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__67; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2173____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1863_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__8; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2599____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__94; static lean_object* l_Lean_Lsp_instFromJsonCompletionItem___closed__1; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____spec__1(size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__139; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__37; static lean_object* l_Lean_Lsp_instFromJsonSemanticTokens___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2514_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3415_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2513_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3414_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonDeclarationParams___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolKind_noConfusion(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2173____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__57; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__44; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____rarg___closed__3; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__3; static lean_object* l_Lean_Lsp_instToJsonFoldingRangeKind___closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__87; static lean_object* l_Lean_Lsp_instToJsonTypeDefinitionParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__44; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1370____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__27; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__87; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2172____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__131; LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenType_names; lean_object* l_List_join___rarg(lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonHover___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__9; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1698____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__27; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3346____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__4; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCompletionItem; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2122_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__139; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2600____spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2121_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2172____closed__1; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3345____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__55; lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentSymbolParams; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__37; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__14; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1697____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonCompletionOptions___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__55; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__28; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1698____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1448____spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSymbolTag(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__106; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__37; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3856____closed__2; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__5; LEAN_EXPORT lean_object* l_Lean_Lsp_DocumentHighlightKind_toCtorIdx___boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__63; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDocumentHighlightParams; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3857____closed__2; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3346____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__7; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__71; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__108; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__18; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1099____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__142; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__15; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__63; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1449____closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__71; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__7; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1370____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_112_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__24; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3345____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1448____closed__2; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__10; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__32; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__108; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3664____closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3346_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3345_(lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____spec__2(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3663____closed__1; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__20; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____spec__2(lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__30; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonFoldingRange; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1100____boxed(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2173_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__24; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__106; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2172_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__5; LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolKind_toCtorIdx(uint8_t); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__43; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1449____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__30; LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolInformation_containerName_x3f___default; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1448____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionOptions_triggerCharacters_x3f___default; LEAN_EXPORT lean_object* l_Lean_Lsp_instReprCompletionItemKind; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentHighlightParams; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3227_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1665_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3226_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__53; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1664_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__42; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__125; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__100; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__113; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1449_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____spec__3___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1448_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonReferenceContext; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1371____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__18; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__15; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__105; static lean_object* l_Lean_Lsp_instToJsonDocumentHighlightKind___closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__105; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__53; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__42; lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_1974____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonInsertReplaceEdit___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__113; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__16; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__76; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonDocumentSymbol_go___spec__3(size_t, size_t, lean_object*); static lean_object* l_Lean_Lsp_instToJsonReferenceParams___closed__1; lean_object* lean_nat_sub(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__18; LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenType_toNat(uint8_t); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3346____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__15; LEAN_EXPORT lean_object* l_Lean_Lsp_Hover_range_x3f___default; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__76; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__30; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__100; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3345____closed__1; static lean_object* l_Lean_Lsp_instToJsonCompletionItem___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__125; static lean_object* l_Lean_Lsp_instToJsonDocumentHighlightKind___closed__4; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1604_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__59; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1448____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__112; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1603_(lean_object*); lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_2125____spec__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Lsp_instInhabitedCompletionItemKind; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__24; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__135; LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolTag_toCtorIdx(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__11; static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____spec__1___closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__148; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3698_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionOptions_allCommitCharacters_x3f___default; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__82; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonDocumentSymbol_go___spec__3___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__75; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonInsertReplaceEdit; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1449____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolInformation_tags___default; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__26; static lean_object* l_Lean_Lsp_instFromJsonFoldingRangeParams___closed__1; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__7; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__11; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__82; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__26; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3739_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__114; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__148; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__58; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3856____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenType_noConfusion(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1100____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__114; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3857____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__112; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1370____spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_112____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1099____closed__1; static lean_object* l_Lean_Lsp_instFromJsonDefinitionParams___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonFoldingRangeKind(uint8_t); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3560_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__135; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3559_(lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1500____spec__1(size_t, size_t, lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__6; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1916_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonReferenceParams; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1967_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1966_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__34; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__75; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____at_Lean_Lsp_instToJsonDocumentSymbol_go___spec__1(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1501____spec__1(size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__61; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__127; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSymbolKind(uint8_t); lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_2095____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDocumentSymbolParams; static lean_object* l_Lean_Lsp_instToJsonCompletionParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__68; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__126; LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenType_toNat___boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__128; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSemanticTokensOptions; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__116; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3698____spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonSemanticTokensLegend___closed__1; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1371____spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonTypeDefinitionParams; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__136; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__126; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__68; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__12; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__128; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__23; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__142; lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__116; static lean_object* l_Lean_Lsp_instToJsonDocumentHighlightKind___closed__2; static lean_object* l_Lean_Lsp_instFromJsonWorkspaceSymbolParams___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3477_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3476_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__13; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__12; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__61; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__41; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__136; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__102; LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolKind_toCtorIdx___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____spec__3___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__127; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__1; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2256____spec__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3227____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3226____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____boxed(lean_object*); static lean_object* l_Lean_Lsp_instFromJsonTypeDefinitionParams___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentSymbolAux(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2070____boxed(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2019_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2648_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2600_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__41; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__13; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2018_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2069____boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2647_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2599_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__110; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__80; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__6; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3856____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3226____closed__1; static lean_object* l_Lean_Lsp_instToJsonDocumentHighlightKind___closed__5; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__50; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__96; static lean_object* l_Lean_Lsp_instFromJsonCompletionParams___closed__1; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__38; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__23; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__21; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__141; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__20; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__88; LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenType_noConfusion___rarg(uint8_t, uint8_t, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__20; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__110; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3227____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__13; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__41; static lean_object* l_Lean_Lsp_instInhabitedCompletionItem___closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolTag_toCtorIdx___boxed(lean_object*); static lean_object* l_Lean_Lsp_instToJsonHoverParams___closed__1; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDeclarationParams; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__23; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSemanticTokensLegend; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3088_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3087_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__22; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3857____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__141; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____spec__1(lean_object*); LEAN_EXPORT uint8_t l_Lean_Lsp_instDecidableEqCompletionItemKind(uint8_t, uint8_t); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3227____closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__111; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____spec__2(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__50; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonFoldingRangeParams; LEAN_EXPORT lean_object* l_Lean_Lsp_FoldingRange_kind_x3f___default; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____spec__3___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__80; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__111; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3226____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_FoldingRangeKind_noConfusion___rarg(uint8_t, uint8_t, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__96; size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonReferenceParams; static lean_object* l_Lean_Lsp_instFromJsonSemanticTokensRangeParams___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3612_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__48; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3227____boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3611_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCompletionItemKind___boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__34; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__93; LEAN_EXPORT lean_object* l_Lean_Lsp_DocumentHighlightKind_noConfusion(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3856____spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSemanticTokensOptions___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__72; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__107; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__21; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__38; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__22; lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_1451_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3226____boxed(lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__33; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__34; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__22; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1501_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__107; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1500_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__25; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__93; static lean_object* l_Lean_Lsp_instFromJsonDeclarationParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1100____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1099____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__38; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__21; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSemanticTokensParams; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__72; static lean_object* l_Lean_Lsp_instFromJsonCompletionOptions___closed__1; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____spec__1___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__29; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267_(lean_object*); lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_2478_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemKind_toCtorIdx(uint8_t); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3476____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__48; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCompletionParams; LEAN_EXPORT lean_object* l_Lean_Lsp_DocumentSymbolAux_detail_x3f___default; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemKind_noConfusion(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__98; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCompletionList; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCompletionItem; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__79; static lean_object* l_Lean_Lsp_instToJsonFoldingRangeKind___closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__81; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__120; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____boxed(lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__46; static lean_object* l_Lean_Lsp_instToJsonSymbolInformation___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_DocumentHighlight_kind_x3f___default; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__73; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__32; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__10; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__101; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__43; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonFoldingRangeParams; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__133; lean_object* l_Lean_JsonNumber_fromNat(lean_object*); lean_object* l_Lean_Json_getObjValD(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__101; static lean_object* l_Lean_Lsp_instToJsonFoldingRangeParams___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_FoldingRangeKind_noConfusion(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_FoldingRangeKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__81; LEAN_EXPORT lean_object* l_Lean_Lsp_DocumentHighlightKind_noConfusion___rarg(uint8_t, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonHoverParams; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__43; lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1074____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__73; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__32; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__10; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3226____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____spec__3(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__98; uint8_t lean_nat_dec_le(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3739____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__120; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1813_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1100____closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__123; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__133; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3345____spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3087____spec__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2599____spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1812_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__66; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__115; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__85; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__35; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1449____spec__2___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1099____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__99; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__56; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCompletionOptions; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__16; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3088____spec__1___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__49; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonHover; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentSymbol_go(lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__36; uint8_t lean_nat_dec_le(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__66; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__117; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__56; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__149; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1169_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__77; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__99; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3346____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3088____closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__88; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1168_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3087____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__49; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentHighlightKind(uint8_t); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3559____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__16; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__19; lean_object* l_Lean_Json_mkObj(lean_object*); lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1019____spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3560____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__123; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__77; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__117; static lean_object* l_Lean_Lsp_instToJsonDocumentHighlightKind___closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__115; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__85; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__149; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItem_textEdit_x3f___default; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__104; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__39; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__35; lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_2438_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__11; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__129; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__83; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2647____boxed(lean_object*); static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__5; lean_object* l_Lean_Json_pretty(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__35; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____at_Lean_Lsp_instToJsonDocumentSymbol_go___spec__1(lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__36; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2648____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__140; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2461____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__45; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1371_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____rarg___closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1370_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__145; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__78; LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenType_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonFoldingRange___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__51; LEAN_EXPORT lean_object* l_Lean_Lsp_DocumentHighlightKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__118; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__146; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____rarg___closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__19; static lean_object* l_Lean_Lsp_instToJsonCompletionList___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3279_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3278_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__36; static lean_object* l_Lean_Lsp_instToJsonDocumentSymbol___closed__1; static lean_object* l_Lean_Lsp_instToJsonSemanticTokensParams___closed__1; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1501____spec__1___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__51; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemKind_toCtorIdx___boxed(lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__2; static lean_object* l_Lean_Lsp_SymbolInformation_tags___default___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__145; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__45; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2462____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__104; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__129; static lean_object* l_Lean_Lsp_instToJsonFoldingRangeKind___closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__83; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255_(uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____spec__1___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254_(uint8_t, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSemanticTokensParams___closed__1; static lean_object* l_Lean_Lsp_instFromJsonReferenceContext___closed__1; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__44; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSemanticTokensOptions; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__70; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonWorkspaceSymbolParams; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3774_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3857____spec__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2070_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3512_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__109; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3773_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2069_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3511_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__65; lean_object* l_Lean_Json_getBool_x3f(lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__151; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__91; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__39; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__122; lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_931____spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentHighlight; LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolTag_noConfusion___rarg___boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__78; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSymbolKind___boxed(lean_object*); static lean_object* l_Lean_Lsp_instToJsonFoldingRangeKind___closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__65; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_112____spec__2___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__109; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_112____spec__2(size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__39; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__91; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__151; LEAN_EXPORT lean_object* l_Lean_Lsp_FoldingRangeKind_toCtorIdx___boxed(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1864____boxed(lean_object*); static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____spec__1___closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1698_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__17; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2256____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1697_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2255____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2255____boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1863____boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonWorkspaceSymbolParams; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2256____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__69; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__17; LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__70; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__8; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2683_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__54; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3477____boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2682_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__69; lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_617_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonReferenceContext; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__62; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonHover; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__45; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__137; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__54; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__62; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSemanticTokensLegend; LEAN_EXPORT lean_object* l_Lean_Lsp_DocumentHighlightKind_toCtorIdx(uint8_t); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____spec__2___rarg(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonHover___closed__1; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__9; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____spec__2(size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3857_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3664_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__33; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3856_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__103; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__132; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3663_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__1; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__9; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__121; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__40; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__14; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____rarg___closed__1; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__27; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1370____spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolTag_noConfusion(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1966____boxed(lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__4; static lean_object* l_Lean_Lsp_instToJsonDocumentSymbolParams___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____boxed(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2208____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__137; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2255____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSemanticTokensRangeParams; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__121; static lean_object* l_Lean_Lsp_instFromJsonReferenceParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__103; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__132; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__33; static lean_object* l_Lean_Lsp_instToJsonWorkspaceSymbolParams___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentSymbolResult(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__28; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCompletionParams; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__14; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__28; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__84; lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_359_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2207____boxed(lean_object*); static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__84; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonTypeDefinitionParams; static lean_object* l_Lean_Lsp_instToJsonDocumentHighlightParams___closed__1; static lean_object* l_Lean_Lsp_instFromJsonDocumentSymbolParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__46; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2256_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__29; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2255_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__46; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolKind_noConfusion___rarg(uint8_t, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenType_toCtorIdx___boxed(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____rarg(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__25; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3087____spec__1(size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__152; static lean_object* l_Lean_Lsp_instInhabitedCompletionItem___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonFoldingRangeKind___boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__119; lean_object* lean_nat_to_int(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__52; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__4; static lean_object* l_Lean_Lsp_instToJsonFoldingRangeKind___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__86; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____spec__2___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCompletionItemKind(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__130; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____spec__1(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2387_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1967____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__140; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2386_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__130; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__8; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2387____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__118; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2386____boxed(lean_object*); static lean_object* l_Lean_Lsp_instToJsonFoldingRangeKind___closed__6; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSymbolInformation; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__146; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__147; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__52; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__86; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__19; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__25; static lean_object* l_Lean_Lsp_instFromJsonCompletionList___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__152; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2422_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__29; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2421_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonSemanticTokensOptions___closed__1; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1371____spec__2(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2462_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2386____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2461_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCompletionList; lean_object* l_Repr_addAppParen(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__119; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2387____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__147; static lean_object* _init_l_Lean_Lsp_CompletionOptions_triggerCharacters_x3f___default() { _start: { @@ -1673,7 +1673,7 @@ x_6 = lean_box(x_5); return x_6; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__1() { _start: { lean_object* x_1; @@ -1681,17 +1681,17 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.text", 32); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__1; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__3() { _start: { lean_object* x_1; lean_object* x_2; @@ -1700,23 +1700,23 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__5() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__4; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__4; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -1724,7 +1724,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__6() { _start: { lean_object* x_1; lean_object* x_2; @@ -1733,23 +1733,23 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__8() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__7; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -1757,7 +1757,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__9() { _start: { lean_object* x_1; @@ -1765,33 +1765,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.method", 34); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__9; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__9; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__10; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__10; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__12() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__11; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__11; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -1799,23 +1799,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__10; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__10; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__14() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__13; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__13; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -1823,7 +1823,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__15() { _start: { lean_object* x_1; @@ -1831,33 +1831,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.function", 36); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__16() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__15; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__15; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__17() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__16; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__16; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__18() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__18() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__17; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__17; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -1865,23 +1865,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__19() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__16; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__16; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__20() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__20() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__19; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__19; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -1889,7 +1889,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__21() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__21() { _start: { lean_object* x_1; @@ -1897,33 +1897,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.constructor", 39); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__22() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__22() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__21; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__21; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__23() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__22; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__22; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__24() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__24() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__23; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__23; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -1931,23 +1931,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__25() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__22; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__22; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__26() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__26() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__25; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__25; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -1955,7 +1955,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__27() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__27() { _start: { lean_object* x_1; @@ -1963,33 +1963,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.field", 33); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__28() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__28() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__27; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__27; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__29() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__29() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__28; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__28; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__30() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__30() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__29; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__29; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -1997,23 +1997,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__31() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__31() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__28; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__28; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__32() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__32() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__31; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__31; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2021,7 +2021,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__33() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__33() { _start: { lean_object* x_1; @@ -2029,33 +2029,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.variable", 36); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__34() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__34() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__33; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__33; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__35() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__35() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__34; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__34; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__36() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__36() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__35; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__35; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2063,23 +2063,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__37() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__37() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__34; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__34; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__38() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__38() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__37; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__37; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2087,7 +2087,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__39() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__39() { _start: { lean_object* x_1; @@ -2095,33 +2095,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.class", 33); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__40() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__40() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__39; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__39; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__41() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__41() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__40; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__40; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__42() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__42() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__41; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__41; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2129,23 +2129,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__43() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__43() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__40; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__40; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__44() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__44() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__43; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__43; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2153,7 +2153,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__45() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__45() { _start: { lean_object* x_1; @@ -2161,33 +2161,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.interface", 37); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__46() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__46() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__45; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__45; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__47() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__47() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__46; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__46; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__48() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__48() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__47; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__47; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2195,23 +2195,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__49() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__49() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__46; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__46; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__50() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__50() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__49; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__49; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2219,7 +2219,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__51() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__51() { _start: { lean_object* x_1; @@ -2227,33 +2227,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.module", 34); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__52() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__52() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__51; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__51; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__53() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__53() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__52; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__52; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__54() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__54() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__53; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__53; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2261,23 +2261,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__55() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__55() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__52; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__52; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__56() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__56() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__55; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__55; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2285,7 +2285,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__57() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__57() { _start: { lean_object* x_1; @@ -2293,33 +2293,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.property", 36); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__58() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__58() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__57; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__57; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__59() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__59() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__58; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__58; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__60() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__60() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__59; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__59; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2327,23 +2327,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__61() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__61() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__58; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__58; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__62() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__62() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__61; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__61; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2351,7 +2351,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__63() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__63() { _start: { lean_object* x_1; @@ -2359,33 +2359,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.unit", 32); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__64() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__64() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__63; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__63; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__65() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__65() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__64; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__64; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__66() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__66() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__65; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__65; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2393,23 +2393,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__67() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__67() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__64; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__64; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__68() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__68() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__67; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__67; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2417,7 +2417,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__69() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__69() { _start: { lean_object* x_1; @@ -2425,33 +2425,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.value", 33); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__70() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__70() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__69; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__69; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__71() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__71() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__70; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__70; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__72() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__72() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__71; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__71; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2459,23 +2459,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__73() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__73() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__70; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__70; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__74() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__74() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__73; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__73; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2483,7 +2483,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__75() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__75() { _start: { lean_object* x_1; @@ -2491,33 +2491,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.enum", 32); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__76() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__76() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__75; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__75; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__77() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__77() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__76; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__76; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__78() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__78() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__77; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__77; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2525,23 +2525,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__79() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__79() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__76; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__76; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__80() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__80() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__79; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__79; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2549,7 +2549,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__81() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__81() { _start: { lean_object* x_1; @@ -2557,33 +2557,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.keyword", 35); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__82() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__82() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__81; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__81; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__83() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__83() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__82; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__82; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__84() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__84() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__83; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__83; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2591,23 +2591,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__85() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__85() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__82; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__82; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__86() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__86() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__85; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__85; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2615,7 +2615,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__87() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__87() { _start: { lean_object* x_1; @@ -2623,33 +2623,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.snippet", 35); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__88() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__88() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__87; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__87; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__89() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__89() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__88; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__88; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__90() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__90() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__89; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__89; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2657,23 +2657,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__91() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__91() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__88; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__88; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__92() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__92() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__91; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__91; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2681,7 +2681,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__93() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__93() { _start: { lean_object* x_1; @@ -2689,33 +2689,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.color", 33); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__94() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__94() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__93; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__93; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__95() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__95() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__94; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__94; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__96() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__96() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__95; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__95; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2723,23 +2723,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__97() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__97() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__94; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__94; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__98() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__98() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__97; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__97; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2747,7 +2747,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__99() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__99() { _start: { lean_object* x_1; @@ -2755,33 +2755,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.file", 32); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__100() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__100() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__99; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__99; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__101() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__101() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__100; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__100; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__102() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__102() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__101; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__101; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2789,23 +2789,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__103() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__103() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__100; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__100; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__104() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__104() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__103; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__103; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2813,7 +2813,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__105() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__105() { _start: { lean_object* x_1; @@ -2821,33 +2821,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.reference", 37); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__106() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__106() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__105; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__105; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__107() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__107() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__106; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__106; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__108() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__108() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__107; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__107; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2855,23 +2855,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__109() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__109() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__106; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__106; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__110() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__110() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__109; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__109; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2879,7 +2879,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__111() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__111() { _start: { lean_object* x_1; @@ -2887,33 +2887,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.folder", 34); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__112() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__112() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__111; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__111; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__113() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__113() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__112; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__112; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__114() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__114() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__113; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__113; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2921,23 +2921,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__115() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__115() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__112; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__112; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__116() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__116() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__115; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__115; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2945,7 +2945,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__117() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__117() { _start: { lean_object* x_1; @@ -2953,33 +2953,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.enumMember", 38); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__118() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__118() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__117; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__117; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__119() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__119() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__118; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__118; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__120() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__120() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__119; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__119; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2987,23 +2987,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__121() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__121() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__118; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__118; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__122() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__122() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__121; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__121; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3011,7 +3011,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__123() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__123() { _start: { lean_object* x_1; @@ -3019,33 +3019,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.constant", 36); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__124() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__124() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__123; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__123; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__125() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__125() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__124; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__124; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__126() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__126() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__125; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__125; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3053,23 +3053,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__127() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__127() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__124; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__124; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__128() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__128() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__127; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__127; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3077,7 +3077,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__129() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__129() { _start: { lean_object* x_1; @@ -3085,33 +3085,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.struct", 34); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__130() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__130() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__129; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__129; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__131() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__131() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__130; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__130; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__132() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__132() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__131; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__131; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3119,23 +3119,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__133() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__133() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__130; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__130; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__134() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__134() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__133; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__133; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3143,7 +3143,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__135() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__135() { _start: { lean_object* x_1; @@ -3151,33 +3151,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.event", 33); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__136() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__136() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__135; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__135; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__137() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__137() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__136; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__136; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__138() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__138() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__137; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__137; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3185,23 +3185,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__139() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__139() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__136; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__136; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__140() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__140() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__139; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__139; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3209,7 +3209,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__141() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__141() { _start: { lean_object* x_1; @@ -3217,33 +3217,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.operator", 36); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__142() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__142() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__141; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__141; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__143() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__143() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__142; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__142; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__144() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__144() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__143; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__143; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3251,23 +3251,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__145() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__145() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__142; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__142; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__146() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__146() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__145; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__145; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3275,7 +3275,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__147() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__147() { _start: { lean_object* x_1; @@ -3283,33 +3283,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.typeParameter", 41) return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__148() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__148() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__147; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__147; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__149() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__149() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__148; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__148; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__150() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__150() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__149; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__149; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3317,23 +3317,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__151() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__151() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__148; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__148; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__152() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__152() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__151; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__151; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3341,7 +3341,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255_(uint8_t x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254_(uint8_t x_1, lean_object* x_2) { _start: { switch (x_1) { @@ -3353,14 +3353,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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__5; +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__8; +x_7 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__8; x_8 = l_Repr_addAppParen(x_7, x_2); return x_8; } @@ -3373,14 +3373,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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__12; +x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__12; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__14; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__14; x_14 = l_Repr_addAppParen(x_13, x_2); return x_14; } @@ -3393,14 +3393,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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__18; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__18; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__20; +x_19 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__20; x_20 = l_Repr_addAppParen(x_19, x_2); return x_20; } @@ -3413,14 +3413,14 @@ x_22 = lean_nat_dec_le(x_21, x_2); if (x_22 == 0) { lean_object* x_23; lean_object* x_24; -x_23 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__24; +x_23 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__24; x_24 = l_Repr_addAppParen(x_23, x_2); return x_24; } else { lean_object* x_25; lean_object* x_26; -x_25 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__26; +x_25 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__26; x_26 = l_Repr_addAppParen(x_25, x_2); return x_26; } @@ -3433,14 +3433,14 @@ x_28 = lean_nat_dec_le(x_27, x_2); if (x_28 == 0) { lean_object* x_29; lean_object* x_30; -x_29 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__30; +x_29 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__30; x_30 = l_Repr_addAppParen(x_29, x_2); return x_30; } else { lean_object* x_31; lean_object* x_32; -x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__32; +x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__32; x_32 = l_Repr_addAppParen(x_31, x_2); return x_32; } @@ -3453,14 +3453,14 @@ x_34 = lean_nat_dec_le(x_33, x_2); if (x_34 == 0) { lean_object* x_35; lean_object* x_36; -x_35 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__36; +x_35 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__36; x_36 = l_Repr_addAppParen(x_35, x_2); return x_36; } else { lean_object* x_37; lean_object* x_38; -x_37 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__38; +x_37 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__38; x_38 = l_Repr_addAppParen(x_37, x_2); return x_38; } @@ -3473,14 +3473,14 @@ x_40 = lean_nat_dec_le(x_39, x_2); if (x_40 == 0) { lean_object* x_41; lean_object* x_42; -x_41 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__42; +x_41 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__42; x_42 = l_Repr_addAppParen(x_41, x_2); return x_42; } else { lean_object* x_43; lean_object* x_44; -x_43 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__44; +x_43 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__44; x_44 = l_Repr_addAppParen(x_43, x_2); return x_44; } @@ -3493,14 +3493,14 @@ x_46 = lean_nat_dec_le(x_45, x_2); if (x_46 == 0) { lean_object* x_47; lean_object* x_48; -x_47 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__48; +x_47 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__48; x_48 = l_Repr_addAppParen(x_47, x_2); return x_48; } else { lean_object* x_49; lean_object* x_50; -x_49 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__50; +x_49 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__50; x_50 = l_Repr_addAppParen(x_49, x_2); return x_50; } @@ -3513,14 +3513,14 @@ x_52 = lean_nat_dec_le(x_51, x_2); if (x_52 == 0) { lean_object* x_53; lean_object* x_54; -x_53 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__54; +x_53 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__54; x_54 = l_Repr_addAppParen(x_53, x_2); return x_54; } else { lean_object* x_55; lean_object* x_56; -x_55 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__56; +x_55 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__56; x_56 = l_Repr_addAppParen(x_55, x_2); return x_56; } @@ -3533,14 +3533,14 @@ x_58 = lean_nat_dec_le(x_57, x_2); if (x_58 == 0) { lean_object* x_59; lean_object* x_60; -x_59 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__60; +x_59 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__60; x_60 = l_Repr_addAppParen(x_59, x_2); return x_60; } else { lean_object* x_61; lean_object* x_62; -x_61 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__62; +x_61 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__62; x_62 = l_Repr_addAppParen(x_61, x_2); return x_62; } @@ -3553,14 +3553,14 @@ x_64 = lean_nat_dec_le(x_63, x_2); if (x_64 == 0) { lean_object* x_65; lean_object* x_66; -x_65 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__66; +x_65 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__66; x_66 = l_Repr_addAppParen(x_65, x_2); return x_66; } else { lean_object* x_67; lean_object* x_68; -x_67 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__68; +x_67 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__68; x_68 = l_Repr_addAppParen(x_67, x_2); return x_68; } @@ -3573,14 +3573,14 @@ x_70 = lean_nat_dec_le(x_69, x_2); if (x_70 == 0) { lean_object* x_71; lean_object* x_72; -x_71 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__72; +x_71 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__72; x_72 = l_Repr_addAppParen(x_71, x_2); return x_72; } else { lean_object* x_73; lean_object* x_74; -x_73 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__74; +x_73 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__74; x_74 = l_Repr_addAppParen(x_73, x_2); return x_74; } @@ -3593,14 +3593,14 @@ x_76 = lean_nat_dec_le(x_75, x_2); if (x_76 == 0) { lean_object* x_77; lean_object* x_78; -x_77 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__78; +x_77 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__78; x_78 = l_Repr_addAppParen(x_77, x_2); return x_78; } else { lean_object* x_79; lean_object* x_80; -x_79 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__80; +x_79 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__80; x_80 = l_Repr_addAppParen(x_79, x_2); return x_80; } @@ -3613,14 +3613,14 @@ x_82 = lean_nat_dec_le(x_81, x_2); if (x_82 == 0) { lean_object* x_83; lean_object* x_84; -x_83 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__84; +x_83 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__84; x_84 = l_Repr_addAppParen(x_83, x_2); return x_84; } else { lean_object* x_85; lean_object* x_86; -x_85 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__86; +x_85 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__86; x_86 = l_Repr_addAppParen(x_85, x_2); return x_86; } @@ -3633,14 +3633,14 @@ x_88 = lean_nat_dec_le(x_87, x_2); if (x_88 == 0) { lean_object* x_89; lean_object* x_90; -x_89 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__90; +x_89 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__90; x_90 = l_Repr_addAppParen(x_89, x_2); return x_90; } else { lean_object* x_91; lean_object* x_92; -x_91 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__92; +x_91 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__92; x_92 = l_Repr_addAppParen(x_91, x_2); return x_92; } @@ -3653,14 +3653,14 @@ x_94 = lean_nat_dec_le(x_93, x_2); if (x_94 == 0) { lean_object* x_95; lean_object* x_96; -x_95 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__96; +x_95 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__96; x_96 = l_Repr_addAppParen(x_95, x_2); return x_96; } else { lean_object* x_97; lean_object* x_98; -x_97 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__98; +x_97 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__98; x_98 = l_Repr_addAppParen(x_97, x_2); return x_98; } @@ -3673,14 +3673,14 @@ x_100 = lean_nat_dec_le(x_99, x_2); if (x_100 == 0) { lean_object* x_101; lean_object* x_102; -x_101 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__102; +x_101 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__102; x_102 = l_Repr_addAppParen(x_101, x_2); return x_102; } else { lean_object* x_103; lean_object* x_104; -x_103 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__104; +x_103 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__104; x_104 = l_Repr_addAppParen(x_103, x_2); return x_104; } @@ -3693,14 +3693,14 @@ x_106 = lean_nat_dec_le(x_105, x_2); if (x_106 == 0) { lean_object* x_107; lean_object* x_108; -x_107 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__108; +x_107 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__108; x_108 = l_Repr_addAppParen(x_107, x_2); return x_108; } else { lean_object* x_109; lean_object* x_110; -x_109 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__110; +x_109 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__110; x_110 = l_Repr_addAppParen(x_109, x_2); return x_110; } @@ -3713,14 +3713,14 @@ x_112 = lean_nat_dec_le(x_111, x_2); if (x_112 == 0) { lean_object* x_113; lean_object* x_114; -x_113 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__114; +x_113 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__114; x_114 = l_Repr_addAppParen(x_113, x_2); return x_114; } else { lean_object* x_115; lean_object* x_116; -x_115 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__116; +x_115 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__116; x_116 = l_Repr_addAppParen(x_115, x_2); return x_116; } @@ -3733,14 +3733,14 @@ x_118 = lean_nat_dec_le(x_117, x_2); if (x_118 == 0) { lean_object* x_119; lean_object* x_120; -x_119 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__120; +x_119 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__120; x_120 = l_Repr_addAppParen(x_119, x_2); return x_120; } else { lean_object* x_121; lean_object* x_122; -x_121 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__122; +x_121 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__122; x_122 = l_Repr_addAppParen(x_121, x_2); return x_122; } @@ -3753,14 +3753,14 @@ x_124 = lean_nat_dec_le(x_123, x_2); if (x_124 == 0) { lean_object* x_125; lean_object* x_126; -x_125 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__126; +x_125 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__126; x_126 = l_Repr_addAppParen(x_125, x_2); return x_126; } else { lean_object* x_127; lean_object* x_128; -x_127 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__128; +x_127 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__128; x_128 = l_Repr_addAppParen(x_127, x_2); return x_128; } @@ -3773,14 +3773,14 @@ x_130 = lean_nat_dec_le(x_129, x_2); if (x_130 == 0) { lean_object* x_131; lean_object* x_132; -x_131 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__132; +x_131 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__132; x_132 = l_Repr_addAppParen(x_131, x_2); return x_132; } else { lean_object* x_133; lean_object* x_134; -x_133 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__134; +x_133 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__134; x_134 = l_Repr_addAppParen(x_133, x_2); return x_134; } @@ -3793,14 +3793,14 @@ x_136 = lean_nat_dec_le(x_135, x_2); if (x_136 == 0) { lean_object* x_137; lean_object* x_138; -x_137 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__138; +x_137 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__138; x_138 = l_Repr_addAppParen(x_137, x_2); return x_138; } else { lean_object* x_139; lean_object* x_140; -x_139 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__140; +x_139 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__140; x_140 = l_Repr_addAppParen(x_139, x_2); return x_140; } @@ -3813,14 +3813,14 @@ x_142 = lean_nat_dec_le(x_141, x_2); if (x_142 == 0) { lean_object* x_143; lean_object* x_144; -x_143 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__144; +x_143 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__144; x_144 = l_Repr_addAppParen(x_143, x_2); return x_144; } else { lean_object* x_145; lean_object* x_146; -x_145 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__146; +x_145 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__146; x_146 = l_Repr_addAppParen(x_145, x_2); return x_146; } @@ -3833,14 +3833,14 @@ x_148 = lean_nat_dec_le(x_147, x_2); if (x_148 == 0) { lean_object* x_149; lean_object* x_150; -x_149 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__150; +x_149 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__150; x_150 = l_Repr_addAppParen(x_149, x_2); return x_150; } else { lean_object* x_151; lean_object* x_152; -x_151 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__152; +x_151 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__152; x_152 = l_Repr_addAppParen(x_151, x_2); return x_152; } @@ -3848,13 +3848,13 @@ return x_152; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255_(x_3, x_2); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254_(x_3, x_2); lean_dec(x_2); return x_4; } @@ -3863,7 +3863,7 @@ static lean_object* _init_l_Lean_Lsp_instReprCompletionItemKind___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____boxed), 2, 0); return x_1; } } @@ -3968,7 +3968,7 @@ lean_dec(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1100____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1099____closed__1() { _start: { lean_object* x_1; @@ -3976,7 +3976,7 @@ x_1 = lean_mk_string_from_bytes("newText", 7); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1100____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1099____closed__2() { _start: { lean_object* x_1; @@ -3984,7 +3984,7 @@ x_1 = lean_mk_string_from_bytes("insert", 6); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1100____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1099____closed__3() { _start: { lean_object* x_1; @@ -3992,11 +3992,11 @@ x_1 = lean_mk_string_from_bytes("replace", 7); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1100_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1099_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1100____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1099____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -4023,7 +4023,7 @@ lean_object* x_7; lean_object* x_8; lean_object* x_9; x_7 = lean_ctor_get(x_3, 0); lean_inc(x_7); lean_dec(x_3); -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1100____closed__2; +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1099____closed__2; x_9 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_931____spec__2(x_1, x_8); if (lean_obj_tag(x_9) == 0) { @@ -4051,7 +4051,7 @@ lean_object* x_13; lean_object* x_14; lean_object* x_15; x_13 = lean_ctor_get(x_9, 0); lean_inc(x_13); lean_dec(x_9); -x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1100____closed__3; +x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1099____closed__3; x_15 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_931____spec__2(x_1, x_14); if (lean_obj_tag(x_15) == 0) { @@ -4108,11 +4108,11 @@ return x_24; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1100____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1099____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1100_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1099_(x_1); lean_dec(x_1); return x_2; } @@ -4121,7 +4121,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonInsertReplaceEdit___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1100____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1099____boxed), 1, 0); return x_1; } } @@ -4133,7 +4133,7 @@ x_1 = l_Lean_Lsp_instFromJsonInsertReplaceEdit___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1169_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1168_(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_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; 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; @@ -4146,7 +4146,7 @@ lean_inc(x_4); lean_dec(x_1); x_5 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_5, 0, x_2); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1100____closed__1; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1099____closed__1; x_7 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); @@ -4155,7 +4155,7 @@ x_9 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_9, 0, x_7); lean_ctor_set(x_9, 1, x_8); x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_617_(x_3); -x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1100____closed__2; +x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1099____closed__2; x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_10); @@ -4163,7 +4163,7 @@ x_13 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_13, 0, x_12); lean_ctor_set(x_13, 1, x_8); x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_617_(x_4); -x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1100____closed__3; +x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1099____closed__3; x_16 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_16, 0, x_15); lean_ctor_set(x_16, 1, x_14); @@ -4188,7 +4188,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonInsertReplaceEdit___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1169_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1168_), 1, 0); return x_1; } } @@ -4208,7 +4208,7 @@ x_1 = lean_box(0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -4272,7 +4272,7 @@ return x_14; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____spec__2(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -4348,7 +4348,7 @@ return x_22; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____spec__3(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____spec__3(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -4362,7 +4362,7 @@ return x_4; else { lean_object* x_5; -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1100_(x_3); +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1099_(x_3); lean_dec(x_3); if (lean_obj_tag(x_5) == 0) { @@ -4412,7 +4412,7 @@ return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__1() { _start: { lean_object* x_1; @@ -4420,7 +4420,7 @@ x_1 = lean_mk_string_from_bytes("label", 5); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__2() { _start: { lean_object* x_1; @@ -4428,7 +4428,7 @@ x_1 = lean_mk_string_from_bytes("detail", 6); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__3() { _start: { lean_object* x_1; @@ -4436,7 +4436,7 @@ x_1 = lean_mk_string_from_bytes("documentation", 13); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__4() { _start: { lean_object* x_1; @@ -4444,7 +4444,7 @@ x_1 = lean_mk_string_from_bytes("kind", 4); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__5() { _start: { lean_object* x_1; @@ -4452,11 +4452,11 @@ x_1 = lean_mk_string_from_bytes("textEdit", 8); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -4483,7 +4483,7 @@ lean_object* x_7; lean_object* x_8; lean_object* x_9; x_7 = lean_ctor_get(x_3, 0); lean_inc(x_7); lean_dec(x_3); -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__2; +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__2; x_9 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_2125____spec__1(x_1, x_8); if (lean_obj_tag(x_9) == 0) { @@ -4511,8 +4511,8 @@ lean_object* x_13; lean_object* x_14; lean_object* x_15; x_13 = lean_ctor_get(x_9, 0); lean_inc(x_13); lean_dec(x_9); -x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__3; -x_15 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____spec__1(x_1, x_14); +x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__3; +x_15 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____spec__1(x_1, x_14); if (lean_obj_tag(x_15) == 0) { uint8_t x_16; @@ -4540,8 +4540,8 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; x_19 = lean_ctor_get(x_15, 0); lean_inc(x_19); lean_dec(x_15); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__4; -x_21 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____spec__2(x_1, x_20); +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__4; +x_21 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____spec__2(x_1, x_20); if (lean_obj_tag(x_21) == 0) { uint8_t x_22; @@ -4570,8 +4570,8 @@ lean_object* x_25; lean_object* x_26; lean_object* x_27; x_25 = lean_ctor_get(x_21, 0); lean_inc(x_25); lean_dec(x_21); -x_26 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__5; -x_27 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____spec__3(x_1, x_26); +x_26 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__5; +x_27 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____spec__3(x_1, x_26); if (lean_obj_tag(x_27) == 0) { uint8_t x_28; @@ -4635,41 +4635,41 @@ return x_36; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____spec__1(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____spec__2___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____spec__2___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____spec__2(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____spec__2(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____spec__3___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____spec__3___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____spec__3(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____spec__3(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267_(x_1); lean_dec(x_1); return x_2; } @@ -4678,7 +4678,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonCompletionItem___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____boxed), 1, 0); return x_1; } } @@ -4690,7 +4690,7 @@ x_1 = l_Lean_Lsp_instFromJsonCompletionItem___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1371____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1370____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -4716,7 +4716,7 @@ return x_8; } } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1371____spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1370____spec__2(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -4752,7 +4752,7 @@ return x_13; } } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1371____spec__3(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1370____spec__3(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -4768,7 +4768,7 @@ lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_obj x_4 = lean_ctor_get(x_2, 0); lean_inc(x_4); lean_dec(x_2); -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1169_(x_4); +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1168_(x_4); x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_1); lean_ctor_set(x_6, 1, x_5); @@ -4780,7 +4780,7 @@ return x_8; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1371_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1370_(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_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; 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; lean_object* x_25; lean_object* x_26; @@ -4797,7 +4797,7 @@ lean_inc(x_6); lean_dec(x_1); x_7 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_7, 0, x_2); -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__1; +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__1; x_9 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_9, 0, x_8); lean_ctor_set(x_9, 1, x_7); @@ -4805,16 +4805,16 @@ x_10 = lean_box(0); x_11 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_11, 0, x_9); lean_ctor_set(x_11, 1, x_10); -x_12 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__2; +x_12 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__2; x_13 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_2095____spec__1(x_12, x_3); lean_dec(x_3); -x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__3; -x_15 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1371____spec__1(x_14, x_4); +x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__3; +x_15 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1370____spec__1(x_14, x_4); lean_dec(x_4); -x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__4; -x_17 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1371____spec__2(x_16, x_5); -x_18 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__5; -x_19 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1371____spec__3(x_18, x_6); +x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__4; +x_17 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1370____spec__2(x_16, x_5); +x_18 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__5; +x_19 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1370____spec__3(x_18, x_6); x_20 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_20, 0, x_19); lean_ctor_set(x_20, 1, x_10); @@ -4835,11 +4835,11 @@ x_26 = l_Lean_Json_mkObj(x_25); return x_26; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1371____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1370____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1371____spec__1(x_1, x_2); +x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1370____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -4848,7 +4848,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonCompletionItem___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1371_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1370_), 1, 0); return x_1; } } @@ -4891,7 +4891,7 @@ x_1 = l_Lean_Lsp_instInhabitedCompletionItem___closed__2; return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1449____spec__2(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1448____spec__2(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -4909,7 +4909,7 @@ lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; x_6 = lean_array_uget(x_3, x_2); x_7 = lean_unsigned_to_nat(0u); x_8 = lean_array_uset(x_3, x_2, x_7); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268_(x_6); +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267_(x_6); lean_dec(x_6); if (lean_obj_tag(x_9) == 0) { @@ -4947,7 +4947,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1449____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1448____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -4964,7 +4964,7 @@ x_5 = lean_array_get_size(x_4); x_6 = lean_usize_of_nat(x_5); lean_dec(x_5); x_7 = 0; -x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1449____spec__2(x_6, x_7, x_4); +x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1448____spec__2(x_6, x_7, x_4); return x_8; } else @@ -4983,7 +4983,7 @@ return x_15; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1449____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1448____closed__1() { _start: { lean_object* x_1; @@ -4991,7 +4991,7 @@ x_1 = lean_mk_string_from_bytes("isIncomplete", 12); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1449____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1448____closed__2() { _start: { lean_object* x_1; @@ -4999,11 +4999,11 @@ x_1 = lean_mk_string_from_bytes("items", 5); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1449_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1448_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1449____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1448____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____spec__3(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -5031,8 +5031,8 @@ lean_object* x_7; lean_object* x_8; lean_object* x_9; x_7 = lean_ctor_get(x_3, 0); lean_inc(x_7); lean_dec(x_3); -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1449____closed__2; -x_9 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1449____spec__1(x_1, x_8); +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1448____closed__2; +x_9 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1448____spec__1(x_1, x_8); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; @@ -5088,7 +5088,7 @@ return x_20; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1449____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1448____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; size_t x_5; lean_object* x_6; @@ -5096,7 +5096,7 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1449____spec__2(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1448____spec__2(x_4, x_5, x_3); return x_6; } } @@ -5104,7 +5104,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonCompletionList___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1449_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1448_), 1, 0); return x_1; } } @@ -5116,7 +5116,7 @@ x_1 = l_Lean_Lsp_instFromJsonCompletionList___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1501____spec__1(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1500____spec__1(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -5131,7 +5131,7 @@ lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x x_5 = lean_array_uget(x_3, x_2); x_6 = lean_unsigned_to_nat(0u); x_7 = lean_array_uset(x_3, x_2, x_6); -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1371_(x_5); +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1370_(x_5); x_9 = 1; x_10 = lean_usize_add(x_2, x_9); x_11 = lean_array_uset(x_7, x_2, x_8); @@ -5141,7 +5141,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1501_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1500_(lean_object* x_1) { _start: { uint8_t 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; size_t x_10; size_t 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; lean_object* x_20; @@ -5151,7 +5151,7 @@ lean_inc(x_3); lean_dec(x_1); x_4 = lean_alloc_ctor(1, 0, 1); lean_ctor_set_uint8(x_4, 0, x_2); -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1449____closed__1; +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1448____closed__1; x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_5); lean_ctor_set(x_6, 1, x_4); @@ -5163,10 +5163,10 @@ x_9 = lean_array_get_size(x_3); x_10 = lean_usize_of_nat(x_9); lean_dec(x_9); x_11 = 0; -x_12 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1501____spec__1(x_10, x_11, x_3); +x_12 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1500____spec__1(x_10, x_11, x_3); x_13 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_13, 0, x_12); -x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1449____closed__2; +x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1448____closed__2; x_15 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_15, 0, x_14); lean_ctor_set(x_15, 1, x_13); @@ -5184,7 +5184,7 @@ x_20 = l_Lean_Json_mkObj(x_19); return x_20; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1501____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1500____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; size_t x_5; lean_object* x_6; @@ -5192,7 +5192,7 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1501____spec__1(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1500____spec__1(x_4, x_5, x_3); return x_6; } } @@ -5200,7 +5200,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonCompletionList___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1501_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1500_), 1, 0); return x_1; } } @@ -5212,7 +5212,7 @@ x_1 = l_Lean_Lsp_instToJsonCompletionList___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__1() { _start: { lean_object* x_1; @@ -5220,7 +5220,7 @@ x_1 = lean_mk_string_from_bytes("textDocument", 12); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__2() { _start: { lean_object* x_1; @@ -5228,11 +5228,11 @@ x_1 = lean_mk_string_from_bytes("position", 8); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_1974____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -5259,7 +5259,7 @@ lean_object* x_7; lean_object* x_8; lean_object* x_9; x_7 = lean_ctor_get(x_3, 0); lean_inc(x_7); lean_dec(x_3); -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__2; +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__2; x_9 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_657____spec__1(x_1, x_8); if (lean_obj_tag(x_9) == 0) { @@ -5312,11 +5312,11 @@ return x_18; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551_(x_1); lean_dec(x_1); return x_2; } @@ -5325,7 +5325,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonCompletionParams___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____boxed), 1, 0); return x_1; } } @@ -5337,14 +5337,14 @@ x_1 = l_Lean_Lsp_instFromJsonCompletionParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1604_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1603_(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_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; lean_object* x_15; lean_object* x_16; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_1451_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -5356,7 +5356,7 @@ x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_359_(x_8); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__2; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__2; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -5378,7 +5378,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonCompletionParams___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1604_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1603_), 1, 0); return x_1; } } @@ -5398,7 +5398,7 @@ x_1 = lean_box(0); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1665____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1664____closed__1() { _start: { lean_object* x_1; @@ -5406,7 +5406,7 @@ x_1 = lean_mk_string_from_bytes("contents", 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1665____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1664____closed__2() { _start: { lean_object* x_1; @@ -5414,7 +5414,7 @@ x_1 = lean_mk_string_from_bytes("range", 5); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1665_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1664_(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_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; @@ -5422,7 +5422,7 @@ x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_2438_(x_2); lean_dec(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1665____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1664____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -5433,7 +5433,7 @@ lean_ctor_set(x_7, 1, x_6); x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1665____closed__2; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1664____closed__2; x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1019____spec__1(x_9, x_8); x_11 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -5450,7 +5450,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonHover___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1665_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1664_), 1, 0); return x_1; } } @@ -5462,7 +5462,7 @@ x_1 = l_Lean_Lsp_instToJsonHover___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1698____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1697____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; @@ -5472,12 +5472,12 @@ lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1698_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1697_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1665____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1698____spec__1(x_1, x_2); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1664____closed__1; +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1697____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -5503,7 +5503,7 @@ lean_object* x_7; lean_object* x_8; lean_object* x_9; x_7 = lean_ctor_get(x_3, 0); lean_inc(x_7); lean_dec(x_3); -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1665____closed__2; +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1664____closed__2; x_9 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1074____spec__1(x_1, x_8); if (lean_obj_tag(x_9) == 0) { @@ -5556,21 +5556,21 @@ return x_18; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1698____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1697____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1698____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1697____spec__1(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1698____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1697____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1698_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1697_(x_1); lean_dec(x_1); return x_2; } @@ -5579,7 +5579,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonHover___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1698____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1697____boxed), 1, 0); return x_1; } } @@ -5591,11 +5591,11 @@ x_1 = l_Lean_Lsp_instFromJsonHover___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1761_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1760_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_1974____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -5622,7 +5622,7 @@ lean_object* x_7; lean_object* x_8; lean_object* x_9; x_7 = lean_ctor_get(x_3, 0); lean_inc(x_7); lean_dec(x_3); -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__2; +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__2; x_9 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_657____spec__1(x_1, x_8); if (lean_obj_tag(x_9) == 0) { @@ -5675,11 +5675,11 @@ return x_18; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1761____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1760____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1761_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1760_(x_1); lean_dec(x_1); return x_2; } @@ -5688,7 +5688,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonHoverParams___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1761____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1760____boxed), 1, 0); return x_1; } } @@ -5700,14 +5700,14 @@ x_1 = l_Lean_Lsp_instFromJsonHoverParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1813_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1812_(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_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; lean_object* x_15; lean_object* x_16; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_1451_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -5719,7 +5719,7 @@ x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_359_(x_8); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__2; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__2; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -5741,7 +5741,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonHoverParams___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1813_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1812_), 1, 0); return x_1; } } @@ -5753,11 +5753,11 @@ x_1 = l_Lean_Lsp_instToJsonHoverParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1864_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1863_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_1974____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -5784,7 +5784,7 @@ lean_object* x_7; lean_object* x_8; lean_object* x_9; x_7 = lean_ctor_get(x_3, 0); lean_inc(x_7); lean_dec(x_3); -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__2; +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__2; x_9 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_657____spec__1(x_1, x_8); if (lean_obj_tag(x_9) == 0) { @@ -5837,11 +5837,11 @@ return x_18; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1864____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1863____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1864_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1863_(x_1); lean_dec(x_1); return x_2; } @@ -5850,7 +5850,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonDeclarationParams___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1864____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1863____boxed), 1, 0); return x_1; } } @@ -5862,14 +5862,14 @@ x_1 = l_Lean_Lsp_instFromJsonDeclarationParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1916_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915_(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_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; lean_object* x_15; lean_object* x_16; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_1451_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -5881,7 +5881,7 @@ x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_359_(x_8); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__2; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__2; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -5903,7 +5903,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonDeclarationParams___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1916_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915_), 1, 0); return x_1; } } @@ -5915,11 +5915,11 @@ x_1 = l_Lean_Lsp_instToJsonDeclarationParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1967_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1966_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_1974____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -5946,7 +5946,7 @@ lean_object* x_7; lean_object* x_8; lean_object* x_9; x_7 = lean_ctor_get(x_3, 0); lean_inc(x_7); lean_dec(x_3); -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__2; +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__2; x_9 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_657____spec__1(x_1, x_8); if (lean_obj_tag(x_9) == 0) { @@ -5999,11 +5999,11 @@ return x_18; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1967____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1966____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1967_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1966_(x_1); lean_dec(x_1); return x_2; } @@ -6012,7 +6012,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonDefinitionParams___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1967____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1966____boxed), 1, 0); return x_1; } } @@ -6024,14 +6024,14 @@ x_1 = l_Lean_Lsp_instFromJsonDefinitionParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2019_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2018_(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_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; lean_object* x_15; lean_object* x_16; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_1451_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -6043,7 +6043,7 @@ x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_359_(x_8); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__2; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__2; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -6065,7 +6065,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonDefinitionParams___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2019_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2018_), 1, 0); return x_1; } } @@ -6077,11 +6077,11 @@ x_1 = l_Lean_Lsp_instToJsonDefinitionParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2070_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2069_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_1974____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -6108,7 +6108,7 @@ lean_object* x_7; lean_object* x_8; lean_object* x_9; x_7 = lean_ctor_get(x_3, 0); lean_inc(x_7); lean_dec(x_3); -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__2; +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__2; x_9 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_657____spec__1(x_1, x_8); if (lean_obj_tag(x_9) == 0) { @@ -6161,11 +6161,11 @@ return x_18; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2070____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2069____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2070_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2069_(x_1); lean_dec(x_1); return x_2; } @@ -6174,7 +6174,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonTypeDefinitionParams___closed__ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2070____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2069____boxed), 1, 0); return x_1; } } @@ -6186,14 +6186,14 @@ x_1 = l_Lean_Lsp_instFromJsonTypeDefinitionParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2122_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2121_(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_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; lean_object* x_15; lean_object* x_16; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_1451_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -6205,7 +6205,7 @@ x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_359_(x_8); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__2; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__2; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -6227,7 +6227,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonTypeDefinitionParams___closed__1( _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2122_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2121_), 1, 0); return x_1; } } @@ -6239,7 +6239,7 @@ x_1 = l_Lean_Lsp_instToJsonTypeDefinitionParams___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2173____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2172____closed__1() { _start: { lean_object* x_1; @@ -6247,11 +6247,11 @@ x_1 = lean_mk_string_from_bytes("includeDeclaration", 18); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2173_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2172_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2173____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2172____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____spec__3(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -6293,11 +6293,11 @@ return x_9; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2173____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2172____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2173_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2172_(x_1); lean_dec(x_1); return x_2; } @@ -6306,7 +6306,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonReferenceContext___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2173____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2172____boxed), 1, 0); return x_1; } } @@ -6318,13 +6318,13 @@ x_1 = l_Lean_Lsp_instFromJsonReferenceContext___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2208_(uint8_t x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2207_(uint8_t 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_8; lean_object* x_9; x_2 = lean_alloc_ctor(1, 0, 1); lean_ctor_set_uint8(x_2, 0, x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2173____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2172____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); lean_ctor_set(x_4, 1, x_2); @@ -6340,13 +6340,13 @@ x_9 = l_Lean_Json_mkObj(x_8); return x_9; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2208____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2207____boxed(lean_object* x_1) { _start: { uint8_t x_2; lean_object* x_3; x_2 = lean_unbox(x_1); lean_dec(x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2208_(x_2); +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2207_(x_2); return x_3; } } @@ -6354,7 +6354,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonReferenceContext___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2208____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2207____boxed), 1, 0); return x_1; } } @@ -6366,17 +6366,17 @@ x_1 = l_Lean_Lsp_instToJsonReferenceContext___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2256____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2255____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; x_3 = l_Lean_Json_getObjValD(x_1, x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2173_(x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2172_(x_3); lean_dec(x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2256____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2255____closed__1() { _start: { lean_object* x_1; @@ -6384,11 +6384,11 @@ x_1 = lean_mk_string_from_bytes("context", 7); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2256_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2255_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_1974____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -6415,7 +6415,7 @@ lean_object* x_7; lean_object* x_8; lean_object* x_9; x_7 = lean_ctor_get(x_3, 0); lean_inc(x_7); lean_dec(x_3); -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__2; +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__2; x_9 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_657____spec__1(x_1, x_8); if (lean_obj_tag(x_9) == 0) { @@ -6443,8 +6443,8 @@ lean_object* x_13; lean_object* x_14; lean_object* x_15; x_13 = lean_ctor_get(x_9, 0); lean_inc(x_13); lean_dec(x_9); -x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2256____closed__1; -x_15 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2256____spec__1(x_1, x_14); +x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2255____closed__1; +x_15 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2255____spec__1(x_1, x_14); if (lean_obj_tag(x_15) == 0) { uint8_t x_16; @@ -6504,21 +6504,21 @@ return x_26; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2256____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2255____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2256____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2255____spec__1(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2256____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2255____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2256_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2255_(x_1); lean_dec(x_1); return x_2; } @@ -6527,7 +6527,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonReferenceParams___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2256____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2255____boxed), 1, 0); return x_1; } } @@ -6539,7 +6539,7 @@ x_1 = l_Lean_Lsp_instFromJsonReferenceParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2325_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2324_(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_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; uint8_t 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; @@ -6548,7 +6548,7 @@ lean_inc(x_2); x_3 = lean_ctor_get(x_2, 0); lean_inc(x_3); x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_1451_(x_3); -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__1; +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__1; x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_5); lean_ctor_set(x_6, 1, x_4); @@ -6560,7 +6560,7 @@ x_9 = lean_ctor_get(x_2, 1); lean_inc(x_9); lean_dec(x_2); x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_359_(x_9); -x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__2; +x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__2; x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_10); @@ -6572,8 +6572,8 @@ lean_inc(x_14); lean_dec(x_1); x_15 = lean_unbox(x_14); lean_dec(x_14); -x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2208_(x_15); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2256____closed__1; +x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2207_(x_15); +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2255____closed__1; x_18 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_18, 0, x_17); lean_ctor_set(x_18, 1, x_16); @@ -6598,7 +6598,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonReferenceParams___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2325_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2324_), 1, 0); return x_1; } } @@ -6610,7 +6610,7 @@ x_1 = l_Lean_Lsp_instToJsonReferenceParams___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2387____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2386____closed__1() { _start: { lean_object* x_1; @@ -6618,11 +6618,11 @@ x_1 = lean_mk_string_from_bytes("query", 5); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2387_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2386_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2387____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2386____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -6664,11 +6664,11 @@ return x_9; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2387____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2386____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2387_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2386_(x_1); lean_dec(x_1); return x_2; } @@ -6677,7 +6677,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonWorkspaceSymbolParams___closed_ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2387____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2386____boxed), 1, 0); return x_1; } } @@ -6689,13 +6689,13 @@ x_1 = l_Lean_Lsp_instFromJsonWorkspaceSymbolParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2422_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2421_(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_8; lean_object* x_9; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2387____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2386____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); lean_ctor_set(x_4, 1, x_2); @@ -6715,7 +6715,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonWorkspaceSymbolParams___closed__1 _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2422_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2421_), 1, 0); return x_1; } } @@ -6727,11 +6727,11 @@ x_1 = l_Lean_Lsp_instToJsonWorkspaceSymbolParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2462_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2461_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_1974____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -6758,7 +6758,7 @@ lean_object* x_7; lean_object* x_8; lean_object* x_9; x_7 = lean_ctor_get(x_3, 0); lean_inc(x_7); lean_dec(x_3); -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__2; +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__2; x_9 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_657____spec__1(x_1, x_8); if (lean_obj_tag(x_9) == 0) { @@ -6811,11 +6811,11 @@ return x_18; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2462____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2461____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2462_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2461_(x_1); lean_dec(x_1); return x_2; } @@ -6824,7 +6824,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonDocumentHighlightParams___close _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2462____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2461____boxed), 1, 0); return x_1; } } @@ -6836,14 +6836,14 @@ x_1 = l_Lean_Lsp_instFromJsonDocumentHighlightParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2514_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2513_(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_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; lean_object* x_15; lean_object* x_16; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_1451_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -6855,7 +6855,7 @@ x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_359_(x_8); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__2; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__2; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -6877,7 +6877,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonDocumentHighlightParams___closed_ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2514_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2513_), 1, 0); return x_1; } } @@ -7052,7 +7052,7 @@ x_1 = lean_box(0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2600____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2599____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -7109,14 +7109,14 @@ return x_15; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2600_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2599_(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_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_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_617_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1665____closed__2; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1664____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); @@ -7127,8 +7127,8 @@ lean_ctor_set(x_7, 1, x_6); x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__4; -x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2600____spec__1(x_9, x_8); +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__4; +x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2599____spec__1(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -7141,11 +7141,11 @@ x_14 = l_Lean_Json_mkObj(x_13); return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2600____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2599____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2600____spec__1(x_1, x_2); +x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2599____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -7154,7 +7154,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonDocumentHighlight___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2600_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2599_), 1, 0); return x_1; } } @@ -7166,11 +7166,11 @@ x_1 = l_Lean_Lsp_instToJsonDocumentHighlight___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2648_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2647_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_1974____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -7212,11 +7212,11 @@ return x_9; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2648____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2647____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2648_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2647_(x_1); lean_dec(x_1); return x_2; } @@ -7225,7 +7225,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonDocumentSymbolParams___closed__ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2648____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2647____boxed), 1, 0); return x_1; } } @@ -7237,12 +7237,12 @@ x_1 = l_Lean_Lsp_instFromJsonDocumentSymbolParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2683_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2682_(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_8; lean_object* x_9; x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_1451_(x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); lean_ctor_set(x_4, 1, x_2); @@ -7262,7 +7262,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonDocumentSymbolParams___closed__1( _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2683_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2682_), 1, 0); return x_1; } } @@ -8101,7 +8101,7 @@ x_2 = lean_box(0); return x_2; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____spec__2___rarg(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____spec__2___rarg(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -8128,15 +8128,15 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____spec__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____spec__2(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____spec__2___rarg___boxed), 4, 0); +x_2 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____spec__2___rarg___boxed), 4, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_3) == 0) @@ -8157,7 +8157,7 @@ x_6 = lean_array_get_size(x_5); x_7 = lean_usize_of_nat(x_6); lean_dec(x_6); x_8 = 0; -x_9 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____spec__2___rarg(x_1, x_7, x_8, x_5); +x_9 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____spec__2___rarg(x_1, x_7, x_8, x_5); x_10 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_10, 0, x_9); x_11 = lean_alloc_ctor(0, 2, 0); @@ -8171,15 +8171,15 @@ return x_13; } } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____spec__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____spec__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____spec__1___rarg), 3, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____spec__1___rarg), 3, 0); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____rarg___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____rarg___closed__1() { _start: { lean_object* x_1; @@ -8187,7 +8187,7 @@ x_1 = lean_mk_string_from_bytes("name", 4); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____rarg___closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____rarg___closed__2() { _start: { lean_object* x_1; @@ -8195,7 +8195,7 @@ x_1 = lean_mk_string_from_bytes("selectionRange", 14); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____rarg___closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____rarg___closed__3() { _start: { lean_object* x_1; @@ -8203,7 +8203,7 @@ x_1 = lean_mk_string_from_bytes("children", 8); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____rarg(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____rarg(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; 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; 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; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; @@ -8221,7 +8221,7 @@ lean_inc(x_8); lean_dec(x_2); x_9 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_9, 0, x_3); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____rarg___closed__1; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____rarg___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); @@ -8229,11 +8229,11 @@ x_12 = lean_box(0); x_13 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_13, 0, x_11); lean_ctor_set(x_13, 1, x_12); -x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__2; +x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__2; x_15 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_2095____spec__1(x_14, x_4); lean_dec(x_4); x_16 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_617_(x_6); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1665____closed__2; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1664____closed__2; x_18 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_18, 0, x_17); lean_ctor_set(x_18, 1, x_16); @@ -8241,15 +8241,15 @@ x_19 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_19, 0, x_18); lean_ctor_set(x_19, 1, x_12); x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_617_(x_7); -x_21 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____rarg___closed__2; +x_21 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____rarg___closed__2; x_22 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_22, 0, x_21); lean_ctor_set(x_22, 1, x_20); x_23 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_23, 0, x_22); lean_ctor_set(x_23, 1, x_12); -x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____rarg___closed__3; -x_25 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____spec__1___rarg(x_1, x_24, x_8); +x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____rarg___closed__3; +x_25 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____spec__1___rarg(x_1, x_24, x_8); x_26 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_26, 0, x_25); lean_ctor_set(x_26, 1, x_12); @@ -8446,7 +8446,7 @@ goto block_38; block_38: { 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; -x_30 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__4; +x_30 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__4; x_31 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_31, 0, x_30); lean_ctor_set(x_31, 1, x_29); @@ -8468,15 +8468,15 @@ return x_37; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867_(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____rarg), 2, 0); +x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____rarg), 2, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____spec__2___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____spec__2___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { size_t x_5; size_t x_6; lean_object* x_7; @@ -8484,7 +8484,7 @@ x_5 = lean_unbox_usize(x_2); lean_dec(x_2); x_6 = lean_unbox_usize(x_3); lean_dec(x_3); -x_7 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____spec__2___rarg(x_1, x_5, x_6, x_4); +x_7 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____spec__2___rarg(x_1, x_5, x_6, x_4); return x_7; } } @@ -8492,7 +8492,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentSymbolAux___rarg(lean_obje _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____rarg), 2, 1); +x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } @@ -8564,7 +8564,7 @@ return x_12; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____at_Lean_Lsp_instToJsonDocumentSymbol_go___spec__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____at_Lean_Lsp_instToJsonDocumentSymbol_go___spec__1(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t 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; 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; 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; @@ -8582,7 +8582,7 @@ lean_inc(x_7); lean_dec(x_1); x_8 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_8, 0, x_2); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____rarg___closed__1; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____rarg___closed__1; x_10 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_10, 0, x_9); lean_ctor_set(x_10, 1, x_8); @@ -8590,11 +8590,11 @@ x_11 = lean_box(0); x_12 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_12, 0, x_10); lean_ctor_set(x_12, 1, x_11); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__2; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__2; x_14 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_2095____spec__1(x_13, x_3); lean_dec(x_3); x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_617_(x_5); -x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1665____closed__2; +x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1664____closed__2; x_17 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_17, 0, x_16); lean_ctor_set(x_17, 1, x_15); @@ -8602,14 +8602,14 @@ x_18 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_18, 0, x_17); lean_ctor_set(x_18, 1, x_11); x_19 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_617_(x_6); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____rarg___closed__2; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____rarg___closed__2; x_21 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_21, 0, x_20); lean_ctor_set(x_21, 1, x_19); x_22 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_22, 0, x_21); lean_ctor_set(x_22, 1, x_11); -x_23 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____rarg___closed__3; +x_23 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____rarg___closed__3; x_24 = l_Lean_Json_opt___at_Lean_Lsp_instToJsonDocumentSymbol_go___spec__2(x_23, x_7); x_25 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_25, 0, x_24); @@ -8807,7 +8807,7 @@ goto block_37; block_37: { lean_object* x_29; 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; -x_29 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__4; +x_29 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__4; x_30 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_30, 0, x_29); lean_ctor_set(x_30, 1, x_28); @@ -8836,7 +8836,7 @@ lean_object* x_2; lean_object* x_3; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); lean_dec(x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____at_Lean_Lsp_instToJsonDocumentSymbol_go___spec__1(x_2); +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____at_Lean_Lsp_instToJsonDocumentSymbol_go___spec__1(x_2); return x_3; } } @@ -8975,7 +8975,7 @@ x_1 = lean_box(0); return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3088____spec__1(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3087____spec__1(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -8999,7 +8999,7 @@ goto _start; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3088____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3087____closed__1() { _start: { lean_object* x_1; @@ -9007,7 +9007,7 @@ x_1 = lean_mk_string_from_bytes("tags", 4); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3088____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3087____closed__2() { _start: { lean_object* x_1; @@ -9015,7 +9015,7 @@ x_1 = lean_mk_string_from_bytes("location", 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3088____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3087____closed__3() { _start: { lean_object* x_1; @@ -9023,7 +9023,7 @@ x_1 = lean_mk_string_from_bytes("containerName", 13); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3088_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3087_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t 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; lean_object* x_11; lean_object* x_12; size_t x_13; size_t 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; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; @@ -9039,7 +9039,7 @@ lean_inc(x_6); lean_dec(x_1); x_7 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_7, 0, x_2); -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____rarg___closed__1; +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____rarg___closed__1; x_9 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_9, 0, x_8); lean_ctor_set(x_9, 1, x_7); @@ -9051,10 +9051,10 @@ x_12 = lean_array_get_size(x_4); x_13 = lean_usize_of_nat(x_12); lean_dec(x_12); x_14 = 0; -x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3088____spec__1(x_13, x_14, x_4); +x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3087____spec__1(x_13, x_14, x_4); x_16 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_16, 0, x_15); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3088____closed__1; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3087____closed__1; x_18 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_18, 0, x_17); lean_ctor_set(x_18, 1, x_16); @@ -9062,14 +9062,14 @@ x_19 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_19, 0, x_18); lean_ctor_set(x_19, 1, x_10); x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_891_(x_5); -x_21 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3088____closed__2; +x_21 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3087____closed__2; x_22 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_22, 0, x_21); lean_ctor_set(x_22, 1, x_20); x_23 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_23, 0, x_22); lean_ctor_set(x_23, 1, x_10); -x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3088____closed__3; +x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3087____closed__3; x_25 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_2095____spec__1(x_24, x_6); lean_dec(x_6); x_26 = lean_alloc_ctor(1, 2, 0); @@ -9268,7 +9268,7 @@ goto block_37; block_37: { 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; -x_30 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__4; +x_30 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__4; x_31 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_31, 0, x_30); lean_ctor_set(x_31, 1, x_29); @@ -9287,7 +9287,7 @@ return x_36; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3088____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3087____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; size_t x_5; lean_object* x_6; @@ -9295,7 +9295,7 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3088____spec__1(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3087____spec__1(x_4, x_5, x_3); return x_6; } } @@ -9303,7 +9303,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonSymbolInformation___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3088_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3087_), 1, 0); return x_1; } } @@ -9514,7 +9514,7 @@ x_3 = l_Lean_Lsp_SemanticTokenType_toNat(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3227____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3226____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -9548,7 +9548,7 @@ return x_15; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3227____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3226____closed__1() { _start: { lean_object* x_1; @@ -9556,7 +9556,7 @@ x_1 = lean_mk_string_from_bytes("tokenTypes", 10); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3227____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3226____closed__2() { _start: { lean_object* x_1; @@ -9564,12 +9564,12 @@ x_1 = lean_mk_string_from_bytes("tokenModifiers", 14); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3227_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3226_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3227____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3227____spec__1(x_1, x_2); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3226____closed__1; +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3226____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -9595,8 +9595,8 @@ lean_object* x_7; lean_object* x_8; lean_object* x_9; x_7 = lean_ctor_get(x_3, 0); lean_inc(x_7); lean_dec(x_3); -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3227____closed__2; -x_9 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3227____spec__1(x_1, x_8); +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3226____closed__2; +x_9 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3226____spec__1(x_1, x_8); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; @@ -9648,21 +9648,21 @@ return x_18; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3227____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3226____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3227____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3226____spec__1(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3227____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3226____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3227_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3226_(x_1); lean_dec(x_1); return x_2; } @@ -9671,7 +9671,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonSemanticTokensLegend___closed__ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3227____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3226____boxed), 1, 0); return x_1; } } @@ -9683,7 +9683,7 @@ x_1 = l_Lean_Lsp_instFromJsonSemanticTokensLegend___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3279_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3278_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; size_t x_4; size_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; size_t 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; @@ -9696,7 +9696,7 @@ x_5 = 0; x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_112____spec__2(x_4, x_5, x_2); x_7 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_7, 0, x_6); -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3227____closed__1; +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3226____closed__1; x_9 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_9, 0, x_8); lean_ctor_set(x_9, 1, x_7); @@ -9713,7 +9713,7 @@ lean_dec(x_13); x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_112____spec__2(x_14, x_5, x_12); x_16 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_16, 0, x_15); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3227____closed__2; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3226____closed__2; x_18 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_18, 0, x_17); lean_ctor_set(x_18, 1, x_16); @@ -9735,7 +9735,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonSemanticTokensLegend___closed__1( _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3279_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3278_), 1, 0); return x_1; } } @@ -9747,17 +9747,17 @@ x_1 = l_Lean_Lsp_instToJsonSemanticTokensLegend___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3346____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3345____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; x_3 = l_Lean_Json_getObjValD(x_1, x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3227_(x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3226_(x_3); lean_dec(x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3346____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3345____closed__1() { _start: { lean_object* x_1; @@ -9765,7 +9765,7 @@ x_1 = lean_mk_string_from_bytes("legend", 6); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3346____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3345____closed__2() { _start: { lean_object* x_1; @@ -9773,12 +9773,12 @@ x_1 = lean_mk_string_from_bytes("full", 4); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3346_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3345_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3346____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3346____spec__1(x_1, x_2); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3345____closed__1; +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3345____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -9804,7 +9804,7 @@ lean_object* x_7; lean_object* x_8; lean_object* x_9; x_7 = lean_ctor_get(x_3, 0); lean_inc(x_7); lean_dec(x_3); -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1665____closed__2; +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1664____closed__2; x_9 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____spec__3(x_1, x_8); if (lean_obj_tag(x_9) == 0) { @@ -9832,7 +9832,7 @@ lean_object* x_13; lean_object* x_14; lean_object* x_15; x_13 = lean_ctor_get(x_9, 0); lean_inc(x_13); lean_dec(x_9); -x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3346____closed__2; +x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3345____closed__2; x_15 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____spec__3(x_1, x_14); if (lean_obj_tag(x_15) == 0) { @@ -9897,21 +9897,21 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3346____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3345____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3346____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3345____spec__1(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3346____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3345____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3346_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3345_(x_1); lean_dec(x_1); return x_2; } @@ -9920,7 +9920,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonSemanticTokensOptions___closed_ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3346____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3345____boxed), 1, 0); return x_1; } } @@ -9932,14 +9932,14 @@ x_1 = l_Lean_Lsp_instFromJsonSemanticTokensOptions___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3415_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3414_(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; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t 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; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3279_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3346____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3278_(x_2); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3345____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -9950,7 +9950,7 @@ lean_ctor_set(x_7, 1, x_6); x_8 = lean_ctor_get_uint8(x_1, sizeof(void*)*1); x_9 = lean_alloc_ctor(1, 0, 1); lean_ctor_set_uint8(x_9, 0, x_8); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1665____closed__2; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1664____closed__2; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -9961,7 +9961,7 @@ x_13 = lean_ctor_get_uint8(x_1, sizeof(void*)*1 + 1); lean_dec(x_1); x_14 = lean_alloc_ctor(1, 0, 1); lean_ctor_set_uint8(x_14, 0, x_13); -x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3346____closed__2; +x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3345____closed__2; x_16 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_16, 0, x_15); lean_ctor_set(x_16, 1, x_14); @@ -9986,7 +9986,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonSemanticTokensOptions___closed__1 _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3415_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3414_), 1, 0); return x_1; } } @@ -9998,11 +9998,11 @@ x_1 = l_Lean_Lsp_instToJsonSemanticTokensOptions___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3477_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3476_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_1974____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -10044,11 +10044,11 @@ return x_9; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3477____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3476____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3477_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3476_(x_1); lean_dec(x_1); return x_2; } @@ -10057,7 +10057,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonSemanticTokensParams___closed__ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3477____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3476____boxed), 1, 0); return x_1; } } @@ -10069,12 +10069,12 @@ x_1 = l_Lean_Lsp_instFromJsonSemanticTokensParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3512_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3511_(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_8; lean_object* x_9; x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_1451_(x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); lean_ctor_set(x_4, 1, x_2); @@ -10094,7 +10094,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonSemanticTokensParams___closed__1( _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3512_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3511_), 1, 0); return x_1; } } @@ -10106,11 +10106,11 @@ x_1 = l_Lean_Lsp_instToJsonSemanticTokensParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3560_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3559_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_1974____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -10137,7 +10137,7 @@ lean_object* x_7; lean_object* x_8; lean_object* x_9; x_7 = lean_ctor_get(x_3, 0); lean_inc(x_7); lean_dec(x_3); -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1665____closed__2; +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1664____closed__2; x_9 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_931____spec__2(x_1, x_8); if (lean_obj_tag(x_9) == 0) { @@ -10190,11 +10190,11 @@ return x_18; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3560____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3559____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3560_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3559_(x_1); lean_dec(x_1); return x_2; } @@ -10203,7 +10203,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonSemanticTokensRangeParams___clo _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3560____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3559____boxed), 1, 0); return x_1; } } @@ -10215,14 +10215,14 @@ x_1 = l_Lean_Lsp_instFromJsonSemanticTokensRangeParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3612_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3611_(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_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; lean_object* x_15; lean_object* x_16; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_1451_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -10234,7 +10234,7 @@ x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_617_(x_8); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1665____closed__2; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1664____closed__2; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -10256,7 +10256,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonSemanticTokensRangeParams___close _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3612_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3611_), 1, 0); return x_1; } } @@ -10268,7 +10268,7 @@ x_1 = l_Lean_Lsp_instToJsonSemanticTokensRangeParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3664____spec__2(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3663____spec__2(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -10324,7 +10324,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3664____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3663____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -10341,7 +10341,7 @@ x_5 = lean_array_get_size(x_4); x_6 = lean_usize_of_nat(x_5); lean_dec(x_5); x_7 = 0; -x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3664____spec__2(x_6, x_7, x_4); +x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3663____spec__2(x_6, x_7, x_4); return x_8; } else @@ -10360,7 +10360,7 @@ return x_15; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3664____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3663____closed__1() { _start: { lean_object* x_1; @@ -10368,12 +10368,12 @@ x_1 = lean_mk_string_from_bytes("data", 4); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3664_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3663_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3664____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3664____spec__1(x_1, x_2); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3663____closed__1; +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3663____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -10414,7 +10414,7 @@ return x_9; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3664____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3663____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; size_t x_5; lean_object* x_6; @@ -10422,7 +10422,7 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3664____spec__2(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3663____spec__2(x_4, x_5, x_3); return x_6; } } @@ -10430,7 +10430,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonSemanticTokens___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3664_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3663_), 1, 0); return x_1; } } @@ -10442,7 +10442,7 @@ x_1 = l_Lean_Lsp_instFromJsonSemanticTokens___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____spec__1(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3698____spec__1(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -10469,7 +10469,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3698_(lean_object* x_1) { _start: { 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_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; @@ -10477,10 +10477,10 @@ x_2 = lean_array_get_size(x_1); x_3 = lean_usize_of_nat(x_2); lean_dec(x_2); x_4 = 0; -x_5 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____spec__1(x_3, x_4, x_1); +x_5 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3698____spec__1(x_3, x_4, x_1); x_6 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_6, 0, x_5); -x_7 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3664____closed__1; +x_7 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3663____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); @@ -10496,7 +10496,7 @@ x_13 = l_Lean_Json_mkObj(x_12); return x_13; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3698____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; size_t x_5; lean_object* x_6; @@ -10504,7 +10504,7 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____spec__1(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3698____spec__1(x_4, x_5, x_3); return x_6; } } @@ -10512,7 +10512,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonSemanticTokens___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3698_), 1, 0); return x_1; } } @@ -10524,11 +10524,11 @@ x_1 = l_Lean_Lsp_instToJsonSemanticTokens___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3739_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_1974____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -10570,11 +10570,11 @@ return x_9; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3739____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3739_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738_(x_1); lean_dec(x_1); return x_2; } @@ -10583,7 +10583,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonFoldingRangeParams___closed__1( _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3739____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____boxed), 1, 0); return x_1; } } @@ -10595,12 +10595,12 @@ x_1 = l_Lean_Lsp_instFromJsonFoldingRangeParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3774_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3773_(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_8; lean_object* x_9; x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_1451_(x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); lean_ctor_set(x_4, 1, x_2); @@ -10620,7 +10620,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonFoldingRangeParams___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3774_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3773_), 1, 0); return x_1; } } @@ -10792,7 +10792,7 @@ x_1 = lean_box(0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3857____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3856____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -10849,7 +10849,7 @@ return x_15; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3857____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3856____closed__1() { _start: { lean_object* x_1; @@ -10857,7 +10857,7 @@ x_1 = lean_mk_string_from_bytes("startLine", 9); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3857____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3856____closed__2() { _start: { lean_object* x_1; @@ -10865,7 +10865,7 @@ x_1 = lean_mk_string_from_bytes("endLine", 7); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3857_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3856_(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_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; 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; @@ -10874,7 +10874,7 @@ lean_inc(x_2); x_3 = l_Lean_JsonNumber_fromNat(x_2); x_4 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_4, 0, x_3); -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3857____closed__1; +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3856____closed__1; x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_5); lean_ctor_set(x_6, 1, x_4); @@ -10887,7 +10887,7 @@ lean_inc(x_9); x_10 = l_Lean_JsonNumber_fromNat(x_9); x_11 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_11, 0, x_10); -x_12 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3857____closed__2; +x_12 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3856____closed__2; x_13 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_13, 0, x_12); lean_ctor_set(x_13, 1, x_11); @@ -10897,8 +10897,8 @@ lean_ctor_set(x_14, 1, x_7); x_15 = lean_ctor_get(x_1, 2); lean_inc(x_15); lean_dec(x_1); -x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__4; -x_17 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3857____spec__1(x_16, x_15); +x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__4; +x_17 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3856____spec__1(x_16, x_15); lean_dec(x_15); x_18 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_18, 0, x_17); @@ -10914,11 +10914,11 @@ x_22 = l_Lean_Json_mkObj(x_21); return x_22; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3857____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3856____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3857____spec__1(x_1, x_2); +x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3856____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -10927,7 +10927,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonFoldingRange___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3857_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3856_), 1, 0); return x_1; } } @@ -10984,320 +10984,320 @@ lean_mark_persistent(l_Lean_Lsp_instToJsonCompletionOptions); l_Lean_Lsp_CompletionItemKind_noConfusion___rarg___closed__1 = _init_l_Lean_Lsp_CompletionItemKind_noConfusion___rarg___closed__1(); lean_mark_persistent(l_Lean_Lsp_CompletionItemKind_noConfusion___rarg___closed__1); l_Lean_Lsp_instInhabitedCompletionItemKind = _init_l_Lean_Lsp_instInhabitedCompletionItemKind(); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__11); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__12); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__13); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__14); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__15); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__16); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__17 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__17(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__17); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__18 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__18(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__18); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__19 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__19(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__19); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__20 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__20(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__20); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__21 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__21(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__21); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__22 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__22(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__22); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__23 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__23(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__23); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__24 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__24(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__24); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__25 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__25(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__25); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__26 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__26(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__26); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__27 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__27(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__27); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__28 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__28(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__28); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__29 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__29(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__29); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__30 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__30(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__30); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__31 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__31(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__31); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__32 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__32(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__32); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__33 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__33(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__33); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__34 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__34(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__34); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__35 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__35(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__35); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__36 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__36(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__36); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__37 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__37(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__37); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__38 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__38(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__38); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__39 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__39(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__39); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__40 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__40(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__40); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__41 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__41(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__41); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__42 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__42(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__42); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__43 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__43(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__43); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__44 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__44(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__44); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__45 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__45(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__45); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__46 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__46(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__46); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__47 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__47(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__47); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__48 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__48(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__48); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__49 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__49(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__49); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__50 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__50(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__50); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__51 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__51(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__51); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__52 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__52(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__52); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__53 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__53(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__53); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__54 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__54(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__54); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__55 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__55(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__55); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__56 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__56(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__56); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__57 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__57(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__57); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__58 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__58(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__58); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__59 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__59(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__59); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__60 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__60(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__60); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__61 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__61(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__61); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__62 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__62(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__62); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__63 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__63(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__63); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__64 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__64(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__64); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__65 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__65(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__65); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__66 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__66(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__66); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__67 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__67(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__67); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__68 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__68(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__68); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__69 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__69(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__69); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__70 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__70(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__70); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__71 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__71(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__71); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__72 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__72(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__72); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__73 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__73(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__73); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__74 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__74(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__74); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__75 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__75(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__75); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__76 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__76(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__76); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__77 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__77(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__77); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__78 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__78(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__78); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__79 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__79(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__79); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__80 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__80(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__80); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__81 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__81(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__81); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__82 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__82(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__82); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__83 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__83(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__83); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__84 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__84(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__84); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__85 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__85(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__85); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__86 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__86(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__86); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__87 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__87(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__87); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__88 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__88(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__88); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__89 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__89(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__89); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__90 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__90(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__90); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__91 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__91(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__91); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__92 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__92(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__92); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__93 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__93(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__93); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__94 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__94(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__94); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__95 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__95(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__95); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__96 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__96(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__96); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__97 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__97(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__97); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__98 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__98(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__98); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__99 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__99(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__99); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__100 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__100(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__100); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__101 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__101(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__101); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__102 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__102(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__102); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__103 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__103(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__103); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__104 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__104(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__104); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__105 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__105(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__105); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__106 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__106(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__106); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__107 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__107(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__107); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__108 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__108(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__108); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__109 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__109(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__109); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__110 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__110(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__110); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__111 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__111(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__111); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__112 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__112(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__112); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__113 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__113(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__113); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__114 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__114(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__114); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__115 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__115(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__115); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__116 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__116(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__116); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__117 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__117(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__117); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__118 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__118(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__118); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__119 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__119(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__119); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__120 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__120(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__120); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__121 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__121(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__121); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__122 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__122(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__122); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__123 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__123(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__123); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__124 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__124(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__124); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__125 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__125(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__125); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__126 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__126(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__126); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__127 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__127(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__127); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__128 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__128(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__128); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__129 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__129(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__129); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__130 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__130(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__130); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__131 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__131(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__131); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__132 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__132(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__132); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__133 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__133(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__133); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__134 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__134(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__134); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__135 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__135(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__135); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__136 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__136(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__136); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__137 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__137(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__137); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__138 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__138(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__138); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__139 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__139(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__139); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__140 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__140(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__140); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__141 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__141(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__141); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__142 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__142(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__142); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__143 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__143(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__143); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__144 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__144(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__144); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__145 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__145(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__145); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__146 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__146(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__146); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__147 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__147(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__147); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__148 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__148(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__148); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__149 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__149(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__149); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__150 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__150(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__150); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__151 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__151(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__151); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__152 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__152(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_255____closed__152); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__12); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__13); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__14); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__15); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__16); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__17 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__17(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__17); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__18 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__18(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__18); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__19 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__19(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__19); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__20 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__20(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__20); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__21 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__21(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__21); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__22 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__22(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__22); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__23 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__23(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__23); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__24 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__24(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__24); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__25 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__25(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__25); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__26 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__26(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__26); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__27 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__27(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__27); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__28 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__28(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__28); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__29 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__29(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__29); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__30 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__30(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__30); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__31 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__31(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__31); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__32 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__32(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__32); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__33 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__33(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__33); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__34 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__34(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__34); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__35 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__35(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__35); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__36 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__36(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__36); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__37 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__37(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__37); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__38 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__38(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__38); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__39 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__39(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__39); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__40 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__40(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__40); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__41 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__41(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__41); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__42 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__42(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__42); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__43 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__43(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__43); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__44 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__44(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__44); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__45 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__45(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__45); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__46 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__46(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__46); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__47 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__47(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__47); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__48 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__48(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__48); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__49 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__49(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__49); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__50 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__50(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__50); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__51 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__51(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__51); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__52 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__52(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__52); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__53 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__53(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__53); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__54 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__54(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__54); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__55 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__55(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__55); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__56 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__56(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__56); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__57 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__57(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__57); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__58 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__58(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__58); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__59 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__59(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__59); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__60 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__60(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__60); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__61 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__61(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__61); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__62 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__62(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__62); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__63 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__63(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__63); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__64 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__64(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__64); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__65 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__65(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__65); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__66 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__66(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__66); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__67 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__67(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__67); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__68 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__68(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__68); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__69 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__69(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__69); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__70 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__70(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__70); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__71 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__71(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__71); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__72 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__72(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__72); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__73 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__73(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__73); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__74 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__74(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__74); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__75 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__75(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__75); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__76 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__76(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__76); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__77 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__77(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__77); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__78 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__78(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__78); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__79 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__79(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__79); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__80 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__80(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__80); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__81 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__81(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__81); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__82 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__82(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__82); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__83 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__83(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__83); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__84 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__84(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__84); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__85 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__85(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__85); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__86 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__86(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__86); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__87 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__87(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__87); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__88 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__88(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__88); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__89 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__89(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__89); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__90 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__90(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__90); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__91 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__91(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__91); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__92 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__92(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__92); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__93 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__93(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__93); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__94 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__94(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__94); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__95 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__95(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__95); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__96 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__96(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__96); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__97 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__97(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__97); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__98 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__98(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__98); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__99 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__99(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__99); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__100 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__100(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__100); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__101 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__101(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__101); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__102 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__102(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__102); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__103 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__103(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__103); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__104 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__104(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__104); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__105 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__105(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__105); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__106 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__106(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__106); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__107 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__107(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__107); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__108 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__108(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__108); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__109 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__109(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__109); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__110 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__110(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__110); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__111 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__111(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__111); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__112 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__112(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__112); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__113 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__113(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__113); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__114 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__114(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__114); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__115 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__115(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__115); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__116 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__116(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__116); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__117 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__117(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__117); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__118 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__118(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__118); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__119 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__119(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__119); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__120 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__120(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__120); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__121 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__121(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__121); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__122 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__122(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__122); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__123 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__123(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__123); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__124 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__124(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__124); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__125 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__125(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__125); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__126 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__126(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__126); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__127 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__127(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__127); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__128 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__128(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__128); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__129 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__129(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__129); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__130 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__130(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__130); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__131 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__131(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__131); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__132 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__132(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__132); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__133 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__133(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__133); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__134 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__134(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__134); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__135 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__135(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__135); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__136 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__136(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__136); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__137 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__137(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__137); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__138 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__138(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__138); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__139 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__139(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__139); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__140 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__140(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__140); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__141 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__141(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__141); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__142 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__142(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__142); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__143 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__143(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__143); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__144 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__144(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__144); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__145 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__145(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__145); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__146 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__146(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__146); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__147 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__147(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__147); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__148 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__148(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__148); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__149 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__149(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__149); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__150 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__150(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__150); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__151 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__151(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__151); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__152 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__152(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_254____closed__152); l_Lean_Lsp_instReprCompletionItemKind___closed__1 = _init_l_Lean_Lsp_instReprCompletionItemKind___closed__1(); lean_mark_persistent(l_Lean_Lsp_instReprCompletionItemKind___closed__1); l_Lean_Lsp_instReprCompletionItemKind = _init_l_Lean_Lsp_instReprCompletionItemKind(); lean_mark_persistent(l_Lean_Lsp_instReprCompletionItemKind); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1100____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1100____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1100____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1100____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1100____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1100____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1100____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1100____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1100____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1099____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1099____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1099____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1099____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1099____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1099____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1099____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1099____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1099____closed__3); l_Lean_Lsp_instFromJsonInsertReplaceEdit___closed__1 = _init_l_Lean_Lsp_instFromJsonInsertReplaceEdit___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonInsertReplaceEdit___closed__1); l_Lean_Lsp_instFromJsonInsertReplaceEdit = _init_l_Lean_Lsp_instFromJsonInsertReplaceEdit(); @@ -11308,16 +11308,16 @@ l_Lean_Lsp_instToJsonInsertReplaceEdit = _init_l_Lean_Lsp_instToJsonInsertReplac lean_mark_persistent(l_Lean_Lsp_instToJsonInsertReplaceEdit); l_Lean_Lsp_CompletionItem_textEdit_x3f___default = _init_l_Lean_Lsp_CompletionItem_textEdit_x3f___default(); lean_mark_persistent(l_Lean_Lsp_CompletionItem_textEdit_x3f___default); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1268____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1267____closed__5); l_Lean_Lsp_instFromJsonCompletionItem___closed__1 = _init_l_Lean_Lsp_instFromJsonCompletionItem___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonCompletionItem___closed__1); l_Lean_Lsp_instFromJsonCompletionItem = _init_l_Lean_Lsp_instFromJsonCompletionItem(); @@ -11332,10 +11332,10 @@ l_Lean_Lsp_instInhabitedCompletionItem___closed__2 = _init_l_Lean_Lsp_instInhabi lean_mark_persistent(l_Lean_Lsp_instInhabitedCompletionItem___closed__2); l_Lean_Lsp_instInhabitedCompletionItem = _init_l_Lean_Lsp_instInhabitedCompletionItem(); lean_mark_persistent(l_Lean_Lsp_instInhabitedCompletionItem); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1449____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1449____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1449____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1449____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1449____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1449____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1448____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1448____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1448____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1448____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1448____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1448____closed__2); l_Lean_Lsp_instFromJsonCompletionList___closed__1 = _init_l_Lean_Lsp_instFromJsonCompletionList___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonCompletionList___closed__1); l_Lean_Lsp_instFromJsonCompletionList = _init_l_Lean_Lsp_instFromJsonCompletionList(); @@ -11344,10 +11344,10 @@ l_Lean_Lsp_instToJsonCompletionList___closed__1 = _init_l_Lean_Lsp_instToJsonCom lean_mark_persistent(l_Lean_Lsp_instToJsonCompletionList___closed__1); l_Lean_Lsp_instToJsonCompletionList = _init_l_Lean_Lsp_instToJsonCompletionList(); lean_mark_persistent(l_Lean_Lsp_instToJsonCompletionList); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551____closed__2); l_Lean_Lsp_instFromJsonCompletionParams___closed__1 = _init_l_Lean_Lsp_instFromJsonCompletionParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonCompletionParams___closed__1); l_Lean_Lsp_instFromJsonCompletionParams = _init_l_Lean_Lsp_instFromJsonCompletionParams(); @@ -11358,10 +11358,10 @@ l_Lean_Lsp_instToJsonCompletionParams = _init_l_Lean_Lsp_instToJsonCompletionPar lean_mark_persistent(l_Lean_Lsp_instToJsonCompletionParams); l_Lean_Lsp_Hover_range_x3f___default = _init_l_Lean_Lsp_Hover_range_x3f___default(); lean_mark_persistent(l_Lean_Lsp_Hover_range_x3f___default); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1665____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1665____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1665____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1665____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1665____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1665____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1664____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1664____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1664____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1664____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1664____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1664____closed__2); l_Lean_Lsp_instToJsonHover___closed__1 = _init_l_Lean_Lsp_instToJsonHover___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonHover___closed__1); l_Lean_Lsp_instToJsonHover = _init_l_Lean_Lsp_instToJsonHover(); @@ -11402,8 +11402,8 @@ l_Lean_Lsp_instToJsonTypeDefinitionParams___closed__1 = _init_l_Lean_Lsp_instToJ lean_mark_persistent(l_Lean_Lsp_instToJsonTypeDefinitionParams___closed__1); l_Lean_Lsp_instToJsonTypeDefinitionParams = _init_l_Lean_Lsp_instToJsonTypeDefinitionParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonTypeDefinitionParams); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2173____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2173____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2173____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2172____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2172____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2172____closed__1); l_Lean_Lsp_instFromJsonReferenceContext___closed__1 = _init_l_Lean_Lsp_instFromJsonReferenceContext___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonReferenceContext___closed__1); l_Lean_Lsp_instFromJsonReferenceContext = _init_l_Lean_Lsp_instFromJsonReferenceContext(); @@ -11412,8 +11412,8 @@ l_Lean_Lsp_instToJsonReferenceContext___closed__1 = _init_l_Lean_Lsp_instToJsonR lean_mark_persistent(l_Lean_Lsp_instToJsonReferenceContext___closed__1); l_Lean_Lsp_instToJsonReferenceContext = _init_l_Lean_Lsp_instToJsonReferenceContext(); lean_mark_persistent(l_Lean_Lsp_instToJsonReferenceContext); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2256____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2256____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2256____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2255____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2255____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2255____closed__1); l_Lean_Lsp_instFromJsonReferenceParams___closed__1 = _init_l_Lean_Lsp_instFromJsonReferenceParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonReferenceParams___closed__1); l_Lean_Lsp_instFromJsonReferenceParams = _init_l_Lean_Lsp_instFromJsonReferenceParams(); @@ -11422,8 +11422,8 @@ l_Lean_Lsp_instToJsonReferenceParams___closed__1 = _init_l_Lean_Lsp_instToJsonRe lean_mark_persistent(l_Lean_Lsp_instToJsonReferenceParams___closed__1); l_Lean_Lsp_instToJsonReferenceParams = _init_l_Lean_Lsp_instToJsonReferenceParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonReferenceParams); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2387____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2387____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2387____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2386____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2386____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2386____closed__1); l_Lean_Lsp_instFromJsonWorkspaceSymbolParams___closed__1 = _init_l_Lean_Lsp_instFromJsonWorkspaceSymbolParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonWorkspaceSymbolParams___closed__1); l_Lean_Lsp_instFromJsonWorkspaceSymbolParams = _init_l_Lean_Lsp_instFromJsonWorkspaceSymbolParams(); @@ -11560,12 +11560,12 @@ l_Lean_Lsp_instToJsonSymbolKind___closed__46 = _init_l_Lean_Lsp_instToJsonSymbol lean_mark_persistent(l_Lean_Lsp_instToJsonSymbolKind___closed__46); l_Lean_Lsp_DocumentSymbolAux_detail_x3f___default = _init_l_Lean_Lsp_DocumentSymbolAux_detail_x3f___default(); lean_mark_persistent(l_Lean_Lsp_DocumentSymbolAux_detail_x3f___default); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____rarg___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____rarg___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____rarg___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____rarg___closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____rarg___closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____rarg___closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____rarg___closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____rarg___closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2868____rarg___closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____rarg___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____rarg___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____rarg___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____rarg___closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____rarg___closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____rarg___closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____rarg___closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____rarg___closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2867____rarg___closed__3); l_Lean_Lsp_instToJsonDocumentSymbol___closed__1 = _init_l_Lean_Lsp_instToJsonDocumentSymbol___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonDocumentSymbol___closed__1); l_Lean_Lsp_instToJsonDocumentSymbol = _init_l_Lean_Lsp_instToJsonDocumentSymbol(); @@ -11576,12 +11576,12 @@ l_Lean_Lsp_SymbolInformation_tags___default = _init_l_Lean_Lsp_SymbolInformation lean_mark_persistent(l_Lean_Lsp_SymbolInformation_tags___default); l_Lean_Lsp_SymbolInformation_containerName_x3f___default = _init_l_Lean_Lsp_SymbolInformation_containerName_x3f___default(); lean_mark_persistent(l_Lean_Lsp_SymbolInformation_containerName_x3f___default); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3088____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3088____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3088____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3088____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3088____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3088____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3088____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3088____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3088____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3087____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3087____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3087____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3087____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3087____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3087____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3087____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3087____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3087____closed__3); l_Lean_Lsp_instToJsonSymbolInformation___closed__1 = _init_l_Lean_Lsp_instToJsonSymbolInformation___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonSymbolInformation___closed__1); l_Lean_Lsp_instToJsonSymbolInformation = _init_l_Lean_Lsp_instToJsonSymbolInformation(); @@ -11606,10 +11606,10 @@ l_Lean_Lsp_SemanticTokenType_names___closed__9 = _init_l_Lean_Lsp_SemanticTokenT lean_mark_persistent(l_Lean_Lsp_SemanticTokenType_names___closed__9); l_Lean_Lsp_SemanticTokenType_names = _init_l_Lean_Lsp_SemanticTokenType_names(); lean_mark_persistent(l_Lean_Lsp_SemanticTokenType_names); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3227____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3227____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3227____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3227____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3227____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3227____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3226____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3226____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3226____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3226____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3226____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3226____closed__2); l_Lean_Lsp_instFromJsonSemanticTokensLegend___closed__1 = _init_l_Lean_Lsp_instFromJsonSemanticTokensLegend___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonSemanticTokensLegend___closed__1); l_Lean_Lsp_instFromJsonSemanticTokensLegend = _init_l_Lean_Lsp_instFromJsonSemanticTokensLegend(); @@ -11618,10 +11618,10 @@ l_Lean_Lsp_instToJsonSemanticTokensLegend___closed__1 = _init_l_Lean_Lsp_instToJ lean_mark_persistent(l_Lean_Lsp_instToJsonSemanticTokensLegend___closed__1); l_Lean_Lsp_instToJsonSemanticTokensLegend = _init_l_Lean_Lsp_instToJsonSemanticTokensLegend(); lean_mark_persistent(l_Lean_Lsp_instToJsonSemanticTokensLegend); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3346____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3346____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3346____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3346____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3346____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3346____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3345____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3345____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3345____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3345____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3345____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3345____closed__2); l_Lean_Lsp_instFromJsonSemanticTokensOptions___closed__1 = _init_l_Lean_Lsp_instFromJsonSemanticTokensOptions___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonSemanticTokensOptions___closed__1); l_Lean_Lsp_instFromJsonSemanticTokensOptions = _init_l_Lean_Lsp_instFromJsonSemanticTokensOptions(); @@ -11646,8 +11646,8 @@ l_Lean_Lsp_instToJsonSemanticTokensRangeParams___closed__1 = _init_l_Lean_Lsp_in lean_mark_persistent(l_Lean_Lsp_instToJsonSemanticTokensRangeParams___closed__1); l_Lean_Lsp_instToJsonSemanticTokensRangeParams = _init_l_Lean_Lsp_instToJsonSemanticTokensRangeParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonSemanticTokensRangeParams); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3664____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3664____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3664____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3663____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3663____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3663____closed__1); l_Lean_Lsp_instFromJsonSemanticTokens___closed__1 = _init_l_Lean_Lsp_instFromJsonSemanticTokens___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonSemanticTokens___closed__1); l_Lean_Lsp_instFromJsonSemanticTokens = _init_l_Lean_Lsp_instFromJsonSemanticTokens(); @@ -11678,10 +11678,10 @@ l_Lean_Lsp_instToJsonFoldingRangeKind___closed__6 = _init_l_Lean_Lsp_instToJsonF lean_mark_persistent(l_Lean_Lsp_instToJsonFoldingRangeKind___closed__6); l_Lean_Lsp_FoldingRange_kind_x3f___default = _init_l_Lean_Lsp_FoldingRange_kind_x3f___default(); lean_mark_persistent(l_Lean_Lsp_FoldingRange_kind_x3f___default); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3857____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3857____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3857____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3857____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3857____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3857____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3856____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3856____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3856____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3856____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3856____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3856____closed__2); l_Lean_Lsp_instToJsonFoldingRange___closed__1 = _init_l_Lean_Lsp_instToJsonFoldingRange___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonFoldingRange___closed__1); l_Lean_Lsp_instToJsonFoldingRange = _init_l_Lean_Lsp_instToJsonFoldingRange(); diff --git a/stage0/stdlib/Lean/Data/Name.c b/stage0/stdlib/Lean/Data/Name.c index c97772ab507..59a5b3019c5 100644 --- a/stage0/stdlib/Lean/Data/Name.c +++ b/stage0/stdlib/Lean/Data/Name.c @@ -19,7 +19,6 @@ lean_object* l_Lean_Name_mkSimple(lean_object*); size_t lean_usize_add(size_t, size_t); static lean_object* l_Lean_Name_getString_x21___closed__3; LEAN_EXPORT lean_object* l_Lean_NameSSet_empty; -extern lean_object* l_String_instInhabitedString; uint8_t lean_uint64_dec_eq(uint64_t, uint64_t); LEAN_EXPORT lean_object* l_Std_RBNode_insert___at_Lean_NameSet_insert___spec__1(lean_object*, lean_object*, lean_object*); lean_object* l___private_Std_Data_HashMap_0__Std_numBucketsForCapacity(lean_object*); @@ -100,6 +99,7 @@ LEAN_EXPORT uint8_t l_Std_HashMapImp_contains___at_Lean_NameSSet_contains___spec LEAN_EXPORT lean_object* l_Lean_Name_hashEx___boxed(lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint64_t lean_name_hash_exported(lean_object*); +lean_object* l_panic___at_Lean_TSyntax_getString___spec__1(lean_object*); LEAN_EXPORT uint8_t l_Std_AssocList_contains___at_Lean_NameSSet_insert___spec__7(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_RBNode_find___at_Lean_NameSet_contains___spec__1___boxed(lean_object*, lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); @@ -172,7 +172,6 @@ uint8_t lean_nat_dec_eq(lean_object*, lean_object*); lean_object* lean_name_mk_numeral(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_RBNode_ins___at_Lean_NameMap_insert___spec__3___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_RBNode_ins___at_Lean_NameMap_insert___spec__3(lean_object*); -lean_object* lean_panic_fn(lean_object*, lean_object*); lean_object* l_String_trim(lean_object*); static lean_object* l_Lean_Name_getString_x21___closed__2; LEAN_EXPORT lean_object* l_Lean_Name_quickCmp___boxed(lean_object*, lean_object*); @@ -180,6 +179,7 @@ LEAN_EXPORT lean_object* l_Lean_NameSet_instForInNameSetName(lean_object*); static lean_object* l_Lean_SMap_empty___at_Lean_NameSSet_empty___spec__1___closed__1; static lean_object* l_Lean_Name_getString_x21___closed__1; lean_object* lean_mk_array(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Name_getString_x21___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_NameMap_instInhabitedNameMap(lean_object*); LEAN_EXPORT lean_object* l_Lean_Name_isNum___boxed(lean_object*); LEAN_EXPORT lean_object* l_Std_mkHashSet___at_Lean_NameHashSet_empty___spec__1(lean_object*); @@ -192,7 +192,6 @@ LEAN_EXPORT lean_object* l_Std_RBNode_ins___at_Lean_NameSet_insert___spec__2(lea static lean_object* l_Lean_Name_getString_x21___closed__4; LEAN_EXPORT lean_object* l_Lean_NameSSet_instEmptyCollectionNameSSet; LEAN_EXPORT lean_object* l_Lean_mkNameMap(lean_object*); -LEAN_EXPORT lean_object* l_panic___at_Lean_Name_getString_x21___spec__1(lean_object*); lean_object* lean_usize_to_nat(size_t); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_NameSSet_insert___spec__5(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SMap_insert___at_Lean_NameSSet_insert___spec__1(lean_object*, lean_object*, lean_object*); @@ -269,15 +268,6 @@ lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_panic___at_Lean_Name_getString_x21___spec__1(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; -x_2 = l_String_instInhabitedString; -x_3 = lean_panic_fn(x_2, x_1); -return x_3; -} -} static lean_object* _init_l_Lean_Name_getString_x21___closed__1() { _start: { @@ -323,19 +313,26 @@ if (lean_obj_tag(x_1) == 1) lean_object* x_2; x_2 = lean_ctor_get(x_1, 1); lean_inc(x_2); -lean_dec(x_1); return x_2; } else { lean_object* x_3; lean_object* x_4; -lean_dec(x_1); x_3 = l_Lean_Name_getString_x21___closed__4; -x_4 = l_panic___at_Lean_Name_getString_x21___spec__1(x_3); +x_4 = l_panic___at_Lean_TSyntax_getString___spec__1(x_3); return x_4; } } } +LEAN_EXPORT lean_object* l_Lean_Name_getString_x21___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Name_getString_x21(x_1); +lean_dec(x_1); +return x_2; +} +} LEAN_EXPORT lean_object* l_Lean_Name_getNumParts(lean_object* x_1) { _start: { diff --git a/stage0/stdlib/Lean/Data/Options.c b/stage0/stdlib/Lean/Data/Options.c index 2c19439ac4b..0d46294427c 100644 --- a/stage0/stdlib/Lean/Data/Options.c +++ b/stage0/stdlib/Lean/Data/Options.c @@ -13,222 +13,221 @@ #ifdef __cplusplus extern "C" { #endif +static lean_object* l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__9; LEAN_EXPORT lean_object* l_Lean_getBoolOption___rarg(lean_object*, lean_object*, lean_object*, uint8_t); lean_object* l_List_reverse___rarg(lean_object*); static lean_object* l_Lean_setOptionFromString___closed__12; static lean_object* l_Lean_setOptionFromString___closed__11; lean_object* l_Lean_initializing(lean_object*); -static lean_object* l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__26; lean_object* l_Lean_KVMap_setBool(lean_object*, lean_object*, uint8_t); -static lean_object* l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__12; -static lean_object* l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__6; +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__29; +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__25; +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__22; LEAN_EXPORT lean_object* l_Lean_Option_register(lean_object*); -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__42; +static lean_object* l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__17; +static lean_object* l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__8; lean_object* lean_mk_empty_array_with_capacity(lean_object*); -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__37; -LEAN_EXPORT lean_object* l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d__; +static lean_object* l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__4; lean_object* lean_name_mk_string(lean_object*, lean_object*); -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__15; +LEAN_EXPORT lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option___x3a___x3a_x3d____1___closed__1; LEAN_EXPORT lean_object* l_Lean_OptionDecl_group___default; -static lean_object* l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__11; -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__3; LEAN_EXPORT lean_object* l_Std_RBNode_find___at_Lean_getOptionDecl___spec__1___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__28; -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__5; +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__6; LEAN_EXPORT lean_object* l_Lean_instMonadWithOptions___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__30; LEAN_EXPORT lean_object* l_Lean_registerOption___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__14; -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__18; lean_object* l_Lean_KVMap_setNat(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__3; -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__1; +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__19; LEAN_EXPORT lean_object* l_Lean_Option_Decl_group___default; -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option_____x3a___x3a_x3d____1___closed__1; lean_object* lean_st_ref_get(lean_object*, lean_object*); -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__2; lean_object* l_Lean_KVMap_setString(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__18; +static lean_object* l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__1; +LEAN_EXPORT lean_object* l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d__; uint8_t lean_name_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withInPattern___rarg___lambda__1(lean_object*); +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__35; LEAN_EXPORT lean_object* l_Lean_withInPattern___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__19; static lean_object* l_Lean_getOptionDeclsArray___closed__1; -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__44; -static lean_object* l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__15; LEAN_EXPORT lean_object* l_Lean_Option_get___rarg___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withInPattern(lean_object*, lean_object*); +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__39; LEAN_EXPORT lean_object* l_Lean_getOptionDefaulValue(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getOptionDecls(lean_object*); LEAN_EXPORT lean_object* l_Lean_Option_get_x3f(lean_object*); static lean_object* l_Lean_instInhabitedOptionDecl___closed__1; lean_object* lean_array_push(lean_object*, lean_object*); lean_object* l_String_toInt_x3f(lean_object*); -static lean_object* l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__14; lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instMonadWithOptions(lean_object*, lean_object*); static lean_object* l_Lean_instForInOptionsProdNameDataValue___closed__1; -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__16; -static lean_object* l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__7; +static lean_object* l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__1; static lean_object* l_Lean_instInhabitedOptionDecl___closed__2; +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__33; lean_object* l_Lean_KVMap_instToStringKVMap(lean_object*); +static lean_object* l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__9; lean_object* l_String_splitOn(lean_object*, lean_object*); lean_object* lean_string_utf8_byte_size(lean_object*); -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__7; -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__24; +static lean_object* l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__2; static lean_object* l_Lean_getOptionDecl___closed__1; uint8_t l_Lean_NameMap_contains___rarg(lean_object*, lean_object*); -extern lean_object* l_Lean_nameLitKind; +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__11; static lean_object* l_Lean_setOptionFromString___closed__9; -static lean_object* l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__9; LEAN_EXPORT lean_object* l_Lean_instInhabitedOptions; static lean_object* l_Lean_setOptionFromString___closed__10; -static lean_object* l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__16; +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__16; +static lean_object* l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__14; +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__32; +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__10; LEAN_EXPORT lean_object* l_Lean_getNatOption(lean_object*); +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__23; LEAN_EXPORT lean_object* l_Lean_getBoolOption___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Option_get___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__24; -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__34; -static lean_object* l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__2; LEAN_EXPORT lean_object* l_Lean_Option_setIfNotSet___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_registerOption___closed__1; static lean_object* l_Lean_registerOption___closed__2; +static lean_object* l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__25; LEAN_EXPORT lean_object* l_Lean_Option_setIfNotSet(lean_object*); -static lean_object* l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__5; lean_object* l_Lean_KVMap_instForInKVMapProdNameDataValue(lean_object*, lean_object*); +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__24; LEAN_EXPORT lean_object* l___private_Lean_Data_Options_0__Lean_optionDeclsRef; -static lean_object* l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__10; -static lean_object* l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__8; -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__23; -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__41; -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__13; +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__36; lean_object* l_String_toName(lean_object*); static lean_object* l_Lean_setOptionFromString___closed__4; -static lean_object* l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__20; LEAN_EXPORT lean_object* l_Lean_instMonadOptions(lean_object*, lean_object*); lean_object* l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(lean_object*, lean_object*); +static lean_object* l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__7; LEAN_EXPORT lean_object* l_Lean_getOptionDecl(lean_object*, lean_object*); -static lean_object* l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__4; +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__7; static lean_object* l_Lean_withInPattern___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_OptionDecl_descr___default; -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__38; LEAN_EXPORT lean_object* l_Lean_instToStringOptions; -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__21; uint8_t l_Lean_KVMap_getBool(lean_object*, lean_object*, uint8_t); +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__13; LEAN_EXPORT lean_object* l_Lean_Options_getInPattern___boxed(lean_object*); -static lean_object* l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__23; -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__31; -static lean_object* l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__21; +static lean_object* l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__21; +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__41; +static lean_object* l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__6; LEAN_EXPORT lean_object* l_Lean_instMonadOptions___rarg(lean_object*, lean_object*); lean_object* l_Lean_Name_toString(lean_object*, uint8_t); -static lean_object* l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__13; +static lean_object* l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__13; LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Data_Options___hyg_121_(lean_object*); static lean_object* l_Lean_setOptionFromString___closed__8; +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__34; uint8_t l_Lean_Name_quickCmp(lean_object*, lean_object*); lean_object* l_Std_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__3; -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__25; -static lean_object* l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__6; +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__20; lean_object* lean_st_mk_ref(lean_object*, lean_object*); -static lean_object* l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__1; lean_object* l_Lean_Syntax_getId(lean_object*); -lean_object* l___private_Init_Meta_0__Lean_quoteNameMk(lean_object*); LEAN_EXPORT lean_object* lean_get_option_decls_array(lean_object*); -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__29; -LEAN_EXPORT lean_object* l_Lean_Option_commandRegister__option_____x3a___x3a_x3d__; -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__6; -static lean_object* l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__5; LEAN_EXPORT lean_object* l_Lean_getNatOption___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_KVMap_setInt(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__21; +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__38; LEAN_EXPORT lean_object* l_Lean_instForInOptionsProdNameDataValue(lean_object*); uint8_t l_Lean_KVMap_contains(lean_object*, lean_object*); -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__36; +static lean_object* l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__23; +static lean_object* l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__22; LEAN_EXPORT lean_object* l_Lean_Options_empty; static lean_object* l_Lean_OptionDecl_group___default___closed__1; -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__39; -static lean_object* l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__25; -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__35; +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__31; static lean_object* l_Lean_setOptionFromString___closed__5; LEAN_EXPORT lean_object* l_Lean_Option_set___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__22; +static lean_object* l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__16; LEAN_EXPORT lean_object* l_Lean_instInhabitedOptionDecl; LEAN_EXPORT lean_object* l_Lean_getBoolOption___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__19; lean_object* l_Lean_KVMap_getNat(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__11; -LEAN_EXPORT lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option_____x3a___x3a_x3d____1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__6; +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__12; LEAN_EXPORT lean_object* l_Std_RBNode_fold___at_Lean_getOptionDeclsArray___spec__1(lean_object*, lean_object*); lean_object* l_Lean_KVMap_setName(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Option_register___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__32; -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__10; +LEAN_EXPORT lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option___x3a___x3a_x3d____1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__5; +static lean_object* l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__18; lean_object* l_String_intercalate(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_RBNode_find___at_Lean_getOptionDecl___spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__20; +static lean_object* l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__7; +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__15; LEAN_EXPORT lean_object* l_Lean_Option_get_x3f___rarg___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__19; LEAN_EXPORT lean_object* l_Lean_instMonadWithOptions___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_registerOption___lambda__2___closed__1; +static lean_object* l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__24; LEAN_EXPORT lean_object* l_Lean_setOptionFromString(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__37; +static lean_object* l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__15; static lean_object* l_Lean_setOptionFromString___closed__6; -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__4; +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__42; +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__28; LEAN_EXPORT lean_object* l_Lean_registerOption___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_registerOption___lambda__1___closed__1; -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__17; +static lean_object* l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__10; +lean_object* l_Lean_quoteNameMk(lean_object*); static lean_object* l_Lean_setOptionFromString___closed__1; -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__8; +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__3; lean_object* l_Lean_KVMap_insertCore(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getBoolOption___rarg___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_getNatOption___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__5; +static lean_object* l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__20; LEAN_EXPORT lean_object* l_Lean_Option_get_x3f___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__30; LEAN_EXPORT lean_object* lean_register_option(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__8; LEAN_EXPORT uint8_t l_Lean_Options_getInPattern(lean_object*); -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__40; -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__9; -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__27; +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__18; +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__14; +static lean_object* l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__3; static lean_object* l_Lean_registerOption___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_instInhabitedOption(lean_object*); -static lean_object* l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__10; +static lean_object* l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__10; +lean_object* l_Lean_Syntax_mkNameLit(lean_object*, lean_object*); +static lean_object* l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__4; +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__17; +static lean_object* l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__12; +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__8; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__4; +static lean_object* l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__11; LEAN_EXPORT lean_object* l_Lean_Option_Decl_descr___default; -static lean_object* l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__17; -static lean_object* l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__7; LEAN_EXPORT lean_object* l_Lean_getBoolOption(lean_object*); lean_object* l_Lean_KVMap_findCore(lean_object*, lean_object*); +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__1; LEAN_EXPORT lean_object* l_Lean_Option_set(lean_object*); -static lean_object* l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__4; lean_object* l_String_trim(lean_object*); -static lean_object* l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__9; +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option___x3a___x3a_x3d____1___closed__2; +static lean_object* l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__2; +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__44; +static lean_object* l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__8; uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); static lean_object* l_Lean_instToStringOptions___closed__1; static lean_object* l_Lean_getOptionDecl___closed__2; -static lean_object* l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__2; LEAN_EXPORT lean_object* l_Std_RBNode_fold___at_Lean_getOptionDeclsArray___spec__1___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__43; +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__9; LEAN_EXPORT lean_object* l_Lean_Option_get(lean_object*); +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__40; static lean_object* l_Lean_setOptionFromString___closed__2; -static lean_object* l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__1; +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__27; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option_____x3a___x3a_x3d____1___closed__2; -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__33; +static lean_object* l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__26; +LEAN_EXPORT lean_object* l_Lean_Option_commandRegister__option___x3a___x3a_x3d__; static lean_object* l_Lean_withInPattern___rarg___lambda__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1(lean_object*, lean_object*, lean_object*); lean_object* l_String_toNat_x3f(lean_object*); -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__22; +static lean_object* l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__3; +static lean_object* l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__5; LEAN_EXPORT lean_object* l_Lean_registerOption___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__26; static lean_object* l_Lean_withInPattern___rarg___lambda__1___closed__2; static lean_object* l_Lean_setOptionFromString___closed__7; LEAN_EXPORT lean_object* l_Lean_instInhabitedOptionDecls; +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__2; LEAN_EXPORT lean_object* l_Lean_getOptionDescr(lean_object*, lean_object*); -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__26; static lean_object* l_Lean_setOptionFromString___closed__3; LEAN_EXPORT lean_object* l_Lean_getNatOption___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTRAux___at_Lean_setOptionFromString___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedOption___rarg(lean_object*); -lean_object* l_Lean_Syntax_mkLit(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__12; +static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__43; static lean_object* _init_l_Lean_Options_empty() { _start: { @@ -2058,7 +2057,7 @@ x_2 = lean_alloc_closure((void*)(l_Lean_Option_register___rarg), 4, 0); return x_2; } } -static lean_object* _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__1() { +static lean_object* _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__1() { _start: { lean_object* x_1; @@ -2066,17 +2065,17 @@ x_1 = lean_mk_string_from_bytes("Lean", 4); return x_1; } } -static lean_object* _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__2() { +static lean_object* _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__1; +x_2 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__3() { +static lean_object* _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__3() { _start: { lean_object* x_1; @@ -2084,35 +2083,35 @@ x_1 = lean_mk_string_from_bytes("Option", 6); return x_1; } } -static lean_object* _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__4() { +static lean_object* _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__2; -x_2 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__3; +x_1 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__2; +x_2 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__5() { +static lean_object* _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__5() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("commandRegister_builtin_option__:_:=_", 37); +x_1 = lean_mk_string_from_bytes("commandRegister_builtin_option_:_:=_", 36); return x_1; } } -static lean_object* _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__6() { +static lean_object* _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__4; -x_2 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__5; +x_1 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__4; +x_2 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__7() { +static lean_object* _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__7() { _start: { lean_object* x_1; @@ -2120,17 +2119,17 @@ x_1 = lean_mk_string_from_bytes("andthen", 7); return x_1; } } -static lean_object* _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__8() { +static lean_object* _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__7; +x_2 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__7; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__9() { +static lean_object* _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__9() { _start: { lean_object* x_1; @@ -2138,17 +2137,17 @@ x_1 = lean_mk_string_from_bytes("register_builtin_option", 23); return x_1; } } -static lean_object* _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__10() { +static lean_object* _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__9; +x_1 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__9; x_2 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__11() { +static lean_object* _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__11() { _start: { lean_object* x_1; @@ -2156,33 +2155,33 @@ x_1 = lean_mk_string_from_bytes("ident", 5); return x_1; } } -static lean_object* _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__12() { +static lean_object* _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__11; +x_2 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__11; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__13() { +static lean_object* _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__13() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__12; +x_1 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__12; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__14() { +static lean_object* _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__8; -x_2 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__10; -x_3 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__13; +x_1 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__8; +x_2 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__10; +x_3 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__13; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2190,7 +2189,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__15() { +static lean_object* _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__15() { _start: { lean_object* x_1; @@ -2198,23 +2197,23 @@ x_1 = lean_mk_string_from_bytes(" : ", 3); return x_1; } } -static lean_object* _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__16() { +static lean_object* _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__16() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__15; +x_1 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__15; x_2 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__17() { +static lean_object* _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__8; -x_2 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__14; -x_3 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__16; +x_1 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__8; +x_2 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__14; +x_3 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__16; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2222,7 +2221,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__18() { +static lean_object* _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__18() { _start: { lean_object* x_1; @@ -2230,21 +2229,21 @@ x_1 = lean_mk_string_from_bytes("term", 4); return x_1; } } -static lean_object* _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__19() { +static lean_object* _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__18; +x_2 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__18; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__20() { +static lean_object* _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__19; +x_1 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__19; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2252,13 +2251,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__21() { +static lean_object* _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__8; -x_2 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__17; -x_3 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__20; +x_1 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__8; +x_2 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__17; +x_3 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__20; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2266,7 +2265,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__22() { +static lean_object* _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__22() { _start: { lean_object* x_1; @@ -2274,23 +2273,23 @@ x_1 = lean_mk_string_from_bytes(" := ", 4); return x_1; } } -static lean_object* _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__23() { +static lean_object* _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__23() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__22; +x_1 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__22; x_2 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__24() { +static lean_object* _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__8; -x_2 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__21; -x_3 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__23; +x_1 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__8; +x_2 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__21; +x_3 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__23; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2298,13 +2297,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__25() { +static lean_object* _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__8; -x_2 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__24; -x_3 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__20; +x_1 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__8; +x_2 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__24; +x_3 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__20; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2312,13 +2311,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__26() { +static lean_object* _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__6; +x_1 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__6; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__25; +x_3 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__25; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2326,15 +2325,15 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d__() { +static lean_object* _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d__() { _start: { lean_object* x_1; -x_1 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__26; +x_1 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__26; return x_1; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__1() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__1() { _start: { lean_object* x_1; @@ -2342,17 +2341,17 @@ x_1 = lean_mk_string_from_bytes("Parser", 6); return x_1; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__2() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__2; -x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__1; +x_1 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__2; +x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__3() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__3() { _start: { lean_object* x_1; @@ -2360,17 +2359,17 @@ x_1 = lean_mk_string_from_bytes("Command", 7); return x_1; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__4() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__2; -x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__3; +x_1 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__2; +x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__5() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__5() { _start: { lean_object* x_1; @@ -2378,17 +2377,17 @@ x_1 = lean_mk_string_from_bytes("builtin_initialize", 18); return x_1; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__6() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__4; -x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__5; +x_1 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__4; +x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__7() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__7() { _start: { lean_object* x_1; @@ -2396,22 +2395,22 @@ x_1 = lean_mk_string_from_bytes("null", 4); return x_1; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__8() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__7; +x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__7; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__9() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__8; +x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__8; x_3 = l_Lean_getOptionDeclsArray___closed__1; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -2420,7 +2419,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__10() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__10() { _start: { lean_object* x_1; @@ -2428,17 +2427,17 @@ x_1 = lean_mk_string_from_bytes("Term", 4); return x_1; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__11() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__2; -x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__10; +x_1 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__2; +x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__10; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__12() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__12() { _start: { lean_object* x_1; @@ -2446,17 +2445,17 @@ x_1 = lean_mk_string_from_bytes("typeSpec", 8); return x_1; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__13() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__11; -x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__12; +x_1 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__11; +x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__12; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__14() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__14() { _start: { lean_object* x_1; @@ -2464,7 +2463,7 @@ x_1 = lean_mk_string_from_bytes(":", 1); return x_1; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__15() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__15() { _start: { lean_object* x_1; @@ -2472,17 +2471,17 @@ x_1 = lean_mk_string_from_bytes("app", 3); return x_1; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__16() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__11; -x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__15; +x_1 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__11; +x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__15; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__17() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__17() { _start: { lean_object* x_1; @@ -2490,22 +2489,22 @@ x_1 = lean_mk_string_from_bytes("Lean.Option", 11); return x_1; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__18() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__18() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__17; +x_1 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__17; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__19() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__17; +x_1 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__17; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__18; +x_3 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__18; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2513,31 +2512,31 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__20() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__4; +x_2 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__4; 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_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__21() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__20; +x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__20; x_3 = lean_alloc_ctor(1, 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_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__22() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__22() { _start: { lean_object* x_1; lean_object* x_2; @@ -2546,7 +2545,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__23() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__23() { _start: { lean_object* x_1; lean_object* x_2; @@ -2555,7 +2554,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__24() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__24() { _start: { lean_object* x_1; @@ -2563,7 +2562,7 @@ x_1 = lean_mk_string_from_bytes("←", 3); return x_1; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__25() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__25() { _start: { lean_object* x_1; lean_object* x_2; @@ -2572,7 +2571,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__26() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__26() { _start: { lean_object* x_1; @@ -2580,17 +2579,17 @@ x_1 = lean_mk_string_from_bytes("doSeqIndent", 11); return x_1; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__27() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__27() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__11; -x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__26; +x_1 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__11; +x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__26; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__28() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__28() { _start: { lean_object* x_1; @@ -2598,17 +2597,17 @@ x_1 = lean_mk_string_from_bytes("doSeqItem", 9); return x_1; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__29() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__29() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__11; -x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__28; +x_1 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__11; +x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__28; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__30() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__30() { _start: { lean_object* x_1; @@ -2616,17 +2615,17 @@ x_1 = lean_mk_string_from_bytes("doExpr", 6); return x_1; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__31() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__31() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__11; -x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__30; +x_1 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__11; +x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__30; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__32() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__32() { _start: { lean_object* x_1; @@ -2634,22 +2633,22 @@ x_1 = lean_mk_string_from_bytes("Lean.Option.register", 20); return x_1; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__33() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__33() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__32; +x_1 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__32; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__34() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__34() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__32; +x_1 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__32; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__33; +x_3 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__33; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2657,7 +2656,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__35() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__35() { _start: { lean_object* x_1; @@ -2665,41 +2664,41 @@ x_1 = lean_mk_string_from_bytes("register", 8); return x_1; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__36() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__36() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__4; -x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__35; +x_1 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__4; +x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__35; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__37() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__37() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__36; +x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__36; 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_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__38() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__38() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__37; +x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__37; x_3 = lean_alloc_ctor(1, 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_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__39() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__39() { _start: { lean_object* x_1; lean_object* x_2; @@ -2708,17 +2707,17 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__40() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__40() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__39; -x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__9; +x_1 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__39; +x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__9; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__41() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__41() { _start: { lean_object* x_1; @@ -2726,17 +2725,17 @@ x_1 = lean_mk_string_from_bytes("quotedName", 10); return x_1; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__42() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__42() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__11; -x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__41; +x_1 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__11; +x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__41; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__43() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__43() { _start: { lean_object* x_1; @@ -2744,7 +2743,7 @@ x_1 = lean_mk_string_from_bytes(".", 1); return x_1; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__44() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__44() { _start: { lean_object* x_1; @@ -2752,11 +2751,11 @@ x_1 = lean_mk_string_from_bytes("`", 1); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; uint8_t x_5; -x_4 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__6; +x_4 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__6; lean_inc(x_1); x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); if (x_5 == 0) @@ -2792,58 +2791,58 @@ lean_inc(x_17); x_18 = lean_ctor_get(x_2, 1); lean_inc(x_18); lean_dec(x_2); -x_19 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__5; +x_19 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__5; lean_inc(x_16); x_20 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_20, 0, x_16); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__14; +x_21 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__14; lean_inc(x_16); x_22 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_22, 0, x_16); lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__4; +x_23 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__4; lean_inc(x_17); lean_inc(x_18); x_24 = l_Lean_addMacroScope(x_18, x_23, x_17); x_25 = lean_box(0); -x_26 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__19; -x_27 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__21; +x_26 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__19; +x_27 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__21; lean_inc(x_16); x_28 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_28, 0, x_16); lean_ctor_set(x_28, 1, x_26); lean_ctor_set(x_28, 2, x_24); lean_ctor_set(x_28, 3, x_27); -x_29 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__22; +x_29 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__22; x_30 = lean_array_push(x_29, x_11); x_31 = lean_box(2); -x_32 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__8; +x_32 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__8; x_33 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_33, 0, x_31); lean_ctor_set(x_33, 1, x_32); lean_ctor_set(x_33, 2, x_30); -x_34 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__23; +x_34 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__23; x_35 = lean_array_push(x_34, x_28); x_36 = lean_array_push(x_35, x_33); -x_37 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__16; +x_37 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__16; x_38 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_38, 0, x_31); lean_ctor_set(x_38, 1, x_37); lean_ctor_set(x_38, 2, x_36); x_39 = lean_array_push(x_34, x_22); x_40 = lean_array_push(x_39, x_38); -x_41 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__13; +x_41 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__13; x_42 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_42, 0, x_31); lean_ctor_set(x_42, 1, x_41); lean_ctor_set(x_42, 2, x_40); -x_43 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__24; +x_43 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__24; lean_inc(x_16); x_44 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_44, 0, x_16); lean_ctor_set(x_44, 1, x_43); -x_45 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__25; +x_45 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__25; lean_inc(x_9); x_46 = lean_array_push(x_45, x_9); x_47 = lean_array_push(x_46, x_42); @@ -2852,10 +2851,10 @@ x_49 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_49, 0, x_31); lean_ctor_set(x_49, 1, x_32); lean_ctor_set(x_49, 2, x_48); -x_50 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__36; +x_50 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__36; x_51 = l_Lean_addMacroScope(x_18, x_50, x_17); -x_52 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__34; -x_53 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__38; +x_52 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__34; +x_53 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__38; x_54 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_54, 0, x_16); lean_ctor_set(x_54, 1, x_52); @@ -2866,13 +2865,13 @@ lean_dec(x_9); lean_inc(x_55); x_56 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_25, x_55); x_57 = lean_array_push(x_34, x_54); -x_58 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__40; +x_58 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__40; x_59 = lean_array_push(x_58, x_20); x_60 = lean_array_push(x_59, x_49); if (lean_obj_tag(x_56) == 0) { lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; -x_61 = l___private_Init_Meta_0__Lean_quoteNameMk(x_55); +x_61 = l_Lean_quoteNameMk(x_55); x_62 = lean_array_push(x_34, x_61); x_63 = lean_array_push(x_62, x_13); x_64 = lean_alloc_ctor(1, 3, 0); @@ -2885,15 +2884,15 @@ lean_ctor_set(x_66, 0, x_31); lean_ctor_set(x_66, 1, x_37); lean_ctor_set(x_66, 2, x_65); x_67 = lean_array_push(x_29, x_66); -x_68 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__31; +x_68 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__31; x_69 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_69, 0, x_31); lean_ctor_set(x_69, 1, x_68); lean_ctor_set(x_69, 2, x_67); x_70 = lean_array_push(x_34, x_69); -x_71 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__9; +x_71 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__9; x_72 = lean_array_push(x_70, x_71); -x_73 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__29; +x_73 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__29; x_74 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_74, 0, x_31); lean_ctor_set(x_74, 1, x_73); @@ -2904,13 +2903,13 @@ lean_ctor_set(x_76, 0, x_31); lean_ctor_set(x_76, 1, x_32); lean_ctor_set(x_76, 2, x_75); x_77 = lean_array_push(x_29, x_76); -x_78 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__27; +x_78 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__27; x_79 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_79, 0, x_31); lean_ctor_set(x_79, 1, x_78); lean_ctor_set(x_79, 2, x_77); x_80 = lean_array_push(x_60, x_79); -x_81 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__6; +x_81 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__6; x_82 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_82, 0, x_31); lean_ctor_set(x_82, 1, x_81); @@ -2920,301 +2919,299 @@ return x_14; } else { -lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_dec(x_55); x_83 = lean_ctor_get(x_56, 0); lean_inc(x_83); lean_dec(x_56); -x_84 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__43; +x_84 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__43; x_85 = l_String_intercalate(x_84, x_83); -x_86 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__44; +x_86 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__44; x_87 = lean_string_append(x_86, x_85); lean_dec(x_85); -x_88 = l_Lean_nameLitKind; -x_89 = l_Lean_Syntax_mkLit(x_88, x_87, x_31); -x_90 = lean_array_push(x_29, x_89); -x_91 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__42; -x_92 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_92, 0, x_31); -lean_ctor_set(x_92, 1, x_91); -lean_ctor_set(x_92, 2, x_90); -x_93 = lean_array_push(x_34, x_92); -x_94 = lean_array_push(x_93, x_13); -x_95 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_95, 0, x_31); -lean_ctor_set(x_95, 1, x_32); -lean_ctor_set(x_95, 2, x_94); -x_96 = lean_array_push(x_57, x_95); -x_97 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_97, 0, x_31); -lean_ctor_set(x_97, 1, x_37); -lean_ctor_set(x_97, 2, x_96); -x_98 = lean_array_push(x_29, x_97); -x_99 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__31; -x_100 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_100, 0, x_31); -lean_ctor_set(x_100, 1, x_99); -lean_ctor_set(x_100, 2, x_98); -x_101 = lean_array_push(x_34, x_100); -x_102 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__9; -x_103 = lean_array_push(x_101, x_102); -x_104 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__29; -x_105 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_105, 0, x_31); -lean_ctor_set(x_105, 1, x_104); -lean_ctor_set(x_105, 2, x_103); -x_106 = lean_array_push(x_29, x_105); -x_107 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_107, 0, x_31); -lean_ctor_set(x_107, 1, x_32); -lean_ctor_set(x_107, 2, x_106); -x_108 = lean_array_push(x_29, x_107); -x_109 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__27; -x_110 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_110, 0, x_31); -lean_ctor_set(x_110, 1, x_109); -lean_ctor_set(x_110, 2, x_108); -x_111 = lean_array_push(x_60, x_110); -x_112 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__6; -x_113 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_113, 0, x_31); -lean_ctor_set(x_113, 1, x_112); -lean_ctor_set(x_113, 2, x_111); -lean_ctor_set(x_14, 0, x_113); +x_88 = l_Lean_Syntax_mkNameLit(x_87, x_31); +x_89 = lean_array_push(x_29, x_88); +x_90 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__42; +x_91 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_91, 0, x_31); +lean_ctor_set(x_91, 1, x_90); +lean_ctor_set(x_91, 2, x_89); +x_92 = lean_array_push(x_34, x_91); +x_93 = lean_array_push(x_92, x_13); +x_94 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_94, 0, x_31); +lean_ctor_set(x_94, 1, x_32); +lean_ctor_set(x_94, 2, x_93); +x_95 = lean_array_push(x_57, x_94); +x_96 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_96, 0, x_31); +lean_ctor_set(x_96, 1, x_37); +lean_ctor_set(x_96, 2, x_95); +x_97 = lean_array_push(x_29, x_96); +x_98 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__31; +x_99 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_99, 0, x_31); +lean_ctor_set(x_99, 1, x_98); +lean_ctor_set(x_99, 2, x_97); +x_100 = lean_array_push(x_34, x_99); +x_101 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__9; +x_102 = lean_array_push(x_100, x_101); +x_103 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__29; +x_104 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_104, 0, x_31); +lean_ctor_set(x_104, 1, x_103); +lean_ctor_set(x_104, 2, x_102); +x_105 = lean_array_push(x_29, x_104); +x_106 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_106, 0, x_31); +lean_ctor_set(x_106, 1, x_32); +lean_ctor_set(x_106, 2, x_105); +x_107 = lean_array_push(x_29, x_106); +x_108 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__27; +x_109 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_109, 0, x_31); +lean_ctor_set(x_109, 1, x_108); +lean_ctor_set(x_109, 2, x_107); +x_110 = lean_array_push(x_60, x_109); +x_111 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__6; +x_112 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_112, 0, x_31); +lean_ctor_set(x_112, 1, x_111); +lean_ctor_set(x_112, 2, x_110); +lean_ctor_set(x_14, 0, x_112); return x_14; } } else { -lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; -x_114 = lean_ctor_get(x_14, 0); -x_115 = lean_ctor_get(x_14, 1); -lean_inc(x_115); +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; +x_113 = lean_ctor_get(x_14, 0); +x_114 = lean_ctor_get(x_14, 1); lean_inc(x_114); +lean_inc(x_113); lean_dec(x_14); -x_116 = lean_ctor_get(x_2, 2); +x_115 = lean_ctor_get(x_2, 2); +lean_inc(x_115); +x_116 = lean_ctor_get(x_2, 1); lean_inc(x_116); -x_117 = lean_ctor_get(x_2, 1); -lean_inc(x_117); lean_dec(x_2); -x_118 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__5; -lean_inc(x_114); -x_119 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_119, 0, x_114); -lean_ctor_set(x_119, 1, x_118); -x_120 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__14; -lean_inc(x_114); -x_121 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_121, 0, x_114); -lean_ctor_set(x_121, 1, x_120); -x_122 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__4; +x_117 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__5; +lean_inc(x_113); +x_118 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_118, 0, x_113); +lean_ctor_set(x_118, 1, x_117); +x_119 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__14; +lean_inc(x_113); +x_120 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_120, 0, x_113); +lean_ctor_set(x_120, 1, x_119); +x_121 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__4; +lean_inc(x_115); lean_inc(x_116); -lean_inc(x_117); -x_123 = l_Lean_addMacroScope(x_117, x_122, x_116); -x_124 = lean_box(0); -x_125 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__19; -x_126 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__21; -lean_inc(x_114); -x_127 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_127, 0, x_114); -lean_ctor_set(x_127, 1, x_125); -lean_ctor_set(x_127, 2, x_123); -lean_ctor_set(x_127, 3, x_126); -x_128 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__22; -x_129 = lean_array_push(x_128, x_11); -x_130 = lean_box(2); -x_131 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__8; -x_132 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_132, 0, x_130); -lean_ctor_set(x_132, 1, x_131); -lean_ctor_set(x_132, 2, x_129); -x_133 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__23; -x_134 = lean_array_push(x_133, x_127); -x_135 = lean_array_push(x_134, x_132); -x_136 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__16; -x_137 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_137, 0, x_130); -lean_ctor_set(x_137, 1, x_136); -lean_ctor_set(x_137, 2, x_135); -x_138 = lean_array_push(x_133, x_121); -x_139 = lean_array_push(x_138, x_137); -x_140 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__13; -x_141 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_141, 0, x_130); -lean_ctor_set(x_141, 1, x_140); -lean_ctor_set(x_141, 2, x_139); -x_142 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__24; -lean_inc(x_114); -x_143 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_143, 0, x_114); -lean_ctor_set(x_143, 1, x_142); -x_144 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__25; +x_122 = l_Lean_addMacroScope(x_116, x_121, x_115); +x_123 = lean_box(0); +x_124 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__19; +x_125 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__21; +lean_inc(x_113); +x_126 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_126, 0, x_113); +lean_ctor_set(x_126, 1, x_124); +lean_ctor_set(x_126, 2, x_122); +lean_ctor_set(x_126, 3, x_125); +x_127 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__22; +x_128 = lean_array_push(x_127, x_11); +x_129 = lean_box(2); +x_130 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__8; +x_131 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_131, 0, x_129); +lean_ctor_set(x_131, 1, x_130); +lean_ctor_set(x_131, 2, x_128); +x_132 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__23; +x_133 = lean_array_push(x_132, x_126); +x_134 = lean_array_push(x_133, x_131); +x_135 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__16; +x_136 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_136, 0, x_129); +lean_ctor_set(x_136, 1, x_135); +lean_ctor_set(x_136, 2, x_134); +x_137 = lean_array_push(x_132, x_120); +x_138 = lean_array_push(x_137, x_136); +x_139 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__13; +x_140 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_140, 0, x_129); +lean_ctor_set(x_140, 1, x_139); +lean_ctor_set(x_140, 2, x_138); +x_141 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__24; +lean_inc(x_113); +x_142 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_142, 0, x_113); +lean_ctor_set(x_142, 1, x_141); +x_143 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__25; lean_inc(x_9); -x_145 = lean_array_push(x_144, x_9); -x_146 = lean_array_push(x_145, x_141); -x_147 = lean_array_push(x_146, x_143); -x_148 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_148, 0, x_130); -lean_ctor_set(x_148, 1, x_131); -lean_ctor_set(x_148, 2, x_147); -x_149 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__36; -x_150 = l_Lean_addMacroScope(x_117, x_149, x_116); -x_151 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__34; -x_152 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__38; -x_153 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_153, 0, x_114); -lean_ctor_set(x_153, 1, x_151); -lean_ctor_set(x_153, 2, x_150); -lean_ctor_set(x_153, 3, x_152); -x_154 = l_Lean_Syntax_getId(x_9); +x_144 = lean_array_push(x_143, x_9); +x_145 = lean_array_push(x_144, x_140); +x_146 = lean_array_push(x_145, x_142); +x_147 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_147, 0, x_129); +lean_ctor_set(x_147, 1, x_130); +lean_ctor_set(x_147, 2, x_146); +x_148 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__36; +x_149 = l_Lean_addMacroScope(x_116, x_148, x_115); +x_150 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__34; +x_151 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__38; +x_152 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_152, 0, x_113); +lean_ctor_set(x_152, 1, x_150); +lean_ctor_set(x_152, 2, x_149); +lean_ctor_set(x_152, 3, x_151); +x_153 = l_Lean_Syntax_getId(x_9); lean_dec(x_9); -lean_inc(x_154); -x_155 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_124, x_154); -x_156 = lean_array_push(x_133, x_153); -x_157 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__40; -x_158 = lean_array_push(x_157, x_119); -x_159 = lean_array_push(x_158, x_148); -if (lean_obj_tag(x_155) == 0) -{ -lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; -x_160 = l___private_Init_Meta_0__Lean_quoteNameMk(x_154); -x_161 = lean_array_push(x_133, x_160); -x_162 = lean_array_push(x_161, x_13); -x_163 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_163, 0, x_130); -lean_ctor_set(x_163, 1, x_131); -lean_ctor_set(x_163, 2, x_162); -x_164 = lean_array_push(x_156, x_163); -x_165 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_165, 0, x_130); -lean_ctor_set(x_165, 1, x_136); -lean_ctor_set(x_165, 2, x_164); -x_166 = lean_array_push(x_128, x_165); -x_167 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__31; -x_168 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_168, 0, x_130); -lean_ctor_set(x_168, 1, x_167); -lean_ctor_set(x_168, 2, x_166); -x_169 = lean_array_push(x_133, x_168); -x_170 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__9; -x_171 = lean_array_push(x_169, x_170); -x_172 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__29; -x_173 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_173, 0, x_130); -lean_ctor_set(x_173, 1, x_172); -lean_ctor_set(x_173, 2, x_171); -x_174 = lean_array_push(x_128, x_173); -x_175 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_175, 0, x_130); -lean_ctor_set(x_175, 1, x_131); -lean_ctor_set(x_175, 2, x_174); -x_176 = lean_array_push(x_128, x_175); -x_177 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__27; -x_178 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_178, 0, x_130); -lean_ctor_set(x_178, 1, x_177); -lean_ctor_set(x_178, 2, x_176); -x_179 = lean_array_push(x_159, x_178); -x_180 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__6; -x_181 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_181, 0, x_130); -lean_ctor_set(x_181, 1, x_180); -lean_ctor_set(x_181, 2, x_179); -x_182 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_182, 0, x_181); -lean_ctor_set(x_182, 1, x_115); -return x_182; +lean_inc(x_153); +x_154 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_123, x_153); +x_155 = lean_array_push(x_132, x_152); +x_156 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__40; +x_157 = lean_array_push(x_156, x_118); +x_158 = lean_array_push(x_157, x_147); +if (lean_obj_tag(x_154) == 0) +{ +lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; +x_159 = l_Lean_quoteNameMk(x_153); +x_160 = lean_array_push(x_132, x_159); +x_161 = lean_array_push(x_160, x_13); +x_162 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_162, 0, x_129); +lean_ctor_set(x_162, 1, x_130); +lean_ctor_set(x_162, 2, x_161); +x_163 = lean_array_push(x_155, x_162); +x_164 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_164, 0, x_129); +lean_ctor_set(x_164, 1, x_135); +lean_ctor_set(x_164, 2, x_163); +x_165 = lean_array_push(x_127, x_164); +x_166 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__31; +x_167 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_167, 0, x_129); +lean_ctor_set(x_167, 1, x_166); +lean_ctor_set(x_167, 2, x_165); +x_168 = lean_array_push(x_132, x_167); +x_169 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__9; +x_170 = lean_array_push(x_168, x_169); +x_171 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__29; +x_172 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_172, 0, x_129); +lean_ctor_set(x_172, 1, x_171); +lean_ctor_set(x_172, 2, x_170); +x_173 = lean_array_push(x_127, x_172); +x_174 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_174, 0, x_129); +lean_ctor_set(x_174, 1, x_130); +lean_ctor_set(x_174, 2, x_173); +x_175 = lean_array_push(x_127, x_174); +x_176 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__27; +x_177 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_177, 0, x_129); +lean_ctor_set(x_177, 1, x_176); +lean_ctor_set(x_177, 2, x_175); +x_178 = lean_array_push(x_158, x_177); +x_179 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__6; +x_180 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_180, 0, x_129); +lean_ctor_set(x_180, 1, x_179); +lean_ctor_set(x_180, 2, x_178); +x_181 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_181, 0, x_180); +lean_ctor_set(x_181, 1, x_114); +return x_181; } else { -lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; +lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; +lean_dec(x_153); +x_182 = lean_ctor_get(x_154, 0); +lean_inc(x_182); lean_dec(x_154); -x_183 = lean_ctor_get(x_155, 0); -lean_inc(x_183); -lean_dec(x_155); -x_184 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__43; -x_185 = l_String_intercalate(x_184, x_183); -x_186 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__44; -x_187 = lean_string_append(x_186, x_185); -lean_dec(x_185); -x_188 = l_Lean_nameLitKind; -x_189 = l_Lean_Syntax_mkLit(x_188, x_187, x_130); -x_190 = lean_array_push(x_128, x_189); -x_191 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__42; -x_192 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_192, 0, x_130); -lean_ctor_set(x_192, 1, x_191); -lean_ctor_set(x_192, 2, x_190); -x_193 = lean_array_push(x_133, x_192); -x_194 = lean_array_push(x_193, x_13); +x_183 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__43; +x_184 = l_String_intercalate(x_183, x_182); +x_185 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__44; +x_186 = lean_string_append(x_185, x_184); +lean_dec(x_184); +x_187 = l_Lean_Syntax_mkNameLit(x_186, x_129); +x_188 = lean_array_push(x_127, x_187); +x_189 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__42; +x_190 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_190, 0, x_129); +lean_ctor_set(x_190, 1, x_189); +lean_ctor_set(x_190, 2, x_188); +x_191 = lean_array_push(x_132, x_190); +x_192 = lean_array_push(x_191, x_13); +x_193 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_193, 0, x_129); +lean_ctor_set(x_193, 1, x_130); +lean_ctor_set(x_193, 2, x_192); +x_194 = lean_array_push(x_155, x_193); x_195 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_195, 0, x_130); -lean_ctor_set(x_195, 1, x_131); +lean_ctor_set(x_195, 0, x_129); +lean_ctor_set(x_195, 1, x_135); lean_ctor_set(x_195, 2, x_194); -x_196 = lean_array_push(x_156, x_195); -x_197 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_197, 0, x_130); -lean_ctor_set(x_197, 1, x_136); -lean_ctor_set(x_197, 2, x_196); -x_198 = lean_array_push(x_128, x_197); -x_199 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__31; -x_200 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_200, 0, x_130); -lean_ctor_set(x_200, 1, x_199); -lean_ctor_set(x_200, 2, x_198); -x_201 = lean_array_push(x_133, x_200); -x_202 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__9; -x_203 = lean_array_push(x_201, x_202); -x_204 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__29; +x_196 = lean_array_push(x_127, x_195); +x_197 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__31; +x_198 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_198, 0, x_129); +lean_ctor_set(x_198, 1, x_197); +lean_ctor_set(x_198, 2, x_196); +x_199 = lean_array_push(x_132, x_198); +x_200 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__9; +x_201 = lean_array_push(x_199, x_200); +x_202 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__29; +x_203 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_203, 0, x_129); +lean_ctor_set(x_203, 1, x_202); +lean_ctor_set(x_203, 2, x_201); +x_204 = lean_array_push(x_127, x_203); x_205 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_205, 0, x_130); -lean_ctor_set(x_205, 1, x_204); -lean_ctor_set(x_205, 2, x_203); -x_206 = lean_array_push(x_128, x_205); -x_207 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_207, 0, x_130); -lean_ctor_set(x_207, 1, x_131); -lean_ctor_set(x_207, 2, x_206); -x_208 = lean_array_push(x_128, x_207); -x_209 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__27; -x_210 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_210, 0, x_130); -lean_ctor_set(x_210, 1, x_209); -lean_ctor_set(x_210, 2, x_208); -x_211 = lean_array_push(x_159, x_210); -x_212 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__6; -x_213 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_213, 0, x_130); -lean_ctor_set(x_213, 1, x_212); -lean_ctor_set(x_213, 2, x_211); -x_214 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_214, 0, x_213); -lean_ctor_set(x_214, 1, x_115); -return x_214; -} -} -} -} -} -static lean_object* _init_l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__1() { +lean_ctor_set(x_205, 0, x_129); +lean_ctor_set(x_205, 1, x_130); +lean_ctor_set(x_205, 2, x_204); +x_206 = lean_array_push(x_127, x_205); +x_207 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__27; +x_208 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_208, 0, x_129); +lean_ctor_set(x_208, 1, x_207); +lean_ctor_set(x_208, 2, x_206); +x_209 = lean_array_push(x_158, x_208); +x_210 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__6; +x_211 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_211, 0, x_129); +lean_ctor_set(x_211, 1, x_210); +lean_ctor_set(x_211, 2, x_209); +x_212 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_212, 0, x_211); +lean_ctor_set(x_212, 1, x_114); +return x_212; +} +} +} +} +} +static lean_object* _init_l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("commandRegister_option__:_:=_", 29); +x_1 = lean_mk_string_from_bytes("commandRegister_option_:_:=_", 28); return x_1; } } -static lean_object* _init_l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__2() { +static lean_object* _init_l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__4; -x_2 = l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__1; +x_1 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__4; +x_2 = l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__3() { +static lean_object* _init_l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__3() { _start: { lean_object* x_1; @@ -3222,23 +3219,23 @@ x_1 = lean_mk_string_from_bytes("register_option", 15); return x_1; } } -static lean_object* _init_l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__4() { +static lean_object* _init_l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__3; +x_1 = l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__3; x_2 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__5() { +static lean_object* _init_l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__8; -x_2 = l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__4; -x_3 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__13; +x_1 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__8; +x_2 = l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__4; +x_3 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__13; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3246,13 +3243,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__6() { +static lean_object* _init_l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__8; -x_2 = l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__5; -x_3 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__16; +x_1 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__8; +x_2 = l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__5; +x_3 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__16; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3260,13 +3257,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__7() { +static lean_object* _init_l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__8; -x_2 = l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__6; -x_3 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__20; +x_1 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__8; +x_2 = l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__6; +x_3 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__20; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3274,13 +3271,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__8() { +static lean_object* _init_l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__8; -x_2 = l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__7; -x_3 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__23; +x_1 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__8; +x_2 = l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__7; +x_3 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__23; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3288,13 +3285,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__9() { +static lean_object* _init_l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__8; -x_2 = l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__8; -x_3 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__20; +x_1 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__8; +x_2 = l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__8; +x_3 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__20; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3302,13 +3299,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__10() { +static lean_object* _init_l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__2; +x_1 = l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__9; +x_3 = l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__9; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3316,15 +3313,15 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Option_commandRegister__option_____x3a___x3a_x3d__() { +static lean_object* _init_l_Lean_Option_commandRegister__option___x3a___x3a_x3d__() { _start: { lean_object* x_1; -x_1 = l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__10; +x_1 = l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__10; return x_1; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option_____x3a___x3a_x3d____1___closed__1() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option___x3a___x3a_x3d____1___closed__1() { _start: { lean_object* x_1; @@ -3332,21 +3329,21 @@ x_1 = lean_mk_string_from_bytes("initialize", 10); return x_1; } } -static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option_____x3a___x3a_x3d____1___closed__2() { +static lean_object* _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option___x3a___x3a_x3d____1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__4; -x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option_____x3a___x3a_x3d____1___closed__1; +x_1 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__4; +x_2 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option___x3a___x3a_x3d____1___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option_____x3a___x3a_x3d____1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option___x3a___x3a_x3d____1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; uint8_t x_5; -x_4 = l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__2; +x_4 = l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__2; lean_inc(x_1); x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); if (x_5 == 0) @@ -3382,58 +3379,58 @@ lean_inc(x_17); x_18 = lean_ctor_get(x_2, 1); lean_inc(x_18); lean_dec(x_2); -x_19 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option_____x3a___x3a_x3d____1___closed__1; +x_19 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option___x3a___x3a_x3d____1___closed__1; lean_inc(x_16); x_20 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_20, 0, x_16); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__14; +x_21 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__14; lean_inc(x_16); x_22 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_22, 0, x_16); lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__4; +x_23 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__4; lean_inc(x_17); lean_inc(x_18); x_24 = l_Lean_addMacroScope(x_18, x_23, x_17); x_25 = lean_box(0); -x_26 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__19; -x_27 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__21; +x_26 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__19; +x_27 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__21; lean_inc(x_16); x_28 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_28, 0, x_16); lean_ctor_set(x_28, 1, x_26); lean_ctor_set(x_28, 2, x_24); lean_ctor_set(x_28, 3, x_27); -x_29 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__22; +x_29 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__22; x_30 = lean_array_push(x_29, x_11); x_31 = lean_box(2); -x_32 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__8; +x_32 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__8; x_33 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_33, 0, x_31); lean_ctor_set(x_33, 1, x_32); lean_ctor_set(x_33, 2, x_30); -x_34 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__23; +x_34 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__23; x_35 = lean_array_push(x_34, x_28); x_36 = lean_array_push(x_35, x_33); -x_37 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__16; +x_37 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__16; x_38 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_38, 0, x_31); lean_ctor_set(x_38, 1, x_37); lean_ctor_set(x_38, 2, x_36); x_39 = lean_array_push(x_34, x_22); x_40 = lean_array_push(x_39, x_38); -x_41 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__13; +x_41 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__13; x_42 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_42, 0, x_31); lean_ctor_set(x_42, 1, x_41); lean_ctor_set(x_42, 2, x_40); -x_43 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__24; +x_43 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__24; lean_inc(x_16); x_44 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_44, 0, x_16); lean_ctor_set(x_44, 1, x_43); -x_45 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__25; +x_45 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__25; lean_inc(x_9); x_46 = lean_array_push(x_45, x_9); x_47 = lean_array_push(x_46, x_42); @@ -3442,10 +3439,10 @@ x_49 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_49, 0, x_31); lean_ctor_set(x_49, 1, x_32); lean_ctor_set(x_49, 2, x_48); -x_50 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__36; +x_50 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__36; x_51 = l_Lean_addMacroScope(x_18, x_50, x_17); -x_52 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__34; -x_53 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__38; +x_52 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__34; +x_53 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__38; x_54 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_54, 0, x_16); lean_ctor_set(x_54, 1, x_52); @@ -3456,13 +3453,13 @@ lean_dec(x_9); lean_inc(x_55); x_56 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_25, x_55); x_57 = lean_array_push(x_34, x_54); -x_58 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__40; +x_58 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__40; x_59 = lean_array_push(x_58, x_20); x_60 = lean_array_push(x_59, x_49); if (lean_obj_tag(x_56) == 0) { lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; -x_61 = l___private_Init_Meta_0__Lean_quoteNameMk(x_55); +x_61 = l_Lean_quoteNameMk(x_55); x_62 = lean_array_push(x_34, x_61); x_63 = lean_array_push(x_62, x_13); x_64 = lean_alloc_ctor(1, 3, 0); @@ -3475,15 +3472,15 @@ lean_ctor_set(x_66, 0, x_31); lean_ctor_set(x_66, 1, x_37); lean_ctor_set(x_66, 2, x_65); x_67 = lean_array_push(x_29, x_66); -x_68 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__31; +x_68 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__31; x_69 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_69, 0, x_31); lean_ctor_set(x_69, 1, x_68); lean_ctor_set(x_69, 2, x_67); x_70 = lean_array_push(x_34, x_69); -x_71 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__9; +x_71 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__9; x_72 = lean_array_push(x_70, x_71); -x_73 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__29; +x_73 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__29; x_74 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_74, 0, x_31); lean_ctor_set(x_74, 1, x_73); @@ -3494,13 +3491,13 @@ lean_ctor_set(x_76, 0, x_31); lean_ctor_set(x_76, 1, x_32); lean_ctor_set(x_76, 2, x_75); x_77 = lean_array_push(x_29, x_76); -x_78 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__27; +x_78 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__27; x_79 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_79, 0, x_31); lean_ctor_set(x_79, 1, x_78); lean_ctor_set(x_79, 2, x_77); x_80 = lean_array_push(x_60, x_79); -x_81 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option_____x3a___x3a_x3d____1___closed__2; +x_81 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option___x3a___x3a_x3d____1___closed__2; x_82 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_82, 0, x_31); lean_ctor_set(x_82, 1, x_81); @@ -3510,277 +3507,275 @@ return x_14; } else { -lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_dec(x_55); x_83 = lean_ctor_get(x_56, 0); lean_inc(x_83); lean_dec(x_56); -x_84 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__43; +x_84 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__43; x_85 = l_String_intercalate(x_84, x_83); -x_86 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__44; +x_86 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__44; x_87 = lean_string_append(x_86, x_85); lean_dec(x_85); -x_88 = l_Lean_nameLitKind; -x_89 = l_Lean_Syntax_mkLit(x_88, x_87, x_31); -x_90 = lean_array_push(x_29, x_89); -x_91 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__42; -x_92 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_92, 0, x_31); -lean_ctor_set(x_92, 1, x_91); -lean_ctor_set(x_92, 2, x_90); -x_93 = lean_array_push(x_34, x_92); -x_94 = lean_array_push(x_93, x_13); -x_95 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_95, 0, x_31); -lean_ctor_set(x_95, 1, x_32); -lean_ctor_set(x_95, 2, x_94); -x_96 = lean_array_push(x_57, x_95); -x_97 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_97, 0, x_31); -lean_ctor_set(x_97, 1, x_37); -lean_ctor_set(x_97, 2, x_96); -x_98 = lean_array_push(x_29, x_97); -x_99 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__31; -x_100 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_100, 0, x_31); -lean_ctor_set(x_100, 1, x_99); -lean_ctor_set(x_100, 2, x_98); -x_101 = lean_array_push(x_34, x_100); -x_102 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__9; -x_103 = lean_array_push(x_101, x_102); -x_104 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__29; -x_105 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_105, 0, x_31); -lean_ctor_set(x_105, 1, x_104); -lean_ctor_set(x_105, 2, x_103); -x_106 = lean_array_push(x_29, x_105); -x_107 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_107, 0, x_31); -lean_ctor_set(x_107, 1, x_32); -lean_ctor_set(x_107, 2, x_106); -x_108 = lean_array_push(x_29, x_107); -x_109 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__27; -x_110 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_110, 0, x_31); -lean_ctor_set(x_110, 1, x_109); -lean_ctor_set(x_110, 2, x_108); -x_111 = lean_array_push(x_60, x_110); -x_112 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option_____x3a___x3a_x3d____1___closed__2; -x_113 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_113, 0, x_31); -lean_ctor_set(x_113, 1, x_112); -lean_ctor_set(x_113, 2, x_111); -lean_ctor_set(x_14, 0, x_113); +x_88 = l_Lean_Syntax_mkNameLit(x_87, x_31); +x_89 = lean_array_push(x_29, x_88); +x_90 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__42; +x_91 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_91, 0, x_31); +lean_ctor_set(x_91, 1, x_90); +lean_ctor_set(x_91, 2, x_89); +x_92 = lean_array_push(x_34, x_91); +x_93 = lean_array_push(x_92, x_13); +x_94 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_94, 0, x_31); +lean_ctor_set(x_94, 1, x_32); +lean_ctor_set(x_94, 2, x_93); +x_95 = lean_array_push(x_57, x_94); +x_96 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_96, 0, x_31); +lean_ctor_set(x_96, 1, x_37); +lean_ctor_set(x_96, 2, x_95); +x_97 = lean_array_push(x_29, x_96); +x_98 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__31; +x_99 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_99, 0, x_31); +lean_ctor_set(x_99, 1, x_98); +lean_ctor_set(x_99, 2, x_97); +x_100 = lean_array_push(x_34, x_99); +x_101 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__9; +x_102 = lean_array_push(x_100, x_101); +x_103 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__29; +x_104 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_104, 0, x_31); +lean_ctor_set(x_104, 1, x_103); +lean_ctor_set(x_104, 2, x_102); +x_105 = lean_array_push(x_29, x_104); +x_106 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_106, 0, x_31); +lean_ctor_set(x_106, 1, x_32); +lean_ctor_set(x_106, 2, x_105); +x_107 = lean_array_push(x_29, x_106); +x_108 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__27; +x_109 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_109, 0, x_31); +lean_ctor_set(x_109, 1, x_108); +lean_ctor_set(x_109, 2, x_107); +x_110 = lean_array_push(x_60, x_109); +x_111 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option___x3a___x3a_x3d____1___closed__2; +x_112 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_112, 0, x_31); +lean_ctor_set(x_112, 1, x_111); +lean_ctor_set(x_112, 2, x_110); +lean_ctor_set(x_14, 0, x_112); return x_14; } } else { -lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; -x_114 = lean_ctor_get(x_14, 0); -x_115 = lean_ctor_get(x_14, 1); -lean_inc(x_115); +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; +x_113 = lean_ctor_get(x_14, 0); +x_114 = lean_ctor_get(x_14, 1); lean_inc(x_114); +lean_inc(x_113); lean_dec(x_14); -x_116 = lean_ctor_get(x_2, 2); +x_115 = lean_ctor_get(x_2, 2); +lean_inc(x_115); +x_116 = lean_ctor_get(x_2, 1); lean_inc(x_116); -x_117 = lean_ctor_get(x_2, 1); -lean_inc(x_117); lean_dec(x_2); -x_118 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option_____x3a___x3a_x3d____1___closed__1; -lean_inc(x_114); -x_119 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_119, 0, x_114); -lean_ctor_set(x_119, 1, x_118); -x_120 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__14; -lean_inc(x_114); -x_121 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_121, 0, x_114); -lean_ctor_set(x_121, 1, x_120); -x_122 = l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__4; +x_117 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option___x3a___x3a_x3d____1___closed__1; +lean_inc(x_113); +x_118 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_118, 0, x_113); +lean_ctor_set(x_118, 1, x_117); +x_119 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__14; +lean_inc(x_113); +x_120 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_120, 0, x_113); +lean_ctor_set(x_120, 1, x_119); +x_121 = l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__4; +lean_inc(x_115); lean_inc(x_116); -lean_inc(x_117); -x_123 = l_Lean_addMacroScope(x_117, x_122, x_116); -x_124 = lean_box(0); -x_125 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__19; -x_126 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__21; -lean_inc(x_114); -x_127 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_127, 0, x_114); -lean_ctor_set(x_127, 1, x_125); -lean_ctor_set(x_127, 2, x_123); -lean_ctor_set(x_127, 3, x_126); -x_128 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__22; -x_129 = lean_array_push(x_128, x_11); -x_130 = lean_box(2); -x_131 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__8; -x_132 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_132, 0, x_130); -lean_ctor_set(x_132, 1, x_131); -lean_ctor_set(x_132, 2, x_129); -x_133 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__23; -x_134 = lean_array_push(x_133, x_127); -x_135 = lean_array_push(x_134, x_132); -x_136 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__16; -x_137 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_137, 0, x_130); -lean_ctor_set(x_137, 1, x_136); -lean_ctor_set(x_137, 2, x_135); -x_138 = lean_array_push(x_133, x_121); -x_139 = lean_array_push(x_138, x_137); -x_140 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__13; -x_141 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_141, 0, x_130); -lean_ctor_set(x_141, 1, x_140); -lean_ctor_set(x_141, 2, x_139); -x_142 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__24; -lean_inc(x_114); -x_143 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_143, 0, x_114); -lean_ctor_set(x_143, 1, x_142); -x_144 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__25; +x_122 = l_Lean_addMacroScope(x_116, x_121, x_115); +x_123 = lean_box(0); +x_124 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__19; +x_125 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__21; +lean_inc(x_113); +x_126 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_126, 0, x_113); +lean_ctor_set(x_126, 1, x_124); +lean_ctor_set(x_126, 2, x_122); +lean_ctor_set(x_126, 3, x_125); +x_127 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__22; +x_128 = lean_array_push(x_127, x_11); +x_129 = lean_box(2); +x_130 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__8; +x_131 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_131, 0, x_129); +lean_ctor_set(x_131, 1, x_130); +lean_ctor_set(x_131, 2, x_128); +x_132 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__23; +x_133 = lean_array_push(x_132, x_126); +x_134 = lean_array_push(x_133, x_131); +x_135 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__16; +x_136 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_136, 0, x_129); +lean_ctor_set(x_136, 1, x_135); +lean_ctor_set(x_136, 2, x_134); +x_137 = lean_array_push(x_132, x_120); +x_138 = lean_array_push(x_137, x_136); +x_139 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__13; +x_140 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_140, 0, x_129); +lean_ctor_set(x_140, 1, x_139); +lean_ctor_set(x_140, 2, x_138); +x_141 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__24; +lean_inc(x_113); +x_142 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_142, 0, x_113); +lean_ctor_set(x_142, 1, x_141); +x_143 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__25; lean_inc(x_9); -x_145 = lean_array_push(x_144, x_9); -x_146 = lean_array_push(x_145, x_141); -x_147 = lean_array_push(x_146, x_143); -x_148 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_148, 0, x_130); -lean_ctor_set(x_148, 1, x_131); -lean_ctor_set(x_148, 2, x_147); -x_149 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__36; -x_150 = l_Lean_addMacroScope(x_117, x_149, x_116); -x_151 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__34; -x_152 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__38; -x_153 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_153, 0, x_114); -lean_ctor_set(x_153, 1, x_151); -lean_ctor_set(x_153, 2, x_150); -lean_ctor_set(x_153, 3, x_152); -x_154 = l_Lean_Syntax_getId(x_9); +x_144 = lean_array_push(x_143, x_9); +x_145 = lean_array_push(x_144, x_140); +x_146 = lean_array_push(x_145, x_142); +x_147 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_147, 0, x_129); +lean_ctor_set(x_147, 1, x_130); +lean_ctor_set(x_147, 2, x_146); +x_148 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__36; +x_149 = l_Lean_addMacroScope(x_116, x_148, x_115); +x_150 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__34; +x_151 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__38; +x_152 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_152, 0, x_113); +lean_ctor_set(x_152, 1, x_150); +lean_ctor_set(x_152, 2, x_149); +lean_ctor_set(x_152, 3, x_151); +x_153 = l_Lean_Syntax_getId(x_9); lean_dec(x_9); -lean_inc(x_154); -x_155 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_124, x_154); -x_156 = lean_array_push(x_133, x_153); -x_157 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__40; -x_158 = lean_array_push(x_157, x_119); -x_159 = lean_array_push(x_158, x_148); -if (lean_obj_tag(x_155) == 0) -{ -lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; -x_160 = l___private_Init_Meta_0__Lean_quoteNameMk(x_154); -x_161 = lean_array_push(x_133, x_160); -x_162 = lean_array_push(x_161, x_13); -x_163 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_163, 0, x_130); -lean_ctor_set(x_163, 1, x_131); -lean_ctor_set(x_163, 2, x_162); -x_164 = lean_array_push(x_156, x_163); -x_165 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_165, 0, x_130); -lean_ctor_set(x_165, 1, x_136); -lean_ctor_set(x_165, 2, x_164); -x_166 = lean_array_push(x_128, x_165); -x_167 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__31; -x_168 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_168, 0, x_130); -lean_ctor_set(x_168, 1, x_167); -lean_ctor_set(x_168, 2, x_166); -x_169 = lean_array_push(x_133, x_168); -x_170 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__9; -x_171 = lean_array_push(x_169, x_170); -x_172 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__29; -x_173 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_173, 0, x_130); -lean_ctor_set(x_173, 1, x_172); -lean_ctor_set(x_173, 2, x_171); -x_174 = lean_array_push(x_128, x_173); -x_175 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_175, 0, x_130); -lean_ctor_set(x_175, 1, x_131); -lean_ctor_set(x_175, 2, x_174); -x_176 = lean_array_push(x_128, x_175); -x_177 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__27; -x_178 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_178, 0, x_130); -lean_ctor_set(x_178, 1, x_177); -lean_ctor_set(x_178, 2, x_176); -x_179 = lean_array_push(x_159, x_178); -x_180 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option_____x3a___x3a_x3d____1___closed__2; -x_181 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_181, 0, x_130); -lean_ctor_set(x_181, 1, x_180); -lean_ctor_set(x_181, 2, x_179); -x_182 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_182, 0, x_181); -lean_ctor_set(x_182, 1, x_115); -return x_182; +lean_inc(x_153); +x_154 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_123, x_153); +x_155 = lean_array_push(x_132, x_152); +x_156 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__40; +x_157 = lean_array_push(x_156, x_118); +x_158 = lean_array_push(x_157, x_147); +if (lean_obj_tag(x_154) == 0) +{ +lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; +x_159 = l_Lean_quoteNameMk(x_153); +x_160 = lean_array_push(x_132, x_159); +x_161 = lean_array_push(x_160, x_13); +x_162 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_162, 0, x_129); +lean_ctor_set(x_162, 1, x_130); +lean_ctor_set(x_162, 2, x_161); +x_163 = lean_array_push(x_155, x_162); +x_164 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_164, 0, x_129); +lean_ctor_set(x_164, 1, x_135); +lean_ctor_set(x_164, 2, x_163); +x_165 = lean_array_push(x_127, x_164); +x_166 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__31; +x_167 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_167, 0, x_129); +lean_ctor_set(x_167, 1, x_166); +lean_ctor_set(x_167, 2, x_165); +x_168 = lean_array_push(x_132, x_167); +x_169 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__9; +x_170 = lean_array_push(x_168, x_169); +x_171 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__29; +x_172 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_172, 0, x_129); +lean_ctor_set(x_172, 1, x_171); +lean_ctor_set(x_172, 2, x_170); +x_173 = lean_array_push(x_127, x_172); +x_174 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_174, 0, x_129); +lean_ctor_set(x_174, 1, x_130); +lean_ctor_set(x_174, 2, x_173); +x_175 = lean_array_push(x_127, x_174); +x_176 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__27; +x_177 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_177, 0, x_129); +lean_ctor_set(x_177, 1, x_176); +lean_ctor_set(x_177, 2, x_175); +x_178 = lean_array_push(x_158, x_177); +x_179 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option___x3a___x3a_x3d____1___closed__2; +x_180 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_180, 0, x_129); +lean_ctor_set(x_180, 1, x_179); +lean_ctor_set(x_180, 2, x_178); +x_181 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_181, 0, x_180); +lean_ctor_set(x_181, 1, x_114); +return x_181; } else { -lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; +lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; +lean_dec(x_153); +x_182 = lean_ctor_get(x_154, 0); +lean_inc(x_182); lean_dec(x_154); -x_183 = lean_ctor_get(x_155, 0); -lean_inc(x_183); -lean_dec(x_155); -x_184 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__43; -x_185 = l_String_intercalate(x_184, x_183); -x_186 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__44; -x_187 = lean_string_append(x_186, x_185); -lean_dec(x_185); -x_188 = l_Lean_nameLitKind; -x_189 = l_Lean_Syntax_mkLit(x_188, x_187, x_130); -x_190 = lean_array_push(x_128, x_189); -x_191 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__42; -x_192 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_192, 0, x_130); -lean_ctor_set(x_192, 1, x_191); -lean_ctor_set(x_192, 2, x_190); -x_193 = lean_array_push(x_133, x_192); -x_194 = lean_array_push(x_193, x_13); +x_183 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__43; +x_184 = l_String_intercalate(x_183, x_182); +x_185 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__44; +x_186 = lean_string_append(x_185, x_184); +lean_dec(x_184); +x_187 = l_Lean_Syntax_mkNameLit(x_186, x_129); +x_188 = lean_array_push(x_127, x_187); +x_189 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__42; +x_190 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_190, 0, x_129); +lean_ctor_set(x_190, 1, x_189); +lean_ctor_set(x_190, 2, x_188); +x_191 = lean_array_push(x_132, x_190); +x_192 = lean_array_push(x_191, x_13); +x_193 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_193, 0, x_129); +lean_ctor_set(x_193, 1, x_130); +lean_ctor_set(x_193, 2, x_192); +x_194 = lean_array_push(x_155, x_193); x_195 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_195, 0, x_130); -lean_ctor_set(x_195, 1, x_131); +lean_ctor_set(x_195, 0, x_129); +lean_ctor_set(x_195, 1, x_135); lean_ctor_set(x_195, 2, x_194); -x_196 = lean_array_push(x_156, x_195); -x_197 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_197, 0, x_130); -lean_ctor_set(x_197, 1, x_136); -lean_ctor_set(x_197, 2, x_196); -x_198 = lean_array_push(x_128, x_197); -x_199 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__31; -x_200 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_200, 0, x_130); -lean_ctor_set(x_200, 1, x_199); -lean_ctor_set(x_200, 2, x_198); -x_201 = lean_array_push(x_133, x_200); -x_202 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__9; -x_203 = lean_array_push(x_201, x_202); -x_204 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__29; +x_196 = lean_array_push(x_127, x_195); +x_197 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__31; +x_198 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_198, 0, x_129); +lean_ctor_set(x_198, 1, x_197); +lean_ctor_set(x_198, 2, x_196); +x_199 = lean_array_push(x_132, x_198); +x_200 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__9; +x_201 = lean_array_push(x_199, x_200); +x_202 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__29; +x_203 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_203, 0, x_129); +lean_ctor_set(x_203, 1, x_202); +lean_ctor_set(x_203, 2, x_201); +x_204 = lean_array_push(x_127, x_203); x_205 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_205, 0, x_130); -lean_ctor_set(x_205, 1, x_204); -lean_ctor_set(x_205, 2, x_203); -x_206 = lean_array_push(x_128, x_205); -x_207 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_207, 0, x_130); -lean_ctor_set(x_207, 1, x_131); -lean_ctor_set(x_207, 2, x_206); -x_208 = lean_array_push(x_128, x_207); -x_209 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__27; -x_210 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_210, 0, x_130); -lean_ctor_set(x_210, 1, x_209); -lean_ctor_set(x_210, 2, x_208); -x_211 = lean_array_push(x_159, x_210); -x_212 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option_____x3a___x3a_x3d____1___closed__2; -x_213 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_213, 0, x_130); -lean_ctor_set(x_213, 1, x_212); -lean_ctor_set(x_213, 2, x_211); -x_214 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_214, 0, x_213); -lean_ctor_set(x_214, 1, x_115); -return x_214; +lean_ctor_set(x_205, 0, x_129); +lean_ctor_set(x_205, 1, x_130); +lean_ctor_set(x_205, 2, x_204); +x_206 = lean_array_push(x_127, x_205); +x_207 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__27; +x_208 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_208, 0, x_129); +lean_ctor_set(x_208, 1, x_207); +lean_ctor_set(x_208, 2, x_206); +x_209 = lean_array_push(x_158, x_208); +x_210 = l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option___x3a___x3a_x3d____1___closed__2; +x_211 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_211, 0, x_129); +lean_ctor_set(x_211, 1, x_210); +lean_ctor_set(x_211, 2, x_209); +x_212 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_212, 0, x_211); +lean_ctor_set(x_212, 1, x_114); +return x_212; } } } @@ -3882,174 +3877,174 @@ l_Lean_Option_Decl_group___default = _init_l_Lean_Option_Decl_group___default(); lean_mark_persistent(l_Lean_Option_Decl_group___default); l_Lean_Option_Decl_descr___default = _init_l_Lean_Option_Decl_descr___default(); lean_mark_persistent(l_Lean_Option_Decl_descr___default); -l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__1 = _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__1(); -lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__1); -l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__2 = _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__2(); -lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__2); -l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__3 = _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__3(); -lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__3); -l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__4 = _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__4(); -lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__4); -l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__5 = _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__5(); -lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__5); -l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__6 = _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__6(); -lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__6); -l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__7 = _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__7(); -lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__7); -l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__8 = _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__8(); -lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__8); -l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__9 = _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__9(); -lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__9); -l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__10 = _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__10(); -lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__10); -l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__11 = _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__11(); -lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__11); -l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__12 = _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__12(); -lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__12); -l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__13 = _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__13(); -lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__13); -l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__14 = _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__14(); -lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__14); -l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__15 = _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__15(); -lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__15); -l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__16 = _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__16(); -lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__16); -l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__17 = _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__17(); -lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__17); -l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__18 = _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__18(); -lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__18); -l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__19 = _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__19(); -lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__19); -l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__20 = _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__20(); -lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__20); -l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__21 = _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__21(); -lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__21); -l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__22 = _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__22(); -lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__22); -l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__23 = _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__23(); -lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__23); -l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__24 = _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__24(); -lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__24); -l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__25 = _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__25(); -lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__25); -l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__26 = _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__26(); -lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d_____closed__26); -l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d__ = _init_l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d__(); -lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option_____x3a___x3a_x3d__); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__1 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__1(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__1); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__2 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__2(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__2); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__3 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__3(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__3); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__4 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__4(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__4); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__5 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__5(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__5); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__6 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__6(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__6); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__7 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__7(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__7); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__8 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__8(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__8); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__9 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__9(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__9); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__10 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__10(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__10); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__11 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__11(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__11); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__12 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__12(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__12); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__13 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__13(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__13); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__14 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__14(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__14); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__15 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__15(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__15); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__16 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__16(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__16); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__17 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__17(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__17); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__18 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__18(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__18); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__19 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__19(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__19); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__20 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__20(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__20); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__21 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__21(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__21); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__22 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__22(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__22); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__23 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__23(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__23); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__24 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__24(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__24); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__25 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__25(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__25); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__26 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__26(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__26); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__27 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__27(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__27); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__28 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__28(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__28); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__29 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__29(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__29); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__30 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__30(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__30); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__31 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__31(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__31); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__32 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__32(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__32); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__33 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__33(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__33); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__34 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__34(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__34); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__35 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__35(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__35); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__36 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__36(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__36); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__37 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__37(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__37); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__38 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__38(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__38); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__39 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__39(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__39); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__40 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__40(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__40); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__41 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__41(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__41); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__42 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__42(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__42); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__43 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__43(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__43); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__44 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__44(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option_____x3a___x3a_x3d____1___closed__44); -l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__1 = _init_l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__1(); -lean_mark_persistent(l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__1); -l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__2 = _init_l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__2(); -lean_mark_persistent(l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__2); -l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__3 = _init_l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__3(); -lean_mark_persistent(l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__3); -l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__4 = _init_l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__4(); -lean_mark_persistent(l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__4); -l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__5 = _init_l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__5(); -lean_mark_persistent(l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__5); -l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__6 = _init_l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__6(); -lean_mark_persistent(l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__6); -l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__7 = _init_l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__7(); -lean_mark_persistent(l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__7); -l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__8 = _init_l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__8(); -lean_mark_persistent(l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__8); -l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__9 = _init_l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__9(); -lean_mark_persistent(l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__9); -l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__10 = _init_l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__10(); -lean_mark_persistent(l_Lean_Option_commandRegister__option_____x3a___x3a_x3d_____closed__10); -l_Lean_Option_commandRegister__option_____x3a___x3a_x3d__ = _init_l_Lean_Option_commandRegister__option_____x3a___x3a_x3d__(); -lean_mark_persistent(l_Lean_Option_commandRegister__option_____x3a___x3a_x3d__); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option_____x3a___x3a_x3d____1___closed__1 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option_____x3a___x3a_x3d____1___closed__1(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option_____x3a___x3a_x3d____1___closed__1); -l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option_____x3a___x3a_x3d____1___closed__2 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option_____x3a___x3a_x3d____1___closed__2(); -lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option_____x3a___x3a_x3d____1___closed__2); +l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__1 = _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__1(); +lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__1); +l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__2 = _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__2(); +lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__2); +l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__3 = _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__3(); +lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__3); +l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__4 = _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__4(); +lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__4); +l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__5 = _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__5(); +lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__5); +l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__6 = _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__6(); +lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__6); +l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__7 = _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__7(); +lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__7); +l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__8 = _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__8(); +lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__8); +l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__9 = _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__9(); +lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__9); +l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__10 = _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__10(); +lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__10); +l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__11 = _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__11(); +lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__11); +l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__12 = _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__12(); +lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__12); +l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__13 = _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__13(); +lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__13); +l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__14 = _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__14(); +lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__14); +l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__15 = _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__15(); +lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__15); +l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__16 = _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__16(); +lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__16); +l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__17 = _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__17(); +lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__17); +l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__18 = _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__18(); +lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__18); +l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__19 = _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__19(); +lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__19); +l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__20 = _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__20(); +lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__20); +l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__21 = _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__21(); +lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__21); +l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__22 = _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__22(); +lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__22); +l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__23 = _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__23(); +lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__23); +l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__24 = _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__24(); +lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__24); +l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__25 = _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__25(); +lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__25); +l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__26 = _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__26(); +lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d_____closed__26); +l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d__ = _init_l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d__(); +lean_mark_persistent(l_Lean_Option_commandRegister__builtin__option___x3a___x3a_x3d__); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__1 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__1(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__1); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__2 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__2(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__2); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__3 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__3(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__3); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__4 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__4(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__4); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__5 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__5(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__5); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__6 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__6(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__6); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__7 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__7(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__7); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__8 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__8(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__8); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__9 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__9(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__9); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__10 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__10(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__10); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__11 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__11(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__11); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__12 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__12(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__12); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__13 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__13(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__13); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__14 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__14(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__14); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__15 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__15(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__15); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__16 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__16(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__16); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__17 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__17(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__17); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__18 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__18(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__18); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__19 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__19(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__19); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__20 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__20(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__20); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__21 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__21(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__21); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__22 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__22(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__22); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__23 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__23(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__23); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__24 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__24(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__24); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__25 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__25(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__25); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__26 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__26(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__26); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__27 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__27(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__27); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__28 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__28(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__28); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__29 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__29(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__29); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__30 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__30(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__30); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__31 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__31(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__31); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__32 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__32(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__32); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__33 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__33(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__33); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__34 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__34(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__34); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__35 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__35(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__35); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__36 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__36(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__36); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__37 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__37(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__37); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__38 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__38(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__38); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__39 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__39(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__39); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__40 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__40(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__40); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__41 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__41(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__41); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__42 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__42(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__42); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__43 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__43(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__43); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__44 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__44(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__builtin__option___x3a___x3a_x3d____1___closed__44); +l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__1 = _init_l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__1(); +lean_mark_persistent(l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__1); +l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__2 = _init_l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__2(); +lean_mark_persistent(l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__2); +l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__3 = _init_l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__3(); +lean_mark_persistent(l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__3); +l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__4 = _init_l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__4(); +lean_mark_persistent(l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__4); +l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__5 = _init_l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__5(); +lean_mark_persistent(l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__5); +l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__6 = _init_l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__6(); +lean_mark_persistent(l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__6); +l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__7 = _init_l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__7(); +lean_mark_persistent(l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__7); +l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__8 = _init_l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__8(); +lean_mark_persistent(l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__8); +l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__9 = _init_l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__9(); +lean_mark_persistent(l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__9); +l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__10 = _init_l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__10(); +lean_mark_persistent(l_Lean_Option_commandRegister__option___x3a___x3a_x3d_____closed__10); +l_Lean_Option_commandRegister__option___x3a___x3a_x3d__ = _init_l_Lean_Option_commandRegister__option___x3a___x3a_x3d__(); +lean_mark_persistent(l_Lean_Option_commandRegister__option___x3a___x3a_x3d__); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option___x3a___x3a_x3d____1___closed__1 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option___x3a___x3a_x3d____1___closed__1(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option___x3a___x3a_x3d____1___closed__1); +l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option___x3a___x3a_x3d____1___closed__2 = _init_l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option___x3a___x3a_x3d____1___closed__2(); +lean_mark_persistent(l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__commandRegister__option___x3a___x3a_x3d____1___closed__2); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Elab/App.c b/stage0/stdlib/Lean/Elab/App.c index 275bc8b5aec..7550577a6f8 100644 --- a/stage0/stdlib/Lean/Elab/App.c +++ b/stage0/stdlib/Lean/Elab/App.c @@ -29,7 +29,6 @@ static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux__ lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_elabForall___spec__1___rarg(lean_object*); size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_shouldPropagateExpectedTypeFor___boxed(lean_object*); -extern lean_object* l_Lean_fieldIdxKind; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_finalize___lambda__4(lean_object*, lean_object*, 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___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_propagateExpectedType___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_addLValArg___spec__1___closed__4; @@ -45,7 +44,6 @@ static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux__ lean_object* lean_erase_macro_scopes(lean_object*); static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_5____closed__4; static lean_object* l_Lean_Elab_Term_elabAppArgs___closed__12; -static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__4; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_addLValArg___spec__2___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); static lean_object* l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__3___rarg___closed__2; @@ -54,15 +52,14 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabExplicit_declRange___close lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppAux___spec__4___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkSort(lean_object*); -static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__8; static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_shouldPropagateExpectedTypeFor___closed__2; static lean_object* l_Lean_addTrace___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_propagateExpectedType___spec__1___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Term_elabPipeProj___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_addNewArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabNamedPattern___closed__2; lean_object* l_Lean_LocalDecl_userName(lean_object*); -extern lean_object* l_Lean_nullKind; static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__8; +LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_findMethodAlias_x3f(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_whnfForall(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppAux___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -114,6 +111,7 @@ lean_object* l_Lean_Elab_Term_instInhabitedTermElabResult___rarg(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabNamedPattern_declRange___closed__5; lean_object* l_Lean_Expr_getAutoParamTactic_x3f(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabIdent_declRange(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_findMethodAlias_x3f___boxed(lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabApp_declRange___closed__6; static lean_object* l_Lean_Elab_Term_throwInvalidNamedArg___rarg___closed__2; @@ -130,7 +128,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppAux___closed__3; static lean_object* l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__2; LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_registerMVarErrorHoleInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_checkTraceOption(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_getForallBody___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -211,10 +208,10 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabExplicit_declRange___close LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppAux___spec__4(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_whnfEasyCases___at_Lean_Meta_whnfCore___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValLoop___boxed(lean_object*, lean_object*, 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_List_filterTRAux___at_Lean_Elab_Term_ElabAppArgs_eraseNamedArgCore___spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_propagateExpectedType___closed__5; lean_object* l_Lean_Elab_Term_registerMVarErrorInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_throwInvalidFieldNotation___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__7; static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppLVals___closed__1; static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_App_0__Lean_Elab_Term_getSuccesses___spec__1___closed__3; lean_object* lean_nat_add(lean_object*, lean_object*); @@ -256,7 +253,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabChoice___closed__3; static lean_object* l_Lean_Elab_Term_elabAppArgs___closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabAppArgs___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Elab_Util_0__Lean_Elab_evalSyntaxConstantUnsafe(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Syntax_SepArray_getElems___rarg(lean_object*); +static lean_object* l_List_forIn_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_mkBaseProjections___spec__1___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_getArgExpectedType___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabAppArgs___closed__7; static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___closed__1; @@ -278,7 +275,6 @@ lean_object* l_Lean_Elab_Term_elabLevel(lean_object*, lean_object*, lean_object* static lean_object* l_Lean_Elab_Term_elabPipeProj___lambda__1___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_addInstMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getFieldInfo_x3f(lean_object*, lean_object*, lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_elabChoice___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Term_elabArrayRef_declRange___closed__3; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_throwInvalidNamedArg___spec__2(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppLVals___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -292,6 +288,7 @@ static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_tryCoeFun_x3f___ lean_object* lean_expr_consume_type_annotations(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_getBindingName(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_synthesizePendingAndNormalizeFunType___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_shouldPropagateExpectedTypeFor___closed__7; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__2___closed__1; lean_object* lean_st_ref_take(lean_object*, lean_object*); @@ -313,7 +310,6 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_App_0 static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_addLValArg___spec__1___closed__6; static lean_object* l_Lean_Elab_Term_elabPipeProj___lambda__1___closed__4; static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppLVals___closed__2; -LEAN_EXPORT lean_object* l_List_filterAux___at_Lean_Elab_Term_ElabAppArgs_eraseNamedArgCore___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__1(lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFnId___spec__1___lambda__1___boxed(lean_object**); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_propagateExpectedTypeFor(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -323,6 +319,7 @@ static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_prop LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_normalizeFunType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__1___boxed(lean_object*, lean_object*, lean_object*, 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___regBuiltin_Lean_Elab_Term_elabProj_declRange(lean_object*); +lean_object* l_Lean_Syntax_isFieldIdx_x3f(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValLoop___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabExplicit_declRange___closed__5; lean_object* l_Lean_Meta_expandCoe(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -350,6 +347,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabApp_declRange___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFnId___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabAppArgs(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabDotIdent(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Name_eraseSuffix_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppAux___spec__4___boxed(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_Meta_withMCtx___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppAux___spec__2(lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_mkProjAndCheck___closed__4; @@ -390,6 +388,7 @@ static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppLValsAux_ static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___closed__2; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_addLValArg___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Array_foldlMUnsafe_fold___at___private_Lean_Elab_App_0__Lean_Elab_Term_getSuccesses___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_filterTRAux___at_Lean_Elab_Term_ElabAppArgs_eraseNamedArgCore___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_addLValArg___lambda__1___closed__1; lean_object* l_Lean_Syntax_getId(lean_object*); lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -401,11 +400,9 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_addLValArg_ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabPipeProj_declRange(lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_mkProjAndCheck___closed__1; lean_object* lean_format_pretty(lean_object*, lean_object*); -extern lean_object* l_Lean_choiceKind; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveDotName_go___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabNamedPattern_declRange___closed__4; -lean_object* l_panic___at_String_toNat_x21___spec__1(lean_object*); static lean_object* l_List_forIn_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_mkBaseProjections___spec__1___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_elabExplicit_declRange___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_elabApp___closed__5; @@ -487,6 +484,7 @@ static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_mkProjAndCheck__ static lean_object* l_Lean_Elab_Term_instToStringNamedArg___closed__3; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_throwInvalidNamedArg___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_shouldPropagateExpectedTypeFor___closed__10; +lean_object* l_Lean_getAliases(lean_object*, lean_object*, uint8_t); uint8_t l_Lean_Expr_isForall(lean_object*); uint8_t l_Lean_BinderInfo_isInstImplicit(uint8_t); static lean_object* l___regBuiltin_Lean_Elab_Term_elabExplicitUniv_declRange___closed__7; @@ -521,6 +519,7 @@ static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_ LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppLVals(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__2(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_hasOptAutoParams(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__28; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_addEtaArg___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_BinderInfo_isImplicit(uint8_t); lean_object* l_Lean_Elab_Term_mkConst(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -536,9 +535,9 @@ static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___clos LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Elab_Term_termElabAttribute; static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___closed__3; -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFnId(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_finalize___lambda__4___closed__1; +lean_object* l_Lean_Syntax_TSepArray_getElems___rarg(lean_object*); lean_object* l_Lean_LocalDecl_type(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_addLValArg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, 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_Array_anyMUnsafe_any___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_hasOptAutoParams___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -555,7 +554,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_instToStringArg(lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabWithoutExpectedTypeAttr; static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppLValsAux_loop___lambda__3___closed__5; -LEAN_EXPORT lean_object* l_List_filterAux___at_Lean_Elab_Term_ElabAppArgs_eraseNamedArgCore___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_App_0__Lean_Elab_Term_getSuccesses___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_PersistentArray_push___rarg(lean_object*, lean_object*); @@ -628,7 +626,6 @@ static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppLValsAux_ static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveDotName_go___closed__3; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValLoop___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabExplicitUniv(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__6; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_propagateExpectedType___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__17; static lean_object* l___regBuiltin_Lean_Elab_Term_elabDotIdent_declRange___closed__3; @@ -643,6 +640,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs lean_object* l_Lean_Elab_Term_tryPostponeIfMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabAppArgs___closed__8; lean_object* l_Lean_Elab_Term_resolveName_x27(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_List_filterMap___at___private_Lean_Elab_App_0__Lean_Elab_Term_findMethodAlias_x3f___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_getSuccesses___lambda__1(lean_object*, 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_Term_ElabAppArgs_eraseNamedArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_anyNamedArgDependsOnCurrent___spec__2___closed__2; @@ -654,7 +652,7 @@ uint8_t l_Lean_Expr_isFVar(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_finalize___lambda__1(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___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_getForallBody___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_mkProjAndCheck___closed__7; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_13490_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_13748_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_5_(lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppAux(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -671,6 +669,7 @@ static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_addLValArg___lam lean_object* l_Lean_Meta_mkArrow(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_constName_x3f(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabApp_declRange___closed__4; +lean_object* l_panic___at_Lean_TSyntax_getNat___spec__1(lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_mkProjAndCheck___closed__2; static lean_object* l_Lean_Elab_Term_elabExplicit___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_annotateIfRec(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -747,7 +746,6 @@ LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_El lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_elabChoice___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_elabExplicit_declRange___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppLVals___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -778,6 +776,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ensureArgTy LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabAppArgs___lambda__1(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_throwInvalidNamedArg___spec__1(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabPipeProj_declRange___closed__4; +static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__27; static lean_object* l___regBuiltin_Lean_Elab_Term_elabExplicitUniv___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_elabProj_declRange___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -832,7 +831,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppLVal LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_Elab_Term_synthesizeInst___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_getForallBody(uint8_t, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__5; static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_synthesizePendingAndNormalizeFunType___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabExplicitUnivs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_finalize___lambda__3___closed__2; @@ -3777,7 +3775,7 @@ lean_dec(x_1); return x_9; } } -LEAN_EXPORT lean_object* l_List_filterAux___at_Lean_Elab_Term_ElabAppArgs_eraseNamedArgCore___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_filterTRAux___at_Lean_Elab_Term_ElabAppArgs_eraseNamedArgCore___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_2) == 0) @@ -3855,16 +3853,16 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_ElabAppArgs_eraseNamedArgCore(lean_obj { lean_object* x_3; lean_object* x_4; x_3 = lean_box(0); -x_4 = l_List_filterAux___at_Lean_Elab_Term_ElabAppArgs_eraseNamedArgCore___spec__1(x_2, x_1, x_3); +x_4 = l_List_filterTRAux___at_Lean_Elab_Term_ElabAppArgs_eraseNamedArgCore___spec__1(x_2, x_1, x_3); lean_dec(x_2); return x_4; } } -LEAN_EXPORT lean_object* l_List_filterAux___at_Lean_Elab_Term_ElabAppArgs_eraseNamedArgCore___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_filterTRAux___at_Lean_Elab_Term_ElabAppArgs_eraseNamedArgCore___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_List_filterAux___at_Lean_Elab_Term_ElabAppArgs_eraseNamedArgCore___spec__1(x_1, x_2, x_3); +x_4 = l_List_filterTRAux___at_Lean_Elab_Term_ElabAppArgs_eraseNamedArgCore___spec__1(x_1, x_2, x_3); lean_dec(x_1); return x_4; } @@ -9525,51 +9523,6 @@ x_1 = lean_mk_string_from_bytes("by", 2); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Tactic", 6); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_shouldPropagateExpectedTypeFor___closed__4; -x_2 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__4; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__6() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("tacticSeq", 9); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__7() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__5; -x_2 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__6; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(1u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; -} -} LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg(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_object* x_9) { _start: { @@ -9825,7 +9778,7 @@ return x_65; } else { -lean_object* x_66; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +lean_object* x_66; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; x_66 = lean_ctor_get(x_61, 0); lean_inc(x_66); lean_dec(x_61); @@ -9844,24 +9797,17 @@ x_72 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg x_73 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_73, 0, x_68); lean_ctor_set(x_73, 1, x_72); -x_74 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__8; -x_75 = lean_array_push(x_74, x_66); -x_76 = lean_box(2); -x_77 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__7; -x_78 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_78, 0, x_76); -lean_ctor_set(x_78, 1, x_77); -lean_ctor_set(x_78, 2, x_75); -x_79 = l___private_Lean_Elab_App_0__Lean_Elab_Term_tryCoeFun_x3f___closed__3; -x_80 = lean_array_push(x_79, x_73); -x_81 = lean_array_push(x_80, x_78); -x_82 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_shouldPropagateExpectedTypeFor___closed__12; -x_83 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_83, 0, x_76); -lean_ctor_set(x_83, 1, x_82); -lean_ctor_set(x_83, 2, x_81); -x_84 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_84, 0, x_83); +x_74 = l___private_Lean_Elab_App_0__Lean_Elab_Term_tryCoeFun_x3f___closed__3; +x_75 = lean_array_push(x_74, x_73); +x_76 = lean_array_push(x_75, x_66); +x_77 = lean_box(2); +x_78 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_shouldPropagateExpectedTypeFor___closed__12; +x_79 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_79, 0, x_77); +lean_ctor_set(x_79, 1, x_78); +lean_ctor_set(x_79, 2, x_76); +x_80 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_80, 0, x_79); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); @@ -9869,33 +9815,33 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -lean_inc(x_84); -x_85 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_propagateExpectedType(x_84, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_71); -if (lean_obj_tag(x_85) == 0) +lean_inc(x_80); +x_81 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_propagateExpectedType(x_80, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_71); +if (lean_obj_tag(x_81) == 0) { -lean_object* x_86; lean_object* x_87; -x_86 = lean_ctor_get(x_85, 1); -lean_inc(x_86); -lean_dec(x_85); +lean_object* x_82; lean_object* x_83; +x_82 = lean_ctor_get(x_81, 1); +lean_inc(x_82); +lean_dec(x_81); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_87 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_elabAndAddNewArg(x_1, x_84, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_86); -if (lean_obj_tag(x_87) == 0) +x_83 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_elabAndAddNewArg(x_1, x_80, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_82); +if (lean_obj_tag(x_83) == 0) { -lean_object* x_88; lean_object* x_89; -x_88 = lean_ctor_get(x_87, 1); -lean_inc(x_88); -lean_dec(x_87); -x_89 = l_Lean_Elab_Term_ElabAppArgs_main(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_88); -return x_89; +lean_object* x_84; lean_object* x_85; +x_84 = lean_ctor_get(x_83, 1); +lean_inc(x_84); +lean_dec(x_83); +x_85 = l_Lean_Elab_Term_ElabAppArgs_main(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_84); +return x_85; } else { -uint8_t x_90; +uint8_t x_86; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -9903,30 +9849,30 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_90 = !lean_is_exclusive(x_87); -if (x_90 == 0) +x_86 = !lean_is_exclusive(x_83); +if (x_86 == 0) { -return x_87; +return x_83; } else { -lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_91 = lean_ctor_get(x_87, 0); -x_92 = lean_ctor_get(x_87, 1); -lean_inc(x_92); -lean_inc(x_91); -lean_dec(x_87); -x_93 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_93, 0, x_91); -lean_ctor_set(x_93, 1, x_92); -return x_93; +lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_87 = lean_ctor_get(x_83, 0); +x_88 = lean_ctor_get(x_83, 1); +lean_inc(x_88); +lean_inc(x_87); +lean_dec(x_83); +x_89 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_89, 0, x_87); +lean_ctor_set(x_89, 1, x_88); +return x_89; } } } else { -uint8_t x_94; -lean_dec(x_84); +uint8_t x_90; +lean_dec(x_80); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -9935,34 +9881,34 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_94 = !lean_is_exclusive(x_85); -if (x_94 == 0) +x_90 = !lean_is_exclusive(x_81); +if (x_90 == 0) { -return x_85; +return x_81; } else { -lean_object* x_95; lean_object* x_96; lean_object* x_97; -x_95 = lean_ctor_get(x_85, 0); -x_96 = lean_ctor_get(x_85, 1); -lean_inc(x_96); -lean_inc(x_95); -lean_dec(x_85); -x_97 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_97, 0, x_95); -lean_ctor_set(x_97, 1, x_96); -return x_97; +lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_91 = lean_ctor_get(x_81, 0); +x_92 = lean_ctor_get(x_81, 1); +lean_inc(x_92); +lean_inc(x_91); +lean_dec(x_81); +x_93 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +return x_93; } } } } else { -lean_object* x_98; lean_object* x_99; +lean_object* x_94; lean_object* x_95; lean_dec(x_54); lean_dec(x_1); -x_98 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__2; -x_99 = l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___spec__1(x_98, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_19); +x_94 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__2; +x_95 = l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___spec__1(x_94, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_19); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -9970,40 +9916,40 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_99; +return x_95; } } } else { -lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; +lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_dec(x_18); lean_dec(x_13); -x_100 = lean_ctor_get(x_20, 0); -lean_inc(x_100); +x_96 = lean_ctor_get(x_20, 0); +lean_inc(x_96); lean_dec(x_20); -x_101 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_addNewArg(x_1, x_100, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_19); -x_102 = lean_ctor_get(x_101, 1); -lean_inc(x_102); -lean_dec(x_101); -x_103 = l_Lean_Elab_Term_ElabAppArgs_main(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_102); -return x_103; +x_97 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_addNewArg(x_1, x_96, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_19); +x_98 = lean_ctor_get(x_97, 1); +lean_inc(x_98); +lean_dec(x_97); +x_99 = l_Lean_Elab_Term_ElabAppArgs_main(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_98); +return x_99; } } else { -lean_object* x_104; lean_object* x_105; uint8_t x_106; -x_104 = lean_ctor_get(x_16, 1); -lean_inc(x_104); +lean_object* x_100; lean_object* x_101; uint8_t x_102; +x_100 = lean_ctor_get(x_16, 1); +lean_inc(x_100); lean_dec(x_16); -x_105 = lean_ctor_get(x_13, 3); -lean_inc(x_105); +x_101 = lean_ctor_get(x_13, 3); +lean_inc(x_101); lean_dec(x_13); -x_106 = l_List_isEmpty___rarg(x_105); -lean_dec(x_105); -if (x_106 == 0) +x_102 = l_List_isEmpty___rarg(x_101); +lean_dec(x_101); +if (x_102 == 0) { -lean_object* x_107; +lean_object* x_103; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); @@ -10011,36 +9957,36 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_107 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_anyNamedArgDependsOnCurrent(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_104); -if (lean_obj_tag(x_107) == 0) +x_103 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_anyNamedArgDependsOnCurrent(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_100); +if (lean_obj_tag(x_103) == 0) { -lean_object* x_108; uint8_t x_109; -x_108 = lean_ctor_get(x_107, 0); -lean_inc(x_108); -x_109 = lean_unbox(x_108); -lean_dec(x_108); -if (x_109 == 0) +lean_object* x_104; uint8_t x_105; +x_104 = lean_ctor_get(x_103, 0); +lean_inc(x_104); +x_105 = lean_unbox(x_104); +lean_dec(x_104); +if (x_105 == 0) { -lean_object* x_110; lean_object* x_111; -x_110 = lean_ctor_get(x_107, 1); -lean_inc(x_110); -lean_dec(x_107); -x_111 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_addEtaArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_110); -return x_111; +lean_object* x_106; lean_object* x_107; +x_106 = lean_ctor_get(x_103, 1); +lean_inc(x_106); +lean_dec(x_103); +x_107 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_addEtaArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_106); +return x_107; } else { -lean_object* x_112; lean_object* x_113; -x_112 = lean_ctor_get(x_107, 1); -lean_inc(x_112); -lean_dec(x_107); -x_113 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_addImplicitArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_112); -return x_113; +lean_object* x_108; lean_object* x_109; +x_108 = lean_ctor_get(x_103, 1); +lean_inc(x_108); +lean_dec(x_103); +x_109 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_addImplicitArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_108); +return x_109; } } else { -uint8_t x_114; +uint8_t x_110; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -10049,46 +9995,46 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_114 = !lean_is_exclusive(x_107); -if (x_114 == 0) +x_110 = !lean_is_exclusive(x_103); +if (x_110 == 0) { -return x_107; +return x_103; } else { -lean_object* x_115; lean_object* x_116; lean_object* x_117; -x_115 = lean_ctor_get(x_107, 0); -x_116 = lean_ctor_get(x_107, 1); -lean_inc(x_116); -lean_inc(x_115); -lean_dec(x_107); -x_117 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_117, 0, x_115); -lean_ctor_set(x_117, 1, x_116); -return x_117; +lean_object* x_111; lean_object* x_112; lean_object* x_113; +x_111 = lean_ctor_get(x_103, 0); +x_112 = lean_ctor_get(x_103, 1); +lean_inc(x_112); +lean_inc(x_111); +lean_dec(x_103); +x_113 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_113, 0, x_111); +lean_ctor_set(x_113, 1, x_112); +return x_113; } } } else { -lean_object* x_118; +lean_object* x_114; lean_dec(x_1); -x_118 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_finalize(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_104); -return x_118; +x_114 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_finalize(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_100); +return x_114; } } } else { -lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; +lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_dec(x_13); -x_119 = lean_ctor_get(x_12, 1); -lean_inc(x_119); +x_115 = lean_ctor_get(x_12, 1); +lean_inc(x_115); lean_dec(x_12); -x_120 = lean_ctor_get(x_14, 0); -lean_inc(x_120); -x_121 = lean_ctor_get(x_14, 1); -lean_inc(x_121); +x_116 = lean_ctor_get(x_14, 0); +lean_inc(x_116); +x_117 = lean_ctor_get(x_14, 1); +lean_inc(x_117); lean_dec(x_14); lean_inc(x_8); lean_inc(x_7); @@ -10097,54 +10043,54 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -lean_inc(x_120); -x_122 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_propagateExpectedType(x_120, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_119); -if (lean_obj_tag(x_122) == 0) +lean_inc(x_116); +x_118 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_propagateExpectedType(x_116, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_115); +if (lean_obj_tag(x_118) == 0) { -lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; uint8_t x_129; -x_123 = lean_ctor_get(x_122, 1); +lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; uint8_t x_125; +x_119 = lean_ctor_get(x_118, 1); +lean_inc(x_119); +lean_dec(x_118); +x_120 = lean_st_ref_get(x_8, x_119); +x_121 = lean_ctor_get(x_120, 1); +lean_inc(x_121); +lean_dec(x_120); +x_122 = lean_st_ref_take(x_2, x_121); +x_123 = lean_ctor_get(x_122, 0); lean_inc(x_123); +x_124 = lean_ctor_get(x_122, 1); +lean_inc(x_124); lean_dec(x_122); -x_124 = lean_st_ref_get(x_8, x_123); -x_125 = lean_ctor_get(x_124, 1); -lean_inc(x_125); -lean_dec(x_124); -x_126 = lean_st_ref_take(x_2, x_125); -x_127 = lean_ctor_get(x_126, 0); -lean_inc(x_127); -x_128 = lean_ctor_get(x_126, 1); -lean_inc(x_128); -lean_dec(x_126); -x_129 = !lean_is_exclusive(x_127); -if (x_129 == 0) +x_125 = !lean_is_exclusive(x_123); +if (x_125 == 0) { -lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; -x_130 = lean_ctor_get(x_127, 2); -lean_dec(x_130); -lean_ctor_set(x_127, 2, x_121); -x_131 = lean_st_ref_set(x_2, x_127, x_128); -x_132 = lean_ctor_get(x_131, 1); -lean_inc(x_132); -lean_dec(x_131); +lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; +x_126 = lean_ctor_get(x_123, 2); +lean_dec(x_126); +lean_ctor_set(x_123, 2, x_117); +x_127 = lean_st_ref_set(x_2, x_123, x_124); +x_128 = lean_ctor_get(x_127, 1); +lean_inc(x_128); +lean_dec(x_127); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_133 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_elabAndAddNewArg(x_1, x_120, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_132); -if (lean_obj_tag(x_133) == 0) +x_129 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_elabAndAddNewArg(x_1, x_116, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_128); +if (lean_obj_tag(x_129) == 0) { -lean_object* x_134; lean_object* x_135; -x_134 = lean_ctor_get(x_133, 1); -lean_inc(x_134); -lean_dec(x_133); -x_135 = l_Lean_Elab_Term_ElabAppArgs_main(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_134); -return x_135; +lean_object* x_130; lean_object* x_131; +x_130 = lean_ctor_get(x_129, 1); +lean_inc(x_130); +lean_dec(x_129); +x_131 = l_Lean_Elab_Term_ElabAppArgs_main(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_130); +return x_131; } else { -uint8_t x_136; +uint8_t x_132; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -10152,82 +10098,82 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_136 = !lean_is_exclusive(x_133); -if (x_136 == 0) +x_132 = !lean_is_exclusive(x_129); +if (x_132 == 0) { -return x_133; +return x_129; } else { -lean_object* x_137; lean_object* x_138; lean_object* x_139; -x_137 = lean_ctor_get(x_133, 0); -x_138 = lean_ctor_get(x_133, 1); -lean_inc(x_138); -lean_inc(x_137); -lean_dec(x_133); -x_139 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_139, 0, x_137); -lean_ctor_set(x_139, 1, x_138); -return x_139; +lean_object* x_133; lean_object* x_134; lean_object* x_135; +x_133 = lean_ctor_get(x_129, 0); +x_134 = lean_ctor_get(x_129, 1); +lean_inc(x_134); +lean_inc(x_133); +lean_dec(x_129); +x_135 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_135, 0, x_133); +lean_ctor_set(x_135, 1, x_134); +return x_135; } } } else { -uint8_t x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; uint8_t x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; uint8_t x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; -x_140 = lean_ctor_get_uint8(x_127, sizeof(void*)*8); -x_141 = lean_ctor_get(x_127, 0); -x_142 = lean_ctor_get(x_127, 1); -x_143 = lean_ctor_get(x_127, 3); -x_144 = lean_ctor_get_uint8(x_127, sizeof(void*)*8 + 1); -x_145 = lean_ctor_get(x_127, 4); -x_146 = lean_ctor_get(x_127, 5); -x_147 = lean_ctor_get(x_127, 6); -x_148 = lean_ctor_get(x_127, 7); -x_149 = lean_ctor_get_uint8(x_127, sizeof(void*)*8 + 2); -lean_inc(x_148); -lean_inc(x_147); -lean_inc(x_146); -lean_inc(x_145); +uint8_t x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; uint8_t x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; uint8_t x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; +x_136 = lean_ctor_get_uint8(x_123, sizeof(void*)*8); +x_137 = lean_ctor_get(x_123, 0); +x_138 = lean_ctor_get(x_123, 1); +x_139 = lean_ctor_get(x_123, 3); +x_140 = lean_ctor_get_uint8(x_123, sizeof(void*)*8 + 1); +x_141 = lean_ctor_get(x_123, 4); +x_142 = lean_ctor_get(x_123, 5); +x_143 = lean_ctor_get(x_123, 6); +x_144 = lean_ctor_get(x_123, 7); +x_145 = lean_ctor_get_uint8(x_123, sizeof(void*)*8 + 2); +lean_inc(x_144); lean_inc(x_143); lean_inc(x_142); lean_inc(x_141); -lean_dec(x_127); -x_150 = lean_alloc_ctor(0, 8, 3); -lean_ctor_set(x_150, 0, x_141); -lean_ctor_set(x_150, 1, x_142); -lean_ctor_set(x_150, 2, x_121); -lean_ctor_set(x_150, 3, x_143); -lean_ctor_set(x_150, 4, x_145); -lean_ctor_set(x_150, 5, x_146); -lean_ctor_set(x_150, 6, x_147); -lean_ctor_set(x_150, 7, x_148); -lean_ctor_set_uint8(x_150, sizeof(void*)*8, x_140); -lean_ctor_set_uint8(x_150, sizeof(void*)*8 + 1, x_144); -lean_ctor_set_uint8(x_150, sizeof(void*)*8 + 2, x_149); -x_151 = lean_st_ref_set(x_2, x_150, x_128); -x_152 = lean_ctor_get(x_151, 1); -lean_inc(x_152); -lean_dec(x_151); +lean_inc(x_139); +lean_inc(x_138); +lean_inc(x_137); +lean_dec(x_123); +x_146 = lean_alloc_ctor(0, 8, 3); +lean_ctor_set(x_146, 0, x_137); +lean_ctor_set(x_146, 1, x_138); +lean_ctor_set(x_146, 2, x_117); +lean_ctor_set(x_146, 3, x_139); +lean_ctor_set(x_146, 4, x_141); +lean_ctor_set(x_146, 5, x_142); +lean_ctor_set(x_146, 6, x_143); +lean_ctor_set(x_146, 7, x_144); +lean_ctor_set_uint8(x_146, sizeof(void*)*8, x_136); +lean_ctor_set_uint8(x_146, sizeof(void*)*8 + 1, x_140); +lean_ctor_set_uint8(x_146, sizeof(void*)*8 + 2, x_145); +x_147 = lean_st_ref_set(x_2, x_146, x_124); +x_148 = lean_ctor_get(x_147, 1); +lean_inc(x_148); +lean_dec(x_147); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_153 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_elabAndAddNewArg(x_1, x_120, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_152); -if (lean_obj_tag(x_153) == 0) +x_149 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_elabAndAddNewArg(x_1, x_116, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_148); +if (lean_obj_tag(x_149) == 0) { -lean_object* x_154; lean_object* x_155; -x_154 = lean_ctor_get(x_153, 1); -lean_inc(x_154); -lean_dec(x_153); -x_155 = l_Lean_Elab_Term_ElabAppArgs_main(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_154); -return x_155; +lean_object* x_150; lean_object* x_151; +x_150 = lean_ctor_get(x_149, 1); +lean_inc(x_150); +lean_dec(x_149); +x_151 = l_Lean_Elab_Term_ElabAppArgs_main(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_150); +return x_151; } else { -lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; +lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -10235,34 +10181,34 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_156 = lean_ctor_get(x_153, 0); -lean_inc(x_156); -x_157 = lean_ctor_get(x_153, 1); -lean_inc(x_157); -if (lean_is_exclusive(x_153)) { - lean_ctor_release(x_153, 0); - lean_ctor_release(x_153, 1); - x_158 = x_153; +x_152 = lean_ctor_get(x_149, 0); +lean_inc(x_152); +x_153 = lean_ctor_get(x_149, 1); +lean_inc(x_153); +if (lean_is_exclusive(x_149)) { + lean_ctor_release(x_149, 0); + lean_ctor_release(x_149, 1); + x_154 = x_149; } else { - lean_dec_ref(x_153); - x_158 = lean_box(0); + lean_dec_ref(x_149); + x_154 = lean_box(0); } -if (lean_is_scalar(x_158)) { - x_159 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_154)) { + x_155 = lean_alloc_ctor(1, 2, 0); } else { - x_159 = x_158; + x_155 = x_154; } -lean_ctor_set(x_159, 0, x_156); -lean_ctor_set(x_159, 1, x_157); -return x_159; +lean_ctor_set(x_155, 0, x_152); +lean_ctor_set(x_155, 1, x_153); +return x_155; } } } else { -uint8_t x_160; -lean_dec(x_121); -lean_dec(x_120); +uint8_t x_156; +lean_dec(x_117); +lean_dec(x_116); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -10271,23 +10217,23 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_160 = !lean_is_exclusive(x_122); -if (x_160 == 0) +x_156 = !lean_is_exclusive(x_118); +if (x_156 == 0) { -return x_122; +return x_118; } else { -lean_object* x_161; lean_object* x_162; lean_object* x_163; -x_161 = lean_ctor_get(x_122, 0); -x_162 = lean_ctor_get(x_122, 1); -lean_inc(x_162); -lean_inc(x_161); -lean_dec(x_122); -x_163 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_163, 0, x_161); -lean_ctor_set(x_163, 1, x_162); -return x_163; +lean_object* x_157; lean_object* x_158; lean_object* x_159; +x_157 = lean_ctor_get(x_118, 0); +x_158 = lean_ctor_get(x_118, 1); +lean_inc(x_158); +lean_inc(x_157); +lean_dec(x_118); +x_159 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_159, 0, x_157); +lean_ctor_set(x_159, 1, x_158); +return x_159; } } } @@ -11779,29 +11725,159 @@ lean_dec(x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_throwInvalidFieldNotation___rarg___closed__1() { +LEAN_EXPORT lean_object* l_List_filterMap___at___private_Lean_Elab_App_0__Lean_Elab_Term_findMethodAlias_x3f___spec__1(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("invalid field notation, type is not of the form (C ...) where C is a constant", 77); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_throwInvalidFieldNotation___rarg___closed__2() { -_start: +if (lean_obj_tag(x_2) == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_App_0__Lean_Elab_Term_throwInvalidFieldNotation___rarg___closed__1; -x_2 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} +lean_object* x_3; +lean_dec(x_1); +x_3 = lean_box(0); +return x_3; } -static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_throwInvalidFieldNotation___rarg___closed__3() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_App_0__Lean_Elab_Term_throwInvalidFieldNotation___rarg___closed__2; +uint8_t x_4; +x_4 = !lean_is_exclusive(x_2); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_2, 0); +x_6 = lean_ctor_get(x_2, 1); +lean_inc(x_1); +lean_inc(x_5); +x_7 = l_Lean_Name_eraseSuffix_x3f(x_5, x_1); +if (lean_obj_tag(x_7) == 0) +{ +lean_free_object(x_2); +lean_dec(x_5); +x_2 = x_6; +goto _start; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = lean_ctor_get(x_7, 0); +lean_inc(x_9); +lean_dec(x_7); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_5); +x_11 = l_List_filterMap___at___private_Lean_Elab_App_0__Lean_Elab_Term_findMethodAlias_x3f___spec__1(x_1, x_6); +lean_ctor_set(x_2, 1, x_11); +lean_ctor_set(x_2, 0, x_10); +return x_2; +} +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_2, 0); +x_13 = lean_ctor_get(x_2, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_2); +lean_inc(x_1); +lean_inc(x_12); +x_14 = l_Lean_Name_eraseSuffix_x3f(x_12, x_1); +if (lean_obj_tag(x_14) == 0) +{ +lean_dec(x_12); +x_2 = x_13; +goto _start; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_14, 0); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_12); +x_18 = l_List_filterMap___at___private_Lean_Elab_App_0__Lean_Elab_Term_findMethodAlias_x3f___spec__1(x_1, x_13); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +return x_19; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_findMethodAlias_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; +lean_inc(x_3); +x_4 = l_Lean_Name_append(x_2, x_3); +x_5 = 0; +x_6 = l_Lean_getAliases(x_1, x_4, x_5); +x_7 = l_List_filterMap___at___private_Lean_Elab_App_0__Lean_Elab_Term_findMethodAlias_x3f___spec__1(x_3, x_6); +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_8; +x_8 = lean_box(0); +return x_8; +} +else +{ +lean_object* x_9; +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_7, 0); +lean_inc(x_10); +lean_dec(x_7); +x_11 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_11, 0, x_10); +return x_11; +} +else +{ +lean_object* x_12; +lean_dec(x_9); +lean_dec(x_7); +x_12 = lean_box(0); +return x_12; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_findMethodAlias_x3f___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Lean_Elab_App_0__Lean_Elab_Term_findMethodAlias_x3f(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_throwInvalidFieldNotation___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("invalid field notation, type is not of the form (C ...) where C is a constant", 77); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_throwInvalidFieldNotation___rarg___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Elab_App_0__Lean_Elab_Term_throwInvalidFieldNotation___rarg___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_Elab_App_0__Lean_Elab_Term_throwInvalidFieldNotation___rarg___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Elab_App_0__Lean_Elab_Term_throwInvalidFieldNotation___rarg___closed__2; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; @@ -12364,41 +12440,47 @@ if (lean_obj_tag(x_48) == 0) lean_object* x_49; lean_object* x_50; lean_inc(x_36); x_49 = lean_name_mk_string(x_45, x_36); +lean_inc(x_49); lean_inc(x_35); +lean_inc(x_41); x_50 = l___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f(x_41, x_35, x_49); if (lean_obj_tag(x_50) == 0) { -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; -lean_free_object(x_37); +lean_object* x_51; +x_51 = l___private_Lean_Elab_App_0__Lean_Elab_Term_findMethodAlias_x3f(x_41, x_35, x_49); lean_dec(x_35); -x_51 = l_Lean_stringToMessageData(x_36); +if (lean_obj_tag(x_51) == 0) +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +lean_free_object(x_37); +x_52 = l_Lean_stringToMessageData(x_36); lean_dec(x_36); -x_52 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__4; -x_53 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_53, 0, x_52); -lean_ctor_set(x_53, 1, x_51); -x_54 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__6; -x_55 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_55, 0, x_53); -lean_ctor_set(x_55, 1, x_54); -x_56 = lean_alloc_ctor(4, 1, 0); -lean_ctor_set(x_56, 0, x_43); -x_57 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_57, 0, x_55); -lean_ctor_set(x_57, 1, x_56); -x_58 = l_Lean_Elab_Term_throwInvalidNamedArg___rarg___closed__8; -x_59 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_59, 0, x_57); -lean_ctor_set(x_59, 1, x_58); -x_60 = l___private_Lean_Elab_App_0__Lean_Elab_Term_throwLValError___rarg(x_3, x_1, x_59, x_5, x_6, x_7, x_8, x_9, x_10, x_40); +x_53 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__4; +x_54 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_54, 0, x_53); +lean_ctor_set(x_54, 1, x_52); +x_55 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__6; +x_56 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_56, 0, x_54); +lean_ctor_set(x_56, 1, x_55); +x_57 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_57, 0, x_43); +x_58 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_58, 0, x_56); +lean_ctor_set(x_58, 1, x_57); +x_59 = l_Lean_Elab_Term_throwInvalidNamedArg___rarg___closed__8; +x_60 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +x_61 = l___private_Lean_Elab_App_0__Lean_Elab_Term_throwLValError___rarg(x_3, x_1, x_60, x_5, x_6, x_7, x_8, x_9, x_10, x_40); lean_dec(x_10); lean_dec(x_8); lean_dec(x_6); -return x_60; +return x_61; } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_dec(x_43); lean_dec(x_36); lean_dec(x_10); @@ -12409,72 +12491,110 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_1); -x_61 = lean_ctor_get(x_50, 0); -lean_inc(x_61); -lean_dec(x_50); -x_62 = lean_ctor_get(x_61, 0); +x_62 = lean_ctor_get(x_51, 0); lean_inc(x_62); -x_63 = lean_ctor_get(x_61, 1); +lean_dec(x_51); +x_63 = lean_ctor_get(x_62, 0); lean_inc(x_63); -lean_dec(x_61); -x_64 = lean_alloc_ctor(2, 3, 0); -lean_ctor_set(x_64, 0, x_62); -lean_ctor_set(x_64, 1, x_35); -lean_ctor_set(x_64, 2, x_63); -lean_ctor_set(x_37, 0, x_64); +x_64 = lean_ctor_get(x_62, 1); +lean_inc(x_64); +lean_dec(x_62); +lean_inc(x_63); +x_65 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_63); +lean_ctor_set(x_65, 2, x_64); +lean_ctor_set(x_37, 0, x_65); return x_37; } } else { -lean_object* x_65; uint8_t x_66; uint8_t x_67; uint8_t x_68; -x_65 = lean_ctor_get(x_48, 0); -lean_inc(x_65); +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +lean_dec(x_49); +lean_dec(x_43); +lean_dec(x_41); +lean_dec(x_36); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_66 = lean_ctor_get(x_50, 0); +lean_inc(x_66); +lean_dec(x_50); +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +lean_dec(x_66); +x_69 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_69, 0, x_67); +lean_ctor_set(x_69, 1, x_35); +lean_ctor_set(x_69, 2, x_68); +lean_ctor_set(x_37, 0, x_69); +return x_37; +} +} +else +{ +lean_object* x_70; uint8_t x_71; uint8_t x_72; uint8_t x_73; +x_70 = lean_ctor_get(x_48, 0); +lean_inc(x_70); lean_dec(x_48); -x_66 = l_Lean_LocalDecl_binderInfo(x_65); -x_67 = 4; -x_68 = l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_362_(x_66, x_67); -if (x_68 == 0) +x_71 = l_Lean_LocalDecl_binderInfo(x_70); +x_72 = 4; +x_73 = l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_362_(x_71, x_72); +if (x_73 == 0) { -lean_object* x_69; lean_object* x_70; -lean_dec(x_65); +lean_object* x_74; lean_object* x_75; +lean_dec(x_70); lean_inc(x_36); -x_69 = lean_name_mk_string(x_45, x_36); +x_74 = lean_name_mk_string(x_45, x_36); +lean_inc(x_74); lean_inc(x_35); -x_70 = l___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f(x_41, x_35, x_69); -if (lean_obj_tag(x_70) == 0) +lean_inc(x_41); +x_75 = l___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f(x_41, x_35, x_74); +if (lean_obj_tag(x_75) == 0) { -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; lean_object* x_79; lean_object* x_80; -lean_free_object(x_37); +lean_object* x_76; +x_76 = l___private_Lean_Elab_App_0__Lean_Elab_Term_findMethodAlias_x3f(x_41, x_35, x_74); lean_dec(x_35); -x_71 = l_Lean_stringToMessageData(x_36); +if (lean_obj_tag(x_76) == 0) +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +lean_free_object(x_37); +x_77 = l_Lean_stringToMessageData(x_36); lean_dec(x_36); -x_72 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__4; -x_73 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_71); -x_74 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__6; -x_75 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_75, 0, x_73); -lean_ctor_set(x_75, 1, x_74); -x_76 = lean_alloc_ctor(4, 1, 0); -lean_ctor_set(x_76, 0, x_43); -x_77 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_77, 0, x_75); -lean_ctor_set(x_77, 1, x_76); -x_78 = l_Lean_Elab_Term_throwInvalidNamedArg___rarg___closed__8; +x_78 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__4; x_79 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_79, 0, x_77); -lean_ctor_set(x_79, 1, x_78); -x_80 = l___private_Lean_Elab_App_0__Lean_Elab_Term_throwLValError___rarg(x_3, x_1, x_79, x_5, x_6, x_7, x_8, x_9, x_10, x_40); +lean_ctor_set(x_79, 0, x_78); +lean_ctor_set(x_79, 1, x_77); +x_80 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__6; +x_81 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_81, 0, x_79); +lean_ctor_set(x_81, 1, x_80); +x_82 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_82, 0, x_43); +x_83 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_83, 0, x_81); +lean_ctor_set(x_83, 1, x_82); +x_84 = l_Lean_Elab_Term_throwInvalidNamedArg___rarg___closed__8; +x_85 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_85, 0, x_83); +lean_ctor_set(x_85, 1, x_84); +x_86 = l___private_Lean_Elab_App_0__Lean_Elab_Term_throwLValError___rarg(x_3, x_1, x_85, x_5, x_6, x_7, x_8, x_9, x_10, x_40); lean_dec(x_10); lean_dec(x_8); lean_dec(x_6); -return x_80; +return x_86; } else { -lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; +lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_dec(x_43); lean_dec(x_36); lean_dec(x_10); @@ -12485,25 +12605,28 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_1); -x_81 = lean_ctor_get(x_70, 0); -lean_inc(x_81); -lean_dec(x_70); -x_82 = lean_ctor_get(x_81, 0); -lean_inc(x_82); -x_83 = lean_ctor_get(x_81, 1); -lean_inc(x_83); -lean_dec(x_81); -x_84 = lean_alloc_ctor(2, 3, 0); -lean_ctor_set(x_84, 0, x_82); -lean_ctor_set(x_84, 1, x_35); -lean_ctor_set(x_84, 2, x_83); -lean_ctor_set(x_37, 0, x_84); +x_87 = lean_ctor_get(x_76, 0); +lean_inc(x_87); +lean_dec(x_76); +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); +lean_inc(x_88); +x_90 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_90, 0, x_88); +lean_ctor_set(x_90, 1, x_88); +lean_ctor_set(x_90, 2, x_89); +lean_ctor_set(x_37, 0, x_90); return x_37; } } else { -lean_object* x_85; lean_object* x_86; +lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +lean_dec(x_74); +lean_dec(x_43); lean_dec(x_41); lean_dec(x_36); lean_dec(x_10); @@ -12514,79 +12637,114 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_1); -x_85 = l_Lean_LocalDecl_toExpr(x_65); -lean_dec(x_65); -x_86 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_86, 0, x_35); -lean_ctor_set(x_86, 1, x_43); -lean_ctor_set(x_86, 2, x_85); -lean_ctor_set(x_37, 0, x_86); +x_91 = lean_ctor_get(x_75, 0); +lean_inc(x_91); +lean_dec(x_75); +x_92 = lean_ctor_get(x_91, 0); +lean_inc(x_92); +x_93 = lean_ctor_get(x_91, 1); +lean_inc(x_93); +lean_dec(x_91); +x_94 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_94, 0, x_92); +lean_ctor_set(x_94, 1, x_35); +lean_ctor_set(x_94, 2, x_93); +lean_ctor_set(x_37, 0, x_94); +return x_37; +} +} +else +{ +lean_object* x_95; lean_object* x_96; +lean_dec(x_41); +lean_dec(x_36); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_95 = l_Lean_LocalDecl_toExpr(x_70); +lean_dec(x_70); +x_96 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_96, 0, x_35); +lean_ctor_set(x_96, 1, x_43); +lean_ctor_set(x_96, 2, x_95); +lean_ctor_set(x_37, 0, x_96); return x_37; } } } else { -lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_87 = lean_box(0); +lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_97 = lean_box(0); lean_inc(x_36); -x_88 = lean_name_mk_string(x_87, x_36); +x_98 = lean_name_mk_string(x_97, x_36); lean_inc(x_35); lean_inc(x_41); -x_89 = l_Lean_findField_x3f(x_41, x_35, x_88); -if (lean_obj_tag(x_89) == 0) +x_99 = l_Lean_findField_x3f(x_41, x_35, x_98); +if (lean_obj_tag(x_99) == 0) { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_inc(x_36); lean_inc(x_35); -x_90 = lean_name_mk_string(x_35, x_36); -x_91 = lean_ctor_get(x_9, 6); -lean_inc(x_91); -lean_inc(x_90); -x_92 = l_Lean_Name_replacePrefix(x_90, x_91, x_87); -lean_dec(x_91); -x_93 = lean_ctor_get(x_7, 1); -lean_inc(x_93); -x_94 = lean_local_ctx_find_from_user_name(x_93, x_92); -if (lean_obj_tag(x_94) == 0) +x_100 = lean_name_mk_string(x_35, x_36); +x_101 = lean_ctor_get(x_9, 6); +lean_inc(x_101); +lean_inc(x_100); +x_102 = l_Lean_Name_replacePrefix(x_100, x_101, x_97); +lean_dec(x_101); +x_103 = lean_ctor_get(x_7, 1); +lean_inc(x_103); +x_104 = lean_local_ctx_find_from_user_name(x_103, x_102); +if (lean_obj_tag(x_104) == 0) { -lean_object* x_95; +lean_object* x_105; +lean_inc(x_98); lean_inc(x_35); -x_95 = l___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f(x_41, x_35, x_88); -if (lean_obj_tag(x_95) == 0) +lean_inc(x_41); +x_105 = l___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f(x_41, x_35, x_98); +if (lean_obj_tag(x_105) == 0) { -lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; -lean_free_object(x_37); +lean_object* x_106; +x_106 = l___private_Lean_Elab_App_0__Lean_Elab_Term_findMethodAlias_x3f(x_41, x_35, x_98); lean_dec(x_35); -x_96 = l_Lean_stringToMessageData(x_36); +if (lean_obj_tag(x_106) == 0) +{ +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_free_object(x_37); +x_107 = l_Lean_stringToMessageData(x_36); lean_dec(x_36); -x_97 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__4; -x_98 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_98, 0, x_97); -lean_ctor_set(x_98, 1, x_96); -x_99 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__6; -x_100 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_100, 0, x_98); -lean_ctor_set(x_100, 1, x_99); -x_101 = lean_alloc_ctor(4, 1, 0); -lean_ctor_set(x_101, 0, x_90); -x_102 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_102, 0, x_100); -lean_ctor_set(x_102, 1, x_101); -x_103 = l_Lean_Elab_Term_throwInvalidNamedArg___rarg___closed__8; -x_104 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_104, 0, x_102); -lean_ctor_set(x_104, 1, x_103); -x_105 = l___private_Lean_Elab_App_0__Lean_Elab_Term_throwLValError___rarg(x_3, x_1, x_104, x_5, x_6, x_7, x_8, x_9, x_10, x_40); +x_108 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__4; +x_109 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_109, 0, x_108); +lean_ctor_set(x_109, 1, x_107); +x_110 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__6; +x_111 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_111, 0, x_109); +lean_ctor_set(x_111, 1, x_110); +x_112 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_112, 0, x_100); +x_113 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_113, 0, x_111); +lean_ctor_set(x_113, 1, x_112); +x_114 = l_Lean_Elab_Term_throwInvalidNamedArg___rarg___closed__8; +x_115 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_115, 0, x_113); +lean_ctor_set(x_115, 1, x_114); +x_116 = l___private_Lean_Elab_App_0__Lean_Elab_Term_throwLValError___rarg(x_3, x_1, x_115, x_5, x_6, x_7, x_8, x_9, x_10, x_40); lean_dec(x_10); lean_dec(x_8); lean_dec(x_6); -return x_105; +return x_116; } else { -lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; -lean_dec(x_90); +lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; +lean_dec(x_100); lean_dec(x_36); lean_dec(x_10); lean_dec(x_9); @@ -12596,71 +12754,109 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_1); -x_106 = lean_ctor_get(x_95, 0); -lean_inc(x_106); -lean_dec(x_95); -x_107 = lean_ctor_get(x_106, 0); -lean_inc(x_107); -x_108 = lean_ctor_get(x_106, 1); -lean_inc(x_108); +x_117 = lean_ctor_get(x_106, 0); +lean_inc(x_117); lean_dec(x_106); -x_109 = lean_alloc_ctor(2, 3, 0); -lean_ctor_set(x_109, 0, x_107); -lean_ctor_set(x_109, 1, x_35); -lean_ctor_set(x_109, 2, x_108); -lean_ctor_set(x_37, 0, x_109); +x_118 = lean_ctor_get(x_117, 0); +lean_inc(x_118); +x_119 = lean_ctor_get(x_117, 1); +lean_inc(x_119); +lean_dec(x_117); +lean_inc(x_118); +x_120 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_120, 0, x_118); +lean_ctor_set(x_120, 1, x_118); +lean_ctor_set(x_120, 2, x_119); +lean_ctor_set(x_37, 0, x_120); return x_37; } } else { -lean_object* x_110; uint8_t x_111; uint8_t x_112; uint8_t x_113; -x_110 = lean_ctor_get(x_94, 0); -lean_inc(x_110); -lean_dec(x_94); -x_111 = l_Lean_LocalDecl_binderInfo(x_110); -x_112 = 4; -x_113 = l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_362_(x_111, x_112); -if (x_113 == 0) +lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; +lean_dec(x_100); +lean_dec(x_98); +lean_dec(x_41); +lean_dec(x_36); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_121 = lean_ctor_get(x_105, 0); +lean_inc(x_121); +lean_dec(x_105); +x_122 = lean_ctor_get(x_121, 0); +lean_inc(x_122); +x_123 = lean_ctor_get(x_121, 1); +lean_inc(x_123); +lean_dec(x_121); +x_124 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_124, 0, x_122); +lean_ctor_set(x_124, 1, x_35); +lean_ctor_set(x_124, 2, x_123); +lean_ctor_set(x_37, 0, x_124); +return x_37; +} +} +else { -lean_object* x_114; -lean_dec(x_110); +lean_object* x_125; uint8_t x_126; uint8_t x_127; uint8_t x_128; +x_125 = lean_ctor_get(x_104, 0); +lean_inc(x_125); +lean_dec(x_104); +x_126 = l_Lean_LocalDecl_binderInfo(x_125); +x_127 = 4; +x_128 = l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_362_(x_126, x_127); +if (x_128 == 0) +{ +lean_object* x_129; +lean_dec(x_125); +lean_inc(x_98); lean_inc(x_35); -x_114 = l___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f(x_41, x_35, x_88); -if (lean_obj_tag(x_114) == 0) +lean_inc(x_41); +x_129 = l___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f(x_41, x_35, x_98); +if (lean_obj_tag(x_129) == 0) { -lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; -lean_free_object(x_37); +lean_object* x_130; +x_130 = l___private_Lean_Elab_App_0__Lean_Elab_Term_findMethodAlias_x3f(x_41, x_35, x_98); lean_dec(x_35); -x_115 = l_Lean_stringToMessageData(x_36); +if (lean_obj_tag(x_130) == 0) +{ +lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; +lean_free_object(x_37); +x_131 = l_Lean_stringToMessageData(x_36); lean_dec(x_36); -x_116 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__4; -x_117 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_117, 0, x_116); -lean_ctor_set(x_117, 1, x_115); -x_118 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__6; -x_119 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_119, 0, x_117); -lean_ctor_set(x_119, 1, x_118); -x_120 = lean_alloc_ctor(4, 1, 0); -lean_ctor_set(x_120, 0, x_90); -x_121 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_121, 0, x_119); -lean_ctor_set(x_121, 1, x_120); -x_122 = l_Lean_Elab_Term_throwInvalidNamedArg___rarg___closed__8; -x_123 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_123, 0, x_121); -lean_ctor_set(x_123, 1, x_122); -x_124 = l___private_Lean_Elab_App_0__Lean_Elab_Term_throwLValError___rarg(x_3, x_1, x_123, x_5, x_6, x_7, x_8, x_9, x_10, x_40); +x_132 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__4; +x_133 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_133, 0, x_132); +lean_ctor_set(x_133, 1, x_131); +x_134 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__6; +x_135 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_135, 0, x_133); +lean_ctor_set(x_135, 1, x_134); +x_136 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_136, 0, x_100); +x_137 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_137, 0, x_135); +lean_ctor_set(x_137, 1, x_136); +x_138 = l_Lean_Elab_Term_throwInvalidNamedArg___rarg___closed__8; +x_139 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_139, 0, x_137); +lean_ctor_set(x_139, 1, x_138); +x_140 = l___private_Lean_Elab_App_0__Lean_Elab_Term_throwLValError___rarg(x_3, x_1, x_139, x_5, x_6, x_7, x_8, x_9, x_10, x_40); lean_dec(x_10); lean_dec(x_8); lean_dec(x_6); -return x_124; +return x_140; } else { -lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; -lean_dec(x_90); +lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; +lean_dec(x_100); lean_dec(x_36); lean_dec(x_10); lean_dec(x_9); @@ -12670,26 +12866,28 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_1); -x_125 = lean_ctor_get(x_114, 0); -lean_inc(x_125); -lean_dec(x_114); -x_126 = lean_ctor_get(x_125, 0); -lean_inc(x_126); -x_127 = lean_ctor_get(x_125, 1); -lean_inc(x_127); -lean_dec(x_125); -x_128 = lean_alloc_ctor(2, 3, 0); -lean_ctor_set(x_128, 0, x_126); -lean_ctor_set(x_128, 1, x_35); -lean_ctor_set(x_128, 2, x_127); -lean_ctor_set(x_37, 0, x_128); +x_141 = lean_ctor_get(x_130, 0); +lean_inc(x_141); +lean_dec(x_130); +x_142 = lean_ctor_get(x_141, 0); +lean_inc(x_142); +x_143 = lean_ctor_get(x_141, 1); +lean_inc(x_143); +lean_dec(x_141); +lean_inc(x_142); +x_144 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_144, 0, x_142); +lean_ctor_set(x_144, 1, x_142); +lean_ctor_set(x_144, 2, x_143); +lean_ctor_set(x_37, 0, x_144); return x_37; } } else { -lean_object* x_129; lean_object* x_130; -lean_dec(x_88); +lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; +lean_dec(x_100); +lean_dec(x_98); lean_dec(x_41); lean_dec(x_36); lean_dec(x_10); @@ -12700,20 +12898,26 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_1); -x_129 = l_Lean_LocalDecl_toExpr(x_110); -lean_dec(x_110); -x_130 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_130, 0, x_35); -lean_ctor_set(x_130, 1, x_90); -lean_ctor_set(x_130, 2, x_129); -lean_ctor_set(x_37, 0, x_130); +x_145 = lean_ctor_get(x_129, 0); +lean_inc(x_145); +lean_dec(x_129); +x_146 = lean_ctor_get(x_145, 0); +lean_inc(x_146); +x_147 = lean_ctor_get(x_145, 1); +lean_inc(x_147); +lean_dec(x_145); +x_148 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_148, 0, x_146); +lean_ctor_set(x_148, 1, x_35); +lean_ctor_set(x_148, 2, x_147); +lean_ctor_set(x_37, 0, x_148); return x_37; } } -} else { -lean_object* x_131; lean_object* x_132; +lean_object* x_149; lean_object* x_150; +lean_dec(x_98); lean_dec(x_41); lean_dec(x_36); lean_dec(x_10); @@ -12724,87 +12928,21 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_1); -x_131 = lean_ctor_get(x_89, 0); -lean_inc(x_131); -lean_dec(x_89); -x_132 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_132, 0, x_131); -lean_ctor_set(x_132, 1, x_35); -lean_ctor_set(x_132, 2, x_88); -lean_ctor_set(x_37, 0, x_132); +x_149 = l_Lean_LocalDecl_toExpr(x_125); +lean_dec(x_125); +x_150 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_150, 0, x_35); +lean_ctor_set(x_150, 1, x_100); +lean_ctor_set(x_150, 2, x_149); +lean_ctor_set(x_37, 0, x_150); return x_37; } } } else { -lean_object* x_133; lean_object* x_134; lean_object* x_135; uint8_t x_136; -x_133 = lean_ctor_get(x_37, 0); -x_134 = lean_ctor_get(x_37, 1); -lean_inc(x_134); -lean_inc(x_133); -lean_dec(x_37); -x_135 = lean_ctor_get(x_133, 0); -lean_inc(x_135); -lean_dec(x_133); -lean_inc(x_35); -lean_inc(x_135); -x_136 = l_Lean_isStructure(x_135, x_35); -if (x_136 == 0) -{ -lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; -lean_inc(x_36); -lean_inc(x_35); -x_137 = lean_name_mk_string(x_35, x_36); -x_138 = lean_ctor_get(x_9, 6); -lean_inc(x_138); -x_139 = lean_box(0); -lean_inc(x_137); -x_140 = l_Lean_Name_replacePrefix(x_137, x_138, x_139); -lean_dec(x_138); -x_141 = lean_ctor_get(x_7, 1); -lean_inc(x_141); -x_142 = lean_local_ctx_find_from_user_name(x_141, x_140); -if (lean_obj_tag(x_142) == 0) -{ -lean_object* x_143; lean_object* x_144; -lean_inc(x_36); -x_143 = lean_name_mk_string(x_139, x_36); -lean_inc(x_35); -x_144 = l___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f(x_135, x_35, x_143); -if (lean_obj_tag(x_144) == 0) -{ -lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; -lean_dec(x_35); -x_145 = l_Lean_stringToMessageData(x_36); -lean_dec(x_36); -x_146 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__4; -x_147 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_147, 0, x_146); -lean_ctor_set(x_147, 1, x_145); -x_148 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__6; -x_149 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_149, 0, x_147); -lean_ctor_set(x_149, 1, x_148); -x_150 = lean_alloc_ctor(4, 1, 0); -lean_ctor_set(x_150, 0, x_137); -x_151 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_151, 0, x_149); -lean_ctor_set(x_151, 1, x_150); -x_152 = l_Lean_Elab_Term_throwInvalidNamedArg___rarg___closed__8; -x_153 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_153, 0, x_151); -lean_ctor_set(x_153, 1, x_152); -x_154 = l___private_Lean_Elab_App_0__Lean_Elab_Term_throwLValError___rarg(x_3, x_1, x_153, x_5, x_6, x_7, x_8, x_9, x_10, x_134); -lean_dec(x_10); -lean_dec(x_8); -lean_dec(x_6); -return x_154; -} -else -{ -lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; -lean_dec(x_137); +lean_object* x_151; lean_object* x_152; +lean_dec(x_41); lean_dec(x_36); lean_dec(x_10); lean_dec(x_9); @@ -12814,45 +12952,64 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_1); -x_155 = lean_ctor_get(x_144, 0); -lean_inc(x_155); -lean_dec(x_144); -x_156 = lean_ctor_get(x_155, 0); -lean_inc(x_156); -x_157 = lean_ctor_get(x_155, 1); -lean_inc(x_157); -lean_dec(x_155); -x_158 = lean_alloc_ctor(2, 3, 0); -lean_ctor_set(x_158, 0, x_156); -lean_ctor_set(x_158, 1, x_35); -lean_ctor_set(x_158, 2, x_157); -x_159 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_159, 0, x_158); -lean_ctor_set(x_159, 1, x_134); -return x_159; +x_151 = lean_ctor_get(x_99, 0); +lean_inc(x_151); +lean_dec(x_99); +x_152 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_152, 0, x_151); +lean_ctor_set(x_152, 1, x_35); +lean_ctor_set(x_152, 2, x_98); +lean_ctor_set(x_37, 0, x_152); +return x_37; +} } } else { -lean_object* x_160; uint8_t x_161; uint8_t x_162; uint8_t x_163; -x_160 = lean_ctor_get(x_142, 0); -lean_inc(x_160); -lean_dec(x_142); -x_161 = l_Lean_LocalDecl_binderInfo(x_160); -x_162 = 4; -x_163 = l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_362_(x_161, x_162); -if (x_163 == 0) +lean_object* x_153; lean_object* x_154; lean_object* x_155; uint8_t x_156; +x_153 = lean_ctor_get(x_37, 0); +x_154 = lean_ctor_get(x_37, 1); +lean_inc(x_154); +lean_inc(x_153); +lean_dec(x_37); +x_155 = lean_ctor_get(x_153, 0); +lean_inc(x_155); +lean_dec(x_153); +lean_inc(x_35); +lean_inc(x_155); +x_156 = l_Lean_isStructure(x_155, x_35); +if (x_156 == 0) { -lean_object* x_164; lean_object* x_165; -lean_dec(x_160); +lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; +lean_inc(x_36); +lean_inc(x_35); +x_157 = lean_name_mk_string(x_35, x_36); +x_158 = lean_ctor_get(x_9, 6); +lean_inc(x_158); +x_159 = lean_box(0); +lean_inc(x_157); +x_160 = l_Lean_Name_replacePrefix(x_157, x_158, x_159); +lean_dec(x_158); +x_161 = lean_ctor_get(x_7, 1); +lean_inc(x_161); +x_162 = lean_local_ctx_find_from_user_name(x_161, x_160); +if (lean_obj_tag(x_162) == 0) +{ +lean_object* x_163; lean_object* x_164; lean_inc(x_36); -x_164 = lean_name_mk_string(x_139, x_36); +x_163 = lean_name_mk_string(x_159, x_36); +lean_inc(x_163); lean_inc(x_35); -x_165 = l___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f(x_135, x_35, x_164); +lean_inc(x_155); +x_164 = l___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f(x_155, x_35, x_163); +if (lean_obj_tag(x_164) == 0) +{ +lean_object* x_165; +x_165 = l___private_Lean_Elab_App_0__Lean_Elab_Term_findMethodAlias_x3f(x_155, x_35, x_163); +lean_dec(x_35); if (lean_obj_tag(x_165) == 0) { lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; -lean_dec(x_35); x_166 = l_Lean_stringToMessageData(x_36); lean_dec(x_36); x_167 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__4; @@ -12864,7 +13021,7 @@ x_170 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_170, 0, x_168); lean_ctor_set(x_170, 1, x_169); x_171 = lean_alloc_ctor(4, 1, 0); -lean_ctor_set(x_171, 0, x_137); +lean_ctor_set(x_171, 0, x_157); x_172 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_172, 0, x_170); lean_ctor_set(x_172, 1, x_171); @@ -12872,7 +13029,7 @@ x_173 = l_Lean_Elab_Term_throwInvalidNamedArg___rarg___closed__8; x_174 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_174, 0, x_172); lean_ctor_set(x_174, 1, x_173); -x_175 = l___private_Lean_Elab_App_0__Lean_Elab_Term_throwLValError___rarg(x_3, x_1, x_174, x_5, x_6, x_7, x_8, x_9, x_10, x_134); +x_175 = l___private_Lean_Elab_App_0__Lean_Elab_Term_throwLValError___rarg(x_3, x_1, x_174, x_5, x_6, x_7, x_8, x_9, x_10, x_154); lean_dec(x_10); lean_dec(x_8); lean_dec(x_6); @@ -12881,7 +13038,7 @@ return x_175; else { lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; -lean_dec(x_137); +lean_dec(x_157); lean_dec(x_36); lean_dec(x_10); lean_dec(x_9); @@ -12899,20 +13056,23 @@ lean_inc(x_177); x_178 = lean_ctor_get(x_176, 1); lean_inc(x_178); lean_dec(x_176); +lean_inc(x_177); x_179 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_179, 0, x_177); -lean_ctor_set(x_179, 1, x_35); +lean_ctor_set(x_179, 1, x_177); lean_ctor_set(x_179, 2, x_178); x_180 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_180, 0, x_179); -lean_ctor_set(x_180, 1, x_134); +lean_ctor_set(x_180, 1, x_154); return x_180; } } else { -lean_object* x_181; lean_object* x_182; lean_object* x_183; -lean_dec(x_135); +lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; +lean_dec(x_163); +lean_dec(x_157); +lean_dec(x_155); lean_dec(x_36); lean_dec(x_10); lean_dec(x_9); @@ -12922,51 +13082,51 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_1); -x_181 = l_Lean_LocalDecl_toExpr(x_160); -lean_dec(x_160); -x_182 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_182, 0, x_35); -lean_ctor_set(x_182, 1, x_137); -lean_ctor_set(x_182, 2, x_181); -x_183 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_183, 0, x_182); -lean_ctor_set(x_183, 1, x_134); -return x_183; -} +x_181 = lean_ctor_get(x_164, 0); +lean_inc(x_181); +lean_dec(x_164); +x_182 = lean_ctor_get(x_181, 0); +lean_inc(x_182); +x_183 = lean_ctor_get(x_181, 1); +lean_inc(x_183); +lean_dec(x_181); +x_184 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_184, 0, x_182); +lean_ctor_set(x_184, 1, x_35); +lean_ctor_set(x_184, 2, x_183); +x_185 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_185, 0, x_184); +lean_ctor_set(x_185, 1, x_154); +return x_185; } } else { -lean_object* x_184; lean_object* x_185; lean_object* x_186; -x_184 = lean_box(0); -lean_inc(x_36); -x_185 = lean_name_mk_string(x_184, x_36); -lean_inc(x_35); -lean_inc(x_135); -x_186 = l_Lean_findField_x3f(x_135, x_35, x_185); -if (lean_obj_tag(x_186) == 0) +lean_object* x_186; uint8_t x_187; uint8_t x_188; uint8_t x_189; +x_186 = lean_ctor_get(x_162, 0); +lean_inc(x_186); +lean_dec(x_162); +x_187 = l_Lean_LocalDecl_binderInfo(x_186); +x_188 = 4; +x_189 = l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_362_(x_187, x_188); +if (x_189 == 0) { -lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; +lean_object* x_190; lean_object* x_191; +lean_dec(x_186); lean_inc(x_36); -lean_inc(x_35); -x_187 = lean_name_mk_string(x_35, x_36); -x_188 = lean_ctor_get(x_9, 6); -lean_inc(x_188); -lean_inc(x_187); -x_189 = l_Lean_Name_replacePrefix(x_187, x_188, x_184); -lean_dec(x_188); -x_190 = lean_ctor_get(x_7, 1); +x_190 = lean_name_mk_string(x_159, x_36); lean_inc(x_190); -x_191 = lean_local_ctx_find_from_user_name(x_190, x_189); +lean_inc(x_35); +lean_inc(x_155); +x_191 = l___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f(x_155, x_35, x_190); if (lean_obj_tag(x_191) == 0) { lean_object* x_192; -lean_inc(x_35); -x_192 = l___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f(x_135, x_35, x_185); +x_192 = l___private_Lean_Elab_App_0__Lean_Elab_Term_findMethodAlias_x3f(x_155, x_35, x_190); +lean_dec(x_35); if (lean_obj_tag(x_192) == 0) { lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; -lean_dec(x_35); x_193 = l_Lean_stringToMessageData(x_36); lean_dec(x_36); x_194 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__4; @@ -12978,7 +13138,7 @@ x_197 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_197, 0, x_195); lean_ctor_set(x_197, 1, x_196); x_198 = lean_alloc_ctor(4, 1, 0); -lean_ctor_set(x_198, 0, x_187); +lean_ctor_set(x_198, 0, x_157); x_199 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_199, 0, x_197); lean_ctor_set(x_199, 1, x_198); @@ -12986,7 +13146,7 @@ x_200 = l_Lean_Elab_Term_throwInvalidNamedArg___rarg___closed__8; x_201 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_201, 0, x_199); lean_ctor_set(x_201, 1, x_200); -x_202 = l___private_Lean_Elab_App_0__Lean_Elab_Term_throwLValError___rarg(x_3, x_1, x_201, x_5, x_6, x_7, x_8, x_9, x_10, x_134); +x_202 = l___private_Lean_Elab_App_0__Lean_Elab_Term_throwLValError___rarg(x_3, x_1, x_201, x_5, x_6, x_7, x_8, x_9, x_10, x_154); lean_dec(x_10); lean_dec(x_8); lean_dec(x_6); @@ -12995,7 +13155,7 @@ return x_202; else { lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; -lean_dec(x_187); +lean_dec(x_157); lean_dec(x_36); lean_dec(x_10); lean_dec(x_9); @@ -13013,64 +13173,292 @@ lean_inc(x_204); x_205 = lean_ctor_get(x_203, 1); lean_inc(x_205); lean_dec(x_203); +lean_inc(x_204); x_206 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_206, 0, x_204); -lean_ctor_set(x_206, 1, x_35); +lean_ctor_set(x_206, 1, x_204); lean_ctor_set(x_206, 2, x_205); x_207 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_207, 0, x_206); -lean_ctor_set(x_207, 1, x_134); +lean_ctor_set(x_207, 1, x_154); return x_207; } } else { -lean_object* x_208; uint8_t x_209; uint8_t x_210; uint8_t x_211; +lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; +lean_dec(x_190); +lean_dec(x_157); +lean_dec(x_155); +lean_dec(x_36); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); x_208 = lean_ctor_get(x_191, 0); lean_inc(x_208); lean_dec(x_191); -x_209 = l_Lean_LocalDecl_binderInfo(x_208); -x_210 = 4; -x_211 = l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_362_(x_209, x_210); -if (x_211 == 0) -{ -lean_object* x_212; +x_209 = lean_ctor_get(x_208, 0); +lean_inc(x_209); +x_210 = lean_ctor_get(x_208, 1); +lean_inc(x_210); lean_dec(x_208); +x_211 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_211, 0, x_209); +lean_ctor_set(x_211, 1, x_35); +lean_ctor_set(x_211, 2, x_210); +x_212 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_212, 0, x_211); +lean_ctor_set(x_212, 1, x_154); +return x_212; +} +} +else +{ +lean_object* x_213; lean_object* x_214; lean_object* x_215; +lean_dec(x_155); +lean_dec(x_36); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_213 = l_Lean_LocalDecl_toExpr(x_186); +lean_dec(x_186); +x_214 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_214, 0, x_35); +lean_ctor_set(x_214, 1, x_157); +lean_ctor_set(x_214, 2, x_213); +x_215 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_215, 0, x_214); +lean_ctor_set(x_215, 1, x_154); +return x_215; +} +} +} +else +{ +lean_object* x_216; lean_object* x_217; lean_object* x_218; +x_216 = lean_box(0); +lean_inc(x_36); +x_217 = lean_name_mk_string(x_216, x_36); +lean_inc(x_35); +lean_inc(x_155); +x_218 = l_Lean_findField_x3f(x_155, x_35, x_217); +if (lean_obj_tag(x_218) == 0) +{ +lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; +lean_inc(x_36); +lean_inc(x_35); +x_219 = lean_name_mk_string(x_35, x_36); +x_220 = lean_ctor_get(x_9, 6); +lean_inc(x_220); +lean_inc(x_219); +x_221 = l_Lean_Name_replacePrefix(x_219, x_220, x_216); +lean_dec(x_220); +x_222 = lean_ctor_get(x_7, 1); +lean_inc(x_222); +x_223 = lean_local_ctx_find_from_user_name(x_222, x_221); +if (lean_obj_tag(x_223) == 0) +{ +lean_object* x_224; +lean_inc(x_217); lean_inc(x_35); -x_212 = l___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f(x_135, x_35, x_185); -if (lean_obj_tag(x_212) == 0) +lean_inc(x_155); +x_224 = l___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f(x_155, x_35, x_217); +if (lean_obj_tag(x_224) == 0) { -lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; +lean_object* x_225; +x_225 = l___private_Lean_Elab_App_0__Lean_Elab_Term_findMethodAlias_x3f(x_155, x_35, x_217); lean_dec(x_35); -x_213 = l_Lean_stringToMessageData(x_36); +if (lean_obj_tag(x_225) == 0) +{ +lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; +x_226 = l_Lean_stringToMessageData(x_36); lean_dec(x_36); -x_214 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__4; -x_215 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_215, 0, x_214); -lean_ctor_set(x_215, 1, x_213); -x_216 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__6; -x_217 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_217, 0, x_215); -lean_ctor_set(x_217, 1, x_216); -x_218 = lean_alloc_ctor(4, 1, 0); -lean_ctor_set(x_218, 0, x_187); -x_219 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_219, 0, x_217); -lean_ctor_set(x_219, 1, x_218); -x_220 = l_Lean_Elab_Term_throwInvalidNamedArg___rarg___closed__8; -x_221 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_221, 0, x_219); -lean_ctor_set(x_221, 1, x_220); -x_222 = l___private_Lean_Elab_App_0__Lean_Elab_Term_throwLValError___rarg(x_3, x_1, x_221, x_5, x_6, x_7, x_8, x_9, x_10, x_134); +x_227 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__4; +x_228 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_228, 0, x_227); +lean_ctor_set(x_228, 1, x_226); +x_229 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__6; +x_230 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_230, 0, x_228); +lean_ctor_set(x_230, 1, x_229); +x_231 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_231, 0, x_219); +x_232 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_232, 0, x_230); +lean_ctor_set(x_232, 1, x_231); +x_233 = l_Lean_Elab_Term_throwInvalidNamedArg___rarg___closed__8; +x_234 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_234, 0, x_232); +lean_ctor_set(x_234, 1, x_233); +x_235 = l___private_Lean_Elab_App_0__Lean_Elab_Term_throwLValError___rarg(x_3, x_1, x_234, x_5, x_6, x_7, x_8, x_9, x_10, x_154); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_6); +return x_235; +} +else +{ +lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; +lean_dec(x_219); +lean_dec(x_36); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_236 = lean_ctor_get(x_225, 0); +lean_inc(x_236); +lean_dec(x_225); +x_237 = lean_ctor_get(x_236, 0); +lean_inc(x_237); +x_238 = lean_ctor_get(x_236, 1); +lean_inc(x_238); +lean_dec(x_236); +lean_inc(x_237); +x_239 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_239, 0, x_237); +lean_ctor_set(x_239, 1, x_237); +lean_ctor_set(x_239, 2, x_238); +x_240 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_240, 0, x_239); +lean_ctor_set(x_240, 1, x_154); +return x_240; +} +} +else +{ +lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; +lean_dec(x_219); +lean_dec(x_217); +lean_dec(x_155); +lean_dec(x_36); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_241 = lean_ctor_get(x_224, 0); +lean_inc(x_241); +lean_dec(x_224); +x_242 = lean_ctor_get(x_241, 0); +lean_inc(x_242); +x_243 = lean_ctor_get(x_241, 1); +lean_inc(x_243); +lean_dec(x_241); +x_244 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_244, 0, x_242); +lean_ctor_set(x_244, 1, x_35); +lean_ctor_set(x_244, 2, x_243); +x_245 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_245, 0, x_244); +lean_ctor_set(x_245, 1, x_154); +return x_245; +} +} +else +{ +lean_object* x_246; uint8_t x_247; uint8_t x_248; uint8_t x_249; +x_246 = lean_ctor_get(x_223, 0); +lean_inc(x_246); +lean_dec(x_223); +x_247 = l_Lean_LocalDecl_binderInfo(x_246); +x_248 = 4; +x_249 = l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_362_(x_247, x_248); +if (x_249 == 0) +{ +lean_object* x_250; +lean_dec(x_246); +lean_inc(x_217); +lean_inc(x_35); +lean_inc(x_155); +x_250 = l___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f(x_155, x_35, x_217); +if (lean_obj_tag(x_250) == 0) +{ +lean_object* x_251; +x_251 = l___private_Lean_Elab_App_0__Lean_Elab_Term_findMethodAlias_x3f(x_155, x_35, x_217); +lean_dec(x_35); +if (lean_obj_tag(x_251) == 0) +{ +lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; +x_252 = l_Lean_stringToMessageData(x_36); +lean_dec(x_36); +x_253 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__4; +x_254 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_254, 0, x_253); +lean_ctor_set(x_254, 1, x_252); +x_255 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__6; +x_256 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_256, 0, x_254); +lean_ctor_set(x_256, 1, x_255); +x_257 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_257, 0, x_219); +x_258 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_258, 0, x_256); +lean_ctor_set(x_258, 1, x_257); +x_259 = l_Lean_Elab_Term_throwInvalidNamedArg___rarg___closed__8; +x_260 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_260, 0, x_258); +lean_ctor_set(x_260, 1, x_259); +x_261 = l___private_Lean_Elab_App_0__Lean_Elab_Term_throwLValError___rarg(x_3, x_1, x_260, x_5, x_6, x_7, x_8, x_9, x_10, x_154); lean_dec(x_10); lean_dec(x_8); lean_dec(x_6); -return x_222; +return x_261; +} +else +{ +lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; +lean_dec(x_219); +lean_dec(x_36); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_262 = lean_ctor_get(x_251, 0); +lean_inc(x_262); +lean_dec(x_251); +x_263 = lean_ctor_get(x_262, 0); +lean_inc(x_263); +x_264 = lean_ctor_get(x_262, 1); +lean_inc(x_264); +lean_dec(x_262); +lean_inc(x_263); +x_265 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_265, 0, x_263); +lean_ctor_set(x_265, 1, x_263); +lean_ctor_set(x_265, 2, x_264); +x_266 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_266, 0, x_265); +lean_ctor_set(x_266, 1, x_154); +return x_266; +} } else { -lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; -lean_dec(x_187); +lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; +lean_dec(x_219); +lean_dec(x_217); +lean_dec(x_155); lean_dec(x_36); lean_dec(x_10); lean_dec(x_9); @@ -13080,29 +13468,29 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_1); -x_223 = lean_ctor_get(x_212, 0); -lean_inc(x_223); -lean_dec(x_212); -x_224 = lean_ctor_get(x_223, 0); -lean_inc(x_224); -x_225 = lean_ctor_get(x_223, 1); -lean_inc(x_225); -lean_dec(x_223); -x_226 = lean_alloc_ctor(2, 3, 0); -lean_ctor_set(x_226, 0, x_224); -lean_ctor_set(x_226, 1, x_35); -lean_ctor_set(x_226, 2, x_225); -x_227 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_227, 0, x_226); -lean_ctor_set(x_227, 1, x_134); -return x_227; +x_267 = lean_ctor_get(x_250, 0); +lean_inc(x_267); +lean_dec(x_250); +x_268 = lean_ctor_get(x_267, 0); +lean_inc(x_268); +x_269 = lean_ctor_get(x_267, 1); +lean_inc(x_269); +lean_dec(x_267); +x_270 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_270, 0, x_268); +lean_ctor_set(x_270, 1, x_35); +lean_ctor_set(x_270, 2, x_269); +x_271 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_271, 0, x_270); +lean_ctor_set(x_271, 1, x_154); +return x_271; } } else { -lean_object* x_228; lean_object* x_229; lean_object* x_230; -lean_dec(x_185); -lean_dec(x_135); +lean_object* x_272; lean_object* x_273; lean_object* x_274; +lean_dec(x_217); +lean_dec(x_155); lean_dec(x_36); lean_dec(x_10); lean_dec(x_9); @@ -13112,23 +13500,23 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_1); -x_228 = l_Lean_LocalDecl_toExpr(x_208); -lean_dec(x_208); -x_229 = lean_alloc_ctor(3, 3, 0); -lean_ctor_set(x_229, 0, x_35); -lean_ctor_set(x_229, 1, x_187); -lean_ctor_set(x_229, 2, x_228); -x_230 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_230, 0, x_229); -lean_ctor_set(x_230, 1, x_134); -return x_230; +x_272 = l_Lean_LocalDecl_toExpr(x_246); +lean_dec(x_246); +x_273 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_273, 0, x_35); +lean_ctor_set(x_273, 1, x_219); +lean_ctor_set(x_273, 2, x_272); +x_274 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_274, 0, x_273); +lean_ctor_set(x_274, 1, x_154); +return x_274; } } } else { -lean_object* x_231; lean_object* x_232; lean_object* x_233; -lean_dec(x_135); +lean_object* x_275; lean_object* x_276; lean_object* x_277; +lean_dec(x_155); lean_dec(x_36); lean_dec(x_10); lean_dec(x_9); @@ -13138,69 +13526,69 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_1); -x_231 = lean_ctor_get(x_186, 0); -lean_inc(x_231); -lean_dec(x_186); -x_232 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_232, 0, x_231); -lean_ctor_set(x_232, 1, x_35); -lean_ctor_set(x_232, 2, x_185); -x_233 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_233, 0, x_232); -lean_ctor_set(x_233, 1, x_134); -return x_233; +x_275 = lean_ctor_get(x_218, 0); +lean_inc(x_275); +lean_dec(x_218); +x_276 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_276, 0, x_275); +lean_ctor_set(x_276, 1, x_35); +lean_ctor_set(x_276, 2, x_217); +x_277 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_277, 0, x_276); +lean_ctor_set(x_277, 1, x_154); +return x_277; } } } } default: { -lean_object* x_234; lean_object* x_235; lean_object* x_236; uint8_t x_237; -x_234 = lean_ctor_get(x_13, 0); -lean_inc(x_234); +lean_object* x_278; lean_object* x_279; lean_object* x_280; uint8_t x_281; +x_278 = lean_ctor_get(x_13, 0); +lean_inc(x_278); lean_dec(x_13); -x_235 = lean_ctor_get(x_2, 1); -lean_inc(x_235); +x_279 = lean_ctor_get(x_2, 1); +lean_inc(x_279); lean_dec(x_2); -x_236 = lean_st_ref_get(x_10, x_11); -x_237 = !lean_is_exclusive(x_236); -if (x_237 == 0) -{ -lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; -x_238 = lean_ctor_get(x_236, 0); -x_239 = lean_ctor_get(x_236, 1); -x_240 = lean_ctor_get(x_238, 0); -lean_inc(x_240); -lean_dec(x_238); -x_241 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__7; -x_242 = lean_name_mk_string(x_234, x_241); -lean_inc(x_242); -x_243 = lean_environment_find(x_240, x_242); -if (lean_obj_tag(x_243) == 0) -{ -lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; -lean_free_object(x_236); -lean_dec(x_235); -x_244 = lean_alloc_ctor(4, 1, 0); -lean_ctor_set(x_244, 0, x_242); -x_245 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__9; -x_246 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_246, 0, x_245); -lean_ctor_set(x_246, 1, x_244); -x_247 = l_Lean_Elab_Term_throwInvalidNamedArg___rarg___closed__8; -x_248 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_248, 0, x_246); -lean_ctor_set(x_248, 1, x_247); -x_249 = l___private_Lean_Elab_App_0__Lean_Elab_Term_throwLValError___rarg(x_3, x_1, x_248, x_5, x_6, x_7, x_8, x_9, x_10, x_239); +x_280 = lean_st_ref_get(x_10, x_11); +x_281 = !lean_is_exclusive(x_280); +if (x_281 == 0) +{ +lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; +x_282 = lean_ctor_get(x_280, 0); +x_283 = lean_ctor_get(x_280, 1); +x_284 = lean_ctor_get(x_282, 0); +lean_inc(x_284); +lean_dec(x_282); +x_285 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__7; +x_286 = lean_name_mk_string(x_278, x_285); +lean_inc(x_286); +x_287 = lean_environment_find(x_284, x_286); +if (lean_obj_tag(x_287) == 0) +{ +lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; +lean_free_object(x_280); +lean_dec(x_279); +x_288 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_288, 0, x_286); +x_289 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__9; +x_290 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_290, 0, x_289); +lean_ctor_set(x_290, 1, x_288); +x_291 = l_Lean_Elab_Term_throwInvalidNamedArg___rarg___closed__8; +x_292 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_292, 0, x_290); +lean_ctor_set(x_292, 1, x_291); +x_293 = l___private_Lean_Elab_App_0__Lean_Elab_Term_throwLValError___rarg(x_3, x_1, x_292, x_5, x_6, x_7, x_8, x_9, x_10, x_283); lean_dec(x_10); lean_dec(x_8); lean_dec(x_6); -return x_249; +return x_293; } else { -lean_object* x_250; -lean_dec(x_243); +lean_object* x_294; +lean_dec(x_287); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -13209,52 +13597,52 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_1); -x_250 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_250, 0, x_242); -lean_ctor_set(x_250, 1, x_235); -lean_ctor_set(x_236, 0, x_250); -return x_236; +x_294 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_294, 0, x_286); +lean_ctor_set(x_294, 1, x_279); +lean_ctor_set(x_280, 0, x_294); +return x_280; } } else { -lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; -x_251 = lean_ctor_get(x_236, 0); -x_252 = lean_ctor_get(x_236, 1); -lean_inc(x_252); -lean_inc(x_251); -lean_dec(x_236); -x_253 = lean_ctor_get(x_251, 0); -lean_inc(x_253); -lean_dec(x_251); -x_254 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__7; -x_255 = lean_name_mk_string(x_234, x_254); -lean_inc(x_255); -x_256 = lean_environment_find(x_253, x_255); -if (lean_obj_tag(x_256) == 0) -{ -lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; -lean_dec(x_235); -x_257 = lean_alloc_ctor(4, 1, 0); -lean_ctor_set(x_257, 0, x_255); -x_258 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__9; -x_259 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_259, 0, x_258); -lean_ctor_set(x_259, 1, x_257); -x_260 = l_Lean_Elab_Term_throwInvalidNamedArg___rarg___closed__8; -x_261 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_261, 0, x_259); -lean_ctor_set(x_261, 1, x_260); -x_262 = l___private_Lean_Elab_App_0__Lean_Elab_Term_throwLValError___rarg(x_3, x_1, x_261, x_5, x_6, x_7, x_8, x_9, x_10, x_252); +lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; +x_295 = lean_ctor_get(x_280, 0); +x_296 = lean_ctor_get(x_280, 1); +lean_inc(x_296); +lean_inc(x_295); +lean_dec(x_280); +x_297 = lean_ctor_get(x_295, 0); +lean_inc(x_297); +lean_dec(x_295); +x_298 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__7; +x_299 = lean_name_mk_string(x_278, x_298); +lean_inc(x_299); +x_300 = lean_environment_find(x_297, x_299); +if (lean_obj_tag(x_300) == 0) +{ +lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; +lean_dec(x_279); +x_301 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_301, 0, x_299); +x_302 = l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__9; +x_303 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_303, 0, x_302); +lean_ctor_set(x_303, 1, x_301); +x_304 = l_Lean_Elab_Term_throwInvalidNamedArg___rarg___closed__8; +x_305 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_305, 0, x_303); +lean_ctor_set(x_305, 1, x_304); +x_306 = l___private_Lean_Elab_App_0__Lean_Elab_Term_throwLValError___rarg(x_3, x_1, x_305, x_5, x_6, x_7, x_8, x_9, x_10, x_296); lean_dec(x_10); lean_dec(x_8); lean_dec(x_6); -return x_262; +return x_306; } else { -lean_object* x_263; lean_object* x_264; -lean_dec(x_256); +lean_object* x_307; lean_object* x_308; +lean_dec(x_300); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -13263,13 +13651,13 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_1); -x_263 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_263, 0, x_255); -lean_ctor_set(x_263, 1, x_235); -x_264 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_264, 0, x_263); -lean_ctor_set(x_264, 1, x_252); -return x_264; +x_307 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_307, 0, x_299); +lean_ctor_set(x_307, 1, x_279); +x_308 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_308, 0, x_307); +lean_ctor_set(x_308, 1, x_296); +return x_308; } } } @@ -14739,6 +15127,15 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } +static lean_object* _init_l_List_forIn_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_mkBaseProjections___spec__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} LEAN_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_mkBaseProjections___spec__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, lean_object* x_9) { _start: { @@ -14787,7 +15184,7 @@ x_20 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_20, 0, x_18); lean_ctor_set(x_20, 1, x_19); lean_ctor_set(x_20, 2, x_17); -x_21 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__8; +x_21 = l_List_forIn_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_mkBaseProjections___spec__1___closed__3; x_22 = lean_array_push(x_21, x_20); x_23 = lean_box(0); x_24 = l_Lean_Elab_Term_ElabAppArgs_State_etaArgs___default___closed__1; @@ -16785,7 +17182,7 @@ x_34 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_34, 0, x_32); lean_ctor_set(x_34, 1, x_33); lean_ctor_set(x_34, 2, x_31); -x_35 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__8; +x_35 = l_List_forIn_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_mkBaseProjections___spec__1___closed__3; x_36 = lean_array_push(x_35, x_34); x_37 = l_Lean_Elab_Term_ElabAppArgs_State_etaArgs___default___closed__1; lean_inc(x_16); @@ -17018,7 +17415,7 @@ lean_dec(x_9); lean_dec(x_1); x_30 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_30, 0, x_10); -x_31 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__8; +x_31 = l_List_forIn_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_mkBaseProjections___spec__1___closed__3; x_32 = lean_array_push(x_31, x_30); x_33 = l_Lean_Elab_Term_ElabAppArgs_State_etaArgs___default___closed__1; lean_inc(x_16); @@ -17285,7 +17682,7 @@ static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppLVa 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___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppLValsAux_loop___lambda__3___closed__1; x_2 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppLValsAux_loop___lambda__3___closed__2; -x_3 = lean_unsigned_to_nat(744u); +x_3 = lean_unsigned_to_nat(765u); x_4 = lean_unsigned_to_nat(8u); x_5 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppLValsAux_loop___lambda__3___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -17833,7 +18230,7 @@ lean_dec(x_104); lean_dec(x_103); x_114 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_114, 0, x_102); -x_115 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__8; +x_115 = l_List_forIn_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_mkBaseProjections___spec__1___closed__3; x_116 = lean_array_push(x_115, x_114); x_117 = l_Lean_Elab_Term_ElabAppArgs_State_etaArgs___default___closed__1; lean_inc(x_15); @@ -21338,7 +21735,7 @@ static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn_ _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("proj", 4); +x_1 = lean_mk_string_from_bytes("choice", 6); return x_1; } } @@ -21346,7 +21743,7 @@ static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_shouldPropagateExpectedTypeFor___closed__6; +x_1 = lean_box(0); x_2 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -21356,7 +21753,7 @@ static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn_ _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("pipeProj", 8); +x_1 = lean_mk_string_from_bytes("proj", 4); return x_1; } } @@ -21374,7 +21771,7 @@ static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn_ _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("arrayRef", 8); +x_1 = lean_mk_string_from_bytes("pipeProj", 8); return x_1; } } @@ -21392,7 +21789,7 @@ static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn_ _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("namedPattern", 12); +x_1 = lean_mk_string_from_bytes("arrayRef", 8); return x_1; } } @@ -21410,7 +21807,7 @@ static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn_ _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("ident", 5); +x_1 = lean_mk_string_from_bytes("namedPattern", 12); return x_1; } } @@ -21418,7 +21815,7 @@ static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); +x_1 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_shouldPropagateExpectedTypeFor___closed__6; x_2 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__9; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -21428,7 +21825,7 @@ static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn_ _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("explicitUniv", 12); +x_1 = lean_mk_string_from_bytes("ident", 5); return x_1; } } @@ -21436,7 +21833,7 @@ static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_shouldPropagateExpectedTypeFor___closed__6; +x_1 = lean_box(0); x_2 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__11; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -21446,7 +21843,7 @@ static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn_ _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("explicit", 8); +x_1 = lean_mk_string_from_bytes("explicitUniv", 12); return x_1; } } @@ -21464,7 +21861,7 @@ static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn_ _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("dotIdent", 8); +x_1 = lean_mk_string_from_bytes("explicit", 8); return x_1; } } @@ -21482,24 +21879,25 @@ static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn_ _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("placeholders '_' cannot be used where a function is expected", 60); +x_1 = lean_mk_string_from_bytes("dotIdent", 8); return x_1; } } static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__18() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__17; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_shouldPropagateExpectedTypeFor___closed__6; +x_2 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__17; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; } } static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__19() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("unexpected occurrence of named pattern", 38); +x_1 = lean_mk_string_from_bytes("placeholders '_' cannot be used where a function is expected", 60); return x_1; } } @@ -21516,21 +21914,38 @@ static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn_ _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("fieldIdx", 8); +x_1 = lean_mk_string_from_bytes("unexpected occurrence of named pattern", 38); return x_1; } } static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__22() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__21; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__23() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("fieldIdx", 8); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__24() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__21; +x_2 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__23; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__23() { +static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__25() { _start: { lean_object* x_1; @@ -21538,7 +21953,7 @@ x_1 = lean_mk_string_from_bytes("Init.Data.Option.BasicAux", 25); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__24() { +static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__26() { _start: { lean_object* x_1; @@ -21546,7 +21961,7 @@ x_1 = lean_mk_string_from_bytes("Option.get!", 11); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__25() { +static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__27() { _start: { lean_object* x_1; @@ -21554,15 +21969,15 @@ x_1 = lean_mk_string_from_bytes("value is none", 13); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__26() { +static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__28() { _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___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__23; -x_2 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__24; +x_1 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__25; +x_2 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__26; x_3 = lean_unsigned_to_nat(16u); x_4 = lean_unsigned_to_nat(14u); -x_5 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__25; +x_5 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__27; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); return x_6; } @@ -21573,49 +21988,49 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn(l lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_inc(x_1); x_17 = l_Lean_Syntax_getKind(x_1); -x_18 = l_Lean_choiceKind; +x_18 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__2; x_19 = lean_name_eq(x_17, x_18); lean_dec(x_17); if (x_19 == 0) { lean_object* x_20; uint8_t x_21; -x_20 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__2; +x_20 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__4; lean_inc(x_1); x_21 = l_Lean_Syntax_isOfKind(x_1, x_20); if (x_21 == 0) { lean_object* x_22; uint8_t x_23; -x_22 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__4; +x_22 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__6; lean_inc(x_1); x_23 = l_Lean_Syntax_isOfKind(x_1, x_22); if (x_23 == 0) { lean_object* x_24; uint8_t x_25; -x_24 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__6; +x_24 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__8; lean_inc(x_1); x_25 = l_Lean_Syntax_isOfKind(x_1, x_24); if (x_25 == 0) { lean_object* x_26; uint8_t x_27; -x_26 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__8; +x_26 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__10; lean_inc(x_1); x_27 = l_Lean_Syntax_isOfKind(x_1, x_26); if (x_27 == 0) { lean_object* x_28; uint8_t x_29; -x_28 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__10; +x_28 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__12; lean_inc(x_1); x_29 = l_Lean_Syntax_isOfKind(x_1, x_28); if (x_29 == 0) { lean_object* x_30; uint8_t x_31; -x_30 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__12; +x_30 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__14; lean_inc(x_1); x_31 = l_Lean_Syntax_isOfKind(x_1, x_30); if (x_31 == 0) { lean_object* x_32; uint8_t x_33; -x_32 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__14; +x_32 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__16; lean_inc(x_1); x_33 = l_Lean_Syntax_isOfKind(x_1, x_32); if (x_33 == 0) @@ -21627,7 +22042,7 @@ x_35 = l_Lean_Syntax_isOfKind(x_1, x_34); if (x_35 == 0) { lean_object* x_36; uint8_t x_37; -x_36 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__16; +x_36 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__18; lean_inc(x_1); x_37 = l_Lean_Syntax_isOfKind(x_1, x_36); if (x_37 == 0) @@ -22628,7 +23043,7 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_270 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__18; +x_270 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__20; x_271 = l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___spec__1(x_270, x_10, x_11, x_12, x_13, x_14, x_15, x_16); lean_dec(x_15); lean_dec(x_14); @@ -23112,7 +23527,7 @@ x_384 = l_Lean_Syntax_getArg(x_1, x_383); lean_dec(x_1); x_385 = l_Lean_Syntax_getArgs(x_384); lean_dec(x_384); -x_386 = l_Lean_Syntax_SepArray_getElems___rarg(x_385); +x_386 = l_Lean_Syntax_TSepArray_getElems___rarg(x_385); lean_dec(x_385); lean_inc(x_15); lean_inc(x_14); @@ -23183,7 +23598,7 @@ else lean_object* x_397; lean_object* x_398; lean_object* x_399; uint8_t x_400; x_397 = lean_unsigned_to_nat(0u); x_398 = l_Lean_Syntax_getArg(x_1, x_397); -x_399 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__10; +x_399 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__12; x_400 = l_Lean_Syntax_isOfKind(x_398, x_399); if (x_400 == 0) { @@ -23566,249 +23981,248 @@ return x_491; } else { -lean_object* x_495; lean_object* x_496; lean_object* x_497; uint8_t x_498; +lean_object* x_495; lean_object* x_496; uint8_t x_497; x_495 = lean_unsigned_to_nat(2u); x_496 = l_Lean_Syntax_getArg(x_1, x_495); -x_497 = l_Lean_nullKind; -x_498 = l_Lean_Syntax_isNodeOf(x_496, x_497, x_397); -if (x_498 == 0) +x_497 = l_Lean_Syntax_matchesNull(x_496, x_397); +if (x_497 == 0) { -uint8_t x_499; uint8_t x_500; -x_499 = l_List_isEmpty___rarg(x_2); +uint8_t x_498; uint8_t x_499; +x_498 = l_List_isEmpty___rarg(x_2); if (x_8 == 0) { +uint8_t x_590; +x_590 = 1; +x_499 = x_590; +goto block_589; +} +else +{ uint8_t x_591; -x_591 = 1; -x_500 = x_591; -goto block_590; -} -else -{ -uint8_t x_592; -x_592 = 0; -x_500 = x_592; -goto block_590; -} -block_590: -{ -if (x_499 == 0) -{ -lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; -x_501 = lean_box(0); -x_502 = lean_box(x_500); -x_503 = lean_box(x_6); -x_504 = lean_box(x_7); -x_505 = lean_box(x_8); -x_506 = lean_alloc_closure((void*)(l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___lambda__1___boxed), 17, 10); -lean_closure_set(x_506, 0, x_1); -lean_closure_set(x_506, 1, x_501); -lean_closure_set(x_506, 2, x_502); -lean_closure_set(x_506, 3, x_2); -lean_closure_set(x_506, 4, x_3); -lean_closure_set(x_506, 5, x_4); -lean_closure_set(x_506, 6, x_5); -lean_closure_set(x_506, 7, x_503); -lean_closure_set(x_506, 8, x_504); -lean_closure_set(x_506, 9, x_505); -x_507 = l_Lean_Elab_Term_observing___rarg(x_506, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_507) == 0) -{ -uint8_t x_508; -x_508 = !lean_is_exclusive(x_507); -if (x_508 == 0) -{ -lean_object* x_509; lean_object* x_510; -x_509 = lean_ctor_get(x_507, 0); -x_510 = lean_array_push(x_9, x_509); -lean_ctor_set(x_507, 0, x_510); -return x_507; -} -else -{ -lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; -x_511 = lean_ctor_get(x_507, 0); -x_512 = lean_ctor_get(x_507, 1); -lean_inc(x_512); +x_591 = 0; +x_499 = x_591; +goto block_589; +} +block_589: +{ +if (x_498 == 0) +{ +lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; +x_500 = lean_box(0); +x_501 = lean_box(x_499); +x_502 = lean_box(x_6); +x_503 = lean_box(x_7); +x_504 = lean_box(x_8); +x_505 = lean_alloc_closure((void*)(l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___lambda__1___boxed), 17, 10); +lean_closure_set(x_505, 0, x_1); +lean_closure_set(x_505, 1, x_500); +lean_closure_set(x_505, 2, x_501); +lean_closure_set(x_505, 3, x_2); +lean_closure_set(x_505, 4, x_3); +lean_closure_set(x_505, 5, x_4); +lean_closure_set(x_505, 6, x_5); +lean_closure_set(x_505, 7, x_502); +lean_closure_set(x_505, 8, x_503); +lean_closure_set(x_505, 9, x_504); +x_506 = l_Lean_Elab_Term_observing___rarg(x_505, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_506) == 0) +{ +uint8_t x_507; +x_507 = !lean_is_exclusive(x_506); +if (x_507 == 0) +{ +lean_object* x_508; lean_object* x_509; +x_508 = lean_ctor_get(x_506, 0); +x_509 = lean_array_push(x_9, x_508); +lean_ctor_set(x_506, 0, x_509); +return x_506; +} +else +{ +lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; +x_510 = lean_ctor_get(x_506, 0); +x_511 = lean_ctor_get(x_506, 1); lean_inc(x_511); -lean_dec(x_507); -x_513 = lean_array_push(x_9, x_511); -x_514 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_514, 0, x_513); -lean_ctor_set(x_514, 1, x_512); -return x_514; +lean_inc(x_510); +lean_dec(x_506); +x_512 = lean_array_push(x_9, x_510); +x_513 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_513, 0, x_512); +lean_ctor_set(x_513, 1, x_511); +return x_513; } } else { -uint8_t x_515; +uint8_t x_514; lean_dec(x_9); -x_515 = !lean_is_exclusive(x_507); -if (x_515 == 0) +x_514 = !lean_is_exclusive(x_506); +if (x_514 == 0) { -return x_507; +return x_506; } else { -lean_object* x_516; lean_object* x_517; lean_object* x_518; -x_516 = lean_ctor_get(x_507, 0); -x_517 = lean_ctor_get(x_507, 1); -lean_inc(x_517); +lean_object* x_515; lean_object* x_516; lean_object* x_517; +x_515 = lean_ctor_get(x_506, 0); +x_516 = lean_ctor_get(x_506, 1); lean_inc(x_516); -lean_dec(x_507); -x_518 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_518, 0, x_516); -lean_ctor_set(x_518, 1, x_517); -return x_518; -} -} -} -else -{ -uint8_t x_519; -x_519 = l_Array_isEmpty___rarg(x_3); -if (x_519 == 0) -{ -lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; -x_520 = lean_box(0); -x_521 = lean_box(x_500); -x_522 = lean_box(x_6); -x_523 = lean_box(x_7); -x_524 = lean_box(x_8); -x_525 = lean_alloc_closure((void*)(l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___lambda__1___boxed), 17, 10); -lean_closure_set(x_525, 0, x_1); -lean_closure_set(x_525, 1, x_520); -lean_closure_set(x_525, 2, x_521); -lean_closure_set(x_525, 3, x_2); -lean_closure_set(x_525, 4, x_3); -lean_closure_set(x_525, 5, x_4); -lean_closure_set(x_525, 6, x_5); -lean_closure_set(x_525, 7, x_522); -lean_closure_set(x_525, 8, x_523); -lean_closure_set(x_525, 9, x_524); -x_526 = l_Lean_Elab_Term_observing___rarg(x_525, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_526) == 0) -{ -uint8_t x_527; -x_527 = !lean_is_exclusive(x_526); -if (x_527 == 0) -{ -lean_object* x_528; lean_object* x_529; -x_528 = lean_ctor_get(x_526, 0); -x_529 = lean_array_push(x_9, x_528); -lean_ctor_set(x_526, 0, x_529); -return x_526; -} -else -{ -lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; -x_530 = lean_ctor_get(x_526, 0); -x_531 = lean_ctor_get(x_526, 1); -lean_inc(x_531); +lean_inc(x_515); +lean_dec(x_506); +x_517 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_517, 0, x_515); +lean_ctor_set(x_517, 1, x_516); +return x_517; +} +} +} +else +{ +uint8_t x_518; +x_518 = l_Array_isEmpty___rarg(x_3); +if (x_518 == 0) +{ +lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; +x_519 = lean_box(0); +x_520 = lean_box(x_499); +x_521 = lean_box(x_6); +x_522 = lean_box(x_7); +x_523 = lean_box(x_8); +x_524 = lean_alloc_closure((void*)(l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___lambda__1___boxed), 17, 10); +lean_closure_set(x_524, 0, x_1); +lean_closure_set(x_524, 1, x_519); +lean_closure_set(x_524, 2, x_520); +lean_closure_set(x_524, 3, x_2); +lean_closure_set(x_524, 4, x_3); +lean_closure_set(x_524, 5, x_4); +lean_closure_set(x_524, 6, x_5); +lean_closure_set(x_524, 7, x_521); +lean_closure_set(x_524, 8, x_522); +lean_closure_set(x_524, 9, x_523); +x_525 = l_Lean_Elab_Term_observing___rarg(x_524, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_525) == 0) +{ +uint8_t x_526; +x_526 = !lean_is_exclusive(x_525); +if (x_526 == 0) +{ +lean_object* x_527; lean_object* x_528; +x_527 = lean_ctor_get(x_525, 0); +x_528 = lean_array_push(x_9, x_527); +lean_ctor_set(x_525, 0, x_528); +return x_525; +} +else +{ +lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; +x_529 = lean_ctor_get(x_525, 0); +x_530 = lean_ctor_get(x_525, 1); lean_inc(x_530); -lean_dec(x_526); -x_532 = lean_array_push(x_9, x_530); -x_533 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_533, 0, x_532); -lean_ctor_set(x_533, 1, x_531); -return x_533; +lean_inc(x_529); +lean_dec(x_525); +x_531 = lean_array_push(x_9, x_529); +x_532 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_532, 0, x_531); +lean_ctor_set(x_532, 1, x_530); +return x_532; } } else { -uint8_t x_534; +uint8_t x_533; lean_dec(x_9); -x_534 = !lean_is_exclusive(x_526); -if (x_534 == 0) +x_533 = !lean_is_exclusive(x_525); +if (x_533 == 0) { -return x_526; +return x_525; } else { -lean_object* x_535; lean_object* x_536; lean_object* x_537; -x_535 = lean_ctor_get(x_526, 0); -x_536 = lean_ctor_get(x_526, 1); -lean_inc(x_536); +lean_object* x_534; lean_object* x_535; lean_object* x_536; +x_534 = lean_ctor_get(x_525, 0); +x_535 = lean_ctor_get(x_525, 1); lean_inc(x_535); -lean_dec(x_526); -x_537 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_537, 0, x_535); -lean_ctor_set(x_537, 1, x_536); -return x_537; -} -} -} -else -{ -uint8_t x_538; -x_538 = l_Array_isEmpty___rarg(x_4); -if (x_538 == 0) -{ -lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; -x_539 = lean_box(0); -x_540 = lean_box(x_500); -x_541 = lean_box(x_6); -x_542 = lean_box(x_7); -x_543 = lean_box(x_8); -x_544 = lean_alloc_closure((void*)(l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___lambda__1___boxed), 17, 10); -lean_closure_set(x_544, 0, x_1); -lean_closure_set(x_544, 1, x_539); -lean_closure_set(x_544, 2, x_540); -lean_closure_set(x_544, 3, x_2); -lean_closure_set(x_544, 4, x_3); -lean_closure_set(x_544, 5, x_4); -lean_closure_set(x_544, 6, x_5); -lean_closure_set(x_544, 7, x_541); -lean_closure_set(x_544, 8, x_542); -lean_closure_set(x_544, 9, x_543); -x_545 = l_Lean_Elab_Term_observing___rarg(x_544, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_545) == 0) -{ -uint8_t x_546; -x_546 = !lean_is_exclusive(x_545); -if (x_546 == 0) -{ -lean_object* x_547; lean_object* x_548; -x_547 = lean_ctor_get(x_545, 0); -x_548 = lean_array_push(x_9, x_547); -lean_ctor_set(x_545, 0, x_548); -return x_545; -} -else -{ -lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; -x_549 = lean_ctor_get(x_545, 0); -x_550 = lean_ctor_get(x_545, 1); -lean_inc(x_550); +lean_inc(x_534); +lean_dec(x_525); +x_536 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_536, 0, x_534); +lean_ctor_set(x_536, 1, x_535); +return x_536; +} +} +} +else +{ +uint8_t x_537; +x_537 = l_Array_isEmpty___rarg(x_4); +if (x_537 == 0) +{ +lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; +x_538 = lean_box(0); +x_539 = lean_box(x_499); +x_540 = lean_box(x_6); +x_541 = lean_box(x_7); +x_542 = lean_box(x_8); +x_543 = lean_alloc_closure((void*)(l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___lambda__1___boxed), 17, 10); +lean_closure_set(x_543, 0, x_1); +lean_closure_set(x_543, 1, x_538); +lean_closure_set(x_543, 2, x_539); +lean_closure_set(x_543, 3, x_2); +lean_closure_set(x_543, 4, x_3); +lean_closure_set(x_543, 5, x_4); +lean_closure_set(x_543, 6, x_5); +lean_closure_set(x_543, 7, x_540); +lean_closure_set(x_543, 8, x_541); +lean_closure_set(x_543, 9, x_542); +x_544 = l_Lean_Elab_Term_observing___rarg(x_543, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_544) == 0) +{ +uint8_t x_545; +x_545 = !lean_is_exclusive(x_544); +if (x_545 == 0) +{ +lean_object* x_546; lean_object* x_547; +x_546 = lean_ctor_get(x_544, 0); +x_547 = lean_array_push(x_9, x_546); +lean_ctor_set(x_544, 0, x_547); +return x_544; +} +else +{ +lean_object* x_548; lean_object* x_549; lean_object* x_550; lean_object* x_551; +x_548 = lean_ctor_get(x_544, 0); +x_549 = lean_ctor_get(x_544, 1); lean_inc(x_549); -lean_dec(x_545); -x_551 = lean_array_push(x_9, x_549); -x_552 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_552, 0, x_551); -lean_ctor_set(x_552, 1, x_550); -return x_552; +lean_inc(x_548); +lean_dec(x_544); +x_550 = lean_array_push(x_9, x_548); +x_551 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_551, 0, x_550); +lean_ctor_set(x_551, 1, x_549); +return x_551; } } else { -uint8_t x_553; +uint8_t x_552; lean_dec(x_9); -x_553 = !lean_is_exclusive(x_545); -if (x_553 == 0) +x_552 = !lean_is_exclusive(x_544); +if (x_552 == 0) { -return x_545; +return x_544; } else { -lean_object* x_554; lean_object* x_555; lean_object* x_556; -x_554 = lean_ctor_get(x_545, 0); -x_555 = lean_ctor_get(x_545, 1); -lean_inc(x_555); +lean_object* x_553; lean_object* x_554; lean_object* x_555; +x_553 = lean_ctor_get(x_544, 0); +x_554 = lean_ctor_get(x_544, 1); lean_inc(x_554); -lean_dec(x_545); -x_556 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_556, 0, x_554); -lean_ctor_set(x_556, 1, x_555); -return x_556; +lean_inc(x_553); +lean_dec(x_544); +x_555 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_555, 0, x_553); +lean_ctor_set(x_555, 1, x_554); +return x_555; } } } @@ -23819,129 +24233,129 @@ lean_dec(x_3); lean_dec(x_2); if (x_8 == 0) { -uint8_t x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; -x_557 = 1; -x_558 = lean_box(x_557); -x_559 = lean_box(x_557); -x_560 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTerm___boxed), 11, 4); -lean_closure_set(x_560, 0, x_1); -lean_closure_set(x_560, 1, x_5); -lean_closure_set(x_560, 2, x_558); -lean_closure_set(x_560, 3, x_559); -x_561 = l_Lean_Elab_Term_observing___rarg(x_560, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_561) == 0) -{ -uint8_t x_562; -x_562 = !lean_is_exclusive(x_561); -if (x_562 == 0) -{ -lean_object* x_563; lean_object* x_564; -x_563 = lean_ctor_get(x_561, 0); -x_564 = lean_array_push(x_9, x_563); -lean_ctor_set(x_561, 0, x_564); -return x_561; -} -else -{ -lean_object* x_565; lean_object* x_566; lean_object* x_567; lean_object* x_568; -x_565 = lean_ctor_get(x_561, 0); -x_566 = lean_ctor_get(x_561, 1); -lean_inc(x_566); +uint8_t x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; +x_556 = 1; +x_557 = lean_box(x_556); +x_558 = lean_box(x_556); +x_559 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTerm___boxed), 11, 4); +lean_closure_set(x_559, 0, x_1); +lean_closure_set(x_559, 1, x_5); +lean_closure_set(x_559, 2, x_557); +lean_closure_set(x_559, 3, x_558); +x_560 = l_Lean_Elab_Term_observing___rarg(x_559, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_560) == 0) +{ +uint8_t x_561; +x_561 = !lean_is_exclusive(x_560); +if (x_561 == 0) +{ +lean_object* x_562; lean_object* x_563; +x_562 = lean_ctor_get(x_560, 0); +x_563 = lean_array_push(x_9, x_562); +lean_ctor_set(x_560, 0, x_563); +return x_560; +} +else +{ +lean_object* x_564; lean_object* x_565; lean_object* x_566; lean_object* x_567; +x_564 = lean_ctor_get(x_560, 0); +x_565 = lean_ctor_get(x_560, 1); lean_inc(x_565); -lean_dec(x_561); -x_567 = lean_array_push(x_9, x_565); -x_568 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_568, 0, x_567); -lean_ctor_set(x_568, 1, x_566); -return x_568; +lean_inc(x_564); +lean_dec(x_560); +x_566 = lean_array_push(x_9, x_564); +x_567 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_567, 0, x_566); +lean_ctor_set(x_567, 1, x_565); +return x_567; } } else { -uint8_t x_569; +uint8_t x_568; lean_dec(x_9); -x_569 = !lean_is_exclusive(x_561); -if (x_569 == 0) +x_568 = !lean_is_exclusive(x_560); +if (x_568 == 0) { -return x_561; +return x_560; } else { -lean_object* x_570; lean_object* x_571; lean_object* x_572; -x_570 = lean_ctor_get(x_561, 0); -x_571 = lean_ctor_get(x_561, 1); -lean_inc(x_571); +lean_object* x_569; lean_object* x_570; lean_object* x_571; +x_569 = lean_ctor_get(x_560, 0); +x_570 = lean_ctor_get(x_560, 1); lean_inc(x_570); -lean_dec(x_561); -x_572 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_572, 0, x_570); -lean_ctor_set(x_572, 1, x_571); -return x_572; +lean_inc(x_569); +lean_dec(x_560); +x_571 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_571, 0, x_569); +lean_ctor_set(x_571, 1, x_570); +return x_571; } } } else { -lean_object* x_573; uint8_t x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; -x_573 = lean_box(0); -x_574 = 1; -x_575 = lean_box(x_500); -x_576 = lean_box(x_574); -x_577 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTermEnsuringType___boxed), 12, 5); -lean_closure_set(x_577, 0, x_1); -lean_closure_set(x_577, 1, x_5); -lean_closure_set(x_577, 2, x_575); -lean_closure_set(x_577, 3, x_576); -lean_closure_set(x_577, 4, x_573); -x_578 = l_Lean_Elab_Term_observing___rarg(x_577, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_578) == 0) +lean_object* x_572; uint8_t x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; +x_572 = lean_box(0); +x_573 = 1; +x_574 = lean_box(x_499); +x_575 = lean_box(x_573); +x_576 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTermEnsuringType___boxed), 12, 5); +lean_closure_set(x_576, 0, x_1); +lean_closure_set(x_576, 1, x_5); +lean_closure_set(x_576, 2, x_574); +lean_closure_set(x_576, 3, x_575); +lean_closure_set(x_576, 4, x_572); +x_577 = l_Lean_Elab_Term_observing___rarg(x_576, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_577) == 0) { -uint8_t x_579; -x_579 = !lean_is_exclusive(x_578); -if (x_579 == 0) +uint8_t x_578; +x_578 = !lean_is_exclusive(x_577); +if (x_578 == 0) { -lean_object* x_580; lean_object* x_581; -x_580 = lean_ctor_get(x_578, 0); -x_581 = lean_array_push(x_9, x_580); -lean_ctor_set(x_578, 0, x_581); -return x_578; +lean_object* x_579; lean_object* x_580; +x_579 = lean_ctor_get(x_577, 0); +x_580 = lean_array_push(x_9, x_579); +lean_ctor_set(x_577, 0, x_580); +return x_577; } else { -lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; -x_582 = lean_ctor_get(x_578, 0); -x_583 = lean_ctor_get(x_578, 1); -lean_inc(x_583); +lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; +x_581 = lean_ctor_get(x_577, 0); +x_582 = lean_ctor_get(x_577, 1); lean_inc(x_582); -lean_dec(x_578); -x_584 = lean_array_push(x_9, x_582); -x_585 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_585, 0, x_584); -lean_ctor_set(x_585, 1, x_583); -return x_585; +lean_inc(x_581); +lean_dec(x_577); +x_583 = lean_array_push(x_9, x_581); +x_584 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_584, 0, x_583); +lean_ctor_set(x_584, 1, x_582); +return x_584; } } else { -uint8_t x_586; +uint8_t x_585; lean_dec(x_9); -x_586 = !lean_is_exclusive(x_578); -if (x_586 == 0) +x_585 = !lean_is_exclusive(x_577); +if (x_585 == 0) { -return x_578; +return x_577; } else { -lean_object* x_587; lean_object* x_588; lean_object* x_589; -x_587 = lean_ctor_get(x_578, 0); -x_588 = lean_ctor_get(x_578, 1); -lean_inc(x_588); +lean_object* x_586; lean_object* x_587; lean_object* x_588; +x_586 = lean_ctor_get(x_577, 0); +x_587 = lean_ctor_get(x_577, 1); lean_inc(x_587); -lean_dec(x_578); -x_589 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_589, 0, x_587); -lean_ctor_set(x_589, 1, x_588); -return x_589; +lean_inc(x_586); +lean_dec(x_577); +x_588 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_588, 0, x_586); +lean_ctor_set(x_588, 1, x_587); +return x_588; } } } @@ -23952,302 +24366,302 @@ return x_589; } else { -lean_object* x_593; lean_object* x_594; +lean_object* x_592; lean_object* x_593; lean_dec(x_9); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_593 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__20; -x_594 = l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___spec__1(x_593, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +x_592 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__22; +x_593 = l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___spec__1(x_592, x_10, x_11, x_12, x_13, x_14, x_15, x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -return x_594; +return x_593; } } } } else { -lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; -x_595 = lean_unsigned_to_nat(0u); -x_596 = l_Lean_Syntax_getArg(x_1, x_595); -x_597 = lean_unsigned_to_nat(1u); -x_598 = l_Lean_Syntax_getArg(x_1, x_597); -x_599 = lean_unsigned_to_nat(2u); -x_600 = l_Lean_Syntax_getArg(x_1, x_599); +lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; +x_594 = lean_unsigned_to_nat(0u); +x_595 = l_Lean_Syntax_getArg(x_1, x_594); +x_596 = lean_unsigned_to_nat(1u); +x_597 = l_Lean_Syntax_getArg(x_1, x_596); +x_598 = lean_unsigned_to_nat(2u); +x_599 = l_Lean_Syntax_getArg(x_1, x_598); lean_dec(x_1); -x_601 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_601, 0, x_598); -lean_ctor_set(x_601, 1, x_600); -x_602 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_602, 0, x_601); -lean_ctor_set(x_602, 1, x_2); -x_1 = x_596; -x_2 = x_602; +x_600 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_600, 0, x_597); +lean_ctor_set(x_600, 1, x_599); +x_601 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_601, 0, x_600); +lean_ctor_set(x_601, 1, x_2); +x_1 = x_595; +x_2 = x_601; goto _start; } } else { -lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; uint8_t x_609; -x_604 = lean_unsigned_to_nat(0u); -x_605 = l_Lean_Syntax_getArg(x_1, x_604); -x_606 = lean_unsigned_to_nat(2u); -x_607 = l_Lean_Syntax_getArg(x_1, x_606); -x_608 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__22; -lean_inc(x_607); -x_609 = l_Lean_Syntax_isOfKind(x_607, x_608); -if (x_609 == 0) +lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; uint8_t x_608; +x_603 = lean_unsigned_to_nat(0u); +x_604 = l_Lean_Syntax_getArg(x_1, x_603); +x_605 = lean_unsigned_to_nat(2u); +x_606 = l_Lean_Syntax_getArg(x_1, x_605); +x_607 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__24; +lean_inc(x_606); +x_608 = l_Lean_Syntax_isOfKind(x_606, x_607); +if (x_608 == 0) { -lean_object* x_610; uint8_t x_611; -x_610 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__10; -lean_inc(x_607); -x_611 = l_Lean_Syntax_isOfKind(x_607, x_610); -if (x_611 == 0) +lean_object* x_609; uint8_t x_610; +x_609 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__12; +lean_inc(x_606); +x_610 = l_Lean_Syntax_isOfKind(x_606, x_609); +if (x_610 == 0) { -uint8_t x_612; uint8_t x_613; -lean_dec(x_607); -lean_dec(x_605); -x_612 = l_List_isEmpty___rarg(x_2); +uint8_t x_611; uint8_t x_612; +lean_dec(x_606); +lean_dec(x_604); +x_611 = l_List_isEmpty___rarg(x_2); if (x_8 == 0) { +uint8_t x_703; +x_703 = 1; +x_612 = x_703; +goto block_702; +} +else +{ uint8_t x_704; -x_704 = 1; -x_613 = x_704; -goto block_703; -} -else -{ -uint8_t x_705; -x_705 = 0; -x_613 = x_705; -goto block_703; -} -block_703: -{ -if (x_612 == 0) -{ -lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; lean_object* x_620; -x_614 = lean_box(0); -x_615 = lean_box(x_613); -x_616 = lean_box(x_6); -x_617 = lean_box(x_7); -x_618 = lean_box(x_8); -x_619 = lean_alloc_closure((void*)(l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___lambda__1___boxed), 17, 10); -lean_closure_set(x_619, 0, x_1); -lean_closure_set(x_619, 1, x_614); -lean_closure_set(x_619, 2, x_615); -lean_closure_set(x_619, 3, x_2); -lean_closure_set(x_619, 4, x_3); -lean_closure_set(x_619, 5, x_4); -lean_closure_set(x_619, 6, x_5); -lean_closure_set(x_619, 7, x_616); -lean_closure_set(x_619, 8, x_617); -lean_closure_set(x_619, 9, x_618); -x_620 = l_Lean_Elab_Term_observing___rarg(x_619, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_620) == 0) -{ -uint8_t x_621; -x_621 = !lean_is_exclusive(x_620); -if (x_621 == 0) -{ -lean_object* x_622; lean_object* x_623; -x_622 = lean_ctor_get(x_620, 0); -x_623 = lean_array_push(x_9, x_622); -lean_ctor_set(x_620, 0, x_623); -return x_620; -} -else -{ -lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; -x_624 = lean_ctor_get(x_620, 0); -x_625 = lean_ctor_get(x_620, 1); -lean_inc(x_625); +x_704 = 0; +x_612 = x_704; +goto block_702; +} +block_702: +{ +if (x_611 == 0) +{ +lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; +x_613 = lean_box(0); +x_614 = lean_box(x_612); +x_615 = lean_box(x_6); +x_616 = lean_box(x_7); +x_617 = lean_box(x_8); +x_618 = lean_alloc_closure((void*)(l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___lambda__1___boxed), 17, 10); +lean_closure_set(x_618, 0, x_1); +lean_closure_set(x_618, 1, x_613); +lean_closure_set(x_618, 2, x_614); +lean_closure_set(x_618, 3, x_2); +lean_closure_set(x_618, 4, x_3); +lean_closure_set(x_618, 5, x_4); +lean_closure_set(x_618, 6, x_5); +lean_closure_set(x_618, 7, x_615); +lean_closure_set(x_618, 8, x_616); +lean_closure_set(x_618, 9, x_617); +x_619 = l_Lean_Elab_Term_observing___rarg(x_618, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_619) == 0) +{ +uint8_t x_620; +x_620 = !lean_is_exclusive(x_619); +if (x_620 == 0) +{ +lean_object* x_621; lean_object* x_622; +x_621 = lean_ctor_get(x_619, 0); +x_622 = lean_array_push(x_9, x_621); +lean_ctor_set(x_619, 0, x_622); +return x_619; +} +else +{ +lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; +x_623 = lean_ctor_get(x_619, 0); +x_624 = lean_ctor_get(x_619, 1); lean_inc(x_624); -lean_dec(x_620); -x_626 = lean_array_push(x_9, x_624); -x_627 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_627, 0, x_626); -lean_ctor_set(x_627, 1, x_625); -return x_627; +lean_inc(x_623); +lean_dec(x_619); +x_625 = lean_array_push(x_9, x_623); +x_626 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_626, 0, x_625); +lean_ctor_set(x_626, 1, x_624); +return x_626; } } else { -uint8_t x_628; +uint8_t x_627; lean_dec(x_9); -x_628 = !lean_is_exclusive(x_620); -if (x_628 == 0) +x_627 = !lean_is_exclusive(x_619); +if (x_627 == 0) { -return x_620; +return x_619; } else { -lean_object* x_629; lean_object* x_630; lean_object* x_631; -x_629 = lean_ctor_get(x_620, 0); -x_630 = lean_ctor_get(x_620, 1); -lean_inc(x_630); +lean_object* x_628; lean_object* x_629; lean_object* x_630; +x_628 = lean_ctor_get(x_619, 0); +x_629 = lean_ctor_get(x_619, 1); lean_inc(x_629); -lean_dec(x_620); -x_631 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_631, 0, x_629); -lean_ctor_set(x_631, 1, x_630); -return x_631; -} -} -} -else -{ -uint8_t x_632; -x_632 = l_Array_isEmpty___rarg(x_3); -if (x_632 == 0) -{ -lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; -x_633 = lean_box(0); -x_634 = lean_box(x_613); -x_635 = lean_box(x_6); -x_636 = lean_box(x_7); -x_637 = lean_box(x_8); -x_638 = lean_alloc_closure((void*)(l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___lambda__1___boxed), 17, 10); -lean_closure_set(x_638, 0, x_1); -lean_closure_set(x_638, 1, x_633); -lean_closure_set(x_638, 2, x_634); -lean_closure_set(x_638, 3, x_2); -lean_closure_set(x_638, 4, x_3); -lean_closure_set(x_638, 5, x_4); -lean_closure_set(x_638, 6, x_5); -lean_closure_set(x_638, 7, x_635); -lean_closure_set(x_638, 8, x_636); -lean_closure_set(x_638, 9, x_637); -x_639 = l_Lean_Elab_Term_observing___rarg(x_638, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_639) == 0) -{ -uint8_t x_640; -x_640 = !lean_is_exclusive(x_639); -if (x_640 == 0) -{ -lean_object* x_641; lean_object* x_642; -x_641 = lean_ctor_get(x_639, 0); -x_642 = lean_array_push(x_9, x_641); -lean_ctor_set(x_639, 0, x_642); -return x_639; -} -else -{ -lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; -x_643 = lean_ctor_get(x_639, 0); -x_644 = lean_ctor_get(x_639, 1); -lean_inc(x_644); +lean_inc(x_628); +lean_dec(x_619); +x_630 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_630, 0, x_628); +lean_ctor_set(x_630, 1, x_629); +return x_630; +} +} +} +else +{ +uint8_t x_631; +x_631 = l_Array_isEmpty___rarg(x_3); +if (x_631 == 0) +{ +lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; +x_632 = lean_box(0); +x_633 = lean_box(x_612); +x_634 = lean_box(x_6); +x_635 = lean_box(x_7); +x_636 = lean_box(x_8); +x_637 = lean_alloc_closure((void*)(l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___lambda__1___boxed), 17, 10); +lean_closure_set(x_637, 0, x_1); +lean_closure_set(x_637, 1, x_632); +lean_closure_set(x_637, 2, x_633); +lean_closure_set(x_637, 3, x_2); +lean_closure_set(x_637, 4, x_3); +lean_closure_set(x_637, 5, x_4); +lean_closure_set(x_637, 6, x_5); +lean_closure_set(x_637, 7, x_634); +lean_closure_set(x_637, 8, x_635); +lean_closure_set(x_637, 9, x_636); +x_638 = l_Lean_Elab_Term_observing___rarg(x_637, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_638) == 0) +{ +uint8_t x_639; +x_639 = !lean_is_exclusive(x_638); +if (x_639 == 0) +{ +lean_object* x_640; lean_object* x_641; +x_640 = lean_ctor_get(x_638, 0); +x_641 = lean_array_push(x_9, x_640); +lean_ctor_set(x_638, 0, x_641); +return x_638; +} +else +{ +lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; +x_642 = lean_ctor_get(x_638, 0); +x_643 = lean_ctor_get(x_638, 1); lean_inc(x_643); -lean_dec(x_639); -x_645 = lean_array_push(x_9, x_643); -x_646 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_646, 0, x_645); -lean_ctor_set(x_646, 1, x_644); -return x_646; +lean_inc(x_642); +lean_dec(x_638); +x_644 = lean_array_push(x_9, x_642); +x_645 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_645, 0, x_644); +lean_ctor_set(x_645, 1, x_643); +return x_645; } } else { -uint8_t x_647; +uint8_t x_646; lean_dec(x_9); -x_647 = !lean_is_exclusive(x_639); -if (x_647 == 0) +x_646 = !lean_is_exclusive(x_638); +if (x_646 == 0) { -return x_639; +return x_638; } else { -lean_object* x_648; lean_object* x_649; lean_object* x_650; -x_648 = lean_ctor_get(x_639, 0); -x_649 = lean_ctor_get(x_639, 1); -lean_inc(x_649); +lean_object* x_647; lean_object* x_648; lean_object* x_649; +x_647 = lean_ctor_get(x_638, 0); +x_648 = lean_ctor_get(x_638, 1); lean_inc(x_648); -lean_dec(x_639); -x_650 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_650, 0, x_648); -lean_ctor_set(x_650, 1, x_649); -return x_650; -} -} -} -else -{ -uint8_t x_651; -x_651 = l_Array_isEmpty___rarg(x_4); -if (x_651 == 0) -{ -lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; lean_object* x_658; -x_652 = lean_box(0); -x_653 = lean_box(x_613); -x_654 = lean_box(x_6); -x_655 = lean_box(x_7); -x_656 = lean_box(x_8); -x_657 = lean_alloc_closure((void*)(l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___lambda__1___boxed), 17, 10); -lean_closure_set(x_657, 0, x_1); -lean_closure_set(x_657, 1, x_652); -lean_closure_set(x_657, 2, x_653); -lean_closure_set(x_657, 3, x_2); -lean_closure_set(x_657, 4, x_3); -lean_closure_set(x_657, 5, x_4); -lean_closure_set(x_657, 6, x_5); -lean_closure_set(x_657, 7, x_654); -lean_closure_set(x_657, 8, x_655); -lean_closure_set(x_657, 9, x_656); -x_658 = l_Lean_Elab_Term_observing___rarg(x_657, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_658) == 0) -{ -uint8_t x_659; -x_659 = !lean_is_exclusive(x_658); -if (x_659 == 0) -{ -lean_object* x_660; lean_object* x_661; -x_660 = lean_ctor_get(x_658, 0); -x_661 = lean_array_push(x_9, x_660); -lean_ctor_set(x_658, 0, x_661); -return x_658; -} -else -{ -lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; -x_662 = lean_ctor_get(x_658, 0); -x_663 = lean_ctor_get(x_658, 1); -lean_inc(x_663); +lean_inc(x_647); +lean_dec(x_638); +x_649 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_649, 0, x_647); +lean_ctor_set(x_649, 1, x_648); +return x_649; +} +} +} +else +{ +uint8_t x_650; +x_650 = l_Array_isEmpty___rarg(x_4); +if (x_650 == 0) +{ +lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; +x_651 = lean_box(0); +x_652 = lean_box(x_612); +x_653 = lean_box(x_6); +x_654 = lean_box(x_7); +x_655 = lean_box(x_8); +x_656 = lean_alloc_closure((void*)(l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___lambda__1___boxed), 17, 10); +lean_closure_set(x_656, 0, x_1); +lean_closure_set(x_656, 1, x_651); +lean_closure_set(x_656, 2, x_652); +lean_closure_set(x_656, 3, x_2); +lean_closure_set(x_656, 4, x_3); +lean_closure_set(x_656, 5, x_4); +lean_closure_set(x_656, 6, x_5); +lean_closure_set(x_656, 7, x_653); +lean_closure_set(x_656, 8, x_654); +lean_closure_set(x_656, 9, x_655); +x_657 = l_Lean_Elab_Term_observing___rarg(x_656, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_657) == 0) +{ +uint8_t x_658; +x_658 = !lean_is_exclusive(x_657); +if (x_658 == 0) +{ +lean_object* x_659; lean_object* x_660; +x_659 = lean_ctor_get(x_657, 0); +x_660 = lean_array_push(x_9, x_659); +lean_ctor_set(x_657, 0, x_660); +return x_657; +} +else +{ +lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; +x_661 = lean_ctor_get(x_657, 0); +x_662 = lean_ctor_get(x_657, 1); lean_inc(x_662); -lean_dec(x_658); -x_664 = lean_array_push(x_9, x_662); -x_665 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_665, 0, x_664); -lean_ctor_set(x_665, 1, x_663); -return x_665; +lean_inc(x_661); +lean_dec(x_657); +x_663 = lean_array_push(x_9, x_661); +x_664 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_664, 0, x_663); +lean_ctor_set(x_664, 1, x_662); +return x_664; } } else { -uint8_t x_666; +uint8_t x_665; lean_dec(x_9); -x_666 = !lean_is_exclusive(x_658); -if (x_666 == 0) +x_665 = !lean_is_exclusive(x_657); +if (x_665 == 0) { -return x_658; +return x_657; } else { -lean_object* x_667; lean_object* x_668; lean_object* x_669; -x_667 = lean_ctor_get(x_658, 0); -x_668 = lean_ctor_get(x_658, 1); -lean_inc(x_668); +lean_object* x_666; lean_object* x_667; lean_object* x_668; +x_666 = lean_ctor_get(x_657, 0); +x_667 = lean_ctor_get(x_657, 1); lean_inc(x_667); -lean_dec(x_658); -x_669 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_669, 0, x_667); -lean_ctor_set(x_669, 1, x_668); -return x_669; +lean_inc(x_666); +lean_dec(x_657); +x_668 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_668, 0, x_666); +lean_ctor_set(x_668, 1, x_667); +return x_668; } } } @@ -24258,129 +24672,129 @@ lean_dec(x_3); lean_dec(x_2); if (x_8 == 0) { -uint8_t x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; lean_object* x_674; -x_670 = 1; -x_671 = lean_box(x_670); -x_672 = lean_box(x_670); -x_673 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTerm___boxed), 11, 4); -lean_closure_set(x_673, 0, x_1); -lean_closure_set(x_673, 1, x_5); -lean_closure_set(x_673, 2, x_671); -lean_closure_set(x_673, 3, x_672); -x_674 = l_Lean_Elab_Term_observing___rarg(x_673, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_674) == 0) -{ -uint8_t x_675; -x_675 = !lean_is_exclusive(x_674); -if (x_675 == 0) -{ -lean_object* x_676; lean_object* x_677; -x_676 = lean_ctor_get(x_674, 0); -x_677 = lean_array_push(x_9, x_676); -lean_ctor_set(x_674, 0, x_677); -return x_674; -} -else -{ -lean_object* x_678; lean_object* x_679; lean_object* x_680; lean_object* x_681; -x_678 = lean_ctor_get(x_674, 0); -x_679 = lean_ctor_get(x_674, 1); -lean_inc(x_679); +uint8_t x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; +x_669 = 1; +x_670 = lean_box(x_669); +x_671 = lean_box(x_669); +x_672 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTerm___boxed), 11, 4); +lean_closure_set(x_672, 0, x_1); +lean_closure_set(x_672, 1, x_5); +lean_closure_set(x_672, 2, x_670); +lean_closure_set(x_672, 3, x_671); +x_673 = l_Lean_Elab_Term_observing___rarg(x_672, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_673) == 0) +{ +uint8_t x_674; +x_674 = !lean_is_exclusive(x_673); +if (x_674 == 0) +{ +lean_object* x_675; lean_object* x_676; +x_675 = lean_ctor_get(x_673, 0); +x_676 = lean_array_push(x_9, x_675); +lean_ctor_set(x_673, 0, x_676); +return x_673; +} +else +{ +lean_object* x_677; lean_object* x_678; lean_object* x_679; lean_object* x_680; +x_677 = lean_ctor_get(x_673, 0); +x_678 = lean_ctor_get(x_673, 1); lean_inc(x_678); -lean_dec(x_674); -x_680 = lean_array_push(x_9, x_678); -x_681 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_681, 0, x_680); -lean_ctor_set(x_681, 1, x_679); -return x_681; +lean_inc(x_677); +lean_dec(x_673); +x_679 = lean_array_push(x_9, x_677); +x_680 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_680, 0, x_679); +lean_ctor_set(x_680, 1, x_678); +return x_680; } } else { -uint8_t x_682; +uint8_t x_681; lean_dec(x_9); -x_682 = !lean_is_exclusive(x_674); -if (x_682 == 0) +x_681 = !lean_is_exclusive(x_673); +if (x_681 == 0) { -return x_674; +return x_673; } else { -lean_object* x_683; lean_object* x_684; lean_object* x_685; -x_683 = lean_ctor_get(x_674, 0); -x_684 = lean_ctor_get(x_674, 1); -lean_inc(x_684); +lean_object* x_682; lean_object* x_683; lean_object* x_684; +x_682 = lean_ctor_get(x_673, 0); +x_683 = lean_ctor_get(x_673, 1); lean_inc(x_683); -lean_dec(x_674); -x_685 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_685, 0, x_683); -lean_ctor_set(x_685, 1, x_684); -return x_685; +lean_inc(x_682); +lean_dec(x_673); +x_684 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_684, 0, x_682); +lean_ctor_set(x_684, 1, x_683); +return x_684; } } } else { -lean_object* x_686; uint8_t x_687; lean_object* x_688; lean_object* x_689; lean_object* x_690; lean_object* x_691; -x_686 = lean_box(0); -x_687 = 1; -x_688 = lean_box(x_613); -x_689 = lean_box(x_687); -x_690 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTermEnsuringType___boxed), 12, 5); -lean_closure_set(x_690, 0, x_1); -lean_closure_set(x_690, 1, x_5); -lean_closure_set(x_690, 2, x_688); -lean_closure_set(x_690, 3, x_689); -lean_closure_set(x_690, 4, x_686); -x_691 = l_Lean_Elab_Term_observing___rarg(x_690, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_691) == 0) +lean_object* x_685; uint8_t x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; lean_object* x_690; +x_685 = lean_box(0); +x_686 = 1; +x_687 = lean_box(x_612); +x_688 = lean_box(x_686); +x_689 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTermEnsuringType___boxed), 12, 5); +lean_closure_set(x_689, 0, x_1); +lean_closure_set(x_689, 1, x_5); +lean_closure_set(x_689, 2, x_687); +lean_closure_set(x_689, 3, x_688); +lean_closure_set(x_689, 4, x_685); +x_690 = l_Lean_Elab_Term_observing___rarg(x_689, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_690) == 0) { -uint8_t x_692; -x_692 = !lean_is_exclusive(x_691); -if (x_692 == 0) +uint8_t x_691; +x_691 = !lean_is_exclusive(x_690); +if (x_691 == 0) { -lean_object* x_693; lean_object* x_694; -x_693 = lean_ctor_get(x_691, 0); -x_694 = lean_array_push(x_9, x_693); -lean_ctor_set(x_691, 0, x_694); -return x_691; +lean_object* x_692; lean_object* x_693; +x_692 = lean_ctor_get(x_690, 0); +x_693 = lean_array_push(x_9, x_692); +lean_ctor_set(x_690, 0, x_693); +return x_690; } else { -lean_object* x_695; lean_object* x_696; lean_object* x_697; lean_object* x_698; -x_695 = lean_ctor_get(x_691, 0); -x_696 = lean_ctor_get(x_691, 1); -lean_inc(x_696); +lean_object* x_694; lean_object* x_695; lean_object* x_696; lean_object* x_697; +x_694 = lean_ctor_get(x_690, 0); +x_695 = lean_ctor_get(x_690, 1); lean_inc(x_695); -lean_dec(x_691); -x_697 = lean_array_push(x_9, x_695); -x_698 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_698, 0, x_697); -lean_ctor_set(x_698, 1, x_696); -return x_698; +lean_inc(x_694); +lean_dec(x_690); +x_696 = lean_array_push(x_9, x_694); +x_697 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_697, 0, x_696); +lean_ctor_set(x_697, 1, x_695); +return x_697; } } else { -uint8_t x_699; +uint8_t x_698; lean_dec(x_9); -x_699 = !lean_is_exclusive(x_691); -if (x_699 == 0) +x_698 = !lean_is_exclusive(x_690); +if (x_698 == 0) { -return x_691; +return x_690; } else { -lean_object* x_700; lean_object* x_701; lean_object* x_702; -x_700 = lean_ctor_get(x_691, 0); -x_701 = lean_ctor_get(x_691, 1); -lean_inc(x_701); +lean_object* x_699; lean_object* x_700; lean_object* x_701; +x_699 = lean_ctor_get(x_690, 0); +x_700 = lean_ctor_get(x_690, 1); lean_inc(x_700); -lean_dec(x_691); -x_702 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_702, 0, x_700); -lean_ctor_set(x_702, 1, x_701); -return x_702; +lean_inc(x_699); +lean_dec(x_690); +x_701 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_701, 0, x_699); +lean_ctor_set(x_701, 1, x_700); +return x_701; } } } @@ -24391,251 +24805,250 @@ return x_702; } else { -lean_object* x_706; lean_object* x_707; lean_object* x_708; uint8_t x_709; -x_706 = lean_unsigned_to_nat(3u); -x_707 = l_Lean_Syntax_getArg(x_1, x_706); -x_708 = l_Lean_nullKind; -x_709 = l_Lean_Syntax_isNodeOf(x_707, x_708, x_604); -if (x_709 == 0) +lean_object* x_705; lean_object* x_706; uint8_t x_707; +x_705 = lean_unsigned_to_nat(3u); +x_706 = l_Lean_Syntax_getArg(x_1, x_705); +x_707 = l_Lean_Syntax_matchesNull(x_706, x_603); +if (x_707 == 0) { -uint8_t x_710; uint8_t x_711; -lean_dec(x_607); -lean_dec(x_605); -x_710 = l_List_isEmpty___rarg(x_2); +uint8_t x_708; uint8_t x_709; +lean_dec(x_606); +lean_dec(x_604); +x_708 = l_List_isEmpty___rarg(x_2); if (x_8 == 0) { -uint8_t x_802; -x_802 = 1; -x_711 = x_802; -goto block_801; +uint8_t x_800; +x_800 = 1; +x_709 = x_800; +goto block_799; } else { -uint8_t x_803; -x_803 = 0; -x_711 = x_803; -goto block_801; +uint8_t x_801; +x_801 = 0; +x_709 = x_801; +goto block_799; } -block_801: +block_799: { -if (x_710 == 0) +if (x_708 == 0) { -lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; lean_object* x_718; -x_712 = lean_box(0); -x_713 = lean_box(x_711); -x_714 = lean_box(x_6); -x_715 = lean_box(x_7); -x_716 = lean_box(x_8); -x_717 = lean_alloc_closure((void*)(l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___lambda__1___boxed), 17, 10); -lean_closure_set(x_717, 0, x_1); -lean_closure_set(x_717, 1, x_712); -lean_closure_set(x_717, 2, x_713); -lean_closure_set(x_717, 3, x_2); -lean_closure_set(x_717, 4, x_3); -lean_closure_set(x_717, 5, x_4); -lean_closure_set(x_717, 6, x_5); -lean_closure_set(x_717, 7, x_714); -lean_closure_set(x_717, 8, x_715); -lean_closure_set(x_717, 9, x_716); -x_718 = l_Lean_Elab_Term_observing___rarg(x_717, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_718) == 0) +lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; +x_710 = lean_box(0); +x_711 = lean_box(x_709); +x_712 = lean_box(x_6); +x_713 = lean_box(x_7); +x_714 = lean_box(x_8); +x_715 = lean_alloc_closure((void*)(l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___lambda__1___boxed), 17, 10); +lean_closure_set(x_715, 0, x_1); +lean_closure_set(x_715, 1, x_710); +lean_closure_set(x_715, 2, x_711); +lean_closure_set(x_715, 3, x_2); +lean_closure_set(x_715, 4, x_3); +lean_closure_set(x_715, 5, x_4); +lean_closure_set(x_715, 6, x_5); +lean_closure_set(x_715, 7, x_712); +lean_closure_set(x_715, 8, x_713); +lean_closure_set(x_715, 9, x_714); +x_716 = l_Lean_Elab_Term_observing___rarg(x_715, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_716) == 0) { -uint8_t x_719; -x_719 = !lean_is_exclusive(x_718); -if (x_719 == 0) +uint8_t x_717; +x_717 = !lean_is_exclusive(x_716); +if (x_717 == 0) { -lean_object* x_720; lean_object* x_721; -x_720 = lean_ctor_get(x_718, 0); -x_721 = lean_array_push(x_9, x_720); -lean_ctor_set(x_718, 0, x_721); -return x_718; +lean_object* x_718; lean_object* x_719; +x_718 = lean_ctor_get(x_716, 0); +x_719 = lean_array_push(x_9, x_718); +lean_ctor_set(x_716, 0, x_719); +return x_716; } else { -lean_object* x_722; lean_object* x_723; lean_object* x_724; lean_object* x_725; -x_722 = lean_ctor_get(x_718, 0); -x_723 = lean_ctor_get(x_718, 1); -lean_inc(x_723); -lean_inc(x_722); -lean_dec(x_718); -x_724 = lean_array_push(x_9, x_722); -x_725 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_725, 0, x_724); -lean_ctor_set(x_725, 1, x_723); -return x_725; +lean_object* x_720; lean_object* x_721; lean_object* x_722; lean_object* x_723; +x_720 = lean_ctor_get(x_716, 0); +x_721 = lean_ctor_get(x_716, 1); +lean_inc(x_721); +lean_inc(x_720); +lean_dec(x_716); +x_722 = lean_array_push(x_9, x_720); +x_723 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_723, 0, x_722); +lean_ctor_set(x_723, 1, x_721); +return x_723; } } else { -uint8_t x_726; +uint8_t x_724; lean_dec(x_9); -x_726 = !lean_is_exclusive(x_718); -if (x_726 == 0) +x_724 = !lean_is_exclusive(x_716); +if (x_724 == 0) { -return x_718; +return x_716; } else { -lean_object* x_727; lean_object* x_728; lean_object* x_729; -x_727 = lean_ctor_get(x_718, 0); -x_728 = lean_ctor_get(x_718, 1); -lean_inc(x_728); -lean_inc(x_727); -lean_dec(x_718); -x_729 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_729, 0, x_727); -lean_ctor_set(x_729, 1, x_728); -return x_729; +lean_object* x_725; lean_object* x_726; lean_object* x_727; +x_725 = lean_ctor_get(x_716, 0); +x_726 = lean_ctor_get(x_716, 1); +lean_inc(x_726); +lean_inc(x_725); +lean_dec(x_716); +x_727 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_727, 0, x_725); +lean_ctor_set(x_727, 1, x_726); +return x_727; } } } else { -uint8_t x_730; -x_730 = l_Array_isEmpty___rarg(x_3); -if (x_730 == 0) +uint8_t x_728; +x_728 = l_Array_isEmpty___rarg(x_3); +if (x_728 == 0) { -lean_object* x_731; lean_object* x_732; lean_object* x_733; lean_object* x_734; lean_object* x_735; lean_object* x_736; lean_object* x_737; -x_731 = lean_box(0); -x_732 = lean_box(x_711); -x_733 = lean_box(x_6); -x_734 = lean_box(x_7); -x_735 = lean_box(x_8); -x_736 = lean_alloc_closure((void*)(l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___lambda__1___boxed), 17, 10); -lean_closure_set(x_736, 0, x_1); -lean_closure_set(x_736, 1, x_731); -lean_closure_set(x_736, 2, x_732); -lean_closure_set(x_736, 3, x_2); -lean_closure_set(x_736, 4, x_3); -lean_closure_set(x_736, 5, x_4); -lean_closure_set(x_736, 6, x_5); -lean_closure_set(x_736, 7, x_733); -lean_closure_set(x_736, 8, x_734); -lean_closure_set(x_736, 9, x_735); -x_737 = l_Lean_Elab_Term_observing___rarg(x_736, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_737) == 0) +lean_object* x_729; lean_object* x_730; lean_object* x_731; lean_object* x_732; lean_object* x_733; lean_object* x_734; lean_object* x_735; +x_729 = lean_box(0); +x_730 = lean_box(x_709); +x_731 = lean_box(x_6); +x_732 = lean_box(x_7); +x_733 = lean_box(x_8); +x_734 = lean_alloc_closure((void*)(l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___lambda__1___boxed), 17, 10); +lean_closure_set(x_734, 0, x_1); +lean_closure_set(x_734, 1, x_729); +lean_closure_set(x_734, 2, x_730); +lean_closure_set(x_734, 3, x_2); +lean_closure_set(x_734, 4, x_3); +lean_closure_set(x_734, 5, x_4); +lean_closure_set(x_734, 6, x_5); +lean_closure_set(x_734, 7, x_731); +lean_closure_set(x_734, 8, x_732); +lean_closure_set(x_734, 9, x_733); +x_735 = l_Lean_Elab_Term_observing___rarg(x_734, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_735) == 0) { -uint8_t x_738; -x_738 = !lean_is_exclusive(x_737); -if (x_738 == 0) +uint8_t x_736; +x_736 = !lean_is_exclusive(x_735); +if (x_736 == 0) { -lean_object* x_739; lean_object* x_740; -x_739 = lean_ctor_get(x_737, 0); -x_740 = lean_array_push(x_9, x_739); -lean_ctor_set(x_737, 0, x_740); -return x_737; +lean_object* x_737; lean_object* x_738; +x_737 = lean_ctor_get(x_735, 0); +x_738 = lean_array_push(x_9, x_737); +lean_ctor_set(x_735, 0, x_738); +return x_735; } else { -lean_object* x_741; lean_object* x_742; lean_object* x_743; lean_object* x_744; -x_741 = lean_ctor_get(x_737, 0); -x_742 = lean_ctor_get(x_737, 1); -lean_inc(x_742); -lean_inc(x_741); -lean_dec(x_737); -x_743 = lean_array_push(x_9, x_741); -x_744 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_744, 0, x_743); -lean_ctor_set(x_744, 1, x_742); -return x_744; +lean_object* x_739; lean_object* x_740; lean_object* x_741; lean_object* x_742; +x_739 = lean_ctor_get(x_735, 0); +x_740 = lean_ctor_get(x_735, 1); +lean_inc(x_740); +lean_inc(x_739); +lean_dec(x_735); +x_741 = lean_array_push(x_9, x_739); +x_742 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_742, 0, x_741); +lean_ctor_set(x_742, 1, x_740); +return x_742; } } else { -uint8_t x_745; +uint8_t x_743; lean_dec(x_9); -x_745 = !lean_is_exclusive(x_737); -if (x_745 == 0) +x_743 = !lean_is_exclusive(x_735); +if (x_743 == 0) { -return x_737; +return x_735; } else { -lean_object* x_746; lean_object* x_747; lean_object* x_748; -x_746 = lean_ctor_get(x_737, 0); -x_747 = lean_ctor_get(x_737, 1); -lean_inc(x_747); -lean_inc(x_746); -lean_dec(x_737); -x_748 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_748, 0, x_746); -lean_ctor_set(x_748, 1, x_747); -return x_748; +lean_object* x_744; lean_object* x_745; lean_object* x_746; +x_744 = lean_ctor_get(x_735, 0); +x_745 = lean_ctor_get(x_735, 1); +lean_inc(x_745); +lean_inc(x_744); +lean_dec(x_735); +x_746 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_746, 0, x_744); +lean_ctor_set(x_746, 1, x_745); +return x_746; } } } else { -uint8_t x_749; -x_749 = l_Array_isEmpty___rarg(x_4); -if (x_749 == 0) +uint8_t x_747; +x_747 = l_Array_isEmpty___rarg(x_4); +if (x_747 == 0) { -lean_object* x_750; lean_object* x_751; lean_object* x_752; lean_object* x_753; lean_object* x_754; lean_object* x_755; lean_object* x_756; -x_750 = lean_box(0); -x_751 = lean_box(x_711); -x_752 = lean_box(x_6); -x_753 = lean_box(x_7); -x_754 = lean_box(x_8); -x_755 = lean_alloc_closure((void*)(l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___lambda__1___boxed), 17, 10); -lean_closure_set(x_755, 0, x_1); -lean_closure_set(x_755, 1, x_750); -lean_closure_set(x_755, 2, x_751); -lean_closure_set(x_755, 3, x_2); -lean_closure_set(x_755, 4, x_3); -lean_closure_set(x_755, 5, x_4); -lean_closure_set(x_755, 6, x_5); -lean_closure_set(x_755, 7, x_752); -lean_closure_set(x_755, 8, x_753); -lean_closure_set(x_755, 9, x_754); -x_756 = l_Lean_Elab_Term_observing___rarg(x_755, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_756) == 0) +lean_object* x_748; lean_object* x_749; lean_object* x_750; lean_object* x_751; lean_object* x_752; lean_object* x_753; lean_object* x_754; +x_748 = lean_box(0); +x_749 = lean_box(x_709); +x_750 = lean_box(x_6); +x_751 = lean_box(x_7); +x_752 = lean_box(x_8); +x_753 = lean_alloc_closure((void*)(l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___lambda__1___boxed), 17, 10); +lean_closure_set(x_753, 0, x_1); +lean_closure_set(x_753, 1, x_748); +lean_closure_set(x_753, 2, x_749); +lean_closure_set(x_753, 3, x_2); +lean_closure_set(x_753, 4, x_3); +lean_closure_set(x_753, 5, x_4); +lean_closure_set(x_753, 6, x_5); +lean_closure_set(x_753, 7, x_750); +lean_closure_set(x_753, 8, x_751); +lean_closure_set(x_753, 9, x_752); +x_754 = l_Lean_Elab_Term_observing___rarg(x_753, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_754) == 0) { -uint8_t x_757; -x_757 = !lean_is_exclusive(x_756); -if (x_757 == 0) +uint8_t x_755; +x_755 = !lean_is_exclusive(x_754); +if (x_755 == 0) { -lean_object* x_758; lean_object* x_759; -x_758 = lean_ctor_get(x_756, 0); -x_759 = lean_array_push(x_9, x_758); -lean_ctor_set(x_756, 0, x_759); -return x_756; +lean_object* x_756; lean_object* x_757; +x_756 = lean_ctor_get(x_754, 0); +x_757 = lean_array_push(x_9, x_756); +lean_ctor_set(x_754, 0, x_757); +return x_754; } else { -lean_object* x_760; lean_object* x_761; lean_object* x_762; lean_object* x_763; -x_760 = lean_ctor_get(x_756, 0); -x_761 = lean_ctor_get(x_756, 1); -lean_inc(x_761); -lean_inc(x_760); -lean_dec(x_756); -x_762 = lean_array_push(x_9, x_760); -x_763 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_763, 0, x_762); -lean_ctor_set(x_763, 1, x_761); -return x_763; +lean_object* x_758; lean_object* x_759; lean_object* x_760; lean_object* x_761; +x_758 = lean_ctor_get(x_754, 0); +x_759 = lean_ctor_get(x_754, 1); +lean_inc(x_759); +lean_inc(x_758); +lean_dec(x_754); +x_760 = lean_array_push(x_9, x_758); +x_761 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_761, 0, x_760); +lean_ctor_set(x_761, 1, x_759); +return x_761; } } else { -uint8_t x_764; +uint8_t x_762; lean_dec(x_9); -x_764 = !lean_is_exclusive(x_756); -if (x_764 == 0) +x_762 = !lean_is_exclusive(x_754); +if (x_762 == 0) { -return x_756; +return x_754; } else { -lean_object* x_765; lean_object* x_766; lean_object* x_767; -x_765 = lean_ctor_get(x_756, 0); -x_766 = lean_ctor_get(x_756, 1); -lean_inc(x_766); -lean_inc(x_765); -lean_dec(x_756); -x_767 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_767, 0, x_765); -lean_ctor_set(x_767, 1, x_766); -return x_767; +lean_object* x_763; lean_object* x_764; lean_object* x_765; +x_763 = lean_ctor_get(x_754, 0); +x_764 = lean_ctor_get(x_754, 1); +lean_inc(x_764); +lean_inc(x_763); +lean_dec(x_754); +x_765 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_765, 0, x_763); +lean_ctor_set(x_765, 1, x_764); +return x_765; } } } @@ -24646,129 +25059,129 @@ lean_dec(x_3); lean_dec(x_2); if (x_8 == 0) { -uint8_t x_768; lean_object* x_769; lean_object* x_770; lean_object* x_771; lean_object* x_772; -x_768 = 1; -x_769 = lean_box(x_768); -x_770 = lean_box(x_768); -x_771 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTerm___boxed), 11, 4); -lean_closure_set(x_771, 0, x_1); -lean_closure_set(x_771, 1, x_5); -lean_closure_set(x_771, 2, x_769); -lean_closure_set(x_771, 3, x_770); -x_772 = l_Lean_Elab_Term_observing___rarg(x_771, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_772) == 0) +uint8_t x_766; lean_object* x_767; lean_object* x_768; lean_object* x_769; lean_object* x_770; +x_766 = 1; +x_767 = lean_box(x_766); +x_768 = lean_box(x_766); +x_769 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTerm___boxed), 11, 4); +lean_closure_set(x_769, 0, x_1); +lean_closure_set(x_769, 1, x_5); +lean_closure_set(x_769, 2, x_767); +lean_closure_set(x_769, 3, x_768); +x_770 = l_Lean_Elab_Term_observing___rarg(x_769, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_770) == 0) { -uint8_t x_773; -x_773 = !lean_is_exclusive(x_772); -if (x_773 == 0) +uint8_t x_771; +x_771 = !lean_is_exclusive(x_770); +if (x_771 == 0) { -lean_object* x_774; lean_object* x_775; -x_774 = lean_ctor_get(x_772, 0); -x_775 = lean_array_push(x_9, x_774); -lean_ctor_set(x_772, 0, x_775); -return x_772; +lean_object* x_772; lean_object* x_773; +x_772 = lean_ctor_get(x_770, 0); +x_773 = lean_array_push(x_9, x_772); +lean_ctor_set(x_770, 0, x_773); +return x_770; } else { -lean_object* x_776; lean_object* x_777; lean_object* x_778; lean_object* x_779; -x_776 = lean_ctor_get(x_772, 0); -x_777 = lean_ctor_get(x_772, 1); -lean_inc(x_777); -lean_inc(x_776); -lean_dec(x_772); -x_778 = lean_array_push(x_9, x_776); -x_779 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_779, 0, x_778); -lean_ctor_set(x_779, 1, x_777); -return x_779; +lean_object* x_774; lean_object* x_775; lean_object* x_776; lean_object* x_777; +x_774 = lean_ctor_get(x_770, 0); +x_775 = lean_ctor_get(x_770, 1); +lean_inc(x_775); +lean_inc(x_774); +lean_dec(x_770); +x_776 = lean_array_push(x_9, x_774); +x_777 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_777, 0, x_776); +lean_ctor_set(x_777, 1, x_775); +return x_777; } } else { -uint8_t x_780; +uint8_t x_778; lean_dec(x_9); -x_780 = !lean_is_exclusive(x_772); -if (x_780 == 0) +x_778 = !lean_is_exclusive(x_770); +if (x_778 == 0) { -return x_772; +return x_770; } else { -lean_object* x_781; lean_object* x_782; lean_object* x_783; -x_781 = lean_ctor_get(x_772, 0); -x_782 = lean_ctor_get(x_772, 1); -lean_inc(x_782); -lean_inc(x_781); -lean_dec(x_772); -x_783 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_783, 0, x_781); -lean_ctor_set(x_783, 1, x_782); -return x_783; +lean_object* x_779; lean_object* x_780; lean_object* x_781; +x_779 = lean_ctor_get(x_770, 0); +x_780 = lean_ctor_get(x_770, 1); +lean_inc(x_780); +lean_inc(x_779); +lean_dec(x_770); +x_781 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_781, 0, x_779); +lean_ctor_set(x_781, 1, x_780); +return x_781; } } } else { -lean_object* x_784; uint8_t x_785; lean_object* x_786; lean_object* x_787; lean_object* x_788; lean_object* x_789; -x_784 = lean_box(0); -x_785 = 1; -x_786 = lean_box(x_711); -x_787 = lean_box(x_785); -x_788 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTermEnsuringType___boxed), 12, 5); -lean_closure_set(x_788, 0, x_1); -lean_closure_set(x_788, 1, x_5); -lean_closure_set(x_788, 2, x_786); -lean_closure_set(x_788, 3, x_787); -lean_closure_set(x_788, 4, x_784); -x_789 = l_Lean_Elab_Term_observing___rarg(x_788, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_789) == 0) +lean_object* x_782; uint8_t x_783; lean_object* x_784; lean_object* x_785; lean_object* x_786; lean_object* x_787; +x_782 = lean_box(0); +x_783 = 1; +x_784 = lean_box(x_709); +x_785 = lean_box(x_783); +x_786 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTermEnsuringType___boxed), 12, 5); +lean_closure_set(x_786, 0, x_1); +lean_closure_set(x_786, 1, x_5); +lean_closure_set(x_786, 2, x_784); +lean_closure_set(x_786, 3, x_785); +lean_closure_set(x_786, 4, x_782); +x_787 = l_Lean_Elab_Term_observing___rarg(x_786, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_787) == 0) { -uint8_t x_790; -x_790 = !lean_is_exclusive(x_789); -if (x_790 == 0) +uint8_t x_788; +x_788 = !lean_is_exclusive(x_787); +if (x_788 == 0) { -lean_object* x_791; lean_object* x_792; -x_791 = lean_ctor_get(x_789, 0); -x_792 = lean_array_push(x_9, x_791); -lean_ctor_set(x_789, 0, x_792); -return x_789; +lean_object* x_789; lean_object* x_790; +x_789 = lean_ctor_get(x_787, 0); +x_790 = lean_array_push(x_9, x_789); +lean_ctor_set(x_787, 0, x_790); +return x_787; } else { -lean_object* x_793; lean_object* x_794; lean_object* x_795; lean_object* x_796; -x_793 = lean_ctor_get(x_789, 0); -x_794 = lean_ctor_get(x_789, 1); -lean_inc(x_794); -lean_inc(x_793); -lean_dec(x_789); -x_795 = lean_array_push(x_9, x_793); -x_796 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_796, 0, x_795); -lean_ctor_set(x_796, 1, x_794); -return x_796; +lean_object* x_791; lean_object* x_792; lean_object* x_793; lean_object* x_794; +x_791 = lean_ctor_get(x_787, 0); +x_792 = lean_ctor_get(x_787, 1); +lean_inc(x_792); +lean_inc(x_791); +lean_dec(x_787); +x_793 = lean_array_push(x_9, x_791); +x_794 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_794, 0, x_793); +lean_ctor_set(x_794, 1, x_792); +return x_794; } } else { -uint8_t x_797; +uint8_t x_795; lean_dec(x_9); -x_797 = !lean_is_exclusive(x_789); -if (x_797 == 0) +x_795 = !lean_is_exclusive(x_787); +if (x_795 == 0) { -return x_789; +return x_787; } else { -lean_object* x_798; lean_object* x_799; lean_object* x_800; -x_798 = lean_ctor_get(x_789, 0); -x_799 = lean_ctor_get(x_789, 1); -lean_inc(x_799); -lean_inc(x_798); -lean_dec(x_789); -x_800 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_800, 0, x_798); -lean_ctor_set(x_800, 1, x_799); -return x_800; +lean_object* x_796; lean_object* x_797; lean_object* x_798; +x_796 = lean_ctor_get(x_787, 0); +x_797 = lean_ctor_get(x_787, 1); +lean_inc(x_797); +lean_inc(x_796); +lean_dec(x_787); +x_798 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_798, 0, x_796); +lean_ctor_set(x_798, 1, x_797); +return x_798; } } } @@ -24779,267 +25192,266 @@ return x_800; } else { -lean_object* x_804; lean_object* x_805; lean_object* x_806; lean_object* x_807; lean_object* x_808; +lean_object* x_802; lean_object* x_803; lean_object* x_804; lean_object* x_805; lean_object* x_806; lean_dec(x_1); +x_802 = lean_box(0); +x_803 = l_Lean_Syntax_identComponents(x_606, x_802); x_804 = lean_box(0); -x_805 = l_Lean_Syntax_identComponents(x_607, x_804); -x_806 = lean_box(0); -lean_inc(x_605); -x_807 = l_List_mapTRAux___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___spec__3(x_605, x_805, x_806); -x_808 = l_List_appendTR___rarg(x_807, x_2); -x_1 = x_605; -x_2 = x_808; +lean_inc(x_604); +x_805 = l_List_mapTRAux___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___spec__3(x_604, x_803, x_804); +x_806 = l_List_appendTR___rarg(x_805, x_2); +x_1 = x_604; +x_2 = x_806; goto _start; } } } else { -lean_object* x_810; lean_object* x_811; lean_object* x_812; uint8_t x_813; -x_810 = lean_unsigned_to_nat(3u); -x_811 = l_Lean_Syntax_getArg(x_1, x_810); -x_812 = l_Lean_nullKind; -x_813 = l_Lean_Syntax_isNodeOf(x_811, x_812, x_604); -if (x_813 == 0) +lean_object* x_808; lean_object* x_809; uint8_t x_810; +x_808 = lean_unsigned_to_nat(3u); +x_809 = l_Lean_Syntax_getArg(x_1, x_808); +x_810 = l_Lean_Syntax_matchesNull(x_809, x_603); +if (x_810 == 0) { -uint8_t x_814; uint8_t x_815; -lean_dec(x_607); -lean_dec(x_605); -x_814 = l_List_isEmpty___rarg(x_2); +uint8_t x_811; uint8_t x_812; +lean_dec(x_606); +lean_dec(x_604); +x_811 = l_List_isEmpty___rarg(x_2); if (x_8 == 0) { -uint8_t x_906; -x_906 = 1; -x_815 = x_906; -goto block_905; +uint8_t x_903; +x_903 = 1; +x_812 = x_903; +goto block_902; } else { -uint8_t x_907; -x_907 = 0; -x_815 = x_907; -goto block_905; +uint8_t x_904; +x_904 = 0; +x_812 = x_904; +goto block_902; } -block_905: +block_902: { -if (x_814 == 0) +if (x_811 == 0) { -lean_object* x_816; lean_object* x_817; lean_object* x_818; lean_object* x_819; lean_object* x_820; lean_object* x_821; lean_object* x_822; -x_816 = lean_box(0); -x_817 = lean_box(x_815); -x_818 = lean_box(x_6); -x_819 = lean_box(x_7); -x_820 = lean_box(x_8); -x_821 = lean_alloc_closure((void*)(l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___lambda__1___boxed), 17, 10); -lean_closure_set(x_821, 0, x_1); -lean_closure_set(x_821, 1, x_816); -lean_closure_set(x_821, 2, x_817); -lean_closure_set(x_821, 3, x_2); -lean_closure_set(x_821, 4, x_3); -lean_closure_set(x_821, 5, x_4); -lean_closure_set(x_821, 6, x_5); -lean_closure_set(x_821, 7, x_818); -lean_closure_set(x_821, 8, x_819); -lean_closure_set(x_821, 9, x_820); -x_822 = l_Lean_Elab_Term_observing___rarg(x_821, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_822) == 0) +lean_object* x_813; lean_object* x_814; lean_object* x_815; lean_object* x_816; lean_object* x_817; lean_object* x_818; lean_object* x_819; +x_813 = lean_box(0); +x_814 = lean_box(x_812); +x_815 = lean_box(x_6); +x_816 = lean_box(x_7); +x_817 = lean_box(x_8); +x_818 = lean_alloc_closure((void*)(l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___lambda__1___boxed), 17, 10); +lean_closure_set(x_818, 0, x_1); +lean_closure_set(x_818, 1, x_813); +lean_closure_set(x_818, 2, x_814); +lean_closure_set(x_818, 3, x_2); +lean_closure_set(x_818, 4, x_3); +lean_closure_set(x_818, 5, x_4); +lean_closure_set(x_818, 6, x_5); +lean_closure_set(x_818, 7, x_815); +lean_closure_set(x_818, 8, x_816); +lean_closure_set(x_818, 9, x_817); +x_819 = l_Lean_Elab_Term_observing___rarg(x_818, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_819) == 0) { -uint8_t x_823; -x_823 = !lean_is_exclusive(x_822); -if (x_823 == 0) +uint8_t x_820; +x_820 = !lean_is_exclusive(x_819); +if (x_820 == 0) { -lean_object* x_824; lean_object* x_825; -x_824 = lean_ctor_get(x_822, 0); -x_825 = lean_array_push(x_9, x_824); -lean_ctor_set(x_822, 0, x_825); -return x_822; +lean_object* x_821; lean_object* x_822; +x_821 = lean_ctor_get(x_819, 0); +x_822 = lean_array_push(x_9, x_821); +lean_ctor_set(x_819, 0, x_822); +return x_819; } else { -lean_object* x_826; lean_object* x_827; lean_object* x_828; lean_object* x_829; -x_826 = lean_ctor_get(x_822, 0); -x_827 = lean_ctor_get(x_822, 1); -lean_inc(x_827); -lean_inc(x_826); -lean_dec(x_822); -x_828 = lean_array_push(x_9, x_826); -x_829 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_829, 0, x_828); -lean_ctor_set(x_829, 1, x_827); -return x_829; +lean_object* x_823; lean_object* x_824; lean_object* x_825; lean_object* x_826; +x_823 = lean_ctor_get(x_819, 0); +x_824 = lean_ctor_get(x_819, 1); +lean_inc(x_824); +lean_inc(x_823); +lean_dec(x_819); +x_825 = lean_array_push(x_9, x_823); +x_826 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_826, 0, x_825); +lean_ctor_set(x_826, 1, x_824); +return x_826; } } else { -uint8_t x_830; +uint8_t x_827; lean_dec(x_9); -x_830 = !lean_is_exclusive(x_822); -if (x_830 == 0) +x_827 = !lean_is_exclusive(x_819); +if (x_827 == 0) { -return x_822; +return x_819; } else { -lean_object* x_831; lean_object* x_832; lean_object* x_833; -x_831 = lean_ctor_get(x_822, 0); -x_832 = lean_ctor_get(x_822, 1); -lean_inc(x_832); -lean_inc(x_831); -lean_dec(x_822); -x_833 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_833, 0, x_831); -lean_ctor_set(x_833, 1, x_832); -return x_833; +lean_object* x_828; lean_object* x_829; lean_object* x_830; +x_828 = lean_ctor_get(x_819, 0); +x_829 = lean_ctor_get(x_819, 1); +lean_inc(x_829); +lean_inc(x_828); +lean_dec(x_819); +x_830 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_830, 0, x_828); +lean_ctor_set(x_830, 1, x_829); +return x_830; } } } else { -uint8_t x_834; -x_834 = l_Array_isEmpty___rarg(x_3); -if (x_834 == 0) +uint8_t x_831; +x_831 = l_Array_isEmpty___rarg(x_3); +if (x_831 == 0) { -lean_object* x_835; lean_object* x_836; lean_object* x_837; lean_object* x_838; lean_object* x_839; lean_object* x_840; lean_object* x_841; -x_835 = lean_box(0); -x_836 = lean_box(x_815); -x_837 = lean_box(x_6); -x_838 = lean_box(x_7); -x_839 = lean_box(x_8); -x_840 = lean_alloc_closure((void*)(l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___lambda__1___boxed), 17, 10); -lean_closure_set(x_840, 0, x_1); -lean_closure_set(x_840, 1, x_835); -lean_closure_set(x_840, 2, x_836); -lean_closure_set(x_840, 3, x_2); -lean_closure_set(x_840, 4, x_3); -lean_closure_set(x_840, 5, x_4); -lean_closure_set(x_840, 6, x_5); -lean_closure_set(x_840, 7, x_837); -lean_closure_set(x_840, 8, x_838); -lean_closure_set(x_840, 9, x_839); -x_841 = l_Lean_Elab_Term_observing___rarg(x_840, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_841) == 0) +lean_object* x_832; lean_object* x_833; lean_object* x_834; lean_object* x_835; lean_object* x_836; lean_object* x_837; lean_object* x_838; +x_832 = lean_box(0); +x_833 = lean_box(x_812); +x_834 = lean_box(x_6); +x_835 = lean_box(x_7); +x_836 = lean_box(x_8); +x_837 = lean_alloc_closure((void*)(l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___lambda__1___boxed), 17, 10); +lean_closure_set(x_837, 0, x_1); +lean_closure_set(x_837, 1, x_832); +lean_closure_set(x_837, 2, x_833); +lean_closure_set(x_837, 3, x_2); +lean_closure_set(x_837, 4, x_3); +lean_closure_set(x_837, 5, x_4); +lean_closure_set(x_837, 6, x_5); +lean_closure_set(x_837, 7, x_834); +lean_closure_set(x_837, 8, x_835); +lean_closure_set(x_837, 9, x_836); +x_838 = l_Lean_Elab_Term_observing___rarg(x_837, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_838) == 0) { -uint8_t x_842; -x_842 = !lean_is_exclusive(x_841); -if (x_842 == 0) +uint8_t x_839; +x_839 = !lean_is_exclusive(x_838); +if (x_839 == 0) { -lean_object* x_843; lean_object* x_844; -x_843 = lean_ctor_get(x_841, 0); -x_844 = lean_array_push(x_9, x_843); -lean_ctor_set(x_841, 0, x_844); -return x_841; +lean_object* x_840; lean_object* x_841; +x_840 = lean_ctor_get(x_838, 0); +x_841 = lean_array_push(x_9, x_840); +lean_ctor_set(x_838, 0, x_841); +return x_838; } else { -lean_object* x_845; lean_object* x_846; lean_object* x_847; lean_object* x_848; -x_845 = lean_ctor_get(x_841, 0); -x_846 = lean_ctor_get(x_841, 1); -lean_inc(x_846); -lean_inc(x_845); -lean_dec(x_841); -x_847 = lean_array_push(x_9, x_845); -x_848 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_848, 0, x_847); -lean_ctor_set(x_848, 1, x_846); -return x_848; +lean_object* x_842; lean_object* x_843; lean_object* x_844; lean_object* x_845; +x_842 = lean_ctor_get(x_838, 0); +x_843 = lean_ctor_get(x_838, 1); +lean_inc(x_843); +lean_inc(x_842); +lean_dec(x_838); +x_844 = lean_array_push(x_9, x_842); +x_845 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_845, 0, x_844); +lean_ctor_set(x_845, 1, x_843); +return x_845; } } else { -uint8_t x_849; +uint8_t x_846; lean_dec(x_9); -x_849 = !lean_is_exclusive(x_841); -if (x_849 == 0) +x_846 = !lean_is_exclusive(x_838); +if (x_846 == 0) { -return x_841; +return x_838; } else { -lean_object* x_850; lean_object* x_851; lean_object* x_852; -x_850 = lean_ctor_get(x_841, 0); -x_851 = lean_ctor_get(x_841, 1); -lean_inc(x_851); -lean_inc(x_850); -lean_dec(x_841); -x_852 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_852, 0, x_850); -lean_ctor_set(x_852, 1, x_851); -return x_852; +lean_object* x_847; lean_object* x_848; lean_object* x_849; +x_847 = lean_ctor_get(x_838, 0); +x_848 = lean_ctor_get(x_838, 1); +lean_inc(x_848); +lean_inc(x_847); +lean_dec(x_838); +x_849 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_849, 0, x_847); +lean_ctor_set(x_849, 1, x_848); +return x_849; } } } else { -uint8_t x_853; -x_853 = l_Array_isEmpty___rarg(x_4); -if (x_853 == 0) +uint8_t x_850; +x_850 = l_Array_isEmpty___rarg(x_4); +if (x_850 == 0) { -lean_object* x_854; lean_object* x_855; lean_object* x_856; lean_object* x_857; lean_object* x_858; lean_object* x_859; lean_object* x_860; -x_854 = lean_box(0); -x_855 = lean_box(x_815); -x_856 = lean_box(x_6); -x_857 = lean_box(x_7); -x_858 = lean_box(x_8); -x_859 = lean_alloc_closure((void*)(l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___lambda__1___boxed), 17, 10); -lean_closure_set(x_859, 0, x_1); -lean_closure_set(x_859, 1, x_854); -lean_closure_set(x_859, 2, x_855); -lean_closure_set(x_859, 3, x_2); -lean_closure_set(x_859, 4, x_3); -lean_closure_set(x_859, 5, x_4); -lean_closure_set(x_859, 6, x_5); -lean_closure_set(x_859, 7, x_856); -lean_closure_set(x_859, 8, x_857); -lean_closure_set(x_859, 9, x_858); -x_860 = l_Lean_Elab_Term_observing___rarg(x_859, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_860) == 0) +lean_object* x_851; lean_object* x_852; lean_object* x_853; lean_object* x_854; lean_object* x_855; lean_object* x_856; lean_object* x_857; +x_851 = lean_box(0); +x_852 = lean_box(x_812); +x_853 = lean_box(x_6); +x_854 = lean_box(x_7); +x_855 = lean_box(x_8); +x_856 = lean_alloc_closure((void*)(l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___lambda__1___boxed), 17, 10); +lean_closure_set(x_856, 0, x_1); +lean_closure_set(x_856, 1, x_851); +lean_closure_set(x_856, 2, x_852); +lean_closure_set(x_856, 3, x_2); +lean_closure_set(x_856, 4, x_3); +lean_closure_set(x_856, 5, x_4); +lean_closure_set(x_856, 6, x_5); +lean_closure_set(x_856, 7, x_853); +lean_closure_set(x_856, 8, x_854); +lean_closure_set(x_856, 9, x_855); +x_857 = l_Lean_Elab_Term_observing___rarg(x_856, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_857) == 0) { -uint8_t x_861; -x_861 = !lean_is_exclusive(x_860); -if (x_861 == 0) +uint8_t x_858; +x_858 = !lean_is_exclusive(x_857); +if (x_858 == 0) { -lean_object* x_862; lean_object* x_863; -x_862 = lean_ctor_get(x_860, 0); -x_863 = lean_array_push(x_9, x_862); -lean_ctor_set(x_860, 0, x_863); -return x_860; +lean_object* x_859; lean_object* x_860; +x_859 = lean_ctor_get(x_857, 0); +x_860 = lean_array_push(x_9, x_859); +lean_ctor_set(x_857, 0, x_860); +return x_857; } else { -lean_object* x_864; lean_object* x_865; lean_object* x_866; lean_object* x_867; -x_864 = lean_ctor_get(x_860, 0); -x_865 = lean_ctor_get(x_860, 1); -lean_inc(x_865); -lean_inc(x_864); -lean_dec(x_860); -x_866 = lean_array_push(x_9, x_864); -x_867 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_867, 0, x_866); -lean_ctor_set(x_867, 1, x_865); -return x_867; +lean_object* x_861; lean_object* x_862; lean_object* x_863; lean_object* x_864; +x_861 = lean_ctor_get(x_857, 0); +x_862 = lean_ctor_get(x_857, 1); +lean_inc(x_862); +lean_inc(x_861); +lean_dec(x_857); +x_863 = lean_array_push(x_9, x_861); +x_864 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_864, 0, x_863); +lean_ctor_set(x_864, 1, x_862); +return x_864; } } else { -uint8_t x_868; +uint8_t x_865; lean_dec(x_9); -x_868 = !lean_is_exclusive(x_860); -if (x_868 == 0) +x_865 = !lean_is_exclusive(x_857); +if (x_865 == 0) { -return x_860; +return x_857; } else { -lean_object* x_869; lean_object* x_870; lean_object* x_871; -x_869 = lean_ctor_get(x_860, 0); -x_870 = lean_ctor_get(x_860, 1); -lean_inc(x_870); -lean_inc(x_869); -lean_dec(x_860); -x_871 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_871, 0, x_869); -lean_ctor_set(x_871, 1, x_870); -return x_871; +lean_object* x_866; lean_object* x_867; lean_object* x_868; +x_866 = lean_ctor_get(x_857, 0); +x_867 = lean_ctor_get(x_857, 1); +lean_inc(x_867); +lean_inc(x_866); +lean_dec(x_857); +x_868 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_868, 0, x_866); +lean_ctor_set(x_868, 1, x_867); +return x_868; } } } @@ -25050,129 +25462,129 @@ lean_dec(x_3); lean_dec(x_2); if (x_8 == 0) { -uint8_t x_872; lean_object* x_873; lean_object* x_874; lean_object* x_875; lean_object* x_876; -x_872 = 1; -x_873 = lean_box(x_872); -x_874 = lean_box(x_872); -x_875 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTerm___boxed), 11, 4); -lean_closure_set(x_875, 0, x_1); -lean_closure_set(x_875, 1, x_5); -lean_closure_set(x_875, 2, x_873); -lean_closure_set(x_875, 3, x_874); -x_876 = l_Lean_Elab_Term_observing___rarg(x_875, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_876) == 0) +uint8_t x_869; lean_object* x_870; lean_object* x_871; lean_object* x_872; lean_object* x_873; +x_869 = 1; +x_870 = lean_box(x_869); +x_871 = lean_box(x_869); +x_872 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTerm___boxed), 11, 4); +lean_closure_set(x_872, 0, x_1); +lean_closure_set(x_872, 1, x_5); +lean_closure_set(x_872, 2, x_870); +lean_closure_set(x_872, 3, x_871); +x_873 = l_Lean_Elab_Term_observing___rarg(x_872, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_873) == 0) { -uint8_t x_877; -x_877 = !lean_is_exclusive(x_876); -if (x_877 == 0) +uint8_t x_874; +x_874 = !lean_is_exclusive(x_873); +if (x_874 == 0) { -lean_object* x_878; lean_object* x_879; -x_878 = lean_ctor_get(x_876, 0); -x_879 = lean_array_push(x_9, x_878); -lean_ctor_set(x_876, 0, x_879); -return x_876; +lean_object* x_875; lean_object* x_876; +x_875 = lean_ctor_get(x_873, 0); +x_876 = lean_array_push(x_9, x_875); +lean_ctor_set(x_873, 0, x_876); +return x_873; } else { -lean_object* x_880; lean_object* x_881; lean_object* x_882; lean_object* x_883; -x_880 = lean_ctor_get(x_876, 0); -x_881 = lean_ctor_get(x_876, 1); -lean_inc(x_881); -lean_inc(x_880); -lean_dec(x_876); -x_882 = lean_array_push(x_9, x_880); -x_883 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_883, 0, x_882); -lean_ctor_set(x_883, 1, x_881); -return x_883; +lean_object* x_877; lean_object* x_878; lean_object* x_879; lean_object* x_880; +x_877 = lean_ctor_get(x_873, 0); +x_878 = lean_ctor_get(x_873, 1); +lean_inc(x_878); +lean_inc(x_877); +lean_dec(x_873); +x_879 = lean_array_push(x_9, x_877); +x_880 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_880, 0, x_879); +lean_ctor_set(x_880, 1, x_878); +return x_880; } } else { -uint8_t x_884; +uint8_t x_881; lean_dec(x_9); -x_884 = !lean_is_exclusive(x_876); -if (x_884 == 0) +x_881 = !lean_is_exclusive(x_873); +if (x_881 == 0) { -return x_876; +return x_873; } else { -lean_object* x_885; lean_object* x_886; lean_object* x_887; -x_885 = lean_ctor_get(x_876, 0); -x_886 = lean_ctor_get(x_876, 1); -lean_inc(x_886); -lean_inc(x_885); -lean_dec(x_876); -x_887 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_887, 0, x_885); -lean_ctor_set(x_887, 1, x_886); -return x_887; +lean_object* x_882; lean_object* x_883; lean_object* x_884; +x_882 = lean_ctor_get(x_873, 0); +x_883 = lean_ctor_get(x_873, 1); +lean_inc(x_883); +lean_inc(x_882); +lean_dec(x_873); +x_884 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_884, 0, x_882); +lean_ctor_set(x_884, 1, x_883); +return x_884; } } } else { -lean_object* x_888; uint8_t x_889; lean_object* x_890; lean_object* x_891; lean_object* x_892; lean_object* x_893; -x_888 = lean_box(0); -x_889 = 1; -x_890 = lean_box(x_815); -x_891 = lean_box(x_889); -x_892 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTermEnsuringType___boxed), 12, 5); -lean_closure_set(x_892, 0, x_1); -lean_closure_set(x_892, 1, x_5); -lean_closure_set(x_892, 2, x_890); -lean_closure_set(x_892, 3, x_891); -lean_closure_set(x_892, 4, x_888); -x_893 = l_Lean_Elab_Term_observing___rarg(x_892, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_893) == 0) +lean_object* x_885; uint8_t x_886; lean_object* x_887; lean_object* x_888; lean_object* x_889; lean_object* x_890; +x_885 = lean_box(0); +x_886 = 1; +x_887 = lean_box(x_812); +x_888 = lean_box(x_886); +x_889 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTermEnsuringType___boxed), 12, 5); +lean_closure_set(x_889, 0, x_1); +lean_closure_set(x_889, 1, x_5); +lean_closure_set(x_889, 2, x_887); +lean_closure_set(x_889, 3, x_888); +lean_closure_set(x_889, 4, x_885); +x_890 = l_Lean_Elab_Term_observing___rarg(x_889, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_890) == 0) { -uint8_t x_894; -x_894 = !lean_is_exclusive(x_893); -if (x_894 == 0) +uint8_t x_891; +x_891 = !lean_is_exclusive(x_890); +if (x_891 == 0) { -lean_object* x_895; lean_object* x_896; -x_895 = lean_ctor_get(x_893, 0); -x_896 = lean_array_push(x_9, x_895); -lean_ctor_set(x_893, 0, x_896); -return x_893; +lean_object* x_892; lean_object* x_893; +x_892 = lean_ctor_get(x_890, 0); +x_893 = lean_array_push(x_9, x_892); +lean_ctor_set(x_890, 0, x_893); +return x_890; } else { -lean_object* x_897; lean_object* x_898; lean_object* x_899; lean_object* x_900; -x_897 = lean_ctor_get(x_893, 0); -x_898 = lean_ctor_get(x_893, 1); -lean_inc(x_898); -lean_inc(x_897); -lean_dec(x_893); -x_899 = lean_array_push(x_9, x_897); -x_900 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_900, 0, x_899); -lean_ctor_set(x_900, 1, x_898); -return x_900; +lean_object* x_894; lean_object* x_895; lean_object* x_896; lean_object* x_897; +x_894 = lean_ctor_get(x_890, 0); +x_895 = lean_ctor_get(x_890, 1); +lean_inc(x_895); +lean_inc(x_894); +lean_dec(x_890); +x_896 = lean_array_push(x_9, x_894); +x_897 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_897, 0, x_896); +lean_ctor_set(x_897, 1, x_895); +return x_897; } } else { -uint8_t x_901; +uint8_t x_898; lean_dec(x_9); -x_901 = !lean_is_exclusive(x_893); -if (x_901 == 0) +x_898 = !lean_is_exclusive(x_890); +if (x_898 == 0) { -return x_893; +return x_890; } else { -lean_object* x_902; lean_object* x_903; lean_object* x_904; -x_902 = lean_ctor_get(x_893, 0); -x_903 = lean_ctor_get(x_893, 1); -lean_inc(x_903); -lean_inc(x_902); -lean_dec(x_893); -x_904 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_904, 0, x_902); -lean_ctor_set(x_904, 1, x_903); -return x_904; +lean_object* x_899; lean_object* x_900; lean_object* x_901; +x_899 = lean_ctor_get(x_890, 0); +x_900 = lean_ctor_get(x_890, 1); +lean_inc(x_900); +lean_inc(x_899); +lean_dec(x_890); +x_901 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_901, 0, x_899); +lean_ctor_set(x_901, 1, x_900); +return x_901; } } } @@ -25183,301 +25595,300 @@ return x_904; } else { -lean_object* x_908; lean_object* x_909; +lean_object* x_905; lean_dec(x_1); -x_908 = l_Lean_fieldIdxKind; -x_909 = l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(x_908, x_607); -if (lean_obj_tag(x_909) == 0) +x_905 = l_Lean_Syntax_isFieldIdx_x3f(x_606); +if (lean_obj_tag(x_905) == 0) +{ +lean_object* x_906; lean_object* x_907; lean_object* x_908; lean_object* x_909; +x_906 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__28; +x_907 = l_panic___at_Lean_TSyntax_getNat___spec__1(x_906); +x_908 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_908, 0, x_606); +lean_ctor_set(x_908, 1, x_907); +x_909 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_909, 0, x_908); +lean_ctor_set(x_909, 1, x_2); +x_1 = x_604; +x_2 = x_909; +goto _start; +} +else { -lean_object* x_910; lean_object* x_911; lean_object* x_912; lean_object* x_913; -x_910 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__26; -x_911 = l_panic___at_String_toNat_x21___spec__1(x_910); +lean_object* x_911; lean_object* x_912; lean_object* x_913; +x_911 = lean_ctor_get(x_905, 0); +lean_inc(x_911); +lean_dec(x_905); x_912 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_912, 0, x_607); +lean_ctor_set(x_912, 0, x_606); lean_ctor_set(x_912, 1, x_911); x_913 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_913, 0, x_912); lean_ctor_set(x_913, 1, x_2); -x_1 = x_605; +x_1 = x_604; x_2 = x_913; goto _start; } -else -{ -lean_object* x_915; lean_object* x_916; lean_object* x_917; -x_915 = lean_ctor_get(x_909, 0); -lean_inc(x_915); -lean_dec(x_909); -x_916 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_916, 0, x_607); -lean_ctor_set(x_916, 1, x_915); -x_917 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_917, 0, x_916); -lean_ctor_set(x_917, 1, x_2); -x_1 = x_605; -x_2 = x_917; -goto _start; -} } } } } else { -lean_object* x_919; lean_object* x_920; lean_object* x_921; lean_object* x_922; lean_object* x_923; uint8_t x_924; -x_919 = lean_unsigned_to_nat(0u); -x_920 = l_Lean_Syntax_getArg(x_1, x_919); -x_921 = lean_unsigned_to_nat(2u); -x_922 = l_Lean_Syntax_getArg(x_1, x_921); -x_923 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__22; -lean_inc(x_922); -x_924 = l_Lean_Syntax_isOfKind(x_922, x_923); -if (x_924 == 0) +lean_object* x_915; lean_object* x_916; lean_object* x_917; lean_object* x_918; lean_object* x_919; uint8_t x_920; +x_915 = lean_unsigned_to_nat(0u); +x_916 = l_Lean_Syntax_getArg(x_1, x_915); +x_917 = lean_unsigned_to_nat(2u); +x_918 = l_Lean_Syntax_getArg(x_1, x_917); +x_919 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__24; +lean_inc(x_918); +x_920 = l_Lean_Syntax_isOfKind(x_918, x_919); +if (x_920 == 0) { -lean_object* x_925; uint8_t x_926; -x_925 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__10; -lean_inc(x_922); -x_926 = l_Lean_Syntax_isOfKind(x_922, x_925); -if (x_926 == 0) +lean_object* x_921; uint8_t x_922; +x_921 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__12; +lean_inc(x_918); +x_922 = l_Lean_Syntax_isOfKind(x_918, x_921); +if (x_922 == 0) { -uint8_t x_927; uint8_t x_928; -lean_dec(x_922); -lean_dec(x_920); -x_927 = l_List_isEmpty___rarg(x_2); +uint8_t x_923; uint8_t x_924; +lean_dec(x_918); +lean_dec(x_916); +x_923 = l_List_isEmpty___rarg(x_2); if (x_8 == 0) { -uint8_t x_1019; -x_1019 = 1; -x_928 = x_1019; -goto block_1018; -} -else -{ -uint8_t x_1020; -x_1020 = 0; -x_928 = x_1020; -goto block_1018; -} -block_1018: -{ -if (x_927 == 0) -{ -lean_object* x_929; lean_object* x_930; lean_object* x_931; lean_object* x_932; lean_object* x_933; lean_object* x_934; lean_object* x_935; -x_929 = lean_box(0); -x_930 = lean_box(x_928); -x_931 = lean_box(x_6); -x_932 = lean_box(x_7); -x_933 = lean_box(x_8); -x_934 = lean_alloc_closure((void*)(l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___lambda__1___boxed), 17, 10); -lean_closure_set(x_934, 0, x_1); -lean_closure_set(x_934, 1, x_929); -lean_closure_set(x_934, 2, x_930); -lean_closure_set(x_934, 3, x_2); -lean_closure_set(x_934, 4, x_3); -lean_closure_set(x_934, 5, x_4); -lean_closure_set(x_934, 6, x_5); -lean_closure_set(x_934, 7, x_931); -lean_closure_set(x_934, 8, x_932); -lean_closure_set(x_934, 9, x_933); -x_935 = l_Lean_Elab_Term_observing___rarg(x_934, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_935) == 0) -{ -uint8_t x_936; -x_936 = !lean_is_exclusive(x_935); -if (x_936 == 0) -{ -lean_object* x_937; lean_object* x_938; -x_937 = lean_ctor_get(x_935, 0); -x_938 = lean_array_push(x_9, x_937); -lean_ctor_set(x_935, 0, x_938); -return x_935; -} -else -{ -lean_object* x_939; lean_object* x_940; lean_object* x_941; lean_object* x_942; -x_939 = lean_ctor_get(x_935, 0); -x_940 = lean_ctor_get(x_935, 1); -lean_inc(x_940); -lean_inc(x_939); -lean_dec(x_935); -x_941 = lean_array_push(x_9, x_939); -x_942 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_942, 0, x_941); -lean_ctor_set(x_942, 1, x_940); -return x_942; -} -} -else -{ -uint8_t x_943; -lean_dec(x_9); -x_943 = !lean_is_exclusive(x_935); -if (x_943 == 0) -{ -return x_935; +uint8_t x_1015; +x_1015 = 1; +x_924 = x_1015; +goto block_1014; } else { -lean_object* x_944; lean_object* x_945; lean_object* x_946; -x_944 = lean_ctor_get(x_935, 0); -x_945 = lean_ctor_get(x_935, 1); -lean_inc(x_945); -lean_inc(x_944); -lean_dec(x_935); -x_946 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_946, 0, x_944); -lean_ctor_set(x_946, 1, x_945); -return x_946; -} +uint8_t x_1016; +x_1016 = 0; +x_924 = x_1016; +goto block_1014; } -} -else +block_1014: { -uint8_t x_947; -x_947 = l_Array_isEmpty___rarg(x_3); -if (x_947 == 0) +if (x_923 == 0) { -lean_object* x_948; lean_object* x_949; lean_object* x_950; lean_object* x_951; lean_object* x_952; lean_object* x_953; lean_object* x_954; -x_948 = lean_box(0); -x_949 = lean_box(x_928); -x_950 = lean_box(x_6); -x_951 = lean_box(x_7); -x_952 = lean_box(x_8); -x_953 = lean_alloc_closure((void*)(l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___lambda__1___boxed), 17, 10); -lean_closure_set(x_953, 0, x_1); -lean_closure_set(x_953, 1, x_948); -lean_closure_set(x_953, 2, x_949); -lean_closure_set(x_953, 3, x_2); -lean_closure_set(x_953, 4, x_3); -lean_closure_set(x_953, 5, x_4); -lean_closure_set(x_953, 6, x_5); -lean_closure_set(x_953, 7, x_950); -lean_closure_set(x_953, 8, x_951); -lean_closure_set(x_953, 9, x_952); -x_954 = l_Lean_Elab_Term_observing___rarg(x_953, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_954) == 0) +lean_object* x_925; lean_object* x_926; lean_object* x_927; lean_object* x_928; lean_object* x_929; lean_object* x_930; lean_object* x_931; +x_925 = lean_box(0); +x_926 = lean_box(x_924); +x_927 = lean_box(x_6); +x_928 = lean_box(x_7); +x_929 = lean_box(x_8); +x_930 = lean_alloc_closure((void*)(l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___lambda__1___boxed), 17, 10); +lean_closure_set(x_930, 0, x_1); +lean_closure_set(x_930, 1, x_925); +lean_closure_set(x_930, 2, x_926); +lean_closure_set(x_930, 3, x_2); +lean_closure_set(x_930, 4, x_3); +lean_closure_set(x_930, 5, x_4); +lean_closure_set(x_930, 6, x_5); +lean_closure_set(x_930, 7, x_927); +lean_closure_set(x_930, 8, x_928); +lean_closure_set(x_930, 9, x_929); +x_931 = l_Lean_Elab_Term_observing___rarg(x_930, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_931) == 0) { -uint8_t x_955; -x_955 = !lean_is_exclusive(x_954); -if (x_955 == 0) +uint8_t x_932; +x_932 = !lean_is_exclusive(x_931); +if (x_932 == 0) { -lean_object* x_956; lean_object* x_957; -x_956 = lean_ctor_get(x_954, 0); -x_957 = lean_array_push(x_9, x_956); -lean_ctor_set(x_954, 0, x_957); -return x_954; +lean_object* x_933; lean_object* x_934; +x_933 = lean_ctor_get(x_931, 0); +x_934 = lean_array_push(x_9, x_933); +lean_ctor_set(x_931, 0, x_934); +return x_931; } else { -lean_object* x_958; lean_object* x_959; lean_object* x_960; lean_object* x_961; -x_958 = lean_ctor_get(x_954, 0); -x_959 = lean_ctor_get(x_954, 1); -lean_inc(x_959); -lean_inc(x_958); -lean_dec(x_954); -x_960 = lean_array_push(x_9, x_958); -x_961 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_961, 0, x_960); -lean_ctor_set(x_961, 1, x_959); -return x_961; +lean_object* x_935; lean_object* x_936; lean_object* x_937; lean_object* x_938; +x_935 = lean_ctor_get(x_931, 0); +x_936 = lean_ctor_get(x_931, 1); +lean_inc(x_936); +lean_inc(x_935); +lean_dec(x_931); +x_937 = lean_array_push(x_9, x_935); +x_938 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_938, 0, x_937); +lean_ctor_set(x_938, 1, x_936); +return x_938; } } else { -uint8_t x_962; +uint8_t x_939; lean_dec(x_9); -x_962 = !lean_is_exclusive(x_954); -if (x_962 == 0) +x_939 = !lean_is_exclusive(x_931); +if (x_939 == 0) { -return x_954; +return x_931; } else { -lean_object* x_963; lean_object* x_964; lean_object* x_965; -x_963 = lean_ctor_get(x_954, 0); -x_964 = lean_ctor_get(x_954, 1); -lean_inc(x_964); -lean_inc(x_963); -lean_dec(x_954); -x_965 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_965, 0, x_963); -lean_ctor_set(x_965, 1, x_964); -return x_965; +lean_object* x_940; lean_object* x_941; lean_object* x_942; +x_940 = lean_ctor_get(x_931, 0); +x_941 = lean_ctor_get(x_931, 1); +lean_inc(x_941); +lean_inc(x_940); +lean_dec(x_931); +x_942 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_942, 0, x_940); +lean_ctor_set(x_942, 1, x_941); +return x_942; } } } else { -uint8_t x_966; -x_966 = l_Array_isEmpty___rarg(x_4); -if (x_966 == 0) -{ -lean_object* x_967; lean_object* x_968; lean_object* x_969; lean_object* x_970; lean_object* x_971; lean_object* x_972; lean_object* x_973; -x_967 = lean_box(0); -x_968 = lean_box(x_928); -x_969 = lean_box(x_6); -x_970 = lean_box(x_7); -x_971 = lean_box(x_8); -x_972 = lean_alloc_closure((void*)(l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___lambda__1___boxed), 17, 10); -lean_closure_set(x_972, 0, x_1); -lean_closure_set(x_972, 1, x_967); -lean_closure_set(x_972, 2, x_968); -lean_closure_set(x_972, 3, x_2); -lean_closure_set(x_972, 4, x_3); -lean_closure_set(x_972, 5, x_4); -lean_closure_set(x_972, 6, x_5); -lean_closure_set(x_972, 7, x_969); -lean_closure_set(x_972, 8, x_970); -lean_closure_set(x_972, 9, x_971); -x_973 = l_Lean_Elab_Term_observing___rarg(x_972, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_973) == 0) +uint8_t x_943; +x_943 = l_Array_isEmpty___rarg(x_3); +if (x_943 == 0) { -uint8_t x_974; -x_974 = !lean_is_exclusive(x_973); -if (x_974 == 0) +lean_object* x_944; lean_object* x_945; lean_object* x_946; lean_object* x_947; lean_object* x_948; lean_object* x_949; lean_object* x_950; +x_944 = lean_box(0); +x_945 = lean_box(x_924); +x_946 = lean_box(x_6); +x_947 = lean_box(x_7); +x_948 = lean_box(x_8); +x_949 = lean_alloc_closure((void*)(l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___lambda__1___boxed), 17, 10); +lean_closure_set(x_949, 0, x_1); +lean_closure_set(x_949, 1, x_944); +lean_closure_set(x_949, 2, x_945); +lean_closure_set(x_949, 3, x_2); +lean_closure_set(x_949, 4, x_3); +lean_closure_set(x_949, 5, x_4); +lean_closure_set(x_949, 6, x_5); +lean_closure_set(x_949, 7, x_946); +lean_closure_set(x_949, 8, x_947); +lean_closure_set(x_949, 9, x_948); +x_950 = l_Lean_Elab_Term_observing___rarg(x_949, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_950) == 0) +{ +uint8_t x_951; +x_951 = !lean_is_exclusive(x_950); +if (x_951 == 0) +{ +lean_object* x_952; lean_object* x_953; +x_952 = lean_ctor_get(x_950, 0); +x_953 = lean_array_push(x_9, x_952); +lean_ctor_set(x_950, 0, x_953); +return x_950; +} +else +{ +lean_object* x_954; lean_object* x_955; lean_object* x_956; lean_object* x_957; +x_954 = lean_ctor_get(x_950, 0); +x_955 = lean_ctor_get(x_950, 1); +lean_inc(x_955); +lean_inc(x_954); +lean_dec(x_950); +x_956 = lean_array_push(x_9, x_954); +x_957 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_957, 0, x_956); +lean_ctor_set(x_957, 1, x_955); +return x_957; +} +} +else +{ +uint8_t x_958; +lean_dec(x_9); +x_958 = !lean_is_exclusive(x_950); +if (x_958 == 0) { -lean_object* x_975; lean_object* x_976; -x_975 = lean_ctor_get(x_973, 0); -x_976 = lean_array_push(x_9, x_975); -lean_ctor_set(x_973, 0, x_976); -return x_973; +return x_950; } else { -lean_object* x_977; lean_object* x_978; lean_object* x_979; lean_object* x_980; -x_977 = lean_ctor_get(x_973, 0); -x_978 = lean_ctor_get(x_973, 1); -lean_inc(x_978); -lean_inc(x_977); -lean_dec(x_973); -x_979 = lean_array_push(x_9, x_977); -x_980 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_980, 0, x_979); -lean_ctor_set(x_980, 1, x_978); -return x_980; +lean_object* x_959; lean_object* x_960; lean_object* x_961; +x_959 = lean_ctor_get(x_950, 0); +x_960 = lean_ctor_get(x_950, 1); +lean_inc(x_960); +lean_inc(x_959); +lean_dec(x_950); +x_961 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_961, 0, x_959); +lean_ctor_set(x_961, 1, x_960); +return x_961; +} } } else { -uint8_t x_981; +uint8_t x_962; +x_962 = l_Array_isEmpty___rarg(x_4); +if (x_962 == 0) +{ +lean_object* x_963; lean_object* x_964; lean_object* x_965; lean_object* x_966; lean_object* x_967; lean_object* x_968; lean_object* x_969; +x_963 = lean_box(0); +x_964 = lean_box(x_924); +x_965 = lean_box(x_6); +x_966 = lean_box(x_7); +x_967 = lean_box(x_8); +x_968 = lean_alloc_closure((void*)(l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___lambda__1___boxed), 17, 10); +lean_closure_set(x_968, 0, x_1); +lean_closure_set(x_968, 1, x_963); +lean_closure_set(x_968, 2, x_964); +lean_closure_set(x_968, 3, x_2); +lean_closure_set(x_968, 4, x_3); +lean_closure_set(x_968, 5, x_4); +lean_closure_set(x_968, 6, x_5); +lean_closure_set(x_968, 7, x_965); +lean_closure_set(x_968, 8, x_966); +lean_closure_set(x_968, 9, x_967); +x_969 = l_Lean_Elab_Term_observing___rarg(x_968, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_969) == 0) +{ +uint8_t x_970; +x_970 = !lean_is_exclusive(x_969); +if (x_970 == 0) +{ +lean_object* x_971; lean_object* x_972; +x_971 = lean_ctor_get(x_969, 0); +x_972 = lean_array_push(x_9, x_971); +lean_ctor_set(x_969, 0, x_972); +return x_969; +} +else +{ +lean_object* x_973; lean_object* x_974; lean_object* x_975; lean_object* x_976; +x_973 = lean_ctor_get(x_969, 0); +x_974 = lean_ctor_get(x_969, 1); +lean_inc(x_974); +lean_inc(x_973); +lean_dec(x_969); +x_975 = lean_array_push(x_9, x_973); +x_976 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_976, 0, x_975); +lean_ctor_set(x_976, 1, x_974); +return x_976; +} +} +else +{ +uint8_t x_977; lean_dec(x_9); -x_981 = !lean_is_exclusive(x_973); -if (x_981 == 0) +x_977 = !lean_is_exclusive(x_969); +if (x_977 == 0) { -return x_973; +return x_969; } else { -lean_object* x_982; lean_object* x_983; lean_object* x_984; -x_982 = lean_ctor_get(x_973, 0); -x_983 = lean_ctor_get(x_973, 1); -lean_inc(x_983); -lean_inc(x_982); -lean_dec(x_973); -x_984 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_984, 0, x_982); -lean_ctor_set(x_984, 1, x_983); -return x_984; +lean_object* x_978; lean_object* x_979; lean_object* x_980; +x_978 = lean_ctor_get(x_969, 0); +x_979 = lean_ctor_get(x_969, 1); +lean_inc(x_979); +lean_inc(x_978); +lean_dec(x_969); +x_980 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_980, 0, x_978); +lean_ctor_set(x_980, 1, x_979); +return x_980; } } } @@ -25488,129 +25899,129 @@ lean_dec(x_3); lean_dec(x_2); if (x_8 == 0) { -uint8_t x_985; lean_object* x_986; lean_object* x_987; lean_object* x_988; lean_object* x_989; -x_985 = 1; -x_986 = lean_box(x_985); -x_987 = lean_box(x_985); -x_988 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTerm___boxed), 11, 4); -lean_closure_set(x_988, 0, x_1); -lean_closure_set(x_988, 1, x_5); -lean_closure_set(x_988, 2, x_986); -lean_closure_set(x_988, 3, x_987); -x_989 = l_Lean_Elab_Term_observing___rarg(x_988, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_989) == 0) -{ -uint8_t x_990; -x_990 = !lean_is_exclusive(x_989); -if (x_990 == 0) -{ -lean_object* x_991; lean_object* x_992; -x_991 = lean_ctor_get(x_989, 0); -x_992 = lean_array_push(x_9, x_991); -lean_ctor_set(x_989, 0, x_992); -return x_989; -} -else -{ -lean_object* x_993; lean_object* x_994; lean_object* x_995; lean_object* x_996; -x_993 = lean_ctor_get(x_989, 0); -x_994 = lean_ctor_get(x_989, 1); -lean_inc(x_994); -lean_inc(x_993); -lean_dec(x_989); -x_995 = lean_array_push(x_9, x_993); -x_996 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_996, 0, x_995); -lean_ctor_set(x_996, 1, x_994); -return x_996; +uint8_t x_981; lean_object* x_982; lean_object* x_983; lean_object* x_984; lean_object* x_985; +x_981 = 1; +x_982 = lean_box(x_981); +x_983 = lean_box(x_981); +x_984 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTerm___boxed), 11, 4); +lean_closure_set(x_984, 0, x_1); +lean_closure_set(x_984, 1, x_5); +lean_closure_set(x_984, 2, x_982); +lean_closure_set(x_984, 3, x_983); +x_985 = l_Lean_Elab_Term_observing___rarg(x_984, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_985) == 0) +{ +uint8_t x_986; +x_986 = !lean_is_exclusive(x_985); +if (x_986 == 0) +{ +lean_object* x_987; lean_object* x_988; +x_987 = lean_ctor_get(x_985, 0); +x_988 = lean_array_push(x_9, x_987); +lean_ctor_set(x_985, 0, x_988); +return x_985; +} +else +{ +lean_object* x_989; lean_object* x_990; lean_object* x_991; lean_object* x_992; +x_989 = lean_ctor_get(x_985, 0); +x_990 = lean_ctor_get(x_985, 1); +lean_inc(x_990); +lean_inc(x_989); +lean_dec(x_985); +x_991 = lean_array_push(x_9, x_989); +x_992 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_992, 0, x_991); +lean_ctor_set(x_992, 1, x_990); +return x_992; } } else { -uint8_t x_997; +uint8_t x_993; lean_dec(x_9); -x_997 = !lean_is_exclusive(x_989); -if (x_997 == 0) +x_993 = !lean_is_exclusive(x_985); +if (x_993 == 0) { -return x_989; +return x_985; } else { -lean_object* x_998; lean_object* x_999; lean_object* x_1000; -x_998 = lean_ctor_get(x_989, 0); -x_999 = lean_ctor_get(x_989, 1); -lean_inc(x_999); -lean_inc(x_998); -lean_dec(x_989); -x_1000 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1000, 0, x_998); -lean_ctor_set(x_1000, 1, x_999); -return x_1000; +lean_object* x_994; lean_object* x_995; lean_object* x_996; +x_994 = lean_ctor_get(x_985, 0); +x_995 = lean_ctor_get(x_985, 1); +lean_inc(x_995); +lean_inc(x_994); +lean_dec(x_985); +x_996 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_996, 0, x_994); +lean_ctor_set(x_996, 1, x_995); +return x_996; } } } else { -lean_object* x_1001; uint8_t x_1002; lean_object* x_1003; lean_object* x_1004; lean_object* x_1005; lean_object* x_1006; -x_1001 = lean_box(0); -x_1002 = 1; -x_1003 = lean_box(x_928); -x_1004 = lean_box(x_1002); -x_1005 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTermEnsuringType___boxed), 12, 5); -lean_closure_set(x_1005, 0, x_1); -lean_closure_set(x_1005, 1, x_5); -lean_closure_set(x_1005, 2, x_1003); -lean_closure_set(x_1005, 3, x_1004); -lean_closure_set(x_1005, 4, x_1001); -x_1006 = l_Lean_Elab_Term_observing___rarg(x_1005, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_1006) == 0) +lean_object* x_997; uint8_t x_998; lean_object* x_999; lean_object* x_1000; lean_object* x_1001; lean_object* x_1002; +x_997 = lean_box(0); +x_998 = 1; +x_999 = lean_box(x_924); +x_1000 = lean_box(x_998); +x_1001 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTermEnsuringType___boxed), 12, 5); +lean_closure_set(x_1001, 0, x_1); +lean_closure_set(x_1001, 1, x_5); +lean_closure_set(x_1001, 2, x_999); +lean_closure_set(x_1001, 3, x_1000); +lean_closure_set(x_1001, 4, x_997); +x_1002 = l_Lean_Elab_Term_observing___rarg(x_1001, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_1002) == 0) { -uint8_t x_1007; -x_1007 = !lean_is_exclusive(x_1006); -if (x_1007 == 0) +uint8_t x_1003; +x_1003 = !lean_is_exclusive(x_1002); +if (x_1003 == 0) { -lean_object* x_1008; lean_object* x_1009; -x_1008 = lean_ctor_get(x_1006, 0); -x_1009 = lean_array_push(x_9, x_1008); -lean_ctor_set(x_1006, 0, x_1009); -return x_1006; +lean_object* x_1004; lean_object* x_1005; +x_1004 = lean_ctor_get(x_1002, 0); +x_1005 = lean_array_push(x_9, x_1004); +lean_ctor_set(x_1002, 0, x_1005); +return x_1002; } else { -lean_object* x_1010; lean_object* x_1011; lean_object* x_1012; lean_object* x_1013; -x_1010 = lean_ctor_get(x_1006, 0); -x_1011 = lean_ctor_get(x_1006, 1); -lean_inc(x_1011); -lean_inc(x_1010); -lean_dec(x_1006); -x_1012 = lean_array_push(x_9, x_1010); -x_1013 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1013, 0, x_1012); -lean_ctor_set(x_1013, 1, x_1011); -return x_1013; +lean_object* x_1006; lean_object* x_1007; lean_object* x_1008; lean_object* x_1009; +x_1006 = lean_ctor_get(x_1002, 0); +x_1007 = lean_ctor_get(x_1002, 1); +lean_inc(x_1007); +lean_inc(x_1006); +lean_dec(x_1002); +x_1008 = lean_array_push(x_9, x_1006); +x_1009 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1009, 0, x_1008); +lean_ctor_set(x_1009, 1, x_1007); +return x_1009; } } else { -uint8_t x_1014; +uint8_t x_1010; lean_dec(x_9); -x_1014 = !lean_is_exclusive(x_1006); -if (x_1014 == 0) +x_1010 = !lean_is_exclusive(x_1002); +if (x_1010 == 0) { -return x_1006; +return x_1002; } else { -lean_object* x_1015; lean_object* x_1016; lean_object* x_1017; -x_1015 = lean_ctor_get(x_1006, 0); -x_1016 = lean_ctor_get(x_1006, 1); -lean_inc(x_1016); -lean_inc(x_1015); -lean_dec(x_1006); -x_1017 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1017, 0, x_1015); -lean_ctor_set(x_1017, 1, x_1016); -return x_1017; +lean_object* x_1011; lean_object* x_1012; lean_object* x_1013; +x_1011 = lean_ctor_get(x_1002, 0); +x_1012 = lean_ctor_get(x_1002, 1); +lean_inc(x_1012); +lean_inc(x_1011); +lean_dec(x_1002); +x_1013 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1013, 0, x_1011); +lean_ctor_set(x_1013, 1, x_1012); +return x_1013; } } } @@ -25621,54 +26032,53 @@ return x_1017; } else { -lean_object* x_1021; lean_object* x_1022; lean_object* x_1023; lean_object* x_1024; lean_object* x_1025; +lean_object* x_1017; lean_object* x_1018; lean_object* x_1019; lean_object* x_1020; lean_object* x_1021; lean_dec(x_1); -x_1021 = lean_box(0); -x_1022 = l_Lean_Syntax_identComponents(x_922, x_1021); -x_1023 = lean_box(0); -lean_inc(x_920); -x_1024 = l_List_mapTRAux___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___spec__4(x_920, x_1022, x_1023); -x_1025 = l_List_appendTR___rarg(x_1024, x_2); -x_1 = x_920; -x_2 = x_1025; +x_1017 = lean_box(0); +x_1018 = l_Lean_Syntax_identComponents(x_918, x_1017); +x_1019 = lean_box(0); +lean_inc(x_916); +x_1020 = l_List_mapTRAux___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___spec__4(x_916, x_1018, x_1019); +x_1021 = l_List_appendTR___rarg(x_1020, x_2); +x_1 = x_916; +x_2 = x_1021; goto _start; } } else { -lean_object* x_1027; lean_object* x_1028; +lean_object* x_1023; lean_dec(x_1); -x_1027 = l_Lean_fieldIdxKind; -x_1028 = l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(x_1027, x_922); -if (lean_obj_tag(x_1028) == 0) -{ -lean_object* x_1029; lean_object* x_1030; lean_object* x_1031; lean_object* x_1032; -x_1029 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__26; -x_1030 = l_panic___at_String_toNat_x21___spec__1(x_1029); -x_1031 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1031, 0, x_922); -lean_ctor_set(x_1031, 1, x_1030); -x_1032 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1032, 0, x_1031); -lean_ctor_set(x_1032, 1, x_2); -x_1 = x_920; -x_2 = x_1032; +x_1023 = l_Lean_Syntax_isFieldIdx_x3f(x_918); +if (lean_obj_tag(x_1023) == 0) +{ +lean_object* x_1024; lean_object* x_1025; lean_object* x_1026; lean_object* x_1027; +x_1024 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__28; +x_1025 = l_panic___at_Lean_TSyntax_getNat___spec__1(x_1024); +x_1026 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1026, 0, x_918); +lean_ctor_set(x_1026, 1, x_1025); +x_1027 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1027, 0, x_1026); +lean_ctor_set(x_1027, 1, x_2); +x_1 = x_916; +x_2 = x_1027; goto _start; } else { -lean_object* x_1034; lean_object* x_1035; lean_object* x_1036; -x_1034 = lean_ctor_get(x_1028, 0); -lean_inc(x_1034); -lean_dec(x_1028); -x_1035 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1035, 0, x_922); -lean_ctor_set(x_1035, 1, x_1034); -x_1036 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1036, 0, x_1035); -lean_ctor_set(x_1036, 1, x_2); -x_1 = x_920; -x_2 = x_1036; +lean_object* x_1029; lean_object* x_1030; lean_object* x_1031; +x_1029 = lean_ctor_get(x_1023, 0); +lean_inc(x_1029); +lean_dec(x_1023); +x_1030 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1030, 0, x_918); +lean_ctor_set(x_1030, 1, x_1029); +x_1031 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1031, 0, x_1030); +lean_ctor_set(x_1031, 1, x_2); +x_1 = x_916; +x_2 = x_1031; goto _start; } } @@ -25676,17 +26086,17 @@ goto _start; } else { -lean_object* x_1038; lean_object* x_1039; lean_object* x_1040; uint8_t x_1041; -x_1038 = l_Lean_Syntax_getArgs(x_1); +lean_object* x_1033; lean_object* x_1034; lean_object* x_1035; uint8_t x_1036; +x_1033 = l_Lean_Syntax_getArgs(x_1); lean_dec(x_1); -x_1039 = lean_array_get_size(x_1038); -x_1040 = lean_unsigned_to_nat(0u); -x_1041 = lean_nat_dec_lt(x_1040, x_1039); -if (x_1041 == 0) -{ -lean_object* x_1042; -lean_dec(x_1039); -lean_dec(x_1038); +x_1034 = lean_array_get_size(x_1033); +x_1035 = lean_unsigned_to_nat(0u); +x_1036 = lean_nat_dec_lt(x_1035, x_1034); +if (x_1036 == 0) +{ +lean_object* x_1037; +lean_dec(x_1034); +lean_dec(x_1033); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); @@ -25697,27 +26107,27 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_1042 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1042, 0, x_9); -lean_ctor_set(x_1042, 1, x_16); -return x_1042; +x_1037 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1037, 0, x_9); +lean_ctor_set(x_1037, 1, x_16); +return x_1037; } else { -uint8_t x_1043; -x_1043 = !lean_is_exclusive(x_10); -if (x_1043 == 0) +uint8_t x_1038; +x_1038 = !lean_is_exclusive(x_10); +if (x_1038 == 0) { -uint8_t x_1044; uint8_t x_1045; -x_1044 = 0; -lean_ctor_set_uint8(x_10, sizeof(void*)*7 + 1, x_1044); -x_1045 = lean_nat_dec_le(x_1039, x_1039); -if (x_1045 == 0) +uint8_t x_1039; uint8_t x_1040; +x_1039 = 0; +lean_ctor_set_uint8(x_10, sizeof(void*)*7 + 1, x_1039); +x_1040 = lean_nat_dec_le(x_1034, x_1034); +if (x_1040 == 0) { -lean_object* x_1046; +lean_object* x_1041; lean_dec(x_10); -lean_dec(x_1039); -lean_dec(x_1038); +lean_dec(x_1034); +lean_dec(x_1033); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); @@ -25727,71 +26137,71 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_1046 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1046, 0, x_9); -lean_ctor_set(x_1046, 1, x_16); -return x_1046; -} -else -{ -size_t x_1047; size_t x_1048; lean_object* x_1049; -x_1047 = 0; -x_1048 = lean_usize_of_nat(x_1039); -lean_dec(x_1039); -x_1049 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___spec__5(x_2, x_3, x_4, x_5, x_6, x_7, x_1038, x_1047, x_1048, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -lean_dec(x_1038); -return x_1049; -} -} -else -{ -lean_object* x_1050; lean_object* x_1051; uint8_t x_1052; uint8_t x_1053; lean_object* x_1054; lean_object* x_1055; lean_object* x_1056; lean_object* x_1057; uint8_t x_1058; uint8_t x_1059; uint8_t x_1060; uint8_t x_1061; lean_object* x_1062; uint8_t x_1063; uint8_t x_1064; lean_object* x_1065; uint8_t x_1066; -x_1050 = lean_ctor_get(x_10, 0); -x_1051 = lean_ctor_get(x_10, 1); -x_1052 = lean_ctor_get_uint8(x_10, sizeof(void*)*7); -x_1053 = lean_ctor_get_uint8(x_10, sizeof(void*)*7 + 2); -x_1054 = lean_ctor_get(x_10, 2); -x_1055 = lean_ctor_get(x_10, 3); -x_1056 = lean_ctor_get(x_10, 4); -x_1057 = lean_ctor_get(x_10, 5); -x_1058 = lean_ctor_get_uint8(x_10, sizeof(void*)*7 + 3); -x_1059 = lean_ctor_get_uint8(x_10, sizeof(void*)*7 + 4); -x_1060 = lean_ctor_get_uint8(x_10, sizeof(void*)*7 + 5); -x_1061 = lean_ctor_get_uint8(x_10, sizeof(void*)*7 + 6); -x_1062 = lean_ctor_get(x_10, 6); -x_1063 = lean_ctor_get_uint8(x_10, sizeof(void*)*7 + 7); -lean_inc(x_1062); +x_1041 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1041, 0, x_9); +lean_ctor_set(x_1041, 1, x_16); +return x_1041; +} +else +{ +size_t x_1042; size_t x_1043; lean_object* x_1044; +x_1042 = 0; +x_1043 = lean_usize_of_nat(x_1034); +lean_dec(x_1034); +x_1044 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___spec__5(x_2, x_3, x_4, x_5, x_6, x_7, x_1033, x_1042, x_1043, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_1033); +return x_1044; +} +} +else +{ +lean_object* x_1045; lean_object* x_1046; uint8_t x_1047; uint8_t x_1048; lean_object* x_1049; lean_object* x_1050; lean_object* x_1051; lean_object* x_1052; uint8_t x_1053; uint8_t x_1054; uint8_t x_1055; uint8_t x_1056; lean_object* x_1057; uint8_t x_1058; uint8_t x_1059; lean_object* x_1060; uint8_t x_1061; +x_1045 = lean_ctor_get(x_10, 0); +x_1046 = lean_ctor_get(x_10, 1); +x_1047 = lean_ctor_get_uint8(x_10, sizeof(void*)*7); +x_1048 = lean_ctor_get_uint8(x_10, sizeof(void*)*7 + 2); +x_1049 = lean_ctor_get(x_10, 2); +x_1050 = lean_ctor_get(x_10, 3); +x_1051 = lean_ctor_get(x_10, 4); +x_1052 = lean_ctor_get(x_10, 5); +x_1053 = lean_ctor_get_uint8(x_10, sizeof(void*)*7 + 3); +x_1054 = lean_ctor_get_uint8(x_10, sizeof(void*)*7 + 4); +x_1055 = lean_ctor_get_uint8(x_10, sizeof(void*)*7 + 5); +x_1056 = lean_ctor_get_uint8(x_10, sizeof(void*)*7 + 6); +x_1057 = lean_ctor_get(x_10, 6); +x_1058 = lean_ctor_get_uint8(x_10, sizeof(void*)*7 + 7); lean_inc(x_1057); -lean_inc(x_1056); -lean_inc(x_1055); -lean_inc(x_1054); +lean_inc(x_1052); lean_inc(x_1051); lean_inc(x_1050); +lean_inc(x_1049); +lean_inc(x_1046); +lean_inc(x_1045); lean_dec(x_10); -x_1064 = 0; -x_1065 = lean_alloc_ctor(0, 7, 8); -lean_ctor_set(x_1065, 0, x_1050); -lean_ctor_set(x_1065, 1, x_1051); -lean_ctor_set(x_1065, 2, x_1054); -lean_ctor_set(x_1065, 3, x_1055); -lean_ctor_set(x_1065, 4, x_1056); -lean_ctor_set(x_1065, 5, x_1057); -lean_ctor_set(x_1065, 6, x_1062); -lean_ctor_set_uint8(x_1065, sizeof(void*)*7, x_1052); -lean_ctor_set_uint8(x_1065, sizeof(void*)*7 + 1, x_1064); -lean_ctor_set_uint8(x_1065, sizeof(void*)*7 + 2, x_1053); -lean_ctor_set_uint8(x_1065, sizeof(void*)*7 + 3, x_1058); -lean_ctor_set_uint8(x_1065, sizeof(void*)*7 + 4, x_1059); -lean_ctor_set_uint8(x_1065, sizeof(void*)*7 + 5, x_1060); -lean_ctor_set_uint8(x_1065, sizeof(void*)*7 + 6, x_1061); -lean_ctor_set_uint8(x_1065, sizeof(void*)*7 + 7, x_1063); -x_1066 = lean_nat_dec_le(x_1039, x_1039); -if (x_1066 == 0) -{ -lean_object* x_1067; -lean_dec(x_1065); -lean_dec(x_1039); -lean_dec(x_1038); +x_1059 = 0; +x_1060 = lean_alloc_ctor(0, 7, 8); +lean_ctor_set(x_1060, 0, x_1045); +lean_ctor_set(x_1060, 1, x_1046); +lean_ctor_set(x_1060, 2, x_1049); +lean_ctor_set(x_1060, 3, x_1050); +lean_ctor_set(x_1060, 4, x_1051); +lean_ctor_set(x_1060, 5, x_1052); +lean_ctor_set(x_1060, 6, x_1057); +lean_ctor_set_uint8(x_1060, sizeof(void*)*7, x_1047); +lean_ctor_set_uint8(x_1060, sizeof(void*)*7 + 1, x_1059); +lean_ctor_set_uint8(x_1060, sizeof(void*)*7 + 2, x_1048); +lean_ctor_set_uint8(x_1060, sizeof(void*)*7 + 3, x_1053); +lean_ctor_set_uint8(x_1060, sizeof(void*)*7 + 4, x_1054); +lean_ctor_set_uint8(x_1060, sizeof(void*)*7 + 5, x_1055); +lean_ctor_set_uint8(x_1060, sizeof(void*)*7 + 6, x_1056); +lean_ctor_set_uint8(x_1060, sizeof(void*)*7 + 7, x_1058); +x_1061 = lean_nat_dec_le(x_1034, x_1034); +if (x_1061 == 0) +{ +lean_object* x_1062; +lean_dec(x_1060); +lean_dec(x_1034); +lean_dec(x_1033); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); @@ -25801,20 +26211,20 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_1067 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1067, 0, x_9); -lean_ctor_set(x_1067, 1, x_16); -return x_1067; +x_1062 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1062, 0, x_9); +lean_ctor_set(x_1062, 1, x_16); +return x_1062; } else { -size_t x_1068; size_t x_1069; lean_object* x_1070; -x_1068 = 0; -x_1069 = lean_usize_of_nat(x_1039); -lean_dec(x_1039); -x_1070 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___spec__5(x_2, x_3, x_4, x_5, x_6, x_7, x_1038, x_1068, x_1069, x_9, x_1065, x_11, x_12, x_13, x_14, x_15, x_16); -lean_dec(x_1038); -return x_1070; +size_t x_1063; size_t x_1064; lean_object* x_1065; +x_1063 = 0; +x_1064 = lean_usize_of_nat(x_1034); +lean_dec(x_1034); +x_1065 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___spec__5(x_2, x_3, x_4, x_5, x_6, x_7, x_1033, x_1063, x_1064, x_9, x_1060, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_1033); +return x_1065; } } } @@ -26696,7 +27106,7 @@ static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0_ 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___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppLValsAux_loop___lambda__3___closed__1; x_2 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__2___closed__1; -x_3 = lean_unsigned_to_nat(949u); +x_3 = lean_unsigned_to_nat(970u); x_4 = lean_unsigned_to_nat(57u); x_5 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppLValsAux_loop___lambda__3___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -27532,7 +27942,7 @@ static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0_ 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___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppLValsAux_loop___lambda__3___closed__1; x_2 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppAux___spec__4___closed__1; -x_3 = lean_unsigned_to_nat(964u); +x_3 = lean_unsigned_to_nat(985u); x_4 = lean_unsigned_to_nat(35u); x_5 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppLValsAux_loop___lambda__3___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -28489,7 +28899,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabApp_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(978u); +x_1 = lean_unsigned_to_nat(999u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28501,7 +28911,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabApp_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(981u); +x_1 = lean_unsigned_to_nat(1002u); x_2 = lean_unsigned_to_nat(90u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28529,7 +28939,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabApp_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(978u); +x_1 = lean_unsigned_to_nat(999u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28541,7 +28951,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabApp_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(978u); +x_1 = lean_unsigned_to_nat(999u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28687,7 +29097,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabIdent(lean_object* x_ { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_2 = l___regBuiltin_Lean_Elab_Term_elabApp___closed__6; -x_3 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__10; +x_3 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__12; x_4 = l___regBuiltin_Lean_Elab_Term_elabIdent___closed__2; x_5 = l___regBuiltin_Lean_Elab_Term_elabIdent___closed__3; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); @@ -28698,7 +29108,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabIdent_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(986u); +x_1 = lean_unsigned_to_nat(1007u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28710,7 +29120,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabIdent_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(986u); +x_1 = lean_unsigned_to_nat(1007u); x_2 = lean_unsigned_to_nat(61u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28738,7 +29148,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabIdent_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(986u); +x_1 = lean_unsigned_to_nat(1007u); x_2 = lean_unsigned_to_nat(29u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28750,7 +29160,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabIdent_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(986u); +x_1 = lean_unsigned_to_nat(1007u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28809,7 +29219,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabNamedPattern___close { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__7; +x_2 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__9; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } @@ -28874,7 +29284,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabNamedPattern_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(988u); +x_1 = lean_unsigned_to_nat(1009u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28886,7 +29296,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabNamedPattern_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(988u); +x_1 = lean_unsigned_to_nat(1009u); x_2 = lean_unsigned_to_nat(75u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28914,7 +29324,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabNamedPattern_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(988u); +x_1 = lean_unsigned_to_nat(1009u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28926,7 +29336,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabNamedPattern_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(988u); +x_1 = lean_unsigned_to_nat(1009u); x_2 = lean_unsigned_to_nat(52u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29011,7 +29421,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabDotIdent(lean_object* { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_2 = l___regBuiltin_Lean_Elab_Term_elabApp___closed__6; -x_3 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__16; +x_3 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__18; x_4 = l___regBuiltin_Lean_Elab_Term_elabDotIdent___closed__2; x_5 = l___regBuiltin_Lean_Elab_Term_elabDotIdent___closed__3; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); @@ -29022,7 +29432,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabDotIdent_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(989u); +x_1 = lean_unsigned_to_nat(1010u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29034,7 +29444,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabDotIdent_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(989u); +x_1 = lean_unsigned_to_nat(1010u); x_2 = lean_unsigned_to_nat(67u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29062,7 +29472,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabDotIdent_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(989u); +x_1 = lean_unsigned_to_nat(1010u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29074,7 +29484,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabDotIdent_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(989u); +x_1 = lean_unsigned_to_nat(1010u); x_2 = lean_unsigned_to_nat(44u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29159,7 +29569,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabExplicitUniv(lean_obj { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_2 = l___regBuiltin_Lean_Elab_Term_elabApp___closed__6; -x_3 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__12; +x_3 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__14; x_4 = l___regBuiltin_Lean_Elab_Term_elabExplicitUniv___closed__2; x_5 = l___regBuiltin_Lean_Elab_Term_elabExplicitUniv___closed__3; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); @@ -29188,7 +29598,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabExplicitUniv_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(991u); +x_1 = lean_unsigned_to_nat(1012u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29200,7 +29610,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabExplicitUniv_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(991u); +x_1 = lean_unsigned_to_nat(1012u); x_2 = lean_unsigned_to_nat(75u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29228,7 +29638,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabExplicitUniv_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(991u); +x_1 = lean_unsigned_to_nat(1012u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29240,7 +29650,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabExplicitUniv_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(991u); +x_1 = lean_unsigned_to_nat(1012u); x_2 = lean_unsigned_to_nat(52u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29423,7 +29833,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabPipeProj(lean_object* x_1, lean_ob _start: { lean_object* x_10; uint8_t x_11; -x_10 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__4; +x_10 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__6; lean_inc(x_1); x_11 = l_Lean_Syntax_isOfKind(x_1, x_10); if (x_11 == 0) @@ -29494,7 +29904,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabPipeProj(lean_object* { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_2 = l___regBuiltin_Lean_Elab_Term_elabApp___closed__6; -x_3 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__4; +x_3 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__6; x_4 = l___regBuiltin_Lean_Elab_Term_elabPipeProj___closed__2; x_5 = l___regBuiltin_Lean_Elab_Term_elabPipeProj___closed__3; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); @@ -29523,7 +29933,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabPipeProj_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(993u); +x_1 = lean_unsigned_to_nat(1014u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29535,7 +29945,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabPipeProj_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(998u); +x_1 = lean_unsigned_to_nat(1019u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29563,7 +29973,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabPipeProj_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(993u); +x_1 = lean_unsigned_to_nat(1014u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29575,7 +29985,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabPipeProj_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(993u); +x_1 = lean_unsigned_to_nat(1014u); x_2 = lean_unsigned_to_nat(44u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29643,7 +30053,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabExplicit(lean_object* x_1, lean_ob _start: { lean_object* x_10; uint8_t x_11; -x_10 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__14; +x_10 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__16; lean_inc(x_1); x_11 = l_Lean_Syntax_isOfKind(x_1, x_10); if (x_11 == 0) @@ -29665,13 +30075,13 @@ else lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; x_13 = lean_unsigned_to_nat(1u); x_14 = l_Lean_Syntax_getArg(x_1, x_13); -x_15 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__10; +x_15 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__12; lean_inc(x_14); x_16 = l_Lean_Syntax_isOfKind(x_14, x_15); if (x_16 == 0) { lean_object* x_17; uint8_t x_18; -x_17 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__12; +x_17 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__14; lean_inc(x_14); x_18 = l_Lean_Syntax_isOfKind(x_14, x_17); if (x_18 == 0) @@ -29691,80 +30101,79 @@ return x_23; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; +lean_object* x_24; lean_object* x_25; uint8_t x_26; x_24 = l_Lean_Syntax_getArg(x_14, x_13); -x_25 = l_Lean_nullKind; -x_26 = lean_unsigned_to_nat(2u); +x_25 = lean_unsigned_to_nat(2u); lean_inc(x_24); -x_27 = l_Lean_Syntax_isNodeOf(x_24, x_25, x_26); -if (x_27 == 0) +x_26 = l_Lean_Syntax_matchesNull(x_24, x_25); +if (x_26 == 0) { -uint8_t x_28; uint8_t x_29; lean_object* x_30; +uint8_t x_27; uint8_t x_28; lean_object* x_29; lean_dec(x_24); -x_28 = 1; -x_29 = 0; -x_30 = l_Lean_Elab_Term_elabTerm(x_14, x_2, x_28, x_29, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_30; +x_27 = 1; +x_28 = 0; +x_29 = l_Lean_Elab_Term_elabTerm(x_14, x_2, x_27, x_28, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_29; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; -x_31 = lean_unsigned_to_nat(0u); -x_32 = l_Lean_Syntax_getArg(x_24, x_31); -x_33 = l_Lean_Syntax_getArg(x_24, x_13); +lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_30 = lean_unsigned_to_nat(0u); +x_31 = l_Lean_Syntax_getArg(x_24, x_30); +x_32 = l_Lean_Syntax_getArg(x_24, x_13); lean_dec(x_24); -x_34 = l_Lean_Syntax_isNodeOf(x_33, x_25, x_31); -if (x_34 == 0) +x_33 = l_Lean_Syntax_matchesNull(x_32, x_30); +if (x_33 == 0) { -uint8_t x_35; uint8_t x_36; lean_object* x_37; -lean_dec(x_32); -x_35 = 1; -x_36 = 0; -x_37 = l_Lean_Elab_Term_elabTerm(x_14, x_2, x_35, x_36, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_37; +uint8_t x_34; uint8_t x_35; lean_object* x_36; +lean_dec(x_31); +x_34 = 1; +x_35 = 0; +x_36 = l_Lean_Elab_Term_elabTerm(x_14, x_2, x_34, x_35, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_36; } else { -uint8_t x_38; uint8_t x_39; lean_object* x_40; +uint8_t x_37; uint8_t x_38; lean_object* x_39; lean_dec(x_14); -x_38 = 1; -x_39 = 0; -x_40 = l_Lean_Elab_Term_elabTerm(x_32, x_2, x_38, x_39, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_40; +x_37 = 1; +x_38 = 0; +x_39 = l_Lean_Elab_Term_elabTerm(x_31, x_2, x_37, x_38, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_39; } } } } else { -lean_object* x_41; lean_object* x_42; uint8_t x_43; -x_41 = lean_unsigned_to_nat(0u); -x_42 = l_Lean_Syntax_getArg(x_14, x_41); -x_43 = l_Lean_Syntax_isOfKind(x_42, x_15); -if (x_43 == 0) +lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_40 = lean_unsigned_to_nat(0u); +x_41 = l_Lean_Syntax_getArg(x_14, x_40); +x_42 = l_Lean_Syntax_isOfKind(x_41, x_15); +if (x_42 == 0) { -uint8_t x_44; uint8_t x_45; lean_object* x_46; +uint8_t x_43; uint8_t x_44; lean_object* x_45; lean_dec(x_1); -x_44 = 1; -x_45 = 0; -x_46 = l_Lean_Elab_Term_elabTerm(x_14, x_2, x_44, x_45, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_46; +x_43 = 1; +x_44 = 0; +x_45 = l_Lean_Elab_Term_elabTerm(x_14, x_2, x_43, x_44, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_45; } else { -lean_object* x_47; +lean_object* x_46; lean_dec(x_14); -x_47 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAtom(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_47; +x_46 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAtom(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_46; } } } else { -lean_object* x_48; +lean_object* x_47; lean_dec(x_14); -x_48 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAtom(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_48; +x_47 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAtom(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_47; } } } @@ -29800,7 +30209,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabExplicit(lean_object* { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_2 = l___regBuiltin_Lean_Elab_Term_elabApp___closed__6; -x_3 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__14; +x_3 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__16; x_4 = l___regBuiltin_Lean_Elab_Term_elabExplicit___closed__2; x_5 = l___regBuiltin_Lean_Elab_Term_elabExplicit___closed__3; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); @@ -29829,7 +30238,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabExplicit_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1003u); +x_1 = lean_unsigned_to_nat(1024u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29841,7 +30250,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabExplicit_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1009u); +x_1 = lean_unsigned_to_nat(1030u); x_2 = lean_unsigned_to_nat(50u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29869,7 +30278,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabExplicit_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1003u); +x_1 = lean_unsigned_to_nat(1024u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29881,7 +30290,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabExplicit_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1003u); +x_1 = lean_unsigned_to_nat(1024u); x_2 = lean_unsigned_to_nat(44u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29939,7 +30348,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabChoice___closed__1() _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("choice", 6); +x_1 = lean_mk_string_from_bytes("elabChoice", 10); return x_1; } } @@ -29947,7 +30356,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabChoice___closed__2() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); +x_1 = l___regBuiltin_Lean_Elab_Term_elabApp___closed__3; x_2 = l___regBuiltin_Lean_Elab_Term_elabChoice___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -29957,24 +30366,6 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabChoice___closed__3() _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("elabChoice", 10); -return x_1; -} -} -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabChoice___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_elabApp___closed__3; -x_2 = l___regBuiltin_Lean_Elab_Term_elabChoice___closed__3; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabChoice___closed__5() { -_start: -{ -lean_object* x_1; x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabChoice), 9, 0); return x_1; } @@ -29984,9 +30375,9 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabChoice(lean_object* x { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_2 = l___regBuiltin_Lean_Elab_Term_elabApp___closed__6; -x_3 = l___regBuiltin_Lean_Elab_Term_elabChoice___closed__2; -x_4 = l___regBuiltin_Lean_Elab_Term_elabChoice___closed__4; -x_5 = l___regBuiltin_Lean_Elab_Term_elabChoice___closed__5; +x_3 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__2; +x_4 = l___regBuiltin_Lean_Elab_Term_elabChoice___closed__2; +x_5 = l___regBuiltin_Lean_Elab_Term_elabChoice___closed__3; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); return x_6; } @@ -29995,7 +30386,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabChoice_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1011u); +x_1 = lean_unsigned_to_nat(1032u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30007,7 +30398,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabChoice_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1011u); +x_1 = lean_unsigned_to_nat(1032u); x_2 = lean_unsigned_to_nat(63u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30035,7 +30426,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabChoice_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1011u); +x_1 = lean_unsigned_to_nat(1032u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30047,7 +30438,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabChoice_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1011u); +x_1 = lean_unsigned_to_nat(1032u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30087,7 +30478,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabChoice_declRange(lean _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_elabChoice___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_elabChoice___closed__2; x_3 = l___regBuiltin_Lean_Elab_Term_elabChoice_declRange___closed__7; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; @@ -30132,7 +30523,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabProj(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_2 = l___regBuiltin_Lean_Elab_Term_elabApp___closed__6; -x_3 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__2; +x_3 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__4; x_4 = l___regBuiltin_Lean_Elab_Term_elabProj___closed__2; x_5 = l___regBuiltin_Lean_Elab_Term_elabProj___closed__3; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); @@ -30143,7 +30534,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabProj_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1012u); +x_1 = lean_unsigned_to_nat(1033u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30155,7 +30546,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabProj_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1012u); +x_1 = lean_unsigned_to_nat(1033u); x_2 = lean_unsigned_to_nat(59u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30183,7 +30574,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabProj_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1012u); +x_1 = lean_unsigned_to_nat(1033u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30195,7 +30586,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabProj_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1012u); +x_1 = lean_unsigned_to_nat(1033u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30280,7 +30671,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabArrayRef(lean_object* { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_2 = l___regBuiltin_Lean_Elab_Term_elabApp___closed__6; -x_3 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__6; +x_3 = l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__8; x_4 = l___regBuiltin_Lean_Elab_Term_elabArrayRef___closed__2; x_5 = l___regBuiltin_Lean_Elab_Term_elabArrayRef___closed__3; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); @@ -30291,7 +30682,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabArrayRef_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1013u); +x_1 = lean_unsigned_to_nat(1034u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30303,7 +30694,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabArrayRef_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1013u); +x_1 = lean_unsigned_to_nat(1034u); x_2 = lean_unsigned_to_nat(67u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30331,7 +30722,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabArrayRef_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1013u); +x_1 = lean_unsigned_to_nat(1034u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30343,7 +30734,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabArrayRef_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1013u); +x_1 = lean_unsigned_to_nat(1034u); x_2 = lean_unsigned_to_nat(44u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30389,7 +30780,7 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_13490_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_13748_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; @@ -30613,16 +31004,6 @@ l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___clos lean_mark_persistent(l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__2); l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__3 = _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__3(); lean_mark_persistent(l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__3); -l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__4 = _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__4(); -lean_mark_persistent(l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__4); -l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__5 = _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__5(); -lean_mark_persistent(l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__5); -l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__6 = _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__6(); -lean_mark_persistent(l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__6); -l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__7 = _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__7(); -lean_mark_persistent(l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__7); -l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__8 = _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__8(); -lean_mark_persistent(l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__8); l_Lean_Elab_Term_elabAppArgs___closed__1 = _init_l_Lean_Elab_Term_elabAppArgs___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_elabAppArgs___closed__1); l_Lean_Elab_Term_elabAppArgs___closed__2 = _init_l_Lean_Elab_Term_elabAppArgs___closed__2(); @@ -30709,6 +31090,8 @@ l_List_forIn_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_mkBaseProjectio lean_mark_persistent(l_List_forIn_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_mkBaseProjections___spec__1___closed__1); l_List_forIn_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_mkBaseProjections___spec__1___closed__2 = _init_l_List_forIn_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_mkBaseProjections___spec__1___closed__2(); lean_mark_persistent(l_List_forIn_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_mkBaseProjections___spec__1___closed__2); +l_List_forIn_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_mkBaseProjections___spec__1___closed__3 = _init_l_List_forIn_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_mkBaseProjections___spec__1___closed__3(); +lean_mark_persistent(l_List_forIn_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_mkBaseProjections___spec__1___closed__3); l___private_Lean_Elab_App_0__Lean_Elab_Term_mkBaseProjections___closed__1 = _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_mkBaseProjections___closed__1(); lean_mark_persistent(l___private_Lean_Elab_App_0__Lean_Elab_Term_mkBaseProjections___closed__1); l___private_Lean_Elab_App_0__Lean_Elab_Term_mkBaseProjections___closed__2 = _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_mkBaseProjections___closed__2(); @@ -30835,6 +31218,10 @@ l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__25 = _init_l___p lean_mark_persistent(l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__25); l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__26 = _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__26(); lean_mark_persistent(l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__26); +l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__27 = _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__27(); +lean_mark_persistent(l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__27); +l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__28 = _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__28(); +lean_mark_persistent(l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__28); l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_App_0__Lean_Elab_Term_getSuccesses___spec__1___closed__1 = _init_l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_App_0__Lean_Elab_Term_getSuccesses___spec__1___closed__1(); lean_mark_persistent(l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_App_0__Lean_Elab_Term_getSuccesses___spec__1___closed__1); l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_App_0__Lean_Elab_Term_getSuccesses___spec__1___closed__2 = _init_l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_App_0__Lean_Elab_Term_getSuccesses___spec__1___closed__2(); @@ -31106,10 +31493,6 @@ l___regBuiltin_Lean_Elab_Term_elabChoice___closed__2 = _init_l___regBuiltin_Lean lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_elabChoice___closed__2); l___regBuiltin_Lean_Elab_Term_elabChoice___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_elabChoice___closed__3(); lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_elabChoice___closed__3); -l___regBuiltin_Lean_Elab_Term_elabChoice___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_elabChoice___closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_elabChoice___closed__4); -l___regBuiltin_Lean_Elab_Term_elabChoice___closed__5 = _init_l___regBuiltin_Lean_Elab_Term_elabChoice___closed__5(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_elabChoice___closed__5); res = l___regBuiltin_Lean_Elab_Term_elabChoice(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); @@ -31182,7 +31565,7 @@ lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_elabArrayRef_declRange___clos res = l___regBuiltin_Lean_Elab_Term_elabArrayRef_declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_13490_(lean_io_mk_world()); +res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_13748_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/Attributes.c b/stage0/stdlib/Lean/Elab/Attributes.c index 50a197277ab..23170ef8414 100644 --- a/stage0/stdlib/Lean/Elab/Attributes.c +++ b/stage0/stdlib/Lean/Elab/Attributes.c @@ -15,6 +15,7 @@ extern "C" { #endif lean_object* l_List_reverse___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Attribute_stx___default; +static lean_object* l_Lean_Elab_mkAttrKindGlobal___closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_elabAttr___spec__1___rarg___lambda__5(lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_elabAttrs___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); @@ -25,7 +26,6 @@ lean_object* l_Lean_throwMaxRecDepthAt___rarg(lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_elabAttr___spec__1___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_elabAttr___spec__1___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_nullKind; lean_object* l_Lean_Macro_getCurrNamespace(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_elabAttr___spec__3___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_name_mk_string(lean_object*, lean_object*); @@ -167,6 +167,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___rarg___lambda__1(lean_object*, u static lean_object* l_Lean_Elab_mkAttrKindGlobal___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_elabAttr___spec__3___rarg(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_liftMacroM___at_Lean_Elab_elabAttr___spec__1___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_mkAttrKindGlobal___closed__9; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_elabAttr___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_toAttributeKind___closed__7; lean_object* lean_nat_to_int(lean_object*); @@ -681,9 +682,27 @@ return x_2; static lean_object* _init_l_Lean_Elab_mkAttrKindGlobal___closed__4() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("null", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_mkAttrKindGlobal___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_mkAttrKindGlobal___closed__4; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_mkAttrKindGlobal___closed__6() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean_nullKind; +x_2 = l_Lean_Elab_mkAttrKindGlobal___closed__5; x_3 = l_Lean_Elab_mkAttrKindGlobal___closed__3; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -692,7 +711,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_mkAttrKindGlobal___closed__5() { +static lean_object* _init_l_Lean_Elab_mkAttrKindGlobal___closed__7() { _start: { lean_object* x_1; lean_object* x_2; @@ -701,23 +720,23 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_mkAttrKindGlobal___closed__6() { +static lean_object* _init_l_Lean_Elab_mkAttrKindGlobal___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_mkAttrKindGlobal___closed__5; -x_2 = l_Lean_Elab_mkAttrKindGlobal___closed__4; +x_1 = l_Lean_Elab_mkAttrKindGlobal___closed__7; +x_2 = l_Lean_Elab_mkAttrKindGlobal___closed__6; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_mkAttrKindGlobal___closed__7() { +static lean_object* _init_l_Lean_Elab_mkAttrKindGlobal___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); x_2 = l_Lean_Elab_mkAttrKindGlobal___closed__2; -x_3 = l_Lean_Elab_mkAttrKindGlobal___closed__6; +x_3 = l_Lean_Elab_mkAttrKindGlobal___closed__8; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -729,7 +748,7 @@ static lean_object* _init_l_Lean_Elab_mkAttrKindGlobal() { _start: { lean_object* x_1; -x_1 = l_Lean_Elab_mkAttrKindGlobal___closed__7; +x_1 = l_Lean_Elab_mkAttrKindGlobal___closed__9; return x_1; } } @@ -2735,6 +2754,10 @@ l_Lean_Elab_mkAttrKindGlobal___closed__6 = _init_l_Lean_Elab_mkAttrKindGlobal___ lean_mark_persistent(l_Lean_Elab_mkAttrKindGlobal___closed__6); l_Lean_Elab_mkAttrKindGlobal___closed__7 = _init_l_Lean_Elab_mkAttrKindGlobal___closed__7(); lean_mark_persistent(l_Lean_Elab_mkAttrKindGlobal___closed__7); +l_Lean_Elab_mkAttrKindGlobal___closed__8 = _init_l_Lean_Elab_mkAttrKindGlobal___closed__8(); +lean_mark_persistent(l_Lean_Elab_mkAttrKindGlobal___closed__8); +l_Lean_Elab_mkAttrKindGlobal___closed__9 = _init_l_Lean_Elab_mkAttrKindGlobal___closed__9(); +lean_mark_persistent(l_Lean_Elab_mkAttrKindGlobal___closed__9); l_Lean_Elab_mkAttrKindGlobal = _init_l_Lean_Elab_mkAttrKindGlobal(); lean_mark_persistent(l_Lean_Elab_mkAttrKindGlobal); l_Lean_Elab_elabAttr___rarg___lambda__2___closed__1 = _init_l_Lean_Elab_elabAttr___rarg___lambda__2___closed__1(); diff --git a/stage0/stdlib/Lean/Elab/AuxDef.c b/stage0/stdlib/Lean/Elab/AuxDef.c index a894c4dae40..108ccd1b3b7 100644 --- a/stage0/stdlib/Lean/Elab/AuxDef.c +++ b/stage0/stdlib/Lean/Elab/AuxDef.c @@ -14,16 +14,16 @@ extern "C" { #endif lean_object* l_List_reverse___rarg(lean_object*); -static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__11; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabAuxDef_declRange(lean_object*); size_t lean_usize_add(size_t, size_t); lean_object* lean_erase_macro_scopes(lean_object*); -static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__19; +static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__6; lean_object* lean_mk_empty_array_with_capacity(lean_object*); -extern lean_object* l_Lean_nullKind; +static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__5; lean_object* lean_name_mk_string(lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); lean_object* lean_array_uget(lean_object*, size_t); +static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__15; static lean_object* l_Lean_Elab_Command_aux__def___closed__24; lean_object* l_Array_append___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___boxed(lean_object*, lean_object*); @@ -31,83 +31,82 @@ static lean_object* l___regBuiltin_Lean_Elab_Command_elabAuxDef_declRange___clos static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__1; lean_object* l_Lean_SourceInfo_fromRef(lean_object*); extern lean_object* l_Lean_Elab_Command_commandElabAttribute; +static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__13; static lean_object* l_Lean_Elab_Command_aux__def___closed__43; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__23; static lean_object* l_Lean_Elab_Command_aux__def___closed__10; static lean_object* l_Lean_Elab_Command_aux__def___closed__32; static lean_object* l_Lean_Elab_Command_aux__def___closed__20; +static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__18; static lean_object* l_Lean_Elab_Command_aux__def___closed__36; LEAN_EXPORT lean_object* l_Lean_mkAuxName___at_Lean_Elab_Command_elabAuxDef___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__24; -static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__7; LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Command_elabAuxDef___spec__4(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__8; lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_aux__def___closed__35; lean_object* l_Lean_mkIdentFrom(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__10; static lean_object* l_Lean_Elab_Command_aux__def___closed__16; +LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); +static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__17; static lean_object* l___regBuiltin_Lean_Elab_Command_elabAuxDef_declRange___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Command_elabAuxDef_declRange___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_aux__def; static lean_object* l___regBuiltin_Lean_Elab_Command_elabAuxDef_declRange___closed__5; -static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__16; lean_object* l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l_Lean_Elab_Command_aux__def___closed__27; static lean_object* l___regBuiltin_Lean_Elab_Command_elabAuxDef_declRange___closed__3; static lean_object* l_Lean_Elab_Command_aux__def___closed__39; static lean_object* l_Lean_Elab_Command_aux__def___closed__19; +static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__19; static lean_object* l_Lean_Elab_Command_aux__def___closed__12; static lean_object* l___regBuiltin_Lean_Elab_Command_elabAuxDef_declRange___closed__6; -static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__20; +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_aux__def___closed__11; static lean_object* l___regBuiltin_Lean_Elab_Command_elabAuxDef_declRange___closed__4; static lean_object* l_Lean_Elab_Command_aux__def___closed__31; -static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__6; +static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__2; +static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__21; lean_object* l_Lean_Name_toString(lean_object*, uint8_t); +static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__9; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabAuxDef(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__1(lean_object*); -static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__22; +LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__11; static lean_object* l_Lean_Elab_Command_aux__def___closed__26; static lean_object* l_Lean_Elab_Command_aux__def___closed__38; static lean_object* l_Lean_Elab_Command_aux__def___closed__21; +static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__4; lean_object* l_Lean_Syntax_getId(lean_object*); static lean_object* l_Lean_Elab_Command_aux__def___closed__6; lean_object* l_Lean_addBuiltinDeclarationRanges(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__14; static lean_object* l___regBuiltin_Lean_Elab_Command_elabAuxDef___closed__3; lean_object* l_Lean_Elab_Command_elabCommand(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_aux__def___closed__34; static lean_object* l_Lean_Elab_Command_aux__def___closed__23; static lean_object* l_Lean_Elab_Command_aux__def___closed__13; -static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__9; static lean_object* l_Lean_Elab_Command_aux__def___closed__41; -static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__5; size_t lean_usize_of_nat(lean_object*); static lean_object* l_Lean_Elab_Command_aux__def___closed__4; static lean_object* l_Lean_Elab_Command_aux__def___closed__45; static lean_object* l_Lean_Elab_Command_aux__def___closed__9; static lean_object* l_Lean_Elab_Command_aux__def___closed__40; static lean_object* l_Lean_Elab_Command_aux__def___closed__47; -static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__3; -static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__15; -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_aux__def___closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabAuxDef___spec__6(lean_object*, size_t, size_t, lean_object*); -static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__18; static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg___closed__1; +static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__10; lean_object* l_String_intercalate(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabAuxDef___spec__5___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_aux__def___closed__37; -static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__8; +static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__12; static lean_object* l___regBuiltin_Lean_Elab_Command_elabAuxDef___closed__4; -static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__17; +static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__20; static lean_object* l_Lean_Elab_Command_aux__def___closed__8; static lean_object* l_Lean_Elab_Command_aux__def___closed__44; static lean_object* l_Lean_Elab_Command_aux__def___closed__17; -LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_getCurrMacroScope(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_aux__def___closed__1; @@ -119,43 +118,39 @@ LEAN_EXPORT lean_object* l_Lean_mkAuxName___at_Lean_Elab_Command_elabAuxDef___sp static lean_object* l_Lean_Elab_Command_aux__def___closed__28; static lean_object* l_Lean_Elab_Command_aux__def___closed__14; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg(lean_object*); +static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__7; static lean_object* l_Lean_Elab_Command_aux__def___closed__30; static lean_object* l_Lean_Elab_Command_aux__def___closed__5; static lean_object* l_Lean_Elab_Command_aux__def___closed__7; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabAuxDef___spec__5(size_t, size_t, lean_object*); -lean_object* l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Command_elabAuxDef___spec__4___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__4; lean_object* l_Lean_Elab_Command_getMainModule___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_getRef(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_aux__def___closed__33; -static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__13; uint8_t l_Lean_Syntax_isNone(lean_object*); static lean_object* l_Lean_Elab_Command_aux__def___closed__42; static lean_object* l_Lean_Elab_Command_aux__def___closed__3; -static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__12; LEAN_EXPORT lean_object* l_List_mapTRAux___at_Lean_Elab_Command_elabAuxDef___spec__2(lean_object*, lean_object*); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); lean_object* l___private_Lean_MonadEnv_0__Lean_mkAuxNameAux(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_aux__def___closed__25; -LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabAuxDef___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_aux__def___closed__46; -LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Elab_unsupportedSyntaxExceptionId; -static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__2; -static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__23; static lean_object* l___regBuiltin_Lean_Elab_Command_elabAuxDef___closed__1; +static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Command_elabAuxDef___closed__2; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_aux__def___closed__22; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabAuxDef(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg___closed__2; -static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__21; lean_object* l_Lean_Elab_Command_getScope___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__22; static lean_object* l_Lean_Elab_Command_aux__def___closed__29; +static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__14; +static lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__16; lean_object* l_Lean_Name_replacePrefix(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_components(lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); @@ -855,63 +850,22 @@ return x_4; static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_Command_aux__def___closed__31; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__1(lean_object* x_1) { -_start: -{ -lean_object* x_2; uint8_t x_3; -x_2 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__1; -lean_inc(x_1); -x_3 = l_Lean_Syntax_isOfKind(x_1, x_2); -if (x_3 == 0) -{ -lean_object* x_4; -lean_dec(x_1); -x_4 = lean_box(0); -return x_4; -} -else -{ -lean_object* x_5; -x_5 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_5, 0, x_1); -return x_5; -} -} -} -static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabAuxDef___lambda__1), 1, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__2() { -_start: -{ lean_object* x_1; x_1 = lean_mk_string_from_bytes("_aux", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__3() { +static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__2; +x_2 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__4() { +static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__3() { _start: { lean_object* x_1; @@ -919,17 +873,17 @@ x_1 = lean_mk_string_from_bytes("_", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__5() { +static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__4; +x_2 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__6() { +static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__5() { _start: { lean_object* x_1; @@ -937,7 +891,7 @@ x_1 = lean_mk_string_from_bytes("declaration", 11); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__7() { +static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__6() { _start: { lean_object* x_1; @@ -945,7 +899,7 @@ x_1 = lean_mk_string_from_bytes("declModifiers", 13); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__8() { +static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__7() { _start: { lean_object* x_1; @@ -953,17 +907,17 @@ x_1 = lean_mk_string_from_bytes("null", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__9() { +static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__8; +x_2 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__7; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__10() { +static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__9() { _start: { lean_object* x_1; lean_object* x_2; @@ -972,13 +926,13 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__11() { +static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__9; -x_3 = l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__10; +x_2 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__8; +x_3 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__9; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -986,7 +940,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__12() { +static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__11() { _start: { lean_object* x_1; lean_object* x_2; @@ -995,7 +949,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__13() { +static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__12() { _start: { lean_object* x_1; @@ -1003,7 +957,7 @@ x_1 = lean_mk_string_from_bytes("def", 3); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__14() { +static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__13() { _start: { lean_object* x_1; @@ -1011,7 +965,7 @@ x_1 = lean_mk_string_from_bytes("declId", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__15() { +static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__14() { _start: { lean_object* x_1; lean_object* x_2; @@ -1020,7 +974,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__16() { +static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__15() { _start: { lean_object* x_1; @@ -1028,7 +982,7 @@ x_1 = lean_mk_string_from_bytes("optDeclSig", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__17() { +static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__16() { _start: { lean_object* x_1; @@ -1036,7 +990,7 @@ x_1 = lean_mk_string_from_bytes("typeSpec", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__18() { +static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__17() { _start: { lean_object* x_1; lean_object* x_2; @@ -1045,17 +999,17 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__19() { +static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__15; -x_2 = l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__11; +x_1 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__14; +x_2 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__10; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__20() { +static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__19() { _start: { lean_object* x_1; @@ -1063,7 +1017,7 @@ x_1 = lean_mk_string_from_bytes("declValSimple", 13); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__21() { +static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__20() { _start: { lean_object* x_1; lean_object* x_2; @@ -1072,7 +1026,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__22() { +static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__21() { _start: { lean_object* x_1; lean_object* x_2; @@ -1081,22 +1035,22 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__23() { +static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__22() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__10; +x_1 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__9; x_2 = l_Array_append___rarg(x_1, x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__24() { +static lean_object* _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__9; -x_3 = l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__23; +x_2 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__8; +x_3 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__22; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1104,318 +1058,310 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabAuxDef___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_Elab_Command_elabAuxDef___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; lean_object* x_13; +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; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_140; size_t x_141; size_t x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; uint8_t x_146; lean_object* x_147; x_9 = lean_unsigned_to_nat(3u); x_10 = l_Lean_Syntax_getArg(x_1, x_9); -x_11 = l_Lean_Syntax_getArgs(x_10); +x_11 = lean_unsigned_to_nat(5u); +x_12 = l_Lean_Syntax_getArg(x_1, x_11); +x_13 = lean_unsigned_to_nat(7u); +x_14 = l_Lean_Syntax_getArg(x_1, x_13); +x_15 = l_Lean_Syntax_getArgs(x_10); lean_dec(x_10); -x_12 = l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__1; -x_13 = l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__1(x_11, x_12); -lean_dec(x_11); -if (lean_obj_tag(x_13) == 0) -{ -lean_object* x_14; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_14 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg(x_8); -return x_14; -} -else -{ -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_146; size_t x_147; size_t x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; uint8_t x_152; -x_15 = lean_ctor_get(x_13, 0); -lean_inc(x_15); -lean_dec(x_13); -x_16 = lean_unsigned_to_nat(5u); -x_17 = l_Lean_Syntax_getArg(x_1, x_16); -x_18 = lean_unsigned_to_nat(7u); -x_19 = l_Lean_Syntax_getArg(x_1, x_18); -lean_dec(x_1); -x_146 = lean_array_get_size(x_15); -x_147 = lean_usize_of_nat(x_146); -lean_dec(x_146); -x_148 = 0; +x_140 = lean_array_get_size(x_15); +x_141 = lean_usize_of_nat(x_140); +lean_dec(x_140); +x_142 = 0; lean_inc(x_15); -x_149 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabAuxDef___spec__5(x_147, x_148, x_15); -x_150 = lean_array_get_size(x_149); -x_151 = lean_unsigned_to_nat(0u); -x_152 = lean_nat_dec_lt(x_151, x_150); -if (x_152 == 0) -{ -lean_object* x_153; -lean_dec(x_150); -lean_dec(x_149); -x_153 = lean_box(0); -x_20 = x_153; -goto block_145; +x_143 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabAuxDef___spec__5(x_141, x_142, x_15); +x_144 = lean_array_get_size(x_143); +x_145 = lean_unsigned_to_nat(0u); +x_146 = lean_nat_dec_lt(x_145, x_144); +x_147 = l_Lean_Elab_Command_getMainModule___rarg(x_7, x_8); +if (x_146 == 0) +{ +lean_object* x_148; lean_object* x_149; lean_object* x_150; +lean_dec(x_144); +lean_dec(x_143); +x_148 = lean_ctor_get(x_147, 0); +lean_inc(x_148); +x_149 = lean_ctor_get(x_147, 1); +lean_inc(x_149); +lean_dec(x_147); +x_150 = lean_box(0); +x_16 = x_150; +x_17 = x_148; +x_18 = x_149; +goto block_139; } else { -uint8_t x_154; -x_154 = lean_nat_dec_le(x_150, x_150); -if (x_154 == 0) -{ -lean_object* x_155; -lean_dec(x_150); -lean_dec(x_149); -x_155 = lean_box(0); -x_20 = x_155; -goto block_145; +uint8_t x_151; +x_151 = lean_nat_dec_le(x_144, x_144); +if (x_151 == 0) +{ +lean_object* x_152; lean_object* x_153; lean_object* x_154; +lean_dec(x_144); +lean_dec(x_143); +x_152 = lean_ctor_get(x_147, 0); +lean_inc(x_152); +x_153 = lean_ctor_get(x_147, 1); +lean_inc(x_153); +lean_dec(x_147); +x_154 = lean_box(0); +x_16 = x_154; +x_17 = x_152; +x_18 = x_153; +goto block_139; } else { -size_t x_156; lean_object* x_157; lean_object* x_158; -x_156 = lean_usize_of_nat(x_150); -lean_dec(x_150); -x_157 = lean_box(0); -x_158 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabAuxDef___spec__6(x_149, x_148, x_156, x_157); -lean_dec(x_149); -x_20 = x_158; -goto block_145; -} -} -block_145: -{ -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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_21 = l_Lean_Elab_Command_getMainModule___rarg(x_7, x_8); -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_21, 1); -lean_inc(x_23); -lean_dec(x_21); -x_24 = l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__3; -x_25 = l_Lean_Name_append(x_24, x_22); -x_26 = l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__5; -x_27 = l_Lean_Name_append(x_25, x_26); -lean_dec(x_25); -x_28 = l_Lean_Name_append(x_27, x_20); -lean_dec(x_27); -x_29 = l_Lean_Name_components(x_28); -x_30 = lean_box(0); -x_31 = l_List_mapTRAux___at_Lean_Elab_Command_elabAuxDef___spec__2(x_29, x_30); -x_32 = l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__4; -x_33 = l_String_intercalate(x_32, x_31); -x_34 = l_Lean_Elab_Command_getScope___rarg(x_7, x_23); -x_35 = lean_ctor_get(x_34, 0); -lean_inc(x_35); -x_36 = lean_ctor_get(x_34, 1); +lean_object* x_155; lean_object* x_156; size_t x_157; lean_object* x_158; lean_object* x_159; +x_155 = lean_ctor_get(x_147, 0); +lean_inc(x_155); +x_156 = lean_ctor_get(x_147, 1); +lean_inc(x_156); +lean_dec(x_147); +x_157 = lean_usize_of_nat(x_144); +lean_dec(x_144); +x_158 = lean_box(0); +x_159 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabAuxDef___spec__6(x_143, x_142, x_157, x_158); +lean_dec(x_143); +x_16 = x_159; +x_17 = x_155; +x_18 = x_156; +goto block_139; +} +} +block_139: +{ +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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_19 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__2; +x_20 = l_Lean_Name_append(x_19, x_17); +x_21 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__4; +x_22 = l_Lean_Name_append(x_20, x_21); +lean_dec(x_20); +x_23 = l_Lean_Name_append(x_22, x_16); +lean_dec(x_22); +x_24 = l_Lean_Name_components(x_23); +x_25 = lean_box(0); +x_26 = l_List_mapTRAux___at_Lean_Elab_Command_elabAuxDef___spec__2(x_24, x_25); +x_27 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__3; +x_28 = l_String_intercalate(x_27, x_26); +x_29 = l_Lean_Elab_Command_getScope___rarg(x_7, x_18); +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_29, 1); +lean_inc(x_31); +lean_dec(x_29); +x_32 = lean_ctor_get(x_30, 2); +lean_inc(x_32); +lean_dec(x_30); +lean_inc(x_32); +x_33 = lean_name_mk_string(x_32, x_28); +x_34 = lean_unsigned_to_nat(1u); +x_35 = l_Lean_mkAuxName___at_Lean_Elab_Command_elabAuxDef___spec__3(x_33, x_34, x_6, x_7, x_31); +x_36 = lean_ctor_get(x_35, 0); lean_inc(x_36); -lean_dec(x_34); -x_37 = lean_ctor_get(x_35, 2); +x_37 = lean_ctor_get(x_35, 1); lean_inc(x_37); lean_dec(x_35); -lean_inc(x_37); -x_38 = lean_name_mk_string(x_37, x_33); -x_39 = lean_unsigned_to_nat(1u); -x_40 = l_Lean_mkAuxName___at_Lean_Elab_Command_elabAuxDef___spec__3(x_38, x_39, x_6, x_7, x_36); +x_38 = lean_box(0); +x_39 = l_Lean_Name_replacePrefix(x_36, x_32, x_38); +lean_dec(x_32); +x_40 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Command_elabAuxDef___spec__4(x_6, x_7, x_37); x_41 = lean_ctor_get(x_40, 0); lean_inc(x_41); x_42 = lean_ctor_get(x_40, 1); lean_inc(x_42); lean_dec(x_40); -x_43 = lean_box(0); -x_44 = l_Lean_Name_replacePrefix(x_41, x_37, x_43); -lean_dec(x_37); -x_45 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Command_elabAuxDef___spec__4(x_6, x_7, x_42); -x_46 = lean_ctor_get(x_45, 0); +x_43 = l_Lean_Elab_Command_getCurrMacroScope(x_6, x_7, x_42); +x_44 = lean_ctor_get(x_43, 1); +lean_inc(x_44); +lean_dec(x_43); +x_45 = l_Lean_Elab_Command_getMainModule___rarg(x_7, x_44); +x_46 = lean_ctor_get(x_45, 1); lean_inc(x_46); -x_47 = lean_ctor_get(x_45, 1); -lean_inc(x_47); lean_dec(x_45); -x_48 = l_Lean_Elab_Command_getCurrMacroScope(x_6, x_7, x_47); -x_49 = lean_ctor_get(x_48, 1); -lean_inc(x_49); -lean_dec(x_48); -x_50 = l_Lean_Elab_Command_getMainModule___rarg(x_7, x_49); -x_51 = lean_ctor_get(x_50, 1); -lean_inc(x_51); -lean_dec(x_50); -x_52 = l_Lean_Elab_Command_aux__def___closed__13; -x_53 = lean_name_mk_string(x_2, x_52); -x_54 = l_Lean_Elab_Command_aux__def___closed__5; -lean_inc(x_53); -x_55 = lean_name_mk_string(x_53, x_54); -x_56 = l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__6; -lean_inc(x_55); -x_57 = lean_name_mk_string(x_55, x_56); -x_58 = l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__7; -lean_inc(x_55); -x_59 = lean_name_mk_string(x_55, x_58); -x_60 = l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__13; -lean_inc(x_55); -x_61 = lean_name_mk_string(x_55, x_60); -lean_inc(x_46); -x_62 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_62, 0, x_46); -lean_ctor_set(x_62, 1, x_60); -x_63 = l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__14; -lean_inc(x_55); -x_64 = lean_name_mk_string(x_55, x_63); -x_65 = lean_box(2); -x_66 = l_Lean_nullKind; -x_67 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_67, 0, x_65); -lean_ctor_set(x_67, 1, x_66); -lean_ctor_set(x_67, 2, x_15); -x_68 = l_Lean_mkIdentFrom(x_67, x_44); -x_69 = l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__15; -x_70 = lean_array_push(x_69, x_68); -x_71 = l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__11; -x_72 = lean_array_push(x_70, x_71); -x_73 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_73, 0, x_65); -lean_ctor_set(x_73, 1, x_64); -lean_ctor_set(x_73, 2, x_72); -x_74 = l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__16; -lean_inc(x_55); -x_75 = lean_name_mk_string(x_55, x_74); -x_76 = l_Lean_Elab_Command_aux__def___closed__20; -x_77 = lean_name_mk_string(x_53, x_76); -x_78 = l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__17; -x_79 = lean_name_mk_string(x_77, x_78); -x_80 = l_Lean_Elab_Command_aux__def___closed__36; -lean_inc(x_46); -x_81 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_81, 0, x_46); -lean_ctor_set(x_81, 1, x_80); -x_82 = lean_array_push(x_69, x_81); -x_83 = lean_array_push(x_82, x_17); -x_84 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_84, 0, x_65); -lean_ctor_set(x_84, 1, x_79); -lean_ctor_set(x_84, 2, x_83); -x_85 = l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__18; -x_86 = lean_array_push(x_85, x_84); -x_87 = l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__9; -x_88 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_88, 0, x_65); -lean_ctor_set(x_88, 1, x_87); -lean_ctor_set(x_88, 2, x_86); -x_89 = l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__19; -x_90 = lean_array_push(x_89, x_88); -x_91 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_91, 0, x_65); -lean_ctor_set(x_91, 1, x_75); -lean_ctor_set(x_91, 2, x_90); -x_92 = l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__20; -x_93 = lean_name_mk_string(x_55, x_92); -x_94 = l_Lean_Elab_Command_aux__def___closed__43; -x_95 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_95, 0, x_46); -lean_ctor_set(x_95, 1, x_94); -x_96 = l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__21; -x_97 = lean_array_push(x_96, x_95); -x_98 = lean_array_push(x_97, x_19); -x_99 = lean_array_push(x_98, x_71); -x_100 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_100, 0, x_65); -lean_ctor_set(x_100, 1, x_93); -lean_ctor_set(x_100, 2, x_99); -x_101 = l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__22; -x_102 = lean_array_push(x_101, x_62); -x_103 = lean_array_push(x_102, x_73); -x_104 = lean_array_push(x_103, x_91); -x_105 = lean_array_push(x_104, x_100); -x_106 = lean_array_push(x_105, x_71); -x_107 = lean_array_push(x_106, x_71); -x_108 = lean_array_push(x_107, x_71); -x_109 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_109, 0, x_65); -lean_ctor_set(x_109, 1, x_61); -lean_ctor_set(x_109, 2, x_108); +x_47 = l_Lean_Elab_Command_aux__def___closed__13; +x_48 = lean_name_mk_string(x_2, x_47); +x_49 = l_Lean_Elab_Command_aux__def___closed__5; +lean_inc(x_48); +x_50 = lean_name_mk_string(x_48, x_49); +x_51 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__5; +lean_inc(x_50); +x_52 = lean_name_mk_string(x_50, x_51); +x_53 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__6; +lean_inc(x_50); +x_54 = lean_name_mk_string(x_50, x_53); +x_55 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__12; +lean_inc(x_50); +x_56 = lean_name_mk_string(x_50, x_55); +lean_inc(x_41); +x_57 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_57, 0, x_41); +lean_ctor_set(x_57, 1, x_55); +x_58 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__13; +lean_inc(x_50); +x_59 = lean_name_mk_string(x_50, x_58); +x_60 = lean_box(2); +x_61 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__8; +x_62 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_62, 0, x_60); +lean_ctor_set(x_62, 1, x_61); +lean_ctor_set(x_62, 2, x_15); +x_63 = l_Lean_mkIdentFrom(x_62, x_39); +x_64 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__14; +x_65 = lean_array_push(x_64, x_63); +x_66 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__10; +x_67 = lean_array_push(x_65, x_66); +x_68 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_68, 0, x_60); +lean_ctor_set(x_68, 1, x_59); +lean_ctor_set(x_68, 2, x_67); +x_69 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__15; +lean_inc(x_50); +x_70 = lean_name_mk_string(x_50, x_69); +x_71 = l_Lean_Elab_Command_aux__def___closed__20; +x_72 = lean_name_mk_string(x_48, x_71); +x_73 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__16; +x_74 = lean_name_mk_string(x_72, x_73); +x_75 = l_Lean_Elab_Command_aux__def___closed__36; +lean_inc(x_41); +x_76 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_76, 0, x_41); +lean_ctor_set(x_76, 1, x_75); +x_77 = lean_array_push(x_64, x_76); +x_78 = lean_array_push(x_77, x_12); +x_79 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_79, 0, x_60); +lean_ctor_set(x_79, 1, x_74); +lean_ctor_set(x_79, 2, x_78); +x_80 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__17; +x_81 = lean_array_push(x_80, x_79); +x_82 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_82, 0, x_60); +lean_ctor_set(x_82, 1, x_61); +lean_ctor_set(x_82, 2, x_81); +x_83 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__18; +x_84 = lean_array_push(x_83, x_82); +x_85 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_85, 0, x_60); +lean_ctor_set(x_85, 1, x_70); +lean_ctor_set(x_85, 2, x_84); +x_86 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__19; +x_87 = lean_name_mk_string(x_50, x_86); +x_88 = l_Lean_Elab_Command_aux__def___closed__43; +x_89 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_89, 0, x_41); +lean_ctor_set(x_89, 1, x_88); +x_90 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__20; +x_91 = lean_array_push(x_90, x_89); +x_92 = lean_array_push(x_91, x_14); +x_93 = lean_array_push(x_92, x_66); +x_94 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_94, 0, x_60); +lean_ctor_set(x_94, 1, x_87); +lean_ctor_set(x_94, 2, x_93); +x_95 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__21; +x_96 = lean_array_push(x_95, x_57); +x_97 = lean_array_push(x_96, x_68); +x_98 = lean_array_push(x_97, x_85); +x_99 = lean_array_push(x_98, x_94); +x_100 = lean_array_push(x_99, x_66); +x_101 = lean_array_push(x_100, x_66); +x_102 = lean_array_push(x_101, x_66); +x_103 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_103, 0, x_60); +lean_ctor_set(x_103, 1, x_56); +lean_ctor_set(x_103, 2, x_102); if (lean_obj_tag(x_3) == 0) { -lean_object* x_142; -x_142 = l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__10; -x_110 = x_142; -goto block_141; +lean_object* x_136; +x_136 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__9; +x_104 = x_136; +goto block_135; } else { -lean_object* x_143; lean_object* x_144; -x_143 = lean_ctor_get(x_3, 0); -lean_inc(x_143); +lean_object* x_137; lean_object* x_138; +x_137 = lean_ctor_get(x_3, 0); +lean_inc(x_137); lean_dec(x_3); -x_144 = lean_array_push(x_85, x_143); -x_110 = x_144; -goto block_141; -} -block_141: -{ -lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; -x_111 = l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__10; -x_112 = l_Array_append___rarg(x_111, x_110); -x_113 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_113, 0, x_65); -lean_ctor_set(x_113, 1, x_87); -lean_ctor_set(x_113, 2, x_112); -x_114 = l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__12; -x_115 = lean_array_push(x_114, x_113); +x_138 = lean_array_push(x_80, x_137); +x_104 = x_138; +goto block_135; +} +block_135: +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; +x_105 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__9; +x_106 = l_Array_append___rarg(x_105, x_104); +x_107 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_107, 0, x_60); +lean_ctor_set(x_107, 1, x_61); +lean_ctor_set(x_107, 2, x_106); +x_108 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__11; +x_109 = lean_array_push(x_108, x_107); if (lean_obj_tag(x_5) == 0) { -lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; -x_116 = l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__24; -x_117 = lean_array_push(x_115, x_116); -x_118 = lean_array_push(x_117, x_71); -x_119 = lean_array_push(x_118, x_71); -x_120 = lean_array_push(x_119, x_71); -x_121 = lean_array_push(x_120, x_71); -x_122 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_122, 0, x_65); -lean_ctor_set(x_122, 1, x_59); -lean_ctor_set(x_122, 2, x_121); -x_123 = lean_array_push(x_69, x_122); -x_124 = lean_array_push(x_123, x_109); -x_125 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_125, 0, x_65); -lean_ctor_set(x_125, 1, x_57); -lean_ctor_set(x_125, 2, x_124); -x_126 = l_Lean_Elab_Command_elabCommand(x_125, x_6, x_7, x_51); -return x_126; +lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; +x_110 = l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__23; +x_111 = lean_array_push(x_109, x_110); +x_112 = lean_array_push(x_111, x_66); +x_113 = lean_array_push(x_112, x_66); +x_114 = lean_array_push(x_113, x_66); +x_115 = lean_array_push(x_114, x_66); +x_116 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_116, 0, x_60); +lean_ctor_set(x_116, 1, x_54); +lean_ctor_set(x_116, 2, x_115); +x_117 = lean_array_push(x_64, x_116); +x_118 = lean_array_push(x_117, x_103); +x_119 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_119, 0, x_60); +lean_ctor_set(x_119, 1, x_52); +lean_ctor_set(x_119, 2, x_118); +x_120 = l_Lean_Elab_Command_elabCommand(x_119, x_6, x_7, x_46); +return x_120; } else { -lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; -x_127 = lean_ctor_get(x_5, 0); -lean_inc(x_127); +lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; +x_121 = lean_ctor_get(x_5, 0); +lean_inc(x_121); lean_dec(x_5); -x_128 = lean_array_push(x_85, x_127); -x_129 = l_Array_append___rarg(x_111, x_128); +x_122 = lean_array_push(x_80, x_121); +x_123 = l_Array_append___rarg(x_105, x_122); +x_124 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_124, 0, x_60); +lean_ctor_set(x_124, 1, x_61); +lean_ctor_set(x_124, 2, x_123); +x_125 = lean_array_push(x_109, x_124); +x_126 = lean_array_push(x_125, x_66); +x_127 = lean_array_push(x_126, x_66); +x_128 = lean_array_push(x_127, x_66); +x_129 = lean_array_push(x_128, x_66); x_130 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_130, 0, x_65); -lean_ctor_set(x_130, 1, x_87); +lean_ctor_set(x_130, 0, x_60); +lean_ctor_set(x_130, 1, x_54); lean_ctor_set(x_130, 2, x_129); -x_131 = lean_array_push(x_115, x_130); -x_132 = lean_array_push(x_131, x_71); -x_133 = lean_array_push(x_132, x_71); -x_134 = lean_array_push(x_133, x_71); -x_135 = lean_array_push(x_134, x_71); -x_136 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_136, 0, x_65); -lean_ctor_set(x_136, 1, x_59); -lean_ctor_set(x_136, 2, x_135); -x_137 = lean_array_push(x_69, x_136); -x_138 = lean_array_push(x_137, x_109); -x_139 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_139, 0, x_65); -lean_ctor_set(x_139, 1, x_57); -lean_ctor_set(x_139, 2, x_138); -x_140 = l_Lean_Elab_Command_elabCommand(x_139, x_6, x_7, x_51); -return x_140; +x_131 = lean_array_push(x_64, x_130); +x_132 = lean_array_push(x_131, x_103); +x_133 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_133, 0, x_60); +lean_ctor_set(x_133, 1, x_52); +lean_ctor_set(x_133, 2, x_132); +x_134 = l_Lean_Elab_Command_elabCommand(x_133, x_6, x_7, x_46); +return x_134; } } } } } -} -LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabAuxDef___lambda__3(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_EXPORT lean_object* l_Lean_Elab_Command_elabAuxDef___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) { _start: { lean_object* x_8; lean_object* x_9; uint8_t x_10; @@ -1425,69 +1371,70 @@ x_9 = l_Lean_Syntax_getArg(x_1, x_8); x_10 = l_Lean_Syntax_isNone(x_9); if (x_10 == 0) { -lean_object* x_11; uint8_t x_12; -x_11 = l_Lean_nullKind; +uint8_t x_11; lean_inc(x_9); -x_12 = l_Lean_Syntax_isNodeOf(x_9, x_11, x_8); -if (x_12 == 0) +x_11 = l_Lean_Syntax_matchesNull(x_9, x_8); +if (x_11 == 0) { -lean_object* x_13; +lean_object* x_12; lean_dec(x_9); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_13 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg(x_7); -return x_13; +x_12 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg(x_7); +return x_12; } 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; lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_14 = lean_unsigned_to_nat(0u); -x_15 = l_Lean_Syntax_getArg(x_9, x_14); +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; lean_object* x_20; uint8_t x_21; +x_13 = lean_unsigned_to_nat(0u); +x_14 = l_Lean_Syntax_getArg(x_9, x_13); lean_dec(x_9); -x_16 = l_Lean_Elab_Command_aux__def___closed__13; +x_15 = l_Lean_Elab_Command_aux__def___closed__13; lean_inc(x_2); -x_17 = lean_name_mk_string(x_2, x_16); -x_18 = l_Lean_Elab_Command_aux__def___closed__20; -x_19 = lean_name_mk_string(x_17, x_18); -x_20 = l_Lean_Elab_Command_aux__def___closed__22; -x_21 = lean_name_mk_string(x_19, x_20); -lean_inc(x_15); -x_22 = l_Lean_Syntax_isOfKind(x_15, x_21); -lean_dec(x_21); -if (x_22 == 0) -{ -lean_object* x_23; -lean_dec(x_15); +x_16 = lean_name_mk_string(x_2, x_15); +x_17 = l_Lean_Elab_Command_aux__def___closed__20; +x_18 = lean_name_mk_string(x_16, x_17); +x_19 = l_Lean_Elab_Command_aux__def___closed__22; +x_20 = lean_name_mk_string(x_18, x_19); +lean_inc(x_14); +x_21 = l_Lean_Syntax_isOfKind(x_14, x_20); +lean_dec(x_20); +if (x_21 == 0) +{ +lean_object* x_22; +lean_dec(x_14); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_23 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg(x_7); -return x_23; +x_22 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg(x_7); +return x_22; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_24, 0, x_15); -x_25 = lean_box(0); -x_26 = l_Lean_Elab_Command_elabAuxDef___lambda__2(x_1, x_2, x_4, x_25, x_24, x_5, x_6, x_7); -return x_26; +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_23, 0, x_14); +x_24 = lean_box(0); +x_25 = l_Lean_Elab_Command_elabAuxDef___lambda__1(x_1, x_2, x_4, x_24, x_23, x_5, x_6, x_7); +lean_dec(x_1); +return x_25; } } } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_dec(x_9); +x_26 = lean_box(0); x_27 = lean_box(0); -x_28 = lean_box(0); -x_29 = l_Lean_Elab_Command_elabAuxDef___lambda__2(x_1, x_2, x_4, x_28, x_27, x_5, x_6, x_7); -return x_29; +x_28 = l_Lean_Elab_Command_elabAuxDef___lambda__1(x_1, x_2, x_4, x_27, x_26, x_5, x_6, x_7); +lean_dec(x_1); +return x_28; } } } @@ -1515,60 +1462,59 @@ x_9 = l_Lean_Syntax_getArg(x_1, x_8); x_10 = l_Lean_Syntax_isNone(x_9); if (x_10 == 0) { -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = l_Lean_nullKind; -x_12 = lean_unsigned_to_nat(1u); +lean_object* x_11; uint8_t x_12; +x_11 = lean_unsigned_to_nat(1u); lean_inc(x_9); -x_13 = l_Lean_Syntax_isNodeOf(x_9, x_11, x_12); -if (x_13 == 0) +x_12 = l_Lean_Syntax_matchesNull(x_9, x_11); +if (x_12 == 0) { -lean_object* x_14; +lean_object* x_13; lean_dec(x_9); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_14 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg(x_4); -return x_14; +x_13 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg(x_4); +return x_13; } else { -lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_15 = l_Lean_Syntax_getArg(x_9, x_8); +lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_14 = l_Lean_Syntax_getArg(x_9, x_8); lean_dec(x_9); -x_16 = l_Lean_Elab_Command_aux__def___closed__17; -lean_inc(x_15); -x_17 = l_Lean_Syntax_isOfKind(x_15, x_16); -if (x_17 == 0) +x_15 = l_Lean_Elab_Command_aux__def___closed__17; +lean_inc(x_14); +x_16 = l_Lean_Syntax_isOfKind(x_14, x_15); +if (x_16 == 0) { -lean_object* x_18; -lean_dec(x_15); +lean_object* x_17; +lean_dec(x_14); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_18 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg(x_4); -return x_18; +x_17 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg(x_4); +return x_17; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_19 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_19, 0, x_15); -x_20 = l_Lean_Elab_Command_aux__def___closed__2; -x_21 = lean_box(0); -x_22 = l_Lean_Elab_Command_elabAuxDef___lambda__3(x_1, x_20, x_21, x_19, x_2, x_3, x_4); -return x_22; +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_18 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_18, 0, x_14); +x_19 = l_Lean_Elab_Command_aux__def___closed__2; +x_20 = lean_box(0); +x_21 = l_Lean_Elab_Command_elabAuxDef___lambda__2(x_1, x_19, x_20, x_18, x_2, x_3, x_4); +return x_21; } } } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_dec(x_9); -x_23 = lean_box(0); -x_24 = l_Lean_Elab_Command_aux__def___closed__2; -x_25 = lean_box(0); -x_26 = l_Lean_Elab_Command_elabAuxDef___lambda__3(x_1, x_24, x_25, x_23, x_2, x_3, x_4); -return x_26; +x_22 = lean_box(0); +x_23 = l_Lean_Elab_Command_aux__def___closed__2; +x_24 = lean_box(0); +x_25 = l_Lean_Elab_Command_elabAuxDef___lambda__2(x_1, x_23, x_24, x_22, x_2, x_3, x_4); +return x_25; } } } @@ -1628,12 +1574,13 @@ lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabAuxDef___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, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabAuxDef___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_Elab_Command_elabAuxDef___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_Elab_Command_elabAuxDef___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_4); +lean_dec(x_1); return x_9; } } @@ -1900,54 +1847,50 @@ l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1__ lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg___closed__2); l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__1 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__1); -l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__1 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__1); -l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__2 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__2(); -lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__2); -l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__3 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__3(); -lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__3); -l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__4 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__4(); -lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__4); -l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__5 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__5(); -lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__5); -l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__6 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__6(); -lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__6); -l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__7 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__7(); -lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__7); -l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__8 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__8(); -lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__8); -l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__9 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__9(); -lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__9); -l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__10 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__10(); -lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__10); -l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__11 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__11(); -lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__11); -l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__12 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__12(); -lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__12); -l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__13 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__13(); -lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__13); -l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__14 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__14(); -lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__14); -l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__15 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__15(); -lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__15); -l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__16 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__16(); -lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__16); -l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__17 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__17(); -lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__17); -l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__18 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__18(); -lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__18); -l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__19 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__19(); -lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__19); -l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__20 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__20(); -lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__20); -l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__21 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__21(); -lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__21); -l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__22 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__22(); -lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__22); -l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__23 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__23(); -lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__23); -l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__24 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__24(); -lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__2___closed__24); +l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__2 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__2); +l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__3 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__3(); +lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__3); +l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__4 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__4(); +lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__4); +l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__5 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__5(); +lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__5); +l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__6 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__6(); +lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__6); +l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__7 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__7(); +lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__7); +l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__8 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__8(); +lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__8); +l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__9 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__9(); +lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__9); +l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__10 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__10(); +lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__10); +l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__11 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__11(); +lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__11); +l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__12 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__12(); +lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__12); +l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__13 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__13(); +lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__13); +l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__14 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__14(); +lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__14); +l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__15 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__15(); +lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__15); +l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__16 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__16(); +lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__16); +l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__17 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__17(); +lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__17); +l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__18 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__18(); +lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__18); +l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__19 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__19(); +lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__19); +l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__20 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__20(); +lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__20); +l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__21 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__21(); +lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__21); +l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__22 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__22(); +lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__22); +l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__23 = _init_l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__23(); +lean_mark_persistent(l_Lean_Elab_Command_elabAuxDef___lambda__1___closed__23); l___regBuiltin_Lean_Elab_Command_elabAuxDef___closed__1 = _init_l___regBuiltin_Lean_Elab_Command_elabAuxDef___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Elab_Command_elabAuxDef___closed__1); l___regBuiltin_Lean_Elab_Command_elabAuxDef___closed__2 = _init_l___regBuiltin_Lean_Elab_Command_elabAuxDef___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Binders.c b/stage0/stdlib/Lean/Elab/Binders.c index 5e9e0d62cef..03ad39859b0 100644 --- a/stage0/stdlib/Lean/Elab/Binders.c +++ b/stage0/stdlib/Lean/Elab/Binders.c @@ -14,6 +14,7 @@ extern "C" { #endif lean_object* l_List_reverse___rarg(lean_object*); +static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__62; static lean_object* l_Lean_Elab_Term_elabLetDeclAux___lambda__3___closed__1; static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__54; static lean_object* l_Lean_Elab_Term_expandWhereDecls___closed__9; @@ -60,7 +61,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabForall___closed__1; static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__25; lean_object* l_Lean_throwError___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_matchBinder(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_nullKind; lean_object* l_Lean_Meta_whnfForall(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetFunDecl_declRange___closed__4; LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*); @@ -97,7 +97,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDelayedDecl_declRange__ static lean_object* l_Lean_Elab_Term_elabArrow___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclAux___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__3; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1437____closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandOptIdent___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabBinder___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandOptIdent___closed__1; @@ -106,7 +105,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclAux___lambda__6___boxed(lea lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabFun(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDelayedDecl___closed__5; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1437____closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_elabDepArrow___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_elabForall___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabBindersEx___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -129,7 +127,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDecl_declRange___closed static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__8; static lean_object* l___regBuiltin_Lean_Elab_Term_elabForall_declRange___closed__2; static lean_object* l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__5; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux(lean_object*); static lean_object* l_Lean_Elab_Term_expandWhereDecls___closed__11; static lean_object* l_Lean_Elab_Term_elabLetDeclCore___lambda__2___closed__4; @@ -169,11 +167,11 @@ lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______ma static lean_object* l_Lean_Elab_Term_expandFunBinders_loop___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetFunDecl_declRange___closed__5; static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__15; +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Elab_Term_expandWhereDecls___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_expr_instantiate1(lean_object*, lean_object*); lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1437____closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandLetEqnsDecl(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDelayedDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__53; @@ -205,7 +203,6 @@ uint8_t lean_usize_dec_lt(size_t, size_t); lean_object* l_liftExcept___at_Lean_Elab_liftMacroM___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderViews(lean_object*); LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_nameLitKind; LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandMatchAltsIntoMatchTactic(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandFunBinders_loop(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDelayedDecl_declRange___closed__1; @@ -218,7 +215,6 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Binde static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl_declRange___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__3; -lean_object* l_Lean_Syntax_SepArray_ofElems(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_FunBinders_elabFunBinderViews(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___regBuiltin_Lean_Elab_Term_elabArrow(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__5; @@ -246,6 +242,7 @@ static lean_object* l_Lean_Elab_Term_expandWhereDecls___closed__6; lean_object* l_Lean_Elab_Term_Quotation_precheck(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_restoreSynthInstanceCache(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__49; +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_FunBinders_elabFunBindersAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__13; @@ -315,6 +312,7 @@ static lean_object* l_Lean_Elab_Term_elabForall___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_elabFun___closed__2; static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_precheckFun___spec__4___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_precheckFun___spec__1___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1498____closed__2; static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__47; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabForall___boxed(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_Term_elabLetDeclAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -330,8 +328,8 @@ lean_object* l_Lean_Meta_mkLambdaFVars(lean_object*, lean_object*, uint8_t, uint LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandFunBinders_loop___spec__4(lean_object*, size_t, size_t, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_expandFun_declRange___closed__4; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_7233_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1437_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_7537_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1498_(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabDepArrow___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclAux___lambda__6(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__20; @@ -347,7 +345,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_Term_elabLetDecl extern lean_object* l_Lean_Elab_Term_Quotation_precheckAttribute; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclAux___lambda__7___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_matchBinder___spec__1___rarg(lean_object*); -static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__15; LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_precheckFun___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabLetDeclCore___lambda__2___closed__2; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__9; @@ -355,6 +352,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabFun_declRange___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandMatchAltsWhereDecls(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabLetDeclCore___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_matchBinder___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Elab_Term_expandWhereDecls___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Term_addLocalVarInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshId___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_FunBinders_elabFunBinderViews___spec__2___rarg___boxed(lean_object*, lean_object*); @@ -371,7 +369,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabForall___lambda__1(lean_object*, l lean_object* l_Lean_Syntax_getId(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandWhereDecls___lambda__1(lean_object*, lean_object*); lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Init_Meta_0__Lean_quoteNameMk(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabFunBinders(lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__1; static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__32; @@ -394,8 +391,10 @@ static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_matchBinder_ uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandFunBinders_loop___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshBinderName___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1498____closed__1; lean_object* l_Lean_Elab_Term_expandOptType(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_registerFailedToInferBinderTypeInfo___closed__3; +static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__63; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_precheckFun___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabForall___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclAux___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -407,13 +406,13 @@ static lean_object* l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__4 static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDecl___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabForall(lean_object*); lean_object* l_Lean_Syntax_mkStrLit(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__61; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux_loop___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_instInhabitedExpr; static lean_object* l___regBuiltin_Lean_Elab_Term_elabArrow___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_declareTacticSyntax___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandFunBinders_loop___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__19; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1437____closed__5; static lean_object* l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabBinders___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -476,7 +475,6 @@ lean_object* l_Lean_Elab_Term_getMatchAltsNumPatterns(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetFunDecl___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabForall_declRange(lean_object*); -extern lean_object* l_Lean_identKind; static lean_object* l_Lean_Elab_Term_expandWhereDecls___closed__5; static lean_object* l_Lean_Elab_Term_expandFunBinders_loop___closed__12; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*); @@ -484,6 +482,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabBinders_ lean_object* l_String_intercalate(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabArrow(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_matchBinder___closed__5; +static lean_object* l_Lean_Elab_Term_expandWhereDecls___closed__14; LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkBinderAnnotations; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDelayedDecl___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_FunBinders_State_expectedType_x3f___default; @@ -500,6 +499,7 @@ lean_object* l_Lean_mkHole(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclCore___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabArrow_declRange(lean_object*); extern lean_object* l_Lean_Elab_macroAttribute; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1498____closed__4; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabFunBinders___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_setArg(lean_object*, lean_object*, lean_object*); @@ -525,10 +525,10 @@ static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic lean_object* l_Lean_Syntax_getKind(lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___closed__1; extern lean_object* l_Lean_Core_instMonadCoreM; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1437____closed__3; lean_object* l_Lean_Meta_saveAndResetSynthInstanceCache___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDecl_declRange___closed__4; +lean_object* l_Lean_quoteNameMk(lean_object*); static lean_object* l_Lean_Elab_Term_mkFreshBinderName___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___spec__2___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_precheckFun___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__6; @@ -550,7 +550,6 @@ static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinder LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_expandFun_declRange(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDelayedDecl___closed__1; -lean_object* l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabLetDeclAux___lambda__1___closed__2; static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_precheckFun___spec__4___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_precheckArrow(lean_object*); @@ -560,6 +559,7 @@ static lean_object* l_Lean_Elab_Term_elabLetDeclCore___lambda__2___closed__3; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_FunBinders_elabFunBinderViews___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_instMonadQuotation___rarg(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_mkNameLit(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_precheckFun___spec__1___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_FunBinders_elabFunBindersAux___boxed(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_liftMacroM___at_Lean_Elab_Term_precheckFun___spec__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -601,10 +601,11 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabFun_declRange___closed__1; static lean_object* l_Lean_Elab_Term_mkFreshBinderName___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___spec__2___rarg___closed__1; static lean_object* l_Lean_Elab_Term_expandFunBinders_loop___closed__15; static lean_object* l_Lean_Elab_Term_expandFunBinders_loop___closed__18; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getFunBinderIds_x3f___spec__3___closed__2; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1498____closed__5; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderViews_loop___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_instantiateMVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl___closed__3; +LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Elab_Term_expandWhereDecls___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclAux___lambda__4(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandFunBinders_loop___closed__3; uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); @@ -622,6 +623,7 @@ extern lean_object* l_Lean_Elab_unsupportedSyntaxExceptionId; LEAN_EXPORT lean_object* l_Lean_Elab_Term_precheckArrow(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl___closed__5; static lean_object* l_Lean_Elab_Term_expandWhereDecls___closed__12; +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__3; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderViews_loop___rarg___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_elabForall___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -631,7 +633,6 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Binde LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclAux___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addBuiltinDocString(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetFunDecl_declRange___closed__1; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getFunBinderIds_x3f___spec__3___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandFunBinders_loop___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDecl___closed__4; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); @@ -660,6 +661,7 @@ static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandFunBinders_loop___closed__8; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetFunDecl___closed__3; +LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Elab_Term_expandWhereDecls___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__38; static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__21; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclCore(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -693,6 +695,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabFun_declRange(lean_ob LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclCore___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_matchBinder___closed__1; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__6; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1498____closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDecl___closed__1; lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Term_elabLetDeclCore___closed__5; @@ -718,7 +721,6 @@ static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__36; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_elabFun___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_precheckFun___closed__3; uint8_t lean_string_dec_eq(lean_object*, lean_object*); -lean_object* l_Lean_Syntax_mkLit(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandFunBinders_loop___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); @@ -1267,243 +1269,243 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___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_object* x_7, size_t x_8, size_t 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, lean_object* x_16, lean_object* x_17) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___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_object* x_7, lean_object* x_8, size_t x_9, size_t 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) { _start: { -uint8_t x_18; -x_18 = lean_usize_dec_lt(x_9, x_8); -if (x_18 == 0) +uint8_t x_19; +x_19 = lean_usize_dec_lt(x_10, x_9); +if (x_19 == 0) { -lean_object* x_19; +lean_object* x_20; +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); -lean_dec(x_11); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_10); -lean_ctor_set(x_19, 1, x_17); -return x_19; +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_11); +lean_ctor_set(x_20, 1, x_18); +return x_20; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_28; lean_object* x_66; lean_object* x_74; uint8_t x_75; -x_20 = lean_array_uget(x_7, x_9); -x_74 = l_Lean_nullKind; -x_75 = lean_name_eq(x_1, x_74); -if (x_75 == 0) +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_29; lean_object* x_67; lean_object* x_75; uint8_t x_76; +x_21 = lean_array_uget(x_8, x_10); +x_75 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; +x_76 = lean_name_eq(x_1, x_75); +if (x_76 == 0) { -lean_object* x_76; -x_76 = lean_box(0); -x_28 = x_76; -goto block_65; +lean_object* x_77; +x_77 = lean_box(0); +x_29 = x_77; +goto block_66; } else { -uint8_t x_77; -x_77 = l_Lean_Syntax_isAntiquotSuffixSplice(x_20); -if (x_77 == 0) -{ uint8_t x_78; -x_78 = l_Lean_Syntax_isAntiquotSplice(x_20); +x_78 = l_Lean_Syntax_isAntiquotSuffixSplice(x_21); if (x_78 == 0) { -lean_object* x_79; -x_79 = lean_box(0); -x_28 = x_79; -goto block_65; +uint8_t x_79; +x_79 = l_Lean_Syntax_isAntiquotSplice(x_21); +if (x_79 == 0) +{ +lean_object* x_80; +x_80 = lean_box(0); +x_29 = x_80; +goto block_66; } else { -lean_object* x_80; -lean_dec(x_10); +lean_object* x_81; +lean_dec(x_11); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -x_80 = lean_box(0); -x_66 = x_80; -goto block_73; +x_81 = lean_box(0); +x_67 = x_81; +goto block_74; } } else { -lean_object* x_81; -lean_dec(x_10); +lean_object* x_82; +lean_dec(x_11); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -x_81 = lean_box(0); -x_66 = x_81; -goto block_73; +x_82 = lean_box(0); +x_67 = x_82; +goto block_74; } } -block_27: +block_28: { -lean_object* x_23; size_t x_24; size_t x_25; -x_23 = lean_ctor_get(x_21, 0); -lean_inc(x_23); -lean_dec(x_21); -x_24 = 1; -x_25 = lean_usize_add(x_9, x_24); -x_9 = x_25; -x_10 = x_23; -x_17 = x_22; +lean_object* x_24; size_t x_25; size_t x_26; +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = 1; +x_26 = lean_usize_add(x_10, x_25); +x_10 = x_26; +x_11 = x_24; +x_18 = x_23; goto _start; } -block_65: +block_66: { -lean_object* x_29; -lean_dec(x_28); +lean_object* x_30; +lean_dec(x_29); +lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -lean_inc(x_11); -x_29 = l_Lean_Elab_Term_quoteAutoTactic(x_20, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -if (lean_obj_tag(x_29) == 0) +x_30 = l_Lean_Elab_Term_quoteAutoTactic(x_21, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_30) == 0) { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_29, 1); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_31 = lean_ctor_get(x_30, 0); lean_inc(x_31); -lean_dec(x_29); -lean_inc(x_15); -x_32 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_15, x_16, x_31); -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_32, 1); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +lean_inc(x_16); +x_33 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_16, x_17, x_32); +x_34 = lean_ctor_get(x_33, 0); lean_inc(x_34); -lean_dec(x_32); -x_35 = lean_ctor_get(x_15, 10); +x_35 = lean_ctor_get(x_33, 1); lean_inc(x_35); -x_36 = lean_st_ref_get(x_16, x_34); -x_37 = lean_ctor_get(x_36, 0); -lean_inc(x_37); -x_38 = lean_ctor_get(x_36, 1); +lean_dec(x_33); +x_36 = lean_ctor_get(x_16, 10); +lean_inc(x_36); +x_37 = lean_st_ref_get(x_17, x_35); +x_38 = lean_ctor_get(x_37, 0); lean_inc(x_38); -lean_dec(x_36); -x_39 = lean_ctor_get(x_37, 0); +x_39 = lean_ctor_get(x_37, 1); lean_inc(x_39); lean_dec(x_37); -x_40 = lean_environment_main_module(x_39); -x_41 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__3; -x_42 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__4; -lean_inc(x_3); -x_43 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_43, 0, x_41); -lean_ctor_set(x_43, 1, x_3); -lean_ctor_set(x_43, 2, x_42); -x_44 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__5; +x_40 = lean_ctor_get(x_38, 0); +lean_inc(x_40); +lean_dec(x_38); +x_41 = lean_environment_main_module(x_40); +x_42 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__3; +x_43 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__4; lean_inc(x_4); -x_45 = lean_name_mk_string(x_4, x_44); -lean_inc(x_45); -x_46 = l_Lean_addMacroScope(x_40, x_45, x_35); +x_44 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_44, 0, x_42); +lean_ctor_set(x_44, 1, x_4); +lean_ctor_set(x_44, 2, x_43); +x_45 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__5; lean_inc(x_5); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_5); +x_46 = lean_name_mk_string(x_5, x_45); +lean_inc(x_46); +x_47 = l_Lean_addMacroScope(x_41, x_46, x_36); lean_inc(x_6); -x_48 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_48, 0, x_47); +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_46); lean_ctor_set(x_48, 1, x_6); -x_49 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_49, 0, x_33); -lean_ctor_set(x_49, 1, x_43); -lean_ctor_set(x_49, 2, x_46); -lean_ctor_set(x_49, 3, x_48); -x_50 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__8; -x_51 = lean_array_push(x_50, x_10); -x_52 = lean_array_push(x_51, x_30); -x_53 = lean_box(2); -x_54 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; -x_55 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_55, 0, x_53); -lean_ctor_set(x_55, 1, x_54); -lean_ctor_set(x_55, 2, x_52); -x_56 = lean_array_push(x_50, x_49); -x_57 = lean_array_push(x_56, x_55); -x_58 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__2; -x_59 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_59, 0, x_53); -lean_ctor_set(x_59, 1, x_58); -lean_ctor_set(x_59, 2, x_57); -x_60 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_60, 0, x_59); -x_21 = x_60; -x_22 = x_38; -goto block_27; +lean_inc(x_7); +x_49 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_7); +x_50 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_50, 0, x_34); +lean_ctor_set(x_50, 1, x_44); +lean_ctor_set(x_50, 2, x_47); +lean_ctor_set(x_50, 3, x_49); +x_51 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__8; +x_52 = lean_array_push(x_51, x_11); +x_53 = lean_array_push(x_52, x_31); +x_54 = lean_box(2); +x_55 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; +x_56 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_56, 0, x_54); +lean_ctor_set(x_56, 1, x_55); +lean_ctor_set(x_56, 2, x_53); +x_57 = lean_array_push(x_51, x_50); +x_58 = lean_array_push(x_57, x_56); +x_59 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__2; +x_60 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_60, 0, x_54); +lean_ctor_set(x_60, 1, x_59); +lean_ctor_set(x_60, 2, x_58); +x_61 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_61, 0, x_60); +x_22 = x_61; +x_23 = x_39; +goto block_28; } else { -uint8_t x_61; +uint8_t x_62; +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -x_61 = !lean_is_exclusive(x_29); -if (x_61 == 0) +x_62 = !lean_is_exclusive(x_30); +if (x_62 == 0) { -return x_29; +return x_30; } else { -lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_62 = lean_ctor_get(x_29, 0); -x_63 = lean_ctor_get(x_29, 1); +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_30, 0); +x_64 = lean_ctor_get(x_30, 1); +lean_inc(x_64); lean_inc(x_63); -lean_inc(x_62); -lean_dec(x_29); -x_64 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_64, 0, x_62); -lean_ctor_set(x_64, 1, x_63); -return x_64; +lean_dec(x_30); +x_65 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_64); +return x_65; } } } -block_73: +block_74: { -lean_object* x_67; lean_object* x_68; uint8_t x_69; -lean_dec(x_66); -x_67 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__10; -x_68 = l_Lean_throwErrorAt___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__1(x_20, x_67, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -lean_dec(x_16); +lean_object* x_68; lean_object* x_69; uint8_t x_70; +lean_dec(x_67); +x_68 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__10; +x_69 = l_Lean_throwErrorAt___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__1(x_21, x_68, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_17); +lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_20); -x_69 = !lean_is_exclusive(x_68); -if (x_69 == 0) +lean_dec(x_21); +x_70 = !lean_is_exclusive(x_69); +if (x_70 == 0) { -return x_68; +return x_69; } else { -lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_70 = lean_ctor_get(x_68, 0); -x_71 = lean_ctor_get(x_68, 1); +lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_71 = lean_ctor_get(x_69, 0); +x_72 = lean_ctor_get(x_69, 1); +lean_inc(x_72); lean_inc(x_71); -lean_inc(x_70); -lean_dec(x_68); -x_72 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_72, 0, x_70); -lean_ctor_set(x_72, 1, x_71); -return x_72; +lean_dec(x_69); +x_73 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_73, 0, x_71); +lean_ctor_set(x_73, 1, x_72); +return x_73; } } } @@ -1661,26 +1663,56 @@ static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__15() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Array.empty", 11); +x_1 = lean_mk_string_from_bytes("term", 4); return x_1; } } static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__16() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__15; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__16; +x_3 = lean_alloc_ctor(1, 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_Elab_Term_quoteAutoTactic___closed__18() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Array.empty", 11); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__19() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__15; +x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__18; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__17() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__15; +x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__18; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Term_quoteAutoTactic___closed__16; +x_3 = l_Lean_Elab_Term_quoteAutoTactic___closed__19; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1688,7 +1720,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__18() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__21() { _start: { lean_object* x_1; @@ -1696,17 +1728,17 @@ x_1 = lean_mk_string_from_bytes("Array", 5); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__19() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__18; +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__21; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__20() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__23() { _start: { lean_object* x_1; @@ -1714,41 +1746,41 @@ x_1 = lean_mk_string_from_bytes("empty", 5); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__21() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__19; -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__20; +x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__22; +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__23; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__22() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__21; +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__24; 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_Elab_Term_quoteAutoTactic___closed__23() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__22; +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__25; x_3 = lean_alloc_ctor(1, 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_Elab_Term_quoteAutoTactic___closed__24() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__27() { _start: { lean_object* x_1; @@ -1756,22 +1788,22 @@ x_1 = lean_mk_string_from_bytes("Syntax.node", 11); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__25() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__28() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__24; +x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__27; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__26() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__29() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__24; +x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__27; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Term_quoteAutoTactic___closed__25; +x_3 = l_Lean_Elab_Term_quoteAutoTactic___closed__28; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1779,7 +1811,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__27() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__30() { _start: { lean_object* x_1; @@ -1787,17 +1819,17 @@ x_1 = lean_mk_string_from_bytes("Syntax", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__28() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__31() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__27; +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__30; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__29() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__32() { _start: { lean_object* x_1; @@ -1805,61 +1837,61 @@ x_1 = lean_mk_string_from_bytes("node", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__30() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__33() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__28; -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__29; +x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__31; +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__32; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__31() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__34() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__2; -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__27; +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__30; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__32() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__35() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__31; -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__29; +x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__34; +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__32; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__33() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__36() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__32; +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__35; 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_Elab_Term_quoteAutoTactic___closed__34() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__37() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__33; +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__36; x_3 = lean_alloc_ctor(1, 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_Elab_Term_quoteAutoTactic___closed__35() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__38() { _start: { lean_object* x_1; @@ -1867,22 +1899,22 @@ x_1 = lean_mk_string_from_bytes("SourceInfo.none", 15); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__36() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__39() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__35; +x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__38; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__37() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__40() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__35; +x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__38; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Term_quoteAutoTactic___closed__36; +x_3 = l_Lean_Elab_Term_quoteAutoTactic___closed__39; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1890,7 +1922,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__38() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__41() { _start: { lean_object* x_1; @@ -1898,17 +1930,17 @@ x_1 = lean_mk_string_from_bytes("SourceInfo", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__39() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__42() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__38; +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__41; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__40() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__43() { _start: { lean_object* x_1; @@ -1916,61 +1948,61 @@ x_1 = lean_mk_string_from_bytes("none", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__41() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__44() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__39; -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__40; +x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__42; +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__43; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__42() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__45() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__2; -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__38; +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__41; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__43() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__46() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__42; -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__40; +x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__45; +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__43; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__44() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__47() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__43; +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__46; 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_Elab_Term_quoteAutoTactic___closed__45() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__48() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__44; +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__47; x_3 = lean_alloc_ctor(1, 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_Elab_Term_quoteAutoTactic___closed__46() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__49() { _start: { lean_object* x_1; lean_object* x_2; @@ -1979,7 +2011,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__47() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__50() { _start: { lean_object* x_1; @@ -1987,17 +2019,17 @@ x_1 = lean_mk_string_from_bytes("quotedName", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__48() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__51() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__6; -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__47; +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__50; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__49() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__52() { _start: { lean_object* x_1; @@ -2005,7 +2037,7 @@ x_1 = lean_mk_string_from_bytes(".", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__50() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__53() { _start: { lean_object* x_1; @@ -2013,7 +2045,7 @@ x_1 = lean_mk_string_from_bytes("`", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__51() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__54() { _start: { lean_object* x_1; lean_object* x_2; @@ -2022,7 +2054,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__52() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__55() { _start: { lean_object* x_1; @@ -2030,22 +2062,22 @@ x_1 = lean_mk_string_from_bytes("mkAtom", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__53() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__56() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__52; +x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__55; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__54() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__57() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__52; +x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__55; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Term_quoteAutoTactic___closed__53; +x_3 = l_Lean_Elab_Term_quoteAutoTactic___closed__56; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2053,51 +2085,51 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__55() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__58() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__52; +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__55; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__56() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__59() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__2; -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__52; +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__55; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__57() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__60() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__56; +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__59; 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_Elab_Term_quoteAutoTactic___closed__58() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__61() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__57; +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__60; x_3 = lean_alloc_ctor(1, 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_Elab_Term_quoteAutoTactic___closed__59() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__62() { _start: { lean_object* x_1; @@ -2105,11 +2137,11 @@ x_1 = lean_mk_string_from_bytes("invalid auto tactic, identifier is not allowed" return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__60() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__63() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__59; +x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__62; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } @@ -2144,7 +2176,7 @@ uint8_t x_14; x_14 = !lean_is_exclusive(x_1); 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; lean_object* 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +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; 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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; x_15 = lean_ctor_get(x_1, 2); lean_dec(x_15); x_16 = lean_ctor_get(x_1, 1); @@ -2173,183 +2205,183 @@ x_27 = lean_ctor_get(x_25, 0); lean_inc(x_27); lean_dec(x_25); x_28 = lean_environment_main_module(x_27); -x_29 = l_Lean_Elab_Term_quoteAutoTactic___closed__21; +x_29 = lean_box(0); +x_30 = l_Lean_Elab_Term_quoteAutoTactic___closed__24; lean_inc(x_23); -x_30 = l_Lean_addMacroScope(x_28, x_29, x_23); -x_31 = lean_box(0); -x_32 = l_Lean_Elab_Term_quoteAutoTactic___closed__17; -x_33 = l_Lean_Elab_Term_quoteAutoTactic___closed__23; +x_31 = l_Lean_addMacroScope(x_28, x_30, x_23); +x_32 = l_Lean_Elab_Term_quoteAutoTactic___closed__20; +x_33 = l_Lean_Elab_Term_quoteAutoTactic___closed__26; x_34 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_34, 0, x_21); lean_ctor_set(x_34, 1, x_32); -lean_ctor_set(x_34, 2, x_30); +lean_ctor_set(x_34, 2, x_31); lean_ctor_set(x_34, 3, x_33); x_35 = lean_array_get_size(x_12); x_36 = lean_usize_of_nat(x_35); lean_dec(x_35); x_37 = 0; -x_38 = lean_unsigned_to_nat(0u); -x_39 = l_Lean_Elab_Term_quoteAutoTactic___closed__19; +x_38 = l_Lean_Elab_Term_quoteAutoTactic___closed__17; +x_39 = lean_unsigned_to_nat(0u); +x_40 = l_Lean_Elab_Term_quoteAutoTactic___closed__22; lean_inc(x_7); lean_inc(x_6); -x_40 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2(x_11, x_19, x_38, x_39, x_31, x_31, x_12, x_36, x_37, x_34, x_2, x_3, x_4, x_5, x_6, x_7, x_26); +x_41 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2(x_11, x_19, x_38, x_39, x_40, x_29, x_29, x_12, x_36, x_37, x_34, x_2, x_3, x_4, x_5, x_6, x_7, x_26); lean_dec(x_12); lean_dec(x_19); -if (lean_obj_tag(x_40) == 0) +if (lean_obj_tag(x_41) == 0) { -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; -x_41 = lean_ctor_get(x_40, 0); -lean_inc(x_41); -x_42 = lean_ctor_get(x_40, 1); +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; +x_42 = lean_ctor_get(x_41, 0); lean_inc(x_42); -lean_dec(x_40); -x_43 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_6, x_7, x_42); -x_44 = lean_ctor_get(x_43, 0); -lean_inc(x_44); -x_45 = lean_ctor_get(x_43, 1); +x_43 = lean_ctor_get(x_41, 1); +lean_inc(x_43); +lean_dec(x_41); +x_44 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_6, x_7, x_43); +x_45 = lean_ctor_get(x_44, 0); lean_inc(x_45); -lean_dec(x_43); -x_46 = lean_st_ref_get(x_7, x_45); +x_46 = lean_ctor_get(x_44, 1); +lean_inc(x_46); +lean_dec(x_44); +x_47 = lean_st_ref_get(x_7, x_46); lean_dec(x_7); -x_47 = !lean_is_exclusive(x_46); -if (x_47 == 0) +x_48 = !lean_is_exclusive(x_47); +if (x_48 == 0) { -lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_48 = lean_ctor_get(x_46, 0); -x_49 = lean_ctor_get(x_48, 0); -lean_inc(x_49); -lean_dec(x_48); -x_50 = lean_environment_main_module(x_49); -x_51 = l_Lean_Elab_Term_quoteAutoTactic___closed__30; -lean_inc(x_23); +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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_49 = lean_ctor_get(x_47, 0); +x_50 = lean_ctor_get(x_49, 0); lean_inc(x_50); -x_52 = l_Lean_addMacroScope(x_50, x_51, x_23); -x_53 = l_Lean_Elab_Term_quoteAutoTactic___closed__26; -x_54 = l_Lean_Elab_Term_quoteAutoTactic___closed__34; -lean_inc(x_44); -x_55 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_55, 0, x_44); -lean_ctor_set(x_55, 1, x_53); -lean_ctor_set(x_55, 2, x_52); -lean_ctor_set(x_55, 3, x_54); -x_56 = l_Lean_Elab_Term_quoteAutoTactic___closed__41; -x_57 = l_Lean_addMacroScope(x_50, x_56, x_23); -x_58 = l_Lean_Elab_Term_quoteAutoTactic___closed__37; -x_59 = l_Lean_Elab_Term_quoteAutoTactic___closed__45; -x_60 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_60, 0, x_44); -lean_ctor_set(x_60, 1, x_58); -lean_ctor_set(x_60, 2, x_57); -lean_ctor_set(x_60, 3, x_59); +lean_dec(x_49); +x_51 = lean_environment_main_module(x_50); +x_52 = l_Lean_Elab_Term_quoteAutoTactic___closed__33; +lean_inc(x_23); +lean_inc(x_51); +x_53 = l_Lean_addMacroScope(x_51, x_52, x_23); +x_54 = l_Lean_Elab_Term_quoteAutoTactic___closed__29; +x_55 = l_Lean_Elab_Term_quoteAutoTactic___closed__37; +lean_inc(x_45); +x_56 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_56, 0, x_45); +lean_ctor_set(x_56, 1, x_54); +lean_ctor_set(x_56, 2, x_53); +lean_ctor_set(x_56, 3, x_55); +x_57 = l_Lean_Elab_Term_quoteAutoTactic___closed__44; +x_58 = l_Lean_addMacroScope(x_51, x_57, x_23); +x_59 = l_Lean_Elab_Term_quoteAutoTactic___closed__40; +x_60 = l_Lean_Elab_Term_quoteAutoTactic___closed__48; +x_61 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_61, 0, x_45); +lean_ctor_set(x_61, 1, x_59); +lean_ctor_set(x_61, 2, x_58); +lean_ctor_set(x_61, 3, x_60); lean_inc(x_11); -x_61 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_31, x_11); -x_62 = l_Lean_Elab_Term_quoteAutoTactic___closed__46; -x_63 = lean_array_push(x_62, x_60); -x_64 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__8; -x_65 = lean_array_push(x_64, x_55); -if (lean_obj_tag(x_61) == 0) -{ -lean_object* x_66; 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; -x_66 = l___private_Init_Meta_0__Lean_quoteNameMk(x_11); -x_67 = lean_array_push(x_63, x_66); -x_68 = lean_array_push(x_67, x_41); -x_69 = lean_box(2); -x_70 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; -lean_ctor_set(x_1, 2, x_68); -lean_ctor_set(x_1, 1, x_70); -lean_ctor_set(x_1, 0, x_69); -x_71 = lean_array_push(x_65, x_1); -x_72 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__2; -x_73 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_73, 0, x_69); -lean_ctor_set(x_73, 1, x_72); -lean_ctor_set(x_73, 2, x_71); -lean_ctor_set(x_46, 0, x_73); -return x_46; +x_62 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_29, x_11); +x_63 = l_Lean_Elab_Term_quoteAutoTactic___closed__49; +x_64 = lean_array_push(x_63, x_61); +x_65 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__8; +x_66 = lean_array_push(x_65, x_56); +if (lean_obj_tag(x_62) == 0) +{ +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; +x_67 = l_Lean_quoteNameMk(x_11); +x_68 = lean_array_push(x_64, x_67); +x_69 = lean_array_push(x_68, x_42); +x_70 = lean_box(2); +x_71 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; +lean_ctor_set(x_1, 2, x_69); +lean_ctor_set(x_1, 1, x_71); +lean_ctor_set(x_1, 0, x_70); +x_72 = lean_array_push(x_66, x_1); +x_73 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__2; +x_74 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_74, 0, x_70); +lean_ctor_set(x_74, 1, x_73); +lean_ctor_set(x_74, 2, x_72); +lean_ctor_set(x_47, 0, x_74); +return x_47; } else { -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_dec(x_11); -x_74 = lean_ctor_get(x_61, 0); -lean_inc(x_74); -lean_dec(x_61); -x_75 = l_Lean_Elab_Term_quoteAutoTactic___closed__49; -x_76 = l_String_intercalate(x_75, x_74); -x_77 = l_Lean_Elab_Term_quoteAutoTactic___closed__50; -x_78 = lean_string_append(x_77, x_76); -lean_dec(x_76); -x_79 = l_Lean_nameLitKind; +x_75 = lean_ctor_get(x_62, 0); +lean_inc(x_75); +lean_dec(x_62); +x_76 = l_Lean_Elab_Term_quoteAutoTactic___closed__52; +x_77 = l_String_intercalate(x_76, x_75); +x_78 = l_Lean_Elab_Term_quoteAutoTactic___closed__53; +x_79 = lean_string_append(x_78, x_77); +lean_dec(x_77); x_80 = lean_box(2); -x_81 = l_Lean_Syntax_mkLit(x_79, x_78, x_80); -x_82 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_81 = l_Lean_Syntax_mkNameLit(x_79, x_80); +x_82 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_83 = lean_array_push(x_82, x_81); -x_84 = l_Lean_Elab_Term_quoteAutoTactic___closed__48; +x_84 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; lean_ctor_set(x_1, 2, x_83); lean_ctor_set(x_1, 1, x_84); lean_ctor_set(x_1, 0, x_80); -x_85 = lean_array_push(x_63, x_1); -x_86 = lean_array_push(x_85, x_41); +x_85 = lean_array_push(x_64, x_1); +x_86 = lean_array_push(x_85, x_42); x_87 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_88 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_88, 0, x_80); lean_ctor_set(x_88, 1, x_87); lean_ctor_set(x_88, 2, x_86); -x_89 = lean_array_push(x_65, x_88); +x_89 = lean_array_push(x_66, x_88); x_90 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__2; x_91 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_91, 0, x_80); lean_ctor_set(x_91, 1, x_90); lean_ctor_set(x_91, 2, x_89); -lean_ctor_set(x_46, 0, x_91); -return x_46; +lean_ctor_set(x_47, 0, x_91); +return x_47; } } else { lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_92 = lean_ctor_get(x_46, 0); -x_93 = lean_ctor_get(x_46, 1); +x_92 = lean_ctor_get(x_47, 0); +x_93 = lean_ctor_get(x_47, 1); lean_inc(x_93); lean_inc(x_92); -lean_dec(x_46); +lean_dec(x_47); x_94 = lean_ctor_get(x_92, 0); lean_inc(x_94); lean_dec(x_92); x_95 = lean_environment_main_module(x_94); -x_96 = l_Lean_Elab_Term_quoteAutoTactic___closed__30; +x_96 = l_Lean_Elab_Term_quoteAutoTactic___closed__33; lean_inc(x_23); lean_inc(x_95); x_97 = l_Lean_addMacroScope(x_95, x_96, x_23); -x_98 = l_Lean_Elab_Term_quoteAutoTactic___closed__26; -x_99 = l_Lean_Elab_Term_quoteAutoTactic___closed__34; -lean_inc(x_44); +x_98 = l_Lean_Elab_Term_quoteAutoTactic___closed__29; +x_99 = l_Lean_Elab_Term_quoteAutoTactic___closed__37; +lean_inc(x_45); x_100 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_100, 0, x_44); +lean_ctor_set(x_100, 0, x_45); lean_ctor_set(x_100, 1, x_98); lean_ctor_set(x_100, 2, x_97); lean_ctor_set(x_100, 3, x_99); -x_101 = l_Lean_Elab_Term_quoteAutoTactic___closed__41; +x_101 = l_Lean_Elab_Term_quoteAutoTactic___closed__44; x_102 = l_Lean_addMacroScope(x_95, x_101, x_23); -x_103 = l_Lean_Elab_Term_quoteAutoTactic___closed__37; -x_104 = l_Lean_Elab_Term_quoteAutoTactic___closed__45; +x_103 = l_Lean_Elab_Term_quoteAutoTactic___closed__40; +x_104 = l_Lean_Elab_Term_quoteAutoTactic___closed__48; x_105 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_105, 0, x_44); +lean_ctor_set(x_105, 0, x_45); lean_ctor_set(x_105, 1, x_103); lean_ctor_set(x_105, 2, x_102); lean_ctor_set(x_105, 3, x_104); lean_inc(x_11); -x_106 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_31, x_11); -x_107 = l_Lean_Elab_Term_quoteAutoTactic___closed__46; +x_106 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_29, x_11); +x_107 = l_Lean_Elab_Term_quoteAutoTactic___closed__49; x_108 = lean_array_push(x_107, x_105); x_109 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__8; x_110 = lean_array_push(x_109, x_100); if (lean_obj_tag(x_106) == 0) { lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; -x_111 = l___private_Init_Meta_0__Lean_quoteNameMk(x_11); +x_111 = l_Lean_quoteNameMk(x_11); x_112 = lean_array_push(x_108, x_111); -x_113 = lean_array_push(x_112, x_41); +x_113 = lean_array_push(x_112, x_42); x_114 = lean_box(2); x_115 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; lean_ctor_set(x_1, 2, x_113); @@ -2368,121 +2400,121 @@ return x_119; } else { -lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; +lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_dec(x_11); x_120 = lean_ctor_get(x_106, 0); lean_inc(x_120); lean_dec(x_106); -x_121 = l_Lean_Elab_Term_quoteAutoTactic___closed__49; +x_121 = l_Lean_Elab_Term_quoteAutoTactic___closed__52; x_122 = l_String_intercalate(x_121, x_120); -x_123 = l_Lean_Elab_Term_quoteAutoTactic___closed__50; +x_123 = l_Lean_Elab_Term_quoteAutoTactic___closed__53; x_124 = lean_string_append(x_123, x_122); lean_dec(x_122); -x_125 = l_Lean_nameLitKind; -x_126 = lean_box(2); -x_127 = l_Lean_Syntax_mkLit(x_125, x_124, x_126); -x_128 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; -x_129 = lean_array_push(x_128, x_127); -x_130 = l_Lean_Elab_Term_quoteAutoTactic___closed__48; -lean_ctor_set(x_1, 2, x_129); -lean_ctor_set(x_1, 1, x_130); -lean_ctor_set(x_1, 0, x_126); -x_131 = lean_array_push(x_108, x_1); -x_132 = lean_array_push(x_131, x_41); -x_133 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; -x_134 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_134, 0, x_126); -lean_ctor_set(x_134, 1, x_133); -lean_ctor_set(x_134, 2, x_132); -x_135 = lean_array_push(x_110, x_134); -x_136 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__2; -x_137 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_137, 0, x_126); -lean_ctor_set(x_137, 1, x_136); -lean_ctor_set(x_137, 2, x_135); -x_138 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_138, 0, x_137); -lean_ctor_set(x_138, 1, x_93); -return x_138; +x_125 = lean_box(2); +x_126 = l_Lean_Syntax_mkNameLit(x_124, x_125); +x_127 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; +x_128 = lean_array_push(x_127, x_126); +x_129 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +lean_ctor_set(x_1, 2, x_128); +lean_ctor_set(x_1, 1, x_129); +lean_ctor_set(x_1, 0, x_125); +x_130 = lean_array_push(x_108, x_1); +x_131 = lean_array_push(x_130, x_42); +x_132 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; +x_133 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_133, 0, x_125); +lean_ctor_set(x_133, 1, x_132); +lean_ctor_set(x_133, 2, x_131); +x_134 = lean_array_push(x_110, x_133); +x_135 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__2; +x_136 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_136, 0, x_125); +lean_ctor_set(x_136, 1, x_135); +lean_ctor_set(x_136, 2, x_134); +x_137 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_137, 0, x_136); +lean_ctor_set(x_137, 1, x_93); +return x_137; } } } else { -uint8_t x_139; +uint8_t x_138; lean_dec(x_23); lean_free_object(x_1); lean_dec(x_11); lean_dec(x_7); lean_dec(x_6); -x_139 = !lean_is_exclusive(x_40); -if (x_139 == 0) +x_138 = !lean_is_exclusive(x_41); +if (x_138 == 0) { -return x_40; +return x_41; } else { -lean_object* x_140; lean_object* x_141; lean_object* x_142; -x_140 = lean_ctor_get(x_40, 0); -x_141 = lean_ctor_get(x_40, 1); -lean_inc(x_141); +lean_object* x_139; lean_object* x_140; lean_object* x_141; +x_139 = lean_ctor_get(x_41, 0); +x_140 = lean_ctor_get(x_41, 1); lean_inc(x_140); -lean_dec(x_40); -x_142 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_142, 0, x_140); -lean_ctor_set(x_142, 1, x_141); -return x_142; +lean_inc(x_139); +lean_dec(x_41); +x_141 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_141, 0, x_139); +lean_ctor_set(x_141, 1, x_140); +return x_141; } } } else { -lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; size_t x_161; size_t x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; +lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; size_t x_160; size_t x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_dec(x_1); -x_143 = l_Lean_Elab_Term_quoteAutoTactic___closed__14; -x_144 = lean_ctor_get(x_143, 0); -lean_inc(x_144); +x_142 = l_Lean_Elab_Term_quoteAutoTactic___closed__14; +x_143 = lean_ctor_get(x_142, 0); +lean_inc(x_143); lean_inc(x_6); -x_145 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_6, x_7, x_8); -x_146 = lean_ctor_get(x_145, 0); +x_144 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_6, x_7, x_8); +x_145 = lean_ctor_get(x_144, 0); +lean_inc(x_145); +x_146 = lean_ctor_get(x_144, 1); lean_inc(x_146); -x_147 = lean_ctor_get(x_145, 1); +lean_dec(x_144); +x_147 = lean_ctor_get(x_6, 10); lean_inc(x_147); -lean_dec(x_145); -x_148 = lean_ctor_get(x_6, 10); -lean_inc(x_148); -x_149 = lean_st_ref_get(x_7, x_147); -x_150 = lean_ctor_get(x_149, 0); +x_148 = lean_st_ref_get(x_7, x_146); +x_149 = lean_ctor_get(x_148, 0); +lean_inc(x_149); +x_150 = lean_ctor_get(x_148, 1); lean_inc(x_150); -x_151 = lean_ctor_get(x_149, 1); +lean_dec(x_148); +x_151 = lean_ctor_get(x_149, 0); lean_inc(x_151); lean_dec(x_149); -x_152 = lean_ctor_get(x_150, 0); -lean_inc(x_152); -lean_dec(x_150); -x_153 = lean_environment_main_module(x_152); -x_154 = l_Lean_Elab_Term_quoteAutoTactic___closed__21; -lean_inc(x_148); -x_155 = l_Lean_addMacroScope(x_153, x_154, x_148); -x_156 = lean_box(0); -x_157 = l_Lean_Elab_Term_quoteAutoTactic___closed__17; -x_158 = l_Lean_Elab_Term_quoteAutoTactic___closed__23; -x_159 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_159, 0, x_146); -lean_ctor_set(x_159, 1, x_157); -lean_ctor_set(x_159, 2, x_155); -lean_ctor_set(x_159, 3, x_158); -x_160 = lean_array_get_size(x_12); -x_161 = lean_usize_of_nat(x_160); -lean_dec(x_160); -x_162 = 0; +x_152 = lean_environment_main_module(x_151); +x_153 = lean_box(0); +x_154 = l_Lean_Elab_Term_quoteAutoTactic___closed__24; +lean_inc(x_147); +x_155 = l_Lean_addMacroScope(x_152, x_154, x_147); +x_156 = l_Lean_Elab_Term_quoteAutoTactic___closed__20; +x_157 = l_Lean_Elab_Term_quoteAutoTactic___closed__26; +x_158 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_158, 0, x_145); +lean_ctor_set(x_158, 1, x_156); +lean_ctor_set(x_158, 2, x_155); +lean_ctor_set(x_158, 3, x_157); +x_159 = lean_array_get_size(x_12); +x_160 = lean_usize_of_nat(x_159); +lean_dec(x_159); +x_161 = 0; +x_162 = l_Lean_Elab_Term_quoteAutoTactic___closed__17; x_163 = lean_unsigned_to_nat(0u); -x_164 = l_Lean_Elab_Term_quoteAutoTactic___closed__19; +x_164 = l_Lean_Elab_Term_quoteAutoTactic___closed__22; lean_inc(x_7); lean_inc(x_6); -x_165 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2(x_11, x_144, x_163, x_164, x_156, x_156, x_12, x_161, x_162, x_159, x_2, x_3, x_4, x_5, x_6, x_7, x_151); +x_165 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2(x_11, x_143, x_162, x_163, x_164, x_153, x_153, x_12, x_160, x_161, x_158, x_2, x_3, x_4, x_5, x_6, x_7, x_150); lean_dec(x_12); -lean_dec(x_144); +lean_dec(x_143); if (lean_obj_tag(x_165) == 0) { lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; @@ -2515,37 +2547,37 @@ x_175 = lean_ctor_get(x_172, 0); lean_inc(x_175); lean_dec(x_172); x_176 = lean_environment_main_module(x_175); -x_177 = l_Lean_Elab_Term_quoteAutoTactic___closed__30; -lean_inc(x_148); +x_177 = l_Lean_Elab_Term_quoteAutoTactic___closed__33; +lean_inc(x_147); lean_inc(x_176); -x_178 = l_Lean_addMacroScope(x_176, x_177, x_148); -x_179 = l_Lean_Elab_Term_quoteAutoTactic___closed__26; -x_180 = l_Lean_Elab_Term_quoteAutoTactic___closed__34; +x_178 = l_Lean_addMacroScope(x_176, x_177, x_147); +x_179 = l_Lean_Elab_Term_quoteAutoTactic___closed__29; +x_180 = l_Lean_Elab_Term_quoteAutoTactic___closed__37; lean_inc(x_169); x_181 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_181, 0, x_169); lean_ctor_set(x_181, 1, x_179); lean_ctor_set(x_181, 2, x_178); lean_ctor_set(x_181, 3, x_180); -x_182 = l_Lean_Elab_Term_quoteAutoTactic___closed__41; -x_183 = l_Lean_addMacroScope(x_176, x_182, x_148); -x_184 = l_Lean_Elab_Term_quoteAutoTactic___closed__37; -x_185 = l_Lean_Elab_Term_quoteAutoTactic___closed__45; +x_182 = l_Lean_Elab_Term_quoteAutoTactic___closed__44; +x_183 = l_Lean_addMacroScope(x_176, x_182, x_147); +x_184 = l_Lean_Elab_Term_quoteAutoTactic___closed__40; +x_185 = l_Lean_Elab_Term_quoteAutoTactic___closed__48; x_186 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_186, 0, x_169); lean_ctor_set(x_186, 1, x_184); lean_ctor_set(x_186, 2, x_183); lean_ctor_set(x_186, 3, x_185); lean_inc(x_11); -x_187 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_156, x_11); -x_188 = l_Lean_Elab_Term_quoteAutoTactic___closed__46; +x_187 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_153, x_11); +x_188 = l_Lean_Elab_Term_quoteAutoTactic___closed__49; x_189 = lean_array_push(x_188, x_186); x_190 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__8; x_191 = lean_array_push(x_190, x_181); if (lean_obj_tag(x_187) == 0) { lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; -x_192 = l___private_Init_Meta_0__Lean_quoteNameMk(x_11); +x_192 = l_Lean_quoteNameMk(x_11); x_193 = lean_array_push(x_189, x_192); x_194 = lean_array_push(x_193, x_166); x_195 = lean_box(2); @@ -2571,201 +2603,200 @@ return x_201; } else { -lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; +lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_dec(x_11); x_202 = lean_ctor_get(x_187, 0); lean_inc(x_202); lean_dec(x_187); -x_203 = l_Lean_Elab_Term_quoteAutoTactic___closed__49; +x_203 = l_Lean_Elab_Term_quoteAutoTactic___closed__52; x_204 = l_String_intercalate(x_203, x_202); -x_205 = l_Lean_Elab_Term_quoteAutoTactic___closed__50; +x_205 = l_Lean_Elab_Term_quoteAutoTactic___closed__53; x_206 = lean_string_append(x_205, x_204); lean_dec(x_204); -x_207 = l_Lean_nameLitKind; -x_208 = lean_box(2); -x_209 = l_Lean_Syntax_mkLit(x_207, x_206, x_208); -x_210 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; -x_211 = lean_array_push(x_210, x_209); -x_212 = l_Lean_Elab_Term_quoteAutoTactic___closed__48; -x_213 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_213, 0, x_208); -lean_ctor_set(x_213, 1, x_212); -lean_ctor_set(x_213, 2, x_211); -x_214 = lean_array_push(x_189, x_213); -x_215 = lean_array_push(x_214, x_166); -x_216 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; -x_217 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_217, 0, x_208); -lean_ctor_set(x_217, 1, x_216); -lean_ctor_set(x_217, 2, x_215); -x_218 = lean_array_push(x_191, x_217); -x_219 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__2; -x_220 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_220, 0, x_208); -lean_ctor_set(x_220, 1, x_219); -lean_ctor_set(x_220, 2, x_218); +x_207 = lean_box(2); +x_208 = l_Lean_Syntax_mkNameLit(x_206, x_207); +x_209 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; +x_210 = lean_array_push(x_209, x_208); +x_211 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_212 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_212, 0, x_207); +lean_ctor_set(x_212, 1, x_211); +lean_ctor_set(x_212, 2, x_210); +x_213 = lean_array_push(x_189, x_212); +x_214 = lean_array_push(x_213, x_166); +x_215 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; +x_216 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_216, 0, x_207); +lean_ctor_set(x_216, 1, x_215); +lean_ctor_set(x_216, 2, x_214); +x_217 = lean_array_push(x_191, x_216); +x_218 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__2; +x_219 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_219, 0, x_207); +lean_ctor_set(x_219, 1, x_218); +lean_ctor_set(x_219, 2, x_217); if (lean_is_scalar(x_174)) { - x_221 = lean_alloc_ctor(0, 2, 0); + x_220 = lean_alloc_ctor(0, 2, 0); } else { - x_221 = x_174; + x_220 = x_174; } -lean_ctor_set(x_221, 0, x_220); -lean_ctor_set(x_221, 1, x_173); -return x_221; +lean_ctor_set(x_220, 0, x_219); +lean_ctor_set(x_220, 1, x_173); +return x_220; } } else { -lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; -lean_dec(x_148); +lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; +lean_dec(x_147); lean_dec(x_11); lean_dec(x_7); lean_dec(x_6); -x_222 = lean_ctor_get(x_165, 0); +x_221 = lean_ctor_get(x_165, 0); +lean_inc(x_221); +x_222 = lean_ctor_get(x_165, 1); lean_inc(x_222); -x_223 = lean_ctor_get(x_165, 1); -lean_inc(x_223); if (lean_is_exclusive(x_165)) { lean_ctor_release(x_165, 0); lean_ctor_release(x_165, 1); - x_224 = x_165; + x_223 = x_165; } else { lean_dec_ref(x_165); - x_224 = lean_box(0); + x_223 = lean_box(0); } -if (lean_is_scalar(x_224)) { - x_225 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_223)) { + x_224 = lean_alloc_ctor(1, 2, 0); } else { - x_225 = x_224; + x_224 = x_223; } -lean_ctor_set(x_225, 0, x_222); -lean_ctor_set(x_225, 1, x_223); -return x_225; +lean_ctor_set(x_224, 0, x_221); +lean_ctor_set(x_224, 1, x_222); +return x_224; } } } else { -lean_object* x_226; lean_object* x_227; +lean_object* x_225; lean_object* x_226; lean_dec(x_12); lean_dec(x_11); -x_226 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__10; -x_227 = l_Lean_throwErrorAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__10(x_1, x_226, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_227; +x_225 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__10; +x_226 = l_Lean_throwErrorAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__10(x_1, x_225, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_226; } } case 2: { -lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; uint8_t x_234; +lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; uint8_t x_233; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_228 = lean_ctor_get(x_1, 1); -lean_inc(x_228); +x_227 = lean_ctor_get(x_1, 1); +lean_inc(x_227); lean_dec(x_1); lean_inc(x_6); -x_229 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_6, x_7, x_8); -x_230 = lean_ctor_get(x_229, 0); +x_228 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_6, x_7, x_8); +x_229 = lean_ctor_get(x_228, 0); +lean_inc(x_229); +x_230 = lean_ctor_get(x_228, 1); lean_inc(x_230); -x_231 = lean_ctor_get(x_229, 1); +lean_dec(x_228); +x_231 = lean_ctor_get(x_6, 10); lean_inc(x_231); -lean_dec(x_229); -x_232 = lean_ctor_get(x_6, 10); -lean_inc(x_232); lean_dec(x_6); -x_233 = lean_st_ref_get(x_7, x_231); +x_232 = lean_st_ref_get(x_7, x_230); lean_dec(x_7); -x_234 = !lean_is_exclusive(x_233); -if (x_234 == 0) -{ -lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; -x_235 = lean_ctor_get(x_233, 0); -x_236 = lean_ctor_get(x_235, 0); -lean_inc(x_236); -lean_dec(x_235); -x_237 = lean_environment_main_module(x_236); -x_238 = l_Lean_Elab_Term_quoteAutoTactic___closed__55; -x_239 = l_Lean_addMacroScope(x_237, x_238, x_232); -x_240 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; -x_241 = l_Lean_Elab_Term_quoteAutoTactic___closed__58; -x_242 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_242, 0, x_230); -lean_ctor_set(x_242, 1, x_240); -lean_ctor_set(x_242, 2, x_239); -lean_ctor_set(x_242, 3, x_241); -x_243 = lean_box(2); -x_244 = l_Lean_Syntax_mkStrLit(x_228, x_243); -lean_dec(x_228); -x_245 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; -x_246 = lean_array_push(x_245, x_244); -x_247 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; -x_248 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_248, 0, x_243); -lean_ctor_set(x_248, 1, x_247); -lean_ctor_set(x_248, 2, x_246); -x_249 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__8; -x_250 = lean_array_push(x_249, x_242); -x_251 = lean_array_push(x_250, x_248); -x_252 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__2; -x_253 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_253, 0, x_243); -lean_ctor_set(x_253, 1, x_252); -lean_ctor_set(x_253, 2, x_251); -lean_ctor_set(x_233, 0, x_253); -return x_233; -} -else +x_233 = !lean_is_exclusive(x_232); +if (x_233 == 0) { -lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; -x_254 = lean_ctor_get(x_233, 0); -x_255 = lean_ctor_get(x_233, 1); -lean_inc(x_255); +lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; +x_234 = lean_ctor_get(x_232, 0); +x_235 = lean_ctor_get(x_234, 0); +lean_inc(x_235); +lean_dec(x_234); +x_236 = lean_environment_main_module(x_235); +x_237 = l_Lean_Elab_Term_quoteAutoTactic___closed__58; +x_238 = l_Lean_addMacroScope(x_236, x_237, x_231); +x_239 = l_Lean_Elab_Term_quoteAutoTactic___closed__57; +x_240 = l_Lean_Elab_Term_quoteAutoTactic___closed__61; +x_241 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_241, 0, x_229); +lean_ctor_set(x_241, 1, x_239); +lean_ctor_set(x_241, 2, x_238); +lean_ctor_set(x_241, 3, x_240); +x_242 = lean_box(2); +x_243 = l_Lean_Syntax_mkStrLit(x_227, x_242); +lean_dec(x_227); +x_244 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; +x_245 = lean_array_push(x_244, x_243); +x_246 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; +x_247 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_247, 0, x_242); +lean_ctor_set(x_247, 1, x_246); +lean_ctor_set(x_247, 2, x_245); +x_248 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__8; +x_249 = lean_array_push(x_248, x_241); +x_250 = lean_array_push(x_249, x_247); +x_251 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__2; +x_252 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_252, 0, x_242); +lean_ctor_set(x_252, 1, x_251); +lean_ctor_set(x_252, 2, x_250); +lean_ctor_set(x_232, 0, x_252); +return x_232; +} +else +{ +lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; +x_253 = lean_ctor_get(x_232, 0); +x_254 = lean_ctor_get(x_232, 1); lean_inc(x_254); -lean_dec(x_233); -x_256 = lean_ctor_get(x_254, 0); -lean_inc(x_256); -lean_dec(x_254); -x_257 = lean_environment_main_module(x_256); -x_258 = l_Lean_Elab_Term_quoteAutoTactic___closed__55; -x_259 = l_Lean_addMacroScope(x_257, x_258, x_232); -x_260 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; -x_261 = l_Lean_Elab_Term_quoteAutoTactic___closed__58; -x_262 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_262, 0, x_230); -lean_ctor_set(x_262, 1, x_260); -lean_ctor_set(x_262, 2, x_259); -lean_ctor_set(x_262, 3, x_261); -x_263 = lean_box(2); -x_264 = l_Lean_Syntax_mkStrLit(x_228, x_263); -lean_dec(x_228); -x_265 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; -x_266 = lean_array_push(x_265, x_264); -x_267 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; -x_268 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_268, 0, x_263); -lean_ctor_set(x_268, 1, x_267); -lean_ctor_set(x_268, 2, x_266); -x_269 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__8; -x_270 = lean_array_push(x_269, x_262); -x_271 = lean_array_push(x_270, x_268); -x_272 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__2; -x_273 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_273, 0, x_263); -lean_ctor_set(x_273, 1, x_272); -lean_ctor_set(x_273, 2, x_271); -x_274 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_274, 0, x_273); -lean_ctor_set(x_274, 1, x_255); -return x_274; +lean_inc(x_253); +lean_dec(x_232); +x_255 = lean_ctor_get(x_253, 0); +lean_inc(x_255); +lean_dec(x_253); +x_256 = lean_environment_main_module(x_255); +x_257 = l_Lean_Elab_Term_quoteAutoTactic___closed__58; +x_258 = l_Lean_addMacroScope(x_256, x_257, x_231); +x_259 = l_Lean_Elab_Term_quoteAutoTactic___closed__57; +x_260 = l_Lean_Elab_Term_quoteAutoTactic___closed__61; +x_261 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_261, 0, x_229); +lean_ctor_set(x_261, 1, x_259); +lean_ctor_set(x_261, 2, x_258); +lean_ctor_set(x_261, 3, x_260); +x_262 = lean_box(2); +x_263 = l_Lean_Syntax_mkStrLit(x_227, x_262); +lean_dec(x_227); +x_264 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; +x_265 = lean_array_push(x_264, x_263); +x_266 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; +x_267 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_267, 0, x_262); +lean_ctor_set(x_267, 1, x_266); +lean_ctor_set(x_267, 2, x_265); +x_268 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__8; +x_269 = lean_array_push(x_268, x_261); +x_270 = lean_array_push(x_269, x_267); +x_271 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__2; +x_272 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_272, 0, x_262); +lean_ctor_set(x_272, 1, x_271); +lean_ctor_set(x_272, 2, x_270); +x_273 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_273, 0, x_272); +lean_ctor_set(x_273, 1, x_254); +return x_273; } } default: { -lean_object* x_275; lean_object* x_276; -x_275 = l_Lean_Elab_Term_quoteAutoTactic___closed__60; -x_276 = l_Lean_throwErrorAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__10(x_1, x_275, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_276; +lean_object* x_274; lean_object* x_275; +x_274 = l_Lean_Elab_Term_quoteAutoTactic___closed__63; +x_275 = l_Lean_throwErrorAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__10(x_1, x_274, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_275; } } } @@ -2809,18 +2840,20 @@ lean_object* x_14 = _args[13]; lean_object* x_15 = _args[14]; lean_object* x_16 = _args[15]; lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; _start: { -size_t x_18; size_t x_19; lean_object* x_20; -x_18 = lean_unbox_usize(x_8); -lean_dec(x_8); +size_t x_19; size_t x_20; lean_object* x_21; x_19 = lean_unbox_usize(x_9); lean_dec(x_9); -x_20 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_18, x_19, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -lean_dec(x_7); +x_20 = lean_unbox_usize(x_10); +lean_dec(x_10); +x_21 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_19, x_20, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_8); +lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_20; +return x_21; } } LEAN_EXPORT lean_object* l_Lean_Elab_Term_declareTacticSyntax___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, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { @@ -2947,7 +2980,7 @@ static lean_object* _init_l_Lean_Elab_Term_declareTacticSyntax___lambda__2___clo { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__31; +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__34; x_3 = l_Lean_mkConst(x_2, x_1); return x_3; } @@ -3691,15 +3724,33 @@ static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binder _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("identifier or `_` expected", 26); +x_1 = lean_mk_string_from_bytes("ident", 5); return x_1; } } static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__2() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("identifier or `_` expected", 26); +return x_1; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__4() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__1; +x_1 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__3; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } @@ -3731,7 +3782,7 @@ x_14 = lean_unsigned_to_nat(0u); x_15 = lean_array_uset(x_3, x_2, x_14); lean_inc(x_13); x_16 = l_Lean_Syntax_getKind(x_13); -x_17 = l_Lean_identKind; +x_17 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__2; x_18 = lean_name_eq(x_16, x_17); if (x_18 == 0) { @@ -3743,7 +3794,7 @@ if (x_20 == 0) { lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_dec(x_15); -x_21 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__2; +x_21 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__4; x_22 = l_Lean_throwErrorAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__10(x_13, x_21, x_4, x_5, x_6, x_7, x_8, x_9, x_10); x_23 = !lean_is_exclusive(x_22); if (x_23 == 0) @@ -4370,7 +4421,7 @@ x_29 = lean_alloc_ctor(0, 2, 1); lean_ctor_set(x_29, 0, x_25); lean_ctor_set(x_29, 1, x_27); lean_ctor_set_uint8(x_29, sizeof(void*)*2, x_28); -x_30 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_30 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_31 = lean_array_push(x_30, x_29); lean_ctor_set(x_23, 0, x_31); return x_23; @@ -4391,7 +4442,7 @@ x_37 = lean_alloc_ctor(0, 2, 1); lean_ctor_set(x_37, 0, x_32); lean_ctor_set(x_37, 1, x_35); lean_ctor_set_uint8(x_37, sizeof(void*)*2, x_36); -x_38 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_38 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_39 = lean_array_push(x_38, x_37); x_40 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_40, 0, x_39); @@ -4993,7 +5044,7 @@ lean_dec(x_1); return x_10; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1437____closed__1() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1498____closed__1() { _start: { lean_object* x_1; @@ -5001,17 +5052,17 @@ x_1 = lean_mk_string_from_bytes("checkBinderAnnotations", 22); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1437____closed__2() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1498____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1437____closed__1; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1498____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1437____closed__3() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1498____closed__3() { _start: { lean_object* x_1; @@ -5019,7 +5070,7 @@ x_1 = lean_mk_string_from_bytes("", 0); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1437____closed__4() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1498____closed__4() { _start: { lean_object* x_1; @@ -5027,13 +5078,13 @@ x_1 = lean_mk_string_from_bytes("check whether type is a class instance whenever return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1437____closed__5() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1498____closed__5() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = 1; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1437____closed__3; -x_3 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1437____closed__4; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1498____closed__3; +x_3 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1498____closed__4; x_4 = lean_box(x_1); x_5 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_5, 0, x_4); @@ -5042,12 +5093,12 @@ lean_ctor_set(x_5, 2, x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1437_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1498_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1437____closed__2; -x_3 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1437____closed__5; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1498____closed__2; +x_3 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1498____closed__5; x_4 = l_Lean_Option_register___at_Lean_Elab_initFn____x40_Lean_Elab_AutoBound___hyg_6____spec__1(x_2, x_3, x_1); return x_4; } @@ -5716,7 +5767,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabBinder___rarg(lean_object* x_1, le _start: { lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_10 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_10 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_11 = lean_array_push(x_10, x_1); x_12 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabBinder___rarg___lambda__1___boxed), 9, 1); lean_closure_set(x_12, 0, x_2); @@ -5964,7 +6015,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabForall_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(208u); +x_1 = lean_unsigned_to_nat(209u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -5976,7 +6027,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabForall_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(214u); +x_1 = lean_unsigned_to_nat(215u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6004,7 +6055,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabForall_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(208u); +x_1 = lean_unsigned_to_nat(209u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6016,7 +6067,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabForall_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(208u); +x_1 = lean_unsigned_to_nat(209u); x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6432,7 +6483,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabArrow_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(223u); +x_1 = lean_unsigned_to_nat(224u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6444,7 +6495,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabArrow_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(230u); +x_1 = lean_unsigned_to_nat(231u); x_2 = lean_unsigned_to_nat(50u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6472,7 +6523,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabArrow_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(223u); +x_1 = lean_unsigned_to_nat(224u); x_2 = lean_unsigned_to_nat(29u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6484,7 +6535,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabArrow_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(223u); +x_1 = lean_unsigned_to_nat(224u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6538,7 +6589,7 @@ x_10 = lean_unsigned_to_nat(0u); x_11 = l_Lean_Syntax_getArg(x_1, x_10); x_12 = lean_unsigned_to_nat(2u); x_13 = l_Lean_Syntax_getArg(x_1, x_12); -x_14 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_14 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_15 = lean_array_push(x_14, x_11); x_16 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabForall___lambda__1), 9, 1); lean_closure_set(x_16, 0, x_13); @@ -6634,7 +6685,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabDepArrow_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(236u); +x_1 = lean_unsigned_to_nat(237u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6646,7 +6697,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabDepArrow_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(241u); +x_1 = lean_unsigned_to_nat(242u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6674,7 +6725,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabDepArrow_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(236u); +x_1 = lean_unsigned_to_nat(237u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6686,7 +6737,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabDepArrow_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(236u); +x_1 = lean_unsigned_to_nat(237u); x_2 = lean_unsigned_to_nat(44u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6846,24 +6897,6 @@ return x_19; } } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getFunBinderIds_x3f___spec__3___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("ident", 5); -return x_1; -} -} -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getFunBinderIds_x3f___spec__3___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getFunBinderIds_x3f___spec__3___closed__1; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getFunBinderIds_x3f___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) { _start: { @@ -6894,7 +6927,7 @@ lean_dec(x_23); if (x_24 == 0) { lean_object* x_25; uint8_t x_26; -x_25 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getFunBinderIds_x3f___spec__3___closed__2; +x_25 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__2; lean_inc(x_21); x_26 = l_Lean_Syntax_isOfKind(x_21, x_25); if (x_26 == 0) @@ -6995,7 +7028,7 @@ if (x_22 == 0) { lean_object* x_23; uint8_t x_24; lean_dec(x_2); -x_23 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getFunBinderIds_x3f___spec__3___closed__2; +x_23 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__2; lean_inc(x_1); x_24 = l_Lean_Syntax_isOfKind(x_1, x_23); if (x_24 == 0) @@ -7063,7 +7096,7 @@ x_69 = l_Lean_Syntax_isOfKind(x_34, x_68); if (x_69 == 0) { lean_object* x_70; uint8_t x_71; -x_70 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getFunBinderIds_x3f___spec__3___closed__2; +x_70 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__2; lean_inc(x_34); x_71 = l_Lean_Syntax_isOfKind(x_34, x_70); if (x_71 == 0) @@ -7842,7 +7875,7 @@ x_41 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_41, 0, x_39); lean_ctor_set(x_41, 1, x_40); lean_ctor_set(x_41, 2, x_38); -x_42 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_42 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_43 = lean_array_push(x_42, x_41); x_44 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_45 = lean_alloc_ctor(1, 3, 0); @@ -7935,7 +7968,7 @@ x_88 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_88, 0, x_86); lean_ctor_set(x_88, 1, x_87); lean_ctor_set(x_88, 2, x_85); -x_89 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_89 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_90 = lean_array_push(x_89, x_88); x_91 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_92 = lean_alloc_ctor(1, 3, 0); @@ -8042,7 +8075,7 @@ x_139 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_139, 0, x_137); lean_ctor_set(x_139, 1, x_138); lean_ctor_set(x_139, 2, x_136); -x_140 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_140 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_141 = lean_array_push(x_140, x_139); x_142 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_143 = lean_alloc_ctor(1, 3, 0); @@ -8165,7 +8198,7 @@ x_193 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_193, 0, x_191); lean_ctor_set(x_193, 1, x_192); lean_ctor_set(x_193, 2, x_190); -x_194 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_194 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_195 = lean_array_push(x_194, x_193); x_196 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_197 = lean_alloc_ctor(1, 3, 0); @@ -8386,7 +8419,7 @@ x_276 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_276, 0, x_274); lean_ctor_set(x_276, 1, x_275); lean_ctor_set(x_276, 2, x_273); -x_277 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_277 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_278 = lean_array_push(x_277, x_276); x_279 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_280 = lean_alloc_ctor(1, 3, 0); @@ -8548,7 +8581,7 @@ x_355 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_355, 0, x_353); lean_ctor_set(x_355, 1, x_354); lean_ctor_set(x_355, 2, x_352); -x_356 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_356 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_357 = lean_array_push(x_356, x_355); x_358 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_359 = lean_alloc_ctor(1, 3, 0); @@ -8669,7 +8702,7 @@ x_407 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_407, 0, x_405); lean_ctor_set(x_407, 1, x_406); lean_ctor_set(x_407, 2, x_404); -x_408 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_408 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_409 = lean_array_push(x_408, x_407); x_410 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_411 = lean_alloc_ctor(1, 3, 0); @@ -8806,7 +8839,7 @@ x_462 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_462, 0, x_460); lean_ctor_set(x_462, 1, x_461); lean_ctor_set(x_462, 2, x_459); -x_463 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_463 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_464 = lean_array_push(x_463, x_462); x_465 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_466 = lean_alloc_ctor(1, 3, 0); @@ -9004,7 +9037,7 @@ x_537 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_537, 0, x_535); lean_ctor_set(x_537, 1, x_536); lean_ctor_set(x_537, 2, x_534); -x_538 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_538 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_539 = lean_array_push(x_538, x_537); x_540 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_541 = lean_alloc_ctor(1, 3, 0); @@ -9166,7 +9199,7 @@ x_616 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_616, 0, x_614); lean_ctor_set(x_616, 1, x_615); lean_ctor_set(x_616, 2, x_613); -x_617 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_617 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_618 = lean_array_push(x_617, x_616); x_619 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_620 = lean_alloc_ctor(1, 3, 0); @@ -9287,7 +9320,7 @@ x_668 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_668, 0, x_666); lean_ctor_set(x_668, 1, x_667); lean_ctor_set(x_668, 2, x_665); -x_669 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_669 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_670 = lean_array_push(x_669, x_668); x_671 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_672 = lean_alloc_ctor(1, 3, 0); @@ -9424,7 +9457,7 @@ x_723 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_723, 0, x_721); lean_ctor_set(x_723, 1, x_722); lean_ctor_set(x_723, 2, x_720); -x_724 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_724 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_725 = lean_array_push(x_724, x_723); x_726 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_727 = lean_alloc_ctor(1, 3, 0); @@ -9621,7 +9654,7 @@ x_798 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_798, 0, x_796); lean_ctor_set(x_798, 1, x_797); lean_ctor_set(x_798, 2, x_795); -x_799 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_799 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_800 = lean_array_push(x_799, x_798); x_801 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_802 = lean_alloc_ctor(1, 3, 0); @@ -9783,7 +9816,7 @@ x_877 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_877, 0, x_875); lean_ctor_set(x_877, 1, x_876); lean_ctor_set(x_877, 2, x_874); -x_878 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_878 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_879 = lean_array_push(x_878, x_877); x_880 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_881 = lean_alloc_ctor(1, 3, 0); @@ -9904,7 +9937,7 @@ x_929 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_929, 0, x_927); lean_ctor_set(x_929, 1, x_928); lean_ctor_set(x_929, 2, x_926); -x_930 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_930 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_931 = lean_array_push(x_930, x_929); x_932 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_933 = lean_alloc_ctor(1, 3, 0); @@ -10041,7 +10074,7 @@ x_984 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_984, 0, x_982); lean_ctor_set(x_984, 1, x_983); lean_ctor_set(x_984, 2, x_981); -x_985 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_985 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_986 = lean_array_push(x_985, x_984); x_987 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_988 = lean_alloc_ctor(1, 3, 0); @@ -10267,7 +10300,7 @@ x_1071 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_1071, 0, x_1069); lean_ctor_set(x_1071, 1, x_1070); lean_ctor_set(x_1071, 2, x_1068); -x_1072 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_1072 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_1073 = lean_array_push(x_1072, x_1071); x_1074 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_1075 = lean_alloc_ctor(1, 3, 0); @@ -10429,7 +10462,7 @@ x_1150 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_1150, 0, x_1148); lean_ctor_set(x_1150, 1, x_1149); lean_ctor_set(x_1150, 2, x_1147); -x_1151 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_1151 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_1152 = lean_array_push(x_1151, x_1150); x_1153 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_1154 = lean_alloc_ctor(1, 3, 0); @@ -10550,7 +10583,7 @@ x_1202 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_1202, 0, x_1200); lean_ctor_set(x_1202, 1, x_1201); lean_ctor_set(x_1202, 2, x_1199); -x_1203 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_1203 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_1204 = lean_array_push(x_1203, x_1202); x_1205 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_1206 = lean_alloc_ctor(1, 3, 0); @@ -10687,7 +10720,7 @@ x_1257 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_1257, 0, x_1255); lean_ctor_set(x_1257, 1, x_1256); lean_ctor_set(x_1257, 2, x_1254); -x_1258 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_1258 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_1259 = lean_array_push(x_1258, x_1257); x_1260 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_1261 = lean_alloc_ctor(1, 3, 0); @@ -10902,7 +10935,7 @@ x_1340 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_1340, 0, x_1338); lean_ctor_set(x_1340, 1, x_1339); lean_ctor_set(x_1340, 2, x_1337); -x_1341 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_1341 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_1342 = lean_array_push(x_1341, x_1340); x_1343 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_1344 = lean_alloc_ctor(1, 3, 0); @@ -11064,7 +11097,7 @@ x_1419 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_1419, 0, x_1417); lean_ctor_set(x_1419, 1, x_1418); lean_ctor_set(x_1419, 2, x_1416); -x_1420 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_1420 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_1421 = lean_array_push(x_1420, x_1419); x_1422 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_1423 = lean_alloc_ctor(1, 3, 0); @@ -11185,7 +11218,7 @@ x_1471 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_1471, 0, x_1469); lean_ctor_set(x_1471, 1, x_1470); lean_ctor_set(x_1471, 2, x_1468); -x_1472 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_1472 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_1473 = lean_array_push(x_1472, x_1471); x_1474 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_1475 = lean_alloc_ctor(1, 3, 0); @@ -11322,7 +11355,7 @@ x_1526 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_1526, 0, x_1524); lean_ctor_set(x_1526, 1, x_1525); lean_ctor_set(x_1526, 2, x_1523); -x_1527 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_1527 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_1528 = lean_array_push(x_1527, x_1526); x_1529 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_1530 = lean_alloc_ctor(1, 3, 0); @@ -11524,7 +11557,7 @@ x_1602 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_1602, 0, x_1600); lean_ctor_set(x_1602, 1, x_1601); lean_ctor_set(x_1602, 2, x_1599); -x_1603 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_1603 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_1604 = lean_array_push(x_1603, x_1602); x_1605 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_1606 = lean_alloc_ctor(1, 3, 0); @@ -11686,7 +11719,7 @@ x_1681 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_1681, 0, x_1679); lean_ctor_set(x_1681, 1, x_1680); lean_ctor_set(x_1681, 2, x_1678); -x_1682 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_1682 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_1683 = lean_array_push(x_1682, x_1681); x_1684 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_1685 = lean_alloc_ctor(1, 3, 0); @@ -11807,7 +11840,7 @@ x_1733 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_1733, 0, x_1731); lean_ctor_set(x_1733, 1, x_1732); lean_ctor_set(x_1733, 2, x_1730); -x_1734 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_1734 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_1735 = lean_array_push(x_1734, x_1733); x_1736 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_1737 = lean_alloc_ctor(1, 3, 0); @@ -11944,7 +11977,7 @@ x_1788 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_1788, 0, x_1786); lean_ctor_set(x_1788, 1, x_1787); lean_ctor_set(x_1788, 2, x_1785); -x_1789 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_1789 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_1790 = lean_array_push(x_1789, x_1788); x_1791 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_1792 = lean_alloc_ctor(1, 3, 0); @@ -12169,7 +12202,7 @@ x_1872 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_1872, 0, x_1870); lean_ctor_set(x_1872, 1, x_1871); lean_ctor_set(x_1872, 2, x_1869); -x_1873 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_1873 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_1874 = lean_array_push(x_1873, x_1872); x_1875 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_1876 = lean_alloc_ctor(1, 3, 0); @@ -12331,7 +12364,7 @@ x_1951 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_1951, 0, x_1949); lean_ctor_set(x_1951, 1, x_1950); lean_ctor_set(x_1951, 2, x_1948); -x_1952 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_1952 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_1953 = lean_array_push(x_1952, x_1951); x_1954 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_1955 = lean_alloc_ctor(1, 3, 0); @@ -12452,7 +12485,7 @@ x_2003 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_2003, 0, x_2001); lean_ctor_set(x_2003, 1, x_2002); lean_ctor_set(x_2003, 2, x_2000); -x_2004 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_2004 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_2005 = lean_array_push(x_2004, x_2003); x_2006 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_2007 = lean_alloc_ctor(1, 3, 0); @@ -12589,7 +12622,7 @@ x_2058 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_2058, 0, x_2056); lean_ctor_set(x_2058, 1, x_2057); lean_ctor_set(x_2058, 2, x_2055); -x_2059 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_2059 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_2060 = lean_array_push(x_2059, x_2058); x_2061 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_2062 = lean_alloc_ctor(1, 3, 0); @@ -12894,7 +12927,7 @@ x_2134 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_2134, 0, x_2132); lean_ctor_set(x_2134, 1, x_2133); lean_ctor_set(x_2134, 2, x_2131); -x_2135 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_2135 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_2136 = lean_array_push(x_2135, x_2134); x_2137 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_2138 = lean_alloc_ctor(1, 3, 0); @@ -13056,7 +13089,7 @@ x_2213 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_2213, 0, x_2211); lean_ctor_set(x_2213, 1, x_2212); lean_ctor_set(x_2213, 2, x_2210); -x_2214 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_2214 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_2215 = lean_array_push(x_2214, x_2213); x_2216 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_2217 = lean_alloc_ctor(1, 3, 0); @@ -13177,7 +13210,7 @@ x_2265 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_2265, 0, x_2263); lean_ctor_set(x_2265, 1, x_2264); lean_ctor_set(x_2265, 2, x_2262); -x_2266 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_2266 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_2267 = lean_array_push(x_2266, x_2265); x_2268 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_2269 = lean_alloc_ctor(1, 3, 0); @@ -13314,7 +13347,7 @@ x_2320 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_2320, 0, x_2318); lean_ctor_set(x_2320, 1, x_2319); lean_ctor_set(x_2320, 2, x_2317); -x_2321 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_2321 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_2322 = lean_array_push(x_2321, x_2320); x_2323 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_2324 = lean_alloc_ctor(1, 3, 0); @@ -13524,7 +13557,7 @@ x_2428 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_2428, 0, x_2426); lean_ctor_set(x_2428, 1, x_2427); lean_ctor_set(x_2428, 2, x_2425); -x_2429 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_2429 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_2430 = lean_array_push(x_2429, x_2428); x_2431 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_2432 = lean_alloc_ctor(1, 3, 0); @@ -13686,7 +13719,7 @@ x_2507 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_2507, 0, x_2505); lean_ctor_set(x_2507, 1, x_2506); lean_ctor_set(x_2507, 2, x_2504); -x_2508 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_2508 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_2509 = lean_array_push(x_2508, x_2507); x_2510 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_2511 = lean_alloc_ctor(1, 3, 0); @@ -13807,7 +13840,7 @@ x_2559 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_2559, 0, x_2557); lean_ctor_set(x_2559, 1, x_2558); lean_ctor_set(x_2559, 2, x_2556); -x_2560 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_2560 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_2561 = lean_array_push(x_2560, x_2559); x_2562 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_2563 = lean_alloc_ctor(1, 3, 0); @@ -13944,7 +13977,7 @@ x_2614 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_2614, 0, x_2612); lean_ctor_set(x_2614, 1, x_2613); lean_ctor_set(x_2614, 2, x_2611); -x_2615 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_2615 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_2616 = lean_array_push(x_2615, x_2614); x_2617 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_2618 = lean_alloc_ctor(1, 3, 0); @@ -14232,7 +14265,7 @@ x_2715 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_2715, 0, x_2713); lean_ctor_set(x_2715, 1, x_2714); lean_ctor_set(x_2715, 2, x_2712); -x_2716 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_2716 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_2717 = lean_array_push(x_2716, x_2715); x_2718 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_2719 = lean_alloc_ctor(1, 3, 0); @@ -14394,7 +14427,7 @@ x_2794 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_2794, 0, x_2792); lean_ctor_set(x_2794, 1, x_2793); lean_ctor_set(x_2794, 2, x_2791); -x_2795 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_2795 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_2796 = lean_array_push(x_2795, x_2794); x_2797 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_2798 = lean_alloc_ctor(1, 3, 0); @@ -14515,7 +14548,7 @@ x_2846 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_2846, 0, x_2844); lean_ctor_set(x_2846, 1, x_2845); lean_ctor_set(x_2846, 2, x_2843); -x_2847 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_2847 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_2848 = lean_array_push(x_2847, x_2846); x_2849 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_2850 = lean_alloc_ctor(1, 3, 0); @@ -14652,7 +14685,7 @@ x_2901 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_2901, 0, x_2899); lean_ctor_set(x_2901, 1, x_2900); lean_ctor_set(x_2901, 2, x_2898); -x_2902 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_2902 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_2903 = lean_array_push(x_2902, x_2901); x_2904 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_2905 = lean_alloc_ctor(1, 3, 0); @@ -14847,7 +14880,7 @@ x_2974 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_2974, 0, x_2972); lean_ctor_set(x_2974, 1, x_2973); lean_ctor_set(x_2974, 2, x_2971); -x_2975 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_2975 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_2976 = lean_array_push(x_2975, x_2974); x_2977 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_2978 = lean_alloc_ctor(1, 3, 0); @@ -15009,7 +15042,7 @@ x_3053 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_3053, 0, x_3051); lean_ctor_set(x_3053, 1, x_3052); lean_ctor_set(x_3053, 2, x_3050); -x_3054 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_3054 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_3055 = lean_array_push(x_3054, x_3053); x_3056 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_3057 = lean_alloc_ctor(1, 3, 0); @@ -15130,7 +15163,7 @@ x_3105 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_3105, 0, x_3103); lean_ctor_set(x_3105, 1, x_3104); lean_ctor_set(x_3105, 2, x_3102); -x_3106 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_3106 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_3107 = lean_array_push(x_3106, x_3105); x_3108 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_3109 = lean_alloc_ctor(1, 3, 0); @@ -15267,7 +15300,7 @@ x_3160 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_3160, 0, x_3158); lean_ctor_set(x_3160, 1, x_3159); lean_ctor_set(x_3160, 2, x_3157); -x_3161 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_3161 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_3162 = lean_array_push(x_3161, x_3160); x_3163 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_3164 = lean_alloc_ctor(1, 3, 0); @@ -15461,7 +15494,7 @@ x_3233 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_3233, 0, x_3231); lean_ctor_set(x_3233, 1, x_3232); lean_ctor_set(x_3233, 2, x_3230); -x_3234 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_3234 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_3235 = lean_array_push(x_3234, x_3233); x_3236 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_3237 = lean_alloc_ctor(1, 3, 0); @@ -15623,7 +15656,7 @@ x_3312 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_3312, 0, x_3310); lean_ctor_set(x_3312, 1, x_3311); lean_ctor_set(x_3312, 2, x_3309); -x_3313 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_3313 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_3314 = lean_array_push(x_3313, x_3312); x_3315 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_3316 = lean_alloc_ctor(1, 3, 0); @@ -15744,7 +15777,7 @@ x_3364 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_3364, 0, x_3362); lean_ctor_set(x_3364, 1, x_3363); lean_ctor_set(x_3364, 2, x_3361); -x_3365 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_3365 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_3366 = lean_array_push(x_3365, x_3364); x_3367 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_3368 = lean_alloc_ctor(1, 3, 0); @@ -15881,7 +15914,7 @@ x_3419 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_3419, 0, x_3417); lean_ctor_set(x_3419, 1, x_3418); lean_ctor_set(x_3419, 2, x_3416); -x_3420 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_3420 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_3421 = lean_array_push(x_3420, x_3419); x_3422 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_3423 = lean_alloc_ctor(1, 3, 0); @@ -16074,7 +16107,7 @@ x_3492 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_3492, 0, x_3490); lean_ctor_set(x_3492, 1, x_3491); lean_ctor_set(x_3492, 2, x_3489); -x_3493 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_3493 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_3494 = lean_array_push(x_3493, x_3492); x_3495 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_3496 = lean_alloc_ctor(1, 3, 0); @@ -16236,7 +16269,7 @@ x_3571 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_3571, 0, x_3569); lean_ctor_set(x_3571, 1, x_3570); lean_ctor_set(x_3571, 2, x_3568); -x_3572 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_3572 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_3573 = lean_array_push(x_3572, x_3571); x_3574 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_3575 = lean_alloc_ctor(1, 3, 0); @@ -16357,7 +16390,7 @@ x_3623 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_3623, 0, x_3621); lean_ctor_set(x_3623, 1, x_3622); lean_ctor_set(x_3623, 2, x_3620); -x_3624 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_3624 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_3625 = lean_array_push(x_3624, x_3623); x_3626 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_3627 = lean_alloc_ctor(1, 3, 0); @@ -16494,7 +16527,7 @@ x_3678 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_3678, 0, x_3676); lean_ctor_set(x_3678, 1, x_3677); lean_ctor_set(x_3678, 2, x_3675); -x_3679 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_3679 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_3680 = lean_array_push(x_3679, x_3678); x_3681 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_3682 = lean_alloc_ctor(1, 3, 0); @@ -16686,7 +16719,7 @@ x_3751 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_3751, 0, x_3749); lean_ctor_set(x_3751, 1, x_3750); lean_ctor_set(x_3751, 2, x_3748); -x_3752 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_3752 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_3753 = lean_array_push(x_3752, x_3751); x_3754 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_3755 = lean_alloc_ctor(1, 3, 0); @@ -16848,7 +16881,7 @@ x_3830 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_3830, 0, x_3828); lean_ctor_set(x_3830, 1, x_3829); lean_ctor_set(x_3830, 2, x_3827); -x_3831 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_3831 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_3832 = lean_array_push(x_3831, x_3830); x_3833 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_3834 = lean_alloc_ctor(1, 3, 0); @@ -16969,7 +17002,7 @@ x_3882 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_3882, 0, x_3880); lean_ctor_set(x_3882, 1, x_3881); lean_ctor_set(x_3882, 2, x_3879); -x_3883 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_3883 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_3884 = lean_array_push(x_3883, x_3882); x_3885 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_3886 = lean_alloc_ctor(1, 3, 0); @@ -17106,7 +17139,7 @@ x_3937 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_3937, 0, x_3935); lean_ctor_set(x_3937, 1, x_3936); lean_ctor_set(x_3937, 2, x_3934); -x_3938 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_3938 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_3939 = lean_array_push(x_3938, x_3937); x_3940 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_3941 = lean_alloc_ctor(1, 3, 0); @@ -17297,7 +17330,7 @@ x_4010 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4010, 0, x_4008); lean_ctor_set(x_4010, 1, x_4009); lean_ctor_set(x_4010, 2, x_4007); -x_4011 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_4011 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_4012 = lean_array_push(x_4011, x_4010); x_4013 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_4014 = lean_alloc_ctor(1, 3, 0); @@ -17457,7 +17490,7 @@ x_4088 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4088, 0, x_4086); lean_ctor_set(x_4088, 1, x_4087); lean_ctor_set(x_4088, 2, x_4085); -x_4089 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_4089 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_4090 = lean_array_push(x_4089, x_4088); x_4091 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_4092 = lean_alloc_ctor(1, 3, 0); @@ -17577,7 +17610,7 @@ x_4140 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4140, 0, x_4138); lean_ctor_set(x_4140, 1, x_4139); lean_ctor_set(x_4140, 2, x_4137); -x_4141 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_4141 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_4142 = lean_array_push(x_4141, x_4140); x_4143 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_4144 = lean_alloc_ctor(1, 3, 0); @@ -17713,7 +17746,7 @@ x_4195 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4195, 0, x_4193); lean_ctor_set(x_4195, 1, x_4194); lean_ctor_set(x_4195, 2, x_4192); -x_4196 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_4196 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_4197 = lean_array_push(x_4196, x_4195); x_4198 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_4199 = lean_alloc_ctor(1, 3, 0); @@ -19887,6 +19920,86 @@ lean_dec(x_1); return x_11; } } +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Elab_Term_expandWhereDecls___spec__2(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; uint8_t x_7; +x_6 = lean_array_get_size(x_1); +x_7 = lean_nat_dec_lt(x_4, x_6); +lean_dec(x_6); +if (x_7 == 0) +{ +lean_object* x_8; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_5); +return x_8; +} +else +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_unsigned_to_nat(0u); +x_10 = lean_nat_dec_eq(x_3, x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_sub(x_3, x_11); +lean_dec(x_3); +x_13 = lean_array_fget(x_1, x_4); +lean_inc(x_2); +x_14 = lean_apply_1(x_2, x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; +lean_dec(x_12); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_15 = lean_box(0); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_14, 0); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_nat_add(x_4, x_11); +lean_dec(x_4); +x_18 = lean_array_push(x_5, x_16); +x_3 = x_12; +x_4 = x_17; +x_5 = x_18; +goto _start; +} +} +else +{ +lean_object* x_20; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_20 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_20, 0, x_5); +return x_20; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Elab_Term_expandWhereDecls___spec__1(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; +x_3 = lean_array_get_size(x_1); +x_4 = lean_mk_empty_array_with_capacity(x_3); +x_5 = lean_unsigned_to_nat(0u); +x_6 = l_Array_sequenceMap_loop___at_Lean_Elab_Term_expandWhereDecls___spec__2(x_1, x_2, x_3, x_5, x_4); +return x_6; +} +} static lean_object* _init_l_Lean_Elab_Term_expandWhereDecls___lambda__1___closed__1() { _start: { @@ -20032,6 +20145,18 @@ return x_1; static lean_object* _init_l_Lean_Elab_Term_expandWhereDecls___closed__13() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(2); +x_2 = l_Lean_Elab_Term_expandWhereDecls___closed__12; +x_3 = lean_alloc_ctor(2, 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_Lean_Elab_Term_expandWhereDecls___closed__14() { +_start: +{ lean_object* x_1; x_1 = lean_mk_string_from_bytes(";", 1); return x_1; @@ -20111,7 +20236,7 @@ goto block_86; block_86: { lean_object* x_18; -x_18 = l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__1(x_17, x_16); +x_18 = l_Array_sequenceMap___at_Lean_Elab_Term_expandWhereDecls___spec__1(x_17, x_16); lean_dec(x_17); if (lean_obj_tag(x_18) == 0) { @@ -20155,8 +20280,8 @@ x_34 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_34, 0, x_32); lean_ctor_set(x_34, 1, x_33); lean_ctor_set(x_34, 2, x_31); -x_35 = l_Lean_Elab_Term_expandWhereDecls___closed__12; -x_36 = l_Lean_Syntax_SepArray_ofElems(x_35, x_21); +x_35 = l_Lean_Elab_Term_expandWhereDecls___closed__13; +x_36 = l_Lean_mkSepArray(x_21, x_35); lean_dec(x_21); x_37 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; x_38 = l_Array_append___rarg(x_37, x_36); @@ -20165,14 +20290,14 @@ x_40 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_40, 0, x_32); lean_ctor_set(x_40, 1, x_39); lean_ctor_set(x_40, 2, x_38); -x_41 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_41 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_42 = lean_array_push(x_41, x_40); x_43 = l_Lean_Elab_Term_expandWhereDecls___closed__11; x_44 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_44, 0, x_32); lean_ctor_set(x_44, 1, x_43); lean_ctor_set(x_44, 2, x_42); -x_45 = l_Lean_Elab_Term_expandWhereDecls___closed__13; +x_45 = l_Lean_Elab_Term_expandWhereDecls___closed__14; x_46 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_46, 0, x_24); lean_ctor_set(x_46, 1, x_45); @@ -20216,8 +20341,8 @@ x_65 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_65, 0, x_63); lean_ctor_set(x_65, 1, x_64); lean_ctor_set(x_65, 2, x_62); -x_66 = l_Lean_Elab_Term_expandWhereDecls___closed__12; -x_67 = l_Lean_Syntax_SepArray_ofElems(x_66, x_21); +x_66 = l_Lean_Elab_Term_expandWhereDecls___closed__13; +x_67 = l_Lean_mkSepArray(x_21, x_66); lean_dec(x_21); x_68 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; x_69 = l_Array_append___rarg(x_68, x_67); @@ -20226,14 +20351,14 @@ x_71 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_71, 0, x_63); lean_ctor_set(x_71, 1, x_70); lean_ctor_set(x_71, 2, x_69); -x_72 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_72 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_73 = lean_array_push(x_72, x_71); x_74 = l_Lean_Elab_Term_expandWhereDecls___closed__11; x_75 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_75, 0, x_63); lean_ctor_set(x_75, 1, x_74); lean_ctor_set(x_75, 2, x_73); -x_76 = l_Lean_Elab_Term_expandWhereDecls___closed__13; +x_76 = l_Lean_Elab_Term_expandWhereDecls___closed__14; x_77 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_77, 0, x_54); lean_ctor_set(x_77, 1, x_76); @@ -20257,6 +20382,24 @@ return x_85; } } } +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Elab_Term_expandWhereDecls___spec__2___boxed(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; +x_6 = l_Array_sequenceMap_loop___at_Lean_Elab_Term_expandWhereDecls___spec__2(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Elab_Term_expandWhereDecls___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Array_sequenceMap___at_Lean_Elab_Term_expandWhereDecls___spec__1(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandWhereDeclsOpt(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { @@ -20586,18 +20729,6 @@ static lean_object* _init_l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expand _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(2); -x_2 = l_Lean_Elab_Term_expandWhereDecls___closed__12; -x_3 = lean_alloc_ctor(2, 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_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__15() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__9; x_2 = l_Lean_Elab_Term_expandFunBinders_loop___closed__1; x_3 = lean_name_mk_string(x_1, x_2); @@ -20680,7 +20811,7 @@ lean_inc(x_36); x_38 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_38, 0, x_36); lean_ctor_set(x_38, 1, x_37); -x_39 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_39 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_40 = lean_array_push(x_39, x_17); x_41 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_42 = lean_alloc_ctor(1, 3, 0); @@ -20691,7 +20822,7 @@ x_43 = l_Lean_Elab_Term_expandFunBinders_loop___closed__13; x_44 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_44, 0, x_36); lean_ctor_set(x_44, 1, x_43); -x_45 = l_Lean_Elab_Term_quoteAutoTactic___closed__46; +x_45 = l_Lean_Elab_Term_quoteAutoTactic___closed__49; x_46 = lean_array_push(x_45, x_42); x_47 = lean_array_push(x_46, x_44); x_48 = lean_array_push(x_47, x_32); @@ -20723,7 +20854,7 @@ lean_inc(x_55); x_58 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_58, 0, x_55); lean_ctor_set(x_58, 1, x_57); -x_59 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_59 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_60 = lean_array_push(x_59, x_17); x_61 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_62 = lean_alloc_ctor(1, 3, 0); @@ -20734,7 +20865,7 @@ x_63 = l_Lean_Elab_Term_expandFunBinders_loop___closed__13; x_64 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_64, 0, x_55); lean_ctor_set(x_64, 1, x_63); -x_65 = l_Lean_Elab_Term_quoteAutoTactic___closed__46; +x_65 = l_Lean_Elab_Term_quoteAutoTactic___closed__49; x_66 = lean_array_push(x_65, x_62); x_67 = lean_array_push(x_66, x_64); x_68 = lean_array_push(x_67, x_32); @@ -20775,7 +20906,7 @@ lean_inc(x_80); x_82 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_82, 0, x_80); lean_ctor_set(x_82, 1, x_81); -x_83 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_83 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_84 = lean_array_push(x_83, x_17); x_85 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_86 = lean_alloc_ctor(1, 3, 0); @@ -20789,11 +20920,11 @@ x_90 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_90, 0, x_27); lean_ctor_set(x_90, 1, x_89); lean_ctor_set(x_90, 2, x_88); -x_91 = l_Lean_Elab_Term_expandWhereDecls___closed__13; +x_91 = l_Lean_Elab_Term_expandWhereDecls___closed__14; x_92 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_92, 0, x_80); lean_ctor_set(x_92, 1, x_91); -x_93 = l_Lean_Elab_Term_quoteAutoTactic___closed__46; +x_93 = l_Lean_Elab_Term_quoteAutoTactic___closed__49; x_94 = lean_array_push(x_93, x_90); x_95 = lean_array_push(x_94, x_92); x_96 = lean_array_push(x_95, x_76); @@ -20823,7 +20954,7 @@ lean_inc(x_101); x_104 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_104, 0, x_101); lean_ctor_set(x_104, 1, x_103); -x_105 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_105 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_106 = lean_array_push(x_105, x_17); x_107 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_108 = lean_alloc_ctor(1, 3, 0); @@ -20837,11 +20968,11 @@ x_112 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_112, 0, x_27); lean_ctor_set(x_112, 1, x_111); lean_ctor_set(x_112, 2, x_110); -x_113 = l_Lean_Elab_Term_expandWhereDecls___closed__13; +x_113 = l_Lean_Elab_Term_expandWhereDecls___closed__14; x_114 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_114, 0, x_101); lean_ctor_set(x_114, 1, x_113); -x_115 = l_Lean_Elab_Term_quoteAutoTactic___closed__46; +x_115 = l_Lean_Elab_Term_quoteAutoTactic___closed__49; x_116 = lean_array_push(x_115, x_112); x_117 = lean_array_push(x_116, x_114); x_118 = lean_array_push(x_117, x_76); @@ -20941,7 +21072,7 @@ lean_inc(x_149); x_153 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_153, 0, x_149); lean_ctor_set(x_153, 1, x_152); -x_154 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_154 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_155 = lean_array_push(x_154, x_131); x_156 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_157 = lean_alloc_ctor(1, 3, 0); @@ -20952,7 +21083,7 @@ x_158 = l_Lean_Elab_Term_expandFunBinders_loop___closed__13; x_159 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_159, 0, x_149); lean_ctor_set(x_159, 1, x_158); -x_160 = l_Lean_Elab_Term_quoteAutoTactic___closed__46; +x_160 = l_Lean_Elab_Term_quoteAutoTactic___closed__49; x_161 = lean_array_push(x_160, x_157); x_162 = lean_array_push(x_161, x_159); x_163 = lean_array_push(x_162, x_146); @@ -21003,7 +21134,7 @@ lean_inc(x_174); x_178 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_178, 0, x_174); lean_ctor_set(x_178, 1, x_177); -x_179 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_179 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_180 = lean_array_push(x_179, x_131); x_181 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_182 = lean_alloc_ctor(1, 3, 0); @@ -21017,11 +21148,11 @@ x_186 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_186, 0, x_141); lean_ctor_set(x_186, 1, x_185); lean_ctor_set(x_186, 2, x_184); -x_187 = l_Lean_Elab_Term_expandWhereDecls___closed__13; +x_187 = l_Lean_Elab_Term_expandWhereDecls___closed__14; x_188 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_188, 0, x_174); lean_ctor_set(x_188, 1, x_187); -x_189 = l_Lean_Elab_Term_quoteAutoTactic___closed__46; +x_189 = l_Lean_Elab_Term_quoteAutoTactic___closed__49; x_190 = lean_array_push(x_189, x_186); x_191 = lean_array_push(x_190, x_188); x_192 = lean_array_push(x_191, x_171); @@ -21149,7 +21280,7 @@ lean_inc(x_228); x_232 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_232, 0, x_228); lean_ctor_set(x_232, 1, x_231); -x_233 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_233 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_234 = lean_array_push(x_233, x_210); x_235 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_236 = lean_alloc_ctor(1, 3, 0); @@ -21160,7 +21291,7 @@ x_237 = l_Lean_Elab_Term_expandFunBinders_loop___closed__13; x_238 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_238, 0, x_228); lean_ctor_set(x_238, 1, x_237); -x_239 = l_Lean_Elab_Term_quoteAutoTactic___closed__46; +x_239 = l_Lean_Elab_Term_quoteAutoTactic___closed__49; x_240 = lean_array_push(x_239, x_236); x_241 = lean_array_push(x_240, x_238); x_242 = lean_array_push(x_241, x_225); @@ -21211,7 +21342,7 @@ lean_inc(x_253); x_257 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_257, 0, x_253); lean_ctor_set(x_257, 1, x_256); -x_258 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_258 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_259 = lean_array_push(x_258, x_210); x_260 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_261 = lean_alloc_ctor(1, 3, 0); @@ -21225,11 +21356,11 @@ x_265 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_265, 0, x_220); lean_ctor_set(x_265, 1, x_264); lean_ctor_set(x_265, 2, x_263); -x_266 = l_Lean_Elab_Term_expandWhereDecls___closed__13; +x_266 = l_Lean_Elab_Term_expandWhereDecls___closed__14; x_267 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_267, 0, x_253); lean_ctor_set(x_267, 1, x_266); -x_268 = l_Lean_Elab_Term_quoteAutoTactic___closed__46; +x_268 = l_Lean_Elab_Term_quoteAutoTactic___closed__49; x_269 = lean_array_push(x_268, x_265); x_270 = lean_array_push(x_269, x_267); x_271 = lean_array_push(x_270, x_250); @@ -21277,7 +21408,7 @@ x_284 = 0; x_285 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__6; x_286 = l_Lean_Elab_Term_expandFunBinders_loop___closed__3; x_287 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___spec__1(x_285, x_286, x_283, x_284, x_4); -x_288 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__14; +x_288 = l_Lean_Elab_Term_expandWhereDecls___closed__13; x_289 = l_Lean_mkSepArray(x_287, x_288); lean_dec(x_287); x_290 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; @@ -21327,7 +21458,7 @@ x_312 = 0; x_313 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__6; x_314 = l_Lean_Elab_Term_expandFunBinders_loop___closed__3; x_315 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___spec__1(x_313, x_314, x_311, x_312, x_4); -x_316 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__14; +x_316 = l_Lean_Elab_Term_expandWhereDecls___closed__13; x_317 = l_Lean_mkSepArray(x_315, x_316); lean_dec(x_315); x_318 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; @@ -21381,7 +21512,7 @@ x_342 = 0; x_343 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__4; x_344 = l_Lean_Elab_Term_expandFunBinders_loop___closed__3; x_345 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___spec__2(x_343, x_344, x_341, x_342, x_4); -x_346 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__14; +x_346 = l_Lean_Elab_Term_expandWhereDecls___closed__13; x_347 = l_Lean_mkSepArray(x_345, x_346); lean_dec(x_345); x_348 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; @@ -21403,7 +21534,7 @@ x_358 = lean_array_push(x_357, x_344); x_359 = lean_array_push(x_358, x_352); x_360 = lean_array_push(x_359, x_354); x_361 = lean_array_push(x_360, x_1); -x_362 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__15; +x_362 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__14; x_363 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_363, 0, x_350); lean_ctor_set(x_363, 1, x_362); @@ -21431,7 +21562,7 @@ x_370 = 0; x_371 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__4; x_372 = l_Lean_Elab_Term_expandFunBinders_loop___closed__3; x_373 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___spec__2(x_371, x_372, x_369, x_370, x_4); -x_374 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__14; +x_374 = l_Lean_Elab_Term_expandWhereDecls___closed__13; x_375 = l_Lean_mkSepArray(x_373, x_374); lean_dec(x_373); x_376 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; @@ -21453,7 +21584,7 @@ x_386 = lean_array_push(x_385, x_372); x_387 = lean_array_push(x_386, x_380); x_388 = lean_array_push(x_387, x_382); x_389 = lean_array_push(x_388, x_1); -x_390 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__15; +x_390 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__14; x_391 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_391, 0, x_378); lean_ctor_set(x_391, 1, x_390); @@ -21742,7 +21873,7 @@ lean_inc(x_36); x_39 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_39, 0, x_36); lean_ctor_set(x_39, 1, x_38); -x_40 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_40 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_41 = lean_array_push(x_40, x_17); x_42 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_43 = lean_alloc_ctor(1, 3, 0); @@ -21753,7 +21884,7 @@ x_44 = l_Lean_Elab_Term_expandFunBinders_loop___closed__13; x_45 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_45, 0, x_36); lean_ctor_set(x_45, 1, x_44); -x_46 = l_Lean_Elab_Term_quoteAutoTactic___closed__46; +x_46 = l_Lean_Elab_Term_quoteAutoTactic___closed__49; x_47 = lean_array_push(x_46, x_43); x_48 = lean_array_push(x_47, x_45); x_49 = lean_array_push(x_48, x_32); @@ -21795,7 +21926,7 @@ lean_inc(x_59); x_63 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_63, 0, x_59); lean_ctor_set(x_63, 1, x_62); -x_64 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_64 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_65 = lean_array_push(x_64, x_17); x_66 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_67 = lean_alloc_ctor(1, 3, 0); @@ -21806,7 +21937,7 @@ x_68 = l_Lean_Elab_Term_expandFunBinders_loop___closed__13; x_69 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_69, 0, x_59); lean_ctor_set(x_69, 1, x_68); -x_70 = l_Lean_Elab_Term_quoteAutoTactic___closed__46; +x_70 = l_Lean_Elab_Term_quoteAutoTactic___closed__49; x_71 = lean_array_push(x_70, x_67); x_72 = lean_array_push(x_71, x_69); x_73 = lean_array_push(x_72, x_32); @@ -21942,7 +22073,7 @@ lean_inc(x_113); x_118 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_118, 0, x_113); lean_ctor_set(x_118, 1, x_117); -x_119 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_119 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_120 = lean_array_push(x_119, x_95); x_121 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_122 = lean_alloc_ctor(1, 3, 0); @@ -21953,7 +22084,7 @@ x_123 = l_Lean_Elab_Term_expandFunBinders_loop___closed__13; x_124 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_124, 0, x_113); lean_ctor_set(x_124, 1, x_123); -x_125 = l_Lean_Elab_Term_quoteAutoTactic___closed__46; +x_125 = l_Lean_Elab_Term_quoteAutoTactic___closed__49; x_126 = lean_array_push(x_125, x_122); x_127 = lean_array_push(x_126, x_124); x_128 = lean_array_push(x_127, x_110); @@ -22119,7 +22250,7 @@ lean_inc(x_173); x_178 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_178, 0, x_173); lean_ctor_set(x_178, 1, x_177); -x_179 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_179 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_180 = lean_array_push(x_179, x_155); x_181 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_182 = lean_alloc_ctor(1, 3, 0); @@ -22130,7 +22261,7 @@ x_183 = l_Lean_Elab_Term_expandFunBinders_loop___closed__13; x_184 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_184, 0, x_173); lean_ctor_set(x_184, 1, x_183); -x_185 = l_Lean_Elab_Term_quoteAutoTactic___closed__46; +x_185 = l_Lean_Elab_Term_quoteAutoTactic___closed__49; x_186 = lean_array_push(x_185, x_182); x_187 = lean_array_push(x_186, x_184); x_188 = lean_array_push(x_187, x_170); @@ -22212,7 +22343,7 @@ x_211 = 0; x_212 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__6; x_213 = l_Lean_Elab_Term_expandFunBinders_loop___closed__3; x_214 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAltsWhereDecls_loop___spec__1(x_212, x_213, x_210, x_211, x_4); -x_215 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__14; +x_215 = l_Lean_Elab_Term_expandWhereDecls___closed__13; x_216 = l_Lean_mkSepArray(x_214, x_215); lean_dec(x_214); x_217 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; @@ -22274,7 +22405,7 @@ x_241 = 0; x_242 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__6; x_243 = l_Lean_Elab_Term_expandFunBinders_loop___closed__3; x_244 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAltsWhereDecls_loop___spec__1(x_242, x_243, x_240, x_241, x_4); -x_245 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__14; +x_245 = l_Lean_Elab_Term_expandWhereDecls___closed__13; x_246 = l_Lean_mkSepArray(x_244, x_245); lean_dec(x_244); x_247 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; @@ -22511,7 +22642,7 @@ x_47 = l_Lean_Elab_Term_expandFunBinders_loop___closed__13; x_48 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_48, 0, x_39); lean_ctor_set(x_48, 1, x_47); -x_49 = l_Lean_Elab_Term_quoteAutoTactic___closed__46; +x_49 = l_Lean_Elab_Term_quoteAutoTactic___closed__49; x_50 = lean_array_push(x_49, x_46); x_51 = lean_array_push(x_50, x_48); x_52 = lean_array_push(x_51, x_36); @@ -22554,7 +22685,7 @@ x_67 = l_Lean_Elab_Term_expandFunBinders_loop___closed__13; x_68 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_68, 0, x_58); lean_ctor_set(x_68, 1, x_67); -x_69 = l_Lean_Elab_Term_quoteAutoTactic___closed__46; +x_69 = l_Lean_Elab_Term_quoteAutoTactic___closed__49; x_70 = lean_array_push(x_69, x_66); x_71 = lean_array_push(x_70, x_68); x_72 = lean_array_push(x_71, x_36); @@ -22653,7 +22784,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandFun_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(542u); +x_1 = lean_unsigned_to_nat(543u); x_2 = lean_unsigned_to_nat(45u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22665,7 +22796,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandFun_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(550u); +x_1 = lean_unsigned_to_nat(551u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22693,7 +22824,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandFun_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(542u); +x_1 = lean_unsigned_to_nat(543u); x_2 = lean_unsigned_to_nat(49u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22705,7 +22836,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandFun_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(542u); +x_1 = lean_unsigned_to_nat(543u); x_2 = lean_unsigned_to_nat(58u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -24563,7 +24694,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabFun_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(564u); +x_1 = lean_unsigned_to_nat(565u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -24575,7 +24706,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabFun_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(577u); +x_1 = lean_unsigned_to_nat(578u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -24603,7 +24734,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabFun_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(564u); +x_1 = lean_unsigned_to_nat(565u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -24615,7 +24746,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabFun_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(564u); +x_1 = lean_unsigned_to_nat(565u); x_2 = lean_unsigned_to_nat(44u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -25963,7 +26094,7 @@ lean_inc(x_20); x_21 = lean_ctor_get(x_19, 1); lean_inc(x_21); lean_dec(x_19); -x_22 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_22 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_23 = lean_array_push(x_22, x_4); x_24 = 0; x_25 = 1; @@ -26075,7 +26206,7 @@ lean_inc(x_21); x_22 = lean_ctor_get(x_20, 1); lean_inc(x_22); lean_dec(x_20); -x_23 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_23 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_24 = lean_array_push(x_23, x_5); x_25 = 0; x_26 = 1; @@ -26312,7 +26443,7 @@ static lean_object* _init_l_Lean_Elab_Term_elabLetDeclAux___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1437____closed__3; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1498____closed__3; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } @@ -26894,7 +27025,7 @@ x_53 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_53, 0, x_52); lean_ctor_set(x_53, 1, x_46); lean_ctor_set(x_53, 2, x_51); -x_54 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_54 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_55 = lean_array_push(x_54, x_53); x_56 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_57 = lean_alloc_ctor(1, 3, 0); @@ -26912,7 +27043,7 @@ lean_inc(x_32); x_62 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_62, 0, x_32); lean_ctor_set(x_62, 1, x_61); -x_63 = l_Lean_Elab_Term_quoteAutoTactic___closed__46; +x_63 = l_Lean_Elab_Term_quoteAutoTactic___closed__49; x_64 = lean_array_push(x_63, x_44); x_65 = lean_array_push(x_64, x_60); x_66 = lean_array_push(x_65, x_62); @@ -27026,7 +27157,7 @@ x_120 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_120, 0, x_119); lean_ctor_set(x_120, 1, x_116); lean_ctor_set(x_120, 2, x_118); -x_121 = l_Lean_Elab_Term_quoteAutoTactic___closed__51; +x_121 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; x_122 = lean_array_push(x_121, x_120); x_123 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__2___closed__7; x_124 = lean_alloc_ctor(1, 3, 0); @@ -27435,7 +27566,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetDecl_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(688u); +x_1 = lean_unsigned_to_nat(689u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27447,7 +27578,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetDecl_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(689u); +x_1 = lean_unsigned_to_nat(690u); x_2 = lean_unsigned_to_nat(129u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27475,7 +27606,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetDecl_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(688u); +x_1 = lean_unsigned_to_nat(689u); x_2 = lean_unsigned_to_nat(29u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27487,7 +27618,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetDecl_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(688u); +x_1 = lean_unsigned_to_nat(689u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27602,7 +27733,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetFunDecl_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(691u); +x_1 = lean_unsigned_to_nat(692u); x_2 = lean_unsigned_to_nat(29u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27614,7 +27745,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetFunDecl_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(692u); +x_1 = lean_unsigned_to_nat(693u); x_2 = lean_unsigned_to_nat(130u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27642,7 +27773,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetFunDecl_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(691u); +x_1 = lean_unsigned_to_nat(692u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27654,7 +27785,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetFunDecl_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(691u); +x_1 = lean_unsigned_to_nat(692u); x_2 = lean_unsigned_to_nat(47u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27770,7 +27901,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetDelayedDecl_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(694u); +x_1 = lean_unsigned_to_nat(695u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27782,7 +27913,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetDelayedDecl_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(695u); +x_1 = lean_unsigned_to_nat(696u); x_2 = lean_unsigned_to_nat(128u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27810,7 +27941,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetDelayedDecl_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(694u); +x_1 = lean_unsigned_to_nat(695u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27822,7 +27953,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetDelayedDecl_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(694u); +x_1 = lean_unsigned_to_nat(695u); x_2 = lean_unsigned_to_nat(55u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27938,7 +28069,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(697u); +x_1 = lean_unsigned_to_nat(698u); x_2 = lean_unsigned_to_nat(29u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27950,7 +28081,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(698u); +x_1 = lean_unsigned_to_nat(699u); x_2 = lean_unsigned_to_nat(128u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27978,7 +28109,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(697u); +x_1 = lean_unsigned_to_nat(698u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27990,7 +28121,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(697u); +x_1 = lean_unsigned_to_nat(698u); x_2 = lean_unsigned_to_nat(47u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28036,7 +28167,7 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_7233_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_7537_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; @@ -28240,6 +28371,12 @@ l_Lean_Elab_Term_quoteAutoTactic___closed__59 = _init_l_Lean_Elab_Term_quoteAuto lean_mark_persistent(l_Lean_Elab_Term_quoteAutoTactic___closed__59); l_Lean_Elab_Term_quoteAutoTactic___closed__60 = _init_l_Lean_Elab_Term_quoteAutoTactic___closed__60(); lean_mark_persistent(l_Lean_Elab_Term_quoteAutoTactic___closed__60); +l_Lean_Elab_Term_quoteAutoTactic___closed__61 = _init_l_Lean_Elab_Term_quoteAutoTactic___closed__61(); +lean_mark_persistent(l_Lean_Elab_Term_quoteAutoTactic___closed__61); +l_Lean_Elab_Term_quoteAutoTactic___closed__62 = _init_l_Lean_Elab_Term_quoteAutoTactic___closed__62(); +lean_mark_persistent(l_Lean_Elab_Term_quoteAutoTactic___closed__62); +l_Lean_Elab_Term_quoteAutoTactic___closed__63 = _init_l_Lean_Elab_Term_quoteAutoTactic___closed__63(); +lean_mark_persistent(l_Lean_Elab_Term_quoteAutoTactic___closed__63); l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__1 = _init_l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__1); l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__2 = _init_l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__2(); @@ -28290,6 +28427,10 @@ l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBin lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__1); l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__2 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__2(); lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__2); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__3 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__3(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__3); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__4 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__4(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__4); l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_matchBinder___spec__1___rarg___closed__1 = _init_l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_matchBinder___spec__1___rarg___closed__1(); lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_matchBinder___spec__1___rarg___closed__1); l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_matchBinder___spec__1___rarg___closed__2 = _init_l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_matchBinder___spec__1___rarg___closed__2(); @@ -28328,17 +28469,17 @@ l___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___closed_ lean_mark_persistent(l___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___closed__3); l___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___closed__4 = _init_l___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___closed__4(); lean_mark_persistent(l___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___closed__4); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1437____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1437____closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1437____closed__1); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1437____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1437____closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1437____closed__2); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1437____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1437____closed__3(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1437____closed__3); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1437____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1437____closed__4(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1437____closed__4); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1437____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1437____closed__5(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1437____closed__5); -if (builtin) {res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1437_(lean_io_mk_world()); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1498____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1498____closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1498____closed__1); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1498____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1498____closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1498____closed__2); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1498____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1498____closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1498____closed__3); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1498____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1498____closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1498____closed__4); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1498____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1498____closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1498____closed__5); +if (builtin) {res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1498_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Elab_Term_checkBinderAnnotations = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Elab_Term_checkBinderAnnotations); @@ -28471,10 +28612,6 @@ lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_elabDepArrow_declRange___clos res = l___regBuiltin_Lean_Elab_Term_elabDepArrow_declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getFunBinderIds_x3f___spec__3___closed__1 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getFunBinderIds_x3f___spec__3___closed__1(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getFunBinderIds_x3f___spec__3___closed__1); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getFunBinderIds_x3f___spec__3___closed__2 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getFunBinderIds_x3f___spec__3___closed__2(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getFunBinderIds_x3f___spec__3___closed__2); l_Lean_Elab_Term_expandFunBinders_loop___closed__1 = _init_l_Lean_Elab_Term_expandFunBinders_loop___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_expandFunBinders_loop___closed__1); l_Lean_Elab_Term_expandFunBinders_loop___closed__2 = _init_l_Lean_Elab_Term_expandFunBinders_loop___closed__2(); @@ -28543,6 +28680,8 @@ l_Lean_Elab_Term_expandWhereDecls___closed__12 = _init_l_Lean_Elab_Term_expandWh lean_mark_persistent(l_Lean_Elab_Term_expandWhereDecls___closed__12); l_Lean_Elab_Term_expandWhereDecls___closed__13 = _init_l_Lean_Elab_Term_expandWhereDecls___closed__13(); lean_mark_persistent(l_Lean_Elab_Term_expandWhereDecls___closed__13); +l_Lean_Elab_Term_expandWhereDecls___closed__14 = _init_l_Lean_Elab_Term_expandWhereDecls___closed__14(); +lean_mark_persistent(l_Lean_Elab_Term_expandWhereDecls___closed__14); l_Lean_mkAuxFunDiscr___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___spec__3___closed__1 = _init_l_Lean_mkAuxFunDiscr___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___spec__3___closed__1(); lean_mark_persistent(l_Lean_mkAuxFunDiscr___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___spec__3___closed__1); l_Lean_mkAuxFunDiscr___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___spec__3___closed__2 = _init_l_Lean_mkAuxFunDiscr___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___spec__3___closed__2(); @@ -28579,8 +28718,6 @@ l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___cl lean_mark_persistent(l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__13); l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__14 = _init_l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__14(); lean_mark_persistent(l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__14); -l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__15 = _init_l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__15(); -lean_mark_persistent(l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__15); l___regBuiltin_Lean_Elab_Term_expandFun___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_expandFun___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_expandFun___closed__1); l___regBuiltin_Lean_Elab_Term_expandFun___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_expandFun___closed__2(); @@ -28820,7 +28957,7 @@ lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl_declRange___cl res = l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl_declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_7233_(lean_io_mk_world()); +res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_7537_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/BindersUtil.c b/stage0/stdlib/Lean/Elab/BindersUtil.c index 3c37970785d..69f228b07b2 100644 --- a/stage0/stdlib/Lean/Elab/BindersUtil.c +++ b/stage0/stdlib/Lean/Elab/BindersUtil.c @@ -14,45 +14,60 @@ extern "C" { #endif size_t lean_usize_add(size_t, size_t); +static lean_object* l_Lean_Elab_Term_expandMatchAlt___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandMatchAlts_x3f(lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); -extern lean_object* l_Lean_nullKind; +static lean_object* l_Lean_Elab_Term_expandMatchAlt___closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_Term_getMatchAltsNumPatterns___boxed(lean_object*); lean_object* lean_name_mk_string(lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); lean_object* lean_array_uget(lean_object*, size_t); +static lean_object* l_Lean_Elab_Term_expandMatchAlt___closed__9; lean_object* l_Array_append___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_expandMatchAlt___closed__4; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); -static lean_object* l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__7; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_expandMatchAlts_x3f___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_expandMatchAlt___closed__8; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__3; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_expandMatchAlts_x3f___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); static lean_object* l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__5; +static uint8_t l_Lean_Elab_Term_shouldExpandMatchAlt___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); +static lean_object* l_Lean_Elab_Term_shouldExpandMatchAlt___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_shouldExpandMatchAlt___boxed(lean_object*); +lean_object* lean_nat_add(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__3; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__1___closed__1; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__1(lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Elab_Term_expandMatchAlt___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_fget(lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__5; +uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); +lean_object* lean_nat_sub(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_expandMatchAlt___closed__7; +static lean_object* l_Lean_Elab_Term_expandMatchAlt___closed__5; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandMatchAlt___lambda__1___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Elab_Term_expandMatchAlt___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_expandMatchAlt___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandOptType(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_expandMatchAlts_x3f___spec__2(lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_expandMatchAlts_x3f___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_expandMatchAlt___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandOptType___boxed(lean_object*, lean_object*); -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Elab_Term_expandMatchAlt___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_getMatchAltsNumPatterns(lean_object*); -static lean_object* l_Lean_Elab_Term_expandMatchAlts_x3f___closed__3; LEAN_EXPORT uint8_t l_Lean_Elab_Term_shouldExpandMatchAlt(lean_object*); lean_object* l_Lean_Syntax_getSepArgs(lean_object*); lean_object* l_Lean_mkHole(lean_object*); -static lean_object* l_Lean_Elab_Term_expandMatchAlts_x3f___closed__5; -lean_object* l_Lean_Syntax_setArg(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_expandMatchAlts_x3f___closed__7; uint8_t lean_nat_dec_le(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_expandMatchAlt___closed__6; lean_object* l_Lean_Syntax_getArgs(lean_object*); -static lean_object* l_Lean_Elab_Term_expandMatchAlts_x3f___closed__4; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandMatchAlt___lambda__1(lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__4; uint8_t l_Lean_Syntax_isNone(lean_object*); static lean_object* l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_expandMatchAlts_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*); @@ -60,16 +75,19 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___boxe static lean_object* l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__1; uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_expandMatchAlts_x3f___spec__1(lean_object*, size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandMatchAlt(lean_object*); -static lean_object* l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__8; -static lean_object* l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__9; +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Elab_Term_expandMatchAlt___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandMatchAlt(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__6; -static lean_object* l_Lean_Elab_Term_expandMatchAlts_x3f___closed__6; -static lean_object* l_Lean_Elab_Term_expandMatchAlts_x3f___closed__2; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__6; +static lean_object* l_Lean_Elab_Term_expandMatchAlt___closed__11; +lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Term_expandMatchAlts_x3f___closed__1; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__2; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__1; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandOptType(lean_object* x_1, lean_object* x_2) { _start: @@ -133,7 +151,113 @@ lean_dec(x_1); return x_2; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__1___closed__1() { +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Elab_Term_expandMatchAlt___spec__2(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; uint8_t x_7; +x_6 = lean_array_get_size(x_1); +x_7 = lean_nat_dec_lt(x_4, x_6); +lean_dec(x_6); +if (x_7 == 0) +{ +lean_object* x_8; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_5); +return x_8; +} +else +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_unsigned_to_nat(0u); +x_10 = lean_nat_dec_eq(x_3, x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_sub(x_3, x_11); +lean_dec(x_3); +x_13 = lean_array_fget(x_1, x_4); +lean_inc(x_2); +x_14 = lean_apply_1(x_2, x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; +lean_dec(x_12); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_15 = lean_box(0); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_14, 0); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_nat_add(x_4, x_11); +lean_dec(x_4); +x_18 = lean_array_push(x_5, x_16); +x_3 = x_12; +x_4 = x_17; +x_5 = x_18; +goto _start; +} +} +else +{ +lean_object* x_20; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_20 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_20, 0, x_5); +return x_20; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Elab_Term_expandMatchAlt___spec__1(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; +x_3 = lean_array_get_size(x_1); +x_4 = lean_mk_empty_array_with_capacity(x_3); +x_5 = lean_unsigned_to_nat(0u); +x_6 = l_Array_sequenceMap_loop___at_Lean_Elab_Term_expandMatchAlt___spec__2(x_1, x_2, x_3, x_5, x_4); +return x_6; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("|", 1); +return x_1; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("null", 4); +return x_1; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__2; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__4() { _start: { lean_object* x_1; lean_object* x_2; @@ -142,108 +266,496 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__1(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__5() { _start: { -uint8_t x_5; -x_5 = lean_usize_dec_lt(x_3, x_2); -if (x_5 == 0) +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("=>", 2); +return x_1; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(4u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3(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, lean_object* x_8, lean_object* x_9) { +_start: +{ +uint8_t x_10; +x_10 = lean_usize_dec_lt(x_6, x_5); +if (x_10 == 0) { +lean_object* x_11; +lean_dec(x_8); +lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -return x_4; +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_7); +lean_ctor_set(x_11, 1, x_9); +return x_11; } 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; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; size_t x_16; size_t x_17; lean_object* x_18; -x_6 = lean_array_uget(x_4, x_3); -x_7 = lean_unsigned_to_nat(0u); -x_8 = lean_array_uset(x_4, x_3, x_7); -x_9 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__1___closed__1; -x_10 = lean_array_push(x_9, x_6); -x_11 = lean_box(2); -x_12 = l_Lean_nullKind; -x_13 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_13, 0, x_11); -lean_ctor_set(x_13, 1, x_12); -lean_ctor_set(x_13, 2, x_10); -x_14 = lean_unsigned_to_nat(1u); +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; 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; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; size_t x_35; size_t x_36; lean_object* x_37; +x_12 = lean_array_uget(x_7, x_6); +x_13 = lean_unsigned_to_nat(0u); +x_14 = lean_array_uset(x_7, x_6, x_13); +lean_inc(x_8); +x_15 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_8, x_9); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__1; +lean_inc(x_16); +x_19 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_19, 0, x_16); +lean_ctor_set(x_19, 1, x_18); +lean_inc(x_2); +x_20 = l_Array_append___rarg(x_2, x_12); +x_21 = lean_box(2); +x_22 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__3; +x_23 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +lean_ctor_set(x_23, 2, x_20); +x_24 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__4; +x_25 = lean_array_push(x_24, x_23); +x_26 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_26, 0, x_21); +lean_ctor_set(x_26, 1, x_22); +lean_ctor_set(x_26, 2, x_25); +x_27 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__5; +x_28 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_28, 0, x_16); +lean_ctor_set(x_28, 1, x_27); +x_29 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__6; +x_30 = lean_array_push(x_29, x_19); +x_31 = lean_array_push(x_30, x_26); +x_32 = lean_array_push(x_31, x_28); +lean_inc(x_3); +x_33 = lean_array_push(x_32, x_3); lean_inc(x_1); -x_15 = l_Lean_Syntax_setArg(x_1, x_14, x_13); -x_16 = 1; -x_17 = lean_usize_add(x_3, x_16); -x_18 = lean_array_uset(x_8, x_3, x_15); -x_3 = x_17; -x_4 = x_18; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_21); +lean_ctor_set(x_34, 1, x_1); +lean_ctor_set(x_34, 2, x_33); +x_35 = 1; +x_36 = lean_usize_add(x_6, x_35); +x_37 = lean_array_uset(x_14, x_6, x_34); +x_6 = x_36; +x_7 = x_37; +x_9 = x_17; goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandMatchAlt(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandMatchAlt___lambda__1(lean_object* x_1) { _start: { -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; -x_2 = lean_unsigned_to_nat(1u); -x_3 = l_Lean_Syntax_getArg(x_1, x_2); -x_4 = l_Lean_Syntax_getArgs(x_3); -x_5 = lean_array_get_size(x_4); -lean_dec(x_4); -x_6 = lean_nat_dec_le(x_5, x_2); -lean_dec(x_5); -if (x_6 == 0) +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_Syntax_getArgs(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_expandMatchAlt___closed__1() { +_start: { -lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; -x_7 = l_Lean_Syntax_getSepArgs(x_3); -lean_dec(x_3); -x_8 = lean_array_get_size(x_7); -x_9 = lean_usize_of_nat(x_8); -lean_dec(x_8); -x_10 = 0; -x_11 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__1(x_1, x_9, x_10, x_7); -return x_11; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Lean", 4); +return x_1; } -else +} +static lean_object* _init_l_Lean_Elab_Term_expandMatchAlt___closed__2() { +_start: { -lean_object* x_12; lean_object* x_13; -lean_dec(x_3); -x_12 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__1___closed__1; -x_13 = lean_array_push(x_12, x_1); -return x_13; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_expandMatchAlt___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_expandMatchAlt___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Parser", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_expandMatchAlt___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Term_expandMatchAlt___closed__2; +x_2 = l_Lean_Elab_Term_expandMatchAlt___closed__3; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_expandMatchAlt___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Term", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_expandMatchAlt___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Term_expandMatchAlt___closed__4; +x_2 = l_Lean_Elab_Term_expandMatchAlt___closed__5; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_expandMatchAlt___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("matchAlt", 8); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_expandMatchAlt___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Term_expandMatchAlt___closed__6; +x_2 = l_Lean_Elab_Term_expandMatchAlt___closed__7; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_expandMatchAlt___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_Term_expandMatchAlt___closed__10() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l_Lean_Elab_Term_expandMatchAlt___closed__9; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; } } +static lean_object* _init_l_Lean_Elab_Term_expandMatchAlt___closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_expandMatchAlt___lambda__1___boxed), 1, 0); +return x_1; +} } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandMatchAlt(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -size_t x_5; size_t x_6; lean_object* x_7; -x_5 = lean_unbox_usize(x_2); +lean_object* x_4; uint8_t x_5; +x_4 = l_Lean_Elab_Term_expandMatchAlt___closed__8; +lean_inc(x_1); +x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_dec(x_2); -x_6 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_7 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__1(x_1, x_5, x_6, x_4); -return x_7; +x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__4; +x_7 = lean_array_push(x_6, x_1); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_3); +return x_8; } +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; +x_9 = lean_unsigned_to_nat(1u); +x_10 = l_Lean_Syntax_getArg(x_1, x_9); +x_11 = l_Lean_Syntax_getArgs(x_10); +lean_dec(x_10); +x_12 = lean_array_get_size(x_11); +x_13 = lean_unsigned_to_nat(0u); +x_14 = lean_nat_dec_lt(x_13, x_12); +if (x_14 == 0) +{ +lean_object* x_35; +lean_dec(x_12); +lean_dec(x_11); +x_35 = l_Lean_Elab_Term_expandMatchAlt___closed__9; +x_15 = x_35; +goto block_34; } -LEAN_EXPORT uint8_t l_Lean_Elab_Term_shouldExpandMatchAlt(lean_object* x_1) { +else +{ +uint8_t x_36; +x_36 = lean_nat_dec_le(x_12, x_12); +if (x_36 == 0) +{ +lean_object* x_37; +lean_dec(x_12); +lean_dec(x_11); +x_37 = l_Lean_Elab_Term_expandMatchAlt___closed__9; +x_15 = x_37; +goto block_34; +} +else +{ +size_t x_38; size_t x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_38 = 0; +x_39 = lean_usize_of_nat(x_12); +lean_dec(x_12); +x_40 = l_Lean_Elab_Term_expandMatchAlt___closed__10; +x_41 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(x_11, x_38, x_39, x_40); +lean_dec(x_11); +x_42 = lean_ctor_get(x_41, 1); +lean_inc(x_42); +lean_dec(x_41); +x_15 = x_42; +goto block_34; +} +} +block_34: +{ +lean_object* x_16; lean_object* x_17; +x_16 = l_Lean_Elab_Term_expandMatchAlt___closed__11; +x_17 = l_Array_sequenceMap___at_Lean_Elab_Term_expandMatchAlt___spec__1(x_15, x_16); +lean_dec(x_15); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +lean_dec(x_2); +x_18 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__4; +x_19 = lean_array_push(x_18, x_1); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_3); +return x_20; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_21 = lean_ctor_get(x_17, 0); +lean_inc(x_21); +lean_dec(x_17); +x_22 = lean_unsigned_to_nat(3u); +x_23 = l_Lean_Syntax_getArg(x_1, x_22); +x_24 = lean_box(0); +x_25 = lean_array_get_size(x_21); +x_26 = lean_nat_dec_le(x_25, x_9); +if (x_26 == 0) +{ +size_t x_27; size_t x_28; lean_object* x_29; lean_object* x_30; +lean_dec(x_1); +x_27 = lean_usize_of_nat(x_25); +lean_dec(x_25); +x_28 = 0; +x_29 = l_Lean_Elab_Term_expandMatchAlt___closed__9; +x_30 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3(x_4, x_29, x_23, x_24, x_27, x_28, x_21, x_2, x_3); +return x_30; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +lean_dec(x_25); +lean_dec(x_23); +lean_dec(x_21); +lean_dec(x_2); +x_31 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__4; +x_32 = lean_array_push(x_31, x_1); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_3); +return x_33; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Elab_Term_expandMatchAlt___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; -x_2 = lean_unsigned_to_nat(1u); -x_3 = l_Lean_Syntax_getArg(x_1, x_2); -x_4 = l_Lean_Syntax_getArgs(x_3); +lean_object* x_6; +x_6 = l_Array_sequenceMap_loop___at_Lean_Elab_Term_expandMatchAlt___spec__2(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Elab_Term_expandMatchAlt___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Array_sequenceMap___at_Lean_Elab_Term_expandMatchAlt___spec__1(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___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_object* x_9) { +_start: +{ +size_t x_10; size_t x_11; lean_object* x_12; +x_10 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_11 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_12 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3(x_1, x_2, x_3, x_4, x_10, x_11, x_7, x_8, x_9); +lean_dec(x_4); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandMatchAlt___lambda__1___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Elab_Term_expandMatchAlt___lambda__1(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_Term_shouldExpandMatchAlt___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Term_expandMatchAlt___closed__9; +x_2 = l_Lean_Elab_Term_expandMatchAlt___closed__11; +x_3 = l_Array_sequenceMap___at_Lean_Elab_Term_expandMatchAlt___spec__1(x_1, x_2); +return x_3; +} +} +static uint8_t _init_l_Lean_Elab_Term_shouldExpandMatchAlt___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Elab_Term_shouldExpandMatchAlt___closed__1; +if (lean_obj_tag(x_1) == 0) +{ +uint8_t x_2; +x_2 = 0; +return x_2; +} +else +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = lean_array_get_size(x_3); lean_dec(x_3); -x_5 = lean_array_get_size(x_4); +x_5 = lean_unsigned_to_nat(1u); +x_6 = lean_nat_dec_lt(x_5, x_4); lean_dec(x_4); -x_6 = lean_nat_dec_lt(x_2, x_5); -lean_dec(x_5); return x_6; } } +} +LEAN_EXPORT uint8_t l_Lean_Elab_Term_shouldExpandMatchAlt(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; +x_2 = l_Lean_Elab_Term_expandMatchAlt___closed__8; +lean_inc(x_1); +x_3 = l_Lean_Syntax_isOfKind(x_1, x_2); +if (x_3 == 0) +{ +uint8_t x_4; +lean_dec(x_1); +x_4 = 0; +return x_4; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_5 = lean_unsigned_to_nat(1u); +x_6 = l_Lean_Syntax_getArg(x_1, x_5); +lean_dec(x_1); +x_7 = l_Lean_Syntax_getArgs(x_6); +lean_dec(x_6); +x_8 = lean_array_get_size(x_7); +x_9 = lean_unsigned_to_nat(0u); +x_10 = lean_nat_dec_lt(x_9, x_8); +if (x_10 == 0) +{ +uint8_t x_11; +lean_dec(x_8); +lean_dec(x_7); +x_11 = l_Lean_Elab_Term_shouldExpandMatchAlt___closed__2; +return x_11; +} +else +{ +uint8_t x_12; +x_12 = lean_nat_dec_le(x_8, x_8); +if (x_12 == 0) +{ +uint8_t x_13; +lean_dec(x_8); +lean_dec(x_7); +x_13 = l_Lean_Elab_Term_shouldExpandMatchAlt___closed__2; +return x_13; +} +else +{ +size_t x_14; size_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_14 = 0; +x_15 = lean_usize_of_nat(x_8); +lean_dec(x_8); +x_16 = l_Lean_Elab_Term_expandMatchAlt___closed__10; +x_17 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(x_7, x_14, x_15, x_16); +lean_dec(x_7); +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); +lean_dec(x_17); +x_19 = l_Lean_Elab_Term_expandMatchAlt___closed__11; +x_20 = l_Array_sequenceMap___at_Lean_Elab_Term_expandMatchAlt___spec__1(x_18, x_19); +lean_dec(x_18); +if (lean_obj_tag(x_20) == 0) +{ +uint8_t x_21; +x_21 = 0; +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_22 = lean_ctor_get(x_20, 0); +lean_inc(x_22); +lean_dec(x_20); +x_23 = lean_array_get_size(x_22); +lean_dec(x_22); +x_24 = lean_nat_dec_lt(x_5, x_23); +lean_dec(x_23); +return x_24; +} +} +} +} +} +} LEAN_EXPORT lean_object* l_Lean_Elab_Term_shouldExpandMatchAlt___boxed(lean_object* x_1) { _start: { uint8_t x_2; lean_object* x_3; x_2 = l_Lean_Elab_Term_shouldExpandMatchAlt(x_1); -lean_dec(x_1); x_3 = lean_box(x_2); return x_3; } @@ -258,7 +770,6 @@ if (x_4 == 0) lean_object* x_5; uint8_t x_6; x_5 = lean_array_uget(x_1, x_2); x_6 = l_Lean_Elab_Term_shouldExpandMatchAlt(x_5); -lean_dec(x_5); if (x_6 == 0) { size_t x_7; size_t x_8; @@ -282,26 +793,38 @@ return x_11; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_expandMatchAlts_x3f___spec__2(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_expandMatchAlts_x3f___spec__2(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -uint8_t x_5; -x_5 = lean_usize_dec_eq(x_2, x_3); -if (x_5 == 0) -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; -x_6 = lean_array_uget(x_1, x_2); -x_7 = l_Lean_Elab_Term_expandMatchAlt(x_6); -x_8 = l_Array_append___rarg(x_4, x_7); -x_9 = 1; -x_10 = lean_usize_add(x_2, x_9); -x_2 = x_10; -x_4 = x_8; +uint8_t x_7; +x_7 = lean_usize_dec_eq(x_2, x_3); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; size_t x_13; size_t x_14; +x_8 = lean_array_uget(x_1, x_2); +lean_inc(x_5); +x_9 = l_Lean_Elab_Term_expandMatchAlt(x_8, x_5, x_6); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = l_Array_append___rarg(x_4, x_10); +x_13 = 1; +x_14 = lean_usize_add(x_2, x_13); +x_2 = x_14; +x_4 = x_12; +x_6 = x_11; goto _start; } else { -return x_4; +lean_object* x_16; +lean_dec(x_5); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_4); +lean_ctor_set(x_16, 1, x_6); +return x_16; } } } @@ -316,39 +839,12 @@ return x_1; static lean_object* _init_l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(0u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__3() { -_start: -{ lean_object* x_1; x_1 = lean_mk_string_from_bytes("match", 5); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("null", 4); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__4; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__6() { +static lean_object* _init_l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__3() { _start: { lean_object* x_1; @@ -356,7 +852,7 @@ x_1 = lean_mk_string_from_bytes("with", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__7() { +static lean_object* _init_l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; @@ -365,22 +861,22 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__8() { +static lean_object* _init_l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__2; +x_1 = l_Lean_Elab_Term_expandMatchAlt___closed__9; x_2 = l_Array_append___rarg(x_1, x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__9() { +static lean_object* _init_l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__5; -x_3 = l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__8; +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__3; +x_3 = l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__5; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -491,161 +987,167 @@ return x_33; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_34 = l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__2; -x_35 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_expandMatchAlts_x3f___spec__2(x_20, x_29, x_30, x_34); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_34 = l_Lean_Elab_Term_expandMatchAlt___closed__9; +lean_inc(x_7); +x_35 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_expandMatchAlts_x3f___spec__2(x_20, x_29, x_30, x_34, x_7, x_8); lean_dec(x_20); -x_36 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_7, x_8); -x_37 = lean_ctor_get(x_36, 0); +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_35, 1); lean_inc(x_37); -x_38 = lean_ctor_get(x_36, 1); -lean_inc(x_38); -if (lean_is_exclusive(x_36)) { - lean_ctor_release(x_36, 0); - lean_ctor_release(x_36, 1); - x_39 = x_36; +lean_dec(x_35); +x_38 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_7, x_37); +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_38, 1); +lean_inc(x_40); +if (lean_is_exclusive(x_38)) { + lean_ctor_release(x_38, 0); + lean_ctor_release(x_38, 1); + x_41 = x_38; } else { - lean_dec_ref(x_36); - x_39 = lean_box(0); -} -x_40 = l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__3; -lean_inc(x_37); -x_41 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_41, 0, x_37); -lean_ctor_set(x_41, 1, x_40); -x_42 = l_Array_append___rarg(x_34, x_21); -x_43 = lean_box(2); -x_44 = l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__5; -x_45 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_45, 0, x_43); -lean_ctor_set(x_45, 1, x_44); -lean_ctor_set(x_45, 2, x_42); -x_46 = l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__6; -x_47 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_47, 0, x_37); + lean_dec_ref(x_38); + x_41 = lean_box(0); +} +x_42 = l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__2; +lean_inc(x_39); +x_43 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_43, 0, x_39); +lean_ctor_set(x_43, 1, x_42); +x_44 = l_Array_append___rarg(x_34, x_21); +x_45 = lean_box(2); +x_46 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__3; +x_47 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_47, 0, x_45); lean_ctor_set(x_47, 1, x_46); -x_48 = l_Array_append___rarg(x_34, x_35); -x_49 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_49, 0, x_43); -lean_ctor_set(x_49, 1, x_44); -lean_ctor_set(x_49, 2, x_48); -x_50 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__1___closed__1; -x_51 = lean_array_push(x_50, x_49); -x_52 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_52, 0, x_43); -lean_ctor_set(x_52, 1, x_14); -lean_ctor_set(x_52, 2, x_51); -x_53 = l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__7; -x_54 = lean_array_push(x_53, x_41); +lean_ctor_set(x_47, 2, x_44); +x_48 = l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__3; +x_49 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_49, 0, x_39); +lean_ctor_set(x_49, 1, x_48); +x_50 = l_Array_append___rarg(x_34, x_36); +x_51 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_51, 0, x_45); +lean_ctor_set(x_51, 1, x_46); +lean_ctor_set(x_51, 2, x_50); +x_52 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__4; +x_53 = lean_array_push(x_52, x_51); +x_54 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_54, 0, x_45); +lean_ctor_set(x_54, 1, x_14); +lean_ctor_set(x_54, 2, x_53); +x_55 = l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__4; +x_56 = lean_array_push(x_55, x_43); if (lean_obj_tag(x_4) == 0) { -x_55 = x_34; -goto block_89; +x_57 = x_34; +goto block_91; } else { -lean_object* x_90; lean_object* x_91; -x_90 = lean_ctor_get(x_4, 0); -lean_inc(x_90); +lean_object* x_92; lean_object* x_93; +x_92 = lean_ctor_get(x_4, 0); +lean_inc(x_92); lean_dec(x_4); -x_91 = lean_array_push(x_50, x_90); -x_55 = x_91; -goto block_89; -} -block_89: -{ -lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_56 = l_Array_append___rarg(x_34, x_55); -x_57 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_57, 0, x_43); -lean_ctor_set(x_57, 1, x_44); -lean_ctor_set(x_57, 2, x_56); -x_58 = lean_array_push(x_54, x_57); +x_93 = lean_array_push(x_52, x_92); +x_57 = x_93; +goto block_91; +} +block_91: +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = l_Array_append___rarg(x_34, x_57); +x_59 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_59, 0, x_45); +lean_ctor_set(x_59, 1, x_46); +lean_ctor_set(x_59, 2, x_58); +x_60 = lean_array_push(x_56, x_59); if (lean_obj_tag(x_6) == 0) { -lean_object* x_59; 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_59 = l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__9; -x_60 = lean_array_push(x_58, x_59); -x_61 = lean_array_push(x_60, x_45); -x_62 = lean_array_push(x_61, x_47); -x_63 = lean_array_push(x_62, x_52); -x_64 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_64, 0, x_43); -lean_ctor_set(x_64, 1, x_3); -lean_ctor_set(x_64, 2, x_63); -x_65 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_65, 0, x_64); -if (lean_is_scalar(x_39)) { - x_66 = lean_alloc_ctor(0, 2, 0); +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_61 = l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__6; +x_62 = lean_array_push(x_60, x_61); +x_63 = lean_array_push(x_62, x_47); +x_64 = lean_array_push(x_63, x_49); +x_65 = lean_array_push(x_64, x_54); +x_66 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_66, 0, x_45); +lean_ctor_set(x_66, 1, x_3); +lean_ctor_set(x_66, 2, x_65); +x_67 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_67, 0, x_66); +if (lean_is_scalar(x_41)) { + x_68 = lean_alloc_ctor(0, 2, 0); } else { - x_66 = x_39; + x_68 = x_41; } -lean_ctor_set(x_66, 0, x_65); -lean_ctor_set(x_66, 1, x_38); -return x_66; +lean_ctor_set(x_68, 0, x_67); +lean_ctor_set(x_68, 1, x_40); +return x_68; } else { -uint8_t x_67; -x_67 = !lean_is_exclusive(x_6); -if (x_67 == 0) -{ -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; -x_68 = lean_ctor_get(x_6, 0); -x_69 = lean_array_push(x_50, x_68); -x_70 = l_Array_append___rarg(x_34, x_69); -x_71 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_71, 0, x_43); -lean_ctor_set(x_71, 1, x_44); -lean_ctor_set(x_71, 2, x_70); -x_72 = lean_array_push(x_58, x_71); -x_73 = lean_array_push(x_72, x_45); -x_74 = lean_array_push(x_73, x_47); -x_75 = lean_array_push(x_74, x_52); -x_76 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_76, 0, x_43); -lean_ctor_set(x_76, 1, x_3); -lean_ctor_set(x_76, 2, x_75); -lean_ctor_set(x_6, 0, x_76); -if (lean_is_scalar(x_39)) { - x_77 = lean_alloc_ctor(0, 2, 0); +uint8_t x_69; +x_69 = !lean_is_exclusive(x_6); +if (x_69 == 0) +{ +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; lean_object* x_79; +x_70 = lean_ctor_get(x_6, 0); +x_71 = lean_array_push(x_52, x_70); +x_72 = l_Array_append___rarg(x_34, x_71); +x_73 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_73, 0, x_45); +lean_ctor_set(x_73, 1, x_46); +lean_ctor_set(x_73, 2, x_72); +x_74 = lean_array_push(x_60, x_73); +x_75 = lean_array_push(x_74, x_47); +x_76 = lean_array_push(x_75, x_49); +x_77 = lean_array_push(x_76, x_54); +x_78 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_78, 0, x_45); +lean_ctor_set(x_78, 1, x_3); +lean_ctor_set(x_78, 2, x_77); +lean_ctor_set(x_6, 0, x_78); +if (lean_is_scalar(x_41)) { + x_79 = lean_alloc_ctor(0, 2, 0); } else { - x_77 = x_39; + x_79 = x_41; } -lean_ctor_set(x_77, 0, x_6); -lean_ctor_set(x_77, 1, x_38); -return x_77; +lean_ctor_set(x_79, 0, x_6); +lean_ctor_set(x_79, 1, x_40); +return x_79; } else { -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_78 = lean_ctor_get(x_6, 0); -lean_inc(x_78); +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_80 = lean_ctor_get(x_6, 0); +lean_inc(x_80); lean_dec(x_6); -x_79 = lean_array_push(x_50, x_78); -x_80 = l_Array_append___rarg(x_34, x_79); -x_81 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_81, 0, x_43); -lean_ctor_set(x_81, 1, x_44); -lean_ctor_set(x_81, 2, x_80); -x_82 = lean_array_push(x_58, x_81); -x_83 = lean_array_push(x_82, x_45); -x_84 = lean_array_push(x_83, x_47); -x_85 = lean_array_push(x_84, x_52); -x_86 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_86, 0, x_43); -lean_ctor_set(x_86, 1, x_3); -lean_ctor_set(x_86, 2, x_85); -x_87 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_87, 0, x_86); -if (lean_is_scalar(x_39)) { - x_88 = lean_alloc_ctor(0, 2, 0); +x_81 = lean_array_push(x_52, x_80); +x_82 = l_Array_append___rarg(x_34, x_81); +x_83 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_83, 0, x_45); +lean_ctor_set(x_83, 1, x_46); +lean_ctor_set(x_83, 2, x_82); +x_84 = lean_array_push(x_60, x_83); +x_85 = lean_array_push(x_84, x_47); +x_86 = lean_array_push(x_85, x_49); +x_87 = lean_array_push(x_86, x_54); +x_88 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_88, 0, x_45); +lean_ctor_set(x_88, 1, x_3); +lean_ctor_set(x_88, 2, x_87); +x_89 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_89, 0, x_88); +if (lean_is_scalar(x_41)) { + x_90 = lean_alloc_ctor(0, 2, 0); } else { - x_88 = x_39; + x_90 = x_41; } -lean_ctor_set(x_88, 0, x_87); -lean_ctor_set(x_88, 1, x_38); -return x_88; +lean_ctor_set(x_90, 0, x_89); +lean_ctor_set(x_90, 1, x_40); +return x_90; } } } @@ -665,109 +1167,54 @@ x_9 = l_Lean_Syntax_getArg(x_1, x_8); x_10 = l_Lean_Syntax_isNone(x_9); if (x_10 == 0) { -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = l_Lean_nullKind; -x_12 = lean_unsigned_to_nat(1u); +lean_object* x_11; uint8_t x_12; +x_11 = lean_unsigned_to_nat(1u); lean_inc(x_9); -x_13 = l_Lean_Syntax_isNodeOf(x_9, x_11, x_12); -if (x_13 == 0) +x_12 = l_Lean_Syntax_matchesNull(x_9, x_11); +if (x_12 == 0) { -lean_object* x_14; lean_object* x_15; +lean_object* x_13; lean_object* x_14; lean_dec(x_9); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); -x_14 = lean_box(0); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_15, 1, x_7); -return x_15; +x_13 = lean_box(0); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_7); +return x_14; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_16 = lean_unsigned_to_nat(0u); -x_17 = l_Lean_Syntax_getArg(x_9, x_16); +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_15 = lean_unsigned_to_nat(0u); +x_16 = l_Lean_Syntax_getArg(x_9, x_15); lean_dec(x_9); -x_18 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_18, 0, x_17); -x_19 = lean_box(0); -x_20 = l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1(x_1, x_2, x_3, x_5, x_19, x_18, x_6, x_7); -return x_20; +x_17 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_17, 0, x_16); +x_18 = lean_box(0); +x_19 = l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1(x_1, x_2, x_3, x_5, x_18, x_17, x_6, x_7); +return x_19; } } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_dec(x_9); +x_20 = lean_box(0); x_21 = lean_box(0); -x_22 = lean_box(0); -x_23 = l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1(x_1, x_2, x_3, x_5, x_22, x_21, x_6, x_7); -return x_23; +x_22 = l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1(x_1, x_2, x_3, x_5, x_21, x_20, x_6, x_7); +return x_22; } } } static lean_object* _init_l_Lean_Elab_Term_expandMatchAlts_x3f___closed__1() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Lean", 4); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Term_expandMatchAlts_x3f___closed__2() { -_start: -{ lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_expandMatchAlts_x3f___closed__1; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Elab_Term_expandMatchAlts_x3f___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Parser", 6); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Term_expandMatchAlts_x3f___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_expandMatchAlts_x3f___closed__2; -x_2 = l_Lean_Elab_Term_expandMatchAlts_x3f___closed__3; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Elab_Term_expandMatchAlts_x3f___closed__5() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Term", 4); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Term_expandMatchAlts_x3f___closed__6() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_expandMatchAlts_x3f___closed__4; -x_2 = l_Lean_Elab_Term_expandMatchAlts_x3f___closed__5; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Elab_Term_expandMatchAlts_x3f___closed__7() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_expandMatchAlts_x3f___closed__6; -x_2 = l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__3; +x_1 = l_Lean_Elab_Term_expandMatchAlt___closed__6; +x_2 = l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__2; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } @@ -776,7 +1223,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandMatchAlts_x3f(lean_object* x_1, _start: { lean_object* x_4; uint8_t x_5; -x_4 = l_Lean_Elab_Term_expandMatchAlts_x3f___closed__7; +x_4 = l_Lean_Elab_Term_expandMatchAlts_x3f___closed__1; lean_inc(x_1); x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); if (x_5 == 0) @@ -798,47 +1245,46 @@ x_9 = l_Lean_Syntax_getArg(x_1, x_8); x_10 = l_Lean_Syntax_isNone(x_9); if (x_10 == 0) { -lean_object* x_11; uint8_t x_12; -x_11 = l_Lean_nullKind; +uint8_t x_11; lean_inc(x_9); -x_12 = l_Lean_Syntax_isNodeOf(x_9, x_11, x_8); -if (x_12 == 0) +x_11 = l_Lean_Syntax_matchesNull(x_9, x_8); +if (x_11 == 0) { -lean_object* x_13; lean_object* x_14; +lean_object* x_12; lean_object* x_13; lean_dec(x_9); lean_dec(x_2); lean_dec(x_1); -x_13 = lean_box(0); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_3); -return x_14; +x_12 = lean_box(0); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_3); +return x_13; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_15 = lean_unsigned_to_nat(0u); -x_16 = l_Lean_Syntax_getArg(x_9, x_15); +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_14 = lean_unsigned_to_nat(0u); +x_15 = l_Lean_Syntax_getArg(x_9, x_14); lean_dec(x_9); -x_17 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_17, 0, x_16); -x_18 = l_Lean_Elab_Term_expandMatchAlts_x3f___closed__6; -x_19 = lean_box(0); -x_20 = l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__2(x_1, x_18, x_4, x_19, x_17, x_2, x_3); +x_16 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_16, 0, x_15); +x_17 = l_Lean_Elab_Term_expandMatchAlt___closed__6; +x_18 = lean_box(0); +x_19 = l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__2(x_1, x_17, x_4, x_18, x_16, x_2, x_3); lean_dec(x_1); -return x_20; +return x_19; } } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_dec(x_9); -x_21 = lean_box(0); -x_22 = l_Lean_Elab_Term_expandMatchAlts_x3f___closed__6; -x_23 = lean_box(0); -x_24 = l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__2(x_1, x_22, x_4, x_23, x_21, x_2, x_3); +x_20 = lean_box(0); +x_21 = l_Lean_Elab_Term_expandMatchAlt___closed__6; +x_22 = lean_box(0); +x_23 = l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__2(x_1, x_21, x_4, x_22, x_20, x_2, x_3); lean_dec(x_1); -return x_24; +return x_23; } } } @@ -857,17 +1303,17 @@ x_7 = lean_box(x_6); return x_7; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_expandMatchAlts_x3f___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_expandMatchAlts_x3f___spec__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: { -size_t x_5; size_t x_6; lean_object* x_7; -x_5 = lean_unbox_usize(x_2); +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = lean_unbox_usize(x_3); +x_8 = lean_unbox_usize(x_3); lean_dec(x_3); -x_7 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_expandMatchAlts_x3f___spec__2(x_1, x_5, x_6, x_4); +x_9 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_expandMatchAlts_x3f___spec__2(x_1, x_7, x_8, x_4, x_5, x_6); lean_dec(x_1); -return x_7; +return x_9; } } LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandMatchAlts_x3f___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) { @@ -902,8 +1348,43 @@ lean_dec_ref(res); res = initialize_Lean_Parser_Term(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__1___closed__1 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__1___closed__1(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__1___closed__1); +l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__1 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__1(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__1); +l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__2 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__2(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__2); +l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__3 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__3(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__3); +l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__4 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__4(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__4); +l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__5 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__5(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__5); +l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__6 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__6(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAlt___spec__3___closed__6); +l_Lean_Elab_Term_expandMatchAlt___closed__1 = _init_l_Lean_Elab_Term_expandMatchAlt___closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_expandMatchAlt___closed__1); +l_Lean_Elab_Term_expandMatchAlt___closed__2 = _init_l_Lean_Elab_Term_expandMatchAlt___closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_expandMatchAlt___closed__2); +l_Lean_Elab_Term_expandMatchAlt___closed__3 = _init_l_Lean_Elab_Term_expandMatchAlt___closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_expandMatchAlt___closed__3); +l_Lean_Elab_Term_expandMatchAlt___closed__4 = _init_l_Lean_Elab_Term_expandMatchAlt___closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_expandMatchAlt___closed__4); +l_Lean_Elab_Term_expandMatchAlt___closed__5 = _init_l_Lean_Elab_Term_expandMatchAlt___closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_expandMatchAlt___closed__5); +l_Lean_Elab_Term_expandMatchAlt___closed__6 = _init_l_Lean_Elab_Term_expandMatchAlt___closed__6(); +lean_mark_persistent(l_Lean_Elab_Term_expandMatchAlt___closed__6); +l_Lean_Elab_Term_expandMatchAlt___closed__7 = _init_l_Lean_Elab_Term_expandMatchAlt___closed__7(); +lean_mark_persistent(l_Lean_Elab_Term_expandMatchAlt___closed__7); +l_Lean_Elab_Term_expandMatchAlt___closed__8 = _init_l_Lean_Elab_Term_expandMatchAlt___closed__8(); +lean_mark_persistent(l_Lean_Elab_Term_expandMatchAlt___closed__8); +l_Lean_Elab_Term_expandMatchAlt___closed__9 = _init_l_Lean_Elab_Term_expandMatchAlt___closed__9(); +lean_mark_persistent(l_Lean_Elab_Term_expandMatchAlt___closed__9); +l_Lean_Elab_Term_expandMatchAlt___closed__10 = _init_l_Lean_Elab_Term_expandMatchAlt___closed__10(); +lean_mark_persistent(l_Lean_Elab_Term_expandMatchAlt___closed__10); +l_Lean_Elab_Term_expandMatchAlt___closed__11 = _init_l_Lean_Elab_Term_expandMatchAlt___closed__11(); +lean_mark_persistent(l_Lean_Elab_Term_expandMatchAlt___closed__11); +l_Lean_Elab_Term_shouldExpandMatchAlt___closed__1 = _init_l_Lean_Elab_Term_shouldExpandMatchAlt___closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_shouldExpandMatchAlt___closed__1); +l_Lean_Elab_Term_shouldExpandMatchAlt___closed__2 = _init_l_Lean_Elab_Term_shouldExpandMatchAlt___closed__2(); l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__1 = _init_l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__1); l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__2 = _init_l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__2(); @@ -916,26 +1397,8 @@ l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__5 = _init_l_Lean_Elab lean_mark_persistent(l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__5); l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__6 = _init_l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__6(); lean_mark_persistent(l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__6); -l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__7 = _init_l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__7(); -lean_mark_persistent(l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__7); -l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__8 = _init_l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__8(); -lean_mark_persistent(l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__8); -l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__9 = _init_l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__9(); -lean_mark_persistent(l_Lean_Elab_Term_expandMatchAlts_x3f___lambda__1___closed__9); l_Lean_Elab_Term_expandMatchAlts_x3f___closed__1 = _init_l_Lean_Elab_Term_expandMatchAlts_x3f___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_expandMatchAlts_x3f___closed__1); -l_Lean_Elab_Term_expandMatchAlts_x3f___closed__2 = _init_l_Lean_Elab_Term_expandMatchAlts_x3f___closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_expandMatchAlts_x3f___closed__2); -l_Lean_Elab_Term_expandMatchAlts_x3f___closed__3 = _init_l_Lean_Elab_Term_expandMatchAlts_x3f___closed__3(); -lean_mark_persistent(l_Lean_Elab_Term_expandMatchAlts_x3f___closed__3); -l_Lean_Elab_Term_expandMatchAlts_x3f___closed__4 = _init_l_Lean_Elab_Term_expandMatchAlts_x3f___closed__4(); -lean_mark_persistent(l_Lean_Elab_Term_expandMatchAlts_x3f___closed__4); -l_Lean_Elab_Term_expandMatchAlts_x3f___closed__5 = _init_l_Lean_Elab_Term_expandMatchAlts_x3f___closed__5(); -lean_mark_persistent(l_Lean_Elab_Term_expandMatchAlts_x3f___closed__5); -l_Lean_Elab_Term_expandMatchAlts_x3f___closed__6 = _init_l_Lean_Elab_Term_expandMatchAlts_x3f___closed__6(); -lean_mark_persistent(l_Lean_Elab_Term_expandMatchAlts_x3f___closed__6); -l_Lean_Elab_Term_expandMatchAlts_x3f___closed__7 = _init_l_Lean_Elab_Term_expandMatchAlts_x3f___closed__7(); -lean_mark_persistent(l_Lean_Elab_Term_expandMatchAlts_x3f___closed__7); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Elab/BuiltinCommand.c b/stage0/stdlib/Lean/Elab/BuiltinCommand.c index 32955eea713..843ebfc1e52 100644 --- a/stage0/stdlib/Lean/Elab/BuiltinCommand.c +++ b/stage0/stdlib/Lean/Elab/BuiltinCommand.c @@ -53,8 +53,8 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabExp LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCheckCore___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_nestedExceptionToMessageData___at_Lean_Elab_Command_elabExport___spec__13___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Command_elabModuleDoc___closed__4; -extern lean_object* l_Lean_nullKind; LEAN_EXPORT lean_object* l_Lean_Elab_Command_withNamespace(lean_object*); +lean_object* l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabSynth_declRange(lean_object*); static lean_object* l_Lean_Elab_Command_elabEvalUnsafe___lambda__2___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Command_elabCheck_declRange___closed__4; @@ -63,6 +63,7 @@ lean_object* l_Lean_Elab_Command_liftTermElabM___rarg(lean_object*, lean_object* static lean_object* l_Lean_Elab_Command_elabCheckCore___lambda__2___closed__1; static lean_object* l_Lean_Elab_Command_elabCheckCore___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabEnd_declRange(lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__13; static lean_object* l_Lean_Elab_throwErrorWithNestedErrors___at_Lean_Elab_Command_elabExport___spec__12___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabReduce___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_name_mk_string(lean_object*, lean_object*); @@ -129,7 +130,6 @@ lean_object* l_Lean_logAt___at_Lean_Elab_Term_traceAtCmdPos___spec__3(lean_objec LEAN_EXPORT lean_object* l_Lean_Elab_Command_hasNoErrorMessages___rarg___boxed(lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_withFreshMacroScope___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabOpen___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabVariable___spec__4___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_matchBinderNames(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Command_elabEnd___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); @@ -268,6 +268,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_OpenDecl_resolveId___at_Lean_Elab_Command_e static lean_object* l___regBuiltin_Lean_Elab_Command_elabNamespace_declRange___closed__1; static lean_object* l_Lean_Elab_Command_elabCheckFailure___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabOpen___lambda__1(lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__11; LEAN_EXPORT lean_object* l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_mkRunMetaEval___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_mkEvalInstCore___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Command_elabNamespace_declRange___closed__4; @@ -280,23 +281,24 @@ LEAN_EXPORT lean_object* l_Lean_Elab_elabSetOption_setOption___at_Lean_Elab_Comm static lean_object* l___regBuiltin_Lean_Elab_Command_elabUniverse___closed__4; lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Command_elabExport___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_isNatLit_x3f(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_matchBinderNames___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_mkRunMetaEval___closed__2; lean_object* l_Lean_Elab_Term_levelMVarToParam(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___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_addScopes(uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCheck___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabEvalUnsafe___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabUniverse_declRange___closed__7; lean_object* lean_st_ref_take(lean_object*, lean_object*); lean_object* l_Lean_getOptionDecl(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_mkRunMetaEval___closed__1; -extern lean_object* l_Lean_numLitKind; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabEvalUnsafe___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__1; lean_object* l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabVariable___closed__2; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___boxed(lean_object*, lean_object*, 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_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___boxed(lean_object*, 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_resolveNamespace___at_Lean_Elab_Command_elabOpen___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getRefPos___at_Lean_Elab_Command_elabExport___spec__14(lean_object*); @@ -351,6 +353,7 @@ lean_object* l_Lean_ScopedEnvExtension_pushScope___rarg(lean_object*, lean_objec static lean_object* l___regBuiltin_Lean_Elab_Command_elabCheckFailure_declRange___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Command_elabOpen___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCheckCore___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__12; lean_object* l_Lean_Syntax_setArgs(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabOpen___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_mkEvalInstCore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -448,7 +451,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Command_elabModuleDoc___closed__10; static lean_object* l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_typelessBinder_x3f___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSetOption___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___boxed(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_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabSetOption___at_Lean_Elab_Command_elabSetOption___spec__1___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Command_elabChoice_declRange___closed__4; @@ -521,12 +524,10 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabOpe lean_object* l_Lean_Elab_Command_addUnivLevel(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabSection_declRange___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Command_elabUniverse_declRange___closed__2; -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); lean_object* l_List_drop___rarg(lean_object*, lean_object*); extern lean_object* l_Lean_instInhabitedException; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabNonComputableSection(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_matchBinderNames___closed__1; -lean_object* l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_failIfSucceeds___closed__1; lean_object* l_Lean_Syntax_getPos_x3f(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabNamespace___spec__1___rarg(lean_object*); @@ -556,7 +557,7 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabExport___s LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabCheckFailure_declRange(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabNonComputableSection_declRange___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Command_elabCheckFailure_declRange___closed__3; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2(lean_object*, uint8_t, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2(lean_object*, uint8_t, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabModuleDoc___closed__12; LEAN_EXPORT lean_object* l_Lean_Elab_nestedExceptionToMessageData___at_Lean_Elab_Command_elabExport___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_logException___at_Lean_Elab_Command_runLinters___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); @@ -657,7 +658,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Command_elabNonComputableSection___ uint8_t l_Lean_Syntax_isNone(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabNamespace_declRange___closed__7; lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandInCmd___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabNamespace(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabEvalUnsafe___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabModuleDoc___closed__5; @@ -672,7 +672,6 @@ LEAN_EXPORT lean_object* l_Lean_resolveNamespace___at_Lean_Elab_Command_elabExpo static lean_object* l___regBuiltin_Lean_Elab_Command_elabInitQuot___closed__1; static lean_object* l_Lean_Elab_Command_elabCheckFailure___closed__3; lean_object* l_Lean_Meta_instantiateMVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Command_expandInCmd___closed__6; static lean_object* l_Lean_resolveUniqueNamespace___at_Lean_Elab_Command_elabOpen___spec__3___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Command_elabModuleDoc___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Command_elabChoice_declRange___closed__6; @@ -704,7 +703,6 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabOpe static lean_object* l___regBuiltin_Lean_Elab_Command_elabSynth_declRange___closed__2; lean_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabEval(lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Command_expandInCmd___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Command_elabSynth_declRange___closed__4; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabReduce_declRange___closed__1; @@ -784,6 +782,7 @@ lean_object* l_Lean_mkConst(lean_object*, lean_object*); lean_object* l_Lean_mkSimpleThunk(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabEnd___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabNamespace___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_expandInCmd___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Open_0__Lean_Elab_OpenDecl_resolveNameUsingNamespacesCore___at_Lean_Elab_Command_elabExport___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabModuleDoc_declRange(lean_object*); @@ -791,7 +790,7 @@ static lean_object* l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_ela static lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Command_elabExport___spec__8___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Command_elabSection___closed__3; static lean_object* l_Lean_Elab_Command_elabNamespace___closed__1; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ScopedEnvExtension_activateScoped___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabVariable___lambda__3(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_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabOpen___spec__7(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -3135,69 +3134,68 @@ return x_7; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +lean_object* x_8; lean_object* x_9; uint8_t x_10; x_8 = lean_unsigned_to_nat(1u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); lean_dec(x_1); -x_10 = l_Lean_nullKind; lean_inc(x_9); -x_11 = l_Lean_Syntax_isNodeOf(x_9, x_10, x_8); -if (x_11 == 0) +x_10 = l_Lean_Syntax_matchesNull(x_9, x_8); +if (x_10 == 0) { -lean_object* x_12; uint8_t x_13; -x_12 = lean_unsigned_to_nat(0u); -x_13 = l_Lean_Syntax_isNodeOf(x_9, x_10, x_12); -if (x_13 == 0) +lean_object* x_11; uint8_t x_12; +x_11 = lean_unsigned_to_nat(0u); +x_12 = l_Lean_Syntax_matchesNull(x_9, x_11); +if (x_12 == 0) { -lean_object* x_14; +lean_object* x_13; lean_dec(x_2); -x_14 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabNamespace___spec__1___rarg(x_4); -return x_14; +x_13 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabNamespace___spec__1___rarg(x_4); +return x_13; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_object* x_20; lean_object* x_21; -x_15 = l_Lean_Elab_Command_getScope___rarg(x_3, x_4); -x_16 = lean_ctor_get(x_15, 0); +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; lean_object* x_20; +x_14 = l_Lean_Elab_Command_getScope___rarg(x_3, x_4); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +lean_dec(x_14); +x_17 = lean_ctor_get(x_15, 2); lean_inc(x_17); lean_dec(x_15); -x_18 = lean_ctor_get(x_16, 2); -lean_inc(x_18); -lean_dec(x_16); -x_19 = 0; -x_20 = l_Lean_Elab_Command_elabModuleDoc___closed__3; -x_21 = l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_addScope(x_19, x_19, x_20, x_18, x_2, x_3, x_17); +x_18 = 0; +x_19 = l_Lean_Elab_Command_elabModuleDoc___closed__3; +x_20 = l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_addScope(x_18, x_18, x_19, x_17, x_2, x_3, x_16); lean_dec(x_2); -return x_21; +return x_20; } } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_22 = lean_unsigned_to_nat(0u); -x_23 = l_Lean_Syntax_getArg(x_9, x_22); +lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_21 = lean_unsigned_to_nat(0u); +x_22 = l_Lean_Syntax_getArg(x_9, x_21); lean_dec(x_9); -x_24 = l_Lean_Elab_Command_elabSection___closed__4; -lean_inc(x_23); -x_25 = l_Lean_Syntax_isOfKind(x_23, x_24); -if (x_25 == 0) +x_23 = l_Lean_Elab_Command_elabSection___closed__4; +lean_inc(x_22); +x_24 = l_Lean_Syntax_isOfKind(x_22, x_23); +if (x_24 == 0) { -lean_object* x_26; -lean_dec(x_23); +lean_object* x_25; +lean_dec(x_22); lean_dec(x_2); -x_26 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabNamespace___spec__1___rarg(x_4); -return x_26; +x_25 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabNamespace___spec__1___rarg(x_4); +return x_25; } else { -lean_object* x_27; uint8_t x_28; lean_object* x_29; -x_27 = l_Lean_Syntax_getId(x_23); -lean_dec(x_23); -x_28 = 0; -x_29 = l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_addScopes(x_28, x_28, x_27, x_2, x_3, x_4); -return x_29; +lean_object* x_26; uint8_t x_27; lean_object* x_28; +x_26 = l_Lean_Syntax_getId(x_22); +lean_dec(x_22); +x_27 = 0; +x_28 = l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_addScopes(x_27, x_27, x_26, x_2, x_3, x_4); +return x_28; } } } @@ -3387,72 +3385,71 @@ return x_7; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; x_8 = lean_unsigned_to_nat(2u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); lean_dec(x_1); -x_10 = l_Lean_nullKind; -x_11 = lean_unsigned_to_nat(1u); +x_10 = lean_unsigned_to_nat(1u); lean_inc(x_9); -x_12 = l_Lean_Syntax_isNodeOf(x_9, x_10, x_11); -if (x_12 == 0) +x_11 = l_Lean_Syntax_matchesNull(x_9, x_10); +if (x_11 == 0) { -lean_object* x_13; uint8_t x_14; -x_13 = lean_unsigned_to_nat(0u); -x_14 = l_Lean_Syntax_isNodeOf(x_9, x_10, x_13); -if (x_14 == 0) +lean_object* x_12; uint8_t x_13; +x_12 = lean_unsigned_to_nat(0u); +x_13 = l_Lean_Syntax_matchesNull(x_9, x_12); +if (x_13 == 0) { -lean_object* x_15; +lean_object* x_14; lean_dec(x_2); -x_15 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabNamespace___spec__1___rarg(x_4); -return x_15; +x_14 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabNamespace___spec__1___rarg(x_4); +return x_14; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; uint8_t x_21; lean_object* x_22; lean_object* x_23; -x_16 = l_Lean_Elab_Command_getScope___rarg(x_3, x_4); -x_17 = lean_ctor_get(x_16, 0); +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; +x_15 = l_Lean_Elab_Command_getScope___rarg(x_3, x_4); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); +lean_dec(x_15); +x_18 = lean_ctor_get(x_16, 2); lean_inc(x_18); lean_dec(x_16); -x_19 = lean_ctor_get(x_17, 2); -lean_inc(x_19); -lean_dec(x_17); -x_20 = 0; -x_21 = 1; -x_22 = l_Lean_Elab_Command_elabModuleDoc___closed__3; -x_23 = l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_addScope(x_20, x_21, x_22, x_19, x_2, x_3, x_18); +x_19 = 0; +x_20 = 1; +x_21 = l_Lean_Elab_Command_elabModuleDoc___closed__3; +x_22 = l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_addScope(x_19, x_20, x_21, x_18, x_2, x_3, x_17); lean_dec(x_2); -return x_23; +return x_22; } } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; -x_24 = lean_unsigned_to_nat(0u); -x_25 = l_Lean_Syntax_getArg(x_9, x_24); +lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_23 = lean_unsigned_to_nat(0u); +x_24 = l_Lean_Syntax_getArg(x_9, x_23); lean_dec(x_9); -x_26 = l_Lean_Elab_Command_elabSection___closed__4; -lean_inc(x_25); -x_27 = l_Lean_Syntax_isOfKind(x_25, x_26); -if (x_27 == 0) +x_25 = l_Lean_Elab_Command_elabSection___closed__4; +lean_inc(x_24); +x_26 = l_Lean_Syntax_isOfKind(x_24, x_25); +if (x_26 == 0) { -lean_object* x_28; -lean_dec(x_25); +lean_object* x_27; +lean_dec(x_24); lean_dec(x_2); -x_28 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabNamespace___spec__1___rarg(x_4); -return x_28; +x_27 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabNamespace___spec__1___rarg(x_4); +return x_27; } else { -lean_object* x_29; uint8_t x_30; uint8_t x_31; lean_object* x_32; -x_29 = l_Lean_Syntax_getId(x_25); -lean_dec(x_25); -x_30 = 0; -x_31 = 1; -x_32 = l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_addScopes(x_30, x_31, x_29, x_2, x_3, x_4); -return x_32; +lean_object* x_28; uint8_t x_29; uint8_t x_30; lean_object* x_31; +x_28 = l_Lean_Syntax_getId(x_24); +lean_dec(x_24); +x_29 = 0; +x_30 = 1; +x_31 = l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_addScopes(x_29, x_30, x_28, x_2, x_3, x_4); +return x_31; } } } @@ -5756,7 +5753,7 @@ x_8 = lean_ctor_get(x_6, 1); lean_inc(x_8); lean_dec(x_6); x_9 = lean_box(0); -x_10 = l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(x_7, x_9); +x_10 = l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(x_7, x_9); x_11 = l_List_isEmpty___rarg(x_10); if (x_11 == 0) { @@ -10562,93 +10559,91 @@ return x_6; } else { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; x_7 = lean_unsigned_to_nat(1u); x_8 = l_Lean_Syntax_getArg(x_1, x_7); x_9 = lean_unsigned_to_nat(2u); x_10 = l_Lean_Syntax_getArg(x_1, x_9); lean_dec(x_1); -x_11 = l_Lean_nullKind; -x_12 = lean_unsigned_to_nat(0u); -x_13 = l_Lean_Syntax_isNodeOf(x_10, x_11, x_12); -if (x_13 == 0) +x_11 = lean_unsigned_to_nat(0u); +x_12 = l_Lean_Syntax_matchesNull(x_10, x_11); +if (x_12 == 0) { -lean_object* x_14; +lean_object* x_13; lean_dec(x_8); -x_14 = lean_box(0); -return x_14; +x_13 = lean_box(0); +return x_13; } else { -lean_object* x_15; lean_object* x_16; size_t x_17; size_t x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_15 = l_Lean_Syntax_getArgs(x_8); +lean_object* x_14; lean_object* x_15; size_t x_16; size_t x_17; lean_object* x_18; uint8_t x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_14 = l_Lean_Syntax_getArgs(x_8); lean_dec(x_8); -x_16 = lean_array_get_size(x_15); -x_17 = lean_usize_of_nat(x_16); -lean_dec(x_16); -x_18 = 0; -x_19 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_getBracketedBinderIds___spec__1(x_17, x_18, x_15); -x_20 = 0; -x_21 = lean_box(x_20); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_19); -lean_ctor_set(x_22, 1, x_21); -x_23 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_23, 0, x_22); -return x_23; +x_15 = lean_array_get_size(x_14); +x_16 = lean_usize_of_nat(x_15); +lean_dec(x_15); +x_17 = 0; +x_18 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_getBracketedBinderIds___spec__1(x_16, x_17, x_14); +x_19 = 0; +x_20 = lean_box(x_19); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_18); +lean_ctor_set(x_21, 1, x_20); +x_22 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_22, 0, x_21); +return x_22; } } } else { -lean_object* x_24; 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_24 = lean_unsigned_to_nat(1u); -x_25 = l_Lean_Syntax_getArg(x_1, x_24); -x_26 = lean_unsigned_to_nat(2u); -x_27 = l_Lean_Syntax_getArg(x_1, x_26); -x_28 = l_Lean_nullKind; -x_29 = lean_unsigned_to_nat(0u); -x_30 = l_Lean_Syntax_isNodeOf(x_27, x_28, x_29); -if (x_30 == 0) +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_23 = lean_unsigned_to_nat(1u); +x_24 = l_Lean_Syntax_getArg(x_1, x_23); +x_25 = lean_unsigned_to_nat(2u); +x_26 = l_Lean_Syntax_getArg(x_1, x_25); +x_27 = lean_unsigned_to_nat(0u); +x_28 = l_Lean_Syntax_matchesNull(x_26, x_27); +if (x_28 == 0) { -lean_object* x_31; -lean_dec(x_25); +lean_object* x_29; +lean_dec(x_24); lean_dec(x_1); -x_31 = lean_box(0); -return x_31; +x_29 = lean_box(0); +return x_29; } else { -lean_object* x_32; lean_object* x_33; uint8_t x_34; -x_32 = lean_unsigned_to_nat(3u); -x_33 = l_Lean_Syntax_getArg(x_1, x_32); +lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_30 = lean_unsigned_to_nat(3u); +x_31 = l_Lean_Syntax_getArg(x_1, x_30); lean_dec(x_1); -x_34 = l_Lean_Syntax_isNodeOf(x_33, x_28, x_29); -if (x_34 == 0) +x_32 = l_Lean_Syntax_matchesNull(x_31, x_27); +if (x_32 == 0) { -lean_object* x_35; -lean_dec(x_25); -x_35 = lean_box(0); -return x_35; +lean_object* x_33; +lean_dec(x_24); +x_33 = lean_box(0); +return x_33; } else { -lean_object* x_36; lean_object* x_37; size_t x_38; size_t x_39; lean_object* x_40; uint8_t x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_36 = l_Lean_Syntax_getArgs(x_25); -lean_dec(x_25); -x_37 = lean_array_get_size(x_36); -x_38 = lean_usize_of_nat(x_37); -lean_dec(x_37); -x_39 = 0; -x_40 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_getBracketedBinderIds___spec__1(x_38, x_39, x_36); -x_41 = 1; -x_42 = lean_box(x_41); -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_40); -lean_ctor_set(x_43, 1, x_42); -x_44 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_44, 0, x_43); -return x_44; +lean_object* x_34; lean_object* x_35; size_t x_36; size_t x_37; lean_object* x_38; uint8_t x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_34 = l_Lean_Syntax_getArgs(x_24); +lean_dec(x_24); +x_35 = lean_array_get_size(x_34); +x_36 = lean_usize_of_nat(x_35); +lean_dec(x_35); +x_37 = 0; +x_38 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_getBracketedBinderIds___spec__1(x_36, x_37, x_34); +x_39 = 1; +x_40 = lean_box(x_39); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_38); +lean_ctor_set(x_41, 1, x_40); +x_42 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_42, 0, x_41); +return x_42; } } } @@ -11167,12 +11162,35 @@ static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Buil _start: { lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Elab_Open_0__Lean_Elab_OpenDecl_resolveNameUsingNamespacesCore___at_Lean_Elab_Command_elabExport___spec__3___closed__1; +x_2 = l_Array_append___rarg(x_1, x_1); +return x_2; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(2); +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__4; +x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__7; +x_4 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; x_1 = lean_unsigned_to_nat(2u); x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__8() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__10() { _start: { lean_object* x_1; @@ -11180,7 +11198,21 @@ x_1 = lean_mk_string_from_bytes("(", 1); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__9() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(2); +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__4; +x_3 = l___private_Lean_Elab_Open_0__Lean_Elab_OpenDecl_resolveNameUsingNamespacesCore___at_Lean_Elab_Command_elabExport___spec__3___closed__1; +x_4 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__12() { _start: { lean_object* x_1; @@ -11188,7 +11220,7 @@ x_1 = lean_mk_string_from_bytes(")", 1); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__10() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__13() { _start: { lean_object* x_1; lean_object* x_2; @@ -11197,38 +11229,38 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, uint8_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_13; -x_13 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__1; +lean_object* x_12; +x_12 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__1; if (x_1 == 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; 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_object* x_32; lean_object* x_33; -lean_dec(x_6); -x_14 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__1(x_10, x_11, x_12); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); -lean_inc(x_16); -lean_dec(x_14); -x_17 = l_Lean_Elab_Command_getCurrMacroScope(x_10, x_11, x_16); -x_18 = lean_ctor_get(x_17, 1); -lean_inc(x_18); -lean_dec(x_17); -x_19 = l_Lean_Elab_Command_getMainModule___rarg(x_11, x_18); -x_20 = lean_ctor_get(x_19, 1); -lean_inc(x_20); -lean_dec(x_19); -x_21 = l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_typelessBinder_x3f___closed__5; -x_22 = lean_name_mk_string(x_2, x_21); -x_23 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__2; +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; 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; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +lean_dec(x_5); +x_13 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__1(x_9, x_10, x_11); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); lean_inc(x_15); -x_24 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_24, 0, x_15); -lean_ctor_set(x_24, 1, x_23); -lean_inc(x_3); -x_25 = l_Array_append___rarg(x_3, x_4); +lean_dec(x_13); +x_16 = l_Lean_Elab_Command_getCurrMacroScope(x_9, x_10, x_15); +x_17 = lean_ctor_get(x_16, 1); +lean_inc(x_17); +lean_dec(x_16); +x_18 = l_Lean_Elab_Command_getMainModule___rarg(x_10, x_17); +x_19 = lean_ctor_get(x_18, 1); +lean_inc(x_19); +lean_dec(x_18); +x_20 = l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_typelessBinder_x3f___closed__5; +x_21 = lean_name_mk_string(x_2, x_20); +x_22 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__2; +lean_inc(x_14); +x_23 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_23, 0, x_14); +lean_ctor_set(x_23, 1, x_22); +x_24 = l___private_Lean_Elab_Open_0__Lean_Elab_OpenDecl_resolveNameUsingNamespacesCore___at_Lean_Elab_Command_elabExport___spec__3___closed__1; +x_25 = l_Array_append___rarg(x_24, x_3); x_26 = lean_box(2); x_27 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__4; x_28 = lean_alloc_ctor(1, 3, 0); @@ -11236,162 +11268,149 @@ lean_ctor_set(x_28, 0, x_26); lean_ctor_set(x_28, 1, x_27); lean_ctor_set(x_28, 2, x_25); x_29 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__5; -lean_inc(x_15); +lean_inc(x_14); x_30 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_30, 0, x_15); +lean_ctor_set(x_30, 0, x_14); lean_ctor_set(x_30, 1, x_29); x_31 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__6; -x_32 = lean_array_push(x_31, x_24); +x_32 = lean_array_push(x_31, x_23); x_33 = lean_array_push(x_32, x_28); -if (lean_obj_tag(x_5) == 0) +if (lean_obj_tag(x_4) == 0) { -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_dec(x_15); -lean_inc(x_3); -x_34 = l_Array_append___rarg(x_3, x_3); -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_26); -lean_ctor_set(x_35, 1, x_27); -lean_ctor_set(x_35, 2, x_34); -x_36 = lean_array_push(x_33, x_35); -x_37 = lean_array_push(x_36, x_30); -x_38 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_38, 0, x_26); -lean_ctor_set(x_38, 1, x_22); -lean_ctor_set(x_38, 2, x_37); -x_39 = lean_array_push(x_8, x_38); -x_40 = lean_box(0); -x_41 = lean_box(x_7); -x_42 = lean_apply_6(x_13, x_41, x_39, x_40, x_10, x_11, x_20); -return x_42; +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_dec(x_14); +x_34 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__8; +x_35 = lean_array_push(x_33, x_34); +x_36 = lean_array_push(x_35, x_30); +x_37 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_37, 0, x_26); +lean_ctor_set(x_37, 1, x_21); +lean_ctor_set(x_37, 2, x_36); +x_38 = lean_array_push(x_7, x_37); +x_39 = lean_box(0); +x_40 = lean_box(x_6); +x_41 = lean_apply_6(x_12, x_40, x_38, x_39, x_9, x_10, x_19); +return x_41; } else { -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_43 = lean_ctor_get(x_5, 0); -lean_inc(x_43); -lean_dec(x_5); -x_44 = l_Lean_Elab_nestedExceptionToMessageData___at_Lean_Elab_Command_elabExport___spec__13___closed__1; -x_45 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_45, 0, x_15); -lean_ctor_set(x_45, 1, x_44); -x_46 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__7; -x_47 = lean_array_push(x_46, x_45); -x_48 = lean_array_push(x_47, x_43); -x_49 = l_Array_append___rarg(x_3, x_48); -x_50 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_50, 0, x_26); -lean_ctor_set(x_50, 1, x_27); -lean_ctor_set(x_50, 2, x_49); -x_51 = lean_array_push(x_33, x_50); -x_52 = lean_array_push(x_51, x_30); -x_53 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_53, 0, x_26); -lean_ctor_set(x_53, 1, x_22); -lean_ctor_set(x_53, 2, x_52); -x_54 = lean_array_push(x_8, x_53); -x_55 = lean_box(0); -x_56 = lean_box(x_7); -x_57 = lean_apply_6(x_13, x_56, x_54, x_55, x_10, x_11, x_20); -return x_57; +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_42 = lean_ctor_get(x_4, 0); +lean_inc(x_42); +lean_dec(x_4); +x_43 = l_Lean_Elab_nestedExceptionToMessageData___at_Lean_Elab_Command_elabExport___spec__13___closed__1; +x_44 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_44, 0, x_14); +lean_ctor_set(x_44, 1, x_43); +x_45 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__9; +x_46 = lean_array_push(x_45, x_44); +x_47 = lean_array_push(x_46, x_42); +x_48 = l_Array_append___rarg(x_24, x_47); +x_49 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_49, 0, x_26); +lean_ctor_set(x_49, 1, x_27); +lean_ctor_set(x_49, 2, x_48); +x_50 = lean_array_push(x_33, x_49); +x_51 = lean_array_push(x_50, x_30); +x_52 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_52, 0, x_26); +lean_ctor_set(x_52, 1, x_21); +lean_ctor_set(x_52, 2, x_51); +x_53 = lean_array_push(x_7, x_52); +x_54 = lean_box(0); +x_55 = lean_box(x_6); +x_56 = lean_apply_6(x_12, x_55, x_53, x_54, x_9, x_10, x_19); +return x_56; } } else { -lean_object* x_58; lean_object* x_59; 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; 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_57; lean_object* x_58; lean_object* x_59; 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; 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_dec(x_2); -x_58 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__1(x_10, x_11, x_12); -x_59 = lean_ctor_get(x_58, 0); +x_57 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__1(x_9, x_10, x_11); +x_58 = lean_ctor_get(x_57, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_57, 1); lean_inc(x_59); -x_60 = lean_ctor_get(x_58, 1); -lean_inc(x_60); -lean_dec(x_58); -x_61 = l_Lean_Elab_Command_getCurrMacroScope(x_10, x_11, x_60); -x_62 = lean_ctor_get(x_61, 1); -lean_inc(x_62); -lean_dec(x_61); -x_63 = l_Lean_Elab_Command_getMainModule___rarg(x_11, x_62); -x_64 = lean_ctor_get(x_63, 1); -lean_inc(x_64); -lean_dec(x_63); -x_65 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__8; -lean_inc(x_59); -x_66 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_66, 0, x_59); -lean_ctor_set(x_66, 1, x_65); -lean_inc(x_3); -x_67 = l_Array_append___rarg(x_3, x_4); +lean_dec(x_57); +x_60 = l_Lean_Elab_Command_getCurrMacroScope(x_9, x_10, x_59); +x_61 = lean_ctor_get(x_60, 1); +lean_inc(x_61); +lean_dec(x_60); +x_62 = l_Lean_Elab_Command_getMainModule___rarg(x_10, x_61); +x_63 = lean_ctor_get(x_62, 1); +lean_inc(x_63); +lean_dec(x_62); +x_64 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__10; +lean_inc(x_58); +x_65 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_65, 0, x_58); +lean_ctor_set(x_65, 1, x_64); +x_66 = l___private_Lean_Elab_Open_0__Lean_Elab_OpenDecl_resolveNameUsingNamespacesCore___at_Lean_Elab_Command_elabExport___spec__3___closed__1; +x_67 = l_Array_append___rarg(x_66, x_3); x_68 = lean_box(2); x_69 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__4; x_70 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_70, 0, x_68); lean_ctor_set(x_70, 1, x_69); lean_ctor_set(x_70, 2, x_67); -lean_inc(x_3); -x_71 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_71, 0, x_68); -lean_ctor_set(x_71, 1, x_69); -lean_ctor_set(x_71, 2, x_3); -x_72 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__9; -lean_inc(x_59); -x_73 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_73, 0, x_59); -lean_ctor_set(x_73, 1, x_72); -x_74 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__10; -x_75 = lean_array_push(x_74, x_66); -x_76 = lean_array_push(x_75, x_70); -if (lean_obj_tag(x_5) == 0) -{ -lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; -lean_dec(x_59); -lean_inc(x_3); -x_77 = l_Array_append___rarg(x_3, x_3); -x_78 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_78, 0, x_68); -lean_ctor_set(x_78, 1, x_69); -lean_ctor_set(x_78, 2, x_77); -x_79 = lean_array_push(x_76, x_78); -x_80 = lean_array_push(x_79, x_71); -x_81 = lean_array_push(x_80, x_73); -x_82 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_82, 0, x_68); -lean_ctor_set(x_82, 1, x_6); -lean_ctor_set(x_82, 2, x_81); -x_83 = lean_array_push(x_8, x_82); -x_84 = lean_box(0); -x_85 = lean_box(x_7); -x_86 = lean_apply_6(x_13, x_85, x_83, x_84, x_10, x_11, x_64); -return x_86; +x_71 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__12; +lean_inc(x_58); +x_72 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_72, 0, x_58); +lean_ctor_set(x_72, 1, x_71); +x_73 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__13; +x_74 = lean_array_push(x_73, x_65); +x_75 = lean_array_push(x_74, x_70); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +lean_dec(x_58); +x_76 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__8; +x_77 = lean_array_push(x_75, x_76); +x_78 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__11; +x_79 = lean_array_push(x_77, x_78); +x_80 = lean_array_push(x_79, x_72); +x_81 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_81, 0, x_68); +lean_ctor_set(x_81, 1, x_5); +lean_ctor_set(x_81, 2, x_80); +x_82 = lean_array_push(x_7, x_81); +x_83 = lean_box(0); +x_84 = lean_box(x_6); +x_85 = lean_apply_6(x_12, x_84, x_82, x_83, x_9, x_10, x_63); +return x_85; } else { -lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; -x_87 = lean_ctor_get(x_5, 0); -lean_inc(x_87); -lean_dec(x_5); -x_88 = l_Lean_Elab_nestedExceptionToMessageData___at_Lean_Elab_Command_elabExport___spec__13___closed__1; -x_89 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_89, 0, x_59); -lean_ctor_set(x_89, 1, x_88); -x_90 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__7; -x_91 = lean_array_push(x_90, x_89); -x_92 = lean_array_push(x_91, x_87); -x_93 = l_Array_append___rarg(x_3, x_92); -x_94 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_94, 0, x_68); -lean_ctor_set(x_94, 1, x_69); -lean_ctor_set(x_94, 2, x_93); -x_95 = lean_array_push(x_76, x_94); -x_96 = lean_array_push(x_95, x_71); -x_97 = lean_array_push(x_96, x_73); +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_86 = lean_ctor_get(x_4, 0); +lean_inc(x_86); +lean_dec(x_4); +x_87 = l_Lean_Elab_nestedExceptionToMessageData___at_Lean_Elab_Command_elabExport___spec__13___closed__1; +x_88 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_88, 0, x_58); +lean_ctor_set(x_88, 1, x_87); +x_89 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__9; +x_90 = lean_array_push(x_89, x_88); +x_91 = lean_array_push(x_90, x_86); +x_92 = l_Array_append___rarg(x_66, x_91); +x_93 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_93, 0, x_68); +lean_ctor_set(x_93, 1, x_69); +lean_ctor_set(x_93, 2, x_92); +x_94 = lean_array_push(x_75, x_93); +x_95 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__11; +x_96 = lean_array_push(x_94, x_95); +x_97 = lean_array_push(x_96, x_72); x_98 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_98, 0, x_68); -lean_ctor_set(x_98, 1, x_6); +lean_ctor_set(x_98, 1, x_5); lean_ctor_set(x_98, 2, x_97); -x_99 = lean_array_push(x_8, x_98); +x_99 = lean_array_push(x_7, x_98); x_100 = lean_box(0); -x_101 = lean_box(x_7); -x_102 = lean_apply_6(x_13, x_101, x_99, x_100, x_10, x_11, x_64); +x_101 = lean_box(x_6); +x_102 = lean_apply_6(x_12, x_101, x_99, x_100, x_9, x_10, x_63); return x_102; } } @@ -11417,21 +11436,62 @@ return x_8; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__4(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; lean_object* x_11; +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; x_5 = lean_unsigned_to_nat(3u); x_6 = l_Lean_Syntax_getArg(x_1, x_5); x_7 = l_Lean_Syntax_getOptional_x3f(x_6); lean_dec(x_6); x_8 = l_Lean_Syntax_getArgs(x_2); -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_4); -lean_ctor_set(x_9, 1, x_7); +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_box(0); x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_8); +lean_ctor_set(x_10, 0, x_4); lean_ctor_set(x_10, 1, x_9); -x_11 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_11, 0, x_10); -return x_11; +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_8); +lean_ctor_set(x_11, 1, x_10); +x_12 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_12, 0, x_11); +return x_12; +} +else +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_7); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_4); +lean_ctor_set(x_14, 1, x_7); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_8); +lean_ctor_set(x_15, 1, x_14); +x_16 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_16, 0, x_15); +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_17 = lean_ctor_get(x_7, 0); +lean_inc(x_17); +lean_dec(x_7); +x_18 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_18, 0, x_17); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_4); +lean_ctor_set(x_19, 1, x_18); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_8); +lean_ctor_set(x_20, 1, x_19); +x_21 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_21, 0, x_20); +return x_21; +} +} } } static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___closed__1() { @@ -11478,434 +11538,425 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2(lean_object* x_1, uint8_t x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -uint8_t x_11; -x_11 = lean_usize_dec_lt(x_6, x_5); -if (x_11 == 0) +uint8_t x_10; +x_10 = lean_usize_dec_lt(x_5, x_4); +if (x_10 == 0) { -lean_object* x_12; -lean_dec(x_9); +lean_object* x_11; lean_dec(x_8); -lean_dec(x_3); +lean_dec(x_7); lean_dec(x_1); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_7); -lean_ctor_set(x_12, 1, x_10); -return x_12; +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_6); +lean_ctor_set(x_11, 1, x_9); +return x_11; } else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_31; lean_object* x_67; uint8_t x_68; -x_13 = lean_array_uget(x_4, x_6); -x_23 = lean_ctor_get(x_7, 0); +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_30; lean_object* x_66; uint8_t x_67; +x_12 = lean_array_uget(x_3, x_5); +x_22 = lean_ctor_get(x_6, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_6, 1); lean_inc(x_23); -x_24 = lean_ctor_get(x_7, 1); -lean_inc(x_24); -if (lean_is_exclusive(x_7)) { - lean_ctor_release(x_7, 0); - lean_ctor_release(x_7, 1); - x_25 = x_7; +if (lean_is_exclusive(x_6)) { + lean_ctor_release(x_6, 0); + lean_ctor_release(x_6, 1); + x_24 = x_6; } else { - lean_dec_ref(x_7); - x_25 = lean_box(0); + lean_dec_ref(x_6); + x_24 = lean_box(0); } -x_67 = l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_typelessBinder_x3f___closed__4; -lean_inc(x_13); -x_68 = l_Lean_Syntax_isOfKind(x_13, x_67); -if (x_68 == 0) +x_66 = l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_typelessBinder_x3f___closed__4; +lean_inc(x_12); +x_67 = l_Lean_Syntax_isOfKind(x_12, x_66); +if (x_67 == 0) { -lean_object* x_69; uint8_t x_70; -x_69 = l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_typelessBinder_x3f___closed__6; -lean_inc(x_13); -x_70 = l_Lean_Syntax_isOfKind(x_13, x_69); -if (x_70 == 0) +lean_object* x_68; uint8_t x_69; +x_68 = l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_typelessBinder_x3f___closed__6; +lean_inc(x_12); +x_69 = l_Lean_Syntax_isOfKind(x_12, x_68); +if (x_69 == 0) { -lean_object* x_71; uint8_t x_72; -x_71 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___closed__4; -lean_inc(x_13); -x_72 = l_Lean_Syntax_isOfKind(x_13, x_71); -if (x_72 == 0) +lean_object* x_70; uint8_t x_71; +x_70 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___closed__4; +lean_inc(x_12); +x_71 = l_Lean_Syntax_isOfKind(x_12, x_70); +if (x_71 == 0) { -lean_object* x_73; -x_73 = lean_box(0); -x_26 = x_73; -goto block_30; +lean_object* x_72; +x_72 = lean_box(0); +x_25 = x_72; +goto block_29; } else { -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; -x_74 = lean_unsigned_to_nat(1u); -x_75 = l_Lean_Syntax_getArg(x_13, x_74); -x_76 = l_Lean_nullKind; -x_77 = lean_unsigned_to_nat(2u); -lean_inc(x_75); -x_78 = l_Lean_Syntax_isNodeOf(x_75, x_76, x_77); -if (x_78 == 0) +lean_object* x_73; lean_object* x_74; lean_object* x_75; uint8_t x_76; +x_73 = lean_unsigned_to_nat(1u); +x_74 = l_Lean_Syntax_getArg(x_12, x_73); +x_75 = lean_unsigned_to_nat(2u); +lean_inc(x_74); +x_76 = l_Lean_Syntax_matchesNull(x_74, x_75); +if (x_76 == 0) { -lean_object* x_79; -lean_dec(x_75); -x_79 = lean_box(0); -x_26 = x_79; -goto block_30; +lean_object* x_77; +lean_dec(x_74); +x_77 = lean_box(0); +x_25 = x_77; +goto block_29; } else { -lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; -lean_dec(x_25); -x_80 = lean_unsigned_to_nat(0u); -x_81 = l_Lean_Syntax_getArg(x_75, x_80); -lean_dec(x_75); -x_82 = l_Lean_Syntax_getArg(x_13, x_77); -x_83 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___closed__5; -x_84 = lean_array_push(x_83, x_81); -x_85 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_85, 0, x_82); -x_86 = lean_box(0); -x_87 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_87, 0, x_85); -lean_ctor_set(x_87, 1, x_86); -x_88 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_88, 0, x_84); -lean_ctor_set(x_88, 1, x_87); -x_31 = x_88; -goto block_66; +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +lean_dec(x_24); +x_78 = lean_unsigned_to_nat(0u); +x_79 = l_Lean_Syntax_getArg(x_74, x_78); +lean_dec(x_74); +x_80 = l_Lean_Syntax_getArg(x_12, x_75); +x_81 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___closed__5; +x_82 = lean_array_push(x_81, x_79); +x_83 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_83, 0, x_80); +x_84 = lean_box(0); +x_85 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_85, 0, x_83); +lean_ctor_set(x_85, 1, x_84); +x_86 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_86, 0, x_82); +lean_ctor_set(x_86, 1, x_85); +x_30 = x_86; +goto block_65; } } } else { -lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; uint8_t x_93; -x_89 = lean_unsigned_to_nat(1u); -x_90 = l_Lean_Syntax_getArg(x_13, x_89); -x_91 = lean_unsigned_to_nat(2u); -x_92 = l_Lean_Syntax_getArg(x_13, x_91); -x_93 = l_Lean_Syntax_isNone(x_92); -if (x_93 == 0) +lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; uint8_t x_91; +x_87 = lean_unsigned_to_nat(1u); +x_88 = l_Lean_Syntax_getArg(x_12, x_87); +x_89 = lean_unsigned_to_nat(2u); +x_90 = l_Lean_Syntax_getArg(x_12, x_89); +x_91 = l_Lean_Syntax_isNone(x_90); +if (x_91 == 0) { -lean_object* x_94; uint8_t x_95; -x_94 = l_Lean_nullKind; -lean_inc(x_92); -x_95 = l_Lean_Syntax_isNodeOf(x_92, x_94, x_91); -if (x_95 == 0) +uint8_t x_92; +lean_inc(x_90); +x_92 = l_Lean_Syntax_matchesNull(x_90, x_89); +if (x_92 == 0) { -lean_object* x_96; -lean_dec(x_92); +lean_object* x_93; lean_dec(x_90); -x_96 = lean_box(0); -x_26 = x_96; -goto block_30; +lean_dec(x_88); +x_93 = lean_box(0); +x_25 = x_93; +goto block_29; } else { -lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; -lean_dec(x_25); -x_97 = l_Lean_Syntax_getArg(x_92, x_89); -lean_dec(x_92); -x_98 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_98, 0, x_97); -x_99 = lean_box(0); -x_100 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__3(x_90, x_99, x_98); +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; +lean_dec(x_24); +x_94 = l_Lean_Syntax_getArg(x_90, x_87); lean_dec(x_90); -x_101 = lean_ctor_get(x_100, 0); -lean_inc(x_101); -lean_dec(x_100); -x_31 = x_101; -goto block_66; +x_95 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_95, 0, x_94); +x_96 = lean_box(0); +x_97 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__3(x_88, x_96, x_95); +lean_dec(x_88); +x_98 = lean_ctor_get(x_97, 0); +lean_inc(x_98); +lean_dec(x_97); +x_30 = x_98; +goto block_65; } } else { -lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; -lean_dec(x_92); -lean_dec(x_25); -x_102 = lean_box(0); -x_103 = lean_box(0); -x_104 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__3(x_90, x_103, x_102); +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_dec(x_90); -x_105 = lean_ctor_get(x_104, 0); -lean_inc(x_105); -lean_dec(x_104); -x_31 = x_105; -goto block_66; +lean_dec(x_24); +x_99 = lean_box(0); +x_100 = lean_box(0); +x_101 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__3(x_88, x_100, x_99); +lean_dec(x_88); +x_102 = lean_ctor_get(x_101, 0); +lean_inc(x_102); +lean_dec(x_101); +x_30 = x_102; +goto block_65; } } } else { -lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; uint8_t x_110; -x_106 = lean_unsigned_to_nat(1u); -x_107 = l_Lean_Syntax_getArg(x_13, x_106); -x_108 = lean_unsigned_to_nat(2u); -x_109 = l_Lean_Syntax_getArg(x_13, x_108); -x_110 = l_Lean_Syntax_isNone(x_109); -if (x_110 == 0) +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; uint8_t x_107; +x_103 = lean_unsigned_to_nat(1u); +x_104 = l_Lean_Syntax_getArg(x_12, x_103); +x_105 = lean_unsigned_to_nat(2u); +x_106 = l_Lean_Syntax_getArg(x_12, x_105); +x_107 = l_Lean_Syntax_isNone(x_106); +if (x_107 == 0) { -lean_object* x_111; uint8_t x_112; -x_111 = l_Lean_nullKind; -lean_inc(x_109); -x_112 = l_Lean_Syntax_isNodeOf(x_109, x_111, x_108); -if (x_112 == 0) +uint8_t x_108; +lean_inc(x_106); +x_108 = l_Lean_Syntax_matchesNull(x_106, x_105); +if (x_108 == 0) { -lean_object* x_113; -lean_dec(x_109); -lean_dec(x_107); -x_113 = lean_box(0); -x_26 = x_113; -goto block_30; +lean_object* x_109; +lean_dec(x_106); +lean_dec(x_104); +x_109 = lean_box(0); +x_25 = x_109; +goto block_29; } else { -lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; -lean_dec(x_25); -x_114 = l_Lean_Syntax_getArg(x_109, x_106); -lean_dec(x_109); -x_115 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_115, 0, x_114); -x_116 = lean_box(0); -x_117 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__4(x_13, x_107, x_116, x_115); -lean_dec(x_107); -x_118 = lean_ctor_get(x_117, 0); -lean_inc(x_118); -lean_dec(x_117); -x_31 = x_118; -goto block_66; +lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; +lean_dec(x_24); +x_110 = l_Lean_Syntax_getArg(x_106, x_103); +lean_dec(x_106); +x_111 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_111, 0, x_110); +x_112 = lean_box(0); +x_113 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__4(x_12, x_104, x_112, x_111); +lean_dec(x_104); +x_114 = lean_ctor_get(x_113, 0); +lean_inc(x_114); +lean_dec(x_113); +x_30 = x_114; +goto block_65; } } else { -lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; -lean_dec(x_109); -lean_dec(x_25); -x_119 = lean_box(0); -x_120 = lean_box(0); -x_121 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__4(x_13, x_107, x_120, x_119); -lean_dec(x_107); -x_122 = lean_ctor_get(x_121, 0); -lean_inc(x_122); -lean_dec(x_121); -x_31 = x_122; -goto block_66; +lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; +lean_dec(x_106); +lean_dec(x_24); +x_115 = lean_box(0); +x_116 = lean_box(0); +x_117 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__4(x_12, x_104, x_116, x_115); +lean_dec(x_104); +x_118 = lean_ctor_get(x_117, 0); +lean_inc(x_118); +lean_dec(x_117); +x_30 = x_118; +goto block_65; } } -block_22: +block_21: { -if (lean_obj_tag(x_14) == 0) +if (lean_obj_tag(x_13) == 0) { -lean_object* x_16; lean_object* x_17; -lean_dec(x_9); +lean_object* x_15; lean_object* x_16; lean_dec(x_8); -lean_dec(x_3); +lean_dec(x_7); lean_dec(x_1); -x_16 = lean_ctor_get(x_14, 0); -lean_inc(x_16); -lean_dec(x_14); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_15); -return x_17; +x_15 = lean_ctor_get(x_13, 0); +lean_inc(x_15); +lean_dec(x_13); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_14); +return x_16; } else { -lean_object* x_18; size_t x_19; size_t x_20; -x_18 = lean_ctor_get(x_14, 0); -lean_inc(x_18); -lean_dec(x_14); -x_19 = 1; -x_20 = lean_usize_add(x_6, x_19); -x_6 = x_20; -x_7 = x_18; -x_10 = x_15; +lean_object* x_17; size_t x_18; size_t x_19; +x_17 = lean_ctor_get(x_13, 0); +lean_inc(x_17); +lean_dec(x_13); +x_18 = 1; +x_19 = lean_usize_add(x_5, x_18); +x_5 = x_19; +x_6 = x_17; +x_9 = x_14; goto _start; } } -block_30: +block_29: { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -lean_dec(x_26); -x_27 = lean_array_push(x_24, x_13); -if (lean_is_scalar(x_25)) { - x_28 = lean_alloc_ctor(0, 2, 0); +lean_object* x_26; lean_object* x_27; lean_object* x_28; +lean_dec(x_25); +x_26 = lean_array_push(x_23, x_12); +if (lean_is_scalar(x_24)) { + x_27 = lean_alloc_ctor(0, 2, 0); } else { - x_28 = x_25; + x_27 = x_24; } -lean_ctor_set(x_28, 0, x_23); -lean_ctor_set(x_28, 1, x_27); -x_29 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_29, 0, x_28); -x_14 = x_29; -x_15 = x_10; -goto block_22; +lean_ctor_set(x_27, 0, x_22); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_28, 0, x_27); +x_13 = x_28; +x_14 = x_9; +goto block_21; } -block_66: +block_65: { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_32 = lean_ctor_get(x_31, 1); +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 0); lean_inc(x_32); +lean_dec(x_30); x_33 = lean_ctor_get(x_31, 0); lean_inc(x_33); -lean_dec(x_31); -x_34 = lean_ctor_get(x_32, 0); +x_34 = lean_ctor_get(x_31, 1); lean_inc(x_34); -x_35 = lean_ctor_get(x_32, 1); -lean_inc(x_35); -lean_dec(x_32); -lean_inc(x_9); +lean_dec(x_31); lean_inc(x_8); +lean_inc(x_7); lean_inc(x_1); -lean_inc(x_33); -x_36 = l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_matchBinderNames(x_33, x_1, x_8, x_9, x_10); -if (lean_obj_tag(x_36) == 0) +lean_inc(x_32); +x_35 = l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_matchBinderNames(x_32, x_1, x_7, x_8, x_9); +if (lean_obj_tag(x_35) == 0) { -lean_object* x_37; uint8_t x_38; -x_37 = lean_ctor_get(x_36, 0); -lean_inc(x_37); -x_38 = lean_unbox(x_37); -lean_dec(x_37); -if (x_38 == 0) +lean_object* x_36; uint8_t x_37; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +x_37 = lean_unbox(x_36); +lean_dec(x_36); +if (x_37 == 0) { -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; -lean_dec(x_35); +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_dec(x_34); lean_dec(x_33); -x_39 = lean_ctor_get(x_36, 1); -lean_inc(x_39); -lean_dec(x_36); -x_40 = lean_array_push(x_24, x_13); -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_23); -lean_ctor_set(x_41, 1, x_40); -x_42 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_42, 0, x_41); -x_14 = x_42; -x_15 = x_39; -goto block_22; +lean_dec(x_32); +x_38 = lean_ctor_get(x_35, 1); +lean_inc(x_38); +lean_dec(x_35); +x_39 = lean_array_push(x_23, x_12); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_22); +lean_ctor_set(x_40, 1, x_39); +x_41 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_41, 0, x_40); +x_13 = x_41; +x_14 = x_38; +goto block_21; } else { -lean_dec(x_13); -if (lean_obj_tag(x_35) == 0) +lean_dec(x_12); +if (lean_obj_tag(x_34) == 0) { -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; lean_object* x_48; -x_43 = lean_ctor_get(x_36, 1); -lean_inc(x_43); -lean_dec(x_36); -x_44 = l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_typelessBinder_x3f___closed__2; -x_45 = l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_typelessBinder_x3f___closed__4; -x_46 = lean_box(0); -x_47 = lean_unbox(x_23); -lean_dec(x_23); -lean_inc(x_9); +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; lean_object* x_47; +x_42 = lean_ctor_get(x_35, 1); +lean_inc(x_42); +lean_dec(x_35); +x_43 = l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_typelessBinder_x3f___closed__2; +x_44 = l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_typelessBinder_x3f___closed__4; +x_45 = lean_box(0); +x_46 = lean_unbox(x_22); +lean_dec(x_22); lean_inc(x_8); -lean_inc(x_3); -x_48 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2(x_2, x_44, x_3, x_33, x_34, x_45, x_47, x_24, x_46, x_8, x_9, x_43); -if (lean_obj_tag(x_48) == 0) +lean_inc(x_7); +x_47 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2(x_2, x_43, x_32, x_33, x_44, x_46, x_23, x_45, x_7, x_8, x_42); +if (lean_obj_tag(x_47) == 0) { -lean_object* x_49; lean_object* x_50; -x_49 = lean_ctor_get(x_48, 0); +lean_object* x_48; lean_object* x_49; +x_48 = lean_ctor_get(x_47, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_47, 1); lean_inc(x_49); -x_50 = lean_ctor_get(x_48, 1); -lean_inc(x_50); -lean_dec(x_48); +lean_dec(x_47); +x_13 = x_48; x_14 = x_49; -x_15 = x_50; -goto block_22; +goto block_21; } else { -uint8_t x_51; -lean_dec(x_9); +uint8_t x_50; lean_dec(x_8); -lean_dec(x_3); +lean_dec(x_7); lean_dec(x_1); -x_51 = !lean_is_exclusive(x_48); -if (x_51 == 0) +x_50 = !lean_is_exclusive(x_47); +if (x_50 == 0) { -return x_48; +return x_47; } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_48, 0); -x_53 = lean_ctor_get(x_48, 1); -lean_inc(x_53); +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_47, 0); +x_52 = lean_ctor_get(x_47, 1); lean_inc(x_52); -lean_dec(x_48); -x_54 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_54, 0, x_52); -lean_ctor_set(x_54, 1, x_53); -return x_54; +lean_inc(x_51); +lean_dec(x_47); +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_52); +return x_53; } } } else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; uint8_t x_58; -lean_dec(x_35); +lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; lean_dec(x_34); lean_dec(x_33); -lean_dec(x_24); +lean_dec(x_32); lean_dec(x_23); -lean_dec(x_3); +lean_dec(x_22); lean_dec(x_1); -x_55 = lean_ctor_get(x_36, 1); -lean_inc(x_55); -lean_dec(x_36); -x_56 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___closed__2; -x_57 = l_Lean_throwError___at_Lean_Elab_Command_expandDeclId___spec__4(x_56, x_8, x_9, x_55); -lean_dec(x_9); -x_58 = !lean_is_exclusive(x_57); -if (x_58 == 0) +x_54 = lean_ctor_get(x_35, 1); +lean_inc(x_54); +lean_dec(x_35); +x_55 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___closed__2; +x_56 = l_Lean_throwError___at_Lean_Elab_Command_expandDeclId___spec__4(x_55, x_7, x_8, x_54); +lean_dec(x_8); +x_57 = !lean_is_exclusive(x_56); +if (x_57 == 0) { -return x_57; +return x_56; } else { -lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_59 = lean_ctor_get(x_57, 0); -x_60 = lean_ctor_get(x_57, 1); -lean_inc(x_60); +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_56, 0); +x_59 = lean_ctor_get(x_56, 1); lean_inc(x_59); -lean_dec(x_57); -x_61 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_61, 0, x_59); -lean_ctor_set(x_61, 1, x_60); -return x_61; +lean_inc(x_58); +lean_dec(x_56); +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +return x_60; } } } } else { -uint8_t x_62; -lean_dec(x_35); +uint8_t x_61; lean_dec(x_34); lean_dec(x_33); -lean_dec(x_24); +lean_dec(x_32); lean_dec(x_23); -lean_dec(x_13); -lean_dec(x_9); +lean_dec(x_22); +lean_dec(x_12); lean_dec(x_8); -lean_dec(x_3); +lean_dec(x_7); lean_dec(x_1); -x_62 = !lean_is_exclusive(x_36); -if (x_62 == 0) +x_61 = !lean_is_exclusive(x_35); +if (x_61 == 0) { -return x_36; +return x_35; } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_63 = lean_ctor_get(x_36, 0); -x_64 = lean_ctor_get(x_36, 1); -lean_inc(x_64); +lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_62 = lean_ctor_get(x_35, 0); +x_63 = lean_ctor_get(x_35, 1); lean_inc(x_63); -lean_dec(x_36); -x_65 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_65, 0, x_63); -lean_ctor_set(x_65, 1, x_64); -return x_65; +lean_inc(x_62); +lean_dec(x_35); +x_64 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_64, 0, x_62); +lean_ctor_set(x_64, 1, x_63); +return x_64; } } } @@ -11987,7 +12038,7 @@ return x_8; } else { -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; lean_object* x_16; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_object* x_22; +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; lean_object* x_16; size_t x_17; size_t x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; x_9 = lean_ctor_get(x_5, 0); lean_inc(x_9); lean_dec(x_5); @@ -12009,116 +12060,115 @@ x_16 = lean_array_get_size(x_15); x_17 = lean_usize_of_nat(x_16); lean_dec(x_16); x_18 = 0; -x_19 = l___private_Lean_Elab_Open_0__Lean_Elab_OpenDecl_resolveNameUsingNamespacesCore___at_Lean_Elab_Command_elabExport___spec__3___closed__1; -x_20 = l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___closed__1; -x_21 = lean_unbox(x_11); +x_19 = l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___closed__1; +x_20 = lean_unbox(x_11); lean_dec(x_11); lean_inc(x_3); lean_inc(x_2); -x_22 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2(x_10, x_21, x_19, x_15, x_17, x_18, x_20, x_2, x_3, x_14); +x_21 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2(x_10, x_20, x_15, x_17, x_18, x_19, x_2, x_3, x_14); lean_dec(x_15); -if (lean_obj_tag(x_22) == 0) +if (lean_obj_tag(x_21) == 0) { -lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_unbox(x_24); -lean_dec(x_24); -if (x_25 == 0) -{ -uint8_t x_26; +lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_unbox(x_23); lean_dec(x_23); +if (x_24 == 0) +{ +uint8_t x_25; +lean_dec(x_22); lean_dec(x_3); lean_dec(x_2); -x_26 = !lean_is_exclusive(x_22); -if (x_26 == 0) +x_25 = !lean_is_exclusive(x_21); +if (x_25 == 0) { -lean_object* x_27; uint8_t x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_22, 0); -lean_dec(x_27); -x_28 = 0; -x_29 = lean_box(x_28); -lean_ctor_set(x_22, 0, x_29); -return x_22; +lean_object* x_26; uint8_t x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_21, 0); +lean_dec(x_26); +x_27 = 0; +x_28 = lean_box(x_27); +lean_ctor_set(x_21, 0, x_28); +return x_21; } else { -lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; -x_30 = lean_ctor_get(x_22, 1); -lean_inc(x_30); -lean_dec(x_22); -x_31 = 0; -x_32 = lean_box(x_31); -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_30); -return x_33; +lean_object* x_29; uint8_t x_30; lean_object* x_31; lean_object* x_32; +x_29 = lean_ctor_get(x_21, 1); +lean_inc(x_29); +lean_dec(x_21); +x_30 = 0; +x_31 = lean_box(x_30); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_29); +return x_32; } } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_33 = lean_ctor_get(x_21, 1); +lean_inc(x_33); +lean_dec(x_21); x_34 = lean_ctor_get(x_22, 1); lean_inc(x_34); lean_dec(x_22); -x_35 = lean_ctor_get(x_23, 1); -lean_inc(x_35); -lean_dec(x_23); -x_36 = lean_alloc_closure((void*)(l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___lambda__1), 2, 1); -lean_closure_set(x_36, 0, x_35); -x_37 = l_Lean_Elab_Command_modifyScope(x_36, x_2, x_3, x_34); +x_35 = lean_alloc_closure((void*)(l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___lambda__1), 2, 1); +lean_closure_set(x_35, 0, x_34); +x_36 = l_Lean_Elab_Command_modifyScope(x_35, x_2, x_3, x_33); lean_dec(x_3); lean_dec(x_2); -x_38 = !lean_is_exclusive(x_37); -if (x_38 == 0) +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) { -lean_object* x_39; uint8_t x_40; lean_object* x_41; -x_39 = lean_ctor_get(x_37, 0); -lean_dec(x_39); -x_40 = 1; -x_41 = lean_box(x_40); -lean_ctor_set(x_37, 0, x_41); -return x_37; +lean_object* x_38; uint8_t x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_36, 0); +lean_dec(x_38); +x_39 = 1; +x_40 = lean_box(x_39); +lean_ctor_set(x_36, 0, x_40); +return x_36; } else { -lean_object* x_42; uint8_t x_43; lean_object* x_44; lean_object* x_45; -x_42 = lean_ctor_get(x_37, 1); -lean_inc(x_42); -lean_dec(x_37); -x_43 = 1; -x_44 = lean_box(x_43); -x_45 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_45, 0, x_44); -lean_ctor_set(x_45, 1, x_42); -return x_45; +lean_object* x_41; uint8_t x_42; lean_object* x_43; lean_object* x_44; +x_41 = lean_ctor_get(x_36, 1); +lean_inc(x_41); +lean_dec(x_36); +x_42 = 1; +x_43 = lean_box(x_42); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_41); +return x_44; } } } else { -uint8_t x_46; +uint8_t x_45; lean_dec(x_3); lean_dec(x_2); -x_46 = !lean_is_exclusive(x_22); -if (x_46 == 0) +x_45 = !lean_is_exclusive(x_21); +if (x_45 == 0) { -return x_22; +return x_21; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_22, 0); -x_48 = lean_ctor_get(x_22, 1); -lean_inc(x_48); +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_21, 0); +x_47 = lean_ctor_get(x_21, 1); lean_inc(x_47); -lean_dec(x_22); -x_49 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_49, 0, x_47); -lean_ctor_set(x_49, 1, x_48); -return x_49; +lean_inc(x_46); +lean_dec(x_21); +x_48 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_48, 0, x_46); +lean_ctor_set(x_48, 1, x_47); +return x_48; } } } @@ -12147,17 +12197,17 @@ lean_dec(x_3); return x_8; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___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, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___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, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -uint8_t x_13; uint8_t x_14; lean_object* x_15; -x_13 = lean_unbox(x_1); +uint8_t x_12; uint8_t x_13; lean_object* x_14; +x_12 = lean_unbox(x_1); lean_dec(x_1); -x_14 = lean_unbox(x_7); -lean_dec(x_7); -x_15 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2(x_13, x_2, x_3, x_4, x_5, x_6, x_14, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_9); -return x_15; +x_13 = lean_unbox(x_6); +lean_dec(x_6); +x_14 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2(x_12, x_2, x_3, x_4, x_5, x_13, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_8); +return x_14; } } LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { @@ -12181,19 +12231,19 @@ lean_dec(x_1); return x_5; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__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, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__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, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -uint8_t x_11; size_t x_12; size_t x_13; lean_object* x_14; -x_11 = lean_unbox(x_2); +uint8_t x_10; size_t x_11; size_t x_12; lean_object* x_13; +x_10 = lean_unbox(x_2); lean_dec(x_2); +x_11 = lean_unbox_usize(x_4); +lean_dec(x_4); x_12 = lean_unbox_usize(x_5); lean_dec(x_5); -x_13 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_14 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2(x_1, x_11, x_3, x_4, x_12, x_13, x_7, x_8, x_9, x_10); -lean_dec(x_4); -return x_14; +x_13 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2(x_1, x_10, x_3, x_11, x_12, x_6, x_7, x_8, x_9); +lean_dec(x_3); +return x_13; } } LEAN_EXPORT lean_object* l_Lean_MonadQuotation_addMacroScope___at_Lean_Elab_Command_elabVariable___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { @@ -15802,7 +15852,7 @@ x_17 = l_Lean_Elab_Command_elabCheckFailure___closed__3; x_18 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_18, 0, x_11); lean_ctor_set(x_18, 1, x_17); -x_19 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__7; +x_19 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__9; x_20 = lean_array_push(x_19, x_18); x_21 = lean_array_push(x_20, x_9); x_22 = lean_box(2); @@ -16433,7 +16483,7 @@ x_24 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_24, 0, x_15); lean_ctor_set(x_24, 1, x_3); x_25 = l_Lean_mkConst(x_23, x_24); -x_26 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__10; +x_26 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__13; x_27 = lean_array_push(x_26, x_12); x_28 = lean_array_push(x_27, x_20); lean_inc(x_4); @@ -16442,7 +16492,7 @@ lean_inc(x_5); x_30 = lean_array_push(x_29, x_5); x_31 = lean_array_push(x_30, x_1); x_32 = l_Lean_mkAppN(x_25, x_31); -x_33 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__7; +x_33 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__9; x_34 = lean_array_push(x_33, x_4); x_35 = lean_array_push(x_34, x_5); x_36 = 0; @@ -19891,104 +19941,103 @@ x_19 = lean_erase_macro_scopes(x_18); x_20 = l_Lean_Syntax_isStrLit_x3f(x_2); if (lean_obj_tag(x_20) == 0) { -lean_object* x_21; lean_object* x_22; -x_21 = l_Lean_numLitKind; -x_22 = l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(x_21, x_2); -if (lean_obj_tag(x_22) == 0) +lean_object* x_21; +x_21 = l_Lean_Syntax_isNatLit_x3f(x_2); +if (lean_obj_tag(x_21) == 0) { if (lean_obj_tag(x_2) == 2) { -lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_23 = lean_ctor_get(x_2, 1); -lean_inc(x_23); -x_24 = l_Lean_Elab_elabSetOption___at_Lean_Elab_Command_elabSetOption___spec__1___closed__3; -x_25 = lean_string_dec_eq(x_23, x_24); -if (x_25 == 0) +lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_22 = lean_ctor_get(x_2, 1); +lean_inc(x_22); +x_23 = l_Lean_Elab_elabSetOption___at_Lean_Elab_Command_elabSetOption___spec__1___closed__3; +x_24 = lean_string_dec_eq(x_22, x_23); +if (x_24 == 0) { -lean_object* x_26; uint8_t x_27; -x_26 = l_Lean_Elab_elabSetOption___at_Lean_Elab_Command_elabSetOption___spec__1___closed__4; -x_27 = lean_string_dec_eq(x_23, x_26); -lean_dec(x_23); -if (x_27 == 0) +lean_object* x_25; uint8_t x_26; +x_25 = l_Lean_Elab_elabSetOption___at_Lean_Elab_Command_elabSetOption___spec__1___closed__4; +x_26 = lean_string_dec_eq(x_22, x_25); +lean_dec(x_22); +if (x_26 == 0) { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_dec(x_19); -x_28 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_28, 0, x_2); -x_29 = l_Lean_Elab_elabSetOption___at_Lean_Elab_Command_elabSetOption___spec__1___closed__2; -x_30 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_28); -x_31 = l_Lean_Elab_Command_elabModuleDoc___closed__4; -x_32 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -x_33 = l_Lean_throwError___at_Lean_Elab_Command_elabSetOption___spec__2(x_32, x_3, x_4, x_17); +x_27 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_27, 0, x_2); +x_28 = l_Lean_Elab_elabSetOption___at_Lean_Elab_Command_elabSetOption___spec__1___closed__2; +x_29 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_27); +x_30 = l_Lean_Elab_Command_elabModuleDoc___closed__4; +x_31 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +x_32 = l_Lean_throwError___at_Lean_Elab_Command_elabSetOption___spec__2(x_31, x_3, x_4, x_17); lean_dec(x_4); -return x_33; +return x_32; } else { -lean_object* x_34; lean_object* x_35; +lean_object* x_33; lean_object* x_34; lean_dec(x_2); -x_34 = l_Lean_Elab_elabSetOption___at_Lean_Elab_Command_elabSetOption___spec__1___closed__5; -x_35 = l_Lean_Elab_elabSetOption_setOption___at_Lean_Elab_Command_elabSetOption___spec__3(x_19, x_34, x_3, x_4, x_17); -return x_35; +x_33 = l_Lean_Elab_elabSetOption___at_Lean_Elab_Command_elabSetOption___spec__1___closed__5; +x_34 = l_Lean_Elab_elabSetOption_setOption___at_Lean_Elab_Command_elabSetOption___spec__3(x_19, x_33, x_3, x_4, x_17); +return x_34; } } else { -lean_object* x_36; lean_object* x_37; -lean_dec(x_23); +lean_object* x_35; lean_object* x_36; +lean_dec(x_22); lean_dec(x_2); -x_36 = l_Lean_Elab_elabSetOption___at_Lean_Elab_Command_elabSetOption___spec__1___closed__6; -x_37 = l_Lean_Elab_elabSetOption_setOption___at_Lean_Elab_Command_elabSetOption___spec__3(x_19, x_36, x_3, x_4, x_17); -return x_37; +x_35 = l_Lean_Elab_elabSetOption___at_Lean_Elab_Command_elabSetOption___spec__1___closed__6; +x_36 = l_Lean_Elab_elabSetOption_setOption___at_Lean_Elab_Command_elabSetOption___spec__3(x_19, x_35, x_3, x_4, x_17); +return x_36; } } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +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_dec(x_19); -x_38 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_38, 0, x_2); -x_39 = l_Lean_Elab_elabSetOption___at_Lean_Elab_Command_elabSetOption___spec__1___closed__2; -x_40 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_38); -x_41 = l_Lean_Elab_Command_elabModuleDoc___closed__4; -x_42 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_42, 0, x_40); -lean_ctor_set(x_42, 1, x_41); -x_43 = l_Lean_throwError___at_Lean_Elab_Command_elabSetOption___spec__2(x_42, x_3, x_4, x_17); +x_37 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_37, 0, x_2); +x_38 = l_Lean_Elab_elabSetOption___at_Lean_Elab_Command_elabSetOption___spec__1___closed__2; +x_39 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_37); +x_40 = l_Lean_Elab_Command_elabModuleDoc___closed__4; +x_41 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +x_42 = l_Lean_throwError___at_Lean_Elab_Command_elabSetOption___spec__2(x_41, x_3, x_4, x_17); lean_dec(x_4); -return x_43; +return x_42; } } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_dec(x_2); -x_44 = lean_ctor_get(x_22, 0); -lean_inc(x_44); -lean_dec(x_22); -x_45 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_45, 0, x_44); -x_46 = l_Lean_Elab_elabSetOption_setOption___at_Lean_Elab_Command_elabSetOption___spec__3(x_19, x_45, x_3, x_4, x_17); -return x_46; +x_43 = lean_ctor_get(x_21, 0); +lean_inc(x_43); +lean_dec(x_21); +x_44 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_44, 0, x_43); +x_45 = l_Lean_Elab_elabSetOption_setOption___at_Lean_Elab_Command_elabSetOption___spec__3(x_19, x_44, x_3, x_4, x_17); +return x_45; } } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_dec(x_2); -x_47 = lean_ctor_get(x_20, 0); -lean_inc(x_47); +x_46 = lean_ctor_get(x_20, 0); +lean_inc(x_46); lean_dec(x_20); -x_48 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_48, 0, x_47); -x_49 = l_Lean_Elab_elabSetOption_setOption___at_Lean_Elab_Command_elabSetOption___spec__3(x_19, x_48, x_3, x_4, x_17); -return x_49; +x_47 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_47, 0, x_46); +x_48 = l_Lean_Elab_elabSetOption_setOption___at_Lean_Elab_Command_elabSetOption___spec__3(x_19, x_47, x_3, x_4, x_17); +return x_48; } } } @@ -20348,135 +20397,149 @@ return x_4; static lean_object* _init_l_Lean_Elab_Command_expandInCmd___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = lean_box(2); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__4; -x_3 = l___private_Lean_Elab_Open_0__Lean_Elab_OpenDecl_resolveNameUsingNamespacesCore___at_Lean_Elab_Command_elabExport___spec__3___closed__1; -x_4 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_4, 0, x_1); -lean_ctor_set(x_4, 1, x_2); -lean_ctor_set(x_4, 2, x_3); -return x_4; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("in", 2); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Command_expandInCmd___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___regBuiltin_Lean_Elab_Command_elabModuleDoc___closed__6; +x_2 = l_Lean_Elab_Command_expandInCmd___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; } } LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandInCmd(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -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_unsigned_to_nat(0u); -x_5 = l_Lean_Syntax_getArg(x_1, x_4); -x_6 = lean_unsigned_to_nat(2u); -x_7 = l_Lean_Syntax_getArg(x_1, x_6); -x_8 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_9 = !lean_is_exclusive(x_8); -if (x_9 == 0) +lean_object* x_4; uint8_t x_5; +x_4 = l_Lean_Elab_Command_expandInCmd___closed__2; +lean_inc(x_1); +x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); +if (x_5 == 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; 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; lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_10 = lean_ctor_get(x_8, 0); -x_11 = l_Lean_Elab_Command_elabSection___closed__1; -lean_inc(x_10); -x_12 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_12, 0, x_10); -lean_ctor_set(x_12, 1, x_11); -x_13 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__7; -x_14 = lean_array_push(x_13, x_12); -x_15 = l_Lean_Elab_Command_expandInCmd___closed__1; -x_16 = lean_array_push(x_14, x_15); -x_17 = lean_box(2); -x_18 = l_Lean_Elab_Command_elabSection___closed__2; -x_19 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set(x_19, 1, x_18); -lean_ctor_set(x_19, 2, x_16); -x_20 = l___regBuiltin_Lean_Elab_Command_elabEnd___closed__1; -x_21 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_21, 0, x_10); -lean_ctor_set(x_21, 1, x_20); -x_22 = lean_array_push(x_13, x_21); -x_23 = lean_array_push(x_22, x_15); -x_24 = l___regBuiltin_Lean_Elab_Command_elabEnd___closed__2; -x_25 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_25, 0, x_17); +lean_object* x_6; lean_object* x_7; +lean_dec(x_2); +lean_dec(x_1); +x_6 = lean_box(1); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_3); +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_8 = lean_unsigned_to_nat(0u); +x_9 = l_Lean_Syntax_getArg(x_1, x_8); +x_10 = lean_unsigned_to_nat(2u); +x_11 = l_Lean_Syntax_getArg(x_1, x_10); +lean_dec(x_1); +x_12 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_13 = !lean_is_exclusive(x_12); +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; 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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_14 = lean_ctor_get(x_12, 0); +x_15 = l_Lean_Elab_Command_elabSection___closed__1; +lean_inc(x_14); +x_16 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +x_17 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__9; +x_18 = lean_array_push(x_17, x_16); +x_19 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__11; +x_20 = lean_array_push(x_18, x_19); +x_21 = lean_box(2); +x_22 = l_Lean_Elab_Command_elabSection___closed__2; +x_23 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +lean_ctor_set(x_23, 2, x_20); +x_24 = l___regBuiltin_Lean_Elab_Command_elabEnd___closed__1; +x_25 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_25, 0, x_14); lean_ctor_set(x_25, 1, x_24); -lean_ctor_set(x_25, 2, x_23); -x_26 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__6; +x_26 = lean_array_push(x_17, x_25); x_27 = lean_array_push(x_26, x_19); -x_28 = lean_array_push(x_27, x_5); -x_29 = lean_array_push(x_28, x_7); -x_30 = lean_array_push(x_29, x_25); -x_31 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__4; -x_32 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_32, 0, x_17); -lean_ctor_set(x_32, 1, x_31); -lean_ctor_set(x_32, 2, x_30); -lean_ctor_set(x_8, 0, x_32); -return x_8; +x_28 = l___regBuiltin_Lean_Elab_Command_elabEnd___closed__2; +x_29 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_29, 0, x_21); +lean_ctor_set(x_29, 1, x_28); +lean_ctor_set(x_29, 2, x_27); +x_30 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__6; +x_31 = lean_array_push(x_30, x_23); +x_32 = lean_array_push(x_31, x_9); +x_33 = lean_array_push(x_32, x_11); +x_34 = lean_array_push(x_33, x_29); +x_35 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__4; +x_36 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_36, 0, x_21); +lean_ctor_set(x_36, 1, x_35); +lean_ctor_set(x_36, 2, x_34); +lean_ctor_set(x_12, 0, x_36); +return x_12; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_33 = lean_ctor_get(x_8, 0); -x_34 = lean_ctor_get(x_8, 1); -lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_8); -x_35 = l_Lean_Elab_Command_elabSection___closed__1; -lean_inc(x_33); -x_36 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_36, 0, x_33); -lean_ctor_set(x_36, 1, x_35); -x_37 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__7; -x_38 = lean_array_push(x_37, x_36); -x_39 = l_Lean_Elab_Command_expandInCmd___closed__1; -x_40 = lean_array_push(x_38, x_39); -x_41 = lean_box(2); -x_42 = l_Lean_Elab_Command_elabSection___closed__2; -x_43 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_43, 0, x_41); -lean_ctor_set(x_43, 1, x_42); -lean_ctor_set(x_43, 2, x_40); -x_44 = l___regBuiltin_Lean_Elab_Command_elabEnd___closed__1; -x_45 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_45, 0, x_33); -lean_ctor_set(x_45, 1, x_44); -x_46 = lean_array_push(x_37, x_45); -x_47 = lean_array_push(x_46, x_39); -x_48 = l___regBuiltin_Lean_Elab_Command_elabEnd___closed__2; -x_49 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_49, 0, x_41); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_37 = lean_ctor_get(x_12, 0); +x_38 = lean_ctor_get(x_12, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_12); +x_39 = l_Lean_Elab_Command_elabSection___closed__1; +lean_inc(x_37); +x_40 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_40, 0, x_37); +lean_ctor_set(x_40, 1, x_39); +x_41 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__9; +x_42 = lean_array_push(x_41, x_40); +x_43 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__11; +x_44 = lean_array_push(x_42, x_43); +x_45 = lean_box(2); +x_46 = l_Lean_Elab_Command_elabSection___closed__2; +x_47 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +lean_ctor_set(x_47, 2, x_44); +x_48 = l___regBuiltin_Lean_Elab_Command_elabEnd___closed__1; +x_49 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_49, 0, x_37); lean_ctor_set(x_49, 1, x_48); -lean_ctor_set(x_49, 2, x_47); -x_50 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__6; +x_50 = lean_array_push(x_41, x_49); x_51 = lean_array_push(x_50, x_43); -x_52 = lean_array_push(x_51, x_5); -x_53 = lean_array_push(x_52, x_7); -x_54 = lean_array_push(x_53, x_49); -x_55 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__4; -x_56 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_56, 0, x_41); -lean_ctor_set(x_56, 1, x_55); -lean_ctor_set(x_56, 2, x_54); -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_34); -return x_57; -} +x_52 = l___regBuiltin_Lean_Elab_Command_elabEnd___closed__2; +x_53 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_53, 0, x_45); +lean_ctor_set(x_53, 1, x_52); +lean_ctor_set(x_53, 2, x_51); +x_54 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__6; +x_55 = lean_array_push(x_54, x_47); +x_56 = lean_array_push(x_55, x_9); +x_57 = lean_array_push(x_56, x_11); +x_58 = lean_array_push(x_57, x_53); +x_59 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__4; +x_60 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_60, 0, x_45); +lean_ctor_set(x_60, 1, x_59); +lean_ctor_set(x_60, 2, x_58); +x_61 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set(x_61, 1, x_38); +return x_61; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandInCmd___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = l_Lean_Elab_Command_expandInCmd(x_1, x_2, x_3); -lean_dec(x_1); -return x_4; } } static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandInCmd___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("in", 2); +x_1 = lean_mk_string_from_bytes("expandInCmd", 11); return x_1; } } @@ -20484,7 +20547,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandInCmd___closed_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Command_elabModuleDoc___closed__6; +x_1 = l___regBuiltin_Lean_Elab_Command_elabModuleDoc___closed__11; x_2 = l___regBuiltin_Lean_Elab_Command_expandInCmd___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -20494,33 +20557,15 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandInCmd___closed_ _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("expandInCmd", 11); -return x_1; -} -} -static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandInCmd___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Command_elabModuleDoc___closed__11; -x_2 = l___regBuiltin_Lean_Elab_Command_expandInCmd___closed__3; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandInCmd___closed__5() { -_start: -{ -lean_object* x_1; x_1 = l_Lean_Elab_macroAttribute; return x_1; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandInCmd___closed__6() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandInCmd___closed__4() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Command_expandInCmd___boxed), 3, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Command_expandInCmd), 3, 0); return x_1; } } @@ -20528,10 +20573,10 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_expandInCmd(lean_objec _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_2 = l___regBuiltin_Lean_Elab_Command_expandInCmd___closed__5; -x_3 = l___regBuiltin_Lean_Elab_Command_expandInCmd___closed__2; -x_4 = l___regBuiltin_Lean_Elab_Command_expandInCmd___closed__4; -x_5 = l___regBuiltin_Lean_Elab_Command_expandInCmd___closed__6; +x_2 = l___regBuiltin_Lean_Elab_Command_expandInCmd___closed__3; +x_3 = l_Lean_Elab_Command_expandInCmd___closed__2; +x_4 = l___regBuiltin_Lean_Elab_Command_expandInCmd___closed__2; +x_5 = l___regBuiltin_Lean_Elab_Command_expandInCmd___closed__4; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); return x_6; } @@ -20552,8 +20597,8 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandInCmd_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(379u); -x_2 = lean_unsigned_to_nat(44u); +x_1 = lean_unsigned_to_nat(378u); +x_2 = lean_unsigned_to_nat(47u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -20567,7 +20612,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Elab_Command_expandInCmd_declRange___closed__1; x_2 = lean_unsigned_to_nat(41u); x_3 = l___regBuiltin_Lean_Elab_Command_expandInCmd_declRange___closed__2; -x_4 = lean_unsigned_to_nat(44u); +x_4 = lean_unsigned_to_nat(47u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -20632,7 +20677,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_expandInCmd_declRange( _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Command_expandInCmd___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Command_expandInCmd___closed__2; x_3 = l___regBuiltin_Lean_Elab_Command_expandInCmd_declRange___closed__7; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; @@ -21131,6 +21176,12 @@ l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Co lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__9); l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__10 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__10(); lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__10); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__11 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__11(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__11); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__12 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__12(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__12); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__13 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__13(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___lambda__2___closed__13); l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___closed__1 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___closed__1(); lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___closed__1); l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___closed__2 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__2___closed__2(); @@ -21469,6 +21520,8 @@ if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); l_Lean_Elab_Command_expandInCmd___closed__1 = _init_l_Lean_Elab_Command_expandInCmd___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_expandInCmd___closed__1); +l_Lean_Elab_Command_expandInCmd___closed__2 = _init_l_Lean_Elab_Command_expandInCmd___closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_expandInCmd___closed__2); l___regBuiltin_Lean_Elab_Command_expandInCmd___closed__1 = _init_l___regBuiltin_Lean_Elab_Command_expandInCmd___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Elab_Command_expandInCmd___closed__1); l___regBuiltin_Lean_Elab_Command_expandInCmd___closed__2 = _init_l___regBuiltin_Lean_Elab_Command_expandInCmd___closed__2(); @@ -21477,10 +21530,6 @@ l___regBuiltin_Lean_Elab_Command_expandInCmd___closed__3 = _init_l___regBuiltin_ lean_mark_persistent(l___regBuiltin_Lean_Elab_Command_expandInCmd___closed__3); l___regBuiltin_Lean_Elab_Command_expandInCmd___closed__4 = _init_l___regBuiltin_Lean_Elab_Command_expandInCmd___closed__4(); lean_mark_persistent(l___regBuiltin_Lean_Elab_Command_expandInCmd___closed__4); -l___regBuiltin_Lean_Elab_Command_expandInCmd___closed__5 = _init_l___regBuiltin_Lean_Elab_Command_expandInCmd___closed__5(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Command_expandInCmd___closed__5); -l___regBuiltin_Lean_Elab_Command_expandInCmd___closed__6 = _init_l___regBuiltin_Lean_Elab_Command_expandInCmd___closed__6(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Command_expandInCmd___closed__6); res = l___regBuiltin_Lean_Elab_Command_expandInCmd(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); diff --git a/stage0/stdlib/Lean/Elab/BuiltinNotation.c b/stage0/stdlib/Lean/Elab/BuiltinNotation.c index f69eb7d5f2d..55169761014 100644 --- a/stage0/stdlib/Lean/Elab/BuiltinNotation.c +++ b/stage0/stdlib/Lean/Elab/BuiltinNotation.c @@ -19,6 +19,7 @@ LEAN_EXPORT lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Term_elabAnonymo static lean_object* l_Lean_Elab_Term_expandDbgTrace___closed__9; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_expandUnreachable_docString(lean_object*); lean_object* l_Lean_mkCIdentFrom(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_expandDbgTrace___closed__27; static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__26; static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_withLocalIdentFor___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_expandAssert_declRange(lean_object*); @@ -55,7 +56,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandSuffices___lambda__1___boxed(lea static lean_object* l_Lean_Elab_Term_expandHave___closed__2; static lean_object* l_Lean_Elab_Term_expandDbgTrace___closed__4; static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__36; -extern lean_object* l_Lean_nullKind; static lean_object* l_Lean_Elab_Term_expandAssert___closed__10; static lean_object* l___regBuiltin_Lean_Elab_Term_elabBorrowed_declRange___closed__5; static lean_object* l_Lean_Elab_Term_expandSuffices___lambda__1___closed__1; @@ -73,6 +73,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabSorry_declRange(lean_ static lean_object* l_Lean_Elab_Term_expandHave___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_elabSubst_declRange___closed__5; static lean_object* l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__1; +static lean_object* l_Lean_Elab_Term_elabPanic___closed__14; static lean_object* l___regBuiltin_Lean_Elab_Term_elabPanic_declRange___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_expandAssert___closed__1; lean_object* l_Array_append___rarg(lean_object*, lean_object*); @@ -108,6 +109,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLeadingParserMacro___lambda__2(lea static lean_object* l_Lean_Elab_Term_elabSorry___closed__10; static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__12; static lean_object* l___regBuiltin_Lean_Elab_Term_expandShow___closed__3; +static lean_object* l_Lean_Elab_Term_expandAssert___closed__12; static lean_object* l_Lean_Elab_Term_expandHave___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLeadingParserMacro(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabSubst___closed__2; @@ -116,13 +118,12 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandCDot_x LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandCDot_x3f_go___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__19; static lean_object* l_Lean_Elab_Term_expandDbgTrace___closed__21; -static lean_object* l___regBuiltin_Lean_Elab_Term_expandAssert___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Term_elabStateRefT_declRange___closed__2; +static lean_object* l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__6; lean_object* l_Lean_Meta_dependsOn(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabSubst_docString___closed__1; static lean_object* l_Lean_Elab_Term_expandHave___closed__9; -static lean_object* l___regBuiltin_Lean_Elab_Term_elabPanic___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Term_elabTrailingParserMacro_declRange___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_elabCoe_declRange___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -160,7 +161,6 @@ static lean_object* l_Lean_Elab_Term_elabSubst___closed__5; static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__1; static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__11; lean_object* lean_string_append(lean_object*, lean_object*); -extern lean_object* l_Lean_interpolatedStrKind; static lean_object* l___regBuiltin_Lean_Elab_Term_elabSubst___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_expandAssert___closed__3; static lean_object* l_Lean_Elab_Term_mkPairs_loop___closed__2; @@ -172,7 +172,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabSorry_declRange___closed__ static lean_object* l___regBuiltin_Lean_Elab_Term_elabStateRefT_declRange___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___lambda__7___boxed(lean_object**); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabParen(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandDbgTrace___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabSorry_declRange___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_expandShow_declRange___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkPairs_loop___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -181,7 +180,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_expandDbgTrace_declRange___clo static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabTParserMacroAux___closed__9; static lean_object* l_Lean_Elab_Term_expandAssert___closed__11; static lean_object* l___regBuiltin_Lean_Elab_Term_expandShow_declRange___closed__3; -LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandDbgTrace___closed__1; static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__39; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabSorry_docString(lean_object*); @@ -200,7 +199,6 @@ static lean_object* l_Lean_Elab_Term_elabAnonymousCtor___closed__6; uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l___regBuiltin_Lean_Elab_Term_elabTrailingParserMacro_declRange___closed__1; lean_object* l_liftExcept___at_Lean_Elab_liftMacroM___spec__1(lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_nameLitKind; static lean_object* l_Lean_Elab_Term_expandDbgTrace___closed__2; static lean_object* l_Lean_Elab_Term_expandShow___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLeadingParserMacro_declRange___closed__3; @@ -244,6 +242,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabTrailingParserMacro_declRa lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabPanic___closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandHave___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_expandDbgTrace___closed__29; static lean_object* l_Lean_Elab_Term_expandShow___closed__2; static lean_object* l_Lean_Elab_Term_expandShow___closed__21; lean_object* l_Lean_throwError___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -259,7 +258,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_expandParen(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_expandHave(lean_object*); static lean_object* l_Lean_Elab_Term_elabStateRefT___lambda__1___closed__5; -lean_object* l_Lean_Syntax_SepArray_getElems___rarg(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabPanic(lean_object*); lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(size_t, size_t, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabAnonymousCtor_declRange___closed__3; @@ -290,6 +288,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_expandShow___closed__2; lean_object* l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandParen_declRange___closed__5; static lean_object* l_Lean_Elab_Term_expandAssert___closed__7; +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabImplicitLambda_loop___spec__1___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandShow_declRange___closed__6; lean_object* l_Lean_Elab_expandMacroImpl_x3f(lean_object*, lean_object*, lean_object*, lean_object*); @@ -297,7 +296,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabCoe_declRange___closed__3; lean_object* lean_st_ref_take(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandHave___closed__1; static lean_object* l_Lean_Elab_Term_expandDbgTrace___closed__6; -extern lean_object* l_Lean_numLitKind; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLeadingParserMacro_declRange___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_elabBorrowed___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_elabSubst_declRange___closed__2; @@ -312,7 +310,7 @@ static lean_object* l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_expandShow_declRange___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_expandDbgTrace_declRange___closed__6; static lean_object* l_Lean_Elab_Term_mkPairs_loop___closed__1; -LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandUnreachable_declRange___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_elabNoindex___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Term_expandAssert_declRange___closed__1; @@ -326,6 +324,8 @@ lean_object* l_Lean_Elab_Term_withMacroExpansion___rarg(lean_object*, lean_objec LEAN_EXPORT lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_isSubstCandidate___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_expandUnreachable_declRange(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabStateRefT___closed__1; +static lean_object* l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__7; +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_elabAnonymousCtor___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandDbgTrace___closed__18; static lean_object* l_Lean_Elab_Term_expandDbgTrace___closed__15; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLeadingParserMacro_declRange___closed__4; @@ -364,7 +364,7 @@ static lean_object* l_Lean_Elab_Term_expandShow___closed__5; static lean_object* l_Lean_Elab_Term_expandShow___closed__15; static lean_object* l_Lean_Elab_Term_mkPairs_loop___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_expandDbgTrace_declRange___closed__1; -LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabParen_declRange___closed__5; LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_expandCDot_x3f_go___spec__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandDbgTrace___closed__3; @@ -386,7 +386,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabStateRefT___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__1___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabParen___closed__1; -lean_object* l___private_Init_Meta_0__Lean_quoteNameMk(lean_object*); static lean_object* l_Lean_Elab_Term_elabLeadingParserMacro___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_elabSorry_docString___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_elabPanic_declRange___closed__1; @@ -396,6 +395,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabCoe_declRange___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_expandDbgTrace_declRange___closed__2; static lean_object* l_Lean_Elab_Term_expandDbgTrace___closed__14; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabParen___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_expandDbgTrace___closed__28; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Term_elabAnonymousCtor___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_back___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -431,7 +431,6 @@ static lean_object* l_Lean_Elab_Term_elabAnonymousCtor___closed__15; static lean_object* l_Lean_Elab_Term_expandDbgTrace___closed__24; static lean_object* l_Lean_Elab_Term_expandDbgTrace___closed__7; static lean_object* l___regBuiltin_Lean_Elab_Term_elabBorrowed_declRange___closed__2; -static lean_object* l___regBuiltin_Lean_Elab_Term_elabPanic___closed__4; lean_object* l_Lean_FileMap_toPosition(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandDbgTrace___closed__1; static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__6; @@ -440,12 +439,12 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLeadingParserMacro___lambda__1___b static lean_object* l___regBuiltin_Lean_Elab_Term_elabTrailingParserMacro_declRange___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabTrailingParserMacro___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabNoindex_declRange___closed__1; -extern lean_object* l_Lean_instInhabitedSyntax; static lean_object* l___regBuiltin_Lean_Elab_Term_elabSorry_declRange___closed__5; static lean_object* l_Lean_Elab_Term_elabStateRefT___closed__1; static lean_object* l_Lean_Elab_Term_expandShow___closed__25; static lean_object* l___regBuiltin_Lean_Elab_Term_elabPanic___closed__2; lean_object* l_Lean_mkSepArray(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_expandAssert___closed__13; static lean_object* l___regBuiltin_Lean_Elab_Term_expandUnreachable_declRange___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Term_elabPanic_declRange___closed__7; static lean_object* l_Lean_Elab_Term_elabSubst___lambda__4___closed__2; @@ -460,7 +459,6 @@ static lean_object* l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; lean_object* lean_whnf(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabSubst_declRange___closed__4; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandAssert___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabPanic_docString___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_isSubstCandidate(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); @@ -470,7 +468,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_expandParen_declRange___closed static lean_object* l_Lean_Elab_Term_elabAnonymousCtor___closed__10; static lean_object* l___regBuiltin_Lean_Elab_Term_expandAssert_declRange___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_expandAssert___closed__4; static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__22; lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__29; @@ -489,9 +486,10 @@ lean_object* l_Lean_Meta_matchEq_x3f(lean_object*, lean_object*, lean_object*, l extern lean_object* l_Lean_Elab_Term_termElabAttribute; static lean_object* l___regBuiltin_Lean_Elab_Term_elabCoe___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_expandDbgTrace_declRange___closed__7; -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_mkNumLit(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabStateRefT___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandParen___closed__11; +lean_object* l_Lean_Syntax_TSepArray_getElems___rarg(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandUnreachable_docString___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_elabSorry___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_elabTrailingParserMacro_declRange___closed__5; @@ -552,6 +550,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLeadingParserMacro___lambda__2___b uint8_t l_Lean_BinderInfo_isExplicit(uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getKind(lean_object*); +lean_object* l_Lean_quoteNameMk(lean_object*); static lean_object* l_Lean_Elab_Term_elabLeadingParserMacro___lambda__1___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_expandHave_declRange___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Term_elabCoe___closed__2; @@ -573,7 +572,6 @@ LEAN_EXPORT lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Term_elabAnonymo static lean_object* l_Lean_Elab_Term_elabLeadingParserMacro___lambda__2___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_elabAnonymousCtor___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_expandUnreachable___closed__5; -static lean_object* l___regBuiltin_Lean_Elab_Term_expandDbgTrace___closed__4; static lean_object* l_Lean_Elab_Term_expandDbgTrace___closed__10; static lean_object* l_Lean_Elab_Term_elabSorry___closed__5; static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; @@ -589,6 +587,7 @@ lean_object* l_Lean_Elab_Term_adaptExpander(lean_object*, lean_object*, lean_obj static lean_object* l_Lean_Elab_Term_expandShow___closed__8; static lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Term_elabAnonymousCtor___spec__1___closed__2; static lean_object* l_Lean_Elab_Term_expandDbgTrace___closed__12; +lean_object* l_Lean_Syntax_mkNameLit(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandParen___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_expandDbgTrace_docString(lean_object*); @@ -659,7 +658,9 @@ static lean_object* l_Lean_Elab_Term_elabSorry___closed__9; static lean_object* l_Lean_Elab_Term_elabBorrowed___closed__1; static lean_object* l_Lean_Elab_Term_elabSubst___lambda__8___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__8; static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__4___closed__1; +static lean_object* l_Lean_Elab_Term_elabPanic___closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkPairs___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabStateRefT_declRange___closed__3; LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -769,6 +770,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabNoindex_declRange(lea static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabTParserMacroAux___closed__4; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_elabAnonymousCtor___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabNoindex_declRange___closed__4; +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_elabAnonymousCtor___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandHave_declRange___closed__1; LEAN_EXPORT lean_object* l_Lean_getRefPosition___at_Lean_Elab_Term_elabPanic___spec__1___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabSorry___closed__5; @@ -794,7 +796,7 @@ static lean_object* l_Lean_Elab_Term_elabPanic___closed__4; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandSuffices_declRange___closed__7; static lean_object* l_Lean_Elab_Term_expandShow___closed__19; -LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabTParserMacroAux___closed__8; static lean_object* l_Lean_Elab_Term_expandAssert___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -805,7 +807,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabCoe(lean_object*); lean_object* l_Lean_markBorrowed(lean_object*); uint8_t lean_string_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandParen___closed__6; -lean_object* l_Lean_Syntax_mkLit(lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); lean_object* l_Lean_Meta_getFVarLocalDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabSorry___closed__3; @@ -1596,6 +1597,54 @@ return x_36; } } } +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_elabAnonymousCtor___spec__4(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; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_9 = lean_ctor_get(x_6, 5); +x_10 = lean_ctor_get(x_2, 1); +lean_inc(x_10); +lean_inc(x_10); +x_11 = l_Lean_Elab_getBetterRef(x_9, x_10); +x_12 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_4, x_5, x_6, x_7, x_8); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(x_13, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_14); +lean_dec(x_2); +lean_dec(x_10); +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_15, 0); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_11); +lean_ctor_set(x_18, 1, x_17); +lean_ctor_set_tag(x_15, 1); +lean_ctor_set(x_15, 0, x_18); +return x_15; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_19 = lean_ctor_get(x_15, 0); +x_20 = lean_ctor_get(x_15, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_15); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_11); +lean_ctor_set(x_21, 1, x_19); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_20); +return x_22; +} +} +} LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabAnonymousCtor___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, lean_object* x_9, lean_object* x_10) { _start: { @@ -1940,7 +1989,7 @@ x_75 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__16; x_76 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_76, 0, x_74); lean_ctor_set(x_76, 1, x_75); -x_77 = l_Lean_throwError___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__11(x_76, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_77 = l_Lean_throwError___at_Lean_Elab_Term_elabAnonymousCtor___spec__4(x_76, x_9, x_10, x_11, x_12, x_13, x_14, x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); @@ -2351,7 +2400,7 @@ lean_inc(x_56); x_57 = lean_ctor_get(x_55, 1); lean_inc(x_57); lean_dec(x_55); -x_58 = l_Lean_Syntax_SepArray_getElems___rarg(x_15); +x_58 = l_Lean_Syntax_TSepArray_getElems___rarg(x_15); lean_dec(x_15); x_59 = lean_array_get_size(x_58); x_60 = lean_nat_dec_lt(x_59, x_56); @@ -2683,6 +2732,19 @@ lean_dec(x_1); return x_14; } } +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_elabAnonymousCtor___spec__4___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_throwError___at_Lean_Elab_Term_elabAnonymousCtor___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_9; +} +} LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabAnonymousCtor___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, lean_object* x_9, lean_object* x_10) { _start: { @@ -5400,470 +5462,467 @@ return x_25; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; +lean_object* x_26; lean_object* x_27; uint8_t x_28; x_26 = l_Lean_Syntax_getArg(x_17, x_16); x_27 = l_Lean_Syntax_getArg(x_17, x_10); -x_28 = l_Lean_nullKind; -x_29 = l_Lean_Syntax_isNodeOf(x_27, x_28, x_16); -if (x_29 == 0) +x_28 = l_Lean_Syntax_matchesNull(x_27, x_16); +if (x_28 == 0) { -lean_object* x_30; lean_object* x_31; +lean_object* x_29; lean_object* x_30; lean_dec(x_26); lean_dec(x_17); lean_dec(x_2); lean_dec(x_1); -x_30 = lean_box(1); -x_31 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_3); -return x_31; +x_29 = lean_box(1); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_3); +return x_30; } else { -lean_object* x_32; lean_object* x_33; uint8_t x_34; -x_32 = lean_unsigned_to_nat(2u); -x_33 = l_Lean_Syntax_getArg(x_17, x_32); -x_34 = l_Lean_Syntax_isNone(x_33); -if (x_34 == 0) +lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_31 = lean_unsigned_to_nat(2u); +x_32 = l_Lean_Syntax_getArg(x_17, x_31); +x_33 = l_Lean_Syntax_isNone(x_32); +if (x_33 == 0) { -uint8_t x_35; -lean_inc(x_33); -x_35 = l_Lean_Syntax_isNodeOf(x_33, x_28, x_10); -if (x_35 == 0) +uint8_t x_34; +lean_inc(x_32); +x_34 = l_Lean_Syntax_matchesNull(x_32, x_10); +if (x_34 == 0) { -lean_object* x_36; lean_object* x_37; -lean_dec(x_33); +lean_object* x_35; lean_object* x_36; +lean_dec(x_32); lean_dec(x_26); lean_dec(x_17); lean_dec(x_2); lean_dec(x_1); -x_36 = lean_box(1); -x_37 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_3); -return x_37; +x_35 = lean_box(1); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_3); +return x_36; } else { -lean_object* x_38; lean_object* x_39; uint8_t x_40; -x_38 = l_Lean_Syntax_getArg(x_33, x_16); -lean_dec(x_33); -x_39 = l_Lean_Elab_Term_expandShow___closed__25; -lean_inc(x_38); -x_40 = l_Lean_Syntax_isOfKind(x_38, x_39); -if (x_40 == 0) +lean_object* x_37; lean_object* x_38; uint8_t x_39; +x_37 = l_Lean_Syntax_getArg(x_32, x_16); +lean_dec(x_32); +x_38 = l_Lean_Elab_Term_expandShow___closed__25; +lean_inc(x_37); +x_39 = l_Lean_Syntax_isOfKind(x_37, x_38); +if (x_39 == 0) { -lean_object* x_41; lean_object* x_42; -lean_dec(x_38); +lean_object* x_40; lean_object* x_41; +lean_dec(x_37); lean_dec(x_26); lean_dec(x_17); lean_dec(x_2); lean_dec(x_1); -x_41 = lean_box(1); -x_42 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_3); -return x_42; +x_40 = lean_box(1); +x_41 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_3); +return x_41; } else { -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_43 = l_Lean_Syntax_getArg(x_38, x_10); -lean_dec(x_38); -x_44 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_44, 0, x_43); -x_45 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; -x_46 = lean_box(0); -x_47 = l_Lean_Elab_Term_expandHave___lambda__1(x_17, x_1, x_45, x_26, x_22, x_46, x_44, x_2, x_3); +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_42 = l_Lean_Syntax_getArg(x_37, x_10); +lean_dec(x_37); +x_43 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_43, 0, x_42); +x_44 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; +x_45 = lean_box(0); +x_46 = l_Lean_Elab_Term_expandHave___lambda__1(x_17, x_1, x_44, x_26, x_22, x_45, x_43, x_2, x_3); lean_dec(x_1); lean_dec(x_17); -return x_47; +return x_46; } } } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -lean_dec(x_33); -x_48 = lean_box(0); -x_49 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; -x_50 = lean_box(0); -x_51 = l_Lean_Elab_Term_expandHave___lambda__1(x_17, x_1, x_49, x_26, x_22, x_50, x_48, x_2, x_3); +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +lean_dec(x_32); +x_47 = lean_box(0); +x_48 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; +x_49 = lean_box(0); +x_50 = l_Lean_Elab_Term_expandHave___lambda__1(x_17, x_1, x_48, x_26, x_22, x_49, x_47, x_2, x_3); lean_dec(x_1); lean_dec(x_17); -return x_51; +return x_50; } } } } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; -x_52 = l_Lean_Syntax_getArg(x_17, x_16); -x_53 = l_Lean_nullKind; -x_54 = lean_unsigned_to_nat(2u); -lean_inc(x_52); -x_55 = l_Lean_Syntax_isNodeOf(x_52, x_53, x_54); -if (x_55 == 0) +lean_object* x_51; lean_object* x_52; uint8_t x_53; +x_51 = l_Lean_Syntax_getArg(x_17, x_16); +x_52 = lean_unsigned_to_nat(2u); +lean_inc(x_51); +x_53 = l_Lean_Syntax_matchesNull(x_51, x_52); +if (x_53 == 0) { -uint8_t x_56; -x_56 = l_Lean_Syntax_isNodeOf(x_52, x_53, x_16); -if (x_56 == 0) +uint8_t x_54; +x_54 = l_Lean_Syntax_matchesNull(x_51, x_16); +if (x_54 == 0) { -lean_object* x_57; lean_object* x_58; +lean_object* x_55; lean_object* x_56; lean_dec(x_17); lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); -x_57 = lean_box(1); -x_58 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_58, 0, x_57); -lean_ctor_set(x_58, 1, x_3); -return x_58; +x_55 = lean_box(1); +x_56 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_3); +return x_56; } else { -lean_object* x_59; uint8_t x_60; -x_59 = l_Lean_Syntax_getArg(x_17, x_10); -x_60 = l_Lean_Syntax_isNone(x_59); -if (x_60 == 0) +lean_object* x_57; uint8_t x_58; +x_57 = l_Lean_Syntax_getArg(x_17, x_10); +x_58 = l_Lean_Syntax_isNone(x_57); +if (x_58 == 0) { -uint8_t x_61; -lean_inc(x_59); -x_61 = l_Lean_Syntax_isNodeOf(x_59, x_53, x_10); -if (x_61 == 0) +uint8_t x_59; +lean_inc(x_57); +x_59 = l_Lean_Syntax_matchesNull(x_57, x_10); +if (x_59 == 0) { -lean_object* x_62; lean_object* x_63; -lean_dec(x_59); +lean_object* x_60; lean_object* x_61; +lean_dec(x_57); lean_dec(x_17); lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); -x_62 = lean_box(1); -x_63 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_63, 0, x_62); -lean_ctor_set(x_63, 1, x_3); -return x_63; +x_60 = lean_box(1); +x_61 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set(x_61, 1, x_3); +return x_61; } else { -lean_object* x_64; lean_object* x_65; uint8_t x_66; -x_64 = l_Lean_Syntax_getArg(x_59, x_16); -lean_dec(x_59); -x_65 = l_Lean_Elab_Term_expandShow___closed__25; -lean_inc(x_64); -x_66 = l_Lean_Syntax_isOfKind(x_64, x_65); -if (x_66 == 0) +lean_object* x_62; lean_object* x_63; uint8_t x_64; +x_62 = l_Lean_Syntax_getArg(x_57, x_16); +lean_dec(x_57); +x_63 = l_Lean_Elab_Term_expandShow___closed__25; +lean_inc(x_62); +x_64 = l_Lean_Syntax_isOfKind(x_62, x_63); +if (x_64 == 0) { -lean_object* x_67; lean_object* x_68; -lean_dec(x_64); +lean_object* x_65; lean_object* x_66; +lean_dec(x_62); lean_dec(x_17); lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); -x_67 = lean_box(1); -x_68 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_68, 0, x_67); -lean_ctor_set(x_68, 1, x_3); -return x_68; +x_65 = lean_box(1); +x_66 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_66, 0, x_65); +lean_ctor_set(x_66, 1, x_3); +return x_66; } else { -lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; -x_69 = l_Lean_Syntax_getArg(x_64, x_10); -lean_dec(x_64); -x_70 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_70, 0, x_69); -x_71 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; -x_72 = lean_box(0); -x_73 = l_Lean_Elab_Term_expandHave___lambda__2(x_17, x_71, x_1, x_5, x_20, x_12, x_6, x_72, x_70, x_2, x_3); +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_67 = l_Lean_Syntax_getArg(x_62, x_10); +lean_dec(x_62); +x_68 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_68, 0, x_67); +x_69 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; +x_70 = lean_box(0); +x_71 = l_Lean_Elab_Term_expandHave___lambda__2(x_17, x_69, x_1, x_5, x_20, x_12, x_6, x_70, x_68, x_2, x_3); lean_dec(x_1); lean_dec(x_17); -return x_73; +return x_71; } } } else { -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; -lean_dec(x_59); +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; +lean_dec(x_57); +x_72 = lean_box(0); +x_73 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; x_74 = lean_box(0); -x_75 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; -x_76 = lean_box(0); -x_77 = l_Lean_Elab_Term_expandHave___lambda__2(x_17, x_75, x_1, x_5, x_20, x_12, x_6, x_76, x_74, x_2, x_3); +x_75 = l_Lean_Elab_Term_expandHave___lambda__2(x_17, x_73, x_1, x_5, x_20, x_12, x_6, x_74, x_72, x_2, x_3); lean_dec(x_1); lean_dec(x_17); -return x_77; +return x_75; } } } else { -lean_object* x_78; lean_object* x_79; lean_object* x_80; uint8_t x_81; +lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; lean_dec(x_5); -x_78 = l_Lean_Syntax_getArg(x_52, x_16); -x_79 = l_Lean_Syntax_getArg(x_52, x_10); -lean_dec(x_52); -x_80 = l_Lean_Syntax_getArg(x_17, x_10); -x_81 = l_Lean_Syntax_isNone(x_80); -if (x_81 == 0) +x_76 = l_Lean_Syntax_getArg(x_51, x_16); +x_77 = l_Lean_Syntax_getArg(x_51, x_10); +lean_dec(x_51); +x_78 = l_Lean_Syntax_getArg(x_17, x_10); +x_79 = l_Lean_Syntax_isNone(x_78); +if (x_79 == 0) { -uint8_t x_82; -lean_inc(x_80); -x_82 = l_Lean_Syntax_isNodeOf(x_80, x_53, x_10); -if (x_82 == 0) +uint8_t x_80; +lean_inc(x_78); +x_80 = l_Lean_Syntax_matchesNull(x_78, x_10); +if (x_80 == 0) { -lean_object* x_83; lean_object* x_84; -lean_dec(x_80); -lean_dec(x_79); +lean_object* x_81; lean_object* x_82; lean_dec(x_78); +lean_dec(x_77); +lean_dec(x_76); lean_dec(x_17); lean_dec(x_2); lean_dec(x_1); -x_83 = lean_box(1); -x_84 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_84, 0, x_83); -lean_ctor_set(x_84, 1, x_3); -return x_84; +x_81 = lean_box(1); +x_82 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_82, 0, x_81); +lean_ctor_set(x_82, 1, x_3); +return x_82; } else { -lean_object* x_85; lean_object* x_86; uint8_t x_87; -x_85 = l_Lean_Syntax_getArg(x_80, x_16); -lean_dec(x_80); -x_86 = l_Lean_Elab_Term_expandShow___closed__25; -lean_inc(x_85); -x_87 = l_Lean_Syntax_isOfKind(x_85, x_86); -if (x_87 == 0) -{ -lean_object* x_88; lean_object* x_89; -lean_dec(x_85); -lean_dec(x_79); +lean_object* x_83; lean_object* x_84; uint8_t x_85; +x_83 = l_Lean_Syntax_getArg(x_78, x_16); lean_dec(x_78); +x_84 = l_Lean_Elab_Term_expandShow___closed__25; +lean_inc(x_83); +x_85 = l_Lean_Syntax_isOfKind(x_83, x_84); +if (x_85 == 0) +{ +lean_object* x_86; lean_object* x_87; +lean_dec(x_83); +lean_dec(x_77); +lean_dec(x_76); lean_dec(x_17); lean_dec(x_2); lean_dec(x_1); -x_88 = lean_box(1); -x_89 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_89, 0, x_88); -lean_ctor_set(x_89, 1, x_3); -return x_89; +x_86 = lean_box(1); +x_87 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_87, 0, x_86); +lean_ctor_set(x_87, 1, x_3); +return x_87; } else { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; -x_90 = l_Lean_Syntax_getArg(x_85, x_10); -lean_dec(x_85); -x_91 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_91, 0, x_90); -x_92 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; -x_93 = lean_box(0); -x_94 = l_Lean_Elab_Term_expandHave___lambda__3(x_17, x_92, x_1, x_79, x_78, x_93, x_91, x_2, x_3); -lean_dec(x_79); +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_88 = l_Lean_Syntax_getArg(x_83, x_10); +lean_dec(x_83); +x_89 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_89, 0, x_88); +x_90 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; +x_91 = lean_box(0); +x_92 = l_Lean_Elab_Term_expandHave___lambda__3(x_17, x_90, x_1, x_77, x_76, x_91, x_89, x_2, x_3); +lean_dec(x_77); lean_dec(x_1); lean_dec(x_17); -return x_94; +return x_92; } } } else { -lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; -lean_dec(x_80); +lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; +lean_dec(x_78); +x_93 = lean_box(0); +x_94 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; x_95 = lean_box(0); -x_96 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; -x_97 = lean_box(0); -x_98 = l_Lean_Elab_Term_expandHave___lambda__3(x_17, x_96, x_1, x_79, x_78, x_97, x_95, x_2, x_3); -lean_dec(x_79); +x_96 = l_Lean_Elab_Term_expandHave___lambda__3(x_17, x_94, x_1, x_77, x_76, x_95, x_93, x_2, x_3); +lean_dec(x_77); lean_dec(x_1); lean_dec(x_17); -return x_98; +return x_96; } } } } else { -lean_object* x_99; lean_object* x_100; lean_object* x_101; uint8_t x_102; -x_99 = l_Lean_Syntax_getArg(x_17, x_16); -x_100 = l_Lean_nullKind; -x_101 = lean_unsigned_to_nat(2u); -lean_inc(x_99); -x_102 = l_Lean_Syntax_isNodeOf(x_99, x_100, x_101); -if (x_102 == 0) +lean_object* x_97; lean_object* x_98; uint8_t x_99; +x_97 = l_Lean_Syntax_getArg(x_17, x_16); +x_98 = lean_unsigned_to_nat(2u); +lean_inc(x_97); +x_99 = l_Lean_Syntax_matchesNull(x_97, x_98); +if (x_99 == 0) { -uint8_t x_103; -x_103 = l_Lean_Syntax_isNodeOf(x_99, x_100, x_16); -if (x_103 == 0) +uint8_t x_100; +x_100 = l_Lean_Syntax_matchesNull(x_97, x_16); +if (x_100 == 0) { -lean_object* x_104; lean_object* x_105; +lean_object* x_101; lean_object* x_102; lean_dec(x_17); lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); -x_104 = lean_box(1); -x_105 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_105, 0, x_104); -lean_ctor_set(x_105, 1, x_3); -return x_105; +x_101 = lean_box(1); +x_102 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_102, 0, x_101); +lean_ctor_set(x_102, 1, x_3); +return x_102; } else { -lean_object* x_106; uint8_t x_107; -x_106 = l_Lean_Syntax_getArg(x_17, x_10); -x_107 = l_Lean_Syntax_isNone(x_106); -if (x_107 == 0) +lean_object* x_103; uint8_t x_104; +x_103 = l_Lean_Syntax_getArg(x_17, x_10); +x_104 = l_Lean_Syntax_isNone(x_103); +if (x_104 == 0) { -uint8_t x_108; -lean_inc(x_106); -x_108 = l_Lean_Syntax_isNodeOf(x_106, x_100, x_10); -if (x_108 == 0) +uint8_t x_105; +lean_inc(x_103); +x_105 = l_Lean_Syntax_matchesNull(x_103, x_10); +if (x_105 == 0) { -lean_object* x_109; lean_object* x_110; -lean_dec(x_106); +lean_object* x_106; lean_object* x_107; +lean_dec(x_103); lean_dec(x_17); lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); -x_109 = lean_box(1); -x_110 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_110, 0, x_109); -lean_ctor_set(x_110, 1, x_3); -return x_110; +x_106 = lean_box(1); +x_107 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_107, 0, x_106); +lean_ctor_set(x_107, 1, x_3); +return x_107; } else { -lean_object* x_111; lean_object* x_112; uint8_t x_113; -x_111 = l_Lean_Syntax_getArg(x_106, x_16); -lean_dec(x_106); -x_112 = l_Lean_Elab_Term_expandShow___closed__25; -lean_inc(x_111); -x_113 = l_Lean_Syntax_isOfKind(x_111, x_112); -if (x_113 == 0) +lean_object* x_108; lean_object* x_109; uint8_t x_110; +x_108 = l_Lean_Syntax_getArg(x_103, x_16); +lean_dec(x_103); +x_109 = l_Lean_Elab_Term_expandShow___closed__25; +lean_inc(x_108); +x_110 = l_Lean_Syntax_isOfKind(x_108, x_109); +if (x_110 == 0) { -lean_object* x_114; lean_object* x_115; -lean_dec(x_111); +lean_object* x_111; lean_object* x_112; +lean_dec(x_108); lean_dec(x_17); lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); -x_114 = lean_box(1); -x_115 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_115, 0, x_114); -lean_ctor_set(x_115, 1, x_3); -return x_115; +x_111 = lean_box(1); +x_112 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_112, 0, x_111); +lean_ctor_set(x_112, 1, x_3); +return x_112; } else { -lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; -x_116 = l_Lean_Syntax_getArg(x_111, x_10); -lean_dec(x_111); -x_117 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_117, 0, x_116); -x_118 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; -x_119 = lean_box(0); -x_120 = l_Lean_Elab_Term_expandHave___lambda__4(x_17, x_1, x_5, x_18, x_12, x_6, x_118, x_119, x_117, x_2, x_3); +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; +x_113 = l_Lean_Syntax_getArg(x_108, x_10); +lean_dec(x_108); +x_114 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_114, 0, x_113); +x_115 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; +x_116 = lean_box(0); +x_117 = l_Lean_Elab_Term_expandHave___lambda__4(x_17, x_1, x_5, x_18, x_12, x_6, x_115, x_116, x_114, x_2, x_3); lean_dec(x_1); lean_dec(x_17); -return x_120; +return x_117; } } } else { -lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; -lean_dec(x_106); -x_121 = lean_box(0); -x_122 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; -x_123 = lean_box(0); -x_124 = l_Lean_Elab_Term_expandHave___lambda__4(x_17, x_1, x_5, x_18, x_12, x_6, x_122, x_123, x_121, x_2, x_3); +lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; +lean_dec(x_103); +x_118 = lean_box(0); +x_119 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; +x_120 = lean_box(0); +x_121 = l_Lean_Elab_Term_expandHave___lambda__4(x_17, x_1, x_5, x_18, x_12, x_6, x_119, x_120, x_118, x_2, x_3); lean_dec(x_1); lean_dec(x_17); -return x_124; +return x_121; } } } else { -lean_object* x_125; lean_object* x_126; lean_object* x_127; uint8_t x_128; +lean_object* x_122; lean_object* x_123; lean_object* x_124; uint8_t x_125; lean_dec(x_5); -x_125 = l_Lean_Syntax_getArg(x_99, x_16); -x_126 = l_Lean_Syntax_getArg(x_99, x_10); -lean_dec(x_99); -x_127 = l_Lean_Syntax_getArg(x_17, x_10); -x_128 = l_Lean_Syntax_isNone(x_127); -if (x_128 == 0) -{ -uint8_t x_129; -lean_inc(x_127); -x_129 = l_Lean_Syntax_isNodeOf(x_127, x_100, x_10); -if (x_129 == 0) -{ -lean_object* x_130; lean_object* x_131; -lean_dec(x_127); -lean_dec(x_126); -lean_dec(x_125); +x_122 = l_Lean_Syntax_getArg(x_97, x_16); +x_123 = l_Lean_Syntax_getArg(x_97, x_10); +lean_dec(x_97); +x_124 = l_Lean_Syntax_getArg(x_17, x_10); +x_125 = l_Lean_Syntax_isNone(x_124); +if (x_125 == 0) +{ +uint8_t x_126; +lean_inc(x_124); +x_126 = l_Lean_Syntax_matchesNull(x_124, x_10); +if (x_126 == 0) +{ +lean_object* x_127; lean_object* x_128; +lean_dec(x_124); +lean_dec(x_123); +lean_dec(x_122); lean_dec(x_17); lean_dec(x_2); lean_dec(x_1); -x_130 = lean_box(1); -x_131 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_131, 0, x_130); -lean_ctor_set(x_131, 1, x_3); -return x_131; +x_127 = lean_box(1); +x_128 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_128, 0, x_127); +lean_ctor_set(x_128, 1, x_3); +return x_128; } else { -lean_object* x_132; lean_object* x_133; uint8_t x_134; -x_132 = l_Lean_Syntax_getArg(x_127, x_16); -lean_dec(x_127); -x_133 = l_Lean_Elab_Term_expandShow___closed__25; -lean_inc(x_132); -x_134 = l_Lean_Syntax_isOfKind(x_132, x_133); -if (x_134 == 0) +lean_object* x_129; lean_object* x_130; uint8_t x_131; +x_129 = l_Lean_Syntax_getArg(x_124, x_16); +lean_dec(x_124); +x_130 = l_Lean_Elab_Term_expandShow___closed__25; +lean_inc(x_129); +x_131 = l_Lean_Syntax_isOfKind(x_129, x_130); +if (x_131 == 0) { -lean_object* x_135; lean_object* x_136; -lean_dec(x_132); -lean_dec(x_126); -lean_dec(x_125); +lean_object* x_132; lean_object* x_133; +lean_dec(x_129); +lean_dec(x_123); +lean_dec(x_122); lean_dec(x_17); lean_dec(x_2); lean_dec(x_1); -x_135 = lean_box(1); -x_136 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_136, 0, x_135); -lean_ctor_set(x_136, 1, x_3); -return x_136; -} -else -{ -lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; -x_137 = l_Lean_Syntax_getArg(x_132, x_10); -lean_dec(x_132); -x_138 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_138, 0, x_137); -x_139 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; -x_140 = lean_box(0); -x_141 = l_Lean_Elab_Term_expandHave___lambda__5(x_17, x_1, x_126, x_139, x_125, x_140, x_138, x_2, x_3); -lean_dec(x_126); +x_132 = lean_box(1); +x_133 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_133, 0, x_132); +lean_ctor_set(x_133, 1, x_3); +return x_133; +} +else +{ +lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; +x_134 = l_Lean_Syntax_getArg(x_129, x_10); +lean_dec(x_129); +x_135 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_135, 0, x_134); +x_136 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; +x_137 = lean_box(0); +x_138 = l_Lean_Elab_Term_expandHave___lambda__5(x_17, x_1, x_123, x_136, x_122, x_137, x_135, x_2, x_3); +lean_dec(x_123); lean_dec(x_1); lean_dec(x_17); -return x_141; +return x_138; } } } else { -lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; -lean_dec(x_127); -x_142 = lean_box(0); -x_143 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; -x_144 = lean_box(0); -x_145 = l_Lean_Elab_Term_expandHave___lambda__5(x_17, x_1, x_126, x_143, x_125, x_144, x_142, x_2, x_3); -lean_dec(x_126); +lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; +lean_dec(x_124); +x_139 = lean_box(0); +x_140 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; +x_141 = lean_box(0); +x_142 = l_Lean_Elab_Term_expandHave___lambda__5(x_17, x_1, x_123, x_140, x_122, x_141, x_139, x_2, x_3); +lean_dec(x_123); lean_dec(x_1); lean_dec(x_17); -return x_145; +return x_142; } } } @@ -6653,52 +6712,51 @@ x_15 = l_Lean_Syntax_getArg(x_9, x_14); x_16 = l_Lean_Syntax_isNone(x_15); if (x_16 == 0) { -lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_17 = l_Lean_nullKind; -x_18 = lean_unsigned_to_nat(2u); +lean_object* x_17; uint8_t x_18; +x_17 = lean_unsigned_to_nat(2u); lean_inc(x_15); -x_19 = l_Lean_Syntax_isNodeOf(x_15, x_17, x_18); -if (x_19 == 0) +x_18 = l_Lean_Syntax_matchesNull(x_15, x_17); +if (x_18 == 0) { -lean_object* x_20; lean_object* x_21; +lean_object* x_19; lean_object* x_20; lean_dec(x_15); lean_dec(x_9); lean_dec(x_2); lean_dec(x_1); -x_20 = lean_box(1); -x_21 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_3); -return x_21; +x_19 = lean_box(1); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_3); +return x_20; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_22 = l_Lean_Syntax_getArg(x_15, x_14); +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_21 = l_Lean_Syntax_getArg(x_15, x_14); lean_dec(x_15); -x_23 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_23, 0, x_22); -x_24 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; -x_25 = l_Lean_Elab_Term_elabAnonymousCtor___closed__2; -x_26 = lean_box(0); -x_27 = l_Lean_Elab_Term_expandSuffices___lambda__1(x_9, x_24, x_25, x_1, x_26, x_23, x_2, x_3); +x_22 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_22, 0, x_21); +x_23 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; +x_24 = l_Lean_Elab_Term_elabAnonymousCtor___closed__2; +x_25 = lean_box(0); +x_26 = l_Lean_Elab_Term_expandSuffices___lambda__1(x_9, x_23, x_24, x_1, x_25, x_22, x_2, x_3); lean_dec(x_1); lean_dec(x_9); -return x_27; +return x_26; } } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_dec(x_15); -x_28 = lean_box(0); -x_29 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; -x_30 = l_Lean_Elab_Term_elabAnonymousCtor___closed__2; -x_31 = lean_box(0); -x_32 = l_Lean_Elab_Term_expandSuffices___lambda__1(x_9, x_29, x_30, x_1, x_31, x_28, x_2, x_3); +x_27 = lean_box(0); +x_28 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; +x_29 = l_Lean_Elab_Term_elabAnonymousCtor___closed__2; +x_30 = lean_box(0); +x_31 = l_Lean_Elab_Term_expandSuffices___lambda__1(x_9, x_28, x_29, x_1, x_30, x_27, x_2, x_3); lean_dec(x_1); lean_dec(x_9); -return x_32; +return x_31; } } } @@ -7316,13 +7374,13 @@ lean_dec(x_20); if (lean_obj_tag(x_22) == 0) { lean_object* x_192; -x_192 = l___private_Init_Meta_0__Lean_quoteNameMk(x_17); +x_192 = l_Lean_quoteNameMk(x_17); x_25 = x_192; goto block_191; } else { -lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; +lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_dec(x_17); x_193 = lean_ctor_get(x_22, 0); lean_inc(x_193); @@ -7332,16 +7390,15 @@ x_195 = l_String_intercalate(x_194, x_193); x_196 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__42; x_197 = lean_string_append(x_196, x_195); lean_dec(x_195); -x_198 = l_Lean_nameLitKind; -x_199 = l_Lean_Syntax_mkLit(x_198, x_197, x_23); -x_200 = l_Lean_Elab_Term_expandShow___closed__27; -x_201 = lean_array_push(x_200, x_199); -x_202 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__40; -x_203 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_203, 0, x_23); -lean_ctor_set(x_203, 1, x_202); -lean_ctor_set(x_203, 2, x_201); -x_25 = x_203; +x_198 = l_Lean_Syntax_mkNameLit(x_197, x_23); +x_199 = l_Lean_Elab_Term_expandShow___closed__27; +x_200 = lean_array_push(x_199, x_198); +x_201 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__40; +x_202 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_202, 0, x_23); +lean_ctor_set(x_202, 1, x_201); +lean_ctor_set(x_202, 2, x_200); +x_25 = x_202; goto block_191; } block_191: @@ -7713,15 +7770,15 @@ return x_190; } else { -lean_object* x_204; lean_object* x_205; +lean_object* x_203; lean_object* x_204; lean_dec(x_19); lean_dec(x_17); lean_dec(x_2); lean_dec(x_1); -x_204 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__4; -x_205 = l_Lean_throwError___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__11(x_204, x_4, x_5, x_6, x_7, x_8, x_9, x_16); +x_203 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__4; +x_204 = l_Lean_throwError___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__11(x_203, x_4, x_5, x_6, x_7, x_8, x_9, x_16); lean_dec(x_8); -return x_205; +return x_204; } } } @@ -7752,12 +7809,11 @@ return x_2; static lean_object* _init_l_Lean_Elab_Term_elabLeadingParserMacro___lambda__1___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_numLitKind; -x_2 = l_Lean_Elab_Term_elabLeadingParserMacro___lambda__1___closed__1; -x_3 = lean_box(2); -x_4 = l_Lean_Syntax_mkLit(x_1, x_2, x_3); -return x_4; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Term_elabLeadingParserMacro___lambda__1___closed__1; +x_2 = lean_box(2); +x_3 = l_Lean_Syntax_mkNumLit(x_1, x_2); +return x_3; } } static lean_object* _init_l_Lean_Elab_Term_elabLeadingParserMacro___lambda__1___closed__3() { @@ -7850,14 +7906,13 @@ x_13 = l_Lean_Syntax_getArg(x_1, x_12); x_14 = l_Lean_Syntax_isNone(x_13); if (x_14 == 0) { -lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_15 = l_Lean_nullKind; -x_16 = lean_unsigned_to_nat(1u); +lean_object* x_15; uint8_t x_16; +x_15 = lean_unsigned_to_nat(1u); lean_inc(x_13); -x_17 = l_Lean_Syntax_isNodeOf(x_13, x_15, x_16); -if (x_17 == 0) +x_16 = l_Lean_Syntax_matchesNull(x_13, x_15); +if (x_16 == 0) { -lean_object* x_18; +lean_object* x_17; lean_dec(x_13); lean_dec(x_10); lean_dec(x_9); @@ -7867,25 +7922,25 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_18 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13___rarg(x_11); -return x_18; +x_17 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13___rarg(x_11); +return x_17; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_19 = lean_unsigned_to_nat(0u); -x_20 = l_Lean_Syntax_getArg(x_13, x_19); +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_18 = lean_unsigned_to_nat(0u); +x_19 = l_Lean_Syntax_getArg(x_13, x_18); lean_dec(x_13); -x_21 = l_Lean_Elab_Term_elabLeadingParserMacro___lambda__2___closed__1; +x_20 = l_Lean_Elab_Term_elabLeadingParserMacro___lambda__2___closed__1; lean_inc(x_2); -x_22 = lean_name_mk_string(x_2, x_21); -lean_inc(x_20); -x_23 = l_Lean_Syntax_isOfKind(x_20, x_22); -lean_dec(x_22); -if (x_23 == 0) +x_21 = lean_name_mk_string(x_2, x_20); +lean_inc(x_19); +x_22 = l_Lean_Syntax_isOfKind(x_19, x_21); +lean_dec(x_21); +if (x_22 == 0) { -lean_object* x_24; -lean_dec(x_20); +lean_object* x_23; +lean_dec(x_19); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -7894,39 +7949,39 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_24 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13___rarg(x_11); -return x_24; +x_23 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13___rarg(x_11); +return x_23; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_25 = lean_unsigned_to_nat(3u); -x_26 = l_Lean_Syntax_getArg(x_20, x_25); -lean_dec(x_20); -x_27 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_27, 0, x_26); -x_28 = lean_box(0); -x_29 = l_Lean_Elab_Term_elabLeadingParserMacro___lambda__1(x_1, x_4, x_2, x_28, x_27, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_24 = lean_unsigned_to_nat(3u); +x_25 = l_Lean_Syntax_getArg(x_19, x_24); +lean_dec(x_19); +x_26 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_26, 0, x_25); +x_27 = lean_box(0); +x_28 = l_Lean_Elab_Term_elabLeadingParserMacro___lambda__1(x_1, x_4, x_2, x_27, x_26, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -return x_29; +return x_28; } } } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_dec(x_13); +x_29 = lean_box(0); x_30 = lean_box(0); -x_31 = lean_box(0); -x_32 = l_Lean_Elab_Term_elabLeadingParserMacro___lambda__1(x_1, x_4, x_2, x_31, x_30, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_31 = l_Lean_Elab_Term_elabLeadingParserMacro___lambda__1(x_1, x_4, x_2, x_30, x_29, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -return x_32; +return x_31; } } } @@ -7976,14 +8031,13 @@ x_13 = l_Lean_Syntax_getArg(x_1, x_12); x_14 = l_Lean_Syntax_isNone(x_13); if (x_14 == 0) { -lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_15 = l_Lean_nullKind; -x_16 = lean_unsigned_to_nat(2u); +lean_object* x_15; uint8_t x_16; +x_15 = lean_unsigned_to_nat(2u); lean_inc(x_13); -x_17 = l_Lean_Syntax_isNodeOf(x_13, x_15, x_16); -if (x_17 == 0) +x_16 = l_Lean_Syntax_matchesNull(x_13, x_15); +if (x_16 == 0) { -lean_object* x_18; +lean_object* x_17; lean_dec(x_13); lean_dec(x_7); lean_dec(x_6); @@ -7992,33 +8046,33 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_18 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13___rarg(x_8); -return x_18; +x_17 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13___rarg(x_8); +return x_17; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_19 = l_Lean_Syntax_getArg(x_13, x_12); +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_18 = l_Lean_Syntax_getArg(x_13, x_12); lean_dec(x_13); -x_20 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_20, 0, x_19); -x_21 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; -x_22 = lean_box(0); -x_23 = l_Lean_Elab_Term_elabLeadingParserMacro___lambda__2(x_1, x_21, x_22, x_20, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_19 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_19, 0, x_18); +x_20 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; +x_21 = lean_box(0); +x_22 = l_Lean_Elab_Term_elabLeadingParserMacro___lambda__2(x_1, x_20, x_21, x_19, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_1); -return x_23; +return x_22; } } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_dec(x_13); -x_24 = lean_box(0); -x_25 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; -x_26 = lean_box(0); -x_27 = l_Lean_Elab_Term_elabLeadingParserMacro___lambda__2(x_1, x_25, x_26, x_24, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_23 = lean_box(0); +x_24 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; +x_25 = lean_box(0); +x_26 = l_Lean_Elab_Term_elabLeadingParserMacro___lambda__2(x_1, x_24, x_25, x_23, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_1); -return x_27; +return x_26; } } } @@ -8330,13 +8384,13 @@ x_19 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_18, x_17); if (lean_obj_tag(x_19) == 0) { lean_object* x_72; -x_72 = l___private_Init_Meta_0__Lean_quoteNameMk(x_17); +x_72 = l_Lean_quoteNameMk(x_17); x_20 = x_72; goto block_71; } else { -lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_dec(x_17); x_73 = lean_ctor_get(x_19, 0); lean_inc(x_73); @@ -8346,17 +8400,16 @@ x_75 = l_String_intercalate(x_74, x_73); x_76 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__42; x_77 = lean_string_append(x_76, x_75); lean_dec(x_75); -x_78 = l_Lean_nameLitKind; -x_79 = lean_box(2); -x_80 = l_Lean_Syntax_mkLit(x_78, x_77, x_79); -x_81 = l_Lean_Elab_Term_expandShow___closed__27; -x_82 = lean_array_push(x_81, x_80); -x_83 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__40; -x_84 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_84, 0, x_79); -lean_ctor_set(x_84, 1, x_83); -lean_ctor_set(x_84, 2, x_82); -x_20 = x_84; +x_78 = lean_box(2); +x_79 = l_Lean_Syntax_mkNameLit(x_77, x_78); +x_80 = l_Lean_Elab_Term_expandShow___closed__27; +x_81 = lean_array_push(x_80, x_79); +x_82 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__40; +x_83 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_83, 0, x_78); +lean_ctor_set(x_83, 1, x_82); +lean_ctor_set(x_83, 2, x_81); +x_20 = x_83; goto block_71; } block_71: @@ -8486,12 +8539,11 @@ return x_2; static lean_object* _init_l_Lean_Elab_Term_elabTrailingParserMacro___lambda__1___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_numLitKind; -x_2 = l_Lean_Elab_Term_elabTrailingParserMacro___lambda__1___closed__1; -x_3 = lean_box(2); -x_4 = l_Lean_Syntax_mkLit(x_1, x_2, x_3); -return x_4; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Term_elabTrailingParserMacro___lambda__1___closed__1; +x_2 = lean_box(2); +x_3 = l_Lean_Syntax_mkNumLit(x_1, x_2); +return x_3; } } LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabTrailingParserMacro___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, lean_object* x_9, lean_object* x_10, lean_object* x_11) { @@ -8558,13 +8610,12 @@ x_12 = l_Lean_Syntax_getArg(x_1, x_11); x_13 = l_Lean_Syntax_isNone(x_12); if (x_13 == 0) { -lean_object* x_14; uint8_t x_15; -x_14 = l_Lean_nullKind; +uint8_t x_14; lean_inc(x_12); -x_15 = l_Lean_Syntax_isNodeOf(x_12, x_14, x_11); -if (x_15 == 0) +x_14 = l_Lean_Syntax_matchesNull(x_12, x_11); +if (x_14 == 0) { -lean_object* x_16; +lean_object* x_15; lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); @@ -8573,38 +8624,38 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_16 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13___rarg(x_10); -return x_16; +x_15 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13___rarg(x_10); +return x_15; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_17 = lean_unsigned_to_nat(1u); -x_18 = l_Lean_Syntax_getArg(x_12, x_17); +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_16 = lean_unsigned_to_nat(1u); +x_17 = l_Lean_Syntax_getArg(x_12, x_16); lean_dec(x_12); -x_19 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_19, 0, x_18); -x_20 = lean_box(0); -x_21 = l_Lean_Elab_Term_elabTrailingParserMacro___lambda__1(x_1, x_3, x_20, x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_18 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_18, 0, x_17); +x_19 = lean_box(0); +x_20 = l_Lean_Elab_Term_elabTrailingParserMacro___lambda__1(x_1, x_3, x_19, x_18, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -return x_21; +return x_20; } } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_dec(x_12); +x_21 = lean_box(0); x_22 = lean_box(0); -x_23 = lean_box(0); -x_24 = l_Lean_Elab_Term_elabTrailingParserMacro___lambda__1(x_1, x_3, x_23, x_22, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_23 = l_Lean_Elab_Term_elabTrailingParserMacro___lambda__1(x_1, x_3, x_22, x_21, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -return x_24; +return x_23; } } } @@ -8654,14 +8705,13 @@ x_13 = l_Lean_Syntax_getArg(x_1, x_12); x_14 = l_Lean_Syntax_isNone(x_13); if (x_14 == 0) { -lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_15 = l_Lean_nullKind; -x_16 = lean_unsigned_to_nat(2u); +lean_object* x_15; uint8_t x_16; +x_15 = lean_unsigned_to_nat(2u); lean_inc(x_13); -x_17 = l_Lean_Syntax_isNodeOf(x_13, x_15, x_16); -if (x_17 == 0) +x_16 = l_Lean_Syntax_matchesNull(x_13, x_15); +if (x_16 == 0) { -lean_object* x_18; +lean_object* x_17; lean_dec(x_13); lean_dec(x_7); lean_dec(x_6); @@ -8670,31 +8720,31 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_18 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13___rarg(x_8); -return x_18; +x_17 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13___rarg(x_8); +return x_17; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_19 = l_Lean_Syntax_getArg(x_13, x_12); +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_18 = l_Lean_Syntax_getArg(x_13, x_12); lean_dec(x_13); -x_20 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_20, 0, x_19); -x_21 = lean_box(0); -x_22 = l_Lean_Elab_Term_elabTrailingParserMacro___lambda__2(x_1, x_21, x_20, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_19 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_19, 0, x_18); +x_20 = lean_box(0); +x_21 = l_Lean_Elab_Term_elabTrailingParserMacro___lambda__2(x_1, x_20, x_19, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_1); -return x_22; +return x_21; } } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_dec(x_13); +x_22 = lean_box(0); x_23 = lean_box(0); -x_24 = lean_box(0); -x_25 = l_Lean_Elab_Term_elabTrailingParserMacro___lambda__2(x_1, x_24, x_23, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_24 = l_Lean_Elab_Term_elabTrailingParserMacro___lambda__2(x_1, x_23, x_22, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_1); -return x_25; +return x_24; } } } @@ -8924,26 +8974,44 @@ static lean_object* _init_l_Lean_Elab_Term_elabPanic___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("panicWithPos", 12); +x_1 = lean_mk_string_from_bytes("panic", 5); return x_1; } } static lean_object* _init_l_Lean_Elab_Term_elabPanic___closed__2() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; +x_2 = l_Lean_Elab_Term_elabPanic___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_elabPanic___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("panicWithPos", 12); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_elabPanic___closed__4() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_elabPanic___closed__1; +x_1 = l_Lean_Elab_Term_elabPanic___closed__3; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_elabPanic___closed__3() { +static lean_object* _init_l_Lean_Elab_Term_elabPanic___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Term_elabPanic___closed__1; +x_1 = l_Lean_Elab_Term_elabPanic___closed__3; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Term_elabPanic___closed__2; +x_3 = l_Lean_Elab_Term_elabPanic___closed__4; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -8951,41 +9019,41 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_elabPanic___closed__4() { +static lean_object* _init_l_Lean_Elab_Term_elabPanic___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_elabPanic___closed__1; +x_2 = l_Lean_Elab_Term_elabPanic___closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_elabPanic___closed__5() { +static lean_object* _init_l_Lean_Elab_Term_elabPanic___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_elabPanic___closed__4; +x_2 = l_Lean_Elab_Term_elabPanic___closed__6; 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_Elab_Term_elabPanic___closed__6() { +static lean_object* _init_l_Lean_Elab_Term_elabPanic___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_elabPanic___closed__5; +x_2 = l_Lean_Elab_Term_elabPanic___closed__7; x_3 = lean_alloc_ctor(1, 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_Elab_Term_elabPanic___closed__7() { +static lean_object* _init_l_Lean_Elab_Term_elabPanic___closed__9() { _start: { lean_object* x_1; @@ -8993,22 +9061,22 @@ x_1 = lean_mk_string_from_bytes("panicWithPosWithDecl", 20); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_elabPanic___closed__8() { +static lean_object* _init_l_Lean_Elab_Term_elabPanic___closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_elabPanic___closed__7; +x_1 = l_Lean_Elab_Term_elabPanic___closed__9; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_elabPanic___closed__9() { +static lean_object* _init_l_Lean_Elab_Term_elabPanic___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Term_elabPanic___closed__7; +x_1 = l_Lean_Elab_Term_elabPanic___closed__9; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Term_elabPanic___closed__8; +x_3 = l_Lean_Elab_Term_elabPanic___closed__10; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -9016,34 +9084,34 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_elabPanic___closed__10() { +static lean_object* _init_l_Lean_Elab_Term_elabPanic___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_elabPanic___closed__7; +x_2 = l_Lean_Elab_Term_elabPanic___closed__9; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_elabPanic___closed__11() { +static lean_object* _init_l_Lean_Elab_Term_elabPanic___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_elabPanic___closed__10; +x_2 = l_Lean_Elab_Term_elabPanic___closed__12; 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_Elab_Term_elabPanic___closed__12() { +static lean_object* _init_l_Lean_Elab_Term_elabPanic___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_elabPanic___closed__11; +x_2 = l_Lean_Elab_Term_elabPanic___closed__13; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); @@ -9053,175 +9121,194 @@ return x_3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabPanic(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_object* x_9) { _start: { -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; lean_object* x_19; lean_object* x_20; -x_10 = lean_unsigned_to_nat(1u); -x_11 = l_Lean_Syntax_getArg(x_1, x_10); -x_12 = l_Lean_getRefPosition___at_Lean_Elab_Term_elabPanic___spec__1___rarg(x_7, x_8, x_9); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec(x_12); -x_15 = lean_st_ref_get(x_8, x_14); +lean_object* x_10; uint8_t x_11; +x_10 = l_Lean_Elab_Term_elabPanic___closed__2; +lean_inc(x_1); +x_11 = l_Lean_Syntax_isOfKind(x_1, x_10); +if (x_11 == 0) +{ +lean_object* x_12; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_12 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_elabForall___spec__1___rarg(x_9); +return x_12; +} +else +{ +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; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_13 = lean_unsigned_to_nat(1u); +x_14 = l_Lean_Syntax_getArg(x_1, x_13); +x_15 = l_Lean_getRefPosition___at_Lean_Elab_Term_elabPanic___spec__1___rarg(x_7, x_8, x_9); x_16 = lean_ctor_get(x_15, 0); lean_inc(x_16); x_17 = lean_ctor_get(x_15, 1); lean_inc(x_17); lean_dec(x_15); -x_18 = lean_ctor_get(x_16, 0); -lean_inc(x_18); -lean_dec(x_16); -x_19 = l_Lean_Elab_Term_getDeclName_x3f(x_3, x_4, x_5, x_6, x_7, x_8, x_17); -x_20 = lean_ctor_get(x_19, 0); +x_18 = lean_st_ref_get(x_8, x_17); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); lean_inc(x_20); -if (lean_obj_tag(x_20) == 0) -{ -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; 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; uint8_t 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_21 = lean_ctor_get(x_19, 1); +lean_dec(x_18); +x_21 = lean_ctor_get(x_19, 0); lean_inc(x_21); lean_dec(x_19); -lean_inc(x_7); -x_22 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_7, x_8, x_21); +x_22 = l_Lean_Elab_Term_getDeclName_x3f(x_3, x_4, x_5, x_6, x_7, x_8, x_20); x_23 = lean_ctor_get(x_22, 0); lean_inc(x_23); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* 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; 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; uint8_t x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; x_24 = lean_ctor_get(x_22, 1); lean_inc(x_24); lean_dec(x_22); -x_25 = lean_ctor_get(x_7, 10); -lean_inc(x_25); -x_26 = lean_st_ref_get(x_8, x_24); -x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_7); +x_25 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_7, x_8, x_24); +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_26, 1); +lean_dec(x_25); +x_28 = lean_ctor_get(x_7, 10); lean_inc(x_28); -lean_dec(x_26); -x_29 = lean_ctor_get(x_27, 0); -lean_inc(x_29); -lean_dec(x_27); -x_30 = lean_environment_main_module(x_29); -x_31 = l_Lean_Elab_Term_elabPanic___closed__4; -x_32 = l_Lean_addMacroScope(x_30, x_31, x_25); -x_33 = l_Lean_Elab_Term_elabPanic___closed__3; +x_29 = lean_st_ref_get(x_8, x_27); +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_29, 1); +lean_inc(x_31); +lean_dec(x_29); +x_32 = lean_ctor_get(x_30, 0); +lean_inc(x_32); +lean_dec(x_30); +x_33 = lean_environment_main_module(x_32); x_34 = l_Lean_Elab_Term_elabPanic___closed__6; -x_35 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_35, 0, x_23); -lean_ctor_set(x_35, 1, x_33); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_35, 3, x_34); -x_36 = lean_environment_main_module(x_18); -x_37 = 1; -x_38 = l_Lean_Name_toString(x_36, x_37); -x_39 = lean_box(2); -x_40 = l_Lean_Syntax_mkStrLit(x_38, x_39); -lean_dec(x_38); -x_41 = lean_ctor_get(x_13, 0); -lean_inc(x_41); -x_42 = l_Nat_repr(x_41); -x_43 = l_Lean_numLitKind; -x_44 = l_Lean_Syntax_mkLit(x_43, x_42, x_39); -x_45 = lean_ctor_get(x_13, 1); -lean_inc(x_45); -lean_dec(x_13); -x_46 = l_Nat_repr(x_45); -x_47 = l_Lean_Syntax_mkLit(x_43, x_46, x_39); -x_48 = l_Lean_Elab_Term_expandShow___closed__31; -x_49 = lean_array_push(x_48, x_40); -x_50 = lean_array_push(x_49, x_44); -x_51 = lean_array_push(x_50, x_47); -x_52 = lean_array_push(x_51, x_11); -x_53 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; -x_54 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_54, 0, x_39); -lean_ctor_set(x_54, 1, x_53); -lean_ctor_set(x_54, 2, x_52); -x_55 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; -x_56 = lean_array_push(x_55, x_35); -x_57 = lean_array_push(x_56, x_54); -x_58 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__5; -x_59 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_59, 0, x_39); -lean_ctor_set(x_59, 1, x_58); -lean_ctor_set(x_59, 2, x_57); -x_60 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__2(x_2, x_1, x_59, x_3, x_4, x_5, x_6, x_7, x_8, x_28); -return x_60; -} -else -{ -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; 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; uint8_t x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; -x_61 = lean_ctor_get(x_19, 1); -lean_inc(x_61); -lean_dec(x_19); -x_62 = lean_ctor_get(x_20, 0); -lean_inc(x_62); -lean_dec(x_20); -lean_inc(x_7); -x_63 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_7, x_8, x_61); -x_64 = lean_ctor_get(x_63, 0); +x_35 = l_Lean_addMacroScope(x_33, x_34, x_28); +x_36 = l_Lean_Elab_Term_elabPanic___closed__5; +x_37 = l_Lean_Elab_Term_elabPanic___closed__8; +x_38 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_38, 0, x_26); +lean_ctor_set(x_38, 1, x_36); +lean_ctor_set(x_38, 2, x_35); +lean_ctor_set(x_38, 3, x_37); +x_39 = lean_environment_main_module(x_21); +x_40 = 1; +x_41 = l_Lean_Name_toString(x_39, x_40); +x_42 = lean_box(2); +x_43 = l_Lean_Syntax_mkStrLit(x_41, x_42); +lean_dec(x_41); +x_44 = lean_ctor_get(x_16, 0); +lean_inc(x_44); +x_45 = l_Nat_repr(x_44); +x_46 = l_Lean_Syntax_mkNumLit(x_45, x_42); +x_47 = lean_ctor_get(x_16, 1); +lean_inc(x_47); +lean_dec(x_16); +x_48 = l_Nat_repr(x_47); +x_49 = l_Lean_Syntax_mkNumLit(x_48, x_42); +x_50 = l_Lean_Elab_Term_expandShow___closed__31; +x_51 = lean_array_push(x_50, x_43); +x_52 = lean_array_push(x_51, x_46); +x_53 = lean_array_push(x_52, x_49); +x_54 = lean_array_push(x_53, x_14); +x_55 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; +x_56 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_56, 0, x_42); +lean_ctor_set(x_56, 1, x_55); +lean_ctor_set(x_56, 2, x_54); +x_57 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; +x_58 = lean_array_push(x_57, x_38); +x_59 = lean_array_push(x_58, x_56); +x_60 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__5; +x_61 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_61, 0, x_42); +lean_ctor_set(x_61, 1, x_60); +lean_ctor_set(x_61, 2, x_59); +x_62 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__2(x_2, x_1, x_61, x_3, x_4, x_5, x_6, x_7, x_8, x_31); +return x_62; +} +else +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; 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; lean_object* x_79; uint8_t x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; +x_63 = lean_ctor_get(x_22, 1); +lean_inc(x_63); +lean_dec(x_22); +x_64 = lean_ctor_get(x_23, 0); lean_inc(x_64); -x_65 = lean_ctor_get(x_63, 1); -lean_inc(x_65); -lean_dec(x_63); -x_66 = lean_ctor_get(x_7, 10); +lean_dec(x_23); +lean_inc(x_7); +x_65 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_7, x_8, x_63); +x_66 = lean_ctor_get(x_65, 0); lean_inc(x_66); -x_67 = lean_st_ref_get(x_8, x_65); -x_68 = lean_ctor_get(x_67, 0); +x_67 = lean_ctor_get(x_65, 1); +lean_inc(x_67); +lean_dec(x_65); +x_68 = lean_ctor_get(x_7, 10); lean_inc(x_68); -x_69 = lean_ctor_get(x_67, 1); -lean_inc(x_69); -lean_dec(x_67); -x_70 = lean_ctor_get(x_68, 0); +x_69 = lean_st_ref_get(x_8, x_67); +x_70 = lean_ctor_get(x_69, 0); lean_inc(x_70); -lean_dec(x_68); -x_71 = lean_environment_main_module(x_70); -x_72 = l_Lean_Elab_Term_elabPanic___closed__10; -x_73 = l_Lean_addMacroScope(x_71, x_72, x_66); -x_74 = l_Lean_Elab_Term_elabPanic___closed__9; -x_75 = l_Lean_Elab_Term_elabPanic___closed__12; -x_76 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_76, 0, x_64); -lean_ctor_set(x_76, 1, x_74); -lean_ctor_set(x_76, 2, x_73); -lean_ctor_set(x_76, 3, x_75); -x_77 = lean_environment_main_module(x_18); -x_78 = 1; -x_79 = l_Lean_Name_toString(x_77, x_78); -x_80 = lean_box(2); -x_81 = l_Lean_Syntax_mkStrLit(x_79, x_80); -lean_dec(x_79); -x_82 = l_Lean_Name_toString(x_62, x_78); -x_83 = l_Lean_Syntax_mkStrLit(x_82, x_80); -lean_dec(x_82); -x_84 = lean_ctor_get(x_13, 0); -lean_inc(x_84); -x_85 = l_Nat_repr(x_84); -x_86 = l_Lean_numLitKind; -x_87 = l_Lean_Syntax_mkLit(x_86, x_85, x_80); -x_88 = lean_ctor_get(x_13, 1); -lean_inc(x_88); -lean_dec(x_13); -x_89 = l_Nat_repr(x_88); -x_90 = l_Lean_Syntax_mkLit(x_86, x_89, x_80); -x_91 = l_Lean_Elab_Term_expandShow___closed__29; -x_92 = lean_array_push(x_91, x_81); +x_71 = lean_ctor_get(x_69, 1); +lean_inc(x_71); +lean_dec(x_69); +x_72 = lean_ctor_get(x_70, 0); +lean_inc(x_72); +lean_dec(x_70); +x_73 = lean_environment_main_module(x_72); +x_74 = l_Lean_Elab_Term_elabPanic___closed__12; +x_75 = l_Lean_addMacroScope(x_73, x_74, x_68); +x_76 = l_Lean_Elab_Term_elabPanic___closed__11; +x_77 = l_Lean_Elab_Term_elabPanic___closed__14; +x_78 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_78, 0, x_66); +lean_ctor_set(x_78, 1, x_76); +lean_ctor_set(x_78, 2, x_75); +lean_ctor_set(x_78, 3, x_77); +x_79 = lean_environment_main_module(x_21); +x_80 = 1; +x_81 = l_Lean_Name_toString(x_79, x_80); +x_82 = lean_box(2); +x_83 = l_Lean_Syntax_mkStrLit(x_81, x_82); +lean_dec(x_81); +x_84 = l_Lean_Name_toString(x_64, x_80); +x_85 = l_Lean_Syntax_mkStrLit(x_84, x_82); +lean_dec(x_84); +x_86 = lean_ctor_get(x_16, 0); +lean_inc(x_86); +x_87 = l_Nat_repr(x_86); +x_88 = l_Lean_Syntax_mkNumLit(x_87, x_82); +x_89 = lean_ctor_get(x_16, 1); +lean_inc(x_89); +lean_dec(x_16); +x_90 = l_Nat_repr(x_89); +x_91 = l_Lean_Syntax_mkNumLit(x_90, x_82); +x_92 = l_Lean_Elab_Term_expandShow___closed__29; x_93 = lean_array_push(x_92, x_83); -x_94 = lean_array_push(x_93, x_87); -x_95 = lean_array_push(x_94, x_90); -x_96 = lean_array_push(x_95, x_11); -x_97 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; -x_98 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_98, 0, x_80); -lean_ctor_set(x_98, 1, x_97); -lean_ctor_set(x_98, 2, x_96); -x_99 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; -x_100 = lean_array_push(x_99, x_76); -x_101 = lean_array_push(x_100, x_98); -x_102 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__5; -x_103 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_103, 0, x_80); -lean_ctor_set(x_103, 1, x_102); -lean_ctor_set(x_103, 2, x_101); -x_104 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__2(x_2, x_1, x_103, x_3, x_4, x_5, x_6, x_7, x_8, x_69); -return x_104; +x_94 = lean_array_push(x_93, x_85); +x_95 = lean_array_push(x_94, x_88); +x_96 = lean_array_push(x_95, x_91); +x_97 = lean_array_push(x_96, x_14); +x_98 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; +x_99 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_99, 0, x_82); +lean_ctor_set(x_99, 1, x_98); +lean_ctor_set(x_99, 2, x_97); +x_100 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; +x_101 = lean_array_push(x_100, x_78); +x_102 = lean_array_push(x_101, x_99); +x_103 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__5; +x_104 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_104, 0, x_82); +lean_ctor_set(x_104, 1, x_103); +lean_ctor_set(x_104, 2, x_102); +x_105 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__2(x_2, x_1, x_104, x_3, x_4, x_5, x_6, x_7, x_8, x_71); +return x_105; +} } } } @@ -9251,7 +9338,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabPanic___closed__1() _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("panic", 5); +x_1 = lean_mk_string_from_bytes("elabPanic", 9); return x_1; } } @@ -9259,7 +9346,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabPanic___closed__2() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; +x_1 = l___regBuiltin_Lean_Elab_Term_elabCoe___closed__8; x_2 = l___regBuiltin_Lean_Elab_Term_elabPanic___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -9269,24 +9356,6 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabPanic___closed__3() _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("elabPanic", 9); -return x_1; -} -} -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabPanic___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_elabCoe___closed__8; -x_2 = l___regBuiltin_Lean_Elab_Term_elabPanic___closed__3; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabPanic___closed__5() { -_start: -{ -lean_object* x_1; x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabPanic), 9, 0); return x_1; } @@ -9296,9 +9365,9 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabPanic(lean_object* x_ { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_2 = l___regBuiltin_Lean_Elab_Term_elabCoe___closed__11; -x_3 = l___regBuiltin_Lean_Elab_Term_elabPanic___closed__2; -x_4 = l___regBuiltin_Lean_Elab_Term_elabPanic___closed__4; -x_5 = l___regBuiltin_Lean_Elab_Term_elabPanic___closed__5; +x_3 = l_Lean_Elab_Term_elabPanic___closed__2; +x_4 = l___regBuiltin_Lean_Elab_Term_elabPanic___closed__2; +x_5 = l___regBuiltin_Lean_Elab_Term_elabPanic___closed__3; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); return x_6; } @@ -9315,7 +9384,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabPanic_docString(lean_ _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_elabPanic___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_elabPanic___closed__2; x_3 = l___regBuiltin_Lean_Elab_Term_elabPanic_docString___closed__1; x_4 = l_Lean_addBuiltinDocString(x_2, x_3, x_1); return x_4; @@ -9337,8 +9406,8 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabPanic_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(140u); -x_2 = lean_unsigned_to_nat(63u); +x_1 = lean_unsigned_to_nat(142u); +x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -9352,7 +9421,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Elab_Term_elabPanic_declRange___closed__1; x_2 = lean_unsigned_to_nat(42u); x_3 = l___regBuiltin_Lean_Elab_Term_elabPanic_declRange___closed__2; -x_4 = lean_unsigned_to_nat(63u); +x_4 = lean_unsigned_to_nat(31u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -9417,7 +9486,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabPanic_declRange(lean_ _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_elabPanic___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_elabPanic___closed__2; x_3 = l___regBuiltin_Lean_Elab_Term_elabPanic_declRange___closed__7; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; @@ -9487,7 +9556,7 @@ lean_ctor_set(x_14, 2, x_11); x_15 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; x_16 = lean_array_push(x_15, x_7); x_17 = lean_array_push(x_16, x_14); -x_18 = l___regBuiltin_Lean_Elab_Term_elabPanic___closed__2; +x_18 = l_Lean_Elab_Term_elabPanic___closed__2; x_19 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_19, 0, x_12); lean_ctor_set(x_19, 1, x_18); @@ -9523,7 +9592,7 @@ lean_ctor_set(x_30, 2, x_27); x_31 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; x_32 = lean_array_push(x_31, x_23); x_33 = lean_array_push(x_32, x_30); -x_34 = l___regBuiltin_Lean_Elab_Term_elabPanic___closed__2; +x_34 = l_Lean_Elab_Term_elabPanic___closed__2; x_35 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_35, 0, x_28); lean_ctor_set(x_35, 1, x_34); @@ -9630,7 +9699,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandUnreachable_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(143u); +x_1 = lean_unsigned_to_nat(145u); x_2 = lean_unsigned_to_nat(46u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9642,7 +9711,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandUnreachable_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(144u); +x_1 = lean_unsigned_to_nat(146u); x_2 = lean_unsigned_to_nat(47u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9670,7 +9739,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandUnreachable_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(143u); +x_1 = lean_unsigned_to_nat(145u); x_2 = lean_unsigned_to_nat(50u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9682,7 +9751,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandUnreachable_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(143u); +x_1 = lean_unsigned_to_nat(145u); x_2 = lean_unsigned_to_nat(67u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9732,7 +9801,7 @@ static lean_object* _init_l_Lean_Elab_Term_expandAssert___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("termIfThenElse", 14); +x_1 = lean_mk_string_from_bytes("assert", 6); return x_1; } } @@ -9740,7 +9809,7 @@ static lean_object* _init_l_Lean_Elab_Term_expandAssert___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); +x_1 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; x_2 = l_Lean_Elab_Term_expandAssert___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -9750,19 +9819,37 @@ static lean_object* _init_l_Lean_Elab_Term_expandAssert___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("if", 2); +x_1 = lean_mk_string_from_bytes("termIfThenElse", 14); return x_1; } } static lean_object* _init_l_Lean_Elab_Term_expandAssert___closed__4() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_expandAssert___closed__3; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_expandAssert___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("if", 2); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_expandAssert___closed__6() { +_start: +{ lean_object* x_1; x_1 = lean_mk_string_from_bytes("then", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_expandAssert___closed__5() { +static lean_object* _init_l_Lean_Elab_Term_expandAssert___closed__7() { _start: { lean_object* x_1; @@ -9770,7 +9857,7 @@ x_1 = lean_mk_string_from_bytes("else", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_expandAssert___closed__6() { +static lean_object* _init_l_Lean_Elab_Term_expandAssert___closed__8() { _start: { lean_object* x_1; @@ -9778,7 +9865,7 @@ x_1 = lean_mk_string_from_bytes("\"assertion violation\"", 21); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_expandAssert___closed__7() { +static lean_object* _init_l_Lean_Elab_Term_expandAssert___closed__9() { _start: { lean_object* x_1; lean_object* x_2; @@ -9787,7 +9874,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_expandAssert___closed__8() { +static lean_object* _init_l_Lean_Elab_Term_expandAssert___closed__10() { _start: { lean_object* x_1; @@ -9795,17 +9882,17 @@ x_1 = lean_mk_string_from_bytes("term_++_", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_expandAssert___closed__9() { +static lean_object* _init_l_Lean_Elab_Term_expandAssert___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_expandAssert___closed__8; +x_2 = l_Lean_Elab_Term_expandAssert___closed__10; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_expandAssert___closed__10() { +static lean_object* _init_l_Lean_Elab_Term_expandAssert___closed__12() { _start: { lean_object* x_1; @@ -9813,7 +9900,7 @@ x_1 = lean_mk_string_from_bytes("\"assertion violation: \"", 23); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_expandAssert___closed__11() { +static lean_object* _init_l_Lean_Elab_Term_expandAssert___closed__13() { _start: { lean_object* x_1; @@ -9824,430 +9911,440 @@ return x_1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandAssert(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_4 = lean_unsigned_to_nat(1u); -x_5 = l_Lean_Syntax_getArg(x_1, x_4); -x_6 = lean_unsigned_to_nat(3u); -x_7 = l_Lean_Syntax_getArg(x_1, x_6); -lean_inc(x_5); -x_8 = l_Lean_Syntax_reprint(x_5); -if (lean_obj_tag(x_8) == 0) +lean_object* x_4; uint8_t x_5; +x_4 = l_Lean_Elab_Term_expandAssert___closed__2; +lean_inc(x_1); +x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); +if (x_5 == 0) { -lean_object* x_9; uint8_t x_10; -x_9 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 0) +lean_object* x_6; lean_object* x_7; +lean_dec(x_2); +lean_dec(x_1); +x_6 = lean_box(1); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_3); +return x_7; +} +else { -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; 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_11 = lean_ctor_get(x_9, 0); -x_12 = l_Lean_Elab_Term_expandAssert___closed__3; -lean_inc(x_11); -x_13 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_13, 0, x_11); -lean_ctor_set(x_13, 1, x_12); -x_14 = l_Lean_Elab_Term_expandAssert___closed__4; -lean_inc(x_11); -x_15 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_15, 0, x_11); -lean_ctor_set(x_15, 1, x_14); +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_8 = lean_unsigned_to_nat(1u); +x_9 = l_Lean_Syntax_getArg(x_1, x_8); +x_10 = lean_unsigned_to_nat(3u); +x_11 = l_Lean_Syntax_getArg(x_1, x_10); +lean_dec(x_1); +lean_inc(x_9); +x_12 = l_Lean_Syntax_reprint(x_9); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; uint8_t x_14; +x_13 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_14 = !lean_is_exclusive(x_13); +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; lean_object* 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_15 = lean_ctor_get(x_13, 0); x_16 = l_Lean_Elab_Term_expandAssert___closed__5; -lean_inc(x_11); +lean_inc(x_15); x_17 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_17, 0, x_11); +lean_ctor_set(x_17, 0, x_15); lean_ctor_set(x_17, 1, x_16); -x_18 = l_Lean_Elab_Term_expandUnreachable___rarg___closed__1; -lean_inc(x_11); +x_18 = l_Lean_Elab_Term_expandAssert___closed__6; +lean_inc(x_15); x_19 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_19, 0, x_11); +lean_ctor_set(x_19, 0, x_15); lean_ctor_set(x_19, 1, x_18); -x_20 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; -lean_inc(x_11); +x_20 = l_Lean_Elab_Term_expandAssert___closed__7; +lean_inc(x_15); x_21 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_21, 0, x_11); +lean_ctor_set(x_21, 0, x_15); lean_ctor_set(x_21, 1, x_20); -x_22 = l_Lean_Elab_Term_expandAssert___closed__6; -lean_inc(x_11); +x_22 = l_Lean_Elab_Term_expandUnreachable___rarg___closed__1; +lean_inc(x_15); x_23 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_23, 0, x_11); +lean_ctor_set(x_23, 0, x_15); lean_ctor_set(x_23, 1, x_22); -x_24 = l_Lean_Elab_Term_expandShow___closed__27; -x_25 = lean_array_push(x_24, x_23); -x_26 = lean_box(2); -x_27 = l_Lean_Elab_Term_expandUnreachable___rarg___closed__3; -x_28 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -lean_ctor_set(x_28, 2, x_25); -x_29 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; -x_30 = lean_array_push(x_29, x_28); -x_31 = l_Lean_Elab_Term_expandShow___closed__23; -x_32 = lean_array_push(x_30, x_31); -x_33 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; -x_34 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_34, 0, x_26); -lean_ctor_set(x_34, 1, x_33); -lean_ctor_set(x_34, 2, x_32); -x_35 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__23; -x_36 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_36, 0, x_11); -lean_ctor_set(x_36, 1, x_35); -x_37 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; -x_38 = lean_array_push(x_37, x_21); -x_39 = lean_array_push(x_38, x_34); -x_40 = lean_array_push(x_39, x_36); -x_41 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__14; -x_42 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_42, 0, x_26); -lean_ctor_set(x_42, 1, x_41); -lean_ctor_set(x_42, 2, x_40); -x_43 = lean_array_push(x_29, x_19); -x_44 = lean_array_push(x_43, x_42); -x_45 = l___regBuiltin_Lean_Elab_Term_elabPanic___closed__2; +x_24 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; +lean_inc(x_15); +x_25 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_25, 0, x_15); +lean_ctor_set(x_25, 1, x_24); +x_26 = l_Lean_Elab_Term_expandAssert___closed__8; +lean_inc(x_15); +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_15); +lean_ctor_set(x_27, 1, x_26); +x_28 = l_Lean_Elab_Term_expandShow___closed__27; +x_29 = lean_array_push(x_28, x_27); +x_30 = lean_box(2); +x_31 = l_Lean_Elab_Term_expandUnreachable___rarg___closed__3; +x_32 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set(x_32, 2, x_29); +x_33 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; +x_34 = lean_array_push(x_33, x_32); +x_35 = l_Lean_Elab_Term_expandShow___closed__23; +x_36 = lean_array_push(x_34, x_35); +x_37 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; +x_38 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_38, 0, x_30); +lean_ctor_set(x_38, 1, x_37); +lean_ctor_set(x_38, 2, x_36); +x_39 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__23; +x_40 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_40, 0, x_15); +lean_ctor_set(x_40, 1, x_39); +x_41 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; +x_42 = lean_array_push(x_41, x_25); +x_43 = lean_array_push(x_42, x_38); +x_44 = lean_array_push(x_43, x_40); +x_45 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__14; x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_26); +lean_ctor_set(x_46, 0, x_30); lean_ctor_set(x_46, 1, x_45); lean_ctor_set(x_46, 2, x_44); -x_47 = l_Lean_Elab_Term_expandAssert___closed__7; -x_48 = lean_array_push(x_47, x_13); -x_49 = lean_array_push(x_48, x_5); -x_50 = lean_array_push(x_49, x_15); -x_51 = lean_array_push(x_50, x_7); +x_47 = lean_array_push(x_33, x_23); +x_48 = lean_array_push(x_47, x_46); +x_49 = l_Lean_Elab_Term_elabPanic___closed__2; +x_50 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_50, 0, x_30); +lean_ctor_set(x_50, 1, x_49); +lean_ctor_set(x_50, 2, x_48); +x_51 = l_Lean_Elab_Term_expandAssert___closed__9; x_52 = lean_array_push(x_51, x_17); -x_53 = lean_array_push(x_52, x_46); -x_54 = l_Lean_Elab_Term_expandAssert___closed__2; -x_55 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_55, 0, x_26); -lean_ctor_set(x_55, 1, x_54); -lean_ctor_set(x_55, 2, x_53); -lean_ctor_set(x_9, 0, x_55); -return x_9; +x_53 = lean_array_push(x_52, x_9); +x_54 = lean_array_push(x_53, x_19); +x_55 = lean_array_push(x_54, x_11); +x_56 = lean_array_push(x_55, x_21); +x_57 = lean_array_push(x_56, x_50); +x_58 = l_Lean_Elab_Term_expandAssert___closed__4; +x_59 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_59, 0, x_30); +lean_ctor_set(x_59, 1, x_58); +lean_ctor_set(x_59, 2, x_57); +lean_ctor_set(x_13, 0, x_59); +return x_13; } else { -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; -x_56 = lean_ctor_get(x_9, 0); -x_57 = lean_ctor_get(x_9, 1); -lean_inc(x_57); -lean_inc(x_56); -lean_dec(x_9); -x_58 = l_Lean_Elab_Term_expandAssert___closed__3; -lean_inc(x_56); -x_59 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_59, 0, x_56); -lean_ctor_set(x_59, 1, x_58); -x_60 = l_Lean_Elab_Term_expandAssert___closed__4; -lean_inc(x_56); -x_61 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_61, 0, x_56); -lean_ctor_set(x_61, 1, x_60); +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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_60 = lean_ctor_get(x_13, 0); +x_61 = lean_ctor_get(x_13, 1); +lean_inc(x_61); +lean_inc(x_60); +lean_dec(x_13); x_62 = l_Lean_Elab_Term_expandAssert___closed__5; -lean_inc(x_56); +lean_inc(x_60); x_63 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_63, 0, x_56); +lean_ctor_set(x_63, 0, x_60); lean_ctor_set(x_63, 1, x_62); -x_64 = l_Lean_Elab_Term_expandUnreachable___rarg___closed__1; -lean_inc(x_56); +x_64 = l_Lean_Elab_Term_expandAssert___closed__6; +lean_inc(x_60); x_65 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_65, 0, x_56); +lean_ctor_set(x_65, 0, x_60); lean_ctor_set(x_65, 1, x_64); -x_66 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; -lean_inc(x_56); +x_66 = l_Lean_Elab_Term_expandAssert___closed__7; +lean_inc(x_60); x_67 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_67, 0, x_56); +lean_ctor_set(x_67, 0, x_60); lean_ctor_set(x_67, 1, x_66); -x_68 = l_Lean_Elab_Term_expandAssert___closed__6; -lean_inc(x_56); +x_68 = l_Lean_Elab_Term_expandUnreachable___rarg___closed__1; +lean_inc(x_60); x_69 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_69, 0, x_56); +lean_ctor_set(x_69, 0, x_60); lean_ctor_set(x_69, 1, x_68); -x_70 = l_Lean_Elab_Term_expandShow___closed__27; -x_71 = lean_array_push(x_70, x_69); -x_72 = lean_box(2); -x_73 = l_Lean_Elab_Term_expandUnreachable___rarg___closed__3; -x_74 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_74, 0, x_72); -lean_ctor_set(x_74, 1, x_73); -lean_ctor_set(x_74, 2, x_71); -x_75 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; -x_76 = lean_array_push(x_75, x_74); -x_77 = l_Lean_Elab_Term_expandShow___closed__23; -x_78 = lean_array_push(x_76, x_77); -x_79 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; -x_80 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_80, 0, x_72); -lean_ctor_set(x_80, 1, x_79); -lean_ctor_set(x_80, 2, x_78); -x_81 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__23; -x_82 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_82, 0, x_56); -lean_ctor_set(x_82, 1, x_81); -x_83 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; -x_84 = lean_array_push(x_83, x_67); -x_85 = lean_array_push(x_84, x_80); -x_86 = lean_array_push(x_85, x_82); -x_87 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__14; -x_88 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_88, 0, x_72); -lean_ctor_set(x_88, 1, x_87); -lean_ctor_set(x_88, 2, x_86); -x_89 = lean_array_push(x_75, x_65); -x_90 = lean_array_push(x_89, x_88); -x_91 = l___regBuiltin_Lean_Elab_Term_elabPanic___closed__2; +x_70 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; +lean_inc(x_60); +x_71 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_71, 0, x_60); +lean_ctor_set(x_71, 1, x_70); +x_72 = l_Lean_Elab_Term_expandAssert___closed__8; +lean_inc(x_60); +x_73 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_73, 0, x_60); +lean_ctor_set(x_73, 1, x_72); +x_74 = l_Lean_Elab_Term_expandShow___closed__27; +x_75 = lean_array_push(x_74, x_73); +x_76 = lean_box(2); +x_77 = l_Lean_Elab_Term_expandUnreachable___rarg___closed__3; +x_78 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_78, 0, x_76); +lean_ctor_set(x_78, 1, x_77); +lean_ctor_set(x_78, 2, x_75); +x_79 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; +x_80 = lean_array_push(x_79, x_78); +x_81 = l_Lean_Elab_Term_expandShow___closed__23; +x_82 = lean_array_push(x_80, x_81); +x_83 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; +x_84 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_84, 0, x_76); +lean_ctor_set(x_84, 1, x_83); +lean_ctor_set(x_84, 2, x_82); +x_85 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__23; +x_86 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_86, 0, x_60); +lean_ctor_set(x_86, 1, x_85); +x_87 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; +x_88 = lean_array_push(x_87, x_71); +x_89 = lean_array_push(x_88, x_84); +x_90 = lean_array_push(x_89, x_86); +x_91 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__14; x_92 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_92, 0, x_72); +lean_ctor_set(x_92, 0, x_76); lean_ctor_set(x_92, 1, x_91); lean_ctor_set(x_92, 2, x_90); -x_93 = l_Lean_Elab_Term_expandAssert___closed__7; -x_94 = lean_array_push(x_93, x_59); -x_95 = lean_array_push(x_94, x_5); -x_96 = lean_array_push(x_95, x_61); -x_97 = lean_array_push(x_96, x_7); +x_93 = lean_array_push(x_79, x_69); +x_94 = lean_array_push(x_93, x_92); +x_95 = l_Lean_Elab_Term_elabPanic___closed__2; +x_96 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_96, 0, x_76); +lean_ctor_set(x_96, 1, x_95); +lean_ctor_set(x_96, 2, x_94); +x_97 = l_Lean_Elab_Term_expandAssert___closed__9; x_98 = lean_array_push(x_97, x_63); -x_99 = lean_array_push(x_98, x_92); -x_100 = l_Lean_Elab_Term_expandAssert___closed__2; -x_101 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_101, 0, x_72); -lean_ctor_set(x_101, 1, x_100); -lean_ctor_set(x_101, 2, x_99); -x_102 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_102, 0, x_101); -lean_ctor_set(x_102, 1, x_57); -return x_102; +x_99 = lean_array_push(x_98, x_9); +x_100 = lean_array_push(x_99, x_65); +x_101 = lean_array_push(x_100, x_11); +x_102 = lean_array_push(x_101, x_67); +x_103 = lean_array_push(x_102, x_96); +x_104 = l_Lean_Elab_Term_expandAssert___closed__4; +x_105 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_105, 0, x_76); +lean_ctor_set(x_105, 1, x_104); +lean_ctor_set(x_105, 2, x_103); +x_106 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_106, 0, x_105); +lean_ctor_set(x_106, 1, x_61); +return x_106; } } else { -lean_object* x_103; lean_object* x_104; uint8_t x_105; -x_103 = lean_ctor_get(x_8, 0); -lean_inc(x_103); -lean_dec(x_8); -x_104 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_105 = !lean_is_exclusive(x_104); -if (x_105 == 0) +lean_object* x_107; lean_object* x_108; uint8_t x_109; +x_107 = lean_ctor_get(x_12, 0); +lean_inc(x_107); +lean_dec(x_12); +x_108 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_109 = !lean_is_exclusive(x_108); +if (x_109 == 0) { -lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; -x_106 = lean_ctor_get(x_104, 0); -x_107 = l_Lean_Elab_Term_expandAssert___closed__3; -lean_inc(x_106); -x_108 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_108, 0, x_106); -lean_ctor_set(x_108, 1, x_107); -x_109 = l_Lean_Elab_Term_expandAssert___closed__4; -lean_inc(x_106); -x_110 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_110, 0, x_106); -lean_ctor_set(x_110, 1, x_109); +lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; +x_110 = lean_ctor_get(x_108, 0); x_111 = l_Lean_Elab_Term_expandAssert___closed__5; -lean_inc(x_106); +lean_inc(x_110); x_112 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_112, 0, x_106); +lean_ctor_set(x_112, 0, x_110); lean_ctor_set(x_112, 1, x_111); -x_113 = l_Lean_Elab_Term_expandUnreachable___rarg___closed__1; -lean_inc(x_106); +x_113 = l_Lean_Elab_Term_expandAssert___closed__6; +lean_inc(x_110); x_114 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_114, 0, x_106); +lean_ctor_set(x_114, 0, x_110); lean_ctor_set(x_114, 1, x_113); -x_115 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; -lean_inc(x_106); +x_115 = l_Lean_Elab_Term_expandAssert___closed__7; +lean_inc(x_110); x_116 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_116, 0, x_106); +lean_ctor_set(x_116, 0, x_110); lean_ctor_set(x_116, 1, x_115); -x_117 = l_Lean_Elab_Term_expandAssert___closed__10; -lean_inc(x_106); +x_117 = l_Lean_Elab_Term_expandUnreachable___rarg___closed__1; +lean_inc(x_110); x_118 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_118, 0, x_106); +lean_ctor_set(x_118, 0, x_110); lean_ctor_set(x_118, 1, x_117); -x_119 = l_Lean_Elab_Term_expandShow___closed__27; -x_120 = lean_array_push(x_119, x_118); -x_121 = lean_box(2); -x_122 = l_Lean_Elab_Term_expandUnreachable___rarg___closed__3; -x_123 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_123, 0, x_121); -lean_ctor_set(x_123, 1, x_122); -lean_ctor_set(x_123, 2, x_120); -x_124 = l_Lean_Elab_Term_expandAssert___closed__11; -lean_inc(x_106); -x_125 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_125, 0, x_106); -lean_ctor_set(x_125, 1, x_124); -x_126 = l_Lean_Syntax_mkStrLit(x_103, x_121); -lean_dec(x_103); -x_127 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; -x_128 = lean_array_push(x_127, x_123); -x_129 = lean_array_push(x_128, x_125); -x_130 = lean_array_push(x_129, x_126); -x_131 = l_Lean_Elab_Term_expandAssert___closed__9; -x_132 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_132, 0, x_121); -lean_ctor_set(x_132, 1, x_131); -lean_ctor_set(x_132, 2, x_130); -x_133 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; -x_134 = lean_array_push(x_133, x_132); -x_135 = l_Lean_Elab_Term_expandShow___closed__23; -x_136 = lean_array_push(x_134, x_135); -x_137 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; -x_138 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_138, 0, x_121); -lean_ctor_set(x_138, 1, x_137); -lean_ctor_set(x_138, 2, x_136); -x_139 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__23; -x_140 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_140, 0, x_106); -lean_ctor_set(x_140, 1, x_139); -x_141 = lean_array_push(x_127, x_116); -x_142 = lean_array_push(x_141, x_138); -x_143 = lean_array_push(x_142, x_140); -x_144 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__14; -x_145 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_145, 0, x_121); -lean_ctor_set(x_145, 1, x_144); -lean_ctor_set(x_145, 2, x_143); -x_146 = lean_array_push(x_133, x_114); -x_147 = lean_array_push(x_146, x_145); -x_148 = l___regBuiltin_Lean_Elab_Term_elabPanic___closed__2; +x_119 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; +lean_inc(x_110); +x_120 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_120, 0, x_110); +lean_ctor_set(x_120, 1, x_119); +x_121 = l_Lean_Elab_Term_expandAssert___closed__12; +lean_inc(x_110); +x_122 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_122, 0, x_110); +lean_ctor_set(x_122, 1, x_121); +x_123 = l_Lean_Elab_Term_expandShow___closed__27; +x_124 = lean_array_push(x_123, x_122); +x_125 = lean_box(2); +x_126 = l_Lean_Elab_Term_expandUnreachable___rarg___closed__3; +x_127 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +lean_ctor_set(x_127, 2, x_124); +x_128 = l_Lean_Elab_Term_expandAssert___closed__13; +lean_inc(x_110); +x_129 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_129, 0, x_110); +lean_ctor_set(x_129, 1, x_128); +x_130 = l_Lean_Syntax_mkStrLit(x_107, x_125); +lean_dec(x_107); +x_131 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; +x_132 = lean_array_push(x_131, x_127); +x_133 = lean_array_push(x_132, x_129); +x_134 = lean_array_push(x_133, x_130); +x_135 = l_Lean_Elab_Term_expandAssert___closed__11; +x_136 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_136, 0, x_125); +lean_ctor_set(x_136, 1, x_135); +lean_ctor_set(x_136, 2, x_134); +x_137 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; +x_138 = lean_array_push(x_137, x_136); +x_139 = l_Lean_Elab_Term_expandShow___closed__23; +x_140 = lean_array_push(x_138, x_139); +x_141 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; +x_142 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_142, 0, x_125); +lean_ctor_set(x_142, 1, x_141); +lean_ctor_set(x_142, 2, x_140); +x_143 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__23; +x_144 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_144, 0, x_110); +lean_ctor_set(x_144, 1, x_143); +x_145 = lean_array_push(x_131, x_120); +x_146 = lean_array_push(x_145, x_142); +x_147 = lean_array_push(x_146, x_144); +x_148 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__14; x_149 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_149, 0, x_121); +lean_ctor_set(x_149, 0, x_125); lean_ctor_set(x_149, 1, x_148); lean_ctor_set(x_149, 2, x_147); -x_150 = l_Lean_Elab_Term_expandAssert___closed__7; -x_151 = lean_array_push(x_150, x_108); -x_152 = lean_array_push(x_151, x_5); -x_153 = lean_array_push(x_152, x_110); -x_154 = lean_array_push(x_153, x_7); +x_150 = lean_array_push(x_137, x_118); +x_151 = lean_array_push(x_150, x_149); +x_152 = l_Lean_Elab_Term_elabPanic___closed__2; +x_153 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_153, 0, x_125); +lean_ctor_set(x_153, 1, x_152); +lean_ctor_set(x_153, 2, x_151); +x_154 = l_Lean_Elab_Term_expandAssert___closed__9; x_155 = lean_array_push(x_154, x_112); -x_156 = lean_array_push(x_155, x_149); -x_157 = l_Lean_Elab_Term_expandAssert___closed__2; -x_158 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_158, 0, x_121); -lean_ctor_set(x_158, 1, x_157); -lean_ctor_set(x_158, 2, x_156); -lean_ctor_set(x_104, 0, x_158); -return x_104; +x_156 = lean_array_push(x_155, x_9); +x_157 = lean_array_push(x_156, x_114); +x_158 = lean_array_push(x_157, x_11); +x_159 = lean_array_push(x_158, x_116); +x_160 = lean_array_push(x_159, x_153); +x_161 = l_Lean_Elab_Term_expandAssert___closed__4; +x_162 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_162, 0, x_125); +lean_ctor_set(x_162, 1, x_161); +lean_ctor_set(x_162, 2, x_160); +lean_ctor_set(x_108, 0, x_162); +return x_108; } else { -lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; -x_159 = lean_ctor_get(x_104, 0); -x_160 = lean_ctor_get(x_104, 1); -lean_inc(x_160); -lean_inc(x_159); -lean_dec(x_104); -x_161 = l_Lean_Elab_Term_expandAssert___closed__3; -lean_inc(x_159); -x_162 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_162, 0, x_159); -lean_ctor_set(x_162, 1, x_161); -x_163 = l_Lean_Elab_Term_expandAssert___closed__4; -lean_inc(x_159); -x_164 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_164, 0, x_159); -lean_ctor_set(x_164, 1, x_163); +lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; +x_163 = lean_ctor_get(x_108, 0); +x_164 = lean_ctor_get(x_108, 1); +lean_inc(x_164); +lean_inc(x_163); +lean_dec(x_108); x_165 = l_Lean_Elab_Term_expandAssert___closed__5; -lean_inc(x_159); +lean_inc(x_163); x_166 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_166, 0, x_159); +lean_ctor_set(x_166, 0, x_163); lean_ctor_set(x_166, 1, x_165); -x_167 = l_Lean_Elab_Term_expandUnreachable___rarg___closed__1; -lean_inc(x_159); +x_167 = l_Lean_Elab_Term_expandAssert___closed__6; +lean_inc(x_163); x_168 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_168, 0, x_159); +lean_ctor_set(x_168, 0, x_163); lean_ctor_set(x_168, 1, x_167); -x_169 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; -lean_inc(x_159); +x_169 = l_Lean_Elab_Term_expandAssert___closed__7; +lean_inc(x_163); x_170 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_170, 0, x_159); +lean_ctor_set(x_170, 0, x_163); lean_ctor_set(x_170, 1, x_169); -x_171 = l_Lean_Elab_Term_expandAssert___closed__10; -lean_inc(x_159); +x_171 = l_Lean_Elab_Term_expandUnreachable___rarg___closed__1; +lean_inc(x_163); x_172 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_172, 0, x_159); +lean_ctor_set(x_172, 0, x_163); lean_ctor_set(x_172, 1, x_171); -x_173 = l_Lean_Elab_Term_expandShow___closed__27; -x_174 = lean_array_push(x_173, x_172); -x_175 = lean_box(2); -x_176 = l_Lean_Elab_Term_expandUnreachable___rarg___closed__3; -x_177 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_177, 0, x_175); -lean_ctor_set(x_177, 1, x_176); -lean_ctor_set(x_177, 2, x_174); -x_178 = l_Lean_Elab_Term_expandAssert___closed__11; -lean_inc(x_159); -x_179 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_179, 0, x_159); -lean_ctor_set(x_179, 1, x_178); -x_180 = l_Lean_Syntax_mkStrLit(x_103, x_175); -lean_dec(x_103); -x_181 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; -x_182 = lean_array_push(x_181, x_177); -x_183 = lean_array_push(x_182, x_179); -x_184 = lean_array_push(x_183, x_180); -x_185 = l_Lean_Elab_Term_expandAssert___closed__9; -x_186 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_186, 0, x_175); -lean_ctor_set(x_186, 1, x_185); -lean_ctor_set(x_186, 2, x_184); -x_187 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; -x_188 = lean_array_push(x_187, x_186); -x_189 = l_Lean_Elab_Term_expandShow___closed__23; -x_190 = lean_array_push(x_188, x_189); -x_191 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; -x_192 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_192, 0, x_175); -lean_ctor_set(x_192, 1, x_191); -lean_ctor_set(x_192, 2, x_190); -x_193 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__23; -x_194 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_194, 0, x_159); -lean_ctor_set(x_194, 1, x_193); -x_195 = lean_array_push(x_181, x_170); -x_196 = lean_array_push(x_195, x_192); -x_197 = lean_array_push(x_196, x_194); -x_198 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__14; -x_199 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_199, 0, x_175); -lean_ctor_set(x_199, 1, x_198); -lean_ctor_set(x_199, 2, x_197); -x_200 = lean_array_push(x_187, x_168); -x_201 = lean_array_push(x_200, x_199); -x_202 = l___regBuiltin_Lean_Elab_Term_elabPanic___closed__2; +x_173 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; +lean_inc(x_163); +x_174 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_174, 0, x_163); +lean_ctor_set(x_174, 1, x_173); +x_175 = l_Lean_Elab_Term_expandAssert___closed__12; +lean_inc(x_163); +x_176 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_176, 0, x_163); +lean_ctor_set(x_176, 1, x_175); +x_177 = l_Lean_Elab_Term_expandShow___closed__27; +x_178 = lean_array_push(x_177, x_176); +x_179 = lean_box(2); +x_180 = l_Lean_Elab_Term_expandUnreachable___rarg___closed__3; +x_181 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_181, 0, x_179); +lean_ctor_set(x_181, 1, x_180); +lean_ctor_set(x_181, 2, x_178); +x_182 = l_Lean_Elab_Term_expandAssert___closed__13; +lean_inc(x_163); +x_183 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_183, 0, x_163); +lean_ctor_set(x_183, 1, x_182); +x_184 = l_Lean_Syntax_mkStrLit(x_107, x_179); +lean_dec(x_107); +x_185 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; +x_186 = lean_array_push(x_185, x_181); +x_187 = lean_array_push(x_186, x_183); +x_188 = lean_array_push(x_187, x_184); +x_189 = l_Lean_Elab_Term_expandAssert___closed__11; +x_190 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_190, 0, x_179); +lean_ctor_set(x_190, 1, x_189); +lean_ctor_set(x_190, 2, x_188); +x_191 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; +x_192 = lean_array_push(x_191, x_190); +x_193 = l_Lean_Elab_Term_expandShow___closed__23; +x_194 = lean_array_push(x_192, x_193); +x_195 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; +x_196 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_196, 0, x_179); +lean_ctor_set(x_196, 1, x_195); +lean_ctor_set(x_196, 2, x_194); +x_197 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__23; +x_198 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_198, 0, x_163); +lean_ctor_set(x_198, 1, x_197); +x_199 = lean_array_push(x_185, x_174); +x_200 = lean_array_push(x_199, x_196); +x_201 = lean_array_push(x_200, x_198); +x_202 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__14; x_203 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_203, 0, x_175); +lean_ctor_set(x_203, 0, x_179); lean_ctor_set(x_203, 1, x_202); lean_ctor_set(x_203, 2, x_201); -x_204 = l_Lean_Elab_Term_expandAssert___closed__7; -x_205 = lean_array_push(x_204, x_162); -x_206 = lean_array_push(x_205, x_5); -x_207 = lean_array_push(x_206, x_164); -x_208 = lean_array_push(x_207, x_7); +x_204 = lean_array_push(x_191, x_172); +x_205 = lean_array_push(x_204, x_203); +x_206 = l_Lean_Elab_Term_elabPanic___closed__2; +x_207 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_207, 0, x_179); +lean_ctor_set(x_207, 1, x_206); +lean_ctor_set(x_207, 2, x_205); +x_208 = l_Lean_Elab_Term_expandAssert___closed__9; x_209 = lean_array_push(x_208, x_166); -x_210 = lean_array_push(x_209, x_203); -x_211 = l_Lean_Elab_Term_expandAssert___closed__2; -x_212 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_212, 0, x_175); -lean_ctor_set(x_212, 1, x_211); -lean_ctor_set(x_212, 2, x_210); -x_213 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_213, 0, x_212); -lean_ctor_set(x_213, 1, x_160); -return x_213; -} +x_210 = lean_array_push(x_209, x_9); +x_211 = lean_array_push(x_210, x_168); +x_212 = lean_array_push(x_211, x_11); +x_213 = lean_array_push(x_212, x_170); +x_214 = lean_array_push(x_213, x_207); +x_215 = l_Lean_Elab_Term_expandAssert___closed__4; +x_216 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_216, 0, x_179); +lean_ctor_set(x_216, 1, x_215); +lean_ctor_set(x_216, 2, x_214); +x_217 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_217, 0, x_216); +lean_ctor_set(x_217, 1, x_164); +return x_217; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandAssert___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = l_Lean_Elab_Term_expandAssert(x_1, x_2, x_3); -lean_dec(x_1); -return x_4; } } static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandAssert___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("assert", 6); +x_1 = lean_mk_string_from_bytes("expandAssert", 12); return x_1; } } @@ -10255,7 +10352,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandAssert___closed__2 _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; +x_1 = l___regBuiltin_Lean_Elab_Term_elabCoe___closed__8; x_2 = l___regBuiltin_Lean_Elab_Term_expandAssert___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -10265,25 +10362,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandAssert___closed__3 _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("expandAssert", 12); -return x_1; -} -} -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandAssert___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_elabCoe___closed__8; -x_2 = l___regBuiltin_Lean_Elab_Term_expandAssert___closed__3; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandAssert___closed__5() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_expandAssert___boxed), 3, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_expandAssert), 3, 0); return x_1; } } @@ -10292,9 +10371,9 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_expandAssert(lean_object* { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_2 = l___regBuiltin_Lean_Elab_Term_expandShow___closed__3; -x_3 = l___regBuiltin_Lean_Elab_Term_expandAssert___closed__2; -x_4 = l___regBuiltin_Lean_Elab_Term_expandAssert___closed__4; -x_5 = l___regBuiltin_Lean_Elab_Term_expandAssert___closed__5; +x_3 = l_Lean_Elab_Term_expandAssert___closed__2; +x_4 = l___regBuiltin_Lean_Elab_Term_expandAssert___closed__2; +x_5 = l___regBuiltin_Lean_Elab_Term_expandAssert___closed__3; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); return x_6; } @@ -10311,7 +10390,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_expandAssert_docString(le _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_expandAssert___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_expandAssert___closed__2; x_3 = l___regBuiltin_Lean_Elab_Term_expandAssert_docString___closed__1; x_4 = l_Lean_addBuiltinDocString(x_2, x_3, x_1); return x_4; @@ -10321,7 +10400,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandAssert_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(147u); +x_1 = lean_unsigned_to_nat(149u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10333,8 +10412,8 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandAssert_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(153u); -x_2 = lean_unsigned_to_nat(70u); +x_1 = lean_unsigned_to_nat(155u); +x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -10348,7 +10427,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Elab_Term_expandAssert_declRange___closed__1; x_2 = lean_unsigned_to_nat(41u); x_3 = l___regBuiltin_Lean_Elab_Term_expandAssert_declRange___closed__2; -x_4 = lean_unsigned_to_nat(70u); +x_4 = lean_unsigned_to_nat(31u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -10361,7 +10440,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandAssert_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(147u); +x_1 = lean_unsigned_to_nat(149u); x_2 = lean_unsigned_to_nat(45u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10373,7 +10452,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandAssert_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(147u); +x_1 = lean_unsigned_to_nat(149u); x_2 = lean_unsigned_to_nat(57u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10413,7 +10492,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_expandAssert_declRange(le _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_expandAssert___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_expandAssert___closed__2; x_3 = l___regBuiltin_Lean_Elab_Term_expandAssert_declRange___closed__7; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; @@ -10430,19 +10509,47 @@ return x_1; static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__2() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; +x_2 = l_Lean_Elab_Term_expandDbgTrace___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("interpolatedStrKind", 19); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_expandDbgTrace___closed__3; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__5() { +_start: +{ lean_object* x_1; lean_object* x_2; x_1 = l_Lean_Elab_Term_expandDbgTrace___closed__1; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__3() { +static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Elab_Term_expandDbgTrace___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Term_expandDbgTrace___closed__2; +x_3 = l_Lean_Elab_Term_expandDbgTrace___closed__5; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -10450,7 +10557,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__4() { +static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -10460,31 +10567,31 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__5() { +static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_expandDbgTrace___closed__4; +x_2 = l_Lean_Elab_Term_expandDbgTrace___closed__7; 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_Elab_Term_expandDbgTrace___closed__6() { +static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_expandDbgTrace___closed__5; +x_2 = l_Lean_Elab_Term_expandDbgTrace___closed__8; x_3 = lean_alloc_ctor(1, 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_Elab_Term_expandDbgTrace___closed__7() { +static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__10() { _start: { lean_object* x_1; @@ -10492,22 +10599,22 @@ x_1 = lean_mk_string_from_bytes("toString", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__8() { +static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_expandDbgTrace___closed__7; +x_1 = l_Lean_Elab_Term_expandDbgTrace___closed__10; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__9() { +static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Term_expandDbgTrace___closed__7; +x_1 = l_Lean_Elab_Term_expandDbgTrace___closed__10; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Term_expandDbgTrace___closed__8; +x_3 = l_Lean_Elab_Term_expandDbgTrace___closed__11; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -10515,17 +10622,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__10() { +static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_expandDbgTrace___closed__7; +x_2 = l_Lean_Elab_Term_expandDbgTrace___closed__10; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__11() { +static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__14() { _start: { lean_object* x_1; @@ -10533,51 +10640,51 @@ x_1 = lean_mk_string_from_bytes("ToString", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__12() { +static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_expandDbgTrace___closed__11; +x_2 = l_Lean_Elab_Term_expandDbgTrace___closed__14; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__13() { +static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_expandDbgTrace___closed__12; -x_2 = l_Lean_Elab_Term_expandDbgTrace___closed__7; +x_1 = l_Lean_Elab_Term_expandDbgTrace___closed__15; +x_2 = l_Lean_Elab_Term_expandDbgTrace___closed__10; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__14() { +static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_expandDbgTrace___closed__13; +x_2 = l_Lean_Elab_Term_expandDbgTrace___closed__16; 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_Elab_Term_expandDbgTrace___closed__15() { +static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_expandDbgTrace___closed__14; +x_2 = l_Lean_Elab_Term_expandDbgTrace___closed__17; x_3 = lean_alloc_ctor(1, 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_Elab_Term_expandDbgTrace___closed__16() { +static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__19() { _start: { lean_object* x_1; @@ -10585,17 +10692,17 @@ x_1 = lean_mk_string_from_bytes("fun", 3); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__17() { +static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; -x_2 = l_Lean_Elab_Term_expandDbgTrace___closed__16; +x_2 = l_Lean_Elab_Term_expandDbgTrace___closed__19; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__18() { +static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__21() { _start: { lean_object* x_1; @@ -10603,17 +10710,17 @@ x_1 = lean_mk_string_from_bytes("basicFun", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__19() { +static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; -x_2 = l_Lean_Elab_Term_expandDbgTrace___closed__18; +x_2 = l_Lean_Elab_Term_expandDbgTrace___closed__21; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__20() { +static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__23() { _start: { lean_object* x_1; @@ -10621,17 +10728,17 @@ x_1 = lean_mk_string_from_bytes("hole", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__21() { +static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; -x_2 = l_Lean_Elab_Term_expandDbgTrace___closed__20; +x_2 = l_Lean_Elab_Term_expandDbgTrace___closed__23; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__22() { +static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__25() { _start: { lean_object* x_1; @@ -10639,7 +10746,7 @@ x_1 = lean_mk_string_from_bytes("_", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__23() { +static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__26() { _start: { lean_object* x_1; @@ -10647,7 +10754,7 @@ x_1 = lean_mk_string_from_bytes("=>", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__24() { +static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__27() { _start: { lean_object* x_1; @@ -10655,17 +10762,17 @@ x_1 = lean_mk_string_from_bytes("termS!_", 7); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__25() { +static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_expandDbgTrace___closed__24; +x_2 = l_Lean_Elab_Term_expandDbgTrace___closed__27; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__26() { +static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__29() { _start: { lean_object* x_1; @@ -10676,554 +10783,555 @@ return x_1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandDbgTrace(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; -x_4 = lean_unsigned_to_nat(1u); -x_5 = l_Lean_Syntax_getArg(x_1, x_4); -x_6 = lean_unsigned_to_nat(3u); -x_7 = l_Lean_Syntax_getArg(x_1, x_6); -lean_inc(x_5); -x_8 = l_Lean_Syntax_getKind(x_5); -x_9 = l_Lean_interpolatedStrKind; -x_10 = lean_name_eq(x_8, x_9); -lean_dec(x_8); -if (x_10 == 0) +lean_object* x_4; uint8_t x_5; +x_4 = l_Lean_Elab_Term_expandDbgTrace___closed__2; +lean_inc(x_1); +x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); +if (x_5 == 0) { -lean_object* x_11; uint8_t x_12; +lean_object* x_6; lean_object* x_7; +lean_dec(x_2); +lean_dec(x_1); +x_6 = lean_box(1); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_3); +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_unsigned_to_nat(1u); +x_9 = l_Lean_Syntax_getArg(x_1, x_8); +x_10 = l_Lean_Elab_Term_expandDbgTrace___closed__4; +lean_inc(x_9); +x_11 = l_Lean_Syntax_isOfKind(x_9, x_10); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = lean_unsigned_to_nat(3u); +x_13 = l_Lean_Syntax_getArg(x_1, x_12); +lean_dec(x_1); lean_inc(x_2); -x_11 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) +x_14 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) { -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; 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; -x_13 = lean_ctor_get(x_11, 0); -x_14 = lean_ctor_get(x_2, 2); -lean_inc(x_14); -x_15 = lean_ctor_get(x_2, 1); -lean_inc(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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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_16 = lean_ctor_get(x_14, 0); +x_17 = lean_ctor_get(x_2, 2); +lean_inc(x_17); +x_18 = lean_ctor_get(x_2, 1); +lean_inc(x_18); lean_dec(x_2); -x_16 = l_Lean_Elab_Term_expandDbgTrace___closed__4; -lean_inc(x_14); -lean_inc(x_15); -x_17 = l_Lean_addMacroScope(x_15, x_16, x_14); -x_18 = l_Lean_Elab_Term_expandDbgTrace___closed__3; -x_19 = l_Lean_Elab_Term_expandDbgTrace___closed__6; -lean_inc(x_13); -x_20 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_20, 0, x_13); -lean_ctor_set(x_20, 1, x_18); -lean_ctor_set(x_20, 2, x_17); -lean_ctor_set(x_20, 3, x_19); -x_21 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; -lean_inc(x_13); -x_22 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_22, 0, x_13); -lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean_Elab_Term_expandDbgTrace___closed__10; -x_24 = l_Lean_addMacroScope(x_15, x_23, x_14); -x_25 = l_Lean_Elab_Term_expandDbgTrace___closed__9; -x_26 = l_Lean_Elab_Term_expandDbgTrace___closed__15; -lean_inc(x_13); -x_27 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_27, 0, x_13); -lean_ctor_set(x_27, 1, x_25); -lean_ctor_set(x_27, 2, x_24); -lean_ctor_set(x_27, 3, x_26); -x_28 = l_Lean_Elab_Term_expandShow___closed__27; -x_29 = lean_array_push(x_28, x_5); -x_30 = lean_box(2); -x_31 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; -x_32 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -lean_ctor_set(x_32, 2, x_29); -x_33 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; -x_34 = lean_array_push(x_33, x_27); -x_35 = lean_array_push(x_34, x_32); -x_36 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__5; -x_37 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_37, 0, x_30); -lean_ctor_set(x_37, 1, x_36); -lean_ctor_set(x_37, 2, x_35); -x_38 = lean_array_push(x_33, x_37); -x_39 = l_Lean_Elab_Term_expandShow___closed__23; -x_40 = lean_array_push(x_38, x_39); -x_41 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_41, 0, x_30); -lean_ctor_set(x_41, 1, x_31); -lean_ctor_set(x_41, 2, x_40); -x_42 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__23; -lean_inc(x_13); -x_43 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_43, 0, x_13); -lean_ctor_set(x_43, 1, x_42); -x_44 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; -x_45 = lean_array_push(x_44, x_22); -x_46 = lean_array_push(x_45, x_41); -x_47 = lean_array_push(x_46, x_43); -x_48 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__14; -x_49 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_49, 0, x_30); -lean_ctor_set(x_49, 1, x_48); -lean_ctor_set(x_49, 2, x_47); -x_50 = l_Lean_Elab_Term_expandDbgTrace___closed__16; -lean_inc(x_13); -x_51 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_51, 0, x_13); -lean_ctor_set(x_51, 1, x_50); -x_52 = l_Lean_Elab_Term_expandDbgTrace___closed__22; -lean_inc(x_13); -x_53 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_53, 0, x_13); -lean_ctor_set(x_53, 1, x_52); -x_54 = lean_array_push(x_28, x_53); -x_55 = l_Lean_Elab_Term_expandDbgTrace___closed__21; -x_56 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_56, 0, x_30); +x_19 = l_Lean_Elab_Term_expandDbgTrace___closed__7; +lean_inc(x_17); +lean_inc(x_18); +x_20 = l_Lean_addMacroScope(x_18, x_19, x_17); +x_21 = l_Lean_Elab_Term_expandDbgTrace___closed__6; +x_22 = l_Lean_Elab_Term_expandDbgTrace___closed__9; +lean_inc(x_16); +x_23 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_23, 0, x_16); +lean_ctor_set(x_23, 1, x_21); +lean_ctor_set(x_23, 2, x_20); +lean_ctor_set(x_23, 3, x_22); +x_24 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; +lean_inc(x_16); +x_25 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_25, 0, x_16); +lean_ctor_set(x_25, 1, x_24); +x_26 = l_Lean_Elab_Term_expandDbgTrace___closed__13; +x_27 = l_Lean_addMacroScope(x_18, x_26, x_17); +x_28 = l_Lean_Elab_Term_expandDbgTrace___closed__12; +x_29 = l_Lean_Elab_Term_expandDbgTrace___closed__18; +lean_inc(x_16); +x_30 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_30, 0, x_16); +lean_ctor_set(x_30, 1, x_28); +lean_ctor_set(x_30, 2, x_27); +lean_ctor_set(x_30, 3, x_29); +x_31 = l_Lean_Elab_Term_expandShow___closed__27; +x_32 = lean_array_push(x_31, x_9); +x_33 = lean_box(2); +x_34 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; +x_35 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +lean_ctor_set(x_35, 2, x_32); +x_36 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; +x_37 = lean_array_push(x_36, x_30); +x_38 = lean_array_push(x_37, x_35); +x_39 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__5; +x_40 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_40, 0, x_33); +lean_ctor_set(x_40, 1, x_39); +lean_ctor_set(x_40, 2, x_38); +x_41 = lean_array_push(x_36, x_40); +x_42 = l_Lean_Elab_Term_expandShow___closed__23; +x_43 = lean_array_push(x_41, x_42); +x_44 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_44, 0, x_33); +lean_ctor_set(x_44, 1, x_34); +lean_ctor_set(x_44, 2, x_43); +x_45 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__23; +lean_inc(x_16); +x_46 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_46, 0, x_16); +lean_ctor_set(x_46, 1, x_45); +x_47 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; +x_48 = lean_array_push(x_47, x_25); +x_49 = lean_array_push(x_48, x_44); +x_50 = lean_array_push(x_49, x_46); +x_51 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__14; +x_52 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_52, 0, x_33); +lean_ctor_set(x_52, 1, x_51); +lean_ctor_set(x_52, 2, x_50); +x_53 = l_Lean_Elab_Term_expandDbgTrace___closed__19; +lean_inc(x_16); +x_54 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_54, 0, x_16); +lean_ctor_set(x_54, 1, x_53); +x_55 = l_Lean_Elab_Term_expandDbgTrace___closed__25; +lean_inc(x_16); +x_56 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_56, 0, x_16); lean_ctor_set(x_56, 1, x_55); -lean_ctor_set(x_56, 2, x_54); -x_57 = lean_array_push(x_28, x_56); -x_58 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_58, 0, x_30); -lean_ctor_set(x_58, 1, x_31); -lean_ctor_set(x_58, 2, x_57); -x_59 = l_Lean_Elab_Term_expandDbgTrace___closed__23; -x_60 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_60, 0, x_13); -lean_ctor_set(x_60, 1, x_59); -x_61 = lean_array_push(x_44, x_58); -x_62 = lean_array_push(x_61, x_60); -x_63 = lean_array_push(x_62, x_7); -x_64 = l_Lean_Elab_Term_expandDbgTrace___closed__19; -x_65 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_65, 0, x_30); -lean_ctor_set(x_65, 1, x_64); -lean_ctor_set(x_65, 2, x_63); -x_66 = lean_array_push(x_33, x_51); -x_67 = lean_array_push(x_66, x_65); -x_68 = l_Lean_Elab_Term_expandDbgTrace___closed__17; -x_69 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_69, 0, x_30); -lean_ctor_set(x_69, 1, x_68); -lean_ctor_set(x_69, 2, x_67); -x_70 = lean_array_push(x_33, x_49); -x_71 = lean_array_push(x_70, x_69); +x_57 = lean_array_push(x_31, x_56); +x_58 = l_Lean_Elab_Term_expandDbgTrace___closed__24; +x_59 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_59, 0, x_33); +lean_ctor_set(x_59, 1, x_58); +lean_ctor_set(x_59, 2, x_57); +x_60 = lean_array_push(x_31, x_59); +x_61 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_61, 0, x_33); +lean_ctor_set(x_61, 1, x_34); +lean_ctor_set(x_61, 2, x_60); +x_62 = l_Lean_Elab_Term_expandDbgTrace___closed__26; +x_63 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_63, 0, x_16); +lean_ctor_set(x_63, 1, x_62); +x_64 = lean_array_push(x_47, x_61); +x_65 = lean_array_push(x_64, x_63); +x_66 = lean_array_push(x_65, x_13); +x_67 = l_Lean_Elab_Term_expandDbgTrace___closed__22; +x_68 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_68, 0, x_33); +lean_ctor_set(x_68, 1, x_67); +lean_ctor_set(x_68, 2, x_66); +x_69 = lean_array_push(x_36, x_54); +x_70 = lean_array_push(x_69, x_68); +x_71 = l_Lean_Elab_Term_expandDbgTrace___closed__20; x_72 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_72, 0, x_30); -lean_ctor_set(x_72, 1, x_31); -lean_ctor_set(x_72, 2, x_71); -x_73 = lean_array_push(x_33, x_20); +lean_ctor_set(x_72, 0, x_33); +lean_ctor_set(x_72, 1, x_71); +lean_ctor_set(x_72, 2, x_70); +x_73 = lean_array_push(x_36, x_52); x_74 = lean_array_push(x_73, x_72); x_75 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_75, 0, x_30); -lean_ctor_set(x_75, 1, x_36); +lean_ctor_set(x_75, 0, x_33); +lean_ctor_set(x_75, 1, x_34); lean_ctor_set(x_75, 2, x_74); -lean_ctor_set(x_11, 0, x_75); -return x_11; +x_76 = lean_array_push(x_36, x_23); +x_77 = lean_array_push(x_76, x_75); +x_78 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_78, 0, x_33); +lean_ctor_set(x_78, 1, x_39); +lean_ctor_set(x_78, 2, x_77); +lean_ctor_set(x_14, 0, x_78); +return x_14; } else { -lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; -x_76 = lean_ctor_get(x_11, 0); -x_77 = lean_ctor_get(x_11, 1); -lean_inc(x_77); -lean_inc(x_76); -lean_dec(x_11); -x_78 = lean_ctor_get(x_2, 2); -lean_inc(x_78); -x_79 = lean_ctor_get(x_2, 1); +lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; +x_79 = lean_ctor_get(x_14, 0); +x_80 = lean_ctor_get(x_14, 1); +lean_inc(x_80); lean_inc(x_79); +lean_dec(x_14); +x_81 = lean_ctor_get(x_2, 2); +lean_inc(x_81); +x_82 = lean_ctor_get(x_2, 1); +lean_inc(x_82); lean_dec(x_2); -x_80 = l_Lean_Elab_Term_expandDbgTrace___closed__4; -lean_inc(x_78); +x_83 = l_Lean_Elab_Term_expandDbgTrace___closed__7; +lean_inc(x_81); +lean_inc(x_82); +x_84 = l_Lean_addMacroScope(x_82, x_83, x_81); +x_85 = l_Lean_Elab_Term_expandDbgTrace___closed__6; +x_86 = l_Lean_Elab_Term_expandDbgTrace___closed__9; lean_inc(x_79); -x_81 = l_Lean_addMacroScope(x_79, x_80, x_78); -x_82 = l_Lean_Elab_Term_expandDbgTrace___closed__3; -x_83 = l_Lean_Elab_Term_expandDbgTrace___closed__6; -lean_inc(x_76); -x_84 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_84, 0, x_76); -lean_ctor_set(x_84, 1, x_82); -lean_ctor_set(x_84, 2, x_81); -lean_ctor_set(x_84, 3, x_83); -x_85 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; -lean_inc(x_76); -x_86 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_86, 0, x_76); -lean_ctor_set(x_86, 1, x_85); -x_87 = l_Lean_Elab_Term_expandDbgTrace___closed__10; -x_88 = l_Lean_addMacroScope(x_79, x_87, x_78); -x_89 = l_Lean_Elab_Term_expandDbgTrace___closed__9; -x_90 = l_Lean_Elab_Term_expandDbgTrace___closed__15; -lean_inc(x_76); -x_91 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_91, 0, x_76); -lean_ctor_set(x_91, 1, x_89); -lean_ctor_set(x_91, 2, x_88); -lean_ctor_set(x_91, 3, x_90); -x_92 = l_Lean_Elab_Term_expandShow___closed__27; -x_93 = lean_array_push(x_92, x_5); -x_94 = lean_box(2); -x_95 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; -x_96 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_96, 0, x_94); -lean_ctor_set(x_96, 1, x_95); -lean_ctor_set(x_96, 2, x_93); -x_97 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; -x_98 = lean_array_push(x_97, x_91); -x_99 = lean_array_push(x_98, x_96); -x_100 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__5; -x_101 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_101, 0, x_94); -lean_ctor_set(x_101, 1, x_100); -lean_ctor_set(x_101, 2, x_99); -x_102 = lean_array_push(x_97, x_101); -x_103 = l_Lean_Elab_Term_expandShow___closed__23; -x_104 = lean_array_push(x_102, x_103); -x_105 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_105, 0, x_94); -lean_ctor_set(x_105, 1, x_95); -lean_ctor_set(x_105, 2, x_104); -x_106 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__23; -lean_inc(x_76); -x_107 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_107, 0, x_76); -lean_ctor_set(x_107, 1, x_106); -x_108 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; -x_109 = lean_array_push(x_108, x_86); -x_110 = lean_array_push(x_109, x_105); -x_111 = lean_array_push(x_110, x_107); -x_112 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__14; -x_113 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_113, 0, x_94); -lean_ctor_set(x_113, 1, x_112); -lean_ctor_set(x_113, 2, x_111); -x_114 = l_Lean_Elab_Term_expandDbgTrace___closed__16; -lean_inc(x_76); -x_115 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_115, 0, x_76); -lean_ctor_set(x_115, 1, x_114); -x_116 = l_Lean_Elab_Term_expandDbgTrace___closed__22; -lean_inc(x_76); -x_117 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_117, 0, x_76); -lean_ctor_set(x_117, 1, x_116); -x_118 = lean_array_push(x_92, x_117); -x_119 = l_Lean_Elab_Term_expandDbgTrace___closed__21; -x_120 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_120, 0, x_94); +x_87 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_87, 0, x_79); +lean_ctor_set(x_87, 1, x_85); +lean_ctor_set(x_87, 2, x_84); +lean_ctor_set(x_87, 3, x_86); +x_88 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; +lean_inc(x_79); +x_89 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_89, 0, x_79); +lean_ctor_set(x_89, 1, x_88); +x_90 = l_Lean_Elab_Term_expandDbgTrace___closed__13; +x_91 = l_Lean_addMacroScope(x_82, x_90, x_81); +x_92 = l_Lean_Elab_Term_expandDbgTrace___closed__12; +x_93 = l_Lean_Elab_Term_expandDbgTrace___closed__18; +lean_inc(x_79); +x_94 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_94, 0, x_79); +lean_ctor_set(x_94, 1, x_92); +lean_ctor_set(x_94, 2, x_91); +lean_ctor_set(x_94, 3, x_93); +x_95 = l_Lean_Elab_Term_expandShow___closed__27; +x_96 = lean_array_push(x_95, x_9); +x_97 = lean_box(2); +x_98 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; +x_99 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_99, 0, x_97); +lean_ctor_set(x_99, 1, x_98); +lean_ctor_set(x_99, 2, x_96); +x_100 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; +x_101 = lean_array_push(x_100, x_94); +x_102 = lean_array_push(x_101, x_99); +x_103 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__5; +x_104 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_104, 0, x_97); +lean_ctor_set(x_104, 1, x_103); +lean_ctor_set(x_104, 2, x_102); +x_105 = lean_array_push(x_100, x_104); +x_106 = l_Lean_Elab_Term_expandShow___closed__23; +x_107 = lean_array_push(x_105, x_106); +x_108 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_108, 0, x_97); +lean_ctor_set(x_108, 1, x_98); +lean_ctor_set(x_108, 2, x_107); +x_109 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__23; +lean_inc(x_79); +x_110 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_110, 0, x_79); +lean_ctor_set(x_110, 1, x_109); +x_111 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; +x_112 = lean_array_push(x_111, x_89); +x_113 = lean_array_push(x_112, x_108); +x_114 = lean_array_push(x_113, x_110); +x_115 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__14; +x_116 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_116, 0, x_97); +lean_ctor_set(x_116, 1, x_115); +lean_ctor_set(x_116, 2, x_114); +x_117 = l_Lean_Elab_Term_expandDbgTrace___closed__19; +lean_inc(x_79); +x_118 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_118, 0, x_79); +lean_ctor_set(x_118, 1, x_117); +x_119 = l_Lean_Elab_Term_expandDbgTrace___closed__25; +lean_inc(x_79); +x_120 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_120, 0, x_79); lean_ctor_set(x_120, 1, x_119); -lean_ctor_set(x_120, 2, x_118); -x_121 = lean_array_push(x_92, x_120); -x_122 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_122, 0, x_94); -lean_ctor_set(x_122, 1, x_95); -lean_ctor_set(x_122, 2, x_121); -x_123 = l_Lean_Elab_Term_expandDbgTrace___closed__23; -x_124 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_124, 0, x_76); -lean_ctor_set(x_124, 1, x_123); -x_125 = lean_array_push(x_108, x_122); -x_126 = lean_array_push(x_125, x_124); -x_127 = lean_array_push(x_126, x_7); -x_128 = l_Lean_Elab_Term_expandDbgTrace___closed__19; -x_129 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_129, 0, x_94); -lean_ctor_set(x_129, 1, x_128); -lean_ctor_set(x_129, 2, x_127); -x_130 = lean_array_push(x_97, x_115); -x_131 = lean_array_push(x_130, x_129); -x_132 = l_Lean_Elab_Term_expandDbgTrace___closed__17; -x_133 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_133, 0, x_94); -lean_ctor_set(x_133, 1, x_132); -lean_ctor_set(x_133, 2, x_131); -x_134 = lean_array_push(x_97, x_113); -x_135 = lean_array_push(x_134, x_133); +x_121 = lean_array_push(x_95, x_120); +x_122 = l_Lean_Elab_Term_expandDbgTrace___closed__24; +x_123 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_123, 0, x_97); +lean_ctor_set(x_123, 1, x_122); +lean_ctor_set(x_123, 2, x_121); +x_124 = lean_array_push(x_95, x_123); +x_125 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_125, 0, x_97); +lean_ctor_set(x_125, 1, x_98); +lean_ctor_set(x_125, 2, x_124); +x_126 = l_Lean_Elab_Term_expandDbgTrace___closed__26; +x_127 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_127, 0, x_79); +lean_ctor_set(x_127, 1, x_126); +x_128 = lean_array_push(x_111, x_125); +x_129 = lean_array_push(x_128, x_127); +x_130 = lean_array_push(x_129, x_13); +x_131 = l_Lean_Elab_Term_expandDbgTrace___closed__22; +x_132 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_132, 0, x_97); +lean_ctor_set(x_132, 1, x_131); +lean_ctor_set(x_132, 2, x_130); +x_133 = lean_array_push(x_100, x_118); +x_134 = lean_array_push(x_133, x_132); +x_135 = l_Lean_Elab_Term_expandDbgTrace___closed__20; x_136 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_136, 0, x_94); -lean_ctor_set(x_136, 1, x_95); -lean_ctor_set(x_136, 2, x_135); -x_137 = lean_array_push(x_97, x_84); +lean_ctor_set(x_136, 0, x_97); +lean_ctor_set(x_136, 1, x_135); +lean_ctor_set(x_136, 2, x_134); +x_137 = lean_array_push(x_100, x_116); x_138 = lean_array_push(x_137, x_136); x_139 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_139, 0, x_94); -lean_ctor_set(x_139, 1, x_100); +lean_ctor_set(x_139, 0, x_97); +lean_ctor_set(x_139, 1, x_98); lean_ctor_set(x_139, 2, x_138); -x_140 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_140, 0, x_139); -lean_ctor_set(x_140, 1, x_77); -return x_140; +x_140 = lean_array_push(x_100, x_87); +x_141 = lean_array_push(x_140, x_139); +x_142 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_142, 0, x_97); +lean_ctor_set(x_142, 1, x_103); +lean_ctor_set(x_142, 2, x_141); +x_143 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_143, 0, x_142); +lean_ctor_set(x_143, 1, x_80); +return x_143; } } else { -lean_object* x_141; uint8_t x_142; +lean_object* x_144; lean_object* x_145; lean_object* x_146; uint8_t x_147; +x_144 = lean_unsigned_to_nat(3u); +x_145 = l_Lean_Syntax_getArg(x_1, x_144); +lean_dec(x_1); lean_inc(x_2); -x_141 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_142 = !lean_is_exclusive(x_141); -if (x_142 == 0) +x_146 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_147 = !lean_is_exclusive(x_146); +if (x_147 == 0) { -lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; -x_143 = lean_ctor_get(x_141, 0); -x_144 = lean_ctor_get(x_2, 2); -lean_inc(x_144); -x_145 = lean_ctor_get(x_2, 1); -lean_inc(x_145); +lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; +x_148 = lean_ctor_get(x_146, 0); +x_149 = lean_ctor_get(x_2, 2); +lean_inc(x_149); +x_150 = lean_ctor_get(x_2, 1); +lean_inc(x_150); lean_dec(x_2); -x_146 = l_Lean_Elab_Term_expandDbgTrace___closed__4; -x_147 = l_Lean_addMacroScope(x_145, x_146, x_144); -x_148 = l_Lean_Elab_Term_expandDbgTrace___closed__3; -x_149 = l_Lean_Elab_Term_expandDbgTrace___closed__6; -lean_inc(x_143); -x_150 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_150, 0, x_143); -lean_ctor_set(x_150, 1, x_148); -lean_ctor_set(x_150, 2, x_147); -lean_ctor_set(x_150, 3, x_149); -x_151 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; -lean_inc(x_143); -x_152 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_152, 0, x_143); -lean_ctor_set(x_152, 1, x_151); -x_153 = l_Lean_Elab_Term_expandDbgTrace___closed__26; -lean_inc(x_143); -x_154 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_154, 0, x_143); -lean_ctor_set(x_154, 1, x_153); -x_155 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; -x_156 = lean_array_push(x_155, x_154); -x_157 = lean_array_push(x_156, x_5); -x_158 = lean_box(2); -x_159 = l_Lean_Elab_Term_expandDbgTrace___closed__25; -x_160 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_160, 0, x_158); -lean_ctor_set(x_160, 1, x_159); -lean_ctor_set(x_160, 2, x_157); -x_161 = lean_array_push(x_155, x_160); -x_162 = l_Lean_Elab_Term_expandShow___closed__23; -x_163 = lean_array_push(x_161, x_162); -x_164 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; +x_151 = l_Lean_Elab_Term_expandDbgTrace___closed__7; +x_152 = l_Lean_addMacroScope(x_150, x_151, x_149); +x_153 = l_Lean_Elab_Term_expandDbgTrace___closed__6; +x_154 = l_Lean_Elab_Term_expandDbgTrace___closed__9; +lean_inc(x_148); +x_155 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_155, 0, x_148); +lean_ctor_set(x_155, 1, x_153); +lean_ctor_set(x_155, 2, x_152); +lean_ctor_set(x_155, 3, x_154); +x_156 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; +lean_inc(x_148); +x_157 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_157, 0, x_148); +lean_ctor_set(x_157, 1, x_156); +x_158 = l_Lean_Elab_Term_expandDbgTrace___closed__29; +lean_inc(x_148); +x_159 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_159, 0, x_148); +lean_ctor_set(x_159, 1, x_158); +x_160 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; +x_161 = lean_array_push(x_160, x_159); +x_162 = lean_array_push(x_161, x_9); +x_163 = lean_box(2); +x_164 = l_Lean_Elab_Term_expandDbgTrace___closed__28; x_165 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_165, 0, x_158); +lean_ctor_set(x_165, 0, x_163); lean_ctor_set(x_165, 1, x_164); -lean_ctor_set(x_165, 2, x_163); -x_166 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__23; -lean_inc(x_143); -x_167 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_167, 0, x_143); -lean_ctor_set(x_167, 1, x_166); -x_168 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; -x_169 = lean_array_push(x_168, x_152); -x_170 = lean_array_push(x_169, x_165); -x_171 = lean_array_push(x_170, x_167); -x_172 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__14; -x_173 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_173, 0, x_158); -lean_ctor_set(x_173, 1, x_172); -lean_ctor_set(x_173, 2, x_171); -x_174 = l_Lean_Elab_Term_expandDbgTrace___closed__16; -lean_inc(x_143); -x_175 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_175, 0, x_143); -lean_ctor_set(x_175, 1, x_174); -x_176 = l_Lean_Elab_Term_expandDbgTrace___closed__22; -lean_inc(x_143); -x_177 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_177, 0, x_143); -lean_ctor_set(x_177, 1, x_176); -x_178 = l_Lean_Elab_Term_expandShow___closed__27; -x_179 = lean_array_push(x_178, x_177); -x_180 = l_Lean_Elab_Term_expandDbgTrace___closed__21; -x_181 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_181, 0, x_158); -lean_ctor_set(x_181, 1, x_180); -lean_ctor_set(x_181, 2, x_179); -x_182 = lean_array_push(x_178, x_181); -x_183 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_183, 0, x_158); -lean_ctor_set(x_183, 1, x_164); -lean_ctor_set(x_183, 2, x_182); -x_184 = l_Lean_Elab_Term_expandDbgTrace___closed__23; -x_185 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_185, 0, x_143); -lean_ctor_set(x_185, 1, x_184); -x_186 = lean_array_push(x_168, x_183); -x_187 = lean_array_push(x_186, x_185); -x_188 = lean_array_push(x_187, x_7); -x_189 = l_Lean_Elab_Term_expandDbgTrace___closed__19; -x_190 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_190, 0, x_158); +lean_ctor_set(x_165, 2, x_162); +x_166 = lean_array_push(x_160, x_165); +x_167 = l_Lean_Elab_Term_expandShow___closed__23; +x_168 = lean_array_push(x_166, x_167); +x_169 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; +x_170 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_170, 0, x_163); +lean_ctor_set(x_170, 1, x_169); +lean_ctor_set(x_170, 2, x_168); +x_171 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__23; +lean_inc(x_148); +x_172 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_172, 0, x_148); +lean_ctor_set(x_172, 1, x_171); +x_173 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; +x_174 = lean_array_push(x_173, x_157); +x_175 = lean_array_push(x_174, x_170); +x_176 = lean_array_push(x_175, x_172); +x_177 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__14; +x_178 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_178, 0, x_163); +lean_ctor_set(x_178, 1, x_177); +lean_ctor_set(x_178, 2, x_176); +x_179 = l_Lean_Elab_Term_expandDbgTrace___closed__19; +lean_inc(x_148); +x_180 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_180, 0, x_148); +lean_ctor_set(x_180, 1, x_179); +x_181 = l_Lean_Elab_Term_expandDbgTrace___closed__25; +lean_inc(x_148); +x_182 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_182, 0, x_148); +lean_ctor_set(x_182, 1, x_181); +x_183 = l_Lean_Elab_Term_expandShow___closed__27; +x_184 = lean_array_push(x_183, x_182); +x_185 = l_Lean_Elab_Term_expandDbgTrace___closed__24; +x_186 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_186, 0, x_163); +lean_ctor_set(x_186, 1, x_185); +lean_ctor_set(x_186, 2, x_184); +x_187 = lean_array_push(x_183, x_186); +x_188 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_188, 0, x_163); +lean_ctor_set(x_188, 1, x_169); +lean_ctor_set(x_188, 2, x_187); +x_189 = l_Lean_Elab_Term_expandDbgTrace___closed__26; +x_190 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_190, 0, x_148); lean_ctor_set(x_190, 1, x_189); -lean_ctor_set(x_190, 2, x_188); -x_191 = lean_array_push(x_155, x_175); +x_191 = lean_array_push(x_173, x_188); x_192 = lean_array_push(x_191, x_190); -x_193 = l_Lean_Elab_Term_expandDbgTrace___closed__17; -x_194 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_194, 0, x_158); -lean_ctor_set(x_194, 1, x_193); -lean_ctor_set(x_194, 2, x_192); -x_195 = lean_array_push(x_155, x_173); -x_196 = lean_array_push(x_195, x_194); -x_197 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_197, 0, x_158); -lean_ctor_set(x_197, 1, x_164); -lean_ctor_set(x_197, 2, x_196); -x_198 = lean_array_push(x_155, x_150); -x_199 = lean_array_push(x_198, x_197); -x_200 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__5; -x_201 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_201, 0, x_158); -lean_ctor_set(x_201, 1, x_200); -lean_ctor_set(x_201, 2, x_199); -lean_ctor_set(x_141, 0, x_201); -return x_141; +x_193 = lean_array_push(x_192, x_145); +x_194 = l_Lean_Elab_Term_expandDbgTrace___closed__22; +x_195 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_195, 0, x_163); +lean_ctor_set(x_195, 1, x_194); +lean_ctor_set(x_195, 2, x_193); +x_196 = lean_array_push(x_160, x_180); +x_197 = lean_array_push(x_196, x_195); +x_198 = l_Lean_Elab_Term_expandDbgTrace___closed__20; +x_199 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_199, 0, x_163); +lean_ctor_set(x_199, 1, x_198); +lean_ctor_set(x_199, 2, x_197); +x_200 = lean_array_push(x_160, x_178); +x_201 = lean_array_push(x_200, x_199); +x_202 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_202, 0, x_163); +lean_ctor_set(x_202, 1, x_169); +lean_ctor_set(x_202, 2, x_201); +x_203 = lean_array_push(x_160, x_155); +x_204 = lean_array_push(x_203, x_202); +x_205 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__5; +x_206 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_206, 0, x_163); +lean_ctor_set(x_206, 1, x_205); +lean_ctor_set(x_206, 2, x_204); +lean_ctor_set(x_146, 0, x_206); +return x_146; } else { -lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; -x_202 = lean_ctor_get(x_141, 0); -x_203 = lean_ctor_get(x_141, 1); -lean_inc(x_203); -lean_inc(x_202); -lean_dec(x_141); -x_204 = lean_ctor_get(x_2, 2); -lean_inc(x_204); -x_205 = lean_ctor_get(x_2, 1); -lean_inc(x_205); +lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; +x_207 = lean_ctor_get(x_146, 0); +x_208 = lean_ctor_get(x_146, 1); +lean_inc(x_208); +lean_inc(x_207); +lean_dec(x_146); +x_209 = lean_ctor_get(x_2, 2); +lean_inc(x_209); +x_210 = lean_ctor_get(x_2, 1); +lean_inc(x_210); lean_dec(x_2); -x_206 = l_Lean_Elab_Term_expandDbgTrace___closed__4; -x_207 = l_Lean_addMacroScope(x_205, x_206, x_204); -x_208 = l_Lean_Elab_Term_expandDbgTrace___closed__3; -x_209 = l_Lean_Elab_Term_expandDbgTrace___closed__6; -lean_inc(x_202); -x_210 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_210, 0, x_202); -lean_ctor_set(x_210, 1, x_208); -lean_ctor_set(x_210, 2, x_207); -lean_ctor_set(x_210, 3, x_209); -x_211 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; -lean_inc(x_202); -x_212 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_212, 0, x_202); -lean_ctor_set(x_212, 1, x_211); -x_213 = l_Lean_Elab_Term_expandDbgTrace___closed__26; -lean_inc(x_202); -x_214 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_214, 0, x_202); -lean_ctor_set(x_214, 1, x_213); -x_215 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; -x_216 = lean_array_push(x_215, x_214); -x_217 = lean_array_push(x_216, x_5); -x_218 = lean_box(2); -x_219 = l_Lean_Elab_Term_expandDbgTrace___closed__25; -x_220 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_220, 0, x_218); -lean_ctor_set(x_220, 1, x_219); -lean_ctor_set(x_220, 2, x_217); -x_221 = lean_array_push(x_215, x_220); -x_222 = l_Lean_Elab_Term_expandShow___closed__23; -x_223 = lean_array_push(x_221, x_222); -x_224 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; +x_211 = l_Lean_Elab_Term_expandDbgTrace___closed__7; +x_212 = l_Lean_addMacroScope(x_210, x_211, x_209); +x_213 = l_Lean_Elab_Term_expandDbgTrace___closed__6; +x_214 = l_Lean_Elab_Term_expandDbgTrace___closed__9; +lean_inc(x_207); +x_215 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_215, 0, x_207); +lean_ctor_set(x_215, 1, x_213); +lean_ctor_set(x_215, 2, x_212); +lean_ctor_set(x_215, 3, x_214); +x_216 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; +lean_inc(x_207); +x_217 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_217, 0, x_207); +lean_ctor_set(x_217, 1, x_216); +x_218 = l_Lean_Elab_Term_expandDbgTrace___closed__29; +lean_inc(x_207); +x_219 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_219, 0, x_207); +lean_ctor_set(x_219, 1, x_218); +x_220 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; +x_221 = lean_array_push(x_220, x_219); +x_222 = lean_array_push(x_221, x_9); +x_223 = lean_box(2); +x_224 = l_Lean_Elab_Term_expandDbgTrace___closed__28; x_225 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_225, 0, x_218); +lean_ctor_set(x_225, 0, x_223); lean_ctor_set(x_225, 1, x_224); -lean_ctor_set(x_225, 2, x_223); -x_226 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__23; -lean_inc(x_202); -x_227 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_227, 0, x_202); -lean_ctor_set(x_227, 1, x_226); -x_228 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; -x_229 = lean_array_push(x_228, x_212); -x_230 = lean_array_push(x_229, x_225); -x_231 = lean_array_push(x_230, x_227); -x_232 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__14; -x_233 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_233, 0, x_218); -lean_ctor_set(x_233, 1, x_232); -lean_ctor_set(x_233, 2, x_231); -x_234 = l_Lean_Elab_Term_expandDbgTrace___closed__16; -lean_inc(x_202); -x_235 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_235, 0, x_202); -lean_ctor_set(x_235, 1, x_234); -x_236 = l_Lean_Elab_Term_expandDbgTrace___closed__22; -lean_inc(x_202); -x_237 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_237, 0, x_202); -lean_ctor_set(x_237, 1, x_236); -x_238 = l_Lean_Elab_Term_expandShow___closed__27; -x_239 = lean_array_push(x_238, x_237); -x_240 = l_Lean_Elab_Term_expandDbgTrace___closed__21; -x_241 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_241, 0, x_218); -lean_ctor_set(x_241, 1, x_240); -lean_ctor_set(x_241, 2, x_239); -x_242 = lean_array_push(x_238, x_241); -x_243 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_243, 0, x_218); -lean_ctor_set(x_243, 1, x_224); -lean_ctor_set(x_243, 2, x_242); -x_244 = l_Lean_Elab_Term_expandDbgTrace___closed__23; -x_245 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_245, 0, x_202); -lean_ctor_set(x_245, 1, x_244); -x_246 = lean_array_push(x_228, x_243); -x_247 = lean_array_push(x_246, x_245); -x_248 = lean_array_push(x_247, x_7); -x_249 = l_Lean_Elab_Term_expandDbgTrace___closed__19; -x_250 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_250, 0, x_218); +lean_ctor_set(x_225, 2, x_222); +x_226 = lean_array_push(x_220, x_225); +x_227 = l_Lean_Elab_Term_expandShow___closed__23; +x_228 = lean_array_push(x_226, x_227); +x_229 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; +x_230 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_230, 0, x_223); +lean_ctor_set(x_230, 1, x_229); +lean_ctor_set(x_230, 2, x_228); +x_231 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__23; +lean_inc(x_207); +x_232 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_232, 0, x_207); +lean_ctor_set(x_232, 1, x_231); +x_233 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; +x_234 = lean_array_push(x_233, x_217); +x_235 = lean_array_push(x_234, x_230); +x_236 = lean_array_push(x_235, x_232); +x_237 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__14; +x_238 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_238, 0, x_223); +lean_ctor_set(x_238, 1, x_237); +lean_ctor_set(x_238, 2, x_236); +x_239 = l_Lean_Elab_Term_expandDbgTrace___closed__19; +lean_inc(x_207); +x_240 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_240, 0, x_207); +lean_ctor_set(x_240, 1, x_239); +x_241 = l_Lean_Elab_Term_expandDbgTrace___closed__25; +lean_inc(x_207); +x_242 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_242, 0, x_207); +lean_ctor_set(x_242, 1, x_241); +x_243 = l_Lean_Elab_Term_expandShow___closed__27; +x_244 = lean_array_push(x_243, x_242); +x_245 = l_Lean_Elab_Term_expandDbgTrace___closed__24; +x_246 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_246, 0, x_223); +lean_ctor_set(x_246, 1, x_245); +lean_ctor_set(x_246, 2, x_244); +x_247 = lean_array_push(x_243, x_246); +x_248 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_248, 0, x_223); +lean_ctor_set(x_248, 1, x_229); +lean_ctor_set(x_248, 2, x_247); +x_249 = l_Lean_Elab_Term_expandDbgTrace___closed__26; +x_250 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_250, 0, x_207); lean_ctor_set(x_250, 1, x_249); -lean_ctor_set(x_250, 2, x_248); -x_251 = lean_array_push(x_215, x_235); +x_251 = lean_array_push(x_233, x_248); x_252 = lean_array_push(x_251, x_250); -x_253 = l_Lean_Elab_Term_expandDbgTrace___closed__17; -x_254 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_254, 0, x_218); -lean_ctor_set(x_254, 1, x_253); -lean_ctor_set(x_254, 2, x_252); -x_255 = lean_array_push(x_215, x_233); -x_256 = lean_array_push(x_255, x_254); -x_257 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_257, 0, x_218); -lean_ctor_set(x_257, 1, x_224); -lean_ctor_set(x_257, 2, x_256); -x_258 = lean_array_push(x_215, x_210); -x_259 = lean_array_push(x_258, x_257); -x_260 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__5; -x_261 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_261, 0, x_218); -lean_ctor_set(x_261, 1, x_260); -lean_ctor_set(x_261, 2, x_259); -x_262 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_262, 0, x_261); -lean_ctor_set(x_262, 1, x_203); -return x_262; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandDbgTrace___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = l_Lean_Elab_Term_expandDbgTrace(x_1, x_2, x_3); -lean_dec(x_1); -return x_4; +x_253 = lean_array_push(x_252, x_145); +x_254 = l_Lean_Elab_Term_expandDbgTrace___closed__22; +x_255 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_255, 0, x_223); +lean_ctor_set(x_255, 1, x_254); +lean_ctor_set(x_255, 2, x_253); +x_256 = lean_array_push(x_220, x_240); +x_257 = lean_array_push(x_256, x_255); +x_258 = l_Lean_Elab_Term_expandDbgTrace___closed__20; +x_259 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_259, 0, x_223); +lean_ctor_set(x_259, 1, x_258); +lean_ctor_set(x_259, 2, x_257); +x_260 = lean_array_push(x_220, x_238); +x_261 = lean_array_push(x_260, x_259); +x_262 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_262, 0, x_223); +lean_ctor_set(x_262, 1, x_229); +lean_ctor_set(x_262, 2, x_261); +x_263 = lean_array_push(x_220, x_215); +x_264 = lean_array_push(x_263, x_262); +x_265 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__5; +x_266 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_266, 0, x_223); +lean_ctor_set(x_266, 1, x_265); +lean_ctor_set(x_266, 2, x_264); +x_267 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_267, 0, x_266); +lean_ctor_set(x_267, 1, x_208); +return x_267; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandDbgTrace___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; -x_2 = l_Lean_Elab_Term_expandDbgTrace___closed__1; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandDbgTrace___closed__2() { +} +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandDbgTrace___closed__1() { _start: { lean_object* x_1; @@ -11231,21 +11339,21 @@ x_1 = lean_mk_string_from_bytes("expandDbgTrace", 14); return x_1; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandDbgTrace___closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandDbgTrace___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Elab_Term_elabCoe___closed__8; -x_2 = l___regBuiltin_Lean_Elab_Term_expandDbgTrace___closed__2; +x_2 = l___regBuiltin_Lean_Elab_Term_expandDbgTrace___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandDbgTrace___closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandDbgTrace___closed__3() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_expandDbgTrace___boxed), 3, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_expandDbgTrace), 3, 0); return x_1; } } @@ -11254,9 +11362,9 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_expandDbgTrace(lean_objec { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_2 = l___regBuiltin_Lean_Elab_Term_expandShow___closed__3; -x_3 = l___regBuiltin_Lean_Elab_Term_expandDbgTrace___closed__1; -x_4 = l___regBuiltin_Lean_Elab_Term_expandDbgTrace___closed__3; -x_5 = l___regBuiltin_Lean_Elab_Term_expandDbgTrace___closed__4; +x_3 = l_Lean_Elab_Term_expandDbgTrace___closed__2; +x_4 = l___regBuiltin_Lean_Elab_Term_expandDbgTrace___closed__2; +x_5 = l___regBuiltin_Lean_Elab_Term_expandDbgTrace___closed__3; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); return x_6; } @@ -11273,7 +11381,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_expandDbgTrace_docString( _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_expandDbgTrace___closed__3; +x_2 = l___regBuiltin_Lean_Elab_Term_expandDbgTrace___closed__2; x_3 = l___regBuiltin_Lean_Elab_Term_expandDbgTrace_docString___closed__1; x_4 = l_Lean_addBuiltinDocString(x_2, x_3, x_1); return x_4; @@ -11283,7 +11391,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandDbgTrace_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(158u); +x_1 = lean_unsigned_to_nat(160u); x_2 = lean_unsigned_to_nat(43u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11295,8 +11403,8 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandDbgTrace_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(164u); -x_2 = lean_unsigned_to_nat(46u); +x_1 = lean_unsigned_to_nat(163u); +x_2 = lean_unsigned_to_nat(70u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -11310,7 +11418,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Elab_Term_expandDbgTrace_declRange___closed__1; x_2 = lean_unsigned_to_nat(43u); x_3 = l___regBuiltin_Lean_Elab_Term_expandDbgTrace_declRange___closed__2; -x_4 = lean_unsigned_to_nat(46u); +x_4 = lean_unsigned_to_nat(70u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -11323,7 +11431,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandDbgTrace_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(158u); +x_1 = lean_unsigned_to_nat(160u); x_2 = lean_unsigned_to_nat(47u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11335,7 +11443,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandDbgTrace_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(158u); +x_1 = lean_unsigned_to_nat(160u); x_2 = lean_unsigned_to_nat(61u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11375,7 +11483,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_expandDbgTrace_declRange( _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_expandDbgTrace___closed__3; +x_2 = l___regBuiltin_Lean_Elab_Term_expandDbgTrace___closed__2; x_3 = l___regBuiltin_Lean_Elab_Term_expandDbgTrace_declRange___closed__7; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; @@ -11538,7 +11646,7 @@ lean_ctor_set(x_23, 0, x_11); lean_ctor_set(x_23, 1, x_21); lean_ctor_set(x_23, 2, x_20); lean_ctor_set(x_23, 3, x_22); -x_24 = l_Lean_Elab_Term_expandDbgTrace___closed__22; +x_24 = l_Lean_Elab_Term_expandDbgTrace___closed__25; lean_inc(x_11); x_25 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_25, 0, x_11); @@ -11546,7 +11654,7 @@ lean_ctor_set(x_25, 1, x_24); x_26 = l_Lean_Elab_Term_expandShow___closed__27; x_27 = lean_array_push(x_26, x_25); x_28 = lean_box(2); -x_29 = l_Lean_Elab_Term_expandDbgTrace___closed__21; +x_29 = l_Lean_Elab_Term_expandDbgTrace___closed__24; x_30 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_30, 0, x_28); lean_ctor_set(x_30, 1, x_29); @@ -11666,7 +11774,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabSorry_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(167u); +x_1 = lean_unsigned_to_nat(166u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11678,7 +11786,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabSorry_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(169u); +x_1 = lean_unsigned_to_nat(168u); x_2 = lean_unsigned_to_nat(64u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11706,7 +11814,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabSorry_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(167u); +x_1 = lean_unsigned_to_nat(166u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11718,7 +11826,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabSorry_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(167u); +x_1 = lean_unsigned_to_nat(166u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11877,7 +11985,7 @@ lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_ x_9 = lean_unsigned_to_nat(1u); x_10 = lean_nat_sub(x_2, x_9); lean_dec(x_2); -x_11 = l_Lean_instInhabitedSyntax; +x_11 = lean_box(0); x_12 = lean_array_get(x_11, x_1, x_10); lean_inc(x_4); x_13 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_4, x_5); @@ -11939,7 +12047,7 @@ x_4 = lean_array_get_size(x_1); x_5 = lean_unsigned_to_nat(1u); x_6 = lean_nat_sub(x_4, x_5); lean_dec(x_4); -x_7 = l_Lean_instInhabitedSyntax; +x_7 = lean_box(0); x_8 = l_Array_back___rarg(x_7, x_1); x_9 = l_Lean_Elab_Term_mkPairs_loop(x_1, x_6, x_8, x_2, x_3); return x_9; @@ -12218,325 +12326,386 @@ return x_3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandCDot_x3f_go(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___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__14; +lean_inc(x_1); +x_6 = l_Lean_Syntax_isOfKind(x_1, x_5); +if (x_6 == 0) +{ +lean_object* x_7; uint8_t x_8; +x_7 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_hasCDot___closed__2; +lean_inc(x_1); +x_8 = l_Lean_Syntax_isOfKind(x_1, x_7); +if (x_8 == 0) +{ if (lean_obj_tag(x_1) == 1) { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_5 = lean_ctor_get(x_1, 0); -lean_inc(x_5); -x_6 = lean_ctor_get(x_1, 1); -lean_inc(x_6); -x_7 = lean_ctor_get(x_1, 2); -lean_inc(x_7); -x_8 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__14; -x_9 = lean_name_eq(x_6, x_8); +uint8_t x_9; +x_9 = !lean_is_exclusive(x_1); if (x_9 == 0) { -uint8_t x_10; -x_10 = !lean_is_exclusive(x_1); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_11 = lean_ctor_get(x_1, 2); +lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; lean_object* x_14; uint8_t x_15; +x_10 = lean_ctor_get(x_1, 2); +x_11 = lean_array_get_size(x_10); +x_12 = lean_usize_of_nat(x_11); lean_dec(x_11); -x_12 = lean_ctor_get(x_1, 1); -lean_dec(x_12); -x_13 = lean_ctor_get(x_1, 0); -lean_dec(x_13); -x_14 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_hasCDot___closed__2; -x_15 = lean_name_eq(x_6, x_14); +x_13 = 0; +x_14 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandCDot_x3f_go___spec__1(x_12, x_13, x_10, x_2, x_3, x_4); +x_15 = !lean_is_exclusive(x_14); if (x_15 == 0) { -lean_object* x_16; size_t x_17; size_t x_18; lean_object* x_19; uint8_t x_20; -x_16 = lean_array_get_size(x_7); -x_17 = lean_usize_of_nat(x_16); -lean_dec(x_16); -x_18 = 0; -x_19 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandCDot_x3f_go___spec__1(x_17, x_18, x_7, x_2, x_3, x_4); -x_20 = !lean_is_exclusive(x_19); -if (x_20 == 0) +lean_object* x_16; uint8_t x_17; +x_16 = lean_ctor_get(x_14, 0); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) { -lean_object* x_21; uint8_t x_22; -x_21 = lean_ctor_get(x_19, 0); -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) +lean_object* x_18; +x_18 = lean_ctor_get(x_16, 0); +lean_ctor_set(x_1, 2, x_18); +lean_ctor_set(x_16, 0, x_1); +return x_14; +} +else { -lean_object* x_23; -x_23 = lean_ctor_get(x_21, 0); -lean_ctor_set(x_1, 2, x_23); +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_16, 0); +x_20 = lean_ctor_get(x_16, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_16); +lean_ctor_set(x_1, 2, x_19); +x_21 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_21, 0, x_1); -return x_19; +lean_ctor_set(x_21, 1, x_20); +lean_ctor_set(x_14, 0, x_21); +return x_14; +} } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_21, 0); -x_25 = lean_ctor_get(x_21, 1); -lean_inc(x_25); +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; +x_22 = lean_ctor_get(x_14, 0); +x_23 = lean_ctor_get(x_14, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_14); +x_24 = lean_ctor_get(x_22, 0); lean_inc(x_24); -lean_dec(x_21); +x_25 = lean_ctor_get(x_22, 1); +lean_inc(x_25); +if (lean_is_exclusive(x_22)) { + lean_ctor_release(x_22, 0); + lean_ctor_release(x_22, 1); + x_26 = x_22; +} else { + lean_dec_ref(x_22); + x_26 = lean_box(0); +} lean_ctor_set(x_1, 2, x_24); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_1); -lean_ctor_set(x_26, 1, x_25); -lean_ctor_set(x_19, 0, x_26); -return x_19; +if (lean_is_scalar(x_26)) { + x_27 = lean_alloc_ctor(0, 2, 0); +} else { + x_27 = x_26; +} +lean_ctor_set(x_27, 0, x_1); +lean_ctor_set(x_27, 1, x_25); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_23); +return x_28; } } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_27 = lean_ctor_get(x_19, 0); -x_28 = lean_ctor_get(x_19, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_19); -x_29 = lean_ctor_get(x_27, 0); -lean_inc(x_29); -x_30 = lean_ctor_get(x_27, 1); +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; size_t x_33; size_t 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; lean_object* x_44; +x_29 = lean_ctor_get(x_1, 0); +x_30 = lean_ctor_get(x_1, 1); +x_31 = lean_ctor_get(x_1, 2); +lean_inc(x_31); lean_inc(x_30); -if (lean_is_exclusive(x_27)) { - lean_ctor_release(x_27, 0); - lean_ctor_release(x_27, 1); - x_31 = x_27; +lean_inc(x_29); +lean_dec(x_1); +x_32 = lean_array_get_size(x_31); +x_33 = lean_usize_of_nat(x_32); +lean_dec(x_32); +x_34 = 0; +x_35 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandCDot_x3f_go___spec__1(x_33, x_34, x_31, x_2, x_3, x_4); +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +if (lean_is_exclusive(x_35)) { + lean_ctor_release(x_35, 0); + lean_ctor_release(x_35, 1); + x_38 = x_35; } else { - lean_dec_ref(x_27); - x_31 = lean_box(0); + lean_dec_ref(x_35); + x_38 = lean_box(0); } -lean_ctor_set(x_1, 2, x_29); -if (lean_is_scalar(x_31)) { - x_32 = lean_alloc_ctor(0, 2, 0); +x_39 = lean_ctor_get(x_36, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_36, 1); +lean_inc(x_40); +if (lean_is_exclusive(x_36)) { + lean_ctor_release(x_36, 0); + lean_ctor_release(x_36, 1); + x_41 = x_36; } else { - x_32 = x_31; + lean_dec_ref(x_36); + x_41 = lean_box(0); } -lean_ctor_set(x_32, 0, x_1); -lean_ctor_set(x_32, 1, x_30); -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_28); -return x_33; +x_42 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_42, 0, x_29); +lean_ctor_set(x_42, 1, x_30); +lean_ctor_set(x_42, 2, x_39); +if (lean_is_scalar(x_41)) { + x_43 = lean_alloc_ctor(0, 2, 0); +} else { + x_43 = x_41; +} +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_40); +if (lean_is_scalar(x_38)) { + x_44 = lean_alloc_ctor(0, 2, 0); +} else { + x_44 = x_38; +} +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_37); +return x_44; } } else { -uint8_t x_34; -lean_free_object(x_1); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -x_34 = !lean_is_exclusive(x_4); -if (x_34 == 0) +lean_object* x_45; lean_object* x_46; +lean_dec(x_3); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_1); +lean_ctor_set(x_45, 1, x_2); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_4); +return x_46; +} +} +else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; -x_35 = lean_ctor_get(x_4, 0); -x_36 = lean_unsigned_to_nat(1u); -x_37 = lean_nat_add(x_35, x_36); -lean_ctor_set(x_4, 0, x_37); -x_38 = !lean_is_exclusive(x_3); -if (x_38 == 0) +uint8_t x_47; +lean_dec(x_1); +x_47 = !lean_is_exclusive(x_4); +if (x_47 == 0) { -lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; -x_39 = lean_ctor_get(x_3, 1); -x_40 = lean_ctor_get(x_3, 2); -lean_dec(x_40); -lean_inc(x_35); -lean_inc(x_39); -lean_ctor_set(x_3, 2, x_35); -x_41 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_expandCDot_x3f_go___spec__2(x_2, x_3, x_4); -x_42 = !lean_is_exclusive(x_41); -if (x_42 == 0) -{ -lean_object* x_43; uint8_t x_44; -x_43 = lean_ctor_get(x_41, 0); -x_44 = !lean_is_exclusive(x_43); -if (x_44 == 0) +lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; +x_48 = lean_ctor_get(x_4, 0); +x_49 = lean_unsigned_to_nat(1u); +x_50 = lean_nat_add(x_48, x_49); +lean_ctor_set(x_4, 0, x_50); +x_51 = !lean_is_exclusive(x_3); +if (x_51 == 0) { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_45 = lean_ctor_get(x_43, 0); -x_46 = lean_ctor_get(x_43, 1); -x_47 = l_Lean_Elab_Term_expandCDot_x3f_go___closed__4; -x_48 = l_Lean_addMacroScope(x_39, x_47, x_35); -x_49 = lean_box(0); -x_50 = l_Lean_Elab_Term_expandCDot_x3f_go___closed__3; -x_51 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_51, 0, x_45); -lean_ctor_set(x_51, 1, x_50); -lean_ctor_set(x_51, 2, x_48); -lean_ctor_set(x_51, 3, x_49); -lean_inc(x_51); -x_52 = lean_array_push(x_46, x_51); -lean_ctor_set(x_43, 1, x_52); -lean_ctor_set(x_43, 0, x_51); -return x_41; +lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_52 = lean_ctor_get(x_3, 1); +x_53 = lean_ctor_get(x_3, 2); +lean_dec(x_53); +lean_inc(x_48); +lean_inc(x_52); +lean_ctor_set(x_3, 2, x_48); +x_54 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_expandCDot_x3f_go___spec__2(x_2, x_3, x_4); +x_55 = !lean_is_exclusive(x_54); +if (x_55 == 0) +{ +lean_object* x_56; uint8_t x_57; +x_56 = lean_ctor_get(x_54, 0); +x_57 = !lean_is_exclusive(x_56); +if (x_57 == 0) +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_58 = lean_ctor_get(x_56, 0); +x_59 = lean_ctor_get(x_56, 1); +x_60 = l_Lean_Elab_Term_expandCDot_x3f_go___closed__4; +x_61 = l_Lean_addMacroScope(x_52, x_60, x_48); +x_62 = lean_box(0); +x_63 = l_Lean_Elab_Term_expandCDot_x3f_go___closed__3; +x_64 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_64, 0, x_58); +lean_ctor_set(x_64, 1, x_63); +lean_ctor_set(x_64, 2, x_61); +lean_ctor_set(x_64, 3, x_62); +lean_inc(x_64); +x_65 = lean_array_push(x_59, x_64); +lean_ctor_set(x_56, 1, x_65); +lean_ctor_set(x_56, 0, x_64); +return x_54; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_53 = lean_ctor_get(x_43, 0); -x_54 = lean_ctor_get(x_43, 1); -lean_inc(x_54); -lean_inc(x_53); -lean_dec(x_43); -x_55 = l_Lean_Elab_Term_expandCDot_x3f_go___closed__4; -x_56 = l_Lean_addMacroScope(x_39, x_55, x_35); -x_57 = lean_box(0); -x_58 = l_Lean_Elab_Term_expandCDot_x3f_go___closed__3; -x_59 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_59, 0, x_53); -lean_ctor_set(x_59, 1, x_58); -lean_ctor_set(x_59, 2, x_56); -lean_ctor_set(x_59, 3, x_57); -lean_inc(x_59); -x_60 = lean_array_push(x_54, x_59); -x_61 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_61, 0, x_59); -lean_ctor_set(x_61, 1, x_60); -lean_ctor_set(x_41, 0, x_61); -return x_41; +lean_object* x_66; 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; +x_66 = lean_ctor_get(x_56, 0); +x_67 = lean_ctor_get(x_56, 1); +lean_inc(x_67); +lean_inc(x_66); +lean_dec(x_56); +x_68 = l_Lean_Elab_Term_expandCDot_x3f_go___closed__4; +x_69 = l_Lean_addMacroScope(x_52, x_68, x_48); +x_70 = lean_box(0); +x_71 = l_Lean_Elab_Term_expandCDot_x3f_go___closed__3; +x_72 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_72, 0, x_66); +lean_ctor_set(x_72, 1, x_71); +lean_ctor_set(x_72, 2, x_69); +lean_ctor_set(x_72, 3, x_70); +lean_inc(x_72); +x_73 = lean_array_push(x_67, x_72); +x_74 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_74, 0, x_72); +lean_ctor_set(x_74, 1, x_73); +lean_ctor_set(x_54, 0, x_74); +return x_54; } } else { -lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; 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; -x_62 = lean_ctor_get(x_41, 0); -x_63 = lean_ctor_get(x_41, 1); -lean_inc(x_63); -lean_inc(x_62); -lean_dec(x_41); -x_64 = lean_ctor_get(x_62, 0); -lean_inc(x_64); -x_65 = lean_ctor_get(x_62, 1); -lean_inc(x_65); -if (lean_is_exclusive(x_62)) { - lean_ctor_release(x_62, 0); - lean_ctor_release(x_62, 1); - x_66 = x_62; +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_75 = lean_ctor_get(x_54, 0); +x_76 = lean_ctor_get(x_54, 1); +lean_inc(x_76); +lean_inc(x_75); +lean_dec(x_54); +x_77 = lean_ctor_get(x_75, 0); +lean_inc(x_77); +x_78 = lean_ctor_get(x_75, 1); +lean_inc(x_78); +if (lean_is_exclusive(x_75)) { + lean_ctor_release(x_75, 0); + lean_ctor_release(x_75, 1); + x_79 = x_75; } else { - lean_dec_ref(x_62); - x_66 = lean_box(0); + lean_dec_ref(x_75); + x_79 = lean_box(0); } -x_67 = l_Lean_Elab_Term_expandCDot_x3f_go___closed__4; -x_68 = l_Lean_addMacroScope(x_39, x_67, x_35); -x_69 = lean_box(0); -x_70 = l_Lean_Elab_Term_expandCDot_x3f_go___closed__3; -x_71 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_71, 0, x_64); -lean_ctor_set(x_71, 1, x_70); -lean_ctor_set(x_71, 2, x_68); -lean_ctor_set(x_71, 3, x_69); -lean_inc(x_71); -x_72 = lean_array_push(x_65, x_71); -if (lean_is_scalar(x_66)) { - x_73 = lean_alloc_ctor(0, 2, 0); +x_80 = l_Lean_Elab_Term_expandCDot_x3f_go___closed__4; +x_81 = l_Lean_addMacroScope(x_52, x_80, x_48); +x_82 = lean_box(0); +x_83 = l_Lean_Elab_Term_expandCDot_x3f_go___closed__3; +x_84 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_84, 0, x_77); +lean_ctor_set(x_84, 1, x_83); +lean_ctor_set(x_84, 2, x_81); +lean_ctor_set(x_84, 3, x_82); +lean_inc(x_84); +x_85 = lean_array_push(x_78, x_84); +if (lean_is_scalar(x_79)) { + x_86 = lean_alloc_ctor(0, 2, 0); } else { - x_73 = x_66; + x_86 = x_79; } -lean_ctor_set(x_73, 0, x_71); -lean_ctor_set(x_73, 1, x_72); -x_74 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_74, 0, x_73); -lean_ctor_set(x_74, 1, x_63); -return x_74; +lean_ctor_set(x_86, 0, x_84); +lean_ctor_set(x_86, 1, x_85); +x_87 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_87, 0, x_86); +lean_ctor_set(x_87, 1, x_76); +return x_87; } } else { -lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; -x_75 = lean_ctor_get(x_3, 0); -x_76 = lean_ctor_get(x_3, 1); -x_77 = lean_ctor_get(x_3, 3); -x_78 = lean_ctor_get(x_3, 4); -x_79 = lean_ctor_get(x_3, 5); -lean_inc(x_79); -lean_inc(x_78); -lean_inc(x_77); -lean_inc(x_76); -lean_inc(x_75); +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; +x_88 = lean_ctor_get(x_3, 0); +x_89 = lean_ctor_get(x_3, 1); +x_90 = lean_ctor_get(x_3, 3); +x_91 = lean_ctor_get(x_3, 4); +x_92 = lean_ctor_get(x_3, 5); +lean_inc(x_92); +lean_inc(x_91); +lean_inc(x_90); +lean_inc(x_89); +lean_inc(x_88); lean_dec(x_3); -lean_inc(x_35); -lean_inc(x_76); -x_80 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_80, 0, x_75); -lean_ctor_set(x_80, 1, x_76); -lean_ctor_set(x_80, 2, x_35); -lean_ctor_set(x_80, 3, x_77); -lean_ctor_set(x_80, 4, x_78); -lean_ctor_set(x_80, 5, x_79); -x_81 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_expandCDot_x3f_go___spec__2(x_2, x_80, x_4); -x_82 = lean_ctor_get(x_81, 0); -lean_inc(x_82); -x_83 = lean_ctor_get(x_81, 1); -lean_inc(x_83); -if (lean_is_exclusive(x_81)) { - lean_ctor_release(x_81, 0); - lean_ctor_release(x_81, 1); - x_84 = x_81; +lean_inc(x_48); +lean_inc(x_89); +x_93 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_93, 0, x_88); +lean_ctor_set(x_93, 1, x_89); +lean_ctor_set(x_93, 2, x_48); +lean_ctor_set(x_93, 3, x_90); +lean_ctor_set(x_93, 4, x_91); +lean_ctor_set(x_93, 5, x_92); +x_94 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_expandCDot_x3f_go___spec__2(x_2, x_93, x_4); +x_95 = lean_ctor_get(x_94, 0); +lean_inc(x_95); +x_96 = lean_ctor_get(x_94, 1); +lean_inc(x_96); +if (lean_is_exclusive(x_94)) { + lean_ctor_release(x_94, 0); + lean_ctor_release(x_94, 1); + x_97 = x_94; } else { - lean_dec_ref(x_81); - x_84 = lean_box(0); + lean_dec_ref(x_94); + x_97 = lean_box(0); } -x_85 = lean_ctor_get(x_82, 0); -lean_inc(x_85); -x_86 = lean_ctor_get(x_82, 1); -lean_inc(x_86); -if (lean_is_exclusive(x_82)) { - lean_ctor_release(x_82, 0); - lean_ctor_release(x_82, 1); - x_87 = x_82; +x_98 = lean_ctor_get(x_95, 0); +lean_inc(x_98); +x_99 = lean_ctor_get(x_95, 1); +lean_inc(x_99); +if (lean_is_exclusive(x_95)) { + lean_ctor_release(x_95, 0); + lean_ctor_release(x_95, 1); + x_100 = x_95; } else { - lean_dec_ref(x_82); - x_87 = lean_box(0); -} -x_88 = l_Lean_Elab_Term_expandCDot_x3f_go___closed__4; -x_89 = l_Lean_addMacroScope(x_76, x_88, x_35); -x_90 = lean_box(0); -x_91 = l_Lean_Elab_Term_expandCDot_x3f_go___closed__3; -x_92 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_92, 0, x_85); -lean_ctor_set(x_92, 1, x_91); -lean_ctor_set(x_92, 2, x_89); -lean_ctor_set(x_92, 3, x_90); -lean_inc(x_92); -x_93 = lean_array_push(x_86, x_92); -if (lean_is_scalar(x_87)) { - x_94 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_95); + x_100 = lean_box(0); +} +x_101 = l_Lean_Elab_Term_expandCDot_x3f_go___closed__4; +x_102 = l_Lean_addMacroScope(x_89, x_101, x_48); +x_103 = lean_box(0); +x_104 = l_Lean_Elab_Term_expandCDot_x3f_go___closed__3; +x_105 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_105, 0, x_98); +lean_ctor_set(x_105, 1, x_104); +lean_ctor_set(x_105, 2, x_102); +lean_ctor_set(x_105, 3, x_103); +lean_inc(x_105); +x_106 = lean_array_push(x_99, x_105); +if (lean_is_scalar(x_100)) { + x_107 = lean_alloc_ctor(0, 2, 0); } else { - x_94 = x_87; + x_107 = x_100; } -lean_ctor_set(x_94, 0, x_92); -lean_ctor_set(x_94, 1, x_93); -if (lean_is_scalar(x_84)) { - x_95 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_107, 0, x_105); +lean_ctor_set(x_107, 1, x_106); +if (lean_is_scalar(x_97)) { + x_108 = lean_alloc_ctor(0, 2, 0); } else { - x_95 = x_84; + x_108 = x_97; } -lean_ctor_set(x_95, 0, x_94); -lean_ctor_set(x_95, 1, x_83); -return x_95; +lean_ctor_set(x_108, 0, x_107); +lean_ctor_set(x_108, 1, x_96); +return x_108; } } else { -lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; -x_96 = lean_ctor_get(x_4, 0); -x_97 = lean_ctor_get(x_4, 1); -lean_inc(x_97); -lean_inc(x_96); +lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; +x_109 = lean_ctor_get(x_4, 0); +x_110 = lean_ctor_get(x_4, 1); +lean_inc(x_110); +lean_inc(x_109); lean_dec(x_4); -x_98 = lean_unsigned_to_nat(1u); -x_99 = lean_nat_add(x_96, x_98); -x_100 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_100, 0, x_99); -lean_ctor_set(x_100, 1, x_97); -x_101 = lean_ctor_get(x_3, 0); -lean_inc(x_101); -x_102 = lean_ctor_get(x_3, 1); -lean_inc(x_102); -x_103 = lean_ctor_get(x_3, 3); -lean_inc(x_103); -x_104 = lean_ctor_get(x_3, 4); -lean_inc(x_104); -x_105 = lean_ctor_get(x_3, 5); -lean_inc(x_105); +x_111 = lean_unsigned_to_nat(1u); +x_112 = lean_nat_add(x_109, x_111); +x_113 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_113, 0, x_112); +lean_ctor_set(x_113, 1, x_110); +x_114 = lean_ctor_get(x_3, 0); +lean_inc(x_114); +x_115 = lean_ctor_get(x_3, 1); +lean_inc(x_115); +x_116 = lean_ctor_get(x_3, 3); +lean_inc(x_116); +x_117 = lean_ctor_get(x_3, 4); +lean_inc(x_117); +x_118 = lean_ctor_get(x_3, 5); +lean_inc(x_118); if (lean_is_exclusive(x_3)) { lean_ctor_release(x_3, 0); lean_ctor_release(x_3, 1); @@ -12544,279 +12713,411 @@ if (lean_is_exclusive(x_3)) { lean_ctor_release(x_3, 3); lean_ctor_release(x_3, 4); lean_ctor_release(x_3, 5); - x_106 = x_3; + x_119 = x_3; } else { lean_dec_ref(x_3); - x_106 = lean_box(0); -} -lean_inc(x_96); -lean_inc(x_102); -if (lean_is_scalar(x_106)) { - x_107 = lean_alloc_ctor(0, 6, 0); -} else { - x_107 = x_106; + x_119 = lean_box(0); } -lean_ctor_set(x_107, 0, x_101); -lean_ctor_set(x_107, 1, x_102); -lean_ctor_set(x_107, 2, x_96); -lean_ctor_set(x_107, 3, x_103); -lean_ctor_set(x_107, 4, x_104); -lean_ctor_set(x_107, 5, x_105); -x_108 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_expandCDot_x3f_go___spec__2(x_2, x_107, x_100); -x_109 = lean_ctor_get(x_108, 0); lean_inc(x_109); -x_110 = lean_ctor_get(x_108, 1); -lean_inc(x_110); -if (lean_is_exclusive(x_108)) { - lean_ctor_release(x_108, 0); - lean_ctor_release(x_108, 1); - x_111 = x_108; +lean_inc(x_115); +if (lean_is_scalar(x_119)) { + x_120 = lean_alloc_ctor(0, 6, 0); } else { - lean_dec_ref(x_108); - x_111 = lean_box(0); + x_120 = x_119; +} +lean_ctor_set(x_120, 0, x_114); +lean_ctor_set(x_120, 1, x_115); +lean_ctor_set(x_120, 2, x_109); +lean_ctor_set(x_120, 3, x_116); +lean_ctor_set(x_120, 4, x_117); +lean_ctor_set(x_120, 5, x_118); +x_121 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_expandCDot_x3f_go___spec__2(x_2, x_120, x_113); +x_122 = lean_ctor_get(x_121, 0); +lean_inc(x_122); +x_123 = lean_ctor_get(x_121, 1); +lean_inc(x_123); +if (lean_is_exclusive(x_121)) { + lean_ctor_release(x_121, 0); + lean_ctor_release(x_121, 1); + x_124 = x_121; +} else { + lean_dec_ref(x_121); + x_124 = lean_box(0); } -x_112 = lean_ctor_get(x_109, 0); -lean_inc(x_112); -x_113 = lean_ctor_get(x_109, 1); -lean_inc(x_113); -if (lean_is_exclusive(x_109)) { - lean_ctor_release(x_109, 0); - lean_ctor_release(x_109, 1); - x_114 = x_109; +x_125 = lean_ctor_get(x_122, 0); +lean_inc(x_125); +x_126 = lean_ctor_get(x_122, 1); +lean_inc(x_126); +if (lean_is_exclusive(x_122)) { + lean_ctor_release(x_122, 0); + lean_ctor_release(x_122, 1); + x_127 = x_122; } else { - lean_dec_ref(x_109); - x_114 = lean_box(0); -} -x_115 = l_Lean_Elab_Term_expandCDot_x3f_go___closed__4; -x_116 = l_Lean_addMacroScope(x_102, x_115, x_96); -x_117 = lean_box(0); -x_118 = l_Lean_Elab_Term_expandCDot_x3f_go___closed__3; -x_119 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_119, 0, x_112); -lean_ctor_set(x_119, 1, x_118); -lean_ctor_set(x_119, 2, x_116); -lean_ctor_set(x_119, 3, x_117); -lean_inc(x_119); -x_120 = lean_array_push(x_113, x_119); -if (lean_is_scalar(x_114)) { - x_121 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_122); + x_127 = lean_box(0); +} +x_128 = l_Lean_Elab_Term_expandCDot_x3f_go___closed__4; +x_129 = l_Lean_addMacroScope(x_115, x_128, x_109); +x_130 = lean_box(0); +x_131 = l_Lean_Elab_Term_expandCDot_x3f_go___closed__3; +x_132 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_132, 0, x_125); +lean_ctor_set(x_132, 1, x_131); +lean_ctor_set(x_132, 2, x_129); +lean_ctor_set(x_132, 3, x_130); +lean_inc(x_132); +x_133 = lean_array_push(x_126, x_132); +if (lean_is_scalar(x_127)) { + x_134 = lean_alloc_ctor(0, 2, 0); } else { - x_121 = x_114; + x_134 = x_127; } -lean_ctor_set(x_121, 0, x_119); -lean_ctor_set(x_121, 1, x_120); -if (lean_is_scalar(x_111)) { - x_122 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_134, 0, x_132); +lean_ctor_set(x_134, 1, x_133); +if (lean_is_scalar(x_124)) { + x_135 = lean_alloc_ctor(0, 2, 0); } else { - x_122 = x_111; + x_135 = x_124; } -lean_ctor_set(x_122, 0, x_121); -lean_ctor_set(x_122, 1, x_110); -return x_122; +lean_ctor_set(x_135, 0, x_134); +lean_ctor_set(x_135, 1, x_123); +return x_135; } } } else { -lean_object* x_123; uint8_t x_124; -lean_dec(x_1); -x_123 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_hasCDot___closed__2; -x_124 = lean_name_eq(x_6, x_123); -if (x_124 == 0) +lean_object* x_136; lean_object* x_137; lean_object* x_138; uint8_t x_139; +x_136 = lean_unsigned_to_nat(1u); +x_137 = l_Lean_Syntax_getArg(x_1, x_136); +x_138 = lean_unsigned_to_nat(2u); +lean_inc(x_137); +x_139 = l_Lean_Syntax_matchesNull(x_137, x_138); +if (x_139 == 0) { -lean_object* x_125; size_t x_126; size_t x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; -x_125 = lean_array_get_size(x_7); -x_126 = lean_usize_of_nat(x_125); -lean_dec(x_125); -x_127 = 0; -x_128 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandCDot_x3f_go___spec__1(x_126, x_127, x_7, x_2, x_3, x_4); -x_129 = lean_ctor_get(x_128, 0); -lean_inc(x_129); -x_130 = lean_ctor_get(x_128, 1); -lean_inc(x_130); -if (lean_is_exclusive(x_128)) { - lean_ctor_release(x_128, 0); - lean_ctor_release(x_128, 1); - x_131 = x_128; -} else { - lean_dec_ref(x_128); - x_131 = lean_box(0); +lean_dec(x_137); +if (lean_obj_tag(x_1) == 1) +{ +uint8_t x_140; +x_140 = !lean_is_exclusive(x_1); +if (x_140 == 0) +{ +lean_object* x_141; lean_object* x_142; size_t x_143; size_t x_144; lean_object* x_145; uint8_t x_146; +x_141 = lean_ctor_get(x_1, 2); +x_142 = lean_array_get_size(x_141); +x_143 = lean_usize_of_nat(x_142); +lean_dec(x_142); +x_144 = 0; +x_145 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandCDot_x3f_go___spec__1(x_143, x_144, x_141, x_2, x_3, x_4); +x_146 = !lean_is_exclusive(x_145); +if (x_146 == 0) +{ +lean_object* x_147; uint8_t x_148; +x_147 = lean_ctor_get(x_145, 0); +x_148 = !lean_is_exclusive(x_147); +if (x_148 == 0) +{ +lean_object* x_149; +x_149 = lean_ctor_get(x_147, 0); +lean_ctor_set(x_1, 2, x_149); +lean_ctor_set(x_147, 0, x_1); +return x_145; } -x_132 = lean_ctor_get(x_129, 0); -lean_inc(x_132); -x_133 = lean_ctor_get(x_129, 1); -lean_inc(x_133); -if (lean_is_exclusive(x_129)) { - lean_ctor_release(x_129, 0); - lean_ctor_release(x_129, 1); - x_134 = x_129; -} else { - lean_dec_ref(x_129); - x_134 = lean_box(0); -} -x_135 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_135, 0, x_5); -lean_ctor_set(x_135, 1, x_6); -lean_ctor_set(x_135, 2, x_132); -if (lean_is_scalar(x_134)) { - x_136 = lean_alloc_ctor(0, 2, 0); +else +{ +lean_object* x_150; lean_object* x_151; lean_object* x_152; +x_150 = lean_ctor_get(x_147, 0); +x_151 = lean_ctor_get(x_147, 1); +lean_inc(x_151); +lean_inc(x_150); +lean_dec(x_147); +lean_ctor_set(x_1, 2, x_150); +x_152 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_152, 0, x_1); +lean_ctor_set(x_152, 1, x_151); +lean_ctor_set(x_145, 0, x_152); +return x_145; +} +} +else +{ +lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; +x_153 = lean_ctor_get(x_145, 0); +x_154 = lean_ctor_get(x_145, 1); +lean_inc(x_154); +lean_inc(x_153); +lean_dec(x_145); +x_155 = lean_ctor_get(x_153, 0); +lean_inc(x_155); +x_156 = lean_ctor_get(x_153, 1); +lean_inc(x_156); +if (lean_is_exclusive(x_153)) { + lean_ctor_release(x_153, 0); + lean_ctor_release(x_153, 1); + x_157 = x_153; } else { - x_136 = x_134; + lean_dec_ref(x_153); + x_157 = lean_box(0); } -lean_ctor_set(x_136, 0, x_135); -lean_ctor_set(x_136, 1, x_133); -if (lean_is_scalar(x_131)) { - x_137 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1, 2, x_155); +if (lean_is_scalar(x_157)) { + x_158 = lean_alloc_ctor(0, 2, 0); } else { - x_137 = x_131; + x_158 = x_157; +} +lean_ctor_set(x_158, 0, x_1); +lean_ctor_set(x_158, 1, x_156); +x_159 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_159, 0, x_158); +lean_ctor_set(x_159, 1, x_154); +return x_159; } -lean_ctor_set(x_137, 0, x_136); -lean_ctor_set(x_137, 1, x_130); -return x_137; } else { -lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -x_138 = lean_ctor_get(x_4, 0); -lean_inc(x_138); -x_139 = lean_ctor_get(x_4, 1); -lean_inc(x_139); -if (lean_is_exclusive(x_4)) { - lean_ctor_release(x_4, 0); - lean_ctor_release(x_4, 1); - x_140 = x_4; +lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; size_t x_164; size_t x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; +x_160 = lean_ctor_get(x_1, 0); +x_161 = lean_ctor_get(x_1, 1); +x_162 = lean_ctor_get(x_1, 2); +lean_inc(x_162); +lean_inc(x_161); +lean_inc(x_160); +lean_dec(x_1); +x_163 = lean_array_get_size(x_162); +x_164 = lean_usize_of_nat(x_163); +lean_dec(x_163); +x_165 = 0; +x_166 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandCDot_x3f_go___spec__1(x_164, x_165, x_162, x_2, x_3, x_4); +x_167 = lean_ctor_get(x_166, 0); +lean_inc(x_167); +x_168 = lean_ctor_get(x_166, 1); +lean_inc(x_168); +if (lean_is_exclusive(x_166)) { + lean_ctor_release(x_166, 0); + lean_ctor_release(x_166, 1); + x_169 = x_166; } else { - lean_dec_ref(x_4); - x_140 = lean_box(0); + lean_dec_ref(x_166); + x_169 = lean_box(0); } -x_141 = lean_unsigned_to_nat(1u); -x_142 = lean_nat_add(x_138, x_141); -if (lean_is_scalar(x_140)) { - x_143 = lean_alloc_ctor(0, 2, 0); +x_170 = lean_ctor_get(x_167, 0); +lean_inc(x_170); +x_171 = lean_ctor_get(x_167, 1); +lean_inc(x_171); +if (lean_is_exclusive(x_167)) { + lean_ctor_release(x_167, 0); + lean_ctor_release(x_167, 1); + x_172 = x_167; } else { - x_143 = x_140; + lean_dec_ref(x_167); + x_172 = lean_box(0); } -lean_ctor_set(x_143, 0, x_142); -lean_ctor_set(x_143, 1, x_139); -x_144 = lean_ctor_get(x_3, 0); -lean_inc(x_144); -x_145 = lean_ctor_get(x_3, 1); -lean_inc(x_145); -x_146 = lean_ctor_get(x_3, 3); -lean_inc(x_146); -x_147 = lean_ctor_get(x_3, 4); -lean_inc(x_147); -x_148 = lean_ctor_get(x_3, 5); -lean_inc(x_148); -if (lean_is_exclusive(x_3)) { - lean_ctor_release(x_3, 0); - lean_ctor_release(x_3, 1); - lean_ctor_release(x_3, 2); - lean_ctor_release(x_3, 3); - lean_ctor_release(x_3, 4); - lean_ctor_release(x_3, 5); - x_149 = x_3; +x_173 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_173, 0, x_160); +lean_ctor_set(x_173, 1, x_161); +lean_ctor_set(x_173, 2, x_170); +if (lean_is_scalar(x_172)) { + x_174 = lean_alloc_ctor(0, 2, 0); } else { - lean_dec_ref(x_3); - x_149 = lean_box(0); + x_174 = x_172; } -lean_inc(x_138); -lean_inc(x_145); -if (lean_is_scalar(x_149)) { - x_150 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_174, 0, x_173); +lean_ctor_set(x_174, 1, x_171); +if (lean_is_scalar(x_169)) { + x_175 = lean_alloc_ctor(0, 2, 0); } else { - x_150 = x_149; -} -lean_ctor_set(x_150, 0, x_144); -lean_ctor_set(x_150, 1, x_145); -lean_ctor_set(x_150, 2, x_138); -lean_ctor_set(x_150, 3, x_146); -lean_ctor_set(x_150, 4, x_147); -lean_ctor_set(x_150, 5, x_148); -x_151 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_expandCDot_x3f_go___spec__2(x_2, x_150, x_143); -x_152 = lean_ctor_get(x_151, 0); -lean_inc(x_152); -x_153 = lean_ctor_get(x_151, 1); -lean_inc(x_153); -if (lean_is_exclusive(x_151)) { - lean_ctor_release(x_151, 0); - lean_ctor_release(x_151, 1); - x_154 = x_151; + x_175 = x_169; +} +lean_ctor_set(x_175, 0, x_174); +lean_ctor_set(x_175, 1, x_168); +return x_175; +} +} +else +{ +lean_object* x_176; lean_object* x_177; +lean_dec(x_3); +x_176 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_176, 0, x_1); +lean_ctor_set(x_176, 1, x_2); +x_177 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_177, 0, x_176); +lean_ctor_set(x_177, 1, x_4); +return x_177; +} +} +else +{ +lean_object* x_178; lean_object* x_179; uint8_t x_180; +x_178 = l_Lean_Syntax_getArg(x_137, x_136); +lean_dec(x_137); +x_179 = lean_unsigned_to_nat(0u); +x_180 = l_Lean_Syntax_matchesNull(x_178, x_179); +if (x_180 == 0) +{ +if (lean_obj_tag(x_1) == 1) +{ +uint8_t x_181; +x_181 = !lean_is_exclusive(x_1); +if (x_181 == 0) +{ +lean_object* x_182; lean_object* x_183; size_t x_184; size_t x_185; lean_object* x_186; uint8_t x_187; +x_182 = lean_ctor_get(x_1, 2); +x_183 = lean_array_get_size(x_182); +x_184 = lean_usize_of_nat(x_183); +lean_dec(x_183); +x_185 = 0; +x_186 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandCDot_x3f_go___spec__1(x_184, x_185, x_182, x_2, x_3, x_4); +x_187 = !lean_is_exclusive(x_186); +if (x_187 == 0) +{ +lean_object* x_188; uint8_t x_189; +x_188 = lean_ctor_get(x_186, 0); +x_189 = !lean_is_exclusive(x_188); +if (x_189 == 0) +{ +lean_object* x_190; +x_190 = lean_ctor_get(x_188, 0); +lean_ctor_set(x_1, 2, x_190); +lean_ctor_set(x_188, 0, x_1); +return x_186; +} +else +{ +lean_object* x_191; lean_object* x_192; lean_object* x_193; +x_191 = lean_ctor_get(x_188, 0); +x_192 = lean_ctor_get(x_188, 1); +lean_inc(x_192); +lean_inc(x_191); +lean_dec(x_188); +lean_ctor_set(x_1, 2, x_191); +x_193 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_193, 0, x_1); +lean_ctor_set(x_193, 1, x_192); +lean_ctor_set(x_186, 0, x_193); +return x_186; +} +} +else +{ +lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; +x_194 = lean_ctor_get(x_186, 0); +x_195 = lean_ctor_get(x_186, 1); +lean_inc(x_195); +lean_inc(x_194); +lean_dec(x_186); +x_196 = lean_ctor_get(x_194, 0); +lean_inc(x_196); +x_197 = lean_ctor_get(x_194, 1); +lean_inc(x_197); +if (lean_is_exclusive(x_194)) { + lean_ctor_release(x_194, 0); + lean_ctor_release(x_194, 1); + x_198 = x_194; } else { - lean_dec_ref(x_151); - x_154 = lean_box(0); + lean_dec_ref(x_194); + x_198 = lean_box(0); } -x_155 = lean_ctor_get(x_152, 0); -lean_inc(x_155); -x_156 = lean_ctor_get(x_152, 1); -lean_inc(x_156); -if (lean_is_exclusive(x_152)) { - lean_ctor_release(x_152, 0); - lean_ctor_release(x_152, 1); - x_157 = x_152; +lean_ctor_set(x_1, 2, x_196); +if (lean_is_scalar(x_198)) { + x_199 = lean_alloc_ctor(0, 2, 0); } else { - lean_dec_ref(x_152); - x_157 = lean_box(0); + x_199 = x_198; } -x_158 = l_Lean_Elab_Term_expandCDot_x3f_go___closed__4; -x_159 = l_Lean_addMacroScope(x_145, x_158, x_138); -x_160 = lean_box(0); -x_161 = l_Lean_Elab_Term_expandCDot_x3f_go___closed__3; -x_162 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_162, 0, x_155); -lean_ctor_set(x_162, 1, x_161); -lean_ctor_set(x_162, 2, x_159); -lean_ctor_set(x_162, 3, x_160); -lean_inc(x_162); -x_163 = lean_array_push(x_156, x_162); -if (lean_is_scalar(x_157)) { - x_164 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_199, 0, x_1); +lean_ctor_set(x_199, 1, x_197); +x_200 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_200, 0, x_199); +lean_ctor_set(x_200, 1, x_195); +return x_200; +} +} +else +{ +lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; size_t x_205; size_t x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; +x_201 = lean_ctor_get(x_1, 0); +x_202 = lean_ctor_get(x_1, 1); +x_203 = lean_ctor_get(x_1, 2); +lean_inc(x_203); +lean_inc(x_202); +lean_inc(x_201); +lean_dec(x_1); +x_204 = lean_array_get_size(x_203); +x_205 = lean_usize_of_nat(x_204); +lean_dec(x_204); +x_206 = 0; +x_207 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandCDot_x3f_go___spec__1(x_205, x_206, x_203, x_2, x_3, x_4); +x_208 = lean_ctor_get(x_207, 0); +lean_inc(x_208); +x_209 = lean_ctor_get(x_207, 1); +lean_inc(x_209); +if (lean_is_exclusive(x_207)) { + lean_ctor_release(x_207, 0); + lean_ctor_release(x_207, 1); + x_210 = x_207; +} else { + lean_dec_ref(x_207); + x_210 = lean_box(0); +} +x_211 = lean_ctor_get(x_208, 0); +lean_inc(x_211); +x_212 = lean_ctor_get(x_208, 1); +lean_inc(x_212); +if (lean_is_exclusive(x_208)) { + lean_ctor_release(x_208, 0); + lean_ctor_release(x_208, 1); + x_213 = x_208; } else { - x_164 = x_157; + lean_dec_ref(x_208); + x_213 = lean_box(0); } -lean_ctor_set(x_164, 0, x_162); -lean_ctor_set(x_164, 1, x_163); -if (lean_is_scalar(x_154)) { - x_165 = lean_alloc_ctor(0, 2, 0); +x_214 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_214, 0, x_201); +lean_ctor_set(x_214, 1, x_202); +lean_ctor_set(x_214, 2, x_211); +if (lean_is_scalar(x_213)) { + x_215 = lean_alloc_ctor(0, 2, 0); } else { - x_165 = x_154; + x_215 = x_213; } -lean_ctor_set(x_165, 0, x_164); -lean_ctor_set(x_165, 1, x_153); -return x_165; +lean_ctor_set(x_215, 0, x_214); +lean_ctor_set(x_215, 1, x_212); +if (lean_is_scalar(x_210)) { + x_216 = lean_alloc_ctor(0, 2, 0); +} else { + x_216 = x_210; } +lean_ctor_set(x_216, 0, x_215); +lean_ctor_set(x_216, 1, x_209); +return x_216; } } else { -lean_object* x_166; lean_object* x_167; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_217; lean_object* x_218; lean_dec(x_3); -x_166 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_166, 0, x_1); -lean_ctor_set(x_166, 1, x_2); -x_167 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_167, 0, x_166); -lean_ctor_set(x_167, 1, x_4); -return x_167; +x_217 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_217, 0, x_1); +lean_ctor_set(x_217, 1, x_2); +x_218 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_218, 0, x_217); +lean_ctor_set(x_218, 1, x_4); +return x_218; } } else { -lean_object* x_168; lean_object* x_169; +lean_object* x_219; lean_object* x_220; lean_dec(x_3); -x_168 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_168, 0, x_1); -lean_ctor_set(x_168, 1, x_2); -x_169 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_169, 0, x_168); -lean_ctor_set(x_169, 1, x_4); -return x_169; +x_219 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_219, 0, x_1); +lean_ctor_set(x_219, 1, x_2); +x_220 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_220, 0, x_219); +lean_ctor_set(x_220, 1, x_4); +return x_220; +} +} } } } @@ -12871,7 +13172,7 @@ 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; lean_object* 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; x_15 = lean_ctor_get(x_13, 0); -x_16 = l_Lean_Elab_Term_expandDbgTrace___closed__16; +x_16 = l_Lean_Elab_Term_expandDbgTrace___closed__19; lean_inc(x_15); x_17 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_17, 0, x_15); @@ -12883,7 +13184,7 @@ x_21 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_21, 0, x_19); lean_ctor_set(x_21, 1, x_20); lean_ctor_set(x_21, 2, x_18); -x_22 = l_Lean_Elab_Term_expandDbgTrace___closed__23; +x_22 = l_Lean_Elab_Term_expandDbgTrace___closed__26; x_23 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_23, 0, x_15); lean_ctor_set(x_23, 1, x_22); @@ -12891,7 +13192,7 @@ x_24 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; x_25 = lean_array_push(x_24, x_21); x_26 = lean_array_push(x_25, x_23); x_27 = lean_array_push(x_26, x_11); -x_28 = l_Lean_Elab_Term_expandDbgTrace___closed__19; +x_28 = l_Lean_Elab_Term_expandDbgTrace___closed__22; x_29 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_29, 0, x_19); lean_ctor_set(x_29, 1, x_28); @@ -12899,7 +13200,7 @@ lean_ctor_set(x_29, 2, x_27); x_30 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; x_31 = lean_array_push(x_30, x_17); x_32 = lean_array_push(x_31, x_29); -x_33 = l_Lean_Elab_Term_expandDbgTrace___closed__17; +x_33 = l_Lean_Elab_Term_expandDbgTrace___closed__20; x_34 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_34, 0, x_19); lean_ctor_set(x_34, 1, x_33); @@ -12917,7 +13218,7 @@ x_37 = lean_ctor_get(x_13, 1); lean_inc(x_37); lean_inc(x_36); lean_dec(x_13); -x_38 = l_Lean_Elab_Term_expandDbgTrace___closed__16; +x_38 = l_Lean_Elab_Term_expandDbgTrace___closed__19; lean_inc(x_36); x_39 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_39, 0, x_36); @@ -12929,7 +13230,7 @@ x_43 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_43, 0, x_41); lean_ctor_set(x_43, 1, x_42); lean_ctor_set(x_43, 2, x_40); -x_44 = l_Lean_Elab_Term_expandDbgTrace___closed__23; +x_44 = l_Lean_Elab_Term_expandDbgTrace___closed__26; x_45 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_45, 0, x_36); lean_ctor_set(x_45, 1, x_44); @@ -12937,7 +13238,7 @@ x_46 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; x_47 = lean_array_push(x_46, x_43); x_48 = lean_array_push(x_47, x_45); x_49 = lean_array_push(x_48, x_11); -x_50 = l_Lean_Elab_Term_expandDbgTrace___closed__19; +x_50 = l_Lean_Elab_Term_expandDbgTrace___closed__22; x_51 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_51, 0, x_41); lean_ctor_set(x_51, 1, x_50); @@ -12945,7 +13246,7 @@ lean_ctor_set(x_51, 2, x_49); x_52 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; x_53 = lean_array_push(x_52, x_39); x_54 = lean_array_push(x_53, x_51); -x_55 = l_Lean_Elab_Term_expandDbgTrace___closed__17; +x_55 = l_Lean_Elab_Term_expandDbgTrace___closed__20; x_56 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_56, 0, x_41); lean_ctor_set(x_56, 1, x_55); @@ -12975,41 +13276,40 @@ return x_6; } else { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; x_7 = lean_unsigned_to_nat(1u); x_8 = l_Lean_Syntax_getArg(x_1, x_7); -x_9 = l_Lean_nullKind; -x_10 = lean_unsigned_to_nat(2u); +x_9 = lean_unsigned_to_nat(2u); lean_inc(x_8); -x_11 = l_Lean_Syntax_isNodeOf(x_8, x_9, x_10); -if (x_11 == 0) +x_10 = l_Lean_Syntax_matchesNull(x_8, x_9); +if (x_10 == 0) { -lean_object* x_12; +lean_object* x_11; lean_dec(x_8); -x_12 = l_Lean_Elab_Term_expandCDot_x3f(x_1, x_2, x_3); -return x_12; +x_11 = l_Lean_Elab_Term_expandCDot_x3f(x_1, x_2, x_3); +return x_11; } else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_13 = lean_unsigned_to_nat(0u); -x_14 = l_Lean_Syntax_getArg(x_8, x_13); -x_15 = l_Lean_Syntax_getArg(x_8, x_7); +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = lean_unsigned_to_nat(0u); +x_13 = l_Lean_Syntax_getArg(x_8, x_12); +x_14 = l_Lean_Syntax_getArg(x_8, x_7); lean_dec(x_8); -x_16 = l_Lean_Syntax_isNodeOf(x_15, x_9, x_13); -if (x_16 == 0) +x_15 = l_Lean_Syntax_matchesNull(x_14, x_12); +if (x_15 == 0) { -lean_object* x_17; -lean_dec(x_14); -x_17 = l_Lean_Elab_Term_expandCDot_x3f(x_1, x_2, x_3); -return x_17; +lean_object* x_16; +lean_dec(x_13); +x_16 = l_Lean_Elab_Term_expandCDot_x3f(x_1, x_2, x_3); +return x_16; } else { -lean_object* x_18; +lean_object* x_17; lean_dec(x_1); -x_18 = l_Lean_Elab_Term_expandCDot_x3f(x_14, x_2, x_3); -return x_18; +x_17 = l_Lean_Elab_Term_expandCDot_x3f(x_13, x_2, x_3); +return x_17; } } } @@ -13655,81 +13955,81 @@ return x_74; } } } -LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__6(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_6; uint8_t x_7; -x_6 = lean_array_get_size(x_3); -x_7 = lean_nat_dec_lt(x_5, x_6); +lean_object* x_7; uint8_t x_8; +x_7 = lean_array_get_size(x_4); +x_8 = lean_nat_dec_lt(x_6, x_7); +lean_dec(x_7); +if (x_8 == 0) +{ +uint8_t x_9; lean_dec(x_6); -if (x_7 == 0) -{ -uint8_t x_8; -lean_dec(x_5); -x_8 = 1; -return x_8; +x_9 = 1; +return x_9; } else { -lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_9 = lean_array_fget(x_3, x_5); -x_10 = lean_array_fget(x_4, x_5); -x_11 = l_Lean_Syntax_structEq(x_9, x_10); -if (x_11 == 0) +lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_10 = lean_array_fget(x_4, x_6); +x_11 = lean_array_fget(x_5, x_6); +x_12 = l_Lean_Syntax_structEq(x_10, x_11); +if (x_12 == 0) { -uint8_t x_12; -lean_dec(x_5); -x_12 = 0; -return x_12; +uint8_t x_13; +lean_dec(x_6); +x_13 = 0; +return x_13; } else { -lean_object* x_13; lean_object* x_14; -x_13 = lean_unsigned_to_nat(1u); -x_14 = lean_nat_add(x_5, x_13); -lean_dec(x_5); -x_2 = lean_box(0); -x_5 = x_14; +lean_object* x_14; lean_object* x_15; +x_14 = lean_unsigned_to_nat(1u); +x_15 = lean_nat_add(x_6, x_14); +lean_dec(x_6); +x_3 = lean_box(0); +x_6 = x_15; goto _start; } } } } -LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__7(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_6; uint8_t x_7; -x_6 = lean_array_get_size(x_3); -x_7 = lean_nat_dec_lt(x_5, x_6); -lean_dec(x_6); -if (x_7 == 0) +lean_object* x_7; uint8_t x_8; +x_7 = lean_array_get_size(x_4); +x_8 = lean_nat_dec_lt(x_6, x_7); +lean_dec(x_7); +if (x_8 == 0) { -uint8_t x_8; -lean_dec(x_5); -x_8 = 1; -return x_8; +uint8_t x_9; +lean_dec(x_6); +x_9 = 1; +return x_9; } else { -lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_9 = lean_array_fget(x_3, x_5); -x_10 = lean_array_fget(x_4, x_5); -x_11 = l_Lean_Syntax_structEq(x_9, x_10); -if (x_11 == 0) +lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_10 = lean_array_fget(x_4, x_6); +x_11 = lean_array_fget(x_5, x_6); +x_12 = l_Lean_Syntax_structEq(x_10, x_11); +if (x_12 == 0) { -uint8_t x_12; -lean_dec(x_5); -x_12 = 0; -return x_12; +uint8_t x_13; +lean_dec(x_6); +x_13 = 0; +return x_13; } else { -lean_object* x_13; lean_object* x_14; -x_13 = lean_unsigned_to_nat(1u); -x_14 = lean_nat_add(x_5, x_13); -lean_dec(x_5); -x_2 = lean_box(0); -x_5 = x_14; +lean_object* x_14; lean_object* x_15; +x_14 = lean_unsigned_to_nat(1u); +x_15 = lean_nat_add(x_6, x_14); +lean_dec(x_6); +x_3 = lean_box(0); +x_6 = x_15; goto _start; } } @@ -13775,6 +14075,36 @@ static lean_object* _init_l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__5 _start: { lean_object* x_1; +x_1 = lean_mk_string_from_bytes("funBinder", 9); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; +x_2 = l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__5; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__6; +x_3 = lean_alloc_ctor(1, 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_Elab_Term_elabCDotFunctionAlias_x3f___closed__8() { +_start: +{ +lean_object* x_1; x_1 = lean_mk_string_from_bytes("term", 4); return x_1; } @@ -13856,7 +14186,7 @@ if (x_22 == 0) lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; x_23 = lean_ctor_get(x_21, 0); x_24 = lean_ctor_get(x_21, 1); -x_25 = l_Lean_Elab_Term_expandDbgTrace___closed__17; +x_25 = l_Lean_Elab_Term_expandDbgTrace___closed__20; lean_inc(x_23); x_26 = l_Lean_Syntax_isOfKind(x_23, x_25); if (x_26 == 0) @@ -13878,7 +14208,7 @@ else lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; x_28 = lean_unsigned_to_nat(1u); x_29 = l_Lean_Syntax_getArg(x_23, x_28); -x_30 = l_Lean_Elab_Term_expandDbgTrace___closed__19; +x_30 = l_Lean_Elab_Term_expandDbgTrace___closed__22; lean_inc(x_29); x_31 = l_Lean_Syntax_isOfKind(x_29, x_30); if (x_31 == 0) @@ -13989,14 +14319,15 @@ return x_21; } else { -uint8_t x_57; -x_57 = l_Array_isEqvAux___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__6(x_23, lean_box(0), x_49, x_52, x_33); +lean_object* x_57; uint8_t x_58; +x_57 = l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__7; +x_58 = l_Array_isEqvAux___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__6(x_23, x_57, lean_box(0), x_49, x_52, x_33); lean_dec(x_52); lean_dec(x_49); lean_dec(x_23); -if (x_57 == 0) +if (x_58 == 0) { -lean_object* x_58; +lean_object* x_59; lean_dec(x_42); lean_dec(x_7); lean_dec(x_6); @@ -14004,64 +14335,64 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_58 = lean_box(0); -lean_ctor_set(x_21, 0, x_58); +x_59 = lean_box(0); +lean_ctor_set(x_21, 0, x_59); return x_21; } else { -lean_object* x_59; uint8_t x_60; lean_object* x_61; +lean_object* x_60; uint8_t x_61; lean_object* x_62; lean_free_object(x_21); -x_59 = l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__5; -x_60 = 0; -x_61 = l_Lean_Elab_Term_resolveId_x3f(x_42, x_59, x_60, x_2, x_3, x_4, x_5, x_6, x_7, x_24); -if (lean_obj_tag(x_61) == 0) +x_60 = l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__8; +x_61 = 0; +x_62 = l_Lean_Elab_Term_resolveId_x3f(x_42, x_60, x_61, x_2, x_3, x_4, x_5, x_6, x_7, x_24); +if (lean_obj_tag(x_62) == 0) { -uint8_t x_62; -x_62 = !lean_is_exclusive(x_61); -if (x_62 == 0) +uint8_t x_63; +x_63 = !lean_is_exclusive(x_62); +if (x_63 == 0) { -return x_61; +return x_62; } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_63 = lean_ctor_get(x_61, 0); -x_64 = lean_ctor_get(x_61, 1); +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_62, 0); +x_65 = lean_ctor_get(x_62, 1); +lean_inc(x_65); lean_inc(x_64); -lean_inc(x_63); -lean_dec(x_61); -x_65 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_65, 0, x_63); -lean_ctor_set(x_65, 1, x_64); -return x_65; +lean_dec(x_62); +x_66 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +return x_66; } } else { -uint8_t x_66; -x_66 = !lean_is_exclusive(x_61); -if (x_66 == 0) +uint8_t x_67; +x_67 = !lean_is_exclusive(x_62); +if (x_67 == 0) { -lean_object* x_67; lean_object* x_68; -x_67 = lean_ctor_get(x_61, 0); -lean_dec(x_67); -x_68 = lean_box(0); -lean_ctor_set_tag(x_61, 0); -lean_ctor_set(x_61, 0, x_68); -return x_61; +lean_object* x_68; lean_object* x_69; +x_68 = lean_ctor_get(x_62, 0); +lean_dec(x_68); +x_69 = lean_box(0); +lean_ctor_set_tag(x_62, 0); +lean_ctor_set(x_62, 0, x_69); +return x_62; } else { -lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_69 = lean_ctor_get(x_61, 1); -lean_inc(x_69); -lean_dec(x_61); -x_70 = lean_box(0); -x_71 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_69); -return x_71; +lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_70 = lean_ctor_get(x_62, 1); +lean_inc(x_70); +lean_dec(x_62); +x_71 = lean_box(0); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set(x_72, 1, x_70); +return x_72; } } } @@ -14071,15 +14402,15 @@ return x_71; } else { -lean_object* x_72; lean_object* x_73; uint8_t x_74; -x_72 = l_Lean_Syntax_getArg(x_36, x_33); -x_73 = l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__4; -lean_inc(x_72); -x_74 = l_Lean_Syntax_isOfKind(x_72, x_73); -if (x_74 == 0) +lean_object* x_73; lean_object* x_74; uint8_t x_75; +x_73 = l_Lean_Syntax_getArg(x_36, x_33); +x_74 = l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__4; +lean_inc(x_73); +x_75 = l_Lean_Syntax_isOfKind(x_73, x_74); +if (x_75 == 0) { -lean_object* x_75; -lean_dec(x_72); +lean_object* x_76; +lean_dec(x_73); lean_dec(x_36); lean_dec(x_34); lean_dec(x_23); @@ -14089,30 +14420,30 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_75 = lean_box(0); -lean_ctor_set(x_21, 0, x_75); +x_76 = lean_box(0); +lean_ctor_set(x_21, 0, x_76); return x_21; } else { -lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; uint8_t x_81; -x_76 = l_Lean_Syntax_getArg(x_36, x_28); +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; uint8_t x_82; +x_77 = l_Lean_Syntax_getArg(x_36, x_28); lean_dec(x_36); -x_77 = l_Lean_Syntax_getArgs(x_76); -lean_dec(x_76); -x_78 = l_Lean_Syntax_getArgs(x_34); +x_78 = l_Lean_Syntax_getArgs(x_77); +lean_dec(x_77); +x_79 = l_Lean_Syntax_getArgs(x_34); lean_dec(x_34); -x_79 = lean_array_get_size(x_78); -x_80 = lean_array_get_size(x_77); -x_81 = lean_nat_dec_eq(x_79, x_80); +x_80 = lean_array_get_size(x_79); +x_81 = lean_array_get_size(x_78); +x_82 = lean_nat_dec_eq(x_80, x_81); +lean_dec(x_81); lean_dec(x_80); -lean_dec(x_79); -if (x_81 == 0) +if (x_82 == 0) { -lean_object* x_82; +lean_object* x_83; +lean_dec(x_79); lean_dec(x_78); -lean_dec(x_77); -lean_dec(x_72); +lean_dec(x_73); lean_dec(x_23); lean_dec(x_7); lean_dec(x_6); @@ -14120,85 +14451,86 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_82 = lean_box(0); -lean_ctor_set(x_21, 0, x_82); +x_83 = lean_box(0); +lean_ctor_set(x_21, 0, x_83); return x_21; } else { -uint8_t x_83; -x_83 = l_Array_isEqvAux___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__7(x_23, lean_box(0), x_78, x_77, x_33); -lean_dec(x_77); +lean_object* x_84; uint8_t x_85; +x_84 = l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__7; +x_85 = l_Array_isEqvAux___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__7(x_23, x_84, lean_box(0), x_79, x_78, x_33); lean_dec(x_78); +lean_dec(x_79); lean_dec(x_23); -if (x_83 == 0) +if (x_85 == 0) { -lean_object* x_84; -lean_dec(x_72); +lean_object* x_86; +lean_dec(x_73); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_84 = lean_box(0); -lean_ctor_set(x_21, 0, x_84); +x_86 = lean_box(0); +lean_ctor_set(x_21, 0, x_86); return x_21; } else { -lean_object* x_85; uint8_t x_86; lean_object* x_87; +lean_object* x_87; uint8_t x_88; lean_object* x_89; lean_free_object(x_21); -x_85 = l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__5; -x_86 = 0; -x_87 = l_Lean_Elab_Term_resolveId_x3f(x_72, x_85, x_86, x_2, x_3, x_4, x_5, x_6, x_7, x_24); -if (lean_obj_tag(x_87) == 0) +x_87 = l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__8; +x_88 = 0; +x_89 = l_Lean_Elab_Term_resolveId_x3f(x_73, x_87, x_88, x_2, x_3, x_4, x_5, x_6, x_7, x_24); +if (lean_obj_tag(x_89) == 0) { -uint8_t x_88; -x_88 = !lean_is_exclusive(x_87); -if (x_88 == 0) +uint8_t x_90; +x_90 = !lean_is_exclusive(x_89); +if (x_90 == 0) { -return x_87; +return x_89; } else { -lean_object* x_89; lean_object* x_90; lean_object* x_91; -x_89 = lean_ctor_get(x_87, 0); -x_90 = lean_ctor_get(x_87, 1); -lean_inc(x_90); -lean_inc(x_89); -lean_dec(x_87); -x_91 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_91, 0, x_89); -lean_ctor_set(x_91, 1, x_90); -return x_91; +lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_91 = lean_ctor_get(x_89, 0); +x_92 = lean_ctor_get(x_89, 1); +lean_inc(x_92); +lean_inc(x_91); +lean_dec(x_89); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +return x_93; } } else { -uint8_t x_92; -x_92 = !lean_is_exclusive(x_87); -if (x_92 == 0) +uint8_t x_94; +x_94 = !lean_is_exclusive(x_89); +if (x_94 == 0) { -lean_object* x_93; lean_object* x_94; -x_93 = lean_ctor_get(x_87, 0); -lean_dec(x_93); -x_94 = lean_box(0); -lean_ctor_set_tag(x_87, 0); -lean_ctor_set(x_87, 0, x_94); -return x_87; +lean_object* x_95; lean_object* x_96; +x_95 = lean_ctor_get(x_89, 0); +lean_dec(x_95); +x_96 = lean_box(0); +lean_ctor_set_tag(x_89, 0); +lean_ctor_set(x_89, 0, x_96); +return x_89; } else { -lean_object* x_95; lean_object* x_96; lean_object* x_97; -x_95 = lean_ctor_get(x_87, 1); -lean_inc(x_95); -lean_dec(x_87); -x_96 = lean_box(0); -x_97 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_97, 0, x_96); -lean_ctor_set(x_97, 1, x_95); -return x_97; +lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_97 = lean_ctor_get(x_89, 1); +lean_inc(x_97); +lean_dec(x_89); +x_98 = lean_box(0); +x_99 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_99, 0, x_98); +lean_ctor_set(x_99, 1, x_97); +return x_99; } } } @@ -14210,229 +14542,230 @@ return x_97; } else { -lean_object* x_98; lean_object* x_99; lean_object* x_100; uint8_t x_101; -x_98 = lean_ctor_get(x_21, 0); -x_99 = lean_ctor_get(x_21, 1); -lean_inc(x_99); -lean_inc(x_98); +lean_object* x_100; lean_object* x_101; lean_object* x_102; uint8_t x_103; +x_100 = lean_ctor_get(x_21, 0); +x_101 = lean_ctor_get(x_21, 1); +lean_inc(x_101); +lean_inc(x_100); lean_dec(x_21); -x_100 = l_Lean_Elab_Term_expandDbgTrace___closed__17; -lean_inc(x_98); -x_101 = l_Lean_Syntax_isOfKind(x_98, x_100); -if (x_101 == 0) +x_102 = l_Lean_Elab_Term_expandDbgTrace___closed__20; +lean_inc(x_100); +x_103 = l_Lean_Syntax_isOfKind(x_100, x_102); +if (x_103 == 0) { -lean_object* x_102; lean_object* x_103; -lean_dec(x_98); +lean_object* x_104; lean_object* x_105; +lean_dec(x_100); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_102 = lean_box(0); -x_103 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_103, 0, x_102); -lean_ctor_set(x_103, 1, x_99); -return x_103; +x_104 = lean_box(0); +x_105 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_105, 0, x_104); +lean_ctor_set(x_105, 1, x_101); +return x_105; } else { -lean_object* x_104; lean_object* x_105; lean_object* x_106; uint8_t x_107; -x_104 = lean_unsigned_to_nat(1u); -x_105 = l_Lean_Syntax_getArg(x_98, x_104); -x_106 = l_Lean_Elab_Term_expandDbgTrace___closed__19; -lean_inc(x_105); -x_107 = l_Lean_Syntax_isOfKind(x_105, x_106); -if (x_107 == 0) +lean_object* x_106; lean_object* x_107; lean_object* x_108; uint8_t x_109; +x_106 = lean_unsigned_to_nat(1u); +x_107 = l_Lean_Syntax_getArg(x_100, x_106); +x_108 = l_Lean_Elab_Term_expandDbgTrace___closed__22; +lean_inc(x_107); +x_109 = l_Lean_Syntax_isOfKind(x_107, x_108); +if (x_109 == 0) { -lean_object* x_108; lean_object* x_109; -lean_dec(x_105); -lean_dec(x_98); +lean_object* x_110; lean_object* x_111; +lean_dec(x_107); +lean_dec(x_100); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_108 = lean_box(0); -x_109 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_109, 0, x_108); -lean_ctor_set(x_109, 1, x_99); -return x_109; +x_110 = lean_box(0); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_110); +lean_ctor_set(x_111, 1, x_101); +return x_111; } else { -lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; uint8_t x_115; -x_110 = lean_unsigned_to_nat(0u); -x_111 = l_Lean_Syntax_getArg(x_105, x_110); -x_112 = lean_unsigned_to_nat(2u); -x_113 = l_Lean_Syntax_getArg(x_105, x_112); -lean_dec(x_105); -x_114 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__5; -lean_inc(x_113); -x_115 = l_Lean_Syntax_isOfKind(x_113, x_114); -if (x_115 == 0) -{ -lean_object* x_116; uint8_t x_117; -x_116 = l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__2; -lean_inc(x_113); -x_117 = l_Lean_Syntax_isOfKind(x_113, x_116); +lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; uint8_t x_117; +x_112 = lean_unsigned_to_nat(0u); +x_113 = l_Lean_Syntax_getArg(x_107, x_112); +x_114 = lean_unsigned_to_nat(2u); +x_115 = l_Lean_Syntax_getArg(x_107, x_114); +lean_dec(x_107); +x_116 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__5; +lean_inc(x_115); +x_117 = l_Lean_Syntax_isOfKind(x_115, x_116); if (x_117 == 0) { -lean_object* x_118; lean_object* x_119; +lean_object* x_118; uint8_t x_119; +x_118 = l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__2; +lean_inc(x_115); +x_119 = l_Lean_Syntax_isOfKind(x_115, x_118); +if (x_119 == 0) +{ +lean_object* x_120; lean_object* x_121; +lean_dec(x_115); lean_dec(x_113); -lean_dec(x_111); -lean_dec(x_98); +lean_dec(x_100); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_118 = lean_box(0); -x_119 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_119, 0, x_118); -lean_ctor_set(x_119, 1, x_99); -return x_119; +x_120 = lean_box(0); +x_121 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_121, 0, x_120); +lean_ctor_set(x_121, 1, x_101); +return x_121; } else { -lean_object* x_120; lean_object* x_121; uint8_t x_122; -x_120 = l_Lean_Syntax_getArg(x_113, x_104); -x_121 = l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__4; -lean_inc(x_120); -x_122 = l_Lean_Syntax_isOfKind(x_120, x_121); -if (x_122 == 0) +lean_object* x_122; lean_object* x_123; uint8_t x_124; +x_122 = l_Lean_Syntax_getArg(x_115, x_106); +x_123 = l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__4; +lean_inc(x_122); +x_124 = l_Lean_Syntax_isOfKind(x_122, x_123); +if (x_124 == 0) { -lean_object* x_123; lean_object* x_124; -lean_dec(x_120); +lean_object* x_125; lean_object* x_126; +lean_dec(x_122); +lean_dec(x_115); lean_dec(x_113); -lean_dec(x_111); -lean_dec(x_98); +lean_dec(x_100); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_123 = lean_box(0); -x_124 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_124, 0, x_123); -lean_ctor_set(x_124, 1, x_99); -return x_124; +x_125 = lean_box(0); +x_126 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_126, 0, x_125); +lean_ctor_set(x_126, 1, x_101); +return x_126; } else { -lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; uint8_t x_134; -x_125 = l_Lean_Syntax_getArg(x_113, x_112); -x_126 = lean_unsigned_to_nat(3u); -x_127 = l_Lean_Syntax_getArg(x_113, x_126); +lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; uint8_t x_136; +x_127 = l_Lean_Syntax_getArg(x_115, x_114); +x_128 = lean_unsigned_to_nat(3u); +x_129 = l_Lean_Syntax_getArg(x_115, x_128); +lean_dec(x_115); +x_130 = l_Lean_Syntax_getArgs(x_113); lean_dec(x_113); -x_128 = l_Lean_Syntax_getArgs(x_111); -lean_dec(x_111); -x_129 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; -x_130 = lean_array_push(x_129, x_125); -x_131 = lean_array_push(x_130, x_127); -x_132 = lean_array_get_size(x_128); -x_133 = lean_array_get_size(x_131); -x_134 = lean_nat_dec_eq(x_132, x_133); +x_131 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; +x_132 = lean_array_push(x_131, x_127); +x_133 = lean_array_push(x_132, x_129); +x_134 = lean_array_get_size(x_130); +x_135 = lean_array_get_size(x_133); +x_136 = lean_nat_dec_eq(x_134, x_135); +lean_dec(x_135); +lean_dec(x_134); +if (x_136 == 0) +{ +lean_object* x_137; lean_object* x_138; lean_dec(x_133); -lean_dec(x_132); -if (x_134 == 0) -{ -lean_object* x_135; lean_object* x_136; -lean_dec(x_131); -lean_dec(x_128); -lean_dec(x_120); -lean_dec(x_98); +lean_dec(x_130); +lean_dec(x_122); +lean_dec(x_100); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_135 = lean_box(0); -x_136 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_136, 0, x_135); -lean_ctor_set(x_136, 1, x_99); -return x_136; +x_137 = lean_box(0); +x_138 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_138, 0, x_137); +lean_ctor_set(x_138, 1, x_101); +return x_138; } else { -uint8_t x_137; -x_137 = l_Array_isEqvAux___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__6(x_98, lean_box(0), x_128, x_131, x_110); -lean_dec(x_131); -lean_dec(x_128); -lean_dec(x_98); -if (x_137 == 0) +lean_object* x_139; uint8_t x_140; +x_139 = l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__7; +x_140 = l_Array_isEqvAux___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__6(x_100, x_139, lean_box(0), x_130, x_133, x_112); +lean_dec(x_133); +lean_dec(x_130); +lean_dec(x_100); +if (x_140 == 0) { -lean_object* x_138; lean_object* x_139; -lean_dec(x_120); +lean_object* x_141; lean_object* x_142; +lean_dec(x_122); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_138 = lean_box(0); -x_139 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_139, 0, x_138); -lean_ctor_set(x_139, 1, x_99); -return x_139; +x_141 = lean_box(0); +x_142 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_142, 0, x_141); +lean_ctor_set(x_142, 1, x_101); +return x_142; } else { -lean_object* x_140; uint8_t x_141; lean_object* x_142; -x_140 = l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__5; -x_141 = 0; -x_142 = l_Lean_Elab_Term_resolveId_x3f(x_120, x_140, x_141, x_2, x_3, x_4, x_5, x_6, x_7, x_99); -if (lean_obj_tag(x_142) == 0) +lean_object* x_143; uint8_t x_144; lean_object* x_145; +x_143 = l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__8; +x_144 = 0; +x_145 = l_Lean_Elab_Term_resolveId_x3f(x_122, x_143, x_144, x_2, x_3, x_4, x_5, x_6, x_7, x_101); +if (lean_obj_tag(x_145) == 0) { -lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; -x_143 = lean_ctor_get(x_142, 0); -lean_inc(x_143); -x_144 = lean_ctor_get(x_142, 1); -lean_inc(x_144); -if (lean_is_exclusive(x_142)) { - lean_ctor_release(x_142, 0); - lean_ctor_release(x_142, 1); - x_145 = x_142; +lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; +x_146 = lean_ctor_get(x_145, 0); +lean_inc(x_146); +x_147 = lean_ctor_get(x_145, 1); +lean_inc(x_147); +if (lean_is_exclusive(x_145)) { + lean_ctor_release(x_145, 0); + lean_ctor_release(x_145, 1); + x_148 = x_145; } else { - lean_dec_ref(x_142); - x_145 = lean_box(0); + lean_dec_ref(x_145); + x_148 = lean_box(0); } -if (lean_is_scalar(x_145)) { - x_146 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_148)) { + x_149 = lean_alloc_ctor(0, 2, 0); } else { - x_146 = x_145; + x_149 = x_148; } -lean_ctor_set(x_146, 0, x_143); -lean_ctor_set(x_146, 1, x_144); -return x_146; +lean_ctor_set(x_149, 0, x_146); +lean_ctor_set(x_149, 1, x_147); +return x_149; } else { -lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; -x_147 = lean_ctor_get(x_142, 1); -lean_inc(x_147); -if (lean_is_exclusive(x_142)) { - lean_ctor_release(x_142, 0); - lean_ctor_release(x_142, 1); - x_148 = x_142; +lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; +x_150 = lean_ctor_get(x_145, 1); +lean_inc(x_150); +if (lean_is_exclusive(x_145)) { + lean_ctor_release(x_145, 0); + lean_ctor_release(x_145, 1); + x_151 = x_145; } else { - lean_dec_ref(x_142); - x_148 = lean_box(0); + lean_dec_ref(x_145); + x_151 = lean_box(0); } -x_149 = lean_box(0); -if (lean_is_scalar(x_148)) { - x_150 = lean_alloc_ctor(0, 2, 0); +x_152 = lean_box(0); +if (lean_is_scalar(x_151)) { + x_153 = lean_alloc_ctor(0, 2, 0); } else { - x_150 = x_148; - lean_ctor_set_tag(x_150, 0); + x_153 = x_151; + lean_ctor_set_tag(x_153, 0); } -lean_ctor_set(x_150, 0, x_149); -lean_ctor_set(x_150, 1, x_147); -return x_150; +lean_ctor_set(x_153, 0, x_152); +lean_ctor_set(x_153, 1, x_150); +return x_153; } } } @@ -14441,74 +14774,51 @@ return x_150; } else { -lean_object* x_151; lean_object* x_152; uint8_t x_153; -x_151 = l_Lean_Syntax_getArg(x_113, x_110); -x_152 = l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__4; -lean_inc(x_151); -x_153 = l_Lean_Syntax_isOfKind(x_151, x_152); -if (x_153 == 0) -{ -lean_object* x_154; lean_object* x_155; -lean_dec(x_151); -lean_dec(x_113); -lean_dec(x_111); -lean_dec(x_98); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_154 = lean_box(0); -x_155 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_155, 0, x_154); -lean_ctor_set(x_155, 1, x_99); -return x_155; -} -else +lean_object* x_154; lean_object* x_155; uint8_t x_156; +x_154 = l_Lean_Syntax_getArg(x_115, x_112); +x_155 = l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__4; +lean_inc(x_154); +x_156 = l_Lean_Syntax_isOfKind(x_154, x_155); +if (x_156 == 0) { -lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; uint8_t x_161; -x_156 = l_Lean_Syntax_getArg(x_113, x_104); +lean_object* x_157; lean_object* x_158; +lean_dec(x_154); +lean_dec(x_115); lean_dec(x_113); -x_157 = l_Lean_Syntax_getArgs(x_156); -lean_dec(x_156); -x_158 = l_Lean_Syntax_getArgs(x_111); -lean_dec(x_111); -x_159 = lean_array_get_size(x_158); -x_160 = lean_array_get_size(x_157); -x_161 = lean_nat_dec_eq(x_159, x_160); -lean_dec(x_160); -lean_dec(x_159); -if (x_161 == 0) -{ -lean_object* x_162; lean_object* x_163; -lean_dec(x_158); -lean_dec(x_157); -lean_dec(x_151); -lean_dec(x_98); +lean_dec(x_100); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_162 = lean_box(0); -x_163 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_163, 0, x_162); -lean_ctor_set(x_163, 1, x_99); -return x_163; +x_157 = lean_box(0); +x_158 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_158, 0, x_157); +lean_ctor_set(x_158, 1, x_101); +return x_158; } else { -uint8_t x_164; -x_164 = l_Array_isEqvAux___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__7(x_98, lean_box(0), x_158, x_157, x_110); -lean_dec(x_157); -lean_dec(x_158); -lean_dec(x_98); +lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; uint8_t x_164; +x_159 = l_Lean_Syntax_getArg(x_115, x_106); +lean_dec(x_115); +x_160 = l_Lean_Syntax_getArgs(x_159); +lean_dec(x_159); +x_161 = l_Lean_Syntax_getArgs(x_113); +lean_dec(x_113); +x_162 = lean_array_get_size(x_161); +x_163 = lean_array_get_size(x_160); +x_164 = lean_nat_dec_eq(x_162, x_163); +lean_dec(x_163); +lean_dec(x_162); if (x_164 == 0) { lean_object* x_165; lean_object* x_166; -lean_dec(x_151); +lean_dec(x_161); +lean_dec(x_160); +lean_dec(x_154); +lean_dec(x_100); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); @@ -14518,63 +14828,87 @@ lean_dec(x_2); x_165 = lean_box(0); x_166 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_166, 0, x_165); -lean_ctor_set(x_166, 1, x_99); +lean_ctor_set(x_166, 1, x_101); return x_166; } else { -lean_object* x_167; uint8_t x_168; lean_object* x_169; -x_167 = l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__5; -x_168 = 0; -x_169 = l_Lean_Elab_Term_resolveId_x3f(x_151, x_167, x_168, x_2, x_3, x_4, x_5, x_6, x_7, x_99); -if (lean_obj_tag(x_169) == 0) +lean_object* x_167; uint8_t x_168; +x_167 = l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__7; +x_168 = l_Array_isEqvAux___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__7(x_100, x_167, lean_box(0), x_161, x_160, x_112); +lean_dec(x_160); +lean_dec(x_161); +lean_dec(x_100); +if (x_168 == 0) { -lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; -x_170 = lean_ctor_get(x_169, 0); -lean_inc(x_170); -x_171 = lean_ctor_get(x_169, 1); -lean_inc(x_171); -if (lean_is_exclusive(x_169)) { - lean_ctor_release(x_169, 0); - lean_ctor_release(x_169, 1); - x_172 = x_169; -} else { - lean_dec_ref(x_169); - x_172 = lean_box(0); -} -if (lean_is_scalar(x_172)) { - x_173 = lean_alloc_ctor(0, 2, 0); -} else { - x_173 = x_172; -} -lean_ctor_set(x_173, 0, x_170); -lean_ctor_set(x_173, 1, x_171); -return x_173; +lean_object* x_169; lean_object* x_170; +lean_dec(x_154); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_169 = lean_box(0); +x_170 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_170, 0, x_169); +lean_ctor_set(x_170, 1, x_101); +return x_170; } else { +lean_object* x_171; uint8_t x_172; lean_object* x_173; +x_171 = l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__8; +x_172 = 0; +x_173 = l_Lean_Elab_Term_resolveId_x3f(x_154, x_171, x_172, x_2, x_3, x_4, x_5, x_6, x_7, x_101); +if (lean_obj_tag(x_173) == 0) +{ lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; -x_174 = lean_ctor_get(x_169, 1); +x_174 = lean_ctor_get(x_173, 0); lean_inc(x_174); -if (lean_is_exclusive(x_169)) { - lean_ctor_release(x_169, 0); - lean_ctor_release(x_169, 1); - x_175 = x_169; +x_175 = lean_ctor_get(x_173, 1); +lean_inc(x_175); +if (lean_is_exclusive(x_173)) { + lean_ctor_release(x_173, 0); + lean_ctor_release(x_173, 1); + x_176 = x_173; } else { - lean_dec_ref(x_169); - x_175 = lean_box(0); + lean_dec_ref(x_173); + x_176 = lean_box(0); } -x_176 = lean_box(0); -if (lean_is_scalar(x_175)) { +if (lean_is_scalar(x_176)) { x_177 = lean_alloc_ctor(0, 2, 0); } else { - x_177 = x_175; - lean_ctor_set_tag(x_177, 0); + x_177 = x_176; } -lean_ctor_set(x_177, 0, x_176); -lean_ctor_set(x_177, 1, x_174); +lean_ctor_set(x_177, 0, x_174); +lean_ctor_set(x_177, 1, x_175); return x_177; } +else +{ +lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; +x_178 = lean_ctor_get(x_173, 1); +lean_inc(x_178); +if (lean_is_exclusive(x_173)) { + lean_ctor_release(x_173, 0); + lean_ctor_release(x_173, 1); + x_179 = x_173; +} else { + lean_dec_ref(x_173); + x_179 = lean_box(0); +} +x_180 = lean_box(0); +if (lean_is_scalar(x_179)) { + x_181 = lean_alloc_ctor(0, 2, 0); +} else { + x_181 = x_179; + lean_ctor_set_tag(x_181, 0); +} +lean_ctor_set(x_181, 0, x_180); +lean_ctor_set(x_181, 1, x_178); +return x_181; +} } } } @@ -14585,60 +14919,60 @@ return x_177; } else { -uint8_t x_178; +uint8_t x_182; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_178 = !lean_is_exclusive(x_21); -if (x_178 == 0) +x_182 = !lean_is_exclusive(x_21); +if (x_182 == 0) { return x_21; } else { -lean_object* x_179; lean_object* x_180; lean_object* x_181; -x_179 = lean_ctor_get(x_21, 0); -x_180 = lean_ctor_get(x_21, 1); -lean_inc(x_180); -lean_inc(x_179); +lean_object* x_183; lean_object* x_184; lean_object* x_185; +x_183 = lean_ctor_get(x_21, 0); +x_184 = lean_ctor_get(x_21, 1); +lean_inc(x_184); +lean_inc(x_183); lean_dec(x_21); -x_181 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_181, 0, x_179); -lean_ctor_set(x_181, 1, x_180); -return x_181; +x_185 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_185, 0, x_183); +lean_ctor_set(x_185, 1, x_184); +return x_185; } } } } else { -uint8_t x_182; +uint8_t x_186; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_182 = !lean_is_exclusive(x_10); -if (x_182 == 0) +x_186 = !lean_is_exclusive(x_10); +if (x_186 == 0) { return x_10; } else { -lean_object* x_183; lean_object* x_184; lean_object* x_185; -x_183 = lean_ctor_get(x_10, 0); -x_184 = lean_ctor_get(x_10, 1); -lean_inc(x_184); -lean_inc(x_183); +lean_object* x_187; lean_object* x_188; lean_object* x_189; +x_187 = lean_ctor_get(x_10, 0); +x_188 = lean_ctor_get(x_10, 1); +lean_inc(x_188); +lean_inc(x_187); lean_dec(x_10); -x_185 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_185, 0, x_183); -lean_ctor_set(x_185, 1, x_184); -return x_185; +x_189 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_189, 0, x_187); +lean_ctor_set(x_189, 1, x_188); +return x_189; } } } @@ -14721,28 +15055,30 @@ lean_dec(x_5); return x_7; } } -LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__6___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: { -uint8_t x_6; lean_object* x_7; -x_6 = l_Array_isEqvAux___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__6(x_1, x_2, x_3, x_4, x_5); +uint8_t x_7; lean_object* x_8; +x_7 = l_Array_isEqvAux___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__6(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -x_7 = lean_box(x_6); -return x_7; +x_8 = lean_box(x_7); +return x_8; } } -LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__7___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: { -uint8_t x_6; lean_object* x_7; -x_6 = l_Array_isEqvAux___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__7(x_1, x_2, x_3, x_4, x_5); +uint8_t x_7; lean_object* x_8; +x_7 = l_Array_isEqvAux___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__7(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -x_7 = lean_box(x_6); -return x_7; +x_8 = lean_box(x_7); +return x_8; } } static lean_object* _init_l_Lean_Elab_Term_expandParen___closed__1() { @@ -15016,785 +15352,784 @@ return x_56; } else { -lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; +lean_object* x_57; lean_object* x_58; lean_object* x_59; uint8_t x_60; x_57 = lean_unsigned_to_nat(1u); x_58 = l_Lean_Syntax_getArg(x_1, x_57); -x_59 = l_Lean_nullKind; -x_60 = lean_unsigned_to_nat(0u); +x_59 = lean_unsigned_to_nat(0u); lean_inc(x_58); -x_61 = l_Lean_Syntax_isNodeOf(x_58, x_59, x_60); -if (x_61 == 0) +x_60 = l_Lean_Syntax_matchesNull(x_58, x_59); +if (x_60 == 0) { -lean_object* x_62; uint8_t x_63; -x_62 = lean_unsigned_to_nat(2u); +lean_object* x_61; uint8_t x_62; +x_61 = lean_unsigned_to_nat(2u); lean_inc(x_58); -x_63 = l_Lean_Syntax_isNodeOf(x_58, x_59, x_62); -if (x_63 == 0) +x_62 = l_Lean_Syntax_matchesNull(x_58, x_61); +if (x_62 == 0) { -lean_object* x_64; uint8_t x_65; -x_64 = l_Lean_Syntax_getArg(x_58, x_60); -x_65 = l_Lean_Syntax_isMissing(x_64); -if (x_65 == 0) +lean_object* x_63; uint8_t x_64; +x_63 = l_Lean_Syntax_getArg(x_58, x_59); +x_64 = l_Lean_Syntax_isMissing(x_63); +if (x_64 == 0) { -lean_object* x_66; uint8_t x_67; -x_66 = l_Lean_Syntax_getArg(x_58, x_57); +lean_object* x_65; uint8_t x_66; +x_65 = l_Lean_Syntax_getArg(x_58, x_57); lean_dec(x_58); -x_67 = l_Lean_Syntax_isMissing(x_66); -lean_dec(x_66); -if (x_67 == 0) +x_66 = l_Lean_Syntax_isMissing(x_65); +lean_dec(x_65); +if (x_66 == 0) { -lean_object* x_68; lean_object* x_69; lean_object* x_70; -lean_dec(x_64); +lean_object* x_67; lean_object* x_68; lean_object* x_69; +lean_dec(x_63); lean_dec(x_2); -x_68 = l_Lean_Elab_Term_expandParen___closed__1; -x_69 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_69, 0, x_1); -lean_ctor_set(x_69, 1, x_68); -x_70 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_3); -return x_70; +x_67 = l_Lean_Elab_Term_expandParen___closed__1; +x_68 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_68, 0, x_1); +lean_ctor_set(x_68, 1, x_67); +x_69 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_69, 0, x_68); +lean_ctor_set(x_69, 1, x_3); +return x_69; } else { -lean_object* x_71; uint8_t x_72; +lean_object* x_70; uint8_t x_71; lean_dec(x_1); -x_71 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_72 = !lean_is_exclusive(x_71); -if (x_72 == 0) +x_70 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_71 = !lean_is_exclusive(x_70); +if (x_71 == 0) { -lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_73 = lean_ctor_get(x_71, 0); -x_74 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; -lean_inc(x_73); -x_75 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_75, 0, x_73); -lean_ctor_set(x_75, 1, x_74); -x_76 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; -x_77 = lean_array_push(x_76, x_64); -x_78 = l_Lean_Elab_Term_expandShow___closed__23; -x_79 = lean_array_push(x_77, x_78); -x_80 = lean_box(2); -x_81 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; -x_82 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_82, 0, x_80); -lean_ctor_set(x_82, 1, x_81); -lean_ctor_set(x_82, 2, x_79); -x_83 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__23; -x_84 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_84, 0, x_73); -lean_ctor_set(x_84, 1, x_83); -x_85 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; -x_86 = lean_array_push(x_85, x_75); -x_87 = lean_array_push(x_86, x_82); -x_88 = lean_array_push(x_87, x_84); -x_89 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_89, 0, x_80); -lean_ctor_set(x_89, 1, x_4); -lean_ctor_set(x_89, 2, x_88); -lean_ctor_set(x_71, 0, x_89); -return 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_72 = lean_ctor_get(x_70, 0); +x_73 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; +lean_inc(x_72); +x_74 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_74, 0, x_72); +lean_ctor_set(x_74, 1, x_73); +x_75 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; +x_76 = lean_array_push(x_75, x_63); +x_77 = l_Lean_Elab_Term_expandShow___closed__23; +x_78 = lean_array_push(x_76, x_77); +x_79 = lean_box(2); +x_80 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; +x_81 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_81, 0, x_79); +lean_ctor_set(x_81, 1, x_80); +lean_ctor_set(x_81, 2, x_78); +x_82 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__23; +x_83 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_83, 0, x_72); +lean_ctor_set(x_83, 1, x_82); +x_84 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; +x_85 = lean_array_push(x_84, x_74); +x_86 = lean_array_push(x_85, x_81); +x_87 = lean_array_push(x_86, x_83); +x_88 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_88, 0, x_79); +lean_ctor_set(x_88, 1, x_4); +lean_ctor_set(x_88, 2, x_87); +lean_ctor_set(x_70, 0, x_88); +return x_70; } else { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; -x_90 = lean_ctor_get(x_71, 0); -x_91 = lean_ctor_get(x_71, 1); -lean_inc(x_91); +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_89 = lean_ctor_get(x_70, 0); +x_90 = lean_ctor_get(x_70, 1); lean_inc(x_90); -lean_dec(x_71); -x_92 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; -lean_inc(x_90); -x_93 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_93, 0, x_90); -lean_ctor_set(x_93, 1, x_92); -x_94 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; -x_95 = lean_array_push(x_94, x_64); -x_96 = l_Lean_Elab_Term_expandShow___closed__23; -x_97 = lean_array_push(x_95, x_96); -x_98 = lean_box(2); -x_99 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; -x_100 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_100, 0, x_98); -lean_ctor_set(x_100, 1, x_99); -lean_ctor_set(x_100, 2, x_97); -x_101 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__23; -x_102 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_102, 0, x_90); -lean_ctor_set(x_102, 1, x_101); -x_103 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; -x_104 = lean_array_push(x_103, x_93); -x_105 = lean_array_push(x_104, x_100); -x_106 = lean_array_push(x_105, x_102); -x_107 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_107, 0, x_98); -lean_ctor_set(x_107, 1, x_4); -lean_ctor_set(x_107, 2, x_106); -x_108 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_108, 0, x_107); -lean_ctor_set(x_108, 1, x_91); -return x_108; +lean_inc(x_89); +lean_dec(x_70); +x_91 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; +lean_inc(x_89); +x_92 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_92, 0, x_89); +lean_ctor_set(x_92, 1, x_91); +x_93 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; +x_94 = lean_array_push(x_93, x_63); +x_95 = l_Lean_Elab_Term_expandShow___closed__23; +x_96 = lean_array_push(x_94, x_95); +x_97 = lean_box(2); +x_98 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; +x_99 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_99, 0, x_97); +lean_ctor_set(x_99, 1, x_98); +lean_ctor_set(x_99, 2, x_96); +x_100 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__23; +x_101 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_101, 0, x_89); +lean_ctor_set(x_101, 1, x_100); +x_102 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; +x_103 = lean_array_push(x_102, x_92); +x_104 = lean_array_push(x_103, x_99); +x_105 = lean_array_push(x_104, x_101); +x_106 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_106, 0, x_97); +lean_ctor_set(x_106, 1, x_4); +lean_ctor_set(x_106, 2, x_105); +x_107 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_107, 0, x_106); +lean_ctor_set(x_107, 1, x_90); +return x_107; } } } else { -lean_object* x_109; lean_object* x_110; lean_object* x_111; -lean_dec(x_64); +lean_object* x_108; lean_object* x_109; lean_object* x_110; +lean_dec(x_63); lean_dec(x_58); lean_dec(x_2); -x_109 = l_Lean_Elab_Term_expandParen___closed__1; -x_110 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_110, 0, x_1); -lean_ctor_set(x_110, 1, x_109); -x_111 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_111, 0, x_110); -lean_ctor_set(x_111, 1, x_3); -return x_111; +x_108 = l_Lean_Elab_Term_expandParen___closed__1; +x_109 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_109, 0, x_1); +lean_ctor_set(x_109, 1, x_108); +x_110 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_110, 0, x_109); +lean_ctor_set(x_110, 1, x_3); +return x_110; } } else { -lean_object* x_112; lean_object* x_113; uint8_t x_114; -x_112 = l_Lean_Syntax_getArg(x_58, x_60); -x_113 = l_Lean_Syntax_getArg(x_58, x_57); -lean_inc(x_113); -x_114 = l_Lean_Syntax_isNodeOf(x_113, x_59, x_57); -if (x_114 == 0) +lean_object* x_111; lean_object* x_112; uint8_t x_113; +x_111 = l_Lean_Syntax_getArg(x_58, x_59); +x_112 = l_Lean_Syntax_getArg(x_58, x_57); +lean_inc(x_112); +x_113 = l_Lean_Syntax_matchesNull(x_112, x_57); +if (x_113 == 0) { -uint8_t x_115; -x_115 = l_Lean_Syntax_isNodeOf(x_113, x_59, x_60); -if (x_115 == 0) +uint8_t x_114; +x_114 = l_Lean_Syntax_matchesNull(x_112, x_59); +if (x_114 == 0) { -lean_object* x_116; uint8_t x_117; -lean_dec(x_112); -x_116 = l_Lean_Syntax_getArg(x_58, x_60); -x_117 = l_Lean_Syntax_isMissing(x_116); -if (x_117 == 0) +lean_object* x_115; uint8_t x_116; +lean_dec(x_111); +x_115 = l_Lean_Syntax_getArg(x_58, x_59); +x_116 = l_Lean_Syntax_isMissing(x_115); +if (x_116 == 0) { -lean_object* x_118; uint8_t x_119; -x_118 = l_Lean_Syntax_getArg(x_58, x_57); +lean_object* x_117; uint8_t x_118; +x_117 = l_Lean_Syntax_getArg(x_58, x_57); lean_dec(x_58); -x_119 = l_Lean_Syntax_isMissing(x_118); -lean_dec(x_118); -if (x_119 == 0) +x_118 = l_Lean_Syntax_isMissing(x_117); +lean_dec(x_117); +if (x_118 == 0) { -lean_object* x_120; lean_object* x_121; lean_object* x_122; -lean_dec(x_116); +lean_object* x_119; lean_object* x_120; lean_object* x_121; +lean_dec(x_115); lean_dec(x_2); -x_120 = l_Lean_Elab_Term_expandParen___closed__1; -x_121 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_121, 0, x_1); -lean_ctor_set(x_121, 1, x_120); -x_122 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_122, 0, x_121); -lean_ctor_set(x_122, 1, x_3); -return x_122; +x_119 = l_Lean_Elab_Term_expandParen___closed__1; +x_120 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_120, 0, x_1); +lean_ctor_set(x_120, 1, x_119); +x_121 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_121, 0, x_120); +lean_ctor_set(x_121, 1, x_3); +return x_121; } else { -lean_object* x_123; uint8_t x_124; +lean_object* x_122; uint8_t x_123; lean_dec(x_1); -x_123 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_124 = !lean_is_exclusive(x_123); -if (x_124 == 0) +x_122 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_123 = !lean_is_exclusive(x_122); +if (x_123 == 0) { -lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; -x_125 = lean_ctor_get(x_123, 0); -x_126 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; -lean_inc(x_125); -x_127 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_127, 0, x_125); -lean_ctor_set(x_127, 1, x_126); -x_128 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; -x_129 = lean_array_push(x_128, x_116); -x_130 = l_Lean_Elab_Term_expandShow___closed__23; -x_131 = lean_array_push(x_129, x_130); -x_132 = lean_box(2); -x_133 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; -x_134 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_134, 0, x_132); -lean_ctor_set(x_134, 1, x_133); -lean_ctor_set(x_134, 2, x_131); -x_135 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__23; -x_136 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_136, 0, x_125); -lean_ctor_set(x_136, 1, x_135); -x_137 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; -x_138 = lean_array_push(x_137, x_127); -x_139 = lean_array_push(x_138, x_134); -x_140 = lean_array_push(x_139, x_136); -x_141 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_141, 0, x_132); -lean_ctor_set(x_141, 1, x_4); -lean_ctor_set(x_141, 2, x_140); -lean_ctor_set(x_123, 0, x_141); -return x_123; +lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; +x_124 = lean_ctor_get(x_122, 0); +x_125 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; +lean_inc(x_124); +x_126 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_126, 0, x_124); +lean_ctor_set(x_126, 1, x_125); +x_127 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; +x_128 = lean_array_push(x_127, x_115); +x_129 = l_Lean_Elab_Term_expandShow___closed__23; +x_130 = lean_array_push(x_128, x_129); +x_131 = lean_box(2); +x_132 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; +x_133 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_133, 0, x_131); +lean_ctor_set(x_133, 1, x_132); +lean_ctor_set(x_133, 2, x_130); +x_134 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__23; +x_135 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_135, 0, x_124); +lean_ctor_set(x_135, 1, x_134); +x_136 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; +x_137 = lean_array_push(x_136, x_126); +x_138 = lean_array_push(x_137, x_133); +x_139 = lean_array_push(x_138, x_135); +x_140 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_140, 0, x_131); +lean_ctor_set(x_140, 1, x_4); +lean_ctor_set(x_140, 2, x_139); +lean_ctor_set(x_122, 0, x_140); +return x_122; } else { -lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; -x_142 = lean_ctor_get(x_123, 0); -x_143 = lean_ctor_get(x_123, 1); -lean_inc(x_143); -lean_inc(x_142); -lean_dec(x_123); -x_144 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; +lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; +x_141 = lean_ctor_get(x_122, 0); +x_142 = lean_ctor_get(x_122, 1); lean_inc(x_142); -x_145 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_145, 0, x_142); -lean_ctor_set(x_145, 1, x_144); -x_146 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; -x_147 = lean_array_push(x_146, x_116); -x_148 = l_Lean_Elab_Term_expandShow___closed__23; -x_149 = lean_array_push(x_147, x_148); -x_150 = lean_box(2); -x_151 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; -x_152 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_152, 0, x_150); -lean_ctor_set(x_152, 1, x_151); -lean_ctor_set(x_152, 2, x_149); -x_153 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__23; -x_154 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_154, 0, x_142); -lean_ctor_set(x_154, 1, x_153); -x_155 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; -x_156 = lean_array_push(x_155, x_145); -x_157 = lean_array_push(x_156, x_152); -x_158 = lean_array_push(x_157, x_154); -x_159 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_159, 0, x_150); -lean_ctor_set(x_159, 1, x_4); -lean_ctor_set(x_159, 2, x_158); -x_160 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_160, 0, x_159); -lean_ctor_set(x_160, 1, x_143); -return x_160; +lean_inc(x_141); +lean_dec(x_122); +x_143 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; +lean_inc(x_141); +x_144 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_144, 0, x_141); +lean_ctor_set(x_144, 1, x_143); +x_145 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; +x_146 = lean_array_push(x_145, x_115); +x_147 = l_Lean_Elab_Term_expandShow___closed__23; +x_148 = lean_array_push(x_146, x_147); +x_149 = lean_box(2); +x_150 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; +x_151 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_151, 0, x_149); +lean_ctor_set(x_151, 1, x_150); +lean_ctor_set(x_151, 2, x_148); +x_152 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__23; +x_153 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_153, 0, x_141); +lean_ctor_set(x_153, 1, x_152); +x_154 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; +x_155 = lean_array_push(x_154, x_144); +x_156 = lean_array_push(x_155, x_151); +x_157 = lean_array_push(x_156, x_153); +x_158 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_158, 0, x_149); +lean_ctor_set(x_158, 1, x_4); +lean_ctor_set(x_158, 2, x_157); +x_159 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_159, 0, x_158); +lean_ctor_set(x_159, 1, x_142); +return x_159; } } } else { -lean_object* x_161; lean_object* x_162; lean_object* x_163; -lean_dec(x_116); +lean_object* x_160; lean_object* x_161; lean_object* x_162; +lean_dec(x_115); lean_dec(x_58); lean_dec(x_2); -x_161 = l_Lean_Elab_Term_expandParen___closed__1; -x_162 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_162, 0, x_1); -lean_ctor_set(x_162, 1, x_161); -x_163 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_163, 0, x_162); -lean_ctor_set(x_163, 1, x_3); -return x_163; +x_160 = l_Lean_Elab_Term_expandParen___closed__1; +x_161 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_161, 0, x_1); +lean_ctor_set(x_161, 1, x_160); +x_162 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_162, 0, x_161); +lean_ctor_set(x_162, 1, x_3); +return x_162; } } else { -lean_object* x_164; lean_object* x_165; +lean_object* x_163; lean_object* x_164; lean_dec(x_58); lean_dec(x_1); -lean_inc(x_112); -x_164 = l_Lean_Elab_Term_expandCDot_x3f(x_112, x_2, x_3); -x_165 = lean_ctor_get(x_164, 0); -lean_inc(x_165); -if (lean_obj_tag(x_165) == 0) +lean_inc(x_111); +x_163 = l_Lean_Elab_Term_expandCDot_x3f(x_111, x_2, x_3); +x_164 = lean_ctor_get(x_163, 0); +lean_inc(x_164); +if (lean_obj_tag(x_164) == 0) { -uint8_t x_166; -x_166 = !lean_is_exclusive(x_164); -if (x_166 == 0) +uint8_t x_165; +x_165 = !lean_is_exclusive(x_163); +if (x_165 == 0) { -lean_object* x_167; -x_167 = lean_ctor_get(x_164, 0); -lean_dec(x_167); -lean_ctor_set(x_164, 0, x_112); -return x_164; +lean_object* x_166; +x_166 = lean_ctor_get(x_163, 0); +lean_dec(x_166); +lean_ctor_set(x_163, 0, x_111); +return x_163; } else { -lean_object* x_168; lean_object* x_169; -x_168 = lean_ctor_get(x_164, 1); -lean_inc(x_168); -lean_dec(x_164); -x_169 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_169, 0, x_112); -lean_ctor_set(x_169, 1, x_168); -return x_169; +lean_object* x_167; lean_object* x_168; +x_167 = lean_ctor_get(x_163, 1); +lean_inc(x_167); +lean_dec(x_163); +x_168 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_168, 0, x_111); +lean_ctor_set(x_168, 1, x_167); +return x_168; } } else { -uint8_t x_170; -lean_dec(x_112); -x_170 = !lean_is_exclusive(x_164); -if (x_170 == 0) +uint8_t x_169; +lean_dec(x_111); +x_169 = !lean_is_exclusive(x_163); +if (x_169 == 0) { -lean_object* x_171; lean_object* x_172; +lean_object* x_170; lean_object* x_171; +x_170 = lean_ctor_get(x_163, 0); +lean_dec(x_170); x_171 = lean_ctor_get(x_164, 0); -lean_dec(x_171); -x_172 = lean_ctor_get(x_165, 0); -lean_inc(x_172); -lean_dec(x_165); -lean_ctor_set(x_164, 0, x_172); -return x_164; +lean_inc(x_171); +lean_dec(x_164); +lean_ctor_set(x_163, 0, x_171); +return x_163; } else { -lean_object* x_173; lean_object* x_174; lean_object* x_175; -x_173 = lean_ctor_get(x_164, 1); +lean_object* x_172; lean_object* x_173; lean_object* x_174; +x_172 = lean_ctor_get(x_163, 1); +lean_inc(x_172); +lean_dec(x_163); +x_173 = lean_ctor_get(x_164, 0); lean_inc(x_173); lean_dec(x_164); -x_174 = lean_ctor_get(x_165, 0); -lean_inc(x_174); -lean_dec(x_165); -x_175 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_175, 0, x_174); -lean_ctor_set(x_175, 1, x_173); -return x_175; +x_174 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_174, 0, x_173); +lean_ctor_set(x_174, 1, x_172); +return x_174; } } } } else { -lean_object* x_176; lean_object* x_177; uint8_t x_178; -x_176 = l_Lean_Syntax_getArg(x_113, x_60); -lean_dec(x_113); -x_177 = l_Lean_Elab_Term_expandParen___closed__3; -lean_inc(x_176); -x_178 = l_Lean_Syntax_isOfKind(x_176, x_177); -if (x_178 == 0) -{ -lean_object* x_179; uint8_t x_180; -x_179 = l_Lean_Elab_Term_expandParen___closed__5; -lean_inc(x_176); -x_180 = l_Lean_Syntax_isOfKind(x_176, x_179); -if (x_180 == 0) -{ -lean_object* x_181; uint8_t x_182; -lean_dec(x_176); +lean_object* x_175; lean_object* x_176; uint8_t x_177; +x_175 = l_Lean_Syntax_getArg(x_112, x_59); lean_dec(x_112); -x_181 = l_Lean_Syntax_getArg(x_58, x_60); -x_182 = l_Lean_Syntax_isMissing(x_181); -if (x_182 == 0) +x_176 = l_Lean_Elab_Term_expandParen___closed__3; +lean_inc(x_175); +x_177 = l_Lean_Syntax_isOfKind(x_175, x_176); +if (x_177 == 0) +{ +lean_object* x_178; uint8_t x_179; +x_178 = l_Lean_Elab_Term_expandParen___closed__5; +lean_inc(x_175); +x_179 = l_Lean_Syntax_isOfKind(x_175, x_178); +if (x_179 == 0) +{ +lean_object* x_180; uint8_t x_181; +lean_dec(x_175); +lean_dec(x_111); +x_180 = l_Lean_Syntax_getArg(x_58, x_59); +x_181 = l_Lean_Syntax_isMissing(x_180); +if (x_181 == 0) { -lean_object* x_183; uint8_t x_184; -x_183 = l_Lean_Syntax_getArg(x_58, x_57); +lean_object* x_182; uint8_t x_183; +x_182 = l_Lean_Syntax_getArg(x_58, x_57); lean_dec(x_58); -x_184 = l_Lean_Syntax_isMissing(x_183); -lean_dec(x_183); -if (x_184 == 0) +x_183 = l_Lean_Syntax_isMissing(x_182); +lean_dec(x_182); +if (x_183 == 0) { -lean_object* x_185; lean_object* x_186; lean_object* x_187; -lean_dec(x_181); +lean_object* x_184; lean_object* x_185; lean_object* x_186; +lean_dec(x_180); lean_dec(x_2); -x_185 = l_Lean_Elab_Term_expandParen___closed__1; -x_186 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_186, 0, x_1); -lean_ctor_set(x_186, 1, x_185); -x_187 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_187, 0, x_186); -lean_ctor_set(x_187, 1, x_3); -return x_187; -} -else -{ -lean_object* x_188; uint8_t x_189; -lean_dec(x_1); -x_188 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_189 = !lean_is_exclusive(x_188); -if (x_189 == 0) -{ -lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; -x_190 = lean_ctor_get(x_188, 0); -x_191 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; -lean_inc(x_190); -x_192 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_192, 0, x_190); -lean_ctor_set(x_192, 1, x_191); -x_193 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; -x_194 = lean_array_push(x_193, x_181); -x_195 = l_Lean_Elab_Term_expandShow___closed__23; -x_196 = lean_array_push(x_194, x_195); -x_197 = lean_box(2); -x_198 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; -x_199 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_199, 0, x_197); -lean_ctor_set(x_199, 1, x_198); -lean_ctor_set(x_199, 2, x_196); -x_200 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__23; -x_201 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_201, 0, x_190); -lean_ctor_set(x_201, 1, x_200); -x_202 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; -x_203 = lean_array_push(x_202, x_192); -x_204 = lean_array_push(x_203, x_199); -x_205 = lean_array_push(x_204, x_201); -x_206 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_206, 0, x_197); -lean_ctor_set(x_206, 1, x_4); -lean_ctor_set(x_206, 2, x_205); -lean_ctor_set(x_188, 0, x_206); -return x_188; -} -else -{ -lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; -x_207 = lean_ctor_get(x_188, 0); -x_208 = lean_ctor_get(x_188, 1); -lean_inc(x_208); -lean_inc(x_207); -lean_dec(x_188); -x_209 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; -lean_inc(x_207); -x_210 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_210, 0, x_207); -lean_ctor_set(x_210, 1, x_209); -x_211 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; -x_212 = lean_array_push(x_211, x_181); -x_213 = l_Lean_Elab_Term_expandShow___closed__23; -x_214 = lean_array_push(x_212, x_213); -x_215 = lean_box(2); -x_216 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; -x_217 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_217, 0, x_215); -lean_ctor_set(x_217, 1, x_216); -lean_ctor_set(x_217, 2, x_214); -x_218 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__23; -x_219 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_219, 0, x_207); -lean_ctor_set(x_219, 1, x_218); -x_220 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; -x_221 = lean_array_push(x_220, x_210); -x_222 = lean_array_push(x_221, x_217); -x_223 = lean_array_push(x_222, x_219); -x_224 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_224, 0, x_215); -lean_ctor_set(x_224, 1, x_4); -lean_ctor_set(x_224, 2, x_223); -x_225 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_225, 0, x_224); -lean_ctor_set(x_225, 1, x_208); -return x_225; -} +x_184 = l_Lean_Elab_Term_expandParen___closed__1; +x_185 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_185, 0, x_1); +lean_ctor_set(x_185, 1, x_184); +x_186 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_186, 0, x_185); +lean_ctor_set(x_186, 1, x_3); +return x_186; } +else +{ +lean_object* x_187; uint8_t x_188; +lean_dec(x_1); +x_187 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_188 = !lean_is_exclusive(x_187); +if (x_188 == 0) +{ +lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; +x_189 = lean_ctor_get(x_187, 0); +x_190 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; +lean_inc(x_189); +x_191 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_191, 0, x_189); +lean_ctor_set(x_191, 1, x_190); +x_192 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; +x_193 = lean_array_push(x_192, x_180); +x_194 = l_Lean_Elab_Term_expandShow___closed__23; +x_195 = lean_array_push(x_193, x_194); +x_196 = lean_box(2); +x_197 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; +x_198 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_198, 0, x_196); +lean_ctor_set(x_198, 1, x_197); +lean_ctor_set(x_198, 2, x_195); +x_199 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__23; +x_200 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_200, 0, x_189); +lean_ctor_set(x_200, 1, x_199); +x_201 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; +x_202 = lean_array_push(x_201, x_191); +x_203 = lean_array_push(x_202, x_198); +x_204 = lean_array_push(x_203, x_200); +x_205 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_205, 0, x_196); +lean_ctor_set(x_205, 1, x_4); +lean_ctor_set(x_205, 2, x_204); +lean_ctor_set(x_187, 0, x_205); +return x_187; } else { -lean_object* x_226; lean_object* x_227; lean_object* x_228; -lean_dec(x_181); +lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; +x_206 = lean_ctor_get(x_187, 0); +x_207 = lean_ctor_get(x_187, 1); +lean_inc(x_207); +lean_inc(x_206); +lean_dec(x_187); +x_208 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; +lean_inc(x_206); +x_209 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_209, 0, x_206); +lean_ctor_set(x_209, 1, x_208); +x_210 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; +x_211 = lean_array_push(x_210, x_180); +x_212 = l_Lean_Elab_Term_expandShow___closed__23; +x_213 = lean_array_push(x_211, x_212); +x_214 = lean_box(2); +x_215 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; +x_216 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_216, 0, x_214); +lean_ctor_set(x_216, 1, x_215); +lean_ctor_set(x_216, 2, x_213); +x_217 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__23; +x_218 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_218, 0, x_206); +lean_ctor_set(x_218, 1, x_217); +x_219 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; +x_220 = lean_array_push(x_219, x_209); +x_221 = lean_array_push(x_220, x_216); +x_222 = lean_array_push(x_221, x_218); +x_223 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_223, 0, x_214); +lean_ctor_set(x_223, 1, x_4); +lean_ctor_set(x_223, 2, x_222); +x_224 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_224, 0, x_223); +lean_ctor_set(x_224, 1, x_207); +return x_224; +} +} +} +else +{ +lean_object* x_225; lean_object* x_226; lean_object* x_227; +lean_dec(x_180); lean_dec(x_58); lean_dec(x_2); -x_226 = l_Lean_Elab_Term_expandParen___closed__1; -x_227 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_227, 0, x_1); -lean_ctor_set(x_227, 1, x_226); -x_228 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_228, 0, x_227); -lean_ctor_set(x_228, 1, x_3); -return x_228; +x_225 = l_Lean_Elab_Term_expandParen___closed__1; +x_226 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_226, 0, x_1); +lean_ctor_set(x_226, 1, x_225); +x_227 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_227, 0, x_226); +lean_ctor_set(x_227, 1, x_3); +return x_227; } } else { -lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; +lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_dec(x_58); lean_dec(x_1); -x_229 = l_Lean_Syntax_getArg(x_176, x_57); -lean_dec(x_176); -x_230 = l_Lean_Syntax_getArgs(x_229); +x_228 = l_Lean_Syntax_getArg(x_175, x_57); +lean_dec(x_175); +x_229 = l_Lean_Syntax_getArgs(x_228); +lean_dec(x_228); +x_230 = l_Lean_Elab_Term_expandShow___closed__27; +x_231 = lean_array_push(x_230, x_111); +x_232 = l_Lean_Syntax_TSepArray_getElems___rarg(x_229); lean_dec(x_229); -x_231 = l_Lean_Elab_Term_expandShow___closed__27; -x_232 = lean_array_push(x_231, x_112); -x_233 = l_Lean_Syntax_SepArray_getElems___rarg(x_230); -lean_dec(x_230); -x_234 = l_Array_append___rarg(x_232, x_233); +x_233 = l_Array_append___rarg(x_231, x_232); lean_inc(x_2); -x_235 = l_Lean_Elab_Term_mkPairs(x_234, x_2, x_3); -lean_dec(x_234); -x_236 = lean_ctor_get(x_235, 0); -lean_inc(x_236); -x_237 = lean_ctor_get(x_235, 1); -lean_inc(x_237); -lean_dec(x_235); +x_234 = l_Lean_Elab_Term_mkPairs(x_233, x_2, x_3); +lean_dec(x_233); +x_235 = lean_ctor_get(x_234, 0); +lean_inc(x_235); +x_236 = lean_ctor_get(x_234, 1); lean_inc(x_236); -x_238 = l_Lean_Elab_Term_expandCDot_x3f(x_236, x_2, x_237); -x_239 = lean_ctor_get(x_238, 0); -lean_inc(x_239); -if (lean_obj_tag(x_239) == 0) +lean_dec(x_234); +lean_inc(x_235); +x_237 = l_Lean_Elab_Term_expandCDot_x3f(x_235, x_2, x_236); +x_238 = lean_ctor_get(x_237, 0); +lean_inc(x_238); +if (lean_obj_tag(x_238) == 0) { -uint8_t x_240; -x_240 = !lean_is_exclusive(x_238); -if (x_240 == 0) +uint8_t x_239; +x_239 = !lean_is_exclusive(x_237); +if (x_239 == 0) { -lean_object* x_241; -x_241 = lean_ctor_get(x_238, 0); -lean_dec(x_241); -lean_ctor_set(x_238, 0, x_236); -return x_238; +lean_object* x_240; +x_240 = lean_ctor_get(x_237, 0); +lean_dec(x_240); +lean_ctor_set(x_237, 0, x_235); +return x_237; } else { -lean_object* x_242; lean_object* x_243; -x_242 = lean_ctor_get(x_238, 1); -lean_inc(x_242); -lean_dec(x_238); -x_243 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_243, 0, x_236); -lean_ctor_set(x_243, 1, x_242); -return x_243; +lean_object* x_241; lean_object* x_242; +x_241 = lean_ctor_get(x_237, 1); +lean_inc(x_241); +lean_dec(x_237); +x_242 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_242, 0, x_235); +lean_ctor_set(x_242, 1, x_241); +return x_242; } } else { -uint8_t x_244; -lean_dec(x_236); -x_244 = !lean_is_exclusive(x_238); -if (x_244 == 0) +uint8_t x_243; +lean_dec(x_235); +x_243 = !lean_is_exclusive(x_237); +if (x_243 == 0) { -lean_object* x_245; lean_object* x_246; +lean_object* x_244; lean_object* x_245; +x_244 = lean_ctor_get(x_237, 0); +lean_dec(x_244); x_245 = lean_ctor_get(x_238, 0); -lean_dec(x_245); -x_246 = lean_ctor_get(x_239, 0); -lean_inc(x_246); -lean_dec(x_239); -lean_ctor_set(x_238, 0, x_246); -return x_238; +lean_inc(x_245); +lean_dec(x_238); +lean_ctor_set(x_237, 0, x_245); +return x_237; } else { -lean_object* x_247; lean_object* x_248; lean_object* x_249; -x_247 = lean_ctor_get(x_238, 1); +lean_object* x_246; lean_object* x_247; lean_object* x_248; +x_246 = lean_ctor_get(x_237, 1); +lean_inc(x_246); +lean_dec(x_237); +x_247 = lean_ctor_get(x_238, 0); lean_inc(x_247); lean_dec(x_238); -x_248 = lean_ctor_get(x_239, 0); -lean_inc(x_248); -lean_dec(x_239); -x_249 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_249, 0, x_248); -lean_ctor_set(x_249, 1, x_247); -return x_249; +x_248 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_248, 0, x_247); +lean_ctor_set(x_248, 1, x_246); +return x_248; } } } } else { -lean_object* x_250; lean_object* x_251; lean_object* x_252; +lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_dec(x_58); lean_dec(x_1); -x_250 = l_Lean_Syntax_getArg(x_176, x_57); -lean_dec(x_176); +x_249 = l_Lean_Syntax_getArg(x_175, x_57); +lean_dec(x_175); lean_inc(x_2); -x_251 = l_Lean_Elab_Term_expandCDot_x3f(x_112, x_2, x_3); -x_252 = lean_ctor_get(x_251, 0); -lean_inc(x_252); -if (lean_obj_tag(x_252) == 0) +x_250 = l_Lean_Elab_Term_expandCDot_x3f(x_111, x_2, x_3); +x_251 = lean_ctor_get(x_250, 0); +lean_inc(x_251); +if (lean_obj_tag(x_251) == 0) { -uint8_t x_253; -lean_dec(x_250); +uint8_t x_252; +lean_dec(x_249); lean_dec(x_2); -x_253 = !lean_is_exclusive(x_251); -if (x_253 == 0) +x_252 = !lean_is_exclusive(x_250); +if (x_252 == 0) { -lean_object* x_254; lean_object* x_255; -x_254 = lean_ctor_get(x_251, 0); -lean_dec(x_254); -x_255 = lean_box(1); -lean_ctor_set_tag(x_251, 1); -lean_ctor_set(x_251, 0, x_255); -return x_251; +lean_object* x_253; lean_object* x_254; +x_253 = lean_ctor_get(x_250, 0); +lean_dec(x_253); +x_254 = lean_box(1); +lean_ctor_set_tag(x_250, 1); +lean_ctor_set(x_250, 0, x_254); +return x_250; } else { -lean_object* x_256; lean_object* x_257; lean_object* x_258; -x_256 = lean_ctor_get(x_251, 1); -lean_inc(x_256); -lean_dec(x_251); -x_257 = lean_box(1); -x_258 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_258, 0, x_257); -lean_ctor_set(x_258, 1, x_256); -return x_258; +lean_object* x_255; lean_object* x_256; lean_object* x_257; +x_255 = lean_ctor_get(x_250, 1); +lean_inc(x_255); +lean_dec(x_250); +x_256 = lean_box(1); +x_257 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_257, 0, x_256); +lean_ctor_set(x_257, 1, x_255); +return x_257; } } else { -lean_object* x_259; lean_object* x_260; lean_object* x_261; uint8_t x_262; -x_259 = lean_ctor_get(x_251, 1); +lean_object* x_258; lean_object* x_259; lean_object* x_260; uint8_t x_261; +x_258 = lean_ctor_get(x_250, 1); +lean_inc(x_258); +lean_dec(x_250); +x_259 = lean_ctor_get(x_251, 0); lean_inc(x_259); lean_dec(x_251); -x_260 = lean_ctor_get(x_252, 0); -lean_inc(x_260); -lean_dec(x_252); -x_261 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_259); -x_262 = !lean_is_exclusive(x_261); -if (x_262 == 0) -{ -lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; -x_263 = lean_ctor_get(x_261, 0); -x_264 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; -lean_inc(x_263); -x_265 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_265, 0, x_263); -lean_ctor_set(x_265, 1, x_264); -x_266 = l_Lean_Elab_Term_expandShow___closed__26; -lean_inc(x_263); -x_267 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_267, 0, x_263); -lean_ctor_set(x_267, 1, x_266); -x_268 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; -x_269 = lean_array_push(x_268, x_267); -x_270 = lean_array_push(x_269, x_250); -x_271 = lean_box(2); -x_272 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_272, 0, x_271); -lean_ctor_set(x_272, 1, x_177); -lean_ctor_set(x_272, 2, x_270); -x_273 = l_Lean_Elab_Term_expandShow___closed__27; -x_274 = lean_array_push(x_273, x_272); -x_275 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; -x_276 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_276, 0, x_271); -lean_ctor_set(x_276, 1, x_275); -lean_ctor_set(x_276, 2, x_274); -x_277 = lean_array_push(x_268, x_260); -x_278 = lean_array_push(x_277, x_276); -x_279 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_279, 0, x_271); -lean_ctor_set(x_279, 1, x_275); -lean_ctor_set(x_279, 2, x_278); -x_280 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__23; -x_281 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_281, 0, x_263); -lean_ctor_set(x_281, 1, x_280); -x_282 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; -x_283 = lean_array_push(x_282, x_265); -x_284 = lean_array_push(x_283, x_279); -x_285 = lean_array_push(x_284, x_281); -x_286 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_286, 0, x_271); -lean_ctor_set(x_286, 1, x_4); -lean_ctor_set(x_286, 2, x_285); -lean_ctor_set(x_261, 0, x_286); -return x_261; -} -else -{ -lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; -x_287 = lean_ctor_get(x_261, 0); -x_288 = lean_ctor_get(x_261, 1); -lean_inc(x_288); -lean_inc(x_287); -lean_dec(x_261); -x_289 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; +x_260 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_258); +x_261 = !lean_is_exclusive(x_260); +if (x_261 == 0) +{ +lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; +x_262 = lean_ctor_get(x_260, 0); +x_263 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; +lean_inc(x_262); +x_264 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_264, 0, x_262); +lean_ctor_set(x_264, 1, x_263); +x_265 = l_Lean_Elab_Term_expandShow___closed__26; +lean_inc(x_262); +x_266 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_266, 0, x_262); +lean_ctor_set(x_266, 1, x_265); +x_267 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; +x_268 = lean_array_push(x_267, x_266); +x_269 = lean_array_push(x_268, x_249); +x_270 = lean_box(2); +x_271 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_271, 0, x_270); +lean_ctor_set(x_271, 1, x_176); +lean_ctor_set(x_271, 2, x_269); +x_272 = l_Lean_Elab_Term_expandShow___closed__27; +x_273 = lean_array_push(x_272, x_271); +x_274 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; +x_275 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_275, 0, x_270); +lean_ctor_set(x_275, 1, x_274); +lean_ctor_set(x_275, 2, x_273); +x_276 = lean_array_push(x_267, x_259); +x_277 = lean_array_push(x_276, x_275); +x_278 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_278, 0, x_270); +lean_ctor_set(x_278, 1, x_274); +lean_ctor_set(x_278, 2, x_277); +x_279 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__23; +x_280 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_280, 0, x_262); +lean_ctor_set(x_280, 1, x_279); +x_281 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; +x_282 = lean_array_push(x_281, x_264); +x_283 = lean_array_push(x_282, x_278); +x_284 = lean_array_push(x_283, x_280); +x_285 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_285, 0, x_270); +lean_ctor_set(x_285, 1, x_4); +lean_ctor_set(x_285, 2, x_284); +lean_ctor_set(x_260, 0, x_285); +return x_260; +} +else +{ +lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; +x_286 = lean_ctor_get(x_260, 0); +x_287 = lean_ctor_get(x_260, 1); lean_inc(x_287); -x_290 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_290, 0, x_287); -lean_ctor_set(x_290, 1, x_289); -x_291 = l_Lean_Elab_Term_expandShow___closed__26; -lean_inc(x_287); -x_292 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_292, 0, x_287); -lean_ctor_set(x_292, 1, x_291); -x_293 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; -x_294 = lean_array_push(x_293, x_292); -x_295 = lean_array_push(x_294, x_250); -x_296 = lean_box(2); -x_297 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_297, 0, x_296); -lean_ctor_set(x_297, 1, x_177); -lean_ctor_set(x_297, 2, x_295); -x_298 = l_Lean_Elab_Term_expandShow___closed__27; -x_299 = lean_array_push(x_298, x_297); -x_300 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; -x_301 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_301, 0, x_296); -lean_ctor_set(x_301, 1, x_300); -lean_ctor_set(x_301, 2, x_299); -x_302 = lean_array_push(x_293, x_260); -x_303 = lean_array_push(x_302, x_301); -x_304 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_304, 0, x_296); -lean_ctor_set(x_304, 1, x_300); -lean_ctor_set(x_304, 2, x_303); -x_305 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__23; -x_306 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_306, 0, x_287); -lean_ctor_set(x_306, 1, x_305); -x_307 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; -x_308 = lean_array_push(x_307, x_290); -x_309 = lean_array_push(x_308, x_304); -x_310 = lean_array_push(x_309, x_306); -x_311 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_311, 0, x_296); -lean_ctor_set(x_311, 1, x_4); -lean_ctor_set(x_311, 2, x_310); -x_312 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_312, 0, x_311); -lean_ctor_set(x_312, 1, x_288); -return x_312; -} -} -} -} -} -} -else -{ -lean_object* x_313; uint8_t x_314; +lean_inc(x_286); +lean_dec(x_260); +x_288 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__15; +lean_inc(x_286); +x_289 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_289, 0, x_286); +lean_ctor_set(x_289, 1, x_288); +x_290 = l_Lean_Elab_Term_expandShow___closed__26; +lean_inc(x_286); +x_291 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_291, 0, x_286); +lean_ctor_set(x_291, 1, x_290); +x_292 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__10; +x_293 = lean_array_push(x_292, x_291); +x_294 = lean_array_push(x_293, x_249); +x_295 = lean_box(2); +x_296 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_296, 0, x_295); +lean_ctor_set(x_296, 1, x_176); +lean_ctor_set(x_296, 2, x_294); +x_297 = l_Lean_Elab_Term_expandShow___closed__27; +x_298 = lean_array_push(x_297, x_296); +x_299 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; +x_300 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_300, 0, x_295); +lean_ctor_set(x_300, 1, x_299); +lean_ctor_set(x_300, 2, x_298); +x_301 = lean_array_push(x_292, x_259); +x_302 = lean_array_push(x_301, x_300); +x_303 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_303, 0, x_295); +lean_ctor_set(x_303, 1, x_299); +lean_ctor_set(x_303, 2, x_302); +x_304 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__23; +x_305 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_305, 0, x_286); +lean_ctor_set(x_305, 1, x_304); +x_306 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; +x_307 = lean_array_push(x_306, x_289); +x_308 = lean_array_push(x_307, x_303); +x_309 = lean_array_push(x_308, x_305); +x_310 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_310, 0, x_295); +lean_ctor_set(x_310, 1, x_4); +lean_ctor_set(x_310, 2, x_309); +x_311 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_311, 0, x_310); +lean_ctor_set(x_311, 1, x_287); +return x_311; +} +} +} +} +} +} +else +{ +lean_object* x_312; uint8_t x_313; lean_dec(x_58); lean_dec(x_1); lean_inc(x_2); -x_313 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_314 = !lean_is_exclusive(x_313); -if (x_314 == 0) -{ -lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; -x_315 = lean_ctor_get(x_313, 0); -x_316 = lean_ctor_get(x_2, 2); +x_312 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_313 = !lean_is_exclusive(x_312); +if (x_313 == 0) +{ +lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; +x_314 = lean_ctor_get(x_312, 0); +x_315 = lean_ctor_get(x_2, 2); +lean_inc(x_315); +x_316 = lean_ctor_get(x_2, 1); lean_inc(x_316); -x_317 = lean_ctor_get(x_2, 1); -lean_inc(x_317); lean_dec(x_2); -x_318 = l_Lean_Elab_Term_expandParen___closed__12; -x_319 = l_Lean_addMacroScope(x_317, x_318, x_316); -x_320 = l_Lean_Elab_Term_expandParen___closed__8; -x_321 = l_Lean_Elab_Term_expandParen___closed__14; -x_322 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_322, 0, x_315); -lean_ctor_set(x_322, 1, x_320); -lean_ctor_set(x_322, 2, x_319); -lean_ctor_set(x_322, 3, x_321); -lean_ctor_set(x_313, 0, x_322); -return x_313; -} -else -{ -lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; -x_323 = lean_ctor_get(x_313, 0); -x_324 = lean_ctor_get(x_313, 1); -lean_inc(x_324); +x_317 = l_Lean_Elab_Term_expandParen___closed__12; +x_318 = l_Lean_addMacroScope(x_316, x_317, x_315); +x_319 = l_Lean_Elab_Term_expandParen___closed__8; +x_320 = l_Lean_Elab_Term_expandParen___closed__14; +x_321 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_321, 0, x_314); +lean_ctor_set(x_321, 1, x_319); +lean_ctor_set(x_321, 2, x_318); +lean_ctor_set(x_321, 3, x_320); +lean_ctor_set(x_312, 0, x_321); +return x_312; +} +else +{ +lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; +x_322 = lean_ctor_get(x_312, 0); +x_323 = lean_ctor_get(x_312, 1); lean_inc(x_323); -lean_dec(x_313); -x_325 = lean_ctor_get(x_2, 2); +lean_inc(x_322); +lean_dec(x_312); +x_324 = lean_ctor_get(x_2, 2); +lean_inc(x_324); +x_325 = lean_ctor_get(x_2, 1); lean_inc(x_325); -x_326 = lean_ctor_get(x_2, 1); -lean_inc(x_326); lean_dec(x_2); -x_327 = l_Lean_Elab_Term_expandParen___closed__12; -x_328 = l_Lean_addMacroScope(x_326, x_327, x_325); -x_329 = l_Lean_Elab_Term_expandParen___closed__8; -x_330 = l_Lean_Elab_Term_expandParen___closed__14; -x_331 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_331, 0, x_323); -lean_ctor_set(x_331, 1, x_329); -lean_ctor_set(x_331, 2, x_328); -lean_ctor_set(x_331, 3, x_330); -x_332 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_332, 0, x_331); -lean_ctor_set(x_332, 1, x_324); -return x_332; +x_326 = l_Lean_Elab_Term_expandParen___closed__12; +x_327 = l_Lean_addMacroScope(x_325, x_326, x_324); +x_328 = l_Lean_Elab_Term_expandParen___closed__8; +x_329 = l_Lean_Elab_Term_expandParen___closed__14; +x_330 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_330, 0, x_322); +lean_ctor_set(x_330, 1, x_328); +lean_ctor_set(x_330, 2, x_327); +lean_ctor_set(x_330, 3, x_329); +x_331 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_331, 0, x_330); +lean_ctor_set(x_331, 1, x_323); +return x_331; } } } @@ -15842,7 +16177,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandParen_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(254u); +x_1 = lean_unsigned_to_nat(253u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15854,7 +16189,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandParen_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(269u); +x_1 = lean_unsigned_to_nat(268u); x_2 = lean_unsigned_to_nat(74u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15882,7 +16217,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandParen_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(254u); +x_1 = lean_unsigned_to_nat(253u); x_2 = lean_unsigned_to_nat(43u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15894,7 +16229,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandParen_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(254u); +x_1 = lean_unsigned_to_nat(253u); x_2 = lean_unsigned_to_nat(54u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15962,17 +16297,16 @@ return x_12; } else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; x_13 = lean_unsigned_to_nat(1u); x_14 = l_Lean_Syntax_getArg(x_1, x_13); lean_dec(x_1); -x_15 = l_Lean_nullKind; -x_16 = lean_unsigned_to_nat(2u); +x_15 = lean_unsigned_to_nat(2u); lean_inc(x_14); -x_17 = l_Lean_Syntax_isNodeOf(x_14, x_15, x_16); -if (x_17 == 0) +x_16 = l_Lean_Syntax_matchesNull(x_14, x_15); +if (x_16 == 0) { -lean_object* x_18; +lean_object* x_17; lean_dec(x_14); lean_dec(x_8); lean_dec(x_7); @@ -15980,156 +16314,156 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_18 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_elabForall___spec__1___rarg(x_9); -return x_18; +x_17 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_elabForall___spec__1___rarg(x_9); +return x_17; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_19 = lean_unsigned_to_nat(0u); -x_20 = l_Lean_Syntax_getArg(x_14, x_19); -x_21 = l_Lean_Syntax_getArg(x_14, x_13); +lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_18 = lean_unsigned_to_nat(0u); +x_19 = l_Lean_Syntax_getArg(x_14, x_18); +x_20 = l_Lean_Syntax_getArg(x_14, x_13); lean_dec(x_14); -lean_inc(x_21); -x_22 = l_Lean_Syntax_isNodeOf(x_21, x_15, x_13); -if (x_22 == 0) +lean_inc(x_20); +x_21 = l_Lean_Syntax_matchesNull(x_20, x_13); +if (x_21 == 0) { -lean_object* x_23; -lean_dec(x_21); +lean_object* x_22; lean_dec(x_20); +lean_dec(x_19); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_23 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_elabForall___spec__1___rarg(x_9); -return x_23; +x_22 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_elabForall___spec__1___rarg(x_9); +return x_22; } else { -lean_object* x_24; lean_object* x_25; uint8_t x_26; -x_24 = l_Lean_Syntax_getArg(x_21, x_19); -lean_dec(x_21); -x_25 = l_Lean_Elab_Term_expandParen___closed__3; -lean_inc(x_24); -x_26 = l_Lean_Syntax_isOfKind(x_24, x_25); -if (x_26 == 0) -{ -lean_object* x_27; -lean_dec(x_24); +lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_23 = l_Lean_Syntax_getArg(x_20, x_18); lean_dec(x_20); +x_24 = l_Lean_Elab_Term_expandParen___closed__3; +lean_inc(x_23); +x_25 = l_Lean_Syntax_isOfKind(x_23, x_24); +if (x_25 == 0) +{ +lean_object* x_26; +lean_dec(x_23); +lean_dec(x_19); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_27 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_elabForall___spec__1___rarg(x_9); -return x_27; +x_26 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_elabForall___spec__1___rarg(x_9); +return x_26; } else { -lean_object* x_28; lean_object* x_29; uint8_t x_30; lean_object* x_31; -x_28 = l_Lean_Syntax_getArg(x_24, x_13); -lean_dec(x_24); -x_29 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabType), 8, 1); -lean_closure_set(x_29, 0, x_28); -x_30 = 1; +lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; +x_27 = l_Lean_Syntax_getArg(x_23, x_13); +lean_dec(x_23); +x_28 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabType), 8, 1); +lean_closure_set(x_28, 0, x_27); +x_29 = 1; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_31 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_29, x_30, x_30, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_31) == 0) +x_30 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_28, x_29, x_29, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_30) == 0) { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_32 = lean_ctor_get(x_31, 0); +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); lean_inc(x_32); -x_33 = lean_ctor_get(x_31, 1); -lean_inc(x_33); -lean_dec(x_31); -x_34 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_34, 0, x_32); +lean_dec(x_30); +x_33 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_33, 0, x_31); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -lean_inc(x_34); -x_35 = l_Lean_Elab_Term_elabTerm(x_20, x_34, x_30, x_30, x_3, x_4, x_5, x_6, x_7, x_8, x_33); -if (lean_obj_tag(x_35) == 0) +lean_inc(x_33); +x_34 = l_Lean_Elab_Term_elabTerm(x_19, x_33, x_29, x_29, x_3, x_4, x_5, x_6, x_7, x_8, x_32); +if (lean_obj_tag(x_34) == 0) { -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_36 = lean_ctor_get(x_35, 0); +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_34, 1); lean_inc(x_36); -x_37 = lean_ctor_get(x_35, 1); -lean_inc(x_37); -lean_dec(x_35); -x_38 = lean_box(0); -x_39 = l_Lean_Elab_Term_ensureHasType(x_34, x_36, x_38, x_3, x_4, x_5, x_6, x_7, x_8, x_37); -return x_39; +lean_dec(x_34); +x_37 = lean_box(0); +x_38 = l_Lean_Elab_Term_ensureHasType(x_33, x_35, x_37, x_3, x_4, x_5, x_6, x_7, x_8, x_36); +return x_38; } else { -uint8_t x_40; -lean_dec(x_34); +uint8_t x_39; +lean_dec(x_33); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_40 = !lean_is_exclusive(x_35); -if (x_40 == 0) +x_39 = !lean_is_exclusive(x_34); +if (x_39 == 0) { -return x_35; +return x_34; } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_35, 0); -x_42 = lean_ctor_get(x_35, 1); -lean_inc(x_42); +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_34, 0); +x_41 = lean_ctor_get(x_34, 1); lean_inc(x_41); -lean_dec(x_35); -x_43 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_43, 0, x_41); -lean_ctor_set(x_43, 1, x_42); -return x_43; +lean_inc(x_40); +lean_dec(x_34); +x_42 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_41); +return x_42; } } } else { -uint8_t x_44; -lean_dec(x_20); +uint8_t x_43; +lean_dec(x_19); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_44 = !lean_is_exclusive(x_31); -if (x_44 == 0) +x_43 = !lean_is_exclusive(x_30); +if (x_43 == 0) { -return x_31; +return x_30; } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_31, 0); -x_46 = lean_ctor_get(x_31, 1); -lean_inc(x_46); +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_30, 0); +x_45 = lean_ctor_get(x_30, 1); lean_inc(x_45); -lean_dec(x_31); -x_47 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_46); -return x_47; +lean_inc(x_44); +lean_dec(x_30); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; } } } @@ -16189,7 +16523,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabParen_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(271u); +x_1 = lean_unsigned_to_nat(270u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -16201,7 +16535,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabParen_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(277u); +x_1 = lean_unsigned_to_nat(276u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -16229,7 +16563,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabParen_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(271u); +x_1 = lean_unsigned_to_nat(270u); x_2 = lean_unsigned_to_nat(29u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -16241,7 +16575,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabParen_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(271u); +x_1 = lean_unsigned_to_nat(270u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -18158,18 +18492,17 @@ return x_15; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +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_16 = lean_unsigned_to_nat(0u); x_17 = l_Lean_Syntax_getArg(x_1, x_16); x_18 = lean_unsigned_to_nat(2u); x_19 = l_Lean_Syntax_getArg(x_1, x_18); -x_20 = l_Lean_nullKind; -x_21 = lean_unsigned_to_nat(1u); +x_20 = lean_unsigned_to_nat(1u); lean_inc(x_19); -x_22 = l_Lean_Syntax_isNodeOf(x_19, x_20, x_21); -if (x_22 == 0) +x_21 = l_Lean_Syntax_matchesNull(x_19, x_20); +if (x_21 == 0) { -lean_object* x_23; +lean_object* x_22; lean_dec(x_19); lean_dec(x_17); lean_dec(x_11); @@ -18180,16 +18513,16 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_23 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_elabForall___spec__1___rarg(x_12); -return x_23; +x_22 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_elabForall___spec__1___rarg(x_12); +return x_22; } else { -lean_object* x_24; lean_object* x_25; uint8_t x_26; lean_object* x_27; -x_24 = l_Lean_Syntax_getArg(x_19, x_16); +lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; +x_23 = l_Lean_Syntax_getArg(x_19, x_16); lean_dec(x_19); -x_25 = lean_box(0); -x_26 = 1; +x_24 = lean_box(0); +x_25 = 1; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); @@ -18197,585 +18530,585 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_17); -x_27 = l_Lean_Elab_Term_elabTerm(x_17, x_25, x_26, x_26, x_3, x_4, x_5, x_6, x_7, x_8, x_12); -if (lean_obj_tag(x_27) == 0) +x_26 = l_Lean_Elab_Term_elabTerm(x_17, x_24, x_25, x_25, x_3, x_4, x_5, x_6, x_7, x_8, x_12); +if (lean_obj_tag(x_26) == 0) { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_27, 0); +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); -lean_dec(x_27); +lean_dec(x_26); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_28); -x_30 = lean_infer_type(x_28, x_5, x_6, x_7, x_8, x_29); -if (lean_obj_tag(x_30) == 0) +lean_inc(x_27); +x_29 = lean_infer_type(x_27, x_5, x_6, x_7, x_8, x_28); +if (lean_obj_tag(x_29) == 0) { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_31 = lean_ctor_get(x_30, 0); +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_29, 1); lean_inc(x_31); -x_32 = lean_ctor_get(x_30, 1); -lean_inc(x_32); -lean_dec(x_30); +lean_dec(x_29); lean_inc(x_6); -x_33 = l_Lean_Meta_instantiateMVars(x_31, x_5, x_6, x_7, x_8, x_32); -x_34 = lean_ctor_get(x_33, 0); +x_32 = l_Lean_Meta_instantiateMVars(x_30, x_5, x_6, x_7, x_8, x_31); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); lean_inc(x_34); -x_35 = lean_ctor_get(x_33, 1); -lean_inc(x_35); -lean_dec(x_33); +lean_dec(x_32); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_34); -x_36 = l_Lean_Meta_matchEq_x3f(x_34, x_5, x_6, x_7, x_8, x_35); -if (lean_obj_tag(x_36) == 0) +lean_inc(x_33); +x_35 = l_Lean_Meta_matchEq_x3f(x_33, x_5, x_6, x_7, x_8, x_34); +if (lean_obj_tag(x_35) == 0) { -lean_object* x_37; -x_37 = lean_ctor_get(x_36, 0); -lean_inc(x_37); -if (lean_obj_tag(x_37) == 0) +lean_object* x_36; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +if (lean_obj_tag(x_36) == 0) { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -lean_dec(x_24); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +lean_dec(x_23); lean_dec(x_17); lean_dec(x_11); lean_dec(x_1); -x_38 = lean_ctor_get(x_36, 1); -lean_inc(x_38); -lean_dec(x_36); -x_39 = l_Lean_indentExpr(x_28); -x_40 = l_Lean_Elab_Term_elabSubst___closed__3; -x_41 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set(x_41, 1, x_39); -x_42 = l_Lean_Elab_Term_elabSubst___closed__5; -x_43 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_43, 0, x_41); -lean_ctor_set(x_43, 1, x_42); -x_44 = l_Lean_indentExpr(x_34); -x_45 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_45, 0, x_43); -lean_ctor_set(x_45, 1, x_44); -x_46 = l_Lean_Elab_Term_elabSubst___closed__7; -x_47 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_46); -x_48 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInst___spec__1(x_47, x_3, x_4, x_5, x_6, x_7, x_8, x_38); +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = l_Lean_indentExpr(x_27); +x_39 = l_Lean_Elab_Term_elabSubst___closed__3; +x_40 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_38); +x_41 = l_Lean_Elab_Term_elabSubst___closed__5; +x_42 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_41); +x_43 = l_Lean_indentExpr(x_33); +x_44 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_44, 0, x_42); +lean_ctor_set(x_44, 1, x_43); +x_45 = l_Lean_Elab_Term_elabSubst___closed__7; +x_46 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +x_47 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInst___spec__1(x_46, x_3, x_4, x_5, x_6, x_7, x_8, x_37); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -return x_48; +return x_47; } else { -lean_object* x_49; lean_object* x_50; -x_49 = lean_ctor_get(x_37, 0); +lean_object* x_48; lean_object* x_49; +x_48 = lean_ctor_get(x_36, 0); +lean_inc(x_48); +lean_dec(x_36); +x_49 = lean_ctor_get(x_48, 1); lean_inc(x_49); -lean_dec(x_37); -x_50 = lean_ctor_get(x_49, 1); -lean_inc(x_50); if (lean_obj_tag(x_11) == 0) { -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; -lean_dec(x_34); +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +lean_dec(x_33); lean_dec(x_17); lean_dec(x_1); -x_51 = lean_ctor_get(x_36, 1); +x_50 = lean_ctor_get(x_35, 1); +lean_inc(x_50); +lean_dec(x_35); +x_51 = lean_ctor_get(x_48, 0); lean_inc(x_51); -lean_dec(x_36); +lean_dec(x_48); x_52 = lean_ctor_get(x_49, 0); lean_inc(x_52); lean_dec(x_49); -x_53 = lean_ctor_get(x_50, 0); -lean_inc(x_53); -lean_dec(x_50); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_54 = l_Lean_Elab_Term_elabTerm(x_24, x_25, x_26, x_26, x_3, x_4, x_5, x_6, x_7, x_8, x_51); -if (lean_obj_tag(x_54) == 0) +x_53 = l_Lean_Elab_Term_elabTerm(x_23, x_24, x_25, x_25, x_3, x_4, x_5, x_6, x_7, x_8, x_50); +if (lean_obj_tag(x_53) == 0) { -lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_55 = lean_ctor_get(x_54, 0); +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_ctor_get(x_53, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_53, 1); lean_inc(x_55); -x_56 = lean_ctor_get(x_54, 1); -lean_inc(x_56); -lean_dec(x_54); +lean_dec(x_53); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_55); -x_57 = lean_infer_type(x_55, x_5, x_6, x_7, x_8, x_56); -if (lean_obj_tag(x_57) == 0) +lean_inc(x_54); +x_56 = lean_infer_type(x_54, x_5, x_6, x_7, x_8, x_55); +if (lean_obj_tag(x_56) == 0) { -lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_58 = lean_ctor_get(x_57, 0); +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_56, 1); lean_inc(x_58); -x_59 = lean_ctor_get(x_57, 1); -lean_inc(x_59); -lean_dec(x_57); -x_60 = lean_box(0); +lean_dec(x_56); +x_59 = lean_box(0); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_53); -x_61 = l_Lean_Meta_kabstract(x_58, x_53, x_60, x_5, x_6, x_7, x_8, x_59); -if (lean_obj_tag(x_61) == 0) +lean_inc(x_52); +x_60 = l_Lean_Meta_kabstract(x_57, x_52, x_59, x_5, x_6, x_7, x_8, x_58); +if (lean_obj_tag(x_60) == 0) { -lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; uint8_t x_69; lean_object* x_70; -x_62 = lean_ctor_get(x_61, 0); +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; lean_object* x_69; +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_60, 1); lean_inc(x_62); -x_63 = lean_ctor_get(x_61, 1); -lean_inc(x_63); -lean_dec(x_61); -x_64 = l_Lean_Elab_Term_elabSubst___lambda__6___closed__2; -x_65 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_64, x_7, x_8, x_63); -x_66 = lean_ctor_get(x_65, 0); +lean_dec(x_60); +x_63 = l_Lean_Elab_Term_elabSubst___lambda__6___closed__2; +x_64 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_63, x_7, x_8, x_62); +x_65 = lean_ctor_get(x_64, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_64, 1); lean_inc(x_66); -x_67 = lean_ctor_get(x_65, 1); -lean_inc(x_67); -lean_dec(x_65); -x_68 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabSubst___lambda__2), 10, 2); -lean_closure_set(x_68, 0, x_53); -lean_closure_set(x_68, 1, x_62); -x_69 = 0; +lean_dec(x_64); +x_67 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabSubst___lambda__2), 10, 2); +lean_closure_set(x_67, 0, x_52); +lean_closure_set(x_67, 1, x_61); +x_68 = 0; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_70 = l_Lean_Meta_withLocalDecl___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabImplicitLambda_loop___spec__1___rarg(x_66, x_69, x_52, x_68, x_3, x_4, x_5, x_6, x_7, x_8, x_67); -if (lean_obj_tag(x_70) == 0) +x_69 = l_Lean_Meta_withLocalDecl___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabImplicitLambda_loop___spec__1___rarg(x_65, x_68, x_51, x_67, x_3, x_4, x_5, x_6, x_7, x_8, x_66); +if (lean_obj_tag(x_69) == 0) { -lean_object* x_71; lean_object* x_72; lean_object* x_73; -x_71 = lean_ctor_get(x_70, 0); +lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_70 = lean_ctor_get(x_69, 0); +lean_inc(x_70); +x_71 = lean_ctor_get(x_69, 1); lean_inc(x_71); -x_72 = lean_ctor_get(x_70, 1); -lean_inc(x_72); -lean_dec(x_70); +lean_dec(x_69); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_71); -x_73 = l_Lean_Meta_isTypeCorrect(x_71, x_5, x_6, x_7, x_8, x_72); -if (lean_obj_tag(x_73) == 0) -{ -lean_object* x_74; uint8_t x_75; -x_74 = lean_ctor_get(x_73, 0); -lean_inc(x_74); -x_75 = lean_unbox(x_74); -lean_dec(x_74); -if (x_75 == 0) -{ -lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; -lean_dec(x_71); -lean_dec(x_55); -lean_dec(x_28); -x_76 = lean_ctor_get(x_73, 1); -lean_inc(x_76); +lean_inc(x_70); +x_72 = l_Lean_Meta_isTypeCorrect(x_70, x_5, x_6, x_7, x_8, x_71); +if (lean_obj_tag(x_72) == 0) +{ +lean_object* x_73; uint8_t x_74; +x_73 = lean_ctor_get(x_72, 0); +lean_inc(x_73); +x_74 = lean_unbox(x_73); lean_dec(x_73); -x_77 = l_Lean_Elab_Term_elabSubst___lambda__8___closed__2; -x_78 = l_Lean_throwError___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__1(x_77, x_3, x_4, x_5, x_6, x_7, x_8, x_76); +if (x_74 == 0) +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; +lean_dec(x_70); +lean_dec(x_54); +lean_dec(x_27); +x_75 = lean_ctor_get(x_72, 1); +lean_inc(x_75); +lean_dec(x_72); +x_76 = l_Lean_Elab_Term_elabSubst___lambda__8___closed__2; +x_77 = l_Lean_throwError___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__1(x_76, x_3, x_4, x_5, x_6, x_7, x_8, x_75); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_79 = !lean_is_exclusive(x_78); -if (x_79 == 0) +x_78 = !lean_is_exclusive(x_77); +if (x_78 == 0) { -return x_78; +return x_77; } else { -lean_object* x_80; lean_object* x_81; lean_object* x_82; -x_80 = lean_ctor_get(x_78, 0); -x_81 = lean_ctor_get(x_78, 1); -lean_inc(x_81); +lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_79 = lean_ctor_get(x_77, 0); +x_80 = lean_ctor_get(x_77, 1); lean_inc(x_80); -lean_dec(x_78); -x_82 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_82, 0, x_80); -lean_ctor_set(x_82, 1, x_81); -return x_82; +lean_inc(x_79); +lean_dec(x_77); +x_81 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_81, 0, x_79); +lean_ctor_set(x_81, 1, x_80); +return x_81; } } else { -lean_object* x_83; lean_object* x_84; +lean_object* x_82; lean_object* x_83; lean_dec(x_4); lean_dec(x_3); -x_83 = lean_ctor_get(x_73, 1); -lean_inc(x_83); -lean_dec(x_73); -x_84 = l_Lean_Meta_mkEqRec(x_71, x_55, x_28, x_5, x_6, x_7, x_8, x_83); -return x_84; +x_82 = lean_ctor_get(x_72, 1); +lean_inc(x_82); +lean_dec(x_72); +x_83 = l_Lean_Meta_mkEqRec(x_70, x_54, x_27, x_5, x_6, x_7, x_8, x_82); +return x_83; } } else { -uint8_t x_85; -lean_dec(x_71); -lean_dec(x_55); -lean_dec(x_28); +uint8_t x_84; +lean_dec(x_70); +lean_dec(x_54); +lean_dec(x_27); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_85 = !lean_is_exclusive(x_73); -if (x_85 == 0) +x_84 = !lean_is_exclusive(x_72); +if (x_84 == 0) { -return x_73; +return x_72; } else { -lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_86 = lean_ctor_get(x_73, 0); -x_87 = lean_ctor_get(x_73, 1); -lean_inc(x_87); +lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_85 = lean_ctor_get(x_72, 0); +x_86 = lean_ctor_get(x_72, 1); lean_inc(x_86); -lean_dec(x_73); -x_88 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_88, 0, x_86); -lean_ctor_set(x_88, 1, x_87); -return x_88; +lean_inc(x_85); +lean_dec(x_72); +x_87 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_87, 0, x_85); +lean_ctor_set(x_87, 1, x_86); +return x_87; } } } else { -uint8_t x_89; -lean_dec(x_55); -lean_dec(x_28); +uint8_t x_88; +lean_dec(x_54); +lean_dec(x_27); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_89 = !lean_is_exclusive(x_70); -if (x_89 == 0) +x_88 = !lean_is_exclusive(x_69); +if (x_88 == 0) { -return x_70; +return x_69; } else { -lean_object* x_90; lean_object* x_91; lean_object* x_92; -x_90 = lean_ctor_get(x_70, 0); -x_91 = lean_ctor_get(x_70, 1); -lean_inc(x_91); +lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_89 = lean_ctor_get(x_69, 0); +x_90 = lean_ctor_get(x_69, 1); lean_inc(x_90); -lean_dec(x_70); -x_92 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_92, 0, x_90); -lean_ctor_set(x_92, 1, x_91); -return x_92; +lean_inc(x_89); +lean_dec(x_69); +x_91 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_91, 0, x_89); +lean_ctor_set(x_91, 1, x_90); +return x_91; } } } else { -uint8_t x_93; -lean_dec(x_55); -lean_dec(x_53); +uint8_t x_92; +lean_dec(x_54); lean_dec(x_52); -lean_dec(x_28); +lean_dec(x_51); +lean_dec(x_27); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_93 = !lean_is_exclusive(x_61); -if (x_93 == 0) +x_92 = !lean_is_exclusive(x_60); +if (x_92 == 0) { -return x_61; +return x_60; } else { -lean_object* x_94; lean_object* x_95; lean_object* x_96; -x_94 = lean_ctor_get(x_61, 0); -x_95 = lean_ctor_get(x_61, 1); -lean_inc(x_95); +lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_93 = lean_ctor_get(x_60, 0); +x_94 = lean_ctor_get(x_60, 1); lean_inc(x_94); -lean_dec(x_61); -x_96 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_96, 0, x_94); -lean_ctor_set(x_96, 1, x_95); -return x_96; +lean_inc(x_93); +lean_dec(x_60); +x_95 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_95, 0, x_93); +lean_ctor_set(x_95, 1, x_94); +return x_95; } } } else { -uint8_t x_97; -lean_dec(x_55); -lean_dec(x_53); +uint8_t x_96; +lean_dec(x_54); lean_dec(x_52); -lean_dec(x_28); +lean_dec(x_51); +lean_dec(x_27); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_97 = !lean_is_exclusive(x_57); -if (x_97 == 0) +x_96 = !lean_is_exclusive(x_56); +if (x_96 == 0) { -return x_57; +return x_56; } else { -lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_98 = lean_ctor_get(x_57, 0); -x_99 = lean_ctor_get(x_57, 1); -lean_inc(x_99); +lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_97 = lean_ctor_get(x_56, 0); +x_98 = lean_ctor_get(x_56, 1); lean_inc(x_98); -lean_dec(x_57); -x_100 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_100, 0, x_98); -lean_ctor_set(x_100, 1, x_99); -return x_100; +lean_inc(x_97); +lean_dec(x_56); +x_99 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_99, 0, x_97); +lean_ctor_set(x_99, 1, x_98); +return x_99; } } } else { -uint8_t x_101; -lean_dec(x_53); +uint8_t x_100; lean_dec(x_52); -lean_dec(x_28); +lean_dec(x_51); +lean_dec(x_27); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_101 = !lean_is_exclusive(x_54); -if (x_101 == 0) +x_100 = !lean_is_exclusive(x_53); +if (x_100 == 0) { -return x_54; +return x_53; } else { -lean_object* x_102; lean_object* x_103; lean_object* x_104; -x_102 = lean_ctor_get(x_54, 0); -x_103 = lean_ctor_get(x_54, 1); -lean_inc(x_103); +lean_object* x_101; lean_object* x_102; lean_object* x_103; +x_101 = lean_ctor_get(x_53, 0); +x_102 = lean_ctor_get(x_53, 1); lean_inc(x_102); -lean_dec(x_54); -x_104 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_104, 0, x_102); -lean_ctor_set(x_104, 1, x_103); -return x_104; +lean_inc(x_101); +lean_dec(x_53); +x_103 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_103, 0, x_101); +lean_ctor_set(x_103, 1, x_102); +return x_103; } } } else { -lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; -x_105 = lean_ctor_get(x_36, 1); +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; +x_104 = lean_ctor_get(x_35, 1); +lean_inc(x_104); +lean_dec(x_35); +x_105 = lean_ctor_get(x_48, 0); lean_inc(x_105); -lean_dec(x_36); +lean_dec(x_48); x_106 = lean_ctor_get(x_49, 0); lean_inc(x_106); -lean_dec(x_49); -x_107 = lean_ctor_get(x_50, 0); +x_107 = lean_ctor_get(x_49, 1); lean_inc(x_107); -x_108 = lean_ctor_get(x_50, 1); +lean_dec(x_49); +x_108 = lean_ctor_get(x_11, 0); lean_inc(x_108); -lean_dec(x_50); -x_109 = lean_ctor_get(x_11, 0); -lean_inc(x_109); lean_dec(x_11); -x_110 = lean_box(0); +x_109 = lean_box(0); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); +lean_inc(x_107); lean_inc(x_108); -lean_inc(x_109); -x_111 = l_Lean_Meta_kabstract(x_109, x_108, x_110, x_5, x_6, x_7, x_8, x_105); -if (lean_obj_tag(x_111) == 0) +x_110 = l_Lean_Meta_kabstract(x_108, x_107, x_109, x_5, x_6, x_7, x_8, x_104); +if (lean_obj_tag(x_110) == 0) { -lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; uint8_t x_117; -x_112 = lean_ctor_get(x_111, 0); +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; uint8_t x_116; +x_111 = lean_ctor_get(x_110, 0); +lean_inc(x_111); +x_112 = lean_ctor_get(x_110, 1); lean_inc(x_112); -x_113 = lean_ctor_get(x_111, 1); -lean_inc(x_113); -lean_dec(x_111); -x_114 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; -x_115 = l_Lean_Elab_Term_elabAnonymousCtor___closed__2; +lean_dec(x_110); +x_113 = l_Lean_Elab_Term_elabAnonymousCtor___closed__3; +x_114 = l_Lean_Elab_Term_elabAnonymousCtor___closed__2; lean_inc(x_17); -lean_inc(x_24); +lean_inc(x_23); lean_inc(x_1); -lean_inc(x_109); -lean_inc(x_106); -x_116 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabSubst___lambda__8___boxed), 20, 8); -lean_closure_set(x_116, 0, x_106); -lean_closure_set(x_116, 1, x_114); -lean_closure_set(x_116, 2, x_115); -lean_closure_set(x_116, 3, x_109); -lean_closure_set(x_116, 4, x_1); -lean_closure_set(x_116, 5, x_24); -lean_closure_set(x_116, 6, x_17); -lean_closure_set(x_116, 7, x_25); -x_117 = l_Lean_Expr_hasLooseBVars(x_112); -if (x_117 == 0) +lean_inc(x_108); +lean_inc(x_105); +x_115 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabSubst___lambda__8___boxed), 20, 8); +lean_closure_set(x_115, 0, x_105); +lean_closure_set(x_115, 1, x_113); +lean_closure_set(x_115, 2, x_114); +lean_closure_set(x_115, 3, x_108); +lean_closure_set(x_115, 4, x_1); +lean_closure_set(x_115, 5, x_23); +lean_closure_set(x_115, 6, x_17); +lean_closure_set(x_115, 7, x_24); +x_116 = l_Lean_Expr_hasLooseBVars(x_111); +if (x_116 == 0) { -lean_object* x_118; -lean_dec(x_112); -lean_dec(x_106); -lean_dec(x_24); +lean_object* x_117; +lean_dec(x_111); +lean_dec(x_105); +lean_dec(x_23); lean_dec(x_17); lean_dec(x_1); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_107); -lean_inc(x_109); -x_118 = l_Lean_Meta_kabstract(x_109, x_107, x_110, x_5, x_6, x_7, x_8, x_113); -if (lean_obj_tag(x_118) == 0) +lean_inc(x_106); +lean_inc(x_108); +x_117 = l_Lean_Meta_kabstract(x_108, x_106, x_109, x_5, x_6, x_7, x_8, x_112); +if (lean_obj_tag(x_117) == 0) { -lean_object* x_119; lean_object* x_120; uint8_t x_121; -x_119 = lean_ctor_get(x_118, 0); +lean_object* x_118; lean_object* x_119; uint8_t x_120; +x_118 = lean_ctor_get(x_117, 0); +lean_inc(x_118); +x_119 = lean_ctor_get(x_117, 1); lean_inc(x_119); -x_120 = lean_ctor_get(x_118, 1); -lean_inc(x_120); -lean_dec(x_118); -x_121 = l_Lean_Expr_hasLooseBVars(x_119); -if (x_121 == 0) +lean_dec(x_117); +x_120 = l_Lean_Expr_hasLooseBVars(x_118); +if (x_120 == 0) { -lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; uint8_t x_132; -lean_dec(x_119); -lean_dec(x_116); -lean_dec(x_108); +lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; uint8_t x_131; +lean_dec(x_118); +lean_dec(x_115); lean_dec(x_107); -lean_dec(x_28); -x_122 = l_Lean_indentExpr(x_109); -x_123 = l_Lean_Elab_Term_elabSubst___closed__9; -x_124 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_124, 0, x_123); -lean_ctor_set(x_124, 1, x_122); -x_125 = l_Lean_Elab_Term_elabSubst___closed__11; -x_126 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_126, 0, x_124); -lean_ctor_set(x_126, 1, x_125); -x_127 = l_Lean_indentExpr(x_34); -x_128 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_128, 0, x_126); -lean_ctor_set(x_128, 1, x_127); -x_129 = l_Lean_Elab_Term_elabAnonymousCtor___closed__11; -x_130 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_130, 0, x_128); -lean_ctor_set(x_130, 1, x_129); -x_131 = l_Lean_throwError___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__1(x_130, x_3, x_4, x_5, x_6, x_7, x_8, x_120); +lean_dec(x_106); +lean_dec(x_27); +x_121 = l_Lean_indentExpr(x_108); +x_122 = l_Lean_Elab_Term_elabSubst___closed__9; +x_123 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_123, 0, x_122); +lean_ctor_set(x_123, 1, x_121); +x_124 = l_Lean_Elab_Term_elabSubst___closed__11; +x_125 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_125, 0, x_123); +lean_ctor_set(x_125, 1, x_124); +x_126 = l_Lean_indentExpr(x_33); +x_127 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +x_128 = l_Lean_Elab_Term_elabAnonymousCtor___closed__11; +x_129 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_129, 0, x_127); +lean_ctor_set(x_129, 1, x_128); +x_130 = l_Lean_throwError___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__1(x_129, x_3, x_4, x_5, x_6, x_7, x_8, x_119); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_132 = !lean_is_exclusive(x_131); -if (x_132 == 0) +x_131 = !lean_is_exclusive(x_130); +if (x_131 == 0) { -return x_131; +return x_130; } else { -lean_object* x_133; lean_object* x_134; lean_object* x_135; -x_133 = lean_ctor_get(x_131, 0); -x_134 = lean_ctor_get(x_131, 1); -lean_inc(x_134); +lean_object* x_132; lean_object* x_133; lean_object* x_134; +x_132 = lean_ctor_get(x_130, 0); +x_133 = lean_ctor_get(x_130, 1); lean_inc(x_133); -lean_dec(x_131); -x_135 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_135, 0, x_133); -lean_ctor_set(x_135, 1, x_134); -return x_135; +lean_inc(x_132); +lean_dec(x_130); +x_134 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_134, 0, x_132); +lean_ctor_set(x_134, 1, x_133); +return x_134; } } else { -lean_object* x_136; lean_object* x_137; -lean_dec(x_109); -lean_dec(x_34); -x_136 = lean_box(0); -x_137 = l_Lean_Elab_Term_elabSubst___lambda__9(x_116, x_119, x_28, x_107, x_108, x_136, x_3, x_4, x_5, x_6, x_7, x_8, x_120); -return x_137; +lean_object* x_135; lean_object* x_136; +lean_dec(x_108); +lean_dec(x_33); +x_135 = lean_box(0); +x_136 = l_Lean_Elab_Term_elabSubst___lambda__9(x_115, x_118, x_27, x_106, x_107, x_135, x_3, x_4, x_5, x_6, x_7, x_8, x_119); +return x_136; } } else { -uint8_t x_138; -lean_dec(x_116); -lean_dec(x_109); +uint8_t x_137; +lean_dec(x_115); lean_dec(x_108); lean_dec(x_107); -lean_dec(x_34); -lean_dec(x_28); +lean_dec(x_106); +lean_dec(x_33); +lean_dec(x_27); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_138 = !lean_is_exclusive(x_118); -if (x_138 == 0) +x_137 = !lean_is_exclusive(x_117); +if (x_137 == 0) { -return x_118; +return x_117; } else { -lean_object* x_139; lean_object* x_140; lean_object* x_141; -x_139 = lean_ctor_get(x_118, 0); -x_140 = lean_ctor_get(x_118, 1); -lean_inc(x_140); +lean_object* x_138; lean_object* x_139; lean_object* x_140; +x_138 = lean_ctor_get(x_117, 0); +x_139 = lean_ctor_get(x_117, 1); lean_inc(x_139); -lean_dec(x_118); -x_141 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_141, 0, x_139); -lean_ctor_set(x_141, 1, x_140); -return x_141; +lean_inc(x_138); +lean_dec(x_117); +x_140 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_140, 0, x_138); +lean_ctor_set(x_140, 1, x_139); +return x_140; } } } else { -lean_object* x_142; lean_object* x_143; -lean_dec(x_116); -lean_dec(x_34); -x_142 = lean_box(0); -x_143 = l_Lean_Elab_Term_elabSubst___lambda__8(x_106, x_114, x_115, x_109, x_1, x_24, x_17, x_25, x_112, x_28, x_107, x_108, x_142, x_3, x_4, x_5, x_6, x_7, x_8, x_113); -return x_143; +lean_object* x_141; lean_object* x_142; +lean_dec(x_115); +lean_dec(x_33); +x_141 = lean_box(0); +x_142 = l_Lean_Elab_Term_elabSubst___lambda__8(x_105, x_113, x_114, x_108, x_1, x_23, x_17, x_24, x_111, x_27, x_106, x_107, x_141, x_3, x_4, x_5, x_6, x_7, x_8, x_112); +return x_142; } } else { -uint8_t x_144; -lean_dec(x_109); +uint8_t x_143; lean_dec(x_108); lean_dec(x_107); lean_dec(x_106); -lean_dec(x_34); -lean_dec(x_28); -lean_dec(x_24); +lean_dec(x_105); +lean_dec(x_33); +lean_dec(x_27); +lean_dec(x_23); lean_dec(x_17); lean_dec(x_8); lean_dec(x_7); @@ -18784,23 +19117,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_144 = !lean_is_exclusive(x_111); -if (x_144 == 0) +x_143 = !lean_is_exclusive(x_110); +if (x_143 == 0) { -return x_111; +return x_110; } else { -lean_object* x_145; lean_object* x_146; lean_object* x_147; -x_145 = lean_ctor_get(x_111, 0); -x_146 = lean_ctor_get(x_111, 1); -lean_inc(x_146); +lean_object* x_144; lean_object* x_145; lean_object* x_146; +x_144 = lean_ctor_get(x_110, 0); +x_145 = lean_ctor_get(x_110, 1); lean_inc(x_145); -lean_dec(x_111); -x_147 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_147, 0, x_145); -lean_ctor_set(x_147, 1, x_146); -return x_147; +lean_inc(x_144); +lean_dec(x_110); +x_146 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_146, 0, x_144); +lean_ctor_set(x_146, 1, x_145); +return x_146; } } } @@ -18808,10 +19141,10 @@ return x_147; } else { -uint8_t x_148; -lean_dec(x_34); -lean_dec(x_28); -lean_dec(x_24); +uint8_t x_147; +lean_dec(x_33); +lean_dec(x_27); +lean_dec(x_23); lean_dec(x_17); lean_dec(x_11); lean_dec(x_8); @@ -18821,31 +19154,31 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_148 = !lean_is_exclusive(x_36); -if (x_148 == 0) +x_147 = !lean_is_exclusive(x_35); +if (x_147 == 0) { -return x_36; +return x_35; } else { -lean_object* x_149; lean_object* x_150; lean_object* x_151; -x_149 = lean_ctor_get(x_36, 0); -x_150 = lean_ctor_get(x_36, 1); -lean_inc(x_150); +lean_object* x_148; lean_object* x_149; lean_object* x_150; +x_148 = lean_ctor_get(x_35, 0); +x_149 = lean_ctor_get(x_35, 1); lean_inc(x_149); -lean_dec(x_36); -x_151 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_151, 0, x_149); -lean_ctor_set(x_151, 1, x_150); -return x_151; +lean_inc(x_148); +lean_dec(x_35); +x_150 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_150, 0, x_148); +lean_ctor_set(x_150, 1, x_149); +return x_150; } } } else { -uint8_t x_152; -lean_dec(x_28); -lean_dec(x_24); +uint8_t x_151; +lean_dec(x_27); +lean_dec(x_23); lean_dec(x_17); lean_dec(x_11); lean_dec(x_8); @@ -18855,30 +19188,30 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_152 = !lean_is_exclusive(x_30); -if (x_152 == 0) +x_151 = !lean_is_exclusive(x_29); +if (x_151 == 0) { -return x_30; +return x_29; } else { -lean_object* x_153; lean_object* x_154; lean_object* x_155; -x_153 = lean_ctor_get(x_30, 0); -x_154 = lean_ctor_get(x_30, 1); -lean_inc(x_154); +lean_object* x_152; lean_object* x_153; lean_object* x_154; +x_152 = lean_ctor_get(x_29, 0); +x_153 = lean_ctor_get(x_29, 1); lean_inc(x_153); -lean_dec(x_30); -x_155 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_155, 0, x_153); -lean_ctor_set(x_155, 1, x_154); -return x_155; +lean_inc(x_152); +lean_dec(x_29); +x_154 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_154, 0, x_152); +lean_ctor_set(x_154, 1, x_153); +return x_154; } } } else { -uint8_t x_156; -lean_dec(x_24); +uint8_t x_155; +lean_dec(x_23); lean_dec(x_17); lean_dec(x_11); lean_dec(x_8); @@ -18888,23 +19221,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_156 = !lean_is_exclusive(x_27); -if (x_156 == 0) +x_155 = !lean_is_exclusive(x_26); +if (x_155 == 0) { -return x_27; +return x_26; } else { -lean_object* x_157; lean_object* x_158; lean_object* x_159; -x_157 = lean_ctor_get(x_27, 0); -x_158 = lean_ctor_get(x_27, 1); -lean_inc(x_158); +lean_object* x_156; lean_object* x_157; lean_object* x_158; +x_156 = lean_ctor_get(x_26, 0); +x_157 = lean_ctor_get(x_26, 1); lean_inc(x_157); -lean_dec(x_27); -x_159 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_159, 0, x_157); -lean_ctor_set(x_159, 1, x_158); -return x_159; +lean_inc(x_156); +lean_dec(x_26); +x_158 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_158, 0, x_156); +lean_ctor_set(x_158, 1, x_157); +return x_158; } } } @@ -18912,7 +19245,7 @@ return x_159; } else { -uint8_t x_160; +uint8_t x_159; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -18920,23 +19253,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_160 = !lean_is_exclusive(x_10); -if (x_160 == 0) +x_159 = !lean_is_exclusive(x_10); +if (x_159 == 0) { return x_10; } else { -lean_object* x_161; lean_object* x_162; lean_object* x_163; -x_161 = lean_ctor_get(x_10, 0); -x_162 = lean_ctor_get(x_10, 1); -lean_inc(x_162); +lean_object* x_160; lean_object* x_161; lean_object* x_162; +x_160 = lean_ctor_get(x_10, 0); +x_161 = lean_ctor_get(x_10, 1); lean_inc(x_161); +lean_inc(x_160); lean_dec(x_10); -x_163 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_163, 0, x_161); -lean_ctor_set(x_163, 1, x_162); -return x_163; +x_162 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_162, 0, x_160); +lean_ctor_set(x_162, 1, x_161); +return x_162; } } } @@ -19099,7 +19432,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabSubst_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(304u); +x_1 = lean_unsigned_to_nat(303u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19111,7 +19444,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabSubst_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(368u); +x_1 = lean_unsigned_to_nat(367u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19139,7 +19472,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabSubst_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(304u); +x_1 = lean_unsigned_to_nat(303u); x_2 = lean_unsigned_to_nat(29u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19151,7 +19484,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabSubst_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(304u); +x_1 = lean_unsigned_to_nat(303u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19587,7 +19920,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabStateRefT_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(370u); +x_1 = lean_unsigned_to_nat(369u); x_2 = lean_unsigned_to_nat(29u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19599,7 +19932,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabStateRefT_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(379u); +x_1 = lean_unsigned_to_nat(378u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19627,7 +19960,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabStateRefT_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(370u); +x_1 = lean_unsigned_to_nat(369u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19639,7 +19972,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabStateRefT_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(370u); +x_1 = lean_unsigned_to_nat(369u); x_2 = lean_unsigned_to_nat(46u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19813,7 +20146,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabNoindex_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(381u); +x_1 = lean_unsigned_to_nat(380u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19825,7 +20158,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabNoindex_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(383u); +x_1 = lean_unsigned_to_nat(382u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19853,7 +20186,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabNoindex_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(381u); +x_1 = lean_unsigned_to_nat(380u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19865,7 +20198,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabNoindex_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(381u); +x_1 = lean_unsigned_to_nat(380u); x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -20509,16 +20842,16 @@ l_Lean_Elab_Term_elabPanic___closed__11 = _init_l_Lean_Elab_Term_elabPanic___clo lean_mark_persistent(l_Lean_Elab_Term_elabPanic___closed__11); l_Lean_Elab_Term_elabPanic___closed__12 = _init_l_Lean_Elab_Term_elabPanic___closed__12(); lean_mark_persistent(l_Lean_Elab_Term_elabPanic___closed__12); +l_Lean_Elab_Term_elabPanic___closed__13 = _init_l_Lean_Elab_Term_elabPanic___closed__13(); +lean_mark_persistent(l_Lean_Elab_Term_elabPanic___closed__13); +l_Lean_Elab_Term_elabPanic___closed__14 = _init_l_Lean_Elab_Term_elabPanic___closed__14(); +lean_mark_persistent(l_Lean_Elab_Term_elabPanic___closed__14); l___regBuiltin_Lean_Elab_Term_elabPanic___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_elabPanic___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_elabPanic___closed__1); l___regBuiltin_Lean_Elab_Term_elabPanic___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_elabPanic___closed__2(); lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_elabPanic___closed__2); l___regBuiltin_Lean_Elab_Term_elabPanic___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_elabPanic___closed__3(); lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_elabPanic___closed__3); -l___regBuiltin_Lean_Elab_Term_elabPanic___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_elabPanic___closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_elabPanic___closed__4); -l___regBuiltin_Lean_Elab_Term_elabPanic___closed__5 = _init_l___regBuiltin_Lean_Elab_Term_elabPanic___closed__5(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_elabPanic___closed__5); res = l___regBuiltin_Lean_Elab_Term_elabPanic(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); @@ -20609,16 +20942,16 @@ l_Lean_Elab_Term_expandAssert___closed__10 = _init_l_Lean_Elab_Term_expandAssert lean_mark_persistent(l_Lean_Elab_Term_expandAssert___closed__10); l_Lean_Elab_Term_expandAssert___closed__11 = _init_l_Lean_Elab_Term_expandAssert___closed__11(); lean_mark_persistent(l_Lean_Elab_Term_expandAssert___closed__11); +l_Lean_Elab_Term_expandAssert___closed__12 = _init_l_Lean_Elab_Term_expandAssert___closed__12(); +lean_mark_persistent(l_Lean_Elab_Term_expandAssert___closed__12); +l_Lean_Elab_Term_expandAssert___closed__13 = _init_l_Lean_Elab_Term_expandAssert___closed__13(); +lean_mark_persistent(l_Lean_Elab_Term_expandAssert___closed__13); l___regBuiltin_Lean_Elab_Term_expandAssert___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_expandAssert___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_expandAssert___closed__1); l___regBuiltin_Lean_Elab_Term_expandAssert___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_expandAssert___closed__2(); lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_expandAssert___closed__2); l___regBuiltin_Lean_Elab_Term_expandAssert___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_expandAssert___closed__3(); lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_expandAssert___closed__3); -l___regBuiltin_Lean_Elab_Term_expandAssert___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_expandAssert___closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_expandAssert___closed__4); -l___regBuiltin_Lean_Elab_Term_expandAssert___closed__5 = _init_l___regBuiltin_Lean_Elab_Term_expandAssert___closed__5(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_expandAssert___closed__5); res = l___regBuiltin_Lean_Elab_Term_expandAssert(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); @@ -20696,14 +21029,18 @@ l_Lean_Elab_Term_expandDbgTrace___closed__25 = _init_l_Lean_Elab_Term_expandDbgT lean_mark_persistent(l_Lean_Elab_Term_expandDbgTrace___closed__25); l_Lean_Elab_Term_expandDbgTrace___closed__26 = _init_l_Lean_Elab_Term_expandDbgTrace___closed__26(); lean_mark_persistent(l_Lean_Elab_Term_expandDbgTrace___closed__26); +l_Lean_Elab_Term_expandDbgTrace___closed__27 = _init_l_Lean_Elab_Term_expandDbgTrace___closed__27(); +lean_mark_persistent(l_Lean_Elab_Term_expandDbgTrace___closed__27); +l_Lean_Elab_Term_expandDbgTrace___closed__28 = _init_l_Lean_Elab_Term_expandDbgTrace___closed__28(); +lean_mark_persistent(l_Lean_Elab_Term_expandDbgTrace___closed__28); +l_Lean_Elab_Term_expandDbgTrace___closed__29 = _init_l_Lean_Elab_Term_expandDbgTrace___closed__29(); +lean_mark_persistent(l_Lean_Elab_Term_expandDbgTrace___closed__29); l___regBuiltin_Lean_Elab_Term_expandDbgTrace___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_expandDbgTrace___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_expandDbgTrace___closed__1); l___regBuiltin_Lean_Elab_Term_expandDbgTrace___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_expandDbgTrace___closed__2(); lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_expandDbgTrace___closed__2); l___regBuiltin_Lean_Elab_Term_expandDbgTrace___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_expandDbgTrace___closed__3(); lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_expandDbgTrace___closed__3); -l___regBuiltin_Lean_Elab_Term_expandDbgTrace___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_expandDbgTrace___closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_expandDbgTrace___closed__4); res = l___regBuiltin_Lean_Elab_Term_expandDbgTrace(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); @@ -20834,6 +21171,12 @@ l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__4 = _init_l_Lean_Elab_Term_ lean_mark_persistent(l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__4); l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__5 = _init_l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__5(); lean_mark_persistent(l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__5); +l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__6 = _init_l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__6(); +lean_mark_persistent(l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__6); +l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__7 = _init_l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__7(); +lean_mark_persistent(l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__7); +l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__8 = _init_l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__8(); +lean_mark_persistent(l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__8); l_Lean_Elab_Term_expandParen___closed__1 = _init_l_Lean_Elab_Term_expandParen___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_expandParen___closed__1); l_Lean_Elab_Term_expandParen___closed__2 = _init_l_Lean_Elab_Term_expandParen___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/BuiltinTerm.c b/stage0/stdlib/Lean/Elab/BuiltinTerm.c index 5491ee85b99..fbe301439b5 100644 --- a/stage0/stdlib/Lean/Elab/BuiltinTerm.c +++ b/stage0/stdlib/Lean/Elab/BuiltinTerm.c @@ -62,6 +62,7 @@ static lean_object* l_Lean_Elab_Term_elabLetMVar___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_elabTypeStx_declRange___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_elabProp___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Term_elabEnsureTypeOf___closed__4; +lean_object* l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabProp___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Open_0__Lean_Elab_OpenDecl_elabOpenScoped___at_Lean_Elab_Term_elabOpen___spec__27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabSetOption_docString(lean_object*); @@ -133,7 +134,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabProp_declRange___closed__2 uint8_t l_Lean_MetavarContext_isDelayedAssigned(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabPipeCompletion_declRange___closed__5; static lean_object* l_Lean_Elab_Term_elabScientificLit___closed__7; -lean_object* l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_registerMVarErrorHoleInfo(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_Term_elabTypeStx___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabEnsureExpectedType_declRange___closed__3; @@ -279,6 +279,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabTypeOf_declRange(lean lean_object* l_Lean_Elab_Term_elabTermEnsuringType(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabOpen___spec__16(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabStrLit_declRange___closed__1; +lean_object* l_Lean_Syntax_isNatLit_x3f(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabEnsureExpectedType___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabOpen___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_elabNoImplicitLambda_declRange___closed__4; @@ -300,7 +301,6 @@ lean_object* l_Lean_getOptionDecl(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabPipeCompletion___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabOpen___spec__25___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_resolveUniqueNamespace___at_Lean_Elab_Term_elabOpen___spec__5___closed__1; -extern lean_object* l_Lean_numLitKind; static lean_object* l___regBuiltin_Lean_Elab_Term_elabSetOption_declRange___closed__1; LEAN_EXPORT lean_object* l_Lean_resolveUniqueNamespace___at_Lean_Elab_Term_elabOpen___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_resolveNamespace___at_Lean_Elab_Term_elabOpen___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -506,7 +506,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabEnsureTypeOf___boxed(lean_object*, static lean_object* l_Lean_Elab_Term_elabWaitIfContainsMVar___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_elabLetMVar___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabStrLit_declRange(lean_object*); -lean_object* l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabOpen___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_BuiltinTerm_0__Lean_Elab_Term_getMVarFromUserName(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_resolveGlobalConst___at_Lean_Elab_Term_elabDoubleQuotedName___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -6538,42 +6537,41 @@ return x_46; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabNumLit(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_object* x_9) { _start: { -lean_object* x_10; lean_object* x_11; -x_10 = l_Lean_numLitKind; -x_11 = l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(x_10, x_1); -if (lean_obj_tag(x_11) == 0) +lean_object* x_10; +x_10 = l_Lean_Syntax_isNatLit_x3f(x_1); +if (lean_obj_tag(x_10) == 0) { -lean_object* x_12; uint8_t x_13; +lean_object* x_11; uint8_t x_12; lean_dec(x_2); lean_dec(x_1); -x_12 = l_Lean_Elab_throwIllFormedSyntax___at_Lean_Elab_Term_elabNumLit___spec__1(x_3, x_4, x_5, x_6, x_7, x_8, x_9); -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) +x_11 = l_Lean_Elab_throwIllFormedSyntax___at_Lean_Elab_Term_elabNumLit___spec__1(x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) { -return x_12; +return x_11; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_12, 0); -x_15 = lean_ctor_get(x_12, 1); -lean_inc(x_15); +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_11, 1); lean_inc(x_14); -lean_dec(x_12); -x_16 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_16, 0, x_14); -lean_ctor_set(x_16, 1, x_15); -return x_16; +lean_inc(x_13); +lean_dec(x_11); +x_15 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_14); +return x_15; } } else { -lean_object* x_17; lean_object* x_18; -x_17 = lean_ctor_get(x_11, 0); -lean_inc(x_17); -lean_dec(x_11); -x_18 = l_Lean_Elab_Term_elabNumLit___lambda__1(x_2, x_1, x_17, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_18; +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_10, 0); +lean_inc(x_16); +lean_dec(x_10); +x_17 = l_Lean_Elab_Term_elabNumLit___lambda__1(x_2, x_1, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_17; } } } @@ -6751,30 +6749,29 @@ return x_4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabRawNatLit(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_object* x_9) { _start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +lean_object* x_10; lean_object* x_11; lean_object* x_12; x_10 = lean_unsigned_to_nat(1u); x_11 = l_Lean_Syntax_getArg(x_1, x_10); -x_12 = l_Lean_numLitKind; -x_13 = l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(x_12, x_11); +x_12 = l_Lean_Syntax_isNatLit_x3f(x_11); lean_dec(x_11); -if (lean_obj_tag(x_13) == 0) +if (lean_obj_tag(x_12) == 0) { -lean_object* x_14; -x_14 = l_Lean_Elab_throwIllFormedSyntax___at_Lean_Elab_Term_elabStrLit___spec__1(x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_14; +lean_object* x_13; +x_13 = l_Lean_Elab_throwIllFormedSyntax___at_Lean_Elab_Term_elabStrLit___spec__1(x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_13; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_dec(x_3); -x_15 = lean_ctor_get(x_13, 0); -lean_inc(x_15); -lean_dec(x_13); -x_16 = l_Lean_mkRawNatLit(x_15); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_9); -return x_17; +x_14 = lean_ctor_get(x_12, 0); +lean_inc(x_14); +lean_dec(x_12); +x_15 = l_Lean_mkRawNatLit(x_14); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_9); +return x_16; } } } @@ -8051,7 +8048,7 @@ x_11 = lean_ctor_get(x_9, 1); lean_inc(x_11); lean_dec(x_9); x_12 = lean_box(0); -x_13 = l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(x_10, x_12); +x_13 = l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(x_10, x_12); x_14 = l_List_isEmpty___rarg(x_13); if (x_14 == 0) { @@ -10933,7 +10930,7 @@ x_12 = lean_ctor_get(x_10, 1); lean_inc(x_12); lean_dec(x_10); x_13 = lean_box(0); -x_14 = l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(x_11, x_13); +x_14 = l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(x_11, x_13); x_15 = l_List_isEmpty___rarg(x_14); if (x_15 == 0) { @@ -15764,112 +15761,111 @@ x_21 = lean_erase_macro_scopes(x_20); x_22 = l_Lean_Syntax_isStrLit_x3f(x_2); if (lean_obj_tag(x_22) == 0) { -lean_object* x_23; lean_object* x_24; -x_23 = l_Lean_numLitKind; -x_24 = l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(x_23, x_2); -if (lean_obj_tag(x_24) == 0) +lean_object* x_23; +x_23 = l_Lean_Syntax_isNatLit_x3f(x_2); +if (lean_obj_tag(x_23) == 0) { if (lean_obj_tag(x_2) == 2) { -lean_object* x_25; lean_object* x_26; uint8_t x_27; -x_25 = lean_ctor_get(x_2, 1); -lean_inc(x_25); -x_26 = l_Lean_Elab_Term_elabScientificLit___closed__10; -x_27 = lean_string_dec_eq(x_25, x_26); -if (x_27 == 0) +lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_24 = lean_ctor_get(x_2, 1); +lean_inc(x_24); +x_25 = l_Lean_Elab_Term_elabScientificLit___closed__10; +x_26 = lean_string_dec_eq(x_24, x_25); +if (x_26 == 0) { -lean_object* x_28; uint8_t x_29; -x_28 = l_Lean_Elab_Term_elabScientificLit___closed__7; -x_29 = lean_string_dec_eq(x_25, x_28); -lean_dec(x_25); -if (x_29 == 0) +lean_object* x_27; uint8_t x_28; +x_27 = l_Lean_Elab_Term_elabScientificLit___closed__7; +x_28 = lean_string_dec_eq(x_24, x_27); +lean_dec(x_24); +if (x_28 == 0) { -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_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_dec(x_21); -x_30 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_30, 0, x_2); -x_31 = l_Lean_Elab_elabSetOption___at_Lean_Elab_Term_elabSetOption___spec__1___closed__2; -x_32 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_30); -x_33 = l_Lean_Elab_Term_elabSyntheticHole___closed__9; -x_34 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_34, 0, x_32); -lean_ctor_set(x_34, 1, x_33); -x_35 = l_Lean_throwError___at_Lean_Elab_Term_elabSetOption___spec__2(x_34, x_3, x_4, x_5, x_6, x_7, x_8, x_19); +x_29 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_29, 0, x_2); +x_30 = l_Lean_Elab_elabSetOption___at_Lean_Elab_Term_elabSetOption___spec__1___closed__2; +x_31 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_29); +x_32 = l_Lean_Elab_Term_elabSyntheticHole___closed__9; +x_33 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +x_34 = l_Lean_throwError___at_Lean_Elab_Term_elabSetOption___spec__2(x_33, x_3, x_4, x_5, x_6, x_7, x_8, x_19); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -return x_35; +return x_34; } else { -lean_object* x_36; lean_object* x_37; +lean_object* x_35; lean_object* x_36; lean_dec(x_2); -x_36 = l_Lean_Elab_elabSetOption___at_Lean_Elab_Term_elabSetOption___spec__1___closed__3; -x_37 = l_Lean_Elab_elabSetOption_setOption___at_Lean_Elab_Term_elabSetOption___spec__3(x_21, x_36, x_3, x_4, x_5, x_6, x_7, x_8, x_19); -return x_37; +x_35 = l_Lean_Elab_elabSetOption___at_Lean_Elab_Term_elabSetOption___spec__1___closed__3; +x_36 = l_Lean_Elab_elabSetOption_setOption___at_Lean_Elab_Term_elabSetOption___spec__3(x_21, x_35, x_3, x_4, x_5, x_6, x_7, x_8, x_19); +return x_36; } } else { -lean_object* x_38; lean_object* x_39; -lean_dec(x_25); +lean_object* x_37; lean_object* x_38; +lean_dec(x_24); lean_dec(x_2); -x_38 = l_Lean_Elab_elabSetOption___at_Lean_Elab_Term_elabSetOption___spec__1___closed__4; -x_39 = l_Lean_Elab_elabSetOption_setOption___at_Lean_Elab_Term_elabSetOption___spec__3(x_21, x_38, x_3, x_4, x_5, x_6, x_7, x_8, x_19); -return x_39; +x_37 = l_Lean_Elab_elabSetOption___at_Lean_Elab_Term_elabSetOption___spec__1___closed__4; +x_38 = l_Lean_Elab_elabSetOption_setOption___at_Lean_Elab_Term_elabSetOption___spec__3(x_21, x_37, x_3, x_4, x_5, x_6, x_7, x_8, x_19); +return x_38; } } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_dec(x_21); -x_40 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_40, 0, x_2); -x_41 = l_Lean_Elab_elabSetOption___at_Lean_Elab_Term_elabSetOption___spec__1___closed__2; -x_42 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_40); -x_43 = l_Lean_Elab_Term_elabSyntheticHole___closed__9; -x_44 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_44, 1, x_43); -x_45 = l_Lean_throwError___at_Lean_Elab_Term_elabSetOption___spec__2(x_44, x_3, x_4, x_5, x_6, x_7, x_8, x_19); +x_39 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_39, 0, x_2); +x_40 = l_Lean_Elab_elabSetOption___at_Lean_Elab_Term_elabSetOption___spec__1___closed__2; +x_41 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_39); +x_42 = l_Lean_Elab_Term_elabSyntheticHole___closed__9; +x_43 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +x_44 = l_Lean_throwError___at_Lean_Elab_Term_elabSetOption___spec__2(x_43, x_3, x_4, x_5, x_6, x_7, x_8, x_19); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -return x_45; +return x_44; } } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_dec(x_2); -x_46 = lean_ctor_get(x_24, 0); -lean_inc(x_46); -lean_dec(x_24); -x_47 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_47, 0, x_46); -x_48 = l_Lean_Elab_elabSetOption_setOption___at_Lean_Elab_Term_elabSetOption___spec__3(x_21, x_47, x_3, x_4, x_5, x_6, x_7, x_8, x_19); -return x_48; +x_45 = lean_ctor_get(x_23, 0); +lean_inc(x_45); +lean_dec(x_23); +x_46 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_46, 0, x_45); +x_47 = l_Lean_Elab_elabSetOption_setOption___at_Lean_Elab_Term_elabSetOption___spec__3(x_21, x_46, x_3, x_4, x_5, x_6, x_7, x_8, x_19); +return x_47; } } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_dec(x_2); -x_49 = lean_ctor_get(x_22, 0); -lean_inc(x_49); +x_48 = lean_ctor_get(x_22, 0); +lean_inc(x_48); lean_dec(x_22); -x_50 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_50, 0, x_49); -x_51 = l_Lean_Elab_elabSetOption_setOption___at_Lean_Elab_Term_elabSetOption___spec__3(x_21, x_50, x_3, x_4, x_5, x_6, x_7, x_8, x_19); -return x_51; +x_49 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_49, 0, x_48); +x_50 = l_Lean_Elab_elabSetOption_setOption___at_Lean_Elab_Term_elabSetOption___spec__3(x_21, x_49, x_3, x_4, x_5, x_6, x_7, x_8, x_19); +return x_50; } } } diff --git a/stage0/stdlib/Lean/Elab/Command.c b/stage0/stdlib/Lean/Elab/Command.c index bb9e4fa92c3..1dee4fb86e9 100644 --- a/stage0/stdlib/Lean/Elab/Command.c +++ b/stage0/stdlib/Lean/Elab/Command.c @@ -59,7 +59,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommand___boxed__const__1; LEAN_EXPORT lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instAddMessageContextCommandElabM___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_getInfoTrees___at_Lean_Elab_Command_elabCommandTopLevel___spec__1___boxed(lean_object*); -extern lean_object* l_Lean_nullKind; LEAN_EXPORT lean_object* l_Lean_Elab_Command_getScope___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabCommand___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadMacroAdapterCommandElabM___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); @@ -246,7 +245,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadResolveNameCommandElabM___ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOptionM___at_Lean_Elab_Command_elabCommand___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadRecDepthCommandElabM___lambda__3(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2106____closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runTermElabM___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Command_expandDeclId___spec__2___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -255,6 +253,7 @@ static lean_object* l_Lean_Elab_Command_getBracketedBinderIds___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_getInfoTrees___at_Lean_Elab_Command_elabCommandTopLevel___spec__1___rarg___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_expandDeclId___closed__1; +static lean_object* l_Lean_Elab_Command_elabCommand___closed__6; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Command_0__Lean_Elab_Command_mkTermContext___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentArray_mapMAux___at_Lean_Elab_Command_liftTermElabM___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -271,22 +270,25 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_runLinters(lean_object*, lean_objec LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOptionM___at_Lean_Elab_Command_elabCommand___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Std_Data_PersistentArray_0__Std_PersistentArray_foldlFromMAux___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*); +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__11___lambda__1___closed__2; static lean_object* l_Lean_addDocString___at_Lean_Elab_Command_expandDeclId___spec__11___closed__1; lean_object* l_Lean_Elab_expandMacroImpl_x3f(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_mkState___closed__1; lean_object* lean_st_ref_take(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2109____closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadInfoTreeCommandElabM___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instMonadEnvCommandElabM___closed__2; LEAN_EXPORT lean_object* l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___closed__3; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2226____closed__2; static lean_object* l_Lean_Elab_Command_instMonadResolveNameCommandElabM___closed__2; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2109____closed__1; static lean_object* l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1___closed__2; LEAN_EXPORT lean_object* l___private_Std_Data_PersistentArray_0__Std_PersistentArray_foldlMAux___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__3(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instInhabitedState___closed__1; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2223____closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadCommandElabM___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__9; LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -294,7 +296,7 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Comm LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentArray_mapM___at_Lean_Elab_Command_liftTermElabM___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_modifyScope___closed__4; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2483____closed__2; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2486____closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__6(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadTraceCommandElabM___lambda__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_getBracketedBinderIds___closed__3; @@ -302,11 +304,11 @@ static lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Command_expandDeclId__ static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_withFreshMacroScope(lean_object*); lean_object* l_Lean_Elab_mkMessageCore(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2223____closed__1; lean_object* l_Lean_getStructureFieldsFlattened(lean_object*, lean_object*, uint8_t); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2226____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadQuotationCommandElabM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2486____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandDeclId(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2483____closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkTermState(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Command_0__Lean_Elab_Command_mkTermContext___spec__2(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Command_getBracketedBinderIds___closed__4; @@ -316,7 +318,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_Scope_varDecls___default; static lean_object* l_Lean_Elab_Command_elabCommand___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_Scope_openDecls___default; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__5___closed__1; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2106____closed__4; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__7(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Command_elabCommand___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_State_nextInstIdx___default; @@ -344,7 +345,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftCoreM___rarg___boxed(lean_objec LEAN_EXPORT lean_object* l_Lean_logTrace___at_Lean_Elab_Command_elabCommand___spec__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabCommand___lambda__3___closed__3; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__11___lambda__1___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2106____closed__1; static lean_object* l_Lean_Elab_Command_liftCoreM___rarg___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__5___boxed(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_Command_instMonadTraceCommandElabM; @@ -361,14 +361,15 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_addUnivLevel___lambda__1(lean_objec LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instInhabitedState___closed__6; lean_object* lean_st_mk_ref(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2109____closed__4; static lean_object* l_Lean_Elab_Command_instInhabitedState___closed__13; lean_object* l_Lean_Core_getMaxHeartbeats(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_expandDeclId___spec__14(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_414_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2106_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_1745_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2223_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2483_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_417_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2109_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_1748_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2226_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2486_(lean_object*); lean_object* l_Lean_Syntax_getId(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_State_ngen___default; extern lean_object* l_Lean_Expr_instHashableExpr; @@ -399,7 +400,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_Context_currRecDepth___default; LEAN_EXPORT lean_object* l_Lean_Elab_withInfoTreeContext___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__11(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_liftAttrM(lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2106____closed__3; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__10___boxed(lean_object*, 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___private_Lean_Elab_Command_0__Lean_Elab_Command_liftAttrM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__2___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -438,6 +438,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadInfoTreeCommandElabM___lam LEAN_EXPORT lean_object* l_Lean_Elab_Command_modifyScope___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ResolveName_resolveGlobalName(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkCoreContext___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2109____closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Command_Scope_levelNames___default; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandDeclId___spec__9(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_modifyScope___closed__1; @@ -471,7 +472,6 @@ static lean_object* l_Lean_Elab_Command_getBracketedBinderIds___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Command_withScope___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_addUnivLevel(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__5(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instMonadRecDepthCommandElabM___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadTraceCommandElabM___lambda__2___boxed(lean_object*, lean_object*, lean_object*); @@ -700,6 +700,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_Scope_varUIds___default; LEAN_EXPORT lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Command_expandDeclId___spec__7___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Command_expandDeclId___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Command_expandDeclId___spec__2___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_elabCommand___closed__5; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandDeclId___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Command_expandDeclId___spec__2___lambda__1___closed__2; @@ -1469,7 +1470,7 @@ lean_ctor_set(x_18, 8, x_17); return x_18; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_414_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_417_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; @@ -7767,7 +7768,7 @@ x_8 = l_Lean_Elab_mkElabAttribute___rarg(x_2, x_3, x_4, x_5, x_6, x_7, x_1); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_1745_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_1748_(lean_object* x_1) { _start: { lean_object* x_2; @@ -9236,7 +9237,7 @@ lean_dec(x_1); return x_4; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2106____closed__1() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2109____closed__1() { _start: { lean_object* x_1; @@ -9244,17 +9245,17 @@ x_1 = lean_mk_string_from_bytes("showPartialSyntaxErrors", 23); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2106____closed__2() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2109____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2106____closed__1; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2109____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2106____closed__3() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2109____closed__3() { _start: { lean_object* x_1; @@ -9262,13 +9263,13 @@ x_1 = lean_mk_string_from_bytes("show elaboration errors from partial syntax tre return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2106____closed__4() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2109____closed__4() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = 0; x_2 = l_Lean_Elab_Command_instInhabitedScope___closed__1; -x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2106____closed__3; +x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2109____closed__3; x_4 = lean_box(x_1); x_5 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_5, 0, x_4); @@ -9277,12 +9278,12 @@ lean_ctor_set(x_5, 2, x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2106_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2109_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2106____closed__2; -x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2106____closed__4; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2109____closed__2; +x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2109____closed__4; x_4 = l_Lean_Option_register___at_Lean_Elab_initFn____x40_Lean_Elab_AutoBound___hyg_6____spec__1(x_2, x_3, x_1); return x_4; } @@ -9630,7 +9631,7 @@ return x_86; } } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2223____closed__1() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2226____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -9640,21 +9641,21 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2223____closed__2() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2226____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2223____closed__1; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2226____closed__1; x_2 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__18; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2223_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2226_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2223____closed__2; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2226____closed__2; x_3 = l_Lean_registerTraceClass(x_2, x_1); return x_3; } @@ -11740,6 +11741,24 @@ static lean_object* _init_l_Lean_Elab_Command_elabCommand___closed__4() { _start: { lean_object* x_1; +x_1 = lean_mk_string_from_bytes("null", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Command_elabCommand___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Command_elabCommand___closed__4; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Command_elabCommand___closed__6() { +_start: +{ +lean_object* x_1; x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand___lambda__4___boxed), 3, 0); return x_1; } @@ -11763,13 +11782,13 @@ x_5 = lean_ctor_get(x_1, 1); lean_inc(x_5); x_6 = lean_ctor_get(x_1, 2); lean_inc(x_6); -x_7 = l_Lean_nullKind; +x_7 = l_Lean_Elab_Command_elabCommand___closed__5; x_8 = lean_name_eq(x_5, x_7); if (x_8 == 0) { lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_dec(x_6); -x_9 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2223____closed__2; +x_9 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2226____closed__2; lean_inc(x_1); x_10 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand___lambda__3), 6, 3); lean_closure_set(x_10, 0, x_1); @@ -11793,7 +11812,7 @@ if (x_15 == 0) lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_dec(x_13); lean_dec(x_6); -x_16 = l_Lean_Elab_Command_elabCommand___closed__4; +x_16 = l_Lean_Elab_Command_elabCommand___closed__6; x_17 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand___lambda__1), 5, 2); lean_closure_set(x_17, 0, x_1); lean_closure_set(x_17, 1, x_16); @@ -11809,7 +11828,7 @@ if (x_19 == 0) lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_dec(x_13); lean_dec(x_6); -x_20 = l_Lean_Elab_Command_elabCommand___closed__4; +x_20 = l_Lean_Elab_Command_elabCommand___closed__6; x_21 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand___lambda__1), 5, 2); lean_closure_set(x_21, 0, x_1); lean_closure_set(x_21, 1, x_20); @@ -12008,7 +12027,7 @@ lean_dec(x_1); return x_4; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2483____closed__1() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2486____closed__1() { _start: { lean_object* x_1; @@ -12016,21 +12035,21 @@ x_1 = lean_mk_string_from_bytes("input", 5); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2483____closed__2() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2486____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2223____closed__1; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2483____closed__1; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2226____closed__1; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2486____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2483_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2486_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2483____closed__2; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2486____closed__2; x_3 = l_Lean_registerTraceClass(x_2, x_1); return x_3; } @@ -14508,7 +14527,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommandTopLevel(lean_object* x_ _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; -x_5 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2483____closed__2; +x_5 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2486____closed__2; x_6 = l_Lean_Elab_Command_getRef(x_2, x_3, x_4); x_7 = lean_ctor_get(x_6, 0); lean_inc(x_7); @@ -14566,7 +14585,7 @@ goto block_23; if (x_12 == 0) { lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2223____closed__1; +x_14 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2226____closed__1; x_15 = lean_box(0); x_16 = l_Lean_Elab_Command_elabCommandTopLevel___lambda__2(x_1, x_14, x_15, x_2, x_3, x_13); return x_16; @@ -14583,7 +14602,7 @@ lean_inc(x_19); x_20 = lean_ctor_get(x_18, 1); lean_inc(x_20); lean_dec(x_18); -x_21 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2223____closed__1; +x_21 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2226____closed__1; x_22 = l_Lean_Elab_Command_elabCommandTopLevel___lambda__2(x_1, x_21, x_19, x_2, x_3, x_20); return x_22; } @@ -14658,7 +14677,7 @@ goto block_54; if (x_43 == 0) { lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_45 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2223____closed__1; +x_45 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2226____closed__1; x_46 = lean_box(0); x_47 = l_Lean_Elab_Command_elabCommandTopLevel___lambda__2(x_1, x_45, x_46, x_42, x_3, x_44); return x_47; @@ -14675,7 +14694,7 @@ lean_inc(x_50); x_51 = lean_ctor_get(x_49, 1); lean_inc(x_51); lean_dec(x_49); -x_52 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2223____closed__1; +x_52 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2226____closed__1; x_53 = l_Lean_Elab_Command_elabCommandTopLevel___lambda__2(x_1, x_52, x_50, x_42, x_3, x_51); return x_53; } @@ -15256,167 +15275,164 @@ return x_8; } else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; x_9 = lean_unsigned_to_nat(1u); x_10 = l_Lean_Syntax_getArg(x_1, x_9); lean_dec(x_1); -x_11 = l_Lean_nullKind; -x_12 = lean_unsigned_to_nat(2u); +x_11 = lean_unsigned_to_nat(2u); lean_inc(x_10); -x_13 = l_Lean_Syntax_isNodeOf(x_10, x_11, x_12); -if (x_13 == 0) +x_12 = l_Lean_Syntax_matchesNull(x_10, x_11); +if (x_12 == 0) { -lean_object* x_14; uint8_t x_15; -x_14 = lean_unsigned_to_nat(0u); -x_15 = l_Lean_Syntax_isNodeOf(x_10, x_11, x_14); -if (x_15 == 0) +lean_object* x_13; uint8_t x_14; +x_13 = lean_unsigned_to_nat(0u); +x_14 = l_Lean_Syntax_matchesNull(x_10, x_13); +if (x_14 == 0) { -lean_object* x_16; -x_16 = l_Lean_Elab_Command_Scope_varDecls___default___closed__1; -return x_16; +lean_object* x_15; +x_15 = l_Lean_Elab_Command_Scope_varDecls___default___closed__1; +return x_15; } else { -lean_object* x_17; -x_17 = l_Lean_Elab_Command_getBracketedBinderIds___closed__10; -return x_17; +lean_object* x_16; +x_16 = l_Lean_Elab_Command_getBracketedBinderIds___closed__10; +return x_16; } } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_18 = lean_unsigned_to_nat(0u); -x_19 = l_Lean_Syntax_getArg(x_10, x_18); +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_17 = lean_unsigned_to_nat(0u); +x_18 = l_Lean_Syntax_getArg(x_10, x_17); lean_dec(x_10); -x_20 = l_Lean_Syntax_getId(x_19); -lean_dec(x_19); -x_21 = l_Lean_Elab_Command_getBracketedBinderIds___closed__9; -x_22 = lean_array_push(x_21, x_20); -return x_22; +x_19 = l_Lean_Syntax_getId(x_18); +lean_dec(x_18); +x_20 = l_Lean_Elab_Command_getBracketedBinderIds___closed__9; +x_21 = lean_array_push(x_20, x_19); +return x_21; } } } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_35; -x_23 = lean_unsigned_to_nat(1u); -x_24 = l_Lean_Syntax_getArg(x_1, x_23); -x_25 = lean_unsigned_to_nat(2u); -x_26 = l_Lean_Syntax_getArg(x_1, x_25); +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_34; +x_22 = lean_unsigned_to_nat(1u); +x_23 = l_Lean_Syntax_getArg(x_1, x_22); +x_24 = lean_unsigned_to_nat(2u); +x_25 = l_Lean_Syntax_getArg(x_1, x_24); lean_dec(x_1); -x_35 = l_Lean_Syntax_isNone(x_26); -if (x_35 == 0) +x_34 = l_Lean_Syntax_isNone(x_25); +if (x_34 == 0) { -lean_object* x_36; uint8_t x_37; -x_36 = l_Lean_nullKind; -lean_inc(x_26); -x_37 = l_Lean_Syntax_isNodeOf(x_26, x_36, x_25); -if (x_37 == 0) +uint8_t x_35; +lean_inc(x_25); +x_35 = l_Lean_Syntax_matchesNull(x_25, x_24); +if (x_35 == 0) { -lean_object* x_38; -lean_dec(x_26); -lean_dec(x_24); -x_38 = l_Lean_Elab_Command_Scope_varDecls___default___closed__1; -return x_38; +lean_object* x_36; +lean_dec(x_25); +lean_dec(x_23); +x_36 = l_Lean_Elab_Command_Scope_varDecls___default___closed__1; +return x_36; } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_39 = l_Lean_Syntax_getArg(x_26, x_23); -lean_dec(x_26); -x_40 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_40, 0, x_39); -x_41 = lean_box(0); -x_27 = x_41; -x_28 = x_40; -goto block_34; +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = l_Lean_Syntax_getArg(x_25, x_22); +lean_dec(x_25); +x_38 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_38, 0, x_37); +x_39 = lean_box(0); +x_26 = x_39; +x_27 = x_38; +goto block_33; } } else { -lean_object* x_42; lean_object* x_43; -lean_dec(x_26); -x_42 = lean_box(0); -x_43 = lean_box(0); -x_27 = x_43; -x_28 = x_42; -goto block_34; +lean_object* x_40; lean_object* x_41; +lean_dec(x_25); +x_40 = lean_box(0); +x_41 = lean_box(0); +x_26 = x_41; +x_27 = x_40; +goto block_33; } -block_34: +block_33: { -lean_object* x_29; lean_object* x_30; size_t x_31; size_t x_32; lean_object* x_33; -lean_dec(x_28); +lean_object* x_28; lean_object* x_29; size_t x_30; size_t x_31; lean_object* x_32; lean_dec(x_27); -x_29 = l_Lean_Syntax_getArgs(x_24); -lean_dec(x_24); -x_30 = lean_array_get_size(x_29); -x_31 = lean_usize_of_nat(x_30); -lean_dec(x_30); -x_32 = 0; -x_33 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_getBracketedBinderIds___spec__1(x_31, x_32, x_29); -return x_33; +lean_dec(x_26); +x_28 = l_Lean_Syntax_getArgs(x_23); +lean_dec(x_23); +x_29 = lean_array_get_size(x_28); +x_30 = lean_usize_of_nat(x_29); +lean_dec(x_29); +x_31 = 0; +x_32 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_getBracketedBinderIds___spec__1(x_30, x_31, x_28); +return x_32; } } } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_56; -x_44 = lean_unsigned_to_nat(1u); +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_54; +x_42 = lean_unsigned_to_nat(1u); +x_43 = l_Lean_Syntax_getArg(x_1, x_42); +x_44 = lean_unsigned_to_nat(2u); x_45 = l_Lean_Syntax_getArg(x_1, x_44); -x_46 = lean_unsigned_to_nat(2u); -x_47 = l_Lean_Syntax_getArg(x_1, x_46); lean_dec(x_1); -x_56 = l_Lean_Syntax_isNone(x_47); -if (x_56 == 0) +x_54 = l_Lean_Syntax_isNone(x_45); +if (x_54 == 0) { -lean_object* x_57; uint8_t x_58; -x_57 = l_Lean_nullKind; -lean_inc(x_47); -x_58 = l_Lean_Syntax_isNodeOf(x_47, x_57, x_46); -if (x_58 == 0) +uint8_t x_55; +lean_inc(x_45); +x_55 = l_Lean_Syntax_matchesNull(x_45, x_44); +if (x_55 == 0) { -lean_object* x_59; -lean_dec(x_47); +lean_object* x_56; lean_dec(x_45); -x_59 = l_Lean_Elab_Command_Scope_varDecls___default___closed__1; -return x_59; +lean_dec(x_43); +x_56 = l_Lean_Elab_Command_Scope_varDecls___default___closed__1; +return x_56; } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = l_Lean_Syntax_getArg(x_47, x_44); -lean_dec(x_47); -x_61 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_61, 0, x_60); -x_62 = lean_box(0); -x_48 = x_62; -x_49 = x_61; -goto block_55; +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = l_Lean_Syntax_getArg(x_45, x_42); +lean_dec(x_45); +x_58 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_58, 0, x_57); +x_59 = lean_box(0); +x_46 = x_59; +x_47 = x_58; +goto block_53; } } else { -lean_object* x_63; lean_object* x_64; -lean_dec(x_47); -x_63 = lean_box(0); -x_64 = lean_box(0); -x_48 = x_64; -x_49 = x_63; -goto block_55; +lean_object* x_60; lean_object* x_61; +lean_dec(x_45); +x_60 = lean_box(0); +x_61 = lean_box(0); +x_46 = x_61; +x_47 = x_60; +goto block_53; } -block_55: +block_53: { -lean_object* x_50; lean_object* x_51; size_t x_52; size_t x_53; lean_object* x_54; +lean_object* x_48; lean_object* x_49; size_t x_50; size_t x_51; lean_object* x_52; +lean_dec(x_47); +lean_dec(x_46); +x_48 = l_Lean_Syntax_getArgs(x_43); +lean_dec(x_43); +x_49 = lean_array_get_size(x_48); +x_50 = lean_usize_of_nat(x_49); lean_dec(x_49); -lean_dec(x_48); -x_50 = l_Lean_Syntax_getArgs(x_45); -lean_dec(x_45); -x_51 = lean_array_get_size(x_50); -x_52 = lean_usize_of_nat(x_51); -lean_dec(x_51); -x_53 = 0; -x_54 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_getBracketedBinderIds___spec__1(x_52, x_53, x_50); -return x_54; +x_51 = 0; +x_52 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_getBracketedBinderIds___spec__1(x_50, x_51, x_48); +return x_52; } } } @@ -22631,7 +22647,7 @@ l_Lean_Elab_Command_instMonadCommandElabM = _init_l_Lean_Elab_Command_instMonadC lean_mark_persistent(l_Lean_Elab_Command_instMonadCommandElabM); l_Lean_Elab_Command_mkState___closed__1 = _init_l_Lean_Elab_Command_mkState___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_mkState___closed__1); -if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_414_(lean_io_mk_world()); +if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_417_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Elab_Command_lintersRef = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Elab_Command_lintersRef); @@ -22768,7 +22784,7 @@ l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__17 = _init_l_Lean_Ela lean_mark_persistent(l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__17); l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__18 = _init_l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__18(); lean_mark_persistent(l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__18); -if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_1745_(lean_io_mk_world()); +if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_1748_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Elab_Command_commandElabAttribute = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Elab_Command_commandElabAttribute); @@ -22801,15 +22817,15 @@ l_Lean_Elab_Command_instMonadRecDepthCommandElabM___closed__4 = _init_l_Lean_Ela lean_mark_persistent(l_Lean_Elab_Command_instMonadRecDepthCommandElabM___closed__4); l_Lean_Elab_Command_instMonadRecDepthCommandElabM = _init_l_Lean_Elab_Command_instMonadRecDepthCommandElabM(); lean_mark_persistent(l_Lean_Elab_Command_instMonadRecDepthCommandElabM); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2106____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2106____closed__1(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2106____closed__1); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2106____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2106____closed__2(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2106____closed__2); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2106____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2106____closed__3(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2106____closed__3); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2106____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2106____closed__4(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2106____closed__4); -if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2106_(lean_io_mk_world()); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2109____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2109____closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2109____closed__1); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2109____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2109____closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2109____closed__2); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2109____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2109____closed__3(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2109____closed__3); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2109____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2109____closed__4(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2109____closed__4); +if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2109_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Elab_Command_showPartialSyntaxErrors = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Elab_Command_showPartialSyntaxErrors); @@ -22818,11 +22834,11 @@ lean_dec_ref(res); lean_mark_persistent(l_Lean_Elab_Command_withLogging___closed__1); l_Lean_Elab_Command_withLogging___closed__2 = _init_l_Lean_Elab_Command_withLogging___closed__2(); lean_mark_persistent(l_Lean_Elab_Command_withLogging___closed__2); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2223____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2223____closed__1(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2223____closed__1); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2223____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2223____closed__2(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2223____closed__2); -res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2223_(lean_io_mk_world()); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2226____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2226____closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2226____closed__1); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2226____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2226____closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2226____closed__2); +res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2226_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__1___closed__1 = _init_l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__1___closed__1(); @@ -22861,13 +22877,17 @@ l_Lean_Elab_Command_elabCommand___closed__3 = _init_l_Lean_Elab_Command_elabComm lean_mark_persistent(l_Lean_Elab_Command_elabCommand___closed__3); l_Lean_Elab_Command_elabCommand___closed__4 = _init_l_Lean_Elab_Command_elabCommand___closed__4(); lean_mark_persistent(l_Lean_Elab_Command_elabCommand___closed__4); +l_Lean_Elab_Command_elabCommand___closed__5 = _init_l_Lean_Elab_Command_elabCommand___closed__5(); +lean_mark_persistent(l_Lean_Elab_Command_elabCommand___closed__5); +l_Lean_Elab_Command_elabCommand___closed__6 = _init_l_Lean_Elab_Command_elabCommand___closed__6(); +lean_mark_persistent(l_Lean_Elab_Command_elabCommand___closed__6); l_Lean_Elab_Command_elabCommand___boxed__const__1 = _init_l_Lean_Elab_Command_elabCommand___boxed__const__1(); lean_mark_persistent(l_Lean_Elab_Command_elabCommand___boxed__const__1); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2483____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2483____closed__1(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2483____closed__1); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2483____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2483____closed__2(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2483____closed__2); -res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2483_(lean_io_mk_world()); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2486____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2486____closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2486____closed__1); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2486____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2486____closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2486____closed__2); +res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2486_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__5___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__5___closed__1(); diff --git a/stage0/stdlib/Lean/Elab/DeclModifiers.c b/stage0/stdlib/Lean/Elab/DeclModifiers.c index e49feac6f8d..3de07fc6d24 100644 --- a/stage0/stdlib/Lean/Elab/DeclModifiers.c +++ b/stage0/stdlib/Lean/Elab/DeclModifiers.c @@ -33,7 +33,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_expandOptDocComment_x3f(lean_object*); static lean_object* l_List_mapTRAux___at_Lean_Elab_instToFormatModifiers___spec__1___closed__8; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_checkIfShadowingStructureField___spec__1___rarg___lambda__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_mkDeclName___rarg___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_nullKind; lean_object* l_Lean_addDocString_x27___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___rarg___lambda__3___closed__1; LEAN_EXPORT uint8_t l_Lean_Elab_Modifiers_isPartial(lean_object*); @@ -219,6 +218,7 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_checkIfShadowin LEAN_EXPORT lean_object* l_Lean_Elab_Visibility_toCtorIdx(uint8_t); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_checkIfShadowingStructureField___spec__1___rarg___closed__2; static lean_object* l_Lean_Elab_instToFormatModifiers___closed__7; +static lean_object* l_Lean_Elab_expandDeclIdCore___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_applyVisibility___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_string_length(lean_object*); lean_object* l_Lean_indentD(lean_object*); @@ -230,6 +230,7 @@ LEAN_EXPORT uint8_t l_Lean_Elab_Modifiers_isPrivate(lean_object*); static lean_object* l_Lean_Elab_expandDeclIdCore___closed__1; static lean_object* l_Lean_Elab_mkDeclName___rarg___closed__1; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_expandDeclIdCore___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_checkIfShadowingStructureField(lean_object*); static lean_object* l_Lean_Elab_expandOptDocComment_x3f___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_expandOptDocComment_x3f___rarg___boxed(lean_object*, lean_object*, lean_object*); @@ -3874,9 +3875,27 @@ return x_13; static lean_object* _init_l_Lean_Elab_expandDeclIdCore___closed__1() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("null", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_expandDeclIdCore___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_expandDeclIdCore___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_expandDeclIdCore___closed__3() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean_nullKind; +x_2 = l_Lean_Elab_expandDeclIdCore___closed__2; x_3 = l_Lean_Elab_Modifiers_attrs___default___closed__1; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -3908,7 +3927,7 @@ else { lean_object* x_9; lean_object* x_10; lean_object* x_11; x_9 = l_Lean_Syntax_getId(x_1); -x_10 = l_Lean_Elab_expandDeclIdCore___closed__1; +x_10 = l_Lean_Elab_expandDeclIdCore___closed__3; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_9); lean_ctor_set(x_11, 1, x_10); @@ -4542,6 +4561,10 @@ l_Lean_Elab_mkDeclName___rarg___closed__4 = _init_l_Lean_Elab_mkDeclName___rarg_ lean_mark_persistent(l_Lean_Elab_mkDeclName___rarg___closed__4); l_Lean_Elab_expandDeclIdCore___closed__1 = _init_l_Lean_Elab_expandDeclIdCore___closed__1(); lean_mark_persistent(l_Lean_Elab_expandDeclIdCore___closed__1); +l_Lean_Elab_expandDeclIdCore___closed__2 = _init_l_Lean_Elab_expandDeclIdCore___closed__2(); +lean_mark_persistent(l_Lean_Elab_expandDeclIdCore___closed__2); +l_Lean_Elab_expandDeclIdCore___closed__3 = _init_l_Lean_Elab_expandDeclIdCore___closed__3(); +lean_mark_persistent(l_Lean_Elab_expandDeclIdCore___closed__3); l_Lean_Elab_expandDeclId___rarg___closed__1 = _init_l_Lean_Elab_expandDeclId___rarg___closed__1(); lean_mark_persistent(l_Lean_Elab_expandDeclId___rarg___closed__1); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/Declaration.c b/stage0/stdlib/Lean/Elab/Declaration.c index f5f5a9227e7..03130939273 100644 --- a/stage0/stdlib/Lean/Elab/Declaration.c +++ b/stage0/stdlib/Lean/Elab/Declaration.c @@ -49,7 +49,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Command_expandMutualNamespace_declR LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabDeclaration___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_expandDeclNamespace_x3f___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabDeclaration___spec__5___rarg(lean_object*); -extern lean_object* l_Lean_nullKind; static lean_object* l_Lean_Elab_Command_elabDeclaration___closed__8; static lean_object* l_Lean_Elab_Command_expandInitCmd___closed__7; lean_object* l_Lean_Elab_Command_elabMutualDef(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -6960,7 +6959,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabDeclaration_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(167u); +x_1 = lean_unsigned_to_nat(168u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6972,7 +6971,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabDeclaration_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(188u); +x_1 = lean_unsigned_to_nat(189u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7000,7 +6999,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabDeclaration_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(167u); +x_1 = lean_unsigned_to_nat(168u); x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7012,7 +7011,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabDeclaration_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(167u); +x_1 = lean_unsigned_to_nat(168u); x_2 = lean_unsigned_to_nat(19u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8236,7 +8235,7 @@ lean_dec(x_13); lean_inc(x_1); x_23 = l_Lean_mkIdentFrom(x_1, x_22); x_24 = lean_box(2); -x_25 = l_Lean_nullKind; +x_25 = l_Lean_Elab_Command_elabDeclaration___closed__4; x_26 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_26, 0, x_24); lean_ctor_set(x_26, 1, x_25); @@ -8246,7 +8245,7 @@ x_28 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRul x_29 = !lean_is_exclusive(x_28); if (x_29 == 0) { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; x_30 = lean_ctor_get(x_28, 0); x_31 = l_Lean_Elab_Command_elabDeclaration___closed__5; lean_inc(x_30); @@ -8268,106 +8267,104 @@ lean_ctor_set(x_39, 0, x_30); lean_ctor_set(x_39, 1, x_38); x_40 = l_Lean_Elab_Command_elabInductive___closed__1; x_41 = lean_array_push(x_40, x_23); -x_42 = l_Lean_Elab_Command_elabDeclaration___closed__4; -x_43 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_43, 0, x_24); -lean_ctor_set(x_43, 1, x_42); -lean_ctor_set(x_43, 2, x_41); -x_44 = lean_array_push(x_33, x_39); -x_45 = lean_array_push(x_44, x_43); -x_46 = l_Lean_Elab_Command_elabDeclaration___closed__9; -x_47 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_47, 0, x_24); -lean_ctor_set(x_47, 1, x_46); -lean_ctor_set(x_47, 2, x_45); -x_48 = l_Lean_Elab_Command_elabDeclaration___closed__10; -x_49 = lean_array_push(x_48, x_37); -x_50 = lean_array_push(x_49, x_27); -x_51 = lean_array_push(x_50, x_47); -x_52 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_52, 0, x_24); -lean_ctor_set(x_52, 1, x_42); -lean_ctor_set(x_52, 2, x_51); -lean_ctor_set(x_28, 0, x_52); +x_42 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_42, 0, x_24); +lean_ctor_set(x_42, 1, x_25); +lean_ctor_set(x_42, 2, x_41); +x_43 = lean_array_push(x_33, x_39); +x_44 = lean_array_push(x_43, x_42); +x_45 = l_Lean_Elab_Command_elabDeclaration___closed__9; +x_46 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_46, 0, x_24); +lean_ctor_set(x_46, 1, x_45); +lean_ctor_set(x_46, 2, x_44); +x_47 = l_Lean_Elab_Command_elabDeclaration___closed__10; +x_48 = lean_array_push(x_47, x_37); +x_49 = lean_array_push(x_48, x_27); +x_50 = lean_array_push(x_49, x_46); +x_51 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_51, 0, x_24); +lean_ctor_set(x_51, 1, x_25); +lean_ctor_set(x_51, 2, x_50); +lean_ctor_set(x_28, 0, x_51); return x_28; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; -x_53 = lean_ctor_get(x_28, 0); -x_54 = lean_ctor_get(x_28, 1); -lean_inc(x_54); +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; +x_52 = lean_ctor_get(x_28, 0); +x_53 = lean_ctor_get(x_28, 1); lean_inc(x_53); +lean_inc(x_52); lean_dec(x_28); -x_55 = l_Lean_Elab_Command_elabDeclaration___closed__5; -lean_inc(x_53); -x_56 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_56, 0, x_53); -lean_ctor_set(x_56, 1, x_55); -x_57 = l_Lean_Elab_Command_elabDeclaration___closed__7; -x_58 = lean_array_push(x_57, x_56); +x_54 = l_Lean_Elab_Command_elabDeclaration___closed__5; +lean_inc(x_52); +x_55 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_55, 0, x_52); +lean_ctor_set(x_55, 1, x_54); +x_56 = l_Lean_Elab_Command_elabDeclaration___closed__7; +x_57 = lean_array_push(x_56, x_55); lean_inc(x_23); -x_59 = lean_array_push(x_58, x_23); -x_60 = l_Lean_Elab_Command_elabDeclaration___closed__6; -x_61 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_61, 0, x_24); -lean_ctor_set(x_61, 1, x_60); -lean_ctor_set(x_61, 2, x_59); -x_62 = l_Lean_Elab_Command_elabDeclaration___closed__8; -x_63 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_63, 0, x_53); -lean_ctor_set(x_63, 1, x_62); -x_64 = l_Lean_Elab_Command_elabInductive___closed__1; -x_65 = lean_array_push(x_64, x_23); -x_66 = l_Lean_Elab_Command_elabDeclaration___closed__4; -x_67 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_67, 0, x_24); -lean_ctor_set(x_67, 1, x_66); -lean_ctor_set(x_67, 2, x_65); -x_68 = lean_array_push(x_57, x_63); -x_69 = lean_array_push(x_68, x_67); -x_70 = l_Lean_Elab_Command_elabDeclaration___closed__9; -x_71 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_71, 0, x_24); -lean_ctor_set(x_71, 1, x_70); -lean_ctor_set(x_71, 2, x_69); -x_72 = l_Lean_Elab_Command_elabDeclaration___closed__10; -x_73 = lean_array_push(x_72, x_61); -x_74 = lean_array_push(x_73, x_27); -x_75 = lean_array_push(x_74, x_71); -x_76 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_76, 0, x_24); -lean_ctor_set(x_76, 1, x_66); -lean_ctor_set(x_76, 2, x_75); -x_77 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_77, 0, x_76); -lean_ctor_set(x_77, 1, x_54); -return x_77; +x_58 = lean_array_push(x_57, x_23); +x_59 = l_Lean_Elab_Command_elabDeclaration___closed__6; +x_60 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_60, 0, x_24); +lean_ctor_set(x_60, 1, x_59); +lean_ctor_set(x_60, 2, x_58); +x_61 = l_Lean_Elab_Command_elabDeclaration___closed__8; +x_62 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_62, 0, x_52); +lean_ctor_set(x_62, 1, x_61); +x_63 = l_Lean_Elab_Command_elabInductive___closed__1; +x_64 = lean_array_push(x_63, x_23); +x_65 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_65, 0, x_24); +lean_ctor_set(x_65, 1, x_25); +lean_ctor_set(x_65, 2, x_64); +x_66 = lean_array_push(x_56, x_62); +x_67 = lean_array_push(x_66, x_65); +x_68 = l_Lean_Elab_Command_elabDeclaration___closed__9; +x_69 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_69, 0, x_24); +lean_ctor_set(x_69, 1, x_68); +lean_ctor_set(x_69, 2, x_67); +x_70 = l_Lean_Elab_Command_elabDeclaration___closed__10; +x_71 = lean_array_push(x_70, x_60); +x_72 = lean_array_push(x_71, x_27); +x_73 = lean_array_push(x_72, x_69); +x_74 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_74, 0, x_24); +lean_ctor_set(x_74, 1, x_25); +lean_ctor_set(x_74, 2, x_73); +x_75 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_75, 0, x_74); +lean_ctor_set(x_75, 1, x_53); +return x_75; } } } else { -uint8_t x_78; +uint8_t x_76; lean_dec(x_2); lean_dec(x_1); -x_78 = !lean_is_exclusive(x_11); -if (x_78 == 0) +x_76 = !lean_is_exclusive(x_11); +if (x_76 == 0) { return x_11; } else { -lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_79 = lean_ctor_get(x_11, 0); -x_80 = lean_ctor_get(x_11, 1); -lean_inc(x_80); -lean_inc(x_79); +lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_77 = lean_ctor_get(x_11, 0); +x_78 = lean_ctor_get(x_11, 1); +lean_inc(x_78); +lean_inc(x_77); lean_dec(x_11); -x_81 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_81, 0, x_79); -lean_ctor_set(x_81, 1, x_80); -return x_81; +x_79 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_79, 0, x_77); +lean_ctor_set(x_79, 1, x_78); +return x_79; } } } @@ -8453,7 +8450,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualNamespace _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(232u); +x_1 = lean_unsigned_to_nat(233u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8465,7 +8462,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualNamespace _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(249u); +x_1 = lean_unsigned_to_nat(250u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8493,7 +8490,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualNamespace _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(232u); +x_1 = lean_unsigned_to_nat(233u); x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8505,7 +8502,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualNamespace _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(232u); +x_1 = lean_unsigned_to_nat(233u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8817,7 +8814,7 @@ x_23 = lean_ctor_get(x_12, 0); lean_inc(x_23); lean_dec(x_12); x_24 = lean_box(2); -x_25 = l_Lean_nullKind; +x_25 = l_Lean_Elab_Command_elabDeclaration___closed__4; x_26 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_26, 0, x_24); lean_ctor_set(x_26, 1, x_25); @@ -8836,7 +8833,7 @@ x_29 = lean_ctor_get(x_12, 0); lean_inc(x_29); lean_dec(x_12); x_30 = lean_box(2); -x_31 = l_Lean_nullKind; +x_31 = l_Lean_Elab_Command_elabDeclaration___closed__4; x_32 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_32, 0, x_30); lean_ctor_set(x_32, 1, x_31); @@ -8929,7 +8926,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualElement_d _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(252u); +x_1 = lean_unsigned_to_nat(253u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8941,7 +8938,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualElement_d _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(262u); +x_1 = lean_unsigned_to_nat(263u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8969,7 +8966,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualElement_d _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(252u); +x_1 = lean_unsigned_to_nat(253u); x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8981,7 +8978,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualElement_d _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(252u); +x_1 = lean_unsigned_to_nat(253u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9112,7 +9109,7 @@ x_25 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_25, 0, x_23); lean_ctor_set(x_25, 1, x_24); lean_ctor_set(x_25, 2, x_22); -x_26 = l_Lean_nullKind; +x_26 = l_Lean_Elab_Command_elabDeclaration___closed__4; x_27 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_27, 0, x_23); lean_ctor_set(x_27, 1, x_26); @@ -9229,7 +9226,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualPreamble_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(265u); +x_1 = lean_unsigned_to_nat(266u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9241,7 +9238,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualPreamble_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(272u); +x_1 = lean_unsigned_to_nat(273u); x_2 = lean_unsigned_to_nat(74u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9269,7 +9266,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualPreamble_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(265u); +x_1 = lean_unsigned_to_nat(266u); x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9281,7 +9278,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMutualPreamble_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(265u); +x_1 = lean_unsigned_to_nat(266u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9844,7 +9841,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabMutual_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(275u); +x_1 = lean_unsigned_to_nat(276u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9856,7 +9853,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabMutual_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(292u); +x_1 = lean_unsigned_to_nat(293u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9884,7 +9881,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabMutual_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(275u); +x_1 = lean_unsigned_to_nat(276u); x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9896,7 +9893,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabMutual_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(275u); +x_1 = lean_unsigned_to_nat(276u); x_2 = lean_unsigned_to_nat(14u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10858,7 +10855,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabAttr_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(295u); +x_1 = lean_unsigned_to_nat(296u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10870,7 +10867,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabAttr_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(312u); +x_1 = lean_unsigned_to_nat(313u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10898,7 +10895,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabAttr_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(295u); +x_1 = lean_unsigned_to_nat(296u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10910,7 +10907,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabAttr_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(295u); +x_1 = lean_unsigned_to_nat(296u); x_2 = lean_unsigned_to_nat(46u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13752,7 +13749,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandInitialize_decl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(338u); +x_1 = lean_unsigned_to_nat(339u); x_2 = lean_unsigned_to_nat(49u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13764,7 +13761,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandInitialize_decl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(339u); +x_1 = lean_unsigned_to_nat(340u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13792,7 +13789,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandInitialize_decl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(338u); +x_1 = lean_unsigned_to_nat(339u); x_2 = lean_unsigned_to_nat(53u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13804,7 +13801,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandInitialize_decl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(338u); +x_1 = lean_unsigned_to_nat(339u); x_2 = lean_unsigned_to_nat(69u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13919,7 +13916,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandBuiltinInitiali _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(341u); +x_1 = lean_unsigned_to_nat(342u); x_2 = lean_unsigned_to_nat(57u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13931,7 +13928,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandBuiltinInitiali _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(342u); +x_1 = lean_unsigned_to_nat(343u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13959,7 +13956,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandBuiltinInitiali _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(341u); +x_1 = lean_unsigned_to_nat(342u); x_2 = lean_unsigned_to_nat(61u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13971,7 +13968,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandBuiltinInitiali _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(341u); +x_1 = lean_unsigned_to_nat(342u); x_2 = lean_unsigned_to_nat(84u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); diff --git a/stage0/stdlib/Lean/Elab/DefView.c b/stage0/stdlib/Lean/Elab/DefView.c index fce6494899e..c24da4960b4 100644 --- a/stage0/stdlib/Lean/Elab/DefView.c +++ b/stage0/stdlib/Lean/Elab/DefView.c @@ -25,7 +25,6 @@ lean_object* lean_mk_empty_array_with_capacity(lean_object*); lean_object* l_Lean_mkSort(lean_object*); static lean_object* l_Lean_Elab_Command_mkInstanceName___lambda__1___closed__1; LEAN_EXPORT uint8_t l_Lean_Elab_DefKind_isDefOrAbbrevOrOpaque(uint8_t); -extern lean_object* l_Lean_nullKind; lean_object* l_Lean_Elab_Command_liftTermElabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_mkDefViewOfInstance___spec__5___rarg(lean_object*); lean_object* lean_name_mk_string(lean_object*, lean_object*); @@ -38,7 +37,9 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkDefViewOfInstance___lambda__1(lea extern lean_object* l_Lean_maxRecDepthErrorMessage; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_mkDefViewOfInstance___spec__5___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_mkDefViewOfInstance___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1485____closed__1; static lean_object* l_Lean_Elab_Command_mkDefViewOfInstance___closed__13; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1485____closed__4; static lean_object* l_Lean_Elab_Command_mkDefViewOfOpaque___closed__3; LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Command_mkDefViewOfInstance___spec__11(lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); @@ -55,6 +56,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkInstanceName___lambda__3___boxed( static lean_object* l_Lean_Elab_Command_mkDefViewOfAbbrev___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkInstanceName___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkDefViewOfOpaque(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1485____closed__3; lean_object* l_Lean_Elab_Term_synthesizeSyntheticMVars(uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_mkDefViewOfInstance___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkDefViewOfDef(lean_object*, lean_object*); @@ -68,6 +70,7 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_mkInstanceName static lean_object* l_Lean_Elab_Command_mkDefViewOfAbbrev___closed__2; static lean_object* l_Lean_Elab_Command_mkDefViewOfInstance___closed__10; static lean_object* l_Lean_Elab_Command_mkInstanceName___lambda__1___closed__4; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1485____closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_DefKind_noConfusion___rarg(uint8_t, uint8_t, lean_object*); static lean_object* l_Lean_Elab_Command_mkDefViewOfAbbrev___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkDefViewOfExample(lean_object*, lean_object*); @@ -85,7 +88,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Comma static lean_object* l_Lean_Elab_Command_mkDefViewOfAbbrev___closed__3; lean_object* l_Lean_Elab_expandMacroImpl_x3f(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); -extern lean_object* l_Lean_numLitKind; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_mkInstanceName___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_forEachExpr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_mkInstanceName___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); @@ -143,14 +145,15 @@ LEAN_EXPORT uint8_t l_Lean_Elab_DefKind_isTheorem(uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkFreshInstanceName(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkInstanceName___lambda__4(lean_object*, 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_mkDefViewOfInstance___spec__5(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_mkDefViewOfOpaque___closed__7; LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_mkDefViewOfInstance___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkAtomFrom(lean_object*, lean_object*); +lean_object* l_Lean_Syntax_mkNumLit(lean_object*, lean_object*); static lean_object* l_Lean_Elab_DefKind_noConfusion___rarg___closed__1; static lean_object* l_Lean_Elab_Command_mkDefViewOfOpaque___closed__4; uint8_t l_Lean_Expr_isType(lean_object*); lean_object* l_Lean_Syntax_getSepArgs(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_DefKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_mkDefViewOfExample___closed__3; lean_object* lean_environment_main_module(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_mkDefViewOfInstance___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -166,13 +169,12 @@ lean_object* l_Lean_Syntax_getKind(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkInstanceName(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkFreshInstanceName___boxed(lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_mkInstanceName___spec__4___closed__1; -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1469_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1485_(lean_object*); static lean_object* l_Lean_Elab_Command_mkInstanceName___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_mkInstanceName___spec__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_mkInstanceName___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_pure___at_Lean_Elab_liftMacroM___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_mkInstanceName___spec__4___closed__2; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1469____closed__2; static lean_object* l_Lean_Elab_Command_isDefLike___closed__6; lean_object* l_Lean_Elab_Command_getMainModule___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_mkInstanceName___lambda__1___closed__3; @@ -188,7 +190,6 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_mkDefView___sp uint8_t l_Lean_Syntax_isNone(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_mkDefViewOfInstance___spec__10(lean_object*, lean_object*); lean_object* l_Lean_Elab_Modifiers_addAttribute(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1469____closed__3; static lean_object* l_Lean_Elab_Command_mkInstanceName___lambda__2___closed__1; static lean_object* l_Lean_Elab_Command_mkDefViewOfInstance___closed__9; lean_object* l_Lean_Meta_instantiateMVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -214,7 +215,6 @@ static lean_object* l_Lean_Elab_Command_mkDefViewOfOpaque___lambda__2___closed__ static lean_object* l_Lean_Elab_Command_mkDefViewOfInstance___closed__18; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_mkInstanceName___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1469____closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkDefView(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_mkInstanceName___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_mkInstanceName___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); @@ -229,7 +229,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Comma static lean_object* l_Lean_Elab_Command_mkDefViewOfOpaque___closed__6; lean_object* l_Lean_Elab_Command_getScope___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_mkDefViewOfInstance___closed__14; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1469____closed__1; static lean_object* l_Lean_Elab_instInhabitedDefView___closed__2; static lean_object* l_Lean_Elab_Command_mkDefViewOfInstance___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_DefView_deriving_x3f___default; @@ -249,7 +248,6 @@ uint8_t l_Lean_Expr_isSort(lean_object*); lean_object* l_Lean_Elab_mkUnusedBaseName(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); uint8_t lean_string_dec_eq(lean_object*, lean_object*); -lean_object* l_Lean_Syntax_mkLit(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_mkDefViewOfExample___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_DefKind_toCtorIdx(uint8_t x_1) { _start: @@ -3428,7 +3426,7 @@ static lean_object* _init_l_Lean_Elab_Command_mkDefViewOfInstance___closed__18() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean_nullKind; +x_2 = l_Lean_Elab_Command_mkDefViewOfInstance___closed__10; x_3 = l_Lean_Elab_instInhabitedDefView___closed__1; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -3465,7 +3463,7 @@ lean_inc(x_3); x_15 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_mkDefViewOfInstance___spec__6(x_14, x_3, x_4, x_11); if (lean_obj_tag(x_15) == 0) { -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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +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; 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_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; lean_object* x_44; lean_object* x_45; uint8_t x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; x_16 = lean_ctor_get(x_15, 0); lean_inc(x_16); x_17 = lean_ctor_get(x_15, 1); @@ -3490,169 +3488,168 @@ x_26 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_26, 0, x_19); lean_ctor_set(x_26, 1, x_25); x_27 = l_Nat_repr(x_16); -x_28 = l_Lean_numLitKind; -x_29 = lean_box(2); -x_30 = l_Lean_Syntax_mkLit(x_28, x_27, x_29); -x_31 = l_Lean_Elab_Command_mkDefViewOfInstance___closed__11; -x_32 = lean_array_push(x_31, x_30); -x_33 = l_Lean_Elab_Command_mkDefViewOfInstance___closed__10; -x_34 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_34, 0, x_29); -lean_ctor_set(x_34, 1, x_33); -lean_ctor_set(x_34, 2, x_32); -x_35 = l_Lean_Elab_Command_mkDefViewOfInstance___closed__12; -x_36 = lean_array_push(x_35, x_26); -x_37 = lean_array_push(x_36, x_34); -x_38 = l_Lean_Elab_Command_mkDefViewOfInstance___closed__8; -x_39 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_39, 0, x_29); -lean_ctor_set(x_39, 1, x_38); -lean_ctor_set(x_39, 2, x_37); -x_40 = lean_unsigned_to_nat(4u); -x_41 = l_Lean_Syntax_getArg(x_2, x_40); -x_42 = l_Lean_Elab_expandDeclSig(x_41); -lean_dec(x_41); -x_43 = lean_ctor_get(x_42, 0); +x_28 = lean_box(2); +x_29 = l_Lean_Syntax_mkNumLit(x_27, x_28); +x_30 = l_Lean_Elab_Command_mkDefViewOfInstance___closed__11; +x_31 = lean_array_push(x_30, x_29); +x_32 = l_Lean_Elab_Command_mkDefViewOfInstance___closed__10; +x_33 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_33, 0, x_28); +lean_ctor_set(x_33, 1, x_32); +lean_ctor_set(x_33, 2, x_31); +x_34 = l_Lean_Elab_Command_mkDefViewOfInstance___closed__12; +x_35 = lean_array_push(x_34, x_26); +x_36 = lean_array_push(x_35, x_33); +x_37 = l_Lean_Elab_Command_mkDefViewOfInstance___closed__8; +x_38 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_38, 0, x_28); +lean_ctor_set(x_38, 1, x_37); +lean_ctor_set(x_38, 2, x_36); +x_39 = lean_unsigned_to_nat(4u); +x_40 = l_Lean_Syntax_getArg(x_2, x_39); +x_41 = l_Lean_Elab_expandDeclSig(x_40); +lean_dec(x_40); +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_41, 1); lean_inc(x_43); -x_44 = lean_ctor_get(x_42, 1); -lean_inc(x_44); -lean_dec(x_42); -x_45 = l_Lean_Elab_Command_mkDefViewOfInstance___closed__13; -x_46 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_46, 0, x_45); -lean_ctor_set(x_46, 1, x_39); -x_47 = lean_unbox(x_10); +lean_dec(x_41); +x_44 = l_Lean_Elab_Command_mkDefViewOfInstance___closed__13; +x_45 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_38); +x_46 = lean_unbox(x_10); lean_dec(x_10); -lean_ctor_set_uint8(x_46, sizeof(void*)*2, x_47); -x_48 = l_Lean_Elab_Modifiers_addAttribute(x_1, x_46); -x_49 = lean_unsigned_to_nat(3u); -x_50 = l_Lean_Syntax_getArg(x_2, x_49); -x_51 = l_Lean_Syntax_getOptional_x3f(x_50); -lean_dec(x_50); -if (lean_obj_tag(x_51) == 0) -{ -lean_object* x_52; lean_object* x_53; -x_52 = l_Lean_Syntax_getArgs(x_43); +lean_ctor_set_uint8(x_45, sizeof(void*)*2, x_46); +x_47 = l_Lean_Elab_Modifiers_addAttribute(x_1, x_45); +x_48 = lean_unsigned_to_nat(3u); +x_49 = l_Lean_Syntax_getArg(x_2, x_48); +x_50 = l_Lean_Syntax_getOptional_x3f(x_49); +lean_dec(x_49); +if (lean_obj_tag(x_50) == 0) +{ +lean_object* x_51; lean_object* x_52; +x_51 = l_Lean_Syntax_getArgs(x_42); lean_inc(x_4); lean_inc(x_3); -lean_inc(x_44); -x_53 = l_Lean_Elab_Command_mkInstanceName(x_52, x_44, x_3, x_4, x_24); -if (lean_obj_tag(x_53) == 0) +lean_inc(x_43); +x_52 = l_Lean_Elab_Command_mkInstanceName(x_51, x_43, x_3, x_4, x_24); +if (lean_obj_tag(x_52) == 0) { -lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_54 = lean_ctor_get(x_53, 0); +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_53 = lean_ctor_get(x_52, 0); +lean_inc(x_53); +x_54 = lean_ctor_get(x_52, 1); lean_inc(x_54); -x_55 = lean_ctor_get(x_53, 1); -lean_inc(x_55); -lean_dec(x_53); +lean_dec(x_52); lean_inc(x_2); -x_56 = l_Lean_mkIdentFrom(x_2, x_54); -x_57 = lean_array_push(x_35, x_56); -x_58 = l_Lean_Elab_Command_mkDefViewOfInstance___closed__18; -x_59 = lean_array_push(x_57, x_58); -x_60 = l_Lean_Elab_Command_mkDefViewOfInstance___closed__17; -x_61 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_61, 0, x_29); -lean_ctor_set(x_61, 1, x_60); -lean_ctor_set(x_61, 2, x_59); -x_62 = l_Lean_Elab_Command_mkDefViewOfInstance___lambda__1(x_44, x_2, x_48, x_43, x_61, x_3, x_4, x_55); +x_55 = l_Lean_mkIdentFrom(x_2, x_53); +x_56 = lean_array_push(x_34, x_55); +x_57 = l_Lean_Elab_Command_mkDefViewOfInstance___closed__18; +x_58 = lean_array_push(x_56, x_57); +x_59 = l_Lean_Elab_Command_mkDefViewOfInstance___closed__17; +x_60 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_60, 0, x_28); +lean_ctor_set(x_60, 1, x_59); +lean_ctor_set(x_60, 2, x_58); +x_61 = l_Lean_Elab_Command_mkDefViewOfInstance___lambda__1(x_43, x_2, x_47, x_42, x_60, x_3, x_4, x_54); lean_dec(x_4); lean_dec(x_3); -return x_62; +return x_61; } else { -uint8_t x_63; -lean_dec(x_48); -lean_dec(x_44); +uint8_t x_62; +lean_dec(x_47); lean_dec(x_43); +lean_dec(x_42); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_63 = !lean_is_exclusive(x_53); -if (x_63 == 0) +x_62 = !lean_is_exclusive(x_52); +if (x_62 == 0) { -return x_53; +return x_52; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_ctor_get(x_53, 0); -x_65 = lean_ctor_get(x_53, 1); -lean_inc(x_65); +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_52, 0); +x_64 = lean_ctor_get(x_52, 1); lean_inc(x_64); -lean_dec(x_53); -x_66 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_66, 0, x_64); -lean_ctor_set(x_66, 1, x_65); -return x_66; +lean_inc(x_63); +lean_dec(x_52); +x_65 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_64); +return x_65; } } } else { -lean_object* x_67; lean_object* x_68; -x_67 = lean_ctor_get(x_51, 0); -lean_inc(x_67); -lean_dec(x_51); -x_68 = l_Lean_Elab_Command_mkDefViewOfInstance___lambda__1(x_44, x_2, x_48, x_43, x_67, x_3, x_4, x_24); +lean_object* x_66; lean_object* x_67; +x_66 = lean_ctor_get(x_50, 0); +lean_inc(x_66); +lean_dec(x_50); +x_67 = l_Lean_Elab_Command_mkDefViewOfInstance___lambda__1(x_43, x_2, x_47, x_42, x_66, x_3, x_4, x_24); lean_dec(x_4); lean_dec(x_3); -return x_68; +return x_67; } } else { -uint8_t x_69; +uint8_t x_68; lean_dec(x_10); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_69 = !lean_is_exclusive(x_15); -if (x_69 == 0) +x_68 = !lean_is_exclusive(x_15); +if (x_68 == 0) { return x_15; } else { -lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_70 = lean_ctor_get(x_15, 0); -x_71 = lean_ctor_get(x_15, 1); -lean_inc(x_71); +lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_69 = lean_ctor_get(x_15, 0); +x_70 = lean_ctor_get(x_15, 1); lean_inc(x_70); +lean_inc(x_69); lean_dec(x_15); -x_72 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_72, 0, x_70); -lean_ctor_set(x_72, 1, x_71); -return x_72; +x_71 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_71, 0, x_69); +lean_ctor_set(x_71, 1, x_70); +return x_71; } } } else { -uint8_t x_73; +uint8_t x_72; lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_73 = !lean_is_exclusive(x_9); -if (x_73 == 0) +x_72 = !lean_is_exclusive(x_9); +if (x_72 == 0) { return x_9; } else { -lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_74 = lean_ctor_get(x_9, 0); -x_75 = lean_ctor_get(x_9, 1); -lean_inc(x_75); +lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_73 = lean_ctor_get(x_9, 0); +x_74 = lean_ctor_get(x_9, 1); lean_inc(x_74); +lean_inc(x_73); lean_dec(x_9); -x_76 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_76, 0, x_74); -lean_ctor_set(x_76, 1, x_75); -return x_76; +x_75 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set(x_75, 1, x_74); +return x_75; } } } @@ -3852,20 +3849,6 @@ return x_1; static lean_object* _init_l_Lean_Elab_Command_mkDefViewOfOpaque___closed__6() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = lean_box(2); -x_2 = l_Lean_Elab_Command_mkDefViewOfInstance___closed__10; -x_3 = l_Lean_Elab_instInhabitedDefView___closed__1; -x_4 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_4, 0, x_1); -lean_ctor_set(x_4, 1, x_2); -lean_ctor_set(x_4, 2, x_3); -return x_4; -} -} -static lean_object* _init_l_Lean_Elab_Command_mkDefViewOfOpaque___closed__7() { -_start: -{ lean_object* x_1; x_1 = lean_mk_string_from_bytes("unsafe", 6); return x_1; @@ -3927,7 +3910,7 @@ lean_ctor_set(x_24, 0, x_17); lean_ctor_set(x_24, 1, x_23); x_25 = l_Lean_Elab_Command_mkDefViewOfInstance___closed__12; x_26 = lean_array_push(x_25, x_24); -x_27 = l_Lean_Elab_Command_mkDefViewOfOpaque___closed__6; +x_27 = l_Lean_Elab_Command_mkDefViewOfInstance___closed__18; x_28 = lean_array_push(x_26, x_27); x_29 = lean_box(2); x_30 = l_Lean_Elab_Command_mkDefViewOfOpaque___closed__4; @@ -3960,7 +3943,7 @@ lean_inc(x_34); x_41 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_41, 0, x_34); lean_ctor_set(x_41, 1, x_40); -x_42 = l_Lean_Elab_Command_mkDefViewOfOpaque___closed__7; +x_42 = l_Lean_Elab_Command_mkDefViewOfOpaque___closed__6; x_43 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_43, 0, x_34); lean_ctor_set(x_43, 1, x_42); @@ -4026,6 +4009,20 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } +static lean_object* _init_l_Lean_Elab_Command_mkDefViewOfExample___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(2); +x_2 = l_Lean_Elab_Command_mkDefViewOfInstance___closed__10; +x_3 = l_Lean_Elab_instInhabitedDefView___closed__1; +x_4 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkDefViewOfExample(lean_object* x_1, lean_object* x_2) { _start: { @@ -4044,7 +4041,7 @@ lean_inc(x_2); x_9 = l_Lean_mkIdentFrom(x_2, x_8); x_10 = l_Lean_Elab_Command_mkDefViewOfInstance___closed__12; x_11 = lean_array_push(x_10, x_9); -x_12 = l_Lean_Elab_Command_mkDefViewOfInstance___closed__18; +x_12 = l_Lean_Elab_Command_mkDefViewOfExample___closed__3; x_13 = lean_array_push(x_11, x_12); x_14 = lean_box(2); x_15 = l_Lean_Elab_Command_mkDefViewOfInstance___closed__17; @@ -4447,7 +4444,7 @@ lean_dec(x_3); return x_5; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1469____closed__1() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1485____closed__1() { _start: { lean_object* x_1; @@ -4455,17 +4452,17 @@ x_1 = lean_mk_string_from_bytes("Elab", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1469____closed__2() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1485____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1469____closed__1; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1485____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1469____closed__3() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1485____closed__3() { _start: { lean_object* x_1; @@ -4473,21 +4470,21 @@ x_1 = lean_mk_string_from_bytes("definition", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1469____closed__4() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1485____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1469____closed__2; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1469____closed__3; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1485____closed__2; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1485____closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1469_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1485_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1469____closed__4; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1485____closed__4; x_3 = l_Lean_registerTraceClass(x_2, x_1); return x_3; } @@ -4650,12 +4647,12 @@ l_Lean_Elab_Command_mkDefViewOfOpaque___closed__5 = _init_l_Lean_Elab_Command_mk lean_mark_persistent(l_Lean_Elab_Command_mkDefViewOfOpaque___closed__5); l_Lean_Elab_Command_mkDefViewOfOpaque___closed__6 = _init_l_Lean_Elab_Command_mkDefViewOfOpaque___closed__6(); lean_mark_persistent(l_Lean_Elab_Command_mkDefViewOfOpaque___closed__6); -l_Lean_Elab_Command_mkDefViewOfOpaque___closed__7 = _init_l_Lean_Elab_Command_mkDefViewOfOpaque___closed__7(); -lean_mark_persistent(l_Lean_Elab_Command_mkDefViewOfOpaque___closed__7); l_Lean_Elab_Command_mkDefViewOfExample___closed__1 = _init_l_Lean_Elab_Command_mkDefViewOfExample___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_mkDefViewOfExample___closed__1); l_Lean_Elab_Command_mkDefViewOfExample___closed__2 = _init_l_Lean_Elab_Command_mkDefViewOfExample___closed__2(); lean_mark_persistent(l_Lean_Elab_Command_mkDefViewOfExample___closed__2); +l_Lean_Elab_Command_mkDefViewOfExample___closed__3 = _init_l_Lean_Elab_Command_mkDefViewOfExample___closed__3(); +lean_mark_persistent(l_Lean_Elab_Command_mkDefViewOfExample___closed__3); l_Lean_Elab_Command_isDefLike___closed__1 = _init_l_Lean_Elab_Command_isDefLike___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_isDefLike___closed__1); l_Lean_Elab_Command_isDefLike___closed__2 = _init_l_Lean_Elab_Command_isDefLike___closed__2(); @@ -4682,15 +4679,15 @@ l_Lean_Elab_Command_mkDefView___closed__1 = _init_l_Lean_Elab_Command_mkDefView_ lean_mark_persistent(l_Lean_Elab_Command_mkDefView___closed__1); l_Lean_Elab_Command_mkDefView___closed__2 = _init_l_Lean_Elab_Command_mkDefView___closed__2(); lean_mark_persistent(l_Lean_Elab_Command_mkDefView___closed__2); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1469____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1469____closed__1(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1469____closed__1); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1469____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1469____closed__2(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1469____closed__2); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1469____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1469____closed__3(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1469____closed__3); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1469____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1469____closed__4(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1469____closed__4); -res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1469_(lean_io_mk_world()); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1485____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1485____closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1485____closed__1); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1485____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1485____closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1485____closed__2); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1485____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1485____closed__3(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1485____closed__3); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1485____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1485____closed__4(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1485____closed__4); +res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_DefView___hyg_1485_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/Deriving/BEq.c b/stage0/stdlib/Lean/Elab/Deriving/BEq.c index d76876dbd02..ab41570bf21 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/BEq.c +++ b/stage0/stdlib/Lean/Elab/Deriving/BEq.c @@ -27,6 +27,7 @@ static lean_object* l_Lean_Elab_Deriving_BEq_mkMutualBlock___closed__14; lean_object* l_Lean_stringToMessageData(lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); static lean_object* l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqEnumFun___closed__23; +lean_object* l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__14(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqInstanceCmds(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqEnumFun___closed__13; lean_object* l_Lean_Elab_Command_liftTermElabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -44,9 +45,10 @@ static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts static lean_object* l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqInstanceCmds___closed__9; static lean_object* l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqInstanceCmds___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_uset(lean_object*, size_t, lean_object*); static lean_object* l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqEnumFun___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__3(lean_object*, lean_object*, 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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__1; lean_object* lean_environment_find(lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); @@ -60,7 +62,7 @@ static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__3; lean_object* lean_array_get_size(lean_object*); static lean_object* l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqEnumFun___closed__17; -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__3___boxed(lean_object*, lean_object*, 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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Core_betaReduce___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofList(lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__1; @@ -69,6 +71,7 @@ static lean_object* l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed_ static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__4; static lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__1___closed__4; lean_object* lean_string_utf8_byte_size(lean_object*); +uint8_t lean_usize_dec_lt(size_t, size_t); lean_object* l_Lean_Elab_Deriving_mkLocalInstanceLetDecls(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_BEq_mkBEqHeader___closed__1; static lean_object* l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqEnumFun___closed__2; @@ -77,7 +80,7 @@ lean_object* lean_nat_add(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_BEq_mkMutualBlock___closed__10; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__8; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__14; -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, 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_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Deriving_mkDiscrs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__6; LEAN_EXPORT lean_object* l_List_mapTRAux___at___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqInstanceCmds___spec__1(lean_object*, lean_object*); @@ -96,13 +99,13 @@ static lean_object* l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed_ lean_object* l_Lean_Elab_Deriving_mkInstanceCmds(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__11; -static lean_object* l_Lean_Elab_Deriving_BEq_initFn____x40_Lean_Elab_Deriving_BEq___hyg_3222____closed__1; static lean_object* l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__10; -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, 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_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__12; +static lean_object* l_Lean_Elab_Deriving_BEq_initFn____x40_Lean_Elab_Deriving_BEq___hyg_3410____closed__1; static lean_object* l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__2; static lean_object* l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__15; -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__18; static lean_object* l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqInstanceCmds___closed__8; lean_object* l_Lean_Elab_Deriving_mkLet(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -150,16 +153,18 @@ lean_object* l_Lean_addTrace___at___private_Lean_Elab_Term_0__Lean_Elab_Term_pos static lean_object* l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__23; lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqEnumFun___closed__16; +static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__14; static lean_object* l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__13; lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__16(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Deriving_BEq_mkBEqInstanceHandler___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__15; static lean_object* l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__21; static lean_object* l_Lean_Elab_Deriving_BEq_mkMutualBlock___closed__8; lean_object* l_Lean_Elab_Deriving_mkHeader(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_Deriving_BEq_mkMatch_mkAlts(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__5; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__8; -LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_BEq_initFn____x40_Lean_Elab_Deriving_BEq___hyg_3222_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_BEq_initFn____x40_Lean_Elab_Deriving_BEq___hyg_3410_(lean_object*); static lean_object* l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__29; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMutualBlock___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -177,6 +182,7 @@ static lean_object* l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq static lean_object* l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__16; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__4; uint8_t lean_nat_dec_le(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_BEq_mkMatch___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Deriving_BEq_mkBEqInstanceHandler___spec__4(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__3; static lean_object* l_Lean_Elab_Deriving_BEq_mkMatch___closed__5; @@ -190,6 +196,7 @@ static lean_object* l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed_ static lean_object* l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqEnumFun___closed__18; static lean_object* l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__12; static lean_object* l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqInstanceCmds___closed__3; +static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__16; static lean_object* l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqEnumFun___closed__15; static lean_object* l_Lean_Elab_Deriving_BEq_mkMutualBlock___closed__13; static lean_object* l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqEnumFun___closed__20; @@ -237,15 +244,16 @@ static lean_object* l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed_ static lean_object* l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__20; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__8; static lean_object* l_Lean_Elab_Deriving_BEq_mkMutualBlock___closed__6; +static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__13; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_BEq_mkBEqInstanceHandler(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqInstanceCmds___closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_BEq_mkAuxFunction(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__1; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_BEq_mkMatch___spec__1(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqEnumFun___closed__11; lean_object* l_Lean_mkConst(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__8; -lean_object* l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__10(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__12; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__6; static lean_object* l_Lean_Elab_Deriving_BEq_mkMutualBlock___closed__5; @@ -256,7 +264,7 @@ lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiv static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__2; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__13; static lean_object* l_Lean_Elab_Deriving_BEq_mkMutualBlock___closed__3; -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5(lean_object*, lean_object*, 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_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5(lean_object*, 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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__9; @@ -1036,104 +1044,104 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__3(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__3(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -uint8_t x_13; -x_13 = lean_nat_dec_le(x_3, x_2); -if (x_13 == 0) +uint8_t x_14; +x_14 = lean_nat_dec_le(x_4, x_3); +if (x_14 == 0) { -lean_object* x_14; uint8_t x_15; -x_14 = lean_unsigned_to_nat(0u); -x_15 = lean_nat_dec_eq(x_1, x_14); -if (x_15 == 0) +lean_object* x_15; uint8_t x_16; +x_15 = lean_unsigned_to_nat(0u); +x_16 = lean_nat_dec_eq(x_2, x_15); +if (x_16 == 0) { -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; 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_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_16 = lean_unsigned_to_nat(1u); -x_17 = lean_nat_sub(x_1, x_16); -lean_dec(x_1); -x_18 = lean_ctor_get(x_5, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_5, 1); +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; 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_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; lean_object* x_44; +x_17 = lean_unsigned_to_nat(1u); +x_18 = lean_nat_sub(x_2, x_17); +lean_dec(x_2); +x_19 = lean_ctor_get(x_6, 0); lean_inc(x_19); -lean_dec(x_5); -lean_inc(x_10); -x_20 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_10, x_11, x_12); -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_20, 1); +x_20 = lean_ctor_get(x_6, 1); +lean_inc(x_20); +lean_dec(x_6); +lean_inc(x_11); +x_21 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_11, x_12, x_13); +x_22 = lean_ctor_get(x_21, 0); lean_inc(x_22); -lean_dec(x_20); -x_23 = lean_st_ref_get(x_11, x_22); -x_24 = lean_ctor_get(x_23, 1); -lean_inc(x_24); -lean_dec(x_23); -x_25 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__9; -x_26 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_26, 0, x_21); -lean_ctor_set(x_26, 1, x_25); -x_27 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__10; -x_28 = lean_array_push(x_27, x_26); -x_29 = lean_box(2); -x_30 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__8; -x_31 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_30); -lean_ctor_set(x_31, 2, x_28); -x_32 = lean_array_push(x_18, x_31); -lean_inc(x_10); -x_33 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_10, x_11, x_24); -x_34 = lean_ctor_get(x_33, 0); -lean_inc(x_34); -x_35 = lean_ctor_get(x_33, 1); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = lean_st_ref_get(x_12, x_23); +x_25 = lean_ctor_get(x_24, 1); +lean_inc(x_25); +lean_dec(x_24); +x_26 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__9; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_22); +lean_ctor_set(x_27, 1, x_26); +x_28 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__10; +x_29 = lean_array_push(x_28, x_27); +x_30 = lean_box(2); +x_31 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__8; +x_32 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set(x_32, 2, x_29); +x_33 = lean_array_push(x_19, x_32); +lean_inc(x_11); +x_34 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_11, x_12, x_25); +x_35 = lean_ctor_get(x_34, 0); lean_inc(x_35); -lean_dec(x_33); -x_36 = lean_st_ref_get(x_11, x_35); -x_37 = lean_ctor_get(x_36, 1); -lean_inc(x_37); -lean_dec(x_36); -x_38 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_38, 0, x_34); -lean_ctor_set(x_38, 1, x_25); -x_39 = lean_array_push(x_27, x_38); -x_40 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_40, 0, x_29); -lean_ctor_set(x_40, 1, x_30); -lean_ctor_set(x_40, 2, x_39); -x_41 = lean_array_push(x_19, x_40); -x_42 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_42, 0, x_32); -lean_ctor_set(x_42, 1, x_41); -x_43 = lean_nat_add(x_2, x_4); -lean_dec(x_2); -x_1 = x_17; -x_2 = x_43; -x_5 = x_42; -x_12 = x_37; +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +lean_dec(x_34); +x_37 = lean_st_ref_get(x_12, x_36); +x_38 = lean_ctor_get(x_37, 1); +lean_inc(x_38); +lean_dec(x_37); +x_39 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_39, 0, x_35); +lean_ctor_set(x_39, 1, x_26); +x_40 = lean_array_push(x_28, x_39); +x_41 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_41, 0, x_30); +lean_ctor_set(x_41, 1, x_31); +lean_ctor_set(x_41, 2, x_40); +x_42 = lean_array_push(x_20, x_41); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_33); +lean_ctor_set(x_43, 1, x_42); +x_44 = lean_nat_add(x_3, x_5); +lean_dec(x_3); +x_2 = x_18; +x_3 = x_44; +x_6 = x_43; +x_13 = x_38; goto _start; } else { -lean_object* x_45; -lean_dec(x_10); +lean_object* x_46; +lean_dec(x_11); +lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_45 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_45, 0, x_5); -lean_ctor_set(x_45, 1, x_12); -return x_45; +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_6); +lean_ctor_set(x_46, 1, x_13); +return x_46; } } else { -lean_object* x_46; -lean_dec(x_10); +lean_object* x_47; +lean_dec(x_11); +lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_5); -lean_ctor_set(x_46, 1, x_12); -return x_46; +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_6); +lean_ctor_set(x_47, 1, x_13); +return x_47; } } } @@ -1261,595 +1269,595 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4(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_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, lean_object* x_16, lean_object* x_17) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4(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_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, lean_object* x_16, lean_object* x_17, lean_object* x_18) { _start: { -uint8_t x_18; -x_18 = lean_nat_dec_le(x_8, x_7); -if (x_18 == 0) +uint8_t x_19; +x_19 = lean_nat_dec_le(x_9, x_8); +if (x_19 == 0) { -lean_object* x_19; uint8_t x_20; -x_19 = lean_unsigned_to_nat(0u); -x_20 = lean_nat_dec_eq(x_6, x_19); -if (x_20 == 0) +lean_object* x_20; uint8_t x_21; +x_20 = lean_unsigned_to_nat(0u); +x_21 = lean_nat_dec_eq(x_7, x_20); +if (x_21 == 0) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_29; lean_object* x_30; uint8_t x_31; -x_21 = lean_unsigned_to_nat(1u); -x_22 = lean_nat_sub(x_6, x_21); -lean_dec(x_6); -x_29 = lean_ctor_get(x_10, 1); -lean_inc(x_29); -x_30 = lean_ctor_get(x_10, 0); +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_22 = lean_unsigned_to_nat(1u); +x_23 = lean_nat_sub(x_7, x_22); +lean_dec(x_7); +x_30 = lean_ctor_get(x_11, 1); lean_inc(x_30); -lean_dec(x_10); -x_31 = !lean_is_exclusive(x_29); -if (x_31 == 0) +x_31 = lean_ctor_get(x_11, 0); +lean_inc(x_31); +lean_dec(x_11); +x_32 = !lean_is_exclusive(x_30); +if (x_32 == 0) { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; -x_32 = lean_ctor_get(x_29, 0); -x_33 = lean_ctor_get(x_29, 1); -x_34 = lean_nat_add(x_5, x_7); -x_35 = l_Lean_instInhabitedExpr; -x_36 = lean_array_get(x_35, x_3, x_34); -lean_dec(x_34); -lean_inc(x_36); -x_37 = l_Lean_Expr_fvarId_x21(x_36); -x_38 = l_Lean_Expr_hasAnyFVar_visit___at_Lean_Expr_containsFVar___spec__1(x_37, x_4); -lean_dec(x_37); -if (x_38 == 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; uint8_t x_39; +x_33 = lean_ctor_get(x_30, 0); +x_34 = lean_ctor_get(x_30, 1); +x_35 = lean_nat_add(x_6, x_8); +x_36 = l_Lean_instInhabitedExpr; +x_37 = lean_array_get(x_36, x_3, x_35); +lean_dec(x_35); +lean_inc(x_37); +x_38 = l_Lean_Expr_fvarId_x21(x_37); +x_39 = l_Lean_Expr_hasAnyFVar_visit___at_Lean_Expr_containsFVar___spec__1(x_38, x_4); +lean_dec(x_38); +if (x_39 == 0) { -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_39 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__2; -x_40 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_39, x_15, x_16, x_17); -x_41 = lean_ctor_get(x_40, 0); -lean_inc(x_41); -x_42 = lean_ctor_get(x_40, 1); +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_40 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__2; +x_41 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_40, x_16, x_17, x_18); +x_42 = lean_ctor_get(x_41, 0); lean_inc(x_42); -lean_dec(x_40); -x_43 = lean_mk_syntax_ident(x_41); -x_44 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__4; -x_45 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_44, x_15, x_16, x_42); -x_46 = lean_ctor_get(x_45, 0); -lean_inc(x_46); -x_47 = lean_ctor_get(x_45, 1); -lean_inc(x_47); -lean_dec(x_45); -x_48 = lean_mk_syntax_ident(x_46); +x_43 = lean_ctor_get(x_41, 1); lean_inc(x_43); -x_49 = lean_array_push(x_30, x_43); +lean_dec(x_41); +x_44 = lean_mk_syntax_ident(x_42); +x_45 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__4; +x_46 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_45, x_16, x_17, x_43); +x_47 = lean_ctor_get(x_46, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_46, 1); lean_inc(x_48); -x_50 = lean_array_push(x_32, x_48); +lean_dec(x_46); +x_49 = lean_mk_syntax_ident(x_47); +lean_inc(x_44); +x_50 = lean_array_push(x_31, x_44); +lean_inc(x_49); +x_51 = lean_array_push(x_33, x_49); +lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -lean_inc(x_13); -x_51 = lean_infer_type(x_36, x_13, x_14, x_15, x_16, x_47); -if (lean_obj_tag(x_51) == 0) +x_52 = lean_infer_type(x_37, x_14, x_15, x_16, x_17, x_48); +if (lean_obj_tag(x_52) == 0) { -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; -x_52 = lean_ctor_get(x_51, 0); -lean_inc(x_52); -x_53 = lean_ctor_get(x_51, 1); +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; +x_53 = lean_ctor_get(x_52, 0); lean_inc(x_53); -lean_dec(x_51); -x_54 = lean_ctor_get(x_1, 0); -x_55 = lean_ctor_get(x_54, 0); -x_56 = l_Lean_Expr_isAppOf(x_52, x_55); +x_54 = lean_ctor_get(x_52, 1); +lean_inc(x_54); lean_dec(x_52); -if (x_56 == 0) +x_55 = lean_ctor_get(x_1, 0); +x_56 = lean_ctor_get(x_55, 0); +x_57 = l_Lean_Expr_isAppOf(x_53, x_56); +lean_dec(x_53); +if (x_57 == 0) { -lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; -lean_inc(x_15); -x_57 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_15, x_16, x_53); -x_58 = lean_ctor_get(x_57, 0); -lean_inc(x_58); -x_59 = lean_ctor_get(x_57, 1); +lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; +lean_inc(x_16); +x_58 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_16, x_17, x_54); +x_59 = lean_ctor_get(x_58, 0); lean_inc(x_59); -lean_dec(x_57); -x_60 = lean_st_ref_get(x_16, x_59); -x_61 = lean_ctor_get(x_60, 1); -lean_inc(x_61); -lean_dec(x_60); -x_62 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__7; -lean_inc(x_58); -x_63 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_63, 0, x_58); -lean_ctor_set(x_63, 1, x_62); -x_64 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__10; -x_65 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_65, 0, x_58); -lean_ctor_set(x_65, 1, x_64); -x_66 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__11; -x_67 = lean_array_push(x_66, x_43); -x_68 = lean_array_push(x_67, x_65); -x_69 = lean_array_push(x_68, x_48); -x_70 = lean_box(2); -x_71 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__9; -x_72 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_72, 0, x_70); -lean_ctor_set(x_72, 1, x_71); -lean_ctor_set(x_72, 2, x_69); -x_73 = lean_array_push(x_66, x_33); -x_74 = lean_array_push(x_73, x_63); -x_75 = lean_array_push(x_74, x_72); -x_76 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__6; -x_77 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_77, 0, x_70); -lean_ctor_set(x_77, 1, x_76); -lean_ctor_set(x_77, 2, x_75); -lean_ctor_set(x_29, 1, x_77); -lean_ctor_set(x_29, 0, x_50); -x_78 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_78, 0, x_49); -lean_ctor_set(x_78, 1, x_29); -x_79 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_79, 0, x_78); -x_23 = x_79; -x_24 = x_61; -goto block_28; +x_60 = lean_ctor_get(x_58, 1); +lean_inc(x_60); +lean_dec(x_58); +x_61 = lean_st_ref_get(x_17, x_60); +x_62 = lean_ctor_get(x_61, 1); +lean_inc(x_62); +lean_dec(x_61); +x_63 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__7; +lean_inc(x_59); +x_64 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_64, 0, x_59); +lean_ctor_set(x_64, 1, x_63); +x_65 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__10; +x_66 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_66, 0, x_59); +lean_ctor_set(x_66, 1, x_65); +x_67 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__11; +x_68 = lean_array_push(x_67, x_44); +x_69 = lean_array_push(x_68, x_66); +x_70 = lean_array_push(x_69, x_49); +x_71 = lean_box(2); +x_72 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__9; +x_73 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_73, 0, x_71); +lean_ctor_set(x_73, 1, x_72); +lean_ctor_set(x_73, 2, x_70); +x_74 = lean_array_push(x_67, x_34); +x_75 = lean_array_push(x_74, x_64); +x_76 = lean_array_push(x_75, x_73); +x_77 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__6; +x_78 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_78, 0, x_71); +lean_ctor_set(x_78, 1, x_77); +lean_ctor_set(x_78, 2, x_76); +lean_ctor_set(x_30, 1, x_78); +lean_ctor_set(x_30, 0, x_51); +x_79 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_79, 0, x_50); +lean_ctor_set(x_79, 1, x_30); +x_80 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_80, 0, x_79); +x_24 = x_80; +x_25 = x_62; +goto block_29; } else { -lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; -lean_inc(x_15); -x_80 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_15, x_16, x_53); -x_81 = lean_ctor_get(x_80, 0); -lean_inc(x_81); -x_82 = lean_ctor_get(x_80, 1); +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +lean_inc(x_16); +x_81 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_16, x_17, x_54); +x_82 = lean_ctor_get(x_81, 0); lean_inc(x_82); -lean_dec(x_80); -x_83 = lean_st_ref_get(x_16, x_82); -x_84 = lean_ctor_get(x_83, 1); -lean_inc(x_84); -lean_dec(x_83); -x_85 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__7; -x_86 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_86, 0, x_81); -lean_ctor_set(x_86, 1, x_85); +x_83 = lean_ctor_get(x_81, 1); +lean_inc(x_83); +lean_dec(x_81); +x_84 = lean_st_ref_get(x_17, x_83); +x_85 = lean_ctor_get(x_84, 1); +lean_inc(x_85); +lean_dec(x_84); +x_86 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__7; +x_87 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_87, 0, x_82); +lean_ctor_set(x_87, 1, x_86); lean_inc(x_2); -x_87 = lean_mk_syntax_ident(x_2); -x_88 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__14; -x_89 = lean_array_push(x_88, x_43); -x_90 = lean_array_push(x_89, x_48); -x_91 = lean_box(2); -x_92 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__15; -x_93 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_93, 0, x_91); -lean_ctor_set(x_93, 1, x_92); -lean_ctor_set(x_93, 2, x_90); -x_94 = lean_array_push(x_88, x_87); -x_95 = lean_array_push(x_94, x_93); -x_96 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__13; -x_97 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_97, 0, x_91); -lean_ctor_set(x_97, 1, x_96); -lean_ctor_set(x_97, 2, x_95); -x_98 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__11; -x_99 = lean_array_push(x_98, x_33); -x_100 = lean_array_push(x_99, x_86); -x_101 = lean_array_push(x_100, x_97); -x_102 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__6; -x_103 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_103, 0, x_91); -lean_ctor_set(x_103, 1, x_102); -lean_ctor_set(x_103, 2, x_101); -lean_ctor_set(x_29, 1, x_103); -lean_ctor_set(x_29, 0, x_50); -x_104 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_104, 0, x_49); -lean_ctor_set(x_104, 1, x_29); -x_105 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_105, 0, x_104); -x_23 = x_105; -x_24 = x_84; -goto block_28; -} -} -else -{ -uint8_t x_106; +x_88 = lean_mk_syntax_ident(x_2); +x_89 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__14; +x_90 = lean_array_push(x_89, x_44); +x_91 = lean_array_push(x_90, x_49); +x_92 = lean_box(2); +x_93 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__15; +x_94 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_94, 0, x_92); +lean_ctor_set(x_94, 1, x_93); +lean_ctor_set(x_94, 2, x_91); +x_95 = lean_array_push(x_89, x_88); +x_96 = lean_array_push(x_95, x_94); +x_97 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__13; +x_98 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_98, 0, x_92); +lean_ctor_set(x_98, 1, x_97); +lean_ctor_set(x_98, 2, x_96); +x_99 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__11; +x_100 = lean_array_push(x_99, x_34); +x_101 = lean_array_push(x_100, x_87); +x_102 = lean_array_push(x_101, x_98); +x_103 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__6; +x_104 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_104, 0, x_92); +lean_ctor_set(x_104, 1, x_103); +lean_ctor_set(x_104, 2, x_102); +lean_ctor_set(x_30, 1, x_104); +lean_ctor_set(x_30, 0, x_51); +x_105 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_105, 0, x_50); +lean_ctor_set(x_105, 1, x_30); +x_106 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_106, 0, x_105); +x_24 = x_106; +x_25 = x_85; +goto block_29; +} +} +else +{ +uint8_t x_107; +lean_dec(x_51); lean_dec(x_50); lean_dec(x_49); -lean_dec(x_48); -lean_dec(x_43); -lean_free_object(x_29); -lean_dec(x_33); -lean_dec(x_22); +lean_dec(x_44); +lean_free_object(x_30); +lean_dec(x_34); +lean_dec(x_23); +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_7); +lean_dec(x_8); lean_dec(x_2); -x_106 = !lean_is_exclusive(x_51); -if (x_106 == 0) +x_107 = !lean_is_exclusive(x_52); +if (x_107 == 0) { -return x_51; +return x_52; } else { -lean_object* x_107; lean_object* x_108; lean_object* x_109; -x_107 = lean_ctor_get(x_51, 0); -x_108 = lean_ctor_get(x_51, 1); +lean_object* x_108; lean_object* x_109; lean_object* x_110; +x_108 = lean_ctor_get(x_52, 0); +x_109 = lean_ctor_get(x_52, 1); +lean_inc(x_109); lean_inc(x_108); -lean_inc(x_107); -lean_dec(x_51); -x_109 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_109, 0, x_107); -lean_ctor_set(x_109, 1, x_108); -return x_109; +lean_dec(x_52); +x_110 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_110, 0, x_108); +lean_ctor_set(x_110, 1, x_109); +return x_110; } } } else { -lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; -lean_dec(x_36); -lean_inc(x_15); -x_110 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_15, x_16, x_17); -x_111 = lean_ctor_get(x_110, 0); -lean_inc(x_111); -x_112 = lean_ctor_get(x_110, 1); +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; +lean_dec(x_37); +lean_inc(x_16); +x_111 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_16, x_17, x_18); +x_112 = lean_ctor_get(x_111, 0); lean_inc(x_112); -lean_dec(x_110); -x_113 = lean_st_ref_get(x_16, x_112); -x_114 = lean_ctor_get(x_113, 1); -lean_inc(x_114); -lean_dec(x_113); -x_115 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__9; -x_116 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_116, 0, x_111); -lean_ctor_set(x_116, 1, x_115); -x_117 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__10; -x_118 = lean_array_push(x_117, x_116); -x_119 = lean_box(2); -x_120 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__8; -x_121 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_121, 0, x_119); -lean_ctor_set(x_121, 1, x_120); -lean_ctor_set(x_121, 2, x_118); -x_122 = lean_array_push(x_30, x_121); -lean_inc(x_15); -x_123 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_15, x_16, x_114); -x_124 = lean_ctor_get(x_123, 0); -lean_inc(x_124); -x_125 = lean_ctor_get(x_123, 1); +x_113 = lean_ctor_get(x_111, 1); +lean_inc(x_113); +lean_dec(x_111); +x_114 = lean_st_ref_get(x_17, x_113); +x_115 = lean_ctor_get(x_114, 1); +lean_inc(x_115); +lean_dec(x_114); +x_116 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__9; +x_117 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_117, 0, x_112); +lean_ctor_set(x_117, 1, x_116); +x_118 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__10; +x_119 = lean_array_push(x_118, x_117); +x_120 = lean_box(2); +x_121 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__8; +x_122 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_122, 0, x_120); +lean_ctor_set(x_122, 1, x_121); +lean_ctor_set(x_122, 2, x_119); +x_123 = lean_array_push(x_31, x_122); +lean_inc(x_16); +x_124 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_16, x_17, x_115); +x_125 = lean_ctor_get(x_124, 0); lean_inc(x_125); -lean_dec(x_123); -x_126 = lean_st_ref_get(x_16, x_125); -x_127 = lean_ctor_get(x_126, 1); -lean_inc(x_127); -lean_dec(x_126); -x_128 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_128, 0, x_124); -lean_ctor_set(x_128, 1, x_115); -x_129 = lean_array_push(x_117, x_128); -x_130 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_130, 0, x_119); -lean_ctor_set(x_130, 1, x_120); -lean_ctor_set(x_130, 2, x_129); -x_131 = lean_array_push(x_32, x_130); -lean_ctor_set(x_29, 0, x_131); -x_132 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_132, 0, x_122); -lean_ctor_set(x_132, 1, x_29); -x_133 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_133, 0, x_132); -x_23 = x_133; -x_24 = x_127; -goto block_28; -} -} -else -{ -lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; uint8_t x_140; -x_134 = lean_ctor_get(x_29, 0); -x_135 = lean_ctor_get(x_29, 1); -lean_inc(x_135); -lean_inc(x_134); -lean_dec(x_29); -x_136 = lean_nat_add(x_5, x_7); -x_137 = l_Lean_instInhabitedExpr; -x_138 = lean_array_get(x_137, x_3, x_136); -lean_dec(x_136); -lean_inc(x_138); -x_139 = l_Lean_Expr_fvarId_x21(x_138); -x_140 = l_Lean_Expr_hasAnyFVar_visit___at_Lean_Expr_containsFVar___spec__1(x_139, x_4); -lean_dec(x_139); -if (x_140 == 0) +x_126 = lean_ctor_get(x_124, 1); +lean_inc(x_126); +lean_dec(x_124); +x_127 = lean_st_ref_get(x_17, x_126); +x_128 = lean_ctor_get(x_127, 1); +lean_inc(x_128); +lean_dec(x_127); +x_129 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_129, 0, x_125); +lean_ctor_set(x_129, 1, x_116); +x_130 = lean_array_push(x_118, x_129); +x_131 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_131, 0, x_120); +lean_ctor_set(x_131, 1, x_121); +lean_ctor_set(x_131, 2, x_130); +x_132 = lean_array_push(x_33, x_131); +lean_ctor_set(x_30, 0, x_132); +x_133 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_133, 0, x_123); +lean_ctor_set(x_133, 1, x_30); +x_134 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_134, 0, x_133); +x_24 = x_134; +x_25 = x_128; +goto block_29; +} +} +else { -lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; -x_141 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__2; -x_142 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_141, x_15, x_16, x_17); -x_143 = lean_ctor_get(x_142, 0); -lean_inc(x_143); -x_144 = lean_ctor_get(x_142, 1); +lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; uint8_t x_141; +x_135 = lean_ctor_get(x_30, 0); +x_136 = lean_ctor_get(x_30, 1); +lean_inc(x_136); +lean_inc(x_135); +lean_dec(x_30); +x_137 = lean_nat_add(x_6, x_8); +x_138 = l_Lean_instInhabitedExpr; +x_139 = lean_array_get(x_138, x_3, x_137); +lean_dec(x_137); +lean_inc(x_139); +x_140 = l_Lean_Expr_fvarId_x21(x_139); +x_141 = l_Lean_Expr_hasAnyFVar_visit___at_Lean_Expr_containsFVar___spec__1(x_140, x_4); +lean_dec(x_140); +if (x_141 == 0) +{ +lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; +x_142 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__2; +x_143 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_142, x_16, x_17, x_18); +x_144 = lean_ctor_get(x_143, 0); lean_inc(x_144); -lean_dec(x_142); -x_145 = lean_mk_syntax_ident(x_143); -x_146 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__4; -x_147 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_146, x_15, x_16, x_144); -x_148 = lean_ctor_get(x_147, 0); -lean_inc(x_148); -x_149 = lean_ctor_get(x_147, 1); -lean_inc(x_149); -lean_dec(x_147); -x_150 = lean_mk_syntax_ident(x_148); +x_145 = lean_ctor_get(x_143, 1); lean_inc(x_145); -x_151 = lean_array_push(x_30, x_145); +lean_dec(x_143); +x_146 = lean_mk_syntax_ident(x_144); +x_147 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__4; +x_148 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_147, x_16, x_17, x_145); +x_149 = lean_ctor_get(x_148, 0); +lean_inc(x_149); +x_150 = lean_ctor_get(x_148, 1); lean_inc(x_150); -x_152 = lean_array_push(x_134, x_150); +lean_dec(x_148); +x_151 = lean_mk_syntax_ident(x_149); +lean_inc(x_146); +x_152 = lean_array_push(x_31, x_146); +lean_inc(x_151); +x_153 = lean_array_push(x_135, x_151); +lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -lean_inc(x_13); -x_153 = lean_infer_type(x_138, x_13, x_14, x_15, x_16, x_149); -if (lean_obj_tag(x_153) == 0) +x_154 = lean_infer_type(x_139, x_14, x_15, x_16, x_17, x_150); +if (lean_obj_tag(x_154) == 0) { -lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; uint8_t x_158; -x_154 = lean_ctor_get(x_153, 0); -lean_inc(x_154); -x_155 = lean_ctor_get(x_153, 1); +lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; uint8_t x_159; +x_155 = lean_ctor_get(x_154, 0); lean_inc(x_155); -lean_dec(x_153); -x_156 = lean_ctor_get(x_1, 0); -x_157 = lean_ctor_get(x_156, 0); -x_158 = l_Lean_Expr_isAppOf(x_154, x_157); +x_156 = lean_ctor_get(x_154, 1); +lean_inc(x_156); lean_dec(x_154); -if (x_158 == 0) +x_157 = lean_ctor_get(x_1, 0); +x_158 = lean_ctor_get(x_157, 0); +x_159 = l_Lean_Expr_isAppOf(x_155, x_158); +lean_dec(x_155); +if (x_159 == 0) { -lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; -lean_inc(x_15); -x_159 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_15, x_16, x_155); -x_160 = lean_ctor_get(x_159, 0); -lean_inc(x_160); -x_161 = lean_ctor_get(x_159, 1); +lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; +lean_inc(x_16); +x_160 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_16, x_17, x_156); +x_161 = lean_ctor_get(x_160, 0); lean_inc(x_161); -lean_dec(x_159); -x_162 = lean_st_ref_get(x_16, x_161); -x_163 = lean_ctor_get(x_162, 1); -lean_inc(x_163); -lean_dec(x_162); -x_164 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__7; -lean_inc(x_160); -x_165 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_165, 0, x_160); -lean_ctor_set(x_165, 1, x_164); -x_166 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__10; -x_167 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_167, 0, x_160); -lean_ctor_set(x_167, 1, x_166); -x_168 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__11; -x_169 = lean_array_push(x_168, x_145); -x_170 = lean_array_push(x_169, x_167); -x_171 = lean_array_push(x_170, x_150); -x_172 = lean_box(2); -x_173 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__9; -x_174 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_174, 0, x_172); -lean_ctor_set(x_174, 1, x_173); -lean_ctor_set(x_174, 2, x_171); -x_175 = lean_array_push(x_168, x_135); -x_176 = lean_array_push(x_175, x_165); -x_177 = lean_array_push(x_176, x_174); -x_178 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__6; -x_179 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_179, 0, x_172); -lean_ctor_set(x_179, 1, x_178); -lean_ctor_set(x_179, 2, x_177); -x_180 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_180, 0, x_152); +x_162 = lean_ctor_get(x_160, 1); +lean_inc(x_162); +lean_dec(x_160); +x_163 = lean_st_ref_get(x_17, x_162); +x_164 = lean_ctor_get(x_163, 1); +lean_inc(x_164); +lean_dec(x_163); +x_165 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__7; +lean_inc(x_161); +x_166 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_166, 0, x_161); +lean_ctor_set(x_166, 1, x_165); +x_167 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__10; +x_168 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_168, 0, x_161); +lean_ctor_set(x_168, 1, x_167); +x_169 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__11; +x_170 = lean_array_push(x_169, x_146); +x_171 = lean_array_push(x_170, x_168); +x_172 = lean_array_push(x_171, x_151); +x_173 = lean_box(2); +x_174 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__9; +x_175 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_175, 0, x_173); +lean_ctor_set(x_175, 1, x_174); +lean_ctor_set(x_175, 2, x_172); +x_176 = lean_array_push(x_169, x_136); +x_177 = lean_array_push(x_176, x_166); +x_178 = lean_array_push(x_177, x_175); +x_179 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__6; +x_180 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_180, 0, x_173); lean_ctor_set(x_180, 1, x_179); +lean_ctor_set(x_180, 2, x_178); x_181 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_181, 0, x_151); +lean_ctor_set(x_181, 0, x_153); lean_ctor_set(x_181, 1, x_180); -x_182 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_182, 0, x_181); -x_23 = x_182; -x_24 = x_163; -goto block_28; +x_182 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_182, 0, x_152); +lean_ctor_set(x_182, 1, x_181); +x_183 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_183, 0, x_182); +x_24 = x_183; +x_25 = x_164; +goto block_29; } else { -lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; -lean_inc(x_15); -x_183 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_15, x_16, x_155); -x_184 = lean_ctor_get(x_183, 0); -lean_inc(x_184); -x_185 = lean_ctor_get(x_183, 1); +lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; +lean_inc(x_16); +x_184 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_16, x_17, x_156); +x_185 = lean_ctor_get(x_184, 0); lean_inc(x_185); -lean_dec(x_183); -x_186 = lean_st_ref_get(x_16, x_185); -x_187 = lean_ctor_get(x_186, 1); -lean_inc(x_187); -lean_dec(x_186); -x_188 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__7; -x_189 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_189, 0, x_184); -lean_ctor_set(x_189, 1, x_188); +x_186 = lean_ctor_get(x_184, 1); +lean_inc(x_186); +lean_dec(x_184); +x_187 = lean_st_ref_get(x_17, x_186); +x_188 = lean_ctor_get(x_187, 1); +lean_inc(x_188); +lean_dec(x_187); +x_189 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__7; +x_190 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_190, 0, x_185); +lean_ctor_set(x_190, 1, x_189); lean_inc(x_2); -x_190 = lean_mk_syntax_ident(x_2); -x_191 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__14; -x_192 = lean_array_push(x_191, x_145); -x_193 = lean_array_push(x_192, x_150); -x_194 = lean_box(2); -x_195 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__15; -x_196 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_196, 0, x_194); -lean_ctor_set(x_196, 1, x_195); -lean_ctor_set(x_196, 2, x_193); -x_197 = lean_array_push(x_191, x_190); -x_198 = lean_array_push(x_197, x_196); -x_199 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__13; -x_200 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_200, 0, x_194); -lean_ctor_set(x_200, 1, x_199); -lean_ctor_set(x_200, 2, x_198); -x_201 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__11; -x_202 = lean_array_push(x_201, x_135); -x_203 = lean_array_push(x_202, x_189); -x_204 = lean_array_push(x_203, x_200); -x_205 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__6; -x_206 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_206, 0, x_194); -lean_ctor_set(x_206, 1, x_205); -lean_ctor_set(x_206, 2, x_204); -x_207 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_207, 0, x_152); +x_191 = lean_mk_syntax_ident(x_2); +x_192 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__14; +x_193 = lean_array_push(x_192, x_146); +x_194 = lean_array_push(x_193, x_151); +x_195 = lean_box(2); +x_196 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__15; +x_197 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_197, 0, x_195); +lean_ctor_set(x_197, 1, x_196); +lean_ctor_set(x_197, 2, x_194); +x_198 = lean_array_push(x_192, x_191); +x_199 = lean_array_push(x_198, x_197); +x_200 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__13; +x_201 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_201, 0, x_195); +lean_ctor_set(x_201, 1, x_200); +lean_ctor_set(x_201, 2, x_199); +x_202 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__11; +x_203 = lean_array_push(x_202, x_136); +x_204 = lean_array_push(x_203, x_190); +x_205 = lean_array_push(x_204, x_201); +x_206 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__6; +x_207 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_207, 0, x_195); lean_ctor_set(x_207, 1, x_206); +lean_ctor_set(x_207, 2, x_205); x_208 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_208, 0, x_151); +lean_ctor_set(x_208, 0, x_153); lean_ctor_set(x_208, 1, x_207); -x_209 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_209, 0, x_208); -x_23 = x_209; -x_24 = x_187; -goto block_28; +x_209 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_209, 0, x_152); +lean_ctor_set(x_209, 1, x_208); +x_210 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_210, 0, x_209); +x_24 = x_210; +x_25 = x_188; +goto block_29; } } else { -lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; +lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; +lean_dec(x_153); lean_dec(x_152); lean_dec(x_151); -lean_dec(x_150); -lean_dec(x_145); -lean_dec(x_135); -lean_dec(x_22); +lean_dec(x_146); +lean_dec(x_136); +lean_dec(x_23); +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_7); +lean_dec(x_8); lean_dec(x_2); -x_210 = lean_ctor_get(x_153, 0); -lean_inc(x_210); -x_211 = lean_ctor_get(x_153, 1); +x_211 = lean_ctor_get(x_154, 0); lean_inc(x_211); -if (lean_is_exclusive(x_153)) { - lean_ctor_release(x_153, 0); - lean_ctor_release(x_153, 1); - x_212 = x_153; +x_212 = lean_ctor_get(x_154, 1); +lean_inc(x_212); +if (lean_is_exclusive(x_154)) { + lean_ctor_release(x_154, 0); + lean_ctor_release(x_154, 1); + x_213 = x_154; } else { - lean_dec_ref(x_153); - x_212 = lean_box(0); + lean_dec_ref(x_154); + x_213 = lean_box(0); } -if (lean_is_scalar(x_212)) { - x_213 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_213)) { + x_214 = lean_alloc_ctor(1, 2, 0); } else { - x_213 = x_212; + x_214 = x_213; } -lean_ctor_set(x_213, 0, x_210); -lean_ctor_set(x_213, 1, x_211); -return x_213; +lean_ctor_set(x_214, 0, x_211); +lean_ctor_set(x_214, 1, x_212); +return x_214; } } else { -lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; -lean_dec(x_138); -lean_inc(x_15); -x_214 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_15, x_16, x_17); -x_215 = lean_ctor_get(x_214, 0); -lean_inc(x_215); -x_216 = lean_ctor_get(x_214, 1); +lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; +lean_dec(x_139); +lean_inc(x_16); +x_215 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_16, x_17, x_18); +x_216 = lean_ctor_get(x_215, 0); lean_inc(x_216); -lean_dec(x_214); -x_217 = lean_st_ref_get(x_16, x_216); -x_218 = lean_ctor_get(x_217, 1); -lean_inc(x_218); -lean_dec(x_217); -x_219 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__9; -x_220 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_220, 0, x_215); -lean_ctor_set(x_220, 1, x_219); -x_221 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__10; -x_222 = lean_array_push(x_221, x_220); -x_223 = lean_box(2); -x_224 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__8; -x_225 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_225, 0, x_223); -lean_ctor_set(x_225, 1, x_224); -lean_ctor_set(x_225, 2, x_222); -x_226 = lean_array_push(x_30, x_225); -lean_inc(x_15); -x_227 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_15, x_16, x_218); -x_228 = lean_ctor_get(x_227, 0); -lean_inc(x_228); -x_229 = lean_ctor_get(x_227, 1); +x_217 = lean_ctor_get(x_215, 1); +lean_inc(x_217); +lean_dec(x_215); +x_218 = lean_st_ref_get(x_17, x_217); +x_219 = lean_ctor_get(x_218, 1); +lean_inc(x_219); +lean_dec(x_218); +x_220 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__9; +x_221 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_221, 0, x_216); +lean_ctor_set(x_221, 1, x_220); +x_222 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__10; +x_223 = lean_array_push(x_222, x_221); +x_224 = lean_box(2); +x_225 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__8; +x_226 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_226, 0, x_224); +lean_ctor_set(x_226, 1, x_225); +lean_ctor_set(x_226, 2, x_223); +x_227 = lean_array_push(x_31, x_226); +lean_inc(x_16); +x_228 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_16, x_17, x_219); +x_229 = lean_ctor_get(x_228, 0); lean_inc(x_229); -lean_dec(x_227); -x_230 = lean_st_ref_get(x_16, x_229); -x_231 = lean_ctor_get(x_230, 1); -lean_inc(x_231); -lean_dec(x_230); -x_232 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_232, 0, x_228); -lean_ctor_set(x_232, 1, x_219); -x_233 = lean_array_push(x_221, x_232); -x_234 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_234, 0, x_223); -lean_ctor_set(x_234, 1, x_224); -lean_ctor_set(x_234, 2, x_233); -x_235 = lean_array_push(x_134, x_234); -x_236 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_236, 0, x_235); -lean_ctor_set(x_236, 1, x_135); +x_230 = lean_ctor_get(x_228, 1); +lean_inc(x_230); +lean_dec(x_228); +x_231 = lean_st_ref_get(x_17, x_230); +x_232 = lean_ctor_get(x_231, 1); +lean_inc(x_232); +lean_dec(x_231); +x_233 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_233, 0, x_229); +lean_ctor_set(x_233, 1, x_220); +x_234 = lean_array_push(x_222, x_233); +x_235 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_235, 0, x_224); +lean_ctor_set(x_235, 1, x_225); +lean_ctor_set(x_235, 2, x_234); +x_236 = lean_array_push(x_135, x_235); x_237 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_237, 0, x_226); -lean_ctor_set(x_237, 1, x_236); -x_238 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_238, 0, x_237); -x_23 = x_238; -x_24 = x_231; -goto block_28; -} -} -block_28: -{ -lean_object* x_25; lean_object* x_26; -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -x_26 = lean_nat_add(x_7, x_9); -lean_dec(x_7); -x_6 = x_22; -x_7 = x_26; -x_10 = x_25; -x_17 = x_24; +lean_ctor_set(x_237, 0, x_236); +lean_ctor_set(x_237, 1, x_136); +x_238 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_238, 0, x_227); +lean_ctor_set(x_238, 1, x_237); +x_239 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_239, 0, x_238); +x_24 = x_239; +x_25 = x_232; +goto block_29; +} +} +block_29: +{ +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_24, 0); +lean_inc(x_26); +lean_dec(x_24); +x_27 = lean_nat_add(x_8, x_10); +lean_dec(x_8); +x_7 = x_23; +x_8 = x_27; +x_11 = x_26; +x_18 = x_25; goto _start; } } else { -lean_object* x_239; +lean_object* x_240; +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); -lean_dec(x_13); +lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); lean_dec(x_2); -x_239 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_239, 0, x_10); -lean_ctor_set(x_239, 1, x_17); -return x_239; +x_240 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_240, 0, x_11); +lean_ctor_set(x_240, 1, x_18); +return x_240; } } else { -lean_object* x_240; +lean_object* x_241; +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); -lean_dec(x_13); +lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); lean_dec(x_2); -x_240 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_240, 0, x_10); -lean_ctor_set(x_240, 1, x_17); -return x_240; +x_241 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_241, 0, x_11); +lean_ctor_set(x_241, 1, x_18); +return x_241; } } } @@ -1873,26 +1881,56 @@ static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_ _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("true", 4); +x_1 = lean_mk_string_from_bytes("term", 4); return x_1; } } static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__4() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__3; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__4; +x_3 = lean_alloc_ctor(1, 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_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("true", 4); +return x_1; +} +} +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__7() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__3; +x_1 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__6; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__5() { +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__3; +x_1 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__6; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__4; +x_3 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__7; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1900,51 +1938,62 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__6() { +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__3; +x_2 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__6; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__7() { +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__7; -x_2 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__3; +x_2 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__6; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__8() { +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__7; +x_2 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__10; 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_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__9() { +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__8; +x_2 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__11; x_3 = lean_alloc_ctor(1, 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_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__10() { +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__1; +x_2 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +return x_2; +} +} +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__14() { _start: { lean_object* x_1; @@ -1952,17 +2001,17 @@ x_1 = lean_mk_string_from_bytes("explicit", 8); return x_1; } } -static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__11() { +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__6; -x_2 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__10; +x_2 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__14; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__12() { +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__16() { _start: { lean_object* x_1; @@ -1970,47 +2019,47 @@ x_1 = lean_mk_string_from_bytes("@", 1); return x_1; } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___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, 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_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___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, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__1; -x_16 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__2; -lean_inc(x_13); +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__1; +x_15 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__2; lean_inc(x_12); -x_17 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_7, x_15, x_16, x_12, x_13, x_14); -if (lean_obj_tag(x_17) == 0) +lean_inc(x_11); +x_16 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_6, x_14, x_15, x_11, x_12, x_13); +if (lean_obj_tag(x_16) == 0) { -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; 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_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; lean_object* x_44; uint8_t x_45; -x_18 = lean_ctor_get(x_17, 0); +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; 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_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; lean_object* x_44; lean_object* x_45; uint8_t x_46; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); +lean_dec(x_16); +x_19 = lean_ctor_get(x_1, 2); lean_inc(x_19); -lean_dec(x_17); -x_20 = lean_ctor_get(x_1, 2); -lean_inc(x_20); -x_21 = lean_unsigned_to_nat(0u); -x_22 = lean_unsigned_to_nat(1u); -lean_inc(x_12); -lean_inc(x_2); -lean_inc(x_20); -x_23 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1(x_20, x_21, x_20, x_22, x_2, x_8, x_9, x_10, x_11, x_12, x_13, x_19); -lean_dec(x_20); +x_20 = lean_unsigned_to_nat(0u); +x_21 = lean_unsigned_to_nat(1u); +x_22 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__1; +lean_inc(x_11); +lean_inc(x_19); +x_23 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1(x_19, x_20, x_19, x_21, x_22, x_7, x_8, x_9, x_10, x_11, x_12, x_18); +lean_dec(x_19); x_24 = lean_ctor_get(x_23, 0); lean_inc(x_24); x_25 = lean_ctor_get(x_23, 1); lean_inc(x_25); lean_dec(x_23); -lean_inc(x_12); -x_26 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_12, x_13, x_25); +lean_inc(x_11); +x_26 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_11, x_12, x_25); x_27 = lean_ctor_get(x_26, 0); lean_inc(x_27); x_28 = lean_ctor_get(x_26, 1); lean_inc(x_28); lean_dec(x_26); -x_29 = lean_ctor_get(x_12, 10); +x_29 = lean_ctor_get(x_11, 10); lean_inc(x_29); -x_30 = lean_st_ref_get(x_13, x_28); +x_30 = lean_st_ref_get(x_12, x_28); x_31 = lean_ctor_get(x_30, 0); lean_inc(x_31); x_32 = lean_ctor_get(x_30, 1); @@ -2020,10 +2069,10 @@ x_33 = lean_ctor_get(x_31, 0); lean_inc(x_33); lean_dec(x_31); x_34 = lean_environment_main_module(x_33); -x_35 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__6; +x_35 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__9; x_36 = l_Lean_addMacroScope(x_34, x_35, x_29); -x_37 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__5; -x_38 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__9; +x_37 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__8; +x_38 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__12; x_39 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_39, 0, x_27); lean_ctor_set(x_39, 1, x_37); @@ -2031,651 +2080,637 @@ lean_ctor_set(x_39, 2, x_36); lean_ctor_set(x_39, 3, x_38); x_40 = lean_ctor_get(x_1, 1); lean_inc(x_40); -lean_inc_n(x_2, 2); -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_2); -lean_ctor_set(x_41, 1, x_2); -lean_inc(x_12); +x_41 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__5; +x_42 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__13; +lean_inc(x_11); lean_inc(x_40); -x_42 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__3(x_40, x_21, x_40, x_22, x_41, x_8, x_9, x_10, x_11, x_12, x_13, x_32); -x_43 = lean_ctor_get(x_42, 0); -lean_inc(x_43); -x_44 = lean_ctor_get(x_42, 1); +x_43 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__3(x_41, x_40, x_20, x_40, x_21, x_42, x_7, x_8, x_9, x_10, x_11, x_12, x_32); +x_44 = lean_ctor_get(x_43, 0); lean_inc(x_44); -lean_dec(x_42); -x_45 = !lean_is_exclusive(x_43); -if (x_45 == 0) +x_45 = lean_ctor_get(x_43, 1); +lean_inc(x_45); +lean_dec(x_43); +x_46 = !lean_is_exclusive(x_44); +if (x_46 == 0) { -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_46 = lean_ctor_get(x_43, 0); -x_47 = lean_ctor_get(x_43, 1); -x_48 = lean_ctor_get(x_3, 4); -lean_inc(x_48); -lean_dec(x_3); -lean_ctor_set(x_43, 1, x_39); -lean_ctor_set(x_43, 0, x_47); -x_49 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_49, 0, x_46); -lean_ctor_set(x_49, 1, x_43); -lean_inc(x_13); +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_47 = lean_ctor_get(x_44, 0); +x_48 = lean_ctor_get(x_44, 1); +x_49 = lean_ctor_get(x_2, 4); +lean_inc(x_49); +lean_dec(x_2); +lean_ctor_set(x_44, 1, x_39); +lean_ctor_set(x_44, 0, x_48); +x_50 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_50, 0, x_47); +lean_ctor_set(x_50, 1, x_44); lean_inc(x_12); -lean_inc(x_48); -x_50 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4(x_1, x_4, x_6, x_18, x_40, x_48, x_21, x_48, x_22, x_49, x_8, x_9, x_10, x_11, x_12, x_13, x_44); -lean_dec(x_48); +lean_inc(x_11); +lean_inc(x_49); +x_51 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4(x_1, x_3, x_5, x_17, x_41, x_40, x_49, x_20, x_49, x_21, x_50, x_7, x_8, x_9, x_10, x_11, x_12, x_45); +lean_dec(x_49); lean_dec(x_40); -lean_dec(x_18); +lean_dec(x_17); lean_dec(x_1); -if (lean_obj_tag(x_50) == 0) +if (lean_obj_tag(x_51) == 0) { -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; uint8_t x_98; -x_51 = lean_ctor_get(x_50, 0); -lean_inc(x_51); -x_52 = lean_ctor_get(x_51, 1); +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; uint8_t x_99; +x_52 = lean_ctor_get(x_51, 0); lean_inc(x_52); -x_53 = lean_ctor_get(x_50, 1); +x_53 = lean_ctor_get(x_52, 1); lean_inc(x_53); -lean_dec(x_50); -x_54 = lean_ctor_get(x_51, 0); +x_54 = lean_ctor_get(x_51, 1); lean_inc(x_54); lean_dec(x_51); x_55 = lean_ctor_get(x_52, 0); lean_inc(x_55); -x_56 = lean_ctor_get(x_52, 1); -lean_inc(x_56); lean_dec(x_52); -lean_inc(x_12); -x_57 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_12, x_13, x_53); -x_58 = lean_ctor_get(x_57, 0); -lean_inc(x_58); -x_59 = lean_ctor_get(x_57, 1); +x_56 = lean_ctor_get(x_53, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_53, 1); +lean_inc(x_57); +lean_dec(x_53); +lean_inc(x_11); +x_58 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_11, x_12, x_54); +x_59 = lean_ctor_get(x_58, 0); lean_inc(x_59); -lean_dec(x_57); -x_60 = lean_st_ref_get(x_13, x_59); -x_61 = lean_ctor_get(x_60, 1); -lean_inc(x_61); -lean_dec(x_60); -x_62 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__12; -x_63 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_63, 0, x_58); -lean_ctor_set(x_63, 1, x_62); -x_64 = lean_mk_syntax_ident(x_5); -x_65 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__14; -x_66 = lean_array_push(x_65, x_63); -lean_inc(x_64); +x_60 = lean_ctor_get(x_58, 1); +lean_inc(x_60); +lean_dec(x_58); +x_61 = lean_st_ref_get(x_12, x_60); +x_62 = lean_ctor_get(x_61, 1); +lean_inc(x_62); +lean_dec(x_61); +x_63 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__16; +x_64 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_64, 0, x_59); +lean_ctor_set(x_64, 1, x_63); +x_65 = lean_mk_syntax_ident(x_4); +x_66 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__14; x_67 = lean_array_push(x_66, x_64); -x_68 = lean_box(2); -x_69 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__11; -x_70 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_69); -lean_ctor_set(x_70, 2, x_67); -lean_inc(x_2); -x_71 = l_Array_append___rarg(x_2, x_54); -x_72 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__15; -x_73 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_73, 0, x_68); -lean_ctor_set(x_73, 1, x_72); -lean_ctor_set(x_73, 2, x_71); -x_74 = lean_array_push(x_65, x_70); -x_75 = lean_array_push(x_74, x_73); -x_76 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__13; -x_77 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_77, 0, x_68); -lean_ctor_set(x_77, 1, x_76); -lean_ctor_set(x_77, 2, x_75); -x_78 = lean_array_push(x_24, x_77); -lean_inc(x_12); -x_79 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_12, x_13, x_61); -x_80 = lean_ctor_get(x_79, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_79, 1); +lean_inc(x_65); +x_68 = lean_array_push(x_67, x_65); +x_69 = lean_box(2); +x_70 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__15; +x_71 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_71, 0, x_69); +lean_ctor_set(x_71, 1, x_70); +lean_ctor_set(x_71, 2, x_68); +x_72 = l_Array_append___rarg(x_22, x_55); +x_73 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__15; +x_74 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_74, 0, x_69); +lean_ctor_set(x_74, 1, x_73); +lean_ctor_set(x_74, 2, x_72); +x_75 = lean_array_push(x_66, x_71); +x_76 = lean_array_push(x_75, x_74); +x_77 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__13; +x_78 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_78, 0, x_69); +lean_ctor_set(x_78, 1, x_77); +lean_ctor_set(x_78, 2, x_76); +x_79 = lean_array_push(x_24, x_78); +lean_inc(x_11); +x_80 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_11, x_12, x_62); +x_81 = lean_ctor_get(x_80, 0); lean_inc(x_81); -lean_dec(x_79); -x_82 = lean_st_ref_get(x_13, x_81); -x_83 = lean_ctor_get(x_82, 1); -lean_inc(x_83); -lean_dec(x_82); -x_84 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_84, 0, x_80); -lean_ctor_set(x_84, 1, x_62); -x_85 = lean_array_push(x_65, x_84); -x_86 = lean_array_push(x_85, x_64); -x_87 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_87, 0, x_68); -lean_ctor_set(x_87, 1, x_69); -lean_ctor_set(x_87, 2, x_86); -lean_inc(x_2); -x_88 = l_Array_append___rarg(x_2, x_55); -x_89 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_89, 0, x_68); -lean_ctor_set(x_89, 1, x_72); -lean_ctor_set(x_89, 2, x_88); -x_90 = lean_array_push(x_65, x_87); -x_91 = lean_array_push(x_90, x_89); -x_92 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_92, 0, x_68); -lean_ctor_set(x_92, 1, x_76); -lean_ctor_set(x_92, 2, x_91); -x_93 = lean_array_push(x_78, x_92); -x_94 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_12, x_13, x_83); -x_95 = lean_ctor_get(x_94, 0); -lean_inc(x_95); -x_96 = lean_ctor_get(x_94, 1); +x_82 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +lean_dec(x_80); +x_83 = lean_st_ref_get(x_12, x_82); +x_84 = lean_ctor_get(x_83, 1); +lean_inc(x_84); +lean_dec(x_83); +x_85 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_85, 0, x_81); +lean_ctor_set(x_85, 1, x_63); +x_86 = lean_array_push(x_66, x_85); +x_87 = lean_array_push(x_86, x_65); +x_88 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_88, 0, x_69); +lean_ctor_set(x_88, 1, x_70); +lean_ctor_set(x_88, 2, x_87); +x_89 = l_Array_append___rarg(x_22, x_56); +x_90 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_90, 0, x_69); +lean_ctor_set(x_90, 1, x_73); +lean_ctor_set(x_90, 2, x_89); +x_91 = lean_array_push(x_66, x_88); +x_92 = lean_array_push(x_91, x_90); +x_93 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_93, 0, x_69); +lean_ctor_set(x_93, 1, x_77); +lean_ctor_set(x_93, 2, x_92); +x_94 = lean_array_push(x_79, x_93); +x_95 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_11, x_12, x_84); +x_96 = lean_ctor_get(x_95, 0); lean_inc(x_96); -lean_dec(x_94); -x_97 = lean_st_ref_get(x_13, x_96); -lean_dec(x_13); -x_98 = !lean_is_exclusive(x_97); -if (x_98 == 0) -{ -lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; size_t x_103; size_t x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; -x_99 = lean_ctor_get(x_97, 0); -lean_dec(x_99); -x_100 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__13; -lean_inc(x_95); -x_101 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_101, 0, x_95); -lean_ctor_set(x_101, 1, x_100); -x_102 = lean_array_get_size(x_93); -x_103 = lean_usize_of_nat(x_102); -lean_dec(x_102); -x_104 = 0; -x_105 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_103, x_104, x_93); -x_106 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__17; -x_107 = l_Lean_mkSepArray(x_105, x_106); -lean_dec(x_105); -x_108 = l_Array_append___rarg(x_2, x_107); -x_109 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_109, 0, x_68); -lean_ctor_set(x_109, 1, x_72); -lean_ctor_set(x_109, 2, x_108); -x_110 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__10; -x_111 = lean_array_push(x_110, x_109); -x_112 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_112, 0, x_68); -lean_ctor_set(x_112, 1, x_72); -lean_ctor_set(x_112, 2, x_111); -x_113 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__18; -x_114 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_114, 0, x_95); -lean_ctor_set(x_114, 1, x_113); -x_115 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__19; -x_116 = lean_array_push(x_115, x_101); -x_117 = lean_array_push(x_116, x_112); -x_118 = lean_array_push(x_117, x_114); -x_119 = lean_array_push(x_118, x_56); -x_120 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__12; -x_121 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_121, 0, x_68); -lean_ctor_set(x_121, 1, x_120); -lean_ctor_set(x_121, 2, x_119); -lean_ctor_set(x_97, 0, x_121); -return x_97; -} -else -{ -lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; size_t x_126; size_t x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; -x_122 = lean_ctor_get(x_97, 1); -lean_inc(x_122); -lean_dec(x_97); -x_123 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__13; -lean_inc(x_95); -x_124 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_124, 0, x_95); -lean_ctor_set(x_124, 1, x_123); -x_125 = lean_array_get_size(x_93); -x_126 = lean_usize_of_nat(x_125); -lean_dec(x_125); -x_127 = 0; -x_128 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_126, x_127, x_93); -x_129 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__17; -x_130 = l_Lean_mkSepArray(x_128, x_129); -lean_dec(x_128); -x_131 = l_Array_append___rarg(x_2, x_130); -x_132 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_132, 0, x_68); -lean_ctor_set(x_132, 1, x_72); -lean_ctor_set(x_132, 2, x_131); -x_133 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__10; -x_134 = lean_array_push(x_133, x_132); -x_135 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_135, 0, x_68); -lean_ctor_set(x_135, 1, x_72); -lean_ctor_set(x_135, 2, x_134); -x_136 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__18; -x_137 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_137, 0, x_95); -lean_ctor_set(x_137, 1, x_136); -x_138 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__19; -x_139 = lean_array_push(x_138, x_124); -x_140 = lean_array_push(x_139, x_135); -x_141 = lean_array_push(x_140, x_137); -x_142 = lean_array_push(x_141, x_56); -x_143 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__12; -x_144 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_144, 0, x_68); -lean_ctor_set(x_144, 1, x_143); -lean_ctor_set(x_144, 2, x_142); -x_145 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_145, 0, x_144); -lean_ctor_set(x_145, 1, x_122); -return x_145; +x_97 = lean_ctor_get(x_95, 1); +lean_inc(x_97); +lean_dec(x_95); +x_98 = lean_st_ref_get(x_12, x_97); +lean_dec(x_12); +x_99 = !lean_is_exclusive(x_98); +if (x_99 == 0) +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; size_t x_104; size_t x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; +x_100 = lean_ctor_get(x_98, 0); +lean_dec(x_100); +x_101 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__13; +lean_inc(x_96); +x_102 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_102, 0, x_96); +lean_ctor_set(x_102, 1, x_101); +x_103 = lean_array_get_size(x_94); +x_104 = lean_usize_of_nat(x_103); +lean_dec(x_103); +x_105 = 0; +x_106 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_104, x_105, x_94); +x_107 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__17; +x_108 = l_Lean_mkSepArray(x_106, x_107); +lean_dec(x_106); +x_109 = l_Array_append___rarg(x_22, x_108); +x_110 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_110, 0, x_69); +lean_ctor_set(x_110, 1, x_73); +lean_ctor_set(x_110, 2, x_109); +x_111 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__10; +x_112 = lean_array_push(x_111, x_110); +x_113 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_113, 0, x_69); +lean_ctor_set(x_113, 1, x_73); +lean_ctor_set(x_113, 2, x_112); +x_114 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__18; +x_115 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_115, 0, x_96); +lean_ctor_set(x_115, 1, x_114); +x_116 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__19; +x_117 = lean_array_push(x_116, x_102); +x_118 = lean_array_push(x_117, x_113); +x_119 = lean_array_push(x_118, x_115); +x_120 = lean_array_push(x_119, x_57); +x_121 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__12; +x_122 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_122, 0, x_69); +lean_ctor_set(x_122, 1, x_121); +lean_ctor_set(x_122, 2, x_120); +lean_ctor_set(x_98, 0, x_122); +return x_98; +} +else +{ +lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; size_t x_127; size_t x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; +x_123 = lean_ctor_get(x_98, 1); +lean_inc(x_123); +lean_dec(x_98); +x_124 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__13; +lean_inc(x_96); +x_125 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_125, 0, x_96); +lean_ctor_set(x_125, 1, x_124); +x_126 = lean_array_get_size(x_94); +x_127 = lean_usize_of_nat(x_126); +lean_dec(x_126); +x_128 = 0; +x_129 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_127, x_128, x_94); +x_130 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__17; +x_131 = l_Lean_mkSepArray(x_129, x_130); +lean_dec(x_129); +x_132 = l_Array_append___rarg(x_22, x_131); +x_133 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_133, 0, x_69); +lean_ctor_set(x_133, 1, x_73); +lean_ctor_set(x_133, 2, x_132); +x_134 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__10; +x_135 = lean_array_push(x_134, x_133); +x_136 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_136, 0, x_69); +lean_ctor_set(x_136, 1, x_73); +lean_ctor_set(x_136, 2, x_135); +x_137 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__18; +x_138 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_138, 0, x_96); +lean_ctor_set(x_138, 1, x_137); +x_139 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__19; +x_140 = lean_array_push(x_139, x_125); +x_141 = lean_array_push(x_140, x_136); +x_142 = lean_array_push(x_141, x_138); +x_143 = lean_array_push(x_142, x_57); +x_144 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__12; +x_145 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_145, 0, x_69); +lean_ctor_set(x_145, 1, x_144); +lean_ctor_set(x_145, 2, x_143); +x_146 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_146, 0, x_145); +lean_ctor_set(x_146, 1, x_123); +return x_146; } } else { -uint8_t x_146; +uint8_t x_147; lean_dec(x_24); -lean_dec(x_13); lean_dec(x_12); -lean_dec(x_5); -lean_dec(x_2); -x_146 = !lean_is_exclusive(x_50); -if (x_146 == 0) +lean_dec(x_11); +lean_dec(x_4); +x_147 = !lean_is_exclusive(x_51); +if (x_147 == 0) { -return x_50; +return x_51; } else { -lean_object* x_147; lean_object* x_148; lean_object* x_149; -x_147 = lean_ctor_get(x_50, 0); -x_148 = lean_ctor_get(x_50, 1); +lean_object* x_148; lean_object* x_149; lean_object* x_150; +x_148 = lean_ctor_get(x_51, 0); +x_149 = lean_ctor_get(x_51, 1); +lean_inc(x_149); lean_inc(x_148); -lean_inc(x_147); -lean_dec(x_50); -x_149 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_149, 0, x_147); -lean_ctor_set(x_149, 1, x_148); -return x_149; +lean_dec(x_51); +x_150 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_150, 0, x_148); +lean_ctor_set(x_150, 1, x_149); +return x_150; } } } else { -lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; -x_150 = lean_ctor_get(x_43, 0); -x_151 = lean_ctor_get(x_43, 1); -lean_inc(x_151); -lean_inc(x_150); -lean_dec(x_43); -x_152 = lean_ctor_get(x_3, 4); +lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; +x_151 = lean_ctor_get(x_44, 0); +x_152 = lean_ctor_get(x_44, 1); lean_inc(x_152); -lean_dec(x_3); -x_153 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_153, 0, x_151); -lean_ctor_set(x_153, 1, x_39); +lean_inc(x_151); +lean_dec(x_44); +x_153 = lean_ctor_get(x_2, 4); +lean_inc(x_153); +lean_dec(x_2); x_154 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_154, 0, x_150); -lean_ctor_set(x_154, 1, x_153); -lean_inc(x_13); +lean_ctor_set(x_154, 0, x_152); +lean_ctor_set(x_154, 1, x_39); +x_155 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_155, 0, x_151); +lean_ctor_set(x_155, 1, x_154); lean_inc(x_12); -lean_inc(x_152); -x_155 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4(x_1, x_4, x_6, x_18, x_40, x_152, x_21, x_152, x_22, x_154, x_8, x_9, x_10, x_11, x_12, x_13, x_44); -lean_dec(x_152); +lean_inc(x_11); +lean_inc(x_153); +x_156 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4(x_1, x_3, x_5, x_17, x_41, x_40, x_153, x_20, x_153, x_21, x_155, x_7, x_8, x_9, x_10, x_11, x_12, x_45); +lean_dec(x_153); lean_dec(x_40); -lean_dec(x_18); +lean_dec(x_17); lean_dec(x_1); -if (lean_obj_tag(x_155) == 0) +if (lean_obj_tag(x_156) == 0) { -lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; size_t x_208; size_t x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; -x_156 = lean_ctor_get(x_155, 0); -lean_inc(x_156); -x_157 = lean_ctor_get(x_156, 1); +lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; size_t x_209; size_t x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; +x_157 = lean_ctor_get(x_156, 0); lean_inc(x_157); -x_158 = lean_ctor_get(x_155, 1); +x_158 = lean_ctor_get(x_157, 1); lean_inc(x_158); -lean_dec(x_155); -x_159 = lean_ctor_get(x_156, 0); +x_159 = lean_ctor_get(x_156, 1); lean_inc(x_159); lean_dec(x_156); x_160 = lean_ctor_get(x_157, 0); lean_inc(x_160); -x_161 = lean_ctor_get(x_157, 1); -lean_inc(x_161); lean_dec(x_157); -lean_inc(x_12); -x_162 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_12, x_13, x_158); -x_163 = lean_ctor_get(x_162, 0); -lean_inc(x_163); -x_164 = lean_ctor_get(x_162, 1); +x_161 = lean_ctor_get(x_158, 0); +lean_inc(x_161); +x_162 = lean_ctor_get(x_158, 1); +lean_inc(x_162); +lean_dec(x_158); +lean_inc(x_11); +x_163 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_11, x_12, x_159); +x_164 = lean_ctor_get(x_163, 0); lean_inc(x_164); -lean_dec(x_162); -x_165 = lean_st_ref_get(x_13, x_164); -x_166 = lean_ctor_get(x_165, 1); -lean_inc(x_166); -lean_dec(x_165); -x_167 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__12; -x_168 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_168, 0, x_163); -lean_ctor_set(x_168, 1, x_167); -x_169 = lean_mk_syntax_ident(x_5); -x_170 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__14; -x_171 = lean_array_push(x_170, x_168); -lean_inc(x_169); +x_165 = lean_ctor_get(x_163, 1); +lean_inc(x_165); +lean_dec(x_163); +x_166 = lean_st_ref_get(x_12, x_165); +x_167 = lean_ctor_get(x_166, 1); +lean_inc(x_167); +lean_dec(x_166); +x_168 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__16; +x_169 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_169, 0, x_164); +lean_ctor_set(x_169, 1, x_168); +x_170 = lean_mk_syntax_ident(x_4); +x_171 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__14; x_172 = lean_array_push(x_171, x_169); -x_173 = lean_box(2); -x_174 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__11; -x_175 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_175, 0, x_173); -lean_ctor_set(x_175, 1, x_174); -lean_ctor_set(x_175, 2, x_172); -lean_inc(x_2); -x_176 = l_Array_append___rarg(x_2, x_159); -x_177 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__15; -x_178 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_178, 0, x_173); -lean_ctor_set(x_178, 1, x_177); -lean_ctor_set(x_178, 2, x_176); -x_179 = lean_array_push(x_170, x_175); -x_180 = lean_array_push(x_179, x_178); -x_181 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__13; -x_182 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_182, 0, x_173); -lean_ctor_set(x_182, 1, x_181); -lean_ctor_set(x_182, 2, x_180); -x_183 = lean_array_push(x_24, x_182); -lean_inc(x_12); -x_184 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_12, x_13, x_166); -x_185 = lean_ctor_get(x_184, 0); -lean_inc(x_185); -x_186 = lean_ctor_get(x_184, 1); +lean_inc(x_170); +x_173 = lean_array_push(x_172, x_170); +x_174 = lean_box(2); +x_175 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__15; +x_176 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_176, 0, x_174); +lean_ctor_set(x_176, 1, x_175); +lean_ctor_set(x_176, 2, x_173); +x_177 = l_Array_append___rarg(x_22, x_160); +x_178 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__15; +x_179 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_179, 0, x_174); +lean_ctor_set(x_179, 1, x_178); +lean_ctor_set(x_179, 2, x_177); +x_180 = lean_array_push(x_171, x_176); +x_181 = lean_array_push(x_180, x_179); +x_182 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__13; +x_183 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_183, 0, x_174); +lean_ctor_set(x_183, 1, x_182); +lean_ctor_set(x_183, 2, x_181); +x_184 = lean_array_push(x_24, x_183); +lean_inc(x_11); +x_185 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_11, x_12, x_167); +x_186 = lean_ctor_get(x_185, 0); lean_inc(x_186); -lean_dec(x_184); -x_187 = lean_st_ref_get(x_13, x_186); -x_188 = lean_ctor_get(x_187, 1); -lean_inc(x_188); -lean_dec(x_187); -x_189 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_189, 0, x_185); -lean_ctor_set(x_189, 1, x_167); -x_190 = lean_array_push(x_170, x_189); -x_191 = lean_array_push(x_190, x_169); -x_192 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_192, 0, x_173); -lean_ctor_set(x_192, 1, x_174); -lean_ctor_set(x_192, 2, x_191); -lean_inc(x_2); -x_193 = l_Array_append___rarg(x_2, x_160); -x_194 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_194, 0, x_173); -lean_ctor_set(x_194, 1, x_177); -lean_ctor_set(x_194, 2, x_193); -x_195 = lean_array_push(x_170, x_192); -x_196 = lean_array_push(x_195, x_194); -x_197 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_197, 0, x_173); -lean_ctor_set(x_197, 1, x_181); -lean_ctor_set(x_197, 2, x_196); -x_198 = lean_array_push(x_183, x_197); -x_199 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_12, x_13, x_188); -x_200 = lean_ctor_get(x_199, 0); -lean_inc(x_200); -x_201 = lean_ctor_get(x_199, 1); +x_187 = lean_ctor_get(x_185, 1); +lean_inc(x_187); +lean_dec(x_185); +x_188 = lean_st_ref_get(x_12, x_187); +x_189 = lean_ctor_get(x_188, 1); +lean_inc(x_189); +lean_dec(x_188); +x_190 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_190, 0, x_186); +lean_ctor_set(x_190, 1, x_168); +x_191 = lean_array_push(x_171, x_190); +x_192 = lean_array_push(x_191, x_170); +x_193 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_193, 0, x_174); +lean_ctor_set(x_193, 1, x_175); +lean_ctor_set(x_193, 2, x_192); +x_194 = l_Array_append___rarg(x_22, x_161); +x_195 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_195, 0, x_174); +lean_ctor_set(x_195, 1, x_178); +lean_ctor_set(x_195, 2, x_194); +x_196 = lean_array_push(x_171, x_193); +x_197 = lean_array_push(x_196, x_195); +x_198 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_198, 0, x_174); +lean_ctor_set(x_198, 1, x_182); +lean_ctor_set(x_198, 2, x_197); +x_199 = lean_array_push(x_184, x_198); +x_200 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_11, x_12, x_189); +x_201 = lean_ctor_get(x_200, 0); lean_inc(x_201); -lean_dec(x_199); -x_202 = lean_st_ref_get(x_13, x_201); -lean_dec(x_13); -x_203 = lean_ctor_get(x_202, 1); -lean_inc(x_203); -if (lean_is_exclusive(x_202)) { - lean_ctor_release(x_202, 0); - lean_ctor_release(x_202, 1); - x_204 = x_202; +x_202 = lean_ctor_get(x_200, 1); +lean_inc(x_202); +lean_dec(x_200); +x_203 = lean_st_ref_get(x_12, x_202); +lean_dec(x_12); +x_204 = lean_ctor_get(x_203, 1); +lean_inc(x_204); +if (lean_is_exclusive(x_203)) { + lean_ctor_release(x_203, 0); + lean_ctor_release(x_203, 1); + x_205 = x_203; } else { - lean_dec_ref(x_202); - x_204 = lean_box(0); -} -x_205 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__13; -lean_inc(x_200); -x_206 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_206, 0, x_200); -lean_ctor_set(x_206, 1, x_205); -x_207 = lean_array_get_size(x_198); -x_208 = lean_usize_of_nat(x_207); -lean_dec(x_207); -x_209 = 0; -x_210 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_208, x_209, x_198); -x_211 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__17; -x_212 = l_Lean_mkSepArray(x_210, x_211); -lean_dec(x_210); -x_213 = l_Array_append___rarg(x_2, x_212); -x_214 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_214, 0, x_173); -lean_ctor_set(x_214, 1, x_177); -lean_ctor_set(x_214, 2, x_213); -x_215 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__10; -x_216 = lean_array_push(x_215, x_214); -x_217 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_217, 0, x_173); -lean_ctor_set(x_217, 1, x_177); -lean_ctor_set(x_217, 2, x_216); -x_218 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__18; -x_219 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_219, 0, x_200); -lean_ctor_set(x_219, 1, x_218); -x_220 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__19; -x_221 = lean_array_push(x_220, x_206); -x_222 = lean_array_push(x_221, x_217); -x_223 = lean_array_push(x_222, x_219); -x_224 = lean_array_push(x_223, x_161); -x_225 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__12; -x_226 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_226, 0, x_173); -lean_ctor_set(x_226, 1, x_225); -lean_ctor_set(x_226, 2, x_224); -if (lean_is_scalar(x_204)) { - x_227 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_203); + x_205 = lean_box(0); +} +x_206 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__13; +lean_inc(x_201); +x_207 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_207, 0, x_201); +lean_ctor_set(x_207, 1, x_206); +x_208 = lean_array_get_size(x_199); +x_209 = lean_usize_of_nat(x_208); +lean_dec(x_208); +x_210 = 0; +x_211 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_209, x_210, x_199); +x_212 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__17; +x_213 = l_Lean_mkSepArray(x_211, x_212); +lean_dec(x_211); +x_214 = l_Array_append___rarg(x_22, x_213); +x_215 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_215, 0, x_174); +lean_ctor_set(x_215, 1, x_178); +lean_ctor_set(x_215, 2, x_214); +x_216 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___spec__1___closed__10; +x_217 = lean_array_push(x_216, x_215); +x_218 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_218, 0, x_174); +lean_ctor_set(x_218, 1, x_178); +lean_ctor_set(x_218, 2, x_217); +x_219 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__18; +x_220 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_220, 0, x_201); +lean_ctor_set(x_220, 1, x_219); +x_221 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__19; +x_222 = lean_array_push(x_221, x_207); +x_223 = lean_array_push(x_222, x_218); +x_224 = lean_array_push(x_223, x_220); +x_225 = lean_array_push(x_224, x_162); +x_226 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__12; +x_227 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_227, 0, x_174); +lean_ctor_set(x_227, 1, x_226); +lean_ctor_set(x_227, 2, x_225); +if (lean_is_scalar(x_205)) { + x_228 = lean_alloc_ctor(0, 2, 0); } else { - x_227 = x_204; + x_228 = x_205; } -lean_ctor_set(x_227, 0, x_226); -lean_ctor_set(x_227, 1, x_203); -return x_227; +lean_ctor_set(x_228, 0, x_227); +lean_ctor_set(x_228, 1, x_204); +return x_228; } else { -lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; +lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_dec(x_24); -lean_dec(x_13); lean_dec(x_12); -lean_dec(x_5); -lean_dec(x_2); -x_228 = lean_ctor_get(x_155, 0); -lean_inc(x_228); -x_229 = lean_ctor_get(x_155, 1); +lean_dec(x_11); +lean_dec(x_4); +x_229 = lean_ctor_get(x_156, 0); lean_inc(x_229); -if (lean_is_exclusive(x_155)) { - lean_ctor_release(x_155, 0); - lean_ctor_release(x_155, 1); - x_230 = x_155; +x_230 = lean_ctor_get(x_156, 1); +lean_inc(x_230); +if (lean_is_exclusive(x_156)) { + lean_ctor_release(x_156, 0); + lean_ctor_release(x_156, 1); + x_231 = x_156; } else { - lean_dec_ref(x_155); - x_230 = lean_box(0); + lean_dec_ref(x_156); + x_231 = lean_box(0); } -if (lean_is_scalar(x_230)) { - x_231 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_231)) { + x_232 = lean_alloc_ctor(1, 2, 0); } else { - x_231 = x_230; + x_232 = x_231; } -lean_ctor_set(x_231, 0, x_228); -lean_ctor_set(x_231, 1, x_229); -return x_231; +lean_ctor_set(x_232, 0, x_229); +lean_ctor_set(x_232, 1, x_230); +return x_232; } } } else { -uint8_t x_232; -lean_dec(x_13); +uint8_t x_233; lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_5); +lean_dec(x_9); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_232 = !lean_is_exclusive(x_17); -if (x_232 == 0) +x_233 = !lean_is_exclusive(x_16); +if (x_233 == 0) { -return x_17; +return x_16; } else { -lean_object* x_233; lean_object* x_234; lean_object* x_235; -x_233 = lean_ctor_get(x_17, 0); -x_234 = lean_ctor_get(x_17, 1); +lean_object* x_234; lean_object* x_235; lean_object* x_236; +x_234 = lean_ctor_get(x_16, 0); +x_235 = lean_ctor_get(x_16, 1); +lean_inc(x_235); lean_inc(x_234); -lean_inc(x_233); -lean_dec(x_17); -x_235 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_235, 0, x_233); -lean_ctor_set(x_235, 1, x_234); -return x_235; +lean_dec(x_16); +x_236 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_236, 0, x_234); +lean_ctor_set(x_236, 1, x_235); +return x_236; } } } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5(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_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -if (lean_obj_tag(x_4) == 0) +if (lean_obj_tag(x_3) == 0) { -lean_object* x_13; -lean_dec(x_11); +lean_object* x_12; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_3); +lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_5); -lean_ctor_set(x_13, 1, x_12); -return x_13; +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_4); +lean_ctor_set(x_12, 1, x_11); +return x_12; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_4, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_4, 1); -lean_inc(x_15); -lean_dec(x_4); -lean_inc(x_6); +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_3, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_3, 1); lean_inc(x_14); -x_16 = l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__1(x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_16) == 0) +lean_dec(x_3); +lean_inc(x_5); +lean_inc(x_13); +x_15 = l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__1(x_13, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_17 = lean_ctor_get(x_16, 0); +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); +lean_dec(x_15); +x_18 = lean_ctor_get(x_16, 0); lean_inc(x_18); -lean_dec(x_16); -x_19 = lean_ctor_get(x_17, 0); +x_19 = lean_ctor_get(x_18, 2); lean_inc(x_19); -x_20 = lean_ctor_get(x_19, 2); -lean_inc(x_20); -lean_dec(x_19); +lean_dec(x_18); lean_inc(x_2); -lean_inc(x_3); lean_inc(x_1); -x_21 = lean_alloc_closure((void*)(l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___boxed), 14, 5); -lean_closure_set(x_21, 0, x_1); -lean_closure_set(x_21, 1, x_3); -lean_closure_set(x_21, 2, x_17); -lean_closure_set(x_21, 3, x_2); -lean_closure_set(x_21, 4, x_14); -lean_inc(x_11); +x_20 = lean_alloc_closure((void*)(l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___boxed), 13, 4); +lean_closure_set(x_20, 0, x_1); +lean_closure_set(x_20, 1, x_16); +lean_closure_set(x_20, 2, x_2); +lean_closure_set(x_20, 3, x_13); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_22 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Elab_Deriving_mkInductArgNames___spec__2___rarg(x_20, x_21, x_6, x_7, x_8, x_9, x_10, x_11, x_18); -if (lean_obj_tag(x_22) == 0) +lean_inc(x_5); +x_21 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Elab_Deriving_mkInductArgNames___spec__2___rarg(x_19, x_20, x_5, x_6, x_7, x_8, x_9, x_10, x_17); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_22, 0); +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_21, 1); lean_inc(x_23); -x_24 = lean_ctor_get(x_22, 1); -lean_inc(x_24); -lean_dec(x_22); -x_25 = lean_array_push(x_5, x_23); -x_4 = x_15; -x_5 = x_25; -x_12 = x_24; +lean_dec(x_21); +x_24 = lean_array_push(x_4, x_22); +x_3 = x_14; +x_4 = x_24; +x_11 = x_23; goto _start; } else { -uint8_t x_27; -lean_dec(x_15); -lean_dec(x_11); +uint8_t x_26; +lean_dec(x_14); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_3); +lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_27 = !lean_is_exclusive(x_22); -if (x_27 == 0) +x_26 = !lean_is_exclusive(x_21); +if (x_26 == 0) { -return x_22; +return x_21; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_22, 0); -x_29 = lean_ctor_get(x_22, 1); -lean_inc(x_29); +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_21, 0); +x_28 = lean_ctor_get(x_21, 1); lean_inc(x_28); -lean_dec(x_22); -x_30 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -return x_30; +lean_inc(x_27); +lean_dec(x_21); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; } } } else { -uint8_t x_31; -lean_dec(x_15); +uint8_t x_30; lean_dec(x_14); -lean_dec(x_11); +lean_dec(x_13); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_3); +lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_31 = !lean_is_exclusive(x_16); -if (x_31 == 0) +x_30 = !lean_is_exclusive(x_15); +if (x_30 == 0) { -return x_16; +return x_15; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_16, 0); -x_33 = lean_ctor_get(x_16, 1); -lean_inc(x_33); +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_15, 0); +x_32 = lean_ctor_get(x_15, 1); lean_inc(x_32); -lean_dec(x_16); -x_34 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_34, 0, x_32); -lean_ctor_set(x_34, 1, x_33); -return x_34; +lean_inc(x_31); +lean_dec(x_15); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; } } } @@ -2695,7 +2730,7 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_1); -x_12 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5(x_1, x_2, x_11, x_10, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_12 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5(x_1, x_2, x_10, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9); if (lean_obj_tag(x_12) == 0) { lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; @@ -2791,19 +2826,20 @@ lean_dec(x_3); return x_9; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_13; -x_13 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); +lean_object* x_14; +x_14 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -return x_13; +lean_dec(x_1); +return x_14; } } LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___boxed(lean_object** _args) { @@ -2824,30 +2860,56 @@ lean_object* x_14 = _args[13]; lean_object* x_15 = _args[14]; lean_object* x_16 = _args[15]; lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; _start: { -lean_object* x_18; -x_18 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_object* x_19; +x_19 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_13); lean_dec(x_12); -lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -return x_18; +return x_19; } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___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, 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_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___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, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_15; -x_15 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_9); +lean_object* x_14; +x_14 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_8); -lean_dec(x_6); -return x_15; +lean_dec(x_7); +lean_dec(x_5); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_BEq_mkMatch___spec__1(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = 1; +x_9 = lean_usize_add(x_2, x_8); +x_10 = lean_array_uset(x_7, x_2, x_5); +x_2 = x_9; +x_3 = x_10; +goto _start; +} } } static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkMatch___closed__1() { @@ -2963,7 +3025,7 @@ x_25 = lean_array_get_size(x_12); x_26 = lean_usize_of_nat(x_25); lean_dec(x_25); x_27 = 0; -x_28 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_26, x_27, x_12); +x_28 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_BEq_mkMatch___spec__1(x_26, x_27, x_12); x_29 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__17; x_30 = l_Lean_mkSepArray(x_28, x_29); lean_dec(x_28); @@ -3022,7 +3084,7 @@ x_57 = lean_array_get_size(x_12); x_58 = lean_usize_of_nat(x_57); lean_dec(x_57); x_59 = 0; -x_60 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_58, x_59, x_12); +x_60 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_BEq_mkMatch___spec__1(x_58, x_59, x_12); x_61 = l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__17; x_62 = l_Lean_mkSepArray(x_60, x_61); lean_dec(x_60); @@ -3096,6 +3158,18 @@ return x_90; } } } +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_BEq_mkMatch___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_BEq_mkMatch___spec__1(x_4, x_5, x_3); +return x_6; +} +} static lean_object* _init_l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__1() { _start: { @@ -4533,7 +4607,7 @@ lean_ctor_set(x_34, 0, x_18); lean_ctor_set(x_34, 1, x_32); lean_ctor_set(x_34, 2, x_31); lean_ctor_set(x_34, 3, x_33); -x_35 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__3; +x_35 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__6; lean_inc(x_18); x_36 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_36, 0, x_18); @@ -4607,7 +4681,7 @@ lean_ctor_set(x_72, 0, x_18); lean_ctor_set(x_72, 1, x_70); lean_ctor_set(x_72, 2, x_69); lean_ctor_set(x_72, 3, x_71); -x_73 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__3; +x_73 = l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__6; lean_inc(x_18); x_74 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_74, 0, x_18); @@ -6082,7 +6156,7 @@ x_9 = lean_ctor_get(x_1, 1); lean_inc(x_9); lean_dec(x_1); lean_inc(x_2); -x_17 = l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__10(x_8, x_2, x_3, x_4); +x_17 = l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__14(x_8, x_2, x_3, x_4); if (lean_obj_tag(x_17) == 0) { lean_object* x_18; @@ -6173,7 +6247,7 @@ LEAN_EXPORT lean_object* l_Lean_isEnumType___at_Lean_Elab_Deriving_BEq_mkBEqInst { lean_object* x_5; lean_inc(x_2); -x_5 = l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__10(x_1, x_2, x_3, x_4); +x_5 = l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__14(x_1, x_2, x_3, x_4); if (lean_obj_tag(x_5) == 0) { lean_object* x_6; @@ -7446,7 +7520,7 @@ lean_dec(x_1); return x_9; } } -static lean_object* _init_l_Lean_Elab_Deriving_BEq_initFn____x40_Lean_Elab_Deriving_BEq___hyg_3222____closed__1() { +static lean_object* _init_l_Lean_Elab_Deriving_BEq_initFn____x40_Lean_Elab_Deriving_BEq___hyg_3410____closed__1() { _start: { lean_object* x_1; @@ -7454,12 +7528,12 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Deriving_BEq_mkBEqInstanceHandler), return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_BEq_initFn____x40_Lean_Elab_Deriving_BEq___hyg_3222_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_BEq_initFn____x40_Lean_Elab_Deriving_BEq___hyg_3410_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; x_2 = l_Lean_Elab_Deriving_BEq_mkBEqHeader___closed__2; -x_3 = l_Lean_Elab_Deriving_BEq_initFn____x40_Lean_Elab_Deriving_BEq___hyg_3222____closed__1; +x_3 = l_Lean_Elab_Deriving_BEq_initFn____x40_Lean_Elab_Deriving_BEq___hyg_3410____closed__1; x_4 = l_Lean_Elab_registerBuiltinDerivingHandler(x_2, x_3, x_1); if (lean_obj_tag(x_4) == 0) { @@ -7638,6 +7712,14 @@ l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda_ lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__11); l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__12 = _init_l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__12(); lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__12); +l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__13 = _init_l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__13(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__13); +l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__14 = _init_l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__14(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__14); +l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__15 = _init_l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__15(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__15); +l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__16 = _init_l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__16(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__16); l_Lean_Elab_Deriving_BEq_mkMatch___closed__1 = _init_l_Lean_Elab_Deriving_BEq_mkMatch___closed__1(); lean_mark_persistent(l_Lean_Elab_Deriving_BEq_mkMatch___closed__1); l_Lean_Elab_Deriving_BEq_mkMatch___closed__2 = _init_l_Lean_Elab_Deriving_BEq_mkMatch___closed__2(); @@ -7808,9 +7890,9 @@ l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqEnumFun___clos lean_mark_persistent(l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqEnumFun___closed__22); l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqEnumFun___closed__23 = _init_l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqEnumFun___closed__23(); lean_mark_persistent(l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqEnumFun___closed__23); -l_Lean_Elab_Deriving_BEq_initFn____x40_Lean_Elab_Deriving_BEq___hyg_3222____closed__1 = _init_l_Lean_Elab_Deriving_BEq_initFn____x40_Lean_Elab_Deriving_BEq___hyg_3222____closed__1(); -lean_mark_persistent(l_Lean_Elab_Deriving_BEq_initFn____x40_Lean_Elab_Deriving_BEq___hyg_3222____closed__1); -res = l_Lean_Elab_Deriving_BEq_initFn____x40_Lean_Elab_Deriving_BEq___hyg_3222_(lean_io_mk_world()); +l_Lean_Elab_Deriving_BEq_initFn____x40_Lean_Elab_Deriving_BEq___hyg_3410____closed__1 = _init_l_Lean_Elab_Deriving_BEq_initFn____x40_Lean_Elab_Deriving_BEq___hyg_3410____closed__1(); +lean_mark_persistent(l_Lean_Elab_Deriving_BEq_initFn____x40_Lean_Elab_Deriving_BEq___hyg_3410____closed__1); +res = l_Lean_Elab_Deriving_BEq_initFn____x40_Lean_Elab_Deriving_BEq___hyg_3410_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/Deriving/Basic.c b/stage0/stdlib/Lean/Elab/Deriving/Basic.c index 7c24ea33cf6..cb8bc9efe3f 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/Basic.c +++ b/stage0/stdlib/Lean/Elab/Deriving/Basic.c @@ -14,6 +14,7 @@ extern "C" { #endif lean_object* l_Lean_initializing(lean_object*); +LEAN_EXPORT lean_object* l_Lean_resolveGlobalName___at_Lean_Elab_elabDeriving___spec__11(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_derivingHandlersRef; size_t lean_usize_add(size_t, size_t); lean_object* l_Lean_registerTraceClass(lean_object*, lean_object*); @@ -23,199 +24,204 @@ lean_object* l_Lean_stringToMessageData(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_elabDeriving_declRange___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_registerBuiltinDerivingHandlerWithArgs___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); -extern lean_object* l_Lean_nullKind; +LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__14(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1519____closed__2; lean_object* l_Lean_Elab_Command_liftTermElabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_name_mk_string(lean_object*, lean_object*); +static lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__12___closed__2; lean_object* lean_array_uget(lean_object*, size_t); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_getOptDerivingClasses___spec__1(lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1519____closed__1; static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_elabDeriving___spec__1___rarg___closed__2; lean_object* l_List_mapTRAux___at_Lean_resolveGlobalConstCore___spec__2(lean_object*, lean_object*); extern lean_object* l_Std_Format_defWidth; -LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_elabDeriving_declRange___closed__7; -LEAN_EXPORT lean_object* l_Lean_resolveGlobalName___at_Lean_Elab_elabDeriving___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_elabDeriving___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Elab_Command_commandElabAttribute; +LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_elabDeriving___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_filterMap___at_Lean_resolveGlobalConst___spec__1(lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_elabDeriving_declRange___closed__4; -LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at_Lean_Elab_elabDeriving___spec__6___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_environment_find(lean_object*, lean_object*); -lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___spec__4(size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_defaultHandler___closed__5; lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l_Lean_Elab_registerBuiltinDerivingHandlerWithArgs___closed__1; static lean_object* l_Lean_Elab_defaultHandler___closed__3; +LEAN_EXPORT lean_object* l_Lean_Elab_resolveGlobalConstNoOverloadWithInfo___at_Lean_Elab_elabDeriving___spec__6___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_elabDeriving_declRange___closed__1; -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1490____closed__3; static lean_object* l___regBuiltin_Lean_Elab_elabDeriving___closed__3; +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Elab_elabDeriving___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); +static lean_object* l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_elabDeriving___spec__7___closed__2; lean_object* lean_string_append(lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofList(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_getOptDerivingClasses(lean_object*); +static lean_object* l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_elabDeriving___spec__7___closed__1; static lean_object* l_Lean_Elab_elabDeriving___lambda__2___closed__1; static lean_object* l___regBuiltin_Lean_Elab_elabDeriving_declRange___closed__2; static lean_object* l_Lean_Elab_registerBuiltinDerivingHandlerWithArgs___lambda__2___closed__1; static lean_object* l___regBuiltin_Lean_Elab_elabDeriving___closed__6; lean_object* l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at_Lean_Elab_elabDeriving___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_NameMap_contains___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_getOptDerivingClasses___spec__1___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_elabDeriving___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_getOptDerivingClasses___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabDeriving___closed__7; static lean_object* l_Lean_Elab_defaultHandler___closed__6; lean_object* lean_nat_add(lean_object*, lean_object*); -lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___spec__3(size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at_Lean_Elab_elabDeriving___spec__13(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandDeclId___spec__13(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_resolveGlobalConstNoOverloadWithInfo___at_Lean_Elab_elabDeriving___spec__6___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_LocalContext_empty; +LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at_Lean_Elab_elabDeriving___spec__10___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_mapTRAux___at_Lean_resolveGlobalConstNoOverload___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_elabDeriving___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__4___closed__3; static lean_object* l_Lean_Elab_elabDeriving___closed__3; static lean_object* l_Lean_Elab_elabDeriving___lambda__2___closed__2; +static lean_object* l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__8___closed__2; lean_object* l_Lean_ConstantInfo_levelParams(lean_object*); +static lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__16___closed__3; static lean_object* l_Lean_Elab_elabDeriving___closed__5; static lean_object* l___regBuiltin_Lean_Elab_elabDeriving___closed__5; lean_object* lean_array_fget(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__12(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__8___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_applyDerivingHandlers(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_registerBuiltinDerivingHandlerWithArgs___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_getOptDerivingClasses___spec__1___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_elabDeriving___spec__5___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabDeriving___closed__2; +lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_elabDeriving___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1490_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_22_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_resolveGlobalName___at_Lean_Elab_elabDeriving___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__12___closed__2; +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1519_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_25_(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_elabDeriving_declRange___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_registerBuiltinDerivingHandler___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_toString(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_defaultHandler___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_resolveGlobalConstNoOverloadWithInfo___at_Lean_Elab_elabDeriving___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_replaceRef(lean_object*, lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoTree___at_Lean_Elab_elabDeriving___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_elabDeriving___spec__3___closed__1; -static lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__12___closed__1; +LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Elab_elabDeriving___spec__2(lean_object*, lean_object*); +static lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__12___closed__1; static lean_object* l_Lean_Elab_defaultHandler___closed__2; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_elabDeriving___spec__4___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabDeriving___closed__1; uint8_t l_Lean_Name_quickCmp(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_elabDeriving___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_registerBuiltinDerivingHandlerWithArgs___closed__2; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_elabDeriving___spec__4(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_elabDeriving_declRange(lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_elabDeriving___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabDeriving___closed__8; lean_object* lean_format_pretty(lean_object*, lean_object*); static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_elabDeriving___spec__1___rarg___closed__1; -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_elabDeriving___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabDeriving___closed__4; -static lean_object* l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_elabDeriving___spec__3___closed__2; +LEAN_EXPORT lean_object* l_Lean_resolveGlobalName___at_Lean_Elab_elabDeriving___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addBuiltinDeclarationRanges(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_defaultHandler___closed__1; lean_object* l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabDeriving___closed__9; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_elabDeriving___spec__15(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_RBNode_find___at_Lean_Elab_applyDerivingHandlers___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at_Lean_Elab_elabDeriving___spec__6(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_resolveGlobalConstNoOverloadWithInfo___at_Lean_Elab_elabDeriving___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_registerBuiltinDerivingHandlerWithArgs___lambda__2___closed__2; static lean_object* l_Lean_Elab_elabDeriving___closed__6; -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_elabDeriving___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoTree___at_Lean_Elab_elabDeriving___spec__13(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_elabDeriving___spec__15(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Array_sequenceMap___at_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convIntro______1___spec__1(lean_object*, lean_object*); lean_object* l_Lean_ResolveName_resolveGlobalName(lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_elabDeriving___spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at_Lean_Elab_elabDeriving___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_elabDeriving___spec__19(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_DerivingClassView_applyHandlers(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_resolveGlobalConstNoOverloadWithInfo___at_Lean_Elab_elabDeriving___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_RBNode_find___at_Lean_Elab_applyDerivingHandlers___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_getOptDerivingClasses___spec__1___rarg___lambda__1(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoTree___at_Lean_Elab_elabDeriving___spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabDeriving___lambda__1(lean_object*, lean_object*, lean_object*); -lean_object* l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabDeriving___lambda__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Std_PersistentArray_push___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabDeriving___lambda__3(lean_object*); -LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_getOptDerivingClasses___rarg___lambda__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_elabDeriving___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Deriving_Basic_0__Lean_Elab_tryApplyDefHandler___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_resolveGlobalConstNoOverloadWithInfo___at_Lean_Elab_elabDeriving___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_logException___at_Lean_Elab_Command_runLinters___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_getOptDerivingClasses___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_elabDeriving___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_mapTRAux___at_Lean_Elab_Term_CollectPatternVars_collect_processExplicitArg___spec__2(lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArgs(lean_object*); +LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Elab_elabDeriving___spec__2___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_elabDeriving___spec__5(size_t, size_t, lean_object*); lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandDeclId___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at_Lean_Elab_elabDeriving___spec__6___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Elab_elabDeriving___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_elabDeriving___spec__18(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_defaultHandler(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__8___closed__3; +LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__16(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1519____closed__3; lean_object* l_Lean_Elab_Term_processDefDeriving(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_resolveGlobalConstNoOverloadWithInfo___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_mapTRAux___at_Lean_mkConstWithLevelParams___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabDeriving___closed__12; lean_object* l_Lean_Elab_Command_getRef(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_elabDeriving_declRange___closed__6; +static lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__16___closed__1; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabDeriving___lambda__2(lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); extern lean_object* l_Lean_instInhabitedName; +LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isNone(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabDeriving(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_registerBuiltinDerivingHandler(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_getOptDerivingClasses___spec__1___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_elabDeriving___spec__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_elabDeriving___spec__15___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__12___closed__3; static lean_object* l___regBuiltin_Lean_Elab_elabDeriving___closed__4; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_elabDeriving___spec__14(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__12(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwError___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabDeriving___closed__11; static lean_object* l_Lean_Elab_getOptDerivingClasses___rarg___closed__2; +LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at_Lean_Elab_elabDeriving___spec__10(lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Elab_unsupportedSyntaxExceptionId; static lean_object* l_Lean_Elab_elabDeriving___closed__10; -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1490____closed__1; -lean_object* l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_registerBuiltinDerivingHandlerWithArgs___lambda__1___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_elabDeriving(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_registerBuiltinDerivingHandlerWithArgs___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1490____closed__2; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_elabDeriving___spec__15___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__16___closed__2; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_elabDeriving___spec__19___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_List_isEmpty___rarg(lean_object*); lean_object* l_Lean_Elab_Command_getScope___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_registerBuiltinDerivingHandlerWithArgs(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__8___closed__3; +LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoTree___at_Lean_Elab_elabDeriving___spec__17(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_elabDeriving___spec__19___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_toString___at_Lean_resolveGlobalConstNoOverloadCore___spec__2(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_elabDeriving___closed__1; -static lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__8___closed__1; -LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at_Lean_Elab_elabDeriving___spec__9(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_getOptDerivingClasses___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_mkConst(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__10(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__4___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Deriving_Basic_0__Lean_Elab_tryApplyDefHandler(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__12___closed__3; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_elabDeriving___spec__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_registerBuiltinDerivingHandler___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_formatStxAux(lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_elabDeriving___spec__1___rarg(lean_object*); lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at_Lean_Elab_elabDeriving___spec__10___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_getOptDerivingClasses___rarg___closed__1; static lean_object* l___regBuiltin_Lean_Elab_elabDeriving___closed__2; -static lean_object* l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__4___closed__1; lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_getOptDerivingClasses___spec__1___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_elabDeriving___spec__11(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__8___closed__2; -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_22_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_25_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; @@ -789,7 +795,141 @@ x_3 = lean_alloc_closure((void*)(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_El return x_3; } } -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_elabDeriving___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Elab_elabDeriving___spec__3(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; uint8_t x_7; +x_6 = lean_array_get_size(x_1); +x_7 = lean_nat_dec_lt(x_4, x_6); +lean_dec(x_6); +if (x_7 == 0) +{ +lean_object* x_8; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_5); +return x_8; +} +else +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_unsigned_to_nat(0u); +x_10 = lean_nat_dec_eq(x_3, x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_sub(x_3, x_11); +lean_dec(x_3); +x_13 = lean_array_fget(x_1, x_4); +lean_inc(x_2); +x_14 = lean_apply_1(x_2, x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; +lean_dec(x_12); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_15 = lean_box(0); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_14, 0); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_nat_add(x_4, x_11); +lean_dec(x_4); +x_18 = lean_array_push(x_5, x_16); +x_3 = x_12; +x_4 = x_17; +x_5 = x_18; +goto _start; +} +} +else +{ +lean_object* x_20; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_20 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_20, 0, x_5); +return x_20; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Elab_elabDeriving___spec__2(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; +x_3 = lean_array_get_size(x_1); +x_4 = lean_mk_empty_array_with_capacity(x_3); +x_5 = lean_unsigned_to_nat(0u); +x_6 = l_Array_sequenceMap_loop___at_Lean_Elab_elabDeriving___spec__3(x_1, x_2, x_3, x_5, x_4); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_elabDeriving___spec__4(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; lean_object* x_11; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = 1; +x_9 = lean_usize_add(x_2, x_8); +x_10 = lean_ctor_get(x_5, 1); +lean_inc(x_10); +lean_dec(x_5); +x_11 = lean_array_uset(x_7, x_2, x_10); +x_2 = x_9; +x_3 = x_11; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_elabDeriving___spec__5(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; lean_object* x_11; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = 1; +x_9 = lean_usize_add(x_2, x_8); +x_10 = lean_ctor_get(x_5, 0); +lean_inc(x_10); +lean_dec(x_5); +x_11 = lean_array_uset(x_7, x_2, x_10); +x_2 = x_9; +x_3 = x_11; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_elabDeriving___spec__9(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; lean_object* x_9; uint8_t x_10; @@ -843,7 +983,7 @@ return x_21; } } } -LEAN_EXPORT lean_object* l_Lean_resolveGlobalName___at_Lean_Elab_elabDeriving___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_resolveGlobalName___at_Lean_Elab_elabDeriving___spec__11(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; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; @@ -897,7 +1037,7 @@ return x_22; } } } -static lean_object* _init_l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__8___closed__1() { +static lean_object* _init_l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__12___closed__1() { _start: { lean_object* x_1; @@ -905,16 +1045,16 @@ x_1 = lean_mk_string_from_bytes("unknown constant '", 18); return x_1; } } -static lean_object* _init_l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__8___closed__2() { +static lean_object* _init_l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__12___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__8___closed__1; +x_1 = l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__12___closed__1; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__8___closed__3() { +static lean_object* _init_l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__12___closed__3() { _start: { lean_object* x_1; lean_object* x_2; @@ -923,7 +1063,7 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__12(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; lean_object* x_11; lean_object* x_12; @@ -931,11 +1071,11 @@ x_5 = lean_box(0); x_6 = l_Lean_mkConst(x_1, x_5); x_7 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_7, 0, x_6); -x_8 = l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__8___closed__2; +x_8 = l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__12___closed__2; x_9 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_9, 0, x_8); lean_ctor_set(x_9, 1, x_7); -x_10 = l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__8___closed__3; +x_10 = l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__12___closed__3; x_11 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_11, 0, x_9); lean_ctor_set(x_11, 1, x_10); @@ -943,7 +1083,7 @@ x_12 = l_Lean_throwError___at_Lean_Elab_Command_expandDeclId___spec__4(x_11, x_2 return x_12; } } -LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at_Lean_Elab_elabDeriving___spec__6___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_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at_Lean_Elab_elabDeriving___spec__10___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: { lean_object* x_7; lean_object* x_8; @@ -954,26 +1094,26 @@ lean_ctor_set(x_8, 1, x_6); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at_Lean_Elab_elabDeriving___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at_Lean_Elab_elabDeriving___spec__10(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; uint8_t x_10; lean_inc(x_1); -x_5 = l_Lean_resolveGlobalName___at_Lean_Elab_elabDeriving___spec__7(x_1, x_2, x_3, x_4); +x_5 = l_Lean_resolveGlobalName___at_Lean_Elab_elabDeriving___spec__11(x_1, x_2, x_3, x_4); x_6 = lean_ctor_get(x_5, 0); lean_inc(x_6); x_7 = lean_ctor_get(x_5, 1); lean_inc(x_7); lean_dec(x_5); x_8 = lean_box(0); -x_9 = l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(x_6, x_8); +x_9 = l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(x_6, x_8); x_10 = l_List_isEmpty___rarg(x_9); if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_dec(x_1); x_11 = lean_box(0); -x_12 = l_Lean_resolveGlobalConstCore___at_Lean_Elab_elabDeriving___spec__6___lambda__1(x_9, x_8, x_11, x_2, x_3, x_7); +x_12 = l_Lean_resolveGlobalConstCore___at_Lean_Elab_elabDeriving___spec__10___lambda__1(x_9, x_8, x_11, x_2, x_3, x_7); lean_dec(x_3); lean_dec(x_2); return x_12; @@ -982,7 +1122,7 @@ else { lean_object* x_13; uint8_t x_14; lean_dec(x_9); -x_13 = l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__8(x_1, x_2, x_3, x_7); +x_13 = l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__12(x_1, x_2, x_3, x_7); lean_dec(x_3); x_14 = !lean_is_exclusive(x_13); if (x_14 == 0) @@ -1005,7 +1145,7 @@ return x_17; } } } -static lean_object* _init_l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__4___closed__1() { +static lean_object* _init_l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__8___closed__1() { _start: { lean_object* x_1; @@ -1013,27 +1153,27 @@ x_1 = lean_mk_string_from_bytes("expected identifier", 19); return x_1; } } -static lean_object* _init_l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__4___closed__2() { +static lean_object* _init_l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__8___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__4___closed__1; +x_1 = l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__8___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_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__4___closed__3() { +static lean_object* _init_l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__8___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__4___closed__2; +x_1 = l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__8___closed__2; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { if (lean_obj_tag(x_1) == 3) @@ -1077,7 +1217,7 @@ lean_object* x_15; lean_object* x_16; x_15 = lean_ctor_get(x_2, 6); lean_dec(x_15); lean_ctor_set(x_2, 6, x_13); -x_16 = l_Lean_resolveGlobalConstCore___at_Lean_Elab_elabDeriving___spec__6(x_5, x_2, x_3, x_12); +x_16 = l_Lean_resolveGlobalConstCore___at_Lean_Elab_elabDeriving___spec__10(x_5, x_2, x_3, x_12); return x_16; } else @@ -1107,7 +1247,7 @@ lean_ctor_set(x_24, 4, x_21); lean_ctor_set(x_24, 5, x_22); lean_ctor_set(x_24, 6, x_13); lean_ctor_set(x_24, 7, x_23); -x_25 = l_Lean_resolveGlobalConstCore___at_Lean_Elab_elabDeriving___spec__6(x_5, x_24, x_3, x_12); +x_25 = l_Lean_resolveGlobalConstCore___at_Lean_Elab_elabDeriving___spec__10(x_5, x_24, x_3, x_12); return x_25; } } @@ -1115,15 +1255,15 @@ return x_25; else { lean_object* x_26; lean_object* x_27; -x_26 = l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__4___closed__3; -x_27 = l_Lean_throwErrorAt___at_Lean_Elab_elabDeriving___spec__5(x_1, x_26, x_2, x_3, x_4); +x_26 = l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__8___closed__3; +x_27 = l_Lean_throwErrorAt___at_Lean_Elab_elabDeriving___spec__9(x_1, x_26, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_1); return x_27; } } } -static lean_object* _init_l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_elabDeriving___spec__3___closed__1() { +static lean_object* _init_l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_elabDeriving___spec__7___closed__1() { _start: { lean_object* x_1; @@ -1131,7 +1271,7 @@ x_1 = lean_mk_string_from_bytes("ambiguous identifier '", 22); return x_1; } } -static lean_object* _init_l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_elabDeriving___spec__3___closed__2() { +static lean_object* _init_l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_elabDeriving___spec__7___closed__2() { _start: { lean_object* x_1; @@ -1139,14 +1279,14 @@ x_1 = lean_mk_string_from_bytes("', possible interpretations: ", 29); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_elabDeriving___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_elabDeriving___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_5 = l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__4(x_1, x_2, x_3, x_4); +x_5 = l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__8(x_1, x_2, x_3, x_4); if (lean_obj_tag(x_5) == 0) { lean_object* x_6; @@ -1165,10 +1305,10 @@ lean_inc(x_1); x_11 = l_Lean_Syntax_formatStxAux(x_8, x_9, x_10, x_1); x_12 = l_Std_Format_defWidth; x_13 = lean_format_pretty(x_11, x_12); -x_14 = l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_elabDeriving___spec__3___closed__1; +x_14 = l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_elabDeriving___spec__7___closed__1; x_15 = lean_string_append(x_14, x_13); lean_dec(x_13); -x_16 = l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_elabDeriving___spec__3___closed__2; +x_16 = l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_elabDeriving___spec__7___closed__2; x_17 = lean_string_append(x_15, x_16); x_18 = lean_box(0); x_19 = l_List_mapTRAux___at_Lean_resolveGlobalConstNoOverload___spec__1(x_6, x_18); @@ -1236,10 +1376,10 @@ lean_inc(x_1); x_38 = l_Lean_Syntax_formatStxAux(x_35, x_36, x_37, x_1); x_39 = l_Std_Format_defWidth; x_40 = lean_format_pretty(x_38, x_39); -x_41 = l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_elabDeriving___spec__3___closed__1; +x_41 = l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_elabDeriving___spec__7___closed__1; x_42 = lean_string_append(x_41, x_40); lean_dec(x_40); -x_43 = l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_elabDeriving___spec__3___closed__2; +x_43 = l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_elabDeriving___spec__7___closed__2; x_44 = lean_string_append(x_42, x_43); x_45 = lean_box(0); x_46 = l_List_mapTRAux___at_Lean_resolveGlobalConstNoOverload___spec__1(x_6, x_45); @@ -1284,7 +1424,7 @@ return x_57; } } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_elabDeriving___spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_elabDeriving___spec__15(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; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; @@ -1338,7 +1478,7 @@ return x_20; } } } -LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__14(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; @@ -1362,15 +1502,15 @@ x_11 = lean_box(0); x_12 = l_Lean_mkConst(x_1, x_11); x_13 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_13, 0, x_12); -x_14 = l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__8___closed__2; +x_14 = l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__12___closed__2; x_15 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_15, 0, x_14); lean_ctor_set(x_15, 1, x_13); -x_16 = l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__8___closed__3; +x_16 = l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__12___closed__3; x_17 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_17, 0, x_15); lean_ctor_set(x_17, 1, x_16); -x_18 = l_Lean_throwError___at_Lean_Elab_elabDeriving___spec__11(x_17, x_2, x_3, x_8); +x_18 = l_Lean_throwError___at_Lean_Elab_elabDeriving___spec__15(x_17, x_2, x_3, x_8); return x_18; } else @@ -1405,15 +1545,15 @@ x_24 = lean_box(0); x_25 = l_Lean_mkConst(x_1, x_24); x_26 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_26, 0, x_25); -x_27 = l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__8___closed__2; +x_27 = l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__12___closed__2; x_28 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_28, 0, x_27); lean_ctor_set(x_28, 1, x_26); -x_29 = l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__8___closed__3; +x_29 = l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__12___closed__3; x_30 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_30, 0, x_28); lean_ctor_set(x_30, 1, x_29); -x_31 = l_Lean_throwError___at_Lean_Elab_elabDeriving___spec__11(x_30, x_2, x_3, x_21); +x_31 = l_Lean_throwError___at_Lean_Elab_elabDeriving___spec__15(x_30, x_2, x_3, x_21); return x_31; } else @@ -1432,12 +1572,12 @@ return x_33; } } } -LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at_Lean_Elab_elabDeriving___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at_Lean_Elab_elabDeriving___spec__13(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_inc(x_1); -x_5 = l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__10(x_1, x_2, x_3, x_4); +x_5 = l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__14(x_1, x_2, x_3, x_4); if (lean_obj_tag(x_5) == 0) { uint8_t x_6; @@ -1498,7 +1638,7 @@ return x_22; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoTree___at_Lean_Elab_elabDeriving___spec__13(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoTree___at_Lean_Elab_elabDeriving___spec__17(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; uint8_t x_8; @@ -1701,7 +1841,7 @@ return x_61; } } } -static lean_object* _init_l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__12___closed__1() { +static lean_object* _init_l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__16___closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -1710,23 +1850,23 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__12___closed__2() { +static lean_object* _init_l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__16___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__12___closed__1; +x_1 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__16___closed__1; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__12___closed__3() { +static lean_object* _init_l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__16___closed__3() { _start: { size_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = 5; -x_2 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__12___closed__2; -x_3 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__12___closed__1; +x_2 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__16___closed__2; +x_3 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__16___closed__1; x_4 = lean_unsigned_to_nat(0u); x_5 = lean_alloc_ctor(0, 4, sizeof(size_t)*1); lean_ctor_set(x_5, 0, x_2); @@ -1737,7 +1877,7 @@ lean_ctor_set_usize(x_5, 4, x_1); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__16(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; uint8_t x_8; @@ -1782,16 +1922,16 @@ lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; x_15 = lean_ctor_get(x_5, 1); lean_inc(x_15); lean_dec(x_5); -x_16 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__12___closed__3; +x_16 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__16___closed__3; x_17 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_17, 0, x_1); lean_ctor_set(x_17, 1, x_16); -x_18 = l_Lean_Elab_pushInfoTree___at_Lean_Elab_elabDeriving___spec__13(x_17, x_2, x_3, x_15); +x_18 = l_Lean_Elab_pushInfoTree___at_Lean_Elab_elabDeriving___spec__17(x_17, x_2, x_3, x_15); return x_18; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_resolveGlobalConstNoOverloadWithInfo___at_Lean_Elab_elabDeriving___spec__2___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Elab_resolveGlobalConstNoOverloadWithInfo___at_Lean_Elab_elabDeriving___spec__6___lambda__1(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; @@ -1801,14 +1941,14 @@ lean_ctor_set(x_6, 1, x_5); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Elab_resolveGlobalConstNoOverloadWithInfo___at_Lean_Elab_elabDeriving___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Elab_resolveGlobalConstNoOverloadWithInfo___at_Lean_Elab_elabDeriving___spec__6(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_inc(x_4); lean_inc(x_3); lean_inc(x_1); -x_6 = l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_elabDeriving___spec__3(x_1, x_3, x_4, x_5); +x_6 = l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_elabDeriving___spec__7(x_1, x_3, x_4, x_5); if (lean_obj_tag(x_6) == 0) { lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; @@ -1861,7 +2001,7 @@ lean_inc(x_17); lean_dec(x_9); lean_inc(x_3); lean_inc(x_7); -x_18 = l_Lean_mkConstWithLevelParams___at_Lean_Elab_elabDeriving___spec__9(x_7, x_3, x_4, x_17); +x_18 = l_Lean_mkConstWithLevelParams___at_Lean_Elab_elabDeriving___spec__13(x_7, x_3, x_4, x_17); if (lean_obj_tag(x_18) == 0) { lean_object* x_19; 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; uint8_t x_28; @@ -1884,7 +2024,7 @@ lean_ctor_set(x_25, 3, x_19); lean_ctor_set_uint8(x_25, sizeof(void*)*4, x_24); x_26 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_26, 0, x_25); -x_27 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__12(x_26, x_3, x_4, x_20); +x_27 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__16(x_26, x_3, x_4, x_20); lean_dec(x_4); lean_dec(x_3); x_28 = !lean_is_exclusive(x_27); @@ -1965,7 +2105,7 @@ return x_39; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_elabDeriving___spec__14(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_elabDeriving___spec__18(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { uint8_t x_7; @@ -1989,7 +2129,7 @@ x_11 = lean_array_uset(x_3, x_2, x_10); x_12 = lean_box(0); lean_inc(x_5); lean_inc(x_4); -x_13 = l_Lean_Elab_resolveGlobalConstNoOverloadWithInfo___at_Lean_Elab_elabDeriving___spec__2(x_9, x_12, x_4, x_5, x_6); +x_13 = l_Lean_Elab_resolveGlobalConstNoOverloadWithInfo___at_Lean_Elab_elabDeriving___spec__6(x_9, x_12, x_4, x_5, x_6); if (lean_obj_tag(x_13) == 0) { lean_object* x_14; lean_object* x_15; size_t x_16; size_t x_17; lean_object* x_18; @@ -2034,7 +2174,7 @@ return x_23; } } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_elabDeriving___spec__15___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_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_elabDeriving___spec__19___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) { _start: { lean_object* x_8; @@ -2042,7 +2182,7 @@ x_8 = l_Lean_Elab_applyDerivingHandlers(x_1, x_2, x_3, x_5, x_6, x_7); return x_8; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_elabDeriving___spec__15(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_forInUnsafe_loop___at_Lean_Elab_elabDeriving___spec__19(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; @@ -2106,7 +2246,7 @@ x_55 = lean_box(0); lean_inc(x_7); lean_inc(x_6); lean_inc(x_11); -x_56 = l_Lean_Elab_resolveGlobalConstNoOverloadWithInfo___at_Lean_Elab_elabDeriving___spec__2(x_11, x_55, x_6, x_7, x_8); +x_56 = l_Lean_Elab_resolveGlobalConstNoOverloadWithInfo___at_Lean_Elab_elabDeriving___spec__6(x_11, x_55, x_6, x_7, x_8); if (lean_obj_tag(x_56) == 0) { lean_object* x_57; lean_object* x_58; lean_object* x_59; uint8_t x_60; lean_object* x_61; lean_object* x_62; @@ -2463,7 +2603,7 @@ x_139 = lean_box(0); lean_inc(x_7); lean_inc(x_6); lean_inc(x_11); -x_140 = l_Lean_Elab_resolveGlobalConstNoOverloadWithInfo___at_Lean_Elab_elabDeriving___spec__2(x_11, x_139, x_6, x_7, x_8); +x_140 = l_Lean_Elab_resolveGlobalConstNoOverloadWithInfo___at_Lean_Elab_elabDeriving___spec__6(x_11, x_139, x_6, x_7, x_8); if (lean_obj_tag(x_140) == 0) { lean_object* x_141; lean_object* x_142; lean_object* x_143; uint8_t x_144; lean_object* x_145; lean_object* x_146; @@ -2934,39 +3074,38 @@ lean_dec(x_1); x_9 = l_Lean_Syntax_isNone(x_8); if (x_9 == 0) { -lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l_Lean_nullKind; -x_11 = lean_unsigned_to_nat(2u); +lean_object* x_10; uint8_t x_11; +x_10 = lean_unsigned_to_nat(2u); lean_inc(x_8); -x_12 = l_Lean_Syntax_isNodeOf(x_8, x_10, x_11); -if (x_12 == 0) +x_11 = l_Lean_Syntax_matchesNull(x_8, x_10); +if (x_11 == 0) { -lean_object* x_13; +lean_object* x_12; lean_dec(x_8); lean_dec(x_6); -x_13 = lean_box(0); -return x_13; +x_12 = lean_box(0); +return x_12; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_14 = l_Lean_Syntax_getArg(x_8, x_7); +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = l_Lean_Syntax_getArg(x_8, x_7); lean_dec(x_8); -x_15 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_15, 0, x_14); -x_16 = lean_box(0); -x_17 = l_Lean_Elab_elabDeriving___lambda__1(x_6, x_16, x_15); -return x_17; +x_14 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, x_13); +x_15 = lean_box(0); +x_16 = l_Lean_Elab_elabDeriving___lambda__1(x_6, x_15, x_14); +return x_16; } } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_dec(x_8); +x_17 = lean_box(0); x_18 = lean_box(0); -x_19 = lean_box(0); -x_20 = l_Lean_Elab_elabDeriving___lambda__1(x_6, x_19, x_18); -return x_20; +x_19 = l_Lean_Elab_elabDeriving___lambda__1(x_6, x_18, x_17); +return x_19; } } } @@ -3158,7 +3297,7 @@ goto block_66; { lean_object* x_15; lean_object* x_16; x_15 = l_Lean_Elab_elabDeriving___closed__11; -x_16 = l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___spec__1(x_14, x_15); +x_16 = l_Array_sequenceMap___at_Lean_Elab_elabDeriving___spec__2(x_14, x_15); lean_dec(x_14); if (lean_obj_tag(x_16) == 0) { @@ -3180,8 +3319,8 @@ x_20 = lean_usize_of_nat(x_19); lean_dec(x_19); x_21 = 0; lean_inc(x_18); -x_22 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___spec__3(x_20, x_21, x_18); -x_23 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___spec__4(x_20, x_21, x_18); +x_22 = l_Array_mapMUnsafe_map___at_Lean_Elab_elabDeriving___spec__4(x_20, x_21, x_18); +x_23 = l_Array_mapMUnsafe_map___at_Lean_Elab_elabDeriving___spec__5(x_20, x_21, x_18); x_24 = lean_unsigned_to_nat(4u); x_25 = l_Lean_Syntax_getArg(x_1, x_24); lean_dec(x_1); @@ -3230,7 +3369,7 @@ goto block_58; { lean_object* x_30; lean_object* x_31; x_30 = l_Lean_Elab_elabDeriving___closed__12; -x_31 = l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__1(x_29, x_30); +x_31 = l_Array_sequenceMap___at_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convIntro______1___spec__1(x_29, x_30); lean_dec(x_29); if (lean_obj_tag(x_31) == 0) { @@ -3253,7 +3392,7 @@ x_35 = lean_usize_of_nat(x_34); lean_dec(x_34); lean_inc(x_3); lean_inc(x_2); -x_36 = l_Array_mapMUnsafe_map___at_Lean_Elab_elabDeriving___spec__14(x_35, x_21, x_33, x_2, x_3, x_4); +x_36 = l_Array_mapMUnsafe_map___at_Lean_Elab_elabDeriving___spec__18(x_35, x_21, x_33, x_2, x_3, x_4); if (lean_obj_tag(x_36) == 0) { lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; size_t x_42; lean_object* x_43; @@ -3267,7 +3406,7 @@ x_40 = l_Array_toSubarray___rarg(x_22, x_12, x_39); x_41 = lean_array_get_size(x_23); x_42 = lean_usize_of_nat(x_41); lean_dec(x_41); -x_43 = l_Array_forInUnsafe_loop___at_Lean_Elab_elabDeriving___spec__15(x_37, x_23, x_42, x_21, x_40, x_2, x_3, x_38); +x_43 = l_Array_forInUnsafe_loop___at_Lean_Elab_elabDeriving___spec__19(x_37, x_23, x_42, x_21, x_40, x_2, x_3, x_38); lean_dec(x_23); if (lean_obj_tag(x_43) == 0) { @@ -3361,105 +3500,147 @@ lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_elabDeriving___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Elab_elabDeriving___spec__3___boxed(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; +x_6 = l_Array_sequenceMap_loop___at_Lean_Elab_elabDeriving___spec__3(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Elab_elabDeriving___spec__2___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Array_sequenceMap___at_Lean_Elab_elabDeriving___spec__2(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_elabDeriving___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_elabDeriving___spec__4(x_4, x_5, x_3); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_elabDeriving___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_elabDeriving___spec__5(x_4, x_5, x_3); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_elabDeriving___spec__9___boxed(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; -x_6 = l_Lean_throwErrorAt___at_Lean_Elab_elabDeriving___spec__5(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_throwErrorAt___at_Lean_Elab_elabDeriving___spec__9(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); lean_dec(x_1); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_resolveGlobalName___at_Lean_Elab_elabDeriving___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_resolveGlobalName___at_Lean_Elab_elabDeriving___spec__11___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_resolveGlobalName___at_Lean_Elab_elabDeriving___spec__7(x_1, x_2, x_3, x_4); +x_5 = l_Lean_resolveGlobalName___at_Lean_Elab_elabDeriving___spec__11(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_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__12___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_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__8(x_1, x_2, x_3, x_4); +x_5 = l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__12(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at_Lean_Elab_elabDeriving___spec__6___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_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at_Lean_Elab_elabDeriving___spec__10___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) { _start: { lean_object* x_7; -x_7 = l_Lean_resolveGlobalConstCore___at_Lean_Elab_elabDeriving___spec__6___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_Lean_resolveGlobalConstCore___at_Lean_Elab_elabDeriving___spec__10___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_elabDeriving___spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_elabDeriving___spec__15___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_throwError___at_Lean_Elab_elabDeriving___spec__11(x_1, x_2, x_3, x_4); +x_5 = l_Lean_throwError___at_Lean_Elab_elabDeriving___spec__15(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__14___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_getConstInfo___at_Lean_Elab_elabDeriving___spec__10(x_1, x_2, x_3, x_4); +x_5 = l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__14(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at_Lean_Elab_elabDeriving___spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at_Lean_Elab_elabDeriving___spec__13___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_mkConstWithLevelParams___at_Lean_Elab_elabDeriving___spec__9(x_1, x_2, x_3, x_4); +x_5 = l_Lean_mkConstWithLevelParams___at_Lean_Elab_elabDeriving___spec__13(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoTree___at_Lean_Elab_elabDeriving___spec__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoTree___at_Lean_Elab_elabDeriving___spec__17___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_Elab_pushInfoTree___at_Lean_Elab_elabDeriving___spec__13(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Elab_pushInfoTree___at_Lean_Elab_elabDeriving___spec__17(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_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__16___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_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__12(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__16(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_Elab_resolveGlobalConstNoOverloadWithInfo___at_Lean_Elab_elabDeriving___spec__2___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Elab_resolveGlobalConstNoOverloadWithInfo___at_Lean_Elab_elabDeriving___spec__6___lambda__1___boxed(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; -x_6 = l_Lean_Elab_resolveGlobalConstNoOverloadWithInfo___at_Lean_Elab_elabDeriving___spec__2___lambda__1(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Elab_resolveGlobalConstNoOverloadWithInfo___at_Lean_Elab_elabDeriving___spec__6___lambda__1(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); return x_6; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_elabDeriving___spec__14___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_mapMUnsafe_map___at_Lean_Elab_elabDeriving___spec__18___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: { size_t x_7; size_t x_8; lean_object* x_9; @@ -3467,20 +3648,20 @@ x_7 = lean_unbox_usize(x_1); lean_dec(x_1); x_8 = lean_unbox_usize(x_2); lean_dec(x_2); -x_9 = l_Array_mapMUnsafe_map___at_Lean_Elab_elabDeriving___spec__14(x_7, x_8, x_3, x_4, x_5, x_6); +x_9 = l_Array_mapMUnsafe_map___at_Lean_Elab_elabDeriving___spec__18(x_7, x_8, x_3, x_4, x_5, x_6); return x_9; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_elabDeriving___spec__15___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_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_elabDeriving___spec__19___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) { _start: { lean_object* x_8; -x_8 = l_Array_forInUnsafe_loop___at_Lean_Elab_elabDeriving___spec__15___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_8 = l_Array_forInUnsafe_loop___at_Lean_Elab_elabDeriving___spec__19___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_4); return x_8; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_elabDeriving___spec__15___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_forInUnsafe_loop___at_Lean_Elab_elabDeriving___spec__19___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; @@ -3488,7 +3669,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_forInUnsafe_loop___at_Lean_Elab_elabDeriving___spec__15(x_1, x_2, x_9, x_10, x_5, x_6, x_7, x_8); +x_11 = l_Array_forInUnsafe_loop___at_Lean_Elab_elabDeriving___spec__19(x_1, x_2, x_9, x_10, x_5, x_6, x_7, x_8); lean_dec(x_2); return x_11; } @@ -4068,134 +4249,133 @@ return x_12; } else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; x_13 = lean_unsigned_to_nat(0u); x_14 = l_Lean_Syntax_getArg(x_6, x_13); lean_dec(x_6); -x_15 = l_Lean_nullKind; -x_16 = lean_unsigned_to_nat(2u); +x_15 = lean_unsigned_to_nat(2u); lean_inc(x_14); -x_17 = l_Lean_Syntax_isNodeOf(x_14, x_15, x_16); -if (x_17 == 0) +x_16 = l_Lean_Syntax_matchesNull(x_14, x_15); +if (x_16 == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_dec(x_14); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_18 = lean_ctor_get(x_1, 0); -lean_inc(x_18); +x_17 = lean_ctor_get(x_1, 0); +lean_inc(x_17); lean_dec(x_1); -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -lean_dec(x_18); -x_20 = l_Lean_Elab_elabDeriving___closed__9; -x_21 = lean_apply_2(x_19, lean_box(0), x_20); -return x_21; +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); +lean_dec(x_17); +x_19 = l_Lean_Elab_elabDeriving___closed__9; +x_20 = lean_apply_2(x_18, lean_box(0), x_19); +return x_20; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; lean_object* x_27; -x_22 = lean_unsigned_to_nat(1u); -x_23 = l_Lean_Syntax_getArg(x_14, x_22); +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; +x_21 = lean_unsigned_to_nat(1u); +x_22 = l_Lean_Syntax_getArg(x_14, x_21); lean_dec(x_14); -x_24 = l_Lean_Syntax_getArgs(x_23); -lean_dec(x_23); -x_25 = lean_array_get_size(x_24); -x_26 = lean_nat_dec_lt(x_13, x_25); -if (x_26 == 0) +x_23 = l_Lean_Syntax_getArgs(x_22); +lean_dec(x_22); +x_24 = lean_array_get_size(x_23); +x_25 = lean_nat_dec_lt(x_13, x_24); +if (x_25 == 0) { -lean_object* x_51; -lean_dec(x_25); +lean_object* x_50; lean_dec(x_24); -x_51 = l_Lean_Elab_elabDeriving___closed__9; -x_27 = x_51; -goto block_50; +lean_dec(x_23); +x_50 = l_Lean_Elab_elabDeriving___closed__9; +x_26 = x_50; +goto block_49; } else { -uint8_t x_52; -x_52 = lean_nat_dec_le(x_25, x_25); -if (x_52 == 0) +uint8_t x_51; +x_51 = lean_nat_dec_le(x_24, x_24); +if (x_51 == 0) { -lean_object* x_53; -lean_dec(x_25); +lean_object* x_52; lean_dec(x_24); -x_53 = l_Lean_Elab_elabDeriving___closed__9; -x_27 = x_53; -goto block_50; +lean_dec(x_23); +x_52 = l_Lean_Elab_elabDeriving___closed__9; +x_26 = x_52; +goto block_49; } else { -size_t x_54; size_t x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_54 = 0; -x_55 = lean_usize_of_nat(x_25); -lean_dec(x_25); -x_56 = l_Lean_Elab_elabDeriving___closed__10; -x_57 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(x_24, x_54, x_55, x_56); +size_t x_53; size_t x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_53 = 0; +x_54 = lean_usize_of_nat(x_24); lean_dec(x_24); -x_58 = lean_ctor_get(x_57, 1); -lean_inc(x_58); -lean_dec(x_57); -x_27 = x_58; -goto block_50; +x_55 = l_Lean_Elab_elabDeriving___closed__10; +x_56 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(x_23, x_53, x_54, x_55); +lean_dec(x_23); +x_57 = lean_ctor_get(x_56, 1); +lean_inc(x_57); +lean_dec(x_56); +x_26 = x_57; +goto block_49; } } -block_50: +block_49: { -lean_object* x_28; lean_object* x_29; -x_28 = l_Lean_Elab_elabDeriving___closed__11; -x_29 = l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___spec__1(x_27, x_28); -lean_dec(x_27); -if (lean_obj_tag(x_29) == 0) +lean_object* x_27; lean_object* x_28; +x_27 = l_Lean_Elab_elabDeriving___closed__11; +x_28 = l_Array_sequenceMap___at_Lean_Elab_elabDeriving___spec__2(x_26, x_27); +lean_dec(x_26); +if (lean_obj_tag(x_28) == 0) { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_30 = lean_ctor_get(x_1, 0); -lean_inc(x_30); +x_29 = lean_ctor_get(x_1, 0); +lean_inc(x_29); lean_dec(x_1); -x_31 = lean_ctor_get(x_30, 1); -lean_inc(x_31); -lean_dec(x_30); -x_32 = l_Lean_Elab_elabDeriving___closed__9; -x_33 = lean_apply_2(x_31, lean_box(0), x_32); -return x_33; +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +lean_dec(x_29); +x_31 = l_Lean_Elab_elabDeriving___closed__9; +x_32 = lean_apply_2(x_30, lean_box(0), x_31); +return x_32; } else { -lean_object* x_34; lean_object* x_35; size_t x_36; size_t 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; lean_object* x_44; lean_object* x_45; size_t x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_34 = lean_ctor_get(x_29, 0); -lean_inc(x_34); -lean_dec(x_29); -x_35 = lean_array_get_size(x_34); -x_36 = lean_usize_of_nat(x_35); -lean_dec(x_35); -x_37 = 0; -lean_inc(x_34); -x_38 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___spec__3(x_36, x_37, x_34); -x_39 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___spec__4(x_36, x_37, x_34); -x_40 = lean_array_get_size(x_38); -x_41 = l_Array_toSubarray___rarg(x_38, x_13, x_40); -x_42 = lean_ctor_get(x_1, 1); -lean_inc(x_42); -x_43 = l_Lean_Elab_elabDeriving___closed__9; -x_44 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_44, 0, x_41); -lean_ctor_set(x_44, 1, x_43); -x_45 = lean_array_get_size(x_39); -x_46 = lean_usize_of_nat(x_45); -lean_dec(x_45); -lean_inc(x_42); +lean_object* x_33; lean_object* x_34; size_t x_35; size_t 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; lean_object* x_44; size_t x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_33 = lean_ctor_get(x_28, 0); +lean_inc(x_33); +lean_dec(x_28); +x_34 = lean_array_get_size(x_33); +x_35 = lean_usize_of_nat(x_34); +lean_dec(x_34); +x_36 = 0; +lean_inc(x_33); +x_37 = l_Array_mapMUnsafe_map___at_Lean_Elab_elabDeriving___spec__4(x_35, x_36, x_33); +x_38 = l_Array_mapMUnsafe_map___at_Lean_Elab_elabDeriving___spec__5(x_35, x_36, x_33); +x_39 = lean_array_get_size(x_37); +x_40 = l_Array_toSubarray___rarg(x_37, x_13, x_39); +x_41 = lean_ctor_get(x_1, 1); +lean_inc(x_41); +x_42 = l_Lean_Elab_elabDeriving___closed__9; +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_40); +lean_ctor_set(x_43, 1, x_42); +x_44 = lean_array_get_size(x_38); +x_45 = lean_usize_of_nat(x_44); +lean_dec(x_44); +lean_inc(x_41); lean_inc(x_1); -x_47 = l_Array_forInUnsafe_loop___at_Lean_Elab_getOptDerivingClasses___spec__1___rarg(x_1, x_2, x_3, x_4, x_5, x_42, x_39, x_46, x_37, x_44); -x_48 = lean_alloc_closure((void*)(l_Lean_Elab_getOptDerivingClasses___rarg___lambda__1), 2, 1); -lean_closure_set(x_48, 0, x_1); -x_49 = lean_apply_4(x_42, lean_box(0), lean_box(0), x_47, x_48); -return x_49; +x_46 = l_Array_forInUnsafe_loop___at_Lean_Elab_getOptDerivingClasses___spec__1___rarg(x_1, x_2, x_3, x_4, x_5, x_41, x_38, x_45, x_36, x_43); +x_47 = lean_alloc_closure((void*)(l_Lean_Elab_getOptDerivingClasses___rarg___lambda__1), 2, 1); +lean_closure_set(x_47, 0, x_1); +x_48 = lean_apply_4(x_41, lean_box(0), lean_box(0), x_46, x_47); +return x_48; } } } @@ -4305,7 +4485,7 @@ return x_24; } } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1490____closed__1() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1519____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -4315,7 +4495,7 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1490____closed__2() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1519____closed__2() { _start: { lean_object* x_1; @@ -4323,21 +4503,21 @@ x_1 = lean_mk_string_from_bytes("Deriving", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1490____closed__3() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1519____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1490____closed__1; -x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1490____closed__2; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1519____closed__1; +x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1519____closed__2; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1490_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1519_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1490____closed__3; +x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1519____closed__3; x_3 = l_Lean_registerTraceClass(x_2, x_1); return x_3; } @@ -4359,7 +4539,7 @@ lean_dec_ref(res); res = initialize_Lean_Elab_MutualDef(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -if (builtin) {res = l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_22_(lean_io_mk_world()); +if (builtin) {res = l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_25_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Elab_derivingHandlersRef = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Elab_derivingHandlersRef); @@ -4390,28 +4570,28 @@ l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_elabDeriving___spec__1___rarg_ lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_elabDeriving___spec__1___rarg___closed__1); l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_elabDeriving___spec__1___rarg___closed__2 = _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_elabDeriving___spec__1___rarg___closed__2(); lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_elabDeriving___spec__1___rarg___closed__2); -l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__8___closed__1 = _init_l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__8___closed__1(); -lean_mark_persistent(l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__8___closed__1); -l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__8___closed__2 = _init_l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__8___closed__2(); -lean_mark_persistent(l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__8___closed__2); -l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__8___closed__3 = _init_l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__8___closed__3(); -lean_mark_persistent(l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__8___closed__3); -l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__4___closed__1 = _init_l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__4___closed__1(); -lean_mark_persistent(l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__4___closed__1); -l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__4___closed__2 = _init_l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__4___closed__2(); -lean_mark_persistent(l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__4___closed__2); -l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__4___closed__3 = _init_l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__4___closed__3(); -lean_mark_persistent(l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__4___closed__3); -l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_elabDeriving___spec__3___closed__1 = _init_l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_elabDeriving___spec__3___closed__1(); -lean_mark_persistent(l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_elabDeriving___spec__3___closed__1); -l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_elabDeriving___spec__3___closed__2 = _init_l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_elabDeriving___spec__3___closed__2(); -lean_mark_persistent(l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_elabDeriving___spec__3___closed__2); -l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__12___closed__1 = _init_l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__12___closed__1(); -lean_mark_persistent(l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__12___closed__1); -l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__12___closed__2 = _init_l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__12___closed__2(); -lean_mark_persistent(l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__12___closed__2); -l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__12___closed__3 = _init_l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__12___closed__3(); -lean_mark_persistent(l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__12___closed__3); +l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__12___closed__1 = _init_l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__12___closed__1(); +lean_mark_persistent(l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__12___closed__1); +l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__12___closed__2 = _init_l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__12___closed__2(); +lean_mark_persistent(l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__12___closed__2); +l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__12___closed__3 = _init_l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__12___closed__3(); +lean_mark_persistent(l_Lean_throwUnknownConstant___at_Lean_Elab_elabDeriving___spec__12___closed__3); +l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__8___closed__1 = _init_l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__8___closed__1(); +lean_mark_persistent(l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__8___closed__1); +l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__8___closed__2 = _init_l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__8___closed__2(); +lean_mark_persistent(l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__8___closed__2); +l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__8___closed__3 = _init_l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__8___closed__3(); +lean_mark_persistent(l_Lean_resolveGlobalConst___at_Lean_Elab_elabDeriving___spec__8___closed__3); +l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_elabDeriving___spec__7___closed__1 = _init_l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_elabDeriving___spec__7___closed__1(); +lean_mark_persistent(l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_elabDeriving___spec__7___closed__1); +l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_elabDeriving___spec__7___closed__2 = _init_l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_elabDeriving___spec__7___closed__2(); +lean_mark_persistent(l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_elabDeriving___spec__7___closed__2); +l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__16___closed__1 = _init_l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__16___closed__1(); +lean_mark_persistent(l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__16___closed__1); +l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__16___closed__2 = _init_l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__16___closed__2(); +lean_mark_persistent(l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__16___closed__2); +l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__16___closed__3 = _init_l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__16___closed__3(); +lean_mark_persistent(l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_elabDeriving___spec__16___closed__3); l_Lean_Elab_elabDeriving___lambda__2___closed__1 = _init_l_Lean_Elab_elabDeriving___lambda__2___closed__1(); lean_mark_persistent(l_Lean_Elab_elabDeriving___lambda__2___closed__1); l_Lean_Elab_elabDeriving___lambda__2___closed__2 = _init_l_Lean_Elab_elabDeriving___lambda__2___closed__2(); @@ -4476,13 +4656,13 @@ l_Lean_Elab_getOptDerivingClasses___rarg___closed__1 = _init_l_Lean_Elab_getOptD lean_mark_persistent(l_Lean_Elab_getOptDerivingClasses___rarg___closed__1); l_Lean_Elab_getOptDerivingClasses___rarg___closed__2 = _init_l_Lean_Elab_getOptDerivingClasses___rarg___closed__2(); lean_mark_persistent(l_Lean_Elab_getOptDerivingClasses___rarg___closed__2); -l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1490____closed__1 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1490____closed__1(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1490____closed__1); -l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1490____closed__2 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1490____closed__2(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1490____closed__2); -l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1490____closed__3 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1490____closed__3(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1490____closed__3); -res = l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1490_(lean_io_mk_world()); +l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1519____closed__1 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1519____closed__1(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1519____closed__1); +l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1519____closed__2 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1519____closed__2(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1519____closed__2); +l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1519____closed__3 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1519____closed__3(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1519____closed__3); +res = l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_1519_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/Deriving/DecEq.c b/stage0/stdlib/Lean/Elab/Deriving/DecEq.c index 6a44ab87d23..2df925d2406 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/DecEq.c +++ b/stage0/stdlib/Lean/Elab/Deriving/DecEq.c @@ -22,6 +22,7 @@ uint8_t l_Lean_Expr_hasAnyFVar_visit___at_Lean_Expr_containsFVar___spec__1(lean_ static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqCmds___closed__5; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__15; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3___closed__10; +size_t lean_usize_add(size_t, size_t); static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__25; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__22; static lean_object* l_Lean_Elab_Deriving_DecEq_mkEnumOfNat_mkDecTree___closed__4; @@ -38,6 +39,7 @@ static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAl lean_object* lean_nat_div(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqCmds___closed__2; lean_object* l_Lean_Meta_mkForallFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__14(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__14; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__24; static lean_object* l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__1; @@ -46,6 +48,7 @@ static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2 static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__58; lean_object* lean_name_mk_string(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__50; +lean_object* lean_array_uget(lean_object*, size_t); static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__17; uint8_t l_Lean_Expr_isProp(lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkEnumOfNatThm___closed__9; @@ -59,11 +62,13 @@ static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAl static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__48; static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqCmds___closed__3; static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__46; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_DecEq_mkMatch___spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__12; static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__29; +lean_object* lean_array_uset(lean_object*, size_t, lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__9; static lean_object* l_Lean_Elab_Deriving_DecEq_mkEnumOfNatThm___closed__2; -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__7(lean_object*, lean_object*, lean_object*, 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_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkEnumOfNatThm___closed__8; static lean_object* l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__4; static lean_object* l_Lean_Elab_Deriving_DecEq_mkEnumOfNat_mkDecTree___closed__6; @@ -98,7 +103,7 @@ static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1 lean_object* lean_array_get_size(lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch___closed__4; static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__7; -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, 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_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__22; lean_object* l_Lean_Core_betaReduce___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofList(lean_object*); @@ -108,6 +113,7 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_DecEq_mk static lean_object* l_Lean_Elab_Deriving_DecEq_mkEnumOfNat_mkDecTree___closed__5; static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqHeader___closed__2; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3___closed__11; +static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_5097____closed__1; static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqHeader___closed__1; lean_object* l_Lean_addTrace___at_Lean_Elab_Command_elabCommand___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkEnumOfNat___lambda__1___closed__3; @@ -121,6 +127,7 @@ static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__5; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__30; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__18; LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkEnumOfNatThm___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__53; static lean_object* l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__21; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqInstanceHandler___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -258,7 +265,6 @@ lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__36; LEAN_EXPORT lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__21; -static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4892____closed__1; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__51; lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__16(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__46; @@ -302,7 +308,8 @@ static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__61; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__32; static lean_object* l_Lean_Elab_Deriving_DecEq_mkEnumOfNatThm___lambda__1___closed__1; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__10; -LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4892_(lean_object*); +static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___closed__6; +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_5097_(lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3___closed__14; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__39; lean_object* l_Lean_Elab_Command_getMainModule___rarg(lean_object*, lean_object*); @@ -395,6 +402,7 @@ static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___closed__7 lean_object* l_Lean_mkAppB(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__14; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch___closed__5; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_DecEq_mkMatch___spec__1(size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__6; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__28; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__4; @@ -403,7 +411,6 @@ static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqCmds___closed__6; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3___closed__5; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3___closed__2; lean_object* l_Lean_mkConst(lean_object*, lean_object*); -lean_object* l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__10(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkEnumOfNat___closed__1; static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__19; static lean_object* l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__20; @@ -3683,267 +3690,267 @@ lean_dec(x_25); x_28 = !lean_is_exclusive(x_26); if (x_28 == 0) { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; x_29 = lean_ctor_get(x_26, 0); x_30 = lean_ctor_get(x_26, 1); x_31 = lean_ctor_get(x_3, 4); lean_inc(x_31); lean_dec(x_3); -x_32 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__19; -lean_ctor_set(x_26, 1, x_32); +lean_inc(x_2); +lean_ctor_set(x_26, 1, x_2); lean_ctor_set(x_26, 0, x_30); -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_29); -lean_ctor_set(x_33, 1, x_26); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_29); +lean_ctor_set(x_32, 1, x_26); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_31); -x_34 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__5(x_1, x_7, x_19, x_21, x_31, x_23, x_31, x_24, x_33, x_9, x_10, x_11, x_12, x_13, x_14, x_27); +x_33 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__5(x_1, x_7, x_19, x_21, x_31, x_23, x_31, x_24, x_32, x_9, x_10, x_11, x_12, x_13, x_14, x_27); lean_dec(x_31); lean_dec(x_21); lean_dec(x_19); lean_dec(x_1); -if (lean_obj_tag(x_34) == 0) +if (lean_obj_tag(x_33) == 0) { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; -x_35 = lean_ctor_get(x_34, 0); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_34, 1); lean_inc(x_35); -x_36 = lean_ctor_get(x_35, 1); +x_36 = lean_ctor_get(x_33, 1); lean_inc(x_36); -x_37 = lean_ctor_get(x_34, 1); +lean_dec(x_33); +x_37 = lean_ctor_get(x_34, 0); lean_inc(x_37); lean_dec(x_34); x_38 = lean_ctor_get(x_35, 0); lean_inc(x_38); -lean_dec(x_35); -x_39 = lean_ctor_get(x_36, 0); +x_39 = lean_ctor_get(x_35, 1); lean_inc(x_39); -x_40 = lean_ctor_get(x_36, 1); -lean_inc(x_40); -lean_dec(x_36); +lean_dec(x_35); lean_inc(x_13); -x_41 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_13, x_14, x_37); -x_42 = lean_ctor_get(x_41, 0); +x_40 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_13, x_14, x_36); +x_41 = lean_ctor_get(x_40, 0); +lean_inc(x_41); +x_42 = lean_ctor_get(x_40, 1); lean_inc(x_42); -x_43 = lean_ctor_get(x_41, 1); -lean_inc(x_43); -lean_dec(x_41); -x_44 = lean_st_ref_get(x_14, x_43); -x_45 = lean_ctor_get(x_44, 1); -lean_inc(x_45); -lean_dec(x_44); -x_46 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__5; -x_47 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_47, 0, x_42); -lean_ctor_set(x_47, 1, x_46); -x_48 = lean_mk_syntax_ident(x_4); -x_49 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__24; +lean_dec(x_40); +x_43 = lean_st_ref_get(x_14, x_42); +x_44 = lean_ctor_get(x_43, 1); +lean_inc(x_44); +lean_dec(x_43); +x_45 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__5; +x_46 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_46, 0, x_41); +lean_ctor_set(x_46, 1, x_45); +x_47 = lean_mk_syntax_ident(x_4); +x_48 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__24; +x_49 = lean_array_push(x_48, x_46); +lean_inc(x_47); x_50 = lean_array_push(x_49, x_47); -lean_inc(x_48); -x_51 = lean_array_push(x_50, x_48); -x_52 = lean_box(2); -x_53 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__4; -x_54 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_54, 0, x_52); -lean_ctor_set(x_54, 1, x_53); -lean_ctor_set(x_54, 2, x_51); +x_51 = lean_box(2); +x_52 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__4; +x_53 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_52); +lean_ctor_set(x_53, 2, x_50); lean_inc(x_2); -x_55 = l_Array_append___rarg(x_2, x_38); -x_56 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__18; -x_57 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_57, 0, x_52); -lean_ctor_set(x_57, 1, x_56); -lean_ctor_set(x_57, 2, x_55); -x_58 = lean_array_push(x_49, x_54); -x_59 = lean_array_push(x_58, x_57); -x_60 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__23; -x_61 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_61, 0, x_52); -lean_ctor_set(x_61, 1, x_60); -lean_ctor_set(x_61, 2, x_59); -x_62 = lean_array_push(x_5, x_61); +x_54 = l_Array_append___rarg(x_2, x_37); +x_55 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__18; +x_56 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_56, 0, x_51); +lean_ctor_set(x_56, 1, x_55); +lean_ctor_set(x_56, 2, x_54); +x_57 = lean_array_push(x_48, x_53); +x_58 = lean_array_push(x_57, x_56); +x_59 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__23; +x_60 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_60, 0, x_51); +lean_ctor_set(x_60, 1, x_59); +lean_ctor_set(x_60, 2, x_58); +x_61 = lean_array_push(x_5, x_60); lean_inc(x_13); -x_63 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_13, x_14, x_45); -x_64 = lean_ctor_get(x_63, 0); +x_62 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_13, x_14, x_44); +x_63 = lean_ctor_get(x_62, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_62, 1); lean_inc(x_64); -x_65 = lean_ctor_get(x_63, 1); -lean_inc(x_65); -lean_dec(x_63); -x_66 = lean_st_ref_get(x_14, x_65); -x_67 = lean_ctor_get(x_66, 1); -lean_inc(x_67); -lean_dec(x_66); -x_68 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_68, 0, x_64); -lean_ctor_set(x_68, 1, x_46); -x_69 = lean_array_push(x_49, x_68); -x_70 = lean_array_push(x_69, x_48); -x_71 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_71, 0, x_52); -lean_ctor_set(x_71, 1, x_53); -lean_ctor_set(x_71, 2, x_70); +lean_dec(x_62); +x_65 = lean_st_ref_get(x_14, x_64); +x_66 = lean_ctor_get(x_65, 1); +lean_inc(x_66); +lean_dec(x_65); +x_67 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_67, 0, x_63); +lean_ctor_set(x_67, 1, x_45); +x_68 = lean_array_push(x_48, x_67); +x_69 = lean_array_push(x_68, x_47); +x_70 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_70, 0, x_51); +lean_ctor_set(x_70, 1, x_52); +lean_ctor_set(x_70, 2, x_69); lean_inc(x_2); -x_72 = l_Array_append___rarg(x_2, x_39); -x_73 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_73, 0, x_52); -lean_ctor_set(x_73, 1, x_56); -lean_ctor_set(x_73, 2, x_72); -x_74 = lean_array_push(x_49, x_71); -x_75 = lean_array_push(x_74, x_73); -x_76 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_76, 0, x_52); -lean_ctor_set(x_76, 1, x_60); -lean_ctor_set(x_76, 2, x_75); -x_77 = lean_array_push(x_62, x_76); -x_78 = lean_array_to_list(lean_box(0), x_40); +x_71 = l_Array_append___rarg(x_2, x_38); +x_72 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_72, 0, x_51); +lean_ctor_set(x_72, 1, x_55); +lean_ctor_set(x_72, 2, x_71); +x_73 = lean_array_push(x_48, x_70); +x_74 = lean_array_push(x_73, x_72); +x_75 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_75, 0, x_51); +lean_ctor_set(x_75, 1, x_59); +lean_ctor_set(x_75, 2, x_74); +x_76 = lean_array_push(x_61, x_75); +x_77 = lean_array_to_list(lean_box(0), x_39); lean_inc(x_14); lean_inc(x_13); -x_79 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs(x_6, x_78, x_9, x_10, x_11, x_12, x_13, x_14, x_67); -if (lean_obj_tag(x_79) == 0) +x_78 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs(x_6, x_77, x_9, x_10, x_11, x_12, x_13, x_14, x_66); +if (lean_obj_tag(x_78) == 0) { -lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; uint8_t x_86; -x_80 = lean_ctor_get(x_79, 0); +lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; uint8_t x_85; +x_79 = lean_ctor_get(x_78, 0); +lean_inc(x_79); +x_80 = lean_ctor_get(x_78, 1); lean_inc(x_80); -x_81 = lean_ctor_get(x_79, 1); -lean_inc(x_81); -lean_dec(x_79); -x_82 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_13, x_14, x_81); -x_83 = lean_ctor_get(x_82, 0); +lean_dec(x_78); +x_81 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_13, x_14, x_80); +x_82 = lean_ctor_get(x_81, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_81, 1); lean_inc(x_83); -x_84 = lean_ctor_get(x_82, 1); -lean_inc(x_84); -lean_dec(x_82); -x_85 = lean_st_ref_get(x_14, x_84); +lean_dec(x_81); +x_84 = lean_st_ref_get(x_14, x_83); lean_dec(x_14); -x_86 = !lean_is_exclusive(x_85); -if (x_86 == 0) +x_85 = !lean_is_exclusive(x_84); +if (x_85 == 0) { -lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; size_t x_91; size_t x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; -x_87 = lean_ctor_get(x_85, 0); -lean_dec(x_87); -x_88 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__8; -lean_inc(x_83); -x_89 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_89, 0, x_83); -lean_ctor_set(x_89, 1, x_88); -x_90 = lean_array_get_size(x_77); -x_91 = lean_usize_of_nat(x_90); -lean_dec(x_90); -x_92 = 0; -x_93 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_91, x_92, x_77); -x_94 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__10; -x_95 = l_Lean_mkSepArray(x_93, x_94); -lean_dec(x_93); -x_96 = l_Array_append___rarg(x_2, x_95); -x_97 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_97, 0, x_52); -lean_ctor_set(x_97, 1, x_56); -lean_ctor_set(x_97, 2, x_96); -x_98 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__26; -x_99 = lean_array_push(x_98, x_97); -x_100 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_100, 0, x_52); -lean_ctor_set(x_100, 1, x_56); -lean_ctor_set(x_100, 2, x_99); -x_101 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__11; -x_102 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_102, 0, x_83); -lean_ctor_set(x_102, 1, x_101); -x_103 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__28; -x_104 = lean_array_push(x_103, x_89); -x_105 = lean_array_push(x_104, x_100); -x_106 = lean_array_push(x_105, x_102); -x_107 = lean_array_push(x_106, x_80); -x_108 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__7; -x_109 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_109, 0, x_52); -lean_ctor_set(x_109, 1, x_108); -lean_ctor_set(x_109, 2, x_107); -lean_ctor_set(x_85, 0, x_109); -return x_85; +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; size_t x_90; size_t x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; +x_86 = lean_ctor_get(x_84, 0); +lean_dec(x_86); +x_87 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__8; +lean_inc(x_82); +x_88 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_88, 0, x_82); +lean_ctor_set(x_88, 1, x_87); +x_89 = lean_array_get_size(x_76); +x_90 = lean_usize_of_nat(x_89); +lean_dec(x_89); +x_91 = 0; +x_92 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_90, x_91, x_76); +x_93 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__10; +x_94 = l_Lean_mkSepArray(x_92, x_93); +lean_dec(x_92); +x_95 = l_Array_append___rarg(x_2, x_94); +x_96 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_96, 0, x_51); +lean_ctor_set(x_96, 1, x_55); +lean_ctor_set(x_96, 2, x_95); +x_97 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__26; +x_98 = lean_array_push(x_97, x_96); +x_99 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_99, 0, x_51); +lean_ctor_set(x_99, 1, x_55); +lean_ctor_set(x_99, 2, x_98); +x_100 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__11; +x_101 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_101, 0, x_82); +lean_ctor_set(x_101, 1, x_100); +x_102 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__28; +x_103 = lean_array_push(x_102, x_88); +x_104 = lean_array_push(x_103, x_99); +x_105 = lean_array_push(x_104, x_101); +x_106 = lean_array_push(x_105, x_79); +x_107 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__7; +x_108 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_108, 0, x_51); +lean_ctor_set(x_108, 1, x_107); +lean_ctor_set(x_108, 2, x_106); +lean_ctor_set(x_84, 0, x_108); +return x_84; } else { -lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; size_t x_114; size_t x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; -x_110 = lean_ctor_get(x_85, 1); -lean_inc(x_110); -lean_dec(x_85); -x_111 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__8; -lean_inc(x_83); -x_112 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_112, 0, x_83); -lean_ctor_set(x_112, 1, x_111); -x_113 = lean_array_get_size(x_77); -x_114 = lean_usize_of_nat(x_113); -lean_dec(x_113); -x_115 = 0; -x_116 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_114, x_115, x_77); -x_117 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__10; -x_118 = l_Lean_mkSepArray(x_116, x_117); -lean_dec(x_116); -x_119 = l_Array_append___rarg(x_2, x_118); -x_120 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_120, 0, x_52); -lean_ctor_set(x_120, 1, x_56); -lean_ctor_set(x_120, 2, x_119); -x_121 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__26; -x_122 = lean_array_push(x_121, x_120); -x_123 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_123, 0, x_52); -lean_ctor_set(x_123, 1, x_56); -lean_ctor_set(x_123, 2, x_122); -x_124 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__11; -x_125 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_125, 0, x_83); -lean_ctor_set(x_125, 1, x_124); -x_126 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__28; -x_127 = lean_array_push(x_126, x_112); -x_128 = lean_array_push(x_127, x_123); -x_129 = lean_array_push(x_128, x_125); -x_130 = lean_array_push(x_129, x_80); -x_131 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__7; -x_132 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_132, 0, x_52); -lean_ctor_set(x_132, 1, x_131); -lean_ctor_set(x_132, 2, x_130); -x_133 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_133, 0, x_132); -lean_ctor_set(x_133, 1, x_110); -return x_133; +lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; size_t x_113; size_t x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; +x_109 = lean_ctor_get(x_84, 1); +lean_inc(x_109); +lean_dec(x_84); +x_110 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__8; +lean_inc(x_82); +x_111 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_111, 0, x_82); +lean_ctor_set(x_111, 1, x_110); +x_112 = lean_array_get_size(x_76); +x_113 = lean_usize_of_nat(x_112); +lean_dec(x_112); +x_114 = 0; +x_115 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_113, x_114, x_76); +x_116 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__10; +x_117 = l_Lean_mkSepArray(x_115, x_116); +lean_dec(x_115); +x_118 = l_Array_append___rarg(x_2, x_117); +x_119 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_119, 0, x_51); +lean_ctor_set(x_119, 1, x_55); +lean_ctor_set(x_119, 2, x_118); +x_120 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__26; +x_121 = lean_array_push(x_120, x_119); +x_122 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_122, 0, x_51); +lean_ctor_set(x_122, 1, x_55); +lean_ctor_set(x_122, 2, x_121); +x_123 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__11; +x_124 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_124, 0, x_82); +lean_ctor_set(x_124, 1, x_123); +x_125 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__28; +x_126 = lean_array_push(x_125, x_111); +x_127 = lean_array_push(x_126, x_122); +x_128 = lean_array_push(x_127, x_124); +x_129 = lean_array_push(x_128, x_79); +x_130 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__7; +x_131 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_131, 0, x_51); +lean_ctor_set(x_131, 1, x_130); +lean_ctor_set(x_131, 2, x_129); +x_132 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_132, 0, x_131); +lean_ctor_set(x_132, 1, x_109); +return x_132; } } else { -uint8_t x_134; -lean_dec(x_77); +uint8_t x_133; +lean_dec(x_76); lean_dec(x_14); lean_dec(x_13); lean_dec(x_2); -x_134 = !lean_is_exclusive(x_79); -if (x_134 == 0) +x_133 = !lean_is_exclusive(x_78); +if (x_133 == 0) { -return x_79; +return x_78; } else { -lean_object* x_135; lean_object* x_136; lean_object* x_137; -x_135 = lean_ctor_get(x_79, 0); -x_136 = lean_ctor_get(x_79, 1); -lean_inc(x_136); +lean_object* x_134; lean_object* x_135; lean_object* x_136; +x_134 = lean_ctor_get(x_78, 0); +x_135 = lean_ctor_get(x_78, 1); lean_inc(x_135); -lean_dec(x_79); -x_137 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_137, 0, x_135); -lean_ctor_set(x_137, 1, x_136); -return x_137; +lean_inc(x_134); +lean_dec(x_78); +x_136 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_136, 0, x_134); +lean_ctor_set(x_136, 1, x_135); +return x_136; } } } else { -uint8_t x_138; +uint8_t x_137; lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); @@ -3954,255 +3961,255 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_138 = !lean_is_exclusive(x_34); -if (x_138 == 0) +x_137 = !lean_is_exclusive(x_33); +if (x_137 == 0) { -return x_34; +return x_33; } else { -lean_object* x_139; lean_object* x_140; lean_object* x_141; -x_139 = lean_ctor_get(x_34, 0); -x_140 = lean_ctor_get(x_34, 1); -lean_inc(x_140); +lean_object* x_138; lean_object* x_139; lean_object* x_140; +x_138 = lean_ctor_get(x_33, 0); +x_139 = lean_ctor_get(x_33, 1); lean_inc(x_139); -lean_dec(x_34); -x_141 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_141, 0, x_139); -lean_ctor_set(x_141, 1, x_140); -return x_141; +lean_inc(x_138); +lean_dec(x_33); +x_140 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_140, 0, x_138); +lean_ctor_set(x_140, 1, x_139); +return x_140; } } } else { -lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; -x_142 = lean_ctor_get(x_26, 0); -x_143 = lean_ctor_get(x_26, 1); -lean_inc(x_143); +lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; +x_141 = lean_ctor_get(x_26, 0); +x_142 = lean_ctor_get(x_26, 1); lean_inc(x_142); +lean_inc(x_141); lean_dec(x_26); -x_144 = lean_ctor_get(x_3, 4); -lean_inc(x_144); +x_143 = lean_ctor_get(x_3, 4); +lean_inc(x_143); lean_dec(x_3); -x_145 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__19; -x_146 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_146, 0, x_143); -lean_ctor_set(x_146, 1, x_145); -x_147 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_147, 0, x_142); -lean_ctor_set(x_147, 1, x_146); +lean_inc(x_2); +x_144 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_144, 0, x_142); +lean_ctor_set(x_144, 1, x_2); +x_145 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_145, 0, x_141); +lean_ctor_set(x_145, 1, x_144); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); -lean_inc(x_144); -x_148 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__5(x_1, x_7, x_19, x_21, x_144, x_23, x_144, x_24, x_147, x_9, x_10, x_11, x_12, x_13, x_14, x_27); -lean_dec(x_144); +lean_inc(x_143); +x_146 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__5(x_1, x_7, x_19, x_21, x_143, x_23, x_143, x_24, x_145, x_9, x_10, x_11, x_12, x_13, x_14, x_27); +lean_dec(x_143); lean_dec(x_21); lean_dec(x_19); lean_dec(x_1); -if (lean_obj_tag(x_148) == 0) +if (lean_obj_tag(x_146) == 0) { -lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; -x_149 = lean_ctor_get(x_148, 0); +lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; +x_147 = lean_ctor_get(x_146, 0); +lean_inc(x_147); +x_148 = lean_ctor_get(x_147, 1); +lean_inc(x_148); +x_149 = lean_ctor_get(x_146, 1); lean_inc(x_149); -x_150 = lean_ctor_get(x_149, 1); +lean_dec(x_146); +x_150 = lean_ctor_get(x_147, 0); lean_inc(x_150); -x_151 = lean_ctor_get(x_148, 1); +lean_dec(x_147); +x_151 = lean_ctor_get(x_148, 0); lean_inc(x_151); -lean_dec(x_148); -x_152 = lean_ctor_get(x_149, 0); +x_152 = lean_ctor_get(x_148, 1); lean_inc(x_152); -lean_dec(x_149); -x_153 = lean_ctor_get(x_150, 0); -lean_inc(x_153); -x_154 = lean_ctor_get(x_150, 1); -lean_inc(x_154); -lean_dec(x_150); +lean_dec(x_148); lean_inc(x_13); -x_155 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_13, x_14, x_151); -x_156 = lean_ctor_get(x_155, 0); -lean_inc(x_156); -x_157 = lean_ctor_get(x_155, 1); +x_153 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_13, x_14, x_149); +x_154 = lean_ctor_get(x_153, 0); +lean_inc(x_154); +x_155 = lean_ctor_get(x_153, 1); +lean_inc(x_155); +lean_dec(x_153); +x_156 = lean_st_ref_get(x_14, x_155); +x_157 = lean_ctor_get(x_156, 1); lean_inc(x_157); -lean_dec(x_155); -x_158 = lean_st_ref_get(x_14, x_157); -x_159 = lean_ctor_get(x_158, 1); -lean_inc(x_159); -lean_dec(x_158); -x_160 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__5; -x_161 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_161, 0, x_156); -lean_ctor_set(x_161, 1, x_160); -x_162 = lean_mk_syntax_ident(x_4); -x_163 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__24; -x_164 = lean_array_push(x_163, x_161); -lean_inc(x_162); -x_165 = lean_array_push(x_164, x_162); -x_166 = lean_box(2); -x_167 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__4; -x_168 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_168, 0, x_166); -lean_ctor_set(x_168, 1, x_167); -lean_ctor_set(x_168, 2, x_165); +lean_dec(x_156); +x_158 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__5; +x_159 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_159, 0, x_154); +lean_ctor_set(x_159, 1, x_158); +x_160 = lean_mk_syntax_ident(x_4); +x_161 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__24; +x_162 = lean_array_push(x_161, x_159); +lean_inc(x_160); +x_163 = lean_array_push(x_162, x_160); +x_164 = lean_box(2); +x_165 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__4; +x_166 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_166, 0, x_164); +lean_ctor_set(x_166, 1, x_165); +lean_ctor_set(x_166, 2, x_163); lean_inc(x_2); -x_169 = l_Array_append___rarg(x_2, x_152); -x_170 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__18; -x_171 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_171, 0, x_166); -lean_ctor_set(x_171, 1, x_170); -lean_ctor_set(x_171, 2, x_169); -x_172 = lean_array_push(x_163, x_168); -x_173 = lean_array_push(x_172, x_171); -x_174 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__23; -x_175 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_175, 0, x_166); -lean_ctor_set(x_175, 1, x_174); -lean_ctor_set(x_175, 2, x_173); -x_176 = lean_array_push(x_5, x_175); +x_167 = l_Array_append___rarg(x_2, x_150); +x_168 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__18; +x_169 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_169, 0, x_164); +lean_ctor_set(x_169, 1, x_168); +lean_ctor_set(x_169, 2, x_167); +x_170 = lean_array_push(x_161, x_166); +x_171 = lean_array_push(x_170, x_169); +x_172 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__23; +x_173 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_173, 0, x_164); +lean_ctor_set(x_173, 1, x_172); +lean_ctor_set(x_173, 2, x_171); +x_174 = lean_array_push(x_5, x_173); lean_inc(x_13); -x_177 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_13, x_14, x_159); -x_178 = lean_ctor_get(x_177, 0); -lean_inc(x_178); -x_179 = lean_ctor_get(x_177, 1); +x_175 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_13, x_14, x_157); +x_176 = lean_ctor_get(x_175, 0); +lean_inc(x_176); +x_177 = lean_ctor_get(x_175, 1); +lean_inc(x_177); +lean_dec(x_175); +x_178 = lean_st_ref_get(x_14, x_177); +x_179 = lean_ctor_get(x_178, 1); lean_inc(x_179); -lean_dec(x_177); -x_180 = lean_st_ref_get(x_14, x_179); -x_181 = lean_ctor_get(x_180, 1); -lean_inc(x_181); -lean_dec(x_180); -x_182 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_182, 0, x_178); -lean_ctor_set(x_182, 1, x_160); -x_183 = lean_array_push(x_163, x_182); -x_184 = lean_array_push(x_183, x_162); +lean_dec(x_178); +x_180 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_180, 0, x_176); +lean_ctor_set(x_180, 1, x_158); +x_181 = lean_array_push(x_161, x_180); +x_182 = lean_array_push(x_181, x_160); +x_183 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_183, 0, x_164); +lean_ctor_set(x_183, 1, x_165); +lean_ctor_set(x_183, 2, x_182); +lean_inc(x_2); +x_184 = l_Array_append___rarg(x_2, x_151); x_185 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_185, 0, x_166); -lean_ctor_set(x_185, 1, x_167); +lean_ctor_set(x_185, 0, x_164); +lean_ctor_set(x_185, 1, x_168); lean_ctor_set(x_185, 2, x_184); -lean_inc(x_2); -x_186 = l_Array_append___rarg(x_2, x_153); -x_187 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_187, 0, x_166); -lean_ctor_set(x_187, 1, x_170); -lean_ctor_set(x_187, 2, x_186); -x_188 = lean_array_push(x_163, x_185); -x_189 = lean_array_push(x_188, x_187); -x_190 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_190, 0, x_166); -lean_ctor_set(x_190, 1, x_174); -lean_ctor_set(x_190, 2, x_189); -x_191 = lean_array_push(x_176, x_190); -x_192 = lean_array_to_list(lean_box(0), x_154); +x_186 = lean_array_push(x_161, x_183); +x_187 = lean_array_push(x_186, x_185); +x_188 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_188, 0, x_164); +lean_ctor_set(x_188, 1, x_172); +lean_ctor_set(x_188, 2, x_187); +x_189 = lean_array_push(x_174, x_188); +x_190 = lean_array_to_list(lean_box(0), x_152); lean_inc(x_14); lean_inc(x_13); -x_193 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs(x_6, x_192, x_9, x_10, x_11, x_12, x_13, x_14, x_181); -if (lean_obj_tag(x_193) == 0) +x_191 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs(x_6, x_190, x_9, x_10, x_11, x_12, x_13, x_14, x_179); +if (lean_obj_tag(x_191) == 0) { -lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; size_t x_205; size_t x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; -x_194 = lean_ctor_get(x_193, 0); -lean_inc(x_194); -x_195 = lean_ctor_get(x_193, 1); +lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; size_t x_203; size_t x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; +x_192 = lean_ctor_get(x_191, 0); +lean_inc(x_192); +x_193 = lean_ctor_get(x_191, 1); +lean_inc(x_193); +lean_dec(x_191); +x_194 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_13, x_14, x_193); +x_195 = lean_ctor_get(x_194, 0); lean_inc(x_195); -lean_dec(x_193); -x_196 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_13, x_14, x_195); -x_197 = lean_ctor_get(x_196, 0); -lean_inc(x_197); -x_198 = lean_ctor_get(x_196, 1); -lean_inc(x_198); -lean_dec(x_196); -x_199 = lean_st_ref_get(x_14, x_198); +x_196 = lean_ctor_get(x_194, 1); +lean_inc(x_196); +lean_dec(x_194); +x_197 = lean_st_ref_get(x_14, x_196); lean_dec(x_14); -x_200 = lean_ctor_get(x_199, 1); -lean_inc(x_200); -if (lean_is_exclusive(x_199)) { - lean_ctor_release(x_199, 0); - lean_ctor_release(x_199, 1); - x_201 = x_199; +x_198 = lean_ctor_get(x_197, 1); +lean_inc(x_198); +if (lean_is_exclusive(x_197)) { + lean_ctor_release(x_197, 0); + lean_ctor_release(x_197, 1); + x_199 = x_197; } else { - lean_dec_ref(x_199); - x_201 = lean_box(0); + lean_dec_ref(x_197); + x_199 = lean_box(0); } -x_202 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__8; -lean_inc(x_197); -x_203 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_203, 0, x_197); -lean_ctor_set(x_203, 1, x_202); -x_204 = lean_array_get_size(x_191); -x_205 = lean_usize_of_nat(x_204); -lean_dec(x_204); -x_206 = 0; -x_207 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_205, x_206, x_191); -x_208 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__10; -x_209 = l_Lean_mkSepArray(x_207, x_208); -lean_dec(x_207); -x_210 = l_Array_append___rarg(x_2, x_209); -x_211 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_211, 0, x_166); -lean_ctor_set(x_211, 1, x_170); -lean_ctor_set(x_211, 2, x_210); -x_212 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__26; -x_213 = lean_array_push(x_212, x_211); -x_214 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_214, 0, x_166); -lean_ctor_set(x_214, 1, x_170); -lean_ctor_set(x_214, 2, x_213); -x_215 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__11; -x_216 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_216, 0, x_197); -lean_ctor_set(x_216, 1, x_215); -x_217 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__28; -x_218 = lean_array_push(x_217, x_203); -x_219 = lean_array_push(x_218, x_214); -x_220 = lean_array_push(x_219, x_216); -x_221 = lean_array_push(x_220, x_194); -x_222 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__7; -x_223 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_223, 0, x_166); -lean_ctor_set(x_223, 1, x_222); -lean_ctor_set(x_223, 2, x_221); -if (lean_is_scalar(x_201)) { - x_224 = lean_alloc_ctor(0, 2, 0); +x_200 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__8; +lean_inc(x_195); +x_201 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_201, 0, x_195); +lean_ctor_set(x_201, 1, x_200); +x_202 = lean_array_get_size(x_189); +x_203 = lean_usize_of_nat(x_202); +lean_dec(x_202); +x_204 = 0; +x_205 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_203, x_204, x_189); +x_206 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__10; +x_207 = l_Lean_mkSepArray(x_205, x_206); +lean_dec(x_205); +x_208 = l_Array_append___rarg(x_2, x_207); +x_209 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_209, 0, x_164); +lean_ctor_set(x_209, 1, x_168); +lean_ctor_set(x_209, 2, x_208); +x_210 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__26; +x_211 = lean_array_push(x_210, x_209); +x_212 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_212, 0, x_164); +lean_ctor_set(x_212, 1, x_168); +lean_ctor_set(x_212, 2, x_211); +x_213 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__11; +x_214 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_214, 0, x_195); +lean_ctor_set(x_214, 1, x_213); +x_215 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__28; +x_216 = lean_array_push(x_215, x_201); +x_217 = lean_array_push(x_216, x_212); +x_218 = lean_array_push(x_217, x_214); +x_219 = lean_array_push(x_218, x_192); +x_220 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__7; +x_221 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_221, 0, x_164); +lean_ctor_set(x_221, 1, x_220); +lean_ctor_set(x_221, 2, x_219); +if (lean_is_scalar(x_199)) { + x_222 = lean_alloc_ctor(0, 2, 0); } else { - x_224 = x_201; + x_222 = x_199; } -lean_ctor_set(x_224, 0, x_223); -lean_ctor_set(x_224, 1, x_200); -return x_224; +lean_ctor_set(x_222, 0, x_221); +lean_ctor_set(x_222, 1, x_198); +return x_222; } else { -lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; -lean_dec(x_191); +lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; +lean_dec(x_189); lean_dec(x_14); lean_dec(x_13); lean_dec(x_2); -x_225 = lean_ctor_get(x_193, 0); -lean_inc(x_225); -x_226 = lean_ctor_get(x_193, 1); -lean_inc(x_226); -if (lean_is_exclusive(x_193)) { - lean_ctor_release(x_193, 0); - lean_ctor_release(x_193, 1); - x_227 = x_193; +x_223 = lean_ctor_get(x_191, 0); +lean_inc(x_223); +x_224 = lean_ctor_get(x_191, 1); +lean_inc(x_224); +if (lean_is_exclusive(x_191)) { + lean_ctor_release(x_191, 0); + lean_ctor_release(x_191, 1); + x_225 = x_191; } else { - lean_dec_ref(x_193); - x_227 = lean_box(0); + lean_dec_ref(x_191); + x_225 = lean_box(0); } -if (lean_is_scalar(x_227)) { - x_228 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_225)) { + x_226 = lean_alloc_ctor(1, 2, 0); } else { - x_228 = x_227; + x_226 = x_225; } -lean_ctor_set(x_228, 0, x_225); -lean_ctor_set(x_228, 1, x_226); -return x_228; +lean_ctor_set(x_226, 0, x_223); +lean_ctor_set(x_226, 1, x_224); +return x_226; } } else { -lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; +lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); @@ -4213,32 +4220,32 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_229 = lean_ctor_get(x_148, 0); -lean_inc(x_229); -x_230 = lean_ctor_get(x_148, 1); -lean_inc(x_230); -if (lean_is_exclusive(x_148)) { - lean_ctor_release(x_148, 0); - lean_ctor_release(x_148, 1); - x_231 = x_148; +x_227 = lean_ctor_get(x_146, 0); +lean_inc(x_227); +x_228 = lean_ctor_get(x_146, 1); +lean_inc(x_228); +if (lean_is_exclusive(x_146)) { + lean_ctor_release(x_146, 0); + lean_ctor_release(x_146, 1); + x_229 = x_146; } else { - lean_dec_ref(x_148); - x_231 = lean_box(0); + lean_dec_ref(x_146); + x_229 = lean_box(0); } -if (lean_is_scalar(x_231)) { - x_232 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_229)) { + x_230 = lean_alloc_ctor(1, 2, 0); } else { - x_232 = x_231; + x_230 = x_229; } -lean_ctor_set(x_232, 0, x_229); -lean_ctor_set(x_232, 1, x_230); -return x_232; +lean_ctor_set(x_230, 0, x_227); +lean_ctor_set(x_230, 1, x_228); +return x_230; } } } else { -uint8_t x_233; +uint8_t x_231; lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); @@ -4251,23 +4258,23 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_233 = !lean_is_exclusive(x_18); -if (x_233 == 0) +x_231 = !lean_is_exclusive(x_18); +if (x_231 == 0) { return x_18; } else { -lean_object* x_234; lean_object* x_235; lean_object* x_236; -x_234 = lean_ctor_get(x_18, 0); -x_235 = lean_ctor_get(x_18, 1); -lean_inc(x_235); -lean_inc(x_234); +lean_object* x_232; lean_object* x_233; lean_object* x_234; +x_232 = lean_ctor_get(x_18, 0); +x_233 = lean_ctor_get(x_18, 1); +lean_inc(x_233); +lean_inc(x_232); lean_dec(x_18); -x_236 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_236, 0, x_234); -lean_ctor_set(x_236, 1, x_235); -return x_236; +x_234 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_234, 0, x_232); +lean_ctor_set(x_234, 1, x_233); +return x_234; } } } @@ -4322,61 +4329,74 @@ lean_ctor_set(x_3, 1, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(2); +x_2 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__18; +x_3 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__19; +x_4 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -if (lean_obj_tag(x_6) == 0) +if (lean_obj_tag(x_5) == 0) { -lean_object* x_15; -lean_dec(x_13); +lean_object* x_14; lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_5); +lean_dec(x_7); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_7); -lean_ctor_set(x_15, 1, x_14); -return x_15; +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_6); +lean_ctor_set(x_14, 1, x_13); +return x_14; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_23; 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_16 = lean_ctor_get(x_6, 0); +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; 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; uint8_t x_29; +x_15 = lean_ctor_get(x_5, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_5, 1); lean_inc(x_16); -x_17 = lean_ctor_get(x_6, 1); -lean_inc(x_17); -lean_dec(x_6); -x_23 = lean_ctor_get(x_1, 2); -lean_inc(x_23); -x_24 = lean_unsigned_to_nat(0u); -x_25 = lean_unsigned_to_nat(1u); -lean_inc(x_12); -lean_inc(x_3); -lean_inc(x_23); -x_26 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__3(x_23, x_24, x_23, x_25, x_3, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_23); +lean_dec(x_5); +x_22 = lean_ctor_get(x_1, 2); +lean_inc(x_22); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_unsigned_to_nat(1u); +x_25 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__19; +lean_inc(x_11); +lean_inc(x_22); +x_26 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__3(x_22, x_23, x_22, x_24, x_25, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_22); x_27 = lean_ctor_get(x_26, 0); lean_inc(x_27); x_28 = lean_ctor_get(x_26, 1); lean_inc(x_28); lean_dec(x_26); -x_29 = lean_name_eq(x_4, x_16); +x_29 = lean_name_eq(x_3, x_15); if (x_29 == 0) { lean_object* x_30; -lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); -lean_inc(x_16); -lean_inc(x_4); -x_30 = l_Lean_Meta_compatibleCtors(x_4, x_16, x_10, x_11, x_12, x_13, x_28); +lean_inc(x_9); +lean_inc(x_15); +lean_inc(x_3); +x_30 = l_Lean_Meta_compatibleCtors(x_3, x_15, x_9, x_10, x_11, x_12, x_28); if (lean_obj_tag(x_30) == 0) { lean_object* x_31; uint8_t x_32; @@ -4388,15 +4408,15 @@ if (x_32 == 0) { lean_object* x_33; lean_object* x_34; lean_dec(x_27); -lean_dec(x_16); +lean_dec(x_15); x_33 = lean_ctor_get(x_30, 1); lean_inc(x_33); lean_dec(x_30); x_34 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_34, 0, x_7); -x_18 = x_34; -x_19 = x_33; -goto block_22; +lean_ctor_set(x_34, 0, x_6); +x_17 = x_34; +x_18 = x_33; +goto block_21; } else { @@ -4404,19 +4424,19 @@ lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean x_35 = lean_ctor_get(x_30, 1); lean_inc(x_35); lean_dec(x_30); -lean_inc(x_12); -x_36 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_12, x_13, x_35); +lean_inc(x_11); +x_36 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_11, x_12, x_35); x_37 = lean_ctor_get(x_36, 0); lean_inc(x_37); x_38 = lean_ctor_get(x_36, 1); lean_inc(x_38); lean_dec(x_36); -x_39 = lean_st_ref_get(x_13, x_38); +x_39 = lean_st_ref_get(x_12, x_38); x_40 = lean_ctor_get(x_39, 1); lean_inc(x_40); lean_dec(x_39); -lean_inc(x_4); -x_41 = lean_mk_syntax_ident(x_4); +lean_inc(x_3); +x_41 = lean_mk_syntax_ident(x_3); x_42 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___closed__3; x_43 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_43, 0, x_37); @@ -4443,18 +4463,18 @@ x_56 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_56, 0, x_46); lean_ctor_set(x_56, 1, x_55); lean_ctor_set(x_56, 2, x_54); -lean_inc(x_12); -x_57 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_12, x_13, x_40); +lean_inc(x_11); +x_57 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_11, x_12, x_40); x_58 = lean_ctor_get(x_57, 0); lean_inc(x_58); x_59 = lean_ctor_get(x_57, 1); lean_inc(x_59); lean_dec(x_57); -x_60 = lean_st_ref_get(x_13, x_59); +x_60 = lean_st_ref_get(x_12, x_59); x_61 = lean_ctor_get(x_60, 1); lean_inc(x_61); lean_dec(x_60); -x_62 = lean_mk_syntax_ident(x_16); +x_62 = lean_mk_syntax_ident(x_15); x_63 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_63, 0, x_58); lean_ctor_set(x_63, 1, x_42); @@ -4477,16 +4497,16 @@ lean_ctor_set(x_70, 2, x_69); x_71 = lean_array_push(x_52, x_56); x_72 = lean_array_push(x_71, x_70); x_73 = l_Array_append___rarg(x_27, x_72); -lean_inc(x_12); -x_74 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_12, x_13, x_61); +lean_inc(x_11); +x_74 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_11, x_12, x_61); x_75 = lean_ctor_get(x_74, 0); lean_inc(x_75); x_76 = lean_ctor_get(x_74, 1); lean_inc(x_76); lean_dec(x_74); -x_77 = lean_ctor_get(x_12, 10); +x_77 = lean_ctor_get(x_11, 10); lean_inc(x_77); -x_78 = lean_st_ref_get(x_13, x_76); +x_78 = lean_st_ref_get(x_12, x_76); x_79 = lean_ctor_get(x_78, 0); lean_inc(x_79); x_80 = lean_ctor_get(x_78, 1); @@ -4568,24 +4588,18 @@ lean_inc(x_75); x_114 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_114, 0, x_75); lean_ctor_set(x_114, 1, x_113); -lean_inc(x_3); -x_115 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_115, 0, x_46); -lean_ctor_set(x_115, 1, x_50); -lean_ctor_set(x_115, 2, x_3); -x_116 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__12; -x_117 = lean_array_push(x_116, x_114); -x_118 = lean_array_push(x_117, x_98); -lean_inc(x_115); -x_119 = lean_array_push(x_118, x_115); +x_115 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__12; +x_116 = lean_array_push(x_115, x_114); +x_117 = lean_array_push(x_116, x_98); +x_118 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___closed__6; +x_119 = lean_array_push(x_117, x_118); x_120 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__49; x_121 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_121, 0, x_46); lean_ctor_set(x_121, 1, x_120); lean_ctor_set(x_121, 2, x_119); x_122 = lean_array_push(x_52, x_121); -lean_inc(x_115); -x_123 = lean_array_push(x_122, x_115); +x_123 = lean_array_push(x_122, x_118); x_124 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_124, 0, x_46); lean_ctor_set(x_124, 1, x_111); @@ -4616,7 +4630,7 @@ lean_ctor_set(x_137, 0, x_46); lean_ctor_set(x_137, 1, x_136); lean_ctor_set(x_137, 2, x_135); x_138 = lean_array_push(x_52, x_137); -x_139 = lean_array_push(x_138, x_115); +x_139 = lean_array_push(x_138, x_118); x_140 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_140, 0, x_46); lean_ctor_set(x_140, 1, x_50); @@ -4625,7 +4639,7 @@ x_141 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__5 x_142 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_142, 0, x_75); lean_ctor_set(x_142, 1, x_141); -x_143 = lean_array_push(x_116, x_90); +x_143 = lean_array_push(x_115, x_90); x_144 = lean_array_push(x_143, x_140); x_145 = lean_array_push(x_144, x_142); x_146 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__40; @@ -4644,14 +4658,14 @@ x_152 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_152, 0, x_46); lean_ctor_set(x_152, 1, x_55); lean_ctor_set(x_152, 2, x_151); -lean_inc(x_12); -x_153 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_12, x_13, x_80); +lean_inc(x_11); +x_153 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_11, x_12, x_80); x_154 = lean_ctor_get(x_153, 0); lean_inc(x_154); x_155 = lean_ctor_get(x_153, 1); lean_inc(x_155); lean_dec(x_153); -x_156 = lean_st_ref_get(x_13, x_155); +x_156 = lean_st_ref_get(x_12, x_155); x_157 = lean_ctor_get(x_156, 1); lean_inc(x_157); lean_dec(x_156); @@ -4668,8 +4682,7 @@ x_163 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__ x_164 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__10; x_165 = l_Lean_mkSepArray(x_163, x_164); lean_dec(x_163); -lean_inc(x_3); -x_166 = l_Array_append___rarg(x_3, x_165); +x_166 = l_Array_append___rarg(x_25, x_165); x_167 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_167, 0, x_46); lean_ctor_set(x_167, 1, x_50); @@ -4693,28 +4706,27 @@ x_178 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_178, 0, x_46); lean_ctor_set(x_178, 1, x_177); lean_ctor_set(x_178, 2, x_176); -x_179 = lean_array_push(x_7, x_178); +x_179 = lean_array_push(x_6, x_178); x_180 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_180, 0, x_179); -x_18 = x_180; -x_19 = x_157; -goto block_22; +x_17 = x_180; +x_18 = x_157; +goto block_21; } } else { uint8_t x_181; lean_dec(x_27); -lean_dec(x_17); lean_dec(x_16); -lean_dec(x_13); +lean_dec(x_15); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_5); +lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); @@ -4742,31 +4754,30 @@ return x_184; else { lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; -lean_dec(x_16); -x_185 = lean_ctor_get(x_5, 0); +lean_dec(x_15); +x_185 = lean_ctor_get(x_4, 0); lean_inc(x_185); x_186 = lean_ctor_get(x_185, 2); lean_inc(x_186); lean_dec(x_185); lean_inc(x_2); -lean_inc(x_4); -lean_inc(x_5); lean_inc(x_3); +lean_inc(x_4); lean_inc(x_1); x_187 = lean_alloc_closure((void*)(l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___boxed), 15, 6); lean_closure_set(x_187, 0, x_1); -lean_closure_set(x_187, 1, x_3); -lean_closure_set(x_187, 2, x_5); -lean_closure_set(x_187, 3, x_4); +lean_closure_set(x_187, 1, x_25); +lean_closure_set(x_187, 2, x_4); +lean_closure_set(x_187, 3, x_3); lean_closure_set(x_187, 4, x_27); lean_closure_set(x_187, 5, x_2); -lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -x_188 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Elab_Deriving_mkInductArgNames___spec__2___rarg(x_186, x_187, x_8, x_9, x_10, x_11, x_12, x_13, x_28); +lean_inc(x_7); +x_188 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Elab_Deriving_mkInductArgNames___spec__2___rarg(x_186, x_187, x_7, x_8, x_9, x_10, x_11, x_12, x_28); if (lean_obj_tag(x_188) == 0) { lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; @@ -4775,25 +4786,24 @@ lean_inc(x_189); x_190 = lean_ctor_get(x_188, 1); lean_inc(x_190); lean_dec(x_188); -x_191 = lean_array_push(x_7, x_189); +x_191 = lean_array_push(x_6, x_189); x_192 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_192, 0, x_191); -x_18 = x_192; -x_19 = x_190; -goto block_22; +x_17 = x_192; +x_18 = x_190; +goto block_21; } else { uint8_t x_193; -lean_dec(x_17); -lean_dec(x_13); +lean_dec(x_16); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_5); +lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); @@ -4818,151 +4828,147 @@ return x_196; } } } -block_22: +block_21: { -lean_object* x_20; -x_20 = lean_ctor_get(x_18, 0); -lean_inc(x_20); -lean_dec(x_18); -x_6 = x_17; -x_7 = x_20; -x_14 = x_19; +lean_object* x_19; +x_19 = lean_ctor_get(x_17, 0); +lean_inc(x_19); +lean_dec(x_17); +x_5 = x_16; +x_6 = x_19; +x_13 = x_18; goto _start; } } } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__7(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__7(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -if (lean_obj_tag(x_5) == 0) +if (lean_obj_tag(x_4) == 0) { -lean_object* x_14; -lean_dec(x_12); +lean_object* x_13; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_4); +lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_6); -lean_ctor_set(x_14, 1, x_13); -return x_14; +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_5); +lean_ctor_set(x_13, 1, x_12); +return x_13; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_5, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_5, 1); -lean_inc(x_16); -lean_dec(x_5); -lean_inc(x_7); +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_4, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_4, 1); lean_inc(x_15); -x_17 = l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__1(x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_17) == 0) +lean_dec(x_4); +lean_inc(x_6); +lean_inc(x_14); +x_16 = l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__1(x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_17, 0); +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); -lean_inc(x_19); -lean_dec(x_17); -lean_inc(x_12); +lean_dec(x_16); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -lean_inc(x_4); +lean_inc(x_6); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_20 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6(x_1, x_2, x_3, x_15, x_18, x_4, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_19); -if (lean_obj_tag(x_20) == 0) +x_19 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6(x_1, x_2, x_14, x_17, x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_18); +if (lean_obj_tag(x_19) == 0) { -lean_object* x_21; lean_object* x_22; -x_21 = lean_ctor_get(x_20, 0); +lean_object* x_20; lean_object* x_21; +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_20, 1); -lean_inc(x_22); -lean_dec(x_20); -x_5 = x_16; -x_6 = x_21; -x_13 = x_22; +lean_dec(x_19); +x_4 = x_15; +x_5 = x_20; +x_12 = x_21; goto _start; } else { -uint8_t x_24; -lean_dec(x_16); -lean_dec(x_12); +uint8_t x_23; +lean_dec(x_15); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_4); +lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_24 = !lean_is_exclusive(x_20); -if (x_24 == 0) +x_23 = !lean_is_exclusive(x_19); +if (x_23 == 0) { -return x_20; +return x_19; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_20, 0); -x_26 = lean_ctor_get(x_20, 1); -lean_inc(x_26); +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_19, 0); +x_25 = lean_ctor_get(x_19, 1); lean_inc(x_25); -lean_dec(x_20); -x_27 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_27, 0, x_25); -lean_ctor_set(x_27, 1, x_26); -return x_27; +lean_inc(x_24); +lean_dec(x_19); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; } } } else { -uint8_t x_28; -lean_dec(x_16); +uint8_t x_27; lean_dec(x_15); -lean_dec(x_12); +lean_dec(x_14); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_4); +lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_28 = !lean_is_exclusive(x_17); -if (x_28 == 0) +x_27 = !lean_is_exclusive(x_16); +if (x_27 == 0) { -return x_17; +return x_16; } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_17, 0); -x_30 = lean_ctor_get(x_17, 1); -lean_inc(x_30); +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_16, 0); +x_29 = lean_ctor_get(x_16, 1); lean_inc(x_29); -lean_dec(x_17); -x_31 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_30); -return x_31; +lean_inc(x_28); +lean_dec(x_16); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; } } } @@ -4976,7 +4982,7 @@ x_10 = lean_ctor_get(x_1, 4); lean_inc(x_10); x_11 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__19; lean_inc(x_10); -x_12 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__7(x_1, x_2, x_11, x_10, x_10, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_12 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__7(x_1, x_2, x_10, x_10, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9); if (lean_obj_tag(x_12) == 0) { uint8_t x_13; @@ -5104,6 +5110,30 @@ lean_dec(x_7); return x_16; } } +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_DecEq_mkMatch___spec__1(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = 1; +x_9 = lean_usize_add(x_2, x_8); +x_10 = lean_array_uset(x_7, x_2, x_5); +x_2 = x_9; +x_3 = x_10; +goto _start; +} +} +} static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkMatch___closed__1() { _start: { @@ -5203,7 +5233,7 @@ x_25 = lean_array_get_size(x_12); x_26 = lean_usize_of_nat(x_25); lean_dec(x_25); x_27 = 0; -x_28 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_26, x_27, x_12); +x_28 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_DecEq_mkMatch___spec__1(x_26, x_27, x_12); x_29 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__10; x_30 = l_Lean_mkSepArray(x_28, x_29); lean_dec(x_28); @@ -5262,7 +5292,7 @@ x_57 = lean_array_get_size(x_12); x_58 = lean_usize_of_nat(x_57); lean_dec(x_57); x_59 = 0; -x_60 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_58, x_59, x_12); +x_60 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_DecEq_mkMatch___spec__1(x_58, x_59, x_12); x_61 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__10; x_62 = l_Lean_mkSepArray(x_60, x_61); lean_dec(x_60); @@ -5336,6 +5366,18 @@ return x_90; } } } +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_DecEq_mkMatch___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_DecEq_mkMatch___spec__1(x_4, x_5, x_3); +return x_6; +} +} static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__1() { _start: { @@ -6498,7 +6540,7 @@ LEAN_EXPORT lean_object* l_Lean_getConstInfoInduct___at_Lean_Elab_Deriving_DecEq lean_object* x_5; lean_inc(x_2); lean_inc(x_1); -x_5 = l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__10(x_1, x_2, x_3, x_4); +x_5 = l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__14(x_1, x_2, x_3, x_4); if (lean_obj_tag(x_5) == 0) { lean_object* x_6; @@ -9297,7 +9339,7 @@ x_9 = lean_ctor_get(x_1, 1); lean_inc(x_9); lean_dec(x_1); lean_inc(x_2); -x_17 = l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__10(x_8, x_2, x_3, x_4); +x_17 = l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__14(x_8, x_2, x_3, x_4); if (lean_obj_tag(x_17) == 0) { lean_object* x_18; @@ -9388,7 +9430,7 @@ LEAN_EXPORT lean_object* l_Lean_isEnumType___at_Lean_Elab_Deriving_DecEq_mkDecEq { lean_object* x_5; lean_inc(x_2); -x_5 = l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__10(x_1, x_2, x_3, x_4); +x_5 = l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__14(x_1, x_2, x_3, x_4); if (lean_obj_tag(x_5) == 0) { lean_object* x_6; @@ -9945,7 +9987,7 @@ lean_dec(x_1); return x_5; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4892____closed__1() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_5097____closed__1() { _start: { lean_object* x_1; @@ -9953,12 +9995,12 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Deriving_DecEq_mkDecEqInstanceHandl return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4892_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_5097_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; x_2 = l_Lean_Elab_Deriving_DecEq_mkDecEqHeader___closed__2; -x_3 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4892____closed__1; +x_3 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_5097____closed__1; x_4 = l_Lean_Elab_registerBuiltinDerivingHandler(x_2, x_3, x_1); if (lean_obj_tag(x_4) == 0) { @@ -10285,6 +10327,8 @@ l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___close lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___closed__4); l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___closed__5 = _init_l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___closed__5(); lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___closed__5); +l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___closed__6 = _init_l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___closed__6(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___closed__6); l_Lean_Elab_Deriving_DecEq_mkMatch___closed__1 = _init_l_Lean_Elab_Deriving_DecEq_mkMatch___closed__1(); lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_mkMatch___closed__1); l_Lean_Elab_Deriving_DecEq_mkMatch___closed__2 = _init_l_Lean_Elab_Deriving_DecEq_mkMatch___closed__2(); @@ -10557,9 +10601,9 @@ l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__71 = _init_l_Lean_Elab_Deriving lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__71); l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__72 = _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__72(); lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__72); -l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4892____closed__1 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4892____closed__1(); -lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4892____closed__1); -res = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4892_(lean_io_mk_world()); +l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_5097____closed__1 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_5097____closed__1(); +lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_5097____closed__1); +res = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_5097_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/Deriving/FromToJson.c b/stage0/stdlib/Lean/Elab/Deriving/FromToJson.c index 967cdcbbc3c..1bce8c7417e 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/FromToJson.c +++ b/stage0/stdlib/Lean/Elab/Deriving/FromToJson.c @@ -27,13 +27,13 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJs static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__4___closed__12; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__13; lean_object* l_Lean_stringToMessageData(lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__8; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__5___closed__2; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__20; lean_object* lean_mk_empty_array_with_capacity(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__21; lean_object* lean_nat_div(lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__13; +lean_object* l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__14(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__11; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__23; lean_object* l_Lean_LocalDecl_userName(lean_object*); @@ -42,6 +42,7 @@ lean_object* l_Lean_Elab_Command_liftTermElabM___rarg(lean_object*, lean_object* lean_object* lean_name_mk_string(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__1___closed__13; uint8_t lean_usize_dec_eq(size_t, size_t); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__18; lean_object* lean_array_uget(lean_object*, size_t); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__4___closed__6; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__15; @@ -60,13 +61,13 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__4___closed__13; size_t lean_usize_sub(size_t, size_t); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__24; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__5(lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__5(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__4; lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__4___closed__3; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___boxed(lean_object*, lean_object*, 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_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__22; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__2; @@ -103,11 +104,10 @@ static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___la lean_object* lean_string_utf8_byte_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); -extern lean_object* l_Lean_nameLitKind; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__2___closed__12; lean_object* l_Lean_Elab_Deriving_mkLocalInstanceLetDecls(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__3; -LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__2; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__1; @@ -119,12 +119,11 @@ lean_object* lean_nat_add(lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__1; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__3; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__2___closed__7; -lean_object* l_Lean_Syntax_SepArray_ofElems(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__3; lean_object* l_Array_zip___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__8___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__34; static lean_object* l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__3___closed__16; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__4___closed__9; @@ -137,7 +136,6 @@ static lean_object* l_Lean_getConstInfoInduct___at_Lean_Elab_Deriving_FromToJson static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__19; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__4___closed__4; lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(size_t, size_t, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__5; LEAN_EXPORT lean_object* l_Lean_getConstInfoInduct___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__9; @@ -157,14 +155,14 @@ static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mk static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__12; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__2; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__9; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__3___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__4(size_t, size_t, lean_object*); lean_object* l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__8; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__2___closed__5; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__3___closed__3; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__8; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__5___closed__1; -extern lean_object* l_Lean_numLitKind; static lean_object* l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__3___closed__10; lean_object* l_Lean_Elab_Deriving_mkInstanceCmds(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__2___boxed(lean_object**); @@ -174,13 +172,12 @@ static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mk lean_object* lean_array_swap(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getStructureFieldsFlattened(lean_object*, lean_object*, uint8_t); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__8; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__1; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__1; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__4___closed__14; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__2; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__3___closed__10; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3(lean_object*, lean_object*, 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_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___boxed(lean_object*, lean_object*, lean_object*, 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_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_toString(lean_object*, uint8_t); lean_object* l_Lean_Elab_Deriving_mkLet(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); @@ -192,19 +189,20 @@ lean_object* l_Lean_Expr_fvarId_x21(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__7; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__4___closed__15; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__2; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__26; static lean_object* l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__3___closed__5; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__7; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__11; static lean_object* l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__3___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__2___boxed(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_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__4___closed__5; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__24; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__2___closed__1; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__7; static lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__1___closed__2; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__9; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__18; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__16; lean_object* l_Nat_repr(lean_object*); @@ -213,18 +211,16 @@ static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mk static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__9; static lean_object* l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__3___closed__7; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__5; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__4; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__4; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__3___closed__16; -LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6045_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6284_(lean_object*); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__2___closed__8; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__1___closed__7; lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Init_Meta_0__Lean_quoteNameMk(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__5; static lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__1___closed__1; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__2___closed__9; -LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__2___boxed(lean_object*, lean_object*, lean_object*, 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_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__2___boxed(lean_object*, lean_object*, 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_getConstInfoCtor___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__4___closed__5; @@ -241,7 +237,6 @@ extern lean_object* l_Lean_instInhabitedExpr; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__4___lambda__1___closed__2; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__1; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__11; -extern lean_object* l_Lean_instInhabitedSyntax; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__4; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__4___lambda__1___boxed(lean_object*, lean_object*, 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_getConstInfoInduct___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); @@ -259,27 +254,30 @@ static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mk static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__3___closed__13; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__12; size_t lean_usize_of_nat(lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__19; lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__27; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__2; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__4___lambda__1___closed__1; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__6___closed__1; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__6; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__4; lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__16(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__1___closed__2; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__3___closed__9; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__6; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__8; -static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6045____closed__1; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__4; static lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__1___closed__4; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__1; +lean_object* l_Lean_Syntax_mkNumLit(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__9; lean_object* l_Lean_LocalDecl_type(lean_object*); lean_object* l_Lean_Elab_Deriving_mkHeader(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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__1___lambda__1(lean_object*, 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_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__3___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__4___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__1; static lean_object* l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__3___closed__15; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__6; uint8_t lean_uint32_dec_eq(uint32_t, uint32_t); @@ -292,6 +290,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__31; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__1___closed__11; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__6; +static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6284____closed__1; lean_object* l_Lean_Meta_getLocalDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_environment_main_module(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__27; @@ -305,26 +304,31 @@ static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mk uint8_t lean_nat_dec_le(lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__5___closed__2; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__10; +static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6284____closed__2; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_append(lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__5; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__4___closed__2; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__10; static lean_object* l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__3___closed__6; +lean_object* l_Lean_quoteNameMk(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkJsonField(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__3(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__12; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__3; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__19; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__1___closed__12; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__4___closed__8; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6(size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__28; lean_object* l_Lean_Elab_Command_getRef(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__1___closed__2; +lean_object* l_Lean_Syntax_mkNameLit(lean_object*, lean_object*); static lean_object* l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__3; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__3; LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -354,6 +358,7 @@ static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___la static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__14; static lean_object* l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__3___closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__20; LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__3___closed__17; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__5; @@ -364,6 +369,7 @@ static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mk static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__29; lean_object* l_Lean_Meta_forallTelescopeReducing___at_Lean_Elab_Deriving_mkInductArgNames___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__15; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__7; LEAN_EXPORT uint8_t l_Array_qsort_sort___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__7___lambda__1(lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__2; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__3___closed__19; @@ -373,7 +379,6 @@ static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___ static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__20; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__2___closed__11; lean_object* lean_mk_syntax_ident(lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__7; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__4___closed__16; extern lean_object* l_Lean_instInhabitedInductiveVal; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__10; @@ -391,7 +396,6 @@ static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mk static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__3___closed__2; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__3; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__6; -static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6045____closed__2; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__9; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__6___closed__2; @@ -401,11 +405,10 @@ static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mk LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkConst(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__3___closed__8; -lean_object* l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__10(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__4___closed__11; LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__2(lean_object*, lean_object*, lean_object*, 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_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__17; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__8; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__35; @@ -425,7 +428,6 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJs static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__1___closed__4; lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); uint8_t lean_string_dec_eq(lean_object*, lean_object*); -lean_object* l_Lean_Syntax_mkLit(lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__3___closed__2; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__1; @@ -1826,7 +1828,7 @@ LEAN_EXPORT lean_object* l_Lean_getConstInfoInduct___at_Lean_Elab_Deriving_FromT lean_object* x_5; lean_inc(x_2); lean_inc(x_1); -x_5 = l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__10(x_1, x_2, x_3, x_4); +x_5 = l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__14(x_1, x_2, x_3, x_4); if (lean_obj_tag(x_5) == 0) { lean_object* x_6; @@ -2344,6 +2346,7 @@ lean_ctor_set(x_38, 1, x_37); x_39 = l_Lean_instInhabitedName; x_40 = lean_array_get(x_39, x_6, x_9); x_41 = l_Lean_Name_getString_x21(x_40); +lean_dec(x_40); x_42 = lean_box(2); x_43 = l_Lean_Syntax_mkStrLit(x_41, x_42); lean_dec(x_41); @@ -2604,6 +2607,7 @@ lean_ctor_set(x_38, 1, x_37); x_39 = l_Lean_instInhabitedName; x_40 = lean_array_get(x_39, x_5, x_9); x_41 = l_Lean_Name_getString_x21(x_40); +lean_dec(x_40); x_42 = lean_box(2); x_43 = l_Lean_Syntax_mkStrLit(x_41, x_42); lean_dec(x_41); @@ -2791,7 +2795,31 @@ return x_134; } } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__1() { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = 1; +x_9 = lean_usize_add(x_2, x_8); +x_10 = lean_array_uset(x_7, x_2, x_5); +x_2 = x_9; +x_3 = x_10; +goto _start; +} +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__1() { _start: { lean_object* x_1; @@ -2799,17 +2827,17 @@ x_1 = lean_mk_string_from_bytes("term[_]", 7); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__2() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__1; +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__3() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__3() { _start: { lean_object* x_1; @@ -2817,7 +2845,7 @@ x_1 = lean_mk_string_from_bytes("[", 1); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__4() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__4() { _start: { lean_object* x_1; @@ -2825,7 +2853,7 @@ x_1 = lean_mk_string_from_bytes("]", 1); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__5() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__5() { _start: { lean_object* x_1; @@ -2833,22 +2861,22 @@ x_1 = lean_mk_string_from_bytes("opt", 3); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__6() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__5; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__5; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__7() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__5; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__5; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__6; +x_3 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__6; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2856,17 +2884,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__8() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__5; +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__9() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__9() { _start: { lean_object* x_1; @@ -2874,7 +2902,7 @@ x_1 = lean_mk_string_from_bytes("Json", 4); return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6(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_object* x_7, 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_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7(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_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { uint8_t x_14; @@ -2931,7 +2959,7 @@ x_32 = lean_ctor_get(x_30, 0); lean_inc(x_32); lean_dec(x_30); x_33 = lean_environment_main_module(x_32); -x_34 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__3; +x_34 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__3; lean_inc(x_26); x_35 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_35, 0, x_26); @@ -3038,14 +3066,14 @@ x_90 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_90, 0, x_65); lean_ctor_set(x_90, 1, x_66); lean_ctor_set(x_90, 2, x_89); -x_91 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__4; +x_91 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__4; x_92 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_92, 0, x_26); lean_ctor_set(x_92, 1, x_91); x_93 = lean_array_push(x_84, x_35); x_94 = lean_array_push(x_93, x_90); x_95 = lean_array_push(x_94, x_92); -x_96 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__2; +x_96 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__2; x_97 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_97, 0, x_65); lean_ctor_set(x_97, 1, x_96); @@ -3083,7 +3111,7 @@ x_110 = lean_ctor_get(x_108, 0); lean_inc(x_110); lean_dec(x_108); x_111 = lean_environment_main_module(x_110); -x_112 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__3; +x_112 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__3; lean_inc(x_104); x_113 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_113, 0, x_104); @@ -3191,14 +3219,14 @@ x_169 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_169, 0, x_144); lean_ctor_set(x_169, 1, x_145); lean_ctor_set(x_169, 2, x_168); -x_170 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__4; +x_170 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__4; x_171 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_171, 0, x_104); lean_ctor_set(x_171, 1, x_170); x_172 = lean_array_push(x_163, x_113); x_173 = lean_array_push(x_172, x_169); x_174 = lean_array_push(x_173, x_171); -x_175 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__2; +x_175 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__2; x_176 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_176, 0, x_144); lean_ctor_set(x_176, 1, x_175); @@ -3248,12 +3276,12 @@ x_195 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstan x_196 = lean_name_mk_string(x_194, x_195); x_197 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__2; x_198 = lean_name_mk_string(x_196, x_197); -x_199 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__8; +x_199 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__8; x_200 = l_Lean_addMacroScope(x_192, x_199, x_187); -x_201 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__9; +x_201 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__9; lean_inc(x_1); x_202 = lean_name_mk_string(x_1, x_201); -x_203 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__5; +x_203 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__5; x_204 = lean_name_mk_string(x_202, x_203); x_205 = lean_box(0); lean_ctor_set(x_19, 1, x_205); @@ -3261,7 +3289,7 @@ lean_ctor_set(x_19, 0, x_204); x_206 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_206, 0, x_19); lean_ctor_set(x_206, 1, x_205); -x_207 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__7; +x_207 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__7; x_208 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_208, 0, x_185); lean_ctor_set(x_208, 1, x_207); @@ -3328,12 +3356,12 @@ x_239 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstan x_240 = lean_name_mk_string(x_238, x_239); x_241 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__2; x_242 = lean_name_mk_string(x_240, x_241); -x_243 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__8; +x_243 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__8; x_244 = l_Lean_addMacroScope(x_236, x_243, x_231); -x_245 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__9; +x_245 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__9; lean_inc(x_1); x_246 = lean_name_mk_string(x_1, x_245); -x_247 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__5; +x_247 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__5; x_248 = lean_name_mk_string(x_246, x_247); x_249 = lean_box(0); x_250 = lean_alloc_ctor(0, 2, 0); @@ -3342,7 +3370,7 @@ lean_ctor_set(x_250, 1, x_249); x_251 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_251, 0, x_250); lean_ctor_set(x_251, 1, x_249); -x_252 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__7; +x_252 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__7; x_253 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_253, 0, x_229); lean_ctor_set(x_253, 1, x_252); @@ -3458,7 +3486,7 @@ static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandle { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__9; +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__9; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } @@ -3567,7 +3595,7 @@ x_40 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___clo lean_inc(x_28); lean_inc(x_33); x_41 = l_Lean_addMacroScope(x_33, x_40, x_28); -x_42 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__9; +x_42 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__9; x_43 = lean_name_mk_string(x_3, x_42); x_44 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__1; lean_inc(x_43); @@ -3586,7 +3614,7 @@ lean_ctor_set(x_50, 0, x_26); lean_ctor_set(x_50, 1, x_49); lean_ctor_set(x_50, 2, x_41); lean_ctor_set(x_50, 3, x_48); -x_51 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__3; +x_51 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__3; lean_inc(x_26); x_52 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_52, 0, x_26); @@ -3600,11 +3628,7 @@ x_56 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_56, 0, x_26); lean_ctor_set(x_56, 1, x_55); x_57 = lean_ctor_get(x_5, 0); -lean_inc(x_57); -lean_dec(x_5); x_58 = lean_ctor_get(x_57, 0); -lean_inc(x_58); -lean_dec(x_57); x_59 = l_Lean_Name_getString_x21(x_58); x_60 = lean_box(2); x_61 = l_Lean_Syntax_mkStrLit(x_59, x_60); @@ -3652,7 +3676,7 @@ x_84 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_84, 0, x_60); lean_ctor_set(x_84, 1, x_83); lean_ctor_set(x_84, 2, x_82); -x_85 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__4; +x_85 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__4; lean_inc(x_26); x_86 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_86, 0, x_26); @@ -3722,7 +3746,7 @@ lean_ctor_set(x_117, 2, x_116); x_118 = lean_array_push(x_87, x_52); x_119 = lean_array_push(x_118, x_117); x_120 = lean_array_push(x_119, x_86); -x_121 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__2; +x_121 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__2; x_122 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_122, 0, x_60); lean_ctor_set(x_122, 1, x_121); @@ -3765,7 +3789,7 @@ x_138 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___cl lean_inc(x_28); lean_inc(x_131); x_139 = l_Lean_addMacroScope(x_131, x_138, x_28); -x_140 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__9; +x_140 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__9; x_141 = lean_name_mk_string(x_3, x_140); x_142 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__1; lean_inc(x_141); @@ -3784,7 +3808,7 @@ lean_ctor_set(x_148, 0, x_26); lean_ctor_set(x_148, 1, x_147); lean_ctor_set(x_148, 2, x_139); lean_ctor_set(x_148, 3, x_146); -x_149 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__3; +x_149 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__3; lean_inc(x_26); x_150 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_150, 0, x_26); @@ -3798,11 +3822,7 @@ x_154 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_154, 0, x_26); lean_ctor_set(x_154, 1, x_153); x_155 = lean_ctor_get(x_5, 0); -lean_inc(x_155); -lean_dec(x_5); x_156 = lean_ctor_get(x_155, 0); -lean_inc(x_156); -lean_dec(x_155); x_157 = l_Lean_Name_getString_x21(x_156); x_158 = lean_box(2); x_159 = l_Lean_Syntax_mkStrLit(x_157, x_158); @@ -3850,7 +3870,7 @@ x_182 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_182, 0, x_158); lean_ctor_set(x_182, 1, x_181); lean_ctor_set(x_182, 2, x_180); -x_183 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__4; +x_183 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__4; lean_inc(x_26); x_184 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_184, 0, x_26); @@ -3920,7 +3940,7 @@ lean_ctor_set(x_215, 2, x_214); x_216 = lean_array_push(x_185, x_150); x_217 = lean_array_push(x_216, x_215); x_218 = lean_array_push(x_217, x_184); -x_219 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__2; +x_219 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__2; x_220 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_220, 0, x_158); lean_ctor_set(x_220, 1, x_219); @@ -3986,7 +4006,7 @@ lean_inc(x_244); x_246 = lean_name_mk_string(x_244, x_245); x_247 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__4; x_248 = l_Lean_addMacroScope(x_240, x_247, x_235); -x_249 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__9; +x_249 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__9; x_250 = lean_name_mk_string(x_3, x_249); x_251 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__1; x_252 = lean_name_mk_string(x_250, x_251); @@ -4004,7 +4024,7 @@ lean_ctor_set(x_257, 0, x_233); lean_ctor_set(x_257, 1, x_256); lean_ctor_set(x_257, 2, x_248); lean_ctor_set(x_257, 3, x_255); -x_258 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__3; +x_258 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__3; lean_inc(x_233); x_259 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_259, 0, x_233); @@ -4018,11 +4038,7 @@ x_263 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_263, 0, x_233); lean_ctor_set(x_263, 1, x_262); x_264 = lean_ctor_get(x_5, 0); -lean_inc(x_264); -lean_dec(x_5); x_265 = lean_ctor_get(x_264, 0); -lean_inc(x_265); -lean_dec(x_264); x_266 = l_Lean_Name_getString_x21(x_265); x_267 = lean_box(2); x_268 = l_Lean_Syntax_mkStrLit(x_266, x_267); @@ -4049,7 +4065,7 @@ x_282 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_282, 0, x_267); lean_ctor_set(x_282, 1, x_281); lean_ctor_set(x_282, 2, x_280); -x_283 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__4; +x_283 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__4; lean_inc(x_233); x_284 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_284, 0, x_233); @@ -4060,7 +4076,7 @@ lean_inc(x_286); x_287 = lean_array_push(x_286, x_282); lean_inc(x_284); x_288 = lean_array_push(x_287, x_284); -x_289 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__2; +x_289 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__2; x_290 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_290, 0, x_267); lean_ctor_set(x_290, 1, x_289); @@ -4159,7 +4175,7 @@ lean_inc(x_330); x_332 = lean_name_mk_string(x_330, x_331); x_333 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__4; x_334 = l_Lean_addMacroScope(x_326, x_333, x_235); -x_335 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__9; +x_335 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__9; x_336 = lean_name_mk_string(x_3, x_335); x_337 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__1; x_338 = lean_name_mk_string(x_336, x_337); @@ -4177,7 +4193,7 @@ lean_ctor_set(x_343, 0, x_233); lean_ctor_set(x_343, 1, x_342); lean_ctor_set(x_343, 2, x_334); lean_ctor_set(x_343, 3, x_341); -x_344 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__3; +x_344 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__3; lean_inc(x_233); x_345 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_345, 0, x_233); @@ -4191,11 +4207,7 @@ x_349 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_349, 0, x_233); lean_ctor_set(x_349, 1, x_348); x_350 = lean_ctor_get(x_5, 0); -lean_inc(x_350); -lean_dec(x_5); x_351 = lean_ctor_get(x_350, 0); -lean_inc(x_351); -lean_dec(x_350); x_352 = l_Lean_Name_getString_x21(x_351); x_353 = lean_box(2); x_354 = l_Lean_Syntax_mkStrLit(x_352, x_353); @@ -4222,7 +4234,7 @@ x_368 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_368, 0, x_353); lean_ctor_set(x_368, 1, x_367); lean_ctor_set(x_368, 2, x_366); -x_369 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__4; +x_369 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__4; lean_inc(x_233); x_370 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_370, 0, x_233); @@ -4233,7 +4245,7 @@ lean_inc(x_372); x_373 = lean_array_push(x_372, x_368); lean_inc(x_370); x_374 = lean_array_push(x_373, x_370); -x_375 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__2; +x_375 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__2; x_376 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_376, 0, x_353); lean_ctor_set(x_376, 1, x_375); @@ -4494,7 +4506,7 @@ lean_inc(x_432); x_434 = lean_name_mk_string(x_432, x_433); x_435 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__4; x_436 = l_Lean_addMacroScope(x_428, x_435, x_423); -x_437 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__9; +x_437 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__9; x_438 = lean_name_mk_string(x_3, x_437); x_439 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__1; x_440 = lean_name_mk_string(x_438, x_439); @@ -4516,7 +4528,7 @@ lean_ctor_set(x_445, 0, x_421); lean_ctor_set(x_445, 1, x_444); lean_ctor_set(x_445, 2, x_436); lean_ctor_set(x_445, 3, x_443); -x_446 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__3; +x_446 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__3; lean_inc(x_421); x_447 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_447, 0, x_421); @@ -4530,11 +4542,7 @@ x_451 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_451, 0, x_421); lean_ctor_set(x_451, 1, x_450); x_452 = lean_ctor_get(x_5, 0); -lean_inc(x_452); -lean_dec(x_5); x_453 = lean_ctor_get(x_452, 0); -lean_inc(x_453); -lean_dec(x_452); x_454 = l_Lean_Name_getString_x21(x_453); x_455 = lean_box(2); x_456 = l_Lean_Syntax_mkStrLit(x_454, x_455); @@ -4589,14 +4597,14 @@ x_482 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_482, 0, x_455); lean_ctor_set(x_482, 1, x_463); lean_ctor_set(x_482, 2, x_481); -x_483 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__4; +x_483 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__4; x_484 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_484, 0, x_421); lean_ctor_set(x_484, 1, x_483); x_485 = lean_array_push(x_476, x_447); x_486 = lean_array_push(x_485, x_482); x_487 = lean_array_push(x_486, x_484); -x_488 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__2; +x_488 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__2; x_489 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_489, 0, x_455); lean_ctor_set(x_489, 1, x_488); @@ -4637,7 +4645,7 @@ lean_inc(x_502); x_504 = lean_name_mk_string(x_502, x_503); x_505 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__4; x_506 = l_Lean_addMacroScope(x_498, x_505, x_423); -x_507 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__9; +x_507 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__9; x_508 = lean_name_mk_string(x_3, x_507); x_509 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__1; x_510 = lean_name_mk_string(x_508, x_509); @@ -4659,7 +4667,7 @@ lean_ctor_set(x_515, 0, x_421); lean_ctor_set(x_515, 1, x_514); lean_ctor_set(x_515, 2, x_506); lean_ctor_set(x_515, 3, x_513); -x_516 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__3; +x_516 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__3; lean_inc(x_421); x_517 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_517, 0, x_421); @@ -4673,11 +4681,7 @@ x_521 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_521, 0, x_421); lean_ctor_set(x_521, 1, x_520); x_522 = lean_ctor_get(x_5, 0); -lean_inc(x_522); -lean_dec(x_5); x_523 = lean_ctor_get(x_522, 0); -lean_inc(x_523); -lean_dec(x_522); x_524 = l_Lean_Name_getString_x21(x_523); x_525 = lean_box(2); x_526 = l_Lean_Syntax_mkStrLit(x_524, x_525); @@ -4732,14 +4736,14 @@ x_552 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_552, 0, x_525); lean_ctor_set(x_552, 1, x_533); lean_ctor_set(x_552, 2, x_551); -x_553 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__4; +x_553 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__4; x_554 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_554, 0, x_421); lean_ctor_set(x_554, 1, x_553); x_555 = lean_array_push(x_546, x_517); x_556 = lean_array_push(x_555, x_552); x_557 = lean_array_push(x_556, x_554); -x_558 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__2; +x_558 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__2; x_559 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_559, 0, x_525); lean_ctor_set(x_559, 1, x_558); @@ -4819,7 +4823,7 @@ lean_inc(x_645); x_647 = lean_name_mk_string(x_645, x_646); x_648 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__4; x_649 = l_Lean_addMacroScope(x_641, x_648, x_636); -x_650 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__9; +x_650 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__9; x_651 = lean_name_mk_string(x_3, x_650); x_652 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__1; x_653 = lean_name_mk_string(x_651, x_652); @@ -4836,7 +4840,7 @@ lean_ctor_set(x_657, 0, x_634); lean_ctor_set(x_657, 1, x_656); lean_ctor_set(x_657, 2, x_649); lean_ctor_set(x_657, 3, x_655); -x_658 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__3; +x_658 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__3; lean_inc(x_634); x_659 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_659, 0, x_634); @@ -4850,11 +4854,7 @@ x_663 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_663, 0, x_634); lean_ctor_set(x_663, 1, x_662); x_664 = lean_ctor_get(x_5, 0); -lean_inc(x_664); -lean_dec(x_5); x_665 = lean_ctor_get(x_664, 0); -lean_inc(x_665); -lean_dec(x_664); x_666 = l_Lean_Name_getString_x21(x_665); x_667 = lean_box(2); x_668 = l_Lean_Syntax_mkStrLit(x_666, x_667); @@ -4881,7 +4881,7 @@ x_682 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_682, 0, x_667); lean_ctor_set(x_682, 1, x_681); lean_ctor_set(x_682, 2, x_680); -x_683 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__4; +x_683 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__4; lean_inc(x_634); x_684 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_684, 0, x_634); @@ -4892,7 +4892,7 @@ lean_inc(x_686); x_687 = lean_array_push(x_686, x_682); lean_inc(x_684); x_688 = lean_array_push(x_687, x_684); -x_689 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__2; +x_689 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__2; x_690 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_690, 0, x_667); lean_ctor_set(x_690, 1, x_689); @@ -4990,7 +4990,7 @@ lean_inc(x_729); x_731 = lean_name_mk_string(x_729, x_730); x_732 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__4; x_733 = l_Lean_addMacroScope(x_725, x_732, x_636); -x_734 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__9; +x_734 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__9; x_735 = lean_name_mk_string(x_3, x_734); x_736 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__1; x_737 = lean_name_mk_string(x_735, x_736); @@ -5007,7 +5007,7 @@ lean_ctor_set(x_741, 0, x_634); lean_ctor_set(x_741, 1, x_740); lean_ctor_set(x_741, 2, x_733); lean_ctor_set(x_741, 3, x_739); -x_742 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__3; +x_742 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__3; lean_inc(x_634); x_743 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_743, 0, x_634); @@ -5021,11 +5021,7 @@ x_747 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_747, 0, x_634); lean_ctor_set(x_747, 1, x_746); x_748 = lean_ctor_get(x_5, 0); -lean_inc(x_748); -lean_dec(x_5); x_749 = lean_ctor_get(x_748, 0); -lean_inc(x_749); -lean_dec(x_748); x_750 = l_Lean_Name_getString_x21(x_749); x_751 = lean_box(2); x_752 = l_Lean_Syntax_mkStrLit(x_750, x_751); @@ -5052,7 +5048,7 @@ x_766 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_766, 0, x_751); lean_ctor_set(x_766, 1, x_765); lean_ctor_set(x_766, 2, x_764); -x_767 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__4; +x_767 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__4; lean_inc(x_634); x_768 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_768, 0, x_634); @@ -5063,7 +5059,7 @@ lean_inc(x_770); x_771 = lean_array_push(x_770, x_766); lean_inc(x_768); x_772 = lean_array_push(x_771, x_768); -x_773 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__2; +x_773 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__2; x_774 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_774, 0, x_751); lean_ctor_set(x_774, 1, x_773); @@ -5194,7 +5190,7 @@ lean_inc(x_825); x_827 = lean_name_mk_string(x_825, x_826); x_828 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__4; x_829 = l_Lean_addMacroScope(x_821, x_828, x_815); -x_830 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__9; +x_830 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__9; x_831 = lean_name_mk_string(x_3, x_830); x_832 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__1; x_833 = lean_name_mk_string(x_831, x_832); @@ -5212,7 +5208,7 @@ lean_ctor_set(x_838, 0, x_813); lean_ctor_set(x_838, 1, x_837); lean_ctor_set(x_838, 2, x_829); lean_ctor_set(x_838, 3, x_836); -x_839 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__3; +x_839 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__3; lean_inc(x_813); x_840 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_840, 0, x_813); @@ -5226,11 +5222,7 @@ x_844 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_844, 0, x_813); lean_ctor_set(x_844, 1, x_843); x_845 = lean_ctor_get(x_5, 0); -lean_inc(x_845); -lean_dec(x_5); x_846 = lean_ctor_get(x_845, 0); -lean_inc(x_846); -lean_dec(x_845); x_847 = l_Lean_Name_getString_x21(x_846); x_848 = lean_box(2); x_849 = l_Lean_Syntax_mkStrLit(x_847, x_848); @@ -5257,7 +5249,7 @@ x_863 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_863, 0, x_848); lean_ctor_set(x_863, 1, x_862); lean_ctor_set(x_863, 2, x_861); -x_864 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__4; +x_864 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__4; lean_inc(x_813); x_865 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_865, 0, x_813); @@ -5268,7 +5260,7 @@ lean_inc(x_867); x_868 = lean_array_push(x_867, x_863); lean_inc(x_865); x_869 = lean_array_push(x_868, x_865); -x_870 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__2; +x_870 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__2; x_871 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_871, 0, x_848); lean_ctor_set(x_871, 1, x_870); @@ -5403,11 +5395,7 @@ lean_ctor_set(x_927, 1, x_926); lean_ctor_set(x_927, 2, x_920); lean_ctor_set(x_927, 3, x_925); x_928 = lean_ctor_get(x_5, 0); -lean_inc(x_928); -lean_dec(x_5); x_929 = lean_ctor_get(x_928, 0); -lean_inc(x_929); -lean_dec(x_928); x_930 = l_Lean_Name_getString_x21(x_929); x_931 = lean_box(2); x_932 = l_Lean_Syntax_mkStrLit(x_930, x_931); @@ -5465,11 +5453,7 @@ lean_ctor_set(x_959, 1, x_958); lean_ctor_set(x_959, 2, x_952); lean_ctor_set(x_959, 3, x_957); x_960 = lean_ctor_get(x_5, 0); -lean_inc(x_960); -lean_dec(x_5); x_961 = lean_ctor_get(x_960, 0); -lean_inc(x_961); -lean_dec(x_960); x_962 = l_Lean_Name_getString_x21(x_961); x_963 = lean_box(2); x_964 = l_Lean_Syntax_mkStrLit(x_962, x_963); @@ -5496,63 +5480,63 @@ return x_973; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -uint8_t x_14; lean_object* x_15; -x_14 = 1; -x_15 = l_Lean_Elab_Deriving_mkInstanceCmds(x_1, x_2, x_3, x_14, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_15) == 0) +uint8_t x_13; lean_object* x_14; +x_13 = 1; +x_14 = l_Lean_Elab_Deriving_mkInstanceCmds(x_1, x_2, x_3, x_13, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_14) == 0) { -uint8_t x_16; -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) +uint8_t x_15; +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_15, 0); -x_18 = lean_array_push(x_4, x_5); -x_19 = l_Array_append___rarg(x_18, x_17); -lean_ctor_set(x_15, 0, x_19); -return x_15; +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_14, 0); +x_17 = lean_array_push(x_4, x_5); +x_18 = l_Array_append___rarg(x_17, x_16); +lean_ctor_set(x_14, 0, x_18); +return x_14; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_20 = lean_ctor_get(x_15, 0); -x_21 = lean_ctor_get(x_15, 1); -lean_inc(x_21); +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_19 = lean_ctor_get(x_14, 0); +x_20 = lean_ctor_get(x_14, 1); lean_inc(x_20); -lean_dec(x_15); -x_22 = lean_array_push(x_4, x_5); -x_23 = l_Array_append___rarg(x_22, x_20); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_21); -return x_24; +lean_inc(x_19); +lean_dec(x_14); +x_21 = lean_array_push(x_4, x_5); +x_22 = l_Array_append___rarg(x_21, x_19); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_20); +return x_23; } } else { -uint8_t x_25; +uint8_t x_24; lean_dec(x_5); lean_dec(x_4); -x_25 = !lean_is_exclusive(x_15); -if (x_25 == 0) +x_24 = !lean_is_exclusive(x_14); +if (x_24 == 0) { -return x_15; +return x_14; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_15, 0); -x_27 = lean_ctor_get(x_15, 1); -lean_inc(x_27); +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_14, 0); +x_26 = lean_ctor_get(x_14, 1); lean_inc(x_26); -lean_dec(x_15); -x_28 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -return x_28; +lean_inc(x_25); +lean_dec(x_14); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; } } } @@ -5942,7 +5926,7 @@ x_43 = lean_array_get_size(x_29); x_44 = lean_usize_of_nat(x_43); lean_dec(x_43); x_45 = 0; -x_46 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_44, x_45, x_29); +x_46 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6(x_44, x_45, x_29); x_47 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__14; x_48 = l_Lean_mkSepArray(x_46, x_47); lean_dec(x_46); @@ -5986,7 +5970,7 @@ lean_ctor_set(x_71, 2, x_69); x_72 = lean_ctor_get_uint8(x_13, sizeof(void*)*2); if (x_72 == 0) { -lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_inc(x_8); x_73 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_8, x_9, x_40); x_74 = lean_ctor_get(x_73, 0); @@ -6085,166 +6069,164 @@ x_127 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_127, 0, x_51); lean_ctor_set(x_127, 1, x_126); lean_ctor_set(x_127, 2, x_125); -x_128 = lean_box(0); -x_129 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__2(x_13, x_23, x_3, x_58, x_127, x_128, x_4, x_5, x_6, x_7, x_8, x_9, x_77); +x_128 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__2(x_13, x_23, x_3, x_58, x_127, x_4, x_5, x_6, x_7, x_8, x_9, x_77); lean_dec(x_13); -return x_129; +return x_128; } else { -lean_object* x_130; lean_object* x_131; -x_130 = lean_ctor_get(x_26, 1); -lean_inc(x_130); +lean_object* x_129; lean_object* x_130; +x_129 = lean_ctor_get(x_26, 1); +lean_inc(x_129); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_131 = l_Lean_Elab_Deriving_mkLocalInstanceLetDecls(x_13, x_23, x_130, x_4, x_5, x_6, x_7, x_8, x_9, x_40); -if (lean_obj_tag(x_131) == 0) +x_130 = l_Lean_Elab_Deriving_mkLocalInstanceLetDecls(x_13, x_23, x_129, x_4, x_5, x_6, x_7, x_8, x_9, x_40); +if (lean_obj_tag(x_130) == 0) { -lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; -x_132 = lean_ctor_get(x_131, 0); +lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; +x_131 = lean_ctor_get(x_130, 0); +lean_inc(x_131); +x_132 = lean_ctor_get(x_130, 1); lean_inc(x_132); -x_133 = lean_ctor_get(x_131, 1); -lean_inc(x_133); -lean_dec(x_131); +lean_dec(x_130); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_134 = l_Lean_Elab_Deriving_mkLet(x_132, x_71, x_4, x_5, x_6, x_7, x_8, x_9, x_133); -lean_dec(x_132); -x_135 = lean_ctor_get(x_134, 0); +x_133 = l_Lean_Elab_Deriving_mkLet(x_131, x_71, x_4, x_5, x_6, x_7, x_8, x_9, x_132); +lean_dec(x_131); +x_134 = lean_ctor_get(x_133, 0); +lean_inc(x_134); +x_135 = lean_ctor_get(x_133, 1); lean_inc(x_135); -x_136 = lean_ctor_get(x_134, 1); -lean_inc(x_136); -lean_dec(x_134); +lean_dec(x_133); lean_inc(x_8); -x_137 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_8, x_9, x_136); -x_138 = lean_ctor_get(x_137, 0); -lean_inc(x_138); -x_139 = lean_ctor_get(x_137, 1); -lean_inc(x_139); -lean_dec(x_137); -x_140 = lean_st_ref_get(x_9, x_139); -x_141 = lean_ctor_get(x_140, 1); -lean_inc(x_141); -lean_dec(x_140); -x_142 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__16; -lean_inc(x_138); -x_143 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_143, 0, x_138); -lean_ctor_set(x_143, 1, x_142); -x_144 = lean_array_push(x_58, x_143); -x_145 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__17; -x_146 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_146, 0, x_51); -lean_ctor_set(x_146, 1, x_145); -lean_ctor_set(x_146, 2, x_144); -x_147 = lean_array_push(x_58, x_146); -x_148 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_148, 0, x_51); -lean_ctor_set(x_148, 1, x_52); -lean_ctor_set(x_148, 2, x_147); -x_149 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__30; +x_136 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_8, x_9, x_135); +x_137 = lean_ctor_get(x_136, 0); +lean_inc(x_137); +x_138 = lean_ctor_get(x_136, 1); lean_inc(x_138); -x_150 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_150, 0, x_138); -lean_ctor_set(x_150, 1, x_149); -x_151 = lean_array_push(x_58, x_150); -x_152 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__31; -x_153 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_153, 0, x_51); -lean_ctor_set(x_153, 1, x_152); -lean_ctor_set(x_153, 2, x_151); -x_154 = lean_array_push(x_58, x_153); -x_155 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_155, 0, x_51); -lean_ctor_set(x_155, 1, x_52); -lean_ctor_set(x_155, 2, x_154); -x_156 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__19; -x_157 = lean_array_push(x_156, x_148); +lean_dec(x_136); +x_139 = lean_st_ref_get(x_9, x_138); +x_140 = lean_ctor_get(x_139, 1); +lean_inc(x_140); +lean_dec(x_139); +x_141 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__16; +lean_inc(x_137); +x_142 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_142, 0, x_137); +lean_ctor_set(x_142, 1, x_141); +x_143 = lean_array_push(x_58, x_142); +x_144 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__17; +x_145 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_145, 0, x_51); +lean_ctor_set(x_145, 1, x_144); +lean_ctor_set(x_145, 2, x_143); +x_146 = lean_array_push(x_58, x_145); +x_147 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_147, 0, x_51); +lean_ctor_set(x_147, 1, x_52); +lean_ctor_set(x_147, 2, x_146); +x_148 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__30; +lean_inc(x_137); +x_149 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_149, 0, x_137); +lean_ctor_set(x_149, 1, x_148); +x_150 = lean_array_push(x_58, x_149); +x_151 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__31; +x_152 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_152, 0, x_51); +lean_ctor_set(x_152, 1, x_151); +lean_ctor_set(x_152, 2, x_150); +x_153 = lean_array_push(x_58, x_152); +x_154 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_154, 0, x_51); +lean_ctor_set(x_154, 1, x_52); +lean_ctor_set(x_154, 2, x_153); +x_155 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__19; +x_156 = lean_array_push(x_155, x_147); +x_157 = lean_array_push(x_156, x_64); x_158 = lean_array_push(x_157, x_64); -x_159 = lean_array_push(x_158, x_64); -x_160 = lean_array_push(x_159, x_155); -x_161 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__15; -x_162 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_162, 0, x_51); -lean_ctor_set(x_162, 1, x_161); -lean_ctor_set(x_162, 2, x_160); -x_163 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__20; -lean_inc(x_138); -x_164 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_164, 0, x_138); -lean_ctor_set(x_164, 1, x_163); -x_165 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__7; -x_166 = lean_array_push(x_165, x_19); -x_167 = lean_array_push(x_166, x_64); -x_168 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__23; -x_169 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_169, 0, x_51); -lean_ctor_set(x_169, 1, x_168); -lean_ctor_set(x_169, 2, x_167); -x_170 = lean_ctor_get(x_26, 0); -lean_inc(x_170); +x_159 = lean_array_push(x_158, x_154); +x_160 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__15; +x_161 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_161, 0, x_51); +lean_ctor_set(x_161, 1, x_160); +lean_ctor_set(x_161, 2, x_159); +x_162 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__20; +lean_inc(x_137); +x_163 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_163, 0, x_137); +lean_ctor_set(x_163, 1, x_162); +x_164 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__7; +x_165 = lean_array_push(x_164, x_19); +x_166 = lean_array_push(x_165, x_64); +x_167 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__23; +x_168 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_168, 0, x_51); +lean_ctor_set(x_168, 1, x_167); +lean_ctor_set(x_168, 2, x_166); +x_169 = lean_ctor_get(x_26, 0); +lean_inc(x_169); lean_dec(x_26); -x_171 = l_Array_append___rarg(x_49, x_170); -x_172 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_172, 0, x_51); -lean_ctor_set(x_172, 1, x_52); -lean_ctor_set(x_172, 2, x_171); -x_173 = lean_array_push(x_165, x_172); -x_174 = lean_array_push(x_173, x_64); -x_175 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__25; -x_176 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_176, 0, x_51); -lean_ctor_set(x_176, 1, x_175); -lean_ctor_set(x_176, 2, x_174); -x_177 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__28; -x_178 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_178, 0, x_138); -lean_ctor_set(x_178, 1, x_177); -x_179 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__5; -x_180 = lean_array_push(x_179, x_178); -x_181 = lean_array_push(x_180, x_135); -x_182 = lean_array_push(x_181, x_64); -x_183 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__27; -x_184 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_184, 0, x_51); -lean_ctor_set(x_184, 1, x_183); -lean_ctor_set(x_184, 2, x_182); -x_185 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__29; -x_186 = lean_array_push(x_185, x_164); -x_187 = lean_array_push(x_186, x_169); -x_188 = lean_array_push(x_187, x_176); -x_189 = lean_array_push(x_188, x_184); +x_170 = l_Array_append___rarg(x_49, x_169); +x_171 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_171, 0, x_51); +lean_ctor_set(x_171, 1, x_52); +lean_ctor_set(x_171, 2, x_170); +x_172 = lean_array_push(x_164, x_171); +x_173 = lean_array_push(x_172, x_64); +x_174 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__25; +x_175 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_175, 0, x_51); +lean_ctor_set(x_175, 1, x_174); +lean_ctor_set(x_175, 2, x_173); +x_176 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__28; +x_177 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_177, 0, x_137); +lean_ctor_set(x_177, 1, x_176); +x_178 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__5; +x_179 = lean_array_push(x_178, x_177); +x_180 = lean_array_push(x_179, x_134); +x_181 = lean_array_push(x_180, x_64); +x_182 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__27; +x_183 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_183, 0, x_51); +lean_ctor_set(x_183, 1, x_182); +lean_ctor_set(x_183, 2, x_181); +x_184 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__29; +x_185 = lean_array_push(x_184, x_163); +x_186 = lean_array_push(x_185, x_168); +x_187 = lean_array_push(x_186, x_175); +x_188 = lean_array_push(x_187, x_183); +x_189 = lean_array_push(x_188, x_64); x_190 = lean_array_push(x_189, x_64); x_191 = lean_array_push(x_190, x_64); -x_192 = lean_array_push(x_191, x_64); -x_193 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__21; -x_194 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_194, 0, x_51); -lean_ctor_set(x_194, 1, x_193); -lean_ctor_set(x_194, 2, x_192); -x_195 = lean_array_push(x_165, x_162); -x_196 = lean_array_push(x_195, x_194); -x_197 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__13; -x_198 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_198, 0, x_51); -lean_ctor_set(x_198, 1, x_197); -lean_ctor_set(x_198, 2, x_196); -x_199 = lean_box(0); -x_200 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__2(x_13, x_23, x_3, x_58, x_198, x_199, x_4, x_5, x_6, x_7, x_8, x_9, x_141); +x_192 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__21; +x_193 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_193, 0, x_51); +lean_ctor_set(x_193, 1, x_192); +lean_ctor_set(x_193, 2, x_191); +x_194 = lean_array_push(x_164, x_161); +x_195 = lean_array_push(x_194, x_193); +x_196 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__13; +x_197 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_197, 0, x_51); +lean_ctor_set(x_197, 1, x_196); +lean_ctor_set(x_197, 2, x_195); +x_198 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__2(x_13, x_23, x_3, x_58, x_197, x_4, x_5, x_6, x_7, x_8, x_9, x_140); lean_dec(x_13); -return x_200; +return x_198; } else { -uint8_t x_201; +uint8_t x_199; lean_dec(x_71); lean_dec(x_26); lean_dec(x_19); @@ -6255,30 +6237,30 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_201 = !lean_is_exclusive(x_131); -if (x_201 == 0) +x_199 = !lean_is_exclusive(x_130); +if (x_199 == 0) { -return x_131; +return x_130; } else { -lean_object* x_202; lean_object* x_203; lean_object* x_204; -x_202 = lean_ctor_get(x_131, 0); -x_203 = lean_ctor_get(x_131, 1); -lean_inc(x_203); -lean_inc(x_202); -lean_dec(x_131); -x_204 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_204, 0, x_202); -lean_ctor_set(x_204, 1, x_203); -return x_204; +lean_object* x_200; lean_object* x_201; lean_object* x_202; +x_200 = lean_ctor_get(x_130, 0); +x_201 = lean_ctor_get(x_130, 1); +lean_inc(x_201); +lean_inc(x_200); +lean_dec(x_130); +x_202 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_202, 0, x_200); +lean_ctor_set(x_202, 1, x_201); +return x_202; } } } } else { -uint8_t x_205; +uint8_t x_203; lean_dec(x_29); lean_dec(x_26); lean_dec(x_19); @@ -6289,29 +6271,29 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_205 = !lean_is_exclusive(x_33); -if (x_205 == 0) +x_203 = !lean_is_exclusive(x_33); +if (x_203 == 0) { return x_33; } else { -lean_object* x_206; lean_object* x_207; lean_object* x_208; -x_206 = lean_ctor_get(x_33, 0); -x_207 = lean_ctor_get(x_33, 1); -lean_inc(x_207); -lean_inc(x_206); +lean_object* x_204; lean_object* x_205; lean_object* x_206; +x_204 = lean_ctor_get(x_33, 0); +x_205 = lean_ctor_get(x_33, 1); +lean_inc(x_205); +lean_inc(x_204); lean_dec(x_33); -x_208 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_208, 0, x_206); -lean_ctor_set(x_208, 1, x_207); -return x_208; +x_206 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_206, 0, x_204); +lean_ctor_set(x_206, 1, x_205); +return x_206; } } } else { -uint8_t x_209; +uint8_t x_207; lean_dec(x_19); lean_dec(x_13); lean_dec(x_9); @@ -6321,29 +6303,29 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_209 = !lean_is_exclusive(x_25); -if (x_209 == 0) +x_207 = !lean_is_exclusive(x_25); +if (x_207 == 0) { return x_25; } else { -lean_object* x_210; lean_object* x_211; lean_object* x_212; -x_210 = lean_ctor_get(x_25, 0); -x_211 = lean_ctor_get(x_25, 1); -lean_inc(x_211); -lean_inc(x_210); +lean_object* x_208; lean_object* x_209; lean_object* x_210; +x_208 = lean_ctor_get(x_25, 0); +x_209 = lean_ctor_get(x_25, 1); +lean_inc(x_209); +lean_inc(x_208); lean_dec(x_25); -x_212 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_212, 0, x_210); -lean_ctor_set(x_212, 1, x_211); -return x_212; +x_210 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_210, 0, x_208); +lean_ctor_set(x_210, 1, x_209); +return x_210; } } } else { -uint8_t x_213; +uint8_t x_211; lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -6351,23 +6333,23 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_213 = !lean_is_exclusive(x_12); -if (x_213 == 0) +x_211 = !lean_is_exclusive(x_12); +if (x_211 == 0) { return x_12; } else { -lean_object* x_214; lean_object* x_215; lean_object* x_216; -x_214 = lean_ctor_get(x_12, 0); -x_215 = lean_ctor_get(x_12, 1); -lean_inc(x_215); -lean_inc(x_214); +lean_object* x_212; lean_object* x_213; lean_object* x_214; +x_212 = lean_ctor_get(x_12, 0); +x_213 = lean_ctor_get(x_12, 1); +lean_inc(x_213); +lean_inc(x_212); lean_dec(x_12); -x_216 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_216, 0, x_214); -lean_ctor_set(x_216, 1, x_215); -return x_216; +x_214 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_214, 0, x_212); +lean_ctor_set(x_214, 1, x_213); +return x_214; } } } @@ -6395,7 +6377,7 @@ static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandle { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__2; -x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__9; +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__9; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } @@ -6594,7 +6576,7 @@ lean_dec(x_29); x_31 = 0; x_32 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__2; lean_inc(x_7); -x_33 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6(x_32, x_18, x_21, x_30, x_31, x_28, x_3, x_4, x_5, x_6, x_7, x_8, x_25); +x_33 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7(x_32, x_18, x_21, x_30, x_31, x_28, x_3, x_4, x_5, x_6, x_7, x_8, x_25); x_34 = lean_ctor_get(x_33, 0); lean_inc(x_34); x_35 = lean_ctor_get(x_33, 1); @@ -6716,20 +6698,20 @@ lean_ctor_set(x_95, 0, x_37); lean_ctor_set(x_95, 1, x_93); lean_ctor_set(x_95, 2, x_92); lean_ctor_set(x_95, 3, x_94); -x_96 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__3; +x_96 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__3; lean_inc(x_37); x_97 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_97, 0, x_37); lean_ctor_set(x_97, 1, x_96); -x_98 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__13; -x_99 = l_Lean_Syntax_SepArray_ofElems(x_98, x_34); +x_98 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__14; +x_99 = l_Lean_mkSepArray(x_34, x_98); lean_dec(x_34); x_100 = l_Array_append___rarg(x_75, x_99); x_101 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_101, 0, x_49); lean_ctor_set(x_101, 1, x_53); lean_ctor_set(x_101, 2, x_100); -x_102 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__4; +x_102 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__4; x_103 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_103, 0, x_37); lean_ctor_set(x_103, 1, x_102); @@ -6737,7 +6719,7 @@ x_104 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceH x_105 = lean_array_push(x_104, x_97); x_106 = lean_array_push(x_105, x_101); x_107 = lean_array_push(x_106, x_103); -x_108 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__2; +x_108 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__2; x_109 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_109, 0, x_49); lean_ctor_set(x_109, 1, x_108); @@ -7572,7 +7554,19 @@ lean_dec(x_1); return x_19; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6(x_4, x_5, x_3); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { size_t x_14; size_t x_15; lean_object* x_16; @@ -7580,7 +7574,7 @@ x_14 = lean_unbox_usize(x_4); lean_dec(x_4); x_15 = lean_unbox_usize(x_5); lean_dec(x_5); -x_16 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6(x_1, x_2, x_3, x_14, x_15, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +x_16 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7(x_1, x_2, x_3, x_14, x_15, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_12); lean_dec(x_10); lean_dec(x_9); @@ -7601,19 +7595,19 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); +lean_dec(x_5); lean_dec(x_1); return x_15; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___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, lean_object* x_7, 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_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___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, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_14; -x_14 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_6); +lean_object* x_13; +x_13 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_3); lean_dec(x_1); -return x_14; +return x_13; } } LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__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_object* x_9, lean_object* x_10) { @@ -8043,7 +8037,7 @@ x_24 = l_Lean_Expr_isAppOf(x_21, x_23); lean_dec(x_21); if (x_24 == 0) { -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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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_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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_inc(x_13); x_25 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_13, x_14, x_15); x_26 = lean_ctor_get(x_25, 0); @@ -8085,161 +8079,159 @@ lean_ctor_set(x_43, 0, x_26); lean_ctor_set(x_43, 1, x_42); lean_ctor_set(x_43, 2, x_41); lean_ctor_set(x_43, 3, x_36); -x_44 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__3; +x_44 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__3; lean_inc(x_26); x_45 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_45, 0, x_26); lean_ctor_set(x_45, 1, x_44); lean_inc(x_6); x_46 = l_Nat_repr(x_6); -x_47 = l_Lean_numLitKind; -x_48 = lean_box(2); -x_49 = l_Lean_Syntax_mkLit(x_47, x_46, x_48); -x_50 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__4; -x_51 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_51, 0, x_26); -lean_ctor_set(x_51, 1, x_50); -x_52 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__16; -x_53 = lean_array_push(x_52, x_43); -x_54 = lean_array_push(x_53, x_45); -x_55 = lean_array_push(x_54, x_49); -x_56 = lean_array_push(x_55, x_51); -x_57 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__3___closed__13; -x_58 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_58, 0, x_48); -lean_ctor_set(x_58, 1, x_57); -lean_ctor_set(x_58, 2, x_56); -x_59 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__10; -x_60 = lean_array_push(x_59, x_58); -x_61 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__9; -x_62 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_62, 0, x_48); -lean_ctor_set(x_62, 1, x_61); -lean_ctor_set(x_62, 2, x_60); -x_63 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__7; -x_64 = lean_array_push(x_63, x_39); -x_65 = lean_array_push(x_64, x_62); -x_66 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__3; -x_67 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_67, 0, x_48); -lean_ctor_set(x_67, 1, x_66); -lean_ctor_set(x_67, 2, x_65); -x_68 = lean_array_push(x_59, x_67); -x_69 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__3___closed__2; -x_70 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_70, 0, x_48); -lean_ctor_set(x_70, 1, x_69); -lean_ctor_set(x_70, 2, x_68); -x_71 = lean_nat_add(x_6, x_18); +x_47 = lean_box(2); +x_48 = l_Lean_Syntax_mkNumLit(x_46, x_47); +x_49 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__4; +x_50 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_50, 0, x_26); +lean_ctor_set(x_50, 1, x_49); +x_51 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__16; +x_52 = lean_array_push(x_51, x_43); +x_53 = lean_array_push(x_52, x_45); +x_54 = lean_array_push(x_53, x_48); +x_55 = lean_array_push(x_54, x_50); +x_56 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__3___closed__13; +x_57 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_57, 0, x_47); +lean_ctor_set(x_57, 1, x_56); +lean_ctor_set(x_57, 2, x_55); +x_58 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__10; +x_59 = lean_array_push(x_58, x_57); +x_60 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__9; +x_61 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_61, 0, x_47); +lean_ctor_set(x_61, 1, x_60); +lean_ctor_set(x_61, 2, x_59); +x_62 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__7; +x_63 = lean_array_push(x_62, x_39); +x_64 = lean_array_push(x_63, x_61); +x_65 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__3; +x_66 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_66, 0, x_47); +lean_ctor_set(x_66, 1, x_65); +lean_ctor_set(x_66, 2, x_64); +x_67 = lean_array_push(x_58, x_66); +x_68 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__3___closed__2; +x_69 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_69, 0, x_47); +lean_ctor_set(x_69, 1, x_68); +lean_ctor_set(x_69, 2, x_67); +x_70 = lean_nat_add(x_6, x_18); lean_dec(x_6); -x_72 = lean_array_push(x_8, x_70); +x_71 = lean_array_push(x_8, x_69); x_5 = x_19; -x_6 = x_71; +x_6 = x_70; x_7 = lean_box(0); -x_8 = x_72; +x_8 = x_71; x_15 = x_31; goto _start; } else { -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_inc(x_13); -x_74 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_13, x_14, x_15); -x_75 = lean_ctor_get(x_74, 0); +x_73 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_13, x_14, x_15); +x_74 = lean_ctor_get(x_73, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_73, 1); lean_inc(x_75); -x_76 = lean_ctor_get(x_74, 1); +lean_dec(x_73); +x_76 = lean_ctor_get(x_13, 10); lean_inc(x_76); -lean_dec(x_74); -x_77 = lean_ctor_get(x_13, 10); -lean_inc(x_77); -x_78 = lean_st_ref_get(x_14, x_76); -x_79 = lean_ctor_get(x_78, 0); +x_77 = lean_st_ref_get(x_14, x_75); +x_78 = lean_ctor_get(x_77, 0); +lean_inc(x_78); +x_79 = lean_ctor_get(x_77, 1); lean_inc(x_79); -x_80 = lean_ctor_get(x_78, 1); +lean_dec(x_77); +x_80 = lean_ctor_get(x_78, 0); lean_inc(x_80); lean_dec(x_78); -x_81 = lean_ctor_get(x_79, 0); -lean_inc(x_81); -lean_dec(x_79); -x_82 = lean_environment_main_module(x_81); -x_83 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__3___closed__17; -x_84 = l_Lean_addMacroScope(x_82, x_83, x_77); -x_85 = lean_box(0); -x_86 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__3___closed__16; -lean_inc(x_75); -x_87 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_87, 0, x_75); -lean_ctor_set(x_87, 1, x_86); -lean_ctor_set(x_87, 2, x_84); -lean_ctor_set(x_87, 3, x_85); -x_88 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__3; -lean_inc(x_75); -x_89 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_89, 0, x_75); -lean_ctor_set(x_89, 1, x_88); +x_81 = lean_environment_main_module(x_80); +x_82 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__3___closed__17; +x_83 = l_Lean_addMacroScope(x_81, x_82, x_76); +x_84 = lean_box(0); +x_85 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__3___closed__16; +lean_inc(x_74); +x_86 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_86, 0, x_74); +lean_ctor_set(x_86, 1, x_85); +lean_ctor_set(x_86, 2, x_83); +lean_ctor_set(x_86, 3, x_84); +x_87 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__3; +lean_inc(x_74); +x_88 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_88, 0, x_74); +lean_ctor_set(x_88, 1, x_87); lean_inc(x_6); -x_90 = l_Nat_repr(x_6); -x_91 = l_Lean_numLitKind; -x_92 = lean_box(2); -x_93 = l_Lean_Syntax_mkLit(x_91, x_90, x_92); -x_94 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__4; -x_95 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_95, 0, x_75); -lean_ctor_set(x_95, 1, x_94); -x_96 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__16; -x_97 = lean_array_push(x_96, x_87); -x_98 = lean_array_push(x_97, x_89); -x_99 = lean_array_push(x_98, x_93); -x_100 = lean_array_push(x_99, x_95); -x_101 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__3___closed__13; -x_102 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_102, 0, x_92); -lean_ctor_set(x_102, 1, x_101); -lean_ctor_set(x_102, 2, x_100); -x_103 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__10; -x_104 = lean_array_push(x_103, x_102); -x_105 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__9; -x_106 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_106, 0, x_92); -lean_ctor_set(x_106, 1, x_105); -lean_ctor_set(x_106, 2, x_104); -x_107 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__7; +x_89 = l_Nat_repr(x_6); +x_90 = lean_box(2); +x_91 = l_Lean_Syntax_mkNumLit(x_89, x_90); +x_92 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__4; +x_93 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_93, 0, x_74); +lean_ctor_set(x_93, 1, x_92); +x_94 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__16; +x_95 = lean_array_push(x_94, x_86); +x_96 = lean_array_push(x_95, x_88); +x_97 = lean_array_push(x_96, x_91); +x_98 = lean_array_push(x_97, x_93); +x_99 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__3___closed__13; +x_100 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_100, 0, x_90); +lean_ctor_set(x_100, 1, x_99); +lean_ctor_set(x_100, 2, x_98); +x_101 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__10; +x_102 = lean_array_push(x_101, x_100); +x_103 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__9; +x_104 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_104, 0, x_90); +lean_ctor_set(x_104, 1, x_103); +lean_ctor_set(x_104, 2, x_102); +x_105 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__7; lean_inc(x_2); -x_108 = lean_array_push(x_107, x_2); -x_109 = lean_array_push(x_108, x_106); -x_110 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__3; -x_111 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_111, 0, x_92); -lean_ctor_set(x_111, 1, x_110); -lean_ctor_set(x_111, 2, x_109); -x_112 = lean_array_push(x_103, x_111); -x_113 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__3___closed__2; -x_114 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_114, 0, x_92); -lean_ctor_set(x_114, 1, x_113); -lean_ctor_set(x_114, 2, x_112); -x_115 = lean_nat_add(x_6, x_18); +x_106 = lean_array_push(x_105, x_2); +x_107 = lean_array_push(x_106, x_104); +x_108 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__3; +x_109 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_109, 0, x_90); +lean_ctor_set(x_109, 1, x_108); +lean_ctor_set(x_109, 2, x_107); +x_110 = lean_array_push(x_101, x_109); +x_111 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__3___closed__2; +x_112 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_112, 0, x_90); +lean_ctor_set(x_112, 1, x_111); +lean_ctor_set(x_112, 2, x_110); +x_113 = lean_nat_add(x_6, x_18); lean_dec(x_6); -x_116 = lean_array_push(x_8, x_114); +x_114 = lean_array_push(x_8, x_112); x_5 = x_19; -x_6 = x_115; +x_6 = x_113; x_7 = lean_box(0); -x_8 = x_116; -x_15 = x_80; +x_8 = x_114; +x_15 = x_79; goto _start; } } else { -lean_object* x_118; +lean_object* x_116; lean_dec(x_13); lean_dec(x_6); lean_dec(x_5); lean_dec(x_2); -x_118 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_118, 0, x_8); -lean_ctor_set(x_118, 1, x_15); -return x_118; +x_116 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_116, 0, x_8); +lean_ctor_set(x_116, 1, x_15); +return x_116; } } } @@ -8389,64 +8381,63 @@ x_1 = lean_mk_string_from_bytes("`", 1); return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__5(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { _start: { -uint8_t x_6; -x_6 = lean_usize_dec_lt(x_4, x_3); -if (x_6 == 0) +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) { +lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -return x_5; +return x_6; } else { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; size_t x_11; size_t x_12; -x_7 = lean_array_uget(x_5, x_4); -x_8 = lean_unsigned_to_nat(0u); -x_9 = lean_array_uset(x_5, x_4, x_8); -lean_inc(x_7); -lean_inc(x_2); -x_10 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_2, x_7); -x_11 = 1; -x_12 = lean_usize_add(x_4, x_11); -if (lean_obj_tag(x_10) == 0) -{ -lean_object* x_13; lean_object* x_14; -x_13 = l___private_Init_Meta_0__Lean_quoteNameMk(x_7); -x_14 = lean_array_uset(x_9, x_4, x_13); -x_4 = x_12; -x_5 = x_14; +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; +x_8 = lean_array_uget(x_6, x_5); +x_9 = lean_unsigned_to_nat(0u); +x_10 = lean_array_uset(x_6, x_5, x_9); +lean_inc(x_8); +lean_inc(x_3); +x_11 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_3, x_8); +x_12 = 1; +x_13 = lean_usize_add(x_5, x_12); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = l_Lean_quoteNameMk(x_8); +x_15 = lean_array_uset(x_10, x_5, x_14); +x_5 = x_13; +x_6 = x_15; goto _start; } else { -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; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -lean_dec(x_7); -x_16 = lean_ctor_get(x_10, 0); -lean_inc(x_16); -lean_dec(x_10); -x_17 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__5___closed__1; -lean_inc(x_1); -x_18 = lean_name_mk_string(x_1, x_17); -x_19 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__5___closed__2; -x_20 = l_String_intercalate(x_19, x_16); -x_21 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__5___closed__3; -x_22 = lean_string_append(x_21, x_20); -lean_dec(x_20); -x_23 = l_Lean_nameLitKind; +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; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_dec(x_8); +x_17 = lean_ctor_get(x_11, 0); +lean_inc(x_17); +lean_dec(x_11); +x_18 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__5___closed__1; +lean_inc(x_2); +x_19 = lean_name_mk_string(x_2, x_18); +x_20 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__5___closed__2; +x_21 = l_String_intercalate(x_20, x_17); +x_22 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__5___closed__3; +x_23 = lean_string_append(x_22, x_21); +lean_dec(x_21); x_24 = lean_box(2); -x_25 = l_Lean_Syntax_mkLit(x_23, x_22, x_24); +x_25 = l_Lean_Syntax_mkNameLit(x_23, x_24); x_26 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__10; x_27 = lean_array_push(x_26, x_25); x_28 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_28, 0, x_24); -lean_ctor_set(x_28, 1, x_18); +lean_ctor_set(x_28, 1, x_19); lean_ctor_set(x_28, 2, x_27); -x_29 = lean_array_uset(x_9, x_4, x_28); -x_4 = x_12; -x_5 = x_29; +x_29 = lean_array_uset(x_10, x_5, x_28); +x_5 = x_13; +x_6 = x_29; goto _start; } } @@ -8823,7 +8814,7 @@ x_19 = lean_st_ref_get(x_13, x_17); x_20 = !lean_is_exclusive(x_19); if (x_20 == 0) { -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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; size_t x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; +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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; size_t x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; x_21 = lean_ctor_get(x_19, 0); x_22 = lean_ctor_get(x_21, 0); lean_inc(x_22); @@ -8858,484 +8849,480 @@ lean_ctor_set(x_35, 0, x_16); lean_ctor_set(x_35, 1, x_34); lean_ctor_set(x_35, 2, x_33); lean_ctor_set(x_35, 3, x_28); -lean_inc(x_1); x_36 = l_Lean_Name_getString_x21(x_1); x_37 = lean_box(2); x_38 = l_Lean_Syntax_mkStrLit(x_36, x_37); lean_dec(x_36); lean_inc(x_2); x_39 = l_Nat_repr(x_2); -x_40 = l_Lean_numLitKind; -x_41 = l_Lean_Syntax_mkLit(x_40, x_39, x_37); -x_42 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__16; -x_43 = lean_array_push(x_42, x_35); -x_44 = lean_array_push(x_43, x_38); -x_45 = lean_array_push(x_44, x_41); -x_46 = lean_array_push(x_45, x_7); -x_47 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__9; -x_48 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_48, 0, x_37); -lean_ctor_set(x_48, 1, x_47); -lean_ctor_set(x_48, 2, x_46); -x_49 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__7; -x_50 = lean_array_push(x_49, x_31); -x_51 = lean_array_push(x_50, x_48); -x_52 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__3; +x_40 = l_Lean_Syntax_mkNumLit(x_39, x_37); +x_41 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__16; +x_42 = lean_array_push(x_41, x_35); +x_43 = lean_array_push(x_42, x_38); +x_44 = lean_array_push(x_43, x_40); +x_45 = lean_array_push(x_44, x_7); +x_46 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__9; +x_47 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_47, 0, x_37); +lean_ctor_set(x_47, 1, x_46); +lean_ctor_set(x_47, 2, x_45); +x_48 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__7; +x_49 = lean_array_push(x_48, x_31); +x_50 = lean_array_push(x_49, x_47); +x_51 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__3; +x_52 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_52, 0, x_37); +lean_ctor_set(x_52, 1, x_51); +lean_ctor_set(x_52, 2, x_50); +lean_inc(x_3); x_53 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_53, 0, x_37); -lean_ctor_set(x_53, 1, x_52); -lean_ctor_set(x_53, 2, x_51); -lean_inc(x_3); -x_54 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_54, 0, x_37); -lean_ctor_set(x_54, 1, x_47); -lean_ctor_set(x_54, 2, x_3); -x_55 = lean_array_push(x_49, x_53); -lean_inc(x_54); -x_56 = lean_array_push(x_55, x_54); -x_57 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_57, 0, x_37); -lean_ctor_set(x_57, 1, x_47); -lean_ctor_set(x_57, 2, x_56); -x_58 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__4; +lean_ctor_set(x_53, 1, x_46); +lean_ctor_set(x_53, 2, x_3); +x_54 = lean_array_push(x_48, x_52); +lean_inc(x_53); +x_55 = lean_array_push(x_54, x_53); +x_56 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_56, 0, x_37); +lean_ctor_set(x_56, 1, x_46); +lean_ctor_set(x_56, 2, x_55); +x_57 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__4; lean_inc(x_16); -x_59 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_59, 0, x_16); -lean_ctor_set(x_59, 1, x_58); -x_60 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__5; -x_61 = lean_array_push(x_60, x_25); -lean_inc(x_61); -x_62 = lean_array_push(x_61, x_57); -lean_inc(x_59); -x_63 = lean_array_push(x_62, x_59); -x_64 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__3; -x_65 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_65, 0, x_37); -lean_ctor_set(x_65, 1, x_64); -lean_ctor_set(x_65, 2, x_63); -x_66 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__5___closed__2; +x_58 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_58, 0, x_16); +lean_ctor_set(x_58, 1, x_57); +x_59 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__5; +x_60 = lean_array_push(x_59, x_25); +lean_inc(x_60); +x_61 = lean_array_push(x_60, x_56); +lean_inc(x_58); +x_62 = lean_array_push(x_61, x_58); +x_63 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__3; +x_64 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_64, 0, x_37); +lean_ctor_set(x_64, 1, x_63); +lean_ctor_set(x_64, 2, x_62); +x_65 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__5___closed__2; lean_inc(x_16); -x_67 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_67, 0, x_16); -lean_ctor_set(x_67, 1, x_66); -x_68 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__19; +x_66 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_66, 0, x_16); +lean_ctor_set(x_66, 1, x_65); +x_67 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__19; lean_inc(x_18); lean_inc(x_23); -x_69 = l_Lean_addMacroScope(x_23, x_68, x_18); -x_70 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__18; -x_71 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__24; +x_68 = l_Lean_addMacroScope(x_23, x_67, x_18); +x_69 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__18; +x_70 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__24; lean_inc(x_16); -x_72 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_72, 0, x_16); -lean_ctor_set(x_72, 1, x_70); -lean_ctor_set(x_72, 2, x_69); -lean_ctor_set(x_72, 3, x_71); -x_73 = lean_array_push(x_60, x_65); -x_74 = lean_array_push(x_73, x_67); -x_75 = lean_array_push(x_74, x_72); -x_76 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__2; -x_77 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_77, 0, x_37); -lean_ctor_set(x_77, 1, x_76); -lean_ctor_set(x_77, 2, x_75); -x_78 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__25; +x_71 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_71, 0, x_16); +lean_ctor_set(x_71, 1, x_69); +lean_ctor_set(x_71, 2, x_68); +lean_ctor_set(x_71, 3, x_70); +x_72 = lean_array_push(x_59, x_64); +x_73 = lean_array_push(x_72, x_66); +x_74 = lean_array_push(x_73, x_71); +x_75 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__2; +x_76 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_76, 0, x_37); +lean_ctor_set(x_76, 1, x_75); +lean_ctor_set(x_76, 2, x_74); +x_77 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__25; lean_inc(x_16); -x_79 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_79, 0, x_16); -lean_ctor_set(x_79, 1, x_78); -x_80 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__3___closed__17; -x_81 = l_Lean_addMacroScope(x_23, x_80, x_18); -x_82 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__3___closed__16; +x_78 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_78, 0, x_16); +lean_ctor_set(x_78, 1, x_77); +x_79 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__3___closed__17; +x_80 = l_Lean_addMacroScope(x_23, x_79, x_18); +x_81 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__3___closed__16; lean_inc(x_16); -x_83 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_83, 0, x_16); -lean_ctor_set(x_83, 1, x_82); -lean_ctor_set(x_83, 2, x_81); -lean_ctor_set(x_83, 3, x_28); -x_84 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__10; -x_85 = lean_array_push(x_84, x_83); -x_86 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_86, 0, x_37); -lean_ctor_set(x_86, 1, x_47); -lean_ctor_set(x_86, 2, x_85); -x_87 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__15; +x_82 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_82, 0, x_16); +lean_ctor_set(x_82, 1, x_81); +lean_ctor_set(x_82, 2, x_80); +lean_ctor_set(x_82, 3, x_28); +x_83 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__10; +x_84 = lean_array_push(x_83, x_82); +x_85 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_85, 0, x_37); +lean_ctor_set(x_85, 1, x_46); +lean_ctor_set(x_85, 2, x_84); +x_86 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__15; lean_inc(x_16); -x_88 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_88, 0, x_16); -lean_ctor_set(x_88, 1, x_87); -x_89 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__29; +x_87 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_87, 0, x_16); +lean_ctor_set(x_87, 1, x_86); +x_88 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__29; lean_inc(x_16); -x_90 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_90, 0, x_16); -lean_ctor_set(x_90, 1, x_89); -x_91 = l_Array_zip___rarg(x_4, x_5); -x_92 = lean_array_get_size(x_91); -x_93 = lean_usize_of_nat(x_92); -lean_dec(x_92); -x_94 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__6; -lean_inc(x_54); +x_89 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_89, 0, x_16); +lean_ctor_set(x_89, 1, x_88); +x_90 = l_Array_zip___rarg(x_4, x_5); +x_91 = lean_array_get_size(x_90); +x_92 = lean_usize_of_nat(x_91); +lean_dec(x_91); +x_93 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__6; +lean_inc(x_53); lean_inc(x_16); -x_95 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__4(x_16, x_94, x_42, x_49, x_54, x_60, x_93, x_6, x_91); +x_94 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__4(x_16, x_93, x_41, x_48, x_53, x_59, x_92, x_6, x_90); lean_inc(x_3); -x_96 = l_Array_append___rarg(x_3, x_95); -x_97 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__36; -x_98 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_98, 0, x_16); -lean_ctor_set(x_98, 1, x_97); -x_99 = lean_mk_syntax_ident(x_1); -x_100 = l_Array_append___rarg(x_3, x_4); -x_101 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_101, 0, x_37); -lean_ctor_set(x_101, 1, x_47); -lean_ctor_set(x_101, 2, x_100); -x_102 = lean_array_push(x_49, x_99); -x_103 = lean_array_push(x_102, x_101); -x_104 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_104, 0, x_37); -lean_ctor_set(x_104, 1, x_52); -lean_ctor_set(x_104, 2, x_103); -x_105 = lean_array_push(x_84, x_104); -x_106 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_106, 0, x_37); -lean_ctor_set(x_106, 1, x_47); -lean_ctor_set(x_106, 2, x_105); -x_107 = lean_array_push(x_49, x_98); -x_108 = lean_array_push(x_107, x_106); -x_109 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__35; -x_110 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_110, 0, x_37); -lean_ctor_set(x_110, 1, x_109); -lean_ctor_set(x_110, 2, x_108); -x_111 = lean_array_push(x_49, x_110); -lean_inc(x_54); -x_112 = lean_array_push(x_111, x_54); -x_113 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__33; -x_114 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_114, 0, x_37); -lean_ctor_set(x_114, 1, x_113); -lean_ctor_set(x_114, 2, x_112); -x_115 = lean_array_push(x_96, x_114); -x_116 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_116, 0, x_37); -lean_ctor_set(x_116, 1, x_47); -lean_ctor_set(x_116, 2, x_115); -x_117 = lean_array_push(x_84, x_116); -x_118 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__32; -x_119 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_119, 0, x_37); -lean_ctor_set(x_119, 1, x_118); -lean_ctor_set(x_119, 2, x_117); -x_120 = lean_array_push(x_49, x_90); -x_121 = lean_array_push(x_120, x_119); -x_122 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__30; -x_123 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_123, 0, x_37); -lean_ctor_set(x_123, 1, x_122); -lean_ctor_set(x_123, 2, x_121); -x_124 = lean_array_push(x_60, x_86); -x_125 = lean_array_push(x_124, x_88); -x_126 = lean_array_push(x_125, x_123); -x_127 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__28; -x_128 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_128, 0, x_37); -lean_ctor_set(x_128, 1, x_127); -lean_ctor_set(x_128, 2, x_126); -x_129 = lean_array_push(x_49, x_79); -x_130 = lean_array_push(x_129, x_128); -x_131 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__26; -x_132 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_132, 0, x_37); -lean_ctor_set(x_132, 1, x_131); -lean_ctor_set(x_132, 2, x_130); -x_133 = lean_array_push(x_49, x_132); -x_134 = lean_array_push(x_133, x_54); -x_135 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_135, 0, x_37); -lean_ctor_set(x_135, 1, x_47); -lean_ctor_set(x_135, 2, x_134); -x_136 = lean_array_push(x_61, x_135); -x_137 = lean_array_push(x_136, x_59); -x_138 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_138, 0, x_37); -lean_ctor_set(x_138, 1, x_64); -lean_ctor_set(x_138, 2, x_137); -x_139 = lean_array_push(x_84, x_138); -x_140 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_140, 0, x_37); -lean_ctor_set(x_140, 1, x_47); -lean_ctor_set(x_140, 2, x_139); -x_141 = lean_array_push(x_49, x_77); -x_142 = lean_array_push(x_141, x_140); -x_143 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_143, 0, x_37); -lean_ctor_set(x_143, 1, x_52); -lean_ctor_set(x_143, 2, x_142); -x_144 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_144, 0, x_143); -lean_ctor_set(x_144, 1, x_2); -lean_ctor_set(x_19, 0, x_144); +x_95 = l_Array_append___rarg(x_3, x_94); +x_96 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__36; +x_97 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_97, 0, x_16); +lean_ctor_set(x_97, 1, x_96); +x_98 = lean_mk_syntax_ident(x_1); +x_99 = l_Array_append___rarg(x_3, x_4); +x_100 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_100, 0, x_37); +lean_ctor_set(x_100, 1, x_46); +lean_ctor_set(x_100, 2, x_99); +x_101 = lean_array_push(x_48, x_98); +x_102 = lean_array_push(x_101, x_100); +x_103 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_103, 0, x_37); +lean_ctor_set(x_103, 1, x_51); +lean_ctor_set(x_103, 2, x_102); +x_104 = lean_array_push(x_83, x_103); +x_105 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_105, 0, x_37); +lean_ctor_set(x_105, 1, x_46); +lean_ctor_set(x_105, 2, x_104); +x_106 = lean_array_push(x_48, x_97); +x_107 = lean_array_push(x_106, x_105); +x_108 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__35; +x_109 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_109, 0, x_37); +lean_ctor_set(x_109, 1, x_108); +lean_ctor_set(x_109, 2, x_107); +x_110 = lean_array_push(x_48, x_109); +lean_inc(x_53); +x_111 = lean_array_push(x_110, x_53); +x_112 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__33; +x_113 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_113, 0, x_37); +lean_ctor_set(x_113, 1, x_112); +lean_ctor_set(x_113, 2, x_111); +x_114 = lean_array_push(x_95, x_113); +x_115 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_115, 0, x_37); +lean_ctor_set(x_115, 1, x_46); +lean_ctor_set(x_115, 2, x_114); +x_116 = lean_array_push(x_83, x_115); +x_117 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__32; +x_118 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_118, 0, x_37); +lean_ctor_set(x_118, 1, x_117); +lean_ctor_set(x_118, 2, x_116); +x_119 = lean_array_push(x_48, x_89); +x_120 = lean_array_push(x_119, x_118); +x_121 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__30; +x_122 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_122, 0, x_37); +lean_ctor_set(x_122, 1, x_121); +lean_ctor_set(x_122, 2, x_120); +x_123 = lean_array_push(x_59, x_85); +x_124 = lean_array_push(x_123, x_87); +x_125 = lean_array_push(x_124, x_122); +x_126 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__28; +x_127 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_127, 0, x_37); +lean_ctor_set(x_127, 1, x_126); +lean_ctor_set(x_127, 2, x_125); +x_128 = lean_array_push(x_48, x_78); +x_129 = lean_array_push(x_128, x_127); +x_130 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__26; +x_131 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_131, 0, x_37); +lean_ctor_set(x_131, 1, x_130); +lean_ctor_set(x_131, 2, x_129); +x_132 = lean_array_push(x_48, x_131); +x_133 = lean_array_push(x_132, x_53); +x_134 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_134, 0, x_37); +lean_ctor_set(x_134, 1, x_46); +lean_ctor_set(x_134, 2, x_133); +x_135 = lean_array_push(x_60, x_134); +x_136 = lean_array_push(x_135, x_58); +x_137 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_137, 0, x_37); +lean_ctor_set(x_137, 1, x_63); +lean_ctor_set(x_137, 2, x_136); +x_138 = lean_array_push(x_83, x_137); +x_139 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_139, 0, x_37); +lean_ctor_set(x_139, 1, x_46); +lean_ctor_set(x_139, 2, x_138); +x_140 = lean_array_push(x_48, x_76); +x_141 = lean_array_push(x_140, x_139); +x_142 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_142, 0, x_37); +lean_ctor_set(x_142, 1, x_51); +lean_ctor_set(x_142, 2, x_141); +x_143 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_143, 0, x_142); +lean_ctor_set(x_143, 1, x_2); +lean_ctor_set(x_19, 0, x_143); return x_19; } else { -lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; size_t x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; -x_145 = lean_ctor_get(x_19, 0); -x_146 = lean_ctor_get(x_19, 1); -lean_inc(x_146); +lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; size_t x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; +x_144 = lean_ctor_get(x_19, 0); +x_145 = lean_ctor_get(x_19, 1); lean_inc(x_145); +lean_inc(x_144); lean_dec(x_19); -x_147 = lean_ctor_get(x_145, 0); -lean_inc(x_147); -lean_dec(x_145); -x_148 = lean_environment_main_module(x_147); -x_149 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__2; +x_146 = lean_ctor_get(x_144, 0); +lean_inc(x_146); +lean_dec(x_144); +x_147 = lean_environment_main_module(x_146); +x_148 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__2; lean_inc(x_16); -x_150 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_150, 0, x_16); -lean_ctor_set(x_150, 1, x_149); -x_151 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__8; +x_149 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_149, 0, x_16); +lean_ctor_set(x_149, 1, x_148); +x_150 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__8; lean_inc(x_18); -lean_inc(x_148); -x_152 = l_Lean_addMacroScope(x_148, x_151, x_18); -x_153 = lean_box(0); -x_154 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__6; -x_155 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__11; +lean_inc(x_147); +x_151 = l_Lean_addMacroScope(x_147, x_150, x_18); +x_152 = lean_box(0); +x_153 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__6; +x_154 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__11; lean_inc(x_16); -x_156 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_156, 0, x_16); -lean_ctor_set(x_156, 1, x_154); -lean_ctor_set(x_156, 2, x_152); -lean_ctor_set(x_156, 3, x_155); -x_157 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__15; +x_155 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_155, 0, x_16); +lean_ctor_set(x_155, 1, x_153); +lean_ctor_set(x_155, 2, x_151); +lean_ctor_set(x_155, 3, x_154); +x_156 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__15; lean_inc(x_18); -lean_inc(x_148); -x_158 = l_Lean_addMacroScope(x_148, x_157, x_18); -x_159 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__14; +lean_inc(x_147); +x_157 = l_Lean_addMacroScope(x_147, x_156, x_18); +x_158 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__14; lean_inc(x_16); -x_160 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_160, 0, x_16); -lean_ctor_set(x_160, 1, x_159); -lean_ctor_set(x_160, 2, x_158); -lean_ctor_set(x_160, 3, x_153); -lean_inc(x_1); -x_161 = l_Lean_Name_getString_x21(x_1); -x_162 = lean_box(2); -x_163 = l_Lean_Syntax_mkStrLit(x_161, x_162); -lean_dec(x_161); +x_159 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_159, 0, x_16); +lean_ctor_set(x_159, 1, x_158); +lean_ctor_set(x_159, 2, x_157); +lean_ctor_set(x_159, 3, x_152); +x_160 = l_Lean_Name_getString_x21(x_1); +x_161 = lean_box(2); +x_162 = l_Lean_Syntax_mkStrLit(x_160, x_161); +lean_dec(x_160); lean_inc(x_2); -x_164 = l_Nat_repr(x_2); -x_165 = l_Lean_numLitKind; -x_166 = l_Lean_Syntax_mkLit(x_165, x_164, x_162); -x_167 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__16; -x_168 = lean_array_push(x_167, x_160); -x_169 = lean_array_push(x_168, x_163); -x_170 = lean_array_push(x_169, x_166); -x_171 = lean_array_push(x_170, x_7); -x_172 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__9; -x_173 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_173, 0, x_162); -lean_ctor_set(x_173, 1, x_172); -lean_ctor_set(x_173, 2, x_171); -x_174 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__7; -x_175 = lean_array_push(x_174, x_156); -x_176 = lean_array_push(x_175, x_173); -x_177 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__3; -x_178 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_178, 0, x_162); -lean_ctor_set(x_178, 1, x_177); -lean_ctor_set(x_178, 2, x_176); +x_163 = l_Nat_repr(x_2); +x_164 = l_Lean_Syntax_mkNumLit(x_163, x_161); +x_165 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__16; +x_166 = lean_array_push(x_165, x_159); +x_167 = lean_array_push(x_166, x_162); +x_168 = lean_array_push(x_167, x_164); +x_169 = lean_array_push(x_168, x_7); +x_170 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__9; +x_171 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_171, 0, x_161); +lean_ctor_set(x_171, 1, x_170); +lean_ctor_set(x_171, 2, x_169); +x_172 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__7; +x_173 = lean_array_push(x_172, x_155); +x_174 = lean_array_push(x_173, x_171); +x_175 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__3; +x_176 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_176, 0, x_161); +lean_ctor_set(x_176, 1, x_175); +lean_ctor_set(x_176, 2, x_174); lean_inc(x_3); -x_179 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_179, 0, x_162); -lean_ctor_set(x_179, 1, x_172); -lean_ctor_set(x_179, 2, x_3); -x_180 = lean_array_push(x_174, x_178); -lean_inc(x_179); -x_181 = lean_array_push(x_180, x_179); -x_182 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_182, 0, x_162); -lean_ctor_set(x_182, 1, x_172); -lean_ctor_set(x_182, 2, x_181); -x_183 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__4; +x_177 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_177, 0, x_161); +lean_ctor_set(x_177, 1, x_170); +lean_ctor_set(x_177, 2, x_3); +x_178 = lean_array_push(x_172, x_176); +lean_inc(x_177); +x_179 = lean_array_push(x_178, x_177); +x_180 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_180, 0, x_161); +lean_ctor_set(x_180, 1, x_170); +lean_ctor_set(x_180, 2, x_179); +x_181 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__4; lean_inc(x_16); -x_184 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_184, 0, x_16); -lean_ctor_set(x_184, 1, x_183); -x_185 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__5; -x_186 = lean_array_push(x_185, x_150); -lean_inc(x_186); -x_187 = lean_array_push(x_186, x_182); +x_182 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_182, 0, x_16); +lean_ctor_set(x_182, 1, x_181); +x_183 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__5; +x_184 = lean_array_push(x_183, x_149); lean_inc(x_184); -x_188 = lean_array_push(x_187, x_184); -x_189 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__3; -x_190 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_190, 0, x_162); -lean_ctor_set(x_190, 1, x_189); -lean_ctor_set(x_190, 2, x_188); -x_191 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__5___closed__2; +x_185 = lean_array_push(x_184, x_180); +lean_inc(x_182); +x_186 = lean_array_push(x_185, x_182); +x_187 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__3; +x_188 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_188, 0, x_161); +lean_ctor_set(x_188, 1, x_187); +lean_ctor_set(x_188, 2, x_186); +x_189 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__5___closed__2; lean_inc(x_16); -x_192 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_192, 0, x_16); -lean_ctor_set(x_192, 1, x_191); -x_193 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__19; +x_190 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_190, 0, x_16); +lean_ctor_set(x_190, 1, x_189); +x_191 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__19; lean_inc(x_18); -lean_inc(x_148); -x_194 = l_Lean_addMacroScope(x_148, x_193, x_18); -x_195 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__18; -x_196 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__24; +lean_inc(x_147); +x_192 = l_Lean_addMacroScope(x_147, x_191, x_18); +x_193 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__18; +x_194 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__24; lean_inc(x_16); -x_197 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_197, 0, x_16); -lean_ctor_set(x_197, 1, x_195); -lean_ctor_set(x_197, 2, x_194); -lean_ctor_set(x_197, 3, x_196); -x_198 = lean_array_push(x_185, x_190); -x_199 = lean_array_push(x_198, x_192); -x_200 = lean_array_push(x_199, x_197); -x_201 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__2; -x_202 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_202, 0, x_162); +x_195 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_195, 0, x_16); +lean_ctor_set(x_195, 1, x_193); +lean_ctor_set(x_195, 2, x_192); +lean_ctor_set(x_195, 3, x_194); +x_196 = lean_array_push(x_183, x_188); +x_197 = lean_array_push(x_196, x_190); +x_198 = lean_array_push(x_197, x_195); +x_199 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__2; +x_200 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_200, 0, x_161); +lean_ctor_set(x_200, 1, x_199); +lean_ctor_set(x_200, 2, x_198); +x_201 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__25; +lean_inc(x_16); +x_202 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_202, 0, x_16); lean_ctor_set(x_202, 1, x_201); -lean_ctor_set(x_202, 2, x_200); -x_203 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__25; +x_203 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__3___closed__17; +x_204 = l_Lean_addMacroScope(x_147, x_203, x_18); +x_205 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__3___closed__16; lean_inc(x_16); -x_204 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_204, 0, x_16); -lean_ctor_set(x_204, 1, x_203); -x_205 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__3___closed__17; -x_206 = l_Lean_addMacroScope(x_148, x_205, x_18); -x_207 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__3___closed__16; +x_206 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_206, 0, x_16); +lean_ctor_set(x_206, 1, x_205); +lean_ctor_set(x_206, 2, x_204); +lean_ctor_set(x_206, 3, x_152); +x_207 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__10; +x_208 = lean_array_push(x_207, x_206); +x_209 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_209, 0, x_161); +lean_ctor_set(x_209, 1, x_170); +lean_ctor_set(x_209, 2, x_208); +x_210 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__15; lean_inc(x_16); -x_208 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_208, 0, x_16); -lean_ctor_set(x_208, 1, x_207); -lean_ctor_set(x_208, 2, x_206); -lean_ctor_set(x_208, 3, x_153); -x_209 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__10; -x_210 = lean_array_push(x_209, x_208); -x_211 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_211, 0, x_162); -lean_ctor_set(x_211, 1, x_172); -lean_ctor_set(x_211, 2, x_210); -x_212 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__15; +x_211 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_211, 0, x_16); +lean_ctor_set(x_211, 1, x_210); +x_212 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__29; lean_inc(x_16); x_213 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_213, 0, x_16); lean_ctor_set(x_213, 1, x_212); -x_214 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__29; -lean_inc(x_16); -x_215 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_215, 0, x_16); -lean_ctor_set(x_215, 1, x_214); -x_216 = l_Array_zip___rarg(x_4, x_5); -x_217 = lean_array_get_size(x_216); -x_218 = lean_usize_of_nat(x_217); -lean_dec(x_217); -x_219 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__6; -lean_inc(x_179); +x_214 = l_Array_zip___rarg(x_4, x_5); +x_215 = lean_array_get_size(x_214); +x_216 = lean_usize_of_nat(x_215); +lean_dec(x_215); +x_217 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__6; +lean_inc(x_177); lean_inc(x_16); -x_220 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__4(x_16, x_219, x_167, x_174, x_179, x_185, x_218, x_6, x_216); +x_218 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__4(x_16, x_217, x_165, x_172, x_177, x_183, x_216, x_6, x_214); lean_inc(x_3); -x_221 = l_Array_append___rarg(x_3, x_220); -x_222 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__36; -x_223 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_223, 0, x_16); -lean_ctor_set(x_223, 1, x_222); -x_224 = lean_mk_syntax_ident(x_1); -x_225 = l_Array_append___rarg(x_3, x_4); -x_226 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_226, 0, x_162); -lean_ctor_set(x_226, 1, x_172); -lean_ctor_set(x_226, 2, x_225); -x_227 = lean_array_push(x_174, x_224); -x_228 = lean_array_push(x_227, x_226); +x_219 = l_Array_append___rarg(x_3, x_218); +x_220 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__36; +x_221 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_221, 0, x_16); +lean_ctor_set(x_221, 1, x_220); +x_222 = lean_mk_syntax_ident(x_1); +x_223 = l_Array_append___rarg(x_3, x_4); +x_224 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_224, 0, x_161); +lean_ctor_set(x_224, 1, x_170); +lean_ctor_set(x_224, 2, x_223); +x_225 = lean_array_push(x_172, x_222); +x_226 = lean_array_push(x_225, x_224); +x_227 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_227, 0, x_161); +lean_ctor_set(x_227, 1, x_175); +lean_ctor_set(x_227, 2, x_226); +x_228 = lean_array_push(x_207, x_227); x_229 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_229, 0, x_162); -lean_ctor_set(x_229, 1, x_177); +lean_ctor_set(x_229, 0, x_161); +lean_ctor_set(x_229, 1, x_170); lean_ctor_set(x_229, 2, x_228); -x_230 = lean_array_push(x_209, x_229); -x_231 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_231, 0, x_162); -lean_ctor_set(x_231, 1, x_172); -lean_ctor_set(x_231, 2, x_230); -x_232 = lean_array_push(x_174, x_223); -x_233 = lean_array_push(x_232, x_231); -x_234 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__35; -x_235 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_235, 0, x_162); -lean_ctor_set(x_235, 1, x_234); -lean_ctor_set(x_235, 2, x_233); -x_236 = lean_array_push(x_174, x_235); -lean_inc(x_179); -x_237 = lean_array_push(x_236, x_179); -x_238 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__33; +x_230 = lean_array_push(x_172, x_221); +x_231 = lean_array_push(x_230, x_229); +x_232 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__35; +x_233 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_233, 0, x_161); +lean_ctor_set(x_233, 1, x_232); +lean_ctor_set(x_233, 2, x_231); +x_234 = lean_array_push(x_172, x_233); +lean_inc(x_177); +x_235 = lean_array_push(x_234, x_177); +x_236 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__33; +x_237 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_237, 0, x_161); +lean_ctor_set(x_237, 1, x_236); +lean_ctor_set(x_237, 2, x_235); +x_238 = lean_array_push(x_219, x_237); x_239 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_239, 0, x_162); -lean_ctor_set(x_239, 1, x_238); -lean_ctor_set(x_239, 2, x_237); -x_240 = lean_array_push(x_221, x_239); -x_241 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_241, 0, x_162); -lean_ctor_set(x_241, 1, x_172); -lean_ctor_set(x_241, 2, x_240); -x_242 = lean_array_push(x_209, x_241); -x_243 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__32; -x_244 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_244, 0, x_162); -lean_ctor_set(x_244, 1, x_243); -lean_ctor_set(x_244, 2, x_242); -x_245 = lean_array_push(x_174, x_215); -x_246 = lean_array_push(x_245, x_244); -x_247 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__30; -x_248 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_248, 0, x_162); -lean_ctor_set(x_248, 1, x_247); -lean_ctor_set(x_248, 2, x_246); -x_249 = lean_array_push(x_185, x_211); -x_250 = lean_array_push(x_249, x_213); -x_251 = lean_array_push(x_250, x_248); -x_252 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__28; -x_253 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_253, 0, x_162); -lean_ctor_set(x_253, 1, x_252); -lean_ctor_set(x_253, 2, x_251); -x_254 = lean_array_push(x_174, x_204); -x_255 = lean_array_push(x_254, x_253); -x_256 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__26; -x_257 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_257, 0, x_162); -lean_ctor_set(x_257, 1, x_256); -lean_ctor_set(x_257, 2, x_255); -x_258 = lean_array_push(x_174, x_257); -x_259 = lean_array_push(x_258, x_179); -x_260 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_260, 0, x_162); -lean_ctor_set(x_260, 1, x_172); -lean_ctor_set(x_260, 2, x_259); -x_261 = lean_array_push(x_186, x_260); -x_262 = lean_array_push(x_261, x_184); +lean_ctor_set(x_239, 0, x_161); +lean_ctor_set(x_239, 1, x_170); +lean_ctor_set(x_239, 2, x_238); +x_240 = lean_array_push(x_207, x_239); +x_241 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__32; +x_242 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_242, 0, x_161); +lean_ctor_set(x_242, 1, x_241); +lean_ctor_set(x_242, 2, x_240); +x_243 = lean_array_push(x_172, x_213); +x_244 = lean_array_push(x_243, x_242); +x_245 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__30; +x_246 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_246, 0, x_161); +lean_ctor_set(x_246, 1, x_245); +lean_ctor_set(x_246, 2, x_244); +x_247 = lean_array_push(x_183, x_209); +x_248 = lean_array_push(x_247, x_211); +x_249 = lean_array_push(x_248, x_246); +x_250 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__28; +x_251 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_251, 0, x_161); +lean_ctor_set(x_251, 1, x_250); +lean_ctor_set(x_251, 2, x_249); +x_252 = lean_array_push(x_172, x_202); +x_253 = lean_array_push(x_252, x_251); +x_254 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__26; +x_255 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_255, 0, x_161); +lean_ctor_set(x_255, 1, x_254); +lean_ctor_set(x_255, 2, x_253); +x_256 = lean_array_push(x_172, x_255); +x_257 = lean_array_push(x_256, x_177); +x_258 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_258, 0, x_161); +lean_ctor_set(x_258, 1, x_170); +lean_ctor_set(x_258, 2, x_257); +x_259 = lean_array_push(x_184, x_258); +x_260 = lean_array_push(x_259, x_182); +x_261 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_261, 0, x_161); +lean_ctor_set(x_261, 1, x_187); +lean_ctor_set(x_261, 2, x_260); +x_262 = lean_array_push(x_207, x_261); x_263 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_263, 0, x_162); -lean_ctor_set(x_263, 1, x_189); +lean_ctor_set(x_263, 0, x_161); +lean_ctor_set(x_263, 1, x_170); lean_ctor_set(x_263, 2, x_262); -x_264 = lean_array_push(x_209, x_263); -x_265 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_265, 0, x_162); -lean_ctor_set(x_265, 1, x_172); -lean_ctor_set(x_265, 2, x_264); -x_266 = lean_array_push(x_174, x_202); -x_267 = lean_array_push(x_266, x_265); -x_268 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_268, 0, x_162); -lean_ctor_set(x_268, 1, x_177); -lean_ctor_set(x_268, 2, x_267); -x_269 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_269, 0, x_268); -lean_ctor_set(x_269, 1, x_2); -x_270 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_270, 0, x_269); -lean_ctor_set(x_270, 1, x_146); -return x_270; +x_264 = lean_array_push(x_172, x_200); +x_265 = lean_array_push(x_264, x_263); +x_266 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_266, 0, x_161); +lean_ctor_set(x_266, 1, x_175); +lean_ctor_set(x_266, 2, x_265); +x_267 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_267, 0, x_266); +lean_ctor_set(x_267, 1, x_2); +x_268 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_268, 0, x_267); +lean_ctor_set(x_268, 1, x_145); +return x_268; } } } @@ -9447,26 +9434,56 @@ static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJ _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("some", 4); +x_1 = lean_mk_string_from_bytes("term", 4); return x_1; } } static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__12() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__11; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__12; +x_3 = lean_alloc_ctor(1, 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_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__14() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("some", 4); +return x_1; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__15() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__11; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__14; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__13() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__11; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__14; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__12; +x_3 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__15; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -9474,44 +9491,44 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__14() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__11; +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__14; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__15() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__7; -x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__11; +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__14; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__16() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__15; +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__18; 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_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__17() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__16; +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__19; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); @@ -9610,7 +9627,7 @@ return x_48; } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; size_t x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; size_t 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; size_t x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; size_t 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_inc(x_11); x_49 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_11, x_12, x_30); x_50 = lean_ctor_get(x_49, 0); @@ -9630,16 +9647,16 @@ x_56 = lean_ctor_get(x_54, 0); lean_inc(x_56); lean_dec(x_54); x_57 = lean_environment_main_module(x_56); -x_58 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__14; -x_59 = l_Lean_addMacroScope(x_57, x_58, x_52); -x_60 = lean_box(0); -x_61 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__13; -x_62 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__17; +x_58 = lean_box(0); +x_59 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__17; +x_60 = l_Lean_addMacroScope(x_57, x_59, x_52); +x_61 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__16; +x_62 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__20; lean_inc(x_50); x_63 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_63, 0, x_50); lean_ctor_set(x_63, 1, x_61); -lean_ctor_set(x_63, 2, x_59); +lean_ctor_set(x_63, 2, x_60); lean_ctor_set(x_63, 3, x_62); x_64 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__13; lean_inc(x_50); @@ -9648,63 +9665,64 @@ lean_ctor_set(x_65, 0, x_50); lean_ctor_set(x_65, 1, x_64); x_66 = lean_usize_of_nat(x_31); lean_dec(x_31); -x_67 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__6; -x_68 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__5(x_67, x_60, x_66, x_25, x_22); -x_69 = lean_array_get_size(x_68); -x_70 = lean_usize_of_nat(x_69); -lean_dec(x_69); -x_71 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_70, x_25, x_68); -x_72 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__14; -x_73 = l_Lean_mkSepArray(x_71, x_72); -lean_dec(x_71); -x_74 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__1; -x_75 = l_Array_append___rarg(x_74, x_73); -x_76 = lean_box(2); -x_77 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__9; -x_78 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_78, 0, x_76); -lean_ctor_set(x_78, 1, x_77); -lean_ctor_set(x_78, 2, x_75); -x_79 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__4; -x_80 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_80, 0, x_50); -lean_ctor_set(x_80, 1, x_79); -x_81 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__5; -x_82 = lean_array_push(x_81, x_65); -x_83 = lean_array_push(x_82, x_78); -x_84 = lean_array_push(x_83, x_80); -x_85 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__12; -x_86 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_86, 0, x_76); -lean_ctor_set(x_86, 1, x_85); -lean_ctor_set(x_86, 2, x_84); -x_87 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__10; -x_88 = lean_array_push(x_87, x_86); -x_89 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_89, 0, x_76); -lean_ctor_set(x_89, 1, x_77); -lean_ctor_set(x_89, 2, x_88); -x_90 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__7; -x_91 = lean_array_push(x_90, x_63); -x_92 = lean_array_push(x_91, x_89); -x_93 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__3; -x_94 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_94, 0, x_76); -lean_ctor_set(x_94, 1, x_93); -lean_ctor_set(x_94, 2, x_92); -x_95 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1(x_4, x_14, x_74, x_26, x_29, x_25, x_94, x_7, x_8, x_9, x_10, x_11, x_12, x_55); +x_67 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__13; +x_68 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__6; +x_69 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__5(x_67, x_68, x_58, x_66, x_25, x_22); +x_70 = lean_array_get_size(x_69); +x_71 = lean_usize_of_nat(x_70); +lean_dec(x_70); +x_72 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_71, x_25, x_69); +x_73 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__14; +x_74 = l_Lean_mkSepArray(x_72, x_73); +lean_dec(x_72); +x_75 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__1; +x_76 = l_Array_append___rarg(x_75, x_74); +x_77 = lean_box(2); +x_78 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__9; +x_79 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_79, 0, x_77); +lean_ctor_set(x_79, 1, x_78); +lean_ctor_set(x_79, 2, x_76); +x_80 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__4; +x_81 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_81, 0, x_50); +lean_ctor_set(x_81, 1, x_80); +x_82 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__5; +x_83 = lean_array_push(x_82, x_65); +x_84 = lean_array_push(x_83, x_79); +x_85 = lean_array_push(x_84, x_81); +x_86 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__12; +x_87 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_87, 0, x_77); +lean_ctor_set(x_87, 1, x_86); +lean_ctor_set(x_87, 2, x_85); +x_88 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__10; +x_89 = lean_array_push(x_88, x_87); +x_90 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_90, 0, x_77); +lean_ctor_set(x_90, 1, x_78); +lean_ctor_set(x_90, 2, x_89); +x_91 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__7; +x_92 = lean_array_push(x_91, x_63); +x_93 = lean_array_push(x_92, x_90); +x_94 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__3; +x_95 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_95, 0, x_77); +lean_ctor_set(x_95, 1, x_94); +lean_ctor_set(x_95, 2, x_93); +x_96 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1(x_4, x_14, x_75, x_26, x_29, x_25, x_95, x_7, x_8, x_9, x_10, x_11, x_12, x_55); lean_dec(x_12); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_29); -return x_95; +return x_96; } } else { -uint8_t x_96; +uint8_t x_97; lean_dec(x_14); lean_dec(x_12); lean_dec(x_11); @@ -9715,23 +9733,23 @@ lean_dec(x_7); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_96 = !lean_is_exclusive(x_18); -if (x_96 == 0) +x_97 = !lean_is_exclusive(x_18); +if (x_97 == 0) { return x_18; } else { -lean_object* x_97; lean_object* x_98; lean_object* x_99; -x_97 = lean_ctor_get(x_18, 0); -x_98 = lean_ctor_get(x_18, 1); +lean_object* x_98; lean_object* x_99; lean_object* x_100; +x_98 = lean_ctor_get(x_18, 0); +x_99 = lean_ctor_get(x_18, 1); +lean_inc(x_99); lean_inc(x_98); -lean_inc(x_97); lean_dec(x_18); -x_99 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_99, 0, x_97); -lean_ctor_set(x_99, 1, x_98); -return x_99; +x_100 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_100, 0, x_98); +lean_ctor_set(x_100, 1, x_99); +return x_100; } } } @@ -10098,7 +10116,7 @@ static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHand _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_instInhabitedSyntax; +x_1 = lean_box(0); x_2 = l_instInhabitedNat; x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10267,16 +10285,17 @@ x_12 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInsta return x_12; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__5___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: { -size_t x_6; size_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_3); -lean_dec(x_3); +size_t x_7; size_t x_8; lean_object* x_9; x_7 = lean_unbox_usize(x_4); lean_dec(x_4); -x_8 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__5(x_1, x_2, x_6, x_7, x_5); -return x_8; +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__5(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_1); +return x_9; } } LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___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, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { @@ -10364,172 +10383,173 @@ x_1 = lean_mk_string_from_bytes("orElseLazy", 10); return x_1; } } -LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__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, lean_object* x_9, lean_object* x_10, size_t x_11, size_t 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, lean_object* x_20) { +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__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, lean_object* x_9, lean_object* x_10, lean_object* x_11, size_t x_12, size_t 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, lean_object* x_20, lean_object* x_21) { _start: { -uint8_t x_21; -x_21 = lean_usize_dec_eq(x_11, x_12); -if (x_21 == 0) +uint8_t x_22; +x_22 = lean_usize_dec_eq(x_12, x_13); +if (x_22 == 0) { -size_t x_22; size_t 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; -x_22 = 1; -x_23 = lean_usize_sub(x_11, x_22); -x_24 = lean_array_uget(x_10, x_23); -lean_inc(x_18); -x_25 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_18, x_19, x_20); -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 1); +size_t x_23; size_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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_23 = 1; +x_24 = lean_usize_sub(x_12, x_23); +x_25 = lean_array_uget(x_11, x_24); +lean_inc(x_19); +x_26 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_19, x_20, x_21); +x_27 = lean_ctor_get(x_26, 0); lean_inc(x_27); -lean_dec(x_25); -x_28 = lean_ctor_get(x_18, 10); +x_28 = lean_ctor_get(x_26, 1); lean_inc(x_28); -x_29 = lean_st_ref_get(x_19, x_27); -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_29, 1); +lean_dec(x_26); +x_29 = lean_ctor_get(x_19, 10); +lean_inc(x_29); +x_30 = lean_st_ref_get(x_20, x_28); +x_31 = lean_ctor_get(x_30, 0); lean_inc(x_31); -lean_dec(x_29); -x_32 = lean_ctor_get(x_30, 0); +x_32 = lean_ctor_get(x_30, 1); lean_inc(x_32); lean_dec(x_30); -x_33 = lean_environment_main_module(x_32); -x_34 = l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__1___closed__1; -x_35 = l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__1___closed__2; -lean_inc(x_3); -x_36 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_3); -lean_ctor_set(x_36, 2, x_35); -x_37 = l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__1___closed__3; +x_33 = lean_ctor_get(x_31, 0); +lean_inc(x_33); +lean_dec(x_31); +x_34 = lean_environment_main_module(x_33); +x_35 = l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__1___closed__1; +x_36 = l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__1___closed__2; lean_inc(x_4); -x_38 = lean_name_mk_string(x_4, x_37); -lean_inc(x_38); -x_39 = l_Lean_addMacroScope(x_33, x_38, x_28); +x_37 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_4); +lean_ctor_set(x_37, 2, x_36); +x_38 = l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__1___closed__3; lean_inc(x_5); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_5); +x_39 = lean_name_mk_string(x_5, x_38); +lean_inc(x_39); +x_40 = l_Lean_addMacroScope(x_34, x_39, x_29); lean_inc(x_6); -x_41 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_41, 0, x_40); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_39); lean_ctor_set(x_41, 1, x_6); -lean_inc(x_26); -x_42 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_42, 0, x_26); -lean_ctor_set(x_42, 1, x_36); -lean_ctor_set(x_42, 2, x_39); -lean_ctor_set(x_42, 3, x_41); -x_43 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__1; -lean_inc(x_1); -x_44 = lean_name_mk_string(x_1, x_43); -x_45 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__2; -lean_inc(x_26); -x_46 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_46, 0, x_26); -lean_ctor_set(x_46, 1, x_45); -x_47 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__25; -lean_inc(x_1); -x_48 = lean_name_mk_string(x_1, x_47); -lean_inc(x_26); -x_49 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_49, 0, x_26); -lean_ctor_set(x_49, 1, x_47); -x_50 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__27; -lean_inc(x_1); -x_51 = lean_name_mk_string(x_1, x_50); -x_52 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__7; -lean_inc(x_1); -x_53 = lean_name_mk_string(x_1, x_52); -x_54 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__9; -lean_inc(x_26); -x_55 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_55, 0, x_26); -lean_ctor_set(x_55, 1, x_54); -lean_inc(x_8); -x_56 = lean_array_push(x_8, x_55); -x_57 = lean_box(2); -x_58 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_58, 0, x_57); -lean_ctor_set(x_58, 1, x_53); -lean_ctor_set(x_58, 2, x_56); -lean_inc(x_8); -x_59 = lean_array_push(x_8, x_58); lean_inc(x_7); -x_60 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_60, 0, x_57); -lean_ctor_set(x_60, 1, x_7); -lean_ctor_set(x_60, 2, x_59); -x_61 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__15; -lean_inc(x_26); -x_62 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_62, 0, x_26); -lean_ctor_set(x_62, 1, x_61); -x_63 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__5; -x_64 = lean_array_push(x_63, x_60); -x_65 = lean_array_push(x_64, x_62); -x_66 = lean_array_push(x_65, x_13); -x_67 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_67, 0, x_57); -lean_ctor_set(x_67, 1, x_51); -lean_ctor_set(x_67, 2, x_66); -lean_inc(x_9); -x_68 = lean_array_push(x_9, x_49); -x_69 = lean_array_push(x_68, x_67); -x_70 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_70, 0, x_57); -lean_ctor_set(x_70, 1, x_48); -lean_ctor_set(x_70, 2, x_69); -x_71 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__1; -lean_inc(x_7); -x_72 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_72, 0, x_57); -lean_ctor_set(x_72, 1, x_7); -lean_ctor_set(x_72, 2, x_71); -lean_inc(x_9); -x_73 = lean_array_push(x_9, x_70); -x_74 = lean_array_push(x_73, x_72); -lean_inc(x_7); -x_75 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_75, 0, x_57); -lean_ctor_set(x_75, 1, x_7); -lean_ctor_set(x_75, 2, x_74); -x_76 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__4; -x_77 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_77, 0, x_26); -lean_ctor_set(x_77, 1, x_76); -x_78 = lean_array_push(x_63, x_46); -x_79 = lean_array_push(x_78, x_75); -x_80 = lean_array_push(x_79, x_77); -x_81 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_81, 0, x_57); -lean_ctor_set(x_81, 1, x_44); -lean_ctor_set(x_81, 2, x_80); +x_42 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_7); +lean_inc(x_27); +x_43 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_43, 0, x_27); +lean_ctor_set(x_43, 1, x_37); +lean_ctor_set(x_43, 2, x_40); +lean_ctor_set(x_43, 3, x_42); +x_44 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__1; +lean_inc(x_2); +x_45 = lean_name_mk_string(x_2, x_44); +x_46 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__2; +lean_inc(x_27); +x_47 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_47, 0, x_27); +lean_ctor_set(x_47, 1, x_46); +x_48 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__25; +lean_inc(x_2); +x_49 = lean_name_mk_string(x_2, x_48); +lean_inc(x_27); +x_50 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_50, 0, x_27); +lean_ctor_set(x_50, 1, x_48); +x_51 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__27; +lean_inc(x_2); +x_52 = lean_name_mk_string(x_2, x_51); +x_53 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__7; +lean_inc(x_2); +x_54 = lean_name_mk_string(x_2, x_53); +x_55 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__9; +lean_inc(x_27); +x_56 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_56, 0, x_27); +lean_ctor_set(x_56, 1, x_55); lean_inc(x_9); -x_82 = lean_array_push(x_9, x_24); -x_83 = lean_array_push(x_82, x_81); -lean_inc(x_7); -x_84 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_84, 0, x_57); -lean_ctor_set(x_84, 1, x_7); -lean_ctor_set(x_84, 2, x_83); +x_57 = lean_array_push(x_9, x_56); +x_58 = lean_box(2); +x_59 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_54); +lean_ctor_set(x_59, 2, x_57); lean_inc(x_9); -x_85 = lean_array_push(x_9, x_42); -x_86 = lean_array_push(x_85, x_84); -lean_inc(x_2); -x_87 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_87, 0, x_57); -lean_ctor_set(x_87, 1, x_2); -lean_ctor_set(x_87, 2, x_86); -x_11 = x_23; -x_13 = x_87; -x_20 = x_31; +x_60 = lean_array_push(x_9, x_59); +lean_inc(x_8); +x_61 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_61, 0, x_58); +lean_ctor_set(x_61, 1, x_8); +lean_ctor_set(x_61, 2, x_60); +x_62 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__15; +lean_inc(x_27); +x_63 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_63, 0, x_27); +lean_ctor_set(x_63, 1, x_62); +x_64 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__5; +x_65 = lean_array_push(x_64, x_61); +x_66 = lean_array_push(x_65, x_63); +x_67 = lean_array_push(x_66, x_14); +x_68 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_68, 0, x_58); +lean_ctor_set(x_68, 1, x_52); +lean_ctor_set(x_68, 2, x_67); +lean_inc(x_10); +x_69 = lean_array_push(x_10, x_50); +x_70 = lean_array_push(x_69, x_68); +x_71 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_71, 0, x_58); +lean_ctor_set(x_71, 1, x_49); +lean_ctor_set(x_71, 2, x_70); +x_72 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__1; +lean_inc(x_8); +x_73 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_73, 0, x_58); +lean_ctor_set(x_73, 1, x_8); +lean_ctor_set(x_73, 2, x_72); +lean_inc(x_10); +x_74 = lean_array_push(x_10, x_71); +x_75 = lean_array_push(x_74, x_73); +lean_inc(x_8); +x_76 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_76, 0, x_58); +lean_ctor_set(x_76, 1, x_8); +lean_ctor_set(x_76, 2, x_75); +x_77 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__4; +x_78 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_78, 0, x_27); +lean_ctor_set(x_78, 1, x_77); +x_79 = lean_array_push(x_64, x_47); +x_80 = lean_array_push(x_79, x_76); +x_81 = lean_array_push(x_80, x_78); +x_82 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_82, 0, x_58); +lean_ctor_set(x_82, 1, x_45); +lean_ctor_set(x_82, 2, x_81); +lean_inc(x_10); +x_83 = lean_array_push(x_10, x_25); +x_84 = lean_array_push(x_83, x_82); +lean_inc(x_8); +x_85 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_85, 0, x_58); +lean_ctor_set(x_85, 1, x_8); +lean_ctor_set(x_85, 2, x_84); +lean_inc(x_10); +x_86 = lean_array_push(x_10, x_43); +x_87 = lean_array_push(x_86, x_85); +lean_inc(x_3); +x_88 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_88, 0, x_58); +lean_ctor_set(x_88, 1, x_3); +lean_ctor_set(x_88, 2, x_87); +x_12 = x_24; +x_14 = x_88; +x_21 = x_32; goto _start; } else { -lean_object* x_89; -lean_dec(x_18); +lean_object* x_90; +lean_dec(x_19); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -10538,180 +10558,180 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_89 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_89, 0, x_13); -lean_ctor_set(x_89, 1, x_20); -return x_89; +x_90 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_90, 0, x_14); +lean_ctor_set(x_90, 1, x_21); +return x_90; } } } -LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___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_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, size_t x_11, size_t 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, lean_object* x_20) { +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___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_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, size_t x_12, size_t 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, lean_object* x_20, lean_object* x_21) { _start: { -uint8_t x_21; -x_21 = lean_usize_dec_eq(x_11, x_12); -if (x_21 == 0) +uint8_t x_22; +x_22 = lean_usize_dec_eq(x_12, x_13); +if (x_22 == 0) { -size_t x_22; size_t 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; -x_22 = 1; -x_23 = lean_usize_sub(x_11, x_22); -x_24 = lean_array_uget(x_10, x_23); -lean_inc(x_18); -x_25 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_18, x_19, x_20); -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 1); +size_t x_23; size_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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_23 = 1; +x_24 = lean_usize_sub(x_12, x_23); +x_25 = lean_array_uget(x_11, x_24); +lean_inc(x_19); +x_26 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_19, x_20, x_21); +x_27 = lean_ctor_get(x_26, 0); lean_inc(x_27); -lean_dec(x_25); -x_28 = lean_ctor_get(x_18, 10); +x_28 = lean_ctor_get(x_26, 1); lean_inc(x_28); -x_29 = lean_st_ref_get(x_19, x_27); -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_29, 1); +lean_dec(x_26); +x_29 = lean_ctor_get(x_19, 10); +lean_inc(x_29); +x_30 = lean_st_ref_get(x_20, x_28); +x_31 = lean_ctor_get(x_30, 0); lean_inc(x_31); -lean_dec(x_29); -x_32 = lean_ctor_get(x_30, 0); +x_32 = lean_ctor_get(x_30, 1); lean_inc(x_32); lean_dec(x_30); -x_33 = lean_environment_main_module(x_32); -x_34 = l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__1___closed__1; -x_35 = l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__1___closed__2; -lean_inc(x_3); -x_36 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_3); -lean_ctor_set(x_36, 2, x_35); -x_37 = l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__1___closed__3; +x_33 = lean_ctor_get(x_31, 0); +lean_inc(x_33); +lean_dec(x_31); +x_34 = lean_environment_main_module(x_33); +x_35 = l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__1___closed__1; +x_36 = l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__1___closed__2; lean_inc(x_4); -x_38 = lean_name_mk_string(x_4, x_37); -lean_inc(x_38); -x_39 = l_Lean_addMacroScope(x_33, x_38, x_28); +x_37 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_4); +lean_ctor_set(x_37, 2, x_36); +x_38 = l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__1___closed__3; lean_inc(x_5); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_5); +x_39 = lean_name_mk_string(x_5, x_38); +lean_inc(x_39); +x_40 = l_Lean_addMacroScope(x_34, x_39, x_29); lean_inc(x_6); -x_41 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_41, 0, x_40); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_39); lean_ctor_set(x_41, 1, x_6); -lean_inc(x_26); -x_42 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_42, 0, x_26); -lean_ctor_set(x_42, 1, x_36); -lean_ctor_set(x_42, 2, x_39); -lean_ctor_set(x_42, 3, x_41); -x_43 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__1; -lean_inc(x_1); -x_44 = lean_name_mk_string(x_1, x_43); -x_45 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__2; -lean_inc(x_26); -x_46 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_46, 0, x_26); -lean_ctor_set(x_46, 1, x_45); -x_47 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__25; -lean_inc(x_1); -x_48 = lean_name_mk_string(x_1, x_47); -lean_inc(x_26); -x_49 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_49, 0, x_26); -lean_ctor_set(x_49, 1, x_47); -x_50 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__27; -lean_inc(x_1); -x_51 = lean_name_mk_string(x_1, x_50); -x_52 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__7; -lean_inc(x_1); -x_53 = lean_name_mk_string(x_1, x_52); -x_54 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__9; -lean_inc(x_26); -x_55 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_55, 0, x_26); -lean_ctor_set(x_55, 1, x_54); -lean_inc(x_8); -x_56 = lean_array_push(x_8, x_55); -x_57 = lean_box(2); -x_58 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_58, 0, x_57); -lean_ctor_set(x_58, 1, x_53); -lean_ctor_set(x_58, 2, x_56); -lean_inc(x_8); -x_59 = lean_array_push(x_8, x_58); -lean_inc(x_7); -x_60 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_60, 0, x_57); -lean_ctor_set(x_60, 1, x_7); -lean_ctor_set(x_60, 2, x_59); -x_61 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__15; -lean_inc(x_26); -x_62 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_62, 0, x_26); -lean_ctor_set(x_62, 1, x_61); -x_63 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__5; -x_64 = lean_array_push(x_63, x_60); -x_65 = lean_array_push(x_64, x_62); -x_66 = lean_array_push(x_65, x_13); -x_67 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_67, 0, x_57); -lean_ctor_set(x_67, 1, x_51); -lean_ctor_set(x_67, 2, x_66); -lean_inc(x_9); -x_68 = lean_array_push(x_9, x_49); -x_69 = lean_array_push(x_68, x_67); -x_70 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_70, 0, x_57); -lean_ctor_set(x_70, 1, x_48); -lean_ctor_set(x_70, 2, x_69); -x_71 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__1; -lean_inc(x_7); -x_72 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_72, 0, x_57); -lean_ctor_set(x_72, 1, x_7); -lean_ctor_set(x_72, 2, x_71); -lean_inc(x_9); -x_73 = lean_array_push(x_9, x_70); -x_74 = lean_array_push(x_73, x_72); lean_inc(x_7); -x_75 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_75, 0, x_57); -lean_ctor_set(x_75, 1, x_7); -lean_ctor_set(x_75, 2, x_74); -x_76 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__4; -x_77 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_77, 0, x_26); -lean_ctor_set(x_77, 1, x_76); -x_78 = lean_array_push(x_63, x_46); -x_79 = lean_array_push(x_78, x_75); -x_80 = lean_array_push(x_79, x_77); -x_81 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_81, 0, x_57); -lean_ctor_set(x_81, 1, x_44); -lean_ctor_set(x_81, 2, x_80); +x_42 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_7); +lean_inc(x_27); +x_43 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_43, 0, x_27); +lean_ctor_set(x_43, 1, x_37); +lean_ctor_set(x_43, 2, x_40); +lean_ctor_set(x_43, 3, x_42); +x_44 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__1; +lean_inc(x_2); +x_45 = lean_name_mk_string(x_2, x_44); +x_46 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__2; +lean_inc(x_27); +x_47 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_47, 0, x_27); +lean_ctor_set(x_47, 1, x_46); +x_48 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__25; +lean_inc(x_2); +x_49 = lean_name_mk_string(x_2, x_48); +lean_inc(x_27); +x_50 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_50, 0, x_27); +lean_ctor_set(x_50, 1, x_48); +x_51 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__27; +lean_inc(x_2); +x_52 = lean_name_mk_string(x_2, x_51); +x_53 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__7; +lean_inc(x_2); +x_54 = lean_name_mk_string(x_2, x_53); +x_55 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__9; +lean_inc(x_27); +x_56 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_56, 0, x_27); +lean_ctor_set(x_56, 1, x_55); lean_inc(x_9); -x_82 = lean_array_push(x_9, x_24); -x_83 = lean_array_push(x_82, x_81); -lean_inc(x_7); -x_84 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_84, 0, x_57); -lean_ctor_set(x_84, 1, x_7); -lean_ctor_set(x_84, 2, x_83); +x_57 = lean_array_push(x_9, x_56); +x_58 = lean_box(2); +x_59 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_54); +lean_ctor_set(x_59, 2, x_57); lean_inc(x_9); -x_85 = lean_array_push(x_9, x_42); -x_86 = lean_array_push(x_85, x_84); -lean_inc(x_2); -x_87 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_87, 0, x_57); -lean_ctor_set(x_87, 1, x_2); -lean_ctor_set(x_87, 2, x_86); -x_11 = x_23; -x_13 = x_87; -x_20 = x_31; +x_60 = lean_array_push(x_9, x_59); +lean_inc(x_8); +x_61 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_61, 0, x_58); +lean_ctor_set(x_61, 1, x_8); +lean_ctor_set(x_61, 2, x_60); +x_62 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__15; +lean_inc(x_27); +x_63 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_63, 0, x_27); +lean_ctor_set(x_63, 1, x_62); +x_64 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__5; +x_65 = lean_array_push(x_64, x_61); +x_66 = lean_array_push(x_65, x_63); +x_67 = lean_array_push(x_66, x_14); +x_68 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_68, 0, x_58); +lean_ctor_set(x_68, 1, x_52); +lean_ctor_set(x_68, 2, x_67); +lean_inc(x_10); +x_69 = lean_array_push(x_10, x_50); +x_70 = lean_array_push(x_69, x_68); +x_71 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_71, 0, x_58); +lean_ctor_set(x_71, 1, x_49); +lean_ctor_set(x_71, 2, x_70); +x_72 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__1; +lean_inc(x_8); +x_73 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_73, 0, x_58); +lean_ctor_set(x_73, 1, x_8); +lean_ctor_set(x_73, 2, x_72); +lean_inc(x_10); +x_74 = lean_array_push(x_10, x_71); +x_75 = lean_array_push(x_74, x_73); +lean_inc(x_8); +x_76 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_76, 0, x_58); +lean_ctor_set(x_76, 1, x_8); +lean_ctor_set(x_76, 2, x_75); +x_77 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__4; +x_78 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_78, 0, x_27); +lean_ctor_set(x_78, 1, x_77); +x_79 = lean_array_push(x_64, x_47); +x_80 = lean_array_push(x_79, x_76); +x_81 = lean_array_push(x_80, x_78); +x_82 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_82, 0, x_58); +lean_ctor_set(x_82, 1, x_45); +lean_ctor_set(x_82, 2, x_81); +lean_inc(x_10); +x_83 = lean_array_push(x_10, x_25); +x_84 = lean_array_push(x_83, x_82); +lean_inc(x_8); +x_85 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_85, 0, x_58); +lean_ctor_set(x_85, 1, x_8); +lean_ctor_set(x_85, 2, x_84); +lean_inc(x_10); +x_86 = lean_array_push(x_10, x_43); +x_87 = lean_array_push(x_86, x_85); +lean_inc(x_3); +x_88 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_88, 0, x_58); +lean_ctor_set(x_88, 1, x_3); +lean_ctor_set(x_88, 2, x_87); +x_12 = x_24; +x_14 = x_88; +x_21 = x_32; goto _start; } else { -lean_object* x_89; -lean_dec(x_18); +lean_object* x_90; +lean_dec(x_19); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -10720,11 +10740,10 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_89 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_89, 0, x_13); -lean_ctor_set(x_89, 1, x_20); -return x_89; +x_90 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_90, 0, x_14); +lean_ctor_set(x_90, 1, x_21); +return x_90; } } } @@ -11189,7 +11208,7 @@ static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHand _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__9; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__9; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } @@ -11198,7 +11217,7 @@ static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHand _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__9; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__9; x_2 = lean_unsigned_to_nat(0u); x_3 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__1___closed__3; x_4 = lean_alloc_ctor(0, 3, 0); @@ -11300,330 +11319,329 @@ return x_3; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___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, 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, 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, lean_object* x_25) { _start: { -lean_object* x_26; uint8_t x_167; -x_167 = lean_ctor_get_uint8(x_1, sizeof(void*)*2); -if (x_167 == 0) +lean_object* x_26; uint8_t x_166; +x_166 = lean_ctor_get_uint8(x_1, sizeof(void*)*2); +if (x_166 == 0) { -uint8_t x_168; -x_168 = lean_ctor_get_uint8(x_16, sizeof(void*)*5); -if (x_168 == 0) +uint8_t x_167; +x_167 = lean_ctor_get_uint8(x_16, sizeof(void*)*5); +if (x_167 == 0) { -lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; -x_169 = lean_ctor_get(x_5, 1); -lean_inc(x_169); +lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; +x_168 = lean_ctor_get(x_5, 1); +lean_inc(x_168); lean_inc(x_23); -x_170 = l_Lean_Elab_Deriving_mkInductiveApp(x_6, x_169, x_19, x_20, x_21, x_22, x_23, x_24, x_25); -x_171 = lean_ctor_get(x_170, 0); +x_169 = l_Lean_Elab_Deriving_mkInductiveApp(x_6, x_168, x_19, x_20, x_21, x_22, x_23, x_24, x_25); +x_170 = lean_ctor_get(x_169, 0); +lean_inc(x_170); +x_171 = lean_ctor_get(x_169, 1); lean_inc(x_171); -x_172 = lean_ctor_get(x_170, 1); -lean_inc(x_172); -lean_dec(x_170); +lean_dec(x_169); lean_inc(x_23); -x_173 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_23, x_24, x_172); -x_174 = lean_ctor_get(x_173, 0); +x_172 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_23, x_24, x_171); +x_173 = lean_ctor_get(x_172, 0); +lean_inc(x_173); +x_174 = lean_ctor_get(x_172, 1); lean_inc(x_174); -x_175 = lean_ctor_get(x_173, 1); +lean_dec(x_172); +x_175 = lean_ctor_get(x_23, 10); lean_inc(x_175); -lean_dec(x_173); -x_176 = lean_ctor_get(x_23, 10); -lean_inc(x_176); -x_177 = lean_st_ref_get(x_24, x_175); -x_178 = lean_ctor_get(x_177, 0); +x_176 = lean_st_ref_get(x_24, x_174); +x_177 = lean_ctor_get(x_176, 0); +lean_inc(x_177); +x_178 = lean_ctor_get(x_176, 1); lean_inc(x_178); -x_179 = lean_ctor_get(x_177, 1); +lean_dec(x_176); +x_179 = lean_ctor_get(x_177, 0); lean_inc(x_179); lean_dec(x_177); -x_180 = lean_ctor_get(x_178, 0); -lean_inc(x_180); -lean_dec(x_178); -x_181 = lean_environment_main_module(x_180); -x_182 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__10; -x_183 = lean_name_mk_string(x_7, x_182); -x_184 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__12; -lean_inc(x_183); -x_185 = lean_name_mk_string(x_183, x_184); -x_186 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__14; -lean_inc(x_183); -x_187 = lean_name_mk_string(x_183, x_186); -x_188 = lean_box(2); -x_189 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__1; +x_180 = lean_environment_main_module(x_179); +x_181 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__10; +x_182 = lean_name_mk_string(x_7, x_181); +x_183 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__12; +lean_inc(x_182); +x_184 = lean_name_mk_string(x_182, x_183); +x_185 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__14; +lean_inc(x_182); +x_186 = lean_name_mk_string(x_182, x_185); +x_187 = lean_box(2); +x_188 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__5___lambda__1___closed__1; lean_inc(x_8); -x_190 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_190, 0, x_188); -lean_ctor_set(x_190, 1, x_8); -lean_ctor_set(x_190, 2, x_189); -x_191 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__16; -lean_inc(x_183); -x_192 = lean_name_mk_string(x_183, x_191); -lean_inc(x_174); -x_193 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_193, 0, x_174); -lean_ctor_set(x_193, 1, x_191); +x_189 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_189, 0, x_187); +lean_ctor_set(x_189, 1, x_8); +lean_ctor_set(x_189, 2, x_188); +x_190 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__16; +lean_inc(x_182); +x_191 = lean_name_mk_string(x_182, x_190); +lean_inc(x_173); +x_192 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_192, 0, x_173); +lean_ctor_set(x_192, 1, x_190); lean_inc(x_4); -x_194 = lean_array_push(x_4, x_193); -x_195 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_195, 0, x_188); -lean_ctor_set(x_195, 1, x_192); -lean_ctor_set(x_195, 2, x_194); +x_193 = lean_array_push(x_4, x_192); +x_194 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_194, 0, x_187); +lean_ctor_set(x_194, 1, x_191); +lean_ctor_set(x_194, 2, x_193); lean_inc(x_4); -x_196 = lean_array_push(x_4, x_195); +x_195 = lean_array_push(x_4, x_194); lean_inc(x_8); -x_197 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_197, 0, x_188); -lean_ctor_set(x_197, 1, x_8); -lean_ctor_set(x_197, 2, x_196); -x_198 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__9; -lean_inc(x_190); -x_199 = lean_array_push(x_198, x_190); -lean_inc(x_190); -x_200 = lean_array_push(x_199, x_190); -x_201 = lean_array_push(x_200, x_197); -lean_inc(x_190); -x_202 = lean_array_push(x_201, x_190); -lean_inc(x_190); -x_203 = lean_array_push(x_202, x_190); -lean_inc(x_190); -x_204 = lean_array_push(x_203, x_190); -x_205 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_205, 0, x_188); -lean_ctor_set(x_205, 1, x_187); -lean_ctor_set(x_205, 2, x_204); -x_206 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__20; -lean_inc(x_183); -x_207 = lean_name_mk_string(x_183, x_206); -lean_inc(x_174); -x_208 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_208, 0, x_174); -lean_ctor_set(x_208, 1, x_206); -x_209 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__22; -lean_inc(x_183); -x_210 = lean_name_mk_string(x_183, x_209); +x_196 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_196, 0, x_187); +lean_ctor_set(x_196, 1, x_8); +lean_ctor_set(x_196, 2, x_195); +x_197 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__9; +lean_inc(x_189); +x_198 = lean_array_push(x_197, x_189); +lean_inc(x_189); +x_199 = lean_array_push(x_198, x_189); +x_200 = lean_array_push(x_199, x_196); +lean_inc(x_189); +x_201 = lean_array_push(x_200, x_189); +lean_inc(x_189); +x_202 = lean_array_push(x_201, x_189); +lean_inc(x_189); +x_203 = lean_array_push(x_202, x_189); +x_204 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_204, 0, x_187); +lean_ctor_set(x_204, 1, x_186); +lean_ctor_set(x_204, 2, x_203); +x_205 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__20; +lean_inc(x_182); +x_206 = lean_name_mk_string(x_182, x_205); +lean_inc(x_173); +x_207 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_207, 0, x_173); +lean_ctor_set(x_207, 1, x_205); +x_208 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__22; +lean_inc(x_182); +x_209 = lean_name_mk_string(x_182, x_208); lean_inc(x_9); -x_211 = lean_array_push(x_9, x_10); -lean_inc(x_190); -x_212 = lean_array_push(x_211, x_190); -x_213 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_213, 0, x_188); -lean_ctor_set(x_213, 1, x_210); -lean_ctor_set(x_213, 2, x_212); -x_214 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__24; -lean_inc(x_183); -x_215 = lean_name_mk_string(x_183, x_214); -x_216 = lean_ctor_get(x_5, 0); -lean_inc(x_216); +x_210 = lean_array_push(x_9, x_10); +lean_inc(x_189); +x_211 = lean_array_push(x_210, x_189); +x_212 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_212, 0, x_187); +lean_ctor_set(x_212, 1, x_209); +lean_ctor_set(x_212, 2, x_211); +x_213 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__24; +lean_inc(x_182); +x_214 = lean_name_mk_string(x_182, x_213); +x_215 = lean_ctor_get(x_5, 0); +lean_inc(x_215); lean_dec(x_5); -x_217 = l_Array_append___rarg(x_189, x_216); -x_218 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__1___closed__1; +x_216 = l_Array_append___rarg(x_188, x_215); +x_217 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__1___closed__1; lean_inc(x_11); -x_219 = lean_name_mk_string(x_11, x_218); -x_220 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__2; -lean_inc(x_174); -x_221 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_221, 0, x_174); -lean_ctor_set(x_221, 1, x_220); -x_222 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__15; -lean_inc(x_176); -lean_inc(x_181); -x_223 = l_Lean_addMacroScope(x_181, x_222, x_176); -x_224 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__14; +x_218 = lean_name_mk_string(x_11, x_217); +x_219 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__2; +lean_inc(x_173); +x_220 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_220, 0, x_173); +lean_ctor_set(x_220, 1, x_219); +x_221 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__15; +lean_inc(x_175); +lean_inc(x_180); +x_222 = l_Lean_addMacroScope(x_180, x_221, x_175); +x_223 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__1___closed__14; lean_inc(x_12); -lean_inc(x_174); -x_225 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_225, 0, x_174); -lean_ctor_set(x_225, 1, x_224); -lean_ctor_set(x_225, 2, x_223); -lean_ctor_set(x_225, 3, x_12); +lean_inc(x_173); +x_224 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_224, 0, x_173); +lean_ctor_set(x_224, 1, x_223); +lean_ctor_set(x_224, 2, x_222); +lean_ctor_set(x_224, 3, x_12); lean_inc(x_4); -x_226 = lean_array_push(x_4, x_225); +x_225 = lean_array_push(x_4, x_224); lean_inc(x_8); -x_227 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_227, 0, x_188); -lean_ctor_set(x_227, 1, x_8); -lean_ctor_set(x_227, 2, x_226); -x_228 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__1___closed__2; -lean_inc(x_174); -x_229 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_229, 0, x_174); -lean_ctor_set(x_229, 1, x_228); -x_230 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__8; -lean_inc(x_176); -lean_inc(x_181); -x_231 = l_Lean_addMacroScope(x_181, x_230, x_176); -x_232 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__9; -x_233 = lean_name_mk_string(x_13, x_232); +x_226 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_226, 0, x_187); +lean_ctor_set(x_226, 1, x_8); +lean_ctor_set(x_226, 2, x_225); +x_227 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__1___closed__2; +lean_inc(x_173); +x_228 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_228, 0, x_173); +lean_ctor_set(x_228, 1, x_227); +x_229 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__8; +lean_inc(x_175); +lean_inc(x_180); +x_230 = l_Lean_addMacroScope(x_180, x_229, x_175); +x_231 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__9; +x_232 = lean_name_mk_string(x_13, x_231); lean_inc(x_12); -x_234 = lean_alloc_ctor(0, 2, 0); +x_233 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_233, 0, x_232); +lean_ctor_set(x_233, 1, x_12); +lean_inc(x_12); +x_234 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_234, 0, x_233); lean_ctor_set(x_234, 1, x_12); -lean_inc(x_12); -x_235 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_235, 0, x_234); -lean_ctor_set(x_235, 1, x_12); -x_236 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__1___closed__4; -lean_inc(x_174); -x_237 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_237, 0, x_174); -lean_ctor_set(x_237, 1, x_236); -lean_ctor_set(x_237, 2, x_231); -lean_ctor_set(x_237, 3, x_235); +x_235 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__1___closed__4; +lean_inc(x_173); +x_236 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_236, 0, x_173); +lean_ctor_set(x_236, 1, x_235); +lean_ctor_set(x_236, 2, x_230); +lean_ctor_set(x_236, 3, x_234); lean_inc(x_9); -x_238 = lean_array_push(x_9, x_229); -lean_inc(x_238); -x_239 = lean_array_push(x_238, x_237); +x_237 = lean_array_push(x_9, x_228); +lean_inc(x_237); +x_238 = lean_array_push(x_237, x_236); lean_inc(x_8); -x_240 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_240, 0, x_188); -lean_ctor_set(x_240, 1, x_8); -lean_ctor_set(x_240, 2, x_239); -x_241 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__4; -lean_inc(x_174); -x_242 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_242, 0, x_174); -lean_ctor_set(x_242, 1, x_241); -x_243 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__1___closed__5; -x_244 = lean_array_push(x_243, x_221); -x_245 = lean_array_push(x_244, x_227); -x_246 = lean_array_push(x_245, x_240); -lean_inc(x_190); -x_247 = lean_array_push(x_246, x_190); -x_248 = lean_array_push(x_247, x_242); -x_249 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_249, 0, x_188); -lean_ctor_set(x_249, 1, x_219); -lean_ctor_set(x_249, 2, x_248); -x_250 = lean_array_push(x_217, x_249); +x_239 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_239, 0, x_187); +lean_ctor_set(x_239, 1, x_8); +lean_ctor_set(x_239, 2, x_238); +x_240 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__4; +lean_inc(x_173); +x_241 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_241, 0, x_173); +lean_ctor_set(x_241, 1, x_240); +x_242 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__1___closed__5; +x_243 = lean_array_push(x_242, x_220); +x_244 = lean_array_push(x_243, x_226); +x_245 = lean_array_push(x_244, x_239); +lean_inc(x_189); +x_246 = lean_array_push(x_245, x_189); +x_247 = lean_array_push(x_246, x_241); +x_248 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_248, 0, x_187); +lean_ctor_set(x_248, 1, x_218); +lean_ctor_set(x_248, 2, x_247); +x_249 = lean_array_push(x_216, x_248); lean_inc(x_8); -x_251 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_251, 0, x_188); -lean_ctor_set(x_251, 1, x_8); -lean_ctor_set(x_251, 2, x_250); -x_252 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__1___closed__6; -x_253 = lean_name_mk_string(x_11, x_252); -lean_inc(x_176); +x_250 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_250, 0, x_187); +lean_ctor_set(x_250, 1, x_8); +lean_ctor_set(x_250, 2, x_249); +x_251 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__1___closed__6; +x_252 = lean_name_mk_string(x_11, x_251); +lean_inc(x_175); lean_inc(x_14); -lean_inc(x_181); -x_254 = l_Lean_addMacroScope(x_181, x_14, x_176); +lean_inc(x_180); +x_253 = l_Lean_addMacroScope(x_180, x_14, x_175); lean_inc(x_12); -x_255 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_255, 0, x_14); -lean_ctor_set(x_255, 1, x_12); +x_254 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_254, 0, x_14); +lean_ctor_set(x_254, 1, x_12); lean_inc(x_12); -x_256 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_256, 0, x_255); -lean_ctor_set(x_256, 1, x_12); -x_257 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__1___closed__9; -lean_inc(x_174); -x_258 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_258, 0, x_174); -lean_ctor_set(x_258, 1, x_257); -lean_ctor_set(x_258, 2, x_254); -lean_ctor_set(x_258, 3, x_256); -x_259 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__1___closed__13; -x_260 = l_Lean_addMacroScope(x_181, x_259, x_176); +x_255 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_255, 0, x_254); +lean_ctor_set(x_255, 1, x_12); +x_256 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__1___closed__9; +lean_inc(x_173); +x_257 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_257, 0, x_173); +lean_ctor_set(x_257, 1, x_256); +lean_ctor_set(x_257, 2, x_253); +lean_ctor_set(x_257, 3, x_255); +x_258 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__1___closed__13; +x_259 = l_Lean_addMacroScope(x_180, x_258, x_175); lean_inc(x_12); -x_261 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_261, 0, x_259); +x_260 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_260, 0, x_258); +lean_ctor_set(x_260, 1, x_12); +x_261 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_261, 0, x_260); lean_ctor_set(x_261, 1, x_12); -x_262 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_262, 0, x_261); -lean_ctor_set(x_262, 1, x_12); -x_263 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__1___closed__12; -lean_inc(x_174); -x_264 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_264, 0, x_174); -lean_ctor_set(x_264, 1, x_263); -lean_ctor_set(x_264, 2, x_260); -lean_ctor_set(x_264, 3, x_262); +x_262 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__1___closed__12; +lean_inc(x_173); +x_263 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_263, 0, x_173); +lean_ctor_set(x_263, 1, x_262); +lean_ctor_set(x_263, 2, x_259); +lean_ctor_set(x_263, 3, x_261); lean_inc(x_9); -x_265 = lean_array_push(x_9, x_264); -x_266 = lean_array_push(x_265, x_171); +x_264 = lean_array_push(x_9, x_263); +x_265 = lean_array_push(x_264, x_170); lean_inc(x_8); -x_267 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_267, 0, x_188); -lean_ctor_set(x_267, 1, x_8); -lean_ctor_set(x_267, 2, x_266); +x_266 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_266, 0, x_187); +lean_ctor_set(x_266, 1, x_8); +lean_ctor_set(x_266, 2, x_265); lean_inc(x_9); -x_268 = lean_array_push(x_9, x_258); -x_269 = lean_array_push(x_268, x_267); -x_270 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_270, 0, x_188); -lean_ctor_set(x_270, 1, x_15); -lean_ctor_set(x_270, 2, x_269); -x_271 = lean_array_push(x_238, x_270); -x_272 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_272, 0, x_188); -lean_ctor_set(x_272, 1, x_253); -lean_ctor_set(x_272, 2, x_271); +x_267 = lean_array_push(x_9, x_257); +x_268 = lean_array_push(x_267, x_266); +x_269 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_269, 0, x_187); +lean_ctor_set(x_269, 1, x_15); +lean_ctor_set(x_269, 2, x_268); +x_270 = lean_array_push(x_237, x_269); +x_271 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_271, 0, x_187); +lean_ctor_set(x_271, 1, x_252); +lean_ctor_set(x_271, 2, x_270); lean_inc(x_4); -x_273 = lean_array_push(x_4, x_272); -x_274 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_274, 0, x_188); -lean_ctor_set(x_274, 1, x_8); -lean_ctor_set(x_274, 2, x_273); +x_272 = lean_array_push(x_4, x_271); +x_273 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_273, 0, x_187); +lean_ctor_set(x_273, 1, x_8); +lean_ctor_set(x_273, 2, x_272); lean_inc(x_9); -x_275 = lean_array_push(x_9, x_251); -x_276 = lean_array_push(x_275, x_274); -x_277 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_277, 0, x_188); -lean_ctor_set(x_277, 1, x_215); -lean_ctor_set(x_277, 2, x_276); -x_278 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__26; -x_279 = lean_name_mk_string(x_183, x_278); -x_280 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__28; -x_281 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_281, 0, x_174); -lean_ctor_set(x_281, 1, x_280); -x_282 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__5; -x_283 = lean_array_push(x_282, x_281); -x_284 = lean_array_push(x_283, x_17); -lean_inc(x_190); -x_285 = lean_array_push(x_284, x_190); -x_286 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_286, 0, x_188); -lean_ctor_set(x_286, 1, x_279); -lean_ctor_set(x_286, 2, x_285); -x_287 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__29; -x_288 = lean_array_push(x_287, x_208); -x_289 = lean_array_push(x_288, x_213); -x_290 = lean_array_push(x_289, x_277); -x_291 = lean_array_push(x_290, x_286); -lean_inc(x_190); -x_292 = lean_array_push(x_291, x_190); -lean_inc(x_190); -x_293 = lean_array_push(x_292, x_190); -x_294 = lean_array_push(x_293, x_190); -x_295 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_295, 0, x_188); -lean_ctor_set(x_295, 1, x_207); -lean_ctor_set(x_295, 2, x_294); -x_296 = lean_array_push(x_9, x_205); -x_297 = lean_array_push(x_296, x_295); -x_298 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_298, 0, x_188); -lean_ctor_set(x_298, 1, x_185); -lean_ctor_set(x_298, 2, x_297); -x_299 = lean_box(0); -x_300 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__2(x_1, x_2, x_3, x_4, x_298, x_299, x_19, x_20, x_21, x_22, x_23, x_24, x_179); -return x_300; +x_274 = lean_array_push(x_9, x_250); +x_275 = lean_array_push(x_274, x_273); +x_276 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_276, 0, x_187); +lean_ctor_set(x_276, 1, x_214); +lean_ctor_set(x_276, 2, x_275); +x_277 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__26; +x_278 = lean_name_mk_string(x_182, x_277); +x_279 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__28; +x_280 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_280, 0, x_173); +lean_ctor_set(x_280, 1, x_279); +x_281 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__5; +x_282 = lean_array_push(x_281, x_280); +x_283 = lean_array_push(x_282, x_17); +lean_inc(x_189); +x_284 = lean_array_push(x_283, x_189); +x_285 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_285, 0, x_187); +lean_ctor_set(x_285, 1, x_278); +lean_ctor_set(x_285, 2, x_284); +x_286 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__29; +x_287 = lean_array_push(x_286, x_207); +x_288 = lean_array_push(x_287, x_212); +x_289 = lean_array_push(x_288, x_276); +x_290 = lean_array_push(x_289, x_285); +lean_inc(x_189); +x_291 = lean_array_push(x_290, x_189); +lean_inc(x_189); +x_292 = lean_array_push(x_291, x_189); +x_293 = lean_array_push(x_292, x_189); +x_294 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_294, 0, x_187); +lean_ctor_set(x_294, 1, x_206); +lean_ctor_set(x_294, 2, x_293); +x_295 = lean_array_push(x_9, x_204); +x_296 = lean_array_push(x_295, x_294); +x_297 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_297, 0, x_187); +lean_ctor_set(x_297, 1, x_184); +lean_ctor_set(x_297, 2, x_296); +x_298 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__2(x_1, x_2, x_3, x_4, x_297, x_19, x_20, x_21, x_22, x_23, x_24, x_178); +return x_298; } else { -lean_object* x_301; -x_301 = lean_box(0); -x_26 = x_301; -goto block_166; +lean_object* x_299; +x_299 = lean_box(0); +x_26 = x_299; +goto block_165; } } else { -lean_object* x_302; -x_302 = lean_box(0); -x_26 = x_302; -goto block_166; +lean_object* x_300; +x_300 = lean_box(0); +x_26 = x_300; +goto block_165; } -block_166: +block_165: { -lean_object* x_27; lean_object* x_28; lean_object* x_29; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; +lean_object* x_27; lean_object* x_28; lean_object* x_29; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_dec(x_26); x_27 = lean_ctor_get(x_5, 1); lean_inc(x_27); @@ -11784,7 +11802,7 @@ x_95 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___clo lean_inc(x_34); lean_inc(x_39); x_96 = l_Lean_addMacroScope(x_39, x_95, x_34); -x_97 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__9; +x_97 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__9; x_98 = lean_name_mk_string(x_13, x_97); lean_inc(x_12); x_99 = lean_alloc_ctor(0, 2, 0); @@ -11937,9 +11955,8 @@ x_163 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_163, 0, x_46); lean_ctor_set(x_163, 1, x_43); lean_ctor_set(x_163, 2, x_162); -x_164 = lean_box(0); -x_165 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__2(x_1, x_2, x_3, x_4, x_163, x_164, x_19, x_20, x_21, x_22, x_23, x_24, x_37); -return x_165; +x_164 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__2(x_1, x_2, x_3, x_4, x_163, x_19, x_20, x_21, x_22, x_23, x_24, x_37); +return x_164; } } } @@ -12145,16 +12162,16 @@ x_37 = lean_ctor_get(x_35, 0); lean_inc(x_37); lean_dec(x_35); x_38 = lean_environment_main_module(x_37); -x_39 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__2___closed__7; -x_40 = l_Lean_addMacroScope(x_38, x_39, x_33); -x_41 = lean_box(0); +x_39 = lean_box(0); +x_40 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__2___closed__7; +x_41 = l_Lean_addMacroScope(x_38, x_40, x_33); x_42 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__2___closed__4; x_43 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__2___closed__9; lean_inc(x_31); x_44 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_44, 0, x_31); lean_ctor_set(x_44, 1, x_42); -lean_ctor_set(x_44, 2, x_40); +lean_ctor_set(x_44, 2, x_41); lean_ctor_set(x_44, 3, x_43); x_45 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__2___closed__12; x_46 = lean_alloc_ctor(2, 2, 0); @@ -12198,30 +12215,31 @@ goto block_88; } else { -size_t x_90; size_t x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; +size_t x_90; size_t x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; x_90 = lean_usize_of_nat(x_60); lean_dec(x_60); x_91 = 0; -x_92 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__6; -x_93 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__2___closed__5; +x_92 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__13; +x_93 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__6; +x_94 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__2___closed__5; lean_inc(x_8); -x_94 = l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__1(x_92, x_58, x_17, x_93, x_41, x_41, x_53, x_47, x_55, x_28, x_90, x_91, x_59, x_4, x_5, x_6, x_7, x_8, x_9, x_36); +x_95 = l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__1(x_92, x_93, x_58, x_17, x_94, x_39, x_39, x_53, x_47, x_55, x_28, x_90, x_91, x_59, x_4, x_5, x_6, x_7, x_8, x_9, x_36); lean_dec(x_28); -x_95 = lean_ctor_get(x_94, 0); -lean_inc(x_95); -x_96 = lean_ctor_get(x_94, 1); +x_96 = lean_ctor_get(x_95, 0); lean_inc(x_96); -lean_dec(x_94); -x_62 = x_95; -x_63 = x_96; +x_97 = lean_ctor_get(x_95, 1); +lean_inc(x_97); +lean_dec(x_95); +x_62 = x_96; +x_63 = x_97; goto block_88; } } else { -uint8_t x_97; -x_97 = lean_nat_dec_lt(x_17, x_60); -if (x_97 == 0) +uint8_t x_98; +x_98 = lean_nat_dec_lt(x_17, x_60); +if (x_98 == 0) { lean_dec(x_60); lean_dec(x_28); @@ -12231,22 +12249,23 @@ goto block_88; } else { -size_t x_98; size_t x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; -x_98 = lean_usize_of_nat(x_60); +size_t x_99; size_t x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_99 = lean_usize_of_nat(x_60); lean_dec(x_60); -x_99 = 0; -x_100 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__6; -x_101 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__2___closed__5; +x_100 = 0; +x_101 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__13; +x_102 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__6; +x_103 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__2___closed__5; lean_inc(x_8); -x_102 = l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__2(x_100, x_58, x_17, x_101, x_41, x_41, x_53, x_47, x_55, x_28, x_98, x_99, x_59, x_4, x_5, x_6, x_7, x_8, x_9, x_36); +x_104 = l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__2(x_101, x_102, x_58, x_17, x_103, x_39, x_39, x_53, x_47, x_55, x_28, x_99, x_100, x_59, x_4, x_5, x_6, x_7, x_8, x_9, x_36); lean_dec(x_28); -x_103 = lean_ctor_get(x_102, 0); -lean_inc(x_103); -x_104 = lean_ctor_get(x_102, 1); -lean_inc(x_104); -lean_dec(x_102); -x_62 = x_103; -x_63 = x_104; +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +lean_dec(x_104); +x_62 = x_105; +x_63 = x_106; goto block_88; } } @@ -12262,7 +12281,7 @@ x_66 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanc x_67 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__2; x_68 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__2___closed__5; x_69 = lean_box(0); -x_70 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__1(x_13, x_19, x_3, x_47, x_21, x_18, x_65, x_53, x_55, x_26, x_66, x_41, x_67, x_68, x_58, x_2, x_62, x_69, x_4, x_5, x_6, x_7, x_8, x_9, x_63); +x_70 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__1(x_13, x_19, x_3, x_47, x_21, x_18, x_65, x_53, x_55, x_26, x_66, x_39, x_67, x_68, x_58, x_2, x_62, x_69, x_4, x_5, x_6, x_7, x_8, x_9, x_63); lean_dec(x_2); lean_dec(x_13); return x_70; @@ -12305,7 +12324,7 @@ x_79 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanc x_80 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__3___closed__2; x_81 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__2___closed__5; x_82 = lean_box(0); -x_83 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__1(x_13, x_19, x_3, x_47, x_21, x_18, x_78, x_53, x_55, x_26, x_79, x_41, x_80, x_81, x_58, x_2, x_76, x_82, x_4, x_5, x_6, x_7, x_8, x_9, x_77); +x_83 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__1(x_13, x_19, x_3, x_47, x_21, x_18, x_78, x_53, x_55, x_26, x_79, x_39, x_80, x_81, x_58, x_2, x_76, x_82, x_4, x_5, x_6, x_7, x_8, x_9, x_77); lean_dec(x_2); lean_dec(x_13); return x_83; @@ -12349,7 +12368,7 @@ return x_87; } else { -uint8_t x_105; +uint8_t x_107; lean_dec(x_26); lean_dec(x_21); lean_dec(x_18); @@ -12361,29 +12380,29 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_105 = !lean_is_exclusive(x_27); -if (x_105 == 0) +x_107 = !lean_is_exclusive(x_27); +if (x_107 == 0) { return x_27; } else { -lean_object* x_106; lean_object* x_107; lean_object* x_108; -x_106 = lean_ctor_get(x_27, 0); -x_107 = lean_ctor_get(x_27, 1); -lean_inc(x_107); -lean_inc(x_106); +lean_object* x_108; lean_object* x_109; lean_object* x_110; +x_108 = lean_ctor_get(x_27, 0); +x_109 = lean_ctor_get(x_27, 1); +lean_inc(x_109); +lean_inc(x_108); lean_dec(x_27); -x_108 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_108, 0, x_106); -lean_ctor_set(x_108, 1, x_107); -return x_108; +x_110 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_110, 0, x_108); +lean_ctor_set(x_110, 1, x_109); +return x_110; } } } else { -uint8_t x_109; +uint8_t x_111; lean_dec(x_18); lean_dec(x_13); lean_dec(x_9); @@ -12393,29 +12412,29 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_109 = !lean_is_exclusive(x_20); -if (x_109 == 0) +x_111 = !lean_is_exclusive(x_20); +if (x_111 == 0) { return x_20; } else { -lean_object* x_110; lean_object* x_111; lean_object* x_112; -x_110 = lean_ctor_get(x_20, 0); -x_111 = lean_ctor_get(x_20, 1); -lean_inc(x_111); -lean_inc(x_110); +lean_object* x_112; lean_object* x_113; lean_object* x_114; +x_112 = lean_ctor_get(x_20, 0); +x_113 = lean_ctor_get(x_20, 1); +lean_inc(x_113); +lean_inc(x_112); lean_dec(x_20); -x_112 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_112, 0, x_110); -lean_ctor_set(x_112, 1, x_111); -return x_112; +x_114 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_114, 0, x_112); +lean_ctor_set(x_114, 1, x_113); +return x_114; } } } else { -uint8_t x_113; +uint8_t x_115; lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -12423,23 +12442,23 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_113 = !lean_is_exclusive(x_12); -if (x_113 == 0) +x_115 = !lean_is_exclusive(x_12); +if (x_115 == 0) { return x_12; } else { -lean_object* x_114; lean_object* x_115; lean_object* x_116; -x_114 = lean_ctor_get(x_12, 0); -x_115 = lean_ctor_get(x_12, 1); -lean_inc(x_115); -lean_inc(x_114); +lean_object* x_116; lean_object* x_117; lean_object* x_118; +x_116 = lean_ctor_get(x_12, 0); +x_117 = lean_ctor_get(x_12, 1); +lean_inc(x_117); +lean_inc(x_116); lean_dec(x_12); -x_116 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_116, 0, x_114); -lean_ctor_set(x_116, 1, x_115); -return x_116; +x_118 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_118, 0, x_116); +lean_ctor_set(x_118, 1, x_117); +return x_118; } } } @@ -13746,21 +13765,23 @@ lean_object* x_17 = _args[16]; lean_object* x_18 = _args[17]; lean_object* x_19 = _args[18]; lean_object* x_20 = _args[19]; +lean_object* x_21 = _args[20]; _start: { -size_t x_21; size_t x_22; lean_object* x_23; -x_21 = lean_unbox_usize(x_11); -lean_dec(x_11); +size_t x_22; size_t x_23; lean_object* x_24; x_22 = lean_unbox_usize(x_12); lean_dec(x_12); -x_23 = l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21, x_22, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); -lean_dec(x_19); +x_23 = lean_unbox_usize(x_13); +lean_dec(x_13); +x_24 = l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_22, x_23, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21); +lean_dec(x_20); +lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_10); -return x_23; +lean_dec(x_11); +lean_dec(x_1); +return x_24; } } LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__2___boxed(lean_object** _args) { @@ -13784,21 +13805,23 @@ lean_object* x_17 = _args[16]; lean_object* x_18 = _args[17]; lean_object* x_19 = _args[18]; lean_object* x_20 = _args[19]; +lean_object* x_21 = _args[20]; _start: { -size_t x_21; size_t x_22; lean_object* x_23; -x_21 = lean_unbox_usize(x_11); -lean_dec(x_11); +size_t x_22; size_t x_23; lean_object* x_24; x_22 = lean_unbox_usize(x_12); lean_dec(x_12); -x_23 = l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21, x_22, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); -lean_dec(x_19); +x_23 = lean_unbox_usize(x_13); +lean_dec(x_13); +x_24 = l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_22, x_23, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21); +lean_dec(x_20); +lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_10); -return x_23; +lean_dec(x_11); +lean_dec(x_1); +return x_24; } } LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { @@ -13921,7 +13944,7 @@ lean_dec(x_2); return x_10; } } -static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6045____closed__1() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6284____closed__1() { _start: { lean_object* x_1; @@ -13929,7 +13952,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanc return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6045____closed__2() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6284____closed__2() { _start: { lean_object* x_1; @@ -13937,12 +13960,12 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Deriving_FromToJson_mkFromJsonInsta return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6045_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6284_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; x_2 = l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__2; -x_3 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6045____closed__1; +x_3 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6284____closed__1; x_4 = l_Lean_Elab_registerBuiltinDerivingHandler(x_2, x_3, x_1); if (lean_obj_tag(x_4) == 0) { @@ -13951,7 +13974,7 @@ x_5 = lean_ctor_get(x_4, 1); lean_inc(x_5); lean_dec(x_4); x_6 = l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__3___closed__8; -x_7 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6045____closed__2; +x_7 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6284____closed__2; x_8 = l_Lean_Elab_registerBuiltinDerivingHandler(x_6, x_7, x_5); return x_8; } @@ -14092,24 +14115,24 @@ l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler__ lean_mark_persistent(l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__4); l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__5 = _init_l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__5(); lean_mark_persistent(l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4___closed__5); -l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__1 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__1(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__1); -l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__2 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__2(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__2); -l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__3 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__3(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__3); -l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__4 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__4(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__4); -l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__5 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__5(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__5); -l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__6 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__6(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__6); -l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__7 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__7(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__7); -l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__8 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__8(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__8); -l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__9 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__9(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__6___closed__9); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__1 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__1(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__1); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__2 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__2(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__2); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__3 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__3(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__3); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__4 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__4(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__4); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__5 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__5(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__5); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__6 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__6(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__6); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__7 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__7(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__7); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__8 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__8(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__8); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__9 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__9(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__7___closed__9); l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__1 = _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__1(); lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__1); l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__2 = _init_l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__1___closed__2(); @@ -14386,6 +14409,12 @@ l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHand lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__16); l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__17 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__17(); lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__17); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__18 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__18(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__18); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__19 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__19(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__19); +l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__20 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__20(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__6___lambda__2___closed__20); l_Array_qsort_sort___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__7___closed__1 = _init_l_Array_qsort_sort___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__7___closed__1(); lean_mark_persistent(l_Array_qsort_sort___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__7___closed__1); l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___closed__1 = _init_l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___closed__1(); @@ -14496,11 +14525,11 @@ l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__3___closed__ lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__3___closed__19); l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__3___closed__20 = _init_l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__3___closed__20(); lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__3___closed__20); -l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6045____closed__1 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6045____closed__1(); -lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6045____closed__1); -l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6045____closed__2 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6045____closed__2(); -lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6045____closed__2); -res = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6045_(lean_io_mk_world()); +l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6284____closed__1 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6284____closed__1(); +lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6284____closed__1); +l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6284____closed__2 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6284____closed__2(); +lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6284____closed__2); +res = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6284_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/Deriving/Hashable.c b/stage0/stdlib/Lean/Elab/Deriving/Hashable.c index 8898e7f2548..98c9e4547e0 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/Hashable.c +++ b/stage0/stdlib/Lean/Elab/Deriving/Hashable.c @@ -15,7 +15,7 @@ extern "C" { #endif lean_object* l_List_reverse___rarg(lean_object*); lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); lean_object* l_Lean_registerTraceClass(lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); @@ -23,6 +23,7 @@ lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Hashable_mkHashableHandler(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__10; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__9; +static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__26; lean_object* l_Lean_Elab_Command_liftTermElabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__24; static lean_object* l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__24; @@ -31,12 +32,11 @@ static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMa uint8_t lean_usize_dec_eq(size_t, size_t); lean_object* lean_array_uget(lean_object*, size_t); lean_object* l_Array_append___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___closed__2; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__10; static lean_object* l_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___closed__1; +lean_object* lean_array_uset(lean_object*, size_t, lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__6; static lean_object* l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__16; -static lean_object* l_Lean_Elab_Deriving_Hashable_mkMatch___closed__7; lean_object* lean_environment_find(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Deriving_Hashable_0__Lean_Elab_Deriving_Hashable_mkHashableInstanceCmds(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); @@ -60,6 +60,7 @@ static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_m lean_object* lean_string_utf8_byte_size(lean_object*); static lean_object* l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__19; static lean_object* l_Lean_Elab_Deriving_Hashable_mkMatch___closed__5; +uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_isInductive___at_Lean_Elab_Deriving_Hashable_mkHashableHandler___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__30; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__20; @@ -83,7 +84,6 @@ static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMa uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__22; static lean_object* l___private_Lean_Elab_Deriving_Hashable_0__Lean_Elab_Deriving_Hashable_mkHashableInstanceCmds___closed__1; -extern lean_object* l_Lean_numLitKind; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Hashable_mkAuxFunction(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Deriving_mkInstanceCmds(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__7; @@ -97,6 +97,7 @@ static lean_object* l___private_Lean_Elab_Deriving_Hashable_0__Lean_Elab_Derivin static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__6; LEAN_EXPORT lean_object* l_Lean_isInductive___at_Lean_Elab_Deriving_Hashable_mkHashableHandler___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__23; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_Hashable_mkMatch___spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__14; lean_object* l_Nat_repr(lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__3___closed__7; @@ -109,7 +110,7 @@ static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMa extern lean_object* l_Lean_instInhabitedExpr; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__15; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkHashFuncs___spec__1(lean_object*, lean_object*, lean_object*, 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_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5(lean_object*, lean_object*, lean_object*, 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_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkSepArray(lean_object*, lean_object*); static lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__1___closed__3; static lean_object* l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__2; @@ -130,13 +131,14 @@ static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMa lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__16(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__17; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__1; +lean_object* l_Lean_Syntax_mkNumLit(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__27; lean_object* l_Lean_Elab_Deriving_mkHeader(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__4; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__3___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Elab_Deriving_Hashable_0__Lean_Elab_Deriving_Hashable_mkHashableInstanceCmds___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__4; lean_object* l_List_redLength___rarg(lean_object*); static lean_object* l___private_Lean_Elab_Deriving_Hashable_0__Lean_Elab_Deriving_Hashable_mkHashableInstanceCmds___closed__2; @@ -147,6 +149,7 @@ static lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_Hashable_mkM static lean_object* l___private_Lean_Elab_Deriving_Hashable_0__Lean_Elab_Deriving_Hashable_mkHashableInstanceCmds___closed__6; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___lambda__1___boxed(lean_object*, lean_object*); lean_object* lean_environment_main_module(lean_object*); +static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__27; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__9; uint8_t lean_nat_dec_le(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Hashable_mkHashableHeader(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -157,7 +160,7 @@ static lean_object* l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__11; static lean_object* l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__7; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__3___closed__1; static lean_object* l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__3; -LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Hashable_initFn____x40_Lean_Elab_Deriving_Hashable___hyg_2199_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Hashable_initFn____x40_Lean_Elab_Deriving_Hashable___hyg_2330_(lean_object*); static lean_object* l_Lean_Elab_Deriving_Hashable_mkHashFuncs___closed__4; static lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__1___closed__2; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__10; @@ -168,12 +171,14 @@ static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMa lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_instInhabitedName; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__8; +static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__12; lean_object* l_List_toArrayAux___rarg(lean_object*, lean_object*); lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_registerBuiltinDerivingHandler(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__31; static lean_object* l_Lean_Elab_Deriving_Hashable_mkHashFuncs___closed__1; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__17; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_Hashable_mkMatch___spec__1(size_t, size_t, lean_object*); lean_object* l_Lean_Meta_forallTelescopeReducing___at_Lean_Elab_Deriving_mkInductArgNames___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTRAux___at___private_Lean_Elab_Deriving_Hashable_0__Lean_Elab_Deriving_Hashable_mkHashableInstanceCmds___spec__1(lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__3___closed__2; @@ -182,12 +187,13 @@ static lean_object* l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__26; lean_object* lean_mk_syntax_ident(lean_object*); extern lean_object* l_Lean_instInhabitedInductiveVal; static lean_object* l___private_Lean_Elab_Deriving_Hashable_0__Lean_Elab_Deriving_Hashable_mkHashableInstanceCmds___closed__10; -static lean_object* l_Lean_Elab_Deriving_Hashable_initFn____x40_Lean_Elab_Deriving_Hashable___hyg_2199____closed__1; +static lean_object* l_Lean_Elab_Deriving_Hashable_initFn____x40_Lean_Elab_Deriving_Hashable___hyg_2330____closed__1; static lean_object* l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__12; static lean_object* l_Lean_Elab_Deriving_Hashable_mkHashFuncs___closed__5; lean_object* l_Lean_Expr_getAppFn(lean_object*); lean_object* l_Array_findIdx_x3f_loop___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Hashable_mkMatch(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__13; static lean_object* l_Lean_Elab_Deriving_Hashable_mkHashFuncs___closed__3; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Deriving_Hashable_0__Lean_Elab_Deriving_Hashable_mkHashableInstanceCmds___closed__8; @@ -197,6 +203,7 @@ lean_object* l_Lean_mkConst(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Deriving_Hashable_0__Lean_Elab_Deriving_Hashable_mkHashableInstanceCmds___closed__4; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__11; static lean_object* l_Lean_Elab_Deriving_Hashable_mkHashableHeader___closed__2; +static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__11; static lean_object* l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(lean_object*, lean_object*, lean_object*); @@ -206,7 +213,6 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable static lean_object* l_Lean_Elab_Deriving_Hashable_mkHashableHeader___closed__1; lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__13; -lean_object* l_Lean_Syntax_mkLit(lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__3___closed__9; static lean_object* l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__20; @@ -815,12 +821,35 @@ return x_2; static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__24() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__25() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(2); +x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__12; +x_3 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__24; +x_4 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__26() { +_start: +{ lean_object* x_1; x_1 = lean_mk_string_from_bytes(")", 1); return x_1; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__25() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__27() { _start: { lean_object* x_1; lean_object* x_2; @@ -852,7 +881,7 @@ lean_inc(x_41); lean_dec(x_10); x_42 = lean_nat_add(x_5, x_7); x_43 = l_Lean_instInhabitedExpr; -x_44 = lean_array_get(x_43, x_4, x_42); +x_44 = lean_array_get(x_43, x_3, x_42); lean_dec(x_42); x_45 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__2; x_46 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_45, x_15, x_16, x_17); @@ -900,8 +929,8 @@ lean_inc(x_58); lean_dec(x_57); x_59 = lean_alloc_closure((void*)(l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___lambda__1___boxed), 2, 1); lean_closure_set(x_59, 0, x_58); -x_60 = lean_array_get_size(x_3); -x_61 = l_Array_findIdx_x3f_loop___rarg(x_3, x_59, x_60, x_19, lean_box(0)); +x_60 = lean_array_get_size(x_2); +x_61 = l_Array_findIdx_x3f_loop___rarg(x_2, x_59, x_60, x_19, lean_box(0)); if (lean_obj_tag(x_61) == 0) { lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; @@ -967,22 +996,18 @@ x_92 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_92, 0, x_85); lean_ctor_set(x_92, 1, x_91); lean_ctor_set(x_92, 2, x_90); -lean_inc(x_2); -x_93 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_93, 0, x_85); -lean_ctor_set(x_93, 1, x_86); -lean_ctor_set(x_93, 2, x_2); -x_94 = lean_array_push(x_88, x_92); -x_95 = lean_array_push(x_94, x_93); +x_93 = lean_array_push(x_88, x_92); +x_94 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__25; +x_95 = lean_array_push(x_93, x_94); x_96 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_96, 0, x_85); lean_ctor_set(x_96, 1, x_86); lean_ctor_set(x_96, 2, x_95); -x_97 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__24; +x_97 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__26; x_98 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_98, 0, x_63); lean_ctor_set(x_98, 1, x_97); -x_99 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__25; +x_99 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__27; x_100 = lean_array_push(x_99, x_77); x_101 = lean_array_push(x_100, x_96); x_102 = lean_array_push(x_101, x_98); @@ -1067,22 +1092,18 @@ x_138 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_138, 0, x_131); lean_ctor_set(x_138, 1, x_137); lean_ctor_set(x_138, 2, x_136); -lean_inc(x_2); -x_139 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_139, 0, x_131); -lean_ctor_set(x_139, 1, x_132); -lean_ctor_set(x_139, 2, x_2); -x_140 = lean_array_push(x_134, x_138); -x_141 = lean_array_push(x_140, x_139); +x_139 = lean_array_push(x_134, x_138); +x_140 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__25; +x_141 = lean_array_push(x_139, x_140); x_142 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_142, 0, x_131); lean_ctor_set(x_142, 1, x_132); lean_ctor_set(x_142, 2, x_141); -x_143 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__24; +x_143 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__26; x_144 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_144, 0, x_63); lean_ctor_set(x_144, 1, x_143); -x_145 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__25; +x_145 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__27; x_146 = lean_array_push(x_145, x_123); x_147 = lean_array_push(x_146, x_142); x_148 = lean_array_push(x_147, x_144); @@ -1176,22 +1197,18 @@ x_190 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_190, 0, x_183); lean_ctor_set(x_190, 1, x_189); lean_ctor_set(x_190, 2, x_188); -lean_inc(x_2); -x_191 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_191, 0, x_183); -lean_ctor_set(x_191, 1, x_184); -lean_ctor_set(x_191, 2, x_2); -x_192 = lean_array_push(x_186, x_190); -x_193 = lean_array_push(x_192, x_191); +x_191 = lean_array_push(x_186, x_190); +x_192 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__25; +x_193 = lean_array_push(x_191, x_192); x_194 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_194, 0, x_183); lean_ctor_set(x_194, 1, x_184); lean_ctor_set(x_194, 2, x_193); -x_195 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__24; +x_195 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__26; x_196 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_196, 0, x_162); lean_ctor_set(x_196, 1, x_195); -x_197 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__25; +x_197 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__27; x_198 = lean_array_push(x_197, x_176); x_199 = lean_array_push(x_198, x_194); x_200 = lean_array_push(x_199, x_196); @@ -1269,22 +1286,18 @@ x_235 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_235, 0, x_228); lean_ctor_set(x_235, 1, x_234); lean_ctor_set(x_235, 2, x_233); -lean_inc(x_2); -x_236 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_236, 0, x_228); -lean_ctor_set(x_236, 1, x_229); -lean_ctor_set(x_236, 2, x_2); -x_237 = lean_array_push(x_231, x_235); -x_238 = lean_array_push(x_237, x_236); +x_236 = lean_array_push(x_231, x_235); +x_237 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__25; +x_238 = lean_array_push(x_236, x_237); x_239 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_239, 0, x_228); lean_ctor_set(x_239, 1, x_229); lean_ctor_set(x_239, 2, x_238); -x_240 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__24; +x_240 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__26; x_241 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_241, 0, x_162); lean_ctor_set(x_241, 1, x_240); -x_242 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__25; +x_242 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__27; x_243 = lean_array_push(x_242, x_221); x_244 = lean_array_push(x_243, x_239); x_245 = lean_array_push(x_244, x_241); @@ -1384,22 +1397,18 @@ x_287 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_287, 0, x_280); lean_ctor_set(x_287, 1, x_286); lean_ctor_set(x_287, 2, x_285); -lean_inc(x_2); -x_288 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_288, 0, x_280); -lean_ctor_set(x_288, 1, x_281); -lean_ctor_set(x_288, 2, x_2); -x_289 = lean_array_push(x_283, x_287); -x_290 = lean_array_push(x_289, x_288); +x_288 = lean_array_push(x_283, x_287); +x_289 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__25; +x_290 = lean_array_push(x_288, x_289); x_291 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_291, 0, x_280); lean_ctor_set(x_291, 1, x_281); lean_ctor_set(x_291, 2, x_290); -x_292 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__24; +x_292 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__26; x_293 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_293, 0, x_258); lean_ctor_set(x_293, 1, x_292); -x_294 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__25; +x_294 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__27; x_295 = lean_array_push(x_294, x_272); x_296 = lean_array_push(x_295, x_291); x_297 = lean_array_push(x_296, x_293); @@ -1484,22 +1493,18 @@ x_333 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_333, 0, x_326); lean_ctor_set(x_333, 1, x_332); lean_ctor_set(x_333, 2, x_331); -lean_inc(x_2); -x_334 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_334, 0, x_326); -lean_ctor_set(x_334, 1, x_327); -lean_ctor_set(x_334, 2, x_2); -x_335 = lean_array_push(x_329, x_333); -x_336 = lean_array_push(x_335, x_334); +x_334 = lean_array_push(x_329, x_333); +x_335 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__25; +x_336 = lean_array_push(x_334, x_335); x_337 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_337, 0, x_326); lean_ctor_set(x_337, 1, x_327); lean_ctor_set(x_337, 2, x_336); -x_338 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__24; +x_338 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__26; x_339 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_339, 0, x_258); lean_ctor_set(x_339, 1, x_338); -x_340 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__25; +x_340 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__27; x_341 = lean_array_push(x_340, x_318); x_342 = lean_array_push(x_341, x_337); x_343 = lean_array_push(x_342, x_339); @@ -1604,7 +1609,6 @@ lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_7); -lean_dec(x_2); x_25 = !lean_is_exclusive(x_23); if (x_25 == 0) { @@ -1659,7 +1663,6 @@ lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_7); -lean_dec(x_2); x_35 = !lean_is_exclusive(x_23); if (x_35 == 0) { @@ -1690,7 +1693,6 @@ lean_dec(x_14); lean_dec(x_13); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_2); x_363 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_363, 0, x_10); lean_ctor_set(x_363, 1, x_17); @@ -1706,7 +1708,6 @@ lean_dec(x_14); lean_dec(x_13); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_2); x_364 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_364, 0, x_10); lean_ctor_set(x_364, 1, x_17); @@ -1718,7 +1719,7 @@ static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkM _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("explicit", 8); +x_1 = lean_mk_string_from_bytes("term", 4); return x_1; } } @@ -1726,7 +1727,7 @@ static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkM _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__3___closed__6; +x_1 = lean_box(0); x_2 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -1735,16 +1736,20 @@ return x_3; static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__3() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("@", 1); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__2; +x_3 = lean_alloc_ctor(1, 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_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__4() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("matchAlt", 8); +x_1 = lean_mk_string_from_bytes("explicit", 8); return x_1; } } @@ -1762,7 +1767,7 @@ static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkM _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("|", 1); +x_1 = lean_mk_string_from_bytes("@", 1); return x_1; } } @@ -1770,7 +1775,7 @@ static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkM _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes(",", 1); +x_1 = lean_mk_string_from_bytes("matchAlt", 8); return x_1; } } @@ -1778,15 +1783,41 @@ static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkM _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(2); +x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__3___closed__6; x_2 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__7; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("|", 1); +return x_1; +} +} +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(",", 1); +return x_1; +} +} +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(2); +x_2 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__10; x_3 = lean_alloc_ctor(2, 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_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__9() { +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__12() { _start: { lean_object* x_1; @@ -1794,7 +1825,7 @@ x_1 = lean_mk_string_from_bytes("=>", 2); return x_1; } } -static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__10() { +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__13() { _start: { lean_object* x_1; lean_object* x_2; @@ -1803,62 +1834,60 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___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, 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, lean_object* x_16) { +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___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, 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) { _start: { -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; 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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_17 = lean_ctor_get(x_1, 2); -lean_inc(x_17); -x_18 = lean_unsigned_to_nat(0u); -x_19 = lean_unsigned_to_nat(1u); -lean_inc(x_14); -lean_inc(x_2); -lean_inc(x_17); -x_20 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__3(x_17, x_18, x_17, x_19, x_2, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -lean_dec(x_17); +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; 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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_16 = lean_ctor_get(x_1, 2); +lean_inc(x_16); +x_17 = lean_unsigned_to_nat(0u); +x_18 = lean_unsigned_to_nat(1u); +x_19 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__24; +lean_inc(x_13); +lean_inc(x_16); +x_20 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__3(x_16, x_17, x_16, x_18, x_19, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_16); x_21 = lean_ctor_get(x_20, 0); lean_inc(x_21); x_22 = lean_ctor_get(x_20, 1); lean_inc(x_22); lean_dec(x_20); -lean_inc(x_14); -x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_14, x_15, x_22); +lean_inc(x_13); +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_13, x_14, x_22); x_24 = lean_ctor_get(x_23, 1); lean_inc(x_24); lean_dec(x_23); -x_25 = lean_st_ref_get(x_15, x_24); +x_25 = lean_st_ref_get(x_14, x_24); x_26 = lean_ctor_get(x_25, 1); lean_inc(x_26); lean_dec(x_25); -x_27 = l_Nat_repr(x_3); -x_28 = l_Lean_numLitKind; -x_29 = lean_box(2); -x_30 = l_Lean_Syntax_mkLit(x_28, x_27, x_29); -x_31 = lean_ctor_get(x_1, 1); -lean_inc(x_31); +x_27 = l_Nat_repr(x_2); +x_28 = lean_box(2); +x_29 = l_Lean_Syntax_mkNumLit(x_27, x_28); +x_30 = lean_ctor_get(x_1, 1); +lean_inc(x_30); lean_dec(x_1); -lean_inc(x_14); -lean_inc(x_2); -lean_inc(x_31); -x_32 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__3(x_31, x_18, x_31, x_19, x_2, x_10, x_11, x_12, x_13, x_14, x_15, x_26); -x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_13); +lean_inc(x_30); +x_31 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__3(x_30, x_17, x_30, x_18, x_19, x_9, x_10, x_11, x_12, x_13, x_14, x_26); +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_31, 1); lean_inc(x_33); -x_34 = lean_ctor_get(x_32, 1); +lean_dec(x_31); +x_34 = lean_ctor_get(x_3, 4); lean_inc(x_34); -lean_dec(x_32); -x_35 = lean_ctor_get(x_4, 4); -lean_inc(x_35); -lean_dec(x_4); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_33); -lean_ctor_set(x_36, 1, x_30); -lean_inc(x_15); +lean_dec(x_3); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_32); +lean_ctor_set(x_35, 1, x_29); +x_36 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__3; lean_inc(x_14); -lean_inc(x_35); -lean_inc(x_2); -x_37 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4(x_5, x_2, x_6, x_8, x_31, x_35, x_18, x_35, x_19, x_36, x_10, x_11, x_12, x_13, x_14, x_15, x_34); -lean_dec(x_35); -lean_dec(x_31); +lean_inc(x_13); +lean_inc(x_34); +x_37 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4(x_4, x_5, x_7, x_36, x_30, x_34, x_17, x_34, x_18, x_35, x_9, x_10, x_11, x_12, x_13, x_14, x_33); +lean_dec(x_34); +lean_dec(x_30); if (lean_obj_tag(x_37) == 0) { lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; uint8_t x_67; @@ -1872,60 +1901,59 @@ lean_inc(x_40); x_41 = lean_ctor_get(x_38, 1); lean_inc(x_41); lean_dec(x_38); -lean_inc(x_14); -x_42 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_14, x_15, x_39); +lean_inc(x_13); +x_42 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_13, x_14, x_39); x_43 = lean_ctor_get(x_42, 0); lean_inc(x_43); x_44 = lean_ctor_get(x_42, 1); lean_inc(x_44); lean_dec(x_42); -x_45 = lean_st_ref_get(x_15, x_44); +x_45 = lean_st_ref_get(x_14, x_44); x_46 = lean_ctor_get(x_45, 1); lean_inc(x_46); lean_dec(x_45); -x_47 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__3; +x_47 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__6; x_48 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_48, 0, x_43); lean_ctor_set(x_48, 1, x_47); -x_49 = lean_mk_syntax_ident(x_7); +x_49 = lean_mk_syntax_ident(x_6); x_50 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__23; x_51 = lean_array_push(x_50, x_48); x_52 = lean_array_push(x_51, x_49); -x_53 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__2; +x_53 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__5; x_54 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_54, 0, x_29); +lean_ctor_set(x_54, 0, x_28); lean_ctor_set(x_54, 1, x_53); lean_ctor_set(x_54, 2, x_52); -lean_inc(x_2); -x_55 = l_Array_append___rarg(x_2, x_40); +x_55 = l_Array_append___rarg(x_19, x_40); x_56 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__12; x_57 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_57, 0, x_29); +lean_ctor_set(x_57, 0, x_28); lean_ctor_set(x_57, 1, x_56); lean_ctor_set(x_57, 2, x_55); x_58 = lean_array_push(x_50, x_54); x_59 = lean_array_push(x_58, x_57); x_60 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__4; x_61 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_61, 0, x_29); +lean_ctor_set(x_61, 0, x_28); lean_ctor_set(x_61, 1, x_60); lean_ctor_set(x_61, 2, x_59); x_62 = lean_array_push(x_21, x_61); -x_63 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_14, x_15, x_46); +x_63 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_13, x_14, x_46); x_64 = lean_ctor_get(x_63, 0); lean_inc(x_64); x_65 = lean_ctor_get(x_63, 1); lean_inc(x_65); lean_dec(x_63); -x_66 = lean_st_ref_get(x_15, x_65); -lean_dec(x_15); +x_66 = lean_st_ref_get(x_14, x_65); +lean_dec(x_14); x_67 = !lean_is_exclusive(x_66); if (x_67 == 0) { lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; size_t x_72; size_t x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; x_68 = lean_ctor_get(x_66, 0); lean_dec(x_68); -x_69 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__6; +x_69 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__9; lean_inc(x_64); x_70 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_70, 0, x_64); @@ -1935,32 +1963,32 @@ x_72 = lean_usize_of_nat(x_71); lean_dec(x_71); x_73 = 0; x_74 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_72, x_73, x_62); -x_75 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__8; +x_75 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__11; x_76 = l_Lean_mkSepArray(x_74, x_75); lean_dec(x_74); -x_77 = l_Array_append___rarg(x_2, x_76); +x_77 = l_Array_append___rarg(x_19, x_76); x_78 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_78, 0, x_29); +lean_ctor_set(x_78, 0, x_28); lean_ctor_set(x_78, 1, x_56); lean_ctor_set(x_78, 2, x_77); x_79 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__3___closed__10; x_80 = lean_array_push(x_79, x_78); x_81 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_81, 0, x_29); +lean_ctor_set(x_81, 0, x_28); lean_ctor_set(x_81, 1, x_56); lean_ctor_set(x_81, 2, x_80); -x_82 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__9; +x_82 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__12; x_83 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_83, 0, x_64); lean_ctor_set(x_83, 1, x_82); -x_84 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__10; +x_84 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__13; x_85 = lean_array_push(x_84, x_70); x_86 = lean_array_push(x_85, x_81); x_87 = lean_array_push(x_86, x_83); x_88 = lean_array_push(x_87, x_41); -x_89 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__5; +x_89 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__8; x_90 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_90, 0, x_29); +lean_ctor_set(x_90, 0, x_28); lean_ctor_set(x_90, 1, x_89); lean_ctor_set(x_90, 2, x_88); lean_ctor_set(x_66, 0, x_90); @@ -1972,7 +2000,7 @@ lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; size x_91 = lean_ctor_get(x_66, 1); lean_inc(x_91); lean_dec(x_66); -x_92 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__6; +x_92 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__9; lean_inc(x_64); x_93 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_93, 0, x_64); @@ -1982,32 +2010,32 @@ x_95 = lean_usize_of_nat(x_94); lean_dec(x_94); x_96 = 0; x_97 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_95, x_96, x_62); -x_98 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__8; +x_98 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__11; x_99 = l_Lean_mkSepArray(x_97, x_98); lean_dec(x_97); -x_100 = l_Array_append___rarg(x_2, x_99); +x_100 = l_Array_append___rarg(x_19, x_99); x_101 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_101, 0, x_29); +lean_ctor_set(x_101, 0, x_28); lean_ctor_set(x_101, 1, x_56); lean_ctor_set(x_101, 2, x_100); x_102 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__3___closed__10; x_103 = lean_array_push(x_102, x_101); x_104 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_104, 0, x_29); +lean_ctor_set(x_104, 0, x_28); lean_ctor_set(x_104, 1, x_56); lean_ctor_set(x_104, 2, x_103); -x_105 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__9; +x_105 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__12; x_106 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_106, 0, x_64); lean_ctor_set(x_106, 1, x_105); -x_107 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__10; +x_107 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__13; x_108 = lean_array_push(x_107, x_93); x_109 = lean_array_push(x_108, x_104); x_110 = lean_array_push(x_109, x_106); x_111 = lean_array_push(x_110, x_41); -x_112 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__5; +x_112 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__8; x_113 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_113, 0, x_29); +lean_ctor_set(x_113, 0, x_28); lean_ctor_set(x_113, 1, x_112); lean_ctor_set(x_113, 2, x_111); x_114 = lean_alloc_ctor(0, 2, 0); @@ -2020,10 +2048,9 @@ else { uint8_t x_115; lean_dec(x_21); -lean_dec(x_15); lean_dec(x_14); -lean_dec(x_7); -lean_dec(x_2); +lean_dec(x_13); +lean_dec(x_6); x_115 = !lean_is_exclusive(x_37); if (x_115 == 0) { @@ -2045,166 +2072,161 @@ return x_118; } } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -if (lean_obj_tag(x_5) == 0) +if (lean_obj_tag(x_4) == 0) { -lean_object* x_14; -lean_dec(x_12); +lean_object* x_13; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_4); +lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_6); -lean_ctor_set(x_14, 1, x_13); -return x_14; +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_5); +lean_ctor_set(x_13, 1, x_12); +return x_13; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_15 = lean_ctor_get(x_5, 0); +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_14 = lean_ctor_get(x_4, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_4, 1); lean_inc(x_15); -x_16 = lean_ctor_get(x_5, 1); +lean_dec(x_4); +x_16 = lean_ctor_get(x_5, 0); lean_inc(x_16); -lean_dec(x_5); -x_17 = lean_ctor_get(x_6, 0); +x_17 = lean_ctor_get(x_5, 1); lean_inc(x_17); -x_18 = lean_ctor_get(x_6, 1); -lean_inc(x_18); -lean_dec(x_6); -lean_inc(x_7); -lean_inc(x_15); -x_19 = l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__1(x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_19) == 0) +lean_dec(x_5); +lean_inc(x_6); +lean_inc(x_14); +x_18 = l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__1(x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_20 = lean_ctor_get(x_19, 0); +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_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); lean_inc(x_20); -x_21 = lean_ctor_get(x_19, 1); +lean_dec(x_18); +x_21 = lean_ctor_get(x_19, 0); lean_inc(x_21); -lean_dec(x_19); -x_22 = lean_ctor_get(x_20, 0); +x_22 = lean_ctor_get(x_21, 2); lean_inc(x_22); -x_23 = lean_ctor_get(x_22, 2); -lean_inc(x_23); -lean_dec(x_22); -lean_inc(x_4); -lean_inc(x_1); -lean_inc(x_18); +lean_dec(x_21); lean_inc(x_3); +lean_inc(x_1); +lean_inc(x_17); lean_inc(x_2); -x_24 = lean_alloc_closure((void*)(l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___boxed), 16, 7); -lean_closure_set(x_24, 0, x_2); -lean_closure_set(x_24, 1, x_3); -lean_closure_set(x_24, 2, x_18); -lean_closure_set(x_24, 3, x_20); -lean_closure_set(x_24, 4, x_1); -lean_closure_set(x_24, 5, x_4); -lean_closure_set(x_24, 6, x_15); -lean_inc(x_12); +x_23 = lean_alloc_closure((void*)(l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___boxed), 15, 6); +lean_closure_set(x_23, 0, x_2); +lean_closure_set(x_23, 1, x_17); +lean_closure_set(x_23, 2, x_19); +lean_closure_set(x_23, 3, x_1); +lean_closure_set(x_23, 4, x_3); +lean_closure_set(x_23, 5, x_14); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -x_25 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Elab_Deriving_mkInductArgNames___spec__2___rarg(x_23, x_24, x_7, x_8, x_9, x_10, x_11, x_12, x_21); -if (lean_obj_tag(x_25) == 0) +lean_inc(x_6); +x_24 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Elab_Deriving_mkInductArgNames___spec__2___rarg(x_22, x_23, x_6, x_7, x_8, x_9, x_10, x_11, x_20); +if (lean_obj_tag(x_24) == 0) { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_26 = lean_ctor_get(x_25, 0); +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 1); -lean_inc(x_27); -lean_dec(x_25); -x_28 = lean_array_push(x_17, x_26); -x_29 = lean_unsigned_to_nat(1u); -x_30 = lean_nat_add(x_18, x_29); -lean_dec(x_18); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_28); -lean_ctor_set(x_31, 1, x_30); -x_5 = x_16; -x_6 = x_31; -x_13 = x_27; +lean_dec(x_24); +x_27 = lean_array_push(x_16, x_25); +x_28 = lean_unsigned_to_nat(1u); +x_29 = lean_nat_add(x_17, x_28); +lean_dec(x_17); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_27); +lean_ctor_set(x_30, 1, x_29); +x_4 = x_15; +x_5 = x_30; +x_12 = x_26; goto _start; } else { -uint8_t x_33; -lean_dec(x_18); +uint8_t x_32; lean_dec(x_17); lean_dec(x_16); -lean_dec(x_12); +lean_dec(x_15); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_4); +lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_33 = !lean_is_exclusive(x_25); -if (x_33 == 0) +x_32 = !lean_is_exclusive(x_24); +if (x_32 == 0) { -return x_25; +return x_24; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_25, 0); -x_35 = lean_ctor_get(x_25, 1); -lean_inc(x_35); +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_24, 0); +x_34 = lean_ctor_get(x_24, 1); lean_inc(x_34); -lean_dec(x_25); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -return x_36; +lean_inc(x_33); +lean_dec(x_24); +x_35 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; } } } else { -uint8_t x_37; -lean_dec(x_18); +uint8_t x_36; lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); -lean_dec(x_12); +lean_dec(x_14); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_4); +lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_37 = !lean_is_exclusive(x_19); -if (x_37 == 0) +x_36 = !lean_is_exclusive(x_18); +if (x_36 == 0) { -return x_19; +return x_18; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_19, 0); -x_39 = lean_ctor_get(x_19, 1); -lean_inc(x_39); +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_18, 0); +x_38 = lean_ctor_get(x_18, 1); lean_inc(x_38); -lean_dec(x_19); -x_40 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -return x_40; +lean_inc(x_37); +lean_dec(x_18); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; } } } @@ -2213,17 +2235,8 @@ return x_40; static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(0u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___closed__2() { -_start: -{ lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___closed__1; +x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__24; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2234,7 +2247,7 @@ return x_3; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts(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_object* x_9) { _start: { -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_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; x_10 = lean_ctor_get(x_2, 3); lean_inc(x_10); x_11 = l_List_redLength___rarg(x_10); @@ -2244,59 +2257,58 @@ x_13 = l_List_toArrayAux___rarg(x_10, x_12); x_14 = lean_ctor_get(x_2, 4); lean_inc(x_14); x_15 = l_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___closed__1; -x_16 = l_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___closed__2; -x_17 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5(x_1, x_2, x_15, x_13, x_14, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_17) == 0) +x_16 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5(x_1, x_2, x_13, x_14, x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_16) == 0) { -uint8_t x_18; -x_18 = !lean_is_exclusive(x_17); -if (x_18 == 0) +uint8_t x_17; +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_17, 0); -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -lean_dec(x_19); -lean_ctor_set(x_17, 0, x_20); -return x_17; +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_16, 0); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +lean_dec(x_18); +lean_ctor_set(x_16, 0, x_19); +return x_16; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_21 = lean_ctor_get(x_17, 0); -x_22 = lean_ctor_get(x_17, 1); -lean_inc(x_22); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_20 = lean_ctor_get(x_16, 0); +x_21 = lean_ctor_get(x_16, 1); lean_inc(x_21); -lean_dec(x_17); -x_23 = lean_ctor_get(x_21, 0); -lean_inc(x_23); -lean_dec(x_21); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_22); -return x_24; +lean_inc(x_20); +lean_dec(x_16); +x_22 = lean_ctor_get(x_20, 0); +lean_inc(x_22); +lean_dec(x_20); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_21); +return x_23; } } else { -uint8_t x_25; -x_25 = !lean_is_exclusive(x_17); -if (x_25 == 0) +uint8_t x_24; +x_24 = !lean_is_exclusive(x_16); +if (x_24 == 0) { -return x_17; +return x_16; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_17, 0); -x_27 = lean_ctor_get(x_17, 1); -lean_inc(x_27); +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_16, 0); +x_26 = lean_ctor_get(x_16, 1); lean_inc(x_26); -lean_dec(x_17); -x_28 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -return x_28; +lean_inc(x_25); +lean_dec(x_16); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; } } } @@ -2382,22 +2394,47 @@ lean_dec(x_8); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); return x_18; } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___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, 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, lean_object* x_16) { +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___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, 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) { _start: { -lean_object* x_17; -x_17 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -lean_dec(x_11); +lean_object* x_16; +x_16 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_6); +lean_dec(x_7); lean_dec(x_5); -return x_17; +lean_dec(x_4); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_Hashable_mkMatch___spec__1(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = 1; +x_9 = lean_usize_add(x_2, x_8); +x_10 = lean_array_uset(x_7, x_2, x_5); +x_2 = x_9; +x_3 = x_10; +goto _start; +} } } static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkMatch___closed__1() { @@ -2421,26 +2458,12 @@ return x_3; static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkMatch___closed__3() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = lean_box(2); -x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__12; -x_3 = l_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___closed__1; -x_4 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_4, 0, x_1); -lean_ctor_set(x_4, 1, x_2); -lean_ctor_set(x_4, 2, x_3); -return x_4; -} -} -static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkMatch___closed__4() { -_start: -{ lean_object* x_1; x_1 = lean_mk_string_from_bytes("with", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkMatch___closed__5() { +static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkMatch___closed__4() { _start: { lean_object* x_1; @@ -2448,17 +2471,17 @@ x_1 = lean_mk_string_from_bytes("matchAlts", 9); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkMatch___closed__6() { +static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkMatch___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__3___closed__6; -x_2 = l_Lean_Elab_Deriving_Hashable_mkMatch___closed__5; +x_2 = l_Lean_Elab_Deriving_Hashable_mkMatch___closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkMatch___closed__7() { +static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkMatch___closed__6() { _start: { lean_object* x_1; lean_object* x_2; @@ -2513,11 +2536,11 @@ x_25 = lean_array_get_size(x_12); x_26 = lean_usize_of_nat(x_25); lean_dec(x_25); x_27 = 0; -x_28 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_26, x_27, x_12); -x_29 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__8; +x_28 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_Hashable_mkMatch___spec__1(x_26, x_27, x_12); +x_29 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__11; x_30 = l_Lean_mkSepArray(x_28, x_29); lean_dec(x_28); -x_31 = l_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___closed__1; +x_31 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__24; x_32 = l_Array_append___rarg(x_31, x_30); x_33 = lean_box(2); x_34 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__12; @@ -2525,7 +2548,7 @@ x_35 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_35, 0, x_33); lean_ctor_set(x_35, 1, x_34); lean_ctor_set(x_35, 2, x_32); -x_36 = l_Lean_Elab_Deriving_Hashable_mkMatch___closed__4; +x_36 = l_Lean_Elab_Deriving_Hashable_mkMatch___closed__3; x_37 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_37, 0, x_18); lean_ctor_set(x_37, 1, x_36); @@ -2536,14 +2559,14 @@ lean_ctor_set(x_39, 1, x_34); lean_ctor_set(x_39, 2, x_38); x_40 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__3___closed__10; x_41 = lean_array_push(x_40, x_39); -x_42 = l_Lean_Elab_Deriving_Hashable_mkMatch___closed__6; +x_42 = l_Lean_Elab_Deriving_Hashable_mkMatch___closed__5; x_43 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_43, 0, x_33); lean_ctor_set(x_43, 1, x_42); lean_ctor_set(x_43, 2, x_41); -x_44 = l_Lean_Elab_Deriving_Hashable_mkMatch___closed__7; +x_44 = l_Lean_Elab_Deriving_Hashable_mkMatch___closed__6; x_45 = lean_array_push(x_44, x_24); -x_46 = l_Lean_Elab_Deriving_Hashable_mkMatch___closed__3; +x_46 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__25; x_47 = lean_array_push(x_45, x_46); x_48 = lean_array_push(x_47, x_46); x_49 = lean_array_push(x_48, x_35); @@ -2572,11 +2595,11 @@ x_57 = lean_array_get_size(x_12); x_58 = lean_usize_of_nat(x_57); lean_dec(x_57); x_59 = 0; -x_60 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_58, x_59, x_12); -x_61 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__8; +x_60 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_Hashable_mkMatch___spec__1(x_58, x_59, x_12); +x_61 = l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__11; x_62 = l_Lean_mkSepArray(x_60, x_61); lean_dec(x_60); -x_63 = l_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___closed__1; +x_63 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__24; x_64 = l_Array_append___rarg(x_63, x_62); x_65 = lean_box(2); x_66 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__12; @@ -2584,7 +2607,7 @@ x_67 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_67, 0, x_65); lean_ctor_set(x_67, 1, x_66); lean_ctor_set(x_67, 2, x_64); -x_68 = l_Lean_Elab_Deriving_Hashable_mkMatch___closed__4; +x_68 = l_Lean_Elab_Deriving_Hashable_mkMatch___closed__3; x_69 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_69, 0, x_18); lean_ctor_set(x_69, 1, x_68); @@ -2595,14 +2618,14 @@ lean_ctor_set(x_71, 1, x_66); lean_ctor_set(x_71, 2, x_70); x_72 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__3___closed__10; x_73 = lean_array_push(x_72, x_71); -x_74 = l_Lean_Elab_Deriving_Hashable_mkMatch___closed__6; +x_74 = l_Lean_Elab_Deriving_Hashable_mkMatch___closed__5; x_75 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_75, 0, x_65); lean_ctor_set(x_75, 1, x_74); lean_ctor_set(x_75, 2, x_73); -x_76 = l_Lean_Elab_Deriving_Hashable_mkMatch___closed__7; +x_76 = l_Lean_Elab_Deriving_Hashable_mkMatch___closed__6; x_77 = lean_array_push(x_76, x_56); -x_78 = l_Lean_Elab_Deriving_Hashable_mkMatch___closed__3; +x_78 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__25; x_79 = lean_array_push(x_77, x_78); x_80 = lean_array_push(x_79, x_78); x_81 = lean_array_push(x_80, x_67); @@ -2646,6 +2669,18 @@ return x_90; } } } +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_Hashable_mkMatch___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_Hashable_mkMatch___spec__1(x_4, x_5, x_3); +return x_6; +} +} static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__1() { _start: { @@ -2722,8 +2757,8 @@ static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_Hashable_mkMatch___closed__7; -x_2 = l_Lean_Elab_Deriving_Hashable_mkMatch___closed__3; +x_1 = l_Lean_Elab_Deriving_Hashable_mkMatch___closed__6; +x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__25; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -2733,7 +2768,7 @@ static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__9; -x_2 = l_Lean_Elab_Deriving_Hashable_mkMatch___closed__3; +x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__25; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -3029,7 +3064,7 @@ lean_ctor_set(x_42, 1, x_41); lean_ctor_set(x_42, 2, x_40); x_43 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__10; x_44 = lean_array_push(x_43, x_42); -x_45 = l_Lean_Elab_Deriving_Hashable_mkMatch___closed__3; +x_45 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__25; x_46 = lean_array_push(x_44, x_45); x_47 = lean_array_push(x_46, x_45); x_48 = lean_array_push(x_47, x_45); @@ -3052,7 +3087,7 @@ x_58 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_58, 0, x_37); lean_ctor_set(x_58, 1, x_57); lean_ctor_set(x_58, 2, x_56); -x_59 = l_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___closed__1; +x_59 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__24; x_60 = l_Array_append___rarg(x_59, x_23); x_61 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_61, 0, x_37); @@ -3096,7 +3131,7 @@ x_79 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__28; x_80 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_80, 0, x_25); lean_ctor_set(x_80, 1, x_79); -x_81 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__25; +x_81 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__27; x_82 = lean_array_push(x_81, x_80); x_83 = lean_array_push(x_82, x_21); x_84 = lean_array_push(x_83, x_45); @@ -3161,7 +3196,7 @@ lean_ctor_set(x_114, 1, x_113); lean_ctor_set(x_114, 2, x_112); x_115 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__10; x_116 = lean_array_push(x_115, x_114); -x_117 = l_Lean_Elab_Deriving_Hashable_mkMatch___closed__3; +x_117 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__25; x_118 = lean_array_push(x_116, x_117); x_119 = lean_array_push(x_118, x_117); x_120 = lean_array_push(x_119, x_117); @@ -3184,7 +3219,7 @@ x_130 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_130, 0, x_109); lean_ctor_set(x_130, 1, x_129); lean_ctor_set(x_130, 2, x_128); -x_131 = l_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___closed__1; +x_131 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__24; x_132 = l_Array_append___rarg(x_131, x_23); x_133 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_133, 0, x_109); @@ -3228,7 +3263,7 @@ x_151 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__28; x_152 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_152, 0, x_25); lean_ctor_set(x_152, 1, x_151); -x_153 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__25; +x_153 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__27; x_154 = lean_array_push(x_153, x_152); x_155 = lean_array_push(x_154, x_21); x_156 = lean_array_push(x_155, x_117); @@ -3332,7 +3367,7 @@ lean_ctor_set(x_202, 1, x_194); lean_ctor_set(x_202, 2, x_201); x_203 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__10; x_204 = lean_array_push(x_203, x_195); -x_205 = l_Lean_Elab_Deriving_Hashable_mkMatch___closed__3; +x_205 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__25; x_206 = lean_array_push(x_204, x_205); x_207 = lean_array_push(x_206, x_205); x_208 = lean_array_push(x_207, x_202); @@ -3355,7 +3390,7 @@ x_218 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_218, 0, x_190); lean_ctor_set(x_218, 1, x_217); lean_ctor_set(x_218, 2, x_216); -x_219 = l_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___closed__1; +x_219 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__24; x_220 = l_Array_append___rarg(x_219, x_176); x_221 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_221, 0, x_190); @@ -3399,7 +3434,7 @@ x_239 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__28; x_240 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_240, 0, x_178); lean_ctor_set(x_240, 1, x_239); -x_241 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__25; +x_241 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__27; x_242 = lean_array_push(x_241, x_240); x_243 = lean_array_push(x_242, x_174); x_244 = lean_array_push(x_243, x_205); @@ -3480,7 +3515,7 @@ lean_ctor_set(x_281, 1, x_273); lean_ctor_set(x_281, 2, x_280); x_282 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__10; x_283 = lean_array_push(x_282, x_274); -x_284 = l_Lean_Elab_Deriving_Hashable_mkMatch___closed__3; +x_284 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__25; x_285 = lean_array_push(x_283, x_284); x_286 = lean_array_push(x_285, x_284); x_287 = lean_array_push(x_286, x_281); @@ -3503,7 +3538,7 @@ x_297 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_297, 0, x_269); lean_ctor_set(x_297, 1, x_296); lean_ctor_set(x_297, 2, x_295); -x_298 = l_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___closed__1; +x_298 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__24; x_299 = l_Array_append___rarg(x_298, x_176); x_300 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_300, 0, x_269); @@ -3547,7 +3582,7 @@ x_318 = l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__28; x_319 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_319, 0, x_178); lean_ctor_set(x_319, 1, x_318); -x_320 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__25; +x_320 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__27; x_321 = lean_array_push(x_320, x_319); x_322 = lean_array_push(x_321, x_174); x_323 = lean_array_push(x_322, x_284); @@ -3796,7 +3831,7 @@ static lean_object* _init_l_Lean_Elab_Deriving_Hashable_mkHashFuncs___closed__4( lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__12; -x_3 = l_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___closed__1; +x_3 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__24; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3823,7 +3858,7 @@ x_10 = lean_array_get_size(x_9); lean_dec(x_9); x_11 = lean_unsigned_to_nat(0u); x_12 = lean_unsigned_to_nat(1u); -x_13 = l_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___closed__1; +x_13 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__24; lean_inc(x_7); lean_inc(x_6); lean_inc(x_10); @@ -4925,7 +4960,7 @@ lean_dec(x_1); return x_9; } } -static lean_object* _init_l_Lean_Elab_Deriving_Hashable_initFn____x40_Lean_Elab_Deriving_Hashable___hyg_2199____closed__1() { +static lean_object* _init_l_Lean_Elab_Deriving_Hashable_initFn____x40_Lean_Elab_Deriving_Hashable___hyg_2330____closed__1() { _start: { lean_object* x_1; @@ -4933,12 +4968,12 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Deriving_Hashable_mkHashableHandler return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Hashable_initFn____x40_Lean_Elab_Deriving_Hashable___hyg_2199_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Hashable_initFn____x40_Lean_Elab_Deriving_Hashable___hyg_2330_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; x_2 = l_Lean_Elab_Deriving_Hashable_mkHashableHeader___closed__2; -x_3 = l_Lean_Elab_Deriving_Hashable_initFn____x40_Lean_Elab_Deriving_Hashable___hyg_2199____closed__1; +x_3 = l_Lean_Elab_Deriving_Hashable_initFn____x40_Lean_Elab_Deriving_Hashable___hyg_2330____closed__1; x_4 = l_Lean_Elab_registerBuiltinDerivingHandler(x_2, x_3, x_1); if (lean_obj_tag(x_4) == 0) { @@ -5077,6 +5112,10 @@ l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4 lean_mark_persistent(l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__24); l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__25 = _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__25(); lean_mark_persistent(l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__25); +l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__26 = _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__26(); +lean_mark_persistent(l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__26); +l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__27 = _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__27(); +lean_mark_persistent(l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__27); l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__1 = _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__1(); lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__1); l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__2 = _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__2(); @@ -5097,10 +5136,14 @@ l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___la lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__9); l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__10 = _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__10(); lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__10); +l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__11 = _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__11(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__11); +l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__12 = _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__12(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__12); +l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__13 = _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__13(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__13); l_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___closed__1 = _init_l_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___closed__1(); lean_mark_persistent(l_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___closed__1); -l_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___closed__2 = _init_l_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___closed__2(); -lean_mark_persistent(l_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___closed__2); l_Lean_Elab_Deriving_Hashable_mkMatch___closed__1 = _init_l_Lean_Elab_Deriving_Hashable_mkMatch___closed__1(); lean_mark_persistent(l_Lean_Elab_Deriving_Hashable_mkMatch___closed__1); l_Lean_Elab_Deriving_Hashable_mkMatch___closed__2 = _init_l_Lean_Elab_Deriving_Hashable_mkMatch___closed__2(); @@ -5113,8 +5156,6 @@ l_Lean_Elab_Deriving_Hashable_mkMatch___closed__5 = _init_l_Lean_Elab_Deriving_H lean_mark_persistent(l_Lean_Elab_Deriving_Hashable_mkMatch___closed__5); l_Lean_Elab_Deriving_Hashable_mkMatch___closed__6 = _init_l_Lean_Elab_Deriving_Hashable_mkMatch___closed__6(); lean_mark_persistent(l_Lean_Elab_Deriving_Hashable_mkMatch___closed__6); -l_Lean_Elab_Deriving_Hashable_mkMatch___closed__7 = _init_l_Lean_Elab_Deriving_Hashable_mkMatch___closed__7(); -lean_mark_persistent(l_Lean_Elab_Deriving_Hashable_mkMatch___closed__7); l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__1 = _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__1(); lean_mark_persistent(l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__1); l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__2 = _init_l_Lean_Elab_Deriving_Hashable_mkAuxFunction___closed__2(); @@ -5207,9 +5248,9 @@ l___private_Lean_Elab_Deriving_Hashable_0__Lean_Elab_Deriving_Hashable_mkHashabl lean_mark_persistent(l___private_Lean_Elab_Deriving_Hashable_0__Lean_Elab_Deriving_Hashable_mkHashableInstanceCmds___closed__9); l___private_Lean_Elab_Deriving_Hashable_0__Lean_Elab_Deriving_Hashable_mkHashableInstanceCmds___closed__10 = _init_l___private_Lean_Elab_Deriving_Hashable_0__Lean_Elab_Deriving_Hashable_mkHashableInstanceCmds___closed__10(); lean_mark_persistent(l___private_Lean_Elab_Deriving_Hashable_0__Lean_Elab_Deriving_Hashable_mkHashableInstanceCmds___closed__10); -l_Lean_Elab_Deriving_Hashable_initFn____x40_Lean_Elab_Deriving_Hashable___hyg_2199____closed__1 = _init_l_Lean_Elab_Deriving_Hashable_initFn____x40_Lean_Elab_Deriving_Hashable___hyg_2199____closed__1(); -lean_mark_persistent(l_Lean_Elab_Deriving_Hashable_initFn____x40_Lean_Elab_Deriving_Hashable___hyg_2199____closed__1); -res = l_Lean_Elab_Deriving_Hashable_initFn____x40_Lean_Elab_Deriving_Hashable___hyg_2199_(lean_io_mk_world()); +l_Lean_Elab_Deriving_Hashable_initFn____x40_Lean_Elab_Deriving_Hashable___hyg_2330____closed__1 = _init_l_Lean_Elab_Deriving_Hashable_initFn____x40_Lean_Elab_Deriving_Hashable___hyg_2330____closed__1(); +lean_mark_persistent(l_Lean_Elab_Deriving_Hashable_initFn____x40_Lean_Elab_Deriving_Hashable___hyg_2330____closed__1); +res = l_Lean_Elab_Deriving_Hashable_initFn____x40_Lean_Elab_Deriving_Hashable___hyg_2330_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/Deriving/Inhabited.c b/stage0/stdlib/Lean/Elab/Deriving/Inhabited.c index 00afe89c2ba..67f7af5b96b 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/Inhabited.c +++ b/stage0/stdlib/Lean/Elab/Deriving/Inhabited.c @@ -30,6 +30,7 @@ static lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInha static lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__10; lean_object* lean_mk_empty_array_with_capacity(lean_object*); static lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__20; +lean_object* l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__14(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstance___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_addLocalInstancesForParamsAux___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Parser_Term_instBinder; @@ -47,13 +48,14 @@ lean_object* l_Lean_Meta_mkAppM(lean_object*, lean_object*, lean_object*, lean_o LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmd_x3f___spec__1___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmd_x3f___lambda__3___closed__2; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__13; -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, 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_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__23; LEAN_EXPORT lean_object* l_Lean_getConstInfoInduct___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstance___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmd_x3f___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmd_x3f___lambda__3___closed__1; lean_object* lean_environment_find(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__42; +static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__27; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__9; static lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmd_x3f___lambda__3___closed__3; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__17; @@ -62,6 +64,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_m static lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__7; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__37; +static lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__49; lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); @@ -84,13 +87,13 @@ static lean_object* l_Lean_getConstInfoCtor___at___private_Lean_Elab_Deriving_In static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__4; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__18; LEAN_EXPORT lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_instBinderF; +static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__28; lean_object* lean_nat_add(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Inhabited___hyg_2506____closed__1; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__8___closed__1; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__14; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__5; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmd_x3f___spec__1___closed__4; -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, 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_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmd_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__31; LEAN_EXPORT lean_object* l_Lean_ForEachExpr_visit___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_collectUsedLocalsInsts___spec__4___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -98,7 +101,7 @@ lean_object* l_Std_RBNode_setBlack___rarg(lean_object*); LEAN_EXPORT lean_object* l_Std_RBNode_ins___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_collectUsedLocalsInsts___spec__3(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__11; lean_object* l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, 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_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__2; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmd_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -115,10 +118,11 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_i LEAN_EXPORT lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_addLocalInstancesForParams___boxed(lean_object*, 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___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmd_x3f(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmd_x3f___lambda__2___closed__2; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Inhabited___hyg_2570____closed__1; lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__6; lean_object* l_Lean_Expr_fvarId_x21(lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__8___boxed(lean_object**); lean_object* l_Std_mkHashMapImp___rarg(lean_object*); static lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_addLocalInstancesForParamsAux___rarg___lambda__2___closed__7; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__7___closed__1; @@ -157,6 +161,7 @@ static lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInha static lean_object* l_Lean_getConstInfoCtor___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__3___closed__1; static lean_object* l_Lean_getConstInfoInduct___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__1___closed__2; static lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstance___closed__3; +static lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__47; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_addLocalInstancesForParamsAux___rarg___closed__2; lean_object* l_Lean_Elab_Command_elabCommand(lean_object*, lean_object*, lean_object*, lean_object*); @@ -188,7 +193,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_mkInhabitedInstanceHandler(lean_object*, le static lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_addLocalInstancesForParamsAux___rarg___lambda__2___closed__9; static lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__36; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmd_x3f___spec__1___lambda__2___closed__7; -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_isInductive___at_Lean_Elab_mkInhabitedInstanceHandler___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmd_x3f___spec__1___closed__3; static lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__19; @@ -235,7 +240,7 @@ static lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInha static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__19; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmd_x3f___spec__1___lambda__2___closed__5; lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, 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_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__33; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__11; LEAN_EXPORT lean_object* l_Std_RBNode_insert___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_addLocalInstancesForParamsAux___spec__1(lean_object*, lean_object*, lean_object*); @@ -263,7 +268,9 @@ static lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInha LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmd_x3f___spec__7___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmd_x3f___spec__1___lambda__2___closed__6; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__12; +static lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__48; static lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_addLocalInstancesForParamsAux___rarg___lambda__2___closed__8; +static lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__50; LEAN_EXPORT lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_collectUsedLocalsInsts___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfoInduct___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__1___closed__1; lean_object* l_Lean_indentExpr(lean_object*); @@ -272,7 +279,6 @@ static lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInha LEAN_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstance___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__6; lean_object* l_Lean_mkConst(lean_object*, lean_object*); -lean_object* l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__10(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_RBNode_findCore___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__5___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_mkInhabitedInstanceHandler___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ForEachExpr_visit___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_collectUsedLocalsInsts___spec__4___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -284,7 +290,7 @@ LEAN_EXPORT lean_object* l_List_mapTRAux___at___private_Lean_Elab_Deriving_Inhab LEAN_EXPORT lean_object* l_Std_RBNode_find___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_collectUsedLocalsInsts___spec__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__46; lean_object* l_runST___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Inhabited___hyg_2506_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Inhabited___hyg_2570_(lean_object*); static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__26; lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); @@ -7915,12 +7921,35 @@ return x_3; static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__14() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(2); +x_2 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__13; +x_3 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__14; +x_4 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__16() { +_start: +{ lean_object* x_1; x_1 = lean_mk_string_from_bytes("}", 1); return x_1; } } -static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__15() { +static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__17() { _start: { lean_object* x_1; lean_object* x_2; @@ -7929,7 +7958,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__16() { +static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__18() { _start: { lean_object* x_1; @@ -7937,17 +7966,17 @@ x_1 = lean_mk_string_from_bytes("instBinder", 10); return x_1; } } -static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__17() { +static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__8; -x_2 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__16; +x_2 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__18; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__18() { +static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__20() { _start: { lean_object* x_1; @@ -7955,7 +7984,7 @@ x_1 = lean_mk_string_from_bytes("[", 1); return x_1; } } -static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__19() { +static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__21() { _start: { lean_object* x_1; @@ -7963,17 +7992,17 @@ x_1 = lean_mk_string_from_bytes("app", 3); return x_1; } } -static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__20() { +static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__8; -x_2 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__19; +x_2 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__21; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__21() { +static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__23() { _start: { lean_object* x_1; lean_object* x_2; @@ -7982,13 +8011,13 @@ x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__22() { +static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_addLocalInstancesForParamsAux___rarg___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__21; +x_3 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__23; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -7996,7 +8025,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__23() { +static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -8008,19 +8037,19 @@ lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__24() { +static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__23; +x_2 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__25; x_3 = lean_alloc_ctor(1, 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_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__25() { +static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__27() { _start: { lean_object* x_1; lean_object* x_2; @@ -8029,7 +8058,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__26() { +static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__28() { _start: { lean_object* x_1; @@ -8037,206 +8066,198 @@ x_1 = lean_mk_string_from_bytes("]", 1); return x_1; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -uint8_t x_15; -x_15 = lean_nat_dec_le(x_5, x_4); -if (x_15 == 0) +uint8_t x_14; +x_14 = lean_nat_dec_le(x_4, x_3); +if (x_14 == 0) { -lean_object* x_16; uint8_t x_17; -x_16 = lean_unsigned_to_nat(0u); -x_17 = lean_nat_dec_eq(x_3, x_16); -if (x_17 == 0) +lean_object* x_15; uint8_t x_16; +x_15 = lean_unsigned_to_nat(0u); +x_16 = lean_nat_dec_eq(x_2, x_15); +if (x_16 == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_18 = lean_unsigned_to_nat(1u); -x_19 = lean_nat_sub(x_3, x_18); -lean_dec(x_3); -x_26 = lean_ctor_get(x_7, 0); +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_17 = lean_unsigned_to_nat(1u); +x_18 = lean_nat_sub(x_2, x_17); +lean_dec(x_2); +x_25 = lean_ctor_get(x_6, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_6, 1); lean_inc(x_26); -x_27 = lean_ctor_get(x_7, 1); -lean_inc(x_27); -lean_dec(x_7); -x_28 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__2; -x_29 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_28, x_12, x_13, x_14); -x_30 = lean_ctor_get(x_29, 0); +lean_dec(x_6); +x_27 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__2; +x_28 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_27, x_11, x_12, x_13); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); lean_inc(x_30); -x_31 = lean_ctor_get(x_29, 1); +lean_dec(x_28); +x_31 = lean_mk_syntax_ident(x_29); lean_inc(x_31); -lean_dec(x_29); -x_32 = lean_mk_syntax_ident(x_30); -lean_inc(x_32); -x_33 = lean_array_push(x_27, x_32); -lean_inc(x_12); -x_34 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___spec__1___rarg(x_12, x_13, x_31); -x_35 = lean_ctor_get(x_34, 0); -lean_inc(x_35); -x_36 = lean_ctor_get(x_34, 1); -lean_inc(x_36); -lean_dec(x_34); -x_37 = lean_st_ref_get(x_13, x_36); -x_38 = lean_ctor_get(x_37, 1); -lean_inc(x_38); -lean_dec(x_37); -x_39 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__11; +x_32 = lean_array_push(x_26, x_31); +lean_inc(x_11); +x_33 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___spec__1___rarg(x_11, x_12, x_30); +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_33, 1); lean_inc(x_35); -x_40 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_40, 0, x_35); -lean_ctor_set(x_40, 1, x_39); -x_41 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_addLocalInstancesForParamsAux___rarg___closed__3; -x_42 = lean_array_push(x_41, x_32); -x_43 = lean_box(2); -x_44 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__13; -x_45 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_45, 0, x_43); -lean_ctor_set(x_45, 1, x_44); -lean_ctor_set(x_45, 2, x_42); -lean_inc(x_2); -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_43); -lean_ctor_set(x_46, 1, x_44); -lean_ctor_set(x_46, 2, x_2); -x_47 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__14; -x_48 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_48, 0, x_35); -lean_ctor_set(x_48, 1, x_47); -x_49 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__15; -x_50 = lean_array_push(x_49, x_40); -lean_inc(x_45); -x_51 = lean_array_push(x_50, x_45); -lean_inc(x_46); +lean_dec(x_33); +x_36 = lean_st_ref_get(x_12, x_35); +x_37 = lean_ctor_get(x_36, 1); +lean_inc(x_37); +lean_dec(x_36); +x_38 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__11; +lean_inc(x_34); +x_39 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_39, 0, x_34); +lean_ctor_set(x_39, 1, x_38); +x_40 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_addLocalInstancesForParamsAux___rarg___closed__3; +x_41 = lean_array_push(x_40, x_31); +x_42 = lean_box(2); +x_43 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__13; +x_44 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_44, 0, x_42); +lean_ctor_set(x_44, 1, x_43); +lean_ctor_set(x_44, 2, x_41); +x_45 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__16; +x_46 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_46, 0, x_34); +lean_ctor_set(x_46, 1, x_45); +x_47 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__17; +x_48 = lean_array_push(x_47, x_39); +lean_inc(x_44); +x_49 = lean_array_push(x_48, x_44); +x_50 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__15; +x_51 = lean_array_push(x_49, x_50); x_52 = lean_array_push(x_51, x_46); -x_53 = lean_array_push(x_52, x_48); -x_54 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__10; -x_55 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_55, 0, x_43); -lean_ctor_set(x_55, 1, x_54); -lean_ctor_set(x_55, 2, x_53); -x_56 = lean_array_push(x_26, x_55); -x_57 = l_Std_RBNode_findCore___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__5(x_1, x_4); -if (lean_obj_tag(x_57) == 0) +x_53 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__10; +x_54 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_54, 0, x_42); +lean_ctor_set(x_54, 1, x_53); +lean_ctor_set(x_54, 2, x_52); +x_55 = lean_array_push(x_25, x_54); +x_56 = l_Std_RBNode_findCore___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__5(x_1, x_3); +if (lean_obj_tag(x_56) == 0) { -lean_object* x_58; lean_object* x_59; -lean_dec(x_46); -lean_dec(x_45); -x_58 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_58, 0, x_56); -lean_ctor_set(x_58, 1, x_33); -x_59 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_59, 0, x_58); -x_20 = x_59; -x_21 = x_38; -goto block_25; +lean_object* x_57; lean_object* x_58; +lean_dec(x_44); +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_32); +x_58 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_58, 0, x_57); +x_19 = x_58; +x_20 = x_37; +goto block_24; } 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; -lean_dec(x_57); -lean_inc(x_12); -x_60 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___spec__1___rarg(x_12, x_13, x_38); -x_61 = lean_ctor_get(x_60, 0); +lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; +lean_dec(x_56); +lean_inc(x_11); +x_59 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___spec__1___rarg(x_11, x_12, x_37); +x_60 = lean_ctor_get(x_59, 0); +lean_inc(x_60); +x_61 = lean_ctor_get(x_59, 1); lean_inc(x_61); -x_62 = lean_ctor_get(x_60, 1); +lean_dec(x_59); +x_62 = lean_ctor_get(x_11, 10); lean_inc(x_62); -lean_dec(x_60); -x_63 = lean_ctor_get(x_12, 10); -lean_inc(x_63); -x_64 = lean_st_ref_get(x_13, x_62); -x_65 = lean_ctor_get(x_64, 0); +x_63 = lean_st_ref_get(x_12, x_61); +x_64 = lean_ctor_get(x_63, 0); +lean_inc(x_64); +x_65 = lean_ctor_get(x_63, 1); lean_inc(x_65); -x_66 = lean_ctor_get(x_64, 1); +lean_dec(x_63); +x_66 = lean_ctor_get(x_64, 0); lean_inc(x_66); lean_dec(x_64); -x_67 = lean_ctor_get(x_65, 0); -lean_inc(x_67); -lean_dec(x_65); -x_68 = lean_environment_main_module(x_67); -x_69 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__18; -lean_inc(x_61); -x_70 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_70, 0, x_61); -lean_ctor_set(x_70, 1, x_69); -x_71 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_addLocalInstancesForParamsAux___rarg___closed__2; -x_72 = l_Lean_addMacroScope(x_68, x_71, x_63); -x_73 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__22; -x_74 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__24; -lean_inc(x_61); -x_75 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_75, 0, x_61); -lean_ctor_set(x_75, 1, x_73); -lean_ctor_set(x_75, 2, x_72); -lean_ctor_set(x_75, 3, x_74); -x_76 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__25; -x_77 = lean_array_push(x_76, x_75); -x_78 = lean_array_push(x_77, x_45); -x_79 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__20; -x_80 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_80, 0, x_43); -lean_ctor_set(x_80, 1, x_79); -lean_ctor_set(x_80, 2, x_78); -x_81 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__26; -x_82 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_82, 0, x_61); -lean_ctor_set(x_82, 1, x_81); -x_83 = lean_array_push(x_49, x_70); -x_84 = lean_array_push(x_83, x_46); -x_85 = lean_array_push(x_84, x_80); -x_86 = lean_array_push(x_85, x_82); -x_87 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__17; -x_88 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_88, 0, x_43); -lean_ctor_set(x_88, 1, x_87); -lean_ctor_set(x_88, 2, x_86); -x_89 = lean_array_push(x_56, x_88); -x_90 = lean_alloc_ctor(0, 2, 0); +x_67 = lean_environment_main_module(x_66); +x_68 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__20; +lean_inc(x_60); +x_69 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_69, 0, x_60); +lean_ctor_set(x_69, 1, x_68); +x_70 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_addLocalInstancesForParamsAux___rarg___closed__2; +x_71 = l_Lean_addMacroScope(x_67, x_70, x_62); +x_72 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__24; +x_73 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__26; +lean_inc(x_60); +x_74 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_74, 0, x_60); +lean_ctor_set(x_74, 1, x_72); +lean_ctor_set(x_74, 2, x_71); +lean_ctor_set(x_74, 3, x_73); +x_75 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__27; +x_76 = lean_array_push(x_75, x_74); +x_77 = lean_array_push(x_76, x_44); +x_78 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__22; +x_79 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_79, 0, x_42); +lean_ctor_set(x_79, 1, x_78); +lean_ctor_set(x_79, 2, x_77); +x_80 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__28; +x_81 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_81, 0, x_60); +lean_ctor_set(x_81, 1, x_80); +x_82 = lean_array_push(x_47, x_69); +x_83 = lean_array_push(x_82, x_50); +x_84 = lean_array_push(x_83, x_79); +x_85 = lean_array_push(x_84, x_81); +x_86 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__19; +x_87 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_87, 0, x_42); +lean_ctor_set(x_87, 1, x_86); +lean_ctor_set(x_87, 2, x_85); +x_88 = lean_array_push(x_55, x_87); +x_89 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_89, 0, x_88); +lean_ctor_set(x_89, 1, x_32); +x_90 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_90, 0, x_89); -lean_ctor_set(x_90, 1, x_33); -x_91 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_91, 0, x_90); -x_20 = x_91; -x_21 = x_66; -goto block_25; +x_19 = x_90; +x_20 = x_65; +goto block_24; } -block_25: +block_24: { -lean_object* x_22; lean_object* x_23; -x_22 = lean_ctor_get(x_20, 0); -lean_inc(x_22); -lean_dec(x_20); -x_23 = lean_nat_add(x_4, x_6); -lean_dec(x_4); -x_3 = x_19; -x_4 = x_23; -x_7 = x_22; -x_14 = x_21; +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_19, 0); +lean_inc(x_21); +lean_dec(x_19); +x_22 = lean_nat_add(x_3, x_5); +lean_dec(x_3); +x_2 = x_18; +x_3 = x_22; +x_6 = x_21; +x_13 = x_20; goto _start; } } else { -lean_object* x_92; -lean_dec(x_12); -lean_dec(x_4); +lean_object* x_91; +lean_dec(x_11); lean_dec(x_3); lean_dec(x_2); -x_92 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_92, 0, x_7); -lean_ctor_set(x_92, 1, x_14); -return x_92; +x_91 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_91, 0, x_6); +lean_ctor_set(x_91, 1, x_13); +return x_91; } } else { -lean_object* x_93; -lean_dec(x_12); -lean_dec(x_4); +lean_object* x_92; +lean_dec(x_11); lean_dec(x_3); lean_dec(x_2); -x_93 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_93, 0, x_7); -lean_ctor_set(x_93, 1, x_14); -return x_93; +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_6); +lean_ctor_set(x_92, 1, x_13); +return x_92; } } } @@ -8256,82 +8277,82 @@ x_1 = lean_mk_string_from_bytes("_", 1); return x_1; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__7(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__7(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_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) { _start: { -uint8_t x_15; -x_15 = lean_nat_dec_le(x_5, x_4); -if (x_15 == 0) +uint8_t x_16; +x_16 = lean_nat_dec_le(x_6, x_5); +if (x_16 == 0) { -lean_object* x_16; uint8_t x_17; -x_16 = lean_unsigned_to_nat(0u); -x_17 = lean_nat_dec_eq(x_3, x_16); -if (x_17 == 0) +lean_object* x_17; uint8_t x_18; +x_17 = lean_unsigned_to_nat(0u); +x_18 = lean_nat_dec_eq(x_4, x_17); +if (x_18 == 0) { -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; 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_object* x_32; lean_object* x_33; -x_18 = lean_unsigned_to_nat(1u); -x_19 = lean_nat_sub(x_3, x_18); -lean_dec(x_3); -lean_inc(x_12); -x_20 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___spec__1___rarg(x_12, x_13, x_14); -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_20, 1); +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; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_nat_sub(x_4, x_19); +lean_dec(x_4); +lean_inc(x_13); +x_21 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___spec__1___rarg(x_13, x_14, x_15); +x_22 = lean_ctor_get(x_21, 0); lean_inc(x_22); -lean_dec(x_20); -x_23 = lean_st_ref_get(x_13, x_22); -x_24 = lean_ctor_get(x_23, 1); -lean_inc(x_24); -lean_dec(x_23); -x_25 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__7___closed__1; -lean_inc(x_1); -x_26 = lean_name_mk_string(x_1, x_25); -x_27 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__7___closed__2; -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_21); -lean_ctor_set(x_28, 1, x_27); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = lean_st_ref_get(x_14, x_23); +x_25 = lean_ctor_get(x_24, 1); +lean_inc(x_25); +lean_dec(x_24); +x_26 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__7___closed__1; lean_inc(x_2); -x_29 = lean_array_push(x_2, x_28); -x_30 = lean_box(2); -x_31 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_26); -lean_ctor_set(x_31, 2, x_29); -x_32 = lean_array_push(x_7, x_31); -x_33 = lean_nat_add(x_4, x_6); -lean_dec(x_4); -x_3 = x_19; -x_4 = x_33; -x_7 = x_32; -x_14 = x_24; +x_27 = lean_name_mk_string(x_2, x_26); +x_28 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__7___closed__2; +x_29 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_29, 0, x_22); +lean_ctor_set(x_29, 1, x_28); +lean_inc(x_3); +x_30 = lean_array_push(x_3, x_29); +x_31 = lean_box(2); +x_32 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_27); +lean_ctor_set(x_32, 2, x_30); +x_33 = lean_array_push(x_8, x_32); +x_34 = lean_nat_add(x_5, x_7); +lean_dec(x_5); +x_4 = x_20; +x_5 = x_34; +x_8 = x_33; +x_15 = x_25; goto _start; } else { -lean_object* x_35; -lean_dec(x_12); +lean_object* x_36; +lean_dec(x_13); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_7); -lean_ctor_set(x_35, 1, x_14); -return x_35; +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_8); +lean_ctor_set(x_36, 1, x_15); +return x_36; } } else { -lean_object* x_36; -lean_dec(x_12); +lean_object* x_37; +lean_dec(x_13); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_7); -lean_ctor_set(x_36, 1, x_14); -return x_36; +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_8); +lean_ctor_set(x_37, 1, x_15); +return x_37; } } } @@ -8360,105 +8381,105 @@ x_1 = lean_mk_string_from_bytes("default", 7); return x_1; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__8(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_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, lean_object* x_16) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__8(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_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, lean_object* x_16, lean_object* x_17) { _start: { -uint8_t x_17; -x_17 = lean_nat_dec_le(x_7, x_6); -if (x_17 == 0) +uint8_t x_18; +x_18 = lean_nat_dec_le(x_8, x_7); +if (x_18 == 0) { -lean_object* x_18; uint8_t x_19; -x_18 = lean_unsigned_to_nat(0u); -x_19 = lean_nat_dec_eq(x_5, x_18); -if (x_19 == 0) +lean_object* x_19; uint8_t x_20; +x_19 = lean_unsigned_to_nat(0u); +x_20 = lean_nat_dec_eq(x_6, x_19); +if (x_20 == 0) { -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; 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; -x_20 = lean_unsigned_to_nat(1u); -x_21 = lean_nat_sub(x_5, x_20); -lean_dec(x_5); -lean_inc(x_14); -x_22 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___spec__1___rarg(x_14, x_15, x_16); -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_22, 1); +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; 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; +x_21 = lean_unsigned_to_nat(1u); +x_22 = lean_nat_sub(x_6, x_21); +lean_dec(x_6); +lean_inc(x_15); +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___spec__1___rarg(x_15, x_16, x_17); +x_24 = lean_ctor_get(x_23, 0); lean_inc(x_24); -lean_dec(x_22); -x_25 = lean_ctor_get(x_14, 10); +x_25 = lean_ctor_get(x_23, 1); lean_inc(x_25); -x_26 = lean_st_ref_get(x_15, x_24); -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); +lean_dec(x_23); +x_26 = lean_ctor_get(x_15, 10); +lean_inc(x_26); +x_27 = lean_st_ref_get(x_16, x_25); +x_28 = lean_ctor_get(x_27, 0); lean_inc(x_28); -lean_dec(x_26); -x_29 = lean_ctor_get(x_27, 0); +x_29 = lean_ctor_get(x_27, 1); lean_inc(x_29); lean_dec(x_27); -x_30 = lean_environment_main_module(x_29); -x_31 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__8___closed__1; -x_32 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__8___closed__2; -lean_inc(x_1); -x_33 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_1); -lean_ctor_set(x_33, 2, x_32); -x_34 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__8___closed__3; +x_30 = lean_ctor_get(x_28, 0); +lean_inc(x_30); +lean_dec(x_28); +x_31 = lean_environment_main_module(x_30); +x_32 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__8___closed__1; +x_33 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__8___closed__2; lean_inc(x_2); -x_35 = lean_name_mk_string(x_2, x_34); -lean_inc(x_35); -x_36 = l_Lean_addMacroScope(x_30, x_35, x_25); +x_34 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_2); +lean_ctor_set(x_34, 2, x_33); +x_35 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__8___closed__3; lean_inc(x_3); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_35); -lean_ctor_set(x_37, 1, x_3); +x_36 = lean_name_mk_string(x_3, x_35); +lean_inc(x_36); +x_37 = l_Lean_addMacroScope(x_31, x_36, x_26); lean_inc(x_4); -x_38 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_38, 0, x_37); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_36); lean_ctor_set(x_38, 1, x_4); -x_39 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_39, 0, x_23); -lean_ctor_set(x_39, 1, x_33); -lean_ctor_set(x_39, 2, x_36); -lean_ctor_set(x_39, 3, x_38); -x_40 = lean_array_push(x_9, x_39); -x_41 = lean_nat_add(x_6, x_8); -lean_dec(x_6); -x_5 = x_21; -x_6 = x_41; -x_9 = x_40; -x_16 = x_28; +lean_inc(x_5); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_5); +x_40 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_40, 0, x_24); +lean_ctor_set(x_40, 1, x_34); +lean_ctor_set(x_40, 2, x_37); +lean_ctor_set(x_40, 3, x_39); +x_41 = lean_array_push(x_10, x_40); +x_42 = lean_nat_add(x_7, x_9); +lean_dec(x_7); +x_6 = x_22; +x_7 = x_42; +x_10 = x_41; +x_17 = x_29; goto _start; } else { -lean_object* x_43; -lean_dec(x_14); +lean_object* x_44; +lean_dec(x_15); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_9); -lean_ctor_set(x_43, 1, x_16); -return x_43; +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_10); +lean_ctor_set(x_44, 1, x_17); +return x_44; } } else { -lean_object* x_44; -lean_dec(x_14); +lean_object* x_45; +lean_dec(x_15); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_44 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_44, 0, x_9); -lean_ctor_set(x_44, 1, x_16); -return x_44; +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_10); +lean_ctor_set(x_45, 1, x_17); +return x_45; } } } @@ -8466,41 +8487,86 @@ static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_ _start: { lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(0u); -x_2 = lean_mk_empty_array_with_capacity(x_1); +x_1 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__14; +x_2 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); return x_2; } } static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__1; -x_2 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_2, 0, x_1); -lean_ctor_set(x_2, 1, x_1); -return x_2; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("term", 4); +return x_1; } } static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__3() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__2; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__3; +x_3 = lean_alloc_ctor(1, 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___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_addLocalInstancesForParamsAux___rarg___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___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__5; +x_3 = lean_alloc_ctor(1, 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___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__7() { +_start: +{ lean_object* x_1; x_1 = lean_mk_string_from_bytes("paren", 5); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__4() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__8; -x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__3; +x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__7; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__5() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__9() { _start: { lean_object* x_1; @@ -8508,7 +8574,7 @@ x_1 = lean_mk_string_from_bytes("(", 1); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__6() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__10() { _start: { lean_object* x_1; @@ -8516,17 +8582,17 @@ x_1 = lean_mk_string_from_bytes("explicit", 8); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__7() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__8; -x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__6; +x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__10; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__8() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__12() { _start: { lean_object* x_1; @@ -8534,13 +8600,13 @@ x_1 = lean_mk_string_from_bytes("@", 1); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__9() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); x_2 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__13; -x_3 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__1; +x_3 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__14; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -8548,7 +8614,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__10() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__14() { _start: { lean_object* x_1; @@ -8556,7 +8622,7 @@ x_1 = lean_mk_string_from_bytes(")", 1); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__11() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__15() { _start: { lean_object* x_1; lean_object* x_2; @@ -8565,7 +8631,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__12() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__16() { _start: { lean_object* x_1; @@ -8573,17 +8639,17 @@ x_1 = lean_mk_string_from_bytes("anonymousCtor", 13); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__13() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__8; -x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__12; +x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__16; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__14() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__18() { _start: { lean_object* x_1; @@ -8591,7 +8657,7 @@ x_1 = lean_mk_string_from_bytes("⟨", 3); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__15() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__19() { _start: { lean_object* x_1; @@ -8599,7 +8665,7 @@ x_1 = lean_mk_string_from_bytes("⟩", 3); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__16() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__20() { _start: { lean_object* x_1; @@ -8607,17 +8673,17 @@ x_1 = lean_mk_string_from_bytes("Command", 7); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__17() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__6; -x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__16; +x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__20; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__18() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__22() { _start: { lean_object* x_1; @@ -8625,17 +8691,17 @@ x_1 = lean_mk_string_from_bytes("declaration", 11); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__19() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__17; -x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__18; +x_1 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__21; +x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__22; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__20() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__24() { _start: { lean_object* x_1; @@ -8643,17 +8709,17 @@ x_1 = lean_mk_string_from_bytes("declModifiers", 13); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__21() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__17; -x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__20; +x_1 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__21; +x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__24; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__22() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__26() { _start: { lean_object* x_1; lean_object* x_2; @@ -8662,73 +8728,73 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__23() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__27() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__22; -x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__9; +x_1 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__26; +x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__13; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__24() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__23; -x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__9; +x_1 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__27; +x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__13; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__25() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__29() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__24; -x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__9; +x_1 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__28; +x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__13; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__26() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__30() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__25; -x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__9; +x_1 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__29; +x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__13; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__27() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__31() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__26; -x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__9; +x_1 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__30; +x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__13; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__28() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__32() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__27; -x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__9; +x_1 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__31; +x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__13; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__29() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__33() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__21; -x_3 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__28; +x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__25; +x_3 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__32; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -8736,7 +8802,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__30() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__34() { _start: { lean_object* x_1; @@ -8744,17 +8810,17 @@ x_1 = lean_mk_string_from_bytes("instance", 8); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__31() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__35() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__17; -x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__30; +x_1 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__21; +x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__34; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__32() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__36() { _start: { lean_object* x_1; @@ -8762,33 +8828,33 @@ x_1 = lean_mk_string_from_bytes("attrKind", 8); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__33() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__37() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__8; -x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__32; +x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__36; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__34() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__38() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_addLocalInstancesForParamsAux___rarg___closed__3; -x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__9; +x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__13; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__35() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__39() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__33; -x_3 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__34; +x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__37; +x_3 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__38; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -8796,7 +8862,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__36() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__40() { _start: { lean_object* x_1; @@ -8804,17 +8870,17 @@ x_1 = lean_mk_string_from_bytes("declSig", 7); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__37() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__41() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__17; -x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__36; +x_1 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__21; +x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__40; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__38() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__42() { _start: { lean_object* x_1; @@ -8822,17 +8888,17 @@ x_1 = lean_mk_string_from_bytes("typeSpec", 8); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__39() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__43() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__8; -x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__38; +x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__42; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__40() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__44() { _start: { lean_object* x_1; @@ -8840,7 +8906,7 @@ x_1 = lean_mk_string_from_bytes(":", 1); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__41() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__45() { _start: { lean_object* x_1; @@ -8848,17 +8914,17 @@ x_1 = lean_mk_string_from_bytes("declValSimple", 13); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__42() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__46() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__17; -x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__41; +x_1 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__21; +x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__45; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__43() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__47() { _start: { lean_object* x_1; @@ -8866,7 +8932,7 @@ x_1 = lean_mk_string_from_bytes(":=", 2); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__44() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__48() { _start: { lean_object* x_1; lean_object* x_2; @@ -8875,22 +8941,22 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__45() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__49() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__44; -x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__35; +x_1 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__48; +x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__39; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__46() { +static lean_object* _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__50() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__25; -x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__29; +x_1 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__27; +x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__33; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -8915,7 +8981,7 @@ lean_inc(x_2); x_14 = l_Lean_getConstInfoCtor___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__3(x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_13); if (lean_obj_tag(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; lean_object* 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; uint8_t x_119; +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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); x_16 = lean_ctor_get(x_14, 1); @@ -8929,401 +8995,402 @@ lean_dec(x_12); x_19 = lean_nat_add(x_17, x_18); lean_dec(x_18); lean_dec(x_17); -x_20 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__1; -x_21 = lean_unsigned_to_nat(0u); -x_22 = lean_unsigned_to_nat(1u); -x_23 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__2; +x_20 = lean_unsigned_to_nat(0u); +x_21 = lean_unsigned_to_nat(1u); +x_22 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__1; lean_inc(x_8); lean_inc(x_19); -x_24 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6(x_3, x_20, x_19, x_21, x_19, x_22, x_23, x_4, x_5, x_6, x_7, x_8, x_9, x_16); +x_23 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6(x_3, x_19, x_20, x_19, x_21, x_22, x_4, x_5, x_6, x_7, x_8, x_9, x_16); lean_dec(x_19); -x_25 = lean_ctor_get(x_24, 0); +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); lean_inc(x_25); -x_26 = lean_ctor_get(x_24, 1); +lean_dec(x_23); +x_26 = lean_ctor_get(x_24, 0); lean_inc(x_26); -lean_dec(x_24); -x_27 = lean_ctor_get(x_25, 0); +x_27 = lean_ctor_get(x_24, 1); lean_inc(x_27); -x_28 = lean_ctor_get(x_25, 1); -lean_inc(x_28); -lean_dec(x_25); +lean_dec(x_24); lean_inc(x_8); -x_29 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___spec__1___rarg(x_8, x_9, x_26); -x_30 = lean_ctor_get(x_29, 0); +x_28 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___spec__1___rarg(x_8, x_9, x_25); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); lean_inc(x_30); -x_31 = lean_ctor_get(x_29, 1); +lean_dec(x_28); +x_31 = lean_ctor_get(x_8, 10); lean_inc(x_31); -lean_dec(x_29); -x_32 = lean_ctor_get(x_8, 10); -lean_inc(x_32); -x_33 = lean_st_ref_get(x_9, x_31); -x_34 = lean_ctor_get(x_33, 0); +x_32 = lean_st_ref_get(x_9, x_30); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); lean_inc(x_34); -x_35 = lean_ctor_get(x_33, 1); +lean_dec(x_32); +x_35 = lean_ctor_get(x_33, 0); lean_inc(x_35); lean_dec(x_33); -x_36 = lean_ctor_get(x_34, 0); -lean_inc(x_36); -lean_dec(x_34); -x_37 = lean_environment_main_module(x_36); +x_36 = lean_environment_main_module(x_35); +x_37 = lean_box(0); x_38 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_addLocalInstancesForParamsAux___rarg___closed__2; -x_39 = l_Lean_addMacroScope(x_37, x_38, x_32); -x_40 = lean_box(0); -x_41 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__22; -x_42 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__24; -lean_inc(x_30); -x_43 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_43, 0, x_30); -lean_ctor_set(x_43, 1, x_41); -lean_ctor_set(x_43, 2, x_39); -lean_ctor_set(x_43, 3, x_42); -x_44 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__5; -lean_inc(x_30); -x_45 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_45, 0, x_30); -lean_ctor_set(x_45, 1, x_44); -x_46 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__8; -lean_inc(x_30); -x_47 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_47, 0, x_30); -lean_ctor_set(x_47, 1, x_46); -x_48 = lean_mk_syntax_ident(x_1); -x_49 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__25; +x_39 = l_Lean_addMacroScope(x_36, x_38, x_31); +x_40 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__24; +x_41 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__6; +lean_inc(x_29); +x_42 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_42, 0, x_29); +lean_ctor_set(x_42, 1, x_40); +lean_ctor_set(x_42, 2, x_39); +lean_ctor_set(x_42, 3, x_41); +x_43 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__9; +lean_inc(x_29); +x_44 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_44, 0, x_29); +lean_ctor_set(x_44, 1, x_43); +x_45 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__12; +lean_inc(x_29); +x_46 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_46, 0, x_29); +lean_ctor_set(x_46, 1, x_45); +x_47 = lean_mk_syntax_ident(x_1); +x_48 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__27; +x_49 = lean_array_push(x_48, x_46); x_50 = lean_array_push(x_49, x_47); -x_51 = lean_array_push(x_50, x_48); -x_52 = lean_box(2); -x_53 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__7; -x_54 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_54, 0, x_52); -lean_ctor_set(x_54, 1, x_53); -lean_ctor_set(x_54, 2, x_51); -x_55 = l_Array_append___rarg(x_20, x_28); +x_51 = lean_box(2); +x_52 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__11; +x_53 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_52); +lean_ctor_set(x_53, 2, x_50); +x_54 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__14; +x_55 = l_Array_append___rarg(x_54, x_27); x_56 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__13; x_57 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_57, 0, x_52); +lean_ctor_set(x_57, 0, x_51); lean_ctor_set(x_57, 1, x_56); lean_ctor_set(x_57, 2, x_55); -x_58 = lean_array_push(x_49, x_54); +x_58 = lean_array_push(x_48, x_53); x_59 = lean_array_push(x_58, x_57); -x_60 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__20; +x_60 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__22; x_61 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_61, 0, x_52); +lean_ctor_set(x_61, 0, x_51); lean_ctor_set(x_61, 1, x_60); lean_ctor_set(x_61, 2, x_59); -x_62 = lean_array_push(x_49, x_61); -x_63 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__9; +x_62 = lean_array_push(x_48, x_61); +x_63 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__13; x_64 = lean_array_push(x_62, x_63); x_65 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_65, 0, x_52); +lean_ctor_set(x_65, 0, x_51); lean_ctor_set(x_65, 1, x_56); lean_ctor_set(x_65, 2, x_64); -x_66 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__10; +x_66 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__14; x_67 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_67, 0, x_30); +lean_ctor_set(x_67, 0, x_29); lean_ctor_set(x_67, 1, x_66); -x_68 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__11; -x_69 = lean_array_push(x_68, x_45); +x_68 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__15; +x_69 = lean_array_push(x_68, x_44); x_70 = lean_array_push(x_69, x_65); x_71 = lean_array_push(x_70, x_67); -x_72 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__4; +x_72 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__8; x_73 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_73, 0, x_52); +lean_ctor_set(x_73, 0, x_51); lean_ctor_set(x_73, 1, x_72); lean_ctor_set(x_73, 2, x_71); x_74 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_addLocalInstancesForParamsAux___rarg___closed__3; x_75 = lean_array_push(x_74, x_73); x_76 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_76, 0, x_52); +lean_ctor_set(x_76, 0, x_51); lean_ctor_set(x_76, 1, x_56); lean_ctor_set(x_76, 2, x_75); -x_77 = lean_array_push(x_49, x_43); +x_77 = lean_array_push(x_48, x_42); x_78 = lean_array_push(x_77, x_76); x_79 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_79, 0, x_52); +lean_ctor_set(x_79, 0, x_51); lean_ctor_set(x_79, 1, x_60); lean_ctor_set(x_79, 2, x_78); x_80 = lean_ctor_get(x_15, 3); lean_inc(x_80); -x_81 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__8; +x_81 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__4; +x_82 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__8; lean_inc(x_8); lean_inc(x_80); -x_82 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__7(x_81, x_74, x_80, x_21, x_80, x_22, x_20, x_4, x_5, x_6, x_7, x_8, x_9, x_35); +x_83 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__7(x_81, x_82, x_74, x_80, x_20, x_80, x_21, x_54, x_4, x_5, x_6, x_7, x_8, x_9, x_34); lean_dec(x_80); -x_83 = lean_ctor_get(x_82, 0); -lean_inc(x_83); -x_84 = lean_ctor_get(x_82, 1); +x_84 = lean_ctor_get(x_83, 0); lean_inc(x_84); -lean_dec(x_82); -x_85 = lean_ctor_get(x_15, 4); +x_85 = lean_ctor_get(x_83, 1); lean_inc(x_85); +lean_dec(x_83); +x_86 = lean_ctor_get(x_15, 4); +lean_inc(x_86); lean_dec(x_15); lean_inc(x_8); -lean_inc(x_85); -x_86 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__8(x_21, x_38, x_40, x_40, x_85, x_21, x_85, x_22, x_83, x_4, x_5, x_6, x_7, x_8, x_9, x_84); +lean_inc(x_86); +x_87 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__8(x_81, x_20, x_38, x_37, x_37, x_86, x_20, x_86, x_21, x_84, x_4, x_5, x_6, x_7, x_8, x_9, x_85); lean_dec(x_4); -lean_dec(x_85); -x_87 = lean_ctor_get(x_86, 0); -lean_inc(x_87); -x_88 = lean_ctor_get(x_86, 1); -lean_inc(x_88); lean_dec(x_86); +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +lean_dec(x_87); lean_inc(x_8); -x_89 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___spec__1___rarg(x_8, x_9, x_88); -x_90 = lean_ctor_get(x_89, 0); -lean_inc(x_90); -x_91 = lean_ctor_get(x_89, 1); +x_90 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___spec__1___rarg(x_8, x_9, x_89); +x_91 = lean_ctor_get(x_90, 0); +lean_inc(x_91); +x_92 = lean_ctor_get(x_90, 1); +lean_inc(x_92); +lean_dec(x_90); +x_93 = lean_st_ref_get(x_9, x_92); +x_94 = lean_ctor_get(x_93, 1); +lean_inc(x_94); +lean_dec(x_93); +x_95 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__18; lean_inc(x_91); -lean_dec(x_89); -x_92 = lean_st_ref_get(x_9, x_91); -x_93 = lean_ctor_get(x_92, 1); -lean_inc(x_93); -lean_dec(x_92); -x_94 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__14; -lean_inc(x_90); -x_95 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_95, 0, x_90); -lean_ctor_set(x_95, 1, x_94); -lean_inc(x_90); x_96 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_96, 0, x_90); -lean_ctor_set(x_96, 1, x_46); -x_97 = lean_mk_syntax_ident(x_2); -x_98 = lean_array_push(x_49, x_96); -x_99 = lean_array_push(x_98, x_97); -x_100 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_100, 0, x_52); -lean_ctor_set(x_100, 1, x_53); -lean_ctor_set(x_100, 2, x_99); -x_101 = l_Array_append___rarg(x_20, x_87); -x_102 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_102, 0, x_52); -lean_ctor_set(x_102, 1, x_56); -lean_ctor_set(x_102, 2, x_101); -x_103 = lean_array_push(x_49, x_100); -x_104 = lean_array_push(x_103, x_102); -x_105 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_105, 0, x_52); -lean_ctor_set(x_105, 1, x_60); -lean_ctor_set(x_105, 2, x_104); -x_106 = lean_array_push(x_74, x_105); -x_107 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_107, 0, x_52); -lean_ctor_set(x_107, 1, x_56); -lean_ctor_set(x_107, 2, x_106); -x_108 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__15; -x_109 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_109, 0, x_90); -lean_ctor_set(x_109, 1, x_108); -x_110 = lean_array_push(x_68, x_95); -x_111 = lean_array_push(x_110, x_107); -x_112 = lean_array_push(x_111, x_109); -x_113 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__13; -x_114 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_114, 0, x_52); -lean_ctor_set(x_114, 1, x_113); -lean_ctor_set(x_114, 2, x_112); -x_115 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___spec__1___rarg(x_8, x_9, x_93); -x_116 = lean_ctor_get(x_115, 0); -lean_inc(x_116); -x_117 = lean_ctor_get(x_115, 1); +lean_ctor_set(x_96, 0, x_91); +lean_ctor_set(x_96, 1, x_95); +lean_inc(x_91); +x_97 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_97, 0, x_91); +lean_ctor_set(x_97, 1, x_45); +x_98 = lean_mk_syntax_ident(x_2); +x_99 = lean_array_push(x_48, x_97); +x_100 = lean_array_push(x_99, x_98); +x_101 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_101, 0, x_51); +lean_ctor_set(x_101, 1, x_52); +lean_ctor_set(x_101, 2, x_100); +x_102 = l_Array_append___rarg(x_54, x_88); +x_103 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_103, 0, x_51); +lean_ctor_set(x_103, 1, x_56); +lean_ctor_set(x_103, 2, x_102); +x_104 = lean_array_push(x_48, x_101); +x_105 = lean_array_push(x_104, x_103); +x_106 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_106, 0, x_51); +lean_ctor_set(x_106, 1, x_60); +lean_ctor_set(x_106, 2, x_105); +x_107 = lean_array_push(x_74, x_106); +x_108 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_108, 0, x_51); +lean_ctor_set(x_108, 1, x_56); +lean_ctor_set(x_108, 2, x_107); +x_109 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__19; +x_110 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_110, 0, x_91); +lean_ctor_set(x_110, 1, x_109); +x_111 = lean_array_push(x_68, x_96); +x_112 = lean_array_push(x_111, x_108); +x_113 = lean_array_push(x_112, x_110); +x_114 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__17; +x_115 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_115, 0, x_51); +lean_ctor_set(x_115, 1, x_114); +lean_ctor_set(x_115, 2, x_113); +x_116 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___spec__1___rarg(x_8, x_9, x_94); +x_117 = lean_ctor_get(x_116, 0); lean_inc(x_117); -lean_dec(x_115); -x_118 = lean_st_ref_get(x_9, x_117); -x_119 = !lean_is_exclusive(x_118); -if (x_119 == 0) -{ -lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; -x_120 = lean_ctor_get(x_118, 0); -lean_dec(x_120); -x_121 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__30; -lean_inc(x_116); -x_122 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_122, 0, x_116); -lean_ctor_set(x_122, 1, x_121); -x_123 = l_Array_append___rarg(x_20, x_27); -x_124 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_124, 0, x_52); -lean_ctor_set(x_124, 1, x_56); -lean_ctor_set(x_124, 2, x_123); -x_125 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__40; -lean_inc(x_116); -x_126 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_126, 0, x_116); -lean_ctor_set(x_126, 1, x_125); -x_127 = lean_array_push(x_49, x_126); -x_128 = lean_array_push(x_127, x_79); -x_129 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__39; -x_130 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_130, 0, x_52); -lean_ctor_set(x_130, 1, x_129); -lean_ctor_set(x_130, 2, x_128); -x_131 = lean_array_push(x_49, x_124); -x_132 = lean_array_push(x_131, x_130); -x_133 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__37; -x_134 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_134, 0, x_52); -lean_ctor_set(x_134, 1, x_133); -lean_ctor_set(x_134, 2, x_132); -x_135 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__43; -x_136 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_136, 0, x_116); -lean_ctor_set(x_136, 1, x_135); -x_137 = lean_array_push(x_68, x_136); -x_138 = lean_array_push(x_137, x_114); -x_139 = lean_array_push(x_138, x_63); -x_140 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__42; -x_141 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_141, 0, x_52); -lean_ctor_set(x_141, 1, x_140); -lean_ctor_set(x_141, 2, x_139); -x_142 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__45; -x_143 = lean_array_push(x_142, x_122); -x_144 = lean_array_push(x_143, x_63); +x_118 = lean_ctor_get(x_116, 1); +lean_inc(x_118); +lean_dec(x_116); +x_119 = lean_st_ref_get(x_9, x_118); +x_120 = !lean_is_exclusive(x_119); +if (x_120 == 0) +{ +lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; +x_121 = lean_ctor_get(x_119, 0); +lean_dec(x_121); +x_122 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__34; +lean_inc(x_117); +x_123 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_123, 0, x_117); +lean_ctor_set(x_123, 1, x_122); +x_124 = l_Array_append___rarg(x_54, x_26); +x_125 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_125, 0, x_51); +lean_ctor_set(x_125, 1, x_56); +lean_ctor_set(x_125, 2, x_124); +x_126 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__44; +lean_inc(x_117); +x_127 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_127, 0, x_117); +lean_ctor_set(x_127, 1, x_126); +x_128 = lean_array_push(x_48, x_127); +x_129 = lean_array_push(x_128, x_79); +x_130 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__43; +x_131 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_131, 0, x_51); +lean_ctor_set(x_131, 1, x_130); +lean_ctor_set(x_131, 2, x_129); +x_132 = lean_array_push(x_48, x_125); +x_133 = lean_array_push(x_132, x_131); +x_134 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__41; +x_135 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_135, 0, x_51); +lean_ctor_set(x_135, 1, x_134); +lean_ctor_set(x_135, 2, x_133); +x_136 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__47; +x_137 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_137, 0, x_117); +lean_ctor_set(x_137, 1, x_136); +x_138 = lean_array_push(x_68, x_137); +x_139 = lean_array_push(x_138, x_115); +x_140 = lean_array_push(x_139, x_63); +x_141 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__46; +x_142 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_142, 0, x_51); +lean_ctor_set(x_142, 1, x_141); +lean_ctor_set(x_142, 2, x_140); +x_143 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__49; +x_144 = lean_array_push(x_143, x_123); x_145 = lean_array_push(x_144, x_63); -x_146 = lean_array_push(x_145, x_134); -x_147 = lean_array_push(x_146, x_141); -x_148 = lean_array_push(x_147, x_63); +x_146 = lean_array_push(x_145, x_63); +x_147 = lean_array_push(x_146, x_135); +x_148 = lean_array_push(x_147, x_142); x_149 = lean_array_push(x_148, x_63); -x_150 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__31; -x_151 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_151, 0, x_52); -lean_ctor_set(x_151, 1, x_150); -lean_ctor_set(x_151, 2, x_149); -x_152 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__46; -x_153 = lean_array_push(x_152, x_151); -x_154 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__19; -x_155 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_155, 0, x_52); -lean_ctor_set(x_155, 1, x_154); -lean_ctor_set(x_155, 2, x_153); -lean_ctor_set(x_118, 0, x_155); -return x_118; +x_150 = lean_array_push(x_149, x_63); +x_151 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__35; +x_152 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_152, 0, x_51); +lean_ctor_set(x_152, 1, x_151); +lean_ctor_set(x_152, 2, x_150); +x_153 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__50; +x_154 = lean_array_push(x_153, x_152); +x_155 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__23; +x_156 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_156, 0, x_51); +lean_ctor_set(x_156, 1, x_155); +lean_ctor_set(x_156, 2, x_154); +lean_ctor_set(x_119, 0, x_156); +return x_119; } else { -lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; -x_156 = lean_ctor_get(x_118, 1); -lean_inc(x_156); -lean_dec(x_118); -x_157 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__30; -lean_inc(x_116); -x_158 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_158, 0, x_116); -lean_ctor_set(x_158, 1, x_157); -x_159 = l_Array_append___rarg(x_20, x_27); -x_160 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_160, 0, x_52); -lean_ctor_set(x_160, 1, x_56); -lean_ctor_set(x_160, 2, x_159); -x_161 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__40; -lean_inc(x_116); -x_162 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_162, 0, x_116); -lean_ctor_set(x_162, 1, x_161); -x_163 = lean_array_push(x_49, x_162); -x_164 = lean_array_push(x_163, x_79); -x_165 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__39; -x_166 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_166, 0, x_52); -lean_ctor_set(x_166, 1, x_165); -lean_ctor_set(x_166, 2, x_164); -x_167 = lean_array_push(x_49, x_160); -x_168 = lean_array_push(x_167, x_166); -x_169 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__37; -x_170 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_170, 0, x_52); -lean_ctor_set(x_170, 1, x_169); -lean_ctor_set(x_170, 2, x_168); -x_171 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__43; -x_172 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_172, 0, x_116); -lean_ctor_set(x_172, 1, x_171); -x_173 = lean_array_push(x_68, x_172); -x_174 = lean_array_push(x_173, x_114); -x_175 = lean_array_push(x_174, x_63); -x_176 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__42; -x_177 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_177, 0, x_52); -lean_ctor_set(x_177, 1, x_176); -lean_ctor_set(x_177, 2, x_175); -x_178 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__45; -x_179 = lean_array_push(x_178, x_158); -x_180 = lean_array_push(x_179, x_63); +lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; +x_157 = lean_ctor_get(x_119, 1); +lean_inc(x_157); +lean_dec(x_119); +x_158 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__34; +lean_inc(x_117); +x_159 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_159, 0, x_117); +lean_ctor_set(x_159, 1, x_158); +x_160 = l_Array_append___rarg(x_54, x_26); +x_161 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_161, 0, x_51); +lean_ctor_set(x_161, 1, x_56); +lean_ctor_set(x_161, 2, x_160); +x_162 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__44; +lean_inc(x_117); +x_163 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_163, 0, x_117); +lean_ctor_set(x_163, 1, x_162); +x_164 = lean_array_push(x_48, x_163); +x_165 = lean_array_push(x_164, x_79); +x_166 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__43; +x_167 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_167, 0, x_51); +lean_ctor_set(x_167, 1, x_166); +lean_ctor_set(x_167, 2, x_165); +x_168 = lean_array_push(x_48, x_161); +x_169 = lean_array_push(x_168, x_167); +x_170 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__41; +x_171 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_171, 0, x_51); +lean_ctor_set(x_171, 1, x_170); +lean_ctor_set(x_171, 2, x_169); +x_172 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__47; +x_173 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_173, 0, x_117); +lean_ctor_set(x_173, 1, x_172); +x_174 = lean_array_push(x_68, x_173); +x_175 = lean_array_push(x_174, x_115); +x_176 = lean_array_push(x_175, x_63); +x_177 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__46; +x_178 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_178, 0, x_51); +lean_ctor_set(x_178, 1, x_177); +lean_ctor_set(x_178, 2, x_176); +x_179 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__49; +x_180 = lean_array_push(x_179, x_159); x_181 = lean_array_push(x_180, x_63); -x_182 = lean_array_push(x_181, x_170); -x_183 = lean_array_push(x_182, x_177); -x_184 = lean_array_push(x_183, x_63); +x_182 = lean_array_push(x_181, x_63); +x_183 = lean_array_push(x_182, x_171); +x_184 = lean_array_push(x_183, x_178); x_185 = lean_array_push(x_184, x_63); -x_186 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__31; -x_187 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_187, 0, x_52); -lean_ctor_set(x_187, 1, x_186); -lean_ctor_set(x_187, 2, x_185); -x_188 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__46; -x_189 = lean_array_push(x_188, x_187); -x_190 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__19; -x_191 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_191, 0, x_52); -lean_ctor_set(x_191, 1, x_190); -lean_ctor_set(x_191, 2, x_189); -x_192 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_192, 0, x_191); -lean_ctor_set(x_192, 1, x_156); -return x_192; -} -} -else -{ -uint8_t x_193; +x_186 = lean_array_push(x_185, x_63); +x_187 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__35; +x_188 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_188, 0, x_51); +lean_ctor_set(x_188, 1, x_187); +lean_ctor_set(x_188, 2, x_186); +x_189 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__50; +x_190 = lean_array_push(x_189, x_188); +x_191 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__23; +x_192 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_192, 0, x_51); +lean_ctor_set(x_192, 1, x_191); +lean_ctor_set(x_192, 2, x_190); +x_193 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_193, 0, x_192); +lean_ctor_set(x_193, 1, x_157); +return x_193; +} +} +else +{ +uint8_t x_194; lean_dec(x_12); lean_dec(x_8); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_193 = !lean_is_exclusive(x_14); -if (x_193 == 0) +x_194 = !lean_is_exclusive(x_14); +if (x_194 == 0) { return x_14; } else { -lean_object* x_194; lean_object* x_195; lean_object* x_196; -x_194 = lean_ctor_get(x_14, 0); -x_195 = lean_ctor_get(x_14, 1); +lean_object* x_195; lean_object* x_196; lean_object* x_197; +x_195 = lean_ctor_get(x_14, 0); +x_196 = lean_ctor_get(x_14, 1); +lean_inc(x_196); lean_inc(x_195); -lean_inc(x_194); lean_dec(x_14); -x_196 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_196, 0, x_194); -lean_ctor_set(x_196, 1, x_195); -return x_196; +x_197 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_197, 0, x_195); +lean_ctor_set(x_197, 1, x_196); +return x_197; } } } else { -uint8_t x_197; +uint8_t x_198; lean_dec(x_8); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_197 = !lean_is_exclusive(x_11); -if (x_197 == 0) +x_198 = !lean_is_exclusive(x_11); +if (x_198 == 0) { return x_11; } else { -lean_object* x_198; lean_object* x_199; lean_object* x_200; -x_198 = lean_ctor_get(x_11, 0); -x_199 = lean_ctor_get(x_11, 1); +lean_object* x_199; lean_object* x_200; lean_object* x_201; +x_199 = lean_ctor_get(x_11, 0); +x_200 = lean_ctor_get(x_11, 1); +lean_inc(x_200); lean_inc(x_199); -lean_inc(x_198); lean_dec(x_11); -x_200 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_200, 0, x_198); -lean_ctor_set(x_200, 1, x_199); -return x_200; +x_201 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_201, 0, x_199); +lean_ctor_set(x_201, 1, x_200); +return x_201; } } } @@ -9390,50 +9457,69 @@ lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_15; -x_15 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_13); -lean_dec(x_11); +lean_object* x_14; +x_14 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_6); +lean_dec(x_7); lean_dec(x_5); +lean_dec(x_4); lean_dec(x_1); -return x_15; +return x_14; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__7___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__7___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_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) { _start: { -lean_object* x_15; -x_15 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_13); +lean_object* x_16; +x_16 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_14); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -return x_15; +lean_dec(x_1); +return x_16; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__8___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_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, lean_object* x_16) { -_start: -{ -lean_object* x_17; -x_17 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -lean_dec(x_15); +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__8___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +_start: +{ +lean_object* x_18; +x_18 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_dec(x_16); +lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); -return x_17; +lean_dec(x_1); +return x_18; } } LEAN_EXPORT lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___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_object* x_9, lean_object* x_10) { @@ -11606,7 +11692,7 @@ LEAN_EXPORT lean_object* l_Lean_getConstInfoInduct___at___private_Lean_Elab_Deri lean_object* x_5; lean_inc(x_2); lean_inc(x_1); -x_5 = l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__10(x_1, x_2, x_3, x_4); +x_5 = l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__14(x_1, x_2, x_3, x_4); if (lean_obj_tag(x_5) == 0) { lean_object* x_6; @@ -12801,7 +12887,7 @@ lean_dec(x_1); return x_5; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Inhabited___hyg_2506____closed__1() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Inhabited___hyg_2570____closed__1() { _start: { lean_object* x_1; @@ -12809,12 +12895,12 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Elab_mkInhabitedInstanceHandler___boxed) return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Inhabited___hyg_2506_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Inhabited___hyg_2570_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; x_2 = l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_addLocalInstancesForParamsAux___rarg___closed__2; -x_3 = l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Inhabited___hyg_2506____closed__1; +x_3 = l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Inhabited___hyg_2570____closed__1; x_4 = l_Lean_Elab_registerBuiltinDerivingHandler(x_2, x_3, x_1); if (lean_obj_tag(x_4) == 0) { @@ -12965,6 +13051,10 @@ l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_ lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__25); l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__26 = _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__26(); lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__26); +l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__27 = _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__27(); +lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__27); +l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__28 = _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__28(); +lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__6___closed__28); l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__7___closed__1 = _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__7___closed__1(); lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__7___closed__1); l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__7___closed__2 = _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__7___closed__2(); @@ -13067,6 +13157,14 @@ l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_m lean_mark_persistent(l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__45); l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__46 = _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__46(); lean_mark_persistent(l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__46); +l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__47 = _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__47(); +lean_mark_persistent(l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__47); +l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__48 = _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__48(); +lean_mark_persistent(l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__48); +l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__49 = _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__49(); +lean_mark_persistent(l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__49); +l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__50 = _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__50(); +lean_mark_persistent(l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__50); l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmd_x3f___spec__1___lambda__2___closed__1 = _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmd_x3f___spec__1___lambda__2___closed__1(); lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmd_x3f___spec__1___lambda__2___closed__1); l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmd_x3f___spec__1___lambda__2___closed__2 = _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmd_x3f___spec__1___lambda__2___closed__2(); @@ -13117,9 +13215,9 @@ l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstance___clos lean_mark_persistent(l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstance___closed__3); l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstance___closed__4 = _init_l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstance___closed__4(); lean_mark_persistent(l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstance___closed__4); -l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Inhabited___hyg_2506____closed__1 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Inhabited___hyg_2506____closed__1(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Inhabited___hyg_2506____closed__1); -res = l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Inhabited___hyg_2506_(lean_io_mk_world()); +l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Inhabited___hyg_2570____closed__1 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Inhabited___hyg_2570____closed__1(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Inhabited___hyg_2570____closed__1); +res = l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Inhabited___hyg_2570_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/Deriving/Ord.c b/stage0/stdlib/Lean/Elab/Deriving/Ord.c index c386963f55a..2f4060be76d 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/Ord.c +++ b/stage0/stdlib/Lean/Elab/Deriving/Ord.c @@ -19,6 +19,7 @@ lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(lean_obje static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3___closed__6; uint8_t l_Lean_Expr_hasAnyFVar_visit___at_Lean_Expr_containsFVar___spec__1(lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); +static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__52; lean_object* l_Lean_registerTraceClass(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTRAux___at___private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkOrdInstanceCmds___spec__1(lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); @@ -37,11 +38,13 @@ lean_object* l_Array_append___rarg(lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__43; static lean_object* l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__12; static lean_object* l_Lean_Elab_Deriving_Ord_mkMutualBlock___closed__2; +lean_object* lean_array_uset(lean_object*, size_t, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkOrdInstanceCmds(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkOrdInstanceCmds___closed__2; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__35; lean_object* lean_environment_find(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_Ord_initFn____x40_Lean_Elab_Deriving_Ord___hyg_2936____closed__1; static lean_object* l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__16; lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__39; @@ -66,11 +69,13 @@ static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_m lean_object* lean_string_utf8_byte_size(lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__29; static lean_object* l___private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkOrdInstanceCmds___closed__8; +uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Deriving_Ord_mkOrdInstanceHandler___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__25; lean_object* l_Lean_Elab_Deriving_mkLocalInstanceLetDecls(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___private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkOrdInstanceCmds___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1(lean_object*, lean_object*, 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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__11; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3___closed__1; lean_object* lean_nat_add(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_Ord_mkMutualBlock___closed__3; @@ -87,7 +92,6 @@ static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_m LEAN_EXPORT lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__19; lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(size_t, size_t, lean_object*); -static lean_object* l_Lean_Elab_Deriving_Ord_initFn____x40_Lean_Elab_Deriving_Ord___hyg_2782____closed__1; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3___closed__8; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__32; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -103,9 +107,11 @@ static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3___closed__4; lean_object* l_Lean_Elab_Deriving_mkInstanceCmds(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3___boxed(lean_object*, lean_object*, 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_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_Ord_mkMatch___spec__1(size_t, size_t, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___closed__4; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3___closed__2; +static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__51; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__34; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__4; lean_object* l_Lean_Elab_Deriving_mkLet(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -113,10 +119,11 @@ lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Ord_mkOrdInstanceHandler(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__48; lean_object* l_Lean_Expr_fvarId_x21(lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___boxed(lean_object*, lean_object*, lean_object*, 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_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__28; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__1; static lean_object* l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__5; +static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__50; static lean_object* l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__29; static lean_object* l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__25; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Ord_mkMutualBlock(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -132,7 +139,7 @@ lean_object* lean_array_to_list(lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__16; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__8; extern lean_object* l_Lean_instInhabitedExpr; -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2(lean_object*, lean_object*, lean_object*, 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_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__6; lean_object* l_Lean_mkSepArray(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -148,7 +155,7 @@ static lean_object* l___private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__20; lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__16(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__41; -LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Ord_initFn____x40_Lean_Elab_Deriving_Ord___hyg_2782_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Ord_initFn____x40_Lean_Elab_Deriving_Ord___hyg_2936_(lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__6; lean_object* l_Lean_Elab_Deriving_mkHeader(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__23; @@ -160,7 +167,6 @@ static lean_object* l___private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__42; lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOptionM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTerm___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMutualBlock___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___closed__1; lean_object* lean_environment_main_module(lean_object*); static lean_object* l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__17; static lean_object* l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__8; @@ -173,7 +179,7 @@ static lean_object* l_Lean_Elab_Deriving_Ord_mkOrdHeader___closed__1; static lean_object* l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__2; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__26; static lean_object* l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__3; -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__1; static lean_object* l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__15; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3___closed__10; @@ -185,7 +191,7 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMa lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__2; lean_object* l_Lean_Elab_registerBuiltinDerivingHandler(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__47; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___closed__1; lean_object* lean_array_pop(lean_object*); @@ -215,6 +221,7 @@ LEAN_EXPORT lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_Ord_mkM static lean_object* l_Lean_Elab_Deriving_Ord_mkMutualBlock___closed__5; lean_object* l_Lean_Core_betaReduce___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Ord_mkMatch_mkAlts(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_Ord_mkMatch___spec__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkConst(lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__40; @@ -222,11 +229,10 @@ static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_m static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__27; static lean_object* l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__18; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__45; -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6(lean_object*, 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_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkOrdInstanceCmds___closed__3; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__17; lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Deriving_Ord_mkMatch___closed__1; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__8; static lean_object* l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__11; lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); @@ -738,22 +744,45 @@ return x_3; static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__5() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(2); +x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__4; +x_3 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__5; +x_4 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__7() { +_start: +{ lean_object* x_1; x_1 = lean_mk_string_from_bytes("matchDiscr", 10); return x_1; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__6() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3___closed__6; -x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__5; +x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__7; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__7() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__9() { _start: { lean_object* x_1; @@ -761,17 +790,17 @@ x_1 = lean_mk_string_from_bytes("app", 3); return x_1; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__8() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3___closed__6; -x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__7; +x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__9; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__9() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__11() { _start: { lean_object* x_1; @@ -779,22 +808,22 @@ x_1 = lean_mk_string_from_bytes("compare", 7); return x_1; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__10() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__9; +x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__11; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__11() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__9; +x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__11; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__10; +x_3 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__12; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -802,51 +831,51 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__12() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__9; +x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__11; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__13() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Deriving_Ord_mkOrdHeader___closed__2; -x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__9; +x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__11; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__14() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__13; +x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__15; 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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__15() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__14; +x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__16; x_3 = lean_alloc_ctor(1, 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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__16() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__18() { _start: { lean_object* x_1; lean_object* x_2; @@ -855,7 +884,17 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__17() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__18; +x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__6; +x_3 = lean_array_push(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__20() { _start: { lean_object* x_1; @@ -863,7 +902,7 @@ x_1 = lean_mk_string_from_bytes("with", 4); return x_1; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__18() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__21() { _start: { lean_object* x_1; @@ -871,17 +910,17 @@ x_1 = lean_mk_string_from_bytes("matchAlts", 9); return x_1; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__19() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3___closed__6; -x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__18; +x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__21; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__20() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__23() { _start: { lean_object* x_1; @@ -889,17 +928,17 @@ x_1 = lean_mk_string_from_bytes("matchAlt", 8); return x_1; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__21() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3___closed__6; -x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__20; +x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__23; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__22() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__25() { _start: { lean_object* x_1; @@ -907,7 +946,7 @@ x_1 = lean_mk_string_from_bytes("|", 1); return x_1; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__23() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__26() { _start: { lean_object* x_1; @@ -915,22 +954,22 @@ x_1 = lean_mk_string_from_bytes("Ordering.lt", 11); return x_1; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__24() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__27() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__23; +x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__26; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__25() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__23; +x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__26; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__24; +x_3 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__27; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -938,7 +977,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__26() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__29() { _start: { lean_object* x_1; @@ -946,17 +985,17 @@ x_1 = lean_mk_string_from_bytes("Ordering", 8); return x_1; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__27() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__30() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__26; +x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__29; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__28() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__31() { _start: { lean_object* x_1; @@ -964,41 +1003,41 @@ x_1 = lean_mk_string_from_bytes("lt", 2); return x_1; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__29() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__32() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__27; -x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__28; +x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__30; +x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__31; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__30() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__33() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__29; +x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__32; 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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__31() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__34() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__30; +x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__33; x_3 = lean_alloc_ctor(1, 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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__32() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__35() { _start: { lean_object* x_1; @@ -1006,7 +1045,7 @@ x_1 = lean_mk_string_from_bytes("=>", 2); return x_1; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__33() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__36() { _start: { lean_object* x_1; lean_object* x_2; @@ -1015,7 +1054,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__34() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__37() { _start: { lean_object* x_1; @@ -1023,22 +1062,22 @@ x_1 = lean_mk_string_from_bytes("Ordering.gt", 11); return x_1; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__35() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__38() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__34; +x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__37; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__36() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__39() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__34; +x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__37; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__35; +x_3 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__38; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1046,7 +1085,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__37() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__40() { _start: { lean_object* x_1; @@ -1054,41 +1093,41 @@ x_1 = lean_mk_string_from_bytes("gt", 2); return x_1; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__38() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__41() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__27; -x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__37; +x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__30; +x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__40; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__39() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__42() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__38; +x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__41; 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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__40() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__43() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__39; +x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__42; x_3 = lean_alloc_ctor(1, 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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__41() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__44() { _start: { lean_object* x_1; @@ -1096,22 +1135,22 @@ x_1 = lean_mk_string_from_bytes("Ordering.eq", 11); return x_1; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__42() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__45() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__41; +x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__44; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__43() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__46() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__41; +x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__44; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__42; +x_3 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__45; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1119,7 +1158,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__44() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__47() { _start: { lean_object* x_1; @@ -1127,41 +1166,41 @@ x_1 = lean_mk_string_from_bytes("eq", 2); return x_1; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__45() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__48() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__27; -x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__44; +x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__30; +x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__47; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__46() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__49() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__45; +x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__48; 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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__47() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__50() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__46; +x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__49; x_3 = lean_alloc_ctor(1, 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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__48() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__51() { _start: { lean_object* x_1; lean_object* x_2; @@ -1170,7 +1209,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__49() { +static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__52() { _start: { lean_object* x_1; lean_object* x_2; @@ -1179,217 +1218,212 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___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, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___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, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -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; 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; -lean_inc(x_10); -x_13 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_10, x_11, x_12); -x_14 = lean_ctor_get(x_13, 0); +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; 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; +lean_inc(x_9); +x_12 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_9, x_10, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); +lean_dec(x_12); +x_15 = lean_ctor_get(x_9, 10); lean_inc(x_15); -lean_dec(x_13); -x_16 = lean_ctor_get(x_10, 10); -lean_inc(x_16); -x_17 = lean_st_ref_get(x_11, x_15); -x_18 = lean_ctor_get(x_17, 0); +x_16 = lean_st_ref_get(x_10, x_14); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); +lean_dec(x_16); +x_19 = lean_ctor_get(x_17, 0); lean_inc(x_19); lean_dec(x_17); -x_20 = lean_ctor_get(x_18, 0); +x_20 = lean_environment_main_module(x_19); +x_21 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__1; +lean_inc(x_13); +x_22 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_22, 0, x_13); +lean_ctor_set(x_22, 1, x_21); +x_23 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__14; +lean_inc(x_15); lean_inc(x_20); -lean_dec(x_18); -x_21 = lean_environment_main_module(x_20); -x_22 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__1; -lean_inc(x_14); -x_23 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_23, 0, x_14); -lean_ctor_set(x_23, 1, x_22); -x_24 = lean_box(2); -x_25 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__4; -x_26 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -lean_ctor_set(x_26, 2, x_1); -x_27 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__12; -lean_inc(x_16); -lean_inc(x_21); -x_28 = l_Lean_addMacroScope(x_21, x_27, x_16); -x_29 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__11; -x_30 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__15; -lean_inc(x_14); -x_31 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_31, 0, x_14); -lean_ctor_set(x_31, 1, x_29); -lean_ctor_set(x_31, 2, x_28); -lean_ctor_set(x_31, 3, x_30); -x_32 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__16; -x_33 = lean_array_push(x_32, x_2); -x_34 = lean_array_push(x_33, x_3); -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_24); -lean_ctor_set(x_35, 1, x_25); -lean_ctor_set(x_35, 2, x_34); -x_36 = lean_array_push(x_32, x_31); -x_37 = lean_array_push(x_36, x_35); -x_38 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__8; -x_39 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_39, 0, x_24); -lean_ctor_set(x_39, 1, x_38); -lean_ctor_set(x_39, 2, x_37); -lean_inc(x_26); -x_40 = lean_array_push(x_32, x_26); -x_41 = lean_array_push(x_40, x_39); -x_42 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__6; -x_43 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_43, 0, x_24); -lean_ctor_set(x_43, 1, x_42); -lean_ctor_set(x_43, 2, x_41); -x_44 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3___closed__10; -x_45 = lean_array_push(x_44, x_43); -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_24); -lean_ctor_set(x_46, 1, x_25); -lean_ctor_set(x_46, 2, x_45); -x_47 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__17; -lean_inc(x_14); +x_24 = l_Lean_addMacroScope(x_20, x_23, x_15); +x_25 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__13; +x_26 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__17; +lean_inc(x_13); +x_27 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_27, 0, x_13); +lean_ctor_set(x_27, 1, x_25); +lean_ctor_set(x_27, 2, x_24); +lean_ctor_set(x_27, 3, x_26); +x_28 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__18; +x_29 = lean_array_push(x_28, x_1); +x_30 = lean_array_push(x_29, x_2); +x_31 = lean_box(2); +x_32 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__4; +x_33 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +lean_ctor_set(x_33, 2, x_30); +x_34 = lean_array_push(x_28, x_27); +x_35 = lean_array_push(x_34, x_33); +x_36 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__10; +x_37 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_37, 0, x_31); +lean_ctor_set(x_37, 1, x_36); +lean_ctor_set(x_37, 2, x_35); +x_38 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__19; +x_39 = lean_array_push(x_38, x_37); +x_40 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__8; +x_41 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_41, 0, x_31); +lean_ctor_set(x_41, 1, x_40); +lean_ctor_set(x_41, 2, x_39); +x_42 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3___closed__10; +x_43 = lean_array_push(x_42, x_41); +x_44 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_44, 0, x_31); +lean_ctor_set(x_44, 1, x_32); +lean_ctor_set(x_44, 2, x_43); +x_45 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__20; +lean_inc(x_13); +x_46 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_46, 0, x_13); +lean_ctor_set(x_46, 1, x_45); +x_47 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__25; +lean_inc(x_13); x_48 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_48, 0, x_14); +lean_ctor_set(x_48, 0, x_13); lean_ctor_set(x_48, 1, x_47); -x_49 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__22; -lean_inc(x_14); -x_50 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_50, 0, x_14); -lean_ctor_set(x_50, 1, x_49); -x_51 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__29; -lean_inc(x_16); -lean_inc(x_21); -x_52 = l_Lean_addMacroScope(x_21, x_51, x_16); -x_53 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__25; -x_54 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__31; -lean_inc(x_14); -x_55 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_55, 0, x_14); -lean_ctor_set(x_55, 1, x_53); -lean_ctor_set(x_55, 2, x_52); -lean_ctor_set(x_55, 3, x_54); -lean_inc(x_55); -x_56 = lean_array_push(x_44, x_55); +x_49 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__32; +lean_inc(x_15); +lean_inc(x_20); +x_50 = l_Lean_addMacroScope(x_20, x_49, x_15); +x_51 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__28; +x_52 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__34; +lean_inc(x_13); +x_53 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_53, 0, x_13); +lean_ctor_set(x_53, 1, x_51); +lean_ctor_set(x_53, 2, x_50); +lean_ctor_set(x_53, 3, x_52); +lean_inc(x_53); +x_54 = lean_array_push(x_42, x_53); +x_55 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_55, 0, x_31); +lean_ctor_set(x_55, 1, x_32); +lean_ctor_set(x_55, 2, x_54); +x_56 = lean_array_push(x_42, x_55); x_57 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_57, 0, x_24); -lean_ctor_set(x_57, 1, x_25); +lean_ctor_set(x_57, 0, x_31); +lean_ctor_set(x_57, 1, x_32); lean_ctor_set(x_57, 2, x_56); -x_58 = lean_array_push(x_44, x_57); -x_59 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_59, 0, x_24); -lean_ctor_set(x_59, 1, x_25); -lean_ctor_set(x_59, 2, x_58); -x_60 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__32; -lean_inc(x_14); -x_61 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_61, 0, x_14); -lean_ctor_set(x_61, 1, x_60); -x_62 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__33; -x_63 = lean_array_push(x_62, x_50); -lean_inc(x_63); -x_64 = lean_array_push(x_63, x_59); +x_58 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__35; +lean_inc(x_13); +x_59 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_59, 0, x_13); +lean_ctor_set(x_59, 1, x_58); +x_60 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__36; +x_61 = lean_array_push(x_60, x_48); lean_inc(x_61); -x_65 = lean_array_push(x_64, x_61); -x_66 = lean_array_push(x_65, x_55); -x_67 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__21; -x_68 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_68, 0, x_24); -lean_ctor_set(x_68, 1, x_67); -lean_ctor_set(x_68, 2, x_66); -x_69 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__38; -lean_inc(x_16); -lean_inc(x_21); -x_70 = l_Lean_addMacroScope(x_21, x_69, x_16); -x_71 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__36; -x_72 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__40; -lean_inc(x_14); -x_73 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_73, 0, x_14); -lean_ctor_set(x_73, 1, x_71); -lean_ctor_set(x_73, 2, x_70); -lean_ctor_set(x_73, 3, x_72); -lean_inc(x_73); -x_74 = lean_array_push(x_44, x_73); +x_62 = lean_array_push(x_61, x_57); +lean_inc(x_59); +x_63 = lean_array_push(x_62, x_59); +x_64 = lean_array_push(x_63, x_53); +x_65 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__24; +x_66 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_66, 0, x_31); +lean_ctor_set(x_66, 1, x_65); +lean_ctor_set(x_66, 2, x_64); +x_67 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__41; +lean_inc(x_15); +lean_inc(x_20); +x_68 = l_Lean_addMacroScope(x_20, x_67, x_15); +x_69 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__39; +x_70 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__43; +lean_inc(x_13); +x_71 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_71, 0, x_13); +lean_ctor_set(x_71, 1, x_69); +lean_ctor_set(x_71, 2, x_68); +lean_ctor_set(x_71, 3, x_70); +lean_inc(x_71); +x_72 = lean_array_push(x_42, x_71); +x_73 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_73, 0, x_31); +lean_ctor_set(x_73, 1, x_32); +lean_ctor_set(x_73, 2, x_72); +x_74 = lean_array_push(x_42, x_73); x_75 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_75, 0, x_24); -lean_ctor_set(x_75, 1, x_25); +lean_ctor_set(x_75, 0, x_31); +lean_ctor_set(x_75, 1, x_32); lean_ctor_set(x_75, 2, x_74); -x_76 = lean_array_push(x_44, x_75); -x_77 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_77, 0, x_24); -lean_ctor_set(x_77, 1, x_25); -lean_ctor_set(x_77, 2, x_76); -lean_inc(x_63); -x_78 = lean_array_push(x_63, x_77); lean_inc(x_61); -x_79 = lean_array_push(x_78, x_61); -x_80 = lean_array_push(x_79, x_73); -x_81 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_81, 0, x_24); -lean_ctor_set(x_81, 1, x_67); -lean_ctor_set(x_81, 2, x_80); -x_82 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__45; -x_83 = l_Lean_addMacroScope(x_21, x_82, x_16); -x_84 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__43; -x_85 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__47; -x_86 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_86, 0, x_14); -lean_ctor_set(x_86, 1, x_84); -lean_ctor_set(x_86, 2, x_83); -lean_ctor_set(x_86, 3, x_85); -x_87 = lean_array_push(x_44, x_86); +x_76 = lean_array_push(x_61, x_75); +lean_inc(x_59); +x_77 = lean_array_push(x_76, x_59); +x_78 = lean_array_push(x_77, x_71); +x_79 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_79, 0, x_31); +lean_ctor_set(x_79, 1, x_65); +lean_ctor_set(x_79, 2, x_78); +x_80 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__48; +x_81 = l_Lean_addMacroScope(x_20, x_80, x_15); +x_82 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__46; +x_83 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__50; +x_84 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_84, 0, x_13); +lean_ctor_set(x_84, 1, x_82); +lean_ctor_set(x_84, 2, x_81); +lean_ctor_set(x_84, 3, x_83); +x_85 = lean_array_push(x_42, x_84); +x_86 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_86, 0, x_31); +lean_ctor_set(x_86, 1, x_32); +lean_ctor_set(x_86, 2, x_85); +x_87 = lean_array_push(x_42, x_86); x_88 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_88, 0, x_24); -lean_ctor_set(x_88, 1, x_25); +lean_ctor_set(x_88, 0, x_31); +lean_ctor_set(x_88, 1, x_32); lean_ctor_set(x_88, 2, x_87); -x_89 = lean_array_push(x_44, x_88); -x_90 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_90, 0, x_24); -lean_ctor_set(x_90, 1, x_25); -lean_ctor_set(x_90, 2, x_89); -x_91 = lean_array_push(x_63, x_90); -x_92 = lean_array_push(x_91, x_61); -x_93 = lean_array_push(x_92, x_5); -x_94 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_94, 0, x_24); -lean_ctor_set(x_94, 1, x_67); -lean_ctor_set(x_94, 2, x_93); -x_95 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__48; -x_96 = lean_array_push(x_95, x_68); -x_97 = lean_array_push(x_96, x_81); -x_98 = lean_array_push(x_97, x_94); -x_99 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_99, 0, x_24); -lean_ctor_set(x_99, 1, x_25); -lean_ctor_set(x_99, 2, x_98); -x_100 = lean_array_push(x_44, x_99); -x_101 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__19; -x_102 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_102, 0, x_24); -lean_ctor_set(x_102, 1, x_101); -lean_ctor_set(x_102, 2, x_100); -x_103 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__49; -x_104 = lean_array_push(x_103, x_23); -lean_inc(x_26); -x_105 = lean_array_push(x_104, x_26); -x_106 = lean_array_push(x_105, x_26); +x_89 = lean_array_push(x_61, x_88); +x_90 = lean_array_push(x_89, x_59); +x_91 = lean_array_push(x_90, x_4); +x_92 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_92, 0, x_31); +lean_ctor_set(x_92, 1, x_65); +lean_ctor_set(x_92, 2, x_91); +x_93 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__51; +x_94 = lean_array_push(x_93, x_66); +x_95 = lean_array_push(x_94, x_79); +x_96 = lean_array_push(x_95, x_92); +x_97 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_97, 0, x_31); +lean_ctor_set(x_97, 1, x_32); +lean_ctor_set(x_97, 2, x_96); +x_98 = lean_array_push(x_42, x_97); +x_99 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__22; +x_100 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_100, 0, x_31); +lean_ctor_set(x_100, 1, x_99); +lean_ctor_set(x_100, 2, x_98); +x_101 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__52; +x_102 = lean_array_push(x_101, x_22); +x_103 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__6; +x_104 = lean_array_push(x_102, x_103); +x_105 = lean_array_push(x_104, x_103); +x_106 = lean_array_push(x_105, x_44); x_107 = lean_array_push(x_106, x_46); -x_108 = lean_array_push(x_107, x_48); -x_109 = lean_array_push(x_108, x_102); -x_110 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__2; -x_111 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_111, 0, x_24); -lean_ctor_set(x_111, 1, x_110); -lean_ctor_set(x_111, 2, x_109); -x_112 = lean_apply_8(x_4, x_111, x_6, x_7, x_8, x_9, x_10, x_11, x_19); -return x_112; +x_108 = lean_array_push(x_107, x_100); +x_109 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__2; +x_110 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_110, 0, x_31); +lean_ctor_set(x_110, 1, x_109); +lean_ctor_set(x_110, 2, x_108); +x_111 = lean_apply_8(x_3, x_110, x_5, x_6, x_7, x_8, x_9, x_10, x_18); +return x_111; } } static lean_object* _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___closed__1() { @@ -1428,315 +1462,309 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5(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_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, lean_object* x_16) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5(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_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) { _start: { -uint8_t x_17; -x_17 = lean_nat_dec_le(x_7, x_6); -if (x_17 == 0) +uint8_t x_16; +x_16 = lean_nat_dec_le(x_6, x_5); +if (x_16 == 0) { -lean_object* x_18; uint8_t x_19; -x_18 = lean_unsigned_to_nat(0u); -x_19 = lean_nat_dec_eq(x_5, x_18); -if (x_19 == 0) +lean_object* x_17; uint8_t x_18; +x_17 = lean_unsigned_to_nat(0u); +x_18 = lean_nat_dec_eq(x_4, x_17); +if (x_18 == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_28; lean_object* x_29; 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; -x_20 = lean_unsigned_to_nat(1u); -x_21 = lean_nat_sub(x_5, x_20); -lean_dec(x_5); -x_28 = lean_ctor_get(x_9, 1); +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_nat_sub(x_4, x_19); +lean_dec(x_4); +x_27 = lean_ctor_get(x_8, 1); +lean_inc(x_27); +x_28 = lean_ctor_get(x_8, 0); lean_inc(x_28); -x_29 = lean_ctor_get(x_9, 0); +lean_dec(x_8); +x_29 = lean_ctor_get(x_27, 0); lean_inc(x_29); -lean_dec(x_9); -x_30 = lean_ctor_get(x_28, 0); +x_30 = lean_ctor_get(x_27, 1); lean_inc(x_30); -x_31 = lean_ctor_get(x_28, 1); -lean_inc(x_31); -if (lean_is_exclusive(x_28)) { - lean_ctor_release(x_28, 0); - lean_ctor_release(x_28, 1); - x_32 = x_28; +if (lean_is_exclusive(x_27)) { + lean_ctor_release(x_27, 0); + lean_ctor_release(x_27, 1); + x_31 = x_27; } else { - lean_dec_ref(x_28); - x_32 = lean_box(0); + lean_dec_ref(x_27); + x_31 = lean_box(0); } -x_33 = lean_nat_add(x_4, x_6); -x_34 = l_Lean_instInhabitedExpr; -x_35 = lean_array_get(x_34, x_2, x_33); -lean_dec(x_33); -lean_inc(x_15); +x_32 = lean_nat_add(x_3, x_5); +x_33 = l_Lean_instInhabitedExpr; +x_34 = lean_array_get(x_33, x_1, x_32); +lean_dec(x_32); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -lean_inc(x_35); -x_36 = lean_infer_type(x_35, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_36) == 0) +lean_inc(x_11); +lean_inc(x_34); +x_35 = lean_infer_type(x_34, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_35) == 0) { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_36, 0); +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_35, 1); lean_inc(x_37); -x_38 = lean_ctor_get(x_36, 1); -lean_inc(x_38); -lean_dec(x_36); -lean_inc(x_15); +lean_dec(x_35); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -x_39 = l_Lean_Meta_isProp(x_37, x_12, x_13, x_14, x_15, x_38); -if (lean_obj_tag(x_39) == 0) +lean_inc(x_11); +x_38 = l_Lean_Meta_isProp(x_36, x_11, x_12, x_13, x_14, x_37); +if (lean_obj_tag(x_38) == 0) { -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_69; uint8_t x_70; -x_40 = lean_ctor_get(x_39, 0); +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_68; uint8_t x_69; +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_38, 1); lean_inc(x_40); -x_41 = lean_ctor_get(x_39, 1); -lean_inc(x_41); +lean_dec(x_38); +x_68 = l_Lean_Expr_fvarId_x21(x_34); +x_69 = l_Lean_Expr_hasAnyFVar_visit___at_Lean_Expr_containsFVar___spec__1(x_68, x_2); +lean_dec(x_68); +if (x_69 == 0) +{ +uint8_t x_70; +x_70 = lean_unbox(x_39); lean_dec(x_39); -x_69 = l_Lean_Expr_fvarId_x21(x_35); -x_70 = l_Lean_Expr_hasAnyFVar_visit___at_Lean_Expr_containsFVar___spec__1(x_69, x_3); -lean_dec(x_69); if (x_70 == 0) { -uint8_t x_71; -x_71 = lean_unbox(x_40); -lean_dec(x_40); -if (x_71 == 0) -{ -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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; -lean_dec(x_32); -x_72 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___closed__2; -x_73 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_72, x_14, x_15, x_41); -x_74 = lean_ctor_get(x_73, 0); +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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +lean_dec(x_31); +x_71 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___closed__2; +x_72 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_71, x_13, x_14, x_40); +x_73 = lean_ctor_get(x_72, 0); +lean_inc(x_73); +x_74 = lean_ctor_get(x_72, 1); lean_inc(x_74); -x_75 = lean_ctor_get(x_73, 1); -lean_inc(x_75); -lean_dec(x_73); -x_76 = lean_mk_syntax_ident(x_74); -x_77 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___closed__4; -x_78 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_77, x_14, x_15, x_75); -x_79 = lean_ctor_get(x_78, 0); +lean_dec(x_72); +x_75 = lean_mk_syntax_ident(x_73); +x_76 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___closed__4; +x_77 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_76, x_13, x_14, x_74); +x_78 = lean_ctor_get(x_77, 0); +lean_inc(x_78); +x_79 = lean_ctor_get(x_77, 1); lean_inc(x_79); -x_80 = lean_ctor_get(x_78, 1); +lean_dec(x_77); +x_80 = lean_mk_syntax_ident(x_78); +lean_inc(x_75); +x_81 = lean_array_push(x_28, x_75); lean_inc(x_80); -lean_dec(x_78); -x_81 = lean_mk_syntax_ident(x_79); -lean_inc(x_76); -x_82 = lean_array_push(x_29, x_76); -lean_inc(x_81); -x_83 = lean_array_push(x_30, x_81); -lean_inc(x_1); -x_84 = lean_alloc_closure((void*)(l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1), 12, 4); -lean_closure_set(x_84, 0, x_1); -lean_closure_set(x_84, 1, x_76); -lean_closure_set(x_84, 2, x_81); -lean_closure_set(x_84, 3, x_31); +x_82 = lean_array_push(x_29, x_80); +x_83 = lean_alloc_closure((void*)(l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1), 11, 3); +lean_closure_set(x_83, 0, x_75); +lean_closure_set(x_83, 1, x_80); +lean_closure_set(x_83, 2, x_30); +x_84 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_84, 0, x_82); +lean_ctor_set(x_84, 1, x_83); x_85 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_85, 0, x_83); +lean_ctor_set(x_85, 0, x_81); lean_ctor_set(x_85, 1, x_84); -x_86 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_86, 0, x_82); -lean_ctor_set(x_86, 1, x_85); -x_87 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_87, 0, x_86); -x_22 = x_87; -x_23 = x_80; -goto block_27; +x_86 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_86, 0, x_85); +x_21 = x_86; +x_22 = x_79; +goto block_26; } else { -lean_object* x_88; -x_88 = lean_box(0); -x_42 = x_88; -goto block_68; +lean_object* x_87; +x_87 = lean_box(0); +x_41 = x_87; +goto block_67; } } else { -lean_object* x_89; -lean_dec(x_40); -x_89 = lean_box(0); -x_42 = x_89; -goto block_68; +lean_object* x_88; +lean_dec(x_39); +x_88 = lean_box(0); +x_41 = x_88; +goto block_67; } -block_68: +block_67: { -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; lean_object* x_67; -lean_dec(x_42); -lean_inc(x_14); -x_43 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_14, x_15, x_41); -x_44 = lean_ctor_get(x_43, 0); +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; +lean_dec(x_41); +lean_inc(x_13); +x_42 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_13, x_14, x_40); +x_43 = lean_ctor_get(x_42, 0); +lean_inc(x_43); +x_44 = lean_ctor_get(x_42, 1); lean_inc(x_44); -x_45 = lean_ctor_get(x_43, 1); -lean_inc(x_45); -lean_dec(x_43); -x_46 = lean_st_ref_get(x_15, x_45); -x_47 = lean_ctor_get(x_46, 1); -lean_inc(x_47); -lean_dec(x_46); -x_48 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3___closed__9; -x_49 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_49, 0, x_44); -lean_ctor_set(x_49, 1, x_48); -x_50 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3___closed__10; -x_51 = lean_array_push(x_50, x_49); -x_52 = lean_box(2); -x_53 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3___closed__8; -x_54 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_54, 0, x_52); -lean_ctor_set(x_54, 1, x_53); -lean_ctor_set(x_54, 2, x_51); -x_55 = lean_array_push(x_29, x_54); -lean_inc(x_14); -x_56 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_14, x_15, x_47); -x_57 = lean_ctor_get(x_56, 0); +lean_dec(x_42); +x_45 = lean_st_ref_get(x_14, x_44); +x_46 = lean_ctor_get(x_45, 1); +lean_inc(x_46); +lean_dec(x_45); +x_47 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3___closed__9; +x_48 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_48, 0, x_43); +lean_ctor_set(x_48, 1, x_47); +x_49 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3___closed__10; +x_50 = lean_array_push(x_49, x_48); +x_51 = lean_box(2); +x_52 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3___closed__8; +x_53 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_52); +lean_ctor_set(x_53, 2, x_50); +x_54 = lean_array_push(x_28, x_53); +lean_inc(x_13); +x_55 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_13, x_14, x_46); +x_56 = lean_ctor_get(x_55, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_55, 1); lean_inc(x_57); -x_58 = lean_ctor_get(x_56, 1); -lean_inc(x_58); -lean_dec(x_56); -x_59 = lean_st_ref_get(x_15, x_58); -x_60 = lean_ctor_get(x_59, 1); -lean_inc(x_60); -lean_dec(x_59); -x_61 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_61, 0, x_57); -lean_ctor_set(x_61, 1, x_48); -x_62 = lean_array_push(x_50, x_61); -x_63 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_63, 0, x_52); -lean_ctor_set(x_63, 1, x_53); -lean_ctor_set(x_63, 2, x_62); -x_64 = lean_array_push(x_30, x_63); -if (lean_is_scalar(x_32)) { - x_65 = lean_alloc_ctor(0, 2, 0); +lean_dec(x_55); +x_58 = lean_st_ref_get(x_14, x_57); +x_59 = lean_ctor_get(x_58, 1); +lean_inc(x_59); +lean_dec(x_58); +x_60 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_60, 0, x_56); +lean_ctor_set(x_60, 1, x_47); +x_61 = lean_array_push(x_49, x_60); +x_62 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_62, 0, x_51); +lean_ctor_set(x_62, 1, x_52); +lean_ctor_set(x_62, 2, x_61); +x_63 = lean_array_push(x_29, x_62); +if (lean_is_scalar(x_31)) { + x_64 = lean_alloc_ctor(0, 2, 0); } else { - x_65 = x_32; + x_64 = x_31; } -lean_ctor_set(x_65, 0, x_64); -lean_ctor_set(x_65, 1, x_31); -x_66 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_66, 0, x_55); -lean_ctor_set(x_66, 1, x_65); -x_67 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_67, 0, x_66); -x_22 = x_67; -x_23 = x_60; -goto block_27; +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set(x_64, 1, x_30); +x_65 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_65, 0, x_54); +lean_ctor_set(x_65, 1, x_64); +x_66 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_66, 0, x_65); +x_21 = x_66; +x_22 = x_59; +goto block_26; } } else { -uint8_t x_90; -lean_dec(x_35); -lean_dec(x_32); +uint8_t x_89; +lean_dec(x_34); lean_dec(x_31); lean_dec(x_30); lean_dec(x_29); -lean_dec(x_21); -lean_dec(x_15); +lean_dec(x_28); +lean_dec(x_20); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); -lean_dec(x_6); -lean_dec(x_1); -x_90 = !lean_is_exclusive(x_39); -if (x_90 == 0) +lean_dec(x_11); +lean_dec(x_5); +x_89 = !lean_is_exclusive(x_38); +if (x_89 == 0) { -return x_39; +return x_38; } else { -lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_91 = lean_ctor_get(x_39, 0); -x_92 = lean_ctor_get(x_39, 1); -lean_inc(x_92); +lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_90 = lean_ctor_get(x_38, 0); +x_91 = lean_ctor_get(x_38, 1); lean_inc(x_91); -lean_dec(x_39); -x_93 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_93, 0, x_91); -lean_ctor_set(x_93, 1, x_92); -return x_93; +lean_inc(x_90); +lean_dec(x_38); +x_92 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_91); +return x_92; } } } else { -uint8_t x_94; -lean_dec(x_35); -lean_dec(x_32); +uint8_t x_93; +lean_dec(x_34); lean_dec(x_31); lean_dec(x_30); lean_dec(x_29); -lean_dec(x_21); -lean_dec(x_15); +lean_dec(x_28); +lean_dec(x_20); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); -lean_dec(x_6); -lean_dec(x_1); -x_94 = !lean_is_exclusive(x_36); -if (x_94 == 0) +lean_dec(x_11); +lean_dec(x_5); +x_93 = !lean_is_exclusive(x_35); +if (x_93 == 0) { -return x_36; +return x_35; } else { -lean_object* x_95; lean_object* x_96; lean_object* x_97; -x_95 = lean_ctor_get(x_36, 0); -x_96 = lean_ctor_get(x_36, 1); -lean_inc(x_96); +lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_94 = lean_ctor_get(x_35, 0); +x_95 = lean_ctor_get(x_35, 1); lean_inc(x_95); -lean_dec(x_36); -x_97 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_97, 0, x_95); -lean_ctor_set(x_97, 1, x_96); -return x_97; +lean_inc(x_94); +lean_dec(x_35); +x_96 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_96, 0, x_94); +lean_ctor_set(x_96, 1, x_95); +return x_96; } } -block_27: +block_26: { -lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_22, 0); -lean_inc(x_24); -lean_dec(x_22); -x_25 = lean_nat_add(x_6, x_8); -lean_dec(x_6); -x_5 = x_21; -x_6 = x_25; -x_9 = x_24; -x_16 = x_23; +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_21, 0); +lean_inc(x_23); +lean_dec(x_21); +x_24 = lean_nat_add(x_5, x_7); +lean_dec(x_5); +x_4 = x_20; +x_5 = x_24; +x_8 = x_23; +x_15 = x_22; goto _start; } } else { -lean_object* x_98; -lean_dec(x_15); +lean_object* x_97; lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); -lean_dec(x_6); +lean_dec(x_11); lean_dec(x_5); -lean_dec(x_1); -x_98 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_98, 0, x_9); -lean_ctor_set(x_98, 1, x_16); -return x_98; +lean_dec(x_4); +x_97 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_97, 0, x_8); +lean_ctor_set(x_97, 1, x_15); +return x_97; } } else { -lean_object* x_99; -lean_dec(x_15); +lean_object* x_98; lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); -lean_dec(x_6); +lean_dec(x_11); lean_dec(x_5); -lean_dec(x_1); -x_99 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_99, 0, x_9); -lean_ctor_set(x_99, 1, x_16); -return x_99; +lean_dec(x_4); +x_98 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_98, 0, x_8); +lean_ctor_set(x_98, 1, x_15); +return x_98; } } } @@ -1766,7 +1794,18 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Core_betaReduce___lambda__2___boxed), 4, return x_1; } } -static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__3() { +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__5; +x_2 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +return x_2; +} +} +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__4() { _start: { lean_object* x_1; @@ -1774,7 +1813,7 @@ x_1 = lean_alloc_closure((void*)(l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_m return x_1; } } -static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__4() { +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__5() { _start: { lean_object* x_1; @@ -1782,17 +1821,17 @@ x_1 = lean_mk_string_from_bytes("explicit", 8); return x_1; } } -static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__5() { +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3___closed__6; -x_2 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__4; +x_2 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__6() { +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__7() { _start: { lean_object* x_1; @@ -1800,31 +1839,31 @@ x_1 = lean_mk_string_from_bytes("@", 1); return x_1; } } -static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__7() { +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__45; +x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__48; 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_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__8() { +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__7; +x_2 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__8; x_3 = lean_alloc_ctor(1, 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_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__9() { +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__10() { _start: { lean_object* x_1; @@ -1832,44 +1871,44 @@ x_1 = lean_mk_string_from_bytes(",", 1); return x_1; } } -static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__10() { +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(2); -x_2 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__9; +x_2 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__10; x_3 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__1; -x_15 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__2; -lean_inc(x_12); +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__1; +x_14 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__2; lean_inc(x_11); -x_16 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_6, x_14, x_15, x_11, x_12, x_13); -if (lean_obj_tag(x_16) == 0) +lean_inc(x_10); +x_15 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_5, x_13, x_14, x_10, x_11, x_12); +if (lean_obj_tag(x_15) == 0) { -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; 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_17 = lean_ctor_get(x_16, 0); +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; 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_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); +lean_dec(x_15); +x_18 = lean_ctor_get(x_1, 2); lean_inc(x_18); -lean_dec(x_16); -x_19 = lean_ctor_get(x_1, 2); -lean_inc(x_19); -x_20 = lean_unsigned_to_nat(0u); -x_21 = lean_unsigned_to_nat(1u); -lean_inc(x_11); -lean_inc(x_2); -lean_inc(x_19); -x_22 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3(x_19, x_20, x_19, x_21, x_2, x_7, x_8, x_9, x_10, x_11, x_12, x_18); -lean_dec(x_19); +x_19 = lean_unsigned_to_nat(0u); +x_20 = lean_unsigned_to_nat(1u); +x_21 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__5; +lean_inc(x_10); +lean_inc(x_18); +x_22 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3(x_18, x_19, x_18, x_20, x_21, x_6, x_7, x_8, x_9, x_10, x_11, x_17); +lean_dec(x_18); x_23 = lean_ctor_get(x_22, 0); lean_inc(x_23); x_24 = lean_ctor_get(x_22, 1); @@ -1878,13 +1917,10 @@ lean_dec(x_22); x_25 = lean_ctor_get(x_1, 1); lean_inc(x_25); lean_dec(x_1); -lean_inc_n(x_2, 2); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_2); -lean_ctor_set(x_26, 1, x_2); -lean_inc(x_11); +x_26 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__3; +lean_inc(x_10); lean_inc(x_25); -x_27 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__4(x_25, x_20, x_25, x_21, x_26, x_7, x_8, x_9, x_10, x_11, x_12, x_24); +x_27 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__4(x_25, x_19, x_25, x_20, x_26, x_6, x_7, x_8, x_9, x_10, x_11, x_24); x_28 = lean_ctor_get(x_27, 0); lean_inc(x_28); x_29 = lean_ctor_get(x_27, 1); @@ -1896,25 +1932,24 @@ if (x_30 == 0) lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; x_31 = lean_ctor_get(x_28, 0); x_32 = lean_ctor_get(x_28, 1); -x_33 = lean_ctor_get(x_3, 4); +x_33 = lean_ctor_get(x_2, 4); lean_inc(x_33); -lean_dec(x_3); -x_34 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__3; +lean_dec(x_2); +x_34 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__4; lean_ctor_set(x_28, 1, x_34); lean_ctor_set(x_28, 0, x_32); x_35 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_35, 0, x_31); lean_ctor_set(x_35, 1, x_28); -lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); +lean_inc(x_8); lean_inc(x_33); -lean_inc(x_2); -x_36 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5(x_2, x_5, x_17, x_25, x_33, x_20, x_33, x_21, x_35, x_7, x_8, x_9, x_10, x_11, x_12, x_29); +x_36 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5(x_4, x_16, x_25, x_33, x_19, x_33, x_20, x_35, x_6, x_7, x_8, x_9, x_10, x_11, x_29); lean_dec(x_33); lean_dec(x_25); -lean_dec(x_17); +lean_dec(x_16); if (lean_obj_tag(x_36) == 0) { 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; @@ -1933,34 +1968,33 @@ lean_inc(x_41); x_42 = lean_ctor_get(x_38, 1); lean_inc(x_42); lean_dec(x_38); -lean_inc(x_11); -x_43 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_11, x_12, x_39); +lean_inc(x_10); +x_43 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_10, x_11, x_39); x_44 = lean_ctor_get(x_43, 0); lean_inc(x_44); x_45 = lean_ctor_get(x_43, 1); lean_inc(x_45); lean_dec(x_43); -x_46 = lean_st_ref_get(x_12, x_45); +x_46 = lean_st_ref_get(x_11, x_45); x_47 = lean_ctor_get(x_46, 1); lean_inc(x_47); lean_dec(x_46); -x_48 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__6; +x_48 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__7; x_49 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_49, 0, x_44); lean_ctor_set(x_49, 1, x_48); -x_50 = lean_mk_syntax_ident(x_4); -x_51 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__16; +x_50 = lean_mk_syntax_ident(x_3); +x_51 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__18; x_52 = lean_array_push(x_51, x_49); lean_inc(x_50); x_53 = lean_array_push(x_52, x_50); x_54 = lean_box(2); -x_55 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__5; +x_55 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__6; x_56 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_56, 0, x_54); lean_ctor_set(x_56, 1, x_55); lean_ctor_set(x_56, 2, x_53); -lean_inc(x_2); -x_57 = l_Array_append___rarg(x_2, x_40); +x_57 = l_Array_append___rarg(x_21, x_40); x_58 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__4; x_59 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_59, 0, x_54); @@ -1968,19 +2002,19 @@ lean_ctor_set(x_59, 1, x_58); lean_ctor_set(x_59, 2, x_57); x_60 = lean_array_push(x_51, x_56); x_61 = lean_array_push(x_60, x_59); -x_62 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__8; +x_62 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__10; x_63 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_63, 0, x_54); lean_ctor_set(x_63, 1, x_62); lean_ctor_set(x_63, 2, x_61); -lean_inc(x_11); -x_64 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_11, x_12, x_47); +lean_inc(x_10); +x_64 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_10, x_11, x_47); x_65 = lean_ctor_get(x_64, 0); lean_inc(x_65); x_66 = lean_ctor_get(x_64, 1); lean_inc(x_66); lean_dec(x_64); -x_67 = lean_st_ref_get(x_12, x_66); +x_67 = lean_st_ref_get(x_11, x_66); x_68 = lean_ctor_get(x_67, 1); lean_inc(x_68); lean_dec(x_67); @@ -1993,8 +2027,7 @@ x_72 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_72, 0, x_54); lean_ctor_set(x_72, 1, x_55); lean_ctor_set(x_72, 2, x_71); -lean_inc(x_2); -x_73 = l_Array_append___rarg(x_2, x_41); +x_73 = l_Array_append___rarg(x_21, x_41); x_74 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_74, 0, x_54); lean_ctor_set(x_74, 1, x_58); @@ -2011,14 +2044,14 @@ lean_inc(x_78); x_79 = lean_array_push(x_78, x_77); lean_inc(x_23); x_80 = l_Array_append___rarg(x_23, x_79); -lean_inc(x_11); -x_81 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_11, x_12, x_68); +lean_inc(x_10); +x_81 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_10, x_11, x_68); x_82 = lean_ctor_get(x_81, 0); lean_inc(x_82); x_83 = lean_ctor_get(x_81, 1); lean_inc(x_83); lean_dec(x_81); -x_84 = lean_st_ref_get(x_12, x_83); +x_84 = lean_st_ref_get(x_11, x_83); x_85 = lean_ctor_get(x_84, 1); lean_inc(x_85); lean_dec(x_84); @@ -2036,14 +2069,14 @@ lean_ctor_set(x_91, 2, x_89); x_92 = lean_array_push(x_78, x_91); lean_inc(x_23); x_93 = l_Array_append___rarg(x_23, x_92); -lean_inc(x_11); -x_94 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_11, x_12, x_85); +lean_inc(x_10); +x_94 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_10, x_11, x_85); x_95 = lean_ctor_get(x_94, 0); lean_inc(x_95); x_96 = lean_ctor_get(x_94, 1); lean_inc(x_96); lean_dec(x_94); -x_97 = lean_st_ref_get(x_12, x_96); +x_97 = lean_st_ref_get(x_11, x_96); x_98 = lean_ctor_get(x_97, 1); lean_inc(x_98); lean_dec(x_97); @@ -2058,16 +2091,16 @@ lean_ctor_set(x_101, 2, x_100); x_102 = lean_array_push(x_51, x_101); x_103 = lean_array_push(x_102, x_77); x_104 = l_Array_append___rarg(x_23, x_103); -lean_inc(x_11); -x_105 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_11, x_12, x_98); +lean_inc(x_10); +x_105 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_10, x_11, x_98); x_106 = lean_ctor_get(x_105, 0); lean_inc(x_106); x_107 = lean_ctor_get(x_105, 1); lean_inc(x_107); lean_dec(x_105); -x_108 = lean_ctor_get(x_11, 10); +x_108 = lean_ctor_get(x_10, 10); lean_inc(x_108); -x_109 = lean_st_ref_get(x_12, x_107); +x_109 = lean_st_ref_get(x_11, x_107); x_110 = lean_ctor_get(x_109, 0); lean_inc(x_110); x_111 = lean_ctor_get(x_109, 1); @@ -2077,19 +2110,19 @@ x_112 = lean_ctor_get(x_110, 0); lean_inc(x_112); lean_dec(x_110); x_113 = lean_environment_main_module(x_112); -x_114 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__45; +x_114 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__48; lean_inc(x_108); x_115 = l_Lean_addMacroScope(x_113, x_114, x_108); -x_116 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__43; -x_117 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__8; +x_116 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__46; +x_117 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__9; x_118 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_118, 0, x_106); lean_ctor_set(x_118, 1, x_116); lean_ctor_set(x_118, 2, x_115); lean_ctor_set(x_118, 3, x_117); -lean_inc(x_12); lean_inc(x_11); -x_119 = lean_apply_8(x_42, x_118, x_7, x_8, x_9, x_10, x_11, x_12, x_111); +lean_inc(x_10); +x_119 = lean_apply_8(x_42, x_118, x_6, x_7, x_8, x_9, x_10, x_11, x_111); if (lean_obj_tag(x_119) == 0) { lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; size_t x_130; size_t x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; size_t x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; uint8_t x_180; @@ -2098,18 +2131,18 @@ lean_inc(x_120); x_121 = lean_ctor_get(x_119, 1); lean_inc(x_121); lean_dec(x_119); -lean_inc(x_11); -x_122 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_11, x_12, x_121); +lean_inc(x_10); +x_122 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_10, x_11, x_121); x_123 = lean_ctor_get(x_122, 0); lean_inc(x_123); x_124 = lean_ctor_get(x_122, 1); lean_inc(x_124); lean_dec(x_122); -x_125 = lean_st_ref_get(x_12, x_124); +x_125 = lean_st_ref_get(x_11, x_124); x_126 = lean_ctor_get(x_125, 1); lean_inc(x_126); lean_dec(x_125); -x_127 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__22; +x_127 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__25; lean_inc(x_123); x_128 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_128, 0, x_123); @@ -2119,11 +2152,10 @@ x_130 = lean_usize_of_nat(x_129); lean_dec(x_129); x_131 = 0; x_132 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_130, x_131, x_80); -x_133 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__10; +x_133 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__11; x_134 = l_Lean_mkSepArray(x_132, x_133); lean_dec(x_132); -lean_inc(x_2); -x_135 = l_Array_append___rarg(x_2, x_134); +x_135 = l_Array_append___rarg(x_21, x_134); x_136 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_136, 0, x_54); lean_ctor_set(x_136, 1, x_58); @@ -2133,28 +2165,28 @@ x_138 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_138, 0, x_54); lean_ctor_set(x_138, 1, x_58); lean_ctor_set(x_138, 2, x_137); -x_139 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__32; +x_139 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__35; x_140 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_140, 0, x_123); lean_ctor_set(x_140, 1, x_139); -x_141 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__33; +x_141 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__36; x_142 = lean_array_push(x_141, x_128); x_143 = lean_array_push(x_142, x_138); x_144 = lean_array_push(x_143, x_140); x_145 = lean_array_push(x_144, x_120); -x_146 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__21; +x_146 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__24; x_147 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_147, 0, x_54); lean_ctor_set(x_147, 1, x_146); lean_ctor_set(x_147, 2, x_145); -lean_inc(x_11); -x_148 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_11, x_12, x_126); +lean_inc(x_10); +x_148 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_10, x_11, x_126); x_149 = lean_ctor_get(x_148, 0); lean_inc(x_149); x_150 = lean_ctor_get(x_148, 1); lean_inc(x_150); lean_dec(x_148); -x_151 = lean_st_ref_get(x_12, x_150); +x_151 = lean_st_ref_get(x_11, x_150); x_152 = lean_ctor_get(x_151, 0); lean_inc(x_152); x_153 = lean_ctor_get(x_151, 1); @@ -2174,8 +2206,7 @@ lean_dec(x_157); x_159 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_158, x_131, x_93); x_160 = l_Lean_mkSepArray(x_159, x_133); lean_dec(x_159); -lean_inc(x_2); -x_161 = l_Array_append___rarg(x_2, x_160); +x_161 = l_Array_append___rarg(x_21, x_160); x_162 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_162, 0, x_54); lean_ctor_set(x_162, 1, x_58); @@ -2189,11 +2220,11 @@ lean_inc(x_149); x_165 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_165, 0, x_149); lean_ctor_set(x_165, 1, x_139); -x_166 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__29; +x_166 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__32; lean_inc(x_108); x_167 = l_Lean_addMacroScope(x_155, x_166, x_108); -x_168 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__25; -x_169 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__31; +x_168 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__28; +x_169 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__34; x_170 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_170, 0, x_149); lean_ctor_set(x_170, 1, x_168); @@ -2207,14 +2238,14 @@ x_175 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_175, 0, x_54); lean_ctor_set(x_175, 1, x_146); lean_ctor_set(x_175, 2, x_174); -x_176 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_11, x_12, x_153); +x_176 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_10, x_11, x_153); x_177 = lean_ctor_get(x_176, 0); lean_inc(x_177); x_178 = lean_ctor_get(x_176, 1); lean_inc(x_178); lean_dec(x_176); -x_179 = lean_st_ref_get(x_12, x_178); -lean_dec(x_12); +x_179 = lean_st_ref_get(x_11, x_178); +lean_dec(x_11); x_180 = !lean_is_exclusive(x_179); if (x_180 == 0) { @@ -2234,7 +2265,7 @@ lean_dec(x_185); x_187 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_186, x_131, x_104); x_188 = l_Lean_mkSepArray(x_187, x_133); lean_dec(x_187); -x_189 = l_Array_append___rarg(x_2, x_188); +x_189 = l_Array_append___rarg(x_21, x_188); x_190 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_190, 0, x_54); lean_ctor_set(x_190, 1, x_58); @@ -2248,10 +2279,10 @@ lean_inc(x_177); x_193 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_193, 0, x_177); lean_ctor_set(x_193, 1, x_139); -x_194 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__38; +x_194 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__41; x_195 = l_Lean_addMacroScope(x_183, x_194, x_108); -x_196 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__36; -x_197 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__40; +x_196 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__39; +x_197 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__43; x_198 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_198, 0, x_177); lean_ctor_set(x_198, 1, x_196); @@ -2265,7 +2296,7 @@ x_203 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_203, 0, x_54); lean_ctor_set(x_203, 1, x_146); lean_ctor_set(x_203, 2, x_202); -x_204 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__48; +x_204 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__51; x_205 = lean_array_push(x_204, x_147); x_206 = lean_array_push(x_205, x_175); x_207 = lean_array_push(x_206, x_203); @@ -2294,7 +2325,7 @@ lean_dec(x_213); x_215 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_214, x_131, x_104); x_216 = l_Lean_mkSepArray(x_215, x_133); lean_dec(x_215); -x_217 = l_Array_append___rarg(x_2, x_216); +x_217 = l_Array_append___rarg(x_21, x_216); x_218 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_218, 0, x_54); lean_ctor_set(x_218, 1, x_58); @@ -2308,10 +2339,10 @@ lean_inc(x_177); x_221 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_221, 0, x_177); lean_ctor_set(x_221, 1, x_139); -x_222 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__38; +x_222 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__41; x_223 = l_Lean_addMacroScope(x_211, x_222, x_108); -x_224 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__36; -x_225 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__40; +x_224 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__39; +x_225 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__43; x_226 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_226, 0, x_177); lean_ctor_set(x_226, 1, x_224); @@ -2325,7 +2356,7 @@ x_231 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_231, 0, x_54); lean_ctor_set(x_231, 1, x_146); lean_ctor_set(x_231, 2, x_230); -x_232 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__48; +x_232 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__51; x_233 = lean_array_push(x_232, x_147); x_234 = lean_array_push(x_233, x_175); x_235 = lean_array_push(x_234, x_231); @@ -2342,9 +2373,8 @@ lean_dec(x_108); lean_dec(x_104); lean_dec(x_93); lean_dec(x_80); -lean_dec(x_12); lean_dec(x_11); -lean_dec(x_2); +lean_dec(x_10); x_237 = !lean_is_exclusive(x_119); if (x_237 == 0) { @@ -2369,14 +2399,13 @@ else { uint8_t x_241; lean_dec(x_23); -lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_4); -lean_dec(x_2); +lean_dec(x_6); +lean_dec(x_3); x_241 = !lean_is_exclusive(x_36); if (x_241 == 0) { @@ -2405,26 +2434,25 @@ x_246 = lean_ctor_get(x_28, 1); lean_inc(x_246); lean_inc(x_245); lean_dec(x_28); -x_247 = lean_ctor_get(x_3, 4); +x_247 = lean_ctor_get(x_2, 4); lean_inc(x_247); -lean_dec(x_3); -x_248 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__3; +lean_dec(x_2); +x_248 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__4; x_249 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_249, 0, x_246); lean_ctor_set(x_249, 1, x_248); x_250 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_250, 0, x_245); lean_ctor_set(x_250, 1, x_249); -lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); +lean_inc(x_8); lean_inc(x_247); -lean_inc(x_2); -x_251 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5(x_2, x_5, x_17, x_25, x_247, x_20, x_247, x_21, x_250, x_7, x_8, x_9, x_10, x_11, x_12, x_29); +x_251 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5(x_4, x_16, x_25, x_247, x_19, x_247, x_20, x_250, x_6, x_7, x_8, x_9, x_10, x_11, x_29); lean_dec(x_247); lean_dec(x_25); -lean_dec(x_17); +lean_dec(x_16); if (lean_obj_tag(x_251) == 0) { lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; @@ -2443,34 +2471,33 @@ lean_inc(x_256); x_257 = lean_ctor_get(x_253, 1); lean_inc(x_257); lean_dec(x_253); -lean_inc(x_11); -x_258 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_11, x_12, x_254); +lean_inc(x_10); +x_258 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_10, x_11, x_254); x_259 = lean_ctor_get(x_258, 0); lean_inc(x_259); x_260 = lean_ctor_get(x_258, 1); lean_inc(x_260); lean_dec(x_258); -x_261 = lean_st_ref_get(x_12, x_260); +x_261 = lean_st_ref_get(x_11, x_260); x_262 = lean_ctor_get(x_261, 1); lean_inc(x_262); lean_dec(x_261); -x_263 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__6; +x_263 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__7; x_264 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_264, 0, x_259); lean_ctor_set(x_264, 1, x_263); -x_265 = lean_mk_syntax_ident(x_4); -x_266 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__16; +x_265 = lean_mk_syntax_ident(x_3); +x_266 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__18; x_267 = lean_array_push(x_266, x_264); lean_inc(x_265); x_268 = lean_array_push(x_267, x_265); x_269 = lean_box(2); -x_270 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__5; +x_270 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__6; x_271 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_271, 0, x_269); lean_ctor_set(x_271, 1, x_270); lean_ctor_set(x_271, 2, x_268); -lean_inc(x_2); -x_272 = l_Array_append___rarg(x_2, x_255); +x_272 = l_Array_append___rarg(x_21, x_255); x_273 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__4; x_274 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_274, 0, x_269); @@ -2478,19 +2505,19 @@ lean_ctor_set(x_274, 1, x_273); lean_ctor_set(x_274, 2, x_272); x_275 = lean_array_push(x_266, x_271); x_276 = lean_array_push(x_275, x_274); -x_277 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__8; +x_277 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__10; x_278 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_278, 0, x_269); lean_ctor_set(x_278, 1, x_277); lean_ctor_set(x_278, 2, x_276); -lean_inc(x_11); -x_279 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_11, x_12, x_262); +lean_inc(x_10); +x_279 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_10, x_11, x_262); x_280 = lean_ctor_get(x_279, 0); lean_inc(x_280); x_281 = lean_ctor_get(x_279, 1); lean_inc(x_281); lean_dec(x_279); -x_282 = lean_st_ref_get(x_12, x_281); +x_282 = lean_st_ref_get(x_11, x_281); x_283 = lean_ctor_get(x_282, 1); lean_inc(x_283); lean_dec(x_282); @@ -2503,8 +2530,7 @@ x_287 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_287, 0, x_269); lean_ctor_set(x_287, 1, x_270); lean_ctor_set(x_287, 2, x_286); -lean_inc(x_2); -x_288 = l_Array_append___rarg(x_2, x_256); +x_288 = l_Array_append___rarg(x_21, x_256); x_289 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_289, 0, x_269); lean_ctor_set(x_289, 1, x_273); @@ -2521,14 +2547,14 @@ lean_inc(x_293); x_294 = lean_array_push(x_293, x_292); lean_inc(x_23); x_295 = l_Array_append___rarg(x_23, x_294); -lean_inc(x_11); -x_296 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_11, x_12, x_283); +lean_inc(x_10); +x_296 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_10, x_11, x_283); x_297 = lean_ctor_get(x_296, 0); lean_inc(x_297); x_298 = lean_ctor_get(x_296, 1); lean_inc(x_298); lean_dec(x_296); -x_299 = lean_st_ref_get(x_12, x_298); +x_299 = lean_st_ref_get(x_11, x_298); x_300 = lean_ctor_get(x_299, 1); lean_inc(x_300); lean_dec(x_299); @@ -2546,14 +2572,14 @@ lean_ctor_set(x_306, 2, x_304); x_307 = lean_array_push(x_293, x_306); lean_inc(x_23); x_308 = l_Array_append___rarg(x_23, x_307); -lean_inc(x_11); -x_309 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_11, x_12, x_300); +lean_inc(x_10); +x_309 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_10, x_11, x_300); x_310 = lean_ctor_get(x_309, 0); lean_inc(x_310); x_311 = lean_ctor_get(x_309, 1); lean_inc(x_311); lean_dec(x_309); -x_312 = lean_st_ref_get(x_12, x_311); +x_312 = lean_st_ref_get(x_11, x_311); x_313 = lean_ctor_get(x_312, 1); lean_inc(x_313); lean_dec(x_312); @@ -2568,16 +2594,16 @@ lean_ctor_set(x_316, 2, x_315); x_317 = lean_array_push(x_266, x_316); x_318 = lean_array_push(x_317, x_292); x_319 = l_Array_append___rarg(x_23, x_318); -lean_inc(x_11); -x_320 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_11, x_12, x_313); +lean_inc(x_10); +x_320 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_10, x_11, x_313); x_321 = lean_ctor_get(x_320, 0); lean_inc(x_321); x_322 = lean_ctor_get(x_320, 1); lean_inc(x_322); lean_dec(x_320); -x_323 = lean_ctor_get(x_11, 10); +x_323 = lean_ctor_get(x_10, 10); lean_inc(x_323); -x_324 = lean_st_ref_get(x_12, x_322); +x_324 = lean_st_ref_get(x_11, x_322); x_325 = lean_ctor_get(x_324, 0); lean_inc(x_325); x_326 = lean_ctor_get(x_324, 1); @@ -2587,19 +2613,19 @@ x_327 = lean_ctor_get(x_325, 0); lean_inc(x_327); lean_dec(x_325); x_328 = lean_environment_main_module(x_327); -x_329 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__45; +x_329 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__48; lean_inc(x_323); x_330 = l_Lean_addMacroScope(x_328, x_329, x_323); -x_331 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__43; -x_332 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__8; +x_331 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__46; +x_332 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__9; x_333 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_333, 0, x_321); lean_ctor_set(x_333, 1, x_331); lean_ctor_set(x_333, 2, x_330); lean_ctor_set(x_333, 3, x_332); -lean_inc(x_12); lean_inc(x_11); -x_334 = lean_apply_8(x_257, x_333, x_7, x_8, x_9, x_10, x_11, x_12, x_326); +lean_inc(x_10); +x_334 = lean_apply_8(x_257, x_333, x_6, x_7, x_8, x_9, x_10, x_11, x_326); if (lean_obj_tag(x_334) == 0) { lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; size_t x_345; size_t x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; size_t x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; size_t x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; @@ -2608,18 +2634,18 @@ lean_inc(x_335); x_336 = lean_ctor_get(x_334, 1); lean_inc(x_336); lean_dec(x_334); -lean_inc(x_11); -x_337 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_11, x_12, x_336); +lean_inc(x_10); +x_337 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_10, x_11, x_336); x_338 = lean_ctor_get(x_337, 0); lean_inc(x_338); x_339 = lean_ctor_get(x_337, 1); lean_inc(x_339); lean_dec(x_337); -x_340 = lean_st_ref_get(x_12, x_339); +x_340 = lean_st_ref_get(x_11, x_339); x_341 = lean_ctor_get(x_340, 1); lean_inc(x_341); lean_dec(x_340); -x_342 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__22; +x_342 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__25; lean_inc(x_338); x_343 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_343, 0, x_338); @@ -2629,11 +2655,10 @@ x_345 = lean_usize_of_nat(x_344); lean_dec(x_344); x_346 = 0; x_347 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_345, x_346, x_295); -x_348 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__10; +x_348 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__11; x_349 = l_Lean_mkSepArray(x_347, x_348); lean_dec(x_347); -lean_inc(x_2); -x_350 = l_Array_append___rarg(x_2, x_349); +x_350 = l_Array_append___rarg(x_21, x_349); x_351 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_351, 0, x_269); lean_ctor_set(x_351, 1, x_273); @@ -2643,28 +2668,28 @@ x_353 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_353, 0, x_269); lean_ctor_set(x_353, 1, x_273); lean_ctor_set(x_353, 2, x_352); -x_354 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__32; +x_354 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__35; x_355 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_355, 0, x_338); lean_ctor_set(x_355, 1, x_354); -x_356 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__33; +x_356 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__36; x_357 = lean_array_push(x_356, x_343); x_358 = lean_array_push(x_357, x_353); x_359 = lean_array_push(x_358, x_355); x_360 = lean_array_push(x_359, x_335); -x_361 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__21; +x_361 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__24; x_362 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_362, 0, x_269); lean_ctor_set(x_362, 1, x_361); lean_ctor_set(x_362, 2, x_360); -lean_inc(x_11); -x_363 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_11, x_12, x_341); +lean_inc(x_10); +x_363 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_10, x_11, x_341); x_364 = lean_ctor_get(x_363, 0); lean_inc(x_364); x_365 = lean_ctor_get(x_363, 1); lean_inc(x_365); lean_dec(x_363); -x_366 = lean_st_ref_get(x_12, x_365); +x_366 = lean_st_ref_get(x_11, x_365); x_367 = lean_ctor_get(x_366, 0); lean_inc(x_367); x_368 = lean_ctor_get(x_366, 1); @@ -2684,8 +2709,7 @@ lean_dec(x_372); x_374 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_373, x_346, x_308); x_375 = l_Lean_mkSepArray(x_374, x_348); lean_dec(x_374); -lean_inc(x_2); -x_376 = l_Array_append___rarg(x_2, x_375); +x_376 = l_Array_append___rarg(x_21, x_375); x_377 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_377, 0, x_269); lean_ctor_set(x_377, 1, x_273); @@ -2699,11 +2723,11 @@ lean_inc(x_364); x_380 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_380, 0, x_364); lean_ctor_set(x_380, 1, x_354); -x_381 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__29; +x_381 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__32; lean_inc(x_323); x_382 = l_Lean_addMacroScope(x_370, x_381, x_323); -x_383 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__25; -x_384 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__31; +x_383 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__28; +x_384 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__34; x_385 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_385, 0, x_364); lean_ctor_set(x_385, 1, x_383); @@ -2717,14 +2741,14 @@ x_390 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_390, 0, x_269); lean_ctor_set(x_390, 1, x_361); lean_ctor_set(x_390, 2, x_389); -x_391 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_11, x_12, x_368); +x_391 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_10, x_11, x_368); x_392 = lean_ctor_get(x_391, 0); lean_inc(x_392); x_393 = lean_ctor_get(x_391, 1); lean_inc(x_393); lean_dec(x_391); -x_394 = lean_st_ref_get(x_12, x_393); -lean_dec(x_12); +x_394 = lean_st_ref_get(x_11, x_393); +lean_dec(x_11); x_395 = lean_ctor_get(x_394, 0); lean_inc(x_395); x_396 = lean_ctor_get(x_394, 1); @@ -2751,7 +2775,7 @@ lean_dec(x_401); x_403 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_402, x_346, x_319); x_404 = l_Lean_mkSepArray(x_403, x_348); lean_dec(x_403); -x_405 = l_Array_append___rarg(x_2, x_404); +x_405 = l_Array_append___rarg(x_21, x_404); x_406 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_406, 0, x_269); lean_ctor_set(x_406, 1, x_273); @@ -2765,10 +2789,10 @@ lean_inc(x_392); x_409 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_409, 0, x_392); lean_ctor_set(x_409, 1, x_354); -x_410 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__38; +x_410 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__41; x_411 = l_Lean_addMacroScope(x_399, x_410, x_323); -x_412 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__36; -x_413 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__40; +x_412 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__39; +x_413 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__43; x_414 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_414, 0, x_392); lean_ctor_set(x_414, 1, x_412); @@ -2782,7 +2806,7 @@ x_419 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_419, 0, x_269); lean_ctor_set(x_419, 1, x_361); lean_ctor_set(x_419, 2, x_418); -x_420 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__48; +x_420 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__51; x_421 = lean_array_push(x_420, x_362); x_422 = lean_array_push(x_421, x_390); x_423 = lean_array_push(x_422, x_419); @@ -2802,9 +2826,8 @@ lean_dec(x_323); lean_dec(x_319); lean_dec(x_308); lean_dec(x_295); -lean_dec(x_12); lean_dec(x_11); -lean_dec(x_2); +lean_dec(x_10); x_425 = lean_ctor_get(x_334, 0); lean_inc(x_425); x_426 = lean_ctor_get(x_334, 1); @@ -2831,14 +2854,13 @@ else { lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_dec(x_23); -lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_4); -lean_dec(x_2); +lean_dec(x_6); +lean_dec(x_3); x_429 = lean_ctor_get(x_251, 0); lean_inc(x_429); x_430 = lean_ctor_get(x_251, 1); @@ -2865,29 +2887,28 @@ return x_432; else { uint8_t x_433; -lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_4); +lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_433 = !lean_is_exclusive(x_16); +x_433 = !lean_is_exclusive(x_15); if (x_433 == 0) { -return x_16; +return x_15; } else { lean_object* x_434; lean_object* x_435; lean_object* x_436; -x_434 = lean_ctor_get(x_16, 0); -x_435 = lean_ctor_get(x_16, 1); +x_434 = lean_ctor_get(x_15, 0); +x_435 = lean_ctor_get(x_15, 1); lean_inc(x_435); lean_inc(x_434); -lean_dec(x_16); +lean_dec(x_15); x_436 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_436, 0, x_434); lean_ctor_set(x_436, 1, x_435); @@ -2896,153 +2917,139 @@ return x_436; } } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6(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_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6(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_object* x_9, lean_object* x_10) { _start: { -if (lean_obj_tag(x_3) == 0) +if (lean_obj_tag(x_2) == 0) { -lean_object* x_12; -lean_dec(x_10); +lean_object* x_11; lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_2); +lean_dec(x_4); lean_dec(x_1); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_4); -lean_ctor_set(x_12, 1, x_11); -return x_12; +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_3); +lean_ctor_set(x_11, 1, x_10); +return x_11; } else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = lean_ctor_get(x_3, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_3, 1); -lean_inc(x_14); -lean_dec(x_3); -lean_inc(x_5); +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_2, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_2, 1); lean_inc(x_13); -x_15 = l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__1(x_13, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_15) == 0) +lean_dec(x_2); +lean_inc(x_4); +lean_inc(x_12); +x_14 = l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__1(x_12, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_16 = lean_ctor_get(x_15, 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; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +lean_dec(x_14); +x_17 = lean_ctor_get(x_15, 0); lean_inc(x_17); -lean_dec(x_15); -x_18 = lean_ctor_get(x_16, 0); +x_18 = lean_ctor_get(x_17, 2); lean_inc(x_18); -x_19 = lean_ctor_get(x_18, 2); -lean_inc(x_19); -lean_dec(x_18); -lean_inc(x_2); +lean_dec(x_17); lean_inc(x_1); -x_20 = lean_alloc_closure((void*)(l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___boxed), 13, 4); -lean_closure_set(x_20, 0, x_1); -lean_closure_set(x_20, 1, x_2); -lean_closure_set(x_20, 2, x_16); -lean_closure_set(x_20, 3, x_13); -lean_inc(x_10); +x_19 = lean_alloc_closure((void*)(l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___boxed), 12, 3); +lean_closure_set(x_19, 0, x_1); +lean_closure_set(x_19, 1, x_15); +lean_closure_set(x_19, 2, x_12); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_21 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Elab_Deriving_mkInductArgNames___spec__2___rarg(x_19, x_20, x_5, x_6, x_7, x_8, x_9, x_10, x_17); -if (lean_obj_tag(x_21) == 0) +lean_inc(x_4); +x_20 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Elab_Deriving_mkInductArgNames___spec__2___rarg(x_18, x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_16); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_21, 0); +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); lean_inc(x_22); -x_23 = lean_ctor_get(x_21, 1); -lean_inc(x_23); -lean_dec(x_21); -x_24 = l_Array_append___rarg(x_4, x_22); -x_3 = x_14; -x_4 = x_24; -x_11 = x_23; +lean_dec(x_20); +x_23 = l_Array_append___rarg(x_3, x_21); +x_2 = x_13; +x_3 = x_23; +x_10 = x_22; goto _start; } else { -uint8_t x_26; -lean_dec(x_14); -lean_dec(x_10); +uint8_t x_25; +lean_dec(x_13); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_2); +lean_dec(x_3); lean_dec(x_1); -x_26 = !lean_is_exclusive(x_21); -if (x_26 == 0) +x_25 = !lean_is_exclusive(x_20); +if (x_25 == 0) { -return x_21; +return x_20; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_21, 0); -x_28 = lean_ctor_get(x_21, 1); -lean_inc(x_28); +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_20, 0); +x_27 = lean_ctor_get(x_20, 1); lean_inc(x_27); -lean_dec(x_21); -x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -return x_29; +lean_inc(x_26); +lean_dec(x_20); +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +return x_28; } } } else { -uint8_t x_30; -lean_dec(x_14); +uint8_t x_29; lean_dec(x_13); -lean_dec(x_10); +lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_2); +lean_dec(x_3); lean_dec(x_1); -x_30 = !lean_is_exclusive(x_15); -if (x_30 == 0) +x_29 = !lean_is_exclusive(x_14); +if (x_29 == 0) { -return x_15; +return x_14; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_15, 0); -x_32 = lean_ctor_get(x_15, 1); -lean_inc(x_32); +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_14, 0); +x_31 = lean_ctor_get(x_14, 1); lean_inc(x_31); -lean_dec(x_15); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -return x_33; -} -} +lean_inc(x_30); +lean_dec(x_14); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +return x_32; } } } -static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(0u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; } } LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Ord_mkMatch_mkAlts(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) { @@ -3051,8 +3058,8 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Ord_mkMatch_mkAlts(lean_object* x_ lean_object* x_9; lean_object* x_10; lean_object* x_11; x_9 = lean_ctor_get(x_1, 4); lean_inc(x_9); -x_10 = l_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___closed__1; -x_11 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6(x_1, x_10, x_9, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_10 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__5; +x_11 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6(x_1, x_9, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8); if (lean_obj_tag(x_11) == 0) { uint8_t x_12; @@ -3162,19 +3169,19 @@ lean_dec(x_3); return x_13; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___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_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, lean_object* x_16) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___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_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) { _start: { -lean_object* x_17; -x_17 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -lean_dec(x_11); +lean_object* x_16; +x_16 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); lean_dec(x_10); -lean_dec(x_8); +lean_dec(x_9); lean_dec(x_7); -lean_dec(x_4); +lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); -return x_17; +lean_dec(x_1); +return x_16; } } LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___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) { @@ -3191,27 +3198,37 @@ lean_dec(x_2); return x_9; } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___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, lean_object* x_7, 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_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___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, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_14; -x_14 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_5); -return x_14; +lean_object* x_13; +x_13 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_4); +return x_13; } } -static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkMatch___closed__1() { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_Ord_mkMatch___spec__1(size_t x_1, size_t x_2, lean_object* x_3) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = lean_box(2); -x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__4; -x_3 = l_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___closed__1; -x_4 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_4, 0, x_1); -lean_ctor_set(x_4, 1, x_2); -lean_ctor_set(x_4, 2, x_3); -return x_4; +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = 1; +x_9 = lean_usize_add(x_2, x_8); +x_10 = lean_array_uset(x_7, x_2, x_5); +x_2 = x_9; +x_3 = x_10; +goto _start; +} } } LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Ord_mkMatch(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_object* x_9) { @@ -3260,11 +3277,11 @@ x_24 = lean_array_get_size(x_11); x_25 = lean_usize_of_nat(x_24); lean_dec(x_24); x_26 = 0; -x_27 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_25, x_26, x_11); -x_28 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__10; +x_27 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_Ord_mkMatch___spec__1(x_25, x_26, x_11); +x_28 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__11; x_29 = l_Lean_mkSepArray(x_27, x_28); lean_dec(x_27); -x_30 = l_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___closed__1; +x_30 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__5; x_31 = l_Array_append___rarg(x_30, x_29); x_32 = lean_box(2); x_33 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__4; @@ -3272,7 +3289,7 @@ x_34 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_34, 0, x_32); lean_ctor_set(x_34, 1, x_33); lean_ctor_set(x_34, 2, x_31); -x_35 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__17; +x_35 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__20; x_36 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_36, 0, x_17); lean_ctor_set(x_36, 1, x_35); @@ -3283,14 +3300,14 @@ lean_ctor_set(x_38, 1, x_33); lean_ctor_set(x_38, 2, x_37); x_39 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3___closed__10; x_40 = lean_array_push(x_39, x_38); -x_41 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__19; +x_41 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__22; x_42 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_42, 0, x_32); lean_ctor_set(x_42, 1, x_41); lean_ctor_set(x_42, 2, x_40); -x_43 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__49; +x_43 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__52; x_44 = lean_array_push(x_43, x_23); -x_45 = l_Lean_Elab_Deriving_Ord_mkMatch___closed__1; +x_45 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__6; x_46 = lean_array_push(x_44, x_45); x_47 = lean_array_push(x_46, x_45); x_48 = lean_array_push(x_47, x_34); @@ -3319,11 +3336,11 @@ x_56 = lean_array_get_size(x_11); x_57 = lean_usize_of_nat(x_56); lean_dec(x_56); x_58 = 0; -x_59 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_57, x_58, x_11); -x_60 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__10; +x_59 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_Ord_mkMatch___spec__1(x_57, x_58, x_11); +x_60 = l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__11; x_61 = l_Lean_mkSepArray(x_59, x_60); lean_dec(x_59); -x_62 = l_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___closed__1; +x_62 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__5; x_63 = l_Array_append___rarg(x_62, x_61); x_64 = lean_box(2); x_65 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__4; @@ -3331,7 +3348,7 @@ x_66 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_66, 0, x_64); lean_ctor_set(x_66, 1, x_65); lean_ctor_set(x_66, 2, x_63); -x_67 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__17; +x_67 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__20; x_68 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_68, 0, x_17); lean_ctor_set(x_68, 1, x_67); @@ -3342,14 +3359,14 @@ lean_ctor_set(x_70, 1, x_65); lean_ctor_set(x_70, 2, x_69); x_71 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__3___closed__10; x_72 = lean_array_push(x_71, x_70); -x_73 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__19; +x_73 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__22; x_74 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_74, 0, x_64); lean_ctor_set(x_74, 1, x_73); lean_ctor_set(x_74, 2, x_72); -x_75 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__49; +x_75 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__52; x_76 = lean_array_push(x_75, x_55); -x_77 = l_Lean_Elab_Deriving_Ord_mkMatch___closed__1; +x_77 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__6; x_78 = lean_array_push(x_76, x_77); x_79 = lean_array_push(x_78, x_77); x_80 = lean_array_push(x_79, x_66); @@ -3393,6 +3410,18 @@ return x_89; } } } +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_Ord_mkMatch___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_Ord_mkMatch___spec__1(x_4, x_5, x_3); +return x_6; +} +} static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__1() { _start: { @@ -3487,8 +3516,8 @@ static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__49; -x_2 = l_Lean_Elab_Deriving_Ord_mkMatch___closed__1; +x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__52; +x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__6; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -3498,7 +3527,7 @@ static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___c { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__11; -x_2 = l_Lean_Elab_Deriving_Ord_mkMatch___closed__1; +x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__6; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -3587,7 +3616,7 @@ static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___c _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__26; +x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__29; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } @@ -3596,7 +3625,7 @@ static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__26; +x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__29; x_2 = lean_unsigned_to_nat(0u); x_3 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__22; x_4 = lean_alloc_ctor(0, 3, 0); @@ -3611,7 +3640,7 @@ static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___c { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__27; +x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__30; x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); @@ -3721,7 +3750,7 @@ lean_ctor_set(x_201, 1, x_200); lean_ctor_set(x_201, 2, x_199); x_202 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__12; x_203 = lean_array_push(x_202, x_201); -x_204 = l_Lean_Elab_Deriving_Ord_mkMatch___closed__1; +x_204 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__6; x_205 = lean_array_push(x_203, x_204); x_206 = lean_array_push(x_205, x_204); x_207 = lean_array_push(x_206, x_204); @@ -3736,7 +3765,7 @@ x_211 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_211, 0, x_184); lean_ctor_set(x_211, 1, x_210); x_212 = lean_mk_syntax_ident(x_2); -x_213 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__16; +x_213 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__18; x_214 = lean_array_push(x_213, x_212); x_215 = lean_array_push(x_214, x_204); x_216 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__16; @@ -3744,7 +3773,7 @@ x_217 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_217, 0, x_196); lean_ctor_set(x_217, 1, x_216); lean_ctor_set(x_217, 2, x_215); -x_218 = l_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___closed__1; +x_218 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__5; x_219 = l_Array_append___rarg(x_218, x_14); x_220 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_220, 0, x_196); @@ -3755,7 +3784,7 @@ lean_inc(x_184); x_222 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_222, 0, x_184); lean_ctor_set(x_222, 1, x_221); -x_223 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__27; +x_223 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__30; x_224 = l_Lean_addMacroScope(x_191, x_223, x_186); x_225 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__23; x_226 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__25; @@ -3788,7 +3817,7 @@ x_238 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__28; x_239 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_239, 0, x_184); lean_ctor_set(x_239, 1, x_238); -x_240 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__48; +x_240 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__51; x_241 = lean_array_push(x_240, x_239); x_242 = lean_array_push(x_241, x_5); x_243 = lean_array_push(x_242, x_204); @@ -3853,7 +3882,7 @@ lean_ctor_set(x_273, 1, x_272); lean_ctor_set(x_273, 2, x_271); x_274 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__12; x_275 = lean_array_push(x_274, x_273); -x_276 = l_Lean_Elab_Deriving_Ord_mkMatch___closed__1; +x_276 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__6; x_277 = lean_array_push(x_275, x_276); x_278 = lean_array_push(x_277, x_276); x_279 = lean_array_push(x_278, x_276); @@ -3868,7 +3897,7 @@ x_283 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_283, 0, x_184); lean_ctor_set(x_283, 1, x_282); x_284 = lean_mk_syntax_ident(x_2); -x_285 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__16; +x_285 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__18; x_286 = lean_array_push(x_285, x_284); x_287 = lean_array_push(x_286, x_276); x_288 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__16; @@ -3876,7 +3905,7 @@ x_289 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_289, 0, x_268); lean_ctor_set(x_289, 1, x_288); lean_ctor_set(x_289, 2, x_287); -x_290 = l_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___closed__1; +x_290 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__5; x_291 = l_Array_append___rarg(x_290, x_14); x_292 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_292, 0, x_268); @@ -3887,7 +3916,7 @@ lean_inc(x_184); x_294 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_294, 0, x_184); lean_ctor_set(x_294, 1, x_293); -x_295 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__27; +x_295 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__30; x_296 = l_Lean_addMacroScope(x_263, x_295, x_186); x_297 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__23; x_298 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__25; @@ -3920,7 +3949,7 @@ x_310 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__28; x_311 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_311, 0, x_184); lean_ctor_set(x_311, 1, x_310); -x_312 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__48; +x_312 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__51; x_313 = lean_array_push(x_312, x_311); x_314 = lean_array_push(x_313, x_5); x_315 = lean_array_push(x_314, x_276); @@ -4031,7 +4060,7 @@ lean_ctor_set(x_41, 1, x_33); lean_ctor_set(x_41, 2, x_40); x_42 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__12; x_43 = lean_array_push(x_42, x_34); -x_44 = l_Lean_Elab_Deriving_Ord_mkMatch___closed__1; +x_44 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__6; x_45 = lean_array_push(x_43, x_44); x_46 = lean_array_push(x_45, x_44); x_47 = lean_array_push(x_46, x_41); @@ -4046,7 +4075,7 @@ x_51 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_51, 0, x_17); lean_ctor_set(x_51, 1, x_50); x_52 = lean_mk_syntax_ident(x_2); -x_53 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__16; +x_53 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__18; x_54 = lean_array_push(x_53, x_52); x_55 = lean_array_push(x_54, x_44); x_56 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__16; @@ -4054,7 +4083,7 @@ x_57 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_57, 0, x_29); lean_ctor_set(x_57, 1, x_56); lean_ctor_set(x_57, 2, x_55); -x_58 = l_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___closed__1; +x_58 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__5; x_59 = l_Array_append___rarg(x_58, x_14); x_60 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_60, 0, x_29); @@ -4065,7 +4094,7 @@ lean_inc(x_17); x_62 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_62, 0, x_17); lean_ctor_set(x_62, 1, x_61); -x_63 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__27; +x_63 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__30; x_64 = l_Lean_addMacroScope(x_24, x_63, x_19); x_65 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__23; x_66 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__25; @@ -4098,7 +4127,7 @@ x_78 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__28; x_79 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_79, 0, x_17); lean_ctor_set(x_79, 1, x_78); -x_80 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__48; +x_80 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__51; x_81 = lean_array_push(x_80, x_79); x_82 = lean_array_push(x_81, x_5); x_83 = lean_array_push(x_82, x_44); @@ -4179,7 +4208,7 @@ lean_ctor_set(x_120, 1, x_112); lean_ctor_set(x_120, 2, x_119); x_121 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__12; x_122 = lean_array_push(x_121, x_113); -x_123 = l_Lean_Elab_Deriving_Ord_mkMatch___closed__1; +x_123 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__6; x_124 = lean_array_push(x_122, x_123); x_125 = lean_array_push(x_124, x_123); x_126 = lean_array_push(x_125, x_120); @@ -4194,7 +4223,7 @@ x_130 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_130, 0, x_17); lean_ctor_set(x_130, 1, x_129); x_131 = lean_mk_syntax_ident(x_2); -x_132 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__16; +x_132 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__18; x_133 = lean_array_push(x_132, x_131); x_134 = lean_array_push(x_133, x_123); x_135 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__16; @@ -4202,7 +4231,7 @@ x_136 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_136, 0, x_108); lean_ctor_set(x_136, 1, x_135); lean_ctor_set(x_136, 2, x_134); -x_137 = l_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___closed__1; +x_137 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__5; x_138 = l_Array_append___rarg(x_137, x_14); x_139 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_139, 0, x_108); @@ -4213,7 +4242,7 @@ lean_inc(x_17); x_141 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_141, 0, x_17); lean_ctor_set(x_141, 1, x_140); -x_142 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__27; +x_142 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__30; x_143 = l_Lean_addMacroScope(x_103, x_142, x_19); x_144 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__23; x_145 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__25; @@ -4246,7 +4275,7 @@ x_157 = l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__28; x_158 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_158, 0, x_17); lean_ctor_set(x_158, 1, x_157); -x_159 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__48; +x_159 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__51; x_160 = lean_array_push(x_159, x_158); x_161 = lean_array_push(x_160, x_5); x_162 = lean_array_push(x_161, x_123); @@ -4746,7 +4775,7 @@ static lean_object* _init_l_Lean_Elab_Deriving_Ord_mkMutualBlock___closed__4() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); x_2 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__4; -x_3 = l_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___closed__1; +x_3 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__5; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -4773,7 +4802,7 @@ x_10 = lean_array_get_size(x_9); lean_dec(x_9); x_11 = lean_unsigned_to_nat(0u); x_12 = lean_unsigned_to_nat(1u); -x_13 = l_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___closed__1; +x_13 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__5; lean_inc(x_7); lean_inc(x_6); lean_inc(x_10); @@ -5875,7 +5904,7 @@ lean_dec(x_1); return x_9; } } -static lean_object* _init_l_Lean_Elab_Deriving_Ord_initFn____x40_Lean_Elab_Deriving_Ord___hyg_2782____closed__1() { +static lean_object* _init_l_Lean_Elab_Deriving_Ord_initFn____x40_Lean_Elab_Deriving_Ord___hyg_2936____closed__1() { _start: { lean_object* x_1; @@ -5883,12 +5912,12 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Deriving_Ord_mkOrdInstanceHandler), return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Ord_initFn____x40_Lean_Elab_Deriving_Ord___hyg_2782_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Ord_initFn____x40_Lean_Elab_Deriving_Ord___hyg_2936_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; x_2 = l_Lean_Elab_Deriving_Ord_mkOrdHeader___closed__2; -x_3 = l_Lean_Elab_Deriving_Ord_initFn____x40_Lean_Elab_Deriving_Ord___hyg_2782____closed__1; +x_3 = l_Lean_Elab_Deriving_Ord_initFn____x40_Lean_Elab_Deriving_Ord___hyg_2936____closed__1; x_4 = l_Lean_Elab_registerBuiltinDerivingHandler(x_2, x_3, x_1); if (lean_obj_tag(x_4) == 0) { @@ -6075,6 +6104,12 @@ l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___la lean_mark_persistent(l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__48); l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__49 = _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__49(); lean_mark_persistent(l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__49); +l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__50 = _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__50(); +lean_mark_persistent(l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__50); +l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__51 = _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__51(); +lean_mark_persistent(l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__51); +l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__52 = _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__52(); +lean_mark_persistent(l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__52); l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___closed__1 = _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___closed__1(); lean_mark_persistent(l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___closed__1); l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___closed__2 = _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___closed__2(); @@ -6103,10 +6138,8 @@ l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda_ lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__9); l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__10 = _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__10(); lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__10); -l_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___closed__1 = _init_l_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___closed__1(); -lean_mark_persistent(l_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___closed__1); -l_Lean_Elab_Deriving_Ord_mkMatch___closed__1 = _init_l_Lean_Elab_Deriving_Ord_mkMatch___closed__1(); -lean_mark_persistent(l_Lean_Elab_Deriving_Ord_mkMatch___closed__1); +l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__11 = _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__11(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__6___lambda__2___closed__11); l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__1 = _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__1(); lean_mark_persistent(l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__1); l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__2 = _init_l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__2(); @@ -6195,9 +6228,9 @@ l___private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkOrdInstanceCmds__ lean_mark_persistent(l___private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkOrdInstanceCmds___closed__9); l___private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkOrdInstanceCmds___closed__10 = _init_l___private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkOrdInstanceCmds___closed__10(); lean_mark_persistent(l___private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkOrdInstanceCmds___closed__10); -l_Lean_Elab_Deriving_Ord_initFn____x40_Lean_Elab_Deriving_Ord___hyg_2782____closed__1 = _init_l_Lean_Elab_Deriving_Ord_initFn____x40_Lean_Elab_Deriving_Ord___hyg_2782____closed__1(); -lean_mark_persistent(l_Lean_Elab_Deriving_Ord_initFn____x40_Lean_Elab_Deriving_Ord___hyg_2782____closed__1); -res = l_Lean_Elab_Deriving_Ord_initFn____x40_Lean_Elab_Deriving_Ord___hyg_2782_(lean_io_mk_world()); +l_Lean_Elab_Deriving_Ord_initFn____x40_Lean_Elab_Deriving_Ord___hyg_2936____closed__1 = _init_l_Lean_Elab_Deriving_Ord_initFn____x40_Lean_Elab_Deriving_Ord___hyg_2936____closed__1(); +lean_mark_persistent(l_Lean_Elab_Deriving_Ord_initFn____x40_Lean_Elab_Deriving_Ord___hyg_2936____closed__1); +res = l_Lean_Elab_Deriving_Ord_initFn____x40_Lean_Elab_Deriving_Ord___hyg_2936_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/Deriving/Repr.c b/stage0/stdlib/Lean/Elab/Deriving/Repr.c index fd3d71e62d6..f2f0991dbb0 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/Repr.c +++ b/stage0/stdlib/Lean/Elab/Deriving/Repr.c @@ -26,6 +26,7 @@ lean_object* l_Lean_registerTraceClass(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Repr_mkReprInstanceHandler(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__5; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_Repr_mkBodyForInduct___spec__1(size_t, size_t, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__14; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__18; @@ -46,6 +47,7 @@ static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForIndu static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__1; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Deriving_Repr_mkReprInstanceHandler___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__1; +lean_object* lean_array_uset(lean_object*, size_t, lean_object*); static lean_object* l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__17; static lean_object* l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__4; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__4; @@ -70,6 +72,7 @@ static lean_object* l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__25; lean_object* l_Lean_MessageData_ofList(lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__37; +static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__55; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__13; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__6; static lean_object* l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__9; @@ -80,6 +83,7 @@ lean_object* l_List_head_x21___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__12; static lean_object* l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__19; +uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__1(lean_object*, lean_object*, 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_throwError___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Deriving_mkLocalInstanceLetDecls(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -91,12 +95,11 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Repr_mkMutualBlock(lean_object*, l static lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForInduct___closed__1; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__30; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForInduct(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_Deriving_Repr_mkBodyForStruct___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isInductive___at_Lean_Elab_Deriving_Repr_mkReprInstanceHandler___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__4; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__42; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__53; -static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3319____closed__1; static lean_object* l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__11; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__3; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__18; @@ -110,12 +113,13 @@ static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForIndu static lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForInduct___closed__3; static lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForInduct___closed__5; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__1___closed__3; +static lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__16; lean_object* l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__28; static lean_object* l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__5; static lean_object* l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__22; -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__18; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__14; static lean_object* l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__28; @@ -127,8 +131,9 @@ lean_object* l_Lean_Elab_Deriving_mkInstanceCmds(lean_object*, lean_object*, lea LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForStruct(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__29; lean_object* lean_nat_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_Repr_mkBodyForInduct___spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__25; -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__4; static lean_object* l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__20; static lean_object* l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__14; @@ -193,6 +198,7 @@ static lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___clos static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__21; static lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__7; static lean_object* l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__23; +static lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__18; lean_object* l_Lean_Elab_Deriving_mkContext(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__15; size_t lean_usize_of_nat(lean_object*); @@ -256,8 +262,9 @@ static lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___clos static lean_object* l_Lean_Elab_Deriving_Repr_mkMutualBlock___closed__1; static lean_object* l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__3; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__24; -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3(lean_object*, lean_object*, 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_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_forallTelescopeReducing___at_Lean_Elab_Deriving_mkInductArgNames___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3494____closed__1; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__11; static lean_object* l___private_Lean_Elab_Deriving_Repr_0__Lean_Elab_Deriving_Repr_mkReprInstanceCmds___closed__5; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__31; @@ -275,6 +282,7 @@ LEAN_EXPORT lean_object* l_Lean_isInductive___at_Lean_Elab_Deriving_Repr_mkReprI static lean_object* l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__11; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__5; static lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__1; +static lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__17; static lean_object* l___private_Lean_Elab_Deriving_Repr_0__Lean_Elab_Deriving_Repr_mkReprInstanceCmds___closed__3; static lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -290,9 +298,8 @@ static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForIndu static lean_object* l_Lean_Elab_Deriving_Repr_mkMutualBlock___closed__2; lean_object* l_Lean_mkConst(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__16; -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Deriving_Repr_0__Lean_Elab_Deriving_Repr_mkReprInstanceCmds___closed__4; -static lean_object* l_Lean_Elab_Deriving_Repr_mkMutualBlock___closed__4; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__3; static lean_object* l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__8; static lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__4; @@ -304,11 +311,11 @@ static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyFo static lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__9; static lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__1___closed__4; lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_isStructure(lean_object*, lean_object*); lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3319_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3494_(lean_object*); static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__1() { _start: { @@ -1798,390 +1805,390 @@ x_1 = lean_mk_string_from_bytes("line", 4); return x_1; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3(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_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, 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_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3(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_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, 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) { _start: { -uint8_t x_22; -x_22 = lean_nat_dec_le(x_12, x_11); -if (x_22 == 0) +uint8_t x_23; +x_23 = lean_nat_dec_le(x_13, x_12); +if (x_23 == 0) { -lean_object* x_23; uint8_t x_24; -x_23 = lean_unsigned_to_nat(0u); -x_24 = lean_nat_dec_eq(x_10, x_23); -if (x_24 == 0) +lean_object* x_24; uint8_t x_25; +x_24 = lean_unsigned_to_nat(0u); +x_25 = lean_nat_dec_eq(x_11, x_24); +if (x_25 == 0) { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t 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; uint8_t x_38; -x_25 = lean_unsigned_to_nat(1u); -x_26 = lean_nat_sub(x_10, x_25); -lean_dec(x_10); -x_27 = lean_ctor_get(x_14, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_14, 1); +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; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; +x_26 = lean_unsigned_to_nat(1u); +x_27 = lean_nat_sub(x_11, x_26); +lean_dec(x_11); +x_28 = lean_ctor_get(x_15, 0); lean_inc(x_28); -lean_dec(x_14); -x_29 = l_Lean_instInhabitedName; -x_30 = lean_array_get(x_29, x_1, x_11); -x_31 = 1; -lean_inc(x_30); -x_32 = l_Lean_Name_toString(x_30, x_31); -x_33 = lean_box(2); -x_34 = l_Lean_Syntax_mkStrLit(x_32, x_33); -lean_dec(x_32); -x_35 = lean_nat_add(x_2, x_11); -x_36 = l_Lean_instInhabitedExpr; -x_37 = lean_array_get(x_36, x_4, x_35); -lean_dec(x_35); -x_38 = lean_unbox(x_28); -if (x_38 == 0) -{ -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; uint8_t x_77; lean_object* x_78; -lean_inc(x_19); -x_39 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_19, x_20, x_21); -x_40 = lean_ctor_get(x_39, 0); -lean_inc(x_40); -x_41 = lean_ctor_get(x_39, 1); +x_29 = lean_ctor_get(x_15, 1); +lean_inc(x_29); +lean_dec(x_15); +x_30 = l_Lean_instInhabitedName; +x_31 = lean_array_get(x_30, x_1, x_12); +x_32 = 1; +lean_inc(x_31); +x_33 = l_Lean_Name_toString(x_31, x_32); +x_34 = lean_box(2); +x_35 = l_Lean_Syntax_mkStrLit(x_33, x_34); +lean_dec(x_33); +x_36 = lean_nat_add(x_2, x_12); +x_37 = l_Lean_instInhabitedExpr; +x_38 = lean_array_get(x_37, x_4, x_36); +lean_dec(x_36); +x_39 = lean_unbox(x_29); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; uint8_t x_78; lean_object* x_79; +lean_inc(x_20); +x_40 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_20, x_21, x_22); +x_41 = lean_ctor_get(x_40, 0); lean_inc(x_41); -lean_dec(x_39); -x_42 = lean_ctor_get(x_19, 10); +x_42 = lean_ctor_get(x_40, 1); lean_inc(x_42); -x_43 = lean_st_ref_get(x_20, x_41); -x_44 = lean_ctor_get(x_43, 0); -lean_inc(x_44); -x_45 = lean_ctor_get(x_43, 1); +lean_dec(x_40); +x_43 = lean_ctor_get(x_20, 10); +lean_inc(x_43); +x_44 = lean_st_ref_get(x_21, x_42); +x_45 = lean_ctor_get(x_44, 0); lean_inc(x_45); -lean_dec(x_43); -x_46 = lean_ctor_get(x_44, 0); +x_46 = lean_ctor_get(x_44, 1); lean_inc(x_46); lean_dec(x_44); -x_47 = lean_environment_main_module(x_46); -x_48 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__3; -lean_inc(x_40); -x_49 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_49, 0, x_40); -lean_ctor_set(x_49, 1, x_48); -x_50 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__1; -lean_inc(x_40); -x_51 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_51, 0, x_40); -lean_ctor_set(x_51, 1, x_50); -x_52 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__18; -x_53 = lean_array_push(x_52, x_51); -x_54 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__6; -x_55 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_55, 0, x_33); -lean_ctor_set(x_55, 1, x_54); -lean_ctor_set(x_55, 2, x_53); -x_56 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__4; -x_57 = lean_array_push(x_56, x_27); -lean_inc(x_49); -x_58 = lean_array_push(x_57, x_49); -x_59 = lean_array_push(x_58, x_55); -x_60 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__2; -x_61 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_61, 0, x_33); -lean_ctor_set(x_61, 1, x_60); -lean_ctor_set(x_61, 2, x_59); -x_62 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__2; -x_63 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__3; -lean_inc(x_5); -x_64 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_64, 0, x_62); -lean_ctor_set(x_64, 1, x_5); -lean_ctor_set(x_64, 2, x_63); -x_65 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__4; +x_47 = lean_ctor_get(x_45, 0); +lean_inc(x_47); +lean_dec(x_45); +x_48 = lean_environment_main_module(x_47); +x_49 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__3; +lean_inc(x_41); +x_50 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_50, 0, x_41); +lean_ctor_set(x_50, 1, x_49); +x_51 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__1; +lean_inc(x_41); +x_52 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_52, 0, x_41); +lean_ctor_set(x_52, 1, x_51); +x_53 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__18; +x_54 = lean_array_push(x_53, x_52); +x_55 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__6; +x_56 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_56, 0, x_34); +lean_ctor_set(x_56, 1, x_55); +lean_ctor_set(x_56, 2, x_54); +x_57 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__4; +x_58 = lean_array_push(x_57, x_28); +lean_inc(x_50); +x_59 = lean_array_push(x_58, x_50); +x_60 = lean_array_push(x_59, x_56); +x_61 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__2; +x_62 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_62, 0, x_34); +lean_ctor_set(x_62, 1, x_61); +lean_ctor_set(x_62, 2, x_60); +x_63 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__2; +x_64 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__3; lean_inc(x_6); -x_66 = lean_name_mk_string(x_6, x_65); -x_67 = l_Lean_addMacroScope(x_47, x_66, x_42); +x_65 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_6); +lean_ctor_set(x_65, 2, x_64); +x_66 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__4; lean_inc(x_7); -x_68 = lean_name_mk_string(x_7, x_65); +x_67 = lean_name_mk_string(x_7, x_66); +x_68 = l_Lean_addMacroScope(x_48, x_67, x_43); lean_inc(x_8); -x_69 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_69, 0, x_68); -lean_ctor_set(x_69, 1, x_8); +x_69 = lean_name_mk_string(x_8, x_66); lean_inc(x_9); -x_70 = lean_alloc_ctor(1, 2, 0); +x_70 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_70, 0, x_69); lean_ctor_set(x_70, 1, x_9); -x_71 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_71, 0, x_40); -lean_ctor_set(x_71, 1, x_64); -lean_ctor_set(x_71, 2, x_67); -lean_ctor_set(x_71, 3, x_70); -x_72 = lean_array_push(x_56, x_61); -x_73 = lean_array_push(x_72, x_49); -x_74 = lean_array_push(x_73, x_71); -x_75 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_75, 0, x_33); -lean_ctor_set(x_75, 1, x_60); -lean_ctor_set(x_75, 2, x_74); -x_76 = lean_box(0); -x_77 = lean_unbox(x_28); -lean_dec(x_28); +lean_inc(x_10); +x_71 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set(x_71, 1, x_10); +x_72 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_72, 0, x_41); +lean_ctor_set(x_72, 1, x_65); +lean_ctor_set(x_72, 2, x_68); +lean_ctor_set(x_72, 3, x_71); +x_73 = lean_array_push(x_57, x_62); +x_74 = lean_array_push(x_73, x_50); +x_75 = lean_array_push(x_74, x_72); +x_76 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_76, 0, x_34); +lean_ctor_set(x_76, 1, x_61); +lean_ctor_set(x_76, 2, x_75); +x_77 = lean_box(0); +x_78 = lean_unbox(x_29); +lean_dec(x_29); +lean_inc(x_21); lean_inc(x_20); lean_inc(x_19); lean_inc(x_18); -lean_inc(x_17); lean_inc(x_3); +lean_inc(x_10); lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_5); -x_78 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1(x_34, x_5, x_8, x_9, x_30, x_3, x_37, x_75, x_77, x_76, x_15, x_16, x_17, x_18, x_19, x_20, x_45); -if (lean_obj_tag(x_78) == 0) -{ -lean_object* x_79; -x_79 = lean_ctor_get(x_78, 0); -lean_inc(x_79); +lean_inc(x_6); +x_79 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1(x_35, x_6, x_9, x_10, x_31, x_3, x_38, x_76, x_78, x_77, x_16, x_17, x_18, x_19, x_20, x_21, x_46); if (lean_obj_tag(x_79) == 0) { -uint8_t x_80; -lean_dec(x_26); +lean_object* x_80; +x_80 = lean_ctor_get(x_79, 0); +lean_inc(x_80); +if (lean_obj_tag(x_80) == 0) +{ +uint8_t x_81; +lean_dec(x_27); +lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_11); +lean_dec(x_12); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); -x_80 = !lean_is_exclusive(x_78); -if (x_80 == 0) +x_81 = !lean_is_exclusive(x_79); +if (x_81 == 0) { -lean_object* x_81; lean_object* x_82; -x_81 = lean_ctor_get(x_78, 0); -lean_dec(x_81); +lean_object* x_82; lean_object* x_83; x_82 = lean_ctor_get(x_79, 0); -lean_inc(x_82); -lean_dec(x_79); -lean_ctor_set(x_78, 0, x_82); -return x_78; +lean_dec(x_82); +x_83 = lean_ctor_get(x_80, 0); +lean_inc(x_83); +lean_dec(x_80); +lean_ctor_set(x_79, 0, x_83); +return x_79; } else { -lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_83 = lean_ctor_get(x_78, 1); -lean_inc(x_83); -lean_dec(x_78); -x_84 = lean_ctor_get(x_79, 0); +lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_84 = lean_ctor_get(x_79, 1); lean_inc(x_84); lean_dec(x_79); -x_85 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_85, 0, x_84); -lean_ctor_set(x_85, 1, x_83); -return x_85; +x_85 = lean_ctor_get(x_80, 0); +lean_inc(x_85); +lean_dec(x_80); +x_86 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_86, 0, x_85); +lean_ctor_set(x_86, 1, x_84); +return x_86; } } else { -lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_86 = lean_ctor_get(x_78, 1); -lean_inc(x_86); -lean_dec(x_78); -x_87 = lean_ctor_get(x_79, 0); +lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_87 = lean_ctor_get(x_79, 1); lean_inc(x_87); lean_dec(x_79); -x_88 = lean_nat_add(x_11, x_13); -lean_dec(x_11); -x_10 = x_26; -x_11 = x_88; -x_14 = x_87; -x_21 = x_86; +x_88 = lean_ctor_get(x_80, 0); +lean_inc(x_88); +lean_dec(x_80); +x_89 = lean_nat_add(x_12, x_14); +lean_dec(x_12); +x_11 = x_27; +x_12 = x_89; +x_15 = x_88; +x_22 = x_87; goto _start; } } else { -uint8_t x_90; -lean_dec(x_26); +uint8_t x_91; +lean_dec(x_27); +lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_11); +lean_dec(x_12); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); -x_90 = !lean_is_exclusive(x_78); -if (x_90 == 0) +x_91 = !lean_is_exclusive(x_79); +if (x_91 == 0) { -return x_78; +return x_79; } else { -lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_91 = lean_ctor_get(x_78, 0); -x_92 = lean_ctor_get(x_78, 1); +lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_92 = lean_ctor_get(x_79, 0); +x_93 = lean_ctor_get(x_79, 1); +lean_inc(x_93); lean_inc(x_92); -lean_inc(x_91); -lean_dec(x_78); -x_93 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_93, 0, x_91); -lean_ctor_set(x_93, 1, x_92); -return x_93; +lean_dec(x_79); +x_94 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_94, 0, x_92); +lean_ctor_set(x_94, 1, x_93); +return x_94; } } } else { -uint8_t x_94; lean_object* x_95; lean_object* x_96; -lean_dec(x_28); -x_94 = 0; -x_95 = lean_box(0); +uint8_t x_95; lean_object* x_96; lean_object* x_97; +lean_dec(x_29); +x_95 = 0; +x_96 = lean_box(0); +lean_inc(x_21); lean_inc(x_20); lean_inc(x_19); lean_inc(x_18); -lean_inc(x_17); lean_inc(x_3); +lean_inc(x_10); lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_5); -x_96 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1(x_34, x_5, x_8, x_9, x_30, x_3, x_37, x_27, x_94, x_95, x_15, x_16, x_17, x_18, x_19, x_20, x_21); -if (lean_obj_tag(x_96) == 0) -{ -lean_object* x_97; -x_97 = lean_ctor_get(x_96, 0); -lean_inc(x_97); +lean_inc(x_6); +x_97 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1(x_35, x_6, x_9, x_10, x_31, x_3, x_38, x_28, x_95, x_96, x_16, x_17, x_18, x_19, x_20, x_21, x_22); if (lean_obj_tag(x_97) == 0) { -uint8_t x_98; -lean_dec(x_26); +lean_object* x_98; +x_98 = lean_ctor_get(x_97, 0); +lean_inc(x_98); +if (lean_obj_tag(x_98) == 0) +{ +uint8_t x_99; +lean_dec(x_27); +lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_11); +lean_dec(x_12); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); -x_98 = !lean_is_exclusive(x_96); -if (x_98 == 0) +x_99 = !lean_is_exclusive(x_97); +if (x_99 == 0) { -lean_object* x_99; lean_object* x_100; -x_99 = lean_ctor_get(x_96, 0); -lean_dec(x_99); +lean_object* x_100; lean_object* x_101; x_100 = lean_ctor_get(x_97, 0); -lean_inc(x_100); -lean_dec(x_97); -lean_ctor_set(x_96, 0, x_100); -return x_96; +lean_dec(x_100); +x_101 = lean_ctor_get(x_98, 0); +lean_inc(x_101); +lean_dec(x_98); +lean_ctor_set(x_97, 0, x_101); +return x_97; } else { -lean_object* x_101; lean_object* x_102; lean_object* x_103; -x_101 = lean_ctor_get(x_96, 1); -lean_inc(x_101); -lean_dec(x_96); -x_102 = lean_ctor_get(x_97, 0); +lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_102 = lean_ctor_get(x_97, 1); lean_inc(x_102); lean_dec(x_97); -x_103 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_103, 0, x_102); -lean_ctor_set(x_103, 1, x_101); -return x_103; +x_103 = lean_ctor_get(x_98, 0); +lean_inc(x_103); +lean_dec(x_98); +x_104 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_104, 0, x_103); +lean_ctor_set(x_104, 1, x_102); +return x_104; } } else { -lean_object* x_104; lean_object* x_105; lean_object* x_106; -x_104 = lean_ctor_get(x_96, 1); -lean_inc(x_104); -lean_dec(x_96); -x_105 = lean_ctor_get(x_97, 0); +lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_105 = lean_ctor_get(x_97, 1); lean_inc(x_105); lean_dec(x_97); -x_106 = lean_nat_add(x_11, x_13); -lean_dec(x_11); -x_10 = x_26; -x_11 = x_106; -x_14 = x_105; -x_21 = x_104; +x_106 = lean_ctor_get(x_98, 0); +lean_inc(x_106); +lean_dec(x_98); +x_107 = lean_nat_add(x_12, x_14); +lean_dec(x_12); +x_11 = x_27; +x_12 = x_107; +x_15 = x_106; +x_22 = x_105; goto _start; } } else { -uint8_t x_108; -lean_dec(x_26); +uint8_t x_109; +lean_dec(x_27); +lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_11); +lean_dec(x_12); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); -x_108 = !lean_is_exclusive(x_96); -if (x_108 == 0) +x_109 = !lean_is_exclusive(x_97); +if (x_109 == 0) { -return x_96; +return x_97; } else { -lean_object* x_109; lean_object* x_110; lean_object* x_111; -x_109 = lean_ctor_get(x_96, 0); -x_110 = lean_ctor_get(x_96, 1); +lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_110 = lean_ctor_get(x_97, 0); +x_111 = lean_ctor_get(x_97, 1); +lean_inc(x_111); lean_inc(x_110); -lean_inc(x_109); -lean_dec(x_96); -x_111 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_111, 0, x_109); -lean_ctor_set(x_111, 1, x_110); -return x_111; +lean_dec(x_97); +x_112 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_112, 0, x_110); +lean_ctor_set(x_112, 1, x_111); +return x_112; } } } } else { -lean_object* x_112; +lean_object* x_113; +lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); -lean_dec(x_17); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); -x_112 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_112, 0, x_14); -lean_ctor_set(x_112, 1, x_21); -return x_112; +x_113 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_113, 0, x_15); +lean_ctor_set(x_113, 1, x_22); +return x_113; } } else { -lean_object* x_113; +lean_object* x_114; +lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); -lean_dec(x_17); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); -x_113 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_113, 0, x_14); -lean_ctor_set(x_113, 1, x_21); -return x_113; +x_114 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_114, 0, x_15); +lean_ctor_set(x_114, 1, x_22); +return x_114; } } } @@ -2240,217 +2247,217 @@ x_1 = lean_mk_string_from_bytes("\" }\"", 4); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForStruct___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, uint8_t 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, lean_object* x_16, lean_object* x_17) { +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForStruct___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, lean_object* x_9, uint8_t 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) { _start: { -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_18 = lean_array_get_size(x_1); -x_19 = lean_box(x_9); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_8); -lean_ctor_set(x_20, 1, x_19); -x_21 = lean_unsigned_to_nat(0u); -x_22 = lean_unsigned_to_nat(1u); +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_19 = lean_array_get_size(x_1); +x_20 = lean_box(x_10); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_9); +lean_ctor_set(x_21, 1, x_20); +x_22 = lean_unsigned_to_nat(0u); +x_23 = lean_unsigned_to_nat(1u); +lean_inc(x_17); lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_18); -lean_inc_n(x_7, 2); +lean_inc(x_19); +lean_inc_n(x_8, 2); +lean_inc(x_7); lean_inc(x_6); -lean_inc(x_5); -x_23 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3(x_1, x_2, x_3, x_4, x_21, x_5, x_6, x_7, x_7, x_18, x_21, x_18, x_22, x_20, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -lean_dec(x_18); -if (lean_obj_tag(x_23) == 0) +x_24 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3(x_1, x_2, x_3, x_4, x_5, x_22, x_6, x_7, x_8, x_8, x_19, x_22, x_19, x_23, x_21, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_19); +if (lean_obj_tag(x_24) == 0) { -lean_object* 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_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); +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_object* x_32; uint8_t x_33; +x_25 = lean_ctor_get(x_24, 0); lean_inc(x_25); -lean_dec(x_23); -x_26 = lean_ctor_get(x_24, 0); +x_26 = lean_ctor_get(x_24, 1); lean_inc(x_26); lean_dec(x_24); -lean_inc(x_15); -x_27 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_15, x_16, x_25); -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); +x_27 = lean_ctor_get(x_25, 0); +lean_inc(x_27); +lean_dec(x_25); +lean_inc(x_16); +x_28 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_16, x_17, x_26); +x_29 = lean_ctor_get(x_28, 0); lean_inc(x_29); -lean_dec(x_27); -x_30 = lean_ctor_get(x_15, 10); +x_30 = lean_ctor_get(x_28, 1); lean_inc(x_30); -lean_dec(x_15); -x_31 = lean_st_ref_get(x_16, x_29); +lean_dec(x_28); +x_31 = lean_ctor_get(x_16, 10); +lean_inc(x_31); lean_dec(x_16); -x_32 = !lean_is_exclusive(x_31); -if (x_32 == 0) +x_32 = lean_st_ref_get(x_17, x_30); +lean_dec(x_17); +x_33 = !lean_is_exclusive(x_32); +if (x_33 == 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; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_33 = lean_ctor_get(x_31, 0); -x_34 = lean_ctor_get(x_33, 0); -lean_inc(x_34); -lean_dec(x_33); -x_35 = lean_environment_main_module(x_34); -x_36 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__1___closed__4; -x_37 = lean_name_mk_string(x_5, x_36); -x_38 = l_Lean_addMacroScope(x_35, x_37, x_30); -x_39 = lean_name_mk_string(x_6, x_36); -lean_inc(x_7); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_7); -x_41 = lean_alloc_ctor(1, 2, 0); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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_34 = lean_ctor_get(x_32, 0); +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +lean_dec(x_34); +x_36 = lean_environment_main_module(x_35); +x_37 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__1___closed__4; +x_38 = lean_name_mk_string(x_6, x_37); +x_39 = l_Lean_addMacroScope(x_36, x_38, x_31); +x_40 = lean_name_mk_string(x_7, x_37); +lean_inc(x_8); +x_41 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_41, 0, x_40); -lean_ctor_set(x_41, 1, x_7); -x_42 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__1___closed__3; -lean_inc(x_28); -x_43 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_43, 0, x_28); -lean_ctor_set(x_43, 1, x_42); -lean_ctor_set(x_43, 2, x_38); -lean_ctor_set(x_43, 3, x_41); -x_44 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__1___closed__5; -lean_inc(x_28); -x_45 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_45, 0, x_28); -lean_ctor_set(x_45, 1, x_44); -x_46 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__18; -x_47 = lean_array_push(x_46, x_45); -x_48 = lean_box(2); -x_49 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__6; -x_50 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_50, 0, x_48); -lean_ctor_set(x_50, 1, x_49); -lean_ctor_set(x_50, 2, x_47); -x_51 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__1___closed__6; -x_52 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_52, 0, x_28); -lean_ctor_set(x_52, 1, x_51); -x_53 = lean_array_push(x_46, x_52); -x_54 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_54, 0, x_48); -lean_ctor_set(x_54, 1, x_49); -lean_ctor_set(x_54, 2, x_53); -x_55 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__4; -x_56 = lean_array_push(x_55, x_50); -x_57 = lean_array_push(x_56, x_26); -x_58 = lean_array_push(x_57, x_54); -x_59 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__13; -x_60 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_60, 0, x_48); -lean_ctor_set(x_60, 1, x_59); -lean_ctor_set(x_60, 2, x_58); -x_61 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__26; -x_62 = lean_array_push(x_61, x_43); -x_63 = lean_array_push(x_62, x_60); -x_64 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__9; -x_65 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_65, 0, x_48); -lean_ctor_set(x_65, 1, x_64); -lean_ctor_set(x_65, 2, x_63); -lean_ctor_set(x_31, 0, x_65); -return x_31; +lean_ctor_set(x_41, 1, x_8); +x_42 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_8); +x_43 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__1___closed__3; +lean_inc(x_29); +x_44 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_44, 0, x_29); +lean_ctor_set(x_44, 1, x_43); +lean_ctor_set(x_44, 2, x_39); +lean_ctor_set(x_44, 3, x_42); +x_45 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__1___closed__5; +lean_inc(x_29); +x_46 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_46, 0, x_29); +lean_ctor_set(x_46, 1, x_45); +x_47 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__18; +x_48 = lean_array_push(x_47, x_46); +x_49 = lean_box(2); +x_50 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__6; +x_51 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +lean_ctor_set(x_51, 2, x_48); +x_52 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__1___closed__6; +x_53 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_53, 0, x_29); +lean_ctor_set(x_53, 1, x_52); +x_54 = lean_array_push(x_47, x_53); +x_55 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_55, 0, x_49); +lean_ctor_set(x_55, 1, x_50); +lean_ctor_set(x_55, 2, x_54); +x_56 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__4; +x_57 = lean_array_push(x_56, x_51); +x_58 = lean_array_push(x_57, x_27); +x_59 = lean_array_push(x_58, x_55); +x_60 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__13; +x_61 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_61, 0, x_49); +lean_ctor_set(x_61, 1, x_60); +lean_ctor_set(x_61, 2, x_59); +x_62 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__26; +x_63 = lean_array_push(x_62, x_44); +x_64 = lean_array_push(x_63, x_61); +x_65 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__9; +x_66 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_66, 0, x_49); +lean_ctor_set(x_66, 1, x_65); +lean_ctor_set(x_66, 2, x_64); +lean_ctor_set(x_32, 0, x_66); +return x_32; } else { -lean_object* x_66; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_66 = lean_ctor_get(x_31, 0); -x_67 = lean_ctor_get(x_31, 1); -lean_inc(x_67); -lean_inc(x_66); -lean_dec(x_31); -x_68 = lean_ctor_get(x_66, 0); +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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_67 = lean_ctor_get(x_32, 0); +x_68 = lean_ctor_get(x_32, 1); lean_inc(x_68); -lean_dec(x_66); -x_69 = lean_environment_main_module(x_68); -x_70 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__1___closed__4; -x_71 = lean_name_mk_string(x_5, x_70); -x_72 = l_Lean_addMacroScope(x_69, x_71, x_30); -x_73 = lean_name_mk_string(x_6, x_70); -lean_inc(x_7); -x_74 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_74, 0, x_73); -lean_ctor_set(x_74, 1, x_7); -x_75 = lean_alloc_ctor(1, 2, 0); +lean_inc(x_67); +lean_dec(x_32); +x_69 = lean_ctor_get(x_67, 0); +lean_inc(x_69); +lean_dec(x_67); +x_70 = lean_environment_main_module(x_69); +x_71 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__1___closed__4; +x_72 = lean_name_mk_string(x_6, x_71); +x_73 = l_Lean_addMacroScope(x_70, x_72, x_31); +x_74 = lean_name_mk_string(x_7, x_71); +lean_inc(x_8); +x_75 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_75, 0, x_74); -lean_ctor_set(x_75, 1, x_7); -x_76 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__1___closed__3; -lean_inc(x_28); -x_77 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_77, 0, x_28); -lean_ctor_set(x_77, 1, x_76); -lean_ctor_set(x_77, 2, x_72); -lean_ctor_set(x_77, 3, x_75); -x_78 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__1___closed__5; -lean_inc(x_28); -x_79 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_79, 0, x_28); -lean_ctor_set(x_79, 1, x_78); -x_80 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__18; -x_81 = lean_array_push(x_80, x_79); -x_82 = lean_box(2); -x_83 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__6; -x_84 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_84, 0, x_82); -lean_ctor_set(x_84, 1, x_83); -lean_ctor_set(x_84, 2, x_81); -x_85 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__1___closed__6; -x_86 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_86, 0, x_28); -lean_ctor_set(x_86, 1, x_85); -x_87 = lean_array_push(x_80, x_86); -x_88 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_88, 0, x_82); -lean_ctor_set(x_88, 1, x_83); -lean_ctor_set(x_88, 2, x_87); -x_89 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__4; -x_90 = lean_array_push(x_89, x_84); -x_91 = lean_array_push(x_90, x_26); -x_92 = lean_array_push(x_91, x_88); -x_93 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__13; -x_94 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_94, 0, x_82); -lean_ctor_set(x_94, 1, x_93); -lean_ctor_set(x_94, 2, x_92); -x_95 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__26; -x_96 = lean_array_push(x_95, x_77); -x_97 = lean_array_push(x_96, x_94); -x_98 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__9; -x_99 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_99, 0, x_82); -lean_ctor_set(x_99, 1, x_98); -lean_ctor_set(x_99, 2, x_97); -x_100 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_100, 0, x_99); -lean_ctor_set(x_100, 1, x_67); -return x_100; +lean_ctor_set(x_75, 1, x_8); +x_76 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_76, 0, x_75); +lean_ctor_set(x_76, 1, x_8); +x_77 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__1___closed__3; +lean_inc(x_29); +x_78 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_78, 0, x_29); +lean_ctor_set(x_78, 1, x_77); +lean_ctor_set(x_78, 2, x_73); +lean_ctor_set(x_78, 3, x_76); +x_79 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__1___closed__5; +lean_inc(x_29); +x_80 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_80, 0, x_29); +lean_ctor_set(x_80, 1, x_79); +x_81 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__18; +x_82 = lean_array_push(x_81, x_80); +x_83 = lean_box(2); +x_84 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__6; +x_85 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_85, 0, x_83); +lean_ctor_set(x_85, 1, x_84); +lean_ctor_set(x_85, 2, x_82); +x_86 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__1___closed__6; +x_87 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_87, 0, x_29); +lean_ctor_set(x_87, 1, x_86); +x_88 = lean_array_push(x_81, x_87); +x_89 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_89, 0, x_83); +lean_ctor_set(x_89, 1, x_84); +lean_ctor_set(x_89, 2, x_88); +x_90 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__4; +x_91 = lean_array_push(x_90, x_85); +x_92 = lean_array_push(x_91, x_27); +x_93 = lean_array_push(x_92, x_89); +x_94 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__13; +x_95 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_95, 0, x_83); +lean_ctor_set(x_95, 1, x_94); +lean_ctor_set(x_95, 2, x_93); +x_96 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__26; +x_97 = lean_array_push(x_96, x_78); +x_98 = lean_array_push(x_97, x_95); +x_99 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__9; +x_100 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_100, 0, x_83); +lean_ctor_set(x_100, 1, x_99); +lean_ctor_set(x_100, 2, x_98); +x_101 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_101, 0, x_100); +lean_ctor_set(x_101, 1, x_68); +return x_101; } } else { -uint8_t x_101; +uint8_t x_102; +lean_dec(x_17); lean_dec(x_16); -lean_dec(x_15); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -x_101 = !lean_is_exclusive(x_23); -if (x_101 == 0) +x_102 = !lean_is_exclusive(x_24); +if (x_102 == 0) { -return x_23; +return x_24; } else { -lean_object* x_102; lean_object* x_103; lean_object* x_104; -x_102 = lean_ctor_get(x_23, 0); -x_103 = lean_ctor_get(x_23, 1); +lean_object* x_103; lean_object* x_104; lean_object* x_105; +x_103 = lean_ctor_get(x_24, 0); +x_104 = lean_ctor_get(x_24, 1); +lean_inc(x_104); lean_inc(x_103); -lean_inc(x_102); -lean_dec(x_23); -x_104 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_104, 0, x_102); -lean_ctor_set(x_104, 1, x_103); -return x_104; +lean_dec(x_24); +x_105 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_105, 0, x_103); +lean_ctor_set(x_105, 1, x_104); +return x_105; } } } @@ -2459,26 +2466,56 @@ static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2_ _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Format.nil", 10); +x_1 = lean_mk_string_from_bytes("term", 4); return x_1; } } static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__2() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__2; +x_3 = lean_alloc_ctor(1, 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_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Format.nil", 10); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__5() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__1; +x_1 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__4; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__3() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__1; +x_1 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__4; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__2; +x_3 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__5; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2486,7 +2523,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__4() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__7() { _start: { lean_object* x_1; @@ -2494,17 +2531,17 @@ x_1 = lean_mk_string_from_bytes("Format", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__5() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__4; +x_2 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__7; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__6() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__9() { _start: { lean_object* x_1; @@ -2512,17 +2549,17 @@ x_1 = lean_mk_string_from_bytes("nil", 3); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__7() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__5; -x_2 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__6; +x_1 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__8; +x_2 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__9; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__8() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__11() { _start: { lean_object* x_1; @@ -2530,61 +2567,61 @@ x_1 = lean_mk_string_from_bytes("Std", 3); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__9() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__8; +x_2 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__11; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__10() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__9; -x_2 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__4; +x_1 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__12; +x_2 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__7; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__11() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__10; -x_2 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__6; +x_1 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__13; +x_2 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__9; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__12() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__11; +x_2 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__14; 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_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__13() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__12; +x_2 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__15; x_3 = lean_alloc_ctor(1, 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_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__14() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__17() { _start: { lean_object* x_1; @@ -2592,11 +2629,11 @@ x_1 = lean_mk_string_from_bytes("'deriving Repr' failed, unexpected number of fi return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__15() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__18() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__14; +x_1 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__17; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } @@ -2625,15 +2662,15 @@ x_20 = lean_ctor_get(x_18, 0); lean_inc(x_20); lean_dec(x_18); x_21 = lean_environment_main_module(x_20); -x_22 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__7; -x_23 = l_Lean_addMacroScope(x_21, x_22, x_16); -x_24 = lean_box(0); -x_25 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__3; -x_26 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__13; +x_22 = lean_box(0); +x_23 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__10; +x_24 = l_Lean_addMacroScope(x_21, x_23, x_16); +x_25 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__6; +x_26 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__16; x_27 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_27, 0, x_14); lean_ctor_set(x_27, 1, x_25); -lean_ctor_set(x_27, 2, x_23); +lean_ctor_set(x_27, 2, x_24); lean_ctor_set(x_27, 3, x_26); x_28 = lean_array_get_size(x_4); x_29 = lean_array_get_size(x_1); @@ -2647,7 +2684,7 @@ if (x_31 == 0) lean_object* x_32; lean_object* x_33; uint8_t x_34; lean_dec(x_27); lean_dec(x_3); -x_32 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__15; +x_32 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__18; x_33 = l_Lean_throwError___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__1(x_32, x_6, x_7, x_8, x_9, x_10, x_11, x_19); lean_dec(x_11); lean_dec(x_10); @@ -2675,15 +2712,16 @@ return x_37; } else { -lean_object* x_38; lean_object* x_39; uint8_t x_40; lean_object* x_41; lean_object* x_42; -x_38 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__5; -x_39 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__10; -x_40 = 1; -x_41 = lean_box(0); -x_42 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__1(x_1, x_2, x_3, x_4, x_38, x_39, x_24, x_27, x_40, x_41, x_6, x_7, x_8, x_9, x_10, x_11, x_19); +lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; lean_object* x_42; lean_object* x_43; +x_38 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__3; +x_39 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__8; +x_40 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__13; +x_41 = 1; +x_42 = lean_box(0); +x_43 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__1(x_1, x_2, x_3, x_4, x_38, x_39, x_40, x_22, x_27, x_41, x_42, x_6, x_7, x_8, x_9, x_10, x_11, x_19); lean_dec(x_7); lean_dec(x_6); -return x_42; +return x_43; } } } @@ -2850,18 +2888,20 @@ lean_object* x_18 = _args[17]; lean_object* x_19 = _args[18]; lean_object* x_20 = _args[19]; lean_object* x_21 = _args[20]; +lean_object* x_22 = _args[21]; _start: { -lean_object* x_22; -x_22 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21); +lean_object* x_23; +x_23 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22); +lean_dec(x_17); lean_dec(x_16); -lean_dec(x_15); +lean_dec(x_14); lean_dec(x_13); -lean_dec(x_12); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -return x_22; +return x_23; } } LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__1___boxed(lean_object** _args) { @@ -2882,19 +2922,21 @@ lean_object* x_14 = _args[13]; lean_object* x_15 = _args[14]; lean_object* x_16 = _args[15]; lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; _start: { -uint8_t x_18; lean_object* x_19; -x_18 = lean_unbox(x_9); -lean_dec(x_9); -x_19 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_18, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +uint8_t x_19; lean_object* x_20; +x_19 = lean_unbox(x_10); +lean_dec(x_10); +x_20 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_19, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -return x_19; +return x_20; } } LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForStruct___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, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { @@ -3087,331 +3129,332 @@ x_1 = lean_mk_string_from_bytes("max_prec", 8); return x_1; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___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_object* x_7, 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, 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, lean_object* x_25) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___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_object* x_7, 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, 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, lean_object* x_25, lean_object* x_26) { _start: { -uint8_t x_26; -x_26 = lean_nat_dec_le(x_16, x_15); -if (x_26 == 0) +uint8_t x_27; +x_27 = lean_nat_dec_le(x_17, x_16); +if (x_27 == 0) { -lean_object* x_27; uint8_t x_28; -x_27 = lean_unsigned_to_nat(0u); -x_28 = lean_nat_dec_eq(x_14, x_27); -if (x_28 == 0) +lean_object* x_28; uint8_t x_29; +x_28 = lean_unsigned_to_nat(0u); +x_29 = lean_nat_dec_eq(x_15, x_28); +if (x_29 == 0) { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_37; -x_29 = lean_unsigned_to_nat(1u); -x_30 = lean_nat_sub(x_14, x_29); -lean_dec(x_14); -x_37 = !lean_is_exclusive(x_18); -if (x_37 == 0) +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_38; +x_30 = lean_unsigned_to_nat(1u); +x_31 = lean_nat_sub(x_15, x_30); +lean_dec(x_15); +x_38 = !lean_is_exclusive(x_19); +if (x_38 == 0) { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_38 = lean_ctor_get(x_18, 0); -x_39 = lean_ctor_get(x_18, 1); -x_40 = lean_nat_add(x_13, x_15); -x_41 = l_Lean_instInhabitedExpr; -x_42 = lean_array_get(x_41, x_3, x_40); -lean_dec(x_40); -x_43 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__2; -x_44 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_43, x_23, x_24, x_25); -x_45 = lean_ctor_get(x_44, 0); -lean_inc(x_45); -x_46 = lean_ctor_get(x_44, 1); +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_39 = lean_ctor_get(x_19, 0); +x_40 = lean_ctor_get(x_19, 1); +x_41 = lean_nat_add(x_14, x_16); +x_42 = l_Lean_instInhabitedExpr; +x_43 = lean_array_get(x_42, x_3, x_41); +lean_dec(x_41); +x_44 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__2; +x_45 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_44, x_24, x_25, x_26); +x_46 = lean_ctor_get(x_45, 0); lean_inc(x_46); -lean_dec(x_44); -x_47 = lean_mk_syntax_ident(x_45); +x_47 = lean_ctor_get(x_45, 1); lean_inc(x_47); -x_48 = lean_array_push(x_38, x_47); -lean_inc(x_42); -x_49 = l_Lean_Expr_fvarId_x21(x_42); -lean_inc(x_21); -x_50 = l_Lean_Meta_getLocalDecl(x_49, x_21, x_22, x_23, x_24, x_46); -if (lean_obj_tag(x_50) == 0) +lean_dec(x_45); +x_48 = lean_mk_syntax_ident(x_46); +lean_inc(x_48); +x_49 = lean_array_push(x_39, x_48); +lean_inc(x_43); +x_50 = l_Lean_Expr_fvarId_x21(x_43); +lean_inc(x_22); +x_51 = l_Lean_Meta_getLocalDecl(x_50, x_22, x_23, x_24, x_25, x_47); +if (lean_obj_tag(x_51) == 0) { -lean_object* x_51; lean_object* x_52; uint8_t x_53; uint8_t x_54; -x_51 = lean_ctor_get(x_50, 0); -lean_inc(x_51); -x_52 = lean_ctor_get(x_50, 1); +lean_object* x_52; lean_object* x_53; uint8_t x_54; uint8_t x_55; +x_52 = lean_ctor_get(x_51, 0); lean_inc(x_52); -lean_dec(x_50); -x_53 = l_Lean_LocalDecl_binderInfo(x_51); +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); lean_dec(x_51); -x_54 = l_Lean_BinderInfo_isExplicit(x_53); -if (x_54 == 0) +x_54 = l_Lean_LocalDecl_binderInfo(x_52); +lean_dec(x_52); +x_55 = l_Lean_BinderInfo_isExplicit(x_54); +if (x_55 == 0) { -lean_object* x_55; -lean_dec(x_47); -lean_dec(x_42); -lean_ctor_set(x_18, 0, x_48); -x_55 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_55, 0, x_18); -x_31 = x_55; -x_32 = x_52; -goto block_36; +lean_object* x_56; +lean_dec(x_48); +lean_dec(x_43); +lean_ctor_set(x_19, 0, x_49); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_19); +x_32 = x_56; +x_33 = x_53; +goto block_37; } else { -lean_object* x_56; +lean_object* x_57; +lean_inc(x_25); lean_inc(x_24); lean_inc(x_23); lean_inc(x_22); -lean_inc(x_21); -x_56 = lean_infer_type(x_42, x_21, x_22, x_23, x_24, x_52); -if (lean_obj_tag(x_56) == 0) +x_57 = lean_infer_type(x_43, x_22, x_23, x_24, x_25, x_53); +if (lean_obj_tag(x_57) == 0) { -lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; -x_57 = lean_ctor_get(x_56, 0); -lean_inc(x_57); -x_58 = lean_ctor_get(x_56, 1); +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; +x_58 = lean_ctor_get(x_57, 0); lean_inc(x_58); -lean_dec(x_56); -x_59 = lean_ctor_get(x_1, 0); -x_60 = lean_ctor_get(x_59, 0); -x_61 = l_Lean_Expr_isAppOf(x_57, x_60); +x_59 = lean_ctor_get(x_57, 1); +lean_inc(x_59); lean_dec(x_57); -if (x_61 == 0) +x_60 = lean_ctor_get(x_1, 0); +x_61 = lean_ctor_get(x_60, 0); +x_62 = l_Lean_Expr_isAppOf(x_58, x_61); +lean_dec(x_58); +if (x_62 == 0) { -lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -lean_inc(x_23); -x_62 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_23, x_24, x_58); -x_63 = lean_ctor_get(x_62, 0); -lean_inc(x_63); -x_64 = lean_ctor_get(x_62, 1); +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; +lean_inc(x_24); +x_63 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_24, x_25, x_59); +x_64 = lean_ctor_get(x_63, 0); lean_inc(x_64); -lean_dec(x_62); -x_65 = lean_ctor_get(x_23, 10); +x_65 = lean_ctor_get(x_63, 1); lean_inc(x_65); -x_66 = lean_st_ref_get(x_24, x_64); -x_67 = lean_ctor_get(x_66, 0); -lean_inc(x_67); -x_68 = lean_ctor_get(x_66, 1); +lean_dec(x_63); +x_66 = lean_ctor_get(x_24, 10); +lean_inc(x_66); +x_67 = lean_st_ref_get(x_25, x_65); +x_68 = lean_ctor_get(x_67, 0); lean_inc(x_68); -lean_dec(x_66); -x_69 = lean_ctor_get(x_67, 0); +x_69 = lean_ctor_get(x_67, 1); lean_inc(x_69); lean_dec(x_67); -x_70 = lean_environment_main_module(x_69); -x_71 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__3; -lean_inc(x_63); -x_72 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_72, 0, x_63); -lean_ctor_set(x_72, 1, x_71); -x_73 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__2; -x_74 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__3; -lean_inc(x_5); -x_75 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_75, 0, x_73); -lean_ctor_set(x_75, 1, x_5); -lean_ctor_set(x_75, 2, x_74); -x_76 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__4; -lean_inc(x_6); -x_77 = lean_name_mk_string(x_6, x_76); -lean_inc(x_65); +x_70 = lean_ctor_get(x_68, 0); lean_inc(x_70); -x_78 = l_Lean_addMacroScope(x_70, x_77, x_65); +lean_dec(x_68); +x_71 = lean_environment_main_module(x_70); +x_72 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__3; +lean_inc(x_64); +x_73 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_73, 0, x_64); +lean_ctor_set(x_73, 1, x_72); +x_74 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__2; +x_75 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__3; +lean_inc(x_6); +x_76 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_76, 0, x_74); +lean_ctor_set(x_76, 1, x_6); +lean_ctor_set(x_76, 2, x_75); +x_77 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__4; lean_inc(x_7); -x_79 = lean_name_mk_string(x_7, x_76); +x_78 = lean_name_mk_string(x_7, x_77); +lean_inc(x_66); +lean_inc(x_71); +x_79 = l_Lean_addMacroScope(x_71, x_78, x_66); lean_inc(x_8); -x_80 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_80, 0, x_79); -lean_ctor_set(x_80, 1, x_8); +x_80 = lean_name_mk_string(x_8, x_77); lean_inc(x_9); -x_81 = lean_alloc_ctor(1, 2, 0); +x_81 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_81, 0, x_80); lean_ctor_set(x_81, 1, x_9); -lean_inc(x_63); -x_82 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_82, 0, x_63); -lean_ctor_set(x_82, 1, x_75); -lean_ctor_set(x_82, 2, x_78); -lean_ctor_set(x_82, 3, x_81); -x_83 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__4; -x_84 = lean_array_push(x_83, x_39); -lean_inc(x_72); -x_85 = lean_array_push(x_84, x_72); -x_86 = lean_array_push(x_85, x_82); -x_87 = lean_box(2); -x_88 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__2; -x_89 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_89, 0, x_87); -lean_ctor_set(x_89, 1, x_88); -lean_ctor_set(x_89, 2, x_86); -x_90 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__3; -x_91 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__4; -lean_inc(x_5); -x_92 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_92, 0, x_90); -lean_ctor_set(x_92, 1, x_5); -lean_ctor_set(x_92, 2, x_91); -x_93 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__5; -x_94 = l_Lean_addMacroScope(x_70, x_93, x_65); -lean_inc(x_8); -x_95 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_95, 0, x_93); -lean_ctor_set(x_95, 1, x_8); +lean_inc(x_10); +x_82 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_82, 0, x_81); +lean_ctor_set(x_82, 1, x_10); +lean_inc(x_64); +x_83 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_83, 0, x_64); +lean_ctor_set(x_83, 1, x_76); +lean_ctor_set(x_83, 2, x_79); +lean_ctor_set(x_83, 3, x_82); +x_84 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__4; +x_85 = lean_array_push(x_84, x_40); +lean_inc(x_73); +x_86 = lean_array_push(x_85, x_73); +x_87 = lean_array_push(x_86, x_83); +x_88 = lean_box(2); +x_89 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__2; +x_90 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_90, 0, x_88); +lean_ctor_set(x_90, 1, x_89); +lean_ctor_set(x_90, 2, x_87); +x_91 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__3; +x_92 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__4; +lean_inc(x_6); +x_93 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_6); +lean_ctor_set(x_93, 2, x_92); +x_94 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__5; +x_95 = l_Lean_addMacroScope(x_71, x_94, x_66); lean_inc(x_9); -x_96 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_96, 0, x_95); +x_96 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_96, 0, x_94); lean_ctor_set(x_96, 1, x_9); -x_97 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_97, 0, x_63); -lean_ctor_set(x_97, 1, x_92); -lean_ctor_set(x_97, 2, x_94); -lean_ctor_set(x_97, 3, x_96); -lean_inc(x_11); -x_98 = lean_array_push(x_11, x_47); lean_inc(x_10); -x_99 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_99, 0, x_87); -lean_ctor_set(x_99, 1, x_10); -lean_ctor_set(x_99, 2, x_98); +x_97 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_97, 0, x_96); +lean_ctor_set(x_97, 1, x_10); +x_98 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_98, 0, x_64); +lean_ctor_set(x_98, 1, x_93); +lean_ctor_set(x_98, 2, x_95); +lean_ctor_set(x_98, 3, x_97); lean_inc(x_12); -x_100 = lean_array_push(x_12, x_97); -x_101 = lean_array_push(x_100, x_99); -lean_inc(x_4); -x_102 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_102, 0, x_87); -lean_ctor_set(x_102, 1, x_4); -lean_ctor_set(x_102, 2, x_101); -x_103 = lean_array_push(x_83, x_89); -x_104 = lean_array_push(x_103, x_72); -x_105 = lean_array_push(x_104, x_102); -x_106 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_106, 0, x_87); -lean_ctor_set(x_106, 1, x_88); -lean_ctor_set(x_106, 2, x_105); -lean_ctor_set(x_18, 1, x_106); -lean_ctor_set(x_18, 0, x_48); -x_107 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_107, 0, x_18); -x_31 = x_107; -x_32 = x_68; -goto block_36; +x_99 = lean_array_push(x_12, x_48); +lean_inc(x_11); +x_100 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_100, 0, x_88); +lean_ctor_set(x_100, 1, x_11); +lean_ctor_set(x_100, 2, x_99); +lean_inc(x_13); +x_101 = lean_array_push(x_13, x_98); +x_102 = lean_array_push(x_101, x_100); +lean_inc(x_5); +x_103 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_103, 0, x_88); +lean_ctor_set(x_103, 1, x_5); +lean_ctor_set(x_103, 2, x_102); +x_104 = lean_array_push(x_84, x_90); +x_105 = lean_array_push(x_104, x_73); +x_106 = lean_array_push(x_105, x_103); +x_107 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_107, 0, x_88); +lean_ctor_set(x_107, 1, x_89); +lean_ctor_set(x_107, 2, x_106); +lean_ctor_set(x_19, 1, x_107); +lean_ctor_set(x_19, 0, x_49); +x_108 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_108, 0, x_19); +x_32 = x_108; +x_33 = x_69; +goto block_37; } else { -lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; -lean_inc(x_23); -x_108 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_23, x_24, x_58); -x_109 = lean_ctor_get(x_108, 0); -lean_inc(x_109); -x_110 = lean_ctor_get(x_108, 1); +lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; +lean_inc(x_24); +x_109 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_24, x_25, x_59); +x_110 = lean_ctor_get(x_109, 0); lean_inc(x_110); -lean_dec(x_108); -x_111 = lean_ctor_get(x_23, 10); +x_111 = lean_ctor_get(x_109, 1); lean_inc(x_111); -x_112 = lean_st_ref_get(x_24, x_110); -x_113 = lean_ctor_get(x_112, 0); -lean_inc(x_113); -x_114 = lean_ctor_get(x_112, 1); +lean_dec(x_109); +x_112 = lean_ctor_get(x_24, 10); +lean_inc(x_112); +x_113 = lean_st_ref_get(x_25, x_111); +x_114 = lean_ctor_get(x_113, 0); lean_inc(x_114); -lean_dec(x_112); -x_115 = lean_ctor_get(x_113, 0); +x_115 = lean_ctor_get(x_113, 1); lean_inc(x_115); lean_dec(x_113); -x_116 = lean_environment_main_module(x_115); -x_117 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__3; -lean_inc(x_109); -x_118 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_118, 0, x_109); -lean_ctor_set(x_118, 1, x_117); -x_119 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__2; -x_120 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__3; -lean_inc(x_5); -x_121 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_121, 0, x_119); -lean_ctor_set(x_121, 1, x_5); -lean_ctor_set(x_121, 2, x_120); -x_122 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__4; +x_116 = lean_ctor_get(x_114, 0); +lean_inc(x_116); +lean_dec(x_114); +x_117 = lean_environment_main_module(x_116); +x_118 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__3; +lean_inc(x_110); +x_119 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_119, 0, x_110); +lean_ctor_set(x_119, 1, x_118); +x_120 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__2; +x_121 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__3; lean_inc(x_6); -x_123 = lean_name_mk_string(x_6, x_122); -x_124 = l_Lean_addMacroScope(x_116, x_123, x_111); +x_122 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_122, 0, x_120); +lean_ctor_set(x_122, 1, x_6); +lean_ctor_set(x_122, 2, x_121); +x_123 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__4; lean_inc(x_7); -x_125 = lean_name_mk_string(x_7, x_122); +x_124 = lean_name_mk_string(x_7, x_123); +x_125 = l_Lean_addMacroScope(x_117, x_124, x_112); lean_inc(x_8); -x_126 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_126, 0, x_125); -lean_ctor_set(x_126, 1, x_8); +x_126 = lean_name_mk_string(x_8, x_123); lean_inc(x_9); -x_127 = lean_alloc_ctor(1, 2, 0); +x_127 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_127, 0, x_126); lean_ctor_set(x_127, 1, x_9); -lean_inc(x_109); -x_128 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_128, 0, x_109); -lean_ctor_set(x_128, 1, x_121); -lean_ctor_set(x_128, 2, x_124); -lean_ctor_set(x_128, 3, x_127); -x_129 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__4; -x_130 = lean_array_push(x_129, x_39); -lean_inc(x_118); -x_131 = lean_array_push(x_130, x_118); -x_132 = lean_array_push(x_131, x_128); -x_133 = lean_box(2); -x_134 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__2; -x_135 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_135, 0, x_133); -lean_ctor_set(x_135, 1, x_134); -lean_ctor_set(x_135, 2, x_132); -lean_inc(x_2); -x_136 = lean_mk_syntax_ident(x_2); -x_137 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__8; -x_138 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_138, 0, x_109); -lean_ctor_set(x_138, 1, x_137); -lean_inc(x_11); -x_139 = lean_array_push(x_11, x_138); -x_140 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__7; -x_141 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_141, 0, x_133); -lean_ctor_set(x_141, 1, x_140); -lean_ctor_set(x_141, 2, x_139); -lean_inc(x_12); -x_142 = lean_array_push(x_12, x_47); -x_143 = lean_array_push(x_142, x_141); lean_inc(x_10); -x_144 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_144, 0, x_133); -lean_ctor_set(x_144, 1, x_10); -lean_ctor_set(x_144, 2, x_143); +x_128 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_128, 0, x_127); +lean_ctor_set(x_128, 1, x_10); +lean_inc(x_110); +x_129 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_129, 0, x_110); +lean_ctor_set(x_129, 1, x_122); +lean_ctor_set(x_129, 2, x_125); +lean_ctor_set(x_129, 3, x_128); +x_130 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__4; +x_131 = lean_array_push(x_130, x_40); +lean_inc(x_119); +x_132 = lean_array_push(x_131, x_119); +x_133 = lean_array_push(x_132, x_129); +x_134 = lean_box(2); +x_135 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__2; +x_136 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_136, 0, x_134); +lean_ctor_set(x_136, 1, x_135); +lean_ctor_set(x_136, 2, x_133); +lean_inc(x_2); +x_137 = lean_mk_syntax_ident(x_2); +x_138 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__8; +x_139 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_139, 0, x_110); +lean_ctor_set(x_139, 1, x_138); lean_inc(x_12); -x_145 = lean_array_push(x_12, x_136); -x_146 = lean_array_push(x_145, x_144); -lean_inc(x_4); -x_147 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_147, 0, x_133); -lean_ctor_set(x_147, 1, x_4); -lean_ctor_set(x_147, 2, x_146); -x_148 = lean_array_push(x_129, x_135); -x_149 = lean_array_push(x_148, x_118); -x_150 = lean_array_push(x_149, x_147); -x_151 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_151, 0, x_133); -lean_ctor_set(x_151, 1, x_134); -lean_ctor_set(x_151, 2, x_150); -lean_ctor_set(x_18, 1, x_151); -lean_ctor_set(x_18, 0, x_48); -x_152 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_152, 0, x_18); -x_31 = x_152; -x_32 = x_114; -goto block_36; +x_140 = lean_array_push(x_12, x_139); +x_141 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__7; +x_142 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_142, 0, x_134); +lean_ctor_set(x_142, 1, x_141); +lean_ctor_set(x_142, 2, x_140); +lean_inc(x_13); +x_143 = lean_array_push(x_13, x_48); +x_144 = lean_array_push(x_143, x_142); +lean_inc(x_11); +x_145 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_145, 0, x_134); +lean_ctor_set(x_145, 1, x_11); +lean_ctor_set(x_145, 2, x_144); +lean_inc(x_13); +x_146 = lean_array_push(x_13, x_137); +x_147 = lean_array_push(x_146, x_145); +lean_inc(x_5); +x_148 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_148, 0, x_134); +lean_ctor_set(x_148, 1, x_5); +lean_ctor_set(x_148, 2, x_147); +x_149 = lean_array_push(x_130, x_136); +x_150 = lean_array_push(x_149, x_119); +x_151 = lean_array_push(x_150, x_148); +x_152 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_152, 0, x_134); +lean_ctor_set(x_152, 1, x_135); +lean_ctor_set(x_152, 2, x_151); +lean_ctor_set(x_19, 1, x_152); +lean_ctor_set(x_19, 0, x_49); +x_153 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_153, 0, x_19); +x_32 = x_153; +x_33 = x_115; +goto block_37; } } else { -uint8_t x_153; +uint8_t x_154; +lean_dec(x_49); lean_dec(x_48); -lean_dec(x_47); -lean_free_object(x_18); -lean_dec(x_39); -lean_dec(x_30); +lean_free_object(x_19); +lean_dec(x_40); +lean_dec(x_31); +lean_dec(x_25); lean_dec(x_24); lean_dec(x_23); lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_15); +lean_dec(x_16); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -3420,43 +3463,43 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); lean_dec(x_2); -x_153 = !lean_is_exclusive(x_56); -if (x_153 == 0) +x_154 = !lean_is_exclusive(x_57); +if (x_154 == 0) { -return x_56; +return x_57; } else { -lean_object* x_154; lean_object* x_155; lean_object* x_156; -x_154 = lean_ctor_get(x_56, 0); -x_155 = lean_ctor_get(x_56, 1); +lean_object* x_155; lean_object* x_156; lean_object* x_157; +x_155 = lean_ctor_get(x_57, 0); +x_156 = lean_ctor_get(x_57, 1); +lean_inc(x_156); lean_inc(x_155); -lean_inc(x_154); -lean_dec(x_56); -x_156 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_156, 0, x_154); -lean_ctor_set(x_156, 1, x_155); -return x_156; +lean_dec(x_57); +x_157 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_157, 0, x_155); +lean_ctor_set(x_157, 1, x_156); +return x_157; } } } } else { -uint8_t x_157; +uint8_t x_158; +lean_dec(x_49); lean_dec(x_48); -lean_dec(x_47); -lean_dec(x_42); -lean_free_object(x_18); -lean_dec(x_39); -lean_dec(x_30); +lean_dec(x_43); +lean_free_object(x_19); +lean_dec(x_40); +lean_dec(x_31); +lean_dec(x_25); lean_dec(x_24); lean_dec(x_23); lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_15); +lean_dec(x_16); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -3465,342 +3508,342 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); lean_dec(x_2); -x_157 = !lean_is_exclusive(x_50); -if (x_157 == 0) +x_158 = !lean_is_exclusive(x_51); +if (x_158 == 0) { -return x_50; +return x_51; } else { -lean_object* x_158; lean_object* x_159; lean_object* x_160; -x_158 = lean_ctor_get(x_50, 0); -x_159 = lean_ctor_get(x_50, 1); +lean_object* x_159; lean_object* x_160; lean_object* x_161; +x_159 = lean_ctor_get(x_51, 0); +x_160 = lean_ctor_get(x_51, 1); +lean_inc(x_160); lean_inc(x_159); -lean_inc(x_158); -lean_dec(x_50); -x_160 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_160, 0, x_158); -lean_ctor_set(x_160, 1, x_159); -return x_160; +lean_dec(x_51); +x_161 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_161, 0, x_159); +lean_ctor_set(x_161, 1, x_160); +return x_161; } } } else { -lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; -x_161 = lean_ctor_get(x_18, 0); -x_162 = lean_ctor_get(x_18, 1); +lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; +x_162 = lean_ctor_get(x_19, 0); +x_163 = lean_ctor_get(x_19, 1); +lean_inc(x_163); lean_inc(x_162); -lean_inc(x_161); -lean_dec(x_18); -x_163 = lean_nat_add(x_13, x_15); -x_164 = l_Lean_instInhabitedExpr; -x_165 = lean_array_get(x_164, x_3, x_163); -lean_dec(x_163); -x_166 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__2; -x_167 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_166, x_23, x_24, x_25); -x_168 = lean_ctor_get(x_167, 0); -lean_inc(x_168); -x_169 = lean_ctor_get(x_167, 1); +lean_dec(x_19); +x_164 = lean_nat_add(x_14, x_16); +x_165 = l_Lean_instInhabitedExpr; +x_166 = lean_array_get(x_165, x_3, x_164); +lean_dec(x_164); +x_167 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__2; +x_168 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_167, x_24, x_25, x_26); +x_169 = lean_ctor_get(x_168, 0); lean_inc(x_169); -lean_dec(x_167); -x_170 = lean_mk_syntax_ident(x_168); +x_170 = lean_ctor_get(x_168, 1); lean_inc(x_170); -x_171 = lean_array_push(x_161, x_170); -lean_inc(x_165); -x_172 = l_Lean_Expr_fvarId_x21(x_165); -lean_inc(x_21); -x_173 = l_Lean_Meta_getLocalDecl(x_172, x_21, x_22, x_23, x_24, x_169); -if (lean_obj_tag(x_173) == 0) +lean_dec(x_168); +x_171 = lean_mk_syntax_ident(x_169); +lean_inc(x_171); +x_172 = lean_array_push(x_162, x_171); +lean_inc(x_166); +x_173 = l_Lean_Expr_fvarId_x21(x_166); +lean_inc(x_22); +x_174 = l_Lean_Meta_getLocalDecl(x_173, x_22, x_23, x_24, x_25, x_170); +if (lean_obj_tag(x_174) == 0) { -lean_object* x_174; lean_object* x_175; uint8_t x_176; uint8_t x_177; -x_174 = lean_ctor_get(x_173, 0); -lean_inc(x_174); -x_175 = lean_ctor_get(x_173, 1); +lean_object* x_175; lean_object* x_176; uint8_t x_177; uint8_t x_178; +x_175 = lean_ctor_get(x_174, 0); lean_inc(x_175); -lean_dec(x_173); -x_176 = l_Lean_LocalDecl_binderInfo(x_174); +x_176 = lean_ctor_get(x_174, 1); +lean_inc(x_176); lean_dec(x_174); -x_177 = l_Lean_BinderInfo_isExplicit(x_176); -if (x_177 == 0) +x_177 = l_Lean_LocalDecl_binderInfo(x_175); +lean_dec(x_175); +x_178 = l_Lean_BinderInfo_isExplicit(x_177); +if (x_178 == 0) { -lean_object* x_178; lean_object* x_179; -lean_dec(x_170); -lean_dec(x_165); -x_178 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_178, 0, x_171); -lean_ctor_set(x_178, 1, x_162); -x_179 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_179, 0, x_178); -x_31 = x_179; -x_32 = x_175; -goto block_36; +lean_object* x_179; lean_object* x_180; +lean_dec(x_171); +lean_dec(x_166); +x_179 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_179, 0, x_172); +lean_ctor_set(x_179, 1, x_163); +x_180 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_180, 0, x_179); +x_32 = x_180; +x_33 = x_176; +goto block_37; } else { -lean_object* x_180; +lean_object* x_181; +lean_inc(x_25); lean_inc(x_24); lean_inc(x_23); lean_inc(x_22); -lean_inc(x_21); -x_180 = lean_infer_type(x_165, x_21, x_22, x_23, x_24, x_175); -if (lean_obj_tag(x_180) == 0) +x_181 = lean_infer_type(x_166, x_22, x_23, x_24, x_25, x_176); +if (lean_obj_tag(x_181) == 0) { -lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; uint8_t x_185; -x_181 = lean_ctor_get(x_180, 0); -lean_inc(x_181); -x_182 = lean_ctor_get(x_180, 1); +lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; uint8_t x_186; +x_182 = lean_ctor_get(x_181, 0); lean_inc(x_182); -lean_dec(x_180); -x_183 = lean_ctor_get(x_1, 0); -x_184 = lean_ctor_get(x_183, 0); -x_185 = l_Lean_Expr_isAppOf(x_181, x_184); +x_183 = lean_ctor_get(x_181, 1); +lean_inc(x_183); lean_dec(x_181); -if (x_185 == 0) +x_184 = lean_ctor_get(x_1, 0); +x_185 = lean_ctor_get(x_184, 0); +x_186 = l_Lean_Expr_isAppOf(x_182, x_185); +lean_dec(x_182); +if (x_186 == 0) { -lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; -lean_inc(x_23); -x_186 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_23, x_24, x_182); -x_187 = lean_ctor_get(x_186, 0); -lean_inc(x_187); -x_188 = lean_ctor_get(x_186, 1); +lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; +lean_inc(x_24); +x_187 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_24, x_25, x_183); +x_188 = lean_ctor_get(x_187, 0); lean_inc(x_188); -lean_dec(x_186); -x_189 = lean_ctor_get(x_23, 10); +x_189 = lean_ctor_get(x_187, 1); lean_inc(x_189); -x_190 = lean_st_ref_get(x_24, x_188); -x_191 = lean_ctor_get(x_190, 0); -lean_inc(x_191); -x_192 = lean_ctor_get(x_190, 1); +lean_dec(x_187); +x_190 = lean_ctor_get(x_24, 10); +lean_inc(x_190); +x_191 = lean_st_ref_get(x_25, x_189); +x_192 = lean_ctor_get(x_191, 0); lean_inc(x_192); -lean_dec(x_190); -x_193 = lean_ctor_get(x_191, 0); +x_193 = lean_ctor_get(x_191, 1); lean_inc(x_193); lean_dec(x_191); -x_194 = lean_environment_main_module(x_193); -x_195 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__3; -lean_inc(x_187); -x_196 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_196, 0, x_187); -lean_ctor_set(x_196, 1, x_195); -x_197 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__2; -x_198 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__3; -lean_inc(x_5); -x_199 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_199, 0, x_197); -lean_ctor_set(x_199, 1, x_5); -lean_ctor_set(x_199, 2, x_198); -x_200 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__4; -lean_inc(x_6); -x_201 = lean_name_mk_string(x_6, x_200); -lean_inc(x_189); +x_194 = lean_ctor_get(x_192, 0); lean_inc(x_194); -x_202 = l_Lean_addMacroScope(x_194, x_201, x_189); +lean_dec(x_192); +x_195 = lean_environment_main_module(x_194); +x_196 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__3; +lean_inc(x_188); +x_197 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_197, 0, x_188); +lean_ctor_set(x_197, 1, x_196); +x_198 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__2; +x_199 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__3; +lean_inc(x_6); +x_200 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_200, 0, x_198); +lean_ctor_set(x_200, 1, x_6); +lean_ctor_set(x_200, 2, x_199); +x_201 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__4; lean_inc(x_7); -x_203 = lean_name_mk_string(x_7, x_200); +x_202 = lean_name_mk_string(x_7, x_201); +lean_inc(x_190); +lean_inc(x_195); +x_203 = l_Lean_addMacroScope(x_195, x_202, x_190); lean_inc(x_8); -x_204 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_204, 0, x_203); -lean_ctor_set(x_204, 1, x_8); +x_204 = lean_name_mk_string(x_8, x_201); lean_inc(x_9); -x_205 = lean_alloc_ctor(1, 2, 0); +x_205 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_205, 0, x_204); lean_ctor_set(x_205, 1, x_9); -lean_inc(x_187); -x_206 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_206, 0, x_187); -lean_ctor_set(x_206, 1, x_199); -lean_ctor_set(x_206, 2, x_202); -lean_ctor_set(x_206, 3, x_205); -x_207 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__4; -x_208 = lean_array_push(x_207, x_162); -lean_inc(x_196); -x_209 = lean_array_push(x_208, x_196); -x_210 = lean_array_push(x_209, x_206); -x_211 = lean_box(2); -x_212 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__2; -x_213 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_213, 0, x_211); -lean_ctor_set(x_213, 1, x_212); -lean_ctor_set(x_213, 2, x_210); -x_214 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__3; -x_215 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__4; -lean_inc(x_5); -x_216 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_216, 0, x_214); -lean_ctor_set(x_216, 1, x_5); -lean_ctor_set(x_216, 2, x_215); -x_217 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__5; -x_218 = l_Lean_addMacroScope(x_194, x_217, x_189); -lean_inc(x_8); -x_219 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_219, 0, x_217); -lean_ctor_set(x_219, 1, x_8); +lean_inc(x_10); +x_206 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_206, 0, x_205); +lean_ctor_set(x_206, 1, x_10); +lean_inc(x_188); +x_207 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_207, 0, x_188); +lean_ctor_set(x_207, 1, x_200); +lean_ctor_set(x_207, 2, x_203); +lean_ctor_set(x_207, 3, x_206); +x_208 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__4; +x_209 = lean_array_push(x_208, x_163); +lean_inc(x_197); +x_210 = lean_array_push(x_209, x_197); +x_211 = lean_array_push(x_210, x_207); +x_212 = lean_box(2); +x_213 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__2; +x_214 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_214, 0, x_212); +lean_ctor_set(x_214, 1, x_213); +lean_ctor_set(x_214, 2, x_211); +x_215 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__3; +x_216 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__4; +lean_inc(x_6); +x_217 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_217, 0, x_215); +lean_ctor_set(x_217, 1, x_6); +lean_ctor_set(x_217, 2, x_216); +x_218 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__5; +x_219 = l_Lean_addMacroScope(x_195, x_218, x_190); lean_inc(x_9); -x_220 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_220, 0, x_219); +x_220 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_220, 0, x_218); lean_ctor_set(x_220, 1, x_9); -x_221 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_221, 0, x_187); -lean_ctor_set(x_221, 1, x_216); -lean_ctor_set(x_221, 2, x_218); -lean_ctor_set(x_221, 3, x_220); -lean_inc(x_11); -x_222 = lean_array_push(x_11, x_170); lean_inc(x_10); -x_223 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_223, 0, x_211); -lean_ctor_set(x_223, 1, x_10); -lean_ctor_set(x_223, 2, x_222); +x_221 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_221, 0, x_220); +lean_ctor_set(x_221, 1, x_10); +x_222 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_222, 0, x_188); +lean_ctor_set(x_222, 1, x_217); +lean_ctor_set(x_222, 2, x_219); +lean_ctor_set(x_222, 3, x_221); lean_inc(x_12); -x_224 = lean_array_push(x_12, x_221); -x_225 = lean_array_push(x_224, x_223); -lean_inc(x_4); -x_226 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_226, 0, x_211); -lean_ctor_set(x_226, 1, x_4); -lean_ctor_set(x_226, 2, x_225); -x_227 = lean_array_push(x_207, x_213); -x_228 = lean_array_push(x_227, x_196); -x_229 = lean_array_push(x_228, x_226); -x_230 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_230, 0, x_211); -lean_ctor_set(x_230, 1, x_212); -lean_ctor_set(x_230, 2, x_229); -x_231 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_231, 0, x_171); -lean_ctor_set(x_231, 1, x_230); -x_232 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_232, 0, x_231); -x_31 = x_232; -x_32 = x_192; -goto block_36; +x_223 = lean_array_push(x_12, x_171); +lean_inc(x_11); +x_224 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_224, 0, x_212); +lean_ctor_set(x_224, 1, x_11); +lean_ctor_set(x_224, 2, x_223); +lean_inc(x_13); +x_225 = lean_array_push(x_13, x_222); +x_226 = lean_array_push(x_225, x_224); +lean_inc(x_5); +x_227 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_227, 0, x_212); +lean_ctor_set(x_227, 1, x_5); +lean_ctor_set(x_227, 2, x_226); +x_228 = lean_array_push(x_208, x_214); +x_229 = lean_array_push(x_228, x_197); +x_230 = lean_array_push(x_229, x_227); +x_231 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_231, 0, x_212); +lean_ctor_set(x_231, 1, x_213); +lean_ctor_set(x_231, 2, x_230); +x_232 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_232, 0, x_172); +lean_ctor_set(x_232, 1, x_231); +x_233 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_233, 0, x_232); +x_32 = x_233; +x_33 = x_193; +goto block_37; } else { -lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; -lean_inc(x_23); -x_233 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_23, x_24, x_182); -x_234 = lean_ctor_get(x_233, 0); -lean_inc(x_234); -x_235 = lean_ctor_get(x_233, 1); +lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; +lean_inc(x_24); +x_234 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_24, x_25, x_183); +x_235 = lean_ctor_get(x_234, 0); lean_inc(x_235); -lean_dec(x_233); -x_236 = lean_ctor_get(x_23, 10); +x_236 = lean_ctor_get(x_234, 1); lean_inc(x_236); -x_237 = lean_st_ref_get(x_24, x_235); -x_238 = lean_ctor_get(x_237, 0); -lean_inc(x_238); -x_239 = lean_ctor_get(x_237, 1); +lean_dec(x_234); +x_237 = lean_ctor_get(x_24, 10); +lean_inc(x_237); +x_238 = lean_st_ref_get(x_25, x_236); +x_239 = lean_ctor_get(x_238, 0); lean_inc(x_239); -lean_dec(x_237); -x_240 = lean_ctor_get(x_238, 0); +x_240 = lean_ctor_get(x_238, 1); lean_inc(x_240); lean_dec(x_238); -x_241 = lean_environment_main_module(x_240); -x_242 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__3; -lean_inc(x_234); -x_243 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_243, 0, x_234); -lean_ctor_set(x_243, 1, x_242); -x_244 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__2; -x_245 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__3; -lean_inc(x_5); -x_246 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_246, 0, x_244); -lean_ctor_set(x_246, 1, x_5); -lean_ctor_set(x_246, 2, x_245); -x_247 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__4; +x_241 = lean_ctor_get(x_239, 0); +lean_inc(x_241); +lean_dec(x_239); +x_242 = lean_environment_main_module(x_241); +x_243 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__3; +lean_inc(x_235); +x_244 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_244, 0, x_235); +lean_ctor_set(x_244, 1, x_243); +x_245 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__2; +x_246 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__3; lean_inc(x_6); -x_248 = lean_name_mk_string(x_6, x_247); -x_249 = l_Lean_addMacroScope(x_241, x_248, x_236); +x_247 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_247, 0, x_245); +lean_ctor_set(x_247, 1, x_6); +lean_ctor_set(x_247, 2, x_246); +x_248 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__4; lean_inc(x_7); -x_250 = lean_name_mk_string(x_7, x_247); +x_249 = lean_name_mk_string(x_7, x_248); +x_250 = l_Lean_addMacroScope(x_242, x_249, x_237); lean_inc(x_8); -x_251 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_251, 0, x_250); -lean_ctor_set(x_251, 1, x_8); +x_251 = lean_name_mk_string(x_8, x_248); lean_inc(x_9); -x_252 = lean_alloc_ctor(1, 2, 0); +x_252 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_252, 0, x_251); lean_ctor_set(x_252, 1, x_9); -lean_inc(x_234); -x_253 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_253, 0, x_234); -lean_ctor_set(x_253, 1, x_246); -lean_ctor_set(x_253, 2, x_249); -lean_ctor_set(x_253, 3, x_252); -x_254 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__4; -x_255 = lean_array_push(x_254, x_162); -lean_inc(x_243); -x_256 = lean_array_push(x_255, x_243); -x_257 = lean_array_push(x_256, x_253); -x_258 = lean_box(2); -x_259 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__2; -x_260 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_260, 0, x_258); -lean_ctor_set(x_260, 1, x_259); -lean_ctor_set(x_260, 2, x_257); -lean_inc(x_2); -x_261 = lean_mk_syntax_ident(x_2); -x_262 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__8; -x_263 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_263, 0, x_234); -lean_ctor_set(x_263, 1, x_262); -lean_inc(x_11); -x_264 = lean_array_push(x_11, x_263); -x_265 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__7; -x_266 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_266, 0, x_258); -lean_ctor_set(x_266, 1, x_265); -lean_ctor_set(x_266, 2, x_264); -lean_inc(x_12); -x_267 = lean_array_push(x_12, x_170); -x_268 = lean_array_push(x_267, x_266); lean_inc(x_10); -x_269 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_269, 0, x_258); -lean_ctor_set(x_269, 1, x_10); -lean_ctor_set(x_269, 2, x_268); +x_253 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_253, 0, x_252); +lean_ctor_set(x_253, 1, x_10); +lean_inc(x_235); +x_254 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_254, 0, x_235); +lean_ctor_set(x_254, 1, x_247); +lean_ctor_set(x_254, 2, x_250); +lean_ctor_set(x_254, 3, x_253); +x_255 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__4; +x_256 = lean_array_push(x_255, x_163); +lean_inc(x_244); +x_257 = lean_array_push(x_256, x_244); +x_258 = lean_array_push(x_257, x_254); +x_259 = lean_box(2); +x_260 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__2; +x_261 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_261, 0, x_259); +lean_ctor_set(x_261, 1, x_260); +lean_ctor_set(x_261, 2, x_258); +lean_inc(x_2); +x_262 = lean_mk_syntax_ident(x_2); +x_263 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__8; +x_264 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_264, 0, x_235); +lean_ctor_set(x_264, 1, x_263); lean_inc(x_12); -x_270 = lean_array_push(x_12, x_261); -x_271 = lean_array_push(x_270, x_269); -lean_inc(x_4); -x_272 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_272, 0, x_258); -lean_ctor_set(x_272, 1, x_4); -lean_ctor_set(x_272, 2, x_271); -x_273 = lean_array_push(x_254, x_260); -x_274 = lean_array_push(x_273, x_243); -x_275 = lean_array_push(x_274, x_272); -x_276 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_276, 0, x_258); -lean_ctor_set(x_276, 1, x_259); -lean_ctor_set(x_276, 2, x_275); -x_277 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_277, 0, x_171); -lean_ctor_set(x_277, 1, x_276); -x_278 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_278, 0, x_277); -x_31 = x_278; -x_32 = x_239; -goto block_36; +x_265 = lean_array_push(x_12, x_264); +x_266 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__7; +x_267 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_267, 0, x_259); +lean_ctor_set(x_267, 1, x_266); +lean_ctor_set(x_267, 2, x_265); +lean_inc(x_13); +x_268 = lean_array_push(x_13, x_171); +x_269 = lean_array_push(x_268, x_267); +lean_inc(x_11); +x_270 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_270, 0, x_259); +lean_ctor_set(x_270, 1, x_11); +lean_ctor_set(x_270, 2, x_269); +lean_inc(x_13); +x_271 = lean_array_push(x_13, x_262); +x_272 = lean_array_push(x_271, x_270); +lean_inc(x_5); +x_273 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_273, 0, x_259); +lean_ctor_set(x_273, 1, x_5); +lean_ctor_set(x_273, 2, x_272); +x_274 = lean_array_push(x_255, x_261); +x_275 = lean_array_push(x_274, x_244); +x_276 = lean_array_push(x_275, x_273); +x_277 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_277, 0, x_259); +lean_ctor_set(x_277, 1, x_260); +lean_ctor_set(x_277, 2, x_276); +x_278 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_278, 0, x_172); +lean_ctor_set(x_278, 1, x_277); +x_279 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_279, 0, x_278); +x_32 = x_279; +x_33 = x_240; +goto block_37; } } else { -lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; +lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; +lean_dec(x_172); lean_dec(x_171); -lean_dec(x_170); -lean_dec(x_162); -lean_dec(x_30); +lean_dec(x_163); +lean_dec(x_31); +lean_dec(x_25); lean_dec(x_24); lean_dec(x_23); lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_15); +lean_dec(x_16); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -3809,44 +3852,44 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); lean_dec(x_2); -x_279 = lean_ctor_get(x_180, 0); -lean_inc(x_279); -x_280 = lean_ctor_get(x_180, 1); +x_280 = lean_ctor_get(x_181, 0); lean_inc(x_280); -if (lean_is_exclusive(x_180)) { - lean_ctor_release(x_180, 0); - lean_ctor_release(x_180, 1); - x_281 = x_180; +x_281 = lean_ctor_get(x_181, 1); +lean_inc(x_281); +if (lean_is_exclusive(x_181)) { + lean_ctor_release(x_181, 0); + lean_ctor_release(x_181, 1); + x_282 = x_181; } else { - lean_dec_ref(x_180); - x_281 = lean_box(0); + lean_dec_ref(x_181); + x_282 = lean_box(0); } -if (lean_is_scalar(x_281)) { - x_282 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_282)) { + x_283 = lean_alloc_ctor(1, 2, 0); } else { - x_282 = x_281; + x_283 = x_282; } -lean_ctor_set(x_282, 0, x_279); -lean_ctor_set(x_282, 1, x_280); -return x_282; +lean_ctor_set(x_283, 0, x_280); +lean_ctor_set(x_283, 1, x_281); +return x_283; } } } else { -lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; +lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; +lean_dec(x_172); lean_dec(x_171); -lean_dec(x_170); -lean_dec(x_165); -lean_dec(x_162); -lean_dec(x_30); +lean_dec(x_166); +lean_dec(x_163); +lean_dec(x_31); +lean_dec(x_25); lean_dec(x_24); lean_dec(x_23); lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_15); +lean_dec(x_16); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -3855,54 +3898,54 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); lean_dec(x_2); -x_283 = lean_ctor_get(x_173, 0); -lean_inc(x_283); -x_284 = lean_ctor_get(x_173, 1); +x_284 = lean_ctor_get(x_174, 0); lean_inc(x_284); -if (lean_is_exclusive(x_173)) { - lean_ctor_release(x_173, 0); - lean_ctor_release(x_173, 1); - x_285 = x_173; +x_285 = lean_ctor_get(x_174, 1); +lean_inc(x_285); +if (lean_is_exclusive(x_174)) { + lean_ctor_release(x_174, 0); + lean_ctor_release(x_174, 1); + x_286 = x_174; } else { - lean_dec_ref(x_173); - x_285 = lean_box(0); + lean_dec_ref(x_174); + x_286 = lean_box(0); } -if (lean_is_scalar(x_285)) { - x_286 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_286)) { + x_287 = lean_alloc_ctor(1, 2, 0); } else { - x_286 = x_285; + x_287 = x_286; } -lean_ctor_set(x_286, 0, x_283); -lean_ctor_set(x_286, 1, x_284); -return x_286; +lean_ctor_set(x_287, 0, x_284); +lean_ctor_set(x_287, 1, x_285); +return x_287; } } -block_36: +block_37: { -lean_object* x_33; lean_object* x_34; -x_33 = lean_ctor_get(x_31, 0); -lean_inc(x_33); -lean_dec(x_31); -x_34 = lean_nat_add(x_15, x_17); -lean_dec(x_15); -x_14 = x_30; -x_15 = x_34; -x_18 = x_33; -x_25 = x_32; +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_32, 0); +lean_inc(x_34); +lean_dec(x_32); +x_35 = lean_nat_add(x_16, x_18); +lean_dec(x_16); +x_15 = x_31; +x_16 = x_35; +x_19 = x_34; +x_26 = x_33; goto _start; } } else { -lean_object* x_287; +lean_object* x_288; +lean_dec(x_25); lean_dec(x_24); lean_dec(x_23); lean_dec(x_22); -lean_dec(x_21); +lean_dec(x_16); lean_dec(x_15); -lean_dec(x_14); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -3911,23 +3954,23 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); lean_dec(x_2); -x_287 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_287, 0, x_18); -lean_ctor_set(x_287, 1, x_25); -return x_287; +x_288 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_288, 0, x_19); +lean_ctor_set(x_288, 1, x_26); +return x_288; } } else { -lean_object* x_288; +lean_object* x_289; +lean_dec(x_25); lean_dec(x_24); lean_dec(x_23); lean_dec(x_22); -lean_dec(x_21); +lean_dec(x_16); lean_dec(x_15); -lean_dec(x_14); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -3936,12 +3979,11 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); lean_dec(x_2); -x_288 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_288, 0, x_18); -lean_ctor_set(x_288, 1, x_25); -return x_288; +x_289 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_289, 0, x_19); +lean_ctor_set(x_289, 1, x_26); +return x_289; } } } @@ -3988,7 +4030,7 @@ static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyF _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__5; +x_1 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__8; x_2 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -3998,7 +4040,7 @@ static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyF _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__10; +x_1 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__13; x_2 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -4224,7 +4266,7 @@ static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyF _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__5; +x_1 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__8; x_2 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__28; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -4234,7 +4276,7 @@ static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyF _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__10; +x_1 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__13; x_2 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__28; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -4307,7 +4349,7 @@ static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyF _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__5; +x_1 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__8; x_2 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__36; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -4317,7 +4359,7 @@ static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyF _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__10; +x_1 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__13; x_2 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__36; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -4461,73 +4503,87 @@ return x_2; static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__54() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(2); +x_2 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__13; +x_3 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__27; +x_4 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__55() { +_start: +{ lean_object* x_1; lean_object* x_2; x_1 = lean_unsigned_to_nat(4u); x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___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, 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) { +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___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, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -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; uint8_t 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_16 = lean_ctor_get(x_1, 2); -lean_inc(x_16); -x_17 = lean_unsigned_to_nat(0u); -x_18 = lean_unsigned_to_nat(1u); -lean_inc(x_13); -lean_inc(x_2); -lean_inc(x_16); -x_19 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__1(x_16, x_17, x_16, x_18, x_2, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -lean_dec(x_16); +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; 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_15 = lean_ctor_get(x_1, 2); +lean_inc(x_15); +x_16 = lean_unsigned_to_nat(0u); +x_17 = lean_unsigned_to_nat(1u); +x_18 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__27; +lean_inc(x_12); +lean_inc(x_15); +x_19 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__1(x_15, x_16, x_15, x_17, x_18, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_15); x_20 = lean_ctor_get(x_19, 0); lean_inc(x_20); x_21 = lean_ctor_get(x_19, 1); lean_inc(x_21); lean_dec(x_19); -x_22 = lean_ctor_get(x_3, 0); -lean_inc(x_22); -lean_dec(x_3); -x_23 = 1; -x_24 = l_Lean_Name_toString(x_22, x_23); -x_25 = lean_box(2); -x_26 = l_Lean_Syntax_mkStrLit(x_24, x_25); -lean_dec(x_24); -lean_inc(x_13); -x_27 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_13, x_14, x_21); -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); +x_22 = lean_box(0); +x_23 = lean_ctor_get(x_2, 0); +lean_inc(x_23); +lean_dec(x_2); +x_24 = 1; +x_25 = l_Lean_Name_toString(x_23, x_24); +x_26 = lean_box(2); +x_27 = l_Lean_Syntax_mkStrLit(x_25, x_26); +lean_dec(x_25); +lean_inc(x_12); +x_28 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_12, x_13, x_21); +x_29 = lean_ctor_get(x_28, 0); lean_inc(x_29); -lean_dec(x_27); -x_30 = lean_ctor_get(x_13, 10); +x_30 = lean_ctor_get(x_28, 1); lean_inc(x_30); -x_31 = lean_st_ref_get(x_14, x_29); -x_32 = lean_ctor_get(x_31, 0); -lean_inc(x_32); -x_33 = lean_ctor_get(x_31, 1); +lean_dec(x_28); +x_31 = lean_ctor_get(x_12, 10); +lean_inc(x_31); +x_32 = lean_st_ref_get(x_13, x_30); +x_33 = lean_ctor_get(x_32, 0); lean_inc(x_33); -lean_dec(x_31); -x_34 = lean_ctor_get(x_32, 0); +x_34 = lean_ctor_get(x_32, 1); lean_inc(x_34); lean_dec(x_32); -x_35 = lean_environment_main_module(x_34); -x_36 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__5; -lean_inc(x_30); -x_37 = l_Lean_addMacroScope(x_35, x_36, x_30); -x_38 = lean_box(0); +x_35 = lean_ctor_get(x_33, 0); +lean_inc(x_35); +lean_dec(x_33); +x_36 = lean_environment_main_module(x_35); +x_37 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__5; +lean_inc(x_31); +x_38 = l_Lean_addMacroScope(x_36, x_37, x_31); x_39 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__3; x_40 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__8; x_41 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_41, 0, x_28); +lean_ctor_set(x_41, 0, x_29); lean_ctor_set(x_41, 1, x_39); -lean_ctor_set(x_41, 2, x_37); +lean_ctor_set(x_41, 2, x_38); lean_ctor_set(x_41, 3, x_40); x_42 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__18; -x_43 = lean_array_push(x_42, x_26); +x_43 = lean_array_push(x_42, x_27); x_44 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__13; x_45 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_45, 0, x_25); +lean_ctor_set(x_45, 0, x_26); lean_ctor_set(x_45, 1, x_44); lean_ctor_set(x_45, 2, x_43); x_46 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__26; @@ -4535,822 +4591,801 @@ x_47 = lean_array_push(x_46, x_41); x_48 = lean_array_push(x_47, x_45); x_49 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__9; x_50 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_50, 0, x_25); +lean_ctor_set(x_50, 0, x_26); lean_ctor_set(x_50, 1, x_49); lean_ctor_set(x_50, 2, x_48); x_51 = lean_ctor_get(x_1, 1); lean_inc(x_51); -lean_inc(x_13); -lean_inc(x_2); +lean_inc(x_12); lean_inc(x_51); -x_52 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__1(x_51, x_17, x_51, x_18, x_2, x_9, x_10, x_11, x_12, x_13, x_14, x_33); +x_52 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__1(x_51, x_16, x_51, x_17, x_18, x_8, x_9, x_10, x_11, x_12, x_13, x_34); x_53 = lean_ctor_get(x_52, 0); lean_inc(x_53); x_54 = lean_ctor_get(x_52, 1); lean_inc(x_54); lean_dec(x_52); -x_55 = lean_ctor_get(x_4, 4); +x_55 = lean_ctor_get(x_3, 4); lean_inc(x_55); -lean_dec(x_4); +lean_dec(x_3); x_56 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_56, 0, x_53); lean_ctor_set(x_56, 1, x_50); -x_57 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__5; -x_58 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__10; -lean_inc(x_14); +x_57 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__3; +x_58 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__8; +x_59 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__13; lean_inc(x_13); +lean_inc(x_12); lean_inc(x_55); -x_59 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2(x_1, x_5, x_7, x_49, x_17, x_57, x_58, x_38, x_38, x_44, x_42, x_46, x_51, x_55, x_17, x_55, x_18, x_56, x_9, x_10, x_11, x_12, x_13, x_14, x_54); +x_60 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2(x_1, x_4, x_6, x_57, x_49, x_16, x_58, x_59, x_22, x_22, x_44, x_42, x_46, x_51, x_55, x_16, x_55, x_17, x_56, x_8, x_9, x_10, x_11, x_12, x_13, x_54); lean_dec(x_55); lean_dec(x_51); lean_dec(x_1); -if (lean_obj_tag(x_59) == 0) +if (lean_obj_tag(x_60) == 0) { -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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; uint8_t x_86; -x_60 = lean_ctor_get(x_59, 0); -lean_inc(x_60); -x_61 = lean_ctor_get(x_59, 1); +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; uint8_t x_87; +x_61 = lean_ctor_get(x_60, 0); lean_inc(x_61); -lean_dec(x_59); -x_62 = lean_ctor_get(x_60, 0); +x_62 = lean_ctor_get(x_60, 1); lean_inc(x_62); -x_63 = lean_ctor_get(x_60, 1); -lean_inc(x_63); lean_dec(x_60); -lean_inc(x_13); -x_64 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_13, x_14, x_61); -x_65 = lean_ctor_get(x_64, 0); -lean_inc(x_65); -x_66 = lean_ctor_get(x_64, 1); +x_63 = lean_ctor_get(x_61, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_61, 1); +lean_inc(x_64); +lean_dec(x_61); +lean_inc(x_12); +x_65 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_12, x_13, x_62); +x_66 = lean_ctor_get(x_65, 0); lean_inc(x_66); -lean_dec(x_64); -x_67 = lean_st_ref_get(x_14, x_66); -x_68 = lean_ctor_get(x_67, 1); -lean_inc(x_68); -lean_dec(x_67); -x_69 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__11; -x_70 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_70, 0, x_65); -lean_ctor_set(x_70, 1, x_69); -x_71 = lean_mk_syntax_ident(x_6); -x_72 = lean_array_push(x_46, x_70); -x_73 = lean_array_push(x_72, x_71); -x_74 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__10; -x_75 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_75, 0, x_25); -lean_ctor_set(x_75, 1, x_74); -lean_ctor_set(x_75, 2, x_73); -lean_inc(x_2); -x_76 = l_Array_append___rarg(x_2, x_62); -x_77 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_77, 0, x_25); -lean_ctor_set(x_77, 1, x_44); -lean_ctor_set(x_77, 2, x_76); -x_78 = lean_array_push(x_46, x_75); -x_79 = lean_array_push(x_78, x_77); -x_80 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_80, 0, x_25); -lean_ctor_set(x_80, 1, x_49); -lean_ctor_set(x_80, 2, x_79); -x_81 = lean_array_push(x_20, x_80); -x_82 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_13, x_14, x_68); -x_83 = lean_ctor_get(x_82, 0); -lean_inc(x_83); -x_84 = lean_ctor_get(x_82, 1); -lean_inc(x_84); -lean_dec(x_82); -x_85 = lean_st_ref_get(x_14, x_84); -lean_dec(x_14); -x_86 = !lean_is_exclusive(x_85); -if (x_86 == 0) -{ -lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; size_t x_93; size_t x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; -x_87 = lean_ctor_get(x_85, 0); -x_88 = lean_ctor_get(x_87, 0); -lean_inc(x_88); -lean_dec(x_87); -x_89 = lean_environment_main_module(x_88); -x_90 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__14; -lean_inc(x_83); -x_91 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_91, 0, x_83); -lean_ctor_set(x_91, 1, x_90); -x_92 = lean_array_get_size(x_81); -x_93 = lean_usize_of_nat(x_92); -lean_dec(x_92); -x_94 = 0; -x_95 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_93, x_94, x_81); -x_96 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__16; -x_97 = l_Lean_mkSepArray(x_95, x_96); -lean_dec(x_95); -lean_inc(x_2); -x_98 = l_Array_append___rarg(x_2, x_97); -x_99 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_99, 0, x_25); -lean_ctor_set(x_99, 1, x_44); -lean_ctor_set(x_99, 2, x_98); -x_100 = lean_array_push(x_42, x_99); -x_101 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_101, 0, x_25); -lean_ctor_set(x_101, 1, x_44); -lean_ctor_set(x_101, 2, x_100); -x_102 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__17; -lean_inc(x_83); -x_103 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_103, 0, x_83); -lean_ctor_set(x_103, 1, x_102); -x_104 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__22; -lean_inc(x_30); -lean_inc(x_89); -x_105 = l_Lean_addMacroScope(x_89, x_104, x_30); -x_106 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__20; -x_107 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__24; -lean_inc(x_83); -x_108 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_108, 0, x_83); -lean_ctor_set(x_108, 1, x_106); -lean_ctor_set(x_108, 2, x_105); -lean_ctor_set(x_108, 3, x_107); -x_109 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__11; -lean_inc(x_83); -x_110 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_110, 0, x_83); -lean_ctor_set(x_110, 1, x_109); -x_111 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__29; -lean_inc(x_30); -lean_inc(x_89); -x_112 = l_Lean_addMacroScope(x_89, x_111, x_30); -x_113 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__27; -x_114 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__32; -lean_inc(x_83); -x_115 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_115, 0, x_83); -lean_ctor_set(x_115, 1, x_113); -lean_ctor_set(x_115, 2, x_112); -lean_ctor_set(x_115, 3, x_114); -x_116 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__37; -lean_inc(x_30); -lean_inc(x_89); -x_117 = l_Lean_addMacroScope(x_89, x_116, x_30); -x_118 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__35; -x_119 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__40; -lean_inc(x_83); -x_120 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_120, 0, x_83); -lean_ctor_set(x_120, 1, x_118); -lean_ctor_set(x_120, 2, x_117); -lean_ctor_set(x_120, 3, x_119); -x_121 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__43; -lean_inc(x_83); -x_122 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_122, 0, x_83); -lean_ctor_set(x_122, 1, x_121); -x_123 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__17; -x_124 = l_Lean_addMacroScope(x_89, x_123, x_30); -x_125 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__16; -lean_inc(x_83); -x_126 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_126, 0, x_83); -lean_ctor_set(x_126, 1, x_125); -lean_ctor_set(x_126, 2, x_124); -lean_ctor_set(x_126, 3, x_38); -x_127 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__46; -lean_inc(x_83); -x_128 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_128, 0, x_83); -lean_ctor_set(x_128, 1, x_127); -x_129 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__8; -lean_inc(x_83); -x_130 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_130, 0, x_83); -lean_ctor_set(x_130, 1, x_129); -x_131 = lean_array_push(x_42, x_130); -x_132 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__7; -x_133 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_133, 0, x_25); -lean_ctor_set(x_133, 1, x_132); -lean_ctor_set(x_133, 2, x_131); -x_134 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__4; -lean_inc(x_126); -x_135 = lean_array_push(x_134, x_126); -x_136 = lean_array_push(x_135, x_128); -x_137 = lean_array_push(x_136, x_133); -x_138 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__45; -x_139 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_139, 0, x_25); -lean_ctor_set(x_139, 1, x_138); -lean_ctor_set(x_139, 2, x_137); -x_140 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__47; -lean_inc(x_83); -x_141 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_141, 0, x_83); -lean_ctor_set(x_141, 1, x_140); -x_142 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__50; -lean_inc(x_83); -x_143 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_143, 0, x_83); -lean_ctor_set(x_143, 1, x_142); -x_144 = lean_array_push(x_42, x_143); -x_145 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__49; -x_146 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_146, 0, x_25); -lean_ctor_set(x_146, 1, x_145); -lean_ctor_set(x_146, 2, x_144); -x_147 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__51; -lean_inc(x_83); -x_148 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_148, 0, x_83); -lean_ctor_set(x_148, 1, x_147); -x_149 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__52; -lean_inc(x_83); -x_150 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_150, 0, x_83); -lean_ctor_set(x_150, 1, x_149); -x_151 = lean_array_push(x_42, x_150); -x_152 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_152, 0, x_25); -lean_ctor_set(x_152, 1, x_145); -lean_ctor_set(x_152, 2, x_151); -x_153 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__53; -x_154 = lean_array_push(x_153, x_122); -x_155 = lean_array_push(x_154, x_139); -x_156 = lean_array_push(x_155, x_141); -x_157 = lean_array_push(x_156, x_146); -x_158 = lean_array_push(x_157, x_148); -x_159 = lean_array_push(x_158, x_152); -x_160 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__42; -x_161 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_161, 0, x_25); -lean_ctor_set(x_161, 1, x_160); -lean_ctor_set(x_161, 2, x_159); +x_67 = lean_ctor_get(x_65, 1); +lean_inc(x_67); +lean_dec(x_65); +x_68 = lean_st_ref_get(x_13, x_67); +x_69 = lean_ctor_get(x_68, 1); +lean_inc(x_69); +lean_dec(x_68); +x_70 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__11; +x_71 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_71, 0, x_66); +lean_ctor_set(x_71, 1, x_70); +x_72 = lean_mk_syntax_ident(x_5); +x_73 = lean_array_push(x_46, x_71); +x_74 = lean_array_push(x_73, x_72); +x_75 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__10; +x_76 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_76, 0, x_26); +lean_ctor_set(x_76, 1, x_75); +lean_ctor_set(x_76, 2, x_74); +x_77 = l_Array_append___rarg(x_18, x_63); +x_78 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_78, 0, x_26); +lean_ctor_set(x_78, 1, x_44); +lean_ctor_set(x_78, 2, x_77); +x_79 = lean_array_push(x_46, x_76); +x_80 = lean_array_push(x_79, x_78); +x_81 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_81, 0, x_26); +lean_ctor_set(x_81, 1, x_49); +lean_ctor_set(x_81, 2, x_80); +x_82 = lean_array_push(x_20, x_81); +x_83 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_12, x_13, x_69); +x_84 = lean_ctor_get(x_83, 0); +lean_inc(x_84); +x_85 = lean_ctor_get(x_83, 1); +lean_inc(x_85); +lean_dec(x_83); +x_86 = lean_st_ref_get(x_13, x_85); +lean_dec(x_13); +x_87 = !lean_is_exclusive(x_86); +if (x_87 == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; size_t x_94; size_t x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; +x_88 = lean_ctor_get(x_86, 0); +x_89 = lean_ctor_get(x_88, 0); +lean_inc(x_89); +lean_dec(x_88); +x_90 = lean_environment_main_module(x_89); +x_91 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__14; +lean_inc(x_84); +x_92 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_92, 0, x_84); +lean_ctor_set(x_92, 1, x_91); +x_93 = lean_array_get_size(x_82); +x_94 = lean_usize_of_nat(x_93); +lean_dec(x_93); +x_95 = 0; +x_96 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_94, x_95, x_82); +x_97 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__16; +x_98 = l_Lean_mkSepArray(x_96, x_97); +lean_dec(x_96); +x_99 = l_Array_append___rarg(x_18, x_98); +x_100 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_100, 0, x_26); +lean_ctor_set(x_100, 1, x_44); +lean_ctor_set(x_100, 2, x_99); +x_101 = lean_array_push(x_42, x_100); +x_102 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_102, 0, x_26); +lean_ctor_set(x_102, 1, x_44); +lean_ctor_set(x_102, 2, x_101); +x_103 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__17; +lean_inc(x_84); +x_104 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_104, 0, x_84); +lean_ctor_set(x_104, 1, x_103); +x_105 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__22; +lean_inc(x_31); +lean_inc(x_90); +x_106 = l_Lean_addMacroScope(x_90, x_105, x_31); +x_107 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__20; +x_108 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__24; +lean_inc(x_84); +x_109 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_109, 0, x_84); +lean_ctor_set(x_109, 1, x_107); +lean_ctor_set(x_109, 2, x_106); +lean_ctor_set(x_109, 3, x_108); +x_110 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__11; +lean_inc(x_84); +x_111 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_111, 0, x_84); +lean_ctor_set(x_111, 1, x_110); +x_112 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__29; +lean_inc(x_31); +lean_inc(x_90); +x_113 = l_Lean_addMacroScope(x_90, x_112, x_31); +x_114 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__27; +x_115 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__32; +lean_inc(x_84); +x_116 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_116, 0, x_84); +lean_ctor_set(x_116, 1, x_114); +lean_ctor_set(x_116, 2, x_113); +lean_ctor_set(x_116, 3, x_115); +x_117 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__37; +lean_inc(x_31); +lean_inc(x_90); +x_118 = l_Lean_addMacroScope(x_90, x_117, x_31); +x_119 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__35; +x_120 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__40; +lean_inc(x_84); +x_121 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_121, 0, x_84); +lean_ctor_set(x_121, 1, x_119); +lean_ctor_set(x_121, 2, x_118); +lean_ctor_set(x_121, 3, x_120); +x_122 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__43; +lean_inc(x_84); +x_123 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_123, 0, x_84); +lean_ctor_set(x_123, 1, x_122); +x_124 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__17; +x_125 = l_Lean_addMacroScope(x_90, x_124, x_31); +x_126 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__16; +lean_inc(x_84); +x_127 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_127, 0, x_84); +lean_ctor_set(x_127, 1, x_126); +lean_ctor_set(x_127, 2, x_125); +lean_ctor_set(x_127, 3, x_22); +x_128 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__46; +lean_inc(x_84); +x_129 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_129, 0, x_84); +lean_ctor_set(x_129, 1, x_128); +x_130 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__8; +lean_inc(x_84); +x_131 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_131, 0, x_84); +lean_ctor_set(x_131, 1, x_130); +x_132 = lean_array_push(x_42, x_131); +x_133 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__7; +x_134 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_134, 0, x_26); +lean_ctor_set(x_134, 1, x_133); +lean_ctor_set(x_134, 2, x_132); +x_135 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__4; +lean_inc(x_127); +x_136 = lean_array_push(x_135, x_127); +x_137 = lean_array_push(x_136, x_129); +x_138 = lean_array_push(x_137, x_134); +x_139 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__45; +x_140 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_140, 0, x_26); +lean_ctor_set(x_140, 1, x_139); +lean_ctor_set(x_140, 2, x_138); +x_141 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__47; +lean_inc(x_84); +x_142 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_142, 0, x_84); +lean_ctor_set(x_142, 1, x_141); +x_143 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__50; +lean_inc(x_84); +x_144 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_144, 0, x_84); +lean_ctor_set(x_144, 1, x_143); +x_145 = lean_array_push(x_42, x_144); +x_146 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__49; +x_147 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_147, 0, x_26); +lean_ctor_set(x_147, 1, x_146); +lean_ctor_set(x_147, 2, x_145); +x_148 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__51; +lean_inc(x_84); +x_149 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_149, 0, x_84); +lean_ctor_set(x_149, 1, x_148); +x_150 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__52; +lean_inc(x_84); +x_151 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_151, 0, x_84); +lean_ctor_set(x_151, 1, x_150); +x_152 = lean_array_push(x_42, x_151); +x_153 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_153, 0, x_26); +lean_ctor_set(x_153, 1, x_146); +lean_ctor_set(x_153, 2, x_152); +x_154 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__53; +x_155 = lean_array_push(x_154, x_123); +x_156 = lean_array_push(x_155, x_140); +x_157 = lean_array_push(x_156, x_142); +x_158 = lean_array_push(x_157, x_147); +x_159 = lean_array_push(x_158, x_149); +x_160 = lean_array_push(x_159, x_153); +x_161 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__42; x_162 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_162, 0, x_25); -lean_ctor_set(x_162, 1, x_44); -lean_ctor_set(x_162, 2, x_2); -x_163 = lean_array_push(x_46, x_161); -lean_inc(x_162); -x_164 = lean_array_push(x_163, x_162); -x_165 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_165, 0, x_25); -lean_ctor_set(x_165, 1, x_44); -lean_ctor_set(x_165, 2, x_164); -x_166 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__29; -x_167 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_167, 0, x_83); -lean_ctor_set(x_167, 1, x_166); -x_168 = lean_array_push(x_134, x_110); +lean_ctor_set(x_162, 0, x_26); +lean_ctor_set(x_162, 1, x_161); +lean_ctor_set(x_162, 2, x_160); +x_163 = lean_array_push(x_46, x_162); +x_164 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__54; +x_165 = lean_array_push(x_163, x_164); +x_166 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_166, 0, x_26); +lean_ctor_set(x_166, 1, x_44); +lean_ctor_set(x_166, 2, x_165); +x_167 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__29; +x_168 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_168, 0, x_84); +lean_ctor_set(x_168, 1, x_167); +x_169 = lean_array_push(x_135, x_111); +lean_inc(x_169); +x_170 = lean_array_push(x_169, x_166); lean_inc(x_168); -x_169 = lean_array_push(x_168, x_165); -lean_inc(x_167); -x_170 = lean_array_push(x_169, x_167); -x_171 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__14; -x_172 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_172, 0, x_25); -lean_ctor_set(x_172, 1, x_171); -lean_ctor_set(x_172, 2, x_170); -x_173 = lean_array_push(x_46, x_63); -lean_inc(x_162); -x_174 = lean_array_push(x_173, x_162); -x_175 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_175, 0, x_25); -lean_ctor_set(x_175, 1, x_44); -lean_ctor_set(x_175, 2, x_174); +x_171 = lean_array_push(x_170, x_168); +x_172 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__14; +x_173 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_173, 0, x_26); +lean_ctor_set(x_173, 1, x_172); +lean_ctor_set(x_173, 2, x_171); +x_174 = lean_array_push(x_46, x_64); +x_175 = lean_array_push(x_174, x_164); +x_176 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_176, 0, x_26); +lean_ctor_set(x_176, 1, x_44); +lean_ctor_set(x_176, 2, x_175); +lean_inc(x_169); +x_177 = lean_array_push(x_169, x_176); lean_inc(x_168); -x_176 = lean_array_push(x_168, x_175); -lean_inc(x_167); -x_177 = lean_array_push(x_176, x_167); -x_178 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_178, 0, x_25); -lean_ctor_set(x_178, 1, x_171); -lean_ctor_set(x_178, 2, x_177); -x_179 = lean_array_push(x_46, x_172); -x_180 = lean_array_push(x_179, x_178); -x_181 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_181, 0, x_25); -lean_ctor_set(x_181, 1, x_44); -lean_ctor_set(x_181, 2, x_180); -x_182 = lean_array_push(x_46, x_120); -x_183 = lean_array_push(x_182, x_181); -x_184 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_184, 0, x_25); -lean_ctor_set(x_184, 1, x_49); -lean_ctor_set(x_184, 2, x_183); -x_185 = lean_array_push(x_46, x_184); -lean_inc(x_162); -x_186 = lean_array_push(x_185, x_162); -x_187 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_187, 0, x_25); -lean_ctor_set(x_187, 1, x_44); -lean_ctor_set(x_187, 2, x_186); +x_178 = lean_array_push(x_177, x_168); +x_179 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_179, 0, x_26); +lean_ctor_set(x_179, 1, x_172); +lean_ctor_set(x_179, 2, x_178); +x_180 = lean_array_push(x_46, x_173); +x_181 = lean_array_push(x_180, x_179); +x_182 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_182, 0, x_26); +lean_ctor_set(x_182, 1, x_44); +lean_ctor_set(x_182, 2, x_181); +x_183 = lean_array_push(x_46, x_121); +x_184 = lean_array_push(x_183, x_182); +x_185 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_185, 0, x_26); +lean_ctor_set(x_185, 1, x_49); +lean_ctor_set(x_185, 2, x_184); +x_186 = lean_array_push(x_46, x_185); +x_187 = lean_array_push(x_186, x_164); +x_188 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_188, 0, x_26); +lean_ctor_set(x_188, 1, x_44); +lean_ctor_set(x_188, 2, x_187); +lean_inc(x_169); +x_189 = lean_array_push(x_169, x_188); lean_inc(x_168); -x_188 = lean_array_push(x_168, x_187); -lean_inc(x_167); -x_189 = lean_array_push(x_188, x_167); -x_190 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_190, 0, x_25); -lean_ctor_set(x_190, 1, x_171); -lean_ctor_set(x_190, 2, x_189); -x_191 = lean_array_push(x_42, x_190); -x_192 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_192, 0, x_25); -lean_ctor_set(x_192, 1, x_44); -lean_ctor_set(x_192, 2, x_191); -x_193 = lean_array_push(x_46, x_115); -x_194 = lean_array_push(x_193, x_192); -x_195 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_195, 0, x_25); -lean_ctor_set(x_195, 1, x_49); -lean_ctor_set(x_195, 2, x_194); -x_196 = lean_array_push(x_46, x_195); -x_197 = lean_array_push(x_196, x_162); -x_198 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_198, 0, x_25); -lean_ctor_set(x_198, 1, x_44); -lean_ctor_set(x_198, 2, x_197); -x_199 = lean_array_push(x_168, x_198); -x_200 = lean_array_push(x_199, x_167); -x_201 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_201, 0, x_25); -lean_ctor_set(x_201, 1, x_171); -lean_ctor_set(x_201, 2, x_200); -x_202 = lean_array_push(x_46, x_201); -x_203 = lean_array_push(x_202, x_126); -x_204 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_204, 0, x_25); -lean_ctor_set(x_204, 1, x_44); -lean_ctor_set(x_204, 2, x_203); -x_205 = lean_array_push(x_46, x_108); -x_206 = lean_array_push(x_205, x_204); -x_207 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_207, 0, x_25); -lean_ctor_set(x_207, 1, x_49); -lean_ctor_set(x_207, 2, x_206); -x_208 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__54; -x_209 = lean_array_push(x_208, x_91); -x_210 = lean_array_push(x_209, x_101); -x_211 = lean_array_push(x_210, x_103); -x_212 = lean_array_push(x_211, x_207); -x_213 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__13; -x_214 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_214, 0, x_25); -lean_ctor_set(x_214, 1, x_213); -lean_ctor_set(x_214, 2, x_212); -lean_ctor_set(x_85, 0, x_214); -return x_85; +x_190 = lean_array_push(x_189, x_168); +x_191 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_191, 0, x_26); +lean_ctor_set(x_191, 1, x_172); +lean_ctor_set(x_191, 2, x_190); +x_192 = lean_array_push(x_42, x_191); +x_193 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_193, 0, x_26); +lean_ctor_set(x_193, 1, x_44); +lean_ctor_set(x_193, 2, x_192); +x_194 = lean_array_push(x_46, x_116); +x_195 = lean_array_push(x_194, x_193); +x_196 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_196, 0, x_26); +lean_ctor_set(x_196, 1, x_49); +lean_ctor_set(x_196, 2, x_195); +x_197 = lean_array_push(x_46, x_196); +x_198 = lean_array_push(x_197, x_164); +x_199 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_199, 0, x_26); +lean_ctor_set(x_199, 1, x_44); +lean_ctor_set(x_199, 2, x_198); +x_200 = lean_array_push(x_169, x_199); +x_201 = lean_array_push(x_200, x_168); +x_202 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_202, 0, x_26); +lean_ctor_set(x_202, 1, x_172); +lean_ctor_set(x_202, 2, x_201); +x_203 = lean_array_push(x_46, x_202); +x_204 = lean_array_push(x_203, x_127); +x_205 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_205, 0, x_26); +lean_ctor_set(x_205, 1, x_44); +lean_ctor_set(x_205, 2, x_204); +x_206 = lean_array_push(x_46, x_109); +x_207 = lean_array_push(x_206, x_205); +x_208 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_208, 0, x_26); +lean_ctor_set(x_208, 1, x_49); +lean_ctor_set(x_208, 2, x_207); +x_209 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__55; +x_210 = lean_array_push(x_209, x_92); +x_211 = lean_array_push(x_210, x_102); +x_212 = lean_array_push(x_211, x_104); +x_213 = lean_array_push(x_212, x_208); +x_214 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__13; +x_215 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_215, 0, x_26); +lean_ctor_set(x_215, 1, x_214); +lean_ctor_set(x_215, 2, x_213); +lean_ctor_set(x_86, 0, x_215); +return x_86; } else { -lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; size_t x_222; size_t x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; -x_215 = lean_ctor_get(x_85, 0); -x_216 = lean_ctor_get(x_85, 1); -lean_inc(x_216); -lean_inc(x_215); -lean_dec(x_85); -x_217 = lean_ctor_get(x_215, 0); +lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; size_t x_223; size_t x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; +x_216 = lean_ctor_get(x_86, 0); +x_217 = lean_ctor_get(x_86, 1); lean_inc(x_217); -lean_dec(x_215); -x_218 = lean_environment_main_module(x_217); -x_219 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__14; -lean_inc(x_83); -x_220 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_220, 0, x_83); -lean_ctor_set(x_220, 1, x_219); -x_221 = lean_array_get_size(x_81); -x_222 = lean_usize_of_nat(x_221); -lean_dec(x_221); -x_223 = 0; -x_224 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_222, x_223, x_81); -x_225 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__16; -x_226 = l_Lean_mkSepArray(x_224, x_225); -lean_dec(x_224); -lean_inc(x_2); -x_227 = l_Array_append___rarg(x_2, x_226); -x_228 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_228, 0, x_25); -lean_ctor_set(x_228, 1, x_44); -lean_ctor_set(x_228, 2, x_227); -x_229 = lean_array_push(x_42, x_228); -x_230 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_230, 0, x_25); -lean_ctor_set(x_230, 1, x_44); -lean_ctor_set(x_230, 2, x_229); -x_231 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__17; -lean_inc(x_83); -x_232 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_232, 0, x_83); -lean_ctor_set(x_232, 1, x_231); -x_233 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__22; -lean_inc(x_30); -lean_inc(x_218); -x_234 = l_Lean_addMacroScope(x_218, x_233, x_30); -x_235 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__20; -x_236 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__24; -lean_inc(x_83); -x_237 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_237, 0, x_83); -lean_ctor_set(x_237, 1, x_235); -lean_ctor_set(x_237, 2, x_234); -lean_ctor_set(x_237, 3, x_236); -x_238 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__11; -lean_inc(x_83); -x_239 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_239, 0, x_83); -lean_ctor_set(x_239, 1, x_238); -x_240 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__29; -lean_inc(x_30); -lean_inc(x_218); -x_241 = l_Lean_addMacroScope(x_218, x_240, x_30); -x_242 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__27; -x_243 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__32; -lean_inc(x_83); -x_244 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_244, 0, x_83); -lean_ctor_set(x_244, 1, x_242); -lean_ctor_set(x_244, 2, x_241); -lean_ctor_set(x_244, 3, x_243); -x_245 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__37; -lean_inc(x_30); +lean_inc(x_216); +lean_dec(x_86); +x_218 = lean_ctor_get(x_216, 0); lean_inc(x_218); -x_246 = l_Lean_addMacroScope(x_218, x_245, x_30); -x_247 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__35; -x_248 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__40; -lean_inc(x_83); -x_249 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_249, 0, x_83); -lean_ctor_set(x_249, 1, x_247); -lean_ctor_set(x_249, 2, x_246); -lean_ctor_set(x_249, 3, x_248); -x_250 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__43; -lean_inc(x_83); -x_251 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_251, 0, x_83); -lean_ctor_set(x_251, 1, x_250); -x_252 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__17; -x_253 = l_Lean_addMacroScope(x_218, x_252, x_30); -x_254 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__16; -lean_inc(x_83); -x_255 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_255, 0, x_83); -lean_ctor_set(x_255, 1, x_254); -lean_ctor_set(x_255, 2, x_253); -lean_ctor_set(x_255, 3, x_38); -x_256 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__46; -lean_inc(x_83); -x_257 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_257, 0, x_83); -lean_ctor_set(x_257, 1, x_256); -x_258 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__8; -lean_inc(x_83); -x_259 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_259, 0, x_83); -lean_ctor_set(x_259, 1, x_258); -x_260 = lean_array_push(x_42, x_259); -x_261 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__7; -x_262 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_262, 0, x_25); -lean_ctor_set(x_262, 1, x_261); -lean_ctor_set(x_262, 2, x_260); -x_263 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__4; -lean_inc(x_255); -x_264 = lean_array_push(x_263, x_255); -x_265 = lean_array_push(x_264, x_257); -x_266 = lean_array_push(x_265, x_262); -x_267 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__45; -x_268 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_268, 0, x_25); -lean_ctor_set(x_268, 1, x_267); -lean_ctor_set(x_268, 2, x_266); -x_269 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__47; -lean_inc(x_83); -x_270 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_270, 0, x_83); -lean_ctor_set(x_270, 1, x_269); -x_271 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__50; -lean_inc(x_83); -x_272 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_272, 0, x_83); -lean_ctor_set(x_272, 1, x_271); -x_273 = lean_array_push(x_42, x_272); -x_274 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__49; -x_275 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_275, 0, x_25); -lean_ctor_set(x_275, 1, x_274); -lean_ctor_set(x_275, 2, x_273); -x_276 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__51; -lean_inc(x_83); -x_277 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_277, 0, x_83); -lean_ctor_set(x_277, 1, x_276); -x_278 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__52; -lean_inc(x_83); -x_279 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_279, 0, x_83); -lean_ctor_set(x_279, 1, x_278); -x_280 = lean_array_push(x_42, x_279); -x_281 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_281, 0, x_25); -lean_ctor_set(x_281, 1, x_274); -lean_ctor_set(x_281, 2, x_280); -x_282 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__53; -x_283 = lean_array_push(x_282, x_251); -x_284 = lean_array_push(x_283, x_268); -x_285 = lean_array_push(x_284, x_270); -x_286 = lean_array_push(x_285, x_275); -x_287 = lean_array_push(x_286, x_277); -x_288 = lean_array_push(x_287, x_281); -x_289 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__42; -x_290 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_290, 0, x_25); -lean_ctor_set(x_290, 1, x_289); -lean_ctor_set(x_290, 2, x_288); +lean_dec(x_216); +x_219 = lean_environment_main_module(x_218); +x_220 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__14; +lean_inc(x_84); +x_221 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_221, 0, x_84); +lean_ctor_set(x_221, 1, x_220); +x_222 = lean_array_get_size(x_82); +x_223 = lean_usize_of_nat(x_222); +lean_dec(x_222); +x_224 = 0; +x_225 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_223, x_224, x_82); +x_226 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__16; +x_227 = l_Lean_mkSepArray(x_225, x_226); +lean_dec(x_225); +x_228 = l_Array_append___rarg(x_18, x_227); +x_229 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_229, 0, x_26); +lean_ctor_set(x_229, 1, x_44); +lean_ctor_set(x_229, 2, x_228); +x_230 = lean_array_push(x_42, x_229); +x_231 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_231, 0, x_26); +lean_ctor_set(x_231, 1, x_44); +lean_ctor_set(x_231, 2, x_230); +x_232 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__17; +lean_inc(x_84); +x_233 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_233, 0, x_84); +lean_ctor_set(x_233, 1, x_232); +x_234 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__22; +lean_inc(x_31); +lean_inc(x_219); +x_235 = l_Lean_addMacroScope(x_219, x_234, x_31); +x_236 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__20; +x_237 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__24; +lean_inc(x_84); +x_238 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_238, 0, x_84); +lean_ctor_set(x_238, 1, x_236); +lean_ctor_set(x_238, 2, x_235); +lean_ctor_set(x_238, 3, x_237); +x_239 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__11; +lean_inc(x_84); +x_240 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_240, 0, x_84); +lean_ctor_set(x_240, 1, x_239); +x_241 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__29; +lean_inc(x_31); +lean_inc(x_219); +x_242 = l_Lean_addMacroScope(x_219, x_241, x_31); +x_243 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__27; +x_244 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__32; +lean_inc(x_84); +x_245 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_245, 0, x_84); +lean_ctor_set(x_245, 1, x_243); +lean_ctor_set(x_245, 2, x_242); +lean_ctor_set(x_245, 3, x_244); +x_246 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__37; +lean_inc(x_31); +lean_inc(x_219); +x_247 = l_Lean_addMacroScope(x_219, x_246, x_31); +x_248 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__35; +x_249 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__40; +lean_inc(x_84); +x_250 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_250, 0, x_84); +lean_ctor_set(x_250, 1, x_248); +lean_ctor_set(x_250, 2, x_247); +lean_ctor_set(x_250, 3, x_249); +x_251 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__43; +lean_inc(x_84); +x_252 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_252, 0, x_84); +lean_ctor_set(x_252, 1, x_251); +x_253 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__17; +x_254 = l_Lean_addMacroScope(x_219, x_253, x_31); +x_255 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__16; +lean_inc(x_84); +x_256 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_256, 0, x_84); +lean_ctor_set(x_256, 1, x_255); +lean_ctor_set(x_256, 2, x_254); +lean_ctor_set(x_256, 3, x_22); +x_257 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__46; +lean_inc(x_84); +x_258 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_258, 0, x_84); +lean_ctor_set(x_258, 1, x_257); +x_259 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__8; +lean_inc(x_84); +x_260 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_260, 0, x_84); +lean_ctor_set(x_260, 1, x_259); +x_261 = lean_array_push(x_42, x_260); +x_262 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__7; +x_263 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_263, 0, x_26); +lean_ctor_set(x_263, 1, x_262); +lean_ctor_set(x_263, 2, x_261); +x_264 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__4; +lean_inc(x_256); +x_265 = lean_array_push(x_264, x_256); +x_266 = lean_array_push(x_265, x_258); +x_267 = lean_array_push(x_266, x_263); +x_268 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__45; +x_269 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_269, 0, x_26); +lean_ctor_set(x_269, 1, x_268); +lean_ctor_set(x_269, 2, x_267); +x_270 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__47; +lean_inc(x_84); +x_271 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_271, 0, x_84); +lean_ctor_set(x_271, 1, x_270); +x_272 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__50; +lean_inc(x_84); +x_273 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_273, 0, x_84); +lean_ctor_set(x_273, 1, x_272); +x_274 = lean_array_push(x_42, x_273); +x_275 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__49; +x_276 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_276, 0, x_26); +lean_ctor_set(x_276, 1, x_275); +lean_ctor_set(x_276, 2, x_274); +x_277 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__51; +lean_inc(x_84); +x_278 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_278, 0, x_84); +lean_ctor_set(x_278, 1, x_277); +x_279 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__52; +lean_inc(x_84); +x_280 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_280, 0, x_84); +lean_ctor_set(x_280, 1, x_279); +x_281 = lean_array_push(x_42, x_280); +x_282 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_282, 0, x_26); +lean_ctor_set(x_282, 1, x_275); +lean_ctor_set(x_282, 2, x_281); +x_283 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__53; +x_284 = lean_array_push(x_283, x_252); +x_285 = lean_array_push(x_284, x_269); +x_286 = lean_array_push(x_285, x_271); +x_287 = lean_array_push(x_286, x_276); +x_288 = lean_array_push(x_287, x_278); +x_289 = lean_array_push(x_288, x_282); +x_290 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__42; x_291 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_291, 0, x_25); -lean_ctor_set(x_291, 1, x_44); -lean_ctor_set(x_291, 2, x_2); -x_292 = lean_array_push(x_46, x_290); -lean_inc(x_291); -x_293 = lean_array_push(x_292, x_291); -x_294 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_294, 0, x_25); -lean_ctor_set(x_294, 1, x_44); -lean_ctor_set(x_294, 2, x_293); -x_295 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__29; -x_296 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_296, 0, x_83); -lean_ctor_set(x_296, 1, x_295); -x_297 = lean_array_push(x_263, x_239); +lean_ctor_set(x_291, 0, x_26); +lean_ctor_set(x_291, 1, x_290); +lean_ctor_set(x_291, 2, x_289); +x_292 = lean_array_push(x_46, x_291); +x_293 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__54; +x_294 = lean_array_push(x_292, x_293); +x_295 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_295, 0, x_26); +lean_ctor_set(x_295, 1, x_44); +lean_ctor_set(x_295, 2, x_294); +x_296 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__29; +x_297 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_297, 0, x_84); +lean_ctor_set(x_297, 1, x_296); +x_298 = lean_array_push(x_264, x_240); +lean_inc(x_298); +x_299 = lean_array_push(x_298, x_295); lean_inc(x_297); -x_298 = lean_array_push(x_297, x_294); -lean_inc(x_296); -x_299 = lean_array_push(x_298, x_296); -x_300 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__14; -x_301 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_301, 0, x_25); -lean_ctor_set(x_301, 1, x_300); -lean_ctor_set(x_301, 2, x_299); -x_302 = lean_array_push(x_46, x_63); -lean_inc(x_291); -x_303 = lean_array_push(x_302, x_291); -x_304 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_304, 0, x_25); -lean_ctor_set(x_304, 1, x_44); -lean_ctor_set(x_304, 2, x_303); +x_300 = lean_array_push(x_299, x_297); +x_301 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__14; +x_302 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_302, 0, x_26); +lean_ctor_set(x_302, 1, x_301); +lean_ctor_set(x_302, 2, x_300); +x_303 = lean_array_push(x_46, x_64); +x_304 = lean_array_push(x_303, x_293); +x_305 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_305, 0, x_26); +lean_ctor_set(x_305, 1, x_44); +lean_ctor_set(x_305, 2, x_304); +lean_inc(x_298); +x_306 = lean_array_push(x_298, x_305); lean_inc(x_297); -x_305 = lean_array_push(x_297, x_304); -lean_inc(x_296); -x_306 = lean_array_push(x_305, x_296); -x_307 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_307, 0, x_25); -lean_ctor_set(x_307, 1, x_300); -lean_ctor_set(x_307, 2, x_306); -x_308 = lean_array_push(x_46, x_301); -x_309 = lean_array_push(x_308, x_307); -x_310 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_310, 0, x_25); -lean_ctor_set(x_310, 1, x_44); -lean_ctor_set(x_310, 2, x_309); -x_311 = lean_array_push(x_46, x_249); -x_312 = lean_array_push(x_311, x_310); -x_313 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_313, 0, x_25); -lean_ctor_set(x_313, 1, x_49); -lean_ctor_set(x_313, 2, x_312); -x_314 = lean_array_push(x_46, x_313); -lean_inc(x_291); -x_315 = lean_array_push(x_314, x_291); -x_316 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_316, 0, x_25); -lean_ctor_set(x_316, 1, x_44); -lean_ctor_set(x_316, 2, x_315); +x_307 = lean_array_push(x_306, x_297); +x_308 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_308, 0, x_26); +lean_ctor_set(x_308, 1, x_301); +lean_ctor_set(x_308, 2, x_307); +x_309 = lean_array_push(x_46, x_302); +x_310 = lean_array_push(x_309, x_308); +x_311 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_311, 0, x_26); +lean_ctor_set(x_311, 1, x_44); +lean_ctor_set(x_311, 2, x_310); +x_312 = lean_array_push(x_46, x_250); +x_313 = lean_array_push(x_312, x_311); +x_314 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_314, 0, x_26); +lean_ctor_set(x_314, 1, x_49); +lean_ctor_set(x_314, 2, x_313); +x_315 = lean_array_push(x_46, x_314); +x_316 = lean_array_push(x_315, x_293); +x_317 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_317, 0, x_26); +lean_ctor_set(x_317, 1, x_44); +lean_ctor_set(x_317, 2, x_316); +lean_inc(x_298); +x_318 = lean_array_push(x_298, x_317); lean_inc(x_297); -x_317 = lean_array_push(x_297, x_316); -lean_inc(x_296); -x_318 = lean_array_push(x_317, x_296); -x_319 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_319, 0, x_25); -lean_ctor_set(x_319, 1, x_300); -lean_ctor_set(x_319, 2, x_318); -x_320 = lean_array_push(x_42, x_319); -x_321 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_321, 0, x_25); -lean_ctor_set(x_321, 1, x_44); -lean_ctor_set(x_321, 2, x_320); -x_322 = lean_array_push(x_46, x_244); -x_323 = lean_array_push(x_322, x_321); -x_324 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_324, 0, x_25); -lean_ctor_set(x_324, 1, x_49); -lean_ctor_set(x_324, 2, x_323); -x_325 = lean_array_push(x_46, x_324); -x_326 = lean_array_push(x_325, x_291); -x_327 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_327, 0, x_25); -lean_ctor_set(x_327, 1, x_44); -lean_ctor_set(x_327, 2, x_326); -x_328 = lean_array_push(x_297, x_327); -x_329 = lean_array_push(x_328, x_296); -x_330 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_330, 0, x_25); -lean_ctor_set(x_330, 1, x_300); -lean_ctor_set(x_330, 2, x_329); -x_331 = lean_array_push(x_46, x_330); -x_332 = lean_array_push(x_331, x_255); -x_333 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_333, 0, x_25); -lean_ctor_set(x_333, 1, x_44); -lean_ctor_set(x_333, 2, x_332); -x_334 = lean_array_push(x_46, x_237); -x_335 = lean_array_push(x_334, x_333); -x_336 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_336, 0, x_25); -lean_ctor_set(x_336, 1, x_49); -lean_ctor_set(x_336, 2, x_335); -x_337 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__54; -x_338 = lean_array_push(x_337, x_220); -x_339 = lean_array_push(x_338, x_230); -x_340 = lean_array_push(x_339, x_232); -x_341 = lean_array_push(x_340, x_336); -x_342 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__13; -x_343 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_343, 0, x_25); -lean_ctor_set(x_343, 1, x_342); -lean_ctor_set(x_343, 2, x_341); -x_344 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_344, 0, x_343); -lean_ctor_set(x_344, 1, x_216); -return x_344; +x_319 = lean_array_push(x_318, x_297); +x_320 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_320, 0, x_26); +lean_ctor_set(x_320, 1, x_301); +lean_ctor_set(x_320, 2, x_319); +x_321 = lean_array_push(x_42, x_320); +x_322 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_322, 0, x_26); +lean_ctor_set(x_322, 1, x_44); +lean_ctor_set(x_322, 2, x_321); +x_323 = lean_array_push(x_46, x_245); +x_324 = lean_array_push(x_323, x_322); +x_325 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_325, 0, x_26); +lean_ctor_set(x_325, 1, x_49); +lean_ctor_set(x_325, 2, x_324); +x_326 = lean_array_push(x_46, x_325); +x_327 = lean_array_push(x_326, x_293); +x_328 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_328, 0, x_26); +lean_ctor_set(x_328, 1, x_44); +lean_ctor_set(x_328, 2, x_327); +x_329 = lean_array_push(x_298, x_328); +x_330 = lean_array_push(x_329, x_297); +x_331 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_331, 0, x_26); +lean_ctor_set(x_331, 1, x_301); +lean_ctor_set(x_331, 2, x_330); +x_332 = lean_array_push(x_46, x_331); +x_333 = lean_array_push(x_332, x_256); +x_334 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_334, 0, x_26); +lean_ctor_set(x_334, 1, x_44); +lean_ctor_set(x_334, 2, x_333); +x_335 = lean_array_push(x_46, x_238); +x_336 = lean_array_push(x_335, x_334); +x_337 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_337, 0, x_26); +lean_ctor_set(x_337, 1, x_49); +lean_ctor_set(x_337, 2, x_336); +x_338 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__55; +x_339 = lean_array_push(x_338, x_221); +x_340 = lean_array_push(x_339, x_231); +x_341 = lean_array_push(x_340, x_233); +x_342 = lean_array_push(x_341, x_337); +x_343 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__13; +x_344 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_344, 0, x_26); +lean_ctor_set(x_344, 1, x_343); +lean_ctor_set(x_344, 2, x_342); +x_345 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_345, 0, x_344); +lean_ctor_set(x_345, 1, x_217); +return x_345; } } else { -uint8_t x_345; -lean_dec(x_30); +uint8_t x_346; +lean_dec(x_31); lean_dec(x_20); -lean_dec(x_14); lean_dec(x_13); -lean_dec(x_6); -lean_dec(x_2); -x_345 = !lean_is_exclusive(x_59); -if (x_345 == 0) +lean_dec(x_12); +lean_dec(x_5); +x_346 = !lean_is_exclusive(x_60); +if (x_346 == 0) { -return x_59; +return x_60; } else { -lean_object* x_346; lean_object* x_347; lean_object* x_348; -x_346 = lean_ctor_get(x_59, 0); -x_347 = lean_ctor_get(x_59, 1); +lean_object* x_347; lean_object* x_348; lean_object* x_349; +x_347 = lean_ctor_get(x_60, 0); +x_348 = lean_ctor_get(x_60, 1); +lean_inc(x_348); lean_inc(x_347); -lean_inc(x_346); -lean_dec(x_59); -x_348 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_348, 0, x_346); -lean_ctor_set(x_348, 1, x_347); -return x_348; +lean_dec(x_60); +x_349 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_349, 0, x_347); +lean_ctor_set(x_349, 1, x_348); +return x_349; } } } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3(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_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -if (lean_obj_tag(x_4) == 0) +if (lean_obj_tag(x_3) == 0) { -lean_object* x_13; -lean_dec(x_11); +lean_object* x_12; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_3); +lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_5); -lean_ctor_set(x_13, 1, x_12); -return x_13; +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_4); +lean_ctor_set(x_12, 1, x_11); +return x_12; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_4, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_4, 1); -lean_inc(x_15); -lean_dec(x_4); -lean_inc(x_6); +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_3, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_3, 1); lean_inc(x_14); -x_16 = l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__1(x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_16) == 0) +lean_dec(x_3); +lean_inc(x_5); +lean_inc(x_13); +x_15 = l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__1(x_13, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_17 = lean_ctor_get(x_16, 0); +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); +lean_dec(x_15); +x_18 = lean_ctor_get(x_16, 0); lean_inc(x_18); -lean_dec(x_16); -x_19 = lean_ctor_get(x_17, 0); +x_19 = lean_ctor_get(x_18, 2); lean_inc(x_19); -x_20 = lean_ctor_get(x_19, 2); -lean_inc(x_20); lean_inc(x_2); -lean_inc(x_3); lean_inc(x_1); -x_21 = lean_alloc_closure((void*)(l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___boxed), 15, 6); -lean_closure_set(x_21, 0, x_1); -lean_closure_set(x_21, 1, x_3); -lean_closure_set(x_21, 2, x_19); -lean_closure_set(x_21, 3, x_17); -lean_closure_set(x_21, 4, x_2); -lean_closure_set(x_21, 5, x_14); -lean_inc(x_11); +x_20 = lean_alloc_closure((void*)(l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___boxed), 14, 5); +lean_closure_set(x_20, 0, x_1); +lean_closure_set(x_20, 1, x_18); +lean_closure_set(x_20, 2, x_16); +lean_closure_set(x_20, 3, x_2); +lean_closure_set(x_20, 4, x_13); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_22 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Elab_Deriving_mkInductArgNames___spec__2___rarg(x_20, x_21, x_6, x_7, x_8, x_9, x_10, x_11, x_18); -if (lean_obj_tag(x_22) == 0) +lean_inc(x_5); +x_21 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Elab_Deriving_mkInductArgNames___spec__2___rarg(x_19, x_20, x_5, x_6, x_7, x_8, x_9, x_10, x_17); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_22, 0); +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_21, 1); lean_inc(x_23); -x_24 = lean_ctor_get(x_22, 1); -lean_inc(x_24); -lean_dec(x_22); -x_25 = lean_array_push(x_5, x_23); -x_4 = x_15; -x_5 = x_25; -x_12 = x_24; +lean_dec(x_21); +x_24 = lean_array_push(x_4, x_22); +x_3 = x_14; +x_4 = x_24; +x_11 = x_23; goto _start; } else { -uint8_t x_27; -lean_dec(x_15); -lean_dec(x_11); +uint8_t x_26; +lean_dec(x_14); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_3); +lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_27 = !lean_is_exclusive(x_22); -if (x_27 == 0) +x_26 = !lean_is_exclusive(x_21); +if (x_26 == 0) { -return x_22; +return x_21; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_22, 0); -x_29 = lean_ctor_get(x_22, 1); -lean_inc(x_29); +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_21, 0); +x_28 = lean_ctor_get(x_21, 1); lean_inc(x_28); -lean_dec(x_22); -x_30 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -return x_30; +lean_inc(x_27); +lean_dec(x_21); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; } } } else { -uint8_t x_31; -lean_dec(x_15); +uint8_t x_30; lean_dec(x_14); -lean_dec(x_11); +lean_dec(x_13); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_3); +lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_31 = !lean_is_exclusive(x_16); -if (x_31 == 0) +x_30 = !lean_is_exclusive(x_15); +if (x_30 == 0) { -return x_16; +return x_15; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_16, 0); -x_33 = lean_ctor_get(x_16, 1); -lean_inc(x_33); +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_15, 0); +x_32 = lean_ctor_get(x_15, 1); lean_inc(x_32); -lean_dec(x_16); -x_34 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_34, 0, x_32); -lean_ctor_set(x_34, 1, x_33); -return x_34; +lean_inc(x_31); +lean_dec(x_15); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; } } } @@ -5363,7 +5398,7 @@ lean_object* x_10; lean_object* x_11; lean_object* x_12; x_10 = lean_ctor_get(x_1, 4); lean_inc(x_10); x_11 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__27; -x_12 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3(x_1, x_2, x_11, x_10, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_12 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3(x_1, x_2, x_10, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9); if (lean_obj_tag(x_12) == 0) { uint8_t x_13; @@ -5451,30 +5486,56 @@ lean_object* x_22 = _args[21]; lean_object* x_23 = _args[22]; lean_object* x_24 = _args[23]; lean_object* x_25 = _args[24]; +lean_object* x_26 = _args[25]; _start: { -lean_object* x_26; -x_26 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_24, x_25); +lean_object* x_27; +x_27 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_24, x_25, x_26); +lean_dec(x_21); lean_dec(x_20); -lean_dec(x_19); +lean_dec(x_18); lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_13); +lean_dec(x_14); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -return x_26; +return x_27; } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___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, 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) { +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___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, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_16; -x_16 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -lean_dec(x_10); +lean_object* x_15; +x_15 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -return x_16; +lean_dec(x_6); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_Repr_mkBodyForInduct___spec__1(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = 1; +x_9 = lean_usize_add(x_2, x_8); +x_10 = lean_array_uset(x_7, x_2, x_5); +x_2 = x_9; +x_3 = x_10; +goto _start; +} } } static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkBodyForInduct___closed__1() { @@ -5567,7 +5628,7 @@ x_25 = lean_array_get_size(x_12); x_26 = lean_usize_of_nat(x_25); lean_dec(x_25); x_27 = 0; -x_28 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_26, x_27, x_12); +x_28 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_Repr_mkBodyForInduct___spec__1(x_26, x_27, x_12); x_29 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__16; x_30 = l_Lean_mkSepArray(x_28, x_29); lean_dec(x_28); @@ -5626,7 +5687,7 @@ x_57 = lean_array_get_size(x_12); x_58 = lean_usize_of_nat(x_57); lean_dec(x_57); x_59 = 0; -x_60 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_58, x_59, x_12); +x_60 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_Repr_mkBodyForInduct___spec__1(x_58, x_59, x_12); x_61 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__16; x_62 = l_Lean_mkSepArray(x_60, x_61); lean_dec(x_60); @@ -5700,6 +5761,18 @@ return x_90; } } } +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_Repr_mkBodyForInduct___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_Repr_mkBodyForInduct___spec__1(x_4, x_5, x_3); +return x_6; +} +} LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Repr_mkBody(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_object* x_9, lean_object* x_10) { _start: { @@ -5903,7 +5976,7 @@ static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___ _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__4; +x_1 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__7; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } @@ -5912,7 +5985,7 @@ static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__4; +x_1 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__7; x_2 = lean_unsigned_to_nat(0u); x_3 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__19; x_4 = lean_alloc_ctor(0, 3, 0); @@ -5927,7 +6000,7 @@ static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__10; +x_2 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__13; x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); @@ -6085,7 +6158,7 @@ lean_inc(x_16); x_54 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_54, 0, x_16); lean_ctor_set(x_54, 1, x_53); -x_55 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__5; +x_55 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__8; x_56 = l_Lean_addMacroScope(x_23, x_55, x_18); x_57 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__20; x_58 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__22; @@ -6217,7 +6290,7 @@ lean_inc(x_16); x_126 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_126, 0, x_16); lean_ctor_set(x_126, 1, x_125); -x_127 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__5; +x_127 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__8; x_128 = l_Lean_addMacroScope(x_95, x_127, x_18); x_129 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__20; x_130 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__22; @@ -6382,7 +6455,7 @@ lean_inc(x_167); x_212 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_212, 0, x_167); lean_ctor_set(x_212, 1, x_211); -x_213 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__5; +x_213 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__8; x_214 = l_Lean_addMacroScope(x_174, x_213, x_169); x_215 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__20; x_216 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__22; @@ -6530,7 +6603,7 @@ lean_inc(x_167); x_291 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_291, 0, x_167); lean_ctor_set(x_291, 1, x_290); -x_292 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__5; +x_292 = l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__8; x_293 = l_Lean_addMacroScope(x_253, x_292, x_169); x_294 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__20; x_295 = l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__22; @@ -6963,20 +7036,6 @@ x_1 = lean_mk_string_from_bytes("end", 3); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkMutualBlock___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = lean_box(2); -x_2 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__13; -x_3 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__27; -x_4 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_4, 0, x_1); -lean_ctor_set(x_4, 1, x_2); -lean_ctor_set(x_4, 2, x_3); -return x_4; -} -} LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Repr_mkMutualBlock(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: { @@ -7035,7 +7094,7 @@ x_31 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__30; x_32 = lean_array_push(x_31, x_24); x_33 = lean_array_push(x_32, x_28); x_34 = lean_array_push(x_33, x_30); -x_35 = l_Lean_Elab_Deriving_Repr_mkMutualBlock___closed__4; +x_35 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__54; x_36 = lean_array_push(x_34, x_35); x_37 = lean_array_push(x_36, x_35); x_38 = l_Lean_Elab_Deriving_Repr_mkMutualBlock___closed__2; @@ -7072,7 +7131,7 @@ x_49 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__30; x_50 = lean_array_push(x_49, x_42); x_51 = lean_array_push(x_50, x_46); x_52 = lean_array_push(x_51, x_48); -x_53 = l_Lean_Elab_Deriving_Repr_mkMutualBlock___closed__4; +x_53 = l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__54; x_54 = lean_array_push(x_52, x_53); x_55 = lean_array_push(x_54, x_53); x_56 = l_Lean_Elab_Deriving_Repr_mkMutualBlock___closed__2; @@ -8081,7 +8140,7 @@ lean_dec(x_1); return x_9; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3319____closed__1() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3494____closed__1() { _start: { lean_object* x_1; @@ -8089,12 +8148,12 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Deriving_Repr_mkReprInstanceHandler return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3319_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3494_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; x_2 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__2; -x_3 = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3319____closed__1; +x_3 = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3494____closed__1; x_4 = l_Lean_Elab_registerBuiltinDerivingHandler(x_2, x_3, x_1); if (lean_obj_tag(x_4) == 0) { @@ -8309,6 +8368,12 @@ l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__14 = _init_l_Lea lean_mark_persistent(l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__14); l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__15 = _init_l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__15(); lean_mark_persistent(l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__15); +l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__16 = _init_l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__16(); +lean_mark_persistent(l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__16); +l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__17 = _init_l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__17(); +lean_mark_persistent(l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__17); +l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__18 = _init_l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__18(); +lean_mark_persistent(l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__18); l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__1___closed__1 = _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__1___closed__1(); lean_mark_persistent(l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__1___closed__1); l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__1___closed__2 = _init_l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__1___closed__2(); @@ -8439,6 +8504,8 @@ l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3_ lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__53); l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__54 = _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__54(); lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__54); +l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__55 = _init_l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__55(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__55); l_Lean_Elab_Deriving_Repr_mkBodyForInduct___closed__1 = _init_l_Lean_Elab_Deriving_Repr_mkBodyForInduct___closed__1(); lean_mark_persistent(l_Lean_Elab_Deriving_Repr_mkBodyForInduct___closed__1); l_Lean_Elab_Deriving_Repr_mkBodyForInduct___closed__2 = _init_l_Lean_Elab_Deriving_Repr_mkBodyForInduct___closed__2(); @@ -8511,8 +8578,6 @@ l_Lean_Elab_Deriving_Repr_mkMutualBlock___closed__2 = _init_l_Lean_Elab_Deriving lean_mark_persistent(l_Lean_Elab_Deriving_Repr_mkMutualBlock___closed__2); l_Lean_Elab_Deriving_Repr_mkMutualBlock___closed__3 = _init_l_Lean_Elab_Deriving_Repr_mkMutualBlock___closed__3(); lean_mark_persistent(l_Lean_Elab_Deriving_Repr_mkMutualBlock___closed__3); -l_Lean_Elab_Deriving_Repr_mkMutualBlock___closed__4 = _init_l_Lean_Elab_Deriving_Repr_mkMutualBlock___closed__4(); -lean_mark_persistent(l_Lean_Elab_Deriving_Repr_mkMutualBlock___closed__4); l___private_Lean_Elab_Deriving_Repr_0__Lean_Elab_Deriving_Repr_mkReprInstanceCmds___closed__1 = _init_l___private_Lean_Elab_Deriving_Repr_0__Lean_Elab_Deriving_Repr_mkReprInstanceCmds___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Deriving_Repr_0__Lean_Elab_Deriving_Repr_mkReprInstanceCmds___closed__1); l___private_Lean_Elab_Deriving_Repr_0__Lean_Elab_Deriving_Repr_mkReprInstanceCmds___closed__2 = _init_l___private_Lean_Elab_Deriving_Repr_0__Lean_Elab_Deriving_Repr_mkReprInstanceCmds___closed__2(); @@ -8531,9 +8596,9 @@ l___private_Lean_Elab_Deriving_Repr_0__Lean_Elab_Deriving_Repr_mkReprInstanceCmd lean_mark_persistent(l___private_Lean_Elab_Deriving_Repr_0__Lean_Elab_Deriving_Repr_mkReprInstanceCmds___closed__8); l___private_Lean_Elab_Deriving_Repr_0__Lean_Elab_Deriving_Repr_mkReprInstanceCmds___closed__9 = _init_l___private_Lean_Elab_Deriving_Repr_0__Lean_Elab_Deriving_Repr_mkReprInstanceCmds___closed__9(); lean_mark_persistent(l___private_Lean_Elab_Deriving_Repr_0__Lean_Elab_Deriving_Repr_mkReprInstanceCmds___closed__9); -l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3319____closed__1 = _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3319____closed__1(); -lean_mark_persistent(l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3319____closed__1); -res = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3319_(lean_io_mk_world()); +l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3494____closed__1 = _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3494____closed__1(); +lean_mark_persistent(l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3494____closed__1); +res = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3494_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/Deriving/Util.c b/stage0/stdlib/Lean/Elab/Deriving/Util.c index 4e86a58e614..159b9819fd0 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/Util.c +++ b/stage0/stdlib/Lean/Elab/Deriving/Util.c @@ -21,7 +21,7 @@ size_t lean_usize_add(size_t, size_t); static lean_object* l_Lean_Elab_Deriving_mkInductiveApp___closed__12; lean_object* lean_erase_macro_scopes(lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_mkInstImplicitBinders___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -80,7 +80,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_instBinderF; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_mkContext___spec__4___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_mkDiscrs___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstImplicitBinders___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkHeader___spec__2___closed__4; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_mkInductArgNames(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_mkDiscrs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -113,7 +112,7 @@ lean_object* l_Lean_Expr_fvarId_x21(lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__16; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkDiscrs___spec__2(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstImplicitBinders___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkHeader___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkHeader___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstImplicitBinders___spec__1___closed__1; lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Array_contains___at_Lean_findField_x3f___spec__1(lean_object*, lean_object*); @@ -124,11 +123,11 @@ static lean_object* l_Lean_Elab_Deriving_explicitBinderF___closed__1; lean_object* lean_array_to_list(lean_object*, lean_object*); static lean_object* l_Lean_getConstInfoInduct___at_Lean_Elab_Deriving_mkContext___spec__1___closed__3; lean_object* l_Lean_Parser_Term_implicitBinder(uint8_t); -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfoInduct___at_Lean_Elab_Deriving_mkContext___spec__1___closed__4; extern lean_object* l_Lean_instInhabitedExpr; static lean_object* l_Lean_Elab_Deriving_mkContext___closed__7; -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, 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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_Term_explicitBinder(uint8_t); LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfoInduct___at_Lean_Elab_Deriving_mkContext___spec__1___closed__2; @@ -143,8 +142,8 @@ static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkHeader___sp LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_mkHeader(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_mkInductiveApp___closed__3; LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkHeader___spec__2___boxed(lean_object*, 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_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkHeader___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOptionM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTerm___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkHeader___spec__2___closed__1; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1___closed__7; @@ -186,7 +185,7 @@ static lean_object* l_Lean_Elab_Deriving_implicitBinderF___closed__1; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkImplicitBinders___spec__1___closed__4; static lean_object* l_Lean_Elab_Deriving_mkInductArgNames___lambda__1___closed__1; static lean_object* l_Lean_Elab_Deriving_mkContext___closed__3; -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_mkInductArgNames___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstImplicitBinders___spec__1___closed__3; @@ -2183,240 +2182,239 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__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, 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, lean_object* x_16, lean_object* x_17) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__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, 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, lean_object* x_16) { _start: { -uint8_t x_18; -x_18 = lean_nat_dec_le(x_8, x_7); -if (x_18 == 0) +uint8_t x_17; +x_17 = lean_nat_dec_le(x_7, x_6); +if (x_17 == 0) { -lean_object* x_19; uint8_t x_20; -x_19 = lean_unsigned_to_nat(0u); -x_20 = lean_nat_dec_eq(x_6, x_19); -if (x_20 == 0) +lean_object* x_18; uint8_t x_19; +x_18 = lean_unsigned_to_nat(0u); +x_19 = lean_nat_dec_eq(x_5, x_18); +if (x_19 == 0) { -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; -x_21 = lean_unsigned_to_nat(1u); -x_22 = lean_nat_sub(x_6, x_21); -lean_dec(x_6); -x_23 = l_Lean_instInhabitedInductiveVal; -x_24 = lean_array_get(x_23, x_5, x_7); -x_25 = lean_ctor_get(x_1, 1); -x_26 = l_Lean_instInhabitedName; -x_27 = lean_array_get(x_26, x_25, x_7); -lean_inc(x_16); +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; +x_20 = lean_unsigned_to_nat(1u); +x_21 = lean_nat_sub(x_5, x_20); +lean_dec(x_5); +x_22 = l_Lean_instInhabitedInductiveVal; +x_23 = lean_array_get(x_22, x_4, x_6); +x_24 = lean_ctor_get(x_1, 1); +x_25 = l_Lean_instInhabitedName; +x_26 = lean_array_get(x_25, x_24, x_6); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); -lean_inc(x_24); -x_28 = l_Lean_Elab_Deriving_mkInductArgNames(x_24, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -if (lean_obj_tag(x_28) == 0) +lean_inc(x_10); +lean_inc(x_23); +x_27 = l_Lean_Elab_Deriving_mkInductArgNames(x_23, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_27) == 0) { -lean_object* x_29; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; -x_29 = lean_ctor_get(x_28, 0); +lean_object* x_28; lean_object* x_29; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_27, 1); lean_inc(x_29); -x_30 = lean_ctor_get(x_28, 1); +lean_dec(x_27); +x_30 = lean_ctor_get(x_23, 1); lean_inc(x_30); -lean_dec(x_28); -x_31 = lean_ctor_get(x_24, 1); -lean_inc(x_31); -x_32 = lean_array_get_size(x_29); -lean_inc(x_31); -x_33 = l_Array_toSubarray___rarg(x_29, x_31, x_32); -x_34 = l_Array_ofSubarray___rarg(x_33); -lean_inc(x_16); +x_31 = lean_array_get_size(x_28); +lean_inc(x_30); +x_32 = l_Array_toSubarray___rarg(x_28, x_30, x_31); +x_33 = l_Array_ofSubarray___rarg(x_32); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); -lean_inc(x_34); -x_35 = l_Lean_Elab_Deriving_mkImplicitBinders(x_34, x_11, x_12, x_13, x_14, x_15, x_16, x_30); -x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_10); +lean_inc(x_33); +x_34 = l_Lean_Elab_Deriving_mkImplicitBinders(x_33, x_10, x_11, x_12, x_13, x_14, x_15, x_29); +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_34, 1); lean_inc(x_36); -x_37 = lean_ctor_get(x_35, 1); -lean_inc(x_37); -lean_dec(x_35); +lean_dec(x_34); lean_inc(x_3); -x_38 = l_Array_toSubarray___rarg(x_3, x_19, x_31); -x_39 = l_Array_ofSubarray___rarg(x_38); -x_40 = l_Array_append___rarg(x_39, x_34); -x_41 = lean_array_get_size(x_40); -x_42 = l_Array_toSubarray___rarg(x_40, x_19, x_41); -x_43 = l_Array_ofSubarray___rarg(x_42); -lean_inc(x_15); -x_44 = l_Lean_Elab_Deriving_mkInductiveApp(x_24, x_43, x_11, x_12, x_13, x_14, x_15, x_16, x_37); -x_45 = lean_ctor_get(x_44, 0); +x_37 = l_Array_toSubarray___rarg(x_3, x_18, x_30); +x_38 = l_Array_ofSubarray___rarg(x_37); +x_39 = l_Array_append___rarg(x_38, x_33); +x_40 = lean_array_get_size(x_39); +x_41 = l_Array_toSubarray___rarg(x_39, x_18, x_40); +x_42 = l_Array_ofSubarray___rarg(x_41); +lean_inc(x_14); +x_43 = l_Lean_Elab_Deriving_mkInductiveApp(x_23, x_42, x_10, x_11, x_12, x_13, x_14, x_15, x_36); +x_44 = lean_ctor_get(x_43, 0); +lean_inc(x_44); +x_45 = lean_ctor_get(x_43, 1); lean_inc(x_45); -x_46 = lean_ctor_get(x_44, 1); -lean_inc(x_46); -lean_dec(x_44); -lean_inc(x_15); -x_47 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_15, x_16, x_46); -x_48 = lean_ctor_get(x_47, 1); -lean_inc(x_48); -lean_dec(x_47); -x_49 = lean_st_ref_get(x_16, x_48); -x_50 = lean_ctor_get(x_49, 1); -lean_inc(x_50); -lean_dec(x_49); +lean_dec(x_43); +lean_inc(x_14); +x_46 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_14, x_15, x_45); +x_47 = lean_ctor_get(x_46, 1); +lean_inc(x_47); +lean_dec(x_46); +x_48 = lean_st_ref_get(x_15, x_47); +x_49 = lean_ctor_get(x_48, 1); +lean_inc(x_49); +lean_dec(x_48); lean_inc(x_2); -x_51 = lean_mk_syntax_ident(x_2); -x_52 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkImplicitBinders___spec__1___closed__4; -x_53 = lean_array_push(x_52, x_45); -x_54 = lean_box(2); -x_55 = l_Lean_Elab_Deriving_mkInductiveApp___closed__14; -x_56 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_56, 0, x_54); -lean_ctor_set(x_56, 1, x_55); -lean_ctor_set(x_56, 2, x_53); -x_57 = l_Lean_Elab_Deriving_mkInductiveApp___closed__12; -x_58 = lean_array_push(x_57, x_51); -x_59 = lean_array_push(x_58, x_56); -x_60 = l_Lean_Elab_Deriving_mkInductiveApp___closed__8; -x_61 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_61, 0, x_54); -lean_ctor_set(x_61, 1, x_60); -lean_ctor_set(x_61, 2, x_59); -lean_inc(x_15); -x_62 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_15, x_16, x_50); -x_63 = lean_ctor_get(x_62, 0); -lean_inc(x_63); -x_64 = lean_ctor_get(x_62, 1); -lean_inc(x_64); -lean_dec(x_62); -x_65 = lean_st_ref_get(x_16, x_64); -x_66 = lean_ctor_get(x_65, 1); -lean_inc(x_66); -lean_dec(x_65); -x_67 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__3; +x_50 = lean_mk_syntax_ident(x_2); +x_51 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkImplicitBinders___spec__1___closed__4; +x_52 = lean_array_push(x_51, x_44); +x_53 = lean_box(2); +x_54 = l_Lean_Elab_Deriving_mkInductiveApp___closed__14; +x_55 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +lean_ctor_set(x_55, 2, x_52); +x_56 = l_Lean_Elab_Deriving_mkInductiveApp___closed__12; +x_57 = lean_array_push(x_56, x_50); +x_58 = lean_array_push(x_57, x_55); +x_59 = l_Lean_Elab_Deriving_mkInductiveApp___closed__8; +x_60 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_60, 0, x_53); +lean_ctor_set(x_60, 1, x_59); +lean_ctor_set(x_60, 2, x_58); +lean_inc(x_14); +x_61 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_14, x_15, x_49); +x_62 = lean_ctor_get(x_61, 0); +lean_inc(x_62); +x_63 = lean_ctor_get(x_61, 1); lean_inc(x_63); -x_68 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_68, 0, x_63); -lean_ctor_set(x_68, 1, x_67); -x_69 = lean_mk_syntax_ident(x_27); -x_70 = lean_array_push(x_52, x_69); -x_71 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_71, 0, x_54); -lean_ctor_set(x_71, 1, x_55); -lean_ctor_set(x_71, 2, x_70); -x_72 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__4; -x_73 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_73, 0, x_63); -lean_ctor_set(x_73, 1, x_72); -x_74 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__5; -x_75 = lean_array_push(x_74, x_68); -x_76 = lean_array_push(x_75, x_71); -x_77 = lean_array_push(x_76, x_73); -x_78 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__2; -x_79 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_79, 0, x_54); -lean_ctor_set(x_79, 1, x_78); -lean_ctor_set(x_79, 2, x_77); -x_80 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__7; -x_81 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_80, x_15, x_16, x_66); -x_82 = lean_ctor_get(x_81, 0); +lean_dec(x_61); +x_64 = lean_st_ref_get(x_15, x_63); +x_65 = lean_ctor_get(x_64, 1); +lean_inc(x_65); +lean_dec(x_64); +x_66 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__3; +lean_inc(x_62); +x_67 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_67, 0, x_62); +lean_ctor_set(x_67, 1, x_66); +x_68 = lean_mk_syntax_ident(x_26); +x_69 = lean_array_push(x_51, x_68); +x_70 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_70, 0, x_53); +lean_ctor_set(x_70, 1, x_54); +lean_ctor_set(x_70, 2, x_69); +x_71 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__4; +x_72 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_72, 0, x_62); +lean_ctor_set(x_72, 1, x_71); +x_73 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__5; +x_74 = lean_array_push(x_73, x_67); +x_75 = lean_array_push(x_74, x_70); +x_76 = lean_array_push(x_75, x_72); +x_77 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__2; +x_78 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_78, 0, x_53); +lean_ctor_set(x_78, 1, x_77); +lean_ctor_set(x_78, 2, x_76); +x_79 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__7; +x_80 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_79, x_14, x_15, x_65); +x_81 = lean_ctor_get(x_80, 0); +lean_inc(x_81); +x_82 = lean_ctor_get(x_80, 1); lean_inc(x_82); -x_83 = lean_ctor_get(x_81, 1); -lean_inc(x_83); -lean_dec(x_81); -lean_inc(x_15); -x_84 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_15, x_16, x_83); -x_85 = lean_ctor_get(x_84, 0); +lean_dec(x_80); +lean_inc(x_14); +x_83 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_14, x_15, x_82); +x_84 = lean_ctor_get(x_83, 0); +lean_inc(x_84); +x_85 = lean_ctor_get(x_83, 1); lean_inc(x_85); -x_86 = lean_ctor_get(x_84, 1); -lean_inc(x_86); -lean_dec(x_84); -x_87 = lean_st_ref_get(x_16, x_86); -x_88 = lean_ctor_get(x_87, 1); -lean_inc(x_88); -lean_dec(x_87); -x_89 = lean_mk_syntax_ident(x_82); -lean_inc(x_4); -x_90 = l_Array_append___rarg(x_4, x_36); +lean_dec(x_83); +x_86 = lean_st_ref_get(x_15, x_85); +x_87 = lean_ctor_get(x_86, 1); +lean_inc(x_87); +lean_dec(x_86); +x_88 = lean_mk_syntax_ident(x_81); +x_89 = l_Lean_Elab_Deriving_mkInductArgNames___lambda__1___closed__1; +x_90 = l_Array_append___rarg(x_89, x_35); x_91 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_91, 0, x_54); -lean_ctor_set(x_91, 1, x_55); +lean_ctor_set(x_91, 0, x_53); +lean_ctor_set(x_91, 1, x_54); lean_ctor_set(x_91, 2, x_90); x_92 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__14; -lean_inc(x_85); +lean_inc(x_84); x_93 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_93, 0, x_85); +lean_ctor_set(x_93, 0, x_84); lean_ctor_set(x_93, 1, x_92); -x_94 = lean_array_push(x_57, x_93); -x_95 = lean_array_push(x_94, x_61); +x_94 = lean_array_push(x_56, x_93); +x_95 = lean_array_push(x_94, x_60); x_96 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__13; x_97 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_97, 0, x_54); +lean_ctor_set(x_97, 0, x_53); lean_ctor_set(x_97, 1, x_96); lean_ctor_set(x_97, 2, x_95); -x_98 = lean_array_push(x_52, x_97); +x_98 = lean_array_push(x_51, x_97); x_99 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_99, 0, x_54); -lean_ctor_set(x_99, 1, x_55); +lean_ctor_set(x_99, 0, x_53); +lean_ctor_set(x_99, 1, x_54); lean_ctor_set(x_99, 2, x_98); x_100 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__15; x_101 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_101, 0, x_85); +lean_ctor_set(x_101, 0, x_84); lean_ctor_set(x_101, 1, x_100); x_102 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__16; -x_103 = lean_array_push(x_102, x_89); +x_103 = lean_array_push(x_102, x_88); x_104 = lean_array_push(x_103, x_91); x_105 = lean_array_push(x_104, x_99); x_106 = lean_array_push(x_105, x_101); -x_107 = lean_array_push(x_106, x_79); +x_107 = lean_array_push(x_106, x_78); x_108 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__11; x_109 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_109, 0, x_54); +lean_ctor_set(x_109, 0, x_53); lean_ctor_set(x_109, 1, x_108); lean_ctor_set(x_109, 2, x_107); -x_110 = lean_array_push(x_52, x_109); +x_110 = lean_array_push(x_51, x_109); x_111 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__9; x_112 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_112, 0, x_54); +lean_ctor_set(x_112, 0, x_53); lean_ctor_set(x_112, 1, x_111); lean_ctor_set(x_112, 2, x_110); -x_113 = lean_array_push(x_10, x_112); -x_114 = lean_nat_add(x_7, x_9); -lean_dec(x_7); -x_6 = x_22; -x_7 = x_114; -x_10 = x_113; -x_17 = x_88; +x_113 = lean_array_push(x_9, x_112); +x_114 = lean_nat_add(x_6, x_8); +lean_dec(x_6); +x_5 = x_21; +x_6 = x_114; +x_9 = x_113; +x_16 = x_87; goto _start; } else { uint8_t x_116; -lean_dec(x_27); -lean_dec(x_24); -lean_dec(x_22); -lean_dec(x_16); +lean_dec(x_26); +lean_dec(x_23); +lean_dec(x_21); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_7); -lean_dec(x_4); +lean_dec(x_9); +lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); -x_116 = !lean_is_exclusive(x_28); +x_116 = !lean_is_exclusive(x_27); if (x_116 == 0) { -return x_28; +return x_27; } else { lean_object* x_117; lean_object* x_118; lean_object* x_119; -x_117 = lean_ctor_get(x_28, 0); -x_118 = lean_ctor_get(x_28, 1); +x_117 = lean_ctor_get(x_27, 0); +x_118 = lean_ctor_get(x_27, 1); lean_inc(x_118); lean_inc(x_117); -lean_dec(x_28); +lean_dec(x_27); x_119 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_119, 0, x_117); lean_ctor_set(x_119, 1, x_118); @@ -2427,40 +2425,38 @@ return x_119; else { lean_object* x_120; -lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_7); +lean_dec(x_10); lean_dec(x_6); -lean_dec(x_4); +lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); x_120 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_120, 0, x_10); -lean_ctor_set(x_120, 1, x_17); +lean_ctor_set(x_120, 0, x_9); +lean_ctor_set(x_120, 1, x_16); return x_120; } } else { lean_object* x_121; -lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_7); +lean_dec(x_10); lean_dec(x_6); -lean_dec(x_4); +lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); x_121 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_121, 0, x_10); -lean_ctor_set(x_121, 1, x_17); +lean_ctor_set(x_121, 0, x_9); +lean_ctor_set(x_121, 1, x_16); return x_121; } } @@ -2471,11 +2467,11 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_mkLocalInstanceLetDecls(lean_objec lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; x_11 = lean_ctor_get(x_1, 0); x_12 = lean_array_get_size(x_11); -x_13 = l_Lean_Elab_Deriving_mkInductArgNames___lambda__1___closed__1; -x_14 = lean_unsigned_to_nat(0u); -x_15 = lean_unsigned_to_nat(1u); +x_13 = lean_unsigned_to_nat(0u); +x_14 = lean_unsigned_to_nat(1u); +x_15 = l_Lean_Elab_Deriving_mkInductArgNames___lambda__1___closed__1; lean_inc(x_12); -x_16 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1(x_1, x_2, x_3, x_13, x_11, x_12, x_14, x_12, x_15, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_16 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1(x_1, x_2, x_3, x_11, x_12, x_13, x_12, x_14, x_15, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_12); if (lean_obj_tag(x_16) == 0) { @@ -2523,33 +2519,16 @@ return x_24; } } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__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, 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, lean_object* x_16) { _start: { -lean_object* x_18; -x_18 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -lean_dec(x_9); +lean_object* x_17; +x_17 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); lean_dec(x_8); -lean_dec(x_5); +lean_dec(x_7); +lean_dec(x_4); lean_dec(x_1); -return x_18; +return x_17; } } LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_mkLocalInstanceLetDecls___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_object* x_9, lean_object* x_10) { @@ -2784,38 +2763,38 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___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, lean_object* x_7, 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, lean_object* x_15, lean_object* x_16, lean_object* x_17) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___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, lean_object* x_7, 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, lean_object* x_15, lean_object* x_16) { _start: { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_18 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_15, x_16, x_17); -x_19 = lean_ctor_get(x_18, 0); +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_17 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_14, x_15, x_16); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); -lean_inc(x_20); -lean_dec(x_18); -x_21 = lean_st_ref_get(x_16, x_20); -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) +lean_dec(x_17); +x_20 = lean_st_ref_get(x_15, x_19); +x_21 = !lean_is_exclusive(x_20); +if (x_21 == 0) { -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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; -x_23 = lean_ctor_get(x_21, 0); -lean_dec(x_23); -x_24 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1___closed__1; -x_25 = lean_name_mk_string(x_1, x_24); -x_26 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1___closed__2; -lean_inc(x_25); -x_27 = lean_name_mk_string(x_25, x_26); -x_28 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1___closed__3; -lean_inc(x_25); -x_29 = lean_name_mk_string(x_25, x_28); -x_30 = lean_box(2); -lean_inc(x_3); +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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_22 = lean_ctor_get(x_20, 0); +lean_dec(x_22); +x_23 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1___closed__1; +x_24 = lean_name_mk_string(x_1, x_23); +x_25 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1___closed__2; +lean_inc(x_24); +x_26 = lean_name_mk_string(x_24, x_25); +x_27 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1___closed__3; +lean_inc(x_24); +x_28 = lean_name_mk_string(x_24, x_27); +x_29 = lean_box(2); +x_30 = l_Lean_Elab_Deriving_mkInductArgNames___lambda__1___closed__1; lean_inc(x_2); x_31 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 0, x_29); lean_ctor_set(x_31, 1, x_2); -lean_ctor_set(x_31, 2, x_3); +lean_ctor_set(x_31, 2, x_30); x_32 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1___closed__4; lean_inc(x_31); x_33 = lean_array_push(x_32, x_31); @@ -2830,67 +2809,67 @@ x_37 = lean_array_push(x_36, x_31); lean_inc(x_31); x_38 = lean_array_push(x_37, x_31); x_39 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_39, 0, x_30); -lean_ctor_set(x_39, 1, x_29); +lean_ctor_set(x_39, 0, x_29); +lean_ctor_set(x_39, 1, x_28); lean_ctor_set(x_39, 2, x_38); x_40 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1___closed__5; -lean_inc(x_25); -x_41 = lean_name_mk_string(x_25, x_40); +lean_inc(x_24); +x_41 = lean_name_mk_string(x_24, x_40); x_42 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1___closed__6; -lean_inc(x_4); -x_43 = lean_name_mk_string(x_4, x_42); +lean_inc(x_3); +x_43 = lean_name_mk_string(x_3, x_42); lean_inc(x_31); -x_44 = lean_array_push(x_5, x_31); +x_44 = lean_array_push(x_4, x_31); x_45 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_45, 0, x_30); +lean_ctor_set(x_45, 0, x_29); lean_ctor_set(x_45, 1, x_43); lean_ctor_set(x_45, 2, x_44); -lean_inc(x_19); +lean_inc(x_18); x_46 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_46, 0, x_19); +lean_ctor_set(x_46, 0, x_18); lean_ctor_set(x_46, 1, x_40); x_47 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1___closed__7; -lean_inc(x_25); -x_48 = lean_name_mk_string(x_25, x_47); -x_49 = l_Array_append___rarg(x_3, x_6); +lean_inc(x_24); +x_48 = lean_name_mk_string(x_24, x_47); +x_49 = l_Array_append___rarg(x_30, x_5); x_50 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_50, 0, x_30); +lean_ctor_set(x_50, 0, x_29); lean_ctor_set(x_50, 1, x_2); lean_ctor_set(x_50, 2, x_49); x_51 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__12; -x_52 = lean_name_mk_string(x_4, x_51); +x_52 = lean_name_mk_string(x_3, x_51); x_53 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__14; -lean_inc(x_19); +lean_inc(x_18); x_54 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_54, 0, x_19); +lean_ctor_set(x_54, 0, x_18); lean_ctor_set(x_54, 1, x_53); -lean_inc(x_7); -x_55 = lean_array_push(x_7, x_54); -x_56 = lean_array_push(x_55, x_8); +lean_inc(x_6); +x_55 = lean_array_push(x_6, x_54); +x_56 = lean_array_push(x_55, x_7); x_57 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_57, 0, x_30); +lean_ctor_set(x_57, 0, x_29); lean_ctor_set(x_57, 1, x_52); lean_ctor_set(x_57, 2, x_56); -lean_inc(x_7); -x_58 = lean_array_push(x_7, x_50); +lean_inc(x_6); +x_58 = lean_array_push(x_6, x_50); x_59 = lean_array_push(x_58, x_57); x_60 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_60, 0, x_30); +lean_ctor_set(x_60, 0, x_29); lean_ctor_set(x_60, 1, x_48); lean_ctor_set(x_60, 2, x_59); x_61 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1___closed__8; -x_62 = lean_name_mk_string(x_25, x_61); +x_62 = lean_name_mk_string(x_24, x_61); x_63 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__15; x_64 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_64, 0, x_19); +lean_ctor_set(x_64, 0, x_18); lean_ctor_set(x_64, 1, x_63); x_65 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__5; x_66 = lean_array_push(x_65, x_64); -x_67 = lean_array_push(x_66, x_10); +x_67 = lean_array_push(x_66, x_9); lean_inc(x_31); x_68 = lean_array_push(x_67, x_31); x_69 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_69, 0, x_30); +lean_ctor_set(x_69, 0, x_29); lean_ctor_set(x_69, 1, x_62); lean_ctor_set(x_69, 2, x_68); x_70 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1___closed__9; @@ -2906,27 +2885,27 @@ lean_inc(x_31); x_77 = lean_array_push(x_76, x_31); x_78 = lean_array_push(x_77, x_31); x_79 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_79, 0, x_30); +lean_ctor_set(x_79, 0, x_29); lean_ctor_set(x_79, 1, x_41); lean_ctor_set(x_79, 2, x_78); -x_80 = lean_array_push(x_7, x_39); +x_80 = lean_array_push(x_6, x_39); x_81 = lean_array_push(x_80, x_79); x_82 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_82, 0, x_30); -lean_ctor_set(x_82, 1, x_27); +lean_ctor_set(x_82, 0, x_29); +lean_ctor_set(x_82, 1, x_26); lean_ctor_set(x_82, 2, x_81); -x_83 = lean_array_push(x_9, x_82); +x_83 = lean_array_push(x_8, x_82); x_84 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_84, 0, x_83); -lean_ctor_set(x_21, 0, x_84); -return x_21; +lean_ctor_set(x_20, 0, x_84); +return x_20; } else { -lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; -x_85 = lean_ctor_get(x_21, 1); +lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; +x_85 = lean_ctor_get(x_20, 1); lean_inc(x_85); -lean_dec(x_21); +lean_dec(x_20); x_86 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1___closed__1; x_87 = lean_name_mk_string(x_1, x_86); x_88 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1___closed__2; @@ -2936,440 +2915,434 @@ x_90 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1__ lean_inc(x_87); x_91 = lean_name_mk_string(x_87, x_90); x_92 = lean_box(2); -lean_inc(x_3); +x_93 = l_Lean_Elab_Deriving_mkInductArgNames___lambda__1___closed__1; lean_inc(x_2); -x_93 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_93, 0, x_92); -lean_ctor_set(x_93, 1, x_2); -lean_ctor_set(x_93, 2, x_3); -x_94 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1___closed__4; -lean_inc(x_93); -x_95 = lean_array_push(x_94, x_93); -lean_inc(x_93); -x_96 = lean_array_push(x_95, x_93); -lean_inc(x_93); -x_97 = lean_array_push(x_96, x_93); -lean_inc(x_93); -x_98 = lean_array_push(x_97, x_93); -lean_inc(x_93); -x_99 = lean_array_push(x_98, x_93); -lean_inc(x_93); -x_100 = lean_array_push(x_99, x_93); -x_101 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_101, 0, x_92); -lean_ctor_set(x_101, 1, x_91); -lean_ctor_set(x_101, 2, x_100); -x_102 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1___closed__5; +x_94 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_94, 0, x_92); +lean_ctor_set(x_94, 1, x_2); +lean_ctor_set(x_94, 2, x_93); +x_95 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1___closed__4; +lean_inc(x_94); +x_96 = lean_array_push(x_95, x_94); +lean_inc(x_94); +x_97 = lean_array_push(x_96, x_94); +lean_inc(x_94); +x_98 = lean_array_push(x_97, x_94); +lean_inc(x_94); +x_99 = lean_array_push(x_98, x_94); +lean_inc(x_94); +x_100 = lean_array_push(x_99, x_94); +lean_inc(x_94); +x_101 = lean_array_push(x_100, x_94); +x_102 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_102, 0, x_92); +lean_ctor_set(x_102, 1, x_91); +lean_ctor_set(x_102, 2, x_101); +x_103 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1___closed__5; lean_inc(x_87); -x_103 = lean_name_mk_string(x_87, x_102); -x_104 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1___closed__6; -lean_inc(x_4); -x_105 = lean_name_mk_string(x_4, x_104); -lean_inc(x_93); -x_106 = lean_array_push(x_5, x_93); -x_107 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_107, 0, x_92); -lean_ctor_set(x_107, 1, x_105); -lean_ctor_set(x_107, 2, x_106); -lean_inc(x_19); -x_108 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_108, 0, x_19); -lean_ctor_set(x_108, 1, x_102); -x_109 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1___closed__7; +x_104 = lean_name_mk_string(x_87, x_103); +x_105 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1___closed__6; +lean_inc(x_3); +x_106 = lean_name_mk_string(x_3, x_105); +lean_inc(x_94); +x_107 = lean_array_push(x_4, x_94); +x_108 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_108, 0, x_92); +lean_ctor_set(x_108, 1, x_106); +lean_ctor_set(x_108, 2, x_107); +lean_inc(x_18); +x_109 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_109, 0, x_18); +lean_ctor_set(x_109, 1, x_103); +x_110 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1___closed__7; lean_inc(x_87); -x_110 = lean_name_mk_string(x_87, x_109); -x_111 = l_Array_append___rarg(x_3, x_6); -x_112 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_112, 0, x_92); -lean_ctor_set(x_112, 1, x_2); -lean_ctor_set(x_112, 2, x_111); -x_113 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__12; -x_114 = lean_name_mk_string(x_4, x_113); -x_115 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__14; -lean_inc(x_19); -x_116 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_116, 0, x_19); -lean_ctor_set(x_116, 1, x_115); -lean_inc(x_7); -x_117 = lean_array_push(x_7, x_116); -x_118 = lean_array_push(x_117, x_8); -x_119 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_119, 0, x_92); -lean_ctor_set(x_119, 1, x_114); -lean_ctor_set(x_119, 2, x_118); -lean_inc(x_7); -x_120 = lean_array_push(x_7, x_112); -x_121 = lean_array_push(x_120, x_119); -x_122 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_122, 0, x_92); -lean_ctor_set(x_122, 1, x_110); -lean_ctor_set(x_122, 2, x_121); -x_123 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1___closed__8; -x_124 = lean_name_mk_string(x_87, x_123); -x_125 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__15; -x_126 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_126, 0, x_19); -lean_ctor_set(x_126, 1, x_125); -x_127 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__5; -x_128 = lean_array_push(x_127, x_126); -x_129 = lean_array_push(x_128, x_10); -lean_inc(x_93); -x_130 = lean_array_push(x_129, x_93); -x_131 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_131, 0, x_92); -lean_ctor_set(x_131, 1, x_124); -lean_ctor_set(x_131, 2, x_130); -x_132 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1___closed__9; -x_133 = lean_array_push(x_132, x_107); +x_111 = lean_name_mk_string(x_87, x_110); +x_112 = l_Array_append___rarg(x_93, x_5); +x_113 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_113, 0, x_92); +lean_ctor_set(x_113, 1, x_2); +lean_ctor_set(x_113, 2, x_112); +x_114 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__12; +x_115 = lean_name_mk_string(x_3, x_114); +x_116 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__14; +lean_inc(x_18); +x_117 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_117, 0, x_18); +lean_ctor_set(x_117, 1, x_116); +lean_inc(x_6); +x_118 = lean_array_push(x_6, x_117); +x_119 = lean_array_push(x_118, x_7); +x_120 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_120, 0, x_92); +lean_ctor_set(x_120, 1, x_115); +lean_ctor_set(x_120, 2, x_119); +lean_inc(x_6); +x_121 = lean_array_push(x_6, x_113); +x_122 = lean_array_push(x_121, x_120); +x_123 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_123, 0, x_92); +lean_ctor_set(x_123, 1, x_111); +lean_ctor_set(x_123, 2, x_122); +x_124 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1___closed__8; +x_125 = lean_name_mk_string(x_87, x_124); +x_126 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__15; +x_127 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_127, 0, x_18); +lean_ctor_set(x_127, 1, x_126); +x_128 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__5; +x_129 = lean_array_push(x_128, x_127); +x_130 = lean_array_push(x_129, x_9); +lean_inc(x_94); +x_131 = lean_array_push(x_130, x_94); +x_132 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_132, 0, x_92); +lean_ctor_set(x_132, 1, x_125); +lean_ctor_set(x_132, 2, x_131); +x_133 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1___closed__9; x_134 = lean_array_push(x_133, x_108); -lean_inc(x_93); -x_135 = lean_array_push(x_134, x_93); -lean_inc(x_93); -x_136 = lean_array_push(x_135, x_93); -x_137 = lean_array_push(x_136, x_122); -x_138 = lean_array_push(x_137, x_131); -lean_inc(x_93); -x_139 = lean_array_push(x_138, x_93); -x_140 = lean_array_push(x_139, x_93); -x_141 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_141, 0, x_92); -lean_ctor_set(x_141, 1, x_103); -lean_ctor_set(x_141, 2, x_140); -x_142 = lean_array_push(x_7, x_101); -x_143 = lean_array_push(x_142, x_141); -x_144 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_144, 0, x_92); -lean_ctor_set(x_144, 1, x_89); -lean_ctor_set(x_144, 2, x_143); -x_145 = lean_array_push(x_9, x_144); -x_146 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_146, 0, x_145); -x_147 = lean_alloc_ctor(0, 2, 0); +x_135 = lean_array_push(x_134, x_109); +lean_inc(x_94); +x_136 = lean_array_push(x_135, x_94); +lean_inc(x_94); +x_137 = lean_array_push(x_136, x_94); +x_138 = lean_array_push(x_137, x_123); +x_139 = lean_array_push(x_138, x_132); +lean_inc(x_94); +x_140 = lean_array_push(x_139, x_94); +x_141 = lean_array_push(x_140, x_94); +x_142 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_142, 0, x_92); +lean_ctor_set(x_142, 1, x_104); +lean_ctor_set(x_142, 2, x_141); +x_143 = lean_array_push(x_6, x_102); +x_144 = lean_array_push(x_143, x_142); +x_145 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_145, 0, x_92); +lean_ctor_set(x_145, 1, x_89); +lean_ctor_set(x_145, 2, x_144); +x_146 = lean_array_push(x_8, x_145); +x_147 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_147, 0, x_146); -lean_ctor_set(x_147, 1, x_85); -return x_147; +x_148 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_148, 0, x_147); +lean_ctor_set(x_148, 1, x_85); +return x_148; } } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t 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, 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_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t 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, 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) { _start: { -uint8_t x_19; -x_19 = lean_nat_dec_le(x_9, x_8); -if (x_19 == 0) +uint8_t x_18; +x_18 = lean_nat_dec_le(x_8, x_7); +if (x_18 == 0) { -lean_object* x_20; uint8_t x_21; -x_20 = lean_unsigned_to_nat(0u); -x_21 = lean_nat_dec_eq(x_7, x_20); -if (x_21 == 0) +lean_object* x_19; uint8_t x_20; +x_19 = lean_unsigned_to_nat(0u); +x_20 = lean_nat_dec_eq(x_6, x_19); +if (x_20 == 0) { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; -x_22 = lean_unsigned_to_nat(1u); -x_23 = lean_nat_sub(x_7, x_22); -lean_dec(x_7); -x_30 = l_Lean_instInhabitedInductiveVal; -x_31 = lean_array_get(x_30, x_6, x_8); +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_21 = lean_unsigned_to_nat(1u); +x_22 = lean_nat_sub(x_6, x_21); +lean_dec(x_6); +x_29 = l_Lean_instInhabitedInductiveVal; +x_30 = lean_array_get(x_29, x_5, x_7); +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); x_32 = lean_ctor_get(x_31, 0); lean_inc(x_32); -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); +lean_dec(x_31); +x_33 = l_Array_contains___at_Lean_findField_x3f___spec__1(x_3, x_32); lean_dec(x_32); -x_34 = l_Array_contains___at_Lean_findField_x3f___spec__1(x_3, x_33); -lean_dec(x_33); -if (x_34 == 0) +if (x_33 == 0) { -lean_object* x_35; -lean_dec(x_31); -x_35 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_35, 0, x_11); -x_24 = x_35; -x_25 = x_18; -goto block_29; +lean_object* x_34; +lean_dec(x_30); +x_34 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_34, 0, x_10); +x_23 = x_34; +x_24 = x_17; +goto block_28; } else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_36 = lean_ctor_get(x_1, 1); -x_37 = l_Lean_instInhabitedName; -x_38 = lean_array_get(x_37, x_36, x_8); -lean_inc(x_17); +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_35 = lean_ctor_get(x_1, 1); +x_36 = l_Lean_instInhabitedName; +x_37 = lean_array_get(x_36, x_35, x_7); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -lean_inc(x_31); -x_39 = l_Lean_Elab_Deriving_mkInductArgNames(x_31, x_12, x_13, x_14, x_15, x_16, x_17, x_18); -if (lean_obj_tag(x_39) == 0) +lean_inc(x_11); +lean_inc(x_30); +x_38 = l_Lean_Elab_Deriving_mkInductArgNames(x_30, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +if (lean_obj_tag(x_38) == 0) { -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_40 = lean_ctor_get(x_39, 0); +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_38, 1); lean_inc(x_40); -x_41 = lean_ctor_get(x_39, 1); -lean_inc(x_41); -lean_dec(x_39); -lean_inc(x_17); +lean_dec(x_38); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -lean_inc(x_40); -x_42 = l_Lean_Elab_Deriving_mkImplicitBinders(x_40, x_12, x_13, x_14, x_15, x_16, x_17, x_41); -x_43 = lean_ctor_get(x_42, 0); +lean_inc(x_11); +lean_inc(x_39); +x_41 = l_Lean_Elab_Deriving_mkImplicitBinders(x_39, x_11, x_12, x_13, x_14, x_15, x_16, x_40); +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_41, 1); lean_inc(x_43); -x_44 = lean_ctor_get(x_42, 1); -lean_inc(x_44); -lean_dec(x_42); -lean_inc(x_17); +lean_dec(x_41); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -lean_inc(x_40); -lean_inc(x_31); +lean_inc(x_11); +lean_inc(x_39); +lean_inc(x_30); lean_inc(x_2); -x_45 = l_Lean_Elab_Deriving_mkInstImplicitBinders(x_2, x_31, x_40, x_12, x_13, x_14, x_15, x_16, x_17, x_44); -if (lean_obj_tag(x_45) == 0) +x_44 = l_Lean_Elab_Deriving_mkInstImplicitBinders(x_2, x_30, x_39, x_11, x_12, x_13, x_14, x_15, x_16, x_43); +if (lean_obj_tag(x_44) == 0) { -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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_46 = lean_ctor_get(x_45, 0); +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_45 = lean_ctor_get(x_44, 0); +lean_inc(x_45); +x_46 = lean_ctor_get(x_44, 1); lean_inc(x_46); -x_47 = lean_ctor_get(x_45, 1); -lean_inc(x_47); -lean_dec(x_45); -x_48 = l_Array_append___rarg(x_43, x_46); -lean_inc(x_16); -x_49 = l_Lean_Elab_Deriving_mkInductiveApp(x_31, x_40, x_12, x_13, x_14, x_15, x_16, x_17, x_47); -x_50 = lean_ctor_get(x_49, 0); +lean_dec(x_44); +x_47 = l_Array_append___rarg(x_42, x_45); +lean_inc(x_15); +x_48 = l_Lean_Elab_Deriving_mkInductiveApp(x_30, x_39, x_11, x_12, x_13, x_14, x_15, x_16, x_46); +x_49 = lean_ctor_get(x_48, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_48, 1); lean_inc(x_50); -x_51 = lean_ctor_get(x_49, 1); -lean_inc(x_51); -lean_dec(x_49); -lean_inc(x_16); -x_52 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_16, x_17, x_51); -x_53 = lean_ctor_get(x_52, 1); -lean_inc(x_53); -lean_dec(x_52); -x_54 = lean_st_ref_get(x_17, x_53); -x_55 = lean_ctor_get(x_54, 1); -lean_inc(x_55); -lean_dec(x_54); +lean_dec(x_48); +lean_inc(x_15); +x_51 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_15, x_16, x_50); +x_52 = lean_ctor_get(x_51, 1); +lean_inc(x_52); +lean_dec(x_51); +x_53 = lean_st_ref_get(x_16, x_52); +x_54 = lean_ctor_get(x_53, 1); +lean_inc(x_54); +lean_dec(x_53); lean_inc(x_2); -x_56 = lean_mk_syntax_ident(x_2); -x_57 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkImplicitBinders___spec__1___closed__4; -x_58 = lean_array_push(x_57, x_50); -x_59 = lean_box(2); -x_60 = l_Lean_Elab_Deriving_mkInductiveApp___closed__14; -x_61 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_61, 0, x_59); -lean_ctor_set(x_61, 1, x_60); -lean_ctor_set(x_61, 2, x_58); -x_62 = l_Lean_Elab_Deriving_mkInductiveApp___closed__12; -x_63 = lean_array_push(x_62, x_56); -x_64 = lean_array_push(x_63, x_61); -x_65 = l_Lean_Elab_Deriving_mkInductiveApp___closed__8; -x_66 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_66, 0, x_59); -lean_ctor_set(x_66, 1, x_65); -lean_ctor_set(x_66, 2, x_64); +x_55 = lean_mk_syntax_ident(x_2); +x_56 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkImplicitBinders___spec__1___closed__4; +x_57 = lean_array_push(x_56, x_49); +x_58 = lean_box(2); +x_59 = l_Lean_Elab_Deriving_mkInductiveApp___closed__14; +x_60 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +lean_ctor_set(x_60, 2, x_57); +x_61 = l_Lean_Elab_Deriving_mkInductiveApp___closed__12; +x_62 = lean_array_push(x_61, x_55); +x_63 = lean_array_push(x_62, x_60); +x_64 = l_Lean_Elab_Deriving_mkInductiveApp___closed__8; +x_65 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_65, 0, x_58); +lean_ctor_set(x_65, 1, x_64); +lean_ctor_set(x_65, 2, x_63); if (x_4 == 0) { -lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_67 = lean_mk_syntax_ident(x_38); -x_68 = l_Lean_Elab_Deriving_mkInductiveApp___closed__4; -x_69 = l_Lean_Elab_Deriving_mkInductiveApp___closed__6; -lean_inc(x_16); -lean_inc(x_5); -x_70 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1(x_68, x_60, x_5, x_69, x_57, x_48, x_62, x_66, x_11, x_67, x_12, x_13, x_14, x_15, x_16, x_17, x_55); -x_71 = lean_ctor_get(x_70, 0); +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_66 = lean_mk_syntax_ident(x_37); +x_67 = l_Lean_Elab_Deriving_mkInductiveApp___closed__4; +x_68 = l_Lean_Elab_Deriving_mkInductiveApp___closed__6; +lean_inc(x_15); +x_69 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1(x_67, x_59, x_68, x_56, x_47, x_61, x_65, x_10, x_66, x_11, x_12, x_13, x_14, x_15, x_16, x_54); +x_70 = lean_ctor_get(x_69, 0); +lean_inc(x_70); +x_71 = lean_ctor_get(x_69, 1); lean_inc(x_71); -x_72 = lean_ctor_get(x_70, 1); -lean_inc(x_72); -lean_dec(x_70); +lean_dec(x_69); +x_23 = x_70; x_24 = x_71; -x_25 = x_72; -goto block_29; +goto block_28; } else { -lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; -lean_inc(x_16); -x_73 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_16, x_17, x_55); -x_74 = lean_ctor_get(x_73, 0); -lean_inc(x_74); -x_75 = lean_ctor_get(x_73, 1); -lean_inc(x_75); -lean_dec(x_73); -x_76 = lean_st_ref_get(x_17, x_75); -x_77 = lean_ctor_get(x_76, 1); -lean_inc(x_77); -lean_dec(x_76); -x_78 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__3; +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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +lean_inc(x_15); +x_72 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_15, x_16, x_54); +x_73 = lean_ctor_get(x_72, 0); +lean_inc(x_73); +x_74 = lean_ctor_get(x_72, 1); lean_inc(x_74); -x_79 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_79, 0, x_74); -lean_ctor_set(x_79, 1, x_78); -x_80 = lean_mk_syntax_ident(x_38); -x_81 = lean_array_push(x_57, x_80); -x_82 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_82, 0, x_59); -lean_ctor_set(x_82, 1, x_60); -lean_ctor_set(x_82, 2, x_81); -x_83 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__4; -x_84 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_84, 0, x_74); -lean_ctor_set(x_84, 1, x_83); -x_85 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__5; -x_86 = lean_array_push(x_85, x_79); -x_87 = lean_array_push(x_86, x_82); -x_88 = lean_array_push(x_87, x_84); -x_89 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__2; -x_90 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_90, 0, x_59); -lean_ctor_set(x_90, 1, x_89); -lean_ctor_set(x_90, 2, x_88); -x_91 = l_Lean_Elab_Deriving_mkInductiveApp___closed__4; -x_92 = l_Lean_Elab_Deriving_mkInductiveApp___closed__6; -lean_inc(x_16); -lean_inc(x_5); -x_93 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1(x_91, x_60, x_5, x_92, x_57, x_48, x_62, x_66, x_11, x_90, x_12, x_13, x_14, x_15, x_16, x_17, x_77); -x_94 = lean_ctor_get(x_93, 0); +lean_dec(x_72); +x_75 = lean_st_ref_get(x_16, x_74); +x_76 = lean_ctor_get(x_75, 1); +lean_inc(x_76); +lean_dec(x_75); +x_77 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__3; +lean_inc(x_73); +x_78 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_78, 0, x_73); +lean_ctor_set(x_78, 1, x_77); +x_79 = lean_mk_syntax_ident(x_37); +x_80 = lean_array_push(x_56, x_79); +x_81 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_81, 0, x_58); +lean_ctor_set(x_81, 1, x_59); +lean_ctor_set(x_81, 2, x_80); +x_82 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__4; +x_83 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_83, 0, x_73); +lean_ctor_set(x_83, 1, x_82); +x_84 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__5; +x_85 = lean_array_push(x_84, x_78); +x_86 = lean_array_push(x_85, x_81); +x_87 = lean_array_push(x_86, x_83); +x_88 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__2; +x_89 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_89, 0, x_58); +lean_ctor_set(x_89, 1, x_88); +lean_ctor_set(x_89, 2, x_87); +x_90 = l_Lean_Elab_Deriving_mkInductiveApp___closed__4; +x_91 = l_Lean_Elab_Deriving_mkInductiveApp___closed__6; +lean_inc(x_15); +x_92 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1(x_90, x_59, x_91, x_56, x_47, x_61, x_65, x_10, x_89, x_11, x_12, x_13, x_14, x_15, x_16, x_76); +x_93 = lean_ctor_get(x_92, 0); +lean_inc(x_93); +x_94 = lean_ctor_get(x_92, 1); lean_inc(x_94); -x_95 = lean_ctor_get(x_93, 1); -lean_inc(x_95); -lean_dec(x_93); +lean_dec(x_92); +x_23 = x_93; x_24 = x_94; -x_25 = x_95; -goto block_29; +goto block_28; } } else { -uint8_t x_96; -lean_dec(x_43); -lean_dec(x_40); -lean_dec(x_38); -lean_dec(x_31); -lean_dec(x_23); -lean_dec(x_17); +uint8_t x_95; +lean_dec(x_42); +lean_dec(x_39); +lean_dec(x_37); +lean_dec(x_30); +lean_dec(x_22); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_8); -lean_dec(x_5); +lean_dec(x_10); +lean_dec(x_7); lean_dec(x_2); -x_96 = !lean_is_exclusive(x_45); -if (x_96 == 0) +x_95 = !lean_is_exclusive(x_44); +if (x_95 == 0) { -return x_45; +return x_44; } else { -lean_object* x_97; lean_object* x_98; lean_object* x_99; -x_97 = lean_ctor_get(x_45, 0); -x_98 = lean_ctor_get(x_45, 1); -lean_inc(x_98); +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_44, 0); +x_97 = lean_ctor_get(x_44, 1); lean_inc(x_97); -lean_dec(x_45); -x_99 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_99, 0, x_97); -lean_ctor_set(x_99, 1, x_98); -return x_99; +lean_inc(x_96); +lean_dec(x_44); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; } } } else { -uint8_t x_100; -lean_dec(x_38); -lean_dec(x_31); -lean_dec(x_23); -lean_dec(x_17); +uint8_t x_99; +lean_dec(x_37); +lean_dec(x_30); +lean_dec(x_22); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_8); -lean_dec(x_5); +lean_dec(x_10); +lean_dec(x_7); lean_dec(x_2); -x_100 = !lean_is_exclusive(x_39); -if (x_100 == 0) +x_99 = !lean_is_exclusive(x_38); +if (x_99 == 0) { -return x_39; +return x_38; } else { -lean_object* x_101; lean_object* x_102; lean_object* x_103; -x_101 = lean_ctor_get(x_39, 0); -x_102 = lean_ctor_get(x_39, 1); -lean_inc(x_102); +lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_100 = lean_ctor_get(x_38, 0); +x_101 = lean_ctor_get(x_38, 1); lean_inc(x_101); -lean_dec(x_39); -x_103 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_103, 0, x_101); -lean_ctor_set(x_103, 1, x_102); -return x_103; +lean_inc(x_100); +lean_dec(x_38); +x_102 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_102, 0, x_100); +lean_ctor_set(x_102, 1, x_101); +return x_102; } } } -block_29: +block_28: { -lean_object* x_26; lean_object* x_27; -x_26 = lean_ctor_get(x_24, 0); -lean_inc(x_26); -lean_dec(x_24); -x_27 = lean_nat_add(x_8, x_10); -lean_dec(x_8); -x_7 = x_23; -x_8 = x_27; -x_11 = x_26; -x_18 = x_25; +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_nat_add(x_7, x_9); +lean_dec(x_7); +x_6 = x_22; +x_7 = x_26; +x_10 = x_25; +x_17 = x_24; goto _start; } } else { -lean_object* x_104; -lean_dec(x_17); +lean_object* x_103; lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); -lean_dec(x_8); +lean_dec(x_11); lean_dec(x_7); -lean_dec(x_5); +lean_dec(x_6); lean_dec(x_2); -x_104 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_104, 0, x_11); -lean_ctor_set(x_104, 1, x_18); -return x_104; +x_103 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_103, 0, x_10); +lean_ctor_set(x_103, 1, x_17); +return x_103; } } else { -lean_object* x_105; -lean_dec(x_17); +lean_object* x_104; lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); -lean_dec(x_8); +lean_dec(x_11); lean_dec(x_7); -lean_dec(x_5); +lean_dec(x_6); lean_dec(x_2); -x_105 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_105, 0, x_11); -lean_ctor_set(x_105, 1, x_18); -return x_105; +x_104 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_104, 0, x_10); +lean_ctor_set(x_104, 1, x_17); +return x_104; } } } @@ -3379,11 +3352,11 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_mkInstanceCmds(lean_object* x_1, l lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_12 = lean_ctor_get(x_1, 0); x_13 = lean_array_get_size(x_12); -x_14 = l_Lean_Elab_Deriving_mkInductArgNames___lambda__1___closed__1; -x_15 = lean_unsigned_to_nat(0u); -x_16 = lean_unsigned_to_nat(1u); +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_unsigned_to_nat(1u); +x_16 = l_Lean_Elab_Deriving_mkInductArgNames___lambda__1___closed__1; lean_inc(x_13); -x_17 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1(x_1, x_2, x_3, x_4, x_14, x_12, x_13, x_15, x_13, x_16, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_17 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1(x_1, x_2, x_3, x_4, x_12, x_13, x_14, x_13, x_15, x_16, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_13); if (lean_obj_tag(x_17) == 0) { @@ -3431,34 +3404,17 @@ return x_25; } } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___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, 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, lean_object* x_16) { _start: { -lean_object* x_18; -x_18 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -lean_dec(x_16); -lean_dec(x_14); +lean_object* x_17; +x_17 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_15); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -return x_18; +lean_dec(x_10); +return x_17; } } LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___boxed(lean_object** _args) { @@ -3479,19 +3435,18 @@ lean_object* x_14 = _args[13]; lean_object* x_15 = _args[14]; lean_object* x_16 = _args[15]; lean_object* x_17 = _args[16]; -lean_object* x_18 = _args[17]; _start: { -uint8_t x_19; lean_object* x_20; -x_19 = lean_unbox(x_4); +uint8_t x_18; lean_object* x_19; +x_18 = lean_unbox(x_4); lean_dec(x_4); -x_20 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1(x_1, x_2, x_3, x_19, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); -lean_dec(x_10); +x_19 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1(x_1, x_2, x_3, x_18, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); lean_dec(x_9); -lean_dec(x_6); +lean_dec(x_8); +lean_dec(x_5); lean_dec(x_3); lean_dec(x_1); -return x_20; +return x_19; } } LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_mkInstanceCmds___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_object* x_9, lean_object* x_10, lean_object* x_11) { @@ -3680,22 +3635,12 @@ return x_1; static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkHeader___spec__2___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_mkInductiveApp___closed__6; -x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkHeader___spec__2___closed__1; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkHeader___spec__2___closed__3() { -_start: -{ lean_object* x_1; x_1 = lean_mk_string_from_bytes("(", 1); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkHeader___spec__2___closed__4() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkHeader___spec__2___closed__3() { _start: { lean_object* x_1; @@ -3703,87 +3648,90 @@ x_1 = lean_mk_string_from_bytes(")", 1); return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkHeader___spec__2(lean_object* x_1, size_t x_2, size_t 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, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkHeader___spec__2(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, 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) { _start: { -uint8_t x_12; -x_12 = lean_usize_dec_lt(x_3, x_2); -if (x_12 == 0) +uint8_t x_15; +x_15 = lean_usize_dec_lt(x_6, x_5); +if (x_15 == 0) { -lean_object* x_13; -lean_dec(x_9); +lean_object* x_16; +lean_dec(x_12); +lean_dec(x_3); lean_dec(x_1); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_4); -lean_ctor_set(x_13, 1, x_11); -return x_13; +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_7); +lean_ctor_set(x_16, 1, x_14); +return x_16; } 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; 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; size_t x_47; size_t x_48; lean_object* x_49; -x_14 = lean_array_uget(x_4, x_3); -x_15 = lean_unsigned_to_nat(0u); -x_16 = lean_array_uset(x_4, x_3, x_15); -lean_inc(x_9); -x_17 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_9, x_10, x_11); -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); -lean_inc(x_19); -lean_dec(x_17); -x_20 = lean_st_ref_get(x_10, x_19); -x_21 = lean_ctor_get(x_20, 1); +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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; size_t x_51; size_t x_52; lean_object* x_53; +x_17 = lean_array_uget(x_7, x_6); +x_18 = lean_unsigned_to_nat(0u); +x_19 = lean_array_uset(x_7, x_6, x_18); +lean_inc(x_12); +x_20 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_12, x_13, x_14); +x_21 = lean_ctor_get(x_20, 0); lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); lean_dec(x_20); -x_22 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkHeader___spec__2___closed__3; -lean_inc(x_18); -x_23 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_23, 0, x_18); -lean_ctor_set(x_23, 1, x_22); -x_24 = lean_mk_syntax_ident(x_14); -x_25 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkImplicitBinders___spec__1___closed__4; -x_26 = lean_array_push(x_25, x_24); -x_27 = lean_box(2); -x_28 = l_Lean_Elab_Deriving_mkInductiveApp___closed__14; -x_29 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -lean_ctor_set(x_29, 2, x_26); -x_30 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__14; -lean_inc(x_18); -x_31 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_31, 0, x_18); -lean_ctor_set(x_31, 1, x_30); -x_32 = l_Lean_Elab_Deriving_mkInductiveApp___closed__12; -x_33 = lean_array_push(x_32, x_31); +x_23 = lean_st_ref_get(x_13, x_22); +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); +x_25 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkHeader___spec__2___closed__1; +lean_inc(x_3); +x_26 = lean_name_mk_string(x_3, x_25); +x_27 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkHeader___spec__2___closed__2; +lean_inc(x_21); +x_28 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_28, 0, x_21); +lean_ctor_set(x_28, 1, x_27); +x_29 = lean_mk_syntax_ident(x_17); +x_30 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkImplicitBinders___spec__1___closed__4; +x_31 = lean_array_push(x_30, x_29); +x_32 = lean_box(2); +x_33 = l_Lean_Elab_Deriving_mkInductiveApp___closed__14; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +x_35 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__14; +lean_inc(x_21); +x_36 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_36, 0, x_21); +lean_ctor_set(x_36, 1, x_35); +x_37 = l_Lean_Elab_Deriving_mkInductiveApp___closed__12; +x_38 = lean_array_push(x_37, x_36); lean_inc(x_1); -x_34 = lean_array_push(x_33, x_1); -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_27); -lean_ctor_set(x_35, 1, x_28); -lean_ctor_set(x_35, 2, x_34); -x_36 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkHeader___spec__2___closed__4; -x_37 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_37, 0, x_18); -lean_ctor_set(x_37, 1, x_36); -x_38 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__16; -x_39 = lean_array_push(x_38, x_23); -x_40 = lean_array_push(x_39, x_29); -x_41 = lean_array_push(x_40, x_35); -x_42 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkImplicitBinders___spec__1___closed__5; -x_43 = lean_array_push(x_41, x_42); -x_44 = lean_array_push(x_43, x_37); -x_45 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkHeader___spec__2___closed__2; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_27); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_44); -x_47 = 1; -x_48 = lean_usize_add(x_3, x_47); -x_49 = lean_array_uset(x_16, x_3, x_46); -x_3 = x_48; -x_4 = x_49; -x_11 = x_21; +x_39 = lean_array_push(x_38, x_1); +x_40 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_40, 0, x_32); +lean_ctor_set(x_40, 1, x_33); +lean_ctor_set(x_40, 2, x_39); +x_41 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkHeader___spec__2___closed__3; +x_42 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_42, 0, x_21); +lean_ctor_set(x_42, 1, x_41); +x_43 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__16; +x_44 = lean_array_push(x_43, x_28); +x_45 = lean_array_push(x_44, x_34); +x_46 = lean_array_push(x_45, x_40); +x_47 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkImplicitBinders___spec__1___closed__5; +x_48 = lean_array_push(x_46, x_47); +x_49 = lean_array_push(x_48, x_42); +x_50 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_50, 0, x_32); +lean_ctor_set(x_50, 1, x_26); +lean_ctor_set(x_50, 2, x_49); +x_51 = 1; +x_52 = lean_usize_add(x_6, x_51); +x_53 = lean_array_uset(x_19, x_6, x_50); +x_6 = x_52; +x_7 = x_53; +x_14 = x_24; goto _start; } } @@ -3851,62 +3799,65 @@ lean_inc(x_12); x_26 = l_Lean_Elab_Deriving_mkInstImplicitBinders(x_1, x_3, x_12, x_4, x_5, x_6, x_7, x_8, x_9, x_25); if (lean_obj_tag(x_26) == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; size_t x_31; size_t x_32; lean_object* x_33; uint8_t x_34; +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; size_t x_32; size_t x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; x_27 = lean_ctor_get(x_26, 0); lean_inc(x_27); x_28 = lean_ctor_get(x_26, 1); lean_inc(x_28); lean_dec(x_26); -x_29 = l_Array_append___rarg(x_15, x_27); -x_30 = lean_array_get_size(x_24); -x_31 = lean_usize_of_nat(x_30); -lean_dec(x_30); -x_32 = 0; +x_29 = lean_box(0); +x_30 = l_Array_append___rarg(x_15, x_27); +x_31 = lean_array_get_size(x_24); +x_32 = lean_usize_of_nat(x_31); +lean_dec(x_31); +x_33 = 0; +x_34 = l_Lean_Elab_Deriving_mkInductiveApp___closed__2; +x_35 = l_Lean_Elab_Deriving_mkInductiveApp___closed__6; lean_inc(x_24); lean_inc(x_18); -x_33 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkHeader___spec__2(x_18, x_31, x_32, x_24, x_4, x_5, x_6, x_7, x_8, x_9, x_28); +x_36 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkHeader___spec__2(x_18, x_34, x_35, x_29, x_32, x_33, x_24, x_4, x_5, x_6, x_7, x_8, x_9, x_28); lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_34 = !lean_is_exclusive(x_33); -if (x_34 == 0) -{ -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_33, 0); -x_36 = l_Array_append___rarg(x_29, x_35); -x_37 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_12); -lean_ctor_set(x_37, 2, x_24); -lean_ctor_set(x_37, 3, x_18); -lean_ctor_set(x_33, 0, x_37); -return x_33; +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_36, 0); +x_39 = l_Array_append___rarg(x_30, x_38); +x_40 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_12); +lean_ctor_set(x_40, 2, x_24); +lean_ctor_set(x_40, 3, x_18); +lean_ctor_set(x_36, 0, x_40); +return x_36; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_38 = lean_ctor_get(x_33, 0); -x_39 = lean_ctor_get(x_33, 1); -lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_33); -x_40 = l_Array_append___rarg(x_29, x_38); -x_41 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set(x_41, 1, x_12); -lean_ctor_set(x_41, 2, x_24); -lean_ctor_set(x_41, 3, x_18); -x_42 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_39); -return x_42; +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_41 = lean_ctor_get(x_36, 0); +x_42 = lean_ctor_get(x_36, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_36); +x_43 = l_Array_append___rarg(x_30, x_41); +x_44 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_12); +lean_ctor_set(x_44, 2, x_24); +lean_ctor_set(x_44, 3, x_18); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_42); +return x_45; } } else { -uint8_t x_43; +uint8_t x_46; lean_dec(x_24); lean_dec(x_18); lean_dec(x_15); @@ -3917,29 +3868,29 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_43 = !lean_is_exclusive(x_26); -if (x_43 == 0) +x_46 = !lean_is_exclusive(x_26); +if (x_46 == 0) { return x_26; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_26, 0); -x_45 = lean_ctor_get(x_26, 1); -lean_inc(x_45); -lean_inc(x_44); +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_26, 0); +x_48 = lean_ctor_get(x_26, 1); +lean_inc(x_48); +lean_inc(x_47); lean_dec(x_26); -x_46 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -return x_46; +x_49 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set(x_49, 1, x_48); +return x_49; } } } else { -uint8_t x_47; +uint8_t x_50; lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -3949,23 +3900,23 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_47 = !lean_is_exclusive(x_11); -if (x_47 == 0) +x_50 = !lean_is_exclusive(x_11); +if (x_50 == 0) { return x_11; } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_48 = lean_ctor_get(x_11, 0); -x_49 = lean_ctor_get(x_11, 1); -lean_inc(x_49); -lean_inc(x_48); +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_11, 0); +x_52 = lean_ctor_get(x_11, 1); +lean_inc(x_52); +lean_inc(x_51); lean_dec(x_11); -x_50 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_50, 0, x_48); -lean_ctor_set(x_50, 1, x_49); -return x_50; +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_52); +return x_53; } } } @@ -3986,21 +3937,23 @@ lean_dec(x_3); return x_13; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkHeader___spec__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, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkHeader___spec__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, lean_object* x_7, 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) { _start: { -size_t x_12; size_t x_13; lean_object* x_14; -x_12 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_13 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_14 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkHeader___spec__2(x_1, x_12, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +size_t x_15; size_t x_16; lean_object* x_17; +x_15 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_16 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_17 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkHeader___spec__2(x_1, x_2, x_3, x_4, x_15, x_16, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_13); +lean_dec(x_11); lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -return x_14; +lean_dec(x_4); +lean_dec(x_2); +return x_17; } } LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at_Lean_Elab_Deriving_mkDiscrs___spec__1(lean_object* x_1, size_t x_2, size_t 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, lean_object* x_11) { @@ -4366,8 +4319,6 @@ l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkHeader___spec__2___closed__2 = lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkHeader___spec__2___closed__2); l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkHeader___spec__2___closed__3 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkHeader___spec__2___closed__3(); lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkHeader___spec__2___closed__3); -l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkHeader___spec__2___closed__4 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkHeader___spec__2___closed__4(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkHeader___spec__2___closed__4); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Elab/Do.c b/stage0/stdlib/Lean/Elab/Do.c index 7dfd396d8db..0e15bad9bfd 100644 --- a/stage0/stdlib/Lean/Elab/Do.c +++ b/stage0/stdlib/Lean/Elab/Do.c @@ -92,11 +92,9 @@ static lean_object* l_Lean_Elab_Term_Do_ToTerm_declToTerm___closed__5; static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___closed__8; static lean_object* l___regBuiltin_Lean_Elab_Term_expandTermReturn_declRange___closed__2; static lean_object* l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__14; -static lean_object* l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__28; static lean_object* l_Lean_Elab_Term_Do_ToTerm_breakToTerm___closed__17; static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___spec__10___closed__12; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_ToTerm_mkJmp(lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_nullKind; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLiftMethod___closed__5; static lean_object* l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__15; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_convertTerminalActionIntoJmp_loop___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -244,6 +242,7 @@ static lean_object* l_Lean_Elab_Term_Do_getDoHaveVars___closed__3; LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Elab_Term_Do_hasBreakContinue___lambda__1(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__5___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_isMutableLet___boxed(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandTermReturn___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_extractBind(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -261,6 +260,7 @@ static lean_object* l_Lean_Elab_Term_Do_ToTerm_returnToTerm___closed__16; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_Do_mkJmp___spec__6___boxed(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_Term_Do_mkSimpleJmp(lean_object*, 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_throwErrorAt___at_Lean_Elab_Term_Do_ToCodeBlock_checkLetArrowRHS___spec__1___boxed(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_throwErrorAt___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkUnit___closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_Do_ToCodeBlock_checkNotShadowingMutable___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_varsToMessageData___closed__2; @@ -310,6 +310,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_extractBind_ static lean_object* l_Lean_Elab_Term_Do_ToTerm_returnToTerm___closed__36; static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__10; static lean_object* l_Lean_Elab_Term_Do_ToTerm_matchNestedTermResult___closed__19; +static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__32; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doReassignArrowToCode___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Do_ToTerm_seqToTerm___closed__8; lean_object* l_Lean_MessageData_joinSep(lean_object*, lean_object*); @@ -355,7 +356,9 @@ static lean_object* l_Lean_Elab_Term_elabLiftMethod___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_ToTerm_run___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_hasLiftMethod___boxed(lean_object*); static lean_object* l_Lean_Elab_Term_Do_pullExitPointsAux___lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__6(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___spec__10___closed__7; LEAN_EXPORT uint8_t l_Lean_Elab_Term_Do_isMutableLet(lean_object*); static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkUnit___closed__8; @@ -366,7 +369,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_mkFreshJP(lean_object*, lean_object static lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkTuple___spec__1___closed__5; LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Do_ToTerm_mkJoinPoint___closed__1; -lean_object* l_Lean_Syntax_SepArray_ofElems(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_RBNode_ins___at_Lean_Elab_Term_Do_insertVars___spec__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__3; static lean_object* l_Lean_Elab_Term_Do_ToTerm_returnToTerm___closed__42; @@ -395,7 +397,6 @@ static lean_object* l_Lean_Elab_Term_Do_ToTerm_returnToTerm___closed__3; static lean_object* l_Lean_Elab_Term_Do_ToTerm_declToTerm___closed__1; lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Do_ToTerm_matchNestedTermResult___closed__20; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_ToTerm_mkJoinPoint___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_addFreshJP(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___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkUnit(lean_object*, lean_object*); @@ -440,8 +441,9 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_getLetDeclVars(lean_object*, lean_o LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_expandTermReturn(lean_object*); static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_checkLetArrowRHS___closed__5; static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_liftMethodDelimiter___closed__7; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__8___boxed(lean_object**); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_ToCodeBlock_expandLiftMethodAux___spec__6(uint8_t, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__27; lean_object* lean_array_fget(lean_object*, lean_object*); @@ -479,6 +481,7 @@ static lean_object* l_Lean_Elab_Term_Do_ToTerm_returnToTerm___closed__14; static lean_object* l_Lean_Elab_Term_Do_ToTerm_returnToTerm___closed__28; static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_ToCodeBlock_expandLiftMethodAux___closed__2; static lean_object* l_Lean_Elab_Term_Do_ToTerm_breakToTerm___closed__10; +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandTermFor_declRange___closed__6; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_Do_extendUpdatedVarsAux_update___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_getDoSeqElems___closed__10; @@ -495,6 +498,7 @@ static lean_object* l_Lean_Elab_Term_Do_ToTerm_returnToTerm___closed__33; static lean_object* l_Lean_Elab_Term_Do_ToTerm_returnToTerm___closed__22; LEAN_EXPORT lean_object* l_Lean_mkIdentFromRef___at_Lean_Elab_Term_Do_ToTerm_mkJoinPoint___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_CodeBlocl_toMessageData_loop___spec__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__7___closed__4; static lean_object* l_Lean_Elab_Term_Do_CodeBlocl_toMessageData_loop___closed__14; @@ -516,6 +520,7 @@ static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_ToCodeBlock_do LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___closed__1; static lean_object* l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__4___closed__6; +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Do_ToCodeBlock_getTryCatchUpdatedVars___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Elab_Do_0__Lean_Elab_Term_liftMethodForbiddenBinder(lean_object*); static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkPureUnit___closed__7; @@ -523,7 +528,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__1(lea static lean_object* l___regBuiltin_Lean_Elab_Term_elabLiftMethod___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_Do_mkJmp___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Do_ToTerm_matchNestedTermResult___closed__24; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__7___boxed(lean_object**); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__5(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_List_mapTRAux___at___private_Lean_Elab_Do_0__Lean_Elab_Term_getDoSeqElems___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__4___closed__2; @@ -537,7 +542,7 @@ static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkPureUnit___c LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_withMacroExpansion___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___closed__5; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__6___boxed(lean_object**); static lean_object* l_Lean_Elab_Term_Do_mkSimpleJmp___closed__2; LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___spec__2___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_Do_ToCodeBlock_tryCatchPred___spec__1(lean_object*, lean_object*, size_t, size_t); @@ -676,7 +681,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_ LEAN_EXPORT lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_toDoElem___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_getDoHaveVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_Do_pullExitPointsAux___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_expandOptType(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_extractBind___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_Do_ToCodeBlock_checkNotShadowingMutable___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -706,15 +710,16 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_extendUpdatedVars(lean_object*, lea LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_mkDoSeq___spec__1(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_instInhabitedCode; static lean_object* l_Lean_Elab_Term_Do_ToTerm_breakToTerm___closed__16; -lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_expandMatchAlts_x3f___spec__2(lean_object*, size_t, size_t, lean_object*); static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkPureUnit___closed__3; static lean_object* l_Lean_Elab_Term_Do_ToTerm_mkNestedKind___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Do_ToTerm_seqToTerm___closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_checkLetArrowRHS(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_Term_Do_getHaveIdLhsVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_getLetEqnsDeclVar___boxed(lean_object*); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_reverse___rarg(lean_object*); static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__12; static lean_object* l_Lean_Elab_Term_Do_getDoLetArrowVars___closed__4; @@ -734,9 +739,11 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_mkReassignCore___lambda__1(lean_obj LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_Do_ToCodeBlock_doSeqToCode___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Do_ToTerm_returnToTerm___closed__23; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_Do_concat___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_mkSepArray(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_ToCodeBlock_expandLiftMethodAux___spec__3(uint8_t, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__6___lambda__1___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_throwMaxRecDepthAt___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_Do_ToCodeBlock_checkReassignable___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_checkLetArrowRHS___closed__1; lean_object* l_Std_RBNode_balLeft___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -845,7 +852,6 @@ lean_object* l_Lean_mkAtomFrom(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Do_ToTerm_breakToTerm___closed__14; LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_Do_ToTerm_returnToTerm___spec__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_ToTerm_mkJoinPoint___spec__2(size_t, size_t, lean_object*); -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addTrace___at_Lean_Elab_Term_Do_mkJmp___spec__3___closed__5; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Do_CodeBlocl_toMessageData_loop___spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOptionM___at_Lean_Elab_Term_Do_ToCodeBlock_mkForInBody___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -859,17 +865,19 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_toDoElem(lea LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_ToCodeBlock_expandLiftMethodAux___spec__5(size_t, size_t, 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_Term_Do_pullExitPointsAux___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Do_ToTerm_breakToTerm___closed__15; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_29313_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_30705_(lean_object*); static lean_object* l_Lean_Elab_Term_Do_pullExitPointsAux___lambda__1___closed__9; static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doSeqToCode___closed__3; static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doReassignArrowToCode___closed__3; lean_object* l_Lean_Elab_Term_getPatternVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Do_ToTerm_matchNestedTermResult___closed__18; lean_object* l_Lean_Meta_assignExprMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabLiftMethod(lean_object*); static lean_object* l_Lean_Elab_Term_Do_ToTerm_matchNestedTermResult___closed__5; static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_letDeclArgHasBinders___closed__3; static lean_object* l_Lean_Elab_Term_Do_ToTerm_seqToTerm___closed__12; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__7(lean_object*, size_t, size_t, 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_Term_Do_concat(lean_object*, 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_liftMacroM___at_Lean_Elab_Term_Do_ToCodeBlock_doSeqToCode___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__6___lambda__1___closed__14; @@ -945,7 +953,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__1 uint8_t lean_nat_dec_le(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_convertTerminalActionIntoJmp_loop___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_sequenceMap___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__1___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_ensureEOS(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_Do_mkReassignCore___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_RBNode_ins___at_Lean_Elab_Term_Do_insertVars___spec__2___boxed(lean_object*, lean_object*, lean_object*); @@ -953,7 +960,6 @@ LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_Do static lean_object* l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_getDoReassignVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___lambda__4___closed__2; -static lean_object* l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__27; uint8_t l_Std_RBNode_isRed___rarg(lean_object*); static lean_object* l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__7___closed__1; static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___lambda__2___closed__2; @@ -965,7 +971,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandTermReturn(lean_object*, lean_ob LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_ToTerm_matchNestedTermResult___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Term_Do_mkJmp___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArgs(lean_object*); lean_object* l_Lean_Name_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_elabLiftMethod___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1015,10 +1021,9 @@ static lean_object* l_Lean_Elab_Term_Do_ToTerm_reassignToTerm___closed__6; static lean_object* l_Lean_Elab_Term_Do_ToTerm_continueToTerm___closed__23; lean_object* l_ReaderT_pure___at_Lean_Elab_liftMacroM___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static size_t l_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_varSetToArray(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_getLetEqnsDeclVar(lean_object*); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_ToTerm_mkJoinPoint___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__1___closed__1; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_ToTerm_mkJoinPoint___spec__1___closed__1; @@ -1151,6 +1156,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_extractBind_ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__6___lambda__1___closed__15; static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_extractBind___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkTuple___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Elab_Term_expandMatchAlt(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___spec__1___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doSeqToCode___closed__16; lean_object* l_Lean_indentD(lean_object*); @@ -1175,6 +1181,7 @@ static lean_object* l_Lean_Elab_Term_Do_ToTerm_seqToTerm___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_expandTermFor___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOptionM___at_Lean_Elab_Term_Do_mkJmp___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkMonadAlias___closed__7; +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__3___boxed(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_Term_Do_attachJP(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_getDoSeqElems___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_concat___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1220,6 +1227,7 @@ static lean_object* l_Lean_Elab_Term_Do_ToTerm_seqToTerm___closed__7; static lean_object* l_Lean_Elab_Term_Do_ToTerm_breakToTerm___closed__5; static lean_object* l_Lean_Elab_Term_Do_ToTerm_matchNestedTermResult___closed__21; static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__11; +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__3(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_Term_Do_ToTerm_mkNestedKind(uint8_t, uint8_t, uint8_t); static lean_object* l___regBuiltin_Lean_Elab_Term_expandTermFor_declRange___closed__5; static lean_object* l_Lean_Elab_Term_Do_ToTerm_reassignToTerm___closed__2; @@ -1259,8 +1267,9 @@ static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doReassignArrowToCode___lamb static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___spec__10___closed__15; lean_object* l_Lean_indentExpr(lean_object*); static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__19; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Do_ToCodeBlock_doSeqToCode___spec__10___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkTuple___spec__1___closed__7; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLiftMethod___closed__9; static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__30; @@ -1279,7 +1288,7 @@ lean_object* l_Lean_MonadRef_mkInfoFromRefPos___rarg(lean_object*, lean_object*) static lean_object* l_Lean_Elab_Term_Do_ToTerm_mkIte___closed__4; LEAN_EXPORT uint8_t l___private_Lean_Elab_Do_0__Lean_Elab_Term_liftMethodDelimiter(lean_object*); static lean_object* l_Lean_Elab_Term_Do_getDoReassignVars___closed__1; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__9___boxed(lean_object**); static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__15; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_Do_ToCodeBlock_checkLetArrowRHS___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__25; @@ -1320,6 +1329,7 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_Do_ToCodeBlock_ch static lean_object* l_Lean_Elab_Term_Do_ToTerm_breakToTerm___closed__12; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_mkForInBody___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkTuple(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Do_ToTerm_continueToTerm___closed__17; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLiftMethod_declRange___closed__6; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___spec__10(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1746,19 +1756,20 @@ return x_6; static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLiftMethod_declRange___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(30u); -x_2 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_2, 0, x_1); -lean_ctor_set(x_2, 1, x_1); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(31u); +x_2 = lean_unsigned_to_nat(30u); +x_3 = lean_alloc_ctor(0, 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___regBuiltin_Lean_Elab_Term_elabLiftMethod_declRange___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(31u); +x_1 = lean_unsigned_to_nat(32u); x_2 = lean_unsigned_to_nat(87u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1786,7 +1797,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLiftMethod_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(30u); +x_1 = lean_unsigned_to_nat(31u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1798,7 +1809,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLiftMethod_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(30u); +x_1 = lean_unsigned_to_nat(31u); x_2 = lean_unsigned_to_nat(48u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11498,7 +11509,7 @@ static lean_object* _init_l_Lean_Elab_Term_Do_mkUnless___closed__1() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean_nullKind; +x_2 = l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__4___closed__2; x_3 = l_Lean_Elab_Term_Do_instInhabitedAlt___rarg___closed__1; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -13344,7 +13355,7 @@ x_9 = lean_array_push(x_8, x_5); x_10 = l_Lean_Elab_Term_Do_mkUnless___closed__1; x_11 = lean_array_push(x_9, x_10); x_12 = lean_box(2); -x_13 = l_Lean_nullKind; +x_13 = l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__4___closed__2; x_14 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_14, 0, x_12); lean_ctor_set(x_14, 1, x_13); @@ -13368,7 +13379,7 @@ lean_dec(x_2); x_4 = 0; x_5 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_mkDoSeq___spec__1(x_3, x_4, x_1); x_6 = lean_box(2); -x_7 = l_Lean_nullKind; +x_7 = l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__4___closed__2; x_8 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_8, 0, x_6); lean_ctor_set(x_8, 1, x_7); @@ -13479,9 +13490,9 @@ LEAN_EXPORT lean_object* l_Array_sequenceMap___at___private_Lean_Elab_Do_0__Lean { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_3 = lean_array_get_size(x_1); -x_4 = lean_unsigned_to_nat(0u); -x_5 = l_Lean_Elab_Term_Do_instInhabitedAlt___rarg___closed__1; -x_6 = l_Array_sequenceMap_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__2(x_1, x_2, x_3, x_4, x_5); +x_4 = lean_mk_empty_array_with_capacity(x_3); +x_5 = lean_unsigned_to_nat(0u); +x_6 = l_Array_sequenceMap_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__2(x_1, x_2, x_3, x_5, x_4); return x_6; } } @@ -14405,593 +14416,593 @@ return x_22; } } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__6(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_object* x_9, lean_object* x_10, size_t x_11, size_t x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__6(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, size_t x_13, size_t x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { _start: { -uint8_t x_16; -x_16 = lean_usize_dec_lt(x_12, x_11); -if (x_16 == 0) +uint8_t x_18; +x_18 = lean_usize_dec_lt(x_14, x_13); +if (x_18 == 0) { -lean_object* x_17; -lean_dec(x_14); +lean_object* x_19; +lean_dec(x_16); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_13); -lean_ctor_set(x_17, 1, x_15); -return x_17; +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_15); +lean_ctor_set(x_19, 1, x_17); +return x_19; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_18 = lean_array_uget(x_10, x_12); -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -x_20 = lean_ctor_get(x_13, 1); -lean_inc(x_20); -x_21 = lean_unbox(x_20); -if (x_21 == 0) -{ -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; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; size_t x_41; size_t x_42; -x_22 = lean_ctor_get(x_18, 0); +lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_20 = lean_array_uget(x_12, x_14); +x_21 = lean_ctor_get(x_20, 1); +lean_inc(x_21); +x_22 = lean_ctor_get(x_15, 1); lean_inc(x_22); -lean_dec(x_18); -x_23 = lean_ctor_get(x_19, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_19, 1); +x_23 = lean_unbox(x_22); +if (x_23 == 0) +{ +lean_object* 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; size_t x_43; size_t x_44; +x_24 = lean_ctor_get(x_20, 0); lean_inc(x_24); -lean_dec(x_19); -x_25 = lean_ctor_get(x_13, 0); +lean_dec(x_20); +x_25 = lean_ctor_get(x_21, 0); lean_inc(x_25); -lean_dec(x_13); -lean_inc(x_14); -x_26 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_14, x_15); -x_27 = lean_ctor_get(x_26, 1); +x_26 = lean_ctor_get(x_21, 1); +lean_inc(x_26); +lean_dec(x_21); +x_27 = lean_ctor_get(x_15, 0); lean_inc(x_27); -lean_dec(x_26); -lean_inc(x_25); +lean_dec(x_15); +lean_inc(x_16); +x_28 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_16, x_17); +x_29 = lean_ctor_get(x_28, 1); +lean_inc(x_29); +lean_dec(x_28); +lean_inc(x_27); +lean_inc(x_10); +x_30 = lean_array_push(x_10, x_27); +lean_inc(x_11); +x_31 = lean_array_push(x_30, x_11); +x_32 = lean_box(2); lean_inc(x_8); -x_28 = lean_array_push(x_8, x_25); -lean_inc(x_9); -x_29 = lean_array_push(x_28, x_9); -x_30 = lean_box(2); -lean_inc(x_6); -x_31 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_6); -lean_ctor_set(x_31, 2, x_29); -lean_inc(x_7); -x_32 = lean_array_push(x_7, x_31); -lean_inc(x_5); x_33 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_33, 0, x_30); -lean_ctor_set(x_33, 1, x_5); -lean_ctor_set(x_33, 2, x_32); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_8); +lean_ctor_set(x_33, 2, x_31); +lean_inc(x_9); +x_34 = lean_array_push(x_9, x_33); lean_inc(x_7); -x_34 = lean_array_push(x_7, x_33); -lean_inc(x_4); x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_30); -lean_ctor_set(x_35, 1, x_4); +lean_ctor_set(x_35, 0, x_32); +lean_ctor_set(x_35, 1, x_7); lean_ctor_set(x_35, 2, x_34); -x_36 = lean_unbox(x_20); -lean_dec(x_20); -lean_inc(x_7); -lean_inc(x_2); -lean_inc(x_5); lean_inc(x_9); -lean_inc(x_8); +x_36 = lean_array_push(x_9, x_35); +lean_inc(x_6); +x_37 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_37, 0, x_32); +lean_ctor_set(x_37, 1, x_6); +lean_ctor_set(x_37, 2, x_36); +x_38 = lean_unbox(x_22); +lean_dec(x_22); +lean_inc(x_9); +lean_inc(x_2); +lean_inc(x_7); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_3); lean_inc(x_1); -x_37 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__6___lambda__1(x_23, x_1, x_3, x_22, x_8, x_9, x_5, x_24, x_2, x_7, x_25, x_36, x_35, x_14, x_27); -lean_dec(x_25); -lean_dec(x_22); -x_38 = lean_ctor_get(x_37, 0); -lean_inc(x_38); -x_39 = lean_ctor_get(x_37, 1); -lean_inc(x_39); -lean_dec(x_37); -x_40 = lean_ctor_get(x_38, 0); +x_39 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__6___lambda__1(x_25, x_1, x_3, x_24, x_10, x_11, x_7, x_26, x_2, x_9, x_27, x_38, x_37, x_16, x_29); +lean_dec(x_27); +lean_dec(x_24); +x_40 = lean_ctor_get(x_39, 0); lean_inc(x_40); -lean_dec(x_38); -x_41 = 1; -x_42 = lean_usize_add(x_12, x_41); -x_12 = x_42; -x_13 = x_40; -x_15 = x_39; +x_41 = lean_ctor_get(x_39, 1); +lean_inc(x_41); +lean_dec(x_39); +x_42 = lean_ctor_get(x_40, 0); +lean_inc(x_42); +lean_dec(x_40); +x_43 = 1; +x_44 = lean_usize_add(x_14, x_43); +x_14 = x_44; +x_15 = x_42; +x_17 = x_41; goto _start; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; size_t x_53; size_t x_54; -x_44 = lean_ctor_get(x_18, 0); -lean_inc(x_44); -lean_dec(x_18); -x_45 = lean_ctor_get(x_19, 0); -lean_inc(x_45); -x_46 = lean_ctor_get(x_19, 1); +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; size_t x_55; size_t x_56; +x_46 = lean_ctor_get(x_20, 0); lean_inc(x_46); -lean_dec(x_19); -x_47 = lean_ctor_get(x_13, 0); -lean_inc(x_47); -lean_dec(x_13); -x_48 = lean_unbox(x_20); lean_dec(x_20); +x_47 = lean_ctor_get(x_21, 0); lean_inc(x_47); -lean_inc(x_7); -lean_inc(x_2); -lean_inc(x_5); +x_48 = lean_ctor_get(x_21, 1); +lean_inc(x_48); +lean_dec(x_21); +x_49 = lean_ctor_get(x_15, 0); +lean_inc(x_49); +lean_dec(x_15); +x_50 = lean_unbox(x_22); +lean_dec(x_22); +lean_inc(x_49); lean_inc(x_9); -lean_inc(x_8); +lean_inc(x_2); +lean_inc(x_7); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_3); lean_inc(x_1); -x_49 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__6___lambda__1(x_45, x_1, x_3, x_44, x_8, x_9, x_5, x_46, x_2, x_7, x_47, x_48, x_47, x_14, x_15); -lean_dec(x_47); -lean_dec(x_44); -x_50 = lean_ctor_get(x_49, 0); -lean_inc(x_50); -x_51 = lean_ctor_get(x_49, 1); -lean_inc(x_51); +x_51 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__6___lambda__1(x_47, x_1, x_3, x_46, x_10, x_11, x_7, x_48, x_2, x_9, x_49, x_50, x_49, x_16, x_17); lean_dec(x_49); -x_52 = lean_ctor_get(x_50, 0); +lean_dec(x_46); +x_52 = lean_ctor_get(x_51, 0); lean_inc(x_52); -lean_dec(x_50); -x_53 = 1; -x_54 = lean_usize_add(x_12, x_53); -x_12 = x_54; -x_13 = x_52; -x_15 = x_51; +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +lean_dec(x_51); +x_54 = lean_ctor_get(x_52, 0); +lean_inc(x_54); +lean_dec(x_52); +x_55 = 1; +x_56 = lean_usize_add(x_14, x_55); +x_14 = x_56; +x_15 = x_54; +x_17 = x_53; goto _start; } } } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__7(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_object* x_9, lean_object* x_10, size_t x_11, size_t x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__7(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, size_t x_13, size_t x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { _start: { -uint8_t x_16; -x_16 = lean_usize_dec_lt(x_12, x_11); -if (x_16 == 0) +uint8_t x_18; +x_18 = lean_usize_dec_lt(x_14, x_13); +if (x_18 == 0) { -lean_object* x_17; -lean_dec(x_14); +lean_object* x_19; +lean_dec(x_16); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_13); -lean_ctor_set(x_17, 1, x_15); -return x_17; +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_15); +lean_ctor_set(x_19, 1, x_17); +return x_19; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_18 = lean_array_uget(x_10, x_12); -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -x_20 = lean_ctor_get(x_13, 1); -lean_inc(x_20); -x_21 = lean_unbox(x_20); -if (x_21 == 0) -{ -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; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; size_t x_41; size_t x_42; -x_22 = lean_ctor_get(x_18, 0); +lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_20 = lean_array_uget(x_12, x_14); +x_21 = lean_ctor_get(x_20, 1); +lean_inc(x_21); +x_22 = lean_ctor_get(x_15, 1); lean_inc(x_22); -lean_dec(x_18); -x_23 = lean_ctor_get(x_19, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_19, 1); +x_23 = lean_unbox(x_22); +if (x_23 == 0) +{ +lean_object* 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; size_t x_43; size_t x_44; +x_24 = lean_ctor_get(x_20, 0); lean_inc(x_24); -lean_dec(x_19); -x_25 = lean_ctor_get(x_13, 0); +lean_dec(x_20); +x_25 = lean_ctor_get(x_21, 0); lean_inc(x_25); -lean_dec(x_13); -lean_inc(x_14); -x_26 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_14, x_15); -x_27 = lean_ctor_get(x_26, 1); +x_26 = lean_ctor_get(x_21, 1); +lean_inc(x_26); +lean_dec(x_21); +x_27 = lean_ctor_get(x_15, 0); lean_inc(x_27); -lean_dec(x_26); -lean_inc(x_25); +lean_dec(x_15); +lean_inc(x_16); +x_28 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_16, x_17); +x_29 = lean_ctor_get(x_28, 1); +lean_inc(x_29); +lean_dec(x_28); +lean_inc(x_27); +lean_inc(x_10); +x_30 = lean_array_push(x_10, x_27); +lean_inc(x_11); +x_31 = lean_array_push(x_30, x_11); +x_32 = lean_box(2); lean_inc(x_8); -x_28 = lean_array_push(x_8, x_25); -lean_inc(x_9); -x_29 = lean_array_push(x_28, x_9); -x_30 = lean_box(2); -lean_inc(x_6); -x_31 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_6); -lean_ctor_set(x_31, 2, x_29); -lean_inc(x_7); -x_32 = lean_array_push(x_7, x_31); -lean_inc(x_5); x_33 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_33, 0, x_30); -lean_ctor_set(x_33, 1, x_5); -lean_ctor_set(x_33, 2, x_32); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_8); +lean_ctor_set(x_33, 2, x_31); +lean_inc(x_9); +x_34 = lean_array_push(x_9, x_33); lean_inc(x_7); -x_34 = lean_array_push(x_7, x_33); -lean_inc(x_4); x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_30); -lean_ctor_set(x_35, 1, x_4); +lean_ctor_set(x_35, 0, x_32); +lean_ctor_set(x_35, 1, x_7); lean_ctor_set(x_35, 2, x_34); -x_36 = lean_unbox(x_20); -lean_dec(x_20); -lean_inc(x_7); -lean_inc(x_2); -lean_inc(x_5); lean_inc(x_9); -lean_inc(x_8); +x_36 = lean_array_push(x_9, x_35); +lean_inc(x_6); +x_37 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_37, 0, x_32); +lean_ctor_set(x_37, 1, x_6); +lean_ctor_set(x_37, 2, x_36); +x_38 = lean_unbox(x_22); +lean_dec(x_22); +lean_inc(x_9); +lean_inc(x_2); +lean_inc(x_7); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_3); lean_inc(x_1); -x_37 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__6___lambda__1(x_23, x_1, x_3, x_22, x_8, x_9, x_5, x_24, x_2, x_7, x_25, x_36, x_35, x_14, x_27); -lean_dec(x_25); -lean_dec(x_22); -x_38 = lean_ctor_get(x_37, 0); -lean_inc(x_38); -x_39 = lean_ctor_get(x_37, 1); -lean_inc(x_39); -lean_dec(x_37); -x_40 = lean_ctor_get(x_38, 0); +x_39 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__6___lambda__1(x_25, x_1, x_3, x_24, x_10, x_11, x_7, x_26, x_2, x_9, x_27, x_38, x_37, x_16, x_29); +lean_dec(x_27); +lean_dec(x_24); +x_40 = lean_ctor_get(x_39, 0); lean_inc(x_40); -lean_dec(x_38); -x_41 = 1; -x_42 = lean_usize_add(x_12, x_41); -x_12 = x_42; -x_13 = x_40; -x_15 = x_39; +x_41 = lean_ctor_get(x_39, 1); +lean_inc(x_41); +lean_dec(x_39); +x_42 = lean_ctor_get(x_40, 0); +lean_inc(x_42); +lean_dec(x_40); +x_43 = 1; +x_44 = lean_usize_add(x_14, x_43); +x_14 = x_44; +x_15 = x_42; +x_17 = x_41; goto _start; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; size_t x_53; size_t x_54; -x_44 = lean_ctor_get(x_18, 0); -lean_inc(x_44); -lean_dec(x_18); -x_45 = lean_ctor_get(x_19, 0); -lean_inc(x_45); -x_46 = lean_ctor_get(x_19, 1); +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; size_t x_55; size_t x_56; +x_46 = lean_ctor_get(x_20, 0); lean_inc(x_46); -lean_dec(x_19); -x_47 = lean_ctor_get(x_13, 0); -lean_inc(x_47); -lean_dec(x_13); -x_48 = lean_unbox(x_20); lean_dec(x_20); +x_47 = lean_ctor_get(x_21, 0); lean_inc(x_47); -lean_inc(x_7); -lean_inc(x_2); -lean_inc(x_5); +x_48 = lean_ctor_get(x_21, 1); +lean_inc(x_48); +lean_dec(x_21); +x_49 = lean_ctor_get(x_15, 0); +lean_inc(x_49); +lean_dec(x_15); +x_50 = lean_unbox(x_22); +lean_dec(x_22); +lean_inc(x_49); lean_inc(x_9); -lean_inc(x_8); +lean_inc(x_2); +lean_inc(x_7); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_3); lean_inc(x_1); -x_49 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__6___lambda__1(x_45, x_1, x_3, x_44, x_8, x_9, x_5, x_46, x_2, x_7, x_47, x_48, x_47, x_14, x_15); -lean_dec(x_47); -lean_dec(x_44); -x_50 = lean_ctor_get(x_49, 0); -lean_inc(x_50); -x_51 = lean_ctor_get(x_49, 1); -lean_inc(x_51); +x_51 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__6___lambda__1(x_47, x_1, x_3, x_46, x_10, x_11, x_7, x_48, x_2, x_9, x_49, x_50, x_49, x_16, x_17); lean_dec(x_49); -x_52 = lean_ctor_get(x_50, 0); +lean_dec(x_46); +x_52 = lean_ctor_get(x_51, 0); lean_inc(x_52); -lean_dec(x_50); -x_53 = 1; -x_54 = lean_usize_add(x_12, x_53); -x_12 = x_54; -x_13 = x_52; -x_15 = x_51; +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +lean_dec(x_51); +x_54 = lean_ctor_get(x_52, 0); +lean_inc(x_54); +lean_dec(x_52); +x_55 = 1; +x_56 = lean_usize_add(x_14, x_55); +x_14 = x_56; +x_15 = x_54; +x_17 = x_53; goto _start; } } } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__8(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_object* x_9, lean_object* x_10, size_t x_11, size_t x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__8(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, size_t x_13, size_t x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { _start: { -uint8_t x_16; -x_16 = lean_usize_dec_lt(x_12, x_11); -if (x_16 == 0) +uint8_t x_18; +x_18 = lean_usize_dec_lt(x_14, x_13); +if (x_18 == 0) { -lean_object* x_17; -lean_dec(x_14); +lean_object* x_19; +lean_dec(x_16); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_13); -lean_ctor_set(x_17, 1, x_15); -return x_17; +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_15); +lean_ctor_set(x_19, 1, x_17); +return x_19; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_18 = lean_array_uget(x_10, x_12); -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -x_20 = lean_ctor_get(x_13, 1); -lean_inc(x_20); -x_21 = lean_unbox(x_20); -if (x_21 == 0) -{ -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; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; size_t x_41; size_t x_42; -x_22 = lean_ctor_get(x_18, 0); +lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_20 = lean_array_uget(x_12, x_14); +x_21 = lean_ctor_get(x_20, 1); +lean_inc(x_21); +x_22 = lean_ctor_get(x_15, 1); lean_inc(x_22); -lean_dec(x_18); -x_23 = lean_ctor_get(x_19, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_19, 1); +x_23 = lean_unbox(x_22); +if (x_23 == 0) +{ +lean_object* 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; size_t x_43; size_t x_44; +x_24 = lean_ctor_get(x_20, 0); lean_inc(x_24); -lean_dec(x_19); -x_25 = lean_ctor_get(x_13, 0); +lean_dec(x_20); +x_25 = lean_ctor_get(x_21, 0); lean_inc(x_25); -lean_dec(x_13); -lean_inc(x_14); -x_26 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_14, x_15); -x_27 = lean_ctor_get(x_26, 1); +x_26 = lean_ctor_get(x_21, 1); +lean_inc(x_26); +lean_dec(x_21); +x_27 = lean_ctor_get(x_15, 0); lean_inc(x_27); -lean_dec(x_26); -lean_inc(x_25); +lean_dec(x_15); +lean_inc(x_16); +x_28 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_16, x_17); +x_29 = lean_ctor_get(x_28, 1); +lean_inc(x_29); +lean_dec(x_28); +lean_inc(x_27); +lean_inc(x_10); +x_30 = lean_array_push(x_10, x_27); +lean_inc(x_11); +x_31 = lean_array_push(x_30, x_11); +x_32 = lean_box(2); lean_inc(x_8); -x_28 = lean_array_push(x_8, x_25); -lean_inc(x_9); -x_29 = lean_array_push(x_28, x_9); -x_30 = lean_box(2); -lean_inc(x_6); -x_31 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_6); -lean_ctor_set(x_31, 2, x_29); -lean_inc(x_7); -x_32 = lean_array_push(x_7, x_31); -lean_inc(x_5); x_33 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_33, 0, x_30); -lean_ctor_set(x_33, 1, x_5); -lean_ctor_set(x_33, 2, x_32); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_8); +lean_ctor_set(x_33, 2, x_31); +lean_inc(x_9); +x_34 = lean_array_push(x_9, x_33); lean_inc(x_7); -x_34 = lean_array_push(x_7, x_33); -lean_inc(x_4); x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_30); -lean_ctor_set(x_35, 1, x_4); +lean_ctor_set(x_35, 0, x_32); +lean_ctor_set(x_35, 1, x_7); lean_ctor_set(x_35, 2, x_34); -x_36 = lean_unbox(x_20); -lean_dec(x_20); -lean_inc(x_7); -lean_inc(x_2); -lean_inc(x_5); lean_inc(x_9); -lean_inc(x_8); +x_36 = lean_array_push(x_9, x_35); +lean_inc(x_6); +x_37 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_37, 0, x_32); +lean_ctor_set(x_37, 1, x_6); +lean_ctor_set(x_37, 2, x_36); +x_38 = lean_unbox(x_22); +lean_dec(x_22); +lean_inc(x_9); +lean_inc(x_2); +lean_inc(x_7); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_3); lean_inc(x_1); -x_37 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__6___lambda__1(x_23, x_1, x_3, x_22, x_8, x_9, x_5, x_24, x_2, x_7, x_25, x_36, x_35, x_14, x_27); -lean_dec(x_25); -lean_dec(x_22); -x_38 = lean_ctor_get(x_37, 0); -lean_inc(x_38); -x_39 = lean_ctor_get(x_37, 1); -lean_inc(x_39); -lean_dec(x_37); -x_40 = lean_ctor_get(x_38, 0); +x_39 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__6___lambda__1(x_25, x_1, x_3, x_24, x_10, x_11, x_7, x_26, x_2, x_9, x_27, x_38, x_37, x_16, x_29); +lean_dec(x_27); +lean_dec(x_24); +x_40 = lean_ctor_get(x_39, 0); lean_inc(x_40); -lean_dec(x_38); -x_41 = 1; -x_42 = lean_usize_add(x_12, x_41); -x_12 = x_42; -x_13 = x_40; -x_15 = x_39; +x_41 = lean_ctor_get(x_39, 1); +lean_inc(x_41); +lean_dec(x_39); +x_42 = lean_ctor_get(x_40, 0); +lean_inc(x_42); +lean_dec(x_40); +x_43 = 1; +x_44 = lean_usize_add(x_14, x_43); +x_14 = x_44; +x_15 = x_42; +x_17 = x_41; goto _start; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; size_t x_53; size_t x_54; -x_44 = lean_ctor_get(x_18, 0); -lean_inc(x_44); -lean_dec(x_18); -x_45 = lean_ctor_get(x_19, 0); -lean_inc(x_45); -x_46 = lean_ctor_get(x_19, 1); +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; size_t x_55; size_t x_56; +x_46 = lean_ctor_get(x_20, 0); lean_inc(x_46); -lean_dec(x_19); -x_47 = lean_ctor_get(x_13, 0); -lean_inc(x_47); -lean_dec(x_13); -x_48 = lean_unbox(x_20); lean_dec(x_20); +x_47 = lean_ctor_get(x_21, 0); lean_inc(x_47); -lean_inc(x_7); -lean_inc(x_2); -lean_inc(x_5); +x_48 = lean_ctor_get(x_21, 1); +lean_inc(x_48); +lean_dec(x_21); +x_49 = lean_ctor_get(x_15, 0); +lean_inc(x_49); +lean_dec(x_15); +x_50 = lean_unbox(x_22); +lean_dec(x_22); +lean_inc(x_49); lean_inc(x_9); -lean_inc(x_8); +lean_inc(x_2); +lean_inc(x_7); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_3); lean_inc(x_1); -x_49 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__6___lambda__1(x_45, x_1, x_3, x_44, x_8, x_9, x_5, x_46, x_2, x_7, x_47, x_48, x_47, x_14, x_15); -lean_dec(x_47); -lean_dec(x_44); -x_50 = lean_ctor_get(x_49, 0); -lean_inc(x_50); -x_51 = lean_ctor_get(x_49, 1); -lean_inc(x_51); +x_51 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__6___lambda__1(x_47, x_1, x_3, x_46, x_10, x_11, x_7, x_48, x_2, x_9, x_49, x_50, x_49, x_16, x_17); lean_dec(x_49); -x_52 = lean_ctor_get(x_50, 0); +lean_dec(x_46); +x_52 = lean_ctor_get(x_51, 0); lean_inc(x_52); -lean_dec(x_50); -x_53 = 1; -x_54 = lean_usize_add(x_12, x_53); -x_12 = x_54; -x_13 = x_52; -x_15 = x_51; +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +lean_dec(x_51); +x_54 = lean_ctor_get(x_52, 0); +lean_inc(x_54); +lean_dec(x_52); +x_55 = 1; +x_56 = lean_usize_add(x_14, x_55); +x_14 = x_56; +x_15 = x_54; +x_17 = x_53; goto _start; } } } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__9(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_object* x_9, lean_object* x_10, size_t x_11, size_t x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__9(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, size_t x_13, size_t x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { _start: { -uint8_t x_16; -x_16 = lean_usize_dec_lt(x_12, x_11); -if (x_16 == 0) +uint8_t x_18; +x_18 = lean_usize_dec_lt(x_14, x_13); +if (x_18 == 0) { -lean_object* x_17; -lean_dec(x_14); +lean_object* x_19; +lean_dec(x_16); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_13); -lean_ctor_set(x_17, 1, x_15); -return x_17; +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_15); +lean_ctor_set(x_19, 1, x_17); +return x_19; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_18 = lean_array_uget(x_10, x_12); -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -x_20 = lean_ctor_get(x_13, 1); -lean_inc(x_20); -x_21 = lean_unbox(x_20); -if (x_21 == 0) -{ -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; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; size_t x_41; size_t x_42; -x_22 = lean_ctor_get(x_18, 0); +lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_20 = lean_array_uget(x_12, x_14); +x_21 = lean_ctor_get(x_20, 1); +lean_inc(x_21); +x_22 = lean_ctor_get(x_15, 1); lean_inc(x_22); -lean_dec(x_18); -x_23 = lean_ctor_get(x_19, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_19, 1); +x_23 = lean_unbox(x_22); +if (x_23 == 0) +{ +lean_object* 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; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; size_t x_43; size_t x_44; +x_24 = lean_ctor_get(x_20, 0); lean_inc(x_24); -lean_dec(x_19); -x_25 = lean_ctor_get(x_13, 0); +lean_dec(x_20); +x_25 = lean_ctor_get(x_21, 0); lean_inc(x_25); -lean_dec(x_13); -lean_inc(x_14); -x_26 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_14, x_15); -x_27 = lean_ctor_get(x_26, 1); +x_26 = lean_ctor_get(x_21, 1); +lean_inc(x_26); +lean_dec(x_21); +x_27 = lean_ctor_get(x_15, 0); lean_inc(x_27); -lean_dec(x_26); -lean_inc(x_25); +lean_dec(x_15); +lean_inc(x_16); +x_28 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_16, x_17); +x_29 = lean_ctor_get(x_28, 1); +lean_inc(x_29); +lean_dec(x_28); +lean_inc(x_27); +lean_inc(x_10); +x_30 = lean_array_push(x_10, x_27); +lean_inc(x_11); +x_31 = lean_array_push(x_30, x_11); +x_32 = lean_box(2); lean_inc(x_8); -x_28 = lean_array_push(x_8, x_25); -lean_inc(x_9); -x_29 = lean_array_push(x_28, x_9); -x_30 = lean_box(2); -lean_inc(x_6); -x_31 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_6); -lean_ctor_set(x_31, 2, x_29); -lean_inc(x_7); -x_32 = lean_array_push(x_7, x_31); -lean_inc(x_5); x_33 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_33, 0, x_30); -lean_ctor_set(x_33, 1, x_5); -lean_ctor_set(x_33, 2, x_32); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_8); +lean_ctor_set(x_33, 2, x_31); +lean_inc(x_9); +x_34 = lean_array_push(x_9, x_33); lean_inc(x_7); -x_34 = lean_array_push(x_7, x_33); -lean_inc(x_4); x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_30); -lean_ctor_set(x_35, 1, x_4); +lean_ctor_set(x_35, 0, x_32); +lean_ctor_set(x_35, 1, x_7); lean_ctor_set(x_35, 2, x_34); -x_36 = lean_unbox(x_20); -lean_dec(x_20); -lean_inc(x_7); -lean_inc(x_2); -lean_inc(x_5); lean_inc(x_9); -lean_inc(x_8); +x_36 = lean_array_push(x_9, x_35); +lean_inc(x_6); +x_37 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_37, 0, x_32); +lean_ctor_set(x_37, 1, x_6); +lean_ctor_set(x_37, 2, x_36); +x_38 = lean_unbox(x_22); +lean_dec(x_22); +lean_inc(x_9); +lean_inc(x_2); +lean_inc(x_7); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_3); lean_inc(x_1); -x_37 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__6___lambda__1(x_23, x_1, x_3, x_22, x_8, x_9, x_5, x_24, x_2, x_7, x_25, x_36, x_35, x_14, x_27); -lean_dec(x_25); -lean_dec(x_22); -x_38 = lean_ctor_get(x_37, 0); -lean_inc(x_38); -x_39 = lean_ctor_get(x_37, 1); -lean_inc(x_39); -lean_dec(x_37); -x_40 = lean_ctor_get(x_38, 0); +x_39 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__6___lambda__1(x_25, x_1, x_3, x_24, x_10, x_11, x_7, x_26, x_2, x_9, x_27, x_38, x_37, x_16, x_29); +lean_dec(x_27); +lean_dec(x_24); +x_40 = lean_ctor_get(x_39, 0); lean_inc(x_40); -lean_dec(x_38); -x_41 = 1; -x_42 = lean_usize_add(x_12, x_41); -x_12 = x_42; -x_13 = x_40; -x_15 = x_39; +x_41 = lean_ctor_get(x_39, 1); +lean_inc(x_41); +lean_dec(x_39); +x_42 = lean_ctor_get(x_40, 0); +lean_inc(x_42); +lean_dec(x_40); +x_43 = 1; +x_44 = lean_usize_add(x_14, x_43); +x_14 = x_44; +x_15 = x_42; +x_17 = x_41; goto _start; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; size_t x_53; size_t x_54; -x_44 = lean_ctor_get(x_18, 0); -lean_inc(x_44); -lean_dec(x_18); -x_45 = lean_ctor_get(x_19, 0); -lean_inc(x_45); -x_46 = lean_ctor_get(x_19, 1); +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; size_t x_55; size_t x_56; +x_46 = lean_ctor_get(x_20, 0); lean_inc(x_46); -lean_dec(x_19); -x_47 = lean_ctor_get(x_13, 0); -lean_inc(x_47); -lean_dec(x_13); -x_48 = lean_unbox(x_20); lean_dec(x_20); +x_47 = lean_ctor_get(x_21, 0); lean_inc(x_47); -lean_inc(x_7); -lean_inc(x_2); -lean_inc(x_5); +x_48 = lean_ctor_get(x_21, 1); +lean_inc(x_48); +lean_dec(x_21); +x_49 = lean_ctor_get(x_15, 0); +lean_inc(x_49); +lean_dec(x_15); +x_50 = lean_unbox(x_22); +lean_dec(x_22); +lean_inc(x_49); lean_inc(x_9); -lean_inc(x_8); +lean_inc(x_2); +lean_inc(x_7); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_3); lean_inc(x_1); -x_49 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__6___lambda__1(x_45, x_1, x_3, x_44, x_8, x_9, x_5, x_46, x_2, x_7, x_47, x_48, x_47, x_14, x_15); -lean_dec(x_47); -lean_dec(x_44); -x_50 = lean_ctor_get(x_49, 0); -lean_inc(x_50); -x_51 = lean_ctor_get(x_49, 1); -lean_inc(x_51); +x_51 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__6___lambda__1(x_47, x_1, x_3, x_46, x_10, x_11, x_7, x_48, x_2, x_9, x_49, x_50, x_49, x_16, x_17); lean_dec(x_49); -x_52 = lean_ctor_get(x_50, 0); -lean_inc(x_52); -lean_dec(x_50); -x_53 = 1; -x_54 = lean_usize_add(x_12, x_53); -x_12 = x_54; -x_13 = x_52; -x_15 = x_51; +lean_dec(x_46); +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +lean_dec(x_51); +x_54 = lean_ctor_get(x_52, 0); +lean_inc(x_54); +lean_dec(x_52); +x_55 = 1; +x_56 = lean_usize_add(x_14, x_55); +x_14 = x_56; +x_15 = x_54; +x_17 = x_53; goto _start; } } @@ -15071,362 +15082,412 @@ static lean_object* _init_l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDo _start: { lean_object* x_1; +x_1 = lean_mk_string_from_bytes("doSeq", 5); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; x_1 = lean_mk_string_from_bytes("doSeqItem", 9); return x_1; } } +static lean_object* _init_l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_Do_pullExitPointsAux___lambda__1___closed__9; +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___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__3; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} LEAN_EXPORT lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___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_object* x_9, lean_object* x_10, size_t x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { -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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; size_t x_73; -x_16 = lean_ctor_get(x_14, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_14, 1); -lean_inc(x_17); -x_18 = lean_ctor_get(x_14, 2); -lean_inc(x_18); -x_19 = lean_ctor_get(x_14, 3); -lean_inc(x_19); -x_20 = lean_ctor_get(x_14, 4); +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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; size_t x_77; +x_16 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__1; +lean_inc(x_1); +x_17 = lean_name_mk_string(x_1, x_16); +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +x_20 = lean_ctor_get(x_14, 0); lean_inc(x_20); -x_21 = lean_ctor_get(x_14, 5); +x_21 = lean_ctor_get(x_14, 1); lean_inc(x_21); -lean_dec(x_14); -x_22 = l_Lean_replaceRef(x_1, x_21); -lean_dec(x_21); -lean_inc(x_18); -lean_inc(x_17); -x_23 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_23, 0, x_16); -lean_ctor_set(x_23, 1, x_17); -lean_ctor_set(x_23, 2, x_18); -lean_ctor_set(x_23, 3, x_19); -lean_ctor_set(x_23, 4, x_20); -lean_ctor_set(x_23, 5, x_22); +x_22 = lean_ctor_get(x_14, 2); +lean_inc(x_22); +x_23 = lean_ctor_get(x_14, 3); lean_inc(x_23); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_23, x_15); -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -x_26 = lean_ctor_get(x_24, 1); -lean_inc(x_26); -lean_dec(x_24); -x_27 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_getDoSeqElems___closed__9; -lean_inc(x_2); -x_28 = lean_name_mk_string(x_2, x_27); -x_29 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__1; -lean_inc(x_2); -x_30 = lean_name_mk_string(x_2, x_29); -x_31 = l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__4___closed__7; -lean_inc(x_2); -x_32 = lean_name_mk_string(x_2, x_31); -x_33 = l_Lean_Elab_Term_Do_pullExitPointsAux___lambda__1___closed__1; -lean_inc(x_2); -x_34 = lean_name_mk_string(x_2, x_33); -x_35 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkPureUnit___closed__3; -lean_inc(x_18); -lean_inc(x_17); -x_36 = l_Lean_addMacroScope(x_17, x_35, x_18); -x_37 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkPureUnit___closed__2; -x_38 = l_Lean_Elab_Term_Do_pullExitPointsAux___lambda__1___closed__11; +x_24 = lean_ctor_get(x_14, 4); +lean_inc(x_24); +x_25 = lean_ctor_get(x_14, 5); lean_inc(x_25); -x_39 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_39, 0, x_25); -lean_ctor_set(x_39, 1, x_37); -lean_ctor_set(x_39, 2, x_36); -lean_ctor_set(x_39, 3, x_38); -x_40 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkPureUnit___closed__7; -x_41 = l_Lean_addMacroScope(x_17, x_40, x_18); -x_42 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkPureUnit___closed__6; -x_43 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkPureUnit___closed__9; -x_44 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_44, 0, x_25); -lean_ctor_set(x_44, 1, x_42); -lean_ctor_set(x_44, 2, x_41); -lean_ctor_set(x_44, 3, x_43); -x_45 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_extractBind___lambda__1___closed__3; -x_46 = lean_array_push(x_45, x_44); -x_47 = lean_box(2); -x_48 = l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__4___closed__2; -x_49 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_49, 0, x_47); -lean_ctor_set(x_49, 1, x_48); -lean_ctor_set(x_49, 2, x_46); -x_50 = l_Lean_Elab_Term_Do_pullExitPointsAux___lambda__1___closed__12; -x_51 = lean_array_push(x_50, x_39); -x_52 = lean_array_push(x_51, x_49); -x_53 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_53, 0, x_47); -lean_ctor_set(x_53, 1, x_34); -lean_ctor_set(x_53, 2, x_52); -x_54 = lean_array_push(x_45, x_53); -x_55 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_55, 0, x_47); -lean_ctor_set(x_55, 1, x_32); -lean_ctor_set(x_55, 2, x_54); -x_56 = lean_array_push(x_50, x_55); -x_57 = l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__4___closed__3; -x_58 = lean_array_push(x_56, x_57); +lean_dec(x_14); +x_26 = l_Lean_replaceRef(x_2, x_25); +lean_dec(x_25); +lean_inc(x_22); +lean_inc(x_21); +x_27 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_27, 0, x_20); +lean_ctor_set(x_27, 1, x_21); +lean_ctor_set(x_27, 2, x_22); +lean_ctor_set(x_27, 3, x_23); +lean_ctor_set(x_27, 4, x_24); +lean_ctor_set(x_27, 5, x_26); +lean_inc(x_27); +x_28 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_27, x_15); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); lean_inc(x_30); +lean_dec(x_28); +x_31 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_getDoSeqElems___closed__9; +lean_inc(x_1); +x_32 = lean_name_mk_string(x_1, x_31); +x_33 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__2; +lean_inc(x_1); +x_34 = lean_name_mk_string(x_1, x_33); +x_35 = l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__4___closed__7; +lean_inc(x_1); +x_36 = lean_name_mk_string(x_1, x_35); +x_37 = l_Lean_Elab_Term_Do_pullExitPointsAux___lambda__1___closed__1; +lean_inc(x_1); +x_38 = lean_name_mk_string(x_1, x_37); +x_39 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkPureUnit___closed__3; +lean_inc(x_22); +lean_inc(x_21); +x_40 = l_Lean_addMacroScope(x_21, x_39, x_22); +x_41 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkPureUnit___closed__2; +x_42 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__4; +lean_inc(x_29); +x_43 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_43, 0, x_29); +lean_ctor_set(x_43, 1, x_41); +lean_ctor_set(x_43, 2, x_40); +lean_ctor_set(x_43, 3, x_42); +x_44 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkPureUnit___closed__7; +x_45 = l_Lean_addMacroScope(x_21, x_44, x_22); +x_46 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkPureUnit___closed__6; +x_47 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkPureUnit___closed__9; +x_48 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_48, 0, x_29); +lean_ctor_set(x_48, 1, x_46); +lean_ctor_set(x_48, 2, x_45); +lean_ctor_set(x_48, 3, x_47); +x_49 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_extractBind___lambda__1___closed__3; +x_50 = lean_array_push(x_49, x_48); +x_51 = lean_box(2); +x_52 = l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__4___closed__2; +x_53 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_52); +lean_ctor_set(x_53, 2, x_50); +x_54 = l_Lean_Elab_Term_Do_pullExitPointsAux___lambda__1___closed__12; +x_55 = lean_array_push(x_54, x_43); +x_56 = lean_array_push(x_55, x_53); +x_57 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_57, 0, x_51); +lean_ctor_set(x_57, 1, x_38); +lean_ctor_set(x_57, 2, x_56); +x_58 = lean_array_push(x_49, x_57); x_59 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_59, 0, x_47); -lean_ctor_set(x_59, 1, x_30); +lean_ctor_set(x_59, 0, x_51); +lean_ctor_set(x_59, 1, x_36); lean_ctor_set(x_59, 2, x_58); -x_60 = lean_array_push(x_45, x_59); -x_61 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_61, 0, x_47); -lean_ctor_set(x_61, 1, x_48); -lean_ctor_set(x_61, 2, x_60); -x_62 = lean_array_push(x_45, x_61); -lean_inc(x_28); +x_60 = lean_array_push(x_54, x_59); +x_61 = l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__4___closed__3; +x_62 = lean_array_push(x_60, x_61); +lean_inc(x_34); x_63 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_63, 0, x_47); -lean_ctor_set(x_63, 1, x_28); +lean_ctor_set(x_63, 0, x_51); +lean_ctor_set(x_63, 1, x_34); lean_ctor_set(x_63, 2, x_62); -x_64 = l_Array_reverse___rarg(x_3); -x_65 = lean_array_push(x_64, x_4); -x_66 = l_Array_reverse___rarg(x_5); -x_67 = lean_array_push(x_66, x_6); -x_68 = l_Array_reverse___rarg(x_7); -x_69 = lean_array_push(x_68, x_8); -x_70 = l_Array_zip___rarg(x_67, x_69); +x_64 = lean_array_push(x_49, x_63); +x_65 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_65, 0, x_51); +lean_ctor_set(x_65, 1, x_52); +lean_ctor_set(x_65, 2, x_64); +x_66 = lean_array_push(x_49, x_65); +lean_inc(x_32); +x_67 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_67, 0, x_51); +lean_ctor_set(x_67, 1, x_32); +lean_ctor_set(x_67, 2, x_66); +x_68 = l_Array_reverse___rarg(x_3); +x_69 = lean_array_push(x_68, x_4); +x_70 = l_Array_reverse___rarg(x_5); +x_71 = lean_array_push(x_70, x_6); +x_72 = l_Array_reverse___rarg(x_7); +x_73 = lean_array_push(x_72, x_8); +x_74 = l_Array_zip___rarg(x_71, x_73); +lean_dec(x_73); +lean_dec(x_71); +x_75 = l_Array_zip___rarg(x_69, x_74); +lean_dec(x_74); lean_dec(x_69); -lean_dec(x_67); -x_71 = l_Array_zip___rarg(x_65, x_70); -lean_dec(x_70); -lean_dec(x_65); -x_72 = lean_array_get_size(x_71); -x_73 = lean_usize_of_nat(x_72); -lean_dec(x_72); +x_76 = lean_array_get_size(x_75); +x_77 = lean_usize_of_nat(x_76); +lean_dec(x_76); if (lean_obj_tag(x_13) == 0) { -uint8_t x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; -x_74 = 1; -x_75 = lean_box(x_74); -x_76 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_76, 0, x_63); -lean_ctor_set(x_76, 1, x_75); -x_77 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__6(x_2, x_9, x_10, x_28, x_48, x_30, x_45, x_50, x_57, x_71, x_73, x_11, x_76, x_23, x_26); -lean_dec(x_71); -x_78 = lean_ctor_get(x_77, 0); -lean_inc(x_78); -x_79 = lean_ctor_get(x_77, 1); -lean_inc(x_79); -lean_dec(x_77); -x_80 = lean_ctor_get(x_78, 0); -lean_inc(x_80); -lean_dec(x_78); -x_81 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_81, 0, x_80); -x_82 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_82, 0, x_81); -lean_ctor_set(x_82, 1, x_79); -return x_82; +uint8_t x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_78 = 1; +x_79 = lean_box(x_78); +x_80 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_80, 0, x_67); +lean_ctor_set(x_80, 1, x_79); +x_81 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__6(x_1, x_9, x_10, x_18, x_19, x_32, x_52, x_34, x_49, x_54, x_61, x_75, x_77, x_11, x_80, x_27, x_30); +lean_dec(x_75); +lean_dec(x_19); +x_82 = lean_ctor_get(x_81, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_81, 1); +lean_inc(x_83); +lean_dec(x_81); +x_84 = lean_ctor_get(x_82, 0); +lean_inc(x_84); +lean_dec(x_82); +x_85 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_85, 0, x_84); +x_86 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_86, 0, x_85); +lean_ctor_set(x_86, 1, x_83); +return x_86; } else { -lean_object* x_83; uint8_t x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; -lean_dec(x_63); -x_83 = lean_ctor_get(x_13, 0); -x_84 = 1; -x_85 = lean_box(x_84); -lean_inc(x_83); -x_86 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_86, 0, x_83); -lean_ctor_set(x_86, 1, x_85); -x_87 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__7(x_2, x_9, x_10, x_28, x_48, x_30, x_45, x_50, x_57, x_71, x_73, x_11, x_86, x_23, x_26); -lean_dec(x_71); -x_88 = lean_ctor_get(x_87, 0); -lean_inc(x_88); -x_89 = lean_ctor_get(x_87, 1); -lean_inc(x_89); -lean_dec(x_87); -x_90 = lean_ctor_get(x_88, 0); -lean_inc(x_90); -lean_dec(x_88); -x_91 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_91, 0, x_90); -x_92 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_92, 0, x_91); -lean_ctor_set(x_92, 1, x_89); -return x_92; +lean_object* x_87; uint8_t x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; +lean_dec(x_67); +x_87 = lean_ctor_get(x_13, 0); +x_88 = 1; +x_89 = lean_box(x_88); +lean_inc(x_87); +x_90 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_90, 0, x_87); +lean_ctor_set(x_90, 1, x_89); +x_91 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__7(x_1, x_9, x_10, x_18, x_19, x_32, x_52, x_34, x_49, x_54, x_61, x_75, x_77, x_11, x_90, x_27, x_30); +lean_dec(x_75); +lean_dec(x_19); +x_92 = lean_ctor_get(x_91, 0); +lean_inc(x_92); +x_93 = lean_ctor_get(x_91, 1); +lean_inc(x_93); +lean_dec(x_91); +x_94 = lean_ctor_get(x_92, 0); +lean_inc(x_94); +lean_dec(x_92); +x_95 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_95, 0, x_94); +x_96 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_96, 0, x_95); +lean_ctor_set(x_96, 1, x_93); +return x_96; } } } LEAN_EXPORT lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__3(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; size_t x_71; size_t x_72; -x_14 = lean_ctor_get(x_12, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_12, 1); -lean_inc(x_15); -x_16 = lean_ctor_get(x_12, 2); -lean_inc(x_16); -x_17 = lean_ctor_get(x_12, 3); -lean_inc(x_17); -x_18 = lean_ctor_get(x_12, 4); +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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; size_t x_75; size_t x_76; +x_14 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__1; +lean_inc(x_1); +x_15 = lean_name_mk_string(x_1, x_14); +x_16 = lean_box(0); +x_17 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +x_18 = lean_ctor_get(x_12, 0); lean_inc(x_18); -x_19 = lean_ctor_get(x_12, 5); +x_19 = lean_ctor_get(x_12, 1); lean_inc(x_19); -lean_dec(x_12); -x_20 = l_Lean_replaceRef(x_1, x_19); -lean_dec(x_19); -lean_inc(x_16); -lean_inc(x_15); -x_21 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_21, 0, x_14); -lean_ctor_set(x_21, 1, x_15); -lean_ctor_set(x_21, 2, x_16); -lean_ctor_set(x_21, 3, x_17); -lean_ctor_set(x_21, 4, x_18); -lean_ctor_set(x_21, 5, x_20); +x_20 = lean_ctor_get(x_12, 2); +lean_inc(x_20); +x_21 = lean_ctor_get(x_12, 3); lean_inc(x_21); -x_22 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_21, x_13); -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_22, 1); -lean_inc(x_24); -lean_dec(x_22); -x_25 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_getDoSeqElems___closed__9; -lean_inc(x_2); -x_26 = lean_name_mk_string(x_2, x_25); -x_27 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__1; -lean_inc(x_2); -x_28 = lean_name_mk_string(x_2, x_27); -x_29 = l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__4___closed__7; -lean_inc(x_2); -x_30 = lean_name_mk_string(x_2, x_29); -x_31 = l_Lean_Elab_Term_Do_pullExitPointsAux___lambda__1___closed__1; -lean_inc(x_2); -x_32 = lean_name_mk_string(x_2, x_31); -x_33 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkPureUnit___closed__3; -lean_inc(x_16); -lean_inc(x_15); -x_34 = l_Lean_addMacroScope(x_15, x_33, x_16); -x_35 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkPureUnit___closed__2; -x_36 = l_Lean_Elab_Term_Do_pullExitPointsAux___lambda__1___closed__11; +x_22 = lean_ctor_get(x_12, 4); +lean_inc(x_22); +x_23 = lean_ctor_get(x_12, 5); lean_inc(x_23); -x_37 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_37, 0, x_23); -lean_ctor_set(x_37, 1, x_35); -lean_ctor_set(x_37, 2, x_34); -lean_ctor_set(x_37, 3, x_36); -x_38 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkPureUnit___closed__7; -x_39 = l_Lean_addMacroScope(x_15, x_38, x_16); -x_40 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkPureUnit___closed__6; -x_41 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkPureUnit___closed__9; -x_42 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_42, 0, x_23); -lean_ctor_set(x_42, 1, x_40); -lean_ctor_set(x_42, 2, x_39); -lean_ctor_set(x_42, 3, x_41); -x_43 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_extractBind___lambda__1___closed__3; -x_44 = lean_array_push(x_43, x_42); -x_45 = lean_box(2); -x_46 = l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__4___closed__2; -x_47 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_46); -lean_ctor_set(x_47, 2, x_44); -x_48 = l_Lean_Elab_Term_Do_pullExitPointsAux___lambda__1___closed__12; -x_49 = lean_array_push(x_48, x_37); -x_50 = lean_array_push(x_49, x_47); -x_51 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_51, 0, x_45); -lean_ctor_set(x_51, 1, x_32); -lean_ctor_set(x_51, 2, x_50); -x_52 = lean_array_push(x_43, x_51); -x_53 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_53, 0, x_45); -lean_ctor_set(x_53, 1, x_30); -lean_ctor_set(x_53, 2, x_52); -x_54 = lean_array_push(x_48, x_53); -x_55 = l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__4___closed__3; -x_56 = lean_array_push(x_54, x_55); +lean_dec(x_12); +x_24 = l_Lean_replaceRef(x_2, x_23); +lean_dec(x_23); +lean_inc(x_20); +lean_inc(x_19); +x_25 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_25, 0, x_18); +lean_ctor_set(x_25, 1, x_19); +lean_ctor_set(x_25, 2, x_20); +lean_ctor_set(x_25, 3, x_21); +lean_ctor_set(x_25, 4, x_22); +lean_ctor_set(x_25, 5, x_24); +lean_inc(x_25); +x_26 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_25, x_13); +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); lean_inc(x_28); +lean_dec(x_26); +x_29 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_getDoSeqElems___closed__9; +lean_inc(x_1); +x_30 = lean_name_mk_string(x_1, x_29); +x_31 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__2; +lean_inc(x_1); +x_32 = lean_name_mk_string(x_1, x_31); +x_33 = l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__4___closed__7; +lean_inc(x_1); +x_34 = lean_name_mk_string(x_1, x_33); +x_35 = l_Lean_Elab_Term_Do_pullExitPointsAux___lambda__1___closed__1; +lean_inc(x_1); +x_36 = lean_name_mk_string(x_1, x_35); +x_37 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkPureUnit___closed__3; +lean_inc(x_20); +lean_inc(x_19); +x_38 = l_Lean_addMacroScope(x_19, x_37, x_20); +x_39 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkPureUnit___closed__2; +x_40 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__4; +lean_inc(x_27); +x_41 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_41, 0, x_27); +lean_ctor_set(x_41, 1, x_39); +lean_ctor_set(x_41, 2, x_38); +lean_ctor_set(x_41, 3, x_40); +x_42 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkPureUnit___closed__7; +x_43 = l_Lean_addMacroScope(x_19, x_42, x_20); +x_44 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkPureUnit___closed__6; +x_45 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkPureUnit___closed__9; +x_46 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_46, 0, x_27); +lean_ctor_set(x_46, 1, x_44); +lean_ctor_set(x_46, 2, x_43); +lean_ctor_set(x_46, 3, x_45); +x_47 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_extractBind___lambda__1___closed__3; +x_48 = lean_array_push(x_47, x_46); +x_49 = lean_box(2); +x_50 = l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__4___closed__2; +x_51 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +lean_ctor_set(x_51, 2, x_48); +x_52 = l_Lean_Elab_Term_Do_pullExitPointsAux___lambda__1___closed__12; +x_53 = lean_array_push(x_52, x_41); +x_54 = lean_array_push(x_53, x_51); +x_55 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_55, 0, x_49); +lean_ctor_set(x_55, 1, x_36); +lean_ctor_set(x_55, 2, x_54); +x_56 = lean_array_push(x_47, x_55); x_57 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_57, 0, x_45); -lean_ctor_set(x_57, 1, x_28); +lean_ctor_set(x_57, 0, x_49); +lean_ctor_set(x_57, 1, x_34); lean_ctor_set(x_57, 2, x_56); -x_58 = lean_array_push(x_43, x_57); -x_59 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_59, 0, x_45); -lean_ctor_set(x_59, 1, x_46); -lean_ctor_set(x_59, 2, x_58); -x_60 = lean_array_push(x_43, x_59); -lean_inc(x_26); +x_58 = lean_array_push(x_52, x_57); +x_59 = l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__4___closed__3; +x_60 = lean_array_push(x_58, x_59); +lean_inc(x_32); x_61 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_61, 0, x_45); -lean_ctor_set(x_61, 1, x_26); +lean_ctor_set(x_61, 0, x_49); +lean_ctor_set(x_61, 1, x_32); lean_ctor_set(x_61, 2, x_60); -x_62 = l_Array_reverse___rarg(x_8); -x_63 = lean_array_push(x_62, x_3); -x_64 = l_Array_reverse___rarg(x_9); -x_65 = lean_array_push(x_64, x_4); -x_66 = l_Array_reverse___rarg(x_10); -x_67 = lean_array_push(x_66, x_5); -x_68 = l_Array_zip___rarg(x_65, x_67); +x_62 = lean_array_push(x_47, x_61); +x_63 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_63, 0, x_49); +lean_ctor_set(x_63, 1, x_50); +lean_ctor_set(x_63, 2, x_62); +x_64 = lean_array_push(x_47, x_63); +lean_inc(x_30); +x_65 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_65, 0, x_49); +lean_ctor_set(x_65, 1, x_30); +lean_ctor_set(x_65, 2, x_64); +x_66 = l_Array_reverse___rarg(x_8); +x_67 = lean_array_push(x_66, x_3); +x_68 = l_Array_reverse___rarg(x_9); +x_69 = lean_array_push(x_68, x_4); +x_70 = l_Array_reverse___rarg(x_10); +x_71 = lean_array_push(x_70, x_5); +x_72 = l_Array_zip___rarg(x_69, x_71); +lean_dec(x_71); +lean_dec(x_69); +x_73 = l_Array_zip___rarg(x_67, x_72); +lean_dec(x_72); lean_dec(x_67); -lean_dec(x_65); -x_69 = l_Array_zip___rarg(x_63, x_68); -lean_dec(x_68); -lean_dec(x_63); -x_70 = lean_array_get_size(x_69); -x_71 = lean_usize_of_nat(x_70); -lean_dec(x_70); -x_72 = 0; +x_74 = lean_array_get_size(x_73); +x_75 = lean_usize_of_nat(x_74); +lean_dec(x_74); +x_76 = 0; if (lean_obj_tag(x_11) == 0) { -uint8_t x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_73 = 1; -x_74 = lean_box(x_73); -x_75 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_75, 0, x_61); -lean_ctor_set(x_75, 1, x_74); -x_76 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__8(x_2, x_6, x_7, x_26, x_46, x_28, x_43, x_48, x_55, x_69, x_71, x_72, x_75, x_21, x_24); -lean_dec(x_69); -x_77 = lean_ctor_get(x_76, 0); -lean_inc(x_77); -x_78 = lean_ctor_get(x_76, 1); -lean_inc(x_78); -lean_dec(x_76); -x_79 = lean_ctor_get(x_77, 0); -lean_inc(x_79); -lean_dec(x_77); -x_80 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_80, 0, x_79); -x_81 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_81, 0, x_80); -lean_ctor_set(x_81, 1, x_78); -return x_81; +uint8_t x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_77 = 1; +x_78 = lean_box(x_77); +x_79 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_79, 0, x_65); +lean_ctor_set(x_79, 1, x_78); +x_80 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__8(x_1, x_6, x_7, x_16, x_17, x_30, x_50, x_32, x_47, x_52, x_59, x_73, x_75, x_76, x_79, x_25, x_28); +lean_dec(x_73); +lean_dec(x_17); +x_81 = lean_ctor_get(x_80, 0); +lean_inc(x_81); +x_82 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +lean_dec(x_80); +x_83 = lean_ctor_get(x_81, 0); +lean_inc(x_83); +lean_dec(x_81); +x_84 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_84, 0, x_83); +x_85 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_85, 0, x_84); +lean_ctor_set(x_85, 1, x_82); +return x_85; } else { -lean_object* x_82; uint8_t x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; -lean_dec(x_61); -x_82 = lean_ctor_get(x_11, 0); -x_83 = 1; -x_84 = lean_box(x_83); -lean_inc(x_82); -x_85 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_85, 0, x_82); -lean_ctor_set(x_85, 1, x_84); -x_86 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__9(x_2, x_6, x_7, x_26, x_46, x_28, x_43, x_48, x_55, x_69, x_71, x_72, x_85, x_21, x_24); -lean_dec(x_69); -x_87 = lean_ctor_get(x_86, 0); -lean_inc(x_87); -x_88 = lean_ctor_get(x_86, 1); -lean_inc(x_88); -lean_dec(x_86); -x_89 = lean_ctor_get(x_87, 0); -lean_inc(x_89); -lean_dec(x_87); -x_90 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_90, 0, x_89); -x_91 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_91, 0, x_90); -lean_ctor_set(x_91, 1, x_88); -return x_91; +lean_object* x_86; uint8_t x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; +lean_dec(x_65); +x_86 = lean_ctor_get(x_11, 0); +x_87 = 1; +x_88 = lean_box(x_87); +lean_inc(x_86); +x_89 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_89, 0, x_86); +lean_ctor_set(x_89, 1, x_88); +x_90 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__9(x_1, x_6, x_7, x_16, x_17, x_30, x_50, x_32, x_47, x_52, x_59, x_73, x_75, x_76, x_89, x_25, x_28); +lean_dec(x_73); +lean_dec(x_17); +x_91 = lean_ctor_get(x_90, 0); +lean_inc(x_91); +x_92 = lean_ctor_get(x_90, 1); +lean_inc(x_92); +lean_dec(x_90); +x_93 = lean_ctor_get(x_91, 0); +lean_inc(x_93); +lean_dec(x_91); +x_94 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_94, 0, x_93); +x_95 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_95, 0, x_94); +lean_ctor_set(x_95, 1, x_92); +return x_95; } } } @@ -15556,14 +15617,13 @@ x_31 = l_Lean_Syntax_getArg(x_1, x_30); x_32 = l_Lean_Syntax_isNone(x_31); if (x_32 == 0) { -lean_object* x_33; lean_object* x_34; uint8_t x_35; -x_33 = l_Lean_nullKind; -x_34 = lean_unsigned_to_nat(2u); +lean_object* x_33; uint8_t x_34; +x_33 = lean_unsigned_to_nat(2u); lean_inc(x_31); -x_35 = l_Lean_Syntax_isNodeOf(x_31, x_33, x_34); -if (x_35 == 0) +x_34 = l_Lean_Syntax_matchesNull(x_31, x_33); +if (x_34 == 0) { -lean_object* x_36; lean_object* x_37; +lean_object* x_35; lean_object* x_36; lean_dec(x_31); lean_dec(x_29); lean_dec(x_28); @@ -15573,251 +15633,250 @@ lean_dec(x_11); lean_dec(x_9); lean_dec(x_2); lean_dec(x_1); -x_36 = lean_box(0); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_3); -return x_37; +x_35 = lean_box(0); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_3); +return x_36; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_38 = l_Lean_Syntax_getArg(x_31, x_10); +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_37 = l_Lean_Syntax_getArg(x_31, x_10); lean_dec(x_31); -x_39 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_39, 0, x_38); -x_40 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_getDoSeqElems___closed__6; -x_41 = lean_box(0); -x_42 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2(x_1, x_40, x_29, x_9, x_28, x_11, x_27, x_15, x_4, x_12, x_26, x_41, x_39, x_2, x_3); -lean_dec(x_39); +x_38 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_38, 0, x_37); +x_39 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_getDoSeqElems___closed__6; +x_40 = lean_box(0); +x_41 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2(x_39, x_1, x_29, x_9, x_28, x_11, x_27, x_15, x_4, x_12, x_26, x_40, x_38, x_2, x_3); +lean_dec(x_38); lean_dec(x_1); -return x_42; +return x_41; } } else { -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_dec(x_31); -x_43 = lean_box(0); -x_44 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_getDoSeqElems___closed__6; -x_45 = lean_box(0); -x_46 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2(x_1, x_44, x_29, x_9, x_28, x_11, x_27, x_15, x_4, x_12, x_26, x_45, x_43, x_2, x_3); +x_42 = lean_box(0); +x_43 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_getDoSeqElems___closed__6; +x_44 = lean_box(0); +x_45 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2(x_43, x_1, x_29, x_9, x_28, x_11, x_27, x_15, x_4, x_12, x_26, x_44, x_42, x_2, x_3); lean_dec(x_1); -return x_46; +return x_45; } } } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; -x_47 = lean_unsigned_to_nat(3u); -x_48 = l_Lean_Syntax_getArg(x_1, x_47); -x_49 = lean_unsigned_to_nat(4u); -x_50 = l_Lean_Syntax_getArg(x_1, x_49); -x_51 = l_Lean_nullKind; -lean_inc(x_50); -x_52 = l_Lean_Syntax_isNodeOf(x_50, x_51, x_8); -if (x_52 == 0) +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; +x_46 = lean_unsigned_to_nat(3u); +x_47 = l_Lean_Syntax_getArg(x_1, x_46); +x_48 = lean_unsigned_to_nat(4u); +x_49 = l_Lean_Syntax_getArg(x_1, x_48); +lean_inc(x_49); +x_50 = l_Lean_Syntax_matchesNull(x_49, x_8); +if (x_50 == 0) { -lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_53 = l_Lean_Syntax_getArgs(x_50); -lean_dec(x_50); -x_54 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___closed__5; -x_55 = l_Array_sequenceMap___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__1(x_53, x_54); -lean_dec(x_53); -if (lean_obj_tag(x_55) == 0) +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = l_Lean_Syntax_getArgs(x_49); +lean_dec(x_49); +x_52 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___closed__5; +x_53 = l_Array_sequenceMap___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__1(x_51, x_52); +lean_dec(x_51); +if (lean_obj_tag(x_53) == 0) { -lean_object* x_56; lean_object* x_57; -lean_dec(x_48); +lean_object* x_54; lean_object* x_55; +lean_dec(x_47); lean_dec(x_11); lean_dec(x_9); lean_dec(x_2); lean_dec(x_1); -x_56 = lean_box(0); -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_3); -return x_57; +x_54 = lean_box(0); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_3); +return x_55; } else { -lean_object* x_58; lean_object* x_59; size_t x_60; size_t x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; -x_58 = lean_ctor_get(x_55, 0); -lean_inc(x_58); -lean_dec(x_55); -x_59 = lean_array_get_size(x_58); -x_60 = lean_usize_of_nat(x_59); -lean_dec(x_59); -x_61 = 0; -lean_inc(x_58); -x_62 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__3(x_60, x_61, x_58); -lean_inc(x_58); -x_63 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__4(x_60, x_61, x_58); -x_64 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__5(x_60, x_61, x_58); -x_65 = lean_unsigned_to_nat(5u); -x_66 = l_Lean_Syntax_getArg(x_1, x_65); -x_67 = l_Lean_Syntax_isNone(x_66); +lean_object* x_56; lean_object* x_57; size_t x_58; size_t x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; uint8_t x_65; +x_56 = lean_ctor_get(x_53, 0); +lean_inc(x_56); +lean_dec(x_53); +x_57 = lean_array_get_size(x_56); +x_58 = lean_usize_of_nat(x_57); +lean_dec(x_57); +x_59 = 0; +lean_inc(x_56); +x_60 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__3(x_58, x_59, x_56); +lean_inc(x_56); +x_61 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__4(x_58, x_59, x_56); +x_62 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__5(x_58, x_59, x_56); +x_63 = lean_unsigned_to_nat(5u); +x_64 = l_Lean_Syntax_getArg(x_1, x_63); +x_65 = l_Lean_Syntax_isNone(x_64); +if (x_65 == 0) +{ +lean_object* x_66; uint8_t x_67; +x_66 = lean_unsigned_to_nat(2u); +lean_inc(x_64); +x_67 = l_Lean_Syntax_matchesNull(x_64, x_66); if (x_67 == 0) { -lean_object* x_68; uint8_t x_69; -x_68 = lean_unsigned_to_nat(2u); -lean_inc(x_66); -x_69 = l_Lean_Syntax_isNodeOf(x_66, x_51, x_68); -if (x_69 == 0) -{ -lean_object* x_70; lean_object* x_71; -lean_dec(x_66); +lean_object* x_68; lean_object* x_69; lean_dec(x_64); -lean_dec(x_63); lean_dec(x_62); -lean_dec(x_48); +lean_dec(x_61); +lean_dec(x_60); +lean_dec(x_47); lean_dec(x_11); lean_dec(x_9); lean_dec(x_2); lean_dec(x_1); -x_70 = lean_box(0); -x_71 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_3); -return x_71; +x_68 = lean_box(0); +x_69 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_69, 0, x_68); +lean_ctor_set(x_69, 1, x_3); +return x_69; } else { -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_72 = l_Lean_Syntax_getArg(x_66, x_10); -lean_dec(x_66); -x_73 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_73, 0, x_72); -x_74 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_getDoSeqElems___closed__6; -x_75 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__3(x_1, x_74, x_9, x_11, x_48, x_4, x_12, x_64, x_63, x_62, x_73, x_2, x_3); -lean_dec(x_73); +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_70 = l_Lean_Syntax_getArg(x_64, x_10); +lean_dec(x_64); +x_71 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_71, 0, x_70); +x_72 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_getDoSeqElems___closed__6; +x_73 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__3(x_72, x_1, x_9, x_11, x_47, x_4, x_12, x_62, x_61, x_60, x_71, x_2, x_3); +lean_dec(x_71); lean_dec(x_1); -return x_75; +return x_73; } } else { -lean_object* x_76; lean_object* x_77; lean_object* x_78; -lean_dec(x_66); -x_76 = lean_box(0); -x_77 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_getDoSeqElems___closed__6; -x_78 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__3(x_1, x_77, x_9, x_11, x_48, x_4, x_12, x_64, x_63, x_62, x_76, x_2, x_3); +lean_object* x_74; lean_object* x_75; lean_object* x_76; +lean_dec(x_64); +x_74 = lean_box(0); +x_75 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_getDoSeqElems___closed__6; +x_76 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__3(x_75, x_1, x_9, x_11, x_47, x_4, x_12, x_62, x_61, x_60, x_74, x_2, x_3); lean_dec(x_1); -return x_78; +return x_76; } } } else { -lean_object* x_79; lean_object* x_80; lean_object* x_81; uint8_t x_82; -x_79 = lean_unsigned_to_nat(5u); -x_80 = l_Lean_Syntax_getArg(x_1, x_79); -x_81 = lean_unsigned_to_nat(2u); -lean_inc(x_80); -x_82 = l_Lean_Syntax_isNodeOf(x_80, x_51, x_81); -if (x_82 == 0) +lean_object* x_77; lean_object* x_78; lean_object* x_79; uint8_t x_80; +x_77 = lean_unsigned_to_nat(5u); +x_78 = l_Lean_Syntax_getArg(x_1, x_77); +x_79 = lean_unsigned_to_nat(2u); +lean_inc(x_78); +x_80 = l_Lean_Syntax_matchesNull(x_78, x_79); +if (x_80 == 0) { -lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_83 = l_Lean_Syntax_getArgs(x_50); -lean_dec(x_50); -x_84 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___closed__5; -x_85 = l_Array_sequenceMap___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__1(x_83, x_84); -lean_dec(x_83); -if (lean_obj_tag(x_85) == 0) +lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_81 = l_Lean_Syntax_getArgs(x_49); +lean_dec(x_49); +x_82 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___closed__5; +x_83 = l_Array_sequenceMap___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__1(x_81, x_82); +lean_dec(x_81); +if (lean_obj_tag(x_83) == 0) { -lean_object* x_86; lean_object* x_87; -lean_dec(x_80); -lean_dec(x_48); +lean_object* x_84; lean_object* x_85; +lean_dec(x_78); +lean_dec(x_47); lean_dec(x_11); lean_dec(x_9); lean_dec(x_2); lean_dec(x_1); -x_86 = lean_box(0); -x_87 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_87, 0, x_86); -lean_ctor_set(x_87, 1, x_3); -return x_87; +x_84 = lean_box(0); +x_85 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_85, 0, x_84); +lean_ctor_set(x_85, 1, x_3); +return x_85; } else { -lean_object* x_88; lean_object* x_89; size_t x_90; size_t x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; uint8_t x_95; -x_88 = lean_ctor_get(x_85, 0); -lean_inc(x_88); -lean_dec(x_85); -x_89 = lean_array_get_size(x_88); -x_90 = lean_usize_of_nat(x_89); -lean_dec(x_89); -x_91 = 0; -lean_inc(x_88); -x_92 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__3(x_90, x_91, x_88); -lean_inc(x_88); -x_93 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__4(x_90, x_91, x_88); -x_94 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__5(x_90, x_91, x_88); -x_95 = l_Lean_Syntax_isNone(x_80); -if (x_95 == 0) +lean_object* x_86; lean_object* x_87; size_t x_88; size_t x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; uint8_t x_93; +x_86 = lean_ctor_get(x_83, 0); +lean_inc(x_86); +lean_dec(x_83); +x_87 = lean_array_get_size(x_86); +x_88 = lean_usize_of_nat(x_87); +lean_dec(x_87); +x_89 = 0; +lean_inc(x_86); +x_90 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__3(x_88, x_89, x_86); +lean_inc(x_86); +x_91 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__4(x_88, x_89, x_86); +x_92 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__5(x_88, x_89, x_86); +x_93 = l_Lean_Syntax_isNone(x_78); +if (x_93 == 0) { -uint8_t x_96; -lean_inc(x_80); -x_96 = l_Lean_Syntax_isNodeOf(x_80, x_51, x_81); -if (x_96 == 0) +uint8_t x_94; +lean_inc(x_78); +x_94 = l_Lean_Syntax_matchesNull(x_78, x_79); +if (x_94 == 0) { -lean_object* x_97; lean_object* x_98; -lean_dec(x_94); -lean_dec(x_93); +lean_object* x_95; lean_object* x_96; lean_dec(x_92); -lean_dec(x_80); -lean_dec(x_48); +lean_dec(x_91); +lean_dec(x_90); +lean_dec(x_78); +lean_dec(x_47); lean_dec(x_11); lean_dec(x_9); lean_dec(x_2); lean_dec(x_1); -x_97 = lean_box(0); -x_98 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_98, 0, x_97); -lean_ctor_set(x_98, 1, x_3); -return x_98; +x_95 = lean_box(0); +x_96 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_96, 0, x_95); +lean_ctor_set(x_96, 1, x_3); +return x_96; } else { -lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; -x_99 = l_Lean_Syntax_getArg(x_80, x_10); -lean_dec(x_80); -x_100 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_100, 0, x_99); -x_101 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_getDoSeqElems___closed__6; -x_102 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__3(x_1, x_101, x_9, x_11, x_48, x_4, x_12, x_94, x_93, x_92, x_100, x_2, x_3); -lean_dec(x_100); +lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; +x_97 = l_Lean_Syntax_getArg(x_78, x_10); +lean_dec(x_78); +x_98 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_98, 0, x_97); +x_99 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_getDoSeqElems___closed__6; +x_100 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__3(x_99, x_1, x_9, x_11, x_47, x_4, x_12, x_92, x_91, x_90, x_98, x_2, x_3); +lean_dec(x_98); lean_dec(x_1); -return x_102; +return x_100; } } else { -lean_object* x_103; lean_object* x_104; lean_object* x_105; -lean_dec(x_80); -x_103 = lean_box(0); -x_104 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_getDoSeqElems___closed__6; -x_105 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__3(x_1, x_104, x_9, x_11, x_48, x_4, x_12, x_94, x_93, x_92, x_103, x_2, x_3); +lean_object* x_101; lean_object* x_102; lean_object* x_103; +lean_dec(x_78); +x_101 = lean_box(0); +x_102 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_getDoSeqElems___closed__6; +x_103 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__3(x_102, x_1, x_9, x_11, x_47, x_4, x_12, x_92, x_91, x_90, x_101, x_2, x_3); lean_dec(x_1); -return x_105; +return x_103; } } } else { -lean_object* x_106; lean_object* x_107; -lean_dec(x_80); -lean_dec(x_50); -lean_dec(x_48); +lean_object* x_104; lean_object* x_105; +lean_dec(x_78); +lean_dec(x_49); +lean_dec(x_47); lean_dec(x_11); lean_dec(x_9); lean_dec(x_2); lean_dec(x_1); -x_106 = lean_box(0); -x_107 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_107, 0, x_106); -lean_ctor_set(x_107, 1, x_3); -return x_107; +x_104 = lean_box(0); +x_105 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_105, 0, x_104); +lean_ctor_set(x_105, 1, x_3); +return x_105; } } } @@ -15891,56 +15950,132 @@ lean_dec(x_4); return x_17; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__6___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_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) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__6___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; _start: { -size_t x_16; size_t x_17; lean_object* x_18; -x_16 = lean_unbox_usize(x_11); -lean_dec(x_11); -x_17 = lean_unbox_usize(x_12); +size_t x_18; size_t x_19; lean_object* x_20; +x_18 = lean_unbox_usize(x_13); +lean_dec(x_13); +x_19 = lean_unbox_usize(x_14); +lean_dec(x_14); +x_20 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_18, x_19, x_15, x_16, x_17); lean_dec(x_12); -x_18 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_16, x_17, x_13, x_14, x_15); -lean_dec(x_10); -return x_18; +lean_dec(x_5); +lean_dec(x_4); +return x_20; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__7___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_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) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__7___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; _start: { -size_t x_16; size_t x_17; lean_object* x_18; -x_16 = lean_unbox_usize(x_11); -lean_dec(x_11); -x_17 = lean_unbox_usize(x_12); +size_t x_18; size_t x_19; lean_object* x_20; +x_18 = lean_unbox_usize(x_13); +lean_dec(x_13); +x_19 = lean_unbox_usize(x_14); +lean_dec(x_14); +x_20 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_18, x_19, x_15, x_16, x_17); lean_dec(x_12); -x_18 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_16, x_17, x_13, x_14, x_15); -lean_dec(x_10); -return x_18; +lean_dec(x_5); +lean_dec(x_4); +return x_20; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__8___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_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) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__8___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; _start: { -size_t x_16; size_t x_17; lean_object* x_18; -x_16 = lean_unbox_usize(x_11); -lean_dec(x_11); -x_17 = lean_unbox_usize(x_12); +size_t x_18; size_t x_19; lean_object* x_20; +x_18 = lean_unbox_usize(x_13); +lean_dec(x_13); +x_19 = lean_unbox_usize(x_14); +lean_dec(x_14); +x_20 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_18, x_19, x_15, x_16, x_17); lean_dec(x_12); -x_18 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_16, x_17, x_13, x_14, x_15); -lean_dec(x_10); -return x_18; +lean_dec(x_5); +lean_dec(x_4); +return x_20; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__9___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_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) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__9___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; _start: { -size_t x_16; size_t x_17; lean_object* x_18; -x_16 = lean_unbox_usize(x_11); -lean_dec(x_11); -x_17 = lean_unbox_usize(x_12); +size_t x_18; size_t x_19; lean_object* x_20; +x_18 = lean_unbox_usize(x_13); +lean_dec(x_13); +x_19 = lean_unbox_usize(x_14); +lean_dec(x_14); +x_20 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_18, x_19, x_15, x_16, x_17); lean_dec(x_12); -x_18 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_16, x_17, x_13, x_14, x_15); -lean_dec(x_10); -return x_18; +lean_dec(x_5); +lean_dec(x_4); +return x_20; } } LEAN_EXPORT lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___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, lean_object* x_7, 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, lean_object* x_15) { @@ -15952,7 +16087,7 @@ lean_dec(x_11); x_17 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_16, x_12, x_13, x_14, x_15); lean_dec(x_13); lean_dec(x_12); -lean_dec(x_1); +lean_dec(x_2); return x_17; } } @@ -15962,7 +16097,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoI lean_object* x_14; x_14 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_11); -lean_dec(x_1); +lean_dec(x_2); return x_14; } } @@ -16343,7 +16478,7 @@ static lean_object* _init_l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destruct 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___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__1; x_2 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__2; -x_3 = lean_unsigned_to_nat(761u); +x_3 = lean_unsigned_to_nat(762u); x_4 = lean_unsigned_to_nat(13u); x_5 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -17585,7 +17720,7 @@ static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_returnToTerm___closed__22() 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___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__1; x_2 = l_Lean_Elab_Term_Do_ToTerm_returnToTerm___closed__21; -x_3 = lean_unsigned_to_nat(879u); +x_3 = lean_unsigned_to_nat(880u); x_4 = lean_unsigned_to_nat(28u); x_5 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -18637,7 +18772,7 @@ static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_continueToTerm___closed__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; x_1 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__1; x_2 = l_Lean_Elab_Term_Do_ToTerm_continueToTerm___closed__1; -x_3 = lean_unsigned_to_nat(888u); +x_3 = lean_unsigned_to_nat(889u); x_4 = lean_unsigned_to_nat(28u); x_5 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -18889,7 +19024,7 @@ static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_continueToTerm___closed__26 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___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__1; x_2 = l_Lean_Elab_Term_Do_ToTerm_continueToTerm___closed__1; -x_3 = lean_unsigned_to_nat(892u); +x_3 = lean_unsigned_to_nat(893u); x_4 = lean_unsigned_to_nat(28u); x_5 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -19622,7 +19757,7 @@ static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_breakToTerm___closed__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; x_1 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__1; x_2 = l_Lean_Elab_Term_Do_ToTerm_breakToTerm___closed__1; -x_3 = lean_unsigned_to_nat(900u); +x_3 = lean_unsigned_to_nat(901u); x_4 = lean_unsigned_to_nat(28u); x_5 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -19708,7 +19843,7 @@ static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_breakToTerm___closed__10() 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___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__1; x_2 = l_Lean_Elab_Term_Do_ToTerm_breakToTerm___closed__1; -x_3 = lean_unsigned_to_nat(904u); +x_3 = lean_unsigned_to_nat(905u); x_4 = lean_unsigned_to_nat(28u); x_5 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -20522,37 +20657,13 @@ static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_Do_pullExitPointsAux___lambda__1___closed__9; -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_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__11() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__10; -x_3 = lean_alloc_ctor(1, 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_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__12() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_getDoSeqElems___closed__6; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__6___lambda__1___closed__14; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__13() { +static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -20564,19 +20675,19 @@ lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__14() { +static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__13; +x_2 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__11; x_3 = lean_alloc_ctor(1, 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_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__15() { +static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__13() { _start: { lean_object* x_1; @@ -20584,20 +20695,20 @@ x_1 = lean_mk_string_from_bytes("Lean.Elab.Term.Do.ToTerm.actionTerminalToTerm", return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__16() { +static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__14() { _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___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__1; -x_2 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__15; -x_3 = lean_unsigned_to_nat(915u); +x_2 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__13; +x_3 = lean_unsigned_to_nat(916u); x_4 = lean_unsigned_to_nat(28u); x_5 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); return x_6; } } -static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__17() { +static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__15() { _start: { lean_object* x_1; @@ -20605,22 +20716,22 @@ x_1 = lean_mk_string_from_bytes("DoResultPR.«pure»", 19); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__18() { +static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__16() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__17; +x_1 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__15; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__19() { +static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__17; +x_1 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__15; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__18; +x_3 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__16; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -20628,7 +20739,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__20() { +static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -20638,31 +20749,31 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__21() { +static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__20; +x_2 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__18; 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_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__22() { +static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__21; +x_2 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__19; x_3 = lean_alloc_ctor(1, 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_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__23() { +static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__21() { _start: { lean_object* x_1; @@ -20670,22 +20781,22 @@ x_1 = lean_mk_string_from_bytes("DoResultPRBC.«pure»", 21); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__24() { +static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__22() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__23; +x_1 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__21; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__25() { +static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__23; +x_1 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__21; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__24; +x_3 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__22; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -20693,7 +20804,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__26() { +static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -20703,24 +20814,24 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__27() { +static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__26; +x_2 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__24; 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_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__28() { +static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__27; +x_2 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__25; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); @@ -20838,7 +20949,7 @@ lean_inc(x_11); lean_inc(x_6); x_46 = l_Lean_addMacroScope(x_6, x_45, x_11); x_47 = l_Lean_Elab_Term_Do_pullExitPointsAux___lambda__1___closed__5; -x_48 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__11; +x_48 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__4; lean_inc(x_24); x_49 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_49, 0, x_24); @@ -20995,7 +21106,7 @@ lean_ctor_set(x_114, 1, x_113); x_115 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_extractBind___lambda__1___closed__3; x_116 = lean_array_push(x_115, x_114); x_117 = lean_box(2); -x_118 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__12; +x_118 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__10; x_119 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_119, 0, x_117); lean_ctor_set(x_119, 1, x_118); @@ -21010,7 +21121,7 @@ lean_inc(x_11); lean_inc(x_6); x_123 = l_Lean_addMacroScope(x_6, x_122, x_11); x_124 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkUnit___closed__13; -x_125 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__14; +x_125 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__12; lean_inc(x_102); x_126 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_126, 0, x_102); @@ -21068,7 +21179,7 @@ lean_inc(x_11); lean_inc(x_6); x_151 = l_Lean_addMacroScope(x_6, x_150, x_11); x_152 = l_Lean_Elab_Term_Do_pullExitPointsAux___lambda__1___closed__5; -x_153 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__11; +x_153 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__4; lean_inc(x_102); x_154 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_154, 0, x_102); @@ -21197,7 +21308,7 @@ lean_ctor_set(x_209, 1, x_208); x_210 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_extractBind___lambda__1___closed__3; x_211 = lean_array_push(x_210, x_209); x_212 = lean_box(2); -x_213 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__12; +x_213 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__10; x_214 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_214, 0, x_212); lean_ctor_set(x_214, 1, x_213); @@ -21212,7 +21323,7 @@ lean_inc(x_11); lean_inc(x_6); x_218 = l_Lean_addMacroScope(x_6, x_217, x_11); x_219 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkUnit___closed__13; -x_220 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__14; +x_220 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__12; lean_inc(x_197); x_221 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_221, 0, x_197); @@ -21270,7 +21381,7 @@ lean_inc(x_11); lean_inc(x_6); x_246 = l_Lean_addMacroScope(x_6, x_245, x_11); x_247 = l_Lean_Elab_Term_Do_pullExitPointsAux___lambda__1___closed__5; -x_248 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__11; +x_248 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__4; lean_inc(x_197); x_249 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_249, 0, x_197); @@ -21413,7 +21524,7 @@ lean_dec(x_1); x_311 = lean_ctor_get(x_17, 1); lean_inc(x_311); lean_dec(x_17); -x_312 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__16; +x_312 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__14; x_313 = l_panic___at_Lean_Elab_Term_Do_ToTerm_returnToTerm___spec__2(x_312, x_2, x_16, x_311); return x_313; } @@ -21485,17 +21596,17 @@ lean_inc(x_11); lean_inc(x_6); x_341 = l_Lean_addMacroScope(x_6, x_340, x_11); x_342 = l_Lean_Elab_Term_Do_pullExitPointsAux___lambda__1___closed__5; -x_343 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__11; +x_343 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__4; lean_inc(x_317); x_344 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_344, 0, x_317); lean_ctor_set(x_344, 1, x_342); lean_ctor_set(x_344, 2, x_341); lean_ctor_set(x_344, 3, x_343); -x_345 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__20; +x_345 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__18; x_346 = l_Lean_addMacroScope(x_6, x_345, x_11); -x_347 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__19; -x_348 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__22; +x_347 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__17; +x_348 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__20; lean_inc(x_317); x_349 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_349, 0, x_317); @@ -21661,7 +21772,7 @@ lean_inc(x_11); lean_inc(x_6); x_424 = l_Lean_addMacroScope(x_6, x_423, x_11); x_425 = l_Lean_Elab_Term_Do_pullExitPointsAux___lambda__1___closed__5; -x_426 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__11; +x_426 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__4; lean_inc(x_400); x_427 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_427, 0, x_400); @@ -21837,17 +21948,17 @@ lean_inc(x_11); lean_inc(x_6); x_507 = l_Lean_addMacroScope(x_6, x_506, x_11); x_508 = l_Lean_Elab_Term_Do_pullExitPointsAux___lambda__1___closed__5; -x_509 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__11; +x_509 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__4; lean_inc(x_483); x_510 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_510, 0, x_483); lean_ctor_set(x_510, 1, x_508); lean_ctor_set(x_510, 2, x_507); lean_ctor_set(x_510, 3, x_509); -x_511 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__26; +x_511 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__24; x_512 = l_Lean_addMacroScope(x_6, x_511, x_11); -x_513 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__25; -x_514 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__28; +x_513 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__23; +x_514 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__26; lean_inc(x_483); x_515 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_515, 0, x_483); @@ -22234,7 +22345,7 @@ x_77 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_77, 0, x_66); lean_ctor_set(x_77, 1, x_76); x_78 = lean_array_push(x_38, x_77); -x_79 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__12; +x_79 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__10; x_80 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_80, 0, x_40); lean_ctor_set(x_80, 1, x_79); @@ -23063,247 +23174,245 @@ return x_28; } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; +lean_object* x_29; lean_object* x_30; uint8_t x_31; x_29 = l_Lean_Syntax_getArg(x_22, x_21); x_30 = l_Lean_Syntax_getArg(x_22, x_15); -x_31 = l_Lean_nullKind; -x_32 = l_Lean_Syntax_isNodeOf(x_30, x_31, x_21); -if (x_32 == 0) +x_31 = l_Lean_Syntax_matchesNull(x_30, x_21); +if (x_31 == 0) { -lean_object* x_33; lean_object* x_34; +lean_object* x_32; lean_object* x_33; lean_dec(x_29); lean_dec(x_22); lean_dec(x_2); -x_33 = l_Lean_Elab_Term_Do_ToTerm_reassignToTerm___closed__3; -x_34 = l_Lean_Macro_throwErrorAt___rarg(x_1, x_33, x_18, x_17); +x_32 = l_Lean_Elab_Term_Do_ToTerm_reassignToTerm___closed__3; +x_33 = l_Lean_Macro_throwErrorAt___rarg(x_1, x_32, x_18, x_17); lean_dec(x_1); -return x_34; +return x_33; } else { -lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_35 = lean_unsigned_to_nat(2u); -x_36 = l_Lean_Syntax_getArg(x_22, x_35); -x_37 = l_Lean_Syntax_isNodeOf(x_36, x_31, x_21); -if (x_37 == 0) +lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_34 = lean_unsigned_to_nat(2u); +x_35 = l_Lean_Syntax_getArg(x_22, x_34); +x_36 = l_Lean_Syntax_matchesNull(x_35, x_21); +if (x_36 == 0) { -lean_object* x_38; lean_object* x_39; +lean_object* x_37; lean_object* x_38; lean_dec(x_29); lean_dec(x_22); lean_dec(x_2); -x_38 = l_Lean_Elab_Term_Do_ToTerm_reassignToTerm___closed__3; -x_39 = l_Lean_Macro_throwErrorAt___rarg(x_1, x_38, x_18, x_17); +x_37 = l_Lean_Elab_Term_Do_ToTerm_reassignToTerm___closed__3; +x_38 = l_Lean_Macro_throwErrorAt___rarg(x_1, x_37, x_18, x_17); lean_dec(x_1); -return x_39; +return x_38; } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_dec(x_1); -x_40 = lean_unsigned_to_nat(4u); -x_41 = l_Lean_Syntax_getArg(x_22, x_40); +x_39 = lean_unsigned_to_nat(4u); +x_40 = l_Lean_Syntax_getArg(x_22, x_39); lean_dec(x_22); -x_42 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_18, x_17); -x_43 = lean_ctor_get(x_42, 0); -lean_inc(x_43); -x_44 = lean_ctor_get(x_42, 1); -lean_inc(x_44); -lean_dec(x_42); -x_45 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_liftMethodForbiddenBinder___closed__9; -lean_inc(x_43); -x_46 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_46, 0, x_43); -lean_ctor_set(x_46, 1, x_45); -x_47 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__11; -lean_inc(x_43); -x_48 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_48, 0, x_43); -lean_ctor_set(x_48, 1, x_47); -x_49 = l_Lean_Elab_Term_Do_ToTerm_reassignToTerm___closed__6; +x_41 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_18, x_17); +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_41, 1); lean_inc(x_43); -x_50 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_50, 0, x_43); -lean_ctor_set(x_50, 1, x_49); -x_51 = l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__4___closed__9; -x_52 = lean_array_push(x_51, x_50); +lean_dec(x_41); +x_44 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_liftMethodForbiddenBinder___closed__9; +lean_inc(x_42); +x_45 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_45, 0, x_42); +lean_ctor_set(x_45, 1, x_44); +x_46 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__11; +lean_inc(x_42); +x_47 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_47, 0, x_42); +lean_ctor_set(x_47, 1, x_46); +x_48 = l_Lean_Elab_Term_Do_ToTerm_reassignToTerm___closed__6; +lean_inc(x_42); +x_49 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_49, 0, x_42); +lean_ctor_set(x_49, 1, x_48); +x_50 = l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__4___closed__9; +x_51 = lean_array_push(x_50, x_49); lean_inc(x_29); -x_53 = lean_array_push(x_52, x_29); -x_54 = l_Lean_Elab_Term_Do_ToTerm_reassignToTerm___closed__8; -x_55 = lean_array_push(x_53, x_54); -x_56 = lean_array_push(x_55, x_41); -x_57 = lean_box(2); -x_58 = l_Lean_Elab_Term_Do_ToTerm_reassignToTerm___closed__5; -x_59 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_59, 0, x_57); -lean_ctor_set(x_59, 1, x_58); -lean_ctor_set(x_59, 2, x_56); -x_60 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__18; -x_61 = lean_array_push(x_60, x_29); -x_62 = l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__4___closed__3; -x_63 = lean_array_push(x_61, x_62); -x_64 = lean_array_push(x_63, x_62); -x_65 = lean_array_push(x_64, x_48); -x_66 = lean_array_push(x_65, x_59); -x_67 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_67, 0, x_57); -lean_ctor_set(x_67, 1, x_25); -lean_ctor_set(x_67, 2, x_66); -x_68 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_extractBind___lambda__1___closed__3; -x_69 = lean_array_push(x_68, x_67); -x_70 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__10; -x_71 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_71, 0, x_57); -lean_ctor_set(x_71, 1, x_70); -lean_ctor_set(x_71, 2, x_69); -x_72 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__19; -x_73 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_73, 0, x_43); -lean_ctor_set(x_73, 1, x_72); -x_74 = lean_array_push(x_51, x_46); -x_75 = lean_array_push(x_74, x_71); -x_76 = lean_array_push(x_75, x_73); -x_77 = lean_array_push(x_76, x_2); -x_78 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_liftMethodForbiddenBinder___closed__10; -x_79 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_79, 0, x_57); -lean_ctor_set(x_79, 1, x_78); -lean_ctor_set(x_79, 2, x_77); -x_80 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_80, 0, x_79); -lean_ctor_set(x_80, 1, x_44); -return x_80; +x_52 = lean_array_push(x_51, x_29); +x_53 = l_Lean_Elab_Term_Do_ToTerm_reassignToTerm___closed__8; +x_54 = lean_array_push(x_52, x_53); +x_55 = lean_array_push(x_54, x_40); +x_56 = lean_box(2); +x_57 = l_Lean_Elab_Term_Do_ToTerm_reassignToTerm___closed__5; +x_58 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_58, 0, x_56); +lean_ctor_set(x_58, 1, x_57); +lean_ctor_set(x_58, 2, x_55); +x_59 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__18; +x_60 = lean_array_push(x_59, x_29); +x_61 = l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__4___closed__3; +x_62 = lean_array_push(x_60, x_61); +x_63 = lean_array_push(x_62, x_61); +x_64 = lean_array_push(x_63, x_47); +x_65 = lean_array_push(x_64, x_58); +x_66 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_66, 0, x_56); +lean_ctor_set(x_66, 1, x_25); +lean_ctor_set(x_66, 2, x_65); +x_67 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_extractBind___lambda__1___closed__3; +x_68 = lean_array_push(x_67, x_66); +x_69 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__10; +x_70 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_70, 0, x_56); +lean_ctor_set(x_70, 1, x_69); +lean_ctor_set(x_70, 2, x_68); +x_71 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__19; +x_72 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_72, 0, x_42); +lean_ctor_set(x_72, 1, x_71); +x_73 = lean_array_push(x_50, x_45); +x_74 = lean_array_push(x_73, x_70); +x_75 = lean_array_push(x_74, x_72); +x_76 = lean_array_push(x_75, x_2); +x_77 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_liftMethodForbiddenBinder___closed__10; +x_78 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_78, 0, x_56); +lean_ctor_set(x_78, 1, x_77); +lean_ctor_set(x_78, 2, x_76); +x_79 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_79, 0, x_78); +lean_ctor_set(x_79, 1, x_43); +return x_79; } } } } else { -lean_object* x_81; lean_object* x_82; uint8_t x_83; -x_81 = l_Lean_Syntax_getArg(x_22, x_21); -x_82 = l_Lean_Elab_Term_Do_ToTerm_reassignToTerm___closed__10; -lean_inc(x_81); -x_83 = l_Lean_Syntax_isOfKind(x_81, x_82); -if (x_83 == 0) +lean_object* x_80; lean_object* x_81; uint8_t x_82; +x_80 = l_Lean_Syntax_getArg(x_22, x_21); +x_81 = l_Lean_Elab_Term_Do_ToTerm_reassignToTerm___closed__10; +lean_inc(x_80); +x_82 = l_Lean_Syntax_isOfKind(x_80, x_81); +if (x_82 == 0) { -lean_object* x_84; lean_object* x_85; -lean_dec(x_81); +lean_object* x_83; lean_object* x_84; +lean_dec(x_80); lean_dec(x_22); lean_dec(x_2); -x_84 = l_Lean_Elab_Term_Do_ToTerm_reassignToTerm___closed__3; -x_85 = l_Lean_Macro_throwErrorAt___rarg(x_1, x_84, x_18, x_17); +x_83 = l_Lean_Elab_Term_Do_ToTerm_reassignToTerm___closed__3; +x_84 = l_Lean_Macro_throwErrorAt___rarg(x_1, x_83, x_18, x_17); lean_dec(x_1); -return x_85; +return x_84; } else { -lean_object* x_86; lean_object* x_87; uint8_t x_88; -x_86 = l_Lean_Syntax_getArg(x_22, x_15); -x_87 = l_Lean_nullKind; -x_88 = l_Lean_Syntax_isNodeOf(x_86, x_87, x_21); -if (x_88 == 0) +lean_object* x_85; uint8_t x_86; +x_85 = l_Lean_Syntax_getArg(x_22, x_15); +x_86 = l_Lean_Syntax_matchesNull(x_85, x_21); +if (x_86 == 0) { -lean_object* x_89; lean_object* x_90; -lean_dec(x_81); +lean_object* x_87; lean_object* x_88; +lean_dec(x_80); lean_dec(x_22); lean_dec(x_2); -x_89 = l_Lean_Elab_Term_Do_ToTerm_reassignToTerm___closed__3; -x_90 = l_Lean_Macro_throwErrorAt___rarg(x_1, x_89, x_18, x_17); +x_87 = l_Lean_Elab_Term_Do_ToTerm_reassignToTerm___closed__3; +x_88 = l_Lean_Macro_throwErrorAt___rarg(x_1, x_87, x_18, x_17); lean_dec(x_1); -return x_90; +return x_88; } else { -lean_object* x_91; lean_object* x_92; uint8_t x_93; -x_91 = lean_unsigned_to_nat(2u); -x_92 = l_Lean_Syntax_getArg(x_22, x_91); -x_93 = l_Lean_Syntax_isNodeOf(x_92, x_87, x_21); -if (x_93 == 0) +lean_object* x_89; lean_object* x_90; uint8_t x_91; +x_89 = lean_unsigned_to_nat(2u); +x_90 = l_Lean_Syntax_getArg(x_22, x_89); +x_91 = l_Lean_Syntax_matchesNull(x_90, x_21); +if (x_91 == 0) { -lean_object* x_94; lean_object* x_95; -lean_dec(x_81); +lean_object* x_92; lean_object* x_93; +lean_dec(x_80); lean_dec(x_22); lean_dec(x_2); -x_94 = l_Lean_Elab_Term_Do_ToTerm_reassignToTerm___closed__3; -x_95 = l_Lean_Macro_throwErrorAt___rarg(x_1, x_94, x_18, x_17); +x_92 = l_Lean_Elab_Term_Do_ToTerm_reassignToTerm___closed__3; +x_93 = l_Lean_Macro_throwErrorAt___rarg(x_1, x_92, x_18, x_17); lean_dec(x_1); -return x_95; +return x_93; } else { -lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_dec(x_1); -x_96 = lean_unsigned_to_nat(4u); -x_97 = l_Lean_Syntax_getArg(x_22, x_96); +x_94 = lean_unsigned_to_nat(4u); +x_95 = l_Lean_Syntax_getArg(x_22, x_94); lean_dec(x_22); -x_98 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_18, x_17); -x_99 = lean_ctor_get(x_98, 0); -lean_inc(x_99); -x_100 = lean_ctor_get(x_98, 1); -lean_inc(x_100); -lean_dec(x_98); -x_101 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_liftMethodForbiddenBinder___closed__9; -lean_inc(x_99); +x_96 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_18, x_17); +x_97 = lean_ctor_get(x_96, 0); +lean_inc(x_97); +x_98 = lean_ctor_get(x_96, 1); +lean_inc(x_98); +lean_dec(x_96); +x_99 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_liftMethodForbiddenBinder___closed__9; +lean_inc(x_97); +x_100 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_100, 0, x_97); +lean_ctor_set(x_100, 1, x_99); +x_101 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__11; +lean_inc(x_97); x_102 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_102, 0, x_99); +lean_ctor_set(x_102, 0, x_97); lean_ctor_set(x_102, 1, x_101); -x_103 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__11; -lean_inc(x_99); +x_103 = l_Lean_Elab_Term_Do_ToTerm_reassignToTerm___closed__6; +lean_inc(x_97); x_104 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_104, 0, x_99); +lean_ctor_set(x_104, 0, x_97); lean_ctor_set(x_104, 1, x_103); -x_105 = l_Lean_Elab_Term_Do_ToTerm_reassignToTerm___closed__6; -lean_inc(x_99); -x_106 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_106, 0, x_99); -lean_ctor_set(x_106, 1, x_105); -x_107 = l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__4___closed__9; -x_108 = lean_array_push(x_107, x_106); -lean_inc(x_81); -x_109 = lean_array_push(x_108, x_81); -x_110 = l_Lean_Elab_Term_Do_ToTerm_reassignToTerm___closed__8; -x_111 = lean_array_push(x_109, x_110); -x_112 = lean_array_push(x_111, x_97); -x_113 = lean_box(2); -x_114 = l_Lean_Elab_Term_Do_ToTerm_reassignToTerm___closed__5; -x_115 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_115, 0, x_113); -lean_ctor_set(x_115, 1, x_114); -lean_ctor_set(x_115, 2, x_112); -x_116 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__18; -x_117 = lean_array_push(x_116, x_81); -x_118 = l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__4___closed__3; -x_119 = lean_array_push(x_117, x_118); -x_120 = lean_array_push(x_119, x_118); -x_121 = lean_array_push(x_120, x_104); -x_122 = lean_array_push(x_121, x_115); -x_123 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_123, 0, x_113); -lean_ctor_set(x_123, 1, x_23); -lean_ctor_set(x_123, 2, x_122); -x_124 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_extractBind___lambda__1___closed__3; -x_125 = lean_array_push(x_124, x_123); -x_126 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__10; -x_127 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_127, 0, x_113); +x_105 = l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__4___closed__9; +x_106 = lean_array_push(x_105, x_104); +lean_inc(x_80); +x_107 = lean_array_push(x_106, x_80); +x_108 = l_Lean_Elab_Term_Do_ToTerm_reassignToTerm___closed__8; +x_109 = lean_array_push(x_107, x_108); +x_110 = lean_array_push(x_109, x_95); +x_111 = lean_box(2); +x_112 = l_Lean_Elab_Term_Do_ToTerm_reassignToTerm___closed__5; +x_113 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_113, 0, x_111); +lean_ctor_set(x_113, 1, x_112); +lean_ctor_set(x_113, 2, x_110); +x_114 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__18; +x_115 = lean_array_push(x_114, x_80); +x_116 = l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__4___closed__3; +x_117 = lean_array_push(x_115, x_116); +x_118 = lean_array_push(x_117, x_116); +x_119 = lean_array_push(x_118, x_102); +x_120 = lean_array_push(x_119, x_113); +x_121 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_121, 0, x_111); +lean_ctor_set(x_121, 1, x_23); +lean_ctor_set(x_121, 2, x_120); +x_122 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_extractBind___lambda__1___closed__3; +x_123 = lean_array_push(x_122, x_121); +x_124 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__10; +x_125 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_125, 0, x_111); +lean_ctor_set(x_125, 1, x_124); +lean_ctor_set(x_125, 2, x_123); +x_126 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__19; +x_127 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_127, 0, x_97); lean_ctor_set(x_127, 1, x_126); -lean_ctor_set(x_127, 2, x_125); -x_128 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__19; -x_129 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_129, 0, x_99); -lean_ctor_set(x_129, 1, x_128); -x_130 = lean_array_push(x_107, x_102); -x_131 = lean_array_push(x_130, x_127); -x_132 = lean_array_push(x_131, x_129); -x_133 = lean_array_push(x_132, x_2); -x_134 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_liftMethodForbiddenBinder___closed__10; -x_135 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_135, 0, x_113); -lean_ctor_set(x_135, 1, x_134); -lean_ctor_set(x_135, 2, x_133); -x_136 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_136, 0, x_135); -lean_ctor_set(x_136, 1, x_100); -return x_136; +x_128 = lean_array_push(x_105, x_100); +x_129 = lean_array_push(x_128, x_125); +x_130 = lean_array_push(x_129, x_127); +x_131 = lean_array_push(x_130, x_2); +x_132 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_liftMethodForbiddenBinder___closed__10; +x_133 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_133, 0, x_111); +lean_ctor_set(x_133, 1, x_132); +lean_ctor_set(x_133, 2, x_131); +x_134 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_134, 0, x_133); +lean_ctor_set(x_134, 1, x_98); +return x_134; } } } @@ -23541,7 +23650,7 @@ lean_ctor_set(x_18, 1, x_17); x_19 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_extractBind___lambda__1___closed__3; x_20 = lean_array_push(x_19, x_18); x_21 = lean_box(2); -x_22 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__12; +x_22 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__10; x_23 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_23, 0, x_21); lean_ctor_set(x_23, 1, x_22); @@ -23852,7 +23961,7 @@ lean_ctor_set(x_49, 0, x_31); lean_ctor_set(x_49, 1, x_48); x_50 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_extractBind___lambda__1___closed__3; x_51 = lean_array_push(x_50, x_49); -x_52 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__12; +x_52 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__10; x_53 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_53, 0, x_43); lean_ctor_set(x_53, 1, x_52); @@ -24042,7 +24151,7 @@ lean_dec(x_10); x_19 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_extractBind___lambda__1___closed__3; x_20 = lean_array_push(x_19, x_18); x_21 = lean_box(2); -x_22 = l_Lean_nullKind; +x_22 = l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__4___closed__2; x_23 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_23, 0, x_21); lean_ctor_set(x_23, 1, x_22); @@ -24531,7 +24640,7 @@ x_119 = lean_ctor_get(x_117, 1); lean_inc(x_119); lean_dec(x_117); x_120 = lean_box(2); -x_121 = l_Lean_nullKind; +x_121 = l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__4___closed__2; x_122 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_122, 0, x_120); lean_ctor_set(x_122, 1, x_121); @@ -24663,7 +24772,7 @@ static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_mkNestedKind___closed__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; x_1 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__1; x_2 = l_Lean_Elab_Term_Do_ToTerm_mkNestedKind___closed__1; -x_3 = lean_unsigned_to_nat(1048u); +x_3 = lean_unsigned_to_nat(1049u); x_4 = lean_unsigned_to_nat(27u); x_5 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -24832,7 +24941,7 @@ static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_matchNestedTermResult___clo 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___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__1; x_2 = l_Lean_Elab_Term_Do_ToTerm_matchNestedTermResult___closed__1; -x_3 = lean_unsigned_to_nat(1104u); +x_3 = lean_unsigned_to_nat(1105u); x_4 = lean_unsigned_to_nat(27u); x_5 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -24844,7 +24953,7 @@ static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_matchNestedTermResult___clo { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_getDoSeqElems___closed__6; -x_2 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__1; +x_2 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__2; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } @@ -26456,7 +26565,7 @@ lean_ctor_set(x_654, 2, x_653); x_655 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkPureUnit___closed__3; x_656 = l_Lean_addMacroScope(x_586, x_655, x_585); x_657 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkPureUnit___closed__2; -x_658 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__11; +x_658 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__4; lean_inc(x_583); x_659 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_659, 0, x_583); @@ -26811,7 +26920,7 @@ lean_inc(x_727); lean_inc(x_728); x_824 = l_Lean_addMacroScope(x_728, x_823, x_727); x_825 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkPureUnit___closed__2; -x_826 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__11; +x_826 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__4; lean_inc(x_725); x_827 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_827, 0, x_725); @@ -27179,12 +27288,12 @@ lean_inc(x_943); x_998 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_998, 0, x_943); lean_ctor_set(x_998, 1, x_997); -x_999 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__20; +x_999 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__18; lean_inc(x_945); lean_inc(x_946); x_1000 = l_Lean_addMacroScope(x_946, x_999, x_945); -x_1001 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__19; -x_1002 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__22; +x_1001 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__17; +x_1002 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__20; lean_inc(x_943); x_1003 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_1003, 0, x_943); @@ -27277,7 +27386,7 @@ lean_inc(x_945); lean_inc(x_946); x_1042 = l_Lean_addMacroScope(x_946, x_1041, x_945); x_1043 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkPureUnit___closed__2; -x_1044 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__11; +x_1044 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__4; lean_inc(x_943); x_1045 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_1045, 0, x_943); @@ -27593,12 +27702,12 @@ lean_inc(x_1140); x_1195 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_1195, 0, x_1140); lean_ctor_set(x_1195, 1, x_1194); -x_1196 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__26; +x_1196 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__24; lean_inc(x_1142); lean_inc(x_1143); x_1197 = l_Lean_addMacroScope(x_1143, x_1196, x_1142); -x_1198 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__25; -x_1199 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__28; +x_1198 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__23; +x_1199 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__26; lean_inc(x_1140); x_1200 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_1200, 0, x_1140); @@ -27692,7 +27801,7 @@ lean_inc(x_1142); lean_inc(x_1143); x_1239 = l_Lean_addMacroScope(x_1143, x_1238, x_1142); x_1240 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkPureUnit___closed__2; -x_1241 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__11; +x_1241 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__4; lean_inc(x_1140); x_1242 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_1242, 0, x_1140); @@ -34068,24 +34177,23 @@ lean_inc(x_4); x_101 = l_Lean_Elab_Term_Do_ToCodeBlock_doSeqToCode(x_100, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_55); if (lean_obj_tag(x_101) == 0) { -lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; +lean_object* x_102; lean_object* x_103; lean_object* x_104; x_102 = lean_ctor_get(x_101, 0); lean_inc(x_102); x_103 = lean_ctor_get(x_101, 1); lean_inc(x_103); lean_dec(x_101); -x_104 = l_Lean_Elab_Term_Do_mkUnless___closed__1; -x_105 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_105, 0, x_49); -lean_ctor_set(x_105, 1, x_104); -lean_ctor_set(x_105, 2, x_102); -x_17 = x_105; +x_104 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_104, 0, x_49); +lean_ctor_set(x_104, 1, x_77); +lean_ctor_set(x_104, 2, x_102); +x_17 = x_104; x_18 = x_103; goto block_23; } else { -lean_object* x_106; lean_object* x_107; lean_object* x_108; +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_dec(x_49); lean_dec(x_16); lean_dec(x_10); @@ -34095,29 +34203,29 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_106 = lean_ctor_get(x_101, 0); +x_105 = lean_ctor_get(x_101, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_101, 1); lean_inc(x_106); -x_107 = lean_ctor_get(x_101, 1); -lean_inc(x_107); lean_dec(x_101); -x_108 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_108, 0, x_106); -lean_ctor_set(x_108, 1, x_107); -return x_108; +x_107 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_107, 0, x_105); +lean_ctor_set(x_107, 1, x_106); +return x_107; } } } else { -lean_object* x_109; lean_object* x_110; uint8_t x_111; +lean_object* x_108; lean_object* x_109; uint8_t x_110; lean_dec(x_24); -x_109 = lean_unsigned_to_nat(1u); -x_110 = l_Lean_Syntax_getArg(x_14, x_109); -x_111 = l_Lean_Syntax_isIdent(x_110); -if (x_111 == 0) +x_108 = lean_unsigned_to_nat(1u); +x_109 = l_Lean_Syntax_getArg(x_14, x_108); +x_110 = l_Lean_Syntax_isIdent(x_109); +if (x_110 == 0) { -lean_object* x_112; lean_object* x_113; -x_112 = lean_box(0); +lean_object* x_111; lean_object* x_112; +x_111 = lean_box(0); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); @@ -34125,23 +34233,23 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_113 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___spec__3___lambda__1(x_14, x_110, x_112, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_112 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___spec__3___lambda__1(x_14, x_109, x_111, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_14); -if (lean_obj_tag(x_113) == 0) +if (lean_obj_tag(x_112) == 0) { -lean_object* x_114; lean_object* x_115; -x_114 = lean_ctor_get(x_113, 0); +lean_object* x_113; lean_object* x_114; +x_113 = lean_ctor_get(x_112, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_112, 1); lean_inc(x_114); -x_115 = lean_ctor_get(x_113, 1); -lean_inc(x_115); -lean_dec(x_113); -x_17 = x_114; -x_18 = x_115; +lean_dec(x_112); +x_17 = x_113; +x_18 = x_114; goto block_23; } else { -lean_object* x_116; lean_object* x_117; lean_object* x_118; +lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_dec(x_16); lean_dec(x_10); lean_dec(x_9); @@ -34150,70 +34258,70 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_116 = lean_ctor_get(x_113, 0); +x_115 = lean_ctor_get(x_112, 0); +lean_inc(x_115); +x_116 = lean_ctor_get(x_112, 1); lean_inc(x_116); -x_117 = lean_ctor_get(x_113, 1); -lean_inc(x_117); -lean_dec(x_113); -x_118 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_118, 0, x_116); -lean_ctor_set(x_118, 1, x_117); -return x_118; +lean_dec(x_112); +x_117 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_117, 0, x_115); +lean_ctor_set(x_117, 1, x_116); +return x_117; } } else { -lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; -x_119 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_extractBind___lambda__1___closed__3; -lean_inc(x_110); -x_120 = lean_array_push(x_119, x_110); -x_121 = lean_ctor_get(x_9, 0); +lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; +x_118 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_extractBind___lambda__1___closed__3; +lean_inc(x_109); +x_119 = lean_array_push(x_118, x_109); +x_120 = lean_ctor_get(x_9, 0); +lean_inc(x_120); +x_121 = lean_ctor_get(x_9, 1); lean_inc(x_121); -x_122 = lean_ctor_get(x_9, 1); +x_122 = lean_ctor_get(x_9, 2); lean_inc(x_122); -x_123 = lean_ctor_get(x_9, 2); +x_123 = lean_ctor_get(x_9, 3); lean_inc(x_123); -x_124 = lean_ctor_get(x_9, 3); +x_124 = lean_ctor_get(x_9, 4); lean_inc(x_124); -x_125 = lean_ctor_get(x_9, 4); +x_125 = lean_ctor_get(x_9, 5); lean_inc(x_125); -x_126 = lean_ctor_get(x_9, 5); +x_126 = lean_ctor_get(x_9, 6); lean_inc(x_126); -x_127 = lean_ctor_get(x_9, 6); +x_127 = lean_ctor_get(x_9, 7); lean_inc(x_127); -x_128 = lean_ctor_get(x_9, 7); +x_128 = lean_ctor_get(x_9, 8); lean_inc(x_128); -x_129 = lean_ctor_get(x_9, 8); +x_129 = lean_ctor_get(x_9, 9); lean_inc(x_129); -x_130 = lean_ctor_get(x_9, 9); +x_130 = lean_ctor_get(x_9, 10); lean_inc(x_130); -x_131 = lean_ctor_get(x_9, 10); -lean_inc(x_131); -x_132 = l_Lean_replaceRef(x_110, x_126); -lean_dec(x_126); -x_133 = lean_alloc_ctor(0, 11, 0); -lean_ctor_set(x_133, 0, x_121); -lean_ctor_set(x_133, 1, x_122); -lean_ctor_set(x_133, 2, x_123); -lean_ctor_set(x_133, 3, x_124); -lean_ctor_set(x_133, 4, x_125); -lean_ctor_set(x_133, 5, x_132); -lean_ctor_set(x_133, 6, x_127); -lean_ctor_set(x_133, 7, x_128); -lean_ctor_set(x_133, 8, x_129); -lean_ctor_set(x_133, 9, x_130); -lean_ctor_set(x_133, 10, x_131); -x_134 = l_Lean_Elab_Term_Do_ToCodeBlock_checkNotShadowingMutable(x_120, x_4, x_5, x_6, x_7, x_8, x_133, x_10, x_11); -lean_dec(x_133); -lean_dec(x_120); -if (lean_obj_tag(x_134) == 0) +x_131 = l_Lean_replaceRef(x_109, x_125); +lean_dec(x_125); +x_132 = lean_alloc_ctor(0, 11, 0); +lean_ctor_set(x_132, 0, x_120); +lean_ctor_set(x_132, 1, x_121); +lean_ctor_set(x_132, 2, x_122); +lean_ctor_set(x_132, 3, x_123); +lean_ctor_set(x_132, 4, x_124); +lean_ctor_set(x_132, 5, x_131); +lean_ctor_set(x_132, 6, x_126); +lean_ctor_set(x_132, 7, x_127); +lean_ctor_set(x_132, 8, x_128); +lean_ctor_set(x_132, 9, x_129); +lean_ctor_set(x_132, 10, x_130); +x_133 = l_Lean_Elab_Term_Do_ToCodeBlock_checkNotShadowingMutable(x_119, x_4, x_5, x_6, x_7, x_8, x_132, x_10, x_11); +lean_dec(x_132); +lean_dec(x_119); +if (lean_obj_tag(x_133) == 0) { -lean_object* x_135; lean_object* x_136; lean_object* x_137; -x_135 = lean_ctor_get(x_134, 0); +lean_object* x_134; lean_object* x_135; lean_object* x_136; +x_134 = lean_ctor_get(x_133, 0); +lean_inc(x_134); +x_135 = lean_ctor_get(x_133, 1); lean_inc(x_135); -x_136 = lean_ctor_get(x_134, 1); -lean_inc(x_136); -lean_dec(x_134); +lean_dec(x_133); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); @@ -34221,24 +34329,24 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_137 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___spec__3___lambda__1(x_14, x_110, x_135, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_136); -lean_dec(x_135); +x_136 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___spec__3___lambda__1(x_14, x_109, x_134, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_135); +lean_dec(x_134); lean_dec(x_14); -if (lean_obj_tag(x_137) == 0) +if (lean_obj_tag(x_136) == 0) { -lean_object* x_138; lean_object* x_139; -x_138 = lean_ctor_get(x_137, 0); +lean_object* x_137; lean_object* x_138; +x_137 = lean_ctor_get(x_136, 0); +lean_inc(x_137); +x_138 = lean_ctor_get(x_136, 1); lean_inc(x_138); -x_139 = lean_ctor_get(x_137, 1); -lean_inc(x_139); -lean_dec(x_137); -x_17 = x_138; -x_18 = x_139; +lean_dec(x_136); +x_17 = x_137; +x_18 = x_138; goto block_23; } else { -lean_object* x_140; lean_object* x_141; lean_object* x_142; +lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_dec(x_16); lean_dec(x_10); lean_dec(x_9); @@ -34247,21 +34355,21 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_140 = lean_ctor_get(x_137, 0); +x_139 = lean_ctor_get(x_136, 0); +lean_inc(x_139); +x_140 = lean_ctor_get(x_136, 1); lean_inc(x_140); -x_141 = lean_ctor_get(x_137, 1); -lean_inc(x_141); -lean_dec(x_137); -x_142 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_142, 0, x_140); -lean_ctor_set(x_142, 1, x_141); -return x_142; +lean_dec(x_136); +x_141 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_141, 0, x_139); +lean_ctor_set(x_141, 1, x_140); +return x_141; } } else { -lean_object* x_143; lean_object* x_144; lean_object* x_145; -lean_dec(x_110); +lean_object* x_142; lean_object* x_143; lean_object* x_144; +lean_dec(x_109); lean_dec(x_16); lean_dec(x_14); lean_dec(x_10); @@ -34271,15 +34379,15 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_143 = lean_ctor_get(x_134, 0); +x_142 = lean_ctor_get(x_133, 0); +lean_inc(x_142); +x_143 = lean_ctor_get(x_133, 1); lean_inc(x_143); -x_144 = lean_ctor_get(x_134, 1); -lean_inc(x_144); -lean_dec(x_134); -x_145 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_145, 0, x_143); -lean_ctor_set(x_145, 1, x_144); -return x_145; +lean_dec(x_133); +x_144 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_144, 0, x_142); +lean_ctor_set(x_144, 1, x_143); +return x_144; } } } @@ -36069,7 +36177,303 @@ return x_48; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__1(size_t x_1, size_t 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, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__3(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_object* x_9) { +_start: +{ +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_10 = lean_ctor_get(x_7, 5); +x_11 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_5, x_6, x_7, x_8, x_9); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +lean_inc(x_10); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_10); +lean_ctor_set(x_14, 1, x_12); +x_15 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_13); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___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_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +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; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_11 = lean_ctor_get(x_8, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_8, 1); +lean_inc(x_12); +x_13 = lean_ctor_get(x_8, 2); +lean_inc(x_13); +x_14 = lean_ctor_get(x_8, 3); +lean_inc(x_14); +x_15 = lean_ctor_get(x_8, 4); +lean_inc(x_15); +x_16 = lean_ctor_get(x_8, 5); +lean_inc(x_16); +x_17 = lean_ctor_get(x_8, 6); +lean_inc(x_17); +x_18 = lean_ctor_get(x_8, 7); +lean_inc(x_18); +x_19 = lean_ctor_get(x_8, 8); +lean_inc(x_19); +x_20 = lean_ctor_get(x_8, 9); +lean_inc(x_20); +x_21 = lean_ctor_get(x_8, 10); +lean_inc(x_21); +lean_dec(x_8); +x_22 = l_Lean_replaceRef(x_1, x_16); +lean_dec(x_16); +lean_dec(x_1); +x_23 = lean_alloc_ctor(0, 11, 0); +lean_ctor_set(x_23, 0, x_11); +lean_ctor_set(x_23, 1, x_12); +lean_ctor_set(x_23, 2, x_13); +lean_ctor_set(x_23, 3, x_14); +lean_ctor_set(x_23, 4, x_15); +lean_ctor_set(x_23, 5, x_22); +lean_ctor_set(x_23, 6, x_17); +lean_ctor_set(x_23, 7, x_18); +lean_ctor_set(x_23, 8, x_19); +lean_ctor_set(x_23, 9, x_20); +lean_ctor_set(x_23, 10, x_21); +x_24 = l_Lean_throwError___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__3(x_2, x_3, x_4, x_5, x_6, x_7, x_23, x_9, x_10); +lean_dec(x_9); +lean_dec(x_23); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_24; +} +} +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__4(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_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_Do_mkJmp___spec__7___closed__2; +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_1); +lean_ctor_set(x_11, 1, x_10); +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_9); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__5___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Do_mkJmp___spec__8___rarg___closed__2; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__5(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: +{ +lean_object* x_8; +x_8 = lean_alloc_closure((void*)(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__5___rarg), 1, 0); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__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, lean_object* x_9) { +_start: +{ +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; 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; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_10 = lean_st_ref_get(x_8, x_9); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_ctor_get(x_7, 3); +lean_inc(x_14); +x_15 = lean_ctor_get(x_7, 4); +lean_inc(x_15); +x_16 = lean_ctor_get(x_7, 5); +lean_inc(x_16); +x_17 = lean_ctor_get(x_7, 6); +lean_inc(x_17); +x_18 = lean_ctor_get(x_7, 7); +lean_inc(x_18); +x_19 = lean_ctor_get(x_7, 10); +lean_inc(x_19); +lean_inc(x_13); +x_20 = lean_alloc_closure((void*)(l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_Do_mkJmp___spec__1___lambda__1___boxed), 4, 1); +lean_closure_set(x_20, 0, x_13); +lean_inc(x_17); +x_21 = lean_alloc_closure((void*)(l_ReaderT_pure___at_Lean_Elab_liftMacroM___spec__2___rarg___boxed), 3, 1); +lean_closure_set(x_21, 0, x_17); +lean_inc(x_13); +x_22 = lean_alloc_closure((void*)(l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_Do_mkJmp___spec__1___lambda__2___boxed), 4, 1); +lean_closure_set(x_22, 0, x_13); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_13); +x_23 = lean_alloc_closure((void*)(l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_Do_mkJmp___spec__1___lambda__3___boxed), 6, 3); +lean_closure_set(x_23, 0, x_13); +lean_closure_set(x_23, 1, x_17); +lean_closure_set(x_23, 2, x_18); +lean_inc(x_13); +x_24 = lean_alloc_closure((void*)(l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_Do_mkJmp___spec__1___lambda__4___boxed), 6, 3); +lean_closure_set(x_24, 0, x_13); +lean_closure_set(x_24, 1, x_17); +lean_closure_set(x_24, 2, x_18); +x_25 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_25, 0, x_20); +lean_ctor_set(x_25, 1, x_21); +lean_ctor_set(x_25, 2, x_22); +lean_ctor_set(x_25, 3, x_23); +lean_ctor_set(x_25, 4, x_24); +x_26 = lean_st_ref_get(x_8, x_12); +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +x_29 = lean_ctor_get(x_27, 1); +lean_inc(x_29); +lean_dec(x_27); +x_30 = lean_environment_main_module(x_13); +x_31 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_31, 0, x_25); +lean_ctor_set(x_31, 1, x_30); +lean_ctor_set(x_31, 2, x_19); +lean_ctor_set(x_31, 3, x_14); +lean_ctor_set(x_31, 4, x_15); +lean_ctor_set(x_31, 5, x_16); +x_32 = lean_box(0); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_29); +lean_ctor_set(x_33, 1, x_32); +x_34 = lean_apply_2(x_1, x_31, x_33); +if (lean_obj_tag(x_34) == 0) +{ +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +lean_dec(x_34); +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +x_38 = lean_st_ref_take(x_8, x_28); +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_38, 1); +lean_inc(x_40); +lean_dec(x_38); +x_41 = lean_ctor_get(x_39, 0); +lean_inc(x_41); +x_42 = lean_ctor_get(x_39, 2); +lean_inc(x_42); +x_43 = lean_ctor_get(x_39, 3); +lean_inc(x_43); +x_44 = lean_ctor_get(x_39, 4); +lean_inc(x_44); +x_45 = lean_ctor_get(x_39, 5); +lean_inc(x_45); +lean_dec(x_39); +x_46 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_46, 0, x_41); +lean_ctor_set(x_46, 1, x_37); +lean_ctor_set(x_46, 2, x_42); +lean_ctor_set(x_46, 3, x_43); +lean_ctor_set(x_46, 4, x_44); +lean_ctor_set(x_46, 5, x_45); +x_47 = lean_st_ref_set(x_8, x_46, x_40); +x_48 = lean_ctor_get(x_47, 1); +lean_inc(x_48); +lean_dec(x_47); +x_49 = lean_ctor_get(x_36, 1); +lean_inc(x_49); +lean_dec(x_36); +x_50 = l_List_reverse___rarg(x_49); +x_51 = l_List_forM___at_Lean_Elab_Term_Do_ToCodeBlock_mkForInBody___spec__4(x_50, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_48); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_52 = lean_ctor_get(x_51, 1); +lean_inc(x_52); +lean_dec(x_51); +x_53 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_53, 0, x_35); +lean_ctor_set(x_53, 1, x_52); +return x_53; +} +else +{ +lean_object* x_54; +x_54 = lean_ctor_get(x_34, 0); +lean_inc(x_54); +lean_dec(x_34); +if (lean_obj_tag(x_54) == 0) +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; uint8_t x_58; +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_54, 1); +lean_inc(x_56); +lean_dec(x_54); +x_57 = l_Lean_maxRecDepthErrorMessage; +x_58 = lean_string_dec_eq(x_56, x_57); +if (x_58 == 0) +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_59, 0, x_56); +x_60 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_60, 0, x_59); +x_61 = l_Lean_throwErrorAt___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__2(x_55, x_60, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_28); +return x_61; +} +else +{ +lean_object* x_62; +lean_dec(x_56); +x_62 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__4(x_55, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_28); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_62; +} +} +else +{ +lean_object* x_63; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_63 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__5___rarg(x_28); +return x_63; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__6(size_t x_1, size_t 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, lean_object* x_11) { _start: { uint8_t x_12; @@ -36268,28 +36672,84 @@ return x_58; } } } -static lean_object* _init_l_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___closed__1() { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__7(lean_object* x_1, size_t x_2, size_t 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, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_Do_instInhabitedAlt___rarg___closed__1; -x_2 = lean_array_get_size(x_1); -return x_2; +uint8_t x_13; +x_13 = lean_usize_dec_eq(x_2, x_3); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_array_uget(x_1, x_2); +x_15 = lean_alloc_closure((void*)(l_Lean_Elab_Term_expandMatchAlt), 3, 1); +lean_closure_set(x_15, 0, x_14); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_16 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__1(x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; size_t x_20; size_t x_21; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = l_Array_append___rarg(x_4, x_17); +x_20 = 1; +x_21 = lean_usize_add(x_2, x_20); +x_2 = x_21; +x_4 = x_19; +x_12 = x_18; +goto _start; } +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_23 = lean_ctor_get(x_16, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_16, 1); +lean_inc(x_24); +lean_dec(x_16); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; } -static size_t _init_l_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___closed__2() { -_start: +} +else { -lean_object* x_1; size_t x_2; -x_1 = l_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___closed__1; -x_2 = lean_usize_of_nat(x_1); -return x_2; +lean_object* x_26; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_4); +lean_ctor_set(x_26, 1, x_12); +return x_26; +} } } LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode(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_object* x_9, lean_object* x_10) { _start: { -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; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; size_t x_24; +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; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; x_11 = lean_unsigned_to_nat(1u); x_12 = l_Lean_Syntax_getArg(x_1, x_11); x_13 = lean_unsigned_to_nat(2u); @@ -36305,107 +36765,37 @@ x_21 = l_Lean_Syntax_getArgs(x_20); lean_dec(x_20); x_22 = lean_array_get_size(x_21); x_23 = lean_nat_dec_lt(x_19, x_22); -x_24 = 0; if (x_23 == 0) { -size_t x_25; lean_object* x_26; lean_object* x_27; +lean_object* x_43; lean_dec(x_22); lean_dec(x_21); -x_25 = l_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___closed__2; -x_26 = l_Lean_Elab_Term_Do_instInhabitedAlt___rarg___closed__1; -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_27 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__1(x_25, x_24, x_26, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_27) == 0) -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); -lean_dec(x_27); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_30 = l_Lean_Elab_Term_Do_mkMatch(x_1, x_12, x_16, x_14, x_28, x_4, x_5, x_6, x_7, x_8, x_9, x_29); -if (lean_obj_tag(x_30) == 0) -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_30, 1); -lean_inc(x_32); -lean_dec(x_30); -x_33 = l_Lean_Elab_Term_Do_ToCodeBlock_concatWith(x_31, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_32); -return x_33; +x_43 = l_Lean_Elab_Term_Do_instInhabitedAlt___rarg___closed__1; +x_24 = x_43; +x_25 = x_10; +goto block_42; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_34 = lean_ctor_get(x_30, 0); -lean_inc(x_34); -x_35 = lean_ctor_get(x_30, 1); -lean_inc(x_35); -lean_dec(x_30); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -return x_36; -} -} -else +uint8_t x_44; +x_44 = lean_nat_dec_le(x_22, x_22); +if (x_44 == 0) { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -lean_dec(x_16); -lean_dec(x_14); -lean_dec(x_12); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_37 = lean_ctor_get(x_27, 0); -lean_inc(x_37); -x_38 = lean_ctor_get(x_27, 1); -lean_inc(x_38); -lean_dec(x_27); -x_39 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_39, 0, x_37); -lean_ctor_set(x_39, 1, x_38); -return x_39; -} +lean_object* x_45; +lean_dec(x_22); +lean_dec(x_21); +x_45 = l_Lean_Elab_Term_Do_instInhabitedAlt___rarg___closed__1; +x_24 = x_45; +x_25 = x_10; +goto block_42; } else { -uint8_t x_40; -x_40 = lean_nat_dec_le(x_22, x_22); -if (x_40 == 0) -{ -size_t x_41; lean_object* x_42; lean_object* x_43; +size_t x_46; size_t x_47; lean_object* x_48; lean_object* x_49; +x_46 = 0; +x_47 = lean_usize_of_nat(x_22); lean_dec(x_22); -lean_dec(x_21); -x_41 = l_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___closed__2; -x_42 = l_Lean_Elab_Term_Do_instInhabitedAlt___rarg___closed__1; +x_48 = l_Lean_Elab_Term_Do_instInhabitedAlt___rarg___closed__1; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); @@ -36413,58 +36803,23 @@ lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_43 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__1(x_41, x_24, x_42, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_43) == 0) -{ -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_43, 0); -lean_inc(x_44); -x_45 = lean_ctor_get(x_43, 1); -lean_inc(x_45); -lean_dec(x_43); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_46 = l_Lean_Elab_Term_Do_mkMatch(x_1, x_12, x_16, x_14, x_44, x_4, x_5, x_6, x_7, x_8, x_9, x_45); -if (lean_obj_tag(x_46) == 0) -{ -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_46, 0); -lean_inc(x_47); -x_48 = lean_ctor_get(x_46, 1); -lean_inc(x_48); -lean_dec(x_46); -x_49 = l_Lean_Elab_Term_Do_ToCodeBlock_concatWith(x_47, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_48); -return x_49; -} -else +x_49 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__7(x_21, x_46, x_47, x_48, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_21); +if (lean_obj_tag(x_49) == 0) { -lean_object* x_50; lean_object* x_51; lean_object* x_52; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_50 = lean_ctor_get(x_46, 0); +lean_object* x_50; lean_object* x_51; +x_50 = lean_ctor_get(x_49, 0); lean_inc(x_50); -x_51 = lean_ctor_get(x_46, 1); +x_51 = lean_ctor_get(x_49, 1); lean_inc(x_51); -lean_dec(x_46); -x_52 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_52, 0, x_50); -lean_ctor_set(x_52, 1, x_51); -return x_52; -} +lean_dec(x_49); +x_24 = x_50; +x_25 = x_51; +goto block_42; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_dec(x_16); lean_dec(x_14); lean_dec(x_12); @@ -36477,28 +36832,25 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_53 = lean_ctor_get(x_43, 0); +x_52 = lean_ctor_get(x_49, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_49, 1); lean_inc(x_53); -x_54 = lean_ctor_get(x_43, 1); -lean_inc(x_54); -lean_dec(x_43); -x_55 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_55, 0, x_53); -lean_ctor_set(x_55, 1, x_54); -return x_55; +lean_dec(x_49); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +return x_54; } } -else +} +block_42: { -size_t x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; size_t x_60; lean_object* x_61; -x_56 = lean_usize_of_nat(x_22); -lean_dec(x_22); -x_57 = l_Lean_Elab_Term_Do_instInhabitedAlt___rarg___closed__1; -x_58 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_expandMatchAlts_x3f___spec__2(x_21, x_24, x_56, x_57); -lean_dec(x_21); -x_59 = lean_array_get_size(x_58); -x_60 = lean_usize_of_nat(x_59); -lean_dec(x_59); +lean_object* x_26; size_t x_27; size_t x_28; lean_object* x_29; +x_26 = lean_array_get_size(x_24); +x_27 = lean_usize_of_nat(x_26); +lean_dec(x_26); +x_28 = 0; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); @@ -36506,36 +36858,36 @@ lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_61 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__1(x_60, x_24, x_58, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_61) == 0) +x_29 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__6(x_27, x_28, x_24, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_25); +if (lean_obj_tag(x_29) == 0) { -lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_62 = lean_ctor_get(x_61, 0); -lean_inc(x_62); -x_63 = lean_ctor_get(x_61, 1); -lean_inc(x_63); -lean_dec(x_61); +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_29, 1); +lean_inc(x_31); +lean_dec(x_29); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_64 = l_Lean_Elab_Term_Do_mkMatch(x_1, x_12, x_16, x_14, x_62, x_4, x_5, x_6, x_7, x_8, x_9, x_63); -if (lean_obj_tag(x_64) == 0) +x_32 = l_Lean_Elab_Term_Do_mkMatch(x_1, x_12, x_16, x_14, x_30, x_4, x_5, x_6, x_7, x_8, x_9, x_31); +if (lean_obj_tag(x_32) == 0) { -lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_65 = lean_ctor_get(x_64, 0); -lean_inc(x_65); -x_66 = lean_ctor_get(x_64, 1); -lean_inc(x_66); -lean_dec(x_64); -x_67 = l_Lean_Elab_Term_Do_ToCodeBlock_concatWith(x_65, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_66); -return x_67; +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +x_35 = l_Lean_Elab_Term_Do_ToCodeBlock_concatWith(x_33, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_34); +return x_35; } else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -36544,20 +36896,20 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_68 = lean_ctor_get(x_64, 0); -lean_inc(x_68); -x_69 = lean_ctor_get(x_64, 1); -lean_inc(x_69); -lean_dec(x_64); -x_70 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_69); -return x_70; +x_36 = lean_ctor_get(x_32, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_32, 1); +lean_inc(x_37); +lean_dec(x_32); +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +return x_38; } } else { -lean_object* x_71; lean_object* x_72; lean_object* x_73; +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_dec(x_16); lean_dec(x_14); lean_dec(x_12); @@ -36570,16 +36922,15 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_71 = lean_ctor_get(x_61, 0); -lean_inc(x_71); -x_72 = lean_ctor_get(x_61, 1); -lean_inc(x_72); -lean_dec(x_61); -x_73 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_73, 0, x_71); -lean_ctor_set(x_73, 1, x_72); -return x_73; -} +x_39 = lean_ctor_get(x_29, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_29, 1); +lean_inc(x_40); +lean_dec(x_29); +x_41 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +return x_41; } } } @@ -37772,27 +38123,39 @@ return x_1; static lean_object* _init_l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__17() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(2); +x_2 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__16; +x_3 = lean_alloc_ctor(2, 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_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__18() { +_start: +{ lean_object* x_1; x_1 = lean_mk_string_from_bytes("Stream.next?", 12); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__18() { +static lean_object* _init_l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__19() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__17; +x_1 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__18; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__19() { +static lean_object* _init_l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__17; +x_1 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__18; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__18; +x_3 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__19; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -37800,7 +38163,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__20() { +static lean_object* _init_l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__21() { _start: { lean_object* x_1; @@ -37808,17 +38171,17 @@ x_1 = lean_mk_string_from_bytes("Stream", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__21() { +static lean_object* _init_l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__20; +x_2 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__21; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__22() { +static lean_object* _init_l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__23() { _start: { lean_object* x_1; @@ -37826,41 +38189,41 @@ x_1 = lean_mk_string_from_bytes("next?", 5); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__23() { +static lean_object* _init_l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__21; -x_2 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__22; +x_1 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__22; +x_2 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__23; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__24() { +static lean_object* _init_l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__23; +x_2 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__24; 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_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__25() { +static lean_object* _init_l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__24; +x_2 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__25; x_3 = lean_alloc_ctor(1, 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_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__26() { +static lean_object* _init_l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__27() { _start: { lean_object* x_1; @@ -37868,17 +38231,17 @@ x_1 = lean_mk_string_from_bytes("tupleTail", 9); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__27() { +static lean_object* _init_l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_getDoSeqElems___closed__6; -x_2 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__26; +x_2 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__27; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__28() { +static lean_object* _init_l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__29() { _start: { lean_object* x_1; @@ -37886,22 +38249,22 @@ x_1 = lean_mk_string_from_bytes("s'", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__29() { +static lean_object* _init_l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__30() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__28; +x_1 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__29; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__30() { +static lean_object* _init_l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__31() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__28; +x_1 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__29; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__29; +x_3 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__30; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -37909,12 +38272,12 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__31() { +static lean_object* _init_l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__32() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__28; +x_2 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__29; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } @@ -37922,7 +38285,7 @@ return x_3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; +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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; x_14 = lean_ctor_get(x_11, 0); lean_inc(x_14); x_15 = lean_ctor_get(x_11, 1); @@ -38097,8 +38460,8 @@ lean_inc(x_42); x_96 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_96, 0, x_42); lean_ctor_set(x_96, 1, x_95); -x_97 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__16; -x_98 = l_Lean_Syntax_SepArray_ofElems(x_97, x_7); +x_97 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__17; +x_98 = l_Lean_mkSepArray(x_7, x_97); lean_dec(x_7); x_99 = l_Lean_Elab_Term_Do_instInhabitedAlt___rarg___closed__1; x_100 = l_Array_append___rarg(x_99, x_98); @@ -38111,12 +38474,12 @@ lean_inc(x_42); x_103 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_103, 0, x_42); lean_ctor_set(x_103, 1, x_102); -x_104 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__23; +x_104 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__24; lean_inc(x_24); lean_inc(x_48); x_105 = l_Lean_addMacroScope(x_48, x_104, x_24); -x_106 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__19; -x_107 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__25; +x_106 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__20; +x_107 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__26; lean_inc(x_42); x_108 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_108, 0, x_42); @@ -38240,205 +38603,206 @@ lean_inc(x_42); x_161 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_161, 0, x_42); lean_ctor_set(x_161, 1, x_160); +x_162 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__16; lean_inc(x_42); -x_162 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_162, 0, x_42); -lean_ctor_set(x_162, 1, x_97); -x_163 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__31; -x_164 = l_Lean_addMacroScope(x_48, x_163, x_24); -x_165 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__30; +x_163 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_163, 0, x_42); +lean_ctor_set(x_163, 1, x_162); +x_164 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__32; +x_165 = l_Lean_addMacroScope(x_48, x_164, x_24); +x_166 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__31; lean_inc(x_42); -x_166 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_166, 0, x_42); -lean_ctor_set(x_166, 1, x_165); -lean_ctor_set(x_166, 2, x_164); -lean_ctor_set(x_166, 3, x_37); -lean_inc(x_166); -x_167 = lean_array_push(x_55, x_166); -x_168 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_168, 0, x_57); -lean_ctor_set(x_168, 1, x_58); -lean_ctor_set(x_168, 2, x_167); -x_169 = lean_array_push(x_68, x_162); -x_170 = lean_array_push(x_169, x_168); -x_171 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__27; -x_172 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_172, 0, x_57); -lean_ctor_set(x_172, 1, x_171); -lean_ctor_set(x_172, 2, x_170); -x_173 = lean_array_push(x_55, x_172); -x_174 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_174, 0, x_57); -lean_ctor_set(x_174, 1, x_58); -lean_ctor_set(x_174, 2, x_173); -x_175 = lean_array_push(x_68, x_8); -x_176 = lean_array_push(x_175, x_174); -x_177 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_177, 0, x_57); -lean_ctor_set(x_177, 1, x_58); -lean_ctor_set(x_177, 2, x_176); -x_178 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkUnit___closed__17; -x_179 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_179, 0, x_42); -lean_ctor_set(x_179, 1, x_178); -x_180 = lean_array_push(x_85, x_161); -x_181 = lean_array_push(x_180, x_177); -x_182 = lean_array_push(x_181, x_179); -x_183 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkUnit___closed__2; -x_184 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_184, 0, x_57); -lean_ctor_set(x_184, 1, x_183); -lean_ctor_set(x_184, 2, x_182); -x_185 = lean_array_push(x_55, x_184); -x_186 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_186, 0, x_57); -lean_ctor_set(x_186, 1, x_58); -lean_ctor_set(x_186, 2, x_185); -x_187 = lean_array_push(x_68, x_159); -x_188 = lean_array_push(x_187, x_186); -x_189 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_189, 0, x_57); -lean_ctor_set(x_189, 1, x_71); -lean_ctor_set(x_189, 2, x_188); -x_190 = lean_array_push(x_55, x_189); -x_191 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_191, 0, x_57); -lean_ctor_set(x_191, 1, x_58); -lean_ctor_set(x_191, 2, x_190); -x_192 = lean_array_push(x_55, x_191); -x_193 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_193, 0, x_57); -lean_ctor_set(x_193, 1, x_58); -lean_ctor_set(x_193, 2, x_192); -x_194 = lean_array_push(x_78, x_166); -x_195 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_195, 0, x_57); -lean_ctor_set(x_195, 1, x_80); -lean_ctor_set(x_195, 2, x_194); -x_196 = lean_array_push(x_55, x_195); -x_197 = l_Lean_Elab_Term_Do_ToTerm_reassignToTerm___closed__2; -x_198 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_198, 0, x_57); -lean_ctor_set(x_198, 1, x_197); -lean_ctor_set(x_198, 2, x_196); -x_199 = lean_array_push(x_68, x_198); -x_200 = lean_array_push(x_199, x_75); -x_201 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_201, 0, x_57); -lean_ctor_set(x_201, 1, x_93); -lean_ctor_set(x_201, 2, x_200); +x_167 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_167, 0, x_42); +lean_ctor_set(x_167, 1, x_166); +lean_ctor_set(x_167, 2, x_165); +lean_ctor_set(x_167, 3, x_37); +lean_inc(x_167); +x_168 = lean_array_push(x_55, x_167); +x_169 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_169, 0, x_57); +lean_ctor_set(x_169, 1, x_58); +lean_ctor_set(x_169, 2, x_168); +x_170 = lean_array_push(x_68, x_163); +x_171 = lean_array_push(x_170, x_169); +x_172 = l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__28; +x_173 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_173, 0, x_57); +lean_ctor_set(x_173, 1, x_172); +lean_ctor_set(x_173, 2, x_171); +x_174 = lean_array_push(x_55, x_173); +x_175 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_175, 0, x_57); +lean_ctor_set(x_175, 1, x_58); +lean_ctor_set(x_175, 2, x_174); +x_176 = lean_array_push(x_68, x_8); +x_177 = lean_array_push(x_176, x_175); +x_178 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_178, 0, x_57); +lean_ctor_set(x_178, 1, x_58); +lean_ctor_set(x_178, 2, x_177); +x_179 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkUnit___closed__17; +x_180 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_180, 0, x_42); +lean_ctor_set(x_180, 1, x_179); +x_181 = lean_array_push(x_85, x_161); +x_182 = lean_array_push(x_181, x_178); +x_183 = lean_array_push(x_182, x_180); +x_184 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkUnit___closed__2; +x_185 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_185, 0, x_57); +lean_ctor_set(x_185, 1, x_184); +lean_ctor_set(x_185, 2, x_183); +x_186 = lean_array_push(x_55, x_185); +x_187 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_187, 0, x_57); +lean_ctor_set(x_187, 1, x_58); +lean_ctor_set(x_187, 2, x_186); +x_188 = lean_array_push(x_68, x_159); +x_189 = lean_array_push(x_188, x_187); +x_190 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_190, 0, x_57); +lean_ctor_set(x_190, 1, x_71); +lean_ctor_set(x_190, 2, x_189); +x_191 = lean_array_push(x_55, x_190); +x_192 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_192, 0, x_57); +lean_ctor_set(x_192, 1, x_58); +lean_ctor_set(x_192, 2, x_191); +x_193 = lean_array_push(x_55, x_192); +x_194 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_194, 0, x_57); +lean_ctor_set(x_194, 1, x_58); +lean_ctor_set(x_194, 2, x_193); +x_195 = lean_array_push(x_78, x_167); +x_196 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_196, 0, x_57); +lean_ctor_set(x_196, 1, x_80); +lean_ctor_set(x_196, 2, x_195); +x_197 = lean_array_push(x_55, x_196); +x_198 = l_Lean_Elab_Term_Do_ToTerm_reassignToTerm___closed__2; +x_199 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_199, 0, x_57); +lean_ctor_set(x_199, 1, x_198); +lean_ctor_set(x_199, 2, x_197); +x_200 = lean_array_push(x_68, x_199); +x_201 = lean_array_push(x_200, x_75); +x_202 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_202, 0, x_57); +lean_ctor_set(x_202, 1, x_93); +lean_ctor_set(x_202, 2, x_201); lean_inc(x_50); -x_202 = lean_array_push(x_68, x_50); -lean_inc(x_202); -x_203 = lean_array_push(x_202, x_9); -x_204 = l_Lean_Elab_Term_Do_ToCodeBlock_doSeqToCode___closed__12; -x_205 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_205, 0, x_57); -lean_ctor_set(x_205, 1, x_204); -lean_ctor_set(x_205, 2, x_203); -x_206 = lean_array_push(x_68, x_205); -x_207 = lean_array_push(x_206, x_75); -x_208 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_208, 0, x_57); -lean_ctor_set(x_208, 1, x_93); -lean_ctor_set(x_208, 2, x_207); -x_209 = lean_array_push(x_68, x_201); -x_210 = lean_array_push(x_209, x_208); -x_211 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_211, 0, x_57); -lean_ctor_set(x_211, 1, x_58); -lean_ctor_set(x_211, 2, x_210); -x_212 = lean_array_push(x_55, x_211); -x_213 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_213, 0, x_57); -lean_ctor_set(x_213, 1, x_146); -lean_ctor_set(x_213, 2, x_212); -x_214 = lean_array_push(x_149, x_193); -x_215 = lean_array_push(x_214, x_134); -x_216 = lean_array_push(x_215, x_213); -x_217 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_217, 0, x_57); -lean_ctor_set(x_217, 1, x_153); -lean_ctor_set(x_217, 2, x_216); -x_218 = lean_array_push(x_68, x_154); -x_219 = lean_array_push(x_218, x_217); -x_220 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_220, 0, x_57); -lean_ctor_set(x_220, 1, x_58); -lean_ctor_set(x_220, 2, x_219); -x_221 = lean_array_push(x_55, x_220); -x_222 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_liftMethodForbiddenBinder___closed__4; -x_223 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_223, 0, x_57); -lean_ctor_set(x_223, 1, x_222); -lean_ctor_set(x_223, 2, x_221); -x_224 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__6___lambda__1___closed__4; -x_225 = lean_array_push(x_224, x_103); -x_226 = lean_array_push(x_225, x_75); +x_203 = lean_array_push(x_68, x_50); +lean_inc(x_203); +x_204 = lean_array_push(x_203, x_9); +x_205 = l_Lean_Elab_Term_Do_ToCodeBlock_doSeqToCode___closed__12; +x_206 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_206, 0, x_57); +lean_ctor_set(x_206, 1, x_205); +lean_ctor_set(x_206, 2, x_204); +x_207 = lean_array_push(x_68, x_206); +x_208 = lean_array_push(x_207, x_75); +x_209 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_209, 0, x_57); +lean_ctor_set(x_209, 1, x_93); +lean_ctor_set(x_209, 2, x_208); +x_210 = lean_array_push(x_68, x_202); +x_211 = lean_array_push(x_210, x_209); +x_212 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_212, 0, x_57); +lean_ctor_set(x_212, 1, x_58); +lean_ctor_set(x_212, 2, x_211); +x_213 = lean_array_push(x_55, x_212); +x_214 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_214, 0, x_57); +lean_ctor_set(x_214, 1, x_146); +lean_ctor_set(x_214, 2, x_213); +x_215 = lean_array_push(x_149, x_194); +x_216 = lean_array_push(x_215, x_134); +x_217 = lean_array_push(x_216, x_214); +x_218 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_218, 0, x_57); +lean_ctor_set(x_218, 1, x_153); +lean_ctor_set(x_218, 2, x_217); +x_219 = lean_array_push(x_68, x_154); +x_220 = lean_array_push(x_219, x_218); +x_221 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_221, 0, x_57); +lean_ctor_set(x_221, 1, x_58); +lean_ctor_set(x_221, 2, x_220); +x_222 = lean_array_push(x_55, x_221); +x_223 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_liftMethodForbiddenBinder___closed__4; +x_224 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_224, 0, x_57); +lean_ctor_set(x_224, 1, x_223); +lean_ctor_set(x_224, 2, x_222); +x_225 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__6___lambda__1___closed__4; +x_226 = lean_array_push(x_225, x_103); x_227 = lean_array_push(x_226, x_75); -x_228 = lean_array_push(x_227, x_119); -x_229 = lean_array_push(x_228, x_121); -x_230 = lean_array_push(x_229, x_223); -x_231 = l_Lean_Elab_Term_Do_ToTerm_matchNestedTermResult___closed__8; -x_232 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_232, 0, x_57); -lean_ctor_set(x_232, 1, x_231); -lean_ctor_set(x_232, 2, x_230); -x_233 = lean_array_push(x_68, x_232); -x_234 = lean_array_push(x_233, x_75); -x_235 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_235, 0, x_57); -lean_ctor_set(x_235, 1, x_93); -lean_ctor_set(x_235, 2, x_234); -x_236 = lean_array_push(x_55, x_235); -x_237 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_237, 0, x_57); -lean_ctor_set(x_237, 1, x_58); -lean_ctor_set(x_237, 2, x_236); -x_238 = lean_array_push(x_55, x_237); -x_239 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_239, 0, x_57); -lean_ctor_set(x_239, 1, x_146); -lean_ctor_set(x_239, 2, x_238); -x_240 = lean_array_push(x_148, x_96); -x_241 = lean_array_push(x_240, x_101); -x_242 = lean_array_push(x_241, x_50); -x_243 = lean_array_push(x_242, x_239); -x_244 = l_Lean_Elab_Term_Do_ToCodeBlock_doSeqToCode___closed__8; -x_245 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_245, 0, x_57); -lean_ctor_set(x_245, 1, x_244); -lean_ctor_set(x_245, 2, x_243); -x_246 = lean_array_push(x_68, x_245); -x_247 = lean_array_push(x_246, x_75); -x_248 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_248, 0, x_57); -lean_ctor_set(x_248, 1, x_93); -lean_ctor_set(x_248, 2, x_247); -x_249 = lean_array_push(x_68, x_94); -x_250 = lean_array_push(x_249, x_248); -x_251 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_251, 0, x_57); -lean_ctor_set(x_251, 1, x_58); -lean_ctor_set(x_251, 2, x_250); -x_252 = lean_array_push(x_55, x_251); -x_253 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_253, 0, x_57); -lean_ctor_set(x_253, 1, x_146); -lean_ctor_set(x_253, 2, x_252); -x_254 = lean_array_push(x_202, x_253); -x_255 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_liftMethodDelimiter___closed__2; -x_256 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_256, 0, x_57); -lean_ctor_set(x_256, 1, x_255); -lean_ctor_set(x_256, 2, x_254); -x_257 = lean_unsigned_to_nat(1u); -x_258 = l_Lean_Syntax_getArg(x_256, x_257); -lean_dec(x_256); -x_259 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_getDoSeqElems(x_258); -x_260 = l_List_appendTR___rarg(x_259, x_10); -x_261 = l_Lean_Elab_Term_Do_ToCodeBlock_doSeqToCode(x_260, x_2, x_3, x_4, x_5, x_6, x_11, x_12, x_46); -return x_261; +x_228 = lean_array_push(x_227, x_75); +x_229 = lean_array_push(x_228, x_119); +x_230 = lean_array_push(x_229, x_121); +x_231 = lean_array_push(x_230, x_224); +x_232 = l_Lean_Elab_Term_Do_ToTerm_matchNestedTermResult___closed__8; +x_233 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_233, 0, x_57); +lean_ctor_set(x_233, 1, x_232); +lean_ctor_set(x_233, 2, x_231); +x_234 = lean_array_push(x_68, x_233); +x_235 = lean_array_push(x_234, x_75); +x_236 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_236, 0, x_57); +lean_ctor_set(x_236, 1, x_93); +lean_ctor_set(x_236, 2, x_235); +x_237 = lean_array_push(x_55, x_236); +x_238 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_238, 0, x_57); +lean_ctor_set(x_238, 1, x_58); +lean_ctor_set(x_238, 2, x_237); +x_239 = lean_array_push(x_55, x_238); +x_240 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_240, 0, x_57); +lean_ctor_set(x_240, 1, x_146); +lean_ctor_set(x_240, 2, x_239); +x_241 = lean_array_push(x_148, x_96); +x_242 = lean_array_push(x_241, x_101); +x_243 = lean_array_push(x_242, x_50); +x_244 = lean_array_push(x_243, x_240); +x_245 = l_Lean_Elab_Term_Do_ToCodeBlock_doSeqToCode___closed__8; +x_246 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_246, 0, x_57); +lean_ctor_set(x_246, 1, x_245); +lean_ctor_set(x_246, 2, x_244); +x_247 = lean_array_push(x_68, x_246); +x_248 = lean_array_push(x_247, x_75); +x_249 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_249, 0, x_57); +lean_ctor_set(x_249, 1, x_93); +lean_ctor_set(x_249, 2, x_248); +x_250 = lean_array_push(x_68, x_94); +x_251 = lean_array_push(x_250, x_249); +x_252 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_252, 0, x_57); +lean_ctor_set(x_252, 1, x_58); +lean_ctor_set(x_252, 2, x_251); +x_253 = lean_array_push(x_55, x_252); +x_254 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_254, 0, x_57); +lean_ctor_set(x_254, 1, x_146); +lean_ctor_set(x_254, 2, x_253); +x_255 = lean_array_push(x_203, x_254); +x_256 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_liftMethodDelimiter___closed__2; +x_257 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_257, 0, x_57); +lean_ctor_set(x_257, 1, x_256); +lean_ctor_set(x_257, 2, x_255); +x_258 = lean_unsigned_to_nat(1u); +x_259 = l_Lean_Syntax_getArg(x_257, x_258); +lean_dec(x_257); +x_260 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_getDoSeqElems(x_259); +x_261 = l_List_appendTR___rarg(x_260, x_10); +x_262 = l_Lean_Elab_Term_Do_ToCodeBlock_doSeqToCode(x_261, x_2, x_3, x_4, x_5, x_6, x_11, x_12, x_46); +return x_262; } } LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__4(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { @@ -40405,7 +40769,7 @@ lean_ctor_set(x_25, 1, x_23); x_26 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_getDoSeqElems___closed__9; lean_inc(x_6); x_27 = lean_name_mk_string(x_6, x_26); -x_28 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__1; +x_28 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__2; lean_inc(x_6); x_29 = lean_name_mk_string(x_6, x_28); x_30 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_liftMethodForbiddenBinder___closed__13; @@ -40952,7 +41316,7 @@ x_98 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_98, 0, x_23); lean_ctor_set(x_98, 1, x_97); x_99 = lean_array_push(x_52, x_98); -x_100 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__12; +x_100 = l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__10; x_101 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_101, 0, x_49); lean_ctor_set(x_101, 1, x_100); @@ -41073,7 +41437,7 @@ lean_ctor_set(x_34, 1, x_32); x_35 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_getDoSeqElems___closed__9; lean_inc(x_3); x_36 = lean_name_mk_string(x_3, x_35); -x_37 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__1; +x_37 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__2; lean_inc(x_3); x_38 = lean_name_mk_string(x_3, x_37); x_39 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_liftMethodForbiddenBinder___closed__13; @@ -41329,7 +41693,7 @@ lean_ctor_set(x_25, 1, x_23); x_26 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_getDoSeqElems___closed__9; lean_inc(x_6); x_27 = lean_name_mk_string(x_6, x_26); -x_28 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__1; +x_28 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__2; lean_inc(x_6); x_29 = lean_name_mk_string(x_6, x_28); x_30 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_liftMethodForbiddenBinder___closed__13; @@ -41488,7 +41852,7 @@ lean_ctor_set(x_25, 1, x_23); x_26 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_getDoSeqElems___closed__9; lean_inc(x_6); x_27 = lean_name_mk_string(x_6, x_26); -x_28 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__1; +x_28 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__2; lean_inc(x_6); x_29 = lean_name_mk_string(x_6, x_28); x_30 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_liftMethodForbiddenBinder___closed__13; @@ -42313,7 +42677,52 @@ x_16 = l_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___lambda__5(x_1, x_2, x_3, x_ return x_16; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__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, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___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_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_throwError___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__4___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_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__5___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: +{ +lean_object* x_8; +x_8 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__6___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_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { size_t x_12; size_t x_13; lean_object* x_14; @@ -42321,10 +42730,23 @@ x_12 = lean_unbox_usize(x_1); lean_dec(x_1); x_13 = lean_unbox_usize(x_2); lean_dec(x_2); -x_14 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__1(x_12, x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_14 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__6(x_12, x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); return x_14; } } +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__7___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +size_t x_13; size_t x_14; lean_object* x_15; +x_13 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_14 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___spec__7(x_1, x_13, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_1); +return x_15; +} +} LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { @@ -42996,7 +43418,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Do_elabDo_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1603u); +x_1 = lean_unsigned_to_nat(1604u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -43008,7 +43430,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Do_elabDo_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1610u); +x_1 = lean_unsigned_to_nat(1611u); x_2 = lean_unsigned_to_nat(84u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -43036,7 +43458,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Do_elabDo_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1603u); +x_1 = lean_unsigned_to_nat(1604u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -43048,7 +43470,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Do_elabDo_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1603u); +x_1 = lean_unsigned_to_nat(1604u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -43094,7 +43516,7 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_29313_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_30705_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; @@ -43250,7 +43672,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandTermFor_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1621u); +x_1 = lean_unsigned_to_nat(1622u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -43262,7 +43684,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandTermFor_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1621u); +x_1 = lean_unsigned_to_nat(1622u); x_2 = lean_unsigned_to_nat(62u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -43290,7 +43712,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandTermFor_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1621u); +x_1 = lean_unsigned_to_nat(1622u); x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -43302,7 +43724,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandTermFor_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1621u); +x_1 = lean_unsigned_to_nat(1622u); x_2 = lean_unsigned_to_nat(17u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -43408,7 +43830,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandTermTry_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1624u); +x_1 = lean_unsigned_to_nat(1625u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -43420,7 +43842,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandTermTry_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1624u); +x_1 = lean_unsigned_to_nat(1625u); x_2 = lean_unsigned_to_nat(62u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -43448,7 +43870,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandTermTry_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1624u); +x_1 = lean_unsigned_to_nat(1625u); x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -43460,7 +43882,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandTermTry_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1624u); +x_1 = lean_unsigned_to_nat(1625u); x_2 = lean_unsigned_to_nat(17u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -43566,7 +43988,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandTermUnless_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1627u); +x_1 = lean_unsigned_to_nat(1628u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -43578,7 +44000,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandTermUnless_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1627u); +x_1 = lean_unsigned_to_nat(1628u); x_2 = lean_unsigned_to_nat(68u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -43606,7 +44028,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandTermUnless_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1627u); +x_1 = lean_unsigned_to_nat(1628u); x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -43618,7 +44040,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandTermUnless_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1627u); +x_1 = lean_unsigned_to_nat(1628u); x_2 = lean_unsigned_to_nat(20u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -43724,7 +44146,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandTermReturn_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1630u); +x_1 = lean_unsigned_to_nat(1631u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -43736,7 +44158,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandTermReturn_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1630u); +x_1 = lean_unsigned_to_nat(1631u); x_2 = lean_unsigned_to_nat(68u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -43764,7 +44186,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandTermReturn_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1630u); +x_1 = lean_unsigned_to_nat(1631u); x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -43776,7 +44198,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandTermReturn_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1630u); +x_1 = lean_unsigned_to_nat(1631u); x_2 = lean_unsigned_to_nat(20u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -44319,6 +44741,12 @@ l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__1___close lean_mark_persistent(l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__1___closed__2); l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__1 = _init_l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__1); +l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__2 = _init_l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__2(); +lean_mark_persistent(l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__2); +l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__3 = _init_l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__3(); +lean_mark_persistent(l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__3); +l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__4 = _init_l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__4(); +lean_mark_persistent(l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__4); l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___closed__1 = _init_l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___closed__1); l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___closed__2 = _init_l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___closed__2(); @@ -44666,10 +45094,6 @@ l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__25 = _init_l_Lean_Elab lean_mark_persistent(l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__25); l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__26 = _init_l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__26(); lean_mark_persistent(l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__26); -l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__27 = _init_l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__27(); -lean_mark_persistent(l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__27); -l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__28 = _init_l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__28(); -lean_mark_persistent(l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__28); l_Lean_Elab_Term_Do_ToTerm_seqToTerm___closed__1 = _init_l_Lean_Elab_Term_Do_ToTerm_seqToTerm___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_Do_ToTerm_seqToTerm___closed__1); l_Lean_Elab_Term_Do_ToTerm_seqToTerm___closed__2 = _init_l_Lean_Elab_Term_Do_ToTerm_seqToTerm___closed__2(); @@ -44977,9 +45401,6 @@ l_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___lambda__5___closed__1 = _init_l_Le lean_mark_persistent(l_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___lambda__5___closed__1); l_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___lambda__5___closed__2 = _init_l_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___lambda__5___closed__2(); lean_mark_persistent(l_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___lambda__5___closed__2); -l_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___closed__1 = _init_l_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___closed__1); -l_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___closed__2 = _init_l_Lean_Elab_Term_Do_ToCodeBlock_doMatchToCode___closed__2(); l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__1___closed__1 = _init_l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__1___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__1___closed__1); l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__1 = _init_l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__1(); @@ -45044,6 +45465,8 @@ l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__30 = _init_l_L lean_mark_persistent(l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__30); l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__31 = _init_l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__31(); lean_mark_persistent(l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__31); +l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__32 = _init_l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__32(); +lean_mark_persistent(l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__32); l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___closed__1 = _init_l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___closed__1); l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___closed__2 = _init_l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___closed__2(); @@ -45126,7 +45549,7 @@ lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Do_elabDo_declRange___closed_ res = l___regBuiltin_Lean_Elab_Term_Do_elabDo_declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_29313_(lean_io_mk_world()); +res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_30705_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); l___regBuiltin_Lean_Elab_Term_expandTermFor___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_expandTermFor___closed__1(); diff --git a/stage0/stdlib/Lean/Elab/ElabRules.c b/stage0/stdlib/Lean/Elab/ElabRules.c index f09208c7871..533521bf73e 100644 --- a/stage0/stdlib/Lean/Elab/ElabRules.c +++ b/stage0/stdlib/Lean/Elab/ElabRules.c @@ -18,6 +18,9 @@ static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__3 static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__59; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__82; lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__6; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__2; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElab___spec__2(size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Command_elabElab___lambda__1___closed__6; size_t lean_usize_add(size_t, size_t); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__12; @@ -26,18 +29,16 @@ static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__1 static lean_object* l___regBuiltin_Lean_Elab_Command_elabElabRules___closed__2; lean_object* l_Lean_stringToMessageData(lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__90; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__10; lean_object* l_Lean_Elab_Command_expandMacroArg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabElabRules___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRules___lambda__7___closed__2; lean_object* l_Lean_Elab_Command_expandNoKindMacroRulesAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRules___lambda__3___closed__1; -extern lean_object* l_Lean_nullKind; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabCommand___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_name_mk_string(lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__3; static lean_object* l_Lean_Elab_Command_elabElab___lambda__1___closed__5; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__92; static lean_object* l_Lean_Elab_Command_elabElabRules___lambda__7___closed__1; @@ -50,20 +51,20 @@ static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__6 static lean_object* l_Lean_Elab_Command_elabElab___lambda__1___closed__3; extern lean_object* l_Lean_Elab_Command_commandElabAttribute; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__7; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__68; LEAN_EXPORT lean_object* l_Lean_Elab_Command_withExpectedType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__87; +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabElabRulesAux___spec__1(lean_object*, lean_object*); lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Command_elabAuxDef___spec__4(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__21; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__38; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__107; static lean_object* l_Lean_Elab_Command_elabElabRules___lambda__3___closed__2; uint8_t lean_name_eq(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabElabRules___lambda__5(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_Command_elabElabRules___lambda__3(lean_object*, 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_throwError___at_Lean_Elab_Command_elabElabRulesAux___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkIdentFrom(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__44; static lean_object* l_Lean_Elab_Command_elabElabRules___lambda__6___closed__1; @@ -72,21 +73,17 @@ static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__9 static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__8; lean_object* lean_array_push(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__1___boxed(lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__4; static lean_object* l_Lean_Elab_Command_elabElabRules___lambda__7___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabElab(lean_object*); static lean_object* l_Lean_Elab_Command_elabElab___lambda__1___closed__9; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabElabRulesAux___spec__4(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__57; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__8; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__1; +lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabSyntax___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__111; lean_object* lean_string_utf8_byte_size(lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__48; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElab___closed__1; static lean_object* l_Lean_Elab_Command_elabElabRules___lambda__3___closed__4; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__13; lean_object* l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__89; lean_object* l_Lean_Elab_Command_mkNameFromParserSyntax(lean_object*, lean_object*, lean_object*, lean_object*); @@ -96,11 +93,10 @@ static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__7 uint8_t l_Lean_Elab_Command_checkRuleKind(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__60; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__47; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__12; static lean_object* l_Lean_Elab_Command_elabElab___closed__2; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabElabRulesAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_resolveSyntaxKind(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Syntax_SepArray_ofElems(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabElabRules_declRange___closed__3; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__27; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__67; @@ -112,11 +108,11 @@ static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__4 static lean_object* l___regBuiltin_Lean_Elab_Command_elabElab___closed__2; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__65; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__109; -lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__9; static lean_object* l_Lean_Elab_Command_elabElab___lambda__1___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Command_elabElab___closed__1; lean_object* l_Lean_Elab_Term_tryPostponeIfNoneOrMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElab___spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__51; static lean_object* l___regBuiltin_Lean_Elab_Command_elabElabRules_declRange___closed__7; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__39; @@ -124,110 +120,116 @@ static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__6 static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__97; static lean_object* l_Lean_Elab_Command_elabElabRules___closed__1; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__1; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__17; +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabElabRules(lean_object*); -extern lean_object* l_Lean_numLitKind; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__1; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__11; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__15; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__83; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__6; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElab___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Command_elabElab_declRange___closed__4; static lean_object* l_Lean_Elab_Command_elabElab___lambda__1___closed__10; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__7; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__5; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__76; -lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabSyntax___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__18; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__30; +lean_object* l_Lean_replaceRef(lean_object*, lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabElabRules___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__5; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__14; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__28; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabElabRules___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__15; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Command_elabElab_declRange___closed__5; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__29; static lean_object* l___regBuiltin_Lean_Elab_Command_elabElabRules___closed__1; lean_object* l_Nat_repr(lean_object*); static lean_object* l_Lean_Elab_Command_withExpectedType___closed__1; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__52; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__16; lean_object* l_Lean_Syntax_getId(lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__55; lean_object* l_Lean_throwError___at_Lean_Elab_Command_resolveSyntaxKind___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_mkIdentFromRef___at_Lean_Elab_Command_elabElabRulesAux___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_adaptExpander(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_choiceKind; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__42; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__4; lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabCommand___spec__12(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabElab___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__100; +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabElabRulesAux___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addBuiltinDeclarationRanges(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__103; +lean_object* l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabElab_declRange___closed__3; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__6; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__8; static lean_object* l_Lean_Elab_Command_withExpectedType___closed__2; -LEAN_EXPORT lean_object* l_Lean_mkIdentFromRef___at_Lean_Elab_Command_elabElabRulesAux___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__37; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__11; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__78; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__91; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__93; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__9; +LEAN_EXPORT lean_object* l_Lean_mkIdentFromRef___at_Lean_Elab_Command_elabElabRulesAux___spec__6(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__22; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRules___lambda__3___closed__7; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__25; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__10; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__86; lean_object* l_Array_unzip___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabElab___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabElab_declRange___closed__6; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__46; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__7; extern lean_object* l_Lean_instInhabitedSyntax; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__2; +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabElabRulesAux___spec__1___boxed(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabElabRules_declRange___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Command_elabElab_declRange___closed__1; lean_object* l_Lean_Elab_Command_elabCommand(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__34; lean_object* l_Lean_evalOptPrio(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabElabRules_declRange___closed__6; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__8; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__13; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__41; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__88; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__23; size_t lean_usize_of_nat(lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__96; -lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabSyntax___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__72; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__4; static lean_object* l_Lean_Elab_Command_elabElabRules___lambda__3___closed__3; lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__3; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__84; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__62; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__33; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabElab___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__110; -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__13; +lean_object* l_Lean_Syntax_mkNumLit(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRules___lambda__7___closed__3; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__54; static lean_object* l_Lean_Elab_Command_elabElabRules___lambda__3___closed__5; lean_object* l_Lean_Syntax_getQuotContent(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabElabRules_declRange___closed__1; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__6; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__99; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabElabRulesAux___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__50; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabElab(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__6; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__80; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabElab___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, 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_Command_elabElab___lambda__4(lean_object*, lean_object*, lean_object*, 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_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__79; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__95; lean_object* l_Lean_Syntax_setArg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__66; +lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__56; lean_object* l_Lean_Elab_Command_getCurrMacroScope(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabElabRulesAux___spec__1___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabElabRules___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArgs(lean_object*); lean_object* l_Lean_Name_append(lean_object*, lean_object*); @@ -237,15 +239,16 @@ lean_object* l_Lean_Syntax_getKind(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabElabRules_declRange___closed__2; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__16; static lean_object* l___regBuiltin_Lean_Elab_Command_elabElab_declRange___closed__2; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabElabRulesAux___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_mkIdentFromRef___at_Lean_Elab_Command_elabElabRulesAux___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_inferMacroRulesAltKind___spec__2___rarg(lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__101; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__43; lean_object* l_Lean_Elab_Command_getMainModule___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabElabRules_declRange(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__20; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__7; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__2; lean_object* l_Lean_Elab_Command_getRef(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__73; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabElabRules(lean_object*, lean_object*, lean_object*, lean_object*); @@ -254,6 +257,8 @@ static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__3 static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__10; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__49; static lean_object* l_Lean_Elab_Command_elabElab___lambda__1___closed__7; +static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabElabRulesAux___spec__1___rarg___closed__1; +lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabSyntax___spec__9(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabElab_declRange___closed__7; uint8_t l_Lean_Syntax_isNone(lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__24; @@ -268,13 +273,16 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabElab___lambda__2___boxed(lean_o LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabElabRules___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__98; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__81; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__12; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__1; +extern lean_object* l_Lean_Elab_unsupportedSyntaxExceptionId; +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabElabRulesAux___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__14; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__77; lean_object* lean_mk_syntax_ident(lean_object*); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__108; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElab___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__5; static lean_object* l_Lean_Elab_Command_elabElab___lambda__1___closed__1; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__7; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__71; @@ -285,28 +293,31 @@ lean_object* l_Lean_Elab_Command_getScope___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__113; static lean_object* l_Lean_Elab_Command_elabElab___lambda__1___closed__2; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__11; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabElabRules___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___closed__2; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__58; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__75; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__26; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__114; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__104; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__15; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__1; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__36; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabElab___lambda__5(lean_object*, lean_object*, 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_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabElabRulesAux___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__3; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__19; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__8; lean_object* l_Lean_throwError___at_Lean_Elab_Term_synthesizeInst___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__70; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Command_elabElab___closed__3; +static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabElabRulesAux___spec__1___rarg___closed__2; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__9; +lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabElabRules___lambda__1___boxed(lean_object*, 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_Command_elabElab___lambda__1___boxed(lean_object**); -lean_object* l_Lean_Syntax_mkLit(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabElabRules___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__35; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__14; static lean_object* _init_l_Lean_Elab_Command_withExpectedType___closed__1() { _start: { @@ -398,7 +409,157 @@ return x_20; } } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabElabRulesAux___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +static lean_object* _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabElabRulesAux___spec__1___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Elab_unsupportedSyntaxExceptionId; +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabElabRulesAux___spec__1___rarg___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabElabRulesAux___spec__1___rarg___closed__1; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabElabRulesAux___spec__1___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabElabRulesAux___spec__1___rarg___closed__2; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabElabRulesAux___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabElabRulesAux___spec__1___rarg), 1, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabElabRulesAux___spec__3(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; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_5 = l_Lean_Elab_Command_getRef(x_2, x_3, x_4); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_5, 1); +lean_inc(x_7); +lean_dec(x_5); +x_8 = lean_ctor_get(x_2, 4); +lean_inc(x_8); +lean_inc(x_8); +x_9 = l_Lean_Elab_getBetterRef(x_6, x_8); +lean_dec(x_6); +x_10 = l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1(x_1, x_2, x_3, x_7); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1(x_11, x_8, x_2, x_3, x_12); +lean_dec(x_2); +lean_dec(x_8); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_9); +lean_ctor_set(x_16, 1, x_15); +lean_ctor_set_tag(x_13, 1); +lean_ctor_set(x_13, 0, x_16); +return x_13; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_17 = lean_ctor_get(x_13, 0); +x_18 = lean_ctor_get(x_13, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_13); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_9); +lean_ctor_set(x_19, 1, x_17); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_18); +return x_20; +} +} +} +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabElabRulesAux___spec__2(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; lean_object* x_9; uint8_t x_10; +x_6 = l_Lean_Elab_Command_getRef(x_3, x_4, x_5); +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_6, 1); +lean_inc(x_8); +lean_dec(x_6); +x_9 = l_Lean_replaceRef(x_1, x_7); +lean_dec(x_7); +lean_dec(x_1); +x_10 = !lean_is_exclusive(x_3); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_3, 6); +lean_dec(x_11); +lean_ctor_set(x_3, 6, x_9); +x_12 = l_Lean_throwError___at_Lean_Elab_Command_elabElabRulesAux___spec__3(x_2, x_3, x_4, x_8); +lean_dec(x_4); +return x_12; +} +else +{ +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; lean_object* x_20; lean_object* x_21; +x_13 = lean_ctor_get(x_3, 0); +x_14 = lean_ctor_get(x_3, 1); +x_15 = lean_ctor_get(x_3, 2); +x_16 = lean_ctor_get(x_3, 3); +x_17 = lean_ctor_get(x_3, 4); +x_18 = lean_ctor_get(x_3, 5); +x_19 = lean_ctor_get(x_3, 7); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_3); +x_20 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_20, 0, x_13); +lean_ctor_set(x_20, 1, x_14); +lean_ctor_set(x_20, 2, x_15); +lean_ctor_set(x_20, 3, x_16); +lean_ctor_set(x_20, 4, x_17); +lean_ctor_set(x_20, 5, x_18); +lean_ctor_set(x_20, 6, x_9); +lean_ctor_set(x_20, 7, x_19); +x_21 = l_Lean_throwError___at_Lean_Elab_Command_elabElabRulesAux___spec__3(x_2, x_20, x_4, x_8); +lean_dec(x_4); +return x_21; +} +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabElabRulesAux___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { _start: { uint8_t x_7; @@ -444,14 +605,14 @@ return x_16; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__1(lean_object* x_1, lean_object* x_2) { _start: { lean_inc(x_1); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__1() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__1() { _start: { lean_object* x_1; @@ -459,7 +620,7 @@ x_1 = lean_mk_string_from_bytes("|", 1); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__2() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__2() { _start: { lean_object* x_1; @@ -467,17 +628,17 @@ x_1 = lean_mk_string_from_bytes("null", 4); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__3() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__2; +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__2; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__4() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__4() { _start: { lean_object* x_1; lean_object* x_2; @@ -486,15 +647,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__5() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(",", 1); -return x_1; -} -} -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__6() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__5() { _start: { lean_object* x_1; lean_object* x_2; @@ -503,7 +656,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__7() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__6() { _start: { lean_object* x_1; @@ -511,7 +664,7 @@ x_1 = lean_mk_string_from_bytes("=>", 2); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__8() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__7() { _start: { lean_object* x_1; lean_object* x_2; @@ -520,7 +673,25 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__9() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("choice", 6); +return x_1; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__8; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__10() { _start: { lean_object* x_1; @@ -528,16 +699,16 @@ x_1 = lean_mk_string_from_bytes("invalid elab_rules alternative, unexpected synt return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__10() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__9; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__10; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__11() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__12() { _start: { lean_object* x_1; @@ -545,16 +716,16 @@ x_1 = lean_mk_string_from_bytes("'", 1); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__12() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__13() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__11; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__12; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__13() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -566,7 +737,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__14() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__15() { _start: { lean_object* x_1; @@ -574,16 +745,16 @@ x_1 = lean_mk_string_from_bytes("invalid elab_rules alternative, expected syntax return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__15() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__16() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__14; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__15; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___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_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___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_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; lean_object* x_12; uint8_t x_13; @@ -595,7 +766,7 @@ x_13 = l_Lean_Elab_Command_checkRuleKind(x_12, x_5); if (x_13 == 0) { lean_object* x_14; uint8_t x_15; -x_14 = l_Lean_choiceKind; +x_14 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__9; x_15 = lean_name_eq(x_12, x_14); if (x_15 == 0) { @@ -608,15 +779,15 @@ lean_dec(x_2); lean_dec(x_1); x_16 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_16, 0, x_12); -x_17 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__10; +x_17 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__11; x_18 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_18, 0, x_17); lean_ctor_set(x_18, 1, x_16); -x_19 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__12; +x_19 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__13; x_20 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_20, 0, x_18); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_throwErrorAt___at_Lean_Elab_Command_elabCommand___spec__11(x_6, x_20, x_8, x_9, x_10); +x_21 = l_Lean_throwErrorAt___at_Lean_Elab_Command_elabElabRulesAux___spec__2(x_6, x_20, x_8, x_9, x_10); return x_21; } else @@ -629,8 +800,8 @@ x_23 = lean_array_get_size(x_22); x_24 = lean_usize_of_nat(x_23); lean_dec(x_23); x_25 = 0; -x_26 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__13; -x_27 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabElabRulesAux___spec__1(x_5, x_26, x_22, x_24, x_25, x_26); +x_26 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__14; +x_27 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabElabRulesAux___spec__4(x_5, x_26, x_22, x_24, x_25, x_26); lean_dec(x_22); x_28 = lean_ctor_get(x_27, 0); lean_inc(x_28); @@ -644,15 +815,15 @@ lean_dec(x_2); lean_dec(x_1); x_29 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_29, 0, x_5); -x_30 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__15; +x_30 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__16; x_31 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_31, 0, x_30); lean_ctor_set(x_31, 1, x_29); -x_32 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__12; +x_32 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__13; x_33 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_33, 0, x_31); lean_ctor_set(x_33, 1, x_32); -x_34 = l_Lean_throwErrorAt___at_Lean_Elab_Command_elabCommand___spec__11(x_6, x_33, x_8, x_9, x_10); +x_34 = l_Lean_throwErrorAt___at_Lean_Elab_Command_elabElabRulesAux___spec__2(x_6, x_33, x_8, x_9, x_10); return x_34; } else @@ -683,99 +854,93 @@ lean_dec(x_9); x_46 = !lean_is_exclusive(x_45); if (x_46 == 0) { -lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; lean_object* x_67; +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; x_47 = lean_ctor_get(x_45, 0); lean_dec(x_47); -x_48 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__1; +x_48 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__1; lean_inc(x_41); x_49 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_49, 0, x_41); lean_ctor_set(x_49, 1, x_48); -x_50 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__5; -x_51 = l_Lean_Syntax_SepArray_ofElems(x_50, x_39); -lean_dec(x_39); -x_52 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__4; -x_53 = l_Array_append___rarg(x_52, x_51); -x_54 = lean_box(2); -x_55 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__3; -x_56 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_56, 0, x_54); -lean_ctor_set(x_56, 1, x_55); -lean_ctor_set(x_56, 2, x_53); -x_57 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__6; -x_58 = lean_array_push(x_57, x_56); -x_59 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_59, 0, x_54); -lean_ctor_set(x_59, 1, x_55); -lean_ctor_set(x_59, 2, x_58); -x_60 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__7; -x_61 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_61, 0, x_41); -lean_ctor_set(x_61, 1, x_60); -x_62 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__8; -x_63 = lean_array_push(x_62, x_49); -x_64 = lean_array_push(x_63, x_59); -x_65 = lean_array_push(x_64, x_61); -x_66 = lean_array_push(x_65, x_3); -x_67 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_67, 0, x_54); -lean_ctor_set(x_67, 1, x_4); -lean_ctor_set(x_67, 2, x_66); -lean_ctor_set(x_45, 0, x_67); +x_50 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__4; +x_51 = l_Array_append___rarg(x_50, x_39); +x_52 = lean_box(2); +x_53 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__3; +x_54 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +lean_ctor_set(x_54, 2, x_51); +x_55 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__5; +x_56 = lean_array_push(x_55, x_54); +x_57 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_57, 0, x_52); +lean_ctor_set(x_57, 1, x_53); +lean_ctor_set(x_57, 2, x_56); +x_58 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__6; +x_59 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_59, 0, x_41); +lean_ctor_set(x_59, 1, x_58); +x_60 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__7; +x_61 = lean_array_push(x_60, x_49); +x_62 = lean_array_push(x_61, x_57); +x_63 = lean_array_push(x_62, x_59); +x_64 = lean_array_push(x_63, x_3); +x_65 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_65, 0, x_52); +lean_ctor_set(x_65, 1, x_4); +lean_ctor_set(x_65, 2, x_64); +lean_ctor_set(x_45, 0, x_65); return x_45; } else { -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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_68 = lean_ctor_get(x_45, 1); -lean_inc(x_68); +lean_object* x_66; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_66 = lean_ctor_get(x_45, 1); +lean_inc(x_66); lean_dec(x_45); -x_69 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__1; +x_67 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__1; lean_inc(x_41); -x_70 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_70, 0, x_41); -lean_ctor_set(x_70, 1, x_69); -x_71 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__5; -x_72 = l_Lean_Syntax_SepArray_ofElems(x_71, x_39); -lean_dec(x_39); -x_73 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__4; -x_74 = l_Array_append___rarg(x_73, x_72); -x_75 = lean_box(2); -x_76 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__3; -x_77 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_77, 0, x_75); -lean_ctor_set(x_77, 1, x_76); -lean_ctor_set(x_77, 2, x_74); -x_78 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__6; -x_79 = lean_array_push(x_78, x_77); -x_80 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_80, 0, x_75); -lean_ctor_set(x_80, 1, x_76); -lean_ctor_set(x_80, 2, x_79); -x_81 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__7; -x_82 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_82, 0, x_41); -lean_ctor_set(x_82, 1, x_81); -x_83 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__8; -x_84 = lean_array_push(x_83, x_70); -x_85 = lean_array_push(x_84, x_80); -x_86 = lean_array_push(x_85, x_82); -x_87 = lean_array_push(x_86, x_3); -x_88 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_88, 0, x_75); -lean_ctor_set(x_88, 1, x_4); -lean_ctor_set(x_88, 2, x_87); -x_89 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_89, 0, x_88); -lean_ctor_set(x_89, 1, x_68); -return x_89; +x_68 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_68, 0, x_41); +lean_ctor_set(x_68, 1, x_67); +x_69 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__4; +x_70 = l_Array_append___rarg(x_69, x_39); +x_71 = lean_box(2); +x_72 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__3; +x_73 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_73, 0, x_71); +lean_ctor_set(x_73, 1, x_72); +lean_ctor_set(x_73, 2, x_70); +x_74 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__5; +x_75 = lean_array_push(x_74, x_73); +x_76 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_76, 0, x_71); +lean_ctor_set(x_76, 1, x_72); +lean_ctor_set(x_76, 2, x_75); +x_77 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__6; +x_78 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_78, 0, x_41); +lean_ctor_set(x_78, 1, x_77); +x_79 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__7; +x_80 = lean_array_push(x_79, x_68); +x_81 = lean_array_push(x_80, x_76); +x_82 = lean_array_push(x_81, x_78); +x_83 = lean_array_push(x_82, x_3); +x_84 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_84, 0, x_71); +lean_ctor_set(x_84, 1, x_4); +lean_ctor_set(x_84, 2, x_83); +x_85 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_85, 0, x_84); +lean_ctor_set(x_85, 1, x_66); +return x_85; } } } } else { -lean_object* x_90; +lean_object* x_86; lean_dec(x_12); lean_dec(x_11); lean_dec(x_9); @@ -785,14 +950,14 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_90 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_90, 0, x_6); -lean_ctor_set(x_90, 1, x_10); -return x_90; +x_86 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_86, 0, x_6); +lean_ctor_set(x_86, 1, x_10); +return x_86; } } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__1() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__1() { _start: { lean_object* x_1; @@ -800,17 +965,17 @@ x_1 = lean_mk_string_from_bytes("Lean", 4); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__2() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__1; +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__3() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__3() { _start: { lean_object* x_1; @@ -818,17 +983,17 @@ x_1 = lean_mk_string_from_bytes("Parser", 6); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__4() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__2; -x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__3; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__2; +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__5() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__5() { _start: { lean_object* x_1; @@ -836,17 +1001,17 @@ x_1 = lean_mk_string_from_bytes("Term", 4); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__6() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__4; -x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__5; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__4; +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__7() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__7() { _start: { lean_object* x_1; @@ -854,17 +1019,17 @@ x_1 = lean_mk_string_from_bytes("matchAlt", 8); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__8() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__6; -x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__7; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__6; +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__7; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { uint8_t x_8; @@ -886,7 +1051,7 @@ lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint x_10 = lean_array_uget(x_4, x_3); x_11 = lean_unsigned_to_nat(0u); x_12 = lean_array_uset(x_4, x_3, x_11); -x_13 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__8; +x_13 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__8; lean_inc(x_10); x_14 = l_Lean_Syntax_isOfKind(x_10, x_13); if (x_14 == 0) @@ -897,7 +1062,7 @@ lean_dec(x_10); lean_dec(x_6); lean_dec(x_5); lean_dec(x_1); -x_15 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__14___rarg(x_7); +x_15 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabElabRulesAux___spec__1___rarg(x_7); x_16 = !lean_is_exclusive(x_15); if (x_16 == 0) { @@ -919,132 +1084,131 @@ return x_19; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +lean_object* x_20; lean_object* x_21; uint8_t x_22; x_20 = lean_unsigned_to_nat(1u); x_21 = l_Lean_Syntax_getArg(x_10, x_20); -x_22 = l_Lean_nullKind; lean_inc(x_21); -x_23 = l_Lean_Syntax_isNodeOf(x_21, x_22, x_20); -if (x_23 == 0) +x_22 = l_Lean_Syntax_matchesNull(x_21, x_20); +if (x_22 == 0) { -lean_object* x_24; uint8_t x_25; +lean_object* x_23; uint8_t x_24; lean_dec(x_21); lean_dec(x_12); lean_dec(x_10); lean_dec(x_6); lean_dec(x_5); lean_dec(x_1); -x_24 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__14___rarg(x_7); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_23 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabElabRulesAux___spec__1___rarg(x_7); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -return x_24; +return x_23; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_24, 0); -x_27 = lean_ctor_get(x_24, 1); -lean_inc(x_27); +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_23, 0); +x_26 = lean_ctor_get(x_23, 1); lean_inc(x_26); -lean_dec(x_24); -x_28 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -return x_28; +lean_inc(x_25); +lean_dec(x_23); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; } } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; -x_29 = l_Lean_Syntax_getArg(x_21, x_11); +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; +x_28 = l_Lean_Syntax_getArg(x_21, x_11); lean_dec(x_21); -x_30 = lean_unsigned_to_nat(3u); -x_31 = l_Lean_Syntax_getArg(x_10, x_30); -x_32 = l_Lean_Syntax_getArgs(x_29); -lean_dec(x_29); -x_33 = l_Lean_instInhabitedSyntax; -x_34 = lean_array_get(x_33, x_32, x_11); -x_35 = l_Lean_Syntax_isQuot(x_34); -if (x_35 == 0) +x_29 = lean_unsigned_to_nat(3u); +x_30 = l_Lean_Syntax_getArg(x_10, x_29); +x_31 = l_Lean_Syntax_getArgs(x_28); +lean_dec(x_28); +x_32 = l_Lean_instInhabitedSyntax; +x_33 = lean_array_get(x_32, x_31, x_11); +x_34 = l_Lean_Syntax_isQuot(x_33); +if (x_34 == 0) { -lean_object* x_36; uint8_t x_37; -lean_dec(x_34); -lean_dec(x_32); +lean_object* x_35; uint8_t x_36; +lean_dec(x_33); lean_dec(x_31); +lean_dec(x_30); lean_dec(x_12); lean_dec(x_10); lean_dec(x_6); lean_dec(x_5); lean_dec(x_1); -x_36 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_inferMacroRulesAltKind___spec__2___rarg(x_7); -x_37 = !lean_is_exclusive(x_36); -if (x_37 == 0) +x_35 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_inferMacroRulesAltKind___spec__2___rarg(x_7); +x_36 = !lean_is_exclusive(x_35); +if (x_36 == 0) { -return x_36; +return x_35; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_36, 0); -x_39 = lean_ctor_get(x_36, 1); -lean_inc(x_39); +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_35, 0); +x_38 = lean_ctor_get(x_35, 1); lean_inc(x_38); -lean_dec(x_36); -x_40 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -return x_40; +lean_inc(x_37); +lean_dec(x_35); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; } } else { -lean_object* x_41; lean_object* x_42; -x_41 = lean_box(0); +lean_object* x_40; lean_object* x_41; +x_40 = lean_box(0); lean_inc(x_6); lean_inc(x_5); lean_inc(x_1); -x_42 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2(x_34, x_32, x_31, x_13, x_1, x_10, x_41, x_5, x_6, x_7); -if (lean_obj_tag(x_42) == 0) +x_41 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2(x_33, x_31, x_30, x_13, x_1, x_10, x_40, x_5, x_6, x_7); +if (lean_obj_tag(x_41) == 0) { -lean_object* x_43; lean_object* x_44; size_t x_45; size_t x_46; lean_object* x_47; -x_43 = lean_ctor_get(x_42, 0); +lean_object* x_42; lean_object* x_43; size_t x_44; size_t x_45; lean_object* x_46; +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_41, 1); lean_inc(x_43); -x_44 = lean_ctor_get(x_42, 1); -lean_inc(x_44); -lean_dec(x_42); -x_45 = 1; -x_46 = lean_usize_add(x_3, x_45); -x_47 = lean_array_uset(x_12, x_3, x_43); -x_3 = x_46; -x_4 = x_47; -x_7 = x_44; +lean_dec(x_41); +x_44 = 1; +x_45 = lean_usize_add(x_3, x_44); +x_46 = lean_array_uset(x_12, x_3, x_42); +x_3 = x_45; +x_4 = x_46; +x_7 = x_43; goto _start; } else { -uint8_t x_49; +uint8_t x_48; lean_dec(x_12); lean_dec(x_6); lean_dec(x_5); lean_dec(x_1); -x_49 = !lean_is_exclusive(x_42); -if (x_49 == 0) +x_48 = !lean_is_exclusive(x_41); +if (x_48 == 0) { -return x_42; +return x_41; } else { -lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_50 = lean_ctor_get(x_42, 0); -x_51 = lean_ctor_get(x_42, 1); -lean_inc(x_51); +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_41, 0); +x_50 = lean_ctor_get(x_41, 1); lean_inc(x_50); -lean_dec(x_42); -x_52 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_52, 0, x_50); -lean_ctor_set(x_52, 1, x_51); -return x_52; +lean_inc(x_49); +lean_dec(x_41); +x_51 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +return x_51; } } } @@ -1053,7 +1217,7 @@ return x_52; } } } -LEAN_EXPORT lean_object* l_Lean_mkIdentFromRef___at_Lean_Elab_Command_elabElabRulesAux___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_mkIdentFromRef___at_Lean_Elab_Command_elabElabRulesAux___spec__6(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; @@ -1166,7 +1330,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabElabRulesAux___lambda__1___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__2; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__2; x_2 = l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__9; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -1220,7 +1384,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabElabRulesAux___lambda__1___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__6; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__6; x_2 = l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__15; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -1246,7 +1410,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabElabRulesAux___lambda__1___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__6; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__6; x_2 = l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__18; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -1264,7 +1428,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabElabRulesAux___lambda__1___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__4; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__4; x_2 = l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__20; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -1489,7 +1653,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabElabRulesAux___lambda__1___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__6; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__6; x_2 = l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__43; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -1507,7 +1671,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabElabRulesAux___lambda__1___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__6; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__6; x_2 = l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__45; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -1525,7 +1689,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabElabRulesAux___lambda__1___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__6; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__6; x_2 = l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__47; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -1627,7 +1791,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabElabRulesAux___lambda__1___clo _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__4; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__4; x_2 = l_Array_append___rarg(x_1, x_1); return x_2; } @@ -1637,7 +1801,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabElabRulesAux___lambda__1___clo { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__3; +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__3; x_3 = l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__58; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -1847,7 +2011,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabElabRulesAux___lambda__1___clo { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__10; -x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__5; +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } @@ -1906,7 +2070,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabElabRulesAux___lambda__1___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__6; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__6; x_2 = l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__84; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -1965,7 +2129,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabElabRulesAux___lambda__1___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__6; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__6; x_2 = l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__90; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -1976,8 +2140,8 @@ static lean_object* _init_l_Lean_Elab_Command_elabElabRulesAux___lambda__1___clo { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__3; -x_3 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__4; +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__3; +x_3 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__4; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1997,7 +2161,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabElabRulesAux___lambda__1___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__6; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__6; x_2 = l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__93; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -2117,7 +2281,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabElabRulesAux___lambda__1___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__6; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__6; x_2 = l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__106; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -2227,7 +2391,7 @@ x_17 = l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__8; x_18 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_18, 0, x_17); lean_ctor_set(x_18, 1, x_16); -x_19 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__12; +x_19 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__13; x_20 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_20, 0, x_18); lean_ctor_set(x_20, 1, x_19); @@ -2240,7 +2404,7 @@ else 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; lean_object* x_30; lean_object* x_31; uint8_t x_32; lean_dec(x_6); lean_inc(x_2); -x_22 = l_Lean_mkIdentFromRef___at_Lean_Elab_Command_elabElabRulesAux___spec__3(x_2, x_7, x_8, x_9); +x_22 = l_Lean_mkIdentFromRef___at_Lean_Elab_Command_elabElabRulesAux___spec__6(x_2, x_7, x_8, x_9); x_23 = lean_ctor_get(x_22, 0); lean_inc(x_23); x_24 = lean_ctor_get(x_22, 1); @@ -2282,10 +2446,10 @@ lean_ctor_set(x_39, 0, x_26); lean_ctor_set(x_39, 1, x_38); lean_ctor_set(x_39, 2, x_36); lean_ctor_set(x_39, 3, x_37); -x_40 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__6; +x_40 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__5; x_41 = lean_array_push(x_40, x_23); x_42 = lean_box(2); -x_43 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__3; +x_43 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__3; x_44 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_44, 0, x_42); lean_ctor_set(x_44, 1, x_43); @@ -2379,9 +2543,9 @@ lean_inc(x_26); x_86 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_86, 0, x_26); lean_ctor_set(x_86, 1, x_85); -x_87 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__4; +x_87 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__4; x_88 = l_Array_append___rarg(x_87, x_4); -x_89 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__1; +x_89 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__1; lean_inc(x_26); x_90 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_90, 0, x_26); @@ -2407,7 +2571,7 @@ x_99 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_99, 0, x_42); lean_ctor_set(x_99, 1, x_43); lean_ctor_set(x_99, 2, x_98); -x_100 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__7; +x_100 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__6; lean_inc(x_26); x_101 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_101, 0, x_26); @@ -2421,12 +2585,12 @@ lean_ctor_set(x_106, 0, x_26); lean_ctor_set(x_106, 1, x_104); lean_ctor_set(x_106, 2, x_103); lean_ctor_set(x_106, 3, x_105); -x_107 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__8; +x_107 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__7; x_108 = lean_array_push(x_107, x_90); x_109 = lean_array_push(x_108, x_99); x_110 = lean_array_push(x_109, x_101); x_111 = lean_array_push(x_110, x_106); -x_112 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__8; +x_112 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__8; x_113 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_113, 0, x_42); lean_ctor_set(x_113, 1, x_112); @@ -2522,10 +2686,10 @@ lean_ctor_set(x_155, 0, x_26); lean_ctor_set(x_155, 1, x_154); lean_ctor_set(x_155, 2, x_152); lean_ctor_set(x_155, 3, x_153); -x_156 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__6; +x_156 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__5; x_157 = lean_array_push(x_156, x_23); x_158 = lean_box(2); -x_159 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__3; +x_159 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__3; x_160 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_160, 0, x_158); lean_ctor_set(x_160, 1, x_159); @@ -2619,9 +2783,9 @@ lean_inc(x_26); x_202 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_202, 0, x_26); lean_ctor_set(x_202, 1, x_201); -x_203 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__4; +x_203 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__4; x_204 = l_Array_append___rarg(x_203, x_4); -x_205 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__1; +x_205 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__1; lean_inc(x_26); x_206 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_206, 0, x_26); @@ -2647,7 +2811,7 @@ x_215 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_215, 0, x_158); lean_ctor_set(x_215, 1, x_159); lean_ctor_set(x_215, 2, x_214); -x_216 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__7; +x_216 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__6; lean_inc(x_26); x_217 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_217, 0, x_26); @@ -2661,12 +2825,12 @@ lean_ctor_set(x_222, 0, x_26); lean_ctor_set(x_222, 1, x_220); lean_ctor_set(x_222, 2, x_219); lean_ctor_set(x_222, 3, x_221); -x_223 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__8; +x_223 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__7; x_224 = lean_array_push(x_223, x_206); x_225 = lean_array_push(x_224, x_215); x_226 = lean_array_push(x_225, x_217); x_227 = lean_array_push(x_226, x_222); -x_228 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__8; +x_228 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__8; x_229 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_229, 0, x_158); lean_ctor_set(x_229, 1, x_228); @@ -2749,7 +2913,7 @@ else lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; uint8_t x_276; lean_dec(x_6); lean_inc(x_2); -x_266 = l_Lean_mkIdentFromRef___at_Lean_Elab_Command_elabElabRulesAux___spec__3(x_2, x_7, x_8, x_9); +x_266 = l_Lean_mkIdentFromRef___at_Lean_Elab_Command_elabElabRulesAux___spec__6(x_2, x_7, x_8, x_9); x_267 = lean_ctor_get(x_266, 0); lean_inc(x_267); x_268 = lean_ctor_get(x_266, 1); @@ -2792,10 +2956,10 @@ lean_ctor_set(x_284, 0, x_270); lean_ctor_set(x_284, 1, x_283); lean_ctor_set(x_284, 2, x_281); lean_ctor_set(x_284, 3, x_282); -x_285 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__6; +x_285 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__5; x_286 = lean_array_push(x_285, x_267); x_287 = lean_box(2); -x_288 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__3; +x_288 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__3; x_289 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_289, 0, x_287); lean_ctor_set(x_289, 1, x_288); @@ -2889,9 +3053,9 @@ lean_inc(x_270); x_331 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_331, 0, x_270); lean_ctor_set(x_331, 1, x_330); -x_332 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__4; +x_332 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__4; x_333 = l_Array_append___rarg(x_332, x_4); -x_334 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__1; +x_334 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__1; lean_inc(x_270); x_335 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_335, 0, x_270); @@ -2917,7 +3081,7 @@ x_344 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_344, 0, x_287); lean_ctor_set(x_344, 1, x_288); lean_ctor_set(x_344, 2, x_343); -x_345 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__7; +x_345 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__6; lean_inc(x_270); x_346 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_346, 0, x_270); @@ -2931,12 +3095,12 @@ lean_ctor_set(x_351, 0, x_270); lean_ctor_set(x_351, 1, x_349); lean_ctor_set(x_351, 2, x_348); lean_ctor_set(x_351, 3, x_350); -x_352 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__8; +x_352 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__7; x_353 = lean_array_push(x_352, x_335); x_354 = lean_array_push(x_353, x_344); x_355 = lean_array_push(x_354, x_346); x_356 = lean_array_push(x_355, x_351); -x_357 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__8; +x_357 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__8; x_358 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_358, 0, x_287); lean_ctor_set(x_358, 1, x_357); @@ -3033,10 +3197,10 @@ lean_ctor_set(x_401, 0, x_270); lean_ctor_set(x_401, 1, x_400); lean_ctor_set(x_401, 2, x_398); lean_ctor_set(x_401, 3, x_399); -x_402 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__6; +x_402 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__5; x_403 = lean_array_push(x_402, x_267); x_404 = lean_box(2); -x_405 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__3; +x_405 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__3; x_406 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_406, 0, x_404); lean_ctor_set(x_406, 1, x_405); @@ -3130,9 +3294,9 @@ lean_inc(x_270); x_448 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_448, 0, x_270); lean_ctor_set(x_448, 1, x_447); -x_449 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__4; +x_449 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__4; x_450 = l_Array_append___rarg(x_449, x_4); -x_451 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__1; +x_451 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__1; lean_inc(x_270); x_452 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_452, 0, x_270); @@ -3158,7 +3322,7 @@ x_461 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_461, 0, x_404); lean_ctor_set(x_461, 1, x_405); lean_ctor_set(x_461, 2, x_460); -x_462 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__7; +x_462 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__6; lean_inc(x_270); x_463 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_463, 0, x_270); @@ -3172,12 +3336,12 @@ lean_ctor_set(x_468, 0, x_270); lean_ctor_set(x_468, 1, x_466); lean_ctor_set(x_468, 2, x_465); lean_ctor_set(x_468, 3, x_467); -x_469 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__8; +x_469 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__7; x_470 = lean_array_push(x_469, x_452); x_471 = lean_array_push(x_470, x_461); x_472 = lean_array_push(x_471, x_463); x_473 = lean_array_push(x_472, x_468); -x_474 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__8; +x_474 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__8; x_475 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_475, 0, x_404); lean_ctor_set(x_475, 1, x_474); @@ -3260,7 +3424,7 @@ else lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; uint8_t x_522; lean_dec(x_6); lean_inc(x_2); -x_512 = l_Lean_mkIdentFromRef___at_Lean_Elab_Command_elabElabRulesAux___spec__3(x_2, x_7, x_8, x_9); +x_512 = l_Lean_mkIdentFromRef___at_Lean_Elab_Command_elabElabRulesAux___spec__6(x_2, x_7, x_8, x_9); x_513 = lean_ctor_get(x_512, 0); lean_inc(x_513); x_514 = lean_ctor_get(x_512, 1); @@ -3303,10 +3467,10 @@ lean_ctor_set(x_530, 0, x_516); lean_ctor_set(x_530, 1, x_529); lean_ctor_set(x_530, 2, x_527); lean_ctor_set(x_530, 3, x_528); -x_531 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__6; +x_531 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__5; x_532 = lean_array_push(x_531, x_513); x_533 = lean_box(2); -x_534 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__3; +x_534 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__3; x_535 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_535, 0, x_533); lean_ctor_set(x_535, 1, x_534); @@ -3430,7 +3594,7 @@ x_589 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_589, 0, x_533); lean_ctor_set(x_589, 1, x_534); lean_ctor_set(x_589, 2, x_588); -x_590 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__7; +x_590 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__6; lean_inc(x_516); x_591 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_591, 0, x_516); @@ -3457,9 +3621,9 @@ lean_inc(x_516); x_601 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_601, 0, x_516); lean_ctor_set(x_601, 1, x_600); -x_602 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__4; +x_602 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__4; x_603 = l_Array_append___rarg(x_602, x_4); -x_604 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__1; +x_604 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__1; lean_inc(x_516); x_605 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_605, 0, x_516); @@ -3483,13 +3647,13 @@ lean_ctor_set(x_614, 0, x_516); lean_ctor_set(x_614, 1, x_612); lean_ctor_set(x_614, 2, x_611); lean_ctor_set(x_614, 3, x_613); -x_615 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__8; +x_615 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__7; x_616 = lean_array_push(x_615, x_605); x_617 = lean_array_push(x_616, x_609); lean_inc(x_591); x_618 = lean_array_push(x_617, x_591); x_619 = lean_array_push(x_618, x_614); -x_620 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__8; +x_620 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__8; x_621 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_621, 0, x_533); lean_ctor_set(x_621, 1, x_620); @@ -3607,10 +3771,10 @@ lean_ctor_set(x_679, 0, x_516); lean_ctor_set(x_679, 1, x_678); lean_ctor_set(x_679, 2, x_676); lean_ctor_set(x_679, 3, x_677); -x_680 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__6; +x_680 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__5; x_681 = lean_array_push(x_680, x_513); x_682 = lean_box(2); -x_683 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__3; +x_683 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__3; x_684 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_684, 0, x_682); lean_ctor_set(x_684, 1, x_683); @@ -3734,7 +3898,7 @@ x_738 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_738, 0, x_682); lean_ctor_set(x_738, 1, x_683); lean_ctor_set(x_738, 2, x_737); -x_739 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__7; +x_739 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__6; lean_inc(x_516); x_740 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_740, 0, x_516); @@ -3761,9 +3925,9 @@ lean_inc(x_516); x_750 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_750, 0, x_516); lean_ctor_set(x_750, 1, x_749); -x_751 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__4; +x_751 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__4; x_752 = l_Array_append___rarg(x_751, x_4); -x_753 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__1; +x_753 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__1; lean_inc(x_516); x_754 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_754, 0, x_516); @@ -3787,13 +3951,13 @@ lean_ctor_set(x_763, 0, x_516); lean_ctor_set(x_763, 1, x_761); lean_ctor_set(x_763, 2, x_760); lean_ctor_set(x_763, 3, x_762); -x_764 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__8; +x_764 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__7; x_765 = lean_array_push(x_764, x_754); x_766 = lean_array_push(x_765, x_758); lean_inc(x_740); x_767 = lean_array_push(x_766, x_740); x_768 = lean_array_push(x_767, x_763); -x_769 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__8; +x_769 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__8; x_770 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_770, 0, x_682); lean_ctor_set(x_770, 1, x_769); @@ -3925,7 +4089,7 @@ else lean_object* x_831; lean_object* x_832; lean_object* x_833; lean_object* x_834; lean_object* x_835; lean_object* x_836; lean_object* x_837; lean_object* x_838; lean_object* x_839; lean_object* x_840; uint8_t x_841; lean_dec(x_6); lean_inc(x_2); -x_831 = l_Lean_mkIdentFromRef___at_Lean_Elab_Command_elabElabRulesAux___spec__3(x_2, x_7, x_8, x_9); +x_831 = l_Lean_mkIdentFromRef___at_Lean_Elab_Command_elabElabRulesAux___spec__6(x_2, x_7, x_8, x_9); x_832 = lean_ctor_get(x_831, 0); lean_inc(x_832); x_833 = lean_ctor_get(x_831, 1); @@ -3968,10 +4132,10 @@ lean_ctor_set(x_849, 0, x_835); lean_ctor_set(x_849, 1, x_848); lean_ctor_set(x_849, 2, x_846); lean_ctor_set(x_849, 3, x_847); -x_850 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__6; +x_850 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__5; x_851 = lean_array_push(x_850, x_832); x_852 = lean_box(2); -x_853 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__3; +x_853 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__3; x_854 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_854, 0, x_852); lean_ctor_set(x_854, 1, x_853); @@ -4095,7 +4259,7 @@ x_907 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_907, 0, x_852); lean_ctor_set(x_907, 1, x_853); lean_ctor_set(x_907, 2, x_906); -x_908 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__7; +x_908 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__6; lean_inc(x_835); x_909 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_909, 0, x_835); @@ -4139,9 +4303,9 @@ lean_inc(x_835); x_926 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_926, 0, x_835); lean_ctor_set(x_926, 1, x_925); -x_927 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__4; +x_927 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__4; x_928 = l_Array_append___rarg(x_927, x_4); -x_929 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__1; +x_929 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__1; lean_inc(x_835); x_930 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_930, 0, x_835); @@ -4176,13 +4340,13 @@ lean_ctor_set(x_944, 0, x_835); lean_ctor_set(x_944, 1, x_942); lean_ctor_set(x_944, 2, x_941); lean_ctor_set(x_944, 3, x_943); -x_945 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__8; +x_945 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__7; x_946 = lean_array_push(x_945, x_930); x_947 = lean_array_push(x_946, x_939); lean_inc(x_909); x_948 = lean_array_push(x_947, x_909); x_949 = lean_array_push(x_948, x_944); -x_950 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__8; +x_950 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__8; x_951 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_951, 0, x_852); lean_ctor_set(x_951, 1, x_950); @@ -4327,10 +4491,10 @@ lean_ctor_set(x_1022, 0, x_835); lean_ctor_set(x_1022, 1, x_1021); lean_ctor_set(x_1022, 2, x_1019); lean_ctor_set(x_1022, 3, x_1020); -x_1023 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__6; +x_1023 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__5; x_1024 = lean_array_push(x_1023, x_832); x_1025 = lean_box(2); -x_1026 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__3; +x_1026 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__3; x_1027 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_1027, 0, x_1025); lean_ctor_set(x_1027, 1, x_1026); @@ -4454,7 +4618,7 @@ x_1080 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_1080, 0, x_1025); lean_ctor_set(x_1080, 1, x_1026); lean_ctor_set(x_1080, 2, x_1079); -x_1081 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__7; +x_1081 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__6; lean_inc(x_835); x_1082 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_1082, 0, x_835); @@ -4498,9 +4662,9 @@ lean_inc(x_835); x_1099 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_1099, 0, x_835); lean_ctor_set(x_1099, 1, x_1098); -x_1100 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__4; +x_1100 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__4; x_1101 = l_Array_append___rarg(x_1100, x_4); -x_1102 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__1; +x_1102 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__1; lean_inc(x_835); x_1103 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_1103, 0, x_835); @@ -4535,13 +4699,13 @@ lean_ctor_set(x_1117, 0, x_835); lean_ctor_set(x_1117, 1, x_1115); lean_ctor_set(x_1117, 2, x_1114); lean_ctor_set(x_1117, 3, x_1116); -x_1118 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__8; +x_1118 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__7; x_1119 = lean_array_push(x_1118, x_1103); x_1120 = lean_array_push(x_1119, x_1112); lean_inc(x_1082); x_1121 = lean_array_push(x_1120, x_1082); x_1122 = lean_array_push(x_1121, x_1117); -x_1123 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__8; +x_1123 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__8; x_1124 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_1124, 0, x_1025); lean_ctor_set(x_1124, 1, x_1123); @@ -4697,7 +4861,7 @@ x_12 = 0; lean_inc(x_8); lean_inc(x_7); lean_inc(x_3); -x_13 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2(x_3, x_11, x_12, x_6, x_7, x_8, x_9); +x_13 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5(x_3, x_11, x_12, x_6, x_7, x_8, x_9); if (lean_obj_tag(x_13) == 0) { if (lean_obj_tag(x_4) == 0) @@ -4794,7 +4958,26 @@ return x_33; } } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabElabRulesAux___spec__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_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabElabRulesAux___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabElabRulesAux___spec__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabElabRulesAux___spec__3___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_throwError___at_Lean_Elab_Command_elabElabRulesAux___spec__3(x_1, x_2, x_3, x_4); +lean_dec(x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabElabRulesAux___spec__4___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: { size_t x_7; size_t x_8; lean_object* x_9; @@ -4802,7 +4985,7 @@ x_7 = lean_unbox_usize(x_4); lean_dec(x_4); x_8 = lean_unbox_usize(x_5); lean_dec(x_5); -x_9 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabElabRulesAux___spec__1(x_1, x_2, x_3, x_7, x_8, x_6); +x_9 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabElabRulesAux___spec__4(x_1, x_2, x_3, x_7, x_8, x_6); lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); @@ -4810,26 +4993,26 @@ lean_dec(x_1); return x_9; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__1(x_1, x_2); +x_3 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__1(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___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, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___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, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; -x_11 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_11 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_7); return x_11; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__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, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___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_8; size_t x_9; lean_object* x_10; @@ -4837,15 +5020,15 @@ x_8 = lean_unbox_usize(x_2); lean_dec(x_2); x_9 = lean_unbox_usize(x_3); lean_dec(x_3); -x_10 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2(x_1, x_8, x_9, x_4, x_5, x_6, x_7); +x_10 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5(x_1, x_8, x_9, x_4, x_5, x_6, x_7); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_mkIdentFromRef___at_Lean_Elab_Command_elabElabRulesAux___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_mkIdentFromRef___at_Lean_Elab_Command_elabElabRulesAux___spec__6___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_mkIdentFromRef___at_Lean_Elab_Command_elabElabRulesAux___spec__3(x_1, x_2, x_3, x_4); +x_5 = l_Lean_mkIdentFromRef___at_Lean_Elab_Command_elabElabRulesAux___spec__6(x_1, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_2); return x_5; @@ -4940,14 +5123,13 @@ x_12 = l_Lean_Syntax_getArg(x_1, x_11); x_13 = l_Lean_Syntax_isNone(x_12); if (x_13 == 0) { -lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_14 = l_Lean_nullKind; -x_15 = lean_unsigned_to_nat(2u); +lean_object* x_14; uint8_t x_15; +x_14 = lean_unsigned_to_nat(2u); lean_inc(x_12); -x_16 = l_Lean_Syntax_isNodeOf(x_12, x_14, x_15); -if (x_16 == 0) +x_15 = l_Lean_Syntax_matchesNull(x_12, x_14); +if (x_15 == 0) { -lean_object* x_17; +lean_object* x_16; lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); @@ -4955,30 +5137,30 @@ lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_17 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__14___rarg(x_10); -return x_17; +x_16 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__14___rarg(x_10); +return x_16; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_18 = lean_unsigned_to_nat(1u); -x_19 = l_Lean_Syntax_getArg(x_12, x_18); +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_17 = lean_unsigned_to_nat(1u); +x_18 = l_Lean_Syntax_getArg(x_12, x_17); lean_dec(x_12); -x_20 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_20, 0, x_19); -x_21 = lean_box(0); -x_22 = l_Lean_Elab_Command_elabElabRules___lambda__1(x_1, x_2, x_3, x_4, x_5, x_7, x_21, x_20, x_8, x_9, x_10); -return x_22; +x_19 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_19, 0, x_18); +x_20 = lean_box(0); +x_21 = l_Lean_Elab_Command_elabElabRules___lambda__1(x_1, x_2, x_3, x_4, x_5, x_7, x_20, x_19, x_8, x_9, x_10); +return x_21; } } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_dec(x_12); +x_22 = lean_box(0); x_23 = lean_box(0); -x_24 = lean_box(0); -x_25 = l_Lean_Elab_Command_elabElabRules___lambda__1(x_1, x_2, x_3, x_4, x_5, x_7, x_24, x_23, x_8, x_9, x_10); -return x_25; +x_24 = l_Lean_Elab_Command_elabElabRules___lambda__1(x_1, x_2, x_3, x_4, x_5, x_7, x_23, x_22, x_8, x_9, x_10); +return x_24; } } } @@ -5070,15 +5252,15 @@ lean_inc(x_13); x_21 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_21, 0, x_13); lean_ctor_set(x_21, 1, x_20); -x_22 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__4; +x_22 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__4; x_23 = l_Array_append___rarg(x_22, x_8); x_24 = lean_box(2); -x_25 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__3; +x_25 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__3; x_26 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_26, 0, x_24); lean_ctor_set(x_26, 1, x_25); lean_ctor_set(x_26, 2, x_23); -x_27 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__6; +x_27 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__5; x_28 = lean_array_push(x_27, x_26); x_29 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_29, 0, x_24); @@ -5292,9 +5474,51 @@ lean_closure_set(x_21, 4, x_5); lean_closure_set(x_21, 5, x_6); x_22 = l_Lean_Elab_Command_elabElabRules___lambda__3___closed__1; x_23 = l_Lean_Elab_Command_expandNoKindMacroRulesAux(x_20, x_22, x_21, x_9, x_10, x_11); -lean_dec(x_20); +if (lean_obj_tag(x_23) == 0) +{ +uint8_t x_24; +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) +{ +return x_23; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_23, 0); +x_26 = lean_ctor_get(x_23, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_23); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +else +{ +uint8_t x_28; +x_28 = !lean_is_exclusive(x_23); +if (x_28 == 0) +{ return x_23; } +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_23, 0); +x_30 = lean_ctor_get(x_23, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_23); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +return x_31; +} +} +} } } LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabElabRules___lambda__5(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_object* x_9, lean_object* x_10) { @@ -5307,14 +5531,13 @@ x_12 = l_Lean_Syntax_getArg(x_1, x_11); x_13 = l_Lean_Syntax_isNone(x_12); if (x_13 == 0) { -lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_14 = l_Lean_nullKind; -x_15 = lean_unsigned_to_nat(2u); +lean_object* x_14; uint8_t x_15; +x_14 = lean_unsigned_to_nat(2u); lean_inc(x_12); -x_16 = l_Lean_Syntax_isNodeOf(x_12, x_14, x_15); -if (x_16 == 0) +x_15 = l_Lean_Syntax_matchesNull(x_12, x_14); +if (x_15 == 0) { -lean_object* x_17; +lean_object* x_16; lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); @@ -5323,30 +5546,30 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_17 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__14___rarg(x_10); -return x_17; +x_16 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__14___rarg(x_10); +return x_16; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_18 = lean_unsigned_to_nat(1u); -x_19 = l_Lean_Syntax_getArg(x_12, x_18); +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_17 = lean_unsigned_to_nat(1u); +x_18 = l_Lean_Syntax_getArg(x_12, x_17); lean_dec(x_12); -x_20 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_20, 0, x_19); -x_21 = lean_box(0); -x_22 = l_Lean_Elab_Command_elabElabRules___lambda__4(x_1, x_2, x_3, x_4, x_7, x_5, x_21, x_20, x_8, x_9, x_10); -return x_22; +x_19 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_19, 0, x_18); +x_20 = lean_box(0); +x_21 = l_Lean_Elab_Command_elabElabRules___lambda__4(x_1, x_2, x_3, x_4, x_7, x_5, x_20, x_19, x_8, x_9, x_10); +return x_21; } } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_dec(x_12); +x_22 = lean_box(0); x_23 = lean_box(0); -x_24 = lean_box(0); -x_25 = l_Lean_Elab_Command_elabElabRules___lambda__4(x_1, x_2, x_3, x_4, x_7, x_5, x_24, x_23, x_8, x_9, x_10); -return x_25; +x_24 = l_Lean_Elab_Command_elabElabRules___lambda__4(x_1, x_2, x_3, x_4, x_7, x_5, x_23, x_22, x_8, x_9, x_10); +return x_24; } } } @@ -5365,7 +5588,7 @@ lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_ lean_dec(x_4); x_9 = lean_unsigned_to_nat(1u); x_10 = l_Lean_Syntax_getArg(x_1, x_9); -x_11 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__5; +x_11 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__5; x_12 = lean_name_mk_string(x_2, x_11); x_13 = l_Lean_Elab_Command_elabElabRules___lambda__6___closed__1; lean_inc(x_12); @@ -5387,130 +5610,129 @@ return x_16; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; x_17 = lean_unsigned_to_nat(3u); x_18 = l_Lean_Syntax_getArg(x_1, x_17); -x_19 = l_Lean_nullKind; -x_20 = lean_unsigned_to_nat(0u); +x_19 = lean_unsigned_to_nat(0u); lean_inc(x_18); -x_21 = l_Lean_Syntax_isNodeOf(x_18, x_19, x_20); -if (x_21 == 0) +x_20 = l_Lean_Syntax_matchesNull(x_18, x_19); +if (x_20 == 0) { -lean_object* x_22; uint8_t x_23; +lean_object* x_21; uint8_t x_22; lean_dec(x_3); -x_22 = lean_unsigned_to_nat(5u); +x_21 = lean_unsigned_to_nat(5u); lean_inc(x_18); -x_23 = l_Lean_Syntax_isNodeOf(x_18, x_19, x_22); -if (x_23 == 0) +x_22 = l_Lean_Syntax_matchesNull(x_18, x_21); +if (x_22 == 0) { -lean_object* x_24; +lean_object* x_23; lean_dec(x_18); lean_dec(x_12); lean_dec(x_10); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_24 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__14___rarg(x_8); -return x_24; +x_23 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__14___rarg(x_8); +return x_23; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; -x_25 = l_Lean_Syntax_getArg(x_18, x_17); +lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_24 = l_Lean_Syntax_getArg(x_18, x_17); lean_dec(x_18); -x_26 = lean_unsigned_to_nat(4u); -x_27 = l_Lean_Syntax_getArg(x_1, x_26); -x_28 = l_Lean_Syntax_isNone(x_27); -if (x_28 == 0) +x_25 = lean_unsigned_to_nat(4u); +x_26 = l_Lean_Syntax_getArg(x_1, x_25); +x_27 = l_Lean_Syntax_isNone(x_26); +if (x_27 == 0) { -lean_object* x_29; uint8_t x_30; -x_29 = lean_unsigned_to_nat(2u); -lean_inc(x_27); -x_30 = l_Lean_Syntax_isNodeOf(x_27, x_19, x_29); -if (x_30 == 0) +lean_object* x_28; uint8_t x_29; +x_28 = lean_unsigned_to_nat(2u); +lean_inc(x_26); +x_29 = l_Lean_Syntax_matchesNull(x_26, x_28); +if (x_29 == 0) { -lean_object* x_31; -lean_dec(x_27); -lean_dec(x_25); +lean_object* x_30; +lean_dec(x_26); +lean_dec(x_24); lean_dec(x_12); lean_dec(x_10); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_31 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__14___rarg(x_8); -return x_31; +x_30 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__14___rarg(x_8); +return x_30; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_32 = l_Lean_Syntax_getArg(x_27, x_9); -lean_dec(x_27); -x_33 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_33, 0, x_32); -x_34 = lean_box(0); -x_35 = l_Lean_Elab_Command_elabElabRules___lambda__2(x_1, x_12, x_25, x_5, x_10, x_34, x_33, x_6, x_7, x_8); -lean_dec(x_25); -return x_35; +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_31 = l_Lean_Syntax_getArg(x_26, x_9); +lean_dec(x_26); +x_32 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_32, 0, x_31); +x_33 = lean_box(0); +x_34 = l_Lean_Elab_Command_elabElabRules___lambda__2(x_1, x_12, x_24, x_5, x_10, x_33, x_32, x_6, x_7, x_8); +lean_dec(x_24); +return x_34; } } else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; -lean_dec(x_27); +lean_object* x_35; lean_object* x_36; lean_object* x_37; +lean_dec(x_26); +x_35 = lean_box(0); x_36 = lean_box(0); -x_37 = lean_box(0); -x_38 = l_Lean_Elab_Command_elabElabRules___lambda__2(x_1, x_12, x_25, x_5, x_10, x_37, x_36, x_6, x_7, x_8); -lean_dec(x_25); -return x_38; +x_37 = l_Lean_Elab_Command_elabElabRules___lambda__2(x_1, x_12, x_24, x_5, x_10, x_36, x_35, x_6, x_7, x_8); +lean_dec(x_24); +return x_37; } } } else { -lean_object* x_39; lean_object* x_40; uint8_t x_41; +lean_object* x_38; lean_object* x_39; uint8_t x_40; lean_dec(x_18); -x_39 = lean_unsigned_to_nat(4u); -x_40 = l_Lean_Syntax_getArg(x_1, x_39); -x_41 = l_Lean_Syntax_isNone(x_40); -if (x_41 == 0) +x_38 = lean_unsigned_to_nat(4u); +x_39 = l_Lean_Syntax_getArg(x_1, x_38); +x_40 = l_Lean_Syntax_isNone(x_39); +if (x_40 == 0) { -lean_object* x_42; uint8_t x_43; -x_42 = lean_unsigned_to_nat(2u); -lean_inc(x_40); -x_43 = l_Lean_Syntax_isNodeOf(x_40, x_19, x_42); -if (x_43 == 0) +lean_object* x_41; uint8_t x_42; +x_41 = lean_unsigned_to_nat(2u); +lean_inc(x_39); +x_42 = l_Lean_Syntax_matchesNull(x_39, x_41); +if (x_42 == 0) { -lean_object* x_44; -lean_dec(x_40); +lean_object* x_43; +lean_dec(x_39); lean_dec(x_12); lean_dec(x_10); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_44 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__14___rarg(x_8); -return x_44; +x_43 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__14___rarg(x_8); +return x_43; } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_45 = l_Lean_Syntax_getArg(x_40, x_9); -lean_dec(x_40); -x_46 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_46, 0, x_45); -x_47 = lean_box(0); -x_48 = l_Lean_Elab_Command_elabElabRules___lambda__5(x_1, x_12, x_10, x_3, x_5, x_47, x_46, x_6, x_7, x_8); -return x_48; +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_44 = l_Lean_Syntax_getArg(x_39, x_9); +lean_dec(x_39); +x_45 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_45, 0, x_44); +x_46 = lean_box(0); +x_47 = l_Lean_Elab_Command_elabElabRules___lambda__5(x_1, x_12, x_10, x_3, x_5, x_46, x_45, x_6, x_7, x_8); +return x_47; } } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; -lean_dec(x_40); +lean_object* x_48; lean_object* x_49; lean_object* x_50; +lean_dec(x_39); +x_48 = lean_box(0); x_49 = lean_box(0); -x_50 = lean_box(0); -x_51 = l_Lean_Elab_Command_elabElabRules___lambda__5(x_1, x_12, x_10, x_3, x_5, x_50, x_49, x_6, x_7, x_8); -return x_51; +x_50 = l_Lean_Elab_Command_elabElabRules___lambda__5(x_1, x_12, x_10, x_3, x_5, x_49, x_48, x_6, x_7, x_8); +return x_50; } } } @@ -5520,7 +5742,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabElabRules___lambda__7___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__4; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__4; x_2 = l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__11; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -5578,62 +5800,61 @@ x_9 = l_Lean_Syntax_getArg(x_1, x_8); x_10 = l_Lean_Syntax_isNone(x_9); if (x_10 == 0) { -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = l_Lean_nullKind; -x_12 = lean_unsigned_to_nat(1u); +lean_object* x_11; uint8_t x_12; +x_11 = lean_unsigned_to_nat(1u); lean_inc(x_9); -x_13 = l_Lean_Syntax_isNodeOf(x_9, x_11, x_12); -if (x_13 == 0) +x_12 = l_Lean_Syntax_matchesNull(x_9, x_11); +if (x_12 == 0) { -lean_object* x_14; +lean_object* x_13; lean_dec(x_9); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_14 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__14___rarg(x_4); -return x_14; +x_13 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__14___rarg(x_4); +return x_13; } else { -lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_15 = l_Lean_Syntax_getArg(x_9, x_8); +lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_14 = l_Lean_Syntax_getArg(x_9, x_8); lean_dec(x_9); -x_16 = l_Lean_Elab_Command_elabElabRules___lambda__7___closed__4; -lean_inc(x_15); -x_17 = l_Lean_Syntax_isOfKind(x_15, x_16); -if (x_17 == 0) +x_15 = l_Lean_Elab_Command_elabElabRules___lambda__7___closed__4; +lean_inc(x_14); +x_16 = l_Lean_Syntax_isOfKind(x_14, x_15); +if (x_16 == 0) { -lean_object* x_18; -lean_dec(x_15); +lean_object* x_17; +lean_dec(x_14); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_18 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__14___rarg(x_4); -return x_18; +x_17 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__14___rarg(x_4); +return x_17; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_19 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_19, 0, x_15); -x_20 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__4; -x_21 = lean_box(0); -x_22 = l_Lean_Elab_Command_elabElabRules___lambda__6(x_1, x_20, x_5, x_21, x_19, x_2, x_3, x_4); +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_18 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_18, 0, x_14); +x_19 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__4; +x_20 = lean_box(0); +x_21 = l_Lean_Elab_Command_elabElabRules___lambda__6(x_1, x_19, x_5, x_20, x_18, x_2, x_3, x_4); lean_dec(x_1); -return x_22; +return x_21; } } } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_dec(x_9); -x_23 = lean_box(0); -x_24 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__4; -x_25 = lean_box(0); -x_26 = l_Lean_Elab_Command_elabElabRules___lambda__6(x_1, x_24, x_5, x_25, x_23, x_2, x_3, x_4); +x_22 = lean_box(0); +x_23 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__4; +x_24 = lean_box(0); +x_25 = l_Lean_Elab_Command_elabElabRules___lambda__6(x_1, x_23, x_5, x_24, x_22, x_2, x_3, x_4); lean_dec(x_1); -return x_26; +return x_25; } } } @@ -5764,7 +5985,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabElabRules_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(70u); +x_1 = lean_unsigned_to_nat(71u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -5776,7 +5997,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabElabRules_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(77u); +x_1 = lean_unsigned_to_nat(78u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -5804,7 +6025,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabElabRules_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(70u); +x_1 = lean_unsigned_to_nat(71u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -5816,7 +6037,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabElabRules_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(70u); +x_1 = lean_unsigned_to_nat(71u); x_2 = lean_unsigned_to_nat(52u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -5930,6 +6151,30 @@ return x_22; } } } +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElab___spec__2(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = 1; +x_9 = lean_usize_add(x_2, x_8); +x_10 = lean_array_uset(x_7, x_2, x_5); +x_2 = x_9; +x_3 = x_10; +goto _start; +} +} +} static lean_object* _init_l_Lean_Elab_Command_elabElab___lambda__1___closed__1() { _start: { @@ -5983,7 +6228,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabElab___lambda__1___closed__7() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__6; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__5; x_2 = l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__92; x_3 = lean_array_push(x_1, x_2); return x_3; @@ -6016,7 +6261,7 @@ return x_1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabElab___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, lean_object* x_6, lean_object* x_7, 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, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { _start: { -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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; size_t x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; +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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; size_t x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; x_19 = l_Lean_Elab_Command_getScope___rarg(x_17, x_18); x_20 = lean_ctor_get(x_19, 0); lean_inc(x_20); @@ -6034,7 +6279,7 @@ x_25 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_25, 0, x_24); lean_ctor_set(x_25, 1, x_23); lean_ctor_set(x_25, 2, x_1); -x_26 = l_Lean_mkIdentFromRef___at_Lean_Elab_Command_elabElabRulesAux___spec__3(x_15, x_16, x_17, x_21); +x_26 = l_Lean_mkIdentFromRef___at_Lean_Elab_Command_elabElabRulesAux___spec__6(x_15, x_16, x_17, x_21); x_27 = lean_ctor_get(x_26, 0); lean_inc(x_27); x_28 = lean_ctor_get(x_26, 1); @@ -6097,9 +6342,9 @@ x_55 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_55, 0, x_24); lean_ctor_set(x_55, 1, x_40); lean_ctor_set(x_55, 2, x_54); -x_56 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__6; +x_56 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__5; x_57 = lean_array_push(x_56, x_55); -x_58 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__3; +x_58 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__3; x_59 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_59, 0, x_24); lean_ctor_set(x_59, 1, x_58); @@ -6113,245 +6358,244 @@ x_63 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_63, 0, x_30); lean_ctor_set(x_63, 1, x_62); x_64 = l_Nat_repr(x_3); -x_65 = l_Lean_numLitKind; -x_66 = l_Lean_Syntax_mkLit(x_65, x_64, x_24); -x_67 = lean_array_push(x_50, x_63); -x_68 = lean_array_push(x_67, x_46); -x_69 = lean_array_push(x_68, x_66); +x_65 = l_Lean_Syntax_mkNumLit(x_64, x_24); +x_66 = lean_array_push(x_50, x_63); +x_67 = lean_array_push(x_66, x_46); +x_68 = lean_array_push(x_67, x_65); lean_inc(x_48); -x_70 = lean_array_push(x_69, x_48); -x_71 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_71, 0, x_24); -lean_ctor_set(x_71, 1, x_61); -lean_ctor_set(x_71, 2, x_70); -x_72 = lean_array_push(x_56, x_71); -x_73 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_73, 0, x_24); -lean_ctor_set(x_73, 1, x_58); -lean_ctor_set(x_73, 2, x_72); -x_74 = lean_array_get_size(x_4); -x_75 = lean_usize_of_nat(x_74); -lean_dec(x_74); -x_76 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_75, x_5, x_4); -x_77 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__4; -x_78 = l_Array_append___rarg(x_77, x_76); -x_79 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_79, 0, x_24); -lean_ctor_set(x_79, 1, x_58); -lean_ctor_set(x_79, 2, x_78); -x_80 = l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__33; +x_69 = lean_array_push(x_68, x_48); +x_70 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_70, 0, x_24); +lean_ctor_set(x_70, 1, x_61); +lean_ctor_set(x_70, 2, x_69); +x_71 = lean_array_push(x_56, x_70); +x_72 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_72, 0, x_24); +lean_ctor_set(x_72, 1, x_58); +lean_ctor_set(x_72, 2, x_71); +x_73 = lean_array_get_size(x_4); +x_74 = lean_usize_of_nat(x_73); +lean_dec(x_73); +x_75 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElab___spec__2(x_74, x_5, x_4); +x_76 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__4; +x_77 = l_Array_append___rarg(x_76, x_75); +x_78 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_78, 0, x_24); +lean_ctor_set(x_78, 1, x_58); +lean_ctor_set(x_78, 2, x_77); +x_79 = l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__33; lean_inc(x_30); -x_81 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_81, 0, x_30); -lean_ctor_set(x_81, 1, x_80); -x_82 = l_Lean_Elab_Command_elabElabRules___lambda__3___closed__1; -x_83 = lean_name_mk_string(x_2, x_82); -x_84 = l_Lean_Elab_Command_elabElab___lambda__1___closed__7; -x_85 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_85, 0, x_24); -lean_ctor_set(x_85, 1, x_6); -lean_ctor_set(x_85, 2, x_84); +x_80 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_80, 0, x_30); +lean_ctor_set(x_80, 1, x_79); +x_81 = l_Lean_Elab_Command_elabElabRules___lambda__3___closed__1; +x_82 = lean_name_mk_string(x_2, x_81); +x_83 = l_Lean_Elab_Command_elabElab___lambda__1___closed__7; +x_84 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_84, 0, x_24); +lean_ctor_set(x_84, 1, x_6); +lean_ctor_set(x_84, 2, x_83); lean_inc(x_30); -x_86 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_86, 0, x_30); -lean_ctor_set(x_86, 1, x_82); -x_87 = l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__26; -lean_inc(x_81); -x_88 = lean_array_push(x_87, x_81); +x_85 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_85, 0, x_30); +lean_ctor_set(x_85, 1, x_81); +x_86 = l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__26; +lean_inc(x_80); +x_87 = lean_array_push(x_86, x_80); lean_inc(x_7); -lean_inc(x_88); -x_89 = lean_array_push(x_88, x_7); -x_90 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_90, 0, x_24); -lean_ctor_set(x_90, 1, x_58); -lean_ctor_set(x_90, 2, x_89); -x_91 = l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__45; +lean_inc(x_87); +x_88 = lean_array_push(x_87, x_7); +x_89 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_89, 0, x_24); +lean_ctor_set(x_89, 1, x_58); +lean_ctor_set(x_89, 2, x_88); +x_90 = l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__45; lean_inc(x_8); -x_92 = lean_name_mk_string(x_8, x_91); -x_93 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__7; +x_91 = lean_name_mk_string(x_8, x_90); +x_92 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__7; lean_inc(x_8); -x_94 = lean_name_mk_string(x_8, x_93); -x_95 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__1; +x_93 = lean_name_mk_string(x_8, x_92); +x_94 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__1; lean_inc(x_30); -x_96 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_96, 0, x_30); -lean_ctor_set(x_96, 1, x_95); -x_97 = l_Lean_Elab_Command_elabElab___lambda__1___closed__8; -x_98 = lean_name_mk_string(x_8, x_97); -x_99 = l_Lean_Elab_Command_elabElab___lambda__1___closed__9; +x_95 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_95, 0, x_30); +lean_ctor_set(x_95, 1, x_94); +x_96 = l_Lean_Elab_Command_elabElab___lambda__1___closed__8; +x_97 = lean_name_mk_string(x_8, x_96); +x_98 = l_Lean_Elab_Command_elabElab___lambda__1___closed__9; lean_inc(x_30); -x_100 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_100, 0, x_30); -lean_ctor_set(x_100, 1, x_99); -x_101 = l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__28; -x_102 = lean_array_push(x_101, x_100); -x_103 = lean_array_push(x_102, x_25); -x_104 = lean_array_push(x_103, x_48); -x_105 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_105, 0, x_24); -lean_ctor_set(x_105, 1, x_98); -lean_ctor_set(x_105, 2, x_104); -x_106 = lean_array_push(x_56, x_105); -x_107 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_107, 0, x_24); -lean_ctor_set(x_107, 1, x_58); -lean_ctor_set(x_107, 2, x_106); -x_108 = lean_array_push(x_56, x_107); -x_109 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_109, 0, x_24); -lean_ctor_set(x_109, 1, x_58); -lean_ctor_set(x_109, 2, x_108); -x_110 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__7; +x_99 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_99, 0, x_30); +lean_ctor_set(x_99, 1, x_98); +x_100 = l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__28; +x_101 = lean_array_push(x_100, x_99); +x_102 = lean_array_push(x_101, x_25); +x_103 = lean_array_push(x_102, x_48); +x_104 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_104, 0, x_24); +lean_ctor_set(x_104, 1, x_97); +lean_ctor_set(x_104, 2, x_103); +x_105 = lean_array_push(x_56, x_104); +x_106 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_106, 0, x_24); +lean_ctor_set(x_106, 1, x_58); +lean_ctor_set(x_106, 2, x_105); +x_107 = lean_array_push(x_56, x_106); +x_108 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_108, 0, x_24); +lean_ctor_set(x_108, 1, x_58); +lean_ctor_set(x_108, 2, x_107); +x_109 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__6; lean_inc(x_30); -x_111 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_111, 0, x_30); -lean_ctor_set(x_111, 1, x_110); -x_112 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__8; -x_113 = lean_array_push(x_112, x_96); -x_114 = lean_array_push(x_113, x_109); -x_115 = lean_array_push(x_114, x_111); -x_116 = lean_array_push(x_115, x_9); -x_117 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_117, 0, x_24); -lean_ctor_set(x_117, 1, x_94); -lean_ctor_set(x_117, 2, x_116); -x_118 = lean_array_push(x_56, x_117); -x_119 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_119, 0, x_24); -lean_ctor_set(x_119, 1, x_58); -lean_ctor_set(x_119, 2, x_118); -x_120 = lean_array_push(x_56, x_119); -x_121 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_121, 0, x_24); -lean_ctor_set(x_121, 1, x_92); -lean_ctor_set(x_121, 2, x_120); +x_110 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_110, 0, x_30); +lean_ctor_set(x_110, 1, x_109); +x_111 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__7; +x_112 = lean_array_push(x_111, x_95); +x_113 = lean_array_push(x_112, x_108); +x_114 = lean_array_push(x_113, x_110); +x_115 = lean_array_push(x_114, x_9); +x_116 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_116, 0, x_24); +lean_ctor_set(x_116, 1, x_93); +lean_ctor_set(x_116, 2, x_115); +x_117 = lean_array_push(x_56, x_116); +x_118 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_118, 0, x_24); +lean_ctor_set(x_118, 1, x_58); +lean_ctor_set(x_118, 2, x_117); +x_119 = lean_array_push(x_56, x_118); +x_120 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_120, 0, x_24); +lean_ctor_set(x_120, 1, x_91); +lean_ctor_set(x_120, 2, x_119); if (lean_obj_tag(x_14) == 0) { -x_122 = x_77; -goto block_174; +x_121 = x_76; +goto block_173; } else { -lean_object* x_175; lean_object* x_176; -x_175 = lean_ctor_get(x_14, 0); -lean_inc(x_175); +lean_object* x_174; lean_object* x_175; +x_174 = lean_ctor_get(x_14, 0); +lean_inc(x_174); lean_dec(x_14); -x_176 = lean_array_push(x_56, x_175); -x_122 = x_176; -goto block_174; -} -block_174: -{ -lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; -x_123 = l_Array_append___rarg(x_77, x_122); -x_124 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_124, 0, x_24); -lean_ctor_set(x_124, 1, x_58); -lean_ctor_set(x_124, 2, x_123); -x_125 = l_Lean_Elab_Command_elabElab___lambda__1___closed__6; -lean_inc(x_124); -x_126 = lean_array_push(x_125, x_124); -x_127 = lean_array_push(x_126, x_10); -x_128 = lean_array_push(x_127, x_38); -x_129 = l_Lean_Elab_Command_elabElabRules___lambda__3___closed__2; -x_130 = lean_array_push(x_129, x_124); +x_175 = lean_array_push(x_56, x_174); +x_121 = x_175; +goto block_173; +} +block_173: +{ +lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; +x_122 = l_Array_append___rarg(x_76, x_121); +x_123 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_123, 0, x_24); +lean_ctor_set(x_123, 1, x_58); +lean_ctor_set(x_123, 2, x_122); +x_124 = l_Lean_Elab_Command_elabElab___lambda__1___closed__6; +lean_inc(x_123); +x_125 = lean_array_push(x_124, x_123); +x_126 = lean_array_push(x_125, x_10); +x_127 = lean_array_push(x_126, x_38); +x_128 = l_Lean_Elab_Command_elabElabRules___lambda__3___closed__2; +x_129 = lean_array_push(x_128, x_123); +x_130 = lean_array_push(x_129, x_84); x_131 = lean_array_push(x_130, x_85); -x_132 = lean_array_push(x_131, x_86); -x_133 = l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__92; -x_134 = lean_array_push(x_132, x_133); -x_135 = lean_array_push(x_134, x_90); +x_132 = l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__92; +x_133 = lean_array_push(x_131, x_132); +x_134 = lean_array_push(x_133, x_89); if (lean_obj_tag(x_12) == 0) { -lean_dec(x_88); +lean_dec(x_87); lean_dec(x_13); -x_136 = x_77; -goto block_167; +x_135 = x_76; +goto block_166; } else { -lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; -x_168 = lean_ctor_get(x_12, 0); -lean_inc(x_168); +lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; +x_167 = lean_ctor_get(x_12, 0); +lean_inc(x_167); lean_dec(x_12); -x_169 = l_Lean_Elab_Command_elabElab___lambda__1___closed__10; -x_170 = lean_name_mk_string(x_13, x_169); -x_171 = lean_array_push(x_88, x_168); -x_172 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_172, 0, x_24); -lean_ctor_set(x_172, 1, x_170); -lean_ctor_set(x_172, 2, x_171); -x_173 = lean_array_push(x_56, x_172); -x_136 = x_173; -goto block_167; -} -block_167: -{ -lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; -x_137 = l_Array_append___rarg(x_77, x_136); -x_138 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_138, 0, x_24); -lean_ctor_set(x_138, 1, x_58); -lean_ctor_set(x_138, 2, x_137); -x_139 = lean_array_push(x_128, x_138); -x_140 = lean_array_push(x_139, x_59); -x_141 = lean_array_push(x_140, x_73); -x_142 = lean_array_push(x_141, x_79); -x_143 = lean_array_push(x_142, x_81); -x_144 = lean_array_push(x_143, x_7); -x_145 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_145, 0, x_24); -lean_ctor_set(x_145, 1, x_37); -lean_ctor_set(x_145, 2, x_144); -x_146 = lean_array_push(x_87, x_145); +x_168 = l_Lean_Elab_Command_elabElab___lambda__1___closed__10; +x_169 = lean_name_mk_string(x_13, x_168); +x_170 = lean_array_push(x_87, x_167); +x_171 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_171, 0, x_24); +lean_ctor_set(x_171, 1, x_169); +lean_ctor_set(x_171, 2, x_170); +x_172 = lean_array_push(x_56, x_171); +x_135 = x_172; +goto block_166; +} +block_166: +{ +lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; +x_136 = l_Array_append___rarg(x_76, x_135); +x_137 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_137, 0, x_24); +lean_ctor_set(x_137, 1, x_58); +lean_ctor_set(x_137, 2, x_136); +x_138 = lean_array_push(x_127, x_137); +x_139 = lean_array_push(x_138, x_59); +x_140 = lean_array_push(x_139, x_72); +x_141 = lean_array_push(x_140, x_78); +x_142 = lean_array_push(x_141, x_80); +x_143 = lean_array_push(x_142, x_7); +x_144 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_144, 0, x_24); +lean_ctor_set(x_144, 1, x_37); +lean_ctor_set(x_144, 2, x_143); +x_145 = lean_array_push(x_86, x_144); if (lean_obj_tag(x_11) == 0) { -lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; +lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_dec(x_30); -x_147 = l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__59; -x_148 = lean_array_push(x_135, x_147); -x_149 = lean_array_push(x_148, x_121); -x_150 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_150, 0, x_24); -lean_ctor_set(x_150, 1, x_83); -lean_ctor_set(x_150, 2, x_149); -x_151 = lean_array_push(x_146, x_150); -x_152 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_152, 0, x_24); -lean_ctor_set(x_152, 1, x_58); -lean_ctor_set(x_152, 2, x_151); -x_153 = l_Lean_Elab_Command_elabCommand(x_152, x_16, x_17, x_35); -return x_153; +x_146 = l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__59; +x_147 = lean_array_push(x_134, x_146); +x_148 = lean_array_push(x_147, x_120); +x_149 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_149, 0, x_24); +lean_ctor_set(x_149, 1, x_82); +lean_ctor_set(x_149, 2, x_148); +x_150 = lean_array_push(x_145, x_149); +x_151 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_151, 0, x_24); +lean_ctor_set(x_151, 1, x_58); +lean_ctor_set(x_151, 2, x_150); +x_152 = l_Lean_Elab_Command_elabCommand(x_151, x_16, x_17, x_35); +return x_152; } else { -lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; -x_154 = lean_ctor_get(x_11, 0); -lean_inc(x_154); +lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; +x_153 = lean_ctor_get(x_11, 0); +lean_inc(x_153); lean_dec(x_11); -x_155 = l_Lean_Elab_Command_elabElabRules___lambda__3___closed__3; -x_156 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_156, 0, x_30); -lean_ctor_set(x_156, 1, x_155); -x_157 = lean_array_push(x_87, x_156); -x_158 = lean_array_push(x_157, x_154); -x_159 = l_Array_append___rarg(x_77, x_158); -x_160 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_160, 0, x_24); -lean_ctor_set(x_160, 1, x_58); -lean_ctor_set(x_160, 2, x_159); -x_161 = lean_array_push(x_135, x_160); -x_162 = lean_array_push(x_161, x_121); -x_163 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_163, 0, x_24); -lean_ctor_set(x_163, 1, x_83); -lean_ctor_set(x_163, 2, x_162); -x_164 = lean_array_push(x_146, x_163); -x_165 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_165, 0, x_24); -lean_ctor_set(x_165, 1, x_58); -lean_ctor_set(x_165, 2, x_164); -x_166 = l_Lean_Elab_Command_elabCommand(x_165, x_16, x_17, x_35); -return x_166; +x_154 = l_Lean_Elab_Command_elabElabRules___lambda__3___closed__3; +x_155 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_155, 0, x_30); +lean_ctor_set(x_155, 1, x_154); +x_156 = lean_array_push(x_86, x_155); +x_157 = lean_array_push(x_156, x_153); +x_158 = l_Array_append___rarg(x_76, x_157); +x_159 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_159, 0, x_24); +lean_ctor_set(x_159, 1, x_58); +lean_ctor_set(x_159, 2, x_158); +x_160 = lean_array_push(x_134, x_159); +x_161 = lean_array_push(x_160, x_120); +x_162 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_162, 0, x_24); +lean_ctor_set(x_162, 1, x_82); +lean_ctor_set(x_162, 2, x_161); +x_163 = lean_array_push(x_145, x_162); +x_164 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_164, 0, x_24); +lean_ctor_set(x_164, 1, x_58); +lean_ctor_set(x_164, 2, x_163); +x_165 = l_Lean_Elab_Command_elabCommand(x_164, x_16, x_17, x_35); +return x_165; } } } @@ -6371,7 +6615,7 @@ x_21 = lean_alloc_closure((void*)(l_Lean_evalOptPrio), 3, 1); lean_closure_set(x_21, 0, x_3); lean_inc(x_16); lean_inc(x_15); -x_22 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabSyntax___spec__2(x_21, x_15, x_16, x_17); +x_22 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabSyntax___spec__4(x_21, x_15, x_16, x_17); if (lean_obj_tag(x_22) == 0) { lean_object* x_23; lean_object* x_24; lean_object* x_25; size_t x_26; size_t x_27; lean_object* x_28; @@ -6407,7 +6651,7 @@ lean_inc(x_33); lean_dec(x_31); x_34 = l_Lean_Syntax_getId(x_6); x_35 = lean_box(2); -x_36 = l_Lean_nullKind; +x_36 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__3; lean_inc(x_32); x_37 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_37, 0, x_35); @@ -6418,7 +6662,7 @@ lean_closure_set(x_38, 0, x_34); lean_closure_set(x_38, 1, x_37); lean_inc(x_16); lean_inc(x_15); -x_39 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabSyntax___spec__7(x_38, x_15, x_16, x_30); +x_39 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabSyntax___spec__9(x_38, x_15, x_16, x_30); if (lean_obj_tag(x_39) == 0) { lean_object* x_40; lean_object* x_41; lean_object* x_42; @@ -6613,13 +6857,12 @@ x_26 = l_Lean_Syntax_getArg(x_18, x_25); x_27 = l_Lean_Syntax_isNone(x_26); if (x_27 == 0) { -lean_object* x_28; uint8_t x_29; -x_28 = l_Lean_nullKind; +uint8_t x_28; lean_inc(x_26); -x_29 = l_Lean_Syntax_isNodeOf(x_26, x_28, x_25); -if (x_29 == 0) +x_28 = l_Lean_Syntax_matchesNull(x_26, x_25); +if (x_28 == 0) { -lean_object* x_30; +lean_object* x_29; lean_dec(x_26); lean_dec(x_24); lean_dec(x_18); @@ -6635,29 +6878,29 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_30 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg(x_14); -return x_30; +x_29 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg(x_14); +return x_29; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_31 = l_Lean_Syntax_getArg(x_26, x_23); +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_30 = l_Lean_Syntax_getArg(x_26, x_23); lean_dec(x_26); -x_32 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_32, 0, x_31); -x_33 = lean_box(0); -x_34 = l_Lean_Elab_Command_elabElab___lambda__2(x_18, x_16, x_11, x_2, x_3, x_24, x_4, x_5, x_6, x_7, x_8, x_9, x_33, x_32, x_12, x_13, x_14); -return x_34; +x_31 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_31, 0, x_30); +x_32 = lean_box(0); +x_33 = l_Lean_Elab_Command_elabElab___lambda__2(x_18, x_16, x_11, x_2, x_3, x_24, x_4, x_5, x_6, x_7, x_8, x_9, x_32, x_31, x_12, x_13, x_14); +return x_33; } } else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_dec(x_26); +x_34 = lean_box(0); x_35 = lean_box(0); -x_36 = lean_box(0); -x_37 = l_Lean_Elab_Command_elabElab___lambda__2(x_18, x_16, x_11, x_2, x_3, x_24, x_4, x_5, x_6, x_7, x_8, x_9, x_36, x_35, x_12, x_13, x_14); -return x_37; +x_36 = l_Lean_Elab_Command_elabElab___lambda__2(x_18, x_16, x_11, x_2, x_3, x_24, x_4, x_5, x_6, x_7, x_8, x_9, x_35, x_34, x_12, x_13, x_14); +return x_36; } } } @@ -6672,14 +6915,13 @@ x_15 = l_Lean_Syntax_getArg(x_1, x_14); x_16 = l_Lean_Syntax_isNone(x_15); if (x_16 == 0) { -lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_17 = l_Lean_nullKind; -x_18 = lean_unsigned_to_nat(1u); +lean_object* x_17; uint8_t x_18; +x_17 = lean_unsigned_to_nat(1u); lean_inc(x_15); -x_19 = l_Lean_Syntax_isNodeOf(x_15, x_17, x_18); -if (x_19 == 0) +x_18 = l_Lean_Syntax_matchesNull(x_15, x_17); +if (x_18 == 0) { -lean_object* x_20; +lean_object* x_19; lean_dec(x_15); lean_dec(x_12); lean_dec(x_11); @@ -6692,25 +6934,25 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_20 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg(x_13); -return x_20; +x_19 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg(x_13); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = lean_unsigned_to_nat(0u); -x_22 = l_Lean_Syntax_getArg(x_15, x_21); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = lean_unsigned_to_nat(0u); +x_21 = l_Lean_Syntax_getArg(x_15, x_20); lean_dec(x_15); -x_23 = l_Lean_Elab_Command_elabElab___lambda__1___closed__4; +x_22 = l_Lean_Elab_Command_elabElab___lambda__1___closed__4; lean_inc(x_2); -x_24 = lean_name_mk_string(x_2, x_23); -lean_inc(x_22); -x_25 = l_Lean_Syntax_isOfKind(x_22, x_24); -lean_dec(x_24); -if (x_25 == 0) +x_23 = lean_name_mk_string(x_2, x_22); +lean_inc(x_21); +x_24 = l_Lean_Syntax_isOfKind(x_21, x_23); +lean_dec(x_23); +if (x_24 == 0) { -lean_object* x_26; -lean_dec(x_22); +lean_object* x_25; +lean_dec(x_21); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -6722,31 +6964,31 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_26 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg(x_13); -return x_26; +x_25 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg(x_13); +return x_25; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_27 = lean_unsigned_to_nat(3u); -x_28 = l_Lean_Syntax_getArg(x_22, x_27); -lean_dec(x_22); -x_29 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_29, 0, x_28); -x_30 = lean_box(0); -x_31 = l_Lean_Elab_Command_elabElab___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_10, x_30, x_29, x_11, x_12, x_13); -return x_31; +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_26 = lean_unsigned_to_nat(3u); +x_27 = l_Lean_Syntax_getArg(x_21, x_26); +lean_dec(x_21); +x_28 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_28, 0, x_27); +x_29 = lean_box(0); +x_30 = l_Lean_Elab_Command_elabElab___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_10, x_29, x_28, x_11, x_12, x_13); +return x_30; } } } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_dec(x_15); +x_31 = lean_box(0); x_32 = lean_box(0); -x_33 = lean_box(0); -x_34 = l_Lean_Elab_Command_elabElab___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_10, x_33, x_32, x_11, x_12, x_13); -return x_34; +x_33 = l_Lean_Elab_Command_elabElab___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_10, x_32, x_31, x_11, x_12, x_13); +return x_33; } } } @@ -6760,14 +7002,13 @@ x_14 = l_Lean_Syntax_getArg(x_1, x_13); x_15 = l_Lean_Syntax_isNone(x_14); if (x_15 == 0) { -lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(1u); +lean_object* x_16; uint8_t x_17; +x_16 = lean_unsigned_to_nat(1u); lean_inc(x_14); -x_18 = l_Lean_Syntax_isNodeOf(x_14, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_14, x_16); +if (x_17 == 0) { -lean_object* x_19; +lean_object* x_18; lean_dec(x_14); lean_dec(x_11); lean_dec(x_10); @@ -6779,25 +7020,25 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_19 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg(x_12); -return x_19; +x_18 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg(x_12); +return x_18; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_20 = lean_unsigned_to_nat(0u); -x_21 = l_Lean_Syntax_getArg(x_14, x_20); +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_19 = lean_unsigned_to_nat(0u); +x_20 = l_Lean_Syntax_getArg(x_14, x_19); lean_dec(x_14); -x_22 = l_Lean_Elab_Command_elabElab___lambda__1___closed__2; +x_21 = l_Lean_Elab_Command_elabElab___lambda__1___closed__2; lean_inc(x_2); -x_23 = lean_name_mk_string(x_2, x_22); -lean_inc(x_21); -x_24 = l_Lean_Syntax_isOfKind(x_21, x_23); -lean_dec(x_23); -if (x_24 == 0) +x_22 = lean_name_mk_string(x_2, x_21); +lean_inc(x_20); +x_23 = l_Lean_Syntax_isOfKind(x_20, x_22); +lean_dec(x_22); +if (x_23 == 0) { -lean_object* x_25; -lean_dec(x_21); +lean_object* x_24; +lean_dec(x_20); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -6808,31 +7049,31 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_25 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg(x_12); -return x_25; +x_24 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg(x_12); +return x_24; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_26 = lean_unsigned_to_nat(3u); -x_27 = l_Lean_Syntax_getArg(x_21, x_26); -lean_dec(x_21); -x_28 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_28, 0, x_27); -x_29 = lean_box(0); -x_30 = l_Lean_Elab_Command_elabElab___lambda__4(x_1, x_2, x_3, x_4, x_5, x_9, x_6, x_7, x_29, x_28, x_10, x_11, x_12); -return x_30; +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_25 = lean_unsigned_to_nat(3u); +x_26 = l_Lean_Syntax_getArg(x_20, x_25); +lean_dec(x_20); +x_27 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_27, 0, x_26); +x_28 = lean_box(0); +x_29 = l_Lean_Elab_Command_elabElab___lambda__4(x_1, x_2, x_3, x_4, x_5, x_9, x_6, x_7, x_28, x_27, x_10, x_11, x_12); +return x_29; } } } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_dec(x_14); +x_30 = lean_box(0); x_31 = lean_box(0); -x_32 = lean_box(0); -x_33 = l_Lean_Elab_Command_elabElab___lambda__4(x_1, x_2, x_3, x_4, x_5, x_9, x_6, x_7, x_32, x_31, x_10, x_11, x_12); -return x_33; +x_32 = l_Lean_Elab_Command_elabElab___lambda__4(x_1, x_2, x_3, x_4, x_5, x_9, x_6, x_7, x_31, x_30, x_10, x_11, x_12); +return x_32; } } } @@ -6843,7 +7084,7 @@ lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_ lean_dec(x_4); x_9 = lean_unsigned_to_nat(1u); x_10 = l_Lean_Syntax_getArg(x_1, x_9); -x_11 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__5; +x_11 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__5; lean_inc(x_2); x_12 = lean_name_mk_string(x_2, x_11); x_13 = l_Lean_Elab_Command_elabElabRules___lambda__6___closed__1; @@ -6874,13 +7115,12 @@ x_18 = l_Lean_Syntax_getArg(x_1, x_17); x_19 = l_Lean_Syntax_isNone(x_18); if (x_19 == 0) { -lean_object* x_20; uint8_t x_21; -x_20 = l_Lean_nullKind; +uint8_t x_20; lean_inc(x_18); -x_21 = l_Lean_Syntax_isNodeOf(x_18, x_20, x_9); -if (x_21 == 0) +x_20 = l_Lean_Syntax_matchesNull(x_18, x_9); +if (x_20 == 0) { -lean_object* x_22; +lean_object* x_21; lean_dec(x_18); lean_dec(x_14); lean_dec(x_12); @@ -6891,25 +7131,25 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_22 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg(x_8); -return x_22; +x_21 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg(x_8); +return x_21; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; -x_23 = lean_unsigned_to_nat(0u); -x_24 = l_Lean_Syntax_getArg(x_18, x_23); +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_22 = lean_unsigned_to_nat(0u); +x_23 = l_Lean_Syntax_getArg(x_18, x_22); lean_dec(x_18); -x_25 = l_Lean_Elab_Command_elabElab___lambda__1___closed__10; +x_24 = l_Lean_Elab_Command_elabElab___lambda__1___closed__10; lean_inc(x_2); -x_26 = lean_name_mk_string(x_2, x_25); -lean_inc(x_24); -x_27 = l_Lean_Syntax_isOfKind(x_24, x_26); -lean_dec(x_26); -if (x_27 == 0) +x_25 = lean_name_mk_string(x_2, x_24); +lean_inc(x_23); +x_26 = l_Lean_Syntax_isOfKind(x_23, x_25); +lean_dec(x_25); +if (x_26 == 0) { -lean_object* x_28; -lean_dec(x_24); +lean_object* x_27; +lean_dec(x_23); lean_dec(x_14); lean_dec(x_12); lean_dec(x_10); @@ -6919,30 +7159,30 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_28 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg(x_8); -return x_28; +x_27 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg(x_8); +return x_27; } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_29 = l_Lean_Syntax_getArg(x_24, x_9); -lean_dec(x_24); -x_30 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_30, 0, x_29); -x_31 = lean_box(0); -x_32 = l_Lean_Elab_Command_elabElab___lambda__5(x_1, x_3, x_14, x_12, x_10, x_2, x_5, x_31, x_30, x_6, x_7, x_8); -return x_32; +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_28 = l_Lean_Syntax_getArg(x_23, x_9); +lean_dec(x_23); +x_29 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_29, 0, x_28); +x_30 = lean_box(0); +x_31 = l_Lean_Elab_Command_elabElab___lambda__5(x_1, x_3, x_14, x_12, x_10, x_2, x_5, x_30, x_29, x_6, x_7, x_8); +return x_31; } } } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_dec(x_18); +x_32 = lean_box(0); x_33 = lean_box(0); -x_34 = lean_box(0); -x_35 = l_Lean_Elab_Command_elabElab___lambda__5(x_1, x_3, x_14, x_12, x_10, x_2, x_5, x_34, x_33, x_6, x_7, x_8); -return x_35; +x_34 = l_Lean_Elab_Command_elabElab___lambda__5(x_1, x_3, x_14, x_12, x_10, x_2, x_5, x_33, x_32, x_6, x_7, x_8); +return x_34; } } } @@ -6989,62 +7229,61 @@ x_9 = l_Lean_Syntax_getArg(x_1, x_8); x_10 = l_Lean_Syntax_isNone(x_9); if (x_10 == 0) { -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = l_Lean_nullKind; -x_12 = lean_unsigned_to_nat(1u); +lean_object* x_11; uint8_t x_12; +x_11 = lean_unsigned_to_nat(1u); lean_inc(x_9); -x_13 = l_Lean_Syntax_isNodeOf(x_9, x_11, x_12); -if (x_13 == 0) +x_12 = l_Lean_Syntax_matchesNull(x_9, x_11); +if (x_12 == 0) { -lean_object* x_14; +lean_object* x_13; lean_dec(x_9); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_14 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg(x_4); -return x_14; +x_13 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg(x_4); +return x_13; } else { -lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_15 = l_Lean_Syntax_getArg(x_9, x_8); +lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_14 = l_Lean_Syntax_getArg(x_9, x_8); lean_dec(x_9); -x_16 = l_Lean_Elab_Command_elabElabRules___lambda__7___closed__4; -lean_inc(x_15); -x_17 = l_Lean_Syntax_isOfKind(x_15, x_16); -if (x_17 == 0) +x_15 = l_Lean_Elab_Command_elabElabRules___lambda__7___closed__4; +lean_inc(x_14); +x_16 = l_Lean_Syntax_isOfKind(x_14, x_15); +if (x_16 == 0) { -lean_object* x_18; -lean_dec(x_15); +lean_object* x_17; +lean_dec(x_14); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_18 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg(x_4); -return x_18; +x_17 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg(x_4); +return x_17; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_19 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_19, 0, x_15); -x_20 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__4; -x_21 = l_Lean_Elab_Command_elabElabRules___lambda__7___closed__1; -x_22 = lean_box(0); -x_23 = l_Lean_Elab_Command_elabElab___lambda__6(x_1, x_20, x_21, x_22, x_19, x_2, x_3, x_4); -return x_23; +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_18 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_18, 0, x_14); +x_19 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__4; +x_20 = l_Lean_Elab_Command_elabElabRules___lambda__7___closed__1; +x_21 = lean_box(0); +x_22 = l_Lean_Elab_Command_elabElab___lambda__6(x_1, x_19, x_20, x_21, x_18, x_2, x_3, x_4); +return x_22; } } } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_dec(x_9); -x_24 = lean_box(0); -x_25 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__4; -x_26 = l_Lean_Elab_Command_elabElabRules___lambda__7___closed__1; -x_27 = lean_box(0); -x_28 = l_Lean_Elab_Command_elabElab___lambda__6(x_1, x_25, x_26, x_27, x_24, x_2, x_3, x_4); -return x_28; +x_23 = lean_box(0); +x_24 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__4; +x_25 = l_Lean_Elab_Command_elabElabRules___lambda__7___closed__1; +x_26 = lean_box(0); +x_27 = l_Lean_Elab_Command_elabElab___lambda__6(x_1, x_24, x_25, x_26, x_23, x_2, x_3, x_4); +return x_27; } } } @@ -7061,6 +7300,18 @@ x_9 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElab___spec__1(x_7, x_8, return x_9; } } +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElab___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElab___spec__2(x_4, x_5, x_3); +return x_6; +} +} LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabElab___lambda__1___boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; @@ -7156,7 +7407,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabElab_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(80u); +x_1 = lean_unsigned_to_nat(81u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7168,7 +7419,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabElab_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(93u); +x_1 = lean_unsigned_to_nat(94u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7196,7 +7447,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabElab_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(80u); +x_1 = lean_unsigned_to_nat(81u); x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7208,7 +7459,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabElab_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(80u); +x_1 = lean_unsigned_to_nat(81u); x_2 = lean_unsigned_to_nat(12u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7275,52 +7526,58 @@ l_Lean_Elab_Command_withExpectedType___closed__1 = _init_l_Lean_Elab_Command_wit lean_mark_persistent(l_Lean_Elab_Command_withExpectedType___closed__1); l_Lean_Elab_Command_withExpectedType___closed__2 = _init_l_Lean_Elab_Command_withExpectedType___closed__2(); lean_mark_persistent(l_Lean_Elab_Command_withExpectedType___closed__2); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__1 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__1(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__1); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__2 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__2(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__2); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__3 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__3(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__3); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__4 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__4(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__4); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__5 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__5(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__5); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__6 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__6(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__6); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__7 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__7(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__7); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__8 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__8(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__8); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__9 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__9(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__9); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__10 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__10(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__10); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__11 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__11(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__11); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__12 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__12(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__12); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__13 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__13(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__13); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__14 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__14(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__14); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__15 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__15(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___lambda__2___closed__15); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__1 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__1(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__1); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__2 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__2(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__2); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__3 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__3(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__3); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__4 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__4(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__4); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__5 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__5(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__5); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__6 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__6(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__6); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__7 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__7(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__7); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__8 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__8(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__2___closed__8); +l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabElabRulesAux___spec__1___rarg___closed__1 = _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabElabRulesAux___spec__1___rarg___closed__1(); +lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabElabRulesAux___spec__1___rarg___closed__1); +l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabElabRulesAux___spec__1___rarg___closed__2 = _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabElabRulesAux___spec__1___rarg___closed__2(); +lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabElabRulesAux___spec__1___rarg___closed__2); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__1 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__1(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__1); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__2 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__2(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__2); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__3 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__3(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__3); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__4 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__4(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__4); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__5 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__5(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__5); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__6 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__6(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__6); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__7 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__7(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__7); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__8 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__8(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__8); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__9 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__9(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__9); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__10 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__10(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__10); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__11 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__11(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__11); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__12 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__12(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__12); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__13 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__13(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__13); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__14 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__14(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__14); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__15 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__15(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__15); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__16 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__16(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__16); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__1 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__1(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__1); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__2 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__2(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__2); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__3 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__3(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__3); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__4 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__4(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__4); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__5 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__5(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__5); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__6 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__6(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__6); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__7 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__7(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__7); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__8 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__8(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___closed__8); l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__1 = _init_l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__1); l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__2 = _init_l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Extra.c b/stage0/stdlib/Lean/Elab/Extra.c index 770f565b4b3..0f2c0d6a87d 100644 --- a/stage0/stdlib/Lean/Elab/Extra.c +++ b/stage0/stdlib/Lean/Elab/Extra.c @@ -44,7 +44,6 @@ static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_BinOp_elabBinCa lean_object* l_Lean_mkSort(lean_object*); static lean_object* l_Lean_Elab_Term_elabForIn___closed__10; static lean_object* l_Lean_Elab_Term_elabForIn___lambda__1___closed__4; -extern lean_object* l_Lean_nullKind; static lean_object* l_Lean_Elab_Term_BinOp_elabBinOp___closed__12; LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_BinOp_toTree_processOp___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_name_mk_string(lean_object*, lean_object*); @@ -174,6 +173,7 @@ lean_object* l_Lean_Meta_mkSorry(lean_object*, uint8_t, lean_object*, lean_objec static lean_object* l_Lean_Elab_Term_elabForIn___closed__12; static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_getMonadForIn___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_elabForIn_declRange___closed__3; +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_BinOp_elabBinCalc___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_BinOp_applyCoe_go___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_BinOp_elabBinRelCore_toBoolIfNecessary___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -268,7 +268,7 @@ static lean_object* l_Lean_Elab_Term_BinOp_elabBinRelCore_toBoolIfNecessary___cl static lean_object* l___regBuiltin_Lean_Elab_Term_BinOp_elabBinRelNoProp_declRange___closed__1; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Term_BinOp_elabBinCalc___spec__3___closed__7; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_BinOp_toExpr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_BinOp_initFn____x40_Lean_Elab_Extra___hyg_6137_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_BinOp_initFn____x40_Lean_Elab_Extra___hyg_6221_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabForIn_x27___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_throwForInFailure___closed__2; LEAN_EXPORT uint8_t l___private_Lean_Elab_Extra_0__Lean_Elab_Term_BinOp_isUnknow(lean_object*); @@ -301,7 +301,6 @@ static lean_object* l_Lean_Elab_Term_elabForIn_x27___lambda__1___closed__3; static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_getMonadForIn___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Term_BinOp_elabBinRelNoProp_declRange___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabForIn_x27(lean_object*); -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabForIn___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_elabForIn_x27_declRange___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Term_BinOp_elabBinCalc_declRange___closed__6; @@ -3561,153 +3560,152 @@ return x_29; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; +lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; x_30 = lean_unsigned_to_nat(1u); x_31 = l_Lean_Syntax_getArg(x_2, x_30); -x_32 = l_Lean_nullKind; -x_33 = lean_unsigned_to_nat(2u); +x_32 = lean_unsigned_to_nat(2u); lean_inc(x_31); -x_34 = l_Lean_Syntax_isNodeOf(x_31, x_32, x_33); -if (x_34 == 0) +x_33 = l_Lean_Syntax_matchesNull(x_31, x_32); +if (x_33 == 0) { -lean_object* x_35; uint8_t x_36; lean_object* x_37; +lean_object* x_34; uint8_t x_35; lean_object* x_36; lean_dec(x_31); lean_dec(x_1); -x_35 = lean_box(0); -x_36 = 1; +x_34 = lean_box(0); +x_35 = 1; lean_inc(x_2); -x_37 = l_Lean_Elab_Term_elabTerm(x_2, x_35, x_36, x_36, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_37) == 0) +x_36 = l_Lean_Elab_Term_elabTerm(x_2, x_34, x_35, x_35, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_36) == 0) { -uint8_t x_38; -x_38 = !lean_is_exclusive(x_37); -if (x_38 == 0) +uint8_t x_37; +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) { -lean_object* x_39; lean_object* x_40; -x_39 = lean_ctor_get(x_37, 0); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_2); -lean_ctor_set(x_40, 1, x_39); -lean_ctor_set(x_37, 0, x_40); -return x_37; +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_36, 0); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_2); +lean_ctor_set(x_39, 1, x_38); +lean_ctor_set(x_36, 0, x_39); +return x_36; } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_41 = lean_ctor_get(x_37, 0); -x_42 = lean_ctor_get(x_37, 1); -lean_inc(x_42); +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_40 = lean_ctor_get(x_36, 0); +x_41 = lean_ctor_get(x_36, 1); lean_inc(x_41); -lean_dec(x_37); +lean_inc(x_40); +lean_dec(x_36); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_2); +lean_ctor_set(x_42, 1, x_40); x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_2); +lean_ctor_set(x_43, 0, x_42); lean_ctor_set(x_43, 1, x_41); -x_44 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_44, 1, x_42); -return x_44; +return x_43; } } else { -uint8_t x_45; +uint8_t x_44; lean_dec(x_2); -x_45 = !lean_is_exclusive(x_37); -if (x_45 == 0) +x_44 = !lean_is_exclusive(x_36); +if (x_44 == 0) { -return x_37; +return x_36; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_37, 0); -x_47 = lean_ctor_get(x_37, 1); -lean_inc(x_47); +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_36, 0); +x_46 = lean_ctor_get(x_36, 1); lean_inc(x_46); -lean_dec(x_37); -x_48 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_48, 0, x_46); -lean_ctor_set(x_48, 1, x_47); -return x_48; +lean_inc(x_45); +lean_dec(x_36); +x_47 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +return x_47; } } } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; -x_49 = lean_unsigned_to_nat(0u); -x_50 = l_Lean_Syntax_getArg(x_31, x_49); -x_51 = l_Lean_Syntax_getArg(x_31, x_30); +lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; +x_48 = lean_unsigned_to_nat(0u); +x_49 = l_Lean_Syntax_getArg(x_31, x_48); +x_50 = l_Lean_Syntax_getArg(x_31, x_30); lean_dec(x_31); -x_52 = l_Lean_Syntax_isNodeOf(x_51, x_32, x_49); -if (x_52 == 0) +x_51 = l_Lean_Syntax_matchesNull(x_50, x_48); +if (x_51 == 0) { -lean_object* x_53; uint8_t x_54; lean_object* x_55; -lean_dec(x_50); +lean_object* x_52; uint8_t x_53; lean_object* x_54; +lean_dec(x_49); lean_dec(x_1); -x_53 = lean_box(0); -x_54 = 1; +x_52 = lean_box(0); +x_53 = 1; lean_inc(x_2); -x_55 = l_Lean_Elab_Term_elabTerm(x_2, x_53, x_54, x_54, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_55) == 0) +x_54 = l_Lean_Elab_Term_elabTerm(x_2, x_52, x_53, x_53, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_54) == 0) { -uint8_t x_56; -x_56 = !lean_is_exclusive(x_55); -if (x_56 == 0) +uint8_t x_55; +x_55 = !lean_is_exclusive(x_54); +if (x_55 == 0) { -lean_object* x_57; lean_object* x_58; -x_57 = lean_ctor_get(x_55, 0); -x_58 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_58, 0, x_2); -lean_ctor_set(x_58, 1, x_57); -lean_ctor_set(x_55, 0, x_58); -return x_55; +lean_object* x_56; lean_object* x_57; +x_56 = lean_ctor_get(x_54, 0); +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_2); +lean_ctor_set(x_57, 1, x_56); +lean_ctor_set(x_54, 0, x_57); +return x_54; } else { -lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_59 = lean_ctor_get(x_55, 0); -x_60 = lean_ctor_get(x_55, 1); -lean_inc(x_60); +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_58 = lean_ctor_get(x_54, 0); +x_59 = lean_ctor_get(x_54, 1); lean_inc(x_59); -lean_dec(x_55); +lean_inc(x_58); +lean_dec(x_54); +x_60 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_60, 0, x_2); +lean_ctor_set(x_60, 1, x_58); x_61 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_61, 0, x_2); +lean_ctor_set(x_61, 0, x_60); lean_ctor_set(x_61, 1, x_59); -x_62 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_62, 0, x_61); -lean_ctor_set(x_62, 1, x_60); -return x_62; +return x_61; } } else { -uint8_t x_63; +uint8_t x_62; lean_dec(x_2); -x_63 = !lean_is_exclusive(x_55); -if (x_63 == 0) +x_62 = !lean_is_exclusive(x_54); +if (x_62 == 0) { -return x_55; +return x_54; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_ctor_get(x_55, 0); -x_65 = lean_ctor_get(x_55, 1); -lean_inc(x_65); +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_54, 0); +x_64 = lean_ctor_get(x_54, 1); lean_inc(x_64); -lean_dec(x_55); -x_66 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_66, 0, x_64); -lean_ctor_set(x_66, 1, x_65); -return x_66; +lean_inc(x_63); +lean_dec(x_54); +x_65 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_64); +return x_65; } } } else { lean_dec(x_2); -x_2 = x_50; +x_2 = x_49; goto _start; } } @@ -3715,32 +3713,32 @@ goto _start; } else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; uint8_t x_74; lean_object* x_75; -x_68 = lean_unsigned_to_nat(1u); -x_69 = l_Lean_Syntax_getArg(x_2, x_68); -x_70 = lean_unsigned_to_nat(2u); -x_71 = l_Lean_Syntax_getArg(x_2, x_70); -x_72 = lean_unsigned_to_nat(3u); -x_73 = l_Lean_Syntax_getArg(x_2, x_72); +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; uint8_t x_73; lean_object* x_74; +x_67 = lean_unsigned_to_nat(1u); +x_68 = l_Lean_Syntax_getArg(x_2, x_67); +x_69 = lean_unsigned_to_nat(2u); +x_70 = l_Lean_Syntax_getArg(x_2, x_69); +x_71 = lean_unsigned_to_nat(3u); +x_72 = l_Lean_Syntax_getArg(x_2, x_71); lean_dec(x_2); -x_74 = 1; -x_75 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_BinOp_toTree_processOp(x_1, x_69, x_71, x_73, x_74, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_75; +x_73 = 1; +x_74 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_BinOp_toTree_processOp(x_1, x_68, x_70, x_72, x_73, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_74; } } else { -lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; uint8_t x_82; lean_object* x_83; -x_76 = lean_unsigned_to_nat(1u); -x_77 = l_Lean_Syntax_getArg(x_2, x_76); -x_78 = lean_unsigned_to_nat(2u); -x_79 = l_Lean_Syntax_getArg(x_2, x_78); -x_80 = lean_unsigned_to_nat(3u); -x_81 = l_Lean_Syntax_getArg(x_2, x_80); +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; uint8_t x_81; lean_object* x_82; +x_75 = lean_unsigned_to_nat(1u); +x_76 = l_Lean_Syntax_getArg(x_2, x_75); +x_77 = lean_unsigned_to_nat(2u); +x_78 = l_Lean_Syntax_getArg(x_2, x_77); +x_79 = lean_unsigned_to_nat(3u); +x_80 = l_Lean_Syntax_getArg(x_2, x_79); lean_dec(x_2); -x_82 = 0; -x_83 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_BinOp_toTree_processOp(x_1, x_77, x_79, x_81, x_82, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_83; +x_81 = 0; +x_82 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_BinOp_toTree_processOp(x_1, x_76, x_78, x_80, x_81, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_82; } } } @@ -16560,7 +16558,7 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_BinOp_initFn____x40_Lean_Elab_Extra___hyg_6137_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_BinOp_initFn____x40_Lean_Elab_Extra___hyg_6221_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; @@ -17163,7 +17161,7 @@ lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_BinOp_elabDefaultOrNonempty_d res = l___regBuiltin_Lean_Elab_Term_BinOp_elabDefaultOrNonempty_declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -res = l_Lean_Elab_Term_BinOp_initFn____x40_Lean_Elab_Extra___hyg_6137_(lean_io_mk_world()); +res = l_Lean_Elab_Term_BinOp_initFn____x40_Lean_Elab_Extra___hyg_6221_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/GenInjective.c b/stage0/stdlib/Lean/Elab/GenInjective.c index 7e4a69acf92..8b59fdb5f99 100644 --- a/stage0/stdlib/Lean/Elab/GenInjective.c +++ b/stage0/stdlib/Lean/Elab/GenInjective.c @@ -19,6 +19,7 @@ static lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Command_elabGenIn static lean_object* l___regBuiltin_Lean_Elab_Command_elabGenInjectiveTheorems___closed__11; lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabGenInjectiveTheorems___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_liftTermElabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_name_mk_string(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_resolveGlobalName___at_Lean_Elab_Command_elabGenInjectiveTheorems___spec__5(lean_object*, lean_object*, lean_object*, lean_object*); @@ -64,7 +65,6 @@ LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at_Lean_Elab_Command_el static lean_object* l___regBuiltin_Lean_Elab_Command_elabGenInjectiveTheorems___closed__14; lean_object* l_Lean_ResolveName_resolveGlobalName(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabGenInjectiveTheorems___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabGenInjectiveTheorems_declRange___closed__4; lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabGenInjectiveTheorems___closed__6; @@ -279,7 +279,7 @@ x_7 = lean_ctor_get(x_5, 1); lean_inc(x_7); lean_dec(x_5); x_8 = lean_box(0); -x_9 = l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(x_6, x_8); +x_9 = l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(x_6, x_8); x_10 = l_List_isEmpty___rarg(x_9); if (x_10 == 0) { diff --git a/stage0/stdlib/Lean/Elab/Inductive.c b/stage0/stdlib/Lean/Elab/Inductive.c index 23d69f32ff5..a6d27130c7c 100644 --- a/stage0/stdlib/Lean/Elab/Inductive.c +++ b/stage0/stdlib/Lean/Elab/Inductive.c @@ -19,6 +19,7 @@ static lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_checkPa lean_object* l_Lean_Elab_ContextInfo_save___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabCtors(lean_object*, 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_List_mapM___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateResultingUniverse___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_withCtorRef___spec__2___rarg___lambda__2(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_shouldInferResultUniverse___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); @@ -117,6 +118,7 @@ static lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_getArit LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_checkUnsafe(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_replaceArrowBinderNames_go___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_checkValidCtorModifier___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_withInductiveLocalDecls___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentArray_mapM___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_mkInductiveDecl___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_fixedIndicesToParams_go___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -701,7 +703,6 @@ static lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateR static lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabCtors_elabCtorType___closed__6; extern lean_object* l_Lean_Expr_instBEqExpr; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_mkIndFVar2Const___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___spec__1___boxed(lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_withoutAutoBoundImplicit___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_mkAuxConstructions___closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_withNewBinderInfos___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_withExplicitToImplicit___spec__2(lean_object*); @@ -780,7 +781,6 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Induc static lean_object* l_Lean_mkRecOn___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_mkAuxConstructions___spec__1___closed__10; lean_object* l_Lean_Elab_Term_withAutoBoundImplicit___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_fixedIndicesToParams___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabCtors_checkParamOccs___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_fixedIndicesToParams_go___lambda__4___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_mkInductiveDecl___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -9055,7 +9055,7 @@ lean_dec(x_3); return x_9; } } -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -9098,7 +9098,7 @@ LEAN_EXPORT uint8_t l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMV lean_object* x_3; uint8_t x_4; x_3 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_3, 0, x_2); -x_4 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___spec__1(x_1, x_3); +x_4 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___spec__1(x_1, x_3); lean_dec(x_3); return x_4; } @@ -9113,11 +9113,11 @@ x_12 = l_Lean_Elab_Term_levelMVarToParam_x27(x_2, x_11, x_3, x_4, x_5, x_6, x_7, return x_12; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___spec__1(x_1, x_2); +x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___spec__1(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); diff --git a/stage0/stdlib/Lean/Elab/Level.c b/stage0/stdlib/Lean/Elab/Level.c index 7c16d837711..97edde53103 100644 --- a/stage0/stdlib/Lean/Elab/Level.c +++ b/stage0/stdlib/Lean/Elab/Level.c @@ -27,12 +27,12 @@ static lean_object* l_Lean_Elab_Level_elabLevel___closed__7; lean_object* lean_name_mk_string(lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); lean_object* lean_array_uget(lean_object*, size_t); +static lean_object* l_Lean_Elab_Level_elabLevel___closed__24; lean_object* l_Lean_MetavarContext_addLevelMVarDecl(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Elab_Level_instMonadOptionsLevelElabM___spec__2(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Level_elabLevel___closed__16; static lean_object* l_Lean_Elab_Level_instMonadRefLevelElabM___closed__4; size_t lean_usize_sub(size_t, size_t); -lean_object* l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshMVarId___at_Lean_Elab_Level_mkFreshLevelMVar___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Level_0__Lean_Elab_Level_checkUniverseOffset___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Level_initFn____x40_Lean_Elab_Level___hyg_254____closed__3; @@ -40,6 +40,7 @@ uint8_t lean_name_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Level_instMonadRefLevelElabM___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Level_elabLevel___closed__9; lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Level_elabLevel___closed__27; static lean_object* l_Lean_Elab_Level_initFn____x40_Lean_Elab_Level___hyg_254____closed__5; lean_object* lean_array_get_size(lean_object*); static lean_object* l_Lean_Elab_Level_elabLevel___closed__11; @@ -56,12 +57,13 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Level_instMonadOptionsLevelElabM___lambda__ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Level_elabLevel___spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_throwIllFormedSyntax___at_Lean_Elab_Level_elabLevel___spec__2___closed__1; static lean_object* l_Lean_Elab_Level_elabLevel___closed__6; +lean_object* l_Lean_Syntax_isNatLit_x3f(lean_object*); uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Level_instMonadOptionsLevelElabM___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Level_instMonadOptionsLevelElabM___lambda__1___boxed(lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_numLitKind; lean_object* l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Level_elabLevel___closed__25; LEAN_EXPORT lean_object* l_Lean_Elab_Level_mkFreshLevelMVar(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Level_instMonadOptionsLevelElabM___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Level_maxUniverseOffset; @@ -96,7 +98,6 @@ static lean_object* l_Lean_Elab_Level_elabLevel___closed__4; static lean_object* l_Lean_Elab_Level_elabLevel___closed__17; LEAN_EXPORT lean_object* l_Lean_mkFreshId___at_Lean_Elab_Level_mkFreshLevelMVar___spec__2___rarg(lean_object*); static lean_object* l_Lean_Elab_Level_elabLevel___closed__1; -extern lean_object* l_Lean_identKind; uint8_t l_List_elem___at_Lean_NameHashSet_insert___spec__2(lean_object*, lean_object*); uint8_t l_Lean_Elab_isValidAutoBoundLevelName(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_throwIllFormedSyntax___at_Lean_Elab_Level_elabLevel___spec__2(lean_object*, lean_object*); @@ -128,6 +129,7 @@ LEAN_EXPORT lean_object* l_ReaderT_read___at_Lean_Elab_Level_instMonadOptionsLev static lean_object* l_Lean_Elab_Level_instMonadNameGeneratorLevelElabM___closed__1; LEAN_EXPORT lean_object* l_Lean_mkFreshMVarId___at_Lean_Elab_Level_mkFreshLevelMVar___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Level_instMonadRefLevelElabM___closed__1; +static lean_object* l_Lean_Elab_Level_elabLevel___closed__26; LEAN_EXPORT lean_object* l___private_Lean_Elab_Level_0__Lean_Elab_Level_checkUniverseOffset___at_Lean_Elab_Level_elabLevel___spec__3___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Level_ofNat(lean_object*); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); @@ -1318,7 +1320,7 @@ static lean_object* _init_l_Lean_Elab_Level_elabLevel___closed__15() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("addLit", 6); +x_1 = lean_mk_string_from_bytes("num", 3); return x_1; } } @@ -1326,7 +1328,7 @@ static lean_object* _init_l_Lean_Elab_Level_elabLevel___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Level_elabLevel___closed__6; +x_1 = lean_box(0); x_2 = l_Lean_Elab_Level_elabLevel___closed__15; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -1336,20 +1338,56 @@ static lean_object* _init_l_Lean_Elab_Level_elabLevel___closed__17() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("unexpected universe level syntax kind", 37); +x_1 = lean_mk_string_from_bytes("ident", 5); return x_1; } } static lean_object* _init_l_Lean_Elab_Level_elabLevel___closed__18() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Level_elabLevel___closed__17; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Level_elabLevel___closed__19() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("addLit", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Level_elabLevel___closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Level_elabLevel___closed__6; +x_2 = l_Lean_Elab_Level_elabLevel___closed__19; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Level_elabLevel___closed__21() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("unexpected universe level syntax kind", 37); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Level_elabLevel___closed__22() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Level_elabLevel___closed__17; +x_1 = l_Lean_Elab_Level_elabLevel___closed__21; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Level_elabLevel___closed__19() { +static lean_object* _init_l_Lean_Elab_Level_elabLevel___closed__23() { _start: { lean_object* x_1; @@ -1357,16 +1395,16 @@ x_1 = lean_mk_string_from_bytes("unknown universe level '", 24); return x_1; } } -static lean_object* _init_l_Lean_Elab_Level_elabLevel___closed__20() { +static lean_object* _init_l_Lean_Elab_Level_elabLevel___closed__24() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Level_elabLevel___closed__19; +x_1 = l_Lean_Elab_Level_elabLevel___closed__23; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Level_elabLevel___closed__21() { +static lean_object* _init_l_Lean_Elab_Level_elabLevel___closed__25() { _start: { lean_object* x_1; @@ -1374,16 +1412,16 @@ x_1 = lean_mk_string_from_bytes("'", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Level_elabLevel___closed__22() { +static lean_object* _init_l_Lean_Elab_Level_elabLevel___closed__26() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Level_elabLevel___closed__21; +x_1 = l_Lean_Elab_Level_elabLevel___closed__25; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Level_elabLevel___closed__23() { +static lean_object* _init_l_Lean_Elab_Level_elabLevel___closed__27() { _start: { lean_object* x_1; @@ -1428,25 +1466,25 @@ x_17 = lean_name_eq(x_4, x_16); if (x_17 == 0) { lean_object* x_18; uint8_t x_19; -x_18 = l_Lean_numLitKind; +x_18 = l_Lean_Elab_Level_elabLevel___closed__16; x_19 = lean_name_eq(x_4, x_18); if (x_19 == 0) { lean_object* x_20; uint8_t x_21; -x_20 = l_Lean_identKind; +x_20 = l_Lean_Elab_Level_elabLevel___closed__18; x_21 = lean_name_eq(x_4, x_20); if (x_21 == 0) { lean_object* x_22; uint8_t x_23; lean_dec(x_8); -x_22 = l_Lean_Elab_Level_elabLevel___closed__16; +x_22 = l_Lean_Elab_Level_elabLevel___closed__20; x_23 = lean_name_eq(x_4, x_22); lean_dec(x_4); if (x_23 == 0) { lean_object* x_24; lean_object* x_25; lean_dec(x_1); -x_24 = l_Lean_Elab_Level_elabLevel___closed__18; +x_24 = l_Lean_Elab_Level_elabLevel___closed__22; x_25 = l_Lean_throwError___at_Lean_Elab_Level_elabLevel___spec__1(x_24, x_2, x_3); lean_dec(x_2); return x_25; @@ -1469,7 +1507,7 @@ lean_dec(x_28); x_31 = lean_unsigned_to_nat(2u); x_32 = l_Lean_Syntax_getArg(x_1, x_31); lean_dec(x_1); -x_33 = l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(x_18, x_32); +x_33 = l_Lean_Syntax_isNatLit_x3f(x_32); lean_dec(x_32); if (lean_obj_tag(x_33) == 0) { @@ -1589,11 +1627,11 @@ lean_dec(x_52); lean_dec(x_8); x_56 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_56, 0, x_51); -x_57 = l_Lean_Elab_Level_elabLevel___closed__20; +x_57 = l_Lean_Elab_Level_elabLevel___closed__24; x_58 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_58, 0, x_57); lean_ctor_set(x_58, 1, x_56); -x_59 = l_Lean_Elab_Level_elabLevel___closed__22; +x_59 = l_Lean_Elab_Level_elabLevel___closed__26; x_60 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_60, 0, x_58); lean_ctor_set(x_60, 1, x_59); @@ -1621,7 +1659,7 @@ return x_65; else { lean_object* x_66; uint8_t x_67; uint8_t x_68; -x_66 = l_Lean_Elab_Level_elabLevel___closed__23; +x_66 = l_Lean_Elab_Level_elabLevel___closed__27; x_67 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_8, x_66); lean_dec(x_8); lean_inc(x_51); @@ -1634,11 +1672,11 @@ lean_dec(x_53); lean_dec(x_52); x_69 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_69, 0, x_51); -x_70 = l_Lean_Elab_Level_elabLevel___closed__20; +x_70 = l_Lean_Elab_Level_elabLevel___closed__24; x_71 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_71, 0, x_70); lean_ctor_set(x_71, 1, x_69); -x_72 = l_Lean_Elab_Level_elabLevel___closed__22; +x_72 = l_Lean_Elab_Level_elabLevel___closed__26; x_73 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_73, 0, x_71); lean_ctor_set(x_73, 1, x_72); @@ -1725,7 +1763,7 @@ else lean_object* x_92; lean_dec(x_8); lean_dec(x_4); -x_92 = l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(x_18, x_1); +x_92 = l_Lean_Syntax_isNatLit_x3f(x_1); lean_dec(x_1); if (lean_obj_tag(x_92) == 0) { @@ -2249,25 +2287,25 @@ x_222 = lean_name_eq(x_4, x_221); if (x_222 == 0) { lean_object* x_223; uint8_t x_224; -x_223 = l_Lean_numLitKind; +x_223 = l_Lean_Elab_Level_elabLevel___closed__16; x_224 = lean_name_eq(x_4, x_223); if (x_224 == 0) { lean_object* x_225; uint8_t x_226; -x_225 = l_Lean_identKind; +x_225 = l_Lean_Elab_Level_elabLevel___closed__18; x_226 = lean_name_eq(x_4, x_225); if (x_226 == 0) { lean_object* x_227; uint8_t x_228; lean_dec(x_212); -x_227 = l_Lean_Elab_Level_elabLevel___closed__16; +x_227 = l_Lean_Elab_Level_elabLevel___closed__20; x_228 = lean_name_eq(x_4, x_227); lean_dec(x_4); if (x_228 == 0) { lean_object* x_229; lean_object* x_230; lean_dec(x_1); -x_229 = l_Lean_Elab_Level_elabLevel___closed__18; +x_229 = l_Lean_Elab_Level_elabLevel___closed__22; x_230 = l_Lean_throwError___at_Lean_Elab_Level_elabLevel___spec__1(x_229, x_216, x_3); lean_dec(x_216); return x_230; @@ -2290,7 +2328,7 @@ lean_dec(x_233); x_236 = lean_unsigned_to_nat(2u); x_237 = l_Lean_Syntax_getArg(x_1, x_236); lean_dec(x_1); -x_238 = l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(x_223, x_237); +x_238 = l_Lean_Syntax_isNatLit_x3f(x_237); lean_dec(x_237); if (lean_obj_tag(x_238) == 0) { @@ -2411,11 +2449,11 @@ lean_dec(x_255); lean_dec(x_212); x_259 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_259, 0, x_254); -x_260 = l_Lean_Elab_Level_elabLevel___closed__20; +x_260 = l_Lean_Elab_Level_elabLevel___closed__24; x_261 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_261, 0, x_260); lean_ctor_set(x_261, 1, x_259); -x_262 = l_Lean_Elab_Level_elabLevel___closed__22; +x_262 = l_Lean_Elab_Level_elabLevel___closed__26; x_263 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_263, 0, x_261); lean_ctor_set(x_263, 1, x_262); @@ -2445,7 +2483,7 @@ return x_268; else { lean_object* x_269; uint8_t x_270; uint8_t x_271; -x_269 = l_Lean_Elab_Level_elabLevel___closed__23; +x_269 = l_Lean_Elab_Level_elabLevel___closed__27; x_270 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_212, x_269); lean_dec(x_212); lean_inc(x_254); @@ -2458,11 +2496,11 @@ lean_dec(x_256); lean_dec(x_255); x_272 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_272, 0, x_254); -x_273 = l_Lean_Elab_Level_elabLevel___closed__20; +x_273 = l_Lean_Elab_Level_elabLevel___closed__24; x_274 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_274, 0, x_273); lean_ctor_set(x_274, 1, x_272); -x_275 = l_Lean_Elab_Level_elabLevel___closed__22; +x_275 = l_Lean_Elab_Level_elabLevel___closed__26; x_276 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_276, 0, x_274); lean_ctor_set(x_276, 1, x_275); @@ -2539,7 +2577,7 @@ else lean_object* x_289; lean_dec(x_212); lean_dec(x_4); -x_289 = l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(x_223, x_1); +x_289 = l_Lean_Syntax_isNatLit_x3f(x_1); lean_dec(x_1); if (lean_obj_tag(x_289) == 0) { @@ -3114,6 +3152,14 @@ l_Lean_Elab_Level_elabLevel___closed__22 = _init_l_Lean_Elab_Level_elabLevel___c lean_mark_persistent(l_Lean_Elab_Level_elabLevel___closed__22); l_Lean_Elab_Level_elabLevel___closed__23 = _init_l_Lean_Elab_Level_elabLevel___closed__23(); lean_mark_persistent(l_Lean_Elab_Level_elabLevel___closed__23); +l_Lean_Elab_Level_elabLevel___closed__24 = _init_l_Lean_Elab_Level_elabLevel___closed__24(); +lean_mark_persistent(l_Lean_Elab_Level_elabLevel___closed__24); +l_Lean_Elab_Level_elabLevel___closed__25 = _init_l_Lean_Elab_Level_elabLevel___closed__25(); +lean_mark_persistent(l_Lean_Elab_Level_elabLevel___closed__25); +l_Lean_Elab_Level_elabLevel___closed__26 = _init_l_Lean_Elab_Level_elabLevel___closed__26(); +lean_mark_persistent(l_Lean_Elab_Level_elabLevel___closed__26); +l_Lean_Elab_Level_elabLevel___closed__27 = _init_l_Lean_Elab_Level_elabLevel___closed__27(); +lean_mark_persistent(l_Lean_Elab_Level_elabLevel___closed__27); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Elab/Macro.c b/stage0/stdlib/Lean/Elab/Macro.c index 329f244137e..d6561c47767 100644 --- a/stage0/stdlib/Lean/Elab/Macro.c +++ b/stage0/stdlib/Lean/Elab/Macro.c @@ -15,11 +15,11 @@ extern "C" { #endif static lean_object* l_Lean_Elab_Command_elabMacro___lambda__6___closed__2; size_t lean_usize_add(size_t, size_t); +static lean_object* l_Lean_Elab_Command_elabMacro___lambda__1___closed__3; lean_object* l_Lean_Elab_Command_expandMacroArg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); lean_object* l_Lean_Syntax_getHeadInfo_x3f(lean_object*); static lean_object* l_Lean_Elab_Command_elabMacro___closed__5; -extern lean_object* l_Lean_nullKind; lean_object* lean_name_mk_string(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabMacro___lambda__2___closed__10; lean_object* lean_array_uget(lean_object*, size_t); @@ -32,7 +32,6 @@ static lean_object* l_Lean_Elab_Command_elabMacro___lambda__2___closed__3; static lean_object* l_Lean_Elab_Command_elabMacro___lambda__2___closed__16; LEAN_EXPORT lean_object* l_Lean_mkIdentFromRef___at_Lean_Elab_Command_elabMacro___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabMacro___lambda__2___closed__15; -static lean_object* l_Lean_Elab_Command_elabMacro___lambda__2___closed__30; static lean_object* l_Lean_Elab_Command_elabMacro___lambda__2___closed__28; static lean_object* l_Lean_Elab_Command_elabMacro___closed__3; static lean_object* l_Lean_Elab_Command_elabMacro___lambda__2___closed__5; @@ -43,24 +42,23 @@ lean_object* lean_array_push(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabMacro___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabMacro___closed__4; +lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabSyntax___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_mkNameFromParserSyntax(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l_Lean_Elab_Command_elabMacro___lambda__2___closed__18; static lean_object* l_Lean_Elab_Command_elabMacro___lambda__2___closed__25; static lean_object* l_Lean_Elab_Command_elabMacro___lambda__2___closed__14; -lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Command_elabMacro___lambda__2___closed__22; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabMacro___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacro___spec__3(size_t, size_t, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabMacro_declRange___closed__6; -static lean_object* l_Lean_Elab_Command_elabMacro___lambda__2___closed__29; -extern lean_object* l_Lean_numLitKind; +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabMacro___lambda__2___closed__23; static lean_object* l_Lean_Elab_Command_elabMacro___lambda__2___closed__13; static lean_object* l_Lean_Elab_Command_elabMacro___lambda__2___closed__17; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabMacro(lean_object*); static lean_object* l_Lean_Elab_Command_elabMacro___lambda__2___closed__21; -lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabSyntax___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabMacro___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabMacro___lambda__2___closed__6; static lean_object* l_Lean_Elab_Command_elabMacro___lambda__2___closed__2; @@ -84,11 +82,10 @@ lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0_ LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabMacro___lambda__2___boxed(lean_object**); static lean_object* l___regBuiltin_Lean_Elab_Command_elabMacro_declRange___closed__7; size_t lean_usize_of_nat(lean_object*); -lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabSyntax___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabMacro___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabMacro___lambda__5(lean_object*, lean_object*, lean_object*, 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_Command_elabMacro(lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_mkNumLit(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabMacro___lambda__2___closed__27; static lean_object* l_Lean_Elab_Command_elabMacro___lambda__2___closed__26; static lean_object* l___regBuiltin_Lean_Elab_Command_elabMacro_declRange___closed__1; @@ -105,6 +102,7 @@ lean_object* l_Lean_Elab_Command_getMainModule___rarg(lean_object*, lean_object* LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacro___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_getRef(lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabSyntax___spec__9(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isNone(lean_object*); static lean_object* l_Lean_Elab_Command_elabMacro___lambda__2___closed__11; static lean_object* l___regBuiltin_Lean_Elab_Command_elabMacro___closed__5; @@ -115,6 +113,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Command_elabMacro___closed__3; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabMacro___lambda__2___closed__12; static lean_object* l_Lean_Elab_Command_elabMacro___closed__1; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacro___spec__3___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabMacro___closed__8; lean_object* l_Lean_Elab_Command_getScope___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabMacro___lambda__1___closed__1; @@ -122,10 +121,10 @@ static lean_object* l_Lean_Elab_Command_elabMacro___closed__9; static lean_object* l_Lean_Elab_Command_elabMacro___closed__4; static lean_object* l_Lean_Elab_Command_elabMacro___lambda__2___closed__20; static lean_object* l_Lean_Elab_Command_elabMacro___lambda__3___closed__1; +static lean_object* l_Lean_Elab_Command_elabMacro___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabMacro___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabMacro___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabMacro___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Syntax_mkLit(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacro___spec__1(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { @@ -224,6 +223,30 @@ return x_12; } } } +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacro___spec__3(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = 1; +x_9 = lean_usize_add(x_2, x_8); +x_10 = lean_array_uset(x_7, x_2, x_5); +x_2 = x_9; +x_3 = x_10; +goto _start; +} +} +} static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__1___closed__1() { _start: { @@ -233,6 +256,24 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } +static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("null", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Command_elabMacro___lambda__1___closed__2; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabMacro___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { @@ -241,7 +282,7 @@ x_6 = l_Lean_Elab_Command_elabMacro___lambda__1___closed__1; x_7 = lean_array_push(x_6, x_1); x_8 = lean_array_push(x_7, x_2); x_9 = lean_box(2); -x_10 = l_Lean_nullKind; +x_10 = l_Lean_Elab_Command_elabMacro___lambda__1___closed__3; x_11 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_11, 0, x_9); lean_ctor_set(x_11, 1, x_10); @@ -261,37 +302,19 @@ return x_1; static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__2() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("null", 4); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__2; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__4() { -_start: -{ lean_object* x_1; lean_object* x_2; x_1 = lean_unsigned_to_nat(0u); x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__5() { +static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__3; -x_3 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__4; +x_2 = l_Lean_Elab_Command_elabMacro___lambda__1___closed__3; +x_3 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__2; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -299,7 +322,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__6() { +static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__4() { _start: { lean_object* x_1; lean_object* x_2; @@ -308,17 +331,17 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__7() { +static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__6; -x_2 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__5; +x_1 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__4; +x_2 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__3; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__8() { +static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__6() { _start: { lean_object* x_1; @@ -326,7 +349,7 @@ x_1 = lean_mk_string_from_bytes("matchAlts", 9); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__9() { +static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__7() { _start: { lean_object* x_1; @@ -334,7 +357,7 @@ x_1 = lean_mk_string_from_bytes("matchAlt", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__10() { +static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__8() { _start: { lean_object* x_1; @@ -342,7 +365,7 @@ x_1 = lean_mk_string_from_bytes("|", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__11() { +static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__9() { _start: { lean_object* x_1; @@ -350,7 +373,7 @@ x_1 = lean_mk_string_from_bytes("quot", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__12() { +static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__10() { _start: { lean_object* x_1; @@ -358,7 +381,7 @@ x_1 = lean_mk_string_from_bytes("`(", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__13() { +static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__11() { _start: { lean_object* x_1; @@ -366,7 +389,7 @@ x_1 = lean_mk_string_from_bytes(")", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__14() { +static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__12() { _start: { lean_object* x_1; lean_object* x_2; @@ -375,7 +398,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__15() { +static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__13() { _start: { lean_object* x_1; @@ -383,7 +406,7 @@ x_1 = lean_mk_string_from_bytes("=>", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__16() { +static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__14() { _start: { lean_object* x_1; lean_object* x_2; @@ -392,7 +415,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__17() { +static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__15() { _start: { lean_object* x_1; lean_object* x_2; @@ -401,22 +424,22 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__18() { +static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__16() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__4; +x_1 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__2; x_2 = l_Array_append___rarg(x_1, x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__19() { +static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__3; -x_3 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__18; +x_2 = l_Lean_Elab_Command_elabMacro___lambda__1___closed__3; +x_3 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__16; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -424,17 +447,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__20() { +static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__17; -x_2 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__19; +x_1 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__15; +x_2 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__17; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__21() { +static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__19() { _start: { lean_object* x_1; @@ -442,7 +465,7 @@ x_1 = lean_mk_string_from_bytes("syntax", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__22() { +static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__20() { _start: { lean_object* x_1; @@ -450,7 +473,7 @@ x_1 = lean_mk_string_from_bytes("namedName", 9); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__23() { +static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__21() { _start: { lean_object* x_1; @@ -458,7 +481,7 @@ x_1 = lean_mk_string_from_bytes("(", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__24() { +static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__22() { _start: { lean_object* x_1; @@ -466,7 +489,7 @@ x_1 = lean_mk_string_from_bytes("name", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__25() { +static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__23() { _start: { lean_object* x_1; @@ -474,7 +497,7 @@ x_1 = lean_mk_string_from_bytes(":=", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__26() { +static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__24() { _start: { lean_object* x_1; @@ -482,7 +505,7 @@ x_1 = lean_mk_string_from_bytes("namedPrio", 9); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__27() { +static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__25() { _start: { lean_object* x_1; @@ -490,7 +513,7 @@ x_1 = lean_mk_string_from_bytes("priority", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__28() { +static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__26() { _start: { lean_object* x_1; @@ -498,7 +521,7 @@ x_1 = lean_mk_string_from_bytes(":", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__29() { +static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__27() { _start: { lean_object* x_1; lean_object* x_2; @@ -507,7 +530,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__30() { +static lean_object* _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__28() { _start: { lean_object* x_1; @@ -518,7 +541,7 @@ return x_1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabMacro___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_object* x_9, size_t 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) { _start: { -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; lean_object* x_30; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; size_t x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; +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; lean_object* x_30; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; size_t x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; x_19 = l_Lean_Elab_Command_getScope___rarg(x_17, x_18); x_20 = lean_ctor_get(x_19, 0); lean_inc(x_20); @@ -556,34 +579,34 @@ x_221 = l_Lean_Elab_Command_getMainModule___rarg(x_17, x_220); x_222 = lean_ctor_get(x_221, 1); lean_inc(x_222); lean_dec(x_221); -x_223 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__21; +x_223 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__19; lean_inc(x_3); x_224 = lean_name_mk_string(x_3, x_223); x_225 = l_Lean_Syntax_getHeadInfo_x3f(x_5); -x_226 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__22; +x_226 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__20; lean_inc(x_3); x_227 = lean_name_mk_string(x_3, x_226); -x_228 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__23; +x_228 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__21; lean_inc(x_217); x_229 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_229, 0, x_217); lean_ctor_set(x_229, 1, x_228); -x_230 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__24; +x_230 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__22; lean_inc(x_217); x_231 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_231, 0, x_217); lean_ctor_set(x_231, 1, x_230); -x_232 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__25; +x_232 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__23; lean_inc(x_217); x_233 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_233, 0, x_217); lean_ctor_set(x_233, 1, x_232); -x_234 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__13; +x_234 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__11; lean_inc(x_217); x_235 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_235, 0, x_217); lean_ctor_set(x_235, 1, x_234); -x_236 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__17; +x_236 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__15; x_237 = lean_array_push(x_236, x_229); lean_inc(x_237); x_238 = lean_array_push(x_237, x_231); @@ -596,65 +619,64 @@ x_242 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_242, 0, x_24); lean_ctor_set(x_242, 1, x_227); lean_ctor_set(x_242, 2, x_241); -x_243 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__6; +x_243 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__4; x_244 = lean_array_push(x_243, x_242); -x_245 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__3; +x_245 = l_Lean_Elab_Command_elabMacro___lambda__1___closed__3; x_246 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_246, 0, x_24); lean_ctor_set(x_246, 1, x_245); lean_ctor_set(x_246, 2, x_244); -x_247 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__26; +x_247 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__24; lean_inc(x_3); x_248 = lean_name_mk_string(x_3, x_247); -x_249 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__27; +x_249 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__25; lean_inc(x_217); x_250 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_250, 0, x_217); lean_ctor_set(x_250, 1, x_249); x_251 = l_Nat_repr(x_8); -x_252 = l_Lean_numLitKind; -x_253 = l_Lean_Syntax_mkLit(x_252, x_251, x_24); -x_254 = lean_array_push(x_237, x_250); -x_255 = lean_array_push(x_254, x_233); -x_256 = lean_array_push(x_255, x_253); -x_257 = lean_array_push(x_256, x_235); -x_258 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_258, 0, x_24); -lean_ctor_set(x_258, 1, x_248); -lean_ctor_set(x_258, 2, x_257); -x_259 = lean_array_push(x_243, x_258); -x_260 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_260, 0, x_24); -lean_ctor_set(x_260, 1, x_245); -lean_ctor_set(x_260, 2, x_259); -x_261 = lean_array_get_size(x_9); -x_262 = lean_usize_of_nat(x_261); -lean_dec(x_261); -x_263 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_262, x_10, x_9); -x_264 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__4; -x_265 = l_Array_append___rarg(x_264, x_263); -x_266 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_266, 0, x_24); -lean_ctor_set(x_266, 1, x_245); -lean_ctor_set(x_266, 2, x_265); -x_267 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__28; +x_252 = l_Lean_Syntax_mkNumLit(x_251, x_24); +x_253 = lean_array_push(x_237, x_250); +x_254 = lean_array_push(x_253, x_233); +x_255 = lean_array_push(x_254, x_252); +x_256 = lean_array_push(x_255, x_235); +x_257 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_257, 0, x_24); +lean_ctor_set(x_257, 1, x_248); +lean_ctor_set(x_257, 2, x_256); +x_258 = lean_array_push(x_243, x_257); +x_259 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_259, 0, x_24); +lean_ctor_set(x_259, 1, x_245); +lean_ctor_set(x_259, 2, x_258); +x_260 = lean_array_get_size(x_9); +x_261 = lean_usize_of_nat(x_260); +lean_dec(x_260); +x_262 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacro___spec__3(x_261, x_10, x_9); +x_263 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__2; +x_264 = l_Array_append___rarg(x_263, x_262); +x_265 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_265, 0, x_24); +lean_ctor_set(x_265, 1, x_245); +lean_ctor_set(x_265, 2, x_264); +x_266 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__26; lean_inc(x_217); -x_268 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_268, 0, x_217); -lean_ctor_set(x_268, 1, x_267); +x_267 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_267, 0, x_217); +lean_ctor_set(x_267, 1, x_266); if (lean_obj_tag(x_7) == 0) { -x_269 = x_264; -goto block_305; +x_268 = x_263; +goto block_304; } else { -lean_object* x_306; lean_object* x_307; -x_306 = lean_ctor_get(x_7, 0); -lean_inc(x_306); -x_307 = lean_array_push(x_243, x_306); -x_269 = x_307; -goto block_305; +lean_object* x_305; lean_object* x_306; +x_305 = lean_ctor_get(x_7, 0); +lean_inc(x_305); +x_306 = lean_array_push(x_243, x_305); +x_268 = x_306; +goto block_304; } block_215: { @@ -686,37 +708,37 @@ lean_inc(x_42); lean_dec(x_41); x_43 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__1; x_44 = lean_name_mk_string(x_3, x_43); -x_45 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__7; +x_45 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__5; x_46 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_46, 0, x_24); lean_ctor_set(x_46, 1, x_4); lean_ctor_set(x_46, 2, x_45); x_47 = l_Lean_Syntax_getHeadInfo_x3f(x_5); lean_dec(x_5); -x_48 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__8; +x_48 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__6; lean_inc(x_6); x_49 = lean_name_mk_string(x_6, x_48); -x_50 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__9; +x_50 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__7; lean_inc(x_6); x_51 = lean_name_mk_string(x_6, x_50); -x_52 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__10; +x_52 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__8; lean_inc(x_37); x_53 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_53, 0, x_37); lean_ctor_set(x_53, 1, x_52); -x_54 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__11; +x_54 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__9; x_55 = lean_name_mk_string(x_6, x_54); -x_56 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__12; +x_56 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__10; lean_inc(x_37); x_57 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_57, 0, x_37); lean_ctor_set(x_57, 1, x_56); -x_58 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__13; +x_58 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__11; lean_inc(x_37); x_59 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_59, 0, x_37); lean_ctor_set(x_59, 1, x_58); -x_60 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__14; +x_60 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__12; x_61 = lean_array_push(x_60, x_57); lean_inc(x_61); x_62 = lean_array_push(x_61, x_25); @@ -727,9 +749,9 @@ x_64 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_64, 0, x_24); lean_ctor_set(x_64, 1, x_55); lean_ctor_set(x_64, 2, x_63); -x_65 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__6; +x_65 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__4; x_66 = lean_array_push(x_65, x_64); -x_67 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__3; +x_67 = l_Lean_Elab_Command_elabMacro___lambda__1___closed__3; x_68 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_68, 0, x_24); lean_ctor_set(x_68, 1, x_67); @@ -739,7 +761,7 @@ x_70 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_70, 0, x_24); lean_ctor_set(x_70, 1, x_67); lean_ctor_set(x_70, 2, x_69); -x_71 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__15; +x_71 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__13; lean_inc(x_37); x_72 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_72, 0, x_37); @@ -750,7 +772,7 @@ x_75 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_75, 0, x_24); lean_ctor_set(x_75, 1, x_55); lean_ctor_set(x_75, 2, x_74); -x_76 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__16; +x_76 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__14; x_77 = lean_array_push(x_76, x_53); x_78 = lean_array_push(x_77, x_70); x_79 = lean_array_push(x_78, x_72); @@ -772,7 +794,7 @@ lean_ctor_set(x_85, 2, x_84); if (lean_obj_tag(x_7) == 0) { lean_object* x_86; lean_object* x_87; -x_86 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__20; +x_86 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__18; x_87 = lean_array_push(x_86, x_46); if (lean_obj_tag(x_47) == 0) { @@ -781,7 +803,7 @@ x_88 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_88, 0, x_37); lean_ctor_set(x_88, 1, x_43); x_89 = lean_array_push(x_87, x_88); -x_90 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__5; +x_90 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__3; x_91 = lean_array_push(x_89, x_90); x_92 = lean_array_push(x_91, x_85); x_93 = lean_alloc_ctor(1, 3, 0); @@ -802,7 +824,7 @@ x_96 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_96, 0, x_95); lean_ctor_set(x_96, 1, x_43); x_97 = lean_array_push(x_87, x_96); -x_98 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__5; +x_98 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__3; x_99 = lean_array_push(x_97, x_98); x_100 = lean_array_push(x_99, x_85); x_101 = lean_alloc_ctor(1, 3, 0); @@ -820,13 +842,13 @@ x_103 = lean_ctor_get(x_7, 0); lean_inc(x_103); lean_dec(x_7); x_104 = lean_array_push(x_65, x_103); -x_105 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__4; +x_105 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__2; x_106 = l_Array_append___rarg(x_105, x_104); x_107 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_107, 0, x_24); lean_ctor_set(x_107, 1, x_67); lean_ctor_set(x_107, 2, x_106); -x_108 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__17; +x_108 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__15; x_109 = lean_array_push(x_108, x_107); x_110 = lean_array_push(x_109, x_46); if (lean_obj_tag(x_47) == 0) @@ -836,7 +858,7 @@ x_111 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_111, 0, x_37); lean_ctor_set(x_111, 1, x_43); x_112 = lean_array_push(x_110, x_111); -x_113 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__5; +x_113 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__3; x_114 = lean_array_push(x_112, x_113); x_115 = lean_array_push(x_114, x_85); x_116 = lean_alloc_ctor(1, 3, 0); @@ -857,7 +879,7 @@ x_119 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_119, 0, x_118); lean_ctor_set(x_119, 1, x_43); x_120 = lean_array_push(x_110, x_119); -x_121 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__5; +x_121 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__3; x_122 = lean_array_push(x_120, x_121); x_123 = lean_array_push(x_122, x_85); x_124 = lean_alloc_ctor(1, 3, 0); @@ -891,37 +913,37 @@ lean_inc(x_134); lean_dec(x_133); x_135 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__1; x_136 = lean_name_mk_string(x_3, x_135); -x_137 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__7; +x_137 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__5; x_138 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_138, 0, x_24); lean_ctor_set(x_138, 1, x_4); lean_ctor_set(x_138, 2, x_137); x_139 = l_Lean_Syntax_getHeadInfo_x3f(x_5); lean_dec(x_5); -x_140 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__8; +x_140 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__6; lean_inc(x_6); x_141 = lean_name_mk_string(x_6, x_140); -x_142 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__9; +x_142 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__7; lean_inc(x_6); x_143 = lean_name_mk_string(x_6, x_142); -x_144 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__10; +x_144 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__8; lean_inc(x_129); x_145 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_145, 0, x_129); lean_ctor_set(x_145, 1, x_144); -x_146 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__11; +x_146 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__9; x_147 = lean_name_mk_string(x_6, x_146); -x_148 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__12; +x_148 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__10; lean_inc(x_129); x_149 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_149, 0, x_129); lean_ctor_set(x_149, 1, x_148); -x_150 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__13; +x_150 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__11; lean_inc(x_129); x_151 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_151, 0, x_129); lean_ctor_set(x_151, 1, x_150); -x_152 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__14; +x_152 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__12; x_153 = lean_array_push(x_152, x_149); x_154 = lean_array_push(x_153, x_25); x_155 = lean_array_push(x_154, x_151); @@ -929,9 +951,9 @@ x_156 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_156, 0, x_24); lean_ctor_set(x_156, 1, x_147); lean_ctor_set(x_156, 2, x_155); -x_157 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__6; +x_157 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__4; x_158 = lean_array_push(x_157, x_156); -x_159 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__3; +x_159 = l_Lean_Elab_Command_elabMacro___lambda__1___closed__3; x_160 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_160, 0, x_24); lean_ctor_set(x_160, 1, x_159); @@ -941,12 +963,12 @@ x_162 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_162, 0, x_24); lean_ctor_set(x_162, 1, x_159); lean_ctor_set(x_162, 2, x_161); -x_163 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__15; +x_163 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__13; lean_inc(x_129); x_164 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_164, 0, x_129); lean_ctor_set(x_164, 1, x_163); -x_165 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__16; +x_165 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__14; x_166 = lean_array_push(x_165, x_145); x_167 = lean_array_push(x_166, x_162); x_168 = lean_array_push(x_167, x_164); @@ -968,7 +990,7 @@ lean_ctor_set(x_174, 2, x_173); if (lean_obj_tag(x_7) == 0) { lean_object* x_175; lean_object* x_176; -x_175 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__20; +x_175 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__18; x_176 = lean_array_push(x_175, x_138); if (lean_obj_tag(x_139) == 0) { @@ -977,7 +999,7 @@ x_177 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_177, 0, x_129); lean_ctor_set(x_177, 1, x_135); x_178 = lean_array_push(x_176, x_177); -x_179 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__5; +x_179 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__3; x_180 = lean_array_push(x_178, x_179); x_181 = lean_array_push(x_180, x_174); x_182 = lean_alloc_ctor(1, 3, 0); @@ -998,7 +1020,7 @@ x_185 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_185, 0, x_184); lean_ctor_set(x_185, 1, x_135); x_186 = lean_array_push(x_176, x_185); -x_187 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__5; +x_187 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__3; x_188 = lean_array_push(x_186, x_187); x_189 = lean_array_push(x_188, x_174); x_190 = lean_alloc_ctor(1, 3, 0); @@ -1016,13 +1038,13 @@ x_192 = lean_ctor_get(x_7, 0); lean_inc(x_192); lean_dec(x_7); x_193 = lean_array_push(x_157, x_192); -x_194 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__4; +x_194 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__2; x_195 = l_Array_append___rarg(x_194, x_193); x_196 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_196, 0, x_24); lean_ctor_set(x_196, 1, x_159); lean_ctor_set(x_196, 2, x_195); -x_197 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__17; +x_197 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__15; x_198 = lean_array_push(x_197, x_196); x_199 = lean_array_push(x_198, x_138); if (lean_obj_tag(x_139) == 0) @@ -1032,7 +1054,7 @@ x_200 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_200, 0, x_129); lean_ctor_set(x_200, 1, x_135); x_201 = lean_array_push(x_199, x_200); -x_202 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__5; +x_202 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__3; x_203 = lean_array_push(x_201, x_202); x_204 = lean_array_push(x_203, x_174); x_205 = lean_alloc_ctor(1, 3, 0); @@ -1053,7 +1075,7 @@ x_208 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_208, 0, x_207); lean_ctor_set(x_208, 1, x_135); x_209 = lean_array_push(x_199, x_208); -x_210 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__5; +x_210 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__3; x_211 = lean_array_push(x_209, x_210); x_212 = lean_array_push(x_211, x_174); x_213 = lean_alloc_ctor(1, 3, 0); @@ -1066,91 +1088,91 @@ return x_214; } } } -block_305: +block_304: { -lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; -x_270 = l_Array_append___rarg(x_264, x_269); -x_271 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_271, 0, x_24); -lean_ctor_set(x_271, 1, x_245); -lean_ctor_set(x_271, 2, x_270); -x_272 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__29; -x_273 = lean_array_push(x_272, x_271); -x_274 = lean_array_push(x_273, x_11); +lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; +x_269 = l_Array_append___rarg(x_263, x_268); +x_270 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_270, 0, x_24); +lean_ctor_set(x_270, 1, x_245); +lean_ctor_set(x_270, 2, x_269); +x_271 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__27; +x_272 = lean_array_push(x_271, x_270); +x_273 = lean_array_push(x_272, x_11); if (lean_obj_tag(x_225) == 0) { -x_275 = x_217; -goto block_303; +x_274 = x_217; +goto block_302; } else { -lean_object* x_304; +lean_object* x_303; lean_dec(x_217); -x_304 = lean_ctor_get(x_225, 0); -lean_inc(x_304); +x_303 = lean_ctor_get(x_225, 0); +lean_inc(x_303); lean_dec(x_225); -x_275 = x_304; -goto block_303; +x_274 = x_303; +goto block_302; } -block_303: +block_302: { -lean_object* x_276; lean_object* x_277; -x_276 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_276, 0, x_275); -lean_ctor_set(x_276, 1, x_223); -x_277 = lean_array_push(x_274, x_276); +lean_object* x_275; lean_object* x_276; +x_275 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_275, 0, x_274); +lean_ctor_set(x_275, 1, x_223); +x_276 = lean_array_push(x_273, x_275); if (lean_obj_tag(x_12) == 0) { -lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; +lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_dec(x_14); -x_278 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__19; -x_279 = lean_array_push(x_277, x_278); -x_280 = lean_array_push(x_279, x_246); -x_281 = lean_array_push(x_280, x_260); -x_282 = lean_array_push(x_281, x_266); -x_283 = lean_array_push(x_282, x_268); -x_284 = lean_array_push(x_283, x_13); -x_285 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_285, 0, x_24); -lean_ctor_set(x_285, 1, x_224); -lean_ctor_set(x_285, 2, x_284); -x_29 = x_285; +x_277 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__17; +x_278 = lean_array_push(x_276, x_277); +x_279 = lean_array_push(x_278, x_246); +x_280 = lean_array_push(x_279, x_259); +x_281 = lean_array_push(x_280, x_265); +x_282 = lean_array_push(x_281, x_267); +x_283 = lean_array_push(x_282, x_13); +x_284 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_284, 0, x_24); +lean_ctor_set(x_284, 1, x_224); +lean_ctor_set(x_284, 2, x_283); +x_29 = x_284; x_30 = x_222; goto block_215; } else { -lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; -x_286 = lean_ctor_get(x_12, 0); -lean_inc(x_286); +lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; +x_285 = lean_ctor_get(x_12, 0); +lean_inc(x_285); lean_dec(x_12); -x_287 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__30; -x_288 = lean_name_mk_string(x_14, x_287); -x_289 = l_Lean_Elab_Command_elabMacro___lambda__1___closed__1; -lean_inc(x_268); -x_290 = lean_array_push(x_289, x_268); -x_291 = lean_array_push(x_290, x_286); -x_292 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_292, 0, x_24); -lean_ctor_set(x_292, 1, x_288); -lean_ctor_set(x_292, 2, x_291); -x_293 = lean_array_push(x_243, x_292); -x_294 = l_Array_append___rarg(x_264, x_293); -x_295 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_295, 0, x_24); -lean_ctor_set(x_295, 1, x_245); -lean_ctor_set(x_295, 2, x_294); -x_296 = lean_array_push(x_277, x_295); -x_297 = lean_array_push(x_296, x_246); -x_298 = lean_array_push(x_297, x_260); -x_299 = lean_array_push(x_298, x_266); -x_300 = lean_array_push(x_299, x_268); -x_301 = lean_array_push(x_300, x_13); -x_302 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_302, 0, x_24); -lean_ctor_set(x_302, 1, x_224); -lean_ctor_set(x_302, 2, x_301); -x_29 = x_302; +x_286 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__28; +x_287 = lean_name_mk_string(x_14, x_286); +x_288 = l_Lean_Elab_Command_elabMacro___lambda__1___closed__1; +lean_inc(x_267); +x_289 = lean_array_push(x_288, x_267); +x_290 = lean_array_push(x_289, x_285); +x_291 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_291, 0, x_24); +lean_ctor_set(x_291, 1, x_287); +lean_ctor_set(x_291, 2, x_290); +x_292 = lean_array_push(x_243, x_291); +x_293 = l_Array_append___rarg(x_263, x_292); +x_294 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_294, 0, x_24); +lean_ctor_set(x_294, 1, x_245); +lean_ctor_set(x_294, 2, x_293); +x_295 = lean_array_push(x_276, x_294); +x_296 = lean_array_push(x_295, x_246); +x_297 = lean_array_push(x_296, x_259); +x_298 = lean_array_push(x_297, x_265); +x_299 = lean_array_push(x_298, x_267); +x_300 = lean_array_push(x_299, x_13); +x_301 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_301, 0, x_24); +lean_ctor_set(x_301, 1, x_224); +lean_ctor_set(x_301, 2, x_300); +x_29 = x_301; x_30 = x_222; goto block_215; } @@ -1216,7 +1238,7 @@ x_29 = lean_alloc_closure((void*)(l_Lean_evalOptPrio), 3, 1); lean_closure_set(x_29, 0, x_12); lean_inc(x_14); lean_inc(x_13); -x_30 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabSyntax___spec__2(x_29, x_13, x_14, x_15); +x_30 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabSyntax___spec__4(x_29, x_13, x_14, x_15); if (lean_obj_tag(x_30) == 0) { lean_object* x_31; lean_object* x_32; lean_object* x_33; size_t x_34; size_t x_35; lean_object* x_36; @@ -1252,7 +1274,7 @@ lean_inc(x_41); lean_dec(x_39); x_42 = l_Lean_Syntax_getId(x_25); x_43 = lean_box(2); -x_44 = l_Lean_nullKind; +x_44 = l_Lean_Elab_Command_elabMacro___lambda__1___closed__3; lean_inc(x_40); x_45 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_45, 0, x_43); @@ -1263,7 +1285,7 @@ lean_closure_set(x_46, 0, x_42); lean_closure_set(x_46, 1, x_45); lean_inc(x_14); lean_inc(x_13); -x_47 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabSyntax___spec__7(x_46, x_13, x_14, x_38); +x_47 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabSyntax___spec__9(x_46, x_13, x_14, x_38); if (lean_obj_tag(x_47) == 0) { lean_object* x_48; lean_object* x_49; lean_object* x_50; @@ -1416,14 +1438,13 @@ x_16 = l_Lean_Syntax_getArg(x_1, x_15); x_17 = l_Lean_Syntax_isNone(x_16); if (x_17 == 0) { -lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_18 = l_Lean_nullKind; -x_19 = lean_unsigned_to_nat(1u); +lean_object* x_18; uint8_t x_19; +x_18 = lean_unsigned_to_nat(1u); lean_inc(x_16); -x_20 = l_Lean_Syntax_isNodeOf(x_16, x_18, x_19); -if (x_20 == 0) +x_19 = l_Lean_Syntax_matchesNull(x_16, x_18); +if (x_19 == 0) { -lean_object* x_21; +lean_object* x_20; lean_dec(x_16); lean_dec(x_13); lean_dec(x_12); @@ -1437,25 +1458,25 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_21 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_14); -return x_21; +x_20 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_14); +return x_20; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; -x_22 = lean_unsigned_to_nat(0u); -x_23 = l_Lean_Syntax_getArg(x_16, x_22); +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_21 = lean_unsigned_to_nat(0u); +x_22 = l_Lean_Syntax_getArg(x_16, x_21); lean_dec(x_16); -x_24 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__26; +x_23 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__24; lean_inc(x_2); -x_25 = lean_name_mk_string(x_2, x_24); -lean_inc(x_23); -x_26 = l_Lean_Syntax_isOfKind(x_23, x_25); -lean_dec(x_25); -if (x_26 == 0) +x_24 = lean_name_mk_string(x_2, x_23); +lean_inc(x_22); +x_25 = l_Lean_Syntax_isOfKind(x_22, x_24); +lean_dec(x_24); +if (x_25 == 0) { -lean_object* x_27; -lean_dec(x_23); +lean_object* x_26; +lean_dec(x_22); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -1468,31 +1489,31 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_27 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_14); -return x_27; +x_26 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_14); +return x_26; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_28 = lean_unsigned_to_nat(3u); -x_29 = l_Lean_Syntax_getArg(x_23, x_28); -lean_dec(x_23); -x_30 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_30, 0, x_29); -x_31 = lean_box(0); -x_32 = l_Lean_Elab_Command_elabMacro___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_11, x_31, x_30, x_12, x_13, x_14); -return x_32; +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_27 = lean_unsigned_to_nat(3u); +x_28 = l_Lean_Syntax_getArg(x_22, x_27); +lean_dec(x_22); +x_29 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_29, 0, x_28); +x_30 = lean_box(0); +x_31 = l_Lean_Elab_Command_elabMacro___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_11, x_30, x_29, x_12, x_13, x_14); +return x_31; } } } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_dec(x_16); +x_32 = lean_box(0); x_33 = lean_box(0); -x_34 = lean_box(0); -x_35 = l_Lean_Elab_Command_elabMacro___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_11, x_34, x_33, x_12, x_13, x_14); -return x_35; +x_34 = l_Lean_Elab_Command_elabMacro___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_11, x_33, x_32, x_12, x_13, x_14); +return x_34; } } } @@ -1506,14 +1527,13 @@ x_15 = l_Lean_Syntax_getArg(x_1, x_14); x_16 = l_Lean_Syntax_isNone(x_15); if (x_16 == 0) { -lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_17 = l_Lean_nullKind; -x_18 = lean_unsigned_to_nat(1u); +lean_object* x_17; uint8_t x_18; +x_17 = lean_unsigned_to_nat(1u); lean_inc(x_15); -x_19 = l_Lean_Syntax_isNodeOf(x_15, x_17, x_18); -if (x_19 == 0) +x_18 = l_Lean_Syntax_matchesNull(x_15, x_17); +if (x_18 == 0) { -lean_object* x_20; +lean_object* x_19; lean_dec(x_15); lean_dec(x_12); lean_dec(x_11); @@ -1526,25 +1546,25 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_20 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_13); -return x_20; +x_19 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_13); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = lean_unsigned_to_nat(0u); -x_22 = l_Lean_Syntax_getArg(x_15, x_21); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = lean_unsigned_to_nat(0u); +x_21 = l_Lean_Syntax_getArg(x_15, x_20); lean_dec(x_15); -x_23 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__22; +x_22 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__20; lean_inc(x_2); -x_24 = lean_name_mk_string(x_2, x_23); -lean_inc(x_22); -x_25 = l_Lean_Syntax_isOfKind(x_22, x_24); -lean_dec(x_24); -if (x_25 == 0) +x_23 = lean_name_mk_string(x_2, x_22); +lean_inc(x_21); +x_24 = l_Lean_Syntax_isOfKind(x_21, x_23); +lean_dec(x_23); +if (x_24 == 0) { -lean_object* x_26; -lean_dec(x_22); +lean_object* x_25; +lean_dec(x_21); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -1556,31 +1576,31 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_26 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_13); -return x_26; +x_25 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_13); +return x_25; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_27 = lean_unsigned_to_nat(3u); -x_28 = l_Lean_Syntax_getArg(x_22, x_27); -lean_dec(x_22); -x_29 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_29, 0, x_28); -x_30 = lean_box(0); -x_31 = l_Lean_Elab_Command_elabMacro___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_10, x_8, x_30, x_29, x_11, x_12, x_13); -return x_31; +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_26 = lean_unsigned_to_nat(3u); +x_27 = l_Lean_Syntax_getArg(x_21, x_26); +lean_dec(x_21); +x_28 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_28, 0, x_27); +x_29 = lean_box(0); +x_30 = l_Lean_Elab_Command_elabMacro___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_10, x_8, x_29, x_28, x_11, x_12, x_13); +return x_30; } } } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_dec(x_15); +x_31 = lean_box(0); x_32 = lean_box(0); -x_33 = lean_box(0); -x_34 = l_Lean_Elab_Command_elabMacro___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_10, x_8, x_33, x_32, x_11, x_12, x_13); -return x_34; +x_33 = l_Lean_Elab_Command_elabMacro___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_10, x_8, x_32, x_31, x_11, x_12, x_13); +return x_33; } } } @@ -1640,13 +1660,12 @@ x_20 = l_Lean_Syntax_getArg(x_1, x_19); x_21 = l_Lean_Syntax_isNone(x_20); if (x_21 == 0) { -lean_object* x_22; uint8_t x_23; -x_22 = l_Lean_nullKind; +uint8_t x_22; lean_inc(x_20); -x_23 = l_Lean_Syntax_isNodeOf(x_20, x_22, x_9); -if (x_23 == 0) +x_22 = l_Lean_Syntax_matchesNull(x_20, x_9); +if (x_22 == 0) { -lean_object* x_24; +lean_object* x_23; lean_dec(x_20); lean_dec(x_18); lean_dec(x_14); @@ -1658,25 +1677,25 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_24 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_8); -return x_24; +x_23 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_8); +return x_23; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_25 = lean_unsigned_to_nat(0u); -x_26 = l_Lean_Syntax_getArg(x_20, x_25); +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_24 = lean_unsigned_to_nat(0u); +x_25 = l_Lean_Syntax_getArg(x_20, x_24); lean_dec(x_20); -x_27 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__30; +x_26 = l_Lean_Elab_Command_elabMacro___lambda__2___closed__28; lean_inc(x_2); -x_28 = lean_name_mk_string(x_2, x_27); -lean_inc(x_26); -x_29 = l_Lean_Syntax_isOfKind(x_26, x_28); -lean_dec(x_28); -if (x_29 == 0) +x_27 = lean_name_mk_string(x_2, x_26); +lean_inc(x_25); +x_28 = l_Lean_Syntax_isOfKind(x_25, x_27); +lean_dec(x_27); +if (x_28 == 0) { -lean_object* x_30; -lean_dec(x_26); +lean_object* x_29; +lean_dec(x_25); lean_dec(x_18); lean_dec(x_14); lean_dec(x_12); @@ -1687,30 +1706,30 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_30 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_8); -return x_30; +x_29 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_8); +return x_29; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_31 = l_Lean_Syntax_getArg(x_26, x_9); -lean_dec(x_26); -x_32 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_32, 0, x_31); -x_33 = lean_box(0); -x_34 = l_Lean_Elab_Command_elabMacro___lambda__5(x_1, x_3, x_14, x_18, x_12, x_5, x_10, x_2, x_33, x_32, x_6, x_7, x_8); -return x_34; +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_30 = l_Lean_Syntax_getArg(x_25, x_9); +lean_dec(x_25); +x_31 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_31, 0, x_30); +x_32 = lean_box(0); +x_33 = l_Lean_Elab_Command_elabMacro___lambda__5(x_1, x_3, x_14, x_18, x_12, x_5, x_10, x_2, x_32, x_31, x_6, x_7, x_8); +return x_33; } } } else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_dec(x_20); +x_34 = lean_box(0); x_35 = lean_box(0); -x_36 = lean_box(0); -x_37 = l_Lean_Elab_Command_elabMacro___lambda__5(x_1, x_3, x_14, x_18, x_12, x_5, x_10, x_2, x_36, x_35, x_6, x_7, x_8); -return x_37; +x_36 = l_Lean_Elab_Command_elabMacro___lambda__5(x_1, x_3, x_14, x_18, x_12, x_5, x_10, x_2, x_35, x_34, x_6, x_7, x_8); +return x_36; } } } @@ -1829,62 +1848,61 @@ x_9 = l_Lean_Syntax_getArg(x_1, x_8); x_10 = l_Lean_Syntax_isNone(x_9); if (x_10 == 0) { -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = l_Lean_nullKind; -x_12 = lean_unsigned_to_nat(1u); +lean_object* x_11; uint8_t x_12; +x_11 = lean_unsigned_to_nat(1u); lean_inc(x_9); -x_13 = l_Lean_Syntax_isNodeOf(x_9, x_11, x_12); -if (x_13 == 0) +x_12 = l_Lean_Syntax_matchesNull(x_9, x_11); +if (x_12 == 0) { -lean_object* x_14; +lean_object* x_13; lean_dec(x_9); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_14 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_4); -return x_14; +x_13 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_4); +return x_13; } else { -lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_15 = l_Lean_Syntax_getArg(x_9, x_8); +lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_14 = l_Lean_Syntax_getArg(x_9, x_8); lean_dec(x_9); -x_16 = l_Lean_Elab_Command_elabMacro___closed__10; -lean_inc(x_15); -x_17 = l_Lean_Syntax_isOfKind(x_15, x_16); -if (x_17 == 0) +x_15 = l_Lean_Elab_Command_elabMacro___closed__10; +lean_inc(x_14); +x_16 = l_Lean_Syntax_isOfKind(x_14, x_15); +if (x_16 == 0) { -lean_object* x_18; -lean_dec(x_15); +lean_object* x_17; +lean_dec(x_14); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_18 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_4); -return x_18; +x_17 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_4); +return x_17; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_19 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_19, 0, x_15); -x_20 = l_Lean_Elab_Command_elabMacro___closed__4; -x_21 = l_Lean_Elab_Command_elabMacro___closed__6; -x_22 = lean_box(0); -x_23 = l_Lean_Elab_Command_elabMacro___lambda__6(x_1, x_20, x_21, x_22, x_19, x_2, x_3, x_4); -return x_23; +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_18 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_18, 0, x_14); +x_19 = l_Lean_Elab_Command_elabMacro___closed__4; +x_20 = l_Lean_Elab_Command_elabMacro___closed__6; +x_21 = lean_box(0); +x_22 = l_Lean_Elab_Command_elabMacro___lambda__6(x_1, x_19, x_20, x_21, x_18, x_2, x_3, x_4); +return x_22; } } } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_dec(x_9); -x_24 = lean_box(0); -x_25 = l_Lean_Elab_Command_elabMacro___closed__4; -x_26 = l_Lean_Elab_Command_elabMacro___closed__6; -x_27 = lean_box(0); -x_28 = l_Lean_Elab_Command_elabMacro___lambda__6(x_1, x_25, x_26, x_27, x_24, x_2, x_3, x_4); -return x_28; +x_23 = lean_box(0); +x_24 = l_Lean_Elab_Command_elabMacro___closed__4; +x_25 = l_Lean_Elab_Command_elabMacro___closed__6; +x_26 = lean_box(0); +x_27 = l_Lean_Elab_Command_elabMacro___lambda__6(x_1, x_24, x_25, x_26, x_23, x_2, x_3, x_4); +return x_27; } } } @@ -1911,6 +1929,18 @@ lean_dec(x_2); return x_5; } } +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacro___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacro___spec__3(x_4, x_5, x_3); +return x_6; +} +} LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabMacro___lambda__2___boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; @@ -2029,7 +2059,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabMacro_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(37u); +x_1 = lean_unsigned_to_nat(38u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2130,6 +2160,10 @@ if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); l_Lean_Elab_Command_elabMacro___lambda__1___closed__1 = _init_l_Lean_Elab_Command_elabMacro___lambda__1___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_elabMacro___lambda__1___closed__1); +l_Lean_Elab_Command_elabMacro___lambda__1___closed__2 = _init_l_Lean_Elab_Command_elabMacro___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_elabMacro___lambda__1___closed__2); +l_Lean_Elab_Command_elabMacro___lambda__1___closed__3 = _init_l_Lean_Elab_Command_elabMacro___lambda__1___closed__3(); +lean_mark_persistent(l_Lean_Elab_Command_elabMacro___lambda__1___closed__3); l_Lean_Elab_Command_elabMacro___lambda__2___closed__1 = _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_elabMacro___lambda__2___closed__1); l_Lean_Elab_Command_elabMacro___lambda__2___closed__2 = _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__2(); @@ -2186,10 +2220,6 @@ l_Lean_Elab_Command_elabMacro___lambda__2___closed__27 = _init_l_Lean_Elab_Comma lean_mark_persistent(l_Lean_Elab_Command_elabMacro___lambda__2___closed__27); l_Lean_Elab_Command_elabMacro___lambda__2___closed__28 = _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__28(); lean_mark_persistent(l_Lean_Elab_Command_elabMacro___lambda__2___closed__28); -l_Lean_Elab_Command_elabMacro___lambda__2___closed__29 = _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__29(); -lean_mark_persistent(l_Lean_Elab_Command_elabMacro___lambda__2___closed__29); -l_Lean_Elab_Command_elabMacro___lambda__2___closed__30 = _init_l_Lean_Elab_Command_elabMacro___lambda__2___closed__30(); -lean_mark_persistent(l_Lean_Elab_Command_elabMacro___lambda__2___closed__30); l_Lean_Elab_Command_elabMacro___lambda__3___closed__1 = _init_l_Lean_Elab_Command_elabMacro___lambda__3___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_elabMacro___lambda__3___closed__1); l_Lean_Elab_Command_elabMacro___lambda__6___closed__1 = _init_l_Lean_Elab_Command_elabMacro___lambda__6___closed__1(); diff --git a/stage0/stdlib/Lean/Elab/MacroArgUtil.c b/stage0/stdlib/Lean/Elab/MacroArgUtil.c index 766642bae1b..1dac76f84ea 100644 --- a/stage0/stdlib/Lean/Elab/MacroArgUtil.c +++ b/stage0/stdlib/Lean/Elab/MacroArgUtil.c @@ -17,7 +17,6 @@ LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at_Lean_Elab_Command_ex LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandMacroArg_mkSplicePat___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_resolveGlobalConstWithInfos___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_erase_macro_scopes(lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); @@ -26,10 +25,9 @@ static lean_object* l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__ lean_object* lean_mk_empty_array_with_capacity(lean_object*); static lean_object* l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__9(lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_nullKind; +lean_object* l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___closed__6; static lean_object* l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__6; -static lean_object* l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__8; lean_object* lean_name_mk_string(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; lean_object* l_List_mapTRAux___at_Lean_resolveGlobalConstCore___spec__2(lean_object*, lean_object*); @@ -47,6 +45,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoTree___at_Lean_Elab_Command_expandM static lean_object* l_Lean_Elab_Command_expandMacroArg___closed__9; lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__23; +static lean_object* l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__2; static lean_object* l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__21; LEAN_EXPORT lean_object* l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -56,11 +55,11 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lam LEAN_EXPORT lean_object* l_Lean_Elab_resolveGlobalConstWithInfos___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); -extern lean_object* l_Lean_interpolatedStrKind; lean_object* l_Lean_MessageData_ofList(lean_object*); static lean_object* l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__11; static lean_object* l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__9; -LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__3; +LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandMacroArg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_string_utf8_byte_size(lean_object*); static lean_object* l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__29; @@ -84,12 +83,13 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg static lean_object* l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__24; static lean_object* l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__5; static lean_object* l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__16; lean_object* lean_st_ref_take(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__5(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__7___closed__4; -extern lean_object* l_Lean_strLitKind; lean_object* l_Lean_Syntax_isStrLit_x3f(lean_object*); static lean_object* l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__7; static lean_object* l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__18; @@ -100,7 +100,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Command_expandM static lean_object* l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__30; static lean_object* l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__15; lean_object* l_Lean_replaceRef(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__9; +lean_object* l_panic___at_Lean_TSyntax_getString___spec__1(lean_object*); static lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__11___closed__1; static lean_object* l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__2; lean_object* l_Lean_Elab_Command_strLitToPattern___boxed(lean_object*, lean_object*, lean_object*); @@ -135,14 +135,12 @@ LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoTree___at_Lean_Elab_Command_expandM lean_object* l_Lean_ConstantInfo_type(lean_object*); lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_expandMacroArg___spec__1___boxed(lean_object*, lean_object*); -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); static lean_object* l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__5; static lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__11___closed__2; static lean_object* l_Lean_Elab_Command_expandMacroArg___closed__7; -lean_object* l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); -extern lean_object* l_Lean_identKind; lean_object* l_Std_PersistentArray_push___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__11(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__10; static lean_object* l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__3; static lean_object* l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__6; lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -193,7 +191,6 @@ LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at_Lean_Elab_Command_ex LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_List_isEmpty___rarg(lean_object*); static lean_object* l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___closed__1; -lean_object* l_panic___at_Lean_Name_getString_x21___spec__1(lean_object*); lean_object* l_Lean_Elab_Command_getScope___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; uint8_t l_Lean_Syntax_matchesIdent(lean_object*, lean_object*); @@ -213,6 +210,7 @@ static lean_object* l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__ static lean_object* l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__17; lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); uint8_t lean_string_dec_eq(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__11; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { @@ -398,7 +396,7 @@ x_7 = lean_ctor_get(x_5, 1); lean_inc(x_7); lean_dec(x_5); x_8 = lean_box(0); -x_9 = l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(x_6, x_8); +x_9 = l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(x_6, x_8); x_10 = l_List_isEmpty___rarg(x_9); if (x_10 == 0) { @@ -2072,273 +2070,287 @@ static lean_object* _init_l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___la _start: { lean_object* x_1; +x_1 = lean_mk_string_from_bytes("str", 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__9; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__11() { +_start: +{ +lean_object* x_1; x_1 = lean_mk_string_from_bytes("Term", 4); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___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_EXPORT lean_object* l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___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) { _start: { -lean_object* x_8; lean_object* x_9; +lean_object* x_9; lean_object* x_10; +lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); +lean_inc(x_7); lean_inc(x_6); -lean_inc(x_5); lean_inc(x_2); -x_8 = l_Lean_Elab_Term_resolveParserName___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__1(x_2, x_5, x_6, x_7); -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -if (lean_obj_tag(x_9) == 0) -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_10 = lean_ctor_get(x_8, 1); +x_9 = l_Lean_Elab_Term_resolveParserName___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__1(x_2, x_6, x_7, x_8); +x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); -lean_dec(x_8); -x_11 = l_Lean_Syntax_getId(x_2); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +lean_dec(x_3); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = l_Lean_Syntax_getId(x_2); lean_dec(x_2); -x_12 = lean_erase_macro_scopes(x_11); -x_13 = lean_st_ref_get(x_6, x_10); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) +x_13 = lean_erase_macro_scopes(x_12); +x_14 = lean_st_ref_get(x_7, x_11); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_ctor_get(x_13, 1); -x_17 = lean_ctor_get(x_15, 0); -lean_inc(x_17); -lean_dec(x_15); -lean_inc(x_12); -x_18 = l_Lean_Parser_isParserCategory(x_17, x_12); -lean_dec(x_17); -if (x_18 == 0) +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_16 = lean_ctor_get(x_14, 0); +x_17 = lean_ctor_get(x_14, 1); +x_18 = lean_ctor_get(x_16, 0); +lean_inc(x_18); +lean_dec(x_16); +lean_inc(x_13); +x_19 = l_Lean_Parser_isParserCategory(x_18, x_13); +lean_dec(x_18); +if (x_19 == 0) { -lean_object* x_19; lean_object* x_20; uint8_t x_21; -lean_free_object(x_13); -x_19 = l_Lean_Parser_isParserAlias(x_12, x_16); -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -x_21 = lean_unbox(x_20); -lean_dec(x_20); -if (x_21 == 0) +lean_object* x_20; lean_object* x_21; uint8_t x_22; +lean_free_object(x_14); +x_20 = l_Lean_Parser_isParserAlias(x_13, x_17); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_unbox(x_21); +lean_dec(x_21); +if (x_22 == 0) { -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; uint8_t x_29; +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; uint8_t x_30; lean_dec(x_1); -x_22 = lean_ctor_get(x_19, 1); -lean_inc(x_22); -lean_dec(x_19); -x_23 = lean_alloc_ctor(4, 1, 0); -lean_ctor_set(x_23, 0, x_12); -x_24 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__2; -x_25 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_23); -x_26 = l_Lean_throwUnknownConstant___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__7___closed__4; -x_27 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_27, 0, x_25); -lean_ctor_set(x_27, 1, x_26); -x_28 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_27, x_5, x_6, x_22); -lean_dec(x_6); -x_29 = !lean_is_exclusive(x_28); -if (x_29 == 0) +x_23 = lean_ctor_get(x_20, 1); +lean_inc(x_23); +lean_dec(x_20); +x_24 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_24, 0, x_13); +x_25 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__2; +x_26 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_24); +x_27 = l_Lean_throwUnknownConstant___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__7___closed__4; +x_28 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +x_29 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_28, x_6, x_7, x_23); +lean_dec(x_7); +x_30 = !lean_is_exclusive(x_29); +if (x_30 == 0) { -return x_28; +return x_29; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_28, 0); -x_31 = lean_ctor_get(x_28, 1); +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_29, 0); +x_32 = lean_ctor_get(x_29, 1); +lean_inc(x_32); lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_28); -x_32 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -return x_32; +lean_dec(x_29); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; } } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_19, 1); -lean_inc(x_33); -lean_dec(x_19); -x_34 = l_Lean_Parser_getSyntaxKindOfParserAlias_x3f(x_12, x_33); -lean_dec(x_12); -x_35 = lean_ctor_get(x_34, 0); -lean_inc(x_35); -if (lean_obj_tag(x_35) == 0) -{ -lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_36 = lean_ctor_get(x_34, 1); +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_20, 1); +lean_inc(x_34); +lean_dec(x_20); +x_35 = l_Lean_Parser_getSyntaxKindOfParserAlias_x3f(x_13, x_34); +lean_dec(x_13); +x_36 = lean_ctor_get(x_35, 0); lean_inc(x_36); -lean_dec(x_34); -x_37 = lean_box(0); -x_38 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_37, x_5, x_6, x_36); +if (lean_obj_tag(x_36) == 0) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_box(0); +x_39 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_38, x_6, x_7, x_37); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -return x_38; +return x_39; } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_39 = lean_ctor_get(x_34, 1); -lean_inc(x_39); -lean_dec(x_34); -x_40 = lean_ctor_get(x_35, 0); +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_35, 1); lean_inc(x_40); lean_dec(x_35); -x_41 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_40, x_5, x_6, x_39); +x_41 = lean_ctor_get(x_36, 0); +lean_inc(x_41); +lean_dec(x_36); +x_42 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_41, x_6, x_7, x_40); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_40); -return x_41; +lean_dec(x_41); +return x_42; } } } else { -lean_object* x_42; lean_object* x_43; uint8_t x_44; lean_object* x_45; +lean_object* x_43; lean_object* x_44; uint8_t x_45; lean_object* x_46; +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -x_42 = lean_box(0); -x_43 = lean_unsigned_to_nat(0u); -x_44 = 1; -x_45 = l_Lean_Syntax_mkAntiquotNode(x_12, x_1, x_43, x_42, x_44); -lean_dec(x_12); -lean_ctor_set(x_13, 0, x_45); -return x_13; +x_43 = lean_box(0); +x_44 = lean_unsigned_to_nat(0u); +x_45 = 1; +x_46 = l_Lean_Syntax_mkAntiquotNode(x_13, x_1, x_44, x_43, x_45); +lean_dec(x_13); +lean_ctor_set(x_14, 0, x_46); +return x_14; } } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; -x_46 = lean_ctor_get(x_13, 0); -x_47 = lean_ctor_get(x_13, 1); -lean_inc(x_47); -lean_inc(x_46); -lean_dec(x_13); -x_48 = lean_ctor_get(x_46, 0); +lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; +x_47 = lean_ctor_get(x_14, 0); +x_48 = lean_ctor_get(x_14, 1); lean_inc(x_48); -lean_dec(x_46); -lean_inc(x_12); -x_49 = l_Lean_Parser_isParserCategory(x_48, x_12); -lean_dec(x_48); -if (x_49 == 0) +lean_inc(x_47); +lean_dec(x_14); +x_49 = lean_ctor_get(x_47, 0); +lean_inc(x_49); +lean_dec(x_47); +lean_inc(x_13); +x_50 = l_Lean_Parser_isParserCategory(x_49, x_13); +lean_dec(x_49); +if (x_50 == 0) { -lean_object* x_50; lean_object* x_51; uint8_t x_52; -x_50 = l_Lean_Parser_isParserAlias(x_12, x_47); -x_51 = lean_ctor_get(x_50, 0); -lean_inc(x_51); -x_52 = lean_unbox(x_51); -lean_dec(x_51); -if (x_52 == 0) +lean_object* x_51; lean_object* x_52; uint8_t x_53; +x_51 = l_Lean_Parser_isParserAlias(x_13, x_48); +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_unbox(x_52); +lean_dec(x_52); +if (x_53 == 0) { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_dec(x_1); -x_53 = lean_ctor_get(x_50, 1); -lean_inc(x_53); -lean_dec(x_50); -x_54 = lean_alloc_ctor(4, 1, 0); -lean_ctor_set(x_54, 0, x_12); -x_55 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__2; -x_56 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_56, 0, x_55); -lean_ctor_set(x_56, 1, x_54); -x_57 = l_Lean_throwUnknownConstant___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__7___closed__4; -x_58 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_58, 0, x_56); -lean_ctor_set(x_58, 1, x_57); -x_59 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_58, x_5, x_6, x_53); -lean_dec(x_6); -x_60 = lean_ctor_get(x_59, 0); -lean_inc(x_60); -x_61 = lean_ctor_get(x_59, 1); +x_54 = lean_ctor_get(x_51, 1); +lean_inc(x_54); +lean_dec(x_51); +x_55 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_55, 0, x_13); +x_56 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__2; +x_57 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_55); +x_58 = l_Lean_throwUnknownConstant___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__7___closed__4; +x_59 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_58); +x_60 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_59, x_6, x_7, x_54); +lean_dec(x_7); +x_61 = lean_ctor_get(x_60, 0); lean_inc(x_61); -if (lean_is_exclusive(x_59)) { - lean_ctor_release(x_59, 0); - lean_ctor_release(x_59, 1); - x_62 = x_59; +x_62 = lean_ctor_get(x_60, 1); +lean_inc(x_62); +if (lean_is_exclusive(x_60)) { + lean_ctor_release(x_60, 0); + lean_ctor_release(x_60, 1); + x_63 = x_60; } else { - lean_dec_ref(x_59); - x_62 = lean_box(0); + lean_dec_ref(x_60); + x_63 = lean_box(0); } -if (lean_is_scalar(x_62)) { - x_63 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_63)) { + x_64 = lean_alloc_ctor(1, 2, 0); } else { - x_63 = x_62; + x_64 = x_63; } -lean_ctor_set(x_63, 0, x_60); -lean_ctor_set(x_63, 1, x_61); -return x_63; +lean_ctor_set(x_64, 0, x_61); +lean_ctor_set(x_64, 1, x_62); +return x_64; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_ctor_get(x_50, 1); -lean_inc(x_64); -lean_dec(x_50); -x_65 = l_Lean_Parser_getSyntaxKindOfParserAlias_x3f(x_12, x_64); -lean_dec(x_12); -x_66 = lean_ctor_get(x_65, 0); -lean_inc(x_66); -if (lean_obj_tag(x_66) == 0) -{ -lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_67 = lean_ctor_get(x_65, 1); +lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = lean_ctor_get(x_51, 1); +lean_inc(x_65); +lean_dec(x_51); +x_66 = l_Lean_Parser_getSyntaxKindOfParserAlias_x3f(x_13, x_65); +lean_dec(x_13); +x_67 = lean_ctor_get(x_66, 0); lean_inc(x_67); -lean_dec(x_65); -x_68 = lean_box(0); -x_69 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_68, x_5, x_6, x_67); +if (lean_obj_tag(x_67) == 0) +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +lean_dec(x_66); +x_69 = lean_box(0); +x_70 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_69, x_6, x_7, x_68); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -return x_69; +return x_70; } else { -lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_70 = lean_ctor_get(x_65, 1); -lean_inc(x_70); -lean_dec(x_65); -x_71 = lean_ctor_get(x_66, 0); +lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_71 = lean_ctor_get(x_66, 1); lean_inc(x_71); lean_dec(x_66); -x_72 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_71, x_5, x_6, x_70); +x_72 = lean_ctor_get(x_67, 0); +lean_inc(x_72); +lean_dec(x_67); +x_73 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_72, x_6, x_7, x_71); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_71); -return x_72; +lean_dec(x_72); +return x_73; } } } else { -lean_object* x_73; lean_object* x_74; uint8_t x_75; lean_object* x_76; lean_object* x_77; +lean_object* x_74; lean_object* x_75; uint8_t x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -x_73 = lean_box(0); -x_74 = lean_unsigned_to_nat(0u); -x_75 = 1; -x_76 = l_Lean_Syntax_mkAntiquotNode(x_12, x_1, x_74, x_73, x_75); -lean_dec(x_12); -x_77 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_77, 0, x_76); -lean_ctor_set(x_77, 1, x_47); -return x_77; +x_74 = lean_box(0); +x_75 = lean_unsigned_to_nat(0u); +x_76 = 1; +x_77 = l_Lean_Syntax_mkAntiquotNode(x_13, x_1, x_75, x_74, x_76); +lean_dec(x_13); +x_78 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_78, 0, x_77); +lean_ctor_set(x_78, 1, x_48); +return x_78; } } } else { -lean_object* x_78; lean_object* x_79; +lean_object* x_79; lean_object* x_80; lean_dec(x_2); -x_78 = lean_ctor_get(x_9, 0); -lean_inc(x_78); -x_79 = lean_ctor_get(x_78, 0); +x_79 = lean_ctor_get(x_10, 0); lean_inc(x_79); -if (lean_obj_tag(x_79) == 1) -{ -lean_object* x_80; x_80 = lean_ctor_get(x_79, 0); lean_inc(x_80); if (lean_obj_tag(x_80) == 1) @@ -2351,333 +2363,342 @@ if (lean_obj_tag(x_81) == 1) lean_object* x_82; x_82 = lean_ctor_get(x_81, 0); lean_inc(x_82); -switch (lean_obj_tag(x_82)) { -case 0: +if (lean_obj_tag(x_82) == 1) { -lean_object* x_83; lean_object* x_84; uint8_t x_85; -x_83 = lean_ctor_get(x_8, 1); +lean_object* x_83; +x_83 = lean_ctor_get(x_82, 0); lean_inc(x_83); -lean_dec(x_8); +switch (lean_obj_tag(x_83)) { +case 0: +{ +lean_object* x_84; lean_object* x_85; uint8_t x_86; x_84 = lean_ctor_get(x_9, 1); lean_inc(x_84); -x_85 = !lean_is_exclusive(x_78); -if (x_85 == 0) -{ -lean_object* x_86; lean_object* x_87; lean_object* x_88; uint8_t x_89; -x_86 = lean_ctor_get(x_78, 1); -x_87 = lean_ctor_get(x_78, 0); -lean_dec(x_87); -x_88 = lean_ctor_get(x_79, 1); -lean_inc(x_88); -x_89 = !lean_is_exclusive(x_80); -if (x_89 == 0) -{ -uint64_t x_90; lean_object* x_91; lean_object* x_92; uint8_t x_93; -x_90 = lean_ctor_get_uint64(x_79, sizeof(void*)*2); -x_91 = lean_ctor_get(x_80, 1); -x_92 = lean_ctor_get(x_80, 0); -lean_dec(x_92); -x_93 = !lean_is_exclusive(x_81); -if (x_93 == 0) -{ -lean_object* x_94; lean_object* x_95; lean_object* x_96; uint8_t x_97; -x_94 = lean_ctor_get(x_81, 1); -x_95 = lean_ctor_get(x_81, 0); +lean_dec(x_9); +x_85 = lean_ctor_get(x_10, 1); +lean_inc(x_85); +x_86 = !lean_is_exclusive(x_79); +if (x_86 == 0) +{ +lean_object* x_87; lean_object* x_88; lean_object* x_89; uint8_t x_90; +x_87 = lean_ctor_get(x_79, 1); +x_88 = lean_ctor_get(x_79, 0); +lean_dec(x_88); +x_89 = lean_ctor_get(x_80, 1); +lean_inc(x_89); +x_90 = !lean_is_exclusive(x_81); +if (x_90 == 0) +{ +uint64_t x_91; lean_object* x_92; lean_object* x_93; uint8_t x_94; +x_91 = lean_ctor_get_uint64(x_80, sizeof(void*)*2); +x_92 = lean_ctor_get(x_81, 1); +x_93 = lean_ctor_get(x_81, 0); +lean_dec(x_93); +x_94 = !lean_is_exclusive(x_82); +if (x_94 == 0) +{ +lean_object* x_95; lean_object* x_96; lean_object* x_97; uint8_t x_98; +x_95 = lean_ctor_get(x_82, 1); +x_96 = lean_ctor_get(x_82, 0); +lean_dec(x_96); +x_97 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__1; +x_98 = lean_string_dec_eq(x_95, x_97); lean_dec(x_95); -x_96 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__1; -x_97 = lean_string_dec_eq(x_94, x_96); -lean_dec(x_94); -if (x_97 == 0) +if (x_98 == 0) { +lean_free_object(x_82); lean_free_object(x_81); -lean_free_object(x_80); -lean_dec(x_91); -lean_dec(x_88); -lean_free_object(x_78); -lean_dec(x_86); -if (lean_obj_tag(x_84) == 0) +lean_dec(x_92); +lean_dec(x_89); +lean_free_object(x_79); +lean_dec(x_87); +lean_dec(x_3); +if (lean_obj_tag(x_85) == 0) { -lean_object* x_98; -lean_dec(x_9); -x_98 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_79, x_5, x_6, x_83); +lean_object* x_99; +lean_dec(x_10); +x_99 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_80, x_6, x_7, x_84); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_79); -return x_98; +lean_dec(x_80); +return x_99; } else { -lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; uint8_t x_108; -lean_dec(x_84); -lean_dec(x_79); +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; uint8_t x_109; +lean_dec(x_85); +lean_dec(x_80); lean_dec(x_1); -x_99 = lean_box(0); -x_100 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_9, x_99); -x_101 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_100, x_99); -x_102 = l_Lean_MessageData_ofList(x_101); -lean_dec(x_101); -x_103 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_104 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_104, 0, x_103); -lean_ctor_set(x_104, 1, x_102); -x_105 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_106 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_106, 0, x_104); -lean_ctor_set(x_106, 1, x_105); -x_107 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_106, x_5, x_6, x_83); -lean_dec(x_6); -x_108 = !lean_is_exclusive(x_107); -if (x_108 == 0) +x_100 = lean_box(0); +x_101 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_10, x_100); +x_102 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_101, x_100); +x_103 = l_Lean_MessageData_ofList(x_102); +lean_dec(x_102); +x_104 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_105 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_105, 0, x_104); +lean_ctor_set(x_105, 1, x_103); +x_106 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_107 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_107, 0, x_105); +lean_ctor_set(x_107, 1, x_106); +x_108 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_107, x_6, x_7, x_84); +lean_dec(x_7); +x_109 = !lean_is_exclusive(x_108); +if (x_109 == 0) { -return x_107; +return x_108; } else { -lean_object* x_109; lean_object* x_110; lean_object* x_111; -x_109 = lean_ctor_get(x_107, 0); -x_110 = lean_ctor_get(x_107, 1); +lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_110 = lean_ctor_get(x_108, 0); +x_111 = lean_ctor_get(x_108, 1); +lean_inc(x_111); lean_inc(x_110); -lean_inc(x_109); -lean_dec(x_107); -x_111 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_111, 0, x_109); -lean_ctor_set(x_111, 1, x_110); -return x_111; +lean_dec(x_108); +x_112 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_112, 0, x_110); +lean_ctor_set(x_112, 1, x_111); +return x_112; } } } else { -uint8_t x_112; -x_112 = !lean_is_exclusive(x_9); -if (x_112 == 0) +uint8_t x_113; +x_113 = !lean_is_exclusive(x_10); +if (x_113 == 0) { -lean_object* x_113; lean_object* x_114; uint8_t x_115; -x_113 = lean_ctor_get(x_9, 1); -lean_dec(x_113); -x_114 = lean_ctor_get(x_9, 0); +lean_object* x_114; lean_object* x_115; uint8_t x_116; +x_114 = lean_ctor_get(x_10, 1); lean_dec(x_114); -x_115 = !lean_is_exclusive(x_79); -if (x_115 == 0) +x_115 = lean_ctor_get(x_10, 0); +lean_dec(x_115); +x_116 = !lean_is_exclusive(x_80); +if (x_116 == 0) { -lean_object* x_116; lean_object* x_117; lean_object* x_118; uint8_t x_119; -x_116 = lean_ctor_get(x_79, 1); -lean_dec(x_116); -x_117 = lean_ctor_get(x_79, 0); +lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; +x_117 = lean_ctor_get(x_80, 1); lean_dec(x_117); -x_118 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__4; -x_119 = lean_string_dec_eq(x_91, x_118); -if (x_119 == 0) -{ -if (lean_obj_tag(x_84) == 0) -{ -lean_object* x_120; -lean_free_object(x_9); -lean_free_object(x_78); -lean_dec(x_86); -lean_ctor_set(x_81, 1, x_96); -x_120 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_79, x_5, x_6, x_83); +x_118 = lean_ctor_get(x_80, 0); +lean_dec(x_118); +x_119 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__4; +x_120 = lean_string_dec_eq(x_92, x_119); +if (x_120 == 0) +{ +lean_dec(x_3); +if (lean_obj_tag(x_85) == 0) +{ +lean_object* x_121; +lean_free_object(x_10); +lean_free_object(x_79); +lean_dec(x_87); +lean_ctor_set(x_82, 1, x_97); +x_121 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_80, x_6, x_7, x_84); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_79); -return x_120; +lean_dec(x_80); +return x_121; } else { -lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; uint8_t x_130; +lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; uint8_t x_131; lean_dec(x_1); -lean_ctor_set(x_81, 1, x_96); -x_121 = lean_box(0); -x_122 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_9, x_121); -x_123 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_122, x_121); -x_124 = l_Lean_MessageData_ofList(x_123); -lean_dec(x_123); -x_125 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_126 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_126, 0, x_125); -lean_ctor_set(x_126, 1, x_124); -x_127 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_128 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_128, 0, x_126); -lean_ctor_set(x_128, 1, x_127); -x_129 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_128, x_5, x_6, x_83); -lean_dec(x_6); -x_130 = !lean_is_exclusive(x_129); -if (x_130 == 0) +lean_ctor_set(x_82, 1, x_97); +x_122 = lean_box(0); +x_123 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_10, x_122); +x_124 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_123, x_122); +x_125 = l_Lean_MessageData_ofList(x_124); +lean_dec(x_124); +x_126 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_127 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_127, 0, x_126); +lean_ctor_set(x_127, 1, x_125); +x_128 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_129 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_129, 0, x_127); +lean_ctor_set(x_129, 1, x_128); +x_130 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_129, x_6, x_7, x_84); +lean_dec(x_7); +x_131 = !lean_is_exclusive(x_130); +if (x_131 == 0) { -return x_129; +return x_130; } else { -lean_object* x_131; lean_object* x_132; lean_object* x_133; -x_131 = lean_ctor_get(x_129, 0); -x_132 = lean_ctor_get(x_129, 1); +lean_object* x_132; lean_object* x_133; lean_object* x_134; +x_132 = lean_ctor_get(x_130, 0); +x_133 = lean_ctor_get(x_130, 1); +lean_inc(x_133); lean_inc(x_132); -lean_inc(x_131); -lean_dec(x_129); -x_133 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_133, 0, x_131); -lean_ctor_set(x_133, 1, x_132); -return x_133; +lean_dec(x_130); +x_134 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_134, 0, x_132); +lean_ctor_set(x_134, 1, x_133); +return x_134; } } } else { -lean_object* x_134; uint8_t x_135; -lean_dec(x_91); -x_134 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__7; -x_135 = lean_string_dec_eq(x_88, x_134); -if (x_135 == 0) -{ -lean_object* x_136; uint8_t x_137; -x_136 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__8; -x_137 = lean_string_dec_eq(x_88, x_136); -if (x_137 == 0) -{ -if (lean_obj_tag(x_84) == 0) -{ -lean_object* x_138; -lean_free_object(x_9); -lean_free_object(x_78); -lean_dec(x_86); -lean_ctor_set(x_81, 1, x_96); -lean_ctor_set(x_80, 1, x_118); -x_138 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_79, x_5, x_6, x_83); +lean_object* x_135; uint8_t x_136; +lean_dec(x_92); +x_135 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__7; +x_136 = lean_string_dec_eq(x_89, x_135); +if (x_136 == 0) +{ +lean_object* x_137; uint8_t x_138; +lean_dec(x_3); +x_137 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__8; +x_138 = lean_string_dec_eq(x_89, x_137); +if (x_138 == 0) +{ +if (lean_obj_tag(x_85) == 0) +{ +lean_object* x_139; +lean_free_object(x_10); +lean_free_object(x_79); +lean_dec(x_87); +lean_ctor_set(x_82, 1, x_97); +lean_ctor_set(x_81, 1, x_119); +x_139 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_80, x_6, x_7, x_84); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_79); -return x_138; +lean_dec(x_80); +return x_139; } else { -lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; uint8_t x_148; +lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; uint8_t x_149; lean_dec(x_1); -lean_ctor_set(x_81, 1, x_96); -lean_ctor_set(x_80, 1, x_118); -x_139 = lean_box(0); -x_140 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_9, x_139); -x_141 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_140, x_139); -x_142 = l_Lean_MessageData_ofList(x_141); -lean_dec(x_141); -x_143 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_144 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_144, 0, x_143); -lean_ctor_set(x_144, 1, x_142); -x_145 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_146 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_146, 0, x_144); -lean_ctor_set(x_146, 1, x_145); -x_147 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_146, x_5, x_6, x_83); -lean_dec(x_6); -x_148 = !lean_is_exclusive(x_147); -if (x_148 == 0) +lean_ctor_set(x_82, 1, x_97); +lean_ctor_set(x_81, 1, x_119); +x_140 = lean_box(0); +x_141 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_10, x_140); +x_142 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_141, x_140); +x_143 = l_Lean_MessageData_ofList(x_142); +lean_dec(x_142); +x_144 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_145 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_145, 0, x_144); +lean_ctor_set(x_145, 1, x_143); +x_146 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_147 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_147, 0, x_145); +lean_ctor_set(x_147, 1, x_146); +x_148 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_147, x_6, x_7, x_84); +lean_dec(x_7); +x_149 = !lean_is_exclusive(x_148); +if (x_149 == 0) { -return x_147; +return x_148; } else { -lean_object* x_149; lean_object* x_150; lean_object* x_151; -x_149 = lean_ctor_get(x_147, 0); -x_150 = lean_ctor_get(x_147, 1); +lean_object* x_150; lean_object* x_151; lean_object* x_152; +x_150 = lean_ctor_get(x_148, 0); +x_151 = lean_ctor_get(x_148, 1); +lean_inc(x_151); lean_inc(x_150); -lean_inc(x_149); -lean_dec(x_147); -x_151 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_151, 0, x_149); -lean_ctor_set(x_151, 1, x_150); -return x_151; +lean_dec(x_148); +x_152 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_152, 0, x_150); +lean_ctor_set(x_152, 1, x_151); +return x_152; } } } else { -lean_dec(x_88); -if (lean_obj_tag(x_84) == 0) +lean_dec(x_89); +if (lean_obj_tag(x_85) == 0) { -lean_object* x_152; lean_object* x_153; -lean_free_object(x_79); -lean_free_object(x_9); -lean_free_object(x_81); +lean_object* x_153; lean_object* x_154; lean_free_object(x_80); -lean_free_object(x_78); -lean_dec(x_86); -x_152 = l_Lean_strLitKind; -x_153 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_152, x_5, x_6, x_83); +lean_free_object(x_10); +lean_free_object(x_82); +lean_free_object(x_81); +lean_free_object(x_79); +lean_dec(x_87); +x_153 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__10; +x_154 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_153, x_6, x_7, x_84); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -return x_153; +return x_154; } else { -lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; uint8_t x_163; +lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; uint8_t x_164; lean_dec(x_1); -lean_ctor_set(x_81, 1, x_96); -lean_ctor_set(x_80, 1, x_118); -lean_ctor_set(x_79, 1, x_136); -x_154 = lean_box(0); -x_155 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_9, x_154); -x_156 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_155, x_154); -x_157 = l_Lean_MessageData_ofList(x_156); -lean_dec(x_156); -x_158 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_159 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_159, 0, x_158); -lean_ctor_set(x_159, 1, x_157); -x_160 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_161 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_161, 0, x_159); -lean_ctor_set(x_161, 1, x_160); -x_162 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_161, x_5, x_6, x_83); -lean_dec(x_6); -x_163 = !lean_is_exclusive(x_162); -if (x_163 == 0) +lean_ctor_set(x_82, 1, x_97); +lean_ctor_set(x_81, 1, x_119); +lean_ctor_set(x_80, 1, x_137); +x_155 = lean_box(0); +x_156 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_10, x_155); +x_157 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_156, x_155); +x_158 = l_Lean_MessageData_ofList(x_157); +lean_dec(x_157); +x_159 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_160 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_160, 0, x_159); +lean_ctor_set(x_160, 1, x_158); +x_161 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_162 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_162, 0, x_160); +lean_ctor_set(x_162, 1, x_161); +x_163 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_162, x_6, x_7, x_84); +lean_dec(x_7); +x_164 = !lean_is_exclusive(x_163); +if (x_164 == 0) { -return x_162; +return x_163; } else { -lean_object* x_164; lean_object* x_165; lean_object* x_166; -x_164 = lean_ctor_get(x_162, 0); -x_165 = lean_ctor_get(x_162, 1); +lean_object* x_165; lean_object* x_166; lean_object* x_167; +x_165 = lean_ctor_get(x_163, 0); +x_166 = lean_ctor_get(x_163, 1); +lean_inc(x_166); lean_inc(x_165); -lean_inc(x_164); -lean_dec(x_162); -x_166 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_166, 0, x_164); -lean_ctor_set(x_166, 1, x_165); -return x_166; +lean_dec(x_163); +x_167 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_167, 0, x_165); +lean_ctor_set(x_167, 1, x_166); +return x_167; } } } } else { -lean_dec(x_88); -if (lean_obj_tag(x_84) == 0) +lean_dec(x_89); +if (lean_obj_tag(x_85) == 0) { -lean_object* x_167; lean_object* x_168; -lean_free_object(x_79); -lean_free_object(x_9); -lean_free_object(x_81); +lean_object* x_168; lean_free_object(x_80); -lean_free_object(x_78); -lean_dec(x_86); -x_167 = l_Lean_identKind; -x_168 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_167, x_5, x_6, x_83); +lean_free_object(x_10); +lean_free_object(x_82); +lean_free_object(x_81); +lean_free_object(x_79); +lean_dec(x_87); +x_168 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_3, x_6, x_7, x_84); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); +lean_dec(x_3); return x_168; } else { lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; uint8_t x_178; +lean_dec(x_3); lean_dec(x_1); -lean_ctor_set(x_81, 1, x_96); -lean_ctor_set(x_80, 1, x_118); -lean_ctor_set(x_79, 1, x_134); +lean_ctor_set(x_82, 1, x_97); +lean_ctor_set(x_81, 1, x_119); +lean_ctor_set(x_80, 1, x_135); x_169 = lean_box(0); -x_170 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_9, x_169); +x_170 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_10, x_169); x_171 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_170, x_169); x_172 = l_Lean_MessageData_ofList(x_171); lean_dec(x_171); @@ -2689,8 +2710,8 @@ x_175 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__ x_176 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_176, 0, x_174); lean_ctor_set(x_176, 1, x_175); -x_177 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_176, x_5, x_6, x_83); -lean_dec(x_6); +x_177 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_176, x_6, x_7, x_84); +lean_dec(x_7); x_178 = !lean_is_exclusive(x_177); if (x_178 == 0) { @@ -2716,25 +2737,26 @@ return x_181; else { lean_object* x_182; uint8_t x_183; -lean_dec(x_79); +lean_dec(x_80); x_182 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__4; -x_183 = lean_string_dec_eq(x_91, x_182); +x_183 = lean_string_dec_eq(x_92, x_182); if (x_183 == 0) { -if (lean_obj_tag(x_84) == 0) +lean_dec(x_3); +if (lean_obj_tag(x_85) == 0) { lean_object* x_184; lean_object* x_185; -lean_free_object(x_9); -lean_free_object(x_78); -lean_dec(x_86); -lean_ctor_set(x_81, 1, x_96); +lean_free_object(x_10); +lean_free_object(x_79); +lean_dec(x_87); +lean_ctor_set(x_82, 1, x_97); x_184 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_184, 0, x_80); -lean_ctor_set(x_184, 1, x_88); -lean_ctor_set_uint64(x_184, sizeof(void*)*2, x_90); -x_185 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_184, x_5, x_6, x_83); +lean_ctor_set(x_184, 0, x_81); +lean_ctor_set(x_184, 1, x_89); +lean_ctor_set_uint64(x_184, sizeof(void*)*2, x_91); +x_185 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_184, x_6, x_7, x_84); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_184); return x_185; } @@ -2742,14 +2764,14 @@ else { lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_dec(x_1); -lean_ctor_set(x_81, 1, x_96); +lean_ctor_set(x_82, 1, x_97); x_186 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_186, 0, x_80); -lean_ctor_set(x_186, 1, x_88); -lean_ctor_set_uint64(x_186, sizeof(void*)*2, x_90); -lean_ctor_set(x_78, 0, x_186); +lean_ctor_set(x_186, 0, x_81); +lean_ctor_set(x_186, 1, x_89); +lean_ctor_set_uint64(x_186, sizeof(void*)*2, x_91); +lean_ctor_set(x_79, 0, x_186); x_187 = lean_box(0); -x_188 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_9, x_187); +x_188 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_10, x_187); x_189 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_188, x_187); x_190 = l_Lean_MessageData_ofList(x_189); lean_dec(x_189); @@ -2761,8 +2783,8 @@ x_193 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__ x_194 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_194, 0, x_192); lean_ctor_set(x_194, 1, x_193); -x_195 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_194, x_5, x_6, x_83); -lean_dec(x_6); +x_195 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_194, x_6, x_7, x_84); +lean_dec(x_7); x_196 = lean_ctor_get(x_195, 0); lean_inc(x_196); x_197 = lean_ctor_get(x_195, 1); @@ -2788,31 +2810,32 @@ return x_199; else { lean_object* x_200; uint8_t x_201; -lean_dec(x_91); +lean_dec(x_92); x_200 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__7; -x_201 = lean_string_dec_eq(x_88, x_200); +x_201 = lean_string_dec_eq(x_89, x_200); if (x_201 == 0) { lean_object* x_202; uint8_t x_203; +lean_dec(x_3); x_202 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__8; -x_203 = lean_string_dec_eq(x_88, x_202); +x_203 = lean_string_dec_eq(x_89, x_202); if (x_203 == 0) { -if (lean_obj_tag(x_84) == 0) +if (lean_obj_tag(x_85) == 0) { lean_object* x_204; lean_object* x_205; -lean_free_object(x_9); -lean_free_object(x_78); -lean_dec(x_86); -lean_ctor_set(x_81, 1, x_96); -lean_ctor_set(x_80, 1, x_182); +lean_free_object(x_10); +lean_free_object(x_79); +lean_dec(x_87); +lean_ctor_set(x_82, 1, x_97); +lean_ctor_set(x_81, 1, x_182); x_204 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_204, 0, x_80); -lean_ctor_set(x_204, 1, x_88); -lean_ctor_set_uint64(x_204, sizeof(void*)*2, x_90); -x_205 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_204, x_5, x_6, x_83); +lean_ctor_set(x_204, 0, x_81); +lean_ctor_set(x_204, 1, x_89); +lean_ctor_set_uint64(x_204, sizeof(void*)*2, x_91); +x_205 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_204, x_6, x_7, x_84); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_204); return x_205; } @@ -2820,15 +2843,15 @@ else { lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_dec(x_1); -lean_ctor_set(x_81, 1, x_96); -lean_ctor_set(x_80, 1, x_182); +lean_ctor_set(x_82, 1, x_97); +lean_ctor_set(x_81, 1, x_182); x_206 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_206, 0, x_80); -lean_ctor_set(x_206, 1, x_88); -lean_ctor_set_uint64(x_206, sizeof(void*)*2, x_90); -lean_ctor_set(x_78, 0, x_206); +lean_ctor_set(x_206, 0, x_81); +lean_ctor_set(x_206, 1, x_89); +lean_ctor_set_uint64(x_206, sizeof(void*)*2, x_91); +lean_ctor_set(x_79, 0, x_206); x_207 = lean_box(0); -x_208 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_9, x_207); +x_208 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_10, x_207); x_209 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_208, x_207); x_210 = l_Lean_MessageData_ofList(x_209); lean_dec(x_209); @@ -2840,8 +2863,8 @@ x_213 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__ x_214 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_214, 0, x_212); lean_ctor_set(x_214, 1, x_213); -x_215 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_214, x_5, x_6, x_83); -lean_dec(x_6); +x_215 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_214, x_6, x_7, x_84); +lean_dec(x_7); x_216 = lean_ctor_get(x_215, 0); lean_inc(x_216); x_217 = lean_ctor_get(x_215, 1); @@ -2866,34 +2889,34 @@ return x_219; } else { -lean_dec(x_88); -if (lean_obj_tag(x_84) == 0) +lean_dec(x_89); +if (lean_obj_tag(x_85) == 0) { lean_object* x_220; lean_object* x_221; -lean_free_object(x_9); +lean_free_object(x_10); +lean_free_object(x_82); lean_free_object(x_81); -lean_free_object(x_80); -lean_free_object(x_78); -lean_dec(x_86); -x_220 = l_Lean_strLitKind; -x_221 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_220, x_5, x_6, x_83); +lean_free_object(x_79); +lean_dec(x_87); +x_220 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__10; +x_221 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_220, x_6, x_7, x_84); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); return x_221; } else { lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_dec(x_1); -lean_ctor_set(x_81, 1, x_96); -lean_ctor_set(x_80, 1, x_182); +lean_ctor_set(x_82, 1, x_97); +lean_ctor_set(x_81, 1, x_182); x_222 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_222, 0, x_80); +lean_ctor_set(x_222, 0, x_81); lean_ctor_set(x_222, 1, x_202); -lean_ctor_set_uint64(x_222, sizeof(void*)*2, x_90); -lean_ctor_set(x_78, 0, x_222); +lean_ctor_set_uint64(x_222, sizeof(void*)*2, x_91); +lean_ctor_set(x_79, 0, x_222); x_223 = lean_box(0); -x_224 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_9, x_223); +x_224 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_10, x_223); x_225 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_224, x_223); x_226 = l_Lean_MessageData_ofList(x_225); lean_dec(x_225); @@ -2905,8 +2928,8 @@ x_229 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__ x_230 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_230, 0, x_228); lean_ctor_set(x_230, 1, x_229); -x_231 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_230, x_5, x_6, x_83); -lean_dec(x_6); +x_231 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_230, x_6, x_7, x_84); +lean_dec(x_7); x_232 = lean_ctor_get(x_231, 0); lean_inc(x_232); x_233 = lean_ctor_get(x_231, 1); @@ -2932,67 +2955,68 @@ return x_235; } else { -lean_dec(x_88); -if (lean_obj_tag(x_84) == 0) +lean_dec(x_89); +if (lean_obj_tag(x_85) == 0) { -lean_object* x_236; lean_object* x_237; -lean_free_object(x_9); +lean_object* x_236; +lean_free_object(x_10); +lean_free_object(x_82); lean_free_object(x_81); -lean_free_object(x_80); -lean_free_object(x_78); -lean_dec(x_86); -x_236 = l_Lean_identKind; -x_237 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_236, x_5, x_6, x_83); +lean_free_object(x_79); +lean_dec(x_87); +x_236 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_3, x_6, x_7, x_84); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -return x_237; +lean_dec(x_3); +return x_236; } else { -lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; +lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; +lean_dec(x_3); lean_dec(x_1); -lean_ctor_set(x_81, 1, x_96); -lean_ctor_set(x_80, 1, x_182); -x_238 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_238, 0, x_80); -lean_ctor_set(x_238, 1, x_200); -lean_ctor_set_uint64(x_238, sizeof(void*)*2, x_90); -lean_ctor_set(x_78, 0, x_238); -x_239 = lean_box(0); -x_240 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_9, x_239); -x_241 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_240, x_239); -x_242 = l_Lean_MessageData_ofList(x_241); -lean_dec(x_241); -x_243 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_244 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_244, 0, x_243); -lean_ctor_set(x_244, 1, x_242); -x_245 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_246 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_246, 0, x_244); -lean_ctor_set(x_246, 1, x_245); -x_247 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_246, x_5, x_6, x_83); -lean_dec(x_6); -x_248 = lean_ctor_get(x_247, 0); +lean_ctor_set(x_82, 1, x_97); +lean_ctor_set(x_81, 1, x_182); +x_237 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_237, 0, x_81); +lean_ctor_set(x_237, 1, x_200); +lean_ctor_set_uint64(x_237, sizeof(void*)*2, x_91); +lean_ctor_set(x_79, 0, x_237); +x_238 = lean_box(0); +x_239 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_10, x_238); +x_240 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_239, x_238); +x_241 = l_Lean_MessageData_ofList(x_240); +lean_dec(x_240); +x_242 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_243 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_243, 0, x_242); +lean_ctor_set(x_243, 1, x_241); +x_244 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_245 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_245, 0, x_243); +lean_ctor_set(x_245, 1, x_244); +x_246 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_245, x_6, x_7, x_84); +lean_dec(x_7); +x_247 = lean_ctor_get(x_246, 0); +lean_inc(x_247); +x_248 = lean_ctor_get(x_246, 1); lean_inc(x_248); -x_249 = lean_ctor_get(x_247, 1); -lean_inc(x_249); -if (lean_is_exclusive(x_247)) { - lean_ctor_release(x_247, 0); - lean_ctor_release(x_247, 1); - x_250 = x_247; +if (lean_is_exclusive(x_246)) { + lean_ctor_release(x_246, 0); + lean_ctor_release(x_246, 1); + x_249 = x_246; } else { - lean_dec_ref(x_247); - x_250 = lean_box(0); + lean_dec_ref(x_246); + x_249 = lean_box(0); } -if (lean_is_scalar(x_250)) { - x_251 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_249)) { + x_250 = lean_alloc_ctor(1, 2, 0); } else { - x_251 = x_250; + x_250 = x_249; } -lean_ctor_set(x_251, 0, x_248); -lean_ctor_set(x_251, 1, x_249); -return x_251; +lean_ctor_set(x_250, 0, x_247); +lean_ctor_set(x_250, 1, x_248); +return x_250; } } } @@ -3000,326 +3024,329 @@ return x_251; } else { -lean_object* x_252; lean_object* x_253; uint8_t x_254; -lean_dec(x_9); -if (lean_is_exclusive(x_79)) { - lean_ctor_release(x_79, 0); - lean_ctor_release(x_79, 1); - x_252 = x_79; +lean_object* x_251; lean_object* x_252; uint8_t x_253; +lean_dec(x_10); +if (lean_is_exclusive(x_80)) { + lean_ctor_release(x_80, 0); + lean_ctor_release(x_80, 1); + x_251 = x_80; } else { - lean_dec_ref(x_79); - x_252 = lean_box(0); + lean_dec_ref(x_80); + x_251 = lean_box(0); } -x_253 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__4; -x_254 = lean_string_dec_eq(x_91, x_253); -if (x_254 == 0) +x_252 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__4; +x_253 = lean_string_dec_eq(x_92, x_252); +if (x_253 == 0) { -if (lean_obj_tag(x_84) == 0) +lean_dec(x_3); +if (lean_obj_tag(x_85) == 0) { -lean_object* x_255; lean_object* x_256; -lean_free_object(x_78); -lean_dec(x_86); -lean_ctor_set(x_81, 1, x_96); -if (lean_is_scalar(x_252)) { - x_255 = lean_alloc_ctor(1, 2, 8); +lean_object* x_254; lean_object* x_255; +lean_free_object(x_79); +lean_dec(x_87); +lean_ctor_set(x_82, 1, x_97); +if (lean_is_scalar(x_251)) { + x_254 = lean_alloc_ctor(1, 2, 8); } else { - x_255 = x_252; + x_254 = x_251; } -lean_ctor_set(x_255, 0, x_80); -lean_ctor_set(x_255, 1, x_88); -lean_ctor_set_uint64(x_255, sizeof(void*)*2, x_90); -x_256 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_255, x_5, x_6, x_83); +lean_ctor_set(x_254, 0, x_81); +lean_ctor_set(x_254, 1, x_89); +lean_ctor_set_uint64(x_254, sizeof(void*)*2, x_91); +x_255 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_254, x_6, x_7, x_84); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_255); -return x_256; +lean_dec(x_254); +return x_255; } else { -lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; +lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_dec(x_1); -lean_ctor_set(x_81, 1, x_96); -if (lean_is_scalar(x_252)) { - x_257 = lean_alloc_ctor(1, 2, 8); -} else { - x_257 = x_252; -} -lean_ctor_set(x_257, 0, x_80); -lean_ctor_set(x_257, 1, x_88); -lean_ctor_set_uint64(x_257, sizeof(void*)*2, x_90); -lean_ctor_set(x_78, 0, x_257); -x_258 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_258, 0, x_78); -lean_ctor_set(x_258, 1, x_84); -x_259 = lean_box(0); -x_260 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_258, x_259); -x_261 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_260, x_259); -x_262 = l_Lean_MessageData_ofList(x_261); -lean_dec(x_261); -x_263 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_264 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_264, 0, x_263); -lean_ctor_set(x_264, 1, x_262); -x_265 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_266 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_266, 0, x_264); -lean_ctor_set(x_266, 1, x_265); -x_267 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_266, x_5, x_6, x_83); -lean_dec(x_6); -x_268 = lean_ctor_get(x_267, 0); +lean_ctor_set(x_82, 1, x_97); +if (lean_is_scalar(x_251)) { + x_256 = lean_alloc_ctor(1, 2, 8); +} else { + x_256 = x_251; +} +lean_ctor_set(x_256, 0, x_81); +lean_ctor_set(x_256, 1, x_89); +lean_ctor_set_uint64(x_256, sizeof(void*)*2, x_91); +lean_ctor_set(x_79, 0, x_256); +x_257 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_257, 0, x_79); +lean_ctor_set(x_257, 1, x_85); +x_258 = lean_box(0); +x_259 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_257, x_258); +x_260 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_259, x_258); +x_261 = l_Lean_MessageData_ofList(x_260); +lean_dec(x_260); +x_262 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_263 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_263, 0, x_262); +lean_ctor_set(x_263, 1, x_261); +x_264 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_265 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_265, 0, x_263); +lean_ctor_set(x_265, 1, x_264); +x_266 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_265, x_6, x_7, x_84); +lean_dec(x_7); +x_267 = lean_ctor_get(x_266, 0); +lean_inc(x_267); +x_268 = lean_ctor_get(x_266, 1); lean_inc(x_268); -x_269 = lean_ctor_get(x_267, 1); -lean_inc(x_269); -if (lean_is_exclusive(x_267)) { - lean_ctor_release(x_267, 0); - lean_ctor_release(x_267, 1); - x_270 = x_267; +if (lean_is_exclusive(x_266)) { + lean_ctor_release(x_266, 0); + lean_ctor_release(x_266, 1); + x_269 = x_266; } else { - lean_dec_ref(x_267); - x_270 = lean_box(0); + lean_dec_ref(x_266); + x_269 = lean_box(0); } -if (lean_is_scalar(x_270)) { - x_271 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_269)) { + x_270 = lean_alloc_ctor(1, 2, 0); } else { - x_271 = x_270; + x_270 = x_269; } -lean_ctor_set(x_271, 0, x_268); -lean_ctor_set(x_271, 1, x_269); -return x_271; +lean_ctor_set(x_270, 0, x_267); +lean_ctor_set(x_270, 1, x_268); +return x_270; } } else { -lean_object* x_272; uint8_t x_273; -lean_dec(x_91); -x_272 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__7; -x_273 = lean_string_dec_eq(x_88, x_272); -if (x_273 == 0) -{ -lean_object* x_274; uint8_t x_275; -x_274 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__8; -x_275 = lean_string_dec_eq(x_88, x_274); -if (x_275 == 0) -{ -if (lean_obj_tag(x_84) == 0) -{ -lean_object* x_276; lean_object* x_277; -lean_free_object(x_78); -lean_dec(x_86); -lean_ctor_set(x_81, 1, x_96); -lean_ctor_set(x_80, 1, x_253); -if (lean_is_scalar(x_252)) { - x_276 = lean_alloc_ctor(1, 2, 8); -} else { - x_276 = x_252; -} -lean_ctor_set(x_276, 0, x_80); -lean_ctor_set(x_276, 1, x_88); -lean_ctor_set_uint64(x_276, sizeof(void*)*2, x_90); -x_277 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_276, x_5, x_6, x_83); +lean_object* x_271; uint8_t x_272; +lean_dec(x_92); +x_271 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__7; +x_272 = lean_string_dec_eq(x_89, x_271); +if (x_272 == 0) +{ +lean_object* x_273; uint8_t x_274; +lean_dec(x_3); +x_273 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__8; +x_274 = lean_string_dec_eq(x_89, x_273); +if (x_274 == 0) +{ +if (lean_obj_tag(x_85) == 0) +{ +lean_object* x_275; lean_object* x_276; +lean_free_object(x_79); +lean_dec(x_87); +lean_ctor_set(x_82, 1, x_97); +lean_ctor_set(x_81, 1, x_252); +if (lean_is_scalar(x_251)) { + x_275 = lean_alloc_ctor(1, 2, 8); +} else { + x_275 = x_251; +} +lean_ctor_set(x_275, 0, x_81); +lean_ctor_set(x_275, 1, x_89); +lean_ctor_set_uint64(x_275, sizeof(void*)*2, x_91); +x_276 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_275, x_6, x_7, x_84); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_276); -return x_277; +lean_dec(x_275); +return x_276; } else { -lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; +lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_dec(x_1); -lean_ctor_set(x_81, 1, x_96); -lean_ctor_set(x_80, 1, x_253); -if (lean_is_scalar(x_252)) { - x_278 = lean_alloc_ctor(1, 2, 8); -} else { - x_278 = x_252; -} -lean_ctor_set(x_278, 0, x_80); -lean_ctor_set(x_278, 1, x_88); -lean_ctor_set_uint64(x_278, sizeof(void*)*2, x_90); -lean_ctor_set(x_78, 0, x_278); -x_279 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_279, 0, x_78); -lean_ctor_set(x_279, 1, x_84); -x_280 = lean_box(0); -x_281 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_279, x_280); -x_282 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_281, x_280); -x_283 = l_Lean_MessageData_ofList(x_282); -lean_dec(x_282); -x_284 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_285 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_285, 0, x_284); -lean_ctor_set(x_285, 1, x_283); -x_286 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_287 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_287, 0, x_285); -lean_ctor_set(x_287, 1, x_286); -x_288 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_287, x_5, x_6, x_83); -lean_dec(x_6); -x_289 = lean_ctor_get(x_288, 0); +lean_ctor_set(x_82, 1, x_97); +lean_ctor_set(x_81, 1, x_252); +if (lean_is_scalar(x_251)) { + x_277 = lean_alloc_ctor(1, 2, 8); +} else { + x_277 = x_251; +} +lean_ctor_set(x_277, 0, x_81); +lean_ctor_set(x_277, 1, x_89); +lean_ctor_set_uint64(x_277, sizeof(void*)*2, x_91); +lean_ctor_set(x_79, 0, x_277); +x_278 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_278, 0, x_79); +lean_ctor_set(x_278, 1, x_85); +x_279 = lean_box(0); +x_280 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_278, x_279); +x_281 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_280, x_279); +x_282 = l_Lean_MessageData_ofList(x_281); +lean_dec(x_281); +x_283 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_284 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_284, 0, x_283); +lean_ctor_set(x_284, 1, x_282); +x_285 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_286 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_286, 0, x_284); +lean_ctor_set(x_286, 1, x_285); +x_287 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_286, x_6, x_7, x_84); +lean_dec(x_7); +x_288 = lean_ctor_get(x_287, 0); +lean_inc(x_288); +x_289 = lean_ctor_get(x_287, 1); lean_inc(x_289); -x_290 = lean_ctor_get(x_288, 1); -lean_inc(x_290); -if (lean_is_exclusive(x_288)) { - lean_ctor_release(x_288, 0); - lean_ctor_release(x_288, 1); - x_291 = x_288; +if (lean_is_exclusive(x_287)) { + lean_ctor_release(x_287, 0); + lean_ctor_release(x_287, 1); + x_290 = x_287; } else { - lean_dec_ref(x_288); - x_291 = lean_box(0); + lean_dec_ref(x_287); + x_290 = lean_box(0); } -if (lean_is_scalar(x_291)) { - x_292 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_290)) { + x_291 = lean_alloc_ctor(1, 2, 0); } else { - x_292 = x_291; + x_291 = x_290; } -lean_ctor_set(x_292, 0, x_289); -lean_ctor_set(x_292, 1, x_290); -return x_292; +lean_ctor_set(x_291, 0, x_288); +lean_ctor_set(x_291, 1, x_289); +return x_291; } } else { -lean_dec(x_88); -if (lean_obj_tag(x_84) == 0) +lean_dec(x_89); +if (lean_obj_tag(x_85) == 0) { -lean_object* x_293; lean_object* x_294; -lean_dec(x_252); +lean_object* x_292; lean_object* x_293; +lean_dec(x_251); +lean_free_object(x_82); lean_free_object(x_81); -lean_free_object(x_80); -lean_free_object(x_78); -lean_dec(x_86); -x_293 = l_Lean_strLitKind; -x_294 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_293, x_5, x_6, x_83); +lean_free_object(x_79); +lean_dec(x_87); +x_292 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__10; +x_293 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_292, x_6, x_7, x_84); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -return x_294; +return x_293; } else { -lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; +lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_dec(x_1); -lean_ctor_set(x_81, 1, x_96); -lean_ctor_set(x_80, 1, x_253); -if (lean_is_scalar(x_252)) { - x_295 = lean_alloc_ctor(1, 2, 8); -} else { - x_295 = x_252; -} -lean_ctor_set(x_295, 0, x_80); -lean_ctor_set(x_295, 1, x_274); -lean_ctor_set_uint64(x_295, sizeof(void*)*2, x_90); -lean_ctor_set(x_78, 0, x_295); -x_296 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_296, 0, x_78); -lean_ctor_set(x_296, 1, x_84); -x_297 = lean_box(0); -x_298 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_296, x_297); -x_299 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_298, x_297); -x_300 = l_Lean_MessageData_ofList(x_299); -lean_dec(x_299); -x_301 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_302 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_302, 0, x_301); -lean_ctor_set(x_302, 1, x_300); -x_303 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_304 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_304, 0, x_302); -lean_ctor_set(x_304, 1, x_303); -x_305 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_304, x_5, x_6, x_83); -lean_dec(x_6); -x_306 = lean_ctor_get(x_305, 0); +lean_ctor_set(x_82, 1, x_97); +lean_ctor_set(x_81, 1, x_252); +if (lean_is_scalar(x_251)) { + x_294 = lean_alloc_ctor(1, 2, 8); +} else { + x_294 = x_251; +} +lean_ctor_set(x_294, 0, x_81); +lean_ctor_set(x_294, 1, x_273); +lean_ctor_set_uint64(x_294, sizeof(void*)*2, x_91); +lean_ctor_set(x_79, 0, x_294); +x_295 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_295, 0, x_79); +lean_ctor_set(x_295, 1, x_85); +x_296 = lean_box(0); +x_297 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_295, x_296); +x_298 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_297, x_296); +x_299 = l_Lean_MessageData_ofList(x_298); +lean_dec(x_298); +x_300 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_301 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_301, 0, x_300); +lean_ctor_set(x_301, 1, x_299); +x_302 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_303 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_303, 0, x_301); +lean_ctor_set(x_303, 1, x_302); +x_304 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_303, x_6, x_7, x_84); +lean_dec(x_7); +x_305 = lean_ctor_get(x_304, 0); +lean_inc(x_305); +x_306 = lean_ctor_get(x_304, 1); lean_inc(x_306); -x_307 = lean_ctor_get(x_305, 1); -lean_inc(x_307); -if (lean_is_exclusive(x_305)) { - lean_ctor_release(x_305, 0); - lean_ctor_release(x_305, 1); - x_308 = x_305; +if (lean_is_exclusive(x_304)) { + lean_ctor_release(x_304, 0); + lean_ctor_release(x_304, 1); + x_307 = x_304; } else { - lean_dec_ref(x_305); - x_308 = lean_box(0); + lean_dec_ref(x_304); + x_307 = lean_box(0); } -if (lean_is_scalar(x_308)) { - x_309 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_307)) { + x_308 = lean_alloc_ctor(1, 2, 0); } else { - x_309 = x_308; + x_308 = x_307; } -lean_ctor_set(x_309, 0, x_306); -lean_ctor_set(x_309, 1, x_307); -return x_309; +lean_ctor_set(x_308, 0, x_305); +lean_ctor_set(x_308, 1, x_306); +return x_308; } } } else { -lean_dec(x_88); -if (lean_obj_tag(x_84) == 0) +lean_dec(x_89); +if (lean_obj_tag(x_85) == 0) { -lean_object* x_310; lean_object* x_311; -lean_dec(x_252); +lean_object* x_309; +lean_dec(x_251); +lean_free_object(x_82); lean_free_object(x_81); -lean_free_object(x_80); -lean_free_object(x_78); -lean_dec(x_86); -x_310 = l_Lean_identKind; -x_311 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_310, x_5, x_6, x_83); +lean_free_object(x_79); +lean_dec(x_87); +x_309 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_3, x_6, x_7, x_84); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -return x_311; +lean_dec(x_3); +return x_309; } else { -lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; +lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; +lean_dec(x_3); lean_dec(x_1); -lean_ctor_set(x_81, 1, x_96); -lean_ctor_set(x_80, 1, x_253); -if (lean_is_scalar(x_252)) { - x_312 = lean_alloc_ctor(1, 2, 8); -} else { - x_312 = x_252; -} -lean_ctor_set(x_312, 0, x_80); -lean_ctor_set(x_312, 1, x_272); -lean_ctor_set_uint64(x_312, sizeof(void*)*2, x_90); -lean_ctor_set(x_78, 0, x_312); -x_313 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_313, 0, x_78); -lean_ctor_set(x_313, 1, x_84); -x_314 = lean_box(0); -x_315 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_313, x_314); -x_316 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_315, x_314); -x_317 = l_Lean_MessageData_ofList(x_316); -lean_dec(x_316); -x_318 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +lean_ctor_set(x_82, 1, x_97); +lean_ctor_set(x_81, 1, x_252); +if (lean_is_scalar(x_251)) { + x_310 = lean_alloc_ctor(1, 2, 8); +} else { + x_310 = x_251; +} +lean_ctor_set(x_310, 0, x_81); +lean_ctor_set(x_310, 1, x_271); +lean_ctor_set_uint64(x_310, sizeof(void*)*2, x_91); +lean_ctor_set(x_79, 0, x_310); +x_311 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_311, 0, x_79); +lean_ctor_set(x_311, 1, x_85); +x_312 = lean_box(0); +x_313 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_311, x_312); +x_314 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_313, x_312); +x_315 = l_Lean_MessageData_ofList(x_314); +lean_dec(x_314); +x_316 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_317 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_317, 0, x_316); +lean_ctor_set(x_317, 1, x_315); +x_318 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; x_319 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_319, 0, x_318); -lean_ctor_set(x_319, 1, x_317); -x_320 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_321 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_321, 0, x_319); -lean_ctor_set(x_321, 1, x_320); -x_322 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_321, x_5, x_6, x_83); -lean_dec(x_6); -x_323 = lean_ctor_get(x_322, 0); -lean_inc(x_323); -x_324 = lean_ctor_get(x_322, 1); -lean_inc(x_324); -if (lean_is_exclusive(x_322)) { - lean_ctor_release(x_322, 0); - lean_ctor_release(x_322, 1); - x_325 = x_322; +lean_ctor_set(x_319, 0, x_317); +lean_ctor_set(x_319, 1, x_318); +x_320 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_319, x_6, x_7, x_84); +lean_dec(x_7); +x_321 = lean_ctor_get(x_320, 0); +lean_inc(x_321); +x_322 = lean_ctor_get(x_320, 1); +lean_inc(x_322); +if (lean_is_exclusive(x_320)) { + lean_ctor_release(x_320, 0); + lean_ctor_release(x_320, 1); + x_323 = x_320; } else { - lean_dec_ref(x_322); - x_325 = lean_box(0); + lean_dec_ref(x_320); + x_323 = lean_box(0); } -if (lean_is_scalar(x_325)) { - x_326 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_323)) { + x_324 = lean_alloc_ctor(1, 2, 0); } else { - x_326 = x_325; + x_324 = x_323; } -lean_ctor_set(x_326, 0, x_323); -lean_ctor_set(x_326, 1, x_324); -return x_326; +lean_ctor_set(x_324, 0, x_321); +lean_ctor_set(x_324, 1, x_322); +return x_324; } } } @@ -3328,445 +3355,449 @@ return x_326; } else { -lean_object* x_327; uint64_t x_328; lean_object* x_329; uint8_t x_330; -x_327 = lean_ctor_get(x_81, 1); -x_328 = lean_ctor_get_uint64(x_81, sizeof(void*)*2); -lean_inc(x_327); -lean_dec(x_81); -x_329 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__1; -x_330 = lean_string_dec_eq(x_327, x_329); -lean_dec(x_327); -if (x_330 == 0) +lean_object* x_325; uint64_t x_326; lean_object* x_327; uint8_t x_328; +x_325 = lean_ctor_get(x_82, 1); +x_326 = lean_ctor_get_uint64(x_82, sizeof(void*)*2); +lean_inc(x_325); +lean_dec(x_82); +x_327 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__1; +x_328 = lean_string_dec_eq(x_325, x_327); +lean_dec(x_325); +if (x_328 == 0) { -lean_free_object(x_80); -lean_dec(x_91); -lean_dec(x_88); -lean_free_object(x_78); -lean_dec(x_86); -if (lean_obj_tag(x_84) == 0) +lean_free_object(x_81); +lean_dec(x_92); +lean_dec(x_89); +lean_free_object(x_79); +lean_dec(x_87); +lean_dec(x_3); +if (lean_obj_tag(x_85) == 0) { -lean_object* x_331; -lean_dec(x_9); -x_331 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_79, x_5, x_6, x_83); +lean_object* x_329; +lean_dec(x_10); +x_329 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_80, x_6, x_7, x_84); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_79); -return x_331; +lean_dec(x_80); +return x_329; } else { -lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; -lean_dec(x_84); -lean_dec(x_79); +lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; +lean_dec(x_85); +lean_dec(x_80); lean_dec(x_1); -x_332 = lean_box(0); -x_333 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_9, x_332); -x_334 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_333, x_332); -x_335 = l_Lean_MessageData_ofList(x_334); -lean_dec(x_334); -x_336 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_330 = lean_box(0); +x_331 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_10, x_330); +x_332 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_331, x_330); +x_333 = l_Lean_MessageData_ofList(x_332); +lean_dec(x_332); +x_334 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_335 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_335, 0, x_334); +lean_ctor_set(x_335, 1, x_333); +x_336 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; x_337 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_337, 0, x_336); -lean_ctor_set(x_337, 1, x_335); -x_338 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_339 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_339, 0, x_337); -lean_ctor_set(x_339, 1, x_338); -x_340 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_339, x_5, x_6, x_83); -lean_dec(x_6); -x_341 = lean_ctor_get(x_340, 0); -lean_inc(x_341); -x_342 = lean_ctor_get(x_340, 1); -lean_inc(x_342); -if (lean_is_exclusive(x_340)) { - lean_ctor_release(x_340, 0); - lean_ctor_release(x_340, 1); - x_343 = x_340; +lean_ctor_set(x_337, 0, x_335); +lean_ctor_set(x_337, 1, x_336); +x_338 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_337, x_6, x_7, x_84); +lean_dec(x_7); +x_339 = lean_ctor_get(x_338, 0); +lean_inc(x_339); +x_340 = lean_ctor_get(x_338, 1); +lean_inc(x_340); +if (lean_is_exclusive(x_338)) { + lean_ctor_release(x_338, 0); + lean_ctor_release(x_338, 1); + x_341 = x_338; } else { - lean_dec_ref(x_340); - x_343 = lean_box(0); + lean_dec_ref(x_338); + x_341 = lean_box(0); } -if (lean_is_scalar(x_343)) { - x_344 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_341)) { + x_342 = lean_alloc_ctor(1, 2, 0); } else { - x_344 = x_343; + x_342 = x_341; } -lean_ctor_set(x_344, 0, x_341); -lean_ctor_set(x_344, 1, x_342); -return x_344; +lean_ctor_set(x_342, 0, x_339); +lean_ctor_set(x_342, 1, x_340); +return x_342; } } else { -lean_object* x_345; lean_object* x_346; lean_object* x_347; uint8_t x_348; -if (lean_is_exclusive(x_9)) { - lean_ctor_release(x_9, 0); - lean_ctor_release(x_9, 1); - x_345 = x_9; +lean_object* x_343; lean_object* x_344; lean_object* x_345; uint8_t x_346; +if (lean_is_exclusive(x_10)) { + lean_ctor_release(x_10, 0); + lean_ctor_release(x_10, 1); + x_343 = x_10; } else { - lean_dec_ref(x_9); - x_345 = lean_box(0); + lean_dec_ref(x_10); + x_343 = lean_box(0); } -if (lean_is_exclusive(x_79)) { - lean_ctor_release(x_79, 0); - lean_ctor_release(x_79, 1); - x_346 = x_79; +if (lean_is_exclusive(x_80)) { + lean_ctor_release(x_80, 0); + lean_ctor_release(x_80, 1); + x_344 = x_80; } else { - lean_dec_ref(x_79); - x_346 = lean_box(0); + lean_dec_ref(x_80); + x_344 = lean_box(0); } -x_347 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__4; -x_348 = lean_string_dec_eq(x_91, x_347); -if (x_348 == 0) +x_345 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__4; +x_346 = lean_string_dec_eq(x_92, x_345); +if (x_346 == 0) { -if (lean_obj_tag(x_84) == 0) +lean_dec(x_3); +if (lean_obj_tag(x_85) == 0) { -lean_object* x_349; lean_object* x_350; lean_object* x_351; -lean_dec(x_345); -lean_free_object(x_78); -lean_dec(x_86); -x_349 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_349, 0, x_82); -lean_ctor_set(x_349, 1, x_329); -lean_ctor_set_uint64(x_349, sizeof(void*)*2, x_328); -lean_ctor_set(x_80, 0, x_349); -if (lean_is_scalar(x_346)) { - x_350 = lean_alloc_ctor(1, 2, 8); -} else { - x_350 = x_346; -} -lean_ctor_set(x_350, 0, x_80); -lean_ctor_set(x_350, 1, x_88); -lean_ctor_set_uint64(x_350, sizeof(void*)*2, x_90); -x_351 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_350, x_5, x_6, x_83); +lean_object* x_347; lean_object* x_348; lean_object* x_349; +lean_dec(x_343); +lean_free_object(x_79); +lean_dec(x_87); +x_347 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_347, 0, x_83); +lean_ctor_set(x_347, 1, x_327); +lean_ctor_set_uint64(x_347, sizeof(void*)*2, x_326); +lean_ctor_set(x_81, 0, x_347); +if (lean_is_scalar(x_344)) { + x_348 = lean_alloc_ctor(1, 2, 8); +} else { + x_348 = x_344; +} +lean_ctor_set(x_348, 0, x_81); +lean_ctor_set(x_348, 1, x_89); +lean_ctor_set_uint64(x_348, sizeof(void*)*2, x_91); +x_349 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_348, x_6, x_7, x_84); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_350); -return x_351; +lean_dec(x_348); +return x_349; } else { -lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; +lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_dec(x_1); -x_352 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_352, 0, x_82); -lean_ctor_set(x_352, 1, x_329); -lean_ctor_set_uint64(x_352, sizeof(void*)*2, x_328); -lean_ctor_set(x_80, 0, x_352); -if (lean_is_scalar(x_346)) { - x_353 = lean_alloc_ctor(1, 2, 8); -} else { - x_353 = x_346; -} -lean_ctor_set(x_353, 0, x_80); -lean_ctor_set(x_353, 1, x_88); -lean_ctor_set_uint64(x_353, sizeof(void*)*2, x_90); -lean_ctor_set(x_78, 0, x_353); -if (lean_is_scalar(x_345)) { - x_354 = lean_alloc_ctor(1, 2, 0); -} else { - x_354 = x_345; -} -lean_ctor_set(x_354, 0, x_78); -lean_ctor_set(x_354, 1, x_84); -x_355 = lean_box(0); -x_356 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_354, x_355); -x_357 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_356, x_355); -x_358 = l_Lean_MessageData_ofList(x_357); -lean_dec(x_357); -x_359 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_350 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_350, 0, x_83); +lean_ctor_set(x_350, 1, x_327); +lean_ctor_set_uint64(x_350, sizeof(void*)*2, x_326); +lean_ctor_set(x_81, 0, x_350); +if (lean_is_scalar(x_344)) { + x_351 = lean_alloc_ctor(1, 2, 8); +} else { + x_351 = x_344; +} +lean_ctor_set(x_351, 0, x_81); +lean_ctor_set(x_351, 1, x_89); +lean_ctor_set_uint64(x_351, sizeof(void*)*2, x_91); +lean_ctor_set(x_79, 0, x_351); +if (lean_is_scalar(x_343)) { + x_352 = lean_alloc_ctor(1, 2, 0); +} else { + x_352 = x_343; +} +lean_ctor_set(x_352, 0, x_79); +lean_ctor_set(x_352, 1, x_85); +x_353 = lean_box(0); +x_354 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_352, x_353); +x_355 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_354, x_353); +x_356 = l_Lean_MessageData_ofList(x_355); +lean_dec(x_355); +x_357 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_358 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_358, 0, x_357); +lean_ctor_set(x_358, 1, x_356); +x_359 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; x_360 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_360, 0, x_359); -lean_ctor_set(x_360, 1, x_358); -x_361 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_362 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_362, 0, x_360); -lean_ctor_set(x_362, 1, x_361); -x_363 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_362, x_5, x_6, x_83); -lean_dec(x_6); -x_364 = lean_ctor_get(x_363, 0); -lean_inc(x_364); -x_365 = lean_ctor_get(x_363, 1); -lean_inc(x_365); -if (lean_is_exclusive(x_363)) { - lean_ctor_release(x_363, 0); - lean_ctor_release(x_363, 1); - x_366 = x_363; +lean_ctor_set(x_360, 0, x_358); +lean_ctor_set(x_360, 1, x_359); +x_361 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_360, x_6, x_7, x_84); +lean_dec(x_7); +x_362 = lean_ctor_get(x_361, 0); +lean_inc(x_362); +x_363 = lean_ctor_get(x_361, 1); +lean_inc(x_363); +if (lean_is_exclusive(x_361)) { + lean_ctor_release(x_361, 0); + lean_ctor_release(x_361, 1); + x_364 = x_361; } else { - lean_dec_ref(x_363); - x_366 = lean_box(0); + lean_dec_ref(x_361); + x_364 = lean_box(0); } -if (lean_is_scalar(x_366)) { - x_367 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_364)) { + x_365 = lean_alloc_ctor(1, 2, 0); } else { - x_367 = x_366; + x_365 = x_364; } -lean_ctor_set(x_367, 0, x_364); -lean_ctor_set(x_367, 1, x_365); -return x_367; +lean_ctor_set(x_365, 0, x_362); +lean_ctor_set(x_365, 1, x_363); +return x_365; } } else { +lean_object* x_366; uint8_t x_367; +lean_dec(x_92); +x_366 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__7; +x_367 = lean_string_dec_eq(x_89, x_366); +if (x_367 == 0) +{ lean_object* x_368; uint8_t x_369; -lean_dec(x_91); -x_368 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__7; -x_369 = lean_string_dec_eq(x_88, x_368); +lean_dec(x_3); +x_368 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__8; +x_369 = lean_string_dec_eq(x_89, x_368); if (x_369 == 0) { -lean_object* x_370; uint8_t x_371; -x_370 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__8; -x_371 = lean_string_dec_eq(x_88, x_370); -if (x_371 == 0) +if (lean_obj_tag(x_85) == 0) { -if (lean_obj_tag(x_84) == 0) -{ -lean_object* x_372; lean_object* x_373; lean_object* x_374; -lean_dec(x_345); -lean_free_object(x_78); -lean_dec(x_86); -x_372 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_372, 0, x_82); -lean_ctor_set(x_372, 1, x_329); -lean_ctor_set_uint64(x_372, sizeof(void*)*2, x_328); -lean_ctor_set(x_80, 1, x_347); -lean_ctor_set(x_80, 0, x_372); -if (lean_is_scalar(x_346)) { - x_373 = lean_alloc_ctor(1, 2, 8); -} else { - x_373 = x_346; -} -lean_ctor_set(x_373, 0, x_80); -lean_ctor_set(x_373, 1, x_88); -lean_ctor_set_uint64(x_373, sizeof(void*)*2, x_90); -x_374 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_373, x_5, x_6, x_83); +lean_object* x_370; lean_object* x_371; lean_object* x_372; +lean_dec(x_343); +lean_free_object(x_79); +lean_dec(x_87); +x_370 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_370, 0, x_83); +lean_ctor_set(x_370, 1, x_327); +lean_ctor_set_uint64(x_370, sizeof(void*)*2, x_326); +lean_ctor_set(x_81, 1, x_345); +lean_ctor_set(x_81, 0, x_370); +if (lean_is_scalar(x_344)) { + x_371 = lean_alloc_ctor(1, 2, 8); +} else { + x_371 = x_344; +} +lean_ctor_set(x_371, 0, x_81); +lean_ctor_set(x_371, 1, x_89); +lean_ctor_set_uint64(x_371, sizeof(void*)*2, x_91); +x_372 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_371, x_6, x_7, x_84); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_373); -return x_374; +lean_dec(x_371); +return x_372; } else { -lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; +lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_dec(x_1); -x_375 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_375, 0, x_82); -lean_ctor_set(x_375, 1, x_329); -lean_ctor_set_uint64(x_375, sizeof(void*)*2, x_328); -lean_ctor_set(x_80, 1, x_347); -lean_ctor_set(x_80, 0, x_375); -if (lean_is_scalar(x_346)) { - x_376 = lean_alloc_ctor(1, 2, 8); -} else { - x_376 = x_346; -} -lean_ctor_set(x_376, 0, x_80); -lean_ctor_set(x_376, 1, x_88); -lean_ctor_set_uint64(x_376, sizeof(void*)*2, x_90); -lean_ctor_set(x_78, 0, x_376); -if (lean_is_scalar(x_345)) { - x_377 = lean_alloc_ctor(1, 2, 0); -} else { - x_377 = x_345; -} -lean_ctor_set(x_377, 0, x_78); -lean_ctor_set(x_377, 1, x_84); -x_378 = lean_box(0); -x_379 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_377, x_378); -x_380 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_379, x_378); -x_381 = l_Lean_MessageData_ofList(x_380); -lean_dec(x_380); -x_382 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_373 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_373, 0, x_83); +lean_ctor_set(x_373, 1, x_327); +lean_ctor_set_uint64(x_373, sizeof(void*)*2, x_326); +lean_ctor_set(x_81, 1, x_345); +lean_ctor_set(x_81, 0, x_373); +if (lean_is_scalar(x_344)) { + x_374 = lean_alloc_ctor(1, 2, 8); +} else { + x_374 = x_344; +} +lean_ctor_set(x_374, 0, x_81); +lean_ctor_set(x_374, 1, x_89); +lean_ctor_set_uint64(x_374, sizeof(void*)*2, x_91); +lean_ctor_set(x_79, 0, x_374); +if (lean_is_scalar(x_343)) { + x_375 = lean_alloc_ctor(1, 2, 0); +} else { + x_375 = x_343; +} +lean_ctor_set(x_375, 0, x_79); +lean_ctor_set(x_375, 1, x_85); +x_376 = lean_box(0); +x_377 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_375, x_376); +x_378 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_377, x_376); +x_379 = l_Lean_MessageData_ofList(x_378); +lean_dec(x_378); +x_380 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_381 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_381, 0, x_380); +lean_ctor_set(x_381, 1, x_379); +x_382 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; x_383 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_383, 0, x_382); -lean_ctor_set(x_383, 1, x_381); -x_384 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_385 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_385, 0, x_383); -lean_ctor_set(x_385, 1, x_384); -x_386 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_385, x_5, x_6, x_83); -lean_dec(x_6); -x_387 = lean_ctor_get(x_386, 0); -lean_inc(x_387); -x_388 = lean_ctor_get(x_386, 1); -lean_inc(x_388); -if (lean_is_exclusive(x_386)) { - lean_ctor_release(x_386, 0); - lean_ctor_release(x_386, 1); - x_389 = x_386; +lean_ctor_set(x_383, 0, x_381); +lean_ctor_set(x_383, 1, x_382); +x_384 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_383, x_6, x_7, x_84); +lean_dec(x_7); +x_385 = lean_ctor_get(x_384, 0); +lean_inc(x_385); +x_386 = lean_ctor_get(x_384, 1); +lean_inc(x_386); +if (lean_is_exclusive(x_384)) { + lean_ctor_release(x_384, 0); + lean_ctor_release(x_384, 1); + x_387 = x_384; } else { - lean_dec_ref(x_386); - x_389 = lean_box(0); + lean_dec_ref(x_384); + x_387 = lean_box(0); } -if (lean_is_scalar(x_389)) { - x_390 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_387)) { + x_388 = lean_alloc_ctor(1, 2, 0); } else { - x_390 = x_389; + x_388 = x_387; } -lean_ctor_set(x_390, 0, x_387); -lean_ctor_set(x_390, 1, x_388); -return x_390; +lean_ctor_set(x_388, 0, x_385); +lean_ctor_set(x_388, 1, x_386); +return x_388; } } else { -lean_dec(x_88); -if (lean_obj_tag(x_84) == 0) +lean_dec(x_89); +if (lean_obj_tag(x_85) == 0) { -lean_object* x_391; lean_object* x_392; -lean_dec(x_346); -lean_dec(x_345); -lean_free_object(x_80); -lean_free_object(x_78); -lean_dec(x_86); -x_391 = l_Lean_strLitKind; -x_392 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_391, x_5, x_6, x_83); +lean_object* x_389; lean_object* x_390; +lean_dec(x_344); +lean_dec(x_343); +lean_free_object(x_81); +lean_free_object(x_79); +lean_dec(x_87); +x_389 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__10; +x_390 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_389, x_6, x_7, x_84); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -return x_392; +return x_390; } else { -lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; +lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_dec(x_1); -x_393 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_393, 0, x_82); -lean_ctor_set(x_393, 1, x_329); -lean_ctor_set_uint64(x_393, sizeof(void*)*2, x_328); -lean_ctor_set(x_80, 1, x_347); -lean_ctor_set(x_80, 0, x_393); -if (lean_is_scalar(x_346)) { - x_394 = lean_alloc_ctor(1, 2, 8); -} else { - x_394 = x_346; -} -lean_ctor_set(x_394, 0, x_80); -lean_ctor_set(x_394, 1, x_370); -lean_ctor_set_uint64(x_394, sizeof(void*)*2, x_90); -lean_ctor_set(x_78, 0, x_394); -if (lean_is_scalar(x_345)) { - x_395 = lean_alloc_ctor(1, 2, 0); -} else { - x_395 = x_345; -} -lean_ctor_set(x_395, 0, x_78); -lean_ctor_set(x_395, 1, x_84); -x_396 = lean_box(0); -x_397 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_395, x_396); -x_398 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_397, x_396); -x_399 = l_Lean_MessageData_ofList(x_398); -lean_dec(x_398); -x_400 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_391 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_391, 0, x_83); +lean_ctor_set(x_391, 1, x_327); +lean_ctor_set_uint64(x_391, sizeof(void*)*2, x_326); +lean_ctor_set(x_81, 1, x_345); +lean_ctor_set(x_81, 0, x_391); +if (lean_is_scalar(x_344)) { + x_392 = lean_alloc_ctor(1, 2, 8); +} else { + x_392 = x_344; +} +lean_ctor_set(x_392, 0, x_81); +lean_ctor_set(x_392, 1, x_368); +lean_ctor_set_uint64(x_392, sizeof(void*)*2, x_91); +lean_ctor_set(x_79, 0, x_392); +if (lean_is_scalar(x_343)) { + x_393 = lean_alloc_ctor(1, 2, 0); +} else { + x_393 = x_343; +} +lean_ctor_set(x_393, 0, x_79); +lean_ctor_set(x_393, 1, x_85); +x_394 = lean_box(0); +x_395 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_393, x_394); +x_396 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_395, x_394); +x_397 = l_Lean_MessageData_ofList(x_396); +lean_dec(x_396); +x_398 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_399 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_399, 0, x_398); +lean_ctor_set(x_399, 1, x_397); +x_400 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; x_401 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_401, 0, x_400); -lean_ctor_set(x_401, 1, x_399); -x_402 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_403 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_403, 0, x_401); -lean_ctor_set(x_403, 1, x_402); -x_404 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_403, x_5, x_6, x_83); -lean_dec(x_6); -x_405 = lean_ctor_get(x_404, 0); -lean_inc(x_405); -x_406 = lean_ctor_get(x_404, 1); -lean_inc(x_406); -if (lean_is_exclusive(x_404)) { - lean_ctor_release(x_404, 0); - lean_ctor_release(x_404, 1); - x_407 = x_404; +lean_ctor_set(x_401, 0, x_399); +lean_ctor_set(x_401, 1, x_400); +x_402 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_401, x_6, x_7, x_84); +lean_dec(x_7); +x_403 = lean_ctor_get(x_402, 0); +lean_inc(x_403); +x_404 = lean_ctor_get(x_402, 1); +lean_inc(x_404); +if (lean_is_exclusive(x_402)) { + lean_ctor_release(x_402, 0); + lean_ctor_release(x_402, 1); + x_405 = x_402; } else { - lean_dec_ref(x_404); - x_407 = lean_box(0); + lean_dec_ref(x_402); + x_405 = lean_box(0); } -if (lean_is_scalar(x_407)) { - x_408 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_405)) { + x_406 = lean_alloc_ctor(1, 2, 0); } else { - x_408 = x_407; + x_406 = x_405; } -lean_ctor_set(x_408, 0, x_405); -lean_ctor_set(x_408, 1, x_406); -return x_408; +lean_ctor_set(x_406, 0, x_403); +lean_ctor_set(x_406, 1, x_404); +return x_406; } } } else { -lean_dec(x_88); -if (lean_obj_tag(x_84) == 0) +lean_dec(x_89); +if (lean_obj_tag(x_85) == 0) { -lean_object* x_409; lean_object* x_410; -lean_dec(x_346); -lean_dec(x_345); -lean_free_object(x_80); -lean_free_object(x_78); -lean_dec(x_86); -x_409 = l_Lean_identKind; -x_410 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_409, x_5, x_6, x_83); +lean_object* x_407; +lean_dec(x_344); +lean_dec(x_343); +lean_free_object(x_81); +lean_free_object(x_79); +lean_dec(x_87); +x_407 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_3, x_6, x_7, x_84); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -return x_410; +lean_dec(x_3); +return x_407; } else { -lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; +lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; +lean_dec(x_3); lean_dec(x_1); -x_411 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_411, 0, x_82); -lean_ctor_set(x_411, 1, x_329); -lean_ctor_set_uint64(x_411, sizeof(void*)*2, x_328); -lean_ctor_set(x_80, 1, x_347); -lean_ctor_set(x_80, 0, x_411); -if (lean_is_scalar(x_346)) { - x_412 = lean_alloc_ctor(1, 2, 8); -} else { - x_412 = x_346; -} -lean_ctor_set(x_412, 0, x_80); -lean_ctor_set(x_412, 1, x_368); -lean_ctor_set_uint64(x_412, sizeof(void*)*2, x_90); -lean_ctor_set(x_78, 0, x_412); -if (lean_is_scalar(x_345)) { - x_413 = lean_alloc_ctor(1, 2, 0); -} else { - x_413 = x_345; -} -lean_ctor_set(x_413, 0, x_78); -lean_ctor_set(x_413, 1, x_84); -x_414 = lean_box(0); -x_415 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_413, x_414); -x_416 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_415, x_414); -x_417 = l_Lean_MessageData_ofList(x_416); -lean_dec(x_416); -x_418 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_419 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_419, 0, x_418); -lean_ctor_set(x_419, 1, x_417); -x_420 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_421 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_421, 0, x_419); -lean_ctor_set(x_421, 1, x_420); -x_422 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_421, x_5, x_6, x_83); -lean_dec(x_6); -x_423 = lean_ctor_get(x_422, 0); -lean_inc(x_423); -x_424 = lean_ctor_get(x_422, 1); -lean_inc(x_424); -if (lean_is_exclusive(x_422)) { - lean_ctor_release(x_422, 0); - lean_ctor_release(x_422, 1); - x_425 = x_422; +x_408 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_408, 0, x_83); +lean_ctor_set(x_408, 1, x_327); +lean_ctor_set_uint64(x_408, sizeof(void*)*2, x_326); +lean_ctor_set(x_81, 1, x_345); +lean_ctor_set(x_81, 0, x_408); +if (lean_is_scalar(x_344)) { + x_409 = lean_alloc_ctor(1, 2, 8); +} else { + x_409 = x_344; +} +lean_ctor_set(x_409, 0, x_81); +lean_ctor_set(x_409, 1, x_366); +lean_ctor_set_uint64(x_409, sizeof(void*)*2, x_91); +lean_ctor_set(x_79, 0, x_409); +if (lean_is_scalar(x_343)) { + x_410 = lean_alloc_ctor(1, 2, 0); +} else { + x_410 = x_343; +} +lean_ctor_set(x_410, 0, x_79); +lean_ctor_set(x_410, 1, x_85); +x_411 = lean_box(0); +x_412 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_410, x_411); +x_413 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_412, x_411); +x_414 = l_Lean_MessageData_ofList(x_413); +lean_dec(x_413); +x_415 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_416 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_416, 0, x_415); +lean_ctor_set(x_416, 1, x_414); +x_417 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_418 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_418, 0, x_416); +lean_ctor_set(x_418, 1, x_417); +x_419 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_418, x_6, x_7, x_84); +lean_dec(x_7); +x_420 = lean_ctor_get(x_419, 0); +lean_inc(x_420); +x_421 = lean_ctor_get(x_419, 1); +lean_inc(x_421); +if (lean_is_exclusive(x_419)) { + lean_ctor_release(x_419, 0); + lean_ctor_release(x_419, 1); + x_422 = x_419; } else { - lean_dec_ref(x_422); - x_425 = lean_box(0); + lean_dec_ref(x_419); + x_422 = lean_box(0); } -if (lean_is_scalar(x_425)) { - x_426 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_422)) { + x_423 = lean_alloc_ctor(1, 2, 0); } else { - x_426 = x_425; + x_423 = x_422; } -lean_ctor_set(x_426, 0, x_423); -lean_ctor_set(x_426, 1, x_424); -return x_426; +lean_ctor_set(x_423, 0, x_420); +lean_ctor_set(x_423, 1, x_421); +return x_423; } } } @@ -3775,495 +3806,499 @@ return x_426; } else { -uint64_t x_427; lean_object* x_428; uint64_t x_429; lean_object* x_430; uint64_t x_431; lean_object* x_432; lean_object* x_433; uint8_t x_434; -x_427 = lean_ctor_get_uint64(x_79, sizeof(void*)*2); -x_428 = lean_ctor_get(x_80, 1); -x_429 = lean_ctor_get_uint64(x_80, sizeof(void*)*2); -lean_inc(x_428); -lean_dec(x_80); -x_430 = lean_ctor_get(x_81, 1); -lean_inc(x_430); -x_431 = lean_ctor_get_uint64(x_81, sizeof(void*)*2); -if (lean_is_exclusive(x_81)) { - lean_ctor_release(x_81, 0); - lean_ctor_release(x_81, 1); - x_432 = x_81; +uint64_t x_424; lean_object* x_425; uint64_t x_426; lean_object* x_427; uint64_t x_428; lean_object* x_429; lean_object* x_430; uint8_t x_431; +x_424 = lean_ctor_get_uint64(x_80, sizeof(void*)*2); +x_425 = lean_ctor_get(x_81, 1); +x_426 = lean_ctor_get_uint64(x_81, sizeof(void*)*2); +lean_inc(x_425); +lean_dec(x_81); +x_427 = lean_ctor_get(x_82, 1); +lean_inc(x_427); +x_428 = lean_ctor_get_uint64(x_82, sizeof(void*)*2); +if (lean_is_exclusive(x_82)) { + lean_ctor_release(x_82, 0); + lean_ctor_release(x_82, 1); + x_429 = x_82; } else { - lean_dec_ref(x_81); - x_432 = lean_box(0); + lean_dec_ref(x_82); + x_429 = lean_box(0); } -x_433 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__1; -x_434 = lean_string_dec_eq(x_430, x_433); -lean_dec(x_430); -if (x_434 == 0) +x_430 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__1; +x_431 = lean_string_dec_eq(x_427, x_430); +lean_dec(x_427); +if (x_431 == 0) { -lean_dec(x_432); -lean_dec(x_428); -lean_dec(x_88); -lean_free_object(x_78); -lean_dec(x_86); -if (lean_obj_tag(x_84) == 0) +lean_dec(x_429); +lean_dec(x_425); +lean_dec(x_89); +lean_free_object(x_79); +lean_dec(x_87); +lean_dec(x_3); +if (lean_obj_tag(x_85) == 0) { -lean_object* x_435; -lean_dec(x_9); -x_435 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_79, x_5, x_6, x_83); +lean_object* x_432; +lean_dec(x_10); +x_432 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_80, x_6, x_7, x_84); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_79); -return x_435; +lean_dec(x_80); +return x_432; } else { -lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; -lean_dec(x_84); -lean_dec(x_79); +lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; +lean_dec(x_85); +lean_dec(x_80); lean_dec(x_1); -x_436 = lean_box(0); -x_437 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_9, x_436); -x_438 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_437, x_436); -x_439 = l_Lean_MessageData_ofList(x_438); -lean_dec(x_438); -x_440 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_441 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_441, 0, x_440); -lean_ctor_set(x_441, 1, x_439); -x_442 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_443 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_443, 0, x_441); -lean_ctor_set(x_443, 1, x_442); -x_444 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_443, x_5, x_6, x_83); -lean_dec(x_6); -x_445 = lean_ctor_get(x_444, 0); -lean_inc(x_445); -x_446 = lean_ctor_get(x_444, 1); -lean_inc(x_446); -if (lean_is_exclusive(x_444)) { - lean_ctor_release(x_444, 0); - lean_ctor_release(x_444, 1); - x_447 = x_444; +x_433 = lean_box(0); +x_434 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_10, x_433); +x_435 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_434, x_433); +x_436 = l_Lean_MessageData_ofList(x_435); +lean_dec(x_435); +x_437 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_438 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_438, 0, x_437); +lean_ctor_set(x_438, 1, x_436); +x_439 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_440 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_440, 0, x_438); +lean_ctor_set(x_440, 1, x_439); +x_441 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_440, x_6, x_7, x_84); +lean_dec(x_7); +x_442 = lean_ctor_get(x_441, 0); +lean_inc(x_442); +x_443 = lean_ctor_get(x_441, 1); +lean_inc(x_443); +if (lean_is_exclusive(x_441)) { + lean_ctor_release(x_441, 0); + lean_ctor_release(x_441, 1); + x_444 = x_441; } else { - lean_dec_ref(x_444); - x_447 = lean_box(0); + lean_dec_ref(x_441); + x_444 = lean_box(0); } -if (lean_is_scalar(x_447)) { - x_448 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_444)) { + x_445 = lean_alloc_ctor(1, 2, 0); } else { - x_448 = x_447; + x_445 = x_444; } -lean_ctor_set(x_448, 0, x_445); -lean_ctor_set(x_448, 1, x_446); -return x_448; +lean_ctor_set(x_445, 0, x_442); +lean_ctor_set(x_445, 1, x_443); +return x_445; } } else { -lean_object* x_449; lean_object* x_450; lean_object* x_451; uint8_t x_452; -if (lean_is_exclusive(x_9)) { - lean_ctor_release(x_9, 0); - lean_ctor_release(x_9, 1); - x_449 = x_9; +lean_object* x_446; lean_object* x_447; lean_object* x_448; uint8_t x_449; +if (lean_is_exclusive(x_10)) { + lean_ctor_release(x_10, 0); + lean_ctor_release(x_10, 1); + x_446 = x_10; } else { - lean_dec_ref(x_9); - x_449 = lean_box(0); + lean_dec_ref(x_10); + x_446 = lean_box(0); } -if (lean_is_exclusive(x_79)) { - lean_ctor_release(x_79, 0); - lean_ctor_release(x_79, 1); - x_450 = x_79; +if (lean_is_exclusive(x_80)) { + lean_ctor_release(x_80, 0); + lean_ctor_release(x_80, 1); + x_447 = x_80; } else { - lean_dec_ref(x_79); - x_450 = lean_box(0); + lean_dec_ref(x_80); + x_447 = lean_box(0); } -x_451 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__4; -x_452 = lean_string_dec_eq(x_428, x_451); -if (x_452 == 0) +x_448 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__4; +x_449 = lean_string_dec_eq(x_425, x_448); +if (x_449 == 0) { -if (lean_obj_tag(x_84) == 0) +lean_dec(x_3); +if (lean_obj_tag(x_85) == 0) { -lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; -lean_dec(x_449); -lean_free_object(x_78); -lean_dec(x_86); -if (lean_is_scalar(x_432)) { - x_453 = lean_alloc_ctor(1, 2, 8); -} else { - x_453 = x_432; -} -lean_ctor_set(x_453, 0, x_82); -lean_ctor_set(x_453, 1, x_433); -lean_ctor_set_uint64(x_453, sizeof(void*)*2, x_431); -x_454 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_454, 0, x_453); -lean_ctor_set(x_454, 1, x_428); -lean_ctor_set_uint64(x_454, sizeof(void*)*2, x_429); -if (lean_is_scalar(x_450)) { - x_455 = lean_alloc_ctor(1, 2, 8); +lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; +lean_dec(x_446); +lean_free_object(x_79); +lean_dec(x_87); +if (lean_is_scalar(x_429)) { + x_450 = lean_alloc_ctor(1, 2, 8); +} else { + x_450 = x_429; +} +lean_ctor_set(x_450, 0, x_83); +lean_ctor_set(x_450, 1, x_430); +lean_ctor_set_uint64(x_450, sizeof(void*)*2, x_428); +x_451 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_451, 0, x_450); +lean_ctor_set(x_451, 1, x_425); +lean_ctor_set_uint64(x_451, sizeof(void*)*2, x_426); +if (lean_is_scalar(x_447)) { + x_452 = lean_alloc_ctor(1, 2, 8); } else { - x_455 = x_450; + x_452 = x_447; } -lean_ctor_set(x_455, 0, x_454); -lean_ctor_set(x_455, 1, x_88); -lean_ctor_set_uint64(x_455, sizeof(void*)*2, x_427); -x_456 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_455, x_5, x_6, x_83); +lean_ctor_set(x_452, 0, x_451); +lean_ctor_set(x_452, 1, x_89); +lean_ctor_set_uint64(x_452, sizeof(void*)*2, x_424); +x_453 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_452, x_6, x_7, x_84); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_455); -return x_456; +lean_dec(x_452); +return x_453; } else { -lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; +lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_dec(x_1); -if (lean_is_scalar(x_432)) { - x_457 = lean_alloc_ctor(1, 2, 8); -} else { - x_457 = x_432; -} -lean_ctor_set(x_457, 0, x_82); -lean_ctor_set(x_457, 1, x_433); -lean_ctor_set_uint64(x_457, sizeof(void*)*2, x_431); -x_458 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_458, 0, x_457); -lean_ctor_set(x_458, 1, x_428); -lean_ctor_set_uint64(x_458, sizeof(void*)*2, x_429); -if (lean_is_scalar(x_450)) { - x_459 = lean_alloc_ctor(1, 2, 8); -} else { - x_459 = x_450; -} -lean_ctor_set(x_459, 0, x_458); -lean_ctor_set(x_459, 1, x_88); -lean_ctor_set_uint64(x_459, sizeof(void*)*2, x_427); -lean_ctor_set(x_78, 0, x_459); -if (lean_is_scalar(x_449)) { - x_460 = lean_alloc_ctor(1, 2, 0); -} else { - x_460 = x_449; -} -lean_ctor_set(x_460, 0, x_78); -lean_ctor_set(x_460, 1, x_84); -x_461 = lean_box(0); -x_462 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_460, x_461); -x_463 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_462, x_461); -x_464 = l_Lean_MessageData_ofList(x_463); -lean_dec(x_463); -x_465 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_466 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_466, 0, x_465); -lean_ctor_set(x_466, 1, x_464); -x_467 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_468 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_468, 0, x_466); -lean_ctor_set(x_468, 1, x_467); -x_469 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_468, x_5, x_6, x_83); -lean_dec(x_6); -x_470 = lean_ctor_get(x_469, 0); -lean_inc(x_470); -x_471 = lean_ctor_get(x_469, 1); -lean_inc(x_471); -if (lean_is_exclusive(x_469)) { - lean_ctor_release(x_469, 0); - lean_ctor_release(x_469, 1); - x_472 = x_469; +if (lean_is_scalar(x_429)) { + x_454 = lean_alloc_ctor(1, 2, 8); +} else { + x_454 = x_429; +} +lean_ctor_set(x_454, 0, x_83); +lean_ctor_set(x_454, 1, x_430); +lean_ctor_set_uint64(x_454, sizeof(void*)*2, x_428); +x_455 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_455, 0, x_454); +lean_ctor_set(x_455, 1, x_425); +lean_ctor_set_uint64(x_455, sizeof(void*)*2, x_426); +if (lean_is_scalar(x_447)) { + x_456 = lean_alloc_ctor(1, 2, 8); +} else { + x_456 = x_447; +} +lean_ctor_set(x_456, 0, x_455); +lean_ctor_set(x_456, 1, x_89); +lean_ctor_set_uint64(x_456, sizeof(void*)*2, x_424); +lean_ctor_set(x_79, 0, x_456); +if (lean_is_scalar(x_446)) { + x_457 = lean_alloc_ctor(1, 2, 0); +} else { + x_457 = x_446; +} +lean_ctor_set(x_457, 0, x_79); +lean_ctor_set(x_457, 1, x_85); +x_458 = lean_box(0); +x_459 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_457, x_458); +x_460 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_459, x_458); +x_461 = l_Lean_MessageData_ofList(x_460); +lean_dec(x_460); +x_462 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_463 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_463, 0, x_462); +lean_ctor_set(x_463, 1, x_461); +x_464 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_465 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_465, 0, x_463); +lean_ctor_set(x_465, 1, x_464); +x_466 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_465, x_6, x_7, x_84); +lean_dec(x_7); +x_467 = lean_ctor_get(x_466, 0); +lean_inc(x_467); +x_468 = lean_ctor_get(x_466, 1); +lean_inc(x_468); +if (lean_is_exclusive(x_466)) { + lean_ctor_release(x_466, 0); + lean_ctor_release(x_466, 1); + x_469 = x_466; } else { - lean_dec_ref(x_469); - x_472 = lean_box(0); + lean_dec_ref(x_466); + x_469 = lean_box(0); } -if (lean_is_scalar(x_472)) { - x_473 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_469)) { + x_470 = lean_alloc_ctor(1, 2, 0); } else { - x_473 = x_472; + x_470 = x_469; } -lean_ctor_set(x_473, 0, x_470); -lean_ctor_set(x_473, 1, x_471); -return x_473; +lean_ctor_set(x_470, 0, x_467); +lean_ctor_set(x_470, 1, x_468); +return x_470; } } else { -lean_object* x_474; uint8_t x_475; -lean_dec(x_428); -x_474 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__7; -x_475 = lean_string_dec_eq(x_88, x_474); -if (x_475 == 0) +lean_object* x_471; uint8_t x_472; +lean_dec(x_425); +x_471 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__7; +x_472 = lean_string_dec_eq(x_89, x_471); +if (x_472 == 0) { -lean_object* x_476; uint8_t x_477; -x_476 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__8; -x_477 = lean_string_dec_eq(x_88, x_476); -if (x_477 == 0) +lean_object* x_473; uint8_t x_474; +lean_dec(x_3); +x_473 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__8; +x_474 = lean_string_dec_eq(x_89, x_473); +if (x_474 == 0) { -if (lean_obj_tag(x_84) == 0) +if (lean_obj_tag(x_85) == 0) { -lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; -lean_dec(x_449); -lean_free_object(x_78); -lean_dec(x_86); -if (lean_is_scalar(x_432)) { - x_478 = lean_alloc_ctor(1, 2, 8); -} else { - x_478 = x_432; -} -lean_ctor_set(x_478, 0, x_82); -lean_ctor_set(x_478, 1, x_433); -lean_ctor_set_uint64(x_478, sizeof(void*)*2, x_431); -x_479 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_479, 0, x_478); -lean_ctor_set(x_479, 1, x_451); -lean_ctor_set_uint64(x_479, sizeof(void*)*2, x_429); -if (lean_is_scalar(x_450)) { - x_480 = lean_alloc_ctor(1, 2, 8); +lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; +lean_dec(x_446); +lean_free_object(x_79); +lean_dec(x_87); +if (lean_is_scalar(x_429)) { + x_475 = lean_alloc_ctor(1, 2, 8); +} else { + x_475 = x_429; +} +lean_ctor_set(x_475, 0, x_83); +lean_ctor_set(x_475, 1, x_430); +lean_ctor_set_uint64(x_475, sizeof(void*)*2, x_428); +x_476 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_476, 0, x_475); +lean_ctor_set(x_476, 1, x_448); +lean_ctor_set_uint64(x_476, sizeof(void*)*2, x_426); +if (lean_is_scalar(x_447)) { + x_477 = lean_alloc_ctor(1, 2, 8); } else { - x_480 = x_450; + x_477 = x_447; } -lean_ctor_set(x_480, 0, x_479); -lean_ctor_set(x_480, 1, x_88); -lean_ctor_set_uint64(x_480, sizeof(void*)*2, x_427); -x_481 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_480, x_5, x_6, x_83); +lean_ctor_set(x_477, 0, x_476); +lean_ctor_set(x_477, 1, x_89); +lean_ctor_set_uint64(x_477, sizeof(void*)*2, x_424); +x_478 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_477, x_6, x_7, x_84); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_480); -return x_481; +lean_dec(x_477); +return x_478; } else { -lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; +lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_dec(x_1); -if (lean_is_scalar(x_432)) { - x_482 = lean_alloc_ctor(1, 2, 8); -} else { - x_482 = x_432; -} -lean_ctor_set(x_482, 0, x_82); -lean_ctor_set(x_482, 1, x_433); -lean_ctor_set_uint64(x_482, sizeof(void*)*2, x_431); -x_483 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_483, 0, x_482); -lean_ctor_set(x_483, 1, x_451); -lean_ctor_set_uint64(x_483, sizeof(void*)*2, x_429); -if (lean_is_scalar(x_450)) { - x_484 = lean_alloc_ctor(1, 2, 8); -} else { - x_484 = x_450; -} -lean_ctor_set(x_484, 0, x_483); -lean_ctor_set(x_484, 1, x_88); -lean_ctor_set_uint64(x_484, sizeof(void*)*2, x_427); -lean_ctor_set(x_78, 0, x_484); -if (lean_is_scalar(x_449)) { - x_485 = lean_alloc_ctor(1, 2, 0); -} else { - x_485 = x_449; -} -lean_ctor_set(x_485, 0, x_78); -lean_ctor_set(x_485, 1, x_84); -x_486 = lean_box(0); -x_487 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_485, x_486); -x_488 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_487, x_486); -x_489 = l_Lean_MessageData_ofList(x_488); -lean_dec(x_488); -x_490 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_491 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_491, 0, x_490); -lean_ctor_set(x_491, 1, x_489); -x_492 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_493 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_493, 0, x_491); -lean_ctor_set(x_493, 1, x_492); -x_494 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_493, x_5, x_6, x_83); -lean_dec(x_6); -x_495 = lean_ctor_get(x_494, 0); -lean_inc(x_495); -x_496 = lean_ctor_get(x_494, 1); -lean_inc(x_496); -if (lean_is_exclusive(x_494)) { - lean_ctor_release(x_494, 0); - lean_ctor_release(x_494, 1); - x_497 = x_494; +if (lean_is_scalar(x_429)) { + x_479 = lean_alloc_ctor(1, 2, 8); +} else { + x_479 = x_429; +} +lean_ctor_set(x_479, 0, x_83); +lean_ctor_set(x_479, 1, x_430); +lean_ctor_set_uint64(x_479, sizeof(void*)*2, x_428); +x_480 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_480, 0, x_479); +lean_ctor_set(x_480, 1, x_448); +lean_ctor_set_uint64(x_480, sizeof(void*)*2, x_426); +if (lean_is_scalar(x_447)) { + x_481 = lean_alloc_ctor(1, 2, 8); +} else { + x_481 = x_447; +} +lean_ctor_set(x_481, 0, x_480); +lean_ctor_set(x_481, 1, x_89); +lean_ctor_set_uint64(x_481, sizeof(void*)*2, x_424); +lean_ctor_set(x_79, 0, x_481); +if (lean_is_scalar(x_446)) { + x_482 = lean_alloc_ctor(1, 2, 0); +} else { + x_482 = x_446; +} +lean_ctor_set(x_482, 0, x_79); +lean_ctor_set(x_482, 1, x_85); +x_483 = lean_box(0); +x_484 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_482, x_483); +x_485 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_484, x_483); +x_486 = l_Lean_MessageData_ofList(x_485); +lean_dec(x_485); +x_487 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_488 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_488, 0, x_487); +lean_ctor_set(x_488, 1, x_486); +x_489 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_490 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_490, 0, x_488); +lean_ctor_set(x_490, 1, x_489); +x_491 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_490, x_6, x_7, x_84); +lean_dec(x_7); +x_492 = lean_ctor_get(x_491, 0); +lean_inc(x_492); +x_493 = lean_ctor_get(x_491, 1); +lean_inc(x_493); +if (lean_is_exclusive(x_491)) { + lean_ctor_release(x_491, 0); + lean_ctor_release(x_491, 1); + x_494 = x_491; } else { - lean_dec_ref(x_494); - x_497 = lean_box(0); + lean_dec_ref(x_491); + x_494 = lean_box(0); } -if (lean_is_scalar(x_497)) { - x_498 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_494)) { + x_495 = lean_alloc_ctor(1, 2, 0); } else { - x_498 = x_497; + x_495 = x_494; } -lean_ctor_set(x_498, 0, x_495); -lean_ctor_set(x_498, 1, x_496); -return x_498; +lean_ctor_set(x_495, 0, x_492); +lean_ctor_set(x_495, 1, x_493); +return x_495; } } else { -lean_dec(x_88); -if (lean_obj_tag(x_84) == 0) -{ -lean_object* x_499; lean_object* x_500; -lean_dec(x_450); -lean_dec(x_449); -lean_dec(x_432); -lean_free_object(x_78); -lean_dec(x_86); -x_499 = l_Lean_strLitKind; -x_500 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_499, x_5, x_6, x_83); +lean_dec(x_89); +if (lean_obj_tag(x_85) == 0) +{ +lean_object* x_496; lean_object* x_497; +lean_dec(x_447); +lean_dec(x_446); +lean_dec(x_429); +lean_free_object(x_79); +lean_dec(x_87); +x_496 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__10; +x_497 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_496, x_6, x_7, x_84); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -return x_500; +return x_497; } else { -lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; +lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_dec(x_1); -if (lean_is_scalar(x_432)) { - x_501 = lean_alloc_ctor(1, 2, 8); -} else { - x_501 = x_432; -} -lean_ctor_set(x_501, 0, x_82); -lean_ctor_set(x_501, 1, x_433); -lean_ctor_set_uint64(x_501, sizeof(void*)*2, x_431); -x_502 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_502, 0, x_501); -lean_ctor_set(x_502, 1, x_451); -lean_ctor_set_uint64(x_502, sizeof(void*)*2, x_429); -if (lean_is_scalar(x_450)) { - x_503 = lean_alloc_ctor(1, 2, 8); -} else { - x_503 = x_450; -} -lean_ctor_set(x_503, 0, x_502); -lean_ctor_set(x_503, 1, x_476); -lean_ctor_set_uint64(x_503, sizeof(void*)*2, x_427); -lean_ctor_set(x_78, 0, x_503); -if (lean_is_scalar(x_449)) { - x_504 = lean_alloc_ctor(1, 2, 0); -} else { - x_504 = x_449; -} -lean_ctor_set(x_504, 0, x_78); -lean_ctor_set(x_504, 1, x_84); -x_505 = lean_box(0); -x_506 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_504, x_505); -x_507 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_506, x_505); -x_508 = l_Lean_MessageData_ofList(x_507); -lean_dec(x_507); -x_509 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_510 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_510, 0, x_509); -lean_ctor_set(x_510, 1, x_508); -x_511 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_512 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_512, 0, x_510); -lean_ctor_set(x_512, 1, x_511); -x_513 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_512, x_5, x_6, x_83); -lean_dec(x_6); -x_514 = lean_ctor_get(x_513, 0); -lean_inc(x_514); -x_515 = lean_ctor_get(x_513, 1); -lean_inc(x_515); -if (lean_is_exclusive(x_513)) { - lean_ctor_release(x_513, 0); - lean_ctor_release(x_513, 1); - x_516 = x_513; +if (lean_is_scalar(x_429)) { + x_498 = lean_alloc_ctor(1, 2, 8); +} else { + x_498 = x_429; +} +lean_ctor_set(x_498, 0, x_83); +lean_ctor_set(x_498, 1, x_430); +lean_ctor_set_uint64(x_498, sizeof(void*)*2, x_428); +x_499 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_499, 0, x_498); +lean_ctor_set(x_499, 1, x_448); +lean_ctor_set_uint64(x_499, sizeof(void*)*2, x_426); +if (lean_is_scalar(x_447)) { + x_500 = lean_alloc_ctor(1, 2, 8); +} else { + x_500 = x_447; +} +lean_ctor_set(x_500, 0, x_499); +lean_ctor_set(x_500, 1, x_473); +lean_ctor_set_uint64(x_500, sizeof(void*)*2, x_424); +lean_ctor_set(x_79, 0, x_500); +if (lean_is_scalar(x_446)) { + x_501 = lean_alloc_ctor(1, 2, 0); +} else { + x_501 = x_446; +} +lean_ctor_set(x_501, 0, x_79); +lean_ctor_set(x_501, 1, x_85); +x_502 = lean_box(0); +x_503 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_501, x_502); +x_504 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_503, x_502); +x_505 = l_Lean_MessageData_ofList(x_504); +lean_dec(x_504); +x_506 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_507 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_507, 0, x_506); +lean_ctor_set(x_507, 1, x_505); +x_508 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_509 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_509, 0, x_507); +lean_ctor_set(x_509, 1, x_508); +x_510 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_509, x_6, x_7, x_84); +lean_dec(x_7); +x_511 = lean_ctor_get(x_510, 0); +lean_inc(x_511); +x_512 = lean_ctor_get(x_510, 1); +lean_inc(x_512); +if (lean_is_exclusive(x_510)) { + lean_ctor_release(x_510, 0); + lean_ctor_release(x_510, 1); + x_513 = x_510; } else { - lean_dec_ref(x_513); - x_516 = lean_box(0); + lean_dec_ref(x_510); + x_513 = lean_box(0); } -if (lean_is_scalar(x_516)) { - x_517 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_513)) { + x_514 = lean_alloc_ctor(1, 2, 0); } else { - x_517 = x_516; + x_514 = x_513; } -lean_ctor_set(x_517, 0, x_514); -lean_ctor_set(x_517, 1, x_515); -return x_517; +lean_ctor_set(x_514, 0, x_511); +lean_ctor_set(x_514, 1, x_512); +return x_514; } } } else { -lean_dec(x_88); -if (lean_obj_tag(x_84) == 0) -{ -lean_object* x_518; lean_object* x_519; -lean_dec(x_450); -lean_dec(x_449); -lean_dec(x_432); -lean_free_object(x_78); -lean_dec(x_86); -x_518 = l_Lean_identKind; -x_519 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_518, x_5, x_6, x_83); +lean_dec(x_89); +if (lean_obj_tag(x_85) == 0) +{ +lean_object* x_515; +lean_dec(x_447); +lean_dec(x_446); +lean_dec(x_429); +lean_free_object(x_79); +lean_dec(x_87); +x_515 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_3, x_6, x_7, x_84); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -return x_519; +lean_dec(x_3); +return x_515; } else { -lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; +lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; +lean_dec(x_3); lean_dec(x_1); -if (lean_is_scalar(x_432)) { - x_520 = lean_alloc_ctor(1, 2, 8); -} else { - x_520 = x_432; -} -lean_ctor_set(x_520, 0, x_82); -lean_ctor_set(x_520, 1, x_433); -lean_ctor_set_uint64(x_520, sizeof(void*)*2, x_431); -x_521 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_521, 0, x_520); -lean_ctor_set(x_521, 1, x_451); -lean_ctor_set_uint64(x_521, sizeof(void*)*2, x_429); -if (lean_is_scalar(x_450)) { - x_522 = lean_alloc_ctor(1, 2, 8); -} else { - x_522 = x_450; -} -lean_ctor_set(x_522, 0, x_521); -lean_ctor_set(x_522, 1, x_474); -lean_ctor_set_uint64(x_522, sizeof(void*)*2, x_427); -lean_ctor_set(x_78, 0, x_522); -if (lean_is_scalar(x_449)) { - x_523 = lean_alloc_ctor(1, 2, 0); -} else { - x_523 = x_449; -} -lean_ctor_set(x_523, 0, x_78); -lean_ctor_set(x_523, 1, x_84); -x_524 = lean_box(0); -x_525 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_523, x_524); -x_526 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_525, x_524); -x_527 = l_Lean_MessageData_ofList(x_526); -lean_dec(x_526); -x_528 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_529 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_529, 0, x_528); -lean_ctor_set(x_529, 1, x_527); -x_530 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_531 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_531, 0, x_529); -lean_ctor_set(x_531, 1, x_530); -x_532 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_531, x_5, x_6, x_83); -lean_dec(x_6); -x_533 = lean_ctor_get(x_532, 0); -lean_inc(x_533); -x_534 = lean_ctor_get(x_532, 1); -lean_inc(x_534); -if (lean_is_exclusive(x_532)) { - lean_ctor_release(x_532, 0); - lean_ctor_release(x_532, 1); - x_535 = x_532; +if (lean_is_scalar(x_429)) { + x_516 = lean_alloc_ctor(1, 2, 8); +} else { + x_516 = x_429; +} +lean_ctor_set(x_516, 0, x_83); +lean_ctor_set(x_516, 1, x_430); +lean_ctor_set_uint64(x_516, sizeof(void*)*2, x_428); +x_517 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_517, 0, x_516); +lean_ctor_set(x_517, 1, x_448); +lean_ctor_set_uint64(x_517, sizeof(void*)*2, x_426); +if (lean_is_scalar(x_447)) { + x_518 = lean_alloc_ctor(1, 2, 8); +} else { + x_518 = x_447; +} +lean_ctor_set(x_518, 0, x_517); +lean_ctor_set(x_518, 1, x_471); +lean_ctor_set_uint64(x_518, sizeof(void*)*2, x_424); +lean_ctor_set(x_79, 0, x_518); +if (lean_is_scalar(x_446)) { + x_519 = lean_alloc_ctor(1, 2, 0); +} else { + x_519 = x_446; +} +lean_ctor_set(x_519, 0, x_79); +lean_ctor_set(x_519, 1, x_85); +x_520 = lean_box(0); +x_521 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_519, x_520); +x_522 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_521, x_520); +x_523 = l_Lean_MessageData_ofList(x_522); +lean_dec(x_522); +x_524 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_525 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_525, 0, x_524); +lean_ctor_set(x_525, 1, x_523); +x_526 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_527 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_527, 0, x_525); +lean_ctor_set(x_527, 1, x_526); +x_528 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_527, x_6, x_7, x_84); +lean_dec(x_7); +x_529 = lean_ctor_get(x_528, 0); +lean_inc(x_529); +x_530 = lean_ctor_get(x_528, 1); +lean_inc(x_530); +if (lean_is_exclusive(x_528)) { + lean_ctor_release(x_528, 0); + lean_ctor_release(x_528, 1); + x_531 = x_528; } else { - lean_dec_ref(x_532); - x_535 = lean_box(0); + lean_dec_ref(x_528); + x_531 = lean_box(0); } -if (lean_is_scalar(x_535)) { - x_536 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_531)) { + x_532 = lean_alloc_ctor(1, 2, 0); } else { - x_536 = x_535; + x_532 = x_531; } -lean_ctor_set(x_536, 0, x_533); -lean_ctor_set(x_536, 1, x_534); -return x_536; +lean_ctor_set(x_532, 0, x_529); +lean_ctor_set(x_532, 1, x_530); +return x_532; } } } @@ -4272,537 +4307,541 @@ return x_536; } else { -lean_object* x_537; lean_object* x_538; uint64_t x_539; lean_object* x_540; uint64_t x_541; lean_object* x_542; lean_object* x_543; uint64_t x_544; lean_object* x_545; lean_object* x_546; uint8_t x_547; -x_537 = lean_ctor_get(x_78, 1); -lean_inc(x_537); -lean_dec(x_78); -x_538 = lean_ctor_get(x_79, 1); -lean_inc(x_538); -x_539 = lean_ctor_get_uint64(x_79, sizeof(void*)*2); -x_540 = lean_ctor_get(x_80, 1); -lean_inc(x_540); -x_541 = lean_ctor_get_uint64(x_80, sizeof(void*)*2); -if (lean_is_exclusive(x_80)) { - lean_ctor_release(x_80, 0); - lean_ctor_release(x_80, 1); - x_542 = x_80; -} else { - lean_dec_ref(x_80); - x_542 = lean_box(0); -} -x_543 = lean_ctor_get(x_81, 1); -lean_inc(x_543); -x_544 = lean_ctor_get_uint64(x_81, sizeof(void*)*2); +lean_object* x_533; lean_object* x_534; uint64_t x_535; lean_object* x_536; uint64_t x_537; lean_object* x_538; lean_object* x_539; uint64_t x_540; lean_object* x_541; lean_object* x_542; uint8_t x_543; +x_533 = lean_ctor_get(x_79, 1); +lean_inc(x_533); +lean_dec(x_79); +x_534 = lean_ctor_get(x_80, 1); +lean_inc(x_534); +x_535 = lean_ctor_get_uint64(x_80, sizeof(void*)*2); +x_536 = lean_ctor_get(x_81, 1); +lean_inc(x_536); +x_537 = lean_ctor_get_uint64(x_81, sizeof(void*)*2); if (lean_is_exclusive(x_81)) { lean_ctor_release(x_81, 0); lean_ctor_release(x_81, 1); - x_545 = x_81; + x_538 = x_81; } else { lean_dec_ref(x_81); - x_545 = lean_box(0); + x_538 = lean_box(0); +} +x_539 = lean_ctor_get(x_82, 1); +lean_inc(x_539); +x_540 = lean_ctor_get_uint64(x_82, sizeof(void*)*2); +if (lean_is_exclusive(x_82)) { + lean_ctor_release(x_82, 0); + lean_ctor_release(x_82, 1); + x_541 = x_82; +} else { + lean_dec_ref(x_82); + x_541 = lean_box(0); } -x_546 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__1; -x_547 = lean_string_dec_eq(x_543, x_546); -lean_dec(x_543); -if (x_547 == 0) +x_542 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__1; +x_543 = lean_string_dec_eq(x_539, x_542); +lean_dec(x_539); +if (x_543 == 0) { -lean_dec(x_545); -lean_dec(x_542); -lean_dec(x_540); +lean_dec(x_541); lean_dec(x_538); -lean_dec(x_537); -if (lean_obj_tag(x_84) == 0) +lean_dec(x_536); +lean_dec(x_534); +lean_dec(x_533); +lean_dec(x_3); +if (lean_obj_tag(x_85) == 0) { -lean_object* x_548; -lean_dec(x_9); -x_548 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_79, x_5, x_6, x_83); +lean_object* x_544; +lean_dec(x_10); +x_544 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_80, x_6, x_7, x_84); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_79); -return x_548; +lean_dec(x_80); +return x_544; } else { -lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; -lean_dec(x_84); -lean_dec(x_79); +lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; +lean_dec(x_85); +lean_dec(x_80); lean_dec(x_1); -x_549 = lean_box(0); -x_550 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_9, x_549); -x_551 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_550, x_549); -x_552 = l_Lean_MessageData_ofList(x_551); -lean_dec(x_551); -x_553 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_554 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_554, 0, x_553); -lean_ctor_set(x_554, 1, x_552); -x_555 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_556 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_556, 0, x_554); -lean_ctor_set(x_556, 1, x_555); -x_557 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_556, x_5, x_6, x_83); -lean_dec(x_6); -x_558 = lean_ctor_get(x_557, 0); -lean_inc(x_558); -x_559 = lean_ctor_get(x_557, 1); -lean_inc(x_559); -if (lean_is_exclusive(x_557)) { - lean_ctor_release(x_557, 0); - lean_ctor_release(x_557, 1); - x_560 = x_557; +x_545 = lean_box(0); +x_546 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_10, x_545); +x_547 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_546, x_545); +x_548 = l_Lean_MessageData_ofList(x_547); +lean_dec(x_547); +x_549 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_550 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_550, 0, x_549); +lean_ctor_set(x_550, 1, x_548); +x_551 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_552 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_552, 0, x_550); +lean_ctor_set(x_552, 1, x_551); +x_553 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_552, x_6, x_7, x_84); +lean_dec(x_7); +x_554 = lean_ctor_get(x_553, 0); +lean_inc(x_554); +x_555 = lean_ctor_get(x_553, 1); +lean_inc(x_555); +if (lean_is_exclusive(x_553)) { + lean_ctor_release(x_553, 0); + lean_ctor_release(x_553, 1); + x_556 = x_553; } else { - lean_dec_ref(x_557); - x_560 = lean_box(0); + lean_dec_ref(x_553); + x_556 = lean_box(0); } -if (lean_is_scalar(x_560)) { - x_561 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_556)) { + x_557 = lean_alloc_ctor(1, 2, 0); } else { - x_561 = x_560; + x_557 = x_556; } -lean_ctor_set(x_561, 0, x_558); -lean_ctor_set(x_561, 1, x_559); -return x_561; +lean_ctor_set(x_557, 0, x_554); +lean_ctor_set(x_557, 1, x_555); +return x_557; } } else { -lean_object* x_562; lean_object* x_563; lean_object* x_564; uint8_t x_565; -if (lean_is_exclusive(x_9)) { - lean_ctor_release(x_9, 0); - lean_ctor_release(x_9, 1); - x_562 = x_9; +lean_object* x_558; lean_object* x_559; lean_object* x_560; uint8_t x_561; +if (lean_is_exclusive(x_10)) { + lean_ctor_release(x_10, 0); + lean_ctor_release(x_10, 1); + x_558 = x_10; } else { - lean_dec_ref(x_9); - x_562 = lean_box(0); + lean_dec_ref(x_10); + x_558 = lean_box(0); } -if (lean_is_exclusive(x_79)) { - lean_ctor_release(x_79, 0); - lean_ctor_release(x_79, 1); - x_563 = x_79; +if (lean_is_exclusive(x_80)) { + lean_ctor_release(x_80, 0); + lean_ctor_release(x_80, 1); + x_559 = x_80; } else { - lean_dec_ref(x_79); - x_563 = lean_box(0); + lean_dec_ref(x_80); + x_559 = lean_box(0); } -x_564 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__4; -x_565 = lean_string_dec_eq(x_540, x_564); -if (x_565 == 0) +x_560 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__4; +x_561 = lean_string_dec_eq(x_536, x_560); +if (x_561 == 0) { -if (lean_obj_tag(x_84) == 0) +lean_dec(x_3); +if (lean_obj_tag(x_85) == 0) { -lean_object* x_566; lean_object* x_567; lean_object* x_568; lean_object* x_569; -lean_dec(x_562); -lean_dec(x_537); -if (lean_is_scalar(x_545)) { - x_566 = lean_alloc_ctor(1, 2, 8); +lean_object* x_562; lean_object* x_563; lean_object* x_564; lean_object* x_565; +lean_dec(x_558); +lean_dec(x_533); +if (lean_is_scalar(x_541)) { + x_562 = lean_alloc_ctor(1, 2, 8); } else { - x_566 = x_545; + x_562 = x_541; } -lean_ctor_set(x_566, 0, x_82); -lean_ctor_set(x_566, 1, x_546); -lean_ctor_set_uint64(x_566, sizeof(void*)*2, x_544); -if (lean_is_scalar(x_542)) { - x_567 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_562, 0, x_83); +lean_ctor_set(x_562, 1, x_542); +lean_ctor_set_uint64(x_562, sizeof(void*)*2, x_540); +if (lean_is_scalar(x_538)) { + x_563 = lean_alloc_ctor(1, 2, 8); } else { - x_567 = x_542; + x_563 = x_538; } -lean_ctor_set(x_567, 0, x_566); -lean_ctor_set(x_567, 1, x_540); -lean_ctor_set_uint64(x_567, sizeof(void*)*2, x_541); -if (lean_is_scalar(x_563)) { - x_568 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_563, 0, x_562); +lean_ctor_set(x_563, 1, x_536); +lean_ctor_set_uint64(x_563, sizeof(void*)*2, x_537); +if (lean_is_scalar(x_559)) { + x_564 = lean_alloc_ctor(1, 2, 8); } else { - x_568 = x_563; + x_564 = x_559; } -lean_ctor_set(x_568, 0, x_567); -lean_ctor_set(x_568, 1, x_538); -lean_ctor_set_uint64(x_568, sizeof(void*)*2, x_539); -x_569 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_568, x_5, x_6, x_83); +lean_ctor_set(x_564, 0, x_563); +lean_ctor_set(x_564, 1, x_534); +lean_ctor_set_uint64(x_564, sizeof(void*)*2, x_535); +x_565 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_564, x_6, x_7, x_84); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_568); -return x_569; +lean_dec(x_564); +return x_565; } else { -lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; +lean_object* x_566; lean_object* x_567; lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_dec(x_1); -if (lean_is_scalar(x_545)) { - x_570 = lean_alloc_ctor(1, 2, 8); -} else { - x_570 = x_545; -} -lean_ctor_set(x_570, 0, x_82); -lean_ctor_set(x_570, 1, x_546); -lean_ctor_set_uint64(x_570, sizeof(void*)*2, x_544); -if (lean_is_scalar(x_542)) { - x_571 = lean_alloc_ctor(1, 2, 8); -} else { - x_571 = x_542; -} -lean_ctor_set(x_571, 0, x_570); -lean_ctor_set(x_571, 1, x_540); -lean_ctor_set_uint64(x_571, sizeof(void*)*2, x_541); -if (lean_is_scalar(x_563)) { - x_572 = lean_alloc_ctor(1, 2, 8); -} else { - x_572 = x_563; -} -lean_ctor_set(x_572, 0, x_571); -lean_ctor_set(x_572, 1, x_538); -lean_ctor_set_uint64(x_572, sizeof(void*)*2, x_539); -x_573 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_573, 0, x_572); -lean_ctor_set(x_573, 1, x_537); -if (lean_is_scalar(x_562)) { - x_574 = lean_alloc_ctor(1, 2, 0); -} else { - x_574 = x_562; -} -lean_ctor_set(x_574, 0, x_573); -lean_ctor_set(x_574, 1, x_84); -x_575 = lean_box(0); -x_576 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_574, x_575); -x_577 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_576, x_575); -x_578 = l_Lean_MessageData_ofList(x_577); -lean_dec(x_577); -x_579 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_580 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_580, 0, x_579); -lean_ctor_set(x_580, 1, x_578); -x_581 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_582 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_582, 0, x_580); -lean_ctor_set(x_582, 1, x_581); -x_583 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_582, x_5, x_6, x_83); -lean_dec(x_6); -x_584 = lean_ctor_get(x_583, 0); -lean_inc(x_584); -x_585 = lean_ctor_get(x_583, 1); -lean_inc(x_585); -if (lean_is_exclusive(x_583)) { - lean_ctor_release(x_583, 0); - lean_ctor_release(x_583, 1); - x_586 = x_583; +if (lean_is_scalar(x_541)) { + x_566 = lean_alloc_ctor(1, 2, 8); +} else { + x_566 = x_541; +} +lean_ctor_set(x_566, 0, x_83); +lean_ctor_set(x_566, 1, x_542); +lean_ctor_set_uint64(x_566, sizeof(void*)*2, x_540); +if (lean_is_scalar(x_538)) { + x_567 = lean_alloc_ctor(1, 2, 8); +} else { + x_567 = x_538; +} +lean_ctor_set(x_567, 0, x_566); +lean_ctor_set(x_567, 1, x_536); +lean_ctor_set_uint64(x_567, sizeof(void*)*2, x_537); +if (lean_is_scalar(x_559)) { + x_568 = lean_alloc_ctor(1, 2, 8); +} else { + x_568 = x_559; +} +lean_ctor_set(x_568, 0, x_567); +lean_ctor_set(x_568, 1, x_534); +lean_ctor_set_uint64(x_568, sizeof(void*)*2, x_535); +x_569 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_569, 0, x_568); +lean_ctor_set(x_569, 1, x_533); +if (lean_is_scalar(x_558)) { + x_570 = lean_alloc_ctor(1, 2, 0); +} else { + x_570 = x_558; +} +lean_ctor_set(x_570, 0, x_569); +lean_ctor_set(x_570, 1, x_85); +x_571 = lean_box(0); +x_572 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_570, x_571); +x_573 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_572, x_571); +x_574 = l_Lean_MessageData_ofList(x_573); +lean_dec(x_573); +x_575 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_576 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_576, 0, x_575); +lean_ctor_set(x_576, 1, x_574); +x_577 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_578 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_578, 0, x_576); +lean_ctor_set(x_578, 1, x_577); +x_579 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_578, x_6, x_7, x_84); +lean_dec(x_7); +x_580 = lean_ctor_get(x_579, 0); +lean_inc(x_580); +x_581 = lean_ctor_get(x_579, 1); +lean_inc(x_581); +if (lean_is_exclusive(x_579)) { + lean_ctor_release(x_579, 0); + lean_ctor_release(x_579, 1); + x_582 = x_579; } else { - lean_dec_ref(x_583); - x_586 = lean_box(0); + lean_dec_ref(x_579); + x_582 = lean_box(0); } -if (lean_is_scalar(x_586)) { - x_587 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_582)) { + x_583 = lean_alloc_ctor(1, 2, 0); } else { - x_587 = x_586; + x_583 = x_582; } -lean_ctor_set(x_587, 0, x_584); -lean_ctor_set(x_587, 1, x_585); -return x_587; +lean_ctor_set(x_583, 0, x_580); +lean_ctor_set(x_583, 1, x_581); +return x_583; } } else { -lean_object* x_588; uint8_t x_589; -lean_dec(x_540); -x_588 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__7; -x_589 = lean_string_dec_eq(x_538, x_588); -if (x_589 == 0) +lean_object* x_584; uint8_t x_585; +lean_dec(x_536); +x_584 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__7; +x_585 = lean_string_dec_eq(x_534, x_584); +if (x_585 == 0) { -lean_object* x_590; uint8_t x_591; -x_590 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__8; -x_591 = lean_string_dec_eq(x_538, x_590); -if (x_591 == 0) +lean_object* x_586; uint8_t x_587; +lean_dec(x_3); +x_586 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__8; +x_587 = lean_string_dec_eq(x_534, x_586); +if (x_587 == 0) { -if (lean_obj_tag(x_84) == 0) +if (lean_obj_tag(x_85) == 0) { -lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; -lean_dec(x_562); -lean_dec(x_537); -if (lean_is_scalar(x_545)) { - x_592 = lean_alloc_ctor(1, 2, 8); +lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; +lean_dec(x_558); +lean_dec(x_533); +if (lean_is_scalar(x_541)) { + x_588 = lean_alloc_ctor(1, 2, 8); } else { - x_592 = x_545; + x_588 = x_541; } -lean_ctor_set(x_592, 0, x_82); -lean_ctor_set(x_592, 1, x_546); -lean_ctor_set_uint64(x_592, sizeof(void*)*2, x_544); -if (lean_is_scalar(x_542)) { - x_593 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_588, 0, x_83); +lean_ctor_set(x_588, 1, x_542); +lean_ctor_set_uint64(x_588, sizeof(void*)*2, x_540); +if (lean_is_scalar(x_538)) { + x_589 = lean_alloc_ctor(1, 2, 8); } else { - x_593 = x_542; + x_589 = x_538; } -lean_ctor_set(x_593, 0, x_592); -lean_ctor_set(x_593, 1, x_564); -lean_ctor_set_uint64(x_593, sizeof(void*)*2, x_541); -if (lean_is_scalar(x_563)) { - x_594 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_589, 0, x_588); +lean_ctor_set(x_589, 1, x_560); +lean_ctor_set_uint64(x_589, sizeof(void*)*2, x_537); +if (lean_is_scalar(x_559)) { + x_590 = lean_alloc_ctor(1, 2, 8); } else { - x_594 = x_563; + x_590 = x_559; } -lean_ctor_set(x_594, 0, x_593); -lean_ctor_set(x_594, 1, x_538); -lean_ctor_set_uint64(x_594, sizeof(void*)*2, x_539); -x_595 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_594, x_5, x_6, x_83); +lean_ctor_set(x_590, 0, x_589); +lean_ctor_set(x_590, 1, x_534); +lean_ctor_set_uint64(x_590, sizeof(void*)*2, x_535); +x_591 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_590, x_6, x_7, x_84); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_594); -return x_595; +lean_dec(x_590); +return x_591; } else { -lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; +lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_dec(x_1); -if (lean_is_scalar(x_545)) { - x_596 = lean_alloc_ctor(1, 2, 8); -} else { - x_596 = x_545; -} -lean_ctor_set(x_596, 0, x_82); -lean_ctor_set(x_596, 1, x_546); -lean_ctor_set_uint64(x_596, sizeof(void*)*2, x_544); -if (lean_is_scalar(x_542)) { - x_597 = lean_alloc_ctor(1, 2, 8); -} else { - x_597 = x_542; -} -lean_ctor_set(x_597, 0, x_596); -lean_ctor_set(x_597, 1, x_564); -lean_ctor_set_uint64(x_597, sizeof(void*)*2, x_541); -if (lean_is_scalar(x_563)) { - x_598 = lean_alloc_ctor(1, 2, 8); -} else { - x_598 = x_563; -} -lean_ctor_set(x_598, 0, x_597); -lean_ctor_set(x_598, 1, x_538); -lean_ctor_set_uint64(x_598, sizeof(void*)*2, x_539); -x_599 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_599, 0, x_598); -lean_ctor_set(x_599, 1, x_537); -if (lean_is_scalar(x_562)) { - x_600 = lean_alloc_ctor(1, 2, 0); -} else { - x_600 = x_562; -} -lean_ctor_set(x_600, 0, x_599); -lean_ctor_set(x_600, 1, x_84); -x_601 = lean_box(0); -x_602 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_600, x_601); -x_603 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_602, x_601); -x_604 = l_Lean_MessageData_ofList(x_603); -lean_dec(x_603); -x_605 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_606 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_606, 0, x_605); -lean_ctor_set(x_606, 1, x_604); -x_607 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_608 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_608, 0, x_606); -lean_ctor_set(x_608, 1, x_607); -x_609 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_608, x_5, x_6, x_83); -lean_dec(x_6); -x_610 = lean_ctor_get(x_609, 0); -lean_inc(x_610); -x_611 = lean_ctor_get(x_609, 1); -lean_inc(x_611); -if (lean_is_exclusive(x_609)) { - lean_ctor_release(x_609, 0); - lean_ctor_release(x_609, 1); - x_612 = x_609; +if (lean_is_scalar(x_541)) { + x_592 = lean_alloc_ctor(1, 2, 8); +} else { + x_592 = x_541; +} +lean_ctor_set(x_592, 0, x_83); +lean_ctor_set(x_592, 1, x_542); +lean_ctor_set_uint64(x_592, sizeof(void*)*2, x_540); +if (lean_is_scalar(x_538)) { + x_593 = lean_alloc_ctor(1, 2, 8); +} else { + x_593 = x_538; +} +lean_ctor_set(x_593, 0, x_592); +lean_ctor_set(x_593, 1, x_560); +lean_ctor_set_uint64(x_593, sizeof(void*)*2, x_537); +if (lean_is_scalar(x_559)) { + x_594 = lean_alloc_ctor(1, 2, 8); } else { - lean_dec_ref(x_609); - x_612 = lean_box(0); + x_594 = x_559; } -if (lean_is_scalar(x_612)) { - x_613 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_594, 0, x_593); +lean_ctor_set(x_594, 1, x_534); +lean_ctor_set_uint64(x_594, sizeof(void*)*2, x_535); +x_595 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_595, 0, x_594); +lean_ctor_set(x_595, 1, x_533); +if (lean_is_scalar(x_558)) { + x_596 = lean_alloc_ctor(1, 2, 0); +} else { + x_596 = x_558; +} +lean_ctor_set(x_596, 0, x_595); +lean_ctor_set(x_596, 1, x_85); +x_597 = lean_box(0); +x_598 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_596, x_597); +x_599 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_598, x_597); +x_600 = l_Lean_MessageData_ofList(x_599); +lean_dec(x_599); +x_601 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_602 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_602, 0, x_601); +lean_ctor_set(x_602, 1, x_600); +x_603 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_604 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_604, 0, x_602); +lean_ctor_set(x_604, 1, x_603); +x_605 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_604, x_6, x_7, x_84); +lean_dec(x_7); +x_606 = lean_ctor_get(x_605, 0); +lean_inc(x_606); +x_607 = lean_ctor_get(x_605, 1); +lean_inc(x_607); +if (lean_is_exclusive(x_605)) { + lean_ctor_release(x_605, 0); + lean_ctor_release(x_605, 1); + x_608 = x_605; +} else { + lean_dec_ref(x_605); + x_608 = lean_box(0); +} +if (lean_is_scalar(x_608)) { + x_609 = lean_alloc_ctor(1, 2, 0); } else { - x_613 = x_612; + x_609 = x_608; } -lean_ctor_set(x_613, 0, x_610); -lean_ctor_set(x_613, 1, x_611); -return x_613; +lean_ctor_set(x_609, 0, x_606); +lean_ctor_set(x_609, 1, x_607); +return x_609; } } else { +lean_dec(x_534); +if (lean_obj_tag(x_85) == 0) +{ +lean_object* x_610; lean_object* x_611; +lean_dec(x_559); +lean_dec(x_558); +lean_dec(x_541); lean_dec(x_538); -if (lean_obj_tag(x_84) == 0) -{ -lean_object* x_614; lean_object* x_615; -lean_dec(x_563); -lean_dec(x_562); -lean_dec(x_545); -lean_dec(x_542); -lean_dec(x_537); -x_614 = l_Lean_strLitKind; -x_615 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_614, x_5, x_6, x_83); +lean_dec(x_533); +x_610 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__10; +x_611 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_610, x_6, x_7, x_84); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -return x_615; +return x_611; } else { -lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; lean_object* x_633; +lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_dec(x_1); -if (lean_is_scalar(x_545)) { - x_616 = lean_alloc_ctor(1, 2, 8); -} else { - x_616 = x_545; -} -lean_ctor_set(x_616, 0, x_82); -lean_ctor_set(x_616, 1, x_546); -lean_ctor_set_uint64(x_616, sizeof(void*)*2, x_544); -if (lean_is_scalar(x_542)) { - x_617 = lean_alloc_ctor(1, 2, 8); -} else { - x_617 = x_542; -} -lean_ctor_set(x_617, 0, x_616); -lean_ctor_set(x_617, 1, x_564); -lean_ctor_set_uint64(x_617, sizeof(void*)*2, x_541); -if (lean_is_scalar(x_563)) { - x_618 = lean_alloc_ctor(1, 2, 8); -} else { - x_618 = x_563; -} -lean_ctor_set(x_618, 0, x_617); -lean_ctor_set(x_618, 1, x_590); -lean_ctor_set_uint64(x_618, sizeof(void*)*2, x_539); -x_619 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_619, 0, x_618); -lean_ctor_set(x_619, 1, x_537); -if (lean_is_scalar(x_562)) { - x_620 = lean_alloc_ctor(1, 2, 0); -} else { - x_620 = x_562; -} -lean_ctor_set(x_620, 0, x_619); -lean_ctor_set(x_620, 1, x_84); -x_621 = lean_box(0); -x_622 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_620, x_621); -x_623 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_622, x_621); -x_624 = l_Lean_MessageData_ofList(x_623); -lean_dec(x_623); -x_625 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_626 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_626, 0, x_625); -lean_ctor_set(x_626, 1, x_624); -x_627 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_628 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_628, 0, x_626); -lean_ctor_set(x_628, 1, x_627); -x_629 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_628, x_5, x_6, x_83); -lean_dec(x_6); -x_630 = lean_ctor_get(x_629, 0); -lean_inc(x_630); -x_631 = lean_ctor_get(x_629, 1); -lean_inc(x_631); -if (lean_is_exclusive(x_629)) { - lean_ctor_release(x_629, 0); - lean_ctor_release(x_629, 1); - x_632 = x_629; +if (lean_is_scalar(x_541)) { + x_612 = lean_alloc_ctor(1, 2, 8); +} else { + x_612 = x_541; +} +lean_ctor_set(x_612, 0, x_83); +lean_ctor_set(x_612, 1, x_542); +lean_ctor_set_uint64(x_612, sizeof(void*)*2, x_540); +if (lean_is_scalar(x_538)) { + x_613 = lean_alloc_ctor(1, 2, 8); +} else { + x_613 = x_538; +} +lean_ctor_set(x_613, 0, x_612); +lean_ctor_set(x_613, 1, x_560); +lean_ctor_set_uint64(x_613, sizeof(void*)*2, x_537); +if (lean_is_scalar(x_559)) { + x_614 = lean_alloc_ctor(1, 2, 8); +} else { + x_614 = x_559; +} +lean_ctor_set(x_614, 0, x_613); +lean_ctor_set(x_614, 1, x_586); +lean_ctor_set_uint64(x_614, sizeof(void*)*2, x_535); +x_615 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_615, 0, x_614); +lean_ctor_set(x_615, 1, x_533); +if (lean_is_scalar(x_558)) { + x_616 = lean_alloc_ctor(1, 2, 0); +} else { + x_616 = x_558; +} +lean_ctor_set(x_616, 0, x_615); +lean_ctor_set(x_616, 1, x_85); +x_617 = lean_box(0); +x_618 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_616, x_617); +x_619 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_618, x_617); +x_620 = l_Lean_MessageData_ofList(x_619); +lean_dec(x_619); +x_621 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_622 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_622, 0, x_621); +lean_ctor_set(x_622, 1, x_620); +x_623 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_624 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_624, 0, x_622); +lean_ctor_set(x_624, 1, x_623); +x_625 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_624, x_6, x_7, x_84); +lean_dec(x_7); +x_626 = lean_ctor_get(x_625, 0); +lean_inc(x_626); +x_627 = lean_ctor_get(x_625, 1); +lean_inc(x_627); +if (lean_is_exclusive(x_625)) { + lean_ctor_release(x_625, 0); + lean_ctor_release(x_625, 1); + x_628 = x_625; } else { - lean_dec_ref(x_629); - x_632 = lean_box(0); + lean_dec_ref(x_625); + x_628 = lean_box(0); } -if (lean_is_scalar(x_632)) { - x_633 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_628)) { + x_629 = lean_alloc_ctor(1, 2, 0); } else { - x_633 = x_632; + x_629 = x_628; } -lean_ctor_set(x_633, 0, x_630); -lean_ctor_set(x_633, 1, x_631); -return x_633; +lean_ctor_set(x_629, 0, x_626); +lean_ctor_set(x_629, 1, x_627); +return x_629; } } } else { +lean_dec(x_534); +if (lean_obj_tag(x_85) == 0) +{ +lean_object* x_630; +lean_dec(x_559); +lean_dec(x_558); +lean_dec(x_541); lean_dec(x_538); -if (lean_obj_tag(x_84) == 0) -{ -lean_object* x_634; lean_object* x_635; -lean_dec(x_563); -lean_dec(x_562); -lean_dec(x_545); -lean_dec(x_542); -lean_dec(x_537); -x_634 = l_Lean_identKind; -x_635 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_634, x_5, x_6, x_83); +lean_dec(x_533); +x_630 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_3, x_6, x_7, x_84); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -return x_635; +lean_dec(x_3); +return x_630; } else { -lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; +lean_object* x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; +lean_dec(x_3); lean_dec(x_1); -if (lean_is_scalar(x_545)) { - x_636 = lean_alloc_ctor(1, 2, 8); -} else { - x_636 = x_545; -} -lean_ctor_set(x_636, 0, x_82); -lean_ctor_set(x_636, 1, x_546); -lean_ctor_set_uint64(x_636, sizeof(void*)*2, x_544); -if (lean_is_scalar(x_542)) { - x_637 = lean_alloc_ctor(1, 2, 8); -} else { - x_637 = x_542; -} -lean_ctor_set(x_637, 0, x_636); -lean_ctor_set(x_637, 1, x_564); -lean_ctor_set_uint64(x_637, sizeof(void*)*2, x_541); -if (lean_is_scalar(x_563)) { - x_638 = lean_alloc_ctor(1, 2, 8); -} else { - x_638 = x_563; -} -lean_ctor_set(x_638, 0, x_637); -lean_ctor_set(x_638, 1, x_588); -lean_ctor_set_uint64(x_638, sizeof(void*)*2, x_539); -x_639 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_639, 0, x_638); -lean_ctor_set(x_639, 1, x_537); -if (lean_is_scalar(x_562)) { - x_640 = lean_alloc_ctor(1, 2, 0); -} else { - x_640 = x_562; -} -lean_ctor_set(x_640, 0, x_639); -lean_ctor_set(x_640, 1, x_84); -x_641 = lean_box(0); -x_642 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_640, x_641); -x_643 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_642, x_641); -x_644 = l_Lean_MessageData_ofList(x_643); -lean_dec(x_643); -x_645 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_646 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_646, 0, x_645); -lean_ctor_set(x_646, 1, x_644); -x_647 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_648 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_648, 0, x_646); -lean_ctor_set(x_648, 1, x_647); -x_649 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_648, x_5, x_6, x_83); -lean_dec(x_6); -x_650 = lean_ctor_get(x_649, 0); -lean_inc(x_650); -x_651 = lean_ctor_get(x_649, 1); -lean_inc(x_651); -if (lean_is_exclusive(x_649)) { - lean_ctor_release(x_649, 0); - lean_ctor_release(x_649, 1); - x_652 = x_649; +if (lean_is_scalar(x_541)) { + x_631 = lean_alloc_ctor(1, 2, 8); +} else { + x_631 = x_541; +} +lean_ctor_set(x_631, 0, x_83); +lean_ctor_set(x_631, 1, x_542); +lean_ctor_set_uint64(x_631, sizeof(void*)*2, x_540); +if (lean_is_scalar(x_538)) { + x_632 = lean_alloc_ctor(1, 2, 8); +} else { + x_632 = x_538; +} +lean_ctor_set(x_632, 0, x_631); +lean_ctor_set(x_632, 1, x_560); +lean_ctor_set_uint64(x_632, sizeof(void*)*2, x_537); +if (lean_is_scalar(x_559)) { + x_633 = lean_alloc_ctor(1, 2, 8); +} else { + x_633 = x_559; +} +lean_ctor_set(x_633, 0, x_632); +lean_ctor_set(x_633, 1, x_584); +lean_ctor_set_uint64(x_633, sizeof(void*)*2, x_535); +x_634 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_634, 0, x_633); +lean_ctor_set(x_634, 1, x_533); +if (lean_is_scalar(x_558)) { + x_635 = lean_alloc_ctor(1, 2, 0); +} else { + x_635 = x_558; +} +lean_ctor_set(x_635, 0, x_634); +lean_ctor_set(x_635, 1, x_85); +x_636 = lean_box(0); +x_637 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_635, x_636); +x_638 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_637, x_636); +x_639 = l_Lean_MessageData_ofList(x_638); +lean_dec(x_638); +x_640 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_641 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_641, 0, x_640); +lean_ctor_set(x_641, 1, x_639); +x_642 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_643 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_643, 0, x_641); +lean_ctor_set(x_643, 1, x_642); +x_644 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_643, x_6, x_7, x_84); +lean_dec(x_7); +x_645 = lean_ctor_get(x_644, 0); +lean_inc(x_645); +x_646 = lean_ctor_get(x_644, 1); +lean_inc(x_646); +if (lean_is_exclusive(x_644)) { + lean_ctor_release(x_644, 0); + lean_ctor_release(x_644, 1); + x_647 = x_644; } else { - lean_dec_ref(x_649); - x_652 = lean_box(0); + lean_dec_ref(x_644); + x_647 = lean_box(0); } -if (lean_is_scalar(x_652)) { - x_653 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_647)) { + x_648 = lean_alloc_ctor(1, 2, 0); } else { - x_653 = x_652; + x_648 = x_647; } -lean_ctor_set(x_653, 0, x_650); -lean_ctor_set(x_653, 1, x_651); -return x_653; +lean_ctor_set(x_648, 0, x_645); +lean_ctor_set(x_648, 1, x_646); +return x_648; } } } @@ -4811,375 +4850,380 @@ return x_653; } case 1: { -lean_object* x_654; -x_654 = lean_ctor_get(x_82, 0); -lean_inc(x_654); -if (lean_obj_tag(x_654) == 0) +lean_object* x_649; +x_649 = lean_ctor_get(x_83, 0); +lean_inc(x_649); +if (lean_obj_tag(x_649) == 0) +{ +lean_object* x_650; lean_object* x_651; uint8_t x_652; +x_650 = lean_ctor_get(x_9, 1); +lean_inc(x_650); +lean_dec(x_9); +x_651 = lean_ctor_get(x_10, 1); +lean_inc(x_651); +x_652 = !lean_is_exclusive(x_79); +if (x_652 == 0) { -lean_object* x_655; lean_object* x_656; uint8_t x_657; -x_655 = lean_ctor_get(x_8, 1); +lean_object* x_653; lean_object* x_654; lean_object* x_655; uint8_t x_656; +x_653 = lean_ctor_get(x_79, 1); +x_654 = lean_ctor_get(x_79, 0); +lean_dec(x_654); +x_655 = lean_ctor_get(x_80, 1); lean_inc(x_655); -lean_dec(x_8); -x_656 = lean_ctor_get(x_9, 1); -lean_inc(x_656); -x_657 = !lean_is_exclusive(x_78); -if (x_657 == 0) -{ -lean_object* x_658; lean_object* x_659; lean_object* x_660; uint8_t x_661; -x_658 = lean_ctor_get(x_78, 1); -x_659 = lean_ctor_get(x_78, 0); +x_656 = !lean_is_exclusive(x_81); +if (x_656 == 0) +{ +uint64_t x_657; lean_object* x_658; lean_object* x_659; uint8_t x_660; +x_657 = lean_ctor_get_uint64(x_80, sizeof(void*)*2); +x_658 = lean_ctor_get(x_81, 1); +x_659 = lean_ctor_get(x_81, 0); lean_dec(x_659); -x_660 = lean_ctor_get(x_79, 1); -lean_inc(x_660); -x_661 = !lean_is_exclusive(x_80); -if (x_661 == 0) -{ -uint64_t x_662; lean_object* x_663; lean_object* x_664; uint8_t x_665; -x_662 = lean_ctor_get_uint64(x_79, sizeof(void*)*2); -x_663 = lean_ctor_get(x_80, 1); -x_664 = lean_ctor_get(x_80, 0); +x_660 = !lean_is_exclusive(x_82); +if (x_660 == 0) +{ +lean_object* x_661; lean_object* x_662; uint8_t x_663; +x_661 = lean_ctor_get(x_82, 1); +x_662 = lean_ctor_get(x_82, 0); +lean_dec(x_662); +x_663 = !lean_is_exclusive(x_83); +if (x_663 == 0) +{ +lean_object* x_664; lean_object* x_665; lean_object* x_666; uint8_t x_667; +x_664 = lean_ctor_get(x_83, 1); +x_665 = lean_ctor_get(x_83, 0); +lean_dec(x_665); +x_666 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__1; +x_667 = lean_string_dec_eq(x_664, x_666); lean_dec(x_664); -x_665 = !lean_is_exclusive(x_81); -if (x_665 == 0) -{ -lean_object* x_666; lean_object* x_667; uint8_t x_668; -x_666 = lean_ctor_get(x_81, 1); -x_667 = lean_ctor_get(x_81, 0); -lean_dec(x_667); -x_668 = !lean_is_exclusive(x_82); -if (x_668 == 0) -{ -lean_object* x_669; lean_object* x_670; lean_object* x_671; uint8_t x_672; -x_669 = lean_ctor_get(x_82, 1); -x_670 = lean_ctor_get(x_82, 0); -lean_dec(x_670); -x_671 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__1; -x_672 = lean_string_dec_eq(x_669, x_671); -lean_dec(x_669); -if (x_672 == 0) +if (x_667 == 0) { +lean_free_object(x_83); lean_free_object(x_82); +lean_dec(x_661); lean_free_object(x_81); -lean_dec(x_666); -lean_free_object(x_80); -lean_dec(x_663); -lean_dec(x_660); -lean_free_object(x_78); lean_dec(x_658); -if (lean_obj_tag(x_656) == 0) +lean_dec(x_655); +lean_free_object(x_79); +lean_dec(x_653); +lean_dec(x_3); +if (lean_obj_tag(x_651) == 0) { -lean_object* x_673; -lean_dec(x_9); -x_673 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_79, x_5, x_6, x_655); +lean_object* x_668; +lean_dec(x_10); +x_668 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_80, x_6, x_7, x_650); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_79); -return x_673; +lean_dec(x_80); +return x_668; } else { -lean_object* x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; lean_object* x_678; lean_object* x_679; lean_object* x_680; lean_object* x_681; lean_object* x_682; uint8_t x_683; -lean_dec(x_656); -lean_dec(x_79); +lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; lean_object* x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; uint8_t x_678; +lean_dec(x_651); +lean_dec(x_80); lean_dec(x_1); -x_674 = lean_box(0); -x_675 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_9, x_674); -x_676 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_675, x_674); -x_677 = l_Lean_MessageData_ofList(x_676); -lean_dec(x_676); -x_678 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_679 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_679, 0, x_678); -lean_ctor_set(x_679, 1, x_677); -x_680 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_681 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_681, 0, x_679); -lean_ctor_set(x_681, 1, x_680); -x_682 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_681, x_5, x_6, x_655); -lean_dec(x_6); -x_683 = !lean_is_exclusive(x_682); -if (x_683 == 0) +x_669 = lean_box(0); +x_670 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_10, x_669); +x_671 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_670, x_669); +x_672 = l_Lean_MessageData_ofList(x_671); +lean_dec(x_671); +x_673 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_674 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_674, 0, x_673); +lean_ctor_set(x_674, 1, x_672); +x_675 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_676 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_676, 0, x_674); +lean_ctor_set(x_676, 1, x_675); +x_677 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_676, x_6, x_7, x_650); +lean_dec(x_7); +x_678 = !lean_is_exclusive(x_677); +if (x_678 == 0) { -return x_682; +return x_677; } else { -lean_object* x_684; lean_object* x_685; lean_object* x_686; -x_684 = lean_ctor_get(x_682, 0); -x_685 = lean_ctor_get(x_682, 1); -lean_inc(x_685); -lean_inc(x_684); -lean_dec(x_682); -x_686 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_686, 0, x_684); -lean_ctor_set(x_686, 1, x_685); -return x_686; +lean_object* x_679; lean_object* x_680; lean_object* x_681; +x_679 = lean_ctor_get(x_677, 0); +x_680 = lean_ctor_get(x_677, 1); +lean_inc(x_680); +lean_inc(x_679); +lean_dec(x_677); +x_681 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_681, 0, x_679); +lean_ctor_set(x_681, 1, x_680); +return x_681; } } } else { -uint8_t x_687; -x_687 = !lean_is_exclusive(x_9); -if (x_687 == 0) +uint8_t x_682; +x_682 = !lean_is_exclusive(x_10); +if (x_682 == 0) { -lean_object* x_688; lean_object* x_689; uint8_t x_690; -x_688 = lean_ctor_get(x_9, 1); -lean_dec(x_688); -x_689 = lean_ctor_get(x_9, 0); -lean_dec(x_689); -x_690 = !lean_is_exclusive(x_79); -if (x_690 == 0) +lean_object* x_683; lean_object* x_684; uint8_t x_685; +x_683 = lean_ctor_get(x_10, 1); +lean_dec(x_683); +x_684 = lean_ctor_get(x_10, 0); +lean_dec(x_684); +x_685 = !lean_is_exclusive(x_80); +if (x_685 == 0) { -lean_object* x_691; lean_object* x_692; lean_object* x_693; uint8_t x_694; -x_691 = lean_ctor_get(x_79, 1); -lean_dec(x_691); -x_692 = lean_ctor_get(x_79, 0); -lean_dec(x_692); -x_693 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__4; -x_694 = lean_string_dec_eq(x_666, x_693); -if (x_694 == 0) +lean_object* x_686; lean_object* x_687; lean_object* x_688; uint8_t x_689; +x_686 = lean_ctor_get(x_80, 1); +lean_dec(x_686); +x_687 = lean_ctor_get(x_80, 0); +lean_dec(x_687); +x_688 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__4; +x_689 = lean_string_dec_eq(x_661, x_688); +if (x_689 == 0) { -if (lean_obj_tag(x_656) == 0) +lean_dec(x_3); +if (lean_obj_tag(x_651) == 0) { -lean_object* x_695; -lean_free_object(x_9); -lean_free_object(x_78); -lean_dec(x_658); -lean_ctor_set(x_82, 1, x_671); -x_695 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_79, x_5, x_6, x_655); +lean_object* x_690; +lean_free_object(x_10); +lean_free_object(x_79); +lean_dec(x_653); +lean_ctor_set(x_83, 1, x_666); +x_690 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_80, x_6, x_7, x_650); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_79); -return x_695; +lean_dec(x_80); +return x_690; } else { -lean_object* x_696; lean_object* x_697; lean_object* x_698; lean_object* x_699; lean_object* x_700; lean_object* x_701; lean_object* x_702; lean_object* x_703; lean_object* x_704; uint8_t x_705; +lean_object* x_691; lean_object* x_692; lean_object* x_693; lean_object* x_694; lean_object* x_695; lean_object* x_696; lean_object* x_697; lean_object* x_698; lean_object* x_699; uint8_t x_700; lean_dec(x_1); -lean_ctor_set(x_82, 1, x_671); -x_696 = lean_box(0); -x_697 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_9, x_696); -x_698 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_697, x_696); -x_699 = l_Lean_MessageData_ofList(x_698); -lean_dec(x_698); -x_700 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_701 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_701, 0, x_700); -lean_ctor_set(x_701, 1, x_699); -x_702 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_703 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_703, 0, x_701); -lean_ctor_set(x_703, 1, x_702); -x_704 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_703, x_5, x_6, x_655); -lean_dec(x_6); -x_705 = !lean_is_exclusive(x_704); -if (x_705 == 0) +lean_ctor_set(x_83, 1, x_666); +x_691 = lean_box(0); +x_692 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_10, x_691); +x_693 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_692, x_691); +x_694 = l_Lean_MessageData_ofList(x_693); +lean_dec(x_693); +x_695 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_696 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_696, 0, x_695); +lean_ctor_set(x_696, 1, x_694); +x_697 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_698 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_698, 0, x_696); +lean_ctor_set(x_698, 1, x_697); +x_699 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_698, x_6, x_7, x_650); +lean_dec(x_7); +x_700 = !lean_is_exclusive(x_699); +if (x_700 == 0) { -return x_704; +return x_699; } else { -lean_object* x_706; lean_object* x_707; lean_object* x_708; -x_706 = lean_ctor_get(x_704, 0); -x_707 = lean_ctor_get(x_704, 1); -lean_inc(x_707); -lean_inc(x_706); -lean_dec(x_704); -x_708 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_708, 0, x_706); -lean_ctor_set(x_708, 1, x_707); -return x_708; +lean_object* x_701; lean_object* x_702; lean_object* x_703; +x_701 = lean_ctor_get(x_699, 0); +x_702 = lean_ctor_get(x_699, 1); +lean_inc(x_702); +lean_inc(x_701); +lean_dec(x_699); +x_703 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_703, 0, x_701); +lean_ctor_set(x_703, 1, x_702); +return x_703; } } } else { -lean_object* x_709; uint8_t x_710; -lean_dec(x_666); -x_709 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__9; -x_710 = lean_string_dec_eq(x_663, x_709); -if (x_710 == 0) +lean_object* x_704; uint8_t x_705; +lean_dec(x_661); +x_704 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__11; +x_705 = lean_string_dec_eq(x_658, x_704); +if (x_705 == 0) { -if (lean_obj_tag(x_656) == 0) +lean_dec(x_3); +if (lean_obj_tag(x_651) == 0) { -lean_object* x_711; -lean_free_object(x_9); -lean_free_object(x_78); -lean_dec(x_658); -lean_ctor_set(x_82, 1, x_671); -lean_ctor_set(x_81, 1, x_693); -x_711 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_79, x_5, x_6, x_655); +lean_object* x_706; +lean_free_object(x_10); +lean_free_object(x_79); +lean_dec(x_653); +lean_ctor_set(x_83, 1, x_666); +lean_ctor_set(x_82, 1, x_688); +x_706 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_80, x_6, x_7, x_650); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_79); -return x_711; +lean_dec(x_80); +return x_706; } else { -lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; lean_object* x_718; lean_object* x_719; lean_object* x_720; uint8_t x_721; +lean_object* x_707; lean_object* x_708; lean_object* x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; uint8_t x_716; lean_dec(x_1); -lean_ctor_set(x_82, 1, x_671); -lean_ctor_set(x_81, 1, x_693); -x_712 = lean_box(0); -x_713 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_9, x_712); -x_714 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_713, x_712); -x_715 = l_Lean_MessageData_ofList(x_714); -lean_dec(x_714); -x_716 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_717 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_717, 0, x_716); -lean_ctor_set(x_717, 1, x_715); -x_718 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_719 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_719, 0, x_717); -lean_ctor_set(x_719, 1, x_718); -x_720 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_719, x_5, x_6, x_655); -lean_dec(x_6); -x_721 = !lean_is_exclusive(x_720); -if (x_721 == 0) +lean_ctor_set(x_83, 1, x_666); +lean_ctor_set(x_82, 1, x_688); +x_707 = lean_box(0); +x_708 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_10, x_707); +x_709 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_708, x_707); +x_710 = l_Lean_MessageData_ofList(x_709); +lean_dec(x_709); +x_711 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_712 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_712, 0, x_711); +lean_ctor_set(x_712, 1, x_710); +x_713 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_714 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_714, 0, x_712); +lean_ctor_set(x_714, 1, x_713); +x_715 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_714, x_6, x_7, x_650); +lean_dec(x_7); +x_716 = !lean_is_exclusive(x_715); +if (x_716 == 0) { -return x_720; +return x_715; } else { -lean_object* x_722; lean_object* x_723; lean_object* x_724; -x_722 = lean_ctor_get(x_720, 0); -x_723 = lean_ctor_get(x_720, 1); -lean_inc(x_723); -lean_inc(x_722); -lean_dec(x_720); -x_724 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_724, 0, x_722); -lean_ctor_set(x_724, 1, x_723); -return x_724; +lean_object* x_717; lean_object* x_718; lean_object* x_719; +x_717 = lean_ctor_get(x_715, 0); +x_718 = lean_ctor_get(x_715, 1); +lean_inc(x_718); +lean_inc(x_717); +lean_dec(x_715); +x_719 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_719, 0, x_717); +lean_ctor_set(x_719, 1, x_718); +return x_719; } } } else { -lean_object* x_725; uint8_t x_726; -lean_dec(x_663); -x_725 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__7; -x_726 = lean_string_dec_eq(x_660, x_725); -if (x_726 == 0) +lean_object* x_720; uint8_t x_721; +lean_dec(x_658); +x_720 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__7; +x_721 = lean_string_dec_eq(x_655, x_720); +if (x_721 == 0) { -if (lean_obj_tag(x_656) == 0) +lean_dec(x_3); +if (lean_obj_tag(x_651) == 0) { -lean_object* x_727; -lean_free_object(x_9); -lean_free_object(x_78); -lean_dec(x_658); -lean_ctor_set(x_82, 1, x_671); -lean_ctor_set(x_81, 1, x_693); -lean_ctor_set(x_80, 1, x_709); -x_727 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_79, x_5, x_6, x_655); +lean_object* x_722; +lean_free_object(x_10); +lean_free_object(x_79); +lean_dec(x_653); +lean_ctor_set(x_83, 1, x_666); +lean_ctor_set(x_82, 1, x_688); +lean_ctor_set(x_81, 1, x_704); +x_722 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_80, x_6, x_7, x_650); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_79); -return x_727; +lean_dec(x_80); +return x_722; } else { -lean_object* x_728; lean_object* x_729; lean_object* x_730; lean_object* x_731; lean_object* x_732; lean_object* x_733; lean_object* x_734; lean_object* x_735; lean_object* x_736; uint8_t x_737; +lean_object* x_723; lean_object* x_724; lean_object* x_725; lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; lean_object* x_730; lean_object* x_731; uint8_t x_732; lean_dec(x_1); -lean_ctor_set(x_82, 1, x_671); -lean_ctor_set(x_81, 1, x_693); -lean_ctor_set(x_80, 1, x_709); -x_728 = lean_box(0); -x_729 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_9, x_728); -x_730 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_729, x_728); -x_731 = l_Lean_MessageData_ofList(x_730); -lean_dec(x_730); -x_732 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_733 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_733, 0, x_732); -lean_ctor_set(x_733, 1, x_731); -x_734 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_735 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_735, 0, x_733); -lean_ctor_set(x_735, 1, x_734); -x_736 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_735, x_5, x_6, x_655); -lean_dec(x_6); -x_737 = !lean_is_exclusive(x_736); -if (x_737 == 0) +lean_ctor_set(x_83, 1, x_666); +lean_ctor_set(x_82, 1, x_688); +lean_ctor_set(x_81, 1, x_704); +x_723 = lean_box(0); +x_724 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_10, x_723); +x_725 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_724, x_723); +x_726 = l_Lean_MessageData_ofList(x_725); +lean_dec(x_725); +x_727 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_728 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_728, 0, x_727); +lean_ctor_set(x_728, 1, x_726); +x_729 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_730 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_730, 0, x_728); +lean_ctor_set(x_730, 1, x_729); +x_731 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_730, x_6, x_7, x_650); +lean_dec(x_7); +x_732 = !lean_is_exclusive(x_731); +if (x_732 == 0) { -return x_736; +return x_731; } else { -lean_object* x_738; lean_object* x_739; lean_object* x_740; -x_738 = lean_ctor_get(x_736, 0); -x_739 = lean_ctor_get(x_736, 1); -lean_inc(x_739); -lean_inc(x_738); -lean_dec(x_736); -x_740 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_740, 0, x_738); -lean_ctor_set(x_740, 1, x_739); -return x_740; +lean_object* x_733; lean_object* x_734; lean_object* x_735; +x_733 = lean_ctor_get(x_731, 0); +x_734 = lean_ctor_get(x_731, 1); +lean_inc(x_734); +lean_inc(x_733); +lean_dec(x_731); +x_735 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_735, 0, x_733); +lean_ctor_set(x_735, 1, x_734); +return x_735; } } } else { -lean_dec(x_660); -if (lean_obj_tag(x_656) == 0) +lean_dec(x_655); +if (lean_obj_tag(x_651) == 0) { -lean_object* x_741; lean_object* x_742; -lean_free_object(x_79); -lean_free_object(x_9); +lean_object* x_736; +lean_free_object(x_80); +lean_free_object(x_10); +lean_free_object(x_83); lean_free_object(x_82); lean_free_object(x_81); -lean_free_object(x_80); -lean_free_object(x_78); -lean_dec(x_658); -x_741 = l_Lean_identKind; -x_742 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_741, x_5, x_6, x_655); +lean_free_object(x_79); +lean_dec(x_653); +x_736 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_3, x_6, x_7, x_650); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -return x_742; +lean_dec(x_3); +return x_736; } else { -lean_object* x_743; lean_object* x_744; lean_object* x_745; lean_object* x_746; lean_object* x_747; lean_object* x_748; lean_object* x_749; lean_object* x_750; lean_object* x_751; uint8_t x_752; +lean_object* x_737; lean_object* x_738; lean_object* x_739; lean_object* x_740; lean_object* x_741; lean_object* x_742; lean_object* x_743; lean_object* x_744; lean_object* x_745; uint8_t x_746; +lean_dec(x_3); lean_dec(x_1); -lean_ctor_set(x_82, 1, x_671); -lean_ctor_set(x_81, 1, x_693); -lean_ctor_set(x_80, 1, x_709); -lean_ctor_set(x_79, 1, x_725); -x_743 = lean_box(0); -x_744 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_9, x_743); -x_745 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_744, x_743); -x_746 = l_Lean_MessageData_ofList(x_745); -lean_dec(x_745); -x_747 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_748 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_748, 0, x_747); -lean_ctor_set(x_748, 1, x_746); -x_749 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_750 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_750, 0, x_748); -lean_ctor_set(x_750, 1, x_749); -x_751 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_750, x_5, x_6, x_655); -lean_dec(x_6); -x_752 = !lean_is_exclusive(x_751); -if (x_752 == 0) +lean_ctor_set(x_83, 1, x_666); +lean_ctor_set(x_82, 1, x_688); +lean_ctor_set(x_81, 1, x_704); +lean_ctor_set(x_80, 1, x_720); +x_737 = lean_box(0); +x_738 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_10, x_737); +x_739 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_738, x_737); +x_740 = l_Lean_MessageData_ofList(x_739); +lean_dec(x_739); +x_741 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_742 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_742, 0, x_741); +lean_ctor_set(x_742, 1, x_740); +x_743 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_744 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_744, 0, x_742); +lean_ctor_set(x_744, 1, x_743); +x_745 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_744, x_6, x_7, x_650); +lean_dec(x_7); +x_746 = !lean_is_exclusive(x_745); +if (x_746 == 0) { -return x_751; +return x_745; } else { -lean_object* x_753; lean_object* x_754; lean_object* x_755; -x_753 = lean_ctor_get(x_751, 0); -x_754 = lean_ctor_get(x_751, 1); -lean_inc(x_754); -lean_inc(x_753); -lean_dec(x_751); -x_755 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_755, 0, x_753); -lean_ctor_set(x_755, 1, x_754); -return x_755; +lean_object* x_747; lean_object* x_748; lean_object* x_749; +x_747 = lean_ctor_get(x_745, 0); +x_748 = lean_ctor_get(x_745, 1); +lean_inc(x_748); +lean_inc(x_747); +lean_dec(x_745); +x_749 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_749, 0, x_747); +lean_ctor_set(x_749, 1, x_748); +return x_749; } } } @@ -5188,291 +5232,295 @@ return x_755; } else { -lean_object* x_756; uint8_t x_757; -lean_dec(x_79); -x_756 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__4; -x_757 = lean_string_dec_eq(x_666, x_756); -if (x_757 == 0) +lean_object* x_750; uint8_t x_751; +lean_dec(x_80); +x_750 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__4; +x_751 = lean_string_dec_eq(x_661, x_750); +if (x_751 == 0) { -if (lean_obj_tag(x_656) == 0) +lean_dec(x_3); +if (lean_obj_tag(x_651) == 0) { -lean_object* x_758; lean_object* x_759; -lean_free_object(x_9); -lean_free_object(x_78); -lean_dec(x_658); -lean_ctor_set(x_82, 1, x_671); -x_758 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_758, 0, x_80); -lean_ctor_set(x_758, 1, x_660); -lean_ctor_set_uint64(x_758, sizeof(void*)*2, x_662); -x_759 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_758, x_5, x_6, x_655); +lean_object* x_752; lean_object* x_753; +lean_free_object(x_10); +lean_free_object(x_79); +lean_dec(x_653); +lean_ctor_set(x_83, 1, x_666); +x_752 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_752, 0, x_81); +lean_ctor_set(x_752, 1, x_655); +lean_ctor_set_uint64(x_752, sizeof(void*)*2, x_657); +x_753 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_752, x_6, x_7, x_650); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_758); -return x_759; +lean_dec(x_752); +return x_753; } else { -lean_object* x_760; lean_object* x_761; lean_object* x_762; lean_object* x_763; lean_object* x_764; lean_object* x_765; lean_object* x_766; lean_object* x_767; lean_object* x_768; lean_object* x_769; lean_object* x_770; lean_object* x_771; lean_object* x_772; lean_object* x_773; +lean_object* x_754; lean_object* x_755; lean_object* x_756; lean_object* x_757; lean_object* x_758; lean_object* x_759; lean_object* x_760; lean_object* x_761; lean_object* x_762; lean_object* x_763; lean_object* x_764; lean_object* x_765; lean_object* x_766; lean_object* x_767; lean_dec(x_1); -lean_ctor_set(x_82, 1, x_671); -x_760 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_760, 0, x_80); -lean_ctor_set(x_760, 1, x_660); -lean_ctor_set_uint64(x_760, sizeof(void*)*2, x_662); -lean_ctor_set(x_78, 0, x_760); -x_761 = lean_box(0); -x_762 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_9, x_761); -x_763 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_762, x_761); -x_764 = l_Lean_MessageData_ofList(x_763); -lean_dec(x_763); -x_765 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_766 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_766, 0, x_765); -lean_ctor_set(x_766, 1, x_764); -x_767 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_768 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_768, 0, x_766); -lean_ctor_set(x_768, 1, x_767); -x_769 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_768, x_5, x_6, x_655); -lean_dec(x_6); -x_770 = lean_ctor_get(x_769, 0); -lean_inc(x_770); -x_771 = lean_ctor_get(x_769, 1); -lean_inc(x_771); -if (lean_is_exclusive(x_769)) { - lean_ctor_release(x_769, 0); - lean_ctor_release(x_769, 1); - x_772 = x_769; +lean_ctor_set(x_83, 1, x_666); +x_754 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_754, 0, x_81); +lean_ctor_set(x_754, 1, x_655); +lean_ctor_set_uint64(x_754, sizeof(void*)*2, x_657); +lean_ctor_set(x_79, 0, x_754); +x_755 = lean_box(0); +x_756 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_10, x_755); +x_757 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_756, x_755); +x_758 = l_Lean_MessageData_ofList(x_757); +lean_dec(x_757); +x_759 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_760 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_760, 0, x_759); +lean_ctor_set(x_760, 1, x_758); +x_761 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_762 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_762, 0, x_760); +lean_ctor_set(x_762, 1, x_761); +x_763 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_762, x_6, x_7, x_650); +lean_dec(x_7); +x_764 = lean_ctor_get(x_763, 0); +lean_inc(x_764); +x_765 = lean_ctor_get(x_763, 1); +lean_inc(x_765); +if (lean_is_exclusive(x_763)) { + lean_ctor_release(x_763, 0); + lean_ctor_release(x_763, 1); + x_766 = x_763; } else { - lean_dec_ref(x_769); - x_772 = lean_box(0); + lean_dec_ref(x_763); + x_766 = lean_box(0); } -if (lean_is_scalar(x_772)) { - x_773 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_766)) { + x_767 = lean_alloc_ctor(1, 2, 0); } else { - x_773 = x_772; + x_767 = x_766; } -lean_ctor_set(x_773, 0, x_770); -lean_ctor_set(x_773, 1, x_771); -return x_773; +lean_ctor_set(x_767, 0, x_764); +lean_ctor_set(x_767, 1, x_765); +return x_767; } } else { -lean_object* x_774; uint8_t x_775; -lean_dec(x_666); -x_774 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__9; -x_775 = lean_string_dec_eq(x_663, x_774); -if (x_775 == 0) +lean_object* x_768; uint8_t x_769; +lean_dec(x_661); +x_768 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__11; +x_769 = lean_string_dec_eq(x_658, x_768); +if (x_769 == 0) { -if (lean_obj_tag(x_656) == 0) +lean_dec(x_3); +if (lean_obj_tag(x_651) == 0) { -lean_object* x_776; lean_object* x_777; -lean_free_object(x_9); -lean_free_object(x_78); -lean_dec(x_658); -lean_ctor_set(x_82, 1, x_671); -lean_ctor_set(x_81, 1, x_756); -x_776 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_776, 0, x_80); -lean_ctor_set(x_776, 1, x_660); -lean_ctor_set_uint64(x_776, sizeof(void*)*2, x_662); -x_777 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_776, x_5, x_6, x_655); +lean_object* x_770; lean_object* x_771; +lean_free_object(x_10); +lean_free_object(x_79); +lean_dec(x_653); +lean_ctor_set(x_83, 1, x_666); +lean_ctor_set(x_82, 1, x_750); +x_770 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_770, 0, x_81); +lean_ctor_set(x_770, 1, x_655); +lean_ctor_set_uint64(x_770, sizeof(void*)*2, x_657); +x_771 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_770, x_6, x_7, x_650); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_776); -return x_777; +lean_dec(x_770); +return x_771; } else { -lean_object* x_778; lean_object* x_779; lean_object* x_780; lean_object* x_781; lean_object* x_782; lean_object* x_783; lean_object* x_784; lean_object* x_785; lean_object* x_786; lean_object* x_787; lean_object* x_788; lean_object* x_789; lean_object* x_790; lean_object* x_791; +lean_object* x_772; lean_object* x_773; lean_object* x_774; lean_object* x_775; lean_object* x_776; lean_object* x_777; lean_object* x_778; lean_object* x_779; lean_object* x_780; lean_object* x_781; lean_object* x_782; lean_object* x_783; lean_object* x_784; lean_object* x_785; lean_dec(x_1); -lean_ctor_set(x_82, 1, x_671); -lean_ctor_set(x_81, 1, x_756); -x_778 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_778, 0, x_80); -lean_ctor_set(x_778, 1, x_660); -lean_ctor_set_uint64(x_778, sizeof(void*)*2, x_662); -lean_ctor_set(x_78, 0, x_778); -x_779 = lean_box(0); -x_780 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_9, x_779); -x_781 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_780, x_779); -x_782 = l_Lean_MessageData_ofList(x_781); -lean_dec(x_781); -x_783 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_784 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_784, 0, x_783); -lean_ctor_set(x_784, 1, x_782); -x_785 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_786 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_786, 0, x_784); -lean_ctor_set(x_786, 1, x_785); -x_787 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_786, x_5, x_6, x_655); -lean_dec(x_6); -x_788 = lean_ctor_get(x_787, 0); -lean_inc(x_788); -x_789 = lean_ctor_get(x_787, 1); -lean_inc(x_789); -if (lean_is_exclusive(x_787)) { - lean_ctor_release(x_787, 0); - lean_ctor_release(x_787, 1); - x_790 = x_787; +lean_ctor_set(x_83, 1, x_666); +lean_ctor_set(x_82, 1, x_750); +x_772 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_772, 0, x_81); +lean_ctor_set(x_772, 1, x_655); +lean_ctor_set_uint64(x_772, sizeof(void*)*2, x_657); +lean_ctor_set(x_79, 0, x_772); +x_773 = lean_box(0); +x_774 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_10, x_773); +x_775 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_774, x_773); +x_776 = l_Lean_MessageData_ofList(x_775); +lean_dec(x_775); +x_777 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_778 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_778, 0, x_777); +lean_ctor_set(x_778, 1, x_776); +x_779 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_780 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_780, 0, x_778); +lean_ctor_set(x_780, 1, x_779); +x_781 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_780, x_6, x_7, x_650); +lean_dec(x_7); +x_782 = lean_ctor_get(x_781, 0); +lean_inc(x_782); +x_783 = lean_ctor_get(x_781, 1); +lean_inc(x_783); +if (lean_is_exclusive(x_781)) { + lean_ctor_release(x_781, 0); + lean_ctor_release(x_781, 1); + x_784 = x_781; } else { - lean_dec_ref(x_787); - x_790 = lean_box(0); + lean_dec_ref(x_781); + x_784 = lean_box(0); } -if (lean_is_scalar(x_790)) { - x_791 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_784)) { + x_785 = lean_alloc_ctor(1, 2, 0); } else { - x_791 = x_790; + x_785 = x_784; } -lean_ctor_set(x_791, 0, x_788); -lean_ctor_set(x_791, 1, x_789); -return x_791; +lean_ctor_set(x_785, 0, x_782); +lean_ctor_set(x_785, 1, x_783); +return x_785; } } else { -lean_object* x_792; uint8_t x_793; -lean_dec(x_663); -x_792 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__7; -x_793 = lean_string_dec_eq(x_660, x_792); -if (x_793 == 0) +lean_object* x_786; uint8_t x_787; +lean_dec(x_658); +x_786 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__7; +x_787 = lean_string_dec_eq(x_655, x_786); +if (x_787 == 0) { -if (lean_obj_tag(x_656) == 0) +lean_dec(x_3); +if (lean_obj_tag(x_651) == 0) { -lean_object* x_794; lean_object* x_795; -lean_free_object(x_9); -lean_free_object(x_78); -lean_dec(x_658); -lean_ctor_set(x_82, 1, x_671); -lean_ctor_set(x_81, 1, x_756); -lean_ctor_set(x_80, 1, x_774); -x_794 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_794, 0, x_80); -lean_ctor_set(x_794, 1, x_660); -lean_ctor_set_uint64(x_794, sizeof(void*)*2, x_662); -x_795 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_794, x_5, x_6, x_655); +lean_object* x_788; lean_object* x_789; +lean_free_object(x_10); +lean_free_object(x_79); +lean_dec(x_653); +lean_ctor_set(x_83, 1, x_666); +lean_ctor_set(x_82, 1, x_750); +lean_ctor_set(x_81, 1, x_768); +x_788 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_788, 0, x_81); +lean_ctor_set(x_788, 1, x_655); +lean_ctor_set_uint64(x_788, sizeof(void*)*2, x_657); +x_789 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_788, x_6, x_7, x_650); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_794); -return x_795; +lean_dec(x_788); +return x_789; } else { -lean_object* x_796; lean_object* x_797; lean_object* x_798; lean_object* x_799; lean_object* x_800; lean_object* x_801; lean_object* x_802; lean_object* x_803; lean_object* x_804; lean_object* x_805; lean_object* x_806; lean_object* x_807; lean_object* x_808; lean_object* x_809; +lean_object* x_790; lean_object* x_791; lean_object* x_792; lean_object* x_793; lean_object* x_794; lean_object* x_795; lean_object* x_796; lean_object* x_797; lean_object* x_798; lean_object* x_799; lean_object* x_800; lean_object* x_801; lean_object* x_802; lean_object* x_803; lean_dec(x_1); -lean_ctor_set(x_82, 1, x_671); -lean_ctor_set(x_81, 1, x_756); -lean_ctor_set(x_80, 1, x_774); -x_796 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_796, 0, x_80); -lean_ctor_set(x_796, 1, x_660); -lean_ctor_set_uint64(x_796, sizeof(void*)*2, x_662); -lean_ctor_set(x_78, 0, x_796); -x_797 = lean_box(0); -x_798 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_9, x_797); -x_799 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_798, x_797); -x_800 = l_Lean_MessageData_ofList(x_799); -lean_dec(x_799); -x_801 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_802 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_802, 0, x_801); -lean_ctor_set(x_802, 1, x_800); -x_803 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_804 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_804, 0, x_802); -lean_ctor_set(x_804, 1, x_803); -x_805 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_804, x_5, x_6, x_655); -lean_dec(x_6); -x_806 = lean_ctor_get(x_805, 0); -lean_inc(x_806); -x_807 = lean_ctor_get(x_805, 1); -lean_inc(x_807); -if (lean_is_exclusive(x_805)) { - lean_ctor_release(x_805, 0); - lean_ctor_release(x_805, 1); - x_808 = x_805; +lean_ctor_set(x_83, 1, x_666); +lean_ctor_set(x_82, 1, x_750); +lean_ctor_set(x_81, 1, x_768); +x_790 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_790, 0, x_81); +lean_ctor_set(x_790, 1, x_655); +lean_ctor_set_uint64(x_790, sizeof(void*)*2, x_657); +lean_ctor_set(x_79, 0, x_790); +x_791 = lean_box(0); +x_792 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_10, x_791); +x_793 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_792, x_791); +x_794 = l_Lean_MessageData_ofList(x_793); +lean_dec(x_793); +x_795 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_796 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_796, 0, x_795); +lean_ctor_set(x_796, 1, x_794); +x_797 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_798 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_798, 0, x_796); +lean_ctor_set(x_798, 1, x_797); +x_799 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_798, x_6, x_7, x_650); +lean_dec(x_7); +x_800 = lean_ctor_get(x_799, 0); +lean_inc(x_800); +x_801 = lean_ctor_get(x_799, 1); +lean_inc(x_801); +if (lean_is_exclusive(x_799)) { + lean_ctor_release(x_799, 0); + lean_ctor_release(x_799, 1); + x_802 = x_799; } else { - lean_dec_ref(x_805); - x_808 = lean_box(0); + lean_dec_ref(x_799); + x_802 = lean_box(0); } -if (lean_is_scalar(x_808)) { - x_809 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_802)) { + x_803 = lean_alloc_ctor(1, 2, 0); } else { - x_809 = x_808; + x_803 = x_802; } -lean_ctor_set(x_809, 0, x_806); -lean_ctor_set(x_809, 1, x_807); -return x_809; +lean_ctor_set(x_803, 0, x_800); +lean_ctor_set(x_803, 1, x_801); +return x_803; } } else { -lean_dec(x_660); -if (lean_obj_tag(x_656) == 0) +lean_dec(x_655); +if (lean_obj_tag(x_651) == 0) { -lean_object* x_810; lean_object* x_811; -lean_free_object(x_9); +lean_object* x_804; +lean_free_object(x_10); +lean_free_object(x_83); lean_free_object(x_82); lean_free_object(x_81); -lean_free_object(x_80); -lean_free_object(x_78); -lean_dec(x_658); -x_810 = l_Lean_identKind; -x_811 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_810, x_5, x_6, x_655); +lean_free_object(x_79); +lean_dec(x_653); +x_804 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_3, x_6, x_7, x_650); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -return x_811; +lean_dec(x_3); +return x_804; } else { -lean_object* x_812; lean_object* x_813; lean_object* x_814; lean_object* x_815; lean_object* x_816; lean_object* x_817; lean_object* x_818; lean_object* x_819; lean_object* x_820; lean_object* x_821; lean_object* x_822; lean_object* x_823; lean_object* x_824; lean_object* x_825; +lean_object* x_805; lean_object* x_806; lean_object* x_807; lean_object* x_808; lean_object* x_809; lean_object* x_810; lean_object* x_811; lean_object* x_812; lean_object* x_813; lean_object* x_814; lean_object* x_815; lean_object* x_816; lean_object* x_817; lean_object* x_818; +lean_dec(x_3); lean_dec(x_1); -lean_ctor_set(x_82, 1, x_671); -lean_ctor_set(x_81, 1, x_756); -lean_ctor_set(x_80, 1, x_774); -x_812 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_812, 0, x_80); -lean_ctor_set(x_812, 1, x_792); -lean_ctor_set_uint64(x_812, sizeof(void*)*2, x_662); -lean_ctor_set(x_78, 0, x_812); -x_813 = lean_box(0); -x_814 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_9, x_813); -x_815 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_814, x_813); -x_816 = l_Lean_MessageData_ofList(x_815); -lean_dec(x_815); -x_817 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_818 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_818, 0, x_817); -lean_ctor_set(x_818, 1, x_816); -x_819 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_820 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_820, 0, x_818); -lean_ctor_set(x_820, 1, x_819); -x_821 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_820, x_5, x_6, x_655); -lean_dec(x_6); -x_822 = lean_ctor_get(x_821, 0); -lean_inc(x_822); -x_823 = lean_ctor_get(x_821, 1); -lean_inc(x_823); -if (lean_is_exclusive(x_821)) { - lean_ctor_release(x_821, 0); - lean_ctor_release(x_821, 1); - x_824 = x_821; +lean_ctor_set(x_83, 1, x_666); +lean_ctor_set(x_82, 1, x_750); +lean_ctor_set(x_81, 1, x_768); +x_805 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_805, 0, x_81); +lean_ctor_set(x_805, 1, x_786); +lean_ctor_set_uint64(x_805, sizeof(void*)*2, x_657); +lean_ctor_set(x_79, 0, x_805); +x_806 = lean_box(0); +x_807 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_10, x_806); +x_808 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_807, x_806); +x_809 = l_Lean_MessageData_ofList(x_808); +lean_dec(x_808); +x_810 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_811 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_811, 0, x_810); +lean_ctor_set(x_811, 1, x_809); +x_812 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_813 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_813, 0, x_811); +lean_ctor_set(x_813, 1, x_812); +x_814 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_813, x_6, x_7, x_650); +lean_dec(x_7); +x_815 = lean_ctor_get(x_814, 0); +lean_inc(x_815); +x_816 = lean_ctor_get(x_814, 1); +lean_inc(x_816); +if (lean_is_exclusive(x_814)) { + lean_ctor_release(x_814, 0); + lean_ctor_release(x_814, 1); + x_817 = x_814; } else { - lean_dec_ref(x_821); - x_824 = lean_box(0); + lean_dec_ref(x_814); + x_817 = lean_box(0); } -if (lean_is_scalar(x_824)) { - x_825 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_817)) { + x_818 = lean_alloc_ctor(1, 2, 0); } else { - x_825 = x_824; + x_818 = x_817; } -lean_ctor_set(x_825, 0, x_822); -lean_ctor_set(x_825, 1, x_823); -return x_825; +lean_ctor_set(x_818, 0, x_815); +lean_ctor_set(x_818, 1, x_816); +return x_818; } } } @@ -5481,336 +5529,340 @@ return x_825; } else { -lean_object* x_826; lean_object* x_827; uint8_t x_828; -lean_dec(x_9); -if (lean_is_exclusive(x_79)) { - lean_ctor_release(x_79, 0); - lean_ctor_release(x_79, 1); - x_826 = x_79; +lean_object* x_819; lean_object* x_820; uint8_t x_821; +lean_dec(x_10); +if (lean_is_exclusive(x_80)) { + lean_ctor_release(x_80, 0); + lean_ctor_release(x_80, 1); + x_819 = x_80; } else { - lean_dec_ref(x_79); - x_826 = lean_box(0); + lean_dec_ref(x_80); + x_819 = lean_box(0); } -x_827 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__4; -x_828 = lean_string_dec_eq(x_666, x_827); -if (x_828 == 0) +x_820 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__4; +x_821 = lean_string_dec_eq(x_661, x_820); +if (x_821 == 0) { -if (lean_obj_tag(x_656) == 0) +lean_dec(x_3); +if (lean_obj_tag(x_651) == 0) { -lean_object* x_829; lean_object* x_830; -lean_free_object(x_78); -lean_dec(x_658); -lean_ctor_set(x_82, 1, x_671); -if (lean_is_scalar(x_826)) { - x_829 = lean_alloc_ctor(1, 2, 8); +lean_object* x_822; lean_object* x_823; +lean_free_object(x_79); +lean_dec(x_653); +lean_ctor_set(x_83, 1, x_666); +if (lean_is_scalar(x_819)) { + x_822 = lean_alloc_ctor(1, 2, 8); } else { - x_829 = x_826; + x_822 = x_819; } -lean_ctor_set(x_829, 0, x_80); -lean_ctor_set(x_829, 1, x_660); -lean_ctor_set_uint64(x_829, sizeof(void*)*2, x_662); -x_830 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_829, x_5, x_6, x_655); +lean_ctor_set(x_822, 0, x_81); +lean_ctor_set(x_822, 1, x_655); +lean_ctor_set_uint64(x_822, sizeof(void*)*2, x_657); +x_823 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_822, x_6, x_7, x_650); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_829); -return x_830; +lean_dec(x_822); +return x_823; } else { -lean_object* x_831; lean_object* x_832; lean_object* x_833; lean_object* x_834; lean_object* x_835; lean_object* x_836; lean_object* x_837; lean_object* x_838; lean_object* x_839; lean_object* x_840; lean_object* x_841; lean_object* x_842; lean_object* x_843; lean_object* x_844; lean_object* x_845; +lean_object* x_824; lean_object* x_825; lean_object* x_826; lean_object* x_827; lean_object* x_828; lean_object* x_829; lean_object* x_830; lean_object* x_831; lean_object* x_832; lean_object* x_833; lean_object* x_834; lean_object* x_835; lean_object* x_836; lean_object* x_837; lean_object* x_838; lean_dec(x_1); -lean_ctor_set(x_82, 1, x_671); -if (lean_is_scalar(x_826)) { - x_831 = lean_alloc_ctor(1, 2, 8); -} else { - x_831 = x_826; -} -lean_ctor_set(x_831, 0, x_80); -lean_ctor_set(x_831, 1, x_660); -lean_ctor_set_uint64(x_831, sizeof(void*)*2, x_662); -lean_ctor_set(x_78, 0, x_831); -x_832 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_832, 0, x_78); -lean_ctor_set(x_832, 1, x_656); -x_833 = lean_box(0); -x_834 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_832, x_833); -x_835 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_834, x_833); -x_836 = l_Lean_MessageData_ofList(x_835); -lean_dec(x_835); -x_837 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_838 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_838, 0, x_837); -lean_ctor_set(x_838, 1, x_836); -x_839 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_840 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_840, 0, x_838); -lean_ctor_set(x_840, 1, x_839); -x_841 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_840, x_5, x_6, x_655); -lean_dec(x_6); -x_842 = lean_ctor_get(x_841, 0); -lean_inc(x_842); -x_843 = lean_ctor_get(x_841, 1); -lean_inc(x_843); -if (lean_is_exclusive(x_841)) { - lean_ctor_release(x_841, 0); - lean_ctor_release(x_841, 1); - x_844 = x_841; +lean_ctor_set(x_83, 1, x_666); +if (lean_is_scalar(x_819)) { + x_824 = lean_alloc_ctor(1, 2, 8); +} else { + x_824 = x_819; +} +lean_ctor_set(x_824, 0, x_81); +lean_ctor_set(x_824, 1, x_655); +lean_ctor_set_uint64(x_824, sizeof(void*)*2, x_657); +lean_ctor_set(x_79, 0, x_824); +x_825 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_825, 0, x_79); +lean_ctor_set(x_825, 1, x_651); +x_826 = lean_box(0); +x_827 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_825, x_826); +x_828 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_827, x_826); +x_829 = l_Lean_MessageData_ofList(x_828); +lean_dec(x_828); +x_830 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_831 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_831, 0, x_830); +lean_ctor_set(x_831, 1, x_829); +x_832 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_833 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_833, 0, x_831); +lean_ctor_set(x_833, 1, x_832); +x_834 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_833, x_6, x_7, x_650); +lean_dec(x_7); +x_835 = lean_ctor_get(x_834, 0); +lean_inc(x_835); +x_836 = lean_ctor_get(x_834, 1); +lean_inc(x_836); +if (lean_is_exclusive(x_834)) { + lean_ctor_release(x_834, 0); + lean_ctor_release(x_834, 1); + x_837 = x_834; } else { - lean_dec_ref(x_841); - x_844 = lean_box(0); + lean_dec_ref(x_834); + x_837 = lean_box(0); } -if (lean_is_scalar(x_844)) { - x_845 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_837)) { + x_838 = lean_alloc_ctor(1, 2, 0); } else { - x_845 = x_844; + x_838 = x_837; } -lean_ctor_set(x_845, 0, x_842); -lean_ctor_set(x_845, 1, x_843); -return x_845; +lean_ctor_set(x_838, 0, x_835); +lean_ctor_set(x_838, 1, x_836); +return x_838; } } else { -lean_object* x_846; uint8_t x_847; -lean_dec(x_666); -x_846 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__9; -x_847 = lean_string_dec_eq(x_663, x_846); -if (x_847 == 0) +lean_object* x_839; uint8_t x_840; +lean_dec(x_661); +x_839 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__11; +x_840 = lean_string_dec_eq(x_658, x_839); +if (x_840 == 0) { -if (lean_obj_tag(x_656) == 0) +lean_dec(x_3); +if (lean_obj_tag(x_651) == 0) { -lean_object* x_848; lean_object* x_849; -lean_free_object(x_78); -lean_dec(x_658); -lean_ctor_set(x_82, 1, x_671); -lean_ctor_set(x_81, 1, x_827); -if (lean_is_scalar(x_826)) { - x_848 = lean_alloc_ctor(1, 2, 8); -} else { - x_848 = x_826; -} -lean_ctor_set(x_848, 0, x_80); -lean_ctor_set(x_848, 1, x_660); -lean_ctor_set_uint64(x_848, sizeof(void*)*2, x_662); -x_849 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_848, x_5, x_6, x_655); +lean_object* x_841; lean_object* x_842; +lean_free_object(x_79); +lean_dec(x_653); +lean_ctor_set(x_83, 1, x_666); +lean_ctor_set(x_82, 1, x_820); +if (lean_is_scalar(x_819)) { + x_841 = lean_alloc_ctor(1, 2, 8); +} else { + x_841 = x_819; +} +lean_ctor_set(x_841, 0, x_81); +lean_ctor_set(x_841, 1, x_655); +lean_ctor_set_uint64(x_841, sizeof(void*)*2, x_657); +x_842 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_841, x_6, x_7, x_650); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_848); -return x_849; +lean_dec(x_841); +return x_842; } else { -lean_object* x_850; lean_object* x_851; lean_object* x_852; lean_object* x_853; lean_object* x_854; lean_object* x_855; lean_object* x_856; lean_object* x_857; lean_object* x_858; lean_object* x_859; lean_object* x_860; lean_object* x_861; lean_object* x_862; lean_object* x_863; lean_object* x_864; +lean_object* x_843; lean_object* x_844; lean_object* x_845; lean_object* x_846; lean_object* x_847; lean_object* x_848; lean_object* x_849; lean_object* x_850; lean_object* x_851; lean_object* x_852; lean_object* x_853; lean_object* x_854; lean_object* x_855; lean_object* x_856; lean_object* x_857; lean_dec(x_1); -lean_ctor_set(x_82, 1, x_671); -lean_ctor_set(x_81, 1, x_827); -if (lean_is_scalar(x_826)) { - x_850 = lean_alloc_ctor(1, 2, 8); -} else { - x_850 = x_826; -} -lean_ctor_set(x_850, 0, x_80); -lean_ctor_set(x_850, 1, x_660); -lean_ctor_set_uint64(x_850, sizeof(void*)*2, x_662); -lean_ctor_set(x_78, 0, x_850); -x_851 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_851, 0, x_78); -lean_ctor_set(x_851, 1, x_656); -x_852 = lean_box(0); -x_853 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_851, x_852); -x_854 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_853, x_852); -x_855 = l_Lean_MessageData_ofList(x_854); -lean_dec(x_854); -x_856 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_857 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_857, 0, x_856); -lean_ctor_set(x_857, 1, x_855); -x_858 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_859 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_859, 0, x_857); -lean_ctor_set(x_859, 1, x_858); -x_860 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_859, x_5, x_6, x_655); -lean_dec(x_6); -x_861 = lean_ctor_get(x_860, 0); -lean_inc(x_861); -x_862 = lean_ctor_get(x_860, 1); -lean_inc(x_862); -if (lean_is_exclusive(x_860)) { - lean_ctor_release(x_860, 0); - lean_ctor_release(x_860, 1); - x_863 = x_860; +lean_ctor_set(x_83, 1, x_666); +lean_ctor_set(x_82, 1, x_820); +if (lean_is_scalar(x_819)) { + x_843 = lean_alloc_ctor(1, 2, 8); +} else { + x_843 = x_819; +} +lean_ctor_set(x_843, 0, x_81); +lean_ctor_set(x_843, 1, x_655); +lean_ctor_set_uint64(x_843, sizeof(void*)*2, x_657); +lean_ctor_set(x_79, 0, x_843); +x_844 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_844, 0, x_79); +lean_ctor_set(x_844, 1, x_651); +x_845 = lean_box(0); +x_846 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_844, x_845); +x_847 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_846, x_845); +x_848 = l_Lean_MessageData_ofList(x_847); +lean_dec(x_847); +x_849 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_850 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_850, 0, x_849); +lean_ctor_set(x_850, 1, x_848); +x_851 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_852 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_852, 0, x_850); +lean_ctor_set(x_852, 1, x_851); +x_853 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_852, x_6, x_7, x_650); +lean_dec(x_7); +x_854 = lean_ctor_get(x_853, 0); +lean_inc(x_854); +x_855 = lean_ctor_get(x_853, 1); +lean_inc(x_855); +if (lean_is_exclusive(x_853)) { + lean_ctor_release(x_853, 0); + lean_ctor_release(x_853, 1); + x_856 = x_853; } else { - lean_dec_ref(x_860); - x_863 = lean_box(0); + lean_dec_ref(x_853); + x_856 = lean_box(0); } -if (lean_is_scalar(x_863)) { - x_864 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_856)) { + x_857 = lean_alloc_ctor(1, 2, 0); } else { - x_864 = x_863; + x_857 = x_856; } -lean_ctor_set(x_864, 0, x_861); -lean_ctor_set(x_864, 1, x_862); -return x_864; +lean_ctor_set(x_857, 0, x_854); +lean_ctor_set(x_857, 1, x_855); +return x_857; } } else { -lean_object* x_865; uint8_t x_866; -lean_dec(x_663); -x_865 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__7; -x_866 = lean_string_dec_eq(x_660, x_865); -if (x_866 == 0) +lean_object* x_858; uint8_t x_859; +lean_dec(x_658); +x_858 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__7; +x_859 = lean_string_dec_eq(x_655, x_858); +if (x_859 == 0) { -if (lean_obj_tag(x_656) == 0) +lean_dec(x_3); +if (lean_obj_tag(x_651) == 0) { -lean_object* x_867; lean_object* x_868; -lean_free_object(x_78); -lean_dec(x_658); -lean_ctor_set(x_82, 1, x_671); -lean_ctor_set(x_81, 1, x_827); -lean_ctor_set(x_80, 1, x_846); -if (lean_is_scalar(x_826)) { - x_867 = lean_alloc_ctor(1, 2, 8); -} else { - x_867 = x_826; -} -lean_ctor_set(x_867, 0, x_80); -lean_ctor_set(x_867, 1, x_660); -lean_ctor_set_uint64(x_867, sizeof(void*)*2, x_662); -x_868 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_867, x_5, x_6, x_655); +lean_object* x_860; lean_object* x_861; +lean_free_object(x_79); +lean_dec(x_653); +lean_ctor_set(x_83, 1, x_666); +lean_ctor_set(x_82, 1, x_820); +lean_ctor_set(x_81, 1, x_839); +if (lean_is_scalar(x_819)) { + x_860 = lean_alloc_ctor(1, 2, 8); +} else { + x_860 = x_819; +} +lean_ctor_set(x_860, 0, x_81); +lean_ctor_set(x_860, 1, x_655); +lean_ctor_set_uint64(x_860, sizeof(void*)*2, x_657); +x_861 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_860, x_6, x_7, x_650); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_867); -return x_868; +lean_dec(x_860); +return x_861; } else { -lean_object* x_869; lean_object* x_870; lean_object* x_871; lean_object* x_872; lean_object* x_873; lean_object* x_874; lean_object* x_875; lean_object* x_876; lean_object* x_877; lean_object* x_878; lean_object* x_879; lean_object* x_880; lean_object* x_881; lean_object* x_882; lean_object* x_883; +lean_object* x_862; lean_object* x_863; lean_object* x_864; lean_object* x_865; lean_object* x_866; lean_object* x_867; lean_object* x_868; lean_object* x_869; lean_object* x_870; lean_object* x_871; lean_object* x_872; lean_object* x_873; lean_object* x_874; lean_object* x_875; lean_object* x_876; lean_dec(x_1); -lean_ctor_set(x_82, 1, x_671); -lean_ctor_set(x_81, 1, x_827); -lean_ctor_set(x_80, 1, x_846); -if (lean_is_scalar(x_826)) { - x_869 = lean_alloc_ctor(1, 2, 8); -} else { - x_869 = x_826; -} -lean_ctor_set(x_869, 0, x_80); -lean_ctor_set(x_869, 1, x_660); -lean_ctor_set_uint64(x_869, sizeof(void*)*2, x_662); -lean_ctor_set(x_78, 0, x_869); -x_870 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_870, 0, x_78); -lean_ctor_set(x_870, 1, x_656); -x_871 = lean_box(0); -x_872 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_870, x_871); -x_873 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_872, x_871); -x_874 = l_Lean_MessageData_ofList(x_873); -lean_dec(x_873); -x_875 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_876 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_876, 0, x_875); -lean_ctor_set(x_876, 1, x_874); -x_877 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_878 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_878, 0, x_876); -lean_ctor_set(x_878, 1, x_877); -x_879 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_878, x_5, x_6, x_655); -lean_dec(x_6); -x_880 = lean_ctor_get(x_879, 0); -lean_inc(x_880); -x_881 = lean_ctor_get(x_879, 1); -lean_inc(x_881); -if (lean_is_exclusive(x_879)) { - lean_ctor_release(x_879, 0); - lean_ctor_release(x_879, 1); - x_882 = x_879; +lean_ctor_set(x_83, 1, x_666); +lean_ctor_set(x_82, 1, x_820); +lean_ctor_set(x_81, 1, x_839); +if (lean_is_scalar(x_819)) { + x_862 = lean_alloc_ctor(1, 2, 8); +} else { + x_862 = x_819; +} +lean_ctor_set(x_862, 0, x_81); +lean_ctor_set(x_862, 1, x_655); +lean_ctor_set_uint64(x_862, sizeof(void*)*2, x_657); +lean_ctor_set(x_79, 0, x_862); +x_863 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_863, 0, x_79); +lean_ctor_set(x_863, 1, x_651); +x_864 = lean_box(0); +x_865 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_863, x_864); +x_866 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_865, x_864); +x_867 = l_Lean_MessageData_ofList(x_866); +lean_dec(x_866); +x_868 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_869 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_869, 0, x_868); +lean_ctor_set(x_869, 1, x_867); +x_870 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_871 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_871, 0, x_869); +lean_ctor_set(x_871, 1, x_870); +x_872 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_871, x_6, x_7, x_650); +lean_dec(x_7); +x_873 = lean_ctor_get(x_872, 0); +lean_inc(x_873); +x_874 = lean_ctor_get(x_872, 1); +lean_inc(x_874); +if (lean_is_exclusive(x_872)) { + lean_ctor_release(x_872, 0); + lean_ctor_release(x_872, 1); + x_875 = x_872; } else { - lean_dec_ref(x_879); - x_882 = lean_box(0); + lean_dec_ref(x_872); + x_875 = lean_box(0); } -if (lean_is_scalar(x_882)) { - x_883 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_875)) { + x_876 = lean_alloc_ctor(1, 2, 0); } else { - x_883 = x_882; + x_876 = x_875; } -lean_ctor_set(x_883, 0, x_880); -lean_ctor_set(x_883, 1, x_881); -return x_883; +lean_ctor_set(x_876, 0, x_873); +lean_ctor_set(x_876, 1, x_874); +return x_876; } } else { -lean_dec(x_660); -if (lean_obj_tag(x_656) == 0) +lean_dec(x_655); +if (lean_obj_tag(x_651) == 0) { -lean_object* x_884; lean_object* x_885; -lean_dec(x_826); +lean_object* x_877; +lean_dec(x_819); +lean_free_object(x_83); lean_free_object(x_82); lean_free_object(x_81); -lean_free_object(x_80); -lean_free_object(x_78); -lean_dec(x_658); -x_884 = l_Lean_identKind; -x_885 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_884, x_5, x_6, x_655); +lean_free_object(x_79); +lean_dec(x_653); +x_877 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_3, x_6, x_7, x_650); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -return x_885; +lean_dec(x_3); +return x_877; } else { -lean_object* x_886; lean_object* x_887; lean_object* x_888; lean_object* x_889; lean_object* x_890; lean_object* x_891; lean_object* x_892; lean_object* x_893; lean_object* x_894; lean_object* x_895; lean_object* x_896; lean_object* x_897; lean_object* x_898; lean_object* x_899; lean_object* x_900; +lean_object* x_878; lean_object* x_879; lean_object* x_880; lean_object* x_881; lean_object* x_882; lean_object* x_883; lean_object* x_884; lean_object* x_885; lean_object* x_886; lean_object* x_887; lean_object* x_888; lean_object* x_889; lean_object* x_890; lean_object* x_891; lean_object* x_892; +lean_dec(x_3); lean_dec(x_1); -lean_ctor_set(x_82, 1, x_671); -lean_ctor_set(x_81, 1, x_827); -lean_ctor_set(x_80, 1, x_846); -if (lean_is_scalar(x_826)) { - x_886 = lean_alloc_ctor(1, 2, 8); -} else { - x_886 = x_826; -} -lean_ctor_set(x_886, 0, x_80); -lean_ctor_set(x_886, 1, x_865); -lean_ctor_set_uint64(x_886, sizeof(void*)*2, x_662); -lean_ctor_set(x_78, 0, x_886); -x_887 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_887, 0, x_78); -lean_ctor_set(x_887, 1, x_656); -x_888 = lean_box(0); -x_889 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_887, x_888); -x_890 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_889, x_888); -x_891 = l_Lean_MessageData_ofList(x_890); -lean_dec(x_890); -x_892 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_893 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_893, 0, x_892); -lean_ctor_set(x_893, 1, x_891); -x_894 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_895 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_895, 0, x_893); -lean_ctor_set(x_895, 1, x_894); -x_896 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_895, x_5, x_6, x_655); -lean_dec(x_6); -x_897 = lean_ctor_get(x_896, 0); -lean_inc(x_897); -x_898 = lean_ctor_get(x_896, 1); -lean_inc(x_898); -if (lean_is_exclusive(x_896)) { - lean_ctor_release(x_896, 0); - lean_ctor_release(x_896, 1); - x_899 = x_896; +lean_ctor_set(x_83, 1, x_666); +lean_ctor_set(x_82, 1, x_820); +lean_ctor_set(x_81, 1, x_839); +if (lean_is_scalar(x_819)) { + x_878 = lean_alloc_ctor(1, 2, 8); +} else { + x_878 = x_819; +} +lean_ctor_set(x_878, 0, x_81); +lean_ctor_set(x_878, 1, x_858); +lean_ctor_set_uint64(x_878, sizeof(void*)*2, x_657); +lean_ctor_set(x_79, 0, x_878); +x_879 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_879, 0, x_79); +lean_ctor_set(x_879, 1, x_651); +x_880 = lean_box(0); +x_881 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_879, x_880); +x_882 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_881, x_880); +x_883 = l_Lean_MessageData_ofList(x_882); +lean_dec(x_882); +x_884 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_885 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_885, 0, x_884); +lean_ctor_set(x_885, 1, x_883); +x_886 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_887 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_887, 0, x_885); +lean_ctor_set(x_887, 1, x_886); +x_888 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_887, x_6, x_7, x_650); +lean_dec(x_7); +x_889 = lean_ctor_get(x_888, 0); +lean_inc(x_889); +x_890 = lean_ctor_get(x_888, 1); +lean_inc(x_890); +if (lean_is_exclusive(x_888)) { + lean_ctor_release(x_888, 0); + lean_ctor_release(x_888, 1); + x_891 = x_888; } else { - lean_dec_ref(x_896); - x_899 = lean_box(0); + lean_dec_ref(x_888); + x_891 = lean_box(0); } -if (lean_is_scalar(x_899)) { - x_900 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_891)) { + x_892 = lean_alloc_ctor(1, 2, 0); } else { - x_900 = x_899; + x_892 = x_891; } -lean_ctor_set(x_900, 0, x_897); -lean_ctor_set(x_900, 1, x_898); -return x_900; +lean_ctor_set(x_892, 0, x_889); +lean_ctor_set(x_892, 1, x_890); +return x_892; } } } @@ -5820,462 +5872,467 @@ return x_900; } else { -lean_object* x_901; uint64_t x_902; lean_object* x_903; uint8_t x_904; -x_901 = lean_ctor_get(x_82, 1); -x_902 = lean_ctor_get_uint64(x_82, sizeof(void*)*2); -lean_inc(x_901); -lean_dec(x_82); -x_903 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__1; -x_904 = lean_string_dec_eq(x_901, x_903); -lean_dec(x_901); -if (x_904 == 0) +lean_object* x_893; uint64_t x_894; lean_object* x_895; uint8_t x_896; +x_893 = lean_ctor_get(x_83, 1); +x_894 = lean_ctor_get_uint64(x_83, sizeof(void*)*2); +lean_inc(x_893); +lean_dec(x_83); +x_895 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__1; +x_896 = lean_string_dec_eq(x_893, x_895); +lean_dec(x_893); +if (x_896 == 0) { +lean_free_object(x_82); +lean_dec(x_661); lean_free_object(x_81); -lean_dec(x_666); -lean_free_object(x_80); -lean_dec(x_663); -lean_dec(x_660); -lean_free_object(x_78); lean_dec(x_658); -if (lean_obj_tag(x_656) == 0) +lean_dec(x_655); +lean_free_object(x_79); +lean_dec(x_653); +lean_dec(x_3); +if (lean_obj_tag(x_651) == 0) { -lean_object* x_905; -lean_dec(x_9); -x_905 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_79, x_5, x_6, x_655); +lean_object* x_897; +lean_dec(x_10); +x_897 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_80, x_6, x_7, x_650); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_79); -return x_905; +lean_dec(x_80); +return x_897; } else { -lean_object* x_906; lean_object* x_907; lean_object* x_908; lean_object* x_909; lean_object* x_910; lean_object* x_911; lean_object* x_912; lean_object* x_913; lean_object* x_914; lean_object* x_915; lean_object* x_916; lean_object* x_917; lean_object* x_918; -lean_dec(x_656); -lean_dec(x_79); +lean_object* x_898; lean_object* x_899; lean_object* x_900; lean_object* x_901; lean_object* x_902; lean_object* x_903; lean_object* x_904; lean_object* x_905; lean_object* x_906; lean_object* x_907; lean_object* x_908; lean_object* x_909; lean_object* x_910; +lean_dec(x_651); +lean_dec(x_80); lean_dec(x_1); -x_906 = lean_box(0); -x_907 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_9, x_906); -x_908 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_907, x_906); -x_909 = l_Lean_MessageData_ofList(x_908); -lean_dec(x_908); -x_910 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_911 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_911, 0, x_910); -lean_ctor_set(x_911, 1, x_909); -x_912 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_913 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_913, 0, x_911); -lean_ctor_set(x_913, 1, x_912); -x_914 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_913, x_5, x_6, x_655); -lean_dec(x_6); -x_915 = lean_ctor_get(x_914, 0); -lean_inc(x_915); -x_916 = lean_ctor_get(x_914, 1); -lean_inc(x_916); -if (lean_is_exclusive(x_914)) { - lean_ctor_release(x_914, 0); - lean_ctor_release(x_914, 1); - x_917 = x_914; +x_898 = lean_box(0); +x_899 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_10, x_898); +x_900 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_899, x_898); +x_901 = l_Lean_MessageData_ofList(x_900); +lean_dec(x_900); +x_902 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_903 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_903, 0, x_902); +lean_ctor_set(x_903, 1, x_901); +x_904 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_905 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_905, 0, x_903); +lean_ctor_set(x_905, 1, x_904); +x_906 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_905, x_6, x_7, x_650); +lean_dec(x_7); +x_907 = lean_ctor_get(x_906, 0); +lean_inc(x_907); +x_908 = lean_ctor_get(x_906, 1); +lean_inc(x_908); +if (lean_is_exclusive(x_906)) { + lean_ctor_release(x_906, 0); + lean_ctor_release(x_906, 1); + x_909 = x_906; } else { - lean_dec_ref(x_914); - x_917 = lean_box(0); + lean_dec_ref(x_906); + x_909 = lean_box(0); } -if (lean_is_scalar(x_917)) { - x_918 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_909)) { + x_910 = lean_alloc_ctor(1, 2, 0); } else { - x_918 = x_917; + x_910 = x_909; } -lean_ctor_set(x_918, 0, x_915); -lean_ctor_set(x_918, 1, x_916); -return x_918; +lean_ctor_set(x_910, 0, x_907); +lean_ctor_set(x_910, 1, x_908); +return x_910; } } else { -lean_object* x_919; lean_object* x_920; lean_object* x_921; uint8_t x_922; -if (lean_is_exclusive(x_9)) { - lean_ctor_release(x_9, 0); - lean_ctor_release(x_9, 1); - x_919 = x_9; +lean_object* x_911; lean_object* x_912; lean_object* x_913; uint8_t x_914; +if (lean_is_exclusive(x_10)) { + lean_ctor_release(x_10, 0); + lean_ctor_release(x_10, 1); + x_911 = x_10; } else { - lean_dec_ref(x_9); - x_919 = lean_box(0); + lean_dec_ref(x_10); + x_911 = lean_box(0); } -if (lean_is_exclusive(x_79)) { - lean_ctor_release(x_79, 0); - lean_ctor_release(x_79, 1); - x_920 = x_79; +if (lean_is_exclusive(x_80)) { + lean_ctor_release(x_80, 0); + lean_ctor_release(x_80, 1); + x_912 = x_80; } else { - lean_dec_ref(x_79); - x_920 = lean_box(0); + lean_dec_ref(x_80); + x_912 = lean_box(0); } -x_921 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__4; -x_922 = lean_string_dec_eq(x_666, x_921); -if (x_922 == 0) +x_913 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__4; +x_914 = lean_string_dec_eq(x_661, x_913); +if (x_914 == 0) { -if (lean_obj_tag(x_656) == 0) +lean_dec(x_3); +if (lean_obj_tag(x_651) == 0) { -lean_object* x_923; lean_object* x_924; lean_object* x_925; -lean_dec(x_919); -lean_free_object(x_78); -lean_dec(x_658); -x_923 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_923, 0, x_654); -lean_ctor_set(x_923, 1, x_903); -lean_ctor_set_uint64(x_923, sizeof(void*)*2, x_902); -lean_ctor_set(x_81, 0, x_923); -if (lean_is_scalar(x_920)) { - x_924 = lean_alloc_ctor(1, 2, 8); -} else { - x_924 = x_920; -} -lean_ctor_set(x_924, 0, x_80); -lean_ctor_set(x_924, 1, x_660); -lean_ctor_set_uint64(x_924, sizeof(void*)*2, x_662); -x_925 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_924, x_5, x_6, x_655); +lean_object* x_915; lean_object* x_916; lean_object* x_917; +lean_dec(x_911); +lean_free_object(x_79); +lean_dec(x_653); +x_915 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_915, 0, x_649); +lean_ctor_set(x_915, 1, x_895); +lean_ctor_set_uint64(x_915, sizeof(void*)*2, x_894); +lean_ctor_set(x_82, 0, x_915); +if (lean_is_scalar(x_912)) { + x_916 = lean_alloc_ctor(1, 2, 8); +} else { + x_916 = x_912; +} +lean_ctor_set(x_916, 0, x_81); +lean_ctor_set(x_916, 1, x_655); +lean_ctor_set_uint64(x_916, sizeof(void*)*2, x_657); +x_917 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_916, x_6, x_7, x_650); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_924); -return x_925; +lean_dec(x_916); +return x_917; } else { -lean_object* x_926; lean_object* x_927; lean_object* x_928; lean_object* x_929; lean_object* x_930; lean_object* x_931; lean_object* x_932; lean_object* x_933; lean_object* x_934; lean_object* x_935; lean_object* x_936; lean_object* x_937; lean_object* x_938; lean_object* x_939; lean_object* x_940; lean_object* x_941; +lean_object* x_918; lean_object* x_919; lean_object* x_920; lean_object* x_921; lean_object* x_922; lean_object* x_923; lean_object* x_924; lean_object* x_925; lean_object* x_926; lean_object* x_927; lean_object* x_928; lean_object* x_929; lean_object* x_930; lean_object* x_931; lean_object* x_932; lean_object* x_933; lean_dec(x_1); -x_926 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_926, 0, x_654); -lean_ctor_set(x_926, 1, x_903); -lean_ctor_set_uint64(x_926, sizeof(void*)*2, x_902); -lean_ctor_set(x_81, 0, x_926); -if (lean_is_scalar(x_920)) { - x_927 = lean_alloc_ctor(1, 2, 8); -} else { - x_927 = x_920; -} -lean_ctor_set(x_927, 0, x_80); -lean_ctor_set(x_927, 1, x_660); -lean_ctor_set_uint64(x_927, sizeof(void*)*2, x_662); -lean_ctor_set(x_78, 0, x_927); -if (lean_is_scalar(x_919)) { - x_928 = lean_alloc_ctor(1, 2, 0); -} else { - x_928 = x_919; -} -lean_ctor_set(x_928, 0, x_78); -lean_ctor_set(x_928, 1, x_656); -x_929 = lean_box(0); -x_930 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_928, x_929); -x_931 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_930, x_929); -x_932 = l_Lean_MessageData_ofList(x_931); -lean_dec(x_931); -x_933 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_934 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_934, 0, x_933); -lean_ctor_set(x_934, 1, x_932); -x_935 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_936 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_936, 0, x_934); -lean_ctor_set(x_936, 1, x_935); -x_937 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_936, x_5, x_6, x_655); -lean_dec(x_6); -x_938 = lean_ctor_get(x_937, 0); -lean_inc(x_938); -x_939 = lean_ctor_get(x_937, 1); -lean_inc(x_939); -if (lean_is_exclusive(x_937)) { - lean_ctor_release(x_937, 0); - lean_ctor_release(x_937, 1); - x_940 = x_937; -} else { - lean_dec_ref(x_937); - x_940 = lean_box(0); -} -if (lean_is_scalar(x_940)) { - x_941 = lean_alloc_ctor(1, 2, 0); +x_918 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_918, 0, x_649); +lean_ctor_set(x_918, 1, x_895); +lean_ctor_set_uint64(x_918, sizeof(void*)*2, x_894); +lean_ctor_set(x_82, 0, x_918); +if (lean_is_scalar(x_912)) { + x_919 = lean_alloc_ctor(1, 2, 8); +} else { + x_919 = x_912; +} +lean_ctor_set(x_919, 0, x_81); +lean_ctor_set(x_919, 1, x_655); +lean_ctor_set_uint64(x_919, sizeof(void*)*2, x_657); +lean_ctor_set(x_79, 0, x_919); +if (lean_is_scalar(x_911)) { + x_920 = lean_alloc_ctor(1, 2, 0); +} else { + x_920 = x_911; +} +lean_ctor_set(x_920, 0, x_79); +lean_ctor_set(x_920, 1, x_651); +x_921 = lean_box(0); +x_922 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_920, x_921); +x_923 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_922, x_921); +x_924 = l_Lean_MessageData_ofList(x_923); +lean_dec(x_923); +x_925 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_926 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_926, 0, x_925); +lean_ctor_set(x_926, 1, x_924); +x_927 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_928 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_928, 0, x_926); +lean_ctor_set(x_928, 1, x_927); +x_929 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_928, x_6, x_7, x_650); +lean_dec(x_7); +x_930 = lean_ctor_get(x_929, 0); +lean_inc(x_930); +x_931 = lean_ctor_get(x_929, 1); +lean_inc(x_931); +if (lean_is_exclusive(x_929)) { + lean_ctor_release(x_929, 0); + lean_ctor_release(x_929, 1); + x_932 = x_929; +} else { + lean_dec_ref(x_929); + x_932 = lean_box(0); +} +if (lean_is_scalar(x_932)) { + x_933 = lean_alloc_ctor(1, 2, 0); } else { - x_941 = x_940; + x_933 = x_932; } -lean_ctor_set(x_941, 0, x_938); -lean_ctor_set(x_941, 1, x_939); -return x_941; +lean_ctor_set(x_933, 0, x_930); +lean_ctor_set(x_933, 1, x_931); +return x_933; } } else { -lean_object* x_942; uint8_t x_943; -lean_dec(x_666); -x_942 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__9; -x_943 = lean_string_dec_eq(x_663, x_942); -if (x_943 == 0) +lean_object* x_934; uint8_t x_935; +lean_dec(x_661); +x_934 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__11; +x_935 = lean_string_dec_eq(x_658, x_934); +if (x_935 == 0) { -if (lean_obj_tag(x_656) == 0) +lean_dec(x_3); +if (lean_obj_tag(x_651) == 0) { -lean_object* x_944; lean_object* x_945; lean_object* x_946; -lean_dec(x_919); -lean_free_object(x_78); -lean_dec(x_658); -x_944 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_944, 0, x_654); -lean_ctor_set(x_944, 1, x_903); -lean_ctor_set_uint64(x_944, sizeof(void*)*2, x_902); -lean_ctor_set(x_81, 1, x_921); -lean_ctor_set(x_81, 0, x_944); -if (lean_is_scalar(x_920)) { - x_945 = lean_alloc_ctor(1, 2, 8); -} else { - x_945 = x_920; -} -lean_ctor_set(x_945, 0, x_80); -lean_ctor_set(x_945, 1, x_660); -lean_ctor_set_uint64(x_945, sizeof(void*)*2, x_662); -x_946 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_945, x_5, x_6, x_655); +lean_object* x_936; lean_object* x_937; lean_object* x_938; +lean_dec(x_911); +lean_free_object(x_79); +lean_dec(x_653); +x_936 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_936, 0, x_649); +lean_ctor_set(x_936, 1, x_895); +lean_ctor_set_uint64(x_936, sizeof(void*)*2, x_894); +lean_ctor_set(x_82, 1, x_913); +lean_ctor_set(x_82, 0, x_936); +if (lean_is_scalar(x_912)) { + x_937 = lean_alloc_ctor(1, 2, 8); +} else { + x_937 = x_912; +} +lean_ctor_set(x_937, 0, x_81); +lean_ctor_set(x_937, 1, x_655); +lean_ctor_set_uint64(x_937, sizeof(void*)*2, x_657); +x_938 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_937, x_6, x_7, x_650); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_945); -return x_946; +lean_dec(x_937); +return x_938; } else { -lean_object* x_947; lean_object* x_948; lean_object* x_949; lean_object* x_950; lean_object* x_951; lean_object* x_952; lean_object* x_953; lean_object* x_954; lean_object* x_955; lean_object* x_956; lean_object* x_957; lean_object* x_958; lean_object* x_959; lean_object* x_960; lean_object* x_961; lean_object* x_962; +lean_object* x_939; lean_object* x_940; lean_object* x_941; lean_object* x_942; lean_object* x_943; lean_object* x_944; lean_object* x_945; lean_object* x_946; lean_object* x_947; lean_object* x_948; lean_object* x_949; lean_object* x_950; lean_object* x_951; lean_object* x_952; lean_object* x_953; lean_object* x_954; lean_dec(x_1); -x_947 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_947, 0, x_654); -lean_ctor_set(x_947, 1, x_903); -lean_ctor_set_uint64(x_947, sizeof(void*)*2, x_902); -lean_ctor_set(x_81, 1, x_921); -lean_ctor_set(x_81, 0, x_947); -if (lean_is_scalar(x_920)) { - x_948 = lean_alloc_ctor(1, 2, 8); -} else { - x_948 = x_920; -} -lean_ctor_set(x_948, 0, x_80); -lean_ctor_set(x_948, 1, x_660); -lean_ctor_set_uint64(x_948, sizeof(void*)*2, x_662); -lean_ctor_set(x_78, 0, x_948); -if (lean_is_scalar(x_919)) { - x_949 = lean_alloc_ctor(1, 2, 0); -} else { - x_949 = x_919; -} -lean_ctor_set(x_949, 0, x_78); -lean_ctor_set(x_949, 1, x_656); -x_950 = lean_box(0); -x_951 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_949, x_950); -x_952 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_951, x_950); -x_953 = l_Lean_MessageData_ofList(x_952); -lean_dec(x_952); -x_954 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_955 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_955, 0, x_954); -lean_ctor_set(x_955, 1, x_953); -x_956 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_957 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_957, 0, x_955); -lean_ctor_set(x_957, 1, x_956); -x_958 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_957, x_5, x_6, x_655); -lean_dec(x_6); -x_959 = lean_ctor_get(x_958, 0); -lean_inc(x_959); -x_960 = lean_ctor_get(x_958, 1); -lean_inc(x_960); -if (lean_is_exclusive(x_958)) { - lean_ctor_release(x_958, 0); - lean_ctor_release(x_958, 1); - x_961 = x_958; +x_939 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_939, 0, x_649); +lean_ctor_set(x_939, 1, x_895); +lean_ctor_set_uint64(x_939, sizeof(void*)*2, x_894); +lean_ctor_set(x_82, 1, x_913); +lean_ctor_set(x_82, 0, x_939); +if (lean_is_scalar(x_912)) { + x_940 = lean_alloc_ctor(1, 2, 8); +} else { + x_940 = x_912; +} +lean_ctor_set(x_940, 0, x_81); +lean_ctor_set(x_940, 1, x_655); +lean_ctor_set_uint64(x_940, sizeof(void*)*2, x_657); +lean_ctor_set(x_79, 0, x_940); +if (lean_is_scalar(x_911)) { + x_941 = lean_alloc_ctor(1, 2, 0); } else { - lean_dec_ref(x_958); - x_961 = lean_box(0); + x_941 = x_911; } -if (lean_is_scalar(x_961)) { - x_962 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_941, 0, x_79); +lean_ctor_set(x_941, 1, x_651); +x_942 = lean_box(0); +x_943 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_941, x_942); +x_944 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_943, x_942); +x_945 = l_Lean_MessageData_ofList(x_944); +lean_dec(x_944); +x_946 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_947 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_947, 0, x_946); +lean_ctor_set(x_947, 1, x_945); +x_948 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_949 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_949, 0, x_947); +lean_ctor_set(x_949, 1, x_948); +x_950 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_949, x_6, x_7, x_650); +lean_dec(x_7); +x_951 = lean_ctor_get(x_950, 0); +lean_inc(x_951); +x_952 = lean_ctor_get(x_950, 1); +lean_inc(x_952); +if (lean_is_exclusive(x_950)) { + lean_ctor_release(x_950, 0); + lean_ctor_release(x_950, 1); + x_953 = x_950; +} else { + lean_dec_ref(x_950); + x_953 = lean_box(0); +} +if (lean_is_scalar(x_953)) { + x_954 = lean_alloc_ctor(1, 2, 0); } else { - x_962 = x_961; + x_954 = x_953; } -lean_ctor_set(x_962, 0, x_959); -lean_ctor_set(x_962, 1, x_960); -return x_962; +lean_ctor_set(x_954, 0, x_951); +lean_ctor_set(x_954, 1, x_952); +return x_954; } } else { -lean_object* x_963; uint8_t x_964; -lean_dec(x_663); -x_963 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__7; -x_964 = lean_string_dec_eq(x_660, x_963); -if (x_964 == 0) +lean_object* x_955; uint8_t x_956; +lean_dec(x_658); +x_955 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__7; +x_956 = lean_string_dec_eq(x_655, x_955); +if (x_956 == 0) { -if (lean_obj_tag(x_656) == 0) +lean_dec(x_3); +if (lean_obj_tag(x_651) == 0) { -lean_object* x_965; lean_object* x_966; lean_object* x_967; -lean_dec(x_919); -lean_free_object(x_78); -lean_dec(x_658); -x_965 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_965, 0, x_654); -lean_ctor_set(x_965, 1, x_903); -lean_ctor_set_uint64(x_965, sizeof(void*)*2, x_902); -lean_ctor_set(x_81, 1, x_921); -lean_ctor_set(x_81, 0, x_965); -lean_ctor_set(x_80, 1, x_942); -if (lean_is_scalar(x_920)) { - x_966 = lean_alloc_ctor(1, 2, 8); -} else { - x_966 = x_920; -} -lean_ctor_set(x_966, 0, x_80); -lean_ctor_set(x_966, 1, x_660); -lean_ctor_set_uint64(x_966, sizeof(void*)*2, x_662); -x_967 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_966, x_5, x_6, x_655); +lean_object* x_957; lean_object* x_958; lean_object* x_959; +lean_dec(x_911); +lean_free_object(x_79); +lean_dec(x_653); +x_957 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_957, 0, x_649); +lean_ctor_set(x_957, 1, x_895); +lean_ctor_set_uint64(x_957, sizeof(void*)*2, x_894); +lean_ctor_set(x_82, 1, x_913); +lean_ctor_set(x_82, 0, x_957); +lean_ctor_set(x_81, 1, x_934); +if (lean_is_scalar(x_912)) { + x_958 = lean_alloc_ctor(1, 2, 8); +} else { + x_958 = x_912; +} +lean_ctor_set(x_958, 0, x_81); +lean_ctor_set(x_958, 1, x_655); +lean_ctor_set_uint64(x_958, sizeof(void*)*2, x_657); +x_959 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_958, x_6, x_7, x_650); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_966); -return x_967; +lean_dec(x_958); +return x_959; } else { -lean_object* x_968; lean_object* x_969; lean_object* x_970; lean_object* x_971; lean_object* x_972; lean_object* x_973; lean_object* x_974; lean_object* x_975; lean_object* x_976; lean_object* x_977; lean_object* x_978; lean_object* x_979; lean_object* x_980; lean_object* x_981; lean_object* x_982; lean_object* x_983; +lean_object* x_960; lean_object* x_961; lean_object* x_962; lean_object* x_963; lean_object* x_964; lean_object* x_965; lean_object* x_966; lean_object* x_967; lean_object* x_968; lean_object* x_969; lean_object* x_970; lean_object* x_971; lean_object* x_972; lean_object* x_973; lean_object* x_974; lean_object* x_975; lean_dec(x_1); -x_968 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_968, 0, x_654); -lean_ctor_set(x_968, 1, x_903); -lean_ctor_set_uint64(x_968, sizeof(void*)*2, x_902); -lean_ctor_set(x_81, 1, x_921); -lean_ctor_set(x_81, 0, x_968); -lean_ctor_set(x_80, 1, x_942); -if (lean_is_scalar(x_920)) { - x_969 = lean_alloc_ctor(1, 2, 8); -} else { - x_969 = x_920; -} -lean_ctor_set(x_969, 0, x_80); -lean_ctor_set(x_969, 1, x_660); -lean_ctor_set_uint64(x_969, sizeof(void*)*2, x_662); -lean_ctor_set(x_78, 0, x_969); -if (lean_is_scalar(x_919)) { - x_970 = lean_alloc_ctor(1, 2, 0); -} else { - x_970 = x_919; -} -lean_ctor_set(x_970, 0, x_78); -lean_ctor_set(x_970, 1, x_656); -x_971 = lean_box(0); -x_972 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_970, x_971); -x_973 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_972, x_971); -x_974 = l_Lean_MessageData_ofList(x_973); -lean_dec(x_973); -x_975 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_976 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_976, 0, x_975); -lean_ctor_set(x_976, 1, x_974); -x_977 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_978 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_978, 0, x_976); -lean_ctor_set(x_978, 1, x_977); -x_979 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_978, x_5, x_6, x_655); -lean_dec(x_6); -x_980 = lean_ctor_get(x_979, 0); -lean_inc(x_980); -x_981 = lean_ctor_get(x_979, 1); -lean_inc(x_981); -if (lean_is_exclusive(x_979)) { - lean_ctor_release(x_979, 0); - lean_ctor_release(x_979, 1); - x_982 = x_979; +x_960 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_960, 0, x_649); +lean_ctor_set(x_960, 1, x_895); +lean_ctor_set_uint64(x_960, sizeof(void*)*2, x_894); +lean_ctor_set(x_82, 1, x_913); +lean_ctor_set(x_82, 0, x_960); +lean_ctor_set(x_81, 1, x_934); +if (lean_is_scalar(x_912)) { + x_961 = lean_alloc_ctor(1, 2, 8); +} else { + x_961 = x_912; +} +lean_ctor_set(x_961, 0, x_81); +lean_ctor_set(x_961, 1, x_655); +lean_ctor_set_uint64(x_961, sizeof(void*)*2, x_657); +lean_ctor_set(x_79, 0, x_961); +if (lean_is_scalar(x_911)) { + x_962 = lean_alloc_ctor(1, 2, 0); +} else { + x_962 = x_911; +} +lean_ctor_set(x_962, 0, x_79); +lean_ctor_set(x_962, 1, x_651); +x_963 = lean_box(0); +x_964 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_962, x_963); +x_965 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_964, x_963); +x_966 = l_Lean_MessageData_ofList(x_965); +lean_dec(x_965); +x_967 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_968 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_968, 0, x_967); +lean_ctor_set(x_968, 1, x_966); +x_969 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_970 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_970, 0, x_968); +lean_ctor_set(x_970, 1, x_969); +x_971 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_970, x_6, x_7, x_650); +lean_dec(x_7); +x_972 = lean_ctor_get(x_971, 0); +lean_inc(x_972); +x_973 = lean_ctor_get(x_971, 1); +lean_inc(x_973); +if (lean_is_exclusive(x_971)) { + lean_ctor_release(x_971, 0); + lean_ctor_release(x_971, 1); + x_974 = x_971; } else { - lean_dec_ref(x_979); - x_982 = lean_box(0); + lean_dec_ref(x_971); + x_974 = lean_box(0); } -if (lean_is_scalar(x_982)) { - x_983 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_974)) { + x_975 = lean_alloc_ctor(1, 2, 0); } else { - x_983 = x_982; + x_975 = x_974; } -lean_ctor_set(x_983, 0, x_980); -lean_ctor_set(x_983, 1, x_981); -return x_983; +lean_ctor_set(x_975, 0, x_972); +lean_ctor_set(x_975, 1, x_973); +return x_975; } } else { -lean_dec(x_660); -if (lean_obj_tag(x_656) == 0) +lean_dec(x_655); +if (lean_obj_tag(x_651) == 0) { -lean_object* x_984; lean_object* x_985; -lean_dec(x_920); -lean_dec(x_919); +lean_object* x_976; +lean_dec(x_912); +lean_dec(x_911); +lean_free_object(x_82); lean_free_object(x_81); -lean_free_object(x_80); -lean_free_object(x_78); -lean_dec(x_658); -x_984 = l_Lean_identKind; -x_985 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_984, x_5, x_6, x_655); +lean_free_object(x_79); +lean_dec(x_653); +x_976 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_3, x_6, x_7, x_650); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -return x_985; +lean_dec(x_3); +return x_976; } else { -lean_object* x_986; lean_object* x_987; lean_object* x_988; lean_object* x_989; lean_object* x_990; lean_object* x_991; lean_object* x_992; lean_object* x_993; lean_object* x_994; lean_object* x_995; lean_object* x_996; lean_object* x_997; lean_object* x_998; lean_object* x_999; lean_object* x_1000; lean_object* x_1001; +lean_object* x_977; lean_object* x_978; lean_object* x_979; lean_object* x_980; lean_object* x_981; lean_object* x_982; lean_object* x_983; lean_object* x_984; lean_object* x_985; lean_object* x_986; lean_object* x_987; lean_object* x_988; lean_object* x_989; lean_object* x_990; lean_object* x_991; lean_object* x_992; +lean_dec(x_3); lean_dec(x_1); -x_986 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_986, 0, x_654); -lean_ctor_set(x_986, 1, x_903); -lean_ctor_set_uint64(x_986, sizeof(void*)*2, x_902); -lean_ctor_set(x_81, 1, x_921); -lean_ctor_set(x_81, 0, x_986); -lean_ctor_set(x_80, 1, x_942); -if (lean_is_scalar(x_920)) { - x_987 = lean_alloc_ctor(1, 2, 8); -} else { - x_987 = x_920; -} -lean_ctor_set(x_987, 0, x_80); -lean_ctor_set(x_987, 1, x_963); -lean_ctor_set_uint64(x_987, sizeof(void*)*2, x_662); -lean_ctor_set(x_78, 0, x_987); -if (lean_is_scalar(x_919)) { - x_988 = lean_alloc_ctor(1, 2, 0); -} else { - x_988 = x_919; -} -lean_ctor_set(x_988, 0, x_78); -lean_ctor_set(x_988, 1, x_656); -x_989 = lean_box(0); -x_990 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_988, x_989); -x_991 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_990, x_989); -x_992 = l_Lean_MessageData_ofList(x_991); -lean_dec(x_991); -x_993 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_994 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_994, 0, x_993); -lean_ctor_set(x_994, 1, x_992); -x_995 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_996 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_996, 0, x_994); -lean_ctor_set(x_996, 1, x_995); -x_997 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_996, x_5, x_6, x_655); -lean_dec(x_6); -x_998 = lean_ctor_get(x_997, 0); -lean_inc(x_998); -x_999 = lean_ctor_get(x_997, 1); -lean_inc(x_999); -if (lean_is_exclusive(x_997)) { - lean_ctor_release(x_997, 0); - lean_ctor_release(x_997, 1); - x_1000 = x_997; +x_977 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_977, 0, x_649); +lean_ctor_set(x_977, 1, x_895); +lean_ctor_set_uint64(x_977, sizeof(void*)*2, x_894); +lean_ctor_set(x_82, 1, x_913); +lean_ctor_set(x_82, 0, x_977); +lean_ctor_set(x_81, 1, x_934); +if (lean_is_scalar(x_912)) { + x_978 = lean_alloc_ctor(1, 2, 8); +} else { + x_978 = x_912; +} +lean_ctor_set(x_978, 0, x_81); +lean_ctor_set(x_978, 1, x_955); +lean_ctor_set_uint64(x_978, sizeof(void*)*2, x_657); +lean_ctor_set(x_79, 0, x_978); +if (lean_is_scalar(x_911)) { + x_979 = lean_alloc_ctor(1, 2, 0); +} else { + x_979 = x_911; +} +lean_ctor_set(x_979, 0, x_79); +lean_ctor_set(x_979, 1, x_651); +x_980 = lean_box(0); +x_981 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_979, x_980); +x_982 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_981, x_980); +x_983 = l_Lean_MessageData_ofList(x_982); +lean_dec(x_982); +x_984 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_985 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_985, 0, x_984); +lean_ctor_set(x_985, 1, x_983); +x_986 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_987 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_987, 0, x_985); +lean_ctor_set(x_987, 1, x_986); +x_988 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_987, x_6, x_7, x_650); +lean_dec(x_7); +x_989 = lean_ctor_get(x_988, 0); +lean_inc(x_989); +x_990 = lean_ctor_get(x_988, 1); +lean_inc(x_990); +if (lean_is_exclusive(x_988)) { + lean_ctor_release(x_988, 0); + lean_ctor_release(x_988, 1); + x_991 = x_988; } else { - lean_dec_ref(x_997); - x_1000 = lean_box(0); + lean_dec_ref(x_988); + x_991 = lean_box(0); } -if (lean_is_scalar(x_1000)) { - x_1001 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_991)) { + x_992 = lean_alloc_ctor(1, 2, 0); } else { - x_1001 = x_1000; + x_992 = x_991; } -lean_ctor_set(x_1001, 0, x_998); -lean_ctor_set(x_1001, 1, x_999); -return x_1001; +lean_ctor_set(x_992, 0, x_989); +lean_ctor_set(x_992, 1, x_990); +return x_992; } } } @@ -6285,524 +6342,529 @@ return x_1001; } else { -lean_object* x_1002; uint64_t x_1003; lean_object* x_1004; uint64_t x_1005; lean_object* x_1006; lean_object* x_1007; uint8_t x_1008; -x_1002 = lean_ctor_get(x_81, 1); -x_1003 = lean_ctor_get_uint64(x_81, sizeof(void*)*2); -lean_inc(x_1002); -lean_dec(x_81); -x_1004 = lean_ctor_get(x_82, 1); -lean_inc(x_1004); -x_1005 = lean_ctor_get_uint64(x_82, sizeof(void*)*2); -if (lean_is_exclusive(x_82)) { - lean_ctor_release(x_82, 0); - lean_ctor_release(x_82, 1); - x_1006 = x_82; -} else { - lean_dec_ref(x_82); - x_1006 = lean_box(0); -} -x_1007 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__1; -x_1008 = lean_string_dec_eq(x_1004, x_1007); -lean_dec(x_1004); -if (x_1008 == 0) -{ -lean_dec(x_1006); -lean_dec(x_1002); -lean_free_object(x_80); -lean_dec(x_663); -lean_dec(x_660); -lean_free_object(x_78); +lean_object* x_993; uint64_t x_994; lean_object* x_995; uint64_t x_996; lean_object* x_997; lean_object* x_998; uint8_t x_999; +x_993 = lean_ctor_get(x_82, 1); +x_994 = lean_ctor_get_uint64(x_82, sizeof(void*)*2); +lean_inc(x_993); +lean_dec(x_82); +x_995 = lean_ctor_get(x_83, 1); +lean_inc(x_995); +x_996 = lean_ctor_get_uint64(x_83, sizeof(void*)*2); +if (lean_is_exclusive(x_83)) { + lean_ctor_release(x_83, 0); + lean_ctor_release(x_83, 1); + x_997 = x_83; +} else { + lean_dec_ref(x_83); + x_997 = lean_box(0); +} +x_998 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__1; +x_999 = lean_string_dec_eq(x_995, x_998); +lean_dec(x_995); +if (x_999 == 0) +{ +lean_dec(x_997); +lean_dec(x_993); +lean_free_object(x_81); lean_dec(x_658); -if (lean_obj_tag(x_656) == 0) +lean_dec(x_655); +lean_free_object(x_79); +lean_dec(x_653); +lean_dec(x_3); +if (lean_obj_tag(x_651) == 0) { -lean_object* x_1009; -lean_dec(x_9); -x_1009 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_79, x_5, x_6, x_655); +lean_object* x_1000; +lean_dec(x_10); +x_1000 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_80, x_6, x_7, x_650); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_79); -return x_1009; +lean_dec(x_80); +return x_1000; } else { -lean_object* x_1010; lean_object* x_1011; lean_object* x_1012; lean_object* x_1013; lean_object* x_1014; lean_object* x_1015; lean_object* x_1016; lean_object* x_1017; lean_object* x_1018; lean_object* x_1019; lean_object* x_1020; lean_object* x_1021; lean_object* x_1022; -lean_dec(x_656); -lean_dec(x_79); +lean_object* x_1001; lean_object* x_1002; lean_object* x_1003; lean_object* x_1004; lean_object* x_1005; lean_object* x_1006; lean_object* x_1007; lean_object* x_1008; lean_object* x_1009; lean_object* x_1010; lean_object* x_1011; lean_object* x_1012; lean_object* x_1013; +lean_dec(x_651); +lean_dec(x_80); lean_dec(x_1); -x_1010 = lean_box(0); -x_1011 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_9, x_1010); -x_1012 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1011, x_1010); -x_1013 = l_Lean_MessageData_ofList(x_1012); -lean_dec(x_1012); -x_1014 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_1015 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1015, 0, x_1014); -lean_ctor_set(x_1015, 1, x_1013); -x_1016 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_1017 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1017, 0, x_1015); -lean_ctor_set(x_1017, 1, x_1016); -x_1018 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1017, x_5, x_6, x_655); -lean_dec(x_6); -x_1019 = lean_ctor_get(x_1018, 0); -lean_inc(x_1019); -x_1020 = lean_ctor_get(x_1018, 1); -lean_inc(x_1020); -if (lean_is_exclusive(x_1018)) { - lean_ctor_release(x_1018, 0); - lean_ctor_release(x_1018, 1); - x_1021 = x_1018; +x_1001 = lean_box(0); +x_1002 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_10, x_1001); +x_1003 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1002, x_1001); +x_1004 = l_Lean_MessageData_ofList(x_1003); +lean_dec(x_1003); +x_1005 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_1006 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1006, 0, x_1005); +lean_ctor_set(x_1006, 1, x_1004); +x_1007 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_1008 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1008, 0, x_1006); +lean_ctor_set(x_1008, 1, x_1007); +x_1009 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1008, x_6, x_7, x_650); +lean_dec(x_7); +x_1010 = lean_ctor_get(x_1009, 0); +lean_inc(x_1010); +x_1011 = lean_ctor_get(x_1009, 1); +lean_inc(x_1011); +if (lean_is_exclusive(x_1009)) { + lean_ctor_release(x_1009, 0); + lean_ctor_release(x_1009, 1); + x_1012 = x_1009; } else { - lean_dec_ref(x_1018); - x_1021 = lean_box(0); + lean_dec_ref(x_1009); + x_1012 = lean_box(0); } -if (lean_is_scalar(x_1021)) { - x_1022 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_1012)) { + x_1013 = lean_alloc_ctor(1, 2, 0); } else { - x_1022 = x_1021; + x_1013 = x_1012; } -lean_ctor_set(x_1022, 0, x_1019); -lean_ctor_set(x_1022, 1, x_1020); -return x_1022; +lean_ctor_set(x_1013, 0, x_1010); +lean_ctor_set(x_1013, 1, x_1011); +return x_1013; } } else { -lean_object* x_1023; lean_object* x_1024; lean_object* x_1025; uint8_t x_1026; -if (lean_is_exclusive(x_9)) { - lean_ctor_release(x_9, 0); - lean_ctor_release(x_9, 1); - x_1023 = x_9; +lean_object* x_1014; lean_object* x_1015; lean_object* x_1016; uint8_t x_1017; +if (lean_is_exclusive(x_10)) { + lean_ctor_release(x_10, 0); + lean_ctor_release(x_10, 1); + x_1014 = x_10; } else { - lean_dec_ref(x_9); - x_1023 = lean_box(0); + lean_dec_ref(x_10); + x_1014 = lean_box(0); } -if (lean_is_exclusive(x_79)) { - lean_ctor_release(x_79, 0); - lean_ctor_release(x_79, 1); - x_1024 = x_79; +if (lean_is_exclusive(x_80)) { + lean_ctor_release(x_80, 0); + lean_ctor_release(x_80, 1); + x_1015 = x_80; } else { - lean_dec_ref(x_79); - x_1024 = lean_box(0); + lean_dec_ref(x_80); + x_1015 = lean_box(0); } -x_1025 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__4; -x_1026 = lean_string_dec_eq(x_1002, x_1025); -if (x_1026 == 0) +x_1016 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__4; +x_1017 = lean_string_dec_eq(x_993, x_1016); +if (x_1017 == 0) { -if (lean_obj_tag(x_656) == 0) +lean_dec(x_3); +if (lean_obj_tag(x_651) == 0) { -lean_object* x_1027; lean_object* x_1028; lean_object* x_1029; lean_object* x_1030; -lean_dec(x_1023); -lean_free_object(x_78); -lean_dec(x_658); -if (lean_is_scalar(x_1006)) { - x_1027 = lean_alloc_ctor(1, 2, 8); -} else { - x_1027 = x_1006; -} -lean_ctor_set(x_1027, 0, x_654); -lean_ctor_set(x_1027, 1, x_1007); -lean_ctor_set_uint64(x_1027, sizeof(void*)*2, x_1005); -x_1028 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_1028, 0, x_1027); -lean_ctor_set(x_1028, 1, x_1002); -lean_ctor_set_uint64(x_1028, sizeof(void*)*2, x_1003); -lean_ctor_set(x_80, 0, x_1028); -if (lean_is_scalar(x_1024)) { - x_1029 = lean_alloc_ctor(1, 2, 8); -} else { - x_1029 = x_1024; -} -lean_ctor_set(x_1029, 0, x_80); -lean_ctor_set(x_1029, 1, x_660); -lean_ctor_set_uint64(x_1029, sizeof(void*)*2, x_662); -x_1030 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_1029, x_5, x_6, x_655); +lean_object* x_1018; lean_object* x_1019; lean_object* x_1020; lean_object* x_1021; +lean_dec(x_1014); +lean_free_object(x_79); +lean_dec(x_653); +if (lean_is_scalar(x_997)) { + x_1018 = lean_alloc_ctor(1, 2, 8); +} else { + x_1018 = x_997; +} +lean_ctor_set(x_1018, 0, x_649); +lean_ctor_set(x_1018, 1, x_998); +lean_ctor_set_uint64(x_1018, sizeof(void*)*2, x_996); +x_1019 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_1019, 0, x_1018); +lean_ctor_set(x_1019, 1, x_993); +lean_ctor_set_uint64(x_1019, sizeof(void*)*2, x_994); +lean_ctor_set(x_81, 0, x_1019); +if (lean_is_scalar(x_1015)) { + x_1020 = lean_alloc_ctor(1, 2, 8); +} else { + x_1020 = x_1015; +} +lean_ctor_set(x_1020, 0, x_81); +lean_ctor_set(x_1020, 1, x_655); +lean_ctor_set_uint64(x_1020, sizeof(void*)*2, x_657); +x_1021 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_1020, x_6, x_7, x_650); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_1029); -return x_1030; +lean_dec(x_1020); +return x_1021; } else { -lean_object* x_1031; lean_object* x_1032; lean_object* x_1033; lean_object* x_1034; lean_object* x_1035; lean_object* x_1036; lean_object* x_1037; lean_object* x_1038; lean_object* x_1039; lean_object* x_1040; lean_object* x_1041; lean_object* x_1042; lean_object* x_1043; lean_object* x_1044; lean_object* x_1045; lean_object* x_1046; lean_object* x_1047; +lean_object* x_1022; lean_object* x_1023; lean_object* x_1024; lean_object* x_1025; lean_object* x_1026; lean_object* x_1027; lean_object* x_1028; lean_object* x_1029; lean_object* x_1030; lean_object* x_1031; lean_object* x_1032; lean_object* x_1033; lean_object* x_1034; lean_object* x_1035; lean_object* x_1036; lean_object* x_1037; lean_object* x_1038; lean_dec(x_1); -if (lean_is_scalar(x_1006)) { - x_1031 = lean_alloc_ctor(1, 2, 8); -} else { - x_1031 = x_1006; -} -lean_ctor_set(x_1031, 0, x_654); -lean_ctor_set(x_1031, 1, x_1007); -lean_ctor_set_uint64(x_1031, sizeof(void*)*2, x_1005); -x_1032 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_1032, 0, x_1031); -lean_ctor_set(x_1032, 1, x_1002); -lean_ctor_set_uint64(x_1032, sizeof(void*)*2, x_1003); -lean_ctor_set(x_80, 0, x_1032); -if (lean_is_scalar(x_1024)) { - x_1033 = lean_alloc_ctor(1, 2, 8); -} else { - x_1033 = x_1024; -} -lean_ctor_set(x_1033, 0, x_80); -lean_ctor_set(x_1033, 1, x_660); -lean_ctor_set_uint64(x_1033, sizeof(void*)*2, x_662); -lean_ctor_set(x_78, 0, x_1033); -if (lean_is_scalar(x_1023)) { - x_1034 = lean_alloc_ctor(1, 2, 0); -} else { - x_1034 = x_1023; -} -lean_ctor_set(x_1034, 0, x_78); -lean_ctor_set(x_1034, 1, x_656); -x_1035 = lean_box(0); -x_1036 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_1034, x_1035); -x_1037 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1036, x_1035); -x_1038 = l_Lean_MessageData_ofList(x_1037); -lean_dec(x_1037); -x_1039 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_1040 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1040, 0, x_1039); -lean_ctor_set(x_1040, 1, x_1038); -x_1041 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_1042 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1042, 0, x_1040); -lean_ctor_set(x_1042, 1, x_1041); -x_1043 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1042, x_5, x_6, x_655); -lean_dec(x_6); -x_1044 = lean_ctor_get(x_1043, 0); -lean_inc(x_1044); -x_1045 = lean_ctor_get(x_1043, 1); -lean_inc(x_1045); -if (lean_is_exclusive(x_1043)) { - lean_ctor_release(x_1043, 0); - lean_ctor_release(x_1043, 1); - x_1046 = x_1043; +if (lean_is_scalar(x_997)) { + x_1022 = lean_alloc_ctor(1, 2, 8); +} else { + x_1022 = x_997; +} +lean_ctor_set(x_1022, 0, x_649); +lean_ctor_set(x_1022, 1, x_998); +lean_ctor_set_uint64(x_1022, sizeof(void*)*2, x_996); +x_1023 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_1023, 0, x_1022); +lean_ctor_set(x_1023, 1, x_993); +lean_ctor_set_uint64(x_1023, sizeof(void*)*2, x_994); +lean_ctor_set(x_81, 0, x_1023); +if (lean_is_scalar(x_1015)) { + x_1024 = lean_alloc_ctor(1, 2, 8); +} else { + x_1024 = x_1015; +} +lean_ctor_set(x_1024, 0, x_81); +lean_ctor_set(x_1024, 1, x_655); +lean_ctor_set_uint64(x_1024, sizeof(void*)*2, x_657); +lean_ctor_set(x_79, 0, x_1024); +if (lean_is_scalar(x_1014)) { + x_1025 = lean_alloc_ctor(1, 2, 0); +} else { + x_1025 = x_1014; +} +lean_ctor_set(x_1025, 0, x_79); +lean_ctor_set(x_1025, 1, x_651); +x_1026 = lean_box(0); +x_1027 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_1025, x_1026); +x_1028 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1027, x_1026); +x_1029 = l_Lean_MessageData_ofList(x_1028); +lean_dec(x_1028); +x_1030 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_1031 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1031, 0, x_1030); +lean_ctor_set(x_1031, 1, x_1029); +x_1032 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_1033 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1033, 0, x_1031); +lean_ctor_set(x_1033, 1, x_1032); +x_1034 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1033, x_6, x_7, x_650); +lean_dec(x_7); +x_1035 = lean_ctor_get(x_1034, 0); +lean_inc(x_1035); +x_1036 = lean_ctor_get(x_1034, 1); +lean_inc(x_1036); +if (lean_is_exclusive(x_1034)) { + lean_ctor_release(x_1034, 0); + lean_ctor_release(x_1034, 1); + x_1037 = x_1034; } else { - lean_dec_ref(x_1043); - x_1046 = lean_box(0); + lean_dec_ref(x_1034); + x_1037 = lean_box(0); } -if (lean_is_scalar(x_1046)) { - x_1047 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_1037)) { + x_1038 = lean_alloc_ctor(1, 2, 0); } else { - x_1047 = x_1046; + x_1038 = x_1037; } -lean_ctor_set(x_1047, 0, x_1044); -lean_ctor_set(x_1047, 1, x_1045); -return x_1047; +lean_ctor_set(x_1038, 0, x_1035); +lean_ctor_set(x_1038, 1, x_1036); +return x_1038; } } else { -lean_object* x_1048; uint8_t x_1049; -lean_dec(x_1002); -x_1048 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__9; -x_1049 = lean_string_dec_eq(x_663, x_1048); -if (x_1049 == 0) +lean_object* x_1039; uint8_t x_1040; +lean_dec(x_993); +x_1039 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__11; +x_1040 = lean_string_dec_eq(x_658, x_1039); +if (x_1040 == 0) { -if (lean_obj_tag(x_656) == 0) +lean_dec(x_3); +if (lean_obj_tag(x_651) == 0) { -lean_object* x_1050; lean_object* x_1051; lean_object* x_1052; lean_object* x_1053; -lean_dec(x_1023); -lean_free_object(x_78); -lean_dec(x_658); -if (lean_is_scalar(x_1006)) { - x_1050 = lean_alloc_ctor(1, 2, 8); -} else { - x_1050 = x_1006; -} -lean_ctor_set(x_1050, 0, x_654); -lean_ctor_set(x_1050, 1, x_1007); -lean_ctor_set_uint64(x_1050, sizeof(void*)*2, x_1005); -x_1051 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_1051, 0, x_1050); -lean_ctor_set(x_1051, 1, x_1025); -lean_ctor_set_uint64(x_1051, sizeof(void*)*2, x_1003); -lean_ctor_set(x_80, 0, x_1051); -if (lean_is_scalar(x_1024)) { - x_1052 = lean_alloc_ctor(1, 2, 8); -} else { - x_1052 = x_1024; -} -lean_ctor_set(x_1052, 0, x_80); -lean_ctor_set(x_1052, 1, x_660); -lean_ctor_set_uint64(x_1052, sizeof(void*)*2, x_662); -x_1053 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_1052, x_5, x_6, x_655); +lean_object* x_1041; lean_object* x_1042; lean_object* x_1043; lean_object* x_1044; +lean_dec(x_1014); +lean_free_object(x_79); +lean_dec(x_653); +if (lean_is_scalar(x_997)) { + x_1041 = lean_alloc_ctor(1, 2, 8); +} else { + x_1041 = x_997; +} +lean_ctor_set(x_1041, 0, x_649); +lean_ctor_set(x_1041, 1, x_998); +lean_ctor_set_uint64(x_1041, sizeof(void*)*2, x_996); +x_1042 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_1042, 0, x_1041); +lean_ctor_set(x_1042, 1, x_1016); +lean_ctor_set_uint64(x_1042, sizeof(void*)*2, x_994); +lean_ctor_set(x_81, 0, x_1042); +if (lean_is_scalar(x_1015)) { + x_1043 = lean_alloc_ctor(1, 2, 8); +} else { + x_1043 = x_1015; +} +lean_ctor_set(x_1043, 0, x_81); +lean_ctor_set(x_1043, 1, x_655); +lean_ctor_set_uint64(x_1043, sizeof(void*)*2, x_657); +x_1044 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_1043, x_6, x_7, x_650); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_1052); -return x_1053; +lean_dec(x_1043); +return x_1044; } else { -lean_object* x_1054; lean_object* x_1055; lean_object* x_1056; lean_object* x_1057; lean_object* x_1058; lean_object* x_1059; lean_object* x_1060; lean_object* x_1061; lean_object* x_1062; lean_object* x_1063; lean_object* x_1064; lean_object* x_1065; lean_object* x_1066; lean_object* x_1067; lean_object* x_1068; lean_object* x_1069; lean_object* x_1070; +lean_object* x_1045; lean_object* x_1046; lean_object* x_1047; lean_object* x_1048; lean_object* x_1049; lean_object* x_1050; lean_object* x_1051; lean_object* x_1052; lean_object* x_1053; lean_object* x_1054; lean_object* x_1055; lean_object* x_1056; lean_object* x_1057; lean_object* x_1058; lean_object* x_1059; lean_object* x_1060; lean_object* x_1061; lean_dec(x_1); -if (lean_is_scalar(x_1006)) { - x_1054 = lean_alloc_ctor(1, 2, 8); -} else { - x_1054 = x_1006; -} -lean_ctor_set(x_1054, 0, x_654); -lean_ctor_set(x_1054, 1, x_1007); -lean_ctor_set_uint64(x_1054, sizeof(void*)*2, x_1005); -x_1055 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_1055, 0, x_1054); -lean_ctor_set(x_1055, 1, x_1025); -lean_ctor_set_uint64(x_1055, sizeof(void*)*2, x_1003); -lean_ctor_set(x_80, 0, x_1055); -if (lean_is_scalar(x_1024)) { - x_1056 = lean_alloc_ctor(1, 2, 8); -} else { - x_1056 = x_1024; -} -lean_ctor_set(x_1056, 0, x_80); -lean_ctor_set(x_1056, 1, x_660); -lean_ctor_set_uint64(x_1056, sizeof(void*)*2, x_662); -lean_ctor_set(x_78, 0, x_1056); -if (lean_is_scalar(x_1023)) { - x_1057 = lean_alloc_ctor(1, 2, 0); -} else { - x_1057 = x_1023; -} -lean_ctor_set(x_1057, 0, x_78); -lean_ctor_set(x_1057, 1, x_656); -x_1058 = lean_box(0); -x_1059 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_1057, x_1058); -x_1060 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1059, x_1058); -x_1061 = l_Lean_MessageData_ofList(x_1060); -lean_dec(x_1060); -x_1062 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_1063 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1063, 0, x_1062); -lean_ctor_set(x_1063, 1, x_1061); -x_1064 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_1065 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1065, 0, x_1063); -lean_ctor_set(x_1065, 1, x_1064); -x_1066 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1065, x_5, x_6, x_655); -lean_dec(x_6); -x_1067 = lean_ctor_get(x_1066, 0); -lean_inc(x_1067); -x_1068 = lean_ctor_get(x_1066, 1); -lean_inc(x_1068); -if (lean_is_exclusive(x_1066)) { - lean_ctor_release(x_1066, 0); - lean_ctor_release(x_1066, 1); - x_1069 = x_1066; +if (lean_is_scalar(x_997)) { + x_1045 = lean_alloc_ctor(1, 2, 8); +} else { + x_1045 = x_997; +} +lean_ctor_set(x_1045, 0, x_649); +lean_ctor_set(x_1045, 1, x_998); +lean_ctor_set_uint64(x_1045, sizeof(void*)*2, x_996); +x_1046 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_1046, 0, x_1045); +lean_ctor_set(x_1046, 1, x_1016); +lean_ctor_set_uint64(x_1046, sizeof(void*)*2, x_994); +lean_ctor_set(x_81, 0, x_1046); +if (lean_is_scalar(x_1015)) { + x_1047 = lean_alloc_ctor(1, 2, 8); +} else { + x_1047 = x_1015; +} +lean_ctor_set(x_1047, 0, x_81); +lean_ctor_set(x_1047, 1, x_655); +lean_ctor_set_uint64(x_1047, sizeof(void*)*2, x_657); +lean_ctor_set(x_79, 0, x_1047); +if (lean_is_scalar(x_1014)) { + x_1048 = lean_alloc_ctor(1, 2, 0); +} else { + x_1048 = x_1014; +} +lean_ctor_set(x_1048, 0, x_79); +lean_ctor_set(x_1048, 1, x_651); +x_1049 = lean_box(0); +x_1050 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_1048, x_1049); +x_1051 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1050, x_1049); +x_1052 = l_Lean_MessageData_ofList(x_1051); +lean_dec(x_1051); +x_1053 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_1054 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1054, 0, x_1053); +lean_ctor_set(x_1054, 1, x_1052); +x_1055 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_1056 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1056, 0, x_1054); +lean_ctor_set(x_1056, 1, x_1055); +x_1057 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1056, x_6, x_7, x_650); +lean_dec(x_7); +x_1058 = lean_ctor_get(x_1057, 0); +lean_inc(x_1058); +x_1059 = lean_ctor_get(x_1057, 1); +lean_inc(x_1059); +if (lean_is_exclusive(x_1057)) { + lean_ctor_release(x_1057, 0); + lean_ctor_release(x_1057, 1); + x_1060 = x_1057; } else { - lean_dec_ref(x_1066); - x_1069 = lean_box(0); + lean_dec_ref(x_1057); + x_1060 = lean_box(0); } -if (lean_is_scalar(x_1069)) { - x_1070 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_1060)) { + x_1061 = lean_alloc_ctor(1, 2, 0); } else { - x_1070 = x_1069; + x_1061 = x_1060; } -lean_ctor_set(x_1070, 0, x_1067); -lean_ctor_set(x_1070, 1, x_1068); -return x_1070; +lean_ctor_set(x_1061, 0, x_1058); +lean_ctor_set(x_1061, 1, x_1059); +return x_1061; } } else { -lean_object* x_1071; uint8_t x_1072; -lean_dec(x_663); -x_1071 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__7; -x_1072 = lean_string_dec_eq(x_660, x_1071); -if (x_1072 == 0) +lean_object* x_1062; uint8_t x_1063; +lean_dec(x_658); +x_1062 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__7; +x_1063 = lean_string_dec_eq(x_655, x_1062); +if (x_1063 == 0) { -if (lean_obj_tag(x_656) == 0) +lean_dec(x_3); +if (lean_obj_tag(x_651) == 0) { -lean_object* x_1073; lean_object* x_1074; lean_object* x_1075; lean_object* x_1076; -lean_dec(x_1023); -lean_free_object(x_78); -lean_dec(x_658); -if (lean_is_scalar(x_1006)) { - x_1073 = lean_alloc_ctor(1, 2, 8); -} else { - x_1073 = x_1006; -} -lean_ctor_set(x_1073, 0, x_654); -lean_ctor_set(x_1073, 1, x_1007); -lean_ctor_set_uint64(x_1073, sizeof(void*)*2, x_1005); -x_1074 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_1074, 0, x_1073); -lean_ctor_set(x_1074, 1, x_1025); -lean_ctor_set_uint64(x_1074, sizeof(void*)*2, x_1003); -lean_ctor_set(x_80, 1, x_1048); -lean_ctor_set(x_80, 0, x_1074); -if (lean_is_scalar(x_1024)) { - x_1075 = lean_alloc_ctor(1, 2, 8); -} else { - x_1075 = x_1024; -} -lean_ctor_set(x_1075, 0, x_80); -lean_ctor_set(x_1075, 1, x_660); -lean_ctor_set_uint64(x_1075, sizeof(void*)*2, x_662); -x_1076 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_1075, x_5, x_6, x_655); +lean_object* x_1064; lean_object* x_1065; lean_object* x_1066; lean_object* x_1067; +lean_dec(x_1014); +lean_free_object(x_79); +lean_dec(x_653); +if (lean_is_scalar(x_997)) { + x_1064 = lean_alloc_ctor(1, 2, 8); +} else { + x_1064 = x_997; +} +lean_ctor_set(x_1064, 0, x_649); +lean_ctor_set(x_1064, 1, x_998); +lean_ctor_set_uint64(x_1064, sizeof(void*)*2, x_996); +x_1065 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_1065, 0, x_1064); +lean_ctor_set(x_1065, 1, x_1016); +lean_ctor_set_uint64(x_1065, sizeof(void*)*2, x_994); +lean_ctor_set(x_81, 1, x_1039); +lean_ctor_set(x_81, 0, x_1065); +if (lean_is_scalar(x_1015)) { + x_1066 = lean_alloc_ctor(1, 2, 8); +} else { + x_1066 = x_1015; +} +lean_ctor_set(x_1066, 0, x_81); +lean_ctor_set(x_1066, 1, x_655); +lean_ctor_set_uint64(x_1066, sizeof(void*)*2, x_657); +x_1067 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_1066, x_6, x_7, x_650); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_1075); -return x_1076; +lean_dec(x_1066); +return x_1067; } else { -lean_object* x_1077; lean_object* x_1078; lean_object* x_1079; lean_object* x_1080; lean_object* x_1081; lean_object* x_1082; lean_object* x_1083; lean_object* x_1084; lean_object* x_1085; lean_object* x_1086; lean_object* x_1087; lean_object* x_1088; lean_object* x_1089; lean_object* x_1090; lean_object* x_1091; lean_object* x_1092; lean_object* x_1093; +lean_object* x_1068; lean_object* x_1069; lean_object* x_1070; lean_object* x_1071; lean_object* x_1072; lean_object* x_1073; lean_object* x_1074; lean_object* x_1075; lean_object* x_1076; lean_object* x_1077; lean_object* x_1078; lean_object* x_1079; lean_object* x_1080; lean_object* x_1081; lean_object* x_1082; lean_object* x_1083; lean_object* x_1084; lean_dec(x_1); -if (lean_is_scalar(x_1006)) { - x_1077 = lean_alloc_ctor(1, 2, 8); -} else { - x_1077 = x_1006; -} -lean_ctor_set(x_1077, 0, x_654); -lean_ctor_set(x_1077, 1, x_1007); -lean_ctor_set_uint64(x_1077, sizeof(void*)*2, x_1005); -x_1078 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_1078, 0, x_1077); -lean_ctor_set(x_1078, 1, x_1025); -lean_ctor_set_uint64(x_1078, sizeof(void*)*2, x_1003); -lean_ctor_set(x_80, 1, x_1048); -lean_ctor_set(x_80, 0, x_1078); -if (lean_is_scalar(x_1024)) { - x_1079 = lean_alloc_ctor(1, 2, 8); -} else { - x_1079 = x_1024; -} -lean_ctor_set(x_1079, 0, x_80); -lean_ctor_set(x_1079, 1, x_660); -lean_ctor_set_uint64(x_1079, sizeof(void*)*2, x_662); -lean_ctor_set(x_78, 0, x_1079); -if (lean_is_scalar(x_1023)) { - x_1080 = lean_alloc_ctor(1, 2, 0); -} else { - x_1080 = x_1023; -} -lean_ctor_set(x_1080, 0, x_78); -lean_ctor_set(x_1080, 1, x_656); -x_1081 = lean_box(0); -x_1082 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_1080, x_1081); -x_1083 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1082, x_1081); -x_1084 = l_Lean_MessageData_ofList(x_1083); -lean_dec(x_1083); -x_1085 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_1086 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1086, 0, x_1085); -lean_ctor_set(x_1086, 1, x_1084); -x_1087 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_1088 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1088, 0, x_1086); -lean_ctor_set(x_1088, 1, x_1087); -x_1089 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1088, x_5, x_6, x_655); -lean_dec(x_6); -x_1090 = lean_ctor_get(x_1089, 0); -lean_inc(x_1090); -x_1091 = lean_ctor_get(x_1089, 1); -lean_inc(x_1091); -if (lean_is_exclusive(x_1089)) { - lean_ctor_release(x_1089, 0); - lean_ctor_release(x_1089, 1); - x_1092 = x_1089; +if (lean_is_scalar(x_997)) { + x_1068 = lean_alloc_ctor(1, 2, 8); +} else { + x_1068 = x_997; +} +lean_ctor_set(x_1068, 0, x_649); +lean_ctor_set(x_1068, 1, x_998); +lean_ctor_set_uint64(x_1068, sizeof(void*)*2, x_996); +x_1069 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_1069, 0, x_1068); +lean_ctor_set(x_1069, 1, x_1016); +lean_ctor_set_uint64(x_1069, sizeof(void*)*2, x_994); +lean_ctor_set(x_81, 1, x_1039); +lean_ctor_set(x_81, 0, x_1069); +if (lean_is_scalar(x_1015)) { + x_1070 = lean_alloc_ctor(1, 2, 8); +} else { + x_1070 = x_1015; +} +lean_ctor_set(x_1070, 0, x_81); +lean_ctor_set(x_1070, 1, x_655); +lean_ctor_set_uint64(x_1070, sizeof(void*)*2, x_657); +lean_ctor_set(x_79, 0, x_1070); +if (lean_is_scalar(x_1014)) { + x_1071 = lean_alloc_ctor(1, 2, 0); +} else { + x_1071 = x_1014; +} +lean_ctor_set(x_1071, 0, x_79); +lean_ctor_set(x_1071, 1, x_651); +x_1072 = lean_box(0); +x_1073 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_1071, x_1072); +x_1074 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1073, x_1072); +x_1075 = l_Lean_MessageData_ofList(x_1074); +lean_dec(x_1074); +x_1076 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_1077 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1077, 0, x_1076); +lean_ctor_set(x_1077, 1, x_1075); +x_1078 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_1079 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1079, 0, x_1077); +lean_ctor_set(x_1079, 1, x_1078); +x_1080 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1079, x_6, x_7, x_650); +lean_dec(x_7); +x_1081 = lean_ctor_get(x_1080, 0); +lean_inc(x_1081); +x_1082 = lean_ctor_get(x_1080, 1); +lean_inc(x_1082); +if (lean_is_exclusive(x_1080)) { + lean_ctor_release(x_1080, 0); + lean_ctor_release(x_1080, 1); + x_1083 = x_1080; } else { - lean_dec_ref(x_1089); - x_1092 = lean_box(0); + lean_dec_ref(x_1080); + x_1083 = lean_box(0); } -if (lean_is_scalar(x_1092)) { - x_1093 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_1083)) { + x_1084 = lean_alloc_ctor(1, 2, 0); } else { - x_1093 = x_1092; + x_1084 = x_1083; } -lean_ctor_set(x_1093, 0, x_1090); -lean_ctor_set(x_1093, 1, x_1091); -return x_1093; +lean_ctor_set(x_1084, 0, x_1081); +lean_ctor_set(x_1084, 1, x_1082); +return x_1084; } } else { -lean_dec(x_660); -if (lean_obj_tag(x_656) == 0) +lean_dec(x_655); +if (lean_obj_tag(x_651) == 0) { -lean_object* x_1094; lean_object* x_1095; -lean_dec(x_1024); -lean_dec(x_1023); -lean_dec(x_1006); -lean_free_object(x_80); -lean_free_object(x_78); -lean_dec(x_658); -x_1094 = l_Lean_identKind; -x_1095 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_1094, x_5, x_6, x_655); +lean_object* x_1085; +lean_dec(x_1015); +lean_dec(x_1014); +lean_dec(x_997); +lean_free_object(x_81); +lean_free_object(x_79); +lean_dec(x_653); +x_1085 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_3, x_6, x_7, x_650); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -return x_1095; +lean_dec(x_3); +return x_1085; } else { -lean_object* x_1096; lean_object* x_1097; lean_object* x_1098; lean_object* x_1099; lean_object* x_1100; lean_object* x_1101; lean_object* x_1102; lean_object* x_1103; lean_object* x_1104; lean_object* x_1105; lean_object* x_1106; lean_object* x_1107; lean_object* x_1108; lean_object* x_1109; lean_object* x_1110; lean_object* x_1111; lean_object* x_1112; +lean_object* x_1086; lean_object* x_1087; lean_object* x_1088; lean_object* x_1089; lean_object* x_1090; lean_object* x_1091; lean_object* x_1092; lean_object* x_1093; lean_object* x_1094; lean_object* x_1095; lean_object* x_1096; lean_object* x_1097; lean_object* x_1098; lean_object* x_1099; lean_object* x_1100; lean_object* x_1101; lean_object* x_1102; +lean_dec(x_3); lean_dec(x_1); -if (lean_is_scalar(x_1006)) { - x_1096 = lean_alloc_ctor(1, 2, 8); -} else { - x_1096 = x_1006; -} -lean_ctor_set(x_1096, 0, x_654); -lean_ctor_set(x_1096, 1, x_1007); -lean_ctor_set_uint64(x_1096, sizeof(void*)*2, x_1005); -x_1097 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_1097, 0, x_1096); -lean_ctor_set(x_1097, 1, x_1025); -lean_ctor_set_uint64(x_1097, sizeof(void*)*2, x_1003); -lean_ctor_set(x_80, 1, x_1048); -lean_ctor_set(x_80, 0, x_1097); -if (lean_is_scalar(x_1024)) { - x_1098 = lean_alloc_ctor(1, 2, 8); -} else { - x_1098 = x_1024; -} -lean_ctor_set(x_1098, 0, x_80); -lean_ctor_set(x_1098, 1, x_1071); -lean_ctor_set_uint64(x_1098, sizeof(void*)*2, x_662); -lean_ctor_set(x_78, 0, x_1098); -if (lean_is_scalar(x_1023)) { - x_1099 = lean_alloc_ctor(1, 2, 0); -} else { - x_1099 = x_1023; -} -lean_ctor_set(x_1099, 0, x_78); -lean_ctor_set(x_1099, 1, x_656); -x_1100 = lean_box(0); -x_1101 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_1099, x_1100); -x_1102 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1101, x_1100); -x_1103 = l_Lean_MessageData_ofList(x_1102); -lean_dec(x_1102); -x_1104 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_1105 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1105, 0, x_1104); -lean_ctor_set(x_1105, 1, x_1103); -x_1106 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_1107 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1107, 0, x_1105); -lean_ctor_set(x_1107, 1, x_1106); -x_1108 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1107, x_5, x_6, x_655); -lean_dec(x_6); -x_1109 = lean_ctor_get(x_1108, 0); -lean_inc(x_1109); -x_1110 = lean_ctor_get(x_1108, 1); -lean_inc(x_1110); -if (lean_is_exclusive(x_1108)) { - lean_ctor_release(x_1108, 0); - lean_ctor_release(x_1108, 1); - x_1111 = x_1108; -} else { - lean_dec_ref(x_1108); - x_1111 = lean_box(0); +if (lean_is_scalar(x_997)) { + x_1086 = lean_alloc_ctor(1, 2, 8); +} else { + x_1086 = x_997; +} +lean_ctor_set(x_1086, 0, x_649); +lean_ctor_set(x_1086, 1, x_998); +lean_ctor_set_uint64(x_1086, sizeof(void*)*2, x_996); +x_1087 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_1087, 0, x_1086); +lean_ctor_set(x_1087, 1, x_1016); +lean_ctor_set_uint64(x_1087, sizeof(void*)*2, x_994); +lean_ctor_set(x_81, 1, x_1039); +lean_ctor_set(x_81, 0, x_1087); +if (lean_is_scalar(x_1015)) { + x_1088 = lean_alloc_ctor(1, 2, 8); +} else { + x_1088 = x_1015; +} +lean_ctor_set(x_1088, 0, x_81); +lean_ctor_set(x_1088, 1, x_1062); +lean_ctor_set_uint64(x_1088, sizeof(void*)*2, x_657); +lean_ctor_set(x_79, 0, x_1088); +if (lean_is_scalar(x_1014)) { + x_1089 = lean_alloc_ctor(1, 2, 0); +} else { + x_1089 = x_1014; +} +lean_ctor_set(x_1089, 0, x_79); +lean_ctor_set(x_1089, 1, x_651); +x_1090 = lean_box(0); +x_1091 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_1089, x_1090); +x_1092 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1091, x_1090); +x_1093 = l_Lean_MessageData_ofList(x_1092); +lean_dec(x_1092); +x_1094 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_1095 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1095, 0, x_1094); +lean_ctor_set(x_1095, 1, x_1093); +x_1096 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_1097 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1097, 0, x_1095); +lean_ctor_set(x_1097, 1, x_1096); +x_1098 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1097, x_6, x_7, x_650); +lean_dec(x_7); +x_1099 = lean_ctor_get(x_1098, 0); +lean_inc(x_1099); +x_1100 = lean_ctor_get(x_1098, 1); +lean_inc(x_1100); +if (lean_is_exclusive(x_1098)) { + lean_ctor_release(x_1098, 0); + lean_ctor_release(x_1098, 1); + x_1101 = x_1098; +} else { + lean_dec_ref(x_1098); + x_1101 = lean_box(0); } -if (lean_is_scalar(x_1111)) { - x_1112 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_1101)) { + x_1102 = lean_alloc_ctor(1, 2, 0); } else { - x_1112 = x_1111; + x_1102 = x_1101; } -lean_ctor_set(x_1112, 0, x_1109); -lean_ctor_set(x_1112, 1, x_1110); -return x_1112; +lean_ctor_set(x_1102, 0, x_1099); +lean_ctor_set(x_1102, 1, x_1100); +return x_1102; } } } @@ -6812,582 +6874,587 @@ return x_1112; } else { -uint64_t x_1113; lean_object* x_1114; uint64_t x_1115; lean_object* x_1116; uint64_t x_1117; lean_object* x_1118; lean_object* x_1119; uint64_t x_1120; lean_object* x_1121; lean_object* x_1122; uint8_t x_1123; -x_1113 = lean_ctor_get_uint64(x_79, sizeof(void*)*2); -x_1114 = lean_ctor_get(x_80, 1); -x_1115 = lean_ctor_get_uint64(x_80, sizeof(void*)*2); -lean_inc(x_1114); -lean_dec(x_80); -x_1116 = lean_ctor_get(x_81, 1); -lean_inc(x_1116); -x_1117 = lean_ctor_get_uint64(x_81, sizeof(void*)*2); -if (lean_is_exclusive(x_81)) { - lean_ctor_release(x_81, 0); - lean_ctor_release(x_81, 1); - x_1118 = x_81; -} else { - lean_dec_ref(x_81); - x_1118 = lean_box(0); -} -x_1119 = lean_ctor_get(x_82, 1); -lean_inc(x_1119); -x_1120 = lean_ctor_get_uint64(x_82, sizeof(void*)*2); +uint64_t x_1103; lean_object* x_1104; uint64_t x_1105; lean_object* x_1106; uint64_t x_1107; lean_object* x_1108; lean_object* x_1109; uint64_t x_1110; lean_object* x_1111; lean_object* x_1112; uint8_t x_1113; +x_1103 = lean_ctor_get_uint64(x_80, sizeof(void*)*2); +x_1104 = lean_ctor_get(x_81, 1); +x_1105 = lean_ctor_get_uint64(x_81, sizeof(void*)*2); +lean_inc(x_1104); +lean_dec(x_81); +x_1106 = lean_ctor_get(x_82, 1); +lean_inc(x_1106); +x_1107 = lean_ctor_get_uint64(x_82, sizeof(void*)*2); if (lean_is_exclusive(x_82)) { lean_ctor_release(x_82, 0); lean_ctor_release(x_82, 1); - x_1121 = x_82; + x_1108 = x_82; } else { lean_dec_ref(x_82); - x_1121 = lean_box(0); -} -x_1122 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__1; -x_1123 = lean_string_dec_eq(x_1119, x_1122); -lean_dec(x_1119); -if (x_1123 == 0) -{ -lean_dec(x_1121); -lean_dec(x_1118); -lean_dec(x_1116); -lean_dec(x_1114); -lean_dec(x_660); -lean_free_object(x_78); -lean_dec(x_658); -if (lean_obj_tag(x_656) == 0) + x_1108 = lean_box(0); +} +x_1109 = lean_ctor_get(x_83, 1); +lean_inc(x_1109); +x_1110 = lean_ctor_get_uint64(x_83, sizeof(void*)*2); +if (lean_is_exclusive(x_83)) { + lean_ctor_release(x_83, 0); + lean_ctor_release(x_83, 1); + x_1111 = x_83; +} else { + lean_dec_ref(x_83); + x_1111 = lean_box(0); +} +x_1112 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__1; +x_1113 = lean_string_dec_eq(x_1109, x_1112); +lean_dec(x_1109); +if (x_1113 == 0) { -lean_object* x_1124; -lean_dec(x_9); -x_1124 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_79, x_5, x_6, x_655); +lean_dec(x_1111); +lean_dec(x_1108); +lean_dec(x_1106); +lean_dec(x_1104); +lean_dec(x_655); +lean_free_object(x_79); +lean_dec(x_653); +lean_dec(x_3); +if (lean_obj_tag(x_651) == 0) +{ +lean_object* x_1114; +lean_dec(x_10); +x_1114 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_80, x_6, x_7, x_650); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_79); -return x_1124; +lean_dec(x_80); +return x_1114; } else { -lean_object* x_1125; lean_object* x_1126; lean_object* x_1127; lean_object* x_1128; lean_object* x_1129; lean_object* x_1130; lean_object* x_1131; lean_object* x_1132; lean_object* x_1133; lean_object* x_1134; lean_object* x_1135; lean_object* x_1136; lean_object* x_1137; -lean_dec(x_656); -lean_dec(x_79); +lean_object* x_1115; lean_object* x_1116; lean_object* x_1117; lean_object* x_1118; lean_object* x_1119; lean_object* x_1120; lean_object* x_1121; lean_object* x_1122; lean_object* x_1123; lean_object* x_1124; lean_object* x_1125; lean_object* x_1126; lean_object* x_1127; +lean_dec(x_651); +lean_dec(x_80); lean_dec(x_1); -x_1125 = lean_box(0); -x_1126 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_9, x_1125); -x_1127 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1126, x_1125); -x_1128 = l_Lean_MessageData_ofList(x_1127); -lean_dec(x_1127); -x_1129 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_1130 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1130, 0, x_1129); -lean_ctor_set(x_1130, 1, x_1128); -x_1131 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_1132 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1132, 0, x_1130); -lean_ctor_set(x_1132, 1, x_1131); -x_1133 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1132, x_5, x_6, x_655); -lean_dec(x_6); -x_1134 = lean_ctor_get(x_1133, 0); -lean_inc(x_1134); -x_1135 = lean_ctor_get(x_1133, 1); -lean_inc(x_1135); -if (lean_is_exclusive(x_1133)) { - lean_ctor_release(x_1133, 0); - lean_ctor_release(x_1133, 1); - x_1136 = x_1133; +x_1115 = lean_box(0); +x_1116 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_10, x_1115); +x_1117 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1116, x_1115); +x_1118 = l_Lean_MessageData_ofList(x_1117); +lean_dec(x_1117); +x_1119 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_1120 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1120, 0, x_1119); +lean_ctor_set(x_1120, 1, x_1118); +x_1121 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_1122 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1122, 0, x_1120); +lean_ctor_set(x_1122, 1, x_1121); +x_1123 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1122, x_6, x_7, x_650); +lean_dec(x_7); +x_1124 = lean_ctor_get(x_1123, 0); +lean_inc(x_1124); +x_1125 = lean_ctor_get(x_1123, 1); +lean_inc(x_1125); +if (lean_is_exclusive(x_1123)) { + lean_ctor_release(x_1123, 0); + lean_ctor_release(x_1123, 1); + x_1126 = x_1123; } else { - lean_dec_ref(x_1133); - x_1136 = lean_box(0); + lean_dec_ref(x_1123); + x_1126 = lean_box(0); } -if (lean_is_scalar(x_1136)) { - x_1137 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_1126)) { + x_1127 = lean_alloc_ctor(1, 2, 0); } else { - x_1137 = x_1136; + x_1127 = x_1126; } -lean_ctor_set(x_1137, 0, x_1134); -lean_ctor_set(x_1137, 1, x_1135); -return x_1137; +lean_ctor_set(x_1127, 0, x_1124); +lean_ctor_set(x_1127, 1, x_1125); +return x_1127; } } else { -lean_object* x_1138; lean_object* x_1139; lean_object* x_1140; uint8_t x_1141; -if (lean_is_exclusive(x_9)) { - lean_ctor_release(x_9, 0); - lean_ctor_release(x_9, 1); - x_1138 = x_9; +lean_object* x_1128; lean_object* x_1129; lean_object* x_1130; uint8_t x_1131; +if (lean_is_exclusive(x_10)) { + lean_ctor_release(x_10, 0); + lean_ctor_release(x_10, 1); + x_1128 = x_10; } else { - lean_dec_ref(x_9); - x_1138 = lean_box(0); + lean_dec_ref(x_10); + x_1128 = lean_box(0); } -if (lean_is_exclusive(x_79)) { - lean_ctor_release(x_79, 0); - lean_ctor_release(x_79, 1); - x_1139 = x_79; +if (lean_is_exclusive(x_80)) { + lean_ctor_release(x_80, 0); + lean_ctor_release(x_80, 1); + x_1129 = x_80; } else { - lean_dec_ref(x_79); - x_1139 = lean_box(0); + lean_dec_ref(x_80); + x_1129 = lean_box(0); } -x_1140 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__4; -x_1141 = lean_string_dec_eq(x_1116, x_1140); -if (x_1141 == 0) +x_1130 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__4; +x_1131 = lean_string_dec_eq(x_1106, x_1130); +if (x_1131 == 0) { -if (lean_obj_tag(x_656) == 0) +lean_dec(x_3); +if (lean_obj_tag(x_651) == 0) { -lean_object* x_1142; lean_object* x_1143; lean_object* x_1144; lean_object* x_1145; lean_object* x_1146; -lean_dec(x_1138); -lean_free_object(x_78); -lean_dec(x_658); -if (lean_is_scalar(x_1121)) { - x_1142 = lean_alloc_ctor(1, 2, 8); -} else { - x_1142 = x_1121; -} -lean_ctor_set(x_1142, 0, x_654); -lean_ctor_set(x_1142, 1, x_1122); -lean_ctor_set_uint64(x_1142, sizeof(void*)*2, x_1120); -if (lean_is_scalar(x_1118)) { - x_1143 = lean_alloc_ctor(1, 2, 8); -} else { - x_1143 = x_1118; -} -lean_ctor_set(x_1143, 0, x_1142); -lean_ctor_set(x_1143, 1, x_1116); -lean_ctor_set_uint64(x_1143, sizeof(void*)*2, x_1117); -x_1144 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_1144, 0, x_1143); -lean_ctor_set(x_1144, 1, x_1114); -lean_ctor_set_uint64(x_1144, sizeof(void*)*2, x_1115); -if (lean_is_scalar(x_1139)) { - x_1145 = lean_alloc_ctor(1, 2, 8); -} else { - x_1145 = x_1139; -} -lean_ctor_set(x_1145, 0, x_1144); -lean_ctor_set(x_1145, 1, x_660); -lean_ctor_set_uint64(x_1145, sizeof(void*)*2, x_1113); -x_1146 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_1145, x_5, x_6, x_655); +lean_object* x_1132; lean_object* x_1133; lean_object* x_1134; lean_object* x_1135; lean_object* x_1136; +lean_dec(x_1128); +lean_free_object(x_79); +lean_dec(x_653); +if (lean_is_scalar(x_1111)) { + x_1132 = lean_alloc_ctor(1, 2, 8); +} else { + x_1132 = x_1111; +} +lean_ctor_set(x_1132, 0, x_649); +lean_ctor_set(x_1132, 1, x_1112); +lean_ctor_set_uint64(x_1132, sizeof(void*)*2, x_1110); +if (lean_is_scalar(x_1108)) { + x_1133 = lean_alloc_ctor(1, 2, 8); +} else { + x_1133 = x_1108; +} +lean_ctor_set(x_1133, 0, x_1132); +lean_ctor_set(x_1133, 1, x_1106); +lean_ctor_set_uint64(x_1133, sizeof(void*)*2, x_1107); +x_1134 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_1134, 0, x_1133); +lean_ctor_set(x_1134, 1, x_1104); +lean_ctor_set_uint64(x_1134, sizeof(void*)*2, x_1105); +if (lean_is_scalar(x_1129)) { + x_1135 = lean_alloc_ctor(1, 2, 8); +} else { + x_1135 = x_1129; +} +lean_ctor_set(x_1135, 0, x_1134); +lean_ctor_set(x_1135, 1, x_655); +lean_ctor_set_uint64(x_1135, sizeof(void*)*2, x_1103); +x_1136 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_1135, x_6, x_7, x_650); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_1145); -return x_1146; +lean_dec(x_1135); +return x_1136; } else { -lean_object* x_1147; lean_object* x_1148; lean_object* x_1149; lean_object* x_1150; lean_object* x_1151; lean_object* x_1152; lean_object* x_1153; lean_object* x_1154; lean_object* x_1155; lean_object* x_1156; lean_object* x_1157; lean_object* x_1158; lean_object* x_1159; lean_object* x_1160; lean_object* x_1161; lean_object* x_1162; lean_object* x_1163; lean_object* x_1164; +lean_object* x_1137; lean_object* x_1138; lean_object* x_1139; lean_object* x_1140; lean_object* x_1141; lean_object* x_1142; lean_object* x_1143; lean_object* x_1144; lean_object* x_1145; lean_object* x_1146; lean_object* x_1147; lean_object* x_1148; lean_object* x_1149; lean_object* x_1150; lean_object* x_1151; lean_object* x_1152; lean_object* x_1153; lean_object* x_1154; lean_dec(x_1); -if (lean_is_scalar(x_1121)) { - x_1147 = lean_alloc_ctor(1, 2, 8); -} else { - x_1147 = x_1121; -} -lean_ctor_set(x_1147, 0, x_654); -lean_ctor_set(x_1147, 1, x_1122); -lean_ctor_set_uint64(x_1147, sizeof(void*)*2, x_1120); -if (lean_is_scalar(x_1118)) { - x_1148 = lean_alloc_ctor(1, 2, 8); -} else { - x_1148 = x_1118; -} -lean_ctor_set(x_1148, 0, x_1147); -lean_ctor_set(x_1148, 1, x_1116); -lean_ctor_set_uint64(x_1148, sizeof(void*)*2, x_1117); -x_1149 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_1149, 0, x_1148); -lean_ctor_set(x_1149, 1, x_1114); -lean_ctor_set_uint64(x_1149, sizeof(void*)*2, x_1115); -if (lean_is_scalar(x_1139)) { - x_1150 = lean_alloc_ctor(1, 2, 8); -} else { - x_1150 = x_1139; -} -lean_ctor_set(x_1150, 0, x_1149); -lean_ctor_set(x_1150, 1, x_660); -lean_ctor_set_uint64(x_1150, sizeof(void*)*2, x_1113); -lean_ctor_set(x_78, 0, x_1150); -if (lean_is_scalar(x_1138)) { - x_1151 = lean_alloc_ctor(1, 2, 0); -} else { - x_1151 = x_1138; -} -lean_ctor_set(x_1151, 0, x_78); -lean_ctor_set(x_1151, 1, x_656); -x_1152 = lean_box(0); -x_1153 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_1151, x_1152); -x_1154 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1153, x_1152); -x_1155 = l_Lean_MessageData_ofList(x_1154); -lean_dec(x_1154); -x_1156 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_1157 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1157, 0, x_1156); -lean_ctor_set(x_1157, 1, x_1155); -x_1158 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_1159 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1159, 0, x_1157); -lean_ctor_set(x_1159, 1, x_1158); -x_1160 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1159, x_5, x_6, x_655); -lean_dec(x_6); -x_1161 = lean_ctor_get(x_1160, 0); -lean_inc(x_1161); -x_1162 = lean_ctor_get(x_1160, 1); -lean_inc(x_1162); -if (lean_is_exclusive(x_1160)) { - lean_ctor_release(x_1160, 0); - lean_ctor_release(x_1160, 1); - x_1163 = x_1160; +if (lean_is_scalar(x_1111)) { + x_1137 = lean_alloc_ctor(1, 2, 8); +} else { + x_1137 = x_1111; +} +lean_ctor_set(x_1137, 0, x_649); +lean_ctor_set(x_1137, 1, x_1112); +lean_ctor_set_uint64(x_1137, sizeof(void*)*2, x_1110); +if (lean_is_scalar(x_1108)) { + x_1138 = lean_alloc_ctor(1, 2, 8); +} else { + x_1138 = x_1108; +} +lean_ctor_set(x_1138, 0, x_1137); +lean_ctor_set(x_1138, 1, x_1106); +lean_ctor_set_uint64(x_1138, sizeof(void*)*2, x_1107); +x_1139 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_1139, 0, x_1138); +lean_ctor_set(x_1139, 1, x_1104); +lean_ctor_set_uint64(x_1139, sizeof(void*)*2, x_1105); +if (lean_is_scalar(x_1129)) { + x_1140 = lean_alloc_ctor(1, 2, 8); +} else { + x_1140 = x_1129; +} +lean_ctor_set(x_1140, 0, x_1139); +lean_ctor_set(x_1140, 1, x_655); +lean_ctor_set_uint64(x_1140, sizeof(void*)*2, x_1103); +lean_ctor_set(x_79, 0, x_1140); +if (lean_is_scalar(x_1128)) { + x_1141 = lean_alloc_ctor(1, 2, 0); +} else { + x_1141 = x_1128; +} +lean_ctor_set(x_1141, 0, x_79); +lean_ctor_set(x_1141, 1, x_651); +x_1142 = lean_box(0); +x_1143 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_1141, x_1142); +x_1144 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1143, x_1142); +x_1145 = l_Lean_MessageData_ofList(x_1144); +lean_dec(x_1144); +x_1146 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_1147 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1147, 0, x_1146); +lean_ctor_set(x_1147, 1, x_1145); +x_1148 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_1149 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1149, 0, x_1147); +lean_ctor_set(x_1149, 1, x_1148); +x_1150 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1149, x_6, x_7, x_650); +lean_dec(x_7); +x_1151 = lean_ctor_get(x_1150, 0); +lean_inc(x_1151); +x_1152 = lean_ctor_get(x_1150, 1); +lean_inc(x_1152); +if (lean_is_exclusive(x_1150)) { + lean_ctor_release(x_1150, 0); + lean_ctor_release(x_1150, 1); + x_1153 = x_1150; } else { - lean_dec_ref(x_1160); - x_1163 = lean_box(0); + lean_dec_ref(x_1150); + x_1153 = lean_box(0); } -if (lean_is_scalar(x_1163)) { - x_1164 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_1153)) { + x_1154 = lean_alloc_ctor(1, 2, 0); } else { - x_1164 = x_1163; + x_1154 = x_1153; } -lean_ctor_set(x_1164, 0, x_1161); -lean_ctor_set(x_1164, 1, x_1162); -return x_1164; +lean_ctor_set(x_1154, 0, x_1151); +lean_ctor_set(x_1154, 1, x_1152); +return x_1154; } } else { -lean_object* x_1165; uint8_t x_1166; -lean_dec(x_1116); -x_1165 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__9; -x_1166 = lean_string_dec_eq(x_1114, x_1165); -if (x_1166 == 0) +lean_object* x_1155; uint8_t x_1156; +lean_dec(x_1106); +x_1155 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__11; +x_1156 = lean_string_dec_eq(x_1104, x_1155); +if (x_1156 == 0) { -if (lean_obj_tag(x_656) == 0) +lean_dec(x_3); +if (lean_obj_tag(x_651) == 0) { -lean_object* x_1167; lean_object* x_1168; lean_object* x_1169; lean_object* x_1170; lean_object* x_1171; -lean_dec(x_1138); -lean_free_object(x_78); -lean_dec(x_658); -if (lean_is_scalar(x_1121)) { - x_1167 = lean_alloc_ctor(1, 2, 8); -} else { - x_1167 = x_1121; -} -lean_ctor_set(x_1167, 0, x_654); -lean_ctor_set(x_1167, 1, x_1122); -lean_ctor_set_uint64(x_1167, sizeof(void*)*2, x_1120); -if (lean_is_scalar(x_1118)) { - x_1168 = lean_alloc_ctor(1, 2, 8); -} else { - x_1168 = x_1118; -} -lean_ctor_set(x_1168, 0, x_1167); -lean_ctor_set(x_1168, 1, x_1140); -lean_ctor_set_uint64(x_1168, sizeof(void*)*2, x_1117); -x_1169 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_1169, 0, x_1168); -lean_ctor_set(x_1169, 1, x_1114); -lean_ctor_set_uint64(x_1169, sizeof(void*)*2, x_1115); -if (lean_is_scalar(x_1139)) { - x_1170 = lean_alloc_ctor(1, 2, 8); -} else { - x_1170 = x_1139; -} -lean_ctor_set(x_1170, 0, x_1169); -lean_ctor_set(x_1170, 1, x_660); -lean_ctor_set_uint64(x_1170, sizeof(void*)*2, x_1113); -x_1171 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_1170, x_5, x_6, x_655); +lean_object* x_1157; lean_object* x_1158; lean_object* x_1159; lean_object* x_1160; lean_object* x_1161; +lean_dec(x_1128); +lean_free_object(x_79); +lean_dec(x_653); +if (lean_is_scalar(x_1111)) { + x_1157 = lean_alloc_ctor(1, 2, 8); +} else { + x_1157 = x_1111; +} +lean_ctor_set(x_1157, 0, x_649); +lean_ctor_set(x_1157, 1, x_1112); +lean_ctor_set_uint64(x_1157, sizeof(void*)*2, x_1110); +if (lean_is_scalar(x_1108)) { + x_1158 = lean_alloc_ctor(1, 2, 8); +} else { + x_1158 = x_1108; +} +lean_ctor_set(x_1158, 0, x_1157); +lean_ctor_set(x_1158, 1, x_1130); +lean_ctor_set_uint64(x_1158, sizeof(void*)*2, x_1107); +x_1159 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_1159, 0, x_1158); +lean_ctor_set(x_1159, 1, x_1104); +lean_ctor_set_uint64(x_1159, sizeof(void*)*2, x_1105); +if (lean_is_scalar(x_1129)) { + x_1160 = lean_alloc_ctor(1, 2, 8); +} else { + x_1160 = x_1129; +} +lean_ctor_set(x_1160, 0, x_1159); +lean_ctor_set(x_1160, 1, x_655); +lean_ctor_set_uint64(x_1160, sizeof(void*)*2, x_1103); +x_1161 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_1160, x_6, x_7, x_650); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_1170); -return x_1171; +lean_dec(x_1160); +return x_1161; } else { -lean_object* x_1172; lean_object* x_1173; lean_object* x_1174; lean_object* x_1175; lean_object* x_1176; lean_object* x_1177; lean_object* x_1178; lean_object* x_1179; lean_object* x_1180; lean_object* x_1181; lean_object* x_1182; lean_object* x_1183; lean_object* x_1184; lean_object* x_1185; lean_object* x_1186; lean_object* x_1187; lean_object* x_1188; lean_object* x_1189; +lean_object* x_1162; lean_object* x_1163; lean_object* x_1164; lean_object* x_1165; lean_object* x_1166; lean_object* x_1167; lean_object* x_1168; lean_object* x_1169; lean_object* x_1170; lean_object* x_1171; lean_object* x_1172; lean_object* x_1173; lean_object* x_1174; lean_object* x_1175; lean_object* x_1176; lean_object* x_1177; lean_object* x_1178; lean_object* x_1179; lean_dec(x_1); -if (lean_is_scalar(x_1121)) { - x_1172 = lean_alloc_ctor(1, 2, 8); -} else { - x_1172 = x_1121; -} -lean_ctor_set(x_1172, 0, x_654); -lean_ctor_set(x_1172, 1, x_1122); -lean_ctor_set_uint64(x_1172, sizeof(void*)*2, x_1120); -if (lean_is_scalar(x_1118)) { - x_1173 = lean_alloc_ctor(1, 2, 8); -} else { - x_1173 = x_1118; -} -lean_ctor_set(x_1173, 0, x_1172); -lean_ctor_set(x_1173, 1, x_1140); -lean_ctor_set_uint64(x_1173, sizeof(void*)*2, x_1117); -x_1174 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_1174, 0, x_1173); -lean_ctor_set(x_1174, 1, x_1114); -lean_ctor_set_uint64(x_1174, sizeof(void*)*2, x_1115); -if (lean_is_scalar(x_1139)) { - x_1175 = lean_alloc_ctor(1, 2, 8); -} else { - x_1175 = x_1139; -} -lean_ctor_set(x_1175, 0, x_1174); -lean_ctor_set(x_1175, 1, x_660); -lean_ctor_set_uint64(x_1175, sizeof(void*)*2, x_1113); -lean_ctor_set(x_78, 0, x_1175); -if (lean_is_scalar(x_1138)) { - x_1176 = lean_alloc_ctor(1, 2, 0); -} else { - x_1176 = x_1138; -} -lean_ctor_set(x_1176, 0, x_78); -lean_ctor_set(x_1176, 1, x_656); -x_1177 = lean_box(0); -x_1178 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_1176, x_1177); -x_1179 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1178, x_1177); -x_1180 = l_Lean_MessageData_ofList(x_1179); -lean_dec(x_1179); -x_1181 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_1182 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1182, 0, x_1181); -lean_ctor_set(x_1182, 1, x_1180); -x_1183 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_1184 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1184, 0, x_1182); -lean_ctor_set(x_1184, 1, x_1183); -x_1185 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1184, x_5, x_6, x_655); -lean_dec(x_6); -x_1186 = lean_ctor_get(x_1185, 0); -lean_inc(x_1186); -x_1187 = lean_ctor_get(x_1185, 1); -lean_inc(x_1187); -if (lean_is_exclusive(x_1185)) { - lean_ctor_release(x_1185, 0); - lean_ctor_release(x_1185, 1); - x_1188 = x_1185; +if (lean_is_scalar(x_1111)) { + x_1162 = lean_alloc_ctor(1, 2, 8); +} else { + x_1162 = x_1111; +} +lean_ctor_set(x_1162, 0, x_649); +lean_ctor_set(x_1162, 1, x_1112); +lean_ctor_set_uint64(x_1162, sizeof(void*)*2, x_1110); +if (lean_is_scalar(x_1108)) { + x_1163 = lean_alloc_ctor(1, 2, 8); +} else { + x_1163 = x_1108; +} +lean_ctor_set(x_1163, 0, x_1162); +lean_ctor_set(x_1163, 1, x_1130); +lean_ctor_set_uint64(x_1163, sizeof(void*)*2, x_1107); +x_1164 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_1164, 0, x_1163); +lean_ctor_set(x_1164, 1, x_1104); +lean_ctor_set_uint64(x_1164, sizeof(void*)*2, x_1105); +if (lean_is_scalar(x_1129)) { + x_1165 = lean_alloc_ctor(1, 2, 8); +} else { + x_1165 = x_1129; +} +lean_ctor_set(x_1165, 0, x_1164); +lean_ctor_set(x_1165, 1, x_655); +lean_ctor_set_uint64(x_1165, sizeof(void*)*2, x_1103); +lean_ctor_set(x_79, 0, x_1165); +if (lean_is_scalar(x_1128)) { + x_1166 = lean_alloc_ctor(1, 2, 0); +} else { + x_1166 = x_1128; +} +lean_ctor_set(x_1166, 0, x_79); +lean_ctor_set(x_1166, 1, x_651); +x_1167 = lean_box(0); +x_1168 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_1166, x_1167); +x_1169 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1168, x_1167); +x_1170 = l_Lean_MessageData_ofList(x_1169); +lean_dec(x_1169); +x_1171 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_1172 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1172, 0, x_1171); +lean_ctor_set(x_1172, 1, x_1170); +x_1173 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_1174 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1174, 0, x_1172); +lean_ctor_set(x_1174, 1, x_1173); +x_1175 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1174, x_6, x_7, x_650); +lean_dec(x_7); +x_1176 = lean_ctor_get(x_1175, 0); +lean_inc(x_1176); +x_1177 = lean_ctor_get(x_1175, 1); +lean_inc(x_1177); +if (lean_is_exclusive(x_1175)) { + lean_ctor_release(x_1175, 0); + lean_ctor_release(x_1175, 1); + x_1178 = x_1175; } else { - lean_dec_ref(x_1185); - x_1188 = lean_box(0); + lean_dec_ref(x_1175); + x_1178 = lean_box(0); } -if (lean_is_scalar(x_1188)) { - x_1189 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_1178)) { + x_1179 = lean_alloc_ctor(1, 2, 0); } else { - x_1189 = x_1188; + x_1179 = x_1178; } -lean_ctor_set(x_1189, 0, x_1186); -lean_ctor_set(x_1189, 1, x_1187); -return x_1189; +lean_ctor_set(x_1179, 0, x_1176); +lean_ctor_set(x_1179, 1, x_1177); +return x_1179; } } else { -lean_object* x_1190; uint8_t x_1191; -lean_dec(x_1114); -x_1190 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__7; -x_1191 = lean_string_dec_eq(x_660, x_1190); -if (x_1191 == 0) +lean_object* x_1180; uint8_t x_1181; +lean_dec(x_1104); +x_1180 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__7; +x_1181 = lean_string_dec_eq(x_655, x_1180); +if (x_1181 == 0) { -if (lean_obj_tag(x_656) == 0) +lean_dec(x_3); +if (lean_obj_tag(x_651) == 0) { -lean_object* x_1192; lean_object* x_1193; lean_object* x_1194; lean_object* x_1195; lean_object* x_1196; -lean_dec(x_1138); -lean_free_object(x_78); -lean_dec(x_658); -if (lean_is_scalar(x_1121)) { - x_1192 = lean_alloc_ctor(1, 2, 8); -} else { - x_1192 = x_1121; -} -lean_ctor_set(x_1192, 0, x_654); -lean_ctor_set(x_1192, 1, x_1122); -lean_ctor_set_uint64(x_1192, sizeof(void*)*2, x_1120); -if (lean_is_scalar(x_1118)) { - x_1193 = lean_alloc_ctor(1, 2, 8); -} else { - x_1193 = x_1118; -} -lean_ctor_set(x_1193, 0, x_1192); -lean_ctor_set(x_1193, 1, x_1140); -lean_ctor_set_uint64(x_1193, sizeof(void*)*2, x_1117); -x_1194 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_1194, 0, x_1193); -lean_ctor_set(x_1194, 1, x_1165); -lean_ctor_set_uint64(x_1194, sizeof(void*)*2, x_1115); -if (lean_is_scalar(x_1139)) { - x_1195 = lean_alloc_ctor(1, 2, 8); -} else { - x_1195 = x_1139; -} -lean_ctor_set(x_1195, 0, x_1194); -lean_ctor_set(x_1195, 1, x_660); -lean_ctor_set_uint64(x_1195, sizeof(void*)*2, x_1113); -x_1196 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_1195, x_5, x_6, x_655); +lean_object* x_1182; lean_object* x_1183; lean_object* x_1184; lean_object* x_1185; lean_object* x_1186; +lean_dec(x_1128); +lean_free_object(x_79); +lean_dec(x_653); +if (lean_is_scalar(x_1111)) { + x_1182 = lean_alloc_ctor(1, 2, 8); +} else { + x_1182 = x_1111; +} +lean_ctor_set(x_1182, 0, x_649); +lean_ctor_set(x_1182, 1, x_1112); +lean_ctor_set_uint64(x_1182, sizeof(void*)*2, x_1110); +if (lean_is_scalar(x_1108)) { + x_1183 = lean_alloc_ctor(1, 2, 8); +} else { + x_1183 = x_1108; +} +lean_ctor_set(x_1183, 0, x_1182); +lean_ctor_set(x_1183, 1, x_1130); +lean_ctor_set_uint64(x_1183, sizeof(void*)*2, x_1107); +x_1184 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_1184, 0, x_1183); +lean_ctor_set(x_1184, 1, x_1155); +lean_ctor_set_uint64(x_1184, sizeof(void*)*2, x_1105); +if (lean_is_scalar(x_1129)) { + x_1185 = lean_alloc_ctor(1, 2, 8); +} else { + x_1185 = x_1129; +} +lean_ctor_set(x_1185, 0, x_1184); +lean_ctor_set(x_1185, 1, x_655); +lean_ctor_set_uint64(x_1185, sizeof(void*)*2, x_1103); +x_1186 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_1185, x_6, x_7, x_650); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_1195); -return x_1196; +lean_dec(x_1185); +return x_1186; } else { -lean_object* x_1197; lean_object* x_1198; lean_object* x_1199; lean_object* x_1200; lean_object* x_1201; lean_object* x_1202; lean_object* x_1203; lean_object* x_1204; lean_object* x_1205; lean_object* x_1206; lean_object* x_1207; lean_object* x_1208; lean_object* x_1209; lean_object* x_1210; lean_object* x_1211; lean_object* x_1212; lean_object* x_1213; lean_object* x_1214; +lean_object* x_1187; lean_object* x_1188; lean_object* x_1189; lean_object* x_1190; lean_object* x_1191; lean_object* x_1192; lean_object* x_1193; lean_object* x_1194; lean_object* x_1195; lean_object* x_1196; lean_object* x_1197; lean_object* x_1198; lean_object* x_1199; lean_object* x_1200; lean_object* x_1201; lean_object* x_1202; lean_object* x_1203; lean_object* x_1204; lean_dec(x_1); -if (lean_is_scalar(x_1121)) { - x_1197 = lean_alloc_ctor(1, 2, 8); -} else { - x_1197 = x_1121; -} -lean_ctor_set(x_1197, 0, x_654); -lean_ctor_set(x_1197, 1, x_1122); -lean_ctor_set_uint64(x_1197, sizeof(void*)*2, x_1120); -if (lean_is_scalar(x_1118)) { - x_1198 = lean_alloc_ctor(1, 2, 8); -} else { - x_1198 = x_1118; -} -lean_ctor_set(x_1198, 0, x_1197); -lean_ctor_set(x_1198, 1, x_1140); -lean_ctor_set_uint64(x_1198, sizeof(void*)*2, x_1117); -x_1199 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_1199, 0, x_1198); -lean_ctor_set(x_1199, 1, x_1165); -lean_ctor_set_uint64(x_1199, sizeof(void*)*2, x_1115); -if (lean_is_scalar(x_1139)) { - x_1200 = lean_alloc_ctor(1, 2, 8); -} else { - x_1200 = x_1139; -} -lean_ctor_set(x_1200, 0, x_1199); -lean_ctor_set(x_1200, 1, x_660); -lean_ctor_set_uint64(x_1200, sizeof(void*)*2, x_1113); -lean_ctor_set(x_78, 0, x_1200); -if (lean_is_scalar(x_1138)) { - x_1201 = lean_alloc_ctor(1, 2, 0); -} else { - x_1201 = x_1138; -} -lean_ctor_set(x_1201, 0, x_78); -lean_ctor_set(x_1201, 1, x_656); -x_1202 = lean_box(0); -x_1203 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_1201, x_1202); -x_1204 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1203, x_1202); -x_1205 = l_Lean_MessageData_ofList(x_1204); -lean_dec(x_1204); -x_1206 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_1207 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1207, 0, x_1206); -lean_ctor_set(x_1207, 1, x_1205); -x_1208 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_1209 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1209, 0, x_1207); -lean_ctor_set(x_1209, 1, x_1208); -x_1210 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1209, x_5, x_6, x_655); -lean_dec(x_6); -x_1211 = lean_ctor_get(x_1210, 0); -lean_inc(x_1211); -x_1212 = lean_ctor_get(x_1210, 1); -lean_inc(x_1212); -if (lean_is_exclusive(x_1210)) { - lean_ctor_release(x_1210, 0); - lean_ctor_release(x_1210, 1); - x_1213 = x_1210; +if (lean_is_scalar(x_1111)) { + x_1187 = lean_alloc_ctor(1, 2, 8); +} else { + x_1187 = x_1111; +} +lean_ctor_set(x_1187, 0, x_649); +lean_ctor_set(x_1187, 1, x_1112); +lean_ctor_set_uint64(x_1187, sizeof(void*)*2, x_1110); +if (lean_is_scalar(x_1108)) { + x_1188 = lean_alloc_ctor(1, 2, 8); +} else { + x_1188 = x_1108; +} +lean_ctor_set(x_1188, 0, x_1187); +lean_ctor_set(x_1188, 1, x_1130); +lean_ctor_set_uint64(x_1188, sizeof(void*)*2, x_1107); +x_1189 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_1189, 0, x_1188); +lean_ctor_set(x_1189, 1, x_1155); +lean_ctor_set_uint64(x_1189, sizeof(void*)*2, x_1105); +if (lean_is_scalar(x_1129)) { + x_1190 = lean_alloc_ctor(1, 2, 8); +} else { + x_1190 = x_1129; +} +lean_ctor_set(x_1190, 0, x_1189); +lean_ctor_set(x_1190, 1, x_655); +lean_ctor_set_uint64(x_1190, sizeof(void*)*2, x_1103); +lean_ctor_set(x_79, 0, x_1190); +if (lean_is_scalar(x_1128)) { + x_1191 = lean_alloc_ctor(1, 2, 0); +} else { + x_1191 = x_1128; +} +lean_ctor_set(x_1191, 0, x_79); +lean_ctor_set(x_1191, 1, x_651); +x_1192 = lean_box(0); +x_1193 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_1191, x_1192); +x_1194 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1193, x_1192); +x_1195 = l_Lean_MessageData_ofList(x_1194); +lean_dec(x_1194); +x_1196 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_1197 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1197, 0, x_1196); +lean_ctor_set(x_1197, 1, x_1195); +x_1198 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_1199 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1199, 0, x_1197); +lean_ctor_set(x_1199, 1, x_1198); +x_1200 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1199, x_6, x_7, x_650); +lean_dec(x_7); +x_1201 = lean_ctor_get(x_1200, 0); +lean_inc(x_1201); +x_1202 = lean_ctor_get(x_1200, 1); +lean_inc(x_1202); +if (lean_is_exclusive(x_1200)) { + lean_ctor_release(x_1200, 0); + lean_ctor_release(x_1200, 1); + x_1203 = x_1200; } else { - lean_dec_ref(x_1210); - x_1213 = lean_box(0); + lean_dec_ref(x_1200); + x_1203 = lean_box(0); } -if (lean_is_scalar(x_1213)) { - x_1214 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_1203)) { + x_1204 = lean_alloc_ctor(1, 2, 0); } else { - x_1214 = x_1213; + x_1204 = x_1203; } -lean_ctor_set(x_1214, 0, x_1211); -lean_ctor_set(x_1214, 1, x_1212); -return x_1214; +lean_ctor_set(x_1204, 0, x_1201); +lean_ctor_set(x_1204, 1, x_1202); +return x_1204; } } else { -lean_dec(x_660); -if (lean_obj_tag(x_656) == 0) +lean_dec(x_655); +if (lean_obj_tag(x_651) == 0) { -lean_object* x_1215; lean_object* x_1216; -lean_dec(x_1139); -lean_dec(x_1138); -lean_dec(x_1121); -lean_dec(x_1118); -lean_free_object(x_78); -lean_dec(x_658); -x_1215 = l_Lean_identKind; -x_1216 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_1215, x_5, x_6, x_655); +lean_object* x_1205; +lean_dec(x_1129); +lean_dec(x_1128); +lean_dec(x_1111); +lean_dec(x_1108); +lean_free_object(x_79); +lean_dec(x_653); +x_1205 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_3, x_6, x_7, x_650); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -return x_1216; +lean_dec(x_3); +return x_1205; } else { -lean_object* x_1217; lean_object* x_1218; lean_object* x_1219; lean_object* x_1220; lean_object* x_1221; lean_object* x_1222; lean_object* x_1223; lean_object* x_1224; lean_object* x_1225; lean_object* x_1226; lean_object* x_1227; lean_object* x_1228; lean_object* x_1229; lean_object* x_1230; lean_object* x_1231; lean_object* x_1232; lean_object* x_1233; lean_object* x_1234; +lean_object* x_1206; lean_object* x_1207; lean_object* x_1208; lean_object* x_1209; lean_object* x_1210; lean_object* x_1211; lean_object* x_1212; lean_object* x_1213; lean_object* x_1214; lean_object* x_1215; lean_object* x_1216; lean_object* x_1217; lean_object* x_1218; lean_object* x_1219; lean_object* x_1220; lean_object* x_1221; lean_object* x_1222; lean_object* x_1223; +lean_dec(x_3); lean_dec(x_1); -if (lean_is_scalar(x_1121)) { - x_1217 = lean_alloc_ctor(1, 2, 8); -} else { - x_1217 = x_1121; -} -lean_ctor_set(x_1217, 0, x_654); -lean_ctor_set(x_1217, 1, x_1122); -lean_ctor_set_uint64(x_1217, sizeof(void*)*2, x_1120); -if (lean_is_scalar(x_1118)) { - x_1218 = lean_alloc_ctor(1, 2, 8); -} else { - x_1218 = x_1118; -} -lean_ctor_set(x_1218, 0, x_1217); -lean_ctor_set(x_1218, 1, x_1140); -lean_ctor_set_uint64(x_1218, sizeof(void*)*2, x_1117); -x_1219 = lean_alloc_ctor(1, 2, 8); -lean_ctor_set(x_1219, 0, x_1218); -lean_ctor_set(x_1219, 1, x_1165); -lean_ctor_set_uint64(x_1219, sizeof(void*)*2, x_1115); -if (lean_is_scalar(x_1139)) { - x_1220 = lean_alloc_ctor(1, 2, 8); -} else { - x_1220 = x_1139; -} -lean_ctor_set(x_1220, 0, x_1219); -lean_ctor_set(x_1220, 1, x_1190); -lean_ctor_set_uint64(x_1220, sizeof(void*)*2, x_1113); -lean_ctor_set(x_78, 0, x_1220); -if (lean_is_scalar(x_1138)) { - x_1221 = lean_alloc_ctor(1, 2, 0); -} else { - x_1221 = x_1138; -} -lean_ctor_set(x_1221, 0, x_78); -lean_ctor_set(x_1221, 1, x_656); -x_1222 = lean_box(0); -x_1223 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_1221, x_1222); -x_1224 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1223, x_1222); -x_1225 = l_Lean_MessageData_ofList(x_1224); -lean_dec(x_1224); -x_1226 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_1227 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1227, 0, x_1226); -lean_ctor_set(x_1227, 1, x_1225); -x_1228 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_1229 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1229, 0, x_1227); -lean_ctor_set(x_1229, 1, x_1228); -x_1230 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1229, x_5, x_6, x_655); -lean_dec(x_6); -x_1231 = lean_ctor_get(x_1230, 0); -lean_inc(x_1231); -x_1232 = lean_ctor_get(x_1230, 1); -lean_inc(x_1232); -if (lean_is_exclusive(x_1230)) { - lean_ctor_release(x_1230, 0); - lean_ctor_release(x_1230, 1); - x_1233 = x_1230; +if (lean_is_scalar(x_1111)) { + x_1206 = lean_alloc_ctor(1, 2, 8); } else { - lean_dec_ref(x_1230); - x_1233 = lean_box(0); + x_1206 = x_1111; } -if (lean_is_scalar(x_1233)) { - x_1234 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1206, 0, x_649); +lean_ctor_set(x_1206, 1, x_1112); +lean_ctor_set_uint64(x_1206, sizeof(void*)*2, x_1110); +if (lean_is_scalar(x_1108)) { + x_1207 = lean_alloc_ctor(1, 2, 8); } else { - x_1234 = x_1233; + x_1207 = x_1108; } -lean_ctor_set(x_1234, 0, x_1231); -lean_ctor_set(x_1234, 1, x_1232); -return x_1234; +lean_ctor_set(x_1207, 0, x_1206); +lean_ctor_set(x_1207, 1, x_1130); +lean_ctor_set_uint64(x_1207, sizeof(void*)*2, x_1107); +x_1208 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_1208, 0, x_1207); +lean_ctor_set(x_1208, 1, x_1155); +lean_ctor_set_uint64(x_1208, sizeof(void*)*2, x_1105); +if (lean_is_scalar(x_1129)) { + x_1209 = lean_alloc_ctor(1, 2, 8); +} else { + x_1209 = x_1129; +} +lean_ctor_set(x_1209, 0, x_1208); +lean_ctor_set(x_1209, 1, x_1180); +lean_ctor_set_uint64(x_1209, sizeof(void*)*2, x_1103); +lean_ctor_set(x_79, 0, x_1209); +if (lean_is_scalar(x_1128)) { + x_1210 = lean_alloc_ctor(1, 2, 0); +} else { + x_1210 = x_1128; +} +lean_ctor_set(x_1210, 0, x_79); +lean_ctor_set(x_1210, 1, x_651); +x_1211 = lean_box(0); +x_1212 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_1210, x_1211); +x_1213 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1212, x_1211); +x_1214 = l_Lean_MessageData_ofList(x_1213); +lean_dec(x_1213); +x_1215 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_1216 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1216, 0, x_1215); +lean_ctor_set(x_1216, 1, x_1214); +x_1217 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_1218 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1218, 0, x_1216); +lean_ctor_set(x_1218, 1, x_1217); +x_1219 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1218, x_6, x_7, x_650); +lean_dec(x_7); +x_1220 = lean_ctor_get(x_1219, 0); +lean_inc(x_1220); +x_1221 = lean_ctor_get(x_1219, 1); +lean_inc(x_1221); +if (lean_is_exclusive(x_1219)) { + lean_ctor_release(x_1219, 0); + lean_ctor_release(x_1219, 1); + x_1222 = x_1219; +} else { + lean_dec_ref(x_1219); + x_1222 = lean_box(0); +} +if (lean_is_scalar(x_1222)) { + x_1223 = lean_alloc_ctor(1, 2, 0); +} else { + x_1223 = x_1222; +} +lean_ctor_set(x_1223, 0, x_1220); +lean_ctor_set(x_1223, 1, x_1221); +return x_1223; } } } @@ -7397,627 +7464,632 @@ return x_1234; } else { -lean_object* x_1235; lean_object* x_1236; uint64_t x_1237; lean_object* x_1238; uint64_t x_1239; lean_object* x_1240; lean_object* x_1241; uint64_t x_1242; lean_object* x_1243; lean_object* x_1244; uint64_t x_1245; lean_object* x_1246; lean_object* x_1247; uint8_t x_1248; -x_1235 = lean_ctor_get(x_78, 1); -lean_inc(x_1235); -lean_dec(x_78); -x_1236 = lean_ctor_get(x_79, 1); -lean_inc(x_1236); -x_1237 = lean_ctor_get_uint64(x_79, sizeof(void*)*2); -x_1238 = lean_ctor_get(x_80, 1); -lean_inc(x_1238); -x_1239 = lean_ctor_get_uint64(x_80, sizeof(void*)*2); -if (lean_is_exclusive(x_80)) { - lean_ctor_release(x_80, 0); - lean_ctor_release(x_80, 1); - x_1240 = x_80; -} else { - lean_dec_ref(x_80); - x_1240 = lean_box(0); -} -x_1241 = lean_ctor_get(x_81, 1); -lean_inc(x_1241); -x_1242 = lean_ctor_get_uint64(x_81, sizeof(void*)*2); +lean_object* x_1224; lean_object* x_1225; uint64_t x_1226; lean_object* x_1227; uint64_t x_1228; lean_object* x_1229; lean_object* x_1230; uint64_t x_1231; lean_object* x_1232; lean_object* x_1233; uint64_t x_1234; lean_object* x_1235; lean_object* x_1236; uint8_t x_1237; +x_1224 = lean_ctor_get(x_79, 1); +lean_inc(x_1224); +lean_dec(x_79); +x_1225 = lean_ctor_get(x_80, 1); +lean_inc(x_1225); +x_1226 = lean_ctor_get_uint64(x_80, sizeof(void*)*2); +x_1227 = lean_ctor_get(x_81, 1); +lean_inc(x_1227); +x_1228 = lean_ctor_get_uint64(x_81, sizeof(void*)*2); if (lean_is_exclusive(x_81)) { lean_ctor_release(x_81, 0); lean_ctor_release(x_81, 1); - x_1243 = x_81; + x_1229 = x_81; } else { lean_dec_ref(x_81); - x_1243 = lean_box(0); + x_1229 = lean_box(0); } -x_1244 = lean_ctor_get(x_82, 1); -lean_inc(x_1244); -x_1245 = lean_ctor_get_uint64(x_82, sizeof(void*)*2); +x_1230 = lean_ctor_get(x_82, 1); +lean_inc(x_1230); +x_1231 = lean_ctor_get_uint64(x_82, sizeof(void*)*2); if (lean_is_exclusive(x_82)) { lean_ctor_release(x_82, 0); lean_ctor_release(x_82, 1); - x_1246 = x_82; + x_1232 = x_82; } else { lean_dec_ref(x_82); - x_1246 = lean_box(0); + x_1232 = lean_box(0); +} +x_1233 = lean_ctor_get(x_83, 1); +lean_inc(x_1233); +x_1234 = lean_ctor_get_uint64(x_83, sizeof(void*)*2); +if (lean_is_exclusive(x_83)) { + lean_ctor_release(x_83, 0); + lean_ctor_release(x_83, 1); + x_1235 = x_83; +} else { + lean_dec_ref(x_83); + x_1235 = lean_box(0); } -x_1247 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__1; -x_1248 = lean_string_dec_eq(x_1244, x_1247); -lean_dec(x_1244); -if (x_1248 == 0) +x_1236 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__1; +x_1237 = lean_string_dec_eq(x_1233, x_1236); +lean_dec(x_1233); +if (x_1237 == 0) { -lean_dec(x_1246); -lean_dec(x_1243); -lean_dec(x_1241); -lean_dec(x_1240); -lean_dec(x_1238); -lean_dec(x_1236); lean_dec(x_1235); -if (lean_obj_tag(x_656) == 0) +lean_dec(x_1232); +lean_dec(x_1230); +lean_dec(x_1229); +lean_dec(x_1227); +lean_dec(x_1225); +lean_dec(x_1224); +lean_dec(x_3); +if (lean_obj_tag(x_651) == 0) { -lean_object* x_1249; -lean_dec(x_9); -x_1249 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_79, x_5, x_6, x_655); +lean_object* x_1238; +lean_dec(x_10); +x_1238 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_80, x_6, x_7, x_650); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_79); -return x_1249; +lean_dec(x_80); +return x_1238; } else { -lean_object* x_1250; lean_object* x_1251; lean_object* x_1252; lean_object* x_1253; lean_object* x_1254; lean_object* x_1255; lean_object* x_1256; lean_object* x_1257; lean_object* x_1258; lean_object* x_1259; lean_object* x_1260; lean_object* x_1261; lean_object* x_1262; -lean_dec(x_656); -lean_dec(x_79); +lean_object* x_1239; lean_object* x_1240; lean_object* x_1241; lean_object* x_1242; lean_object* x_1243; lean_object* x_1244; lean_object* x_1245; lean_object* x_1246; lean_object* x_1247; lean_object* x_1248; lean_object* x_1249; lean_object* x_1250; lean_object* x_1251; +lean_dec(x_651); +lean_dec(x_80); lean_dec(x_1); -x_1250 = lean_box(0); -x_1251 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_9, x_1250); -x_1252 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1251, x_1250); -x_1253 = l_Lean_MessageData_ofList(x_1252); -lean_dec(x_1252); -x_1254 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_1255 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1255, 0, x_1254); -lean_ctor_set(x_1255, 1, x_1253); -x_1256 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_1257 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1257, 0, x_1255); -lean_ctor_set(x_1257, 1, x_1256); -x_1258 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1257, x_5, x_6, x_655); -lean_dec(x_6); -x_1259 = lean_ctor_get(x_1258, 0); -lean_inc(x_1259); -x_1260 = lean_ctor_get(x_1258, 1); -lean_inc(x_1260); -if (lean_is_exclusive(x_1258)) { - lean_ctor_release(x_1258, 0); - lean_ctor_release(x_1258, 1); - x_1261 = x_1258; +x_1239 = lean_box(0); +x_1240 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_10, x_1239); +x_1241 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1240, x_1239); +x_1242 = l_Lean_MessageData_ofList(x_1241); +lean_dec(x_1241); +x_1243 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_1244 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1244, 0, x_1243); +lean_ctor_set(x_1244, 1, x_1242); +x_1245 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_1246 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1246, 0, x_1244); +lean_ctor_set(x_1246, 1, x_1245); +x_1247 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1246, x_6, x_7, x_650); +lean_dec(x_7); +x_1248 = lean_ctor_get(x_1247, 0); +lean_inc(x_1248); +x_1249 = lean_ctor_get(x_1247, 1); +lean_inc(x_1249); +if (lean_is_exclusive(x_1247)) { + lean_ctor_release(x_1247, 0); + lean_ctor_release(x_1247, 1); + x_1250 = x_1247; } else { - lean_dec_ref(x_1258); - x_1261 = lean_box(0); + lean_dec_ref(x_1247); + x_1250 = lean_box(0); } -if (lean_is_scalar(x_1261)) { - x_1262 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_1250)) { + x_1251 = lean_alloc_ctor(1, 2, 0); } else { - x_1262 = x_1261; + x_1251 = x_1250; } -lean_ctor_set(x_1262, 0, x_1259); -lean_ctor_set(x_1262, 1, x_1260); -return x_1262; +lean_ctor_set(x_1251, 0, x_1248); +lean_ctor_set(x_1251, 1, x_1249); +return x_1251; } } else { -lean_object* x_1263; lean_object* x_1264; lean_object* x_1265; uint8_t x_1266; -if (lean_is_exclusive(x_9)) { - lean_ctor_release(x_9, 0); - lean_ctor_release(x_9, 1); - x_1263 = x_9; +lean_object* x_1252; lean_object* x_1253; lean_object* x_1254; uint8_t x_1255; +if (lean_is_exclusive(x_10)) { + lean_ctor_release(x_10, 0); + lean_ctor_release(x_10, 1); + x_1252 = x_10; } else { - lean_dec_ref(x_9); - x_1263 = lean_box(0); + lean_dec_ref(x_10); + x_1252 = lean_box(0); } -if (lean_is_exclusive(x_79)) { - lean_ctor_release(x_79, 0); - lean_ctor_release(x_79, 1); - x_1264 = x_79; +if (lean_is_exclusive(x_80)) { + lean_ctor_release(x_80, 0); + lean_ctor_release(x_80, 1); + x_1253 = x_80; } else { - lean_dec_ref(x_79); - x_1264 = lean_box(0); + lean_dec_ref(x_80); + x_1253 = lean_box(0); } -x_1265 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__4; -x_1266 = lean_string_dec_eq(x_1241, x_1265); -if (x_1266 == 0) +x_1254 = l_List_filterMap___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__14___closed__4; +x_1255 = lean_string_dec_eq(x_1230, x_1254); +if (x_1255 == 0) { -if (lean_obj_tag(x_656) == 0) +lean_dec(x_3); +if (lean_obj_tag(x_651) == 0) { -lean_object* x_1267; lean_object* x_1268; lean_object* x_1269; lean_object* x_1270; lean_object* x_1271; -lean_dec(x_1263); -lean_dec(x_1235); -if (lean_is_scalar(x_1246)) { - x_1267 = lean_alloc_ctor(1, 2, 8); +lean_object* x_1256; lean_object* x_1257; lean_object* x_1258; lean_object* x_1259; lean_object* x_1260; +lean_dec(x_1252); +lean_dec(x_1224); +if (lean_is_scalar(x_1235)) { + x_1256 = lean_alloc_ctor(1, 2, 8); } else { - x_1267 = x_1246; + x_1256 = x_1235; } -lean_ctor_set(x_1267, 0, x_654); -lean_ctor_set(x_1267, 1, x_1247); -lean_ctor_set_uint64(x_1267, sizeof(void*)*2, x_1245); -if (lean_is_scalar(x_1243)) { - x_1268 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_1256, 0, x_649); +lean_ctor_set(x_1256, 1, x_1236); +lean_ctor_set_uint64(x_1256, sizeof(void*)*2, x_1234); +if (lean_is_scalar(x_1232)) { + x_1257 = lean_alloc_ctor(1, 2, 8); } else { - x_1268 = x_1243; + x_1257 = x_1232; } -lean_ctor_set(x_1268, 0, x_1267); -lean_ctor_set(x_1268, 1, x_1241); -lean_ctor_set_uint64(x_1268, sizeof(void*)*2, x_1242); -if (lean_is_scalar(x_1240)) { - x_1269 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_1257, 0, x_1256); +lean_ctor_set(x_1257, 1, x_1230); +lean_ctor_set_uint64(x_1257, sizeof(void*)*2, x_1231); +if (lean_is_scalar(x_1229)) { + x_1258 = lean_alloc_ctor(1, 2, 8); } else { - x_1269 = x_1240; + x_1258 = x_1229; } -lean_ctor_set(x_1269, 0, x_1268); -lean_ctor_set(x_1269, 1, x_1238); -lean_ctor_set_uint64(x_1269, sizeof(void*)*2, x_1239); -if (lean_is_scalar(x_1264)) { - x_1270 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_1258, 0, x_1257); +lean_ctor_set(x_1258, 1, x_1227); +lean_ctor_set_uint64(x_1258, sizeof(void*)*2, x_1228); +if (lean_is_scalar(x_1253)) { + x_1259 = lean_alloc_ctor(1, 2, 8); } else { - x_1270 = x_1264; + x_1259 = x_1253; } -lean_ctor_set(x_1270, 0, x_1269); -lean_ctor_set(x_1270, 1, x_1236); -lean_ctor_set_uint64(x_1270, sizeof(void*)*2, x_1237); -x_1271 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_1270, x_5, x_6, x_655); +lean_ctor_set(x_1259, 0, x_1258); +lean_ctor_set(x_1259, 1, x_1225); +lean_ctor_set_uint64(x_1259, sizeof(void*)*2, x_1226); +x_1260 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_1259, x_6, x_7, x_650); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_1270); -return x_1271; +lean_dec(x_1259); +return x_1260; } else { -lean_object* x_1272; lean_object* x_1273; lean_object* x_1274; lean_object* x_1275; lean_object* x_1276; lean_object* x_1277; lean_object* x_1278; lean_object* x_1279; lean_object* x_1280; lean_object* x_1281; lean_object* x_1282; lean_object* x_1283; lean_object* x_1284; lean_object* x_1285; lean_object* x_1286; lean_object* x_1287; lean_object* x_1288; lean_object* x_1289; lean_object* x_1290; +lean_object* x_1261; lean_object* x_1262; lean_object* x_1263; lean_object* x_1264; lean_object* x_1265; lean_object* x_1266; lean_object* x_1267; lean_object* x_1268; lean_object* x_1269; lean_object* x_1270; lean_object* x_1271; lean_object* x_1272; lean_object* x_1273; lean_object* x_1274; lean_object* x_1275; lean_object* x_1276; lean_object* x_1277; lean_object* x_1278; lean_object* x_1279; lean_dec(x_1); -if (lean_is_scalar(x_1246)) { - x_1272 = lean_alloc_ctor(1, 2, 8); -} else { - x_1272 = x_1246; -} -lean_ctor_set(x_1272, 0, x_654); -lean_ctor_set(x_1272, 1, x_1247); -lean_ctor_set_uint64(x_1272, sizeof(void*)*2, x_1245); -if (lean_is_scalar(x_1243)) { - x_1273 = lean_alloc_ctor(1, 2, 8); -} else { - x_1273 = x_1243; -} -lean_ctor_set(x_1273, 0, x_1272); -lean_ctor_set(x_1273, 1, x_1241); -lean_ctor_set_uint64(x_1273, sizeof(void*)*2, x_1242); -if (lean_is_scalar(x_1240)) { - x_1274 = lean_alloc_ctor(1, 2, 8); -} else { - x_1274 = x_1240; -} -lean_ctor_set(x_1274, 0, x_1273); -lean_ctor_set(x_1274, 1, x_1238); -lean_ctor_set_uint64(x_1274, sizeof(void*)*2, x_1239); -if (lean_is_scalar(x_1264)) { - x_1275 = lean_alloc_ctor(1, 2, 8); -} else { - x_1275 = x_1264; -} -lean_ctor_set(x_1275, 0, x_1274); -lean_ctor_set(x_1275, 1, x_1236); -lean_ctor_set_uint64(x_1275, sizeof(void*)*2, x_1237); -x_1276 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1276, 0, x_1275); -lean_ctor_set(x_1276, 1, x_1235); -if (lean_is_scalar(x_1263)) { - x_1277 = lean_alloc_ctor(1, 2, 0); -} else { - x_1277 = x_1263; -} -lean_ctor_set(x_1277, 0, x_1276); -lean_ctor_set(x_1277, 1, x_656); -x_1278 = lean_box(0); -x_1279 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_1277, x_1278); -x_1280 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1279, x_1278); -x_1281 = l_Lean_MessageData_ofList(x_1280); -lean_dec(x_1280); -x_1282 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_1283 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1283, 0, x_1282); -lean_ctor_set(x_1283, 1, x_1281); -x_1284 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_1285 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1285, 0, x_1283); -lean_ctor_set(x_1285, 1, x_1284); -x_1286 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1285, x_5, x_6, x_655); -lean_dec(x_6); -x_1287 = lean_ctor_get(x_1286, 0); -lean_inc(x_1287); -x_1288 = lean_ctor_get(x_1286, 1); -lean_inc(x_1288); -if (lean_is_exclusive(x_1286)) { - lean_ctor_release(x_1286, 0); - lean_ctor_release(x_1286, 1); - x_1289 = x_1286; +if (lean_is_scalar(x_1235)) { + x_1261 = lean_alloc_ctor(1, 2, 8); +} else { + x_1261 = x_1235; +} +lean_ctor_set(x_1261, 0, x_649); +lean_ctor_set(x_1261, 1, x_1236); +lean_ctor_set_uint64(x_1261, sizeof(void*)*2, x_1234); +if (lean_is_scalar(x_1232)) { + x_1262 = lean_alloc_ctor(1, 2, 8); +} else { + x_1262 = x_1232; +} +lean_ctor_set(x_1262, 0, x_1261); +lean_ctor_set(x_1262, 1, x_1230); +lean_ctor_set_uint64(x_1262, sizeof(void*)*2, x_1231); +if (lean_is_scalar(x_1229)) { + x_1263 = lean_alloc_ctor(1, 2, 8); +} else { + x_1263 = x_1229; +} +lean_ctor_set(x_1263, 0, x_1262); +lean_ctor_set(x_1263, 1, x_1227); +lean_ctor_set_uint64(x_1263, sizeof(void*)*2, x_1228); +if (lean_is_scalar(x_1253)) { + x_1264 = lean_alloc_ctor(1, 2, 8); +} else { + x_1264 = x_1253; +} +lean_ctor_set(x_1264, 0, x_1263); +lean_ctor_set(x_1264, 1, x_1225); +lean_ctor_set_uint64(x_1264, sizeof(void*)*2, x_1226); +x_1265 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1265, 0, x_1264); +lean_ctor_set(x_1265, 1, x_1224); +if (lean_is_scalar(x_1252)) { + x_1266 = lean_alloc_ctor(1, 2, 0); +} else { + x_1266 = x_1252; +} +lean_ctor_set(x_1266, 0, x_1265); +lean_ctor_set(x_1266, 1, x_651); +x_1267 = lean_box(0); +x_1268 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_1266, x_1267); +x_1269 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1268, x_1267); +x_1270 = l_Lean_MessageData_ofList(x_1269); +lean_dec(x_1269); +x_1271 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_1272 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1272, 0, x_1271); +lean_ctor_set(x_1272, 1, x_1270); +x_1273 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_1274 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1274, 0, x_1272); +lean_ctor_set(x_1274, 1, x_1273); +x_1275 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1274, x_6, x_7, x_650); +lean_dec(x_7); +x_1276 = lean_ctor_get(x_1275, 0); +lean_inc(x_1276); +x_1277 = lean_ctor_get(x_1275, 1); +lean_inc(x_1277); +if (lean_is_exclusive(x_1275)) { + lean_ctor_release(x_1275, 0); + lean_ctor_release(x_1275, 1); + x_1278 = x_1275; } else { - lean_dec_ref(x_1286); - x_1289 = lean_box(0); + lean_dec_ref(x_1275); + x_1278 = lean_box(0); } -if (lean_is_scalar(x_1289)) { - x_1290 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_1278)) { + x_1279 = lean_alloc_ctor(1, 2, 0); } else { - x_1290 = x_1289; + x_1279 = x_1278; } -lean_ctor_set(x_1290, 0, x_1287); -lean_ctor_set(x_1290, 1, x_1288); -return x_1290; +lean_ctor_set(x_1279, 0, x_1276); +lean_ctor_set(x_1279, 1, x_1277); +return x_1279; } } else { -lean_object* x_1291; uint8_t x_1292; -lean_dec(x_1241); -x_1291 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__9; -x_1292 = lean_string_dec_eq(x_1238, x_1291); -if (x_1292 == 0) +lean_object* x_1280; uint8_t x_1281; +lean_dec(x_1230); +x_1280 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__11; +x_1281 = lean_string_dec_eq(x_1227, x_1280); +if (x_1281 == 0) { -if (lean_obj_tag(x_656) == 0) +lean_dec(x_3); +if (lean_obj_tag(x_651) == 0) { -lean_object* x_1293; lean_object* x_1294; lean_object* x_1295; lean_object* x_1296; lean_object* x_1297; -lean_dec(x_1263); -lean_dec(x_1235); -if (lean_is_scalar(x_1246)) { - x_1293 = lean_alloc_ctor(1, 2, 8); +lean_object* x_1282; lean_object* x_1283; lean_object* x_1284; lean_object* x_1285; lean_object* x_1286; +lean_dec(x_1252); +lean_dec(x_1224); +if (lean_is_scalar(x_1235)) { + x_1282 = lean_alloc_ctor(1, 2, 8); } else { - x_1293 = x_1246; + x_1282 = x_1235; } -lean_ctor_set(x_1293, 0, x_654); -lean_ctor_set(x_1293, 1, x_1247); -lean_ctor_set_uint64(x_1293, sizeof(void*)*2, x_1245); -if (lean_is_scalar(x_1243)) { - x_1294 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_1282, 0, x_649); +lean_ctor_set(x_1282, 1, x_1236); +lean_ctor_set_uint64(x_1282, sizeof(void*)*2, x_1234); +if (lean_is_scalar(x_1232)) { + x_1283 = lean_alloc_ctor(1, 2, 8); } else { - x_1294 = x_1243; + x_1283 = x_1232; } -lean_ctor_set(x_1294, 0, x_1293); -lean_ctor_set(x_1294, 1, x_1265); -lean_ctor_set_uint64(x_1294, sizeof(void*)*2, x_1242); -if (lean_is_scalar(x_1240)) { - x_1295 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_1283, 0, x_1282); +lean_ctor_set(x_1283, 1, x_1254); +lean_ctor_set_uint64(x_1283, sizeof(void*)*2, x_1231); +if (lean_is_scalar(x_1229)) { + x_1284 = lean_alloc_ctor(1, 2, 8); } else { - x_1295 = x_1240; + x_1284 = x_1229; } -lean_ctor_set(x_1295, 0, x_1294); -lean_ctor_set(x_1295, 1, x_1238); -lean_ctor_set_uint64(x_1295, sizeof(void*)*2, x_1239); -if (lean_is_scalar(x_1264)) { - x_1296 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_1284, 0, x_1283); +lean_ctor_set(x_1284, 1, x_1227); +lean_ctor_set_uint64(x_1284, sizeof(void*)*2, x_1228); +if (lean_is_scalar(x_1253)) { + x_1285 = lean_alloc_ctor(1, 2, 8); } else { - x_1296 = x_1264; + x_1285 = x_1253; } -lean_ctor_set(x_1296, 0, x_1295); -lean_ctor_set(x_1296, 1, x_1236); -lean_ctor_set_uint64(x_1296, sizeof(void*)*2, x_1237); -x_1297 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_1296, x_5, x_6, x_655); +lean_ctor_set(x_1285, 0, x_1284); +lean_ctor_set(x_1285, 1, x_1225); +lean_ctor_set_uint64(x_1285, sizeof(void*)*2, x_1226); +x_1286 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_1285, x_6, x_7, x_650); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_1296); -return x_1297; +lean_dec(x_1285); +return x_1286; } else { -lean_object* x_1298; lean_object* x_1299; lean_object* x_1300; lean_object* x_1301; lean_object* x_1302; lean_object* x_1303; lean_object* x_1304; lean_object* x_1305; lean_object* x_1306; lean_object* x_1307; lean_object* x_1308; lean_object* x_1309; lean_object* x_1310; lean_object* x_1311; lean_object* x_1312; lean_object* x_1313; lean_object* x_1314; lean_object* x_1315; lean_object* x_1316; +lean_object* x_1287; lean_object* x_1288; lean_object* x_1289; lean_object* x_1290; lean_object* x_1291; lean_object* x_1292; lean_object* x_1293; lean_object* x_1294; lean_object* x_1295; lean_object* x_1296; lean_object* x_1297; lean_object* x_1298; lean_object* x_1299; lean_object* x_1300; lean_object* x_1301; lean_object* x_1302; lean_object* x_1303; lean_object* x_1304; lean_object* x_1305; lean_dec(x_1); -if (lean_is_scalar(x_1246)) { - x_1298 = lean_alloc_ctor(1, 2, 8); -} else { - x_1298 = x_1246; -} -lean_ctor_set(x_1298, 0, x_654); -lean_ctor_set(x_1298, 1, x_1247); -lean_ctor_set_uint64(x_1298, sizeof(void*)*2, x_1245); -if (lean_is_scalar(x_1243)) { - x_1299 = lean_alloc_ctor(1, 2, 8); -} else { - x_1299 = x_1243; -} -lean_ctor_set(x_1299, 0, x_1298); -lean_ctor_set(x_1299, 1, x_1265); -lean_ctor_set_uint64(x_1299, sizeof(void*)*2, x_1242); -if (lean_is_scalar(x_1240)) { - x_1300 = lean_alloc_ctor(1, 2, 8); -} else { - x_1300 = x_1240; -} -lean_ctor_set(x_1300, 0, x_1299); -lean_ctor_set(x_1300, 1, x_1238); -lean_ctor_set_uint64(x_1300, sizeof(void*)*2, x_1239); -if (lean_is_scalar(x_1264)) { - x_1301 = lean_alloc_ctor(1, 2, 8); -} else { - x_1301 = x_1264; -} -lean_ctor_set(x_1301, 0, x_1300); -lean_ctor_set(x_1301, 1, x_1236); -lean_ctor_set_uint64(x_1301, sizeof(void*)*2, x_1237); -x_1302 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1302, 0, x_1301); -lean_ctor_set(x_1302, 1, x_1235); -if (lean_is_scalar(x_1263)) { - x_1303 = lean_alloc_ctor(1, 2, 0); -} else { - x_1303 = x_1263; -} -lean_ctor_set(x_1303, 0, x_1302); -lean_ctor_set(x_1303, 1, x_656); -x_1304 = lean_box(0); -x_1305 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_1303, x_1304); -x_1306 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1305, x_1304); -x_1307 = l_Lean_MessageData_ofList(x_1306); -lean_dec(x_1306); -x_1308 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_1309 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1309, 0, x_1308); -lean_ctor_set(x_1309, 1, x_1307); -x_1310 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_1311 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1311, 0, x_1309); -lean_ctor_set(x_1311, 1, x_1310); -x_1312 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1311, x_5, x_6, x_655); -lean_dec(x_6); -x_1313 = lean_ctor_get(x_1312, 0); -lean_inc(x_1313); -x_1314 = lean_ctor_get(x_1312, 1); -lean_inc(x_1314); -if (lean_is_exclusive(x_1312)) { - lean_ctor_release(x_1312, 0); - lean_ctor_release(x_1312, 1); - x_1315 = x_1312; +if (lean_is_scalar(x_1235)) { + x_1287 = lean_alloc_ctor(1, 2, 8); +} else { + x_1287 = x_1235; +} +lean_ctor_set(x_1287, 0, x_649); +lean_ctor_set(x_1287, 1, x_1236); +lean_ctor_set_uint64(x_1287, sizeof(void*)*2, x_1234); +if (lean_is_scalar(x_1232)) { + x_1288 = lean_alloc_ctor(1, 2, 8); +} else { + x_1288 = x_1232; +} +lean_ctor_set(x_1288, 0, x_1287); +lean_ctor_set(x_1288, 1, x_1254); +lean_ctor_set_uint64(x_1288, sizeof(void*)*2, x_1231); +if (lean_is_scalar(x_1229)) { + x_1289 = lean_alloc_ctor(1, 2, 8); +} else { + x_1289 = x_1229; +} +lean_ctor_set(x_1289, 0, x_1288); +lean_ctor_set(x_1289, 1, x_1227); +lean_ctor_set_uint64(x_1289, sizeof(void*)*2, x_1228); +if (lean_is_scalar(x_1253)) { + x_1290 = lean_alloc_ctor(1, 2, 8); +} else { + x_1290 = x_1253; +} +lean_ctor_set(x_1290, 0, x_1289); +lean_ctor_set(x_1290, 1, x_1225); +lean_ctor_set_uint64(x_1290, sizeof(void*)*2, x_1226); +x_1291 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1291, 0, x_1290); +lean_ctor_set(x_1291, 1, x_1224); +if (lean_is_scalar(x_1252)) { + x_1292 = lean_alloc_ctor(1, 2, 0); +} else { + x_1292 = x_1252; +} +lean_ctor_set(x_1292, 0, x_1291); +lean_ctor_set(x_1292, 1, x_651); +x_1293 = lean_box(0); +x_1294 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_1292, x_1293); +x_1295 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1294, x_1293); +x_1296 = l_Lean_MessageData_ofList(x_1295); +lean_dec(x_1295); +x_1297 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_1298 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1298, 0, x_1297); +lean_ctor_set(x_1298, 1, x_1296); +x_1299 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_1300 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1300, 0, x_1298); +lean_ctor_set(x_1300, 1, x_1299); +x_1301 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1300, x_6, x_7, x_650); +lean_dec(x_7); +x_1302 = lean_ctor_get(x_1301, 0); +lean_inc(x_1302); +x_1303 = lean_ctor_get(x_1301, 1); +lean_inc(x_1303); +if (lean_is_exclusive(x_1301)) { + lean_ctor_release(x_1301, 0); + lean_ctor_release(x_1301, 1); + x_1304 = x_1301; } else { - lean_dec_ref(x_1312); - x_1315 = lean_box(0); + lean_dec_ref(x_1301); + x_1304 = lean_box(0); } -if (lean_is_scalar(x_1315)) { - x_1316 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_1304)) { + x_1305 = lean_alloc_ctor(1, 2, 0); } else { - x_1316 = x_1315; + x_1305 = x_1304; } -lean_ctor_set(x_1316, 0, x_1313); -lean_ctor_set(x_1316, 1, x_1314); -return x_1316; +lean_ctor_set(x_1305, 0, x_1302); +lean_ctor_set(x_1305, 1, x_1303); +return x_1305; } } else { -lean_object* x_1317; uint8_t x_1318; -lean_dec(x_1238); -x_1317 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__7; -x_1318 = lean_string_dec_eq(x_1236, x_1317); -if (x_1318 == 0) +lean_object* x_1306; uint8_t x_1307; +lean_dec(x_1227); +x_1306 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__7; +x_1307 = lean_string_dec_eq(x_1225, x_1306); +if (x_1307 == 0) { -if (lean_obj_tag(x_656) == 0) +lean_dec(x_3); +if (lean_obj_tag(x_651) == 0) { -lean_object* x_1319; lean_object* x_1320; lean_object* x_1321; lean_object* x_1322; lean_object* x_1323; -lean_dec(x_1263); -lean_dec(x_1235); -if (lean_is_scalar(x_1246)) { - x_1319 = lean_alloc_ctor(1, 2, 8); +lean_object* x_1308; lean_object* x_1309; lean_object* x_1310; lean_object* x_1311; lean_object* x_1312; +lean_dec(x_1252); +lean_dec(x_1224); +if (lean_is_scalar(x_1235)) { + x_1308 = lean_alloc_ctor(1, 2, 8); } else { - x_1319 = x_1246; + x_1308 = x_1235; } -lean_ctor_set(x_1319, 0, x_654); -lean_ctor_set(x_1319, 1, x_1247); -lean_ctor_set_uint64(x_1319, sizeof(void*)*2, x_1245); -if (lean_is_scalar(x_1243)) { - x_1320 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_1308, 0, x_649); +lean_ctor_set(x_1308, 1, x_1236); +lean_ctor_set_uint64(x_1308, sizeof(void*)*2, x_1234); +if (lean_is_scalar(x_1232)) { + x_1309 = lean_alloc_ctor(1, 2, 8); } else { - x_1320 = x_1243; + x_1309 = x_1232; } -lean_ctor_set(x_1320, 0, x_1319); -lean_ctor_set(x_1320, 1, x_1265); -lean_ctor_set_uint64(x_1320, sizeof(void*)*2, x_1242); -if (lean_is_scalar(x_1240)) { - x_1321 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_1309, 0, x_1308); +lean_ctor_set(x_1309, 1, x_1254); +lean_ctor_set_uint64(x_1309, sizeof(void*)*2, x_1231); +if (lean_is_scalar(x_1229)) { + x_1310 = lean_alloc_ctor(1, 2, 8); } else { - x_1321 = x_1240; + x_1310 = x_1229; } -lean_ctor_set(x_1321, 0, x_1320); -lean_ctor_set(x_1321, 1, x_1291); -lean_ctor_set_uint64(x_1321, sizeof(void*)*2, x_1239); -if (lean_is_scalar(x_1264)) { - x_1322 = lean_alloc_ctor(1, 2, 8); +lean_ctor_set(x_1310, 0, x_1309); +lean_ctor_set(x_1310, 1, x_1280); +lean_ctor_set_uint64(x_1310, sizeof(void*)*2, x_1228); +if (lean_is_scalar(x_1253)) { + x_1311 = lean_alloc_ctor(1, 2, 8); } else { - x_1322 = x_1264; + x_1311 = x_1253; } -lean_ctor_set(x_1322, 0, x_1321); -lean_ctor_set(x_1322, 1, x_1236); -lean_ctor_set_uint64(x_1322, sizeof(void*)*2, x_1237); -x_1323 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_1322, x_5, x_6, x_655); +lean_ctor_set(x_1311, 0, x_1310); +lean_ctor_set(x_1311, 1, x_1225); +lean_ctor_set_uint64(x_1311, sizeof(void*)*2, x_1226); +x_1312 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_1311, x_6, x_7, x_650); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_1322); -return x_1323; +lean_dec(x_1311); +return x_1312; } else { -lean_object* x_1324; lean_object* x_1325; lean_object* x_1326; lean_object* x_1327; lean_object* x_1328; lean_object* x_1329; lean_object* x_1330; lean_object* x_1331; lean_object* x_1332; lean_object* x_1333; lean_object* x_1334; lean_object* x_1335; lean_object* x_1336; lean_object* x_1337; lean_object* x_1338; lean_object* x_1339; lean_object* x_1340; lean_object* x_1341; lean_object* x_1342; +lean_object* x_1313; lean_object* x_1314; lean_object* x_1315; lean_object* x_1316; lean_object* x_1317; lean_object* x_1318; lean_object* x_1319; lean_object* x_1320; lean_object* x_1321; lean_object* x_1322; lean_object* x_1323; lean_object* x_1324; lean_object* x_1325; lean_object* x_1326; lean_object* x_1327; lean_object* x_1328; lean_object* x_1329; lean_object* x_1330; lean_object* x_1331; lean_dec(x_1); -if (lean_is_scalar(x_1246)) { - x_1324 = lean_alloc_ctor(1, 2, 8); -} else { - x_1324 = x_1246; -} -lean_ctor_set(x_1324, 0, x_654); -lean_ctor_set(x_1324, 1, x_1247); -lean_ctor_set_uint64(x_1324, sizeof(void*)*2, x_1245); -if (lean_is_scalar(x_1243)) { - x_1325 = lean_alloc_ctor(1, 2, 8); -} else { - x_1325 = x_1243; -} -lean_ctor_set(x_1325, 0, x_1324); -lean_ctor_set(x_1325, 1, x_1265); -lean_ctor_set_uint64(x_1325, sizeof(void*)*2, x_1242); -if (lean_is_scalar(x_1240)) { - x_1326 = lean_alloc_ctor(1, 2, 8); -} else { - x_1326 = x_1240; -} -lean_ctor_set(x_1326, 0, x_1325); -lean_ctor_set(x_1326, 1, x_1291); -lean_ctor_set_uint64(x_1326, sizeof(void*)*2, x_1239); -if (lean_is_scalar(x_1264)) { - x_1327 = lean_alloc_ctor(1, 2, 8); -} else { - x_1327 = x_1264; -} -lean_ctor_set(x_1327, 0, x_1326); -lean_ctor_set(x_1327, 1, x_1236); -lean_ctor_set_uint64(x_1327, sizeof(void*)*2, x_1237); -x_1328 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1328, 0, x_1327); -lean_ctor_set(x_1328, 1, x_1235); -if (lean_is_scalar(x_1263)) { - x_1329 = lean_alloc_ctor(1, 2, 0); -} else { - x_1329 = x_1263; -} -lean_ctor_set(x_1329, 0, x_1328); -lean_ctor_set(x_1329, 1, x_656); -x_1330 = lean_box(0); -x_1331 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_1329, x_1330); -x_1332 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1331, x_1330); -x_1333 = l_Lean_MessageData_ofList(x_1332); -lean_dec(x_1332); -x_1334 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_1335 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1335, 0, x_1334); -lean_ctor_set(x_1335, 1, x_1333); -x_1336 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_1337 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1337, 0, x_1335); -lean_ctor_set(x_1337, 1, x_1336); -x_1338 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1337, x_5, x_6, x_655); -lean_dec(x_6); -x_1339 = lean_ctor_get(x_1338, 0); -lean_inc(x_1339); -x_1340 = lean_ctor_get(x_1338, 1); -lean_inc(x_1340); -if (lean_is_exclusive(x_1338)) { - lean_ctor_release(x_1338, 0); - lean_ctor_release(x_1338, 1); - x_1341 = x_1338; +if (lean_is_scalar(x_1235)) { + x_1313 = lean_alloc_ctor(1, 2, 8); +} else { + x_1313 = x_1235; +} +lean_ctor_set(x_1313, 0, x_649); +lean_ctor_set(x_1313, 1, x_1236); +lean_ctor_set_uint64(x_1313, sizeof(void*)*2, x_1234); +if (lean_is_scalar(x_1232)) { + x_1314 = lean_alloc_ctor(1, 2, 8); +} else { + x_1314 = x_1232; +} +lean_ctor_set(x_1314, 0, x_1313); +lean_ctor_set(x_1314, 1, x_1254); +lean_ctor_set_uint64(x_1314, sizeof(void*)*2, x_1231); +if (lean_is_scalar(x_1229)) { + x_1315 = lean_alloc_ctor(1, 2, 8); +} else { + x_1315 = x_1229; +} +lean_ctor_set(x_1315, 0, x_1314); +lean_ctor_set(x_1315, 1, x_1280); +lean_ctor_set_uint64(x_1315, sizeof(void*)*2, x_1228); +if (lean_is_scalar(x_1253)) { + x_1316 = lean_alloc_ctor(1, 2, 8); +} else { + x_1316 = x_1253; +} +lean_ctor_set(x_1316, 0, x_1315); +lean_ctor_set(x_1316, 1, x_1225); +lean_ctor_set_uint64(x_1316, sizeof(void*)*2, x_1226); +x_1317 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1317, 0, x_1316); +lean_ctor_set(x_1317, 1, x_1224); +if (lean_is_scalar(x_1252)) { + x_1318 = lean_alloc_ctor(1, 2, 0); +} else { + x_1318 = x_1252; +} +lean_ctor_set(x_1318, 0, x_1317); +lean_ctor_set(x_1318, 1, x_651); +x_1319 = lean_box(0); +x_1320 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_1318, x_1319); +x_1321 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1320, x_1319); +x_1322 = l_Lean_MessageData_ofList(x_1321); +lean_dec(x_1321); +x_1323 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_1324 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1324, 0, x_1323); +lean_ctor_set(x_1324, 1, x_1322); +x_1325 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_1326 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1326, 0, x_1324); +lean_ctor_set(x_1326, 1, x_1325); +x_1327 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1326, x_6, x_7, x_650); +lean_dec(x_7); +x_1328 = lean_ctor_get(x_1327, 0); +lean_inc(x_1328); +x_1329 = lean_ctor_get(x_1327, 1); +lean_inc(x_1329); +if (lean_is_exclusive(x_1327)) { + lean_ctor_release(x_1327, 0); + lean_ctor_release(x_1327, 1); + x_1330 = x_1327; } else { - lean_dec_ref(x_1338); - x_1341 = lean_box(0); + lean_dec_ref(x_1327); + x_1330 = lean_box(0); } -if (lean_is_scalar(x_1341)) { - x_1342 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_1330)) { + x_1331 = lean_alloc_ctor(1, 2, 0); } else { - x_1342 = x_1341; + x_1331 = x_1330; } -lean_ctor_set(x_1342, 0, x_1339); -lean_ctor_set(x_1342, 1, x_1340); -return x_1342; +lean_ctor_set(x_1331, 0, x_1328); +lean_ctor_set(x_1331, 1, x_1329); +return x_1331; } } else { -lean_dec(x_1236); -if (lean_obj_tag(x_656) == 0) +lean_dec(x_1225); +if (lean_obj_tag(x_651) == 0) { -lean_object* x_1343; lean_object* x_1344; -lean_dec(x_1264); -lean_dec(x_1263); -lean_dec(x_1246); -lean_dec(x_1243); -lean_dec(x_1240); +lean_object* x_1332; +lean_dec(x_1253); +lean_dec(x_1252); lean_dec(x_1235); -x_1343 = l_Lean_identKind; -x_1344 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_1343, x_5, x_6, x_655); +lean_dec(x_1232); +lean_dec(x_1229); +lean_dec(x_1224); +x_1332 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_3, x_6, x_7, x_650); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -return x_1344; +lean_dec(x_3); +return x_1332; } else { -lean_object* x_1345; lean_object* x_1346; lean_object* x_1347; lean_object* x_1348; lean_object* x_1349; lean_object* x_1350; lean_object* x_1351; lean_object* x_1352; lean_object* x_1353; lean_object* x_1354; lean_object* x_1355; lean_object* x_1356; lean_object* x_1357; lean_object* x_1358; lean_object* x_1359; lean_object* x_1360; lean_object* x_1361; lean_object* x_1362; lean_object* x_1363; +lean_object* x_1333; lean_object* x_1334; lean_object* x_1335; lean_object* x_1336; lean_object* x_1337; lean_object* x_1338; lean_object* x_1339; lean_object* x_1340; lean_object* x_1341; lean_object* x_1342; lean_object* x_1343; lean_object* x_1344; lean_object* x_1345; lean_object* x_1346; lean_object* x_1347; lean_object* x_1348; lean_object* x_1349; lean_object* x_1350; lean_object* x_1351; +lean_dec(x_3); lean_dec(x_1); -if (lean_is_scalar(x_1246)) { - x_1345 = lean_alloc_ctor(1, 2, 8); -} else { - x_1345 = x_1246; -} -lean_ctor_set(x_1345, 0, x_654); -lean_ctor_set(x_1345, 1, x_1247); -lean_ctor_set_uint64(x_1345, sizeof(void*)*2, x_1245); -if (lean_is_scalar(x_1243)) { - x_1346 = lean_alloc_ctor(1, 2, 8); -} else { - x_1346 = x_1243; -} -lean_ctor_set(x_1346, 0, x_1345); -lean_ctor_set(x_1346, 1, x_1265); -lean_ctor_set_uint64(x_1346, sizeof(void*)*2, x_1242); -if (lean_is_scalar(x_1240)) { - x_1347 = lean_alloc_ctor(1, 2, 8); -} else { - x_1347 = x_1240; -} -lean_ctor_set(x_1347, 0, x_1346); -lean_ctor_set(x_1347, 1, x_1291); -lean_ctor_set_uint64(x_1347, sizeof(void*)*2, x_1239); -if (lean_is_scalar(x_1264)) { - x_1348 = lean_alloc_ctor(1, 2, 8); -} else { - x_1348 = x_1264; -} -lean_ctor_set(x_1348, 0, x_1347); -lean_ctor_set(x_1348, 1, x_1317); -lean_ctor_set_uint64(x_1348, sizeof(void*)*2, x_1237); -x_1349 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1349, 0, x_1348); -lean_ctor_set(x_1349, 1, x_1235); -if (lean_is_scalar(x_1263)) { - x_1350 = lean_alloc_ctor(1, 2, 0); -} else { - x_1350 = x_1263; -} -lean_ctor_set(x_1350, 0, x_1349); -lean_ctor_set(x_1350, 1, x_656); -x_1351 = lean_box(0); -x_1352 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_1350, x_1351); -x_1353 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1352, x_1351); -x_1354 = l_Lean_MessageData_ofList(x_1353); -lean_dec(x_1353); -x_1355 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_1356 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1356, 0, x_1355); -lean_ctor_set(x_1356, 1, x_1354); -x_1357 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_1358 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1358, 0, x_1356); -lean_ctor_set(x_1358, 1, x_1357); -x_1359 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1358, x_5, x_6, x_655); -lean_dec(x_6); -x_1360 = lean_ctor_get(x_1359, 0); -lean_inc(x_1360); -x_1361 = lean_ctor_get(x_1359, 1); -lean_inc(x_1361); -if (lean_is_exclusive(x_1359)) { - lean_ctor_release(x_1359, 0); - lean_ctor_release(x_1359, 1); - x_1362 = x_1359; +if (lean_is_scalar(x_1235)) { + x_1333 = lean_alloc_ctor(1, 2, 8); } else { - lean_dec_ref(x_1359); - x_1362 = lean_box(0); + x_1333 = x_1235; } -if (lean_is_scalar(x_1362)) { - x_1363 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1333, 0, x_649); +lean_ctor_set(x_1333, 1, x_1236); +lean_ctor_set_uint64(x_1333, sizeof(void*)*2, x_1234); +if (lean_is_scalar(x_1232)) { + x_1334 = lean_alloc_ctor(1, 2, 8); } else { - x_1363 = x_1362; + x_1334 = x_1232; } -lean_ctor_set(x_1363, 0, x_1360); -lean_ctor_set(x_1363, 1, x_1361); -return x_1363; +lean_ctor_set(x_1334, 0, x_1333); +lean_ctor_set(x_1334, 1, x_1254); +lean_ctor_set_uint64(x_1334, sizeof(void*)*2, x_1231); +if (lean_is_scalar(x_1229)) { + x_1335 = lean_alloc_ctor(1, 2, 8); +} else { + x_1335 = x_1229; +} +lean_ctor_set(x_1335, 0, x_1334); +lean_ctor_set(x_1335, 1, x_1280); +lean_ctor_set_uint64(x_1335, sizeof(void*)*2, x_1228); +if (lean_is_scalar(x_1253)) { + x_1336 = lean_alloc_ctor(1, 2, 8); +} else { + x_1336 = x_1253; +} +lean_ctor_set(x_1336, 0, x_1335); +lean_ctor_set(x_1336, 1, x_1306); +lean_ctor_set_uint64(x_1336, sizeof(void*)*2, x_1226); +x_1337 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1337, 0, x_1336); +lean_ctor_set(x_1337, 1, x_1224); +if (lean_is_scalar(x_1252)) { + x_1338 = lean_alloc_ctor(1, 2, 0); +} else { + x_1338 = x_1252; +} +lean_ctor_set(x_1338, 0, x_1337); +lean_ctor_set(x_1338, 1, x_651); +x_1339 = lean_box(0); +x_1340 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_1338, x_1339); +x_1341 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1340, x_1339); +x_1342 = l_Lean_MessageData_ofList(x_1341); +lean_dec(x_1341); +x_1343 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_1344 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1344, 0, x_1343); +lean_ctor_set(x_1344, 1, x_1342); +x_1345 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_1346 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1346, 0, x_1344); +lean_ctor_set(x_1346, 1, x_1345); +x_1347 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1346, x_6, x_7, x_650); +lean_dec(x_7); +x_1348 = lean_ctor_get(x_1347, 0); +lean_inc(x_1348); +x_1349 = lean_ctor_get(x_1347, 1); +lean_inc(x_1349); +if (lean_is_exclusive(x_1347)) { + lean_ctor_release(x_1347, 0); + lean_ctor_release(x_1347, 1); + x_1350 = x_1347; +} else { + lean_dec_ref(x_1347); + x_1350 = lean_box(0); +} +if (lean_is_scalar(x_1350)) { + x_1351 = lean_alloc_ctor(1, 2, 0); +} else { + x_1351 = x_1350; +} +lean_ctor_set(x_1351, 0, x_1348); +lean_ctor_set(x_1351, 1, x_1349); +return x_1351; } } } @@ -8027,135 +8099,137 @@ return x_1363; } else { -lean_object* x_1364; -lean_dec(x_654); +lean_object* x_1352; +lean_dec(x_649); +lean_dec(x_83); lean_dec(x_82); lean_dec(x_81); -lean_dec(x_80); -lean_dec(x_78); -x_1364 = lean_ctor_get(x_9, 1); -lean_inc(x_1364); -if (lean_obj_tag(x_1364) == 0) +lean_dec(x_79); +lean_dec(x_3); +x_1352 = lean_ctor_get(x_10, 1); +lean_inc(x_1352); +if (lean_obj_tag(x_1352) == 0) { -lean_object* x_1365; lean_object* x_1366; +lean_object* x_1353; lean_object* x_1354; +lean_dec(x_10); +x_1353 = lean_ctor_get(x_9, 1); +lean_inc(x_1353); lean_dec(x_9); -x_1365 = lean_ctor_get(x_8, 1); -lean_inc(x_1365); -lean_dec(x_8); -x_1366 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_79, x_5, x_6, x_1365); +x_1354 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_80, x_6, x_7, x_1353); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_79); -return x_1366; +lean_dec(x_80); +return x_1354; } else { -lean_object* x_1367; lean_object* x_1368; lean_object* x_1369; lean_object* x_1370; lean_object* x_1371; lean_object* x_1372; lean_object* x_1373; lean_object* x_1374; lean_object* x_1375; lean_object* x_1376; uint8_t x_1377; -lean_dec(x_1364); -lean_dec(x_79); +lean_object* x_1355; lean_object* x_1356; lean_object* x_1357; lean_object* x_1358; lean_object* x_1359; lean_object* x_1360; lean_object* x_1361; lean_object* x_1362; lean_object* x_1363; lean_object* x_1364; uint8_t x_1365; +lean_dec(x_1352); +lean_dec(x_80); lean_dec(x_1); -x_1367 = lean_ctor_get(x_8, 1); -lean_inc(x_1367); -lean_dec(x_8); -x_1368 = lean_box(0); -x_1369 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_9, x_1368); -x_1370 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1369, x_1368); -x_1371 = l_Lean_MessageData_ofList(x_1370); -lean_dec(x_1370); -x_1372 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_1373 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1373, 0, x_1372); -lean_ctor_set(x_1373, 1, x_1371); -x_1374 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_1375 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1375, 0, x_1373); -lean_ctor_set(x_1375, 1, x_1374); -x_1376 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1375, x_5, x_6, x_1367); -lean_dec(x_6); -x_1377 = !lean_is_exclusive(x_1376); -if (x_1377 == 0) +x_1355 = lean_ctor_get(x_9, 1); +lean_inc(x_1355); +lean_dec(x_9); +x_1356 = lean_box(0); +x_1357 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_10, x_1356); +x_1358 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1357, x_1356); +x_1359 = l_Lean_MessageData_ofList(x_1358); +lean_dec(x_1358); +x_1360 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_1361 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1361, 0, x_1360); +lean_ctor_set(x_1361, 1, x_1359); +x_1362 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_1363 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1363, 0, x_1361); +lean_ctor_set(x_1363, 1, x_1362); +x_1364 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1363, x_6, x_7, x_1355); +lean_dec(x_7); +x_1365 = !lean_is_exclusive(x_1364); +if (x_1365 == 0) { -return x_1376; +return x_1364; } else { -lean_object* x_1378; lean_object* x_1379; lean_object* x_1380; -x_1378 = lean_ctor_get(x_1376, 0); -x_1379 = lean_ctor_get(x_1376, 1); -lean_inc(x_1379); -lean_inc(x_1378); -lean_dec(x_1376); -x_1380 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1380, 0, x_1378); -lean_ctor_set(x_1380, 1, x_1379); -return x_1380; +lean_object* x_1366; lean_object* x_1367; lean_object* x_1368; +x_1366 = lean_ctor_get(x_1364, 0); +x_1367 = lean_ctor_get(x_1364, 1); +lean_inc(x_1367); +lean_inc(x_1366); +lean_dec(x_1364); +x_1368 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1368, 0, x_1366); +lean_ctor_set(x_1368, 1, x_1367); +return x_1368; } } } } default: { -lean_object* x_1381; +lean_object* x_1369; +lean_dec(x_83); lean_dec(x_82); lean_dec(x_81); -lean_dec(x_80); -lean_dec(x_78); -x_1381 = lean_ctor_get(x_9, 1); -lean_inc(x_1381); -if (lean_obj_tag(x_1381) == 0) +lean_dec(x_79); +lean_dec(x_3); +x_1369 = lean_ctor_get(x_10, 1); +lean_inc(x_1369); +if (lean_obj_tag(x_1369) == 0) { -lean_object* x_1382; lean_object* x_1383; +lean_object* x_1370; lean_object* x_1371; +lean_dec(x_10); +x_1370 = lean_ctor_get(x_9, 1); +lean_inc(x_1370); lean_dec(x_9); -x_1382 = lean_ctor_get(x_8, 1); -lean_inc(x_1382); -lean_dec(x_8); -x_1383 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_79, x_5, x_6, x_1382); +x_1371 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_80, x_6, x_7, x_1370); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_79); -return x_1383; +lean_dec(x_80); +return x_1371; } else { -lean_object* x_1384; lean_object* x_1385; lean_object* x_1386; lean_object* x_1387; lean_object* x_1388; lean_object* x_1389; lean_object* x_1390; lean_object* x_1391; lean_object* x_1392; lean_object* x_1393; uint8_t x_1394; -lean_dec(x_1381); -lean_dec(x_79); +lean_object* x_1372; lean_object* x_1373; lean_object* x_1374; lean_object* x_1375; lean_object* x_1376; lean_object* x_1377; lean_object* x_1378; lean_object* x_1379; lean_object* x_1380; lean_object* x_1381; uint8_t x_1382; +lean_dec(x_1369); +lean_dec(x_80); lean_dec(x_1); -x_1384 = lean_ctor_get(x_8, 1); -lean_inc(x_1384); -lean_dec(x_8); -x_1385 = lean_box(0); -x_1386 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_9, x_1385); -x_1387 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1386, x_1385); -x_1388 = l_Lean_MessageData_ofList(x_1387); -lean_dec(x_1387); -x_1389 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_1390 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1390, 0, x_1389); -lean_ctor_set(x_1390, 1, x_1388); -x_1391 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_1392 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1392, 0, x_1390); -lean_ctor_set(x_1392, 1, x_1391); -x_1393 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1392, x_5, x_6, x_1384); -lean_dec(x_6); -x_1394 = !lean_is_exclusive(x_1393); -if (x_1394 == 0) +x_1372 = lean_ctor_get(x_9, 1); +lean_inc(x_1372); +lean_dec(x_9); +x_1373 = lean_box(0); +x_1374 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_10, x_1373); +x_1375 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1374, x_1373); +x_1376 = l_Lean_MessageData_ofList(x_1375); +lean_dec(x_1375); +x_1377 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_1378 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1378, 0, x_1377); +lean_ctor_set(x_1378, 1, x_1376); +x_1379 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_1380 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1380, 0, x_1378); +lean_ctor_set(x_1380, 1, x_1379); +x_1381 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1380, x_6, x_7, x_1372); +lean_dec(x_7); +x_1382 = !lean_is_exclusive(x_1381); +if (x_1382 == 0) { -return x_1393; +return x_1381; } else { -lean_object* x_1395; lean_object* x_1396; lean_object* x_1397; -x_1395 = lean_ctor_get(x_1393, 0); -x_1396 = lean_ctor_get(x_1393, 1); -lean_inc(x_1396); -lean_inc(x_1395); -lean_dec(x_1393); -x_1397 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1397, 0, x_1395); -lean_ctor_set(x_1397, 1, x_1396); -return x_1397; +lean_object* x_1383; lean_object* x_1384; lean_object* x_1385; +x_1383 = lean_ctor_get(x_1381, 0); +x_1384 = lean_ctor_get(x_1381, 1); +lean_inc(x_1384); +lean_inc(x_1383); +lean_dec(x_1381); +x_1385 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1385, 0, x_1383); +lean_ctor_set(x_1385, 1, x_1384); +return x_1385; } } } @@ -8163,195 +8237,198 @@ return x_1397; } else { -lean_object* x_1398; +lean_object* x_1386; +lean_dec(x_82); lean_dec(x_81); -lean_dec(x_80); -lean_dec(x_78); -x_1398 = lean_ctor_get(x_9, 1); -lean_inc(x_1398); -if (lean_obj_tag(x_1398) == 0) +lean_dec(x_79); +lean_dec(x_3); +x_1386 = lean_ctor_get(x_10, 1); +lean_inc(x_1386); +if (lean_obj_tag(x_1386) == 0) { -lean_object* x_1399; lean_object* x_1400; +lean_object* x_1387; lean_object* x_1388; +lean_dec(x_10); +x_1387 = lean_ctor_get(x_9, 1); +lean_inc(x_1387); lean_dec(x_9); -x_1399 = lean_ctor_get(x_8, 1); -lean_inc(x_1399); -lean_dec(x_8); -x_1400 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_79, x_5, x_6, x_1399); +x_1388 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_80, x_6, x_7, x_1387); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_79); -return x_1400; +lean_dec(x_80); +return x_1388; } else { -lean_object* x_1401; lean_object* x_1402; lean_object* x_1403; lean_object* x_1404; lean_object* x_1405; lean_object* x_1406; lean_object* x_1407; lean_object* x_1408; lean_object* x_1409; lean_object* x_1410; uint8_t x_1411; -lean_dec(x_1398); -lean_dec(x_79); +lean_object* x_1389; lean_object* x_1390; lean_object* x_1391; lean_object* x_1392; lean_object* x_1393; lean_object* x_1394; lean_object* x_1395; lean_object* x_1396; lean_object* x_1397; lean_object* x_1398; uint8_t x_1399; +lean_dec(x_1386); +lean_dec(x_80); lean_dec(x_1); -x_1401 = lean_ctor_get(x_8, 1); -lean_inc(x_1401); -lean_dec(x_8); -x_1402 = lean_box(0); -x_1403 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_9, x_1402); -x_1404 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1403, x_1402); -x_1405 = l_Lean_MessageData_ofList(x_1404); -lean_dec(x_1404); -x_1406 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_1407 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1407, 0, x_1406); -lean_ctor_set(x_1407, 1, x_1405); -x_1408 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_1409 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1409, 0, x_1407); -lean_ctor_set(x_1409, 1, x_1408); -x_1410 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1409, x_5, x_6, x_1401); -lean_dec(x_6); -x_1411 = !lean_is_exclusive(x_1410); -if (x_1411 == 0) +x_1389 = lean_ctor_get(x_9, 1); +lean_inc(x_1389); +lean_dec(x_9); +x_1390 = lean_box(0); +x_1391 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_10, x_1390); +x_1392 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1391, x_1390); +x_1393 = l_Lean_MessageData_ofList(x_1392); +lean_dec(x_1392); +x_1394 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_1395 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1395, 0, x_1394); +lean_ctor_set(x_1395, 1, x_1393); +x_1396 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_1397 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1397, 0, x_1395); +lean_ctor_set(x_1397, 1, x_1396); +x_1398 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1397, x_6, x_7, x_1389); +lean_dec(x_7); +x_1399 = !lean_is_exclusive(x_1398); +if (x_1399 == 0) { -return x_1410; +return x_1398; } else { -lean_object* x_1412; lean_object* x_1413; lean_object* x_1414; -x_1412 = lean_ctor_get(x_1410, 0); -x_1413 = lean_ctor_get(x_1410, 1); -lean_inc(x_1413); -lean_inc(x_1412); -lean_dec(x_1410); -x_1414 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1414, 0, x_1412); -lean_ctor_set(x_1414, 1, x_1413); -return x_1414; +lean_object* x_1400; lean_object* x_1401; lean_object* x_1402; +x_1400 = lean_ctor_get(x_1398, 0); +x_1401 = lean_ctor_get(x_1398, 1); +lean_inc(x_1401); +lean_inc(x_1400); +lean_dec(x_1398); +x_1402 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1402, 0, x_1400); +lean_ctor_set(x_1402, 1, x_1401); +return x_1402; } } } } else { -lean_object* x_1415; -lean_dec(x_80); -lean_dec(x_78); -x_1415 = lean_ctor_get(x_9, 1); -lean_inc(x_1415); -if (lean_obj_tag(x_1415) == 0) +lean_object* x_1403; +lean_dec(x_81); +lean_dec(x_79); +lean_dec(x_3); +x_1403 = lean_ctor_get(x_10, 1); +lean_inc(x_1403); +if (lean_obj_tag(x_1403) == 0) { -lean_object* x_1416; lean_object* x_1417; +lean_object* x_1404; lean_object* x_1405; +lean_dec(x_10); +x_1404 = lean_ctor_get(x_9, 1); +lean_inc(x_1404); lean_dec(x_9); -x_1416 = lean_ctor_get(x_8, 1); -lean_inc(x_1416); -lean_dec(x_8); -x_1417 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_79, x_5, x_6, x_1416); +x_1405 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_80, x_6, x_7, x_1404); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_79); -return x_1417; +lean_dec(x_80); +return x_1405; } else { -lean_object* x_1418; lean_object* x_1419; lean_object* x_1420; lean_object* x_1421; lean_object* x_1422; lean_object* x_1423; lean_object* x_1424; lean_object* x_1425; lean_object* x_1426; lean_object* x_1427; uint8_t x_1428; -lean_dec(x_1415); -lean_dec(x_79); +lean_object* x_1406; lean_object* x_1407; lean_object* x_1408; lean_object* x_1409; lean_object* x_1410; lean_object* x_1411; lean_object* x_1412; lean_object* x_1413; lean_object* x_1414; lean_object* x_1415; uint8_t x_1416; +lean_dec(x_1403); +lean_dec(x_80); lean_dec(x_1); -x_1418 = lean_ctor_get(x_8, 1); -lean_inc(x_1418); -lean_dec(x_8); -x_1419 = lean_box(0); -x_1420 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_9, x_1419); -x_1421 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1420, x_1419); -x_1422 = l_Lean_MessageData_ofList(x_1421); -lean_dec(x_1421); -x_1423 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_1424 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1424, 0, x_1423); -lean_ctor_set(x_1424, 1, x_1422); -x_1425 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_1426 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1426, 0, x_1424); -lean_ctor_set(x_1426, 1, x_1425); -x_1427 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1426, x_5, x_6, x_1418); -lean_dec(x_6); -x_1428 = !lean_is_exclusive(x_1427); -if (x_1428 == 0) +x_1406 = lean_ctor_get(x_9, 1); +lean_inc(x_1406); +lean_dec(x_9); +x_1407 = lean_box(0); +x_1408 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_10, x_1407); +x_1409 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1408, x_1407); +x_1410 = l_Lean_MessageData_ofList(x_1409); +lean_dec(x_1409); +x_1411 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_1412 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1412, 0, x_1411); +lean_ctor_set(x_1412, 1, x_1410); +x_1413 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_1414 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1414, 0, x_1412); +lean_ctor_set(x_1414, 1, x_1413); +x_1415 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1414, x_6, x_7, x_1406); +lean_dec(x_7); +x_1416 = !lean_is_exclusive(x_1415); +if (x_1416 == 0) { -return x_1427; +return x_1415; } else { -lean_object* x_1429; lean_object* x_1430; lean_object* x_1431; -x_1429 = lean_ctor_get(x_1427, 0); -x_1430 = lean_ctor_get(x_1427, 1); -lean_inc(x_1430); -lean_inc(x_1429); -lean_dec(x_1427); -x_1431 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1431, 0, x_1429); -lean_ctor_set(x_1431, 1, x_1430); -return x_1431; +lean_object* x_1417; lean_object* x_1418; lean_object* x_1419; +x_1417 = lean_ctor_get(x_1415, 0); +x_1418 = lean_ctor_get(x_1415, 1); +lean_inc(x_1418); +lean_inc(x_1417); +lean_dec(x_1415); +x_1419 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1419, 0, x_1417); +lean_ctor_set(x_1419, 1, x_1418); +return x_1419; } } } } else { -lean_object* x_1432; -lean_dec(x_78); -x_1432 = lean_ctor_get(x_9, 1); -lean_inc(x_1432); -if (lean_obj_tag(x_1432) == 0) +lean_object* x_1420; +lean_dec(x_79); +lean_dec(x_3); +x_1420 = lean_ctor_get(x_10, 1); +lean_inc(x_1420); +if (lean_obj_tag(x_1420) == 0) { -lean_object* x_1433; lean_object* x_1434; +lean_object* x_1421; lean_object* x_1422; +lean_dec(x_10); +x_1421 = lean_ctor_get(x_9, 1); +lean_inc(x_1421); lean_dec(x_9); -x_1433 = lean_ctor_get(x_8, 1); -lean_inc(x_1433); -lean_dec(x_8); -x_1434 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_79, x_5, x_6, x_1433); +x_1422 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__1(x_1, x_80, x_6, x_7, x_1421); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_79); -return x_1434; +lean_dec(x_80); +return x_1422; } else { -lean_object* x_1435; lean_object* x_1436; lean_object* x_1437; lean_object* x_1438; lean_object* x_1439; lean_object* x_1440; lean_object* x_1441; lean_object* x_1442; lean_object* x_1443; lean_object* x_1444; uint8_t x_1445; -lean_dec(x_1432); -lean_dec(x_79); +lean_object* x_1423; lean_object* x_1424; lean_object* x_1425; lean_object* x_1426; lean_object* x_1427; lean_object* x_1428; lean_object* x_1429; lean_object* x_1430; lean_object* x_1431; lean_object* x_1432; uint8_t x_1433; +lean_dec(x_1420); +lean_dec(x_80); lean_dec(x_1); -x_1435 = lean_ctor_get(x_8, 1); -lean_inc(x_1435); -lean_dec(x_8); -x_1436 = lean_box(0); -x_1437 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_9, x_1436); -x_1438 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1437, x_1436); -x_1439 = l_Lean_MessageData_ofList(x_1438); -lean_dec(x_1438); -x_1440 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; -x_1441 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1441, 0, x_1440); -lean_ctor_set(x_1441, 1, x_1439); -x_1442 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; -x_1443 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_1443, 0, x_1441); -lean_ctor_set(x_1443, 1, x_1442); -x_1444 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1443, x_5, x_6, x_1435); -lean_dec(x_6); -x_1445 = !lean_is_exclusive(x_1444); -if (x_1445 == 0) +x_1423 = lean_ctor_get(x_9, 1); +lean_inc(x_1423); +lean_dec(x_9); +x_1424 = lean_box(0); +x_1425 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_10, x_1424); +x_1426 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_1425, x_1424); +x_1427 = l_Lean_MessageData_ofList(x_1426); +lean_dec(x_1426); +x_1428 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__4; +x_1429 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1429, 0, x_1428); +lean_ctor_set(x_1429, 1, x_1427); +x_1430 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__6; +x_1431 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_1431, 0, x_1429); +lean_ctor_set(x_1431, 1, x_1430); +x_1432 = l_Lean_throwError___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___spec__15(x_1431, x_6, x_7, x_1423); +lean_dec(x_7); +x_1433 = !lean_is_exclusive(x_1432); +if (x_1433 == 0) { -return x_1444; +return x_1432; } else { -lean_object* x_1446; lean_object* x_1447; lean_object* x_1448; -x_1446 = lean_ctor_get(x_1444, 0); -x_1447 = lean_ctor_get(x_1444, 1); -lean_inc(x_1447); -lean_inc(x_1446); -lean_dec(x_1444); -x_1448 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1448, 0, x_1446); -lean_ctor_set(x_1448, 1, x_1447); -return x_1448; +lean_object* x_1434; lean_object* x_1435; lean_object* x_1436; +x_1434 = lean_ctor_get(x_1432, 0); +x_1435 = lean_ctor_get(x_1432, 1); +lean_inc(x_1435); +lean_inc(x_1434); +lean_dec(x_1432); +x_1436 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1436, 0, x_1434); +lean_ctor_set(x_1436, 1, x_1435); +return x_1436; } } } @@ -8522,107 +8599,105 @@ lean_dec(x_1); x_27 = l_Lean_Syntax_isNone(x_26); if (x_27 == 0) { -lean_object* x_28; uint8_t x_29; -x_28 = l_Lean_nullKind; +uint8_t x_28; lean_inc(x_26); -x_29 = l_Lean_Syntax_isNodeOf(x_26, x_28, x_25); -if (x_29 == 0) +x_28 = l_Lean_Syntax_matchesNull(x_26, x_25); +if (x_28 == 0) { -lean_object* x_30; lean_object* x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; +lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_dec(x_26); lean_dec(x_17); lean_dec(x_4); lean_dec(x_3); +x_29 = lean_box(0); x_30 = lean_box(0); -x_31 = lean_box(0); -x_32 = 1; -x_33 = l_Lean_Syntax_mkAntiquotNode(x_31, x_2, x_16, x_30, x_32); -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_5); -return x_34; +x_31 = 1; +x_32 = l_Lean_Syntax_mkAntiquotNode(x_30, x_2, x_16, x_29, x_31); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_5); +return x_33; } else { -lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_35 = l_Lean_Syntax_getArg(x_26, x_16); +lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_34 = l_Lean_Syntax_getArg(x_26, x_16); lean_dec(x_26); -x_36 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___closed__11; -lean_inc(x_35); -x_37 = l_Lean_Syntax_isOfKind(x_35, x_36); -if (x_37 == 0) +x_35 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___closed__11; +lean_inc(x_34); +x_36 = l_Lean_Syntax_isOfKind(x_34, x_35); +if (x_36 == 0) { -lean_object* x_38; lean_object* x_39; uint8_t x_40; lean_object* x_41; lean_object* x_42; -lean_dec(x_35); +lean_object* x_37; lean_object* x_38; uint8_t x_39; lean_object* x_40; lean_object* x_41; +lean_dec(x_34); lean_dec(x_17); lean_dec(x_4); lean_dec(x_3); +x_37 = lean_box(0); x_38 = lean_box(0); -x_39 = lean_box(0); -x_40 = 1; -x_41 = l_Lean_Syntax_mkAntiquotNode(x_39, x_2, x_16, x_38, x_40); -x_42 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_5); -return x_42; +x_39 = 1; +x_40 = l_Lean_Syntax_mkAntiquotNode(x_38, x_2, x_16, x_37, x_39); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_5); +return x_41; } else { -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_43 = l_Lean_Syntax_getArg(x_35, x_25); -lean_dec(x_35); -x_44 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_44, 0, x_43); -x_45 = lean_box(0); -x_46 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2(x_2, x_17, x_45, x_44, x_3, x_4, x_5); -return x_46; +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_42 = l_Lean_Syntax_getArg(x_34, x_25); +lean_dec(x_34); +x_43 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_43, 0, x_42); +x_44 = lean_box(0); +x_45 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2(x_2, x_17, x_18, x_44, x_43, x_3, x_4, x_5); +return x_45; } } } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_dec(x_26); +x_46 = lean_box(0); x_47 = lean_box(0); -x_48 = lean_box(0); -x_49 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2(x_2, x_17, x_48, x_47, x_3, x_4, x_5); -return x_49; +x_48 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2(x_2, x_17, x_18, x_47, x_46, x_3, x_4, x_5); +return x_48; } } } } else { -lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; -x_50 = lean_unsigned_to_nat(1u); -x_51 = l_Lean_Syntax_getArg(x_1, x_50); +lean_object* x_49; lean_object* x_50; uint8_t x_51; +x_49 = lean_unsigned_to_nat(1u); +x_50 = l_Lean_Syntax_getArg(x_1, x_49); lean_dec(x_1); -x_52 = l_Lean_nullKind; -lean_inc(x_51); -x_53 = l_Lean_Syntax_isNodeOf(x_51, x_52, x_50); -if (x_53 == 0) +lean_inc(x_50); +x_51 = l_Lean_Syntax_matchesNull(x_50, x_49); +if (x_51 == 0) { -lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; lean_object* x_58; lean_object* x_59; -lean_dec(x_51); +lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; lean_object* x_56; lean_object* x_57; +lean_dec(x_50); lean_dec(x_4); lean_dec(x_3); -x_54 = lean_box(0); -x_55 = lean_box(0); -x_56 = lean_unsigned_to_nat(0u); -x_57 = 1; -x_58 = l_Lean_Syntax_mkAntiquotNode(x_55, x_2, x_56, x_54, x_57); -x_59 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_59, 0, x_58); -lean_ctor_set(x_59, 1, x_5); -return x_59; +x_52 = lean_box(0); +x_53 = lean_box(0); +x_54 = lean_unsigned_to_nat(0u); +x_55 = 1; +x_56 = l_Lean_Syntax_mkAntiquotNode(x_53, x_2, x_54, x_52, x_55); +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_5); +return x_57; } else { -lean_object* x_60; lean_object* x_61; -x_60 = lean_unsigned_to_nat(0u); -x_61 = l_Lean_Syntax_getArg(x_51, x_60); -lean_dec(x_51); -x_1 = x_61; +lean_object* x_58; lean_object* x_59; +x_58 = lean_unsigned_to_nat(0u); +x_59 = l_Lean_Syntax_getArg(x_50, x_58); +lean_dec(x_50); +x_1 = x_59; goto _start; } } @@ -8764,6 +8839,24 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } +static lean_object* _init_l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("null", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__2; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandMacroArg_mkSplicePat(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: { @@ -8781,7 +8874,7 @@ x_11 = l_Lean_Syntax_mkAntiquotSuffixSpliceNode(x_1, x_10, x_4); x_12 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; x_13 = lean_array_push(x_12, x_11); x_14 = lean_box(2); -x_15 = l_Lean_nullKind; +x_15 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__3; x_16 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_16, 0, x_14); lean_ctor_set(x_16, 1, x_15); @@ -8801,7 +8894,7 @@ x_19 = l_Lean_Syntax_mkAntiquotSuffixSpliceNode(x_1, x_17, x_4); x_20 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; x_21 = lean_array_push(x_20, x_19); x_22 = lean_box(2); -x_23 = l_Lean_nullKind; +x_23 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__3; x_24 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_24, 0, x_22); lean_ctor_set(x_24, 1, x_23); @@ -8931,7 +9024,7 @@ if (lean_obj_tag(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; x_10 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__2___closed__6; -x_11 = l_panic___at_Lean_Name_getString_x21___spec__1(x_10); +x_11 = l_panic___at_Lean_TSyntax_getString___spec__1(x_10); x_12 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__2___closed__7; x_13 = lean_string_append(x_11, x_12); x_14 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__2___closed__2; @@ -9079,29 +9172,11 @@ static lean_object* _init_l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___la _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("null", 4); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__6; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__8() { -_start: -{ -lean_object* x_1; x_1 = lean_mk_string_from_bytes(")", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__9() { +static lean_object* _init_l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7() { _start: { lean_object* x_1; lean_object* x_2; @@ -9120,201 +9195,200 @@ lean_dec(x_1); x_16 = l_Lean_Syntax_isNone(x_15); if (x_16 == 0) { -lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_17 = l_Lean_nullKind; -x_18 = lean_unsigned_to_nat(2u); -x_19 = l_Lean_Syntax_isNodeOf(x_15, x_17, x_18); -if (x_19 == 0) +lean_object* x_17; uint8_t x_18; +x_17 = lean_unsigned_to_nat(2u); +x_18 = l_Lean_Syntax_matchesNull(x_15, x_17); +if (x_18 == 0) { lean_dec(x_3); lean_dec(x_2); if (lean_obj_tag(x_6) == 0) { -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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_dec(x_5); -x_20 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_11, x_12, x_13); -x_21 = lean_ctor_get(x_20, 0); +x_19 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_11, x_12, x_13); +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_20, 1); -lean_inc(x_22); -lean_dec(x_20); -x_23 = l_Lean_Elab_Command_getCurrMacroScope(x_11, x_12, x_22); -x_24 = lean_ctor_get(x_23, 0); +lean_dec(x_19); +x_22 = l_Lean_Elab_Command_getCurrMacroScope(x_11, x_12, x_21); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); -lean_inc(x_25); -lean_dec(x_23); -x_26 = l_Lean_Elab_Command_getMainModule___rarg(x_12, x_25); -x_27 = lean_ctor_get(x_26, 0); +lean_dec(x_22); +x_25 = l_Lean_Elab_Command_getMainModule___rarg(x_12, x_24); +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_26, 1); -lean_inc(x_28); -lean_dec(x_26); -x_29 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; -x_30 = l_Lean_addMacroScope(x_27, x_29, x_24); -x_31 = lean_box(0); -x_32 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; -lean_inc(x_21); -x_33 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_33, 0, x_21); -lean_ctor_set(x_33, 1, x_32); -lean_ctor_set(x_33, 2, x_30); -lean_ctor_set(x_33, 3, x_31); -x_34 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; -lean_inc(x_21); -x_35 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_35, 0, x_21); -lean_ctor_set(x_35, 1, x_34); -x_36 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; +lean_dec(x_25); +x_28 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; +x_29 = l_Lean_addMacroScope(x_26, x_28, x_23); +x_30 = lean_box(0); +x_31 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; +lean_inc(x_20); +x_32 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_32, 0, x_20); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set(x_32, 2, x_29); +lean_ctor_set(x_32, 3, x_30); +x_33 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; +lean_inc(x_20); +x_34 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_34, 0, x_20); +lean_ctor_set(x_34, 1, x_33); +x_35 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; lean_inc(x_7); -x_37 = lean_array_push(x_36, x_7); -x_38 = lean_box(2); -x_39 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; -x_40 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -lean_ctor_set(x_40, 2, x_37); -x_41 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__8; -x_42 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_42, 0, x_21); -lean_ctor_set(x_42, 1, x_41); -x_43 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__9; -x_44 = lean_array_push(x_43, x_33); -x_45 = lean_array_push(x_44, x_35); -x_46 = lean_array_push(x_45, x_40); -x_47 = lean_array_push(x_46, x_42); -x_48 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_48, 0, x_38); -lean_ctor_set(x_48, 1, x_8); -lean_ctor_set(x_48, 2, x_47); -x_49 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_7, x_4, x_11, x_12, x_28); -if (lean_obj_tag(x_49) == 0) -{ -uint8_t x_50; -x_50 = !lean_is_exclusive(x_49); -if (x_50 == 0) +x_36 = lean_array_push(x_35, x_7); +x_37 = lean_box(2); +x_38 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__3; +x_39 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +lean_ctor_set(x_39, 2, x_36); +x_40 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__6; +x_41 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_41, 0, x_20); +lean_ctor_set(x_41, 1, x_40); +x_42 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; +x_43 = lean_array_push(x_42, x_32); +x_44 = lean_array_push(x_43, x_34); +x_45 = lean_array_push(x_44, x_39); +x_46 = lean_array_push(x_45, x_41); +x_47 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_47, 0, x_37); +lean_ctor_set(x_47, 1, x_8); +lean_ctor_set(x_47, 2, x_46); +x_48 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_7, x_4, x_11, x_12, x_27); +if (lean_obj_tag(x_48) == 0) +{ +uint8_t x_49; +x_49 = !lean_is_exclusive(x_48); +if (x_49 == 0) { -lean_object* x_51; lean_object* x_52; -x_51 = lean_ctor_get(x_49, 0); -x_52 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_52, 0, x_48); -lean_ctor_set(x_52, 1, x_51); -lean_ctor_set(x_49, 0, x_52); -return x_49; +lean_object* x_50; lean_object* x_51; +x_50 = lean_ctor_get(x_48, 0); +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_47); +lean_ctor_set(x_51, 1, x_50); +lean_ctor_set(x_48, 0, x_51); +return x_48; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_53 = lean_ctor_get(x_49, 0); -x_54 = lean_ctor_get(x_49, 1); -lean_inc(x_54); +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_52 = lean_ctor_get(x_48, 0); +x_53 = lean_ctor_get(x_48, 1); lean_inc(x_53); -lean_dec(x_49); +lean_inc(x_52); +lean_dec(x_48); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_47); +lean_ctor_set(x_54, 1, x_52); x_55 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_55, 0, x_48); +lean_ctor_set(x_55, 0, x_54); lean_ctor_set(x_55, 1, x_53); -x_56 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_56, 0, x_55); -lean_ctor_set(x_56, 1, x_54); -return x_56; +return x_55; } } else { -uint8_t x_57; -lean_dec(x_48); -x_57 = !lean_is_exclusive(x_49); -if (x_57 == 0) +uint8_t x_56; +lean_dec(x_47); +x_56 = !lean_is_exclusive(x_48); +if (x_56 == 0) { -return x_49; +return x_48; } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_58 = lean_ctor_get(x_49, 0); -x_59 = lean_ctor_get(x_49, 1); -lean_inc(x_59); +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_ctor_get(x_48, 0); +x_58 = lean_ctor_get(x_48, 1); lean_inc(x_58); -lean_dec(x_49); -x_60 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_60, 0, x_58); -lean_ctor_set(x_60, 1, x_59); -return x_60; +lean_inc(x_57); +lean_dec(x_48); +x_59 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_58); +return x_59; } } } else { -lean_object* x_61; lean_object* x_62; +lean_object* x_60; lean_object* x_61; lean_dec(x_8); lean_dec(x_4); -x_61 = lean_ctor_get(x_6, 0); -lean_inc(x_61); +x_60 = lean_ctor_get(x_6, 0); +lean_inc(x_60); lean_dec(x_6); lean_inc(x_12); lean_inc(x_11); -x_62 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_7, x_61, x_11, x_12, x_13); -if (lean_obj_tag(x_62) == 0) +x_61 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_7, x_60, x_11, x_12, x_13); +if (lean_obj_tag(x_61) == 0) { -lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_63 = lean_ctor_get(x_62, 0); +lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_62 = lean_ctor_get(x_61, 0); +lean_inc(x_62); +x_63 = lean_ctor_get(x_61, 1); lean_inc(x_63); -x_64 = lean_ctor_get(x_62, 1); -lean_inc(x_64); -lean_dec(x_62); -x_65 = lean_apply_4(x_5, x_63, x_11, x_12, x_64); -return x_65; +lean_dec(x_61); +x_64 = lean_apply_4(x_5, x_62, x_11, x_12, x_63); +return x_64; } else { -uint8_t x_66; +uint8_t x_65; lean_dec(x_12); lean_dec(x_11); lean_dec(x_5); -x_66 = !lean_is_exclusive(x_62); -if (x_66 == 0) +x_65 = !lean_is_exclusive(x_61); +if (x_65 == 0) { -return x_62; +return x_61; } else { -lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_67 = lean_ctor_get(x_62, 0); -x_68 = lean_ctor_get(x_62, 1); -lean_inc(x_68); +lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_66 = lean_ctor_get(x_61, 0); +x_67 = lean_ctor_get(x_61, 1); lean_inc(x_67); -lean_dec(x_62); -x_69 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_69, 0, x_67); -lean_ctor_set(x_69, 1, x_68); -return x_69; +lean_inc(x_66); +lean_dec(x_61); +x_68 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_68, 0, x_66); +lean_ctor_set(x_68, 1, x_67); +return x_68; } } } } else { -lean_object* x_70; lean_object* x_71; +lean_object* x_69; lean_object* x_70; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -x_70 = lean_box(0); -x_71 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__2(x_2, x_3, x_4, x_5, x_70, x_11, x_12, x_13); +x_69 = lean_box(0); +x_70 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__2(x_2, x_3, x_4, x_5, x_69, x_11, x_12, x_13); lean_dec(x_2); -return x_71; +return x_70; } } else { -lean_object* x_72; lean_object* x_73; +lean_object* x_71; lean_object* x_72; lean_dec(x_15); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -x_72 = lean_box(0); -x_73 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__2(x_2, x_3, x_4, x_5, x_72, x_11, x_12, x_13); +x_71 = lean_box(0); +x_72 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__2(x_2, x_3, x_4, x_5, x_71, x_11, x_12, x_13); lean_dec(x_2); -return x_73; +return x_72; } } } @@ -9404,7 +9478,7 @@ static lean_object* _init_l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___cl _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("str", 3); +x_1 = lean_mk_string_from_bytes("optional", 8); return x_1; } } @@ -9422,7 +9496,7 @@ static lean_object* _init_l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___cl _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("optional", 8); +x_1 = lean_mk_string_from_bytes("many", 4); return x_1; } } @@ -9440,7 +9514,7 @@ static lean_object* _init_l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___cl _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("many", 4); +x_1 = lean_mk_string_from_bytes("many1", 5); return x_1; } } @@ -9458,7 +9532,7 @@ static lean_object* _init_l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___cl _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("many1", 5); +x_1 = lean_mk_string_from_bytes("interpolatedStr", 15); return x_1; } } @@ -9476,7 +9550,7 @@ static lean_object* _init_l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___cl _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("interpolatedStr", 15); +x_1 = lean_mk_string_from_bytes("withPosition", 12); return x_1; } } @@ -9494,7 +9568,7 @@ static lean_object* _init_l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___cl _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("withPosition", 12); +x_1 = lean_mk_string_from_bytes("term", 4); return x_1; } } @@ -9512,7 +9586,7 @@ static lean_object* _init_l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___cl _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("term", 4); +x_1 = lean_mk_string_from_bytes("interpolatedStrKind", 19); return x_1; } } @@ -9669,16 +9743,16 @@ x_34 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; lean_inc(x_3); x_35 = lean_array_push(x_34, x_3); x_36 = lean_box(2); -x_37 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; +x_37 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__3; x_38 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_38, 0, x_36); lean_ctor_set(x_38, 1, x_37); lean_ctor_set(x_38, 2, x_35); -x_39 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__8; +x_39 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__6; x_40 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_40, 0, x_19); lean_ctor_set(x_40, 1, x_39); -x_41 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__9; +x_41 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; x_42 = lean_array_push(x_41, x_31); x_43 = lean_array_push(x_42, x_33); x_44 = lean_array_push(x_43, x_38); @@ -9796,747 +9870,746 @@ return x_67; } else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; +lean_object* x_68; lean_object* x_69; uint8_t x_70; x_68 = lean_unsigned_to_nat(1u); x_69 = l_Lean_Syntax_getArg(x_3, x_68); -x_70 = l_Lean_nullKind; lean_inc(x_69); -x_71 = l_Lean_Syntax_isNodeOf(x_69, x_70, x_68); -if (x_71 == 0) +x_70 = l_Lean_Syntax_matchesNull(x_69, x_68); +if (x_70 == 0) { lean_dec(x_69); lean_dec(x_7); if (lean_obj_tag(x_1) == 0) { -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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; -x_72 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); -x_73 = lean_ctor_get(x_72, 0); +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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; +x_71 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); +x_72 = lean_ctor_get(x_71, 0); +lean_inc(x_72); +x_73 = lean_ctor_get(x_71, 1); lean_inc(x_73); -x_74 = lean_ctor_get(x_72, 1); -lean_inc(x_74); -lean_dec(x_72); -x_75 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_74); -x_76 = lean_ctor_get(x_75, 0); +lean_dec(x_71); +x_74 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_73); +x_75 = lean_ctor_get(x_74, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_74, 1); lean_inc(x_76); -x_77 = lean_ctor_get(x_75, 1); -lean_inc(x_77); -lean_dec(x_75); -x_78 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_77); -x_79 = lean_ctor_get(x_78, 0); +lean_dec(x_74); +x_77 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_76); +x_78 = lean_ctor_get(x_77, 0); +lean_inc(x_78); +x_79 = lean_ctor_get(x_77, 1); lean_inc(x_79); -x_80 = lean_ctor_get(x_78, 1); -lean_inc(x_80); -lean_dec(x_78); -x_81 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; -x_82 = l_Lean_addMacroScope(x_79, x_81, x_76); -x_83 = lean_box(0); -x_84 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; -lean_inc(x_73); -x_85 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_85, 0, x_73); -lean_ctor_set(x_85, 1, x_84); -lean_ctor_set(x_85, 2, x_82); -lean_ctor_set(x_85, 3, x_83); -x_86 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; -lean_inc(x_73); -x_87 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_87, 0, x_73); -lean_ctor_set(x_87, 1, x_86); -x_88 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; +lean_dec(x_77); +x_80 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; +x_81 = l_Lean_addMacroScope(x_78, x_80, x_75); +x_82 = lean_box(0); +x_83 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; +lean_inc(x_72); +x_84 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_84, 0, x_72); +lean_ctor_set(x_84, 1, x_83); +lean_ctor_set(x_84, 2, x_81); +lean_ctor_set(x_84, 3, x_82); +x_85 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; +lean_inc(x_72); +x_86 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_86, 0, x_72); +lean_ctor_set(x_86, 1, x_85); +x_87 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; lean_inc(x_3); -x_89 = lean_array_push(x_88, x_3); -x_90 = lean_box(2); -x_91 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; -x_92 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_92, 0, x_90); -lean_ctor_set(x_92, 1, x_91); -lean_ctor_set(x_92, 2, x_89); -x_93 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__8; -x_94 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_94, 0, x_73); -lean_ctor_set(x_94, 1, x_93); -x_95 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__9; -x_96 = lean_array_push(x_95, x_85); -x_97 = lean_array_push(x_96, x_87); -x_98 = lean_array_push(x_97, x_92); -x_99 = lean_array_push(x_98, x_94); -x_100 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_100, 0, x_90); -lean_ctor_set(x_100, 1, x_12); -lean_ctor_set(x_100, 2, x_99); -x_101 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_80); -if (lean_obj_tag(x_101) == 0) -{ -uint8_t x_102; -x_102 = !lean_is_exclusive(x_101); -if (x_102 == 0) -{ -lean_object* x_103; lean_object* x_104; -x_103 = lean_ctor_get(x_101, 0); -x_104 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_104, 0, x_100); -lean_ctor_set(x_104, 1, x_103); -lean_ctor_set(x_101, 0, x_104); -return x_101; -} -else -{ -lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; -x_105 = lean_ctor_get(x_101, 0); -x_106 = lean_ctor_get(x_101, 1); -lean_inc(x_106); +x_88 = lean_array_push(x_87, x_3); +x_89 = lean_box(2); +x_90 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__3; +x_91 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_91, 0, x_89); +lean_ctor_set(x_91, 1, x_90); +lean_ctor_set(x_91, 2, x_88); +x_92 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__6; +x_93 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_93, 0, x_72); +lean_ctor_set(x_93, 1, x_92); +x_94 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; +x_95 = lean_array_push(x_94, x_84); +x_96 = lean_array_push(x_95, x_86); +x_97 = lean_array_push(x_96, x_91); +x_98 = lean_array_push(x_97, x_93); +x_99 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_99, 0, x_89); +lean_ctor_set(x_99, 1, x_12); +lean_ctor_set(x_99, 2, x_98); +x_100 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_79); +if (lean_obj_tag(x_100) == 0) +{ +uint8_t x_101; +x_101 = !lean_is_exclusive(x_100); +if (x_101 == 0) +{ +lean_object* x_102; lean_object* x_103; +x_102 = lean_ctor_get(x_100, 0); +x_103 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_103, 0, x_99); +lean_ctor_set(x_103, 1, x_102); +lean_ctor_set(x_100, 0, x_103); +return x_100; +} +else +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_104 = lean_ctor_get(x_100, 0); +x_105 = lean_ctor_get(x_100, 1); lean_inc(x_105); -lean_dec(x_101); +lean_inc(x_104); +lean_dec(x_100); +x_106 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_106, 0, x_99); +lean_ctor_set(x_106, 1, x_104); x_107 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_107, 0, x_100); +lean_ctor_set(x_107, 0, x_106); lean_ctor_set(x_107, 1, x_105); -x_108 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_108, 0, x_107); -lean_ctor_set(x_108, 1, x_106); -return x_108; +return x_107; } } else { -uint8_t x_109; -lean_dec(x_100); -x_109 = !lean_is_exclusive(x_101); -if (x_109 == 0) +uint8_t x_108; +lean_dec(x_99); +x_108 = !lean_is_exclusive(x_100); +if (x_108 == 0) { -return x_101; +return x_100; } else { -lean_object* x_110; lean_object* x_111; lean_object* x_112; -x_110 = lean_ctor_get(x_101, 0); -x_111 = lean_ctor_get(x_101, 1); -lean_inc(x_111); +lean_object* x_109; lean_object* x_110; lean_object* x_111; +x_109 = lean_ctor_get(x_100, 0); +x_110 = lean_ctor_get(x_100, 1); lean_inc(x_110); -lean_dec(x_101); -x_112 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_112, 0, x_110); -lean_ctor_set(x_112, 1, x_111); -return x_112; +lean_inc(x_109); +lean_dec(x_100); +x_111 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_111, 0, x_109); +lean_ctor_set(x_111, 1, x_110); +return x_111; } } } else { -lean_object* x_113; lean_object* x_114; +lean_object* x_112; lean_object* x_113; lean_dec(x_2); -x_113 = lean_ctor_get(x_1, 0); -lean_inc(x_113); +x_112 = lean_ctor_get(x_1, 0); +lean_inc(x_112); lean_dec(x_1); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_114 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_113, x_4, x_5, x_6); -if (lean_obj_tag(x_114) == 0) +x_113 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_112, x_4, x_5, x_6); +if (lean_obj_tag(x_113) == 0) { -lean_object* x_115; lean_object* x_116; lean_object* x_117; -x_115 = lean_ctor_get(x_114, 0); +lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_114 = lean_ctor_get(x_113, 0); +lean_inc(x_114); +x_115 = lean_ctor_get(x_113, 1); lean_inc(x_115); -x_116 = lean_ctor_get(x_114, 1); -lean_inc(x_116); -lean_dec(x_114); -x_117 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_115, x_4, x_5, x_116); +lean_dec(x_113); +x_116 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_114, x_4, x_5, x_115); lean_dec(x_5); lean_dec(x_4); -return x_117; +return x_116; } else { -uint8_t x_118; +uint8_t x_117; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_118 = !lean_is_exclusive(x_114); -if (x_118 == 0) +x_117 = !lean_is_exclusive(x_113); +if (x_117 == 0) { -return x_114; +return x_113; } else { -lean_object* x_119; lean_object* x_120; lean_object* x_121; -x_119 = lean_ctor_get(x_114, 0); -x_120 = lean_ctor_get(x_114, 1); -lean_inc(x_120); +lean_object* x_118; lean_object* x_119; lean_object* x_120; +x_118 = lean_ctor_get(x_113, 0); +x_119 = lean_ctor_get(x_113, 1); lean_inc(x_119); -lean_dec(x_114); -x_121 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_121, 0, x_119); -lean_ctor_set(x_121, 1, x_120); -return x_121; +lean_inc(x_118); +lean_dec(x_113); +x_120 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_120, 0, x_118); +lean_ctor_set(x_120, 1, x_119); +return x_120; } } } } else { -lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; uint8_t x_127; -x_122 = lean_unsigned_to_nat(0u); -x_123 = l_Lean_Syntax_getArg(x_69, x_122); +lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; uint8_t x_126; +x_121 = lean_unsigned_to_nat(0u); +x_122 = l_Lean_Syntax_getArg(x_69, x_121); lean_dec(x_69); -x_124 = lean_unsigned_to_nat(3u); -x_125 = l_Lean_Syntax_getArg(x_3, x_124); -x_126 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__11; -lean_inc(x_125); -x_127 = l_Lean_Syntax_isOfKind(x_125, x_126); -if (x_127 == 0) -{ -lean_dec(x_125); -lean_dec(x_123); +x_123 = lean_unsigned_to_nat(3u); +x_124 = l_Lean_Syntax_getArg(x_3, x_123); +x_125 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__10; +lean_inc(x_124); +x_126 = l_Lean_Syntax_isOfKind(x_124, x_125); +if (x_126 == 0) +{ +lean_dec(x_124); +lean_dec(x_122); lean_dec(x_7); if (lean_obj_tag(x_1) == 0) { -lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; -x_128 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); -x_129 = lean_ctor_get(x_128, 0); +lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; +x_127 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); +x_128 = lean_ctor_get(x_127, 0); +lean_inc(x_128); +x_129 = lean_ctor_get(x_127, 1); lean_inc(x_129); -x_130 = lean_ctor_get(x_128, 1); -lean_inc(x_130); -lean_dec(x_128); -x_131 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_130); -x_132 = lean_ctor_get(x_131, 0); +lean_dec(x_127); +x_130 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_129); +x_131 = lean_ctor_get(x_130, 0); +lean_inc(x_131); +x_132 = lean_ctor_get(x_130, 1); lean_inc(x_132); -x_133 = lean_ctor_get(x_131, 1); -lean_inc(x_133); -lean_dec(x_131); -x_134 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_133); -x_135 = lean_ctor_get(x_134, 0); +lean_dec(x_130); +x_133 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_132); +x_134 = lean_ctor_get(x_133, 0); +lean_inc(x_134); +x_135 = lean_ctor_get(x_133, 1); lean_inc(x_135); -x_136 = lean_ctor_get(x_134, 1); -lean_inc(x_136); -lean_dec(x_134); -x_137 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; -x_138 = l_Lean_addMacroScope(x_135, x_137, x_132); -x_139 = lean_box(0); -x_140 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; -lean_inc(x_129); -x_141 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_141, 0, x_129); -lean_ctor_set(x_141, 1, x_140); -lean_ctor_set(x_141, 2, x_138); -lean_ctor_set(x_141, 3, x_139); -x_142 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; -lean_inc(x_129); -x_143 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_143, 0, x_129); -lean_ctor_set(x_143, 1, x_142); -x_144 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; +lean_dec(x_133); +x_136 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; +x_137 = l_Lean_addMacroScope(x_134, x_136, x_131); +x_138 = lean_box(0); +x_139 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; +lean_inc(x_128); +x_140 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_140, 0, x_128); +lean_ctor_set(x_140, 1, x_139); +lean_ctor_set(x_140, 2, x_137); +lean_ctor_set(x_140, 3, x_138); +x_141 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; +lean_inc(x_128); +x_142 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_142, 0, x_128); +lean_ctor_set(x_142, 1, x_141); +x_143 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; lean_inc(x_3); -x_145 = lean_array_push(x_144, x_3); -x_146 = lean_box(2); -x_147 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; -x_148 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_148, 0, x_146); -lean_ctor_set(x_148, 1, x_147); -lean_ctor_set(x_148, 2, x_145); -x_149 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__8; -x_150 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_150, 0, x_129); -lean_ctor_set(x_150, 1, x_149); -x_151 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__9; -x_152 = lean_array_push(x_151, x_141); -x_153 = lean_array_push(x_152, x_143); -x_154 = lean_array_push(x_153, x_148); -x_155 = lean_array_push(x_154, x_150); -x_156 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_156, 0, x_146); -lean_ctor_set(x_156, 1, x_12); -lean_ctor_set(x_156, 2, x_155); -x_157 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_136); -if (lean_obj_tag(x_157) == 0) -{ -uint8_t x_158; -x_158 = !lean_is_exclusive(x_157); -if (x_158 == 0) -{ -lean_object* x_159; lean_object* x_160; -x_159 = lean_ctor_get(x_157, 0); -x_160 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_160, 0, x_156); -lean_ctor_set(x_160, 1, x_159); -lean_ctor_set(x_157, 0, x_160); -return x_157; -} -else -{ -lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; -x_161 = lean_ctor_get(x_157, 0); -x_162 = lean_ctor_get(x_157, 1); -lean_inc(x_162); +x_144 = lean_array_push(x_143, x_3); +x_145 = lean_box(2); +x_146 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__3; +x_147 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_147, 0, x_145); +lean_ctor_set(x_147, 1, x_146); +lean_ctor_set(x_147, 2, x_144); +x_148 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__6; +x_149 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_149, 0, x_128); +lean_ctor_set(x_149, 1, x_148); +x_150 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; +x_151 = lean_array_push(x_150, x_140); +x_152 = lean_array_push(x_151, x_142); +x_153 = lean_array_push(x_152, x_147); +x_154 = lean_array_push(x_153, x_149); +x_155 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_155, 0, x_145); +lean_ctor_set(x_155, 1, x_12); +lean_ctor_set(x_155, 2, x_154); +x_156 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_135); +if (lean_obj_tag(x_156) == 0) +{ +uint8_t x_157; +x_157 = !lean_is_exclusive(x_156); +if (x_157 == 0) +{ +lean_object* x_158; lean_object* x_159; +x_158 = lean_ctor_get(x_156, 0); +x_159 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_159, 0, x_155); +lean_ctor_set(x_159, 1, x_158); +lean_ctor_set(x_156, 0, x_159); +return x_156; +} +else +{ +lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; +x_160 = lean_ctor_get(x_156, 0); +x_161 = lean_ctor_get(x_156, 1); lean_inc(x_161); -lean_dec(x_157); +lean_inc(x_160); +lean_dec(x_156); +x_162 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_162, 0, x_155); +lean_ctor_set(x_162, 1, x_160); x_163 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_163, 0, x_156); +lean_ctor_set(x_163, 0, x_162); lean_ctor_set(x_163, 1, x_161); -x_164 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_164, 0, x_163); -lean_ctor_set(x_164, 1, x_162); -return x_164; +return x_163; } } else { -uint8_t x_165; -lean_dec(x_156); -x_165 = !lean_is_exclusive(x_157); -if (x_165 == 0) +uint8_t x_164; +lean_dec(x_155); +x_164 = !lean_is_exclusive(x_156); +if (x_164 == 0) { -return x_157; +return x_156; } else { -lean_object* x_166; lean_object* x_167; lean_object* x_168; -x_166 = lean_ctor_get(x_157, 0); -x_167 = lean_ctor_get(x_157, 1); -lean_inc(x_167); +lean_object* x_165; lean_object* x_166; lean_object* x_167; +x_165 = lean_ctor_get(x_156, 0); +x_166 = lean_ctor_get(x_156, 1); lean_inc(x_166); -lean_dec(x_157); -x_168 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_168, 0, x_166); -lean_ctor_set(x_168, 1, x_167); -return x_168; +lean_inc(x_165); +lean_dec(x_156); +x_167 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_167, 0, x_165); +lean_ctor_set(x_167, 1, x_166); +return x_167; } } } else { -lean_object* x_169; lean_object* x_170; +lean_object* x_168; lean_object* x_169; lean_dec(x_2); -x_169 = lean_ctor_get(x_1, 0); -lean_inc(x_169); +x_168 = lean_ctor_get(x_1, 0); +lean_inc(x_168); lean_dec(x_1); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_170 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_169, x_4, x_5, x_6); -if (lean_obj_tag(x_170) == 0) +x_169 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_168, x_4, x_5, x_6); +if (lean_obj_tag(x_169) == 0) { -lean_object* x_171; lean_object* x_172; lean_object* x_173; -x_171 = lean_ctor_get(x_170, 0); +lean_object* x_170; lean_object* x_171; lean_object* x_172; +x_170 = lean_ctor_get(x_169, 0); +lean_inc(x_170); +x_171 = lean_ctor_get(x_169, 1); lean_inc(x_171); -x_172 = lean_ctor_get(x_170, 1); -lean_inc(x_172); -lean_dec(x_170); -x_173 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_171, x_4, x_5, x_172); +lean_dec(x_169); +x_172 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_170, x_4, x_5, x_171); lean_dec(x_5); lean_dec(x_4); -return x_173; +return x_172; } else { -uint8_t x_174; +uint8_t x_173; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_174 = !lean_is_exclusive(x_170); -if (x_174 == 0) +x_173 = !lean_is_exclusive(x_169); +if (x_173 == 0) { -return x_170; +return x_169; } else { -lean_object* x_175; lean_object* x_176; lean_object* x_177; -x_175 = lean_ctor_get(x_170, 0); -x_176 = lean_ctor_get(x_170, 1); -lean_inc(x_176); +lean_object* x_174; lean_object* x_175; lean_object* x_176; +x_174 = lean_ctor_get(x_169, 0); +x_175 = lean_ctor_get(x_169, 1); lean_inc(x_175); -lean_dec(x_170); -x_177 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_177, 0, x_175); -lean_ctor_set(x_177, 1, x_176); -return x_177; +lean_inc(x_174); +lean_dec(x_169); +x_176 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_176, 0, x_174); +lean_ctor_set(x_176, 1, x_175); +return x_176; } } } } else { -lean_object* x_178; lean_object* x_179; uint8_t x_180; -x_178 = lean_unsigned_to_nat(4u); -x_179 = l_Lean_Syntax_getArg(x_3, x_178); -x_180 = l_Lean_Syntax_isNone(x_179); -if (x_180 == 0) +lean_object* x_177; lean_object* x_178; uint8_t x_179; +x_177 = lean_unsigned_to_nat(4u); +x_178 = l_Lean_Syntax_getArg(x_3, x_177); +x_179 = l_Lean_Syntax_isNone(x_178); +if (x_179 == 0) { -lean_object* x_181; uint8_t x_182; -x_181 = lean_unsigned_to_nat(2u); -lean_inc(x_179); -x_182 = l_Lean_Syntax_isNodeOf(x_179, x_70, x_181); -if (x_182 == 0) +lean_object* x_180; uint8_t x_181; +x_180 = lean_unsigned_to_nat(2u); +lean_inc(x_178); +x_181 = l_Lean_Syntax_matchesNull(x_178, x_180); +if (x_181 == 0) { -lean_dec(x_179); -lean_dec(x_125); -lean_dec(x_123); +lean_dec(x_178); +lean_dec(x_124); +lean_dec(x_122); lean_dec(x_7); if (lean_obj_tag(x_1) == 0) { -lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; -x_183 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); -x_184 = lean_ctor_get(x_183, 0); +lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; +x_182 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); +x_183 = lean_ctor_get(x_182, 0); +lean_inc(x_183); +x_184 = lean_ctor_get(x_182, 1); lean_inc(x_184); -x_185 = lean_ctor_get(x_183, 1); -lean_inc(x_185); -lean_dec(x_183); -x_186 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_185); -x_187 = lean_ctor_get(x_186, 0); +lean_dec(x_182); +x_185 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_184); +x_186 = lean_ctor_get(x_185, 0); +lean_inc(x_186); +x_187 = lean_ctor_get(x_185, 1); lean_inc(x_187); -x_188 = lean_ctor_get(x_186, 1); -lean_inc(x_188); -lean_dec(x_186); -x_189 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_188); -x_190 = lean_ctor_get(x_189, 0); +lean_dec(x_185); +x_188 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_187); +x_189 = lean_ctor_get(x_188, 0); +lean_inc(x_189); +x_190 = lean_ctor_get(x_188, 1); lean_inc(x_190); -x_191 = lean_ctor_get(x_189, 1); -lean_inc(x_191); -lean_dec(x_189); -x_192 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; -x_193 = l_Lean_addMacroScope(x_190, x_192, x_187); -x_194 = lean_box(0); -x_195 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; -lean_inc(x_184); -x_196 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_196, 0, x_184); -lean_ctor_set(x_196, 1, x_195); -lean_ctor_set(x_196, 2, x_193); -lean_ctor_set(x_196, 3, x_194); -x_197 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; -lean_inc(x_184); -x_198 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_198, 0, x_184); -lean_ctor_set(x_198, 1, x_197); -x_199 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; +lean_dec(x_188); +x_191 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; +x_192 = l_Lean_addMacroScope(x_189, x_191, x_186); +x_193 = lean_box(0); +x_194 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; +lean_inc(x_183); +x_195 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_195, 0, x_183); +lean_ctor_set(x_195, 1, x_194); +lean_ctor_set(x_195, 2, x_192); +lean_ctor_set(x_195, 3, x_193); +x_196 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; +lean_inc(x_183); +x_197 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_197, 0, x_183); +lean_ctor_set(x_197, 1, x_196); +x_198 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; lean_inc(x_3); -x_200 = lean_array_push(x_199, x_3); -x_201 = lean_box(2); -x_202 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; -x_203 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_203, 0, x_201); -lean_ctor_set(x_203, 1, x_202); -lean_ctor_set(x_203, 2, x_200); -x_204 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__8; -x_205 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_205, 0, x_184); -lean_ctor_set(x_205, 1, x_204); -x_206 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__9; -x_207 = lean_array_push(x_206, x_196); -x_208 = lean_array_push(x_207, x_198); -x_209 = lean_array_push(x_208, x_203); -x_210 = lean_array_push(x_209, x_205); -x_211 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_211, 0, x_201); -lean_ctor_set(x_211, 1, x_12); -lean_ctor_set(x_211, 2, x_210); -x_212 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_191); -if (lean_obj_tag(x_212) == 0) -{ -uint8_t x_213; -x_213 = !lean_is_exclusive(x_212); -if (x_213 == 0) -{ -lean_object* x_214; lean_object* x_215; -x_214 = lean_ctor_get(x_212, 0); -x_215 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_215, 0, x_211); -lean_ctor_set(x_215, 1, x_214); -lean_ctor_set(x_212, 0, x_215); -return x_212; -} -else -{ -lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; -x_216 = lean_ctor_get(x_212, 0); -x_217 = lean_ctor_get(x_212, 1); -lean_inc(x_217); +x_199 = lean_array_push(x_198, x_3); +x_200 = lean_box(2); +x_201 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__3; +x_202 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_202, 0, x_200); +lean_ctor_set(x_202, 1, x_201); +lean_ctor_set(x_202, 2, x_199); +x_203 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__6; +x_204 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_204, 0, x_183); +lean_ctor_set(x_204, 1, x_203); +x_205 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; +x_206 = lean_array_push(x_205, x_195); +x_207 = lean_array_push(x_206, x_197); +x_208 = lean_array_push(x_207, x_202); +x_209 = lean_array_push(x_208, x_204); +x_210 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_210, 0, x_200); +lean_ctor_set(x_210, 1, x_12); +lean_ctor_set(x_210, 2, x_209); +x_211 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_190); +if (lean_obj_tag(x_211) == 0) +{ +uint8_t x_212; +x_212 = !lean_is_exclusive(x_211); +if (x_212 == 0) +{ +lean_object* x_213; lean_object* x_214; +x_213 = lean_ctor_get(x_211, 0); +x_214 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_214, 0, x_210); +lean_ctor_set(x_214, 1, x_213); +lean_ctor_set(x_211, 0, x_214); +return x_211; +} +else +{ +lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; +x_215 = lean_ctor_get(x_211, 0); +x_216 = lean_ctor_get(x_211, 1); lean_inc(x_216); -lean_dec(x_212); +lean_inc(x_215); +lean_dec(x_211); +x_217 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_217, 0, x_210); +lean_ctor_set(x_217, 1, x_215); x_218 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_218, 0, x_211); +lean_ctor_set(x_218, 0, x_217); lean_ctor_set(x_218, 1, x_216); -x_219 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_219, 0, x_218); -lean_ctor_set(x_219, 1, x_217); -return x_219; +return x_218; } } else { -uint8_t x_220; -lean_dec(x_211); -x_220 = !lean_is_exclusive(x_212); -if (x_220 == 0) +uint8_t x_219; +lean_dec(x_210); +x_219 = !lean_is_exclusive(x_211); +if (x_219 == 0) { -return x_212; +return x_211; } else { -lean_object* x_221; lean_object* x_222; lean_object* x_223; -x_221 = lean_ctor_get(x_212, 0); -x_222 = lean_ctor_get(x_212, 1); -lean_inc(x_222); +lean_object* x_220; lean_object* x_221; lean_object* x_222; +x_220 = lean_ctor_get(x_211, 0); +x_221 = lean_ctor_get(x_211, 1); lean_inc(x_221); -lean_dec(x_212); -x_223 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_223, 0, x_221); -lean_ctor_set(x_223, 1, x_222); -return x_223; +lean_inc(x_220); +lean_dec(x_211); +x_222 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_222, 0, x_220); +lean_ctor_set(x_222, 1, x_221); +return x_222; } } } else { -lean_object* x_224; lean_object* x_225; +lean_object* x_223; lean_object* x_224; lean_dec(x_2); -x_224 = lean_ctor_get(x_1, 0); -lean_inc(x_224); +x_223 = lean_ctor_get(x_1, 0); +lean_inc(x_223); lean_dec(x_1); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_225 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_224, x_4, x_5, x_6); -if (lean_obj_tag(x_225) == 0) +x_224 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_223, x_4, x_5, x_6); +if (lean_obj_tag(x_224) == 0) { -lean_object* x_226; lean_object* x_227; lean_object* x_228; -x_226 = lean_ctor_get(x_225, 0); +lean_object* x_225; lean_object* x_226; lean_object* x_227; +x_225 = lean_ctor_get(x_224, 0); +lean_inc(x_225); +x_226 = lean_ctor_get(x_224, 1); lean_inc(x_226); -x_227 = lean_ctor_get(x_225, 1); -lean_inc(x_227); -lean_dec(x_225); -x_228 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_226, x_4, x_5, x_227); +lean_dec(x_224); +x_227 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_225, x_4, x_5, x_226); lean_dec(x_5); lean_dec(x_4); -return x_228; +return x_227; } else { -uint8_t x_229; +uint8_t x_228; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_229 = !lean_is_exclusive(x_225); -if (x_229 == 0) +x_228 = !lean_is_exclusive(x_224); +if (x_228 == 0) { -return x_225; +return x_224; } else { -lean_object* x_230; lean_object* x_231; lean_object* x_232; -x_230 = lean_ctor_get(x_225, 0); -x_231 = lean_ctor_get(x_225, 1); -lean_inc(x_231); +lean_object* x_229; lean_object* x_230; lean_object* x_231; +x_229 = lean_ctor_get(x_224, 0); +x_230 = lean_ctor_get(x_224, 1); lean_inc(x_230); -lean_dec(x_225); -x_232 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_232, 0, x_230); -lean_ctor_set(x_232, 1, x_231); -return x_232; +lean_inc(x_229); +lean_dec(x_224); +x_231 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_231, 0, x_229); +lean_ctor_set(x_231, 1, x_230); +return x_231; } } } } else { -lean_object* x_233; uint8_t x_234; -x_233 = l_Lean_Syntax_getArg(x_179, x_68); -lean_dec(x_179); -lean_inc(x_233); -x_234 = l_Lean_Syntax_isNodeOf(x_233, x_70, x_68); -if (x_234 == 0) +lean_object* x_232; uint8_t x_233; +x_232 = l_Lean_Syntax_getArg(x_178, x_68); +lean_dec(x_178); +lean_inc(x_232); +x_233 = l_Lean_Syntax_matchesNull(x_232, x_68); +if (x_233 == 0) { -lean_dec(x_233); -lean_dec(x_125); -lean_dec(x_123); +lean_dec(x_232); +lean_dec(x_124); +lean_dec(x_122); lean_dec(x_7); if (lean_obj_tag(x_1) == 0) { -lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; -x_235 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); -x_236 = lean_ctor_get(x_235, 0); +lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; +x_234 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); +x_235 = lean_ctor_get(x_234, 0); +lean_inc(x_235); +x_236 = lean_ctor_get(x_234, 1); lean_inc(x_236); -x_237 = lean_ctor_get(x_235, 1); -lean_inc(x_237); -lean_dec(x_235); -x_238 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_237); -x_239 = lean_ctor_get(x_238, 0); +lean_dec(x_234); +x_237 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_236); +x_238 = lean_ctor_get(x_237, 0); +lean_inc(x_238); +x_239 = lean_ctor_get(x_237, 1); lean_inc(x_239); -x_240 = lean_ctor_get(x_238, 1); -lean_inc(x_240); -lean_dec(x_238); -x_241 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_240); -x_242 = lean_ctor_get(x_241, 0); +lean_dec(x_237); +x_240 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_239); +x_241 = lean_ctor_get(x_240, 0); +lean_inc(x_241); +x_242 = lean_ctor_get(x_240, 1); lean_inc(x_242); -x_243 = lean_ctor_get(x_241, 1); -lean_inc(x_243); -lean_dec(x_241); -x_244 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; -x_245 = l_Lean_addMacroScope(x_242, x_244, x_239); -x_246 = lean_box(0); -x_247 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; -lean_inc(x_236); -x_248 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_248, 0, x_236); -lean_ctor_set(x_248, 1, x_247); -lean_ctor_set(x_248, 2, x_245); -lean_ctor_set(x_248, 3, x_246); -x_249 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; -lean_inc(x_236); -x_250 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_250, 0, x_236); -lean_ctor_set(x_250, 1, x_249); -x_251 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; +lean_dec(x_240); +x_243 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; +x_244 = l_Lean_addMacroScope(x_241, x_243, x_238); +x_245 = lean_box(0); +x_246 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; +lean_inc(x_235); +x_247 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_247, 0, x_235); +lean_ctor_set(x_247, 1, x_246); +lean_ctor_set(x_247, 2, x_244); +lean_ctor_set(x_247, 3, x_245); +x_248 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; +lean_inc(x_235); +x_249 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_249, 0, x_235); +lean_ctor_set(x_249, 1, x_248); +x_250 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; lean_inc(x_3); -x_252 = lean_array_push(x_251, x_3); -x_253 = lean_box(2); -x_254 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; -x_255 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_255, 0, x_253); -lean_ctor_set(x_255, 1, x_254); -lean_ctor_set(x_255, 2, x_252); -x_256 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__8; -x_257 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_257, 0, x_236); -lean_ctor_set(x_257, 1, x_256); -x_258 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__9; -x_259 = lean_array_push(x_258, x_248); -x_260 = lean_array_push(x_259, x_250); -x_261 = lean_array_push(x_260, x_255); -x_262 = lean_array_push(x_261, x_257); -x_263 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_263, 0, x_253); -lean_ctor_set(x_263, 1, x_12); -lean_ctor_set(x_263, 2, x_262); -x_264 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_243); -if (lean_obj_tag(x_264) == 0) -{ -uint8_t x_265; -x_265 = !lean_is_exclusive(x_264); -if (x_265 == 0) -{ -lean_object* x_266; lean_object* x_267; -x_266 = lean_ctor_get(x_264, 0); -x_267 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_267, 0, x_263); -lean_ctor_set(x_267, 1, x_266); -lean_ctor_set(x_264, 0, x_267); -return x_264; -} -else -{ -lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; -x_268 = lean_ctor_get(x_264, 0); -x_269 = lean_ctor_get(x_264, 1); -lean_inc(x_269); +x_251 = lean_array_push(x_250, x_3); +x_252 = lean_box(2); +x_253 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__3; +x_254 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_254, 0, x_252); +lean_ctor_set(x_254, 1, x_253); +lean_ctor_set(x_254, 2, x_251); +x_255 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__6; +x_256 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_256, 0, x_235); +lean_ctor_set(x_256, 1, x_255); +x_257 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; +x_258 = lean_array_push(x_257, x_247); +x_259 = lean_array_push(x_258, x_249); +x_260 = lean_array_push(x_259, x_254); +x_261 = lean_array_push(x_260, x_256); +x_262 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_262, 0, x_252); +lean_ctor_set(x_262, 1, x_12); +lean_ctor_set(x_262, 2, x_261); +x_263 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_242); +if (lean_obj_tag(x_263) == 0) +{ +uint8_t x_264; +x_264 = !lean_is_exclusive(x_263); +if (x_264 == 0) +{ +lean_object* x_265; lean_object* x_266; +x_265 = lean_ctor_get(x_263, 0); +x_266 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_266, 0, x_262); +lean_ctor_set(x_266, 1, x_265); +lean_ctor_set(x_263, 0, x_266); +return x_263; +} +else +{ +lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; +x_267 = lean_ctor_get(x_263, 0); +x_268 = lean_ctor_get(x_263, 1); lean_inc(x_268); -lean_dec(x_264); +lean_inc(x_267); +lean_dec(x_263); +x_269 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_269, 0, x_262); +lean_ctor_set(x_269, 1, x_267); x_270 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_270, 0, x_263); +lean_ctor_set(x_270, 0, x_269); lean_ctor_set(x_270, 1, x_268); -x_271 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_271, 0, x_270); -lean_ctor_set(x_271, 1, x_269); -return x_271; +return x_270; } } else { -uint8_t x_272; -lean_dec(x_263); -x_272 = !lean_is_exclusive(x_264); -if (x_272 == 0) +uint8_t x_271; +lean_dec(x_262); +x_271 = !lean_is_exclusive(x_263); +if (x_271 == 0) { -return x_264; +return x_263; } else { -lean_object* x_273; lean_object* x_274; lean_object* x_275; -x_273 = lean_ctor_get(x_264, 0); -x_274 = lean_ctor_get(x_264, 1); -lean_inc(x_274); +lean_object* x_272; lean_object* x_273; lean_object* x_274; +x_272 = lean_ctor_get(x_263, 0); +x_273 = lean_ctor_get(x_263, 1); lean_inc(x_273); -lean_dec(x_264); -x_275 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_275, 0, x_273); -lean_ctor_set(x_275, 1, x_274); -return x_275; +lean_inc(x_272); +lean_dec(x_263); +x_274 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_274, 0, x_272); +lean_ctor_set(x_274, 1, x_273); +return x_274; } } } else { -lean_object* x_276; lean_object* x_277; +lean_object* x_275; lean_object* x_276; lean_dec(x_2); -x_276 = lean_ctor_get(x_1, 0); -lean_inc(x_276); +x_275 = lean_ctor_get(x_1, 0); +lean_inc(x_275); lean_dec(x_1); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_277 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_276, x_4, x_5, x_6); -if (lean_obj_tag(x_277) == 0) +x_276 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_275, x_4, x_5, x_6); +if (lean_obj_tag(x_276) == 0) { -lean_object* x_278; lean_object* x_279; lean_object* x_280; -x_278 = lean_ctor_get(x_277, 0); +lean_object* x_277; lean_object* x_278; lean_object* x_279; +x_277 = lean_ctor_get(x_276, 0); +lean_inc(x_277); +x_278 = lean_ctor_get(x_276, 1); lean_inc(x_278); -x_279 = lean_ctor_get(x_277, 1); -lean_inc(x_279); -lean_dec(x_277); -x_280 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_278, x_4, x_5, x_279); +lean_dec(x_276); +x_279 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_277, x_4, x_5, x_278); lean_dec(x_5); lean_dec(x_4); -return x_280; +return x_279; } else { -uint8_t x_281; +uint8_t x_280; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_281 = !lean_is_exclusive(x_277); -if (x_281 == 0) +x_280 = !lean_is_exclusive(x_276); +if (x_280 == 0) { -return x_277; +return x_276; } else { -lean_object* x_282; lean_object* x_283; lean_object* x_284; -x_282 = lean_ctor_get(x_277, 0); -x_283 = lean_ctor_get(x_277, 1); -lean_inc(x_283); +lean_object* x_281; lean_object* x_282; lean_object* x_283; +x_281 = lean_ctor_get(x_276, 0); +x_282 = lean_ctor_get(x_276, 1); lean_inc(x_282); -lean_dec(x_277); -x_284 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_284, 0, x_282); -lean_ctor_set(x_284, 1, x_283); -return x_284; +lean_inc(x_281); +lean_dec(x_276); +x_283 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_283, 0, x_281); +lean_ctor_set(x_283, 1, x_282); +return x_283; } } } } else { -lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; -x_285 = l_Lean_Syntax_getArg(x_233, x_122); -lean_dec(x_233); -x_286 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_286, 0, x_285); -x_287 = lean_box(0); +lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; +x_284 = l_Lean_Syntax_getArg(x_232, x_121); +lean_dec(x_232); +x_285 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_285, 0, x_284); +x_286 = lean_box(0); lean_inc(x_3); -x_288 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3(x_3, x_125, x_123, x_2, x_7, x_1, x_3, x_12, x_287, x_286, x_4, x_5, x_6); -lean_dec(x_286); -return x_288; +x_287 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3(x_3, x_124, x_122, x_2, x_7, x_1, x_3, x_12, x_286, x_285, x_4, x_5, x_6); +lean_dec(x_285); +return x_287; } } } else { -lean_object* x_289; lean_object* x_290; lean_object* x_291; -lean_dec(x_179); +lean_object* x_288; lean_object* x_289; lean_object* x_290; +lean_dec(x_178); +x_288 = lean_box(0); x_289 = lean_box(0); -x_290 = lean_box(0); lean_inc(x_3); -x_291 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3(x_3, x_125, x_123, x_2, x_7, x_1, x_3, x_12, x_290, x_289, x_4, x_5, x_6); -return x_291; +x_290 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3(x_3, x_124, x_122, x_2, x_7, x_1, x_3, x_12, x_289, x_288, x_4, x_5, x_6); +return x_290; } } } @@ -10544,747 +10617,746 @@ return x_291; } else { -lean_object* x_292; lean_object* x_293; lean_object* x_294; uint8_t x_295; -x_292 = lean_unsigned_to_nat(1u); -x_293 = l_Lean_Syntax_getArg(x_3, x_292); -x_294 = l_Lean_nullKind; -lean_inc(x_293); -x_295 = l_Lean_Syntax_isNodeOf(x_293, x_294, x_292); -if (x_295 == 0) +lean_object* x_291; lean_object* x_292; uint8_t x_293; +x_291 = lean_unsigned_to_nat(1u); +x_292 = l_Lean_Syntax_getArg(x_3, x_291); +lean_inc(x_292); +x_293 = l_Lean_Syntax_matchesNull(x_292, x_291); +if (x_293 == 0) { -lean_dec(x_293); +lean_dec(x_292); lean_dec(x_7); if (lean_obj_tag(x_1) == 0) { -lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; -x_296 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); -x_297 = lean_ctor_get(x_296, 0); -lean_inc(x_297); -x_298 = lean_ctor_get(x_296, 1); +lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; +x_294 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); +x_295 = lean_ctor_get(x_294, 0); +lean_inc(x_295); +x_296 = lean_ctor_get(x_294, 1); +lean_inc(x_296); +lean_dec(x_294); +x_297 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_296); +x_298 = lean_ctor_get(x_297, 0); lean_inc(x_298); -lean_dec(x_296); -x_299 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_298); -x_300 = lean_ctor_get(x_299, 0); -lean_inc(x_300); -x_301 = lean_ctor_get(x_299, 1); +x_299 = lean_ctor_get(x_297, 1); +lean_inc(x_299); +lean_dec(x_297); +x_300 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_299); +x_301 = lean_ctor_get(x_300, 0); lean_inc(x_301); -lean_dec(x_299); -x_302 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_301); -x_303 = lean_ctor_get(x_302, 0); -lean_inc(x_303); -x_304 = lean_ctor_get(x_302, 1); -lean_inc(x_304); -lean_dec(x_302); -x_305 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; -x_306 = l_Lean_addMacroScope(x_303, x_305, x_300); -x_307 = lean_box(0); -x_308 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; -lean_inc(x_297); -x_309 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_309, 0, x_297); +x_302 = lean_ctor_get(x_300, 1); +lean_inc(x_302); +lean_dec(x_300); +x_303 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; +x_304 = l_Lean_addMacroScope(x_301, x_303, x_298); +x_305 = lean_box(0); +x_306 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; +lean_inc(x_295); +x_307 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_307, 0, x_295); +lean_ctor_set(x_307, 1, x_306); +lean_ctor_set(x_307, 2, x_304); +lean_ctor_set(x_307, 3, x_305); +x_308 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; +lean_inc(x_295); +x_309 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_309, 0, x_295); lean_ctor_set(x_309, 1, x_308); -lean_ctor_set(x_309, 2, x_306); -lean_ctor_set(x_309, 3, x_307); -x_310 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; -lean_inc(x_297); -x_311 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_311, 0, x_297); -lean_ctor_set(x_311, 1, x_310); -x_312 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; +x_310 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; lean_inc(x_3); -x_313 = lean_array_push(x_312, x_3); -x_314 = lean_box(2); -x_315 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; -x_316 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_316, 0, x_314); +x_311 = lean_array_push(x_310, x_3); +x_312 = lean_box(2); +x_313 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__3; +x_314 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_314, 0, x_312); +lean_ctor_set(x_314, 1, x_313); +lean_ctor_set(x_314, 2, x_311); +x_315 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__6; +x_316 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_316, 0, x_295); lean_ctor_set(x_316, 1, x_315); -lean_ctor_set(x_316, 2, x_313); -x_317 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__8; -x_318 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_318, 0, x_297); -lean_ctor_set(x_318, 1, x_317); -x_319 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__9; -x_320 = lean_array_push(x_319, x_309); -x_321 = lean_array_push(x_320, x_311); -x_322 = lean_array_push(x_321, x_316); -x_323 = lean_array_push(x_322, x_318); -x_324 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_324, 0, x_314); -lean_ctor_set(x_324, 1, x_12); -lean_ctor_set(x_324, 2, x_323); -x_325 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_304); -if (lean_obj_tag(x_325) == 0) -{ -uint8_t x_326; -x_326 = !lean_is_exclusive(x_325); -if (x_326 == 0) -{ -lean_object* x_327; lean_object* x_328; -x_327 = lean_ctor_get(x_325, 0); -x_328 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_328, 0, x_324); -lean_ctor_set(x_328, 1, x_327); -lean_ctor_set(x_325, 0, x_328); -return x_325; -} -else -{ -lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; -x_329 = lean_ctor_get(x_325, 0); -x_330 = lean_ctor_get(x_325, 1); -lean_inc(x_330); -lean_inc(x_329); -lean_dec(x_325); -x_331 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_331, 0, x_324); -lean_ctor_set(x_331, 1, x_329); -x_332 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_332, 0, x_331); -lean_ctor_set(x_332, 1, x_330); -return x_332; +x_317 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; +x_318 = lean_array_push(x_317, x_307); +x_319 = lean_array_push(x_318, x_309); +x_320 = lean_array_push(x_319, x_314); +x_321 = lean_array_push(x_320, x_316); +x_322 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_322, 0, x_312); +lean_ctor_set(x_322, 1, x_12); +lean_ctor_set(x_322, 2, x_321); +x_323 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_302); +if (lean_obj_tag(x_323) == 0) +{ +uint8_t x_324; +x_324 = !lean_is_exclusive(x_323); +if (x_324 == 0) +{ +lean_object* x_325; lean_object* x_326; +x_325 = lean_ctor_get(x_323, 0); +x_326 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_326, 0, x_322); +lean_ctor_set(x_326, 1, x_325); +lean_ctor_set(x_323, 0, x_326); +return x_323; +} +else +{ +lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; +x_327 = lean_ctor_get(x_323, 0); +x_328 = lean_ctor_get(x_323, 1); +lean_inc(x_328); +lean_inc(x_327); +lean_dec(x_323); +x_329 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_329, 0, x_322); +lean_ctor_set(x_329, 1, x_327); +x_330 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_330, 0, x_329); +lean_ctor_set(x_330, 1, x_328); +return x_330; } } else { -uint8_t x_333; -lean_dec(x_324); -x_333 = !lean_is_exclusive(x_325); -if (x_333 == 0) +uint8_t x_331; +lean_dec(x_322); +x_331 = !lean_is_exclusive(x_323); +if (x_331 == 0) { -return x_325; +return x_323; } else { -lean_object* x_334; lean_object* x_335; lean_object* x_336; -x_334 = lean_ctor_get(x_325, 0); -x_335 = lean_ctor_get(x_325, 1); -lean_inc(x_335); -lean_inc(x_334); -lean_dec(x_325); -x_336 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_336, 0, x_334); -lean_ctor_set(x_336, 1, x_335); -return x_336; +lean_object* x_332; lean_object* x_333; lean_object* x_334; +x_332 = lean_ctor_get(x_323, 0); +x_333 = lean_ctor_get(x_323, 1); +lean_inc(x_333); +lean_inc(x_332); +lean_dec(x_323); +x_334 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_334, 0, x_332); +lean_ctor_set(x_334, 1, x_333); +return x_334; } } } else { -lean_object* x_337; lean_object* x_338; +lean_object* x_335; lean_object* x_336; lean_dec(x_2); -x_337 = lean_ctor_get(x_1, 0); -lean_inc(x_337); +x_335 = lean_ctor_get(x_1, 0); +lean_inc(x_335); lean_dec(x_1); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_338 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_337, x_4, x_5, x_6); -if (lean_obj_tag(x_338) == 0) +x_336 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_335, x_4, x_5, x_6); +if (lean_obj_tag(x_336) == 0) { -lean_object* x_339; lean_object* x_340; lean_object* x_341; -x_339 = lean_ctor_get(x_338, 0); -lean_inc(x_339); -x_340 = lean_ctor_get(x_338, 1); -lean_inc(x_340); -lean_dec(x_338); -x_341 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_339, x_4, x_5, x_340); +lean_object* x_337; lean_object* x_338; lean_object* x_339; +x_337 = lean_ctor_get(x_336, 0); +lean_inc(x_337); +x_338 = lean_ctor_get(x_336, 1); +lean_inc(x_338); +lean_dec(x_336); +x_339 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_337, x_4, x_5, x_338); lean_dec(x_5); lean_dec(x_4); -return x_341; +return x_339; } else { -uint8_t x_342; +uint8_t x_340; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_342 = !lean_is_exclusive(x_338); -if (x_342 == 0) +x_340 = !lean_is_exclusive(x_336); +if (x_340 == 0) { -return x_338; +return x_336; } else { -lean_object* x_343; lean_object* x_344; lean_object* x_345; -x_343 = lean_ctor_get(x_338, 0); -x_344 = lean_ctor_get(x_338, 1); -lean_inc(x_344); -lean_inc(x_343); -lean_dec(x_338); -x_345 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_345, 0, x_343); -lean_ctor_set(x_345, 1, x_344); -return x_345; +lean_object* x_341; lean_object* x_342; lean_object* x_343; +x_341 = lean_ctor_get(x_336, 0); +x_342 = lean_ctor_get(x_336, 1); +lean_inc(x_342); +lean_inc(x_341); +lean_dec(x_336); +x_343 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_343, 0, x_341); +lean_ctor_set(x_343, 1, x_342); +return x_343; } } } } else { -lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; uint8_t x_351; -x_346 = lean_unsigned_to_nat(0u); -x_347 = l_Lean_Syntax_getArg(x_293, x_346); -lean_dec(x_293); -x_348 = lean_unsigned_to_nat(3u); -x_349 = l_Lean_Syntax_getArg(x_3, x_348); -x_350 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__11; -lean_inc(x_349); -x_351 = l_Lean_Syntax_isOfKind(x_349, x_350); -if (x_351 == 0) +lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; uint8_t x_349; +x_344 = lean_unsigned_to_nat(0u); +x_345 = l_Lean_Syntax_getArg(x_292, x_344); +lean_dec(x_292); +x_346 = lean_unsigned_to_nat(3u); +x_347 = l_Lean_Syntax_getArg(x_3, x_346); +x_348 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__10; +lean_inc(x_347); +x_349 = l_Lean_Syntax_isOfKind(x_347, x_348); +if (x_349 == 0) { -lean_dec(x_349); lean_dec(x_347); +lean_dec(x_345); lean_dec(x_7); if (lean_obj_tag(x_1) == 0) { -lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; -x_352 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); -x_353 = lean_ctor_get(x_352, 0); -lean_inc(x_353); -x_354 = lean_ctor_get(x_352, 1); +lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; +x_350 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); +x_351 = lean_ctor_get(x_350, 0); +lean_inc(x_351); +x_352 = lean_ctor_get(x_350, 1); +lean_inc(x_352); +lean_dec(x_350); +x_353 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_352); +x_354 = lean_ctor_get(x_353, 0); lean_inc(x_354); -lean_dec(x_352); -x_355 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_354); -x_356 = lean_ctor_get(x_355, 0); -lean_inc(x_356); -x_357 = lean_ctor_get(x_355, 1); +x_355 = lean_ctor_get(x_353, 1); +lean_inc(x_355); +lean_dec(x_353); +x_356 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_355); +x_357 = lean_ctor_get(x_356, 0); lean_inc(x_357); -lean_dec(x_355); -x_358 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_357); -x_359 = lean_ctor_get(x_358, 0); -lean_inc(x_359); -x_360 = lean_ctor_get(x_358, 1); -lean_inc(x_360); -lean_dec(x_358); -x_361 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; -x_362 = l_Lean_addMacroScope(x_359, x_361, x_356); -x_363 = lean_box(0); -x_364 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; -lean_inc(x_353); -x_365 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_365, 0, x_353); +x_358 = lean_ctor_get(x_356, 1); +lean_inc(x_358); +lean_dec(x_356); +x_359 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; +x_360 = l_Lean_addMacroScope(x_357, x_359, x_354); +x_361 = lean_box(0); +x_362 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; +lean_inc(x_351); +x_363 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_363, 0, x_351); +lean_ctor_set(x_363, 1, x_362); +lean_ctor_set(x_363, 2, x_360); +lean_ctor_set(x_363, 3, x_361); +x_364 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; +lean_inc(x_351); +x_365 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_365, 0, x_351); lean_ctor_set(x_365, 1, x_364); -lean_ctor_set(x_365, 2, x_362); -lean_ctor_set(x_365, 3, x_363); -x_366 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; -lean_inc(x_353); -x_367 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_367, 0, x_353); -lean_ctor_set(x_367, 1, x_366); -x_368 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; +x_366 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; lean_inc(x_3); -x_369 = lean_array_push(x_368, x_3); -x_370 = lean_box(2); -x_371 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; -x_372 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_372, 0, x_370); +x_367 = lean_array_push(x_366, x_3); +x_368 = lean_box(2); +x_369 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__3; +x_370 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_370, 0, x_368); +lean_ctor_set(x_370, 1, x_369); +lean_ctor_set(x_370, 2, x_367); +x_371 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__6; +x_372 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_372, 0, x_351); lean_ctor_set(x_372, 1, x_371); -lean_ctor_set(x_372, 2, x_369); -x_373 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__8; -x_374 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_374, 0, x_353); -lean_ctor_set(x_374, 1, x_373); -x_375 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__9; -x_376 = lean_array_push(x_375, x_365); -x_377 = lean_array_push(x_376, x_367); -x_378 = lean_array_push(x_377, x_372); -x_379 = lean_array_push(x_378, x_374); -x_380 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_380, 0, x_370); -lean_ctor_set(x_380, 1, x_12); -lean_ctor_set(x_380, 2, x_379); -x_381 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_360); -if (lean_obj_tag(x_381) == 0) -{ -uint8_t x_382; -x_382 = !lean_is_exclusive(x_381); -if (x_382 == 0) -{ -lean_object* x_383; lean_object* x_384; -x_383 = lean_ctor_get(x_381, 0); -x_384 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_384, 0, x_380); -lean_ctor_set(x_384, 1, x_383); -lean_ctor_set(x_381, 0, x_384); -return x_381; -} -else -{ -lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; -x_385 = lean_ctor_get(x_381, 0); -x_386 = lean_ctor_get(x_381, 1); -lean_inc(x_386); -lean_inc(x_385); -lean_dec(x_381); -x_387 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_387, 0, x_380); -lean_ctor_set(x_387, 1, x_385); -x_388 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_388, 0, x_387); -lean_ctor_set(x_388, 1, x_386); -return x_388; -} -} -else -{ -uint8_t x_389; -lean_dec(x_380); -x_389 = !lean_is_exclusive(x_381); -if (x_389 == 0) -{ -return x_381; -} -else -{ -lean_object* x_390; lean_object* x_391; lean_object* x_392; -x_390 = lean_ctor_get(x_381, 0); -x_391 = lean_ctor_get(x_381, 1); -lean_inc(x_391); -lean_inc(x_390); -lean_dec(x_381); -x_392 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_392, 0, x_390); -lean_ctor_set(x_392, 1, x_391); -return x_392; +x_373 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; +x_374 = lean_array_push(x_373, x_363); +x_375 = lean_array_push(x_374, x_365); +x_376 = lean_array_push(x_375, x_370); +x_377 = lean_array_push(x_376, x_372); +x_378 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_378, 0, x_368); +lean_ctor_set(x_378, 1, x_12); +lean_ctor_set(x_378, 2, x_377); +x_379 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_358); +if (lean_obj_tag(x_379) == 0) +{ +uint8_t x_380; +x_380 = !lean_is_exclusive(x_379); +if (x_380 == 0) +{ +lean_object* x_381; lean_object* x_382; +x_381 = lean_ctor_get(x_379, 0); +x_382 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_382, 0, x_378); +lean_ctor_set(x_382, 1, x_381); +lean_ctor_set(x_379, 0, x_382); +return x_379; +} +else +{ +lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; +x_383 = lean_ctor_get(x_379, 0); +x_384 = lean_ctor_get(x_379, 1); +lean_inc(x_384); +lean_inc(x_383); +lean_dec(x_379); +x_385 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_385, 0, x_378); +lean_ctor_set(x_385, 1, x_383); +x_386 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_386, 0, x_385); +lean_ctor_set(x_386, 1, x_384); +return x_386; +} +} +else +{ +uint8_t x_387; +lean_dec(x_378); +x_387 = !lean_is_exclusive(x_379); +if (x_387 == 0) +{ +return x_379; +} +else +{ +lean_object* x_388; lean_object* x_389; lean_object* x_390; +x_388 = lean_ctor_get(x_379, 0); +x_389 = lean_ctor_get(x_379, 1); +lean_inc(x_389); +lean_inc(x_388); +lean_dec(x_379); +x_390 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_390, 0, x_388); +lean_ctor_set(x_390, 1, x_389); +return x_390; } } } else { -lean_object* x_393; lean_object* x_394; +lean_object* x_391; lean_object* x_392; lean_dec(x_2); -x_393 = lean_ctor_get(x_1, 0); -lean_inc(x_393); +x_391 = lean_ctor_get(x_1, 0); +lean_inc(x_391); lean_dec(x_1); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_394 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_393, x_4, x_5, x_6); -if (lean_obj_tag(x_394) == 0) -{ -lean_object* x_395; lean_object* x_396; lean_object* x_397; -x_395 = lean_ctor_get(x_394, 0); -lean_inc(x_395); -x_396 = lean_ctor_get(x_394, 1); -lean_inc(x_396); -lean_dec(x_394); -x_397 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_395, x_4, x_5, x_396); +x_392 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_391, x_4, x_5, x_6); +if (lean_obj_tag(x_392) == 0) +{ +lean_object* x_393; lean_object* x_394; lean_object* x_395; +x_393 = lean_ctor_get(x_392, 0); +lean_inc(x_393); +x_394 = lean_ctor_get(x_392, 1); +lean_inc(x_394); +lean_dec(x_392); +x_395 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_393, x_4, x_5, x_394); lean_dec(x_5); lean_dec(x_4); -return x_397; +return x_395; } else { -uint8_t x_398; +uint8_t x_396; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_398 = !lean_is_exclusive(x_394); -if (x_398 == 0) +x_396 = !lean_is_exclusive(x_392); +if (x_396 == 0) { -return x_394; +return x_392; } else { -lean_object* x_399; lean_object* x_400; lean_object* x_401; -x_399 = lean_ctor_get(x_394, 0); -x_400 = lean_ctor_get(x_394, 1); -lean_inc(x_400); -lean_inc(x_399); -lean_dec(x_394); -x_401 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_401, 0, x_399); -lean_ctor_set(x_401, 1, x_400); -return x_401; +lean_object* x_397; lean_object* x_398; lean_object* x_399; +x_397 = lean_ctor_get(x_392, 0); +x_398 = lean_ctor_get(x_392, 1); +lean_inc(x_398); +lean_inc(x_397); +lean_dec(x_392); +x_399 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_399, 0, x_397); +lean_ctor_set(x_399, 1, x_398); +return x_399; } } } } else { -lean_object* x_402; lean_object* x_403; uint8_t x_404; -x_402 = lean_unsigned_to_nat(4u); -x_403 = l_Lean_Syntax_getArg(x_3, x_402); -x_404 = l_Lean_Syntax_isNone(x_403); -if (x_404 == 0) +lean_object* x_400; lean_object* x_401; uint8_t x_402; +x_400 = lean_unsigned_to_nat(4u); +x_401 = l_Lean_Syntax_getArg(x_3, x_400); +x_402 = l_Lean_Syntax_isNone(x_401); +if (x_402 == 0) { -lean_object* x_405; uint8_t x_406; -x_405 = lean_unsigned_to_nat(2u); -lean_inc(x_403); -x_406 = l_Lean_Syntax_isNodeOf(x_403, x_294, x_405); -if (x_406 == 0) +lean_object* x_403; uint8_t x_404; +x_403 = lean_unsigned_to_nat(2u); +lean_inc(x_401); +x_404 = l_Lean_Syntax_matchesNull(x_401, x_403); +if (x_404 == 0) { -lean_dec(x_403); -lean_dec(x_349); +lean_dec(x_401); lean_dec(x_347); +lean_dec(x_345); lean_dec(x_7); if (lean_obj_tag(x_1) == 0) { -lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; -x_407 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); -x_408 = lean_ctor_get(x_407, 0); -lean_inc(x_408); -x_409 = lean_ctor_get(x_407, 1); +lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; +x_405 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); +x_406 = lean_ctor_get(x_405, 0); +lean_inc(x_406); +x_407 = lean_ctor_get(x_405, 1); +lean_inc(x_407); +lean_dec(x_405); +x_408 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_407); +x_409 = lean_ctor_get(x_408, 0); lean_inc(x_409); -lean_dec(x_407); -x_410 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_409); -x_411 = lean_ctor_get(x_410, 0); -lean_inc(x_411); -x_412 = lean_ctor_get(x_410, 1); +x_410 = lean_ctor_get(x_408, 1); +lean_inc(x_410); +lean_dec(x_408); +x_411 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_410); +x_412 = lean_ctor_get(x_411, 0); lean_inc(x_412); -lean_dec(x_410); -x_413 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_412); -x_414 = lean_ctor_get(x_413, 0); -lean_inc(x_414); -x_415 = lean_ctor_get(x_413, 1); -lean_inc(x_415); -lean_dec(x_413); -x_416 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; -x_417 = l_Lean_addMacroScope(x_414, x_416, x_411); -x_418 = lean_box(0); -x_419 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; -lean_inc(x_408); -x_420 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_420, 0, x_408); +x_413 = lean_ctor_get(x_411, 1); +lean_inc(x_413); +lean_dec(x_411); +x_414 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; +x_415 = l_Lean_addMacroScope(x_412, x_414, x_409); +x_416 = lean_box(0); +x_417 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; +lean_inc(x_406); +x_418 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_418, 0, x_406); +lean_ctor_set(x_418, 1, x_417); +lean_ctor_set(x_418, 2, x_415); +lean_ctor_set(x_418, 3, x_416); +x_419 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; +lean_inc(x_406); +x_420 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_420, 0, x_406); lean_ctor_set(x_420, 1, x_419); -lean_ctor_set(x_420, 2, x_417); -lean_ctor_set(x_420, 3, x_418); -x_421 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; -lean_inc(x_408); -x_422 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_422, 0, x_408); -lean_ctor_set(x_422, 1, x_421); -x_423 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; +x_421 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; lean_inc(x_3); -x_424 = lean_array_push(x_423, x_3); -x_425 = lean_box(2); -x_426 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; -x_427 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_427, 0, x_425); +x_422 = lean_array_push(x_421, x_3); +x_423 = lean_box(2); +x_424 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__3; +x_425 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_425, 0, x_423); +lean_ctor_set(x_425, 1, x_424); +lean_ctor_set(x_425, 2, x_422); +x_426 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__6; +x_427 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_427, 0, x_406); lean_ctor_set(x_427, 1, x_426); -lean_ctor_set(x_427, 2, x_424); -x_428 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__8; -x_429 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_429, 0, x_408); -lean_ctor_set(x_429, 1, x_428); -x_430 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__9; -x_431 = lean_array_push(x_430, x_420); -x_432 = lean_array_push(x_431, x_422); -x_433 = lean_array_push(x_432, x_427); -x_434 = lean_array_push(x_433, x_429); -x_435 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_435, 0, x_425); -lean_ctor_set(x_435, 1, x_12); -lean_ctor_set(x_435, 2, x_434); -x_436 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_415); -if (lean_obj_tag(x_436) == 0) -{ -uint8_t x_437; -x_437 = !lean_is_exclusive(x_436); -if (x_437 == 0) -{ -lean_object* x_438; lean_object* x_439; -x_438 = lean_ctor_get(x_436, 0); -x_439 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_439, 0, x_435); -lean_ctor_set(x_439, 1, x_438); -lean_ctor_set(x_436, 0, x_439); -return x_436; -} -else -{ -lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; -x_440 = lean_ctor_get(x_436, 0); -x_441 = lean_ctor_get(x_436, 1); -lean_inc(x_441); -lean_inc(x_440); -lean_dec(x_436); -x_442 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_442, 0, x_435); -lean_ctor_set(x_442, 1, x_440); -x_443 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_443, 0, x_442); -lean_ctor_set(x_443, 1, x_441); -return x_443; -} -} -else -{ -uint8_t x_444; -lean_dec(x_435); -x_444 = !lean_is_exclusive(x_436); -if (x_444 == 0) +x_428 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; +x_429 = lean_array_push(x_428, x_418); +x_430 = lean_array_push(x_429, x_420); +x_431 = lean_array_push(x_430, x_425); +x_432 = lean_array_push(x_431, x_427); +x_433 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_433, 0, x_423); +lean_ctor_set(x_433, 1, x_12); +lean_ctor_set(x_433, 2, x_432); +x_434 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_413); +if (lean_obj_tag(x_434) == 0) +{ +uint8_t x_435; +x_435 = !lean_is_exclusive(x_434); +if (x_435 == 0) +{ +lean_object* x_436; lean_object* x_437; +x_436 = lean_ctor_get(x_434, 0); +x_437 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_437, 0, x_433); +lean_ctor_set(x_437, 1, x_436); +lean_ctor_set(x_434, 0, x_437); +return x_434; +} +else +{ +lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; +x_438 = lean_ctor_get(x_434, 0); +x_439 = lean_ctor_get(x_434, 1); +lean_inc(x_439); +lean_inc(x_438); +lean_dec(x_434); +x_440 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_440, 0, x_433); +lean_ctor_set(x_440, 1, x_438); +x_441 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_441, 0, x_440); +lean_ctor_set(x_441, 1, x_439); +return x_441; +} +} +else { -return x_436; +uint8_t x_442; +lean_dec(x_433); +x_442 = !lean_is_exclusive(x_434); +if (x_442 == 0) +{ +return x_434; } else { -lean_object* x_445; lean_object* x_446; lean_object* x_447; -x_445 = lean_ctor_get(x_436, 0); -x_446 = lean_ctor_get(x_436, 1); -lean_inc(x_446); -lean_inc(x_445); -lean_dec(x_436); -x_447 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_447, 0, x_445); -lean_ctor_set(x_447, 1, x_446); -return x_447; +lean_object* x_443; lean_object* x_444; lean_object* x_445; +x_443 = lean_ctor_get(x_434, 0); +x_444 = lean_ctor_get(x_434, 1); +lean_inc(x_444); +lean_inc(x_443); +lean_dec(x_434); +x_445 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_445, 0, x_443); +lean_ctor_set(x_445, 1, x_444); +return x_445; } } } else { -lean_object* x_448; lean_object* x_449; +lean_object* x_446; lean_object* x_447; lean_dec(x_2); -x_448 = lean_ctor_get(x_1, 0); -lean_inc(x_448); +x_446 = lean_ctor_get(x_1, 0); +lean_inc(x_446); lean_dec(x_1); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_449 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_448, x_4, x_5, x_6); -if (lean_obj_tag(x_449) == 0) -{ -lean_object* x_450; lean_object* x_451; lean_object* x_452; -x_450 = lean_ctor_get(x_449, 0); -lean_inc(x_450); -x_451 = lean_ctor_get(x_449, 1); -lean_inc(x_451); -lean_dec(x_449); -x_452 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_450, x_4, x_5, x_451); +x_447 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_446, x_4, x_5, x_6); +if (lean_obj_tag(x_447) == 0) +{ +lean_object* x_448; lean_object* x_449; lean_object* x_450; +x_448 = lean_ctor_get(x_447, 0); +lean_inc(x_448); +x_449 = lean_ctor_get(x_447, 1); +lean_inc(x_449); +lean_dec(x_447); +x_450 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_448, x_4, x_5, x_449); lean_dec(x_5); lean_dec(x_4); -return x_452; +return x_450; } else { -uint8_t x_453; +uint8_t x_451; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_453 = !lean_is_exclusive(x_449); -if (x_453 == 0) +x_451 = !lean_is_exclusive(x_447); +if (x_451 == 0) { -return x_449; +return x_447; } else { -lean_object* x_454; lean_object* x_455; lean_object* x_456; -x_454 = lean_ctor_get(x_449, 0); -x_455 = lean_ctor_get(x_449, 1); -lean_inc(x_455); -lean_inc(x_454); -lean_dec(x_449); -x_456 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_456, 0, x_454); -lean_ctor_set(x_456, 1, x_455); -return x_456; +lean_object* x_452; lean_object* x_453; lean_object* x_454; +x_452 = lean_ctor_get(x_447, 0); +x_453 = lean_ctor_get(x_447, 1); +lean_inc(x_453); +lean_inc(x_452); +lean_dec(x_447); +x_454 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_454, 0, x_452); +lean_ctor_set(x_454, 1, x_453); +return x_454; } } } } else { -lean_object* x_457; uint8_t x_458; -x_457 = l_Lean_Syntax_getArg(x_403, x_292); -lean_dec(x_403); -lean_inc(x_457); -x_458 = l_Lean_Syntax_isNodeOf(x_457, x_294, x_292); -if (x_458 == 0) +lean_object* x_455; uint8_t x_456; +x_455 = l_Lean_Syntax_getArg(x_401, x_291); +lean_dec(x_401); +lean_inc(x_455); +x_456 = l_Lean_Syntax_matchesNull(x_455, x_291); +if (x_456 == 0) { -lean_dec(x_457); -lean_dec(x_349); +lean_dec(x_455); lean_dec(x_347); +lean_dec(x_345); lean_dec(x_7); if (lean_obj_tag(x_1) == 0) { -lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; -x_459 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); -x_460 = lean_ctor_get(x_459, 0); -lean_inc(x_460); -x_461 = lean_ctor_get(x_459, 1); +lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; +x_457 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); +x_458 = lean_ctor_get(x_457, 0); +lean_inc(x_458); +x_459 = lean_ctor_get(x_457, 1); +lean_inc(x_459); +lean_dec(x_457); +x_460 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_459); +x_461 = lean_ctor_get(x_460, 0); lean_inc(x_461); -lean_dec(x_459); -x_462 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_461); -x_463 = lean_ctor_get(x_462, 0); -lean_inc(x_463); -x_464 = lean_ctor_get(x_462, 1); +x_462 = lean_ctor_get(x_460, 1); +lean_inc(x_462); +lean_dec(x_460); +x_463 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_462); +x_464 = lean_ctor_get(x_463, 0); lean_inc(x_464); -lean_dec(x_462); -x_465 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_464); -x_466 = lean_ctor_get(x_465, 0); -lean_inc(x_466); -x_467 = lean_ctor_get(x_465, 1); -lean_inc(x_467); -lean_dec(x_465); -x_468 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; -x_469 = l_Lean_addMacroScope(x_466, x_468, x_463); -x_470 = lean_box(0); -x_471 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; -lean_inc(x_460); -x_472 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_472, 0, x_460); +x_465 = lean_ctor_get(x_463, 1); +lean_inc(x_465); +lean_dec(x_463); +x_466 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; +x_467 = l_Lean_addMacroScope(x_464, x_466, x_461); +x_468 = lean_box(0); +x_469 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; +lean_inc(x_458); +x_470 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_470, 0, x_458); +lean_ctor_set(x_470, 1, x_469); +lean_ctor_set(x_470, 2, x_467); +lean_ctor_set(x_470, 3, x_468); +x_471 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; +lean_inc(x_458); +x_472 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_472, 0, x_458); lean_ctor_set(x_472, 1, x_471); -lean_ctor_set(x_472, 2, x_469); -lean_ctor_set(x_472, 3, x_470); -x_473 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; -lean_inc(x_460); -x_474 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_474, 0, x_460); -lean_ctor_set(x_474, 1, x_473); -x_475 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; +x_473 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; lean_inc(x_3); -x_476 = lean_array_push(x_475, x_3); -x_477 = lean_box(2); -x_478 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; -x_479 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_479, 0, x_477); +x_474 = lean_array_push(x_473, x_3); +x_475 = lean_box(2); +x_476 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__3; +x_477 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_477, 0, x_475); +lean_ctor_set(x_477, 1, x_476); +lean_ctor_set(x_477, 2, x_474); +x_478 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__6; +x_479 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_479, 0, x_458); lean_ctor_set(x_479, 1, x_478); -lean_ctor_set(x_479, 2, x_476); -x_480 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__8; -x_481 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_481, 0, x_460); -lean_ctor_set(x_481, 1, x_480); -x_482 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__9; -x_483 = lean_array_push(x_482, x_472); -x_484 = lean_array_push(x_483, x_474); -x_485 = lean_array_push(x_484, x_479); -x_486 = lean_array_push(x_485, x_481); -x_487 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_487, 0, x_477); -lean_ctor_set(x_487, 1, x_12); -lean_ctor_set(x_487, 2, x_486); -x_488 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_467); -if (lean_obj_tag(x_488) == 0) -{ -uint8_t x_489; -x_489 = !lean_is_exclusive(x_488); -if (x_489 == 0) -{ -lean_object* x_490; lean_object* x_491; -x_490 = lean_ctor_get(x_488, 0); -x_491 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_491, 0, x_487); -lean_ctor_set(x_491, 1, x_490); -lean_ctor_set(x_488, 0, x_491); -return x_488; -} -else -{ -lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; -x_492 = lean_ctor_get(x_488, 0); -x_493 = lean_ctor_get(x_488, 1); -lean_inc(x_493); -lean_inc(x_492); -lean_dec(x_488); -x_494 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_494, 0, x_487); -lean_ctor_set(x_494, 1, x_492); -x_495 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_495, 0, x_494); -lean_ctor_set(x_495, 1, x_493); -return x_495; +x_480 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; +x_481 = lean_array_push(x_480, x_470); +x_482 = lean_array_push(x_481, x_472); +x_483 = lean_array_push(x_482, x_477); +x_484 = lean_array_push(x_483, x_479); +x_485 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_485, 0, x_475); +lean_ctor_set(x_485, 1, x_12); +lean_ctor_set(x_485, 2, x_484); +x_486 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_465); +if (lean_obj_tag(x_486) == 0) +{ +uint8_t x_487; +x_487 = !lean_is_exclusive(x_486); +if (x_487 == 0) +{ +lean_object* x_488; lean_object* x_489; +x_488 = lean_ctor_get(x_486, 0); +x_489 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_489, 0, x_485); +lean_ctor_set(x_489, 1, x_488); +lean_ctor_set(x_486, 0, x_489); +return x_486; +} +else +{ +lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; +x_490 = lean_ctor_get(x_486, 0); +x_491 = lean_ctor_get(x_486, 1); +lean_inc(x_491); +lean_inc(x_490); +lean_dec(x_486); +x_492 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_492, 0, x_485); +lean_ctor_set(x_492, 1, x_490); +x_493 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_493, 0, x_492); +lean_ctor_set(x_493, 1, x_491); +return x_493; } } else { -uint8_t x_496; -lean_dec(x_487); -x_496 = !lean_is_exclusive(x_488); -if (x_496 == 0) +uint8_t x_494; +lean_dec(x_485); +x_494 = !lean_is_exclusive(x_486); +if (x_494 == 0) { -return x_488; +return x_486; } else { -lean_object* x_497; lean_object* x_498; lean_object* x_499; -x_497 = lean_ctor_get(x_488, 0); -x_498 = lean_ctor_get(x_488, 1); -lean_inc(x_498); -lean_inc(x_497); -lean_dec(x_488); -x_499 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_499, 0, x_497); -lean_ctor_set(x_499, 1, x_498); -return x_499; +lean_object* x_495; lean_object* x_496; lean_object* x_497; +x_495 = lean_ctor_get(x_486, 0); +x_496 = lean_ctor_get(x_486, 1); +lean_inc(x_496); +lean_inc(x_495); +lean_dec(x_486); +x_497 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_497, 0, x_495); +lean_ctor_set(x_497, 1, x_496); +return x_497; } } } else { -lean_object* x_500; lean_object* x_501; +lean_object* x_498; lean_object* x_499; lean_dec(x_2); -x_500 = lean_ctor_get(x_1, 0); -lean_inc(x_500); +x_498 = lean_ctor_get(x_1, 0); +lean_inc(x_498); lean_dec(x_1); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_501 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_500, x_4, x_5, x_6); -if (lean_obj_tag(x_501) == 0) -{ -lean_object* x_502; lean_object* x_503; lean_object* x_504; -x_502 = lean_ctor_get(x_501, 0); -lean_inc(x_502); -x_503 = lean_ctor_get(x_501, 1); -lean_inc(x_503); -lean_dec(x_501); -x_504 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_502, x_4, x_5, x_503); +x_499 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_498, x_4, x_5, x_6); +if (lean_obj_tag(x_499) == 0) +{ +lean_object* x_500; lean_object* x_501; lean_object* x_502; +x_500 = lean_ctor_get(x_499, 0); +lean_inc(x_500); +x_501 = lean_ctor_get(x_499, 1); +lean_inc(x_501); +lean_dec(x_499); +x_502 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_500, x_4, x_5, x_501); lean_dec(x_5); lean_dec(x_4); -return x_504; +return x_502; } else { -uint8_t x_505; +uint8_t x_503; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_505 = !lean_is_exclusive(x_501); -if (x_505 == 0) +x_503 = !lean_is_exclusive(x_499); +if (x_503 == 0) { -return x_501; +return x_499; } else { -lean_object* x_506; lean_object* x_507; lean_object* x_508; -x_506 = lean_ctor_get(x_501, 0); -x_507 = lean_ctor_get(x_501, 1); -lean_inc(x_507); -lean_inc(x_506); -lean_dec(x_501); -x_508 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_508, 0, x_506); -lean_ctor_set(x_508, 1, x_507); -return x_508; +lean_object* x_504; lean_object* x_505; lean_object* x_506; +x_504 = lean_ctor_get(x_499, 0); +x_505 = lean_ctor_get(x_499, 1); +lean_inc(x_505); +lean_inc(x_504); +lean_dec(x_499); +x_506 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_506, 0, x_504); +lean_ctor_set(x_506, 1, x_505); +return x_506; } } } } else { -lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; -x_509 = l_Lean_Syntax_getArg(x_457, x_346); -lean_dec(x_457); -x_510 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_510, 0, x_509); -x_511 = lean_box(0); +lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; +x_507 = l_Lean_Syntax_getArg(x_455, x_344); +lean_dec(x_455); +x_508 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_508, 0, x_507); +x_509 = lean_box(0); lean_inc(x_3); -x_512 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3(x_3, x_349, x_347, x_2, x_7, x_1, x_3, x_12, x_511, x_510, x_4, x_5, x_6); -lean_dec(x_510); -return x_512; +x_510 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3(x_3, x_347, x_345, x_2, x_7, x_1, x_3, x_12, x_509, x_508, x_4, x_5, x_6); +lean_dec(x_508); +return x_510; } } } else { -lean_object* x_513; lean_object* x_514; lean_object* x_515; -lean_dec(x_403); -x_513 = lean_box(0); -x_514 = lean_box(0); +lean_object* x_511; lean_object* x_512; lean_object* x_513; +lean_dec(x_401); +x_511 = lean_box(0); +x_512 = lean_box(0); lean_inc(x_3); -x_515 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3(x_3, x_349, x_347, x_2, x_7, x_1, x_3, x_12, x_514, x_513, x_4, x_5, x_6); -return x_515; +x_513 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3(x_3, x_347, x_345, x_2, x_7, x_1, x_3, x_12, x_512, x_511, x_4, x_5, x_6); +return x_513; } } } @@ -11292,426 +11364,425 @@ return x_515; } else { -lean_object* x_516; lean_object* x_517; lean_object* x_518; uint8_t x_519; +lean_object* x_514; lean_object* x_515; lean_object* x_516; uint8_t x_517; lean_dec(x_7); -x_516 = lean_unsigned_to_nat(0u); -x_517 = l_Lean_Syntax_getArg(x_3, x_516); +x_514 = lean_unsigned_to_nat(0u); +x_515 = l_Lean_Syntax_getArg(x_3, x_514); +x_516 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__11; +x_517 = l_Lean_Syntax_matchesIdent(x_515, x_516); +if (x_517 == 0) +{ +lean_object* x_518; uint8_t x_519; x_518 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__13; -x_519 = l_Lean_Syntax_matchesIdent(x_517, x_518); +x_519 = l_Lean_Syntax_matchesIdent(x_515, x_518); if (x_519 == 0) { lean_object* x_520; uint8_t x_521; x_520 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__15; -x_521 = l_Lean_Syntax_matchesIdent(x_517, x_520); +x_521 = l_Lean_Syntax_matchesIdent(x_515, x_520); if (x_521 == 0) { lean_object* x_522; uint8_t x_523; x_522 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__17; -x_523 = l_Lean_Syntax_matchesIdent(x_517, x_522); +x_523 = l_Lean_Syntax_matchesIdent(x_515, x_522); if (x_523 == 0) { lean_object* x_524; uint8_t x_525; x_524 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__19; -x_525 = l_Lean_Syntax_matchesIdent(x_517, x_524); +x_525 = l_Lean_Syntax_matchesIdent(x_515, x_524); +lean_dec(x_515); if (x_525 == 0) { -lean_object* x_526; uint8_t x_527; -x_526 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__21; -x_527 = l_Lean_Syntax_matchesIdent(x_517, x_526); -lean_dec(x_517); -if (x_527 == 0) -{ if (lean_obj_tag(x_1) == 0) { -lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; -x_528 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); -x_529 = lean_ctor_get(x_528, 0); -lean_inc(x_529); -x_530 = lean_ctor_get(x_528, 1); +lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; +x_526 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); +x_527 = lean_ctor_get(x_526, 0); +lean_inc(x_527); +x_528 = lean_ctor_get(x_526, 1); +lean_inc(x_528); +lean_dec(x_526); +x_529 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_528); +x_530 = lean_ctor_get(x_529, 0); lean_inc(x_530); -lean_dec(x_528); -x_531 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_530); -x_532 = lean_ctor_get(x_531, 0); -lean_inc(x_532); -x_533 = lean_ctor_get(x_531, 1); +x_531 = lean_ctor_get(x_529, 1); +lean_inc(x_531); +lean_dec(x_529); +x_532 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_531); +x_533 = lean_ctor_get(x_532, 0); lean_inc(x_533); -lean_dec(x_531); -x_534 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_533); -x_535 = lean_ctor_get(x_534, 0); -lean_inc(x_535); -x_536 = lean_ctor_get(x_534, 1); -lean_inc(x_536); -lean_dec(x_534); -x_537 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; -x_538 = l_Lean_addMacroScope(x_535, x_537, x_532); -x_539 = lean_box(0); -x_540 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; -lean_inc(x_529); -x_541 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_541, 0, x_529); +x_534 = lean_ctor_get(x_532, 1); +lean_inc(x_534); +lean_dec(x_532); +x_535 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; +x_536 = l_Lean_addMacroScope(x_533, x_535, x_530); +x_537 = lean_box(0); +x_538 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; +lean_inc(x_527); +x_539 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_539, 0, x_527); +lean_ctor_set(x_539, 1, x_538); +lean_ctor_set(x_539, 2, x_536); +lean_ctor_set(x_539, 3, x_537); +x_540 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; +lean_inc(x_527); +x_541 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_541, 0, x_527); lean_ctor_set(x_541, 1, x_540); -lean_ctor_set(x_541, 2, x_538); -lean_ctor_set(x_541, 3, x_539); -x_542 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; -lean_inc(x_529); -x_543 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_543, 0, x_529); -lean_ctor_set(x_543, 1, x_542); -x_544 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; +x_542 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; lean_inc(x_3); -x_545 = lean_array_push(x_544, x_3); -x_546 = lean_box(2); -x_547 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; -x_548 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_548, 0, x_546); +x_543 = lean_array_push(x_542, x_3); +x_544 = lean_box(2); +x_545 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__3; +x_546 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_546, 0, x_544); +lean_ctor_set(x_546, 1, x_545); +lean_ctor_set(x_546, 2, x_543); +x_547 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__6; +x_548 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_548, 0, x_527); lean_ctor_set(x_548, 1, x_547); -lean_ctor_set(x_548, 2, x_545); -x_549 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__8; -x_550 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_550, 0, x_529); -lean_ctor_set(x_550, 1, x_549); -x_551 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__9; -x_552 = lean_array_push(x_551, x_541); -x_553 = lean_array_push(x_552, x_543); -x_554 = lean_array_push(x_553, x_548); -x_555 = lean_array_push(x_554, x_550); -x_556 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_556, 0, x_546); -lean_ctor_set(x_556, 1, x_12); -lean_ctor_set(x_556, 2, x_555); -x_557 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_536); -if (lean_obj_tag(x_557) == 0) -{ -uint8_t x_558; -x_558 = !lean_is_exclusive(x_557); -if (x_558 == 0) -{ -lean_object* x_559; lean_object* x_560; -x_559 = lean_ctor_get(x_557, 0); -x_560 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_560, 0, x_556); -lean_ctor_set(x_560, 1, x_559); -lean_ctor_set(x_557, 0, x_560); -return x_557; -} -else -{ -lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; -x_561 = lean_ctor_get(x_557, 0); -x_562 = lean_ctor_get(x_557, 1); -lean_inc(x_562); -lean_inc(x_561); -lean_dec(x_557); -x_563 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_563, 0, x_556); -lean_ctor_set(x_563, 1, x_561); -x_564 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_564, 0, x_563); -lean_ctor_set(x_564, 1, x_562); -return x_564; +x_549 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; +x_550 = lean_array_push(x_549, x_539); +x_551 = lean_array_push(x_550, x_541); +x_552 = lean_array_push(x_551, x_546); +x_553 = lean_array_push(x_552, x_548); +x_554 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_554, 0, x_544); +lean_ctor_set(x_554, 1, x_12); +lean_ctor_set(x_554, 2, x_553); +x_555 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_534); +if (lean_obj_tag(x_555) == 0) +{ +uint8_t x_556; +x_556 = !lean_is_exclusive(x_555); +if (x_556 == 0) +{ +lean_object* x_557; lean_object* x_558; +x_557 = lean_ctor_get(x_555, 0); +x_558 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_558, 0, x_554); +lean_ctor_set(x_558, 1, x_557); +lean_ctor_set(x_555, 0, x_558); +return x_555; +} +else +{ +lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; +x_559 = lean_ctor_get(x_555, 0); +x_560 = lean_ctor_get(x_555, 1); +lean_inc(x_560); +lean_inc(x_559); +lean_dec(x_555); +x_561 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_561, 0, x_554); +lean_ctor_set(x_561, 1, x_559); +x_562 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_562, 0, x_561); +lean_ctor_set(x_562, 1, x_560); +return x_562; } } else { -uint8_t x_565; -lean_dec(x_556); -x_565 = !lean_is_exclusive(x_557); -if (x_565 == 0) +uint8_t x_563; +lean_dec(x_554); +x_563 = !lean_is_exclusive(x_555); +if (x_563 == 0) { -return x_557; +return x_555; } else { -lean_object* x_566; lean_object* x_567; lean_object* x_568; -x_566 = lean_ctor_get(x_557, 0); -x_567 = lean_ctor_get(x_557, 1); -lean_inc(x_567); -lean_inc(x_566); -lean_dec(x_557); -x_568 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_568, 0, x_566); -lean_ctor_set(x_568, 1, x_567); -return x_568; +lean_object* x_564; lean_object* x_565; lean_object* x_566; +x_564 = lean_ctor_get(x_555, 0); +x_565 = lean_ctor_get(x_555, 1); +lean_inc(x_565); +lean_inc(x_564); +lean_dec(x_555); +x_566 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_566, 0, x_564); +lean_ctor_set(x_566, 1, x_565); +return x_566; } } } else { -lean_object* x_569; lean_object* x_570; +lean_object* x_567; lean_object* x_568; lean_dec(x_2); -x_569 = lean_ctor_get(x_1, 0); -lean_inc(x_569); +x_567 = lean_ctor_get(x_1, 0); +lean_inc(x_567); lean_dec(x_1); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_570 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_569, x_4, x_5, x_6); -if (lean_obj_tag(x_570) == 0) -{ -lean_object* x_571; lean_object* x_572; lean_object* x_573; -x_571 = lean_ctor_get(x_570, 0); -lean_inc(x_571); -x_572 = lean_ctor_get(x_570, 1); -lean_inc(x_572); -lean_dec(x_570); -x_573 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_571, x_4, x_5, x_572); +x_568 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_567, x_4, x_5, x_6); +if (lean_obj_tag(x_568) == 0) +{ +lean_object* x_569; lean_object* x_570; lean_object* x_571; +x_569 = lean_ctor_get(x_568, 0); +lean_inc(x_569); +x_570 = lean_ctor_get(x_568, 1); +lean_inc(x_570); +lean_dec(x_568); +x_571 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_569, x_4, x_5, x_570); lean_dec(x_5); lean_dec(x_4); -return x_573; +return x_571; } else { -uint8_t x_574; +uint8_t x_572; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_574 = !lean_is_exclusive(x_570); -if (x_574 == 0) +x_572 = !lean_is_exclusive(x_568); +if (x_572 == 0) { -return x_570; +return x_568; } else { -lean_object* x_575; lean_object* x_576; lean_object* x_577; -x_575 = lean_ctor_get(x_570, 0); -x_576 = lean_ctor_get(x_570, 1); -lean_inc(x_576); -lean_inc(x_575); -lean_dec(x_570); -x_577 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_577, 0, x_575); -lean_ctor_set(x_577, 1, x_576); -return x_577; +lean_object* x_573; lean_object* x_574; lean_object* x_575; +x_573 = lean_ctor_get(x_568, 0); +x_574 = lean_ctor_get(x_568, 1); +lean_inc(x_574); +lean_inc(x_573); +lean_dec(x_568); +x_575 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_575, 0, x_573); +lean_ctor_set(x_575, 1, x_574); +return x_575; } } } } else { -lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; uint8_t x_582; -x_578 = lean_unsigned_to_nat(2u); -x_579 = l_Lean_Syntax_getArg(x_3, x_578); -x_580 = l_Lean_nullKind; -x_581 = lean_unsigned_to_nat(1u); -lean_inc(x_579); -x_582 = l_Lean_Syntax_isNodeOf(x_579, x_580, x_581); -if (x_582 == 0) +lean_object* x_576; lean_object* x_577; lean_object* x_578; uint8_t x_579; +x_576 = lean_unsigned_to_nat(2u); +x_577 = l_Lean_Syntax_getArg(x_3, x_576); +x_578 = lean_unsigned_to_nat(1u); +lean_inc(x_577); +x_579 = l_Lean_Syntax_matchesNull(x_577, x_578); +if (x_579 == 0) { -lean_dec(x_579); +lean_dec(x_577); if (lean_obj_tag(x_1) == 0) { -lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; -x_583 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); +lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; +x_580 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); +x_581 = lean_ctor_get(x_580, 0); +lean_inc(x_581); +x_582 = lean_ctor_get(x_580, 1); +lean_inc(x_582); +lean_dec(x_580); +x_583 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_582); x_584 = lean_ctor_get(x_583, 0); lean_inc(x_584); x_585 = lean_ctor_get(x_583, 1); lean_inc(x_585); lean_dec(x_583); -x_586 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_585); +x_586 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_585); x_587 = lean_ctor_get(x_586, 0); lean_inc(x_587); x_588 = lean_ctor_get(x_586, 1); lean_inc(x_588); lean_dec(x_586); -x_589 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_588); -x_590 = lean_ctor_get(x_589, 0); -lean_inc(x_590); -x_591 = lean_ctor_get(x_589, 1); -lean_inc(x_591); -lean_dec(x_589); -x_592 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; -x_593 = l_Lean_addMacroScope(x_590, x_592, x_587); -x_594 = lean_box(0); -x_595 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; -lean_inc(x_584); -x_596 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_596, 0, x_584); -lean_ctor_set(x_596, 1, x_595); -lean_ctor_set(x_596, 2, x_593); -lean_ctor_set(x_596, 3, x_594); -x_597 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; -lean_inc(x_584); -x_598 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_598, 0, x_584); -lean_ctor_set(x_598, 1, x_597); -x_599 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; +x_589 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; +x_590 = l_Lean_addMacroScope(x_587, x_589, x_584); +x_591 = lean_box(0); +x_592 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; +lean_inc(x_581); +x_593 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_593, 0, x_581); +lean_ctor_set(x_593, 1, x_592); +lean_ctor_set(x_593, 2, x_590); +lean_ctor_set(x_593, 3, x_591); +x_594 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; +lean_inc(x_581); +x_595 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_595, 0, x_581); +lean_ctor_set(x_595, 1, x_594); +x_596 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; lean_inc(x_3); -x_600 = lean_array_push(x_599, x_3); -x_601 = lean_box(2); -x_602 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; -x_603 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_603, 0, x_601); -lean_ctor_set(x_603, 1, x_602); -lean_ctor_set(x_603, 2, x_600); -x_604 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__8; -x_605 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_605, 0, x_584); -lean_ctor_set(x_605, 1, x_604); -x_606 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__9; -x_607 = lean_array_push(x_606, x_596); -x_608 = lean_array_push(x_607, x_598); -x_609 = lean_array_push(x_608, x_603); -x_610 = lean_array_push(x_609, x_605); -x_611 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_611, 0, x_601); -lean_ctor_set(x_611, 1, x_12); -lean_ctor_set(x_611, 2, x_610); -x_612 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_591); -if (lean_obj_tag(x_612) == 0) -{ -uint8_t x_613; -x_613 = !lean_is_exclusive(x_612); -if (x_613 == 0) -{ -lean_object* x_614; lean_object* x_615; -x_614 = lean_ctor_get(x_612, 0); +x_597 = lean_array_push(x_596, x_3); +x_598 = lean_box(2); +x_599 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__3; +x_600 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_600, 0, x_598); +lean_ctor_set(x_600, 1, x_599); +lean_ctor_set(x_600, 2, x_597); +x_601 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__6; +x_602 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_602, 0, x_581); +lean_ctor_set(x_602, 1, x_601); +x_603 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; +x_604 = lean_array_push(x_603, x_593); +x_605 = lean_array_push(x_604, x_595); +x_606 = lean_array_push(x_605, x_600); +x_607 = lean_array_push(x_606, x_602); +x_608 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_608, 0, x_598); +lean_ctor_set(x_608, 1, x_12); +lean_ctor_set(x_608, 2, x_607); +x_609 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_588); +if (lean_obj_tag(x_609) == 0) +{ +uint8_t x_610; +x_610 = !lean_is_exclusive(x_609); +if (x_610 == 0) +{ +lean_object* x_611; lean_object* x_612; +x_611 = lean_ctor_get(x_609, 0); +x_612 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_612, 0, x_608); +lean_ctor_set(x_612, 1, x_611); +lean_ctor_set(x_609, 0, x_612); +return x_609; +} +else +{ +lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; +x_613 = lean_ctor_get(x_609, 0); +x_614 = lean_ctor_get(x_609, 1); +lean_inc(x_614); +lean_inc(x_613); +lean_dec(x_609); x_615 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_615, 0, x_611); -lean_ctor_set(x_615, 1, x_614); -lean_ctor_set(x_612, 0, x_615); -return x_612; -} -else -{ -lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; -x_616 = lean_ctor_get(x_612, 0); -x_617 = lean_ctor_get(x_612, 1); -lean_inc(x_617); -lean_inc(x_616); -lean_dec(x_612); -x_618 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_618, 0, x_611); -lean_ctor_set(x_618, 1, x_616); -x_619 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_619, 0, x_618); -lean_ctor_set(x_619, 1, x_617); -return x_619; +lean_ctor_set(x_615, 0, x_608); +lean_ctor_set(x_615, 1, x_613); +x_616 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_616, 0, x_615); +lean_ctor_set(x_616, 1, x_614); +return x_616; } } else { -uint8_t x_620; -lean_dec(x_611); -x_620 = !lean_is_exclusive(x_612); -if (x_620 == 0) +uint8_t x_617; +lean_dec(x_608); +x_617 = !lean_is_exclusive(x_609); +if (x_617 == 0) { -return x_612; +return x_609; } else { -lean_object* x_621; lean_object* x_622; lean_object* x_623; -x_621 = lean_ctor_get(x_612, 0); -x_622 = lean_ctor_get(x_612, 1); -lean_inc(x_622); -lean_inc(x_621); -lean_dec(x_612); -x_623 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_623, 0, x_621); -lean_ctor_set(x_623, 1, x_622); -return x_623; +lean_object* x_618; lean_object* x_619; lean_object* x_620; +x_618 = lean_ctor_get(x_609, 0); +x_619 = lean_ctor_get(x_609, 1); +lean_inc(x_619); +lean_inc(x_618); +lean_dec(x_609); +x_620 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_620, 0, x_618); +lean_ctor_set(x_620, 1, x_619); +return x_620; } } } else { -lean_object* x_624; lean_object* x_625; +lean_object* x_621; lean_object* x_622; lean_dec(x_2); -x_624 = lean_ctor_get(x_1, 0); -lean_inc(x_624); +x_621 = lean_ctor_get(x_1, 0); +lean_inc(x_621); lean_dec(x_1); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_625 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_624, x_4, x_5, x_6); -if (lean_obj_tag(x_625) == 0) +x_622 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_621, x_4, x_5, x_6); +if (lean_obj_tag(x_622) == 0) { -lean_object* x_626; lean_object* x_627; lean_object* x_628; -x_626 = lean_ctor_get(x_625, 0); -lean_inc(x_626); -x_627 = lean_ctor_get(x_625, 1); -lean_inc(x_627); -lean_dec(x_625); -x_628 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_626, x_4, x_5, x_627); +lean_object* x_623; lean_object* x_624; lean_object* x_625; +x_623 = lean_ctor_get(x_622, 0); +lean_inc(x_623); +x_624 = lean_ctor_get(x_622, 1); +lean_inc(x_624); +lean_dec(x_622); +x_625 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_623, x_4, x_5, x_624); lean_dec(x_5); lean_dec(x_4); -return x_628; +return x_625; } else { -uint8_t x_629; +uint8_t x_626; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_629 = !lean_is_exclusive(x_625); -if (x_629 == 0) +x_626 = !lean_is_exclusive(x_622); +if (x_626 == 0) { -return x_625; +return x_622; } else { -lean_object* x_630; lean_object* x_631; lean_object* x_632; -x_630 = lean_ctor_get(x_625, 0); -x_631 = lean_ctor_get(x_625, 1); -lean_inc(x_631); -lean_inc(x_630); -lean_dec(x_625); -x_632 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_632, 0, x_630); -lean_ctor_set(x_632, 1, x_631); -return x_632; +lean_object* x_627; lean_object* x_628; lean_object* x_629; +x_627 = lean_ctor_get(x_622, 0); +x_628 = lean_ctor_get(x_622, 1); +lean_inc(x_628); +lean_inc(x_627); +lean_dec(x_622); +x_629 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_629, 0, x_627); +lean_ctor_set(x_629, 1, x_628); +return x_629; } } } } else { -lean_object* x_633; lean_object* x_634; +lean_object* x_630; lean_object* x_631; lean_dec(x_3); -x_633 = l_Lean_Syntax_getArg(x_579, x_516); -lean_dec(x_579); -x_634 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat(x_1, x_2, x_633, x_4, x_5, x_6); -if (lean_obj_tag(x_634) == 0) +x_630 = l_Lean_Syntax_getArg(x_577, x_514); +lean_dec(x_577); +x_631 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat(x_1, x_2, x_630, x_4, x_5, x_6); +if (lean_obj_tag(x_631) == 0) { -uint8_t x_635; -x_635 = !lean_is_exclusive(x_634); -if (x_635 == 0) +uint8_t x_632; +x_632 = !lean_is_exclusive(x_631); +if (x_632 == 0) { -return x_634; +return x_631; } else { -lean_object* x_636; lean_object* x_637; lean_object* x_638; -x_636 = lean_ctor_get(x_634, 0); -x_637 = lean_ctor_get(x_634, 1); -lean_inc(x_637); -lean_inc(x_636); -lean_dec(x_634); -x_638 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_638, 0, x_636); -lean_ctor_set(x_638, 1, x_637); -return x_638; +lean_object* x_633; lean_object* x_634; lean_object* x_635; +x_633 = lean_ctor_get(x_631, 0); +x_634 = lean_ctor_get(x_631, 1); +lean_inc(x_634); +lean_inc(x_633); +lean_dec(x_631); +x_635 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_635, 0, x_633); +lean_ctor_set(x_635, 1, x_634); +return x_635; } } else { -uint8_t x_639; -x_639 = !lean_is_exclusive(x_634); -if (x_639 == 0) +uint8_t x_636; +x_636 = !lean_is_exclusive(x_631); +if (x_636 == 0) { -return x_634; +return x_631; } else { -lean_object* x_640; lean_object* x_641; lean_object* x_642; -x_640 = lean_ctor_get(x_634, 0); -x_641 = lean_ctor_get(x_634, 1); -lean_inc(x_641); -lean_inc(x_640); -lean_dec(x_634); -x_642 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_642, 0, x_640); -lean_ctor_set(x_642, 1, x_641); -return x_642; +lean_object* x_637; lean_object* x_638; lean_object* x_639; +x_637 = lean_ctor_get(x_631, 0); +x_638 = lean_ctor_get(x_631, 1); +lean_inc(x_638); +lean_inc(x_637); +lean_dec(x_631); +x_639 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_639, 0, x_637); +lean_ctor_set(x_639, 1, x_638); +return x_639; } } } @@ -11719,718 +11790,717 @@ return x_642; } else { -lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; uint8_t x_647; -lean_dec(x_517); -x_643 = lean_unsigned_to_nat(2u); -x_644 = l_Lean_Syntax_getArg(x_3, x_643); -x_645 = l_Lean_nullKind; -x_646 = lean_unsigned_to_nat(1u); -lean_inc(x_644); -x_647 = l_Lean_Syntax_isNodeOf(x_644, x_645, x_646); -if (x_647 == 0) +lean_object* x_640; lean_object* x_641; lean_object* x_642; uint8_t x_643; +lean_dec(x_515); +x_640 = lean_unsigned_to_nat(2u); +x_641 = l_Lean_Syntax_getArg(x_3, x_640); +x_642 = lean_unsigned_to_nat(1u); +lean_inc(x_641); +x_643 = l_Lean_Syntax_matchesNull(x_641, x_642); +if (x_643 == 0) { -lean_dec(x_644); +lean_dec(x_641); if (lean_obj_tag(x_1) == 0) { -lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; lean_object* x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; -x_648 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); -x_649 = lean_ctor_get(x_648, 0); +lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; +x_644 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); +x_645 = lean_ctor_get(x_644, 0); +lean_inc(x_645); +x_646 = lean_ctor_get(x_644, 1); +lean_inc(x_646); +lean_dec(x_644); +x_647 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_646); +x_648 = lean_ctor_get(x_647, 0); +lean_inc(x_648); +x_649 = lean_ctor_get(x_647, 1); lean_inc(x_649); -x_650 = lean_ctor_get(x_648, 1); -lean_inc(x_650); -lean_dec(x_648); -x_651 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_650); -x_652 = lean_ctor_get(x_651, 0); +lean_dec(x_647); +x_650 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_649); +x_651 = lean_ctor_get(x_650, 0); +lean_inc(x_651); +x_652 = lean_ctor_get(x_650, 1); lean_inc(x_652); -x_653 = lean_ctor_get(x_651, 1); -lean_inc(x_653); -lean_dec(x_651); -x_654 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_653); -x_655 = lean_ctor_get(x_654, 0); -lean_inc(x_655); -x_656 = lean_ctor_get(x_654, 1); -lean_inc(x_656); -lean_dec(x_654); -x_657 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; -x_658 = l_Lean_addMacroScope(x_655, x_657, x_652); -x_659 = lean_box(0); -x_660 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; -lean_inc(x_649); -x_661 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_661, 0, x_649); -lean_ctor_set(x_661, 1, x_660); -lean_ctor_set(x_661, 2, x_658); -lean_ctor_set(x_661, 3, x_659); -x_662 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; -lean_inc(x_649); -x_663 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_663, 0, x_649); -lean_ctor_set(x_663, 1, x_662); -x_664 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; +lean_dec(x_650); +x_653 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; +x_654 = l_Lean_addMacroScope(x_651, x_653, x_648); +x_655 = lean_box(0); +x_656 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; +lean_inc(x_645); +x_657 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_657, 0, x_645); +lean_ctor_set(x_657, 1, x_656); +lean_ctor_set(x_657, 2, x_654); +lean_ctor_set(x_657, 3, x_655); +x_658 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; +lean_inc(x_645); +x_659 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_659, 0, x_645); +lean_ctor_set(x_659, 1, x_658); +x_660 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; lean_inc(x_3); -x_665 = lean_array_push(x_664, x_3); -x_666 = lean_box(2); +x_661 = lean_array_push(x_660, x_3); +x_662 = lean_box(2); +x_663 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__3; +x_664 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_664, 0, x_662); +lean_ctor_set(x_664, 1, x_663); +lean_ctor_set(x_664, 2, x_661); +x_665 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__6; +x_666 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_666, 0, x_645); +lean_ctor_set(x_666, 1, x_665); x_667 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; -x_668 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_668, 0, x_666); -lean_ctor_set(x_668, 1, x_667); -lean_ctor_set(x_668, 2, x_665); -x_669 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__8; -x_670 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_670, 0, x_649); -lean_ctor_set(x_670, 1, x_669); -x_671 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__9; -x_672 = lean_array_push(x_671, x_661); -x_673 = lean_array_push(x_672, x_663); -x_674 = lean_array_push(x_673, x_668); -x_675 = lean_array_push(x_674, x_670); -x_676 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_676, 0, x_666); -lean_ctor_set(x_676, 1, x_12); -lean_ctor_set(x_676, 2, x_675); -x_677 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_656); -if (lean_obj_tag(x_677) == 0) -{ -uint8_t x_678; -x_678 = !lean_is_exclusive(x_677); -if (x_678 == 0) -{ -lean_object* x_679; lean_object* x_680; -x_679 = lean_ctor_get(x_677, 0); -x_680 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_680, 0, x_676); -lean_ctor_set(x_680, 1, x_679); -lean_ctor_set(x_677, 0, x_680); -return x_677; +x_668 = lean_array_push(x_667, x_657); +x_669 = lean_array_push(x_668, x_659); +x_670 = lean_array_push(x_669, x_664); +x_671 = lean_array_push(x_670, x_666); +x_672 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_672, 0, x_662); +lean_ctor_set(x_672, 1, x_12); +lean_ctor_set(x_672, 2, x_671); +x_673 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_652); +if (lean_obj_tag(x_673) == 0) +{ +uint8_t x_674; +x_674 = !lean_is_exclusive(x_673); +if (x_674 == 0) +{ +lean_object* x_675; lean_object* x_676; +x_675 = lean_ctor_get(x_673, 0); +x_676 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_676, 0, x_672); +lean_ctor_set(x_676, 1, x_675); +lean_ctor_set(x_673, 0, x_676); +return x_673; } else { -lean_object* x_681; lean_object* x_682; lean_object* x_683; lean_object* x_684; -x_681 = lean_ctor_get(x_677, 0); -x_682 = lean_ctor_get(x_677, 1); -lean_inc(x_682); -lean_inc(x_681); -lean_dec(x_677); -x_683 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_683, 0, x_676); -lean_ctor_set(x_683, 1, x_681); -x_684 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_684, 0, x_683); -lean_ctor_set(x_684, 1, x_682); -return x_684; +lean_object* x_677; lean_object* x_678; lean_object* x_679; lean_object* x_680; +x_677 = lean_ctor_get(x_673, 0); +x_678 = lean_ctor_get(x_673, 1); +lean_inc(x_678); +lean_inc(x_677); +lean_dec(x_673); +x_679 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_679, 0, x_672); +lean_ctor_set(x_679, 1, x_677); +x_680 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_680, 0, x_679); +lean_ctor_set(x_680, 1, x_678); +return x_680; } } else { -uint8_t x_685; -lean_dec(x_676); -x_685 = !lean_is_exclusive(x_677); -if (x_685 == 0) +uint8_t x_681; +lean_dec(x_672); +x_681 = !lean_is_exclusive(x_673); +if (x_681 == 0) { -return x_677; +return x_673; } else { -lean_object* x_686; lean_object* x_687; lean_object* x_688; -x_686 = lean_ctor_get(x_677, 0); -x_687 = lean_ctor_get(x_677, 1); -lean_inc(x_687); -lean_inc(x_686); -lean_dec(x_677); -x_688 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_688, 0, x_686); -lean_ctor_set(x_688, 1, x_687); -return x_688; +lean_object* x_682; lean_object* x_683; lean_object* x_684; +x_682 = lean_ctor_get(x_673, 0); +x_683 = lean_ctor_get(x_673, 1); +lean_inc(x_683); +lean_inc(x_682); +lean_dec(x_673); +x_684 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_684, 0, x_682); +lean_ctor_set(x_684, 1, x_683); +return x_684; } } } else { -lean_object* x_689; lean_object* x_690; +lean_object* x_685; lean_object* x_686; lean_dec(x_2); -x_689 = lean_ctor_get(x_1, 0); -lean_inc(x_689); +x_685 = lean_ctor_get(x_1, 0); +lean_inc(x_685); lean_dec(x_1); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_690 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_689, x_4, x_5, x_6); -if (lean_obj_tag(x_690) == 0) +x_686 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_685, x_4, x_5, x_6); +if (lean_obj_tag(x_686) == 0) { -lean_object* x_691; lean_object* x_692; lean_object* x_693; -x_691 = lean_ctor_get(x_690, 0); -lean_inc(x_691); -x_692 = lean_ctor_get(x_690, 1); -lean_inc(x_692); -lean_dec(x_690); -x_693 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_691, x_4, x_5, x_692); +lean_object* x_687; lean_object* x_688; lean_object* x_689; +x_687 = lean_ctor_get(x_686, 0); +lean_inc(x_687); +x_688 = lean_ctor_get(x_686, 1); +lean_inc(x_688); +lean_dec(x_686); +x_689 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_687, x_4, x_5, x_688); lean_dec(x_5); lean_dec(x_4); -return x_693; +return x_689; } else { -uint8_t x_694; +uint8_t x_690; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_694 = !lean_is_exclusive(x_690); -if (x_694 == 0) +x_690 = !lean_is_exclusive(x_686); +if (x_690 == 0) { -return x_690; +return x_686; } else { -lean_object* x_695; lean_object* x_696; lean_object* x_697; -x_695 = lean_ctor_get(x_690, 0); -x_696 = lean_ctor_get(x_690, 1); -lean_inc(x_696); -lean_inc(x_695); -lean_dec(x_690); -x_697 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_697, 0, x_695); -lean_ctor_set(x_697, 1, x_696); -return x_697; +lean_object* x_691; lean_object* x_692; lean_object* x_693; +x_691 = lean_ctor_get(x_686, 0); +x_692 = lean_ctor_get(x_686, 1); +lean_inc(x_692); +lean_inc(x_691); +lean_dec(x_686); +x_693 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_693, 0, x_691); +lean_ctor_set(x_693, 1, x_692); +return x_693; } } } } else { -lean_object* x_698; lean_object* x_699; uint8_t x_700; -x_698 = l_Lean_Syntax_getArg(x_644, x_516); -lean_dec(x_644); -x_699 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___closed__8; -lean_inc(x_698); -x_700 = l_Lean_Syntax_isOfKind(x_698, x_699); -if (x_700 == 0) +lean_object* x_694; lean_object* x_695; uint8_t x_696; +x_694 = l_Lean_Syntax_getArg(x_641, x_514); +lean_dec(x_641); +x_695 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___closed__8; +lean_inc(x_694); +x_696 = l_Lean_Syntax_isOfKind(x_694, x_695); +if (x_696 == 0) { -lean_dec(x_698); +lean_dec(x_694); if (lean_obj_tag(x_1) == 0) { -lean_object* x_701; lean_object* x_702; lean_object* x_703; lean_object* x_704; lean_object* x_705; lean_object* x_706; lean_object* x_707; lean_object* x_708; lean_object* x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; lean_object* x_718; lean_object* x_719; lean_object* x_720; lean_object* x_721; lean_object* x_722; lean_object* x_723; lean_object* x_724; lean_object* x_725; lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; lean_object* x_730; -x_701 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); -x_702 = lean_ctor_get(x_701, 0); +lean_object* x_697; lean_object* x_698; lean_object* x_699; lean_object* x_700; lean_object* x_701; lean_object* x_702; lean_object* x_703; lean_object* x_704; lean_object* x_705; lean_object* x_706; lean_object* x_707; lean_object* x_708; lean_object* x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; lean_object* x_718; lean_object* x_719; lean_object* x_720; lean_object* x_721; lean_object* x_722; lean_object* x_723; lean_object* x_724; lean_object* x_725; lean_object* x_726; +x_697 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); +x_698 = lean_ctor_get(x_697, 0); +lean_inc(x_698); +x_699 = lean_ctor_get(x_697, 1); +lean_inc(x_699); +lean_dec(x_697); +x_700 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_699); +x_701 = lean_ctor_get(x_700, 0); +lean_inc(x_701); +x_702 = lean_ctor_get(x_700, 1); lean_inc(x_702); -x_703 = lean_ctor_get(x_701, 1); -lean_inc(x_703); -lean_dec(x_701); -x_704 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_703); -x_705 = lean_ctor_get(x_704, 0); +lean_dec(x_700); +x_703 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_702); +x_704 = lean_ctor_get(x_703, 0); +lean_inc(x_704); +x_705 = lean_ctor_get(x_703, 1); lean_inc(x_705); -x_706 = lean_ctor_get(x_704, 1); -lean_inc(x_706); -lean_dec(x_704); -x_707 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_706); -x_708 = lean_ctor_get(x_707, 0); -lean_inc(x_708); -x_709 = lean_ctor_get(x_707, 1); -lean_inc(x_709); -lean_dec(x_707); -x_710 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; -x_711 = l_Lean_addMacroScope(x_708, x_710, x_705); -x_712 = lean_box(0); -x_713 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; -lean_inc(x_702); -x_714 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_714, 0, x_702); -lean_ctor_set(x_714, 1, x_713); -lean_ctor_set(x_714, 2, x_711); -lean_ctor_set(x_714, 3, x_712); -x_715 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; -lean_inc(x_702); -x_716 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_716, 0, x_702); -lean_ctor_set(x_716, 1, x_715); -x_717 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; +lean_dec(x_703); +x_706 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; +x_707 = l_Lean_addMacroScope(x_704, x_706, x_701); +x_708 = lean_box(0); +x_709 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; +lean_inc(x_698); +x_710 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_710, 0, x_698); +lean_ctor_set(x_710, 1, x_709); +lean_ctor_set(x_710, 2, x_707); +lean_ctor_set(x_710, 3, x_708); +x_711 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; +lean_inc(x_698); +x_712 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_712, 0, x_698); +lean_ctor_set(x_712, 1, x_711); +x_713 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; lean_inc(x_3); -x_718 = lean_array_push(x_717, x_3); -x_719 = lean_box(2); +x_714 = lean_array_push(x_713, x_3); +x_715 = lean_box(2); +x_716 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__3; +x_717 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_717, 0, x_715); +lean_ctor_set(x_717, 1, x_716); +lean_ctor_set(x_717, 2, x_714); +x_718 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__6; +x_719 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_719, 0, x_698); +lean_ctor_set(x_719, 1, x_718); x_720 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; -x_721 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_721, 0, x_719); -lean_ctor_set(x_721, 1, x_720); -lean_ctor_set(x_721, 2, x_718); -x_722 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__8; -x_723 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_723, 0, x_702); -lean_ctor_set(x_723, 1, x_722); -x_724 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__9; -x_725 = lean_array_push(x_724, x_714); -x_726 = lean_array_push(x_725, x_716); -x_727 = lean_array_push(x_726, x_721); -x_728 = lean_array_push(x_727, x_723); -x_729 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_729, 0, x_719); -lean_ctor_set(x_729, 1, x_12); -lean_ctor_set(x_729, 2, x_728); -x_730 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_709); -if (lean_obj_tag(x_730) == 0) -{ -uint8_t x_731; -x_731 = !lean_is_exclusive(x_730); -if (x_731 == 0) -{ -lean_object* x_732; lean_object* x_733; -x_732 = lean_ctor_get(x_730, 0); +x_721 = lean_array_push(x_720, x_710); +x_722 = lean_array_push(x_721, x_712); +x_723 = lean_array_push(x_722, x_717); +x_724 = lean_array_push(x_723, x_719); +x_725 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_725, 0, x_715); +lean_ctor_set(x_725, 1, x_12); +lean_ctor_set(x_725, 2, x_724); +x_726 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_705); +if (lean_obj_tag(x_726) == 0) +{ +uint8_t x_727; +x_727 = !lean_is_exclusive(x_726); +if (x_727 == 0) +{ +lean_object* x_728; lean_object* x_729; +x_728 = lean_ctor_get(x_726, 0); +x_729 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_729, 0, x_725); +lean_ctor_set(x_729, 1, x_728); +lean_ctor_set(x_726, 0, x_729); +return x_726; +} +else +{ +lean_object* x_730; lean_object* x_731; lean_object* x_732; lean_object* x_733; +x_730 = lean_ctor_get(x_726, 0); +x_731 = lean_ctor_get(x_726, 1); +lean_inc(x_731); +lean_inc(x_730); +lean_dec(x_726); +x_732 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_732, 0, x_725); +lean_ctor_set(x_732, 1, x_730); x_733 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_733, 0, x_729); -lean_ctor_set(x_733, 1, x_732); -lean_ctor_set(x_730, 0, x_733); -return x_730; -} -else -{ -lean_object* x_734; lean_object* x_735; lean_object* x_736; lean_object* x_737; -x_734 = lean_ctor_get(x_730, 0); -x_735 = lean_ctor_get(x_730, 1); -lean_inc(x_735); -lean_inc(x_734); -lean_dec(x_730); -x_736 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_736, 0, x_729); -lean_ctor_set(x_736, 1, x_734); -x_737 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_737, 0, x_736); -lean_ctor_set(x_737, 1, x_735); -return x_737; +lean_ctor_set(x_733, 0, x_732); +lean_ctor_set(x_733, 1, x_731); +return x_733; } } else { -uint8_t x_738; -lean_dec(x_729); -x_738 = !lean_is_exclusive(x_730); -if (x_738 == 0) +uint8_t x_734; +lean_dec(x_725); +x_734 = !lean_is_exclusive(x_726); +if (x_734 == 0) { -return x_730; +return x_726; } else { -lean_object* x_739; lean_object* x_740; lean_object* x_741; -x_739 = lean_ctor_get(x_730, 0); -x_740 = lean_ctor_get(x_730, 1); -lean_inc(x_740); -lean_inc(x_739); -lean_dec(x_730); -x_741 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_741, 0, x_739); -lean_ctor_set(x_741, 1, x_740); -return x_741; +lean_object* x_735; lean_object* x_736; lean_object* x_737; +x_735 = lean_ctor_get(x_726, 0); +x_736 = lean_ctor_get(x_726, 1); +lean_inc(x_736); +lean_inc(x_735); +lean_dec(x_726); +x_737 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_737, 0, x_735); +lean_ctor_set(x_737, 1, x_736); +return x_737; } } } else { -lean_object* x_742; lean_object* x_743; +lean_object* x_738; lean_object* x_739; lean_dec(x_2); -x_742 = lean_ctor_get(x_1, 0); -lean_inc(x_742); +x_738 = lean_ctor_get(x_1, 0); +lean_inc(x_738); lean_dec(x_1); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_743 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_742, x_4, x_5, x_6); -if (lean_obj_tag(x_743) == 0) +x_739 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_738, x_4, x_5, x_6); +if (lean_obj_tag(x_739) == 0) { -lean_object* x_744; lean_object* x_745; lean_object* x_746; -x_744 = lean_ctor_get(x_743, 0); -lean_inc(x_744); -x_745 = lean_ctor_get(x_743, 1); -lean_inc(x_745); -lean_dec(x_743); -x_746 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_744, x_4, x_5, x_745); +lean_object* x_740; lean_object* x_741; lean_object* x_742; +x_740 = lean_ctor_get(x_739, 0); +lean_inc(x_740); +x_741 = lean_ctor_get(x_739, 1); +lean_inc(x_741); +lean_dec(x_739); +x_742 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_740, x_4, x_5, x_741); lean_dec(x_5); lean_dec(x_4); -return x_746; +return x_742; } else { -uint8_t x_747; +uint8_t x_743; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_747 = !lean_is_exclusive(x_743); -if (x_747 == 0) +x_743 = !lean_is_exclusive(x_739); +if (x_743 == 0) { -return x_743; +return x_739; } else { -lean_object* x_748; lean_object* x_749; lean_object* x_750; -x_748 = lean_ctor_get(x_743, 0); -x_749 = lean_ctor_get(x_743, 1); -lean_inc(x_749); -lean_inc(x_748); -lean_dec(x_743); -x_750 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_750, 0, x_748); -lean_ctor_set(x_750, 1, x_749); -return x_750; +lean_object* x_744; lean_object* x_745; lean_object* x_746; +x_744 = lean_ctor_get(x_739, 0); +x_745 = lean_ctor_get(x_739, 1); +lean_inc(x_745); +lean_inc(x_744); +lean_dec(x_739); +x_746 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_746, 0, x_744); +lean_ctor_set(x_746, 1, x_745); +return x_746; } } } } else { -lean_object* x_751; lean_object* x_752; uint8_t x_753; -x_751 = l_Lean_Syntax_getArg(x_698, x_516); -x_752 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__23; -x_753 = l_Lean_Syntax_matchesIdent(x_751, x_752); -lean_dec(x_751); -if (x_753 == 0) +lean_object* x_747; lean_object* x_748; uint8_t x_749; +x_747 = l_Lean_Syntax_getArg(x_694, x_514); +x_748 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__21; +x_749 = l_Lean_Syntax_matchesIdent(x_747, x_748); +lean_dec(x_747); +if (x_749 == 0) { -lean_dec(x_698); +lean_dec(x_694); if (lean_obj_tag(x_1) == 0) { -lean_object* x_754; lean_object* x_755; lean_object* x_756; lean_object* x_757; lean_object* x_758; lean_object* x_759; lean_object* x_760; lean_object* x_761; lean_object* x_762; lean_object* x_763; lean_object* x_764; lean_object* x_765; lean_object* x_766; lean_object* x_767; lean_object* x_768; lean_object* x_769; lean_object* x_770; lean_object* x_771; lean_object* x_772; lean_object* x_773; lean_object* x_774; lean_object* x_775; lean_object* x_776; lean_object* x_777; lean_object* x_778; lean_object* x_779; lean_object* x_780; lean_object* x_781; lean_object* x_782; lean_object* x_783; -x_754 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); -x_755 = lean_ctor_get(x_754, 0); +lean_object* x_750; lean_object* x_751; lean_object* x_752; lean_object* x_753; lean_object* x_754; lean_object* x_755; lean_object* x_756; lean_object* x_757; lean_object* x_758; lean_object* x_759; lean_object* x_760; lean_object* x_761; lean_object* x_762; lean_object* x_763; lean_object* x_764; lean_object* x_765; lean_object* x_766; lean_object* x_767; lean_object* x_768; lean_object* x_769; lean_object* x_770; lean_object* x_771; lean_object* x_772; lean_object* x_773; lean_object* x_774; lean_object* x_775; lean_object* x_776; lean_object* x_777; lean_object* x_778; lean_object* x_779; +x_750 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); +x_751 = lean_ctor_get(x_750, 0); +lean_inc(x_751); +x_752 = lean_ctor_get(x_750, 1); +lean_inc(x_752); +lean_dec(x_750); +x_753 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_752); +x_754 = lean_ctor_get(x_753, 0); +lean_inc(x_754); +x_755 = lean_ctor_get(x_753, 1); lean_inc(x_755); -x_756 = lean_ctor_get(x_754, 1); -lean_inc(x_756); -lean_dec(x_754); -x_757 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_756); -x_758 = lean_ctor_get(x_757, 0); +lean_dec(x_753); +x_756 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_755); +x_757 = lean_ctor_get(x_756, 0); +lean_inc(x_757); +x_758 = lean_ctor_get(x_756, 1); lean_inc(x_758); -x_759 = lean_ctor_get(x_757, 1); -lean_inc(x_759); -lean_dec(x_757); -x_760 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_759); -x_761 = lean_ctor_get(x_760, 0); -lean_inc(x_761); -x_762 = lean_ctor_get(x_760, 1); -lean_inc(x_762); -lean_dec(x_760); -x_763 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; -x_764 = l_Lean_addMacroScope(x_761, x_763, x_758); -x_765 = lean_box(0); -x_766 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; -lean_inc(x_755); -x_767 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_767, 0, x_755); -lean_ctor_set(x_767, 1, x_766); -lean_ctor_set(x_767, 2, x_764); -lean_ctor_set(x_767, 3, x_765); -x_768 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; -lean_inc(x_755); -x_769 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_769, 0, x_755); -lean_ctor_set(x_769, 1, x_768); -x_770 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; +lean_dec(x_756); +x_759 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; +x_760 = l_Lean_addMacroScope(x_757, x_759, x_754); +x_761 = lean_box(0); +x_762 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; +lean_inc(x_751); +x_763 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_763, 0, x_751); +lean_ctor_set(x_763, 1, x_762); +lean_ctor_set(x_763, 2, x_760); +lean_ctor_set(x_763, 3, x_761); +x_764 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; +lean_inc(x_751); +x_765 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_765, 0, x_751); +lean_ctor_set(x_765, 1, x_764); +x_766 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; lean_inc(x_3); -x_771 = lean_array_push(x_770, x_3); -x_772 = lean_box(2); +x_767 = lean_array_push(x_766, x_3); +x_768 = lean_box(2); +x_769 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__3; +x_770 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_770, 0, x_768); +lean_ctor_set(x_770, 1, x_769); +lean_ctor_set(x_770, 2, x_767); +x_771 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__6; +x_772 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_772, 0, x_751); +lean_ctor_set(x_772, 1, x_771); x_773 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; -x_774 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_774, 0, x_772); -lean_ctor_set(x_774, 1, x_773); -lean_ctor_set(x_774, 2, x_771); -x_775 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__8; -x_776 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_776, 0, x_755); -lean_ctor_set(x_776, 1, x_775); -x_777 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__9; -x_778 = lean_array_push(x_777, x_767); -x_779 = lean_array_push(x_778, x_769); -x_780 = lean_array_push(x_779, x_774); -x_781 = lean_array_push(x_780, x_776); -x_782 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_782, 0, x_772); -lean_ctor_set(x_782, 1, x_12); -lean_ctor_set(x_782, 2, x_781); -x_783 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_762); -if (lean_obj_tag(x_783) == 0) -{ -uint8_t x_784; -x_784 = !lean_is_exclusive(x_783); -if (x_784 == 0) -{ -lean_object* x_785; lean_object* x_786; -x_785 = lean_ctor_get(x_783, 0); +x_774 = lean_array_push(x_773, x_763); +x_775 = lean_array_push(x_774, x_765); +x_776 = lean_array_push(x_775, x_770); +x_777 = lean_array_push(x_776, x_772); +x_778 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_778, 0, x_768); +lean_ctor_set(x_778, 1, x_12); +lean_ctor_set(x_778, 2, x_777); +x_779 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_758); +if (lean_obj_tag(x_779) == 0) +{ +uint8_t x_780; +x_780 = !lean_is_exclusive(x_779); +if (x_780 == 0) +{ +lean_object* x_781; lean_object* x_782; +x_781 = lean_ctor_get(x_779, 0); +x_782 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_782, 0, x_778); +lean_ctor_set(x_782, 1, x_781); +lean_ctor_set(x_779, 0, x_782); +return x_779; +} +else +{ +lean_object* x_783; lean_object* x_784; lean_object* x_785; lean_object* x_786; +x_783 = lean_ctor_get(x_779, 0); +x_784 = lean_ctor_get(x_779, 1); +lean_inc(x_784); +lean_inc(x_783); +lean_dec(x_779); +x_785 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_785, 0, x_778); +lean_ctor_set(x_785, 1, x_783); x_786 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_786, 0, x_782); -lean_ctor_set(x_786, 1, x_785); -lean_ctor_set(x_783, 0, x_786); -return x_783; -} -else -{ -lean_object* x_787; lean_object* x_788; lean_object* x_789; lean_object* x_790; -x_787 = lean_ctor_get(x_783, 0); -x_788 = lean_ctor_get(x_783, 1); -lean_inc(x_788); -lean_inc(x_787); -lean_dec(x_783); -x_789 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_789, 0, x_782); -lean_ctor_set(x_789, 1, x_787); -x_790 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_790, 0, x_789); -lean_ctor_set(x_790, 1, x_788); -return x_790; +lean_ctor_set(x_786, 0, x_785); +lean_ctor_set(x_786, 1, x_784); +return x_786; } } else { -uint8_t x_791; -lean_dec(x_782); -x_791 = !lean_is_exclusive(x_783); -if (x_791 == 0) +uint8_t x_787; +lean_dec(x_778); +x_787 = !lean_is_exclusive(x_779); +if (x_787 == 0) { -return x_783; +return x_779; } else { -lean_object* x_792; lean_object* x_793; lean_object* x_794; -x_792 = lean_ctor_get(x_783, 0); -x_793 = lean_ctor_get(x_783, 1); -lean_inc(x_793); -lean_inc(x_792); -lean_dec(x_783); -x_794 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_794, 0, x_792); -lean_ctor_set(x_794, 1, x_793); -return x_794; +lean_object* x_788; lean_object* x_789; lean_object* x_790; +x_788 = lean_ctor_get(x_779, 0); +x_789 = lean_ctor_get(x_779, 1); +lean_inc(x_789); +lean_inc(x_788); +lean_dec(x_779); +x_790 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_790, 0, x_788); +lean_ctor_set(x_790, 1, x_789); +return x_790; } } } else { -lean_object* x_795; lean_object* x_796; +lean_object* x_791; lean_object* x_792; lean_dec(x_2); -x_795 = lean_ctor_get(x_1, 0); -lean_inc(x_795); +x_791 = lean_ctor_get(x_1, 0); +lean_inc(x_791); lean_dec(x_1); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_796 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_795, x_4, x_5, x_6); -if (lean_obj_tag(x_796) == 0) +x_792 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_791, x_4, x_5, x_6); +if (lean_obj_tag(x_792) == 0) { -lean_object* x_797; lean_object* x_798; lean_object* x_799; -x_797 = lean_ctor_get(x_796, 0); -lean_inc(x_797); -x_798 = lean_ctor_get(x_796, 1); -lean_inc(x_798); -lean_dec(x_796); -x_799 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_797, x_4, x_5, x_798); +lean_object* x_793; lean_object* x_794; lean_object* x_795; +x_793 = lean_ctor_get(x_792, 0); +lean_inc(x_793); +x_794 = lean_ctor_get(x_792, 1); +lean_inc(x_794); +lean_dec(x_792); +x_795 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_793, x_4, x_5, x_794); lean_dec(x_5); lean_dec(x_4); -return x_799; +return x_795; } else { -uint8_t x_800; +uint8_t x_796; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_800 = !lean_is_exclusive(x_796); -if (x_800 == 0) +x_796 = !lean_is_exclusive(x_792); +if (x_796 == 0) { -return x_796; +return x_792; } else { -lean_object* x_801; lean_object* x_802; lean_object* x_803; -x_801 = lean_ctor_get(x_796, 0); -x_802 = lean_ctor_get(x_796, 1); -lean_inc(x_802); -lean_inc(x_801); -lean_dec(x_796); -x_803 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_803, 0, x_801); -lean_ctor_set(x_803, 1, x_802); -return x_803; +lean_object* x_797; lean_object* x_798; lean_object* x_799; +x_797 = lean_ctor_get(x_792, 0); +x_798 = lean_ctor_get(x_792, 1); +lean_inc(x_798); +lean_inc(x_797); +lean_dec(x_792); +x_799 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_799, 0, x_797); +lean_ctor_set(x_799, 1, x_798); +return x_799; } } } } else { -lean_object* x_804; uint8_t x_805; -x_804 = l_Lean_Syntax_getArg(x_698, x_646); -lean_dec(x_698); -x_805 = l_Lean_Syntax_isNodeOf(x_804, x_645, x_516); -if (x_805 == 0) +lean_object* x_800; uint8_t x_801; +x_800 = l_Lean_Syntax_getArg(x_694, x_642); +lean_dec(x_694); +x_801 = l_Lean_Syntax_matchesNull(x_800, x_514); +if (x_801 == 0) { if (lean_obj_tag(x_1) == 0) { -lean_object* x_806; lean_object* x_807; lean_object* x_808; lean_object* x_809; lean_object* x_810; lean_object* x_811; lean_object* x_812; lean_object* x_813; lean_object* x_814; lean_object* x_815; lean_object* x_816; lean_object* x_817; lean_object* x_818; lean_object* x_819; lean_object* x_820; lean_object* x_821; lean_object* x_822; lean_object* x_823; lean_object* x_824; lean_object* x_825; lean_object* x_826; lean_object* x_827; lean_object* x_828; lean_object* x_829; lean_object* x_830; lean_object* x_831; lean_object* x_832; lean_object* x_833; lean_object* x_834; lean_object* x_835; -x_806 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); -x_807 = lean_ctor_get(x_806, 0); -lean_inc(x_807); -x_808 = lean_ctor_get(x_806, 1); -lean_inc(x_808); -lean_dec(x_806); -x_809 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_808); -x_810 = lean_ctor_get(x_809, 0); -lean_inc(x_810); -x_811 = lean_ctor_get(x_809, 1); -lean_inc(x_811); -lean_dec(x_809); -x_812 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_811); -x_813 = lean_ctor_get(x_812, 0); -lean_inc(x_813); -x_814 = lean_ctor_get(x_812, 1); -lean_inc(x_814); -lean_dec(x_812); -x_815 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; -x_816 = l_Lean_addMacroScope(x_813, x_815, x_810); -x_817 = lean_box(0); -x_818 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; -lean_inc(x_807); -x_819 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_819, 0, x_807); -lean_ctor_set(x_819, 1, x_818); -lean_ctor_set(x_819, 2, x_816); -lean_ctor_set(x_819, 3, x_817); -x_820 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; +lean_object* x_802; lean_object* x_803; lean_object* x_804; lean_object* x_805; lean_object* x_806; lean_object* x_807; lean_object* x_808; lean_object* x_809; lean_object* x_810; lean_object* x_811; lean_object* x_812; lean_object* x_813; lean_object* x_814; lean_object* x_815; lean_object* x_816; lean_object* x_817; lean_object* x_818; lean_object* x_819; lean_object* x_820; lean_object* x_821; lean_object* x_822; lean_object* x_823; lean_object* x_824; lean_object* x_825; lean_object* x_826; lean_object* x_827; lean_object* x_828; lean_object* x_829; lean_object* x_830; lean_object* x_831; +x_802 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); +x_803 = lean_ctor_get(x_802, 0); +lean_inc(x_803); +x_804 = lean_ctor_get(x_802, 1); +lean_inc(x_804); +lean_dec(x_802); +x_805 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_804); +x_806 = lean_ctor_get(x_805, 0); +lean_inc(x_806); +x_807 = lean_ctor_get(x_805, 1); lean_inc(x_807); -x_821 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_821, 0, x_807); -lean_ctor_set(x_821, 1, x_820); -x_822 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; +lean_dec(x_805); +x_808 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_807); +x_809 = lean_ctor_get(x_808, 0); +lean_inc(x_809); +x_810 = lean_ctor_get(x_808, 1); +lean_inc(x_810); +lean_dec(x_808); +x_811 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; +x_812 = l_Lean_addMacroScope(x_809, x_811, x_806); +x_813 = lean_box(0); +x_814 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; +lean_inc(x_803); +x_815 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_815, 0, x_803); +lean_ctor_set(x_815, 1, x_814); +lean_ctor_set(x_815, 2, x_812); +lean_ctor_set(x_815, 3, x_813); +x_816 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; +lean_inc(x_803); +x_817 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_817, 0, x_803); +lean_ctor_set(x_817, 1, x_816); +x_818 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; lean_inc(x_3); -x_823 = lean_array_push(x_822, x_3); -x_824 = lean_box(2); +x_819 = lean_array_push(x_818, x_3); +x_820 = lean_box(2); +x_821 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__3; +x_822 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_822, 0, x_820); +lean_ctor_set(x_822, 1, x_821); +lean_ctor_set(x_822, 2, x_819); +x_823 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__6; +x_824 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_824, 0, x_803); +lean_ctor_set(x_824, 1, x_823); x_825 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; -x_826 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_826, 0, x_824); -lean_ctor_set(x_826, 1, x_825); -lean_ctor_set(x_826, 2, x_823); -x_827 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__8; -x_828 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_828, 0, x_807); -lean_ctor_set(x_828, 1, x_827); -x_829 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__9; -x_830 = lean_array_push(x_829, x_819); -x_831 = lean_array_push(x_830, x_821); -x_832 = lean_array_push(x_831, x_826); -x_833 = lean_array_push(x_832, x_828); -x_834 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_834, 0, x_824); -lean_ctor_set(x_834, 1, x_12); -lean_ctor_set(x_834, 2, x_833); -x_835 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_814); -if (lean_obj_tag(x_835) == 0) -{ -uint8_t x_836; -x_836 = !lean_is_exclusive(x_835); -if (x_836 == 0) -{ -lean_object* x_837; lean_object* x_838; -x_837 = lean_ctor_get(x_835, 0); +x_826 = lean_array_push(x_825, x_815); +x_827 = lean_array_push(x_826, x_817); +x_828 = lean_array_push(x_827, x_822); +x_829 = lean_array_push(x_828, x_824); +x_830 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_830, 0, x_820); +lean_ctor_set(x_830, 1, x_12); +lean_ctor_set(x_830, 2, x_829); +x_831 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_810); +if (lean_obj_tag(x_831) == 0) +{ +uint8_t x_832; +x_832 = !lean_is_exclusive(x_831); +if (x_832 == 0) +{ +lean_object* x_833; lean_object* x_834; +x_833 = lean_ctor_get(x_831, 0); +x_834 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_834, 0, x_830); +lean_ctor_set(x_834, 1, x_833); +lean_ctor_set(x_831, 0, x_834); +return x_831; +} +else +{ +lean_object* x_835; lean_object* x_836; lean_object* x_837; lean_object* x_838; +x_835 = lean_ctor_get(x_831, 0); +x_836 = lean_ctor_get(x_831, 1); +lean_inc(x_836); +lean_inc(x_835); +lean_dec(x_831); +x_837 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_837, 0, x_830); +lean_ctor_set(x_837, 1, x_835); x_838 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_838, 0, x_834); -lean_ctor_set(x_838, 1, x_837); -lean_ctor_set(x_835, 0, x_838); -return x_835; -} -else -{ -lean_object* x_839; lean_object* x_840; lean_object* x_841; lean_object* x_842; -x_839 = lean_ctor_get(x_835, 0); -x_840 = lean_ctor_get(x_835, 1); -lean_inc(x_840); -lean_inc(x_839); -lean_dec(x_835); -x_841 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_841, 0, x_834); -lean_ctor_set(x_841, 1, x_839); -x_842 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_842, 0, x_841); -lean_ctor_set(x_842, 1, x_840); -return x_842; +lean_ctor_set(x_838, 0, x_837); +lean_ctor_set(x_838, 1, x_836); +return x_838; } } else { -uint8_t x_843; -lean_dec(x_834); -x_843 = !lean_is_exclusive(x_835); -if (x_843 == 0) +uint8_t x_839; +lean_dec(x_830); +x_839 = !lean_is_exclusive(x_831); +if (x_839 == 0) { -return x_835; +return x_831; } else { -lean_object* x_844; lean_object* x_845; lean_object* x_846; -x_844 = lean_ctor_get(x_835, 0); -x_845 = lean_ctor_get(x_835, 1); -lean_inc(x_845); -lean_inc(x_844); -lean_dec(x_835); -x_846 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_846, 0, x_844); -lean_ctor_set(x_846, 1, x_845); -return x_846; +lean_object* x_840; lean_object* x_841; lean_object* x_842; +x_840 = lean_ctor_get(x_831, 0); +x_841 = lean_ctor_get(x_831, 1); +lean_inc(x_841); +lean_inc(x_840); +lean_dec(x_831); +x_842 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_842, 0, x_840); +lean_ctor_set(x_842, 1, x_841); +return x_842; } } } else { -lean_object* x_847; lean_object* x_848; +lean_object* x_843; lean_object* x_844; lean_dec(x_2); -x_847 = lean_ctor_get(x_1, 0); -lean_inc(x_847); +x_843 = lean_ctor_get(x_1, 0); +lean_inc(x_843); lean_dec(x_1); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_848 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_847, x_4, x_5, x_6); -if (lean_obj_tag(x_848) == 0) +x_844 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_843, x_4, x_5, x_6); +if (lean_obj_tag(x_844) == 0) { -lean_object* x_849; lean_object* x_850; lean_object* x_851; -x_849 = lean_ctor_get(x_848, 0); -lean_inc(x_849); -x_850 = lean_ctor_get(x_848, 1); -lean_inc(x_850); -lean_dec(x_848); -x_851 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_849, x_4, x_5, x_850); +lean_object* x_845; lean_object* x_846; lean_object* x_847; +x_845 = lean_ctor_get(x_844, 0); +lean_inc(x_845); +x_846 = lean_ctor_get(x_844, 1); +lean_inc(x_846); +lean_dec(x_844); +x_847 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_845, x_4, x_5, x_846); lean_dec(x_5); lean_dec(x_4); -return x_851; +return x_847; } else { -uint8_t x_852; +uint8_t x_848; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_852 = !lean_is_exclusive(x_848); -if (x_852 == 0) +x_848 = !lean_is_exclusive(x_844); +if (x_848 == 0) { -return x_848; +return x_844; } else { -lean_object* x_853; lean_object* x_854; lean_object* x_855; -x_853 = lean_ctor_get(x_848, 0); -x_854 = lean_ctor_get(x_848, 1); -lean_inc(x_854); -lean_inc(x_853); -lean_dec(x_848); -x_855 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_855, 0, x_853); -lean_ctor_set(x_855, 1, x_854); -return x_855; +lean_object* x_849; lean_object* x_850; lean_object* x_851; +x_849 = lean_ctor_get(x_844, 0); +x_850 = lean_ctor_get(x_844, 1); +lean_inc(x_850); +lean_inc(x_849); +lean_dec(x_844); +x_851 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_851, 0, x_849); +lean_ctor_set(x_851, 1, x_850); +return x_851; } } } } else { -lean_object* x_856; lean_object* x_857; uint8_t x_858; lean_object* x_859; lean_object* x_860; +lean_object* x_852; lean_object* x_853; uint8_t x_854; lean_object* x_855; lean_object* x_856; lean_dec(x_1); -x_856 = lean_box(0); -x_857 = l_Lean_interpolatedStrKind; -x_858 = 0; -x_859 = l_Lean_Syntax_mkAntiquotNode(x_857, x_2, x_516, x_856, x_858); -x_860 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_859, x_4, x_5, x_6); +x_852 = lean_box(0); +x_853 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__23; +x_854 = 0; +x_855 = l_Lean_Syntax_mkAntiquotNode(x_853, x_2, x_514, x_852, x_854); +x_856 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_855, x_4, x_5, x_6); lean_dec(x_5); lean_dec(x_4); -return x_860; +return x_856; } } } @@ -12439,228 +12509,227 @@ return x_860; } else { -lean_object* x_861; lean_object* x_862; lean_object* x_863; lean_object* x_864; uint8_t x_865; -lean_dec(x_517); -x_861 = lean_unsigned_to_nat(2u); -x_862 = l_Lean_Syntax_getArg(x_3, x_861); -x_863 = l_Lean_nullKind; -x_864 = lean_unsigned_to_nat(1u); -lean_inc(x_862); -x_865 = l_Lean_Syntax_isNodeOf(x_862, x_863, x_864); -if (x_865 == 0) +lean_object* x_857; lean_object* x_858; lean_object* x_859; uint8_t x_860; +lean_dec(x_515); +x_857 = lean_unsigned_to_nat(2u); +x_858 = l_Lean_Syntax_getArg(x_3, x_857); +x_859 = lean_unsigned_to_nat(1u); +lean_inc(x_858); +x_860 = l_Lean_Syntax_matchesNull(x_858, x_859); +if (x_860 == 0) { -lean_dec(x_862); +lean_dec(x_858); if (lean_obj_tag(x_1) == 0) { -lean_object* x_866; lean_object* x_867; lean_object* x_868; lean_object* x_869; lean_object* x_870; lean_object* x_871; lean_object* x_872; lean_object* x_873; lean_object* x_874; lean_object* x_875; lean_object* x_876; lean_object* x_877; lean_object* x_878; lean_object* x_879; lean_object* x_880; lean_object* x_881; lean_object* x_882; lean_object* x_883; lean_object* x_884; lean_object* x_885; lean_object* x_886; lean_object* x_887; lean_object* x_888; lean_object* x_889; lean_object* x_890; lean_object* x_891; lean_object* x_892; lean_object* x_893; lean_object* x_894; lean_object* x_895; -x_866 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); -x_867 = lean_ctor_get(x_866, 0); -lean_inc(x_867); -x_868 = lean_ctor_get(x_866, 1); +lean_object* x_861; lean_object* x_862; lean_object* x_863; lean_object* x_864; lean_object* x_865; lean_object* x_866; lean_object* x_867; lean_object* x_868; lean_object* x_869; lean_object* x_870; lean_object* x_871; lean_object* x_872; lean_object* x_873; lean_object* x_874; lean_object* x_875; lean_object* x_876; lean_object* x_877; lean_object* x_878; lean_object* x_879; lean_object* x_880; lean_object* x_881; lean_object* x_882; lean_object* x_883; lean_object* x_884; lean_object* x_885; lean_object* x_886; lean_object* x_887; lean_object* x_888; lean_object* x_889; lean_object* x_890; +x_861 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); +x_862 = lean_ctor_get(x_861, 0); +lean_inc(x_862); +x_863 = lean_ctor_get(x_861, 1); +lean_inc(x_863); +lean_dec(x_861); +x_864 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_863); +x_865 = lean_ctor_get(x_864, 0); +lean_inc(x_865); +x_866 = lean_ctor_get(x_864, 1); +lean_inc(x_866); +lean_dec(x_864); +x_867 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_866); +x_868 = lean_ctor_get(x_867, 0); lean_inc(x_868); -lean_dec(x_866); -x_869 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_868); -x_870 = lean_ctor_get(x_869, 0); -lean_inc(x_870); -x_871 = lean_ctor_get(x_869, 1); -lean_inc(x_871); -lean_dec(x_869); -x_872 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_871); -x_873 = lean_ctor_get(x_872, 0); -lean_inc(x_873); -x_874 = lean_ctor_get(x_872, 1); -lean_inc(x_874); -lean_dec(x_872); -x_875 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; -x_876 = l_Lean_addMacroScope(x_873, x_875, x_870); -x_877 = lean_box(0); -x_878 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; -lean_inc(x_867); -x_879 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_879, 0, x_867); -lean_ctor_set(x_879, 1, x_878); -lean_ctor_set(x_879, 2, x_876); -lean_ctor_set(x_879, 3, x_877); -x_880 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; -lean_inc(x_867); -x_881 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_881, 0, x_867); -lean_ctor_set(x_881, 1, x_880); -x_882 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; +x_869 = lean_ctor_get(x_867, 1); +lean_inc(x_869); +lean_dec(x_867); +x_870 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; +x_871 = l_Lean_addMacroScope(x_868, x_870, x_865); +x_872 = lean_box(0); +x_873 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; +lean_inc(x_862); +x_874 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_874, 0, x_862); +lean_ctor_set(x_874, 1, x_873); +lean_ctor_set(x_874, 2, x_871); +lean_ctor_set(x_874, 3, x_872); +x_875 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; +lean_inc(x_862); +x_876 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_876, 0, x_862); +lean_ctor_set(x_876, 1, x_875); +x_877 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; lean_inc(x_3); -x_883 = lean_array_push(x_882, x_3); -x_884 = lean_box(2); -x_885 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; -x_886 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_886, 0, x_884); -lean_ctor_set(x_886, 1, x_885); -lean_ctor_set(x_886, 2, x_883); -x_887 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__8; -x_888 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_888, 0, x_867); -lean_ctor_set(x_888, 1, x_887); -x_889 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__9; -x_890 = lean_array_push(x_889, x_879); -x_891 = lean_array_push(x_890, x_881); -x_892 = lean_array_push(x_891, x_886); -x_893 = lean_array_push(x_892, x_888); -x_894 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_894, 0, x_884); -lean_ctor_set(x_894, 1, x_12); -lean_ctor_set(x_894, 2, x_893); -x_895 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_874); -if (lean_obj_tag(x_895) == 0) -{ -uint8_t x_896; -x_896 = !lean_is_exclusive(x_895); -if (x_896 == 0) -{ -lean_object* x_897; lean_object* x_898; -x_897 = lean_ctor_get(x_895, 0); -x_898 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_898, 0, x_894); -lean_ctor_set(x_898, 1, x_897); -lean_ctor_set(x_895, 0, x_898); -return x_895; -} -else -{ -lean_object* x_899; lean_object* x_900; lean_object* x_901; lean_object* x_902; -x_899 = lean_ctor_get(x_895, 0); -x_900 = lean_ctor_get(x_895, 1); -lean_inc(x_900); -lean_inc(x_899); -lean_dec(x_895); -x_901 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_901, 0, x_894); -lean_ctor_set(x_901, 1, x_899); -x_902 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_902, 0, x_901); -lean_ctor_set(x_902, 1, x_900); -return x_902; +x_878 = lean_array_push(x_877, x_3); +x_879 = lean_box(2); +x_880 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__3; +x_881 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_881, 0, x_879); +lean_ctor_set(x_881, 1, x_880); +lean_ctor_set(x_881, 2, x_878); +x_882 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__6; +x_883 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_883, 0, x_862); +lean_ctor_set(x_883, 1, x_882); +x_884 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; +x_885 = lean_array_push(x_884, x_874); +x_886 = lean_array_push(x_885, x_876); +x_887 = lean_array_push(x_886, x_881); +x_888 = lean_array_push(x_887, x_883); +x_889 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_889, 0, x_879); +lean_ctor_set(x_889, 1, x_12); +lean_ctor_set(x_889, 2, x_888); +x_890 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_869); +if (lean_obj_tag(x_890) == 0) +{ +uint8_t x_891; +x_891 = !lean_is_exclusive(x_890); +if (x_891 == 0) +{ +lean_object* x_892; lean_object* x_893; +x_892 = lean_ctor_get(x_890, 0); +x_893 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_893, 0, x_889); +lean_ctor_set(x_893, 1, x_892); +lean_ctor_set(x_890, 0, x_893); +return x_890; +} +else +{ +lean_object* x_894; lean_object* x_895; lean_object* x_896; lean_object* x_897; +x_894 = lean_ctor_get(x_890, 0); +x_895 = lean_ctor_get(x_890, 1); +lean_inc(x_895); +lean_inc(x_894); +lean_dec(x_890); +x_896 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_896, 0, x_889); +lean_ctor_set(x_896, 1, x_894); +x_897 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_897, 0, x_896); +lean_ctor_set(x_897, 1, x_895); +return x_897; } } else { -uint8_t x_903; -lean_dec(x_894); -x_903 = !lean_is_exclusive(x_895); -if (x_903 == 0) +uint8_t x_898; +lean_dec(x_889); +x_898 = !lean_is_exclusive(x_890); +if (x_898 == 0) { -return x_895; +return x_890; } else { -lean_object* x_904; lean_object* x_905; lean_object* x_906; -x_904 = lean_ctor_get(x_895, 0); -x_905 = lean_ctor_get(x_895, 1); -lean_inc(x_905); -lean_inc(x_904); -lean_dec(x_895); -x_906 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_906, 0, x_904); -lean_ctor_set(x_906, 1, x_905); -return x_906; +lean_object* x_899; lean_object* x_900; lean_object* x_901; +x_899 = lean_ctor_get(x_890, 0); +x_900 = lean_ctor_get(x_890, 1); +lean_inc(x_900); +lean_inc(x_899); +lean_dec(x_890); +x_901 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_901, 0, x_899); +lean_ctor_set(x_901, 1, x_900); +return x_901; } } } else { -lean_object* x_907; lean_object* x_908; +lean_object* x_902; lean_object* x_903; lean_dec(x_2); -x_907 = lean_ctor_get(x_1, 0); -lean_inc(x_907); +x_902 = lean_ctor_get(x_1, 0); +lean_inc(x_902); lean_dec(x_1); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_908 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_907, x_4, x_5, x_6); -if (lean_obj_tag(x_908) == 0) +x_903 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_902, x_4, x_5, x_6); +if (lean_obj_tag(x_903) == 0) { -lean_object* x_909; lean_object* x_910; lean_object* x_911; -x_909 = lean_ctor_get(x_908, 0); -lean_inc(x_909); -x_910 = lean_ctor_get(x_908, 1); -lean_inc(x_910); -lean_dec(x_908); -x_911 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_909, x_4, x_5, x_910); +lean_object* x_904; lean_object* x_905; lean_object* x_906; +x_904 = lean_ctor_get(x_903, 0); +lean_inc(x_904); +x_905 = lean_ctor_get(x_903, 1); +lean_inc(x_905); +lean_dec(x_903); +x_906 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_904, x_4, x_5, x_905); lean_dec(x_5); lean_dec(x_4); -return x_911; +return x_906; } else { -uint8_t x_912; +uint8_t x_907; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_912 = !lean_is_exclusive(x_908); -if (x_912 == 0) +x_907 = !lean_is_exclusive(x_903); +if (x_907 == 0) { -return x_908; +return x_903; } else { -lean_object* x_913; lean_object* x_914; lean_object* x_915; -x_913 = lean_ctor_get(x_908, 0); -x_914 = lean_ctor_get(x_908, 1); -lean_inc(x_914); -lean_inc(x_913); -lean_dec(x_908); -x_915 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_915, 0, x_913); -lean_ctor_set(x_915, 1, x_914); -return x_915; +lean_object* x_908; lean_object* x_909; lean_object* x_910; +x_908 = lean_ctor_get(x_903, 0); +x_909 = lean_ctor_get(x_903, 1); +lean_inc(x_909); +lean_inc(x_908); +lean_dec(x_903); +x_910 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_910, 0, x_908); +lean_ctor_set(x_910, 1, x_909); +return x_910; } } } } else { -lean_object* x_916; lean_object* x_917; lean_object* x_918; +lean_object* x_911; lean_object* x_912; lean_object* x_913; lean_dec(x_1); -x_916 = l_Lean_Syntax_getArg(x_862, x_516); -lean_dec(x_862); -x_917 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__2___closed__7; +x_911 = l_Lean_Syntax_getArg(x_858, x_514); +lean_dec(x_858); +x_912 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__2___closed__7; lean_inc(x_5); lean_inc(x_4); -x_918 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat(x_520, x_916, x_2, x_917, x_4, x_5, x_6); -if (lean_obj_tag(x_918) == 0) +x_913 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat(x_518, x_911, x_2, x_912, x_4, x_5, x_6); +if (lean_obj_tag(x_913) == 0) { -lean_object* x_919; lean_object* x_920; lean_object* x_921; -x_919 = lean_ctor_get(x_918, 0); -lean_inc(x_919); -x_920 = lean_ctor_get(x_918, 1); -lean_inc(x_920); -lean_dec(x_918); -x_921 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_919, x_4, x_5, x_920); +lean_object* x_914; lean_object* x_915; lean_object* x_916; +x_914 = lean_ctor_get(x_913, 0); +lean_inc(x_914); +x_915 = lean_ctor_get(x_913, 1); +lean_inc(x_915); +lean_dec(x_913); +x_916 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_914, x_4, x_5, x_915); lean_dec(x_5); lean_dec(x_4); -return x_921; +return x_916; } else { -uint8_t x_922; +uint8_t x_917; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_922 = !lean_is_exclusive(x_918); -if (x_922 == 0) +x_917 = !lean_is_exclusive(x_913); +if (x_917 == 0) { -return x_918; +return x_913; } else { -lean_object* x_923; lean_object* x_924; lean_object* x_925; -x_923 = lean_ctor_get(x_918, 0); -x_924 = lean_ctor_get(x_918, 1); -lean_inc(x_924); -lean_inc(x_923); -lean_dec(x_918); -x_925 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_925, 0, x_923); -lean_ctor_set(x_925, 1, x_924); -return x_925; +lean_object* x_918; lean_object* x_919; lean_object* x_920; +x_918 = lean_ctor_get(x_913, 0); +x_919 = lean_ctor_get(x_913, 1); +lean_inc(x_919); +lean_inc(x_918); +lean_dec(x_913); +x_920 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_920, 0, x_918); +lean_ctor_set(x_920, 1, x_919); +return x_920; } } } @@ -12668,228 +12737,227 @@ return x_925; } else { -lean_object* x_926; lean_object* x_927; lean_object* x_928; lean_object* x_929; uint8_t x_930; -lean_dec(x_517); -x_926 = lean_unsigned_to_nat(2u); -x_927 = l_Lean_Syntax_getArg(x_3, x_926); -x_928 = l_Lean_nullKind; -x_929 = lean_unsigned_to_nat(1u); -lean_inc(x_927); -x_930 = l_Lean_Syntax_isNodeOf(x_927, x_928, x_929); -if (x_930 == 0) +lean_object* x_921; lean_object* x_922; lean_object* x_923; uint8_t x_924; +lean_dec(x_515); +x_921 = lean_unsigned_to_nat(2u); +x_922 = l_Lean_Syntax_getArg(x_3, x_921); +x_923 = lean_unsigned_to_nat(1u); +lean_inc(x_922); +x_924 = l_Lean_Syntax_matchesNull(x_922, x_923); +if (x_924 == 0) { -lean_dec(x_927); +lean_dec(x_922); if (lean_obj_tag(x_1) == 0) { -lean_object* x_931; lean_object* x_932; lean_object* x_933; lean_object* x_934; lean_object* x_935; lean_object* x_936; lean_object* x_937; lean_object* x_938; lean_object* x_939; lean_object* x_940; lean_object* x_941; lean_object* x_942; lean_object* x_943; lean_object* x_944; lean_object* x_945; lean_object* x_946; lean_object* x_947; lean_object* x_948; lean_object* x_949; lean_object* x_950; lean_object* x_951; lean_object* x_952; lean_object* x_953; lean_object* x_954; lean_object* x_955; lean_object* x_956; lean_object* x_957; lean_object* x_958; lean_object* x_959; lean_object* x_960; -x_931 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); +lean_object* x_925; lean_object* x_926; lean_object* x_927; lean_object* x_928; lean_object* x_929; lean_object* x_930; lean_object* x_931; lean_object* x_932; lean_object* x_933; lean_object* x_934; lean_object* x_935; lean_object* x_936; lean_object* x_937; lean_object* x_938; lean_object* x_939; lean_object* x_940; lean_object* x_941; lean_object* x_942; lean_object* x_943; lean_object* x_944; lean_object* x_945; lean_object* x_946; lean_object* x_947; lean_object* x_948; lean_object* x_949; lean_object* x_950; lean_object* x_951; lean_object* x_952; lean_object* x_953; lean_object* x_954; +x_925 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); +x_926 = lean_ctor_get(x_925, 0); +lean_inc(x_926); +x_927 = lean_ctor_get(x_925, 1); +lean_inc(x_927); +lean_dec(x_925); +x_928 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_927); +x_929 = lean_ctor_get(x_928, 0); +lean_inc(x_929); +x_930 = lean_ctor_get(x_928, 1); +lean_inc(x_930); +lean_dec(x_928); +x_931 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_930); x_932 = lean_ctor_get(x_931, 0); lean_inc(x_932); x_933 = lean_ctor_get(x_931, 1); lean_inc(x_933); lean_dec(x_931); -x_934 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_933); -x_935 = lean_ctor_get(x_934, 0); -lean_inc(x_935); -x_936 = lean_ctor_get(x_934, 1); -lean_inc(x_936); -lean_dec(x_934); -x_937 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_936); -x_938 = lean_ctor_get(x_937, 0); -lean_inc(x_938); -x_939 = lean_ctor_get(x_937, 1); -lean_inc(x_939); -lean_dec(x_937); -x_940 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; -x_941 = l_Lean_addMacroScope(x_938, x_940, x_935); -x_942 = lean_box(0); -x_943 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; -lean_inc(x_932); -x_944 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_944, 0, x_932); -lean_ctor_set(x_944, 1, x_943); -lean_ctor_set(x_944, 2, x_941); -lean_ctor_set(x_944, 3, x_942); -x_945 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; -lean_inc(x_932); -x_946 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_946, 0, x_932); -lean_ctor_set(x_946, 1, x_945); -x_947 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; +x_934 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; +x_935 = l_Lean_addMacroScope(x_932, x_934, x_929); +x_936 = lean_box(0); +x_937 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; +lean_inc(x_926); +x_938 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_938, 0, x_926); +lean_ctor_set(x_938, 1, x_937); +lean_ctor_set(x_938, 2, x_935); +lean_ctor_set(x_938, 3, x_936); +x_939 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; +lean_inc(x_926); +x_940 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_940, 0, x_926); +lean_ctor_set(x_940, 1, x_939); +x_941 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; lean_inc(x_3); -x_948 = lean_array_push(x_947, x_3); -x_949 = lean_box(2); -x_950 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; -x_951 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_951, 0, x_949); -lean_ctor_set(x_951, 1, x_950); -lean_ctor_set(x_951, 2, x_948); -x_952 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__8; -x_953 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_953, 0, x_932); -lean_ctor_set(x_953, 1, x_952); -x_954 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__9; -x_955 = lean_array_push(x_954, x_944); -x_956 = lean_array_push(x_955, x_946); -x_957 = lean_array_push(x_956, x_951); -x_958 = lean_array_push(x_957, x_953); -x_959 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_959, 0, x_949); -lean_ctor_set(x_959, 1, x_12); -lean_ctor_set(x_959, 2, x_958); -x_960 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_939); -if (lean_obj_tag(x_960) == 0) -{ -uint8_t x_961; -x_961 = !lean_is_exclusive(x_960); -if (x_961 == 0) -{ -lean_object* x_962; lean_object* x_963; -x_962 = lean_ctor_get(x_960, 0); -x_963 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_963, 0, x_959); -lean_ctor_set(x_963, 1, x_962); -lean_ctor_set(x_960, 0, x_963); -return x_960; -} -else -{ -lean_object* x_964; lean_object* x_965; lean_object* x_966; lean_object* x_967; -x_964 = lean_ctor_get(x_960, 0); -x_965 = lean_ctor_get(x_960, 1); -lean_inc(x_965); -lean_inc(x_964); -lean_dec(x_960); -x_966 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_966, 0, x_959); -lean_ctor_set(x_966, 1, x_964); -x_967 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_967, 0, x_966); -lean_ctor_set(x_967, 1, x_965); -return x_967; +x_942 = lean_array_push(x_941, x_3); +x_943 = lean_box(2); +x_944 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__3; +x_945 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_945, 0, x_943); +lean_ctor_set(x_945, 1, x_944); +lean_ctor_set(x_945, 2, x_942); +x_946 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__6; +x_947 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_947, 0, x_926); +lean_ctor_set(x_947, 1, x_946); +x_948 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; +x_949 = lean_array_push(x_948, x_938); +x_950 = lean_array_push(x_949, x_940); +x_951 = lean_array_push(x_950, x_945); +x_952 = lean_array_push(x_951, x_947); +x_953 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_953, 0, x_943); +lean_ctor_set(x_953, 1, x_12); +lean_ctor_set(x_953, 2, x_952); +x_954 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_933); +if (lean_obj_tag(x_954) == 0) +{ +uint8_t x_955; +x_955 = !lean_is_exclusive(x_954); +if (x_955 == 0) +{ +lean_object* x_956; lean_object* x_957; +x_956 = lean_ctor_get(x_954, 0); +x_957 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_957, 0, x_953); +lean_ctor_set(x_957, 1, x_956); +lean_ctor_set(x_954, 0, x_957); +return x_954; +} +else +{ +lean_object* x_958; lean_object* x_959; lean_object* x_960; lean_object* x_961; +x_958 = lean_ctor_get(x_954, 0); +x_959 = lean_ctor_get(x_954, 1); +lean_inc(x_959); +lean_inc(x_958); +lean_dec(x_954); +x_960 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_960, 0, x_953); +lean_ctor_set(x_960, 1, x_958); +x_961 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_961, 0, x_960); +lean_ctor_set(x_961, 1, x_959); +return x_961; } } else { -uint8_t x_968; -lean_dec(x_959); -x_968 = !lean_is_exclusive(x_960); -if (x_968 == 0) +uint8_t x_962; +lean_dec(x_953); +x_962 = !lean_is_exclusive(x_954); +if (x_962 == 0) { -return x_960; +return x_954; } else { -lean_object* x_969; lean_object* x_970; lean_object* x_971; -x_969 = lean_ctor_get(x_960, 0); -x_970 = lean_ctor_get(x_960, 1); -lean_inc(x_970); -lean_inc(x_969); -lean_dec(x_960); -x_971 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_971, 0, x_969); -lean_ctor_set(x_971, 1, x_970); -return x_971; +lean_object* x_963; lean_object* x_964; lean_object* x_965; +x_963 = lean_ctor_get(x_954, 0); +x_964 = lean_ctor_get(x_954, 1); +lean_inc(x_964); +lean_inc(x_963); +lean_dec(x_954); +x_965 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_965, 0, x_963); +lean_ctor_set(x_965, 1, x_964); +return x_965; } } } else { -lean_object* x_972; lean_object* x_973; +lean_object* x_966; lean_object* x_967; lean_dec(x_2); -x_972 = lean_ctor_get(x_1, 0); -lean_inc(x_972); +x_966 = lean_ctor_get(x_1, 0); +lean_inc(x_966); lean_dec(x_1); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_973 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_972, x_4, x_5, x_6); -if (lean_obj_tag(x_973) == 0) -{ -lean_object* x_974; lean_object* x_975; lean_object* x_976; -x_974 = lean_ctor_get(x_973, 0); -lean_inc(x_974); -x_975 = lean_ctor_get(x_973, 1); -lean_inc(x_975); -lean_dec(x_973); -x_976 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_974, x_4, x_5, x_975); +x_967 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_966, x_4, x_5, x_6); +if (lean_obj_tag(x_967) == 0) +{ +lean_object* x_968; lean_object* x_969; lean_object* x_970; +x_968 = lean_ctor_get(x_967, 0); +lean_inc(x_968); +x_969 = lean_ctor_get(x_967, 1); +lean_inc(x_969); +lean_dec(x_967); +x_970 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_968, x_4, x_5, x_969); lean_dec(x_5); lean_dec(x_4); -return x_976; +return x_970; } else { -uint8_t x_977; +uint8_t x_971; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_977 = !lean_is_exclusive(x_973); -if (x_977 == 0) +x_971 = !lean_is_exclusive(x_967); +if (x_971 == 0) { -return x_973; +return x_967; } else { -lean_object* x_978; lean_object* x_979; lean_object* x_980; -x_978 = lean_ctor_get(x_973, 0); -x_979 = lean_ctor_get(x_973, 1); -lean_inc(x_979); -lean_inc(x_978); -lean_dec(x_973); -x_980 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_980, 0, x_978); -lean_ctor_set(x_980, 1, x_979); -return x_980; +lean_object* x_972; lean_object* x_973; lean_object* x_974; +x_972 = lean_ctor_get(x_967, 0); +x_973 = lean_ctor_get(x_967, 1); +lean_inc(x_973); +lean_inc(x_972); +lean_dec(x_967); +x_974 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_974, 0, x_972); +lean_ctor_set(x_974, 1, x_973); +return x_974; } } } } else { -lean_object* x_981; lean_object* x_982; lean_object* x_983; +lean_object* x_975; lean_object* x_976; lean_object* x_977; lean_dec(x_1); -x_981 = l_Lean_Syntax_getArg(x_927, x_516); -lean_dec(x_927); -x_982 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__2___closed__7; +x_975 = l_Lean_Syntax_getArg(x_922, x_514); +lean_dec(x_922); +x_976 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__2___closed__7; lean_inc(x_5); lean_inc(x_4); -x_983 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat(x_520, x_981, x_2, x_982, x_4, x_5, x_6); -if (lean_obj_tag(x_983) == 0) -{ -lean_object* x_984; lean_object* x_985; lean_object* x_986; -x_984 = lean_ctor_get(x_983, 0); -lean_inc(x_984); -x_985 = lean_ctor_get(x_983, 1); -lean_inc(x_985); -lean_dec(x_983); -x_986 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_984, x_4, x_5, x_985); +x_977 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat(x_518, x_975, x_2, x_976, x_4, x_5, x_6); +if (lean_obj_tag(x_977) == 0) +{ +lean_object* x_978; lean_object* x_979; lean_object* x_980; +x_978 = lean_ctor_get(x_977, 0); +lean_inc(x_978); +x_979 = lean_ctor_get(x_977, 1); +lean_inc(x_979); +lean_dec(x_977); +x_980 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_978, x_4, x_5, x_979); lean_dec(x_5); lean_dec(x_4); -return x_986; +return x_980; } else { -uint8_t x_987; +uint8_t x_981; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_987 = !lean_is_exclusive(x_983); -if (x_987 == 0) +x_981 = !lean_is_exclusive(x_977); +if (x_981 == 0) { -return x_983; +return x_977; } else { -lean_object* x_988; lean_object* x_989; lean_object* x_990; -x_988 = lean_ctor_get(x_983, 0); -x_989 = lean_ctor_get(x_983, 1); -lean_inc(x_989); -lean_inc(x_988); -lean_dec(x_983); -x_990 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_990, 0, x_988); -lean_ctor_set(x_990, 1, x_989); -return x_990; +lean_object* x_982; lean_object* x_983; lean_object* x_984; +x_982 = lean_ctor_get(x_977, 0); +x_983 = lean_ctor_get(x_977, 1); +lean_inc(x_983); +lean_inc(x_982); +lean_dec(x_977); +x_984 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_984, 0, x_982); +lean_ctor_set(x_984, 1, x_983); +return x_984; } } } @@ -12897,228 +12965,227 @@ return x_990; } else { -lean_object* x_991; lean_object* x_992; lean_object* x_993; lean_object* x_994; uint8_t x_995; -lean_dec(x_517); -x_991 = lean_unsigned_to_nat(2u); -x_992 = l_Lean_Syntax_getArg(x_3, x_991); -x_993 = l_Lean_nullKind; -x_994 = lean_unsigned_to_nat(1u); -lean_inc(x_992); -x_995 = l_Lean_Syntax_isNodeOf(x_992, x_993, x_994); -if (x_995 == 0) +lean_object* x_985; lean_object* x_986; lean_object* x_987; uint8_t x_988; +lean_dec(x_515); +x_985 = lean_unsigned_to_nat(2u); +x_986 = l_Lean_Syntax_getArg(x_3, x_985); +x_987 = lean_unsigned_to_nat(1u); +lean_inc(x_986); +x_988 = l_Lean_Syntax_matchesNull(x_986, x_987); +if (x_988 == 0) { -lean_dec(x_992); +lean_dec(x_986); if (lean_obj_tag(x_1) == 0) { -lean_object* x_996; lean_object* x_997; lean_object* x_998; lean_object* x_999; lean_object* x_1000; lean_object* x_1001; lean_object* x_1002; lean_object* x_1003; lean_object* x_1004; lean_object* x_1005; lean_object* x_1006; lean_object* x_1007; lean_object* x_1008; lean_object* x_1009; lean_object* x_1010; lean_object* x_1011; lean_object* x_1012; lean_object* x_1013; lean_object* x_1014; lean_object* x_1015; lean_object* x_1016; lean_object* x_1017; lean_object* x_1018; lean_object* x_1019; lean_object* x_1020; lean_object* x_1021; lean_object* x_1022; lean_object* x_1023; lean_object* x_1024; lean_object* x_1025; -x_996 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); -x_997 = lean_ctor_get(x_996, 0); -lean_inc(x_997); -x_998 = lean_ctor_get(x_996, 1); -lean_inc(x_998); -lean_dec(x_996); -x_999 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_998); -x_1000 = lean_ctor_get(x_999, 0); -lean_inc(x_1000); -x_1001 = lean_ctor_get(x_999, 1); -lean_inc(x_1001); -lean_dec(x_999); -x_1002 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_1001); -x_1003 = lean_ctor_get(x_1002, 0); -lean_inc(x_1003); -x_1004 = lean_ctor_get(x_1002, 1); -lean_inc(x_1004); -lean_dec(x_1002); -x_1005 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; -x_1006 = l_Lean_addMacroScope(x_1003, x_1005, x_1000); -x_1007 = lean_box(0); -x_1008 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; +lean_object* x_989; lean_object* x_990; lean_object* x_991; lean_object* x_992; lean_object* x_993; lean_object* x_994; lean_object* x_995; lean_object* x_996; lean_object* x_997; lean_object* x_998; lean_object* x_999; lean_object* x_1000; lean_object* x_1001; lean_object* x_1002; lean_object* x_1003; lean_object* x_1004; lean_object* x_1005; lean_object* x_1006; lean_object* x_1007; lean_object* x_1008; lean_object* x_1009; lean_object* x_1010; lean_object* x_1011; lean_object* x_1012; lean_object* x_1013; lean_object* x_1014; lean_object* x_1015; lean_object* x_1016; lean_object* x_1017; lean_object* x_1018; +x_989 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); +x_990 = lean_ctor_get(x_989, 0); +lean_inc(x_990); +x_991 = lean_ctor_get(x_989, 1); +lean_inc(x_991); +lean_dec(x_989); +x_992 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_991); +x_993 = lean_ctor_get(x_992, 0); +lean_inc(x_993); +x_994 = lean_ctor_get(x_992, 1); +lean_inc(x_994); +lean_dec(x_992); +x_995 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_994); +x_996 = lean_ctor_get(x_995, 0); +lean_inc(x_996); +x_997 = lean_ctor_get(x_995, 1); lean_inc(x_997); -x_1009 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1009, 0, x_997); +lean_dec(x_995); +x_998 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; +x_999 = l_Lean_addMacroScope(x_996, x_998, x_993); +x_1000 = lean_box(0); +x_1001 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; +lean_inc(x_990); +x_1002 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1002, 0, x_990); +lean_ctor_set(x_1002, 1, x_1001); +lean_ctor_set(x_1002, 2, x_999); +lean_ctor_set(x_1002, 3, x_1000); +x_1003 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; +lean_inc(x_990); +x_1004 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1004, 0, x_990); +lean_ctor_set(x_1004, 1, x_1003); +x_1005 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; +lean_inc(x_3); +x_1006 = lean_array_push(x_1005, x_3); +x_1007 = lean_box(2); +x_1008 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__3; +x_1009 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1009, 0, x_1007); lean_ctor_set(x_1009, 1, x_1008); lean_ctor_set(x_1009, 2, x_1006); -lean_ctor_set(x_1009, 3, x_1007); -x_1010 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; -lean_inc(x_997); +x_1010 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__6; x_1011 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1011, 0, x_997); +lean_ctor_set(x_1011, 0, x_990); lean_ctor_set(x_1011, 1, x_1010); -x_1012 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; -lean_inc(x_3); -x_1013 = lean_array_push(x_1012, x_3); -x_1014 = lean_box(2); -x_1015 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; -x_1016 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1016, 0, x_1014); -lean_ctor_set(x_1016, 1, x_1015); -lean_ctor_set(x_1016, 2, x_1013); -x_1017 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__8; -x_1018 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1018, 0, x_997); -lean_ctor_set(x_1018, 1, x_1017); -x_1019 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__9; -x_1020 = lean_array_push(x_1019, x_1009); -x_1021 = lean_array_push(x_1020, x_1011); -x_1022 = lean_array_push(x_1021, x_1016); -x_1023 = lean_array_push(x_1022, x_1018); -x_1024 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1024, 0, x_1014); -lean_ctor_set(x_1024, 1, x_12); -lean_ctor_set(x_1024, 2, x_1023); -x_1025 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_1004); -if (lean_obj_tag(x_1025) == 0) -{ -uint8_t x_1026; -x_1026 = !lean_is_exclusive(x_1025); -if (x_1026 == 0) -{ -lean_object* x_1027; lean_object* x_1028; -x_1027 = lean_ctor_get(x_1025, 0); -x_1028 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1028, 0, x_1024); -lean_ctor_set(x_1028, 1, x_1027); -lean_ctor_set(x_1025, 0, x_1028); +x_1012 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; +x_1013 = lean_array_push(x_1012, x_1002); +x_1014 = lean_array_push(x_1013, x_1004); +x_1015 = lean_array_push(x_1014, x_1009); +x_1016 = lean_array_push(x_1015, x_1011); +x_1017 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1017, 0, x_1007); +lean_ctor_set(x_1017, 1, x_12); +lean_ctor_set(x_1017, 2, x_1016); +x_1018 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_997); +if (lean_obj_tag(x_1018) == 0) +{ +uint8_t x_1019; +x_1019 = !lean_is_exclusive(x_1018); +if (x_1019 == 0) +{ +lean_object* x_1020; lean_object* x_1021; +x_1020 = lean_ctor_get(x_1018, 0); +x_1021 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1021, 0, x_1017); +lean_ctor_set(x_1021, 1, x_1020); +lean_ctor_set(x_1018, 0, x_1021); +return x_1018; +} +else +{ +lean_object* x_1022; lean_object* x_1023; lean_object* x_1024; lean_object* x_1025; +x_1022 = lean_ctor_get(x_1018, 0); +x_1023 = lean_ctor_get(x_1018, 1); +lean_inc(x_1023); +lean_inc(x_1022); +lean_dec(x_1018); +x_1024 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1024, 0, x_1017); +lean_ctor_set(x_1024, 1, x_1022); +x_1025 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1025, 0, x_1024); +lean_ctor_set(x_1025, 1, x_1023); return x_1025; } -else -{ -lean_object* x_1029; lean_object* x_1030; lean_object* x_1031; lean_object* x_1032; -x_1029 = lean_ctor_get(x_1025, 0); -x_1030 = lean_ctor_get(x_1025, 1); -lean_inc(x_1030); -lean_inc(x_1029); -lean_dec(x_1025); -x_1031 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1031, 0, x_1024); -lean_ctor_set(x_1031, 1, x_1029); -x_1032 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1032, 0, x_1031); -lean_ctor_set(x_1032, 1, x_1030); -return x_1032; -} } else { -uint8_t x_1033; -lean_dec(x_1024); -x_1033 = !lean_is_exclusive(x_1025); -if (x_1033 == 0) +uint8_t x_1026; +lean_dec(x_1017); +x_1026 = !lean_is_exclusive(x_1018); +if (x_1026 == 0) { -return x_1025; +return x_1018; } else { -lean_object* x_1034; lean_object* x_1035; lean_object* x_1036; -x_1034 = lean_ctor_get(x_1025, 0); -x_1035 = lean_ctor_get(x_1025, 1); -lean_inc(x_1035); -lean_inc(x_1034); -lean_dec(x_1025); -x_1036 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1036, 0, x_1034); -lean_ctor_set(x_1036, 1, x_1035); -return x_1036; +lean_object* x_1027; lean_object* x_1028; lean_object* x_1029; +x_1027 = lean_ctor_get(x_1018, 0); +x_1028 = lean_ctor_get(x_1018, 1); +lean_inc(x_1028); +lean_inc(x_1027); +lean_dec(x_1018); +x_1029 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1029, 0, x_1027); +lean_ctor_set(x_1029, 1, x_1028); +return x_1029; } } } else { -lean_object* x_1037; lean_object* x_1038; +lean_object* x_1030; lean_object* x_1031; lean_dec(x_2); -x_1037 = lean_ctor_get(x_1, 0); -lean_inc(x_1037); +x_1030 = lean_ctor_get(x_1, 0); +lean_inc(x_1030); lean_dec(x_1); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_1038 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_1037, x_4, x_5, x_6); -if (lean_obj_tag(x_1038) == 0) +x_1031 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_1030, x_4, x_5, x_6); +if (lean_obj_tag(x_1031) == 0) { -lean_object* x_1039; lean_object* x_1040; lean_object* x_1041; -x_1039 = lean_ctor_get(x_1038, 0); -lean_inc(x_1039); -x_1040 = lean_ctor_get(x_1038, 1); -lean_inc(x_1040); -lean_dec(x_1038); -x_1041 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_1039, x_4, x_5, x_1040); +lean_object* x_1032; lean_object* x_1033; lean_object* x_1034; +x_1032 = lean_ctor_get(x_1031, 0); +lean_inc(x_1032); +x_1033 = lean_ctor_get(x_1031, 1); +lean_inc(x_1033); +lean_dec(x_1031); +x_1034 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_1032, x_4, x_5, x_1033); lean_dec(x_5); lean_dec(x_4); -return x_1041; +return x_1034; } else { -uint8_t x_1042; +uint8_t x_1035; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_1042 = !lean_is_exclusive(x_1038); -if (x_1042 == 0) +x_1035 = !lean_is_exclusive(x_1031); +if (x_1035 == 0) { -return x_1038; +return x_1031; } else { -lean_object* x_1043; lean_object* x_1044; lean_object* x_1045; -x_1043 = lean_ctor_get(x_1038, 0); -x_1044 = lean_ctor_get(x_1038, 1); -lean_inc(x_1044); -lean_inc(x_1043); -lean_dec(x_1038); -x_1045 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1045, 0, x_1043); -lean_ctor_set(x_1045, 1, x_1044); -return x_1045; +lean_object* x_1036; lean_object* x_1037; lean_object* x_1038; +x_1036 = lean_ctor_get(x_1031, 0); +x_1037 = lean_ctor_get(x_1031, 1); +lean_inc(x_1037); +lean_inc(x_1036); +lean_dec(x_1031); +x_1038 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1038, 0, x_1036); +lean_ctor_set(x_1038, 1, x_1037); +return x_1038; } } } } else { -lean_object* x_1046; lean_object* x_1047; lean_object* x_1048; +lean_object* x_1039; lean_object* x_1040; lean_object* x_1041; lean_dec(x_1); -x_1046 = l_Lean_Syntax_getArg(x_992, x_516); -lean_dec(x_992); -x_1047 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__24; +x_1039 = l_Lean_Syntax_getArg(x_986, x_514); +lean_dec(x_986); +x_1040 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__24; lean_inc(x_5); lean_inc(x_4); -x_1048 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat(x_518, x_1046, x_2, x_1047, x_4, x_5, x_6); -if (lean_obj_tag(x_1048) == 0) +x_1041 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat(x_516, x_1039, x_2, x_1040, x_4, x_5, x_6); +if (lean_obj_tag(x_1041) == 0) { -lean_object* x_1049; lean_object* x_1050; lean_object* x_1051; -x_1049 = lean_ctor_get(x_1048, 0); -lean_inc(x_1049); -x_1050 = lean_ctor_get(x_1048, 1); -lean_inc(x_1050); -lean_dec(x_1048); -x_1051 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_1049, x_4, x_5, x_1050); +lean_object* x_1042; lean_object* x_1043; lean_object* x_1044; +x_1042 = lean_ctor_get(x_1041, 0); +lean_inc(x_1042); +x_1043 = lean_ctor_get(x_1041, 1); +lean_inc(x_1043); +lean_dec(x_1041); +x_1044 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_1042, x_4, x_5, x_1043); lean_dec(x_5); lean_dec(x_4); -return x_1051; +return x_1044; } else { -uint8_t x_1052; +uint8_t x_1045; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_1052 = !lean_is_exclusive(x_1048); -if (x_1052 == 0) +x_1045 = !lean_is_exclusive(x_1041); +if (x_1045 == 0) { -return x_1048; +return x_1041; } else { -lean_object* x_1053; lean_object* x_1054; lean_object* x_1055; -x_1053 = lean_ctor_get(x_1048, 0); -x_1054 = lean_ctor_get(x_1048, 1); -lean_inc(x_1054); -lean_inc(x_1053); -lean_dec(x_1048); -x_1055 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1055, 0, x_1053); -lean_ctor_set(x_1055, 1, x_1054); -return x_1055; +lean_object* x_1046; lean_object* x_1047; lean_object* x_1048; +x_1046 = lean_ctor_get(x_1041, 0); +x_1047 = lean_ctor_get(x_1041, 1); +lean_inc(x_1047); +lean_inc(x_1046); +lean_dec(x_1041); +x_1048 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1048, 0, x_1046); +lean_ctor_set(x_1048, 1, x_1047); +return x_1048; } } } @@ -13127,241 +13194,241 @@ return x_1055; } else { -lean_object* x_1056; lean_object* x_1057; lean_object* x_1058; uint8_t x_1059; +lean_object* x_1049; lean_object* x_1050; lean_object* x_1051; uint8_t x_1052; lean_dec(x_7); -x_1056 = lean_unsigned_to_nat(1u); -x_1057 = l_Lean_Syntax_getArg(x_3, x_1056); -x_1058 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__11; -lean_inc(x_1057); -x_1059 = l_Lean_Syntax_isOfKind(x_1057, x_1058); -if (x_1059 == 0) +x_1049 = lean_unsigned_to_nat(1u); +x_1050 = l_Lean_Syntax_getArg(x_3, x_1049); +x_1051 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__10; +lean_inc(x_1050); +x_1052 = l_Lean_Syntax_isOfKind(x_1050, x_1051); +if (x_1052 == 0) { -lean_dec(x_1057); +lean_dec(x_1050); if (lean_obj_tag(x_1) == 0) { -lean_object* x_1060; lean_object* x_1061; lean_object* x_1062; lean_object* x_1063; lean_object* x_1064; lean_object* x_1065; lean_object* x_1066; lean_object* x_1067; lean_object* x_1068; lean_object* x_1069; lean_object* x_1070; lean_object* x_1071; lean_object* x_1072; lean_object* x_1073; lean_object* x_1074; lean_object* x_1075; lean_object* x_1076; lean_object* x_1077; lean_object* x_1078; lean_object* x_1079; lean_object* x_1080; lean_object* x_1081; lean_object* x_1082; lean_object* x_1083; lean_object* x_1084; lean_object* x_1085; lean_object* x_1086; lean_object* x_1087; lean_object* x_1088; lean_object* x_1089; lean_object* x_1090; -x_1060 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); -x_1061 = lean_ctor_get(x_1060, 0); -lean_inc(x_1061); -x_1062 = lean_ctor_get(x_1060, 1); -lean_inc(x_1062); -lean_dec(x_1060); -x_1063 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_1062); -x_1064 = lean_ctor_get(x_1063, 0); -lean_inc(x_1064); -x_1065 = lean_ctor_get(x_1063, 1); -lean_inc(x_1065); -lean_dec(x_1063); -x_1066 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_1065); -x_1067 = lean_ctor_get(x_1066, 0); -lean_inc(x_1067); -x_1068 = lean_ctor_get(x_1066, 1); -lean_inc(x_1068); -lean_dec(x_1066); -x_1069 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; -x_1070 = l_Lean_addMacroScope(x_1067, x_1069, x_1064); -x_1071 = lean_box(0); -x_1072 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; +lean_object* x_1053; lean_object* x_1054; lean_object* x_1055; lean_object* x_1056; lean_object* x_1057; lean_object* x_1058; lean_object* x_1059; lean_object* x_1060; lean_object* x_1061; lean_object* x_1062; lean_object* x_1063; lean_object* x_1064; lean_object* x_1065; lean_object* x_1066; lean_object* x_1067; lean_object* x_1068; lean_object* x_1069; lean_object* x_1070; lean_object* x_1071; lean_object* x_1072; lean_object* x_1073; lean_object* x_1074; lean_object* x_1075; lean_object* x_1076; lean_object* x_1077; lean_object* x_1078; lean_object* x_1079; lean_object* x_1080; lean_object* x_1081; lean_object* x_1082; lean_object* x_1083; +x_1053 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); +x_1054 = lean_ctor_get(x_1053, 0); +lean_inc(x_1054); +x_1055 = lean_ctor_get(x_1053, 1); +lean_inc(x_1055); +lean_dec(x_1053); +x_1056 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_1055); +x_1057 = lean_ctor_get(x_1056, 0); +lean_inc(x_1057); +x_1058 = lean_ctor_get(x_1056, 1); +lean_inc(x_1058); +lean_dec(x_1056); +x_1059 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_1058); +x_1060 = lean_ctor_get(x_1059, 0); +lean_inc(x_1060); +x_1061 = lean_ctor_get(x_1059, 1); lean_inc(x_1061); -x_1073 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1073, 0, x_1061); +lean_dec(x_1059); +x_1062 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; +x_1063 = l_Lean_addMacroScope(x_1060, x_1062, x_1057); +x_1064 = lean_box(0); +x_1065 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; +lean_inc(x_1054); +x_1066 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1066, 0, x_1054); +lean_ctor_set(x_1066, 1, x_1065); +lean_ctor_set(x_1066, 2, x_1063); +lean_ctor_set(x_1066, 3, x_1064); +x_1067 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; +lean_inc(x_1054); +x_1068 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1068, 0, x_1054); +lean_ctor_set(x_1068, 1, x_1067); +x_1069 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; +lean_inc(x_3); +x_1070 = lean_array_push(x_1069, x_3); +x_1071 = lean_box(2); +x_1072 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__3; +x_1073 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1073, 0, x_1071); lean_ctor_set(x_1073, 1, x_1072); lean_ctor_set(x_1073, 2, x_1070); -lean_ctor_set(x_1073, 3, x_1071); -x_1074 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; -lean_inc(x_1061); +x_1074 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__6; x_1075 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1075, 0, x_1061); +lean_ctor_set(x_1075, 0, x_1054); lean_ctor_set(x_1075, 1, x_1074); -x_1076 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; -lean_inc(x_3); -x_1077 = lean_array_push(x_1076, x_3); -x_1078 = lean_box(2); -x_1079 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; -x_1080 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1080, 0, x_1078); -lean_ctor_set(x_1080, 1, x_1079); -lean_ctor_set(x_1080, 2, x_1077); -x_1081 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__8; -x_1082 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1082, 0, x_1061); +x_1076 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; +x_1077 = lean_array_push(x_1076, x_1066); +x_1078 = lean_array_push(x_1077, x_1068); +x_1079 = lean_array_push(x_1078, x_1073); +x_1080 = lean_array_push(x_1079, x_1075); +x_1081 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__6; +x_1082 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1082, 0, x_1071); lean_ctor_set(x_1082, 1, x_1081); -x_1083 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__9; -x_1084 = lean_array_push(x_1083, x_1073); -x_1085 = lean_array_push(x_1084, x_1075); -x_1086 = lean_array_push(x_1085, x_1080); -x_1087 = lean_array_push(x_1086, x_1082); -x_1088 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__6; -x_1089 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1089, 0, x_1078); -lean_ctor_set(x_1089, 1, x_1088); -lean_ctor_set(x_1089, 2, x_1087); -x_1090 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_1068); -if (lean_obj_tag(x_1090) == 0) +lean_ctor_set(x_1082, 2, x_1080); +x_1083 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_1061); +if (lean_obj_tag(x_1083) == 0) { -uint8_t x_1091; -x_1091 = !lean_is_exclusive(x_1090); -if (x_1091 == 0) +uint8_t x_1084; +x_1084 = !lean_is_exclusive(x_1083); +if (x_1084 == 0) { -lean_object* x_1092; lean_object* x_1093; -x_1092 = lean_ctor_get(x_1090, 0); -x_1093 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1093, 0, x_1089); -lean_ctor_set(x_1093, 1, x_1092); -lean_ctor_set(x_1090, 0, x_1093); -return x_1090; +lean_object* x_1085; lean_object* x_1086; +x_1085 = lean_ctor_get(x_1083, 0); +x_1086 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1086, 0, x_1082); +lean_ctor_set(x_1086, 1, x_1085); +lean_ctor_set(x_1083, 0, x_1086); +return x_1083; } else { -lean_object* x_1094; lean_object* x_1095; lean_object* x_1096; lean_object* x_1097; -x_1094 = lean_ctor_get(x_1090, 0); -x_1095 = lean_ctor_get(x_1090, 1); -lean_inc(x_1095); -lean_inc(x_1094); -lean_dec(x_1090); -x_1096 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1096, 0, x_1089); -lean_ctor_set(x_1096, 1, x_1094); -x_1097 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1097, 0, x_1096); -lean_ctor_set(x_1097, 1, x_1095); -return x_1097; +lean_object* x_1087; lean_object* x_1088; lean_object* x_1089; lean_object* x_1090; +x_1087 = lean_ctor_get(x_1083, 0); +x_1088 = lean_ctor_get(x_1083, 1); +lean_inc(x_1088); +lean_inc(x_1087); +lean_dec(x_1083); +x_1089 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1089, 0, x_1082); +lean_ctor_set(x_1089, 1, x_1087); +x_1090 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1090, 0, x_1089); +lean_ctor_set(x_1090, 1, x_1088); +return x_1090; } } else { -uint8_t x_1098; -lean_dec(x_1089); -x_1098 = !lean_is_exclusive(x_1090); -if (x_1098 == 0) +uint8_t x_1091; +lean_dec(x_1082); +x_1091 = !lean_is_exclusive(x_1083); +if (x_1091 == 0) { -return x_1090; +return x_1083; } else { -lean_object* x_1099; lean_object* x_1100; lean_object* x_1101; -x_1099 = lean_ctor_get(x_1090, 0); -x_1100 = lean_ctor_get(x_1090, 1); -lean_inc(x_1100); -lean_inc(x_1099); -lean_dec(x_1090); -x_1101 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1101, 0, x_1099); -lean_ctor_set(x_1101, 1, x_1100); -return x_1101; +lean_object* x_1092; lean_object* x_1093; lean_object* x_1094; +x_1092 = lean_ctor_get(x_1083, 0); +x_1093 = lean_ctor_get(x_1083, 1); +lean_inc(x_1093); +lean_inc(x_1092); +lean_dec(x_1083); +x_1094 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1094, 0, x_1092); +lean_ctor_set(x_1094, 1, x_1093); +return x_1094; } } } else { -lean_object* x_1102; lean_object* x_1103; +lean_object* x_1095; lean_object* x_1096; lean_dec(x_2); -x_1102 = lean_ctor_get(x_1, 0); -lean_inc(x_1102); +x_1095 = lean_ctor_get(x_1, 0); +lean_inc(x_1095); lean_dec(x_1); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_1103 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_1102, x_4, x_5, x_6); -if (lean_obj_tag(x_1103) == 0) +x_1096 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_1095, x_4, x_5, x_6); +if (lean_obj_tag(x_1096) == 0) { -lean_object* x_1104; lean_object* x_1105; lean_object* x_1106; -x_1104 = lean_ctor_get(x_1103, 0); -lean_inc(x_1104); -x_1105 = lean_ctor_get(x_1103, 1); -lean_inc(x_1105); -lean_dec(x_1103); -x_1106 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_1104, x_4, x_5, x_1105); +lean_object* x_1097; lean_object* x_1098; lean_object* x_1099; +x_1097 = lean_ctor_get(x_1096, 0); +lean_inc(x_1097); +x_1098 = lean_ctor_get(x_1096, 1); +lean_inc(x_1098); +lean_dec(x_1096); +x_1099 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_1097, x_4, x_5, x_1098); lean_dec(x_5); lean_dec(x_4); -return x_1106; +return x_1099; } else { -uint8_t x_1107; +uint8_t x_1100; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_1107 = !lean_is_exclusive(x_1103); -if (x_1107 == 0) +x_1100 = !lean_is_exclusive(x_1096); +if (x_1100 == 0) { -return x_1103; +return x_1096; } else { -lean_object* x_1108; lean_object* x_1109; lean_object* x_1110; -x_1108 = lean_ctor_get(x_1103, 0); -x_1109 = lean_ctor_get(x_1103, 1); -lean_inc(x_1109); -lean_inc(x_1108); -lean_dec(x_1103); -x_1110 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1110, 0, x_1108); -lean_ctor_set(x_1110, 1, x_1109); -return x_1110; +lean_object* x_1101; lean_object* x_1102; lean_object* x_1103; +x_1101 = lean_ctor_get(x_1096, 0); +x_1102 = lean_ctor_get(x_1096, 1); +lean_inc(x_1102); +lean_inc(x_1101); +lean_dec(x_1096); +x_1103 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1103, 0, x_1101); +lean_ctor_set(x_1103, 1, x_1102); +return x_1103; } } } } else { -lean_object* x_1111; lean_object* x_1112; +lean_object* x_1104; lean_object* x_1105; lean_dec(x_1); -x_1111 = lean_alloc_closure((void*)(l_Lean_Elab_Command_strLitToPattern___boxed), 3, 1); -lean_closure_set(x_1111, 0, x_1057); +x_1104 = lean_alloc_closure((void*)(l_Lean_Elab_Command_strLitToPattern___boxed), 3, 1); +lean_closure_set(x_1104, 0, x_1050); lean_inc(x_5); lean_inc(x_4); -x_1112 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__10(x_1111, x_4, x_5, x_6); -if (lean_obj_tag(x_1112) == 0) -{ -lean_object* x_1113; lean_object* x_1114; lean_object* x_1115; lean_object* x_1116; lean_object* x_1117; lean_object* x_1118; lean_object* x_1119; lean_object* x_1120; lean_object* x_1121; lean_object* x_1122; lean_object* x_1123; lean_object* x_1124; lean_object* x_1125; -x_1113 = lean_ctor_get(x_1112, 0); -lean_inc(x_1113); -x_1114 = lean_ctor_get(x_1112, 1); -lean_inc(x_1114); -lean_dec(x_1112); -x_1115 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__9; -x_1116 = lean_array_push(x_1115, x_1113); -x_1117 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__28; -x_1118 = lean_array_push(x_1116, x_1117); -x_1119 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__30; -x_1120 = lean_array_push(x_1118, x_1119); -x_1121 = lean_array_push(x_1120, x_2); -x_1122 = lean_box(2); -x_1123 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__26; -x_1124 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1124, 0, x_1122); -lean_ctor_set(x_1124, 1, x_1123); -lean_ctor_set(x_1124, 2, x_1121); -x_1125 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_1124, x_4, x_5, x_1114); +x_1105 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__10(x_1104, x_4, x_5, x_6); +if (lean_obj_tag(x_1105) == 0) +{ +lean_object* x_1106; lean_object* x_1107; lean_object* x_1108; lean_object* x_1109; lean_object* x_1110; lean_object* x_1111; lean_object* x_1112; lean_object* x_1113; lean_object* x_1114; lean_object* x_1115; lean_object* x_1116; lean_object* x_1117; lean_object* x_1118; +x_1106 = lean_ctor_get(x_1105, 0); +lean_inc(x_1106); +x_1107 = lean_ctor_get(x_1105, 1); +lean_inc(x_1107); +lean_dec(x_1105); +x_1108 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; +x_1109 = lean_array_push(x_1108, x_1106); +x_1110 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__28; +x_1111 = lean_array_push(x_1109, x_1110); +x_1112 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__30; +x_1113 = lean_array_push(x_1111, x_1112); +x_1114 = lean_array_push(x_1113, x_2); +x_1115 = lean_box(2); +x_1116 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__26; +x_1117 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1117, 0, x_1115); +lean_ctor_set(x_1117, 1, x_1116); +lean_ctor_set(x_1117, 2, x_1114); +x_1118 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_1117, x_4, x_5, x_1107); lean_dec(x_5); lean_dec(x_4); -return x_1125; +return x_1118; } else { -uint8_t x_1126; +uint8_t x_1119; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_1126 = !lean_is_exclusive(x_1112); -if (x_1126 == 0) +x_1119 = !lean_is_exclusive(x_1105); +if (x_1119 == 0) { -return x_1112; +return x_1105; } else { -lean_object* x_1127; lean_object* x_1128; lean_object* x_1129; -x_1127 = lean_ctor_get(x_1112, 0); -x_1128 = lean_ctor_get(x_1112, 1); -lean_inc(x_1128); -lean_inc(x_1127); -lean_dec(x_1112); -x_1129 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1129, 0, x_1127); -lean_ctor_set(x_1129, 1, x_1128); -return x_1129; +lean_object* x_1120; lean_object* x_1121; lean_object* x_1122; +x_1120 = lean_ctor_get(x_1105, 0); +x_1121 = lean_ctor_get(x_1105, 1); +lean_inc(x_1121); +lean_inc(x_1120); +lean_dec(x_1105); +x_1122 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1122, 0, x_1120); +lean_ctor_set(x_1122, 1, x_1121); +return x_1122; } } } @@ -13369,241 +13436,241 @@ return x_1129; } else { -lean_object* x_1130; lean_object* x_1131; lean_object* x_1132; uint8_t x_1133; +lean_object* x_1123; lean_object* x_1124; lean_object* x_1125; uint8_t x_1126; lean_dec(x_7); -x_1130 = lean_unsigned_to_nat(0u); -x_1131 = l_Lean_Syntax_getArg(x_3, x_1130); -x_1132 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__11; -lean_inc(x_1131); -x_1133 = l_Lean_Syntax_isOfKind(x_1131, x_1132); -if (x_1133 == 0) +x_1123 = lean_unsigned_to_nat(0u); +x_1124 = l_Lean_Syntax_getArg(x_3, x_1123); +x_1125 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__10; +lean_inc(x_1124); +x_1126 = l_Lean_Syntax_isOfKind(x_1124, x_1125); +if (x_1126 == 0) { -lean_dec(x_1131); +lean_dec(x_1124); if (lean_obj_tag(x_1) == 0) { -lean_object* x_1134; lean_object* x_1135; lean_object* x_1136; lean_object* x_1137; lean_object* x_1138; lean_object* x_1139; lean_object* x_1140; lean_object* x_1141; lean_object* x_1142; lean_object* x_1143; lean_object* x_1144; lean_object* x_1145; lean_object* x_1146; lean_object* x_1147; lean_object* x_1148; lean_object* x_1149; lean_object* x_1150; lean_object* x_1151; lean_object* x_1152; lean_object* x_1153; lean_object* x_1154; lean_object* x_1155; lean_object* x_1156; lean_object* x_1157; lean_object* x_1158; lean_object* x_1159; lean_object* x_1160; lean_object* x_1161; lean_object* x_1162; lean_object* x_1163; lean_object* x_1164; -x_1134 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); -x_1135 = lean_ctor_get(x_1134, 0); -lean_inc(x_1135); -x_1136 = lean_ctor_get(x_1134, 1); -lean_inc(x_1136); -lean_dec(x_1134); -x_1137 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_1136); -x_1138 = lean_ctor_get(x_1137, 0); -lean_inc(x_1138); -x_1139 = lean_ctor_get(x_1137, 1); -lean_inc(x_1139); -lean_dec(x_1137); -x_1140 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_1139); -x_1141 = lean_ctor_get(x_1140, 0); -lean_inc(x_1141); -x_1142 = lean_ctor_get(x_1140, 1); -lean_inc(x_1142); -lean_dec(x_1140); -x_1143 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; -x_1144 = l_Lean_addMacroScope(x_1141, x_1143, x_1138); -x_1145 = lean_box(0); -x_1146 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; +lean_object* x_1127; lean_object* x_1128; lean_object* x_1129; lean_object* x_1130; lean_object* x_1131; lean_object* x_1132; lean_object* x_1133; lean_object* x_1134; lean_object* x_1135; lean_object* x_1136; lean_object* x_1137; lean_object* x_1138; lean_object* x_1139; lean_object* x_1140; lean_object* x_1141; lean_object* x_1142; lean_object* x_1143; lean_object* x_1144; lean_object* x_1145; lean_object* x_1146; lean_object* x_1147; lean_object* x_1148; lean_object* x_1149; lean_object* x_1150; lean_object* x_1151; lean_object* x_1152; lean_object* x_1153; lean_object* x_1154; lean_object* x_1155; lean_object* x_1156; lean_object* x_1157; +x_1127 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_4, x_5, x_6); +x_1128 = lean_ctor_get(x_1127, 0); +lean_inc(x_1128); +x_1129 = lean_ctor_get(x_1127, 1); +lean_inc(x_1129); +lean_dec(x_1127); +x_1130 = l_Lean_Elab_Command_getCurrMacroScope(x_4, x_5, x_1129); +x_1131 = lean_ctor_get(x_1130, 0); +lean_inc(x_1131); +x_1132 = lean_ctor_get(x_1130, 1); +lean_inc(x_1132); +lean_dec(x_1130); +x_1133 = l_Lean_Elab_Command_getMainModule___rarg(x_5, x_1132); +x_1134 = lean_ctor_get(x_1133, 0); +lean_inc(x_1134); +x_1135 = lean_ctor_get(x_1133, 1); lean_inc(x_1135); -x_1147 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1147, 0, x_1135); +lean_dec(x_1133); +x_1136 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__4; +x_1137 = l_Lean_addMacroScope(x_1134, x_1136, x_1131); +x_1138 = lean_box(0); +x_1139 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__3; +lean_inc(x_1128); +x_1140 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1140, 0, x_1128); +lean_ctor_set(x_1140, 1, x_1139); +lean_ctor_set(x_1140, 2, x_1137); +lean_ctor_set(x_1140, 3, x_1138); +x_1141 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; +lean_inc(x_1128); +x_1142 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1142, 0, x_1128); +lean_ctor_set(x_1142, 1, x_1141); +x_1143 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; +lean_inc(x_3); +x_1144 = lean_array_push(x_1143, x_3); +x_1145 = lean_box(2); +x_1146 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__3; +x_1147 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1147, 0, x_1145); lean_ctor_set(x_1147, 1, x_1146); lean_ctor_set(x_1147, 2, x_1144); -lean_ctor_set(x_1147, 3, x_1145); -x_1148 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__5; -lean_inc(x_1135); +x_1148 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__6; x_1149 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1149, 0, x_1135); +lean_ctor_set(x_1149, 0, x_1128); lean_ctor_set(x_1149, 1, x_1148); -x_1150 = l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1; -lean_inc(x_3); -x_1151 = lean_array_push(x_1150, x_3); -x_1152 = lean_box(2); -x_1153 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; -x_1154 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1154, 0, x_1152); -lean_ctor_set(x_1154, 1, x_1153); -lean_ctor_set(x_1154, 2, x_1151); -x_1155 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__8; -x_1156 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1156, 0, x_1135); +x_1150 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; +x_1151 = lean_array_push(x_1150, x_1140); +x_1152 = lean_array_push(x_1151, x_1142); +x_1153 = lean_array_push(x_1152, x_1147); +x_1154 = lean_array_push(x_1153, x_1149); +x_1155 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__6; +x_1156 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1156, 0, x_1145); lean_ctor_set(x_1156, 1, x_1155); -x_1157 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__9; -x_1158 = lean_array_push(x_1157, x_1147); -x_1159 = lean_array_push(x_1158, x_1149); -x_1160 = lean_array_push(x_1159, x_1154); -x_1161 = lean_array_push(x_1160, x_1156); -x_1162 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__6; -x_1163 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1163, 0, x_1152); -lean_ctor_set(x_1163, 1, x_1162); -lean_ctor_set(x_1163, 2, x_1161); -x_1164 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_1142); -if (lean_obj_tag(x_1164) == 0) +lean_ctor_set(x_1156, 2, x_1154); +x_1157 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_2, x_4, x_5, x_1135); +if (lean_obj_tag(x_1157) == 0) { -uint8_t x_1165; -x_1165 = !lean_is_exclusive(x_1164); -if (x_1165 == 0) +uint8_t x_1158; +x_1158 = !lean_is_exclusive(x_1157); +if (x_1158 == 0) { -lean_object* x_1166; lean_object* x_1167; -x_1166 = lean_ctor_get(x_1164, 0); -x_1167 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1167, 0, x_1163); -lean_ctor_set(x_1167, 1, x_1166); -lean_ctor_set(x_1164, 0, x_1167); -return x_1164; +lean_object* x_1159; lean_object* x_1160; +x_1159 = lean_ctor_get(x_1157, 0); +x_1160 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1160, 0, x_1156); +lean_ctor_set(x_1160, 1, x_1159); +lean_ctor_set(x_1157, 0, x_1160); +return x_1157; } else { -lean_object* x_1168; lean_object* x_1169; lean_object* x_1170; lean_object* x_1171; -x_1168 = lean_ctor_get(x_1164, 0); -x_1169 = lean_ctor_get(x_1164, 1); -lean_inc(x_1169); -lean_inc(x_1168); -lean_dec(x_1164); -x_1170 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1170, 0, x_1163); -lean_ctor_set(x_1170, 1, x_1168); -x_1171 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1171, 0, x_1170); -lean_ctor_set(x_1171, 1, x_1169); -return x_1171; +lean_object* x_1161; lean_object* x_1162; lean_object* x_1163; lean_object* x_1164; +x_1161 = lean_ctor_get(x_1157, 0); +x_1162 = lean_ctor_get(x_1157, 1); +lean_inc(x_1162); +lean_inc(x_1161); +lean_dec(x_1157); +x_1163 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1163, 0, x_1156); +lean_ctor_set(x_1163, 1, x_1161); +x_1164 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1164, 0, x_1163); +lean_ctor_set(x_1164, 1, x_1162); +return x_1164; } } else { -uint8_t x_1172; -lean_dec(x_1163); -x_1172 = !lean_is_exclusive(x_1164); -if (x_1172 == 0) +uint8_t x_1165; +lean_dec(x_1156); +x_1165 = !lean_is_exclusive(x_1157); +if (x_1165 == 0) { -return x_1164; +return x_1157; } else { -lean_object* x_1173; lean_object* x_1174; lean_object* x_1175; -x_1173 = lean_ctor_get(x_1164, 0); -x_1174 = lean_ctor_get(x_1164, 1); -lean_inc(x_1174); -lean_inc(x_1173); -lean_dec(x_1164); -x_1175 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1175, 0, x_1173); -lean_ctor_set(x_1175, 1, x_1174); -return x_1175; +lean_object* x_1166; lean_object* x_1167; lean_object* x_1168; +x_1166 = lean_ctor_get(x_1157, 0); +x_1167 = lean_ctor_get(x_1157, 1); +lean_inc(x_1167); +lean_inc(x_1166); +lean_dec(x_1157); +x_1168 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1168, 0, x_1166); +lean_ctor_set(x_1168, 1, x_1167); +return x_1168; } } } else { -lean_object* x_1176; lean_object* x_1177; +lean_object* x_1169; lean_object* x_1170; lean_dec(x_2); -x_1176 = lean_ctor_get(x_1, 0); -lean_inc(x_1176); +x_1169 = lean_ctor_get(x_1, 0); +lean_inc(x_1169); lean_dec(x_1); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_1177 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_1176, x_4, x_5, x_6); -if (lean_obj_tag(x_1177) == 0) -{ -lean_object* x_1178; lean_object* x_1179; lean_object* x_1180; -x_1178 = lean_ctor_get(x_1177, 0); -lean_inc(x_1178); -x_1179 = lean_ctor_get(x_1177, 1); -lean_inc(x_1179); -lean_dec(x_1177); -x_1180 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_1178, x_4, x_5, x_1179); +x_1170 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode(x_3, x_1169, x_4, x_5, x_6); +if (lean_obj_tag(x_1170) == 0) +{ +lean_object* x_1171; lean_object* x_1172; lean_object* x_1173; +x_1171 = lean_ctor_get(x_1170, 0); +lean_inc(x_1171); +x_1172 = lean_ctor_get(x_1170, 1); +lean_inc(x_1172); +lean_dec(x_1170); +x_1173 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_1171, x_4, x_5, x_1172); lean_dec(x_5); lean_dec(x_4); -return x_1180; +return x_1173; } else { -uint8_t x_1181; +uint8_t x_1174; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_1181 = !lean_is_exclusive(x_1177); -if (x_1181 == 0) +x_1174 = !lean_is_exclusive(x_1170); +if (x_1174 == 0) { -return x_1177; +return x_1170; } else { -lean_object* x_1182; lean_object* x_1183; lean_object* x_1184; -x_1182 = lean_ctor_get(x_1177, 0); -x_1183 = lean_ctor_get(x_1177, 1); -lean_inc(x_1183); -lean_inc(x_1182); -lean_dec(x_1177); -x_1184 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1184, 0, x_1182); -lean_ctor_set(x_1184, 1, x_1183); -return x_1184; +lean_object* x_1175; lean_object* x_1176; lean_object* x_1177; +x_1175 = lean_ctor_get(x_1170, 0); +x_1176 = lean_ctor_get(x_1170, 1); +lean_inc(x_1176); +lean_inc(x_1175); +lean_dec(x_1170); +x_1177 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1177, 0, x_1175); +lean_ctor_set(x_1177, 1, x_1176); +return x_1177; } } } } else { -lean_object* x_1185; lean_object* x_1186; +lean_object* x_1178; lean_object* x_1179; lean_dec(x_1); -x_1185 = lean_alloc_closure((void*)(l_Lean_Elab_Command_strLitToPattern___boxed), 3, 1); -lean_closure_set(x_1185, 0, x_1131); +x_1178 = lean_alloc_closure((void*)(l_Lean_Elab_Command_strLitToPattern___boxed), 3, 1); +lean_closure_set(x_1178, 0, x_1124); lean_inc(x_5); lean_inc(x_4); -x_1186 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__10(x_1185, x_4, x_5, x_6); -if (lean_obj_tag(x_1186) == 0) -{ -lean_object* x_1187; lean_object* x_1188; lean_object* x_1189; lean_object* x_1190; lean_object* x_1191; lean_object* x_1192; lean_object* x_1193; lean_object* x_1194; lean_object* x_1195; lean_object* x_1196; lean_object* x_1197; lean_object* x_1198; lean_object* x_1199; -x_1187 = lean_ctor_get(x_1186, 0); -lean_inc(x_1187); -x_1188 = lean_ctor_get(x_1186, 1); -lean_inc(x_1188); -lean_dec(x_1186); -x_1189 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__9; -x_1190 = lean_array_push(x_1189, x_1187); -x_1191 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__28; -x_1192 = lean_array_push(x_1190, x_1191); -x_1193 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__30; -x_1194 = lean_array_push(x_1192, x_1193); -x_1195 = lean_array_push(x_1194, x_2); -x_1196 = lean_box(2); -x_1197 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__26; -x_1198 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1198, 0, x_1196); -lean_ctor_set(x_1198, 1, x_1197); -lean_ctor_set(x_1198, 2, x_1195); -x_1199 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_1198, x_4, x_5, x_1188); +x_1179 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__10(x_1178, x_4, x_5, x_6); +if (lean_obj_tag(x_1179) == 0) +{ +lean_object* x_1180; lean_object* x_1181; lean_object* x_1182; lean_object* x_1183; lean_object* x_1184; lean_object* x_1185; lean_object* x_1186; lean_object* x_1187; lean_object* x_1188; lean_object* x_1189; lean_object* x_1190; lean_object* x_1191; lean_object* x_1192; +x_1180 = lean_ctor_get(x_1179, 0); +lean_inc(x_1180); +x_1181 = lean_ctor_get(x_1179, 1); +lean_inc(x_1181); +lean_dec(x_1179); +x_1182 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7; +x_1183 = lean_array_push(x_1182, x_1180); +x_1184 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__28; +x_1185 = lean_array_push(x_1183, x_1184); +x_1186 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__30; +x_1187 = lean_array_push(x_1185, x_1186); +x_1188 = lean_array_push(x_1187, x_2); +x_1189 = lean_box(2); +x_1190 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__26; +x_1191 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1191, 0, x_1189); +lean_ctor_set(x_1191, 1, x_1190); +lean_ctor_set(x_1191, 2, x_1188); +x_1192 = l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__1(x_3, x_1191, x_4, x_5, x_1181); lean_dec(x_5); lean_dec(x_4); -return x_1199; +return x_1192; } else { -uint8_t x_1200; +uint8_t x_1193; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_1200 = !lean_is_exclusive(x_1186); -if (x_1200 == 0) +x_1193 = !lean_is_exclusive(x_1179); +if (x_1193 == 0) { -return x_1186; +return x_1179; } else { -lean_object* x_1201; lean_object* x_1202; lean_object* x_1203; -x_1201 = lean_ctor_get(x_1186, 0); -x_1202 = lean_ctor_get(x_1186, 1); -lean_inc(x_1202); -lean_inc(x_1201); -lean_dec(x_1186); -x_1203 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1203, 0, x_1201); -lean_ctor_set(x_1203, 1, x_1202); -return x_1203; +lean_object* x_1194; lean_object* x_1195; lean_object* x_1196; +x_1194 = lean_ctor_get(x_1179, 0); +x_1195 = lean_ctor_get(x_1179, 1); +lean_inc(x_1195); +lean_inc(x_1194); +lean_dec(x_1179); +x_1196 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1196, 0, x_1194); +lean_ctor_set(x_1196, 1, x_1195); +return x_1196; } } } @@ -13831,165 +13898,164 @@ return x_16; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; x_17 = lean_unsigned_to_nat(0u); x_18 = l_Lean_Syntax_getArg(x_7, x_17); -x_19 = l_Lean_nullKind; -x_20 = lean_unsigned_to_nat(2u); +x_19 = lean_unsigned_to_nat(2u); lean_inc(x_18); -x_21 = l_Lean_Syntax_isNodeOf(x_18, x_19, x_20); -if (x_21 == 0) +x_20 = l_Lean_Syntax_matchesNull(x_18, x_19); +if (x_20 == 0) { -uint8_t x_22; -x_22 = l_Lean_Syntax_isNodeOf(x_18, x_19, x_17); -if (x_22 == 0) +uint8_t x_21; +x_21 = l_Lean_Syntax_matchesNull(x_18, x_17); +if (x_21 == 0) { -lean_object* x_23; uint8_t x_24; +lean_object* x_22; uint8_t x_23; lean_dec(x_7); lean_dec(x_3); lean_dec(x_2); -x_23 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_expandMacroArg___spec__1___rarg(x_8); -x_24 = !lean_is_exclusive(x_23); -if (x_24 == 0) +x_22 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_expandMacroArg___spec__1___rarg(x_8); +x_23 = !lean_is_exclusive(x_22); +if (x_23 == 0) { -return x_23; +return x_22; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_23, 0); -x_26 = lean_ctor_get(x_23, 1); -lean_inc(x_26); +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_22, 0); +x_25 = lean_ctor_get(x_22, 1); lean_inc(x_25); -lean_dec(x_23); -x_27 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_27, 0, x_25); -lean_ctor_set(x_27, 1, x_26); -return x_27; +lean_inc(x_24); +lean_dec(x_22); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; } } else { -lean_object* x_28; lean_object* x_29; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_28 = lean_unsigned_to_nat(1u); -x_29 = l_Lean_Syntax_getArg(x_7, x_28); +lean_object* x_27; lean_object* x_28; lean_object* x_29; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_27 = lean_unsigned_to_nat(1u); +x_28 = l_Lean_Syntax_getArg(x_7, x_27); lean_dec(x_7); -x_30 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_2, x_3, x_8); -x_31 = lean_ctor_get(x_30, 0); +x_29 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_2, x_3, x_8); +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_29, 1); lean_inc(x_31); -x_32 = lean_ctor_get(x_30, 1); -lean_inc(x_32); -lean_dec(x_30); -x_33 = l_Lean_Elab_Command_getCurrMacroScope(x_2, x_3, x_32); -x_34 = lean_ctor_get(x_33, 0); +lean_dec(x_29); +x_32 = l_Lean_Elab_Command_getCurrMacroScope(x_2, x_3, x_31); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); lean_inc(x_34); -x_35 = lean_ctor_get(x_33, 1); -lean_inc(x_35); -lean_dec(x_33); -x_36 = l_Lean_Elab_Command_getMainModule___rarg(x_3, x_35); -x_37 = lean_ctor_get(x_36, 0); +lean_dec(x_32); +x_35 = l_Lean_Elab_Command_getMainModule___rarg(x_3, x_34); +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_35, 1); lean_inc(x_37); -x_38 = lean_ctor_get(x_36, 1); -lean_inc(x_38); -lean_dec(x_36); -x_39 = l_Lean_Elab_Command_expandMacroArg___closed__9; -x_40 = l_Lean_addMacroScope(x_37, x_39, x_34); -x_41 = lean_box(0); -x_42 = l_Lean_Elab_Command_expandMacroArg___closed__8; -x_43 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_43, 0, x_31); -lean_ctor_set(x_43, 1, x_42); -lean_ctor_set(x_43, 2, x_40); -lean_ctor_set(x_43, 3, x_41); -x_44 = lean_box(0); +lean_dec(x_35); +x_38 = l_Lean_Elab_Command_expandMacroArg___closed__9; +x_39 = l_Lean_addMacroScope(x_36, x_38, x_33); +x_40 = lean_box(0); +x_41 = l_Lean_Elab_Command_expandMacroArg___closed__8; +x_42 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_42, 0, x_30); +lean_ctor_set(x_42, 1, x_41); +lean_ctor_set(x_42, 2, x_39); +lean_ctor_set(x_42, 3, x_40); +x_43 = lean_box(0); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_42); +lean_ctor_set(x_44, 1, x_28); x_45 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_45, 0, x_43); -lean_ctor_set(x_45, 1, x_29); -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -x_47 = lean_apply_4(x_9, x_46, x_2, x_3, x_38); -return x_47; +lean_ctor_set(x_45, 1, x_44); +x_46 = lean_apply_4(x_9, x_45, x_2, x_3, x_37); +return x_46; } } else { -lean_object* x_48; lean_object* x_49; uint8_t x_50; -x_48 = l_Lean_Syntax_getArg(x_18, x_17); +lean_object* x_47; lean_object* x_48; uint8_t x_49; +x_47 = l_Lean_Syntax_getArg(x_18, x_17); lean_dec(x_18); -x_49 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___closed__9; -lean_inc(x_48); -x_50 = l_Lean_Syntax_isOfKind(x_48, x_49); -if (x_50 == 0) +x_48 = l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___closed__9; +lean_inc(x_47); +x_49 = l_Lean_Syntax_isOfKind(x_47, x_48); +if (x_49 == 0) { -lean_object* x_51; uint8_t x_52; -lean_dec(x_48); +lean_object* x_50; uint8_t x_51; +lean_dec(x_47); lean_dec(x_7); lean_dec(x_3); lean_dec(x_2); -x_51 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_expandMacroArg___spec__1___rarg(x_8); -x_52 = !lean_is_exclusive(x_51); -if (x_52 == 0) +x_50 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_expandMacroArg___spec__1___rarg(x_8); +x_51 = !lean_is_exclusive(x_50); +if (x_51 == 0) { -return x_51; +return x_50; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_53 = lean_ctor_get(x_51, 0); -x_54 = lean_ctor_get(x_51, 1); -lean_inc(x_54); +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_50, 0); +x_53 = lean_ctor_get(x_50, 1); lean_inc(x_53); -lean_dec(x_51); -x_55 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_55, 0, x_53); -lean_ctor_set(x_55, 1, x_54); -return x_55; +lean_inc(x_52); +lean_dec(x_50); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +return x_54; } } else { -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_56 = lean_unsigned_to_nat(1u); -x_57 = l_Lean_Syntax_getArg(x_7, x_56); +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_55 = lean_unsigned_to_nat(1u); +x_56 = l_Lean_Syntax_getArg(x_7, x_55); lean_dec(x_7); -lean_inc(x_48); -x_58 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_58, 0, x_48); +lean_inc(x_47); +x_57 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_57, 0, x_47); +x_58 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_58, 0, x_47); +lean_ctor_set(x_58, 1, x_56); x_59 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_59, 0, x_48); -lean_ctor_set(x_59, 1, x_57); -x_60 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_60, 0, x_58); -lean_ctor_set(x_60, 1, x_59); -x_61 = lean_apply_4(x_9, x_60, x_2, x_3, x_8); -return x_61; +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_58); +x_60 = lean_apply_4(x_9, x_59, x_2, x_3, x_8); +return x_60; } } } } else { -uint8_t x_62; +uint8_t x_61; lean_dec(x_3); lean_dec(x_2); -x_62 = !lean_is_exclusive(x_6); -if (x_62 == 0) +x_61 = !lean_is_exclusive(x_6); +if (x_61 == 0) { return x_6; } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_63 = lean_ctor_get(x_6, 0); -x_64 = lean_ctor_get(x_6, 1); -lean_inc(x_64); +lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_62 = lean_ctor_get(x_6, 0); +x_63 = lean_ctor_get(x_6, 1); lean_inc(x_63); +lean_inc(x_62); lean_dec(x_6); -x_65 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_65, 0, x_63); -lean_ctor_set(x_65, 1, x_64); -return x_65; +x_64 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_64, 0, x_62); +lean_ctor_set(x_64, 1, x_63); +return x_64; } } } @@ -14065,6 +14131,10 @@ l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__8 = _ini lean_mark_persistent(l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__8); l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__9 = _init_l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__9(); lean_mark_persistent(l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__9); +l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__10 = _init_l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__10(); +lean_mark_persistent(l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__10); +l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__11 = _init_l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__11(); +lean_mark_persistent(l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___lambda__2___closed__11); l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___closed__1 = _init_l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___closed__1); l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___closed__2 = _init_l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___closed__2(); @@ -14089,6 +14159,10 @@ l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___closed__11 = _init_l_Lean_El lean_mark_persistent(l_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___closed__11); l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1 = _init_l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__1); +l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__2 = _init_l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__2); +l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__3 = _init_l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__3(); +lean_mark_persistent(l_Lean_Elab_Command_expandMacroArg_mkSplicePat___closed__3); l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__2___closed__1 = _init_l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__2___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__2___closed__1); l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__2___closed__2 = _init_l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__2___closed__2(); @@ -14117,10 +14191,6 @@ l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__6 = _ini lean_mark_persistent(l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__6); l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7 = _init_l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7(); lean_mark_persistent(l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__7); -l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__8 = _init_l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__8(); -lean_mark_persistent(l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__8); -l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__9 = _init_l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__9(); -lean_mark_persistent(l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__3___closed__9); l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__1 = _init_l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__1); l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__2 = _init_l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/MacroRules.c b/stage0/stdlib/Lean/Elab/MacroRules.c index 5e01b1d2e0e..c35ee16879e 100644 --- a/stage0/stdlib/Lean/Elab/MacroRules.c +++ b/stage0/stdlib/Lean/Elab/MacroRules.c @@ -15,59 +15,66 @@ extern "C" { #endif uint8_t l_Lean_Syntax_isQuot(lean_object*); static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__34; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__3; lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabMacroRules___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__48; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__11; static lean_object* l___regBuiltin_Lean_Elab_Command_elabMacroRules_declRange___closed__7; lean_object* l_Lean_stringToMessageData(lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__15; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__50; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__22; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__18; lean_object* lean_mk_empty_array_with_capacity(lean_object*); static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__15; lean_object* l_Lean_Syntax_getHeadInfo_x3f(lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__5; lean_object* l_Lean_Elab_Command_expandNoKindMacroRulesAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_nullKind; -lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabCommand___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__14; +static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabMacroRulesAux___spec__1___rarg___closed__2; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_name_mk_string(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabMacroRules_declRange(lean_object*); static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__14; lean_object* lean_array_uget(lean_object*, size_t); static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__28; lean_object* l_Array_append___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabMacroRulesAux___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__30; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__5; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__5; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__4; extern lean_object* l_Lean_Elab_Command_commandElabAttribute; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__33; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabMacroRules(lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__9; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__13; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__3; static lean_object* l_Lean_Elab_Command_elabMacroRules___lambda__1___closed__3; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__23; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__8; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__41; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__52; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__8; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__4; lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Command_elabAuxDef___spec__4(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabMacroRules_declRange___closed__2; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__1; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__62; uint8_t lean_name_eq(lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__4; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__54; +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabMacroRulesAux___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkIdentFrom(lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__9; static lean_object* l___regBuiltin_Lean_Elab_Command_elabMacroRules_declRange___closed__1; lean_object* lean_array_push(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__2; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__45; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__2; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__9; +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabMacroRulesAux___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__27; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__7; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__40; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__25; lean_object* lean_string_utf8_byte_size(lean_object*); @@ -77,95 +84,97 @@ uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabMacroRules___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__29; uint8_t l_Lean_Elab_Command_checkRuleKind(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabMacroRules___lambda__2___closed__1; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__37; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__53; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__42; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabMacroRules___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_resolveSyntaxKind(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Syntax_SepArray_ofElems(lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___boxed(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_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__55; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabMacroRules___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabMacroRules___lambda__1___closed__2; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__8; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__2; +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__1; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__3; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__56; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__43; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__13; +lean_object* l_Lean_replaceRef(lean_object*, lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__12; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__32; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__10; static lean_object* l_Lean_Elab_Command_elabMacroRules___closed__1; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__6; +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabMacroRulesAux___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__20; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabMacroRules(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabMacroRulesAux___spec__1___rarg___closed__1; lean_object* l_Lean_Syntax_getId(lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__5; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__15; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__46; lean_object* l_Lean_Elab_Command_adaptExpander(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_choiceKind; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__16; 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*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabMacroRules_declRange___closed__5; static lean_object* l_Lean_Elab_Command_elabMacroRules___lambda__3___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Command_elabMacroRules___closed__4; extern lean_object* l_Lean_instInhabitedSyntax; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__14; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__59; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__11; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__11; size_t lean_usize_of_nat(lean_object*); lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabMacroRulesAux___spec__4(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabMacroRulesAux___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getQuotContent(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabMacroRules_declRange___closed__4; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2(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_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabMacroRulesAux___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__16; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__61; static lean_object* l___regBuiltin_Lean_Elab_Command_elabMacroRules___closed__2; lean_object* l_Lean_Syntax_setArg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__49; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__8; +lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabMacroRules___lambda__3___closed__4; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabMacroRulesAux___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabMacroRules___lambda__2___closed__2; lean_object* l_Lean_Elab_Command_getCurrMacroScope(lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__7; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__31; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__1; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__7; lean_object* l_Lean_Syntax_getArgs(lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__7; lean_object* l_Lean_Syntax_getKind(lean_object*); static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__24; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabMacroRules___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_inferMacroRulesAltKind___spec__2___rarg(lean_object*); static lean_object* l_Lean_Elab_Command_elabMacroRules___lambda__1___closed__1; static lean_object* l_Lean_Elab_Command_elabMacroRules___lambda__3___closed__1; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__2; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__7; lean_object* l_Lean_Elab_Command_getMainModule___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__58; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__26; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__3; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__4; static lean_object* l_Lean_Elab_Command_elabMacroRules___lambda__1___closed__4; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__39; +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabMacroRulesAux___spec__1___rarg(lean_object*); +lean_object* l_Lean_Elab_Command_getRef(lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__3; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__35; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__10; uint8_t l_Lean_Syntax_isNone(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabMacroRules_declRange___closed__6; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__38; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__21; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabMacroRulesAux___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabMacroRules___lambda__2___closed__4; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__1___boxed(lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__12; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__10; +extern lean_object* l_Lean_Elab_unsupportedSyntaxExceptionId; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__5; lean_object* lean_mk_syntax_ident(lean_object*); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabMacroRulesAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__19; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__51; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__44; @@ -174,19 +183,171 @@ static lean_object* l___regBuiltin_Lean_Elab_Command_elabMacroRules___closed__1; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__47; lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__14___rarg(lean_object*); static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__1; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__6; static lean_object* l_Lean_Elab_Command_elabMacroRules___lambda__1___closed__5; static lean_object* l_Lean_Elab_Command_elabMacroRules___lambda__2___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Command_elabMacroRules___closed__3; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__2; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Command_elabMacroRules_declRange___closed__3; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__12; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__17; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__4; static lean_object* l_Lean_Elab_Command_elabMacroRules___lambda__3___closed__2; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__57; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__6; +lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__8; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabMacroRulesAux___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +static lean_object* _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabMacroRulesAux___spec__1___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Elab_unsupportedSyntaxExceptionId; +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabMacroRulesAux___spec__1___rarg___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabMacroRulesAux___spec__1___rarg___closed__1; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabMacroRulesAux___spec__1___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabMacroRulesAux___spec__1___rarg___closed__2; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabMacroRulesAux___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabMacroRulesAux___spec__1___rarg), 1, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabMacroRulesAux___spec__3(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; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_5 = l_Lean_Elab_Command_getRef(x_2, x_3, x_4); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_5, 1); +lean_inc(x_7); +lean_dec(x_5); +x_8 = lean_ctor_get(x_2, 4); +lean_inc(x_8); +lean_inc(x_8); +x_9 = l_Lean_Elab_getBetterRef(x_6, x_8); +lean_dec(x_6); +x_10 = l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1(x_1, x_2, x_3, x_7); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1(x_11, x_8, x_2, x_3, x_12); +lean_dec(x_2); +lean_dec(x_8); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_9); +lean_ctor_set(x_16, 1, x_15); +lean_ctor_set_tag(x_13, 1); +lean_ctor_set(x_13, 0, x_16); +return x_13; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_17 = lean_ctor_get(x_13, 0); +x_18 = lean_ctor_get(x_13, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_13); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_9); +lean_ctor_set(x_19, 1, x_17); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_18); +return x_20; +} +} +} +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabMacroRulesAux___spec__2(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; lean_object* x_9; uint8_t x_10; +x_6 = l_Lean_Elab_Command_getRef(x_3, x_4, x_5); +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_6, 1); +lean_inc(x_8); +lean_dec(x_6); +x_9 = l_Lean_replaceRef(x_1, x_7); +lean_dec(x_7); +lean_dec(x_1); +x_10 = !lean_is_exclusive(x_3); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_3, 6); +lean_dec(x_11); +lean_ctor_set(x_3, 6, x_9); +x_12 = l_Lean_throwError___at_Lean_Elab_Command_elabMacroRulesAux___spec__3(x_2, x_3, x_4, x_8); +lean_dec(x_4); +return x_12; +} +else +{ +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; lean_object* x_20; lean_object* x_21; +x_13 = lean_ctor_get(x_3, 0); +x_14 = lean_ctor_get(x_3, 1); +x_15 = lean_ctor_get(x_3, 2); +x_16 = lean_ctor_get(x_3, 3); +x_17 = lean_ctor_get(x_3, 4); +x_18 = lean_ctor_get(x_3, 5); +x_19 = lean_ctor_get(x_3, 7); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_3); +x_20 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_20, 0, x_13); +lean_ctor_set(x_20, 1, x_14); +lean_ctor_set(x_20, 2, x_15); +lean_ctor_set(x_20, 3, x_16); +lean_ctor_set(x_20, 4, x_17); +lean_ctor_set(x_20, 5, x_18); +lean_ctor_set(x_20, 6, x_9); +lean_ctor_set(x_20, 7, x_19); +x_21 = l_Lean_throwError___at_Lean_Elab_Command_elabMacroRulesAux___spec__3(x_2, x_20, x_4, x_8); +lean_dec(x_4); +return x_21; +} +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabMacroRulesAux___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { _start: { uint8_t x_7; @@ -232,14 +393,14 @@ return x_16; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__1(lean_object* x_1, lean_object* x_2) { _start: { lean_inc(x_1); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__1() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__1() { _start: { lean_object* x_1; @@ -247,7 +408,7 @@ x_1 = lean_mk_string_from_bytes("|", 1); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__2() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__2() { _start: { lean_object* x_1; @@ -255,17 +416,17 @@ x_1 = lean_mk_string_from_bytes("null", 4); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__3() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__2; +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__2; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__4() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__4() { _start: { lean_object* x_1; lean_object* x_2; @@ -274,15 +435,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__5() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(",", 1); -return x_1; -} -} -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__6() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__5() { _start: { lean_object* x_1; lean_object* x_2; @@ -291,7 +444,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__7() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__6() { _start: { lean_object* x_1; @@ -299,7 +452,7 @@ x_1 = lean_mk_string_from_bytes("=>", 2); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__8() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__7() { _start: { lean_object* x_1; lean_object* x_2; @@ -308,7 +461,25 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__9() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("choice", 6); +return x_1; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__8; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__10() { _start: { lean_object* x_1; @@ -316,16 +487,16 @@ x_1 = lean_mk_string_from_bytes("invalid macro_rules alternative, unexpected syn return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__10() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__9; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__10; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__11() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__12() { _start: { lean_object* x_1; @@ -333,16 +504,16 @@ x_1 = lean_mk_string_from_bytes("'", 1); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__12() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__13() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__11; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__12; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__13() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -354,7 +525,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__14() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__15() { _start: { lean_object* x_1; @@ -362,16 +533,16 @@ x_1 = lean_mk_string_from_bytes("invalid macro_rules alternative, expected synta return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__15() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__16() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__14; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__15; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___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_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___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_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; lean_object* x_12; uint8_t x_13; @@ -383,7 +554,7 @@ x_13 = l_Lean_Elab_Command_checkRuleKind(x_12, x_5); if (x_13 == 0) { lean_object* x_14; uint8_t x_15; -x_14 = l_Lean_choiceKind; +x_14 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__9; x_15 = lean_name_eq(x_12, x_14); if (x_15 == 0) { @@ -396,15 +567,15 @@ lean_dec(x_2); lean_dec(x_1); x_16 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_16, 0, x_12); -x_17 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__10; +x_17 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__11; x_18 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_18, 0, x_17); lean_ctor_set(x_18, 1, x_16); -x_19 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__12; +x_19 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__13; x_20 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_20, 0, x_18); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_throwErrorAt___at_Lean_Elab_Command_elabCommand___spec__11(x_6, x_20, x_8, x_9, x_10); +x_21 = l_Lean_throwErrorAt___at_Lean_Elab_Command_elabMacroRulesAux___spec__2(x_6, x_20, x_8, x_9, x_10); return x_21; } else @@ -417,8 +588,8 @@ x_23 = lean_array_get_size(x_22); x_24 = lean_usize_of_nat(x_23); lean_dec(x_23); x_25 = 0; -x_26 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__13; -x_27 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabMacroRulesAux___spec__1(x_5, x_26, x_22, x_24, x_25, x_26); +x_26 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__14; +x_27 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabMacroRulesAux___spec__4(x_5, x_26, x_22, x_24, x_25, x_26); lean_dec(x_22); x_28 = lean_ctor_get(x_27, 0); lean_inc(x_28); @@ -432,15 +603,15 @@ lean_dec(x_2); lean_dec(x_1); x_29 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_29, 0, x_5); -x_30 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__15; +x_30 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__16; x_31 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_31, 0, x_30); lean_ctor_set(x_31, 1, x_29); -x_32 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__12; +x_32 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__13; x_33 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_33, 0, x_31); lean_ctor_set(x_33, 1, x_32); -x_34 = l_Lean_throwErrorAt___at_Lean_Elab_Command_elabCommand___spec__11(x_6, x_33, x_8, x_9, x_10); +x_34 = l_Lean_throwErrorAt___at_Lean_Elab_Command_elabMacroRulesAux___spec__2(x_6, x_33, x_8, x_9, x_10); return x_34; } else @@ -471,99 +642,93 @@ lean_dec(x_9); x_46 = !lean_is_exclusive(x_45); if (x_46 == 0) { -lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; lean_object* x_67; +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; x_47 = lean_ctor_get(x_45, 0); lean_dec(x_47); -x_48 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__1; +x_48 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__1; lean_inc(x_41); x_49 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_49, 0, x_41); lean_ctor_set(x_49, 1, x_48); -x_50 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__5; -x_51 = l_Lean_Syntax_SepArray_ofElems(x_50, x_39); -lean_dec(x_39); -x_52 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__4; -x_53 = l_Array_append___rarg(x_52, x_51); -x_54 = lean_box(2); -x_55 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__3; -x_56 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_56, 0, x_54); -lean_ctor_set(x_56, 1, x_55); -lean_ctor_set(x_56, 2, x_53); -x_57 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__6; -x_58 = lean_array_push(x_57, x_56); -x_59 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_59, 0, x_54); -lean_ctor_set(x_59, 1, x_55); -lean_ctor_set(x_59, 2, x_58); -x_60 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__7; -x_61 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_61, 0, x_41); -lean_ctor_set(x_61, 1, x_60); -x_62 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__8; -x_63 = lean_array_push(x_62, x_49); -x_64 = lean_array_push(x_63, x_59); -x_65 = lean_array_push(x_64, x_61); -x_66 = lean_array_push(x_65, x_3); -x_67 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_67, 0, x_54); -lean_ctor_set(x_67, 1, x_4); -lean_ctor_set(x_67, 2, x_66); -lean_ctor_set(x_45, 0, x_67); +x_50 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__4; +x_51 = l_Array_append___rarg(x_50, x_39); +x_52 = lean_box(2); +x_53 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__3; +x_54 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +lean_ctor_set(x_54, 2, x_51); +x_55 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__5; +x_56 = lean_array_push(x_55, x_54); +x_57 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_57, 0, x_52); +lean_ctor_set(x_57, 1, x_53); +lean_ctor_set(x_57, 2, x_56); +x_58 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__6; +x_59 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_59, 0, x_41); +lean_ctor_set(x_59, 1, x_58); +x_60 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__7; +x_61 = lean_array_push(x_60, x_49); +x_62 = lean_array_push(x_61, x_57); +x_63 = lean_array_push(x_62, x_59); +x_64 = lean_array_push(x_63, x_3); +x_65 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_65, 0, x_52); +lean_ctor_set(x_65, 1, x_4); +lean_ctor_set(x_65, 2, x_64); +lean_ctor_set(x_45, 0, x_65); return x_45; } else { -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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_68 = lean_ctor_get(x_45, 1); -lean_inc(x_68); +lean_object* x_66; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_66 = lean_ctor_get(x_45, 1); +lean_inc(x_66); lean_dec(x_45); -x_69 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__1; +x_67 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__1; lean_inc(x_41); -x_70 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_70, 0, x_41); -lean_ctor_set(x_70, 1, x_69); -x_71 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__5; -x_72 = l_Lean_Syntax_SepArray_ofElems(x_71, x_39); -lean_dec(x_39); -x_73 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__4; -x_74 = l_Array_append___rarg(x_73, x_72); -x_75 = lean_box(2); -x_76 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__3; -x_77 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_77, 0, x_75); -lean_ctor_set(x_77, 1, x_76); -lean_ctor_set(x_77, 2, x_74); -x_78 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__6; -x_79 = lean_array_push(x_78, x_77); -x_80 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_80, 0, x_75); -lean_ctor_set(x_80, 1, x_76); -lean_ctor_set(x_80, 2, x_79); -x_81 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__7; -x_82 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_82, 0, x_41); -lean_ctor_set(x_82, 1, x_81); -x_83 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__8; -x_84 = lean_array_push(x_83, x_70); -x_85 = lean_array_push(x_84, x_80); -x_86 = lean_array_push(x_85, x_82); -x_87 = lean_array_push(x_86, x_3); -x_88 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_88, 0, x_75); -lean_ctor_set(x_88, 1, x_4); -lean_ctor_set(x_88, 2, x_87); -x_89 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_89, 0, x_88); -lean_ctor_set(x_89, 1, x_68); -return x_89; +x_68 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_68, 0, x_41); +lean_ctor_set(x_68, 1, x_67); +x_69 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__4; +x_70 = l_Array_append___rarg(x_69, x_39); +x_71 = lean_box(2); +x_72 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__3; +x_73 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_73, 0, x_71); +lean_ctor_set(x_73, 1, x_72); +lean_ctor_set(x_73, 2, x_70); +x_74 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__5; +x_75 = lean_array_push(x_74, x_73); +x_76 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_76, 0, x_71); +lean_ctor_set(x_76, 1, x_72); +lean_ctor_set(x_76, 2, x_75); +x_77 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__6; +x_78 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_78, 0, x_41); +lean_ctor_set(x_78, 1, x_77); +x_79 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__7; +x_80 = lean_array_push(x_79, x_68); +x_81 = lean_array_push(x_80, x_76); +x_82 = lean_array_push(x_81, x_78); +x_83 = lean_array_push(x_82, x_3); +x_84 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_84, 0, x_71); +lean_ctor_set(x_84, 1, x_4); +lean_ctor_set(x_84, 2, x_83); +x_85 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_85, 0, x_84); +lean_ctor_set(x_85, 1, x_66); +return x_85; } } } } else { -lean_object* x_90; +lean_object* x_86; lean_dec(x_12); lean_dec(x_11); lean_dec(x_9); @@ -573,14 +738,14 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_90 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_90, 0, x_6); -lean_ctor_set(x_90, 1, x_10); -return x_90; +x_86 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_86, 0, x_6); +lean_ctor_set(x_86, 1, x_10); +return x_86; } } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__1() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__1() { _start: { lean_object* x_1; @@ -588,17 +753,17 @@ x_1 = lean_mk_string_from_bytes("Lean", 4); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__2() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__1; +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__3() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__3() { _start: { lean_object* x_1; @@ -606,17 +771,17 @@ x_1 = lean_mk_string_from_bytes("Parser", 6); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__4() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__2; -x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__3; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__2; +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__5() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__5() { _start: { lean_object* x_1; @@ -624,17 +789,17 @@ x_1 = lean_mk_string_from_bytes("Term", 4); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__6() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__4; -x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__5; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__4; +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__7() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__7() { _start: { lean_object* x_1; @@ -642,17 +807,17 @@ x_1 = lean_mk_string_from_bytes("matchAlt", 8); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__8() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__6; -x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__7; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__6; +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__7; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { uint8_t x_8; @@ -674,7 +839,7 @@ lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint x_10 = lean_array_uget(x_4, x_3); x_11 = lean_unsigned_to_nat(0u); x_12 = lean_array_uset(x_4, x_3, x_11); -x_13 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__8; +x_13 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__8; lean_inc(x_10); x_14 = l_Lean_Syntax_isOfKind(x_10, x_13); if (x_14 == 0) @@ -685,7 +850,7 @@ lean_dec(x_10); lean_dec(x_6); lean_dec(x_5); lean_dec(x_1); -x_15 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__14___rarg(x_7); +x_15 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabMacroRulesAux___spec__1___rarg(x_7); x_16 = !lean_is_exclusive(x_15); if (x_16 == 0) { @@ -707,132 +872,131 @@ return x_19; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +lean_object* x_20; lean_object* x_21; uint8_t x_22; x_20 = lean_unsigned_to_nat(1u); x_21 = l_Lean_Syntax_getArg(x_10, x_20); -x_22 = l_Lean_nullKind; lean_inc(x_21); -x_23 = l_Lean_Syntax_isNodeOf(x_21, x_22, x_20); -if (x_23 == 0) +x_22 = l_Lean_Syntax_matchesNull(x_21, x_20); +if (x_22 == 0) { -lean_object* x_24; uint8_t x_25; +lean_object* x_23; uint8_t x_24; lean_dec(x_21); lean_dec(x_12); lean_dec(x_10); lean_dec(x_6); lean_dec(x_5); lean_dec(x_1); -x_24 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__14___rarg(x_7); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_23 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabMacroRulesAux___spec__1___rarg(x_7); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -return x_24; +return x_23; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_24, 0); -x_27 = lean_ctor_get(x_24, 1); -lean_inc(x_27); +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_23, 0); +x_26 = lean_ctor_get(x_23, 1); lean_inc(x_26); -lean_dec(x_24); -x_28 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -return x_28; +lean_inc(x_25); +lean_dec(x_23); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; } } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; -x_29 = l_Lean_Syntax_getArg(x_21, x_11); +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; +x_28 = l_Lean_Syntax_getArg(x_21, x_11); lean_dec(x_21); -x_30 = lean_unsigned_to_nat(3u); -x_31 = l_Lean_Syntax_getArg(x_10, x_30); -x_32 = l_Lean_Syntax_getArgs(x_29); -lean_dec(x_29); -x_33 = l_Lean_instInhabitedSyntax; -x_34 = lean_array_get(x_33, x_32, x_11); -x_35 = l_Lean_Syntax_isQuot(x_34); -if (x_35 == 0) +x_29 = lean_unsigned_to_nat(3u); +x_30 = l_Lean_Syntax_getArg(x_10, x_29); +x_31 = l_Lean_Syntax_getArgs(x_28); +lean_dec(x_28); +x_32 = l_Lean_instInhabitedSyntax; +x_33 = lean_array_get(x_32, x_31, x_11); +x_34 = l_Lean_Syntax_isQuot(x_33); +if (x_34 == 0) { -lean_object* x_36; uint8_t x_37; -lean_dec(x_34); -lean_dec(x_32); +lean_object* x_35; uint8_t x_36; +lean_dec(x_33); lean_dec(x_31); +lean_dec(x_30); lean_dec(x_12); lean_dec(x_10); lean_dec(x_6); lean_dec(x_5); lean_dec(x_1); -x_36 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_inferMacroRulesAltKind___spec__2___rarg(x_7); -x_37 = !lean_is_exclusive(x_36); -if (x_37 == 0) +x_35 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_inferMacroRulesAltKind___spec__2___rarg(x_7); +x_36 = !lean_is_exclusive(x_35); +if (x_36 == 0) { -return x_36; +return x_35; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_36, 0); -x_39 = lean_ctor_get(x_36, 1); -lean_inc(x_39); +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_35, 0); +x_38 = lean_ctor_get(x_35, 1); lean_inc(x_38); -lean_dec(x_36); -x_40 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -return x_40; +lean_inc(x_37); +lean_dec(x_35); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; } } else { -lean_object* x_41; lean_object* x_42; -x_41 = lean_box(0); +lean_object* x_40; lean_object* x_41; +x_40 = lean_box(0); lean_inc(x_6); lean_inc(x_5); lean_inc(x_1); -x_42 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2(x_34, x_32, x_31, x_13, x_1, x_10, x_41, x_5, x_6, x_7); -if (lean_obj_tag(x_42) == 0) +x_41 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2(x_33, x_31, x_30, x_13, x_1, x_10, x_40, x_5, x_6, x_7); +if (lean_obj_tag(x_41) == 0) { -lean_object* x_43; lean_object* x_44; size_t x_45; size_t x_46; lean_object* x_47; -x_43 = lean_ctor_get(x_42, 0); +lean_object* x_42; lean_object* x_43; size_t x_44; size_t x_45; lean_object* x_46; +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_41, 1); lean_inc(x_43); -x_44 = lean_ctor_get(x_42, 1); -lean_inc(x_44); -lean_dec(x_42); -x_45 = 1; -x_46 = lean_usize_add(x_3, x_45); -x_47 = lean_array_uset(x_12, x_3, x_43); -x_3 = x_46; -x_4 = x_47; -x_7 = x_44; +lean_dec(x_41); +x_44 = 1; +x_45 = lean_usize_add(x_3, x_44); +x_46 = lean_array_uset(x_12, x_3, x_42); +x_3 = x_45; +x_4 = x_46; +x_7 = x_43; goto _start; } else { -uint8_t x_49; +uint8_t x_48; lean_dec(x_12); lean_dec(x_6); lean_dec(x_5); lean_dec(x_1); -x_49 = !lean_is_exclusive(x_42); -if (x_49 == 0) +x_48 = !lean_is_exclusive(x_41); +if (x_48 == 0) { -return x_42; +return x_41; } else { -lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_50 = lean_ctor_get(x_42, 0); -x_51 = lean_ctor_get(x_42, 1); -lean_inc(x_51); +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_41, 0); +x_50 = lean_ctor_get(x_41, 1); lean_inc(x_50); -lean_dec(x_42); -x_52 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_52, 0, x_50); -lean_ctor_set(x_52, 1, x_51); -return x_52; +lean_inc(x_49); +lean_dec(x_41); +x_51 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +return x_51; } } } @@ -853,7 +1017,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabMacroRulesAux___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__2; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__2; x_2 = l_Lean_Elab_Command_elabMacroRulesAux___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -907,7 +1071,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabMacroRulesAux___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__6; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__6; x_2 = l_Lean_Elab_Command_elabMacroRulesAux___closed__7; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -933,7 +1097,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabMacroRulesAux___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__6; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__6; x_2 = l_Lean_Elab_Command_elabMacroRulesAux___closed__10; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -951,7 +1115,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabMacroRulesAux___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__4; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__4; x_2 = l_Lean_Elab_Command_elabMacroRulesAux___closed__12; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -1095,7 +1259,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabMacroRulesAux___closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__2; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__2; x_2 = l_Lean_Elab_Command_elabMacroRulesAux___closed__24; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -1145,7 +1309,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabMacroRulesAux___closed__33() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__6; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__6; x_2 = l_Lean_Elab_Command_elabMacroRulesAux___closed__32; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -1163,7 +1327,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabMacroRulesAux___closed__35() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__6; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__6; x_2 = l_Lean_Elab_Command_elabMacroRulesAux___closed__34; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -1181,7 +1345,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabMacroRulesAux___closed__37() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__6; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__6; x_2 = l_Lean_Elab_Command_elabMacroRulesAux___closed__36; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -1207,7 +1371,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabMacroRulesAux___closed__40() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__6; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__6; x_2 = l_Lean_Elab_Command_elabMacroRulesAux___closed__39; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -1410,7 +1574,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabMacroRulesAux___closed__60() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__4; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__4; x_2 = l_Array_append___rarg(x_1, x_1); return x_2; } @@ -1420,7 +1584,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabMacroRulesAux___closed__61() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__3; +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__3; x_3 = l_Lean_Elab_Command_elabMacroRulesAux___closed__60; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -1450,7 +1614,7 @@ x_11 = 0; lean_inc(x_7); lean_inc(x_6); lean_inc(x_4); -x_12 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2(x_4, x_10, x_11, x_5, x_6, x_7, x_8); +x_12 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5(x_4, x_10, x_11, x_5, x_6, x_7, x_8); if (lean_obj_tag(x_12) == 0) { 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; lean_object* x_20; lean_object* x_21; uint8_t x_22; @@ -1507,9 +1671,9 @@ x_38 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_38, 0, x_32); lean_ctor_set(x_38, 1, x_37); lean_ctor_set(x_38, 2, x_36); -x_39 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__6; +x_39 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__5; x_40 = lean_array_push(x_39, x_38); -x_41 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__3; +x_41 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__3; x_42 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_42, 0, x_32); lean_ctor_set(x_42, 1, x_41); @@ -1584,9 +1748,9 @@ lean_inc(x_16); x_74 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_74, 0, x_16); lean_ctor_set(x_74, 1, x_73); -x_75 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__4; +x_75 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__4; x_76 = l_Array_append___rarg(x_75, x_13); -x_77 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__1; +x_77 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__1; lean_inc(x_16); x_78 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_78, 0, x_16); @@ -1612,7 +1776,7 @@ x_87 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_87, 0, x_32); lean_ctor_set(x_87, 1, x_41); lean_ctor_set(x_87, 2, x_86); -x_88 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__7; +x_88 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__6; lean_inc(x_16); x_89 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_89, 0, x_16); @@ -1650,12 +1814,12 @@ x_105 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_105, 0, x_32); lean_ctor_set(x_105, 1, x_104); lean_ctor_set(x_105, 2, x_103); -x_106 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__8; +x_106 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__7; x_107 = lean_array_push(x_106, x_78); x_108 = lean_array_push(x_107, x_87); x_109 = lean_array_push(x_108, x_89); x_110 = lean_array_push(x_109, x_105); -x_111 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__8; +x_111 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__8; x_112 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_112, 0, x_32); lean_ctor_set(x_112, 1, x_111); @@ -1763,9 +1927,9 @@ x_163 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_163, 0, x_157); lean_ctor_set(x_163, 1, x_162); lean_ctor_set(x_163, 2, x_161); -x_164 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__6; +x_164 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__5; x_165 = lean_array_push(x_164, x_163); -x_166 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__3; +x_166 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__3; x_167 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_167, 0, x_157); lean_ctor_set(x_167, 1, x_166); @@ -1840,9 +2004,9 @@ lean_inc(x_16); x_199 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_199, 0, x_16); lean_ctor_set(x_199, 1, x_198); -x_200 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__4; +x_200 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__4; x_201 = l_Array_append___rarg(x_200, x_13); -x_202 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__1; +x_202 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__1; lean_inc(x_16); x_203 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_203, 0, x_16); @@ -1868,7 +2032,7 @@ x_212 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_212, 0, x_157); lean_ctor_set(x_212, 1, x_166); lean_ctor_set(x_212, 2, x_211); -x_213 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__7; +x_213 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__6; lean_inc(x_16); x_214 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_214, 0, x_16); @@ -1906,12 +2070,12 @@ x_230 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_230, 0, x_157); lean_ctor_set(x_230, 1, x_229); lean_ctor_set(x_230, 2, x_228); -x_231 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__8; +x_231 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__7; x_232 = lean_array_push(x_231, x_203); x_233 = lean_array_push(x_232, x_212); x_234 = lean_array_push(x_233, x_214); x_235 = lean_array_push(x_234, x_230); -x_236 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__8; +x_236 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__8; x_237 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_237, 0, x_157); lean_ctor_set(x_237, 1, x_236); @@ -2018,7 +2182,26 @@ return x_277; } } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabMacroRulesAux___spec__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_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabMacroRulesAux___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabMacroRulesAux___spec__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabMacroRulesAux___spec__3___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_throwError___at_Lean_Elab_Command_elabMacroRulesAux___spec__3(x_1, x_2, x_3, x_4); +lean_dec(x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabMacroRulesAux___spec__4___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: { size_t x_7; size_t x_8; lean_object* x_9; @@ -2026,7 +2209,7 @@ x_7 = lean_unbox_usize(x_4); lean_dec(x_4); x_8 = lean_unbox_usize(x_5); lean_dec(x_5); -x_9 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabMacroRulesAux___spec__1(x_1, x_2, x_3, x_7, x_8, x_6); +x_9 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabMacroRulesAux___spec__4(x_1, x_2, x_3, x_7, x_8, x_6); lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); @@ -2034,26 +2217,26 @@ lean_dec(x_1); return x_9; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__1(x_1, x_2); +x_3 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__1(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___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, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___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, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; -x_11 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_11 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_7); return x_11; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__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, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___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_8; size_t x_9; lean_object* x_10; @@ -2061,7 +2244,7 @@ x_8 = lean_unbox_usize(x_2); lean_dec(x_2); x_9 = lean_unbox_usize(x_3); lean_dec(x_3); -x_10 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2(x_1, x_8, x_9, x_4, x_5, x_6, x_7); +x_10 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5(x_1, x_8, x_9, x_4, x_5, x_6, x_7); return x_10; } } @@ -2132,15 +2315,15 @@ if (lean_is_exclusive(x_16)) { x_18 = lean_box(0); } x_19 = l_Lean_Syntax_getHeadInfo_x3f(x_1); -x_20 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__4; +x_20 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__4; x_21 = l_Array_append___rarg(x_20, x_7); x_22 = lean_box(2); -x_23 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__3; +x_23 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__3; x_24 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_24, 0, x_22); lean_ctor_set(x_24, 1, x_23); lean_ctor_set(x_24, 2, x_21); -x_25 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__6; +x_25 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__5; x_26 = lean_array_push(x_25, x_24); x_27 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_27, 0, x_22); @@ -2345,7 +2528,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabMacroRules___lambda__2(lean_obj lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; x_10 = lean_unsigned_to_nat(1u); x_11 = l_Lean_Syntax_getArg(x_1, x_10); -x_12 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__5; +x_12 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__5; lean_inc(x_2); x_13 = lean_name_mk_string(x_2, x_12); x_14 = l_Lean_Elab_Command_elabMacroRules___lambda__2___closed__1; @@ -2370,25 +2553,24 @@ return x_17; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; x_18 = lean_unsigned_to_nat(2u); x_19 = l_Lean_Syntax_getArg(x_1, x_18); x_20 = lean_unsigned_to_nat(3u); x_21 = l_Lean_Syntax_getArg(x_1, x_20); -x_22 = l_Lean_nullKind; -x_23 = lean_unsigned_to_nat(0u); +x_22 = lean_unsigned_to_nat(0u); lean_inc(x_21); -x_24 = l_Lean_Syntax_isNodeOf(x_21, x_22, x_23); -if (x_24 == 0) +x_23 = l_Lean_Syntax_matchesNull(x_21, x_22); +if (x_23 == 0) { -lean_object* x_25; uint8_t x_26; +lean_object* x_24; uint8_t x_25; lean_dec(x_4); -x_25 = lean_unsigned_to_nat(5u); +x_24 = lean_unsigned_to_nat(5u); lean_inc(x_21); -x_26 = l_Lean_Syntax_isNodeOf(x_21, x_22, x_25); -if (x_26 == 0) +x_25 = l_Lean_Syntax_matchesNull(x_21, x_24); +if (x_25 == 0) { -lean_object* x_27; +lean_object* x_26; lean_dec(x_21); lean_dec(x_19); lean_dec(x_13); @@ -2398,27 +2580,27 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); -x_27 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__14___rarg(x_9); -return x_27; +x_26 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__14___rarg(x_9); +return x_26; } else { -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 = l_Lean_Syntax_getArg(x_21, x_20); +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_27 = l_Lean_Syntax_getArg(x_21, x_20); lean_dec(x_21); -x_29 = lean_unsigned_to_nat(4u); -x_30 = l_Lean_Syntax_getArg(x_1, x_29); -x_31 = l_Lean_Elab_Command_elabMacroRulesAux___closed__34; +x_28 = lean_unsigned_to_nat(4u); +x_29 = l_Lean_Syntax_getArg(x_1, x_28); +x_30 = l_Lean_Elab_Command_elabMacroRulesAux___closed__34; lean_inc(x_13); -x_32 = lean_name_mk_string(x_13, x_31); -lean_inc(x_30); -x_33 = l_Lean_Syntax_isOfKind(x_30, x_32); -lean_dec(x_32); -if (x_33 == 0) +x_31 = lean_name_mk_string(x_13, x_30); +lean_inc(x_29); +x_32 = l_Lean_Syntax_isOfKind(x_29, x_31); +lean_dec(x_31); +if (x_32 == 0) { -lean_object* x_34; -lean_dec(x_30); -lean_dec(x_28); +lean_object* x_33; +lean_dec(x_29); +lean_dec(x_27); lean_dec(x_19); lean_dec(x_13); lean_dec(x_11); @@ -2427,725 +2609,725 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); -x_34 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__14___rarg(x_9); -return x_34; +x_33 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__14___rarg(x_9); +return x_33; } else { -lean_object* x_35; uint8_t x_36; -x_35 = l_Lean_Syntax_getArg(x_30, x_23); -lean_dec(x_30); -lean_inc(x_35); -x_36 = l_Lean_Syntax_isNodeOf(x_35, x_22, x_10); -if (x_36 == 0) +lean_object* x_34; uint8_t x_35; +x_34 = l_Lean_Syntax_getArg(x_29, x_22); +lean_dec(x_29); +lean_inc(x_34); +x_35 = l_Lean_Syntax_matchesNull(x_34, x_10); +if (x_35 == 0) { -lean_object* x_37; lean_object* x_38; lean_object* x_39; +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_dec(x_13); lean_dec(x_3); lean_dec(x_2); -x_37 = l_Lean_Syntax_getArgs(x_35); -lean_dec(x_35); -x_38 = l_Lean_Syntax_getId(x_28); -lean_dec(x_28); +x_36 = l_Lean_Syntax_getArgs(x_34); +lean_dec(x_34); +x_37 = l_Lean_Syntax_getId(x_27); +lean_dec(x_27); lean_inc(x_8); lean_inc(x_7); -x_39 = l_Lean_Elab_Command_resolveSyntaxKind(x_38, x_7, x_8, x_9); -if (lean_obj_tag(x_39) == 0) +x_38 = l_Lean_Elab_Command_resolveSyntaxKind(x_37, x_7, x_8, x_9); +if (lean_obj_tag(x_38) == 0) { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_39, 0); +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_38, 1); lean_inc(x_40); -x_41 = lean_ctor_get(x_39, 1); -lean_inc(x_41); -lean_dec(x_39); -x_42 = l_Lean_Elab_Command_elabMacroRulesAux(x_6, x_11, x_19, x_40, x_37, x_7, x_8, x_41); -return x_42; +lean_dec(x_38); +x_41 = l_Lean_Elab_Command_elabMacroRulesAux(x_6, x_11, x_19, x_39, x_36, x_7, x_8, x_40); +return x_41; } else { -uint8_t x_43; -lean_dec(x_37); +uint8_t x_42; +lean_dec(x_36); lean_dec(x_19); lean_dec(x_11); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -x_43 = !lean_is_exclusive(x_39); -if (x_43 == 0) +x_42 = !lean_is_exclusive(x_38); +if (x_42 == 0) { -return x_39; +return x_38; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_39, 0); -x_45 = lean_ctor_get(x_39, 1); -lean_inc(x_45); +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_38, 0); +x_44 = lean_ctor_get(x_38, 1); lean_inc(x_44); -lean_dec(x_39); -x_46 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -return x_46; +lean_inc(x_43); +lean_dec(x_38); +x_45 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +return x_45; } } } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; -x_47 = l_Lean_Syntax_getArg(x_35, x_23); -x_48 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__7; +lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; +x_46 = l_Lean_Syntax_getArg(x_34, x_22); +x_47 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__7; lean_inc(x_13); -x_49 = lean_name_mk_string(x_13, x_48); -lean_inc(x_47); -x_50 = l_Lean_Syntax_isOfKind(x_47, x_49); -lean_dec(x_49); -if (x_50 == 0) +x_48 = lean_name_mk_string(x_13, x_47); +lean_inc(x_46); +x_49 = l_Lean_Syntax_isOfKind(x_46, x_48); +lean_dec(x_48); +if (x_49 == 0) { -lean_object* x_51; lean_object* x_52; lean_object* x_53; -lean_dec(x_47); +lean_object* x_50; lean_object* x_51; lean_object* x_52; +lean_dec(x_46); lean_dec(x_13); lean_dec(x_3); lean_dec(x_2); -x_51 = l_Lean_Syntax_getArgs(x_35); -lean_dec(x_35); -x_52 = l_Lean_Syntax_getId(x_28); -lean_dec(x_28); +x_50 = l_Lean_Syntax_getArgs(x_34); +lean_dec(x_34); +x_51 = l_Lean_Syntax_getId(x_27); +lean_dec(x_27); lean_inc(x_8); lean_inc(x_7); -x_53 = l_Lean_Elab_Command_resolveSyntaxKind(x_52, x_7, x_8, x_9); -if (lean_obj_tag(x_53) == 0) +x_52 = l_Lean_Elab_Command_resolveSyntaxKind(x_51, x_7, x_8, x_9); +if (lean_obj_tag(x_52) == 0) { -lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_54 = lean_ctor_get(x_53, 0); +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_52, 0); +lean_inc(x_53); +x_54 = lean_ctor_get(x_52, 1); lean_inc(x_54); -x_55 = lean_ctor_get(x_53, 1); -lean_inc(x_55); -lean_dec(x_53); -x_56 = l_Lean_Elab_Command_elabMacroRulesAux(x_6, x_11, x_19, x_54, x_51, x_7, x_8, x_55); -return x_56; +lean_dec(x_52); +x_55 = l_Lean_Elab_Command_elabMacroRulesAux(x_6, x_11, x_19, x_53, x_50, x_7, x_8, x_54); +return x_55; } else { -uint8_t x_57; -lean_dec(x_51); +uint8_t x_56; +lean_dec(x_50); lean_dec(x_19); lean_dec(x_11); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -x_57 = !lean_is_exclusive(x_53); -if (x_57 == 0) +x_56 = !lean_is_exclusive(x_52); +if (x_56 == 0) { -return x_53; +return x_52; } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_58 = lean_ctor_get(x_53, 0); -x_59 = lean_ctor_get(x_53, 1); -lean_inc(x_59); +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_ctor_get(x_52, 0); +x_58 = lean_ctor_get(x_52, 1); lean_inc(x_58); -lean_dec(x_53); -x_60 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_60, 0, x_58); -lean_ctor_set(x_60, 1, x_59); -return x_60; +lean_inc(x_57); +lean_dec(x_52); +x_59 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_58); +return x_59; } } } else { -lean_object* x_61; uint8_t x_62; -x_61 = l_Lean_Syntax_getArg(x_47, x_10); -lean_inc(x_61); -x_62 = l_Lean_Syntax_isNodeOf(x_61, x_22, x_10); -if (x_62 == 0) +lean_object* x_60; uint8_t x_61; +x_60 = l_Lean_Syntax_getArg(x_46, x_10); +lean_inc(x_60); +x_61 = l_Lean_Syntax_matchesNull(x_60, x_10); +if (x_61 == 0) { -lean_object* x_63; lean_object* x_64; lean_object* x_65; -lean_dec(x_61); -lean_dec(x_47); +lean_object* x_62; lean_object* x_63; lean_object* x_64; +lean_dec(x_60); +lean_dec(x_46); lean_dec(x_13); lean_dec(x_3); lean_dec(x_2); -x_63 = l_Lean_Syntax_getArgs(x_35); -lean_dec(x_35); -x_64 = l_Lean_Syntax_getId(x_28); -lean_dec(x_28); +x_62 = l_Lean_Syntax_getArgs(x_34); +lean_dec(x_34); +x_63 = l_Lean_Syntax_getId(x_27); +lean_dec(x_27); lean_inc(x_8); lean_inc(x_7); -x_65 = l_Lean_Elab_Command_resolveSyntaxKind(x_64, x_7, x_8, x_9); -if (lean_obj_tag(x_65) == 0) +x_64 = l_Lean_Elab_Command_resolveSyntaxKind(x_63, x_7, x_8, x_9); +if (lean_obj_tag(x_64) == 0) { -lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_66 = lean_ctor_get(x_65, 0); +lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = lean_ctor_get(x_64, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_64, 1); lean_inc(x_66); -x_67 = lean_ctor_get(x_65, 1); -lean_inc(x_67); -lean_dec(x_65); -x_68 = l_Lean_Elab_Command_elabMacroRulesAux(x_6, x_11, x_19, x_66, x_63, x_7, x_8, x_67); -return x_68; +lean_dec(x_64); +x_67 = l_Lean_Elab_Command_elabMacroRulesAux(x_6, x_11, x_19, x_65, x_62, x_7, x_8, x_66); +return x_67; } else { -uint8_t x_69; -lean_dec(x_63); +uint8_t x_68; +lean_dec(x_62); lean_dec(x_19); lean_dec(x_11); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -x_69 = !lean_is_exclusive(x_65); -if (x_69 == 0) +x_68 = !lean_is_exclusive(x_64); +if (x_68 == 0) { -return x_65; +return x_64; } else { -lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_70 = lean_ctor_get(x_65, 0); -x_71 = lean_ctor_get(x_65, 1); -lean_inc(x_71); +lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_69 = lean_ctor_get(x_64, 0); +x_70 = lean_ctor_get(x_64, 1); lean_inc(x_70); -lean_dec(x_65); -x_72 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_72, 0, x_70); -lean_ctor_set(x_72, 1, x_71); -return x_72; +lean_inc(x_69); +lean_dec(x_64); +x_71 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_71, 0, x_69); +lean_ctor_set(x_71, 1, x_70); +return x_71; } } } else { -lean_object* x_73; uint8_t x_74; -x_73 = l_Lean_Syntax_getArg(x_61, x_23); -lean_dec(x_61); -lean_inc(x_73); -x_74 = l_Lean_Syntax_isNodeOf(x_73, x_22, x_10); -if (x_74 == 0) +lean_object* x_72; uint8_t x_73; +x_72 = l_Lean_Syntax_getArg(x_60, x_22); +lean_dec(x_60); +lean_inc(x_72); +x_73 = l_Lean_Syntax_matchesNull(x_72, x_10); +if (x_73 == 0) { -lean_object* x_75; lean_object* x_76; lean_object* x_77; -lean_dec(x_73); -lean_dec(x_47); +lean_object* x_74; lean_object* x_75; lean_object* x_76; +lean_dec(x_72); +lean_dec(x_46); lean_dec(x_13); lean_dec(x_3); lean_dec(x_2); -x_75 = l_Lean_Syntax_getArgs(x_35); -lean_dec(x_35); -x_76 = l_Lean_Syntax_getId(x_28); -lean_dec(x_28); +x_74 = l_Lean_Syntax_getArgs(x_34); +lean_dec(x_34); +x_75 = l_Lean_Syntax_getId(x_27); +lean_dec(x_27); lean_inc(x_8); lean_inc(x_7); -x_77 = l_Lean_Elab_Command_resolveSyntaxKind(x_76, x_7, x_8, x_9); -if (lean_obj_tag(x_77) == 0) +x_76 = l_Lean_Elab_Command_resolveSyntaxKind(x_75, x_7, x_8, x_9); +if (lean_obj_tag(x_76) == 0) { -lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_78 = lean_ctor_get(x_77, 0); +lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_77 = lean_ctor_get(x_76, 0); +lean_inc(x_77); +x_78 = lean_ctor_get(x_76, 1); lean_inc(x_78); -x_79 = lean_ctor_get(x_77, 1); -lean_inc(x_79); -lean_dec(x_77); -x_80 = l_Lean_Elab_Command_elabMacroRulesAux(x_6, x_11, x_19, x_78, x_75, x_7, x_8, x_79); -return x_80; +lean_dec(x_76); +x_79 = l_Lean_Elab_Command_elabMacroRulesAux(x_6, x_11, x_19, x_77, x_74, x_7, x_8, x_78); +return x_79; } else { -uint8_t x_81; -lean_dec(x_75); +uint8_t x_80; +lean_dec(x_74); lean_dec(x_19); lean_dec(x_11); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -x_81 = !lean_is_exclusive(x_77); -if (x_81 == 0) +x_80 = !lean_is_exclusive(x_76); +if (x_80 == 0) { -return x_77; +return x_76; } else { -lean_object* x_82; lean_object* x_83; lean_object* x_84; -x_82 = lean_ctor_get(x_77, 0); -x_83 = lean_ctor_get(x_77, 1); -lean_inc(x_83); +lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_81 = lean_ctor_get(x_76, 0); +x_82 = lean_ctor_get(x_76, 1); lean_inc(x_82); -lean_dec(x_77); -x_84 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_84, 0, x_82); -lean_ctor_set(x_84, 1, x_83); -return x_84; +lean_inc(x_81); +lean_dec(x_76); +x_83 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_83, 0, x_81); +lean_ctor_set(x_83, 1, x_82); +return x_83; } } } else { -lean_object* x_85; lean_object* x_86; uint8_t x_87; -x_85 = l_Lean_Syntax_getArg(x_73, x_23); -lean_dec(x_73); -x_86 = l_Lean_Elab_Command_elabMacroRules___lambda__2___closed__3; -lean_inc(x_85); -x_87 = l_Lean_Syntax_isOfKind(x_85, x_86); -if (x_87 == 0) -{ -lean_object* x_88; lean_object* x_89; lean_object* x_90; -lean_dec(x_85); -lean_dec(x_47); +lean_object* x_84; lean_object* x_85; uint8_t x_86; +x_84 = l_Lean_Syntax_getArg(x_72, x_22); +lean_dec(x_72); +x_85 = l_Lean_Elab_Command_elabMacroRules___lambda__2___closed__3; +lean_inc(x_84); +x_86 = l_Lean_Syntax_isOfKind(x_84, x_85); +if (x_86 == 0) +{ +lean_object* x_87; lean_object* x_88; lean_object* x_89; +lean_dec(x_84); +lean_dec(x_46); lean_dec(x_13); lean_dec(x_3); lean_dec(x_2); -x_88 = l_Lean_Syntax_getArgs(x_35); -lean_dec(x_35); -x_89 = l_Lean_Syntax_getId(x_28); -lean_dec(x_28); +x_87 = l_Lean_Syntax_getArgs(x_34); +lean_dec(x_34); +x_88 = l_Lean_Syntax_getId(x_27); +lean_dec(x_27); lean_inc(x_8); lean_inc(x_7); -x_90 = l_Lean_Elab_Command_resolveSyntaxKind(x_89, x_7, x_8, x_9); -if (lean_obj_tag(x_90) == 0) +x_89 = l_Lean_Elab_Command_resolveSyntaxKind(x_88, x_7, x_8, x_9); +if (lean_obj_tag(x_89) == 0) { -lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_91 = lean_ctor_get(x_90, 0); +lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_90 = lean_ctor_get(x_89, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_89, 1); lean_inc(x_91); -x_92 = lean_ctor_get(x_90, 1); -lean_inc(x_92); -lean_dec(x_90); -x_93 = l_Lean_Elab_Command_elabMacroRulesAux(x_6, x_11, x_19, x_91, x_88, x_7, x_8, x_92); -return x_93; +lean_dec(x_89); +x_92 = l_Lean_Elab_Command_elabMacroRulesAux(x_6, x_11, x_19, x_90, x_87, x_7, x_8, x_91); +return x_92; } else { -uint8_t x_94; -lean_dec(x_88); +uint8_t x_93; +lean_dec(x_87); lean_dec(x_19); lean_dec(x_11); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -x_94 = !lean_is_exclusive(x_90); -if (x_94 == 0) +x_93 = !lean_is_exclusive(x_89); +if (x_93 == 0) { -return x_90; +return x_89; } else { -lean_object* x_95; lean_object* x_96; lean_object* x_97; -x_95 = lean_ctor_get(x_90, 0); -x_96 = lean_ctor_get(x_90, 1); -lean_inc(x_96); +lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_94 = lean_ctor_get(x_89, 0); +x_95 = lean_ctor_get(x_89, 1); lean_inc(x_95); -lean_dec(x_90); -x_97 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_97, 0, x_95); -lean_ctor_set(x_97, 1, x_96); -return x_97; +lean_inc(x_94); +lean_dec(x_89); +x_96 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_96, 0, x_94); +lean_ctor_set(x_96, 1, x_95); +return x_96; } } } else { -lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; uint8_t x_106; -lean_dec(x_35); -x_98 = l_Lean_Syntax_getArg(x_47, x_20); -lean_dec(x_47); -x_99 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Command_elabAuxDef___spec__4(x_7, x_8, x_9); -x_100 = lean_ctor_get(x_99, 0); +lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; uint8_t x_105; +lean_dec(x_34); +x_97 = l_Lean_Syntax_getArg(x_46, x_20); +lean_dec(x_46); +x_98 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Command_elabAuxDef___spec__4(x_7, x_8, x_9); +x_99 = lean_ctor_get(x_98, 0); +lean_inc(x_99); +x_100 = lean_ctor_get(x_98, 1); lean_inc(x_100); -x_101 = lean_ctor_get(x_99, 1); -lean_inc(x_101); -lean_dec(x_99); -x_102 = l_Lean_Elab_Command_getCurrMacroScope(x_7, x_8, x_101); +lean_dec(x_98); +x_101 = l_Lean_Elab_Command_getCurrMacroScope(x_7, x_8, x_100); lean_dec(x_7); -x_103 = lean_ctor_get(x_102, 0); +x_102 = lean_ctor_get(x_101, 0); +lean_inc(x_102); +x_103 = lean_ctor_get(x_101, 1); lean_inc(x_103); -x_104 = lean_ctor_get(x_102, 1); -lean_inc(x_104); -lean_dec(x_102); -x_105 = l_Lean_Elab_Command_getMainModule___rarg(x_8, x_104); +lean_dec(x_101); +x_104 = l_Lean_Elab_Command_getMainModule___rarg(x_8, x_103); lean_dec(x_8); -x_106 = !lean_is_exclusive(x_105); -if (x_106 == 0) +x_105 = !lean_is_exclusive(x_104); +if (x_105 == 0) { -lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; -x_107 = lean_ctor_get(x_105, 0); -x_108 = l_Lean_Elab_Command_elabMacroRulesAux___closed__1; +lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; +x_106 = lean_ctor_get(x_104, 0); +x_107 = l_Lean_Elab_Command_elabMacroRulesAux___closed__1; lean_inc(x_3); -x_109 = lean_name_mk_string(x_3, x_108); -x_110 = l_Lean_Elab_Command_elabMacroRulesAux___closed__3; -x_111 = lean_name_mk_string(x_109, x_110); -x_112 = l_Lean_Elab_Command_elabMacroRulesAux___closed__5; -x_113 = lean_name_mk_string(x_111, x_112); -x_114 = l_Lean_Elab_Command_elabMacroRulesAux___closed__7; +x_108 = lean_name_mk_string(x_3, x_107); +x_109 = l_Lean_Elab_Command_elabMacroRulesAux___closed__3; +x_110 = lean_name_mk_string(x_108, x_109); +x_111 = l_Lean_Elab_Command_elabMacroRulesAux___closed__5; +x_112 = lean_name_mk_string(x_110, x_111); +x_113 = l_Lean_Elab_Command_elabMacroRulesAux___closed__7; lean_inc(x_13); -x_115 = lean_name_mk_string(x_13, x_114); -x_116 = l_Lean_Elab_Command_elabMacroRulesAux___closed__9; -lean_inc(x_100); -x_117 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_117, 0, x_100); -lean_ctor_set(x_117, 1, x_116); -x_118 = l_Lean_Elab_Command_elabMacroRulesAux___closed__10; +x_114 = lean_name_mk_string(x_13, x_113); +x_115 = l_Lean_Elab_Command_elabMacroRulesAux___closed__9; +lean_inc(x_99); +x_116 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_116, 0, x_99); +lean_ctor_set(x_116, 1, x_115); +x_117 = l_Lean_Elab_Command_elabMacroRulesAux___closed__10; lean_inc(x_13); -x_119 = lean_name_mk_string(x_13, x_118); -x_120 = l_Lean_Elab_Command_elabMacroRulesAux___closed__12; -x_121 = lean_name_mk_string(x_2, x_120); -x_122 = l_Lean_Elab_Command_elabMacroRulesAux___closed__14; -x_123 = lean_name_mk_string(x_121, x_122); -lean_inc(x_100); -x_124 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_124, 0, x_100); -lean_ctor_set(x_124, 1, x_122); -x_125 = l_Lean_Elab_Command_elabMacroRulesAux___closed__16; -x_126 = lean_array_push(x_125, x_124); -lean_inc(x_28); -x_127 = lean_array_push(x_126, x_28); -x_128 = lean_box(2); -x_129 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_129, 0, x_128); -lean_ctor_set(x_129, 1, x_123); -lean_ctor_set(x_129, 2, x_127); -x_130 = lean_array_push(x_125, x_11); -x_131 = lean_array_push(x_130, x_129); -x_132 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_132, 0, x_128); -lean_ctor_set(x_132, 1, x_119); -lean_ctor_set(x_132, 2, x_131); -x_133 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__6; -x_134 = lean_array_push(x_133, x_132); -x_135 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__3; -x_136 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_136, 0, x_128); -lean_ctor_set(x_136, 1, x_135); -lean_ctor_set(x_136, 2, x_134); -x_137 = l_Lean_Elab_Command_elabMacroRulesAux___closed__17; -lean_inc(x_100); -x_138 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_138, 0, x_100); -lean_ctor_set(x_138, 1, x_137); -x_139 = l_Lean_Elab_Command_elabMacroRulesAux___closed__18; -x_140 = lean_array_push(x_139, x_117); -x_141 = lean_array_push(x_140, x_136); -x_142 = lean_array_push(x_141, x_138); -x_143 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_143, 0, x_128); -lean_ctor_set(x_143, 1, x_115); -lean_ctor_set(x_143, 2, x_142); -x_144 = lean_array_push(x_133, x_143); -x_145 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_145, 0, x_128); -lean_ctor_set(x_145, 1, x_135); -lean_ctor_set(x_145, 2, x_144); -lean_inc(x_100); -x_146 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_146, 0, x_100); -lean_ctor_set(x_146, 1, x_112); -x_147 = l_Lean_Syntax_getId(x_28); -x_148 = l_Lean_mkIdentFrom(x_19, x_147); -x_149 = lean_array_push(x_125, x_148); -x_150 = lean_array_push(x_149, x_28); -x_151 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_151, 0, x_128); -lean_ctor_set(x_151, 1, x_135); -lean_ctor_set(x_151, 2, x_150); -x_152 = l_Lean_Elab_Command_elabMacroRulesAux___closed__23; -lean_inc(x_100); -x_153 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_153, 0, x_100); -lean_ctor_set(x_153, 1, x_152); -x_154 = l_Lean_Elab_Command_elabMacroRulesAux___closed__27; -x_155 = l_Lean_addMacroScope(x_107, x_154, x_103); -x_156 = l_Lean_Elab_Command_elabMacroRulesAux___closed__24; -x_157 = lean_name_mk_string(x_3, x_156); -x_158 = lean_box(0); -x_159 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_159, 0, x_157); -lean_ctor_set(x_159, 1, x_158); -x_160 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_160, 0, x_159); -lean_ctor_set(x_160, 1, x_158); -x_161 = l_Lean_Elab_Command_elabMacroRulesAux___closed__26; -lean_inc(x_100); -x_162 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_162, 0, x_100); -lean_ctor_set(x_162, 1, x_161); -lean_ctor_set(x_162, 2, x_155); -lean_ctor_set(x_162, 3, x_160); -x_163 = l_Lean_Elab_Command_elabMacroRulesAux___closed__31; -lean_inc(x_100); -x_164 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_164, 0, x_100); -lean_ctor_set(x_164, 1, x_163); -x_165 = l_Lean_Elab_Command_elabMacroRulesAux___closed__32; +x_118 = lean_name_mk_string(x_13, x_117); +x_119 = l_Lean_Elab_Command_elabMacroRulesAux___closed__12; +x_120 = lean_name_mk_string(x_2, x_119); +x_121 = l_Lean_Elab_Command_elabMacroRulesAux___closed__14; +x_122 = lean_name_mk_string(x_120, x_121); +lean_inc(x_99); +x_123 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_123, 0, x_99); +lean_ctor_set(x_123, 1, x_121); +x_124 = l_Lean_Elab_Command_elabMacroRulesAux___closed__16; +x_125 = lean_array_push(x_124, x_123); +lean_inc(x_27); +x_126 = lean_array_push(x_125, x_27); +x_127 = lean_box(2); +x_128 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_128, 0, x_127); +lean_ctor_set(x_128, 1, x_122); +lean_ctor_set(x_128, 2, x_126); +x_129 = lean_array_push(x_124, x_11); +x_130 = lean_array_push(x_129, x_128); +x_131 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_131, 0, x_127); +lean_ctor_set(x_131, 1, x_118); +lean_ctor_set(x_131, 2, x_130); +x_132 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__5; +x_133 = lean_array_push(x_132, x_131); +x_134 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__3; +x_135 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_135, 0, x_127); +lean_ctor_set(x_135, 1, x_134); +lean_ctor_set(x_135, 2, x_133); +x_136 = l_Lean_Elab_Command_elabMacroRulesAux___closed__17; +lean_inc(x_99); +x_137 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_137, 0, x_99); +lean_ctor_set(x_137, 1, x_136); +x_138 = l_Lean_Elab_Command_elabMacroRulesAux___closed__18; +x_139 = lean_array_push(x_138, x_116); +x_140 = lean_array_push(x_139, x_135); +x_141 = lean_array_push(x_140, x_137); +x_142 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_142, 0, x_127); +lean_ctor_set(x_142, 1, x_114); +lean_ctor_set(x_142, 2, x_141); +x_143 = lean_array_push(x_132, x_142); +x_144 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_144, 0, x_127); +lean_ctor_set(x_144, 1, x_134); +lean_ctor_set(x_144, 2, x_143); +lean_inc(x_99); +x_145 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_145, 0, x_99); +lean_ctor_set(x_145, 1, x_111); +x_146 = l_Lean_Syntax_getId(x_27); +x_147 = l_Lean_mkIdentFrom(x_19, x_146); +x_148 = lean_array_push(x_124, x_147); +x_149 = lean_array_push(x_148, x_27); +x_150 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_150, 0, x_127); +lean_ctor_set(x_150, 1, x_134); +lean_ctor_set(x_150, 2, x_149); +x_151 = l_Lean_Elab_Command_elabMacroRulesAux___closed__23; +lean_inc(x_99); +x_152 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_152, 0, x_99); +lean_ctor_set(x_152, 1, x_151); +x_153 = l_Lean_Elab_Command_elabMacroRulesAux___closed__27; +x_154 = l_Lean_addMacroScope(x_106, x_153, x_102); +x_155 = l_Lean_Elab_Command_elabMacroRulesAux___closed__24; +x_156 = lean_name_mk_string(x_3, x_155); +x_157 = lean_box(0); +x_158 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_158, 0, x_156); +lean_ctor_set(x_158, 1, x_157); +x_159 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_159, 0, x_158); +lean_ctor_set(x_159, 1, x_157); +x_160 = l_Lean_Elab_Command_elabMacroRulesAux___closed__26; +lean_inc(x_99); +x_161 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_161, 0, x_99); +lean_ctor_set(x_161, 1, x_160); +lean_ctor_set(x_161, 2, x_154); +lean_ctor_set(x_161, 3, x_159); +x_162 = l_Lean_Elab_Command_elabMacroRulesAux___closed__31; +lean_inc(x_99); +x_163 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_163, 0, x_99); +lean_ctor_set(x_163, 1, x_162); +x_164 = l_Lean_Elab_Command_elabMacroRulesAux___closed__32; lean_inc(x_13); -x_166 = lean_name_mk_string(x_13, x_165); -lean_inc(x_100); -x_167 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_167, 0, x_100); -lean_ctor_set(x_167, 1, x_165); -x_168 = l_Lean_Elab_Command_elabMacroRules___lambda__2___closed__4; -x_169 = lean_name_mk_string(x_13, x_168); -x_170 = lean_array_push(x_133, x_85); -x_171 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_171, 0, x_128); -lean_ctor_set(x_171, 1, x_135); -lean_ctor_set(x_171, 2, x_170); -x_172 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__7; -x_173 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_173, 0, x_100); -lean_ctor_set(x_173, 1, x_172); -x_174 = lean_array_push(x_139, x_171); -x_175 = lean_array_push(x_174, x_173); -x_176 = lean_array_push(x_175, x_98); -x_177 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_177, 0, x_128); -lean_ctor_set(x_177, 1, x_169); -lean_ctor_set(x_177, 2, x_176); -x_178 = lean_array_push(x_125, x_167); -x_179 = lean_array_push(x_178, x_177); -x_180 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_180, 0, x_128); -lean_ctor_set(x_180, 1, x_166); -lean_ctor_set(x_180, 2, x_179); +x_165 = lean_name_mk_string(x_13, x_164); +lean_inc(x_99); +x_166 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_166, 0, x_99); +lean_ctor_set(x_166, 1, x_164); +x_167 = l_Lean_Elab_Command_elabMacroRules___lambda__2___closed__4; +x_168 = lean_name_mk_string(x_13, x_167); +x_169 = lean_array_push(x_132, x_84); +x_170 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_170, 0, x_127); +lean_ctor_set(x_170, 1, x_134); +lean_ctor_set(x_170, 2, x_169); +x_171 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__6; +x_172 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_172, 0, x_99); +lean_ctor_set(x_172, 1, x_171); +x_173 = lean_array_push(x_138, x_170); +x_174 = lean_array_push(x_173, x_172); +x_175 = lean_array_push(x_174, x_97); +x_176 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_176, 0, x_127); +lean_ctor_set(x_176, 1, x_168); +lean_ctor_set(x_176, 2, x_175); +x_177 = lean_array_push(x_124, x_166); +x_178 = lean_array_push(x_177, x_176); +x_179 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_179, 0, x_127); +lean_ctor_set(x_179, 1, x_165); +lean_ctor_set(x_179, 2, x_178); if (lean_obj_tag(x_6) == 0) { -lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; -x_181 = l_Lean_Elab_Command_elabMacroRulesAux___closed__62; +lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; +x_180 = l_Lean_Elab_Command_elabMacroRulesAux___closed__62; +x_181 = lean_array_push(x_180, x_144); x_182 = lean_array_push(x_181, x_145); -x_183 = lean_array_push(x_182, x_146); -x_184 = lean_array_push(x_183, x_151); -x_185 = lean_array_push(x_184, x_153); -x_186 = lean_array_push(x_185, x_162); -x_187 = lean_array_push(x_186, x_164); -x_188 = lean_array_push(x_187, x_180); -x_189 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_189, 0, x_128); -lean_ctor_set(x_189, 1, x_113); -lean_ctor_set(x_189, 2, x_188); -lean_ctor_set(x_105, 0, x_189); -return x_105; +x_183 = lean_array_push(x_182, x_150); +x_184 = lean_array_push(x_183, x_152); +x_185 = lean_array_push(x_184, x_161); +x_186 = lean_array_push(x_185, x_163); +x_187 = lean_array_push(x_186, x_179); +x_188 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_188, 0, x_127); +lean_ctor_set(x_188, 1, x_112); +lean_ctor_set(x_188, 2, x_187); +lean_ctor_set(x_104, 0, x_188); +return x_104; } else { -lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; -x_190 = lean_ctor_get(x_6, 0); -lean_inc(x_190); +lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; +x_189 = lean_ctor_get(x_6, 0); +lean_inc(x_189); lean_dec(x_6); -x_191 = lean_array_push(x_133, x_190); -x_192 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__4; -x_193 = l_Array_append___rarg(x_192, x_191); -x_194 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_194, 0, x_128); -lean_ctor_set(x_194, 1, x_135); -lean_ctor_set(x_194, 2, x_193); -x_195 = l_Lean_Elab_Command_elabMacroRulesAux___closed__59; -x_196 = lean_array_push(x_195, x_194); +x_190 = lean_array_push(x_132, x_189); +x_191 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__4; +x_192 = l_Array_append___rarg(x_191, x_190); +x_193 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_193, 0, x_127); +lean_ctor_set(x_193, 1, x_134); +lean_ctor_set(x_193, 2, x_192); +x_194 = l_Lean_Elab_Command_elabMacroRulesAux___closed__59; +x_195 = lean_array_push(x_194, x_193); +x_196 = lean_array_push(x_195, x_144); x_197 = lean_array_push(x_196, x_145); -x_198 = lean_array_push(x_197, x_146); -x_199 = lean_array_push(x_198, x_151); -x_200 = lean_array_push(x_199, x_153); -x_201 = lean_array_push(x_200, x_162); -x_202 = lean_array_push(x_201, x_164); -x_203 = lean_array_push(x_202, x_180); -x_204 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_204, 0, x_128); -lean_ctor_set(x_204, 1, x_113); -lean_ctor_set(x_204, 2, x_203); -lean_ctor_set(x_105, 0, x_204); -return x_105; +x_198 = lean_array_push(x_197, x_150); +x_199 = lean_array_push(x_198, x_152); +x_200 = lean_array_push(x_199, x_161); +x_201 = lean_array_push(x_200, x_163); +x_202 = lean_array_push(x_201, x_179); +x_203 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_203, 0, x_127); +lean_ctor_set(x_203, 1, x_112); +lean_ctor_set(x_203, 2, x_202); +lean_ctor_set(x_104, 0, x_203); +return x_104; } } else { -lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; -x_205 = lean_ctor_get(x_105, 0); -x_206 = lean_ctor_get(x_105, 1); -lean_inc(x_206); +lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; +x_204 = lean_ctor_get(x_104, 0); +x_205 = lean_ctor_get(x_104, 1); lean_inc(x_205); -lean_dec(x_105); -x_207 = l_Lean_Elab_Command_elabMacroRulesAux___closed__1; +lean_inc(x_204); +lean_dec(x_104); +x_206 = l_Lean_Elab_Command_elabMacroRulesAux___closed__1; lean_inc(x_3); -x_208 = lean_name_mk_string(x_3, x_207); -x_209 = l_Lean_Elab_Command_elabMacroRulesAux___closed__3; -x_210 = lean_name_mk_string(x_208, x_209); -x_211 = l_Lean_Elab_Command_elabMacroRulesAux___closed__5; -x_212 = lean_name_mk_string(x_210, x_211); -x_213 = l_Lean_Elab_Command_elabMacroRulesAux___closed__7; +x_207 = lean_name_mk_string(x_3, x_206); +x_208 = l_Lean_Elab_Command_elabMacroRulesAux___closed__3; +x_209 = lean_name_mk_string(x_207, x_208); +x_210 = l_Lean_Elab_Command_elabMacroRulesAux___closed__5; +x_211 = lean_name_mk_string(x_209, x_210); +x_212 = l_Lean_Elab_Command_elabMacroRulesAux___closed__7; lean_inc(x_13); -x_214 = lean_name_mk_string(x_13, x_213); -x_215 = l_Lean_Elab_Command_elabMacroRulesAux___closed__9; -lean_inc(x_100); -x_216 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_216, 0, x_100); -lean_ctor_set(x_216, 1, x_215); -x_217 = l_Lean_Elab_Command_elabMacroRulesAux___closed__10; +x_213 = lean_name_mk_string(x_13, x_212); +x_214 = l_Lean_Elab_Command_elabMacroRulesAux___closed__9; +lean_inc(x_99); +x_215 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_215, 0, x_99); +lean_ctor_set(x_215, 1, x_214); +x_216 = l_Lean_Elab_Command_elabMacroRulesAux___closed__10; lean_inc(x_13); -x_218 = lean_name_mk_string(x_13, x_217); -x_219 = l_Lean_Elab_Command_elabMacroRulesAux___closed__12; -x_220 = lean_name_mk_string(x_2, x_219); -x_221 = l_Lean_Elab_Command_elabMacroRulesAux___closed__14; -x_222 = lean_name_mk_string(x_220, x_221); -lean_inc(x_100); -x_223 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_223, 0, x_100); -lean_ctor_set(x_223, 1, x_221); -x_224 = l_Lean_Elab_Command_elabMacroRulesAux___closed__16; -x_225 = lean_array_push(x_224, x_223); -lean_inc(x_28); -x_226 = lean_array_push(x_225, x_28); -x_227 = lean_box(2); -x_228 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_228, 0, x_227); -lean_ctor_set(x_228, 1, x_222); -lean_ctor_set(x_228, 2, x_226); -x_229 = lean_array_push(x_224, x_11); -x_230 = lean_array_push(x_229, x_228); -x_231 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_231, 0, x_227); -lean_ctor_set(x_231, 1, x_218); -lean_ctor_set(x_231, 2, x_230); -x_232 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__6; -x_233 = lean_array_push(x_232, x_231); -x_234 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__3; -x_235 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_235, 0, x_227); -lean_ctor_set(x_235, 1, x_234); -lean_ctor_set(x_235, 2, x_233); -x_236 = l_Lean_Elab_Command_elabMacroRulesAux___closed__17; -lean_inc(x_100); -x_237 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_237, 0, x_100); -lean_ctor_set(x_237, 1, x_236); -x_238 = l_Lean_Elab_Command_elabMacroRulesAux___closed__18; -x_239 = lean_array_push(x_238, x_216); -x_240 = lean_array_push(x_239, x_235); -x_241 = lean_array_push(x_240, x_237); -x_242 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_242, 0, x_227); -lean_ctor_set(x_242, 1, x_214); -lean_ctor_set(x_242, 2, x_241); -x_243 = lean_array_push(x_232, x_242); -x_244 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_244, 0, x_227); -lean_ctor_set(x_244, 1, x_234); -lean_ctor_set(x_244, 2, x_243); -lean_inc(x_100); -x_245 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_245, 0, x_100); -lean_ctor_set(x_245, 1, x_211); -x_246 = l_Lean_Syntax_getId(x_28); -x_247 = l_Lean_mkIdentFrom(x_19, x_246); -x_248 = lean_array_push(x_224, x_247); -x_249 = lean_array_push(x_248, x_28); -x_250 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_250, 0, x_227); -lean_ctor_set(x_250, 1, x_234); -lean_ctor_set(x_250, 2, x_249); -x_251 = l_Lean_Elab_Command_elabMacroRulesAux___closed__23; -lean_inc(x_100); -x_252 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_252, 0, x_100); -lean_ctor_set(x_252, 1, x_251); -x_253 = l_Lean_Elab_Command_elabMacroRulesAux___closed__27; -x_254 = l_Lean_addMacroScope(x_205, x_253, x_103); -x_255 = l_Lean_Elab_Command_elabMacroRulesAux___closed__24; -x_256 = lean_name_mk_string(x_3, x_255); -x_257 = lean_box(0); -x_258 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_258, 0, x_256); -lean_ctor_set(x_258, 1, x_257); -x_259 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_259, 0, x_258); -lean_ctor_set(x_259, 1, x_257); -x_260 = l_Lean_Elab_Command_elabMacroRulesAux___closed__26; -lean_inc(x_100); -x_261 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_261, 0, x_100); -lean_ctor_set(x_261, 1, x_260); -lean_ctor_set(x_261, 2, x_254); -lean_ctor_set(x_261, 3, x_259); -x_262 = l_Lean_Elab_Command_elabMacroRulesAux___closed__31; -lean_inc(x_100); -x_263 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_263, 0, x_100); -lean_ctor_set(x_263, 1, x_262); -x_264 = l_Lean_Elab_Command_elabMacroRulesAux___closed__32; +x_217 = lean_name_mk_string(x_13, x_216); +x_218 = l_Lean_Elab_Command_elabMacroRulesAux___closed__12; +x_219 = lean_name_mk_string(x_2, x_218); +x_220 = l_Lean_Elab_Command_elabMacroRulesAux___closed__14; +x_221 = lean_name_mk_string(x_219, x_220); +lean_inc(x_99); +x_222 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_222, 0, x_99); +lean_ctor_set(x_222, 1, x_220); +x_223 = l_Lean_Elab_Command_elabMacroRulesAux___closed__16; +x_224 = lean_array_push(x_223, x_222); +lean_inc(x_27); +x_225 = lean_array_push(x_224, x_27); +x_226 = lean_box(2); +x_227 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_227, 0, x_226); +lean_ctor_set(x_227, 1, x_221); +lean_ctor_set(x_227, 2, x_225); +x_228 = lean_array_push(x_223, x_11); +x_229 = lean_array_push(x_228, x_227); +x_230 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_230, 0, x_226); +lean_ctor_set(x_230, 1, x_217); +lean_ctor_set(x_230, 2, x_229); +x_231 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__5; +x_232 = lean_array_push(x_231, x_230); +x_233 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__3; +x_234 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_234, 0, x_226); +lean_ctor_set(x_234, 1, x_233); +lean_ctor_set(x_234, 2, x_232); +x_235 = l_Lean_Elab_Command_elabMacroRulesAux___closed__17; +lean_inc(x_99); +x_236 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_236, 0, x_99); +lean_ctor_set(x_236, 1, x_235); +x_237 = l_Lean_Elab_Command_elabMacroRulesAux___closed__18; +x_238 = lean_array_push(x_237, x_215); +x_239 = lean_array_push(x_238, x_234); +x_240 = lean_array_push(x_239, x_236); +x_241 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_241, 0, x_226); +lean_ctor_set(x_241, 1, x_213); +lean_ctor_set(x_241, 2, x_240); +x_242 = lean_array_push(x_231, x_241); +x_243 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_243, 0, x_226); +lean_ctor_set(x_243, 1, x_233); +lean_ctor_set(x_243, 2, x_242); +lean_inc(x_99); +x_244 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_244, 0, x_99); +lean_ctor_set(x_244, 1, x_210); +x_245 = l_Lean_Syntax_getId(x_27); +x_246 = l_Lean_mkIdentFrom(x_19, x_245); +x_247 = lean_array_push(x_223, x_246); +x_248 = lean_array_push(x_247, x_27); +x_249 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_249, 0, x_226); +lean_ctor_set(x_249, 1, x_233); +lean_ctor_set(x_249, 2, x_248); +x_250 = l_Lean_Elab_Command_elabMacroRulesAux___closed__23; +lean_inc(x_99); +x_251 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_251, 0, x_99); +lean_ctor_set(x_251, 1, x_250); +x_252 = l_Lean_Elab_Command_elabMacroRulesAux___closed__27; +x_253 = l_Lean_addMacroScope(x_204, x_252, x_102); +x_254 = l_Lean_Elab_Command_elabMacroRulesAux___closed__24; +x_255 = lean_name_mk_string(x_3, x_254); +x_256 = lean_box(0); +x_257 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_257, 0, x_255); +lean_ctor_set(x_257, 1, x_256); +x_258 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_258, 0, x_257); +lean_ctor_set(x_258, 1, x_256); +x_259 = l_Lean_Elab_Command_elabMacroRulesAux___closed__26; +lean_inc(x_99); +x_260 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_260, 0, x_99); +lean_ctor_set(x_260, 1, x_259); +lean_ctor_set(x_260, 2, x_253); +lean_ctor_set(x_260, 3, x_258); +x_261 = l_Lean_Elab_Command_elabMacroRulesAux___closed__31; +lean_inc(x_99); +x_262 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_262, 0, x_99); +lean_ctor_set(x_262, 1, x_261); +x_263 = l_Lean_Elab_Command_elabMacroRulesAux___closed__32; lean_inc(x_13); -x_265 = lean_name_mk_string(x_13, x_264); -lean_inc(x_100); -x_266 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_266, 0, x_100); -lean_ctor_set(x_266, 1, x_264); -x_267 = l_Lean_Elab_Command_elabMacroRules___lambda__2___closed__4; -x_268 = lean_name_mk_string(x_13, x_267); -x_269 = lean_array_push(x_232, x_85); -x_270 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_270, 0, x_227); -lean_ctor_set(x_270, 1, x_234); -lean_ctor_set(x_270, 2, x_269); -x_271 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__7; -x_272 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_272, 0, x_100); -lean_ctor_set(x_272, 1, x_271); -x_273 = lean_array_push(x_238, x_270); -x_274 = lean_array_push(x_273, x_272); -x_275 = lean_array_push(x_274, x_98); -x_276 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_276, 0, x_227); -lean_ctor_set(x_276, 1, x_268); -lean_ctor_set(x_276, 2, x_275); -x_277 = lean_array_push(x_224, x_266); -x_278 = lean_array_push(x_277, x_276); -x_279 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_279, 0, x_227); -lean_ctor_set(x_279, 1, x_265); -lean_ctor_set(x_279, 2, x_278); +x_264 = lean_name_mk_string(x_13, x_263); +lean_inc(x_99); +x_265 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_265, 0, x_99); +lean_ctor_set(x_265, 1, x_263); +x_266 = l_Lean_Elab_Command_elabMacroRules___lambda__2___closed__4; +x_267 = lean_name_mk_string(x_13, x_266); +x_268 = lean_array_push(x_231, x_84); +x_269 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_269, 0, x_226); +lean_ctor_set(x_269, 1, x_233); +lean_ctor_set(x_269, 2, x_268); +x_270 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__6; +x_271 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_271, 0, x_99); +lean_ctor_set(x_271, 1, x_270); +x_272 = lean_array_push(x_237, x_269); +x_273 = lean_array_push(x_272, x_271); +x_274 = lean_array_push(x_273, x_97); +x_275 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_275, 0, x_226); +lean_ctor_set(x_275, 1, x_267); +lean_ctor_set(x_275, 2, x_274); +x_276 = lean_array_push(x_223, x_265); +x_277 = lean_array_push(x_276, x_275); +x_278 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_278, 0, x_226); +lean_ctor_set(x_278, 1, x_264); +lean_ctor_set(x_278, 2, x_277); if (lean_obj_tag(x_6) == 0) { -lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; -x_280 = l_Lean_Elab_Command_elabMacroRulesAux___closed__62; +lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; +x_279 = l_Lean_Elab_Command_elabMacroRulesAux___closed__62; +x_280 = lean_array_push(x_279, x_243); x_281 = lean_array_push(x_280, x_244); -x_282 = lean_array_push(x_281, x_245); -x_283 = lean_array_push(x_282, x_250); -x_284 = lean_array_push(x_283, x_252); -x_285 = lean_array_push(x_284, x_261); -x_286 = lean_array_push(x_285, x_263); -x_287 = lean_array_push(x_286, x_279); -x_288 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_288, 0, x_227); -lean_ctor_set(x_288, 1, x_212); -lean_ctor_set(x_288, 2, x_287); -x_289 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_289, 0, x_288); -lean_ctor_set(x_289, 1, x_206); -return x_289; +x_282 = lean_array_push(x_281, x_249); +x_283 = lean_array_push(x_282, x_251); +x_284 = lean_array_push(x_283, x_260); +x_285 = lean_array_push(x_284, x_262); +x_286 = lean_array_push(x_285, x_278); +x_287 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_287, 0, x_226); +lean_ctor_set(x_287, 1, x_211); +lean_ctor_set(x_287, 2, x_286); +x_288 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_288, 0, x_287); +lean_ctor_set(x_288, 1, x_205); +return x_288; } else { -lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; -x_290 = lean_ctor_get(x_6, 0); -lean_inc(x_290); +lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; +x_289 = lean_ctor_get(x_6, 0); +lean_inc(x_289); lean_dec(x_6); -x_291 = lean_array_push(x_232, x_290); -x_292 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__4; -x_293 = l_Array_append___rarg(x_292, x_291); -x_294 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_294, 0, x_227); -lean_ctor_set(x_294, 1, x_234); -lean_ctor_set(x_294, 2, x_293); -x_295 = l_Lean_Elab_Command_elabMacroRulesAux___closed__59; -x_296 = lean_array_push(x_295, x_294); +x_290 = lean_array_push(x_231, x_289); +x_291 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__4; +x_292 = l_Array_append___rarg(x_291, x_290); +x_293 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_293, 0, x_226); +lean_ctor_set(x_293, 1, x_233); +lean_ctor_set(x_293, 2, x_292); +x_294 = l_Lean_Elab_Command_elabMacroRulesAux___closed__59; +x_295 = lean_array_push(x_294, x_293); +x_296 = lean_array_push(x_295, x_243); x_297 = lean_array_push(x_296, x_244); -x_298 = lean_array_push(x_297, x_245); -x_299 = lean_array_push(x_298, x_250); -x_300 = lean_array_push(x_299, x_252); -x_301 = lean_array_push(x_300, x_261); -x_302 = lean_array_push(x_301, x_263); -x_303 = lean_array_push(x_302, x_279); -x_304 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_304, 0, x_227); -lean_ctor_set(x_304, 1, x_212); -lean_ctor_set(x_304, 2, x_303); -x_305 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_305, 0, x_304); -lean_ctor_set(x_305, 1, x_206); -return x_305; +x_298 = lean_array_push(x_297, x_249); +x_299 = lean_array_push(x_298, x_251); +x_300 = lean_array_push(x_299, x_260); +x_301 = lean_array_push(x_300, x_262); +x_302 = lean_array_push(x_301, x_278); +x_303 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_303, 0, x_226); +lean_ctor_set(x_303, 1, x_211); +lean_ctor_set(x_303, 2, x_302); +x_304 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_304, 0, x_303); +lean_ctor_set(x_304, 1, x_205); +return x_304; } } } @@ -3158,47 +3340,89 @@ return x_305; } else { -lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; uint8_t x_310; +lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; uint8_t x_309; lean_dec(x_21); lean_dec(x_3); lean_dec(x_2); -x_306 = lean_unsigned_to_nat(4u); -x_307 = l_Lean_Syntax_getArg(x_1, x_306); -x_308 = l_Lean_Elab_Command_elabMacroRulesAux___closed__34; -x_309 = lean_name_mk_string(x_13, x_308); -lean_inc(x_307); -x_310 = l_Lean_Syntax_isOfKind(x_307, x_309); -if (x_310 == 0) -{ -lean_object* x_311; -lean_dec(x_309); -lean_dec(x_307); +x_305 = lean_unsigned_to_nat(4u); +x_306 = l_Lean_Syntax_getArg(x_1, x_305); +x_307 = l_Lean_Elab_Command_elabMacroRulesAux___closed__34; +x_308 = lean_name_mk_string(x_13, x_307); +lean_inc(x_306); +x_309 = l_Lean_Syntax_isOfKind(x_306, x_308); +if (x_309 == 0) +{ +lean_object* x_310; +lean_dec(x_308); +lean_dec(x_306); lean_dec(x_19); lean_dec(x_11); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_4); -x_311 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__14___rarg(x_9); -return x_311; +x_310 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__14___rarg(x_9); +return x_310; } else { -lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; -x_312 = l_Lean_Syntax_getArg(x_307, x_23); -lean_dec(x_307); -x_313 = l_Lean_Syntax_getArgs(x_312); -lean_dec(x_312); -x_314 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabMacroRules___lambda__1___boxed), 10, 5); -lean_closure_set(x_314, 0, x_19); -lean_closure_set(x_314, 1, x_309); -lean_closure_set(x_314, 2, x_11); -lean_closure_set(x_314, 3, x_4); -lean_closure_set(x_314, 4, x_6); -x_315 = l_Lean_Elab_Command_elabMacroRules___lambda__1___closed__2; -x_316 = l_Lean_Elab_Command_expandNoKindMacroRulesAux(x_313, x_315, x_314, x_7, x_8, x_9); -lean_dec(x_313); -return x_316; +lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; +x_311 = l_Lean_Syntax_getArg(x_306, x_22); +lean_dec(x_306); +x_312 = l_Lean_Syntax_getArgs(x_311); +lean_dec(x_311); +x_313 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabMacroRules___lambda__1___boxed), 10, 5); +lean_closure_set(x_313, 0, x_19); +lean_closure_set(x_313, 1, x_308); +lean_closure_set(x_313, 2, x_11); +lean_closure_set(x_313, 3, x_4); +lean_closure_set(x_313, 4, x_6); +x_314 = l_Lean_Elab_Command_elabMacroRules___lambda__1___closed__2; +x_315 = l_Lean_Elab_Command_expandNoKindMacroRulesAux(x_312, x_314, x_313, x_7, x_8, x_9); +if (lean_obj_tag(x_315) == 0) +{ +uint8_t x_316; +x_316 = !lean_is_exclusive(x_315); +if (x_316 == 0) +{ +return x_315; +} +else +{ +lean_object* x_317; lean_object* x_318; lean_object* x_319; +x_317 = lean_ctor_get(x_315, 0); +x_318 = lean_ctor_get(x_315, 1); +lean_inc(x_318); +lean_inc(x_317); +lean_dec(x_315); +x_319 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_319, 0, x_317); +lean_ctor_set(x_319, 1, x_318); +return x_319; +} +} +else +{ +uint8_t x_320; +x_320 = !lean_is_exclusive(x_315); +if (x_320 == 0) +{ +return x_315; +} +else +{ +lean_object* x_321; lean_object* x_322; lean_object* x_323; +x_321 = lean_ctor_get(x_315, 0); +x_322 = lean_ctor_get(x_315, 1); +lean_inc(x_322); +lean_inc(x_321); +lean_dec(x_315); +x_323 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_323, 0, x_321); +lean_ctor_set(x_323, 1, x_322); +return x_323; +} +} } } } @@ -3208,7 +3432,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabMacroRules___lambda__3___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__4; +x_1 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__4; x_2 = l_Lean_Elab_Command_elabMacroRulesAux___closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -3266,64 +3490,63 @@ x_9 = l_Lean_Syntax_getArg(x_1, x_8); x_10 = l_Lean_Syntax_isNone(x_9); if (x_10 == 0) { -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = l_Lean_nullKind; -x_12 = lean_unsigned_to_nat(1u); +lean_object* x_11; uint8_t x_12; +x_11 = lean_unsigned_to_nat(1u); lean_inc(x_9); -x_13 = l_Lean_Syntax_isNodeOf(x_9, x_11, x_12); -if (x_13 == 0) +x_12 = l_Lean_Syntax_matchesNull(x_9, x_11); +if (x_12 == 0) { -lean_object* x_14; +lean_object* x_13; lean_dec(x_9); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_14 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__14___rarg(x_4); -return x_14; +x_13 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__14___rarg(x_4); +return x_13; } else { -lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_15 = l_Lean_Syntax_getArg(x_9, x_8); +lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_14 = l_Lean_Syntax_getArg(x_9, x_8); lean_dec(x_9); -x_16 = l_Lean_Elab_Command_elabMacroRules___lambda__3___closed__4; -lean_inc(x_15); -x_17 = l_Lean_Syntax_isOfKind(x_15, x_16); -if (x_17 == 0) +x_15 = l_Lean_Elab_Command_elabMacroRules___lambda__3___closed__4; +lean_inc(x_14); +x_16 = l_Lean_Syntax_isOfKind(x_14, x_15); +if (x_16 == 0) { -lean_object* x_18; -lean_dec(x_15); +lean_object* x_17; +lean_dec(x_14); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_18 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__14___rarg(x_4); -return x_18; +x_17 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__14___rarg(x_4); +return x_17; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_19 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_19, 0, x_15); -x_20 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__4; -x_21 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__2; -x_22 = lean_box(0); -x_23 = l_Lean_Elab_Command_elabMacroRules___lambda__2(x_1, x_20, x_21, x_5, x_22, x_19, x_2, x_3, x_4); +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_18 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_18, 0, x_14); +x_19 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__4; +x_20 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__2; +x_21 = lean_box(0); +x_22 = l_Lean_Elab_Command_elabMacroRules___lambda__2(x_1, x_19, x_20, x_5, x_21, x_18, x_2, x_3, x_4); lean_dec(x_1); -return x_23; +return x_22; } } } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_dec(x_9); -x_24 = lean_box(0); -x_25 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__4; -x_26 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__2; -x_27 = lean_box(0); -x_28 = l_Lean_Elab_Command_elabMacroRules___lambda__2(x_1, x_25, x_26, x_5, x_27, x_24, x_2, x_3, x_4); +x_23 = lean_box(0); +x_24 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__4; +x_25 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__2; +x_26 = lean_box(0); +x_27 = l_Lean_Elab_Command_elabMacroRules___lambda__2(x_1, x_24, x_25, x_5, x_26, x_23, x_2, x_3, x_4); lean_dec(x_1); -return x_28; +return x_27; } } } @@ -3531,52 +3754,58 @@ lean_dec_ref(res); res = initialize_Lean_Elab_AuxDef(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__1 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__1(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__1); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__2 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__2(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__2); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__3 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__3(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__3); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__4 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__4(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__4); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__5 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__5(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__5); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__6 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__6(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__6); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__7 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__7(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__7); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__8 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__8(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__8); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__9 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__9(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__9); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__10 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__10(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__10); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__11 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__11(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__11); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__12 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__12(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__12); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__13 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__13(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__13); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__14 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__14(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__14); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__15 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__15(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___lambda__2___closed__15); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__1 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__1(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__1); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__2 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__2(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__2); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__3 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__3(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__3); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__4 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__4(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__4); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__5 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__5(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__5); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__6 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__6(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__6); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__7 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__7(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__7); -l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__8 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__8(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__2___closed__8); +l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabMacroRulesAux___spec__1___rarg___closed__1 = _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabMacroRulesAux___spec__1___rarg___closed__1(); +lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabMacroRulesAux___spec__1___rarg___closed__1); +l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabMacroRulesAux___spec__1___rarg___closed__2 = _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabMacroRulesAux___spec__1___rarg___closed__2(); +lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabMacroRulesAux___spec__1___rarg___closed__2); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__1 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__1(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__1); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__2 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__2(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__2); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__3 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__3(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__3); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__4 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__4(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__4); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__5 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__5(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__5); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__6 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__6(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__6); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__7 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__7(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__7); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__8 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__8(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__8); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__9 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__9(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__9); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__10 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__10(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__10); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__11 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__11(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__11); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__12 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__12(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__12); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__13 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__13(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__13); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__14 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__14(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__14); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__15 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__15(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__15); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__16 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__16(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__16); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__1 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__1(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__1); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__2 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__2(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__2); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__3 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__3(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__3); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__4 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__4(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__4); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__5 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__5(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__5); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__6 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__6(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__6); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__7 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__7(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__7); +l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__8 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__8(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___closed__8); l_Lean_Elab_Command_elabMacroRulesAux___closed__1 = _init_l_Lean_Elab_Command_elabMacroRulesAux___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_elabMacroRulesAux___closed__1); l_Lean_Elab_Command_elabMacroRulesAux___closed__2 = _init_l_Lean_Elab_Command_elabMacroRulesAux___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Match.c b/stage0/stdlib/Lean/Elab/Match.c index 3dd289b47e9..591d9bb8670 100644 --- a/stage0/stdlib/Lean/Elab/Match.c +++ b/stage0/stdlib/Lean/Elab/Match.c @@ -47,6 +47,7 @@ lean_object* lean_erase_macro_scopes(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_withPatternVars_loop___rarg(lean_object*, lean_object*, 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_Meta_transform___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___closed__2; +static lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4___rarg___closed__3; lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_FunBinders_elabFunBinderViews___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_ToDepElimPattern_main___spec__4(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -64,16 +65,17 @@ static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_findDiscrRefin LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_withEqs(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_elabMatch_elabMatchDefault___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_ToDepElimPattern_normalize___spec__3(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_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchCore___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltView___lambda__8(lean_object*, lean_object*, lean_object*, 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_Array_anyMUnsafe_any___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__4(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalDecl_userName(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMatch(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapM___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_toPattern___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashMapImp_expand___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__4(lean_object*, lean_object*); -extern lean_object* l_Lean_nullKind; LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltView___lambda__3___boxed(lean_object*, lean_object*, lean_object*, 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___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_ToDepElimPattern_main___spec__4___closed__2; +LEAN_EXPORT lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_loop___lambda__4___boxed(lean_object**); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabMatch_declRange(lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___closed__3; @@ -83,7 +85,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_withPatte uint8_t lean_usize_dec_eq(size_t, size_t); lean_object* lean_array_uget(lean_object*, size_t); lean_object* l_Lean_Meta_getFVarsToGeneralize(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_updateMatchType___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__11(size_t, size_t, lean_object*); uint64_t lean_uint64_of_nat(lean_object*); @@ -100,6 +102,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabNoMatch(lean_object*, lean_object* LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_RBTree_toList___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_collectDeps___spec__2(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_topSort___spec__1(lean_object*, size_t, size_t, 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_withInfoContext_x27___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__3___boxed(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___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_whnfPreservingPatternRef(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Elab_Term_precheckMatch___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_normalize_processVar___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -120,7 +123,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitForall___at___private_ LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_CollectFVars_main(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_loop___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13146____closed__3; extern lean_object* l_Lean_maxRecDepthErrorMessage; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_collectDeps___closed__2; @@ -130,7 +132,6 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_Match_0 static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_isMatchUnit_x3f___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_withDepElimPatterns(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__12___boxed(lean_object*, lean_object*); lean_object* l_Array_extract___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_AssocList_replace___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__7(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_toPattern___closed__3; @@ -173,7 +174,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_waitExpec LEAN_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__9___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabInaccessible(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkIdentFrom(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__2(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_reportMatcherResultErrors___lambda__1___boxed(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_Term_ToDepElimPattern_main___rarg(lean_object*, 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_Array_mapMUnsafe_map___at_Lean_Elab_Term_ToDepElimPattern_normalize___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -182,7 +183,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit___at___private_Lean_Elab_Ma uint8_t l_Lean_Expr_isAppOf(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabPatterns___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_topSort_visit___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__2___closed__1; static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__1; lean_object* l_Lean_Elab_Term_getMVarDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_isMatchUnit_x3f___closed__5; @@ -229,6 +229,7 @@ static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandNonAtomi static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___spec__1___closed__12; lean_object* l_Std_mkHashSetImp___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_elabMatch_elabMatchDefault___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4___rarg___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_throwMVarError___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -237,6 +238,7 @@ lean_object* l_Lean_Expr_ReplaceImpl_replaceUnsafe(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandNonAtomicDiscrs_x3f_loop___lambda__2___closed__2; lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_assertAfter___spec__10(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Term_elabMatch_elabMatchDefault___closed__2; +static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___closed__1; LEAN_EXPORT lean_object* l_Std_AssocList_foldlM___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__6(lean_object*, lean_object*); lean_object* lean_string_utf8_byte_size(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_withEqs_go___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -248,6 +250,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit___at___private_Lean_Elab_Ma static lean_object* l_Lean_Elab_Term_ToDepElimPattern_normalize_addVar___closed__2; lean_object* l_Lean_Meta_Match_mkMatcher(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_replaceFVars(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__3(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___rarg___lambda__1___boxed(lean_object*, 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_liftMacroM___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__1___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -264,11 +267,11 @@ static lean_object* l_Lean_Meta_transform_visit___at___private_Lean_Elab_Match_0 extern lean_object* l_Lean_levelZero; LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandNonAtomicDiscrs_x3f_loop(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__12___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwAbortTerm___at_Lean_Elab_Term_isAtomicDiscr_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__8___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_normalize(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__12(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_findDiscrRefinementPath_goIndex___spec__1(lean_object*, lean_object*, 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_Term_precheckMatch___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -278,12 +281,14 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatch static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__2(lean_object*); lean_object* l_Array_zip___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13231____closed__5; static lean_object* l_Lean_Elab_Term_elabNoMatch___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__1___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMatch_elabMatchDefault___lambda__1___boxed(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___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_updateMatchType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withoutAuxDiscrs___spec__4___boxed(lean_object*, lean_object*, 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_Std_PersistentArray_forIn___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withoutAuxDiscrs___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4___rarg___closed__1; lean_object* l_Lean_mkAppN(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_throwInvalidPattern___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_throwInvalidPattern___rarg___closed__1; @@ -304,15 +309,15 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabInaccessible_declRange___c lean_object* l_Lean_Core_withFreshMacroScope___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___lambda__3(lean_object*, lean_object*, 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___private_Lean_Elab_Match_0__Lean_Elab_Term_getDiscrs(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___lambda__3___boxed(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_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMacrosInPatterns___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_loop___spec__7(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_Quotation_precheck(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_packMatchTypePatterns(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_StateRefT_x27_lift___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Syntax_SepArray_getElems___rarg(lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); -lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___closed__2; static lean_object* l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__2___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalDecl_value(lean_object*); @@ -330,6 +335,7 @@ uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lea static lean_object* l___regBuiltin_Lean_Elab_Term_elabMatch_declRange___closed__1; LEAN_EXPORT lean_object* l_Lean_mkAuxDiscr___at___private_Lean_Elab_Match_0__Lean_Elab_Term_expandNonAtomicDiscrs_x3f_loop___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_loop___spec__4___closed__2; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13231____closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_findDiscrRefinementPath_goIndex(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabAtomicDiscr___closed__2; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___spec__1___closed__15; @@ -343,6 +349,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_transform___at___private_Lean_Elab_Match_0_ LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_findDiscrRefinementPath_goType___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_StateRefT_x27_lift(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__5; +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabImplicitLambda_loop___spec__1___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandNonAtomicDiscrs_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_forallMetaTelescopeReducingAux(lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -368,6 +375,7 @@ lean_object* l_Std_RBNode_insert___at_Lean_Level_collectMVars___spec__1(lean_obj LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_ToDepElimPattern_isExplicitPatternVar___spec__1(lean_object*, lean_object*, size_t, size_t); static lean_object* l___regBuiltin_Lean_Elab_Term_elabMatch_declRange___closed__2; static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___lambda__3___closed__1; +LEAN_EXPORT lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_AssocList_find_x3f___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_applyRefMap___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentArray_anyM___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__1(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_Term_ToDepElimPattern_isExplicitPatternVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -381,6 +389,7 @@ static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lea static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__1(lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___closed__11; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13231____closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__1___rarg___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_collectDeps___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -392,7 +401,6 @@ lean_object* l_Lean_Elab_Term_withMacroExpansion___rarg(lean_object*, lean_objec LEAN_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_collectDeps___spec__5(lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_collectDeps___closed__1; LEAN_EXPORT lean_object* l_Std_mkHashMap___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__1(lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13146____closed__4; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_generalize___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_elabTerm(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_mkFreshBinderName___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___spec__2___rarg(lean_object*, lean_object*, lean_object*); @@ -452,7 +460,7 @@ lean_object* l_Lean_Meta_mkArrayLit(lean_object*, lean_object*, lean_object*, le LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_loop___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_panic___at___private_Lean_Meta_WHNF_0__Lean_Meta_cached_x3f___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandMacrosInPatterns(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_filterMapM___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_filterMapM___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_tryPostponeIfDiscrTypeIsMVar___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabNoMatch___closed__2; @@ -514,6 +522,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Elab_Matc lean_object* l_Lean_Elab_Term_resolveId_x3f(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__8___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_main___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13231____closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_tryPostponeIfDiscrTypeIsMVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltView___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static uint64_t l_Lean_Elab_Term_instInhabitedDiscr___closed__1; @@ -567,7 +576,6 @@ static lean_object* l_Lean_Elab_Term_instInhabitedDiscr___closed__2; lean_object* l_Lean_Elab_Term_elabTermEnsuringType___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_whnf(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_Meta_Match_toPattern___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13146____closed__5; lean_object* l_Lean_mkFVar(lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_loop___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_ToDepElimPattern_normalize___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -613,7 +621,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_findDiscr LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_normalize_processInaccessible(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___closed__2; lean_object* l_Lean_throwError___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_elabMatch___closed__2; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_tryPostponeIfDiscrTypeIsMVar___spec__1___closed__1; extern lean_object* l_Lean_Elab_Term_termElabAttribute; static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltView___lambda__4___closed__1; @@ -623,7 +630,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatch lean_object* l_Lean_mkAtomFrom(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_waitExpectedType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_TSepArray_getElems___rarg(lean_object*); static lean_object* l_Lean_Elab_Term_ToDepElimPattern_main___rarg___lambda__2___closed__3; lean_object* l_Lean_LocalDecl_type(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltView___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -643,14 +650,18 @@ LEAN_EXPORT lean_object* l_Lean_Elab_throwAbortTerm___at_Lean_Elab_Term_isAtomic uint8_t l_Array_contains___at___private_Lean_Meta_Match_Value_0__Lean_Meta_isUIntTypeName___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_withPatternElabConfig___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_filterMapM___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_filterMapM___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Term_reportMatcherResultErrors___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13231____closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__7___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_redLength___rarg(lean_object*); +lean_object* l_Std_PersistentArray_push___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchCore___spec__1(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_mkUserNameFor(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_main_unpack_go(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4___boxed(lean_object*, lean_object*); lean_object* l_Lean_Syntax_getSepArgs(lean_object*); static lean_object* l_Lean_Elab_Term_ToDepElimPattern_main___rarg___lambda__4___closed__1; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___spec__1___closed__1; @@ -740,7 +751,6 @@ static lean_object* l_List_forIn_loop___at_Lean_Elab_Term_reportMatcherResultErr LEAN_EXPORT lean_object* l_Lean_Meta_transform___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__1(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentArray_anyM___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_mkAuxDiscr___at___private_Lean_Elab_Match_0__Lean_Elab_Term_expandNonAtomicDiscrs_x3f_loop___spec__1___rarg___closed__2; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13146____closed__2; lean_object* l_Lean_Elab_Term_elabTerm___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_LocalDecl_hasExprMVar(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabMatch_declRange___closed__5; @@ -763,7 +773,6 @@ LEAN_EXPORT lean_object* l_Lean_mkAuxDiscr___at___private_Lean_Elab_Match_0__Lea LEAN_EXPORT lean_object* l_Std_RBNode_revFold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_collectDeps___spec__3(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabMatch___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__7(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13146____closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_toPattern___spec__2(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Meta_Match_Pattern_hasExprMVar(lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___closed__7; @@ -789,6 +798,7 @@ static lean_object* l_Lean_Elab_Term_isAtomicDiscr_x3f___closed__3; static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_collectDeps___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_elabNoMatch_declRange___closed__6; LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_generalize___boxed(lean_object*, 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_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4(uint8_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_withElaboratedLHS___rarg___lambda__1(lean_object*, lean_object*, 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_throwError___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_throwInvalidPattern___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___spec__1___closed__13; @@ -810,6 +820,7 @@ lean_object* l_Lean_mkForall(lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_precheckMatch(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabNoMatch_docString(lean_object*); LEAN_EXPORT lean_object* l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__12(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitForall___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__5(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_generalize___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -860,7 +871,6 @@ static lean_object* l_Lean_Elab_throwAbortTerm___at_Lean_Elab_Term_isAtomicDiscr uint8_t l_Lean_Expr_binderInfo(lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_toPattern___closed__2; uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_precheckMatch___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_eraseInaccessibleAnnotations___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_findDiscrRefinementPath_goIndex___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabMatch_docString(lean_object*); @@ -925,6 +935,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatch lean_object* l_Std_RBNode_findCore___at_Lean_Meta_mkGeneralizationForbiddenSet_visit___spec__1(lean_object*, lean_object*); uint8_t l_List_isEmpty___rarg(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabInaccessible___closed__4; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___lambda__3(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_withElaboratedLHS___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_HashMapImp_find_x3f___at_Lean_MetavarContext_instantiateExprMVars___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_match_ignoreUnusedAlts; @@ -964,8 +975,8 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabNoMatch_declRange(lea lean_object* l_Lean_Elab_Term_isLocalIdent_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Closure_mkBinding___spec__1(size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_16449_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13146_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_16570_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13231_(lean_object*); lean_object* l_Lean_mkSimpleThunk(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_isAtomicDiscr_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabNoMatch___closed__4; @@ -981,6 +992,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabInaccessible_declRange___c LEAN_EXPORT lean_object* l_Std_PersistentArray_anyMAux___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabMatch___closed__1; static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___lambda__4___closed__1; +lean_object* l_Lean_Elab_Term_mkTermInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabNoMatch_declRange___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_precheckMatch___spec__4(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_withElaboratedLHS___rarg___closed__1; @@ -6434,316 +6446,317 @@ return x_4; } else { -lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +lean_object* x_5; lean_object* x_6; uint8_t x_7; x_5 = lean_unsigned_to_nat(1u); x_6 = l_Lean_Syntax_getArg(x_1, x_5); -x_7 = l_Lean_nullKind; lean_inc(x_6); -x_8 = l_Lean_Syntax_isNodeOf(x_6, x_7, x_5); -if (x_8 == 0) +x_7 = l_Lean_Syntax_matchesNull(x_6, x_5); +if (x_7 == 0) { -lean_object* x_9; +lean_object* x_8; lean_dec(x_6); lean_dec(x_1); -x_9 = lean_box(0); -return x_9; +x_8 = lean_box(0); +return x_8; } else { -lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_10 = lean_unsigned_to_nat(0u); -x_11 = l_Lean_Syntax_getArg(x_6, x_10); +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_9 = lean_unsigned_to_nat(0u); +x_10 = l_Lean_Syntax_getArg(x_6, x_9); lean_dec(x_6); -x_12 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___closed__3; -lean_inc(x_11); -x_13 = l_Lean_Syntax_isOfKind(x_11, x_12); -if (x_13 == 0) +x_11 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___closed__3; +lean_inc(x_10); +x_12 = l_Lean_Syntax_isOfKind(x_10, x_11); +if (x_12 == 0) { -lean_object* x_14; -lean_dec(x_11); +lean_object* x_13; +lean_dec(x_10); lean_dec(x_1); -x_14 = lean_box(0); -return x_14; +x_13 = lean_box(0); +return x_13; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_15 = lean_unsigned_to_nat(3u); -x_16 = l_Lean_Syntax_getArg(x_11, x_15); -lean_dec(x_11); -x_17 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___closed__5; -lean_inc(x_16); -x_18 = l_Lean_Syntax_isOfKind(x_16, x_17); -if (x_18 == 0) +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_14 = lean_unsigned_to_nat(3u); +x_15 = l_Lean_Syntax_getArg(x_10, x_14); +lean_dec(x_10); +x_16 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___closed__5; +lean_inc(x_15); +x_17 = l_Lean_Syntax_isOfKind(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; uint8_t x_20; -x_19 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___closed__7; -x_20 = l_Lean_Syntax_isOfKind(x_16, x_19); -if (x_20 == 0) +lean_object* x_18; uint8_t x_19; +x_18 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___closed__7; +x_19 = l_Lean_Syntax_isOfKind(x_15, x_18); +if (x_19 == 0) { -lean_object* x_21; +lean_object* x_20; lean_dec(x_1); -x_21 = lean_box(0); -return x_21; +x_20 = lean_box(0); +return x_20; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_33; -x_22 = lean_unsigned_to_nat(2u); -x_23 = l_Lean_Syntax_getArg(x_1, x_22); -x_33 = l_Lean_Syntax_isNone(x_23); -if (x_33 == 0) +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_32; +x_21 = lean_unsigned_to_nat(2u); +x_22 = l_Lean_Syntax_getArg(x_1, x_21); +x_32 = l_Lean_Syntax_isNone(x_22); +if (x_32 == 0) { -uint8_t x_34; -lean_inc(x_23); -x_34 = l_Lean_Syntax_isNodeOf(x_23, x_7, x_5); -if (x_34 == 0) +uint8_t x_33; +lean_inc(x_22); +x_33 = l_Lean_Syntax_matchesNull(x_22, x_5); +if (x_33 == 0) { -lean_object* x_35; -lean_dec(x_23); +lean_object* x_34; +lean_dec(x_22); lean_dec(x_1); -x_35 = lean_box(0); -return x_35; +x_34 = lean_box(0); +return x_34; } else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_36 = l_Lean_Syntax_getArg(x_23, x_10); -lean_dec(x_23); -x_37 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_37, 0, x_36); -x_38 = lean_box(0); -x_24 = x_38; -x_25 = x_37; -goto block_32; +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = l_Lean_Syntax_getArg(x_22, x_9); +lean_dec(x_22); +x_36 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_36, 0, x_35); +x_37 = lean_box(0); +x_23 = x_37; +x_24 = x_36; +goto block_31; } } else { -lean_object* x_39; lean_object* x_40; -lean_dec(x_23); +lean_object* x_38; lean_object* x_39; +lean_dec(x_22); +x_38 = lean_box(0); x_39 = lean_box(0); -x_40 = lean_box(0); -x_24 = x_40; -x_25 = x_39; -goto block_32; +x_23 = x_39; +x_24 = x_38; +goto block_31; } -block_32: +block_31: { -lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; -lean_dec(x_25); +lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; lean_dec(x_24); -x_26 = lean_unsigned_to_nat(5u); -x_27 = l_Lean_Syntax_getArg(x_1, x_26); +lean_dec(x_23); +x_25 = lean_unsigned_to_nat(5u); +x_26 = l_Lean_Syntax_getArg(x_1, x_25); lean_dec(x_1); -x_28 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___closed__9; -x_29 = l_Lean_Syntax_isOfKind(x_27, x_28); -if (x_29 == 0) +x_27 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___closed__9; +x_28 = l_Lean_Syntax_isOfKind(x_26, x_27); +if (x_28 == 0) { -lean_object* x_30; -x_30 = lean_box(0); -return x_30; +lean_object* x_29; +x_29 = lean_box(0); +return x_29; } else { -lean_object* x_31; -x_31 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___closed__10; -return x_31; +lean_object* x_30; +x_30 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___closed__10; +return x_30; } } } } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_52; -lean_dec(x_16); -x_41 = lean_unsigned_to_nat(2u); -x_42 = l_Lean_Syntax_getArg(x_1, x_41); -x_52 = l_Lean_Syntax_isNone(x_42); -if (x_52 == 0) +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_51; +lean_dec(x_15); +x_40 = lean_unsigned_to_nat(2u); +x_41 = l_Lean_Syntax_getArg(x_1, x_40); +x_51 = l_Lean_Syntax_isNone(x_41); +if (x_51 == 0) { -uint8_t x_53; -lean_inc(x_42); -x_53 = l_Lean_Syntax_isNodeOf(x_42, x_7, x_5); -if (x_53 == 0) +uint8_t x_52; +lean_inc(x_41); +x_52 = l_Lean_Syntax_matchesNull(x_41, x_5); +if (x_52 == 0) { -lean_object* x_54; -lean_dec(x_42); +lean_object* x_53; +lean_dec(x_41); lean_dec(x_1); -x_54 = lean_box(0); -return x_54; +x_53 = lean_box(0); +return x_53; } else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_55 = l_Lean_Syntax_getArg(x_42, x_10); -lean_dec(x_42); -x_56 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_56, 0, x_55); -x_57 = lean_box(0); -x_43 = x_57; -x_44 = x_56; -goto block_51; +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = l_Lean_Syntax_getArg(x_41, x_9); +lean_dec(x_41); +x_55 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_55, 0, x_54); +x_56 = lean_box(0); +x_42 = x_56; +x_43 = x_55; +goto block_50; } } else { -lean_object* x_58; lean_object* x_59; -lean_dec(x_42); +lean_object* x_57; lean_object* x_58; +lean_dec(x_41); +x_57 = lean_box(0); x_58 = lean_box(0); -x_59 = lean_box(0); -x_43 = x_59; -x_44 = x_58; -goto block_51; +x_42 = x_58; +x_43 = x_57; +goto block_50; } -block_51: +block_50: { -lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; -lean_dec(x_44); +lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; lean_dec(x_43); -x_45 = lean_unsigned_to_nat(5u); -x_46 = l_Lean_Syntax_getArg(x_1, x_45); +lean_dec(x_42); +x_44 = lean_unsigned_to_nat(5u); +x_45 = l_Lean_Syntax_getArg(x_1, x_44); lean_dec(x_1); -x_47 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___closed__9; -x_48 = l_Lean_Syntax_isOfKind(x_46, x_47); -if (x_48 == 0) +x_46 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___closed__9; +x_47 = l_Lean_Syntax_isOfKind(x_45, x_46); +if (x_47 == 0) { -lean_object* x_49; -x_49 = lean_box(0); -return x_49; +lean_object* x_48; +x_48 = lean_box(0); +return x_48; } else { -lean_object* x_50; -x_50 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___closed__11; -return x_50; -} -} +lean_object* x_49; +x_49 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___closed__11; +return x_49; } } } } } } -static lean_object* _init_l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__2___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("matchAlt", 8); -return x_1; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__2(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { _start: { -uint8_t x_6; -x_6 = lean_usize_dec_eq(x_3, x_4); -if (x_6 == 0) +uint8_t x_7; +x_7 = lean_usize_dec_eq(x_4, x_5); +if (x_7 == 0) { -lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; size_t x_11; size_t x_12; -x_7 = lean_array_uget(x_2, x_3); -x_8 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__2___closed__1; -lean_inc(x_1); -x_9 = lean_name_mk_string(x_1, x_8); -lean_inc(x_7); -x_10 = l_Lean_Syntax_isOfKind(x_7, x_9); -lean_dec(x_9); -x_11 = 1; -x_12 = lean_usize_add(x_3, x_11); -if (x_10 == 0) +lean_object* x_8; uint8_t x_9; size_t x_10; size_t x_11; +x_8 = lean_array_uget(x_3, x_4); +lean_inc(x_8); +x_9 = l_Lean_Syntax_isOfKind(x_8, x_1); +x_10 = 1; +x_11 = lean_usize_add(x_4, x_10); +if (x_9 == 0) { -lean_dec(x_7); -x_3 = x_12; +lean_dec(x_8); +x_4 = x_11; goto _start; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_14 = lean_unsigned_to_nat(1u); -x_15 = l_Lean_Syntax_getArg(x_7, x_14); -x_16 = l_Lean_nullKind; -lean_inc(x_15); -x_17 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_14); -if (x_17 == 0) +lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_13 = lean_unsigned_to_nat(1u); +x_14 = l_Lean_Syntax_getArg(x_8, x_13); +lean_inc(x_14); +x_15 = l_Lean_Syntax_matchesNull(x_14, x_13); +if (x_15 == 0) { -lean_dec(x_15); -lean_dec(x_7); -x_3 = x_12; +lean_dec(x_14); +lean_dec(x_8); +x_4 = x_11; goto _start; } 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; -x_19 = lean_unsigned_to_nat(0u); -x_20 = l_Lean_Syntax_getArg(x_15, x_19); -lean_dec(x_15); -x_21 = lean_unsigned_to_nat(3u); -x_22 = l_Lean_Syntax_getArg(x_7, x_21); -x_23 = l_Lean_Syntax_getArgs(x_20); -lean_dec(x_20); -x_24 = l_Lean_Syntax_SepArray_getElems___rarg(x_23); -lean_dec(x_23); -x_25 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_25, 0, x_7); -lean_ctor_set(x_25, 1, x_24); -lean_ctor_set(x_25, 2, x_22); -x_26 = lean_array_push(x_5, x_25); -x_3 = x_12; -x_5 = x_26; +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_17 = lean_unsigned_to_nat(0u); +x_18 = l_Lean_Syntax_getArg(x_14, x_17); +lean_dec(x_14); +x_19 = lean_unsigned_to_nat(3u); +x_20 = l_Lean_Syntax_getArg(x_8, x_19); +x_21 = l_Lean_Syntax_getArgs(x_18); +lean_dec(x_18); +x_22 = l_Lean_Syntax_TSepArray_getElems___rarg(x_21); +lean_dec(x_21); +x_23 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_23, 0, x_8); +lean_ctor_set(x_23, 1, x_22); +lean_ctor_set(x_23, 2, x_20); +x_24 = lean_array_push(x_6, x_23); +x_4 = x_11; +x_6 = x_24; goto _start; } } } else { -lean_dec(x_1); -return x_5; +return x_6; } } } -LEAN_EXPORT lean_object* l_Array_filterMapM___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_filterMapM___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -uint8_t x_5; -x_5 = lean_nat_dec_lt(x_3, x_4); -if (x_5 == 0) +uint8_t x_6; +x_6 = lean_nat_dec_lt(x_4, x_5); +if (x_6 == 0) { -lean_object* x_6; +lean_object* x_7; +lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_6 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__15; -return x_6; +x_7 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__15; +return x_7; } else { -lean_object* x_7; uint8_t x_8; -x_7 = lean_array_get_size(x_2); -x_8 = lean_nat_dec_le(x_4, x_7); -lean_dec(x_7); -if (x_8 == 0) +lean_object* x_8; uint8_t x_9; +x_8 = lean_array_get_size(x_3); +x_9 = lean_nat_dec_le(x_5, x_8); +lean_dec(x_8); +if (x_9 == 0) { -lean_object* x_9; +lean_object* x_10; +lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_9 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__15; -return x_9; +x_10 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__15; +return x_10; } else { -size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; -x_10 = lean_usize_of_nat(x_3); -lean_dec(x_3); +size_t x_11; size_t x_12; lean_object* x_13; lean_object* x_14; x_11 = lean_usize_of_nat(x_4); lean_dec(x_4); -x_12 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__15; -x_13 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__2(x_1, x_2, x_10, x_11, x_12); -return x_13; +x_12 = lean_usize_of_nat(x_5); +lean_dec(x_5); +x_13 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__15; +x_14 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__2(x_1, x_2, x_3, x_11, x_12, x_13); +return x_14; } } } } +static lean_object* _init_l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("matchAlt", 8); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__6; +x_2 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts(lean_object* x_1) { _start: { @@ -6766,56 +6779,54 @@ x_6 = l_Lean_Syntax_getArg(x_1, x_5); x_36 = l_Lean_Syntax_isNone(x_6); if (x_36 == 0) { -lean_object* x_37; uint8_t x_38; -x_37 = l_Lean_nullKind; +uint8_t x_37; lean_inc(x_6); -x_38 = l_Lean_Syntax_isNodeOf(x_6, x_37, x_5); -if (x_38 == 0) +x_37 = l_Lean_Syntax_matchesNull(x_6, x_5); +if (x_37 == 0) { -lean_object* x_39; +lean_object* x_38; lean_dec(x_6); lean_dec(x_1); -x_39 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__15; -return x_39; +x_38 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__15; +return x_38; } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_40 = lean_unsigned_to_nat(0u); -x_41 = l_Lean_Syntax_getArg(x_6, x_40); +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_39 = lean_unsigned_to_nat(0u); +x_40 = l_Lean_Syntax_getArg(x_6, x_39); lean_dec(x_6); -x_42 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_42, 0, x_41); -x_43 = lean_box(0); -x_7 = x_43; -x_8 = x_42; +x_41 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_41, 0, x_40); +x_42 = lean_box(0); +x_7 = x_42; +x_8 = x_41; goto block_35; } } else { -lean_object* x_44; lean_object* x_45; +lean_object* x_43; lean_object* x_44; lean_dec(x_6); +x_43 = lean_box(0); x_44 = lean_box(0); -x_45 = lean_box(0); -x_7 = x_45; -x_8 = x_44; +x_7 = x_44; +x_8 = x_43; goto block_35; } block_35: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_25; +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_26; lean_dec(x_8); lean_dec(x_7); x_9 = lean_unsigned_to_nat(2u); x_10 = l_Lean_Syntax_getArg(x_1, x_9); -x_25 = l_Lean_Syntax_isNone(x_10); -if (x_25 == 0) +x_26 = l_Lean_Syntax_isNone(x_10); +if (x_26 == 0) { -lean_object* x_26; uint8_t x_27; -x_26 = l_Lean_nullKind; +uint8_t x_27; lean_inc(x_10); -x_27 = l_Lean_Syntax_isNodeOf(x_10, x_26, x_5); +x_27 = l_Lean_Syntax_matchesNull(x_10, x_5); if (x_27 == 0) { lean_object* x_28; @@ -6835,7 +6846,7 @@ lean_ctor_set(x_31, 0, x_30); x_32 = lean_box(0); x_11 = x_32; x_12 = x_31; -goto block_24; +goto block_25; } } else @@ -6846,9 +6857,9 @@ x_33 = lean_box(0); x_34 = lean_box(0); x_11 = x_34; x_12 = x_33; -goto block_24; +goto block_25; } -block_24: +block_25: { lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_dec(x_12); @@ -6868,43 +6879,48 @@ return x_17; } else { -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_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_18 = lean_unsigned_to_nat(0u); x_19 = l_Lean_Syntax_getArg(x_14, x_18); lean_dec(x_14); -x_20 = l_Lean_Syntax_getArgs(x_19); +x_20 = lean_box(0); +x_21 = l_Lean_Syntax_getArgs(x_19); lean_dec(x_19); -x_21 = lean_array_get_size(x_20); -x_22 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__6; -x_23 = l_Array_filterMapM___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__1(x_22, x_20, x_18, x_21); -lean_dec(x_20); -return x_23; +x_22 = lean_array_get_size(x_21); +x_23 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___closed__2; +x_24 = l_Array_filterMapM___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__1(x_23, x_20, x_21, x_18, x_22); +lean_dec(x_21); +return x_24; } } } } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__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: { -size_t x_6; size_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_3); -lean_dec(x_3); +size_t x_7; size_t x_8; lean_object* x_9; x_7 = lean_unbox_usize(x_4); lean_dec(x_4); -x_8 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__2(x_1, x_2, x_6, x_7, x_5); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__2(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); lean_dec(x_2); -return x_8; +lean_dec(x_1); +return x_9; } } -LEAN_EXPORT lean_object* l_Array_filterMapM___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_filterMapM___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__1___boxed(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; -x_5 = l_Array_filterMapM___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__1(x_1, x_2, x_3, x_4); +lean_object* x_6; +x_6 = l_Array_filterMapM___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__1(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_3); lean_dec(x_2); -return x_5; +lean_dec(x_1); +return x_6; } } LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabInaccessible(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_object* x_9) { @@ -7235,9 +7251,9 @@ LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Elab_Term_precheckMatch__ { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_3 = lean_array_get_size(x_1); -x_4 = lean_unsigned_to_nat(0u); -x_5 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__15; -x_6 = l_Array_sequenceMap_loop___at_Lean_Elab_Term_precheckMatch___spec__2(x_1, x_2, x_3, x_4, x_5); +x_4 = lean_mk_empty_array_with_capacity(x_3); +x_5 = lean_unsigned_to_nat(0u); +x_6 = l_Array_sequenceMap_loop___at_Lean_Elab_Term_precheckMatch___spec__2(x_1, x_2, x_3, x_5, x_4); return x_6; } } @@ -7591,27 +7607,26 @@ return x_6; } else { -lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +lean_object* x_7; lean_object* x_8; uint8_t x_9; x_7 = lean_unsigned_to_nat(0u); x_8 = l_Lean_Syntax_getArg(x_2, x_7); -x_9 = l_Lean_nullKind; -x_10 = l_Lean_Syntax_isNodeOf(x_8, x_9, x_7); -if (x_10 == 0) +x_9 = l_Lean_Syntax_matchesNull(x_8, x_7); +if (x_9 == 0) { -lean_object* x_11; +lean_object* x_10; lean_dec(x_2); -x_11 = lean_box(0); -return x_11; +x_10 = lean_box(0); +return x_10; } else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_unsigned_to_nat(1u); -x_13 = l_Lean_Syntax_getArg(x_2, x_12); +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_unsigned_to_nat(1u); +x_12 = l_Lean_Syntax_getArg(x_2, x_11); lean_dec(x_2); -x_14 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_14, 0, x_13); -return x_14; +x_13 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_13, 0, x_12); +return x_13; } } } @@ -7637,7 +7652,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_precheckMatch___lambda__3(lean_object* _start: { lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_5 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__2___closed__1; +x_5 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___closed__1; x_6 = lean_name_mk_string(x_1, x_5); lean_inc(x_4); x_7 = l_Lean_Syntax_isOfKind(x_4, x_6); @@ -7647,44 +7662,79 @@ if (x_7 == 0) lean_object* x_8; lean_dec(x_4); lean_dec(x_3); +lean_dec(x_2); x_8 = lean_box(0); return x_8; } else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +lean_object* x_9; lean_object* x_10; uint8_t x_11; x_9 = lean_unsigned_to_nat(1u); x_10 = l_Lean_Syntax_getArg(x_4, x_9); -x_11 = l_Lean_nullKind; lean_inc(x_10); -x_12 = l_Lean_Syntax_isNodeOf(x_10, x_11, x_9); -if (x_12 == 0) +x_11 = l_Lean_Syntax_matchesNull(x_10, x_9); +if (x_11 == 0) { -lean_object* x_13; +lean_object* x_12; lean_dec(x_10); lean_dec(x_4); lean_dec(x_3); -x_13 = lean_box(0); -return x_13; +lean_dec(x_2); +x_12 = lean_box(0); +return x_12; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_14 = lean_unsigned_to_nat(0u); -x_15 = l_Lean_Syntax_getArg(x_10, x_14); +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; +x_13 = lean_unsigned_to_nat(0u); +x_14 = l_Lean_Syntax_getArg(x_10, x_13); lean_dec(x_10); -x_16 = l_Lean_Syntax_getArgs(x_15); +x_15 = l_Lean_Syntax_getArgs(x_14); +lean_dec(x_14); +x_16 = lean_array_get_size(x_15); +x_17 = lean_nat_dec_lt(x_13, x_16); +if (x_17 == 0) +{ +lean_dec(x_16); lean_dec(x_15); -x_17 = lean_array_get_size(x_16); -x_18 = lean_nat_dec_lt(x_14, x_17); -if (x_18 == 0) +lean_dec(x_3); +x_18 = x_2; +goto block_32; +} +else +{ +uint8_t x_33; +x_33 = lean_nat_dec_le(x_16, x_16); +if (x_33 == 0) { -lean_object* x_19; lean_object* x_20; -lean_dec(x_17); lean_dec(x_16); +lean_dec(x_15); lean_dec(x_3); +x_18 = x_2; +goto block_32; +} +else +{ +size_t x_34; size_t x_35; lean_object* x_36; lean_object* x_37; +lean_dec(x_2); +x_34 = 0; +x_35 = lean_usize_of_nat(x_16); +lean_dec(x_16); +x_36 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(x_15, x_34, x_35, x_3); +lean_dec(x_15); +x_37 = lean_ctor_get(x_36, 1); +lean_inc(x_37); +lean_dec(x_36); +x_18 = x_37; +goto block_32; +} +} +block_32: +{ +lean_object* x_19; lean_object* x_20; x_19 = l_Lean_Elab_Term_precheckMatch___lambda__3___closed__1; -x_20 = l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__1(x_2, x_19); +x_20 = l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__1(x_18, x_19); +lean_dec(x_18); if (lean_obj_tag(x_20) == 0) { lean_object* x_21; @@ -7727,117 +7777,6 @@ return x_31; } } } -else -{ -uint8_t x_32; -x_32 = lean_nat_dec_le(x_17, x_17); -if (x_32 == 0) -{ -lean_object* x_33; lean_object* x_34; -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_3); -x_33 = l_Lean_Elab_Term_precheckMatch___lambda__3___closed__1; -x_34 = l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__1(x_2, x_33); -if (lean_obj_tag(x_34) == 0) -{ -lean_object* x_35; -lean_dec(x_4); -x_35 = lean_box(0); -return x_35; -} -else -{ -uint8_t x_36; -x_36 = !lean_is_exclusive(x_34); -if (x_36 == 0) -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_37 = lean_ctor_get(x_34, 0); -x_38 = lean_unsigned_to_nat(3u); -x_39 = l_Lean_Syntax_getArg(x_4, x_38); -lean_dec(x_4); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_37); -lean_ctor_set(x_40, 1, x_39); -lean_ctor_set(x_34, 0, x_40); -return x_34; -} -else -{ -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_41 = lean_ctor_get(x_34, 0); -lean_inc(x_41); -lean_dec(x_34); -x_42 = lean_unsigned_to_nat(3u); -x_43 = l_Lean_Syntax_getArg(x_4, x_42); -lean_dec(x_4); -x_44 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_44, 0, x_41); -lean_ctor_set(x_44, 1, x_43); -x_45 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_45, 0, x_44); -return x_45; -} -} -} -else -{ -size_t x_46; size_t x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_46 = 0; -x_47 = lean_usize_of_nat(x_17); -lean_dec(x_17); -x_48 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(x_16, x_46, x_47, x_3); -lean_dec(x_16); -x_49 = lean_ctor_get(x_48, 1); -lean_inc(x_49); -lean_dec(x_48); -x_50 = l_Lean_Elab_Term_precheckMatch___lambda__3___closed__1; -x_51 = l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__1(x_49, x_50); -lean_dec(x_49); -if (lean_obj_tag(x_51) == 0) -{ -lean_object* x_52; -lean_dec(x_4); -x_52 = lean_box(0); -return x_52; -} -else -{ -uint8_t x_53; -x_53 = !lean_is_exclusive(x_51); -if (x_53 == 0) -{ -lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_54 = lean_ctor_get(x_51, 0); -x_55 = lean_unsigned_to_nat(3u); -x_56 = l_Lean_Syntax_getArg(x_4, x_55); -lean_dec(x_4); -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_54); -lean_ctor_set(x_57, 1, x_56); -lean_ctor_set(x_51, 0, x_57); -return x_51; -} -else -{ -lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_58 = lean_ctor_get(x_51, 0); -lean_inc(x_58); -lean_dec(x_51); -x_59 = lean_unsigned_to_nat(3u); -x_60 = l_Lean_Syntax_getArg(x_4, x_59); -lean_dec(x_4); -x_61 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_61, 0, x_58); -lean_ctor_set(x_61, 1, x_60); -x_62 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_62, 0, x_61); -return x_62; -} -} -} -} } } } @@ -7890,15 +7829,14 @@ return x_12; } else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; x_13 = lean_unsigned_to_nat(1u); x_14 = l_Lean_Syntax_getArg(x_1, x_13); -x_15 = l_Lean_nullKind; -x_16 = lean_unsigned_to_nat(0u); -x_17 = l_Lean_Syntax_isNodeOf(x_14, x_15, x_16); -if (x_17 == 0) +x_15 = lean_unsigned_to_nat(0u); +x_16 = l_Lean_Syntax_matchesNull(x_14, x_15); +if (x_16 == 0) { -lean_object* x_18; +lean_object* x_17; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -7907,18 +7845,18 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_18 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_precheckIdent___spec__1___rarg(x_9); -return x_18; +x_17 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_precheckIdent___spec__1___rarg(x_9); +return x_17; } else { -lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_19 = lean_unsigned_to_nat(2u); -x_20 = l_Lean_Syntax_getArg(x_1, x_19); -x_21 = l_Lean_Syntax_isNodeOf(x_20, x_15, x_16); -if (x_21 == 0) +lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_18 = lean_unsigned_to_nat(2u); +x_19 = l_Lean_Syntax_getArg(x_1, x_18); +x_20 = l_Lean_Syntax_matchesNull(x_19, x_15); +if (x_20 == 0) { -lean_object* x_22; +lean_object* x_21; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -7927,67 +7865,67 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_22 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_precheckIdent___spec__1___rarg(x_9); -return x_22; +x_21 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_precheckIdent___spec__1___rarg(x_9); +return x_21; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_23 = lean_unsigned_to_nat(3u); -x_24 = l_Lean_Syntax_getArg(x_1, x_23); -x_25 = l_Lean_Syntax_getArgs(x_24); -lean_dec(x_24); -x_26 = lean_array_get_size(x_25); -x_27 = lean_nat_dec_lt(x_16, x_26); -x_28 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__6; -x_29 = lean_alloc_closure((void*)(l_Lean_Elab_Term_precheckMatch___lambda__1), 2, 1); -lean_closure_set(x_29, 0, x_28); -if (x_27 == 0) +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_22 = lean_unsigned_to_nat(3u); +x_23 = l_Lean_Syntax_getArg(x_1, x_22); +x_24 = l_Lean_Syntax_getArgs(x_23); +lean_dec(x_23); +x_25 = lean_array_get_size(x_24); +x_26 = lean_nat_dec_lt(x_15, x_25); +x_27 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__6; +x_28 = lean_alloc_closure((void*)(l_Lean_Elab_Term_precheckMatch___lambda__1), 2, 1); +lean_closure_set(x_28, 0, x_27); +if (x_26 == 0) { -lean_object* x_131; -lean_dec(x_26); +lean_object* x_130; lean_dec(x_25); -x_131 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__15; -x_30 = x_131; -goto block_130; +lean_dec(x_24); +x_130 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__15; +x_29 = x_130; +goto block_129; } else { -uint8_t x_132; -x_132 = lean_nat_dec_le(x_26, x_26); -if (x_132 == 0) +uint8_t x_131; +x_131 = lean_nat_dec_le(x_25, x_25); +if (x_131 == 0) { -lean_object* x_133; -lean_dec(x_26); +lean_object* x_132; lean_dec(x_25); -x_133 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__15; -x_30 = x_133; -goto block_130; +lean_dec(x_24); +x_132 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__15; +x_29 = x_132; +goto block_129; } else { -size_t x_134; size_t x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; -x_134 = 0; -x_135 = lean_usize_of_nat(x_26); -lean_dec(x_26); -x_136 = l_Lean_Elab_Term_precheckMatch___closed__1; -x_137 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(x_25, x_134, x_135, x_136); +size_t x_133; size_t x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; +x_133 = 0; +x_134 = lean_usize_of_nat(x_25); lean_dec(x_25); -x_138 = lean_ctor_get(x_137, 1); -lean_inc(x_138); -lean_dec(x_137); -x_30 = x_138; -goto block_130; +x_135 = l_Lean_Elab_Term_precheckMatch___closed__1; +x_136 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(x_24, x_133, x_134, x_135); +lean_dec(x_24); +x_137 = lean_ctor_get(x_136, 1); +lean_inc(x_137); +lean_dec(x_136); +x_29 = x_137; +goto block_129; } } -block_130: +block_129: { -lean_object* x_31; -x_31 = l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__1(x_30, x_29); -lean_dec(x_30); -if (lean_obj_tag(x_31) == 0) +lean_object* x_30; +x_30 = l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__1(x_29, x_28); +lean_dec(x_29); +if (lean_obj_tag(x_30) == 0) { -lean_object* x_32; +lean_object* x_31; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -7996,26 +7934,26 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_32 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_precheckIdent___spec__1___rarg(x_9); -return x_32; +x_31 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_precheckIdent___spec__1___rarg(x_9); +return x_31; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_33 = lean_ctor_get(x_31, 0); -lean_inc(x_33); -lean_dec(x_31); -x_34 = lean_unsigned_to_nat(5u); -x_35 = l_Lean_Syntax_getArg(x_1, x_34); +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_32 = lean_ctor_get(x_30, 0); +lean_inc(x_32); +lean_dec(x_30); +x_33 = lean_unsigned_to_nat(5u); +x_34 = l_Lean_Syntax_getArg(x_1, x_33); lean_dec(x_1); -x_36 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___closed__9; -lean_inc(x_35); -x_37 = l_Lean_Syntax_isOfKind(x_35, x_36); -if (x_37 == 0) +x_35 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___closed__9; +lean_inc(x_34); +x_36 = l_Lean_Syntax_isOfKind(x_34, x_35); +if (x_36 == 0) { -lean_object* x_38; -lean_dec(x_35); -lean_dec(x_33); +lean_object* x_37; +lean_dec(x_34); +lean_dec(x_32); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -8023,28 +7961,28 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_38 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_precheckIdent___spec__1___rarg(x_9); -return x_38; +x_37 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_precheckIdent___spec__1___rarg(x_9); +return x_37; } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_39 = l_Lean_Syntax_getArg(x_35, x_16); -lean_dec(x_35); -x_40 = l_Lean_Syntax_getArgs(x_39); +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_38 = l_Lean_Syntax_getArg(x_34, x_15); +lean_dec(x_34); +x_39 = l_Lean_Syntax_getArgs(x_38); +lean_dec(x_38); +x_40 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__15; +x_41 = l_Lean_Elab_Term_precheckMatch___closed__1; +x_42 = lean_alloc_closure((void*)(l_Lean_Elab_Term_precheckMatch___lambda__3), 4, 3); +lean_closure_set(x_42, 0, x_27); +lean_closure_set(x_42, 1, x_40); +lean_closure_set(x_42, 2, x_41); +x_43 = l_Array_sequenceMap___at_Lean_Elab_Term_precheckMatch___spec__1(x_39, x_42); lean_dec(x_39); -x_41 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__15; -x_42 = l_Lean_Elab_Term_precheckMatch___closed__1; -x_43 = lean_alloc_closure((void*)(l_Lean_Elab_Term_precheckMatch___lambda__3___boxed), 4, 3); -lean_closure_set(x_43, 0, x_28); -lean_closure_set(x_43, 1, x_41); -lean_closure_set(x_43, 2, x_42); -x_44 = l_Array_sequenceMap___at_Lean_Elab_Term_precheckMatch___spec__1(x_40, x_43); -lean_dec(x_40); -if (lean_obj_tag(x_44) == 0) +if (lean_obj_tag(x_43) == 0) { -lean_object* x_45; -lean_dec(x_33); +lean_object* x_44; +lean_dec(x_32); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -8052,238 +7990,238 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_45 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_precheckIdent___spec__1___rarg(x_9); -return x_45; +x_44 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_precheckIdent___spec__1___rarg(x_9); +return x_44; } else { -lean_object* x_46; lean_object* x_47; size_t x_48; size_t x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; -x_46 = lean_ctor_get(x_44, 0); -lean_inc(x_46); -lean_dec(x_44); -x_47 = lean_array_get_size(x_46); -x_48 = lean_usize_of_nat(x_47); -lean_dec(x_47); -x_49 = 0; -lean_inc(x_46); -x_50 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_precheckMatch___spec__3(x_48, x_49, x_46); -x_51 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_precheckMatch___spec__4(x_48, x_49, x_46); -x_52 = lean_array_get_size(x_33); -x_53 = lean_nat_dec_lt(x_16, x_52); -if (x_53 == 0) +lean_object* x_45; lean_object* x_46; size_t x_47; size_t x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; +x_45 = lean_ctor_get(x_43, 0); +lean_inc(x_45); +lean_dec(x_43); +x_46 = lean_array_get_size(x_45); +x_47 = lean_usize_of_nat(x_46); +lean_dec(x_46); +x_48 = 0; +lean_inc(x_45); +x_49 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_precheckMatch___spec__3(x_47, x_48, x_45); +x_50 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_precheckMatch___spec__4(x_47, x_48, x_45); +x_51 = lean_array_get_size(x_32); +x_52 = lean_nat_dec_lt(x_15, x_51); +if (x_52 == 0) { -lean_object* x_54; lean_object* x_55; size_t x_56; lean_object* x_57; lean_object* x_58; -lean_dec(x_52); -lean_dec(x_33); -x_54 = l_Array_zip___rarg(x_51, x_50); -lean_dec(x_50); +lean_object* x_53; lean_object* x_54; size_t x_55; lean_object* x_56; lean_object* x_57; lean_dec(x_51); -x_55 = lean_array_get_size(x_54); -x_56 = lean_usize_of_nat(x_55); -lean_dec(x_55); -x_57 = l_Lean_Elab_Term_precheckMatch___closed__2; -x_58 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_precheckMatch___spec__5(x_57, x_54, x_56, x_49, x_57, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_32); +x_53 = l_Array_zip___rarg(x_50, x_49); +lean_dec(x_49); +lean_dec(x_50); +x_54 = lean_array_get_size(x_53); +x_55 = lean_usize_of_nat(x_54); lean_dec(x_54); -if (lean_obj_tag(x_58) == 0) +x_56 = l_Lean_Elab_Term_precheckMatch___closed__2; +x_57 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_precheckMatch___spec__5(x_56, x_53, x_55, x_48, x_56, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_53); +if (lean_obj_tag(x_57) == 0) { -lean_object* x_59; lean_object* x_60; +lean_object* x_58; lean_object* x_59; +x_58 = lean_ctor_get(x_57, 0); +lean_inc(x_58); x_59 = lean_ctor_get(x_58, 0); lean_inc(x_59); -x_60 = lean_ctor_get(x_59, 0); -lean_inc(x_60); -lean_dec(x_59); -if (lean_obj_tag(x_60) == 0) +lean_dec(x_58); +if (lean_obj_tag(x_59) == 0) { -uint8_t x_61; -x_61 = !lean_is_exclusive(x_58); -if (x_61 == 0) +uint8_t x_60; +x_60 = !lean_is_exclusive(x_57); +if (x_60 == 0) { -lean_object* x_62; lean_object* x_63; -x_62 = lean_ctor_get(x_58, 0); -lean_dec(x_62); -x_63 = lean_box(0); -lean_ctor_set(x_58, 0, x_63); -return x_58; +lean_object* x_61; lean_object* x_62; +x_61 = lean_ctor_get(x_57, 0); +lean_dec(x_61); +x_62 = lean_box(0); +lean_ctor_set(x_57, 0, x_62); +return x_57; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_ctor_get(x_58, 1); -lean_inc(x_64); -lean_dec(x_58); -x_65 = lean_box(0); -x_66 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_66, 0, x_65); -lean_ctor_set(x_66, 1, x_64); -return x_66; +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_57, 1); +lean_inc(x_63); +lean_dec(x_57); +x_64 = lean_box(0); +x_65 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_65, 0, x_64); +lean_ctor_set(x_65, 1, x_63); +return x_65; } } else { -uint8_t x_67; -x_67 = !lean_is_exclusive(x_58); -if (x_67 == 0) +uint8_t x_66; +x_66 = !lean_is_exclusive(x_57); +if (x_66 == 0) { -lean_object* x_68; lean_object* x_69; -x_68 = lean_ctor_get(x_58, 0); -lean_dec(x_68); -x_69 = lean_ctor_get(x_60, 0); -lean_inc(x_69); -lean_dec(x_60); -lean_ctor_set(x_58, 0, x_69); -return x_58; +lean_object* x_67; lean_object* x_68; +x_67 = lean_ctor_get(x_57, 0); +lean_dec(x_67); +x_68 = lean_ctor_get(x_59, 0); +lean_inc(x_68); +lean_dec(x_59); +lean_ctor_set(x_57, 0, x_68); +return x_57; } else { -lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_70 = lean_ctor_get(x_58, 1); +lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_69 = lean_ctor_get(x_57, 1); +lean_inc(x_69); +lean_dec(x_57); +x_70 = lean_ctor_get(x_59, 0); lean_inc(x_70); -lean_dec(x_58); -x_71 = lean_ctor_get(x_60, 0); -lean_inc(x_71); -lean_dec(x_60); -x_72 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_72, 0, x_71); -lean_ctor_set(x_72, 1, x_70); -return x_72; +lean_dec(x_59); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set(x_71, 1, x_69); +return x_71; } } } else { -uint8_t x_73; -x_73 = !lean_is_exclusive(x_58); -if (x_73 == 0) +uint8_t x_72; +x_72 = !lean_is_exclusive(x_57); +if (x_72 == 0) { -return x_58; +return x_57; } else { -lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_74 = lean_ctor_get(x_58, 0); -x_75 = lean_ctor_get(x_58, 1); -lean_inc(x_75); +lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_73 = lean_ctor_get(x_57, 0); +x_74 = lean_ctor_get(x_57, 1); lean_inc(x_74); -lean_dec(x_58); -x_76 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_76, 0, x_74); -lean_ctor_set(x_76, 1, x_75); -return x_76; +lean_inc(x_73); +lean_dec(x_57); +x_75 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set(x_75, 1, x_74); +return x_75; } } } else { -uint8_t x_77; -x_77 = lean_nat_dec_le(x_52, x_52); -if (x_77 == 0) +uint8_t x_76; +x_76 = lean_nat_dec_le(x_51, x_51); +if (x_76 == 0) { -lean_object* x_78; lean_object* x_79; size_t x_80; lean_object* x_81; lean_object* x_82; -lean_dec(x_52); -lean_dec(x_33); -x_78 = l_Array_zip___rarg(x_51, x_50); -lean_dec(x_50); +lean_object* x_77; lean_object* x_78; size_t x_79; lean_object* x_80; lean_object* x_81; lean_dec(x_51); -x_79 = lean_array_get_size(x_78); -x_80 = lean_usize_of_nat(x_79); -lean_dec(x_79); -x_81 = l_Lean_Elab_Term_precheckMatch___closed__2; -x_82 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_precheckMatch___spec__5(x_81, x_78, x_80, x_49, x_81, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_32); +x_77 = l_Array_zip___rarg(x_50, x_49); +lean_dec(x_49); +lean_dec(x_50); +x_78 = lean_array_get_size(x_77); +x_79 = lean_usize_of_nat(x_78); lean_dec(x_78); -if (lean_obj_tag(x_82) == 0) +x_80 = l_Lean_Elab_Term_precheckMatch___closed__2; +x_81 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_precheckMatch___spec__5(x_80, x_77, x_79, x_48, x_80, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_77); +if (lean_obj_tag(x_81) == 0) { -lean_object* x_83; lean_object* x_84; +lean_object* x_82; lean_object* x_83; +x_82 = lean_ctor_get(x_81, 0); +lean_inc(x_82); x_83 = lean_ctor_get(x_82, 0); lean_inc(x_83); -x_84 = lean_ctor_get(x_83, 0); -lean_inc(x_84); -lean_dec(x_83); -if (lean_obj_tag(x_84) == 0) +lean_dec(x_82); +if (lean_obj_tag(x_83) == 0) { -uint8_t x_85; -x_85 = !lean_is_exclusive(x_82); -if (x_85 == 0) +uint8_t x_84; +x_84 = !lean_is_exclusive(x_81); +if (x_84 == 0) { -lean_object* x_86; lean_object* x_87; -x_86 = lean_ctor_get(x_82, 0); -lean_dec(x_86); -x_87 = lean_box(0); -lean_ctor_set(x_82, 0, x_87); -return x_82; +lean_object* x_85; lean_object* x_86; +x_85 = lean_ctor_get(x_81, 0); +lean_dec(x_85); +x_86 = lean_box(0); +lean_ctor_set(x_81, 0, x_86); +return x_81; } else { -lean_object* x_88; lean_object* x_89; lean_object* x_90; -x_88 = lean_ctor_get(x_82, 1); -lean_inc(x_88); -lean_dec(x_82); -x_89 = lean_box(0); -x_90 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_90, 0, x_89); -lean_ctor_set(x_90, 1, x_88); -return x_90; +lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_87 = lean_ctor_get(x_81, 1); +lean_inc(x_87); +lean_dec(x_81); +x_88 = lean_box(0); +x_89 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_89, 0, x_88); +lean_ctor_set(x_89, 1, x_87); +return x_89; } } else { -uint8_t x_91; -x_91 = !lean_is_exclusive(x_82); -if (x_91 == 0) +uint8_t x_90; +x_90 = !lean_is_exclusive(x_81); +if (x_90 == 0) { -lean_object* x_92; lean_object* x_93; -x_92 = lean_ctor_get(x_82, 0); -lean_dec(x_92); -x_93 = lean_ctor_get(x_84, 0); -lean_inc(x_93); -lean_dec(x_84); -lean_ctor_set(x_82, 0, x_93); -return x_82; +lean_object* x_91; lean_object* x_92; +x_91 = lean_ctor_get(x_81, 0); +lean_dec(x_91); +x_92 = lean_ctor_get(x_83, 0); +lean_inc(x_92); +lean_dec(x_83); +lean_ctor_set(x_81, 0, x_92); +return x_81; } else { -lean_object* x_94; lean_object* x_95; lean_object* x_96; -x_94 = lean_ctor_get(x_82, 1); +lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_93 = lean_ctor_get(x_81, 1); +lean_inc(x_93); +lean_dec(x_81); +x_94 = lean_ctor_get(x_83, 0); lean_inc(x_94); -lean_dec(x_82); -x_95 = lean_ctor_get(x_84, 0); -lean_inc(x_95); -lean_dec(x_84); -x_96 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_96, 0, x_95); -lean_ctor_set(x_96, 1, x_94); -return x_96; +lean_dec(x_83); +x_95 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_95, 0, x_94); +lean_ctor_set(x_95, 1, x_93); +return x_95; } } } else { -uint8_t x_97; -x_97 = !lean_is_exclusive(x_82); -if (x_97 == 0) +uint8_t x_96; +x_96 = !lean_is_exclusive(x_81); +if (x_96 == 0) { -return x_82; +return x_81; } else { -lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_98 = lean_ctor_get(x_82, 0); -x_99 = lean_ctor_get(x_82, 1); -lean_inc(x_99); +lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_97 = lean_ctor_get(x_81, 0); +x_98 = lean_ctor_get(x_81, 1); lean_inc(x_98); -lean_dec(x_82); -x_100 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_100, 0, x_98); -lean_ctor_set(x_100, 1, x_99); -return x_100; +lean_inc(x_97); +lean_dec(x_81); +x_99 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_99, 0, x_97); +lean_ctor_set(x_99, 1, x_98); +return x_99; } } } else { -size_t x_101; lean_object* x_102; lean_object* x_103; -x_101 = lean_usize_of_nat(x_52); -lean_dec(x_52); -x_102 = lean_box(0); +size_t x_100; lean_object* x_101; lean_object* x_102; +x_100 = lean_usize_of_nat(x_51); +lean_dec(x_51); +x_101 = lean_box(0); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); @@ -8291,114 +8229,114 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_103 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Quotation_precheckParen___spec__1(x_33, x_49, x_101, x_102, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_33); -if (lean_obj_tag(x_103) == 0) +x_102 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Quotation_precheckParen___spec__1(x_32, x_48, x_100, x_101, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_32); +if (lean_obj_tag(x_102) == 0) { -lean_object* x_104; lean_object* x_105; lean_object* x_106; size_t x_107; lean_object* x_108; lean_object* x_109; -x_104 = lean_ctor_get(x_103, 1); -lean_inc(x_104); -lean_dec(x_103); -x_105 = l_Array_zip___rarg(x_51, x_50); +lean_object* x_103; lean_object* x_104; lean_object* x_105; size_t x_106; lean_object* x_107; lean_object* x_108; +x_103 = lean_ctor_get(x_102, 1); +lean_inc(x_103); +lean_dec(x_102); +x_104 = l_Array_zip___rarg(x_50, x_49); +lean_dec(x_49); lean_dec(x_50); -lean_dec(x_51); -x_106 = lean_array_get_size(x_105); -x_107 = lean_usize_of_nat(x_106); -lean_dec(x_106); -x_108 = l_Lean_Elab_Term_precheckMatch___closed__2; -x_109 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_precheckMatch___spec__5(x_108, x_105, x_107, x_49, x_108, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_104); +x_105 = lean_array_get_size(x_104); +x_106 = lean_usize_of_nat(x_105); lean_dec(x_105); -if (lean_obj_tag(x_109) == 0) +x_107 = l_Lean_Elab_Term_precheckMatch___closed__2; +x_108 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_precheckMatch___spec__5(x_107, x_104, x_106, x_48, x_107, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_103); +lean_dec(x_104); +if (lean_obj_tag(x_108) == 0) { -lean_object* x_110; lean_object* x_111; +lean_object* x_109; lean_object* x_110; +x_109 = lean_ctor_get(x_108, 0); +lean_inc(x_109); x_110 = lean_ctor_get(x_109, 0); lean_inc(x_110); -x_111 = lean_ctor_get(x_110, 0); -lean_inc(x_111); -lean_dec(x_110); -if (lean_obj_tag(x_111) == 0) +lean_dec(x_109); +if (lean_obj_tag(x_110) == 0) { -uint8_t x_112; -x_112 = !lean_is_exclusive(x_109); -if (x_112 == 0) +uint8_t x_111; +x_111 = !lean_is_exclusive(x_108); +if (x_111 == 0) { -lean_object* x_113; -x_113 = lean_ctor_get(x_109, 0); -lean_dec(x_113); -lean_ctor_set(x_109, 0, x_102); -return x_109; +lean_object* x_112; +x_112 = lean_ctor_get(x_108, 0); +lean_dec(x_112); +lean_ctor_set(x_108, 0, x_101); +return x_108; } else { -lean_object* x_114; lean_object* x_115; -x_114 = lean_ctor_get(x_109, 1); -lean_inc(x_114); -lean_dec(x_109); -x_115 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_115, 0, x_102); -lean_ctor_set(x_115, 1, x_114); -return x_115; +lean_object* x_113; lean_object* x_114; +x_113 = lean_ctor_get(x_108, 1); +lean_inc(x_113); +lean_dec(x_108); +x_114 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_114, 0, x_101); +lean_ctor_set(x_114, 1, x_113); +return x_114; } } else { -uint8_t x_116; -x_116 = !lean_is_exclusive(x_109); -if (x_116 == 0) +uint8_t x_115; +x_115 = !lean_is_exclusive(x_108); +if (x_115 == 0) { -lean_object* x_117; lean_object* x_118; -x_117 = lean_ctor_get(x_109, 0); -lean_dec(x_117); -x_118 = lean_ctor_get(x_111, 0); -lean_inc(x_118); -lean_dec(x_111); -lean_ctor_set(x_109, 0, x_118); -return x_109; +lean_object* x_116; lean_object* x_117; +x_116 = lean_ctor_get(x_108, 0); +lean_dec(x_116); +x_117 = lean_ctor_get(x_110, 0); +lean_inc(x_117); +lean_dec(x_110); +lean_ctor_set(x_108, 0, x_117); +return x_108; } else { -lean_object* x_119; lean_object* x_120; lean_object* x_121; -x_119 = lean_ctor_get(x_109, 1); +lean_object* x_118; lean_object* x_119; lean_object* x_120; +x_118 = lean_ctor_get(x_108, 1); +lean_inc(x_118); +lean_dec(x_108); +x_119 = lean_ctor_get(x_110, 0); lean_inc(x_119); -lean_dec(x_109); -x_120 = lean_ctor_get(x_111, 0); -lean_inc(x_120); -lean_dec(x_111); -x_121 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_121, 0, x_120); -lean_ctor_set(x_121, 1, x_119); -return x_121; +lean_dec(x_110); +x_120 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_120, 0, x_119); +lean_ctor_set(x_120, 1, x_118); +return x_120; } } } else { -uint8_t x_122; -x_122 = !lean_is_exclusive(x_109); -if (x_122 == 0) +uint8_t x_121; +x_121 = !lean_is_exclusive(x_108); +if (x_121 == 0) { -return x_109; +return x_108; } else { -lean_object* x_123; lean_object* x_124; lean_object* x_125; -x_123 = lean_ctor_get(x_109, 0); -x_124 = lean_ctor_get(x_109, 1); -lean_inc(x_124); +lean_object* x_122; lean_object* x_123; lean_object* x_124; +x_122 = lean_ctor_get(x_108, 0); +x_123 = lean_ctor_get(x_108, 1); lean_inc(x_123); -lean_dec(x_109); -x_125 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_125, 0, x_123); -lean_ctor_set(x_125, 1, x_124); -return x_125; +lean_inc(x_122); +lean_dec(x_108); +x_124 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_124, 0, x_122); +lean_ctor_set(x_124, 1, x_123); +return x_124; } } } else { -uint8_t x_126; -lean_dec(x_51); +uint8_t x_125; lean_dec(x_50); +lean_dec(x_49); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -8406,23 +8344,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_126 = !lean_is_exclusive(x_103); -if (x_126 == 0) +x_125 = !lean_is_exclusive(x_102); +if (x_125 == 0) { -return x_103; +return x_102; } else { -lean_object* x_127; lean_object* x_128; lean_object* x_129; -x_127 = lean_ctor_get(x_103, 0); -x_128 = lean_ctor_get(x_103, 1); -lean_inc(x_128); +lean_object* x_126; lean_object* x_127; lean_object* x_128; +x_126 = lean_ctor_get(x_102, 0); +x_127 = lean_ctor_get(x_102, 1); lean_inc(x_127); -lean_dec(x_103); -x_129 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_129, 0, x_127); -lean_ctor_set(x_129, 1, x_128); -return x_129; +lean_inc(x_126); +lean_dec(x_102); +x_128 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_128, 0, x_126); +lean_ctor_set(x_128, 1, x_127); +return x_128; } } } @@ -8491,15 +8429,6 @@ lean_dec(x_2); return x_16; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_precheckMatch___lambda__3___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_Elab_Term_precheckMatch___lambda__3(x_1, x_2, x_3, x_4); -lean_dec(x_2); -return x_5; -} -} static lean_object* _init_l___regBuiltin_Lean_Elab_Term_precheckMatch___closed__1() { _start: { @@ -22548,213 +22477,1100 @@ return x_22; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_26 = lean_ctor_get(x_22, 0); -x_27 = lean_ctor_get(x_22, 1); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_22); -x_28 = lean_ctor_get(x_26, 2); -lean_inc(x_28); -lean_dec(x_26); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_27); -return x_29; +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_26 = lean_ctor_get(x_22, 0); +x_27 = lean_ctor_get(x_22, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_22); +x_28 = lean_ctor_get(x_26, 2); +lean_inc(x_28); +lean_dec(x_26); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_27); +return x_29; +} +} +else +{ +uint8_t x_30; +lean_dec(x_16); +x_30 = !lean_is_exclusive(x_18); +if (x_30 == 0) +{ +return x_18; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_18, 0); +x_32 = lean_ctor_get(x_18, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_18); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_topSort___spec__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, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +size_t x_13; size_t x_14; lean_object* x_15; +x_13 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_14 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_topSort___spec__1(x_1, x_13, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +return x_15; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_topSort___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___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_topSort(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__1___rarg___lambda__1(lean_object* x_1, uint8_t 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) { +_start: +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_box(x_2); +x_12 = lean_apply_9(x_1, x_5, x_11, x_3, x_4, x_6, x_7, x_8, x_9, x_10); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__1___rarg(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_box(x_5); +x_14 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__1___rarg___lambda__1___boxed), 10, 4); +lean_closure_set(x_14, 0, x_4); +lean_closure_set(x_14, 1, x_13); +lean_closure_set(x_14, 2, x_6); +lean_closure_set(x_14, 3, x_7); +x_15 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp___rarg(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_15) == 0) +{ +uint8_t x_16; +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +return x_15; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_15, 0); +x_18 = lean_ctor_get(x_15, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_15); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +return x_19; +} +} +else +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_15); +if (x_20 == 0) +{ +return x_15; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_15, 0); +x_22 = lean_ctor_get(x_15, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_15); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__1___rarg___boxed), 12, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_box(x_5); +x_14 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__1___rarg___lambda__1___boxed), 10, 4); +lean_closure_set(x_14, 0, x_4); +lean_closure_set(x_14, 1, x_13); +lean_closure_set(x_14, 2, x_6); +lean_closure_set(x_14, 3, x_7); +x_15 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLetDeclImp___rarg(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_15) == 0) +{ +uint8_t x_16; +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +return x_15; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_15, 0); +x_18 = lean_ctor_get(x_15, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_15); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +return x_19; +} +} +else +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_15); +if (x_20 == 0) +{ +return x_15; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_15, 0); +x_22 = lean_ctor_get(x_15, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_15); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__2(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withLetDecl___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__2___rarg___boxed), 12, 0); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4___rarg___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(32u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4___rarg___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4___rarg___closed__1; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4___rarg___closed__3() { +_start: +{ +size_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = 5; +x_2 = l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4___rarg___closed__2; +x_3 = l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4___rarg___closed__1; +x_4 = lean_unsigned_to_nat(0u); +x_5 = lean_alloc_ctor(0, 4, sizeof(size_t)*1); +lean_ctor_set(x_5, 0, x_2); +lean_ctor_set(x_5, 1, x_3); +lean_ctor_set(x_5, 2, x_4); +lean_ctor_set(x_5, 3, x_4); +lean_ctor_set_usize(x_5, 4, x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4___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; 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; lean_object* x_19; uint8_t x_20; +x_7 = lean_st_ref_get(x_5, x_6); +x_8 = lean_ctor_get(x_7, 1); +lean_inc(x_8); +lean_dec(x_7); +x_9 = lean_st_ref_get(x_1, x_8); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = lean_ctor_get(x_10, 4); +lean_inc(x_12); +lean_dec(x_10); +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +lean_dec(x_12); +x_14 = lean_st_ref_get(x_5, x_11); +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +lean_dec(x_14); +x_16 = lean_st_ref_take(x_1, x_15); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_17, 4); +lean_inc(x_18); +x_19 = lean_ctor_get(x_16, 1); +lean_inc(x_19); +lean_dec(x_16); +x_20 = !lean_is_exclusive(x_17); +if (x_20 == 0) +{ +lean_object* x_21; uint8_t x_22; +x_21 = lean_ctor_get(x_17, 4); +lean_dec(x_21); +x_22 = !lean_is_exclusive(x_18); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_23 = lean_ctor_get(x_18, 1); +lean_dec(x_23); +x_24 = l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4___rarg___closed__3; +lean_ctor_set(x_18, 1, x_24); +x_25 = lean_st_ref_set(x_1, x_17, x_19); +x_26 = !lean_is_exclusive(x_25); +if (x_26 == 0) +{ +lean_object* x_27; +x_27 = lean_ctor_get(x_25, 0); +lean_dec(x_27); +lean_ctor_set(x_25, 0, x_13); +return x_25; +} +else +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_25, 1); +lean_inc(x_28); +lean_dec(x_25); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_13); +lean_ctor_set(x_29, 1, x_28); +return x_29; +} +} +else +{ +uint8_t 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; +x_30 = lean_ctor_get_uint8(x_18, sizeof(void*)*2); +x_31 = lean_ctor_get(x_18, 0); +lean_inc(x_31); +lean_dec(x_18); +x_32 = l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4___rarg___closed__3; +x_33 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +lean_ctor_set_uint8(x_33, sizeof(void*)*2, x_30); +lean_ctor_set(x_17, 4, x_33); +x_34 = lean_st_ref_set(x_1, x_17, x_19); +x_35 = lean_ctor_get(x_34, 1); +lean_inc(x_35); +if (lean_is_exclusive(x_34)) { + lean_ctor_release(x_34, 0); + lean_ctor_release(x_34, 1); + x_36 = x_34; +} else { + lean_dec_ref(x_34); + x_36 = lean_box(0); +} +if (lean_is_scalar(x_36)) { + x_37 = lean_alloc_ctor(0, 2, 0); +} else { + x_37 = x_36; +} +lean_ctor_set(x_37, 0, x_13); +lean_ctor_set(x_37, 1, x_35); +return x_37; +} +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_38 = lean_ctor_get(x_17, 0); +x_39 = lean_ctor_get(x_17, 1); +x_40 = lean_ctor_get(x_17, 2); +x_41 = lean_ctor_get(x_17, 3); +lean_inc(x_41); +lean_inc(x_40); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_17); +x_42 = lean_ctor_get_uint8(x_18, sizeof(void*)*2); +x_43 = lean_ctor_get(x_18, 0); +lean_inc(x_43); +if (lean_is_exclusive(x_18)) { + lean_ctor_release(x_18, 0); + lean_ctor_release(x_18, 1); + x_44 = x_18; +} else { + lean_dec_ref(x_18); + x_44 = lean_box(0); +} +x_45 = l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4___rarg___closed__3; +if (lean_is_scalar(x_44)) { + x_46 = lean_alloc_ctor(0, 2, 1); +} else { + x_46 = x_44; +} +lean_ctor_set(x_46, 0, x_43); +lean_ctor_set(x_46, 1, x_45); +lean_ctor_set_uint8(x_46, sizeof(void*)*2, x_42); +x_47 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_47, 0, x_38); +lean_ctor_set(x_47, 1, x_39); +lean_ctor_set(x_47, 2, x_40); +lean_ctor_set(x_47, 3, x_41); +lean_ctor_set(x_47, 4, x_46); +x_48 = lean_st_ref_set(x_1, x_47, x_19); +x_49 = lean_ctor_get(x_48, 1); +lean_inc(x_49); +if (lean_is_exclusive(x_48)) { + lean_ctor_release(x_48, 0); + lean_ctor_release(x_48, 1); + x_50 = x_48; +} else { + lean_dec_ref(x_48); + x_50 = lean_box(0); +} +if (lean_is_scalar(x_50)) { + x_51 = lean_alloc_ctor(0, 2, 0); +} else { + x_51 = x_50; +} +lean_ctor_set(x_51, 0, x_13); +lean_ctor_set(x_51, 1, x_49); +return x_51; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4(uint8_t x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4___rarg___boxed), 6, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__3(lean_object* x_1, lean_object* x_2, uint8_t 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) { +_start: +{ +lean_object* x_11; 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_st_ref_get(x_9, x_10); +x_25 = lean_ctor_get(x_24, 1); +lean_inc(x_25); +lean_dec(x_24); +x_26 = lean_st_ref_get(x_5, x_25); +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_27, 4); +lean_inc(x_28); +lean_dec(x_27); +x_29 = lean_ctor_get_uint8(x_28, sizeof(void*)*2); +lean_dec(x_28); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +lean_dec(x_2); +x_30 = lean_ctor_get(x_26, 1); +lean_inc(x_30); +lean_dec(x_26); +x_31 = lean_box(x_3); +x_32 = lean_apply_8(x_1, x_31, x_4, x_5, x_6, x_7, x_8, x_9, x_30); +return x_32; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_33 = lean_ctor_get(x_26, 1); +lean_inc(x_33); +lean_dec(x_26); +x_34 = l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4___rarg(x_5, x_6, x_7, x_8, x_9, x_33); +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +lean_dec(x_34); +x_37 = lean_box(x_3); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_38 = lean_apply_8(x_1, x_37, x_4, x_5, x_6, x_7, x_8, x_9, x_36); +if (lean_obj_tag(x_38) == 0) +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_38, 1); +lean_inc(x_40); +lean_dec(x_38); +x_41 = lean_box(x_3); +lean_inc(x_9); +lean_inc(x_5); +lean_inc(x_39); +x_42 = lean_apply_9(x_2, x_39, x_41, x_4, x_5, x_6, x_7, x_8, x_9, x_40); +if (lean_obj_tag(x_42) == 0) +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_43 = lean_ctor_get(x_42, 0); +lean_inc(x_43); +x_44 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +lean_dec(x_42); +x_45 = lean_st_ref_get(x_9, x_44); +lean_dec(x_9); +x_46 = lean_ctor_get(x_45, 1); +lean_inc(x_46); +lean_dec(x_45); +x_47 = lean_st_ref_take(x_5, x_46); +x_48 = lean_ctor_get(x_47, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_48, 4); +lean_inc(x_49); +if (lean_obj_tag(x_43) == 0) +{ +lean_object* x_50; uint8_t x_51; +x_50 = lean_ctor_get(x_47, 1); +lean_inc(x_50); +lean_dec(x_47); +x_51 = !lean_is_exclusive(x_48); +if (x_51 == 0) +{ +lean_object* x_52; uint8_t x_53; +x_52 = lean_ctor_get(x_48, 4); +lean_dec(x_52); +x_53 = !lean_is_exclusive(x_49); +if (x_53 == 0) +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; +x_54 = lean_ctor_get(x_49, 1); +x_55 = lean_ctor_get(x_43, 0); +lean_inc(x_55); +lean_dec(x_43); +x_56 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_54); +x_57 = l_Std_PersistentArray_push___rarg(x_35, x_56); +lean_ctor_set(x_49, 1, x_57); +x_58 = lean_st_ref_set(x_5, x_48, x_50); +lean_dec(x_5); +x_59 = !lean_is_exclusive(x_58); +if (x_59 == 0) +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_58, 0); +lean_dec(x_60); +x_61 = lean_box(0); +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_39); +lean_ctor_set(x_62, 1, x_61); +lean_ctor_set(x_58, 0, x_62); +x_11 = x_58; +goto block_23; +} +else +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_63 = lean_ctor_get(x_58, 1); +lean_inc(x_63); +lean_dec(x_58); +x_64 = lean_box(0); +x_65 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_65, 0, x_39); +lean_ctor_set(x_65, 1, x_64); +x_66 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_66, 0, x_65); +lean_ctor_set(x_66, 1, x_63); +x_11 = x_66; +goto block_23; +} +} +else +{ +uint8_t 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; lean_object* x_79; +x_67 = lean_ctor_get_uint8(x_49, sizeof(void*)*2); +x_68 = lean_ctor_get(x_49, 0); +x_69 = lean_ctor_get(x_49, 1); +lean_inc(x_69); +lean_inc(x_68); +lean_dec(x_49); +x_70 = lean_ctor_get(x_43, 0); +lean_inc(x_70); +lean_dec(x_43); +x_71 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set(x_71, 1, x_69); +x_72 = l_Std_PersistentArray_push___rarg(x_35, x_71); +x_73 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_73, 0, x_68); +lean_ctor_set(x_73, 1, x_72); +lean_ctor_set_uint8(x_73, sizeof(void*)*2, x_67); +lean_ctor_set(x_48, 4, x_73); +x_74 = lean_st_ref_set(x_5, x_48, x_50); +lean_dec(x_5); +x_75 = lean_ctor_get(x_74, 1); +lean_inc(x_75); +if (lean_is_exclusive(x_74)) { + lean_ctor_release(x_74, 0); + lean_ctor_release(x_74, 1); + x_76 = x_74; +} else { + lean_dec_ref(x_74); + x_76 = lean_box(0); +} +x_77 = lean_box(0); +x_78 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_78, 0, x_39); +lean_ctor_set(x_78, 1, x_77); +if (lean_is_scalar(x_76)) { + x_79 = lean_alloc_ctor(0, 2, 0); +} else { + x_79 = x_76; +} +lean_ctor_set(x_79, 0, x_78); +lean_ctor_set(x_79, 1, x_75); +x_11 = x_79; +goto block_23; +} +} +else +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; uint8_t x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_80 = lean_ctor_get(x_48, 0); +x_81 = lean_ctor_get(x_48, 1); +x_82 = lean_ctor_get(x_48, 2); +x_83 = lean_ctor_get(x_48, 3); +lean_inc(x_83); +lean_inc(x_82); +lean_inc(x_81); +lean_inc(x_80); +lean_dec(x_48); +x_84 = lean_ctor_get_uint8(x_49, sizeof(void*)*2); +x_85 = lean_ctor_get(x_49, 0); +lean_inc(x_85); +x_86 = lean_ctor_get(x_49, 1); +lean_inc(x_86); +if (lean_is_exclusive(x_49)) { + lean_ctor_release(x_49, 0); + lean_ctor_release(x_49, 1); + x_87 = x_49; +} else { + lean_dec_ref(x_49); + x_87 = lean_box(0); +} +x_88 = lean_ctor_get(x_43, 0); +lean_inc(x_88); +lean_dec(x_43); +x_89 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_89, 0, x_88); +lean_ctor_set(x_89, 1, x_86); +x_90 = l_Std_PersistentArray_push___rarg(x_35, x_89); +if (lean_is_scalar(x_87)) { + x_91 = lean_alloc_ctor(0, 2, 1); +} else { + x_91 = x_87; +} +lean_ctor_set(x_91, 0, x_85); +lean_ctor_set(x_91, 1, x_90); +lean_ctor_set_uint8(x_91, sizeof(void*)*2, x_84); +x_92 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_92, 0, x_80); +lean_ctor_set(x_92, 1, x_81); +lean_ctor_set(x_92, 2, x_82); +lean_ctor_set(x_92, 3, x_83); +lean_ctor_set(x_92, 4, x_91); +x_93 = lean_st_ref_set(x_5, x_92, x_50); +lean_dec(x_5); +x_94 = lean_ctor_get(x_93, 1); +lean_inc(x_94); +if (lean_is_exclusive(x_93)) { + lean_ctor_release(x_93, 0); + lean_ctor_release(x_93, 1); + x_95 = x_93; +} else { + lean_dec_ref(x_93); + x_95 = lean_box(0); +} +x_96 = lean_box(0); +x_97 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_97, 0, x_39); +lean_ctor_set(x_97, 1, x_96); +if (lean_is_scalar(x_95)) { + x_98 = lean_alloc_ctor(0, 2, 0); +} else { + x_98 = x_95; +} +lean_ctor_set(x_98, 0, x_97); +lean_ctor_set(x_98, 1, x_94); +x_11 = x_98; +goto block_23; +} +} +else +{ +lean_object* x_99; uint8_t x_100; +x_99 = lean_ctor_get(x_47, 1); +lean_inc(x_99); +lean_dec(x_47); +x_100 = !lean_is_exclusive(x_48); +if (x_100 == 0) +{ +lean_object* x_101; uint8_t x_102; +x_101 = lean_ctor_get(x_48, 4); +lean_dec(x_101); +x_102 = !lean_is_exclusive(x_49); +if (x_102 == 0) +{ +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; uint8_t x_108; +x_103 = lean_ctor_get(x_49, 1); +lean_dec(x_103); +x_104 = lean_ctor_get(x_43, 0); +lean_inc(x_104); +lean_dec(x_43); +x_105 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_105, 0, x_104); +x_106 = l_Std_PersistentArray_push___rarg(x_35, x_105); +lean_ctor_set(x_49, 1, x_106); +x_107 = lean_st_ref_set(x_5, x_48, x_99); +lean_dec(x_5); +x_108 = !lean_is_exclusive(x_107); +if (x_108 == 0) +{ +lean_object* x_109; lean_object* x_110; lean_object* x_111; +x_109 = lean_ctor_get(x_107, 0); +lean_dec(x_109); +x_110 = lean_box(0); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_39); +lean_ctor_set(x_111, 1, x_110); +lean_ctor_set(x_107, 0, x_111); +x_11 = x_107; +goto block_23; +} +else +{ +lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; +x_112 = lean_ctor_get(x_107, 1); +lean_inc(x_112); +lean_dec(x_107); +x_113 = lean_box(0); +x_114 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_114, 0, x_39); +lean_ctor_set(x_114, 1, x_113); +x_115 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_115, 0, x_114); +lean_ctor_set(x_115, 1, x_112); +x_11 = x_115; +goto block_23; } } else { -uint8_t x_30; -lean_dec(x_16); -x_30 = !lean_is_exclusive(x_18); -if (x_30 == 0) -{ -return x_18; +uint8_t x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_116 = lean_ctor_get_uint8(x_49, sizeof(void*)*2); +x_117 = lean_ctor_get(x_49, 0); +lean_inc(x_117); +lean_dec(x_49); +x_118 = lean_ctor_get(x_43, 0); +lean_inc(x_118); +lean_dec(x_43); +x_119 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_119, 0, x_118); +x_120 = l_Std_PersistentArray_push___rarg(x_35, x_119); +x_121 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_121, 0, x_117); +lean_ctor_set(x_121, 1, x_120); +lean_ctor_set_uint8(x_121, sizeof(void*)*2, x_116); +lean_ctor_set(x_48, 4, x_121); +x_122 = lean_st_ref_set(x_5, x_48, x_99); +lean_dec(x_5); +x_123 = lean_ctor_get(x_122, 1); +lean_inc(x_123); +if (lean_is_exclusive(x_122)) { + lean_ctor_release(x_122, 0); + lean_ctor_release(x_122, 1); + x_124 = x_122; +} else { + lean_dec_ref(x_122); + x_124 = lean_box(0); +} +x_125 = lean_box(0); +x_126 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_126, 0, x_39); +lean_ctor_set(x_126, 1, x_125); +if (lean_is_scalar(x_124)) { + x_127 = lean_alloc_ctor(0, 2, 0); +} else { + x_127 = x_124; +} +lean_ctor_set(x_127, 0, x_126); +lean_ctor_set(x_127, 1, x_123); +x_11 = x_127; +goto block_23; +} } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_18, 0); -x_32 = lean_ctor_get(x_18, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_18); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -return x_33; +lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; uint8_t x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; +x_128 = lean_ctor_get(x_48, 0); +x_129 = lean_ctor_get(x_48, 1); +x_130 = lean_ctor_get(x_48, 2); +x_131 = lean_ctor_get(x_48, 3); +lean_inc(x_131); +lean_inc(x_130); +lean_inc(x_129); +lean_inc(x_128); +lean_dec(x_48); +x_132 = lean_ctor_get_uint8(x_49, sizeof(void*)*2); +x_133 = lean_ctor_get(x_49, 0); +lean_inc(x_133); +if (lean_is_exclusive(x_49)) { + lean_ctor_release(x_49, 0); + lean_ctor_release(x_49, 1); + x_134 = x_49; +} else { + lean_dec_ref(x_49); + x_134 = lean_box(0); +} +x_135 = lean_ctor_get(x_43, 0); +lean_inc(x_135); +lean_dec(x_43); +x_136 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_136, 0, x_135); +x_137 = l_Std_PersistentArray_push___rarg(x_35, x_136); +if (lean_is_scalar(x_134)) { + x_138 = lean_alloc_ctor(0, 2, 1); +} else { + x_138 = x_134; +} +lean_ctor_set(x_138, 0, x_133); +lean_ctor_set(x_138, 1, x_137); +lean_ctor_set_uint8(x_138, sizeof(void*)*2, x_132); +x_139 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_139, 0, x_128); +lean_ctor_set(x_139, 1, x_129); +lean_ctor_set(x_139, 2, x_130); +lean_ctor_set(x_139, 3, x_131); +lean_ctor_set(x_139, 4, x_138); +x_140 = lean_st_ref_set(x_5, x_139, x_99); +lean_dec(x_5); +x_141 = lean_ctor_get(x_140, 1); +lean_inc(x_141); +if (lean_is_exclusive(x_140)) { + lean_ctor_release(x_140, 0); + lean_ctor_release(x_140, 1); + x_142 = x_140; +} else { + lean_dec_ref(x_140); + x_142 = lean_box(0); } +x_143 = lean_box(0); +x_144 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_144, 0, x_39); +lean_ctor_set(x_144, 1, x_143); +if (lean_is_scalar(x_142)) { + x_145 = lean_alloc_ctor(0, 2, 0); +} else { + x_145 = x_142; } +lean_ctor_set(x_145, 0, x_144); +lean_ctor_set(x_145, 1, x_141); +x_11 = x_145; +goto block_23; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_topSort___spec__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, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: +} +else { -size_t x_13; size_t x_14; lean_object* x_15; -x_13 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_14 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_topSort___spec__1(x_1, x_13, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_7); -lean_dec(x_6); +uint8_t x_146; +lean_dec(x_39); +lean_dec(x_35); +lean_dec(x_9); lean_dec(x_5); -lean_dec(x_1); -return x_15; +x_146 = !lean_is_exclusive(x_42); +if (x_146 == 0) +{ +x_11 = x_42; +goto block_23; } +else +{ +lean_object* x_147; lean_object* x_148; lean_object* x_149; +x_147 = lean_ctor_get(x_42, 0); +x_148 = lean_ctor_get(x_42, 1); +lean_inc(x_148); +lean_inc(x_147); +lean_dec(x_42); +x_149 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_149, 0, x_147); +lean_ctor_set(x_149, 1, x_148); +x_11 = x_149; +goto block_23; } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_topSort___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: +} +} +else { -lean_object* x_9; -x_9 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_topSort(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; uint8_t x_158; +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_3); +lean_dec(x_4); lean_dec(x_2); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__1___rarg___lambda__1(lean_object* x_1, uint8_t 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) { -_start: -{ -lean_object* x_11; lean_object* x_12; -x_11 = lean_box(x_2); -x_12 = lean_apply_9(x_1, x_5, x_11, x_3, x_4, x_6, x_7, x_8, x_9, x_10); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__1___rarg(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: +x_150 = lean_ctor_get(x_38, 0); +lean_inc(x_150); +x_151 = lean_ctor_get(x_38, 1); +lean_inc(x_151); +lean_dec(x_38); +x_152 = lean_st_ref_get(x_9, x_151); +lean_dec(x_9); +x_153 = lean_ctor_get(x_152, 1); +lean_inc(x_153); +lean_dec(x_152); +x_154 = lean_st_ref_take(x_5, x_153); +x_155 = lean_ctor_get(x_154, 0); +lean_inc(x_155); +x_156 = lean_ctor_get(x_155, 4); +lean_inc(x_156); +x_157 = lean_ctor_get(x_154, 1); +lean_inc(x_157); +lean_dec(x_154); +x_158 = !lean_is_exclusive(x_155); +if (x_158 == 0) { -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = lean_box(x_5); -x_14 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__1___rarg___lambda__1___boxed), 10, 4); -lean_closure_set(x_14, 0, x_4); -lean_closure_set(x_14, 1, x_13); -lean_closure_set(x_14, 2, x_6); -lean_closure_set(x_14, 3, x_7); -x_15 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp___rarg(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_15) == 0) +lean_object* x_159; uint8_t x_160; +x_159 = lean_ctor_get(x_155, 4); +lean_dec(x_159); +x_160 = !lean_is_exclusive(x_156); +if (x_160 == 0) { -uint8_t x_16; -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) +lean_object* x_161; lean_object* x_162; uint8_t x_163; +x_161 = lean_ctor_get(x_156, 1); +lean_dec(x_161); +lean_ctor_set(x_156, 1, x_35); +x_162 = lean_st_ref_set(x_5, x_155, x_157); +lean_dec(x_5); +x_163 = !lean_is_exclusive(x_162); +if (x_163 == 0) { -return x_15; +lean_object* x_164; +x_164 = lean_ctor_get(x_162, 0); +lean_dec(x_164); +lean_ctor_set_tag(x_162, 1); +lean_ctor_set(x_162, 0, x_150); +x_11 = x_162; +goto block_23; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_15, 0); -x_18 = lean_ctor_get(x_15, 1); -lean_inc(x_18); -lean_inc(x_17); -lean_dec(x_15); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set(x_19, 1, x_18); -return x_19; +lean_object* x_165; lean_object* x_166; +x_165 = lean_ctor_get(x_162, 1); +lean_inc(x_165); +lean_dec(x_162); +x_166 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_166, 0, x_150); +lean_ctor_set(x_166, 1, x_165); +x_11 = x_166; +goto block_23; } } else { -uint8_t x_20; -x_20 = !lean_is_exclusive(x_15); -if (x_20 == 0) -{ -return x_15; +uint8_t x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; +x_167 = lean_ctor_get_uint8(x_156, sizeof(void*)*2); +x_168 = lean_ctor_get(x_156, 0); +lean_inc(x_168); +lean_dec(x_156); +x_169 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_169, 0, x_168); +lean_ctor_set(x_169, 1, x_35); +lean_ctor_set_uint8(x_169, sizeof(void*)*2, x_167); +lean_ctor_set(x_155, 4, x_169); +x_170 = lean_st_ref_set(x_5, x_155, x_157); +lean_dec(x_5); +x_171 = lean_ctor_get(x_170, 1); +lean_inc(x_171); +if (lean_is_exclusive(x_170)) { + lean_ctor_release(x_170, 0); + lean_ctor_release(x_170, 1); + x_172 = x_170; +} else { + lean_dec_ref(x_170); + x_172 = lean_box(0); +} +if (lean_is_scalar(x_172)) { + x_173 = lean_alloc_ctor(1, 2, 0); +} else { + x_173 = x_172; + lean_ctor_set_tag(x_173, 1); +} +lean_ctor_set(x_173, 0, x_150); +lean_ctor_set(x_173, 1, x_171); +x_11 = x_173; +goto block_23; +} } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_15, 0); -x_22 = lean_ctor_get(x_15, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_15); -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); -return x_23; +lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; uint8_t x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; +x_174 = lean_ctor_get(x_155, 0); +x_175 = lean_ctor_get(x_155, 1); +x_176 = lean_ctor_get(x_155, 2); +x_177 = lean_ctor_get(x_155, 3); +lean_inc(x_177); +lean_inc(x_176); +lean_inc(x_175); +lean_inc(x_174); +lean_dec(x_155); +x_178 = lean_ctor_get_uint8(x_156, sizeof(void*)*2); +x_179 = lean_ctor_get(x_156, 0); +lean_inc(x_179); +if (lean_is_exclusive(x_156)) { + lean_ctor_release(x_156, 0); + lean_ctor_release(x_156, 1); + x_180 = x_156; +} else { + lean_dec_ref(x_156); + x_180 = lean_box(0); } +if (lean_is_scalar(x_180)) { + x_181 = lean_alloc_ctor(0, 2, 1); +} else { + x_181 = x_180; } +lean_ctor_set(x_181, 0, x_179); +lean_ctor_set(x_181, 1, x_35); +lean_ctor_set_uint8(x_181, sizeof(void*)*2, x_178); +x_182 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_182, 0, x_174); +lean_ctor_set(x_182, 1, x_175); +lean_ctor_set(x_182, 2, x_176); +lean_ctor_set(x_182, 3, x_177); +lean_ctor_set(x_182, 4, x_181); +x_183 = lean_st_ref_set(x_5, x_182, x_157); +lean_dec(x_5); +x_184 = lean_ctor_get(x_183, 1); +lean_inc(x_184); +if (lean_is_exclusive(x_183)) { + lean_ctor_release(x_183, 0); + lean_ctor_release(x_183, 1); + x_185 = x_183; +} else { + lean_dec_ref(x_183); + x_185 = lean_box(0); } +if (lean_is_scalar(x_185)) { + x_186 = lean_alloc_ctor(1, 2, 0); +} else { + x_186 = x_185; + lean_ctor_set_tag(x_186, 1); } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__1(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__1___rarg___boxed), 12, 0); -return x_2; +lean_ctor_set(x_186, 0, x_150); +lean_ctor_set(x_186, 1, x_184); +x_11 = x_186; +goto block_23; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: +} +block_23: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = lean_box(x_5); -x_14 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__1___rarg___lambda__1___boxed), 10, 4); -lean_closure_set(x_14, 0, x_4); -lean_closure_set(x_14, 1, x_13); -lean_closure_set(x_14, 2, x_6); -lean_closure_set(x_14, 3, x_7); -x_15 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLetDeclImp___rarg(x_1, x_2, x_3, x_14, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_15) == 0) +if (lean_obj_tag(x_11) == 0) { -uint8_t x_16; -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) +uint8_t x_12; +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) { -return x_15; +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +lean_dec(x_13); +lean_ctor_set(x_11, 0, x_14); +return x_11; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_15 = lean_ctor_get(x_11, 0); +x_16 = lean_ctor_get(x_11, 1); +lean_inc(x_16); +lean_inc(x_15); +lean_dec(x_11); x_17 = lean_ctor_get(x_15, 0); -x_18 = lean_ctor_get(x_15, 1); -lean_inc(x_18); lean_inc(x_17); lean_dec(x_15); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set(x_19, 1, x_18); -return x_19; +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_16); +return x_18; } } else { -uint8_t x_20; -x_20 = !lean_is_exclusive(x_15); -if (x_20 == 0) +uint8_t x_19; +x_19 = !lean_is_exclusive(x_11); +if (x_19 == 0) { -return x_15; +return x_11; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_15, 0); -x_22 = lean_ctor_get(x_15, 1); -lean_inc(x_22); +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_11, 0); +x_21 = lean_ctor_get(x_11, 1); lean_inc(x_21); -lean_dec(x_15); -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); -return x_23; -} +lean_inc(x_20); +lean_dec(x_11); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__2(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withLetDecl___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__2___rarg___boxed), 12, 0); -return x_2; } } LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___lambda__1(lean_object* x_1, lean_object* x_2, uint8_t 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) { @@ -22875,6 +23691,41 @@ return x_24; } } } +LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___lambda__3(lean_object* x_1, lean_object* x_2, uint8_t 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) { +_start: +{ +lean_object* x_11; uint8_t x_12; +x_11 = lean_box(0); +x_12 = l_Lean_Expr_isFVar(x_2); +if (x_12 == 0) +{ +lean_object* x_13; uint8_t x_14; lean_object* x_15; +x_13 = lean_box(0); +x_14 = 0; +x_15 = l_Lean_Elab_Term_mkTermInfo(x_13, x_1, x_2, x_11, x_11, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_15; +} +else +{ +if (x_3 == 0) +{ +lean_object* x_16; uint8_t x_17; lean_object* x_18; +x_16 = lean_box(0); +x_17 = 1; +x_18 = l_Lean_Elab_Term_mkTermInfo(x_16, x_1, x_2, x_11, x_11, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_18; +} +else +{ +lean_object* x_19; uint8_t x_20; lean_object* x_21; +x_19 = lean_box(0); +x_20 = 0; +x_21 = l_Lean_Elab_Term_mkTermInfo(x_19, x_1, x_2, x_11, x_11, x_20, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_21; +} +} +} +} LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go(lean_object* x_1, uint8_t 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) { _start: { @@ -23289,7 +24140,7 @@ return x_92; } else { -lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; +lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_dec(x_78); lean_dec(x_77); x_93 = lean_ctor_get(x_80, 0); @@ -23300,294 +24151,224 @@ lean_inc(x_94); x_95 = lean_ctor_get(x_93, 1); lean_inc(x_95); lean_dec(x_93); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_96 = l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go(x_95, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_96) == 0) -{ -lean_object* x_97; lean_object* x_98; uint8_t x_99; -x_97 = lean_ctor_get(x_96, 0); -lean_inc(x_97); -x_98 = lean_ctor_get(x_96, 1); -lean_inc(x_98); -lean_dec(x_96); -x_99 = l_Lean_Expr_isFVar(x_97); -if (x_99 == 0) -{ -lean_object* x_100; lean_object* x_101; uint8_t x_102; lean_object* x_103; -x_100 = lean_box(0); -x_101 = lean_box(0); -x_102 = 0; -x_103 = l_Lean_Elab_Term_addTermInfo(x_94, x_97, x_100, x_100, x_101, x_102, x_3, x_4, x_5, x_6, x_7, x_8, x_98); -return x_103; -} -else -{ -if (x_2 == 0) -{ -lean_object* x_104; lean_object* x_105; uint8_t x_106; lean_object* x_107; -x_104 = lean_box(0); -x_105 = lean_box(0); -x_106 = 1; -x_107 = l_Lean_Elab_Term_addTermInfo(x_94, x_97, x_104, x_104, x_105, x_106, x_3, x_4, x_5, x_6, x_7, x_8, x_98); -return x_107; -} -else -{ -lean_object* x_108; lean_object* x_109; uint8_t x_110; lean_object* x_111; -x_108 = lean_box(0); -x_109 = lean_box(0); -x_110 = 0; -x_111 = l_Lean_Elab_Term_addTermInfo(x_94, x_97, x_108, x_108, x_109, x_110, x_3, x_4, x_5, x_6, x_7, x_8, x_98); -return x_111; -} -} -} -else -{ -uint8_t x_112; -lean_dec(x_94); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_112 = !lean_is_exclusive(x_96); -if (x_112 == 0) -{ -return x_96; -} -else -{ -lean_object* x_113; lean_object* x_114; lean_object* x_115; -x_113 = lean_ctor_get(x_96, 0); -x_114 = lean_ctor_get(x_96, 1); -lean_inc(x_114); -lean_inc(x_113); -lean_dec(x_96); -x_115 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_115, 0, x_113); -lean_ctor_set(x_115, 1, x_114); -return x_115; -} -} +x_96 = lean_alloc_closure((void*)(l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___boxed), 9, 1); +lean_closure_set(x_96, 0, x_95); +x_97 = lean_alloc_closure((void*)(l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___lambda__3___boxed), 10, 1); +lean_closure_set(x_97, 0, x_94); +x_98 = l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__3(x_96, x_97, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_98; } } else { -uint8_t x_116; lean_object* x_117; +uint8_t x_99; lean_object* x_100; lean_dec(x_79); lean_dec(x_1); -x_116 = 0; -x_117 = l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go(x_78, x_116, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_117) == 0) +x_99 = 0; +x_100 = l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go(x_78, x_99, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_100) == 0) { -uint8_t x_118; -x_118 = !lean_is_exclusive(x_117); -if (x_118 == 0) +uint8_t x_101; +x_101 = !lean_is_exclusive(x_100); +if (x_101 == 0) { -lean_object* x_119; lean_object* x_120; -x_119 = lean_ctor_get(x_117, 0); -x_120 = l_Lean_mkMData(x_77, x_119); -lean_ctor_set(x_117, 0, x_120); -return x_117; +lean_object* x_102; lean_object* x_103; +x_102 = lean_ctor_get(x_100, 0); +x_103 = l_Lean_mkMData(x_77, x_102); +lean_ctor_set(x_100, 0, x_103); +return x_100; } else { -lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; -x_121 = lean_ctor_get(x_117, 0); -x_122 = lean_ctor_get(x_117, 1); -lean_inc(x_122); -lean_inc(x_121); -lean_dec(x_117); -x_123 = l_Lean_mkMData(x_77, x_121); -x_124 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_124, 0, x_123); -lean_ctor_set(x_124, 1, x_122); -return x_124; +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_104 = lean_ctor_get(x_100, 0); +x_105 = lean_ctor_get(x_100, 1); +lean_inc(x_105); +lean_inc(x_104); +lean_dec(x_100); +x_106 = l_Lean_mkMData(x_77, x_104); +x_107 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_107, 0, x_106); +lean_ctor_set(x_107, 1, x_105); +return x_107; } } else { -uint8_t x_125; +uint8_t x_108; lean_dec(x_77); -x_125 = !lean_is_exclusive(x_117); -if (x_125 == 0) +x_108 = !lean_is_exclusive(x_100); +if (x_108 == 0) { -return x_117; +return x_100; } else { -lean_object* x_126; lean_object* x_127; lean_object* x_128; -x_126 = lean_ctor_get(x_117, 0); -x_127 = lean_ctor_get(x_117, 1); -lean_inc(x_127); -lean_inc(x_126); -lean_dec(x_117); -x_128 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_128, 0, x_126); -lean_ctor_set(x_128, 1, x_127); -return x_128; +lean_object* x_109; lean_object* x_110; lean_object* x_111; +x_109 = lean_ctor_get(x_100, 0); +x_110 = lean_ctor_get(x_100, 1); +lean_inc(x_110); +lean_inc(x_109); +lean_dec(x_100); +x_111 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_111, 0, x_109); +lean_ctor_set(x_111, 1, x_110); +return x_111; } } } } case 11: { -uint8_t x_129; -x_129 = !lean_is_exclusive(x_1); -if (x_129 == 0) +uint8_t x_112; +x_112 = !lean_is_exclusive(x_1); +if (x_112 == 0) { -lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; -x_130 = lean_ctor_get(x_1, 0); -x_131 = lean_ctor_get(x_1, 1); -x_132 = lean_ctor_get(x_1, 2); -lean_inc(x_132); -x_133 = l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go(x_132, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_133) == 0) +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_113 = lean_ctor_get(x_1, 0); +x_114 = lean_ctor_get(x_1, 1); +x_115 = lean_ctor_get(x_1, 2); +lean_inc(x_115); +x_116 = l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go(x_115, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_116) == 0) { -uint8_t x_134; -x_134 = !lean_is_exclusive(x_133); -if (x_134 == 0) +uint8_t x_117; +x_117 = !lean_is_exclusive(x_116); +if (x_117 == 0) { -lean_object* x_135; lean_object* x_136; -x_135 = lean_ctor_get(x_133, 0); -x_136 = lean_expr_update_proj(x_1, x_135); -lean_ctor_set(x_133, 0, x_136); -return x_133; +lean_object* x_118; lean_object* x_119; +x_118 = lean_ctor_get(x_116, 0); +x_119 = lean_expr_update_proj(x_1, x_118); +lean_ctor_set(x_116, 0, x_119); +return x_116; } else { -lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; -x_137 = lean_ctor_get(x_133, 0); -x_138 = lean_ctor_get(x_133, 1); -lean_inc(x_138); -lean_inc(x_137); -lean_dec(x_133); -x_139 = lean_expr_update_proj(x_1, x_137); -x_140 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_140, 0, x_139); -lean_ctor_set(x_140, 1, x_138); -return x_140; +lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_120 = lean_ctor_get(x_116, 0); +x_121 = lean_ctor_get(x_116, 1); +lean_inc(x_121); +lean_inc(x_120); +lean_dec(x_116); +x_122 = lean_expr_update_proj(x_1, x_120); +x_123 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_123, 0, x_122); +lean_ctor_set(x_123, 1, x_121); +return x_123; } } else { -uint8_t x_141; +uint8_t x_124; lean_free_object(x_1); -lean_dec(x_132); -lean_dec(x_131); -lean_dec(x_130); -x_141 = !lean_is_exclusive(x_133); -if (x_141 == 0) +lean_dec(x_115); +lean_dec(x_114); +lean_dec(x_113); +x_124 = !lean_is_exclusive(x_116); +if (x_124 == 0) { -return x_133; +return x_116; } else { -lean_object* x_142; lean_object* x_143; lean_object* x_144; -x_142 = lean_ctor_get(x_133, 0); -x_143 = lean_ctor_get(x_133, 1); -lean_inc(x_143); -lean_inc(x_142); -lean_dec(x_133); -x_144 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_144, 0, x_142); -lean_ctor_set(x_144, 1, x_143); -return x_144; +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_116, 0); +x_126 = lean_ctor_get(x_116, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_116); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; } } } else { -lean_object* x_145; lean_object* x_146; lean_object* x_147; uint64_t x_148; lean_object* x_149; -x_145 = lean_ctor_get(x_1, 0); -x_146 = lean_ctor_get(x_1, 1); -x_147 = lean_ctor_get(x_1, 2); -x_148 = lean_ctor_get_uint64(x_1, sizeof(void*)*3); -lean_inc(x_147); -lean_inc(x_146); -lean_inc(x_145); +lean_object* x_128; lean_object* x_129; lean_object* x_130; uint64_t x_131; lean_object* x_132; +x_128 = lean_ctor_get(x_1, 0); +x_129 = lean_ctor_get(x_1, 1); +x_130 = lean_ctor_get(x_1, 2); +x_131 = lean_ctor_get_uint64(x_1, sizeof(void*)*3); +lean_inc(x_130); +lean_inc(x_129); +lean_inc(x_128); lean_dec(x_1); -lean_inc(x_147); -x_149 = l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go(x_147, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_149) == 0) +lean_inc(x_130); +x_132 = l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go(x_130, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_132) == 0) { -lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; -x_150 = lean_ctor_get(x_149, 0); -lean_inc(x_150); -x_151 = lean_ctor_get(x_149, 1); -lean_inc(x_151); -if (lean_is_exclusive(x_149)) { - lean_ctor_release(x_149, 0); - lean_ctor_release(x_149, 1); - x_152 = x_149; +lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; +x_133 = lean_ctor_get(x_132, 0); +lean_inc(x_133); +x_134 = lean_ctor_get(x_132, 1); +lean_inc(x_134); +if (lean_is_exclusive(x_132)) { + lean_ctor_release(x_132, 0); + lean_ctor_release(x_132, 1); + x_135 = x_132; } else { - lean_dec_ref(x_149); - x_152 = lean_box(0); + lean_dec_ref(x_132); + x_135 = lean_box(0); } -x_153 = lean_alloc_ctor(11, 3, 8); -lean_ctor_set(x_153, 0, x_145); -lean_ctor_set(x_153, 1, x_146); -lean_ctor_set(x_153, 2, x_147); -lean_ctor_set_uint64(x_153, sizeof(void*)*3, x_148); -x_154 = lean_expr_update_proj(x_153, x_150); -if (lean_is_scalar(x_152)) { - x_155 = lean_alloc_ctor(0, 2, 0); +x_136 = lean_alloc_ctor(11, 3, 8); +lean_ctor_set(x_136, 0, x_128); +lean_ctor_set(x_136, 1, x_129); +lean_ctor_set(x_136, 2, x_130); +lean_ctor_set_uint64(x_136, sizeof(void*)*3, x_131); +x_137 = lean_expr_update_proj(x_136, x_133); +if (lean_is_scalar(x_135)) { + x_138 = lean_alloc_ctor(0, 2, 0); } else { - x_155 = x_152; + x_138 = x_135; } -lean_ctor_set(x_155, 0, x_154); -lean_ctor_set(x_155, 1, x_151); -return x_155; +lean_ctor_set(x_138, 0, x_137); +lean_ctor_set(x_138, 1, x_134); +return x_138; } else { -lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; -lean_dec(x_147); -lean_dec(x_146); -lean_dec(x_145); -x_156 = lean_ctor_get(x_149, 0); -lean_inc(x_156); -x_157 = lean_ctor_get(x_149, 1); -lean_inc(x_157); -if (lean_is_exclusive(x_149)) { - lean_ctor_release(x_149, 0); - lean_ctor_release(x_149, 1); - x_158 = x_149; +lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; +lean_dec(x_130); +lean_dec(x_129); +lean_dec(x_128); +x_139 = lean_ctor_get(x_132, 0); +lean_inc(x_139); +x_140 = lean_ctor_get(x_132, 1); +lean_inc(x_140); +if (lean_is_exclusive(x_132)) { + lean_ctor_release(x_132, 0); + lean_ctor_release(x_132, 1); + x_141 = x_132; } else { - lean_dec_ref(x_149); - x_158 = lean_box(0); + lean_dec_ref(x_132); + x_141 = lean_box(0); } -if (lean_is_scalar(x_158)) { - x_159 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_141)) { + x_142 = lean_alloc_ctor(1, 2, 0); } else { - x_159 = x_158; + x_142 = x_141; } -lean_ctor_set(x_159, 0, x_156); -lean_ctor_set(x_159, 1, x_157); -return x_159; +lean_ctor_set(x_142, 0, x_139); +lean_ctor_set(x_142, 1, x_140); +return x_142; } } } default: { -lean_object* x_160; +lean_object* x_143; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_160 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_160, 0, x_1); -lean_ctor_set(x_160, 1, x_9); -return x_160; +x_143 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_143, 0, x_1); +lean_ctor_set(x_143, 1, x_9); +return x_143; } } } @@ -23624,6 +24405,40 @@ x_14 = l_Lean_Meta_withLetDecl___at_Lean_Elab_Term_ToDepElimPattern_savePatternI return x_14; } } +LEAN_EXPORT lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4___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) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4___rarg(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4___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_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4(x_3, x_2); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___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_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_3); +lean_dec(x_3); +x_12 = l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__3(x_1, x_2, x_11, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_12; +} +} LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___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, lean_object* x_9, lean_object* x_10) { _start: { @@ -23646,6 +24461,22 @@ lean_dec(x_1); return x_12; } } +LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___lambda__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_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_3); +lean_dec(x_3); +x_12 = l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___lambda__3(x_1, x_2, x_11, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_12; +} +} LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___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_object* x_9) { _start: { @@ -34913,7 +35744,7 @@ lean_dec(x_2); return x_9; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13146____closed__1() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13231____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -34923,7 +35754,7 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13146____closed__2() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13231____closed__2() { _start: { lean_object* x_1; @@ -34931,17 +35762,17 @@ x_1 = lean_mk_string_from_bytes("ignoreUnusedAlts", 16); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13146____closed__3() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13231____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13146____closed__1; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13146____closed__2; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13231____closed__1; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13231____closed__2; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13146____closed__4() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13231____closed__4() { _start: { lean_object* x_1; @@ -34949,13 +35780,13 @@ x_1 = lean_mk_string_from_bytes("if true, do not generate error if an alternativ return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13146____closed__5() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13231____closed__5() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = 0; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___spec__1___closed__15; -x_3 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13146____closed__4; +x_3 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13231____closed__4; x_4 = lean_box(x_1); x_5 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_5, 0, x_4); @@ -34964,12 +35795,12 @@ lean_ctor_set(x_5, 2, x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13146_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13231_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13146____closed__3; -x_3 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13146____closed__5; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13231____closed__3; +x_3 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13231____closed__5; x_4 = l_Lean_Option_register___at_Lean_Elab_initFn____x40_Lean_Elab_AutoBound___hyg_6____spec__1(x_2, x_3, x_1); return x_4; } @@ -35623,7 +36454,7 @@ static lean_object* _init_l___private_Lean_Elab_Match_0__Lean_Elab_Term_isMatchU 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___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_toPattern___closed__1; x_2 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_isMatchUnit_x3f___closed__4; -x_3 = lean_unsigned_to_nat(1029u); +x_3 = lean_unsigned_to_nat(1026u); x_4 = lean_unsigned_to_nat(2u); x_5 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_isMatchUnit_x3f___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -37485,7 +38316,7 @@ goto _start; } } } -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__12(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__12(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -38192,7 +39023,7 @@ if (lean_obj_tag(x_28) == 0) lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_dec(x_21); x_30 = lean_array_get_size(x_22); -x_31 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13146____closed__1; +x_31 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13231____closed__1; lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); @@ -38669,7 +39500,7 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); x_15 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___closed__11; -x_16 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__12(x_1, x_15); +x_16 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__12(x_1, x_15); if (x_16 == 0) { lean_object* x_17; lean_object* x_18; @@ -38854,11 +39685,11 @@ x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_el return x_6; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__12___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__12___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__12(x_1, x_2); +x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__12(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -40590,6 +41421,30 @@ lean_dec(x_1); return x_10; } } +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchCore___spec__1(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = 1; +x_9 = lean_usize_add(x_2, x_8); +x_10 = lean_array_uset(x_7, x_2, x_5); +x_2 = x_9; +x_3 = x_10; +goto _start; +} +} +} LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchCore(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_object* x_9) { _start: { @@ -40613,7 +41468,7 @@ x_14 = lean_array_get_size(x_13); x_15 = lean_usize_of_nat(x_14); lean_dec(x_14); x_16 = 0; -x_17 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_15, x_16, x_13); +x_17 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchCore___spec__1(x_15, x_16, x_13); lean_inc(x_1); x_18 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f(x_1); lean_inc(x_1); @@ -40655,6 +41510,18 @@ return x_26; } } } +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchCore___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchCore___spec__1(x_4, x_5, x_3); +return x_6; +} +} LEAN_EXPORT uint8_t l___private_Lean_Elab_Match_0__Lean_Elab_Term_isPatternVar_isAtomicIdent(lean_object* x_1) { _start: { @@ -41805,16 +42672,6 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_elabMatch___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__6; -x_2 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__2___closed__1; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMatch(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_object* x_9) { _start: { @@ -41830,216 +42687,215 @@ return x_12; } else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; x_13 = lean_unsigned_to_nat(1u); x_14 = l_Lean_Syntax_getArg(x_1, x_13); -x_15 = l_Lean_nullKind; -x_16 = lean_unsigned_to_nat(0u); -x_17 = l_Lean_Syntax_isNodeOf(x_14, x_15, x_16); -if (x_17 == 0) +x_15 = lean_unsigned_to_nat(0u); +x_16 = l_Lean_Syntax_matchesNull(x_14, x_15); +if (x_16 == 0) { -lean_object* x_18; -x_18 = l_Lean_Elab_Term_elabMatch_elabMatchDefault(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_18; +lean_object* x_17; +x_17 = l_Lean_Elab_Term_elabMatch_elabMatchDefault(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_17; } else { -lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_19 = lean_unsigned_to_nat(2u); -x_20 = l_Lean_Syntax_getArg(x_1, x_19); -x_21 = l_Lean_Syntax_isNodeOf(x_20, x_15, x_16); -if (x_21 == 0) +lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_18 = lean_unsigned_to_nat(2u); +x_19 = l_Lean_Syntax_getArg(x_1, x_18); +x_20 = l_Lean_Syntax_matchesNull(x_19, x_15); +if (x_20 == 0) { -lean_object* x_22; -x_22 = l_Lean_Elab_Term_elabMatch_elabMatchDefault(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_22; +lean_object* x_21; +x_21 = l_Lean_Elab_Term_elabMatch_elabMatchDefault(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_21; } else { -lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_23 = lean_unsigned_to_nat(3u); -x_24 = l_Lean_Syntax_getArg(x_1, x_23); -lean_inc(x_24); -x_25 = l_Lean_Syntax_isNodeOf(x_24, x_15, x_13); -if (x_25 == 0) +lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_22 = lean_unsigned_to_nat(3u); +x_23 = l_Lean_Syntax_getArg(x_1, x_22); +lean_inc(x_23); +x_24 = l_Lean_Syntax_matchesNull(x_23, x_13); +if (x_24 == 0) { -lean_object* x_26; -lean_dec(x_24); -x_26 = l_Lean_Elab_Term_elabMatch_elabMatchDefault(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_26; +lean_object* x_25; +lean_dec(x_23); +x_25 = l_Lean_Elab_Term_elabMatch_elabMatchDefault(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_25; } else { -lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_27 = l_Lean_Syntax_getArg(x_24, x_16); -lean_dec(x_24); -x_28 = l_Lean_Elab_Term_elabMatch___closed__1; -lean_inc(x_27); -x_29 = l_Lean_Syntax_isOfKind(x_27, x_28); -if (x_29 == 0) +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = l_Lean_Syntax_getArg(x_23, x_15); +lean_dec(x_23); +x_27 = l_Lean_Elab_Term_elabMatch___closed__1; +lean_inc(x_26); +x_28 = l_Lean_Syntax_isOfKind(x_26, x_27); +if (x_28 == 0) { -lean_object* x_30; -lean_dec(x_27); -x_30 = l_Lean_Elab_Term_elabMatch_elabMatchDefault(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_30; +lean_object* x_29; +lean_dec(x_26); +x_29 = l_Lean_Elab_Term_elabMatch_elabMatchDefault(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_29; } else { -lean_object* x_31; uint8_t x_32; -x_31 = l_Lean_Syntax_getArg(x_27, x_16); -x_32 = l_Lean_Syntax_isNodeOf(x_31, x_15, x_16); -if (x_32 == 0) +lean_object* x_30; uint8_t x_31; +x_30 = l_Lean_Syntax_getArg(x_26, x_15); +x_31 = l_Lean_Syntax_matchesNull(x_30, x_15); +if (x_31 == 0) { -lean_object* x_33; -lean_dec(x_27); -x_33 = l_Lean_Elab_Term_elabMatch_elabMatchDefault(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_33; +lean_object* x_32; +lean_dec(x_26); +x_32 = l_Lean_Elab_Term_elabMatch_elabMatchDefault(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_32; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; -x_34 = l_Lean_Syntax_getArg(x_27, x_13); -lean_dec(x_27); -x_35 = lean_unsigned_to_nat(5u); -x_36 = l_Lean_Syntax_getArg(x_1, x_35); -x_37 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___closed__9; -lean_inc(x_36); -x_38 = l_Lean_Syntax_isOfKind(x_36, x_37); -if (x_38 == 0) +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_33 = l_Lean_Syntax_getArg(x_26, x_13); +lean_dec(x_26); +x_34 = lean_unsigned_to_nat(5u); +x_35 = l_Lean_Syntax_getArg(x_1, x_34); +x_36 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___closed__9; +lean_inc(x_35); +x_37 = l_Lean_Syntax_isOfKind(x_35, x_36); +if (x_37 == 0) { -lean_object* x_39; -lean_dec(x_36); -lean_dec(x_34); -x_39 = l_Lean_Elab_Term_elabMatch_elabMatchDefault(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_39; +lean_object* x_38; +lean_dec(x_35); +lean_dec(x_33); +x_38 = l_Lean_Elab_Term_elabMatch_elabMatchDefault(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_38; } else { -lean_object* x_40; uint8_t x_41; -x_40 = l_Lean_Syntax_getArg(x_36, x_16); -lean_dec(x_36); -lean_inc(x_40); -x_41 = l_Lean_Syntax_isNodeOf(x_40, x_15, x_13); -if (x_41 == 0) +lean_object* x_39; uint8_t x_40; +x_39 = l_Lean_Syntax_getArg(x_35, x_15); +lean_dec(x_35); +lean_inc(x_39); +x_40 = l_Lean_Syntax_matchesNull(x_39, x_13); +if (x_40 == 0) { -lean_object* x_42; -lean_dec(x_40); -lean_dec(x_34); -x_42 = l_Lean_Elab_Term_elabMatch_elabMatchDefault(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_42; +lean_object* x_41; +lean_dec(x_39); +lean_dec(x_33); +x_41 = l_Lean_Elab_Term_elabMatch_elabMatchDefault(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_41; } else { -lean_object* x_43; lean_object* x_44; uint8_t x_45; -x_43 = l_Lean_Syntax_getArg(x_40, x_16); -lean_dec(x_40); -x_44 = l_Lean_Elab_Term_elabMatch___closed__2; -lean_inc(x_43); -x_45 = l_Lean_Syntax_isOfKind(x_43, x_44); -if (x_45 == 0) +lean_object* x_42; lean_object* x_43; uint8_t x_44; +x_42 = l_Lean_Syntax_getArg(x_39, x_15); +lean_dec(x_39); +x_43 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___closed__2; +lean_inc(x_42); +x_44 = l_Lean_Syntax_isOfKind(x_42, x_43); +if (x_44 == 0) { -lean_object* x_46; -lean_dec(x_43); -lean_dec(x_34); -x_46 = l_Lean_Elab_Term_elabMatch_elabMatchDefault(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_46; +lean_object* x_45; +lean_dec(x_42); +lean_dec(x_33); +x_45 = l_Lean_Elab_Term_elabMatch_elabMatchDefault(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_45; } else { -lean_object* x_47; uint8_t x_48; -x_47 = l_Lean_Syntax_getArg(x_43, x_13); -lean_inc(x_47); -x_48 = l_Lean_Syntax_isNodeOf(x_47, x_15, x_13); -if (x_48 == 0) +lean_object* x_46; uint8_t x_47; +x_46 = l_Lean_Syntax_getArg(x_42, x_13); +lean_inc(x_46); +x_47 = l_Lean_Syntax_matchesNull(x_46, x_13); +if (x_47 == 0) { -lean_object* x_49; -lean_dec(x_47); -lean_dec(x_43); -lean_dec(x_34); -x_49 = l_Lean_Elab_Term_elabMatch_elabMatchDefault(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_49; +lean_object* x_48; +lean_dec(x_46); +lean_dec(x_42); +lean_dec(x_33); +x_48 = l_Lean_Elab_Term_elabMatch_elabMatchDefault(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_48; } else { -lean_object* x_50; uint8_t x_51; -x_50 = l_Lean_Syntax_getArg(x_47, x_16); -lean_dec(x_47); -lean_inc(x_50); -x_51 = l_Lean_Syntax_isNodeOf(x_50, x_15, x_13); -if (x_51 == 0) +lean_object* x_49; uint8_t x_50; +x_49 = l_Lean_Syntax_getArg(x_46, x_15); +lean_dec(x_46); +lean_inc(x_49); +x_50 = l_Lean_Syntax_matchesNull(x_49, x_13); +if (x_50 == 0) { -lean_object* x_52; -lean_dec(x_50); -lean_dec(x_43); -lean_dec(x_34); -x_52 = l_Lean_Elab_Term_elabMatch_elabMatchDefault(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_52; +lean_object* x_51; +lean_dec(x_49); +lean_dec(x_42); +lean_dec(x_33); +x_51 = l_Lean_Elab_Term_elabMatch_elabMatchDefault(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_51; } else { -lean_object* x_53; lean_object* x_54; uint8_t x_55; -x_53 = l_Lean_Syntax_getArg(x_50, x_16); -lean_dec(x_50); -x_54 = l_Lean_Elab_Term_isAtomicDiscr_x3f___closed__2; -lean_inc(x_53); -x_55 = l_Lean_Syntax_isOfKind(x_53, x_54); -if (x_55 == 0) +lean_object* x_52; lean_object* x_53; uint8_t x_54; +x_52 = l_Lean_Syntax_getArg(x_49, x_15); +lean_dec(x_49); +x_53 = l_Lean_Elab_Term_isAtomicDiscr_x3f___closed__2; +lean_inc(x_52); +x_54 = l_Lean_Syntax_isOfKind(x_52, x_53); +if (x_54 == 0) { -lean_object* x_56; -lean_dec(x_53); -lean_dec(x_43); -lean_dec(x_34); -x_56 = l_Lean_Elab_Term_elabMatch_elabMatchDefault(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_56; +lean_object* x_55; +lean_dec(x_52); +lean_dec(x_42); +lean_dec(x_33); +x_55 = l_Lean_Elab_Term_elabMatch_elabMatchDefault(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_55; } else { -lean_object* x_57; lean_object* x_58; -x_57 = l_Lean_Syntax_getArg(x_43, x_23); -lean_dec(x_43); +lean_object* x_56; lean_object* x_57; +x_56 = l_Lean_Syntax_getArg(x_42, x_22); +lean_dec(x_42); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -lean_inc(x_53); -x_58 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_isPatternVar(x_53, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_58) == 0) +lean_inc(x_52); +x_57 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_isPatternVar(x_52, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_57) == 0) { -lean_object* x_59; uint8_t x_60; -x_59 = lean_ctor_get(x_58, 0); -lean_inc(x_59); -x_60 = lean_unbox(x_59); -lean_dec(x_59); -if (x_60 == 0) +lean_object* x_58; uint8_t x_59; +x_58 = lean_ctor_get(x_57, 0); +lean_inc(x_58); +x_59 = lean_unbox(x_58); +lean_dec(x_58); +if (x_59 == 0) { -lean_object* x_61; lean_object* x_62; +lean_object* x_60; lean_object* x_61; +lean_dec(x_56); +lean_dec(x_52); +lean_dec(x_33); +x_60 = lean_ctor_get(x_57, 1); +lean_inc(x_60); lean_dec(x_57); -lean_dec(x_53); -lean_dec(x_34); -x_61 = lean_ctor_get(x_58, 1); -lean_inc(x_61); -lean_dec(x_58); -x_62 = l_Lean_Elab_Term_elabMatch_elabMatchDefault(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_61); -return x_62; +x_61 = l_Lean_Elab_Term_elabMatch_elabMatchDefault(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_60); +return x_61; } else { -lean_object* x_63; lean_object* x_64; -x_63 = lean_ctor_get(x_58, 1); -lean_inc(x_63); -lean_dec(x_58); -x_64 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch(x_1, x_34, x_53, x_57, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_63); -return x_64; +lean_object* x_62; lean_object* x_63; +x_62 = lean_ctor_get(x_57, 1); +lean_inc(x_62); +lean_dec(x_57); +x_63 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch(x_1, x_33, x_52, x_56, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_62); +return x_63; } } else { -uint8_t x_65; -lean_dec(x_57); -lean_dec(x_53); -lean_dec(x_34); +uint8_t x_64; +lean_dec(x_56); +lean_dec(x_52); +lean_dec(x_33); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -42048,23 +42904,23 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_65 = !lean_is_exclusive(x_58); -if (x_65 == 0) +x_64 = !lean_is_exclusive(x_57); +if (x_64 == 0) { -return x_58; +return x_57; } else { -lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_66 = lean_ctor_get(x_58, 0); -x_67 = lean_ctor_get(x_58, 1); -lean_inc(x_67); +lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = lean_ctor_get(x_57, 0); +x_66 = lean_ctor_get(x_57, 1); lean_inc(x_66); -lean_dec(x_58); -x_68 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_68, 0, x_66); -lean_ctor_set(x_68, 1, x_67); -return x_68; +lean_inc(x_65); +lean_dec(x_57); +x_67 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_67, 0, x_65); +lean_ctor_set(x_67, 1, x_66); +return x_67; } } } @@ -42141,7 +42997,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabMatch_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1254u); +x_1 = lean_unsigned_to_nat(1252u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -42153,7 +43009,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabMatch_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1271u); +x_1 = lean_unsigned_to_nat(1269u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -42181,7 +43037,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabMatch_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1254u); +x_1 = lean_unsigned_to_nat(1252u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -42193,7 +43049,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabMatch_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1254u); +x_1 = lean_unsigned_to_nat(1252u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -42239,7 +43095,7 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_16449_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_16570_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; @@ -42271,7 +43127,7 @@ static lean_object* _init_l_Lean_Elab_Term_elabNoMatch___closed__3() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean_nullKind; +x_2 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__14; x_3 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__15; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -42545,7 +43401,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabNoMatch_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1280u); +x_1 = lean_unsigned_to_nat(1278u); x_2 = lean_unsigned_to_nat(29u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -42557,7 +43413,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabNoMatch_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1292u); +x_1 = lean_unsigned_to_nat(1290u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -42585,7 +43441,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabNoMatch_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1280u); +x_1 = lean_unsigned_to_nat(1278u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -42597,7 +43453,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabNoMatch_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(1280u); +x_1 = lean_unsigned_to_nat(1278u); x_2 = lean_unsigned_to_nat(44u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -42831,8 +43687,10 @@ l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___closed_ lean_mark_persistent(l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___closed__10); l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___closed__11 = _init_l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___closed__11(); lean_mark_persistent(l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___closed__11); -l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__2___closed__1 = _init_l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__2___closed__1(); -lean_mark_persistent(l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__2___closed__1); +l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___closed__1 = _init_l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___closed__1(); +lean_mark_persistent(l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___closed__1); +l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___closed__2 = _init_l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___closed__2(); +lean_mark_persistent(l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___closed__2); l___regBuiltin_Lean_Elab_Term_elabInaccessible___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_elabInaccessible___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_elabInaccessible___closed__1); l___regBuiltin_Lean_Elab_Term_elabInaccessible___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_elabInaccessible___closed__2(); @@ -42946,6 +43804,12 @@ l_Lean_Elab_Term_ToDepElimPattern_TopSort_State_result___default = _init_l_Lean_ lean_mark_persistent(l_Lean_Elab_Term_ToDepElimPattern_TopSort_State_result___default); l___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_topSort___closed__1 = _init_l___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_topSort___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_topSort___closed__1); +l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4___rarg___closed__1 = _init_l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4___rarg___closed__1(); +lean_mark_persistent(l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4___rarg___closed__1); +l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4___rarg___closed__2 = _init_l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4___rarg___closed__2(); +lean_mark_persistent(l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4___rarg___closed__2); +l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4___rarg___closed__3 = _init_l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4___rarg___closed__3(); +lean_mark_persistent(l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4___rarg___closed__3); l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__2___closed__1 = _init_l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__2___closed__1(); lean_mark_persistent(l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__2___closed__1); l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__2___closed__2 = _init_l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__2___closed__2(); @@ -43033,17 +43897,17 @@ l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_loop___closed__1 lean_mark_persistent(l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_loop___closed__1); l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_loop___closed__2 = _init_l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_loop___closed__2(); lean_mark_persistent(l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_loop___closed__2); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13146____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13146____closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13146____closed__1); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13146____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13146____closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13146____closed__2); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13146____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13146____closed__3(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13146____closed__3); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13146____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13146____closed__4(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13146____closed__4); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13146____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13146____closed__5(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13146____closed__5); -if (builtin) {res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13146_(lean_io_mk_world()); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13231____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13231____closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13231____closed__1); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13231____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13231____closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13231____closed__2); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13231____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13231____closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13231____closed__3); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13231____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13231____closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13231____closed__4); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13231____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13231____closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13231____closed__5); +if (builtin) {res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_13231_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Elab_Term_match_ignoreUnusedAlts = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Elab_Term_match_ignoreUnusedAlts); @@ -43136,8 +44000,6 @@ l_Lean_Elab_Term_elabMatch_elabMatchDefault___closed__2 = _init_l_Lean_Elab_Term lean_mark_persistent(l_Lean_Elab_Term_elabMatch_elabMatchDefault___closed__2); l_Lean_Elab_Term_elabMatch___closed__1 = _init_l_Lean_Elab_Term_elabMatch___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_elabMatch___closed__1); -l_Lean_Elab_Term_elabMatch___closed__2 = _init_l_Lean_Elab_Term_elabMatch___closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_elabMatch___closed__2); l___regBuiltin_Lean_Elab_Term_elabMatch___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_elabMatch___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_elabMatch___closed__1); l___regBuiltin_Lean_Elab_Term_elabMatch___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_elabMatch___closed__2(); @@ -43169,7 +44031,7 @@ lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_elabMatch_declRange___closed_ res = l___regBuiltin_Lean_Elab_Term_elabMatch_declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_16449_(lean_io_mk_world()); +res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_16570_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); l_Lean_Elab_Term_elabNoMatch___closed__1 = _init_l_Lean_Elab_Term_elabNoMatch___closed__1(); diff --git a/stage0/stdlib/Lean/Elab/Mixfix.c b/stage0/stdlib/Lean/Elab/Mixfix.c index 0cb13fd2bed..1d4197257c3 100644 --- a/stage0/stdlib/Lean/Elab/Mixfix.c +++ b/stage0/stdlib/Lean/Elab/Mixfix.c @@ -15,7 +15,6 @@ extern "C" { #endif lean_object* lean_mk_empty_array_with_capacity(lean_object*); static lean_object* l_Lean_Elab_Command_expandMixfix___lambda__1___closed__11; -extern lean_object* l_Lean_nullKind; static lean_object* l_Lean_Elab_Command_expandMixfix___lambda__5___closed__6; lean_object* lean_name_mk_string(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_expandMixfix___lambda__11___closed__20; @@ -54,7 +53,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Command_expandMixfix___closed__7; static lean_object* l_Lean_Elab_Command_expandMixfix___lambda__1___closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandMixfix___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_expandMixfix___lambda__1___closed__23; -extern lean_object* l_Lean_numLitKind; +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_expandMixfix___lambda__5___closed__2; static lean_object* l_Lean_Elab_Command_expandMixfix___lambda__1___closed__22; static lean_object* l___regBuiltin_Lean_Elab_Command_expandMixfix_declRange___closed__7; @@ -92,7 +91,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandMixfix(lean_object*, lean_obj static lean_object* l_Lean_Elab_Command_expandMixfix___lambda__1___closed__10; lean_object* l_Lean_evalPrec(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandMixfix___lambda__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_mkNumLit(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_expandMixfix_declRange___closed__6; extern lean_object* l_Lean_Elab_macroAttribute; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_expandMixfix(lean_object*); @@ -135,7 +134,6 @@ static lean_object* l_Lean_Elab_Command_expandMixfix___lambda__11___closed__14; static lean_object* l___regBuiltin_Lean_Elab_Command_expandMixfix_declRange___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Command_expandMixfix___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandMixfix___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Syntax_mkLit(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandMixfix_withAttrKindGlobal(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { @@ -708,14 +706,13 @@ x_12 = l_Lean_Syntax_getArg(x_1, x_11); x_13 = l_Lean_Syntax_isNone(x_12); if (x_13 == 0) { -lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_14 = l_Lean_nullKind; -x_15 = lean_unsigned_to_nat(1u); +lean_object* x_14; uint8_t x_15; +x_14 = lean_unsigned_to_nat(1u); lean_inc(x_12); -x_16 = l_Lean_Syntax_isNodeOf(x_12, x_14, x_15); -if (x_16 == 0) +x_15 = l_Lean_Syntax_matchesNull(x_12, x_14); +if (x_15 == 0) { -lean_object* x_17; lean_object* x_18; +lean_object* x_16; lean_object* x_17; lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); @@ -724,28 +721,28 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_17 = lean_box(1); -x_18 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_10); -return x_18; +x_16 = lean_box(1); +x_17 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_10); +return x_17; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_19 = lean_unsigned_to_nat(0u); -x_20 = l_Lean_Syntax_getArg(x_12, x_19); +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_18 = lean_unsigned_to_nat(0u); +x_19 = l_Lean_Syntax_getArg(x_12, x_18); lean_dec(x_12); -x_21 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__20; +x_20 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__20; lean_inc(x_2); -x_22 = lean_name_mk_string(x_2, x_21); -lean_inc(x_20); -x_23 = l_Lean_Syntax_isOfKind(x_20, x_22); -lean_dec(x_22); -if (x_23 == 0) +x_21 = lean_name_mk_string(x_2, x_20); +lean_inc(x_19); +x_22 = l_Lean_Syntax_isOfKind(x_19, x_21); +lean_dec(x_21); +if (x_22 == 0) { -lean_object* x_24; lean_object* x_25; -lean_dec(x_20); +lean_object* x_23; lean_object* x_24; +lean_dec(x_19); lean_dec(x_9); lean_dec(x_8); lean_dec(x_6); @@ -753,34 +750,34 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_24 = lean_box(1); -x_25 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_10); -return x_25; +x_23 = lean_box(1); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_10); +return x_24; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_26 = lean_unsigned_to_nat(3u); -x_27 = l_Lean_Syntax_getArg(x_20, x_26); -lean_dec(x_20); -x_28 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_28, 0, x_27); -x_29 = lean_box(0); -x_30 = l_Lean_Elab_Command_expandMixfix___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_8, x_29, x_28, x_9, x_10); -return x_30; +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_25 = lean_unsigned_to_nat(3u); +x_26 = l_Lean_Syntax_getArg(x_19, x_25); +lean_dec(x_19); +x_27 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_27, 0, x_26); +x_28 = lean_box(0); +x_29 = l_Lean_Elab_Command_expandMixfix___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_8, x_28, x_27, x_9, x_10); +return x_29; } } } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_dec(x_12); +x_30 = lean_box(0); x_31 = lean_box(0); -x_32 = lean_box(0); -x_33 = l_Lean_Elab_Command_expandMixfix___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_8, x_32, x_31, x_9, x_10); -return x_33; +x_32 = l_Lean_Elab_Command_expandMixfix___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_8, x_31, x_30, x_9, x_10); +return x_32; } } } @@ -1048,14 +1045,13 @@ x_12 = l_Lean_Syntax_getArg(x_1, x_11); x_13 = l_Lean_Syntax_isNone(x_12); if (x_13 == 0) { -lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_14 = l_Lean_nullKind; -x_15 = lean_unsigned_to_nat(1u); +lean_object* x_14; uint8_t x_15; +x_14 = lean_unsigned_to_nat(1u); lean_inc(x_12); -x_16 = l_Lean_Syntax_isNodeOf(x_12, x_14, x_15); -if (x_16 == 0) +x_15 = l_Lean_Syntax_matchesNull(x_12, x_14); +if (x_15 == 0) { -lean_object* x_17; lean_object* x_18; +lean_object* x_16; lean_object* x_17; lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); @@ -1064,28 +1060,28 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_17 = lean_box(1); -x_18 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_10); -return x_18; +x_16 = lean_box(1); +x_17 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_10); +return x_17; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_19 = lean_unsigned_to_nat(0u); -x_20 = l_Lean_Syntax_getArg(x_12, x_19); +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_18 = lean_unsigned_to_nat(0u); +x_19 = l_Lean_Syntax_getArg(x_12, x_18); lean_dec(x_12); -x_21 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__20; +x_20 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__20; lean_inc(x_2); -x_22 = lean_name_mk_string(x_2, x_21); -lean_inc(x_20); -x_23 = l_Lean_Syntax_isOfKind(x_20, x_22); -lean_dec(x_22); -if (x_23 == 0) +x_21 = lean_name_mk_string(x_2, x_20); +lean_inc(x_19); +x_22 = l_Lean_Syntax_isOfKind(x_19, x_21); +lean_dec(x_21); +if (x_22 == 0) { -lean_object* x_24; lean_object* x_25; -lean_dec(x_20); +lean_object* x_23; lean_object* x_24; +lean_dec(x_19); lean_dec(x_9); lean_dec(x_8); lean_dec(x_6); @@ -1093,34 +1089,34 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_24 = lean_box(1); -x_25 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_10); -return x_25; +x_23 = lean_box(1); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_10); +return x_24; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_26 = lean_unsigned_to_nat(3u); -x_27 = l_Lean_Syntax_getArg(x_20, x_26); -lean_dec(x_20); -x_28 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_28, 0, x_27); -x_29 = lean_box(0); -x_30 = l_Lean_Elab_Command_expandMixfix___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_8, x_29, x_28, x_9, x_10); -return x_30; +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_25 = lean_unsigned_to_nat(3u); +x_26 = l_Lean_Syntax_getArg(x_19, x_25); +lean_dec(x_19); +x_27 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_27, 0, x_26); +x_28 = lean_box(0); +x_29 = l_Lean_Elab_Command_expandMixfix___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_8, x_28, x_27, x_9, x_10); +return x_29; } } } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_dec(x_12); +x_30 = lean_box(0); x_31 = lean_box(0); -x_32 = lean_box(0); -x_33 = l_Lean_Elab_Command_expandMixfix___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_8, x_32, x_31, x_9, x_10); -return x_33; +x_32 = l_Lean_Elab_Command_expandMixfix___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_8, x_31, x_30, x_9, x_10); +return x_32; } } } @@ -1228,7 +1224,7 @@ lean_inc(x_2); x_16 = l_Lean_evalPrec(x_2, x_10, x_11); if (lean_obj_tag(x_16) == 0) { -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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); x_18 = lean_ctor_get(x_16, 1); @@ -1238,289 +1234,288 @@ x_19 = lean_unsigned_to_nat(1u); x_20 = lean_nat_add(x_17, x_19); lean_dec(x_17); x_21 = l_Nat_repr(x_20); -x_22 = l_Lean_numLitKind; -x_23 = lean_box(2); -x_24 = l_Lean_Syntax_mkLit(x_22, x_21, x_23); +x_22 = lean_box(2); +x_23 = l_Lean_Syntax_mkNumLit(x_21, x_22); lean_inc(x_10); -x_25 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_10, x_18); -x_26 = lean_ctor_get(x_25, 0); +x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_10, x_18); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 1); -lean_inc(x_27); -if (lean_is_exclusive(x_25)) { - lean_ctor_release(x_25, 0); - lean_ctor_release(x_25, 1); - x_28 = x_25; +if (lean_is_exclusive(x_24)) { + lean_ctor_release(x_24, 0); + lean_ctor_release(x_24, 1); + x_27 = x_24; } else { - lean_dec_ref(x_25); - x_28 = lean_box(0); + lean_dec_ref(x_24); + x_27 = lean_box(0); } -x_29 = lean_ctor_get(x_10, 2); +x_28 = lean_ctor_get(x_10, 2); +lean_inc(x_28); +x_29 = lean_ctor_get(x_10, 1); lean_inc(x_29); -x_30 = lean_ctor_get(x_10, 1); -lean_inc(x_30); lean_dec(x_10); -x_31 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__1; +x_30 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__1; lean_inc(x_3); -x_32 = lean_name_mk_string(x_3, x_31); -x_33 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__7; -x_34 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_34, 0, x_23); -lean_ctor_set(x_34, 1, x_4); -lean_ctor_set(x_34, 2, x_33); -lean_inc(x_26); -x_35 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_35, 0, x_26); -lean_ctor_set(x_35, 1, x_31); -x_36 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__8; -lean_inc(x_26); -x_37 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_37, 0, x_26); -lean_ctor_set(x_37, 1, x_36); -x_38 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__9; -x_39 = lean_array_push(x_38, x_37); -lean_inc(x_39); -x_40 = lean_array_push(x_39, x_2); +x_31 = lean_name_mk_string(x_3, x_30); +x_32 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__7; +x_33 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_33, 0, x_22); +lean_ctor_set(x_33, 1, x_4); +lean_ctor_set(x_33, 2, x_32); +lean_inc(x_25); +x_34 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_34, 0, x_25); +lean_ctor_set(x_34, 1, x_30); +x_35 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__8; +lean_inc(x_25); +x_36 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_36, 0, x_25); +lean_ctor_set(x_36, 1, x_35); +x_37 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__9; +x_38 = lean_array_push(x_37, x_36); +lean_inc(x_38); +x_39 = lean_array_push(x_38, x_2); lean_inc(x_5); -x_41 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_41, 0, x_23); -lean_ctor_set(x_41, 1, x_5); -lean_ctor_set(x_41, 2, x_40); -x_42 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__6; -x_43 = lean_array_push(x_42, x_41); -x_44 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__3; -x_45 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_45, 0, x_23); -lean_ctor_set(x_45, 1, x_44); -lean_ctor_set(x_45, 2, x_43); -x_46 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__10; +x_40 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_40, 0, x_22); +lean_ctor_set(x_40, 1, x_5); +lean_ctor_set(x_40, 2, x_39); +x_41 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__6; +x_42 = lean_array_push(x_41, x_40); +x_43 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__3; +x_44 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_44, 0, x_22); +lean_ctor_set(x_44, 1, x_43); +lean_ctor_set(x_44, 2, x_42); +x_45 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__10; lean_inc(x_3); -x_47 = lean_name_mk_string(x_3, x_46); -x_48 = l_Lean_Elab_Command_expandMixfix___lambda__5___closed__4; +x_46 = lean_name_mk_string(x_3, x_45); +x_47 = l_Lean_Elab_Command_expandMixfix___lambda__5___closed__4; +lean_inc(x_28); lean_inc(x_29); -lean_inc(x_30); -x_49 = l_Lean_addMacroScope(x_30, x_48, x_29); -x_50 = lean_box(0); -x_51 = l_Lean_Elab_Command_expandMixfix___lambda__5___closed__3; -lean_inc(x_26); -x_52 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_52, 0, x_26); -lean_ctor_set(x_52, 1, x_51); -lean_ctor_set(x_52, 2, x_49); -lean_ctor_set(x_52, 3, x_50); -x_53 = lean_array_push(x_39, x_24); -x_54 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_54, 0, x_23); -lean_ctor_set(x_54, 1, x_5); -lean_ctor_set(x_54, 2, x_53); -x_55 = lean_array_push(x_42, x_54); -x_56 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_56, 0, x_23); -lean_ctor_set(x_56, 1, x_44); -lean_ctor_set(x_56, 2, x_55); -x_57 = lean_array_push(x_38, x_52); -lean_inc(x_57); -x_58 = lean_array_push(x_57, x_56); -lean_inc(x_47); -x_59 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_59, 0, x_23); -lean_ctor_set(x_59, 1, x_47); -lean_ctor_set(x_59, 2, x_58); -x_60 = l_Lean_Elab_Command_expandMixfix___lambda__5___closed__8; -x_61 = l_Lean_addMacroScope(x_30, x_60, x_29); -x_62 = l_Lean_Elab_Command_expandMixfix___lambda__5___closed__7; -lean_inc(x_26); -x_63 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_63, 0, x_26); -lean_ctor_set(x_63, 1, x_62); -lean_ctor_set(x_63, 2, x_61); -lean_ctor_set(x_63, 3, x_50); -lean_inc(x_63); -x_64 = lean_array_push(x_38, x_63); -lean_inc(x_45); -x_65 = lean_array_push(x_64, x_45); -x_66 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_66, 0, x_23); -lean_ctor_set(x_66, 1, x_47); -lean_ctor_set(x_66, 2, x_65); -x_67 = l_Lean_Elab_Command_expandMixfix___lambda__5___closed__9; -x_68 = lean_array_push(x_67, x_59); -x_69 = lean_array_push(x_68, x_13); -x_70 = lean_array_push(x_69, x_66); -x_71 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_71, 0, x_23); -lean_ctor_set(x_71, 1, x_44); -lean_ctor_set(x_71, 2, x_70); -x_72 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__15; -lean_inc(x_26); -x_73 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_73, 0, x_26); -lean_ctor_set(x_73, 1, x_72); -x_74 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__16; -x_75 = lean_name_mk_string(x_6, x_74); -x_76 = lean_array_push(x_57, x_63); -x_77 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_77, 0, x_23); -lean_ctor_set(x_77, 1, x_44); -lean_ctor_set(x_77, 2, x_76); -x_78 = lean_array_push(x_38, x_15); -x_79 = lean_array_push(x_78, x_77); -x_80 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_80, 0, x_23); -lean_ctor_set(x_80, 1, x_75); -lean_ctor_set(x_80, 2, x_79); -x_81 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__17; +x_48 = l_Lean_addMacroScope(x_29, x_47, x_28); +x_49 = lean_box(0); +x_50 = l_Lean_Elab_Command_expandMixfix___lambda__5___closed__3; +lean_inc(x_25); +x_51 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_51, 0, x_25); +lean_ctor_set(x_51, 1, x_50); +lean_ctor_set(x_51, 2, x_48); +lean_ctor_set(x_51, 3, x_49); +x_52 = lean_array_push(x_38, x_23); +x_53 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_53, 0, x_22); +lean_ctor_set(x_53, 1, x_5); +lean_ctor_set(x_53, 2, x_52); +x_54 = lean_array_push(x_41, x_53); +x_55 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_55, 0, x_22); +lean_ctor_set(x_55, 1, x_43); +lean_ctor_set(x_55, 2, x_54); +x_56 = lean_array_push(x_37, x_51); +lean_inc(x_56); +x_57 = lean_array_push(x_56, x_55); +lean_inc(x_46); +x_58 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_58, 0, x_22); +lean_ctor_set(x_58, 1, x_46); +lean_ctor_set(x_58, 2, x_57); +x_59 = l_Lean_Elab_Command_expandMixfix___lambda__5___closed__8; +x_60 = l_Lean_addMacroScope(x_29, x_59, x_28); +x_61 = l_Lean_Elab_Command_expandMixfix___lambda__5___closed__7; +lean_inc(x_25); +x_62 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_62, 0, x_25); +lean_ctor_set(x_62, 1, x_61); +lean_ctor_set(x_62, 2, x_60); +lean_ctor_set(x_62, 3, x_49); +lean_inc(x_62); +x_63 = lean_array_push(x_37, x_62); +lean_inc(x_44); +x_64 = lean_array_push(x_63, x_44); +x_65 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_65, 0, x_22); +lean_ctor_set(x_65, 1, x_46); +lean_ctor_set(x_65, 2, x_64); +x_66 = l_Lean_Elab_Command_expandMixfix___lambda__5___closed__9; +x_67 = lean_array_push(x_66, x_58); +x_68 = lean_array_push(x_67, x_13); +x_69 = lean_array_push(x_68, x_65); +x_70 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_70, 0, x_22); +lean_ctor_set(x_70, 1, x_43); +lean_ctor_set(x_70, 2, x_69); +x_71 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__15; +lean_inc(x_25); +x_72 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_72, 0, x_25); +lean_ctor_set(x_72, 1, x_71); +x_73 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__16; +x_74 = lean_name_mk_string(x_6, x_73); +x_75 = lean_array_push(x_56, x_62); +x_76 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_76, 0, x_22); +lean_ctor_set(x_76, 1, x_43); +lean_ctor_set(x_76, 2, x_75); +x_77 = lean_array_push(x_37, x_15); +x_78 = lean_array_push(x_77, x_76); +x_79 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_79, 0, x_22); +lean_ctor_set(x_79, 1, x_74); +lean_ctor_set(x_79, 2, x_78); +x_80 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__17; +x_81 = lean_array_push(x_80, x_33); x_82 = lean_array_push(x_81, x_34); -x_83 = lean_array_push(x_82, x_35); -x_84 = lean_array_push(x_83, x_45); +x_83 = lean_array_push(x_82, x_44); if (lean_obj_tag(x_7) == 0) { -lean_object* x_125; -x_125 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__4; -x_85 = x_125; -goto block_124; +lean_object* x_124; +x_124 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__4; +x_84 = x_124; +goto block_123; } else { -lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; -x_126 = lean_ctor_get(x_7, 0); -lean_inc(x_126); +lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; +x_125 = lean_ctor_get(x_7, 0); +lean_inc(x_125); lean_dec(x_7); -x_127 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__26; +x_126 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__26; lean_inc(x_3); -x_128 = lean_name_mk_string(x_3, x_127); -x_129 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__21; -lean_inc(x_26); -x_130 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_130, 0, x_26); -lean_ctor_set(x_130, 1, x_129); -x_131 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__27; -lean_inc(x_26); -x_132 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_132, 0, x_26); -lean_ctor_set(x_132, 1, x_131); -x_133 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__23; -lean_inc(x_26); -x_134 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_134, 0, x_26); -lean_ctor_set(x_134, 1, x_133); -x_135 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__24; -lean_inc(x_26); -x_136 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_136, 0, x_26); -lean_ctor_set(x_136, 1, x_135); -x_137 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__25; -x_138 = lean_array_push(x_137, x_130); -x_139 = lean_array_push(x_138, x_132); -x_140 = lean_array_push(x_139, x_134); -x_141 = lean_array_push(x_140, x_126); -x_142 = lean_array_push(x_141, x_136); -x_143 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_143, 0, x_23); -lean_ctor_set(x_143, 1, x_128); -lean_ctor_set(x_143, 2, x_142); -x_144 = lean_array_push(x_42, x_143); -x_85 = x_144; -goto block_124; -} -block_124: -{ -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_86 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__4; -x_87 = l_Array_append___rarg(x_86, x_85); -x_88 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_88, 0, x_23); -lean_ctor_set(x_88, 1, x_44); -lean_ctor_set(x_88, 2, x_87); -x_89 = lean_array_push(x_84, x_88); +x_127 = lean_name_mk_string(x_3, x_126); +x_128 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__21; +lean_inc(x_25); +x_129 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_129, 0, x_25); +lean_ctor_set(x_129, 1, x_128); +x_130 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__27; +lean_inc(x_25); +x_131 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_131, 0, x_25); +lean_ctor_set(x_131, 1, x_130); +x_132 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__23; +lean_inc(x_25); +x_133 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_133, 0, x_25); +lean_ctor_set(x_133, 1, x_132); +x_134 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__24; +lean_inc(x_25); +x_135 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_135, 0, x_25); +lean_ctor_set(x_135, 1, x_134); +x_136 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__25; +x_137 = lean_array_push(x_136, x_129); +x_138 = lean_array_push(x_137, x_131); +x_139 = lean_array_push(x_138, x_133); +x_140 = lean_array_push(x_139, x_125); +x_141 = lean_array_push(x_140, x_135); +x_142 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_142, 0, x_22); +lean_ctor_set(x_142, 1, x_127); +lean_ctor_set(x_142, 2, x_141); +x_143 = lean_array_push(x_41, x_142); +x_84 = x_143; +goto block_123; +} +block_123: +{ +lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_85 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__4; +x_86 = l_Array_append___rarg(x_85, x_84); +x_87 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_87, 0, x_22); +lean_ctor_set(x_87, 1, x_43); +lean_ctor_set(x_87, 2, x_86); +x_88 = lean_array_push(x_83, x_87); if (lean_obj_tag(x_9) == 0) { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; -lean_dec(x_26); +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; +lean_dec(x_25); lean_dec(x_3); -x_90 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__19; -x_91 = lean_array_push(x_89, x_90); -x_92 = lean_array_push(x_91, x_71); -x_93 = lean_array_push(x_92, x_73); -x_94 = lean_array_push(x_93, x_80); -x_95 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_95, 0, x_23); -lean_ctor_set(x_95, 1, x_32); -lean_ctor_set(x_95, 2, x_94); -if (lean_is_scalar(x_28)) { - x_96 = lean_alloc_ctor(0, 2, 0); +x_89 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__19; +x_90 = lean_array_push(x_88, x_89); +x_91 = lean_array_push(x_90, x_70); +x_92 = lean_array_push(x_91, x_72); +x_93 = lean_array_push(x_92, x_79); +x_94 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_94, 0, x_22); +lean_ctor_set(x_94, 1, x_31); +lean_ctor_set(x_94, 2, x_93); +if (lean_is_scalar(x_27)) { + x_95 = lean_alloc_ctor(0, 2, 0); } else { - x_96 = x_28; + x_95 = x_27; } -lean_ctor_set(x_96, 0, x_95); -lean_ctor_set(x_96, 1, x_27); -return x_96; +lean_ctor_set(x_95, 0, x_94); +lean_ctor_set(x_95, 1, x_26); +return x_95; } else { -lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; -x_97 = lean_ctor_get(x_9, 0); -lean_inc(x_97); +lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; +x_96 = lean_ctor_get(x_9, 0); +lean_inc(x_96); lean_dec(x_9); -x_98 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__20; -x_99 = lean_name_mk_string(x_3, x_98); -x_100 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__21; -lean_inc(x_26); -x_101 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_101, 0, x_26); -lean_ctor_set(x_101, 1, x_100); -x_102 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__22; -lean_inc(x_26); -x_103 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_103, 0, x_26); -lean_ctor_set(x_103, 1, x_102); -x_104 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__23; -lean_inc(x_26); -x_105 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_105, 0, x_26); -lean_ctor_set(x_105, 1, x_104); -x_106 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__24; -x_107 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_107, 0, x_26); -lean_ctor_set(x_107, 1, x_106); -x_108 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__25; -x_109 = lean_array_push(x_108, x_101); -x_110 = lean_array_push(x_109, x_103); -x_111 = lean_array_push(x_110, x_105); -x_112 = lean_array_push(x_111, x_97); -x_113 = lean_array_push(x_112, x_107); -x_114 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_114, 0, x_23); -lean_ctor_set(x_114, 1, x_99); -lean_ctor_set(x_114, 2, x_113); -x_115 = lean_array_push(x_42, x_114); -x_116 = l_Array_append___rarg(x_86, x_115); -x_117 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_117, 0, x_23); -lean_ctor_set(x_117, 1, x_44); -lean_ctor_set(x_117, 2, x_116); -x_118 = lean_array_push(x_89, x_117); -x_119 = lean_array_push(x_118, x_71); -x_120 = lean_array_push(x_119, x_73); -x_121 = lean_array_push(x_120, x_80); -x_122 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_122, 0, x_23); -lean_ctor_set(x_122, 1, x_32); -lean_ctor_set(x_122, 2, x_121); -if (lean_is_scalar(x_28)) { - x_123 = lean_alloc_ctor(0, 2, 0); +x_97 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__20; +x_98 = lean_name_mk_string(x_3, x_97); +x_99 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__21; +lean_inc(x_25); +x_100 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_100, 0, x_25); +lean_ctor_set(x_100, 1, x_99); +x_101 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__22; +lean_inc(x_25); +x_102 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_102, 0, x_25); +lean_ctor_set(x_102, 1, x_101); +x_103 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__23; +lean_inc(x_25); +x_104 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_104, 0, x_25); +lean_ctor_set(x_104, 1, x_103); +x_105 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__24; +x_106 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_106, 0, x_25); +lean_ctor_set(x_106, 1, x_105); +x_107 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__25; +x_108 = lean_array_push(x_107, x_100); +x_109 = lean_array_push(x_108, x_102); +x_110 = lean_array_push(x_109, x_104); +x_111 = lean_array_push(x_110, x_96); +x_112 = lean_array_push(x_111, x_106); +x_113 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_113, 0, x_22); +lean_ctor_set(x_113, 1, x_98); +lean_ctor_set(x_113, 2, x_112); +x_114 = lean_array_push(x_41, x_113); +x_115 = l_Array_append___rarg(x_85, x_114); +x_116 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_116, 0, x_22); +lean_ctor_set(x_116, 1, x_43); +lean_ctor_set(x_116, 2, x_115); +x_117 = lean_array_push(x_88, x_116); +x_118 = lean_array_push(x_117, x_70); +x_119 = lean_array_push(x_118, x_72); +x_120 = lean_array_push(x_119, x_79); +x_121 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_121, 0, x_22); +lean_ctor_set(x_121, 1, x_31); +lean_ctor_set(x_121, 2, x_120); +if (lean_is_scalar(x_27)) { + x_122 = lean_alloc_ctor(0, 2, 0); } else { - x_123 = x_28; + x_122 = x_27; } -lean_ctor_set(x_123, 0, x_122); -lean_ctor_set(x_123, 1, x_27); -return x_123; +lean_ctor_set(x_122, 0, x_121); +lean_ctor_set(x_122, 1, x_26); +return x_122; } } } else { -uint8_t x_145; +uint8_t x_144; lean_dec(x_15); lean_dec(x_13); lean_dec(x_10); @@ -1531,23 +1526,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_145 = !lean_is_exclusive(x_16); -if (x_145 == 0) +x_144 = !lean_is_exclusive(x_16); +if (x_144 == 0) { return x_16; } else { -lean_object* x_146; lean_object* x_147; lean_object* x_148; -x_146 = lean_ctor_get(x_16, 0); -x_147 = lean_ctor_get(x_16, 1); -lean_inc(x_147); +lean_object* x_145; lean_object* x_146; lean_object* x_147; +x_145 = lean_ctor_get(x_16, 0); +x_146 = lean_ctor_get(x_16, 1); lean_inc(x_146); +lean_inc(x_145); lean_dec(x_16); -x_148 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_148, 0, x_146); -lean_ctor_set(x_148, 1, x_147); -return x_148; +x_147 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_147, 0, x_145); +lean_ctor_set(x_147, 1, x_146); +return x_147; } } } @@ -1562,14 +1557,13 @@ x_12 = l_Lean_Syntax_getArg(x_1, x_11); x_13 = l_Lean_Syntax_isNone(x_12); if (x_13 == 0) { -lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_14 = l_Lean_nullKind; -x_15 = lean_unsigned_to_nat(1u); +lean_object* x_14; uint8_t x_15; +x_14 = lean_unsigned_to_nat(1u); lean_inc(x_12); -x_16 = l_Lean_Syntax_isNodeOf(x_12, x_14, x_15); -if (x_16 == 0) +x_15 = l_Lean_Syntax_matchesNull(x_12, x_14); +if (x_15 == 0) { -lean_object* x_17; lean_object* x_18; +lean_object* x_16; lean_object* x_17; lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); @@ -1578,28 +1572,28 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_17 = lean_box(1); -x_18 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_10); -return x_18; +x_16 = lean_box(1); +x_17 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_10); +return x_17; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_19 = lean_unsigned_to_nat(0u); -x_20 = l_Lean_Syntax_getArg(x_12, x_19); +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_18 = lean_unsigned_to_nat(0u); +x_19 = l_Lean_Syntax_getArg(x_12, x_18); lean_dec(x_12); -x_21 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__20; +x_20 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__20; lean_inc(x_3); -x_22 = lean_name_mk_string(x_3, x_21); -lean_inc(x_20); -x_23 = l_Lean_Syntax_isOfKind(x_20, x_22); -lean_dec(x_22); -if (x_23 == 0) +x_21 = lean_name_mk_string(x_3, x_20); +lean_inc(x_19); +x_22 = l_Lean_Syntax_isOfKind(x_19, x_21); +lean_dec(x_21); +if (x_22 == 0) { -lean_object* x_24; lean_object* x_25; -lean_dec(x_20); +lean_object* x_23; lean_object* x_24; +lean_dec(x_19); lean_dec(x_9); lean_dec(x_8); lean_dec(x_6); @@ -1607,34 +1601,34 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_24 = lean_box(1); -x_25 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_10); -return x_25; +x_23 = lean_box(1); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_10); +return x_24; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_26 = lean_unsigned_to_nat(3u); -x_27 = l_Lean_Syntax_getArg(x_20, x_26); -lean_dec(x_20); -x_28 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_28, 0, x_27); -x_29 = lean_box(0); -x_30 = l_Lean_Elab_Command_expandMixfix___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_8, x_29, x_28, x_9, x_10); -return x_30; +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_25 = lean_unsigned_to_nat(3u); +x_26 = l_Lean_Syntax_getArg(x_19, x_25); +lean_dec(x_19); +x_27 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_27, 0, x_26); +x_28 = lean_box(0); +x_29 = l_Lean_Elab_Command_expandMixfix___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_8, x_28, x_27, x_9, x_10); +return x_29; } } } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_dec(x_12); +x_30 = lean_box(0); x_31 = lean_box(0); -x_32 = lean_box(0); -x_33 = l_Lean_Elab_Command_expandMixfix___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_8, x_32, x_31, x_9, x_10); -return x_33; +x_32 = l_Lean_Elab_Command_expandMixfix___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_8, x_31, x_30, x_9, x_10); +return x_32; } } } @@ -1651,7 +1645,7 @@ lean_inc(x_2); x_16 = l_Lean_evalPrec(x_2, x_10, x_11); if (lean_obj_tag(x_16) == 0) { -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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); x_18 = lean_ctor_get(x_16, 1); @@ -1661,289 +1655,288 @@ x_19 = lean_unsigned_to_nat(1u); x_20 = lean_nat_add(x_17, x_19); lean_dec(x_17); x_21 = l_Nat_repr(x_20); -x_22 = l_Lean_numLitKind; -x_23 = lean_box(2); -x_24 = l_Lean_Syntax_mkLit(x_22, x_21, x_23); +x_22 = lean_box(2); +x_23 = l_Lean_Syntax_mkNumLit(x_21, x_22); lean_inc(x_10); -x_25 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_10, x_18); -x_26 = lean_ctor_get(x_25, 0); +x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_10, x_18); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 1); -lean_inc(x_27); -if (lean_is_exclusive(x_25)) { - lean_ctor_release(x_25, 0); - lean_ctor_release(x_25, 1); - x_28 = x_25; +if (lean_is_exclusive(x_24)) { + lean_ctor_release(x_24, 0); + lean_ctor_release(x_24, 1); + x_27 = x_24; } else { - lean_dec_ref(x_25); - x_28 = lean_box(0); + lean_dec_ref(x_24); + x_27 = lean_box(0); } -x_29 = lean_ctor_get(x_10, 2); +x_28 = lean_ctor_get(x_10, 2); +lean_inc(x_28); +x_29 = lean_ctor_get(x_10, 1); lean_inc(x_29); -x_30 = lean_ctor_get(x_10, 1); -lean_inc(x_30); lean_dec(x_10); -x_31 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__1; +x_30 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__1; lean_inc(x_3); -x_32 = lean_name_mk_string(x_3, x_31); -x_33 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__7; -x_34 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_34, 0, x_23); -lean_ctor_set(x_34, 1, x_4); -lean_ctor_set(x_34, 2, x_33); -lean_inc(x_26); -x_35 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_35, 0, x_26); -lean_ctor_set(x_35, 1, x_31); -x_36 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__8; -lean_inc(x_26); -x_37 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_37, 0, x_26); -lean_ctor_set(x_37, 1, x_36); -x_38 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__9; -x_39 = lean_array_push(x_38, x_37); -lean_inc(x_39); -x_40 = lean_array_push(x_39, x_2); +x_31 = lean_name_mk_string(x_3, x_30); +x_32 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__7; +x_33 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_33, 0, x_22); +lean_ctor_set(x_33, 1, x_4); +lean_ctor_set(x_33, 2, x_32); +lean_inc(x_25); +x_34 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_34, 0, x_25); +lean_ctor_set(x_34, 1, x_30); +x_35 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__8; +lean_inc(x_25); +x_36 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_36, 0, x_25); +lean_ctor_set(x_36, 1, x_35); +x_37 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__9; +x_38 = lean_array_push(x_37, x_36); +lean_inc(x_38); +x_39 = lean_array_push(x_38, x_2); lean_inc(x_5); -x_41 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_41, 0, x_23); -lean_ctor_set(x_41, 1, x_5); -lean_ctor_set(x_41, 2, x_40); -x_42 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__6; -x_43 = lean_array_push(x_42, x_41); -x_44 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__3; -x_45 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_45, 0, x_23); -lean_ctor_set(x_45, 1, x_44); -lean_ctor_set(x_45, 2, x_43); -x_46 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__10; +x_40 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_40, 0, x_22); +lean_ctor_set(x_40, 1, x_5); +lean_ctor_set(x_40, 2, x_39); +x_41 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__6; +x_42 = lean_array_push(x_41, x_40); +x_43 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__3; +x_44 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_44, 0, x_22); +lean_ctor_set(x_44, 1, x_43); +lean_ctor_set(x_44, 2, x_42); +x_45 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__10; lean_inc(x_3); -x_47 = lean_name_mk_string(x_3, x_46); -x_48 = l_Lean_Elab_Command_expandMixfix___lambda__5___closed__4; +x_46 = lean_name_mk_string(x_3, x_45); +x_47 = l_Lean_Elab_Command_expandMixfix___lambda__5___closed__4; +lean_inc(x_28); lean_inc(x_29); -lean_inc(x_30); -x_49 = l_Lean_addMacroScope(x_30, x_48, x_29); -x_50 = lean_box(0); -x_51 = l_Lean_Elab_Command_expandMixfix___lambda__5___closed__3; -lean_inc(x_26); -x_52 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_52, 0, x_26); -lean_ctor_set(x_52, 1, x_51); -lean_ctor_set(x_52, 2, x_49); -lean_ctor_set(x_52, 3, x_50); -x_53 = lean_array_push(x_39, x_24); -x_54 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_54, 0, x_23); -lean_ctor_set(x_54, 1, x_5); -lean_ctor_set(x_54, 2, x_53); -x_55 = lean_array_push(x_42, x_54); -x_56 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_56, 0, x_23); -lean_ctor_set(x_56, 1, x_44); -lean_ctor_set(x_56, 2, x_55); -x_57 = lean_array_push(x_38, x_52); +x_48 = l_Lean_addMacroScope(x_29, x_47, x_28); +x_49 = lean_box(0); +x_50 = l_Lean_Elab_Command_expandMixfix___lambda__5___closed__3; +lean_inc(x_25); +x_51 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_51, 0, x_25); +lean_ctor_set(x_51, 1, x_50); +lean_ctor_set(x_51, 2, x_48); +lean_ctor_set(x_51, 3, x_49); +x_52 = lean_array_push(x_38, x_23); +x_53 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_53, 0, x_22); +lean_ctor_set(x_53, 1, x_5); +lean_ctor_set(x_53, 2, x_52); +x_54 = lean_array_push(x_41, x_53); +x_55 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_55, 0, x_22); +lean_ctor_set(x_55, 1, x_43); +lean_ctor_set(x_55, 2, x_54); +x_56 = lean_array_push(x_37, x_51); +lean_inc(x_55); lean_inc(x_56); -lean_inc(x_57); -x_58 = lean_array_push(x_57, x_56); -lean_inc(x_47); -x_59 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_59, 0, x_23); -lean_ctor_set(x_59, 1, x_47); -lean_ctor_set(x_59, 2, x_58); -x_60 = l_Lean_Elab_Command_expandMixfix___lambda__5___closed__8; -x_61 = l_Lean_addMacroScope(x_30, x_60, x_29); -x_62 = l_Lean_Elab_Command_expandMixfix___lambda__5___closed__7; -lean_inc(x_26); -x_63 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_63, 0, x_26); -lean_ctor_set(x_63, 1, x_62); -lean_ctor_set(x_63, 2, x_61); -lean_ctor_set(x_63, 3, x_50); -lean_inc(x_63); -x_64 = lean_array_push(x_38, x_63); -x_65 = lean_array_push(x_64, x_56); -x_66 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_66, 0, x_23); -lean_ctor_set(x_66, 1, x_47); -lean_ctor_set(x_66, 2, x_65); -x_67 = l_Lean_Elab_Command_expandMixfix___lambda__5___closed__9; -x_68 = lean_array_push(x_67, x_59); -x_69 = lean_array_push(x_68, x_13); -x_70 = lean_array_push(x_69, x_66); -x_71 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_71, 0, x_23); -lean_ctor_set(x_71, 1, x_44); -lean_ctor_set(x_71, 2, x_70); -x_72 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__15; -lean_inc(x_26); -x_73 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_73, 0, x_26); -lean_ctor_set(x_73, 1, x_72); -x_74 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__16; -x_75 = lean_name_mk_string(x_6, x_74); -x_76 = lean_array_push(x_57, x_63); -x_77 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_77, 0, x_23); -lean_ctor_set(x_77, 1, x_44); -lean_ctor_set(x_77, 2, x_76); -x_78 = lean_array_push(x_38, x_15); -x_79 = lean_array_push(x_78, x_77); -x_80 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_80, 0, x_23); -lean_ctor_set(x_80, 1, x_75); -lean_ctor_set(x_80, 2, x_79); -x_81 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__17; +x_57 = lean_array_push(x_56, x_55); +lean_inc(x_46); +x_58 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_58, 0, x_22); +lean_ctor_set(x_58, 1, x_46); +lean_ctor_set(x_58, 2, x_57); +x_59 = l_Lean_Elab_Command_expandMixfix___lambda__5___closed__8; +x_60 = l_Lean_addMacroScope(x_29, x_59, x_28); +x_61 = l_Lean_Elab_Command_expandMixfix___lambda__5___closed__7; +lean_inc(x_25); +x_62 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_62, 0, x_25); +lean_ctor_set(x_62, 1, x_61); +lean_ctor_set(x_62, 2, x_60); +lean_ctor_set(x_62, 3, x_49); +lean_inc(x_62); +x_63 = lean_array_push(x_37, x_62); +x_64 = lean_array_push(x_63, x_55); +x_65 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_65, 0, x_22); +lean_ctor_set(x_65, 1, x_46); +lean_ctor_set(x_65, 2, x_64); +x_66 = l_Lean_Elab_Command_expandMixfix___lambda__5___closed__9; +x_67 = lean_array_push(x_66, x_58); +x_68 = lean_array_push(x_67, x_13); +x_69 = lean_array_push(x_68, x_65); +x_70 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_70, 0, x_22); +lean_ctor_set(x_70, 1, x_43); +lean_ctor_set(x_70, 2, x_69); +x_71 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__15; +lean_inc(x_25); +x_72 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_72, 0, x_25); +lean_ctor_set(x_72, 1, x_71); +x_73 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__16; +x_74 = lean_name_mk_string(x_6, x_73); +x_75 = lean_array_push(x_56, x_62); +x_76 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_76, 0, x_22); +lean_ctor_set(x_76, 1, x_43); +lean_ctor_set(x_76, 2, x_75); +x_77 = lean_array_push(x_37, x_15); +x_78 = lean_array_push(x_77, x_76); +x_79 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_79, 0, x_22); +lean_ctor_set(x_79, 1, x_74); +lean_ctor_set(x_79, 2, x_78); +x_80 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__17; +x_81 = lean_array_push(x_80, x_33); x_82 = lean_array_push(x_81, x_34); -x_83 = lean_array_push(x_82, x_35); -x_84 = lean_array_push(x_83, x_45); +x_83 = lean_array_push(x_82, x_44); if (lean_obj_tag(x_7) == 0) { -lean_object* x_125; -x_125 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__4; -x_85 = x_125; -goto block_124; +lean_object* x_124; +x_124 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__4; +x_84 = x_124; +goto block_123; } else { -lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; -x_126 = lean_ctor_get(x_7, 0); -lean_inc(x_126); +lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; +x_125 = lean_ctor_get(x_7, 0); +lean_inc(x_125); lean_dec(x_7); -x_127 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__26; +x_126 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__26; lean_inc(x_3); -x_128 = lean_name_mk_string(x_3, x_127); -x_129 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__21; -lean_inc(x_26); -x_130 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_130, 0, x_26); -lean_ctor_set(x_130, 1, x_129); -x_131 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__27; -lean_inc(x_26); -x_132 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_132, 0, x_26); -lean_ctor_set(x_132, 1, x_131); -x_133 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__23; -lean_inc(x_26); -x_134 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_134, 0, x_26); -lean_ctor_set(x_134, 1, x_133); -x_135 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__24; -lean_inc(x_26); -x_136 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_136, 0, x_26); -lean_ctor_set(x_136, 1, x_135); -x_137 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__25; -x_138 = lean_array_push(x_137, x_130); -x_139 = lean_array_push(x_138, x_132); -x_140 = lean_array_push(x_139, x_134); -x_141 = lean_array_push(x_140, x_126); -x_142 = lean_array_push(x_141, x_136); -x_143 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_143, 0, x_23); -lean_ctor_set(x_143, 1, x_128); -lean_ctor_set(x_143, 2, x_142); -x_144 = lean_array_push(x_42, x_143); -x_85 = x_144; -goto block_124; -} -block_124: -{ -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_86 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__4; -x_87 = l_Array_append___rarg(x_86, x_85); -x_88 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_88, 0, x_23); -lean_ctor_set(x_88, 1, x_44); -lean_ctor_set(x_88, 2, x_87); -x_89 = lean_array_push(x_84, x_88); +x_127 = lean_name_mk_string(x_3, x_126); +x_128 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__21; +lean_inc(x_25); +x_129 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_129, 0, x_25); +lean_ctor_set(x_129, 1, x_128); +x_130 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__27; +lean_inc(x_25); +x_131 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_131, 0, x_25); +lean_ctor_set(x_131, 1, x_130); +x_132 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__23; +lean_inc(x_25); +x_133 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_133, 0, x_25); +lean_ctor_set(x_133, 1, x_132); +x_134 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__24; +lean_inc(x_25); +x_135 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_135, 0, x_25); +lean_ctor_set(x_135, 1, x_134); +x_136 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__25; +x_137 = lean_array_push(x_136, x_129); +x_138 = lean_array_push(x_137, x_131); +x_139 = lean_array_push(x_138, x_133); +x_140 = lean_array_push(x_139, x_125); +x_141 = lean_array_push(x_140, x_135); +x_142 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_142, 0, x_22); +lean_ctor_set(x_142, 1, x_127); +lean_ctor_set(x_142, 2, x_141); +x_143 = lean_array_push(x_41, x_142); +x_84 = x_143; +goto block_123; +} +block_123: +{ +lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_85 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__4; +x_86 = l_Array_append___rarg(x_85, x_84); +x_87 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_87, 0, x_22); +lean_ctor_set(x_87, 1, x_43); +lean_ctor_set(x_87, 2, x_86); +x_88 = lean_array_push(x_83, x_87); if (lean_obj_tag(x_9) == 0) { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; -lean_dec(x_26); +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; +lean_dec(x_25); lean_dec(x_3); -x_90 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__19; -x_91 = lean_array_push(x_89, x_90); -x_92 = lean_array_push(x_91, x_71); -x_93 = lean_array_push(x_92, x_73); -x_94 = lean_array_push(x_93, x_80); -x_95 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_95, 0, x_23); -lean_ctor_set(x_95, 1, x_32); -lean_ctor_set(x_95, 2, x_94); -if (lean_is_scalar(x_28)) { - x_96 = lean_alloc_ctor(0, 2, 0); +x_89 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__19; +x_90 = lean_array_push(x_88, x_89); +x_91 = lean_array_push(x_90, x_70); +x_92 = lean_array_push(x_91, x_72); +x_93 = lean_array_push(x_92, x_79); +x_94 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_94, 0, x_22); +lean_ctor_set(x_94, 1, x_31); +lean_ctor_set(x_94, 2, x_93); +if (lean_is_scalar(x_27)) { + x_95 = lean_alloc_ctor(0, 2, 0); } else { - x_96 = x_28; + x_95 = x_27; } -lean_ctor_set(x_96, 0, x_95); -lean_ctor_set(x_96, 1, x_27); -return x_96; +lean_ctor_set(x_95, 0, x_94); +lean_ctor_set(x_95, 1, x_26); +return x_95; } else { -lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; -x_97 = lean_ctor_get(x_9, 0); -lean_inc(x_97); +lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; +x_96 = lean_ctor_get(x_9, 0); +lean_inc(x_96); lean_dec(x_9); -x_98 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__20; -x_99 = lean_name_mk_string(x_3, x_98); -x_100 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__21; -lean_inc(x_26); -x_101 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_101, 0, x_26); -lean_ctor_set(x_101, 1, x_100); -x_102 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__22; -lean_inc(x_26); -x_103 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_103, 0, x_26); -lean_ctor_set(x_103, 1, x_102); -x_104 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__23; -lean_inc(x_26); -x_105 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_105, 0, x_26); -lean_ctor_set(x_105, 1, x_104); -x_106 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__24; -x_107 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_107, 0, x_26); -lean_ctor_set(x_107, 1, x_106); -x_108 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__25; -x_109 = lean_array_push(x_108, x_101); -x_110 = lean_array_push(x_109, x_103); -x_111 = lean_array_push(x_110, x_105); -x_112 = lean_array_push(x_111, x_97); -x_113 = lean_array_push(x_112, x_107); -x_114 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_114, 0, x_23); -lean_ctor_set(x_114, 1, x_99); -lean_ctor_set(x_114, 2, x_113); -x_115 = lean_array_push(x_42, x_114); -x_116 = l_Array_append___rarg(x_86, x_115); -x_117 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_117, 0, x_23); -lean_ctor_set(x_117, 1, x_44); -lean_ctor_set(x_117, 2, x_116); -x_118 = lean_array_push(x_89, x_117); -x_119 = lean_array_push(x_118, x_71); -x_120 = lean_array_push(x_119, x_73); -x_121 = lean_array_push(x_120, x_80); -x_122 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_122, 0, x_23); -lean_ctor_set(x_122, 1, x_32); -lean_ctor_set(x_122, 2, x_121); -if (lean_is_scalar(x_28)) { - x_123 = lean_alloc_ctor(0, 2, 0); +x_97 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__20; +x_98 = lean_name_mk_string(x_3, x_97); +x_99 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__21; +lean_inc(x_25); +x_100 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_100, 0, x_25); +lean_ctor_set(x_100, 1, x_99); +x_101 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__22; +lean_inc(x_25); +x_102 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_102, 0, x_25); +lean_ctor_set(x_102, 1, x_101); +x_103 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__23; +lean_inc(x_25); +x_104 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_104, 0, x_25); +lean_ctor_set(x_104, 1, x_103); +x_105 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__24; +x_106 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_106, 0, x_25); +lean_ctor_set(x_106, 1, x_105); +x_107 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__25; +x_108 = lean_array_push(x_107, x_100); +x_109 = lean_array_push(x_108, x_102); +x_110 = lean_array_push(x_109, x_104); +x_111 = lean_array_push(x_110, x_96); +x_112 = lean_array_push(x_111, x_106); +x_113 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_113, 0, x_22); +lean_ctor_set(x_113, 1, x_98); +lean_ctor_set(x_113, 2, x_112); +x_114 = lean_array_push(x_41, x_113); +x_115 = l_Array_append___rarg(x_85, x_114); +x_116 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_116, 0, x_22); +lean_ctor_set(x_116, 1, x_43); +lean_ctor_set(x_116, 2, x_115); +x_117 = lean_array_push(x_88, x_116); +x_118 = lean_array_push(x_117, x_70); +x_119 = lean_array_push(x_118, x_72); +x_120 = lean_array_push(x_119, x_79); +x_121 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_121, 0, x_22); +lean_ctor_set(x_121, 1, x_31); +lean_ctor_set(x_121, 2, x_120); +if (lean_is_scalar(x_27)) { + x_122 = lean_alloc_ctor(0, 2, 0); } else { - x_123 = x_28; + x_122 = x_27; } -lean_ctor_set(x_123, 0, x_122); -lean_ctor_set(x_123, 1, x_27); -return x_123; +lean_ctor_set(x_122, 0, x_121); +lean_ctor_set(x_122, 1, x_26); +return x_122; } } } else { -uint8_t x_145; +uint8_t x_144; lean_dec(x_15); lean_dec(x_13); lean_dec(x_10); @@ -1954,23 +1947,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_145 = !lean_is_exclusive(x_16); -if (x_145 == 0) +x_144 = !lean_is_exclusive(x_16); +if (x_144 == 0) { return x_16; } else { -lean_object* x_146; lean_object* x_147; lean_object* x_148; -x_146 = lean_ctor_get(x_16, 0); -x_147 = lean_ctor_get(x_16, 1); -lean_inc(x_147); +lean_object* x_145; lean_object* x_146; lean_object* x_147; +x_145 = lean_ctor_get(x_16, 0); +x_146 = lean_ctor_get(x_16, 1); lean_inc(x_146); +lean_inc(x_145); lean_dec(x_16); -x_148 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_148, 0, x_146); -lean_ctor_set(x_148, 1, x_147); -return x_148; +x_147 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_147, 0, x_145); +lean_ctor_set(x_147, 1, x_146); +return x_147; } } } @@ -1985,14 +1978,13 @@ x_12 = l_Lean_Syntax_getArg(x_1, x_11); x_13 = l_Lean_Syntax_isNone(x_12); if (x_13 == 0) { -lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_14 = l_Lean_nullKind; -x_15 = lean_unsigned_to_nat(1u); +lean_object* x_14; uint8_t x_15; +x_14 = lean_unsigned_to_nat(1u); lean_inc(x_12); -x_16 = l_Lean_Syntax_isNodeOf(x_12, x_14, x_15); -if (x_16 == 0) +x_15 = l_Lean_Syntax_matchesNull(x_12, x_14); +if (x_15 == 0) { -lean_object* x_17; lean_object* x_18; +lean_object* x_16; lean_object* x_17; lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); @@ -2001,28 +1993,28 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_17 = lean_box(1); -x_18 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_10); -return x_18; +x_16 = lean_box(1); +x_17 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_10); +return x_17; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_19 = lean_unsigned_to_nat(0u); -x_20 = l_Lean_Syntax_getArg(x_12, x_19); +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_18 = lean_unsigned_to_nat(0u); +x_19 = l_Lean_Syntax_getArg(x_12, x_18); lean_dec(x_12); -x_21 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__20; +x_20 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__20; lean_inc(x_3); -x_22 = lean_name_mk_string(x_3, x_21); -lean_inc(x_20); -x_23 = l_Lean_Syntax_isOfKind(x_20, x_22); -lean_dec(x_22); -if (x_23 == 0) +x_21 = lean_name_mk_string(x_3, x_20); +lean_inc(x_19); +x_22 = l_Lean_Syntax_isOfKind(x_19, x_21); +lean_dec(x_21); +if (x_22 == 0) { -lean_object* x_24; lean_object* x_25; -lean_dec(x_20); +lean_object* x_23; lean_object* x_24; +lean_dec(x_19); lean_dec(x_9); lean_dec(x_8); lean_dec(x_6); @@ -2030,34 +2022,34 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_24 = lean_box(1); -x_25 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_10); -return x_25; +x_23 = lean_box(1); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_10); +return x_24; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_26 = lean_unsigned_to_nat(3u); -x_27 = l_Lean_Syntax_getArg(x_20, x_26); -lean_dec(x_20); -x_28 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_28, 0, x_27); -x_29 = lean_box(0); -x_30 = l_Lean_Elab_Command_expandMixfix___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_8, x_29, x_28, x_9, x_10); -return x_30; +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_25 = lean_unsigned_to_nat(3u); +x_26 = l_Lean_Syntax_getArg(x_19, x_25); +lean_dec(x_19); +x_27 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_27, 0, x_26); +x_28 = lean_box(0); +x_29 = l_Lean_Elab_Command_expandMixfix___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_8, x_28, x_27, x_9, x_10); +return x_29; } } } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_dec(x_12); +x_30 = lean_box(0); x_31 = lean_box(0); -x_32 = lean_box(0); -x_33 = l_Lean_Elab_Command_expandMixfix___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_8, x_32, x_31, x_9, x_10); -return x_33; +x_32 = l_Lean_Elab_Command_expandMixfix___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_8, x_31, x_30, x_9, x_10); +return x_32; } } } @@ -2074,7 +2066,7 @@ lean_inc(x_2); x_16 = l_Lean_evalPrec(x_2, x_10, x_11); if (lean_obj_tag(x_16) == 0) { -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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); x_18 = lean_ctor_get(x_16, 1); @@ -2084,289 +2076,288 @@ x_19 = lean_unsigned_to_nat(1u); x_20 = lean_nat_add(x_17, x_19); lean_dec(x_17); x_21 = l_Nat_repr(x_20); -x_22 = l_Lean_numLitKind; -x_23 = lean_box(2); -x_24 = l_Lean_Syntax_mkLit(x_22, x_21, x_23); +x_22 = lean_box(2); +x_23 = l_Lean_Syntax_mkNumLit(x_21, x_22); lean_inc(x_10); -x_25 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_10, x_18); -x_26 = lean_ctor_get(x_25, 0); +x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_10, x_18); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 1); -lean_inc(x_27); -if (lean_is_exclusive(x_25)) { - lean_ctor_release(x_25, 0); - lean_ctor_release(x_25, 1); - x_28 = x_25; +if (lean_is_exclusive(x_24)) { + lean_ctor_release(x_24, 0); + lean_ctor_release(x_24, 1); + x_27 = x_24; } else { - lean_dec_ref(x_25); - x_28 = lean_box(0); + lean_dec_ref(x_24); + x_27 = lean_box(0); } -x_29 = lean_ctor_get(x_10, 2); +x_28 = lean_ctor_get(x_10, 2); +lean_inc(x_28); +x_29 = lean_ctor_get(x_10, 1); lean_inc(x_29); -x_30 = lean_ctor_get(x_10, 1); -lean_inc(x_30); lean_dec(x_10); -x_31 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__1; +x_30 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__1; lean_inc(x_3); -x_32 = lean_name_mk_string(x_3, x_31); -x_33 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__7; -x_34 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_34, 0, x_23); -lean_ctor_set(x_34, 1, x_4); -lean_ctor_set(x_34, 2, x_33); -lean_inc(x_26); -x_35 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_35, 0, x_26); -lean_ctor_set(x_35, 1, x_31); -x_36 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__8; -lean_inc(x_26); -x_37 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_37, 0, x_26); -lean_ctor_set(x_37, 1, x_36); -x_38 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__9; -x_39 = lean_array_push(x_38, x_37); -lean_inc(x_39); -x_40 = lean_array_push(x_39, x_2); +x_31 = lean_name_mk_string(x_3, x_30); +x_32 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__7; +x_33 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_33, 0, x_22); +lean_ctor_set(x_33, 1, x_4); +lean_ctor_set(x_33, 2, x_32); +lean_inc(x_25); +x_34 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_34, 0, x_25); +lean_ctor_set(x_34, 1, x_30); +x_35 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__8; +lean_inc(x_25); +x_36 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_36, 0, x_25); +lean_ctor_set(x_36, 1, x_35); +x_37 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__9; +x_38 = lean_array_push(x_37, x_36); +lean_inc(x_38); +x_39 = lean_array_push(x_38, x_2); lean_inc(x_5); -x_41 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_41, 0, x_23); -lean_ctor_set(x_41, 1, x_5); -lean_ctor_set(x_41, 2, x_40); -x_42 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__6; -x_43 = lean_array_push(x_42, x_41); -x_44 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__3; -x_45 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_45, 0, x_23); -lean_ctor_set(x_45, 1, x_44); -lean_ctor_set(x_45, 2, x_43); -x_46 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__10; +x_40 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_40, 0, x_22); +lean_ctor_set(x_40, 1, x_5); +lean_ctor_set(x_40, 2, x_39); +x_41 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__6; +x_42 = lean_array_push(x_41, x_40); +x_43 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__3; +x_44 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_44, 0, x_22); +lean_ctor_set(x_44, 1, x_43); +lean_ctor_set(x_44, 2, x_42); +x_45 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__10; lean_inc(x_3); -x_47 = lean_name_mk_string(x_3, x_46); -x_48 = l_Lean_Elab_Command_expandMixfix___lambda__5___closed__4; +x_46 = lean_name_mk_string(x_3, x_45); +x_47 = l_Lean_Elab_Command_expandMixfix___lambda__5___closed__4; +lean_inc(x_28); lean_inc(x_29); -lean_inc(x_30); -x_49 = l_Lean_addMacroScope(x_30, x_48, x_29); -x_50 = lean_box(0); -x_51 = l_Lean_Elab_Command_expandMixfix___lambda__5___closed__3; -lean_inc(x_26); -x_52 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_52, 0, x_26); -lean_ctor_set(x_52, 1, x_51); -lean_ctor_set(x_52, 2, x_49); -lean_ctor_set(x_52, 3, x_50); -x_53 = lean_array_push(x_38, x_52); -lean_inc(x_45); -lean_inc(x_53); -x_54 = lean_array_push(x_53, x_45); -lean_inc(x_47); -x_55 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_55, 0, x_23); -lean_ctor_set(x_55, 1, x_47); -lean_ctor_set(x_55, 2, x_54); -x_56 = l_Lean_Elab_Command_expandMixfix___lambda__5___closed__8; -x_57 = l_Lean_addMacroScope(x_30, x_56, x_29); -x_58 = l_Lean_Elab_Command_expandMixfix___lambda__5___closed__7; -lean_inc(x_26); -x_59 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_59, 0, x_26); -lean_ctor_set(x_59, 1, x_58); -lean_ctor_set(x_59, 2, x_57); -lean_ctor_set(x_59, 3, x_50); -x_60 = lean_array_push(x_39, x_24); -x_61 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_61, 0, x_23); -lean_ctor_set(x_61, 1, x_5); -lean_ctor_set(x_61, 2, x_60); -x_62 = lean_array_push(x_42, x_61); -x_63 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_63, 0, x_23); -lean_ctor_set(x_63, 1, x_44); -lean_ctor_set(x_63, 2, x_62); -lean_inc(x_59); -x_64 = lean_array_push(x_38, x_59); -x_65 = lean_array_push(x_64, x_63); -x_66 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_66, 0, x_23); -lean_ctor_set(x_66, 1, x_47); -lean_ctor_set(x_66, 2, x_65); -x_67 = l_Lean_Elab_Command_expandMixfix___lambda__5___closed__9; -x_68 = lean_array_push(x_67, x_55); -x_69 = lean_array_push(x_68, x_13); -x_70 = lean_array_push(x_69, x_66); -x_71 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_71, 0, x_23); -lean_ctor_set(x_71, 1, x_44); -lean_ctor_set(x_71, 2, x_70); -x_72 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__15; -lean_inc(x_26); -x_73 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_73, 0, x_26); -lean_ctor_set(x_73, 1, x_72); -x_74 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__16; -x_75 = lean_name_mk_string(x_6, x_74); -x_76 = lean_array_push(x_53, x_59); -x_77 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_77, 0, x_23); -lean_ctor_set(x_77, 1, x_44); -lean_ctor_set(x_77, 2, x_76); -x_78 = lean_array_push(x_38, x_15); -x_79 = lean_array_push(x_78, x_77); -x_80 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_80, 0, x_23); -lean_ctor_set(x_80, 1, x_75); -lean_ctor_set(x_80, 2, x_79); -x_81 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__17; +x_48 = l_Lean_addMacroScope(x_29, x_47, x_28); +x_49 = lean_box(0); +x_50 = l_Lean_Elab_Command_expandMixfix___lambda__5___closed__3; +lean_inc(x_25); +x_51 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_51, 0, x_25); +lean_ctor_set(x_51, 1, x_50); +lean_ctor_set(x_51, 2, x_48); +lean_ctor_set(x_51, 3, x_49); +x_52 = lean_array_push(x_37, x_51); +lean_inc(x_44); +lean_inc(x_52); +x_53 = lean_array_push(x_52, x_44); +lean_inc(x_46); +x_54 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_54, 0, x_22); +lean_ctor_set(x_54, 1, x_46); +lean_ctor_set(x_54, 2, x_53); +x_55 = l_Lean_Elab_Command_expandMixfix___lambda__5___closed__8; +x_56 = l_Lean_addMacroScope(x_29, x_55, x_28); +x_57 = l_Lean_Elab_Command_expandMixfix___lambda__5___closed__7; +lean_inc(x_25); +x_58 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_58, 0, x_25); +lean_ctor_set(x_58, 1, x_57); +lean_ctor_set(x_58, 2, x_56); +lean_ctor_set(x_58, 3, x_49); +x_59 = lean_array_push(x_38, x_23); +x_60 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_60, 0, x_22); +lean_ctor_set(x_60, 1, x_5); +lean_ctor_set(x_60, 2, x_59); +x_61 = lean_array_push(x_41, x_60); +x_62 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_62, 0, x_22); +lean_ctor_set(x_62, 1, x_43); +lean_ctor_set(x_62, 2, x_61); +lean_inc(x_58); +x_63 = lean_array_push(x_37, x_58); +x_64 = lean_array_push(x_63, x_62); +x_65 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_65, 0, x_22); +lean_ctor_set(x_65, 1, x_46); +lean_ctor_set(x_65, 2, x_64); +x_66 = l_Lean_Elab_Command_expandMixfix___lambda__5___closed__9; +x_67 = lean_array_push(x_66, x_54); +x_68 = lean_array_push(x_67, x_13); +x_69 = lean_array_push(x_68, x_65); +x_70 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_70, 0, x_22); +lean_ctor_set(x_70, 1, x_43); +lean_ctor_set(x_70, 2, x_69); +x_71 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__15; +lean_inc(x_25); +x_72 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_72, 0, x_25); +lean_ctor_set(x_72, 1, x_71); +x_73 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__16; +x_74 = lean_name_mk_string(x_6, x_73); +x_75 = lean_array_push(x_52, x_58); +x_76 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_76, 0, x_22); +lean_ctor_set(x_76, 1, x_43); +lean_ctor_set(x_76, 2, x_75); +x_77 = lean_array_push(x_37, x_15); +x_78 = lean_array_push(x_77, x_76); +x_79 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_79, 0, x_22); +lean_ctor_set(x_79, 1, x_74); +lean_ctor_set(x_79, 2, x_78); +x_80 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__17; +x_81 = lean_array_push(x_80, x_33); x_82 = lean_array_push(x_81, x_34); -x_83 = lean_array_push(x_82, x_35); -x_84 = lean_array_push(x_83, x_45); +x_83 = lean_array_push(x_82, x_44); if (lean_obj_tag(x_7) == 0) { -lean_object* x_125; -x_125 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__4; -x_85 = x_125; -goto block_124; +lean_object* x_124; +x_124 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__4; +x_84 = x_124; +goto block_123; } else { -lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; -x_126 = lean_ctor_get(x_7, 0); -lean_inc(x_126); +lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; +x_125 = lean_ctor_get(x_7, 0); +lean_inc(x_125); lean_dec(x_7); -x_127 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__26; +x_126 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__26; lean_inc(x_3); -x_128 = lean_name_mk_string(x_3, x_127); -x_129 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__21; -lean_inc(x_26); -x_130 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_130, 0, x_26); -lean_ctor_set(x_130, 1, x_129); -x_131 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__27; -lean_inc(x_26); -x_132 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_132, 0, x_26); -lean_ctor_set(x_132, 1, x_131); -x_133 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__23; -lean_inc(x_26); -x_134 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_134, 0, x_26); -lean_ctor_set(x_134, 1, x_133); -x_135 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__24; -lean_inc(x_26); -x_136 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_136, 0, x_26); -lean_ctor_set(x_136, 1, x_135); -x_137 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__25; -x_138 = lean_array_push(x_137, x_130); -x_139 = lean_array_push(x_138, x_132); -x_140 = lean_array_push(x_139, x_134); -x_141 = lean_array_push(x_140, x_126); -x_142 = lean_array_push(x_141, x_136); -x_143 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_143, 0, x_23); -lean_ctor_set(x_143, 1, x_128); -lean_ctor_set(x_143, 2, x_142); -x_144 = lean_array_push(x_42, x_143); -x_85 = x_144; -goto block_124; -} -block_124: -{ -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_86 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__4; -x_87 = l_Array_append___rarg(x_86, x_85); -x_88 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_88, 0, x_23); -lean_ctor_set(x_88, 1, x_44); -lean_ctor_set(x_88, 2, x_87); -x_89 = lean_array_push(x_84, x_88); +x_127 = lean_name_mk_string(x_3, x_126); +x_128 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__21; +lean_inc(x_25); +x_129 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_129, 0, x_25); +lean_ctor_set(x_129, 1, x_128); +x_130 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__27; +lean_inc(x_25); +x_131 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_131, 0, x_25); +lean_ctor_set(x_131, 1, x_130); +x_132 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__23; +lean_inc(x_25); +x_133 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_133, 0, x_25); +lean_ctor_set(x_133, 1, x_132); +x_134 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__24; +lean_inc(x_25); +x_135 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_135, 0, x_25); +lean_ctor_set(x_135, 1, x_134); +x_136 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__25; +x_137 = lean_array_push(x_136, x_129); +x_138 = lean_array_push(x_137, x_131); +x_139 = lean_array_push(x_138, x_133); +x_140 = lean_array_push(x_139, x_125); +x_141 = lean_array_push(x_140, x_135); +x_142 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_142, 0, x_22); +lean_ctor_set(x_142, 1, x_127); +lean_ctor_set(x_142, 2, x_141); +x_143 = lean_array_push(x_41, x_142); +x_84 = x_143; +goto block_123; +} +block_123: +{ +lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_85 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__4; +x_86 = l_Array_append___rarg(x_85, x_84); +x_87 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_87, 0, x_22); +lean_ctor_set(x_87, 1, x_43); +lean_ctor_set(x_87, 2, x_86); +x_88 = lean_array_push(x_83, x_87); if (lean_obj_tag(x_9) == 0) { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; -lean_dec(x_26); +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; +lean_dec(x_25); lean_dec(x_3); -x_90 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__19; -x_91 = lean_array_push(x_89, x_90); -x_92 = lean_array_push(x_91, x_71); -x_93 = lean_array_push(x_92, x_73); -x_94 = lean_array_push(x_93, x_80); -x_95 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_95, 0, x_23); -lean_ctor_set(x_95, 1, x_32); -lean_ctor_set(x_95, 2, x_94); -if (lean_is_scalar(x_28)) { - x_96 = lean_alloc_ctor(0, 2, 0); +x_89 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__19; +x_90 = lean_array_push(x_88, x_89); +x_91 = lean_array_push(x_90, x_70); +x_92 = lean_array_push(x_91, x_72); +x_93 = lean_array_push(x_92, x_79); +x_94 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_94, 0, x_22); +lean_ctor_set(x_94, 1, x_31); +lean_ctor_set(x_94, 2, x_93); +if (lean_is_scalar(x_27)) { + x_95 = lean_alloc_ctor(0, 2, 0); } else { - x_96 = x_28; + x_95 = x_27; } -lean_ctor_set(x_96, 0, x_95); -lean_ctor_set(x_96, 1, x_27); -return x_96; +lean_ctor_set(x_95, 0, x_94); +lean_ctor_set(x_95, 1, x_26); +return x_95; } else { -lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; -x_97 = lean_ctor_get(x_9, 0); -lean_inc(x_97); +lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; +x_96 = lean_ctor_get(x_9, 0); +lean_inc(x_96); lean_dec(x_9); -x_98 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__20; -x_99 = lean_name_mk_string(x_3, x_98); -x_100 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__21; -lean_inc(x_26); -x_101 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_101, 0, x_26); -lean_ctor_set(x_101, 1, x_100); -x_102 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__22; -lean_inc(x_26); -x_103 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_103, 0, x_26); -lean_ctor_set(x_103, 1, x_102); -x_104 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__23; -lean_inc(x_26); -x_105 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_105, 0, x_26); -lean_ctor_set(x_105, 1, x_104); -x_106 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__24; -x_107 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_107, 0, x_26); -lean_ctor_set(x_107, 1, x_106); -x_108 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__25; -x_109 = lean_array_push(x_108, x_101); -x_110 = lean_array_push(x_109, x_103); -x_111 = lean_array_push(x_110, x_105); -x_112 = lean_array_push(x_111, x_97); -x_113 = lean_array_push(x_112, x_107); -x_114 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_114, 0, x_23); -lean_ctor_set(x_114, 1, x_99); -lean_ctor_set(x_114, 2, x_113); -x_115 = lean_array_push(x_42, x_114); -x_116 = l_Array_append___rarg(x_86, x_115); -x_117 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_117, 0, x_23); -lean_ctor_set(x_117, 1, x_44); -lean_ctor_set(x_117, 2, x_116); -x_118 = lean_array_push(x_89, x_117); -x_119 = lean_array_push(x_118, x_71); -x_120 = lean_array_push(x_119, x_73); -x_121 = lean_array_push(x_120, x_80); -x_122 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_122, 0, x_23); -lean_ctor_set(x_122, 1, x_32); -lean_ctor_set(x_122, 2, x_121); -if (lean_is_scalar(x_28)) { - x_123 = lean_alloc_ctor(0, 2, 0); +x_97 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__20; +x_98 = lean_name_mk_string(x_3, x_97); +x_99 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__21; +lean_inc(x_25); +x_100 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_100, 0, x_25); +lean_ctor_set(x_100, 1, x_99); +x_101 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__22; +lean_inc(x_25); +x_102 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_102, 0, x_25); +lean_ctor_set(x_102, 1, x_101); +x_103 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__23; +lean_inc(x_25); +x_104 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_104, 0, x_25); +lean_ctor_set(x_104, 1, x_103); +x_105 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__24; +x_106 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_106, 0, x_25); +lean_ctor_set(x_106, 1, x_105); +x_107 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__25; +x_108 = lean_array_push(x_107, x_100); +x_109 = lean_array_push(x_108, x_102); +x_110 = lean_array_push(x_109, x_104); +x_111 = lean_array_push(x_110, x_96); +x_112 = lean_array_push(x_111, x_106); +x_113 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_113, 0, x_22); +lean_ctor_set(x_113, 1, x_98); +lean_ctor_set(x_113, 2, x_112); +x_114 = lean_array_push(x_41, x_113); +x_115 = l_Array_append___rarg(x_85, x_114); +x_116 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_116, 0, x_22); +lean_ctor_set(x_116, 1, x_43); +lean_ctor_set(x_116, 2, x_115); +x_117 = lean_array_push(x_88, x_116); +x_118 = lean_array_push(x_117, x_70); +x_119 = lean_array_push(x_118, x_72); +x_120 = lean_array_push(x_119, x_79); +x_121 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_121, 0, x_22); +lean_ctor_set(x_121, 1, x_31); +lean_ctor_set(x_121, 2, x_120); +if (lean_is_scalar(x_27)) { + x_122 = lean_alloc_ctor(0, 2, 0); } else { - x_123 = x_28; + x_122 = x_27; } -lean_ctor_set(x_123, 0, x_122); -lean_ctor_set(x_123, 1, x_27); -return x_123; +lean_ctor_set(x_122, 0, x_121); +lean_ctor_set(x_122, 1, x_26); +return x_122; } } } else { -uint8_t x_145; +uint8_t x_144; lean_dec(x_15); lean_dec(x_13); lean_dec(x_10); @@ -2377,23 +2368,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_145 = !lean_is_exclusive(x_16); -if (x_145 == 0) +x_144 = !lean_is_exclusive(x_16); +if (x_144 == 0) { return x_16; } else { -lean_object* x_146; lean_object* x_147; lean_object* x_148; -x_146 = lean_ctor_get(x_16, 0); -x_147 = lean_ctor_get(x_16, 1); -lean_inc(x_147); +lean_object* x_145; lean_object* x_146; lean_object* x_147; +x_145 = lean_ctor_get(x_16, 0); +x_146 = lean_ctor_get(x_16, 1); lean_inc(x_146); +lean_inc(x_145); lean_dec(x_16); -x_148 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_148, 0, x_146); -lean_ctor_set(x_148, 1, x_147); -return x_148; +x_147 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_147, 0, x_145); +lean_ctor_set(x_147, 1, x_146); +return x_147; } } } @@ -2408,14 +2399,13 @@ x_12 = l_Lean_Syntax_getArg(x_1, x_11); x_13 = l_Lean_Syntax_isNone(x_12); if (x_13 == 0) { -lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_14 = l_Lean_nullKind; -x_15 = lean_unsigned_to_nat(1u); +lean_object* x_14; uint8_t x_15; +x_14 = lean_unsigned_to_nat(1u); lean_inc(x_12); -x_16 = l_Lean_Syntax_isNodeOf(x_12, x_14, x_15); -if (x_16 == 0) +x_15 = l_Lean_Syntax_matchesNull(x_12, x_14); +if (x_15 == 0) { -lean_object* x_17; lean_object* x_18; +lean_object* x_16; lean_object* x_17; lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); @@ -2424,28 +2414,28 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_17 = lean_box(1); -x_18 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_10); -return x_18; +x_16 = lean_box(1); +x_17 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_10); +return x_17; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_19 = lean_unsigned_to_nat(0u); -x_20 = l_Lean_Syntax_getArg(x_12, x_19); +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_18 = lean_unsigned_to_nat(0u); +x_19 = l_Lean_Syntax_getArg(x_12, x_18); lean_dec(x_12); -x_21 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__20; +x_20 = l_Lean_Elab_Command_expandMixfix___lambda__1___closed__20; lean_inc(x_3); -x_22 = lean_name_mk_string(x_3, x_21); -lean_inc(x_20); -x_23 = l_Lean_Syntax_isOfKind(x_20, x_22); -lean_dec(x_22); -if (x_23 == 0) +x_21 = lean_name_mk_string(x_3, x_20); +lean_inc(x_19); +x_22 = l_Lean_Syntax_isOfKind(x_19, x_21); +lean_dec(x_21); +if (x_22 == 0) { -lean_object* x_24; lean_object* x_25; -lean_dec(x_20); +lean_object* x_23; lean_object* x_24; +lean_dec(x_19); lean_dec(x_9); lean_dec(x_8); lean_dec(x_6); @@ -2453,34 +2443,34 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_24 = lean_box(1); -x_25 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_10); -return x_25; +x_23 = lean_box(1); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_10); +return x_24; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_26 = lean_unsigned_to_nat(3u); -x_27 = l_Lean_Syntax_getArg(x_20, x_26); -lean_dec(x_20); -x_28 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_28, 0, x_27); -x_29 = lean_box(0); -x_30 = l_Lean_Elab_Command_expandMixfix___lambda__9(x_1, x_2, x_3, x_4, x_5, x_6, x_8, x_29, x_28, x_9, x_10); -return x_30; +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_25 = lean_unsigned_to_nat(3u); +x_26 = l_Lean_Syntax_getArg(x_19, x_25); +lean_dec(x_19); +x_27 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_27, 0, x_26); +x_28 = lean_box(0); +x_29 = l_Lean_Elab_Command_expandMixfix___lambda__9(x_1, x_2, x_3, x_4, x_5, x_6, x_8, x_28, x_27, x_9, x_10); +return x_29; } } } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_dec(x_12); +x_30 = lean_box(0); x_31 = lean_box(0); -x_32 = lean_box(0); -x_33 = l_Lean_Elab_Command_expandMixfix___lambda__9(x_1, x_2, x_3, x_4, x_5, x_6, x_8, x_32, x_31, x_9, x_10); -return x_33; +x_32 = l_Lean_Elab_Command_expandMixfix___lambda__9(x_1, x_2, x_3, x_4, x_5, x_6, x_8, x_31, x_30, x_9, x_10); +return x_32; } } } @@ -2750,554 +2740,553 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; uint8_t x_16; +lean_object* x_14; uint8_t x_15; x_14 = l_Lean_Syntax_getArg(x_9, x_8); lean_dec(x_9); -x_15 = l_Lean_nullKind; -x_16 = l_Lean_Syntax_isNodeOf(x_14, x_15, x_8); -if (x_16 == 0) +x_15 = l_Lean_Syntax_matchesNull(x_14, x_8); +if (x_15 == 0) { -lean_object* x_17; lean_object* x_18; +lean_object* x_16; lean_object* x_17; lean_dec(x_2); lean_dec(x_1); -x_17 = lean_box(1); -x_18 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_3); -return x_18; +x_16 = lean_box(1); +x_17 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_3); +return x_17; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_19 = lean_unsigned_to_nat(1u); -x_20 = l_Lean_Syntax_getArg(x_1, x_19); -x_21 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__14; -lean_inc(x_20); -x_22 = l_Lean_Syntax_isOfKind(x_20, x_21); -if (x_22 == 0) +lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_18 = lean_unsigned_to_nat(1u); +x_19 = l_Lean_Syntax_getArg(x_1, x_18); +x_20 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__14; +lean_inc(x_19); +x_21 = l_Lean_Syntax_isOfKind(x_19, x_20); +if (x_21 == 0) { -lean_object* x_23; uint8_t x_24; -x_23 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__16; -lean_inc(x_20); -x_24 = l_Lean_Syntax_isOfKind(x_20, x_23); -if (x_24 == 0) +lean_object* x_22; uint8_t x_23; +x_22 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__16; +lean_inc(x_19); +x_23 = l_Lean_Syntax_isOfKind(x_19, x_22); +if (x_23 == 0) { -lean_object* x_25; uint8_t x_26; -x_25 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__18; -lean_inc(x_20); -x_26 = l_Lean_Syntax_isOfKind(x_20, x_25); -if (x_26 == 0) +lean_object* x_24; uint8_t x_25; +x_24 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__18; +lean_inc(x_19); +x_25 = l_Lean_Syntax_isOfKind(x_19, x_24); +if (x_25 == 0) { -lean_object* x_27; uint8_t x_28; -x_27 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__20; -lean_inc(x_20); -x_28 = l_Lean_Syntax_isOfKind(x_20, x_27); -if (x_28 == 0) +lean_object* x_26; uint8_t x_27; +x_26 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__20; +lean_inc(x_19); +x_27 = l_Lean_Syntax_isOfKind(x_19, x_26); +if (x_27 == 0) { -lean_object* x_29; uint8_t x_30; -x_29 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__22; -x_30 = l_Lean_Syntax_isOfKind(x_20, x_29); -if (x_30 == 0) +lean_object* x_28; uint8_t x_29; +x_28 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__22; +x_29 = l_Lean_Syntax_isOfKind(x_19, x_28); +if (x_29 == 0) { -lean_object* x_31; lean_object* x_32; +lean_object* x_30; lean_object* x_31; lean_dec(x_2); lean_dec(x_1); -x_31 = lean_box(1); -x_32 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_3); -return x_32; +x_30 = lean_box(1); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_3); +return x_31; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; -x_33 = lean_unsigned_to_nat(2u); -x_34 = l_Lean_Syntax_getArg(x_1, x_33); -x_35 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__24; -lean_inc(x_34); -x_36 = l_Lean_Syntax_isOfKind(x_34, x_35); -if (x_36 == 0) +lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_32 = lean_unsigned_to_nat(2u); +x_33 = l_Lean_Syntax_getArg(x_1, x_32); +x_34 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__24; +lean_inc(x_33); +x_35 = l_Lean_Syntax_isOfKind(x_33, x_34); +if (x_35 == 0) { -lean_object* x_37; lean_object* x_38; -lean_dec(x_34); +lean_object* x_36; lean_object* x_37; +lean_dec(x_33); lean_dec(x_2); lean_dec(x_1); -x_37 = lean_box(1); -x_38 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_38, 0, x_37); -lean_ctor_set(x_38, 1, x_3); -return x_38; +x_36 = lean_box(1); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_3); +return x_37; } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; -x_39 = l_Lean_Syntax_getArg(x_34, x_19); -lean_dec(x_34); -x_40 = lean_unsigned_to_nat(3u); -x_41 = l_Lean_Syntax_getArg(x_1, x_40); -x_42 = l_Lean_Syntax_isNone(x_41); +lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_38 = l_Lean_Syntax_getArg(x_33, x_18); +lean_dec(x_33); +x_39 = lean_unsigned_to_nat(3u); +x_40 = l_Lean_Syntax_getArg(x_1, x_39); +x_41 = l_Lean_Syntax_isNone(x_40); +if (x_41 == 0) +{ +uint8_t x_42; +lean_inc(x_40); +x_42 = l_Lean_Syntax_matchesNull(x_40, x_18); if (x_42 == 0) { -uint8_t x_43; -lean_inc(x_41); -x_43 = l_Lean_Syntax_isNodeOf(x_41, x_15, x_19); -if (x_43 == 0) -{ -lean_object* x_44; lean_object* x_45; -lean_dec(x_41); -lean_dec(x_39); +lean_object* x_43; lean_object* x_44; +lean_dec(x_40); +lean_dec(x_38); lean_dec(x_2); lean_dec(x_1); -x_44 = lean_box(1); -x_45 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_45, 0, x_44); -lean_ctor_set(x_45, 1, x_3); -return x_45; +x_43 = lean_box(1); +x_44 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_3); +return x_44; } else { -lean_object* x_46; lean_object* x_47; uint8_t x_48; -x_46 = l_Lean_Syntax_getArg(x_41, x_8); -lean_dec(x_41); -x_47 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__25; -lean_inc(x_46); -x_48 = l_Lean_Syntax_isOfKind(x_46, x_47); -if (x_48 == 0) +lean_object* x_45; lean_object* x_46; uint8_t x_47; +x_45 = l_Lean_Syntax_getArg(x_40, x_8); +lean_dec(x_40); +x_46 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__25; +lean_inc(x_45); +x_47 = l_Lean_Syntax_isOfKind(x_45, x_46); +if (x_47 == 0) { -lean_object* x_49; lean_object* x_50; -lean_dec(x_46); -lean_dec(x_39); +lean_object* x_48; lean_object* x_49; +lean_dec(x_45); +lean_dec(x_38); lean_dec(x_2); lean_dec(x_1); -x_49 = lean_box(1); -x_50 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_50, 0, x_49); -lean_ctor_set(x_50, 1, x_3); -return x_50; +x_48 = lean_box(1); +x_49 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_3); +return x_49; } else { -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_51 = l_Lean_Syntax_getArg(x_46, x_40); -lean_dec(x_46); -x_52 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_52, 0, x_51); -x_53 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__6; -x_54 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__10; -x_55 = lean_box(0); -x_56 = l_Lean_Elab_Command_expandMixfix___lambda__2(x_1, x_53, x_10, x_39, x_35, x_54, x_55, x_52, x_2, x_3); +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_50 = l_Lean_Syntax_getArg(x_45, x_39); +lean_dec(x_45); +x_51 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_51, 0, x_50); +x_52 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__6; +x_53 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__10; +x_54 = lean_box(0); +x_55 = l_Lean_Elab_Command_expandMixfix___lambda__2(x_1, x_52, x_10, x_38, x_34, x_53, x_54, x_51, x_2, x_3); lean_dec(x_1); -return x_56; +return x_55; } } } else { -lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; -lean_dec(x_41); -x_57 = lean_box(0); -x_58 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__6; -x_59 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__10; -x_60 = lean_box(0); -x_61 = l_Lean_Elab_Command_expandMixfix___lambda__2(x_1, x_58, x_10, x_39, x_35, x_59, x_60, x_57, x_2, x_3); +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +lean_dec(x_40); +x_56 = lean_box(0); +x_57 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__6; +x_58 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__10; +x_59 = lean_box(0); +x_60 = l_Lean_Elab_Command_expandMixfix___lambda__2(x_1, x_57, x_10, x_38, x_34, x_58, x_59, x_56, x_2, x_3); lean_dec(x_1); -return x_61; +return x_60; } } } } else { -lean_object* x_62; lean_object* x_63; lean_object* x_64; uint8_t x_65; -lean_dec(x_20); -x_62 = lean_unsigned_to_nat(2u); -x_63 = l_Lean_Syntax_getArg(x_1, x_62); -x_64 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__24; -lean_inc(x_63); -x_65 = l_Lean_Syntax_isOfKind(x_63, x_64); -if (x_65 == 0) -{ -lean_object* x_66; lean_object* x_67; -lean_dec(x_63); +lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; +lean_dec(x_19); +x_61 = lean_unsigned_to_nat(2u); +x_62 = l_Lean_Syntax_getArg(x_1, x_61); +x_63 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__24; +lean_inc(x_62); +x_64 = l_Lean_Syntax_isOfKind(x_62, x_63); +if (x_64 == 0) +{ +lean_object* x_65; lean_object* x_66; +lean_dec(x_62); lean_dec(x_2); lean_dec(x_1); -x_66 = lean_box(1); -x_67 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_67, 0, x_66); -lean_ctor_set(x_67, 1, x_3); -return x_67; +x_65 = lean_box(1); +x_66 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_66, 0, x_65); +lean_ctor_set(x_66, 1, x_3); +return x_66; } else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; -x_68 = l_Lean_Syntax_getArg(x_63, x_19); -lean_dec(x_63); -x_69 = lean_unsigned_to_nat(3u); -x_70 = l_Lean_Syntax_getArg(x_1, x_69); -x_71 = l_Lean_Syntax_isNone(x_70); +lean_object* x_67; lean_object* x_68; lean_object* x_69; uint8_t x_70; +x_67 = l_Lean_Syntax_getArg(x_62, x_18); +lean_dec(x_62); +x_68 = lean_unsigned_to_nat(3u); +x_69 = l_Lean_Syntax_getArg(x_1, x_68); +x_70 = l_Lean_Syntax_isNone(x_69); +if (x_70 == 0) +{ +uint8_t x_71; +lean_inc(x_69); +x_71 = l_Lean_Syntax_matchesNull(x_69, x_18); if (x_71 == 0) { -uint8_t x_72; -lean_inc(x_70); -x_72 = l_Lean_Syntax_isNodeOf(x_70, x_15, x_19); -if (x_72 == 0) -{ -lean_object* x_73; lean_object* x_74; -lean_dec(x_70); -lean_dec(x_68); +lean_object* x_72; lean_object* x_73; +lean_dec(x_69); +lean_dec(x_67); lean_dec(x_2); lean_dec(x_1); -x_73 = lean_box(1); -x_74 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_74, 0, x_73); -lean_ctor_set(x_74, 1, x_3); -return x_74; +x_72 = lean_box(1); +x_73 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_73, 0, x_72); +lean_ctor_set(x_73, 1, x_3); +return x_73; } else { -lean_object* x_75; lean_object* x_76; uint8_t x_77; -x_75 = l_Lean_Syntax_getArg(x_70, x_8); -lean_dec(x_70); -x_76 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__25; -lean_inc(x_75); -x_77 = l_Lean_Syntax_isOfKind(x_75, x_76); -if (x_77 == 0) -{ -lean_object* x_78; lean_object* x_79; -lean_dec(x_75); -lean_dec(x_68); +lean_object* x_74; lean_object* x_75; uint8_t x_76; +x_74 = l_Lean_Syntax_getArg(x_69, x_8); +lean_dec(x_69); +x_75 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__25; +lean_inc(x_74); +x_76 = l_Lean_Syntax_isOfKind(x_74, x_75); +if (x_76 == 0) +{ +lean_object* x_77; lean_object* x_78; +lean_dec(x_74); +lean_dec(x_67); lean_dec(x_2); lean_dec(x_1); -x_78 = lean_box(1); -x_79 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_79, 0, x_78); -lean_ctor_set(x_79, 1, x_3); -return x_79; +x_77 = lean_box(1); +x_78 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_78, 0, x_77); +lean_ctor_set(x_78, 1, x_3); +return x_78; } else { -lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_80 = l_Lean_Syntax_getArg(x_75, x_69); -lean_dec(x_75); -x_81 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_81, 0, x_80); -x_82 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__6; -x_83 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__10; -x_84 = lean_box(0); -x_85 = l_Lean_Elab_Command_expandMixfix___lambda__4(x_1, x_82, x_10, x_68, x_64, x_83, x_84, x_81, x_2, x_3); +lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_79 = l_Lean_Syntax_getArg(x_74, x_68); +lean_dec(x_74); +x_80 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_80, 0, x_79); +x_81 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__6; +x_82 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__10; +x_83 = lean_box(0); +x_84 = l_Lean_Elab_Command_expandMixfix___lambda__4(x_1, x_81, x_10, x_67, x_63, x_82, x_83, x_80, x_2, x_3); lean_dec(x_1); -return x_85; +return x_84; } } } else { -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; -lean_dec(x_70); -x_86 = lean_box(0); -x_87 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__6; -x_88 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__10; -x_89 = lean_box(0); -x_90 = l_Lean_Elab_Command_expandMixfix___lambda__4(x_1, x_87, x_10, x_68, x_64, x_88, x_89, x_86, x_2, x_3); +lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +lean_dec(x_69); +x_85 = lean_box(0); +x_86 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__6; +x_87 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__10; +x_88 = lean_box(0); +x_89 = l_Lean_Elab_Command_expandMixfix___lambda__4(x_1, x_86, x_10, x_67, x_63, x_87, x_88, x_85, x_2, x_3); lean_dec(x_1); -return x_90; +return x_89; } } } } else { -lean_object* x_91; lean_object* x_92; lean_object* x_93; uint8_t x_94; -lean_dec(x_20); -x_91 = lean_unsigned_to_nat(2u); -x_92 = l_Lean_Syntax_getArg(x_1, x_91); -x_93 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__24; -lean_inc(x_92); -x_94 = l_Lean_Syntax_isOfKind(x_92, x_93); -if (x_94 == 0) -{ -lean_object* x_95; lean_object* x_96; -lean_dec(x_92); +lean_object* x_90; lean_object* x_91; lean_object* x_92; uint8_t x_93; +lean_dec(x_19); +x_90 = lean_unsigned_to_nat(2u); +x_91 = l_Lean_Syntax_getArg(x_1, x_90); +x_92 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__24; +lean_inc(x_91); +x_93 = l_Lean_Syntax_isOfKind(x_91, x_92); +if (x_93 == 0) +{ +lean_object* x_94; lean_object* x_95; +lean_dec(x_91); lean_dec(x_2); lean_dec(x_1); -x_95 = lean_box(1); -x_96 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_96, 0, x_95); -lean_ctor_set(x_96, 1, x_3); -return x_96; +x_94 = lean_box(1); +x_95 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_95, 0, x_94); +lean_ctor_set(x_95, 1, x_3); +return x_95; } else { -lean_object* x_97; lean_object* x_98; lean_object* x_99; uint8_t x_100; -x_97 = l_Lean_Syntax_getArg(x_92, x_19); -lean_dec(x_92); -x_98 = lean_unsigned_to_nat(3u); -x_99 = l_Lean_Syntax_getArg(x_1, x_98); -x_100 = l_Lean_Syntax_isNone(x_99); +lean_object* x_96; lean_object* x_97; lean_object* x_98; uint8_t x_99; +x_96 = l_Lean_Syntax_getArg(x_91, x_18); +lean_dec(x_91); +x_97 = lean_unsigned_to_nat(3u); +x_98 = l_Lean_Syntax_getArg(x_1, x_97); +x_99 = l_Lean_Syntax_isNone(x_98); +if (x_99 == 0) +{ +uint8_t x_100; +lean_inc(x_98); +x_100 = l_Lean_Syntax_matchesNull(x_98, x_18); if (x_100 == 0) { -uint8_t x_101; -lean_inc(x_99); -x_101 = l_Lean_Syntax_isNodeOf(x_99, x_15, x_19); -if (x_101 == 0) -{ -lean_object* x_102; lean_object* x_103; -lean_dec(x_99); -lean_dec(x_97); +lean_object* x_101; lean_object* x_102; +lean_dec(x_98); +lean_dec(x_96); lean_dec(x_2); lean_dec(x_1); -x_102 = lean_box(1); -x_103 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_103, 0, x_102); -lean_ctor_set(x_103, 1, x_3); -return x_103; +x_101 = lean_box(1); +x_102 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_102, 0, x_101); +lean_ctor_set(x_102, 1, x_3); +return x_102; } else { -lean_object* x_104; lean_object* x_105; uint8_t x_106; -x_104 = l_Lean_Syntax_getArg(x_99, x_8); -lean_dec(x_99); -x_105 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__25; -lean_inc(x_104); -x_106 = l_Lean_Syntax_isOfKind(x_104, x_105); -if (x_106 == 0) -{ -lean_object* x_107; lean_object* x_108; -lean_dec(x_104); -lean_dec(x_97); +lean_object* x_103; lean_object* x_104; uint8_t x_105; +x_103 = l_Lean_Syntax_getArg(x_98, x_8); +lean_dec(x_98); +x_104 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__25; +lean_inc(x_103); +x_105 = l_Lean_Syntax_isOfKind(x_103, x_104); +if (x_105 == 0) +{ +lean_object* x_106; lean_object* x_107; +lean_dec(x_103); +lean_dec(x_96); lean_dec(x_2); lean_dec(x_1); -x_107 = lean_box(1); -x_108 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_108, 0, x_107); -lean_ctor_set(x_108, 1, x_3); -return x_108; +x_106 = lean_box(1); +x_107 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_107, 0, x_106); +lean_ctor_set(x_107, 1, x_3); +return x_107; } else { -lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; -x_109 = l_Lean_Syntax_getArg(x_104, x_98); -lean_dec(x_104); -x_110 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_110, 0, x_109); -x_111 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__6; -x_112 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__10; -x_113 = lean_box(0); -x_114 = l_Lean_Elab_Command_expandMixfix___lambda__6(x_1, x_97, x_111, x_10, x_93, x_112, x_113, x_110, x_2, x_3); +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; +x_108 = l_Lean_Syntax_getArg(x_103, x_97); +lean_dec(x_103); +x_109 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_109, 0, x_108); +x_110 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__6; +x_111 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__10; +x_112 = lean_box(0); +x_113 = l_Lean_Elab_Command_expandMixfix___lambda__6(x_1, x_96, x_110, x_10, x_92, x_111, x_112, x_109, x_2, x_3); lean_dec(x_1); -return x_114; +return x_113; } } } else { -lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; -lean_dec(x_99); -x_115 = lean_box(0); -x_116 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__6; -x_117 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__10; -x_118 = lean_box(0); -x_119 = l_Lean_Elab_Command_expandMixfix___lambda__6(x_1, x_97, x_116, x_10, x_93, x_117, x_118, x_115, x_2, x_3); +lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; +lean_dec(x_98); +x_114 = lean_box(0); +x_115 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__6; +x_116 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__10; +x_117 = lean_box(0); +x_118 = l_Lean_Elab_Command_expandMixfix___lambda__6(x_1, x_96, x_115, x_10, x_92, x_116, x_117, x_114, x_2, x_3); lean_dec(x_1); -return x_119; +return x_118; } } } } else { -lean_object* x_120; lean_object* x_121; lean_object* x_122; uint8_t x_123; -lean_dec(x_20); -x_120 = lean_unsigned_to_nat(2u); -x_121 = l_Lean_Syntax_getArg(x_1, x_120); -x_122 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__24; -lean_inc(x_121); -x_123 = l_Lean_Syntax_isOfKind(x_121, x_122); -if (x_123 == 0) -{ -lean_object* x_124; lean_object* x_125; -lean_dec(x_121); +lean_object* x_119; lean_object* x_120; lean_object* x_121; uint8_t x_122; +lean_dec(x_19); +x_119 = lean_unsigned_to_nat(2u); +x_120 = l_Lean_Syntax_getArg(x_1, x_119); +x_121 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__24; +lean_inc(x_120); +x_122 = l_Lean_Syntax_isOfKind(x_120, x_121); +if (x_122 == 0) +{ +lean_object* x_123; lean_object* x_124; +lean_dec(x_120); lean_dec(x_2); lean_dec(x_1); -x_124 = lean_box(1); -x_125 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_125, 0, x_124); -lean_ctor_set(x_125, 1, x_3); -return x_125; +x_123 = lean_box(1); +x_124 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_124, 0, x_123); +lean_ctor_set(x_124, 1, x_3); +return x_124; } else { -lean_object* x_126; lean_object* x_127; lean_object* x_128; uint8_t x_129; -x_126 = l_Lean_Syntax_getArg(x_121, x_19); -lean_dec(x_121); -x_127 = lean_unsigned_to_nat(3u); -x_128 = l_Lean_Syntax_getArg(x_1, x_127); -x_129 = l_Lean_Syntax_isNone(x_128); +lean_object* x_125; lean_object* x_126; lean_object* x_127; uint8_t x_128; +x_125 = l_Lean_Syntax_getArg(x_120, x_18); +lean_dec(x_120); +x_126 = lean_unsigned_to_nat(3u); +x_127 = l_Lean_Syntax_getArg(x_1, x_126); +x_128 = l_Lean_Syntax_isNone(x_127); +if (x_128 == 0) +{ +uint8_t x_129; +lean_inc(x_127); +x_129 = l_Lean_Syntax_matchesNull(x_127, x_18); if (x_129 == 0) { -uint8_t x_130; -lean_inc(x_128); -x_130 = l_Lean_Syntax_isNodeOf(x_128, x_15, x_19); -if (x_130 == 0) -{ -lean_object* x_131; lean_object* x_132; -lean_dec(x_128); -lean_dec(x_126); +lean_object* x_130; lean_object* x_131; +lean_dec(x_127); +lean_dec(x_125); lean_dec(x_2); lean_dec(x_1); -x_131 = lean_box(1); -x_132 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_132, 0, x_131); -lean_ctor_set(x_132, 1, x_3); -return x_132; +x_130 = lean_box(1); +x_131 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_131, 0, x_130); +lean_ctor_set(x_131, 1, x_3); +return x_131; } else { -lean_object* x_133; lean_object* x_134; uint8_t x_135; -x_133 = l_Lean_Syntax_getArg(x_128, x_8); -lean_dec(x_128); -x_134 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__25; -lean_inc(x_133); -x_135 = l_Lean_Syntax_isOfKind(x_133, x_134); -if (x_135 == 0) -{ -lean_object* x_136; lean_object* x_137; -lean_dec(x_133); -lean_dec(x_126); +lean_object* x_132; lean_object* x_133; uint8_t x_134; +x_132 = l_Lean_Syntax_getArg(x_127, x_8); +lean_dec(x_127); +x_133 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__25; +lean_inc(x_132); +x_134 = l_Lean_Syntax_isOfKind(x_132, x_133); +if (x_134 == 0) +{ +lean_object* x_135; lean_object* x_136; +lean_dec(x_132); +lean_dec(x_125); lean_dec(x_2); lean_dec(x_1); -x_136 = lean_box(1); -x_137 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_137, 0, x_136); -lean_ctor_set(x_137, 1, x_3); -return x_137; +x_135 = lean_box(1); +x_136 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_136, 0, x_135); +lean_ctor_set(x_136, 1, x_3); +return x_136; } else { -lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; -x_138 = l_Lean_Syntax_getArg(x_133, x_127); -lean_dec(x_133); -x_139 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_139, 0, x_138); -x_140 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__6; -x_141 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__10; -x_142 = lean_box(0); -x_143 = l_Lean_Elab_Command_expandMixfix___lambda__8(x_1, x_126, x_140, x_10, x_122, x_141, x_142, x_139, x_2, x_3); +lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; +x_137 = l_Lean_Syntax_getArg(x_132, x_126); +lean_dec(x_132); +x_138 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_138, 0, x_137); +x_139 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__6; +x_140 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__10; +x_141 = lean_box(0); +x_142 = l_Lean_Elab_Command_expandMixfix___lambda__8(x_1, x_125, x_139, x_10, x_121, x_140, x_141, x_138, x_2, x_3); lean_dec(x_1); -return x_143; +return x_142; } } } else { -lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; -lean_dec(x_128); -x_144 = lean_box(0); -x_145 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__6; -x_146 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__10; -x_147 = lean_box(0); -x_148 = l_Lean_Elab_Command_expandMixfix___lambda__8(x_1, x_126, x_145, x_10, x_122, x_146, x_147, x_144, x_2, x_3); +lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; +lean_dec(x_127); +x_143 = lean_box(0); +x_144 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__6; +x_145 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__10; +x_146 = lean_box(0); +x_147 = l_Lean_Elab_Command_expandMixfix___lambda__8(x_1, x_125, x_144, x_10, x_121, x_145, x_146, x_143, x_2, x_3); lean_dec(x_1); -return x_148; +return x_147; } } } } else { -lean_object* x_149; lean_object* x_150; lean_object* x_151; uint8_t x_152; -lean_dec(x_20); -x_149 = lean_unsigned_to_nat(2u); -x_150 = l_Lean_Syntax_getArg(x_1, x_149); -x_151 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__24; -lean_inc(x_150); -x_152 = l_Lean_Syntax_isOfKind(x_150, x_151); -if (x_152 == 0) -{ -lean_object* x_153; lean_object* x_154; -lean_dec(x_150); +lean_object* x_148; lean_object* x_149; lean_object* x_150; uint8_t x_151; +lean_dec(x_19); +x_148 = lean_unsigned_to_nat(2u); +x_149 = l_Lean_Syntax_getArg(x_1, x_148); +x_150 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__24; +lean_inc(x_149); +x_151 = l_Lean_Syntax_isOfKind(x_149, x_150); +if (x_151 == 0) +{ +lean_object* x_152; lean_object* x_153; +lean_dec(x_149); lean_dec(x_2); lean_dec(x_1); -x_153 = lean_box(1); -x_154 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_154, 0, x_153); -lean_ctor_set(x_154, 1, x_3); -return x_154; +x_152 = lean_box(1); +x_153 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_153, 0, x_152); +lean_ctor_set(x_153, 1, x_3); +return x_153; } else { -lean_object* x_155; lean_object* x_156; lean_object* x_157; uint8_t x_158; -x_155 = l_Lean_Syntax_getArg(x_150, x_19); -lean_dec(x_150); -x_156 = lean_unsigned_to_nat(3u); -x_157 = l_Lean_Syntax_getArg(x_1, x_156); -x_158 = l_Lean_Syntax_isNone(x_157); +lean_object* x_154; lean_object* x_155; lean_object* x_156; uint8_t x_157; +x_154 = l_Lean_Syntax_getArg(x_149, x_18); +lean_dec(x_149); +x_155 = lean_unsigned_to_nat(3u); +x_156 = l_Lean_Syntax_getArg(x_1, x_155); +x_157 = l_Lean_Syntax_isNone(x_156); +if (x_157 == 0) +{ +uint8_t x_158; +lean_inc(x_156); +x_158 = l_Lean_Syntax_matchesNull(x_156, x_18); if (x_158 == 0) { -uint8_t x_159; -lean_inc(x_157); -x_159 = l_Lean_Syntax_isNodeOf(x_157, x_15, x_19); -if (x_159 == 0) -{ -lean_object* x_160; lean_object* x_161; -lean_dec(x_157); -lean_dec(x_155); +lean_object* x_159; lean_object* x_160; +lean_dec(x_156); +lean_dec(x_154); lean_dec(x_2); lean_dec(x_1); -x_160 = lean_box(1); -x_161 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_161, 0, x_160); -lean_ctor_set(x_161, 1, x_3); -return x_161; +x_159 = lean_box(1); +x_160 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_160, 0, x_159); +lean_ctor_set(x_160, 1, x_3); +return x_160; } else { -lean_object* x_162; lean_object* x_163; uint8_t x_164; -x_162 = l_Lean_Syntax_getArg(x_157, x_8); -lean_dec(x_157); -x_163 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__25; -lean_inc(x_162); -x_164 = l_Lean_Syntax_isOfKind(x_162, x_163); -if (x_164 == 0) -{ -lean_object* x_165; lean_object* x_166; -lean_dec(x_162); -lean_dec(x_155); +lean_object* x_161; lean_object* x_162; uint8_t x_163; +x_161 = l_Lean_Syntax_getArg(x_156, x_8); +lean_dec(x_156); +x_162 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__25; +lean_inc(x_161); +x_163 = l_Lean_Syntax_isOfKind(x_161, x_162); +if (x_163 == 0) +{ +lean_object* x_164; lean_object* x_165; +lean_dec(x_161); +lean_dec(x_154); lean_dec(x_2); lean_dec(x_1); -x_165 = lean_box(1); -x_166 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_166, 0, x_165); -lean_ctor_set(x_166, 1, x_3); -return x_166; +x_164 = lean_box(1); +x_165 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_165, 0, x_164); +lean_ctor_set(x_165, 1, x_3); +return x_165; } else { -lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; -x_167 = l_Lean_Syntax_getArg(x_162, x_156); -lean_dec(x_162); -x_168 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_168, 0, x_167); -x_169 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__6; -x_170 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__10; -x_171 = lean_box(0); -x_172 = l_Lean_Elab_Command_expandMixfix___lambda__10(x_1, x_155, x_169, x_10, x_151, x_170, x_171, x_168, x_2, x_3); +lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; +x_166 = l_Lean_Syntax_getArg(x_161, x_155); +lean_dec(x_161); +x_167 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_167, 0, x_166); +x_168 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__6; +x_169 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__10; +x_170 = lean_box(0); +x_171 = l_Lean_Elab_Command_expandMixfix___lambda__10(x_1, x_154, x_168, x_10, x_150, x_169, x_170, x_167, x_2, x_3); lean_dec(x_1); -return x_172; +return x_171; } } } else { -lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; -lean_dec(x_157); -x_173 = lean_box(0); -x_174 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__6; -x_175 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__10; -x_176 = lean_box(0); -x_177 = l_Lean_Elab_Command_expandMixfix___lambda__10(x_1, x_155, x_174, x_10, x_151, x_175, x_176, x_173, x_2, x_3); +lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; +lean_dec(x_156); +x_172 = lean_box(0); +x_173 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__6; +x_174 = l_Lean_Elab_Command_expandMixfix___lambda__11___closed__10; +x_175 = lean_box(0); +x_176 = l_Lean_Elab_Command_expandMixfix___lambda__10(x_1, x_154, x_173, x_10, x_150, x_174, x_175, x_172, x_2, x_3); lean_dec(x_1); -return x_177; +return x_176; } } } @@ -3509,7 +3498,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandMixfix_declRang { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(33u); -x_2 = lean_unsigned_to_nat(32u); +x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -3523,7 +3512,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Elab_Command_expandMixfix_declRange___closed__1; x_2 = lean_unsigned_to_nat(43u); x_3 = l___regBuiltin_Lean_Elab_Command_expandMixfix_declRange___closed__2; -x_4 = lean_unsigned_to_nat(32u); +x_4 = lean_unsigned_to_nat(36u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); diff --git a/stage0/stdlib/Lean/Elab/MutualDef.c b/stage0/stdlib/Lean/Elab/MutualDef.c index c3c7a8f8ef7..d9e982947c0 100644 --- a/stage0/stdlib/Lean/Elab/MutualDef.c +++ b/stage0/stdlib/Lean/Elab/MutualDef.c @@ -30,8 +30,10 @@ lean_object* l_List_forM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTer static lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getAllUserLevelNames(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_isMultiConstant_x3f___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_isAuxDiscrName(lean_object*); lean_object* l_Lean_Expr_mvarId_x21(lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__3___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_erase_macro_scopes(lean_object*); lean_object* l_Lean_CollectMVars_visit(lean_object*, lean_object*); @@ -57,7 +59,6 @@ static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getPendind LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Term_checkForHiddenUnivLevels_visit___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalDecl_userName(lean_object*); static lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__7; -extern lean_object* l_Lean_nullKind; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__1(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withFunLocalDecls_loop___rarg___lambda__1(lean_object*, lean_object*, 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_Command_elabMutualDef(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -78,6 +79,7 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Mutua LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t lean_uint64_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_bindingDomain_x21(lean_object*); lean_object* l_Array_append___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_levelMVarToParamHeaders___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -105,6 +107,7 @@ static lean_object* l_Lean_addTrace___at___private_Lean_Elab_MutualDef_0__Lean_E extern lean_object* l_Lean_maxRecDepthErrorMessage; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT uint8_t l_Lean_Elab_Term_MutualClosure_FixPoint_State_modified___default; +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkModifiers___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__2___lambda__1(lean_object*, lean_object*, 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___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -114,6 +117,7 @@ lean_object* l_Lean_Elab_Term_elabBindersEx___rarg(lean_object*, lean_object*, l LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_expandWhereDecls(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__5; +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___boxed__const__1; lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_eraseAuxDiscr___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -126,7 +130,6 @@ LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0_ LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkModifiers___lambda__2___boxed(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___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_updateUsedVarsOf___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkLetRecsToLiftTypes___spec__1___boxed(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_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_mkInst_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_checkTraceOption(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_mkInst_x3f_go_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -139,18 +142,20 @@ static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_MutualD LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__8(lean_object*, lean_object*, lean_object*, size_t, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_processDeriving___spec__2___boxed(lean_object*, 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_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_registerFailedToInferDefTypeInfo___closed__3; uint8_t lean_name_eq(lean_object*, lean_object*); lean_object* l_Lean_LocalContext_get_x21(lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_Elab_Term_mkAuxName___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkKinds___closed__2; +LEAN_EXPORT lean_object* l_Array_sequenceMap___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_mkInst_x3f_go_x3f___lambda__2(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_elabAttrs___at_Lean_Elab_Command_elabMutualDef___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__7___closed__2; +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__5; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Std_AssocList_find_x3f___at_Lean_Elab_Term_checkForHiddenUnivLevels_visit___spec__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_processDefDeriving___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___closed__5; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabMutualDef___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_declRangeExt; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___closed__2; @@ -161,7 +166,6 @@ lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______ma LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_checkForHiddenUnivLevels_visit___spec__12___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkKinds___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_processDeriving___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_isMultiConstant_x3f(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_levelMVarToParamHeaders_process___boxed(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___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___lambda__1___boxed(lean_object*, lean_object*, lean_object*); @@ -179,11 +183,12 @@ lean_object* lean_string_append(lean_object*, lean_object*); lean_object* l_Std_RBNode_insert___at_Lean_Elab_Structural_findRecArg___spec__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___closed__4; +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__2___closed__2; lean_object* l_Lean_MessageData_ofList(lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___closed__1; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMutualDef___spec__11___closed__2; lean_object* l_Lean_Elab_Term_applyAttributesAt(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__3; lean_object* lean_string_utf8_extract(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkKinds___lambda__1(uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_MutualClosure_ClosureState_newLocalDecls___default; @@ -198,12 +203,12 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_Mutua static lean_object* l_Lean_Elab_elabModifiers___at_Lean_Elab_Command_elabMutualDef___spec__1___lambda__3___closed__5; lean_object* l_Std_mkHashSetImp___rarg(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getFunName(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___closed__2; static lean_object* l_List_mapTRAux___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getPendindMVarErrorMessage___spec__2___closed__1; lean_object* l_Lean_Elab_Term_getLevelNames___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_RBNode_find___at_Lean_Elab_Term_MutualClosure_Replacement_apply___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_RBNode_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_pushNewVars___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_addTrace___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureForAux___spec__4___closed__3; +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_MutualClosure_main___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_ReplaceImpl_replaceUnsafe(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -248,10 +253,10 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expan LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_processDefDeriving___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_isTheorem___boxed(lean_object*); -lean_object* l_Lean_Syntax_SepArray_ofElems(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_getUsedFVarsMap___rarg(lean_object*); lean_object* l_Std_RBNode_findCore___at_Lean_Elab_Structural_findRecArg_go___spec__11(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_MutualClosure_FixPoint_run___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, size_t, size_t, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__1; lean_object* l_Lean_throwError___at_Lean_Elab_Term_throwErrorIfErrors___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -267,20 +272,18 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabMutualDef___lambda__1(lean_obje lean_object* l_Lean_Expr_FindImpl_findUnsafe_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_processDefDeriving___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureForAux___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_processDeriving___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_MutualClosure_Replacement_apply___lambda__1(lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__5; lean_object* l_Lean_throwError___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__1; -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___closed__4; lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__10(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__2; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureForAux___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getFunName___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withFunLocalDecls_loop(lean_object*); LEAN_EXPORT lean_object* l_Std_AssocList_replace___at_Lean_Elab_Term_checkForHiddenUnivLevels_visit___spec__8(lean_object*, lean_object*, lean_object*); -lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__6; static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__5___closed__3; static lean_object* l_Lean_Elab_elabModifiers___at_Lean_Elab_Command_elabMutualDef___spec__1___lambda__3___closed__4; @@ -291,11 +294,9 @@ static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___la LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withUsed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_RBNode_setBlack___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runTermElabM___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabMutualDef___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkKinds___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_MutualClosure_insertReplacementForMainFns(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__2; lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_doubleQuotedNameToPattern___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -316,6 +317,7 @@ lean_object* l_Lean_Elab_levelMVarToParamPreDecls(lean_object*, lean_object*, le lean_object* l_Lean_Meta_Closure_mkLambda(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_isModified(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabImplicitLambda_loop___spec__1___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addTrace___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureForAux___spec__4___closed__4; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -324,6 +326,7 @@ lean_object* lean_st_ref_take(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_isMultiConstant_x3f___spec__3(lean_object*, size_t, size_t); lean_object* l_Lean_Meta_getZetaFVarIds___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_registerCustomErrorIfMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__2___closed__1; LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Term_checkForHiddenUnivLevels_visit___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, 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_elabModifiers___at_Lean_Elab_Command_elabMutualDef___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___lambda__4___closed__3; @@ -391,6 +394,7 @@ lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); lean_object* l_Std_mkHashMapImp___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__8___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_eraseAuxDiscr___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_instantiateMVarsAtHeader___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_processDeriving___spec__1___closed__5; @@ -415,6 +419,7 @@ LEAN_EXPORT lean_object* l_Std_RBNode_find___at___private_Lean_Elab_MutualDef_0_ lean_object* l_Lean_addAndCompile___at_Lean_Elab_Term_evalExpr___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_addLocalVarInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabModifiers___at_Lean_Elab_Command_elabMutualDef___spec__1___closed__3; +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__6; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerm___closed__1; lean_object* l_List_findSome_x3f___rarg(lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); @@ -432,7 +437,6 @@ uint64_t l_Lean_Expr_hash(lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureForAux___closed__5; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Elab_Term_checkForHiddenUnivLevels_visit___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_List_foldl___at_Lean_Elab_Term_MutualClosure_pushLetRecs___spec__1(uint8_t, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___closed__1; lean_object* l_Lean_Meta_mkForallFVars_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); @@ -442,6 +446,7 @@ LEAN_EXPORT lean_object* l_Std_mkHashMap___at_Lean_Elab_Term_checkForHiddenUnivL LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_levelMVarToParamHeaders_process___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isExprDefEqGuarded(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getPendindMVarErrorMessage___closed__2; +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureForAux___closed__3; LEAN_EXPORT lean_object* l_Std_RBNode_insert___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__4(lean_object*, lean_object*, lean_object*); @@ -468,16 +473,17 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_ static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__2; lean_object* l_Lean_Meta_InfoCacheKey_instHashableInfoCacheKey___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_checkForHiddenUnivLevels_visit___spec__12(lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__6; static lean_object* l_List_mapTRAux___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getPendindMVarErrorMessage___spec__1___closed__1; static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__5___closed__10; LEAN_EXPORT lean_object* l_Std_RBNode_insert___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__3(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_CollectLevelParams_main(lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withFunLocalDecls___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_MutualClosure_getModifiersForLetRecs___spec__2(lean_object*, size_t, size_t); extern lean_object* l_Lean_instInhabitedExpr; LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels___lambda__1___boxed(lean_object*, lean_object*, lean_object*, 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_Std_RBNode_ins___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__4(lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabMutualDef___lambda__2(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabModifiers___at_Lean_Elab_Command_elabMutualDef___spec__1___closed__2; size_t lean_usize_modn(size_t, lean_object*); @@ -501,6 +507,7 @@ LEAN_EXPORT lean_object* l_Std_AssocList_contains___at_Lean_Elab_Term_checkForHi LEAN_EXPORT lean_object* l_List_mapM___at_Lean_Elab_Term_MutualClosure_main___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__5___closed__13; +lean_object* l_Lean_mkSepArray(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_isTheorem___spec__1(lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_resetModified(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__7___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__8(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); @@ -511,9 +518,11 @@ lean_object* l_Lean_ResolveName_resolveGlobalName(lean_object*, lean_object*, le lean_object* l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_resetZetaFVarIds___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkModifiers___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___closed__3; uint8_t l_Lean_Expr_isForall(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_MutualClosure_FixPoint_State_usedFVarsMap___default; static lean_object* l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lambda__2___closed__3; +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__6; lean_object* l_Lean_mkFVar(lean_object*); static lean_object* l_Lean_Elab_elabModifiers___at_Lean_Elab_Command_elabMutualDef___spec__1___closed__4; uint8_t l_Lean_Expr_Data_binderInfo(uint64_t); @@ -530,14 +539,14 @@ lean_object* l_Lean_MetavarContext_getDelayedRoot(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_getMax_x3f___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_pickMaxFVar_x3f___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_addDeclarationRanges___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_isExample___spec__1___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___closed__3; lean_object* l_instBEqProd___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lambda__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addTrace___at___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTerm___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__3___lambda__1___boxed(lean_object*, lean_object*, 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_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__3; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_checkForHiddenUnivLevels___spec__3___closed__2; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_removeUnusedVars___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_checkForHiddenUnivLevels___spec__3(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -546,18 +555,14 @@ LEAN_EXPORT uint8_t l_Std_AssocList_contains___at_Lean_Elab_Term_checkForHiddenU LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceOptions(lean_object*); lean_object* l_Array_feraseIdx___rarg(lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__6; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_processDeriving___spec__1___closed__2; static lean_object* l_Lean_Elab_instInhabitedDefViewElabHeader___closed__4; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkModifiers___closed__1; extern lean_object* l_Lean_pp_funBinderTypes; -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__8___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__9(lean_object*, lean_object*, size_t, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabMutualDef___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_AssocList_find_x3f___at_Lean_Elab_Term_checkForHiddenUnivLevels_visit___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_MutualClosure_ClosureState_localDecls___default; -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_getUsedFVarsMap(lean_object*); static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__5___closed__6; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_MutualClosure_getKindForLetRecs___spec__1(lean_object*, size_t, size_t); @@ -570,7 +575,6 @@ uint8_t l_List_elem___at_Lean_NameHashSet_insert___spec__2(lean_object*, lean_ob LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosureFor___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getPos_x3f(lean_object*, uint8_t); lean_object* l_Lean_Meta_assignExprMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__4; uint8_t l_Array_contains___at___private_Lean_Meta_Match_Value_0__Lean_Meta_isUIntTypeName___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Name_getString_x21(lean_object*); lean_object* l_String_intercalate(lean_object*, lean_object*); @@ -595,6 +599,7 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Mutu LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_levelMVarToParamHeaders_process(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_Array_mapMUnsafe_map___at_Lean_Elab_Term_MutualClosure_main___spec__1___lambda__1(lean_object*, 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_List_forM___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_collectUsed___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__4; LEAN_EXPORT lean_object* l_List_mapTRAux___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getPendindMVarErrorMessage___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getPendindMVarErrorMessage___closed__3; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___lambda__4___closed__2; @@ -608,7 +613,6 @@ static lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___close lean_object* l_Lean_Meta_forallBoundedTelescope___at_Lean_Elab_Term_addAutoBoundImplicits_x27___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_expr_equal(lean_object*, lean_object*); lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2(lean_object*, lean_object*, size_t, size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_removeUnusedVars___boxed(lean_object*, 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_Std_RBNode_foldM___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_updateUsedVarsOf___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_RBNode_ins___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__5(lean_object*, lean_object*, lean_object*); @@ -643,7 +647,6 @@ LEAN_EXPORT lean_object* l_Array_erase___at___private_Lean_Elab_MutualDef_0__Lea lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_instantiateForallAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visit___closed__1; static lean_object* l_List_forM___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkLetRecsToLiftTypes___spec__1___closed__3; -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__4; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__1___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -668,7 +671,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_elabModifiers___at_Lean_Elab_Command_elabMu LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_isMultiConstant_x3f___spec__3___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_elabModifiers___at_Lean_Elab_Command_elabMutualDef___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); -lean_object* l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_levelMVarToParamHeaders___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerm(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___lambda__4___closed__7; @@ -699,12 +701,9 @@ LEAN_EXPORT lean_object* l_Std_RBNode_fold___at___private_Lean_Elab_MutualDef_0_ LEAN_EXPORT lean_object* l_Std_RBNode_find___at_Lean_Elab_Term_MutualClosure_Replacement_apply___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_registerFailedToInferDefTypeInfo___boxed(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_Term_MutualClosure_ClosureState_newLetDecls___default; -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__3___closed__1; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__4; lean_object* l_Lean_Elab_Term_elabType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__3; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_MutualClosure_getModifiersForLetRecs___boxed(lean_object*); static lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_isMultiConstant_x3f___spec__3___closed__8; @@ -712,7 +711,6 @@ static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWher lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabModifiers___at_Lean_Elab_Command_elabMutualDef___spec__1___lambda__3(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_MutualClosure_getModifiersForLetRecs___closed__1; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_isMultiConstant_x3f___spec__3___closed__2; static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__5___closed__11; lean_object* l_List_toArrayAux___rarg(lean_object*, lean_object*); @@ -723,8 +721,10 @@ lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_obje LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureForAux___spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*); uint8_t l_Lean_Elab_Modifiers_isNonrec(lean_object*); lean_object* l_Lean_expandMacros(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarImpl(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__1___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__2(lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef_processDeriving(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_erase___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureForAux___spec__1___boxed(lean_object*, lean_object*); @@ -737,6 +737,7 @@ static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWher static lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_isMultiConstant_x3f___spec__3___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_updateUsedVarsOf___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__7; lean_object* l_Lean_Name_quickCmp___boxed(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_processDeriving___spec__1___closed__6; lean_object* l_instHashableProd___rarg___boxed(lean_object*, lean_object*, lean_object*); @@ -744,8 +745,8 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDe uint8_t l_Array_contains___at_Lean_Meta_getElimInfo___spec__4(lean_object*, lean_object*); lean_object* l_Lean_resolveGlobalConstNoOverload___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_doubleQuotedNameToPattern___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_RBNode_find___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_updateUsedVarsOf___spec__1___boxed(lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__2; lean_object* l_Lean_Meta_instantiateMVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___closed__1; lean_object* lean_mk_array(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_modifyUsedFVars___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_MutualClosure_main___spec__6(lean_object*, size_t, size_t, lean_object*); @@ -753,7 +754,6 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Mutua static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__3___closed__1; lean_object* l_Lean_Elab_Term_logUnassignedUsingErrorInfos(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabDeclAttrs___at_Lean_Elab_Command_elabMutualDef___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getTailPos_x3f(lean_object*, uint8_t); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_mkInst_x3f_go_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -761,34 +761,34 @@ extern lean_object* l_Lean_Expr_instBEqExpr; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_collectUsed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_withoutAutoBoundImplicit___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_fixpoint___boxed(lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___closed__2; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__3___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_MutualClosure_main___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_elabBinders___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkKinds___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__3___boxed(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___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_fixpoint___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_modifyUsedFVars(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_sequenceMap___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_instantiateMVarsAtLetRecToLift(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__2; lean_object* l_Lean_indentD(lean_object*); extern lean_object* l_Lean_Elab_unsupportedSyntaxExceptionId; lean_object* l_Lean_Elab_instantiateMVarsAtPreDecls(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__1; lean_object* l_Lean_Elab_toAttributeKind___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__7; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_pushNewVars(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_fixpoint(lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__2; LEAN_EXPORT lean_object* l_Std_RBNode_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabMutualDef___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_isMultiConstant_x3f___spec__3___closed__4; +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___closed__1; lean_object* l_Lean_Elab_Term_addAutoBoundImplicits(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_whnfD(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__3___closed__2; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lambda__2(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); @@ -796,6 +796,7 @@ lean_object* lean_name_append_before(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_checkForHiddenUnivLevels___spec__3___closed__4; static lean_object* l_Lean_Elab_addDeclarationRanges___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__3___closed__3; LEAN_EXPORT uint8_t l_List_foldr___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_typeHasRecFun___spec__1(lean_object*, uint8_t, lean_object*); +static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__7; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___lambda__1___boxed(lean_object**); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__3___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_updateUsedVarsOf(lean_object*, lean_object*, lean_object*); @@ -812,7 +813,6 @@ LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_processDe lean_object* l_Lean_Core_resetMessageLog(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_mkInst_x3f_go_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__1(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_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_pp_universes; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__2___closed__6; @@ -820,18 +820,15 @@ LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Elab_MutualDef_0__Lean_ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_pickMaxFVar_x3f___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_getScope___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkKinds___lambda__1___closed__2; -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visit___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_fixLevelParams(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_Array_mapMUnsafe_map___at_Lean_Elab_Term_MutualClosure_main___spec__1___lambda__1___boxed(lean_object*, 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_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLetDeclImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__3(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_Term_checkForHiddenUnivLevels___lambda__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__5___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__6(lean_object*, lean_object*, size_t, size_t, lean_object*); uint8_t l_List_beq___at_Lean_OpenDecl_instToStringOpenDecl___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__3___closed__3; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__1; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_processDeriving___spec__1___closed__1; lean_object* l_Lean_indentExpr(lean_object*); @@ -863,7 +860,6 @@ lean_object* l_Lean_Option_set___at_Lean_Meta_withPPInaccessibleNamesImp___spec_ static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_processDeriving___spec__1___closed__4; LEAN_EXPORT lean_object* l_List_mapTRAux___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_isMultiConstant_x3f___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_pickMaxFVar_x3f(lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__5; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_collectUsed___spec__3___boxed(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___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___lambda__2(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_registerFailedToInferDefTypeInfo___closed__2; @@ -872,9 +868,10 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_Mutua static lean_object* l_List_mapTRAux___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getPendindMVarErrorMessage___spec__2___closed__2; uint8_t l___private_Lean_Elab_DefView_0__Lean_Elab_beqDefKind____x40_Lean_Elab_DefView___hyg_15_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_List_mapM___at_Lean_Elab_Term_elabMutualDef_go___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___closed__1; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_MutualClosure_FixPoint_run(lean_object*, lean_object*); lean_object* l_Lean_setEnv___at_Lean_Elab_Term_evalExpr___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__5___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_markModified___boxed(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMutualDef___spec__11___closed__1; @@ -882,6 +879,7 @@ lean_object* l_Lean_Meta_check(lean_object*, lean_object*, lean_object*, lean_ob static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___lambda__4___closed__6; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosureFor(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_checkForHiddenUnivLevels_visit___spec__12___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__2___closed__3; lean_object* l_Lean_Elab_mkUnusedBaseName(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels___closed__2; @@ -5587,7 +5585,87 @@ x_2 = lean_alloc_closure((void*)(l___private_Lean_Elab_MutualDef_0__Lean_Elab_Te return x_2; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___closed__1() { +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2(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; uint8_t x_7; +x_6 = lean_array_get_size(x_1); +x_7 = lean_nat_dec_lt(x_4, x_6); +lean_dec(x_6); +if (x_7 == 0) +{ +lean_object* x_8; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_5); +return x_8; +} +else +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_unsigned_to_nat(0u); +x_10 = lean_nat_dec_eq(x_3, x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_sub(x_3, x_11); +lean_dec(x_3); +x_13 = lean_array_fget(x_1, x_4); +lean_inc(x_2); +x_14 = lean_apply_1(x_2, x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; +lean_dec(x_12); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_15 = lean_box(0); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_14, 0); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_nat_add(x_4, x_11); +lean_dec(x_4); +x_18 = lean_array_push(x_5, x_16); +x_3 = x_12; +x_4 = x_17; +x_5 = x_18; +goto _start; +} +} +else +{ +lean_object* x_20; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_20 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_20, 0, x_5); +return x_20; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_sequenceMap___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1(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; +x_3 = lean_array_get_size(x_1); +x_4 = lean_mk_empty_array_with_capacity(x_3); +x_5 = lean_unsigned_to_nat(0u); +x_6 = l_Array_sequenceMap_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2(x_1, x_2, x_3, x_5, x_4); +return x_6; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__1() { _start: { lean_object* x_1; @@ -5595,7 +5673,7 @@ x_1 = lean_mk_string_from_bytes("letDecl", 7); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___closed__2() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__2() { _start: { lean_object* x_1; @@ -5603,7 +5681,7 @@ x_1 = lean_mk_string_from_bytes("letPatDecl", 10); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___closed__3() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__3() { _start: { lean_object* x_1; @@ -5611,7 +5689,7 @@ x_1 = lean_mk_string_from_bytes("letEqnsDecl", 11); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___closed__4() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__4() { _start: { lean_object* x_1; @@ -5619,7 +5697,7 @@ x_1 = lean_mk_string_from_bytes("letIdDecl", 9); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___closed__5() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__5() { _start: { lean_object* x_1; @@ -5627,7 +5705,7 @@ x_1 = lean_mk_string_from_bytes("patterns are not allowed here", 29); return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { uint8_t x_7; @@ -5651,7 +5729,7 @@ x_11 = lean_array_uset(x_4, x_3, x_10); x_12 = l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_isMultiConstant_x3f___spec__3___closed__5; lean_inc(x_1); x_13 = lean_name_mk_string(x_1, x_12); -x_14 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___closed__1; +x_14 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__1; lean_inc(x_13); x_15 = lean_name_mk_string(x_13, x_14); lean_inc(x_9); @@ -5675,7 +5753,7 @@ else { lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; x_19 = l_Lean_Syntax_getArg(x_9, x_10); -x_20 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___closed__2; +x_20 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__2; lean_inc(x_13); x_21 = lean_name_mk_string(x_13, x_20); lean_inc(x_19); @@ -5685,7 +5763,7 @@ if (x_22 == 0) { lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_dec(x_9); -x_23 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___closed__3; +x_23 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__3; lean_inc(x_13); x_24 = lean_name_mk_string(x_13, x_23); lean_inc(x_19); @@ -5694,7 +5772,7 @@ lean_dec(x_24); if (x_25 == 0) { lean_object* x_26; lean_object* x_27; uint8_t x_28; -x_26 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___closed__4; +x_26 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__4; x_27 = lean_name_mk_string(x_13, x_26); lean_inc(x_19); x_28 = l_Lean_Syntax_isOfKind(x_19, x_27); @@ -5750,7 +5828,7 @@ lean_dec(x_19); lean_dec(x_13); lean_dec(x_11); lean_dec(x_1); -x_42 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___closed__5; +x_42 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__5; x_43 = l_Lean_Macro_throwErrorAt___rarg(x_9, x_42, x_5, x_6); lean_dec(x_9); x_44 = !lean_is_exclusive(x_43); @@ -5776,16 +5854,7 @@ return x_47; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__1(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__1() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__1() { _start: { lean_object* x_1; @@ -5793,7 +5862,7 @@ x_1 = lean_mk_string_from_bytes("structInstField", 15); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__2() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__2() { _start: { lean_object* x_1; @@ -5801,7 +5870,7 @@ x_1 = lean_mk_string_from_bytes("structInstLVal", 14); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__3() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__3() { _start: { lean_object* x_1; @@ -5809,17 +5878,17 @@ x_1 = lean_mk_string_from_bytes("null", 4); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__4() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__3; +x_2 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__5() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__5() { _start: { lean_object* x_1; lean_object* x_2; @@ -5828,7 +5897,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__6() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__6() { _start: { lean_object* x_1; @@ -5836,7 +5905,7 @@ x_1 = lean_mk_string_from_bytes(":=", 2); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__7() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__7() { _start: { lean_object* x_1; lean_object* x_2; @@ -5845,7 +5914,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___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_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___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) { _start: { lean_object* x_8; uint8_t x_9; @@ -5855,29 +5924,29 @@ 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; 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; x_10 = lean_ctor_get(x_8, 0); -x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__1; +x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__1; lean_inc(x_1); x_12 = lean_name_mk_string(x_1, x_11); -x_13 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__2; +x_13 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__2; x_14 = lean_name_mk_string(x_1, x_13); x_15 = lean_box(2); -x_16 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__4; +x_16 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__4; x_17 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_17, 0, x_15); lean_ctor_set(x_17, 1, x_16); lean_ctor_set(x_17, 2, x_2); -x_18 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__5; +x_18 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__5; x_19 = lean_array_push(x_18, x_3); x_20 = lean_array_push(x_19, x_17); x_21 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_21, 0, x_15); lean_ctor_set(x_21, 1, x_14); lean_ctor_set(x_21, 2, x_20); -x_22 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__6; +x_22 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__6; x_23 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_23, 0, x_10); lean_ctor_set(x_23, 1, x_22); -x_24 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__7; +x_24 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__7; x_25 = lean_array_push(x_24, x_21); x_26 = lean_array_push(x_25, x_23); x_27 = lean_array_push(x_26, x_5); @@ -5896,29 +5965,29 @@ x_30 = lean_ctor_get(x_8, 1); lean_inc(x_30); lean_inc(x_29); lean_dec(x_8); -x_31 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__1; +x_31 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__1; lean_inc(x_1); x_32 = lean_name_mk_string(x_1, x_31); -x_33 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__2; +x_33 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__2; x_34 = lean_name_mk_string(x_1, x_33); x_35 = lean_box(2); -x_36 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__4; +x_36 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__4; x_37 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_37, 0, x_35); lean_ctor_set(x_37, 1, x_36); lean_ctor_set(x_37, 2, x_2); -x_38 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__5; +x_38 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__5; x_39 = lean_array_push(x_38, x_3); x_40 = lean_array_push(x_39, x_37); x_41 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_41, 0, x_35); lean_ctor_set(x_41, 1, x_34); lean_ctor_set(x_41, 2, x_40); -x_42 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__6; +x_42 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__6; x_43 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_43, 0, x_29); lean_ctor_set(x_43, 1, x_42); -x_44 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__7; +x_44 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__7; x_45 = lean_array_push(x_44, x_41); x_46 = lean_array_push(x_45, x_43); x_47 = lean_array_push(x_46, x_5); @@ -5933,7 +6002,7 @@ return x_49; } } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__3___closed__1() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__2___closed__1() { _start: { lean_object* x_1; @@ -5941,7 +6010,7 @@ x_1 = lean_mk_string_from_bytes("fun", 3); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__3___closed__2() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__2___closed__2() { _start: { lean_object* x_1; @@ -5949,7 +6018,7 @@ x_1 = lean_mk_string_from_bytes("basicFun", 8); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__3___closed__3() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__2___closed__3() { _start: { lean_object* x_1; @@ -5957,82 +6026,79 @@ x_1 = lean_mk_string_from_bytes("=>", 2); return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___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) { _start: { -lean_object* x_10; lean_object* x_11; uint8_t x_12; -lean_dec(x_7); -x_10 = lean_array_get_size(x_4); -x_11 = lean_unsigned_to_nat(0u); -x_12 = lean_nat_dec_lt(x_11, x_10); -if (x_12 == 0) +lean_object* x_9; lean_object* x_10; uint8_t x_11; +lean_dec(x_6); +x_9 = lean_array_get_size(x_4); +x_10 = lean_unsigned_to_nat(0u); +x_11 = lean_nat_dec_lt(x_10, x_9); +lean_dec(x_9); +if (x_11 == 0) { -lean_object* x_13; -lean_dec(x_10); +lean_object* x_12; lean_dec(x_4); -lean_inc(x_6); -x_13 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2(x_1, x_2, x_3, x_6, x_6, x_8, x_9); -lean_dec(x_6); -return x_13; +lean_inc(x_5); +x_12 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1(x_1, x_2, x_3, x_5, x_5, x_7, x_8); +lean_dec(x_5); +return x_12; } 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; lean_object* x_20; lean_object* x_21; size_t 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; 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_inc(x_8); -x_14 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_8, x_9); -x_15 = lean_ctor_get(x_14, 0); +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; 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; 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_inc(x_7); +x_13 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_7, x_8); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); -lean_inc(x_16); -lean_dec(x_14); -x_17 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__3___closed__1; +lean_dec(x_13); +x_16 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__2___closed__1; lean_inc(x_1); -x_18 = lean_name_mk_string(x_1, x_17); -lean_inc(x_15); -x_19 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_19, 0, x_15); -lean_ctor_set(x_19, 1, x_17); -x_20 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__3___closed__2; +x_17 = lean_name_mk_string(x_1, x_16); +lean_inc(x_14); +x_18 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_18, 0, x_14); +lean_ctor_set(x_18, 1, x_16); +x_19 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__2___closed__2; lean_inc(x_1); -x_21 = lean_name_mk_string(x_1, x_20); -x_22 = lean_usize_of_nat(x_10); -lean_dec(x_10); -x_23 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_22, x_5, x_4); +x_20 = lean_name_mk_string(x_1, x_19); lean_inc(x_2); -x_24 = l_Array_append___rarg(x_2, x_23); -x_25 = lean_box(2); -x_26 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__4; -x_27 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_27, 0, x_25); -lean_ctor_set(x_27, 1, x_26); -lean_ctor_set(x_27, 2, x_24); -x_28 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__3___closed__3; -x_29 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_29, 0, x_15); -lean_ctor_set(x_29, 1, x_28); -x_30 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__7; -x_31 = lean_array_push(x_30, x_27); -x_32 = lean_array_push(x_31, x_29); -lean_inc(x_6); -x_33 = lean_array_push(x_32, x_6); -x_34 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_34, 0, x_25); -lean_ctor_set(x_34, 1, x_21); -lean_ctor_set(x_34, 2, x_33); -x_35 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__5; -x_36 = lean_array_push(x_35, x_19); -x_37 = lean_array_push(x_36, x_34); -x_38 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_38, 0, x_25); -lean_ctor_set(x_38, 1, x_18); -lean_ctor_set(x_38, 2, x_37); -x_39 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2(x_1, x_2, x_3, x_6, x_38, x_8, x_16); -lean_dec(x_6); -return x_39; +x_21 = l_Array_append___rarg(x_2, x_4); +x_22 = lean_box(2); +x_23 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__4; +x_24 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +lean_ctor_set(x_24, 2, x_21); +x_25 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__2___closed__3; +x_26 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_26, 0, x_14); +lean_ctor_set(x_26, 1, x_25); +x_27 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__7; +x_28 = lean_array_push(x_27, x_24); +x_29 = lean_array_push(x_28, x_26); +lean_inc(x_5); +x_30 = lean_array_push(x_29, x_5); +x_31 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_31, 0, x_22); +lean_ctor_set(x_31, 1, x_20); +lean_ctor_set(x_31, 2, x_30); +x_32 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__5; +x_33 = lean_array_push(x_32, x_18); +x_34 = lean_array_push(x_33, x_31); +x_35 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_35, 0, x_22); +lean_ctor_set(x_35, 1, x_17); +lean_ctor_set(x_35, 2, x_34); +x_36 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1(x_1, x_2, x_3, x_5, x_35, x_7, x_15); +lean_dec(x_5); +return x_36; } } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__1() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__1() { _start: { lean_object* x_1; @@ -6040,7 +6106,7 @@ x_1 = lean_mk_string_from_bytes("paren", 5); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__2() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__2() { _start: { lean_object* x_1; @@ -6048,7 +6114,7 @@ x_1 = lean_mk_string_from_bytes("(", 1); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__3() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__3() { _start: { lean_object* x_1; @@ -6056,7 +6122,7 @@ x_1 = lean_mk_string_from_bytes("typeAscription", 14); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__4() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__4() { _start: { lean_object* x_1; @@ -6064,7 +6130,7 @@ x_1 = lean_mk_string_from_bytes(":", 1); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__5() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__5() { _start: { lean_object* x_1; lean_object* x_2; @@ -6073,7 +6139,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__6() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__6() { _start: { lean_object* x_1; @@ -6081,59 +6147,61 @@ x_1 = lean_mk_string_from_bytes(")", 1); return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3(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_object* x_9) { _start: { -lean_object* x_11; lean_object* x_12; uint8_t x_13; -lean_dec(x_7); -x_11 = lean_unsigned_to_nat(4u); -x_12 = l_Lean_Syntax_getArg(x_1, x_11); -x_13 = !lean_is_exclusive(x_9); +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +lean_dec(x_6); +x_10 = lean_unsigned_to_nat(4u); +x_11 = l_Lean_Syntax_getArg(x_1, x_10); +x_12 = l_Lean_Syntax_getArgs(x_2); +lean_dec(x_2); +x_13 = !lean_is_exclusive(x_8); if (x_13 == 0) { lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_9, 5); +x_14 = lean_ctor_get(x_8, 5); x_15 = l_Lean_replaceRef(x_1, x_14); lean_dec(x_14); lean_dec(x_1); -lean_ctor_set(x_9, 5, x_15); -if (lean_obj_tag(x_8) == 0) +lean_ctor_set(x_8, 5, x_15); +if (lean_obj_tag(x_7) == 0) { lean_object* x_16; lean_object* x_17; x_16 = lean_box(0); -x_17 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__3(x_2, x_3, x_4, x_5, x_6, x_12, x_16, x_9, x_10); +x_17 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__2(x_3, x_4, x_5, x_12, x_11, x_16, x_8, x_9); return x_17; } else { 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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_18 = lean_ctor_get(x_8, 0); +x_18 = lean_ctor_get(x_7, 0); lean_inc(x_18); -lean_dec(x_8); -lean_inc(x_9); -x_19 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_9, x_10); +lean_dec(x_7); +lean_inc(x_8); +x_19 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_8, x_9); x_20 = lean_ctor_get(x_19, 0); lean_inc(x_20); x_21 = lean_ctor_get(x_19, 1); lean_inc(x_21); lean_dec(x_19); -x_22 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__1; -lean_inc(x_2); -x_23 = lean_name_mk_string(x_2, x_22); -x_24 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__2; +x_22 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__1; +lean_inc(x_3); +x_23 = lean_name_mk_string(x_3, x_22); +x_24 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__2; lean_inc(x_20); x_25 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_25, 0, x_20); lean_ctor_set(x_25, 1, x_24); -x_26 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__3; -lean_inc(x_2); -x_27 = lean_name_mk_string(x_2, x_26); -x_28 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__4; +x_26 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__3; +lean_inc(x_3); +x_27 = lean_name_mk_string(x_3, x_26); +x_28 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__4; lean_inc(x_20); x_29 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_29, 0, x_20); lean_ctor_set(x_29, 1, x_28); -x_30 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__5; +x_30 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__5; x_31 = lean_array_push(x_30, x_29); x_32 = lean_array_push(x_31, x_18); x_33 = lean_box(2); @@ -6141,24 +6209,24 @@ x_34 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_34, 0, x_33); lean_ctor_set(x_34, 1, x_27); lean_ctor_set(x_34, 2, x_32); -x_35 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__5; +x_35 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__5; x_36 = lean_array_push(x_35, x_34); -x_37 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__4; +x_37 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__4; x_38 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_38, 0, x_33); lean_ctor_set(x_38, 1, x_37); lean_ctor_set(x_38, 2, x_36); -x_39 = lean_array_push(x_30, x_12); +x_39 = lean_array_push(x_30, x_11); x_40 = lean_array_push(x_39, x_38); x_41 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_41, 0, x_33); lean_ctor_set(x_41, 1, x_37); lean_ctor_set(x_41, 2, x_40); -x_42 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__6; +x_42 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__6; x_43 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_43, 0, x_20); lean_ctor_set(x_43, 1, x_42); -x_44 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__7; +x_44 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__7; x_45 = lean_array_push(x_44, x_25); x_46 = lean_array_push(x_45, x_41); x_47 = lean_array_push(x_46, x_43); @@ -6167,26 +6235,26 @@ lean_ctor_set(x_48, 0, x_33); lean_ctor_set(x_48, 1, x_23); lean_ctor_set(x_48, 2, x_47); x_49 = lean_box(0); -x_50 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__3(x_2, x_3, x_4, x_5, x_6, x_48, x_49, x_9, x_21); +x_50 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__2(x_3, x_4, x_5, x_12, x_48, x_49, x_8, x_21); return x_50; } } else { lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_51 = lean_ctor_get(x_9, 0); -x_52 = lean_ctor_get(x_9, 1); -x_53 = lean_ctor_get(x_9, 2); -x_54 = lean_ctor_get(x_9, 3); -x_55 = lean_ctor_get(x_9, 4); -x_56 = lean_ctor_get(x_9, 5); +x_51 = lean_ctor_get(x_8, 0); +x_52 = lean_ctor_get(x_8, 1); +x_53 = lean_ctor_get(x_8, 2); +x_54 = lean_ctor_get(x_8, 3); +x_55 = lean_ctor_get(x_8, 4); +x_56 = lean_ctor_get(x_8, 5); lean_inc(x_56); lean_inc(x_55); lean_inc(x_54); lean_inc(x_53); lean_inc(x_52); lean_inc(x_51); -lean_dec(x_9); +lean_dec(x_8); x_57 = l_Lean_replaceRef(x_1, x_56); lean_dec(x_56); lean_dec(x_1); @@ -6197,43 +6265,43 @@ lean_ctor_set(x_58, 2, x_53); lean_ctor_set(x_58, 3, x_54); lean_ctor_set(x_58, 4, x_55); lean_ctor_set(x_58, 5, x_57); -if (lean_obj_tag(x_8) == 0) +if (lean_obj_tag(x_7) == 0) { lean_object* x_59; lean_object* x_60; x_59 = lean_box(0); -x_60 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__3(x_2, x_3, x_4, x_5, x_6, x_12, x_59, x_58, x_10); +x_60 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__2(x_3, x_4, x_5, x_12, x_11, x_59, x_58, x_9); return x_60; } else { lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_61 = lean_ctor_get(x_8, 0); +x_61 = lean_ctor_get(x_7, 0); lean_inc(x_61); -lean_dec(x_8); +lean_dec(x_7); lean_inc(x_58); -x_62 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_58, x_10); +x_62 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_58, x_9); x_63 = lean_ctor_get(x_62, 0); lean_inc(x_63); x_64 = lean_ctor_get(x_62, 1); lean_inc(x_64); lean_dec(x_62); -x_65 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__1; -lean_inc(x_2); -x_66 = lean_name_mk_string(x_2, x_65); -x_67 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__2; +x_65 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__1; +lean_inc(x_3); +x_66 = lean_name_mk_string(x_3, x_65); +x_67 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__2; lean_inc(x_63); x_68 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_68, 0, x_63); lean_ctor_set(x_68, 1, x_67); -x_69 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__3; -lean_inc(x_2); -x_70 = lean_name_mk_string(x_2, x_69); -x_71 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__4; +x_69 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__3; +lean_inc(x_3); +x_70 = lean_name_mk_string(x_3, x_69); +x_71 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__4; lean_inc(x_63); x_72 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_72, 0, x_63); lean_ctor_set(x_72, 1, x_71); -x_73 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__5; +x_73 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__5; x_74 = lean_array_push(x_73, x_72); x_75 = lean_array_push(x_74, x_61); x_76 = lean_box(2); @@ -6241,24 +6309,24 @@ x_77 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_77, 0, x_76); lean_ctor_set(x_77, 1, x_70); lean_ctor_set(x_77, 2, x_75); -x_78 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__5; +x_78 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__5; x_79 = lean_array_push(x_78, x_77); -x_80 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__4; +x_80 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__4; x_81 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_81, 0, x_76); lean_ctor_set(x_81, 1, x_80); lean_ctor_set(x_81, 2, x_79); -x_82 = lean_array_push(x_73, x_12); +x_82 = lean_array_push(x_73, x_11); x_83 = lean_array_push(x_82, x_81); x_84 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_84, 0, x_76); lean_ctor_set(x_84, 1, x_80); lean_ctor_set(x_84, 2, x_83); -x_85 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__6; +x_85 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__6; x_86 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_86, 0, x_63); lean_ctor_set(x_86, 1, x_85); -x_87 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__7; +x_87 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__7; x_88 = lean_array_push(x_87, x_68); x_89 = lean_array_push(x_88, x_84); x_90 = lean_array_push(x_89, x_86); @@ -6267,13 +6335,13 @@ lean_ctor_set(x_91, 0, x_76); lean_ctor_set(x_91, 1, x_66); lean_ctor_set(x_91, 2, x_90); x_92 = lean_box(0); -x_93 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__3(x_2, x_3, x_4, x_5, x_6, x_91, x_92, x_58, x_64); +x_93 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__2(x_3, x_4, x_5, x_12, x_91, x_92, x_58, x_64); return x_93; } } } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___closed__1() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___closed__1() { _start: { lean_object* x_1; @@ -6281,25 +6349,17 @@ x_1 = lean_mk_string_from_bytes("ident", 5); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___closed__2() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___closed__1; +x_2 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__1), 1, 0); -return x_1; -} -} -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___closed__4() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___closed__3() { _start: { lean_object* x_1; @@ -6307,7 +6367,7 @@ x_1 = lean_mk_string_from_bytes("typeSpec", 8); return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, size_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4(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_object* x_7, lean_object* x_8) { _start: { uint8_t x_9; @@ -6325,306 +6385,161 @@ return x_10; } else { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; x_11 = lean_array_uget(x_6, x_5); x_12 = lean_unsigned_to_nat(0u); x_13 = lean_array_uset(x_6, x_5, x_12); -x_21 = l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_isMultiConstant_x3f___spec__3___closed__5; -lean_inc(x_1); -x_22 = lean_name_mk_string(x_1, x_21); -x_23 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___closed__4; -lean_inc(x_22); -x_24 = lean_name_mk_string(x_22, x_23); +x_14 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__4; +lean_inc(x_2); +x_15 = lean_name_mk_string(x_2, x_14); lean_inc(x_11); -x_25 = l_Lean_Syntax_isOfKind(x_11, x_24); -lean_dec(x_24); -if (x_25 == 0) -{ -lean_object* x_26; lean_object* x_27; -lean_dec(x_22); -lean_dec(x_13); -lean_dec(x_11); -lean_dec(x_7); -lean_dec(x_2); -lean_dec(x_1); -x_26 = lean_box(1); -x_27 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_8); -return x_27; -} -else -{ -lean_object* x_28; lean_object* x_29; uint8_t x_30; -x_28 = l_Lean_Syntax_getArg(x_11, x_12); -x_29 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___closed__2; -lean_inc(x_28); -x_30 = l_Lean_Syntax_isOfKind(x_28, x_29); -if (x_30 == 0) +x_16 = l_Lean_Syntax_isOfKind(x_11, x_15); +lean_dec(x_15); +if (x_16 == 0) { -lean_object* x_31; lean_object* x_32; -lean_dec(x_28); -lean_dec(x_22); +lean_object* x_17; lean_object* x_18; lean_dec(x_13); lean_dec(x_11); lean_dec(x_7); lean_dec(x_2); lean_dec(x_1); -x_31 = lean_box(1); -x_32 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_8); -return x_32; +x_17 = lean_box(1); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_8); +return x_18; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_33 = lean_unsigned_to_nat(1u); -x_34 = l_Lean_Syntax_getArg(x_11, x_33); -x_35 = l_Lean_Syntax_getArgs(x_34); -lean_dec(x_34); -x_36 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___closed__3; -x_37 = l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__1(x_35, x_36); -lean_dec(x_35); -if (lean_obj_tag(x_37) == 0) +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = l_Lean_Syntax_getArg(x_11, x_12); +x_20 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___closed__2; +lean_inc(x_19); +x_21 = l_Lean_Syntax_isOfKind(x_19, x_20); +if (x_21 == 0) { -lean_object* x_38; lean_object* x_39; -lean_dec(x_28); -lean_dec(x_22); +lean_object* x_22; lean_object* x_23; +lean_dec(x_19); lean_dec(x_13); lean_dec(x_11); lean_dec(x_7); lean_dec(x_2); lean_dec(x_1); -x_38 = lean_box(1); -x_39 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_8); -return x_39; +x_22 = lean_box(1); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_8); +return x_23; } else { -uint8_t x_40; -x_40 = !lean_is_exclusive(x_37); -if (x_40 == 0) -{ -lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; -x_41 = lean_ctor_get(x_37, 0); -x_42 = lean_unsigned_to_nat(2u); -x_43 = l_Lean_Syntax_getArg(x_11, x_42); -x_44 = l_Lean_Syntax_isNone(x_43); -if (x_44 == 0) -{ -lean_object* x_45; uint8_t x_46; -x_45 = l_Lean_nullKind; -lean_inc(x_43); -x_46 = l_Lean_Syntax_isNodeOf(x_43, x_45, x_33); -if (x_46 == 0) -{ -lean_object* x_47; lean_object* x_48; -lean_dec(x_43); -lean_free_object(x_37); -lean_dec(x_41); -lean_dec(x_28); -lean_dec(x_22); -lean_dec(x_13); -lean_dec(x_11); -lean_dec(x_7); -lean_dec(x_2); -lean_dec(x_1); -x_47 = lean_box(1); -x_48 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_8); -return x_48; -} -else +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_24 = lean_unsigned_to_nat(1u); +x_25 = l_Lean_Syntax_getArg(x_11, x_24); +x_26 = lean_unsigned_to_nat(2u); +x_27 = l_Lean_Syntax_getArg(x_11, x_26); +x_28 = l_Lean_Syntax_isNone(x_27); +if (x_28 == 0) { -lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; -x_49 = l_Lean_Syntax_getArg(x_43, x_12); -lean_dec(x_43); -x_50 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___closed__4; -lean_inc(x_22); -x_51 = lean_name_mk_string(x_22, x_50); -lean_inc(x_49); -x_52 = l_Lean_Syntax_isOfKind(x_49, x_51); -lean_dec(x_51); -if (x_52 == 0) +uint8_t x_29; +lean_inc(x_27); +x_29 = l_Lean_Syntax_matchesNull(x_27, x_24); +if (x_29 == 0) { -lean_object* x_53; lean_object* x_54; -lean_dec(x_49); -lean_free_object(x_37); -lean_dec(x_41); -lean_dec(x_28); -lean_dec(x_22); +lean_object* x_30; lean_object* x_31; +lean_dec(x_27); +lean_dec(x_25); +lean_dec(x_19); lean_dec(x_13); lean_dec(x_11); lean_dec(x_7); lean_dec(x_2); lean_dec(x_1); -x_53 = lean_box(1); -x_54 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_8); -return x_54; -} -else -{ -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_55 = l_Lean_Syntax_getArg(x_49, x_33); -lean_dec(x_49); -lean_ctor_set(x_37, 0, x_55); -x_56 = lean_box(0); -lean_inc(x_7); -lean_inc(x_2); -x_57 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4(x_11, x_22, x_2, x_28, x_41, x_3, x_56, x_37, x_7, x_8); -x_58 = lean_ctor_get(x_57, 0); -lean_inc(x_58); -x_59 = lean_ctor_get(x_57, 1); -lean_inc(x_59); -lean_dec(x_57); -x_14 = x_58; -x_15 = x_59; -goto block_20; -} -} +x_30 = lean_box(1); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_8); +return x_31; } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; -lean_dec(x_43); -lean_free_object(x_37); -x_60 = lean_box(0); -x_61 = lean_box(0); -lean_inc(x_7); +lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_32 = l_Lean_Syntax_getArg(x_27, x_12); +lean_dec(x_27); +x_33 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___closed__3; lean_inc(x_2); -x_62 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4(x_11, x_22, x_2, x_28, x_41, x_3, x_61, x_60, x_7, x_8); -x_63 = lean_ctor_get(x_62, 0); -lean_inc(x_63); -x_64 = lean_ctor_get(x_62, 1); -lean_inc(x_64); -lean_dec(x_62); -x_14 = x_63; -x_15 = x_64; -goto block_20; -} -} -else -{ -lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; -x_65 = lean_ctor_get(x_37, 0); -lean_inc(x_65); -lean_dec(x_37); -x_66 = lean_unsigned_to_nat(2u); -x_67 = l_Lean_Syntax_getArg(x_11, x_66); -x_68 = l_Lean_Syntax_isNone(x_67); -if (x_68 == 0) -{ -lean_object* x_69; uint8_t x_70; -x_69 = l_Lean_nullKind; -lean_inc(x_67); -x_70 = l_Lean_Syntax_isNodeOf(x_67, x_69, x_33); -if (x_70 == 0) -{ -lean_object* x_71; lean_object* x_72; -lean_dec(x_67); -lean_dec(x_65); -lean_dec(x_28); -lean_dec(x_22); -lean_dec(x_13); -lean_dec(x_11); -lean_dec(x_7); -lean_dec(x_2); -lean_dec(x_1); -x_71 = lean_box(1); -x_72 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_72, 0, x_71); -lean_ctor_set(x_72, 1, x_8); -return x_72; -} -else -{ -lean_object* x_73; lean_object* x_74; lean_object* x_75; uint8_t x_76; -x_73 = l_Lean_Syntax_getArg(x_67, x_12); -lean_dec(x_67); -x_74 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___closed__4; -lean_inc(x_22); -x_75 = lean_name_mk_string(x_22, x_74); -lean_inc(x_73); -x_76 = l_Lean_Syntax_isOfKind(x_73, x_75); -lean_dec(x_75); -if (x_76 == 0) +x_34 = lean_name_mk_string(x_2, x_33); +lean_inc(x_32); +x_35 = l_Lean_Syntax_isOfKind(x_32, x_34); +lean_dec(x_34); +if (x_35 == 0) { -lean_object* x_77; lean_object* x_78; -lean_dec(x_73); -lean_dec(x_65); -lean_dec(x_28); -lean_dec(x_22); +lean_object* x_36; lean_object* x_37; +lean_dec(x_32); +lean_dec(x_25); +lean_dec(x_19); lean_dec(x_13); lean_dec(x_11); lean_dec(x_7); lean_dec(x_2); lean_dec(x_1); -x_77 = lean_box(1); -x_78 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_78, 0, x_77); -lean_ctor_set(x_78, 1, x_8); -return x_78; +x_36 = lean_box(1); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_8); +return x_37; } else { -lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; -x_79 = l_Lean_Syntax_getArg(x_73, x_33); -lean_dec(x_73); -x_80 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_80, 0, x_79); -x_81 = lean_box(0); +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; size_t x_44; size_t x_45; lean_object* x_46; +x_38 = l_Lean_Syntax_getArg(x_32, x_24); +lean_dec(x_32); +x_39 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_39, 0, x_38); +x_40 = lean_box(0); lean_inc(x_7); +lean_inc(x_1); lean_inc(x_2); -x_82 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4(x_11, x_22, x_2, x_28, x_65, x_3, x_81, x_80, x_7, x_8); -x_83 = lean_ctor_get(x_82, 0); -lean_inc(x_83); -x_84 = lean_ctor_get(x_82, 1); -lean_inc(x_84); -lean_dec(x_82); -x_14 = x_83; -x_15 = x_84; -goto block_20; +x_41 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3(x_11, x_25, x_2, x_1, x_19, x_40, x_39, x_7, x_8); +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_41, 1); +lean_inc(x_43); +lean_dec(x_41); +x_44 = 1; +x_45 = lean_usize_add(x_5, x_44); +x_46 = lean_array_uset(x_13, x_5, x_42); +x_5 = x_45; +x_6 = x_46; +x_8 = x_43; +goto _start; } } } else { -lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -lean_dec(x_67); -x_85 = lean_box(0); -x_86 = lean_box(0); +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; size_t x_53; size_t x_54; lean_object* x_55; +lean_dec(x_27); +x_48 = lean_box(0); +x_49 = lean_box(0); lean_inc(x_7); +lean_inc(x_1); lean_inc(x_2); -x_87 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4(x_11, x_22, x_2, x_28, x_65, x_3, x_86, x_85, x_7, x_8); -x_88 = lean_ctor_get(x_87, 0); -lean_inc(x_88); -x_89 = lean_ctor_get(x_87, 1); -lean_inc(x_89); -lean_dec(x_87); -x_14 = x_88; -x_15 = x_89; -goto block_20; -} -} -} +x_50 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3(x_11, x_25, x_2, x_1, x_19, x_49, x_48, x_7, x_8); +x_51 = lean_ctor_get(x_50, 0); +lean_inc(x_51); +x_52 = lean_ctor_get(x_50, 1); +lean_inc(x_52); +lean_dec(x_50); +x_53 = 1; +x_54 = lean_usize_add(x_5, x_53); +x_55 = lean_array_uset(x_13, x_5, x_51); +x_5 = x_54; +x_6 = x_55; +x_8 = x_52; +goto _start; } } -block_20: -{ -size_t x_16; size_t x_17; lean_object* x_18; -x_16 = 1; -x_17 = lean_usize_add(x_5, x_16); -x_18 = lean_array_uset(x_13, x_5, x_14); -x_5 = x_17; -x_6 = x_18; -x_8 = x_15; -goto _start; } } } @@ -6662,7 +6577,7 @@ x_9 = l_Lean_Syntax_getArg(x_3, x_8); lean_dec(x_3); x_10 = l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_isMultiConstant_x3f___spec__3___closed__5; x_11 = lean_name_mk_string(x_2, x_10); -x_12 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___closed__1; +x_12 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__1; x_13 = lean_name_mk_string(x_11, x_12); lean_inc(x_9); x_14 = l_Lean_Syntax_isOfKind(x_9, x_13); @@ -6711,12 +6626,24 @@ return x_1; static lean_object* _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__4() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(2); +x_2 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__3; +x_3 = lean_alloc_ctor(2, 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_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__5() { +_start: +{ lean_object* x_1; x_1 = lean_mk_string_from_bytes("optEllipsis", 11); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__5() { +static lean_object* _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__6() { _start: { lean_object* x_1; @@ -6724,7 +6651,7 @@ x_1 = lean_mk_string_from_bytes("}", 1); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__6() { +static lean_object* _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__7() { _start: { lean_object* x_1; lean_object* x_2; @@ -6736,241 +6663,240 @@ return x_2; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___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) { _start: { -lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; -x_8 = lean_array_get_size(x_1); -x_9 = lean_usize_of_nat(x_8); -lean_dec(x_8); -x_10 = 0; +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; lean_object* x_14; +x_8 = l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_isMultiConstant_x3f___spec__3___closed__5; +lean_inc(x_1); +x_9 = lean_name_mk_string(x_1, x_8); +x_10 = lean_box(0); +x_11 = lean_array_get_size(x_2); +x_12 = lean_usize_of_nat(x_11); +lean_dec(x_11); +x_13 = 0; lean_inc(x_6); -lean_inc(x_2); -x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1(x_2, x_9, x_10, x_1, x_6, x_7); -if (lean_obj_tag(x_11) == 0) +x_14 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3(x_1, x_12, x_13, x_2, x_6, x_7); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_12; lean_object* x_13; lean_object* x_14; size_t x_15; lean_object* x_16; -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = lean_array_get_size(x_12); -x_15 = lean_usize_of_nat(x_14); +lean_object* x_15; lean_object* x_16; lean_object* x_17; size_t x_18; lean_object* x_19; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); lean_dec(x_14); +x_17 = lean_array_get_size(x_15); +x_18 = lean_usize_of_nat(x_17); +lean_dec(x_17); lean_inc(x_6); +lean_inc(x_9); lean_inc(x_3); -lean_inc(x_2); -x_16 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2(x_2, x_3, x_10, x_15, x_10, x_12, x_6, x_13); -if (lean_obj_tag(x_16) == 0) +x_19 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4(x_3, x_9, x_10, x_18, x_13, x_15, x_6, x_16); +if (lean_obj_tag(x_19) == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec(x_16); +lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); lean_inc(x_6); -x_19 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_6, x_18); -x_20 = !lean_is_exclusive(x_19); -if (x_20 == 0) +x_22 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_6, x_21); +x_23 = !lean_is_exclusive(x_22); +if (x_23 == 0) { -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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_21 = lean_ctor_get(x_19, 0); -x_22 = lean_ctor_get(x_19, 1); -x_23 = l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_isMultiConstant_x3f___spec__3___closed__5; -x_24 = lean_name_mk_string(x_2, x_23); -x_25 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__1; +lean_object* 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_24 = lean_ctor_get(x_22, 0); +x_25 = lean_ctor_get(x_22, 1); +x_26 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__1; +lean_inc(x_9); +x_27 = lean_name_mk_string(x_9, x_26); +x_28 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__2; lean_inc(x_24); -x_26 = lean_name_mk_string(x_24, x_25); -x_27 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__2; -lean_inc(x_21); -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_21); -lean_ctor_set(x_28, 1, x_27); -x_29 = lean_box(2); -x_30 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__4; +x_29 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_29, 0, x_24); +lean_ctor_set(x_29, 1, x_28); +x_30 = lean_box(2); +x_31 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__4; lean_inc(x_3); -x_31 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_30); -lean_ctor_set(x_31, 2, x_3); -x_32 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__3; -x_33 = l_Lean_Syntax_SepArray_ofElems(x_32, x_17); -lean_dec(x_17); -x_34 = l_Array_append___rarg(x_3, x_33); -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_29); -lean_ctor_set(x_35, 1, x_30); -lean_ctor_set(x_35, 2, x_34); -x_36 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__4; -x_37 = lean_name_mk_string(x_24, x_36); -x_38 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__5; -lean_inc(x_31); -x_39 = lean_array_push(x_38, x_31); -x_40 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_40, 0, x_29); -lean_ctor_set(x_40, 1, x_37); -lean_ctor_set(x_40, 2, x_39); -x_41 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__5; -x_42 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_42, 0, x_21); -lean_ctor_set(x_42, 1, x_41); -x_43 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__6; -x_44 = lean_array_push(x_43, x_28); -lean_inc(x_31); -x_45 = lean_array_push(x_44, x_31); -x_46 = lean_array_push(x_45, x_35); -x_47 = lean_array_push(x_46, x_40); -x_48 = lean_array_push(x_47, x_31); -x_49 = lean_array_push(x_48, x_42); -x_50 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_50, 0, x_29); -lean_ctor_set(x_50, 1, x_26); -lean_ctor_set(x_50, 2, x_49); +x_32 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set(x_32, 2, x_3); +x_33 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__4; +x_34 = l_Lean_mkSepArray(x_20, x_33); +lean_dec(x_20); +x_35 = l_Array_append___rarg(x_3, x_34); +x_36 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_36, 0, x_30); +lean_ctor_set(x_36, 1, x_31); +lean_ctor_set(x_36, 2, x_35); +x_37 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__5; +x_38 = lean_name_mk_string(x_9, x_37); +x_39 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__5; +lean_inc(x_32); +x_40 = lean_array_push(x_39, x_32); +x_41 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_41, 0, x_30); +lean_ctor_set(x_41, 1, x_38); +lean_ctor_set(x_41, 2, x_40); +x_42 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__6; +x_43 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_43, 0, x_24); +lean_ctor_set(x_43, 1, x_42); +x_44 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__7; +x_45 = lean_array_push(x_44, x_29); +lean_inc(x_32); +x_46 = lean_array_push(x_45, x_32); +x_47 = lean_array_push(x_46, x_36); +x_48 = lean_array_push(x_47, x_41); +x_49 = lean_array_push(x_48, x_32); +x_50 = lean_array_push(x_49, x_43); +x_51 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_51, 0, x_30); +lean_ctor_set(x_51, 1, x_27); +lean_ctor_set(x_51, 2, x_50); if (lean_obj_tag(x_5) == 0) { lean_dec(x_6); -lean_ctor_set(x_19, 0, x_50); -return x_19; +lean_ctor_set(x_22, 0, x_51); +return x_22; } else { -lean_object* x_51; lean_object* x_52; -lean_free_object(x_19); -x_51 = lean_ctor_get(x_5, 0); -lean_inc(x_51); +lean_object* x_52; lean_object* x_53; +lean_free_object(x_22); +x_52 = lean_ctor_get(x_5, 0); +lean_inc(x_52); lean_dec(x_5); -x_52 = l_Lean_Elab_Term_expandWhereDecls(x_51, x_50, x_6, x_22); -return x_52; +x_53 = l_Lean_Elab_Term_expandWhereDecls(x_52, x_51, x_6, x_25); +return x_53; } } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; -x_53 = lean_ctor_get(x_19, 0); -x_54 = lean_ctor_get(x_19, 1); +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_54 = lean_ctor_get(x_22, 0); +x_55 = lean_ctor_get(x_22, 1); +lean_inc(x_55); lean_inc(x_54); -lean_inc(x_53); -lean_dec(x_19); -x_55 = l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_isMultiConstant_x3f___spec__3___closed__5; -x_56 = lean_name_mk_string(x_2, x_55); -x_57 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__1; -lean_inc(x_56); -x_58 = lean_name_mk_string(x_56, x_57); -x_59 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__2; -lean_inc(x_53); -x_60 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_60, 0, x_53); -lean_ctor_set(x_60, 1, x_59); -x_61 = lean_box(2); -x_62 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__4; +lean_dec(x_22); +x_56 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__1; +lean_inc(x_9); +x_57 = lean_name_mk_string(x_9, x_56); +x_58 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__2; +lean_inc(x_54); +x_59 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_59, 0, x_54); +lean_ctor_set(x_59, 1, x_58); +x_60 = lean_box(2); +x_61 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__4; lean_inc(x_3); -x_63 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_63, 0, x_61); -lean_ctor_set(x_63, 1, x_62); -lean_ctor_set(x_63, 2, x_3); -x_64 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__3; -x_65 = l_Lean_Syntax_SepArray_ofElems(x_64, x_17); -lean_dec(x_17); -x_66 = l_Array_append___rarg(x_3, x_65); -x_67 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_67, 0, x_61); -lean_ctor_set(x_67, 1, x_62); -lean_ctor_set(x_67, 2, x_66); -x_68 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__4; -x_69 = lean_name_mk_string(x_56, x_68); -x_70 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__5; -lean_inc(x_63); -x_71 = lean_array_push(x_70, x_63); -x_72 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_72, 0, x_61); -lean_ctor_set(x_72, 1, x_69); -lean_ctor_set(x_72, 2, x_71); -x_73 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__5; -x_74 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_74, 0, x_53); -lean_ctor_set(x_74, 1, x_73); -x_75 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__6; -x_76 = lean_array_push(x_75, x_60); -lean_inc(x_63); -x_77 = lean_array_push(x_76, x_63); -x_78 = lean_array_push(x_77, x_67); -x_79 = lean_array_push(x_78, x_72); -x_80 = lean_array_push(x_79, x_63); -x_81 = lean_array_push(x_80, x_74); -x_82 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_82, 0, x_61); -lean_ctor_set(x_82, 1, x_58); -lean_ctor_set(x_82, 2, x_81); +x_62 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_62, 0, x_60); +lean_ctor_set(x_62, 1, x_61); +lean_ctor_set(x_62, 2, x_3); +x_63 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__4; +x_64 = l_Lean_mkSepArray(x_20, x_63); +lean_dec(x_20); +x_65 = l_Array_append___rarg(x_3, x_64); +x_66 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_66, 0, x_60); +lean_ctor_set(x_66, 1, x_61); +lean_ctor_set(x_66, 2, x_65); +x_67 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__5; +x_68 = lean_name_mk_string(x_9, x_67); +x_69 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__5; +lean_inc(x_62); +x_70 = lean_array_push(x_69, x_62); +x_71 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_71, 0, x_60); +lean_ctor_set(x_71, 1, x_68); +lean_ctor_set(x_71, 2, x_70); +x_72 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__6; +x_73 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_73, 0, x_54); +lean_ctor_set(x_73, 1, x_72); +x_74 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__7; +x_75 = lean_array_push(x_74, x_59); +lean_inc(x_62); +x_76 = lean_array_push(x_75, x_62); +x_77 = lean_array_push(x_76, x_66); +x_78 = lean_array_push(x_77, x_71); +x_79 = lean_array_push(x_78, x_62); +x_80 = lean_array_push(x_79, x_73); +x_81 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_81, 0, x_60); +lean_ctor_set(x_81, 1, x_57); +lean_ctor_set(x_81, 2, x_80); if (lean_obj_tag(x_5) == 0) { -lean_object* x_83; +lean_object* x_82; lean_dec(x_6); -x_83 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_83, 0, x_82); -lean_ctor_set(x_83, 1, x_54); -return x_83; +x_82 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_82, 0, x_81); +lean_ctor_set(x_82, 1, x_55); +return x_82; } else { -lean_object* x_84; lean_object* x_85; -x_84 = lean_ctor_get(x_5, 0); -lean_inc(x_84); +lean_object* x_83; lean_object* x_84; +x_83 = lean_ctor_get(x_5, 0); +lean_inc(x_83); lean_dec(x_5); -x_85 = l_Lean_Elab_Term_expandWhereDecls(x_84, x_82, x_6, x_54); -return x_85; +x_84 = l_Lean_Elab_Term_expandWhereDecls(x_83, x_81, x_6, x_55); +return x_84; } } } else { -uint8_t x_86; +uint8_t x_85; +lean_dec(x_9); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -lean_dec(x_2); -x_86 = !lean_is_exclusive(x_16); -if (x_86 == 0) +x_85 = !lean_is_exclusive(x_19); +if (x_85 == 0) { -return x_16; +return x_19; } else { -lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_87 = lean_ctor_get(x_16, 0); -x_88 = lean_ctor_get(x_16, 1); -lean_inc(x_88); +lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_86 = lean_ctor_get(x_19, 0); +x_87 = lean_ctor_get(x_19, 1); lean_inc(x_87); -lean_dec(x_16); -x_89 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_89, 0, x_87); -lean_ctor_set(x_89, 1, x_88); -return x_89; +lean_inc(x_86); +lean_dec(x_19); +x_88 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_88, 0, x_86); +lean_ctor_set(x_88, 1, x_87); +return x_88; } } } else { -uint8_t x_90; +uint8_t x_89; +lean_dec(x_9); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -lean_dec(x_2); -x_90 = !lean_is_exclusive(x_11); -if (x_90 == 0) +x_89 = !lean_is_exclusive(x_14); +if (x_89 == 0) { -return x_11; +return x_14; } else { -lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_91 = lean_ctor_get(x_11, 0); -x_92 = lean_ctor_get(x_11, 1); -lean_inc(x_92); +lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_90 = lean_ctor_get(x_14, 0); +x_91 = lean_ctor_get(x_14, 1); lean_inc(x_91); -lean_dec(x_11); -x_93 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_93, 0, x_91); -lean_ctor_set(x_93, 1, x_92); -return x_93; +lean_inc(x_90); +lean_dec(x_14); +x_92 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_91); +return x_92; } } } @@ -7059,46 +6985,46 @@ lean_closure_set(x_16, 0, x_14); lean_closure_set(x_16, 1, x_15); if (x_13 == 0) { -lean_object* x_64; +lean_object* x_62; lean_dec(x_11); lean_dec(x_10); -x_64 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; -x_17 = x_64; -goto block_63; +x_62 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; +x_17 = x_62; +goto block_61; } else { -uint8_t x_65; -x_65 = lean_nat_dec_le(x_11, x_11); -if (x_65 == 0) +uint8_t x_63; +x_63 = lean_nat_dec_le(x_11, x_11); +if (x_63 == 0) { -lean_object* x_66; +lean_object* x_64; lean_dec(x_11); lean_dec(x_10); -x_66 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; -x_17 = x_66; -goto block_63; +x_64 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; +x_17 = x_64; +goto block_61; } else { -size_t x_67; size_t x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_67 = 0; -x_68 = lean_usize_of_nat(x_11); +size_t x_65; size_t x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_65 = 0; +x_66 = lean_usize_of_nat(x_11); lean_dec(x_11); -x_69 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__3; -x_70 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(x_10, x_67, x_68, x_69); +x_67 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__3; +x_68 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(x_10, x_65, x_66, x_67); lean_dec(x_10); -x_71 = lean_ctor_get(x_70, 1); -lean_inc(x_71); -lean_dec(x_70); -x_17 = x_71; -goto block_63; +x_69 = lean_ctor_get(x_68, 1); +lean_inc(x_69); +lean_dec(x_68); +x_17 = x_69; +goto block_61; } } -block_63: +block_61: { lean_object* x_18; -x_18 = l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__1(x_17, x_16); +x_18 = l_Array_sequenceMap___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1(x_17, x_16); lean_dec(x_17); if (lean_obj_tag(x_18) == 0) { @@ -7125,144 +7051,160 @@ lean_dec(x_1); x_25 = l_Lean_Syntax_isNone(x_24); if (x_25 == 0) { -lean_object* x_26; uint8_t x_27; -x_26 = l_Lean_nullKind; +uint8_t x_26; lean_inc(x_24); -x_27 = l_Lean_Syntax_isNodeOf(x_24, x_26, x_8); -if (x_27 == 0) +x_26 = l_Lean_Syntax_matchesNull(x_24, x_8); +if (x_26 == 0) { -lean_object* x_28; lean_object* x_29; +lean_object* x_27; lean_object* x_28; lean_dec(x_24); lean_free_object(x_18); lean_dec(x_22); lean_dec(x_2); -x_28 = lean_box(1); -x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_3); -return x_29; +x_27 = lean_box(1); +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_3); +return x_28; } else { -lean_object* x_30; lean_object* x_31; uint8_t x_32; -x_30 = l_Lean_Syntax_getArg(x_24, x_12); +lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_29 = l_Lean_Syntax_getArg(x_24, x_12); lean_dec(x_24); -x_31 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__5; -lean_inc(x_30); -x_32 = l_Lean_Syntax_isOfKind(x_30, x_31); -if (x_32 == 0) +x_30 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__5; +lean_inc(x_29); +x_31 = l_Lean_Syntax_isOfKind(x_29, x_30); +if (x_31 == 0) { -lean_object* x_33; lean_object* x_34; -lean_dec(x_30); +lean_object* x_32; lean_object* x_33; +lean_dec(x_29); lean_free_object(x_18); lean_dec(x_22); lean_dec(x_2); -x_33 = lean_box(1); -x_34 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_3); -return x_34; +x_32 = lean_box(1); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_3); +return x_33; } else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; -lean_ctor_set(x_18, 0, x_30); -x_35 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; -x_36 = lean_box(0); -x_37 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2(x_22, x_15, x_35, x_36, x_18, x_2, x_3); -return x_37; +lean_object* x_34; lean_object* x_35; lean_object* x_36; +lean_ctor_set(x_18, 0, x_29); +x_34 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; +x_35 = lean_box(0); +x_36 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2(x_15, x_22, x_34, x_35, x_18, x_2, x_3); +return x_36; } } } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_dec(x_24); lean_free_object(x_18); -x_38 = lean_box(0); -x_39 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; -x_40 = lean_box(0); -x_41 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2(x_22, x_15, x_39, x_40, x_38, x_2, x_3); -return x_41; +x_37 = lean_box(0); +x_38 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; +x_39 = lean_box(0); +x_40 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2(x_15, x_22, x_38, x_39, x_37, x_2, x_3); +return x_40; } } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; -x_42 = lean_ctor_get(x_18, 0); -lean_inc(x_42); +lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; +x_41 = lean_ctor_get(x_18, 0); +lean_inc(x_41); lean_dec(x_18); -x_43 = lean_unsigned_to_nat(2u); -x_44 = l_Lean_Syntax_getArg(x_1, x_43); +x_42 = lean_unsigned_to_nat(2u); +x_43 = l_Lean_Syntax_getArg(x_1, x_42); lean_dec(x_1); -x_45 = l_Lean_Syntax_isNone(x_44); -if (x_45 == 0) +x_44 = l_Lean_Syntax_isNone(x_43); +if (x_44 == 0) { -lean_object* x_46; uint8_t x_47; -x_46 = l_Lean_nullKind; -lean_inc(x_44); -x_47 = l_Lean_Syntax_isNodeOf(x_44, x_46, x_8); -if (x_47 == 0) +uint8_t x_45; +lean_inc(x_43); +x_45 = l_Lean_Syntax_matchesNull(x_43, x_8); +if (x_45 == 0) { -lean_object* x_48; lean_object* x_49; -lean_dec(x_44); -lean_dec(x_42); +lean_object* x_46; lean_object* x_47; +lean_dec(x_43); +lean_dec(x_41); lean_dec(x_2); -x_48 = lean_box(1); -x_49 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set(x_49, 1, x_3); -return x_49; +x_46 = lean_box(1); +x_47 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_3); +return x_47; } else { -lean_object* x_50; lean_object* x_51; uint8_t x_52; -x_50 = l_Lean_Syntax_getArg(x_44, x_12); -lean_dec(x_44); -x_51 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__5; -lean_inc(x_50); -x_52 = l_Lean_Syntax_isOfKind(x_50, x_51); -if (x_52 == 0) +lean_object* x_48; lean_object* x_49; uint8_t x_50; +x_48 = l_Lean_Syntax_getArg(x_43, x_12); +lean_dec(x_43); +x_49 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__5; +lean_inc(x_48); +x_50 = l_Lean_Syntax_isOfKind(x_48, x_49); +if (x_50 == 0) { -lean_object* x_53; lean_object* x_54; -lean_dec(x_50); -lean_dec(x_42); +lean_object* x_51; lean_object* x_52; +lean_dec(x_48); +lean_dec(x_41); lean_dec(x_2); -x_53 = lean_box(1); -x_54 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_3); -return x_54; +x_51 = lean_box(1); +x_52 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_3); +return x_52; } else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_55 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_55, 0, x_50); -x_56 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; -x_57 = lean_box(0); -x_58 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2(x_42, x_15, x_56, x_57, x_55, x_2, x_3); -return x_58; +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_53 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_53, 0, x_48); +x_54 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; +x_55 = lean_box(0); +x_56 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2(x_15, x_41, x_54, x_55, x_53, x_2, x_3); +return x_56; } } } else { -lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; -lean_dec(x_44); +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +lean_dec(x_43); +x_57 = lean_box(0); +x_58 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; x_59 = lean_box(0); -x_60 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; -x_61 = lean_box(0); -x_62 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2(x_42, x_15, x_60, x_61, x_59, x_2, x_3); -return x_62; +x_60 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2(x_15, x_41, x_58, x_59, x_57, x_2, x_3); +return x_60; +} +} } } } } } +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___boxed(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; +x_6 = l_Array_sequenceMap_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_1); +return x_6; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__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_EXPORT lean_object* l_Array_sequenceMap___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Array_sequenceMap___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___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) { _start: { size_t x_7; size_t x_8; lean_object* x_9; @@ -7270,51 +7212,30 @@ x_7 = lean_unbox_usize(x_2); lean_dec(x_2); x_8 = lean_unbox_usize(x_3); lean_dec(x_3); -x_9 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1(x_1, x_7, x_8, x_4, x_5, x_6); +x_9 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3(x_1, x_7, x_8, x_4, x_5, x_6); return x_9; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___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, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___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) { _start: { lean_object* x_8; -x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_4); return x_8; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__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_object* x_9) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___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_10; lean_object* x_11; +size_t x_9; size_t x_10; lean_object* x_11; +x_9 = lean_unbox_usize(x_4); +lean_dec(x_4); x_10 = lean_unbox_usize(x_5); lean_dec(x_5); -x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__3(x_1, x_2, x_3, x_4, x_10, x_6, x_7, x_8, x_9); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___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_object* x_9, lean_object* x_10) { -_start: -{ -size_t x_11; lean_object* x_12; -x_11 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_12 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4(x_1, x_2, x_3, x_4, x_5, x_11, x_7, x_8, x_9, x_10); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__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, lean_object* x_7, lean_object* x_8) { -_start: -{ -size_t x_9; size_t x_10; size_t x_11; lean_object* x_12; -x_9 = lean_unbox_usize(x_3); +x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4(x_1, x_2, x_3, x_9, x_10, x_6, x_7, x_8); lean_dec(x_3); -x_10 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_11 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_12 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2(x_1, x_2, x_9, x_10, x_11, x_6, x_7, x_8); -return x_12; +return x_11; } } LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___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, lean_object* x_7) { @@ -25673,6 +25594,7 @@ lean_dec(x_41); x_43 = l_Lean_Elab_Term_processDefDeriving___closed__1; x_44 = lean_name_append_before(x_2, x_43); x_45 = l_Lean_Name_getString_x21(x_1); +lean_dec(x_1); x_46 = lean_name_append_after(x_44, x_45); x_47 = lean_alloc_closure((void*)(l_Lean_Elab_mkUnusedBaseName), 3, 1); lean_closure_set(x_47, 0, x_46); @@ -25935,6 +25857,7 @@ lean_dec(x_110); x_112 = l_Lean_Elab_Term_processDefDeriving___closed__1; x_113 = lean_name_append_before(x_2, x_112); x_114 = l_Lean_Name_getString_x21(x_1); +lean_dec(x_1); x_115 = lean_name_append_after(x_113, x_114); x_116 = lean_alloc_closure((void*)(l_Lean_Elab_mkUnusedBaseName), 3, 1); lean_closure_set(x_116, 0, x_115); @@ -26295,6 +26218,7 @@ lean_dec(x_189); x_191 = l_Lean_Elab_Term_processDefDeriving___closed__1; x_192 = lean_name_append_before(x_2, x_191); x_193 = l_Lean_Name_getString_x21(x_1); +lean_dec(x_1); x_194 = lean_name_append_after(x_192, x_193); x_195 = lean_alloc_closure((void*)(l_Lean_Elab_mkUnusedBaseName), 3, 1); lean_closure_set(x_195, 0, x_194); @@ -35254,56 +35178,54 @@ l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_el lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___lambda__4___closed__8); l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___boxed__const__1 = _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___boxed__const__1(); lean_mark_persistent(l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___boxed__const__1); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___closed__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___closed__1(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___closed__1); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___closed__2 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___closed__2(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___closed__2); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___closed__3 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___closed__3(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___closed__3); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___closed__4 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___closed__4(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___closed__4); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___closed__5 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___closed__5(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1___closed__5); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__1(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__1); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__2 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__2(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__2); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__3 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__3(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__3); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__4 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__4(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__4); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__5 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__5(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__5); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__6 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__6(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__6); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__7 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__7(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__2___closed__7); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__3___closed__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__3___closed__1(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__3___closed__1); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__3___closed__2 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__3___closed__2(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__3___closed__2); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__3___closed__3 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__3___closed__3(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__3___closed__3); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__1(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__1); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__2 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__2(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__2); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__3 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__3(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__3); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__4 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__4(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__4); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__5 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__5(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__5); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__6 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__6(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___lambda__4___closed__6); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___closed__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___closed__1(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___closed__1); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___closed__2 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___closed__2(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___closed__2); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___closed__3 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___closed__3(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___closed__3); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___closed__4 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___closed__4(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___closed__4); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__1(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__1); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__2 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__2(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__2); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__3 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__3(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__3); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__4 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__4(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__4); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__5 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__5(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__5); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__1(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__1); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__2 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__2(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__2); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__3 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__3(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__3); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__4 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__4(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__4); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__5 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__5(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__5); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__6 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__6(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__6); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__7 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__7(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__1___closed__7); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__2___closed__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__2___closed__1(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__2___closed__1); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__2___closed__2 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__2___closed__2(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__2___closed__2); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__2___closed__3 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__2___closed__3(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__2___closed__3); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__1(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__1); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__2 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__2(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__2); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__3 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__3(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__3); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__4 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__4(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__4); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__5 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__5(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__5); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__6 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__6(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__3___closed__6); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___closed__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___closed__1(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___closed__1); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___closed__2 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___closed__2(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___closed__2); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___closed__3 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___closed__3(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___closed__3); l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1___closed__1 = _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1___closed__1(); lean_mark_persistent(l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1___closed__1); l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__1 = _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__1(); @@ -35318,6 +35240,8 @@ l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda lean_mark_persistent(l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__5); l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__6 = _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__6(); lean_mark_persistent(l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__6); +l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__7 = _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__7(); +lean_mark_persistent(l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__7); l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__1 = _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__1(); lean_mark_persistent(l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__1); l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__2 = _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Notation.c b/stage0/stdlib/Lean/Elab/Notation.c index 2f19d97a5a0..7c912a3d165 100644 --- a/stage0/stdlib/Lean/Elab/Notation.c +++ b/stage0/stdlib/Lean/Elab/Notation.c @@ -17,287 +17,285 @@ static lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___close static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__19; static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__23; LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandNotation___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__29; size_t lean_usize_add(size_t, size_t); -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__34; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__45; static lean_object* l_Lean_Elab_Command_expandNotation___closed__1; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__48; static lean_object* l_Lean_Elab_Command_expandNotation___closed__2; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__46; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__2; lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2(lean_object*, size_t, size_t, lean_object*, 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_Array_mapMUnsafe_map___at_Lean_Elab_Command_removeParentheses___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*); lean_object* l_Std_RBNode_insert___at_Lean_NameSet_insert___spec__1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__25; +static lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__12; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__64; static lean_object* l_Lean_Elab_Command_removeParentheses___closed__4; -extern lean_object* l_Lean_nullKind; lean_object* l_Lean_Macro_getCurrNamespace(lean_object*, lean_object*); lean_object* lean_name_mk_string(lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); lean_object* lean_array_uget(lean_object*, size_t); -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__78; static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_isLocalAttrKind___closed__3; static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__11; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_append___rarg(lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__36; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__80; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__6; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__53; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__42; LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandNotation___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_SourceInfo_fromRef(lean_object*); -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__62; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__54; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__84; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__17; static lean_object* l_Lean_Elab_Command_removeParentheses___closed__2; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__33; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__57; static lean_object* l_Lean_Elab_Command_removeParentheses___closed__1; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__22; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandNotation___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__27; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__40; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__67; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__9; static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__26; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__79; +static lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__11; uint8_t lean_name_eq(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__11; static lean_object* l_Lean_Elab_Command_removeParentheses___closed__3; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__77; lean_object* l_Lean_mkIdentFrom(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__66; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__56; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__85; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__70; lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__35; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__36; static lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__6; -static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__39; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_mkSimpleDelab___spec__3___at_Lean_Elab_Command_mkSimpleDelab___spec__4(size_t, size_t, lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Command_hasDuplicateAntiquot___spec__1___lambda__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__20; lean_object* l_Lean_Syntax_setTailInfo(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__35; lean_object* l_Lean_Syntax_getAntiquotTerm(lean_object*); +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__51; static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_isLocalAttrKind___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Command_hasDuplicateAntiquot___boxed(lean_object*); static lean_object* l_Lean_Elab_Command_hasDuplicateAntiquot___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___spec__2(size_t, size_t, lean_object*, lean_object*, lean_object*); lean_object* lean_string_utf8_byte_size(lean_object*); +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__19; lean_object* l_Lean_Elab_Command_strLitToPattern(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_setHeadInfo(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__16; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_hasDuplicateAntiquot___spec__3(lean_object*, lean_object*, size_t, size_t, lean_object*); -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__43; lean_object* l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_liftM___at_Lean_Elab_Command_mkSimpleDelab___spec__2___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_mkNameFromParserSyntax(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__7; static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__24; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___spec__3___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__50; lean_object* l_Lean_Syntax_mkAntiquotNode(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__13; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__57; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__4; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__5; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__30; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__28; LEAN_EXPORT lean_object* l_Lean_Elab_Command_removeParenthesesAux___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__3; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__27; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__6; +static lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__2; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__80; static lean_object* l___regBuiltin_Lean_Elab_Command_expandNotation___closed__4; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__45; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__9; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__67; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__40; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__83; +static lean_object* l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Command_hasDuplicateAntiquot___spec__1___closed__4; +static lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__8; static lean_object* l___regBuiltin_Lean_Elab_Command_expandNotation_declRange___closed__6; static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__21; -static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__38; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__70; +LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__11; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__65; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__39; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__78; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__22; LEAN_EXPORT lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__8; -lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(size_t, size_t, lean_object*); +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__33; +lean_object* l_Lean_Syntax_mkApp(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_expandNotation(lean_object*); static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__1; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__17; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__29; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__62; LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Command_mkSimpleDelab___spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__1; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__64; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__54; LEAN_EXPORT lean_object* l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Command_hasDuplicateAntiquot___spec__1___lambda__1___boxed(lean_object*, lean_object*); +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Command_hasDuplicateAntiquot___spec__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_numLitKind; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__47; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__60; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__74; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_strLitKind; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_mkSimpleDelab___spec__3(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__2; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__25; LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandNotationItemIntoPattern(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__46; static lean_object* l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Command_hasDuplicateAntiquot___spec__1___closed__2; lean_object* l_Lean_Syntax_getHeadInfo(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_hasDuplicateAntiquot___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__52; LEAN_EXPORT lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___boxed(lean_object*, lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__5; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__74; static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___closed__2; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__14; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__76; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__18; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__50; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__48; static lean_object* l___regBuiltin_Lean_Elab_Command_expandNotation_declRange___closed__2; static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__34; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__15; static lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__7; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__30; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__3; lean_object* l_Nat_repr(lean_object*); -static lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__16; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__36; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__23; static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__31; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__44; lean_object* l_Lean_Syntax_getId(lean_object*); -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__28; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__35; -extern lean_object* l_Lean_choiceKind; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__38; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__21; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__8; +static lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__4; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__41; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__13; static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___closed__3; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__72; lean_object* l_Lean_addBuiltinDeclarationRanges(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__9; static lean_object* l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Command_hasDuplicateAntiquot___spec__1___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Command_expandNotation_declRange___closed__3; LEAN_EXPORT lean_object* l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Command_hasDuplicateAntiquot___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__12; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__47; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__60; LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkSimpleDelab(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__26; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__75; +static lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Command_expandNotation_declRange___closed__1; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__12; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__7; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__71; static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__18; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__63; static lean_object* l___regBuiltin_Lean_Elab_Command_expandNotation_declRange___closed__7; static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__15; extern lean_object* l_Lean_instInhabitedSyntax; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__20; static lean_object* l___regBuiltin_Lean_Elab_Command_expandNotation___closed__2; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__32; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__10; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___spec__3(size_t, size_t, lean_object*); lean_object* l_Lean_evalOptPrio(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___spec__4(lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__81; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___spec__4___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__13; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__58; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__73; +LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__29; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__31; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__59; size_t lean_usize_of_nat(lean_object*); +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__77; uint8_t l_Lean_Syntax_isAntiquot(lean_object*); static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__14; static lean_object* l___regBuiltin_Lean_Elab_Command_expandNotation_declRange___closed__5; extern lean_object* l_Lean_NameSet_empty; static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__28; +static lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__13; lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__65; static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___closed__4; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__39; static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__5; -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__16; +static lean_object* l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Command_hasDuplicateAntiquot___spec__1___closed__3; +lean_object* l_Lean_Syntax_mkNumLit(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__30; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__49; static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___closed__1; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__23; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__72; static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__8; static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__17; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_hasDuplicateAntiquot___spec__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Command_hasDuplicateAntiquot___closed__3; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__38; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__21; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__5; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__24; static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__4; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__14; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___spec__5(lean_object*, size_t, size_t, lean_object*); extern lean_object* l_Lean_Elab_macroAttribute; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__44; -static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__37; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_hasDuplicateAntiquot___lambda__1___boxed(lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__41; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__13; +LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_expandNotation___closed__1; lean_object* l_Lean_Syntax_getArgs(lean_object*); lean_object* l_Lean_Name_append(lean_object*, lean_object*); lean_object* l_Lean_Macro_resolveGlobalName(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__37; static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__2; lean_object* l_Lean_Syntax_getKind(lean_object*); +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__55; static lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__12; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__3; LEAN_EXPORT uint8_t l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___lambda__1(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__52; LEAN_EXPORT lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__18; static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__27; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_mkSimpleDelab___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__9; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__68; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__16; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__15; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__76; +static lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__7; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_expandNotation_declRange(lean_object*); +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__69; LEAN_EXPORT lean_object* l_Lean_Elab_Command_hasDuplicateAntiquot(lean_object*); LEAN_EXPORT uint8_t l_Lean_Elab_Command_hasDuplicateAntiquot___lambda__1(lean_object*); -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__49; LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandNotationItemIntoPattern___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__53; LEAN_EXPORT lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_isLocalAttrKind___boxed(lean_object*); -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__63; uint8_t l_Lean_Syntax_isNone(lean_object*); static lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__5; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_hasDuplicateAntiquot___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__14; static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__10; static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__32; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__61; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__31; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__59; lean_object* l_Lean_Syntax_getTailInfo(lean_object*); static lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandNotation___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__10; static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_isLocalAttrKind___closed__2; static lean_object* l_Lean_Elab_Command_hasDuplicateAntiquot___closed__2; uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__10; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__32; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_mkSimpleDelab___spec__3___at_Lean_Elab_Command_mkSimpleDelab___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__81; -static lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__15; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__71; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__7; -LEAN_EXPORT lean_object* l_liftM___at_Lean_Elab_Command_mkSimpleDelab___spec__2(lean_object*); +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__26; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__34; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__24; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___spec__4(size_t, size_t, lean_object*); +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__58; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_NameSet_contains(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__3; lean_object* lean_mk_syntax_ident(lean_object*); +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__75; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__73; lean_object* l_Lean_Elab_toAttributeKind(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__20; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__51; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__12; +static lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__1; static lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__2; lean_object* l_Array_findIdx_x3f_loop___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__69; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___spec__1(lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__43; static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__25; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__11; static lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__9; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__61; LEAN_EXPORT lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___lambda__1___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__68; LEAN_EXPORT lean_object* l_Lean_Elab_Command_removeParenthesesAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_removeParentheses___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__19; static lean_object* l___regBuiltin_Lean_Elab_Command_expandNotation_declRange___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Command_removeParentheses(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__22; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__66; +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__56; static lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__4; static lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__1; static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__33; LEAN_EXPORT uint8_t l___private_Lean_Elab_Notation_0__Lean_Elab_Command_isLocalAttrKind(lean_object*); -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__37; LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandNotation(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__42; -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__82; static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_isLocalAttrKind___closed__1; static lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__8; lean_object* l_Lean_Elab_Term_expandCDot_x3f(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__79; static lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__6; -lean_object* l_Lean_Syntax_mkLit(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_expandNotation___closed__3; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_mkSimpleDelab___closed__55; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___spec__1(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { @@ -482,6 +480,302 @@ lean_dec(x_1); return x_3; } } +static lean_object* _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Syntax", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("cat", 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___closed__3; +x_2 = lean_string_utf8_byte_size(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___closed__3; +x_2 = lean_unsigned_to_nat(0u); +x_3 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__3; +x_4 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("null", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__5; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(2u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__7; +x_2 = l_Array_append___rarg(x_1, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(2); +x_2 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__6; +x_3 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__9; +x_4 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("precedence", 10); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(":", 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1(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; uint8_t x_7; +lean_inc(x_4); +x_6 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_4, x_5); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 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; 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; +x_8 = lean_ctor_get(x_6, 0); +x_9 = lean_ctor_get(x_4, 2); +lean_inc(x_9); +x_10 = lean_ctor_get(x_4, 1); +lean_inc(x_10); +lean_dec(x_4); +x_11 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__1; +lean_inc(x_1); +x_12 = lean_name_mk_string(x_1, x_11); +x_13 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__2; +x_14 = lean_name_mk_string(x_12, x_13); +x_15 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___closed__4; +x_16 = l_Lean_addMacroScope(x_10, x_15, x_9); +x_17 = lean_box(0); +x_18 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__4; +lean_inc(x_8); +x_19 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_19, 0, x_8); +lean_ctor_set(x_19, 1, x_18); +lean_ctor_set(x_19, 2, x_16); +lean_ctor_set(x_19, 3, x_17); +x_20 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__8; +x_21 = lean_array_push(x_20, x_19); +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +lean_dec(x_8); +lean_dec(x_1); +x_22 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__10; +x_23 = lean_array_push(x_21, x_22); +x_24 = lean_box(2); +x_25 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_14); +lean_ctor_set(x_25, 2, x_23); +lean_ctor_set(x_6, 0, x_25); +return x_6; +} +else +{ +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_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; +x_26 = lean_ctor_get(x_3, 0); +lean_inc(x_26); +lean_dec(x_3); +x_27 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__11; +x_28 = lean_name_mk_string(x_1, x_27); +x_29 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__12; +x_30 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_30, 0, x_8); +lean_ctor_set(x_30, 1, x_29); +x_31 = lean_array_push(x_20, x_30); +x_32 = lean_array_push(x_31, x_26); +x_33 = lean_box(2); +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_28); +lean_ctor_set(x_34, 2, x_32); +x_35 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__13; +x_36 = lean_array_push(x_35, x_34); +x_37 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__7; +x_38 = l_Array_append___rarg(x_37, x_36); +x_39 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__6; +x_40 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_40, 0, x_33); +lean_ctor_set(x_40, 1, x_39); +lean_ctor_set(x_40, 2, x_38); +x_41 = lean_array_push(x_21, x_40); +x_42 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_42, 0, x_33); +lean_ctor_set(x_42, 1, x_14); +lean_ctor_set(x_42, 2, x_41); +lean_ctor_set(x_6, 0, x_42); +return x_6; +} +} +else +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_43 = lean_ctor_get(x_6, 0); +x_44 = lean_ctor_get(x_6, 1); +lean_inc(x_44); +lean_inc(x_43); +lean_dec(x_6); +x_45 = lean_ctor_get(x_4, 2); +lean_inc(x_45); +x_46 = lean_ctor_get(x_4, 1); +lean_inc(x_46); +lean_dec(x_4); +x_47 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__1; +lean_inc(x_1); +x_48 = lean_name_mk_string(x_1, x_47); +x_49 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__2; +x_50 = lean_name_mk_string(x_48, x_49); +x_51 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___closed__4; +x_52 = l_Lean_addMacroScope(x_46, x_51, x_45); +x_53 = lean_box(0); +x_54 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__4; +lean_inc(x_43); +x_55 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_55, 0, x_43); +lean_ctor_set(x_55, 1, x_54); +lean_ctor_set(x_55, 2, x_52); +lean_ctor_set(x_55, 3, x_53); +x_56 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__8; +x_57 = lean_array_push(x_56, x_55); +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; +lean_dec(x_43); +lean_dec(x_1); +x_58 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__10; +x_59 = lean_array_push(x_57, x_58); +x_60 = lean_box(2); +x_61 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set(x_61, 1, x_50); +lean_ctor_set(x_61, 2, x_59); +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_61); +lean_ctor_set(x_62, 1, x_44); +return x_62; +} +else +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; 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; lean_object* x_79; lean_object* x_80; +x_63 = lean_ctor_get(x_3, 0); +lean_inc(x_63); +lean_dec(x_3); +x_64 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__11; +x_65 = lean_name_mk_string(x_1, x_64); +x_66 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__12; +x_67 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_67, 0, x_43); +lean_ctor_set(x_67, 1, x_66); +x_68 = lean_array_push(x_56, x_67); +x_69 = lean_array_push(x_68, x_63); +x_70 = lean_box(2); +x_71 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set(x_71, 1, x_65); +lean_ctor_set(x_71, 2, x_69); +x_72 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__13; +x_73 = lean_array_push(x_72, x_71); +x_74 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__7; +x_75 = l_Array_append___rarg(x_74, x_73); +x_76 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__6; +x_77 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_77, 0, x_70); +lean_ctor_set(x_77, 1, x_76); +lean_ctor_set(x_77, 2, x_75); +x_78 = lean_array_push(x_57, x_77); +x_79 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_79, 0, x_70); +lean_ctor_set(x_79, 1, x_50); +lean_ctor_set(x_79, 2, x_78); +x_80 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_80, 0, x_79); +lean_ctor_set(x_80, 1, x_44); +return x_80; +} +} +} +} static lean_object* _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__1() { _start: { @@ -558,7 +852,7 @@ static lean_object* _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem__ _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Syntax", 6); +x_1 = lean_mk_string_from_bytes("str", 3); return x_1; } } @@ -566,7 +860,7 @@ static lean_object* _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__4; +x_1 = lean_box(0); x_2 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__9; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -575,131 +869,202 @@ return x_3; static lean_object* _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__11() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("atom", 4); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__12() { -_start: -{ lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__10; -x_2 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__11; +x_1 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__4; +x_2 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__13() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(1u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__14() { +static lean_object* _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__12() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("cat", 3); +x_1 = lean_mk_string_from_bytes("atom", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__15() { +static lean_object* _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__10; -x_2 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__14; +x_1 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__11; +x_2 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__12; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__16() { +static lean_object* _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__14() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(2u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__4; +x_2 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__11; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; } } LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_4; lean_object* x_5; uint8_t x_6; +lean_object* x_4; uint8_t x_5; +x_4 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__8; lean_inc(x_1); -x_4 = l_Lean_Syntax_getKind(x_1); -x_5 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__8; -x_6 = lean_name_eq(x_4, x_5); -if (x_6 == 0) +x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); +if (x_5 == 0) { -lean_object* x_7; uint8_t x_8; -x_7 = l_Lean_strLitKind; -x_8 = lean_name_eq(x_4, x_7); -lean_dec(x_4); -if (x_8 == 0) +lean_object* x_6; uint8_t x_7; +x_6 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__10; +lean_inc(x_1); +x_7 = l_Lean_Syntax_isOfKind(x_1, x_6); +if (x_7 == 0) { -lean_object* x_9; lean_object* x_10; +lean_object* x_8; lean_object* x_9; +lean_dec(x_2); lean_dec(x_1); -x_9 = lean_box(1); -x_10 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_10, 1, x_3); +x_8 = lean_box(1); +x_9 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_3); +return x_9; +} +else +{ +lean_object* x_10; uint8_t x_11; +x_10 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_11 = !lean_is_exclusive(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; lean_object* x_17; +x_12 = lean_ctor_get(x_10, 0); +lean_dec(x_12); +x_13 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__13; +x_14 = lean_array_push(x_13, x_1); +x_15 = lean_box(2); +x_16 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__13; +x_17 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +lean_ctor_set(x_17, 2, x_14); +lean_ctor_set(x_10, 0, x_17); return x_10; } else { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_11 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__13; -x_12 = lean_array_push(x_11, x_1); -x_13 = lean_box(2); -x_14 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__12; -x_15 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_15, 0, x_13); -lean_ctor_set(x_15, 1, x_14); -lean_ctor_set(x_15, 2, x_12); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_3); -return x_16; +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_18 = lean_ctor_get(x_10, 1); +lean_inc(x_18); +lean_dec(x_10); +x_19 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__13; +x_20 = lean_array_push(x_19, x_1); +x_21 = lean_box(2); +x_22 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__13; +x_23 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +lean_ctor_set(x_23, 2, x_20); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_18); +return x_24; +} } } else { -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; lean_object* x_25; lean_object* x_26; lean_object* x_27; -lean_dec(x_4); -x_17 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___closed__4; -lean_inc(x_1); -x_18 = l_Lean_mkIdentFrom(x_1, x_17); -x_19 = lean_unsigned_to_nat(1u); -x_20 = l_Lean_Syntax_getArg(x_1, x_19); +lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_25 = lean_unsigned_to_nat(0u); +x_26 = l_Lean_Syntax_getArg(x_1, x_25); +x_27 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___closed__2; +x_28 = l_Lean_Syntax_isOfKind(x_26, x_27); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; +lean_dec(x_2); lean_dec(x_1); -x_21 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__16; -x_22 = lean_array_push(x_21, x_18); -x_23 = lean_array_push(x_22, x_20); -x_24 = lean_box(2); -x_25 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__15; -x_26 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -lean_ctor_set(x_26, 2, x_23); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_3); -return x_27; +x_29 = lean_box(1); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_3); +return x_30; +} +else +{ +lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_31 = lean_unsigned_to_nat(1u); +x_32 = l_Lean_Syntax_getArg(x_1, x_31); +lean_dec(x_1); +x_33 = l_Lean_Syntax_isNone(x_32); +if (x_33 == 0) +{ +uint8_t x_34; +lean_inc(x_32); +x_34 = l_Lean_Syntax_matchesNull(x_32, x_31); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; +lean_dec(x_32); +lean_dec(x_2); +x_35 = lean_box(1); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_3); +return x_36; +} +else +{ +lean_object* x_37; lean_object* x_38; uint8_t x_39; +x_37 = l_Lean_Syntax_getArg(x_32, x_25); +lean_dec(x_32); +x_38 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__14; +lean_inc(x_37); +x_39 = l_Lean_Syntax_isOfKind(x_37, x_38); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; +lean_dec(x_37); +lean_dec(x_2); +x_40 = lean_box(1); +x_41 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_3); +return x_41; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_42 = l_Lean_Syntax_getArg(x_37, x_31); +lean_dec(x_37); +x_43 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_43, 0, x_42); +x_44 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__4; +x_45 = lean_box(0); +x_46 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1(x_44, x_45, x_43, x_2, x_3); +return x_46; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +lean_dec(x_32); +x_47 = lean_box(0); +x_48 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__4; +x_49 = lean_box(0); +x_50 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1(x_48, x_49, x_47, x_2, x_3); +return x_50; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_4; -x_4 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem(x_1, x_2, x_3); +lean_object* x_6; +x_6 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1(x_1, x_2, x_3, x_4, x_5); lean_dec(x_2); -return x_4; +return x_6; } } LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandNotationItemIntoPattern(lean_object* x_1, lean_object* x_2, lean_object* x_3) { @@ -713,7 +1078,7 @@ x_6 = lean_name_eq(x_4, x_5); if (x_6 == 0) { lean_object* x_7; uint8_t x_8; -x_7 = l_Lean_strLitKind; +x_7 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__10; x_8 = lean_name_eq(x_4, x_7); lean_dec(x_4); if (x_8 == 0) @@ -1092,279 +1457,278 @@ return x_29; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; +lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; x_30 = lean_unsigned_to_nat(1u); x_31 = l_Lean_Syntax_getArg(x_1, x_30); -x_32 = l_Lean_nullKind; -x_33 = lean_unsigned_to_nat(2u); +x_32 = lean_unsigned_to_nat(2u); lean_inc(x_31); -x_34 = l_Lean_Syntax_isNodeOf(x_31, x_32, x_33); -if (x_34 == 0) +x_33 = l_Lean_Syntax_matchesNull(x_31, x_32); +if (x_33 == 0) { lean_dec(x_31); if (lean_obj_tag(x_1) == 1) { -uint8_t x_35; -x_35 = !lean_is_exclusive(x_1); -if (x_35 == 0) +uint8_t x_34; +x_34 = !lean_is_exclusive(x_1); +if (x_34 == 0) { -lean_object* x_36; lean_object* x_37; size_t x_38; size_t x_39; lean_object* x_40; uint8_t x_41; -x_36 = lean_ctor_get(x_1, 2); -x_37 = lean_array_get_size(x_36); -x_38 = lean_usize_of_nat(x_37); -lean_dec(x_37); -x_39 = 0; -x_40 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_removeParentheses___spec__1(x_38, x_39, x_36, x_2, x_3); -x_41 = !lean_is_exclusive(x_40); -if (x_41 == 0) +lean_object* x_35; lean_object* x_36; size_t x_37; size_t x_38; lean_object* x_39; uint8_t x_40; +x_35 = lean_ctor_get(x_1, 2); +x_36 = lean_array_get_size(x_35); +x_37 = lean_usize_of_nat(x_36); +lean_dec(x_36); +x_38 = 0; +x_39 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_removeParentheses___spec__1(x_37, x_38, x_35, x_2, x_3); +x_40 = !lean_is_exclusive(x_39); +if (x_40 == 0) { -lean_object* x_42; -x_42 = lean_ctor_get(x_40, 0); -lean_ctor_set(x_1, 2, x_42); -lean_ctor_set(x_40, 0, x_1); -return x_40; +lean_object* x_41; +x_41 = lean_ctor_get(x_39, 0); +lean_ctor_set(x_1, 2, x_41); +lean_ctor_set(x_39, 0, x_1); +return x_39; } else { -lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_43 = lean_ctor_get(x_40, 0); -x_44 = lean_ctor_get(x_40, 1); -lean_inc(x_44); +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_39, 0); +x_43 = lean_ctor_get(x_39, 1); lean_inc(x_43); -lean_dec(x_40); -lean_ctor_set(x_1, 2, x_43); -x_45 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_45, 0, x_1); -lean_ctor_set(x_45, 1, x_44); -return x_45; +lean_inc(x_42); +lean_dec(x_39); +lean_ctor_set(x_1, 2, x_42); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_1); +lean_ctor_set(x_44, 1, x_43); +return x_44; } } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; size_t x_50; size_t x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_46 = lean_ctor_get(x_1, 0); -x_47 = lean_ctor_get(x_1, 1); -x_48 = lean_ctor_get(x_1, 2); -lean_inc(x_48); +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; size_t x_49; size_t x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_45 = lean_ctor_get(x_1, 0); +x_46 = lean_ctor_get(x_1, 1); +x_47 = lean_ctor_get(x_1, 2); lean_inc(x_47); lean_inc(x_46); +lean_inc(x_45); lean_dec(x_1); -x_49 = lean_array_get_size(x_48); -x_50 = lean_usize_of_nat(x_49); -lean_dec(x_49); -x_51 = 0; -x_52 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_removeParentheses___spec__1(x_50, x_51, x_48, x_2, x_3); -x_53 = lean_ctor_get(x_52, 0); +x_48 = lean_array_get_size(x_47); +x_49 = lean_usize_of_nat(x_48); +lean_dec(x_48); +x_50 = 0; +x_51 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_removeParentheses___spec__1(x_49, x_50, x_47, x_2, x_3); +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_51, 1); lean_inc(x_53); -x_54 = lean_ctor_get(x_52, 1); -lean_inc(x_54); -if (lean_is_exclusive(x_52)) { - lean_ctor_release(x_52, 0); - lean_ctor_release(x_52, 1); - x_55 = x_52; +if (lean_is_exclusive(x_51)) { + lean_ctor_release(x_51, 0); + lean_ctor_release(x_51, 1); + x_54 = x_51; } else { - lean_dec_ref(x_52); - x_55 = lean_box(0); -} -x_56 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_56, 0, x_46); -lean_ctor_set(x_56, 1, x_47); -lean_ctor_set(x_56, 2, x_53); -if (lean_is_scalar(x_55)) { - x_57 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_51); + x_54 = lean_box(0); +} +x_55 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_55, 0, x_45); +lean_ctor_set(x_55, 1, x_46); +lean_ctor_set(x_55, 2, x_52); +if (lean_is_scalar(x_54)) { + x_56 = lean_alloc_ctor(0, 2, 0); } else { - x_57 = x_55; + x_56 = x_54; } -lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_54); -return x_57; +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_53); +return x_56; } } else { -lean_object* x_58; +lean_object* x_57; lean_dec(x_2); -x_58 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_58, 0, x_1); -lean_ctor_set(x_58, 1, x_3); -return x_58; +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_1); +lean_ctor_set(x_57, 1, x_3); +return x_57; } } else { -lean_object* x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; -x_59 = lean_unsigned_to_nat(0u); -x_60 = l_Lean_Syntax_getArg(x_31, x_59); -x_61 = l_Lean_Syntax_getArg(x_31, x_30); -lean_dec(x_31); -x_62 = l_Lean_Syntax_isNodeOf(x_61, x_32, x_59); +lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; +x_58 = lean_unsigned_to_nat(0u); +x_59 = l_Lean_Syntax_getArg(x_31, x_58); +x_60 = l_Lean_Syntax_getArg(x_31, x_30); +lean_dec(x_31); +x_61 = l_Lean_Syntax_matchesNull(x_60, x_58); +if (x_61 == 0) +{ +lean_dec(x_59); +if (lean_obj_tag(x_1) == 1) +{ +uint8_t x_62; +x_62 = !lean_is_exclusive(x_1); if (x_62 == 0) { -lean_dec(x_60); -if (lean_obj_tag(x_1) == 1) +lean_object* x_63; lean_object* x_64; size_t x_65; size_t x_66; lean_object* x_67; uint8_t x_68; +x_63 = lean_ctor_get(x_1, 2); +x_64 = lean_array_get_size(x_63); +x_65 = lean_usize_of_nat(x_64); +lean_dec(x_64); +x_66 = 0; +x_67 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_removeParentheses___spec__1(x_65, x_66, x_63, x_2, x_3); +x_68 = !lean_is_exclusive(x_67); +if (x_68 == 0) { -uint8_t x_63; -x_63 = !lean_is_exclusive(x_1); -if (x_63 == 0) -{ -lean_object* x_64; lean_object* x_65; size_t x_66; size_t x_67; lean_object* x_68; uint8_t x_69; -x_64 = lean_ctor_get(x_1, 2); -x_65 = lean_array_get_size(x_64); -x_66 = lean_usize_of_nat(x_65); -lean_dec(x_65); -x_67 = 0; -x_68 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_removeParentheses___spec__1(x_66, x_67, x_64, x_2, x_3); -x_69 = !lean_is_exclusive(x_68); -if (x_69 == 0) -{ -lean_object* x_70; -x_70 = lean_ctor_get(x_68, 0); -lean_ctor_set(x_1, 2, x_70); -lean_ctor_set(x_68, 0, x_1); -return x_68; +lean_object* x_69; +x_69 = lean_ctor_get(x_67, 0); +lean_ctor_set(x_1, 2, x_69); +lean_ctor_set(x_67, 0, x_1); +return x_67; } else { -lean_object* x_71; lean_object* x_72; lean_object* x_73; -x_71 = lean_ctor_get(x_68, 0); -x_72 = lean_ctor_get(x_68, 1); -lean_inc(x_72); +lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_70 = lean_ctor_get(x_67, 0); +x_71 = lean_ctor_get(x_67, 1); lean_inc(x_71); -lean_dec(x_68); -lean_ctor_set(x_1, 2, x_71); -x_73 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_73, 0, x_1); -lean_ctor_set(x_73, 1, x_72); -return x_73; +lean_inc(x_70); +lean_dec(x_67); +lean_ctor_set(x_1, 2, x_70); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_1); +lean_ctor_set(x_72, 1, x_71); +return x_72; } } else { -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; size_t x_78; size_t x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_74 = lean_ctor_get(x_1, 0); -x_75 = lean_ctor_get(x_1, 1); -x_76 = lean_ctor_get(x_1, 2); -lean_inc(x_76); +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; size_t x_77; size_t x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_73 = lean_ctor_get(x_1, 0); +x_74 = lean_ctor_get(x_1, 1); +x_75 = lean_ctor_get(x_1, 2); lean_inc(x_75); lean_inc(x_74); +lean_inc(x_73); lean_dec(x_1); -x_77 = lean_array_get_size(x_76); -x_78 = lean_usize_of_nat(x_77); -lean_dec(x_77); -x_79 = 0; -x_80 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_removeParentheses___spec__1(x_78, x_79, x_76, x_2, x_3); -x_81 = lean_ctor_get(x_80, 0); +x_76 = lean_array_get_size(x_75); +x_77 = lean_usize_of_nat(x_76); +lean_dec(x_76); +x_78 = 0; +x_79 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_removeParentheses___spec__1(x_77, x_78, x_75, x_2, x_3); +x_80 = lean_ctor_get(x_79, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_79, 1); lean_inc(x_81); -x_82 = lean_ctor_get(x_80, 1); -lean_inc(x_82); -if (lean_is_exclusive(x_80)) { - lean_ctor_release(x_80, 0); - lean_ctor_release(x_80, 1); - x_83 = x_80; +if (lean_is_exclusive(x_79)) { + lean_ctor_release(x_79, 0); + lean_ctor_release(x_79, 1); + x_82 = x_79; } else { - lean_dec_ref(x_80); - x_83 = lean_box(0); + lean_dec_ref(x_79); + x_82 = lean_box(0); } -x_84 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_84, 0, x_74); -lean_ctor_set(x_84, 1, x_75); -lean_ctor_set(x_84, 2, x_81); -if (lean_is_scalar(x_83)) { - x_85 = lean_alloc_ctor(0, 2, 0); +x_83 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_83, 0, x_73); +lean_ctor_set(x_83, 1, x_74); +lean_ctor_set(x_83, 2, x_80); +if (lean_is_scalar(x_82)) { + x_84 = lean_alloc_ctor(0, 2, 0); } else { - x_85 = x_83; + x_84 = x_82; } -lean_ctor_set(x_85, 0, x_84); -lean_ctor_set(x_85, 1, x_82); -return x_85; +lean_ctor_set(x_84, 0, x_83); +lean_ctor_set(x_84, 1, x_81); +return x_84; } } else { -lean_object* x_86; +lean_object* x_85; lean_dec(x_2); -x_86 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_86, 0, x_1); -lean_ctor_set(x_86, 1, x_3); -return x_86; +x_85 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_85, 0, x_1); +lean_ctor_set(x_85, 1, x_3); +return x_85; } } else { -lean_object* x_87; lean_object* x_88; +lean_object* x_86; lean_object* x_87; lean_inc(x_2); -lean_inc(x_60); -x_87 = l_Lean_Elab_Term_expandCDot_x3f(x_60, x_2, x_3); -x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_59); +x_86 = l_Lean_Elab_Term_expandCDot_x3f(x_59, x_2, x_3); +x_87 = lean_ctor_get(x_86, 0); +lean_inc(x_87); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; uint8_t x_90; +x_88 = lean_ctor_get(x_86, 1); lean_inc(x_88); -if (lean_obj_tag(x_88) == 0) -{ -lean_object* x_89; lean_object* x_90; uint8_t x_91; -x_89 = lean_ctor_get(x_87, 1); -lean_inc(x_89); -lean_dec(x_87); -x_90 = l_Lean_Elab_Command_removeParentheses(x_60, x_2, x_89); -x_91 = !lean_is_exclusive(x_90); -if (x_91 == 0) -{ -lean_object* x_92; lean_object* x_93; -x_92 = lean_ctor_get(x_90, 0); -x_93 = l_Lean_Elab_Command_removeParenthesesAux(x_1, x_92); +lean_dec(x_86); +x_89 = l_Lean_Elab_Command_removeParentheses(x_59, x_2, x_88); +x_90 = !lean_is_exclusive(x_89); +if (x_90 == 0) +{ +lean_object* x_91; lean_object* x_92; +x_91 = lean_ctor_get(x_89, 0); +x_92 = l_Lean_Elab_Command_removeParenthesesAux(x_1, x_91); lean_dec(x_1); -lean_ctor_set(x_90, 0, x_93); -return x_90; +lean_ctor_set(x_89, 0, x_92); +return x_89; } else { -lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; -x_94 = lean_ctor_get(x_90, 0); -x_95 = lean_ctor_get(x_90, 1); -lean_inc(x_95); +lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_93 = lean_ctor_get(x_89, 0); +x_94 = lean_ctor_get(x_89, 1); lean_inc(x_94); -lean_dec(x_90); -x_96 = l_Lean_Elab_Command_removeParenthesesAux(x_1, x_94); +lean_inc(x_93); +lean_dec(x_89); +x_95 = l_Lean_Elab_Command_removeParenthesesAux(x_1, x_93); lean_dec(x_1); -x_97 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_97, 0, x_96); -lean_ctor_set(x_97, 1, x_95); -return x_97; +x_96 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_96, 0, x_95); +lean_ctor_set(x_96, 1, x_94); +return x_96; } } else { -lean_object* x_98; lean_object* x_99; lean_object* x_100; uint8_t x_101; -lean_dec(x_60); -x_98 = lean_ctor_get(x_87, 1); +lean_object* x_97; lean_object* x_98; lean_object* x_99; uint8_t x_100; +lean_dec(x_59); +x_97 = lean_ctor_get(x_86, 1); +lean_inc(x_97); +lean_dec(x_86); +x_98 = lean_ctor_get(x_87, 0); lean_inc(x_98); lean_dec(x_87); -x_99 = lean_ctor_get(x_88, 0); -lean_inc(x_99); -lean_dec(x_88); -x_100 = l_Lean_Elab_Command_removeParentheses(x_99, x_2, x_98); -x_101 = !lean_is_exclusive(x_100); -if (x_101 == 0) +x_99 = l_Lean_Elab_Command_removeParentheses(x_98, x_2, x_97); +x_100 = !lean_is_exclusive(x_99); +if (x_100 == 0) { -lean_object* x_102; lean_object* x_103; -x_102 = lean_ctor_get(x_100, 0); -x_103 = l_Lean_Elab_Command_removeParenthesesAux(x_1, x_102); +lean_object* x_101; lean_object* x_102; +x_101 = lean_ctor_get(x_99, 0); +x_102 = l_Lean_Elab_Command_removeParenthesesAux(x_1, x_101); lean_dec(x_1); -lean_ctor_set(x_100, 0, x_103); -return x_100; +lean_ctor_set(x_99, 0, x_102); +return x_99; } else { -lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -x_104 = lean_ctor_get(x_100, 0); -x_105 = lean_ctor_get(x_100, 1); -lean_inc(x_105); +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_103 = lean_ctor_get(x_99, 0); +x_104 = lean_ctor_get(x_99, 1); lean_inc(x_104); -lean_dec(x_100); -x_106 = l_Lean_Elab_Command_removeParenthesesAux(x_1, x_104); +lean_inc(x_103); +lean_dec(x_99); +x_105 = l_Lean_Elab_Command_removeParenthesesAux(x_1, x_103); lean_dec(x_1); -x_107 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_107, 0, x_106); -lean_ctor_set(x_107, 1, x_105); -return x_107; +x_106 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_106, 0, x_105); +lean_ctor_set(x_106, 1, x_104); +return x_106; } } } @@ -1530,6 +1894,24 @@ return x_1; static lean_object* _init_l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Command_hasDuplicateAntiquot___spec__1___closed__2() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("choice", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Command_hasDuplicateAntiquot___spec__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Command_hasDuplicateAntiquot___spec__1___closed__2; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Command_hasDuplicateAntiquot___spec__1___closed__4() { +_start: +{ uint8_t x_1; lean_object* x_2; lean_object* x_3; x_1 = 1; x_2 = lean_box(x_1); @@ -1581,7 +1963,7 @@ lean_object* x_43; lean_object* x_44; lean_dec(x_39); lean_dec(x_3); lean_dec(x_1); -x_43 = l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Command_hasDuplicateAntiquot___spec__1___closed__2; +x_43 = l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Command_hasDuplicateAntiquot___spec__1___closed__4; lean_ctor_set(x_4, 0, x_43); x_44 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_44, 0, x_4); @@ -1631,7 +2013,7 @@ lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_dec(x_49); lean_dec(x_3); lean_dec(x_1); -x_54 = l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Command_hasDuplicateAntiquot___spec__1___closed__2; +x_54 = l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Command_hasDuplicateAntiquot___spec__1___closed__4; x_55 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_55, 0, x_54); lean_ctor_set(x_55, 1, x_45); @@ -1664,7 +2046,7 @@ goto block_22; else { lean_object* x_24; uint8_t x_25; -x_24 = l_Lean_choiceKind; +x_24 = l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Command_hasDuplicateAntiquot___spec__1___closed__3; x_25 = lean_name_eq(x_8, x_24); lean_dec(x_8); if (x_25 == 0) @@ -2097,303 +2479,7 @@ lean_ctor_set(x_6, 1, x_2); return x_6; } } -LEAN_EXPORT lean_object* l_liftM___at_Lean_Elab_Command_mkSimpleDelab___spec__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = lean_apply_2(x_1, x_2, x_3); -if (lean_obj_tag(x_4) == 0) -{ -uint8_t x_5; -x_5 = !lean_is_exclusive(x_4); -if (x_5 == 0) -{ -lean_object* x_6; lean_object* x_7; -x_6 = lean_ctor_get(x_4, 0); -x_7 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_7, 0, x_6); -lean_ctor_set(x_4, 0, x_7); -return x_4; -} -else -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_8 = lean_ctor_get(x_4, 0); -x_9 = lean_ctor_get(x_4, 1); -lean_inc(x_9); -lean_inc(x_8); -lean_dec(x_4); -x_10 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_10, 0, x_8); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_10); -lean_ctor_set(x_11, 1, x_9); -return x_11; -} -} -else -{ -uint8_t x_12; -x_12 = !lean_is_exclusive(x_4); -if (x_12 == 0) -{ -return x_4; -} -else -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = lean_ctor_get(x_4, 0); -x_14 = lean_ctor_get(x_4, 1); -lean_inc(x_14); -lean_inc(x_13); -lean_dec(x_4); -x_15 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_15, 0, x_13); -lean_ctor_set(x_15, 1, x_14); -return x_15; -} -} -} -} -LEAN_EXPORT lean_object* l_liftM___at_Lean_Elab_Command_mkSimpleDelab___spec__2(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_liftM___at_Lean_Elab_Command_mkSimpleDelab___spec__2___rarg), 3, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_mkSimpleDelab___spec__3(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_3, x_2); -if (x_7 == 0) -{ -lean_object* x_8; lean_object* x_9; -lean_dec(x_5); -lean_dec(x_1); -x_8 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_8, 0, x_4); -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_8); -lean_ctor_set(x_9, 1, x_6); -return x_9; -} -else -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_10 = lean_array_uget(x_4, x_3); -x_11 = lean_unsigned_to_nat(0u); -x_12 = lean_array_uset(x_4, x_3, x_11); -x_13 = lean_alloc_closure((void*)(l_Lean_Elab_Command_removeParentheses), 3, 1); -lean_closure_set(x_13, 0, x_10); -lean_inc(x_1); -lean_inc(x_5); -x_14 = lean_apply_3(x_1, x_13, x_5, x_6); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -if (lean_obj_tag(x_15) == 0) -{ -uint8_t x_16; -lean_dec(x_12); -lean_dec(x_5); -lean_dec(x_1); -x_16 = !lean_is_exclusive(x_14); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; -x_17 = lean_ctor_get(x_14, 0); -lean_dec(x_17); -x_18 = lean_box(0); -lean_ctor_set(x_14, 0, x_18); -return x_14; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_14, 1); -lean_inc(x_19); -lean_dec(x_14); -x_20 = lean_box(0); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_19); -return x_21; -} -} -else -{ -lean_object* x_22; lean_object* x_23; size_t x_24; size_t x_25; lean_object* x_26; -x_22 = lean_ctor_get(x_14, 1); -lean_inc(x_22); -lean_dec(x_14); -x_23 = lean_ctor_get(x_15, 0); -lean_inc(x_23); -lean_dec(x_15); -x_24 = 1; -x_25 = lean_usize_add(x_3, x_24); -x_26 = lean_array_uset(x_12, x_3, x_23); -x_3 = x_25; -x_4 = x_26; -x_6 = x_22; -goto _start; -} -} -else -{ -uint8_t x_28; -lean_dec(x_12); -lean_dec(x_5); -lean_dec(x_1); -x_28 = !lean_is_exclusive(x_14); -if (x_28 == 0) -{ -return x_14; -} -else -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_14, 0); -x_30 = lean_ctor_get(x_14, 1); -lean_inc(x_30); -lean_inc(x_29); -lean_dec(x_14); -x_31 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_30); -return x_31; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_mkSimpleDelab___spec__3___at_Lean_Elab_Command_mkSimpleDelab___spec__4(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -uint8_t x_6; -x_6 = lean_usize_dec_lt(x_2, x_1); -if (x_6 == 0) -{ -lean_object* x_7; lean_object* x_8; -lean_dec(x_4); -x_7 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_7, 0, x_3); -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_7); -lean_ctor_set(x_8, 1, x_5); -return x_8; -} -else -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_9 = lean_array_uget(x_3, x_2); -x_10 = lean_unsigned_to_nat(0u); -x_11 = lean_array_uset(x_3, x_2, x_10); -x_12 = lean_alloc_closure((void*)(l_Lean_Elab_Command_removeParentheses), 3, 1); -lean_closure_set(x_12, 0, x_9); -lean_inc(x_4); -x_13 = l_liftM___at_Lean_Elab_Command_mkSimpleDelab___spec__2___rarg(x_12, x_4, x_5); -if (lean_obj_tag(x_13) == 0) -{ -lean_object* x_14; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -if (lean_obj_tag(x_14) == 0) -{ -uint8_t x_15; -lean_dec(x_11); -lean_dec(x_4); -x_15 = !lean_is_exclusive(x_13); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_13, 0); -lean_dec(x_16); -x_17 = lean_box(0); -lean_ctor_set(x_13, 0, x_17); -return x_13; -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_13, 1); -lean_inc(x_18); -lean_dec(x_13); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_18); -return x_20; -} -} -else -{ -lean_object* x_21; lean_object* x_22; size_t x_23; size_t x_24; lean_object* x_25; -x_21 = lean_ctor_get(x_13, 1); -lean_inc(x_21); -lean_dec(x_13); -x_22 = lean_ctor_get(x_14, 0); -lean_inc(x_22); -lean_dec(x_14); -x_23 = 1; -x_24 = lean_usize_add(x_2, x_23); -x_25 = lean_array_uset(x_11, x_2, x_22); -x_2 = x_24; -x_3 = x_25; -x_5 = x_21; -goto _start; -} -} -else -{ -uint8_t x_27; -lean_dec(x_11); -lean_dec(x_4); -x_27 = !lean_is_exclusive(x_13); -if (x_27 == 0) -{ -return x_13; -} -else -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_13, 0); -x_29 = lean_ctor_get(x_13, 1); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_13); -x_30 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -return x_30; -} -} -} -} -} -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("app", 3); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_removeParentheses___closed__2; -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__1; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__3() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__1() { _start: { lean_object* x_1; @@ -2401,27 +2487,27 @@ x_1 = lean_mk_string_from_bytes("Elab", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__4() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__2; -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__3; +x_2 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__5() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_mkSimpleDelab___closed__4; +x_1 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__2; x_2 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__6() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__4() { _start: { lean_object* x_1; @@ -2429,50 +2515,23 @@ x_1 = lean_mk_string_from_bytes("aux_def", 7); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__7() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_mkSimpleDelab___closed__5; -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__6; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__8() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("null", 4); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__9() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__8; +x_1 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__3; +x_2 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__10() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(0u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__11() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__9; -x_3 = l_Lean_Elab_Command_mkSimpleDelab___closed__10; +x_2 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__6; +x_3 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__7; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2480,7 +2539,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__12() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__7() { _start: { lean_object* x_1; @@ -2488,17 +2547,17 @@ x_1 = lean_mk_string_from_bytes("attributes", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__13() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Command_removeParentheses___closed__2; -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__12; +x_2 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__7; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__14() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__9() { _start: { lean_object* x_1; @@ -2506,7 +2565,7 @@ x_1 = lean_mk_string_from_bytes("@[", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__15() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__10() { _start: { lean_object* x_1; @@ -2514,17 +2573,17 @@ x_1 = lean_mk_string_from_bytes("attrInstance", 12); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__16() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Command_removeParentheses___closed__2; -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__15; +x_2 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__10; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__17() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__12() { _start: { lean_object* x_1; @@ -2532,17 +2591,17 @@ x_1 = lean_mk_string_from_bytes("Attr", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__18() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__4; -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__17; +x_2 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__12; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__19() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__14() { _start: { lean_object* x_1; @@ -2550,17 +2609,17 @@ x_1 = lean_mk_string_from_bytes("simple", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__20() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_mkSimpleDelab___closed__18; -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__19; +x_1 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__13; +x_2 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__14; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__21() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__16() { _start: { lean_object* x_1; @@ -2568,22 +2627,22 @@ x_1 = lean_mk_string_from_bytes("appUnexpander", 13); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__22() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__17() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Command_mkSimpleDelab___closed__21; +x_1 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__16; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__23() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Command_mkSimpleDelab___closed__21; +x_1 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__16; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Command_mkSimpleDelab___closed__22; +x_3 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__17; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2591,17 +2650,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__24() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__21; +x_2 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__16; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__25() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__20() { _start: { lean_object* x_1; @@ -2609,7 +2668,7 @@ x_1 = lean_mk_string_from_bytes("]", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__26() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__21() { _start: { lean_object* x_1; lean_object* x_2; @@ -2618,7 +2677,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__27() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__22() { _start: { lean_object* x_1; @@ -2626,22 +2685,22 @@ x_1 = lean_mk_string_from_bytes("unexpand", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__28() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__23() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Command_mkSimpleDelab___closed__27; +x_1 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__22; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__29() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Command_mkSimpleDelab___closed__27; +x_1 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__22; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Command_mkSimpleDelab___closed__28; +x_3 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__23; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2649,25 +2708,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__30() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__27; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__31() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__25() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(":", 1); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__22; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__32() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__26() { _start: { lean_object* x_1; @@ -2675,22 +2726,22 @@ x_1 = lean_mk_string_from_bytes("Lean.PrettyPrinter.Unexpander", 29); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__33() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__27() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Command_mkSimpleDelab___closed__32; +x_1 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__26; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__34() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Command_mkSimpleDelab___closed__32; +x_1 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__26; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Command_mkSimpleDelab___closed__33; +x_3 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__27; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2698,7 +2749,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__35() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__29() { _start: { lean_object* x_1; @@ -2706,17 +2757,17 @@ x_1 = lean_mk_string_from_bytes("PrettyPrinter", 13); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__36() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__30() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__2; -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__35; +x_2 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__29; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__37() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__31() { _start: { lean_object* x_1; @@ -2724,41 +2775,41 @@ x_1 = lean_mk_string_from_bytes("Unexpander", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__38() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__32() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_mkSimpleDelab___closed__36; -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__37; +x_1 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__30; +x_2 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__31; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__39() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__33() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__38; +x_2 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__32; 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_Elab_Command_mkSimpleDelab___closed__40() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__34() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__39; +x_2 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__33; x_3 = lean_alloc_ctor(1, 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_Elab_Command_mkSimpleDelab___closed__41() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__35() { _start: { lean_object* x_1; @@ -2766,7 +2817,7 @@ x_1 = lean_mk_string_from_bytes(":=", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__42() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__36() { _start: { lean_object* x_1; @@ -2774,17 +2825,17 @@ x_1 = lean_mk_string_from_bytes("fun", 3); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__43() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__37() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Command_removeParentheses___closed__2; -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__42; +x_2 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__36; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__44() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__38() { _start: { lean_object* x_1; @@ -2792,17 +2843,17 @@ x_1 = lean_mk_string_from_bytes("matchAlts", 9); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__45() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__39() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Command_removeParentheses___closed__2; -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__44; +x_2 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__38; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__46() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__40() { _start: { lean_object* x_1; @@ -2810,17 +2861,17 @@ x_1 = lean_mk_string_from_bytes("matchAlt", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__47() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__41() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Command_removeParentheses___closed__2; -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__46; +x_2 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__40; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__48() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__42() { _start: { lean_object* x_1; @@ -2828,7 +2879,7 @@ x_1 = lean_mk_string_from_bytes("|", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__49() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__43() { _start: { lean_object* x_1; @@ -2836,17 +2887,17 @@ x_1 = lean_mk_string_from_bytes("quot", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__50() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__44() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Command_removeParentheses___closed__2; -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__49; +x_2 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__43; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__51() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__45() { _start: { lean_object* x_1; @@ -2854,56 +2905,64 @@ x_1 = lean_mk_string_from_bytes("`(", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__52() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__46() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("antiquot", 8); +x_1 = lean_mk_string_from_bytes(")", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__53() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__47() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___closed__2; -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__52; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("=>", 2); +return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__54() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__48() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("$", 1); +x_1 = lean_mk_string_from_bytes("app", 3); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__55() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__49() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Command_removeParentheses___closed__2; +x_2 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__48; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__50() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("f", 1); +x_1 = lean_mk_string_from_bytes("withRef", 7); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__56() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__51() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Command_mkSimpleDelab___closed__55; +x_1 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__50; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__57() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__52() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Command_mkSimpleDelab___closed__55; +x_1 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__50; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Command_mkSimpleDelab___closed__56; +x_3 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__51; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2911,83 +2970,74 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__58() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__53() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__55; +x_2 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__50; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__59() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("antiquotName", 12); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__60() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__54() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__59; +x_1 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__2; +x_2 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__50; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__61() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(4u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__62() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__55() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(")", 1); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__54; +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_Elab_Command_mkSimpleDelab___closed__63() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__56() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("=>", 2); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__55; +x_3 = lean_alloc_ctor(1, 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_Elab_Command_mkSimpleDelab___closed__64() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__57() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("withRef", 7); +x_1 = lean_mk_string_from_bytes("f", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__65() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__58() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Command_mkSimpleDelab___closed__64; +x_1 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__57; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__66() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__59() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Command_mkSimpleDelab___closed__64; +x_1 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__57; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Command_mkSimpleDelab___closed__65; +x_3 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__58; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2995,51 +3045,26 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__67() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__60() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__64; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__68() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__2; -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__64; +x_2 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__57; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__69() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__68; -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_Elab_Command_mkSimpleDelab___closed__70() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__61() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__69; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(4u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__71() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__62() { _start: { lean_object* x_1; @@ -3047,17 +3072,17 @@ x_1 = lean_mk_string_from_bytes("hole", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__72() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__63() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Command_removeParentheses___closed__2; -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__71; +x_2 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__62; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__73() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__64() { _start: { lean_object* x_1; @@ -3065,7 +3090,7 @@ x_1 = lean_mk_string_from_bytes("_", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__74() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__65() { _start: { lean_object* x_1; @@ -3073,22 +3098,22 @@ x_1 = lean_mk_string_from_bytes("throw", 5); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__75() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__66() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Command_mkSimpleDelab___closed__74; +x_1 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__65; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__76() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__67() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Command_mkSimpleDelab___closed__74; +x_1 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__65; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Command_mkSimpleDelab___closed__75; +x_3 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__66; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3096,17 +3121,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__77() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__68() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__74; +x_2 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__65; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__78() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__69() { _start: { lean_object* x_1; @@ -3114,51 +3139,51 @@ x_1 = lean_mk_string_from_bytes("MonadExcept", 11); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__79() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__70() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__78; +x_2 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__69; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__80() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__71() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_mkSimpleDelab___closed__79; -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__74; +x_1 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__70; +x_2 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__65; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__81() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__72() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__80; +x_2 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__71; 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_Elab_Command_mkSimpleDelab___closed__82() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__73() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__81; +x_2 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__72; x_3 = lean_alloc_ctor(1, 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_Elab_Command_mkSimpleDelab___closed__83() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__74() { _start: { lean_object* x_1; @@ -3166,7 +3191,7 @@ x_1 = lean_mk_string_from_bytes("(", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__84() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__75() { _start: { lean_object* x_1; lean_object* x_2; @@ -3175,3092 +3200,1906 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___closed__85() { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__76() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_mkSimpleDelab___closed__84; -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__11; +x_1 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__75; +x_2 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__6; x_3 = lean_array_push(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkSimpleDelab(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__77() { _start: { -lean_object* x_6; uint8_t x_7; -x_6 = l_Lean_Elab_Command_mkSimpleDelab___closed__2; -lean_inc(x_3); -x_7 = l_Lean_Syntax_isOfKind(x_3, x_6); -if (x_7 == 0) +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("antiquot", 8); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__78() { +_start: { -lean_object* x_8; uint8_t x_9; -x_8 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___closed__2; -lean_inc(x_3); -x_9 = l_Lean_Syntax_isOfKind(x_3, x_8); -if (x_9 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___closed__2; +x_2 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__77; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__79() { +_start: { -lean_object* x_10; lean_object* x_11; -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_10 = lean_box(0); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_10); -lean_ctor_set(x_11, 1, x_5); -return x_11; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("$", 1); +return x_1; } -else +} +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__80() { +_start: { -lean_object* x_12; lean_object* x_13; -x_12 = l_Lean_Syntax_getId(x_3); +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("antiquotName", 12); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__81() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__80; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkSimpleDelab___lambda__1(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; lean_object* x_9; +x_6 = lean_ctor_get(x_3, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_3, 1); +lean_inc(x_7); lean_dec(x_3); +x_8 = l_Lean_Syntax_getId(x_6); +lean_dec(x_6); lean_inc(x_4); -x_13 = l_Lean_Macro_resolveGlobalName(x_12, x_4, x_5); -if (lean_obj_tag(x_13) == 0) +x_9 = l_Lean_Macro_resolveGlobalName(x_8, x_4, x_5); +if (lean_obj_tag(x_9) == 0) { -lean_object* x_14; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -if (lean_obj_tag(x_14) == 0) +lean_object* x_10; +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +if (lean_obj_tag(x_10) == 0) { -uint8_t x_15; +uint8_t x_11; +lean_dec(x_7); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_15 = !lean_is_exclusive(x_13); -if (x_15 == 0) +x_11 = !lean_is_exclusive(x_9); +if (x_11 == 0) { -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_13, 0); -lean_dec(x_16); -x_17 = lean_box(0); -lean_ctor_set(x_13, 0, x_17); -return x_13; +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_9, 0); +lean_dec(x_12); +x_13 = lean_box(0); +lean_ctor_set(x_9, 0, x_13); +return x_9; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_13, 1); -lean_inc(x_18); -lean_dec(x_13); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_18); -return x_20; +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_9, 1); +lean_inc(x_14); +lean_dec(x_9); +x_15 = lean_box(0); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_14); +return x_16; } } else { -lean_object* x_21; lean_object* x_22; -x_21 = lean_ctor_get(x_14, 0); +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_10, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; +x_19 = lean_ctor_get(x_10, 1); +lean_inc(x_19); +lean_dec(x_10); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; size_t x_23; size_t x_24; lean_object* x_25; uint8_t x_26; +x_20 = lean_ctor_get(x_9, 1); +lean_inc(x_20); +lean_dec(x_9); +x_21 = lean_ctor_get(x_17, 0); lean_inc(x_21); -x_22 = lean_ctor_get(x_21, 1); -lean_inc(x_22); -if (lean_obj_tag(x_22) == 0) +lean_dec(x_17); +x_22 = lean_array_get_size(x_7); +x_23 = lean_usize_of_nat(x_22); +lean_dec(x_22); +x_24 = 0; +lean_inc(x_4); +x_25 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_removeParentheses___spec__1(x_23, x_24, x_7, x_4, x_20); +x_26 = !lean_is_exclusive(x_25); +if (x_26 == 0) { -lean_object* x_23; -x_23 = lean_ctor_get(x_14, 1); -lean_inc(x_23); -lean_dec(x_14); -if (lean_obj_tag(x_23) == 0) +lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_27 = lean_ctor_get(x_25, 0); +x_28 = lean_ctor_get(x_25, 1); +x_29 = l_Lean_Elab_Command_hasDuplicateAntiquot(x_27); +x_30 = lean_unbox(x_29); +lean_dec(x_29); +if (x_30 == 0) { -lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; -x_24 = lean_ctor_get(x_13, 1); -lean_inc(x_24); -lean_dec(x_13); -x_25 = lean_ctor_get(x_21, 0); -lean_inc(x_25); -lean_dec(x_21); +lean_object* x_31; lean_object* x_32; lean_object* x_499; lean_object* x_500; lean_object* x_501; uint8_t x_502; +lean_free_object(x_25); +lean_inc(x_4); +x_499 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Command_mkSimpleDelab___spec__1(x_4, x_28); +x_500 = lean_ctor_get(x_499, 0); +lean_inc(x_500); +x_501 = lean_ctor_get(x_499, 1); +lean_inc(x_501); +lean_dec(x_499); +x_502 = !lean_is_exclusive(x_500); +if (x_502 == 0) +{ +lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; +x_503 = lean_ctor_get(x_500, 0); +x_504 = lean_ctor_get(x_4, 2); +lean_inc(x_504); +x_505 = lean_ctor_get(x_4, 1); +lean_inc(x_505); +x_506 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__79; +lean_inc(x_503); +x_507 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_507, 0, x_503); +lean_ctor_set(x_507, 1, x_506); +x_508 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__60; +x_509 = l_Lean_addMacroScope(x_505, x_508, x_504); +x_510 = lean_box(0); +x_511 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__59; +lean_inc(x_503); +x_512 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_512, 0, x_503); +lean_ctor_set(x_512, 1, x_511); +lean_ctor_set(x_512, 2, x_509); +lean_ctor_set(x_512, 3, x_510); +x_513 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__12; +lean_inc(x_503); +x_514 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_514, 0, x_503); +lean_ctor_set(x_514, 1, x_513); +x_515 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___closed__1; +x_516 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_516, 0, x_503); +lean_ctor_set(x_516, 1, x_515); +x_517 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__8; +x_518 = lean_array_push(x_517, x_514); +x_519 = lean_array_push(x_518, x_516); +x_520 = lean_box(2); +x_521 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__81; +x_522 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_522, 0, x_520); +lean_ctor_set(x_522, 1, x_521); +lean_ctor_set(x_522, 2, x_519); +x_523 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__61; +x_524 = lean_array_push(x_523, x_507); +x_525 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__6; +x_526 = lean_array_push(x_524, x_525); +x_527 = lean_array_push(x_526, x_512); +x_528 = lean_array_push(x_527, x_522); +x_529 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__78; +x_530 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_530, 0, x_520); +lean_ctor_set(x_530, 1, x_529); +lean_ctor_set(x_530, 2, x_528); +lean_ctor_set(x_500, 0, x_530); +x_31 = x_500; +x_32 = x_501; +goto block_498; +} +else +{ +lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; +x_531 = lean_ctor_get(x_500, 0); +lean_inc(x_531); +lean_dec(x_500); +x_532 = lean_ctor_get(x_4, 2); +lean_inc(x_532); +x_533 = lean_ctor_get(x_4, 1); +lean_inc(x_533); +x_534 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__79; +lean_inc(x_531); +x_535 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_535, 0, x_531); +lean_ctor_set(x_535, 1, x_534); +x_536 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__60; +x_537 = l_Lean_addMacroScope(x_533, x_536, x_532); +x_538 = lean_box(0); +x_539 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__59; +lean_inc(x_531); +x_540 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_540, 0, x_531); +lean_ctor_set(x_540, 1, x_539); +lean_ctor_set(x_540, 2, x_537); +lean_ctor_set(x_540, 3, x_538); +x_541 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__12; +lean_inc(x_531); +x_542 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_542, 0, x_531); +lean_ctor_set(x_542, 1, x_541); +x_543 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___closed__1; +x_544 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_544, 0, x_531); +lean_ctor_set(x_544, 1, x_543); +x_545 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__8; +x_546 = lean_array_push(x_545, x_542); +x_547 = lean_array_push(x_546, x_544); +x_548 = lean_box(2); +x_549 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__81; +x_550 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_550, 0, x_548); +lean_ctor_set(x_550, 1, x_549); +lean_ctor_set(x_550, 2, x_547); +x_551 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__61; +x_552 = lean_array_push(x_551, x_535); +x_553 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__6; +x_554 = lean_array_push(x_552, x_553); +x_555 = lean_array_push(x_554, x_540); +x_556 = lean_array_push(x_555, x_550); +x_557 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__78; +x_558 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_558, 0, x_548); +lean_ctor_set(x_558, 1, x_557); +lean_ctor_set(x_558, 2, x_556); +x_559 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_559, 0, x_558); +x_31 = x_559; +x_32 = x_501; +goto block_498; +} +block_498: +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_33 = lean_ctor_get(x_31, 0); +lean_inc(x_33); +lean_dec(x_31); +x_34 = l_Lean_Syntax_mkApp(x_33, x_27); lean_inc(x_4); -x_26 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Command_mkSimpleDelab___spec__1(x_4, x_24); -x_27 = !lean_is_exclusive(x_26); -if (x_27 == 0) +x_35 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Command_mkSimpleDelab___spec__1(x_4, x_32); +x_36 = !lean_is_exclusive(x_35); +if (x_36 == 0) { -lean_object* x_28; uint8_t x_29; -x_28 = lean_ctor_get(x_26, 0); -x_29 = !lean_is_exclusive(x_28); -if (x_29 == 0) +lean_object* x_37; uint8_t x_38; +x_37 = lean_ctor_get(x_35, 0); +x_38 = !lean_is_exclusive(x_37); +if (x_38 == 0) { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; -x_30 = lean_ctor_get(x_28, 0); -x_31 = lean_ctor_get(x_4, 2); -lean_inc(x_31); -x_32 = lean_ctor_get(x_4, 1); -lean_inc(x_32); +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; +x_39 = lean_ctor_get(x_37, 0); +x_40 = lean_ctor_get(x_4, 2); +lean_inc(x_40); +x_41 = lean_ctor_get(x_4, 1); +lean_inc(x_41); lean_dec(x_4); -x_33 = l_Lean_Elab_Command_mkSimpleDelab___closed__14; -lean_inc(x_30); -x_34 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_34, 0, x_30); -lean_ctor_set(x_34, 1, x_33); -x_35 = l_Lean_Elab_Command_mkSimpleDelab___closed__24; -lean_inc(x_31); -lean_inc(x_32); -x_36 = l_Lean_addMacroScope(x_32, x_35, x_31); -x_37 = lean_box(0); -x_38 = l_Lean_Elab_Command_mkSimpleDelab___closed__23; -lean_inc(x_30); -x_39 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_39, 0, x_30); -lean_ctor_set(x_39, 1, x_38); -lean_ctor_set(x_39, 2, x_36); -lean_ctor_set(x_39, 3, x_37); -x_40 = lean_mk_syntax_ident(x_25); -x_41 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__13; +x_42 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__9; +lean_inc(x_39); +x_43 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_43, 0, x_39); +lean_ctor_set(x_43, 1, x_42); +x_44 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__19; lean_inc(x_40); -x_42 = lean_array_push(x_41, x_40); -x_43 = lean_box(2); -x_44 = l_Lean_Elab_Command_mkSimpleDelab___closed__9; -x_45 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_45, 0, x_43); -lean_ctor_set(x_45, 1, x_44); -lean_ctor_set(x_45, 2, x_42); -x_46 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__16; -x_47 = lean_array_push(x_46, x_39); -x_48 = lean_array_push(x_47, x_45); -x_49 = l_Lean_Elab_Command_mkSimpleDelab___closed__20; -x_50 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_50, 0, x_43); -lean_ctor_set(x_50, 1, x_49); -lean_ctor_set(x_50, 2, x_48); -x_51 = lean_array_push(x_46, x_1); -x_52 = lean_array_push(x_51, x_50); -x_53 = l_Lean_Elab_Command_mkSimpleDelab___closed__16; +lean_inc(x_41); +x_45 = l_Lean_addMacroScope(x_41, x_44, x_40); +x_46 = lean_box(0); +x_47 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__18; +lean_inc(x_39); +x_48 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_48, 0, x_39); +lean_ctor_set(x_48, 1, x_47); +lean_ctor_set(x_48, 2, x_45); +lean_ctor_set(x_48, 3, x_46); +x_49 = lean_mk_syntax_ident(x_21); +x_50 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__13; +lean_inc(x_49); +x_51 = lean_array_push(x_50, x_49); +x_52 = lean_box(2); +x_53 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__6; x_54 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_54, 0, x_43); +lean_ctor_set(x_54, 0, x_52); lean_ctor_set(x_54, 1, x_53); -lean_ctor_set(x_54, 2, x_52); -x_55 = lean_array_push(x_41, x_54); -x_56 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_56, 0, x_43); -lean_ctor_set(x_56, 1, x_44); -lean_ctor_set(x_56, 2, x_55); -x_57 = l_Lean_Elab_Command_mkSimpleDelab___closed__25; -lean_inc(x_30); -x_58 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_58, 0, x_30); -lean_ctor_set(x_58, 1, x_57); -x_59 = l_Lean_Elab_Command_mkSimpleDelab___closed__26; -x_60 = lean_array_push(x_59, x_34); -x_61 = lean_array_push(x_60, x_56); -x_62 = lean_array_push(x_61, x_58); -x_63 = l_Lean_Elab_Command_mkSimpleDelab___closed__13; -x_64 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_64, 0, x_43); -lean_ctor_set(x_64, 1, x_63); -lean_ctor_set(x_64, 2, x_62); -x_65 = lean_array_push(x_41, x_64); -x_66 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_66, 0, x_43); -lean_ctor_set(x_66, 1, x_44); -lean_ctor_set(x_66, 2, x_65); -x_67 = l_Lean_Elab_Command_mkSimpleDelab___closed__6; -lean_inc(x_30); -x_68 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_68, 0, x_30); -lean_ctor_set(x_68, 1, x_67); -x_69 = l_Lean_Elab_Command_mkSimpleDelab___closed__30; -lean_inc(x_31); -lean_inc(x_32); -x_70 = l_Lean_addMacroScope(x_32, x_69, x_31); -x_71 = l_Lean_Elab_Command_mkSimpleDelab___closed__29; -lean_inc(x_30); -x_72 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_72, 0, x_30); -lean_ctor_set(x_72, 1, x_71); -lean_ctor_set(x_72, 2, x_70); -lean_ctor_set(x_72, 3, x_37); -x_73 = lean_array_push(x_46, x_72); -x_74 = lean_array_push(x_73, x_40); +lean_ctor_set(x_54, 2, x_51); +x_55 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__8; +x_56 = lean_array_push(x_55, x_48); +x_57 = lean_array_push(x_56, x_54); +x_58 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__15; +x_59 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_59, 0, x_52); +lean_ctor_set(x_59, 1, x_58); +lean_ctor_set(x_59, 2, x_57); +x_60 = lean_array_push(x_55, x_1); +x_61 = lean_array_push(x_60, x_59); +x_62 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__11; +x_63 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_63, 0, x_52); +lean_ctor_set(x_63, 1, x_62); +lean_ctor_set(x_63, 2, x_61); +x_64 = lean_array_push(x_50, x_63); +x_65 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_65, 0, x_52); +lean_ctor_set(x_65, 1, x_53); +lean_ctor_set(x_65, 2, x_64); +x_66 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__20; +lean_inc(x_39); +x_67 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_67, 0, x_39); +lean_ctor_set(x_67, 1, x_66); +x_68 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__21; +x_69 = lean_array_push(x_68, x_43); +x_70 = lean_array_push(x_69, x_65); +x_71 = lean_array_push(x_70, x_67); +x_72 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__8; +x_73 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_73, 0, x_52); +lean_ctor_set(x_73, 1, x_72); +lean_ctor_set(x_73, 2, x_71); +x_74 = lean_array_push(x_50, x_73); x_75 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_75, 0, x_43); -lean_ctor_set(x_75, 1, x_44); +lean_ctor_set(x_75, 0, x_52); +lean_ctor_set(x_75, 1, x_53); lean_ctor_set(x_75, 2, x_74); -x_76 = l_Lean_Elab_Command_mkSimpleDelab___closed__31; -lean_inc(x_30); +x_76 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__4; +lean_inc(x_39); x_77 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_77, 0, x_30); +lean_ctor_set(x_77, 0, x_39); lean_ctor_set(x_77, 1, x_76); -x_78 = l_Lean_Elab_Command_mkSimpleDelab___closed__38; -lean_inc(x_31); -lean_inc(x_32); -x_79 = l_Lean_addMacroScope(x_32, x_78, x_31); -x_80 = l_Lean_Elab_Command_mkSimpleDelab___closed__34; -x_81 = l_Lean_Elab_Command_mkSimpleDelab___closed__40; -lean_inc(x_30); -x_82 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_82, 0, x_30); -lean_ctor_set(x_82, 1, x_80); -lean_ctor_set(x_82, 2, x_79); -lean_ctor_set(x_82, 3, x_81); -x_83 = l_Lean_Elab_Command_mkSimpleDelab___closed__41; -lean_inc(x_30); -x_84 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_84, 0, x_30); -lean_ctor_set(x_84, 1, x_83); -x_85 = l_Lean_Elab_Command_mkSimpleDelab___closed__42; -lean_inc(x_30); +x_78 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__25; +lean_inc(x_40); +lean_inc(x_41); +x_79 = l_Lean_addMacroScope(x_41, x_78, x_40); +x_80 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__24; +lean_inc(x_39); +x_81 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_81, 0, x_39); +lean_ctor_set(x_81, 1, x_80); +lean_ctor_set(x_81, 2, x_79); +lean_ctor_set(x_81, 3, x_46); +x_82 = lean_array_push(x_55, x_81); +x_83 = lean_array_push(x_82, x_49); +x_84 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_84, 0, x_52); +lean_ctor_set(x_84, 1, x_53); +lean_ctor_set(x_84, 2, x_83); +x_85 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__12; +lean_inc(x_39); x_86 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_86, 0, x_30); +lean_ctor_set(x_86, 0, x_39); lean_ctor_set(x_86, 1, x_85); -x_87 = l_Lean_Elab_Command_mkSimpleDelab___closed__48; -lean_inc(x_30); -x_88 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_88, 0, x_30); -lean_ctor_set(x_88, 1, x_87); -x_89 = l_Lean_Elab_Command_mkSimpleDelab___closed__51; -lean_inc(x_30); -x_90 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_90, 0, x_30); -lean_ctor_set(x_90, 1, x_89); -x_91 = l_Lean_Elab_Command_mkSimpleDelab___closed__54; -lean_inc(x_30); -x_92 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_92, 0, x_30); -lean_ctor_set(x_92, 1, x_91); -x_93 = l_Lean_Elab_Command_mkSimpleDelab___closed__58; -lean_inc(x_31); -lean_inc(x_32); -x_94 = l_Lean_addMacroScope(x_32, x_93, x_31); -x_95 = l_Lean_Elab_Command_mkSimpleDelab___closed__57; -lean_inc(x_30); -x_96 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_96, 0, x_30); -lean_ctor_set(x_96, 1, x_95); -lean_ctor_set(x_96, 2, x_94); -lean_ctor_set(x_96, 3, x_37); -x_97 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___closed__1; -lean_inc(x_30); -x_98 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_98, 0, x_30); -lean_ctor_set(x_98, 1, x_97); -lean_inc(x_77); -x_99 = lean_array_push(x_46, x_77); -x_100 = lean_array_push(x_99, x_98); -x_101 = l_Lean_Elab_Command_mkSimpleDelab___closed__60; -x_102 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_102, 0, x_43); -lean_ctor_set(x_102, 1, x_101); -lean_ctor_set(x_102, 2, x_100); -x_103 = l_Lean_Elab_Command_mkSimpleDelab___closed__61; -x_104 = lean_array_push(x_103, x_92); -x_105 = l_Lean_Elab_Command_mkSimpleDelab___closed__11; -x_106 = lean_array_push(x_104, x_105); -lean_inc(x_96); -x_107 = lean_array_push(x_106, x_96); -x_108 = lean_array_push(x_107, x_102); -x_109 = l_Lean_Elab_Command_mkSimpleDelab___closed__53; +x_87 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__32; +lean_inc(x_40); +lean_inc(x_41); +x_88 = l_Lean_addMacroScope(x_41, x_87, x_40); +x_89 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__28; +x_90 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__34; +lean_inc(x_39); +x_91 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_91, 0, x_39); +lean_ctor_set(x_91, 1, x_89); +lean_ctor_set(x_91, 2, x_88); +lean_ctor_set(x_91, 3, x_90); +x_92 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__35; +lean_inc(x_39); +x_93 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_93, 0, x_39); +lean_ctor_set(x_93, 1, x_92); +x_94 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__36; +lean_inc(x_39); +x_95 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_95, 0, x_39); +lean_ctor_set(x_95, 1, x_94); +x_96 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__42; +lean_inc(x_39); +x_97 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_97, 0, x_39); +lean_ctor_set(x_97, 1, x_96); +x_98 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__45; +lean_inc(x_39); +x_99 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_99, 0, x_39); +lean_ctor_set(x_99, 1, x_98); +x_100 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__46; +lean_inc(x_39); +x_101 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_101, 0, x_39); +lean_ctor_set(x_101, 1, x_100); +x_102 = lean_array_push(x_68, x_99); +lean_inc(x_102); +x_103 = lean_array_push(x_102, x_34); +lean_inc(x_101); +x_104 = lean_array_push(x_103, x_101); +x_105 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__44; +x_106 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_106, 0, x_52); +lean_ctor_set(x_106, 1, x_105); +lean_ctor_set(x_106, 2, x_104); +x_107 = lean_array_push(x_50, x_106); +x_108 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_108, 0, x_52); +lean_ctor_set(x_108, 1, x_53); +lean_ctor_set(x_108, 2, x_107); +x_109 = lean_array_push(x_50, x_108); x_110 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_110, 0, x_43); -lean_ctor_set(x_110, 1, x_109); -lean_ctor_set(x_110, 2, x_108); -x_111 = l_Lean_Elab_Command_mkSimpleDelab___closed__62; -lean_inc(x_30); +lean_ctor_set(x_110, 0, x_52); +lean_ctor_set(x_110, 1, x_53); +lean_ctor_set(x_110, 2, x_109); +x_111 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__47; +lean_inc(x_39); x_112 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_112, 0, x_30); +lean_ctor_set(x_112, 0, x_39); lean_ctor_set(x_112, 1, x_111); -x_113 = lean_array_push(x_59, x_90); -lean_inc(x_113); -x_114 = lean_array_push(x_113, x_110); -lean_inc(x_112); -x_115 = lean_array_push(x_114, x_112); -x_116 = l_Lean_Elab_Command_mkSimpleDelab___closed__50; -x_117 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_117, 0, x_43); -lean_ctor_set(x_117, 1, x_116); -lean_ctor_set(x_117, 2, x_115); -x_118 = lean_array_push(x_41, x_117); -x_119 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_119, 0, x_43); -lean_ctor_set(x_119, 1, x_44); -lean_ctor_set(x_119, 2, x_118); -x_120 = lean_array_push(x_41, x_119); -x_121 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_121, 0, x_43); -lean_ctor_set(x_121, 1, x_44); -lean_ctor_set(x_121, 2, x_120); -x_122 = l_Lean_Elab_Command_mkSimpleDelab___closed__63; -lean_inc(x_30); -x_123 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_123, 0, x_30); -lean_ctor_set(x_123, 1, x_122); -x_124 = l_Lean_Elab_Command_mkSimpleDelab___closed__67; -lean_inc(x_31); -lean_inc(x_32); -x_125 = l_Lean_addMacroScope(x_32, x_124, x_31); -x_126 = l_Lean_Elab_Command_mkSimpleDelab___closed__66; -x_127 = l_Lean_Elab_Command_mkSimpleDelab___closed__70; -lean_inc(x_30); -x_128 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_128, 0, x_30); -lean_ctor_set(x_128, 1, x_126); -lean_ctor_set(x_128, 2, x_125); -lean_ctor_set(x_128, 3, x_127); -x_129 = lean_array_push(x_113, x_2); -lean_inc(x_112); -x_130 = lean_array_push(x_129, x_112); +x_113 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__53; +lean_inc(x_40); +lean_inc(x_41); +x_114 = l_Lean_addMacroScope(x_41, x_113, x_40); +x_115 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__52; +x_116 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__56; +lean_inc(x_39); +x_117 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_117, 0, x_39); +lean_ctor_set(x_117, 1, x_115); +lean_ctor_set(x_117, 2, x_114); +lean_ctor_set(x_117, 3, x_116); +x_118 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__60; +lean_inc(x_40); +lean_inc(x_41); +x_119 = l_Lean_addMacroScope(x_41, x_118, x_40); +x_120 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__59; +lean_inc(x_39); +x_121 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_121, 0, x_39); +lean_ctor_set(x_121, 1, x_120); +lean_ctor_set(x_121, 2, x_119); +lean_ctor_set(x_121, 3, x_46); +x_122 = lean_array_push(x_102, x_2); +lean_inc(x_101); +x_123 = lean_array_push(x_122, x_101); +x_124 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_124, 0, x_52); +lean_ctor_set(x_124, 1, x_105); +lean_ctor_set(x_124, 2, x_123); +x_125 = lean_array_push(x_55, x_121); +x_126 = lean_array_push(x_125, x_124); +x_127 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_127, 0, x_52); +lean_ctor_set(x_127, 1, x_53); +lean_ctor_set(x_127, 2, x_126); +x_128 = lean_array_push(x_55, x_117); +x_129 = lean_array_push(x_128, x_127); +x_130 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__49; x_131 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_131, 0, x_43); -lean_ctor_set(x_131, 1, x_116); -lean_ctor_set(x_131, 2, x_130); -x_132 = lean_array_push(x_46, x_96); -x_133 = lean_array_push(x_132, x_131); -x_134 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_134, 0, x_43); -lean_ctor_set(x_134, 1, x_44); -lean_ctor_set(x_134, 2, x_133); -x_135 = lean_array_push(x_46, x_128); -x_136 = lean_array_push(x_135, x_134); -x_137 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_137, 0, x_43); -lean_ctor_set(x_137, 1, x_6); -lean_ctor_set(x_137, 2, x_136); -x_138 = lean_array_push(x_103, x_88); -lean_inc(x_138); -x_139 = lean_array_push(x_138, x_121); -lean_inc(x_123); -x_140 = lean_array_push(x_139, x_123); -x_141 = lean_array_push(x_140, x_137); -x_142 = l_Lean_Elab_Command_mkSimpleDelab___closed__47; +lean_ctor_set(x_131, 0, x_52); +lean_ctor_set(x_131, 1, x_130); +lean_ctor_set(x_131, 2, x_129); +x_132 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__61; +x_133 = lean_array_push(x_132, x_97); +lean_inc(x_133); +x_134 = lean_array_push(x_133, x_110); +lean_inc(x_112); +x_135 = lean_array_push(x_134, x_112); +x_136 = lean_array_push(x_135, x_131); +x_137 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__41; +x_138 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_138, 0, x_52); +lean_ctor_set(x_138, 1, x_137); +lean_ctor_set(x_138, 2, x_136); +x_139 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__64; +lean_inc(x_39); +x_140 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_140, 0, x_39); +lean_ctor_set(x_140, 1, x_139); +x_141 = lean_array_push(x_50, x_140); +x_142 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__63; x_143 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_143, 0, x_43); +lean_ctor_set(x_143, 0, x_52); lean_ctor_set(x_143, 1, x_142); lean_ctor_set(x_143, 2, x_141); -x_144 = l_Lean_Elab_Command_mkSimpleDelab___closed__73; -lean_inc(x_30); -x_145 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_145, 0, x_30); -lean_ctor_set(x_145, 1, x_144); -x_146 = lean_array_push(x_41, x_145); -x_147 = l_Lean_Elab_Command_mkSimpleDelab___closed__72; -x_148 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_148, 0, x_43); -lean_ctor_set(x_148, 1, x_147); -lean_ctor_set(x_148, 2, x_146); -x_149 = lean_array_push(x_41, x_148); -x_150 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_150, 0, x_43); -lean_ctor_set(x_150, 1, x_44); -lean_ctor_set(x_150, 2, x_149); -x_151 = lean_array_push(x_41, x_150); -x_152 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_152, 0, x_43); -lean_ctor_set(x_152, 1, x_44); -lean_ctor_set(x_152, 2, x_151); -x_153 = l_Lean_Elab_Command_mkSimpleDelab___closed__77; -x_154 = l_Lean_addMacroScope(x_32, x_153, x_31); -x_155 = l_Lean_Elab_Command_mkSimpleDelab___closed__76; -x_156 = l_Lean_Elab_Command_mkSimpleDelab___closed__82; -lean_inc(x_30); -x_157 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_157, 0, x_30); -lean_ctor_set(x_157, 1, x_155); -lean_ctor_set(x_157, 2, x_154); -lean_ctor_set(x_157, 3, x_156); -x_158 = l_Lean_Elab_Command_mkSimpleDelab___closed__83; -x_159 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_159, 0, x_30); -lean_ctor_set(x_159, 1, x_158); -x_160 = lean_array_push(x_59, x_159); -x_161 = lean_array_push(x_160, x_105); -x_162 = lean_array_push(x_161, x_112); -x_163 = l_Lean_Elab_Command_removeParentheses___closed__4; -x_164 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_164, 0, x_43); -lean_ctor_set(x_164, 1, x_163); -lean_ctor_set(x_164, 2, x_162); -x_165 = lean_array_push(x_41, x_164); -x_166 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_166, 0, x_43); -lean_ctor_set(x_166, 1, x_44); -lean_ctor_set(x_166, 2, x_165); -x_167 = lean_array_push(x_46, x_157); -x_168 = lean_array_push(x_167, x_166); +x_144 = lean_array_push(x_50, x_143); +x_145 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_145, 0, x_52); +lean_ctor_set(x_145, 1, x_53); +lean_ctor_set(x_145, 2, x_144); +x_146 = lean_array_push(x_50, x_145); +x_147 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_147, 0, x_52); +lean_ctor_set(x_147, 1, x_53); +lean_ctor_set(x_147, 2, x_146); +x_148 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__68; +x_149 = l_Lean_addMacroScope(x_41, x_148, x_40); +x_150 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__67; +x_151 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__73; +lean_inc(x_39); +x_152 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_152, 0, x_39); +lean_ctor_set(x_152, 1, x_150); +lean_ctor_set(x_152, 2, x_149); +lean_ctor_set(x_152, 3, x_151); +x_153 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__74; +x_154 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_154, 0, x_39); +lean_ctor_set(x_154, 1, x_153); +x_155 = lean_array_push(x_68, x_154); +x_156 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__6; +x_157 = lean_array_push(x_155, x_156); +x_158 = lean_array_push(x_157, x_101); +x_159 = l_Lean_Elab_Command_removeParentheses___closed__4; +x_160 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_160, 0, x_52); +lean_ctor_set(x_160, 1, x_159); +lean_ctor_set(x_160, 2, x_158); +x_161 = lean_array_push(x_50, x_160); +x_162 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_162, 0, x_52); +lean_ctor_set(x_162, 1, x_53); +lean_ctor_set(x_162, 2, x_161); +x_163 = lean_array_push(x_55, x_152); +x_164 = lean_array_push(x_163, x_162); +x_165 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_165, 0, x_52); +lean_ctor_set(x_165, 1, x_130); +lean_ctor_set(x_165, 2, x_164); +x_166 = lean_array_push(x_133, x_147); +x_167 = lean_array_push(x_166, x_112); +x_168 = lean_array_push(x_167, x_165); x_169 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_169, 0, x_43); -lean_ctor_set(x_169, 1, x_6); +lean_ctor_set(x_169, 0, x_52); +lean_ctor_set(x_169, 1, x_137); lean_ctor_set(x_169, 2, x_168); -x_170 = lean_array_push(x_138, x_152); -x_171 = lean_array_push(x_170, x_123); -x_172 = lean_array_push(x_171, x_169); -x_173 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_173, 0, x_43); -lean_ctor_set(x_173, 1, x_142); -lean_ctor_set(x_173, 2, x_172); -x_174 = lean_array_push(x_46, x_143); -x_175 = lean_array_push(x_174, x_173); -x_176 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_176, 0, x_43); -lean_ctor_set(x_176, 1, x_44); -lean_ctor_set(x_176, 2, x_175); -x_177 = lean_array_push(x_41, x_176); -x_178 = l_Lean_Elab_Command_mkSimpleDelab___closed__45; +x_170 = lean_array_push(x_55, x_138); +x_171 = lean_array_push(x_170, x_169); +x_172 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_172, 0, x_52); +lean_ctor_set(x_172, 1, x_53); +lean_ctor_set(x_172, 2, x_171); +x_173 = lean_array_push(x_50, x_172); +x_174 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__39; +x_175 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_175, 0, x_52); +lean_ctor_set(x_175, 1, x_174); +lean_ctor_set(x_175, 2, x_173); +x_176 = lean_array_push(x_55, x_95); +x_177 = lean_array_push(x_176, x_175); +x_178 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__37; x_179 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_179, 0, x_43); +lean_ctor_set(x_179, 0, x_52); lean_ctor_set(x_179, 1, x_178); lean_ctor_set(x_179, 2, x_177); -x_180 = lean_array_push(x_46, x_86); -x_181 = lean_array_push(x_180, x_179); -x_182 = l_Lean_Elab_Command_mkSimpleDelab___closed__43; -x_183 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_183, 0, x_43); -lean_ctor_set(x_183, 1, x_182); -lean_ctor_set(x_183, 2, x_181); -x_184 = l_Lean_Elab_Command_mkSimpleDelab___closed__85; -x_185 = lean_array_push(x_184, x_66); -x_186 = lean_array_push(x_185, x_68); -x_187 = lean_array_push(x_186, x_75); -x_188 = lean_array_push(x_187, x_77); -x_189 = lean_array_push(x_188, x_82); -x_190 = lean_array_push(x_189, x_84); -x_191 = lean_array_push(x_190, x_183); -x_192 = l_Lean_Elab_Command_mkSimpleDelab___closed__7; -x_193 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_193, 0, x_43); -lean_ctor_set(x_193, 1, x_192); -lean_ctor_set(x_193, 2, x_191); -lean_ctor_set(x_28, 0, x_193); -return x_26; +x_180 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__76; +x_181 = lean_array_push(x_180, x_75); +x_182 = lean_array_push(x_181, x_77); +x_183 = lean_array_push(x_182, x_84); +x_184 = lean_array_push(x_183, x_86); +x_185 = lean_array_push(x_184, x_91); +x_186 = lean_array_push(x_185, x_93); +x_187 = lean_array_push(x_186, x_179); +x_188 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__5; +x_189 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_189, 0, x_52); +lean_ctor_set(x_189, 1, x_188); +lean_ctor_set(x_189, 2, x_187); +lean_ctor_set(x_37, 0, x_189); +return x_35; } else { -lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; -x_194 = lean_ctor_get(x_28, 0); -lean_inc(x_194); -lean_dec(x_28); -x_195 = lean_ctor_get(x_4, 2); -lean_inc(x_195); -x_196 = lean_ctor_get(x_4, 1); -lean_inc(x_196); +lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; +x_190 = lean_ctor_get(x_37, 0); +lean_inc(x_190); +lean_dec(x_37); +x_191 = lean_ctor_get(x_4, 2); +lean_inc(x_191); +x_192 = lean_ctor_get(x_4, 1); +lean_inc(x_192); lean_dec(x_4); -x_197 = l_Lean_Elab_Command_mkSimpleDelab___closed__14; -lean_inc(x_194); -x_198 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_198, 0, x_194); -lean_ctor_set(x_198, 1, x_197); -x_199 = l_Lean_Elab_Command_mkSimpleDelab___closed__24; -lean_inc(x_195); -lean_inc(x_196); -x_200 = l_Lean_addMacroScope(x_196, x_199, x_195); -x_201 = lean_box(0); -x_202 = l_Lean_Elab_Command_mkSimpleDelab___closed__23; -lean_inc(x_194); -x_203 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_203, 0, x_194); -lean_ctor_set(x_203, 1, x_202); -lean_ctor_set(x_203, 2, x_200); -lean_ctor_set(x_203, 3, x_201); -x_204 = lean_mk_syntax_ident(x_25); -x_205 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__13; -lean_inc(x_204); -x_206 = lean_array_push(x_205, x_204); -x_207 = lean_box(2); -x_208 = l_Lean_Elab_Command_mkSimpleDelab___closed__9; -x_209 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_209, 0, x_207); -lean_ctor_set(x_209, 1, x_208); -lean_ctor_set(x_209, 2, x_206); -x_210 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__16; -x_211 = lean_array_push(x_210, x_203); -x_212 = lean_array_push(x_211, x_209); -x_213 = l_Lean_Elab_Command_mkSimpleDelab___closed__20; +x_193 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__9; +lean_inc(x_190); +x_194 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_194, 0, x_190); +lean_ctor_set(x_194, 1, x_193); +x_195 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__19; +lean_inc(x_191); +lean_inc(x_192); +x_196 = l_Lean_addMacroScope(x_192, x_195, x_191); +x_197 = lean_box(0); +x_198 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__18; +lean_inc(x_190); +x_199 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_199, 0, x_190); +lean_ctor_set(x_199, 1, x_198); +lean_ctor_set(x_199, 2, x_196); +lean_ctor_set(x_199, 3, x_197); +x_200 = lean_mk_syntax_ident(x_21); +x_201 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__13; +lean_inc(x_200); +x_202 = lean_array_push(x_201, x_200); +x_203 = lean_box(2); +x_204 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__6; +x_205 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_205, 0, x_203); +lean_ctor_set(x_205, 1, x_204); +lean_ctor_set(x_205, 2, x_202); +x_206 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__8; +x_207 = lean_array_push(x_206, x_199); +x_208 = lean_array_push(x_207, x_205); +x_209 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__15; +x_210 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_210, 0, x_203); +lean_ctor_set(x_210, 1, x_209); +lean_ctor_set(x_210, 2, x_208); +x_211 = lean_array_push(x_206, x_1); +x_212 = lean_array_push(x_211, x_210); +x_213 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__11; x_214 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_214, 0, x_207); +lean_ctor_set(x_214, 0, x_203); lean_ctor_set(x_214, 1, x_213); lean_ctor_set(x_214, 2, x_212); -x_215 = lean_array_push(x_210, x_1); -x_216 = lean_array_push(x_215, x_214); -x_217 = l_Lean_Elab_Command_mkSimpleDelab___closed__16; -x_218 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_218, 0, x_207); +x_215 = lean_array_push(x_201, x_214); +x_216 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_216, 0, x_203); +lean_ctor_set(x_216, 1, x_204); +lean_ctor_set(x_216, 2, x_215); +x_217 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__20; +lean_inc(x_190); +x_218 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_218, 0, x_190); lean_ctor_set(x_218, 1, x_217); -lean_ctor_set(x_218, 2, x_216); -x_219 = lean_array_push(x_205, x_218); -x_220 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_220, 0, x_207); -lean_ctor_set(x_220, 1, x_208); -lean_ctor_set(x_220, 2, x_219); -x_221 = l_Lean_Elab_Command_mkSimpleDelab___closed__25; -lean_inc(x_194); -x_222 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_222, 0, x_194); -lean_ctor_set(x_222, 1, x_221); -x_223 = l_Lean_Elab_Command_mkSimpleDelab___closed__26; -x_224 = lean_array_push(x_223, x_198); -x_225 = lean_array_push(x_224, x_220); -x_226 = lean_array_push(x_225, x_222); -x_227 = l_Lean_Elab_Command_mkSimpleDelab___closed__13; -x_228 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_228, 0, x_207); +x_219 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__21; +x_220 = lean_array_push(x_219, x_194); +x_221 = lean_array_push(x_220, x_216); +x_222 = lean_array_push(x_221, x_218); +x_223 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__8; +x_224 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_224, 0, x_203); +lean_ctor_set(x_224, 1, x_223); +lean_ctor_set(x_224, 2, x_222); +x_225 = lean_array_push(x_201, x_224); +x_226 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_226, 0, x_203); +lean_ctor_set(x_226, 1, x_204); +lean_ctor_set(x_226, 2, x_225); +x_227 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__4; +lean_inc(x_190); +x_228 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_228, 0, x_190); lean_ctor_set(x_228, 1, x_227); -lean_ctor_set(x_228, 2, x_226); -x_229 = lean_array_push(x_205, x_228); -x_230 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_230, 0, x_207); -lean_ctor_set(x_230, 1, x_208); -lean_ctor_set(x_230, 2, x_229); -x_231 = l_Lean_Elab_Command_mkSimpleDelab___closed__6; -lean_inc(x_194); -x_232 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_232, 0, x_194); +x_229 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__25; +lean_inc(x_191); +lean_inc(x_192); +x_230 = l_Lean_addMacroScope(x_192, x_229, x_191); +x_231 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__24; +lean_inc(x_190); +x_232 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_232, 0, x_190); lean_ctor_set(x_232, 1, x_231); -x_233 = l_Lean_Elab_Command_mkSimpleDelab___closed__30; -lean_inc(x_195); -lean_inc(x_196); -x_234 = l_Lean_addMacroScope(x_196, x_233, x_195); -x_235 = l_Lean_Elab_Command_mkSimpleDelab___closed__29; -lean_inc(x_194); -x_236 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_236, 0, x_194); -lean_ctor_set(x_236, 1, x_235); -lean_ctor_set(x_236, 2, x_234); -lean_ctor_set(x_236, 3, x_201); -x_237 = lean_array_push(x_210, x_236); -x_238 = lean_array_push(x_237, x_204); -x_239 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_239, 0, x_207); -lean_ctor_set(x_239, 1, x_208); -lean_ctor_set(x_239, 2, x_238); -x_240 = l_Lean_Elab_Command_mkSimpleDelab___closed__31; -lean_inc(x_194); -x_241 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_241, 0, x_194); -lean_ctor_set(x_241, 1, x_240); -x_242 = l_Lean_Elab_Command_mkSimpleDelab___closed__38; -lean_inc(x_195); -lean_inc(x_196); -x_243 = l_Lean_addMacroScope(x_196, x_242, x_195); -x_244 = l_Lean_Elab_Command_mkSimpleDelab___closed__34; -x_245 = l_Lean_Elab_Command_mkSimpleDelab___closed__40; -lean_inc(x_194); -x_246 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_246, 0, x_194); -lean_ctor_set(x_246, 1, x_244); -lean_ctor_set(x_246, 2, x_243); -lean_ctor_set(x_246, 3, x_245); -x_247 = l_Lean_Elab_Command_mkSimpleDelab___closed__41; -lean_inc(x_194); +lean_ctor_set(x_232, 2, x_230); +lean_ctor_set(x_232, 3, x_197); +x_233 = lean_array_push(x_206, x_232); +x_234 = lean_array_push(x_233, x_200); +x_235 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_235, 0, x_203); +lean_ctor_set(x_235, 1, x_204); +lean_ctor_set(x_235, 2, x_234); +x_236 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__12; +lean_inc(x_190); +x_237 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_237, 0, x_190); +lean_ctor_set(x_237, 1, x_236); +x_238 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__32; +lean_inc(x_191); +lean_inc(x_192); +x_239 = l_Lean_addMacroScope(x_192, x_238, x_191); +x_240 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__28; +x_241 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__34; +lean_inc(x_190); +x_242 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_242, 0, x_190); +lean_ctor_set(x_242, 1, x_240); +lean_ctor_set(x_242, 2, x_239); +lean_ctor_set(x_242, 3, x_241); +x_243 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__35; +lean_inc(x_190); +x_244 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_244, 0, x_190); +lean_ctor_set(x_244, 1, x_243); +x_245 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__36; +lean_inc(x_190); +x_246 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_246, 0, x_190); +lean_ctor_set(x_246, 1, x_245); +x_247 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__42; +lean_inc(x_190); x_248 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_248, 0, x_194); +lean_ctor_set(x_248, 0, x_190); lean_ctor_set(x_248, 1, x_247); -x_249 = l_Lean_Elab_Command_mkSimpleDelab___closed__42; -lean_inc(x_194); +x_249 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__45; +lean_inc(x_190); x_250 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_250, 0, x_194); +lean_ctor_set(x_250, 0, x_190); lean_ctor_set(x_250, 1, x_249); -x_251 = l_Lean_Elab_Command_mkSimpleDelab___closed__48; -lean_inc(x_194); +x_251 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__46; +lean_inc(x_190); x_252 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_252, 0, x_194); +lean_ctor_set(x_252, 0, x_190); lean_ctor_set(x_252, 1, x_251); -x_253 = l_Lean_Elab_Command_mkSimpleDelab___closed__51; -lean_inc(x_194); -x_254 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_254, 0, x_194); -lean_ctor_set(x_254, 1, x_253); -x_255 = l_Lean_Elab_Command_mkSimpleDelab___closed__54; -lean_inc(x_194); -x_256 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_256, 0, x_194); -lean_ctor_set(x_256, 1, x_255); -x_257 = l_Lean_Elab_Command_mkSimpleDelab___closed__58; -lean_inc(x_195); -lean_inc(x_196); -x_258 = l_Lean_addMacroScope(x_196, x_257, x_195); -x_259 = l_Lean_Elab_Command_mkSimpleDelab___closed__57; -lean_inc(x_194); -x_260 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_260, 0, x_194); -lean_ctor_set(x_260, 1, x_259); -lean_ctor_set(x_260, 2, x_258); -lean_ctor_set(x_260, 3, x_201); -x_261 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___closed__1; -lean_inc(x_194); -x_262 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_262, 0, x_194); -lean_ctor_set(x_262, 1, x_261); -lean_inc(x_241); -x_263 = lean_array_push(x_210, x_241); -x_264 = lean_array_push(x_263, x_262); -x_265 = l_Lean_Elab_Command_mkSimpleDelab___closed__60; -x_266 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_266, 0, x_207); -lean_ctor_set(x_266, 1, x_265); -lean_ctor_set(x_266, 2, x_264); -x_267 = l_Lean_Elab_Command_mkSimpleDelab___closed__61; -x_268 = lean_array_push(x_267, x_256); -x_269 = l_Lean_Elab_Command_mkSimpleDelab___closed__11; -x_270 = lean_array_push(x_268, x_269); -lean_inc(x_260); -x_271 = lean_array_push(x_270, x_260); -x_272 = lean_array_push(x_271, x_266); -x_273 = l_Lean_Elab_Command_mkSimpleDelab___closed__53; -x_274 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_274, 0, x_207); -lean_ctor_set(x_274, 1, x_273); -lean_ctor_set(x_274, 2, x_272); -x_275 = l_Lean_Elab_Command_mkSimpleDelab___closed__62; -lean_inc(x_194); -x_276 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_276, 0, x_194); -lean_ctor_set(x_276, 1, x_275); -x_277 = lean_array_push(x_223, x_254); -lean_inc(x_277); -x_278 = lean_array_push(x_277, x_274); -lean_inc(x_276); -x_279 = lean_array_push(x_278, x_276); -x_280 = l_Lean_Elab_Command_mkSimpleDelab___closed__50; -x_281 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_281, 0, x_207); -lean_ctor_set(x_281, 1, x_280); -lean_ctor_set(x_281, 2, x_279); -x_282 = lean_array_push(x_205, x_281); -x_283 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_283, 0, x_207); -lean_ctor_set(x_283, 1, x_208); -lean_ctor_set(x_283, 2, x_282); -x_284 = lean_array_push(x_205, x_283); -x_285 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_285, 0, x_207); -lean_ctor_set(x_285, 1, x_208); -lean_ctor_set(x_285, 2, x_284); -x_286 = l_Lean_Elab_Command_mkSimpleDelab___closed__63; -lean_inc(x_194); -x_287 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_287, 0, x_194); -lean_ctor_set(x_287, 1, x_286); -x_288 = l_Lean_Elab_Command_mkSimpleDelab___closed__67; -lean_inc(x_195); -lean_inc(x_196); -x_289 = l_Lean_addMacroScope(x_196, x_288, x_195); -x_290 = l_Lean_Elab_Command_mkSimpleDelab___closed__66; -x_291 = l_Lean_Elab_Command_mkSimpleDelab___closed__70; -lean_inc(x_194); -x_292 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_292, 0, x_194); -lean_ctor_set(x_292, 1, x_290); -lean_ctor_set(x_292, 2, x_289); -lean_ctor_set(x_292, 3, x_291); -x_293 = lean_array_push(x_277, x_2); -lean_inc(x_276); -x_294 = lean_array_push(x_293, x_276); -x_295 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_295, 0, x_207); -lean_ctor_set(x_295, 1, x_280); -lean_ctor_set(x_295, 2, x_294); -x_296 = lean_array_push(x_210, x_260); -x_297 = lean_array_push(x_296, x_295); +x_253 = lean_array_push(x_219, x_250); +lean_inc(x_253); +x_254 = lean_array_push(x_253, x_34); +lean_inc(x_252); +x_255 = lean_array_push(x_254, x_252); +x_256 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__44; +x_257 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_257, 0, x_203); +lean_ctor_set(x_257, 1, x_256); +lean_ctor_set(x_257, 2, x_255); +x_258 = lean_array_push(x_201, x_257); +x_259 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_259, 0, x_203); +lean_ctor_set(x_259, 1, x_204); +lean_ctor_set(x_259, 2, x_258); +x_260 = lean_array_push(x_201, x_259); +x_261 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_261, 0, x_203); +lean_ctor_set(x_261, 1, x_204); +lean_ctor_set(x_261, 2, x_260); +x_262 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__47; +lean_inc(x_190); +x_263 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_263, 0, x_190); +lean_ctor_set(x_263, 1, x_262); +x_264 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__53; +lean_inc(x_191); +lean_inc(x_192); +x_265 = l_Lean_addMacroScope(x_192, x_264, x_191); +x_266 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__52; +x_267 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__56; +lean_inc(x_190); +x_268 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_268, 0, x_190); +lean_ctor_set(x_268, 1, x_266); +lean_ctor_set(x_268, 2, x_265); +lean_ctor_set(x_268, 3, x_267); +x_269 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__60; +lean_inc(x_191); +lean_inc(x_192); +x_270 = l_Lean_addMacroScope(x_192, x_269, x_191); +x_271 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__59; +lean_inc(x_190); +x_272 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_272, 0, x_190); +lean_ctor_set(x_272, 1, x_271); +lean_ctor_set(x_272, 2, x_270); +lean_ctor_set(x_272, 3, x_197); +x_273 = lean_array_push(x_253, x_2); +lean_inc(x_252); +x_274 = lean_array_push(x_273, x_252); +x_275 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_275, 0, x_203); +lean_ctor_set(x_275, 1, x_256); +lean_ctor_set(x_275, 2, x_274); +x_276 = lean_array_push(x_206, x_272); +x_277 = lean_array_push(x_276, x_275); +x_278 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_278, 0, x_203); +lean_ctor_set(x_278, 1, x_204); +lean_ctor_set(x_278, 2, x_277); +x_279 = lean_array_push(x_206, x_268); +x_280 = lean_array_push(x_279, x_278); +x_281 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__49; +x_282 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_282, 0, x_203); +lean_ctor_set(x_282, 1, x_281); +lean_ctor_set(x_282, 2, x_280); +x_283 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__61; +x_284 = lean_array_push(x_283, x_248); +lean_inc(x_284); +x_285 = lean_array_push(x_284, x_261); +lean_inc(x_263); +x_286 = lean_array_push(x_285, x_263); +x_287 = lean_array_push(x_286, x_282); +x_288 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__41; +x_289 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_289, 0, x_203); +lean_ctor_set(x_289, 1, x_288); +lean_ctor_set(x_289, 2, x_287); +x_290 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__64; +lean_inc(x_190); +x_291 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_291, 0, x_190); +lean_ctor_set(x_291, 1, x_290); +x_292 = lean_array_push(x_201, x_291); +x_293 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__63; +x_294 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_294, 0, x_203); +lean_ctor_set(x_294, 1, x_293); +lean_ctor_set(x_294, 2, x_292); +x_295 = lean_array_push(x_201, x_294); +x_296 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_296, 0, x_203); +lean_ctor_set(x_296, 1, x_204); +lean_ctor_set(x_296, 2, x_295); +x_297 = lean_array_push(x_201, x_296); x_298 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_298, 0, x_207); -lean_ctor_set(x_298, 1, x_208); +lean_ctor_set(x_298, 0, x_203); +lean_ctor_set(x_298, 1, x_204); lean_ctor_set(x_298, 2, x_297); -x_299 = lean_array_push(x_210, x_292); -x_300 = lean_array_push(x_299, x_298); -x_301 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_301, 0, x_207); -lean_ctor_set(x_301, 1, x_6); -lean_ctor_set(x_301, 2, x_300); -x_302 = lean_array_push(x_267, x_252); -lean_inc(x_302); -x_303 = lean_array_push(x_302, x_285); -lean_inc(x_287); -x_304 = lean_array_push(x_303, x_287); -x_305 = lean_array_push(x_304, x_301); -x_306 = l_Lean_Elab_Command_mkSimpleDelab___closed__47; -x_307 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_307, 0, x_207); -lean_ctor_set(x_307, 1, x_306); -lean_ctor_set(x_307, 2, x_305); -x_308 = l_Lean_Elab_Command_mkSimpleDelab___closed__73; -lean_inc(x_194); -x_309 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_309, 0, x_194); -lean_ctor_set(x_309, 1, x_308); -x_310 = lean_array_push(x_205, x_309); -x_311 = l_Lean_Elab_Command_mkSimpleDelab___closed__72; -x_312 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_312, 0, x_207); -lean_ctor_set(x_312, 1, x_311); -lean_ctor_set(x_312, 2, x_310); -x_313 = lean_array_push(x_205, x_312); -x_314 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_314, 0, x_207); -lean_ctor_set(x_314, 1, x_208); -lean_ctor_set(x_314, 2, x_313); -x_315 = lean_array_push(x_205, x_314); +x_299 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__68; +x_300 = l_Lean_addMacroScope(x_192, x_299, x_191); +x_301 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__67; +x_302 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__73; +lean_inc(x_190); +x_303 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_303, 0, x_190); +lean_ctor_set(x_303, 1, x_301); +lean_ctor_set(x_303, 2, x_300); +lean_ctor_set(x_303, 3, x_302); +x_304 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__74; +x_305 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_305, 0, x_190); +lean_ctor_set(x_305, 1, x_304); +x_306 = lean_array_push(x_219, x_305); +x_307 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__6; +x_308 = lean_array_push(x_306, x_307); +x_309 = lean_array_push(x_308, x_252); +x_310 = l_Lean_Elab_Command_removeParentheses___closed__4; +x_311 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_311, 0, x_203); +lean_ctor_set(x_311, 1, x_310); +lean_ctor_set(x_311, 2, x_309); +x_312 = lean_array_push(x_201, x_311); +x_313 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_313, 0, x_203); +lean_ctor_set(x_313, 1, x_204); +lean_ctor_set(x_313, 2, x_312); +x_314 = lean_array_push(x_206, x_303); +x_315 = lean_array_push(x_314, x_313); x_316 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_316, 0, x_207); -lean_ctor_set(x_316, 1, x_208); +lean_ctor_set(x_316, 0, x_203); +lean_ctor_set(x_316, 1, x_281); lean_ctor_set(x_316, 2, x_315); -x_317 = l_Lean_Elab_Command_mkSimpleDelab___closed__77; -x_318 = l_Lean_addMacroScope(x_196, x_317, x_195); -x_319 = l_Lean_Elab_Command_mkSimpleDelab___closed__76; -x_320 = l_Lean_Elab_Command_mkSimpleDelab___closed__82; -lean_inc(x_194); -x_321 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_321, 0, x_194); -lean_ctor_set(x_321, 1, x_319); -lean_ctor_set(x_321, 2, x_318); -lean_ctor_set(x_321, 3, x_320); -x_322 = l_Lean_Elab_Command_mkSimpleDelab___closed__83; -x_323 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_323, 0, x_194); -lean_ctor_set(x_323, 1, x_322); -x_324 = lean_array_push(x_223, x_323); -x_325 = lean_array_push(x_324, x_269); -x_326 = lean_array_push(x_325, x_276); -x_327 = l_Lean_Elab_Command_removeParentheses___closed__4; -x_328 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_328, 0, x_207); -lean_ctor_set(x_328, 1, x_327); -lean_ctor_set(x_328, 2, x_326); -x_329 = lean_array_push(x_205, x_328); +x_317 = lean_array_push(x_284, x_298); +x_318 = lean_array_push(x_317, x_263); +x_319 = lean_array_push(x_318, x_316); +x_320 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_320, 0, x_203); +lean_ctor_set(x_320, 1, x_288); +lean_ctor_set(x_320, 2, x_319); +x_321 = lean_array_push(x_206, x_289); +x_322 = lean_array_push(x_321, x_320); +x_323 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_323, 0, x_203); +lean_ctor_set(x_323, 1, x_204); +lean_ctor_set(x_323, 2, x_322); +x_324 = lean_array_push(x_201, x_323); +x_325 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__39; +x_326 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_326, 0, x_203); +lean_ctor_set(x_326, 1, x_325); +lean_ctor_set(x_326, 2, x_324); +x_327 = lean_array_push(x_206, x_246); +x_328 = lean_array_push(x_327, x_326); +x_329 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__37; x_330 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_330, 0, x_207); -lean_ctor_set(x_330, 1, x_208); -lean_ctor_set(x_330, 2, x_329); -x_331 = lean_array_push(x_210, x_321); -x_332 = lean_array_push(x_331, x_330); -x_333 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_333, 0, x_207); -lean_ctor_set(x_333, 1, x_6); -lean_ctor_set(x_333, 2, x_332); -x_334 = lean_array_push(x_302, x_316); -x_335 = lean_array_push(x_334, x_287); -x_336 = lean_array_push(x_335, x_333); -x_337 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_337, 0, x_207); -lean_ctor_set(x_337, 1, x_306); -lean_ctor_set(x_337, 2, x_336); -x_338 = lean_array_push(x_210, x_307); -x_339 = lean_array_push(x_338, x_337); +lean_ctor_set(x_330, 0, x_203); +lean_ctor_set(x_330, 1, x_329); +lean_ctor_set(x_330, 2, x_328); +x_331 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__76; +x_332 = lean_array_push(x_331, x_226); +x_333 = lean_array_push(x_332, x_228); +x_334 = lean_array_push(x_333, x_235); +x_335 = lean_array_push(x_334, x_237); +x_336 = lean_array_push(x_335, x_242); +x_337 = lean_array_push(x_336, x_244); +x_338 = lean_array_push(x_337, x_330); +x_339 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__5; x_340 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_340, 0, x_207); -lean_ctor_set(x_340, 1, x_208); -lean_ctor_set(x_340, 2, x_339); -x_341 = lean_array_push(x_205, x_340); -x_342 = l_Lean_Elab_Command_mkSimpleDelab___closed__45; -x_343 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_343, 0, x_207); -lean_ctor_set(x_343, 1, x_342); -lean_ctor_set(x_343, 2, x_341); -x_344 = lean_array_push(x_210, x_250); -x_345 = lean_array_push(x_344, x_343); -x_346 = l_Lean_Elab_Command_mkSimpleDelab___closed__43; -x_347 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_347, 0, x_207); -lean_ctor_set(x_347, 1, x_346); -lean_ctor_set(x_347, 2, x_345); -x_348 = l_Lean_Elab_Command_mkSimpleDelab___closed__85; -x_349 = lean_array_push(x_348, x_230); -x_350 = lean_array_push(x_349, x_232); -x_351 = lean_array_push(x_350, x_239); -x_352 = lean_array_push(x_351, x_241); -x_353 = lean_array_push(x_352, x_246); -x_354 = lean_array_push(x_353, x_248); -x_355 = lean_array_push(x_354, x_347); -x_356 = l_Lean_Elab_Command_mkSimpleDelab___closed__7; -x_357 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_357, 0, x_207); -lean_ctor_set(x_357, 1, x_356); -lean_ctor_set(x_357, 2, x_355); -x_358 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_358, 0, x_357); -lean_ctor_set(x_26, 0, x_358); -return x_26; +lean_ctor_set(x_340, 0, x_203); +lean_ctor_set(x_340, 1, x_339); +lean_ctor_set(x_340, 2, x_338); +x_341 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_341, 0, x_340); +lean_ctor_set(x_35, 0, x_341); +return x_35; } } else { -lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; -x_359 = lean_ctor_get(x_26, 0); -x_360 = lean_ctor_get(x_26, 1); -lean_inc(x_360); -lean_inc(x_359); -lean_dec(x_26); -x_361 = lean_ctor_get(x_359, 0); -lean_inc(x_361); -if (lean_is_exclusive(x_359)) { - lean_ctor_release(x_359, 0); - x_362 = x_359; +lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; +x_342 = lean_ctor_get(x_35, 0); +x_343 = lean_ctor_get(x_35, 1); +lean_inc(x_343); +lean_inc(x_342); +lean_dec(x_35); +x_344 = lean_ctor_get(x_342, 0); +lean_inc(x_344); +if (lean_is_exclusive(x_342)) { + lean_ctor_release(x_342, 0); + x_345 = x_342; } else { - lean_dec_ref(x_359); - x_362 = lean_box(0); + lean_dec_ref(x_342); + x_345 = lean_box(0); } -x_363 = lean_ctor_get(x_4, 2); -lean_inc(x_363); -x_364 = lean_ctor_get(x_4, 1); -lean_inc(x_364); +x_346 = lean_ctor_get(x_4, 2); +lean_inc(x_346); +x_347 = lean_ctor_get(x_4, 1); +lean_inc(x_347); lean_dec(x_4); -x_365 = l_Lean_Elab_Command_mkSimpleDelab___closed__14; -lean_inc(x_361); -x_366 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_366, 0, x_361); -lean_ctor_set(x_366, 1, x_365); -x_367 = l_Lean_Elab_Command_mkSimpleDelab___closed__24; -lean_inc(x_363); -lean_inc(x_364); -x_368 = l_Lean_addMacroScope(x_364, x_367, x_363); -x_369 = lean_box(0); -x_370 = l_Lean_Elab_Command_mkSimpleDelab___closed__23; -lean_inc(x_361); -x_371 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_371, 0, x_361); -lean_ctor_set(x_371, 1, x_370); -lean_ctor_set(x_371, 2, x_368); -lean_ctor_set(x_371, 3, x_369); -x_372 = lean_mk_syntax_ident(x_25); -x_373 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__13; -lean_inc(x_372); -x_374 = lean_array_push(x_373, x_372); -x_375 = lean_box(2); -x_376 = l_Lean_Elab_Command_mkSimpleDelab___closed__9; -x_377 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_377, 0, x_375); -lean_ctor_set(x_377, 1, x_376); -lean_ctor_set(x_377, 2, x_374); -x_378 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__16; -x_379 = lean_array_push(x_378, x_371); -x_380 = lean_array_push(x_379, x_377); -x_381 = l_Lean_Elab_Command_mkSimpleDelab___closed__20; -x_382 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_382, 0, x_375); -lean_ctor_set(x_382, 1, x_381); -lean_ctor_set(x_382, 2, x_380); -x_383 = lean_array_push(x_378, x_1); -x_384 = lean_array_push(x_383, x_382); -x_385 = l_Lean_Elab_Command_mkSimpleDelab___closed__16; -x_386 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_386, 0, x_375); -lean_ctor_set(x_386, 1, x_385); -lean_ctor_set(x_386, 2, x_384); -x_387 = lean_array_push(x_373, x_386); -x_388 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_388, 0, x_375); -lean_ctor_set(x_388, 1, x_376); -lean_ctor_set(x_388, 2, x_387); -x_389 = l_Lean_Elab_Command_mkSimpleDelab___closed__25; -lean_inc(x_361); -x_390 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_390, 0, x_361); -lean_ctor_set(x_390, 1, x_389); -x_391 = l_Lean_Elab_Command_mkSimpleDelab___closed__26; -x_392 = lean_array_push(x_391, x_366); -x_393 = lean_array_push(x_392, x_388); -x_394 = lean_array_push(x_393, x_390); -x_395 = l_Lean_Elab_Command_mkSimpleDelab___closed__13; -x_396 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_396, 0, x_375); -lean_ctor_set(x_396, 1, x_395); -lean_ctor_set(x_396, 2, x_394); -x_397 = lean_array_push(x_373, x_396); -x_398 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_398, 0, x_375); -lean_ctor_set(x_398, 1, x_376); -lean_ctor_set(x_398, 2, x_397); -x_399 = l_Lean_Elab_Command_mkSimpleDelab___closed__6; -lean_inc(x_361); -x_400 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_400, 0, x_361); -lean_ctor_set(x_400, 1, x_399); -x_401 = l_Lean_Elab_Command_mkSimpleDelab___closed__30; -lean_inc(x_363); -lean_inc(x_364); -x_402 = l_Lean_addMacroScope(x_364, x_401, x_363); -x_403 = l_Lean_Elab_Command_mkSimpleDelab___closed__29; -lean_inc(x_361); -x_404 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_404, 0, x_361); -lean_ctor_set(x_404, 1, x_403); -lean_ctor_set(x_404, 2, x_402); -lean_ctor_set(x_404, 3, x_369); -x_405 = lean_array_push(x_378, x_404); -x_406 = lean_array_push(x_405, x_372); -x_407 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_407, 0, x_375); -lean_ctor_set(x_407, 1, x_376); -lean_ctor_set(x_407, 2, x_406); -x_408 = l_Lean_Elab_Command_mkSimpleDelab___closed__31; -lean_inc(x_361); -x_409 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_409, 0, x_361); -lean_ctor_set(x_409, 1, x_408); -x_410 = l_Lean_Elab_Command_mkSimpleDelab___closed__38; -lean_inc(x_363); -lean_inc(x_364); -x_411 = l_Lean_addMacroScope(x_364, x_410, x_363); -x_412 = l_Lean_Elab_Command_mkSimpleDelab___closed__34; -x_413 = l_Lean_Elab_Command_mkSimpleDelab___closed__40; -lean_inc(x_361); -x_414 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_414, 0, x_361); -lean_ctor_set(x_414, 1, x_412); -lean_ctor_set(x_414, 2, x_411); -lean_ctor_set(x_414, 3, x_413); -x_415 = l_Lean_Elab_Command_mkSimpleDelab___closed__41; -lean_inc(x_361); -x_416 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_416, 0, x_361); -lean_ctor_set(x_416, 1, x_415); -x_417 = l_Lean_Elab_Command_mkSimpleDelab___closed__42; -lean_inc(x_361); +x_348 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__9; +lean_inc(x_344); +x_349 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_349, 0, x_344); +lean_ctor_set(x_349, 1, x_348); +x_350 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__19; +lean_inc(x_346); +lean_inc(x_347); +x_351 = l_Lean_addMacroScope(x_347, x_350, x_346); +x_352 = lean_box(0); +x_353 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__18; +lean_inc(x_344); +x_354 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_354, 0, x_344); +lean_ctor_set(x_354, 1, x_353); +lean_ctor_set(x_354, 2, x_351); +lean_ctor_set(x_354, 3, x_352); +x_355 = lean_mk_syntax_ident(x_21); +x_356 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__13; +lean_inc(x_355); +x_357 = lean_array_push(x_356, x_355); +x_358 = lean_box(2); +x_359 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__6; +x_360 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_360, 0, x_358); +lean_ctor_set(x_360, 1, x_359); +lean_ctor_set(x_360, 2, x_357); +x_361 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__8; +x_362 = lean_array_push(x_361, x_354); +x_363 = lean_array_push(x_362, x_360); +x_364 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__15; +x_365 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_365, 0, x_358); +lean_ctor_set(x_365, 1, x_364); +lean_ctor_set(x_365, 2, x_363); +x_366 = lean_array_push(x_361, x_1); +x_367 = lean_array_push(x_366, x_365); +x_368 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__11; +x_369 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_369, 0, x_358); +lean_ctor_set(x_369, 1, x_368); +lean_ctor_set(x_369, 2, x_367); +x_370 = lean_array_push(x_356, x_369); +x_371 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_371, 0, x_358); +lean_ctor_set(x_371, 1, x_359); +lean_ctor_set(x_371, 2, x_370); +x_372 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__20; +lean_inc(x_344); +x_373 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_373, 0, x_344); +lean_ctor_set(x_373, 1, x_372); +x_374 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__21; +x_375 = lean_array_push(x_374, x_349); +x_376 = lean_array_push(x_375, x_371); +x_377 = lean_array_push(x_376, x_373); +x_378 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__8; +x_379 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_379, 0, x_358); +lean_ctor_set(x_379, 1, x_378); +lean_ctor_set(x_379, 2, x_377); +x_380 = lean_array_push(x_356, x_379); +x_381 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_381, 0, x_358); +lean_ctor_set(x_381, 1, x_359); +lean_ctor_set(x_381, 2, x_380); +x_382 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__4; +lean_inc(x_344); +x_383 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_383, 0, x_344); +lean_ctor_set(x_383, 1, x_382); +x_384 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__25; +lean_inc(x_346); +lean_inc(x_347); +x_385 = l_Lean_addMacroScope(x_347, x_384, x_346); +x_386 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__24; +lean_inc(x_344); +x_387 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_387, 0, x_344); +lean_ctor_set(x_387, 1, x_386); +lean_ctor_set(x_387, 2, x_385); +lean_ctor_set(x_387, 3, x_352); +x_388 = lean_array_push(x_361, x_387); +x_389 = lean_array_push(x_388, x_355); +x_390 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_390, 0, x_358); +lean_ctor_set(x_390, 1, x_359); +lean_ctor_set(x_390, 2, x_389); +x_391 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__12; +lean_inc(x_344); +x_392 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_392, 0, x_344); +lean_ctor_set(x_392, 1, x_391); +x_393 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__32; +lean_inc(x_346); +lean_inc(x_347); +x_394 = l_Lean_addMacroScope(x_347, x_393, x_346); +x_395 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__28; +x_396 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__34; +lean_inc(x_344); +x_397 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_397, 0, x_344); +lean_ctor_set(x_397, 1, x_395); +lean_ctor_set(x_397, 2, x_394); +lean_ctor_set(x_397, 3, x_396); +x_398 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__35; +lean_inc(x_344); +x_399 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_399, 0, x_344); +lean_ctor_set(x_399, 1, x_398); +x_400 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__36; +lean_inc(x_344); +x_401 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_401, 0, x_344); +lean_ctor_set(x_401, 1, x_400); +x_402 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__42; +lean_inc(x_344); +x_403 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_403, 0, x_344); +lean_ctor_set(x_403, 1, x_402); +x_404 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__45; +lean_inc(x_344); +x_405 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_405, 0, x_344); +lean_ctor_set(x_405, 1, x_404); +x_406 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__46; +lean_inc(x_344); +x_407 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_407, 0, x_344); +lean_ctor_set(x_407, 1, x_406); +x_408 = lean_array_push(x_374, x_405); +lean_inc(x_408); +x_409 = lean_array_push(x_408, x_34); +lean_inc(x_407); +x_410 = lean_array_push(x_409, x_407); +x_411 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__44; +x_412 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_412, 0, x_358); +lean_ctor_set(x_412, 1, x_411); +lean_ctor_set(x_412, 2, x_410); +x_413 = lean_array_push(x_356, x_412); +x_414 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_414, 0, x_358); +lean_ctor_set(x_414, 1, x_359); +lean_ctor_set(x_414, 2, x_413); +x_415 = lean_array_push(x_356, x_414); +x_416 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_416, 0, x_358); +lean_ctor_set(x_416, 1, x_359); +lean_ctor_set(x_416, 2, x_415); +x_417 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__47; +lean_inc(x_344); x_418 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_418, 0, x_361); +lean_ctor_set(x_418, 0, x_344); lean_ctor_set(x_418, 1, x_417); -x_419 = l_Lean_Elab_Command_mkSimpleDelab___closed__48; -lean_inc(x_361); -x_420 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_420, 0, x_361); -lean_ctor_set(x_420, 1, x_419); -x_421 = l_Lean_Elab_Command_mkSimpleDelab___closed__51; -lean_inc(x_361); -x_422 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_422, 0, x_361); -lean_ctor_set(x_422, 1, x_421); -x_423 = l_Lean_Elab_Command_mkSimpleDelab___closed__54; -lean_inc(x_361); -x_424 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_424, 0, x_361); -lean_ctor_set(x_424, 1, x_423); -x_425 = l_Lean_Elab_Command_mkSimpleDelab___closed__58; -lean_inc(x_363); -lean_inc(x_364); -x_426 = l_Lean_addMacroScope(x_364, x_425, x_363); -x_427 = l_Lean_Elab_Command_mkSimpleDelab___closed__57; -lean_inc(x_361); -x_428 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_428, 0, x_361); -lean_ctor_set(x_428, 1, x_427); -lean_ctor_set(x_428, 2, x_426); -lean_ctor_set(x_428, 3, x_369); -x_429 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___closed__1; -lean_inc(x_361); -x_430 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_430, 0, x_361); -lean_ctor_set(x_430, 1, x_429); -lean_inc(x_409); -x_431 = lean_array_push(x_378, x_409); +x_419 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__53; +lean_inc(x_346); +lean_inc(x_347); +x_420 = l_Lean_addMacroScope(x_347, x_419, x_346); +x_421 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__52; +x_422 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__56; +lean_inc(x_344); +x_423 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_423, 0, x_344); +lean_ctor_set(x_423, 1, x_421); +lean_ctor_set(x_423, 2, x_420); +lean_ctor_set(x_423, 3, x_422); +x_424 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__60; +lean_inc(x_346); +lean_inc(x_347); +x_425 = l_Lean_addMacroScope(x_347, x_424, x_346); +x_426 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__59; +lean_inc(x_344); +x_427 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_427, 0, x_344); +lean_ctor_set(x_427, 1, x_426); +lean_ctor_set(x_427, 2, x_425); +lean_ctor_set(x_427, 3, x_352); +x_428 = lean_array_push(x_408, x_2); +lean_inc(x_407); +x_429 = lean_array_push(x_428, x_407); +x_430 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_430, 0, x_358); +lean_ctor_set(x_430, 1, x_411); +lean_ctor_set(x_430, 2, x_429); +x_431 = lean_array_push(x_361, x_427); x_432 = lean_array_push(x_431, x_430); -x_433 = l_Lean_Elab_Command_mkSimpleDelab___closed__60; -x_434 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_434, 0, x_375); -lean_ctor_set(x_434, 1, x_433); -lean_ctor_set(x_434, 2, x_432); -x_435 = l_Lean_Elab_Command_mkSimpleDelab___closed__61; -x_436 = lean_array_push(x_435, x_424); -x_437 = l_Lean_Elab_Command_mkSimpleDelab___closed__11; -x_438 = lean_array_push(x_436, x_437); -lean_inc(x_428); -x_439 = lean_array_push(x_438, x_428); -x_440 = lean_array_push(x_439, x_434); -x_441 = l_Lean_Elab_Command_mkSimpleDelab___closed__53; -x_442 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_442, 0, x_375); -lean_ctor_set(x_442, 1, x_441); -lean_ctor_set(x_442, 2, x_440); -x_443 = l_Lean_Elab_Command_mkSimpleDelab___closed__62; -lean_inc(x_361); -x_444 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_444, 0, x_361); +x_433 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_433, 0, x_358); +lean_ctor_set(x_433, 1, x_359); +lean_ctor_set(x_433, 2, x_432); +x_434 = lean_array_push(x_361, x_423); +x_435 = lean_array_push(x_434, x_433); +x_436 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__49; +x_437 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_437, 0, x_358); +lean_ctor_set(x_437, 1, x_436); +lean_ctor_set(x_437, 2, x_435); +x_438 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__61; +x_439 = lean_array_push(x_438, x_403); +lean_inc(x_439); +x_440 = lean_array_push(x_439, x_416); +lean_inc(x_418); +x_441 = lean_array_push(x_440, x_418); +x_442 = lean_array_push(x_441, x_437); +x_443 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__41; +x_444 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_444, 0, x_358); lean_ctor_set(x_444, 1, x_443); -x_445 = lean_array_push(x_391, x_422); -lean_inc(x_445); -x_446 = lean_array_push(x_445, x_442); -lean_inc(x_444); -x_447 = lean_array_push(x_446, x_444); -x_448 = l_Lean_Elab_Command_mkSimpleDelab___closed__50; +lean_ctor_set(x_444, 2, x_442); +x_445 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__64; +lean_inc(x_344); +x_446 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_446, 0, x_344); +lean_ctor_set(x_446, 1, x_445); +x_447 = lean_array_push(x_356, x_446); +x_448 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__63; x_449 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_449, 0, x_375); +lean_ctor_set(x_449, 0, x_358); lean_ctor_set(x_449, 1, x_448); lean_ctor_set(x_449, 2, x_447); -x_450 = lean_array_push(x_373, x_449); +x_450 = lean_array_push(x_356, x_449); x_451 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_451, 0, x_375); -lean_ctor_set(x_451, 1, x_376); +lean_ctor_set(x_451, 0, x_358); +lean_ctor_set(x_451, 1, x_359); lean_ctor_set(x_451, 2, x_450); -x_452 = lean_array_push(x_373, x_451); +x_452 = lean_array_push(x_356, x_451); x_453 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_453, 0, x_375); -lean_ctor_set(x_453, 1, x_376); +lean_ctor_set(x_453, 0, x_358); +lean_ctor_set(x_453, 1, x_359); lean_ctor_set(x_453, 2, x_452); -x_454 = l_Lean_Elab_Command_mkSimpleDelab___closed__63; -lean_inc(x_361); -x_455 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_455, 0, x_361); -lean_ctor_set(x_455, 1, x_454); -x_456 = l_Lean_Elab_Command_mkSimpleDelab___closed__67; -lean_inc(x_363); -lean_inc(x_364); -x_457 = l_Lean_addMacroScope(x_364, x_456, x_363); -x_458 = l_Lean_Elab_Command_mkSimpleDelab___closed__66; -x_459 = l_Lean_Elab_Command_mkSimpleDelab___closed__70; -lean_inc(x_361); -x_460 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_460, 0, x_361); -lean_ctor_set(x_460, 1, x_458); -lean_ctor_set(x_460, 2, x_457); -lean_ctor_set(x_460, 3, x_459); -x_461 = lean_array_push(x_445, x_2); -lean_inc(x_444); -x_462 = lean_array_push(x_461, x_444); -x_463 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_463, 0, x_375); -lean_ctor_set(x_463, 1, x_448); -lean_ctor_set(x_463, 2, x_462); -x_464 = lean_array_push(x_378, x_428); -x_465 = lean_array_push(x_464, x_463); +x_454 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__68; +x_455 = l_Lean_addMacroScope(x_347, x_454, x_346); +x_456 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__67; +x_457 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__73; +lean_inc(x_344); +x_458 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_458, 0, x_344); +lean_ctor_set(x_458, 1, x_456); +lean_ctor_set(x_458, 2, x_455); +lean_ctor_set(x_458, 3, x_457); +x_459 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__74; +x_460 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_460, 0, x_344); +lean_ctor_set(x_460, 1, x_459); +x_461 = lean_array_push(x_374, x_460); +x_462 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__6; +x_463 = lean_array_push(x_461, x_462); +x_464 = lean_array_push(x_463, x_407); +x_465 = l_Lean_Elab_Command_removeParentheses___closed__4; x_466 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_466, 0, x_375); -lean_ctor_set(x_466, 1, x_376); -lean_ctor_set(x_466, 2, x_465); -x_467 = lean_array_push(x_378, x_460); -x_468 = lean_array_push(x_467, x_466); -x_469 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_469, 0, x_375); -lean_ctor_set(x_469, 1, x_6); -lean_ctor_set(x_469, 2, x_468); -x_470 = lean_array_push(x_435, x_420); -lean_inc(x_470); -x_471 = lean_array_push(x_470, x_453); -lean_inc(x_455); -x_472 = lean_array_push(x_471, x_455); -x_473 = lean_array_push(x_472, x_469); -x_474 = l_Lean_Elab_Command_mkSimpleDelab___closed__47; +lean_ctor_set(x_466, 0, x_358); +lean_ctor_set(x_466, 1, x_465); +lean_ctor_set(x_466, 2, x_464); +x_467 = lean_array_push(x_356, x_466); +x_468 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_468, 0, x_358); +lean_ctor_set(x_468, 1, x_359); +lean_ctor_set(x_468, 2, x_467); +x_469 = lean_array_push(x_361, x_458); +x_470 = lean_array_push(x_469, x_468); +x_471 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_471, 0, x_358); +lean_ctor_set(x_471, 1, x_436); +lean_ctor_set(x_471, 2, x_470); +x_472 = lean_array_push(x_439, x_453); +x_473 = lean_array_push(x_472, x_418); +x_474 = lean_array_push(x_473, x_471); x_475 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_475, 0, x_375); -lean_ctor_set(x_475, 1, x_474); -lean_ctor_set(x_475, 2, x_473); -x_476 = l_Lean_Elab_Command_mkSimpleDelab___closed__73; -lean_inc(x_361); -x_477 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_477, 0, x_361); -lean_ctor_set(x_477, 1, x_476); -x_478 = lean_array_push(x_373, x_477); -x_479 = l_Lean_Elab_Command_mkSimpleDelab___closed__72; -x_480 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_480, 0, x_375); -lean_ctor_set(x_480, 1, x_479); -lean_ctor_set(x_480, 2, x_478); -x_481 = lean_array_push(x_373, x_480); -x_482 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_482, 0, x_375); -lean_ctor_set(x_482, 1, x_376); -lean_ctor_set(x_482, 2, x_481); -x_483 = lean_array_push(x_373, x_482); -x_484 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_484, 0, x_375); -lean_ctor_set(x_484, 1, x_376); -lean_ctor_set(x_484, 2, x_483); -x_485 = l_Lean_Elab_Command_mkSimpleDelab___closed__77; -x_486 = l_Lean_addMacroScope(x_364, x_485, x_363); -x_487 = l_Lean_Elab_Command_mkSimpleDelab___closed__76; -x_488 = l_Lean_Elab_Command_mkSimpleDelab___closed__82; -lean_inc(x_361); -x_489 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_489, 0, x_361); -lean_ctor_set(x_489, 1, x_487); -lean_ctor_set(x_489, 2, x_486); -lean_ctor_set(x_489, 3, x_488); -x_490 = l_Lean_Elab_Command_mkSimpleDelab___closed__83; -x_491 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_491, 0, x_361); -lean_ctor_set(x_491, 1, x_490); -x_492 = lean_array_push(x_391, x_491); -x_493 = lean_array_push(x_492, x_437); -x_494 = lean_array_push(x_493, x_444); -x_495 = l_Lean_Elab_Command_removeParentheses___closed__4; -x_496 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_496, 0, x_375); -lean_ctor_set(x_496, 1, x_495); -lean_ctor_set(x_496, 2, x_494); -x_497 = lean_array_push(x_373, x_496); -x_498 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_498, 0, x_375); -lean_ctor_set(x_498, 1, x_376); -lean_ctor_set(x_498, 2, x_497); -x_499 = lean_array_push(x_378, x_489); -x_500 = lean_array_push(x_499, x_498); -x_501 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_501, 0, x_375); -lean_ctor_set(x_501, 1, x_6); -lean_ctor_set(x_501, 2, x_500); -x_502 = lean_array_push(x_470, x_484); -x_503 = lean_array_push(x_502, x_455); -x_504 = lean_array_push(x_503, x_501); -x_505 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_505, 0, x_375); -lean_ctor_set(x_505, 1, x_474); -lean_ctor_set(x_505, 2, x_504); -x_506 = lean_array_push(x_378, x_475); -x_507 = lean_array_push(x_506, x_505); -x_508 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_508, 0, x_375); -lean_ctor_set(x_508, 1, x_376); -lean_ctor_set(x_508, 2, x_507); -x_509 = lean_array_push(x_373, x_508); -x_510 = l_Lean_Elab_Command_mkSimpleDelab___closed__45; -x_511 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_511, 0, x_375); -lean_ctor_set(x_511, 1, x_510); -lean_ctor_set(x_511, 2, x_509); -x_512 = lean_array_push(x_378, x_418); -x_513 = lean_array_push(x_512, x_511); -x_514 = l_Lean_Elab_Command_mkSimpleDelab___closed__43; -x_515 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_515, 0, x_375); -lean_ctor_set(x_515, 1, x_514); -lean_ctor_set(x_515, 2, x_513); -x_516 = l_Lean_Elab_Command_mkSimpleDelab___closed__85; -x_517 = lean_array_push(x_516, x_398); -x_518 = lean_array_push(x_517, x_400); -x_519 = lean_array_push(x_518, x_407); -x_520 = lean_array_push(x_519, x_409); -x_521 = lean_array_push(x_520, x_414); -x_522 = lean_array_push(x_521, x_416); -x_523 = lean_array_push(x_522, x_515); -x_524 = l_Lean_Elab_Command_mkSimpleDelab___closed__7; -x_525 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_525, 0, x_375); -lean_ctor_set(x_525, 1, x_524); -lean_ctor_set(x_525, 2, x_523); -if (lean_is_scalar(x_362)) { - x_526 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_475, 0, x_358); +lean_ctor_set(x_475, 1, x_443); +lean_ctor_set(x_475, 2, x_474); +x_476 = lean_array_push(x_361, x_444); +x_477 = lean_array_push(x_476, x_475); +x_478 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_478, 0, x_358); +lean_ctor_set(x_478, 1, x_359); +lean_ctor_set(x_478, 2, x_477); +x_479 = lean_array_push(x_356, x_478); +x_480 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__39; +x_481 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_481, 0, x_358); +lean_ctor_set(x_481, 1, x_480); +lean_ctor_set(x_481, 2, x_479); +x_482 = lean_array_push(x_361, x_401); +x_483 = lean_array_push(x_482, x_481); +x_484 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__37; +x_485 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_485, 0, x_358); +lean_ctor_set(x_485, 1, x_484); +lean_ctor_set(x_485, 2, x_483); +x_486 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__76; +x_487 = lean_array_push(x_486, x_381); +x_488 = lean_array_push(x_487, x_383); +x_489 = lean_array_push(x_488, x_390); +x_490 = lean_array_push(x_489, x_392); +x_491 = lean_array_push(x_490, x_397); +x_492 = lean_array_push(x_491, x_399); +x_493 = lean_array_push(x_492, x_485); +x_494 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__5; +x_495 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_495, 0, x_358); +lean_ctor_set(x_495, 1, x_494); +lean_ctor_set(x_495, 2, x_493); +if (lean_is_scalar(x_345)) { + x_496 = lean_alloc_ctor(1, 1, 0); } else { - x_526 = x_362; -} -lean_ctor_set(x_526, 0, x_525); -x_527 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_527, 0, x_526); -lean_ctor_set(x_527, 1, x_360); -return x_527; -} -} -else -{ -uint8_t x_528; -lean_dec(x_23); -lean_dec(x_21); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_528 = !lean_is_exclusive(x_13); -if (x_528 == 0) -{ -lean_object* x_529; lean_object* x_530; -x_529 = lean_ctor_get(x_13, 0); -lean_dec(x_529); -x_530 = lean_box(0); -lean_ctor_set(x_13, 0, x_530); -return x_13; + x_496 = x_345; } -else -{ -lean_object* x_531; lean_object* x_532; lean_object* x_533; -x_531 = lean_ctor_get(x_13, 1); -lean_inc(x_531); -lean_dec(x_13); -x_532 = lean_box(0); -x_533 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_533, 0, x_532); -lean_ctor_set(x_533, 1, x_531); -return x_533; +lean_ctor_set(x_496, 0, x_495); +x_497 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_497, 0, x_496); +lean_ctor_set(x_497, 1, x_343); +return x_497; } } } else { -uint8_t x_534; -lean_dec(x_22); +lean_object* x_560; +lean_dec(x_27); lean_dec(x_21); -lean_dec(x_14); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_534 = !lean_is_exclusive(x_13); -if (x_534 == 0) -{ -lean_object* x_535; lean_object* x_536; -x_535 = lean_ctor_get(x_13, 0); -lean_dec(x_535); -x_536 = lean_box(0); -lean_ctor_set(x_13, 0, x_536); -return x_13; -} -else -{ -lean_object* x_537; lean_object* x_538; lean_object* x_539; -x_537 = lean_ctor_get(x_13, 1); -lean_inc(x_537); -lean_dec(x_13); -x_538 = lean_box(0); -x_539 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_539, 0, x_538); -lean_ctor_set(x_539, 1, x_537); -return x_539; -} -} -} -} -else -{ -uint8_t x_540; -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_540 = !lean_is_exclusive(x_13); -if (x_540 == 0) -{ -return x_13; -} -else -{ -lean_object* x_541; lean_object* x_542; lean_object* x_543; -x_541 = lean_ctor_get(x_13, 0); -x_542 = lean_ctor_get(x_13, 1); -lean_inc(x_542); -lean_inc(x_541); -lean_dec(x_13); -x_543 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_543, 0, x_541); -lean_ctor_set(x_543, 1, x_542); -return x_543; -} -} -} -} -else -{ -lean_object* x_544; lean_object* x_545; lean_object* x_546; uint8_t x_547; -x_544 = lean_unsigned_to_nat(0u); -x_545 = l_Lean_Syntax_getArg(x_3, x_544); -x_546 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___closed__2; -lean_inc(x_545); -x_547 = l_Lean_Syntax_isOfKind(x_545, x_546); -if (x_547 == 0) -{ -lean_object* x_548; lean_object* x_549; -lean_dec(x_545); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_548 = lean_box(0); -x_549 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_549, 0, x_548); -lean_ctor_set(x_549, 1, x_5); -return x_549; -} -else -{ -lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; -x_550 = lean_unsigned_to_nat(1u); -x_551 = l_Lean_Syntax_getArg(x_3, x_550); -lean_dec(x_3); -x_552 = l_Lean_Syntax_getArgs(x_551); -lean_dec(x_551); -x_553 = l_Lean_Syntax_getId(x_545); -lean_dec(x_545); -lean_inc(x_4); -x_554 = l_Lean_Macro_resolveGlobalName(x_553, x_4, x_5); -if (lean_obj_tag(x_554) == 0) -{ -lean_object* x_555; -x_555 = lean_ctor_get(x_554, 0); -lean_inc(x_555); -if (lean_obj_tag(x_555) == 0) -{ -uint8_t x_556; -lean_dec(x_552); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_556 = !lean_is_exclusive(x_554); -if (x_556 == 0) -{ -lean_object* x_557; lean_object* x_558; -x_557 = lean_ctor_get(x_554, 0); -lean_dec(x_557); -x_558 = lean_box(0); -lean_ctor_set(x_554, 0, x_558); -return x_554; -} -else -{ -lean_object* x_559; lean_object* x_560; lean_object* x_561; -x_559 = lean_ctor_get(x_554, 1); -lean_inc(x_559); -lean_dec(x_554); x_560 = lean_box(0); -x_561 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_561, 0, x_560); -lean_ctor_set(x_561, 1, x_559); -return x_561; +lean_ctor_set(x_25, 0, x_560); +return x_25; } } else { -lean_object* x_562; lean_object* x_563; -x_562 = lean_ctor_get(x_555, 0); +lean_object* x_561; lean_object* x_562; lean_object* x_563; uint8_t x_564; +x_561 = lean_ctor_get(x_25, 0); +x_562 = lean_ctor_get(x_25, 1); lean_inc(x_562); -x_563 = lean_ctor_get(x_562, 1); -lean_inc(x_563); -if (lean_obj_tag(x_563) == 0) -{ -lean_object* x_564; -x_564 = lean_ctor_get(x_555, 1); -lean_inc(x_564); -lean_dec(x_555); -if (lean_obj_tag(x_564) == 0) -{ -lean_object* x_565; lean_object* x_566; lean_object* x_567; size_t x_568; size_t x_569; lean_object* x_570; -x_565 = lean_ctor_get(x_554, 1); -lean_inc(x_565); -lean_dec(x_554); -x_566 = lean_ctor_get(x_562, 0); -lean_inc(x_566); -lean_dec(x_562); -x_567 = lean_array_get_size(x_552); -x_568 = lean_usize_of_nat(x_567); -lean_dec(x_567); -x_569 = 0; -lean_inc(x_4); -x_570 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_mkSimpleDelab___spec__3___at_Lean_Elab_Command_mkSimpleDelab___spec__4(x_568, x_569, x_552, x_4, x_565); -if (lean_obj_tag(x_570) == 0) +lean_inc(x_561); +lean_dec(x_25); +x_563 = l_Lean_Elab_Command_hasDuplicateAntiquot(x_561); +x_564 = lean_unbox(x_563); +lean_dec(x_563); +if (x_564 == 0) { -lean_object* x_571; -x_571 = lean_ctor_get(x_570, 0); +lean_object* x_565; lean_object* x_566; lean_object* x_728; lean_object* x_729; lean_object* x_730; lean_object* x_731; lean_object* x_732; lean_object* x_733; lean_object* x_734; lean_object* x_735; lean_object* x_736; lean_object* x_737; lean_object* x_738; lean_object* x_739; lean_object* x_740; lean_object* x_741; lean_object* x_742; lean_object* x_743; lean_object* x_744; lean_object* x_745; lean_object* x_746; lean_object* x_747; lean_object* x_748; lean_object* x_749; lean_object* x_750; lean_object* x_751; lean_object* x_752; lean_object* x_753; lean_object* x_754; lean_object* x_755; lean_object* x_756; lean_object* x_757; lean_object* x_758; lean_object* x_759; lean_object* x_760; +lean_inc(x_4); +x_728 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Command_mkSimpleDelab___spec__1(x_4, x_562); +x_729 = lean_ctor_get(x_728, 0); +lean_inc(x_729); +x_730 = lean_ctor_get(x_728, 1); +lean_inc(x_730); +lean_dec(x_728); +x_731 = lean_ctor_get(x_729, 0); +lean_inc(x_731); +if (lean_is_exclusive(x_729)) { + lean_ctor_release(x_729, 0); + x_732 = x_729; +} else { + lean_dec_ref(x_729); + x_732 = lean_box(0); +} +x_733 = lean_ctor_get(x_4, 2); +lean_inc(x_733); +x_734 = lean_ctor_get(x_4, 1); +lean_inc(x_734); +x_735 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__79; +lean_inc(x_731); +x_736 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_736, 0, x_731); +lean_ctor_set(x_736, 1, x_735); +x_737 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__60; +x_738 = l_Lean_addMacroScope(x_734, x_737, x_733); +x_739 = lean_box(0); +x_740 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__59; +lean_inc(x_731); +x_741 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_741, 0, x_731); +lean_ctor_set(x_741, 1, x_740); +lean_ctor_set(x_741, 2, x_738); +lean_ctor_set(x_741, 3, x_739); +x_742 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__12; +lean_inc(x_731); +x_743 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_743, 0, x_731); +lean_ctor_set(x_743, 1, x_742); +x_744 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___closed__1; +x_745 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_745, 0, x_731); +lean_ctor_set(x_745, 1, x_744); +x_746 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__8; +x_747 = lean_array_push(x_746, x_743); +x_748 = lean_array_push(x_747, x_745); +x_749 = lean_box(2); +x_750 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__81; +x_751 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_751, 0, x_749); +lean_ctor_set(x_751, 1, x_750); +lean_ctor_set(x_751, 2, x_748); +x_752 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__61; +x_753 = lean_array_push(x_752, x_736); +x_754 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__6; +x_755 = lean_array_push(x_753, x_754); +x_756 = lean_array_push(x_755, x_741); +x_757 = lean_array_push(x_756, x_751); +x_758 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__78; +x_759 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_759, 0, x_749); +lean_ctor_set(x_759, 1, x_758); +lean_ctor_set(x_759, 2, x_757); +if (lean_is_scalar(x_732)) { + x_760 = lean_alloc_ctor(1, 1, 0); +} else { + x_760 = x_732; +} +lean_ctor_set(x_760, 0, x_759); +x_565 = x_760; +x_566 = x_730; +goto block_727; +block_727: +{ +lean_object* x_567; lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; lean_object* x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; lean_object* x_678; lean_object* x_679; lean_object* x_680; lean_object* x_681; lean_object* x_682; lean_object* x_683; lean_object* x_684; lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; lean_object* x_690; lean_object* x_691; lean_object* x_692; lean_object* x_693; lean_object* x_694; lean_object* x_695; lean_object* x_696; lean_object* x_697; lean_object* x_698; lean_object* x_699; lean_object* x_700; lean_object* x_701; lean_object* x_702; lean_object* x_703; lean_object* x_704; lean_object* x_705; lean_object* x_706; lean_object* x_707; lean_object* x_708; lean_object* x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; lean_object* x_718; lean_object* x_719; lean_object* x_720; lean_object* x_721; lean_object* x_722; lean_object* x_723; lean_object* x_724; lean_object* x_725; lean_object* x_726; +x_567 = lean_ctor_get(x_565, 0); +lean_inc(x_567); +lean_dec(x_565); +x_568 = l_Lean_Syntax_mkApp(x_567, x_561); +lean_inc(x_4); +x_569 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Command_mkSimpleDelab___spec__1(x_4, x_566); +x_570 = lean_ctor_get(x_569, 0); +lean_inc(x_570); +x_571 = lean_ctor_get(x_569, 1); lean_inc(x_571); -if (lean_obj_tag(x_571) == 0) -{ -uint8_t x_572; -lean_dec(x_566); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_572 = !lean_is_exclusive(x_570); -if (x_572 == 0) -{ -lean_object* x_573; lean_object* x_574; +if (lean_is_exclusive(x_569)) { + lean_ctor_release(x_569, 0); + lean_ctor_release(x_569, 1); + x_572 = x_569; +} else { + lean_dec_ref(x_569); + x_572 = lean_box(0); +} x_573 = lean_ctor_get(x_570, 0); -lean_dec(x_573); -x_574 = lean_box(0); -lean_ctor_set(x_570, 0, x_574); -return x_570; +lean_inc(x_573); +if (lean_is_exclusive(x_570)) { + lean_ctor_release(x_570, 0); + x_574 = x_570; +} else { + lean_dec_ref(x_570); + x_574 = lean_box(0); } -else -{ -lean_object* x_575; lean_object* x_576; lean_object* x_577; -x_575 = lean_ctor_get(x_570, 1); +x_575 = lean_ctor_get(x_4, 2); lean_inc(x_575); -lean_dec(x_570); -x_576 = lean_box(0); -x_577 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_577, 0, x_576); -lean_ctor_set(x_577, 1, x_575); -return x_577; -} -} -else -{ -uint8_t x_578; -x_578 = !lean_is_exclusive(x_570); -if (x_578 == 0) -{ -lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; uint8_t x_583; -x_579 = lean_ctor_get(x_570, 1); -x_580 = lean_ctor_get(x_570, 0); -lean_dec(x_580); -x_581 = lean_ctor_get(x_571, 0); -lean_inc(x_581); -lean_dec(x_571); -x_582 = l_Lean_Elab_Command_hasDuplicateAntiquot(x_581); -x_583 = lean_unbox(x_582); -lean_dec(x_582); -if (x_583 == 0) -{ -lean_object* x_584; uint8_t x_585; -lean_free_object(x_570); -lean_inc(x_4); -x_584 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Command_mkSimpleDelab___spec__1(x_4, x_579); -x_585 = !lean_is_exclusive(x_584); -if (x_585 == 0) -{ -lean_object* x_586; uint8_t x_587; -x_586 = lean_ctor_get(x_584, 0); -x_587 = !lean_is_exclusive(x_586); -if (x_587 == 0) -{ -lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; lean_object* x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; lean_object* x_678; lean_object* x_679; lean_object* x_680; lean_object* x_681; lean_object* x_682; lean_object* x_683; lean_object* x_684; lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; lean_object* x_690; lean_object* x_691; lean_object* x_692; lean_object* x_693; lean_object* x_694; lean_object* x_695; lean_object* x_696; lean_object* x_697; lean_object* x_698; lean_object* x_699; lean_object* x_700; lean_object* x_701; lean_object* x_702; lean_object* x_703; lean_object* x_704; lean_object* x_705; lean_object* x_706; lean_object* x_707; lean_object* x_708; lean_object* x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; lean_object* x_718; lean_object* x_719; lean_object* x_720; lean_object* x_721; lean_object* x_722; lean_object* x_723; lean_object* x_724; lean_object* x_725; lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; lean_object* x_730; lean_object* x_731; lean_object* x_732; lean_object* x_733; lean_object* x_734; lean_object* x_735; lean_object* x_736; lean_object* x_737; lean_object* x_738; lean_object* x_739; lean_object* x_740; lean_object* x_741; lean_object* x_742; lean_object* x_743; lean_object* x_744; lean_object* x_745; lean_object* x_746; lean_object* x_747; lean_object* x_748; lean_object* x_749; lean_object* x_750; lean_object* x_751; lean_object* x_752; lean_object* x_753; lean_object* x_754; lean_object* x_755; lean_object* x_756; lean_object* x_757; -x_588 = lean_ctor_get(x_586, 0); -x_589 = lean_ctor_get(x_4, 2); -lean_inc(x_589); -x_590 = lean_ctor_get(x_4, 1); -lean_inc(x_590); +x_576 = lean_ctor_get(x_4, 1); +lean_inc(x_576); lean_dec(x_4); -x_591 = l_Lean_Elab_Command_mkSimpleDelab___closed__14; -lean_inc(x_588); -x_592 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_592, 0, x_588); -lean_ctor_set(x_592, 1, x_591); -x_593 = l_Lean_Elab_Command_mkSimpleDelab___closed__24; -lean_inc(x_589); -lean_inc(x_590); -x_594 = l_Lean_addMacroScope(x_590, x_593, x_589); -x_595 = lean_box(0); -x_596 = l_Lean_Elab_Command_mkSimpleDelab___closed__23; -lean_inc(x_588); -x_597 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_597, 0, x_588); -lean_ctor_set(x_597, 1, x_596); -lean_ctor_set(x_597, 2, x_594); -lean_ctor_set(x_597, 3, x_595); -x_598 = lean_mk_syntax_ident(x_566); -x_599 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__13; -lean_inc(x_598); -x_600 = lean_array_push(x_599, x_598); -x_601 = lean_box(2); -x_602 = l_Lean_Elab_Command_mkSimpleDelab___closed__9; -x_603 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_603, 0, x_601); -lean_ctor_set(x_603, 1, x_602); -lean_ctor_set(x_603, 2, x_600); -x_604 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__16; -x_605 = lean_array_push(x_604, x_597); -x_606 = lean_array_push(x_605, x_603); -x_607 = l_Lean_Elab_Command_mkSimpleDelab___closed__20; +x_577 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__9; +lean_inc(x_573); +x_578 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_578, 0, x_573); +lean_ctor_set(x_578, 1, x_577); +x_579 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__19; +lean_inc(x_575); +lean_inc(x_576); +x_580 = l_Lean_addMacroScope(x_576, x_579, x_575); +x_581 = lean_box(0); +x_582 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__18; +lean_inc(x_573); +x_583 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_583, 0, x_573); +lean_ctor_set(x_583, 1, x_582); +lean_ctor_set(x_583, 2, x_580); +lean_ctor_set(x_583, 3, x_581); +x_584 = lean_mk_syntax_ident(x_21); +x_585 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__13; +lean_inc(x_584); +x_586 = lean_array_push(x_585, x_584); +x_587 = lean_box(2); +x_588 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__6; +x_589 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_589, 0, x_587); +lean_ctor_set(x_589, 1, x_588); +lean_ctor_set(x_589, 2, x_586); +x_590 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__8; +x_591 = lean_array_push(x_590, x_583); +x_592 = lean_array_push(x_591, x_589); +x_593 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__15; +x_594 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_594, 0, x_587); +lean_ctor_set(x_594, 1, x_593); +lean_ctor_set(x_594, 2, x_592); +x_595 = lean_array_push(x_590, x_1); +x_596 = lean_array_push(x_595, x_594); +x_597 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__11; +x_598 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_598, 0, x_587); +lean_ctor_set(x_598, 1, x_597); +lean_ctor_set(x_598, 2, x_596); +x_599 = lean_array_push(x_585, x_598); +x_600 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_600, 0, x_587); +lean_ctor_set(x_600, 1, x_588); +lean_ctor_set(x_600, 2, x_599); +x_601 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__20; +lean_inc(x_573); +x_602 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_602, 0, x_573); +lean_ctor_set(x_602, 1, x_601); +x_603 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__21; +x_604 = lean_array_push(x_603, x_578); +x_605 = lean_array_push(x_604, x_600); +x_606 = lean_array_push(x_605, x_602); +x_607 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__8; x_608 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_608, 0, x_601); +lean_ctor_set(x_608, 0, x_587); lean_ctor_set(x_608, 1, x_607); lean_ctor_set(x_608, 2, x_606); -x_609 = lean_array_push(x_604, x_1); -x_610 = lean_array_push(x_609, x_608); -x_611 = l_Lean_Elab_Command_mkSimpleDelab___closed__16; -x_612 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_612, 0, x_601); +x_609 = lean_array_push(x_585, x_608); +x_610 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_610, 0, x_587); +lean_ctor_set(x_610, 1, x_588); +lean_ctor_set(x_610, 2, x_609); +x_611 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__4; +lean_inc(x_573); +x_612 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_612, 0, x_573); lean_ctor_set(x_612, 1, x_611); -lean_ctor_set(x_612, 2, x_610); -x_613 = lean_array_push(x_599, x_612); -x_614 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_614, 0, x_601); -lean_ctor_set(x_614, 1, x_602); -lean_ctor_set(x_614, 2, x_613); -x_615 = l_Lean_Elab_Command_mkSimpleDelab___closed__25; -lean_inc(x_588); -x_616 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_616, 0, x_588); +x_613 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__25; +lean_inc(x_575); +lean_inc(x_576); +x_614 = l_Lean_addMacroScope(x_576, x_613, x_575); +x_615 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__24; +lean_inc(x_573); +x_616 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_616, 0, x_573); lean_ctor_set(x_616, 1, x_615); -x_617 = l_Lean_Elab_Command_mkSimpleDelab___closed__26; -x_618 = lean_array_push(x_617, x_592); -x_619 = lean_array_push(x_618, x_614); -x_620 = lean_array_push(x_619, x_616); -x_621 = l_Lean_Elab_Command_mkSimpleDelab___closed__13; -x_622 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_622, 0, x_601); -lean_ctor_set(x_622, 1, x_621); -lean_ctor_set(x_622, 2, x_620); -x_623 = lean_array_push(x_599, x_622); -x_624 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_624, 0, x_601); -lean_ctor_set(x_624, 1, x_602); -lean_ctor_set(x_624, 2, x_623); -x_625 = l_Lean_Elab_Command_mkSimpleDelab___closed__6; -lean_inc(x_588); -x_626 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_626, 0, x_588); -lean_ctor_set(x_626, 1, x_625); -x_627 = l_Lean_Elab_Command_mkSimpleDelab___closed__30; -lean_inc(x_589); -lean_inc(x_590); -x_628 = l_Lean_addMacroScope(x_590, x_627, x_589); -x_629 = l_Lean_Elab_Command_mkSimpleDelab___closed__29; -lean_inc(x_588); -x_630 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_630, 0, x_588); +lean_ctor_set(x_616, 2, x_614); +lean_ctor_set(x_616, 3, x_581); +x_617 = lean_array_push(x_590, x_616); +x_618 = lean_array_push(x_617, x_584); +x_619 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_619, 0, x_587); +lean_ctor_set(x_619, 1, x_588); +lean_ctor_set(x_619, 2, x_618); +x_620 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__12; +lean_inc(x_573); +x_621 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_621, 0, x_573); +lean_ctor_set(x_621, 1, x_620); +x_622 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__32; +lean_inc(x_575); +lean_inc(x_576); +x_623 = l_Lean_addMacroScope(x_576, x_622, x_575); +x_624 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__28; +x_625 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__34; +lean_inc(x_573); +x_626 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_626, 0, x_573); +lean_ctor_set(x_626, 1, x_624); +lean_ctor_set(x_626, 2, x_623); +lean_ctor_set(x_626, 3, x_625); +x_627 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__35; +lean_inc(x_573); +x_628 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_628, 0, x_573); +lean_ctor_set(x_628, 1, x_627); +x_629 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__36; +lean_inc(x_573); +x_630 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_630, 0, x_573); lean_ctor_set(x_630, 1, x_629); -lean_ctor_set(x_630, 2, x_628); -lean_ctor_set(x_630, 3, x_595); -x_631 = lean_array_push(x_604, x_630); -x_632 = lean_array_push(x_631, x_598); -x_633 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_633, 0, x_601); -lean_ctor_set(x_633, 1, x_602); -lean_ctor_set(x_633, 2, x_632); -x_634 = l_Lean_Elab_Command_mkSimpleDelab___closed__31; -lean_inc(x_588); -x_635 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_635, 0, x_588); -lean_ctor_set(x_635, 1, x_634); -x_636 = l_Lean_Elab_Command_mkSimpleDelab___closed__38; -lean_inc(x_589); -lean_inc(x_590); -x_637 = l_Lean_addMacroScope(x_590, x_636, x_589); -x_638 = l_Lean_Elab_Command_mkSimpleDelab___closed__34; -x_639 = l_Lean_Elab_Command_mkSimpleDelab___closed__40; -lean_inc(x_588); -x_640 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_640, 0, x_588); -lean_ctor_set(x_640, 1, x_638); -lean_ctor_set(x_640, 2, x_637); -lean_ctor_set(x_640, 3, x_639); -x_641 = l_Lean_Elab_Command_mkSimpleDelab___closed__41; -lean_inc(x_588); -x_642 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_642, 0, x_588); -lean_ctor_set(x_642, 1, x_641); -x_643 = l_Lean_Elab_Command_mkSimpleDelab___closed__42; -lean_inc(x_588); -x_644 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_644, 0, x_588); -lean_ctor_set(x_644, 1, x_643); -x_645 = l_Lean_Elab_Command_mkSimpleDelab___closed__48; -lean_inc(x_588); -x_646 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_646, 0, x_588); -lean_ctor_set(x_646, 1, x_645); -x_647 = l_Lean_Elab_Command_mkSimpleDelab___closed__51; -lean_inc(x_588); -x_648 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_648, 0, x_588); -lean_ctor_set(x_648, 1, x_647); -x_649 = l_Lean_Elab_Command_mkSimpleDelab___closed__54; -lean_inc(x_588); -x_650 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_650, 0, x_588); -lean_ctor_set(x_650, 1, x_649); -x_651 = l_Lean_Elab_Command_mkSimpleDelab___closed__58; -lean_inc(x_589); -lean_inc(x_590); -x_652 = l_Lean_addMacroScope(x_590, x_651, x_589); -x_653 = l_Lean_Elab_Command_mkSimpleDelab___closed__57; -lean_inc(x_588); -x_654 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_654, 0, x_588); -lean_ctor_set(x_654, 1, x_653); -lean_ctor_set(x_654, 2, x_652); -lean_ctor_set(x_654, 3, x_595); -x_655 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___closed__1; -lean_inc(x_588); -x_656 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_656, 0, x_588); +x_631 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__42; +lean_inc(x_573); +x_632 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_632, 0, x_573); +lean_ctor_set(x_632, 1, x_631); +x_633 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__45; +lean_inc(x_573); +x_634 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_634, 0, x_573); +lean_ctor_set(x_634, 1, x_633); +x_635 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__46; +lean_inc(x_573); +x_636 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_636, 0, x_573); +lean_ctor_set(x_636, 1, x_635); +x_637 = lean_array_push(x_603, x_634); +lean_inc(x_637); +x_638 = lean_array_push(x_637, x_568); +lean_inc(x_636); +x_639 = lean_array_push(x_638, x_636); +x_640 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__44; +x_641 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_641, 0, x_587); +lean_ctor_set(x_641, 1, x_640); +lean_ctor_set(x_641, 2, x_639); +x_642 = lean_array_push(x_585, x_641); +x_643 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_643, 0, x_587); +lean_ctor_set(x_643, 1, x_588); +lean_ctor_set(x_643, 2, x_642); +x_644 = lean_array_push(x_585, x_643); +x_645 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_645, 0, x_587); +lean_ctor_set(x_645, 1, x_588); +lean_ctor_set(x_645, 2, x_644); +x_646 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__47; +lean_inc(x_573); +x_647 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_647, 0, x_573); +lean_ctor_set(x_647, 1, x_646); +x_648 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__53; +lean_inc(x_575); +lean_inc(x_576); +x_649 = l_Lean_addMacroScope(x_576, x_648, x_575); +x_650 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__52; +x_651 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__56; +lean_inc(x_573); +x_652 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_652, 0, x_573); +lean_ctor_set(x_652, 1, x_650); +lean_ctor_set(x_652, 2, x_649); +lean_ctor_set(x_652, 3, x_651); +x_653 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__60; +lean_inc(x_575); +lean_inc(x_576); +x_654 = l_Lean_addMacroScope(x_576, x_653, x_575); +x_655 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__59; +lean_inc(x_573); +x_656 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_656, 0, x_573); lean_ctor_set(x_656, 1, x_655); -lean_inc(x_635); -x_657 = lean_array_push(x_604, x_635); -x_658 = lean_array_push(x_657, x_656); -x_659 = l_Lean_Elab_Command_mkSimpleDelab___closed__60; -x_660 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_660, 0, x_601); -lean_ctor_set(x_660, 1, x_659); -lean_ctor_set(x_660, 2, x_658); -x_661 = l_Lean_Elab_Command_mkSimpleDelab___closed__61; -x_662 = lean_array_push(x_661, x_650); -x_663 = l_Lean_Elab_Command_mkSimpleDelab___closed__11; -x_664 = lean_array_push(x_662, x_663); -lean_inc(x_654); -x_665 = lean_array_push(x_664, x_654); -x_666 = lean_array_push(x_665, x_660); -x_667 = l_Lean_Elab_Command_mkSimpleDelab___closed__53; -x_668 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_668, 0, x_601); -lean_ctor_set(x_668, 1, x_667); -lean_ctor_set(x_668, 2, x_666); -x_669 = l_Lean_Elab_Command_mkSimpleDelab___closed__10; -x_670 = l_Array_append___rarg(x_669, x_581); -x_671 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_671, 0, x_601); -lean_ctor_set(x_671, 1, x_602); -lean_ctor_set(x_671, 2, x_670); -x_672 = lean_array_push(x_604, x_668); -x_673 = lean_array_push(x_672, x_671); -x_674 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_674, 0, x_601); -lean_ctor_set(x_674, 1, x_6); -lean_ctor_set(x_674, 2, x_673); -x_675 = l_Lean_Elab_Command_mkSimpleDelab___closed__62; -lean_inc(x_588); -x_676 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_676, 0, x_588); -lean_ctor_set(x_676, 1, x_675); -x_677 = lean_array_push(x_617, x_648); -lean_inc(x_677); -x_678 = lean_array_push(x_677, x_674); -lean_inc(x_676); -x_679 = lean_array_push(x_678, x_676); -x_680 = l_Lean_Elab_Command_mkSimpleDelab___closed__50; -x_681 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_681, 0, x_601); -lean_ctor_set(x_681, 1, x_680); -lean_ctor_set(x_681, 2, x_679); -x_682 = lean_array_push(x_599, x_681); -x_683 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_683, 0, x_601); -lean_ctor_set(x_683, 1, x_602); -lean_ctor_set(x_683, 2, x_682); -x_684 = lean_array_push(x_599, x_683); -x_685 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_685, 0, x_601); -lean_ctor_set(x_685, 1, x_602); -lean_ctor_set(x_685, 2, x_684); -x_686 = l_Lean_Elab_Command_mkSimpleDelab___closed__63; -lean_inc(x_588); -x_687 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_687, 0, x_588); -lean_ctor_set(x_687, 1, x_686); -x_688 = l_Lean_Elab_Command_mkSimpleDelab___closed__67; -lean_inc(x_589); -lean_inc(x_590); -x_689 = l_Lean_addMacroScope(x_590, x_688, x_589); -x_690 = l_Lean_Elab_Command_mkSimpleDelab___closed__66; -x_691 = l_Lean_Elab_Command_mkSimpleDelab___closed__70; -lean_inc(x_588); -x_692 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_692, 0, x_588); -lean_ctor_set(x_692, 1, x_690); -lean_ctor_set(x_692, 2, x_689); -lean_ctor_set(x_692, 3, x_691); -x_693 = lean_array_push(x_677, x_2); -lean_inc(x_676); -x_694 = lean_array_push(x_693, x_676); +lean_ctor_set(x_656, 2, x_654); +lean_ctor_set(x_656, 3, x_581); +x_657 = lean_array_push(x_637, x_2); +lean_inc(x_636); +x_658 = lean_array_push(x_657, x_636); +x_659 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_659, 0, x_587); +lean_ctor_set(x_659, 1, x_640); +lean_ctor_set(x_659, 2, x_658); +x_660 = lean_array_push(x_590, x_656); +x_661 = lean_array_push(x_660, x_659); +x_662 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_662, 0, x_587); +lean_ctor_set(x_662, 1, x_588); +lean_ctor_set(x_662, 2, x_661); +x_663 = lean_array_push(x_590, x_652); +x_664 = lean_array_push(x_663, x_662); +x_665 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__49; +x_666 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_666, 0, x_587); +lean_ctor_set(x_666, 1, x_665); +lean_ctor_set(x_666, 2, x_664); +x_667 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__61; +x_668 = lean_array_push(x_667, x_632); +lean_inc(x_668); +x_669 = lean_array_push(x_668, x_645); +lean_inc(x_647); +x_670 = lean_array_push(x_669, x_647); +x_671 = lean_array_push(x_670, x_666); +x_672 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__41; +x_673 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_673, 0, x_587); +lean_ctor_set(x_673, 1, x_672); +lean_ctor_set(x_673, 2, x_671); +x_674 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__64; +lean_inc(x_573); +x_675 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_675, 0, x_573); +lean_ctor_set(x_675, 1, x_674); +x_676 = lean_array_push(x_585, x_675); +x_677 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__63; +x_678 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_678, 0, x_587); +lean_ctor_set(x_678, 1, x_677); +lean_ctor_set(x_678, 2, x_676); +x_679 = lean_array_push(x_585, x_678); +x_680 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_680, 0, x_587); +lean_ctor_set(x_680, 1, x_588); +lean_ctor_set(x_680, 2, x_679); +x_681 = lean_array_push(x_585, x_680); +x_682 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_682, 0, x_587); +lean_ctor_set(x_682, 1, x_588); +lean_ctor_set(x_682, 2, x_681); +x_683 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__68; +x_684 = l_Lean_addMacroScope(x_576, x_683, x_575); +x_685 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__67; +x_686 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__73; +lean_inc(x_573); +x_687 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_687, 0, x_573); +lean_ctor_set(x_687, 1, x_685); +lean_ctor_set(x_687, 2, x_684); +lean_ctor_set(x_687, 3, x_686); +x_688 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__74; +x_689 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_689, 0, x_573); +lean_ctor_set(x_689, 1, x_688); +x_690 = lean_array_push(x_603, x_689); +x_691 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__6; +x_692 = lean_array_push(x_690, x_691); +x_693 = lean_array_push(x_692, x_636); +x_694 = l_Lean_Elab_Command_removeParentheses___closed__4; x_695 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_695, 0, x_601); -lean_ctor_set(x_695, 1, x_680); -lean_ctor_set(x_695, 2, x_694); -x_696 = lean_array_push(x_604, x_654); -x_697 = lean_array_push(x_696, x_695); -x_698 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_698, 0, x_601); -lean_ctor_set(x_698, 1, x_602); -lean_ctor_set(x_698, 2, x_697); -x_699 = lean_array_push(x_604, x_692); -x_700 = lean_array_push(x_699, x_698); -x_701 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_701, 0, x_601); -lean_ctor_set(x_701, 1, x_6); -lean_ctor_set(x_701, 2, x_700); -x_702 = lean_array_push(x_661, x_646); -lean_inc(x_702); -x_703 = lean_array_push(x_702, x_685); -lean_inc(x_687); -x_704 = lean_array_push(x_703, x_687); -x_705 = lean_array_push(x_704, x_701); -x_706 = l_Lean_Elab_Command_mkSimpleDelab___closed__47; +lean_ctor_set(x_695, 0, x_587); +lean_ctor_set(x_695, 1, x_694); +lean_ctor_set(x_695, 2, x_693); +x_696 = lean_array_push(x_585, x_695); +x_697 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_697, 0, x_587); +lean_ctor_set(x_697, 1, x_588); +lean_ctor_set(x_697, 2, x_696); +x_698 = lean_array_push(x_590, x_687); +x_699 = lean_array_push(x_698, x_697); +x_700 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_700, 0, x_587); +lean_ctor_set(x_700, 1, x_665); +lean_ctor_set(x_700, 2, x_699); +x_701 = lean_array_push(x_668, x_682); +x_702 = lean_array_push(x_701, x_647); +x_703 = lean_array_push(x_702, x_700); +x_704 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_704, 0, x_587); +lean_ctor_set(x_704, 1, x_672); +lean_ctor_set(x_704, 2, x_703); +x_705 = lean_array_push(x_590, x_673); +x_706 = lean_array_push(x_705, x_704); x_707 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_707, 0, x_601); -lean_ctor_set(x_707, 1, x_706); -lean_ctor_set(x_707, 2, x_705); -x_708 = l_Lean_Elab_Command_mkSimpleDelab___closed__73; -lean_inc(x_588); -x_709 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_709, 0, x_588); -lean_ctor_set(x_709, 1, x_708); -x_710 = lean_array_push(x_599, x_709); -x_711 = l_Lean_Elab_Command_mkSimpleDelab___closed__72; -x_712 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_712, 0, x_601); -lean_ctor_set(x_712, 1, x_711); -lean_ctor_set(x_712, 2, x_710); -x_713 = lean_array_push(x_599, x_712); +lean_ctor_set(x_707, 0, x_587); +lean_ctor_set(x_707, 1, x_588); +lean_ctor_set(x_707, 2, x_706); +x_708 = lean_array_push(x_585, x_707); +x_709 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__39; +x_710 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_710, 0, x_587); +lean_ctor_set(x_710, 1, x_709); +lean_ctor_set(x_710, 2, x_708); +x_711 = lean_array_push(x_590, x_630); +x_712 = lean_array_push(x_711, x_710); +x_713 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__37; x_714 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_714, 0, x_601); -lean_ctor_set(x_714, 1, x_602); -lean_ctor_set(x_714, 2, x_713); -x_715 = lean_array_push(x_599, x_714); -x_716 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_716, 0, x_601); -lean_ctor_set(x_716, 1, x_602); -lean_ctor_set(x_716, 2, x_715); -x_717 = l_Lean_Elab_Command_mkSimpleDelab___closed__77; -x_718 = l_Lean_addMacroScope(x_590, x_717, x_589); -x_719 = l_Lean_Elab_Command_mkSimpleDelab___closed__76; -x_720 = l_Lean_Elab_Command_mkSimpleDelab___closed__82; -lean_inc(x_588); -x_721 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_721, 0, x_588); -lean_ctor_set(x_721, 1, x_719); -lean_ctor_set(x_721, 2, x_718); -lean_ctor_set(x_721, 3, x_720); -x_722 = l_Lean_Elab_Command_mkSimpleDelab___closed__83; -x_723 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_723, 0, x_588); -lean_ctor_set(x_723, 1, x_722); -x_724 = lean_array_push(x_617, x_723); -x_725 = lean_array_push(x_724, x_663); -x_726 = lean_array_push(x_725, x_676); -x_727 = l_Lean_Elab_Command_removeParentheses___closed__4; -x_728 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_728, 0, x_601); -lean_ctor_set(x_728, 1, x_727); -lean_ctor_set(x_728, 2, x_726); -x_729 = lean_array_push(x_599, x_728); -x_730 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_730, 0, x_601); -lean_ctor_set(x_730, 1, x_602); -lean_ctor_set(x_730, 2, x_729); -x_731 = lean_array_push(x_604, x_721); -x_732 = lean_array_push(x_731, x_730); -x_733 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_733, 0, x_601); -lean_ctor_set(x_733, 1, x_6); -lean_ctor_set(x_733, 2, x_732); -x_734 = lean_array_push(x_702, x_716); -x_735 = lean_array_push(x_734, x_687); -x_736 = lean_array_push(x_735, x_733); -x_737 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_737, 0, x_601); -lean_ctor_set(x_737, 1, x_706); -lean_ctor_set(x_737, 2, x_736); -x_738 = lean_array_push(x_604, x_707); -x_739 = lean_array_push(x_738, x_737); -x_740 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_740, 0, x_601); -lean_ctor_set(x_740, 1, x_602); -lean_ctor_set(x_740, 2, x_739); -x_741 = lean_array_push(x_599, x_740); -x_742 = l_Lean_Elab_Command_mkSimpleDelab___closed__45; -x_743 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_743, 0, x_601); -lean_ctor_set(x_743, 1, x_742); -lean_ctor_set(x_743, 2, x_741); -x_744 = lean_array_push(x_604, x_644); -x_745 = lean_array_push(x_744, x_743); -x_746 = l_Lean_Elab_Command_mkSimpleDelab___closed__43; -x_747 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_747, 0, x_601); -lean_ctor_set(x_747, 1, x_746); -lean_ctor_set(x_747, 2, x_745); -x_748 = l_Lean_Elab_Command_mkSimpleDelab___closed__85; -x_749 = lean_array_push(x_748, x_624); -x_750 = lean_array_push(x_749, x_626); -x_751 = lean_array_push(x_750, x_633); -x_752 = lean_array_push(x_751, x_635); -x_753 = lean_array_push(x_752, x_640); -x_754 = lean_array_push(x_753, x_642); -x_755 = lean_array_push(x_754, x_747); -x_756 = l_Lean_Elab_Command_mkSimpleDelab___closed__7; -x_757 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_757, 0, x_601); -lean_ctor_set(x_757, 1, x_756); -lean_ctor_set(x_757, 2, x_755); -lean_ctor_set(x_586, 0, x_757); -return x_584; -} -else -{ -lean_object* x_758; lean_object* x_759; lean_object* x_760; lean_object* x_761; lean_object* x_762; lean_object* x_763; lean_object* x_764; lean_object* x_765; lean_object* x_766; lean_object* x_767; lean_object* x_768; lean_object* x_769; lean_object* x_770; lean_object* x_771; lean_object* x_772; lean_object* x_773; lean_object* x_774; lean_object* x_775; lean_object* x_776; lean_object* x_777; lean_object* x_778; lean_object* x_779; lean_object* x_780; lean_object* x_781; lean_object* x_782; lean_object* x_783; lean_object* x_784; lean_object* x_785; lean_object* x_786; lean_object* x_787; lean_object* x_788; lean_object* x_789; lean_object* x_790; lean_object* x_791; lean_object* x_792; lean_object* x_793; lean_object* x_794; lean_object* x_795; lean_object* x_796; lean_object* x_797; lean_object* x_798; lean_object* x_799; lean_object* x_800; lean_object* x_801; lean_object* x_802; lean_object* x_803; lean_object* x_804; lean_object* x_805; lean_object* x_806; lean_object* x_807; lean_object* x_808; lean_object* x_809; lean_object* x_810; lean_object* x_811; lean_object* x_812; lean_object* x_813; lean_object* x_814; lean_object* x_815; lean_object* x_816; lean_object* x_817; lean_object* x_818; lean_object* x_819; lean_object* x_820; lean_object* x_821; lean_object* x_822; lean_object* x_823; lean_object* x_824; lean_object* x_825; lean_object* x_826; lean_object* x_827; lean_object* x_828; lean_object* x_829; lean_object* x_830; lean_object* x_831; lean_object* x_832; lean_object* x_833; lean_object* x_834; lean_object* x_835; lean_object* x_836; lean_object* x_837; lean_object* x_838; lean_object* x_839; lean_object* x_840; lean_object* x_841; lean_object* x_842; lean_object* x_843; lean_object* x_844; lean_object* x_845; lean_object* x_846; lean_object* x_847; lean_object* x_848; lean_object* x_849; lean_object* x_850; lean_object* x_851; lean_object* x_852; lean_object* x_853; lean_object* x_854; lean_object* x_855; lean_object* x_856; lean_object* x_857; lean_object* x_858; lean_object* x_859; lean_object* x_860; lean_object* x_861; lean_object* x_862; lean_object* x_863; lean_object* x_864; lean_object* x_865; lean_object* x_866; lean_object* x_867; lean_object* x_868; lean_object* x_869; lean_object* x_870; lean_object* x_871; lean_object* x_872; lean_object* x_873; lean_object* x_874; lean_object* x_875; lean_object* x_876; lean_object* x_877; lean_object* x_878; lean_object* x_879; lean_object* x_880; lean_object* x_881; lean_object* x_882; lean_object* x_883; lean_object* x_884; lean_object* x_885; lean_object* x_886; lean_object* x_887; lean_object* x_888; lean_object* x_889; lean_object* x_890; lean_object* x_891; lean_object* x_892; lean_object* x_893; lean_object* x_894; lean_object* x_895; lean_object* x_896; lean_object* x_897; lean_object* x_898; lean_object* x_899; lean_object* x_900; lean_object* x_901; lean_object* x_902; lean_object* x_903; lean_object* x_904; lean_object* x_905; lean_object* x_906; lean_object* x_907; lean_object* x_908; lean_object* x_909; lean_object* x_910; lean_object* x_911; lean_object* x_912; lean_object* x_913; lean_object* x_914; lean_object* x_915; lean_object* x_916; lean_object* x_917; lean_object* x_918; lean_object* x_919; lean_object* x_920; lean_object* x_921; lean_object* x_922; lean_object* x_923; lean_object* x_924; lean_object* x_925; lean_object* x_926; lean_object* x_927; lean_object* x_928; -x_758 = lean_ctor_get(x_586, 0); -lean_inc(x_758); -lean_dec(x_586); -x_759 = lean_ctor_get(x_4, 2); -lean_inc(x_759); -x_760 = lean_ctor_get(x_4, 1); -lean_inc(x_760); -lean_dec(x_4); -x_761 = l_Lean_Elab_Command_mkSimpleDelab___closed__14; -lean_inc(x_758); -x_762 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_762, 0, x_758); -lean_ctor_set(x_762, 1, x_761); -x_763 = l_Lean_Elab_Command_mkSimpleDelab___closed__24; -lean_inc(x_759); -lean_inc(x_760); -x_764 = l_Lean_addMacroScope(x_760, x_763, x_759); -x_765 = lean_box(0); -x_766 = l_Lean_Elab_Command_mkSimpleDelab___closed__23; -lean_inc(x_758); -x_767 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_767, 0, x_758); -lean_ctor_set(x_767, 1, x_766); -lean_ctor_set(x_767, 2, x_764); -lean_ctor_set(x_767, 3, x_765); -x_768 = lean_mk_syntax_ident(x_566); -x_769 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__13; -lean_inc(x_768); -x_770 = lean_array_push(x_769, x_768); -x_771 = lean_box(2); -x_772 = l_Lean_Elab_Command_mkSimpleDelab___closed__9; -x_773 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_773, 0, x_771); -lean_ctor_set(x_773, 1, x_772); -lean_ctor_set(x_773, 2, x_770); -x_774 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__16; -x_775 = lean_array_push(x_774, x_767); -x_776 = lean_array_push(x_775, x_773); -x_777 = l_Lean_Elab_Command_mkSimpleDelab___closed__20; -x_778 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_778, 0, x_771); -lean_ctor_set(x_778, 1, x_777); -lean_ctor_set(x_778, 2, x_776); -x_779 = lean_array_push(x_774, x_1); -x_780 = lean_array_push(x_779, x_778); -x_781 = l_Lean_Elab_Command_mkSimpleDelab___closed__16; -x_782 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_782, 0, x_771); -lean_ctor_set(x_782, 1, x_781); -lean_ctor_set(x_782, 2, x_780); -x_783 = lean_array_push(x_769, x_782); -x_784 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_784, 0, x_771); -lean_ctor_set(x_784, 1, x_772); -lean_ctor_set(x_784, 2, x_783); -x_785 = l_Lean_Elab_Command_mkSimpleDelab___closed__25; -lean_inc(x_758); -x_786 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_786, 0, x_758); -lean_ctor_set(x_786, 1, x_785); -x_787 = l_Lean_Elab_Command_mkSimpleDelab___closed__26; -x_788 = lean_array_push(x_787, x_762); -x_789 = lean_array_push(x_788, x_784); -x_790 = lean_array_push(x_789, x_786); -x_791 = l_Lean_Elab_Command_mkSimpleDelab___closed__13; -x_792 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_792, 0, x_771); -lean_ctor_set(x_792, 1, x_791); -lean_ctor_set(x_792, 2, x_790); -x_793 = lean_array_push(x_769, x_792); -x_794 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_794, 0, x_771); -lean_ctor_set(x_794, 1, x_772); -lean_ctor_set(x_794, 2, x_793); -x_795 = l_Lean_Elab_Command_mkSimpleDelab___closed__6; -lean_inc(x_758); -x_796 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_796, 0, x_758); -lean_ctor_set(x_796, 1, x_795); -x_797 = l_Lean_Elab_Command_mkSimpleDelab___closed__30; -lean_inc(x_759); -lean_inc(x_760); -x_798 = l_Lean_addMacroScope(x_760, x_797, x_759); -x_799 = l_Lean_Elab_Command_mkSimpleDelab___closed__29; -lean_inc(x_758); -x_800 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_800, 0, x_758); -lean_ctor_set(x_800, 1, x_799); -lean_ctor_set(x_800, 2, x_798); -lean_ctor_set(x_800, 3, x_765); -x_801 = lean_array_push(x_774, x_800); -x_802 = lean_array_push(x_801, x_768); -x_803 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_803, 0, x_771); -lean_ctor_set(x_803, 1, x_772); -lean_ctor_set(x_803, 2, x_802); -x_804 = l_Lean_Elab_Command_mkSimpleDelab___closed__31; -lean_inc(x_758); -x_805 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_805, 0, x_758); -lean_ctor_set(x_805, 1, x_804); -x_806 = l_Lean_Elab_Command_mkSimpleDelab___closed__38; -lean_inc(x_759); -lean_inc(x_760); -x_807 = l_Lean_addMacroScope(x_760, x_806, x_759); -x_808 = l_Lean_Elab_Command_mkSimpleDelab___closed__34; -x_809 = l_Lean_Elab_Command_mkSimpleDelab___closed__40; -lean_inc(x_758); -x_810 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_810, 0, x_758); -lean_ctor_set(x_810, 1, x_808); -lean_ctor_set(x_810, 2, x_807); -lean_ctor_set(x_810, 3, x_809); -x_811 = l_Lean_Elab_Command_mkSimpleDelab___closed__41; -lean_inc(x_758); -x_812 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_812, 0, x_758); -lean_ctor_set(x_812, 1, x_811); -x_813 = l_Lean_Elab_Command_mkSimpleDelab___closed__42; -lean_inc(x_758); -x_814 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_814, 0, x_758); -lean_ctor_set(x_814, 1, x_813); -x_815 = l_Lean_Elab_Command_mkSimpleDelab___closed__48; -lean_inc(x_758); -x_816 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_816, 0, x_758); -lean_ctor_set(x_816, 1, x_815); -x_817 = l_Lean_Elab_Command_mkSimpleDelab___closed__51; -lean_inc(x_758); -x_818 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_818, 0, x_758); -lean_ctor_set(x_818, 1, x_817); -x_819 = l_Lean_Elab_Command_mkSimpleDelab___closed__54; -lean_inc(x_758); -x_820 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_820, 0, x_758); -lean_ctor_set(x_820, 1, x_819); -x_821 = l_Lean_Elab_Command_mkSimpleDelab___closed__58; -lean_inc(x_759); -lean_inc(x_760); -x_822 = l_Lean_addMacroScope(x_760, x_821, x_759); -x_823 = l_Lean_Elab_Command_mkSimpleDelab___closed__57; -lean_inc(x_758); -x_824 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_824, 0, x_758); -lean_ctor_set(x_824, 1, x_823); -lean_ctor_set(x_824, 2, x_822); -lean_ctor_set(x_824, 3, x_765); -x_825 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___closed__1; -lean_inc(x_758); -x_826 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_826, 0, x_758); -lean_ctor_set(x_826, 1, x_825); -lean_inc(x_805); -x_827 = lean_array_push(x_774, x_805); -x_828 = lean_array_push(x_827, x_826); -x_829 = l_Lean_Elab_Command_mkSimpleDelab___closed__60; -x_830 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_830, 0, x_771); -lean_ctor_set(x_830, 1, x_829); -lean_ctor_set(x_830, 2, x_828); -x_831 = l_Lean_Elab_Command_mkSimpleDelab___closed__61; -x_832 = lean_array_push(x_831, x_820); -x_833 = l_Lean_Elab_Command_mkSimpleDelab___closed__11; -x_834 = lean_array_push(x_832, x_833); -lean_inc(x_824); -x_835 = lean_array_push(x_834, x_824); -x_836 = lean_array_push(x_835, x_830); -x_837 = l_Lean_Elab_Command_mkSimpleDelab___closed__53; -x_838 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_838, 0, x_771); -lean_ctor_set(x_838, 1, x_837); -lean_ctor_set(x_838, 2, x_836); -x_839 = l_Lean_Elab_Command_mkSimpleDelab___closed__10; -x_840 = l_Array_append___rarg(x_839, x_581); -x_841 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_841, 0, x_771); -lean_ctor_set(x_841, 1, x_772); -lean_ctor_set(x_841, 2, x_840); -x_842 = lean_array_push(x_774, x_838); -x_843 = lean_array_push(x_842, x_841); -x_844 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_844, 0, x_771); -lean_ctor_set(x_844, 1, x_6); -lean_ctor_set(x_844, 2, x_843); -x_845 = l_Lean_Elab_Command_mkSimpleDelab___closed__62; -lean_inc(x_758); -x_846 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_846, 0, x_758); -lean_ctor_set(x_846, 1, x_845); -x_847 = lean_array_push(x_787, x_818); -lean_inc(x_847); -x_848 = lean_array_push(x_847, x_844); -lean_inc(x_846); -x_849 = lean_array_push(x_848, x_846); -x_850 = l_Lean_Elab_Command_mkSimpleDelab___closed__50; -x_851 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_851, 0, x_771); -lean_ctor_set(x_851, 1, x_850); -lean_ctor_set(x_851, 2, x_849); -x_852 = lean_array_push(x_769, x_851); -x_853 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_853, 0, x_771); -lean_ctor_set(x_853, 1, x_772); -lean_ctor_set(x_853, 2, x_852); -x_854 = lean_array_push(x_769, x_853); -x_855 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_855, 0, x_771); -lean_ctor_set(x_855, 1, x_772); -lean_ctor_set(x_855, 2, x_854); -x_856 = l_Lean_Elab_Command_mkSimpleDelab___closed__63; -lean_inc(x_758); -x_857 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_857, 0, x_758); -lean_ctor_set(x_857, 1, x_856); -x_858 = l_Lean_Elab_Command_mkSimpleDelab___closed__67; -lean_inc(x_759); -lean_inc(x_760); -x_859 = l_Lean_addMacroScope(x_760, x_858, x_759); -x_860 = l_Lean_Elab_Command_mkSimpleDelab___closed__66; -x_861 = l_Lean_Elab_Command_mkSimpleDelab___closed__70; -lean_inc(x_758); -x_862 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_862, 0, x_758); -lean_ctor_set(x_862, 1, x_860); -lean_ctor_set(x_862, 2, x_859); -lean_ctor_set(x_862, 3, x_861); -x_863 = lean_array_push(x_847, x_2); -lean_inc(x_846); -x_864 = lean_array_push(x_863, x_846); -x_865 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_865, 0, x_771); -lean_ctor_set(x_865, 1, x_850); -lean_ctor_set(x_865, 2, x_864); -x_866 = lean_array_push(x_774, x_824); -x_867 = lean_array_push(x_866, x_865); -x_868 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_868, 0, x_771); -lean_ctor_set(x_868, 1, x_772); -lean_ctor_set(x_868, 2, x_867); -x_869 = lean_array_push(x_774, x_862); -x_870 = lean_array_push(x_869, x_868); -x_871 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_871, 0, x_771); -lean_ctor_set(x_871, 1, x_6); -lean_ctor_set(x_871, 2, x_870); -x_872 = lean_array_push(x_831, x_816); -lean_inc(x_872); -x_873 = lean_array_push(x_872, x_855); -lean_inc(x_857); -x_874 = lean_array_push(x_873, x_857); -x_875 = lean_array_push(x_874, x_871); -x_876 = l_Lean_Elab_Command_mkSimpleDelab___closed__47; -x_877 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_877, 0, x_771); -lean_ctor_set(x_877, 1, x_876); -lean_ctor_set(x_877, 2, x_875); -x_878 = l_Lean_Elab_Command_mkSimpleDelab___closed__73; -lean_inc(x_758); -x_879 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_879, 0, x_758); -lean_ctor_set(x_879, 1, x_878); -x_880 = lean_array_push(x_769, x_879); -x_881 = l_Lean_Elab_Command_mkSimpleDelab___closed__72; -x_882 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_882, 0, x_771); -lean_ctor_set(x_882, 1, x_881); -lean_ctor_set(x_882, 2, x_880); -x_883 = lean_array_push(x_769, x_882); -x_884 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_884, 0, x_771); -lean_ctor_set(x_884, 1, x_772); -lean_ctor_set(x_884, 2, x_883); -x_885 = lean_array_push(x_769, x_884); -x_886 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_886, 0, x_771); -lean_ctor_set(x_886, 1, x_772); -lean_ctor_set(x_886, 2, x_885); -x_887 = l_Lean_Elab_Command_mkSimpleDelab___closed__77; -x_888 = l_Lean_addMacroScope(x_760, x_887, x_759); -x_889 = l_Lean_Elab_Command_mkSimpleDelab___closed__76; -x_890 = l_Lean_Elab_Command_mkSimpleDelab___closed__82; -lean_inc(x_758); -x_891 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_891, 0, x_758); -lean_ctor_set(x_891, 1, x_889); -lean_ctor_set(x_891, 2, x_888); -lean_ctor_set(x_891, 3, x_890); -x_892 = l_Lean_Elab_Command_mkSimpleDelab___closed__83; -x_893 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_893, 0, x_758); -lean_ctor_set(x_893, 1, x_892); -x_894 = lean_array_push(x_787, x_893); -x_895 = lean_array_push(x_894, x_833); -x_896 = lean_array_push(x_895, x_846); -x_897 = l_Lean_Elab_Command_removeParentheses___closed__4; -x_898 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_898, 0, x_771); -lean_ctor_set(x_898, 1, x_897); -lean_ctor_set(x_898, 2, x_896); -x_899 = lean_array_push(x_769, x_898); -x_900 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_900, 0, x_771); -lean_ctor_set(x_900, 1, x_772); -lean_ctor_set(x_900, 2, x_899); -x_901 = lean_array_push(x_774, x_891); -x_902 = lean_array_push(x_901, x_900); -x_903 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_903, 0, x_771); -lean_ctor_set(x_903, 1, x_6); -lean_ctor_set(x_903, 2, x_902); -x_904 = lean_array_push(x_872, x_886); -x_905 = lean_array_push(x_904, x_857); -x_906 = lean_array_push(x_905, x_903); -x_907 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_907, 0, x_771); -lean_ctor_set(x_907, 1, x_876); -lean_ctor_set(x_907, 2, x_906); -x_908 = lean_array_push(x_774, x_877); -x_909 = lean_array_push(x_908, x_907); -x_910 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_910, 0, x_771); -lean_ctor_set(x_910, 1, x_772); -lean_ctor_set(x_910, 2, x_909); -x_911 = lean_array_push(x_769, x_910); -x_912 = l_Lean_Elab_Command_mkSimpleDelab___closed__45; -x_913 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_913, 0, x_771); -lean_ctor_set(x_913, 1, x_912); -lean_ctor_set(x_913, 2, x_911); -x_914 = lean_array_push(x_774, x_814); -x_915 = lean_array_push(x_914, x_913); -x_916 = l_Lean_Elab_Command_mkSimpleDelab___closed__43; -x_917 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_917, 0, x_771); -lean_ctor_set(x_917, 1, x_916); -lean_ctor_set(x_917, 2, x_915); -x_918 = l_Lean_Elab_Command_mkSimpleDelab___closed__85; -x_919 = lean_array_push(x_918, x_794); -x_920 = lean_array_push(x_919, x_796); -x_921 = lean_array_push(x_920, x_803); -x_922 = lean_array_push(x_921, x_805); -x_923 = lean_array_push(x_922, x_810); -x_924 = lean_array_push(x_923, x_812); -x_925 = lean_array_push(x_924, x_917); -x_926 = l_Lean_Elab_Command_mkSimpleDelab___closed__7; -x_927 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_927, 0, x_771); -lean_ctor_set(x_927, 1, x_926); -lean_ctor_set(x_927, 2, x_925); -x_928 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_928, 0, x_927); -lean_ctor_set(x_584, 0, x_928); -return x_584; -} -} -else -{ -lean_object* x_929; lean_object* x_930; lean_object* x_931; lean_object* x_932; lean_object* x_933; lean_object* x_934; lean_object* x_935; lean_object* x_936; lean_object* x_937; lean_object* x_938; lean_object* x_939; lean_object* x_940; lean_object* x_941; lean_object* x_942; lean_object* x_943; lean_object* x_944; lean_object* x_945; lean_object* x_946; lean_object* x_947; lean_object* x_948; lean_object* x_949; lean_object* x_950; lean_object* x_951; lean_object* x_952; lean_object* x_953; lean_object* x_954; lean_object* x_955; lean_object* x_956; lean_object* x_957; lean_object* x_958; lean_object* x_959; lean_object* x_960; lean_object* x_961; lean_object* x_962; lean_object* x_963; lean_object* x_964; lean_object* x_965; lean_object* x_966; lean_object* x_967; lean_object* x_968; lean_object* x_969; lean_object* x_970; lean_object* x_971; lean_object* x_972; lean_object* x_973; lean_object* x_974; lean_object* x_975; lean_object* x_976; lean_object* x_977; lean_object* x_978; lean_object* x_979; lean_object* x_980; lean_object* x_981; lean_object* x_982; lean_object* x_983; lean_object* x_984; lean_object* x_985; lean_object* x_986; lean_object* x_987; lean_object* x_988; lean_object* x_989; lean_object* x_990; lean_object* x_991; lean_object* x_992; lean_object* x_993; lean_object* x_994; lean_object* x_995; lean_object* x_996; lean_object* x_997; lean_object* x_998; lean_object* x_999; lean_object* x_1000; lean_object* x_1001; lean_object* x_1002; lean_object* x_1003; lean_object* x_1004; lean_object* x_1005; lean_object* x_1006; lean_object* x_1007; lean_object* x_1008; lean_object* x_1009; lean_object* x_1010; lean_object* x_1011; lean_object* x_1012; lean_object* x_1013; lean_object* x_1014; lean_object* x_1015; lean_object* x_1016; lean_object* x_1017; lean_object* x_1018; lean_object* x_1019; lean_object* x_1020; lean_object* x_1021; lean_object* x_1022; lean_object* x_1023; lean_object* x_1024; lean_object* x_1025; lean_object* x_1026; lean_object* x_1027; lean_object* x_1028; lean_object* x_1029; lean_object* x_1030; lean_object* x_1031; lean_object* x_1032; lean_object* x_1033; lean_object* x_1034; lean_object* x_1035; lean_object* x_1036; lean_object* x_1037; lean_object* x_1038; lean_object* x_1039; lean_object* x_1040; lean_object* x_1041; lean_object* x_1042; lean_object* x_1043; lean_object* x_1044; lean_object* x_1045; lean_object* x_1046; lean_object* x_1047; lean_object* x_1048; lean_object* x_1049; lean_object* x_1050; lean_object* x_1051; lean_object* x_1052; lean_object* x_1053; lean_object* x_1054; lean_object* x_1055; lean_object* x_1056; lean_object* x_1057; lean_object* x_1058; lean_object* x_1059; lean_object* x_1060; lean_object* x_1061; lean_object* x_1062; lean_object* x_1063; lean_object* x_1064; lean_object* x_1065; lean_object* x_1066; lean_object* x_1067; lean_object* x_1068; lean_object* x_1069; lean_object* x_1070; lean_object* x_1071; lean_object* x_1072; lean_object* x_1073; lean_object* x_1074; lean_object* x_1075; lean_object* x_1076; lean_object* x_1077; lean_object* x_1078; lean_object* x_1079; lean_object* x_1080; lean_object* x_1081; lean_object* x_1082; lean_object* x_1083; lean_object* x_1084; lean_object* x_1085; lean_object* x_1086; lean_object* x_1087; lean_object* x_1088; lean_object* x_1089; lean_object* x_1090; lean_object* x_1091; lean_object* x_1092; lean_object* x_1093; lean_object* x_1094; lean_object* x_1095; lean_object* x_1096; lean_object* x_1097; lean_object* x_1098; lean_object* x_1099; lean_object* x_1100; lean_object* x_1101; lean_object* x_1102; lean_object* x_1103; -x_929 = lean_ctor_get(x_584, 0); -x_930 = lean_ctor_get(x_584, 1); -lean_inc(x_930); -lean_inc(x_929); -lean_dec(x_584); -x_931 = lean_ctor_get(x_929, 0); -lean_inc(x_931); -if (lean_is_exclusive(x_929)) { - lean_ctor_release(x_929, 0); - x_932 = x_929; +lean_ctor_set(x_714, 0, x_587); +lean_ctor_set(x_714, 1, x_713); +lean_ctor_set(x_714, 2, x_712); +x_715 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__76; +x_716 = lean_array_push(x_715, x_610); +x_717 = lean_array_push(x_716, x_612); +x_718 = lean_array_push(x_717, x_619); +x_719 = lean_array_push(x_718, x_621); +x_720 = lean_array_push(x_719, x_626); +x_721 = lean_array_push(x_720, x_628); +x_722 = lean_array_push(x_721, x_714); +x_723 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__5; +x_724 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_724, 0, x_587); +lean_ctor_set(x_724, 1, x_723); +lean_ctor_set(x_724, 2, x_722); +if (lean_is_scalar(x_574)) { + x_725 = lean_alloc_ctor(1, 1, 0); } else { - lean_dec_ref(x_929); - x_932 = lean_box(0); + x_725 = x_574; } -x_933 = lean_ctor_get(x_4, 2); -lean_inc(x_933); -x_934 = lean_ctor_get(x_4, 1); -lean_inc(x_934); -lean_dec(x_4); -x_935 = l_Lean_Elab_Command_mkSimpleDelab___closed__14; -lean_inc(x_931); -x_936 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_936, 0, x_931); -lean_ctor_set(x_936, 1, x_935); -x_937 = l_Lean_Elab_Command_mkSimpleDelab___closed__24; -lean_inc(x_933); -lean_inc(x_934); -x_938 = l_Lean_addMacroScope(x_934, x_937, x_933); -x_939 = lean_box(0); -x_940 = l_Lean_Elab_Command_mkSimpleDelab___closed__23; -lean_inc(x_931); -x_941 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_941, 0, x_931); -lean_ctor_set(x_941, 1, x_940); -lean_ctor_set(x_941, 2, x_938); -lean_ctor_set(x_941, 3, x_939); -x_942 = lean_mk_syntax_ident(x_566); -x_943 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__13; -lean_inc(x_942); -x_944 = lean_array_push(x_943, x_942); -x_945 = lean_box(2); -x_946 = l_Lean_Elab_Command_mkSimpleDelab___closed__9; -x_947 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_947, 0, x_945); -lean_ctor_set(x_947, 1, x_946); -lean_ctor_set(x_947, 2, x_944); -x_948 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__16; -x_949 = lean_array_push(x_948, x_941); -x_950 = lean_array_push(x_949, x_947); -x_951 = l_Lean_Elab_Command_mkSimpleDelab___closed__20; -x_952 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_952, 0, x_945); -lean_ctor_set(x_952, 1, x_951); -lean_ctor_set(x_952, 2, x_950); -x_953 = lean_array_push(x_948, x_1); -x_954 = lean_array_push(x_953, x_952); -x_955 = l_Lean_Elab_Command_mkSimpleDelab___closed__16; -x_956 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_956, 0, x_945); -lean_ctor_set(x_956, 1, x_955); -lean_ctor_set(x_956, 2, x_954); -x_957 = lean_array_push(x_943, x_956); -x_958 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_958, 0, x_945); -lean_ctor_set(x_958, 1, x_946); -lean_ctor_set(x_958, 2, x_957); -x_959 = l_Lean_Elab_Command_mkSimpleDelab___closed__25; -lean_inc(x_931); -x_960 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_960, 0, x_931); -lean_ctor_set(x_960, 1, x_959); -x_961 = l_Lean_Elab_Command_mkSimpleDelab___closed__26; -x_962 = lean_array_push(x_961, x_936); -x_963 = lean_array_push(x_962, x_958); -x_964 = lean_array_push(x_963, x_960); -x_965 = l_Lean_Elab_Command_mkSimpleDelab___closed__13; -x_966 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_966, 0, x_945); -lean_ctor_set(x_966, 1, x_965); -lean_ctor_set(x_966, 2, x_964); -x_967 = lean_array_push(x_943, x_966); -x_968 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_968, 0, x_945); -lean_ctor_set(x_968, 1, x_946); -lean_ctor_set(x_968, 2, x_967); -x_969 = l_Lean_Elab_Command_mkSimpleDelab___closed__6; -lean_inc(x_931); -x_970 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_970, 0, x_931); -lean_ctor_set(x_970, 1, x_969); -x_971 = l_Lean_Elab_Command_mkSimpleDelab___closed__30; -lean_inc(x_933); -lean_inc(x_934); -x_972 = l_Lean_addMacroScope(x_934, x_971, x_933); -x_973 = l_Lean_Elab_Command_mkSimpleDelab___closed__29; -lean_inc(x_931); -x_974 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_974, 0, x_931); -lean_ctor_set(x_974, 1, x_973); -lean_ctor_set(x_974, 2, x_972); -lean_ctor_set(x_974, 3, x_939); -x_975 = lean_array_push(x_948, x_974); -x_976 = lean_array_push(x_975, x_942); -x_977 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_977, 0, x_945); -lean_ctor_set(x_977, 1, x_946); -lean_ctor_set(x_977, 2, x_976); -x_978 = l_Lean_Elab_Command_mkSimpleDelab___closed__31; -lean_inc(x_931); -x_979 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_979, 0, x_931); -lean_ctor_set(x_979, 1, x_978); -x_980 = l_Lean_Elab_Command_mkSimpleDelab___closed__38; -lean_inc(x_933); -lean_inc(x_934); -x_981 = l_Lean_addMacroScope(x_934, x_980, x_933); -x_982 = l_Lean_Elab_Command_mkSimpleDelab___closed__34; -x_983 = l_Lean_Elab_Command_mkSimpleDelab___closed__40; -lean_inc(x_931); -x_984 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_984, 0, x_931); -lean_ctor_set(x_984, 1, x_982); -lean_ctor_set(x_984, 2, x_981); -lean_ctor_set(x_984, 3, x_983); -x_985 = l_Lean_Elab_Command_mkSimpleDelab___closed__41; -lean_inc(x_931); -x_986 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_986, 0, x_931); -lean_ctor_set(x_986, 1, x_985); -x_987 = l_Lean_Elab_Command_mkSimpleDelab___closed__42; -lean_inc(x_931); -x_988 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_988, 0, x_931); -lean_ctor_set(x_988, 1, x_987); -x_989 = l_Lean_Elab_Command_mkSimpleDelab___closed__48; -lean_inc(x_931); -x_990 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_990, 0, x_931); -lean_ctor_set(x_990, 1, x_989); -x_991 = l_Lean_Elab_Command_mkSimpleDelab___closed__51; -lean_inc(x_931); -x_992 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_992, 0, x_931); -lean_ctor_set(x_992, 1, x_991); -x_993 = l_Lean_Elab_Command_mkSimpleDelab___closed__54; -lean_inc(x_931); -x_994 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_994, 0, x_931); -lean_ctor_set(x_994, 1, x_993); -x_995 = l_Lean_Elab_Command_mkSimpleDelab___closed__58; -lean_inc(x_933); -lean_inc(x_934); -x_996 = l_Lean_addMacroScope(x_934, x_995, x_933); -x_997 = l_Lean_Elab_Command_mkSimpleDelab___closed__57; -lean_inc(x_931); -x_998 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_998, 0, x_931); -lean_ctor_set(x_998, 1, x_997); -lean_ctor_set(x_998, 2, x_996); -lean_ctor_set(x_998, 3, x_939); -x_999 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___closed__1; -lean_inc(x_931); -x_1000 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1000, 0, x_931); -lean_ctor_set(x_1000, 1, x_999); -lean_inc(x_979); -x_1001 = lean_array_push(x_948, x_979); -x_1002 = lean_array_push(x_1001, x_1000); -x_1003 = l_Lean_Elab_Command_mkSimpleDelab___closed__60; -x_1004 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1004, 0, x_945); -lean_ctor_set(x_1004, 1, x_1003); -lean_ctor_set(x_1004, 2, x_1002); -x_1005 = l_Lean_Elab_Command_mkSimpleDelab___closed__61; -x_1006 = lean_array_push(x_1005, x_994); -x_1007 = l_Lean_Elab_Command_mkSimpleDelab___closed__11; -x_1008 = lean_array_push(x_1006, x_1007); -lean_inc(x_998); -x_1009 = lean_array_push(x_1008, x_998); -x_1010 = lean_array_push(x_1009, x_1004); -x_1011 = l_Lean_Elab_Command_mkSimpleDelab___closed__53; -x_1012 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1012, 0, x_945); -lean_ctor_set(x_1012, 1, x_1011); -lean_ctor_set(x_1012, 2, x_1010); -x_1013 = l_Lean_Elab_Command_mkSimpleDelab___closed__10; -x_1014 = l_Array_append___rarg(x_1013, x_581); -x_1015 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1015, 0, x_945); -lean_ctor_set(x_1015, 1, x_946); -lean_ctor_set(x_1015, 2, x_1014); -x_1016 = lean_array_push(x_948, x_1012); -x_1017 = lean_array_push(x_1016, x_1015); -x_1018 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1018, 0, x_945); -lean_ctor_set(x_1018, 1, x_6); -lean_ctor_set(x_1018, 2, x_1017); -x_1019 = l_Lean_Elab_Command_mkSimpleDelab___closed__62; -lean_inc(x_931); -x_1020 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1020, 0, x_931); -lean_ctor_set(x_1020, 1, x_1019); -x_1021 = lean_array_push(x_961, x_992); -lean_inc(x_1021); -x_1022 = lean_array_push(x_1021, x_1018); -lean_inc(x_1020); -x_1023 = lean_array_push(x_1022, x_1020); -x_1024 = l_Lean_Elab_Command_mkSimpleDelab___closed__50; -x_1025 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1025, 0, x_945); -lean_ctor_set(x_1025, 1, x_1024); -lean_ctor_set(x_1025, 2, x_1023); -x_1026 = lean_array_push(x_943, x_1025); -x_1027 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1027, 0, x_945); -lean_ctor_set(x_1027, 1, x_946); -lean_ctor_set(x_1027, 2, x_1026); -x_1028 = lean_array_push(x_943, x_1027); -x_1029 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1029, 0, x_945); -lean_ctor_set(x_1029, 1, x_946); -lean_ctor_set(x_1029, 2, x_1028); -x_1030 = l_Lean_Elab_Command_mkSimpleDelab___closed__63; -lean_inc(x_931); -x_1031 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1031, 0, x_931); -lean_ctor_set(x_1031, 1, x_1030); -x_1032 = l_Lean_Elab_Command_mkSimpleDelab___closed__67; -lean_inc(x_933); -lean_inc(x_934); -x_1033 = l_Lean_addMacroScope(x_934, x_1032, x_933); -x_1034 = l_Lean_Elab_Command_mkSimpleDelab___closed__66; -x_1035 = l_Lean_Elab_Command_mkSimpleDelab___closed__70; -lean_inc(x_931); -x_1036 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1036, 0, x_931); -lean_ctor_set(x_1036, 1, x_1034); -lean_ctor_set(x_1036, 2, x_1033); -lean_ctor_set(x_1036, 3, x_1035); -x_1037 = lean_array_push(x_1021, x_2); -lean_inc(x_1020); -x_1038 = lean_array_push(x_1037, x_1020); -x_1039 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1039, 0, x_945); -lean_ctor_set(x_1039, 1, x_1024); -lean_ctor_set(x_1039, 2, x_1038); -x_1040 = lean_array_push(x_948, x_998); -x_1041 = lean_array_push(x_1040, x_1039); -x_1042 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1042, 0, x_945); -lean_ctor_set(x_1042, 1, x_946); -lean_ctor_set(x_1042, 2, x_1041); -x_1043 = lean_array_push(x_948, x_1036); -x_1044 = lean_array_push(x_1043, x_1042); -x_1045 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1045, 0, x_945); -lean_ctor_set(x_1045, 1, x_6); -lean_ctor_set(x_1045, 2, x_1044); -x_1046 = lean_array_push(x_1005, x_990); -lean_inc(x_1046); -x_1047 = lean_array_push(x_1046, x_1029); -lean_inc(x_1031); -x_1048 = lean_array_push(x_1047, x_1031); -x_1049 = lean_array_push(x_1048, x_1045); -x_1050 = l_Lean_Elab_Command_mkSimpleDelab___closed__47; -x_1051 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1051, 0, x_945); -lean_ctor_set(x_1051, 1, x_1050); -lean_ctor_set(x_1051, 2, x_1049); -x_1052 = l_Lean_Elab_Command_mkSimpleDelab___closed__73; -lean_inc(x_931); -x_1053 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1053, 0, x_931); -lean_ctor_set(x_1053, 1, x_1052); -x_1054 = lean_array_push(x_943, x_1053); -x_1055 = l_Lean_Elab_Command_mkSimpleDelab___closed__72; -x_1056 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1056, 0, x_945); -lean_ctor_set(x_1056, 1, x_1055); -lean_ctor_set(x_1056, 2, x_1054); -x_1057 = lean_array_push(x_943, x_1056); -x_1058 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1058, 0, x_945); -lean_ctor_set(x_1058, 1, x_946); -lean_ctor_set(x_1058, 2, x_1057); -x_1059 = lean_array_push(x_943, x_1058); -x_1060 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1060, 0, x_945); -lean_ctor_set(x_1060, 1, x_946); -lean_ctor_set(x_1060, 2, x_1059); -x_1061 = l_Lean_Elab_Command_mkSimpleDelab___closed__77; -x_1062 = l_Lean_addMacroScope(x_934, x_1061, x_933); -x_1063 = l_Lean_Elab_Command_mkSimpleDelab___closed__76; -x_1064 = l_Lean_Elab_Command_mkSimpleDelab___closed__82; -lean_inc(x_931); -x_1065 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1065, 0, x_931); -lean_ctor_set(x_1065, 1, x_1063); -lean_ctor_set(x_1065, 2, x_1062); -lean_ctor_set(x_1065, 3, x_1064); -x_1066 = l_Lean_Elab_Command_mkSimpleDelab___closed__83; -x_1067 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1067, 0, x_931); -lean_ctor_set(x_1067, 1, x_1066); -x_1068 = lean_array_push(x_961, x_1067); -x_1069 = lean_array_push(x_1068, x_1007); -x_1070 = lean_array_push(x_1069, x_1020); -x_1071 = l_Lean_Elab_Command_removeParentheses___closed__4; -x_1072 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1072, 0, x_945); -lean_ctor_set(x_1072, 1, x_1071); -lean_ctor_set(x_1072, 2, x_1070); -x_1073 = lean_array_push(x_943, x_1072); -x_1074 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1074, 0, x_945); -lean_ctor_set(x_1074, 1, x_946); -lean_ctor_set(x_1074, 2, x_1073); -x_1075 = lean_array_push(x_948, x_1065); -x_1076 = lean_array_push(x_1075, x_1074); -x_1077 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1077, 0, x_945); -lean_ctor_set(x_1077, 1, x_6); -lean_ctor_set(x_1077, 2, x_1076); -x_1078 = lean_array_push(x_1046, x_1060); -x_1079 = lean_array_push(x_1078, x_1031); -x_1080 = lean_array_push(x_1079, x_1077); -x_1081 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1081, 0, x_945); -lean_ctor_set(x_1081, 1, x_1050); -lean_ctor_set(x_1081, 2, x_1080); -x_1082 = lean_array_push(x_948, x_1051); -x_1083 = lean_array_push(x_1082, x_1081); -x_1084 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1084, 0, x_945); -lean_ctor_set(x_1084, 1, x_946); -lean_ctor_set(x_1084, 2, x_1083); -x_1085 = lean_array_push(x_943, x_1084); -x_1086 = l_Lean_Elab_Command_mkSimpleDelab___closed__45; -x_1087 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1087, 0, x_945); -lean_ctor_set(x_1087, 1, x_1086); -lean_ctor_set(x_1087, 2, x_1085); -x_1088 = lean_array_push(x_948, x_988); -x_1089 = lean_array_push(x_1088, x_1087); -x_1090 = l_Lean_Elab_Command_mkSimpleDelab___closed__43; -x_1091 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1091, 0, x_945); -lean_ctor_set(x_1091, 1, x_1090); -lean_ctor_set(x_1091, 2, x_1089); -x_1092 = l_Lean_Elab_Command_mkSimpleDelab___closed__85; -x_1093 = lean_array_push(x_1092, x_968); -x_1094 = lean_array_push(x_1093, x_970); -x_1095 = lean_array_push(x_1094, x_977); -x_1096 = lean_array_push(x_1095, x_979); -x_1097 = lean_array_push(x_1096, x_984); -x_1098 = lean_array_push(x_1097, x_986); -x_1099 = lean_array_push(x_1098, x_1091); -x_1100 = l_Lean_Elab_Command_mkSimpleDelab___closed__7; -x_1101 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1101, 0, x_945); -lean_ctor_set(x_1101, 1, x_1100); -lean_ctor_set(x_1101, 2, x_1099); -if (lean_is_scalar(x_932)) { - x_1102 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_725, 0, x_724); +if (lean_is_scalar(x_572)) { + x_726 = lean_alloc_ctor(0, 2, 0); } else { - x_1102 = x_932; + x_726 = x_572; } -lean_ctor_set(x_1102, 0, x_1101); -x_1103 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1103, 0, x_1102); -lean_ctor_set(x_1103, 1, x_930); -return x_1103; +lean_ctor_set(x_726, 0, x_725); +lean_ctor_set(x_726, 1, x_571); +return x_726; } } else { -lean_object* x_1104; -lean_dec(x_581); -lean_dec(x_566); +lean_object* x_761; lean_object* x_762; +lean_dec(x_561); +lean_dec(x_21); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_1104 = lean_box(0); -lean_ctor_set(x_570, 0, x_1104); -return x_570; -} -} -else -{ -lean_object* x_1105; lean_object* x_1106; lean_object* x_1107; uint8_t x_1108; -x_1105 = lean_ctor_get(x_570, 1); -lean_inc(x_1105); -lean_dec(x_570); -x_1106 = lean_ctor_get(x_571, 0); -lean_inc(x_1106); -lean_dec(x_571); -x_1107 = l_Lean_Elab_Command_hasDuplicateAntiquot(x_1106); -x_1108 = lean_unbox(x_1107); -lean_dec(x_1107); -if (x_1108 == 0) -{ -lean_object* x_1109; lean_object* x_1110; lean_object* x_1111; lean_object* x_1112; lean_object* x_1113; lean_object* x_1114; lean_object* x_1115; lean_object* x_1116; lean_object* x_1117; lean_object* x_1118; lean_object* x_1119; lean_object* x_1120; lean_object* x_1121; lean_object* x_1122; lean_object* x_1123; lean_object* x_1124; lean_object* x_1125; lean_object* x_1126; lean_object* x_1127; lean_object* x_1128; lean_object* x_1129; lean_object* x_1130; lean_object* x_1131; lean_object* x_1132; lean_object* x_1133; lean_object* x_1134; lean_object* x_1135; lean_object* x_1136; lean_object* x_1137; lean_object* x_1138; lean_object* x_1139; lean_object* x_1140; lean_object* x_1141; lean_object* x_1142; lean_object* x_1143; lean_object* x_1144; lean_object* x_1145; lean_object* x_1146; lean_object* x_1147; lean_object* x_1148; lean_object* x_1149; lean_object* x_1150; lean_object* x_1151; lean_object* x_1152; lean_object* x_1153; lean_object* x_1154; lean_object* x_1155; lean_object* x_1156; lean_object* x_1157; lean_object* x_1158; lean_object* x_1159; lean_object* x_1160; lean_object* x_1161; lean_object* x_1162; lean_object* x_1163; lean_object* x_1164; lean_object* x_1165; lean_object* x_1166; lean_object* x_1167; lean_object* x_1168; lean_object* x_1169; lean_object* x_1170; lean_object* x_1171; lean_object* x_1172; lean_object* x_1173; lean_object* x_1174; lean_object* x_1175; lean_object* x_1176; lean_object* x_1177; lean_object* x_1178; lean_object* x_1179; lean_object* x_1180; lean_object* x_1181; lean_object* x_1182; lean_object* x_1183; lean_object* x_1184; lean_object* x_1185; lean_object* x_1186; lean_object* x_1187; lean_object* x_1188; lean_object* x_1189; lean_object* x_1190; lean_object* x_1191; lean_object* x_1192; lean_object* x_1193; lean_object* x_1194; lean_object* x_1195; lean_object* x_1196; lean_object* x_1197; lean_object* x_1198; lean_object* x_1199; lean_object* x_1200; lean_object* x_1201; lean_object* x_1202; lean_object* x_1203; lean_object* x_1204; lean_object* x_1205; lean_object* x_1206; lean_object* x_1207; lean_object* x_1208; lean_object* x_1209; lean_object* x_1210; lean_object* x_1211; lean_object* x_1212; lean_object* x_1213; lean_object* x_1214; lean_object* x_1215; lean_object* x_1216; lean_object* x_1217; lean_object* x_1218; lean_object* x_1219; lean_object* x_1220; lean_object* x_1221; lean_object* x_1222; lean_object* x_1223; lean_object* x_1224; lean_object* x_1225; lean_object* x_1226; lean_object* x_1227; lean_object* x_1228; lean_object* x_1229; lean_object* x_1230; lean_object* x_1231; lean_object* x_1232; lean_object* x_1233; lean_object* x_1234; lean_object* x_1235; lean_object* x_1236; lean_object* x_1237; lean_object* x_1238; lean_object* x_1239; lean_object* x_1240; lean_object* x_1241; lean_object* x_1242; lean_object* x_1243; lean_object* x_1244; lean_object* x_1245; lean_object* x_1246; lean_object* x_1247; lean_object* x_1248; lean_object* x_1249; lean_object* x_1250; lean_object* x_1251; lean_object* x_1252; lean_object* x_1253; lean_object* x_1254; lean_object* x_1255; lean_object* x_1256; lean_object* x_1257; lean_object* x_1258; lean_object* x_1259; lean_object* x_1260; lean_object* x_1261; lean_object* x_1262; lean_object* x_1263; lean_object* x_1264; lean_object* x_1265; lean_object* x_1266; lean_object* x_1267; lean_object* x_1268; lean_object* x_1269; lean_object* x_1270; lean_object* x_1271; lean_object* x_1272; lean_object* x_1273; lean_object* x_1274; lean_object* x_1275; lean_object* x_1276; lean_object* x_1277; lean_object* x_1278; lean_object* x_1279; lean_object* x_1280; lean_object* x_1281; lean_object* x_1282; lean_object* x_1283; lean_object* x_1284; lean_object* x_1285; -lean_inc(x_4); -x_1109 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Command_mkSimpleDelab___spec__1(x_4, x_1105); -x_1110 = lean_ctor_get(x_1109, 0); -lean_inc(x_1110); -x_1111 = lean_ctor_get(x_1109, 1); -lean_inc(x_1111); -if (lean_is_exclusive(x_1109)) { - lean_ctor_release(x_1109, 0); - lean_ctor_release(x_1109, 1); - x_1112 = x_1109; -} else { - lean_dec_ref(x_1109); - x_1112 = lean_box(0); -} -x_1113 = lean_ctor_get(x_1110, 0); -lean_inc(x_1113); -if (lean_is_exclusive(x_1110)) { - lean_ctor_release(x_1110, 0); - x_1114 = x_1110; -} else { - lean_dec_ref(x_1110); - x_1114 = lean_box(0); +x_761 = lean_box(0); +x_762 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_762, 0, x_761); +lean_ctor_set(x_762, 1, x_562); +return x_762; } -x_1115 = lean_ctor_get(x_4, 2); -lean_inc(x_1115); -x_1116 = lean_ctor_get(x_4, 1); -lean_inc(x_1116); -lean_dec(x_4); -x_1117 = l_Lean_Elab_Command_mkSimpleDelab___closed__14; -lean_inc(x_1113); -x_1118 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1118, 0, x_1113); -lean_ctor_set(x_1118, 1, x_1117); -x_1119 = l_Lean_Elab_Command_mkSimpleDelab___closed__24; -lean_inc(x_1115); -lean_inc(x_1116); -x_1120 = l_Lean_addMacroScope(x_1116, x_1119, x_1115); -x_1121 = lean_box(0); -x_1122 = l_Lean_Elab_Command_mkSimpleDelab___closed__23; -lean_inc(x_1113); -x_1123 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1123, 0, x_1113); -lean_ctor_set(x_1123, 1, x_1122); -lean_ctor_set(x_1123, 2, x_1120); -lean_ctor_set(x_1123, 3, x_1121); -x_1124 = lean_mk_syntax_ident(x_566); -x_1125 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__13; -lean_inc(x_1124); -x_1126 = lean_array_push(x_1125, x_1124); -x_1127 = lean_box(2); -x_1128 = l_Lean_Elab_Command_mkSimpleDelab___closed__9; -x_1129 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1129, 0, x_1127); -lean_ctor_set(x_1129, 1, x_1128); -lean_ctor_set(x_1129, 2, x_1126); -x_1130 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__16; -x_1131 = lean_array_push(x_1130, x_1123); -x_1132 = lean_array_push(x_1131, x_1129); -x_1133 = l_Lean_Elab_Command_mkSimpleDelab___closed__20; -x_1134 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1134, 0, x_1127); -lean_ctor_set(x_1134, 1, x_1133); -lean_ctor_set(x_1134, 2, x_1132); -x_1135 = lean_array_push(x_1130, x_1); -x_1136 = lean_array_push(x_1135, x_1134); -x_1137 = l_Lean_Elab_Command_mkSimpleDelab___closed__16; -x_1138 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1138, 0, x_1127); -lean_ctor_set(x_1138, 1, x_1137); -lean_ctor_set(x_1138, 2, x_1136); -x_1139 = lean_array_push(x_1125, x_1138); -x_1140 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1140, 0, x_1127); -lean_ctor_set(x_1140, 1, x_1128); -lean_ctor_set(x_1140, 2, x_1139); -x_1141 = l_Lean_Elab_Command_mkSimpleDelab___closed__25; -lean_inc(x_1113); -x_1142 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1142, 0, x_1113); -lean_ctor_set(x_1142, 1, x_1141); -x_1143 = l_Lean_Elab_Command_mkSimpleDelab___closed__26; -x_1144 = lean_array_push(x_1143, x_1118); -x_1145 = lean_array_push(x_1144, x_1140); -x_1146 = lean_array_push(x_1145, x_1142); -x_1147 = l_Lean_Elab_Command_mkSimpleDelab___closed__13; -x_1148 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1148, 0, x_1127); -lean_ctor_set(x_1148, 1, x_1147); -lean_ctor_set(x_1148, 2, x_1146); -x_1149 = lean_array_push(x_1125, x_1148); -x_1150 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1150, 0, x_1127); -lean_ctor_set(x_1150, 1, x_1128); -lean_ctor_set(x_1150, 2, x_1149); -x_1151 = l_Lean_Elab_Command_mkSimpleDelab___closed__6; -lean_inc(x_1113); -x_1152 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1152, 0, x_1113); -lean_ctor_set(x_1152, 1, x_1151); -x_1153 = l_Lean_Elab_Command_mkSimpleDelab___closed__30; -lean_inc(x_1115); -lean_inc(x_1116); -x_1154 = l_Lean_addMacroScope(x_1116, x_1153, x_1115); -x_1155 = l_Lean_Elab_Command_mkSimpleDelab___closed__29; -lean_inc(x_1113); -x_1156 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1156, 0, x_1113); -lean_ctor_set(x_1156, 1, x_1155); -lean_ctor_set(x_1156, 2, x_1154); -lean_ctor_set(x_1156, 3, x_1121); -x_1157 = lean_array_push(x_1130, x_1156); -x_1158 = lean_array_push(x_1157, x_1124); -x_1159 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1159, 0, x_1127); -lean_ctor_set(x_1159, 1, x_1128); -lean_ctor_set(x_1159, 2, x_1158); -x_1160 = l_Lean_Elab_Command_mkSimpleDelab___closed__31; -lean_inc(x_1113); -x_1161 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1161, 0, x_1113); -lean_ctor_set(x_1161, 1, x_1160); -x_1162 = l_Lean_Elab_Command_mkSimpleDelab___closed__38; -lean_inc(x_1115); -lean_inc(x_1116); -x_1163 = l_Lean_addMacroScope(x_1116, x_1162, x_1115); -x_1164 = l_Lean_Elab_Command_mkSimpleDelab___closed__34; -x_1165 = l_Lean_Elab_Command_mkSimpleDelab___closed__40; -lean_inc(x_1113); -x_1166 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1166, 0, x_1113); -lean_ctor_set(x_1166, 1, x_1164); -lean_ctor_set(x_1166, 2, x_1163); -lean_ctor_set(x_1166, 3, x_1165); -x_1167 = l_Lean_Elab_Command_mkSimpleDelab___closed__41; -lean_inc(x_1113); -x_1168 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1168, 0, x_1113); -lean_ctor_set(x_1168, 1, x_1167); -x_1169 = l_Lean_Elab_Command_mkSimpleDelab___closed__42; -lean_inc(x_1113); -x_1170 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1170, 0, x_1113); -lean_ctor_set(x_1170, 1, x_1169); -x_1171 = l_Lean_Elab_Command_mkSimpleDelab___closed__48; -lean_inc(x_1113); -x_1172 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1172, 0, x_1113); -lean_ctor_set(x_1172, 1, x_1171); -x_1173 = l_Lean_Elab_Command_mkSimpleDelab___closed__51; -lean_inc(x_1113); -x_1174 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1174, 0, x_1113); -lean_ctor_set(x_1174, 1, x_1173); -x_1175 = l_Lean_Elab_Command_mkSimpleDelab___closed__54; -lean_inc(x_1113); -x_1176 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1176, 0, x_1113); -lean_ctor_set(x_1176, 1, x_1175); -x_1177 = l_Lean_Elab_Command_mkSimpleDelab___closed__58; -lean_inc(x_1115); -lean_inc(x_1116); -x_1178 = l_Lean_addMacroScope(x_1116, x_1177, x_1115); -x_1179 = l_Lean_Elab_Command_mkSimpleDelab___closed__57; -lean_inc(x_1113); -x_1180 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1180, 0, x_1113); -lean_ctor_set(x_1180, 1, x_1179); -lean_ctor_set(x_1180, 2, x_1178); -lean_ctor_set(x_1180, 3, x_1121); -x_1181 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___closed__1; -lean_inc(x_1113); -x_1182 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1182, 0, x_1113); -lean_ctor_set(x_1182, 1, x_1181); -lean_inc(x_1161); -x_1183 = lean_array_push(x_1130, x_1161); -x_1184 = lean_array_push(x_1183, x_1182); -x_1185 = l_Lean_Elab_Command_mkSimpleDelab___closed__60; -x_1186 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1186, 0, x_1127); -lean_ctor_set(x_1186, 1, x_1185); -lean_ctor_set(x_1186, 2, x_1184); -x_1187 = l_Lean_Elab_Command_mkSimpleDelab___closed__61; -x_1188 = lean_array_push(x_1187, x_1176); -x_1189 = l_Lean_Elab_Command_mkSimpleDelab___closed__11; -x_1190 = lean_array_push(x_1188, x_1189); -lean_inc(x_1180); -x_1191 = lean_array_push(x_1190, x_1180); -x_1192 = lean_array_push(x_1191, x_1186); -x_1193 = l_Lean_Elab_Command_mkSimpleDelab___closed__53; -x_1194 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1194, 0, x_1127); -lean_ctor_set(x_1194, 1, x_1193); -lean_ctor_set(x_1194, 2, x_1192); -x_1195 = l_Lean_Elab_Command_mkSimpleDelab___closed__10; -x_1196 = l_Array_append___rarg(x_1195, x_1106); -x_1197 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1197, 0, x_1127); -lean_ctor_set(x_1197, 1, x_1128); -lean_ctor_set(x_1197, 2, x_1196); -x_1198 = lean_array_push(x_1130, x_1194); -x_1199 = lean_array_push(x_1198, x_1197); -x_1200 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1200, 0, x_1127); -lean_ctor_set(x_1200, 1, x_6); -lean_ctor_set(x_1200, 2, x_1199); -x_1201 = l_Lean_Elab_Command_mkSimpleDelab___closed__62; -lean_inc(x_1113); -x_1202 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1202, 0, x_1113); -lean_ctor_set(x_1202, 1, x_1201); -x_1203 = lean_array_push(x_1143, x_1174); -lean_inc(x_1203); -x_1204 = lean_array_push(x_1203, x_1200); -lean_inc(x_1202); -x_1205 = lean_array_push(x_1204, x_1202); -x_1206 = l_Lean_Elab_Command_mkSimpleDelab___closed__50; -x_1207 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1207, 0, x_1127); -lean_ctor_set(x_1207, 1, x_1206); -lean_ctor_set(x_1207, 2, x_1205); -x_1208 = lean_array_push(x_1125, x_1207); -x_1209 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1209, 0, x_1127); -lean_ctor_set(x_1209, 1, x_1128); -lean_ctor_set(x_1209, 2, x_1208); -x_1210 = lean_array_push(x_1125, x_1209); -x_1211 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1211, 0, x_1127); -lean_ctor_set(x_1211, 1, x_1128); -lean_ctor_set(x_1211, 2, x_1210); -x_1212 = l_Lean_Elab_Command_mkSimpleDelab___closed__63; -lean_inc(x_1113); -x_1213 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1213, 0, x_1113); -lean_ctor_set(x_1213, 1, x_1212); -x_1214 = l_Lean_Elab_Command_mkSimpleDelab___closed__67; -lean_inc(x_1115); -lean_inc(x_1116); -x_1215 = l_Lean_addMacroScope(x_1116, x_1214, x_1115); -x_1216 = l_Lean_Elab_Command_mkSimpleDelab___closed__66; -x_1217 = l_Lean_Elab_Command_mkSimpleDelab___closed__70; -lean_inc(x_1113); -x_1218 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1218, 0, x_1113); -lean_ctor_set(x_1218, 1, x_1216); -lean_ctor_set(x_1218, 2, x_1215); -lean_ctor_set(x_1218, 3, x_1217); -x_1219 = lean_array_push(x_1203, x_2); -lean_inc(x_1202); -x_1220 = lean_array_push(x_1219, x_1202); -x_1221 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1221, 0, x_1127); -lean_ctor_set(x_1221, 1, x_1206); -lean_ctor_set(x_1221, 2, x_1220); -x_1222 = lean_array_push(x_1130, x_1180); -x_1223 = lean_array_push(x_1222, x_1221); -x_1224 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1224, 0, x_1127); -lean_ctor_set(x_1224, 1, x_1128); -lean_ctor_set(x_1224, 2, x_1223); -x_1225 = lean_array_push(x_1130, x_1218); -x_1226 = lean_array_push(x_1225, x_1224); -x_1227 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1227, 0, x_1127); -lean_ctor_set(x_1227, 1, x_6); -lean_ctor_set(x_1227, 2, x_1226); -x_1228 = lean_array_push(x_1187, x_1172); -lean_inc(x_1228); -x_1229 = lean_array_push(x_1228, x_1211); -lean_inc(x_1213); -x_1230 = lean_array_push(x_1229, x_1213); -x_1231 = lean_array_push(x_1230, x_1227); -x_1232 = l_Lean_Elab_Command_mkSimpleDelab___closed__47; -x_1233 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1233, 0, x_1127); -lean_ctor_set(x_1233, 1, x_1232); -lean_ctor_set(x_1233, 2, x_1231); -x_1234 = l_Lean_Elab_Command_mkSimpleDelab___closed__73; -lean_inc(x_1113); -x_1235 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1235, 0, x_1113); -lean_ctor_set(x_1235, 1, x_1234); -x_1236 = lean_array_push(x_1125, x_1235); -x_1237 = l_Lean_Elab_Command_mkSimpleDelab___closed__72; -x_1238 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1238, 0, x_1127); -lean_ctor_set(x_1238, 1, x_1237); -lean_ctor_set(x_1238, 2, x_1236); -x_1239 = lean_array_push(x_1125, x_1238); -x_1240 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1240, 0, x_1127); -lean_ctor_set(x_1240, 1, x_1128); -lean_ctor_set(x_1240, 2, x_1239); -x_1241 = lean_array_push(x_1125, x_1240); -x_1242 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1242, 0, x_1127); -lean_ctor_set(x_1242, 1, x_1128); -lean_ctor_set(x_1242, 2, x_1241); -x_1243 = l_Lean_Elab_Command_mkSimpleDelab___closed__77; -x_1244 = l_Lean_addMacroScope(x_1116, x_1243, x_1115); -x_1245 = l_Lean_Elab_Command_mkSimpleDelab___closed__76; -x_1246 = l_Lean_Elab_Command_mkSimpleDelab___closed__82; -lean_inc(x_1113); -x_1247 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1247, 0, x_1113); -lean_ctor_set(x_1247, 1, x_1245); -lean_ctor_set(x_1247, 2, x_1244); -lean_ctor_set(x_1247, 3, x_1246); -x_1248 = l_Lean_Elab_Command_mkSimpleDelab___closed__83; -x_1249 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1249, 0, x_1113); -lean_ctor_set(x_1249, 1, x_1248); -x_1250 = lean_array_push(x_1143, x_1249); -x_1251 = lean_array_push(x_1250, x_1189); -x_1252 = lean_array_push(x_1251, x_1202); -x_1253 = l_Lean_Elab_Command_removeParentheses___closed__4; -x_1254 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1254, 0, x_1127); -lean_ctor_set(x_1254, 1, x_1253); -lean_ctor_set(x_1254, 2, x_1252); -x_1255 = lean_array_push(x_1125, x_1254); -x_1256 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1256, 0, x_1127); -lean_ctor_set(x_1256, 1, x_1128); -lean_ctor_set(x_1256, 2, x_1255); -x_1257 = lean_array_push(x_1130, x_1247); -x_1258 = lean_array_push(x_1257, x_1256); -x_1259 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1259, 0, x_1127); -lean_ctor_set(x_1259, 1, x_6); -lean_ctor_set(x_1259, 2, x_1258); -x_1260 = lean_array_push(x_1228, x_1242); -x_1261 = lean_array_push(x_1260, x_1213); -x_1262 = lean_array_push(x_1261, x_1259); -x_1263 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1263, 0, x_1127); -lean_ctor_set(x_1263, 1, x_1232); -lean_ctor_set(x_1263, 2, x_1262); -x_1264 = lean_array_push(x_1130, x_1233); -x_1265 = lean_array_push(x_1264, x_1263); -x_1266 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1266, 0, x_1127); -lean_ctor_set(x_1266, 1, x_1128); -lean_ctor_set(x_1266, 2, x_1265); -x_1267 = lean_array_push(x_1125, x_1266); -x_1268 = l_Lean_Elab_Command_mkSimpleDelab___closed__45; -x_1269 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1269, 0, x_1127); -lean_ctor_set(x_1269, 1, x_1268); -lean_ctor_set(x_1269, 2, x_1267); -x_1270 = lean_array_push(x_1130, x_1170); -x_1271 = lean_array_push(x_1270, x_1269); -x_1272 = l_Lean_Elab_Command_mkSimpleDelab___closed__43; -x_1273 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1273, 0, x_1127); -lean_ctor_set(x_1273, 1, x_1272); -lean_ctor_set(x_1273, 2, x_1271); -x_1274 = l_Lean_Elab_Command_mkSimpleDelab___closed__85; -x_1275 = lean_array_push(x_1274, x_1150); -x_1276 = lean_array_push(x_1275, x_1152); -x_1277 = lean_array_push(x_1276, x_1159); -x_1278 = lean_array_push(x_1277, x_1161); -x_1279 = lean_array_push(x_1278, x_1166); -x_1280 = lean_array_push(x_1279, x_1168); -x_1281 = lean_array_push(x_1280, x_1273); -x_1282 = l_Lean_Elab_Command_mkSimpleDelab___closed__7; -x_1283 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1283, 0, x_1127); -lean_ctor_set(x_1283, 1, x_1282); -lean_ctor_set(x_1283, 2, x_1281); -if (lean_is_scalar(x_1114)) { - x_1284 = lean_alloc_ctor(1, 1, 0); -} else { - x_1284 = x_1114; -} -lean_ctor_set(x_1284, 0, x_1283); -if (lean_is_scalar(x_1112)) { - x_1285 = lean_alloc_ctor(0, 2, 0); -} else { - x_1285 = x_1112; } -lean_ctor_set(x_1285, 0, x_1284); -lean_ctor_set(x_1285, 1, x_1111); -return x_1285; } else { -lean_object* x_1286; lean_object* x_1287; -lean_dec(x_1106); -lean_dec(x_566); +uint8_t x_763; +lean_dec(x_19); +lean_dec(x_17); +lean_dec(x_7); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_1286 = lean_box(0); -x_1287 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1287, 0, x_1286); -lean_ctor_set(x_1287, 1, x_1105); -return x_1287; +x_763 = !lean_is_exclusive(x_9); +if (x_763 == 0) +{ +lean_object* x_764; lean_object* x_765; +x_764 = lean_ctor_get(x_9, 0); +lean_dec(x_764); +x_765 = lean_box(0); +lean_ctor_set(x_9, 0, x_765); +return x_9; } +else +{ +lean_object* x_766; lean_object* x_767; lean_object* x_768; +x_766 = lean_ctor_get(x_9, 1); +lean_inc(x_766); +lean_dec(x_9); +x_767 = lean_box(0); +x_768 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_768, 0, x_767); +lean_ctor_set(x_768, 1, x_766); +return x_768; } } } else { -uint8_t x_1288; -lean_dec(x_566); +uint8_t x_769; +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_10); +lean_dec(x_7); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_1288 = !lean_is_exclusive(x_570); -if (x_1288 == 0) -{ -return x_570; +x_769 = !lean_is_exclusive(x_9); +if (x_769 == 0) +{ +lean_object* x_770; lean_object* x_771; +x_770 = lean_ctor_get(x_9, 0); +lean_dec(x_770); +x_771 = lean_box(0); +lean_ctor_set(x_9, 0, x_771); +return x_9; } else { -lean_object* x_1289; lean_object* x_1290; lean_object* x_1291; -x_1289 = lean_ctor_get(x_570, 0); -x_1290 = lean_ctor_get(x_570, 1); -lean_inc(x_1290); -lean_inc(x_1289); -lean_dec(x_570); -x_1291 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1291, 0, x_1289); -lean_ctor_set(x_1291, 1, x_1290); -return x_1291; +lean_object* x_772; lean_object* x_773; lean_object* x_774; +x_772 = lean_ctor_get(x_9, 1); +lean_inc(x_772); +lean_dec(x_9); +x_773 = lean_box(0); +x_774 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_774, 0, x_773); +lean_ctor_set(x_774, 1, x_772); +return x_774; +} } } } else { -uint8_t x_1292; -lean_dec(x_564); -lean_dec(x_562); -lean_dec(x_552); +uint8_t x_775; +lean_dec(x_7); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_1292 = !lean_is_exclusive(x_554); -if (x_1292 == 0) +x_775 = !lean_is_exclusive(x_9); +if (x_775 == 0) { -lean_object* x_1293; lean_object* x_1294; -x_1293 = lean_ctor_get(x_554, 0); -lean_dec(x_1293); -x_1294 = lean_box(0); -lean_ctor_set(x_554, 0, x_1294); -return x_554; +return x_9; } else { -lean_object* x_1295; lean_object* x_1296; lean_object* x_1297; -x_1295 = lean_ctor_get(x_554, 1); -lean_inc(x_1295); -lean_dec(x_554); -x_1296 = lean_box(0); -x_1297 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1297, 0, x_1296); -lean_ctor_set(x_1297, 1, x_1295); -return x_1297; +lean_object* x_776; lean_object* x_777; lean_object* x_778; +x_776 = lean_ctor_get(x_9, 0); +x_777 = lean_ctor_get(x_9, 1); +lean_inc(x_777); +lean_inc(x_776); +lean_dec(x_9); +x_778 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_778, 0, x_776); +lean_ctor_set(x_778, 1, x_777); +return x_778; } } } -else +} +LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkSimpleDelab(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -uint8_t x_1298; -lean_dec(x_563); -lean_dec(x_562); -lean_dec(x_555); -lean_dec(x_552); +lean_object* x_6; uint8_t x_7; +x_6 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__49; +lean_inc(x_3); +x_7 = l_Lean_Syntax_isOfKind(x_3, x_6); +if (x_7 == 0) +{ +lean_object* x_8; uint8_t x_9; +x_8 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___closed__2; +lean_inc(x_3); +x_9 = l_Lean_Syntax_isOfKind(x_3, x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_1298 = !lean_is_exclusive(x_554); -if (x_1298 == 0) -{ -lean_object* x_1299; lean_object* x_1300; -x_1299 = lean_ctor_get(x_554, 0); -lean_dec(x_1299); -x_1300 = lean_box(0); -lean_ctor_set(x_554, 0, x_1300); -return x_554; +x_10 = lean_box(0); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_5); +return x_11; } else { -lean_object* x_1301; lean_object* x_1302; lean_object* x_1303; -x_1301 = lean_ctor_get(x_554, 1); -lean_inc(x_1301); -lean_dec(x_554); -x_1302 = lean_box(0); -x_1303 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_1303, 0, x_1302); -lean_ctor_set(x_1303, 1, x_1301); -return x_1303; -} -} +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__7; +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_3); +lean_ctor_set(x_13, 1, x_12); +x_14 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1(x_1, x_2, x_13, x_4, x_5); +return x_14; } } else { -uint8_t x_1304; -lean_dec(x_552); +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_15 = lean_unsigned_to_nat(0u); +x_16 = l_Lean_Syntax_getArg(x_3, x_15); +x_17 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___closed__2; +lean_inc(x_16); +x_18 = l_Lean_Syntax_isOfKind(x_16, x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +lean_dec(x_16); lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_1304 = !lean_is_exclusive(x_554); -if (x_1304 == 0) -{ -return x_554; +x_19 = lean_box(0); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_5); +return x_20; } else { -lean_object* x_1305; lean_object* x_1306; lean_object* x_1307; -x_1305 = lean_ctor_get(x_554, 0); -x_1306 = lean_ctor_get(x_554, 1); -lean_inc(x_1306); -lean_inc(x_1305); -lean_dec(x_554); -x_1307 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1307, 0, x_1305); -lean_ctor_set(x_1307, 1, x_1306); -return x_1307; -} -} -} -} -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_mkSimpleDelab___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) { -_start: -{ -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_8 = lean_unbox_usize(x_3); +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_21 = lean_unsigned_to_nat(1u); +x_22 = l_Lean_Syntax_getArg(x_3, x_21); lean_dec(x_3); -x_9 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_mkSimpleDelab___spec__3(x_1, x_7, x_8, x_4, x_5, x_6); -return x_9; +x_23 = l_Lean_Syntax_getArgs(x_22); +lean_dec(x_22); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_16); +lean_ctor_set(x_24, 1, x_23); +x_25 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1(x_1, x_2, x_24, x_4, x_5); +return x_25; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_mkSimpleDelab___spec__3___at_Lean_Elab_Command_mkSimpleDelab___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -size_t x_6; size_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_1); -lean_dec(x_1); -x_7 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_8 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_mkSimpleDelab___spec__3___at_Lean_Elab_Command_mkSimpleDelab___spec__4(x_6, x_7, x_3, x_4, x_5); -return x_8; } } static lean_object* _init_l___private_Lean_Elab_Notation_0__Lean_Elab_Command_isLocalAttrKind___closed__1() { @@ -6315,29 +5154,28 @@ return x_4; } else { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Syntax_getArg(x_1, x_5); lean_dec(x_1); -x_7 = l_Lean_nullKind; -x_8 = lean_unsigned_to_nat(1u); +x_7 = lean_unsigned_to_nat(1u); lean_inc(x_6); -x_9 = l_Lean_Syntax_isNodeOf(x_6, x_7, x_8); -if (x_9 == 0) +x_8 = l_Lean_Syntax_matchesNull(x_6, x_7); +if (x_8 == 0) { -uint8_t x_10; +uint8_t x_9; lean_dec(x_6); -x_10 = 0; -return x_10; +x_9 = 0; +return x_9; } else { -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = l_Lean_Syntax_getArg(x_6, x_5); +lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_10 = l_Lean_Syntax_getArg(x_6, x_5); lean_dec(x_6); -x_12 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_isLocalAttrKind___closed__4; -x_13 = l_Lean_Syntax_isOfKind(x_11, x_12); -return x_13; +x_11 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_isLocalAttrKind___closed__4; +x_12 = l_Lean_Syntax_isOfKind(x_10, x_11); +return x_12; } } } @@ -6359,6 +5197,7 @@ x_6 = lean_usize_dec_lt(x_2, x_1); if (x_6 == 0) { lean_object* x_7; +lean_dec(x_4); x_7 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_7, 0, x_3); lean_ctor_set(x_7, 1, x_5); @@ -6370,6 +5209,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_array_uget(x_3, x_2); x_9 = lean_unsigned_to_nat(0u); x_10 = lean_array_uset(x_3, x_2, x_9); +lean_inc(x_4); x_11 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem(x_8, x_4, x_5); if (lean_obj_tag(x_11) == 0) { @@ -6391,6 +5231,7 @@ else { uint8_t x_18; lean_dec(x_10); +lean_dec(x_4); x_18 = !lean_is_exclusive(x_11); if (x_18 == 0) { @@ -6501,7 +5342,31 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___spec__4(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___spec__4(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = 1; +x_9 = lean_usize_add(x_2, x_8); +x_10 = lean_array_uset(x_7, x_2, x_5); +x_2 = x_9; +x_3 = x_10; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___spec__5(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -6551,118 +5416,115 @@ lean_inc(x_12); if (lean_obj_tag(x_12) == 0) { uint8_t x_13; -lean_dec(x_6); +lean_dec(x_7); x_13 = !lean_is_exclusive(x_11); 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_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; x_14 = lean_ctor_get(x_11, 0); lean_dec(x_14); x_15 = lean_array_push(x_4, x_5); -x_16 = lean_array_push(x_15, x_7); +x_16 = lean_array_push(x_15, x_8); x_17 = lean_box(2); -x_18 = l_Lean_nullKind; -x_19 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set(x_19, 1, x_18); -lean_ctor_set(x_19, 2, x_16); -lean_ctor_set(x_11, 0, x_19); +x_18 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_6); +lean_ctor_set(x_18, 2, x_16); +lean_ctor_set(x_11, 0, x_18); return x_11; } else { -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; -x_20 = lean_ctor_get(x_11, 1); -lean_inc(x_20); +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_19 = lean_ctor_get(x_11, 1); +lean_inc(x_19); lean_dec(x_11); -x_21 = lean_array_push(x_4, x_5); -x_22 = lean_array_push(x_21, x_7); -x_23 = lean_box(2); -x_24 = l_Lean_nullKind; -x_25 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_25, 0, x_23); -lean_ctor_set(x_25, 1, x_24); -lean_ctor_set(x_25, 2, x_22); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_20); -return x_26; +x_20 = lean_array_push(x_4, x_5); +x_21 = lean_array_push(x_20, x_8); +x_22 = lean_box(2); +x_23 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_6); +lean_ctor_set(x_23, 2, x_21); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_19); +return x_24; } } else { -uint8_t x_27; +uint8_t x_25; lean_dec(x_4); -x_27 = !lean_is_exclusive(x_11); -if (x_27 == 0) +x_25 = !lean_is_exclusive(x_11); +if (x_25 == 0) { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_28 = lean_ctor_get(x_11, 0); -lean_dec(x_28); -x_29 = lean_ctor_get(x_12, 0); -lean_inc(x_29); +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_object* x_32; +x_26 = lean_ctor_get(x_11, 0); +lean_dec(x_26); +x_27 = lean_ctor_get(x_12, 0); +lean_inc(x_27); lean_dec(x_12); -x_30 = lean_array_push(x_6, x_5); -x_31 = lean_array_push(x_30, x_7); -x_32 = lean_array_push(x_31, x_29); -x_33 = lean_box(2); -x_34 = l_Lean_nullKind; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_11, 0, x_35); +x_28 = lean_array_push(x_7, x_5); +x_29 = lean_array_push(x_28, x_8); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_box(2); +x_32 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_6); +lean_ctor_set(x_32, 2, x_30); +lean_ctor_set(x_11, 0, x_32); return x_11; } else { -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; lean_object* x_44; -x_36 = lean_ctor_get(x_11, 1); -lean_inc(x_36); +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; +x_33 = lean_ctor_get(x_11, 1); +lean_inc(x_33); lean_dec(x_11); -x_37 = lean_ctor_get(x_12, 0); -lean_inc(x_37); +x_34 = lean_ctor_get(x_12, 0); +lean_inc(x_34); lean_dec(x_12); -x_38 = lean_array_push(x_6, x_5); -x_39 = lean_array_push(x_38, x_7); -x_40 = lean_array_push(x_39, x_37); -x_41 = lean_box(2); -x_42 = l_Lean_nullKind; -x_43 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_43, 0, x_41); -lean_ctor_set(x_43, 1, x_42); -lean_ctor_set(x_43, 2, x_40); -x_44 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_44, 1, x_36); -return x_44; +x_35 = lean_array_push(x_7, x_5); +x_36 = lean_array_push(x_35, x_8); +x_37 = lean_array_push(x_36, x_34); +x_38 = lean_box(2); +x_39 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_6); +lean_ctor_set(x_39, 2, x_37); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_33); +return x_40; } } } else { -uint8_t x_45; +uint8_t x_41; +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_45 = !lean_is_exclusive(x_11); -if (x_45 == 0) +x_41 = !lean_is_exclusive(x_11); +if (x_41 == 0) { return x_11; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_11, 0); -x_47 = lean_ctor_get(x_11, 1); -lean_inc(x_47); -lean_inc(x_46); +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_11, 0); +x_43 = lean_ctor_get(x_11, 1); +lean_inc(x_43); +lean_inc(x_42); lean_dec(x_11); -x_48 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_48, 0, x_46); -lean_ctor_set(x_48, 1, x_47); -return x_48; +x_44 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_44, 0, x_42); +lean_ctor_set(x_44, 1, x_43); +return x_44; } } } @@ -6690,8 +5552,8 @@ static lean_object* _init_l___private_Lean_Elab_Notation_0__Lean_Elab_Command_ex { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__9; -x_3 = l_Lean_Elab_Command_mkSimpleDelab___closed__10; +x_2 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__6; +x_3 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__7; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -6703,7 +5565,7 @@ static lean_object* _init_l___private_Lean_Elab_Notation_0__Lean_Elab_Command_ex _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__13; +x_1 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__13; x_2 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__3; x_3 = lean_array_push(x_1, x_2); return x_3; @@ -6996,47 +5858,6 @@ x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__36() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Command_mkSimpleDelab___closed__10; -x_2 = l_Array_append___rarg(x_1, x_1); -return x_2; -} -} -static lean_object* _init_l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__37() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = lean_box(2); -x_2 = l_Lean_Elab_Command_mkSimpleDelab___closed__9; -x_3 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__36; -x_4 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_4, 0, x_1); -lean_ctor_set(x_4, 1, x_2); -lean_ctor_set(x_4, 2, x_3); -return x_4; -} -} -static lean_object* _init_l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__38() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("precedence", 10); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__39() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__4; -x_2 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__38; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} LEAN_EXPORT lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2(lean_object* x_1, size_t x_2, size_t 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, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { @@ -7047,41 +5868,41 @@ lean_inc(x_4); x_17 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___spec__2(x_2, x_3, x_4, x_13, x_14); if (x_16 == 0) { -lean_object* x_207; +lean_object* x_206; lean_dec(x_4); lean_dec(x_1); -x_207 = l_Lean_Elab_Command_mkSimpleDelab___closed__10; -x_18 = x_207; -goto block_206; +x_206 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__7; +x_18 = x_206; +goto block_205; } else { -uint8_t x_208; -x_208 = lean_nat_dec_le(x_1, x_1); +uint8_t x_207; +x_207 = lean_nat_dec_le(x_1, x_1); lean_dec(x_1); -if (x_208 == 0) +if (x_207 == 0) { -lean_object* x_209; +lean_object* x_208; lean_dec(x_4); -x_209 = l_Lean_Elab_Command_mkSimpleDelab___closed__10; -x_18 = x_209; -goto block_206; +x_208 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__7; +x_18 = x_208; +goto block_205; } else { -lean_object* x_210; lean_object* x_211; -x_210 = l_Lean_Elab_Command_mkSimpleDelab___closed__10; -x_211 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___spec__4(x_4, x_3, x_2, x_210); +lean_object* x_209; lean_object* x_210; +x_209 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__7; +x_210 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___spec__5(x_4, x_3, x_2, x_209); lean_dec(x_4); -x_18 = x_211; -goto block_206; +x_18 = x_210; +goto block_205; } } -block_206: +block_205: { if (lean_obj_tag(x_17) == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; size_t 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; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; size_t x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; +lean_object* x_19; lean_object* x_20; lean_object* x_21; size_t 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; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; size_t x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; x_19 = lean_ctor_get(x_17, 0); lean_inc(x_19); x_20 = lean_ctor_get(x_17, 1); @@ -7113,7 +5934,7 @@ lean_inc(x_127); x_130 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_130, 0, x_127); lean_ctor_set(x_130, 1, x_129); -x_131 = l_Lean_Elab_Command_mkSimpleDelab___closed__83; +x_131 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__74; lean_inc(x_127); x_132 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_132, 0, x_127); @@ -7123,13 +5944,13 @@ lean_inc(x_127); x_134 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_134, 0, x_127); lean_ctor_set(x_134, 1, x_133); -x_135 = l_Lean_Elab_Command_mkSimpleDelab___closed__41; +x_135 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__35; lean_inc(x_127); x_136 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_136, 0, x_127); lean_ctor_set(x_136, 1, x_135); x_137 = lean_mk_syntax_ident(x_12); -x_138 = l_Lean_Elab_Command_mkSimpleDelab___closed__62; +x_138 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__46; lean_inc(x_127); x_139 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_139, 0, x_127); @@ -7148,9 +5969,9 @@ x_147 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_147, 0, x_26); lean_ctor_set(x_147, 1, x_146); lean_ctor_set(x_147, 2, x_145); -x_148 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__13; +x_148 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__13; x_149 = lean_array_push(x_148, x_147); -x_150 = l_Lean_Elab_Command_mkSimpleDelab___closed__9; +x_150 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__6; x_151 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_151, 0, x_26); lean_ctor_set(x_151, 1, x_150); @@ -7161,92 +5982,91 @@ x_153 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_153, 0, x_127); lean_ctor_set(x_153, 1, x_152); x_154 = l_Nat_repr(x_8); -x_155 = l_Lean_numLitKind; -x_156 = l_Lean_Syntax_mkLit(x_155, x_154, x_26); -x_157 = lean_array_push(x_141, x_153); -x_158 = lean_array_push(x_157, x_136); -x_159 = lean_array_push(x_158, x_156); -x_160 = lean_array_push(x_159, x_139); -x_161 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__32; -x_162 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_162, 0, x_26); -lean_ctor_set(x_162, 1, x_161); -lean_ctor_set(x_162, 2, x_160); -x_163 = lean_array_push(x_148, x_162); -x_164 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_164, 0, x_26); -lean_ctor_set(x_164, 1, x_150); -lean_ctor_set(x_164, 2, x_163); -x_165 = lean_array_get_size(x_9); -x_166 = lean_usize_of_nat(x_165); -lean_dec(x_165); -x_167 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_166, x_3, x_9); -x_168 = l_Lean_Elab_Command_mkSimpleDelab___closed__10; -x_169 = l_Array_append___rarg(x_168, x_167); -x_170 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_170, 0, x_26); -lean_ctor_set(x_170, 1, x_150); -lean_ctor_set(x_170, 2, x_169); -x_171 = l_Lean_Elab_Command_mkSimpleDelab___closed__31; -x_172 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_172, 0, x_127); -lean_ctor_set(x_172, 1, x_171); -x_173 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__35; +x_155 = l_Lean_Syntax_mkNumLit(x_154, x_26); +x_156 = lean_array_push(x_141, x_153); +x_157 = lean_array_push(x_156, x_136); +x_158 = lean_array_push(x_157, x_155); +x_159 = lean_array_push(x_158, x_139); +x_160 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__32; +x_161 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_161, 0, x_26); +lean_ctor_set(x_161, 1, x_160); +lean_ctor_set(x_161, 2, x_159); +x_162 = lean_array_push(x_148, x_161); +x_163 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_163, 0, x_26); +lean_ctor_set(x_163, 1, x_150); +lean_ctor_set(x_163, 2, x_162); +x_164 = lean_array_get_size(x_9); +x_165 = lean_usize_of_nat(x_164); +lean_dec(x_164); +x_166 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___spec__4(x_165, x_3, x_9); +x_167 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__7; +x_168 = l_Array_append___rarg(x_167, x_166); +x_169 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_169, 0, x_26); +lean_ctor_set(x_169, 1, x_150); +lean_ctor_set(x_169, 2, x_168); +x_170 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__12; +x_171 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_171, 0, x_127); +lean_ctor_set(x_171, 1, x_170); +x_172 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__35; lean_inc(x_7); -x_174 = lean_array_push(x_173, x_7); -x_175 = lean_array_push(x_174, x_130); +x_173 = lean_array_push(x_172, x_7); +x_174 = lean_array_push(x_173, x_130); if (lean_obj_tag(x_10) == 0) { -lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; -x_176 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__37; -x_177 = lean_array_push(x_175, x_176); -x_178 = lean_array_push(x_177, x_151); -x_179 = lean_array_push(x_178, x_164); -x_180 = lean_array_push(x_179, x_170); -x_181 = lean_array_push(x_180, x_172); -x_182 = lean_array_push(x_181, x_11); -x_183 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__27; -x_184 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_184, 0, x_26); -lean_ctor_set(x_184, 1, x_183); -lean_ctor_set(x_184, 2, x_182); -x_28 = x_184; +lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; +x_175 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__10; +x_176 = lean_array_push(x_174, x_175); +x_177 = lean_array_push(x_176, x_151); +x_178 = lean_array_push(x_177, x_163); +x_179 = lean_array_push(x_178, x_169); +x_180 = lean_array_push(x_179, x_171); +x_181 = lean_array_push(x_180, x_11); +x_182 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__27; +x_183 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_183, 0, x_26); +lean_ctor_set(x_183, 1, x_182); +lean_ctor_set(x_183, 2, x_181); +x_28 = x_183; x_29 = x_128; goto block_125; } else { -lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; -x_185 = lean_ctor_get(x_10, 0); -lean_inc(x_185); +lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; +x_184 = lean_ctor_get(x_10, 0); +lean_inc(x_184); lean_dec(x_10); -x_186 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__16; -lean_inc(x_172); -x_187 = lean_array_push(x_186, x_172); -x_188 = lean_array_push(x_187, x_185); -x_189 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__39; -x_190 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_190, 0, x_26); -lean_ctor_set(x_190, 1, x_189); -lean_ctor_set(x_190, 2, x_188); -x_191 = lean_array_push(x_148, x_190); -x_192 = l_Array_append___rarg(x_168, x_191); -x_193 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_193, 0, x_26); -lean_ctor_set(x_193, 1, x_150); -lean_ctor_set(x_193, 2, x_192); -x_194 = lean_array_push(x_175, x_193); -x_195 = lean_array_push(x_194, x_151); -x_196 = lean_array_push(x_195, x_164); -x_197 = lean_array_push(x_196, x_170); -x_198 = lean_array_push(x_197, x_172); -x_199 = lean_array_push(x_198, x_11); -x_200 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__27; -x_201 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_201, 0, x_26); -lean_ctor_set(x_201, 1, x_200); -lean_ctor_set(x_201, 2, x_199); -x_28 = x_201; +x_185 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__8; +lean_inc(x_171); +x_186 = lean_array_push(x_185, x_171); +x_187 = lean_array_push(x_186, x_184); +x_188 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__14; +x_189 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_189, 0, x_26); +lean_ctor_set(x_189, 1, x_188); +lean_ctor_set(x_189, 2, x_187); +x_190 = lean_array_push(x_148, x_189); +x_191 = l_Array_append___rarg(x_167, x_190); +x_192 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_192, 0, x_26); +lean_ctor_set(x_192, 1, x_150); +lean_ctor_set(x_192, 2, x_191); +x_193 = lean_array_push(x_174, x_192); +x_194 = lean_array_push(x_193, x_151); +x_195 = lean_array_push(x_194, x_163); +x_196 = lean_array_push(x_195, x_169); +x_197 = lean_array_push(x_196, x_171); +x_198 = lean_array_push(x_197, x_11); +x_199 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__27; +x_200 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_200, 0, x_26); +lean_ctor_set(x_200, 1, x_199); +lean_ctor_set(x_200, 2, x_198); +x_28 = x_200; x_29 = x_128; goto block_125; } @@ -7265,36 +6085,36 @@ lean_inc(x_31); x_34 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_34, 0, x_31); lean_ctor_set(x_34, 1, x_33); -x_35 = l_Lean_Elab_Command_mkSimpleDelab___closed__48; +x_35 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__42; lean_inc(x_31); x_36 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_36, 0, x_31); lean_ctor_set(x_36, 1, x_35); -x_37 = l_Lean_Elab_Command_mkSimpleDelab___closed__51; +x_37 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__45; lean_inc(x_31); x_38 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_38, 0, x_31); lean_ctor_set(x_38, 1, x_37); -x_39 = l_Lean_Elab_Command_mkSimpleDelab___closed__62; +x_39 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__46; lean_inc(x_31); x_40 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_40, 0, x_31); lean_ctor_set(x_40, 1, x_39); -x_41 = l_Lean_Elab_Command_mkSimpleDelab___closed__26; +x_41 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__21; x_42 = lean_array_push(x_41, x_38); lean_inc(x_27); lean_inc(x_42); x_43 = lean_array_push(x_42, x_27); lean_inc(x_40); x_44 = lean_array_push(x_43, x_40); -x_45 = l_Lean_Elab_Command_mkSimpleDelab___closed__50; +x_45 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__44; x_46 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_46, 0, x_26); lean_ctor_set(x_46, 1, x_45); lean_ctor_set(x_46, 2, x_44); -x_47 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__13; +x_47 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__13; x_48 = lean_array_push(x_47, x_46); -x_49 = l_Lean_Elab_Command_mkSimpleDelab___closed__9; +x_49 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__6; x_50 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_50, 0, x_26); lean_ctor_set(x_50, 1, x_49); @@ -7304,7 +6124,7 @@ x_52 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_52, 0, x_26); lean_ctor_set(x_52, 1, x_49); lean_ctor_set(x_52, 2, x_51); -x_53 = l_Lean_Elab_Command_mkSimpleDelab___closed__63; +x_53 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__47; lean_inc(x_31); x_54 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_54, 0, x_31); @@ -7320,7 +6140,7 @@ x_59 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_59, 0, x_26); lean_ctor_set(x_59, 1, x_45); lean_ctor_set(x_59, 2, x_58); -x_60 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__16; +x_60 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__8; x_61 = lean_array_push(x_60, x_56); x_62 = lean_array_push(x_61, x_59); x_63 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__7; @@ -7328,12 +6148,12 @@ x_64 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_64, 0, x_26); lean_ctor_set(x_64, 1, x_63); lean_ctor_set(x_64, 2, x_62); -x_65 = l_Lean_Elab_Command_mkSimpleDelab___closed__61; +x_65 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__61; x_66 = lean_array_push(x_65, x_36); x_67 = lean_array_push(x_66, x_52); x_68 = lean_array_push(x_67, x_54); x_69 = lean_array_push(x_68, x_64); -x_70 = l_Lean_Elab_Command_mkSimpleDelab___closed__47; +x_70 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__41; x_71 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_71, 0, x_26); lean_ctor_set(x_71, 1, x_70); @@ -7344,7 +6164,7 @@ lean_ctor_set(x_73, 0, x_26); lean_ctor_set(x_73, 1, x_49); lean_ctor_set(x_73, 2, x_72); x_74 = lean_array_push(x_47, x_73); -x_75 = l_Lean_Elab_Command_mkSimpleDelab___closed__45; +x_75 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__39; x_76 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_76, 0, x_26); lean_ctor_set(x_76, 1, x_75); @@ -7363,93 +6183,96 @@ lean_inc(x_7); x_84 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_isLocalAttrKind(x_7); if (x_84 == 0) { -lean_object* x_85; lean_object* x_86; -x_85 = lean_box(0); -x_86 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__1(x_7, x_27, x_24, x_60, x_28, x_41, x_83, x_85, x_13, x_32); -return x_86; +lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_85 = lean_array_push(x_47, x_83); +x_86 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_86, 0, x_26); +lean_ctor_set(x_86, 1, x_49); +lean_ctor_set(x_86, 2, x_85); +x_87 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__1(x_7, x_27, x_24, x_60, x_28, x_49, x_41, x_86, x_13, x_32); +return x_87; } else { -lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_inc(x_13); -x_87 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_13, x_32); -x_88 = lean_ctor_get(x_87, 0); -lean_inc(x_88); -x_89 = lean_ctor_get(x_87, 1); +x_88 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_13, x_32); +x_89 = lean_ctor_get(x_88, 0); lean_inc(x_89); -lean_dec(x_87); -x_90 = lean_ctor_get(x_13, 2); +x_90 = lean_ctor_get(x_88, 1); lean_inc(x_90); -x_91 = lean_ctor_get(x_13, 1); +lean_dec(x_88); +x_91 = lean_ctor_get(x_13, 2); lean_inc(x_91); -x_92 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__12; -lean_inc(x_88); -x_93 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_93, 0, x_88); -lean_ctor_set(x_93, 1, x_92); -x_94 = lean_array_push(x_60, x_93); -x_95 = lean_array_push(x_94, x_79); -x_96 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__13; -x_97 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_97, 0, x_26); -lean_ctor_set(x_97, 1, x_96); -lean_ctor_set(x_97, 2, x_95); -x_98 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__14; -lean_inc(x_88); -x_99 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_99, 0, x_88); -lean_ctor_set(x_99, 1, x_98); -x_100 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__22; -x_101 = l_Lean_addMacroScope(x_91, x_100, x_90); -x_102 = lean_box(0); -x_103 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__18; -lean_inc(x_88); -x_104 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_104, 0, x_88); -lean_ctor_set(x_104, 1, x_103); -lean_ctor_set(x_104, 2, x_101); -lean_ctor_set(x_104, 3, x_102); -x_105 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__23; -lean_inc(x_88); -x_106 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_106, 0, x_88); -lean_ctor_set(x_106, 1, x_105); -x_107 = lean_array_push(x_41, x_99); -x_108 = lean_array_push(x_107, x_104); -x_109 = lean_array_push(x_108, x_106); -x_110 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__15; -x_111 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_111, 0, x_26); -lean_ctor_set(x_111, 1, x_110); -lean_ctor_set(x_111, 2, x_109); -x_112 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__24; -x_113 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_113, 0, x_88); -lean_ctor_set(x_113, 1, x_112); -x_114 = lean_array_push(x_60, x_113); -x_115 = lean_array_push(x_114, x_79); -x_116 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__25; -x_117 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_117, 0, x_26); -lean_ctor_set(x_117, 1, x_116); -lean_ctor_set(x_117, 2, x_115); -x_118 = lean_array_push(x_65, x_97); -x_119 = lean_array_push(x_118, x_111); -x_120 = lean_array_push(x_119, x_83); -x_121 = lean_array_push(x_120, x_117); -x_122 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_122, 0, x_26); -lean_ctor_set(x_122, 1, x_49); -lean_ctor_set(x_122, 2, x_121); -x_123 = lean_box(0); -x_124 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__1(x_7, x_27, x_24, x_60, x_28, x_41, x_122, x_123, x_13, x_89); +x_92 = lean_ctor_get(x_13, 1); +lean_inc(x_92); +x_93 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__12; +lean_inc(x_89); +x_94 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_94, 0, x_89); +lean_ctor_set(x_94, 1, x_93); +x_95 = lean_array_push(x_60, x_94); +x_96 = lean_array_push(x_95, x_79); +x_97 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__13; +x_98 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_98, 0, x_26); +lean_ctor_set(x_98, 1, x_97); +lean_ctor_set(x_98, 2, x_96); +x_99 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__14; +lean_inc(x_89); +x_100 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_100, 0, x_89); +lean_ctor_set(x_100, 1, x_99); +x_101 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__22; +x_102 = l_Lean_addMacroScope(x_92, x_101, x_91); +x_103 = lean_box(0); +x_104 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__18; +lean_inc(x_89); +x_105 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_105, 0, x_89); +lean_ctor_set(x_105, 1, x_104); +lean_ctor_set(x_105, 2, x_102); +lean_ctor_set(x_105, 3, x_103); +x_106 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__23; +lean_inc(x_89); +x_107 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_107, 0, x_89); +lean_ctor_set(x_107, 1, x_106); +x_108 = lean_array_push(x_41, x_100); +x_109 = lean_array_push(x_108, x_105); +x_110 = lean_array_push(x_109, x_107); +x_111 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__15; +x_112 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_112, 0, x_26); +lean_ctor_set(x_112, 1, x_111); +lean_ctor_set(x_112, 2, x_110); +x_113 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__24; +x_114 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_114, 0, x_89); +lean_ctor_set(x_114, 1, x_113); +x_115 = lean_array_push(x_60, x_114); +x_116 = lean_array_push(x_115, x_79); +x_117 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__25; +x_118 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_118, 0, x_26); +lean_ctor_set(x_118, 1, x_117); +lean_ctor_set(x_118, 2, x_116); +x_119 = lean_array_push(x_65, x_98); +x_120 = lean_array_push(x_119, x_112); +x_121 = lean_array_push(x_120, x_83); +x_122 = lean_array_push(x_121, x_118); +x_123 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_123, 0, x_26); +lean_ctor_set(x_123, 1, x_49); +lean_ctor_set(x_123, 2, x_122); +x_124 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__1(x_7, x_27, x_24, x_60, x_28, x_49, x_41, x_123, x_13, x_90); return x_124; } } } else { -uint8_t x_202; +uint8_t x_201; lean_dec(x_18); lean_dec(x_13); lean_dec(x_12); @@ -7460,23 +6283,23 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_202 = !lean_is_exclusive(x_17); -if (x_202 == 0) +x_201 = !lean_is_exclusive(x_17); +if (x_201 == 0) { return x_17; } else { -lean_object* x_203; lean_object* x_204; lean_object* x_205; -x_203 = lean_ctor_get(x_17, 0); -x_204 = lean_ctor_get(x_17, 1); -lean_inc(x_204); +lean_object* x_202; lean_object* x_203; lean_object* x_204; +x_202 = lean_ctor_get(x_17, 0); +x_203 = lean_ctor_get(x_17, 1); lean_inc(x_203); +lean_inc(x_202); lean_dec(x_17); -x_205 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_205, 0, x_203); -lean_ctor_set(x_205, 1, x_204); -return x_205; +x_204 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_204, 0, x_202); +lean_ctor_set(x_204, 1, x_203); +return x_204; } } } @@ -7499,6 +6322,7 @@ lean_dec(x_11); x_14 = lean_array_get_size(x_7); x_15 = lean_usize_of_nat(x_14); x_16 = 0; +lean_inc(x_9); lean_inc(x_7); x_17 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___spec__1(x_15, x_16, x_7, x_9, x_13); if (lean_obj_tag(x_17) == 0) @@ -7515,7 +6339,7 @@ if (lean_obj_tag(x_5) == 0) { lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; x_22 = lean_box(2); -x_23 = l_Lean_nullKind; +x_23 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__6; lean_inc(x_18); x_24 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_24, 0, x_22); @@ -7653,7 +6477,6 @@ lean_dec(x_1); x_7 = lean_unbox_usize(x_2); lean_dec(x_2); x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___spec__1(x_6, x_7, x_3, x_4, x_5); -lean_dec(x_4); return x_8; } } @@ -7682,7 +6505,19 @@ x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Notation_0__Lean_Elab_Comm return x_6; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___spec__4(x_4, x_5, x_3); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { size_t x_5; size_t x_6; lean_object* x_7; @@ -7690,20 +6525,11 @@ x_5 = lean_unbox_usize(x_2); lean_dec(x_2); x_6 = lean_unbox_usize(x_3); lean_dec(x_3); -x_7 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___spec__4(x_1, x_5, x_6, x_4); +x_7 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___spec__5(x_1, x_5, x_6, x_4); lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___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, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_8); -return x_11; -} -} LEAN_EXPORT lean_object* l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___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, lean_object* x_7, 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) { _start: { @@ -7820,14 +6646,13 @@ x_10 = l_Lean_Syntax_getArg(x_1, x_9); x_11 = l_Lean_Syntax_isNone(x_10); if (x_11 == 0) { -lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_12 = l_Lean_nullKind; -x_13 = lean_unsigned_to_nat(1u); +lean_object* x_12; uint8_t x_13; +x_12 = lean_unsigned_to_nat(1u); lean_inc(x_10); -x_14 = l_Lean_Syntax_isNodeOf(x_10, x_12, x_13); -if (x_14 == 0) +x_13 = l_Lean_Syntax_matchesNull(x_10, x_12); +if (x_13 == 0) { -lean_object* x_15; lean_object* x_16; +lean_object* x_14; lean_object* x_15; lean_dec(x_10); lean_dec(x_7); lean_dec(x_6); @@ -7835,61 +6660,61 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_15 = lean_box(1); -x_16 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_8); -return x_16; +x_14 = lean_box(1); +x_15 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_8); +return x_15; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_17 = lean_unsigned_to_nat(0u); -x_18 = l_Lean_Syntax_getArg(x_10, x_17); +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_16 = lean_unsigned_to_nat(0u); +x_17 = l_Lean_Syntax_getArg(x_10, x_16); lean_dec(x_10); -x_19 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__31; -x_20 = lean_name_mk_string(x_4, x_19); -lean_inc(x_18); -x_21 = l_Lean_Syntax_isOfKind(x_18, x_20); -lean_dec(x_20); -if (x_21 == 0) +x_18 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__31; +x_19 = lean_name_mk_string(x_4, x_18); +lean_inc(x_17); +x_20 = l_Lean_Syntax_isOfKind(x_17, x_19); +lean_dec(x_19); +if (x_20 == 0) { -lean_object* x_22; lean_object* x_23; -lean_dec(x_18); +lean_object* x_21; lean_object* x_22; +lean_dec(x_17); lean_dec(x_7); lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_22 = lean_box(1); -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_8); -return x_23; +x_21 = lean_box(1); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_8); +return x_22; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_24 = lean_unsigned_to_nat(3u); -x_25 = l_Lean_Syntax_getArg(x_18, x_24); -lean_dec(x_18); -x_26 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_26, 0, x_25); -x_27 = lean_box(0); -x_28 = l_Lean_Elab_Command_expandNotation___lambda__1(x_1, x_2, x_3, x_6, x_27, x_26, x_7, x_8); -return x_28; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_23 = lean_unsigned_to_nat(3u); +x_24 = l_Lean_Syntax_getArg(x_17, x_23); +lean_dec(x_17); +x_25 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_25, 0, x_24); +x_26 = lean_box(0); +x_27 = l_Lean_Elab_Command_expandNotation___lambda__1(x_1, x_2, x_3, x_6, x_26, x_25, x_7, x_8); +return x_27; } } } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_dec(x_10); lean_dec(x_4); +x_28 = lean_box(0); x_29 = lean_box(0); -x_30 = lean_box(0); -x_31 = l_Lean_Elab_Command_expandNotation___lambda__1(x_1, x_2, x_3, x_6, x_30, x_29, x_7, x_8); -return x_31; +x_30 = l_Lean_Elab_Command_expandNotation___lambda__1(x_1, x_2, x_3, x_6, x_29, x_28, x_7, x_8); +return x_30; } } } @@ -7903,74 +6728,73 @@ x_9 = l_Lean_Syntax_getArg(x_1, x_8); x_10 = l_Lean_Syntax_isNone(x_9); if (x_10 == 0) { -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = l_Lean_nullKind; -x_12 = lean_unsigned_to_nat(1u); +lean_object* x_11; uint8_t x_12; +x_11 = lean_unsigned_to_nat(1u); lean_inc(x_9); -x_13 = l_Lean_Syntax_isNodeOf(x_9, x_11, x_12); -if (x_13 == 0) +x_12 = l_Lean_Syntax_matchesNull(x_9, x_11); +if (x_12 == 0) { -lean_object* x_14; lean_object* x_15; +lean_object* x_13; lean_object* x_14; lean_dec(x_9); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_14 = lean_box(1); -x_15 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_15, 1, x_7); -return x_15; +x_13 = lean_box(1); +x_14 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_7); +return x_14; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_16 = lean_unsigned_to_nat(0u); -x_17 = l_Lean_Syntax_getArg(x_9, x_16); +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_15 = lean_unsigned_to_nat(0u); +x_16 = l_Lean_Syntax_getArg(x_9, x_15); lean_dec(x_9); -x_18 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__28; +x_17 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__28; lean_inc(x_3); -x_19 = lean_name_mk_string(x_3, x_18); -lean_inc(x_17); -x_20 = l_Lean_Syntax_isOfKind(x_17, x_19); -lean_dec(x_19); -if (x_20 == 0) +x_18 = lean_name_mk_string(x_3, x_17); +lean_inc(x_16); +x_19 = l_Lean_Syntax_isOfKind(x_16, x_18); +lean_dec(x_18); +if (x_19 == 0) { -lean_object* x_21; lean_object* x_22; -lean_dec(x_17); +lean_object* x_20; lean_object* x_21; +lean_dec(x_16); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_21 = lean_box(1); -x_22 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_7); -return x_22; +x_20 = lean_box(1); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_7); +return x_21; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_23 = l_Lean_Syntax_getArg(x_17, x_8); -lean_dec(x_17); -x_24 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_24, 0, x_23); -x_25 = lean_box(0); -x_26 = l_Lean_Elab_Command_expandNotation___lambda__2(x_1, x_2, x_5, x_3, x_25, x_24, x_6, x_7); -return x_26; +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_22 = l_Lean_Syntax_getArg(x_16, x_8); +lean_dec(x_16); +x_23 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_23, 0, x_22); +x_24 = lean_box(0); +x_25 = l_Lean_Elab_Command_expandNotation___lambda__2(x_1, x_2, x_5, x_3, x_24, x_23, x_6, x_7); +return x_25; } } } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_dec(x_9); +x_26 = lean_box(0); x_27 = lean_box(0); -x_28 = lean_box(0); -x_29 = l_Lean_Elab_Command_expandNotation___lambda__2(x_1, x_2, x_5, x_3, x_28, x_27, x_6, x_7); -return x_29; +x_28 = l_Lean_Elab_Command_expandNotation___lambda__2(x_1, x_2, x_5, x_3, x_27, x_26, x_6, x_7); +return x_28; } } } @@ -8038,68 +6862,67 @@ x_15 = l_Lean_Syntax_getArg(x_1, x_14); x_16 = l_Lean_Syntax_isNone(x_15); if (x_16 == 0) { -lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_17 = l_Lean_nullKind; -x_18 = lean_unsigned_to_nat(1u); +lean_object* x_17; uint8_t x_18; +x_17 = lean_unsigned_to_nat(1u); lean_inc(x_15); -x_19 = l_Lean_Syntax_isNodeOf(x_15, x_17, x_18); -if (x_19 == 0) +x_18 = l_Lean_Syntax_matchesNull(x_15, x_17); +if (x_18 == 0) { -lean_object* x_20; lean_object* x_21; +lean_object* x_19; lean_object* x_20; lean_dec(x_15); lean_dec(x_9); lean_dec(x_2); lean_dec(x_1); -x_20 = lean_box(1); -x_21 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_3); -return x_21; +x_19 = lean_box(1); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_3); +return x_20; } else { -lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_22 = l_Lean_Syntax_getArg(x_15, x_8); +lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_21 = l_Lean_Syntax_getArg(x_15, x_8); lean_dec(x_15); -x_23 = l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__39; -lean_inc(x_22); -x_24 = l_Lean_Syntax_isOfKind(x_22, x_23); -if (x_24 == 0) +x_22 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__14; +lean_inc(x_21); +x_23 = l_Lean_Syntax_isOfKind(x_21, x_22); +if (x_23 == 0) { -lean_object* x_25; lean_object* x_26; -lean_dec(x_22); +lean_object* x_24; lean_object* x_25; +lean_dec(x_21); lean_dec(x_9); lean_dec(x_2); lean_dec(x_1); -x_25 = lean_box(1); -x_26 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_3); -return x_26; +x_24 = lean_box(1); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_3); +return x_25; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_27 = l_Lean_Syntax_getArg(x_22, x_18); -lean_dec(x_22); -x_28 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_28, 0, x_27); -x_29 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__6; -x_30 = lean_box(0); -x_31 = l_Lean_Elab_Command_expandNotation___lambda__3(x_1, x_9, x_29, x_30, x_28, x_2, x_3); -return x_31; +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_26 = l_Lean_Syntax_getArg(x_21, x_17); +lean_dec(x_21); +x_27 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_27, 0, x_26); +x_28 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__6; +x_29 = lean_box(0); +x_30 = l_Lean_Elab_Command_expandNotation___lambda__3(x_1, x_9, x_28, x_29, x_27, x_2, x_3); +return x_30; } } } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_dec(x_15); -x_32 = lean_box(0); -x_33 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__6; -x_34 = lean_box(0); -x_35 = l_Lean_Elab_Command_expandNotation___lambda__3(x_1, x_9, x_33, x_34, x_32, x_2, x_3); -return x_35; +x_31 = lean_box(0); +x_32 = l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__6; +x_33 = lean_box(0); +x_34 = l_Lean_Elab_Command_expandNotation___lambda__3(x_1, x_9, x_32, x_33, x_31, x_2, x_3); +return x_34; } } } @@ -8126,7 +6949,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandNotation___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_mkSimpleDelab___closed__5; +x_1 = l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__3; x_2 = l___regBuiltin_Lean_Elab_Command_expandNotation___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -8164,7 +6987,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandNotation_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(143u); +x_1 = lean_unsigned_to_nat(138u); x_2 = lean_unsigned_to_nat(45u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8176,7 +6999,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandNotation_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(148u); +x_1 = lean_unsigned_to_nat(143u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8204,7 +7027,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandNotation_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(143u); +x_1 = lean_unsigned_to_nat(138u); x_2 = lean_unsigned_to_nat(49u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8216,7 +7039,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_expandNotation_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(143u); +x_1 = lean_unsigned_to_nat(138u); x_2 = lean_unsigned_to_nat(63u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8291,6 +7114,32 @@ l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___closed__3 = _ini lean_mark_persistent(l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___closed__3); l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___closed__4 = _init_l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___closed__4(); lean_mark_persistent(l___private_Lean_Elab_Notation_0__Lean_Elab_Command_antiquote___closed__4); +l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__1 = _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__1); +l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__2 = _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__2); +l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__3 = _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__3(); +lean_mark_persistent(l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__3); +l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__4 = _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__4(); +lean_mark_persistent(l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__4); +l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__5 = _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__5(); +lean_mark_persistent(l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__5); +l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__6 = _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__6(); +lean_mark_persistent(l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__6); +l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__7 = _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__7(); +lean_mark_persistent(l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__7); +l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__8 = _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__8(); +lean_mark_persistent(l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__8); +l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__9 = _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__9(); +lean_mark_persistent(l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__9); +l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__10 = _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__10(); +lean_mark_persistent(l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__10); +l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__11 = _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__11(); +lean_mark_persistent(l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__11); +l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__12 = _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__12(); +lean_mark_persistent(l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__12); +l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__13 = _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__13(); +lean_mark_persistent(l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___lambda__1___closed__13); l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__1 = _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__1); l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__2 = _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__2(); @@ -8319,10 +7168,6 @@ l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__13 = _init_l_Lean lean_mark_persistent(l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__13); l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__14 = _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__14(); lean_mark_persistent(l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__14); -l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__15 = _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__15(); -lean_mark_persistent(l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__15); -l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__16 = _init_l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__16(); -lean_mark_persistent(l_Lean_Elab_Command_expandNotationItemIntoSyntaxItem___closed__16); l_Lean_Elab_Command_removeParentheses___closed__1 = _init_l_Lean_Elab_Command_removeParentheses___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_removeParentheses___closed__1); l_Lean_Elab_Command_removeParentheses___closed__2 = _init_l_Lean_Elab_Command_removeParentheses___closed__2(); @@ -8335,182 +7180,178 @@ l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Command_hasDuplicateAnt lean_mark_persistent(l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Command_hasDuplicateAntiquot___spec__1___closed__1); l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Command_hasDuplicateAntiquot___spec__1___closed__2 = _init_l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Command_hasDuplicateAntiquot___spec__1___closed__2(); lean_mark_persistent(l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Command_hasDuplicateAntiquot___spec__1___closed__2); +l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Command_hasDuplicateAntiquot___spec__1___closed__3 = _init_l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Command_hasDuplicateAntiquot___spec__1___closed__3(); +lean_mark_persistent(l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Command_hasDuplicateAntiquot___spec__1___closed__3); +l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Command_hasDuplicateAntiquot___spec__1___closed__4 = _init_l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Command_hasDuplicateAntiquot___spec__1___closed__4(); +lean_mark_persistent(l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Command_hasDuplicateAntiquot___spec__1___closed__4); l_Lean_Elab_Command_hasDuplicateAntiquot___closed__1 = _init_l_Lean_Elab_Command_hasDuplicateAntiquot___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_hasDuplicateAntiquot___closed__1); l_Lean_Elab_Command_hasDuplicateAntiquot___closed__2 = _init_l_Lean_Elab_Command_hasDuplicateAntiquot___closed__2(); lean_mark_persistent(l_Lean_Elab_Command_hasDuplicateAntiquot___closed__2); l_Lean_Elab_Command_hasDuplicateAntiquot___closed__3 = _init_l_Lean_Elab_Command_hasDuplicateAntiquot___closed__3(); lean_mark_persistent(l_Lean_Elab_Command_hasDuplicateAntiquot___closed__3); -l_Lean_Elab_Command_mkSimpleDelab___closed__1 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__1(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__1); -l_Lean_Elab_Command_mkSimpleDelab___closed__2 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__2(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__2); -l_Lean_Elab_Command_mkSimpleDelab___closed__3 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__3(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__3); -l_Lean_Elab_Command_mkSimpleDelab___closed__4 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__4(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__4); -l_Lean_Elab_Command_mkSimpleDelab___closed__5 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__5(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__5); -l_Lean_Elab_Command_mkSimpleDelab___closed__6 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__6(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__6); -l_Lean_Elab_Command_mkSimpleDelab___closed__7 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__7(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__7); -l_Lean_Elab_Command_mkSimpleDelab___closed__8 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__8(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__8); -l_Lean_Elab_Command_mkSimpleDelab___closed__9 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__9(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__9); -l_Lean_Elab_Command_mkSimpleDelab___closed__10 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__10(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__10); -l_Lean_Elab_Command_mkSimpleDelab___closed__11 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__11(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__11); -l_Lean_Elab_Command_mkSimpleDelab___closed__12 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__12(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__12); -l_Lean_Elab_Command_mkSimpleDelab___closed__13 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__13(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__13); -l_Lean_Elab_Command_mkSimpleDelab___closed__14 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__14(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__14); -l_Lean_Elab_Command_mkSimpleDelab___closed__15 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__15(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__15); -l_Lean_Elab_Command_mkSimpleDelab___closed__16 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__16(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__16); -l_Lean_Elab_Command_mkSimpleDelab___closed__17 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__17(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__17); -l_Lean_Elab_Command_mkSimpleDelab___closed__18 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__18(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__18); -l_Lean_Elab_Command_mkSimpleDelab___closed__19 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__19(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__19); -l_Lean_Elab_Command_mkSimpleDelab___closed__20 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__20(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__20); -l_Lean_Elab_Command_mkSimpleDelab___closed__21 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__21(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__21); -l_Lean_Elab_Command_mkSimpleDelab___closed__22 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__22(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__22); -l_Lean_Elab_Command_mkSimpleDelab___closed__23 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__23(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__23); -l_Lean_Elab_Command_mkSimpleDelab___closed__24 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__24(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__24); -l_Lean_Elab_Command_mkSimpleDelab___closed__25 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__25(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__25); -l_Lean_Elab_Command_mkSimpleDelab___closed__26 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__26(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__26); -l_Lean_Elab_Command_mkSimpleDelab___closed__27 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__27(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__27); -l_Lean_Elab_Command_mkSimpleDelab___closed__28 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__28(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__28); -l_Lean_Elab_Command_mkSimpleDelab___closed__29 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__29(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__29); -l_Lean_Elab_Command_mkSimpleDelab___closed__30 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__30(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__30); -l_Lean_Elab_Command_mkSimpleDelab___closed__31 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__31(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__31); -l_Lean_Elab_Command_mkSimpleDelab___closed__32 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__32(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__32); -l_Lean_Elab_Command_mkSimpleDelab___closed__33 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__33(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__33); -l_Lean_Elab_Command_mkSimpleDelab___closed__34 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__34(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__34); -l_Lean_Elab_Command_mkSimpleDelab___closed__35 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__35(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__35); -l_Lean_Elab_Command_mkSimpleDelab___closed__36 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__36(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__36); -l_Lean_Elab_Command_mkSimpleDelab___closed__37 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__37(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__37); -l_Lean_Elab_Command_mkSimpleDelab___closed__38 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__38(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__38); -l_Lean_Elab_Command_mkSimpleDelab___closed__39 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__39(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__39); -l_Lean_Elab_Command_mkSimpleDelab___closed__40 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__40(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__40); -l_Lean_Elab_Command_mkSimpleDelab___closed__41 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__41(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__41); -l_Lean_Elab_Command_mkSimpleDelab___closed__42 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__42(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__42); -l_Lean_Elab_Command_mkSimpleDelab___closed__43 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__43(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__43); -l_Lean_Elab_Command_mkSimpleDelab___closed__44 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__44(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__44); -l_Lean_Elab_Command_mkSimpleDelab___closed__45 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__45(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__45); -l_Lean_Elab_Command_mkSimpleDelab___closed__46 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__46(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__46); -l_Lean_Elab_Command_mkSimpleDelab___closed__47 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__47(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__47); -l_Lean_Elab_Command_mkSimpleDelab___closed__48 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__48(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__48); -l_Lean_Elab_Command_mkSimpleDelab___closed__49 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__49(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__49); -l_Lean_Elab_Command_mkSimpleDelab___closed__50 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__50(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__50); -l_Lean_Elab_Command_mkSimpleDelab___closed__51 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__51(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__51); -l_Lean_Elab_Command_mkSimpleDelab___closed__52 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__52(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__52); -l_Lean_Elab_Command_mkSimpleDelab___closed__53 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__53(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__53); -l_Lean_Elab_Command_mkSimpleDelab___closed__54 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__54(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__54); -l_Lean_Elab_Command_mkSimpleDelab___closed__55 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__55(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__55); -l_Lean_Elab_Command_mkSimpleDelab___closed__56 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__56(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__56); -l_Lean_Elab_Command_mkSimpleDelab___closed__57 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__57(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__57); -l_Lean_Elab_Command_mkSimpleDelab___closed__58 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__58(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__58); -l_Lean_Elab_Command_mkSimpleDelab___closed__59 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__59(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__59); -l_Lean_Elab_Command_mkSimpleDelab___closed__60 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__60(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__60); -l_Lean_Elab_Command_mkSimpleDelab___closed__61 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__61(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__61); -l_Lean_Elab_Command_mkSimpleDelab___closed__62 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__62(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__62); -l_Lean_Elab_Command_mkSimpleDelab___closed__63 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__63(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__63); -l_Lean_Elab_Command_mkSimpleDelab___closed__64 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__64(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__64); -l_Lean_Elab_Command_mkSimpleDelab___closed__65 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__65(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__65); -l_Lean_Elab_Command_mkSimpleDelab___closed__66 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__66(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__66); -l_Lean_Elab_Command_mkSimpleDelab___closed__67 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__67(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__67); -l_Lean_Elab_Command_mkSimpleDelab___closed__68 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__68(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__68); -l_Lean_Elab_Command_mkSimpleDelab___closed__69 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__69(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__69); -l_Lean_Elab_Command_mkSimpleDelab___closed__70 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__70(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__70); -l_Lean_Elab_Command_mkSimpleDelab___closed__71 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__71(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__71); -l_Lean_Elab_Command_mkSimpleDelab___closed__72 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__72(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__72); -l_Lean_Elab_Command_mkSimpleDelab___closed__73 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__73(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__73); -l_Lean_Elab_Command_mkSimpleDelab___closed__74 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__74(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__74); -l_Lean_Elab_Command_mkSimpleDelab___closed__75 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__75(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__75); -l_Lean_Elab_Command_mkSimpleDelab___closed__76 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__76(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__76); -l_Lean_Elab_Command_mkSimpleDelab___closed__77 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__77(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__77); -l_Lean_Elab_Command_mkSimpleDelab___closed__78 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__78(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__78); -l_Lean_Elab_Command_mkSimpleDelab___closed__79 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__79(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__79); -l_Lean_Elab_Command_mkSimpleDelab___closed__80 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__80(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__80); -l_Lean_Elab_Command_mkSimpleDelab___closed__81 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__81(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__81); -l_Lean_Elab_Command_mkSimpleDelab___closed__82 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__82(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__82); -l_Lean_Elab_Command_mkSimpleDelab___closed__83 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__83(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__83); -l_Lean_Elab_Command_mkSimpleDelab___closed__84 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__84(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__84); -l_Lean_Elab_Command_mkSimpleDelab___closed__85 = _init_l_Lean_Elab_Command_mkSimpleDelab___closed__85(); -lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___closed__85); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__1 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__1); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__2 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__2); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__3 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__3(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__3); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__4 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__4(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__4); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__5 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__5(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__5); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__6 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__6(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__6); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__7 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__7(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__7); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__8 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__8(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__8); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__9 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__9(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__9); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__10 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__10(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__10); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__11 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__11(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__11); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__12 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__12(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__12); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__13 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__13(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__13); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__14 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__14(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__14); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__15 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__15(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__15); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__16 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__16(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__16); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__17 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__17(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__17); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__18 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__18(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__18); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__19 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__19(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__19); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__20 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__20(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__20); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__21 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__21(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__21); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__22 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__22(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__22); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__23 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__23(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__23); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__24 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__24(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__24); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__25 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__25(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__25); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__26 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__26(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__26); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__27 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__27(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__27); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__28 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__28(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__28); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__29 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__29(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__29); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__30 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__30(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__30); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__31 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__31(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__31); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__32 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__32(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__32); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__33 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__33(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__33); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__34 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__34(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__34); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__35 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__35(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__35); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__36 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__36(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__36); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__37 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__37(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__37); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__38 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__38(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__38); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__39 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__39(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__39); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__40 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__40(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__40); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__41 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__41(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__41); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__42 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__42(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__42); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__43 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__43(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__43); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__44 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__44(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__44); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__45 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__45(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__45); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__46 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__46(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__46); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__47 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__47(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__47); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__48 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__48(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__48); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__49 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__49(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__49); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__50 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__50(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__50); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__51 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__51(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__51); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__52 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__52(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__52); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__53 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__53(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__53); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__54 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__54(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__54); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__55 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__55(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__55); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__56 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__56(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__56); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__57 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__57(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__57); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__58 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__58(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__58); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__59 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__59(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__59); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__60 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__60(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__60); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__61 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__61(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__61); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__62 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__62(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__62); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__63 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__63(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__63); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__64 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__64(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__64); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__65 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__65(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__65); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__66 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__66(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__66); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__67 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__67(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__67); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__68 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__68(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__68); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__69 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__69(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__69); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__70 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__70(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__70); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__71 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__71(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__71); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__72 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__72(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__72); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__73 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__73(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__73); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__74 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__74(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__74); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__75 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__75(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__75); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__76 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__76(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__76); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__77 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__77(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__77); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__78 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__78(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__78); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__79 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__79(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__79); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__80 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__80(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__80); +l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__81 = _init_l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__81(); +lean_mark_persistent(l_Lean_Elab_Command_mkSimpleDelab___lambda__1___closed__81); l___private_Lean_Elab_Notation_0__Lean_Elab_Command_isLocalAttrKind___closed__1 = _init_l___private_Lean_Elab_Notation_0__Lean_Elab_Command_isLocalAttrKind___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Notation_0__Lean_Elab_Command_isLocalAttrKind___closed__1); l___private_Lean_Elab_Notation_0__Lean_Elab_Command_isLocalAttrKind___closed__2 = _init_l___private_Lean_Elab_Notation_0__Lean_Elab_Command_isLocalAttrKind___closed__2(); @@ -8589,14 +7430,6 @@ l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__ lean_mark_persistent(l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__34); l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__35 = _init_l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__35(); lean_mark_persistent(l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__35); -l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__36 = _init_l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__36(); -lean_mark_persistent(l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__36); -l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__37 = _init_l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__37(); -lean_mark_persistent(l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__37); -l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__38 = _init_l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__38(); -lean_mark_persistent(l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__38); -l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__39 = _init_l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__39(); -lean_mark_persistent(l___private_Lean_Elab_Notation_0__Lean_Elab_Command_expandNotationAux___lambda__2___closed__39); l_Lean_Elab_Command_expandNotation___closed__1 = _init_l_Lean_Elab_Command_expandNotation___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_expandNotation___closed__1); l_Lean_Elab_Command_expandNotation___closed__2 = _init_l_Lean_Elab_Command_expandNotation___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Open.c b/stage0/stdlib/Lean/Elab/Open.c index 49d7f00ee5e..2b057407f15 100644 --- a/stage0/stdlib/Lean/Elab/Open.c +++ b/stage0/stdlib/Lean/Elab/Open.c @@ -32,6 +32,7 @@ LEAN_EXPORT lean_object* l_Lean_resolveUniqueNamespace___at___private_Lean_Elab_ lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Lean_resolveNamespace___at___private_Lean_Elab_Open_0__Lean_Elab_OpenDecl_elabOpenRenaming___spec__2___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_OpenDecl_resolveId___spec__4___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_OpenDecl_resolveId___rarg___lambda__1___boxed(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_Array_forInUnsafe_loop___at___private_Lean_Elab_Open_0__Lean_Elab_OpenDecl_elabOpenRenaming___spec__6(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Open_0__Lean_Elab_OpenDecl_resolveNameUsingNamespacesCore___rarg___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -252,7 +253,6 @@ LEAN_EXPORT lean_object* l_Lean_resolveNamespace___at___private_Lean_Elab_Open_0 LEAN_EXPORT lean_object* l___private_Lean_Elab_Open_0__Lean_Elab_OpenDecl_elabOpenScoped___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_OpenDecl_resolveNameUsingNamespaces___rarg(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_Array_forInUnsafe_loop___at___private_Lean_Elab_Open_0__Lean_Elab_OpenDecl_elabOpenScoped___spec__6___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); -lean_object* l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Open_0__Lean_Elab_OpenDecl_elabOpenOnly___spec__2(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Open_0__Lean_Elab_OpenDecl_elabOpenScoped___spec__6___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getPos_x3f(lean_object*, uint8_t); @@ -809,7 +809,7 @@ LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at_Lean_Elab_OpenDecl_r { lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; x_9 = lean_box(0); -x_10 = l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(x_8, x_9); +x_10 = l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(x_8, x_9); lean_inc(x_1); lean_inc(x_10); x_11 = lean_alloc_closure((void*)(l_Lean_resolveGlobalConstCore___at_Lean_Elab_OpenDecl_resolveId___spec__2___rarg___lambda__1___boxed), 5, 3); diff --git a/stage0/stdlib/Lean/Elab/PatternVar.c b/stage0/stdlib/Lean/Elab/PatternVar.c index 9958d710942..e7a0a7a0be0 100644 --- a/stage0/stdlib/Lean/Elab/PatternVar.c +++ b/stage0/stdlib/Lean/Elab/PatternVar.c @@ -17,7 +17,7 @@ lean_object* l_List_reverse___rarg(lean_object*); lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_collectPatternVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_resolveGlobalConstNoOverloadWithInfo___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_doubleQuotedNameToPattern___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___spec__2___boxed(lean_object*, 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_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___spec__2___boxed(lean_object*, lean_object*, 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_Term_CollectPatternVars_collect___lambda__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___closed__8; size_t lean_usize_add(size_t, size_t); @@ -41,8 +41,8 @@ static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___closed__4; static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect_pushNewArg___closed__2; lean_object* l_Lean_LocalDecl_userName(lean_object*); static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___closed__17; -extern lean_object* l_Lean_nullKind; static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__8___closed__8; +lean_object* l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__25; LEAN_EXPORT lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_processVar___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_name_mk_string(lean_object*, lean_object*); @@ -100,11 +100,13 @@ uint8_t lean_name_eq(lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_Elab_Term_mkAuxName___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwCtorExpected___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_getPatternsVars___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___closed__23; static lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_processVar___lambda__2___closed__2; static lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__2; static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_getPatternsVars___spec__6___rarg___closed__1; static lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_processVar___closed__2; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_main___spec__3___closed__6; +static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___closed__21; static lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__30; static lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_getPatternsVars___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -150,11 +152,10 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_Coll static lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__9; LEAN_EXPORT lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_samePatternsVariables___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__3; -lean_object* l_Lean_Syntax_SepArray_ofElems(lean_object*, lean_object*); static lean_object* l_Lean_resolveGlobalConst___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_doubleQuotedNameToPattern___spec__3___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__12(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_addTrace___at_Lean_Elab_Term_getPatternsVars___spec__2(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_Term_CollectPatternVars_collect_processCtorApp___lambda__2(lean_object*, lean_object*, uint8_t, lean_object*, 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_Term_CollectPatternVars_collect_processCtorApp___lambda__2(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__8; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_main___spec__3___closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_CollectPatternVars_collect_processCtorAppCore___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -174,7 +175,6 @@ static lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPa LEAN_EXPORT uint8_t l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_isDone(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect_pushNewArg(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___closed__12; -lean_object* l_Lean_Syntax_SepArray_getElems___rarg(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_main___spec__3___closed__4; lean_object* l_Lean_Syntax_mkApp(lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_levelParams(lean_object*); @@ -193,23 +193,23 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__1 LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOptionM___at_Lean_Elab_Term_CollectPatternVars_main___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect_processExplicitArg___closed__1; +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); lean_object* l_Lean_Elab_expandMacroImpl_x3f(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); -extern lean_object* l_Lean_numLitKind; LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_Context_newArgs___default; static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__5; static lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__17; static lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_getPatternVarNames(lean_object*); -extern lean_object* l_Lean_strLitKind; +static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3___closed__3; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Array_mapSepElemsMAux___at_Lean_Elab_Term_CollectPatternVars_collect___spec__6___boxed(lean_object*, lean_object*, 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_Term_CollectPatternVars_collect_processCtorAppCore___boxed(lean_object*, lean_object*, 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_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect_processCtorAppCore___lambda__1(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_Term_CollectPatternVars_State_found___default; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___spec__2(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_main___spec__3___closed__3; static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___closed__1; @@ -223,7 +223,7 @@ lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__8___closed__1; lean_object* l_Lean_ResolveName_resolveNamespace(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__1; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Term_CollectPatternVars_collect___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__34; LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect_processImplicitArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___closed__2; @@ -256,9 +256,7 @@ lean_object* l_Lean_Syntax_getId(lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Term_CollectPatternVars_main___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_format_pretty(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_doubleQuotedNameToPattern___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_choiceKind; static lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__6; -extern lean_object* l_Lean_charLitKind; static lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_getNextParam___closed__1; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_main___spec__3___closed__5; static lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_processVar___lambda__2___closed__3; @@ -274,6 +272,7 @@ static lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPa LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_doubleQuotedNameToPattern___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*, lean_object*); uint8_t l_Lean_Name_isAtomic(lean_object*); +static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___closed__20; uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_CollectPatternVars_collect___spec__4(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___closed__10; @@ -289,12 +288,14 @@ lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_ uint8_t l_Array_isEmpty___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_doubleQuotedNameToPattern___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_instInhabitedSyntax; +lean_object* l_Lean_mkSepArray(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__8___closed__7; lean_object* l_Lean_ResolveName_resolveGlobalName(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__12; static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___closed__16; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Array_mapSepElemsMAux___at_Lean_Elab_Term_CollectPatternVars_collect___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___closed__24; static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___lambda__2___closed__2; size_t lean_usize_of_nat(lean_object*); extern lean_object* l_Lean_NameSet_empty; @@ -310,13 +311,12 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect(lean_object static lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__26; LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect_processCtorAppContext(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkAtomFrom(lean_object*, lean_object*); -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_mkNumLit(lean_object*, lean_object*); +lean_object* l_Lean_Syntax_TSepArray_getElems___rarg(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___spec__2___closed__3; LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*); -lean_object* l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); extern uint8_t l_Lean_instInhabitedBinderInfo; static lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__11; -extern lean_object* l_Lean_identKind; LEAN_EXPORT lean_object* l_Lean_Elab_Term_getPatternVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_getPatternsVars___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__9___closed__2; @@ -331,7 +331,6 @@ static lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPa extern lean_object* l_Lean_Elab_Term_instMonadTermElabM; static lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__32; LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect_processExplicitArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_scientificLitKind; static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__8___closed__3; static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__9___closed__1; lean_object* l_Lean_Syntax_setArg(lean_object*, lean_object*, lean_object*); @@ -362,6 +361,7 @@ static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect_processExplicitA uint8_t l_Lean_TagAttribute_hasTag(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_processVar___lambda__3___closed__2; lean_object* l_ReaderT_pure___at_Lean_Elab_liftMacroM___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__1___closed__1; static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__11___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect_processImplicitArg(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_getPatternsVars___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -370,9 +370,10 @@ lean_object* l_List_mapTRAux___at_Lean_mkConstWithLevelParams___spec__1(lean_obj static lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__36; LEAN_EXPORT lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_processVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwCtorExpected(lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect___spec__2___closed__3; +static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__1___closed__2; static lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__24; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_getPatternsVars___spec__1___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___closed__22; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___lambda__2___closed__1; LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_Term_CollectPatternVars_collect_pushNewArg___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -383,7 +384,7 @@ LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_Term_CollectPatternVars_collect_ static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect_processExplicitArg___closed__3; extern lean_object* l_Lean_instInhabitedName; LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_Context_paramDeclIdx___default; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Term_CollectPatternVars_collect___lambda__2___boxed(lean_object**); uint8_t l_Lean_Syntax_isNone(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_CollectPatternVars_collect_processExplicitArg___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__1; @@ -399,6 +400,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_Coll LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstNoOverload___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_doubleQuotedNameToPattern___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_CollectPatternVars_collect_processCtorAppCore___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___closed__25; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_getPatternVarNames___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_doubleQuotedNameToPattern___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Term_getPatternsVars___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -444,7 +446,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_Coll lean_object* l_List_toString___at_Lean_resolveGlobalConstNoOverloadCore___spec__2(lean_object*); static lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__22; static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___closed__3; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect___spec__2(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___closed__18; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_getPatternsVars___spec__1___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_processVar___lambda__2___closed__4; @@ -468,14 +470,14 @@ static lean_object* l_Lean_throwUnknownConstant___at___private_Lean_Elab_Pattern LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at_Lean_Elab_Term_CollectPatternVars_collect_processCtorAppCore___spec__4(lean_object*); static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__12___closed__1; static lean_object* l_Lean_resolveGlobalConstNoOverload___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_doubleQuotedNameToPattern___spec__2___closed__2; +static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__9; static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___closed__2; static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_getPatternsVars___spec__6___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect_pushNewArg___lambda__1___boxed(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___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwCtorExpected___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect___spec__2___boxed(lean_object*, lean_object*, 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_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___closed__1; uint8_t lean_string_dec_eq(lean_object*, lean_object*); -lean_object* l_Lean_Syntax_mkLit(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_instInhabitedContext; LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect_processCtorAppContext___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -2307,7 +2309,7 @@ x_109 = lean_st_ref_get(x_7, x_107); x_110 = !lean_is_exclusive(x_109); if (x_110 == 0) { -lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; x_111 = lean_ctor_get(x_109, 0); x_112 = lean_ctor_get(x_111, 0); lean_inc(x_112); @@ -2324,98 +2326,96 @@ lean_ctor_set(x_118, 1, x_116); lean_ctor_set(x_118, 2, x_115); lean_ctor_set(x_118, 3, x_117); x_119 = l_Nat_repr(x_101); -x_120 = l_Lean_numLitKind; -x_121 = lean_box(2); -x_122 = l_Lean_Syntax_mkLit(x_120, x_119, x_121); -x_123 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__26; -x_124 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_124, 0, x_106); -lean_ctor_set(x_124, 1, x_123); -x_125 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__27; -x_126 = lean_array_push(x_125, x_124); -x_127 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__25; -x_128 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_128, 0, x_121); -lean_ctor_set(x_128, 1, x_127); -lean_ctor_set(x_128, 2, x_126); -x_129 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__28; -x_130 = lean_array_push(x_129, x_103); -x_131 = lean_array_push(x_130, x_122); -x_132 = lean_array_push(x_131, x_128); -x_133 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__23; -x_134 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_134, 0, x_121); -lean_ctor_set(x_134, 1, x_133); -lean_ctor_set(x_134, 2, x_132); -x_135 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__12; -x_136 = lean_array_push(x_135, x_118); -x_137 = lean_array_push(x_136, x_134); -x_138 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__13; -x_139 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_139, 0, x_121); -lean_ctor_set(x_139, 1, x_138); -lean_ctor_set(x_139, 2, x_137); -lean_ctor_set(x_109, 0, x_139); +x_120 = lean_box(2); +x_121 = l_Lean_Syntax_mkNumLit(x_119, x_120); +x_122 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__26; +x_123 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_123, 0, x_106); +lean_ctor_set(x_123, 1, x_122); +x_124 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__27; +x_125 = lean_array_push(x_124, x_123); +x_126 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__25; +x_127 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_127, 0, x_120); +lean_ctor_set(x_127, 1, x_126); +lean_ctor_set(x_127, 2, x_125); +x_128 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__28; +x_129 = lean_array_push(x_128, x_103); +x_130 = lean_array_push(x_129, x_121); +x_131 = lean_array_push(x_130, x_127); +x_132 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__23; +x_133 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_133, 0, x_120); +lean_ctor_set(x_133, 1, x_132); +lean_ctor_set(x_133, 2, x_131); +x_134 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__12; +x_135 = lean_array_push(x_134, x_118); +x_136 = lean_array_push(x_135, x_133); +x_137 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__13; +x_138 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_138, 0, x_120); +lean_ctor_set(x_138, 1, x_137); +lean_ctor_set(x_138, 2, x_136); +lean_ctor_set(x_109, 0, x_138); return x_109; } else { -lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; -x_140 = lean_ctor_get(x_109, 0); -x_141 = lean_ctor_get(x_109, 1); -lean_inc(x_141); +lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; +x_139 = lean_ctor_get(x_109, 0); +x_140 = lean_ctor_get(x_109, 1); lean_inc(x_140); +lean_inc(x_139); lean_dec(x_109); -x_142 = lean_ctor_get(x_140, 0); -lean_inc(x_142); -lean_dec(x_140); -x_143 = lean_environment_main_module(x_142); -x_144 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__33; -x_145 = l_Lean_addMacroScope(x_143, x_144, x_108); -x_146 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__31; -x_147 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__36; +x_141 = lean_ctor_get(x_139, 0); +lean_inc(x_141); +lean_dec(x_139); +x_142 = lean_environment_main_module(x_141); +x_143 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__33; +x_144 = l_Lean_addMacroScope(x_142, x_143, x_108); +x_145 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__31; +x_146 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__36; lean_inc(x_106); -x_148 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_148, 0, x_106); -lean_ctor_set(x_148, 1, x_146); -lean_ctor_set(x_148, 2, x_145); -lean_ctor_set(x_148, 3, x_147); -x_149 = l_Nat_repr(x_101); -x_150 = l_Lean_numLitKind; -x_151 = lean_box(2); -x_152 = l_Lean_Syntax_mkLit(x_150, x_149, x_151); -x_153 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__26; -x_154 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_154, 0, x_106); -lean_ctor_set(x_154, 1, x_153); -x_155 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__27; -x_156 = lean_array_push(x_155, x_154); -x_157 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__25; -x_158 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_158, 0, x_151); -lean_ctor_set(x_158, 1, x_157); -lean_ctor_set(x_158, 2, x_156); -x_159 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__28; -x_160 = lean_array_push(x_159, x_103); -x_161 = lean_array_push(x_160, x_152); -x_162 = lean_array_push(x_161, x_158); -x_163 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__23; -x_164 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_164, 0, x_151); -lean_ctor_set(x_164, 1, x_163); -lean_ctor_set(x_164, 2, x_162); -x_165 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__12; -x_166 = lean_array_push(x_165, x_148); -x_167 = lean_array_push(x_166, x_164); -x_168 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__13; -x_169 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_169, 0, x_151); -lean_ctor_set(x_169, 1, x_168); -lean_ctor_set(x_169, 2, x_167); -x_170 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_170, 0, x_169); -lean_ctor_set(x_170, 1, x_141); -return x_170; +x_147 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_147, 0, x_106); +lean_ctor_set(x_147, 1, x_145); +lean_ctor_set(x_147, 2, x_144); +lean_ctor_set(x_147, 3, x_146); +x_148 = l_Nat_repr(x_101); +x_149 = lean_box(2); +x_150 = l_Lean_Syntax_mkNumLit(x_148, x_149); +x_151 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__26; +x_152 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_152, 0, x_106); +lean_ctor_set(x_152, 1, x_151); +x_153 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__27; +x_154 = lean_array_push(x_153, x_152); +x_155 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__25; +x_156 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_156, 0, x_149); +lean_ctor_set(x_156, 1, x_155); +lean_ctor_set(x_156, 2, x_154); +x_157 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__28; +x_158 = lean_array_push(x_157, x_103); +x_159 = lean_array_push(x_158, x_150); +x_160 = lean_array_push(x_159, x_156); +x_161 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__23; +x_162 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_162, 0, x_149); +lean_ctor_set(x_162, 1, x_161); +lean_ctor_set(x_162, 2, x_160); +x_163 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__12; +x_164 = lean_array_push(x_163, x_147); +x_165 = lean_array_push(x_164, x_162); +x_166 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__13; +x_167 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_167, 0, x_149); +lean_ctor_set(x_167, 1, x_166); +lean_ctor_set(x_167, 2, x_165); +x_168 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_168, 0, x_167); +lean_ctor_set(x_168, 1, x_140); +return x_168; } } } @@ -2503,7 +2503,7 @@ return x_12; } else { -lean_object* x_13; lean_object* x_14; +lean_object* x_13; lean_object* x_14; uint8_t x_15; x_13 = lean_ctor_get(x_11, 0); lean_inc(x_13); lean_dec(x_11); @@ -2513,8 +2513,25 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) +{ return x_14; } +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_14, 0); +x_17 = lean_ctor_get(x_14, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_14); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +return x_18; +} +} } } LEAN_EXPORT lean_object* l_Lean_Elab_throwIllFormedSyntax___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_quotedNameToPattern___spec__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) { @@ -2670,7 +2687,7 @@ x_11 = lean_ctor_get(x_9, 1); lean_inc(x_11); lean_dec(x_9); x_12 = lean_box(0); -x_13 = l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(x_10, x_12); +x_13 = l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(x_10, x_12); x_14 = l_List_isEmpty___rarg(x_13); if (x_14 == 0) { @@ -3382,7 +3399,7 @@ lean_inc(x_2); x_12 = l_Lean_Elab_resolveGlobalConstNoOverloadWithInfo___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_doubleQuotedNameToPattern___spec__1(x_10, x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8); if (lean_obj_tag(x_12) == 0) { -lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); x_14 = lean_ctor_get(x_12, 1); @@ -3394,34 +3411,51 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ return x_15; } else { -uint8_t x_16; +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_15, 0); +x_18 = lean_ctor_get(x_15, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_15); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +return x_19; +} +} +else +{ +uint8_t x_20; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_16 = !lean_is_exclusive(x_12); -if (x_16 == 0) +x_20 = !lean_is_exclusive(x_12); +if (x_20 == 0) { return x_12; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_12, 0); -x_18 = lean_ctor_get(x_12, 1); -lean_inc(x_18); -lean_inc(x_17); +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_12, 0); +x_22 = lean_ctor_get(x_12, 1); +lean_inc(x_22); +lean_inc(x_21); lean_dec(x_12); -x_19 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set(x_19, 1, x_18); -return x_19; +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; } } } @@ -3629,7 +3663,7 @@ static lean_object* _init_l_panic___at_Lean_Elab_Term_CollectPatternVars_collect { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_panic___at_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___spec__1___closed__1; -x_2 = l_Lean_instInhabitedSyntax; +x_2 = lean_box(0); x_3 = l_instInhabited___rarg(x_1, x_2); return x_3; } @@ -3674,157 +3708,157 @@ static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatt 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_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___spec__2___closed__1; x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___spec__2___closed__2; -x_3 = lean_unsigned_to_nat(242u); +x_3 = lean_unsigned_to_nat(243u); x_4 = lean_unsigned_to_nat(74u); x_5 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___spec__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_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___spec__2(size_t x_1, size_t 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, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___spec__2(lean_object* x_1, size_t x_2, size_t 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, lean_object* x_11, lean_object* x_12) { _start: { -uint8_t x_12; -x_12 = lean_usize_dec_lt(x_2, x_1); -if (x_12 == 0) +uint8_t x_13; +x_13 = lean_usize_dec_lt(x_3, x_2); +if (x_13 == 0) { -lean_object* x_13; +lean_object* x_14; +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_3); -lean_ctor_set(x_13, 1, x_11); -return x_13; +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_4); +lean_ctor_set(x_14, 1, x_12); +return x_14; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_array_uget(x_3, x_2); -x_15 = lean_unsigned_to_nat(0u); -x_16 = lean_array_uset(x_3, x_2, x_15); -if (lean_obj_tag(x_14) == 0) +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_array_uget(x_4, x_3); +x_16 = lean_unsigned_to_nat(0u); +x_17 = lean_array_uset(x_4, x_3, x_16); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_17; lean_object* x_18; -x_17 = lean_ctor_get(x_14, 0); -lean_inc(x_17); -lean_dec(x_14); +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_15, 0); +lean_inc(x_18); +lean_dec(x_15); +lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_4); -x_18 = l_Lean_Elab_Term_CollectPatternVars_collect(x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_18) == 0) +x_19 = l_Lean_Elab_Term_CollectPatternVars_collect(x_18, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_19) == 0) { -lean_object* x_19; lean_object* x_20; size_t x_21; size_t x_22; lean_object* x_23; -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); +lean_object* x_20; lean_object* x_21; size_t x_22; size_t x_23; lean_object* x_24; +x_20 = lean_ctor_get(x_19, 0); lean_inc(x_20); -lean_dec(x_18); -x_21 = 1; -x_22 = lean_usize_add(x_2, x_21); -x_23 = lean_array_uset(x_16, x_2, x_19); -x_2 = x_22; +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = 1; +x_23 = lean_usize_add(x_3, x_22); +x_24 = lean_array_uset(x_17, x_3, x_20); x_3 = x_23; -x_11 = x_20; +x_4 = x_24; +x_12 = x_21; goto _start; } else { -uint8_t x_25; -lean_dec(x_16); +uint8_t x_26; +lean_dec(x_17); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -x_25 = !lean_is_exclusive(x_18); -if (x_25 == 0) +x_26 = !lean_is_exclusive(x_19); +if (x_26 == 0) { -return x_18; +return x_19; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_18, 0); -x_27 = lean_ctor_get(x_18, 1); +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_19, 0); +x_28 = lean_ctor_get(x_19, 1); +lean_inc(x_28); lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_18); -x_28 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -return x_28; +lean_dec(x_19); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; } } } else { -lean_object* x_29; lean_object* x_30; -lean_dec(x_14); -x_29 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___spec__2___closed__4; +lean_object* x_30; lean_object* x_31; +lean_dec(x_15); +x_30 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___spec__2___closed__4; +lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_4); -x_30 = l_panic___at_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___spec__1(x_29, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_30) == 0) +x_31 = l_panic___at_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___spec__1(x_30, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_31) == 0) { -lean_object* x_31; lean_object* x_32; size_t x_33; size_t x_34; lean_object* x_35; -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_30, 1); +lean_object* x_32; lean_object* x_33; size_t x_34; size_t x_35; lean_object* x_36; +x_32 = lean_ctor_get(x_31, 0); lean_inc(x_32); -lean_dec(x_30); -x_33 = 1; -x_34 = lean_usize_add(x_2, x_33); -x_35 = lean_array_uset(x_16, x_2, x_31); -x_2 = x_34; +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec(x_31); +x_34 = 1; +x_35 = lean_usize_add(x_3, x_34); +x_36 = lean_array_uset(x_17, x_3, x_32); x_3 = x_35; -x_11 = x_32; +x_4 = x_36; +x_12 = x_33; goto _start; } else { -uint8_t x_37; -lean_dec(x_16); +uint8_t x_38; +lean_dec(x_17); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -x_37 = !lean_is_exclusive(x_30); -if (x_37 == 0) +x_38 = !lean_is_exclusive(x_31); +if (x_38 == 0) { -return x_30; +return x_31; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_30, 0); -x_39 = lean_ctor_get(x_30, 1); +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_31, 0); +x_40 = lean_ctor_get(x_31, 1); +lean_inc(x_40); lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_30); -x_40 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -return x_40; +lean_dec(x_31); +x_41 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +return x_41; } } } @@ -3858,7 +3892,7 @@ x_1 = lean_mk_string_from_bytes("..", 2); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___lambda__2(lean_object* x_1, lean_object* x_2, uint8_t 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, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t 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, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { lean_object* x_15; size_t x_16; size_t x_17; lean_object* x_18; @@ -3874,21 +3908,21 @@ lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -x_18 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___spec__2(x_16, x_17, x_1, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +x_18 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___spec__2(x_2, x_16, x_17, x_1, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); if (lean_obj_tag(x_18) == 0) { -if (x_3 == 0) +if (x_4 == 0) { lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_dec(x_5); -lean_dec(x_4); +lean_dec(x_2); x_19 = lean_ctor_get(x_18, 0); lean_inc(x_19); x_20 = lean_ctor_get(x_18, 1); lean_inc(x_20); lean_dec(x_18); x_21 = lean_box(0); -x_22 = l_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___lambda__1(x_2, x_19, x_21, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_20); +x_22 = l_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___lambda__1(x_3, x_19, x_21, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_20); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -3907,7 +3941,7 @@ x_24 = lean_ctor_get(x_18, 1); lean_inc(x_24); lean_dec(x_18); x_25 = l_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___lambda__2___closed__1; -x_26 = lean_name_mk_string(x_4, x_25); +x_26 = lean_name_mk_string(x_2, x_25); x_27 = l_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___lambda__2___closed__2; x_28 = l_Lean_mkAtomFrom(x_5, x_27); x_29 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__27; @@ -3919,7 +3953,7 @@ lean_ctor_set(x_32, 1, x_26); lean_ctor_set(x_32, 2, x_30); x_33 = lean_array_push(x_23, x_32); x_34 = lean_box(0); -x_35 = l_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___lambda__1(x_2, x_33, x_34, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_24); +x_35 = l_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___lambda__1(x_3, x_33, x_34, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_24); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -3941,7 +3975,7 @@ lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); -lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); x_36 = !lean_is_exclusive(x_18); if (x_36 == 0) @@ -4093,7 +4127,7 @@ x_32 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_fin x_33 = lean_box(0); x_34 = lean_unbox(x_19); lean_dec(x_19); -x_35 = l_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___lambda__2(x_18, x_16, x_34, x_32, x_1, x_33, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_15); +x_35 = l_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___lambda__2(x_18, x_32, x_16, x_34, x_1, x_33, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_15); return x_35; } } @@ -5884,7 +5918,7 @@ static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect_pushNewArg 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_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___spec__2___closed__1; x_2 = l_Lean_Elab_Term_CollectPatternVars_collect_pushNewArg___closed__1; -x_3 = lean_unsigned_to_nat(273u); +x_3 = lean_unsigned_to_nat(274u); x_4 = lean_unsigned_to_nat(11u); x_5 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___spec__2___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -6169,19 +6203,11 @@ static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatt _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("structInstField", 15); -return x_1; -} -} -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect___spec__2___closed__2() { -_start: -{ -lean_object* x_1; x_1 = lean_mk_string_from_bytes("structInstLVal", 14); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect___spec__2___closed__3() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect___spec__2___closed__2() { _start: { lean_object* x_1; @@ -6189,46 +6215,44 @@ x_1 = lean_mk_string_from_bytes(":=", 2); return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect___spec__2(lean_object* x_1, size_t x_2, size_t 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, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect___spec__2(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_object* x_7, 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) { _start: { -uint8_t x_13; -x_13 = lean_usize_dec_lt(x_3, x_2); -if (x_13 == 0) +uint8_t x_15; +x_15 = lean_usize_dec_lt(x_5, x_4); +if (x_15 == 0) { -lean_object* x_14; +lean_object* x_16; +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_dec(x_2); lean_dec(x_1); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_4); -lean_ctor_set(x_14, 1, x_12); -return x_14; +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_6); +lean_ctor_set(x_16, 1, x_14); +return x_16; } else { -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_15 = lean_array_uget(x_4, x_3); -x_16 = lean_unsigned_to_nat(0u); -x_17 = lean_array_uset(x_4, x_3, x_16); -x_18 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect___spec__2___closed__1; -lean_inc(x_1); -x_19 = lean_name_mk_string(x_1, x_18); -lean_inc(x_15); -x_20 = l_Lean_Syntax_isOfKind(x_15, x_19); +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_17 = lean_array_uget(x_6, x_5); +x_18 = lean_unsigned_to_nat(0u); +x_19 = lean_array_uset(x_6, x_5, x_18); +lean_inc(x_17); +x_20 = l_Lean_Syntax_isOfKind(x_17, x_2); if (x_20 == 0) { lean_object* x_21; uint8_t x_22; lean_dec(x_19); lean_dec(x_17); -lean_dec(x_15); +lean_dec(x_2); lean_dec(x_1); -x_21 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwInvalidPattern___rarg(x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_21 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwInvalidPattern___rarg(x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); x_22 = !lean_is_exclusive(x_21); if (x_22 == 0) { @@ -6251,8 +6275,8 @@ return x_25; else { lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_26 = l_Lean_Syntax_getArg(x_15, x_16); -x_27 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect___spec__2___closed__2; +x_26 = l_Lean_Syntax_getArg(x_17, x_18); +x_27 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect___spec__2___closed__1; lean_inc(x_1); x_28 = lean_name_mk_string(x_1, x_27); lean_inc(x_26); @@ -6264,9 +6288,9 @@ lean_object* x_30; uint8_t x_31; lean_dec(x_26); lean_dec(x_19); lean_dec(x_17); -lean_dec(x_15); +lean_dec(x_2); lean_dec(x_1); -x_30 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwInvalidPattern___rarg(x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_30 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwInvalidPattern___rarg(x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); x_31 = !lean_is_exclusive(x_30); if (x_31 == 0) { @@ -6290,16 +6314,16 @@ else { lean_object* x_35; lean_object* x_36; lean_object* x_37; x_35 = lean_unsigned_to_nat(2u); -x_36 = l_Lean_Syntax_getArg(x_15, x_35); -lean_dec(x_15); +x_36 = l_Lean_Syntax_getArg(x_17, x_35); +lean_dec(x_17); +lean_inc(x_13); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_37 = l_Lean_Elab_Term_CollectPatternVars_collect(x_36, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_37 = l_Lean_Elab_Term_CollectPatternVars_collect(x_36, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); if (lean_obj_tag(x_37) == 0) { lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; size_t x_53; size_t x_54; lean_object* x_55; @@ -6308,18 +6332,18 @@ lean_inc(x_38); x_39 = lean_ctor_get(x_37, 1); lean_inc(x_39); lean_dec(x_37); -lean_inc(x_10); -x_40 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___spec__2___rarg(x_10, x_11, x_39); +lean_inc(x_12); +x_40 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___spec__2___rarg(x_12, x_13, x_39); x_41 = lean_ctor_get(x_40, 0); lean_inc(x_41); x_42 = lean_ctor_get(x_40, 1); lean_inc(x_42); lean_dec(x_40); -x_43 = lean_st_ref_get(x_11, x_42); +x_43 = lean_st_ref_get(x_13, x_42); x_44 = lean_ctor_get(x_43, 1); lean_inc(x_44); lean_dec(x_43); -x_45 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect___spec__2___closed__3; +x_45 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect___spec__2___closed__2; x_46 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_46, 0, x_41); lean_ctor_set(x_46, 1, x_45); @@ -6328,16 +6352,17 @@ x_48 = lean_array_push(x_47, x_26); x_49 = lean_array_push(x_48, x_46); x_50 = lean_array_push(x_49, x_38); x_51 = lean_box(2); +lean_inc(x_2); x_52 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_52, 0, x_51); -lean_ctor_set(x_52, 1, x_19); +lean_ctor_set(x_52, 1, x_2); lean_ctor_set(x_52, 2, x_50); x_53 = 1; -x_54 = lean_usize_add(x_3, x_53); -x_55 = lean_array_uset(x_17, x_3, x_52); -x_3 = x_54; -x_4 = x_55; -x_12 = x_44; +x_54 = lean_usize_add(x_5, x_53); +x_55 = lean_array_uset(x_19, x_5, x_52); +x_5 = x_54; +x_6 = x_55; +x_14 = x_44; goto _start; } else @@ -6345,14 +6370,14 @@ else uint8_t x_57; lean_dec(x_26); lean_dec(x_19); -lean_dec(x_17); +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_dec(x_2); lean_dec(x_1); x_57 = !lean_is_exclusive(x_37); if (x_57 == 0) @@ -6583,6 +6608,24 @@ x_13 = l___private_Init_Meta_0__Array_mapSepElemsMAux___at_Lean_Elab_Term_Collec return x_13; } } +static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("choice", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__1___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___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, lean_object* x_9) { _start: { @@ -6666,7 +6709,7 @@ lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; x_42 = lean_ctor_get(x_40, 0); lean_dec(x_42); x_43 = lean_box(2); -x_44 = l_Lean_choiceKind; +x_44 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__1___closed__2; x_45 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_45, 0, x_43); lean_ctor_set(x_45, 1, x_44); @@ -6681,7 +6724,7 @@ x_46 = lean_ctor_get(x_40, 1); lean_inc(x_46); lean_dec(x_40); x_47 = lean_box(2); -x_48 = l_Lean_choiceKind; +x_48 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__1___closed__2; x_49 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_49, 0, x_47); lean_ctor_set(x_49, 1, x_48); @@ -6770,12 +6813,24 @@ return x_1; static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__3() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(2); +x_2 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__2; +x_3 = lean_alloc_ctor(2, 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_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__4() { +_start: +{ lean_object* x_1; x_1 = lean_mk_string_from_bytes("}", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__4() { +static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__5() { _start: { lean_object* x_1; lean_object* x_2; @@ -6784,7 +6839,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__5() { +static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__6() { _start: { lean_object* x_1; lean_object* x_2; @@ -6793,13 +6848,13 @@ x_2 = l_Array_append___rarg(x_1, x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__6() { +static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); x_2 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__23; -x_3 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__5; +x_3 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__6; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -6807,7 +6862,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__7() { +static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__8() { _start: { lean_object* x_1; @@ -6815,7 +6870,7 @@ x_1 = lean_mk_string_from_bytes(":", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__8() { +static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__9() { _start: { lean_object* x_1; @@ -6823,240 +6878,240 @@ x_1 = lean_mk_string_from_bytes("with", 4); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___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_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, lean_object* x_16) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___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_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, lean_object* x_16, lean_object* x_17, lean_object* x_18) { _start: { -lean_object* x_17; lean_object* x_18; size_t x_19; size_t x_20; lean_object* x_21; -x_17 = l_Lean_Syntax_SepArray_getElems___rarg(x_1); -x_18 = lean_array_get_size(x_17); -x_19 = lean_usize_of_nat(x_18); -lean_dec(x_18); -x_20 = 0; -lean_inc(x_15); -lean_inc(x_14); -x_21 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect___spec__2(x_2, x_19, x_20, x_17, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_21) == 0) +lean_object* x_19; lean_object* x_20; size_t x_21; size_t x_22; lean_object* x_23; +x_19 = l_Lean_Syntax_TSepArray_getElems___rarg(x_1); +x_20 = lean_array_get_size(x_19); +x_21 = lean_usize_of_nat(x_20); +lean_dec(x_20); +x_22 = 0; +lean_inc(x_17); +lean_inc(x_16); +x_23 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect___spec__2(x_2, x_3, x_4, x_21, x_22, x_19, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_23) == 0) { -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; 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_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_21, 1); -lean_inc(x_23); -lean_dec(x_21); -x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___spec__2___rarg(x_14, x_15, x_23); -x_25 = lean_ctor_get(x_24, 0); +lean_object* 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; 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; lean_object* x_44; lean_object* x_45; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); lean_inc(x_25); -x_26 = lean_ctor_get(x_24, 1); -lean_inc(x_26); -lean_dec(x_24); -x_27 = lean_st_ref_get(x_15, x_26); -lean_dec(x_15); -x_28 = lean_ctor_get(x_27, 1); +lean_dec(x_23); +x_26 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___spec__2___rarg(x_16, x_17, x_25); +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); lean_inc(x_28); -if (lean_is_exclusive(x_27)) { - lean_ctor_release(x_27, 0); - lean_ctor_release(x_27, 1); - x_29 = x_27; +lean_dec(x_26); +x_29 = lean_st_ref_get(x_17, x_28); +lean_dec(x_17); +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +if (lean_is_exclusive(x_29)) { + lean_ctor_release(x_29, 0); + lean_ctor_release(x_29, 1); + x_31 = x_29; } else { - lean_dec_ref(x_27); - x_29 = lean_box(0); + lean_dec_ref(x_29); + x_31 = lean_box(0); } -x_30 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__1; -lean_inc(x_25); -x_31 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_31, 0, x_25); -lean_ctor_set(x_31, 1, x_30); -x_32 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__2; -x_33 = l_Lean_Syntax_SepArray_ofElems(x_32, x_22); -lean_dec(x_22); -x_34 = l_Lean_Elab_Term_CollectPatternVars_State_vars___default___closed__1; -x_35 = l_Array_append___rarg(x_34, x_33); -x_36 = lean_box(2); -x_37 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__23; -x_38 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -lean_ctor_set(x_38, 2, x_35); -x_39 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__3; -lean_inc(x_25); -x_40 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_40, 0, x_25); +x_32 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__1; +lean_inc(x_27); +x_33 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_33, 0, x_27); +lean_ctor_set(x_33, 1, x_32); +x_34 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__3; +x_35 = l_Lean_mkSepArray(x_24, x_34); +lean_dec(x_24); +x_36 = l_Lean_Elab_Term_CollectPatternVars_State_vars___default___closed__1; +x_37 = l_Array_append___rarg(x_36, x_35); +x_38 = lean_box(2); +x_39 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__23; +x_40 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_40, 0, x_38); lean_ctor_set(x_40, 1, x_39); +lean_ctor_set(x_40, 2, x_37); x_41 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__4; -x_42 = lean_array_push(x_41, x_31); -if (lean_obj_tag(x_7) == 0) +lean_inc(x_27); +x_42 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_42, 0, x_27); +lean_ctor_set(x_42, 1, x_41); +x_43 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__5; +x_44 = lean_array_push(x_43, x_33); +if (lean_obj_tag(x_9) == 0) { -x_43 = x_34; -goto block_84; +x_45 = x_36; +goto block_86; } else { -lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; -x_85 = lean_ctor_get(x_7, 0); -lean_inc(x_85); -lean_dec(x_7); -x_86 = l_Array_append___rarg(x_34, x_85); -x_87 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_87, 0, x_36); -lean_ctor_set(x_87, 1, x_37); -lean_ctor_set(x_87, 2, x_86); -x_88 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__8; -lean_inc(x_25); -x_89 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_89, 0, x_25); -lean_ctor_set(x_89, 1, x_88); -x_90 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__12; -x_91 = lean_array_push(x_90, x_87); -x_92 = lean_array_push(x_91, x_89); -x_43 = x_92; -goto block_84; +lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_87 = lean_ctor_get(x_9, 0); +lean_inc(x_87); +lean_dec(x_9); +x_88 = l_Array_append___rarg(x_36, x_87); +x_89 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_89, 0, x_38); +lean_ctor_set(x_89, 1, x_39); +lean_ctor_set(x_89, 2, x_88); +x_90 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__9; +lean_inc(x_27); +x_91 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_91, 0, x_27); +lean_ctor_set(x_91, 1, x_90); +x_92 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__12; +x_93 = lean_array_push(x_92, x_89); +x_94 = lean_array_push(x_93, x_91); +x_45 = x_94; +goto block_86; } -block_84: +block_86: { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_44 = l_Array_append___rarg(x_34, x_43); -x_45 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_45, 0, x_36); -lean_ctor_set(x_45, 1, x_37); -lean_ctor_set(x_45, 2, x_44); -x_46 = lean_array_push(x_42, x_45); -x_47 = lean_array_push(x_46, x_38); -if (lean_obj_tag(x_6) == 0) +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_46 = l_Array_append___rarg(x_36, x_45); +x_47 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_47, 0, x_38); +lean_ctor_set(x_47, 1, x_39); +lean_ctor_set(x_47, 2, x_46); +x_48 = lean_array_push(x_44, x_47); +x_49 = lean_array_push(x_48, x_40); +if (lean_obj_tag(x_8) == 0) { -x_48 = x_34; -goto block_72; +x_50 = x_36; +goto block_74; } else { -lean_object* x_73; lean_object* x_74; -x_73 = lean_ctor_get(x_6, 0); -x_74 = l_Lean_Syntax_getHeadInfo_x3f(x_73); -if (lean_obj_tag(x_74) == 0) +lean_object* x_75; lean_object* x_76; +x_75 = lean_ctor_get(x_8, 0); +x_76 = l_Lean_Syntax_getHeadInfo_x3f(x_75); +if (lean_obj_tag(x_76) == 0) { -lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_75 = l_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___lambda__2___closed__2; -lean_inc(x_25); -x_76 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_76, 0, x_25); -lean_ctor_set(x_76, 1, x_75); -x_77 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__27; -x_78 = lean_array_push(x_77, x_76); -x_48 = x_78; -goto block_72; +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_77 = l_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___lambda__2___closed__2; +lean_inc(x_27); +x_78 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_78, 0, x_27); +lean_ctor_set(x_78, 1, x_77); +x_79 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__27; +x_80 = lean_array_push(x_79, x_78); +x_50 = x_80; +goto block_74; } else { -lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; -x_79 = lean_ctor_get(x_74, 0); -lean_inc(x_79); -lean_dec(x_74); -x_80 = l_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___lambda__2___closed__2; -x_81 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_81, 0, x_79); -lean_ctor_set(x_81, 1, x_80); -x_82 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__27; -x_83 = lean_array_push(x_82, x_81); -x_48 = x_83; -goto block_72; -} -} -block_72: -{ -lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_49 = l_Array_append___rarg(x_34, x_48); -x_50 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_50, 0, x_36); -lean_ctor_set(x_50, 1, x_37); -lean_ctor_set(x_50, 2, x_49); -x_51 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__27; -x_52 = lean_array_push(x_51, x_50); -x_53 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_53, 0, x_36); -lean_ctor_set(x_53, 1, x_3); -lean_ctor_set(x_53, 2, x_52); -x_54 = lean_array_push(x_47, x_53); -if (lean_obj_tag(x_4) == 0) -{ -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -lean_dec(x_25); -x_55 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__6; -x_56 = lean_array_push(x_54, x_55); -x_57 = lean_array_push(x_56, x_40); -x_58 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_58, 0, x_36); -lean_ctor_set(x_58, 1, x_5); -lean_ctor_set(x_58, 2, x_57); -if (lean_is_scalar(x_29)) { - x_59 = lean_alloc_ctor(0, 2, 0); +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_81 = lean_ctor_get(x_76, 0); +lean_inc(x_81); +lean_dec(x_76); +x_82 = l_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___lambda__2___closed__2; +x_83 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_83, 0, x_81); +lean_ctor_set(x_83, 1, x_82); +x_84 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__27; +x_85 = lean_array_push(x_84, x_83); +x_50 = x_85; +goto block_74; +} +} +block_74: +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_51 = l_Array_append___rarg(x_36, x_50); +x_52 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_52, 0, x_38); +lean_ctor_set(x_52, 1, x_39); +lean_ctor_set(x_52, 2, x_51); +x_53 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__27; +x_54 = lean_array_push(x_53, x_52); +x_55 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_55, 0, x_38); +lean_ctor_set(x_55, 1, x_5); +lean_ctor_set(x_55, 2, x_54); +x_56 = lean_array_push(x_49, x_55); +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +lean_dec(x_27); +x_57 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__7; +x_58 = lean_array_push(x_56, x_57); +x_59 = lean_array_push(x_58, x_42); +x_60 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_60, 0, x_38); +lean_ctor_set(x_60, 1, x_7); +lean_ctor_set(x_60, 2, x_59); +if (lean_is_scalar(x_31)) { + x_61 = lean_alloc_ctor(0, 2, 0); } else { - x_59 = x_29; + x_61 = x_31; } -lean_ctor_set(x_59, 0, x_58); -lean_ctor_set(x_59, 1, x_28); -return x_59; +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set(x_61, 1, x_30); +return x_61; } 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; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_60 = lean_ctor_get(x_4, 0); -lean_inc(x_60); -lean_dec(x_4); -x_61 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__7; -x_62 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_62, 0, x_25); -lean_ctor_set(x_62, 1, x_61); -x_63 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__12; -x_64 = lean_array_push(x_63, x_62); -x_65 = lean_array_push(x_64, x_60); -x_66 = l_Array_append___rarg(x_34, x_65); -x_67 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_67, 0, x_36); -lean_ctor_set(x_67, 1, x_37); -lean_ctor_set(x_67, 2, x_66); -x_68 = lean_array_push(x_54, x_67); -x_69 = lean_array_push(x_68, x_40); -x_70 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_70, 0, x_36); -lean_ctor_set(x_70, 1, x_5); -lean_ctor_set(x_70, 2, x_69); -if (lean_is_scalar(x_29)) { - x_71 = lean_alloc_ctor(0, 2, 0); +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; 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; +x_62 = lean_ctor_get(x_6, 0); +lean_inc(x_62); +lean_dec(x_6); +x_63 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__8; +x_64 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_64, 0, x_27); +lean_ctor_set(x_64, 1, x_63); +x_65 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__12; +x_66 = lean_array_push(x_65, x_64); +x_67 = lean_array_push(x_66, x_62); +x_68 = l_Array_append___rarg(x_36, x_67); +x_69 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_69, 0, x_38); +lean_ctor_set(x_69, 1, x_39); +lean_ctor_set(x_69, 2, x_68); +x_70 = lean_array_push(x_56, x_69); +x_71 = lean_array_push(x_70, x_42); +x_72 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_72, 0, x_38); +lean_ctor_set(x_72, 1, x_7); +lean_ctor_set(x_72, 2, x_71); +if (lean_is_scalar(x_31)) { + x_73 = lean_alloc_ctor(0, 2, 0); } else { - x_71 = x_29; + x_73 = x_31; } -lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_28); -return x_71; +lean_ctor_set(x_73, 0, x_72); +lean_ctor_set(x_73, 1, x_30); +return x_73; } } } } else { -uint8_t x_93; -lean_dec(x_15); -lean_dec(x_14); +uint8_t x_95; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_9); lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_93 = !lean_is_exclusive(x_21); -if (x_93 == 0) +x_95 = !lean_is_exclusive(x_23); +if (x_95 == 0) { -return x_21; +return x_23; } else { -lean_object* x_94; lean_object* x_95; lean_object* x_96; -x_94 = lean_ctor_get(x_21, 0); -x_95 = lean_ctor_get(x_21, 1); -lean_inc(x_95); -lean_inc(x_94); -lean_dec(x_21); -x_96 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_96, 0, x_94); -lean_ctor_set(x_96, 1, x_95); -return x_96; +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_23, 0); +x_97 = lean_ctor_get(x_23, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_23); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +return x_98; } } } @@ -7065,15 +7120,23 @@ static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__ _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("invalid struct instance pattern, 'with' is not allowed in patterns", 66); +x_1 = lean_mk_string_from_bytes("structInstField", 15); return x_1; } } static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3___closed__2() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("invalid struct instance pattern, 'with' is not allowed in patterns", 66); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3___closed__3() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3___closed__1; +x_1 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3___closed__2; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } @@ -7081,58 +7144,69 @@ return x_2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3(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_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, lean_object* x_16) { _start: { -lean_object* x_17; +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_dec(x_7); -x_17 = l_Lean_Syntax_getArgs(x_1); -lean_dec(x_1); +x_17 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3___closed__1; +lean_inc(x_1); +x_18 = lean_name_mk_string(x_1, x_17); +x_19 = lean_box(0); +lean_inc(x_18); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +x_21 = l_Lean_Syntax_getArgs(x_2); +lean_dec(x_2); if (lean_obj_tag(x_6) == 0) { -lean_object* x_18; lean_object* x_19; -x_18 = lean_box(0); -x_19 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2(x_17, x_2, x_3, x_8, x_4, x_5, x_6, x_18, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2(x_21, x_1, x_18, x_20, x_3, x_8, x_4, x_5, x_6, x_22, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); lean_dec(x_5); -lean_dec(x_17); -return x_19; +lean_dec(x_20); +lean_dec(x_21); +return x_23; } else { -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; uint8_t x_27; -lean_dec(x_17); +lean_object* 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; uint8_t x_31; +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_18); lean_dec(x_8); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -x_20 = lean_ctor_get(x_6, 0); -lean_inc(x_20); +lean_dec(x_1); +x_24 = lean_ctor_get(x_6, 0); +lean_inc(x_24); lean_dec(x_6); -x_21 = l_Lean_Syntax_SepArray_getElems___rarg(x_20); -lean_dec(x_20); -x_22 = lean_box(2); -x_23 = l_Lean_nullKind; -x_24 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_23); -lean_ctor_set(x_24, 2, x_21); -x_25 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3___closed__2; -x_26 = l_Lean_throwErrorAt___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_processVar___spec__2(x_24, x_25, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -x_27 = !lean_is_exclusive(x_26); -if (x_27 == 0) +x_25 = l_Lean_Syntax_TSepArray_getElems___rarg(x_24); +lean_dec(x_24); +x_26 = lean_box(2); +x_27 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__23; +x_28 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +lean_ctor_set(x_28, 2, x_25); +x_29 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3___closed__3; +x_30 = l_Lean_throwErrorAt___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_processVar___spec__2(x_28, x_29, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +x_31 = !lean_is_exclusive(x_30); +if (x_31 == 0) { -return x_26; +return x_30; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_26, 0); -x_29 = lean_ctor_get(x_26, 1); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_26); -x_30 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -return x_30; +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_30, 0); +x_33 = lean_ctor_get(x_30, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_30); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; } } } @@ -7148,14 +7222,13 @@ lean_dec(x_1); x_19 = l_Lean_Syntax_isNone(x_18); if (x_19 == 0) { -lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_20 = l_Lean_nullKind; -x_21 = lean_unsigned_to_nat(2u); +lean_object* x_20; uint8_t x_21; +x_20 = lean_unsigned_to_nat(2u); lean_inc(x_18); -x_22 = l_Lean_Syntax_isNodeOf(x_18, x_20, x_21); -if (x_22 == 0) +x_21 = l_Lean_Syntax_matchesNull(x_18, x_20); +if (x_21 == 0) { -lean_object* x_23; +lean_object* x_22; lean_dec(x_18); lean_dec(x_8); lean_dec(x_6); @@ -7163,30 +7236,30 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_23 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwInvalidPattern___rarg(x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -return x_23; +x_22 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwInvalidPattern___rarg(x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +return x_22; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_24 = lean_unsigned_to_nat(1u); -x_25 = l_Lean_Syntax_getArg(x_18, x_24); +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_23 = lean_unsigned_to_nat(1u); +x_24 = l_Lean_Syntax_getArg(x_18, x_23); lean_dec(x_18); -x_26 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_26, 0, x_25); -x_27 = lean_box(0); -x_28 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3(x_2, x_3, x_4, x_5, x_8, x_6, x_27, x_26, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -return x_28; +x_25 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_25, 0, x_24); +x_26 = lean_box(0); +x_27 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3(x_2, x_3, x_4, x_5, x_8, x_6, x_26, x_25, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +return x_27; } } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_dec(x_18); +x_28 = lean_box(0); x_29 = lean_box(0); -x_30 = lean_box(0); -x_31 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3(x_2, x_3, x_4, x_5, x_8, x_6, x_30, x_29, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -return x_31; +x_30 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3(x_2, x_3, x_4, x_5, x_8, x_6, x_29, x_28, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +return x_30; } } } @@ -7234,14 +7307,13 @@ lean_dec(x_17); x_24 = l_Lean_Syntax_isNone(x_23); if (x_24 == 0) { -lean_object* x_25; lean_object* x_26; uint8_t x_27; -x_25 = l_Lean_nullKind; -x_26 = lean_unsigned_to_nat(1u); +lean_object* x_25; uint8_t x_26; +x_25 = lean_unsigned_to_nat(1u); lean_inc(x_23); -x_27 = l_Lean_Syntax_isNodeOf(x_23, x_25, x_26); -if (x_27 == 0) +x_26 = l_Lean_Syntax_matchesNull(x_23, x_25); +if (x_26 == 0) { -lean_object* x_28; +lean_object* x_27; lean_dec(x_23); lean_dec(x_19); lean_dec(x_15); @@ -7249,29 +7321,29 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_28 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwInvalidPattern___rarg(x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_28; +x_27 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwInvalidPattern___rarg(x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_27; } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_29 = l_Lean_Syntax_getArg(x_23, x_22); +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_28 = l_Lean_Syntax_getArg(x_23, x_22); lean_dec(x_23); -x_30 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_30, 0, x_29); -x_31 = lean_box(0); -x_32 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__4(x_1, x_15, x_2, x_19, x_3, x_5, x_31, x_30, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_32; +x_29 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_29, 0, x_28); +x_30 = lean_box(0); +x_31 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__4(x_1, x_2, x_15, x_19, x_3, x_5, x_30, x_29, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_31; } } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_dec(x_23); +x_32 = lean_box(0); x_33 = lean_box(0); -x_34 = lean_box(0); -x_35 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__4(x_1, x_15, x_2, x_19, x_3, x_5, x_34, x_33, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_35; +x_34 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__4(x_1, x_2, x_15, x_19, x_3, x_5, x_33, x_32, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_34; } } } @@ -8015,7 +8087,7 @@ 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; x_14 = lean_ctor_get(x_12, 0); x_15 = lean_box(2); -x_16 = l_Lean_nullKind; +x_16 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__23; x_17 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_17, 0, x_15); lean_ctor_set(x_17, 1, x_16); @@ -8034,7 +8106,7 @@ lean_inc(x_21); lean_inc(x_20); lean_dec(x_12); x_22 = lean_box(2); -x_23 = l_Lean_nullKind; +x_23 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__23; x_24 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_24, 0, x_22); lean_ctor_set(x_24, 1, x_23); @@ -8185,22 +8257,78 @@ return x_3; static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___closed__13() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__17; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__32; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___closed__15() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("scientific", 10); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__15; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___closed__17() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("char", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__17; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___closed__19() { +_start: +{ lean_object* x_1; x_1 = lean_mk_string_from_bytes("quotedName", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___closed__14() { +static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__8; -x_2 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__13; +x_2 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__19; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___closed__15() { +static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___closed__21() { _start: { lean_object* x_1; @@ -8208,17 +8336,17 @@ x_1 = lean_mk_string_from_bytes("doubleQuotedName", 16); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___closed__16() { +static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__8; -x_2 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__15; +x_2 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__21; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___closed__17() { +static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -8228,7 +8356,7 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___closed__18() { +static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___closed__24() { _start: { lean_object* x_1; @@ -8236,12 +8364,12 @@ x_1 = lean_mk_string_from_bytes("typeAscription", 14); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___closed__19() { +static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__8; -x_2 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__18; +x_2 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__24; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } @@ -8252,7 +8380,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect(lean_object lean_object* x_10; lean_object* x_11; uint8_t x_12; uint8_t x_13; lean_inc(x_1); x_10 = l_Lean_Syntax_getKind(x_1); -x_11 = l_Lean_identKind; +x_11 = l_Lean_Elab_Term_CollectPatternVars_collect_processCtorAppCore___closed__2; x_12 = lean_name_eq(x_10, x_11); x_13 = !lean_is_exclusive(x_7); if (x_13 == 0) @@ -8315,43 +8443,43 @@ x_35 = lean_name_eq(x_10, x_34); if (x_35 == 0) { lean_object* x_36; uint8_t x_37; -x_36 = l_Lean_strLitKind; +x_36 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__13; x_37 = lean_name_eq(x_10, x_36); if (x_37 == 0) { lean_object* x_38; uint8_t x_39; -x_38 = l_Lean_numLitKind; +x_38 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__14; x_39 = lean_name_eq(x_10, x_38); if (x_39 == 0) { lean_object* x_40; uint8_t x_41; -x_40 = l_Lean_scientificLitKind; +x_40 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__16; x_41 = lean_name_eq(x_10, x_40); if (x_41 == 0) { lean_object* x_42; uint8_t x_43; -x_42 = l_Lean_charLitKind; +x_42 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__18; x_43 = lean_name_eq(x_10, x_42); if (x_43 == 0) { lean_object* x_44; uint8_t x_45; -x_44 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__14; +x_44 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__20; x_45 = lean_name_eq(x_10, x_44); if (x_45 == 0) { lean_object* x_46; uint8_t x_47; -x_46 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__16; +x_46 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__22; x_47 = lean_name_eq(x_10, x_46); if (x_47 == 0) { lean_object* x_48; uint8_t x_49; -x_48 = l_Lean_choiceKind; +x_48 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__1___closed__2; x_49 = lean_name_eq(x_10, x_48); lean_dec(x_10); if (x_49 == 0) { lean_object* x_50; uint8_t x_51; -x_50 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__17; +x_50 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__23; lean_inc(x_1); x_51 = l_Lean_Syntax_isOfKind(x_1, x_50); if (x_51 == 0) @@ -8375,151 +8503,150 @@ x_55 = l_Lean_Syntax_getArg(x_1, x_54); x_56 = l_Lean_Syntax_isNone(x_55); if (x_56 == 0) { -lean_object* x_57; lean_object* x_58; uint8_t x_59; -x_57 = l_Lean_nullKind; -x_58 = lean_unsigned_to_nat(2u); +lean_object* x_57; uint8_t x_58; +x_57 = lean_unsigned_to_nat(2u); lean_inc(x_55); -x_59 = l_Lean_Syntax_isNodeOf(x_55, x_57, x_58); -if (x_59 == 0) +x_58 = l_Lean_Syntax_matchesNull(x_55, x_57); +if (x_58 == 0) { -lean_object* x_60; lean_object* x_61; +lean_object* x_59; lean_object* x_60; lean_dec(x_55); lean_dec(x_1); -x_60 = lean_alloc_closure((void*)(l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwInvalidPattern___rarg), 8, 5); -lean_closure_set(x_60, 0, x_2); -lean_closure_set(x_60, 1, x_3); -lean_closure_set(x_60, 2, x_4); -lean_closure_set(x_60, 3, x_5); -lean_closure_set(x_60, 4, x_6); -x_61 = l_Lean_Core_withFreshMacroScope___rarg(x_60, x_7, x_8, x_9); -return x_61; +x_59 = lean_alloc_closure((void*)(l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwInvalidPattern___rarg), 8, 5); +lean_closure_set(x_59, 0, x_2); +lean_closure_set(x_59, 1, x_3); +lean_closure_set(x_59, 2, x_4); +lean_closure_set(x_59, 3, x_5); +lean_closure_set(x_59, 4, x_6); +x_60 = l_Lean_Core_withFreshMacroScope___rarg(x_59, x_7, x_8, x_9); +return x_60; } else { -lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_62 = lean_unsigned_to_nat(0u); -x_63 = l_Lean_Syntax_getArg(x_55, x_62); +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_61 = lean_unsigned_to_nat(0u); +x_62 = l_Lean_Syntax_getArg(x_55, x_61); lean_dec(x_55); -x_64 = l_Lean_Syntax_getArgs(x_63); -lean_dec(x_63); -x_65 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_65, 0, x_64); -x_66 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__8; -x_67 = lean_box(0); -x_68 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__5), 13, 10); -lean_closure_set(x_68, 0, x_1); -lean_closure_set(x_68, 1, x_66); -lean_closure_set(x_68, 2, x_50); -lean_closure_set(x_68, 3, x_67); -lean_closure_set(x_68, 4, x_65); -lean_closure_set(x_68, 5, x_2); -lean_closure_set(x_68, 6, x_3); -lean_closure_set(x_68, 7, x_4); -lean_closure_set(x_68, 8, x_5); -lean_closure_set(x_68, 9, x_6); -x_69 = l_Lean_Core_withFreshMacroScope___rarg(x_68, x_7, x_8, x_9); -return x_69; -} -} -else -{ -lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_63 = l_Lean_Syntax_getArgs(x_62); +lean_dec(x_62); +x_64 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_64, 0, x_63); +x_65 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__8; +x_66 = lean_box(0); +x_67 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__5), 13, 10); +lean_closure_set(x_67, 0, x_1); +lean_closure_set(x_67, 1, x_65); +lean_closure_set(x_67, 2, x_50); +lean_closure_set(x_67, 3, x_66); +lean_closure_set(x_67, 4, x_64); +lean_closure_set(x_67, 5, x_2); +lean_closure_set(x_67, 6, x_3); +lean_closure_set(x_67, 7, x_4); +lean_closure_set(x_67, 8, x_5); +lean_closure_set(x_67, 9, x_6); +x_68 = l_Lean_Core_withFreshMacroScope___rarg(x_67, x_7, x_8, x_9); +return x_68; +} +} +else +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_dec(x_55); -x_70 = lean_box(0); -x_71 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__8; -x_72 = lean_box(0); -x_73 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__5), 13, 10); -lean_closure_set(x_73, 0, x_1); -lean_closure_set(x_73, 1, x_71); -lean_closure_set(x_73, 2, x_50); -lean_closure_set(x_73, 3, x_72); -lean_closure_set(x_73, 4, x_70); -lean_closure_set(x_73, 5, x_2); -lean_closure_set(x_73, 6, x_3); -lean_closure_set(x_73, 7, x_4); -lean_closure_set(x_73, 8, x_5); -lean_closure_set(x_73, 9, x_6); -x_74 = l_Lean_Core_withFreshMacroScope___rarg(x_73, x_7, x_8, x_9); -return x_74; +x_69 = lean_box(0); +x_70 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__8; +x_71 = lean_box(0); +x_72 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__5), 13, 10); +lean_closure_set(x_72, 0, x_1); +lean_closure_set(x_72, 1, x_70); +lean_closure_set(x_72, 2, x_50); +lean_closure_set(x_72, 3, x_71); +lean_closure_set(x_72, 4, x_69); +lean_closure_set(x_72, 5, x_2); +lean_closure_set(x_72, 6, x_3); +lean_closure_set(x_72, 7, x_4); +lean_closure_set(x_72, 8, x_5); +lean_closure_set(x_72, 9, x_6); +x_73 = l_Lean_Core_withFreshMacroScope___rarg(x_72, x_7, x_8, x_9); +return x_73; } } } else { -lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; -x_75 = l_Lean_Syntax_getArgs(x_1); +lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; +x_74 = l_Lean_Syntax_getArgs(x_1); lean_dec(x_1); -x_76 = lean_array_get_size(x_75); -x_77 = lean_unsigned_to_nat(0u); -x_78 = lean_nat_dec_lt(x_77, x_76); -if (x_78 == 0) +x_75 = lean_array_get_size(x_74); +x_76 = lean_unsigned_to_nat(0u); +x_77 = lean_nat_dec_lt(x_76, x_75); +if (x_77 == 0) { -lean_object* x_79; lean_object* x_80; -lean_dec(x_76); -x_79 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__1), 9, 6); -lean_closure_set(x_79, 0, x_2); -lean_closure_set(x_79, 1, x_75); -lean_closure_set(x_79, 2, x_3); -lean_closure_set(x_79, 3, x_4); -lean_closure_set(x_79, 4, x_5); -lean_closure_set(x_79, 5, x_6); -x_80 = l_Lean_Core_withFreshMacroScope___rarg(x_79, x_7, x_8, x_9); -return x_80; +lean_object* x_78; lean_object* x_79; +lean_dec(x_75); +x_78 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__1), 9, 6); +lean_closure_set(x_78, 0, x_2); +lean_closure_set(x_78, 1, x_74); +lean_closure_set(x_78, 2, x_3); +lean_closure_set(x_78, 3, x_4); +lean_closure_set(x_78, 4, x_5); +lean_closure_set(x_78, 5, x_6); +x_79 = l_Lean_Core_withFreshMacroScope___rarg(x_78, x_7, x_8, x_9); +return x_79; } else { -uint8_t x_81; -x_81 = lean_nat_dec_le(x_76, x_76); -if (x_81 == 0) +uint8_t x_80; +x_80 = lean_nat_dec_le(x_75, x_75); +if (x_80 == 0) { -lean_object* x_82; lean_object* x_83; -lean_dec(x_76); -x_82 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__1), 9, 6); -lean_closure_set(x_82, 0, x_2); -lean_closure_set(x_82, 1, x_75); -lean_closure_set(x_82, 2, x_3); -lean_closure_set(x_82, 3, x_4); -lean_closure_set(x_82, 4, x_5); -lean_closure_set(x_82, 5, x_6); -x_83 = l_Lean_Core_withFreshMacroScope___rarg(x_82, x_7, x_8, x_9); -return x_83; +lean_object* x_81; lean_object* x_82; +lean_dec(x_75); +x_81 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__1), 9, 6); +lean_closure_set(x_81, 0, x_2); +lean_closure_set(x_81, 1, x_74); +lean_closure_set(x_81, 2, x_3); +lean_closure_set(x_81, 3, x_4); +lean_closure_set(x_81, 4, x_5); +lean_closure_set(x_81, 5, x_6); +x_82 = l_Lean_Core_withFreshMacroScope___rarg(x_81, x_7, x_8, x_9); +return x_82; } else { -size_t x_84; size_t x_85; lean_object* x_86; uint8_t x_87; -x_84 = 0; -x_85 = lean_usize_of_nat(x_76); -lean_dec(x_76); -x_86 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__8; -x_87 = l_Array_anyMUnsafe_any___at_Lean_Elab_Term_CollectPatternVars_collect___spec__3(x_86, x_75, x_84, x_85); -if (x_87 == 0) -{ -lean_object* x_88; lean_object* x_89; -x_88 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__1), 9, 6); -lean_closure_set(x_88, 0, x_2); -lean_closure_set(x_88, 1, x_75); -lean_closure_set(x_88, 2, x_3); -lean_closure_set(x_88, 3, x_4); -lean_closure_set(x_88, 4, x_5); -lean_closure_set(x_88, 5, x_6); -x_89 = l_Lean_Core_withFreshMacroScope___rarg(x_88, x_7, x_8, x_9); -return x_89; +size_t x_83; size_t x_84; lean_object* x_85; uint8_t x_86; +x_83 = 0; +x_84 = lean_usize_of_nat(x_75); +lean_dec(x_75); +x_85 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__8; +x_86 = l_Array_anyMUnsafe_any___at_Lean_Elab_Term_CollectPatternVars_collect___spec__3(x_85, x_74, x_83, x_84); +if (x_86 == 0) +{ +lean_object* x_87; lean_object* x_88; +x_87 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__1), 9, 6); +lean_closure_set(x_87, 0, x_2); +lean_closure_set(x_87, 1, x_74); +lean_closure_set(x_87, 2, x_3); +lean_closure_set(x_87, 3, x_4); +lean_closure_set(x_87, 4, x_5); +lean_closure_set(x_87, 5, x_6); +x_88 = l_Lean_Core_withFreshMacroScope___rarg(x_87, x_7, x_8, x_9); +return x_88; } else { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_90 = l_Lean_Elab_Term_CollectPatternVars_State_vars___default___closed__1; -x_91 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_CollectPatternVars_collect___spec__4(x_86, x_75, x_84, x_85, x_90); -lean_dec(x_75); -x_92 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__1), 9, 6); -lean_closure_set(x_92, 0, x_2); -lean_closure_set(x_92, 1, x_91); -lean_closure_set(x_92, 2, x_3); -lean_closure_set(x_92, 3, x_4); -lean_closure_set(x_92, 4, x_5); -lean_closure_set(x_92, 5, x_6); -x_93 = l_Lean_Core_withFreshMacroScope___rarg(x_92, x_7, x_8, x_9); -return x_93; +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_89 = l_Lean_Elab_Term_CollectPatternVars_State_vars___default___closed__1; +x_90 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_CollectPatternVars_collect___spec__4(x_85, x_74, x_83, x_84, x_89); +lean_dec(x_74); +x_91 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__1), 9, 6); +lean_closure_set(x_91, 0, x_2); +lean_closure_set(x_91, 1, x_90); +lean_closure_set(x_91, 2, x_3); +lean_closure_set(x_91, 3, x_4); +lean_closure_set(x_91, 4, x_5); +lean_closure_set(x_91, 5, x_6); +x_92 = l_Lean_Core_withFreshMacroScope___rarg(x_91, x_7, x_8, x_9); +return x_92; } } } @@ -8527,365 +8654,364 @@ return x_93; } else { -lean_object* x_94; lean_object* x_95; +lean_object* x_93; lean_object* x_94; lean_dec(x_10); lean_dec(x_2); -x_94 = lean_alloc_closure((void*)(l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_doubleQuotedNameToPattern___boxed), 8, 5); -lean_closure_set(x_94, 0, x_1); -lean_closure_set(x_94, 1, x_3); -lean_closure_set(x_94, 2, x_4); -lean_closure_set(x_94, 3, x_5); -lean_closure_set(x_94, 4, x_6); -x_95 = l_Lean_Core_withFreshMacroScope___rarg(x_94, x_7, x_8, x_9); -return x_95; +x_93 = lean_alloc_closure((void*)(l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_doubleQuotedNameToPattern___boxed), 8, 5); +lean_closure_set(x_93, 0, x_1); +lean_closure_set(x_93, 1, x_3); +lean_closure_set(x_93, 2, x_4); +lean_closure_set(x_93, 3, x_5); +lean_closure_set(x_93, 4, x_6); +x_94 = l_Lean_Core_withFreshMacroScope___rarg(x_93, x_7, x_8, x_9); +return x_94; } } else { -lean_object* x_96; lean_object* x_97; +lean_object* x_95; lean_object* x_96; lean_dec(x_10); lean_dec(x_2); -x_96 = lean_alloc_closure((void*)(l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_quotedNameToPattern), 8, 5); -lean_closure_set(x_96, 0, x_1); -lean_closure_set(x_96, 1, x_3); -lean_closure_set(x_96, 2, x_4); -lean_closure_set(x_96, 3, x_5); -lean_closure_set(x_96, 4, x_6); -x_97 = l_Lean_Core_withFreshMacroScope___rarg(x_96, x_7, x_8, x_9); -return x_97; +x_95 = lean_alloc_closure((void*)(l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_quotedNameToPattern), 8, 5); +lean_closure_set(x_95, 0, x_1); +lean_closure_set(x_95, 1, x_3); +lean_closure_set(x_95, 2, x_4); +lean_closure_set(x_95, 3, x_5); +lean_closure_set(x_95, 4, x_6); +x_96 = l_Lean_Core_withFreshMacroScope___rarg(x_95, x_7, x_8, x_9); +return x_96; } } else { -lean_object* x_98; lean_object* x_99; +lean_object* x_97; lean_object* x_98; lean_dec(x_10); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_98 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__6___boxed), 4, 1); -lean_closure_set(x_98, 0, x_1); -x_99 = l_Lean_Core_withFreshMacroScope___rarg(x_98, x_7, x_8, x_9); -return x_99; +x_97 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__6___boxed), 4, 1); +lean_closure_set(x_97, 0, x_1); +x_98 = l_Lean_Core_withFreshMacroScope___rarg(x_97, x_7, x_8, x_9); +return x_98; } } else { -lean_object* x_100; lean_object* x_101; +lean_object* x_99; lean_object* x_100; lean_dec(x_10); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_100 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__6___boxed), 4, 1); -lean_closure_set(x_100, 0, x_1); -x_101 = l_Lean_Core_withFreshMacroScope___rarg(x_100, x_7, x_8, x_9); -return x_101; +x_99 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__6___boxed), 4, 1); +lean_closure_set(x_99, 0, x_1); +x_100 = l_Lean_Core_withFreshMacroScope___rarg(x_99, x_7, x_8, x_9); +return x_100; } } else { -lean_object* x_102; lean_object* x_103; +lean_object* x_101; lean_object* x_102; lean_dec(x_10); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_102 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__6___boxed), 4, 1); -lean_closure_set(x_102, 0, x_1); -x_103 = l_Lean_Core_withFreshMacroScope___rarg(x_102, x_7, x_8, x_9); -return x_103; +x_101 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__6___boxed), 4, 1); +lean_closure_set(x_101, 0, x_1); +x_102 = l_Lean_Core_withFreshMacroScope___rarg(x_101, x_7, x_8, x_9); +return x_102; } } else { -lean_object* x_104; lean_object* x_105; +lean_object* x_103; lean_object* x_104; lean_dec(x_10); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_104 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__6___boxed), 4, 1); -lean_closure_set(x_104, 0, x_1); -x_105 = l_Lean_Core_withFreshMacroScope___rarg(x_104, x_7, x_8, x_9); -return x_105; +x_103 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__6___boxed), 4, 1); +lean_closure_set(x_103, 0, x_1); +x_104 = l_Lean_Core_withFreshMacroScope___rarg(x_103, x_7, x_8, x_9); +return x_104; } } else { -lean_object* x_106; lean_object* x_107; +lean_object* x_105; lean_object* x_106; lean_dec(x_10); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_106 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__6___boxed), 4, 1); -lean_closure_set(x_106, 0, x_1); -x_107 = l_Lean_Core_withFreshMacroScope___rarg(x_106, x_7, x_8, x_9); -return x_107; +x_105 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__6___boxed), 4, 1); +lean_closure_set(x_105, 0, x_1); +x_106 = l_Lean_Core_withFreshMacroScope___rarg(x_105, x_7, x_8, x_9); +return x_106; } } else { -lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_dec(x_10); -x_108 = lean_unsigned_to_nat(2u); -x_109 = l_Lean_Syntax_getArg(x_1, x_108); -x_110 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__7), 10, 7); -lean_closure_set(x_110, 0, x_109); -lean_closure_set(x_110, 1, x_2); -lean_closure_set(x_110, 2, x_3); -lean_closure_set(x_110, 3, x_4); -lean_closure_set(x_110, 4, x_5); -lean_closure_set(x_110, 5, x_6); -lean_closure_set(x_110, 6, x_1); -x_111 = l_Lean_Core_withFreshMacroScope___rarg(x_110, x_7, x_8, x_9); -return x_111; +x_107 = lean_unsigned_to_nat(2u); +x_108 = l_Lean_Syntax_getArg(x_1, x_107); +x_109 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__7), 10, 7); +lean_closure_set(x_109, 0, x_108); +lean_closure_set(x_109, 1, x_2); +lean_closure_set(x_109, 2, x_3); +lean_closure_set(x_109, 3, x_4); +lean_closure_set(x_109, 4, x_5); +lean_closure_set(x_109, 5, x_6); +lean_closure_set(x_109, 6, x_1); +x_110 = l_Lean_Core_withFreshMacroScope___rarg(x_109, x_7, x_8, x_9); +return x_110; } } else { -lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_dec(x_10); -x_112 = lean_unsigned_to_nat(0u); -x_113 = l_Lean_Syntax_getArg(x_1, x_112); -x_114 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__9), 11, 8); -lean_closure_set(x_114, 0, x_113); -lean_closure_set(x_114, 1, x_2); -lean_closure_set(x_114, 2, x_3); -lean_closure_set(x_114, 3, x_4); -lean_closure_set(x_114, 4, x_5); -lean_closure_set(x_114, 5, x_6); -lean_closure_set(x_114, 6, x_1); -lean_closure_set(x_114, 7, x_16); -x_115 = l_Lean_Core_withFreshMacroScope___rarg(x_114, x_7, x_8, x_9); -return x_115; +x_111 = lean_unsigned_to_nat(0u); +x_112 = l_Lean_Syntax_getArg(x_1, x_111); +x_113 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__9), 11, 8); +lean_closure_set(x_113, 0, x_112); +lean_closure_set(x_113, 1, x_2); +lean_closure_set(x_113, 2, x_3); +lean_closure_set(x_113, 3, x_4); +lean_closure_set(x_113, 4, x_5); +lean_closure_set(x_113, 5, x_6); +lean_closure_set(x_113, 6, x_1); +lean_closure_set(x_113, 7, x_16); +x_114 = l_Lean_Core_withFreshMacroScope___rarg(x_113, x_7, x_8, x_9); +return x_114; } } else { -lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; +lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_dec(x_10); -x_116 = lean_unsigned_to_nat(0u); -x_117 = l_Lean_Syntax_getArg(x_1, x_116); +x_115 = lean_unsigned_to_nat(0u); +x_116 = l_Lean_Syntax_getArg(x_1, x_115); lean_dec(x_1); -x_118 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect_processCtor), 9, 6); -lean_closure_set(x_118, 0, x_117); -lean_closure_set(x_118, 1, x_2); -lean_closure_set(x_118, 2, x_3); -lean_closure_set(x_118, 3, x_4); -lean_closure_set(x_118, 4, x_5); -lean_closure_set(x_118, 5, x_6); -x_119 = l_Lean_Core_withFreshMacroScope___rarg(x_118, x_7, x_8, x_9); -return x_119; +x_117 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect_processCtor), 9, 6); +lean_closure_set(x_117, 0, x_116); +lean_closure_set(x_117, 1, x_2); +lean_closure_set(x_117, 2, x_3); +lean_closure_set(x_117, 3, x_4); +lean_closure_set(x_117, 4, x_5); +lean_closure_set(x_117, 5, x_6); +x_118 = l_Lean_Core_withFreshMacroScope___rarg(x_117, x_7, x_8, x_9); +return x_118; } } else { -lean_object* x_120; lean_object* x_121; uint8_t x_122; +lean_object* x_119; lean_object* x_120; uint8_t x_121; lean_dec(x_10); -x_120 = lean_unsigned_to_nat(1u); -x_121 = l_Lean_Syntax_getArg(x_1, x_120); -x_122 = l_Lean_Syntax_isNone(x_121); -if (x_122 == 0) -{ -lean_object* x_123; lean_object* x_124; lean_object* x_125; uint8_t x_126; -x_123 = lean_unsigned_to_nat(0u); -x_124 = l_Lean_Syntax_getArg(x_121, x_123); -x_125 = l_Lean_Syntax_getArg(x_121, x_120); -x_126 = l_Lean_Syntax_isNone(x_125); -if (x_126 == 0) -{ -lean_object* x_127; lean_object* x_128; lean_object* x_129; uint8_t x_130; -x_127 = l_Lean_Syntax_getArg(x_125, x_123); -lean_dec(x_125); -x_128 = l_Lean_Syntax_getKind(x_127); -x_129 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__19; -x_130 = lean_name_eq(x_128, x_129); -lean_dec(x_128); -if (x_130 == 0) -{ -lean_object* x_131; lean_object* x_132; +x_119 = lean_unsigned_to_nat(1u); +x_120 = l_Lean_Syntax_getArg(x_1, x_119); +x_121 = l_Lean_Syntax_isNone(x_120); +if (x_121 == 0) +{ +lean_object* x_122; lean_object* x_123; lean_object* x_124; uint8_t x_125; +x_122 = lean_unsigned_to_nat(0u); +x_123 = l_Lean_Syntax_getArg(x_120, x_122); +x_124 = l_Lean_Syntax_getArg(x_120, x_119); +x_125 = l_Lean_Syntax_isNone(x_124); +if (x_125 == 0) +{ +lean_object* x_126; lean_object* x_127; lean_object* x_128; uint8_t x_129; +x_126 = l_Lean_Syntax_getArg(x_124, x_122); lean_dec(x_124); -lean_dec(x_121); +x_127 = l_Lean_Syntax_getKind(x_126); +x_128 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__25; +x_129 = lean_name_eq(x_127, x_128); +lean_dec(x_127); +if (x_129 == 0) +{ +lean_object* x_130; lean_object* x_131; +lean_dec(x_123); +lean_dec(x_120); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_131 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__6___boxed), 4, 1); -lean_closure_set(x_131, 0, x_1); -x_132 = l_Lean_Core_withFreshMacroScope___rarg(x_131, x_7, x_8, x_9); -return x_132; +x_130 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__6___boxed), 4, 1); +lean_closure_set(x_130, 0, x_1); +x_131 = l_Lean_Core_withFreshMacroScope___rarg(x_130, x_7, x_8, x_9); +return x_131; } else { -lean_object* x_133; lean_object* x_134; -x_133 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__10), 11, 8); -lean_closure_set(x_133, 0, x_124); -lean_closure_set(x_133, 1, x_2); -lean_closure_set(x_133, 2, x_3); -lean_closure_set(x_133, 3, x_4); -lean_closure_set(x_133, 4, x_5); -lean_closure_set(x_133, 5, x_6); -lean_closure_set(x_133, 6, x_121); -lean_closure_set(x_133, 7, x_1); -x_134 = l_Lean_Core_withFreshMacroScope___rarg(x_133, x_7, x_8, x_9); -return x_134; +lean_object* x_132; lean_object* x_133; +x_132 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__10), 11, 8); +lean_closure_set(x_132, 0, x_123); +lean_closure_set(x_132, 1, x_2); +lean_closure_set(x_132, 2, x_3); +lean_closure_set(x_132, 3, x_4); +lean_closure_set(x_132, 4, x_5); +lean_closure_set(x_132, 5, x_6); +lean_closure_set(x_132, 6, x_120); +lean_closure_set(x_132, 7, x_1); +x_133 = l_Lean_Core_withFreshMacroScope___rarg(x_132, x_7, x_8, x_9); +return x_133; } } else { -lean_object* x_135; lean_object* x_136; -lean_dec(x_125); -x_135 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__10), 11, 8); -lean_closure_set(x_135, 0, x_124); -lean_closure_set(x_135, 1, x_2); -lean_closure_set(x_135, 2, x_3); -lean_closure_set(x_135, 3, x_4); -lean_closure_set(x_135, 4, x_5); -lean_closure_set(x_135, 5, x_6); -lean_closure_set(x_135, 6, x_121); -lean_closure_set(x_135, 7, x_1); -x_136 = l_Lean_Core_withFreshMacroScope___rarg(x_135, x_7, x_8, x_9); -return x_136; +lean_object* x_134; lean_object* x_135; +lean_dec(x_124); +x_134 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__10), 11, 8); +lean_closure_set(x_134, 0, x_123); +lean_closure_set(x_134, 1, x_2); +lean_closure_set(x_134, 2, x_3); +lean_closure_set(x_134, 3, x_4); +lean_closure_set(x_134, 4, x_5); +lean_closure_set(x_134, 5, x_6); +lean_closure_set(x_134, 6, x_120); +lean_closure_set(x_134, 7, x_1); +x_135 = l_Lean_Core_withFreshMacroScope___rarg(x_134, x_7, x_8, x_9); +return x_135; } } else { -lean_object* x_137; lean_object* x_138; -lean_dec(x_121); +lean_object* x_136; lean_object* x_137; +lean_dec(x_120); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_137 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__6___boxed), 4, 1); -lean_closure_set(x_137, 0, x_1); -x_138 = l_Lean_Core_withFreshMacroScope___rarg(x_137, x_7, x_8, x_9); -return x_138; +x_136 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__6___boxed), 4, 1); +lean_closure_set(x_136, 0, x_1); +x_137 = l_Lean_Core_withFreshMacroScope___rarg(x_136, x_7, x_8, x_9); +return x_137; } } } else { -lean_object* x_139; lean_object* x_140; lean_object* x_141; +lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_dec(x_10); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_139 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__8; -x_140 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__11___boxed), 5, 2); -lean_closure_set(x_140, 0, x_139); -lean_closure_set(x_140, 1, x_1); -x_141 = l_Lean_Core_withFreshMacroScope___rarg(x_140, x_7, x_8, x_9); -return x_141; +x_138 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__8; +x_139 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__11___boxed), 5, 2); +lean_closure_set(x_139, 0, x_138); +lean_closure_set(x_139, 1, x_1); +x_140 = l_Lean_Core_withFreshMacroScope___rarg(x_139, x_7, x_8, x_9); +return x_140; } } else { -lean_object* x_142; lean_object* x_143; lean_object* x_144; +lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_dec(x_10); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_142 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__8; -x_143 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__11___boxed), 5, 2); -lean_closure_set(x_143, 0, x_142); -lean_closure_set(x_143, 1, x_1); -x_144 = l_Lean_Core_withFreshMacroScope___rarg(x_143, x_7, x_8, x_9); -return x_144; +x_141 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__8; +x_142 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__11___boxed), 5, 2); +lean_closure_set(x_142, 0, x_141); +lean_closure_set(x_142, 1, x_1); +x_143 = l_Lean_Core_withFreshMacroScope___rarg(x_142, x_7, x_8, x_9); +return x_143; } } else { -lean_object* x_145; lean_object* x_146; +lean_object* x_144; lean_object* x_145; lean_dec(x_10); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_145 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__6___boxed), 4, 1); -lean_closure_set(x_145, 0, x_1); -x_146 = l_Lean_Core_withFreshMacroScope___rarg(x_145, x_7, x_8, x_9); -return x_146; +x_144 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__6___boxed), 4, 1); +lean_closure_set(x_144, 0, x_1); +x_145 = l_Lean_Core_withFreshMacroScope___rarg(x_144, x_7, x_8, x_9); +return x_145; } } else { -lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; +lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_dec(x_10); -x_147 = lean_unsigned_to_nat(1u); -x_148 = l_Lean_Syntax_getArg(x_1, x_147); -x_149 = l_Lean_Syntax_getArgs(x_148); -lean_dec(x_148); -x_150 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__12___boxed), 10, 7); -lean_closure_set(x_150, 0, x_149); -lean_closure_set(x_150, 1, x_2); -lean_closure_set(x_150, 2, x_3); -lean_closure_set(x_150, 3, x_4); -lean_closure_set(x_150, 4, x_5); -lean_closure_set(x_150, 5, x_6); -lean_closure_set(x_150, 6, x_1); -x_151 = l_Lean_Core_withFreshMacroScope___rarg(x_150, x_7, x_8, x_9); -return x_151; +x_146 = lean_unsigned_to_nat(1u); +x_147 = l_Lean_Syntax_getArg(x_1, x_146); +x_148 = l_Lean_Syntax_getArgs(x_147); +lean_dec(x_147); +x_149 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__12___boxed), 10, 7); +lean_closure_set(x_149, 0, x_148); +lean_closure_set(x_149, 1, x_2); +lean_closure_set(x_149, 2, x_3); +lean_closure_set(x_149, 3, x_4); +lean_closure_set(x_149, 4, x_5); +lean_closure_set(x_149, 5, x_6); +lean_closure_set(x_149, 6, x_1); +x_150 = l_Lean_Core_withFreshMacroScope___rarg(x_149, x_7, x_8, x_9); +return x_150; } } else { -lean_object* x_152; lean_object* x_153; +lean_object* x_151; lean_object* x_152; lean_dec(x_10); -x_152 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp), 9, 6); -lean_closure_set(x_152, 0, x_1); -lean_closure_set(x_152, 1, x_2); -lean_closure_set(x_152, 2, x_3); -lean_closure_set(x_152, 3, x_4); -lean_closure_set(x_152, 4, x_5); -lean_closure_set(x_152, 5, x_6); -x_153 = l_Lean_Core_withFreshMacroScope___rarg(x_152, x_7, x_8, x_9); -return x_153; +x_151 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp), 9, 6); +lean_closure_set(x_151, 0, x_1); +lean_closure_set(x_151, 1, x_2); +lean_closure_set(x_151, 2, x_3); +lean_closure_set(x_151, 3, x_4); +lean_closure_set(x_151, 4, x_5); +lean_closure_set(x_151, 5, x_6); +x_152 = l_Lean_Core_withFreshMacroScope___rarg(x_151, x_7, x_8, x_9); +return x_152; } } else { -lean_object* x_154; lean_object* x_155; +lean_object* x_153; lean_object* x_154; lean_dec(x_10); -x_154 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect_processId), 9, 6); -lean_closure_set(x_154, 0, x_1); -lean_closure_set(x_154, 1, x_2); -lean_closure_set(x_154, 2, x_3); -lean_closure_set(x_154, 3, x_4); -lean_closure_set(x_154, 4, x_5); -lean_closure_set(x_154, 5, x_6); -x_155 = l_Lean_Core_withFreshMacroScope___rarg(x_154, x_7, x_8, x_9); -return x_155; -} -} -else -{ -lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; -x_156 = lean_ctor_get(x_7, 0); -x_157 = lean_ctor_get(x_7, 1); -x_158 = lean_ctor_get(x_7, 2); -x_159 = lean_ctor_get(x_7, 3); -x_160 = lean_ctor_get(x_7, 4); -x_161 = lean_ctor_get(x_7, 5); -x_162 = lean_ctor_get(x_7, 6); -x_163 = lean_ctor_get(x_7, 7); -x_164 = lean_ctor_get(x_7, 8); -x_165 = lean_ctor_get(x_7, 9); -x_166 = lean_ctor_get(x_7, 10); -lean_inc(x_166); +x_153 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect_processId), 9, 6); +lean_closure_set(x_153, 0, x_1); +lean_closure_set(x_153, 1, x_2); +lean_closure_set(x_153, 2, x_3); +lean_closure_set(x_153, 3, x_4); +lean_closure_set(x_153, 4, x_5); +lean_closure_set(x_153, 5, x_6); +x_154 = l_Lean_Core_withFreshMacroScope___rarg(x_153, x_7, x_8, x_9); +return x_154; +} +} +else +{ +lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; +x_155 = lean_ctor_get(x_7, 0); +x_156 = lean_ctor_get(x_7, 1); +x_157 = lean_ctor_get(x_7, 2); +x_158 = lean_ctor_get(x_7, 3); +x_159 = lean_ctor_get(x_7, 4); +x_160 = lean_ctor_get(x_7, 5); +x_161 = lean_ctor_get(x_7, 6); +x_162 = lean_ctor_get(x_7, 7); +x_163 = lean_ctor_get(x_7, 8); +x_164 = lean_ctor_get(x_7, 9); +x_165 = lean_ctor_get(x_7, 10); lean_inc(x_165); lean_inc(x_164); lean_inc(x_163); @@ -8896,279 +9022,279 @@ lean_inc(x_159); lean_inc(x_158); lean_inc(x_157); lean_inc(x_156); +lean_inc(x_155); lean_dec(x_7); -x_167 = l_Lean_replaceRef(x_1, x_161); -lean_dec(x_161); -x_168 = lean_alloc_ctor(0, 11, 0); -lean_ctor_set(x_168, 0, x_156); -lean_ctor_set(x_168, 1, x_157); -lean_ctor_set(x_168, 2, x_158); -lean_ctor_set(x_168, 3, x_159); -lean_ctor_set(x_168, 4, x_160); -lean_ctor_set(x_168, 5, x_167); -lean_ctor_set(x_168, 6, x_162); -lean_ctor_set(x_168, 7, x_163); -lean_ctor_set(x_168, 8, x_164); -lean_ctor_set(x_168, 9, x_165); -lean_ctor_set(x_168, 10, x_166); +x_166 = l_Lean_replaceRef(x_1, x_160); +lean_dec(x_160); +x_167 = lean_alloc_ctor(0, 11, 0); +lean_ctor_set(x_167, 0, x_155); +lean_ctor_set(x_167, 1, x_156); +lean_ctor_set(x_167, 2, x_157); +lean_ctor_set(x_167, 3, x_158); +lean_ctor_set(x_167, 4, x_159); +lean_ctor_set(x_167, 5, x_166); +lean_ctor_set(x_167, 6, x_161); +lean_ctor_set(x_167, 7, x_162); +lean_ctor_set(x_167, 8, x_163); +lean_ctor_set(x_167, 9, x_164); +lean_ctor_set(x_167, 10, x_165); if (x_12 == 0) { -lean_object* x_169; uint8_t x_170; -x_169 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__13; -x_170 = lean_name_eq(x_10, x_169); -if (x_170 == 0) -{ -lean_object* x_171; uint8_t x_172; -x_171 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__2; -x_172 = lean_name_eq(x_10, x_171); -if (x_172 == 0) -{ -lean_object* x_173; uint8_t x_174; -x_173 = l_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___closed__2; -x_174 = lean_name_eq(x_10, x_173); -if (x_174 == 0) -{ -lean_object* x_175; uint8_t x_176; -x_175 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__25; -x_176 = lean_name_eq(x_10, x_175); -if (x_176 == 0) -{ -lean_object* x_177; uint8_t x_178; -x_177 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__4; -x_178 = lean_name_eq(x_10, x_177); -if (x_178 == 0) -{ -lean_object* x_179; uint8_t x_180; -x_179 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__6; -x_180 = lean_name_eq(x_10, x_179); -if (x_180 == 0) -{ -lean_object* x_181; uint8_t x_182; -x_181 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__8; -x_182 = lean_name_eq(x_10, x_181); -if (x_182 == 0) -{ -lean_object* x_183; uint8_t x_184; -x_183 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__9; -x_184 = lean_name_eq(x_10, x_183); -if (x_184 == 0) -{ -lean_object* x_185; uint8_t x_186; -x_185 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__11; -x_186 = lean_name_eq(x_10, x_185); -if (x_186 == 0) -{ -lean_object* x_187; uint8_t x_188; -x_187 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__12; -x_188 = lean_name_eq(x_10, x_187); -if (x_188 == 0) -{ -lean_object* x_189; uint8_t x_190; -x_189 = l_Lean_strLitKind; -x_190 = lean_name_eq(x_10, x_189); -if (x_190 == 0) -{ -lean_object* x_191; uint8_t x_192; -x_191 = l_Lean_numLitKind; -x_192 = lean_name_eq(x_10, x_191); -if (x_192 == 0) -{ -lean_object* x_193; uint8_t x_194; -x_193 = l_Lean_scientificLitKind; -x_194 = lean_name_eq(x_10, x_193); -if (x_194 == 0) -{ -lean_object* x_195; uint8_t x_196; -x_195 = l_Lean_charLitKind; -x_196 = lean_name_eq(x_10, x_195); -if (x_196 == 0) -{ -lean_object* x_197; uint8_t x_198; -x_197 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__14; -x_198 = lean_name_eq(x_10, x_197); -if (x_198 == 0) -{ -lean_object* x_199; uint8_t x_200; -x_199 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__16; -x_200 = lean_name_eq(x_10, x_199); -if (x_200 == 0) -{ -lean_object* x_201; uint8_t x_202; -x_201 = l_Lean_choiceKind; -x_202 = lean_name_eq(x_10, x_201); +lean_object* x_168; uint8_t x_169; +x_168 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__13; +x_169 = lean_name_eq(x_10, x_168); +if (x_169 == 0) +{ +lean_object* x_170; uint8_t x_171; +x_170 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__2; +x_171 = lean_name_eq(x_10, x_170); +if (x_171 == 0) +{ +lean_object* x_172; uint8_t x_173; +x_172 = l_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___closed__2; +x_173 = lean_name_eq(x_10, x_172); +if (x_173 == 0) +{ +lean_object* x_174; uint8_t x_175; +x_174 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___closed__25; +x_175 = lean_name_eq(x_10, x_174); +if (x_175 == 0) +{ +lean_object* x_176; uint8_t x_177; +x_176 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__4; +x_177 = lean_name_eq(x_10, x_176); +if (x_177 == 0) +{ +lean_object* x_178; uint8_t x_179; +x_178 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__6; +x_179 = lean_name_eq(x_10, x_178); +if (x_179 == 0) +{ +lean_object* x_180; uint8_t x_181; +x_180 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__8; +x_181 = lean_name_eq(x_10, x_180); +if (x_181 == 0) +{ +lean_object* x_182; uint8_t x_183; +x_182 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__9; +x_183 = lean_name_eq(x_10, x_182); +if (x_183 == 0) +{ +lean_object* x_184; uint8_t x_185; +x_184 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__11; +x_185 = lean_name_eq(x_10, x_184); +if (x_185 == 0) +{ +lean_object* x_186; uint8_t x_187; +x_186 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__12; +x_187 = lean_name_eq(x_10, x_186); +if (x_187 == 0) +{ +lean_object* x_188; uint8_t x_189; +x_188 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__13; +x_189 = lean_name_eq(x_10, x_188); +if (x_189 == 0) +{ +lean_object* x_190; uint8_t x_191; +x_190 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__14; +x_191 = lean_name_eq(x_10, x_190); +if (x_191 == 0) +{ +lean_object* x_192; uint8_t x_193; +x_192 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__16; +x_193 = lean_name_eq(x_10, x_192); +if (x_193 == 0) +{ +lean_object* x_194; uint8_t x_195; +x_194 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__18; +x_195 = lean_name_eq(x_10, x_194); +if (x_195 == 0) +{ +lean_object* x_196; uint8_t x_197; +x_196 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__20; +x_197 = lean_name_eq(x_10, x_196); +if (x_197 == 0) +{ +lean_object* x_198; uint8_t x_199; +x_198 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__22; +x_199 = lean_name_eq(x_10, x_198); +if (x_199 == 0) +{ +lean_object* x_200; uint8_t x_201; +x_200 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__1___closed__2; +x_201 = lean_name_eq(x_10, x_200); lean_dec(x_10); -if (x_202 == 0) +if (x_201 == 0) { -lean_object* x_203; uint8_t x_204; -x_203 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__17; +lean_object* x_202; uint8_t x_203; +x_202 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__23; lean_inc(x_1); -x_204 = l_Lean_Syntax_isOfKind(x_1, x_203); -if (x_204 == 0) +x_203 = l_Lean_Syntax_isOfKind(x_1, x_202); +if (x_203 == 0) { -lean_object* x_205; lean_object* x_206; +lean_object* x_204; lean_object* x_205; lean_dec(x_1); -x_205 = lean_alloc_closure((void*)(l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwInvalidPattern___rarg), 8, 5); -lean_closure_set(x_205, 0, x_2); -lean_closure_set(x_205, 1, x_3); -lean_closure_set(x_205, 2, x_4); -lean_closure_set(x_205, 3, x_5); -lean_closure_set(x_205, 4, x_6); -x_206 = l_Lean_Core_withFreshMacroScope___rarg(x_205, x_168, x_8, x_9); -return x_206; -} -else -{ -lean_object* x_207; lean_object* x_208; uint8_t x_209; -x_207 = lean_unsigned_to_nat(1u); -x_208 = l_Lean_Syntax_getArg(x_1, x_207); -x_209 = l_Lean_Syntax_isNone(x_208); -if (x_209 == 0) -{ -lean_object* x_210; lean_object* x_211; uint8_t x_212; -x_210 = l_Lean_nullKind; -x_211 = lean_unsigned_to_nat(2u); -lean_inc(x_208); -x_212 = l_Lean_Syntax_isNodeOf(x_208, x_210, x_211); -if (x_212 == 0) -{ -lean_object* x_213; lean_object* x_214; -lean_dec(x_208); +x_204 = lean_alloc_closure((void*)(l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwInvalidPattern___rarg), 8, 5); +lean_closure_set(x_204, 0, x_2); +lean_closure_set(x_204, 1, x_3); +lean_closure_set(x_204, 2, x_4); +lean_closure_set(x_204, 3, x_5); +lean_closure_set(x_204, 4, x_6); +x_205 = l_Lean_Core_withFreshMacroScope___rarg(x_204, x_167, x_8, x_9); +return x_205; +} +else +{ +lean_object* x_206; lean_object* x_207; uint8_t x_208; +x_206 = lean_unsigned_to_nat(1u); +x_207 = l_Lean_Syntax_getArg(x_1, x_206); +x_208 = l_Lean_Syntax_isNone(x_207); +if (x_208 == 0) +{ +lean_object* x_209; uint8_t x_210; +x_209 = lean_unsigned_to_nat(2u); +lean_inc(x_207); +x_210 = l_Lean_Syntax_matchesNull(x_207, x_209); +if (x_210 == 0) +{ +lean_object* x_211; lean_object* x_212; +lean_dec(x_207); lean_dec(x_1); -x_213 = lean_alloc_closure((void*)(l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwInvalidPattern___rarg), 8, 5); -lean_closure_set(x_213, 0, x_2); -lean_closure_set(x_213, 1, x_3); -lean_closure_set(x_213, 2, x_4); -lean_closure_set(x_213, 3, x_5); -lean_closure_set(x_213, 4, x_6); -x_214 = l_Lean_Core_withFreshMacroScope___rarg(x_213, x_168, x_8, x_9); -return x_214; -} -else -{ -lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; -x_215 = lean_unsigned_to_nat(0u); -x_216 = l_Lean_Syntax_getArg(x_208, x_215); -lean_dec(x_208); -x_217 = l_Lean_Syntax_getArgs(x_216); -lean_dec(x_216); -x_218 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_218, 0, x_217); -x_219 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__8; -x_220 = lean_box(0); -x_221 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__5), 13, 10); -lean_closure_set(x_221, 0, x_1); -lean_closure_set(x_221, 1, x_219); -lean_closure_set(x_221, 2, x_203); -lean_closure_set(x_221, 3, x_220); -lean_closure_set(x_221, 4, x_218); -lean_closure_set(x_221, 5, x_2); -lean_closure_set(x_221, 6, x_3); -lean_closure_set(x_221, 7, x_4); -lean_closure_set(x_221, 8, x_5); -lean_closure_set(x_221, 9, x_6); -x_222 = l_Lean_Core_withFreshMacroScope___rarg(x_221, x_168, x_8, x_9); -return x_222; -} -} -else -{ -lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; -lean_dec(x_208); +x_211 = lean_alloc_closure((void*)(l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwInvalidPattern___rarg), 8, 5); +lean_closure_set(x_211, 0, x_2); +lean_closure_set(x_211, 1, x_3); +lean_closure_set(x_211, 2, x_4); +lean_closure_set(x_211, 3, x_5); +lean_closure_set(x_211, 4, x_6); +x_212 = l_Lean_Core_withFreshMacroScope___rarg(x_211, x_167, x_8, x_9); +return x_212; +} +else +{ +lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; +x_213 = lean_unsigned_to_nat(0u); +x_214 = l_Lean_Syntax_getArg(x_207, x_213); +lean_dec(x_207); +x_215 = l_Lean_Syntax_getArgs(x_214); +lean_dec(x_214); +x_216 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_216, 0, x_215); +x_217 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__8; +x_218 = lean_box(0); +x_219 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__5), 13, 10); +lean_closure_set(x_219, 0, x_1); +lean_closure_set(x_219, 1, x_217); +lean_closure_set(x_219, 2, x_202); +lean_closure_set(x_219, 3, x_218); +lean_closure_set(x_219, 4, x_216); +lean_closure_set(x_219, 5, x_2); +lean_closure_set(x_219, 6, x_3); +lean_closure_set(x_219, 7, x_4); +lean_closure_set(x_219, 8, x_5); +lean_closure_set(x_219, 9, x_6); +x_220 = l_Lean_Core_withFreshMacroScope___rarg(x_219, x_167, x_8, x_9); +return x_220; +} +} +else +{ +lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; +lean_dec(x_207); +x_221 = lean_box(0); +x_222 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__8; x_223 = lean_box(0); -x_224 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__8; -x_225 = lean_box(0); -x_226 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__5), 13, 10); -lean_closure_set(x_226, 0, x_1); -lean_closure_set(x_226, 1, x_224); -lean_closure_set(x_226, 2, x_203); -lean_closure_set(x_226, 3, x_225); -lean_closure_set(x_226, 4, x_223); -lean_closure_set(x_226, 5, x_2); -lean_closure_set(x_226, 6, x_3); -lean_closure_set(x_226, 7, x_4); -lean_closure_set(x_226, 8, x_5); -lean_closure_set(x_226, 9, x_6); -x_227 = l_Lean_Core_withFreshMacroScope___rarg(x_226, x_168, x_8, x_9); -return x_227; +x_224 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__5), 13, 10); +lean_closure_set(x_224, 0, x_1); +lean_closure_set(x_224, 1, x_222); +lean_closure_set(x_224, 2, x_202); +lean_closure_set(x_224, 3, x_223); +lean_closure_set(x_224, 4, x_221); +lean_closure_set(x_224, 5, x_2); +lean_closure_set(x_224, 6, x_3); +lean_closure_set(x_224, 7, x_4); +lean_closure_set(x_224, 8, x_5); +lean_closure_set(x_224, 9, x_6); +x_225 = l_Lean_Core_withFreshMacroScope___rarg(x_224, x_167, x_8, x_9); +return x_225; } } } else { -lean_object* x_228; lean_object* x_229; lean_object* x_230; uint8_t x_231; -x_228 = l_Lean_Syntax_getArgs(x_1); +lean_object* x_226; lean_object* x_227; lean_object* x_228; uint8_t x_229; +x_226 = l_Lean_Syntax_getArgs(x_1); lean_dec(x_1); -x_229 = lean_array_get_size(x_228); -x_230 = lean_unsigned_to_nat(0u); -x_231 = lean_nat_dec_lt(x_230, x_229); -if (x_231 == 0) -{ -lean_object* x_232; lean_object* x_233; -lean_dec(x_229); -x_232 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__1), 9, 6); -lean_closure_set(x_232, 0, x_2); -lean_closure_set(x_232, 1, x_228); -lean_closure_set(x_232, 2, x_3); -lean_closure_set(x_232, 3, x_4); -lean_closure_set(x_232, 4, x_5); -lean_closure_set(x_232, 5, x_6); -x_233 = l_Lean_Core_withFreshMacroScope___rarg(x_232, x_168, x_8, x_9); -return x_233; -} -else -{ -uint8_t x_234; -x_234 = lean_nat_dec_le(x_229, x_229); -if (x_234 == 0) -{ -lean_object* x_235; lean_object* x_236; -lean_dec(x_229); -x_235 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__1), 9, 6); -lean_closure_set(x_235, 0, x_2); -lean_closure_set(x_235, 1, x_228); -lean_closure_set(x_235, 2, x_3); -lean_closure_set(x_235, 3, x_4); -lean_closure_set(x_235, 4, x_5); -lean_closure_set(x_235, 5, x_6); -x_236 = l_Lean_Core_withFreshMacroScope___rarg(x_235, x_168, x_8, x_9); -return x_236; -} -else -{ -size_t x_237; size_t x_238; lean_object* x_239; uint8_t x_240; -x_237 = 0; -x_238 = lean_usize_of_nat(x_229); -lean_dec(x_229); -x_239 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__8; -x_240 = l_Array_anyMUnsafe_any___at_Lean_Elab_Term_CollectPatternVars_collect___spec__3(x_239, x_228, x_237, x_238); -if (x_240 == 0) -{ -lean_object* x_241; lean_object* x_242; -x_241 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__1), 9, 6); -lean_closure_set(x_241, 0, x_2); -lean_closure_set(x_241, 1, x_228); -lean_closure_set(x_241, 2, x_3); -lean_closure_set(x_241, 3, x_4); -lean_closure_set(x_241, 4, x_5); -lean_closure_set(x_241, 5, x_6); -x_242 = l_Lean_Core_withFreshMacroScope___rarg(x_241, x_168, x_8, x_9); -return x_242; -} -else -{ -lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; -x_243 = l_Lean_Elab_Term_CollectPatternVars_State_vars___default___closed__1; -x_244 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_CollectPatternVars_collect___spec__4(x_239, x_228, x_237, x_238, x_243); -lean_dec(x_228); -x_245 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__1), 9, 6); -lean_closure_set(x_245, 0, x_2); -lean_closure_set(x_245, 1, x_244); -lean_closure_set(x_245, 2, x_3); -lean_closure_set(x_245, 3, x_4); -lean_closure_set(x_245, 4, x_5); -lean_closure_set(x_245, 5, x_6); -x_246 = l_Lean_Core_withFreshMacroScope___rarg(x_245, x_168, x_8, x_9); -return x_246; +x_227 = lean_array_get_size(x_226); +x_228 = lean_unsigned_to_nat(0u); +x_229 = lean_nat_dec_lt(x_228, x_227); +if (x_229 == 0) +{ +lean_object* x_230; lean_object* x_231; +lean_dec(x_227); +x_230 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__1), 9, 6); +lean_closure_set(x_230, 0, x_2); +lean_closure_set(x_230, 1, x_226); +lean_closure_set(x_230, 2, x_3); +lean_closure_set(x_230, 3, x_4); +lean_closure_set(x_230, 4, x_5); +lean_closure_set(x_230, 5, x_6); +x_231 = l_Lean_Core_withFreshMacroScope___rarg(x_230, x_167, x_8, x_9); +return x_231; +} +else +{ +uint8_t x_232; +x_232 = lean_nat_dec_le(x_227, x_227); +if (x_232 == 0) +{ +lean_object* x_233; lean_object* x_234; +lean_dec(x_227); +x_233 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__1), 9, 6); +lean_closure_set(x_233, 0, x_2); +lean_closure_set(x_233, 1, x_226); +lean_closure_set(x_233, 2, x_3); +lean_closure_set(x_233, 3, x_4); +lean_closure_set(x_233, 4, x_5); +lean_closure_set(x_233, 5, x_6); +x_234 = l_Lean_Core_withFreshMacroScope___rarg(x_233, x_167, x_8, x_9); +return x_234; +} +else +{ +size_t x_235; size_t x_236; lean_object* x_237; uint8_t x_238; +x_235 = 0; +x_236 = lean_usize_of_nat(x_227); +lean_dec(x_227); +x_237 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__8; +x_238 = l_Array_anyMUnsafe_any___at_Lean_Elab_Term_CollectPatternVars_collect___spec__3(x_237, x_226, x_235, x_236); +if (x_238 == 0) +{ +lean_object* x_239; lean_object* x_240; +x_239 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__1), 9, 6); +lean_closure_set(x_239, 0, x_2); +lean_closure_set(x_239, 1, x_226); +lean_closure_set(x_239, 2, x_3); +lean_closure_set(x_239, 3, x_4); +lean_closure_set(x_239, 4, x_5); +lean_closure_set(x_239, 5, x_6); +x_240 = l_Lean_Core_withFreshMacroScope___rarg(x_239, x_167, x_8, x_9); +return x_240; +} +else +{ +lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; +x_241 = l_Lean_Elab_Term_CollectPatternVars_State_vars___default___closed__1; +x_242 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_CollectPatternVars_collect___spec__4(x_237, x_226, x_235, x_236, x_241); +lean_dec(x_226); +x_243 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__1), 9, 6); +lean_closure_set(x_243, 0, x_2); +lean_closure_set(x_243, 1, x_242); +lean_closure_set(x_243, 2, x_3); +lean_closure_set(x_243, 3, x_4); +lean_closure_set(x_243, 4, x_5); +lean_closure_set(x_243, 5, x_6); +x_244 = l_Lean_Core_withFreshMacroScope___rarg(x_243, x_167, x_8, x_9); +return x_244; } } } @@ -9176,16 +9302,31 @@ return x_246; } else { +lean_object* x_245; lean_object* x_246; +lean_dec(x_10); +lean_dec(x_2); +x_245 = lean_alloc_closure((void*)(l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_doubleQuotedNameToPattern___boxed), 8, 5); +lean_closure_set(x_245, 0, x_1); +lean_closure_set(x_245, 1, x_3); +lean_closure_set(x_245, 2, x_4); +lean_closure_set(x_245, 3, x_5); +lean_closure_set(x_245, 4, x_6); +x_246 = l_Lean_Core_withFreshMacroScope___rarg(x_245, x_167, x_8, x_9); +return x_246; +} +} +else +{ lean_object* x_247; lean_object* x_248; lean_dec(x_10); lean_dec(x_2); -x_247 = lean_alloc_closure((void*)(l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_doubleQuotedNameToPattern___boxed), 8, 5); +x_247 = lean_alloc_closure((void*)(l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_quotedNameToPattern), 8, 5); lean_closure_set(x_247, 0, x_1); lean_closure_set(x_247, 1, x_3); lean_closure_set(x_247, 2, x_4); lean_closure_set(x_247, 3, x_5); lean_closure_set(x_247, 4, x_6); -x_248 = l_Lean_Core_withFreshMacroScope___rarg(x_247, x_168, x_8, x_9); +x_248 = l_Lean_Core_withFreshMacroScope___rarg(x_247, x_167, x_8, x_9); return x_248; } } @@ -9193,14 +9334,14 @@ else { lean_object* x_249; lean_object* x_250; lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); -x_249 = lean_alloc_closure((void*)(l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_quotedNameToPattern), 8, 5); +x_249 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__6___boxed), 4, 1); lean_closure_set(x_249, 0, x_1); -lean_closure_set(x_249, 1, x_3); -lean_closure_set(x_249, 2, x_4); -lean_closure_set(x_249, 3, x_5); -lean_closure_set(x_249, 4, x_6); -x_250 = l_Lean_Core_withFreshMacroScope___rarg(x_249, x_168, x_8, x_9); +x_250 = l_Lean_Core_withFreshMacroScope___rarg(x_249, x_167, x_8, x_9); return x_250; } } @@ -9215,7 +9356,7 @@ lean_dec(x_3); lean_dec(x_2); x_251 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__6___boxed), 4, 1); lean_closure_set(x_251, 0, x_1); -x_252 = l_Lean_Core_withFreshMacroScope___rarg(x_251, x_168, x_8, x_9); +x_252 = l_Lean_Core_withFreshMacroScope___rarg(x_251, x_167, x_8, x_9); return x_252; } } @@ -9230,7 +9371,7 @@ lean_dec(x_3); lean_dec(x_2); x_253 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__6___boxed), 4, 1); lean_closure_set(x_253, 0, x_1); -x_254 = l_Lean_Core_withFreshMacroScope___rarg(x_253, x_168, x_8, x_9); +x_254 = l_Lean_Core_withFreshMacroScope___rarg(x_253, x_167, x_8, x_9); return x_254; } } @@ -9245,7 +9386,7 @@ lean_dec(x_3); lean_dec(x_2); x_255 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__6___boxed), 4, 1); lean_closure_set(x_255, 0, x_1); -x_256 = l_Lean_Core_withFreshMacroScope___rarg(x_255, x_168, x_8, x_9); +x_256 = l_Lean_Core_withFreshMacroScope___rarg(x_255, x_167, x_8, x_9); return x_256; } } @@ -9260,233 +9401,233 @@ lean_dec(x_3); lean_dec(x_2); x_257 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__6___boxed), 4, 1); lean_closure_set(x_257, 0, x_1); -x_258 = l_Lean_Core_withFreshMacroScope___rarg(x_257, x_168, x_8, x_9); +x_258 = l_Lean_Core_withFreshMacroScope___rarg(x_257, x_167, x_8, x_9); return x_258; } } else { -lean_object* x_259; lean_object* x_260; -lean_dec(x_10); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_259 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__6___boxed), 4, 1); -lean_closure_set(x_259, 0, x_1); -x_260 = l_Lean_Core_withFreshMacroScope___rarg(x_259, x_168, x_8, x_9); -return x_260; -} -} -else -{ -lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; +lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_dec(x_10); -x_261 = lean_unsigned_to_nat(2u); -x_262 = l_Lean_Syntax_getArg(x_1, x_261); -x_263 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__7), 10, 7); -lean_closure_set(x_263, 0, x_262); -lean_closure_set(x_263, 1, x_2); -lean_closure_set(x_263, 2, x_3); -lean_closure_set(x_263, 3, x_4); -lean_closure_set(x_263, 4, x_5); -lean_closure_set(x_263, 5, x_6); -lean_closure_set(x_263, 6, x_1); -x_264 = l_Lean_Core_withFreshMacroScope___rarg(x_263, x_168, x_8, x_9); -return x_264; +x_259 = lean_unsigned_to_nat(2u); +x_260 = l_Lean_Syntax_getArg(x_1, x_259); +x_261 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__7), 10, 7); +lean_closure_set(x_261, 0, x_260); +lean_closure_set(x_261, 1, x_2); +lean_closure_set(x_261, 2, x_3); +lean_closure_set(x_261, 3, x_4); +lean_closure_set(x_261, 4, x_5); +lean_closure_set(x_261, 5, x_6); +lean_closure_set(x_261, 6, x_1); +x_262 = l_Lean_Core_withFreshMacroScope___rarg(x_261, x_167, x_8, x_9); +return x_262; } } else { -lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; +lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_dec(x_10); -x_265 = lean_unsigned_to_nat(0u); -x_266 = l_Lean_Syntax_getArg(x_1, x_265); -x_267 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__9), 11, 8); -lean_closure_set(x_267, 0, x_266); -lean_closure_set(x_267, 1, x_2); -lean_closure_set(x_267, 2, x_3); -lean_closure_set(x_267, 3, x_4); -lean_closure_set(x_267, 4, x_5); -lean_closure_set(x_267, 5, x_6); -lean_closure_set(x_267, 6, x_1); -lean_closure_set(x_267, 7, x_169); -x_268 = l_Lean_Core_withFreshMacroScope___rarg(x_267, x_168, x_8, x_9); -return x_268; +x_263 = lean_unsigned_to_nat(0u); +x_264 = l_Lean_Syntax_getArg(x_1, x_263); +x_265 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__9), 11, 8); +lean_closure_set(x_265, 0, x_264); +lean_closure_set(x_265, 1, x_2); +lean_closure_set(x_265, 2, x_3); +lean_closure_set(x_265, 3, x_4); +lean_closure_set(x_265, 4, x_5); +lean_closure_set(x_265, 5, x_6); +lean_closure_set(x_265, 6, x_1); +lean_closure_set(x_265, 7, x_168); +x_266 = l_Lean_Core_withFreshMacroScope___rarg(x_265, x_167, x_8, x_9); +return x_266; } } else { -lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; +lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_dec(x_10); -x_269 = lean_unsigned_to_nat(0u); -x_270 = l_Lean_Syntax_getArg(x_1, x_269); +x_267 = lean_unsigned_to_nat(0u); +x_268 = l_Lean_Syntax_getArg(x_1, x_267); lean_dec(x_1); -x_271 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect_processCtor), 9, 6); -lean_closure_set(x_271, 0, x_270); -lean_closure_set(x_271, 1, x_2); -lean_closure_set(x_271, 2, x_3); -lean_closure_set(x_271, 3, x_4); -lean_closure_set(x_271, 4, x_5); -lean_closure_set(x_271, 5, x_6); -x_272 = l_Lean_Core_withFreshMacroScope___rarg(x_271, x_168, x_8, x_9); -return x_272; +x_269 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect_processCtor), 9, 6); +lean_closure_set(x_269, 0, x_268); +lean_closure_set(x_269, 1, x_2); +lean_closure_set(x_269, 2, x_3); +lean_closure_set(x_269, 3, x_4); +lean_closure_set(x_269, 4, x_5); +lean_closure_set(x_269, 5, x_6); +x_270 = l_Lean_Core_withFreshMacroScope___rarg(x_269, x_167, x_8, x_9); +return x_270; } } else { -lean_object* x_273; lean_object* x_274; uint8_t x_275; +lean_object* x_271; lean_object* x_272; uint8_t x_273; lean_dec(x_10); -x_273 = lean_unsigned_to_nat(1u); -x_274 = l_Lean_Syntax_getArg(x_1, x_273); -x_275 = l_Lean_Syntax_isNone(x_274); -if (x_275 == 0) -{ -lean_object* x_276; lean_object* x_277; lean_object* x_278; uint8_t x_279; -x_276 = lean_unsigned_to_nat(0u); -x_277 = l_Lean_Syntax_getArg(x_274, x_276); -x_278 = l_Lean_Syntax_getArg(x_274, x_273); -x_279 = l_Lean_Syntax_isNone(x_278); -if (x_279 == 0) -{ -lean_object* x_280; lean_object* x_281; lean_object* x_282; uint8_t x_283; -x_280 = l_Lean_Syntax_getArg(x_278, x_276); -lean_dec(x_278); -x_281 = l_Lean_Syntax_getKind(x_280); -x_282 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__19; -x_283 = lean_name_eq(x_281, x_282); -lean_dec(x_281); -if (x_283 == 0) -{ -lean_object* x_284; lean_object* x_285; -lean_dec(x_277); -lean_dec(x_274); +x_271 = lean_unsigned_to_nat(1u); +x_272 = l_Lean_Syntax_getArg(x_1, x_271); +x_273 = l_Lean_Syntax_isNone(x_272); +if (x_273 == 0) +{ +lean_object* x_274; lean_object* x_275; lean_object* x_276; uint8_t x_277; +x_274 = lean_unsigned_to_nat(0u); +x_275 = l_Lean_Syntax_getArg(x_272, x_274); +x_276 = l_Lean_Syntax_getArg(x_272, x_271); +x_277 = l_Lean_Syntax_isNone(x_276); +if (x_277 == 0) +{ +lean_object* x_278; lean_object* x_279; lean_object* x_280; uint8_t x_281; +x_278 = l_Lean_Syntax_getArg(x_276, x_274); +lean_dec(x_276); +x_279 = l_Lean_Syntax_getKind(x_278); +x_280 = l_Lean_Elab_Term_CollectPatternVars_collect___closed__25; +x_281 = lean_name_eq(x_279, x_280); +lean_dec(x_279); +if (x_281 == 0) +{ +lean_object* x_282; lean_object* x_283; +lean_dec(x_275); +lean_dec(x_272); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_284 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__6___boxed), 4, 1); -lean_closure_set(x_284, 0, x_1); -x_285 = l_Lean_Core_withFreshMacroScope___rarg(x_284, x_168, x_8, x_9); +x_282 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__6___boxed), 4, 1); +lean_closure_set(x_282, 0, x_1); +x_283 = l_Lean_Core_withFreshMacroScope___rarg(x_282, x_167, x_8, x_9); +return x_283; +} +else +{ +lean_object* x_284; lean_object* x_285; +x_284 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__10), 11, 8); +lean_closure_set(x_284, 0, x_275); +lean_closure_set(x_284, 1, x_2); +lean_closure_set(x_284, 2, x_3); +lean_closure_set(x_284, 3, x_4); +lean_closure_set(x_284, 4, x_5); +lean_closure_set(x_284, 5, x_6); +lean_closure_set(x_284, 6, x_272); +lean_closure_set(x_284, 7, x_1); +x_285 = l_Lean_Core_withFreshMacroScope___rarg(x_284, x_167, x_8, x_9); return x_285; } +} else { lean_object* x_286; lean_object* x_287; +lean_dec(x_276); x_286 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__10), 11, 8); -lean_closure_set(x_286, 0, x_277); +lean_closure_set(x_286, 0, x_275); lean_closure_set(x_286, 1, x_2); lean_closure_set(x_286, 2, x_3); lean_closure_set(x_286, 3, x_4); lean_closure_set(x_286, 4, x_5); lean_closure_set(x_286, 5, x_6); -lean_closure_set(x_286, 6, x_274); +lean_closure_set(x_286, 6, x_272); lean_closure_set(x_286, 7, x_1); -x_287 = l_Lean_Core_withFreshMacroScope___rarg(x_286, x_168, x_8, x_9); +x_287 = l_Lean_Core_withFreshMacroScope___rarg(x_286, x_167, x_8, x_9); return x_287; } } else { lean_object* x_288; lean_object* x_289; -lean_dec(x_278); -x_288 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__10), 11, 8); -lean_closure_set(x_288, 0, x_277); -lean_closure_set(x_288, 1, x_2); -lean_closure_set(x_288, 2, x_3); -lean_closure_set(x_288, 3, x_4); -lean_closure_set(x_288, 4, x_5); -lean_closure_set(x_288, 5, x_6); -lean_closure_set(x_288, 6, x_274); -lean_closure_set(x_288, 7, x_1); -x_289 = l_Lean_Core_withFreshMacroScope___rarg(x_288, x_168, x_8, x_9); -return x_289; -} -} -else -{ -lean_object* x_290; lean_object* x_291; -lean_dec(x_274); +lean_dec(x_272); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_290 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__6___boxed), 4, 1); -lean_closure_set(x_290, 0, x_1); -x_291 = l_Lean_Core_withFreshMacroScope___rarg(x_290, x_168, x_8, x_9); -return x_291; +x_288 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__6___boxed), 4, 1); +lean_closure_set(x_288, 0, x_1); +x_289 = l_Lean_Core_withFreshMacroScope___rarg(x_288, x_167, x_8, x_9); +return x_289; } } } else { -lean_object* x_292; lean_object* x_293; lean_object* x_294; +lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_dec(x_10); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_292 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__8; -x_293 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__11___boxed), 5, 2); -lean_closure_set(x_293, 0, x_292); -lean_closure_set(x_293, 1, x_1); -x_294 = l_Lean_Core_withFreshMacroScope___rarg(x_293, x_168, x_8, x_9); -return x_294; +x_290 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__8; +x_291 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__11___boxed), 5, 2); +lean_closure_set(x_291, 0, x_290); +lean_closure_set(x_291, 1, x_1); +x_292 = l_Lean_Core_withFreshMacroScope___rarg(x_291, x_167, x_8, x_9); +return x_292; } } else { -lean_object* x_295; lean_object* x_296; lean_object* x_297; +lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_dec(x_10); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_295 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__8; -x_296 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__11___boxed), 5, 2); -lean_closure_set(x_296, 0, x_295); -lean_closure_set(x_296, 1, x_1); -x_297 = l_Lean_Core_withFreshMacroScope___rarg(x_296, x_168, x_8, x_9); -return x_297; +x_293 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__8; +x_294 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__11___boxed), 5, 2); +lean_closure_set(x_294, 0, x_293); +lean_closure_set(x_294, 1, x_1); +x_295 = l_Lean_Core_withFreshMacroScope___rarg(x_294, x_167, x_8, x_9); +return x_295; } } else { -lean_object* x_298; lean_object* x_299; +lean_object* x_296; lean_object* x_297; lean_dec(x_10); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_298 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__6___boxed), 4, 1); -lean_closure_set(x_298, 0, x_1); -x_299 = l_Lean_Core_withFreshMacroScope___rarg(x_298, x_168, x_8, x_9); -return x_299; +x_296 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__6___boxed), 4, 1); +lean_closure_set(x_296, 0, x_1); +x_297 = l_Lean_Core_withFreshMacroScope___rarg(x_296, x_167, x_8, x_9); +return x_297; +} +} +else +{ +lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; +lean_dec(x_10); +x_298 = lean_unsigned_to_nat(1u); +x_299 = l_Lean_Syntax_getArg(x_1, x_298); +x_300 = l_Lean_Syntax_getArgs(x_299); +lean_dec(x_299); +x_301 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__12___boxed), 10, 7); +lean_closure_set(x_301, 0, x_300); +lean_closure_set(x_301, 1, x_2); +lean_closure_set(x_301, 2, x_3); +lean_closure_set(x_301, 3, x_4); +lean_closure_set(x_301, 4, x_5); +lean_closure_set(x_301, 5, x_6); +lean_closure_set(x_301, 6, x_1); +x_302 = l_Lean_Core_withFreshMacroScope___rarg(x_301, x_167, x_8, x_9); +return x_302; } } else { -lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; +lean_object* x_303; lean_object* x_304; lean_dec(x_10); -x_300 = lean_unsigned_to_nat(1u); -x_301 = l_Lean_Syntax_getArg(x_1, x_300); -x_302 = l_Lean_Syntax_getArgs(x_301); -lean_dec(x_301); -x_303 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__12___boxed), 10, 7); -lean_closure_set(x_303, 0, x_302); +x_303 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp), 9, 6); +lean_closure_set(x_303, 0, x_1); lean_closure_set(x_303, 1, x_2); lean_closure_set(x_303, 2, x_3); lean_closure_set(x_303, 3, x_4); lean_closure_set(x_303, 4, x_5); lean_closure_set(x_303, 5, x_6); -lean_closure_set(x_303, 6, x_1); -x_304 = l_Lean_Core_withFreshMacroScope___rarg(x_303, x_168, x_8, x_9); +x_304 = l_Lean_Core_withFreshMacroScope___rarg(x_303, x_167, x_8, x_9); return x_304; } } @@ -9494,32 +9635,17 @@ else { lean_object* x_305; lean_object* x_306; lean_dec(x_10); -x_305 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp), 9, 6); +x_305 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect_processId), 9, 6); lean_closure_set(x_305, 0, x_1); lean_closure_set(x_305, 1, x_2); lean_closure_set(x_305, 2, x_3); lean_closure_set(x_305, 3, x_4); lean_closure_set(x_305, 4, x_5); lean_closure_set(x_305, 5, x_6); -x_306 = l_Lean_Core_withFreshMacroScope___rarg(x_305, x_168, x_8, x_9); +x_306 = l_Lean_Core_withFreshMacroScope___rarg(x_305, x_167, x_8, x_9); return x_306; } } -else -{ -lean_object* x_307; lean_object* x_308; -lean_dec(x_10); -x_307 = lean_alloc_closure((void*)(l_Lean_Elab_Term_CollectPatternVars_collect_processId), 9, 6); -lean_closure_set(x_307, 0, x_1); -lean_closure_set(x_307, 1, x_2); -lean_closure_set(x_307, 2, x_3); -lean_closure_set(x_307, 3, x_4); -lean_closure_set(x_307, 4, x_5); -lean_closure_set(x_307, 5, x_6); -x_308 = l_Lean_Core_withFreshMacroScope___rarg(x_307, x_168, x_8, x_9); -return x_308; -} -} } } LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect_processCtor(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_object* x_9) { @@ -9726,16 +9852,17 @@ return x_26; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___spec__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, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___spec__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, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -size_t x_12; size_t x_13; lean_object* x_14; -x_12 = lean_unbox_usize(x_1); -lean_dec(x_1); +size_t x_13; size_t x_14; lean_object* x_15; x_13 = lean_unbox_usize(x_2); lean_dec(x_2); -x_14 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___spec__2(x_12, x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_14; +x_14 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_15 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___spec__2(x_1, x_13, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_1); +return x_15; } } LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___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, lean_object* x_9, lean_object* x_10, lean_object* x_11) { @@ -9758,9 +9885,9 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect_processCtor _start: { uint8_t x_15; lean_object* x_16; -x_15 = lean_unbox(x_3); -lean_dec(x_3); -x_16 = l_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___lambda__2(x_1, x_2, x_15, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +x_15 = lean_unbox(x_4); +lean_dec(x_4); +x_16 = l_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___lambda__2(x_1, x_2, x_3, x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); return x_16; } } @@ -9933,16 +10060,17 @@ lean_dec(x_3); return x_17; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect___spec__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, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect___spec__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, lean_object* x_7, 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) { _start: { -size_t x_13; size_t x_14; lean_object* x_15; -x_13 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_14 = lean_unbox_usize(x_3); +size_t x_15; size_t x_16; lean_object* x_17; +x_15 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_16 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_17 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect___spec__2(x_1, x_2, x_3, x_15, x_16, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); lean_dec(x_3); -x_15 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect___spec__2(x_1, x_13, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_15; +return x_17; } } LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_CollectPatternVars_collect___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { @@ -9990,15 +10118,34 @@ lean_dec(x_1); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___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, lean_object* x_7, 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, lean_object* x_15, lean_object* x_16) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; _start: { -lean_object* x_17; -x_17 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_object* x_19; +x_19 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_10); lean_dec(x_8); -lean_dec(x_6); +lean_dec(x_4); lean_dec(x_1); -return x_17; +return x_19; } } LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { @@ -12579,10 +12726,12 @@ l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect___spec__2_ lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect___spec__2___closed__1); l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect___spec__2___closed__2 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect___spec__2___closed__2(); lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect___spec__2___closed__2); -l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect___spec__2___closed__3 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect___spec__2___closed__3(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect___spec__2___closed__3); l_Array_anyMUnsafe_any___at_Lean_Elab_Term_CollectPatternVars_collect___spec__3___closed__1 = _init_l_Array_anyMUnsafe_any___at_Lean_Elab_Term_CollectPatternVars_collect___spec__3___closed__1(); lean_mark_persistent(l_Array_anyMUnsafe_any___at_Lean_Elab_Term_CollectPatternVars_collect___spec__3___closed__1); +l_Lean_Elab_Term_CollectPatternVars_collect___lambda__1___closed__1 = _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__1___closed__1); +l_Lean_Elab_Term_CollectPatternVars_collect___lambda__1___closed__2 = _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__1___closed__2); l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__1 = _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__1); l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__2 = _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__2(); @@ -12599,10 +12748,14 @@ l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__7 = _init_l_Le lean_mark_persistent(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__7); l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__8 = _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__8(); lean_mark_persistent(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__8); +l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__9 = _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__9(); +lean_mark_persistent(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__9); l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3___closed__1 = _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3___closed__1); l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3___closed__2 = _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3___closed__2(); lean_mark_persistent(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3___closed__2); +l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3___closed__3 = _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3___closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3___closed__3); l_Lean_Elab_Term_CollectPatternVars_collect___lambda__5___closed__1 = _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__5___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__5___closed__1); l_Lean_Elab_Term_CollectPatternVars_collect___lambda__8___closed__1 = _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__8___closed__1(); @@ -12679,6 +12832,18 @@ l_Lean_Elab_Term_CollectPatternVars_collect___closed__18 = _init_l_Lean_Elab_Ter lean_mark_persistent(l_Lean_Elab_Term_CollectPatternVars_collect___closed__18); l_Lean_Elab_Term_CollectPatternVars_collect___closed__19 = _init_l_Lean_Elab_Term_CollectPatternVars_collect___closed__19(); lean_mark_persistent(l_Lean_Elab_Term_CollectPatternVars_collect___closed__19); +l_Lean_Elab_Term_CollectPatternVars_collect___closed__20 = _init_l_Lean_Elab_Term_CollectPatternVars_collect___closed__20(); +lean_mark_persistent(l_Lean_Elab_Term_CollectPatternVars_collect___closed__20); +l_Lean_Elab_Term_CollectPatternVars_collect___closed__21 = _init_l_Lean_Elab_Term_CollectPatternVars_collect___closed__21(); +lean_mark_persistent(l_Lean_Elab_Term_CollectPatternVars_collect___closed__21); +l_Lean_Elab_Term_CollectPatternVars_collect___closed__22 = _init_l_Lean_Elab_Term_CollectPatternVars_collect___closed__22(); +lean_mark_persistent(l_Lean_Elab_Term_CollectPatternVars_collect___closed__22); +l_Lean_Elab_Term_CollectPatternVars_collect___closed__23 = _init_l_Lean_Elab_Term_CollectPatternVars_collect___closed__23(); +lean_mark_persistent(l_Lean_Elab_Term_CollectPatternVars_collect___closed__23); +l_Lean_Elab_Term_CollectPatternVars_collect___closed__24 = _init_l_Lean_Elab_Term_CollectPatternVars_collect___closed__24(); +lean_mark_persistent(l_Lean_Elab_Term_CollectPatternVars_collect___closed__24); +l_Lean_Elab_Term_CollectPatternVars_collect___closed__25 = _init_l_Lean_Elab_Term_CollectPatternVars_collect___closed__25(); +lean_mark_persistent(l_Lean_Elab_Term_CollectPatternVars_collect___closed__25); l_Lean_addTrace___at_Lean_Elab_Term_CollectPatternVars_main___spec__1___closed__1 = _init_l_Lean_addTrace___at_Lean_Elab_Term_CollectPatternVars_main___spec__1___closed__1(); lean_mark_persistent(l_Lean_addTrace___at_Lean_Elab_Term_CollectPatternVars_main___spec__1___closed__1); l_Lean_addTrace___at_Lean_Elab_Term_CollectPatternVars_main___spec__1___closed__2 = _init_l_Lean_addTrace___at_Lean_Elab_Term_CollectPatternVars_main___spec__1___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/Main.c b/stage0/stdlib/Lean/Elab/PreDefinition/Main.c index 436b0528ead..f95206377d2 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/Main.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/Main.c @@ -52,7 +52,6 @@ static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_PreDefinitio lean_object* l_List_mapTRAux___at_Lean_Elab_addAndCompileUnsafe___spec__2(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__1___closed__6; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21(lean_object*, lean_object*, lean_object*, size_t, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__14; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_PreDefinition_Main_0__Lean_Elab_addAndCompilePartial___spec__2___lambda__3___closed__1; lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_PreDefinition_Main_0__Lean_Elab_partitionPreDefs___spec__25(lean_object*, size_t, size_t, lean_object*); @@ -129,12 +128,11 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_PreDe lean_object* l_Lean_throwError___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__1___closed__3; lean_object* l_Lean_Elab_applyAttributesOf(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_Main___hyg_2378_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_Main___hyg_2382_(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_addPreDefinitions___spec__17___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_PreDefinition_Main_0__Lean_Elab_addAndCompilePartial___spec__2___closed__6; lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_addPreDefinitions___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__13; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_PreDefinition_Main_0__Lean_Elab_partitionPreDefs___spec__25___lambda__1(lean_object*, lean_object*); lean_object* l_Lean_Expr_FoldConstsImpl_fold_visit___rarg(lean_object*, size_t, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_mkInhabitantFor(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -146,7 +144,6 @@ lean_object* l_Lean_addDecl___at_Lean_Elab_Term_evalExpr___spec__2(lean_object*, lean_object* l_Lean_Meta_mkSorry(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_PreDefinition_Main_0__Lean_Elab_betaReduceLetRecApps___spec__3___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_expandMacroImpl_x3f(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__12; LEAN_EXPORT lean_object* l___private_Lean_Util_SCC_0__Lean_SCC_push___at___private_Lean_Elab_PreDefinition_Main_0__Lean_Elab_partitionPreDefs___spec__10(lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__19(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -155,6 +152,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_addPreDefinitions___lambda__1___boxed(lean_ lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_Main_0__Lean_Elab_addAndCompilePartial(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_PreDefinition_Main_0__Lean_Elab_partitionPreDefs___spec__4(lean_object*, lean_object*, size_t, size_t); +uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_PreDefinition_Main_0__Lean_Elab_partitionPreDefs___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_addPreDefinitions___spec__26(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_transform___at_Lean_Meta_zetaReduce___spec__1(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -244,7 +242,6 @@ lean_object* l_Lean_setEnv___at_Lean_Meta_setInlineAttribute___spec__2(lean_obje uint8_t l_Lean_Expr_isLambda(lean_object*); static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_addPreDefinitions___spec__8___rarg___closed__2; lean_object* l_Lean_addTrace___at___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTerm___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__11; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__2; static lean_object* l_Lean_Elab_addPreDefinitions___closed__5; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__1; @@ -293,7 +290,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_addPreDefinitions lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_addPreDefinitions___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_addPreDefinitions___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1(lean_object*, lean_object*); lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_SCC_0__Lean_SCC_addSCC_add___at___private_Lean_Elab_PreDefinition_Main_0__Lean_Elab_partitionPreDefs___spec__21(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_logException___at_Lean_Elab_Term_exceptionToSorry___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -2637,7 +2633,7 @@ lean_inc(x_13); x_14 = lean_ctor_get(x_11, 0); lean_inc(x_14); lean_dec(x_11); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1(x_13, x_14); lean_dec(x_14); lean_dec(x_13); if (x_15 == 0) @@ -2670,7 +2666,7 @@ lean_inc(x_20); x_21 = lean_ctor_get(x_18, 0); lean_inc(x_21); lean_dec(x_18); -x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1(x_20, x_21); +x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1(x_20, x_21); lean_dec(x_21); lean_dec(x_20); if (x_22 == 0) @@ -6819,362 +6815,455 @@ return x_70; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t 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, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { -if (lean_obj_tag(x_1) == 0) +lean_object* x_16; +if (lean_obj_tag(x_2) == 0) { if (lean_obj_tag(x_6) == 0) { -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_16 = l_Lean_Elab_instInhabitedPreDefinition; -x_17 = lean_unsigned_to_nat(0u); -x_18 = lean_array_get(x_16, x_2, x_17); -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -lean_dec(x_18); -lean_inc(x_2); -x_20 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__1), 2, 1); -lean_closure_set(x_20, 0, x_2); -x_21 = !lean_is_exclusive(x_13); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_22 = lean_ctor_get(x_13, 5); -x_23 = l_Lean_replaceRef(x_19, x_22); -lean_dec(x_22); -lean_dec(x_19); -lean_ctor_set(x_13, 5, x_23); -x_24 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__2), 9, 4); -lean_closure_set(x_24, 0, x_2); -lean_closure_set(x_24, 1, x_9); -lean_closure_set(x_24, 2, x_10); -lean_closure_set(x_24, 3, x_3); -x_25 = l_Lean_Meta_mapErrorImp___rarg(x_24, x_20, x_11, x_12, x_13, x_14, x_15); -if (lean_obj_tag(x_25) == 0) -{ -uint8_t x_26; -x_26 = !lean_is_exclusive(x_25); -if (x_26 == 0) +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; uint8_t x_75; +x_70 = l_Lean_Elab_instInhabitedPreDefinition; +x_71 = lean_unsigned_to_nat(0u); +x_72 = lean_array_get(x_70, x_1, x_71); +x_73 = lean_ctor_get(x_72, 0); +lean_inc(x_73); +lean_dec(x_72); +lean_inc(x_1); +x_74 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__1), 2, 1); +lean_closure_set(x_74, 0, x_1); +x_75 = !lean_is_exclusive(x_13); +if (x_75 == 0) +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_76 = lean_ctor_get(x_13, 5); +x_77 = l_Lean_replaceRef(x_73, x_76); +lean_dec(x_76); +lean_dec(x_73); +lean_ctor_set(x_13, 5, x_77); +x_78 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__2), 9, 4); +lean_closure_set(x_78, 0, x_1); +lean_closure_set(x_78, 1, x_9); +lean_closure_set(x_78, 2, x_10); +lean_closure_set(x_78, 3, x_3); +x_79 = l_Lean_Meta_mapErrorImp___rarg(x_78, x_74, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_79) == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_27 = lean_ctor_get(x_25, 0); -x_28 = lean_box(x_4); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_5); -x_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_7); -lean_ctor_set(x_30, 1, x_29); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_27); -lean_ctor_set(x_31, 1, x_30); -lean_ctor_set(x_25, 0, x_31); -return x_25; -} -else +uint8_t x_80; +x_80 = !lean_is_exclusive(x_79); +if (x_80 == 0) { -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; -x_32 = lean_ctor_get(x_25, 0); -x_33 = lean_ctor_get(x_25, 1); -lean_inc(x_33); -lean_inc(x_32); -lean_dec(x_25); -x_34 = lean_box(x_4); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_5); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_7); -lean_ctor_set(x_36, 1, x_35); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_32); -lean_ctor_set(x_37, 1, x_36); -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_37); -lean_ctor_set(x_38, 1, x_33); -return x_38; +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_81 = lean_ctor_get(x_79, 0); +x_82 = lean_box(x_4); +x_83 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_83, 0, x_82); +lean_ctor_set(x_83, 1, x_5); +x_84 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_84, 0, x_7); +lean_ctor_set(x_84, 1, x_83); +x_85 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_85, 0, x_81); +lean_ctor_set(x_85, 1, x_84); +lean_ctor_set(x_79, 0, x_85); +return x_79; +} +else +{ +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_86 = lean_ctor_get(x_79, 0); +x_87 = lean_ctor_get(x_79, 1); +lean_inc(x_87); +lean_inc(x_86); +lean_dec(x_79); +x_88 = lean_box(x_4); +x_89 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_89, 0, x_88); +lean_ctor_set(x_89, 1, x_5); +x_90 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_90, 0, x_7); +lean_ctor_set(x_90, 1, x_89); +x_91 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_91, 0, x_86); +lean_ctor_set(x_91, 1, x_90); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_91); +lean_ctor_set(x_92, 1, x_87); +return x_92; } } else { -uint8_t x_39; +uint8_t x_93; lean_dec(x_7); lean_dec(x_5); -x_39 = !lean_is_exclusive(x_25); -if (x_39 == 0) +x_93 = !lean_is_exclusive(x_79); +if (x_93 == 0) { -return x_25; +return x_79; } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_25, 0); -x_41 = lean_ctor_get(x_25, 1); -lean_inc(x_41); -lean_inc(x_40); -lean_dec(x_25); -x_42 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_42, 0, x_40); -lean_ctor_set(x_42, 1, x_41); -return x_42; +lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_94 = lean_ctor_get(x_79, 0); +x_95 = lean_ctor_get(x_79, 1); +lean_inc(x_95); +lean_inc(x_94); +lean_dec(x_79); +x_96 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_96, 0, x_94); +lean_ctor_set(x_96, 1, x_95); +return x_96; } } } else { -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_43 = lean_ctor_get(x_13, 0); -x_44 = lean_ctor_get(x_13, 1); -x_45 = lean_ctor_get(x_13, 2); -x_46 = lean_ctor_get(x_13, 3); -x_47 = lean_ctor_get(x_13, 4); -x_48 = lean_ctor_get(x_13, 5); -x_49 = lean_ctor_get(x_13, 6); -x_50 = lean_ctor_get(x_13, 7); -x_51 = lean_ctor_get(x_13, 8); -x_52 = lean_ctor_get(x_13, 9); -x_53 = lean_ctor_get(x_13, 10); -lean_inc(x_53); -lean_inc(x_52); -lean_inc(x_51); -lean_inc(x_50); -lean_inc(x_49); -lean_inc(x_48); -lean_inc(x_47); -lean_inc(x_46); -lean_inc(x_45); -lean_inc(x_44); -lean_inc(x_43); +lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; +x_97 = lean_ctor_get(x_13, 0); +x_98 = lean_ctor_get(x_13, 1); +x_99 = lean_ctor_get(x_13, 2); +x_100 = lean_ctor_get(x_13, 3); +x_101 = lean_ctor_get(x_13, 4); +x_102 = lean_ctor_get(x_13, 5); +x_103 = lean_ctor_get(x_13, 6); +x_104 = lean_ctor_get(x_13, 7); +x_105 = lean_ctor_get(x_13, 8); +x_106 = lean_ctor_get(x_13, 9); +x_107 = lean_ctor_get(x_13, 10); +lean_inc(x_107); +lean_inc(x_106); +lean_inc(x_105); +lean_inc(x_104); +lean_inc(x_103); +lean_inc(x_102); +lean_inc(x_101); +lean_inc(x_100); +lean_inc(x_99); +lean_inc(x_98); +lean_inc(x_97); lean_dec(x_13); -x_54 = l_Lean_replaceRef(x_19, x_48); -lean_dec(x_48); -lean_dec(x_19); -x_55 = lean_alloc_ctor(0, 11, 0); -lean_ctor_set(x_55, 0, x_43); -lean_ctor_set(x_55, 1, x_44); -lean_ctor_set(x_55, 2, x_45); -lean_ctor_set(x_55, 3, x_46); -lean_ctor_set(x_55, 4, x_47); -lean_ctor_set(x_55, 5, x_54); -lean_ctor_set(x_55, 6, x_49); -lean_ctor_set(x_55, 7, x_50); -lean_ctor_set(x_55, 8, x_51); -lean_ctor_set(x_55, 9, x_52); -lean_ctor_set(x_55, 10, x_53); -x_56 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__2), 9, 4); -lean_closure_set(x_56, 0, x_2); -lean_closure_set(x_56, 1, x_9); -lean_closure_set(x_56, 2, x_10); -lean_closure_set(x_56, 3, x_3); -x_57 = l_Lean_Meta_mapErrorImp___rarg(x_56, x_20, x_11, x_12, x_55, x_14, x_15); -if (lean_obj_tag(x_57) == 0) -{ -lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_58 = lean_ctor_get(x_57, 0); -lean_inc(x_58); -x_59 = lean_ctor_get(x_57, 1); -lean_inc(x_59); -if (lean_is_exclusive(x_57)) { - lean_ctor_release(x_57, 0); - lean_ctor_release(x_57, 1); - x_60 = x_57; +x_108 = l_Lean_replaceRef(x_73, x_102); +lean_dec(x_102); +lean_dec(x_73); +x_109 = lean_alloc_ctor(0, 11, 0); +lean_ctor_set(x_109, 0, x_97); +lean_ctor_set(x_109, 1, x_98); +lean_ctor_set(x_109, 2, x_99); +lean_ctor_set(x_109, 3, x_100); +lean_ctor_set(x_109, 4, x_101); +lean_ctor_set(x_109, 5, x_108); +lean_ctor_set(x_109, 6, x_103); +lean_ctor_set(x_109, 7, x_104); +lean_ctor_set(x_109, 8, x_105); +lean_ctor_set(x_109, 9, x_106); +lean_ctor_set(x_109, 10, x_107); +x_110 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__2), 9, 4); +lean_closure_set(x_110, 0, x_1); +lean_closure_set(x_110, 1, x_9); +lean_closure_set(x_110, 2, x_10); +lean_closure_set(x_110, 3, x_3); +x_111 = l_Lean_Meta_mapErrorImp___rarg(x_110, x_74, x_11, x_12, x_109, x_14, x_15); +if (lean_obj_tag(x_111) == 0) +{ +lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; +x_112 = lean_ctor_get(x_111, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_111, 1); +lean_inc(x_113); +if (lean_is_exclusive(x_111)) { + lean_ctor_release(x_111, 0); + lean_ctor_release(x_111, 1); + x_114 = x_111; } else { - lean_dec_ref(x_57); - x_60 = lean_box(0); -} -x_61 = lean_box(x_4); -x_62 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_62, 0, x_61); -lean_ctor_set(x_62, 1, x_5); -x_63 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_63, 0, x_7); -lean_ctor_set(x_63, 1, x_62); -x_64 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_64, 0, x_58); -lean_ctor_set(x_64, 1, x_63); -if (lean_is_scalar(x_60)) { - x_65 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_111); + x_114 = lean_box(0); +} +x_115 = lean_box(x_4); +x_116 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_116, 0, x_115); +lean_ctor_set(x_116, 1, x_5); +x_117 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_117, 0, x_7); +lean_ctor_set(x_117, 1, x_116); +x_118 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_118, 0, x_112); +lean_ctor_set(x_118, 1, x_117); +if (lean_is_scalar(x_114)) { + x_119 = lean_alloc_ctor(0, 2, 0); } else { - x_65 = x_60; + x_119 = x_114; } -lean_ctor_set(x_65, 0, x_64); -lean_ctor_set(x_65, 1, x_59); -return x_65; +lean_ctor_set(x_119, 0, x_118); +lean_ctor_set(x_119, 1, x_113); +return x_119; } else { -lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_dec(x_7); lean_dec(x_5); -x_66 = lean_ctor_get(x_57, 0); -lean_inc(x_66); -x_67 = lean_ctor_get(x_57, 1); -lean_inc(x_67); -if (lean_is_exclusive(x_57)) { - lean_ctor_release(x_57, 0); - lean_ctor_release(x_57, 1); - x_68 = x_57; +x_120 = lean_ctor_get(x_111, 0); +lean_inc(x_120); +x_121 = lean_ctor_get(x_111, 1); +lean_inc(x_121); +if (lean_is_exclusive(x_111)) { + lean_ctor_release(x_111, 0); + lean_ctor_release(x_111, 1); + x_122 = x_111; } else { - lean_dec_ref(x_57); - x_68 = lean_box(0); + lean_dec_ref(x_111); + x_122 = lean_box(0); } -if (lean_is_scalar(x_68)) { - x_69 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_122)) { + x_123 = lean_alloc_ctor(1, 2, 0); } else { - x_69 = x_68; + x_123 = x_122; } -lean_ctor_set(x_69, 0, x_66); -lean_ctor_set(x_69, 1, x_67); -return x_69; +lean_ctor_set(x_123, 0, x_120); +lean_ctor_set(x_123, 1, x_121); +return x_123; } } } else { -lean_object* x_70; -lean_dec(x_3); -x_70 = l_Lean_Elab_wfRecursion(x_2, x_1, x_6, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -if (lean_obj_tag(x_70) == 0) -{ -uint8_t x_71; -x_71 = !lean_is_exclusive(x_70); -if (x_71 == 0) -{ -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_72 = lean_ctor_get(x_70, 0); -x_73 = lean_box(x_4); -x_74 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_74, 0, x_73); -lean_ctor_set(x_74, 1, x_5); -x_75 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_75, 0, x_7); -lean_ctor_set(x_75, 1, x_74); -x_76 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_76, 0, x_72); -lean_ctor_set(x_76, 1, x_75); -lean_ctor_set(x_70, 0, x_76); -return x_70; +lean_object* x_124; +x_124 = lean_box(0); +x_16 = x_124; +goto block_69; +} } else { -lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; -x_77 = lean_ctor_get(x_70, 0); -x_78 = lean_ctor_get(x_70, 1); -lean_inc(x_78); -lean_inc(x_77); -lean_dec(x_70); -x_79 = lean_box(x_4); -x_80 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_80, 0, x_79); -lean_ctor_set(x_80, 1, x_5); -x_81 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_81, 0, x_7); -lean_ctor_set(x_81, 1, x_80); -x_82 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_82, 0, x_77); -lean_ctor_set(x_82, 1, x_81); -x_83 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_83, 0, x_82); -lean_ctor_set(x_83, 1, x_78); -return x_83; +lean_object* x_125; +x_125 = lean_box(0); +x_16 = x_125; +goto block_69; +} +block_69: +{ +lean_dec(x_16); +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_17; +x_17 = l_Lean_Elab_wfRecursion(x_1, x_2, x_3, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_17) == 0) +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_19 = lean_ctor_get(x_17, 0); +x_20 = lean_box(x_4); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_5); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_7); +lean_ctor_set(x_22, 1, x_21); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_19); +lean_ctor_set(x_23, 1, x_22); +lean_ctor_set(x_17, 0, x_23); +return x_17; +} +else +{ +lean_object* 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; +x_24 = lean_ctor_get(x_17, 0); +x_25 = lean_ctor_get(x_17, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_17); +x_26 = lean_box(x_4); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_5); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_7); +lean_ctor_set(x_28, 1, x_27); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_24); +lean_ctor_set(x_29, 1, x_28); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_25); +return x_30; } } else { -uint8_t x_84; +uint8_t x_31; lean_dec(x_7); lean_dec(x_5); -x_84 = !lean_is_exclusive(x_70); -if (x_84 == 0) +x_31 = !lean_is_exclusive(x_17); +if (x_31 == 0) { -return x_70; +return x_17; } else { -lean_object* x_85; lean_object* x_86; lean_object* x_87; -x_85 = lean_ctor_get(x_70, 0); -x_86 = lean_ctor_get(x_70, 1); -lean_inc(x_86); -lean_inc(x_85); -lean_dec(x_70); -x_87 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_87, 0, x_85); -lean_ctor_set(x_87, 1, x_86); -return x_87; -} +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_17, 0); +x_33 = lean_ctor_get(x_17, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_17); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; } } } else { -lean_object* x_88; +uint8_t x_35; lean_dec(x_3); -x_88 = l_Lean_Elab_wfRecursion(x_2, x_1, x_6, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -if (lean_obj_tag(x_88) == 0) +x_35 = !lean_is_exclusive(x_6); +if (x_35 == 0) { -uint8_t x_89; -x_89 = !lean_is_exclusive(x_88); -if (x_89 == 0) +lean_object* x_36; +x_36 = l_Lean_Elab_wfRecursion(x_1, x_2, x_6, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_36) == 0) { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; -x_90 = lean_ctor_get(x_88, 0); -x_91 = lean_box(x_4); -x_92 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_92, 0, x_91); -lean_ctor_set(x_92, 1, x_5); -x_93 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_93, 0, x_7); -lean_ctor_set(x_93, 1, x_92); -x_94 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_94, 0, x_90); -lean_ctor_set(x_94, 1, x_93); -lean_ctor_set(x_88, 0, x_94); -return x_88; +uint8_t x_37; +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_38 = lean_ctor_get(x_36, 0); +x_39 = lean_box(x_4); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_5); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_7); +lean_ctor_set(x_41, 1, x_40); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_38); +lean_ctor_set(x_42, 1, x_41); +lean_ctor_set(x_36, 0, x_42); +return x_36; } else { -lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; -x_95 = lean_ctor_get(x_88, 0); -x_96 = lean_ctor_get(x_88, 1); -lean_inc(x_96); -lean_inc(x_95); -lean_dec(x_88); -x_97 = lean_box(x_4); -x_98 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_98, 0, x_97); -lean_ctor_set(x_98, 1, x_5); -x_99 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_99, 0, x_7); -lean_ctor_set(x_99, 1, x_98); -x_100 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_100, 0, x_95); -lean_ctor_set(x_100, 1, x_99); -x_101 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_101, 0, x_100); -lean_ctor_set(x_101, 1, x_96); -return x_101; +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_43 = lean_ctor_get(x_36, 0); +x_44 = lean_ctor_get(x_36, 1); +lean_inc(x_44); +lean_inc(x_43); +lean_dec(x_36); +x_45 = lean_box(x_4); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_5); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_7); +lean_ctor_set(x_47, 1, x_46); +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_43); +lean_ctor_set(x_48, 1, x_47); +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_44); +return x_49; } } else { -uint8_t x_102; +uint8_t x_50; lean_dec(x_7); lean_dec(x_5); -x_102 = !lean_is_exclusive(x_88); -if (x_102 == 0) +x_50 = !lean_is_exclusive(x_36); +if (x_50 == 0) { -return x_88; +return x_36; } else { -lean_object* x_103; lean_object* x_104; lean_object* x_105; -x_103 = lean_ctor_get(x_88, 0); -x_104 = lean_ctor_get(x_88, 1); -lean_inc(x_104); -lean_inc(x_103); -lean_dec(x_88); -x_105 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_105, 0, x_103); -lean_ctor_set(x_105, 1, x_104); -return x_105; +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_36, 0); +x_52 = lean_ctor_get(x_36, 1); +lean_inc(x_52); +lean_inc(x_51); +lean_dec(x_36); +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_52); +return x_53; +} +} +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_ctor_get(x_6, 0); +lean_inc(x_54); +lean_dec(x_6); +x_55 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_55, 0, x_54); +x_56 = l_Lean_Elab_wfRecursion(x_1, x_2, x_55, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_56) == 0) +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_56, 1); +lean_inc(x_58); +if (lean_is_exclusive(x_56)) { + lean_ctor_release(x_56, 0); + lean_ctor_release(x_56, 1); + x_59 = x_56; +} else { + lean_dec_ref(x_56); + x_59 = lean_box(0); +} +x_60 = lean_box(x_4); +x_61 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set(x_61, 1, x_5); +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_7); +lean_ctor_set(x_62, 1, x_61); +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_57); +lean_ctor_set(x_63, 1, x_62); +if (lean_is_scalar(x_59)) { + x_64 = lean_alloc_ctor(0, 2, 0); +} else { + x_64 = x_59; +} +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set(x_64, 1, x_58); +return x_64; +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +lean_dec(x_7); +lean_dec(x_5); +x_65 = lean_ctor_get(x_56, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_56, 1); +lean_inc(x_66); +if (lean_is_exclusive(x_56)) { + lean_ctor_release(x_56, 0); + lean_ctor_release(x_56, 1); + x_67 = x_56; +} else { + lean_dec_ref(x_56); + x_67 = lean_box(0); +} +if (lean_is_scalar(x_67)) { + x_68 = lean_alloc_ctor(1, 2, 0); +} else { + x_68 = x_67; +} +lean_ctor_set(x_68, 0, x_65); +lean_ctor_set(x_68, 1, x_66); +return x_68; +} } } } @@ -7263,42 +7352,6 @@ return x_1; static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__10() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Tactic", 6); -return x_1; -} -} -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__11() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__4; -x_2 = l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__10; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__12() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("tacticSeq", 9); -return x_1; -} -} -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__13() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__11; -x_2 = l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__12; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__14() { -_start: -{ lean_object* x_1; lean_object* x_2; x_1 = lean_unsigned_to_nat(2u); x_2 = lean_mk_empty_array_with_capacity(x_1); @@ -7320,7 +7373,7 @@ if (lean_obj_tag(x_20) == 0) lean_object* x_21; lean_object* x_22; lean_dec(x_19); x_21 = lean_box(0); -x_22 = l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__3(x_9, x_1, x_2, x_3, x_8, x_6, x_7, x_21, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +x_22 = l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__3(x_1, x_9, x_2, x_3, x_8, x_6, x_7, x_21, x_11, x_12, x_13, x_14, x_15, x_16, x_17); return x_22; } else @@ -7330,7 +7383,7 @@ lean_dec(x_6); x_23 = !lean_is_exclusive(x_20); if (x_23 == 0) { -lean_object* 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +lean_object* 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; x_24 = lean_ctor_get(x_20, 0); x_25 = lean_ctor_get(x_24, 0); lean_inc(x_25); @@ -7388,114 +7441,100 @@ x_45 = l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___la x_46 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_46, 0, x_41); lean_ctor_set(x_46, 1, x_45); -x_47 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_PreDefinition_Main_0__Lean_Elab_addAsAxioms___spec__1___closed__1; -x_48 = lean_array_push(x_47, x_26); -x_49 = lean_box(2); -x_50 = l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__13; -x_51 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_51, 0, x_49); -lean_ctor_set(x_51, 1, x_50); -lean_ctor_set(x_51, 2, x_48); -x_52 = l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__14; -x_53 = lean_array_push(x_52, x_46); -x_54 = lean_array_push(x_53, x_51); -x_55 = l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__8; -x_56 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_56, 0, x_49); -lean_ctor_set(x_56, 1, x_55); -lean_ctor_set(x_56, 2, x_54); -lean_ctor_set(x_20, 0, x_56); -x_57 = l_Lean_Elab_WF_TerminationHint_markAsUsed(x_7, x_19); +x_47 = l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__10; +x_48 = lean_array_push(x_47, x_46); +x_49 = lean_array_push(x_48, x_26); +x_50 = lean_box(2); +x_51 = l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__8; +x_52 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_52, 0, x_50); +lean_ctor_set(x_52, 1, x_51); +lean_ctor_set(x_52, 2, x_49); +lean_ctor_set(x_20, 0, x_52); +x_53 = l_Lean_Elab_WF_TerminationHint_markAsUsed(x_7, x_19); lean_dec(x_19); -x_58 = lean_box(0); -x_59 = l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__3(x_9, x_1, x_2, x_3, x_8, x_20, x_57, x_58, x_11, x_12, x_13, x_14, x_15, x_16, x_44); -return x_59; +x_54 = lean_box(0); +x_55 = l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__3(x_1, x_9, x_2, x_3, x_8, x_20, x_53, x_54, x_11, x_12, x_13, x_14, x_15, x_16, x_44); +return x_55; } 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; -x_60 = lean_ctor_get(x_20, 0); -lean_inc(x_60); +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_56 = lean_ctor_get(x_20, 0); +lean_inc(x_56); lean_dec(x_20); -x_61 = lean_ctor_get(x_60, 0); +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_56, 1); +lean_inc(x_58); +lean_dec(x_56); +x_59 = lean_ctor_get(x_15, 0); +lean_inc(x_59); +x_60 = lean_ctor_get(x_15, 1); +lean_inc(x_60); +x_61 = lean_ctor_get(x_15, 2); lean_inc(x_61); -x_62 = lean_ctor_get(x_60, 1); +x_62 = lean_ctor_get(x_15, 3); lean_inc(x_62); -lean_dec(x_60); -x_63 = lean_ctor_get(x_15, 0); +x_63 = lean_ctor_get(x_15, 4); lean_inc(x_63); -x_64 = lean_ctor_get(x_15, 1); +x_64 = lean_ctor_get(x_15, 5); lean_inc(x_64); -x_65 = lean_ctor_get(x_15, 2); +x_65 = lean_ctor_get(x_15, 6); lean_inc(x_65); -x_66 = lean_ctor_get(x_15, 3); +x_66 = lean_ctor_get(x_15, 7); lean_inc(x_66); -x_67 = lean_ctor_get(x_15, 4); +x_67 = lean_ctor_get(x_15, 8); lean_inc(x_67); -x_68 = lean_ctor_get(x_15, 5); +x_68 = lean_ctor_get(x_15, 9); lean_inc(x_68); -x_69 = lean_ctor_get(x_15, 6); +x_69 = lean_ctor_get(x_15, 10); lean_inc(x_69); -x_70 = lean_ctor_get(x_15, 7); -lean_inc(x_70); -x_71 = lean_ctor_get(x_15, 8); -lean_inc(x_71); -x_72 = lean_ctor_get(x_15, 9); -lean_inc(x_72); -x_73 = lean_ctor_get(x_15, 10); +x_70 = l_Lean_replaceRef(x_57, x_64); +lean_dec(x_64); +lean_dec(x_57); +x_71 = lean_alloc_ctor(0, 11, 0); +lean_ctor_set(x_71, 0, x_59); +lean_ctor_set(x_71, 1, x_60); +lean_ctor_set(x_71, 2, x_61); +lean_ctor_set(x_71, 3, x_62); +lean_ctor_set(x_71, 4, x_63); +lean_ctor_set(x_71, 5, x_70); +lean_ctor_set(x_71, 6, x_65); +lean_ctor_set(x_71, 7, x_66); +lean_ctor_set(x_71, 8, x_67); +lean_ctor_set(x_71, 9, x_68); +lean_ctor_set(x_71, 10, x_69); +x_72 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_71, x_16, x_17); +x_73 = lean_ctor_get(x_72, 0); lean_inc(x_73); -x_74 = l_Lean_replaceRef(x_61, x_68); -lean_dec(x_68); -lean_dec(x_61); -x_75 = lean_alloc_ctor(0, 11, 0); -lean_ctor_set(x_75, 0, x_63); -lean_ctor_set(x_75, 1, x_64); -lean_ctor_set(x_75, 2, x_65); -lean_ctor_set(x_75, 3, x_66); -lean_ctor_set(x_75, 4, x_67); -lean_ctor_set(x_75, 5, x_74); -lean_ctor_set(x_75, 6, x_69); -lean_ctor_set(x_75, 7, x_70); -lean_ctor_set(x_75, 8, x_71); -lean_ctor_set(x_75, 9, x_72); -lean_ctor_set(x_75, 10, x_73); -x_76 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_75, x_16, x_17); -x_77 = lean_ctor_get(x_76, 0); -lean_inc(x_77); -x_78 = lean_ctor_get(x_76, 1); -lean_inc(x_78); -lean_dec(x_76); -x_79 = lean_st_ref_get(x_16, x_78); -x_80 = lean_ctor_get(x_79, 1); -lean_inc(x_80); -lean_dec(x_79); -x_81 = l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__9; -x_82 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_82, 0, x_77); -lean_ctor_set(x_82, 1, x_81); -x_83 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_PreDefinition_Main_0__Lean_Elab_addAsAxioms___spec__1___closed__1; -x_84 = lean_array_push(x_83, x_62); -x_85 = lean_box(2); -x_86 = l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__13; -x_87 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_87, 0, x_85); -lean_ctor_set(x_87, 1, x_86); -lean_ctor_set(x_87, 2, x_84); -x_88 = l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__14; -x_89 = lean_array_push(x_88, x_82); -x_90 = lean_array_push(x_89, x_87); -x_91 = l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__8; -x_92 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_92, 0, x_85); -lean_ctor_set(x_92, 1, x_91); -lean_ctor_set(x_92, 2, x_90); -x_93 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_93, 0, x_92); -x_94 = l_Lean_Elab_WF_TerminationHint_markAsUsed(x_7, x_19); +x_74 = lean_ctor_get(x_72, 1); +lean_inc(x_74); +lean_dec(x_72); +x_75 = lean_st_ref_get(x_16, x_74); +x_76 = lean_ctor_get(x_75, 1); +lean_inc(x_76); +lean_dec(x_75); +x_77 = l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__9; +x_78 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_78, 0, x_73); +lean_ctor_set(x_78, 1, x_77); +x_79 = l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__10; +x_80 = lean_array_push(x_79, x_78); +x_81 = lean_array_push(x_80, x_58); +x_82 = lean_box(2); +x_83 = l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__8; +x_84 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_84, 0, x_82); +lean_ctor_set(x_84, 1, x_83); +lean_ctor_set(x_84, 2, x_81); +x_85 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_85, 0, x_84); +x_86 = l_Lean_Elab_WF_TerminationHint_markAsUsed(x_7, x_19); lean_dec(x_19); -x_95 = lean_box(0); -x_96 = l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__3(x_9, x_1, x_2, x_3, x_8, x_93, x_94, x_95, x_11, x_12, x_13, x_14, x_15, x_16, x_80); -return x_96; +x_87 = lean_box(0); +x_88 = l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__3(x_1, x_9, x_2, x_3, x_8, x_85, x_86, x_87, x_11, x_12, x_13, x_14, x_15, x_16, x_76); +return x_88; } } } @@ -10469,7 +10508,7 @@ x_16 = l_Lean_Elab_addPreDefinitions___lambda__1(x_1, x_14, x_15, x_4, x_5, x_6, return x_16; } } -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_Main___hyg_2378_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_Main___hyg_2382_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; @@ -10650,14 +10689,6 @@ l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4 lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__9); l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__10 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__10(); lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__10); -l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__11 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__11(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__11); -l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__12 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__12(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__12); -l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__13 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__13(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__13); -l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__14 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__14(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___lambda__4___closed__14); l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___closed__1(); lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___closed__1); l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_addPreDefinitions___spec__21___closed__2(); @@ -10690,7 +10721,7 @@ l_Lean_Elab_addPreDefinitions___closed__13 = _init_l_Lean_Elab_addPreDefinitions lean_mark_persistent(l_Lean_Elab_addPreDefinitions___closed__13); l_Lean_Elab_addPreDefinitions___boxed__const__1 = _init_l_Lean_Elab_addPreDefinitions___boxed__const__1(); lean_mark_persistent(l_Lean_Elab_addPreDefinitions___boxed__const__1); -res = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_Main___hyg_2378_(lean_io_mk_world()); +res = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_Main___hyg_2382_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/Structural/FindRecArg.c b/stage0/stdlib/Lean/Elab/PreDefinition/Structural/FindRecArg.c index b6481d2b792..440771970ca 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/Structural/FindRecArg.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/Structural/FindRecArg.c @@ -110,7 +110,6 @@ LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Structural_findRecArg_go LEAN_EXPORT lean_object* l_Lean_hasConst___at_Lean_Elab_Structural_findRecArg_go___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Structural_findRecArg_go___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Array_forInUnsafe_loop___at_Lean_Elab_Structural_findRecArg___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_panic___at_String_toNat_x21___spec__1(lean_object*); static lean_object* l_Lean_Elab_Structural_findRecArg_go___rarg___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_PreDefinition_Structural_FindRecArg_0__Lean_Elab_Structural_hasBadParamDep_x3f___spec__2(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); @@ -163,6 +162,7 @@ lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isFVar(lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_panic___at_Lean_TSyntax_getNat___spec__1(lean_object*); lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Structural_findRecArg_go___rarg___closed__1; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Structural_findRecArg_go___spec__9(lean_object*, size_t, size_t); @@ -1797,7 +1797,7 @@ if (lean_obj_tag(x_9) == 0) { lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = l_Array_mapMUnsafe_map___at_Lean_Elab_Structural_findRecArg_go___spec__2___closed__4; -x_13 = l_panic___at_String_toNat_x21___spec__1(x_12); +x_13 = l_panic___at_Lean_TSyntax_getNat___spec__1(x_12); x_14 = lean_array_uset(x_8, x_3, x_13); x_3 = x_11; x_4 = x_14; diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/WF/Rel.c b/stage0/stdlib/Lean/Elab/PreDefinition/WF/Rel.c index 395f5c47d5c..30753927b71 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/WF/Rel.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/WF/Rel.c @@ -97,7 +97,6 @@ lean_object* l_Lean_Meta_rename(lean_object*, lean_object*, lean_object*, lean_o lean_object* l_Lean_Meta_getMVarType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_getNumCandidateArgs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_guess___spec__5___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_numLitKind; static lean_object* l_Lean_Elab_WF_elabWFRel_go___rarg___lambda__1___closed__4; lean_object* lean_nat_sub(lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_elabWFRel_generateElements___closed__4; @@ -116,7 +115,7 @@ static lean_object* l___private_Lean_Elab_PreDefinition_WF_Rel_0__Lean_Elab_WF_u LEAN_EXPORT lean_object* l_panic___at___private_Lean_Elab_PreDefinition_WF_Rel_0__Lean_Elab_WF_unpackMutual_go___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_elabWFRel_go___rarg___lambda__1___closed__1; static lean_object* l_Lean_Elab_WF_elabWFRel_guess___rarg___lambda__1___closed__2; -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__1(lean_object*, lean_object*, lean_object*, 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_Std_Range_forIn_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Rel_0__Lean_Elab_WF_unpackUnary_go___closed__12; lean_object* l_Nat_repr(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_PreDefinition_WF_Rel_0__Lean_Elab_WF_unpackUnary___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -131,7 +130,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_WF_elabWFRel_guess___rarg___lambda__1___box lean_object* lean_array_to_list(lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_elabWFRel_guess___rarg___lambda__1___closed__1; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Rel_0__Lean_Elab_WF_unpackUnary_go___closed__10; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_instInhabitedExpr; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_PreDefinition_WF_Rel_0__Lean_Elab_WF_unpackUnary___spec__6___boxed(lean_object*, lean_object*, lean_object*, 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___private_Lean_Elab_PreDefinition_WF_Rel_0__Lean_Elab_WF_getRefFromElems___lambda__1(lean_object*, lean_object*); @@ -162,6 +161,7 @@ lean_object* l_Lean_addTrace___at___private_Lean_Elab_Term_0__Lean_Elab_Term_pos lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalDecl_fvarId(lean_object*); static lean_object* l_panic___at___private_Lean_Elab_PreDefinition_WF_Rel_0__Lean_Elab_WF_unpackMutual_go___spec__1___closed__1; +lean_object* l_Lean_Syntax_mkNumLit(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_WF_elabWFRel_go___spec__2(lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Rel_0__Lean_Elab_WF_unpackUnary_go___lambda__1___closed__2; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__7; @@ -230,6 +230,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Rel_0__Lean_Elab static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_WF_getForbiddenByTrivialSizeOf___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Rel_0__Lean_Elab_WF_unpackUnary___lambda__2___closed__1; +static lean_object* l_Lean_Elab_WF_elabWFRel_generateElements___closed__16; lean_object* l_Lean_Meta_whnfD(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_observing_x3f___at_Lean_Elab_WF_elabWFRel_guess___spec__3(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_go___spec__4___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -238,6 +239,7 @@ static lean_object* l___private_Lean_Elab_PreDefinition_WF_Rel_0__Lean_Elab_WF_u LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_elabWFRel_guess___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Subarray_forInUnsafe_loop___at_Lean_Elab_WF_getForbiddenByTrivialSizeOf___spec__1___closed__4; lean_object* l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_WF_elabWFRel_generateElements___closed__14; lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_lambdaTelescopeImp___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_elabWFRel_generateElements___closed__5; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Rel_0__Lean_Elab_WF_unpackMutual_go___closed__3; @@ -253,16 +255,16 @@ lean_object* l_Lean_mkConst(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_PreDefinition_WF_Rel_0__Lean_Elab_WF_getRefFromElems___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Rel_0__Lean_Elab_WF_unpackUnary_go___closed__6; lean_object* l_Lean_Elab_Term_saveState___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_WF_elabWFRel_generateElements___closed__15; LEAN_EXPORT lean_object* l_Lean_Elab_WF_elabWFRel_guess___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Elab_WF_generateCombinations_x3f_isForbidden(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at_Lean_Elab_WF_getForbiddenByTrivialSizeOf___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_elabWFRel_generateElements___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_WF_generateCombinations_x3f_isForbidden___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__1___boxed(lean_object*, lean_object*, lean_object*, 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_Std_Range_forIn_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Rel_0__Lean_Elab_WF_unpackUnary_go___closed__1; static lean_object* l_Subarray_forInUnsafe_loop___at_Lean_Elab_WF_getForbiddenByTrivialSizeOf___spec__1___closed__2; lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); -lean_object* l_Lean_Syntax_mkLit(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Rel_0__Lean_Elab_WF_unpackUnary___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); uint8_t l_Lean_Syntax_isIdent(lean_object*); @@ -6228,53 +6230,53 @@ lean_dec(x_3); return x_9; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__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, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__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, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -uint8_t x_14; -x_14 = lean_nat_dec_le(x_4, x_3); -if (x_14 == 0) +uint8_t x_15; +x_15 = lean_nat_dec_le(x_5, x_4); +if (x_15 == 0) { -lean_object* x_15; uint8_t x_16; -x_15 = lean_unsigned_to_nat(0u); -x_16 = lean_nat_dec_eq(x_2, x_15); -if (x_16 == 0) +lean_object* x_16; uint8_t x_17; +x_16 = lean_unsigned_to_nat(0u); +x_17 = lean_nat_dec_eq(x_3, x_16); +if (x_17 == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_17 = lean_unsigned_to_nat(1u); -x_18 = lean_nat_sub(x_2, x_17); -lean_dec(x_2); -lean_inc(x_1); -x_19 = lean_array_push(x_6, x_1); -x_20 = lean_nat_add(x_3, x_5); +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_18 = lean_unsigned_to_nat(1u); +x_19 = lean_nat_sub(x_3, x_18); lean_dec(x_3); -x_2 = x_18; -x_3 = x_20; -x_6 = x_19; +lean_inc(x_2); +x_20 = lean_array_push(x_7, x_2); +x_21 = lean_nat_add(x_4, x_6); +lean_dec(x_4); +x_3 = x_19; +x_4 = x_21; +x_7 = x_20; goto _start; } else { -lean_object* x_22; +lean_object* x_23; +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_6); -lean_ctor_set(x_22, 1, x_13); -return x_22; +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_7); +lean_ctor_set(x_23, 1, x_14); +return x_23; } } else { -lean_object* x_23; +lean_object* x_24; +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_6); -lean_ctor_set(x_23, 1, x_13); -return x_23; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_7); +lean_ctor_set(x_24, 1, x_14); +return x_24; } } } @@ -6426,15 +6428,16 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___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_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, size_t x_12, size_t 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, lean_object* x_20, lean_object* x_21) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___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_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, size_t x_13, size_t 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) { _start: { -uint8_t x_22; -x_22 = lean_usize_dec_lt(x_13, x_12); -if (x_22 == 0) +uint8_t x_23; +x_23 = lean_usize_dec_lt(x_14, x_13); +if (x_23 == 0) { -lean_object* x_23; -lean_dec(x_19); +lean_object* x_24; +lean_dec(x_20); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -6442,954 +6445,950 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_14); -lean_ctor_set(x_23, 1, x_21); -return x_23; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_15); +lean_ctor_set(x_24, 1, x_22); +return x_24; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_24 = lean_array_uget(x_11, x_13); -x_34 = lean_ctor_get(x_14, 1); -lean_inc(x_34); -x_35 = lean_ctor_get(x_34, 1); +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; +x_25 = lean_array_uget(x_12, x_14); +x_35 = lean_ctor_get(x_15, 1); lean_inc(x_35); -x_36 = lean_ctor_get(x_35, 0); +x_36 = lean_ctor_get(x_35, 1); lean_inc(x_36); -x_37 = !lean_is_exclusive(x_14); -if (x_37 == 0) +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +x_38 = !lean_is_exclusive(x_15); +if (x_38 == 0) { -lean_object* x_38; lean_object* x_39; uint8_t x_40; -x_38 = lean_ctor_get(x_14, 0); -x_39 = lean_ctor_get(x_14, 1); -lean_dec(x_39); -x_40 = !lean_is_exclusive(x_34); -if (x_40 == 0) +lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_39 = lean_ctor_get(x_15, 0); +x_40 = lean_ctor_get(x_15, 1); +lean_dec(x_40); +x_41 = !lean_is_exclusive(x_35); +if (x_41 == 0) { -lean_object* x_41; lean_object* x_42; uint8_t x_43; -x_41 = lean_ctor_get(x_34, 0); -x_42 = lean_ctor_get(x_34, 1); -lean_dec(x_42); -x_43 = !lean_is_exclusive(x_35); -if (x_43 == 0) +lean_object* x_42; lean_object* x_43; uint8_t x_44; +x_42 = lean_ctor_get(x_35, 0); +x_43 = lean_ctor_get(x_35, 1); +lean_dec(x_43); +x_44 = !lean_is_exclusive(x_36); +if (x_44 == 0) { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; -x_44 = lean_ctor_get(x_35, 1); -x_45 = lean_ctor_get(x_35, 0); -lean_dec(x_45); +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; +x_45 = lean_ctor_get(x_36, 1); x_46 = lean_ctor_get(x_36, 0); -lean_inc(x_46); -x_47 = lean_ctor_get(x_36, 1); +lean_dec(x_46); +x_47 = lean_ctor_get(x_37, 0); lean_inc(x_47); -x_48 = lean_ctor_get(x_36, 2); +x_48 = lean_ctor_get(x_37, 1); lean_inc(x_48); -x_49 = lean_nat_dec_lt(x_46, x_47); -if (x_49 == 0) +x_49 = lean_ctor_get(x_37, 2); +lean_inc(x_49); +x_50 = lean_nat_dec_lt(x_47, x_48); +if (x_50 == 0) { -lean_object* x_50; +lean_object* x_51; +lean_dec(x_49); lean_dec(x_48); lean_dec(x_47); -lean_dec(x_46); -lean_dec(x_24); -x_50 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_50, 0, x_14); -x_25 = x_50; -x_26 = x_21; -goto block_33; +lean_dec(x_25); +x_51 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_51, 0, x_15); +x_26 = x_51; +x_27 = x_22; +goto block_34; } else { -uint8_t x_51; -x_51 = !lean_is_exclusive(x_36); -if (x_51 == 0) +uint8_t x_52; +x_52 = !lean_is_exclusive(x_37); +if (x_52 == 0) { -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; -x_52 = lean_ctor_get(x_36, 2); -lean_dec(x_52); -x_53 = lean_ctor_get(x_36, 1); +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; uint8_t x_60; +x_53 = lean_ctor_get(x_37, 2); lean_dec(x_53); -x_54 = lean_ctor_get(x_36, 0); +x_54 = lean_ctor_get(x_37, 1); lean_dec(x_54); -x_55 = lean_nat_add(x_46, x_48); -lean_ctor_set(x_36, 0, x_55); -x_56 = lean_ctor_get(x_41, 0); -lean_inc(x_56); -x_57 = lean_ctor_get(x_41, 1); +x_55 = lean_ctor_get(x_37, 0); +lean_dec(x_55); +x_56 = lean_nat_add(x_47, x_49); +lean_ctor_set(x_37, 0, x_56); +x_57 = lean_ctor_get(x_42, 0); lean_inc(x_57); -x_58 = lean_ctor_get(x_41, 2); +x_58 = lean_ctor_get(x_42, 1); lean_inc(x_58); -x_59 = lean_nat_dec_lt(x_57, x_58); -if (x_59 == 0) +x_59 = lean_ctor_get(x_42, 2); +lean_inc(x_59); +x_60 = lean_nat_dec_lt(x_58, x_59); +if (x_60 == 0) { -lean_object* x_60; +lean_object* x_61; +lean_dec(x_59); lean_dec(x_58); lean_dec(x_57); -lean_dec(x_56); -lean_dec(x_46); -lean_dec(x_24); -x_60 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_60, 0, x_14); -x_25 = x_60; -x_26 = x_21; -goto block_33; +lean_dec(x_47); +lean_dec(x_25); +x_61 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_61, 0, x_15); +x_26 = x_61; +x_27 = x_22; +goto block_34; } else { -uint8_t x_61; -x_61 = !lean_is_exclusive(x_41); -if (x_61 == 0) +uint8_t x_62; +x_62 = !lean_is_exclusive(x_42); +if (x_62 == 0) { -lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; -x_62 = lean_ctor_get(x_41, 2); -lean_dec(x_62); -x_63 = lean_ctor_get(x_41, 1); +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; +x_63 = lean_ctor_get(x_42, 2); lean_dec(x_63); -x_64 = lean_ctor_get(x_41, 0); +x_64 = lean_ctor_get(x_42, 1); lean_dec(x_64); -x_65 = lean_array_fget(x_56, x_57); -x_66 = lean_unsigned_to_nat(1u); -x_67 = lean_nat_add(x_57, x_66); -lean_dec(x_57); -lean_ctor_set(x_41, 1, x_67); -x_68 = lean_ctor_get(x_38, 0); -lean_inc(x_68); -x_69 = lean_ctor_get(x_38, 1); +x_65 = lean_ctor_get(x_42, 0); +lean_dec(x_65); +x_66 = lean_array_fget(x_57, x_58); +x_67 = lean_unsigned_to_nat(1u); +x_68 = lean_nat_add(x_58, x_67); +lean_dec(x_58); +lean_ctor_set(x_42, 1, x_68); +x_69 = lean_ctor_get(x_39, 0); lean_inc(x_69); -x_70 = lean_ctor_get(x_38, 2); +x_70 = lean_ctor_get(x_39, 1); lean_inc(x_70); -x_71 = lean_nat_dec_lt(x_69, x_70); -if (x_71 == 0) +x_71 = lean_ctor_get(x_39, 2); +lean_inc(x_71); +x_72 = lean_nat_dec_lt(x_70, x_71); +if (x_72 == 0) { -lean_object* x_72; +lean_object* x_73; +lean_dec(x_71); lean_dec(x_70); lean_dec(x_69); -lean_dec(x_68); -lean_dec(x_65); -lean_dec(x_46); -lean_dec(x_24); -x_72 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_72, 0, x_14); -x_25 = x_72; -x_26 = x_21; -goto block_33; +lean_dec(x_66); +lean_dec(x_47); +lean_dec(x_25); +x_73 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_73, 0, x_15); +x_26 = x_73; +x_27 = x_22; +goto block_34; } else { -uint8_t x_73; +uint8_t x_74; +lean_free_object(x_36); lean_free_object(x_35); -lean_free_object(x_34); -lean_free_object(x_14); -x_73 = !lean_is_exclusive(x_38); -if (x_73 == 0) -{ -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; uint8_t x_86; -x_74 = lean_ctor_get(x_38, 2); -lean_dec(x_74); -x_75 = lean_ctor_get(x_38, 1); +lean_free_object(x_15); +x_74 = !lean_is_exclusive(x_39); +if (x_74 == 0) +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; uint8_t x_88; +x_75 = lean_ctor_get(x_39, 2); lean_dec(x_75); -x_76 = lean_ctor_get(x_38, 0); +x_76 = lean_ctor_get(x_39, 1); lean_dec(x_76); -x_77 = lean_array_fget(x_68, x_69); -x_78 = lean_nat_add(x_69, x_66); -lean_dec(x_69); -lean_ctor_set(x_38, 1, x_78); -lean_inc(x_6); -lean_inc(x_8); -x_79 = lean_array_push(x_8, x_6); -x_80 = lean_nat_sub(x_77, x_65); -lean_dec(x_65); +x_77 = lean_ctor_get(x_39, 0); lean_dec(x_77); -x_81 = lean_nat_sub(x_80, x_66); -lean_dec(x_80); -x_82 = lean_unsigned_to_nat(0u); -lean_inc(x_81); -lean_inc(x_9); -x_83 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__1(x_9, x_81, x_82, x_81, x_66, x_79, x_15, x_16, x_17, x_18, x_19, x_20, x_21); -lean_dec(x_81); -x_84 = lean_ctor_get(x_83, 0); -lean_inc(x_84); -x_85 = lean_ctor_get(x_83, 1); -lean_inc(x_85); +x_78 = lean_array_fget(x_69, x_70); +x_79 = lean_nat_add(x_70, x_67); +lean_dec(x_70); +lean_ctor_set(x_39, 1, x_79); +x_80 = l___private_Lean_Elab_PreDefinition_WF_Rel_0__Lean_Elab_WF_unpackUnary_go___lambda__1___closed__1; +lean_inc(x_7); +x_81 = lean_array_push(x_80, x_7); +x_82 = lean_nat_sub(x_78, x_66); +lean_dec(x_66); +lean_dec(x_78); +x_83 = lean_nat_sub(x_82, x_67); +lean_dec(x_82); +x_84 = lean_unsigned_to_nat(0u); +lean_inc(x_83); +lean_inc(x_10); +x_85 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__1(x_2, x_10, x_83, x_84, x_83, x_67, x_81, x_16, x_17, x_18, x_19, x_20, x_21, x_22); lean_dec(x_83); -x_86 = lean_nat_dec_lt(x_66, x_10); -if (x_86 == 0) +x_86 = lean_ctor_get(x_85, 0); +lean_inc(x_86); +x_87 = lean_ctor_get(x_85, 1); +lean_inc(x_87); +lean_dec(x_85); +x_88 = lean_nat_dec_lt(x_67, x_11); +if (x_88 == 0) { -lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; -lean_dec(x_46); -lean_inc(x_19); -x_87 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_19, x_20, x_85); -x_88 = lean_ctor_get(x_87, 0); -lean_inc(x_88); -x_89 = lean_ctor_get(x_87, 1); -lean_inc(x_89); -lean_dec(x_87); -x_90 = lean_ctor_get(x_19, 10); +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; +lean_dec(x_47); +lean_inc(x_20); +x_89 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_20, x_21, x_87); +x_90 = lean_ctor_get(x_89, 0); lean_inc(x_90); -x_91 = lean_st_ref_get(x_20, x_89); -x_92 = lean_ctor_get(x_91, 0); +x_91 = lean_ctor_get(x_89, 1); +lean_inc(x_91); +lean_dec(x_89); +x_92 = lean_ctor_get(x_20, 10); lean_inc(x_92); -x_93 = lean_ctor_get(x_91, 1); -lean_inc(x_93); -lean_dec(x_91); -x_94 = lean_ctor_get(x_92, 0); +x_93 = lean_st_ref_get(x_21, x_91); +x_94 = lean_ctor_get(x_93, 0); lean_inc(x_94); -lean_dec(x_92); -x_95 = lean_environment_main_module(x_94); -x_96 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__1; -lean_inc(x_7); -x_97 = lean_name_mk_string(x_7, x_96); -x_98 = l_Subarray_forInUnsafe_loop___at_Lean_Elab_WF_getForbiddenByTrivialSizeOf___spec__1___closed__3; -x_99 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__2; -lean_inc(x_2); -x_100 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_100, 0, x_98); -lean_ctor_set(x_100, 1, x_2); -lean_ctor_set(x_100, 2, x_99); -x_101 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__3; -lean_inc(x_90); +x_95 = lean_ctor_get(x_93, 1); lean_inc(x_95); -x_102 = l_Lean_addMacroScope(x_95, x_101, x_90); -x_103 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__4; -lean_inc(x_5); -x_104 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_104, 0, x_103); -lean_ctor_set(x_104, 1, x_5); -lean_inc(x_88); -x_105 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_105, 0, x_88); -lean_ctor_set(x_105, 1, x_100); -lean_ctor_set(x_105, 2, x_102); -lean_ctor_set(x_105, 3, x_104); -lean_inc(x_4); -x_106 = l_Lean_addMacroScope(x_95, x_4, x_90); -lean_inc(x_5); +lean_dec(x_93); +x_96 = lean_ctor_get(x_94, 0); +lean_inc(x_96); +lean_dec(x_94); +x_97 = lean_environment_main_module(x_96); +x_98 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__1; +lean_inc(x_8); +x_99 = lean_name_mk_string(x_8, x_98); +x_100 = l_Subarray_forInUnsafe_loop___at_Lean_Elab_WF_getForbiddenByTrivialSizeOf___spec__1___closed__3; +x_101 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__2; lean_inc(x_3); +x_102 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_102, 0, x_100); +lean_ctor_set(x_102, 1, x_3); +lean_ctor_set(x_102, 2, x_101); +x_103 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__3; +lean_inc(x_92); +lean_inc(x_97); +x_104 = l_Lean_addMacroScope(x_97, x_103, x_92); +x_105 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__4; +lean_inc(x_6); +x_106 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_106, 0, x_105); +lean_ctor_set(x_106, 1, x_6); +lean_inc(x_90); x_107 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_107, 0, x_88); -lean_ctor_set(x_107, 1, x_3); -lean_ctor_set(x_107, 2, x_106); -lean_ctor_set(x_107, 3, x_5); -lean_inc(x_8); -x_108 = lean_array_push(x_8, x_107); -x_109 = lean_box(2); -x_110 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__6; -x_111 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_111, 0, x_109); -lean_ctor_set(x_111, 1, x_110); -lean_ctor_set(x_111, 2, x_108); -x_112 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__7; -x_113 = lean_array_push(x_112, x_105); -x_114 = lean_array_push(x_113, x_111); -x_115 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_115, 0, x_109); -lean_ctor_set(x_115, 1, x_97); -lean_ctor_set(x_115, 2, x_114); -x_116 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___lambda__1(x_24, x_84, x_36, x_41, x_38, x_44, x_115, x_15, x_16, x_17, x_18, x_19, x_20, x_93); -lean_dec(x_24); -x_117 = lean_ctor_get(x_116, 0); -lean_inc(x_117); -x_118 = lean_ctor_get(x_116, 1); -lean_inc(x_118); -lean_dec(x_116); -x_25 = x_117; -x_26 = x_118; -goto block_33; +lean_ctor_set(x_107, 0, x_90); +lean_ctor_set(x_107, 1, x_102); +lean_ctor_set(x_107, 2, x_104); +lean_ctor_set(x_107, 3, x_106); +lean_inc(x_5); +x_108 = l_Lean_addMacroScope(x_97, x_5, x_92); +lean_inc(x_6); +lean_inc(x_4); +x_109 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_109, 0, x_90); +lean_ctor_set(x_109, 1, x_4); +lean_ctor_set(x_109, 2, x_108); +lean_ctor_set(x_109, 3, x_6); +lean_inc(x_9); +x_110 = lean_array_push(x_9, x_109); +x_111 = lean_box(2); +x_112 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__6; +x_113 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_113, 0, x_111); +lean_ctor_set(x_113, 1, x_112); +lean_ctor_set(x_113, 2, x_110); +x_114 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__7; +x_115 = lean_array_push(x_114, x_107); +x_116 = lean_array_push(x_115, x_113); +x_117 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_117, 0, x_111); +lean_ctor_set(x_117, 1, x_99); +lean_ctor_set(x_117, 2, x_116); +x_118 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___lambda__1(x_25, x_86, x_37, x_42, x_39, x_45, x_117, x_16, x_17, x_18, x_19, x_20, x_21, x_95); +lean_dec(x_25); +x_119 = lean_ctor_get(x_118, 0); +lean_inc(x_119); +x_120 = lean_ctor_get(x_118, 1); +lean_inc(x_120); +lean_dec(x_118); +x_26 = x_119; +x_27 = x_120; +goto block_34; } else { -lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; -lean_inc(x_19); -x_119 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_19, x_20, x_85); -x_120 = lean_ctor_get(x_119, 0); -lean_inc(x_120); -x_121 = lean_ctor_get(x_119, 1); -lean_inc(x_121); -lean_dec(x_119); -x_122 = lean_ctor_get(x_19, 10); +lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; +lean_inc(x_20); +x_121 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_20, x_21, x_87); +x_122 = lean_ctor_get(x_121, 0); lean_inc(x_122); -x_123 = lean_st_ref_get(x_20, x_121); -x_124 = lean_ctor_get(x_123, 0); +x_123 = lean_ctor_get(x_121, 1); +lean_inc(x_123); +lean_dec(x_121); +x_124 = lean_ctor_get(x_20, 10); lean_inc(x_124); -x_125 = lean_ctor_get(x_123, 1); -lean_inc(x_125); -lean_dec(x_123); -x_126 = lean_ctor_get(x_124, 0); +x_125 = lean_st_ref_get(x_21, x_123); +x_126 = lean_ctor_get(x_125, 0); lean_inc(x_126); -lean_dec(x_124); -x_127 = lean_environment_main_module(x_126); -x_128 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__8; -lean_inc(x_7); -x_129 = lean_name_mk_string(x_7, x_128); -x_130 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__9; -lean_inc(x_120); -x_131 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_131, 0, x_120); -lean_ctor_set(x_131, 1, x_130); -x_132 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__1; -lean_inc(x_7); -x_133 = lean_name_mk_string(x_7, x_132); -x_134 = l_Subarray_forInUnsafe_loop___at_Lean_Elab_WF_getForbiddenByTrivialSizeOf___spec__1___closed__3; -x_135 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__2; -lean_inc(x_2); -x_136 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_136, 0, x_134); -lean_ctor_set(x_136, 1, x_2); -lean_ctor_set(x_136, 2, x_135); -x_137 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__3; -lean_inc(x_122); +x_127 = lean_ctor_get(x_125, 1); lean_inc(x_127); -x_138 = l_Lean_addMacroScope(x_127, x_137, x_122); -x_139 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__4; -lean_inc(x_5); -x_140 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_140, 0, x_139); -lean_ctor_set(x_140, 1, x_5); -lean_inc(x_120); -x_141 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_141, 0, x_120); -lean_ctor_set(x_141, 1, x_136); -lean_ctor_set(x_141, 2, x_138); -lean_ctor_set(x_141, 3, x_140); -lean_inc(x_4); -x_142 = l_Lean_addMacroScope(x_127, x_4, x_122); -lean_inc(x_5); -lean_inc(x_3); -lean_inc(x_120); -x_143 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_143, 0, x_120); -lean_ctor_set(x_143, 1, x_3); -lean_ctor_set(x_143, 2, x_142); -lean_ctor_set(x_143, 3, x_5); +lean_dec(x_125); +x_128 = lean_ctor_get(x_126, 0); +lean_inc(x_128); +lean_dec(x_126); +x_129 = lean_environment_main_module(x_128); +x_130 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__8; lean_inc(x_8); -x_144 = lean_array_push(x_8, x_143); -x_145 = lean_box(2); -x_146 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__6; -x_147 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_147, 0, x_145); -lean_ctor_set(x_147, 1, x_146); -lean_ctor_set(x_147, 2, x_144); -x_148 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__7; -x_149 = lean_array_push(x_148, x_141); -x_150 = lean_array_push(x_149, x_147); -x_151 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_151, 0, x_145); -lean_ctor_set(x_151, 1, x_133); -lean_ctor_set(x_151, 2, x_150); -x_152 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__10; -lean_inc(x_7); -x_153 = lean_name_mk_string(x_7, x_152); -x_154 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__11; -lean_inc(x_120); -x_155 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_155, 0, x_120); -lean_ctor_set(x_155, 1, x_154); -x_156 = l_Nat_repr(x_46); -x_157 = l_Lean_numLitKind; -x_158 = l_Lean_Syntax_mkLit(x_157, x_156, x_145); +x_131 = lean_name_mk_string(x_8, x_130); +x_132 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__9; +lean_inc(x_122); +x_133 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_133, 0, x_122); +lean_ctor_set(x_133, 1, x_132); +x_134 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__1; lean_inc(x_8); -x_159 = lean_array_push(x_8, x_158); -x_160 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_160, 0, x_145); -lean_ctor_set(x_160, 1, x_146); -lean_ctor_set(x_160, 2, x_159); -x_161 = lean_array_push(x_148, x_155); -x_162 = lean_array_push(x_161, x_160); -x_163 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_163, 0, x_145); -lean_ctor_set(x_163, 1, x_153); -lean_ctor_set(x_163, 2, x_162); +x_135 = lean_name_mk_string(x_8, x_134); +x_136 = l_Subarray_forInUnsafe_loop___at_Lean_Elab_WF_getForbiddenByTrivialSizeOf___spec__1___closed__3; +x_137 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__2; +lean_inc(x_3); +x_138 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_138, 0, x_136); +lean_ctor_set(x_138, 1, x_3); +lean_ctor_set(x_138, 2, x_137); +x_139 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__3; +lean_inc(x_124); +lean_inc(x_129); +x_140 = l_Lean_addMacroScope(x_129, x_139, x_124); +x_141 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__4; +lean_inc(x_6); +x_142 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_142, 0, x_141); +lean_ctor_set(x_142, 1, x_6); +lean_inc(x_122); +x_143 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_143, 0, x_122); +lean_ctor_set(x_143, 1, x_138); +lean_ctor_set(x_143, 2, x_140); +lean_ctor_set(x_143, 3, x_142); +lean_inc(x_5); +x_144 = l_Lean_addMacroScope(x_129, x_5, x_124); +lean_inc(x_6); +lean_inc(x_4); +lean_inc(x_122); +x_145 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_145, 0, x_122); +lean_ctor_set(x_145, 1, x_4); +lean_ctor_set(x_145, 2, x_144); +lean_ctor_set(x_145, 3, x_6); +lean_inc(x_9); +x_146 = lean_array_push(x_9, x_145); +x_147 = lean_box(2); +x_148 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__6; +x_149 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_149, 0, x_147); +lean_ctor_set(x_149, 1, x_148); +lean_ctor_set(x_149, 2, x_146); +x_150 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__7; +x_151 = lean_array_push(x_150, x_143); +x_152 = lean_array_push(x_151, x_149); +x_153 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_153, 0, x_147); +lean_ctor_set(x_153, 1, x_135); +lean_ctor_set(x_153, 2, x_152); +x_154 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__10; lean_inc(x_8); -x_164 = lean_array_push(x_8, x_163); -x_165 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_165, 0, x_145); -lean_ctor_set(x_165, 1, x_146); -lean_ctor_set(x_165, 2, x_164); -x_166 = lean_array_push(x_148, x_151); -x_167 = lean_array_push(x_166, x_165); -x_168 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_168, 0, x_145); -lean_ctor_set(x_168, 1, x_146); -lean_ctor_set(x_168, 2, x_167); -x_169 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__12; -x_170 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_170, 0, x_120); -lean_ctor_set(x_170, 1, x_169); -x_171 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__13; -x_172 = lean_array_push(x_171, x_131); -x_173 = lean_array_push(x_172, x_168); -x_174 = lean_array_push(x_173, x_170); -x_175 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_175, 0, x_145); -lean_ctor_set(x_175, 1, x_129); -lean_ctor_set(x_175, 2, x_174); -x_176 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___lambda__1(x_24, x_84, x_36, x_41, x_38, x_44, x_175, x_15, x_16, x_17, x_18, x_19, x_20, x_125); -lean_dec(x_24); -x_177 = lean_ctor_get(x_176, 0); -lean_inc(x_177); -x_178 = lean_ctor_get(x_176, 1); +x_155 = lean_name_mk_string(x_8, x_154); +x_156 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__11; +lean_inc(x_122); +x_157 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_157, 0, x_122); +lean_ctor_set(x_157, 1, x_156); +x_158 = l_Nat_repr(x_47); +x_159 = l_Lean_Syntax_mkNumLit(x_158, x_147); +lean_inc(x_9); +x_160 = lean_array_push(x_9, x_159); +x_161 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_161, 0, x_147); +lean_ctor_set(x_161, 1, x_148); +lean_ctor_set(x_161, 2, x_160); +x_162 = lean_array_push(x_150, x_157); +x_163 = lean_array_push(x_162, x_161); +x_164 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_164, 0, x_147); +lean_ctor_set(x_164, 1, x_155); +lean_ctor_set(x_164, 2, x_163); +lean_inc(x_9); +x_165 = lean_array_push(x_9, x_164); +x_166 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_166, 0, x_147); +lean_ctor_set(x_166, 1, x_148); +lean_ctor_set(x_166, 2, x_165); +x_167 = lean_array_push(x_150, x_153); +x_168 = lean_array_push(x_167, x_166); +x_169 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_169, 0, x_147); +lean_ctor_set(x_169, 1, x_148); +lean_ctor_set(x_169, 2, x_168); +x_170 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__12; +x_171 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_171, 0, x_122); +lean_ctor_set(x_171, 1, x_170); +x_172 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__13; +x_173 = lean_array_push(x_172, x_133); +x_174 = lean_array_push(x_173, x_169); +x_175 = lean_array_push(x_174, x_171); +x_176 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_176, 0, x_147); +lean_ctor_set(x_176, 1, x_131); +lean_ctor_set(x_176, 2, x_175); +x_177 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___lambda__1(x_25, x_86, x_37, x_42, x_39, x_45, x_176, x_16, x_17, x_18, x_19, x_20, x_21, x_127); +lean_dec(x_25); +x_178 = lean_ctor_get(x_177, 0); lean_inc(x_178); -lean_dec(x_176); -x_25 = x_177; +x_179 = lean_ctor_get(x_177, 1); +lean_inc(x_179); +lean_dec(x_177); x_26 = x_178; -goto block_33; +x_27 = x_179; +goto block_34; } } else { -lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; uint8_t x_189; -lean_dec(x_38); -x_179 = lean_array_fget(x_68, x_69); -x_180 = lean_nat_add(x_69, x_66); -lean_dec(x_69); -x_181 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_181, 0, x_68); -lean_ctor_set(x_181, 1, x_180); -lean_ctor_set(x_181, 2, x_70); -lean_inc(x_6); -lean_inc(x_8); -x_182 = lean_array_push(x_8, x_6); -x_183 = lean_nat_sub(x_179, x_65); -lean_dec(x_65); -lean_dec(x_179); -x_184 = lean_nat_sub(x_183, x_66); -lean_dec(x_183); -x_185 = lean_unsigned_to_nat(0u); -lean_inc(x_184); -lean_inc(x_9); -x_186 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__1(x_9, x_184, x_185, x_184, x_66, x_182, x_15, x_16, x_17, x_18, x_19, x_20, x_21); -lean_dec(x_184); -x_187 = lean_ctor_get(x_186, 0); -lean_inc(x_187); -x_188 = lean_ctor_get(x_186, 1); -lean_inc(x_188); +lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; uint8_t x_191; +lean_dec(x_39); +x_180 = lean_array_fget(x_69, x_70); +x_181 = lean_nat_add(x_70, x_67); +lean_dec(x_70); +x_182 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_182, 0, x_69); +lean_ctor_set(x_182, 1, x_181); +lean_ctor_set(x_182, 2, x_71); +x_183 = l___private_Lean_Elab_PreDefinition_WF_Rel_0__Lean_Elab_WF_unpackUnary_go___lambda__1___closed__1; +lean_inc(x_7); +x_184 = lean_array_push(x_183, x_7); +x_185 = lean_nat_sub(x_180, x_66); +lean_dec(x_66); +lean_dec(x_180); +x_186 = lean_nat_sub(x_185, x_67); +lean_dec(x_185); +x_187 = lean_unsigned_to_nat(0u); +lean_inc(x_186); +lean_inc(x_10); +x_188 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__1(x_2, x_10, x_186, x_187, x_186, x_67, x_184, x_16, x_17, x_18, x_19, x_20, x_21, x_22); lean_dec(x_186); -x_189 = lean_nat_dec_lt(x_66, x_10); -if (x_189 == 0) +x_189 = lean_ctor_get(x_188, 0); +lean_inc(x_189); +x_190 = lean_ctor_get(x_188, 1); +lean_inc(x_190); +lean_dec(x_188); +x_191 = lean_nat_dec_lt(x_67, x_11); +if (x_191 == 0) { -lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; -lean_dec(x_46); -lean_inc(x_19); -x_190 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_19, x_20, x_188); -x_191 = lean_ctor_get(x_190, 0); -lean_inc(x_191); -x_192 = lean_ctor_get(x_190, 1); -lean_inc(x_192); -lean_dec(x_190); -x_193 = lean_ctor_get(x_19, 10); +lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; +lean_dec(x_47); +lean_inc(x_20); +x_192 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_20, x_21, x_190); +x_193 = lean_ctor_get(x_192, 0); lean_inc(x_193); -x_194 = lean_st_ref_get(x_20, x_192); -x_195 = lean_ctor_get(x_194, 0); +x_194 = lean_ctor_get(x_192, 1); +lean_inc(x_194); +lean_dec(x_192); +x_195 = lean_ctor_get(x_20, 10); lean_inc(x_195); -x_196 = lean_ctor_get(x_194, 1); -lean_inc(x_196); -lean_dec(x_194); -x_197 = lean_ctor_get(x_195, 0); +x_196 = lean_st_ref_get(x_21, x_194); +x_197 = lean_ctor_get(x_196, 0); lean_inc(x_197); -lean_dec(x_195); -x_198 = lean_environment_main_module(x_197); -x_199 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__1; -lean_inc(x_7); -x_200 = lean_name_mk_string(x_7, x_199); -x_201 = l_Subarray_forInUnsafe_loop___at_Lean_Elab_WF_getForbiddenByTrivialSizeOf___spec__1___closed__3; -x_202 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__2; -lean_inc(x_2); -x_203 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_203, 0, x_201); -lean_ctor_set(x_203, 1, x_2); -lean_ctor_set(x_203, 2, x_202); -x_204 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__3; -lean_inc(x_193); +x_198 = lean_ctor_get(x_196, 1); lean_inc(x_198); -x_205 = l_Lean_addMacroScope(x_198, x_204, x_193); -x_206 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__4; -lean_inc(x_5); -x_207 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_207, 0, x_206); -lean_ctor_set(x_207, 1, x_5); -lean_inc(x_191); -x_208 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_208, 0, x_191); -lean_ctor_set(x_208, 1, x_203); -lean_ctor_set(x_208, 2, x_205); -lean_ctor_set(x_208, 3, x_207); -lean_inc(x_4); -x_209 = l_Lean_addMacroScope(x_198, x_4, x_193); -lean_inc(x_5); +lean_dec(x_196); +x_199 = lean_ctor_get(x_197, 0); +lean_inc(x_199); +lean_dec(x_197); +x_200 = lean_environment_main_module(x_199); +x_201 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__1; +lean_inc(x_8); +x_202 = lean_name_mk_string(x_8, x_201); +x_203 = l_Subarray_forInUnsafe_loop___at_Lean_Elab_WF_getForbiddenByTrivialSizeOf___spec__1___closed__3; +x_204 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__2; lean_inc(x_3); +x_205 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_205, 0, x_203); +lean_ctor_set(x_205, 1, x_3); +lean_ctor_set(x_205, 2, x_204); +x_206 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__3; +lean_inc(x_195); +lean_inc(x_200); +x_207 = l_Lean_addMacroScope(x_200, x_206, x_195); +x_208 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__4; +lean_inc(x_6); +x_209 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_209, 0, x_208); +lean_ctor_set(x_209, 1, x_6); +lean_inc(x_193); x_210 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_210, 0, x_191); -lean_ctor_set(x_210, 1, x_3); -lean_ctor_set(x_210, 2, x_209); -lean_ctor_set(x_210, 3, x_5); -lean_inc(x_8); -x_211 = lean_array_push(x_8, x_210); -x_212 = lean_box(2); -x_213 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__6; -x_214 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_214, 0, x_212); -lean_ctor_set(x_214, 1, x_213); -lean_ctor_set(x_214, 2, x_211); -x_215 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__7; -x_216 = lean_array_push(x_215, x_208); -x_217 = lean_array_push(x_216, x_214); -x_218 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_218, 0, x_212); -lean_ctor_set(x_218, 1, x_200); -lean_ctor_set(x_218, 2, x_217); -x_219 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___lambda__1(x_24, x_187, x_36, x_41, x_181, x_44, x_218, x_15, x_16, x_17, x_18, x_19, x_20, x_196); -lean_dec(x_24); -x_220 = lean_ctor_get(x_219, 0); -lean_inc(x_220); -x_221 = lean_ctor_get(x_219, 1); -lean_inc(x_221); -lean_dec(x_219); -x_25 = x_220; -x_26 = x_221; -goto block_33; +lean_ctor_set(x_210, 0, x_193); +lean_ctor_set(x_210, 1, x_205); +lean_ctor_set(x_210, 2, x_207); +lean_ctor_set(x_210, 3, x_209); +lean_inc(x_5); +x_211 = l_Lean_addMacroScope(x_200, x_5, x_195); +lean_inc(x_6); +lean_inc(x_4); +x_212 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_212, 0, x_193); +lean_ctor_set(x_212, 1, x_4); +lean_ctor_set(x_212, 2, x_211); +lean_ctor_set(x_212, 3, x_6); +lean_inc(x_9); +x_213 = lean_array_push(x_9, x_212); +x_214 = lean_box(2); +x_215 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__6; +x_216 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_216, 0, x_214); +lean_ctor_set(x_216, 1, x_215); +lean_ctor_set(x_216, 2, x_213); +x_217 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__7; +x_218 = lean_array_push(x_217, x_210); +x_219 = lean_array_push(x_218, x_216); +x_220 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_220, 0, x_214); +lean_ctor_set(x_220, 1, x_202); +lean_ctor_set(x_220, 2, x_219); +x_221 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___lambda__1(x_25, x_189, x_37, x_42, x_182, x_45, x_220, x_16, x_17, x_18, x_19, x_20, x_21, x_198); +lean_dec(x_25); +x_222 = lean_ctor_get(x_221, 0); +lean_inc(x_222); +x_223 = lean_ctor_get(x_221, 1); +lean_inc(x_223); +lean_dec(x_221); +x_26 = x_222; +x_27 = x_223; +goto block_34; } else { -lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; -lean_inc(x_19); -x_222 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_19, x_20, x_188); -x_223 = lean_ctor_get(x_222, 0); -lean_inc(x_223); -x_224 = lean_ctor_get(x_222, 1); -lean_inc(x_224); -lean_dec(x_222); -x_225 = lean_ctor_get(x_19, 10); +lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; +lean_inc(x_20); +x_224 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_20, x_21, x_190); +x_225 = lean_ctor_get(x_224, 0); lean_inc(x_225); -x_226 = lean_st_ref_get(x_20, x_224); -x_227 = lean_ctor_get(x_226, 0); +x_226 = lean_ctor_get(x_224, 1); +lean_inc(x_226); +lean_dec(x_224); +x_227 = lean_ctor_get(x_20, 10); lean_inc(x_227); -x_228 = lean_ctor_get(x_226, 1); -lean_inc(x_228); -lean_dec(x_226); -x_229 = lean_ctor_get(x_227, 0); +x_228 = lean_st_ref_get(x_21, x_226); +x_229 = lean_ctor_get(x_228, 0); lean_inc(x_229); -lean_dec(x_227); -x_230 = lean_environment_main_module(x_229); -x_231 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__8; -lean_inc(x_7); -x_232 = lean_name_mk_string(x_7, x_231); -x_233 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__9; -lean_inc(x_223); -x_234 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_234, 0, x_223); -lean_ctor_set(x_234, 1, x_233); -x_235 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__1; -lean_inc(x_7); -x_236 = lean_name_mk_string(x_7, x_235); -x_237 = l_Subarray_forInUnsafe_loop___at_Lean_Elab_WF_getForbiddenByTrivialSizeOf___spec__1___closed__3; -x_238 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__2; -lean_inc(x_2); -x_239 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_239, 0, x_237); -lean_ctor_set(x_239, 1, x_2); -lean_ctor_set(x_239, 2, x_238); -x_240 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__3; -lean_inc(x_225); +x_230 = lean_ctor_get(x_228, 1); lean_inc(x_230); -x_241 = l_Lean_addMacroScope(x_230, x_240, x_225); -x_242 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__4; -lean_inc(x_5); -x_243 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_243, 0, x_242); -lean_ctor_set(x_243, 1, x_5); -lean_inc(x_223); -x_244 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_244, 0, x_223); -lean_ctor_set(x_244, 1, x_239); -lean_ctor_set(x_244, 2, x_241); -lean_ctor_set(x_244, 3, x_243); -lean_inc(x_4); -x_245 = l_Lean_addMacroScope(x_230, x_4, x_225); -lean_inc(x_5); -lean_inc(x_3); -lean_inc(x_223); -x_246 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_246, 0, x_223); -lean_ctor_set(x_246, 1, x_3); -lean_ctor_set(x_246, 2, x_245); -lean_ctor_set(x_246, 3, x_5); +lean_dec(x_228); +x_231 = lean_ctor_get(x_229, 0); +lean_inc(x_231); +lean_dec(x_229); +x_232 = lean_environment_main_module(x_231); +x_233 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__8; lean_inc(x_8); -x_247 = lean_array_push(x_8, x_246); -x_248 = lean_box(2); -x_249 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__6; -x_250 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_250, 0, x_248); -lean_ctor_set(x_250, 1, x_249); -lean_ctor_set(x_250, 2, x_247); -x_251 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__7; -x_252 = lean_array_push(x_251, x_244); -x_253 = lean_array_push(x_252, x_250); -x_254 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_254, 0, x_248); -lean_ctor_set(x_254, 1, x_236); -lean_ctor_set(x_254, 2, x_253); -x_255 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__10; -lean_inc(x_7); -x_256 = lean_name_mk_string(x_7, x_255); -x_257 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__11; -lean_inc(x_223); -x_258 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_258, 0, x_223); -lean_ctor_set(x_258, 1, x_257); -x_259 = l_Nat_repr(x_46); -x_260 = l_Lean_numLitKind; -x_261 = l_Lean_Syntax_mkLit(x_260, x_259, x_248); +x_234 = lean_name_mk_string(x_8, x_233); +x_235 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__9; +lean_inc(x_225); +x_236 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_236, 0, x_225); +lean_ctor_set(x_236, 1, x_235); +x_237 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__1; lean_inc(x_8); -x_262 = lean_array_push(x_8, x_261); -x_263 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_263, 0, x_248); -lean_ctor_set(x_263, 1, x_249); -lean_ctor_set(x_263, 2, x_262); -x_264 = lean_array_push(x_251, x_258); -x_265 = lean_array_push(x_264, x_263); -x_266 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_266, 0, x_248); -lean_ctor_set(x_266, 1, x_256); -lean_ctor_set(x_266, 2, x_265); +x_238 = lean_name_mk_string(x_8, x_237); +x_239 = l_Subarray_forInUnsafe_loop___at_Lean_Elab_WF_getForbiddenByTrivialSizeOf___spec__1___closed__3; +x_240 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__2; +lean_inc(x_3); +x_241 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_241, 0, x_239); +lean_ctor_set(x_241, 1, x_3); +lean_ctor_set(x_241, 2, x_240); +x_242 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__3; +lean_inc(x_227); +lean_inc(x_232); +x_243 = l_Lean_addMacroScope(x_232, x_242, x_227); +x_244 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__4; +lean_inc(x_6); +x_245 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_245, 0, x_244); +lean_ctor_set(x_245, 1, x_6); +lean_inc(x_225); +x_246 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_246, 0, x_225); +lean_ctor_set(x_246, 1, x_241); +lean_ctor_set(x_246, 2, x_243); +lean_ctor_set(x_246, 3, x_245); +lean_inc(x_5); +x_247 = l_Lean_addMacroScope(x_232, x_5, x_227); +lean_inc(x_6); +lean_inc(x_4); +lean_inc(x_225); +x_248 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_248, 0, x_225); +lean_ctor_set(x_248, 1, x_4); +lean_ctor_set(x_248, 2, x_247); +lean_ctor_set(x_248, 3, x_6); +lean_inc(x_9); +x_249 = lean_array_push(x_9, x_248); +x_250 = lean_box(2); +x_251 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__6; +x_252 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_252, 0, x_250); +lean_ctor_set(x_252, 1, x_251); +lean_ctor_set(x_252, 2, x_249); +x_253 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__7; +x_254 = lean_array_push(x_253, x_246); +x_255 = lean_array_push(x_254, x_252); +x_256 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_256, 0, x_250); +lean_ctor_set(x_256, 1, x_238); +lean_ctor_set(x_256, 2, x_255); +x_257 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__10; lean_inc(x_8); -x_267 = lean_array_push(x_8, x_266); -x_268 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_268, 0, x_248); -lean_ctor_set(x_268, 1, x_249); -lean_ctor_set(x_268, 2, x_267); -x_269 = lean_array_push(x_251, x_254); -x_270 = lean_array_push(x_269, x_268); -x_271 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_271, 0, x_248); -lean_ctor_set(x_271, 1, x_249); -lean_ctor_set(x_271, 2, x_270); -x_272 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__12; -x_273 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_273, 0, x_223); -lean_ctor_set(x_273, 1, x_272); -x_274 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__13; -x_275 = lean_array_push(x_274, x_234); -x_276 = lean_array_push(x_275, x_271); -x_277 = lean_array_push(x_276, x_273); -x_278 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_278, 0, x_248); -lean_ctor_set(x_278, 1, x_232); -lean_ctor_set(x_278, 2, x_277); -x_279 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___lambda__1(x_24, x_187, x_36, x_41, x_181, x_44, x_278, x_15, x_16, x_17, x_18, x_19, x_20, x_228); -lean_dec(x_24); -x_280 = lean_ctor_get(x_279, 0); -lean_inc(x_280); -x_281 = lean_ctor_get(x_279, 1); +x_258 = lean_name_mk_string(x_8, x_257); +x_259 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__11; +lean_inc(x_225); +x_260 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_260, 0, x_225); +lean_ctor_set(x_260, 1, x_259); +x_261 = l_Nat_repr(x_47); +x_262 = l_Lean_Syntax_mkNumLit(x_261, x_250); +lean_inc(x_9); +x_263 = lean_array_push(x_9, x_262); +x_264 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_264, 0, x_250); +lean_ctor_set(x_264, 1, x_251); +lean_ctor_set(x_264, 2, x_263); +x_265 = lean_array_push(x_253, x_260); +x_266 = lean_array_push(x_265, x_264); +x_267 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_267, 0, x_250); +lean_ctor_set(x_267, 1, x_258); +lean_ctor_set(x_267, 2, x_266); +lean_inc(x_9); +x_268 = lean_array_push(x_9, x_267); +x_269 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_269, 0, x_250); +lean_ctor_set(x_269, 1, x_251); +lean_ctor_set(x_269, 2, x_268); +x_270 = lean_array_push(x_253, x_256); +x_271 = lean_array_push(x_270, x_269); +x_272 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_272, 0, x_250); +lean_ctor_set(x_272, 1, x_251); +lean_ctor_set(x_272, 2, x_271); +x_273 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__12; +x_274 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_274, 0, x_225); +lean_ctor_set(x_274, 1, x_273); +x_275 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__13; +x_276 = lean_array_push(x_275, x_236); +x_277 = lean_array_push(x_276, x_272); +x_278 = lean_array_push(x_277, x_274); +x_279 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_279, 0, x_250); +lean_ctor_set(x_279, 1, x_234); +lean_ctor_set(x_279, 2, x_278); +x_280 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___lambda__1(x_25, x_189, x_37, x_42, x_182, x_45, x_279, x_16, x_17, x_18, x_19, x_20, x_21, x_230); +lean_dec(x_25); +x_281 = lean_ctor_get(x_280, 0); lean_inc(x_281); -lean_dec(x_279); -x_25 = x_280; +x_282 = lean_ctor_get(x_280, 1); +lean_inc(x_282); +lean_dec(x_280); x_26 = x_281; -goto block_33; +x_27 = x_282; +goto block_34; } } } } else { -lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; uint8_t x_289; -lean_dec(x_41); -x_282 = lean_array_fget(x_56, x_57); -x_283 = lean_unsigned_to_nat(1u); -x_284 = lean_nat_add(x_57, x_283); -lean_dec(x_57); -x_285 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_285, 0, x_56); -lean_ctor_set(x_285, 1, x_284); -lean_ctor_set(x_285, 2, x_58); -x_286 = lean_ctor_get(x_38, 0); -lean_inc(x_286); -x_287 = lean_ctor_get(x_38, 1); +lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; uint8_t x_290; +lean_dec(x_42); +x_283 = lean_array_fget(x_57, x_58); +x_284 = lean_unsigned_to_nat(1u); +x_285 = lean_nat_add(x_58, x_284); +lean_dec(x_58); +x_286 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_286, 0, x_57); +lean_ctor_set(x_286, 1, x_285); +lean_ctor_set(x_286, 2, x_59); +x_287 = lean_ctor_get(x_39, 0); lean_inc(x_287); -x_288 = lean_ctor_get(x_38, 2); +x_288 = lean_ctor_get(x_39, 1); lean_inc(x_288); -x_289 = lean_nat_dec_lt(x_287, x_288); -if (x_289 == 0) +x_289 = lean_ctor_get(x_39, 2); +lean_inc(x_289); +x_290 = lean_nat_dec_lt(x_288, x_289); +if (x_290 == 0) { -lean_object* x_290; +lean_object* x_291; +lean_dec(x_289); lean_dec(x_288); lean_dec(x_287); -lean_dec(x_286); -lean_dec(x_282); -lean_dec(x_46); -lean_dec(x_24); -lean_ctor_set(x_34, 0, x_285); -x_290 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_290, 0, x_14); -x_25 = x_290; -x_26 = x_21; -goto block_33; +lean_dec(x_283); +lean_dec(x_47); +lean_dec(x_25); +lean_ctor_set(x_35, 0, x_286); +x_291 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_291, 0, x_15); +x_26 = x_291; +x_27 = x_22; +goto block_34; } else { -lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; uint8_t x_302; +lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; uint8_t x_304; +lean_free_object(x_36); lean_free_object(x_35); -lean_free_object(x_34); -lean_free_object(x_14); -if (lean_is_exclusive(x_38)) { - lean_ctor_release(x_38, 0); - lean_ctor_release(x_38, 1); - lean_ctor_release(x_38, 2); - x_291 = x_38; +lean_free_object(x_15); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + lean_ctor_release(x_39, 2); + x_292 = x_39; } else { - lean_dec_ref(x_38); - x_291 = lean_box(0); + lean_dec_ref(x_39); + x_292 = lean_box(0); } -x_292 = lean_array_fget(x_286, x_287); -x_293 = lean_nat_add(x_287, x_283); -lean_dec(x_287); -if (lean_is_scalar(x_291)) { - x_294 = lean_alloc_ctor(0, 3, 0); +x_293 = lean_array_fget(x_287, x_288); +x_294 = lean_nat_add(x_288, x_284); +lean_dec(x_288); +if (lean_is_scalar(x_292)) { + x_295 = lean_alloc_ctor(0, 3, 0); } else { - x_294 = x_291; + x_295 = x_292; } -lean_ctor_set(x_294, 0, x_286); -lean_ctor_set(x_294, 1, x_293); -lean_ctor_set(x_294, 2, x_288); -lean_inc(x_6); -lean_inc(x_8); -x_295 = lean_array_push(x_8, x_6); -x_296 = lean_nat_sub(x_292, x_282); -lean_dec(x_282); -lean_dec(x_292); -x_297 = lean_nat_sub(x_296, x_283); -lean_dec(x_296); -x_298 = lean_unsigned_to_nat(0u); -lean_inc(x_297); -lean_inc(x_9); -x_299 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__1(x_9, x_297, x_298, x_297, x_283, x_295, x_15, x_16, x_17, x_18, x_19, x_20, x_21); -lean_dec(x_297); -x_300 = lean_ctor_get(x_299, 0); -lean_inc(x_300); -x_301 = lean_ctor_get(x_299, 1); -lean_inc(x_301); +lean_ctor_set(x_295, 0, x_287); +lean_ctor_set(x_295, 1, x_294); +lean_ctor_set(x_295, 2, x_289); +x_296 = l___private_Lean_Elab_PreDefinition_WF_Rel_0__Lean_Elab_WF_unpackUnary_go___lambda__1___closed__1; +lean_inc(x_7); +x_297 = lean_array_push(x_296, x_7); +x_298 = lean_nat_sub(x_293, x_283); +lean_dec(x_283); +lean_dec(x_293); +x_299 = lean_nat_sub(x_298, x_284); +lean_dec(x_298); +x_300 = lean_unsigned_to_nat(0u); +lean_inc(x_299); +lean_inc(x_10); +x_301 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__1(x_2, x_10, x_299, x_300, x_299, x_284, x_297, x_16, x_17, x_18, x_19, x_20, x_21, x_22); lean_dec(x_299); -x_302 = lean_nat_dec_lt(x_283, x_10); -if (x_302 == 0) -{ -lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; -lean_dec(x_46); -lean_inc(x_19); -x_303 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_19, x_20, x_301); -x_304 = lean_ctor_get(x_303, 0); -lean_inc(x_304); -x_305 = lean_ctor_get(x_303, 1); -lean_inc(x_305); -lean_dec(x_303); -x_306 = lean_ctor_get(x_19, 10); +x_302 = lean_ctor_get(x_301, 0); +lean_inc(x_302); +x_303 = lean_ctor_get(x_301, 1); +lean_inc(x_303); +lean_dec(x_301); +x_304 = lean_nat_dec_lt(x_284, x_11); +if (x_304 == 0) +{ +lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; +lean_dec(x_47); +lean_inc(x_20); +x_305 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_20, x_21, x_303); +x_306 = lean_ctor_get(x_305, 0); lean_inc(x_306); -x_307 = lean_st_ref_get(x_20, x_305); -x_308 = lean_ctor_get(x_307, 0); +x_307 = lean_ctor_get(x_305, 1); +lean_inc(x_307); +lean_dec(x_305); +x_308 = lean_ctor_get(x_20, 10); lean_inc(x_308); -x_309 = lean_ctor_get(x_307, 1); -lean_inc(x_309); -lean_dec(x_307); -x_310 = lean_ctor_get(x_308, 0); +x_309 = lean_st_ref_get(x_21, x_307); +x_310 = lean_ctor_get(x_309, 0); lean_inc(x_310); -lean_dec(x_308); -x_311 = lean_environment_main_module(x_310); -x_312 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__1; -lean_inc(x_7); -x_313 = lean_name_mk_string(x_7, x_312); -x_314 = l_Subarray_forInUnsafe_loop___at_Lean_Elab_WF_getForbiddenByTrivialSizeOf___spec__1___closed__3; -x_315 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__2; -lean_inc(x_2); -x_316 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_316, 0, x_314); -lean_ctor_set(x_316, 1, x_2); -lean_ctor_set(x_316, 2, x_315); -x_317 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__3; -lean_inc(x_306); +x_311 = lean_ctor_get(x_309, 1); lean_inc(x_311); -x_318 = l_Lean_addMacroScope(x_311, x_317, x_306); -x_319 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__4; -lean_inc(x_5); -x_320 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_320, 0, x_319); -lean_ctor_set(x_320, 1, x_5); -lean_inc(x_304); -x_321 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_321, 0, x_304); -lean_ctor_set(x_321, 1, x_316); -lean_ctor_set(x_321, 2, x_318); -lean_ctor_set(x_321, 3, x_320); -lean_inc(x_4); -x_322 = l_Lean_addMacroScope(x_311, x_4, x_306); -lean_inc(x_5); +lean_dec(x_309); +x_312 = lean_ctor_get(x_310, 0); +lean_inc(x_312); +lean_dec(x_310); +x_313 = lean_environment_main_module(x_312); +x_314 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__1; +lean_inc(x_8); +x_315 = lean_name_mk_string(x_8, x_314); +x_316 = l_Subarray_forInUnsafe_loop___at_Lean_Elab_WF_getForbiddenByTrivialSizeOf___spec__1___closed__3; +x_317 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__2; lean_inc(x_3); +x_318 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_318, 0, x_316); +lean_ctor_set(x_318, 1, x_3); +lean_ctor_set(x_318, 2, x_317); +x_319 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__3; +lean_inc(x_308); +lean_inc(x_313); +x_320 = l_Lean_addMacroScope(x_313, x_319, x_308); +x_321 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__4; +lean_inc(x_6); +x_322 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_322, 0, x_321); +lean_ctor_set(x_322, 1, x_6); +lean_inc(x_306); x_323 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_323, 0, x_304); -lean_ctor_set(x_323, 1, x_3); -lean_ctor_set(x_323, 2, x_322); -lean_ctor_set(x_323, 3, x_5); -lean_inc(x_8); -x_324 = lean_array_push(x_8, x_323); -x_325 = lean_box(2); -x_326 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__6; -x_327 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_327, 0, x_325); -lean_ctor_set(x_327, 1, x_326); -lean_ctor_set(x_327, 2, x_324); -x_328 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__7; -x_329 = lean_array_push(x_328, x_321); -x_330 = lean_array_push(x_329, x_327); -x_331 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_331, 0, x_325); -lean_ctor_set(x_331, 1, x_313); -lean_ctor_set(x_331, 2, x_330); -x_332 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___lambda__1(x_24, x_300, x_36, x_285, x_294, x_44, x_331, x_15, x_16, x_17, x_18, x_19, x_20, x_309); -lean_dec(x_24); -x_333 = lean_ctor_get(x_332, 0); -lean_inc(x_333); -x_334 = lean_ctor_get(x_332, 1); -lean_inc(x_334); -lean_dec(x_332); -x_25 = x_333; -x_26 = x_334; -goto block_33; +lean_ctor_set(x_323, 0, x_306); +lean_ctor_set(x_323, 1, x_318); +lean_ctor_set(x_323, 2, x_320); +lean_ctor_set(x_323, 3, x_322); +lean_inc(x_5); +x_324 = l_Lean_addMacroScope(x_313, x_5, x_308); +lean_inc(x_6); +lean_inc(x_4); +x_325 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_325, 0, x_306); +lean_ctor_set(x_325, 1, x_4); +lean_ctor_set(x_325, 2, x_324); +lean_ctor_set(x_325, 3, x_6); +lean_inc(x_9); +x_326 = lean_array_push(x_9, x_325); +x_327 = lean_box(2); +x_328 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__6; +x_329 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_329, 0, x_327); +lean_ctor_set(x_329, 1, x_328); +lean_ctor_set(x_329, 2, x_326); +x_330 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__7; +x_331 = lean_array_push(x_330, x_323); +x_332 = lean_array_push(x_331, x_329); +x_333 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_333, 0, x_327); +lean_ctor_set(x_333, 1, x_315); +lean_ctor_set(x_333, 2, x_332); +x_334 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___lambda__1(x_25, x_302, x_37, x_286, x_295, x_45, x_333, x_16, x_17, x_18, x_19, x_20, x_21, x_311); +lean_dec(x_25); +x_335 = lean_ctor_get(x_334, 0); +lean_inc(x_335); +x_336 = lean_ctor_get(x_334, 1); +lean_inc(x_336); +lean_dec(x_334); +x_26 = x_335; +x_27 = x_336; +goto block_34; } else { -lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; -lean_inc(x_19); -x_335 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_19, x_20, x_301); -x_336 = lean_ctor_get(x_335, 0); -lean_inc(x_336); -x_337 = lean_ctor_get(x_335, 1); -lean_inc(x_337); -lean_dec(x_335); -x_338 = lean_ctor_get(x_19, 10); +lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; +lean_inc(x_20); +x_337 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_20, x_21, x_303); +x_338 = lean_ctor_get(x_337, 0); lean_inc(x_338); -x_339 = lean_st_ref_get(x_20, x_337); -x_340 = lean_ctor_get(x_339, 0); +x_339 = lean_ctor_get(x_337, 1); +lean_inc(x_339); +lean_dec(x_337); +x_340 = lean_ctor_get(x_20, 10); lean_inc(x_340); -x_341 = lean_ctor_get(x_339, 1); -lean_inc(x_341); -lean_dec(x_339); -x_342 = lean_ctor_get(x_340, 0); +x_341 = lean_st_ref_get(x_21, x_339); +x_342 = lean_ctor_get(x_341, 0); lean_inc(x_342); -lean_dec(x_340); -x_343 = lean_environment_main_module(x_342); -x_344 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__8; -lean_inc(x_7); -x_345 = lean_name_mk_string(x_7, x_344); -x_346 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__9; -lean_inc(x_336); -x_347 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_347, 0, x_336); -lean_ctor_set(x_347, 1, x_346); -x_348 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__1; -lean_inc(x_7); -x_349 = lean_name_mk_string(x_7, x_348); -x_350 = l_Subarray_forInUnsafe_loop___at_Lean_Elab_WF_getForbiddenByTrivialSizeOf___spec__1___closed__3; -x_351 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__2; -lean_inc(x_2); -x_352 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_352, 0, x_350); -lean_ctor_set(x_352, 1, x_2); -lean_ctor_set(x_352, 2, x_351); -x_353 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__3; -lean_inc(x_338); +x_343 = lean_ctor_get(x_341, 1); lean_inc(x_343); -x_354 = l_Lean_addMacroScope(x_343, x_353, x_338); -x_355 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__4; -lean_inc(x_5); -x_356 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_356, 0, x_355); -lean_ctor_set(x_356, 1, x_5); -lean_inc(x_336); -x_357 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_357, 0, x_336); -lean_ctor_set(x_357, 1, x_352); -lean_ctor_set(x_357, 2, x_354); -lean_ctor_set(x_357, 3, x_356); -lean_inc(x_4); -x_358 = l_Lean_addMacroScope(x_343, x_4, x_338); -lean_inc(x_5); -lean_inc(x_3); -lean_inc(x_336); -x_359 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_359, 0, x_336); -lean_ctor_set(x_359, 1, x_3); -lean_ctor_set(x_359, 2, x_358); -lean_ctor_set(x_359, 3, x_5); +lean_dec(x_341); +x_344 = lean_ctor_get(x_342, 0); +lean_inc(x_344); +lean_dec(x_342); +x_345 = lean_environment_main_module(x_344); +x_346 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__8; lean_inc(x_8); -x_360 = lean_array_push(x_8, x_359); -x_361 = lean_box(2); -x_362 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__6; -x_363 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_363, 0, x_361); -lean_ctor_set(x_363, 1, x_362); -lean_ctor_set(x_363, 2, x_360); -x_364 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__7; -x_365 = lean_array_push(x_364, x_357); -x_366 = lean_array_push(x_365, x_363); -x_367 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_367, 0, x_361); -lean_ctor_set(x_367, 1, x_349); -lean_ctor_set(x_367, 2, x_366); -x_368 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__10; -lean_inc(x_7); -x_369 = lean_name_mk_string(x_7, x_368); -x_370 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__11; -lean_inc(x_336); -x_371 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_371, 0, x_336); -lean_ctor_set(x_371, 1, x_370); -x_372 = l_Nat_repr(x_46); -x_373 = l_Lean_numLitKind; -x_374 = l_Lean_Syntax_mkLit(x_373, x_372, x_361); +x_347 = lean_name_mk_string(x_8, x_346); +x_348 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__9; +lean_inc(x_338); +x_349 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_349, 0, x_338); +lean_ctor_set(x_349, 1, x_348); +x_350 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__1; lean_inc(x_8); -x_375 = lean_array_push(x_8, x_374); -x_376 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_376, 0, x_361); -lean_ctor_set(x_376, 1, x_362); -lean_ctor_set(x_376, 2, x_375); -x_377 = lean_array_push(x_364, x_371); -x_378 = lean_array_push(x_377, x_376); -x_379 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_379, 0, x_361); -lean_ctor_set(x_379, 1, x_369); -lean_ctor_set(x_379, 2, x_378); +x_351 = lean_name_mk_string(x_8, x_350); +x_352 = l_Subarray_forInUnsafe_loop___at_Lean_Elab_WF_getForbiddenByTrivialSizeOf___spec__1___closed__3; +x_353 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__2; +lean_inc(x_3); +x_354 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_354, 0, x_352); +lean_ctor_set(x_354, 1, x_3); +lean_ctor_set(x_354, 2, x_353); +x_355 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__3; +lean_inc(x_340); +lean_inc(x_345); +x_356 = l_Lean_addMacroScope(x_345, x_355, x_340); +x_357 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__4; +lean_inc(x_6); +x_358 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_358, 0, x_357); +lean_ctor_set(x_358, 1, x_6); +lean_inc(x_338); +x_359 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_359, 0, x_338); +lean_ctor_set(x_359, 1, x_354); +lean_ctor_set(x_359, 2, x_356); +lean_ctor_set(x_359, 3, x_358); +lean_inc(x_5); +x_360 = l_Lean_addMacroScope(x_345, x_5, x_340); +lean_inc(x_6); +lean_inc(x_4); +lean_inc(x_338); +x_361 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_361, 0, x_338); +lean_ctor_set(x_361, 1, x_4); +lean_ctor_set(x_361, 2, x_360); +lean_ctor_set(x_361, 3, x_6); +lean_inc(x_9); +x_362 = lean_array_push(x_9, x_361); +x_363 = lean_box(2); +x_364 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__6; +x_365 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_365, 0, x_363); +lean_ctor_set(x_365, 1, x_364); +lean_ctor_set(x_365, 2, x_362); +x_366 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__7; +x_367 = lean_array_push(x_366, x_359); +x_368 = lean_array_push(x_367, x_365); +x_369 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_369, 0, x_363); +lean_ctor_set(x_369, 1, x_351); +lean_ctor_set(x_369, 2, x_368); +x_370 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__10; lean_inc(x_8); -x_380 = lean_array_push(x_8, x_379); -x_381 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_381, 0, x_361); -lean_ctor_set(x_381, 1, x_362); -lean_ctor_set(x_381, 2, x_380); -x_382 = lean_array_push(x_364, x_367); -x_383 = lean_array_push(x_382, x_381); -x_384 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_384, 0, x_361); -lean_ctor_set(x_384, 1, x_362); -lean_ctor_set(x_384, 2, x_383); -x_385 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__12; -x_386 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_386, 0, x_336); -lean_ctor_set(x_386, 1, x_385); -x_387 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__13; -x_388 = lean_array_push(x_387, x_347); -x_389 = lean_array_push(x_388, x_384); -x_390 = lean_array_push(x_389, x_386); -x_391 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_391, 0, x_361); -lean_ctor_set(x_391, 1, x_345); -lean_ctor_set(x_391, 2, x_390); -x_392 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___lambda__1(x_24, x_300, x_36, x_285, x_294, x_44, x_391, x_15, x_16, x_17, x_18, x_19, x_20, x_341); -lean_dec(x_24); -x_393 = lean_ctor_get(x_392, 0); -lean_inc(x_393); -x_394 = lean_ctor_get(x_392, 1); +x_371 = lean_name_mk_string(x_8, x_370); +x_372 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__11; +lean_inc(x_338); +x_373 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_373, 0, x_338); +lean_ctor_set(x_373, 1, x_372); +x_374 = l_Nat_repr(x_47); +x_375 = l_Lean_Syntax_mkNumLit(x_374, x_363); +lean_inc(x_9); +x_376 = lean_array_push(x_9, x_375); +x_377 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_377, 0, x_363); +lean_ctor_set(x_377, 1, x_364); +lean_ctor_set(x_377, 2, x_376); +x_378 = lean_array_push(x_366, x_373); +x_379 = lean_array_push(x_378, x_377); +x_380 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_380, 0, x_363); +lean_ctor_set(x_380, 1, x_371); +lean_ctor_set(x_380, 2, x_379); +lean_inc(x_9); +x_381 = lean_array_push(x_9, x_380); +x_382 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_382, 0, x_363); +lean_ctor_set(x_382, 1, x_364); +lean_ctor_set(x_382, 2, x_381); +x_383 = lean_array_push(x_366, x_369); +x_384 = lean_array_push(x_383, x_382); +x_385 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_385, 0, x_363); +lean_ctor_set(x_385, 1, x_364); +lean_ctor_set(x_385, 2, x_384); +x_386 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__12; +x_387 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_387, 0, x_338); +lean_ctor_set(x_387, 1, x_386); +x_388 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__13; +x_389 = lean_array_push(x_388, x_349); +x_390 = lean_array_push(x_389, x_385); +x_391 = lean_array_push(x_390, x_387); +x_392 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_392, 0, x_363); +lean_ctor_set(x_392, 1, x_347); +lean_ctor_set(x_392, 2, x_391); +x_393 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___lambda__1(x_25, x_302, x_37, x_286, x_295, x_45, x_392, x_16, x_17, x_18, x_19, x_20, x_21, x_343); +lean_dec(x_25); +x_394 = lean_ctor_get(x_393, 0); lean_inc(x_394); -lean_dec(x_392); -x_25 = x_393; +x_395 = lean_ctor_get(x_393, 1); +lean_inc(x_395); +lean_dec(x_393); x_26 = x_394; -goto block_33; +x_27 = x_395; +goto block_34; } } } @@ -7397,348 +7396,347 @@ goto block_33; } else { -lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; uint8_t x_400; -lean_dec(x_36); -x_395 = lean_nat_add(x_46, x_48); -x_396 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_396, 0, x_395); -lean_ctor_set(x_396, 1, x_47); -lean_ctor_set(x_396, 2, x_48); -x_397 = lean_ctor_get(x_41, 0); -lean_inc(x_397); -x_398 = lean_ctor_get(x_41, 1); +lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; uint8_t x_401; +lean_dec(x_37); +x_396 = lean_nat_add(x_47, x_49); +x_397 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_397, 0, x_396); +lean_ctor_set(x_397, 1, x_48); +lean_ctor_set(x_397, 2, x_49); +x_398 = lean_ctor_get(x_42, 0); lean_inc(x_398); -x_399 = lean_ctor_get(x_41, 2); +x_399 = lean_ctor_get(x_42, 1); lean_inc(x_399); -x_400 = lean_nat_dec_lt(x_398, x_399); -if (x_400 == 0) +x_400 = lean_ctor_get(x_42, 2); +lean_inc(x_400); +x_401 = lean_nat_dec_lt(x_399, x_400); +if (x_401 == 0) { -lean_object* x_401; +lean_object* x_402; +lean_dec(x_400); lean_dec(x_399); lean_dec(x_398); -lean_dec(x_397); -lean_dec(x_46); -lean_dec(x_24); -lean_ctor_set(x_35, 0, x_396); -x_401 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_401, 0, x_14); -x_25 = x_401; -x_26 = x_21; -goto block_33; +lean_dec(x_47); +lean_dec(x_25); +lean_ctor_set(x_36, 0, x_397); +x_402 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_402, 0, x_15); +x_26 = x_402; +x_27 = x_22; +goto block_34; } else { -lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; uint8_t x_410; -if (lean_is_exclusive(x_41)) { - lean_ctor_release(x_41, 0); - lean_ctor_release(x_41, 1); - lean_ctor_release(x_41, 2); - x_402 = x_41; +lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; uint8_t x_411; +if (lean_is_exclusive(x_42)) { + lean_ctor_release(x_42, 0); + lean_ctor_release(x_42, 1); + lean_ctor_release(x_42, 2); + x_403 = x_42; } else { - lean_dec_ref(x_41); - x_402 = lean_box(0); + lean_dec_ref(x_42); + x_403 = lean_box(0); } -x_403 = lean_array_fget(x_397, x_398); -x_404 = lean_unsigned_to_nat(1u); -x_405 = lean_nat_add(x_398, x_404); -lean_dec(x_398); -if (lean_is_scalar(x_402)) { - x_406 = lean_alloc_ctor(0, 3, 0); +x_404 = lean_array_fget(x_398, x_399); +x_405 = lean_unsigned_to_nat(1u); +x_406 = lean_nat_add(x_399, x_405); +lean_dec(x_399); +if (lean_is_scalar(x_403)) { + x_407 = lean_alloc_ctor(0, 3, 0); } else { - x_406 = x_402; -} -lean_ctor_set(x_406, 0, x_397); -lean_ctor_set(x_406, 1, x_405); -lean_ctor_set(x_406, 2, x_399); -x_407 = lean_ctor_get(x_38, 0); -lean_inc(x_407); -x_408 = lean_ctor_get(x_38, 1); + x_407 = x_403; +} +lean_ctor_set(x_407, 0, x_398); +lean_ctor_set(x_407, 1, x_406); +lean_ctor_set(x_407, 2, x_400); +x_408 = lean_ctor_get(x_39, 0); lean_inc(x_408); -x_409 = lean_ctor_get(x_38, 2); +x_409 = lean_ctor_get(x_39, 1); lean_inc(x_409); -x_410 = lean_nat_dec_lt(x_408, x_409); -if (x_410 == 0) +x_410 = lean_ctor_get(x_39, 2); +lean_inc(x_410); +x_411 = lean_nat_dec_lt(x_409, x_410); +if (x_411 == 0) { -lean_object* x_411; +lean_object* x_412; +lean_dec(x_410); lean_dec(x_409); lean_dec(x_408); -lean_dec(x_407); -lean_dec(x_403); -lean_dec(x_46); -lean_dec(x_24); -lean_ctor_set(x_35, 0, x_396); -lean_ctor_set(x_34, 0, x_406); -x_411 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_411, 0, x_14); -x_25 = x_411; -x_26 = x_21; -goto block_33; +lean_dec(x_404); +lean_dec(x_47); +lean_dec(x_25); +lean_ctor_set(x_36, 0, x_397); +lean_ctor_set(x_35, 0, x_407); +x_412 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_412, 0, x_15); +x_26 = x_412; +x_27 = x_22; +goto block_34; } else { -lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; uint8_t x_423; +lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; uint8_t x_425; +lean_free_object(x_36); lean_free_object(x_35); -lean_free_object(x_34); -lean_free_object(x_14); -if (lean_is_exclusive(x_38)) { - lean_ctor_release(x_38, 0); - lean_ctor_release(x_38, 1); - lean_ctor_release(x_38, 2); - x_412 = x_38; +lean_free_object(x_15); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + lean_ctor_release(x_39, 2); + x_413 = x_39; } else { - lean_dec_ref(x_38); - x_412 = lean_box(0); + lean_dec_ref(x_39); + x_413 = lean_box(0); } -x_413 = lean_array_fget(x_407, x_408); -x_414 = lean_nat_add(x_408, x_404); -lean_dec(x_408); -if (lean_is_scalar(x_412)) { - x_415 = lean_alloc_ctor(0, 3, 0); +x_414 = lean_array_fget(x_408, x_409); +x_415 = lean_nat_add(x_409, x_405); +lean_dec(x_409); +if (lean_is_scalar(x_413)) { + x_416 = lean_alloc_ctor(0, 3, 0); } else { - x_415 = x_412; + x_416 = x_413; } -lean_ctor_set(x_415, 0, x_407); -lean_ctor_set(x_415, 1, x_414); -lean_ctor_set(x_415, 2, x_409); -lean_inc(x_6); -lean_inc(x_8); -x_416 = lean_array_push(x_8, x_6); -x_417 = lean_nat_sub(x_413, x_403); -lean_dec(x_403); -lean_dec(x_413); -x_418 = lean_nat_sub(x_417, x_404); -lean_dec(x_417); -x_419 = lean_unsigned_to_nat(0u); -lean_inc(x_418); -lean_inc(x_9); -x_420 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__1(x_9, x_418, x_419, x_418, x_404, x_416, x_15, x_16, x_17, x_18, x_19, x_20, x_21); -lean_dec(x_418); -x_421 = lean_ctor_get(x_420, 0); -lean_inc(x_421); -x_422 = lean_ctor_get(x_420, 1); -lean_inc(x_422); +lean_ctor_set(x_416, 0, x_408); +lean_ctor_set(x_416, 1, x_415); +lean_ctor_set(x_416, 2, x_410); +x_417 = l___private_Lean_Elab_PreDefinition_WF_Rel_0__Lean_Elab_WF_unpackUnary_go___lambda__1___closed__1; +lean_inc(x_7); +x_418 = lean_array_push(x_417, x_7); +x_419 = lean_nat_sub(x_414, x_404); +lean_dec(x_404); +lean_dec(x_414); +x_420 = lean_nat_sub(x_419, x_405); +lean_dec(x_419); +x_421 = lean_unsigned_to_nat(0u); +lean_inc(x_420); +lean_inc(x_10); +x_422 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__1(x_2, x_10, x_420, x_421, x_420, x_405, x_418, x_16, x_17, x_18, x_19, x_20, x_21, x_22); lean_dec(x_420); -x_423 = lean_nat_dec_lt(x_404, x_10); -if (x_423 == 0) -{ -lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; -lean_dec(x_46); -lean_inc(x_19); -x_424 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_19, x_20, x_422); -x_425 = lean_ctor_get(x_424, 0); -lean_inc(x_425); -x_426 = lean_ctor_get(x_424, 1); -lean_inc(x_426); -lean_dec(x_424); -x_427 = lean_ctor_get(x_19, 10); +x_423 = lean_ctor_get(x_422, 0); +lean_inc(x_423); +x_424 = lean_ctor_get(x_422, 1); +lean_inc(x_424); +lean_dec(x_422); +x_425 = lean_nat_dec_lt(x_405, x_11); +if (x_425 == 0) +{ +lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; +lean_dec(x_47); +lean_inc(x_20); +x_426 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_20, x_21, x_424); +x_427 = lean_ctor_get(x_426, 0); lean_inc(x_427); -x_428 = lean_st_ref_get(x_20, x_426); -x_429 = lean_ctor_get(x_428, 0); +x_428 = lean_ctor_get(x_426, 1); +lean_inc(x_428); +lean_dec(x_426); +x_429 = lean_ctor_get(x_20, 10); lean_inc(x_429); -x_430 = lean_ctor_get(x_428, 1); -lean_inc(x_430); -lean_dec(x_428); -x_431 = lean_ctor_get(x_429, 0); +x_430 = lean_st_ref_get(x_21, x_428); +x_431 = lean_ctor_get(x_430, 0); lean_inc(x_431); -lean_dec(x_429); -x_432 = lean_environment_main_module(x_431); -x_433 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__1; -lean_inc(x_7); -x_434 = lean_name_mk_string(x_7, x_433); -x_435 = l_Subarray_forInUnsafe_loop___at_Lean_Elab_WF_getForbiddenByTrivialSizeOf___spec__1___closed__3; -x_436 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__2; -lean_inc(x_2); -x_437 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_437, 0, x_435); -lean_ctor_set(x_437, 1, x_2); -lean_ctor_set(x_437, 2, x_436); -x_438 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__3; -lean_inc(x_427); +x_432 = lean_ctor_get(x_430, 1); lean_inc(x_432); -x_439 = l_Lean_addMacroScope(x_432, x_438, x_427); -x_440 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__4; -lean_inc(x_5); -x_441 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_441, 0, x_440); -lean_ctor_set(x_441, 1, x_5); -lean_inc(x_425); -x_442 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_442, 0, x_425); -lean_ctor_set(x_442, 1, x_437); -lean_ctor_set(x_442, 2, x_439); -lean_ctor_set(x_442, 3, x_441); -lean_inc(x_4); -x_443 = l_Lean_addMacroScope(x_432, x_4, x_427); -lean_inc(x_5); +lean_dec(x_430); +x_433 = lean_ctor_get(x_431, 0); +lean_inc(x_433); +lean_dec(x_431); +x_434 = lean_environment_main_module(x_433); +x_435 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__1; +lean_inc(x_8); +x_436 = lean_name_mk_string(x_8, x_435); +x_437 = l_Subarray_forInUnsafe_loop___at_Lean_Elab_WF_getForbiddenByTrivialSizeOf___spec__1___closed__3; +x_438 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__2; lean_inc(x_3); +x_439 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_439, 0, x_437); +lean_ctor_set(x_439, 1, x_3); +lean_ctor_set(x_439, 2, x_438); +x_440 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__3; +lean_inc(x_429); +lean_inc(x_434); +x_441 = l_Lean_addMacroScope(x_434, x_440, x_429); +x_442 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__4; +lean_inc(x_6); +x_443 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_443, 0, x_442); +lean_ctor_set(x_443, 1, x_6); +lean_inc(x_427); x_444 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_444, 0, x_425); -lean_ctor_set(x_444, 1, x_3); -lean_ctor_set(x_444, 2, x_443); -lean_ctor_set(x_444, 3, x_5); -lean_inc(x_8); -x_445 = lean_array_push(x_8, x_444); -x_446 = lean_box(2); -x_447 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__6; -x_448 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_448, 0, x_446); -lean_ctor_set(x_448, 1, x_447); -lean_ctor_set(x_448, 2, x_445); -x_449 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__7; -x_450 = lean_array_push(x_449, x_442); -x_451 = lean_array_push(x_450, x_448); -x_452 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_452, 0, x_446); -lean_ctor_set(x_452, 1, x_434); -lean_ctor_set(x_452, 2, x_451); -x_453 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___lambda__1(x_24, x_421, x_396, x_406, x_415, x_44, x_452, x_15, x_16, x_17, x_18, x_19, x_20, x_430); -lean_dec(x_24); -x_454 = lean_ctor_get(x_453, 0); -lean_inc(x_454); -x_455 = lean_ctor_get(x_453, 1); -lean_inc(x_455); -lean_dec(x_453); -x_25 = x_454; -x_26 = x_455; -goto block_33; +lean_ctor_set(x_444, 0, x_427); +lean_ctor_set(x_444, 1, x_439); +lean_ctor_set(x_444, 2, x_441); +lean_ctor_set(x_444, 3, x_443); +lean_inc(x_5); +x_445 = l_Lean_addMacroScope(x_434, x_5, x_429); +lean_inc(x_6); +lean_inc(x_4); +x_446 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_446, 0, x_427); +lean_ctor_set(x_446, 1, x_4); +lean_ctor_set(x_446, 2, x_445); +lean_ctor_set(x_446, 3, x_6); +lean_inc(x_9); +x_447 = lean_array_push(x_9, x_446); +x_448 = lean_box(2); +x_449 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__6; +x_450 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_450, 0, x_448); +lean_ctor_set(x_450, 1, x_449); +lean_ctor_set(x_450, 2, x_447); +x_451 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__7; +x_452 = lean_array_push(x_451, x_444); +x_453 = lean_array_push(x_452, x_450); +x_454 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_454, 0, x_448); +lean_ctor_set(x_454, 1, x_436); +lean_ctor_set(x_454, 2, x_453); +x_455 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___lambda__1(x_25, x_423, x_397, x_407, x_416, x_45, x_454, x_16, x_17, x_18, x_19, x_20, x_21, x_432); +lean_dec(x_25); +x_456 = lean_ctor_get(x_455, 0); +lean_inc(x_456); +x_457 = lean_ctor_get(x_455, 1); +lean_inc(x_457); +lean_dec(x_455); +x_26 = x_456; +x_27 = x_457; +goto block_34; } else { -lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; -lean_inc(x_19); -x_456 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_19, x_20, x_422); -x_457 = lean_ctor_get(x_456, 0); -lean_inc(x_457); -x_458 = lean_ctor_get(x_456, 1); -lean_inc(x_458); -lean_dec(x_456); -x_459 = lean_ctor_get(x_19, 10); +lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; +lean_inc(x_20); +x_458 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_20, x_21, x_424); +x_459 = lean_ctor_get(x_458, 0); lean_inc(x_459); -x_460 = lean_st_ref_get(x_20, x_458); -x_461 = lean_ctor_get(x_460, 0); +x_460 = lean_ctor_get(x_458, 1); +lean_inc(x_460); +lean_dec(x_458); +x_461 = lean_ctor_get(x_20, 10); lean_inc(x_461); -x_462 = lean_ctor_get(x_460, 1); -lean_inc(x_462); -lean_dec(x_460); -x_463 = lean_ctor_get(x_461, 0); +x_462 = lean_st_ref_get(x_21, x_460); +x_463 = lean_ctor_get(x_462, 0); lean_inc(x_463); -lean_dec(x_461); -x_464 = lean_environment_main_module(x_463); -x_465 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__8; -lean_inc(x_7); -x_466 = lean_name_mk_string(x_7, x_465); -x_467 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__9; -lean_inc(x_457); -x_468 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_468, 0, x_457); -lean_ctor_set(x_468, 1, x_467); -x_469 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__1; -lean_inc(x_7); -x_470 = lean_name_mk_string(x_7, x_469); -x_471 = l_Subarray_forInUnsafe_loop___at_Lean_Elab_WF_getForbiddenByTrivialSizeOf___spec__1___closed__3; -x_472 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__2; -lean_inc(x_2); -x_473 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_473, 0, x_471); -lean_ctor_set(x_473, 1, x_2); -lean_ctor_set(x_473, 2, x_472); -x_474 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__3; -lean_inc(x_459); +x_464 = lean_ctor_get(x_462, 1); lean_inc(x_464); -x_475 = l_Lean_addMacroScope(x_464, x_474, x_459); -x_476 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__4; -lean_inc(x_5); -x_477 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_477, 0, x_476); -lean_ctor_set(x_477, 1, x_5); -lean_inc(x_457); -x_478 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_478, 0, x_457); -lean_ctor_set(x_478, 1, x_473); -lean_ctor_set(x_478, 2, x_475); -lean_ctor_set(x_478, 3, x_477); -lean_inc(x_4); -x_479 = l_Lean_addMacroScope(x_464, x_4, x_459); -lean_inc(x_5); -lean_inc(x_3); -lean_inc(x_457); -x_480 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_480, 0, x_457); -lean_ctor_set(x_480, 1, x_3); -lean_ctor_set(x_480, 2, x_479); -lean_ctor_set(x_480, 3, x_5); +lean_dec(x_462); +x_465 = lean_ctor_get(x_463, 0); +lean_inc(x_465); +lean_dec(x_463); +x_466 = lean_environment_main_module(x_465); +x_467 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__8; lean_inc(x_8); -x_481 = lean_array_push(x_8, x_480); -x_482 = lean_box(2); -x_483 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__6; -x_484 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_484, 0, x_482); -lean_ctor_set(x_484, 1, x_483); -lean_ctor_set(x_484, 2, x_481); -x_485 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__7; -x_486 = lean_array_push(x_485, x_478); -x_487 = lean_array_push(x_486, x_484); -x_488 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_488, 0, x_482); -lean_ctor_set(x_488, 1, x_470); -lean_ctor_set(x_488, 2, x_487); -x_489 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__10; -lean_inc(x_7); -x_490 = lean_name_mk_string(x_7, x_489); -x_491 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__11; -lean_inc(x_457); -x_492 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_492, 0, x_457); -lean_ctor_set(x_492, 1, x_491); -x_493 = l_Nat_repr(x_46); -x_494 = l_Lean_numLitKind; -x_495 = l_Lean_Syntax_mkLit(x_494, x_493, x_482); +x_468 = lean_name_mk_string(x_8, x_467); +x_469 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__9; +lean_inc(x_459); +x_470 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_470, 0, x_459); +lean_ctor_set(x_470, 1, x_469); +x_471 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__1; lean_inc(x_8); -x_496 = lean_array_push(x_8, x_495); -x_497 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_497, 0, x_482); -lean_ctor_set(x_497, 1, x_483); -lean_ctor_set(x_497, 2, x_496); -x_498 = lean_array_push(x_485, x_492); -x_499 = lean_array_push(x_498, x_497); -x_500 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_500, 0, x_482); -lean_ctor_set(x_500, 1, x_490); -lean_ctor_set(x_500, 2, x_499); +x_472 = lean_name_mk_string(x_8, x_471); +x_473 = l_Subarray_forInUnsafe_loop___at_Lean_Elab_WF_getForbiddenByTrivialSizeOf___spec__1___closed__3; +x_474 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__2; +lean_inc(x_3); +x_475 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_475, 0, x_473); +lean_ctor_set(x_475, 1, x_3); +lean_ctor_set(x_475, 2, x_474); +x_476 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__3; +lean_inc(x_461); +lean_inc(x_466); +x_477 = l_Lean_addMacroScope(x_466, x_476, x_461); +x_478 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__4; +lean_inc(x_6); +x_479 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_479, 0, x_478); +lean_ctor_set(x_479, 1, x_6); +lean_inc(x_459); +x_480 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_480, 0, x_459); +lean_ctor_set(x_480, 1, x_475); +lean_ctor_set(x_480, 2, x_477); +lean_ctor_set(x_480, 3, x_479); +lean_inc(x_5); +x_481 = l_Lean_addMacroScope(x_466, x_5, x_461); +lean_inc(x_6); +lean_inc(x_4); +lean_inc(x_459); +x_482 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_482, 0, x_459); +lean_ctor_set(x_482, 1, x_4); +lean_ctor_set(x_482, 2, x_481); +lean_ctor_set(x_482, 3, x_6); +lean_inc(x_9); +x_483 = lean_array_push(x_9, x_482); +x_484 = lean_box(2); +x_485 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__6; +x_486 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_486, 0, x_484); +lean_ctor_set(x_486, 1, x_485); +lean_ctor_set(x_486, 2, x_483); +x_487 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__7; +x_488 = lean_array_push(x_487, x_480); +x_489 = lean_array_push(x_488, x_486); +x_490 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_490, 0, x_484); +lean_ctor_set(x_490, 1, x_472); +lean_ctor_set(x_490, 2, x_489); +x_491 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__10; lean_inc(x_8); -x_501 = lean_array_push(x_8, x_500); -x_502 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_502, 0, x_482); -lean_ctor_set(x_502, 1, x_483); -lean_ctor_set(x_502, 2, x_501); -x_503 = lean_array_push(x_485, x_488); -x_504 = lean_array_push(x_503, x_502); -x_505 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_505, 0, x_482); -lean_ctor_set(x_505, 1, x_483); -lean_ctor_set(x_505, 2, x_504); -x_506 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__12; -x_507 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_507, 0, x_457); -lean_ctor_set(x_507, 1, x_506); -x_508 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__13; -x_509 = lean_array_push(x_508, x_468); -x_510 = lean_array_push(x_509, x_505); -x_511 = lean_array_push(x_510, x_507); -x_512 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_512, 0, x_482); -lean_ctor_set(x_512, 1, x_466); -lean_ctor_set(x_512, 2, x_511); -x_513 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___lambda__1(x_24, x_421, x_396, x_406, x_415, x_44, x_512, x_15, x_16, x_17, x_18, x_19, x_20, x_462); -lean_dec(x_24); -x_514 = lean_ctor_get(x_513, 0); -lean_inc(x_514); -x_515 = lean_ctor_get(x_513, 1); +x_492 = lean_name_mk_string(x_8, x_491); +x_493 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__11; +lean_inc(x_459); +x_494 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_494, 0, x_459); +lean_ctor_set(x_494, 1, x_493); +x_495 = l_Nat_repr(x_47); +x_496 = l_Lean_Syntax_mkNumLit(x_495, x_484); +lean_inc(x_9); +x_497 = lean_array_push(x_9, x_496); +x_498 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_498, 0, x_484); +lean_ctor_set(x_498, 1, x_485); +lean_ctor_set(x_498, 2, x_497); +x_499 = lean_array_push(x_487, x_494); +x_500 = lean_array_push(x_499, x_498); +x_501 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_501, 0, x_484); +lean_ctor_set(x_501, 1, x_492); +lean_ctor_set(x_501, 2, x_500); +lean_inc(x_9); +x_502 = lean_array_push(x_9, x_501); +x_503 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_503, 0, x_484); +lean_ctor_set(x_503, 1, x_485); +lean_ctor_set(x_503, 2, x_502); +x_504 = lean_array_push(x_487, x_490); +x_505 = lean_array_push(x_504, x_503); +x_506 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_506, 0, x_484); +lean_ctor_set(x_506, 1, x_485); +lean_ctor_set(x_506, 2, x_505); +x_507 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__12; +x_508 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_508, 0, x_459); +lean_ctor_set(x_508, 1, x_507); +x_509 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__13; +x_510 = lean_array_push(x_509, x_470); +x_511 = lean_array_push(x_510, x_506); +x_512 = lean_array_push(x_511, x_508); +x_513 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_513, 0, x_484); +lean_ctor_set(x_513, 1, x_468); +lean_ctor_set(x_513, 2, x_512); +x_514 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___lambda__1(x_25, x_423, x_397, x_407, x_416, x_45, x_513, x_16, x_17, x_18, x_19, x_20, x_21, x_464); +lean_dec(x_25); +x_515 = lean_ctor_get(x_514, 0); lean_inc(x_515); -lean_dec(x_513); -x_25 = x_514; +x_516 = lean_ctor_get(x_514, 1); +lean_inc(x_516); +lean_dec(x_514); x_26 = x_515; -goto block_33; +x_27 = x_516; +goto block_34; } } } @@ -7747,395 +7745,394 @@ goto block_33; } else { -lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; uint8_t x_520; -x_516 = lean_ctor_get(x_35, 1); -lean_inc(x_516); -lean_dec(x_35); -x_517 = lean_ctor_get(x_36, 0); +lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; uint8_t x_521; +x_517 = lean_ctor_get(x_36, 1); lean_inc(x_517); -x_518 = lean_ctor_get(x_36, 1); +lean_dec(x_36); +x_518 = lean_ctor_get(x_37, 0); lean_inc(x_518); -x_519 = lean_ctor_get(x_36, 2); +x_519 = lean_ctor_get(x_37, 1); lean_inc(x_519); -x_520 = lean_nat_dec_lt(x_517, x_518); -if (x_520 == 0) +x_520 = lean_ctor_get(x_37, 2); +lean_inc(x_520); +x_521 = lean_nat_dec_lt(x_518, x_519); +if (x_521 == 0) { -lean_object* x_521; lean_object* x_522; +lean_object* x_522; lean_object* x_523; +lean_dec(x_520); lean_dec(x_519); lean_dec(x_518); -lean_dec(x_517); -lean_dec(x_24); -x_521 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_521, 0, x_36); -lean_ctor_set(x_521, 1, x_516); -lean_ctor_set(x_34, 1, x_521); -x_522 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_522, 0, x_14); -x_25 = x_522; -x_26 = x_21; -goto block_33; -} -else -{ -lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; uint8_t x_529; -if (lean_is_exclusive(x_36)) { - lean_ctor_release(x_36, 0); - lean_ctor_release(x_36, 1); - lean_ctor_release(x_36, 2); - x_523 = x_36; +lean_dec(x_25); +x_522 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_522, 0, x_37); +lean_ctor_set(x_522, 1, x_517); +lean_ctor_set(x_35, 1, x_522); +x_523 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_523, 0, x_15); +x_26 = x_523; +x_27 = x_22; +goto block_34; +} +else +{ +lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; uint8_t x_530; +if (lean_is_exclusive(x_37)) { + lean_ctor_release(x_37, 0); + lean_ctor_release(x_37, 1); + lean_ctor_release(x_37, 2); + x_524 = x_37; } else { - lean_dec_ref(x_36); - x_523 = lean_box(0); + lean_dec_ref(x_37); + x_524 = lean_box(0); } -x_524 = lean_nat_add(x_517, x_519); -if (lean_is_scalar(x_523)) { - x_525 = lean_alloc_ctor(0, 3, 0); +x_525 = lean_nat_add(x_518, x_520); +if (lean_is_scalar(x_524)) { + x_526 = lean_alloc_ctor(0, 3, 0); } else { - x_525 = x_523; -} -lean_ctor_set(x_525, 0, x_524); -lean_ctor_set(x_525, 1, x_518); -lean_ctor_set(x_525, 2, x_519); -x_526 = lean_ctor_get(x_41, 0); -lean_inc(x_526); -x_527 = lean_ctor_get(x_41, 1); + x_526 = x_524; +} +lean_ctor_set(x_526, 0, x_525); +lean_ctor_set(x_526, 1, x_519); +lean_ctor_set(x_526, 2, x_520); +x_527 = lean_ctor_get(x_42, 0); lean_inc(x_527); -x_528 = lean_ctor_get(x_41, 2); +x_528 = lean_ctor_get(x_42, 1); lean_inc(x_528); -x_529 = lean_nat_dec_lt(x_527, x_528); -if (x_529 == 0) +x_529 = lean_ctor_get(x_42, 2); +lean_inc(x_529); +x_530 = lean_nat_dec_lt(x_528, x_529); +if (x_530 == 0) { -lean_object* x_530; lean_object* x_531; +lean_object* x_531; lean_object* x_532; +lean_dec(x_529); lean_dec(x_528); lean_dec(x_527); -lean_dec(x_526); -lean_dec(x_517); -lean_dec(x_24); -x_530 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_530, 0, x_525); -lean_ctor_set(x_530, 1, x_516); -lean_ctor_set(x_34, 1, x_530); -x_531 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_531, 0, x_14); -x_25 = x_531; -x_26 = x_21; -goto block_33; -} -else -{ -lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; uint8_t x_540; -if (lean_is_exclusive(x_41)) { - lean_ctor_release(x_41, 0); - lean_ctor_release(x_41, 1); - lean_ctor_release(x_41, 2); - x_532 = x_41; +lean_dec(x_518); +lean_dec(x_25); +x_531 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_531, 0, x_526); +lean_ctor_set(x_531, 1, x_517); +lean_ctor_set(x_35, 1, x_531); +x_532 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_532, 0, x_15); +x_26 = x_532; +x_27 = x_22; +goto block_34; +} +else +{ +lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; uint8_t x_541; +if (lean_is_exclusive(x_42)) { + lean_ctor_release(x_42, 0); + lean_ctor_release(x_42, 1); + lean_ctor_release(x_42, 2); + x_533 = x_42; } else { - lean_dec_ref(x_41); - x_532 = lean_box(0); + lean_dec_ref(x_42); + x_533 = lean_box(0); } -x_533 = lean_array_fget(x_526, x_527); -x_534 = lean_unsigned_to_nat(1u); -x_535 = lean_nat_add(x_527, x_534); -lean_dec(x_527); -if (lean_is_scalar(x_532)) { - x_536 = lean_alloc_ctor(0, 3, 0); +x_534 = lean_array_fget(x_527, x_528); +x_535 = lean_unsigned_to_nat(1u); +x_536 = lean_nat_add(x_528, x_535); +lean_dec(x_528); +if (lean_is_scalar(x_533)) { + x_537 = lean_alloc_ctor(0, 3, 0); } else { - x_536 = x_532; -} -lean_ctor_set(x_536, 0, x_526); -lean_ctor_set(x_536, 1, x_535); -lean_ctor_set(x_536, 2, x_528); -x_537 = lean_ctor_get(x_38, 0); -lean_inc(x_537); -x_538 = lean_ctor_get(x_38, 1); + x_537 = x_533; +} +lean_ctor_set(x_537, 0, x_527); +lean_ctor_set(x_537, 1, x_536); +lean_ctor_set(x_537, 2, x_529); +x_538 = lean_ctor_get(x_39, 0); lean_inc(x_538); -x_539 = lean_ctor_get(x_38, 2); +x_539 = lean_ctor_get(x_39, 1); lean_inc(x_539); -x_540 = lean_nat_dec_lt(x_538, x_539); -if (x_540 == 0) +x_540 = lean_ctor_get(x_39, 2); +lean_inc(x_540); +x_541 = lean_nat_dec_lt(x_539, x_540); +if (x_541 == 0) { -lean_object* x_541; lean_object* x_542; +lean_object* x_542; lean_object* x_543; +lean_dec(x_540); lean_dec(x_539); lean_dec(x_538); -lean_dec(x_537); -lean_dec(x_533); -lean_dec(x_517); -lean_dec(x_24); -x_541 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_541, 0, x_525); -lean_ctor_set(x_541, 1, x_516); -lean_ctor_set(x_34, 1, x_541); -lean_ctor_set(x_34, 0, x_536); -x_542 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_542, 0, x_14); -x_25 = x_542; -x_26 = x_21; -goto block_33; -} -else -{ -lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; uint8_t x_554; -lean_free_object(x_34); -lean_free_object(x_14); -if (lean_is_exclusive(x_38)) { - lean_ctor_release(x_38, 0); - lean_ctor_release(x_38, 1); - lean_ctor_release(x_38, 2); - x_543 = x_38; +lean_dec(x_534); +lean_dec(x_518); +lean_dec(x_25); +x_542 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_542, 0, x_526); +lean_ctor_set(x_542, 1, x_517); +lean_ctor_set(x_35, 1, x_542); +lean_ctor_set(x_35, 0, x_537); +x_543 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_543, 0, x_15); +x_26 = x_543; +x_27 = x_22; +goto block_34; +} +else +{ +lean_object* x_544; lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; uint8_t x_556; +lean_free_object(x_35); +lean_free_object(x_15); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + lean_ctor_release(x_39, 2); + x_544 = x_39; } else { - lean_dec_ref(x_38); - x_543 = lean_box(0); + lean_dec_ref(x_39); + x_544 = lean_box(0); } -x_544 = lean_array_fget(x_537, x_538); -x_545 = lean_nat_add(x_538, x_534); -lean_dec(x_538); -if (lean_is_scalar(x_543)) { - x_546 = lean_alloc_ctor(0, 3, 0); +x_545 = lean_array_fget(x_538, x_539); +x_546 = lean_nat_add(x_539, x_535); +lean_dec(x_539); +if (lean_is_scalar(x_544)) { + x_547 = lean_alloc_ctor(0, 3, 0); } else { - x_546 = x_543; + x_547 = x_544; } -lean_ctor_set(x_546, 0, x_537); -lean_ctor_set(x_546, 1, x_545); -lean_ctor_set(x_546, 2, x_539); -lean_inc(x_6); -lean_inc(x_8); -x_547 = lean_array_push(x_8, x_6); -x_548 = lean_nat_sub(x_544, x_533); -lean_dec(x_533); -lean_dec(x_544); -x_549 = lean_nat_sub(x_548, x_534); -lean_dec(x_548); -x_550 = lean_unsigned_to_nat(0u); -lean_inc(x_549); -lean_inc(x_9); -x_551 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__1(x_9, x_549, x_550, x_549, x_534, x_547, x_15, x_16, x_17, x_18, x_19, x_20, x_21); -lean_dec(x_549); -x_552 = lean_ctor_get(x_551, 0); -lean_inc(x_552); -x_553 = lean_ctor_get(x_551, 1); -lean_inc(x_553); +lean_ctor_set(x_547, 0, x_538); +lean_ctor_set(x_547, 1, x_546); +lean_ctor_set(x_547, 2, x_540); +x_548 = l___private_Lean_Elab_PreDefinition_WF_Rel_0__Lean_Elab_WF_unpackUnary_go___lambda__1___closed__1; +lean_inc(x_7); +x_549 = lean_array_push(x_548, x_7); +x_550 = lean_nat_sub(x_545, x_534); +lean_dec(x_534); +lean_dec(x_545); +x_551 = lean_nat_sub(x_550, x_535); +lean_dec(x_550); +x_552 = lean_unsigned_to_nat(0u); +lean_inc(x_551); +lean_inc(x_10); +x_553 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__1(x_2, x_10, x_551, x_552, x_551, x_535, x_549, x_16, x_17, x_18, x_19, x_20, x_21, x_22); lean_dec(x_551); -x_554 = lean_nat_dec_lt(x_534, x_10); -if (x_554 == 0) -{ -lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; lean_object* x_565; lean_object* x_566; lean_object* x_567; lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; -lean_dec(x_517); -lean_inc(x_19); -x_555 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_19, x_20, x_553); -x_556 = lean_ctor_get(x_555, 0); -lean_inc(x_556); -x_557 = lean_ctor_get(x_555, 1); -lean_inc(x_557); -lean_dec(x_555); -x_558 = lean_ctor_get(x_19, 10); +x_554 = lean_ctor_get(x_553, 0); +lean_inc(x_554); +x_555 = lean_ctor_get(x_553, 1); +lean_inc(x_555); +lean_dec(x_553); +x_556 = lean_nat_dec_lt(x_535, x_11); +if (x_556 == 0) +{ +lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; lean_object* x_565; lean_object* x_566; lean_object* x_567; lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; +lean_dec(x_518); +lean_inc(x_20); +x_557 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_20, x_21, x_555); +x_558 = lean_ctor_get(x_557, 0); lean_inc(x_558); -x_559 = lean_st_ref_get(x_20, x_557); -x_560 = lean_ctor_get(x_559, 0); +x_559 = lean_ctor_get(x_557, 1); +lean_inc(x_559); +lean_dec(x_557); +x_560 = lean_ctor_get(x_20, 10); lean_inc(x_560); -x_561 = lean_ctor_get(x_559, 1); -lean_inc(x_561); -lean_dec(x_559); -x_562 = lean_ctor_get(x_560, 0); +x_561 = lean_st_ref_get(x_21, x_559); +x_562 = lean_ctor_get(x_561, 0); lean_inc(x_562); -lean_dec(x_560); -x_563 = lean_environment_main_module(x_562); -x_564 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__1; -lean_inc(x_7); -x_565 = lean_name_mk_string(x_7, x_564); -x_566 = l_Subarray_forInUnsafe_loop___at_Lean_Elab_WF_getForbiddenByTrivialSizeOf___spec__1___closed__3; -x_567 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__2; -lean_inc(x_2); -x_568 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_568, 0, x_566); -lean_ctor_set(x_568, 1, x_2); -lean_ctor_set(x_568, 2, x_567); -x_569 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__3; -lean_inc(x_558); +x_563 = lean_ctor_get(x_561, 1); lean_inc(x_563); -x_570 = l_Lean_addMacroScope(x_563, x_569, x_558); -x_571 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__4; -lean_inc(x_5); -x_572 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_572, 0, x_571); -lean_ctor_set(x_572, 1, x_5); -lean_inc(x_556); -x_573 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_573, 0, x_556); -lean_ctor_set(x_573, 1, x_568); -lean_ctor_set(x_573, 2, x_570); -lean_ctor_set(x_573, 3, x_572); -lean_inc(x_4); -x_574 = l_Lean_addMacroScope(x_563, x_4, x_558); -lean_inc(x_5); +lean_dec(x_561); +x_564 = lean_ctor_get(x_562, 0); +lean_inc(x_564); +lean_dec(x_562); +x_565 = lean_environment_main_module(x_564); +x_566 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__1; +lean_inc(x_8); +x_567 = lean_name_mk_string(x_8, x_566); +x_568 = l_Subarray_forInUnsafe_loop___at_Lean_Elab_WF_getForbiddenByTrivialSizeOf___spec__1___closed__3; +x_569 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__2; lean_inc(x_3); +x_570 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_570, 0, x_568); +lean_ctor_set(x_570, 1, x_3); +lean_ctor_set(x_570, 2, x_569); +x_571 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__3; +lean_inc(x_560); +lean_inc(x_565); +x_572 = l_Lean_addMacroScope(x_565, x_571, x_560); +x_573 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__4; +lean_inc(x_6); +x_574 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_574, 0, x_573); +lean_ctor_set(x_574, 1, x_6); +lean_inc(x_558); x_575 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_575, 0, x_556); -lean_ctor_set(x_575, 1, x_3); -lean_ctor_set(x_575, 2, x_574); -lean_ctor_set(x_575, 3, x_5); -lean_inc(x_8); -x_576 = lean_array_push(x_8, x_575); -x_577 = lean_box(2); -x_578 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__6; -x_579 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_579, 0, x_577); -lean_ctor_set(x_579, 1, x_578); -lean_ctor_set(x_579, 2, x_576); -x_580 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__7; -x_581 = lean_array_push(x_580, x_573); -x_582 = lean_array_push(x_581, x_579); -x_583 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_583, 0, x_577); -lean_ctor_set(x_583, 1, x_565); -lean_ctor_set(x_583, 2, x_582); -x_584 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___lambda__1(x_24, x_552, x_525, x_536, x_546, x_516, x_583, x_15, x_16, x_17, x_18, x_19, x_20, x_561); -lean_dec(x_24); -x_585 = lean_ctor_get(x_584, 0); -lean_inc(x_585); -x_586 = lean_ctor_get(x_584, 1); -lean_inc(x_586); -lean_dec(x_584); -x_25 = x_585; -x_26 = x_586; -goto block_33; +lean_ctor_set(x_575, 0, x_558); +lean_ctor_set(x_575, 1, x_570); +lean_ctor_set(x_575, 2, x_572); +lean_ctor_set(x_575, 3, x_574); +lean_inc(x_5); +x_576 = l_Lean_addMacroScope(x_565, x_5, x_560); +lean_inc(x_6); +lean_inc(x_4); +x_577 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_577, 0, x_558); +lean_ctor_set(x_577, 1, x_4); +lean_ctor_set(x_577, 2, x_576); +lean_ctor_set(x_577, 3, x_6); +lean_inc(x_9); +x_578 = lean_array_push(x_9, x_577); +x_579 = lean_box(2); +x_580 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__6; +x_581 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_581, 0, x_579); +lean_ctor_set(x_581, 1, x_580); +lean_ctor_set(x_581, 2, x_578); +x_582 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__7; +x_583 = lean_array_push(x_582, x_575); +x_584 = lean_array_push(x_583, x_581); +x_585 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_585, 0, x_579); +lean_ctor_set(x_585, 1, x_567); +lean_ctor_set(x_585, 2, x_584); +x_586 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___lambda__1(x_25, x_554, x_526, x_537, x_547, x_517, x_585, x_16, x_17, x_18, x_19, x_20, x_21, x_563); +lean_dec(x_25); +x_587 = lean_ctor_get(x_586, 0); +lean_inc(x_587); +x_588 = lean_ctor_get(x_586, 1); +lean_inc(x_588); +lean_dec(x_586); +x_26 = x_587; +x_27 = x_588; +goto block_34; } else { -lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; -lean_inc(x_19); -x_587 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_19, x_20, x_553); -x_588 = lean_ctor_get(x_587, 0); -lean_inc(x_588); -x_589 = lean_ctor_get(x_587, 1); -lean_inc(x_589); -lean_dec(x_587); -x_590 = lean_ctor_get(x_19, 10); +lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; +lean_inc(x_20); +x_589 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_20, x_21, x_555); +x_590 = lean_ctor_get(x_589, 0); lean_inc(x_590); -x_591 = lean_st_ref_get(x_20, x_589); -x_592 = lean_ctor_get(x_591, 0); +x_591 = lean_ctor_get(x_589, 1); +lean_inc(x_591); +lean_dec(x_589); +x_592 = lean_ctor_get(x_20, 10); lean_inc(x_592); -x_593 = lean_ctor_get(x_591, 1); -lean_inc(x_593); -lean_dec(x_591); -x_594 = lean_ctor_get(x_592, 0); +x_593 = lean_st_ref_get(x_21, x_591); +x_594 = lean_ctor_get(x_593, 0); lean_inc(x_594); -lean_dec(x_592); -x_595 = lean_environment_main_module(x_594); -x_596 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__8; -lean_inc(x_7); -x_597 = lean_name_mk_string(x_7, x_596); -x_598 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__9; -lean_inc(x_588); -x_599 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_599, 0, x_588); -lean_ctor_set(x_599, 1, x_598); -x_600 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__1; -lean_inc(x_7); -x_601 = lean_name_mk_string(x_7, x_600); -x_602 = l_Subarray_forInUnsafe_loop___at_Lean_Elab_WF_getForbiddenByTrivialSizeOf___spec__1___closed__3; -x_603 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__2; -lean_inc(x_2); -x_604 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_604, 0, x_602); -lean_ctor_set(x_604, 1, x_2); -lean_ctor_set(x_604, 2, x_603); -x_605 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__3; -lean_inc(x_590); +x_595 = lean_ctor_get(x_593, 1); lean_inc(x_595); -x_606 = l_Lean_addMacroScope(x_595, x_605, x_590); -x_607 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__4; -lean_inc(x_5); -x_608 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_608, 0, x_607); -lean_ctor_set(x_608, 1, x_5); -lean_inc(x_588); -x_609 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_609, 0, x_588); -lean_ctor_set(x_609, 1, x_604); -lean_ctor_set(x_609, 2, x_606); -lean_ctor_set(x_609, 3, x_608); -lean_inc(x_4); -x_610 = l_Lean_addMacroScope(x_595, x_4, x_590); -lean_inc(x_5); -lean_inc(x_3); -lean_inc(x_588); -x_611 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_611, 0, x_588); -lean_ctor_set(x_611, 1, x_3); -lean_ctor_set(x_611, 2, x_610); -lean_ctor_set(x_611, 3, x_5); +lean_dec(x_593); +x_596 = lean_ctor_get(x_594, 0); +lean_inc(x_596); +lean_dec(x_594); +x_597 = lean_environment_main_module(x_596); +x_598 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__8; lean_inc(x_8); -x_612 = lean_array_push(x_8, x_611); -x_613 = lean_box(2); -x_614 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__6; -x_615 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_615, 0, x_613); -lean_ctor_set(x_615, 1, x_614); -lean_ctor_set(x_615, 2, x_612); -x_616 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__7; -x_617 = lean_array_push(x_616, x_609); -x_618 = lean_array_push(x_617, x_615); -x_619 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_619, 0, x_613); -lean_ctor_set(x_619, 1, x_601); -lean_ctor_set(x_619, 2, x_618); -x_620 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__10; -lean_inc(x_7); -x_621 = lean_name_mk_string(x_7, x_620); -x_622 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__11; -lean_inc(x_588); -x_623 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_623, 0, x_588); -lean_ctor_set(x_623, 1, x_622); -x_624 = l_Nat_repr(x_517); -x_625 = l_Lean_numLitKind; -x_626 = l_Lean_Syntax_mkLit(x_625, x_624, x_613); +x_599 = lean_name_mk_string(x_8, x_598); +x_600 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__9; +lean_inc(x_590); +x_601 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_601, 0, x_590); +lean_ctor_set(x_601, 1, x_600); +x_602 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__1; lean_inc(x_8); -x_627 = lean_array_push(x_8, x_626); -x_628 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_628, 0, x_613); -lean_ctor_set(x_628, 1, x_614); -lean_ctor_set(x_628, 2, x_627); -x_629 = lean_array_push(x_616, x_623); -x_630 = lean_array_push(x_629, x_628); -x_631 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_631, 0, x_613); -lean_ctor_set(x_631, 1, x_621); -lean_ctor_set(x_631, 2, x_630); +x_603 = lean_name_mk_string(x_8, x_602); +x_604 = l_Subarray_forInUnsafe_loop___at_Lean_Elab_WF_getForbiddenByTrivialSizeOf___spec__1___closed__3; +x_605 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__2; +lean_inc(x_3); +x_606 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_606, 0, x_604); +lean_ctor_set(x_606, 1, x_3); +lean_ctor_set(x_606, 2, x_605); +x_607 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__3; +lean_inc(x_592); +lean_inc(x_597); +x_608 = l_Lean_addMacroScope(x_597, x_607, x_592); +x_609 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__4; +lean_inc(x_6); +x_610 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_610, 0, x_609); +lean_ctor_set(x_610, 1, x_6); +lean_inc(x_590); +x_611 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_611, 0, x_590); +lean_ctor_set(x_611, 1, x_606); +lean_ctor_set(x_611, 2, x_608); +lean_ctor_set(x_611, 3, x_610); +lean_inc(x_5); +x_612 = l_Lean_addMacroScope(x_597, x_5, x_592); +lean_inc(x_6); +lean_inc(x_4); +lean_inc(x_590); +x_613 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_613, 0, x_590); +lean_ctor_set(x_613, 1, x_4); +lean_ctor_set(x_613, 2, x_612); +lean_ctor_set(x_613, 3, x_6); +lean_inc(x_9); +x_614 = lean_array_push(x_9, x_613); +x_615 = lean_box(2); +x_616 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__6; +x_617 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_617, 0, x_615); +lean_ctor_set(x_617, 1, x_616); +lean_ctor_set(x_617, 2, x_614); +x_618 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__7; +x_619 = lean_array_push(x_618, x_611); +x_620 = lean_array_push(x_619, x_617); +x_621 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_621, 0, x_615); +lean_ctor_set(x_621, 1, x_603); +lean_ctor_set(x_621, 2, x_620); +x_622 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__10; lean_inc(x_8); -x_632 = lean_array_push(x_8, x_631); -x_633 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_633, 0, x_613); -lean_ctor_set(x_633, 1, x_614); -lean_ctor_set(x_633, 2, x_632); -x_634 = lean_array_push(x_616, x_619); -x_635 = lean_array_push(x_634, x_633); -x_636 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_636, 0, x_613); -lean_ctor_set(x_636, 1, x_614); -lean_ctor_set(x_636, 2, x_635); -x_637 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__12; -x_638 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_638, 0, x_588); -lean_ctor_set(x_638, 1, x_637); -x_639 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__13; -x_640 = lean_array_push(x_639, x_599); -x_641 = lean_array_push(x_640, x_636); -x_642 = lean_array_push(x_641, x_638); -x_643 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_643, 0, x_613); -lean_ctor_set(x_643, 1, x_597); -lean_ctor_set(x_643, 2, x_642); -x_644 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___lambda__1(x_24, x_552, x_525, x_536, x_546, x_516, x_643, x_15, x_16, x_17, x_18, x_19, x_20, x_593); -lean_dec(x_24); -x_645 = lean_ctor_get(x_644, 0); -lean_inc(x_645); -x_646 = lean_ctor_get(x_644, 1); +x_623 = lean_name_mk_string(x_8, x_622); +x_624 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__11; +lean_inc(x_590); +x_625 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_625, 0, x_590); +lean_ctor_set(x_625, 1, x_624); +x_626 = l_Nat_repr(x_518); +x_627 = l_Lean_Syntax_mkNumLit(x_626, x_615); +lean_inc(x_9); +x_628 = lean_array_push(x_9, x_627); +x_629 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_629, 0, x_615); +lean_ctor_set(x_629, 1, x_616); +lean_ctor_set(x_629, 2, x_628); +x_630 = lean_array_push(x_618, x_625); +x_631 = lean_array_push(x_630, x_629); +x_632 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_632, 0, x_615); +lean_ctor_set(x_632, 1, x_623); +lean_ctor_set(x_632, 2, x_631); +lean_inc(x_9); +x_633 = lean_array_push(x_9, x_632); +x_634 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_634, 0, x_615); +lean_ctor_set(x_634, 1, x_616); +lean_ctor_set(x_634, 2, x_633); +x_635 = lean_array_push(x_618, x_621); +x_636 = lean_array_push(x_635, x_634); +x_637 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_637, 0, x_615); +lean_ctor_set(x_637, 1, x_616); +lean_ctor_set(x_637, 2, x_636); +x_638 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__12; +x_639 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_639, 0, x_590); +lean_ctor_set(x_639, 1, x_638); +x_640 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__13; +x_641 = lean_array_push(x_640, x_601); +x_642 = lean_array_push(x_641, x_637); +x_643 = lean_array_push(x_642, x_639); +x_644 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_644, 0, x_615); +lean_ctor_set(x_644, 1, x_599); +lean_ctor_set(x_644, 2, x_643); +x_645 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___lambda__1(x_25, x_554, x_526, x_537, x_547, x_517, x_644, x_16, x_17, x_18, x_19, x_20, x_21, x_595); +lean_dec(x_25); +x_646 = lean_ctor_get(x_645, 0); lean_inc(x_646); -lean_dec(x_644); -x_25 = x_645; +x_647 = lean_ctor_get(x_645, 1); +lean_inc(x_647); +lean_dec(x_645); x_26 = x_646; -goto block_33; +x_27 = x_647; +goto block_34; } } } @@ -8144,425 +8141,424 @@ goto block_33; } else { -lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; uint8_t x_653; -x_647 = lean_ctor_get(x_34, 0); -lean_inc(x_647); -lean_dec(x_34); -x_648 = lean_ctor_get(x_35, 1); +lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; uint8_t x_654; +x_648 = lean_ctor_get(x_35, 0); lean_inc(x_648); -if (lean_is_exclusive(x_35)) { - lean_ctor_release(x_35, 0); - lean_ctor_release(x_35, 1); - x_649 = x_35; +lean_dec(x_35); +x_649 = lean_ctor_get(x_36, 1); +lean_inc(x_649); +if (lean_is_exclusive(x_36)) { + lean_ctor_release(x_36, 0); + lean_ctor_release(x_36, 1); + x_650 = x_36; } else { - lean_dec_ref(x_35); - x_649 = lean_box(0); + lean_dec_ref(x_36); + x_650 = lean_box(0); } -x_650 = lean_ctor_get(x_36, 0); -lean_inc(x_650); -x_651 = lean_ctor_get(x_36, 1); +x_651 = lean_ctor_get(x_37, 0); lean_inc(x_651); -x_652 = lean_ctor_get(x_36, 2); +x_652 = lean_ctor_get(x_37, 1); lean_inc(x_652); -x_653 = lean_nat_dec_lt(x_650, x_651); -if (x_653 == 0) +x_653 = lean_ctor_get(x_37, 2); +lean_inc(x_653); +x_654 = lean_nat_dec_lt(x_651, x_652); +if (x_654 == 0) { -lean_object* x_654; lean_object* x_655; lean_object* x_656; +lean_object* x_655; lean_object* x_656; lean_object* x_657; +lean_dec(x_653); lean_dec(x_652); lean_dec(x_651); -lean_dec(x_650); -lean_dec(x_24); -if (lean_is_scalar(x_649)) { - x_654 = lean_alloc_ctor(0, 2, 0); +lean_dec(x_25); +if (lean_is_scalar(x_650)) { + x_655 = lean_alloc_ctor(0, 2, 0); } else { - x_654 = x_649; -} -lean_ctor_set(x_654, 0, x_36); -lean_ctor_set(x_654, 1, x_648); -x_655 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_655, 0, x_647); -lean_ctor_set(x_655, 1, x_654); -lean_ctor_set(x_14, 1, x_655); -x_656 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_656, 0, x_14); -x_25 = x_656; -x_26 = x_21; -goto block_33; -} -else -{ -lean_object* x_657; lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; lean_object* x_662; uint8_t x_663; -if (lean_is_exclusive(x_36)) { - lean_ctor_release(x_36, 0); - lean_ctor_release(x_36, 1); - lean_ctor_release(x_36, 2); - x_657 = x_36; + x_655 = x_650; +} +lean_ctor_set(x_655, 0, x_37); +lean_ctor_set(x_655, 1, x_649); +x_656 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_656, 0, x_648); +lean_ctor_set(x_656, 1, x_655); +lean_ctor_set(x_15, 1, x_656); +x_657 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_657, 0, x_15); +x_26 = x_657; +x_27 = x_22; +goto block_34; +} +else +{ +lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; uint8_t x_664; +if (lean_is_exclusive(x_37)) { + lean_ctor_release(x_37, 0); + lean_ctor_release(x_37, 1); + lean_ctor_release(x_37, 2); + x_658 = x_37; } else { - lean_dec_ref(x_36); - x_657 = lean_box(0); + lean_dec_ref(x_37); + x_658 = lean_box(0); } -x_658 = lean_nat_add(x_650, x_652); -if (lean_is_scalar(x_657)) { - x_659 = lean_alloc_ctor(0, 3, 0); +x_659 = lean_nat_add(x_651, x_653); +if (lean_is_scalar(x_658)) { + x_660 = lean_alloc_ctor(0, 3, 0); } else { - x_659 = x_657; -} -lean_ctor_set(x_659, 0, x_658); -lean_ctor_set(x_659, 1, x_651); -lean_ctor_set(x_659, 2, x_652); -x_660 = lean_ctor_get(x_647, 0); -lean_inc(x_660); -x_661 = lean_ctor_get(x_647, 1); + x_660 = x_658; +} +lean_ctor_set(x_660, 0, x_659); +lean_ctor_set(x_660, 1, x_652); +lean_ctor_set(x_660, 2, x_653); +x_661 = lean_ctor_get(x_648, 0); lean_inc(x_661); -x_662 = lean_ctor_get(x_647, 2); +x_662 = lean_ctor_get(x_648, 1); lean_inc(x_662); -x_663 = lean_nat_dec_lt(x_661, x_662); -if (x_663 == 0) +x_663 = lean_ctor_get(x_648, 2); +lean_inc(x_663); +x_664 = lean_nat_dec_lt(x_662, x_663); +if (x_664 == 0) { -lean_object* x_664; lean_object* x_665; lean_object* x_666; +lean_object* x_665; lean_object* x_666; lean_object* x_667; +lean_dec(x_663); lean_dec(x_662); lean_dec(x_661); -lean_dec(x_660); -lean_dec(x_650); -lean_dec(x_24); -if (lean_is_scalar(x_649)) { - x_664 = lean_alloc_ctor(0, 2, 0); +lean_dec(x_651); +lean_dec(x_25); +if (lean_is_scalar(x_650)) { + x_665 = lean_alloc_ctor(0, 2, 0); } else { - x_664 = x_649; -} -lean_ctor_set(x_664, 0, x_659); -lean_ctor_set(x_664, 1, x_648); -x_665 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_665, 0, x_647); -lean_ctor_set(x_665, 1, x_664); -lean_ctor_set(x_14, 1, x_665); -x_666 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_666, 0, x_14); -x_25 = x_666; -x_26 = x_21; -goto block_33; -} -else -{ -lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; lean_object* x_674; uint8_t x_675; -if (lean_is_exclusive(x_647)) { - lean_ctor_release(x_647, 0); - lean_ctor_release(x_647, 1); - lean_ctor_release(x_647, 2); - x_667 = x_647; + x_665 = x_650; +} +lean_ctor_set(x_665, 0, x_660); +lean_ctor_set(x_665, 1, x_649); +x_666 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_666, 0, x_648); +lean_ctor_set(x_666, 1, x_665); +lean_ctor_set(x_15, 1, x_666); +x_667 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_667, 0, x_15); +x_26 = x_667; +x_27 = x_22; +goto block_34; +} +else +{ +lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; lean_object* x_674; lean_object* x_675; uint8_t x_676; +if (lean_is_exclusive(x_648)) { + lean_ctor_release(x_648, 0); + lean_ctor_release(x_648, 1); + lean_ctor_release(x_648, 2); + x_668 = x_648; } else { - lean_dec_ref(x_647); - x_667 = lean_box(0); + lean_dec_ref(x_648); + x_668 = lean_box(0); } -x_668 = lean_array_fget(x_660, x_661); -x_669 = lean_unsigned_to_nat(1u); -x_670 = lean_nat_add(x_661, x_669); -lean_dec(x_661); -if (lean_is_scalar(x_667)) { - x_671 = lean_alloc_ctor(0, 3, 0); +x_669 = lean_array_fget(x_661, x_662); +x_670 = lean_unsigned_to_nat(1u); +x_671 = lean_nat_add(x_662, x_670); +lean_dec(x_662); +if (lean_is_scalar(x_668)) { + x_672 = lean_alloc_ctor(0, 3, 0); } else { - x_671 = x_667; -} -lean_ctor_set(x_671, 0, x_660); -lean_ctor_set(x_671, 1, x_670); -lean_ctor_set(x_671, 2, x_662); -x_672 = lean_ctor_get(x_38, 0); -lean_inc(x_672); -x_673 = lean_ctor_get(x_38, 1); + x_672 = x_668; +} +lean_ctor_set(x_672, 0, x_661); +lean_ctor_set(x_672, 1, x_671); +lean_ctor_set(x_672, 2, x_663); +x_673 = lean_ctor_get(x_39, 0); lean_inc(x_673); -x_674 = lean_ctor_get(x_38, 2); +x_674 = lean_ctor_get(x_39, 1); lean_inc(x_674); -x_675 = lean_nat_dec_lt(x_673, x_674); -if (x_675 == 0) +x_675 = lean_ctor_get(x_39, 2); +lean_inc(x_675); +x_676 = lean_nat_dec_lt(x_674, x_675); +if (x_676 == 0) { -lean_object* x_676; lean_object* x_677; lean_object* x_678; +lean_object* x_677; lean_object* x_678; lean_object* x_679; +lean_dec(x_675); lean_dec(x_674); lean_dec(x_673); -lean_dec(x_672); -lean_dec(x_668); -lean_dec(x_650); -lean_dec(x_24); -if (lean_is_scalar(x_649)) { - x_676 = lean_alloc_ctor(0, 2, 0); +lean_dec(x_669); +lean_dec(x_651); +lean_dec(x_25); +if (lean_is_scalar(x_650)) { + x_677 = lean_alloc_ctor(0, 2, 0); } else { - x_676 = x_649; + x_677 = x_650; } -lean_ctor_set(x_676, 0, x_659); -lean_ctor_set(x_676, 1, x_648); -x_677 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_677, 0, x_671); -lean_ctor_set(x_677, 1, x_676); -lean_ctor_set(x_14, 1, x_677); -x_678 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_678, 0, x_14); -x_25 = x_678; -x_26 = x_21; -goto block_33; +lean_ctor_set(x_677, 0, x_660); +lean_ctor_set(x_677, 1, x_649); +x_678 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_678, 0, x_672); +lean_ctor_set(x_678, 1, x_677); +lean_ctor_set(x_15, 1, x_678); +x_679 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_679, 0, x_15); +x_26 = x_679; +x_27 = x_22; +goto block_34; } else { -lean_object* x_679; lean_object* x_680; lean_object* x_681; lean_object* x_682; lean_object* x_683; lean_object* x_684; lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; uint8_t x_690; -lean_dec(x_649); -lean_free_object(x_14); -if (lean_is_exclusive(x_38)) { - lean_ctor_release(x_38, 0); - lean_ctor_release(x_38, 1); - lean_ctor_release(x_38, 2); - x_679 = x_38; +lean_object* x_680; lean_object* x_681; lean_object* x_682; lean_object* x_683; lean_object* x_684; lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; lean_object* x_690; lean_object* x_691; uint8_t x_692; +lean_dec(x_650); +lean_free_object(x_15); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + lean_ctor_release(x_39, 2); + x_680 = x_39; } else { - lean_dec_ref(x_38); - x_679 = lean_box(0); + lean_dec_ref(x_39); + x_680 = lean_box(0); } -x_680 = lean_array_fget(x_672, x_673); -x_681 = lean_nat_add(x_673, x_669); -lean_dec(x_673); -if (lean_is_scalar(x_679)) { - x_682 = lean_alloc_ctor(0, 3, 0); +x_681 = lean_array_fget(x_673, x_674); +x_682 = lean_nat_add(x_674, x_670); +lean_dec(x_674); +if (lean_is_scalar(x_680)) { + x_683 = lean_alloc_ctor(0, 3, 0); } else { - x_682 = x_679; + x_683 = x_680; } -lean_ctor_set(x_682, 0, x_672); -lean_ctor_set(x_682, 1, x_681); -lean_ctor_set(x_682, 2, x_674); -lean_inc(x_6); -lean_inc(x_8); -x_683 = lean_array_push(x_8, x_6); -x_684 = lean_nat_sub(x_680, x_668); -lean_dec(x_668); -lean_dec(x_680); -x_685 = lean_nat_sub(x_684, x_669); -lean_dec(x_684); -x_686 = lean_unsigned_to_nat(0u); -lean_inc(x_685); -lean_inc(x_9); -x_687 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__1(x_9, x_685, x_686, x_685, x_669, x_683, x_15, x_16, x_17, x_18, x_19, x_20, x_21); -lean_dec(x_685); -x_688 = lean_ctor_get(x_687, 0); -lean_inc(x_688); -x_689 = lean_ctor_get(x_687, 1); -lean_inc(x_689); +lean_ctor_set(x_683, 0, x_673); +lean_ctor_set(x_683, 1, x_682); +lean_ctor_set(x_683, 2, x_675); +x_684 = l___private_Lean_Elab_PreDefinition_WF_Rel_0__Lean_Elab_WF_unpackUnary_go___lambda__1___closed__1; +lean_inc(x_7); +x_685 = lean_array_push(x_684, x_7); +x_686 = lean_nat_sub(x_681, x_669); +lean_dec(x_669); +lean_dec(x_681); +x_687 = lean_nat_sub(x_686, x_670); +lean_dec(x_686); +x_688 = lean_unsigned_to_nat(0u); +lean_inc(x_687); +lean_inc(x_10); +x_689 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__1(x_2, x_10, x_687, x_688, x_687, x_670, x_685, x_16, x_17, x_18, x_19, x_20, x_21, x_22); lean_dec(x_687); -x_690 = lean_nat_dec_lt(x_669, x_10); -if (x_690 == 0) -{ -lean_object* x_691; lean_object* x_692; lean_object* x_693; lean_object* x_694; lean_object* x_695; lean_object* x_696; lean_object* x_697; lean_object* x_698; lean_object* x_699; lean_object* x_700; lean_object* x_701; lean_object* x_702; lean_object* x_703; lean_object* x_704; lean_object* x_705; lean_object* x_706; lean_object* x_707; lean_object* x_708; lean_object* x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; lean_object* x_718; lean_object* x_719; lean_object* x_720; lean_object* x_721; lean_object* x_722; -lean_dec(x_650); -lean_inc(x_19); -x_691 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_19, x_20, x_689); -x_692 = lean_ctor_get(x_691, 0); -lean_inc(x_692); -x_693 = lean_ctor_get(x_691, 1); -lean_inc(x_693); -lean_dec(x_691); -x_694 = lean_ctor_get(x_19, 10); +x_690 = lean_ctor_get(x_689, 0); +lean_inc(x_690); +x_691 = lean_ctor_get(x_689, 1); +lean_inc(x_691); +lean_dec(x_689); +x_692 = lean_nat_dec_lt(x_670, x_11); +if (x_692 == 0) +{ +lean_object* x_693; lean_object* x_694; lean_object* x_695; lean_object* x_696; lean_object* x_697; lean_object* x_698; lean_object* x_699; lean_object* x_700; lean_object* x_701; lean_object* x_702; lean_object* x_703; lean_object* x_704; lean_object* x_705; lean_object* x_706; lean_object* x_707; lean_object* x_708; lean_object* x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; lean_object* x_718; lean_object* x_719; lean_object* x_720; lean_object* x_721; lean_object* x_722; lean_object* x_723; lean_object* x_724; +lean_dec(x_651); +lean_inc(x_20); +x_693 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_20, x_21, x_691); +x_694 = lean_ctor_get(x_693, 0); lean_inc(x_694); -x_695 = lean_st_ref_get(x_20, x_693); -x_696 = lean_ctor_get(x_695, 0); +x_695 = lean_ctor_get(x_693, 1); +lean_inc(x_695); +lean_dec(x_693); +x_696 = lean_ctor_get(x_20, 10); lean_inc(x_696); -x_697 = lean_ctor_get(x_695, 1); -lean_inc(x_697); -lean_dec(x_695); -x_698 = lean_ctor_get(x_696, 0); +x_697 = lean_st_ref_get(x_21, x_695); +x_698 = lean_ctor_get(x_697, 0); lean_inc(x_698); -lean_dec(x_696); -x_699 = lean_environment_main_module(x_698); -x_700 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__1; -lean_inc(x_7); -x_701 = lean_name_mk_string(x_7, x_700); -x_702 = l_Subarray_forInUnsafe_loop___at_Lean_Elab_WF_getForbiddenByTrivialSizeOf___spec__1___closed__3; -x_703 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__2; -lean_inc(x_2); -x_704 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_704, 0, x_702); -lean_ctor_set(x_704, 1, x_2); -lean_ctor_set(x_704, 2, x_703); -x_705 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__3; -lean_inc(x_694); +x_699 = lean_ctor_get(x_697, 1); lean_inc(x_699); -x_706 = l_Lean_addMacroScope(x_699, x_705, x_694); -x_707 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__4; -lean_inc(x_5); -x_708 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_708, 0, x_707); -lean_ctor_set(x_708, 1, x_5); -lean_inc(x_692); -x_709 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_709, 0, x_692); -lean_ctor_set(x_709, 1, x_704); -lean_ctor_set(x_709, 2, x_706); -lean_ctor_set(x_709, 3, x_708); -lean_inc(x_4); -x_710 = l_Lean_addMacroScope(x_699, x_4, x_694); -lean_inc(x_5); +lean_dec(x_697); +x_700 = lean_ctor_get(x_698, 0); +lean_inc(x_700); +lean_dec(x_698); +x_701 = lean_environment_main_module(x_700); +x_702 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__1; +lean_inc(x_8); +x_703 = lean_name_mk_string(x_8, x_702); +x_704 = l_Subarray_forInUnsafe_loop___at_Lean_Elab_WF_getForbiddenByTrivialSizeOf___spec__1___closed__3; +x_705 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__2; lean_inc(x_3); +x_706 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_706, 0, x_704); +lean_ctor_set(x_706, 1, x_3); +lean_ctor_set(x_706, 2, x_705); +x_707 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__3; +lean_inc(x_696); +lean_inc(x_701); +x_708 = l_Lean_addMacroScope(x_701, x_707, x_696); +x_709 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__4; +lean_inc(x_6); +x_710 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_710, 0, x_709); +lean_ctor_set(x_710, 1, x_6); +lean_inc(x_694); x_711 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_711, 0, x_692); -lean_ctor_set(x_711, 1, x_3); -lean_ctor_set(x_711, 2, x_710); -lean_ctor_set(x_711, 3, x_5); -lean_inc(x_8); -x_712 = lean_array_push(x_8, x_711); -x_713 = lean_box(2); -x_714 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__6; -x_715 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_715, 0, x_713); -lean_ctor_set(x_715, 1, x_714); -lean_ctor_set(x_715, 2, x_712); -x_716 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__7; -x_717 = lean_array_push(x_716, x_709); -x_718 = lean_array_push(x_717, x_715); -x_719 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_719, 0, x_713); -lean_ctor_set(x_719, 1, x_701); -lean_ctor_set(x_719, 2, x_718); -x_720 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___lambda__1(x_24, x_688, x_659, x_671, x_682, x_648, x_719, x_15, x_16, x_17, x_18, x_19, x_20, x_697); -lean_dec(x_24); -x_721 = lean_ctor_get(x_720, 0); -lean_inc(x_721); -x_722 = lean_ctor_get(x_720, 1); -lean_inc(x_722); -lean_dec(x_720); -x_25 = x_721; -x_26 = x_722; -goto block_33; +lean_ctor_set(x_711, 0, x_694); +lean_ctor_set(x_711, 1, x_706); +lean_ctor_set(x_711, 2, x_708); +lean_ctor_set(x_711, 3, x_710); +lean_inc(x_5); +x_712 = l_Lean_addMacroScope(x_701, x_5, x_696); +lean_inc(x_6); +lean_inc(x_4); +x_713 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_713, 0, x_694); +lean_ctor_set(x_713, 1, x_4); +lean_ctor_set(x_713, 2, x_712); +lean_ctor_set(x_713, 3, x_6); +lean_inc(x_9); +x_714 = lean_array_push(x_9, x_713); +x_715 = lean_box(2); +x_716 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__6; +x_717 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_717, 0, x_715); +lean_ctor_set(x_717, 1, x_716); +lean_ctor_set(x_717, 2, x_714); +x_718 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__7; +x_719 = lean_array_push(x_718, x_711); +x_720 = lean_array_push(x_719, x_717); +x_721 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_721, 0, x_715); +lean_ctor_set(x_721, 1, x_703); +lean_ctor_set(x_721, 2, x_720); +x_722 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___lambda__1(x_25, x_690, x_660, x_672, x_683, x_649, x_721, x_16, x_17, x_18, x_19, x_20, x_21, x_699); +lean_dec(x_25); +x_723 = lean_ctor_get(x_722, 0); +lean_inc(x_723); +x_724 = lean_ctor_get(x_722, 1); +lean_inc(x_724); +lean_dec(x_722); +x_26 = x_723; +x_27 = x_724; +goto block_34; } else { -lean_object* x_723; lean_object* x_724; lean_object* x_725; lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; lean_object* x_730; lean_object* x_731; lean_object* x_732; lean_object* x_733; lean_object* x_734; lean_object* x_735; lean_object* x_736; lean_object* x_737; lean_object* x_738; lean_object* x_739; lean_object* x_740; lean_object* x_741; lean_object* x_742; lean_object* x_743; lean_object* x_744; lean_object* x_745; lean_object* x_746; lean_object* x_747; lean_object* x_748; lean_object* x_749; lean_object* x_750; lean_object* x_751; lean_object* x_752; lean_object* x_753; lean_object* x_754; lean_object* x_755; lean_object* x_756; lean_object* x_757; lean_object* x_758; lean_object* x_759; lean_object* x_760; lean_object* x_761; lean_object* x_762; lean_object* x_763; lean_object* x_764; lean_object* x_765; lean_object* x_766; lean_object* x_767; lean_object* x_768; lean_object* x_769; lean_object* x_770; lean_object* x_771; lean_object* x_772; lean_object* x_773; lean_object* x_774; lean_object* x_775; lean_object* x_776; lean_object* x_777; lean_object* x_778; lean_object* x_779; lean_object* x_780; lean_object* x_781; lean_object* x_782; -lean_inc(x_19); -x_723 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_19, x_20, x_689); -x_724 = lean_ctor_get(x_723, 0); -lean_inc(x_724); -x_725 = lean_ctor_get(x_723, 1); -lean_inc(x_725); -lean_dec(x_723); -x_726 = lean_ctor_get(x_19, 10); +lean_object* x_725; lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; lean_object* x_730; lean_object* x_731; lean_object* x_732; lean_object* x_733; lean_object* x_734; lean_object* x_735; lean_object* x_736; lean_object* x_737; lean_object* x_738; lean_object* x_739; lean_object* x_740; lean_object* x_741; lean_object* x_742; lean_object* x_743; lean_object* x_744; lean_object* x_745; lean_object* x_746; lean_object* x_747; lean_object* x_748; lean_object* x_749; lean_object* x_750; lean_object* x_751; lean_object* x_752; lean_object* x_753; lean_object* x_754; lean_object* x_755; lean_object* x_756; lean_object* x_757; lean_object* x_758; lean_object* x_759; lean_object* x_760; lean_object* x_761; lean_object* x_762; lean_object* x_763; lean_object* x_764; lean_object* x_765; lean_object* x_766; lean_object* x_767; lean_object* x_768; lean_object* x_769; lean_object* x_770; lean_object* x_771; lean_object* x_772; lean_object* x_773; lean_object* x_774; lean_object* x_775; lean_object* x_776; lean_object* x_777; lean_object* x_778; lean_object* x_779; lean_object* x_780; lean_object* x_781; lean_object* x_782; lean_object* x_783; +lean_inc(x_20); +x_725 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_20, x_21, x_691); +x_726 = lean_ctor_get(x_725, 0); lean_inc(x_726); -x_727 = lean_st_ref_get(x_20, x_725); -x_728 = lean_ctor_get(x_727, 0); +x_727 = lean_ctor_get(x_725, 1); +lean_inc(x_727); +lean_dec(x_725); +x_728 = lean_ctor_get(x_20, 10); lean_inc(x_728); -x_729 = lean_ctor_get(x_727, 1); -lean_inc(x_729); -lean_dec(x_727); -x_730 = lean_ctor_get(x_728, 0); +x_729 = lean_st_ref_get(x_21, x_727); +x_730 = lean_ctor_get(x_729, 0); lean_inc(x_730); -lean_dec(x_728); -x_731 = lean_environment_main_module(x_730); -x_732 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__8; -lean_inc(x_7); -x_733 = lean_name_mk_string(x_7, x_732); -x_734 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__9; -lean_inc(x_724); -x_735 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_735, 0, x_724); -lean_ctor_set(x_735, 1, x_734); -x_736 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__1; -lean_inc(x_7); -x_737 = lean_name_mk_string(x_7, x_736); -x_738 = l_Subarray_forInUnsafe_loop___at_Lean_Elab_WF_getForbiddenByTrivialSizeOf___spec__1___closed__3; -x_739 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__2; -lean_inc(x_2); -x_740 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_740, 0, x_738); -lean_ctor_set(x_740, 1, x_2); -lean_ctor_set(x_740, 2, x_739); -x_741 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__3; +x_731 = lean_ctor_get(x_729, 1); +lean_inc(x_731); +lean_dec(x_729); +x_732 = lean_ctor_get(x_730, 0); +lean_inc(x_732); +lean_dec(x_730); +x_733 = lean_environment_main_module(x_732); +x_734 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__8; +lean_inc(x_8); +x_735 = lean_name_mk_string(x_8, x_734); +x_736 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__9; +lean_inc(x_726); +x_737 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_737, 0, x_726); +lean_ctor_set(x_737, 1, x_736); +x_738 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__1; +lean_inc(x_8); +x_739 = lean_name_mk_string(x_8, x_738); +x_740 = l_Subarray_forInUnsafe_loop___at_Lean_Elab_WF_getForbiddenByTrivialSizeOf___spec__1___closed__3; +x_741 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__2; +lean_inc(x_3); +x_742 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_742, 0, x_740); +lean_ctor_set(x_742, 1, x_3); +lean_ctor_set(x_742, 2, x_741); +x_743 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__3; +lean_inc(x_728); +lean_inc(x_733); +x_744 = l_Lean_addMacroScope(x_733, x_743, x_728); +x_745 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__4; +lean_inc(x_6); +x_746 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_746, 0, x_745); +lean_ctor_set(x_746, 1, x_6); lean_inc(x_726); -lean_inc(x_731); -x_742 = l_Lean_addMacroScope(x_731, x_741, x_726); -x_743 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__4; +x_747 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_747, 0, x_726); +lean_ctor_set(x_747, 1, x_742); +lean_ctor_set(x_747, 2, x_744); +lean_ctor_set(x_747, 3, x_746); lean_inc(x_5); -x_744 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_744, 0, x_743); -lean_ctor_set(x_744, 1, x_5); -lean_inc(x_724); -x_745 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_745, 0, x_724); -lean_ctor_set(x_745, 1, x_740); -lean_ctor_set(x_745, 2, x_742); -lean_ctor_set(x_745, 3, x_744); +x_748 = l_Lean_addMacroScope(x_733, x_5, x_728); +lean_inc(x_6); lean_inc(x_4); -x_746 = l_Lean_addMacroScope(x_731, x_4, x_726); -lean_inc(x_5); -lean_inc(x_3); -lean_inc(x_724); -x_747 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_747, 0, x_724); -lean_ctor_set(x_747, 1, x_3); -lean_ctor_set(x_747, 2, x_746); -lean_ctor_set(x_747, 3, x_5); -lean_inc(x_8); -x_748 = lean_array_push(x_8, x_747); -x_749 = lean_box(2); -x_750 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__6; -x_751 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_751, 0, x_749); -lean_ctor_set(x_751, 1, x_750); -lean_ctor_set(x_751, 2, x_748); -x_752 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__7; -x_753 = lean_array_push(x_752, x_745); -x_754 = lean_array_push(x_753, x_751); -x_755 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_755, 0, x_749); -lean_ctor_set(x_755, 1, x_737); -lean_ctor_set(x_755, 2, x_754); -x_756 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__10; -lean_inc(x_7); -x_757 = lean_name_mk_string(x_7, x_756); -x_758 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__11; -lean_inc(x_724); -x_759 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_759, 0, x_724); -lean_ctor_set(x_759, 1, x_758); -x_760 = l_Nat_repr(x_650); -x_761 = l_Lean_numLitKind; -x_762 = l_Lean_Syntax_mkLit(x_761, x_760, x_749); -lean_inc(x_8); -x_763 = lean_array_push(x_8, x_762); -x_764 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_764, 0, x_749); -lean_ctor_set(x_764, 1, x_750); -lean_ctor_set(x_764, 2, x_763); -x_765 = lean_array_push(x_752, x_759); -x_766 = lean_array_push(x_765, x_764); -x_767 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_767, 0, x_749); -lean_ctor_set(x_767, 1, x_757); -lean_ctor_set(x_767, 2, x_766); +lean_inc(x_726); +x_749 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_749, 0, x_726); +lean_ctor_set(x_749, 1, x_4); +lean_ctor_set(x_749, 2, x_748); +lean_ctor_set(x_749, 3, x_6); +lean_inc(x_9); +x_750 = lean_array_push(x_9, x_749); +x_751 = lean_box(2); +x_752 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__6; +x_753 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_753, 0, x_751); +lean_ctor_set(x_753, 1, x_752); +lean_ctor_set(x_753, 2, x_750); +x_754 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__7; +x_755 = lean_array_push(x_754, x_747); +x_756 = lean_array_push(x_755, x_753); +x_757 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_757, 0, x_751); +lean_ctor_set(x_757, 1, x_739); +lean_ctor_set(x_757, 2, x_756); +x_758 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__10; lean_inc(x_8); -x_768 = lean_array_push(x_8, x_767); -x_769 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_769, 0, x_749); -lean_ctor_set(x_769, 1, x_750); -lean_ctor_set(x_769, 2, x_768); -x_770 = lean_array_push(x_752, x_755); -x_771 = lean_array_push(x_770, x_769); -x_772 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_772, 0, x_749); -lean_ctor_set(x_772, 1, x_750); -lean_ctor_set(x_772, 2, x_771); -x_773 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__12; -x_774 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_774, 0, x_724); -lean_ctor_set(x_774, 1, x_773); -x_775 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__13; -x_776 = lean_array_push(x_775, x_735); -x_777 = lean_array_push(x_776, x_772); -x_778 = lean_array_push(x_777, x_774); -x_779 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_779, 0, x_749); -lean_ctor_set(x_779, 1, x_733); -lean_ctor_set(x_779, 2, x_778); -x_780 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___lambda__1(x_24, x_688, x_659, x_671, x_682, x_648, x_779, x_15, x_16, x_17, x_18, x_19, x_20, x_729); -lean_dec(x_24); -x_781 = lean_ctor_get(x_780, 0); -lean_inc(x_781); -x_782 = lean_ctor_get(x_780, 1); +x_759 = lean_name_mk_string(x_8, x_758); +x_760 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__11; +lean_inc(x_726); +x_761 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_761, 0, x_726); +lean_ctor_set(x_761, 1, x_760); +x_762 = l_Nat_repr(x_651); +x_763 = l_Lean_Syntax_mkNumLit(x_762, x_751); +lean_inc(x_9); +x_764 = lean_array_push(x_9, x_763); +x_765 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_765, 0, x_751); +lean_ctor_set(x_765, 1, x_752); +lean_ctor_set(x_765, 2, x_764); +x_766 = lean_array_push(x_754, x_761); +x_767 = lean_array_push(x_766, x_765); +x_768 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_768, 0, x_751); +lean_ctor_set(x_768, 1, x_759); +lean_ctor_set(x_768, 2, x_767); +lean_inc(x_9); +x_769 = lean_array_push(x_9, x_768); +x_770 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_770, 0, x_751); +lean_ctor_set(x_770, 1, x_752); +lean_ctor_set(x_770, 2, x_769); +x_771 = lean_array_push(x_754, x_757); +x_772 = lean_array_push(x_771, x_770); +x_773 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_773, 0, x_751); +lean_ctor_set(x_773, 1, x_752); +lean_ctor_set(x_773, 2, x_772); +x_774 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__12; +x_775 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_775, 0, x_726); +lean_ctor_set(x_775, 1, x_774); +x_776 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__13; +x_777 = lean_array_push(x_776, x_737); +x_778 = lean_array_push(x_777, x_773); +x_779 = lean_array_push(x_778, x_775); +x_780 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_780, 0, x_751); +lean_ctor_set(x_780, 1, x_735); +lean_ctor_set(x_780, 2, x_779); +x_781 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___lambda__1(x_25, x_690, x_660, x_672, x_683, x_649, x_780, x_16, x_17, x_18, x_19, x_20, x_21, x_731); +lean_dec(x_25); +x_782 = lean_ctor_get(x_781, 0); lean_inc(x_782); -lean_dec(x_780); -x_25 = x_781; +x_783 = lean_ctor_get(x_781, 1); +lean_inc(x_783); +lean_dec(x_781); x_26 = x_782; -goto block_33; +x_27 = x_783; +goto block_34; } } } @@ -8571,464 +8567,464 @@ goto block_33; } else { -lean_object* x_783; lean_object* x_784; lean_object* x_785; lean_object* x_786; lean_object* x_787; lean_object* x_788; lean_object* x_789; lean_object* x_790; uint8_t x_791; -x_783 = lean_ctor_get(x_14, 0); -lean_inc(x_783); -lean_dec(x_14); -x_784 = lean_ctor_get(x_34, 0); +lean_object* x_784; lean_object* x_785; lean_object* x_786; lean_object* x_787; lean_object* x_788; lean_object* x_789; lean_object* x_790; lean_object* x_791; uint8_t x_792; +x_784 = lean_ctor_get(x_15, 0); lean_inc(x_784); -if (lean_is_exclusive(x_34)) { - lean_ctor_release(x_34, 0); - lean_ctor_release(x_34, 1); - x_785 = x_34; -} else { - lean_dec_ref(x_34); - x_785 = lean_box(0); -} -x_786 = lean_ctor_get(x_35, 1); -lean_inc(x_786); +lean_dec(x_15); +x_785 = lean_ctor_get(x_35, 0); +lean_inc(x_785); if (lean_is_exclusive(x_35)) { lean_ctor_release(x_35, 0); lean_ctor_release(x_35, 1); - x_787 = x_35; + x_786 = x_35; } else { lean_dec_ref(x_35); - x_787 = lean_box(0); + x_786 = lean_box(0); +} +x_787 = lean_ctor_get(x_36, 1); +lean_inc(x_787); +if (lean_is_exclusive(x_36)) { + lean_ctor_release(x_36, 0); + lean_ctor_release(x_36, 1); + x_788 = x_36; +} else { + lean_dec_ref(x_36); + x_788 = lean_box(0); } -x_788 = lean_ctor_get(x_36, 0); -lean_inc(x_788); -x_789 = lean_ctor_get(x_36, 1); +x_789 = lean_ctor_get(x_37, 0); lean_inc(x_789); -x_790 = lean_ctor_get(x_36, 2); +x_790 = lean_ctor_get(x_37, 1); lean_inc(x_790); -x_791 = lean_nat_dec_lt(x_788, x_789); -if (x_791 == 0) +x_791 = lean_ctor_get(x_37, 2); +lean_inc(x_791); +x_792 = lean_nat_dec_lt(x_789, x_790); +if (x_792 == 0) { -lean_object* x_792; lean_object* x_793; lean_object* x_794; lean_object* x_795; +lean_object* x_793; lean_object* x_794; lean_object* x_795; lean_object* x_796; +lean_dec(x_791); lean_dec(x_790); lean_dec(x_789); -lean_dec(x_788); -lean_dec(x_24); -if (lean_is_scalar(x_787)) { - x_792 = lean_alloc_ctor(0, 2, 0); +lean_dec(x_25); +if (lean_is_scalar(x_788)) { + x_793 = lean_alloc_ctor(0, 2, 0); } else { - x_792 = x_787; + x_793 = x_788; } -lean_ctor_set(x_792, 0, x_36); -lean_ctor_set(x_792, 1, x_786); -if (lean_is_scalar(x_785)) { - x_793 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_793, 0, x_37); +lean_ctor_set(x_793, 1, x_787); +if (lean_is_scalar(x_786)) { + x_794 = lean_alloc_ctor(0, 2, 0); } else { - x_793 = x_785; + x_794 = x_786; } -lean_ctor_set(x_793, 0, x_784); -lean_ctor_set(x_793, 1, x_792); -x_794 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_794, 0, x_783); +lean_ctor_set(x_794, 0, x_785); lean_ctor_set(x_794, 1, x_793); -x_795 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_795, 0, x_794); -x_25 = x_795; -x_26 = x_21; -goto block_33; -} -else -{ -lean_object* x_796; lean_object* x_797; lean_object* x_798; lean_object* x_799; lean_object* x_800; lean_object* x_801; uint8_t x_802; -if (lean_is_exclusive(x_36)) { - lean_ctor_release(x_36, 0); - lean_ctor_release(x_36, 1); - lean_ctor_release(x_36, 2); - x_796 = x_36; +x_795 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_795, 0, x_784); +lean_ctor_set(x_795, 1, x_794); +x_796 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_796, 0, x_795); +x_26 = x_796; +x_27 = x_22; +goto block_34; +} +else +{ +lean_object* x_797; lean_object* x_798; lean_object* x_799; lean_object* x_800; lean_object* x_801; lean_object* x_802; uint8_t x_803; +if (lean_is_exclusive(x_37)) { + lean_ctor_release(x_37, 0); + lean_ctor_release(x_37, 1); + lean_ctor_release(x_37, 2); + x_797 = x_37; } else { - lean_dec_ref(x_36); - x_796 = lean_box(0); + lean_dec_ref(x_37); + x_797 = lean_box(0); } -x_797 = lean_nat_add(x_788, x_790); -if (lean_is_scalar(x_796)) { - x_798 = lean_alloc_ctor(0, 3, 0); +x_798 = lean_nat_add(x_789, x_791); +if (lean_is_scalar(x_797)) { + x_799 = lean_alloc_ctor(0, 3, 0); } else { - x_798 = x_796; -} -lean_ctor_set(x_798, 0, x_797); -lean_ctor_set(x_798, 1, x_789); -lean_ctor_set(x_798, 2, x_790); -x_799 = lean_ctor_get(x_784, 0); -lean_inc(x_799); -x_800 = lean_ctor_get(x_784, 1); + x_799 = x_797; +} +lean_ctor_set(x_799, 0, x_798); +lean_ctor_set(x_799, 1, x_790); +lean_ctor_set(x_799, 2, x_791); +x_800 = lean_ctor_get(x_785, 0); lean_inc(x_800); -x_801 = lean_ctor_get(x_784, 2); +x_801 = lean_ctor_get(x_785, 1); lean_inc(x_801); -x_802 = lean_nat_dec_lt(x_800, x_801); -if (x_802 == 0) +x_802 = lean_ctor_get(x_785, 2); +lean_inc(x_802); +x_803 = lean_nat_dec_lt(x_801, x_802); +if (x_803 == 0) { -lean_object* x_803; lean_object* x_804; lean_object* x_805; lean_object* x_806; +lean_object* x_804; lean_object* x_805; lean_object* x_806; lean_object* x_807; +lean_dec(x_802); lean_dec(x_801); lean_dec(x_800); -lean_dec(x_799); -lean_dec(x_788); -lean_dec(x_24); -if (lean_is_scalar(x_787)) { - x_803 = lean_alloc_ctor(0, 2, 0); +lean_dec(x_789); +lean_dec(x_25); +if (lean_is_scalar(x_788)) { + x_804 = lean_alloc_ctor(0, 2, 0); } else { - x_803 = x_787; + x_804 = x_788; } -lean_ctor_set(x_803, 0, x_798); -lean_ctor_set(x_803, 1, x_786); -if (lean_is_scalar(x_785)) { - x_804 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_804, 0, x_799); +lean_ctor_set(x_804, 1, x_787); +if (lean_is_scalar(x_786)) { + x_805 = lean_alloc_ctor(0, 2, 0); } else { - x_804 = x_785; + x_805 = x_786; } -lean_ctor_set(x_804, 0, x_784); -lean_ctor_set(x_804, 1, x_803); -x_805 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_805, 0, x_783); +lean_ctor_set(x_805, 0, x_785); lean_ctor_set(x_805, 1, x_804); -x_806 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_806, 0, x_805); -x_25 = x_806; -x_26 = x_21; -goto block_33; -} -else -{ -lean_object* x_807; lean_object* x_808; lean_object* x_809; lean_object* x_810; lean_object* x_811; lean_object* x_812; lean_object* x_813; lean_object* x_814; uint8_t x_815; -if (lean_is_exclusive(x_784)) { - lean_ctor_release(x_784, 0); - lean_ctor_release(x_784, 1); - lean_ctor_release(x_784, 2); - x_807 = x_784; +x_806 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_806, 0, x_784); +lean_ctor_set(x_806, 1, x_805); +x_807 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_807, 0, x_806); +x_26 = x_807; +x_27 = x_22; +goto block_34; +} +else +{ +lean_object* x_808; lean_object* x_809; lean_object* x_810; lean_object* x_811; lean_object* x_812; lean_object* x_813; lean_object* x_814; lean_object* x_815; uint8_t x_816; +if (lean_is_exclusive(x_785)) { + lean_ctor_release(x_785, 0); + lean_ctor_release(x_785, 1); + lean_ctor_release(x_785, 2); + x_808 = x_785; } else { - lean_dec_ref(x_784); - x_807 = lean_box(0); + lean_dec_ref(x_785); + x_808 = lean_box(0); } -x_808 = lean_array_fget(x_799, x_800); -x_809 = lean_unsigned_to_nat(1u); -x_810 = lean_nat_add(x_800, x_809); -lean_dec(x_800); -if (lean_is_scalar(x_807)) { - x_811 = lean_alloc_ctor(0, 3, 0); +x_809 = lean_array_fget(x_800, x_801); +x_810 = lean_unsigned_to_nat(1u); +x_811 = lean_nat_add(x_801, x_810); +lean_dec(x_801); +if (lean_is_scalar(x_808)) { + x_812 = lean_alloc_ctor(0, 3, 0); } else { - x_811 = x_807; -} -lean_ctor_set(x_811, 0, x_799); -lean_ctor_set(x_811, 1, x_810); -lean_ctor_set(x_811, 2, x_801); -x_812 = lean_ctor_get(x_783, 0); -lean_inc(x_812); -x_813 = lean_ctor_get(x_783, 1); + x_812 = x_808; +} +lean_ctor_set(x_812, 0, x_800); +lean_ctor_set(x_812, 1, x_811); +lean_ctor_set(x_812, 2, x_802); +x_813 = lean_ctor_get(x_784, 0); lean_inc(x_813); -x_814 = lean_ctor_get(x_783, 2); +x_814 = lean_ctor_get(x_784, 1); lean_inc(x_814); -x_815 = lean_nat_dec_lt(x_813, x_814); -if (x_815 == 0) +x_815 = lean_ctor_get(x_784, 2); +lean_inc(x_815); +x_816 = lean_nat_dec_lt(x_814, x_815); +if (x_816 == 0) { -lean_object* x_816; lean_object* x_817; lean_object* x_818; lean_object* x_819; +lean_object* x_817; lean_object* x_818; lean_object* x_819; lean_object* x_820; +lean_dec(x_815); lean_dec(x_814); lean_dec(x_813); -lean_dec(x_812); -lean_dec(x_808); -lean_dec(x_788); -lean_dec(x_24); -if (lean_is_scalar(x_787)) { - x_816 = lean_alloc_ctor(0, 2, 0); +lean_dec(x_809); +lean_dec(x_789); +lean_dec(x_25); +if (lean_is_scalar(x_788)) { + x_817 = lean_alloc_ctor(0, 2, 0); } else { - x_816 = x_787; + x_817 = x_788; } -lean_ctor_set(x_816, 0, x_798); -lean_ctor_set(x_816, 1, x_786); -if (lean_is_scalar(x_785)) { - x_817 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_817, 0, x_799); +lean_ctor_set(x_817, 1, x_787); +if (lean_is_scalar(x_786)) { + x_818 = lean_alloc_ctor(0, 2, 0); } else { - x_817 = x_785; + x_818 = x_786; } -lean_ctor_set(x_817, 0, x_811); -lean_ctor_set(x_817, 1, x_816); -x_818 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_818, 0, x_783); +lean_ctor_set(x_818, 0, x_812); lean_ctor_set(x_818, 1, x_817); -x_819 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_819, 0, x_818); -x_25 = x_819; -x_26 = x_21; -goto block_33; -} -else -{ -lean_object* x_820; lean_object* x_821; lean_object* x_822; lean_object* x_823; lean_object* x_824; lean_object* x_825; lean_object* x_826; lean_object* x_827; lean_object* x_828; lean_object* x_829; lean_object* x_830; uint8_t x_831; -lean_dec(x_787); -lean_dec(x_785); -if (lean_is_exclusive(x_783)) { - lean_ctor_release(x_783, 0); - lean_ctor_release(x_783, 1); - lean_ctor_release(x_783, 2); - x_820 = x_783; +x_819 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_819, 0, x_784); +lean_ctor_set(x_819, 1, x_818); +x_820 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_820, 0, x_819); +x_26 = x_820; +x_27 = x_22; +goto block_34; +} +else +{ +lean_object* x_821; lean_object* x_822; lean_object* x_823; lean_object* x_824; lean_object* x_825; lean_object* x_826; lean_object* x_827; lean_object* x_828; lean_object* x_829; lean_object* x_830; lean_object* x_831; lean_object* x_832; uint8_t x_833; +lean_dec(x_788); +lean_dec(x_786); +if (lean_is_exclusive(x_784)) { + lean_ctor_release(x_784, 0); + lean_ctor_release(x_784, 1); + lean_ctor_release(x_784, 2); + x_821 = x_784; } else { - lean_dec_ref(x_783); - x_820 = lean_box(0); + lean_dec_ref(x_784); + x_821 = lean_box(0); } -x_821 = lean_array_fget(x_812, x_813); -x_822 = lean_nat_add(x_813, x_809); -lean_dec(x_813); -if (lean_is_scalar(x_820)) { - x_823 = lean_alloc_ctor(0, 3, 0); +x_822 = lean_array_fget(x_813, x_814); +x_823 = lean_nat_add(x_814, x_810); +lean_dec(x_814); +if (lean_is_scalar(x_821)) { + x_824 = lean_alloc_ctor(0, 3, 0); } else { - x_823 = x_820; + x_824 = x_821; } -lean_ctor_set(x_823, 0, x_812); -lean_ctor_set(x_823, 1, x_822); -lean_ctor_set(x_823, 2, x_814); -lean_inc(x_6); -lean_inc(x_8); -x_824 = lean_array_push(x_8, x_6); -x_825 = lean_nat_sub(x_821, x_808); -lean_dec(x_808); -lean_dec(x_821); -x_826 = lean_nat_sub(x_825, x_809); -lean_dec(x_825); -x_827 = lean_unsigned_to_nat(0u); -lean_inc(x_826); -lean_inc(x_9); -x_828 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__1(x_9, x_826, x_827, x_826, x_809, x_824, x_15, x_16, x_17, x_18, x_19, x_20, x_21); -lean_dec(x_826); -x_829 = lean_ctor_get(x_828, 0); -lean_inc(x_829); -x_830 = lean_ctor_get(x_828, 1); -lean_inc(x_830); +lean_ctor_set(x_824, 0, x_813); +lean_ctor_set(x_824, 1, x_823); +lean_ctor_set(x_824, 2, x_815); +x_825 = l___private_Lean_Elab_PreDefinition_WF_Rel_0__Lean_Elab_WF_unpackUnary_go___lambda__1___closed__1; +lean_inc(x_7); +x_826 = lean_array_push(x_825, x_7); +x_827 = lean_nat_sub(x_822, x_809); +lean_dec(x_809); +lean_dec(x_822); +x_828 = lean_nat_sub(x_827, x_810); +lean_dec(x_827); +x_829 = lean_unsigned_to_nat(0u); +lean_inc(x_828); +lean_inc(x_10); +x_830 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__1(x_2, x_10, x_828, x_829, x_828, x_810, x_826, x_16, x_17, x_18, x_19, x_20, x_21, x_22); lean_dec(x_828); -x_831 = lean_nat_dec_lt(x_809, x_10); -if (x_831 == 0) -{ -lean_object* x_832; lean_object* x_833; lean_object* x_834; lean_object* x_835; lean_object* x_836; lean_object* x_837; lean_object* x_838; lean_object* x_839; lean_object* x_840; lean_object* x_841; lean_object* x_842; lean_object* x_843; lean_object* x_844; lean_object* x_845; lean_object* x_846; lean_object* x_847; lean_object* x_848; lean_object* x_849; lean_object* x_850; lean_object* x_851; lean_object* x_852; lean_object* x_853; lean_object* x_854; lean_object* x_855; lean_object* x_856; lean_object* x_857; lean_object* x_858; lean_object* x_859; lean_object* x_860; lean_object* x_861; lean_object* x_862; lean_object* x_863; -lean_dec(x_788); -lean_inc(x_19); -x_832 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_19, x_20, x_830); -x_833 = lean_ctor_get(x_832, 0); -lean_inc(x_833); -x_834 = lean_ctor_get(x_832, 1); -lean_inc(x_834); -lean_dec(x_832); -x_835 = lean_ctor_get(x_19, 10); +x_831 = lean_ctor_get(x_830, 0); +lean_inc(x_831); +x_832 = lean_ctor_get(x_830, 1); +lean_inc(x_832); +lean_dec(x_830); +x_833 = lean_nat_dec_lt(x_810, x_11); +if (x_833 == 0) +{ +lean_object* x_834; lean_object* x_835; lean_object* x_836; lean_object* x_837; lean_object* x_838; lean_object* x_839; lean_object* x_840; lean_object* x_841; lean_object* x_842; lean_object* x_843; lean_object* x_844; lean_object* x_845; lean_object* x_846; lean_object* x_847; lean_object* x_848; lean_object* x_849; lean_object* x_850; lean_object* x_851; lean_object* x_852; lean_object* x_853; lean_object* x_854; lean_object* x_855; lean_object* x_856; lean_object* x_857; lean_object* x_858; lean_object* x_859; lean_object* x_860; lean_object* x_861; lean_object* x_862; lean_object* x_863; lean_object* x_864; lean_object* x_865; +lean_dec(x_789); +lean_inc(x_20); +x_834 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_20, x_21, x_832); +x_835 = lean_ctor_get(x_834, 0); lean_inc(x_835); -x_836 = lean_st_ref_get(x_20, x_834); -x_837 = lean_ctor_get(x_836, 0); +x_836 = lean_ctor_get(x_834, 1); +lean_inc(x_836); +lean_dec(x_834); +x_837 = lean_ctor_get(x_20, 10); lean_inc(x_837); -x_838 = lean_ctor_get(x_836, 1); -lean_inc(x_838); -lean_dec(x_836); -x_839 = lean_ctor_get(x_837, 0); +x_838 = lean_st_ref_get(x_21, x_836); +x_839 = lean_ctor_get(x_838, 0); lean_inc(x_839); -lean_dec(x_837); -x_840 = lean_environment_main_module(x_839); -x_841 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__1; -lean_inc(x_7); -x_842 = lean_name_mk_string(x_7, x_841); -x_843 = l_Subarray_forInUnsafe_loop___at_Lean_Elab_WF_getForbiddenByTrivialSizeOf___spec__1___closed__3; -x_844 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__2; -lean_inc(x_2); -x_845 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_845, 0, x_843); -lean_ctor_set(x_845, 1, x_2); -lean_ctor_set(x_845, 2, x_844); -x_846 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__3; -lean_inc(x_835); +x_840 = lean_ctor_get(x_838, 1); lean_inc(x_840); -x_847 = l_Lean_addMacroScope(x_840, x_846, x_835); -x_848 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__4; -lean_inc(x_5); -x_849 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_849, 0, x_848); -lean_ctor_set(x_849, 1, x_5); -lean_inc(x_833); -x_850 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_850, 0, x_833); -lean_ctor_set(x_850, 1, x_845); -lean_ctor_set(x_850, 2, x_847); -lean_ctor_set(x_850, 3, x_849); -lean_inc(x_4); -x_851 = l_Lean_addMacroScope(x_840, x_4, x_835); -lean_inc(x_5); +lean_dec(x_838); +x_841 = lean_ctor_get(x_839, 0); +lean_inc(x_841); +lean_dec(x_839); +x_842 = lean_environment_main_module(x_841); +x_843 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__1; +lean_inc(x_8); +x_844 = lean_name_mk_string(x_8, x_843); +x_845 = l_Subarray_forInUnsafe_loop___at_Lean_Elab_WF_getForbiddenByTrivialSizeOf___spec__1___closed__3; +x_846 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__2; lean_inc(x_3); +x_847 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_847, 0, x_845); +lean_ctor_set(x_847, 1, x_3); +lean_ctor_set(x_847, 2, x_846); +x_848 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__3; +lean_inc(x_837); +lean_inc(x_842); +x_849 = l_Lean_addMacroScope(x_842, x_848, x_837); +x_850 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__4; +lean_inc(x_6); +x_851 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_851, 0, x_850); +lean_ctor_set(x_851, 1, x_6); +lean_inc(x_835); x_852 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_852, 0, x_833); -lean_ctor_set(x_852, 1, x_3); -lean_ctor_set(x_852, 2, x_851); -lean_ctor_set(x_852, 3, x_5); -lean_inc(x_8); -x_853 = lean_array_push(x_8, x_852); -x_854 = lean_box(2); -x_855 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__6; -x_856 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_856, 0, x_854); -lean_ctor_set(x_856, 1, x_855); -lean_ctor_set(x_856, 2, x_853); -x_857 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__7; -x_858 = lean_array_push(x_857, x_850); -x_859 = lean_array_push(x_858, x_856); -x_860 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_860, 0, x_854); -lean_ctor_set(x_860, 1, x_842); -lean_ctor_set(x_860, 2, x_859); -x_861 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___lambda__1(x_24, x_829, x_798, x_811, x_823, x_786, x_860, x_15, x_16, x_17, x_18, x_19, x_20, x_838); -lean_dec(x_24); -x_862 = lean_ctor_get(x_861, 0); -lean_inc(x_862); -x_863 = lean_ctor_get(x_861, 1); -lean_inc(x_863); -lean_dec(x_861); -x_25 = x_862; -x_26 = x_863; -goto block_33; +lean_ctor_set(x_852, 0, x_835); +lean_ctor_set(x_852, 1, x_847); +lean_ctor_set(x_852, 2, x_849); +lean_ctor_set(x_852, 3, x_851); +lean_inc(x_5); +x_853 = l_Lean_addMacroScope(x_842, x_5, x_837); +lean_inc(x_6); +lean_inc(x_4); +x_854 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_854, 0, x_835); +lean_ctor_set(x_854, 1, x_4); +lean_ctor_set(x_854, 2, x_853); +lean_ctor_set(x_854, 3, x_6); +lean_inc(x_9); +x_855 = lean_array_push(x_9, x_854); +x_856 = lean_box(2); +x_857 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__6; +x_858 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_858, 0, x_856); +lean_ctor_set(x_858, 1, x_857); +lean_ctor_set(x_858, 2, x_855); +x_859 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__7; +x_860 = lean_array_push(x_859, x_852); +x_861 = lean_array_push(x_860, x_858); +x_862 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_862, 0, x_856); +lean_ctor_set(x_862, 1, x_844); +lean_ctor_set(x_862, 2, x_861); +x_863 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___lambda__1(x_25, x_831, x_799, x_812, x_824, x_787, x_862, x_16, x_17, x_18, x_19, x_20, x_21, x_840); +lean_dec(x_25); +x_864 = lean_ctor_get(x_863, 0); +lean_inc(x_864); +x_865 = lean_ctor_get(x_863, 1); +lean_inc(x_865); +lean_dec(x_863); +x_26 = x_864; +x_27 = x_865; +goto block_34; } else { -lean_object* x_864; lean_object* x_865; lean_object* x_866; lean_object* x_867; lean_object* x_868; lean_object* x_869; lean_object* x_870; lean_object* x_871; lean_object* x_872; lean_object* x_873; lean_object* x_874; lean_object* x_875; lean_object* x_876; lean_object* x_877; lean_object* x_878; lean_object* x_879; lean_object* x_880; lean_object* x_881; lean_object* x_882; lean_object* x_883; lean_object* x_884; lean_object* x_885; lean_object* x_886; lean_object* x_887; lean_object* x_888; lean_object* x_889; lean_object* x_890; lean_object* x_891; lean_object* x_892; lean_object* x_893; lean_object* x_894; lean_object* x_895; lean_object* x_896; lean_object* x_897; lean_object* x_898; lean_object* x_899; lean_object* x_900; lean_object* x_901; lean_object* x_902; lean_object* x_903; lean_object* x_904; lean_object* x_905; lean_object* x_906; lean_object* x_907; lean_object* x_908; lean_object* x_909; lean_object* x_910; lean_object* x_911; lean_object* x_912; lean_object* x_913; lean_object* x_914; lean_object* x_915; lean_object* x_916; lean_object* x_917; lean_object* x_918; lean_object* x_919; lean_object* x_920; lean_object* x_921; lean_object* x_922; lean_object* x_923; -lean_inc(x_19); -x_864 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_19, x_20, x_830); -x_865 = lean_ctor_get(x_864, 0); -lean_inc(x_865); -x_866 = lean_ctor_get(x_864, 1); -lean_inc(x_866); -lean_dec(x_864); -x_867 = lean_ctor_get(x_19, 10); +lean_object* x_866; lean_object* x_867; lean_object* x_868; lean_object* x_869; lean_object* x_870; lean_object* x_871; lean_object* x_872; lean_object* x_873; lean_object* x_874; lean_object* x_875; lean_object* x_876; lean_object* x_877; lean_object* x_878; lean_object* x_879; lean_object* x_880; lean_object* x_881; lean_object* x_882; lean_object* x_883; lean_object* x_884; lean_object* x_885; lean_object* x_886; lean_object* x_887; lean_object* x_888; lean_object* x_889; lean_object* x_890; lean_object* x_891; lean_object* x_892; lean_object* x_893; lean_object* x_894; lean_object* x_895; lean_object* x_896; lean_object* x_897; lean_object* x_898; lean_object* x_899; lean_object* x_900; lean_object* x_901; lean_object* x_902; lean_object* x_903; lean_object* x_904; lean_object* x_905; lean_object* x_906; lean_object* x_907; lean_object* x_908; lean_object* x_909; lean_object* x_910; lean_object* x_911; lean_object* x_912; lean_object* x_913; lean_object* x_914; lean_object* x_915; lean_object* x_916; lean_object* x_917; lean_object* x_918; lean_object* x_919; lean_object* x_920; lean_object* x_921; lean_object* x_922; lean_object* x_923; lean_object* x_924; +lean_inc(x_20); +x_866 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_20, x_21, x_832); +x_867 = lean_ctor_get(x_866, 0); lean_inc(x_867); -x_868 = lean_st_ref_get(x_20, x_866); -x_869 = lean_ctor_get(x_868, 0); +x_868 = lean_ctor_get(x_866, 1); +lean_inc(x_868); +lean_dec(x_866); +x_869 = lean_ctor_get(x_20, 10); lean_inc(x_869); -x_870 = lean_ctor_get(x_868, 1); -lean_inc(x_870); -lean_dec(x_868); -x_871 = lean_ctor_get(x_869, 0); +x_870 = lean_st_ref_get(x_21, x_868); +x_871 = lean_ctor_get(x_870, 0); lean_inc(x_871); -lean_dec(x_869); -x_872 = lean_environment_main_module(x_871); -x_873 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__8; -lean_inc(x_7); -x_874 = lean_name_mk_string(x_7, x_873); -x_875 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__9; -lean_inc(x_865); -x_876 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_876, 0, x_865); -lean_ctor_set(x_876, 1, x_875); -x_877 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__1; -lean_inc(x_7); -x_878 = lean_name_mk_string(x_7, x_877); -x_879 = l_Subarray_forInUnsafe_loop___at_Lean_Elab_WF_getForbiddenByTrivialSizeOf___spec__1___closed__3; -x_880 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__2; -lean_inc(x_2); -x_881 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_881, 0, x_879); -lean_ctor_set(x_881, 1, x_2); -lean_ctor_set(x_881, 2, x_880); -x_882 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__3; -lean_inc(x_867); +x_872 = lean_ctor_get(x_870, 1); lean_inc(x_872); -x_883 = l_Lean_addMacroScope(x_872, x_882, x_867); -x_884 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__4; -lean_inc(x_5); -x_885 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_885, 0, x_884); -lean_ctor_set(x_885, 1, x_5); -lean_inc(x_865); -x_886 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_886, 0, x_865); -lean_ctor_set(x_886, 1, x_881); -lean_ctor_set(x_886, 2, x_883); -lean_ctor_set(x_886, 3, x_885); -lean_inc(x_4); -x_887 = l_Lean_addMacroScope(x_872, x_4, x_867); -lean_inc(x_5); -lean_inc(x_3); -lean_inc(x_865); -x_888 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_888, 0, x_865); -lean_ctor_set(x_888, 1, x_3); -lean_ctor_set(x_888, 2, x_887); -lean_ctor_set(x_888, 3, x_5); +lean_dec(x_870); +x_873 = lean_ctor_get(x_871, 0); +lean_inc(x_873); +lean_dec(x_871); +x_874 = lean_environment_main_module(x_873); +x_875 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__8; lean_inc(x_8); -x_889 = lean_array_push(x_8, x_888); -x_890 = lean_box(2); -x_891 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__6; -x_892 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_892, 0, x_890); -lean_ctor_set(x_892, 1, x_891); -lean_ctor_set(x_892, 2, x_889); -x_893 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__7; -x_894 = lean_array_push(x_893, x_886); -x_895 = lean_array_push(x_894, x_892); -x_896 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_896, 0, x_890); -lean_ctor_set(x_896, 1, x_878); -lean_ctor_set(x_896, 2, x_895); -x_897 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__10; -lean_inc(x_7); -x_898 = lean_name_mk_string(x_7, x_897); -x_899 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__11; -lean_inc(x_865); -x_900 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_900, 0, x_865); -lean_ctor_set(x_900, 1, x_899); -x_901 = l_Nat_repr(x_788); -x_902 = l_Lean_numLitKind; -x_903 = l_Lean_Syntax_mkLit(x_902, x_901, x_890); +x_876 = lean_name_mk_string(x_8, x_875); +x_877 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__9; +lean_inc(x_867); +x_878 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_878, 0, x_867); +lean_ctor_set(x_878, 1, x_877); +x_879 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__1; lean_inc(x_8); -x_904 = lean_array_push(x_8, x_903); -x_905 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_905, 0, x_890); -lean_ctor_set(x_905, 1, x_891); -lean_ctor_set(x_905, 2, x_904); -x_906 = lean_array_push(x_893, x_900); -x_907 = lean_array_push(x_906, x_905); -x_908 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_908, 0, x_890); -lean_ctor_set(x_908, 1, x_898); -lean_ctor_set(x_908, 2, x_907); +x_880 = lean_name_mk_string(x_8, x_879); +x_881 = l_Subarray_forInUnsafe_loop___at_Lean_Elab_WF_getForbiddenByTrivialSizeOf___spec__1___closed__3; +x_882 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__2; +lean_inc(x_3); +x_883 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_883, 0, x_881); +lean_ctor_set(x_883, 1, x_3); +lean_ctor_set(x_883, 2, x_882); +x_884 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__3; +lean_inc(x_869); +lean_inc(x_874); +x_885 = l_Lean_addMacroScope(x_874, x_884, x_869); +x_886 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__4; +lean_inc(x_6); +x_887 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_887, 0, x_886); +lean_ctor_set(x_887, 1, x_6); +lean_inc(x_867); +x_888 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_888, 0, x_867); +lean_ctor_set(x_888, 1, x_883); +lean_ctor_set(x_888, 2, x_885); +lean_ctor_set(x_888, 3, x_887); +lean_inc(x_5); +x_889 = l_Lean_addMacroScope(x_874, x_5, x_869); +lean_inc(x_6); +lean_inc(x_4); +lean_inc(x_867); +x_890 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_890, 0, x_867); +lean_ctor_set(x_890, 1, x_4); +lean_ctor_set(x_890, 2, x_889); +lean_ctor_set(x_890, 3, x_6); +lean_inc(x_9); +x_891 = lean_array_push(x_9, x_890); +x_892 = lean_box(2); +x_893 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__6; +x_894 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_894, 0, x_892); +lean_ctor_set(x_894, 1, x_893); +lean_ctor_set(x_894, 2, x_891); +x_895 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__7; +x_896 = lean_array_push(x_895, x_888); +x_897 = lean_array_push(x_896, x_894); +x_898 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_898, 0, x_892); +lean_ctor_set(x_898, 1, x_880); +lean_ctor_set(x_898, 2, x_897); +x_899 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__10; lean_inc(x_8); -x_909 = lean_array_push(x_8, x_908); -x_910 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_910, 0, x_890); -lean_ctor_set(x_910, 1, x_891); -lean_ctor_set(x_910, 2, x_909); -x_911 = lean_array_push(x_893, x_896); -x_912 = lean_array_push(x_911, x_910); -x_913 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_913, 0, x_890); -lean_ctor_set(x_913, 1, x_891); -lean_ctor_set(x_913, 2, x_912); -x_914 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__12; -x_915 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_915, 0, x_865); -lean_ctor_set(x_915, 1, x_914); -x_916 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__13; -x_917 = lean_array_push(x_916, x_876); -x_918 = lean_array_push(x_917, x_913); -x_919 = lean_array_push(x_918, x_915); -x_920 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_920, 0, x_890); -lean_ctor_set(x_920, 1, x_874); -lean_ctor_set(x_920, 2, x_919); -x_921 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___lambda__1(x_24, x_829, x_798, x_811, x_823, x_786, x_920, x_15, x_16, x_17, x_18, x_19, x_20, x_870); -lean_dec(x_24); -x_922 = lean_ctor_get(x_921, 0); -lean_inc(x_922); -x_923 = lean_ctor_get(x_921, 1); +x_900 = lean_name_mk_string(x_8, x_899); +x_901 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__11; +lean_inc(x_867); +x_902 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_902, 0, x_867); +lean_ctor_set(x_902, 1, x_901); +x_903 = l_Nat_repr(x_789); +x_904 = l_Lean_Syntax_mkNumLit(x_903, x_892); +lean_inc(x_9); +x_905 = lean_array_push(x_9, x_904); +x_906 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_906, 0, x_892); +lean_ctor_set(x_906, 1, x_893); +lean_ctor_set(x_906, 2, x_905); +x_907 = lean_array_push(x_895, x_902); +x_908 = lean_array_push(x_907, x_906); +x_909 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_909, 0, x_892); +lean_ctor_set(x_909, 1, x_900); +lean_ctor_set(x_909, 2, x_908); +lean_inc(x_9); +x_910 = lean_array_push(x_9, x_909); +x_911 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_911, 0, x_892); +lean_ctor_set(x_911, 1, x_893); +lean_ctor_set(x_911, 2, x_910); +x_912 = lean_array_push(x_895, x_898); +x_913 = lean_array_push(x_912, x_911); +x_914 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_914, 0, x_892); +lean_ctor_set(x_914, 1, x_893); +lean_ctor_set(x_914, 2, x_913); +x_915 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__12; +x_916 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_916, 0, x_867); +lean_ctor_set(x_916, 1, x_915); +x_917 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___closed__13; +x_918 = lean_array_push(x_917, x_878); +x_919 = lean_array_push(x_918, x_914); +x_920 = lean_array_push(x_919, x_916); +x_921 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_921, 0, x_892); +lean_ctor_set(x_921, 1, x_876); +lean_ctor_set(x_921, 2, x_920); +x_922 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___lambda__1(x_25, x_831, x_799, x_812, x_824, x_787, x_921, x_16, x_17, x_18, x_19, x_20, x_21, x_872); +lean_dec(x_25); +x_923 = lean_ctor_get(x_922, 0); lean_inc(x_923); -lean_dec(x_921); -x_25 = x_922; +x_924 = lean_ctor_get(x_922, 1); +lean_inc(x_924); +lean_dec(x_922); x_26 = x_923; -goto block_33; +x_27 = x_924; +goto block_34; } } } } } -block_33: +block_34: { -if (lean_obj_tag(x_25) == 0) +if (lean_obj_tag(x_26) == 0) { -lean_object* x_27; lean_object* x_28; -lean_dec(x_19); +lean_object* x_28; lean_object* x_29; +lean_dec(x_20); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -9036,26 +9032,25 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -x_27 = lean_ctor_get(x_25, 0); -lean_inc(x_27); -lean_dec(x_25); -x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_27); -lean_ctor_set(x_28, 1, x_26); -return x_28; +x_28 = lean_ctor_get(x_26, 0); +lean_inc(x_28); +lean_dec(x_26); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_27); +return x_29; } else { -lean_object* x_29; size_t x_30; size_t x_31; -x_29 = lean_ctor_get(x_25, 0); -lean_inc(x_29); -lean_dec(x_25); -x_30 = 1; -x_31 = lean_usize_add(x_13, x_30); -x_13 = x_31; -x_14 = x_29; -x_21 = x_26; +lean_object* x_30; size_t x_31; size_t x_32; +x_30 = lean_ctor_get(x_26, 0); +lean_inc(x_30); +lean_dec(x_26); +x_31 = 1; +x_32 = lean_usize_add(x_14, x_31); +x_14 = x_32; +x_15 = x_30; +x_22 = x_27; goto _start; } } @@ -9066,26 +9061,56 @@ static lean_object* _init_l_Lean_Elab_WF_elabWFRel_generateElements___closed__1( _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("x", 1); +x_1 = lean_mk_string_from_bytes("term", 4); return x_1; } } static lean_object* _init_l_Lean_Elab_WF_elabWFRel_generateElements___closed__2() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_WF_elabWFRel_generateElements___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__2; +x_3 = lean_alloc_ctor(1, 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_Elab_WF_elabWFRel_generateElements___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("x", 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_WF_elabWFRel_generateElements___closed__5() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__1; +x_1 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__4; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_WF_elabWFRel_generateElements___closed__3() { +static lean_object* _init_l_Lean_Elab_WF_elabWFRel_generateElements___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__1; +x_1 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__4; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__2; +x_3 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__5; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -9093,17 +9118,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_WF_elabWFRel_generateElements___closed__4() { +static lean_object* _init_l_Lean_Elab_WF_elabWFRel_generateElements___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__1; +x_2 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_WF_elabWFRel_generateElements___closed__5() { +static lean_object* _init_l_Lean_Elab_WF_elabWFRel_generateElements___closed__8() { _start: { lean_object* x_1; @@ -9111,17 +9136,17 @@ x_1 = lean_mk_string_from_bytes("Lean", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_WF_elabWFRel_generateElements___closed__6() { +static lean_object* _init_l_Lean_Elab_WF_elabWFRel_generateElements___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__5; +x_2 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__8; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_WF_elabWFRel_generateElements___closed__7() { +static lean_object* _init_l_Lean_Elab_WF_elabWFRel_generateElements___closed__10() { _start: { lean_object* x_1; @@ -9129,17 +9154,17 @@ x_1 = lean_mk_string_from_bytes("Parser", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_WF_elabWFRel_generateElements___closed__8() { +static lean_object* _init_l_Lean_Elab_WF_elabWFRel_generateElements___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__6; -x_2 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__7; +x_1 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__9; +x_2 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__10; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_WF_elabWFRel_generateElements___closed__9() { +static lean_object* _init_l_Lean_Elab_WF_elabWFRel_generateElements___closed__12() { _start: { lean_object* x_1; @@ -9147,17 +9172,17 @@ x_1 = lean_mk_string_from_bytes("Term", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_WF_elabWFRel_generateElements___closed__10() { +static lean_object* _init_l_Lean_Elab_WF_elabWFRel_generateElements___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__8; -x_2 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__9; +x_1 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__11; +x_2 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__12; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_WF_elabWFRel_generateElements___closed__11() { +static lean_object* _init_l_Lean_Elab_WF_elabWFRel_generateElements___closed__14() { _start: { lean_object* x_1; @@ -9165,17 +9190,17 @@ x_1 = lean_mk_string_from_bytes("hole", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_WF_elabWFRel_generateElements___closed__12() { +static lean_object* _init_l_Lean_Elab_WF_elabWFRel_generateElements___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__10; -x_2 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__11; +x_1 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__13; +x_2 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__14; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_WF_elabWFRel_generateElements___closed__13() { +static lean_object* _init_l_Lean_Elab_WF_elabWFRel_generateElements___closed__16() { _start: { lean_object* x_1; @@ -9186,7 +9211,7 @@ return x_1; LEAN_EXPORT lean_object* l_Lean_Elab_WF_elabWFRel_generateElements(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_object* x_9, lean_object* x_10) { _start: { -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; 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; size_t x_49; size_t x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; +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; 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; size_t x_49; size_t x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; lean_inc(x_8); x_11 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_8, x_9, x_10); x_12 = lean_ctor_get(x_11, 0); @@ -9206,15 +9231,15 @@ x_18 = lean_ctor_get(x_16, 0); lean_inc(x_18); lean_dec(x_16); x_19 = lean_environment_main_module(x_18); -x_20 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__4; -x_21 = l_Lean_addMacroScope(x_19, x_20, x_14); -x_22 = lean_box(0); -x_23 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__3; +x_20 = lean_box(0); +x_21 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__7; +x_22 = l_Lean_addMacroScope(x_19, x_21, x_14); +x_23 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__6; x_24 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_24, 0, x_12); lean_ctor_set(x_24, 1, x_23); -lean_ctor_set(x_24, 2, x_21); -lean_ctor_set(x_24, 3, x_22); +lean_ctor_set(x_24, 2, x_22); +lean_ctor_set(x_24, 3, x_20); lean_inc(x_8); x_25 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_8, x_9, x_17); x_26 = lean_ctor_get(x_25, 0); @@ -9226,14 +9251,14 @@ x_28 = lean_st_ref_get(x_9, x_27); x_29 = lean_ctor_get(x_28, 1); lean_inc(x_29); lean_dec(x_28); -x_30 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__13; +x_30 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__16; x_31 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_31, 0, x_26); lean_ctor_set(x_31, 1, x_30); x_32 = l___private_Lean_Elab_PreDefinition_WF_Rel_0__Lean_Elab_WF_unpackUnary_go___lambda__1___closed__1; x_33 = lean_array_push(x_32, x_31); x_34 = lean_box(2); -x_35 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__12; +x_35 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__15; x_36 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_36, 0, x_34); lean_ctor_set(x_36, 1, x_35); @@ -9262,59 +9287,61 @@ lean_ctor_set(x_48, 0, x_39); lean_ctor_set(x_48, 1, x_47); x_49 = lean_usize_of_nat(x_42); x_50 = 0; -x_51 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__10; -x_52 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2(x_1, x_38, x_23, x_20, x_22, x_24, x_51, x_32, x_36, x_42, x_1, x_49, x_50, x_48, x_4, x_5, x_6, x_7, x_8, x_9, x_29); +x_51 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__3; +x_52 = l_Lean_Elab_WF_elabWFRel_generateElements___closed__13; +x_53 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2(x_1, x_51, x_38, x_23, x_21, x_20, x_24, x_52, x_32, x_36, x_42, x_1, x_49, x_50, x_48, x_4, x_5, x_6, x_7, x_8, x_9, x_29); lean_dec(x_42); -x_53 = lean_ctor_get(x_52, 0); -lean_inc(x_53); -x_54 = lean_ctor_get(x_53, 1); +x_54 = lean_ctor_get(x_53, 0); lean_inc(x_54); -lean_dec(x_53); x_55 = lean_ctor_get(x_54, 1); lean_inc(x_55); lean_dec(x_54); -x_56 = !lean_is_exclusive(x_52); -if (x_56 == 0) -{ -lean_object* x_57; lean_object* x_58; -x_57 = lean_ctor_get(x_52, 0); -lean_dec(x_57); -x_58 = lean_ctor_get(x_55, 1); -lean_inc(x_58); +x_56 = lean_ctor_get(x_55, 1); +lean_inc(x_56); lean_dec(x_55); -lean_ctor_set(x_52, 0, x_58); -return x_52; +x_57 = !lean_is_exclusive(x_53); +if (x_57 == 0) +{ +lean_object* x_58; lean_object* x_59; +x_58 = lean_ctor_get(x_53, 0); +lean_dec(x_58); +x_59 = lean_ctor_get(x_56, 1); +lean_inc(x_59); +lean_dec(x_56); +lean_ctor_set(x_53, 0, x_59); +return x_53; } else { -lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_59 = lean_ctor_get(x_52, 1); -lean_inc(x_59); -lean_dec(x_52); -x_60 = lean_ctor_get(x_55, 1); +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_53, 1); lean_inc(x_60); -lean_dec(x_55); -x_61 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_61, 0, x_60); -lean_ctor_set(x_61, 1, x_59); -return x_61; +lean_dec(x_53); +x_61 = lean_ctor_get(x_56, 1); +lean_inc(x_61); +lean_dec(x_56); +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_61); +lean_ctor_set(x_62, 1, x_60); +return x_62; } } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__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, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__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, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_14; -x_14 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_object* x_15; +x_15 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -return x_14; +lean_dec(x_1); +return x_15; } } LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2___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, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { @@ -9354,23 +9381,25 @@ lean_object* x_18 = _args[17]; lean_object* x_19 = _args[18]; lean_object* x_20 = _args[19]; lean_object* x_21 = _args[20]; +lean_object* x_22 = _args[21]; _start: { -size_t x_22; size_t x_23; lean_object* x_24; -x_22 = lean_unbox_usize(x_12); -lean_dec(x_12); +size_t x_23; size_t x_24; lean_object* x_25; x_23 = lean_unbox_usize(x_13); lean_dec(x_13); -x_24 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_22, x_23, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21); -lean_dec(x_20); +x_24 = lean_unbox_usize(x_14); +lean_dec(x_14); +x_25 = l_Array_forInUnsafe_loop___at_Lean_Elab_WF_elabWFRel_generateElements___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_23, x_24, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22); +lean_dec(x_21); +lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); -lean_dec(x_15); +lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); +lean_dec(x_2); lean_dec(x_1); -return x_24; +return x_25; } } LEAN_EXPORT lean_object* l_Lean_Elab_WF_elabWFRel_generateElements___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_object* x_9, lean_object* x_10) { @@ -10743,6 +10772,12 @@ l_Lean_Elab_WF_elabWFRel_generateElements___closed__12 = _init_l_Lean_Elab_WF_el lean_mark_persistent(l_Lean_Elab_WF_elabWFRel_generateElements___closed__12); l_Lean_Elab_WF_elabWFRel_generateElements___closed__13 = _init_l_Lean_Elab_WF_elabWFRel_generateElements___closed__13(); lean_mark_persistent(l_Lean_Elab_WF_elabWFRel_generateElements___closed__13); +l_Lean_Elab_WF_elabWFRel_generateElements___closed__14 = _init_l_Lean_Elab_WF_elabWFRel_generateElements___closed__14(); +lean_mark_persistent(l_Lean_Elab_WF_elabWFRel_generateElements___closed__14); +l_Lean_Elab_WF_elabWFRel_generateElements___closed__15 = _init_l_Lean_Elab_WF_elabWFRel_generateElements___closed__15(); +lean_mark_persistent(l_Lean_Elab_WF_elabWFRel_generateElements___closed__15); +l_Lean_Elab_WF_elabWFRel_generateElements___closed__16 = _init_l_Lean_Elab_WF_elabWFRel_generateElements___closed__16(); +lean_mark_persistent(l_Lean_Elab_WF_elabWFRel_generateElements___closed__16); l_Lean_Elab_WF_elabWFRel_guess___rarg___lambda__1___closed__1 = _init_l_Lean_Elab_WF_elabWFRel_guess___rarg___lambda__1___closed__1(); lean_mark_persistent(l_Lean_Elab_WF_elabWFRel_guess___rarg___lambda__1___closed__1); l_Lean_Elab_WF_elabWFRel_guess___rarg___lambda__1___closed__2 = _init_l_Lean_Elab_WF_elabWFRel_guess___rarg___lambda__1___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Print.c b/stage0/stdlib/Lean/Elab/Print.c index f73c02e4ee3..788d64e586d 100644 --- a/stage0/stdlib/Lean/Elab/Print.c +++ b/stage0/stdlib/Lean/Elab/Print.c @@ -34,6 +34,7 @@ lean_object* l_Std_RBNode_insert___at_Lean_NameSet_insert___spec__1(lean_object* static lean_object* l___regBuiltin_Lean_Elab_Command_elabPrint_declRange___closed__2; LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at___private_Lean_Elab_Print_0__Lean_Elab_Command_printId___spec__10(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Print_0__Lean_Elab_Command_printId___closed__1; +lean_object* l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); lean_object* l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Print_0__Lean_Elab_Command_mkHeader___closed__19; static lean_object* l___regBuiltin_Lean_Elab_Command_elabPrint_declRange___closed__4; @@ -73,7 +74,6 @@ lean_object* lean_array_push(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabPrintAxioms___spec__1___rarg___closed__2; static lean_object* l___private_Lean_Elab_Print_0__Lean_Elab_Command_printAxiomsOf___closed__3; -static lean_object* l_Lean_Elab_Command_elabPrint___closed__17; lean_object* l_Lean_MessageData_ofList(lean_object*); static lean_object* l___private_Lean_Elab_Print_0__Lean_Elab_Command_levelParamsToMessageData___closed__2; static lean_object* l_Lean_Elab_Command_elabPrint___closed__4; @@ -115,7 +115,6 @@ static lean_object* l_Lean_Elab_pushInfoLeaf___at___private_Lean_Elab_Print_0__L static lean_object* l___private_Lean_Elab_Print_0__Lean_Elab_Command_mkHeader___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Command_elabPrintAxioms___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Print_0__Lean_Elab_Command_levelParamsToMessageData(lean_object*); -lean_object* l_Lean_Syntax_isStrLit_x3f(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabPrint(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_replaceRef(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Print_0__Lean_Elab_Command_levelParamsToMessageData___closed__3; @@ -148,7 +147,6 @@ static lean_object* l_Lean_Elab_Command_elabPrint___closed__11; static lean_object* l___private_Lean_Elab_Print_0__Lean_Elab_Command_throwUnknownId___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_CollectAxioms_State_axioms___default; static lean_object* l___private_Lean_Elab_Print_0__Lean_Elab_Command_printInduct___closed__6; -lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Array_isEmpty___rarg(lean_object*); static lean_object* l___private_Lean_Elab_Print_0__Lean_Elab_Command_throwUnknownId___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Command_elabPrint_declRange___closed__7; @@ -176,7 +174,6 @@ LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Print_0__L static lean_object* l___regBuiltin_Lean_Elab_Command_elabPrintAxioms_declRange___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabPrint_declRange(lean_object*); static lean_object* l___private_Lean_Elab_Print_0__Lean_Elab_Command_mkHeader___closed__22; -lean_object* l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabPrint___closed__10; LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at___private_Lean_Elab_Print_0__Lean_Elab_Command_printId___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_PersistentArray_push___rarg(lean_object*, lean_object*); @@ -193,7 +190,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_addCompletionInfo___at___private_Lean_Elab_ lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandDeclId___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabPrintAxioms_declRange___closed__6; static lean_object* l___private_Lean_Elab_Print_0__Lean_Elab_Command_mkHeader___closed__18; -static lean_object* l_Lean_Elab_Command_elabPrint___closed__16; static lean_object* l___private_Lean_Elab_Print_0__Lean_Elab_Command_mkHeader___closed__15; static lean_object* l___private_Lean_Elab_Print_0__Lean_Elab_Command_mkHeader___closed__5; static lean_object* l_Lean_Elab_Command_CollectAxioms_State_axioms___default___closed__1; @@ -227,10 +223,10 @@ LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoLeaf___at___private_Lean_Elab_Print uint8_t l_Lean_TagDeclarationExtension_isTagged(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwError___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); +lean_object* l_Lean_TSyntax_getString(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_resolveGlobalConstWithInfos___at___private_Lean_Elab_Print_0__Lean_Elab_Command_printId___spec__4___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Print_0__Lean_Elab_Command_mkHeader___closed__16; static lean_object* l___regBuiltin_Lean_Elab_Command_elabPrintAxioms_declRange___closed__1; -static lean_object* l_Lean_Elab_Command_elabPrint___closed__15; LEAN_EXPORT lean_object* l_Lean_getConstInfo___at___private_Lean_Elab_Print_0__Lean_Elab_Command_printInduct___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabPrint___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Command_elabPrintAxioms_declRange___closed__2; @@ -241,14 +237,12 @@ LEAN_EXPORT lean_object* l_Lean_getReducibilityStatus___at___private_Lean_Elab_P uint8_t l_Lean_NameSet_contains(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Print_0__Lean_Elab_Command_mkHeader___closed__24; LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at___private_Lean_Elab_Print_0__Lean_Elab_Command_printId___spec__7___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_elabPrint___closed__18; static lean_object* l_List_forIn_loop___at___private_Lean_Elab_Print_0__Lean_Elab_Command_levelParamsToMessageData___spec__1___closed__3; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabPrint___closed__7; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Print_0__Lean_Elab_Command_printInduct___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getUsedConstants(lean_object*); uint8_t l_List_isEmpty___rarg(lean_object*); -lean_object* l_panic___at_Lean_Name_getString_x21___spec__1(lean_object*); lean_object* l_Lean_Elab_Command_getScope___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_resolveGlobalName___at___private_Lean_Elab_Print_0__Lean_Elab_Command_printId___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Print_0__Lean_Elab_Command_printAxiomsOf___closed__4; @@ -2553,7 +2547,7 @@ x_7 = lean_ctor_get(x_5, 1); lean_inc(x_7); lean_dec(x_5); x_8 = lean_box(0); -x_9 = l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(x_6, x_8); +x_9 = l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(x_6, x_8); x_10 = l_List_isEmpty___rarg(x_9); if (x_10 == 0) { @@ -3401,43 +3395,6 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_elabPrint___closed__15() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Init.Data.Option.BasicAux", 25); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Command_elabPrint___closed__16() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Option.get!", 11); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Command_elabPrint___closed__17() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("value is none", 13); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Command_elabPrint___closed__18() { -_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_Elab_Command_elabPrint___closed__15; -x_2 = l_Lean_Elab_Command_elabPrint___closed__16; -x_3 = lean_unsigned_to_nat(16u); -x_4 = lean_unsigned_to_nat(14u); -x_5 = l_Lean_Elab_Command_elabPrint___closed__17; -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_Elab_Command_elabPrint(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { @@ -3483,95 +3440,71 @@ return x_18; } else { -lean_object* x_19; -x_19 = l_Lean_Syntax_isStrLit_x3f(x_12); -lean_dec(x_12); -if (lean_obj_tag(x_19) == 0) -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; -x_20 = l_Lean_Elab_Command_elabPrint___closed__18; -x_21 = l_panic___at_Lean_Name_getString_x21___spec__1(x_20); -x_22 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_22, 0, x_21); -x_23 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_23, 0, x_22); -x_24 = 0; -x_25 = l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2(x_10, x_23, x_24, x_2, x_3, x_4); +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; +x_19 = l_Lean_TSyntax_getString(x_12); +x_20 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_20, 0, x_19); +x_21 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_21, 0, x_20); +x_22 = 0; +x_23 = l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2(x_10, x_21, x_22, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_10); -return x_25; +return x_23; +} } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; -x_26 = lean_ctor_get(x_19, 0); +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_24 = l_Lean_Elab_Command_getRef(x_2, x_3, x_4); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); lean_inc(x_26); -lean_dec(x_19); -x_27 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_27, 0, x_26); -x_28 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_28, 0, x_27); -x_29 = 0; -x_30 = l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2(x_10, x_28, x_29, x_2, x_3, x_4); -lean_dec(x_3); -lean_dec(x_2); +lean_dec(x_24); +x_27 = l_Lean_replaceRef(x_10, x_25); +lean_dec(x_25); lean_dec(x_10); +x_28 = !lean_is_exclusive(x_2); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_2, 6); +lean_dec(x_29); +lean_ctor_set(x_2, 6, x_27); +x_30 = l___private_Lean_Elab_Print_0__Lean_Elab_Command_printId(x_12, x_2, x_3, x_26); return x_30; } -} -} else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; -x_31 = l_Lean_Elab_Command_getRef(x_2, x_3, x_4); -x_32 = lean_ctor_get(x_31, 0); -lean_inc(x_32); -x_33 = lean_ctor_get(x_31, 1); +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; +x_31 = lean_ctor_get(x_2, 0); +x_32 = lean_ctor_get(x_2, 1); +x_33 = lean_ctor_get(x_2, 2); +x_34 = lean_ctor_get(x_2, 3); +x_35 = lean_ctor_get(x_2, 4); +x_36 = lean_ctor_get(x_2, 5); +x_37 = lean_ctor_get(x_2, 7); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); lean_inc(x_33); -lean_dec(x_31); -x_34 = l_Lean_replaceRef(x_10, x_32); -lean_dec(x_32); -lean_dec(x_10); -x_35 = !lean_is_exclusive(x_2); -if (x_35 == 0) -{ -lean_object* x_36; lean_object* x_37; -x_36 = lean_ctor_get(x_2, 6); -lean_dec(x_36); -lean_ctor_set(x_2, 6, x_34); -x_37 = l___private_Lean_Elab_Print_0__Lean_Elab_Command_printId(x_12, x_2, x_3, x_33); -return x_37; -} -else -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_38 = lean_ctor_get(x_2, 0); -x_39 = lean_ctor_get(x_2, 1); -x_40 = lean_ctor_get(x_2, 2); -x_41 = lean_ctor_get(x_2, 3); -x_42 = lean_ctor_get(x_2, 4); -x_43 = lean_ctor_get(x_2, 5); -x_44 = lean_ctor_get(x_2, 7); -lean_inc(x_44); -lean_inc(x_43); -lean_inc(x_42); -lean_inc(x_41); -lean_inc(x_40); -lean_inc(x_39); -lean_inc(x_38); +lean_inc(x_32); +lean_inc(x_31); lean_dec(x_2); -x_45 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_45, 0, x_38); -lean_ctor_set(x_45, 1, x_39); -lean_ctor_set(x_45, 2, x_40); -lean_ctor_set(x_45, 3, x_41); -lean_ctor_set(x_45, 4, x_42); -lean_ctor_set(x_45, 5, x_43); -lean_ctor_set(x_45, 6, x_34); -lean_ctor_set(x_45, 7, x_44); -x_46 = l___private_Lean_Elab_Print_0__Lean_Elab_Command_printId(x_12, x_45, x_3, x_33); -return x_46; +x_38 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_38, 0, x_31); +lean_ctor_set(x_38, 1, x_32); +lean_ctor_set(x_38, 2, x_33); +lean_ctor_set(x_38, 3, x_34); +lean_ctor_set(x_38, 4, x_35); +lean_ctor_set(x_38, 5, x_36); +lean_ctor_set(x_38, 6, x_27); +lean_ctor_set(x_38, 7, x_37); +x_39 = l___private_Lean_Elab_Print_0__Lean_Elab_Command_printId(x_12, x_38, x_3, x_26); +return x_39; } } } @@ -5557,14 +5490,6 @@ l_Lean_Elab_Command_elabPrint___closed__13 = _init_l_Lean_Elab_Command_elabPrint lean_mark_persistent(l_Lean_Elab_Command_elabPrint___closed__13); l_Lean_Elab_Command_elabPrint___closed__14 = _init_l_Lean_Elab_Command_elabPrint___closed__14(); lean_mark_persistent(l_Lean_Elab_Command_elabPrint___closed__14); -l_Lean_Elab_Command_elabPrint___closed__15 = _init_l_Lean_Elab_Command_elabPrint___closed__15(); -lean_mark_persistent(l_Lean_Elab_Command_elabPrint___closed__15); -l_Lean_Elab_Command_elabPrint___closed__16 = _init_l_Lean_Elab_Command_elabPrint___closed__16(); -lean_mark_persistent(l_Lean_Elab_Command_elabPrint___closed__16); -l_Lean_Elab_Command_elabPrint___closed__17 = _init_l_Lean_Elab_Command_elabPrint___closed__17(); -lean_mark_persistent(l_Lean_Elab_Command_elabPrint___closed__17); -l_Lean_Elab_Command_elabPrint___closed__18 = _init_l_Lean_Elab_Command_elabPrint___closed__18(); -lean_mark_persistent(l_Lean_Elab_Command_elabPrint___closed__18); l___regBuiltin_Lean_Elab_Command_elabPrint___closed__1 = _init_l___regBuiltin_Lean_Elab_Command_elabPrint___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Elab_Command_elabPrint___closed__1); l___regBuiltin_Lean_Elab_Command_elabPrint___closed__2 = _init_l___regBuiltin_Lean_Elab_Command_elabPrint___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Quotation.c b/stage0/stdlib/Lean/Elab/Quotation.c index 3184ca89380..1cb958c6108 100644 --- a/stage0/stdlib/Lean/Elab/Quotation.c +++ b/stage0/stdlib/Lean/Elab/Quotation.c @@ -13,45 +13,43 @@ #ifdef __cplusplus extern "C" { #endif +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928____closed__2; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__1; lean_object* l_List_reverse___rarg(lean_object*); uint8_t l_Lean_Syntax_isQuot(lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663__declRange___closed__3; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934____closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__18___boxed(lean_object**); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__44; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__26; +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__61; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__50; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__9; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__20; lean_object* l_Lean_mkCIdentFrom(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__18; +static lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__3; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__2; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__1; +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__19(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__24; LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__1___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__58; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__8; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__8; lean_object* l_Lean_Syntax_unescapeAntiquot(lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__12; static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__13; uint8_t l_Lean_Syntax_isAntiquotSuffixSplice(lean_object*); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__31; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681____closed__4; -lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__38; lean_object* l_Lean_extractMacroScopes(lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__17; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__4; -static lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__1; -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__51; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940__declRange___closed__3; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__3; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928____closed__4; static lean_object* l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__5; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__9; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__6; size_t lean_usize_add(size_t, size_t); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepFromSplice___closed__3; static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__3; LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at_Lean_Elab_Term_Quotation_getQuotKind___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__2; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__21; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__26; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__45; static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__29; lean_object* l_List_tail_x21___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -60,56 +58,55 @@ uint8_t l_Lean_Syntax_isTokenAntiquot(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__6(lean_object*); lean_object* l_Lean_registerTraceClass(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax___closed__2; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633__declRange___closed__1; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__48; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__17; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__9; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__25; static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__14; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__17; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__57; static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__28; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952__declRange___closed__1; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__2; lean_object* l_Lean_stringToMessageData(lean_object*); +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__46; static lean_object* l_List_format___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__9___closed__9; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__7; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__4; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__18(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__58; +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__18(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__26; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__40; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__1; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916____closed__2; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__27; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__6; lean_object* lean_mk_empty_array_with_capacity(lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__1; -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8(lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__2; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__68; +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__22(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__4; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___boxed(lean_object**); -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__7; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__35; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__27; lean_object* l_Lean_Syntax_getHeadInfo_x3f(lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__5; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__65; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__40; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__1; static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__30; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__67; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__75; static lean_object* l_List_format___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__9___closed__4; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__1; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__39; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892____closed__4; static lean_object* l_Lean_Elab_Term_Quotation_ArrayStxBuilder_empty___closed__1; -extern lean_object* l_Lean_nullKind; +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946__declRange(lean_object*); +lean_object* l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__1; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916____closed__4; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__1; static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__15; +static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__6; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_List_beq___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__5(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__13; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_Quotation_getQuotKind___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904____closed__1; uint8_t l_Lean_isLitKind(lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__26; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__17; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651____closed__1; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__44; lean_object* lean_name_mk_string(lean_object*, lean_object*); lean_object* l_Array_eraseIdx___rarg(lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); @@ -119,616 +116,641 @@ static lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_ static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__6; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__16; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__44; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669__declRange___closed__4; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___lambda__1___closed__5; +static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__4(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__8; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_append___rarg(lean_object*, lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651____closed__2; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__24; +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__9; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__1; lean_object* l_List_mapTRAux___at_Lean_resolveGlobalConstCore___spec__2(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__28(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__28(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__18; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651__declRange___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699__declRange___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_resolveSectionVariable___boxed(lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__4; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__2; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__34; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__31; -static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__6; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__17; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675__declRange___closed__1; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__36; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952__declRange___closed__2; +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__20___boxed(lean_object*, lean_object*, lean_object*, 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_Term_Quotation_match__syntax_expand___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__6; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892__declRange___closed__3; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904____closed__3; lean_object* l_Lean_Syntax_antiquotSuffixSplice_x3f(lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__63; LEAN_EXPORT lean_object* l_List_format___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__9(lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__11___boxed(lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886____closed__4; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__3___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__7; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__71; lean_object* l_Lean_SourceInfo_fromRef(lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910____closed__3; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__14; static lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__10; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__37; -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__9; -static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__6___rarg___closed__1; +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952__declRange(lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904__declRange___closed__4; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_noOpMatchAdaptPats___closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946__declRange___closed__1; lean_object* l_List_filterMap___at_Lean_resolveGlobalConst___spec__1(lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__27; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__64; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__4; lean_object* l_List_unzip___rarg(lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__13; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__1; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639____closed__1; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__11(lean_object*); +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__10; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__32; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922____closed__2; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__8; -LEAN_EXPORT lean_object* l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__8(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__8(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__7; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__19___boxed(lean_object**); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645____closed__1; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904__declRange___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__25; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__1(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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940__declRange(lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__5; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__9; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651__declRange___closed__2; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__24; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__2; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_ArrayStxBuilder_append(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_structEq(lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__6; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645____closed__2; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639____closed__2; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__5; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__51; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922____closed__4; static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__16; size_t lean_usize_sub(size_t, size_t); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__30; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__14; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__11; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904__declRange___closed__3; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__4; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__28; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898____closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__4; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__40; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_Quotation_getQuotKind___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__9; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886____closed__3; static lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__7___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___lambda__3___closed__1; +static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__1; lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__20; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916____closed__3; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__4; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__10; +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__2; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__6; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__42; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__3___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__9; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886____closed__1; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__25; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__26; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__53; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__15; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_match__syntax_expand___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___closed__6; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__5(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__27; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892____closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__7; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__24; +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__10(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__7; lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Level_PP_Result_quote___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17___closed__2; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___lambda__1___closed__3; LEAN_EXPORT lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__7___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__32; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__5; lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__12; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_noOpMatchAdaptPats___lambda__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__2; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__11; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__18; +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898____closed__1; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904__declRange___closed__1; +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__8; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__39; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__47; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__76; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669__declRange___closed__3; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__29; lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__51; lean_object* lean_array_get_size(lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__23; +static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__7; +static lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__1; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928____closed__3; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892____closed__3; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874__declRange___closed__4; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898____closed__3; lean_object* lean_string_append(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__48; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__23; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__61; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__10; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___boxed(lean_object*, lean_object*, lean_object*, 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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__19; lean_object* l_List_range(lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__3; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__20; lean_object* l_Lean_MessageData_ofList(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__7___boxed(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__22; +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__15; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__34; lean_object* lean_string_utf8_extract(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__5; +static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__14; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699__declRange___closed__3; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681____closed__1; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__14(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_Syntax_getAntiquotTerm(lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__2; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687__declRange___closed__3; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__43; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922____closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__6___rarg(lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__19; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681____closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___lambda__2___boxed(lean_object*, 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_Array_mapMUnsafe_map___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__3(size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__44; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892____closed__5; static lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__6; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__1; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946__declRange___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax___closed__1; +static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__7; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__31; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__3; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__11; lean_object* l_Lean_throwError___at_Lean_Elab_Term_expandDeclId___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__18; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904____closed__4; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__32; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__10; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__8; +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__28; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___lambda__3___boxed(lean_object*, lean_object*); lean_object* lean_string_utf8_byte_size(lean_object*); static lean_object* l_panic___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__1___closed__1; lean_object* l_List_head_x21___rarg(lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__22; -static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__10; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__59; lean_object* l_Lean_Elab_Term_Quotation_getPatternsVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_mkTuple(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__5; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886____closed__5; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__13; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__20; lean_object* l_Lean_Elab_Term_instInhabitedTermElabM(lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__37; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880____closed__1; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952__declRange___closed__4; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__14; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651__declRange___closed__4; +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__4; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__16; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__18; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__50; +static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__13; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__3; lean_object* l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__41; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__4; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874__declRange___closed__1; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934__declRange___closed__3; uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__55; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__12; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__73; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_match__syntax_expand(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_nameLitKind; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__18; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__12___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__9; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__12; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__15; static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax___closed__3; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880____closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17___boxed(lean_object**); -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__30; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__9; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__5; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__11; -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__27; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__13; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663_(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_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669_(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_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657_(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_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633_(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_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711_(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_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705_(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_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675_(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_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693_(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_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681_(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_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699_(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_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687_(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_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645_(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_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651_(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_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639_(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6; +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916_(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_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904_(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_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910_(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_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898_(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_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922_(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_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886_(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_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874_(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_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892_(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_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880_(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_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934_(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_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940_(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_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928_(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_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952_(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_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946_(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__2; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__63; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___lambda__1___closed__1; lean_object* lean_nat_add(lean_object*, lean_object*); static lean_object* l_List_format___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__9___closed__2; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__24; static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__21; static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__3; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687____closed__4; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__47; static lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Term_Quotation_getQuotKind___spec__5___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__9; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__21; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699____closed__4; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__29; -lean_object* l_Lean_Syntax_SepArray_ofElems(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__24; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__18; static lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__20___boxed(lean_object**); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663____closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__21(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__21(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__21(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__18; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__5; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633__declRange___closed__3; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880__declRange___closed__4; +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__8; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__25; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__23; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__58; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898__declRange___closed__3; static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__23; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__7; lean_object* l_Array_zip___rarg(lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__4; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__66; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__2; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__30; +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__20(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__13; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__15; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__56; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__77; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__22; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657____closed__1; -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__12; +static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__11; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__49; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__16; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__16; +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__42; LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675____closed__1; static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__5; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__7; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__46; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_match__syntax_expand___lambda__2(lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669____closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__9; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__12(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__18; -static lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___closed__1; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__12; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663____closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__25; lean_object* l_Lean_throwError___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880__declRange___closed__1; uint8_t l_Lean_Syntax_isAntiquots(lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__13; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__9___boxed(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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5(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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__9___boxed(lean_object*, 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_Term_Quotation_commandElab__stx__quot__; lean_object* l_Lean_Core_withFreshMacroScope___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669____closed__2; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934__declRange___closed__2; +static lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__1; static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__15; LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Term_Quotation_getQuotKind___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__10___boxed(lean_object**); -static lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___closed__2; +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__49; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___lambda__1___closed__4; lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(size_t, size_t, lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675____closed__3; -LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___boxed(lean_object*, lean_object*, 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934__declRange___closed__4; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__12; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__31; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__2___closed__2; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__59; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__8(lean_object*, size_t, lean_object*, size_t, size_t); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__3; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711____closed__1; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25(lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax_docString___closed__1; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__4; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__55; static lean_object* l___private_Init_Meta_0__Lean_quoteList___at_Lean_Elab_Term_Quotation_ArrayStxBuilder_build___spec__1___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_isLit_x3f(lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__46; +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__18(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__1; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__19; +static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__9; +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__29; lean_object* lean_array_fget(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__29; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__47; uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__2; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__25; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__74; -static lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___closed__3; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946__declRange___closed__3; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__11___boxed(lean_object**); -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__21; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__37; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__8; -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__29; +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__52; +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__8; -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__10; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___boxed(lean_object**); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898__declRange___closed__2; +static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__4; lean_object* l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__7; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711____closed__3; +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898__declRange___closed__1; static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__14; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__61; -LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, 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_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__6; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__8; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693____closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__62; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__11; -extern lean_object* l_Lean_numLitKind; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__9; static lean_object* l_Lean_Elab_Term_Quotation_ArrayStxBuilder_push___closed__4; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__54; lean_object* l_ReaderT_bind___at_Lean_Unhygienic_instMonadQuotationUnhygienic___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__6; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__5; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__7; +static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__9; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__45; +static lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; lean_object* lean_nat_sub(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__8; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669____closed__3; +LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__8; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__2___closed__4; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__2; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__59; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__31; +static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__8; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__30; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_Quotation_getQuotKind___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1(lean_object*, lean_object*, lean_object*); +uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__3___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__9; -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__6; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__4; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___rarg(lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__17; -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__9; lean_object* l_Lean_Syntax_mkCApp(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__8; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__17; static lean_object* l_List_format___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__9___closed__1; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__8; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__43; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657____closed__3; +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__26(lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__17; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__28; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax_declRange(lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__30; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699____closed__3; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__14; LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at_Lean_Elab_Term_Quotation_getQuotKind___spec__4___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__17; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax_docString(lean_object*); -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__28; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__5; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__34; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__76; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__30; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__12; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645__declRange___closed__3; -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__2; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__33; -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__3; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__22; static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__8; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898__declRange___closed__4; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__47; static lean_object* l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__4; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__15; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__18; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910__declRange___closed__4; +static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__12; +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_ArrayStxBuilder_push(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__17; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__3; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__19; -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__28___closed__1; +static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__2___closed__6; static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__25; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687____closed__3; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__38; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__21; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__8; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__38; static lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__7; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699__declRange___closed__2; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__46; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__15; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_log___at_Lean_Elab_Term_traceAtCmdPos___spec__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__2; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__27; static lean_object* l_List_format___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__9___closed__6; LEAN_EXPORT lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___closed__7; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__28; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__37; lean_object* l_Lean_replaceRef(lean_object*, lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699____closed__1; lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Unhygienic_run___rarg(lean_object*); -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__15; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__28___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639__declRange(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__28___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Meta_0__Lean_quoteList___at_Lean_Elab_Term_Quotation_ArrayStxBuilder_build___spec__1___closed__2; -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645__declRange(lean_object*); -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__3; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__4; +lean_object* l_panic___at_Lean_TSyntax_getString___spec__1(lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__42; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__53; static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__13; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__21; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__6; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__8; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__23; static lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Term_Quotation_getQuotKind___spec__5___closed__2; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699__declRange___closed__4; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681__declRange___closed__1; static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__9; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9(lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934__declRange(lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__38; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__6; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__2___closed__2; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940__declRange___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__21; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__78; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__14; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__10; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687____closed__1; -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__28; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705____closed__1; -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__5; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940__declRange___closed__2; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892__declRange___closed__2; +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__27(lean_object*); static lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__4___closed__1; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__1; -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651__declRange(lean_object*); -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__2(lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__60; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922____closed__1; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__41; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_match__syntax_expand___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__10(lean_object*); static lean_object* l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__3; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__48; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__2___closed__1; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__56; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687____closed__2; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__66; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__50; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__13; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_Quotation_getQuotKind___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928__declRange___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__1; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__3; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__12___boxed(lean_object**); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__3; -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__8; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705____closed__3; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940__declRange___closed__4; static lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__3; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__48; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__42; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__41; static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__2; -static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__11; lean_object* l_Nat_repr(lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__79; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__15; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__12; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__4; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__3; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663____closed__4; -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633__declRange(lean_object*); -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__20; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693____closed__1; +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__15(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__33; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__34; -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__28___closed__2; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__15; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__53; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__7; -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__26; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__22___boxed(lean_object**); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675____closed__4; -static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__12; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__10; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940____closed__2; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__55; +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__27___boxed(lean_object*); lean_object* l_Lean_Syntax_getId(lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__31; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__5; lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___lambda__1___closed__6; -lean_object* l___private_Init_Meta_0__Lean_quoteNameMk(lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__10; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910__declRange___closed__1; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928____closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__31; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__13; -extern lean_object* l_Lean_choiceKind; static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__6; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__6; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__21; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__52; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693____closed__2; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__20; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__10; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__32; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__10; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__28; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_match__syntax_expand___lambda__5(lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax(lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__22; LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__37; +static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__2; lean_object* l_List_replicateTR___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__14; lean_object* l_Array_back___rarg(lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928__declRange___closed__3; static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__26; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__3___lambda__1___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__1; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__9; lean_object* l_Lean_addBuiltinDeclarationRanges(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__23; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__49; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639__declRange___closed__3; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880__declRange___closed__3; lean_object* l_Function_comp___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; lean_object* lean_array_to_list(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__7; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Name_isAtomic(lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__4; static lean_object* l___private_Init_Meta_0__Lean_quoteList___at_Lean_Elab_Term_Quotation_ArrayStxBuilder_build___spec__1___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax_declRange___closed__4; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_mkTuple___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__2___closed__4; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_noOpMatchAdaptPats___lambda__1___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__2; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693____closed__3; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928__declRange___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax_declRange___closed__1; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__10; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__57; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__63; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__4; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910__declRange___closed__2; +static lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__10; -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__8; lean_object* l_Lean_Syntax_mkStrLit(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__4; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__2; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__62; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__15; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__56; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__66; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886__declRange___closed__2; static lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__1; -static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__13; +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916__declRange(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_noOpMatchAdaptPats(lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_Quotation_getAntiquotationIds(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__1(lean_object*, 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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__22; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__17; static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__19; -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__4; -LEAN_EXPORT lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__29; +LEAN_EXPORT lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__29; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__1; -static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__1; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__21; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__9; lean_object* l_Std_RBNode_find___at___private_Lean_Hygiene_0__Lean_sanitizeSyntaxAux___spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880__declRange(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__25; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681__declRange___closed__4; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_13012_(lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946____closed__2; +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_13744_(lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910__declRange___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__24; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__20; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__3; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711__declRange___closed__4; lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__8; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693__declRange___closed__4; +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__17(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__54; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__2(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__12; uint8_t l_Array_isEmpty___rarg(lean_object*); -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__9; +static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__2___closed__5; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940____closed__3; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__43; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__24; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__52; +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928__declRange(lean_object*); static lean_object* l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__11___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__12___closed__2; -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__32; extern lean_object* l_Lean_instInhabitedSyntax; static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__11; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__10; static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax_declRange___closed__5; static lean_object* l_Lean_resolveGlobalConst___at_Lean_Elab_Term_Quotation_getQuotKind___spec__2___closed__1; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__16; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__27; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__7; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928__declRange___closed__2; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__7; lean_object* l_Lean_mkSepArray(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__49; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__30(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922__declRange___closed__4; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__3; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__37; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663__declRange___closed__1; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__49; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__55; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_getQuotKind(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__11___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__12___closed__1; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__10; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__42; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__53; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__30___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__32; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__73; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__26; +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__5; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__25; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892__declRange___closed__4; lean_object* l_String_dropRight(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681__declRange(lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__16; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__66; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__5; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__2___closed__5; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__56; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__20; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699____closed__2; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681__declRange___closed__2; -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__7; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940____closed__4; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__40; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952____closed__2; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__4; size_t lean_usize_of_nat(lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693__declRange___closed__2; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946__declRange___closed__2; +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922__declRange(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_ArrayStxBuilder_empty; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711__declRange___closed__2; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__11; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886__declRange___closed__4; static lean_object* l___private_Init_Meta_0__Lean_quoteList___at_Lean_Elab_Term_Quotation_ArrayStxBuilder_build___spec__1___closed__6; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__16; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__3; @@ -737,534 +759,514 @@ uint8_t l_List_foldr___at_List_and___spec__1(uint8_t, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___closed__2; lean_object* l_Lean_addTrace___at___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTerm___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__15(lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__20; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__5___boxed(lean_object*, 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_resolveGlobalConst___at_Lean_Elab_Term_Quotation_getQuotKind___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__36; lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__45; static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax_declRange___closed__7; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__77; -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__29; +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__67; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__22; LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__1___boxed(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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__12; +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952____closed__1; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__64; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__50; +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Term_Quotation_getQuotKind___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__2(lean_object*, lean_object*, 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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isAtom(lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__35; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651__declRange___closed__3; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__13; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__11; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711__declRange___closed__1; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__11; +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892__declRange(lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__7; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__3; extern lean_object* l_Lean_Elab_Term_termElabAttribute; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__1; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705____closed__2; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__41; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_noOpMatchAdaptPats___lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__2; static lean_object* l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__6; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__5; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874__declRange___closed__2; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__5; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__18; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__1___boxed(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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886__declRange(lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892__declRange___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__19(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__19(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__60; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__10; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__11; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__30; -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__7; +lean_object* l_Lean_Syntax_mkNumLit(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_beq___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__5___boxed(lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__11; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__29; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__25; -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__30; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946____closed__1; +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24(lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__45; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__5; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__3; -lean_object* l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepFromSplice(lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__24; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__18; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669__declRange___closed__1; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__14; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getQuotContent(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__5(size_t, size_t, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__65; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__39; +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__14___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__3; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__32; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__3; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__10; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__41; +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910__declRange(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__7(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__7(lean_object*, lean_object*, size_t, size_t); -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__13; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952__declRange___closed__3; +static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__7; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__16; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__21; LEAN_EXPORT lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_String_intercalate(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__28; +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904__declRange(lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__17; lean_object* l_List_redLength___rarg(lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__15; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__5; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__12; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711__declRange___closed__3; +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16___boxed(lean_object**); +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__30; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__15; LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at_Lean_Elab_Term_Quotation_getQuotKind___spec__4___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__6; +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898__declRange(lean_object*); static lean_object* l_List_format___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__9___closed__7; -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__9; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__19; -static lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__1; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__69; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__2; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__49; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711____closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__20(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__20(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922__declRange___closed__1; lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOptionM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTerm___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__48; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__2___closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__29(lean_object*, lean_object*, 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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__29(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__34; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__23; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__7; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__5; lean_object* l_Lean_Syntax_getNumArgs(lean_object*); -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__23; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705__declRange___closed__3; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__4(lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651_(lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639_(lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633_(lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645_(lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681_(lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675_(lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663_(lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657_(lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669_(lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711_(lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705_(lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699_(lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687_(lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693_(lean_object*); -static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__4; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__72; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__74; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__23; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__5; +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940_(lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928_(lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922_(lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934_(lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916_(lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952_(lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946_(lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904_(lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898_(lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910_(lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892_(lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880_(lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886_(lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874_(lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__5; lean_object* l_Lean_Syntax_setArg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__18; lean_object* lean_environment_main_module(lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__14; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___lambda__3(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__18; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__21; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__62; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__54; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__22; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__20; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__30; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__3; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__3; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__2; static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__9; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23(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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886__declRange___closed__3; static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__27; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__38; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__12; -static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__8; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__52; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880__declRange___closed__2; uint8_t lean_nat_dec_le(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__15; +static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__10; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___boxed(lean_object**); static lean_object* l_Lean_Elab_Term_Quotation_match__syntax_expand___lambda__3___closed__1; static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__4; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__14___boxed(lean_object**); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__42; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_match__syntax_expand___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681__declRange___closed__3; -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__3; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687__declRange___closed__1; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__36; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940____closed__1; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922__declRange___closed__2; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__28; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__2; static lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__9; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__1; lean_object* l_Lean_Syntax_getArgs(lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__10; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2(lean_object*, lean_object*, lean_object*, 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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874__declRange(lean_object*); LEAN_EXPORT lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__4___lambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__5; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__58; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_match__syntax_expand___lambda__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__7(lean_object*); lean_object* l_Lean_Syntax_getAntiquotSpliceSuffix(lean_object*); lean_object* l_Lean_Syntax_getKind(lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16___closed__4; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__12(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__32; +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__12(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); lean_object* l_Lean_MacroScopesView_review(lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__20; -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__3; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__11; -static lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__4; +lean_object* l_Lean_quoteNameMk(lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_ArrayStxBuilder_build___closed__1; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__2; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657____closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__33; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__14; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__7(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__10; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__60; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__2___closed__6; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__47; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__19; -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__50; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663____closed__2; static lean_object* l___private_Init_Meta_0__Lean_quoteList___at_Lean_Elab_Term_Quotation_ArrayStxBuilder_build___spec__1___closed__3; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__70; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__33; static lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__8; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__22; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675____closed__2; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__8; LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_match__syntax_expand___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633__declRange___closed__4; +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__11(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886__declRange___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___closed__4; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___closed__1; static lean_object* l_Lean_resolveGlobalConst___at_Lean_Elab_Term_Quotation_getQuotKind___spec__2___closed__3; LEAN_EXPORT lean_object* l_Array_sequenceMap___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___spec__1___boxed(lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__5; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__3(size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__3(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_ArrayStxBuilder_build___closed__2; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__46; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__24; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__11; lean_object* l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__2; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__12; +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__4; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__35; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__6; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__7; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__2; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16___closed__2; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__11; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__7; -uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__14; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__19; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__1; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__64; static lean_object* l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__2; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__12; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__2___closed__3; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__75; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__26; lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934__declRange___closed__1; +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__14___closed__2; lean_object* l_Lean_Elab_Term_adaptExpander(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__19; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__28; -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__16; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___closed__3; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__20; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__45; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__45; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__27; +LEAN_EXPORT lean_object* l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__8___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__2; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__4; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__17; +lean_object* l_Lean_Syntax_mkNameLit(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax_declRange___closed__2; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651____closed__3; LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__11(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__19; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__9; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__29; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__3; lean_object* l_Lean_resolveGlobalName___at_Lean_Elab_Term_resolveName___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Meta_0__Lean_quoteList___at_Lean_Elab_Term_Quotation_ArrayStxBuilder_build___spec__1___closed__7; +static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__25; lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__8; -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__5; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687__declRange___closed__2; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__4; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__9; LEAN_EXPORT lean_object* l_panic___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__5; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___boxed(lean_object**); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645____closed__5; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681____closed__2; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__52; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__44; lean_object* l_List_toArrayAux___rarg(lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__1; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675__declRange___closed__4; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657__declRange___closed__2; +static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__14; LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isNone(lean_object*); lean_object* lean_name_mk_numeral(lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__31; -static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__3; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__5; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__22; +static lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__4; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__27; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___boxed(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_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639____closed__3; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__40; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__9; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874__declRange___closed__3; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__5; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__6; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at_Lean_Elab_Term_Quotation_ArrayStxBuilder_build___spec__1(lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__78; -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__6; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__6; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__16(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__46; -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__6; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705__declRange___closed__1; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__57; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___closed__5; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__12; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__6; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651____closed__5; -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693__declRange(lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_ArrayStxBuilder_push___closed__2; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880____closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_resolveSectionVariable_loop___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__34; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__4; lean_object* l_Lean_Syntax_getAntiquotSpliceContents(lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__25; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__6; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705__declRange___closed__4; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952____closed__3; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__17; +static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__7; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__57; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__8; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__36; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645__declRange___closed__4; lean_object* lean_mk_array(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__48; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__26; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645____closed__3; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__4; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__3; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__24; -LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__11; +LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabMatchSyntax(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__16; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__26; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687__declRange___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot____; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__3; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916__declRange___closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__6; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepFromSplice___closed__2; +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__12___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax_declRange___closed__3; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675__declRange___closed__2; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886____closed__2; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__23; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663__declRange___closed__2; +static lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__2; static lean_object* l_List_format___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__9___closed__3; +static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__3; LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_adaptRhs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__22(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__22(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__17; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16___closed__5; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__7; -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705__declRange(lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__35; -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687__declRange(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax_declRange___closed__6; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916__declRange___closed__2; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__10; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__61; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__32; -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699__declRange(lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__7; lean_object* lean_string_length(lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__33; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__8; extern lean_object* l_Lean_Elab_unsupportedSyntaxExceptionId; -static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__15; static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__11; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16___closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___boxed(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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__39; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657__declRange___closed__1; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_resolveSectionVariable_loop(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isEscapedAntiquot(lean_object*); -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__50; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705__declRange___closed__2; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946____closed__3; +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__4; lean_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__5___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_pure___at_Lean_Elab_Term_addTermInfo___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___lambda__1___closed__2; lean_object* l_Lean_addBuiltinDocString(lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__62; -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__20; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__12; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepFromSplice___boxed(lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__16; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__54; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934____closed__1; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__22; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__19; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__24; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__21; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__38; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__2; -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__2___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__6; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepFromSplice___closed__1; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__3; +static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__9; +static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___rarg___closed__1; +static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__8; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__13; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__41; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__33; LEAN_EXPORT lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__4___lambda__1___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__72; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639__declRange___closed__4; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__29___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__13; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669__declRange___closed__2; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__36; +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__29___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892____closed__2; static lean_object* l___private_Init_Meta_0__Lean_quoteList___at_Lean_Elab_Term_Quotation_ArrayStxBuilder_build___spec__1___closed__5; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__23; LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___rarg___closed__2; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__21; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693__declRange___closed__3; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__11; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__10; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__19; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__70; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__23; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916__declRange___closed__1; +static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__6; lean_object* l_Array_findIdx_x3f_loop___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__13; -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__25; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__31; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_match__syntax_expand___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_antiquotSpliceKind_x3f(lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__14; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14(lean_object*, lean_object*, 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_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15(lean_object*, lean_object*, lean_object*, lean_object*, 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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14(lean_object*, lean_object*, lean_object*, lean_object*, 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_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__28; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675__declRange___closed__3; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__69; static lean_object* l_Lean_Elab_Term_Quotation_match__syntax_expand___lambda__5___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__27; static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__7; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934____closed__4; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__15; uint8_t l_List_isEmpty___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__13___boxed(lean_object**); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__6(lean_object*, lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904____closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__43; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_antiquotKinds(lean_object*); -lean_object* l_panic___at_Lean_Name_getString_x21___spec__1(lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__26; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__8; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__16; LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__10(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__7; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__4; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_sequenceMap___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___spec__1(lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__35; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__20; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__65; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__1; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663__declRange___closed__4; static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__12; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepFromSplice___closed__4; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__3; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__79; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__51; +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__39; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__16; LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_format___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__9___closed__5; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__19; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__1(lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__43; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_noOpMatchAdaptPats___closed__1; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639__declRange___closed__1; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__1; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__51; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__30; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910____closed__2; +static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__6; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922__declRange___closed__3; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_noOpMatchAdaptPats___closed__2; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__11; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__22; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__18; extern lean_object* l_Lean_Elab_Term_Quotation_hygiene; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25(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___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711__declRange(lean_object*); -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__1; -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__1; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657__declRange___closed__4; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__44; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916__declRange___closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__6___closed__1; static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__12; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__2; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4(lean_object*, lean_object*, size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__11; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__2; static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__16; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__68; lean_object* l_Lean_Syntax_getCanonicalAntiquot(lean_object*); -static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__2; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__2; lean_object* l_Lean_mkConst(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__51; -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675__declRange(lean_object*); -static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__14; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__63; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__23; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__2; +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__22; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__47; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639__declRange___closed__2; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__20; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__17; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_ArrayStxBuilder_build(lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669__declRange(lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645____closed__4; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__9; -static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__5; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__60; -static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__6___rarg___closed__2; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645__declRange___closed__2; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934____closed__3; +static lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__1; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__8; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__65; lean_object* lean_nat_to_int(lean_object*); lean_object* l_List_appendTR___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__8; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__2___closed__7; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__8; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633__declRange___closed__2; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__12; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__7; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__39; static lean_object* l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__11___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__12___closed__3; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__13; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__64; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__4(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8(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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_ArrayStxBuilder_push___closed__1; static lean_object* l_List_format___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__9___closed__8; static lean_object* l_Lean_Elab_Term_Quotation_match__syntax_expand___lambda__5___closed__2; +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__3; lean_object* l_Lean_Syntax_formatStxAux(lean_object*, uint8_t, lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__71; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(lean_object*, size_t, size_t, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__13; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657__declRange___closed__3; +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__21___boxed(lean_object**); LEAN_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__10; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651____closed__4; -LEAN_EXPORT lean_object* l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__7(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__11; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__9; +static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__15; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__24; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__35; static lean_object* l_Lean_Elab_Term_Quotation_ArrayStxBuilder_push___closed__3; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__31; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__9; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645__declRange___closed__1; -LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__2; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__40; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693__declRange___closed__1; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__21; -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657__declRange(lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__43; -static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__27; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__59; +static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__3; +static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__36; +static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__6; +static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__5; +static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__5; lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); static lean_object* l_Lean_resolveGlobalConst___at_Lean_Elab_Term_Quotation_getQuotKind___spec__2___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__4___boxed(lean_object*, lean_object*, lean_object*); uint8_t lean_string_dec_eq(lean_object*, lean_object*); -lean_object* l_Lean_Syntax_mkLit(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__19; -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663__declRange(lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_match__syntax_expand___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__6; static lean_object* l_Lean_Elab_Term_Quotation_elabMatchSyntax___closed__1; uint8_t l_Lean_Syntax_isIdent(lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__9; -static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__16; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__10; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__29; uint8_t l_Lean_Syntax_isAntiquotSplice(lean_object*); @@ -1779,6 +1781,24 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__2___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("choice", 6); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__2___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__2___closed__5; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___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_object* x_9, lean_object* x_10) { _start: { @@ -1823,7 +1843,7 @@ lean_inc(x_27); x_28 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___boxed), 10, 2); lean_closure_set(x_28, 0, x_27); lean_closure_set(x_28, 1, x_6); -x_29 = l_Lean_choiceKind; +x_29 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__2___closed__6; lean_inc(x_7); x_30 = l_Lean_Syntax_isOfKind(x_7, x_29); if (x_30 == 0) @@ -1847,7 +1867,7 @@ lean_dec(x_34); x_36 = 0; x_37 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__4(x_27, x_35, x_36, x_33); x_38 = lean_box(2); -x_39 = l_Lean_nullKind; +x_39 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; x_40 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_40, 0, x_38); lean_ctor_set(x_40, 1, x_39); @@ -1883,7 +1903,7 @@ lean_inc(x_49); x_50 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___boxed), 10, 2); lean_closure_set(x_50, 0, x_49); lean_closure_set(x_50, 1, x_6); -x_51 = l_Lean_choiceKind; +x_51 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__2___closed__6; lean_inc(x_7); x_52 = l_Lean_Syntax_isOfKind(x_7, x_51); if (x_52 == 0) @@ -1909,7 +1929,7 @@ lean_dec(x_57); x_59 = 0; x_60 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__4(x_49, x_58, x_59, x_56); x_61 = lean_box(2); -x_62 = l_Lean_nullKind; +x_62 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; x_63 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_63, 0, x_61); lean_ctor_set(x_63, 1, x_62); @@ -1962,7 +1982,7 @@ lean_inc(x_77); x_78 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___boxed), 10, 2); lean_closure_set(x_78, 0, x_77); lean_closure_set(x_78, 1, x_6); -x_79 = l_Lean_choiceKind; +x_79 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__2___closed__6; lean_inc(x_7); x_80 = l_Lean_Syntax_isOfKind(x_7, x_79); if (x_80 == 0) @@ -1993,7 +2013,7 @@ lean_dec(x_86); x_88 = 0; x_89 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__4(x_77, x_87, x_88, x_85); x_90 = lean_box(2); -x_91 = l_Lean_nullKind; +x_91 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; x_92 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_92, 0, x_90); lean_ctor_set(x_92, 1, x_91); @@ -2436,7 +2456,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot 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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepFromSplice___closed__1; x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepFromSplice___closed__2; -x_3 = lean_unsigned_to_nat(44u); +x_3 = lean_unsigned_to_nat(45u); x_4 = lean_unsigned_to_nat(4u); x_5 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepFromSplice___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -2463,7 +2483,7 @@ else lean_object* x_6; lean_object* x_7; lean_dec(x_2); x_6 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepFromSplice___closed__4; -x_7 = l_panic___at_Lean_Name_getString_x21___spec__1(x_6); +x_7 = l_panic___at_Lean_TSyntax_getString___spec__1(x_6); return x_7; } } @@ -3580,7 +3600,166 @@ lean_ctor_set(x_9, 0, x_8); return x_9; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__1(size_t x_1, size_t 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) { +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___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_object* x_7, lean_object* x_8) { +_start: +{ +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; uint8_t x_16; +x_9 = lean_ctor_get(x_6, 5); +x_10 = lean_ctor_get(x_2, 1); +lean_inc(x_10); +lean_inc(x_10); +x_11 = l_Lean_Elab_getBetterRef(x_9, x_10); +x_12 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_4, x_5, x_6, x_7, x_8); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(x_13, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_14); +lean_dec(x_2); +lean_dec(x_10); +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_15, 0); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_11); +lean_ctor_set(x_18, 1, x_17); +lean_ctor_set_tag(x_15, 1); +lean_ctor_set(x_15, 0, x_18); +return x_15; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_19 = lean_ctor_get(x_15, 0); +x_20 = lean_ctor_get(x_15, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_15); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_11); +lean_ctor_set(x_21, 1, x_19); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_20); +return x_22; +} +} +} +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__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, lean_object* x_9) { +_start: +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_7); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_ctor_get(x_7, 5); +x_12 = l_Lean_replaceRef(x_1, x_11); +lean_dec(x_11); +lean_dec(x_1); +lean_ctor_set(x_7, 5, x_12); +x_13 = l_Lean_throwError___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__2(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_13; +} +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; 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; +x_14 = lean_ctor_get(x_7, 0); +x_15 = lean_ctor_get(x_7, 1); +x_16 = lean_ctor_get(x_7, 2); +x_17 = lean_ctor_get(x_7, 3); +x_18 = lean_ctor_get(x_7, 4); +x_19 = lean_ctor_get(x_7, 5); +x_20 = lean_ctor_get(x_7, 6); +x_21 = lean_ctor_get(x_7, 7); +x_22 = lean_ctor_get(x_7, 8); +x_23 = lean_ctor_get(x_7, 9); +x_24 = lean_ctor_get(x_7, 10); +lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_dec(x_7); +x_25 = l_Lean_replaceRef(x_1, x_19); +lean_dec(x_19); +lean_dec(x_1); +x_26 = lean_alloc_ctor(0, 11, 0); +lean_ctor_set(x_26, 0, x_14); +lean_ctor_set(x_26, 1, x_15); +lean_ctor_set(x_26, 2, x_16); +lean_ctor_set(x_26, 3, x_17); +lean_ctor_set(x_26, 4, x_18); +lean_ctor_set(x_26, 5, x_25); +lean_ctor_set(x_26, 6, x_20); +lean_ctor_set(x_26, 7, x_21); +lean_ctor_set(x_26, 8, x_22); +lean_ctor_set(x_26, 9, x_23); +lean_ctor_set(x_26, 10, x_24); +x_27 = l_Lean_throwError___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__2(x_2, x_3, x_4, x_5, x_6, x_26, x_8, x_9); +lean_dec(x_8); +lean_dec(x_26); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_27; +} +} +} +static lean_object* _init_l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Elab_unsupportedSyntaxExceptionId; +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___rarg___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___rarg___closed__1; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___rarg___closed__2; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3(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 = lean_alloc_closure((void*)(l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___rarg), 1, 0); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4(size_t x_1, size_t 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) { _start: { uint8_t x_11; @@ -3660,7 +3839,7 @@ return x_26; } } } -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -3697,7 +3876,7 @@ return x_8; } } } -static lean_object* _init_l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___closed__1() { +static lean_object* _init_l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__1() { _start: { lean_object* x_1; @@ -3705,22 +3884,22 @@ x_1 = lean_mk_string_from_bytes("Array.zip", 9); return x_1; } } -static lean_object* _init_l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___closed__2() { +static lean_object* _init_l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___closed__1; +x_1 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__1; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___closed__3() { +static lean_object* _init_l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___closed__1; +x_1 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___closed__2; +x_3 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__2; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3728,7 +3907,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___closed__4() { +static lean_object* _init_l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__4() { _start: { lean_object* x_1; @@ -3736,88 +3915,499 @@ x_1 = lean_mk_string_from_bytes("zip", 3); return x_1; } } -LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6(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_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_3, x_4); -if (x_13 == 0) +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_4, x_5); +if (x_14 == 0) { -size_t x_14; size_t 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; 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_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_14 = 1; -x_15 = lean_usize_sub(x_3, x_14); -x_16 = lean_array_uget(x_2, x_15); -lean_inc(x_10); -x_17 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_10, x_11, x_12); -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); +size_t x_15; size_t 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; 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_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_15 = 1; +x_16 = lean_usize_sub(x_4, x_15); +x_17 = lean_array_uget(x_3, x_16); +lean_inc(x_11); +x_18 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_11, x_12, x_13); +x_19 = lean_ctor_get(x_18, 0); lean_inc(x_19); -lean_dec(x_17); -x_20 = lean_ctor_get(x_10, 10); +x_20 = lean_ctor_get(x_18, 1); lean_inc(x_20); -x_21 = lean_st_ref_get(x_11, x_19); -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_21, 1); +lean_dec(x_18); +x_21 = lean_ctor_get(x_11, 10); +lean_inc(x_21); +x_22 = lean_st_ref_get(x_12, x_20); +x_23 = lean_ctor_get(x_22, 0); lean_inc(x_23); -lean_dec(x_21); -x_24 = lean_ctor_get(x_22, 0); +x_24 = lean_ctor_get(x_22, 1); lean_inc(x_24); lean_dec(x_22); -x_25 = lean_environment_main_module(x_24); -x_26 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___closed__4; +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_environment_main_module(x_25); +x_27 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__4; lean_inc(x_1); -x_27 = lean_name_mk_string(x_1, x_26); -lean_inc(x_27); -x_28 = l_Lean_addMacroScope(x_25, x_27, x_20); -x_29 = lean_box(0); -x_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_27); -lean_ctor_set(x_30, 1, x_29); -x_31 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_29); -x_32 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___closed__3; -x_33 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_33, 0, x_18); -lean_ctor_set(x_33, 1, x_32); -lean_ctor_set(x_33, 2, x_28); -lean_ctor_set(x_33, 3, x_31); -x_34 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_35 = lean_array_push(x_34, x_16); -x_36 = lean_array_push(x_35, x_5); -x_37 = lean_box(2); -x_38 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_28 = lean_name_mk_string(x_1, x_27); +lean_inc(x_28); +x_29 = l_Lean_addMacroScope(x_26, x_28, x_21); +x_30 = lean_box(0); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_28); +lean_ctor_set(x_31, 1, x_30); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +x_33 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__3; +x_34 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_34, 0, x_19); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_29); +lean_ctor_set(x_34, 3, x_32); +x_35 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_36 = lean_array_push(x_35, x_17); +x_37 = lean_array_push(x_36, x_6); +x_38 = lean_box(2); +lean_inc(x_2); x_39 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_39, 0, x_37); -lean_ctor_set(x_39, 1, x_38); -lean_ctor_set(x_39, 2, x_36); -x_40 = lean_array_push(x_34, x_33); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_2); +lean_ctor_set(x_39, 2, x_37); +x_40 = lean_array_push(x_35, x_34); x_41 = lean_array_push(x_40, x_39); x_42 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; x_43 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_43, 0, x_37); +lean_ctor_set(x_43, 0, x_38); lean_ctor_set(x_43, 1, x_42); lean_ctor_set(x_43, 2, x_41); -x_3 = x_15; -x_5 = x_43; -x_12 = x_23; +x_4 = x_16; +x_6 = x_43; +x_13 = x_24; goto _start; } else { lean_object* x_45; -lean_dec(x_10); +lean_dec(x_11); +lean_dec(x_2); +lean_dec(x_1); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_6); +lean_ctor_set(x_45, 1, x_13); +return x_45; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__7(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_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_4, x_5); +if (x_14 == 0) +{ +size_t x_15; size_t 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; 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_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_15 = 1; +x_16 = lean_usize_sub(x_4, x_15); +x_17 = lean_array_uget(x_3, x_16); +lean_inc(x_11); +x_18 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_11, x_12, x_13); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_ctor_get(x_11, 10); +lean_inc(x_21); +x_22 = lean_st_ref_get(x_12, x_20); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_environment_main_module(x_25); +x_27 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__4; +lean_inc(x_1); +x_28 = lean_name_mk_string(x_1, x_27); +lean_inc(x_28); +x_29 = l_Lean_addMacroScope(x_26, x_28, x_21); +x_30 = lean_box(0); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_28); +lean_ctor_set(x_31, 1, x_30); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +x_33 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__3; +x_34 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_34, 0, x_19); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_29); +lean_ctor_set(x_34, 3, x_32); +x_35 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_36 = lean_array_push(x_35, x_17); +x_37 = lean_array_push(x_36, x_6); +x_38 = lean_box(2); +lean_inc(x_2); +x_39 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_2); +lean_ctor_set(x_39, 2, x_37); +x_40 = lean_array_push(x_35, x_34); +x_41 = lean_array_push(x_40, x_39); +x_42 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_43 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_43, 0, x_38); +lean_ctor_set(x_43, 1, x_42); +lean_ctor_set(x_43, 2, x_41); +x_4 = x_16; +x_6 = x_43; +x_13 = x_24; +goto _start; +} +else +{ +lean_object* x_45; +lean_dec(x_11); +lean_dec(x_2); +lean_dec(x_1); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_6); +lean_ctor_set(x_45, 1, x_13); +return x_45; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8(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_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_4, x_5); +if (x_14 == 0) +{ +size_t x_15; size_t 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; 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_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_15 = 1; +x_16 = lean_usize_sub(x_4, x_15); +x_17 = lean_array_uget(x_3, x_16); +lean_inc(x_11); +x_18 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_11, x_12, x_13); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_ctor_get(x_11, 10); +lean_inc(x_21); +x_22 = lean_st_ref_get(x_12, x_20); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_environment_main_module(x_25); +x_27 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__4; +lean_inc(x_1); +x_28 = lean_name_mk_string(x_1, x_27); +lean_inc(x_28); +x_29 = l_Lean_addMacroScope(x_26, x_28, x_21); +x_30 = lean_box(0); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_28); +lean_ctor_set(x_31, 1, x_30); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +x_33 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__3; +x_34 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_34, 0, x_19); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_29); +lean_ctor_set(x_34, 3, x_32); +x_35 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_36 = lean_array_push(x_35, x_17); +x_37 = lean_array_push(x_36, x_6); +x_38 = lean_box(2); +lean_inc(x_2); +x_39 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_2); +lean_ctor_set(x_39, 2, x_37); +x_40 = lean_array_push(x_35, x_34); +x_41 = lean_array_push(x_40, x_39); +x_42 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_43 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_43, 0, x_38); +lean_ctor_set(x_43, 1, x_42); +lean_ctor_set(x_43, 2, x_41); +x_4 = x_16; +x_6 = x_43; +x_13 = x_24; +goto _start; +} +else +{ +lean_object* x_45; +lean_dec(x_11); +lean_dec(x_2); +lean_dec(x_1); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_6); +lean_ctor_set(x_45, 1, x_13); +return x_45; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9(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_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_4, x_5); +if (x_14 == 0) +{ +size_t x_15; size_t 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; 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_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_15 = 1; +x_16 = lean_usize_sub(x_4, x_15); +x_17 = lean_array_uget(x_3, x_16); +lean_inc(x_11); +x_18 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_11, x_12, x_13); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_ctor_get(x_11, 10); +lean_inc(x_21); +x_22 = lean_st_ref_get(x_12, x_20); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_environment_main_module(x_25); +x_27 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__4; +lean_inc(x_1); +x_28 = lean_name_mk_string(x_1, x_27); +lean_inc(x_28); +x_29 = l_Lean_addMacroScope(x_26, x_28, x_21); +x_30 = lean_box(0); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_28); +lean_ctor_set(x_31, 1, x_30); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +x_33 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__3; +x_34 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_34, 0, x_19); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_29); +lean_ctor_set(x_34, 3, x_32); +x_35 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_36 = lean_array_push(x_35, x_17); +x_37 = lean_array_push(x_36, x_6); +x_38 = lean_box(2); +lean_inc(x_2); +x_39 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_2); +lean_ctor_set(x_39, 2, x_37); +x_40 = lean_array_push(x_35, x_34); +x_41 = lean_array_push(x_40, x_39); +x_42 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_43 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_43, 0, x_38); +lean_ctor_set(x_43, 1, x_42); +lean_ctor_set(x_43, 2, x_41); +x_4 = x_16; +x_6 = x_43; +x_13 = x_24; +goto _start; +} +else +{ +lean_object* x_45; +lean_dec(x_11); +lean_dec(x_2); +lean_dec(x_1); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_6); +lean_ctor_set(x_45, 1, x_13); +return x_45; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__10(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_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_4, x_5); +if (x_14 == 0) +{ +size_t x_15; size_t 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; 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_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_15 = 1; +x_16 = lean_usize_sub(x_4, x_15); +x_17 = lean_array_uget(x_3, x_16); +lean_inc(x_11); +x_18 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_11, x_12, x_13); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_ctor_get(x_11, 10); +lean_inc(x_21); +x_22 = lean_st_ref_get(x_12, x_20); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_environment_main_module(x_25); +x_27 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__4; +lean_inc(x_1); +x_28 = lean_name_mk_string(x_1, x_27); +lean_inc(x_28); +x_29 = l_Lean_addMacroScope(x_26, x_28, x_21); +x_30 = lean_box(0); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_28); +lean_ctor_set(x_31, 1, x_30); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +x_33 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__3; +x_34 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_34, 0, x_19); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_29); +lean_ctor_set(x_34, 3, x_32); +x_35 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_36 = lean_array_push(x_35, x_17); +x_37 = lean_array_push(x_36, x_6); +x_38 = lean_box(2); +lean_inc(x_2); +x_39 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_2); +lean_ctor_set(x_39, 2, x_37); +x_40 = lean_array_push(x_35, x_34); +x_41 = lean_array_push(x_40, x_39); +x_42 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_43 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_43, 0, x_38); +lean_ctor_set(x_43, 1, x_42); +lean_ctor_set(x_43, 2, x_41); +x_4 = x_16; +x_6 = x_43; +x_13 = x_24; +goto _start; +} +else +{ +lean_object* x_45; +lean_dec(x_11); +lean_dec(x_2); +lean_dec(x_1); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_6); +lean_ctor_set(x_45, 1, x_13); +return x_45; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__11(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_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_4, x_5); +if (x_14 == 0) +{ +size_t x_15; size_t 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; 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_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_15 = 1; +x_16 = lean_usize_sub(x_4, x_15); +x_17 = lean_array_uget(x_3, x_16); +lean_inc(x_11); +x_18 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_11, x_12, x_13); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_ctor_get(x_11, 10); +lean_inc(x_21); +x_22 = lean_st_ref_get(x_12, x_20); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_environment_main_module(x_25); +x_27 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__4; +lean_inc(x_1); +x_28 = lean_name_mk_string(x_1, x_27); +lean_inc(x_28); +x_29 = l_Lean_addMacroScope(x_26, x_28, x_21); +x_30 = lean_box(0); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_28); +lean_ctor_set(x_31, 1, x_30); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +x_33 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__3; +x_34 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_34, 0, x_19); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_29); +lean_ctor_set(x_34, 3, x_32); +x_35 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_36 = lean_array_push(x_35, x_17); +x_37 = lean_array_push(x_36, x_6); +x_38 = lean_box(2); +lean_inc(x_2); +x_39 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_2); +lean_ctor_set(x_39, 2, x_37); +x_40 = lean_array_push(x_35, x_34); +x_41 = lean_array_push(x_40, x_39); +x_42 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_43 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_43, 0, x_38); +lean_ctor_set(x_43, 1, x_42); +lean_ctor_set(x_43, 2, x_41); +x_4 = x_16; +x_6 = x_43; +x_13 = x_24; +goto _start; +} +else +{ +lean_object* x_45; +lean_dec(x_11); +lean_dec(x_2); lean_dec(x_1); x_45 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_45, 0, x_5); -lean_ctor_set(x_45, 1, x_12); +lean_ctor_set(x_45, 0, x_6); +lean_ctor_set(x_45, 1, x_13); return x_45; } } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__1() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__12___closed__1() { _start: { lean_object* x_1; @@ -3825,7 +4415,7 @@ x_1 = lean_mk_string_from_bytes("matchDiscr", 10); return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__12(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { _start: { uint8_t x_6; @@ -3842,7 +4432,7 @@ lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_ob x_7 = lean_array_uget(x_5, x_4); x_8 = lean_unsigned_to_nat(0u); x_9 = lean_array_uset(x_5, x_4, x_8); -x_10 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__1; +x_10 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__12___closed__1; lean_inc(x_1); x_11 = lean_name_mk_string(x_1, x_10); x_12 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; @@ -3863,7 +4453,7 @@ goto _start; } } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__1() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__1() { _start: { lean_object* x_1; @@ -3871,22 +4461,22 @@ x_1 = lean_mk_string_from_bytes("some", 4); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__2() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__1; +x_1 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__1; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__3() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__1; +x_1 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__2; +x_3 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__2; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3894,17 +4484,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__4() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__1; +x_2 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__5() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__5() { _start: { lean_object* x_1; @@ -3912,51 +4502,51 @@ x_1 = lean_mk_string_from_bytes("Option", 6); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__6() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__5; +x_2 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__7() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__6; -x_2 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__1; +x_1 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__6; +x_2 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__8() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__7; +x_2 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__7; 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_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__9() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__8; +x_2 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__8; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, size_t x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, size_t x_7, lean_object* x_8) { _start: { uint8_t x_9; @@ -3977,27 +4567,27 @@ x_10 = lean_array_uget(x_8, x_7); x_11 = lean_unsigned_to_nat(0u); x_12 = lean_array_uset(x_8, x_7, x_11); x_13 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__3; +lean_inc(x_5); +x_14 = lean_name_mk_string(x_5, x_13); +x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__4; +lean_inc(x_3); lean_inc(x_4); -x_14 = lean_name_mk_string(x_4, x_13); -x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__4; +x_16 = l_Lean_addMacroScope(x_4, x_15, x_3); +x_17 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__3; +x_18 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__9; lean_inc(x_2); -lean_inc(x_3); -x_16 = l_Lean_addMacroScope(x_3, x_15, x_2); -x_17 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__3; -x_18 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__9; -lean_inc(x_1); x_19 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_19, 0, x_1); +lean_ctor_set(x_19, 0, x_2); lean_ctor_set(x_19, 1, x_17); lean_ctor_set(x_19, 2, x_16); lean_ctor_set(x_19, 3, x_18); x_20 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; x_21 = lean_array_push(x_20, x_10); x_22 = lean_box(2); -lean_inc(x_5); +lean_inc(x_1); x_23 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_5); +lean_ctor_set(x_23, 1, x_1); lean_ctor_set(x_23, 2, x_21); x_24 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; x_25 = lean_array_push(x_24, x_19); @@ -4015,7 +4605,7 @@ goto _start; } } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__1() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__14___closed__1() { _start: { lean_object* x_1; @@ -4023,7 +4613,7 @@ x_1 = lean_mk_string_from_bytes("hole", 4); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__2() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__14___closed__2() { _start: { lean_object* x_1; @@ -4031,7 +4621,7 @@ x_1 = lean_mk_string_from_bytes("_", 1); return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6(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_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__14(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { _start: { uint8_t x_7; @@ -4049,7 +4639,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_o x_8 = lean_array_uget(x_6, x_5); x_9 = lean_unsigned_to_nat(0u); x_10 = lean_array_uset(x_6, x_5, x_9); -x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__1; +x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__14___closed__1; lean_inc(x_2); x_12 = lean_name_mk_string(x_2, x_11); x_13 = l_Lean_Syntax_getHeadInfo_x3f(x_8); @@ -4059,7 +4649,7 @@ x_15 = lean_usize_add(x_5, x_14); if (lean_obj_tag(x_13) == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_16 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__2; +x_16 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__14___closed__2; lean_inc(x_1); x_17 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_17, 0, x_1); @@ -4082,7 +4672,7 @@ lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean x_23 = lean_ctor_get(x_13, 0); lean_inc(x_23); lean_dec(x_13); -x_24 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__2; +x_24 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__14___closed__2; x_25 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_25, 0, x_23); lean_ctor_set(x_25, 1, x_24); @@ -4101,7 +4691,499 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__7(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__15(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_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_4, x_5); +if (x_14 == 0) +{ +size_t x_15; size_t 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; 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_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_15 = 1; +x_16 = lean_usize_sub(x_4, x_15); +x_17 = lean_array_uget(x_3, x_16); +lean_inc(x_11); +x_18 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_11, x_12, x_13); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_ctor_get(x_11, 10); +lean_inc(x_21); +x_22 = lean_st_ref_get(x_12, x_20); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_environment_main_module(x_25); +x_27 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__4; +lean_inc(x_1); +x_28 = lean_name_mk_string(x_1, x_27); +lean_inc(x_28); +x_29 = l_Lean_addMacroScope(x_26, x_28, x_21); +x_30 = lean_box(0); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_28); +lean_ctor_set(x_31, 1, x_30); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +x_33 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__3; +x_34 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_34, 0, x_19); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_29); +lean_ctor_set(x_34, 3, x_32); +x_35 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_36 = lean_array_push(x_35, x_17); +x_37 = lean_array_push(x_36, x_6); +x_38 = lean_box(2); +lean_inc(x_2); +x_39 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_2); +lean_ctor_set(x_39, 2, x_37); +x_40 = lean_array_push(x_35, x_34); +x_41 = lean_array_push(x_40, x_39); +x_42 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_43 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_43, 0, x_38); +lean_ctor_set(x_43, 1, x_42); +lean_ctor_set(x_43, 2, x_41); +x_4 = x_16; +x_6 = x_43; +x_13 = x_24; +goto _start; +} +else +{ +lean_object* x_45; +lean_dec(x_11); +lean_dec(x_2); +lean_dec(x_1); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_6); +lean_ctor_set(x_45, 1, x_13); +return x_45; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__16(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_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_4, x_5); +if (x_14 == 0) +{ +size_t x_15; size_t 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; 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_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_15 = 1; +x_16 = lean_usize_sub(x_4, x_15); +x_17 = lean_array_uget(x_3, x_16); +lean_inc(x_11); +x_18 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_11, x_12, x_13); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_ctor_get(x_11, 10); +lean_inc(x_21); +x_22 = lean_st_ref_get(x_12, x_20); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_environment_main_module(x_25); +x_27 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__4; +lean_inc(x_1); +x_28 = lean_name_mk_string(x_1, x_27); +lean_inc(x_28); +x_29 = l_Lean_addMacroScope(x_26, x_28, x_21); +x_30 = lean_box(0); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_28); +lean_ctor_set(x_31, 1, x_30); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +x_33 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__3; +x_34 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_34, 0, x_19); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_29); +lean_ctor_set(x_34, 3, x_32); +x_35 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_36 = lean_array_push(x_35, x_17); +x_37 = lean_array_push(x_36, x_6); +x_38 = lean_box(2); +lean_inc(x_2); +x_39 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_2); +lean_ctor_set(x_39, 2, x_37); +x_40 = lean_array_push(x_35, x_34); +x_41 = lean_array_push(x_40, x_39); +x_42 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_43 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_43, 0, x_38); +lean_ctor_set(x_43, 1, x_42); +lean_ctor_set(x_43, 2, x_41); +x_4 = x_16; +x_6 = x_43; +x_13 = x_24; +goto _start; +} +else +{ +lean_object* x_45; +lean_dec(x_11); +lean_dec(x_2); +lean_dec(x_1); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_6); +lean_ctor_set(x_45, 1, x_13); +return x_45; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__17(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_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_4, x_5); +if (x_14 == 0) +{ +size_t x_15; size_t 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; 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_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_15 = 1; +x_16 = lean_usize_sub(x_4, x_15); +x_17 = lean_array_uget(x_3, x_16); +lean_inc(x_11); +x_18 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_11, x_12, x_13); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_ctor_get(x_11, 10); +lean_inc(x_21); +x_22 = lean_st_ref_get(x_12, x_20); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_environment_main_module(x_25); +x_27 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__4; +lean_inc(x_1); +x_28 = lean_name_mk_string(x_1, x_27); +lean_inc(x_28); +x_29 = l_Lean_addMacroScope(x_26, x_28, x_21); +x_30 = lean_box(0); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_28); +lean_ctor_set(x_31, 1, x_30); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +x_33 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__3; +x_34 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_34, 0, x_19); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_29); +lean_ctor_set(x_34, 3, x_32); +x_35 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_36 = lean_array_push(x_35, x_17); +x_37 = lean_array_push(x_36, x_6); +x_38 = lean_box(2); +lean_inc(x_2); +x_39 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_2); +lean_ctor_set(x_39, 2, x_37); +x_40 = lean_array_push(x_35, x_34); +x_41 = lean_array_push(x_40, x_39); +x_42 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_43 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_43, 0, x_38); +lean_ctor_set(x_43, 1, x_42); +lean_ctor_set(x_43, 2, x_41); +x_4 = x_16; +x_6 = x_43; +x_13 = x_24; +goto _start; +} +else +{ +lean_object* x_45; +lean_dec(x_11); +lean_dec(x_2); +lean_dec(x_1); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_6); +lean_ctor_set(x_45, 1, x_13); +return x_45; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__18(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_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_4, x_5); +if (x_14 == 0) +{ +size_t x_15; size_t 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; 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_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_15 = 1; +x_16 = lean_usize_sub(x_4, x_15); +x_17 = lean_array_uget(x_3, x_16); +lean_inc(x_11); +x_18 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_11, x_12, x_13); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_ctor_get(x_11, 10); +lean_inc(x_21); +x_22 = lean_st_ref_get(x_12, x_20); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_environment_main_module(x_25); +x_27 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__4; +lean_inc(x_1); +x_28 = lean_name_mk_string(x_1, x_27); +lean_inc(x_28); +x_29 = l_Lean_addMacroScope(x_26, x_28, x_21); +x_30 = lean_box(0); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_28); +lean_ctor_set(x_31, 1, x_30); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +x_33 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__3; +x_34 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_34, 0, x_19); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_29); +lean_ctor_set(x_34, 3, x_32); +x_35 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_36 = lean_array_push(x_35, x_17); +x_37 = lean_array_push(x_36, x_6); +x_38 = lean_box(2); +lean_inc(x_2); +x_39 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_2); +lean_ctor_set(x_39, 2, x_37); +x_40 = lean_array_push(x_35, x_34); +x_41 = lean_array_push(x_40, x_39); +x_42 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_43 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_43, 0, x_38); +lean_ctor_set(x_43, 1, x_42); +lean_ctor_set(x_43, 2, x_41); +x_4 = x_16; +x_6 = x_43; +x_13 = x_24; +goto _start; +} +else +{ +lean_object* x_45; +lean_dec(x_11); +lean_dec(x_2); +lean_dec(x_1); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_6); +lean_ctor_set(x_45, 1, x_13); +return x_45; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__19(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_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_4, x_5); +if (x_14 == 0) +{ +size_t x_15; size_t 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; 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_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_15 = 1; +x_16 = lean_usize_sub(x_4, x_15); +x_17 = lean_array_uget(x_3, x_16); +lean_inc(x_11); +x_18 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_11, x_12, x_13); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_ctor_get(x_11, 10); +lean_inc(x_21); +x_22 = lean_st_ref_get(x_12, x_20); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_environment_main_module(x_25); +x_27 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__4; +lean_inc(x_1); +x_28 = lean_name_mk_string(x_1, x_27); +lean_inc(x_28); +x_29 = l_Lean_addMacroScope(x_26, x_28, x_21); +x_30 = lean_box(0); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_28); +lean_ctor_set(x_31, 1, x_30); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +x_33 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__3; +x_34 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_34, 0, x_19); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_29); +lean_ctor_set(x_34, 3, x_32); +x_35 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_36 = lean_array_push(x_35, x_17); +x_37 = lean_array_push(x_36, x_6); +x_38 = lean_box(2); +lean_inc(x_2); +x_39 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_2); +lean_ctor_set(x_39, 2, x_37); +x_40 = lean_array_push(x_35, x_34); +x_41 = lean_array_push(x_40, x_39); +x_42 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_43 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_43, 0, x_38); +lean_ctor_set(x_43, 1, x_42); +lean_ctor_set(x_43, 2, x_41); +x_4 = x_16; +x_6 = x_43; +x_13 = x_24; +goto _start; +} +else +{ +lean_object* x_45; +lean_dec(x_11); +lean_dec(x_2); +lean_dec(x_1); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_6); +lean_ctor_set(x_45, 1, x_13); +return x_45; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__20(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_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_4, x_5); +if (x_14 == 0) +{ +size_t x_15; size_t 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; 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_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_15 = 1; +x_16 = lean_usize_sub(x_4, x_15); +x_17 = lean_array_uget(x_3, x_16); +lean_inc(x_11); +x_18 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_11, x_12, x_13); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_ctor_get(x_11, 10); +lean_inc(x_21); +x_22 = lean_st_ref_get(x_12, x_20); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_environment_main_module(x_25); +x_27 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__4; +lean_inc(x_1); +x_28 = lean_name_mk_string(x_1, x_27); +lean_inc(x_28); +x_29 = l_Lean_addMacroScope(x_26, x_28, x_21); +x_30 = lean_box(0); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_28); +lean_ctor_set(x_31, 1, x_30); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +x_33 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__3; +x_34 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_34, 0, x_19); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_29); +lean_ctor_set(x_34, 3, x_32); +x_35 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_36 = lean_array_push(x_35, x_17); +x_37 = lean_array_push(x_36, x_6); +x_38 = lean_box(2); +lean_inc(x_2); +x_39 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_2); +lean_ctor_set(x_39, 2, x_37); +x_40 = lean_array_push(x_35, x_34); +x_41 = lean_array_push(x_40, x_39); +x_42 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_43 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_43, 0, x_38); +lean_ctor_set(x_43, 1, x_42); +lean_ctor_set(x_43, 2, x_41); +x_4 = x_16; +x_6 = x_43; +x_13 = x_24; +goto _start; +} +else +{ +lean_object* x_45; +lean_dec(x_11); +lean_dec(x_2); +lean_dec(x_1); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_6); +lean_ctor_set(x_45, 1, x_13); +return x_45; +} +} +} +LEAN_EXPORT lean_object* l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__21(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -4153,7 +5235,127 @@ goto _start; } } } -static lean_object* _init_l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__1() { +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23(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; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_9 = lean_ctor_get(x_6, 5); +x_10 = lean_ctor_get(x_2, 1); +lean_inc(x_10); +lean_inc(x_10); +x_11 = l_Lean_Elab_getBetterRef(x_9, x_10); +x_12 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_4, x_5, x_6, x_7, x_8); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(x_13, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_14); +lean_dec(x_2); +lean_dec(x_10); +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_15, 0); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_11); +lean_ctor_set(x_18, 1, x_17); +lean_ctor_set_tag(x_15, 1); +lean_ctor_set(x_15, 0, x_18); +return x_15; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_19 = lean_ctor_get(x_15, 0); +x_20 = lean_ctor_get(x_15, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_15); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_11); +lean_ctor_set(x_21, 1, x_19); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_20); +return x_22; +} +} +} +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__22(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_object* x_9) { +_start: +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_7); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_ctor_get(x_7, 5); +x_12 = l_Lean_replaceRef(x_1, x_11); +lean_dec(x_11); +lean_dec(x_1); +lean_ctor_set(x_7, 5, x_12); +x_13 = l_Lean_throwError___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_13; +} +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; 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; +x_14 = lean_ctor_get(x_7, 0); +x_15 = lean_ctor_get(x_7, 1); +x_16 = lean_ctor_get(x_7, 2); +x_17 = lean_ctor_get(x_7, 3); +x_18 = lean_ctor_get(x_7, 4); +x_19 = lean_ctor_get(x_7, 5); +x_20 = lean_ctor_get(x_7, 6); +x_21 = lean_ctor_get(x_7, 7); +x_22 = lean_ctor_get(x_7, 8); +x_23 = lean_ctor_get(x_7, 9); +x_24 = lean_ctor_get(x_7, 10); +lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_dec(x_7); +x_25 = l_Lean_replaceRef(x_1, x_19); +lean_dec(x_19); +lean_dec(x_1); +x_26 = lean_alloc_ctor(0, 11, 0); +lean_ctor_set(x_26, 0, x_14); +lean_ctor_set(x_26, 1, x_15); +lean_ctor_set(x_26, 2, x_16); +lean_ctor_set(x_26, 3, x_17); +lean_ctor_set(x_26, 4, x_18); +lean_ctor_set(x_26, 5, x_25); +lean_ctor_set(x_26, 6, x_20); +lean_ctor_set(x_26, 7, x_21); +lean_ctor_set(x_26, 8, x_22); +lean_ctor_set(x_26, 9, x_23); +lean_ctor_set(x_26, 10, x_24); +x_27 = l_Lean_throwError___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23(x_2, x_3, x_4, x_5, x_6, x_26, x_8, x_9); +lean_dec(x_8); +lean_dec(x_26); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_27; +} +} +} +static lean_object* _init_l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__1() { _start: { lean_object* x_1; @@ -4161,17 +5363,17 @@ x_1 = lean_mk_string_from_bytes("quotedName", 10); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__2() { +static lean_object* _init_l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_2 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__1; +x_2 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__3() { +static lean_object* _init_l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3() { _start: { lean_object* x_1; @@ -4179,7 +5381,7 @@ x_1 = lean_mk_string_from_bytes(".", 1); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__4() { +static lean_object* _init_l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4() { _start: { lean_object* x_1; @@ -4187,7 +5389,7 @@ x_1 = lean_mk_string_from_bytes("`", 1); return x_1; } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24(lean_object* x_1) { _start: { if (lean_obj_tag(x_1) == 0) @@ -4207,11 +5409,11 @@ lean_dec(x_1); x_5 = lean_box(0); lean_inc(x_3); x_6 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_5, x_3); -x_7 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8(x_4); +x_7 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24(x_4); if (lean_obj_tag(x_6) == 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; -x_8 = l___private_Init_Meta_0__Lean_quoteNameMk(x_3); +x_8 = l_Lean_quoteNameMk(x_3); x_9 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; x_10 = lean_array_push(x_9, x_8); x_11 = lean_array_push(x_10, x_7); @@ -4221,37 +5423,36 @@ return x_13; } 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; 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; lean_object* x_30; +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; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_dec(x_3); x_14 = lean_ctor_get(x_6, 0); lean_inc(x_14); lean_dec(x_6); -x_15 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__3; +x_15 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; x_16 = l_String_intercalate(x_15, x_14); -x_17 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__4; +x_17 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_nameLitKind; -x_20 = lean_box(2); -x_21 = l_Lean_Syntax_mkLit(x_19, x_18, x_20); -x_22 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; -x_23 = lean_array_push(x_22, x_21); -x_24 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__2; -x_25 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_25, 0, x_20); -lean_ctor_set(x_25, 1, x_24); -lean_ctor_set(x_25, 2, x_23); -x_26 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_27 = lean_array_push(x_26, x_25); -x_28 = lean_array_push(x_27, x_7); -x_29 = l___private_Init_Meta_0__Lean_quoteList___at_Lean_Elab_Term_Quotation_ArrayStxBuilder_build___spec__1___closed__7; -x_30 = l_Lean_Syntax_mkCApp(x_29, x_28); -return x_30; +x_19 = lean_box(2); +x_20 = l_Lean_Syntax_mkNameLit(x_18, x_19); +x_21 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; +x_22 = lean_array_push(x_21, x_20); +x_23 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2; +x_24 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_24, 0, x_19); +lean_ctor_set(x_24, 1, x_23); +lean_ctor_set(x_24, 2, x_22); +x_25 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_26 = lean_array_push(x_25, x_24); +x_27 = lean_array_push(x_26, x_7); +x_28 = l___private_Init_Meta_0__Lean_quoteList___at_Lean_Elab_Term_Quotation_ArrayStxBuilder_build___spec__1___closed__7; +x_29 = l_Lean_Syntax_mkCApp(x_28, x_27); +return x_29; } } } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___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, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___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, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; @@ -4313,7 +5514,7 @@ return x_25; } } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__1() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__1() { _start: { lean_object* x_1; @@ -4321,27 +5522,27 @@ x_1 = lean_mk_string_from_bytes("sepBy", 5); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__2() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__1; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__3() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__2; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__2; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__4() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__4() { _start: { lean_object* x_1; @@ -4349,22 +5550,22 @@ x_1 = lean_mk_string_from_bytes("mkSepArray", 10); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__5() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__4; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__4; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__6() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__4; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__4; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__5; +x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__5; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -4372,119 +5573,119 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__7() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__4; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__8() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__2; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__4; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__9() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__8; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__8; 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__10() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__9; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__9; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_14; uint8_t x_15; -x_14 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__3; -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__2(x_3, x_14); +lean_object* x_15; uint8_t x_16; +x_15 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__3; +x_16 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5(x_3, x_15); lean_dec(x_3); -if (x_15 == 0) +if (x_16 == 0) { -lean_object* x_16; +lean_object* x_17; +lean_dec(x_5); lean_dec(x_4); -x_16 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__1(x_1, x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_16; +x_17 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__1(x_1, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_17; } else { -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; 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_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_inc(x_11); -x_17 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_11, x_12, x_13); -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); +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; 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_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_inc(x_12); +x_18 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_12, x_13, x_14); +x_19 = lean_ctor_get(x_18, 0); lean_inc(x_19); -lean_dec(x_17); -x_20 = lean_ctor_get(x_11, 10); +x_20 = lean_ctor_get(x_18, 1); lean_inc(x_20); -x_21 = lean_st_ref_get(x_12, x_19); -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_21, 1); +lean_dec(x_18); +x_21 = lean_ctor_get(x_12, 10); +lean_inc(x_21); +x_22 = lean_st_ref_get(x_13, x_20); +x_23 = lean_ctor_get(x_22, 0); lean_inc(x_23); -lean_dec(x_21); -x_24 = lean_ctor_get(x_22, 0); +x_24 = lean_ctor_get(x_22, 1); lean_inc(x_24); lean_dec(x_22); -x_25 = lean_environment_main_module(x_24); -x_26 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__7; -x_27 = l_Lean_addMacroScope(x_25, x_26, x_20); -x_28 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__6; -x_29 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__10; -x_30 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_30, 0, x_18); -lean_ctor_set(x_30, 1, x_28); -lean_ctor_set(x_30, 2, x_27); -lean_ctor_set(x_30, 3, x_29); -x_31 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice(x_4); +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_environment_main_module(x_25); +x_27 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__7; +x_28 = l_Lean_addMacroScope(x_26, x_27, x_21); +x_29 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__6; +x_30 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__10; +x_31 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_31, 0, x_19); +lean_ctor_set(x_31, 1, x_29); +lean_ctor_set(x_31, 2, x_28); +lean_ctor_set(x_31, 3, x_30); +x_32 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice(x_4); lean_dec(x_4); -x_32 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_33 = lean_array_push(x_32, x_6); -x_34 = lean_array_push(x_33, x_31); -x_35 = lean_box(2); -x_36 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_33 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_34 = lean_array_push(x_33, x_7); +x_35 = lean_array_push(x_34, x_32); +x_36 = lean_box(2); x_37 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_37, 0, x_35); -lean_ctor_set(x_37, 1, x_36); -lean_ctor_set(x_37, 2, x_34); -x_38 = lean_array_push(x_32, x_30); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_5); +lean_ctor_set(x_37, 2, x_35); +x_38 = lean_array_push(x_33, x_31); x_39 = lean_array_push(x_38, x_37); x_40 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; x_41 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_41, 0, x_35); +lean_ctor_set(x_41, 0, x_36); lean_ctor_set(x_41, 1, x_40); lean_ctor_set(x_41, 2, x_39); -x_42 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__1(x_1, x_2, x_5, x_41, x_7, x_8, x_9, x_10, x_11, x_12, x_23); +x_42 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__1(x_1, x_2, x_6, x_41, x_8, x_9, x_10, x_11, x_12, x_13, x_24); return x_42; } } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__1() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__1() { _start: { lean_object* x_1; @@ -4492,22 +5693,22 @@ x_1 = lean_mk_string_from_bytes("Array.map", 9); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__2() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__1; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__1; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__3() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__1; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__2; +x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__2; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -4515,7 +5716,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__4() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__4() { _start: { lean_object* x_1; @@ -4523,7 +5724,7 @@ x_1 = lean_mk_string_from_bytes("map", 3); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__5() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__5() { _start: { lean_object* x_1; @@ -4531,17 +5732,17 @@ x_1 = lean_mk_string_from_bytes("paren", 5); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__6() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__5; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7() { _start: { lean_object* x_1; @@ -4549,7 +5750,7 @@ x_1 = lean_mk_string_from_bytes("(", 1); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__8() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__8() { _start: { lean_object* x_1; @@ -4557,17 +5758,17 @@ x_1 = lean_mk_string_from_bytes("fun", 3); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__9() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__8; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__8; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__10() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__10() { _start: { lean_object* x_1; @@ -4575,17 +5776,17 @@ x_1 = lean_mk_string_from_bytes("basicFun", 8); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__11() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__10; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__10; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__12() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__12() { _start: { lean_object* x_1; @@ -4593,7 +5794,7 @@ x_1 = lean_mk_string_from_bytes("=>", 2); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13() { _start: { lean_object* x_1; lean_object* x_2; @@ -4602,7 +5803,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14() { _start: { lean_object* x_1; @@ -4610,7 +5811,7 @@ x_1 = lean_mk_string_from_bytes(")", 1); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__15() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__15() { _start: { lean_object* x_1; @@ -4618,7 +5819,7 @@ x_1 = lean_mk_string_from_bytes("optional", 8); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__16() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__16() { _start: { lean_object* x_1; @@ -4626,17 +5827,17 @@ x_1 = lean_mk_string_from_bytes("match", 5); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__17() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__16; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__16; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__18() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__18() { _start: { lean_object* x_1; @@ -4644,19 +5845,19 @@ x_1 = lean_mk_string_from_bytes(",", 1); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__19() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(2); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__18; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__18; x_3 = lean_alloc_ctor(2, 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__20() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__20() { _start: { lean_object* x_1; @@ -4664,7 +5865,7 @@ x_1 = lean_mk_string_from_bytes("with", 4); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__21() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__21() { _start: { lean_object* x_1; @@ -4672,17 +5873,17 @@ x_1 = lean_mk_string_from_bytes("matchAlts", 9); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__22() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__21; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__21; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__23() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__23() { _start: { lean_object* x_1; @@ -4690,17 +5891,17 @@ x_1 = lean_mk_string_from_bytes("matchAlt", 8); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__24() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__23; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__23; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__25() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__25() { _start: { lean_object* x_1; @@ -4708,7 +5909,7 @@ x_1 = lean_mk_string_from_bytes("|", 1); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__26() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__26() { _start: { lean_object* x_1; @@ -4716,22 +5917,22 @@ x_1 = lean_mk_string_from_bytes("Array.empty", 11); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__27() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__27() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__26; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__26; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__28() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__26; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__26; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__27; +x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__27; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -4739,7 +5940,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__29() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__29() { _start: { lean_object* x_1; @@ -4747,7 +5948,7 @@ x_1 = lean_mk_string_from_bytes("empty", 5); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__30() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__30() { _start: { lean_object* x_1; lean_object* x_2; @@ -4756,758 +5957,793 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3(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, size_t 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, lean_object* x_15, lean_object* x_16, lean_object* x_17) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3(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, size_t 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, lean_object* x_16, lean_object* x_17, lean_object* x_18) { _start: { -lean_dec(x_10); +lean_dec(x_11); if (lean_obj_tag(x_3) == 0) { -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; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; lean_object* x_31; -x_18 = l_Lean_instInhabitedSyntax; -x_19 = l_Array_back___rarg(x_18, x_5); -x_20 = lean_array_get_size(x_5); -x_21 = lean_unsigned_to_nat(1u); -x_22 = lean_nat_sub(x_20, x_21); -lean_dec(x_20); -x_23 = lean_unsigned_to_nat(0u); -lean_inc(x_5); -x_24 = l_Array_toSubarray___rarg(x_5, x_23, x_22); -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -x_26 = lean_ctor_get(x_24, 2); -lean_inc(x_26); -x_27 = lean_ctor_get(x_24, 1); -lean_inc(x_27); -lean_dec(x_24); -x_28 = lean_array_get_size(x_25); -x_29 = lean_nat_dec_le(x_26, x_28); -if (x_29 == 0) +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; uint8_t x_30; lean_object* x_31; lean_object* x_32; +x_19 = lean_box(0); +x_20 = l_Array_back___rarg(x_19, x_6); +x_21 = lean_array_get_size(x_6); +x_22 = lean_unsigned_to_nat(1u); +x_23 = lean_nat_sub(x_21, x_22); +lean_dec(x_21); +x_24 = lean_unsigned_to_nat(0u); +lean_inc(x_6); +x_25 = l_Array_toSubarray___rarg(x_6, x_24, x_23); +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_25, 2); +lean_inc(x_27); +x_28 = lean_ctor_get(x_25, 1); +lean_inc(x_28); +lean_dec(x_25); +x_29 = lean_array_get_size(x_26); +x_30 = lean_nat_dec_le(x_27, x_29); +if (x_30 == 0) { -uint8_t x_95; -lean_dec(x_26); -x_95 = lean_nat_dec_lt(x_27, x_28); -if (x_95 == 0) +uint8_t x_96; +lean_dec(x_27); +x_96 = lean_nat_dec_lt(x_28, x_29); +if (x_96 == 0) { +lean_dec(x_29); lean_dec(x_28); -lean_dec(x_27); -lean_dec(x_25); -x_30 = x_19; -x_31 = x_17; -goto block_94; +lean_dec(x_26); +x_31 = x_20; +x_32 = x_18; +goto block_95; } else { -size_t x_96; size_t x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_96 = lean_usize_of_nat(x_28); +size_t x_97; size_t x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_97 = lean_usize_of_nat(x_29); +lean_dec(x_29); +x_98 = lean_usize_of_nat(x_28); lean_dec(x_28); -x_97 = lean_usize_of_nat(x_27); -lean_dec(x_27); -lean_inc(x_15); -lean_inc(x_6); -x_98 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3(x_6, x_25, x_96, x_97, x_19, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -lean_dec(x_25); -x_99 = lean_ctor_get(x_98, 0); -lean_inc(x_99); -x_100 = lean_ctor_get(x_98, 1); +lean_inc(x_16); +lean_inc(x_5); +lean_inc(x_7); +x_99 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6(x_7, x_5, x_26, x_97, x_98, x_20, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_26); +x_100 = lean_ctor_get(x_99, 0); lean_inc(x_100); -lean_dec(x_98); -x_30 = x_99; +x_101 = lean_ctor_get(x_99, 1); +lean_inc(x_101); +lean_dec(x_99); x_31 = x_100; -goto block_94; +x_32 = x_101; +goto block_95; } } else { -uint8_t x_101; -lean_dec(x_28); -x_101 = lean_nat_dec_lt(x_27, x_26); -if (x_101 == 0) +uint8_t x_102; +lean_dec(x_29); +x_102 = lean_nat_dec_lt(x_28, x_27); +if (x_102 == 0) { +lean_dec(x_28); lean_dec(x_27); lean_dec(x_26); -lean_dec(x_25); -x_30 = x_19; -x_31 = x_17; -goto block_94; +x_31 = x_20; +x_32 = x_18; +goto block_95; } else { -size_t x_102; size_t x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; -x_102 = lean_usize_of_nat(x_26); -lean_dec(x_26); +size_t x_103; size_t x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; x_103 = lean_usize_of_nat(x_27); lean_dec(x_27); -lean_inc(x_15); -lean_inc(x_6); -x_104 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3(x_6, x_25, x_102, x_103, x_19, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -lean_dec(x_25); -x_105 = lean_ctor_get(x_104, 0); -lean_inc(x_105); -x_106 = lean_ctor_get(x_104, 1); +x_104 = lean_usize_of_nat(x_28); +lean_dec(x_28); +lean_inc(x_16); +lean_inc(x_5); +lean_inc(x_7); +x_105 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__7(x_7, x_5, x_26, x_103, x_104, x_20, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_26); +x_106 = lean_ctor_get(x_105, 0); lean_inc(x_106); -lean_dec(x_104); -x_30 = x_105; +x_107 = lean_ctor_get(x_105, 1); +lean_inc(x_107); +lean_dec(x_105); x_31 = x_106; -goto block_94; +x_32 = x_107; +goto block_95; } } -block_94: +block_95: { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; -lean_inc(x_15); -x_32 = l_Lean_Elab_Term_Quotation_mkTuple(x_5, x_11, x_12, x_13, x_14, x_15, x_16, x_31); -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_32, 1); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +lean_inc(x_16); +x_33 = l_Lean_Elab_Term_Quotation_mkTuple(x_6, x_12, x_13, x_14, x_15, x_16, x_17, x_32); +x_34 = lean_ctor_get(x_33, 0); lean_inc(x_34); -lean_dec(x_32); -lean_inc(x_15); -x_35 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_15, x_16, x_34); -x_36 = lean_ctor_get(x_35, 0); -lean_inc(x_36); -x_37 = lean_ctor_get(x_35, 1); +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +lean_inc(x_16); +x_36 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_16, x_17, x_35); +x_37 = lean_ctor_get(x_36, 0); lean_inc(x_37); -lean_dec(x_35); -x_38 = lean_ctor_get(x_15, 10); +x_38 = lean_ctor_get(x_36, 1); lean_inc(x_38); -x_39 = lean_st_ref_get(x_16, x_37); -x_40 = lean_ctor_get(x_39, 0); -lean_inc(x_40); -x_41 = lean_ctor_get(x_39, 1); +lean_dec(x_36); +x_39 = lean_ctor_get(x_16, 10); +lean_inc(x_39); +x_40 = lean_st_ref_get(x_17, x_38); +x_41 = lean_ctor_get(x_40, 0); lean_inc(x_41); -lean_dec(x_39); -x_42 = lean_ctor_get(x_40, 0); +x_42 = lean_ctor_get(x_40, 1); lean_inc(x_42); lean_dec(x_40); -x_43 = lean_environment_main_module(x_42); -x_44 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__4; -x_45 = lean_name_mk_string(x_6, x_44); -lean_inc(x_45); -x_46 = l_Lean_addMacroScope(x_43, x_45, x_38); -x_47 = lean_box(0); -x_48 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_48, 0, x_45); -lean_ctor_set(x_48, 1, x_47); -x_49 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set(x_49, 1, x_47); -x_50 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__3; -lean_inc(x_36); -x_51 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_51, 0, x_36); -lean_ctor_set(x_51, 1, x_50); -lean_ctor_set(x_51, 2, x_46); -lean_ctor_set(x_51, 3, x_49); -x_52 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; -lean_inc(x_36); -x_53 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_53, 0, x_36); -lean_ctor_set(x_53, 1, x_52); -x_54 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__8; -lean_inc(x_36); -x_55 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_55, 0, x_36); -lean_ctor_set(x_55, 1, x_54); -x_56 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; -x_57 = lean_array_push(x_56, x_33); -x_58 = lean_box(2); -x_59 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_43 = lean_ctor_get(x_41, 0); +lean_inc(x_43); +lean_dec(x_41); +x_44 = lean_environment_main_module(x_43); +x_45 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__4; +x_46 = lean_name_mk_string(x_7, x_45); +lean_inc(x_46); +x_47 = l_Lean_addMacroScope(x_44, x_46, x_39); +x_48 = lean_box(0); +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_46); +lean_ctor_set(x_49, 1, x_48); +x_50 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_50, 0, x_49); +lean_ctor_set(x_50, 1, x_48); +x_51 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__3; +lean_inc(x_37); +x_52 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_52, 0, x_37); +lean_ctor_set(x_52, 1, x_51); +lean_ctor_set(x_52, 2, x_47); +lean_ctor_set(x_52, 3, x_50); +x_53 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; +lean_inc(x_37); +x_54 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_54, 0, x_37); +lean_ctor_set(x_54, 1, x_53); +x_55 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__8; +lean_inc(x_37); +x_56 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_56, 0, x_37); +lean_ctor_set(x_56, 1, x_55); +x_57 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; +x_58 = lean_array_push(x_57, x_34); +x_59 = lean_box(2); +lean_inc(x_5); x_60 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_60, 0, x_58); -lean_ctor_set(x_60, 1, x_59); -lean_ctor_set(x_60, 2, x_57); -x_61 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__12; -lean_inc(x_36); +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set(x_60, 1, x_5); +lean_ctor_set(x_60, 2, x_58); +x_61 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__12; +lean_inc(x_37); x_62 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_62, 0, x_36); +lean_ctor_set(x_62, 0, x_37); lean_ctor_set(x_62, 1, x_61); -x_63 = lean_array_get(x_18, x_7, x_23); -lean_dec(x_7); -x_64 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; +x_63 = lean_array_get(x_19, x_8, x_24); +lean_dec(x_8); +x_64 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; x_65 = lean_array_push(x_64, x_60); x_66 = lean_array_push(x_65, x_62); x_67 = lean_array_push(x_66, x_63); -x_68 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__11; +x_68 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__11; x_69 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_69, 0, x_58); +lean_ctor_set(x_69, 0, x_59); lean_ctor_set(x_69, 1, x_68); lean_ctor_set(x_69, 2, x_67); x_70 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_71 = lean_array_push(x_70, x_55); +x_71 = lean_array_push(x_70, x_56); x_72 = lean_array_push(x_71, x_69); -x_73 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__9; +x_73 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__9; x_74 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_74, 0, x_58); +lean_ctor_set(x_74, 0, x_59); lean_ctor_set(x_74, 1, x_73); lean_ctor_set(x_74, 2, x_72); -x_75 = lean_array_push(x_70, x_74); -x_76 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; -x_77 = lean_array_push(x_75, x_76); -x_78 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_78, 0, x_58); -lean_ctor_set(x_78, 1, x_59); -lean_ctor_set(x_78, 2, x_77); -x_79 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; -x_80 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_80, 0, x_36); -lean_ctor_set(x_80, 1, x_79); -x_81 = lean_array_push(x_64, x_53); -x_82 = lean_array_push(x_81, x_78); -x_83 = lean_array_push(x_82, x_80); -x_84 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__6; -x_85 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_85, 0, x_58); -lean_ctor_set(x_85, 1, x_84); -lean_ctor_set(x_85, 2, x_83); -x_86 = lean_array_push(x_70, x_85); -x_87 = lean_array_push(x_86, x_30); -x_88 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_88, 0, x_58); -lean_ctor_set(x_88, 1, x_59); -lean_ctor_set(x_88, 2, x_87); -x_89 = lean_array_push(x_70, x_51); -x_90 = lean_array_push(x_89, x_88); -x_91 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_92 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_92, 0, x_58); -lean_ctor_set(x_92, 1, x_91); -lean_ctor_set(x_92, 2, x_90); -x_93 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2(x_1, x_2, x_3, x_4, x_9, x_92, x_11, x_12, x_13, x_14, x_15, x_16, x_41); -return x_93; +x_75 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; +lean_inc(x_5); +x_76 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_76, 0, x_59); +lean_ctor_set(x_76, 1, x_5); +lean_ctor_set(x_76, 2, x_75); +x_77 = lean_array_push(x_70, x_74); +x_78 = lean_array_push(x_77, x_76); +lean_inc(x_5); +x_79 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_79, 0, x_59); +lean_ctor_set(x_79, 1, x_5); +lean_ctor_set(x_79, 2, x_78); +x_80 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; +x_81 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_81, 0, x_37); +lean_ctor_set(x_81, 1, x_80); +x_82 = lean_array_push(x_64, x_54); +x_83 = lean_array_push(x_82, x_79); +x_84 = lean_array_push(x_83, x_81); +x_85 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6; +x_86 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_86, 0, x_59); +lean_ctor_set(x_86, 1, x_85); +lean_ctor_set(x_86, 2, x_84); +x_87 = lean_array_push(x_70, x_86); +x_88 = lean_array_push(x_87, x_31); +lean_inc(x_5); +x_89 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_89, 0, x_59); +lean_ctor_set(x_89, 1, x_5); +lean_ctor_set(x_89, 2, x_88); +x_90 = lean_array_push(x_70, x_52); +x_91 = lean_array_push(x_90, x_89); +x_92 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_93 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_93, 0, x_59); +lean_ctor_set(x_93, 1, x_92); +lean_ctor_set(x_93, 2, x_91); +x_94 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2(x_1, x_2, x_3, x_4, x_5, x_10, x_93, x_12, x_13, x_14, x_15, x_16, x_17, x_42); +return x_94; } } else { -lean_object* x_107; -x_107 = lean_ctor_get(x_3, 0); -lean_inc(x_107); -switch (lean_obj_tag(x_107)) { +lean_object* x_108; +x_108 = lean_ctor_get(x_3, 0); +lean_inc(x_108); +switch (lean_obj_tag(x_108)) { case 0: { -lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; uint8_t x_119; lean_object* x_120; lean_object* x_121; -x_108 = l_Lean_instInhabitedSyntax; -x_109 = l_Array_back___rarg(x_108, x_5); -x_110 = lean_array_get_size(x_5); -x_111 = lean_unsigned_to_nat(1u); -x_112 = lean_nat_sub(x_110, x_111); -lean_dec(x_110); -x_113 = lean_unsigned_to_nat(0u); -lean_inc(x_5); -x_114 = l_Array_toSubarray___rarg(x_5, x_113, x_112); -x_115 = lean_ctor_get(x_114, 0); -lean_inc(x_115); -x_116 = lean_ctor_get(x_114, 2); +lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; lean_object* x_121; lean_object* x_122; +x_109 = lean_box(0); +x_110 = l_Array_back___rarg(x_109, x_6); +x_111 = lean_array_get_size(x_6); +x_112 = lean_unsigned_to_nat(1u); +x_113 = lean_nat_sub(x_111, x_112); +lean_dec(x_111); +x_114 = lean_unsigned_to_nat(0u); +lean_inc(x_6); +x_115 = l_Array_toSubarray___rarg(x_6, x_114, x_113); +x_116 = lean_ctor_get(x_115, 0); lean_inc(x_116); -x_117 = lean_ctor_get(x_114, 1); +x_117 = lean_ctor_get(x_115, 2); lean_inc(x_117); -lean_dec(x_114); -x_118 = lean_array_get_size(x_115); -x_119 = lean_nat_dec_le(x_116, x_118); -if (x_119 == 0) +x_118 = lean_ctor_get(x_115, 1); +lean_inc(x_118); +lean_dec(x_115); +x_119 = lean_array_get_size(x_116); +x_120 = lean_nat_dec_le(x_117, x_119); +if (x_120 == 0) { -uint8_t x_185; -lean_dec(x_116); -x_185 = lean_nat_dec_lt(x_117, x_118); -if (x_185 == 0) +uint8_t x_186; +lean_dec(x_117); +x_186 = lean_nat_dec_lt(x_118, x_119); +if (x_186 == 0) { +lean_dec(x_119); lean_dec(x_118); -lean_dec(x_117); -lean_dec(x_115); -x_120 = x_109; -x_121 = x_17; -goto block_184; +lean_dec(x_116); +x_121 = x_110; +x_122 = x_18; +goto block_185; } else { -size_t x_186; size_t x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; -x_186 = lean_usize_of_nat(x_118); +size_t x_187; size_t x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; +x_187 = lean_usize_of_nat(x_119); +lean_dec(x_119); +x_188 = lean_usize_of_nat(x_118); lean_dec(x_118); -x_187 = lean_usize_of_nat(x_117); -lean_dec(x_117); -lean_inc(x_15); -lean_inc(x_6); -x_188 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3(x_6, x_115, x_186, x_187, x_109, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -lean_dec(x_115); -x_189 = lean_ctor_get(x_188, 0); -lean_inc(x_189); -x_190 = lean_ctor_get(x_188, 1); +lean_inc(x_16); +lean_inc(x_5); +lean_inc(x_7); +x_189 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8(x_7, x_5, x_116, x_187, x_188, x_110, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_116); +x_190 = lean_ctor_get(x_189, 0); lean_inc(x_190); -lean_dec(x_188); -x_120 = x_189; +x_191 = lean_ctor_get(x_189, 1); +lean_inc(x_191); +lean_dec(x_189); x_121 = x_190; -goto block_184; +x_122 = x_191; +goto block_185; } } else { -uint8_t x_191; -lean_dec(x_118); -x_191 = lean_nat_dec_lt(x_117, x_116); -if (x_191 == 0) +uint8_t x_192; +lean_dec(x_119); +x_192 = lean_nat_dec_lt(x_118, x_117); +if (x_192 == 0) { +lean_dec(x_118); lean_dec(x_117); lean_dec(x_116); -lean_dec(x_115); -x_120 = x_109; -x_121 = x_17; -goto block_184; +x_121 = x_110; +x_122 = x_18; +goto block_185; } else { -size_t x_192; size_t x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; -x_192 = lean_usize_of_nat(x_116); -lean_dec(x_116); +size_t x_193; size_t x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; x_193 = lean_usize_of_nat(x_117); lean_dec(x_117); -lean_inc(x_15); -lean_inc(x_6); -x_194 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3(x_6, x_115, x_192, x_193, x_109, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -lean_dec(x_115); -x_195 = lean_ctor_get(x_194, 0); -lean_inc(x_195); -x_196 = lean_ctor_get(x_194, 1); +x_194 = lean_usize_of_nat(x_118); +lean_dec(x_118); +lean_inc(x_16); +lean_inc(x_5); +lean_inc(x_7); +x_195 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9(x_7, x_5, x_116, x_193, x_194, x_110, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_116); +x_196 = lean_ctor_get(x_195, 0); lean_inc(x_196); -lean_dec(x_194); -x_120 = x_195; +x_197 = lean_ctor_get(x_195, 1); +lean_inc(x_197); +lean_dec(x_195); x_121 = x_196; -goto block_184; +x_122 = x_197; +goto block_185; } } -block_184: +block_185: { -lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; -lean_inc(x_15); -x_122 = l_Lean_Elab_Term_Quotation_mkTuple(x_5, x_11, x_12, x_13, x_14, x_15, x_16, x_121); -x_123 = lean_ctor_get(x_122, 0); -lean_inc(x_123); -x_124 = lean_ctor_get(x_122, 1); +lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; +lean_inc(x_16); +x_123 = l_Lean_Elab_Term_Quotation_mkTuple(x_6, x_12, x_13, x_14, x_15, x_16, x_17, x_122); +x_124 = lean_ctor_get(x_123, 0); lean_inc(x_124); -lean_dec(x_122); -lean_inc(x_15); -x_125 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_15, x_16, x_124); -x_126 = lean_ctor_get(x_125, 0); -lean_inc(x_126); -x_127 = lean_ctor_get(x_125, 1); +x_125 = lean_ctor_get(x_123, 1); +lean_inc(x_125); +lean_dec(x_123); +lean_inc(x_16); +x_126 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_16, x_17, x_125); +x_127 = lean_ctor_get(x_126, 0); lean_inc(x_127); -lean_dec(x_125); -x_128 = lean_ctor_get(x_15, 10); +x_128 = lean_ctor_get(x_126, 1); lean_inc(x_128); -x_129 = lean_st_ref_get(x_16, x_127); -x_130 = lean_ctor_get(x_129, 0); -lean_inc(x_130); -x_131 = lean_ctor_get(x_129, 1); +lean_dec(x_126); +x_129 = lean_ctor_get(x_16, 10); +lean_inc(x_129); +x_130 = lean_st_ref_get(x_17, x_128); +x_131 = lean_ctor_get(x_130, 0); lean_inc(x_131); -lean_dec(x_129); -x_132 = lean_ctor_get(x_130, 0); +x_132 = lean_ctor_get(x_130, 1); lean_inc(x_132); lean_dec(x_130); -x_133 = lean_environment_main_module(x_132); -x_134 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__4; -x_135 = lean_name_mk_string(x_6, x_134); -lean_inc(x_135); -x_136 = l_Lean_addMacroScope(x_133, x_135, x_128); -x_137 = lean_box(0); -x_138 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_138, 0, x_135); -lean_ctor_set(x_138, 1, x_137); -x_139 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_139, 0, x_138); -lean_ctor_set(x_139, 1, x_137); -x_140 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__3; -lean_inc(x_126); -x_141 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_141, 0, x_126); -lean_ctor_set(x_141, 1, x_140); -lean_ctor_set(x_141, 2, x_136); -lean_ctor_set(x_141, 3, x_139); -x_142 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; -lean_inc(x_126); -x_143 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_143, 0, x_126); -lean_ctor_set(x_143, 1, x_142); -x_144 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__8; -lean_inc(x_126); -x_145 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_145, 0, x_126); -lean_ctor_set(x_145, 1, x_144); -x_146 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; -x_147 = lean_array_push(x_146, x_123); -x_148 = lean_box(2); -x_149 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_133 = lean_ctor_get(x_131, 0); +lean_inc(x_133); +lean_dec(x_131); +x_134 = lean_environment_main_module(x_133); +x_135 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__4; +x_136 = lean_name_mk_string(x_7, x_135); +lean_inc(x_136); +x_137 = l_Lean_addMacroScope(x_134, x_136, x_129); +x_138 = lean_box(0); +x_139 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_139, 0, x_136); +lean_ctor_set(x_139, 1, x_138); +x_140 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_140, 0, x_139); +lean_ctor_set(x_140, 1, x_138); +x_141 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__3; +lean_inc(x_127); +x_142 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_142, 0, x_127); +lean_ctor_set(x_142, 1, x_141); +lean_ctor_set(x_142, 2, x_137); +lean_ctor_set(x_142, 3, x_140); +x_143 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; +lean_inc(x_127); +x_144 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_144, 0, x_127); +lean_ctor_set(x_144, 1, x_143); +x_145 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__8; +lean_inc(x_127); +x_146 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_146, 0, x_127); +lean_ctor_set(x_146, 1, x_145); +x_147 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; +x_148 = lean_array_push(x_147, x_124); +x_149 = lean_box(2); +lean_inc(x_5); x_150 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_150, 0, x_148); -lean_ctor_set(x_150, 1, x_149); -lean_ctor_set(x_150, 2, x_147); -x_151 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__12; -lean_inc(x_126); +lean_ctor_set(x_150, 0, x_149); +lean_ctor_set(x_150, 1, x_5); +lean_ctor_set(x_150, 2, x_148); +x_151 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__12; +lean_inc(x_127); x_152 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_152, 0, x_126); +lean_ctor_set(x_152, 0, x_127); lean_ctor_set(x_152, 1, x_151); -x_153 = lean_array_get(x_108, x_7, x_113); -lean_dec(x_7); -x_154 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; +x_153 = lean_array_get(x_109, x_8, x_114); +lean_dec(x_8); +x_154 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; x_155 = lean_array_push(x_154, x_150); x_156 = lean_array_push(x_155, x_152); x_157 = lean_array_push(x_156, x_153); -x_158 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__11; +x_158 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__11; x_159 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_159, 0, x_148); +lean_ctor_set(x_159, 0, x_149); lean_ctor_set(x_159, 1, x_158); lean_ctor_set(x_159, 2, x_157); x_160 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_161 = lean_array_push(x_160, x_145); +x_161 = lean_array_push(x_160, x_146); x_162 = lean_array_push(x_161, x_159); -x_163 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__9; +x_163 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__9; x_164 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_164, 0, x_148); +lean_ctor_set(x_164, 0, x_149); lean_ctor_set(x_164, 1, x_163); lean_ctor_set(x_164, 2, x_162); -x_165 = lean_array_push(x_160, x_164); -x_166 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; -x_167 = lean_array_push(x_165, x_166); -x_168 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_168, 0, x_148); -lean_ctor_set(x_168, 1, x_149); -lean_ctor_set(x_168, 2, x_167); -x_169 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; -x_170 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_170, 0, x_126); -lean_ctor_set(x_170, 1, x_169); -x_171 = lean_array_push(x_154, x_143); -x_172 = lean_array_push(x_171, x_168); -x_173 = lean_array_push(x_172, x_170); -x_174 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__6; -x_175 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_175, 0, x_148); -lean_ctor_set(x_175, 1, x_174); -lean_ctor_set(x_175, 2, x_173); -x_176 = lean_array_push(x_160, x_175); -x_177 = lean_array_push(x_176, x_120); -x_178 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_178, 0, x_148); -lean_ctor_set(x_178, 1, x_149); -lean_ctor_set(x_178, 2, x_177); -x_179 = lean_array_push(x_160, x_141); -x_180 = lean_array_push(x_179, x_178); -x_181 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_182 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_182, 0, x_148); -lean_ctor_set(x_182, 1, x_181); -lean_ctor_set(x_182, 2, x_180); -x_183 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2(x_1, x_2, x_3, x_4, x_9, x_182, x_11, x_12, x_13, x_14, x_15, x_16, x_131); -return x_183; +x_165 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; +lean_inc(x_5); +x_166 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_166, 0, x_149); +lean_ctor_set(x_166, 1, x_5); +lean_ctor_set(x_166, 2, x_165); +x_167 = lean_array_push(x_160, x_164); +x_168 = lean_array_push(x_167, x_166); +lean_inc(x_5); +x_169 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_169, 0, x_149); +lean_ctor_set(x_169, 1, x_5); +lean_ctor_set(x_169, 2, x_168); +x_170 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; +x_171 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_171, 0, x_127); +lean_ctor_set(x_171, 1, x_170); +x_172 = lean_array_push(x_154, x_144); +x_173 = lean_array_push(x_172, x_169); +x_174 = lean_array_push(x_173, x_171); +x_175 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6; +x_176 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_176, 0, x_149); +lean_ctor_set(x_176, 1, x_175); +lean_ctor_set(x_176, 2, x_174); +x_177 = lean_array_push(x_160, x_176); +x_178 = lean_array_push(x_177, x_121); +lean_inc(x_5); +x_179 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_179, 0, x_149); +lean_ctor_set(x_179, 1, x_5); +lean_ctor_set(x_179, 2, x_178); +x_180 = lean_array_push(x_160, x_142); +x_181 = lean_array_push(x_180, x_179); +x_182 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_183 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_183, 0, x_149); +lean_ctor_set(x_183, 1, x_182); +lean_ctor_set(x_183, 2, x_181); +x_184 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2(x_1, x_2, x_3, x_4, x_5, x_10, x_183, x_12, x_13, x_14, x_15, x_16, x_17, x_132); +return x_184; } } case 1: { -lean_object* x_197; -x_197 = lean_ctor_get(x_107, 0); -lean_inc(x_197); -switch (lean_obj_tag(x_197)) { +lean_object* x_198; +x_198 = lean_ctor_get(x_108, 0); +lean_inc(x_198); +switch (lean_obj_tag(x_198)) { case 0: { -lean_object* x_198; lean_object* x_199; uint8_t x_200; -x_198 = lean_ctor_get(x_107, 1); -lean_inc(x_198); -lean_dec(x_107); -x_199 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__15; -x_200 = lean_string_dec_eq(x_198, x_199); -lean_dec(x_198); -if (x_200 == 0) -{ -lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; uint8_t x_212; lean_object* x_213; lean_object* x_214; -x_201 = l_Lean_instInhabitedSyntax; -x_202 = l_Array_back___rarg(x_201, x_5); -x_203 = lean_array_get_size(x_5); -x_204 = lean_unsigned_to_nat(1u); -x_205 = lean_nat_sub(x_203, x_204); -lean_dec(x_203); -x_206 = lean_unsigned_to_nat(0u); -lean_inc(x_5); -x_207 = l_Array_toSubarray___rarg(x_5, x_206, x_205); -x_208 = lean_ctor_get(x_207, 0); -lean_inc(x_208); -x_209 = lean_ctor_get(x_207, 2); +lean_object* x_199; lean_object* x_200; uint8_t x_201; +x_199 = lean_ctor_get(x_108, 1); +lean_inc(x_199); +lean_dec(x_108); +x_200 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__15; +x_201 = lean_string_dec_eq(x_199, x_200); +lean_dec(x_199); +if (x_201 == 0) +{ +lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; uint8_t x_213; lean_object* x_214; lean_object* x_215; +x_202 = lean_box(0); +x_203 = l_Array_back___rarg(x_202, x_6); +x_204 = lean_array_get_size(x_6); +x_205 = lean_unsigned_to_nat(1u); +x_206 = lean_nat_sub(x_204, x_205); +lean_dec(x_204); +x_207 = lean_unsigned_to_nat(0u); +lean_inc(x_6); +x_208 = l_Array_toSubarray___rarg(x_6, x_207, x_206); +x_209 = lean_ctor_get(x_208, 0); lean_inc(x_209); -x_210 = lean_ctor_get(x_207, 1); +x_210 = lean_ctor_get(x_208, 2); lean_inc(x_210); -lean_dec(x_207); -x_211 = lean_array_get_size(x_208); -x_212 = lean_nat_dec_le(x_209, x_211); -if (x_212 == 0) +x_211 = lean_ctor_get(x_208, 1); +lean_inc(x_211); +lean_dec(x_208); +x_212 = lean_array_get_size(x_209); +x_213 = lean_nat_dec_le(x_210, x_212); +if (x_213 == 0) { -uint8_t x_278; -lean_dec(x_209); -x_278 = lean_nat_dec_lt(x_210, x_211); -if (x_278 == 0) +uint8_t x_279; +lean_dec(x_210); +x_279 = lean_nat_dec_lt(x_211, x_212); +if (x_279 == 0) { +lean_dec(x_212); lean_dec(x_211); -lean_dec(x_210); -lean_dec(x_208); -x_213 = x_202; -x_214 = x_17; -goto block_277; +lean_dec(x_209); +x_214 = x_203; +x_215 = x_18; +goto block_278; } else { -size_t x_279; size_t x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; -x_279 = lean_usize_of_nat(x_211); +size_t x_280; size_t x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; +x_280 = lean_usize_of_nat(x_212); +lean_dec(x_212); +x_281 = lean_usize_of_nat(x_211); lean_dec(x_211); -x_280 = lean_usize_of_nat(x_210); -lean_dec(x_210); -lean_inc(x_15); -lean_inc(x_6); -x_281 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3(x_6, x_208, x_279, x_280, x_202, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -lean_dec(x_208); -x_282 = lean_ctor_get(x_281, 0); -lean_inc(x_282); -x_283 = lean_ctor_get(x_281, 1); +lean_inc(x_16); +lean_inc(x_5); +lean_inc(x_7); +x_282 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__10(x_7, x_5, x_209, x_280, x_281, x_203, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_209); +x_283 = lean_ctor_get(x_282, 0); lean_inc(x_283); -lean_dec(x_281); -x_213 = x_282; +x_284 = lean_ctor_get(x_282, 1); +lean_inc(x_284); +lean_dec(x_282); x_214 = x_283; -goto block_277; +x_215 = x_284; +goto block_278; } } else { -uint8_t x_284; -lean_dec(x_211); -x_284 = lean_nat_dec_lt(x_210, x_209); -if (x_284 == 0) +uint8_t x_285; +lean_dec(x_212); +x_285 = lean_nat_dec_lt(x_211, x_210); +if (x_285 == 0) { +lean_dec(x_211); lean_dec(x_210); lean_dec(x_209); -lean_dec(x_208); -x_213 = x_202; -x_214 = x_17; -goto block_277; +x_214 = x_203; +x_215 = x_18; +goto block_278; } else { -size_t x_285; size_t x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; -x_285 = lean_usize_of_nat(x_209); -lean_dec(x_209); +size_t x_286; size_t x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; x_286 = lean_usize_of_nat(x_210); lean_dec(x_210); -lean_inc(x_15); -lean_inc(x_6); -x_287 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3(x_6, x_208, x_285, x_286, x_202, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -lean_dec(x_208); -x_288 = lean_ctor_get(x_287, 0); -lean_inc(x_288); -x_289 = lean_ctor_get(x_287, 1); +x_287 = lean_usize_of_nat(x_211); +lean_dec(x_211); +lean_inc(x_16); +lean_inc(x_5); +lean_inc(x_7); +x_288 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__11(x_7, x_5, x_209, x_286, x_287, x_203, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_dec(x_209); +x_289 = lean_ctor_get(x_288, 0); lean_inc(x_289); -lean_dec(x_287); -x_213 = x_288; +x_290 = lean_ctor_get(x_288, 1); +lean_inc(x_290); +lean_dec(x_288); x_214 = x_289; -goto block_277; +x_215 = x_290; +goto block_278; } } -block_277: +block_278: { -lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; -lean_inc(x_15); -x_215 = l_Lean_Elab_Term_Quotation_mkTuple(x_5, x_11, x_12, x_13, x_14, x_15, x_16, x_214); -x_216 = lean_ctor_get(x_215, 0); -lean_inc(x_216); -x_217 = lean_ctor_get(x_215, 1); +lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; +lean_inc(x_16); +x_216 = l_Lean_Elab_Term_Quotation_mkTuple(x_6, x_12, x_13, x_14, x_15, x_16, x_17, x_215); +x_217 = lean_ctor_get(x_216, 0); lean_inc(x_217); -lean_dec(x_215); -lean_inc(x_15); -x_218 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_15, x_16, x_217); -x_219 = lean_ctor_get(x_218, 0); -lean_inc(x_219); -x_220 = lean_ctor_get(x_218, 1); +x_218 = lean_ctor_get(x_216, 1); +lean_inc(x_218); +lean_dec(x_216); +lean_inc(x_16); +x_219 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_16, x_17, x_218); +x_220 = lean_ctor_get(x_219, 0); lean_inc(x_220); -lean_dec(x_218); -x_221 = lean_ctor_get(x_15, 10); +x_221 = lean_ctor_get(x_219, 1); lean_inc(x_221); -x_222 = lean_st_ref_get(x_16, x_220); -x_223 = lean_ctor_get(x_222, 0); -lean_inc(x_223); -x_224 = lean_ctor_get(x_222, 1); +lean_dec(x_219); +x_222 = lean_ctor_get(x_16, 10); +lean_inc(x_222); +x_223 = lean_st_ref_get(x_17, x_221); +x_224 = lean_ctor_get(x_223, 0); lean_inc(x_224); -lean_dec(x_222); -x_225 = lean_ctor_get(x_223, 0); +x_225 = lean_ctor_get(x_223, 1); lean_inc(x_225); lean_dec(x_223); -x_226 = lean_environment_main_module(x_225); -x_227 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__4; -x_228 = lean_name_mk_string(x_6, x_227); -lean_inc(x_228); -x_229 = l_Lean_addMacroScope(x_226, x_228, x_221); -x_230 = lean_box(0); -x_231 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_231, 0, x_228); -lean_ctor_set(x_231, 1, x_230); -x_232 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_232, 0, x_231); -lean_ctor_set(x_232, 1, x_230); -x_233 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__3; -lean_inc(x_219); -x_234 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_234, 0, x_219); -lean_ctor_set(x_234, 1, x_233); -lean_ctor_set(x_234, 2, x_229); -lean_ctor_set(x_234, 3, x_232); -x_235 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; -lean_inc(x_219); -x_236 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_236, 0, x_219); -lean_ctor_set(x_236, 1, x_235); -x_237 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__8; -lean_inc(x_219); -x_238 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_238, 0, x_219); -lean_ctor_set(x_238, 1, x_237); -x_239 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; -x_240 = lean_array_push(x_239, x_216); -x_241 = lean_box(2); -x_242 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_226 = lean_ctor_get(x_224, 0); +lean_inc(x_226); +lean_dec(x_224); +x_227 = lean_environment_main_module(x_226); +x_228 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__4; +x_229 = lean_name_mk_string(x_7, x_228); +lean_inc(x_229); +x_230 = l_Lean_addMacroScope(x_227, x_229, x_222); +x_231 = lean_box(0); +x_232 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_232, 0, x_229); +lean_ctor_set(x_232, 1, x_231); +x_233 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_233, 0, x_232); +lean_ctor_set(x_233, 1, x_231); +x_234 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__3; +lean_inc(x_220); +x_235 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_235, 0, x_220); +lean_ctor_set(x_235, 1, x_234); +lean_ctor_set(x_235, 2, x_230); +lean_ctor_set(x_235, 3, x_233); +x_236 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; +lean_inc(x_220); +x_237 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_237, 0, x_220); +lean_ctor_set(x_237, 1, x_236); +x_238 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__8; +lean_inc(x_220); +x_239 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_239, 0, x_220); +lean_ctor_set(x_239, 1, x_238); +x_240 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; +x_241 = lean_array_push(x_240, x_217); +x_242 = lean_box(2); +lean_inc(x_5); x_243 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_243, 0, x_241); -lean_ctor_set(x_243, 1, x_242); -lean_ctor_set(x_243, 2, x_240); -x_244 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__12; -lean_inc(x_219); +lean_ctor_set(x_243, 0, x_242); +lean_ctor_set(x_243, 1, x_5); +lean_ctor_set(x_243, 2, x_241); +x_244 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__12; +lean_inc(x_220); x_245 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_245, 0, x_219); +lean_ctor_set(x_245, 0, x_220); lean_ctor_set(x_245, 1, x_244); -x_246 = lean_array_get(x_201, x_7, x_206); -lean_dec(x_7); -x_247 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; +x_246 = lean_array_get(x_202, x_8, x_207); +lean_dec(x_8); +x_247 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; x_248 = lean_array_push(x_247, x_243); x_249 = lean_array_push(x_248, x_245); x_250 = lean_array_push(x_249, x_246); -x_251 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__11; +x_251 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__11; x_252 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_252, 0, x_241); +lean_ctor_set(x_252, 0, x_242); lean_ctor_set(x_252, 1, x_251); lean_ctor_set(x_252, 2, x_250); x_253 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_254 = lean_array_push(x_253, x_238); +x_254 = lean_array_push(x_253, x_239); x_255 = lean_array_push(x_254, x_252); -x_256 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__9; +x_256 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__9; x_257 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_257, 0, x_241); +lean_ctor_set(x_257, 0, x_242); lean_ctor_set(x_257, 1, x_256); lean_ctor_set(x_257, 2, x_255); -x_258 = lean_array_push(x_253, x_257); -x_259 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; -x_260 = lean_array_push(x_258, x_259); -x_261 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_261, 0, x_241); -lean_ctor_set(x_261, 1, x_242); -lean_ctor_set(x_261, 2, x_260); -x_262 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; -x_263 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_263, 0, x_219); -lean_ctor_set(x_263, 1, x_262); -x_264 = lean_array_push(x_247, x_236); -x_265 = lean_array_push(x_264, x_261); -x_266 = lean_array_push(x_265, x_263); -x_267 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__6; -x_268 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_268, 0, x_241); -lean_ctor_set(x_268, 1, x_267); -lean_ctor_set(x_268, 2, x_266); -x_269 = lean_array_push(x_253, x_268); -x_270 = lean_array_push(x_269, x_213); -x_271 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_271, 0, x_241); -lean_ctor_set(x_271, 1, x_242); -lean_ctor_set(x_271, 2, x_270); -x_272 = lean_array_push(x_253, x_234); -x_273 = lean_array_push(x_272, x_271); -x_274 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_275 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_275, 0, x_241); -lean_ctor_set(x_275, 1, x_274); -lean_ctor_set(x_275, 2, x_273); -x_276 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2(x_1, x_2, x_3, x_4, x_9, x_275, x_11, x_12, x_13, x_14, x_15, x_16, x_224); -return x_276; +x_258 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; +lean_inc(x_5); +x_259 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_259, 0, x_242); +lean_ctor_set(x_259, 1, x_5); +lean_ctor_set(x_259, 2, x_258); +x_260 = lean_array_push(x_253, x_257); +x_261 = lean_array_push(x_260, x_259); +lean_inc(x_5); +x_262 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_262, 0, x_242); +lean_ctor_set(x_262, 1, x_5); +lean_ctor_set(x_262, 2, x_261); +x_263 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; +x_264 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_264, 0, x_220); +lean_ctor_set(x_264, 1, x_263); +x_265 = lean_array_push(x_247, x_237); +x_266 = lean_array_push(x_265, x_262); +x_267 = lean_array_push(x_266, x_264); +x_268 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6; +x_269 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_269, 0, x_242); +lean_ctor_set(x_269, 1, x_268); +lean_ctor_set(x_269, 2, x_267); +x_270 = lean_array_push(x_253, x_269); +x_271 = lean_array_push(x_270, x_214); +lean_inc(x_5); +x_272 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_272, 0, x_242); +lean_ctor_set(x_272, 1, x_5); +lean_ctor_set(x_272, 2, x_271); +x_273 = lean_array_push(x_253, x_235); +x_274 = lean_array_push(x_273, x_272); +x_275 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_276 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_276, 0, x_242); +lean_ctor_set(x_276, 1, x_275); +lean_ctor_set(x_276, 2, x_274); +x_277 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2(x_1, x_2, x_3, x_4, x_5, x_10, x_276, x_12, x_13, x_14, x_15, x_16, x_17, x_225); +return x_277; } } else { -lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; size_t x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; -lean_inc(x_15); -x_290 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_15, x_16, x_17); -x_291 = lean_ctor_get(x_290, 0); -lean_inc(x_291); -x_292 = lean_ctor_get(x_290, 1); +lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; size_t x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; +lean_inc(x_16); +x_291 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_16, x_17, x_18); +x_292 = lean_ctor_get(x_291, 0); lean_inc(x_292); -lean_dec(x_290); -x_293 = lean_ctor_get(x_15, 10); +x_293 = lean_ctor_get(x_291, 1); lean_inc(x_293); -x_294 = lean_st_ref_get(x_16, x_292); -x_295 = lean_ctor_get(x_294, 0); -lean_inc(x_295); -x_296 = lean_ctor_get(x_294, 1); +lean_dec(x_291); +x_294 = lean_ctor_get(x_16, 10); +lean_inc(x_294); +x_295 = lean_st_ref_get(x_17, x_293); +x_296 = lean_ctor_get(x_295, 0); lean_inc(x_296); -lean_dec(x_294); -x_297 = lean_ctor_get(x_295, 0); +x_297 = lean_ctor_get(x_295, 1); lean_inc(x_297); lean_dec(x_295); -x_298 = lean_environment_main_module(x_297); -x_299 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__16; -lean_inc(x_291); -x_300 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_300, 0, x_291); -lean_ctor_set(x_300, 1, x_299); -x_301 = lean_array_get_size(x_5); -x_302 = lean_usize_of_nat(x_301); -lean_dec(x_301); -x_303 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_304 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; +x_298 = lean_ctor_get(x_296, 0); +lean_inc(x_298); +lean_dec(x_296); +x_299 = lean_environment_main_module(x_298); +x_300 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__16; +lean_inc(x_292); +x_301 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_301, 0, x_292); +lean_ctor_set(x_301, 1, x_300); +x_302 = lean_box(2); +x_303 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; lean_inc(x_5); -x_305 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4(x_303, x_304, x_302, x_8, x_5); -x_306 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__19; -x_307 = l_Lean_mkSepArray(x_305, x_306); +x_304 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_304, 0, x_302); +lean_ctor_set(x_304, 1, x_5); +lean_ctor_set(x_304, 2, x_303); +x_305 = lean_array_get_size(x_6); +x_306 = lean_usize_of_nat(x_305); lean_dec(x_305); -x_308 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; -x_309 = l_Array_append___rarg(x_308, x_307); -x_310 = lean_box(2); -x_311 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_307 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; +lean_inc(x_6); +lean_inc(x_304); +x_308 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__12(x_307, x_304, x_306, x_9, x_6); +x_309 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__19; +x_310 = l_Lean_mkSepArray(x_308, x_309); +lean_dec(x_308); +x_311 = l_Array_append___rarg(x_303, x_310); +lean_inc(x_5); x_312 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_312, 0, x_310); -lean_ctor_set(x_312, 1, x_311); -lean_ctor_set(x_312, 2, x_309); -x_313 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__20; -lean_inc(x_291); +lean_ctor_set(x_312, 0, x_302); +lean_ctor_set(x_312, 1, x_5); +lean_ctor_set(x_312, 2, x_311); +x_313 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__20; +lean_inc(x_292); x_314 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_314, 0, x_291); +lean_ctor_set(x_314, 0, x_292); lean_ctor_set(x_314, 1, x_313); -x_315 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__25; -lean_inc(x_291); +x_315 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__25; +lean_inc(x_292); x_316 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_316, 0, x_291); +lean_ctor_set(x_316, 0, x_292); lean_ctor_set(x_316, 1, x_315); +lean_inc(x_6); +lean_inc(x_299); +lean_inc(x_294); +lean_inc(x_292); lean_inc(x_5); -lean_inc(x_298); -lean_inc(x_293); -lean_inc(x_291); -x_317 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5(x_291, x_293, x_298, x_303, x_311, x_302, x_8, x_5); -x_318 = l_Lean_mkSepArray(x_317, x_306); +x_317 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13(x_5, x_292, x_294, x_299, x_307, x_306, x_9, x_6); +x_318 = l_Lean_mkSepArray(x_317, x_309); lean_dec(x_317); -x_319 = l_Array_append___rarg(x_308, x_318); +x_319 = l_Array_append___rarg(x_303, x_318); +lean_inc(x_5); x_320 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_320, 0, x_310); -lean_ctor_set(x_320, 1, x_311); +lean_ctor_set(x_320, 0, x_302); +lean_ctor_set(x_320, 1, x_5); lean_ctor_set(x_320, 2, x_319); x_321 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; x_322 = lean_array_push(x_321, x_320); +lean_inc(x_5); x_323 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_323, 0, x_310); -lean_ctor_set(x_323, 1, x_311); +lean_ctor_set(x_323, 0, x_302); +lean_ctor_set(x_323, 1, x_5); lean_ctor_set(x_323, 2, x_322); -x_324 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__12; -lean_inc(x_291); +x_324 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__12; +lean_inc(x_292); x_325 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_325, 0, x_291); +lean_ctor_set(x_325, 0, x_292); lean_ctor_set(x_325, 1, x_324); -x_326 = lean_array_to_list(lean_box(0), x_7); +x_326 = lean_array_to_list(lean_box(0), x_8); x_327 = l___private_Init_Meta_0__Lean_quoteList___at_Lean_Elab_Term_Quotation_ArrayStxBuilder_build___spec__1(x_326); x_328 = lean_array_push(x_321, x_327); x_329 = l_Lean_Elab_Term_Quotation_ArrayStxBuilder_build___closed__2; @@ -5519,29 +6755,31 @@ x_333 = lean_array_push(x_332, x_323); lean_inc(x_325); x_334 = lean_array_push(x_333, x_325); x_335 = lean_array_push(x_334, x_330); -x_336 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__24; +x_336 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__24; x_337 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_337, 0, x_310); +lean_ctor_set(x_337, 0, x_302); lean_ctor_set(x_337, 1, x_336); lean_ctor_set(x_337, 2, x_335); -lean_inc(x_291); -x_338 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6(x_291, x_303, x_321, x_302, x_8, x_5); -x_339 = l_Lean_mkSepArray(x_338, x_306); +lean_inc(x_292); +x_338 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__14(x_292, x_307, x_321, x_306, x_9, x_6); +x_339 = l_Lean_mkSepArray(x_338, x_309); lean_dec(x_338); -x_340 = l_Array_append___rarg(x_308, x_339); +x_340 = l_Array_append___rarg(x_303, x_339); +lean_inc(x_5); x_341 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_341, 0, x_310); -lean_ctor_set(x_341, 1, x_311); +lean_ctor_set(x_341, 0, x_302); +lean_ctor_set(x_341, 1, x_5); lean_ctor_set(x_341, 2, x_340); x_342 = lean_array_push(x_321, x_341); +lean_inc(x_5); x_343 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_343, 0, x_310); -lean_ctor_set(x_343, 1, x_311); +lean_ctor_set(x_343, 0, x_302); +lean_ctor_set(x_343, 1, x_5); lean_ctor_set(x_343, 2, x_342); -x_344 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__29; -x_345 = lean_name_mk_string(x_6, x_344); +x_344 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__29; +x_345 = lean_name_mk_string(x_7, x_344); lean_inc(x_345); -x_346 = l_Lean_addMacroScope(x_298, x_345, x_293); +x_346 = l_Lean_addMacroScope(x_299, x_345, x_294); x_347 = lean_box(0); x_348 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_348, 0, x_345); @@ -5549,9 +6787,9 @@ lean_ctor_set(x_348, 1, x_347); x_349 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_349, 0, x_348); lean_ctor_set(x_349, 1, x_347); -x_350 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__28; +x_350 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__28; x_351 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_351, 0, x_291); +lean_ctor_set(x_351, 0, x_292); lean_ctor_set(x_351, 1, x_350); lean_ctor_set(x_351, 2, x_346); lean_ctor_set(x_351, 3, x_349); @@ -5559,52 +6797,54 @@ x_352 = lean_array_push(x_332, x_343); x_353 = lean_array_push(x_352, x_325); x_354 = lean_array_push(x_353, x_351); x_355 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_355, 0, x_310); +lean_ctor_set(x_355, 0, x_302); lean_ctor_set(x_355, 1, x_336); lean_ctor_set(x_355, 2, x_354); x_356 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; x_357 = lean_array_push(x_356, x_337); x_358 = lean_array_push(x_357, x_355); +lean_inc(x_5); x_359 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_359, 0, x_310); -lean_ctor_set(x_359, 1, x_311); +lean_ctor_set(x_359, 0, x_302); +lean_ctor_set(x_359, 1, x_5); lean_ctor_set(x_359, 2, x_358); x_360 = lean_array_push(x_321, x_359); -x_361 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__22; +x_361 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__22; x_362 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_362, 0, x_310); +lean_ctor_set(x_362, 0, x_302); lean_ctor_set(x_362, 1, x_361); lean_ctor_set(x_362, 2, x_360); -x_363 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__30; -x_364 = lean_array_push(x_363, x_300); +x_363 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__30; +x_364 = lean_array_push(x_363, x_301); +lean_inc(x_304); x_365 = lean_array_push(x_364, x_304); x_366 = lean_array_push(x_365, x_304); x_367 = lean_array_push(x_366, x_312); x_368 = lean_array_push(x_367, x_314); x_369 = lean_array_push(x_368, x_362); -x_370 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__17; +x_370 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__17; x_371 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_371, 0, x_310); +lean_ctor_set(x_371, 0, x_302); lean_ctor_set(x_371, 1, x_370); lean_ctor_set(x_371, 2, x_369); -x_372 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2(x_1, x_2, x_3, x_4, x_9, x_371, x_11, x_12, x_13, x_14, x_15, x_16, x_296); +x_372 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2(x_1, x_2, x_3, x_4, x_5, x_10, x_371, x_12, x_13, x_14, x_15, x_16, x_17, x_297); return x_372; } } case 1: { lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; uint8_t x_384; lean_object* x_385; lean_object* x_386; -lean_dec(x_197); -lean_dec(x_107); -x_373 = l_Lean_instInhabitedSyntax; -x_374 = l_Array_back___rarg(x_373, x_5); -x_375 = lean_array_get_size(x_5); +lean_dec(x_198); +lean_dec(x_108); +x_373 = lean_box(0); +x_374 = l_Array_back___rarg(x_373, x_6); +x_375 = lean_array_get_size(x_6); x_376 = lean_unsigned_to_nat(1u); x_377 = lean_nat_sub(x_375, x_376); lean_dec(x_375); x_378 = lean_unsigned_to_nat(0u); -lean_inc(x_5); -x_379 = l_Array_toSubarray___rarg(x_5, x_378, x_377); +lean_inc(x_6); +x_379 = l_Array_toSubarray___rarg(x_6, x_378, x_377); x_380 = lean_ctor_get(x_379, 0); lean_inc(x_380); x_381 = lean_ctor_get(x_379, 2); @@ -5625,7 +6865,7 @@ lean_dec(x_383); lean_dec(x_382); lean_dec(x_380); x_385 = x_374; -x_386 = x_17; +x_386 = x_18; goto block_449; } else @@ -5635,9 +6875,10 @@ x_451 = lean_usize_of_nat(x_383); lean_dec(x_383); x_452 = lean_usize_of_nat(x_382); lean_dec(x_382); -lean_inc(x_15); -lean_inc(x_6); -x_453 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3(x_6, x_380, x_451, x_452, x_374, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_inc(x_16); +lean_inc(x_5); +lean_inc(x_7); +x_453 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__15(x_7, x_5, x_380, x_451, x_452, x_374, x_12, x_13, x_14, x_15, x_16, x_17, x_18); lean_dec(x_380); x_454 = lean_ctor_get(x_453, 0); lean_inc(x_454); @@ -5660,7 +6901,7 @@ lean_dec(x_382); lean_dec(x_381); lean_dec(x_380); x_385 = x_374; -x_386 = x_17; +x_386 = x_18; goto block_449; } else @@ -5670,9 +6911,10 @@ x_457 = lean_usize_of_nat(x_381); lean_dec(x_381); x_458 = lean_usize_of_nat(x_382); lean_dec(x_382); -lean_inc(x_15); -lean_inc(x_6); -x_459 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3(x_6, x_380, x_457, x_458, x_374, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_inc(x_16); +lean_inc(x_5); +lean_inc(x_7); +x_459 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__16(x_7, x_5, x_380, x_457, x_458, x_374, x_12, x_13, x_14, x_15, x_16, x_17, x_18); lean_dec(x_380); x_460 = lean_ctor_get(x_459, 0); lean_inc(x_460); @@ -5687,23 +6929,23 @@ goto block_449; block_449: { lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; -lean_inc(x_15); -x_387 = l_Lean_Elab_Term_Quotation_mkTuple(x_5, x_11, x_12, x_13, x_14, x_15, x_16, x_386); +lean_inc(x_16); +x_387 = l_Lean_Elab_Term_Quotation_mkTuple(x_6, x_12, x_13, x_14, x_15, x_16, x_17, x_386); x_388 = lean_ctor_get(x_387, 0); lean_inc(x_388); x_389 = lean_ctor_get(x_387, 1); lean_inc(x_389); lean_dec(x_387); -lean_inc(x_15); -x_390 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_15, x_16, x_389); +lean_inc(x_16); +x_390 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_16, x_17, x_389); x_391 = lean_ctor_get(x_390, 0); lean_inc(x_391); x_392 = lean_ctor_get(x_390, 1); lean_inc(x_392); lean_dec(x_390); -x_393 = lean_ctor_get(x_15, 10); +x_393 = lean_ctor_get(x_16, 10); lean_inc(x_393); -x_394 = lean_st_ref_get(x_16, x_392); +x_394 = lean_st_ref_get(x_17, x_392); x_395 = lean_ctor_get(x_394, 0); lean_inc(x_395); x_396 = lean_ctor_get(x_394, 1); @@ -5713,8 +6955,8 @@ x_397 = lean_ctor_get(x_395, 0); lean_inc(x_397); lean_dec(x_395); x_398 = lean_environment_main_module(x_397); -x_399 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__4; -x_400 = lean_name_mk_string(x_6, x_399); +x_399 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__4; +x_400 = lean_name_mk_string(x_7, x_399); lean_inc(x_400); x_401 = l_Lean_addMacroScope(x_398, x_400, x_393); x_402 = lean_box(0); @@ -5724,19 +6966,19 @@ lean_ctor_set(x_403, 1, x_402); x_404 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_404, 0, x_403); lean_ctor_set(x_404, 1, x_402); -x_405 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__3; +x_405 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__3; lean_inc(x_391); x_406 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_406, 0, x_391); lean_ctor_set(x_406, 1, x_405); lean_ctor_set(x_406, 2, x_401); lean_ctor_set(x_406, 3, x_404); -x_407 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; +x_407 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; lean_inc(x_391); x_408 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_408, 0, x_391); lean_ctor_set(x_408, 1, x_407); -x_409 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__8; +x_409 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__8; lean_inc(x_391); x_410 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_410, 0, x_391); @@ -5744,85 +6986,92 @@ lean_ctor_set(x_410, 1, x_409); x_411 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; x_412 = lean_array_push(x_411, x_388); x_413 = lean_box(2); -x_414 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_415 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_415, 0, x_413); -lean_ctor_set(x_415, 1, x_414); -lean_ctor_set(x_415, 2, x_412); -x_416 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__12; +lean_inc(x_5); +x_414 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_414, 0, x_413); +lean_ctor_set(x_414, 1, x_5); +lean_ctor_set(x_414, 2, x_412); +x_415 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__12; lean_inc(x_391); -x_417 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_417, 0, x_391); -lean_ctor_set(x_417, 1, x_416); -x_418 = lean_array_get(x_373, x_7, x_378); -lean_dec(x_7); -x_419 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; -x_420 = lean_array_push(x_419, x_415); +x_416 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_416, 0, x_391); +lean_ctor_set(x_416, 1, x_415); +x_417 = lean_array_get(x_373, x_8, x_378); +lean_dec(x_8); +x_418 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; +x_419 = lean_array_push(x_418, x_414); +x_420 = lean_array_push(x_419, x_416); x_421 = lean_array_push(x_420, x_417); -x_422 = lean_array_push(x_421, x_418); -x_423 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__11; -x_424 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_424, 0, x_413); -lean_ctor_set(x_424, 1, x_423); -lean_ctor_set(x_424, 2, x_422); -x_425 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_426 = lean_array_push(x_425, x_410); -x_427 = lean_array_push(x_426, x_424); -x_428 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__9; -x_429 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_429, 0, x_413); -lean_ctor_set(x_429, 1, x_428); -lean_ctor_set(x_429, 2, x_427); -x_430 = lean_array_push(x_425, x_429); -x_431 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; -x_432 = lean_array_push(x_430, x_431); +x_422 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__11; +x_423 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_423, 0, x_413); +lean_ctor_set(x_423, 1, x_422); +lean_ctor_set(x_423, 2, x_421); +x_424 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_425 = lean_array_push(x_424, x_410); +x_426 = lean_array_push(x_425, x_423); +x_427 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__9; +x_428 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_428, 0, x_413); +lean_ctor_set(x_428, 1, x_427); +lean_ctor_set(x_428, 2, x_426); +x_429 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; +lean_inc(x_5); +x_430 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_430, 0, x_413); +lean_ctor_set(x_430, 1, x_5); +lean_ctor_set(x_430, 2, x_429); +x_431 = lean_array_push(x_424, x_428); +x_432 = lean_array_push(x_431, x_430); +lean_inc(x_5); x_433 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_433, 0, x_413); -lean_ctor_set(x_433, 1, x_414); +lean_ctor_set(x_433, 1, x_5); lean_ctor_set(x_433, 2, x_432); -x_434 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; +x_434 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; x_435 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_435, 0, x_391); lean_ctor_set(x_435, 1, x_434); -x_436 = lean_array_push(x_419, x_408); +x_436 = lean_array_push(x_418, x_408); x_437 = lean_array_push(x_436, x_433); x_438 = lean_array_push(x_437, x_435); -x_439 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__6; +x_439 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6; x_440 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_440, 0, x_413); lean_ctor_set(x_440, 1, x_439); lean_ctor_set(x_440, 2, x_438); -x_441 = lean_array_push(x_425, x_440); +x_441 = lean_array_push(x_424, x_440); x_442 = lean_array_push(x_441, x_385); +lean_inc(x_5); x_443 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_443, 0, x_413); -lean_ctor_set(x_443, 1, x_414); +lean_ctor_set(x_443, 1, x_5); lean_ctor_set(x_443, 2, x_442); -x_444 = lean_array_push(x_425, x_406); +x_444 = lean_array_push(x_424, x_406); x_445 = lean_array_push(x_444, x_443); x_446 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; x_447 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_447, 0, x_413); lean_ctor_set(x_447, 1, x_446); lean_ctor_set(x_447, 2, x_445); -x_448 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2(x_1, x_2, x_3, x_4, x_9, x_447, x_11, x_12, x_13, x_14, x_15, x_16, x_396); +x_448 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2(x_1, x_2, x_3, x_4, x_5, x_10, x_447, x_12, x_13, x_14, x_15, x_16, x_17, x_396); return x_448; } } default: { lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; uint8_t x_473; lean_object* x_474; lean_object* x_475; -lean_dec(x_197); -lean_dec(x_107); -x_462 = l_Lean_instInhabitedSyntax; -x_463 = l_Array_back___rarg(x_462, x_5); -x_464 = lean_array_get_size(x_5); +lean_dec(x_198); +lean_dec(x_108); +x_462 = lean_box(0); +x_463 = l_Array_back___rarg(x_462, x_6); +x_464 = lean_array_get_size(x_6); x_465 = lean_unsigned_to_nat(1u); x_466 = lean_nat_sub(x_464, x_465); lean_dec(x_464); x_467 = lean_unsigned_to_nat(0u); -lean_inc(x_5); -x_468 = l_Array_toSubarray___rarg(x_5, x_467, x_466); +lean_inc(x_6); +x_468 = l_Array_toSubarray___rarg(x_6, x_467, x_466); x_469 = lean_ctor_get(x_468, 0); lean_inc(x_469); x_470 = lean_ctor_get(x_468, 2); @@ -5843,7 +7092,7 @@ lean_dec(x_472); lean_dec(x_471); lean_dec(x_469); x_474 = x_463; -x_475 = x_17; +x_475 = x_18; goto block_538; } else @@ -5853,9 +7102,10 @@ x_540 = lean_usize_of_nat(x_472); lean_dec(x_472); x_541 = lean_usize_of_nat(x_471); lean_dec(x_471); -lean_inc(x_15); -lean_inc(x_6); -x_542 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3(x_6, x_469, x_540, x_541, x_463, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_inc(x_16); +lean_inc(x_5); +lean_inc(x_7); +x_542 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__17(x_7, x_5, x_469, x_540, x_541, x_463, x_12, x_13, x_14, x_15, x_16, x_17, x_18); lean_dec(x_469); x_543 = lean_ctor_get(x_542, 0); lean_inc(x_543); @@ -5878,7 +7128,7 @@ lean_dec(x_471); lean_dec(x_470); lean_dec(x_469); x_474 = x_463; -x_475 = x_17; +x_475 = x_18; goto block_538; } else @@ -5888,9 +7138,10 @@ x_546 = lean_usize_of_nat(x_470); lean_dec(x_470); x_547 = lean_usize_of_nat(x_471); lean_dec(x_471); -lean_inc(x_15); -lean_inc(x_6); -x_548 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3(x_6, x_469, x_546, x_547, x_463, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_inc(x_16); +lean_inc(x_5); +lean_inc(x_7); +x_548 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__18(x_7, x_5, x_469, x_546, x_547, x_463, x_12, x_13, x_14, x_15, x_16, x_17, x_18); lean_dec(x_469); x_549 = lean_ctor_get(x_548, 0); lean_inc(x_549); @@ -5905,23 +7156,23 @@ goto block_538; block_538: { lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; -lean_inc(x_15); -x_476 = l_Lean_Elab_Term_Quotation_mkTuple(x_5, x_11, x_12, x_13, x_14, x_15, x_16, x_475); +lean_inc(x_16); +x_476 = l_Lean_Elab_Term_Quotation_mkTuple(x_6, x_12, x_13, x_14, x_15, x_16, x_17, x_475); x_477 = lean_ctor_get(x_476, 0); lean_inc(x_477); x_478 = lean_ctor_get(x_476, 1); lean_inc(x_478); lean_dec(x_476); -lean_inc(x_15); -x_479 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_15, x_16, x_478); +lean_inc(x_16); +x_479 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_16, x_17, x_478); x_480 = lean_ctor_get(x_479, 0); lean_inc(x_480); x_481 = lean_ctor_get(x_479, 1); lean_inc(x_481); lean_dec(x_479); -x_482 = lean_ctor_get(x_15, 10); +x_482 = lean_ctor_get(x_16, 10); lean_inc(x_482); -x_483 = lean_st_ref_get(x_16, x_481); +x_483 = lean_st_ref_get(x_17, x_481); x_484 = lean_ctor_get(x_483, 0); lean_inc(x_484); x_485 = lean_ctor_get(x_483, 1); @@ -5931,8 +7182,8 @@ x_486 = lean_ctor_get(x_484, 0); lean_inc(x_486); lean_dec(x_484); x_487 = lean_environment_main_module(x_486); -x_488 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__4; -x_489 = lean_name_mk_string(x_6, x_488); +x_488 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__4; +x_489 = lean_name_mk_string(x_7, x_488); lean_inc(x_489); x_490 = l_Lean_addMacroScope(x_487, x_489, x_482); x_491 = lean_box(0); @@ -5942,19 +7193,19 @@ lean_ctor_set(x_492, 1, x_491); x_493 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_493, 0, x_492); lean_ctor_set(x_493, 1, x_491); -x_494 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__3; +x_494 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__3; lean_inc(x_480); x_495 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_495, 0, x_480); lean_ctor_set(x_495, 1, x_494); lean_ctor_set(x_495, 2, x_490); lean_ctor_set(x_495, 3, x_493); -x_496 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; +x_496 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; lean_inc(x_480); x_497 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_497, 0, x_480); lean_ctor_set(x_497, 1, x_496); -x_498 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__8; +x_498 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__8; lean_inc(x_480); x_499 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_499, 0, x_480); @@ -5962,68 +7213,75 @@ lean_ctor_set(x_499, 1, x_498); x_500 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; x_501 = lean_array_push(x_500, x_477); x_502 = lean_box(2); -x_503 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_504 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_504, 0, x_502); -lean_ctor_set(x_504, 1, x_503); -lean_ctor_set(x_504, 2, x_501); -x_505 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__12; +lean_inc(x_5); +x_503 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_503, 0, x_502); +lean_ctor_set(x_503, 1, x_5); +lean_ctor_set(x_503, 2, x_501); +x_504 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__12; lean_inc(x_480); -x_506 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_506, 0, x_480); -lean_ctor_set(x_506, 1, x_505); -x_507 = lean_array_get(x_462, x_7, x_467); -lean_dec(x_7); -x_508 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; -x_509 = lean_array_push(x_508, x_504); +x_505 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_505, 0, x_480); +lean_ctor_set(x_505, 1, x_504); +x_506 = lean_array_get(x_462, x_8, x_467); +lean_dec(x_8); +x_507 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; +x_508 = lean_array_push(x_507, x_503); +x_509 = lean_array_push(x_508, x_505); x_510 = lean_array_push(x_509, x_506); -x_511 = lean_array_push(x_510, x_507); -x_512 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__11; -x_513 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_513, 0, x_502); -lean_ctor_set(x_513, 1, x_512); -lean_ctor_set(x_513, 2, x_511); -x_514 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_515 = lean_array_push(x_514, x_499); -x_516 = lean_array_push(x_515, x_513); -x_517 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__9; -x_518 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_518, 0, x_502); -lean_ctor_set(x_518, 1, x_517); -lean_ctor_set(x_518, 2, x_516); -x_519 = lean_array_push(x_514, x_518); -x_520 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; -x_521 = lean_array_push(x_519, x_520); +x_511 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__11; +x_512 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_512, 0, x_502); +lean_ctor_set(x_512, 1, x_511); +lean_ctor_set(x_512, 2, x_510); +x_513 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_514 = lean_array_push(x_513, x_499); +x_515 = lean_array_push(x_514, x_512); +x_516 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__9; +x_517 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_517, 0, x_502); +lean_ctor_set(x_517, 1, x_516); +lean_ctor_set(x_517, 2, x_515); +x_518 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; +lean_inc(x_5); +x_519 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_519, 0, x_502); +lean_ctor_set(x_519, 1, x_5); +lean_ctor_set(x_519, 2, x_518); +x_520 = lean_array_push(x_513, x_517); +x_521 = lean_array_push(x_520, x_519); +lean_inc(x_5); x_522 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_522, 0, x_502); -lean_ctor_set(x_522, 1, x_503); +lean_ctor_set(x_522, 1, x_5); lean_ctor_set(x_522, 2, x_521); -x_523 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; +x_523 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; x_524 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_524, 0, x_480); lean_ctor_set(x_524, 1, x_523); -x_525 = lean_array_push(x_508, x_497); +x_525 = lean_array_push(x_507, x_497); x_526 = lean_array_push(x_525, x_522); x_527 = lean_array_push(x_526, x_524); -x_528 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__6; +x_528 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6; x_529 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_529, 0, x_502); lean_ctor_set(x_529, 1, x_528); lean_ctor_set(x_529, 2, x_527); -x_530 = lean_array_push(x_514, x_529); +x_530 = lean_array_push(x_513, x_529); x_531 = lean_array_push(x_530, x_474); +lean_inc(x_5); x_532 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_532, 0, x_502); -lean_ctor_set(x_532, 1, x_503); +lean_ctor_set(x_532, 1, x_5); lean_ctor_set(x_532, 2, x_531); -x_533 = lean_array_push(x_514, x_495); +x_533 = lean_array_push(x_513, x_495); x_534 = lean_array_push(x_533, x_532); x_535 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; x_536 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_536, 0, x_502); lean_ctor_set(x_536, 1, x_535); lean_ctor_set(x_536, 2, x_534); -x_537 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2(x_1, x_2, x_3, x_4, x_9, x_536, x_11, x_12, x_13, x_14, x_15, x_16, x_485); +x_537 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2(x_1, x_2, x_3, x_4, x_5, x_10, x_536, x_12, x_13, x_14, x_15, x_16, x_17, x_485); return x_537; } } @@ -6032,16 +7290,16 @@ return x_537; default: { lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; uint8_t x_562; lean_object* x_563; lean_object* x_564; -lean_dec(x_107); -x_551 = l_Lean_instInhabitedSyntax; -x_552 = l_Array_back___rarg(x_551, x_5); -x_553 = lean_array_get_size(x_5); +lean_dec(x_108); +x_551 = lean_box(0); +x_552 = l_Array_back___rarg(x_551, x_6); +x_553 = lean_array_get_size(x_6); x_554 = lean_unsigned_to_nat(1u); x_555 = lean_nat_sub(x_553, x_554); lean_dec(x_553); x_556 = lean_unsigned_to_nat(0u); -lean_inc(x_5); -x_557 = l_Array_toSubarray___rarg(x_5, x_556, x_555); +lean_inc(x_6); +x_557 = l_Array_toSubarray___rarg(x_6, x_556, x_555); x_558 = lean_ctor_get(x_557, 0); lean_inc(x_558); x_559 = lean_ctor_get(x_557, 2); @@ -6062,7 +7320,7 @@ lean_dec(x_561); lean_dec(x_560); lean_dec(x_558); x_563 = x_552; -x_564 = x_17; +x_564 = x_18; goto block_627; } else @@ -6072,9 +7330,10 @@ x_629 = lean_usize_of_nat(x_561); lean_dec(x_561); x_630 = lean_usize_of_nat(x_560); lean_dec(x_560); -lean_inc(x_15); -lean_inc(x_6); -x_631 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3(x_6, x_558, x_629, x_630, x_552, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_inc(x_16); +lean_inc(x_5); +lean_inc(x_7); +x_631 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__19(x_7, x_5, x_558, x_629, x_630, x_552, x_12, x_13, x_14, x_15, x_16, x_17, x_18); lean_dec(x_558); x_632 = lean_ctor_get(x_631, 0); lean_inc(x_632); @@ -6097,7 +7356,7 @@ lean_dec(x_560); lean_dec(x_559); lean_dec(x_558); x_563 = x_552; -x_564 = x_17; +x_564 = x_18; goto block_627; } else @@ -6107,9 +7366,10 @@ x_635 = lean_usize_of_nat(x_559); lean_dec(x_559); x_636 = lean_usize_of_nat(x_560); lean_dec(x_560); -lean_inc(x_15); -lean_inc(x_6); -x_637 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3(x_6, x_558, x_635, x_636, x_552, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_inc(x_16); +lean_inc(x_5); +lean_inc(x_7); +x_637 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__20(x_7, x_5, x_558, x_635, x_636, x_552, x_12, x_13, x_14, x_15, x_16, x_17, x_18); lean_dec(x_558); x_638 = lean_ctor_get(x_637, 0); lean_inc(x_638); @@ -6124,23 +7384,23 @@ goto block_627; block_627: { lean_object* x_565; lean_object* x_566; lean_object* x_567; lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; -lean_inc(x_15); -x_565 = l_Lean_Elab_Term_Quotation_mkTuple(x_5, x_11, x_12, x_13, x_14, x_15, x_16, x_564); +lean_inc(x_16); +x_565 = l_Lean_Elab_Term_Quotation_mkTuple(x_6, x_12, x_13, x_14, x_15, x_16, x_17, x_564); x_566 = lean_ctor_get(x_565, 0); lean_inc(x_566); x_567 = lean_ctor_get(x_565, 1); lean_inc(x_567); lean_dec(x_565); -lean_inc(x_15); -x_568 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_15, x_16, x_567); +lean_inc(x_16); +x_568 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_16, x_17, x_567); x_569 = lean_ctor_get(x_568, 0); lean_inc(x_569); x_570 = lean_ctor_get(x_568, 1); lean_inc(x_570); lean_dec(x_568); -x_571 = lean_ctor_get(x_15, 10); +x_571 = lean_ctor_get(x_16, 10); lean_inc(x_571); -x_572 = lean_st_ref_get(x_16, x_570); +x_572 = lean_st_ref_get(x_17, x_570); x_573 = lean_ctor_get(x_572, 0); lean_inc(x_573); x_574 = lean_ctor_get(x_572, 1); @@ -6150,8 +7410,8 @@ x_575 = lean_ctor_get(x_573, 0); lean_inc(x_575); lean_dec(x_573); x_576 = lean_environment_main_module(x_575); -x_577 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__4; -x_578 = lean_name_mk_string(x_6, x_577); +x_577 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__4; +x_578 = lean_name_mk_string(x_7, x_577); lean_inc(x_578); x_579 = l_Lean_addMacroScope(x_576, x_578, x_571); x_580 = lean_box(0); @@ -6161,19 +7421,19 @@ lean_ctor_set(x_581, 1, x_580); x_582 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_582, 0, x_581); lean_ctor_set(x_582, 1, x_580); -x_583 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__3; +x_583 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__3; lean_inc(x_569); x_584 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_584, 0, x_569); lean_ctor_set(x_584, 1, x_583); lean_ctor_set(x_584, 2, x_579); lean_ctor_set(x_584, 3, x_582); -x_585 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; +x_585 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; lean_inc(x_569); x_586 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_586, 0, x_569); lean_ctor_set(x_586, 1, x_585); -x_587 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__8; +x_587 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__8; lean_inc(x_569); x_588 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_588, 0, x_569); @@ -6181,68 +7441,75 @@ lean_ctor_set(x_588, 1, x_587); x_589 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; x_590 = lean_array_push(x_589, x_566); x_591 = lean_box(2); -x_592 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_593 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_593, 0, x_591); -lean_ctor_set(x_593, 1, x_592); -lean_ctor_set(x_593, 2, x_590); -x_594 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__12; +lean_inc(x_5); +x_592 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_592, 0, x_591); +lean_ctor_set(x_592, 1, x_5); +lean_ctor_set(x_592, 2, x_590); +x_593 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__12; lean_inc(x_569); -x_595 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_595, 0, x_569); -lean_ctor_set(x_595, 1, x_594); -x_596 = lean_array_get(x_551, x_7, x_556); -lean_dec(x_7); -x_597 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; -x_598 = lean_array_push(x_597, x_593); +x_594 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_594, 0, x_569); +lean_ctor_set(x_594, 1, x_593); +x_595 = lean_array_get(x_551, x_8, x_556); +lean_dec(x_8); +x_596 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; +x_597 = lean_array_push(x_596, x_592); +x_598 = lean_array_push(x_597, x_594); x_599 = lean_array_push(x_598, x_595); -x_600 = lean_array_push(x_599, x_596); -x_601 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__11; -x_602 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_602, 0, x_591); -lean_ctor_set(x_602, 1, x_601); -lean_ctor_set(x_602, 2, x_600); -x_603 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_604 = lean_array_push(x_603, x_588); -x_605 = lean_array_push(x_604, x_602); -x_606 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__9; -x_607 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_607, 0, x_591); -lean_ctor_set(x_607, 1, x_606); -lean_ctor_set(x_607, 2, x_605); -x_608 = lean_array_push(x_603, x_607); -x_609 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; -x_610 = lean_array_push(x_608, x_609); +x_600 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__11; +x_601 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_601, 0, x_591); +lean_ctor_set(x_601, 1, x_600); +lean_ctor_set(x_601, 2, x_599); +x_602 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_603 = lean_array_push(x_602, x_588); +x_604 = lean_array_push(x_603, x_601); +x_605 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__9; +x_606 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_606, 0, x_591); +lean_ctor_set(x_606, 1, x_605); +lean_ctor_set(x_606, 2, x_604); +x_607 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; +lean_inc(x_5); +x_608 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_608, 0, x_591); +lean_ctor_set(x_608, 1, x_5); +lean_ctor_set(x_608, 2, x_607); +x_609 = lean_array_push(x_602, x_606); +x_610 = lean_array_push(x_609, x_608); +lean_inc(x_5); x_611 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_611, 0, x_591); -lean_ctor_set(x_611, 1, x_592); +lean_ctor_set(x_611, 1, x_5); lean_ctor_set(x_611, 2, x_610); -x_612 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; +x_612 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; x_613 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_613, 0, x_569); lean_ctor_set(x_613, 1, x_612); -x_614 = lean_array_push(x_597, x_586); +x_614 = lean_array_push(x_596, x_586); x_615 = lean_array_push(x_614, x_611); x_616 = lean_array_push(x_615, x_613); -x_617 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__6; +x_617 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6; x_618 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_618, 0, x_591); lean_ctor_set(x_618, 1, x_617); lean_ctor_set(x_618, 2, x_616); -x_619 = lean_array_push(x_603, x_618); +x_619 = lean_array_push(x_602, x_618); x_620 = lean_array_push(x_619, x_563); +lean_inc(x_5); x_621 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_621, 0, x_591); -lean_ctor_set(x_621, 1, x_592); +lean_ctor_set(x_621, 1, x_5); lean_ctor_set(x_621, 2, x_620); -x_622 = lean_array_push(x_603, x_584); +x_622 = lean_array_push(x_602, x_584); x_623 = lean_array_push(x_622, x_621); x_624 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; x_625 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_625, 0, x_591); lean_ctor_set(x_625, 1, x_624); lean_ctor_set(x_625, 2, x_623); -x_626 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2(x_1, x_2, x_3, x_4, x_9, x_625, x_11, x_12, x_13, x_14, x_15, x_16, x_574); +x_626 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2(x_1, x_2, x_3, x_4, x_5, x_10, x_625, x_12, x_13, x_14, x_15, x_16, x_17, x_574); return x_626; } } @@ -6250,7 +7517,7 @@ return x_626; } } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__1() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__1() { _start: { lean_object* x_1; @@ -6258,7 +7525,7 @@ x_1 = lean_alloc_closure((void*)(l_ReaderT_pure___at_Lean_Elab_Term_addTermInfo_ return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__2() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__2() { _start: { lean_object* x_1; @@ -6266,16 +7533,16 @@ x_1 = lean_mk_string_from_bytes("antiquotation splice must contain at least one return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__3() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__2; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__2; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__4() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__4() { _start: { lean_object* x_1; @@ -6283,16 +7550,16 @@ x_1 = lean_mk_string_from_bytes("invalid antiquotation suffix splice kind '", 42 return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__5() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__4; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__4; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__6() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__6() { _start: { lean_object* x_1; @@ -6300,39 +7567,39 @@ x_1 = lean_mk_string_from_bytes("none", 4); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__7() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__6; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__6; x_2 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__8() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__7; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__7; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__9() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__5; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__8; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__5; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__8; x_3 = lean_alloc_ctor(10, 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__10() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__10() { _start: { lean_object* x_1; @@ -6340,28 +7607,28 @@ x_1 = lean_mk_string_from_bytes("'", 1); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__11() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__10; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__10; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__12() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__9; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__11; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__9; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__11; x_3 = lean_alloc_ctor(10, 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__13() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__13() { _start: { lean_object* x_1; @@ -6369,47 +7636,47 @@ x_1 = lean_mk_string_from_bytes("some (", 6); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__14() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__14() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__13; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__13; x_2 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__15() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__15() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__14; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__14; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__16() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__16() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__17() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__17() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__16; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__16; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__18() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__18() { _start: { lean_object* x_1; @@ -6417,7 +7684,7 @@ x_1 = lean_mk_string_from_bytes("many", 4); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__19() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__19() { _start: { lean_object* x_1; @@ -6425,17 +7692,17 @@ x_1 = lean_mk_string_from_bytes("explicit", 8); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__20() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__19; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__19; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__21() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__21() { _start: { lean_object* x_1; @@ -6443,7 +7710,7 @@ x_1 = lean_mk_string_from_bytes("@", 1); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__22() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__22() { _start: { lean_object* x_1; @@ -6451,22 +7718,22 @@ x_1 = lean_mk_string_from_bytes("TSepArray.elemsAndSeps", 22); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__23() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__23() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__22; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__22; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__24() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__22; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__22; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__23; +x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__23; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -6474,7 +7741,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__25() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__25() { _start: { lean_object* x_1; @@ -6482,17 +7749,17 @@ x_1 = lean_mk_string_from_bytes("TSepArray", 9); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__26() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__25; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__25; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__27() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__27() { _start: { lean_object* x_1; @@ -6500,17 +7767,17 @@ x_1 = lean_mk_string_from_bytes("elemsAndSeps", 12); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__28() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__26; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__27; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__26; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__27; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__29() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__29() { _start: { lean_object* x_1; @@ -6518,61 +7785,61 @@ x_1 = lean_mk_string_from_bytes("Syntax", 6); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__30() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__30() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__2; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__29; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__29; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__31() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__31() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__30; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__25; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__30; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__25; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__32() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__32() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__31; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__27; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__31; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__27; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__33() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__33() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__32; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__32; 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__34() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__34() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__33; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__33; x_3 = lean_alloc_ctor(1, 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__35() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__35() { _start: { lean_object* x_1; @@ -6580,22 +7847,22 @@ x_1 = lean_mk_string_from_bytes("TSyntaxArray.raw", 16); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__36() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__36() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__35; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__35; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__37() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__37() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__35; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__35; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__36; +x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__36; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -6603,7 +7870,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__38() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__38() { _start: { lean_object* x_1; @@ -6611,17 +7878,17 @@ x_1 = lean_mk_string_from_bytes("TSyntaxArray", 12); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__39() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__39() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__38; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__38; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__40() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__40() { _start: { lean_object* x_1; @@ -6629,71 +7896,71 @@ x_1 = lean_mk_string_from_bytes("raw", 3); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__41() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__41() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__39; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__40; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__39; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__40; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__42() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__42() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__2; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__38; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__38; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__43() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__43() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__42; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__40; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__42; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__40; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__44() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__44() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__43; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__43; 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__45() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__45() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__44; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__44; x_3 = lean_alloc_ctor(1, 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__46() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__46() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_2 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__1; +x_2 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__12___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__47() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__47() { _start: { lean_object* x_1; @@ -6701,22 +7968,22 @@ x_1 = lean_mk_string_from_bytes("Option.map", 10); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__48() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__48() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__47; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__47; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__49() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__49() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__47; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__47; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__48; +x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__48; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -6724,41 +7991,41 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__50() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__50() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__6; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__4; +x_1 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__6; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__51() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__51() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__50; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__50; 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__52() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__52() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__51; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__51; x_3 = lean_alloc_ctor(1, 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__53() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__53() { _start: { lean_object* x_1; @@ -6766,22 +8033,22 @@ x_1 = lean_mk_string_from_bytes("TSyntax.raw", 11); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__54() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__54() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__53; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__53; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__55() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__55() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__53; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__53; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__54; +x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__54; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -6789,7 +8056,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__56() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__56() { _start: { lean_object* x_1; @@ -6797,71 +8064,71 @@ x_1 = lean_mk_string_from_bytes("TSyntax", 7); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__57() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__57() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__56; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__56; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__58() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__58() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__57; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__40; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__57; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__40; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__59() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__59() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__2; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__56; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__56; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__60() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__60() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__59; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__40; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__59; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__40; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__61() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__61() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__60; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__60; 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__62() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__62() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__61; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__61; x_3 = lean_alloc_ctor(1, 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__63() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__63() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -6871,31 +8138,31 @@ x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__64() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__64() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__7; +x_2 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__7; 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__65() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__65() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__64; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__64; x_3 = lean_alloc_ctor(1, 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__66() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__66() { _start: { lean_object* x_1; @@ -6903,22 +8170,22 @@ x_1 = lean_mk_string_from_bytes("x", 1); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__67() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__67() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__66; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__66; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__68() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__68() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__66; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__66; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__67; +x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__67; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -6926,17 +8193,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__69() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__69() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__66; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__66; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__70() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__70() { _start: { lean_object* x_1; @@ -6944,22 +8211,22 @@ x_1 = lean_mk_string_from_bytes("Array.empty.push", 16); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__71() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__71() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__70; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__70; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__72() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__72() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__70; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__70; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__71; +x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__71; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -6967,7 +8234,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__73() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__73() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -6979,22 +8246,22 @@ lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__74() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__74() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__6; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__6; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__75() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__75() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__6; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__6; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__74; +x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__74; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -7002,51 +8269,51 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__76() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__76() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__6; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__6; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__77() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__77() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__6; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__6; +x_1 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__6; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__6; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__78() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__78() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__77; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__77; 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__79() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__79() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__78; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__78; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, lean_object* x_5, lean_object* x_6, size_t x_7, size_t 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, lean_object* x_15, lean_object* x_16) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, lean_object* x_5, lean_object* x_6, size_t x_7, size_t 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, lean_object* x_15, lean_object* x_16) { _start: { uint8_t x_17; @@ -7071,7 +8338,7 @@ else { lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_29; lean_object* x_39; uint8_t x_40; x_19 = lean_array_uget(x_6, x_8); -x_39 = l_Lean_nullKind; +x_39 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; x_40 = lean_name_eq(x_1, x_39); if (x_40 == 0) { @@ -7199,7 +8466,7 @@ else { lean_object* x_61; lean_object* x_62; lean_object* x_63; x_61 = l_Lean_Syntax_antiquotSpliceKind_x3f(x_19); -x_62 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__1; +x_62 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__1; lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); @@ -7230,7 +8497,7 @@ lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); -x_71 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__1(x_70, x_4, x_68, x_10, x_11, x_12, x_13, x_14, x_15, x_65); +x_71 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4(x_70, x_4, x_68, x_10, x_11, x_12, x_13, x_14, x_15, x_65); if (lean_obj_tag(x_71) == 0) { lean_object* x_72; lean_object* x_73; lean_object* x_74; @@ -7268,7 +8535,7 @@ lean_inc(x_11); lean_inc(x_10); lean_inc(x_3); lean_inc(x_5); -x_79 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3(x_67, x_5, x_61, x_66, x_75, x_3, x_72, x_4, x_9, x_78, x_10, x_11, x_12, x_13, x_14, x_15, x_76); +x_79 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3(x_67, x_5, x_61, x_66, x_39, x_75, x_3, x_72, x_4, x_9, x_78, x_10, x_11, x_12, x_13, x_14, x_15, x_76); if (lean_obj_tag(x_79) == 0) { lean_object* x_80; lean_object* x_81; @@ -7323,7 +8590,7 @@ lean_dec(x_61); lean_dec(x_9); lean_dec(x_5); lean_dec(x_3); -x_86 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__3; +x_86 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__3; x_87 = l_Lean_throwErrorAt___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__1(x_2, x_86, x_10, x_11, x_12, x_13, x_14, x_15, x_76); lean_dec(x_15); lean_dec(x_13); @@ -7462,7 +8729,7 @@ x_105 = l_Lean_Syntax_getArg(x_19, x_104); lean_inc(x_105); x_106 = l_Lean_Syntax_antiquotKinds(x_105); x_107 = lean_box(0); -x_108 = l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__7(x_106, x_107); +x_108 = l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__21(x_106, x_107); x_109 = l_Lean_Syntax_getCanonicalAntiquot(x_105); x_110 = l_Lean_Syntax_getAntiquotTerm(x_109); lean_dec(x_109); @@ -7472,14 +8739,14 @@ if (lean_obj_tag(x_111) == 0) lean_object* x_112; lean_object* x_113; lean_dec(x_110); lean_dec(x_108); -x_112 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__12; +x_112 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__12; lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); -x_113 = l_Lean_throwErrorAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__10(x_19, x_112, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +x_113 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__22(x_19, x_112, x_10, x_11, x_12, x_13, x_14, x_15, x_16); x_29 = x_113; goto block_38; } @@ -7499,17 +8766,17 @@ if (lean_obj_tag(x_115) == 0) lean_object* x_116; lean_object* x_117; uint8_t x_118; x_116 = lean_ctor_get(x_114, 1); lean_inc(x_116); -x_117 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__15; +x_117 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__15; x_118 = lean_string_dec_eq(x_116, x_117); if (x_118 == 0) { lean_object* x_119; uint8_t x_120; -x_119 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__18; +x_119 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__18; x_120 = lean_string_dec_eq(x_116, x_119); if (x_120 == 0) { lean_object* x_121; uint8_t x_122; -x_121 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__1; +x_121 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__1; x_122 = lean_string_dec_eq(x_116, x_121); lean_dec(x_116); if (x_122 == 0) @@ -7519,19 +8786,19 @@ lean_dec(x_110); lean_dec(x_108); x_123 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_123, 0, x_114); -x_124 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__15; +x_124 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__15; x_125 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_125, 0, x_124); lean_ctor_set(x_125, 1, x_123); -x_126 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__17; +x_126 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__17; x_127 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_127, 0, x_125); lean_ctor_set(x_127, 1, x_126); -x_128 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__5; +x_128 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__5; x_129 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_129, 0, x_128); lean_ctor_set(x_129, 1, x_127); -x_130 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__11; +x_130 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__11; x_131 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_131, 0, x_129); lean_ctor_set(x_131, 1, x_130); @@ -7541,939 +8808,928 @@ lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); -x_132 = l_Lean_throwErrorAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__10(x_19, x_131, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +x_132 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__22(x_19, x_131, x_10, x_11, x_12, x_13, x_14, x_15, x_16); x_29 = x_132; goto block_38; } else { -lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; uint8_t x_138; +lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; uint8_t x_141; lean_dec(x_114); -lean_inc(x_14); -x_133 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_14, x_15, x_16); -x_134 = lean_ctor_get(x_133, 0); -lean_inc(x_134); -x_135 = lean_ctor_get(x_133, 1); -lean_inc(x_135); +x_133 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepFromSplice(x_19); +lean_dec(x_19); +x_134 = lean_box(2); +x_135 = l_Lean_Syntax_mkStrLit(x_133, x_134); lean_dec(x_133); -x_136 = lean_ctor_get(x_14, 10); -lean_inc(x_136); -x_137 = lean_st_ref_get(x_15, x_135); -x_138 = !lean_is_exclusive(x_137); -if (x_138 == 0) +lean_inc(x_14); +x_136 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_14, x_15, x_16); +x_137 = lean_ctor_get(x_136, 0); +lean_inc(x_137); +x_138 = lean_ctor_get(x_136, 1); +lean_inc(x_138); +lean_dec(x_136); +x_139 = lean_ctor_get(x_14, 10); +lean_inc(x_139); +x_140 = lean_st_ref_get(x_15, x_138); +x_141 = !lean_is_exclusive(x_140); +if (x_141 == 0) { -lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; -x_139 = lean_ctor_get(x_137, 0); -x_140 = lean_ctor_get(x_139, 0); -lean_inc(x_140); -lean_dec(x_139); -x_141 = lean_environment_main_module(x_140); -x_142 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__21; -lean_inc(x_134); -x_143 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_143, 0, x_134); -lean_ctor_set(x_143, 1, x_142); -x_144 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__28; -x_145 = l_Lean_addMacroScope(x_141, x_144, x_136); -x_146 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__24; -x_147 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__34; -x_148 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_148, 0, x_134); -lean_ctor_set(x_148, 1, x_146); -lean_ctor_set(x_148, 2, x_145); -lean_ctor_set(x_148, 3, x_147); -x_149 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_150 = lean_array_push(x_149, x_143); -x_151 = lean_array_push(x_150, x_148); -x_152 = lean_box(2); -x_153 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__20; -x_154 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_154, 0, x_152); -lean_ctor_set(x_154, 1, x_153); -lean_ctor_set(x_154, 2, x_151); -x_155 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8(x_108); -x_156 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepFromSplice(x_19); -lean_dec(x_19); -x_157 = l_Lean_Syntax_mkStrLit(x_156, x_152); -lean_dec(x_156); -x_158 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; -x_159 = lean_array_push(x_158, x_155); -x_160 = lean_array_push(x_159, x_157); +lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; +x_142 = lean_ctor_get(x_140, 0); +x_143 = lean_ctor_get(x_142, 0); +lean_inc(x_143); +lean_dec(x_142); +x_144 = lean_environment_main_module(x_143); +x_145 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__21; +lean_inc(x_137); +x_146 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_146, 0, x_137); +lean_ctor_set(x_146, 1, x_145); +x_147 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__28; +x_148 = l_Lean_addMacroScope(x_144, x_147, x_139); +x_149 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__24; +x_150 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__34; +x_151 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_151, 0, x_137); +lean_ctor_set(x_151, 1, x_149); +lean_ctor_set(x_151, 2, x_148); +lean_ctor_set(x_151, 3, x_150); +x_152 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_153 = lean_array_push(x_152, x_146); +x_154 = lean_array_push(x_153, x_151); +x_155 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__20; +x_156 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_156, 0, x_134); +lean_ctor_set(x_156, 1, x_155); +lean_ctor_set(x_156, 2, x_154); +x_157 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24(x_108); +x_158 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; +x_159 = lean_array_push(x_158, x_157); +x_160 = lean_array_push(x_159, x_135); x_161 = lean_array_push(x_160, x_110); -x_162 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_163 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_163, 0, x_152); -lean_ctor_set(x_163, 1, x_162); -lean_ctor_set(x_163, 2, x_161); -x_164 = lean_array_push(x_149, x_154); -x_165 = lean_array_push(x_164, x_163); -x_166 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_167 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_167, 0, x_152); -lean_ctor_set(x_167, 1, x_166); -lean_ctor_set(x_167, 2, x_165); -lean_ctor_set(x_137, 0, x_167); -x_29 = x_137; +x_162 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_162, 0, x_134); +lean_ctor_set(x_162, 1, x_39); +lean_ctor_set(x_162, 2, x_161); +x_163 = lean_array_push(x_152, x_156); +x_164 = lean_array_push(x_163, x_162); +x_165 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_166 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_166, 0, x_134); +lean_ctor_set(x_166, 1, x_165); +lean_ctor_set(x_166, 2, x_164); +lean_ctor_set(x_140, 0, x_166); +x_29 = x_140; goto block_38; } else { -lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; -x_168 = lean_ctor_get(x_137, 0); -x_169 = lean_ctor_get(x_137, 1); -lean_inc(x_169); +lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; +x_167 = lean_ctor_get(x_140, 0); +x_168 = lean_ctor_get(x_140, 1); lean_inc(x_168); -lean_dec(x_137); -x_170 = lean_ctor_get(x_168, 0); -lean_inc(x_170); -lean_dec(x_168); -x_171 = lean_environment_main_module(x_170); -x_172 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__21; -lean_inc(x_134); -x_173 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_173, 0, x_134); -lean_ctor_set(x_173, 1, x_172); -x_174 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__28; -x_175 = l_Lean_addMacroScope(x_171, x_174, x_136); -x_176 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__24; -x_177 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__34; -x_178 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_178, 0, x_134); -lean_ctor_set(x_178, 1, x_176); -lean_ctor_set(x_178, 2, x_175); -lean_ctor_set(x_178, 3, x_177); -x_179 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_180 = lean_array_push(x_179, x_173); -x_181 = lean_array_push(x_180, x_178); -x_182 = lean_box(2); -x_183 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__20; -x_184 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_184, 0, x_182); -lean_ctor_set(x_184, 1, x_183); -lean_ctor_set(x_184, 2, x_181); -x_185 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8(x_108); -x_186 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepFromSplice(x_19); -lean_dec(x_19); -x_187 = l_Lean_Syntax_mkStrLit(x_186, x_182); -lean_dec(x_186); -x_188 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; -x_189 = lean_array_push(x_188, x_185); -x_190 = lean_array_push(x_189, x_187); -x_191 = lean_array_push(x_190, x_110); -x_192 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_193 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_193, 0, x_182); -lean_ctor_set(x_193, 1, x_192); -lean_ctor_set(x_193, 2, x_191); -x_194 = lean_array_push(x_179, x_184); -x_195 = lean_array_push(x_194, x_193); -x_196 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_197 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_197, 0, x_182); -lean_ctor_set(x_197, 1, x_196); -lean_ctor_set(x_197, 2, x_195); -x_198 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_198, 0, x_197); -lean_ctor_set(x_198, 1, x_169); -x_29 = x_198; +lean_inc(x_167); +lean_dec(x_140); +x_169 = lean_ctor_get(x_167, 0); +lean_inc(x_169); +lean_dec(x_167); +x_170 = lean_environment_main_module(x_169); +x_171 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__21; +lean_inc(x_137); +x_172 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_172, 0, x_137); +lean_ctor_set(x_172, 1, x_171); +x_173 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__28; +x_174 = l_Lean_addMacroScope(x_170, x_173, x_139); +x_175 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__24; +x_176 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__34; +x_177 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_177, 0, x_137); +lean_ctor_set(x_177, 1, x_175); +lean_ctor_set(x_177, 2, x_174); +lean_ctor_set(x_177, 3, x_176); +x_178 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_179 = lean_array_push(x_178, x_172); +x_180 = lean_array_push(x_179, x_177); +x_181 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__20; +x_182 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_182, 0, x_134); +lean_ctor_set(x_182, 1, x_181); +lean_ctor_set(x_182, 2, x_180); +x_183 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24(x_108); +x_184 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; +x_185 = lean_array_push(x_184, x_183); +x_186 = lean_array_push(x_185, x_135); +x_187 = lean_array_push(x_186, x_110); +x_188 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_188, 0, x_134); +lean_ctor_set(x_188, 1, x_39); +lean_ctor_set(x_188, 2, x_187); +x_189 = lean_array_push(x_178, x_182); +x_190 = lean_array_push(x_189, x_188); +x_191 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_192 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_192, 0, x_134); +lean_ctor_set(x_192, 1, x_191); +lean_ctor_set(x_192, 2, x_190); +x_193 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_193, 0, x_192); +lean_ctor_set(x_193, 1, x_168); +x_29 = x_193; goto block_38; } } } else { -lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; uint8_t x_204; +lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; uint8_t x_199; lean_dec(x_116); lean_dec(x_114); lean_dec(x_19); lean_inc(x_14); -x_199 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_14, x_15, x_16); -x_200 = lean_ctor_get(x_199, 0); -lean_inc(x_200); -x_201 = lean_ctor_get(x_199, 1); -lean_inc(x_201); -lean_dec(x_199); -x_202 = lean_ctor_get(x_14, 10); -lean_inc(x_202); -x_203 = lean_st_ref_get(x_15, x_201); -x_204 = !lean_is_exclusive(x_203); -if (x_204 == 0) +x_194 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_14, x_15, x_16); +x_195 = lean_ctor_get(x_194, 0); +lean_inc(x_195); +x_196 = lean_ctor_get(x_194, 1); +lean_inc(x_196); +lean_dec(x_194); +x_197 = lean_ctor_get(x_14, 10); +lean_inc(x_197); +x_198 = lean_st_ref_get(x_15, x_196); +x_199 = !lean_is_exclusive(x_198); +if (x_199 == 0) { -lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; -x_205 = lean_ctor_get(x_203, 0); -x_206 = lean_ctor_get(x_205, 0); -lean_inc(x_206); -lean_dec(x_205); -x_207 = lean_environment_main_module(x_206); -x_208 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__21; -lean_inc(x_200); -x_209 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_209, 0, x_200); -lean_ctor_set(x_209, 1, x_208); -x_210 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__41; -x_211 = l_Lean_addMacroScope(x_207, x_210, x_202); -x_212 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__37; -x_213 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__45; -x_214 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_214, 0, x_200); -lean_ctor_set(x_214, 1, x_212); -lean_ctor_set(x_214, 2, x_211); -lean_ctor_set(x_214, 3, x_213); -x_215 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_216 = lean_array_push(x_215, x_209); -x_217 = lean_array_push(x_216, x_214); -x_218 = lean_box(2); -x_219 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__20; -x_220 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_220, 0, x_218); -lean_ctor_set(x_220, 1, x_219); -lean_ctor_set(x_220, 2, x_217); -x_221 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8(x_108); -x_222 = lean_array_push(x_215, x_221); -x_223 = lean_array_push(x_222, x_110); -x_224 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_225 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_225, 0, x_218); -lean_ctor_set(x_225, 1, x_224); -lean_ctor_set(x_225, 2, x_223); -x_226 = lean_array_push(x_215, x_220); -x_227 = lean_array_push(x_226, x_225); -x_228 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_229 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_229, 0, x_218); -lean_ctor_set(x_229, 1, x_228); -lean_ctor_set(x_229, 2, x_227); -lean_ctor_set(x_203, 0, x_229); -x_29 = x_203; +lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; +x_200 = lean_ctor_get(x_198, 0); +x_201 = lean_ctor_get(x_200, 0); +lean_inc(x_201); +lean_dec(x_200); +x_202 = lean_environment_main_module(x_201); +x_203 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__21; +lean_inc(x_195); +x_204 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_204, 0, x_195); +lean_ctor_set(x_204, 1, x_203); +x_205 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__41; +x_206 = l_Lean_addMacroScope(x_202, x_205, x_197); +x_207 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__37; +x_208 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__45; +x_209 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_209, 0, x_195); +lean_ctor_set(x_209, 1, x_207); +lean_ctor_set(x_209, 2, x_206); +lean_ctor_set(x_209, 3, x_208); +x_210 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_211 = lean_array_push(x_210, x_204); +x_212 = lean_array_push(x_211, x_209); +x_213 = lean_box(2); +x_214 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__20; +x_215 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_215, 0, x_213); +lean_ctor_set(x_215, 1, x_214); +lean_ctor_set(x_215, 2, x_212); +x_216 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24(x_108); +x_217 = lean_array_push(x_210, x_216); +x_218 = lean_array_push(x_217, x_110); +x_219 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_219, 0, x_213); +lean_ctor_set(x_219, 1, x_39); +lean_ctor_set(x_219, 2, x_218); +x_220 = lean_array_push(x_210, x_215); +x_221 = lean_array_push(x_220, x_219); +x_222 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_223 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_223, 0, x_213); +lean_ctor_set(x_223, 1, x_222); +lean_ctor_set(x_223, 2, x_221); +lean_ctor_set(x_198, 0, x_223); +x_29 = x_198; goto block_38; } else { -lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; -x_230 = lean_ctor_get(x_203, 0); -x_231 = lean_ctor_get(x_203, 1); -lean_inc(x_231); -lean_inc(x_230); -lean_dec(x_203); -x_232 = lean_ctor_get(x_230, 0); -lean_inc(x_232); -lean_dec(x_230); -x_233 = lean_environment_main_module(x_232); -x_234 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__21; -lean_inc(x_200); -x_235 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_235, 0, x_200); -lean_ctor_set(x_235, 1, x_234); -x_236 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__41; -x_237 = l_Lean_addMacroScope(x_233, x_236, x_202); -x_238 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__37; -x_239 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__45; -x_240 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_240, 0, x_200); -lean_ctor_set(x_240, 1, x_238); +lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; +x_224 = lean_ctor_get(x_198, 0); +x_225 = lean_ctor_get(x_198, 1); +lean_inc(x_225); +lean_inc(x_224); +lean_dec(x_198); +x_226 = lean_ctor_get(x_224, 0); +lean_inc(x_226); +lean_dec(x_224); +x_227 = lean_environment_main_module(x_226); +x_228 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__21; +lean_inc(x_195); +x_229 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_229, 0, x_195); +lean_ctor_set(x_229, 1, x_228); +x_230 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__41; +x_231 = l_Lean_addMacroScope(x_227, x_230, x_197); +x_232 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__37; +x_233 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__45; +x_234 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_234, 0, x_195); +lean_ctor_set(x_234, 1, x_232); +lean_ctor_set(x_234, 2, x_231); +lean_ctor_set(x_234, 3, x_233); +x_235 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_236 = lean_array_push(x_235, x_229); +x_237 = lean_array_push(x_236, x_234); +x_238 = lean_box(2); +x_239 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__20; +x_240 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_240, 0, x_238); +lean_ctor_set(x_240, 1, x_239); lean_ctor_set(x_240, 2, x_237); -lean_ctor_set(x_240, 3, x_239); -x_241 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_242 = lean_array_push(x_241, x_235); -x_243 = lean_array_push(x_242, x_240); -x_244 = lean_box(2); -x_245 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__20; -x_246 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_246, 0, x_244); -lean_ctor_set(x_246, 1, x_245); -lean_ctor_set(x_246, 2, x_243); -x_247 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8(x_108); -x_248 = lean_array_push(x_241, x_247); -x_249 = lean_array_push(x_248, x_110); -x_250 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_251 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_251, 0, x_244); -lean_ctor_set(x_251, 1, x_250); -lean_ctor_set(x_251, 2, x_249); -x_252 = lean_array_push(x_241, x_246); -x_253 = lean_array_push(x_252, x_251); -x_254 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_255 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_255, 0, x_244); -lean_ctor_set(x_255, 1, x_254); -lean_ctor_set(x_255, 2, x_253); -x_256 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_256, 0, x_255); -lean_ctor_set(x_256, 1, x_231); -x_29 = x_256; +x_241 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24(x_108); +x_242 = lean_array_push(x_235, x_241); +x_243 = lean_array_push(x_242, x_110); +x_244 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_244, 0, x_238); +lean_ctor_set(x_244, 1, x_39); +lean_ctor_set(x_244, 2, x_243); +x_245 = lean_array_push(x_235, x_240); +x_246 = lean_array_push(x_245, x_244); +x_247 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_248 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_248, 0, x_238); +lean_ctor_set(x_248, 1, x_247); +lean_ctor_set(x_248, 2, x_246); +x_249 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_249, 0, x_248); +lean_ctor_set(x_249, 1, x_225); +x_29 = x_249; goto block_38; } } } else { -lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; uint8_t x_262; +lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; uint8_t x_255; lean_dec(x_116); lean_dec(x_114); lean_dec(x_19); lean_inc(x_14); -x_257 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_14, x_15, x_16); -x_258 = lean_ctor_get(x_257, 0); -lean_inc(x_258); -x_259 = lean_ctor_get(x_257, 1); -lean_inc(x_259); -lean_dec(x_257); -x_260 = lean_ctor_get(x_14, 10); -lean_inc(x_260); -x_261 = lean_st_ref_get(x_15, x_259); -x_262 = !lean_is_exclusive(x_261); -if (x_262 == 0) +x_250 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_14, x_15, x_16); +x_251 = lean_ctor_get(x_250, 0); +lean_inc(x_251); +x_252 = lean_ctor_get(x_250, 1); +lean_inc(x_252); +lean_dec(x_250); +x_253 = lean_ctor_get(x_14, 10); +lean_inc(x_253); +x_254 = lean_st_ref_get(x_15, x_252); +x_255 = !lean_is_exclusive(x_254); +if (x_255 == 0) { -lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; -x_263 = lean_ctor_get(x_261, 0); -x_264 = lean_ctor_get(x_263, 0); -lean_inc(x_264); -lean_dec(x_263); -x_265 = lean_environment_main_module(x_264); -x_266 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__16; +lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; +x_256 = lean_ctor_get(x_254, 0); +x_257 = lean_ctor_get(x_256, 0); +lean_inc(x_257); +lean_dec(x_256); +x_258 = lean_environment_main_module(x_257); +x_259 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__16; +lean_inc(x_251); +x_260 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_260, 0, x_251); +lean_ctor_set(x_260, 1, x_259); +x_261 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__50; +lean_inc(x_253); lean_inc(x_258); +x_262 = l_Lean_addMacroScope(x_258, x_261, x_253); +x_263 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__49; +x_264 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__52; +lean_inc(x_251); +x_265 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_265, 0, x_251); +lean_ctor_set(x_265, 1, x_263); +lean_ctor_set(x_265, 2, x_262); +lean_ctor_set(x_265, 3, x_264); +x_266 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; +lean_inc(x_251); x_267 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_267, 0, x_258); +lean_ctor_set(x_267, 0, x_251); lean_ctor_set(x_267, 1, x_266); -x_268 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__50; -lean_inc(x_260); -lean_inc(x_265); -x_269 = l_Lean_addMacroScope(x_265, x_268, x_260); -x_270 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__49; -x_271 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__52; -lean_inc(x_258); -x_272 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_272, 0, x_258); -lean_ctor_set(x_272, 1, x_270); -lean_ctor_set(x_272, 2, x_269); -lean_ctor_set(x_272, 3, x_271); -x_273 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; -lean_inc(x_258); -x_274 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_274, 0, x_258); -lean_ctor_set(x_274, 1, x_273); -x_275 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__21; -lean_inc(x_258); -x_276 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_276, 0, x_258); -lean_ctor_set(x_276, 1, x_275); -x_277 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__58; -lean_inc(x_260); -lean_inc(x_265); -x_278 = l_Lean_addMacroScope(x_265, x_277, x_260); -x_279 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__55; -x_280 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__62; +x_268 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__21; +lean_inc(x_251); +x_269 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_269, 0, x_251); +lean_ctor_set(x_269, 1, x_268); +x_270 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__58; +lean_inc(x_253); lean_inc(x_258); -x_281 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_281, 0, x_258); -lean_ctor_set(x_281, 1, x_279); -lean_ctor_set(x_281, 2, x_278); -lean_ctor_set(x_281, 3, x_280); -x_282 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_283 = lean_array_push(x_282, x_276); -x_284 = lean_array_push(x_283, x_281); -x_285 = lean_box(2); -x_286 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__20; -x_287 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_287, 0, x_285); -lean_ctor_set(x_287, 1, x_286); -lean_ctor_set(x_287, 2, x_284); -x_288 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8(x_108); -x_289 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; -x_290 = lean_array_push(x_289, x_288); -x_291 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_271 = l_Lean_addMacroScope(x_258, x_270, x_253); +x_272 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__55; +x_273 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__62; +lean_inc(x_251); +x_274 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_274, 0, x_251); +lean_ctor_set(x_274, 1, x_272); +lean_ctor_set(x_274, 2, x_271); +lean_ctor_set(x_274, 3, x_273); +x_275 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_276 = lean_array_push(x_275, x_269); +x_277 = lean_array_push(x_276, x_274); +x_278 = lean_box(2); +x_279 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__20; +x_280 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_280, 0, x_278); +lean_ctor_set(x_280, 1, x_279); +lean_ctor_set(x_280, 2, x_277); +x_281 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24(x_108); +x_282 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; +x_283 = lean_array_push(x_282, x_281); +x_284 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_284, 0, x_278); +lean_ctor_set(x_284, 1, x_39); +lean_ctor_set(x_284, 2, x_283); +x_285 = lean_array_push(x_275, x_280); +x_286 = lean_array_push(x_285, x_284); +x_287 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_288 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_288, 0, x_278); +lean_ctor_set(x_288, 1, x_287); +lean_ctor_set(x_288, 2, x_286); +x_289 = lean_array_push(x_275, x_288); +x_290 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; +x_291 = lean_array_push(x_289, x_290); x_292 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_292, 0, x_285); -lean_ctor_set(x_292, 1, x_291); -lean_ctor_set(x_292, 2, x_290); -x_293 = lean_array_push(x_282, x_287); -x_294 = lean_array_push(x_293, x_292); -x_295 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_296 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_296, 0, x_285); -lean_ctor_set(x_296, 1, x_295); -lean_ctor_set(x_296, 2, x_294); -x_297 = lean_array_push(x_282, x_296); -x_298 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; -x_299 = lean_array_push(x_297, x_298); +lean_ctor_set(x_292, 0, x_278); +lean_ctor_set(x_292, 1, x_39); +lean_ctor_set(x_292, 2, x_291); +x_293 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; +lean_inc(x_251); +x_294 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_294, 0, x_251); +lean_ctor_set(x_294, 1, x_293); +x_295 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; +x_296 = lean_array_push(x_295, x_267); +x_297 = lean_array_push(x_296, x_292); +x_298 = lean_array_push(x_297, x_294); +x_299 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6; x_300 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_300, 0, x_285); -lean_ctor_set(x_300, 1, x_291); -lean_ctor_set(x_300, 2, x_299); -x_301 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; -lean_inc(x_258); -x_302 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_302, 0, x_258); -lean_ctor_set(x_302, 1, x_301); -x_303 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; -x_304 = lean_array_push(x_303, x_274); -x_305 = lean_array_push(x_304, x_300); -x_306 = lean_array_push(x_305, x_302); -x_307 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__6; -x_308 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_308, 0, x_285); -lean_ctor_set(x_308, 1, x_307); -lean_ctor_set(x_308, 2, x_306); -x_309 = lean_array_push(x_282, x_308); -x_310 = lean_array_push(x_309, x_110); -x_311 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_311, 0, x_285); -lean_ctor_set(x_311, 1, x_291); -lean_ctor_set(x_311, 2, x_310); -x_312 = lean_array_push(x_282, x_272); -x_313 = lean_array_push(x_312, x_311); -x_314 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_314, 0, x_285); -lean_ctor_set(x_314, 1, x_295); -lean_ctor_set(x_314, 2, x_313); -x_315 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__63; -x_316 = lean_array_push(x_315, x_314); -x_317 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__46; -x_318 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_318, 0, x_285); -lean_ctor_set(x_318, 1, x_317); -lean_ctor_set(x_318, 2, x_316); -x_319 = lean_array_push(x_289, x_318); -x_320 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_320, 0, x_285); -lean_ctor_set(x_320, 1, x_291); -lean_ctor_set(x_320, 2, x_319); -x_321 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__20; -lean_inc(x_258); -x_322 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_322, 0, x_258); -lean_ctor_set(x_322, 1, x_321); -x_323 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__25; -lean_inc(x_258); -x_324 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_324, 0, x_258); -lean_ctor_set(x_324, 1, x_323); -x_325 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__4; -lean_inc(x_260); -lean_inc(x_265); -x_326 = l_Lean_addMacroScope(x_265, x_325, x_260); -x_327 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__3; -x_328 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__65; -lean_inc(x_258); -x_329 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_329, 0, x_258); -lean_ctor_set(x_329, 1, x_327); -lean_ctor_set(x_329, 2, x_326); -lean_ctor_set(x_329, 3, x_328); -x_330 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__69; -lean_inc(x_260); -lean_inc(x_265); -x_331 = l_Lean_addMacroScope(x_265, x_330, x_260); -x_332 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__68; +lean_ctor_set(x_300, 0, x_278); +lean_ctor_set(x_300, 1, x_299); +lean_ctor_set(x_300, 2, x_298); +x_301 = lean_array_push(x_275, x_300); +x_302 = lean_array_push(x_301, x_110); +x_303 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_303, 0, x_278); +lean_ctor_set(x_303, 1, x_39); +lean_ctor_set(x_303, 2, x_302); +x_304 = lean_array_push(x_275, x_265); +x_305 = lean_array_push(x_304, x_303); +x_306 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_306, 0, x_278); +lean_ctor_set(x_306, 1, x_287); +lean_ctor_set(x_306, 2, x_305); +x_307 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__63; +x_308 = lean_array_push(x_307, x_306); +x_309 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__46; +x_310 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_310, 0, x_278); +lean_ctor_set(x_310, 1, x_309); +lean_ctor_set(x_310, 2, x_308); +x_311 = lean_array_push(x_282, x_310); +x_312 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_312, 0, x_278); +lean_ctor_set(x_312, 1, x_39); +lean_ctor_set(x_312, 2, x_311); +x_313 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__20; +lean_inc(x_251); +x_314 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_314, 0, x_251); +lean_ctor_set(x_314, 1, x_313); +x_315 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__25; +lean_inc(x_251); +x_316 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_316, 0, x_251); +lean_ctor_set(x_316, 1, x_315); +x_317 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__4; +lean_inc(x_253); lean_inc(x_258); -x_333 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_333, 0, x_258); -lean_ctor_set(x_333, 1, x_332); -lean_ctor_set(x_333, 2, x_331); -lean_ctor_set(x_333, 3, x_107); -x_334 = lean_array_push(x_289, x_333); -x_335 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_335, 0, x_285); -lean_ctor_set(x_335, 1, x_291); -lean_ctor_set(x_335, 2, x_334); -x_336 = lean_array_push(x_282, x_329); -lean_inc(x_335); -x_337 = lean_array_push(x_336, x_335); -x_338 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_338, 0, x_285); -lean_ctor_set(x_338, 1, x_295); -lean_ctor_set(x_338, 2, x_337); -x_339 = lean_array_push(x_289, x_338); -x_340 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_340, 0, x_285); -lean_ctor_set(x_340, 1, x_291); -lean_ctor_set(x_340, 2, x_339); -x_341 = lean_array_push(x_289, x_340); -x_342 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_342, 0, x_285); -lean_ctor_set(x_342, 1, x_291); -lean_ctor_set(x_342, 2, x_341); -x_343 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__12; +x_318 = l_Lean_addMacroScope(x_258, x_317, x_253); +x_319 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__3; +x_320 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__65; +lean_inc(x_251); +x_321 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_321, 0, x_251); +lean_ctor_set(x_321, 1, x_319); +lean_ctor_set(x_321, 2, x_318); +lean_ctor_set(x_321, 3, x_320); +x_322 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__69; +lean_inc(x_253); lean_inc(x_258); -x_344 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_344, 0, x_258); -lean_ctor_set(x_344, 1, x_343); -x_345 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__29; +x_323 = l_Lean_addMacroScope(x_258, x_322, x_253); +x_324 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__68; +lean_inc(x_251); +x_325 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_325, 0, x_251); +lean_ctor_set(x_325, 1, x_324); +lean_ctor_set(x_325, 2, x_323); +lean_ctor_set(x_325, 3, x_107); +x_326 = lean_array_push(x_282, x_325); +x_327 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_327, 0, x_278); +lean_ctor_set(x_327, 1, x_39); +lean_ctor_set(x_327, 2, x_326); +x_328 = lean_array_push(x_275, x_321); +lean_inc(x_327); +x_329 = lean_array_push(x_328, x_327); +x_330 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_330, 0, x_278); +lean_ctor_set(x_330, 1, x_287); +lean_ctor_set(x_330, 2, x_329); +x_331 = lean_array_push(x_282, x_330); +x_332 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_332, 0, x_278); +lean_ctor_set(x_332, 1, x_39); +lean_ctor_set(x_332, 2, x_331); +x_333 = lean_array_push(x_282, x_332); +x_334 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_334, 0, x_278); +lean_ctor_set(x_334, 1, x_39); +lean_ctor_set(x_334, 2, x_333); +x_335 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__12; +lean_inc(x_251); +x_336 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_336, 0, x_251); +lean_ctor_set(x_336, 1, x_335); +x_337 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__29; lean_inc(x_3); -x_346 = lean_name_mk_string(x_3, x_345); -x_347 = l_Lean_Elab_Term_Quotation_ArrayStxBuilder_push___closed__3; -lean_inc(x_346); -x_348 = lean_name_mk_string(x_346, x_347); -lean_inc(x_260); -lean_inc(x_265); -x_349 = l_Lean_addMacroScope(x_265, x_348, x_260); -x_350 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__73; -lean_inc(x_346); -x_351 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_351, 0, x_346); -lean_ctor_set(x_351, 1, x_350); -x_352 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_352, 0, x_351); -lean_ctor_set(x_352, 1, x_107); -x_353 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__72; +x_338 = lean_name_mk_string(x_3, x_337); +x_339 = l_Lean_Elab_Term_Quotation_ArrayStxBuilder_push___closed__3; +lean_inc(x_338); +x_340 = lean_name_mk_string(x_338, x_339); +lean_inc(x_253); lean_inc(x_258); -x_354 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_354, 0, x_258); -lean_ctor_set(x_354, 1, x_353); -lean_ctor_set(x_354, 2, x_349); -lean_ctor_set(x_354, 3, x_352); -x_355 = lean_array_push(x_282, x_354); -x_356 = lean_array_push(x_355, x_335); -x_357 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_357, 0, x_285); -lean_ctor_set(x_357, 1, x_295); -lean_ctor_set(x_357, 2, x_356); -x_358 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; -x_359 = lean_array_push(x_358, x_324); -lean_inc(x_359); -x_360 = lean_array_push(x_359, x_342); -lean_inc(x_344); -x_361 = lean_array_push(x_360, x_344); -x_362 = lean_array_push(x_361, x_357); -x_363 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__24; -x_364 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_364, 0, x_285); -lean_ctor_set(x_364, 1, x_363); -lean_ctor_set(x_364, 2, x_362); -x_365 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__76; -lean_inc(x_260); -lean_inc(x_265); -x_366 = l_Lean_addMacroScope(x_265, x_365, x_260); -x_367 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__75; -x_368 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__79; +x_341 = l_Lean_addMacroScope(x_258, x_340, x_253); +x_342 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__73; +lean_inc(x_338); +x_343 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_343, 0, x_338); +lean_ctor_set(x_343, 1, x_342); +x_344 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_344, 0, x_343); +lean_ctor_set(x_344, 1, x_107); +x_345 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__72; +lean_inc(x_251); +x_346 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_346, 0, x_251); +lean_ctor_set(x_346, 1, x_345); +lean_ctor_set(x_346, 2, x_341); +lean_ctor_set(x_346, 3, x_344); +x_347 = lean_array_push(x_275, x_346); +x_348 = lean_array_push(x_347, x_327); +x_349 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_349, 0, x_278); +lean_ctor_set(x_349, 1, x_287); +lean_ctor_set(x_349, 2, x_348); +x_350 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; +x_351 = lean_array_push(x_350, x_316); +lean_inc(x_351); +x_352 = lean_array_push(x_351, x_334); +lean_inc(x_336); +x_353 = lean_array_push(x_352, x_336); +x_354 = lean_array_push(x_353, x_349); +x_355 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__24; +x_356 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_356, 0, x_278); +lean_ctor_set(x_356, 1, x_355); +lean_ctor_set(x_356, 2, x_354); +x_357 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__76; +lean_inc(x_253); lean_inc(x_258); -x_369 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_369, 0, x_258); -lean_ctor_set(x_369, 1, x_367); -lean_ctor_set(x_369, 2, x_366); -lean_ctor_set(x_369, 3, x_368); -x_370 = lean_array_push(x_289, x_369); -x_371 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_371, 0, x_285); -lean_ctor_set(x_371, 1, x_291); -lean_ctor_set(x_371, 2, x_370); -x_372 = lean_array_push(x_289, x_371); -x_373 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_373, 0, x_285); -lean_ctor_set(x_373, 1, x_291); -lean_ctor_set(x_373, 2, x_372); -lean_inc(x_346); -x_374 = l_Lean_addMacroScope(x_265, x_346, x_260); -x_375 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_375, 0, x_346); -lean_ctor_set(x_375, 1, x_107); -x_376 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_376, 0, x_375); -lean_ctor_set(x_376, 1, x_107); -x_377 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__28; -x_378 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_378, 0, x_258); -lean_ctor_set(x_378, 1, x_377); -lean_ctor_set(x_378, 2, x_374); -lean_ctor_set(x_378, 3, x_376); -x_379 = lean_array_push(x_359, x_373); -x_380 = lean_array_push(x_379, x_344); -x_381 = lean_array_push(x_380, x_378); -x_382 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_382, 0, x_285); -lean_ctor_set(x_382, 1, x_363); -lean_ctor_set(x_382, 2, x_381); -x_383 = lean_array_push(x_282, x_364); -x_384 = lean_array_push(x_383, x_382); -x_385 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_385, 0, x_285); -lean_ctor_set(x_385, 1, x_291); -lean_ctor_set(x_385, 2, x_384); -x_386 = lean_array_push(x_289, x_385); -x_387 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__22; -x_388 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_388, 0, x_285); -lean_ctor_set(x_388, 1, x_387); -lean_ctor_set(x_388, 2, x_386); -x_389 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__30; -x_390 = lean_array_push(x_389, x_267); -x_391 = lean_array_push(x_390, x_298); -x_392 = lean_array_push(x_391, x_298); -x_393 = lean_array_push(x_392, x_320); -x_394 = lean_array_push(x_393, x_322); -x_395 = lean_array_push(x_394, x_388); -x_396 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__17; -x_397 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_397, 0, x_285); -lean_ctor_set(x_397, 1, x_396); -lean_ctor_set(x_397, 2, x_395); -lean_ctor_set(x_261, 0, x_397); -x_29 = x_261; +x_358 = l_Lean_addMacroScope(x_258, x_357, x_253); +x_359 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__75; +x_360 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__79; +lean_inc(x_251); +x_361 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_361, 0, x_251); +lean_ctor_set(x_361, 1, x_359); +lean_ctor_set(x_361, 2, x_358); +lean_ctor_set(x_361, 3, x_360); +x_362 = lean_array_push(x_282, x_361); +x_363 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_363, 0, x_278); +lean_ctor_set(x_363, 1, x_39); +lean_ctor_set(x_363, 2, x_362); +x_364 = lean_array_push(x_282, x_363); +x_365 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_365, 0, x_278); +lean_ctor_set(x_365, 1, x_39); +lean_ctor_set(x_365, 2, x_364); +lean_inc(x_338); +x_366 = l_Lean_addMacroScope(x_258, x_338, x_253); +x_367 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_367, 0, x_338); +lean_ctor_set(x_367, 1, x_107); +x_368 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_368, 0, x_367); +lean_ctor_set(x_368, 1, x_107); +x_369 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__28; +x_370 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_370, 0, x_251); +lean_ctor_set(x_370, 1, x_369); +lean_ctor_set(x_370, 2, x_366); +lean_ctor_set(x_370, 3, x_368); +x_371 = lean_array_push(x_351, x_365); +x_372 = lean_array_push(x_371, x_336); +x_373 = lean_array_push(x_372, x_370); +x_374 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_374, 0, x_278); +lean_ctor_set(x_374, 1, x_355); +lean_ctor_set(x_374, 2, x_373); +x_375 = lean_array_push(x_275, x_356); +x_376 = lean_array_push(x_375, x_374); +x_377 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_377, 0, x_278); +lean_ctor_set(x_377, 1, x_39); +lean_ctor_set(x_377, 2, x_376); +x_378 = lean_array_push(x_282, x_377); +x_379 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__22; +x_380 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_380, 0, x_278); +lean_ctor_set(x_380, 1, x_379); +lean_ctor_set(x_380, 2, x_378); +x_381 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__30; +x_382 = lean_array_push(x_381, x_260); +x_383 = lean_array_push(x_382, x_290); +x_384 = lean_array_push(x_383, x_290); +x_385 = lean_array_push(x_384, x_312); +x_386 = lean_array_push(x_385, x_314); +x_387 = lean_array_push(x_386, x_380); +x_388 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__17; +x_389 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_389, 0, x_278); +lean_ctor_set(x_389, 1, x_388); +lean_ctor_set(x_389, 2, x_387); +lean_ctor_set(x_254, 0, x_389); +x_29 = x_254; goto block_38; } else { -lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; -x_398 = lean_ctor_get(x_261, 0); -x_399 = lean_ctor_get(x_261, 1); -lean_inc(x_399); -lean_inc(x_398); -lean_dec(x_261); -x_400 = lean_ctor_get(x_398, 0); -lean_inc(x_400); -lean_dec(x_398); -x_401 = lean_environment_main_module(x_400); -x_402 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__16; -lean_inc(x_258); -x_403 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_403, 0, x_258); -lean_ctor_set(x_403, 1, x_402); -x_404 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__50; -lean_inc(x_260); -lean_inc(x_401); -x_405 = l_Lean_addMacroScope(x_401, x_404, x_260); -x_406 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__49; -x_407 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__52; -lean_inc(x_258); -x_408 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_408, 0, x_258); -lean_ctor_set(x_408, 1, x_406); -lean_ctor_set(x_408, 2, x_405); -lean_ctor_set(x_408, 3, x_407); -x_409 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; -lean_inc(x_258); -x_410 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_410, 0, x_258); -lean_ctor_set(x_410, 1, x_409); -x_411 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__21; -lean_inc(x_258); -x_412 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_412, 0, x_258); -lean_ctor_set(x_412, 1, x_411); -x_413 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__58; -lean_inc(x_260); -lean_inc(x_401); -x_414 = l_Lean_addMacroScope(x_401, x_413, x_260); -x_415 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__55; -x_416 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__62; -lean_inc(x_258); -x_417 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_417, 0, x_258); -lean_ctor_set(x_417, 1, x_415); -lean_ctor_set(x_417, 2, x_414); -lean_ctor_set(x_417, 3, x_416); -x_418 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_419 = lean_array_push(x_418, x_412); -x_420 = lean_array_push(x_419, x_417); -x_421 = lean_box(2); -x_422 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__20; +lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; +x_390 = lean_ctor_get(x_254, 0); +x_391 = lean_ctor_get(x_254, 1); +lean_inc(x_391); +lean_inc(x_390); +lean_dec(x_254); +x_392 = lean_ctor_get(x_390, 0); +lean_inc(x_392); +lean_dec(x_390); +x_393 = lean_environment_main_module(x_392); +x_394 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__16; +lean_inc(x_251); +x_395 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_395, 0, x_251); +lean_ctor_set(x_395, 1, x_394); +x_396 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__50; +lean_inc(x_253); +lean_inc(x_393); +x_397 = l_Lean_addMacroScope(x_393, x_396, x_253); +x_398 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__49; +x_399 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__52; +lean_inc(x_251); +x_400 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_400, 0, x_251); +lean_ctor_set(x_400, 1, x_398); +lean_ctor_set(x_400, 2, x_397); +lean_ctor_set(x_400, 3, x_399); +x_401 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; +lean_inc(x_251); +x_402 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_402, 0, x_251); +lean_ctor_set(x_402, 1, x_401); +x_403 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__21; +lean_inc(x_251); +x_404 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_404, 0, x_251); +lean_ctor_set(x_404, 1, x_403); +x_405 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__58; +lean_inc(x_253); +lean_inc(x_393); +x_406 = l_Lean_addMacroScope(x_393, x_405, x_253); +x_407 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__55; +x_408 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__62; +lean_inc(x_251); +x_409 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_409, 0, x_251); +lean_ctor_set(x_409, 1, x_407); +lean_ctor_set(x_409, 2, x_406); +lean_ctor_set(x_409, 3, x_408); +x_410 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_411 = lean_array_push(x_410, x_404); +x_412 = lean_array_push(x_411, x_409); +x_413 = lean_box(2); +x_414 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__20; +x_415 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_415, 0, x_413); +lean_ctor_set(x_415, 1, x_414); +lean_ctor_set(x_415, 2, x_412); +x_416 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24(x_108); +x_417 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; +x_418 = lean_array_push(x_417, x_416); +x_419 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_419, 0, x_413); +lean_ctor_set(x_419, 1, x_39); +lean_ctor_set(x_419, 2, x_418); +x_420 = lean_array_push(x_410, x_415); +x_421 = lean_array_push(x_420, x_419); +x_422 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; x_423 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_423, 0, x_421); +lean_ctor_set(x_423, 0, x_413); lean_ctor_set(x_423, 1, x_422); -lean_ctor_set(x_423, 2, x_420); -x_424 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8(x_108); -x_425 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; -x_426 = lean_array_push(x_425, x_424); -x_427 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_428 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_428, 0, x_421); -lean_ctor_set(x_428, 1, x_427); -lean_ctor_set(x_428, 2, x_426); -x_429 = lean_array_push(x_418, x_423); -x_430 = lean_array_push(x_429, x_428); -x_431 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_432 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_432, 0, x_421); -lean_ctor_set(x_432, 1, x_431); -lean_ctor_set(x_432, 2, x_430); -x_433 = lean_array_push(x_418, x_432); -x_434 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; -x_435 = lean_array_push(x_433, x_434); -x_436 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_436, 0, x_421); -lean_ctor_set(x_436, 1, x_427); -lean_ctor_set(x_436, 2, x_435); -x_437 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; -lean_inc(x_258); -x_438 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_438, 0, x_258); -lean_ctor_set(x_438, 1, x_437); -x_439 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; -x_440 = lean_array_push(x_439, x_410); -x_441 = lean_array_push(x_440, x_436); -x_442 = lean_array_push(x_441, x_438); -x_443 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__6; -x_444 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_444, 0, x_421); -lean_ctor_set(x_444, 1, x_443); -lean_ctor_set(x_444, 2, x_442); -x_445 = lean_array_push(x_418, x_444); -x_446 = lean_array_push(x_445, x_110); +lean_ctor_set(x_423, 2, x_421); +x_424 = lean_array_push(x_410, x_423); +x_425 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; +x_426 = lean_array_push(x_424, x_425); +x_427 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_427, 0, x_413); +lean_ctor_set(x_427, 1, x_39); +lean_ctor_set(x_427, 2, x_426); +x_428 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; +lean_inc(x_251); +x_429 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_429, 0, x_251); +lean_ctor_set(x_429, 1, x_428); +x_430 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; +x_431 = lean_array_push(x_430, x_402); +x_432 = lean_array_push(x_431, x_427); +x_433 = lean_array_push(x_432, x_429); +x_434 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6; +x_435 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_435, 0, x_413); +lean_ctor_set(x_435, 1, x_434); +lean_ctor_set(x_435, 2, x_433); +x_436 = lean_array_push(x_410, x_435); +x_437 = lean_array_push(x_436, x_110); +x_438 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_438, 0, x_413); +lean_ctor_set(x_438, 1, x_39); +lean_ctor_set(x_438, 2, x_437); +x_439 = lean_array_push(x_410, x_400); +x_440 = lean_array_push(x_439, x_438); +x_441 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_441, 0, x_413); +lean_ctor_set(x_441, 1, x_422); +lean_ctor_set(x_441, 2, x_440); +x_442 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__63; +x_443 = lean_array_push(x_442, x_441); +x_444 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__46; +x_445 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_445, 0, x_413); +lean_ctor_set(x_445, 1, x_444); +lean_ctor_set(x_445, 2, x_443); +x_446 = lean_array_push(x_417, x_445); x_447 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_447, 0, x_421); -lean_ctor_set(x_447, 1, x_427); +lean_ctor_set(x_447, 0, x_413); +lean_ctor_set(x_447, 1, x_39); lean_ctor_set(x_447, 2, x_446); -x_448 = lean_array_push(x_418, x_408); -x_449 = lean_array_push(x_448, x_447); -x_450 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_450, 0, x_421); -lean_ctor_set(x_450, 1, x_431); -lean_ctor_set(x_450, 2, x_449); -x_451 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__63; -x_452 = lean_array_push(x_451, x_450); -x_453 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__46; -x_454 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_454, 0, x_421); -lean_ctor_set(x_454, 1, x_453); -lean_ctor_set(x_454, 2, x_452); -x_455 = lean_array_push(x_425, x_454); -x_456 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_456, 0, x_421); -lean_ctor_set(x_456, 1, x_427); -lean_ctor_set(x_456, 2, x_455); -x_457 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__20; -lean_inc(x_258); -x_458 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_458, 0, x_258); -lean_ctor_set(x_458, 1, x_457); -x_459 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__25; -lean_inc(x_258); -x_460 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_460, 0, x_258); +x_448 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__20; +lean_inc(x_251); +x_449 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_449, 0, x_251); +lean_ctor_set(x_449, 1, x_448); +x_450 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__25; +lean_inc(x_251); +x_451 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_451, 0, x_251); +lean_ctor_set(x_451, 1, x_450); +x_452 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__4; +lean_inc(x_253); +lean_inc(x_393); +x_453 = l_Lean_addMacroScope(x_393, x_452, x_253); +x_454 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__3; +x_455 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__65; +lean_inc(x_251); +x_456 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_456, 0, x_251); +lean_ctor_set(x_456, 1, x_454); +lean_ctor_set(x_456, 2, x_453); +lean_ctor_set(x_456, 3, x_455); +x_457 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__69; +lean_inc(x_253); +lean_inc(x_393); +x_458 = l_Lean_addMacroScope(x_393, x_457, x_253); +x_459 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__68; +lean_inc(x_251); +x_460 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_460, 0, x_251); lean_ctor_set(x_460, 1, x_459); -x_461 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__4; -lean_inc(x_260); -lean_inc(x_401); -x_462 = l_Lean_addMacroScope(x_401, x_461, x_260); -x_463 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__3; -x_464 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__65; -lean_inc(x_258); -x_465 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_465, 0, x_258); -lean_ctor_set(x_465, 1, x_463); -lean_ctor_set(x_465, 2, x_462); -lean_ctor_set(x_465, 3, x_464); -x_466 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__69; -lean_inc(x_260); -lean_inc(x_401); -x_467 = l_Lean_addMacroScope(x_401, x_466, x_260); -x_468 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__68; -lean_inc(x_258); -x_469 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_469, 0, x_258); -lean_ctor_set(x_469, 1, x_468); -lean_ctor_set(x_469, 2, x_467); -lean_ctor_set(x_469, 3, x_107); -x_470 = lean_array_push(x_425, x_469); -x_471 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_471, 0, x_421); -lean_ctor_set(x_471, 1, x_427); -lean_ctor_set(x_471, 2, x_470); -x_472 = lean_array_push(x_418, x_465); -lean_inc(x_471); -x_473 = lean_array_push(x_472, x_471); -x_474 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_474, 0, x_421); -lean_ctor_set(x_474, 1, x_431); -lean_ctor_set(x_474, 2, x_473); -x_475 = lean_array_push(x_425, x_474); -x_476 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_476, 0, x_421); -lean_ctor_set(x_476, 1, x_427); -lean_ctor_set(x_476, 2, x_475); -x_477 = lean_array_push(x_425, x_476); -x_478 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_478, 0, x_421); -lean_ctor_set(x_478, 1, x_427); -lean_ctor_set(x_478, 2, x_477); -x_479 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__12; -lean_inc(x_258); -x_480 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_480, 0, x_258); -lean_ctor_set(x_480, 1, x_479); -x_481 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__29; +lean_ctor_set(x_460, 2, x_458); +lean_ctor_set(x_460, 3, x_107); +x_461 = lean_array_push(x_417, x_460); +x_462 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_462, 0, x_413); +lean_ctor_set(x_462, 1, x_39); +lean_ctor_set(x_462, 2, x_461); +x_463 = lean_array_push(x_410, x_456); +lean_inc(x_462); +x_464 = lean_array_push(x_463, x_462); +x_465 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_465, 0, x_413); +lean_ctor_set(x_465, 1, x_422); +lean_ctor_set(x_465, 2, x_464); +x_466 = lean_array_push(x_417, x_465); +x_467 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_467, 0, x_413); +lean_ctor_set(x_467, 1, x_39); +lean_ctor_set(x_467, 2, x_466); +x_468 = lean_array_push(x_417, x_467); +x_469 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_469, 0, x_413); +lean_ctor_set(x_469, 1, x_39); +lean_ctor_set(x_469, 2, x_468); +x_470 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__12; +lean_inc(x_251); +x_471 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_471, 0, x_251); +lean_ctor_set(x_471, 1, x_470); +x_472 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__29; lean_inc(x_3); -x_482 = lean_name_mk_string(x_3, x_481); -x_483 = l_Lean_Elab_Term_Quotation_ArrayStxBuilder_push___closed__3; -lean_inc(x_482); -x_484 = lean_name_mk_string(x_482, x_483); -lean_inc(x_260); -lean_inc(x_401); -x_485 = l_Lean_addMacroScope(x_401, x_484, x_260); -x_486 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__73; -lean_inc(x_482); -x_487 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_487, 0, x_482); -lean_ctor_set(x_487, 1, x_486); -x_488 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_488, 0, x_487); -lean_ctor_set(x_488, 1, x_107); -x_489 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__72; -lean_inc(x_258); -x_490 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_490, 0, x_258); -lean_ctor_set(x_490, 1, x_489); -lean_ctor_set(x_490, 2, x_485); -lean_ctor_set(x_490, 3, x_488); -x_491 = lean_array_push(x_418, x_490); -x_492 = lean_array_push(x_491, x_471); -x_493 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_493, 0, x_421); -lean_ctor_set(x_493, 1, x_431); -lean_ctor_set(x_493, 2, x_492); -x_494 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; -x_495 = lean_array_push(x_494, x_460); -lean_inc(x_495); -x_496 = lean_array_push(x_495, x_478); -lean_inc(x_480); -x_497 = lean_array_push(x_496, x_480); -x_498 = lean_array_push(x_497, x_493); -x_499 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__24; +x_473 = lean_name_mk_string(x_3, x_472); +x_474 = l_Lean_Elab_Term_Quotation_ArrayStxBuilder_push___closed__3; +lean_inc(x_473); +x_475 = lean_name_mk_string(x_473, x_474); +lean_inc(x_253); +lean_inc(x_393); +x_476 = l_Lean_addMacroScope(x_393, x_475, x_253); +x_477 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__73; +lean_inc(x_473); +x_478 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_478, 0, x_473); +lean_ctor_set(x_478, 1, x_477); +x_479 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_479, 0, x_478); +lean_ctor_set(x_479, 1, x_107); +x_480 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__72; +lean_inc(x_251); +x_481 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_481, 0, x_251); +lean_ctor_set(x_481, 1, x_480); +lean_ctor_set(x_481, 2, x_476); +lean_ctor_set(x_481, 3, x_479); +x_482 = lean_array_push(x_410, x_481); +x_483 = lean_array_push(x_482, x_462); +x_484 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_484, 0, x_413); +lean_ctor_set(x_484, 1, x_422); +lean_ctor_set(x_484, 2, x_483); +x_485 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; +x_486 = lean_array_push(x_485, x_451); +lean_inc(x_486); +x_487 = lean_array_push(x_486, x_469); +lean_inc(x_471); +x_488 = lean_array_push(x_487, x_471); +x_489 = lean_array_push(x_488, x_484); +x_490 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__24; +x_491 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_491, 0, x_413); +lean_ctor_set(x_491, 1, x_490); +lean_ctor_set(x_491, 2, x_489); +x_492 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__76; +lean_inc(x_253); +lean_inc(x_393); +x_493 = l_Lean_addMacroScope(x_393, x_492, x_253); +x_494 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__75; +x_495 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__79; +lean_inc(x_251); +x_496 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_496, 0, x_251); +lean_ctor_set(x_496, 1, x_494); +lean_ctor_set(x_496, 2, x_493); +lean_ctor_set(x_496, 3, x_495); +x_497 = lean_array_push(x_417, x_496); +x_498 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_498, 0, x_413); +lean_ctor_set(x_498, 1, x_39); +lean_ctor_set(x_498, 2, x_497); +x_499 = lean_array_push(x_417, x_498); x_500 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_500, 0, x_421); -lean_ctor_set(x_500, 1, x_499); -lean_ctor_set(x_500, 2, x_498); -x_501 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__76; -lean_inc(x_260); -lean_inc(x_401); -x_502 = l_Lean_addMacroScope(x_401, x_501, x_260); -x_503 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__75; -x_504 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__79; -lean_inc(x_258); +lean_ctor_set(x_500, 0, x_413); +lean_ctor_set(x_500, 1, x_39); +lean_ctor_set(x_500, 2, x_499); +lean_inc(x_473); +x_501 = l_Lean_addMacroScope(x_393, x_473, x_253); +x_502 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_502, 0, x_473); +lean_ctor_set(x_502, 1, x_107); +x_503 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_503, 0, x_502); +lean_ctor_set(x_503, 1, x_107); +x_504 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__28; x_505 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_505, 0, x_258); -lean_ctor_set(x_505, 1, x_503); -lean_ctor_set(x_505, 2, x_502); -lean_ctor_set(x_505, 3, x_504); -x_506 = lean_array_push(x_425, x_505); -x_507 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_507, 0, x_421); -lean_ctor_set(x_507, 1, x_427); -lean_ctor_set(x_507, 2, x_506); -x_508 = lean_array_push(x_425, x_507); +lean_ctor_set(x_505, 0, x_251); +lean_ctor_set(x_505, 1, x_504); +lean_ctor_set(x_505, 2, x_501); +lean_ctor_set(x_505, 3, x_503); +x_506 = lean_array_push(x_486, x_500); +x_507 = lean_array_push(x_506, x_471); +x_508 = lean_array_push(x_507, x_505); x_509 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_509, 0, x_421); -lean_ctor_set(x_509, 1, x_427); +lean_ctor_set(x_509, 0, x_413); +lean_ctor_set(x_509, 1, x_490); lean_ctor_set(x_509, 2, x_508); -lean_inc(x_482); -x_510 = l_Lean_addMacroScope(x_401, x_482, x_260); -x_511 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_511, 0, x_482); -lean_ctor_set(x_511, 1, x_107); -x_512 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_512, 0, x_511); -lean_ctor_set(x_512, 1, x_107); -x_513 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__28; -x_514 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_514, 0, x_258); -lean_ctor_set(x_514, 1, x_513); -lean_ctor_set(x_514, 2, x_510); -lean_ctor_set(x_514, 3, x_512); -x_515 = lean_array_push(x_495, x_509); -x_516 = lean_array_push(x_515, x_480); -x_517 = lean_array_push(x_516, x_514); -x_518 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_518, 0, x_421); -lean_ctor_set(x_518, 1, x_499); -lean_ctor_set(x_518, 2, x_517); -x_519 = lean_array_push(x_418, x_500); -x_520 = lean_array_push(x_519, x_518); -x_521 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_521, 0, x_421); -lean_ctor_set(x_521, 1, x_427); -lean_ctor_set(x_521, 2, x_520); -x_522 = lean_array_push(x_425, x_521); -x_523 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__22; +x_510 = lean_array_push(x_410, x_491); +x_511 = lean_array_push(x_510, x_509); +x_512 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_512, 0, x_413); +lean_ctor_set(x_512, 1, x_39); +lean_ctor_set(x_512, 2, x_511); +x_513 = lean_array_push(x_417, x_512); +x_514 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__22; +x_515 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_515, 0, x_413); +lean_ctor_set(x_515, 1, x_514); +lean_ctor_set(x_515, 2, x_513); +x_516 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__30; +x_517 = lean_array_push(x_516, x_395); +x_518 = lean_array_push(x_517, x_425); +x_519 = lean_array_push(x_518, x_425); +x_520 = lean_array_push(x_519, x_447); +x_521 = lean_array_push(x_520, x_449); +x_522 = lean_array_push(x_521, x_515); +x_523 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__17; x_524 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_524, 0, x_421); +lean_ctor_set(x_524, 0, x_413); lean_ctor_set(x_524, 1, x_523); lean_ctor_set(x_524, 2, x_522); -x_525 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__30; -x_526 = lean_array_push(x_525, x_403); -x_527 = lean_array_push(x_526, x_434); -x_528 = lean_array_push(x_527, x_434); -x_529 = lean_array_push(x_528, x_456); -x_530 = lean_array_push(x_529, x_458); -x_531 = lean_array_push(x_530, x_524); -x_532 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__17; -x_533 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_533, 0, x_421); -lean_ctor_set(x_533, 1, x_532); -lean_ctor_set(x_533, 2, x_531); -x_534 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_534, 0, x_533); -lean_ctor_set(x_534, 1, x_399); -x_29 = x_534; +x_525 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_525, 0, x_524); +lean_ctor_set(x_525, 1, x_391); +x_29 = x_525; goto block_38; } } } else { -lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; +lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_dec(x_115); lean_dec(x_110); lean_dec(x_108); -x_535 = lean_alloc_ctor(4, 1, 0); -lean_ctor_set(x_535, 0, x_114); -x_536 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__15; -x_537 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_537, 0, x_536); -lean_ctor_set(x_537, 1, x_535); -x_538 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__17; -x_539 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_539, 0, x_537); -lean_ctor_set(x_539, 1, x_538); -x_540 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__5; -x_541 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_541, 0, x_540); -lean_ctor_set(x_541, 1, x_539); -x_542 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__11; -x_543 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_543, 0, x_541); -lean_ctor_set(x_543, 1, x_542); +x_526 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_526, 0, x_114); +x_527 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__15; +x_528 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_528, 0, x_527); +lean_ctor_set(x_528, 1, x_526); +x_529 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__17; +x_530 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_530, 0, x_528); +lean_ctor_set(x_530, 1, x_529); +x_531 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__5; +x_532 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_532, 0, x_531); +lean_ctor_set(x_532, 1, x_530); +x_533 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__11; +x_534 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_534, 0, x_532); +lean_ctor_set(x_534, 1, x_533); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); -x_544 = l_Lean_throwErrorAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__10(x_19, x_543, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -x_29 = x_544; +x_535 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__22(x_19, x_534, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +x_29 = x_535; goto block_38; } } else { -lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; +lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_dec(x_110); lean_dec(x_108); -x_545 = lean_alloc_ctor(4, 1, 0); -lean_ctor_set(x_545, 0, x_114); -x_546 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__15; -x_547 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_547, 0, x_546); -lean_ctor_set(x_547, 1, x_545); -x_548 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__17; -x_549 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_549, 0, x_547); -lean_ctor_set(x_549, 1, x_548); -x_550 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__5; -x_551 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_551, 0, x_550); -lean_ctor_set(x_551, 1, x_549); -x_552 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__11; -x_553 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_553, 0, x_551); -lean_ctor_set(x_553, 1, x_552); +x_536 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_536, 0, x_114); +x_537 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__15; +x_538 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_538, 0, x_537); +lean_ctor_set(x_538, 1, x_536); +x_539 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__17; +x_540 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_540, 0, x_538); +lean_ctor_set(x_540, 1, x_539); +x_541 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__5; +x_542 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_542, 0, x_541); +lean_ctor_set(x_542, 1, x_540); +x_543 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__11; +x_544 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_544, 0, x_542); +lean_ctor_set(x_544, 1, x_543); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); -x_554 = l_Lean_throwErrorAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__10(x_19, x_553, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -x_29 = x_554; +x_545 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__22(x_19, x_544, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +x_29 = x_545; goto block_38; } } @@ -8567,7 +9823,7 @@ return x_37; } } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__11(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__27(lean_object* x_1) { _start: { if (lean_obj_tag(x_1) == 0) @@ -8583,7 +9839,7 @@ x_3 = lean_ctor_get(x_1, 0); x_4 = lean_ctor_get(x_1, 1); x_5 = lean_box(2); x_6 = l_Lean_Syntax_mkStrLit(x_3, x_5); -x_7 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__11(x_4); +x_7 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__27(x_4); x_8 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; x_9 = lean_array_push(x_8, x_6); x_10 = lean_array_push(x_9, x_7); @@ -8593,7 +9849,7 @@ return x_12; } } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__10(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__26(lean_object* x_1) { _start: { if (lean_obj_tag(x_1) == 0) @@ -8610,7 +9866,7 @@ lean_inc(x_3); x_4 = lean_ctor_get(x_1, 1); lean_inc(x_4); lean_dec(x_1); -x_5 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__10(x_4); +x_5 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__26(x_4); x_6 = lean_ctor_get(x_3, 0); lean_inc(x_6); x_7 = lean_ctor_get(x_3, 1); @@ -8619,12 +9875,12 @@ lean_dec(x_3); x_8 = lean_box(0); lean_inc(x_6); x_9 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_8, x_6); -x_10 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__11(x_7); +x_10 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__27(x_7); lean_dec(x_7); if (lean_obj_tag(x_9) == 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; lean_object* x_20; -x_11 = l___private_Init_Meta_0__Lean_quoteNameMk(x_6); +x_11 = l_Lean_quoteNameMk(x_6); x_12 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; x_13 = lean_array_push(x_12, x_11); x_14 = lean_array_push(x_13, x_10); @@ -8638,36 +9894,35 @@ return x_20; } else { -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; 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_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; 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_dec(x_6); x_21 = lean_ctor_get(x_9, 0); lean_inc(x_21); lean_dec(x_9); -x_22 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__3; +x_22 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; x_23 = l_String_intercalate(x_22, x_21); -x_24 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__4; +x_24 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4; x_25 = lean_string_append(x_24, x_23); lean_dec(x_23); -x_26 = l_Lean_nameLitKind; -x_27 = lean_box(2); -x_28 = l_Lean_Syntax_mkLit(x_26, x_25, x_27); -x_29 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; -x_30 = lean_array_push(x_29, x_28); -x_31 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__2; -x_32 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_32, 0, x_27); -lean_ctor_set(x_32, 1, x_31); -lean_ctor_set(x_32, 2, x_30); -x_33 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_34 = lean_array_push(x_33, x_32); -x_35 = lean_array_push(x_34, x_10); -x_36 = l_Lean_Elab_Term_Quotation_mkTuple___closed__7; -x_37 = l_Lean_Syntax_mkCApp(x_36, x_35); -x_38 = lean_array_push(x_33, x_37); -x_39 = lean_array_push(x_38, x_5); -x_40 = l___private_Init_Meta_0__Lean_quoteList___at_Lean_Elab_Term_Quotation_ArrayStxBuilder_build___spec__1___closed__7; -x_41 = l_Lean_Syntax_mkCApp(x_40, x_39); -return x_41; +x_26 = lean_box(2); +x_27 = l_Lean_Syntax_mkNameLit(x_25, x_26); +x_28 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; +x_29 = lean_array_push(x_28, x_27); +x_30 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2; +x_31 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_31, 0, x_26); +lean_ctor_set(x_31, 1, x_30); +lean_ctor_set(x_31, 2, x_29); +x_32 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_33 = lean_array_push(x_32, x_31); +x_34 = lean_array_push(x_33, x_10); +x_35 = l_Lean_Elab_Term_Quotation_mkTuple___closed__7; +x_36 = l_Lean_Syntax_mkCApp(x_35, x_34); +x_37 = lean_array_push(x_32, x_36); +x_38 = lean_array_push(x_37, x_5); +x_39 = l___private_Init_Meta_0__Lean_quoteList___at_Lean_Elab_Term_Quotation_ArrayStxBuilder_build___spec__1___closed__7; +x_40 = l_Lean_Syntax_mkCApp(x_39, x_38); +return x_40; } } } @@ -8708,7 +9963,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__29; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__29; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } @@ -8735,7 +9990,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__30; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__30; x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -9029,7 +10284,7 @@ lean_inc(x_171); x_172 = lean_ctor_get(x_170, 1); lean_inc(x_172); lean_dec(x_170); -x_173 = l___private_Init_Meta_0__Lean_quoteNameMk(x_1); +x_173 = l_Lean_quoteNameMk(x_1); x_19 = x_173; x_20 = x_171; x_21 = x_172; @@ -9037,7 +10292,7 @@ goto block_167; } else { -lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; +lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_dec(x_1); x_174 = lean_ctor_get(x_169, 0); lean_inc(x_174); @@ -9047,22 +10302,21 @@ lean_inc(x_175); x_176 = lean_ctor_get(x_170, 1); lean_inc(x_176); lean_dec(x_170); -x_177 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__3; +x_177 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; x_178 = l_String_intercalate(x_177, x_174); -x_179 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__4; +x_179 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4; x_180 = lean_string_append(x_179, x_178); lean_dec(x_178); -x_181 = l_Lean_nameLitKind; -x_182 = lean_box(2); -x_183 = l_Lean_Syntax_mkLit(x_181, x_180, x_182); -x_184 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; -x_185 = lean_array_push(x_184, x_183); -x_186 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__2; -x_187 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_187, 0, x_182); -lean_ctor_set(x_187, 1, x_186); -lean_ctor_set(x_187, 2, x_185); -x_19 = x_187; +x_181 = lean_box(2); +x_182 = l_Lean_Syntax_mkNameLit(x_180, x_181); +x_183 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; +x_184 = lean_array_push(x_183, x_182); +x_185 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2; +x_186 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_186, 0, x_181); +lean_ctor_set(x_186, 1, x_185); +lean_ctor_set(x_186, 2, x_184); +x_19 = x_186; x_20 = x_175; x_21 = x_176; goto block_167; @@ -9107,7 +10361,7 @@ lean_ctor_set(x_37, 0, x_20); lean_ctor_set(x_37, 1, x_36); lean_ctor_set(x_37, 2, x_35); lean_ctor_set(x_37, 3, x_30); -x_38 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; +x_38 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; lean_inc(x_20); x_39 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_39, 0, x_20); @@ -9144,7 +10398,7 @@ lean_ctor_set(x_52, 0, x_20); lean_ctor_set(x_52, 1, x_51); lean_ctor_set(x_52, 2, x_50); lean_ctor_set(x_52, 3, x_30); -x_53 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; +x_53 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; x_54 = lean_array_push(x_53, x_48); x_55 = lean_array_push(x_54, x_19); x_56 = lean_array_push(x_55, x_52); @@ -9169,19 +10423,19 @@ x_68 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_68, 0, x_57); lean_ctor_set(x_68, 1, x_58); lean_ctor_set(x_68, 2, x_67); -x_69 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; +x_69 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; x_70 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_70, 0, x_20); lean_ctor_set(x_70, 1, x_69); x_71 = lean_array_push(x_53, x_39); x_72 = lean_array_push(x_71, x_68); x_73 = lean_array_push(x_72, x_70); -x_74 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__6; +x_74 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6; x_75 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_75, 0, x_57); lean_ctor_set(x_75, 1, x_74); lean_ctor_set(x_75, 2, x_73); -x_76 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__10(x_18); +x_76 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__26(x_18); x_77 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; x_78 = lean_array_push(x_77, x_37); x_79 = lean_array_push(x_60, x_33); @@ -9246,7 +10500,7 @@ lean_ctor_set(x_108, 0, x_20); lean_ctor_set(x_108, 1, x_107); lean_ctor_set(x_108, 2, x_106); lean_ctor_set(x_108, 3, x_101); -x_109 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; +x_109 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; lean_inc(x_20); x_110 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_110, 0, x_20); @@ -9283,7 +10537,7 @@ lean_ctor_set(x_123, 0, x_20); lean_ctor_set(x_123, 1, x_122); lean_ctor_set(x_123, 2, x_121); lean_ctor_set(x_123, 3, x_101); -x_124 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; +x_124 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; x_125 = lean_array_push(x_124, x_119); x_126 = lean_array_push(x_125, x_19); x_127 = lean_array_push(x_126, x_123); @@ -9308,19 +10562,19 @@ x_139 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_139, 0, x_128); lean_ctor_set(x_139, 1, x_129); lean_ctor_set(x_139, 2, x_138); -x_140 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; +x_140 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; x_141 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_141, 0, x_20); lean_ctor_set(x_141, 1, x_140); x_142 = lean_array_push(x_124, x_110); x_143 = lean_array_push(x_142, x_139); x_144 = lean_array_push(x_143, x_141); -x_145 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__6; +x_145 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6; x_146 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_146, 0, x_128); lean_ctor_set(x_146, 1, x_145); lean_ctor_set(x_146, 2, x_144); -x_147 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__10(x_18); +x_147 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__26(x_18); x_148 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; x_149 = lean_array_push(x_148, x_108); x_150 = lean_array_push(x_131, x_104); @@ -9424,7 +10678,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__30; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__30; x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__6; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -9497,7 +10751,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__6; +x_1 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__6; x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__14; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -9572,7 +10826,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__30; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__30; x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__18; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -9673,7 +10927,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__30; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__30; x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__30; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -9757,7 +11011,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__39; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__6; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__6; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } @@ -9777,7 +11031,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__41; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__6; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__6; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } @@ -9844,30 +11098,6 @@ return x_2; static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__49() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__60; -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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__50() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__49; -x_3 = lean_alloc_ctor(1, 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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__51() { -_start: -{ lean_object* x_1; x_1 = l_Lean_Elab_Term_Quotation_hygiene; return x_1; @@ -9887,319 +11117,319 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_234 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13___rarg(x_8); +x_234 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___rarg(x_8); return x_234; } case 1: { -lean_object* x_235; lean_object* x_236; lean_object* x_367; uint8_t x_380; +lean_object* x_235; lean_object* x_236; lean_object* x_365; uint8_t x_378; x_235 = lean_ctor_get(x_1, 1); lean_inc(x_235); lean_inc(x_1); -x_380 = l_Lean_Syntax_isAntiquots(x_1); -if (x_380 == 0) +x_378 = l_Lean_Syntax_isAntiquots(x_1); +if (x_378 == 0) { -uint8_t x_381; +uint8_t x_379; lean_inc(x_1); -x_381 = l_Lean_Syntax_isTokenAntiquot(x_1); -if (x_381 == 0) +x_379 = l_Lean_Syntax_isTokenAntiquot(x_1); +if (x_379 == 0) { -lean_object* x_382; -x_382 = lean_box(0); -x_367 = x_382; -goto block_379; +lean_object* x_380; +x_380 = lean_box(0); +x_365 = x_380; +goto block_377; } else { -uint8_t x_383; -x_383 = l_Lean_Syntax_isEscapedAntiquot(x_1); -if (x_383 == 0) +uint8_t x_381; +x_381 = l_Lean_Syntax_isEscapedAntiquot(x_1); +if (x_381 == 0) { -lean_object* x_384; +lean_object* x_382; lean_dec(x_235); -x_384 = lean_box(0); -x_9 = x_384; +x_382 = lean_box(0); +x_9 = x_382; goto block_233; } else { -lean_object* x_385; -x_385 = lean_box(0); -x_367 = x_385; -goto block_379; +lean_object* x_383; +x_383 = lean_box(0); +x_365 = x_383; +goto block_377; } } } else { -lean_object* x_386; uint8_t x_387; +lean_object* x_384; uint8_t x_385; lean_inc(x_1); -x_386 = l_Lean_Syntax_getCanonicalAntiquot(x_1); -x_387 = l_Lean_Syntax_isEscapedAntiquot(x_386); -if (x_387 == 0) +x_384 = l_Lean_Syntax_getCanonicalAntiquot(x_1); +x_385 = l_Lean_Syntax_isEscapedAntiquot(x_384); +if (x_385 == 0) { -lean_object* x_388; uint8_t x_389; +lean_object* x_386; uint8_t x_387; lean_dec(x_235); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_inc(x_1); -x_388 = l_Lean_Syntax_antiquotKinds(x_1); -x_389 = !lean_is_exclusive(x_1); -if (x_389 == 0) +x_386 = l_Lean_Syntax_antiquotKinds(x_1); +x_387 = !lean_is_exclusive(x_1); +if (x_387 == 0) { -lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; uint8_t x_398; -x_390 = lean_ctor_get(x_1, 2); +lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; uint8_t x_396; +x_388 = lean_ctor_get(x_1, 2); +lean_dec(x_388); +x_389 = lean_ctor_get(x_1, 1); +lean_dec(x_389); +x_390 = lean_ctor_get(x_1, 0); lean_dec(x_390); -x_391 = lean_ctor_get(x_1, 1); -lean_dec(x_391); -x_392 = lean_ctor_get(x_1, 0); -lean_dec(x_392); lean_inc(x_6); -x_393 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_6, x_7, x_8); -x_394 = lean_ctor_get(x_393, 0); +x_391 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_6, x_7, x_8); +x_392 = lean_ctor_get(x_391, 0); +lean_inc(x_392); +x_393 = lean_ctor_get(x_391, 1); +lean_inc(x_393); +lean_dec(x_391); +x_394 = lean_ctor_get(x_6, 10); lean_inc(x_394); -x_395 = lean_ctor_get(x_393, 1); -lean_inc(x_395); -lean_dec(x_393); -x_396 = lean_ctor_get(x_6, 10); -lean_inc(x_396); lean_dec(x_6); -x_397 = lean_st_ref_get(x_7, x_395); +x_395 = lean_st_ref_get(x_7, x_393); lean_dec(x_7); -x_398 = !lean_is_exclusive(x_397); -if (x_398 == 0) +x_396 = !lean_is_exclusive(x_395); +if (x_396 == 0) { -lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; -x_399 = lean_ctor_get(x_397, 0); -x_400 = lean_ctor_get(x_399, 0); -lean_inc(x_400); -lean_dec(x_399); -x_401 = lean_environment_main_module(x_400); -x_402 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__21; -lean_inc(x_394); -x_403 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_403, 0, x_394); -lean_ctor_set(x_403, 1, x_402); -x_404 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__58; -x_405 = l_Lean_addMacroScope(x_401, x_404, x_396); -x_406 = lean_box(0); -x_407 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__55; -x_408 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__50; -x_409 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_409, 0, x_394); -lean_ctor_set(x_409, 1, x_407); -lean_ctor_set(x_409, 2, x_405); -lean_ctor_set(x_409, 3, x_408); -x_410 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_411 = lean_array_push(x_410, x_403); -x_412 = lean_array_push(x_411, x_409); -x_413 = lean_box(2); -x_414 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__20; -lean_ctor_set(x_1, 2, x_412); -lean_ctor_set(x_1, 1, x_414); -lean_ctor_set(x_1, 0, x_413); -x_415 = l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__7(x_388, x_406); -x_416 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8(x_415); -x_417 = l_Lean_Syntax_getAntiquotTerm(x_386); -lean_dec(x_386); -x_418 = lean_array_push(x_410, x_416); -x_419 = lean_array_push(x_418, x_417); -x_420 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_421 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_421, 0, x_413); -lean_ctor_set(x_421, 1, x_420); -lean_ctor_set(x_421, 2, x_419); -x_422 = lean_array_push(x_410, x_1); -x_423 = lean_array_push(x_422, x_421); -x_424 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_425 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_425, 0, x_413); -lean_ctor_set(x_425, 1, x_424); -lean_ctor_set(x_425, 2, x_423); -lean_ctor_set(x_397, 0, x_425); -return x_397; +lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; +x_397 = lean_ctor_get(x_395, 0); +x_398 = lean_ctor_get(x_397, 0); +lean_inc(x_398); +lean_dec(x_397); +x_399 = lean_environment_main_module(x_398); +x_400 = lean_box(0); +x_401 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__21; +lean_inc(x_392); +x_402 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_402, 0, x_392); +lean_ctor_set(x_402, 1, x_401); +x_403 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__58; +x_404 = l_Lean_addMacroScope(x_399, x_403, x_394); +x_405 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__55; +x_406 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__62; +x_407 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_407, 0, x_392); +lean_ctor_set(x_407, 1, x_405); +lean_ctor_set(x_407, 2, x_404); +lean_ctor_set(x_407, 3, x_406); +x_408 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_409 = lean_array_push(x_408, x_402); +x_410 = lean_array_push(x_409, x_407); +x_411 = lean_box(2); +x_412 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__20; +lean_ctor_set(x_1, 2, x_410); +lean_ctor_set(x_1, 1, x_412); +lean_ctor_set(x_1, 0, x_411); +x_413 = l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__21(x_386, x_400); +x_414 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24(x_413); +x_415 = l_Lean_Syntax_getAntiquotTerm(x_384); +lean_dec(x_384); +x_416 = lean_array_push(x_408, x_414); +x_417 = lean_array_push(x_416, x_415); +x_418 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_419 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_419, 0, x_411); +lean_ctor_set(x_419, 1, x_418); +lean_ctor_set(x_419, 2, x_417); +x_420 = lean_array_push(x_408, x_1); +x_421 = lean_array_push(x_420, x_419); +x_422 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_423 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_423, 0, x_411); +lean_ctor_set(x_423, 1, x_422); +lean_ctor_set(x_423, 2, x_421); +lean_ctor_set(x_395, 0, x_423); +return x_395; } else { -lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; -x_426 = lean_ctor_get(x_397, 0); -x_427 = lean_ctor_get(x_397, 1); -lean_inc(x_427); +lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; +x_424 = lean_ctor_get(x_395, 0); +x_425 = lean_ctor_get(x_395, 1); +lean_inc(x_425); +lean_inc(x_424); +lean_dec(x_395); +x_426 = lean_ctor_get(x_424, 0); lean_inc(x_426); -lean_dec(x_397); -x_428 = lean_ctor_get(x_426, 0); -lean_inc(x_428); -lean_dec(x_426); -x_429 = lean_environment_main_module(x_428); -x_430 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__21; -lean_inc(x_394); -x_431 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_431, 0, x_394); -lean_ctor_set(x_431, 1, x_430); -x_432 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__58; -x_433 = l_Lean_addMacroScope(x_429, x_432, x_396); -x_434 = lean_box(0); -x_435 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__55; -x_436 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__50; -x_437 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_437, 0, x_394); -lean_ctor_set(x_437, 1, x_435); -lean_ctor_set(x_437, 2, x_433); -lean_ctor_set(x_437, 3, x_436); -x_438 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_439 = lean_array_push(x_438, x_431); -x_440 = lean_array_push(x_439, x_437); -x_441 = lean_box(2); -x_442 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__20; -lean_ctor_set(x_1, 2, x_440); -lean_ctor_set(x_1, 1, x_442); -lean_ctor_set(x_1, 0, x_441); -x_443 = l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__7(x_388, x_434); -x_444 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8(x_443); -x_445 = l_Lean_Syntax_getAntiquotTerm(x_386); -lean_dec(x_386); -x_446 = lean_array_push(x_438, x_444); -x_447 = lean_array_push(x_446, x_445); -x_448 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_449 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_449, 0, x_441); -lean_ctor_set(x_449, 1, x_448); -lean_ctor_set(x_449, 2, x_447); -x_450 = lean_array_push(x_438, x_1); -x_451 = lean_array_push(x_450, x_449); -x_452 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_453 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_453, 0, x_441); -lean_ctor_set(x_453, 1, x_452); -lean_ctor_set(x_453, 2, x_451); -x_454 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_454, 0, x_453); -lean_ctor_set(x_454, 1, x_427); -return x_454; +lean_dec(x_424); +x_427 = lean_environment_main_module(x_426); +x_428 = lean_box(0); +x_429 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__21; +lean_inc(x_392); +x_430 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_430, 0, x_392); +lean_ctor_set(x_430, 1, x_429); +x_431 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__58; +x_432 = l_Lean_addMacroScope(x_427, x_431, x_394); +x_433 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__55; +x_434 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__62; +x_435 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_435, 0, x_392); +lean_ctor_set(x_435, 1, x_433); +lean_ctor_set(x_435, 2, x_432); +lean_ctor_set(x_435, 3, x_434); +x_436 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_437 = lean_array_push(x_436, x_430); +x_438 = lean_array_push(x_437, x_435); +x_439 = lean_box(2); +x_440 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__20; +lean_ctor_set(x_1, 2, x_438); +lean_ctor_set(x_1, 1, x_440); +lean_ctor_set(x_1, 0, x_439); +x_441 = l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__21(x_386, x_428); +x_442 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24(x_441); +x_443 = l_Lean_Syntax_getAntiquotTerm(x_384); +lean_dec(x_384); +x_444 = lean_array_push(x_436, x_442); +x_445 = lean_array_push(x_444, x_443); +x_446 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_447 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_447, 0, x_439); +lean_ctor_set(x_447, 1, x_446); +lean_ctor_set(x_447, 2, x_445); +x_448 = lean_array_push(x_436, x_1); +x_449 = lean_array_push(x_448, x_447); +x_450 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_451 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_451, 0, x_439); +lean_ctor_set(x_451, 1, x_450); +lean_ctor_set(x_451, 2, x_449); +x_452 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_452, 0, x_451); +lean_ctor_set(x_452, 1, x_425); +return x_452; } } else { -lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; +lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_dec(x_1); lean_inc(x_6); -x_455 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_6, x_7, x_8); -x_456 = lean_ctor_get(x_455, 0); +x_453 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_6, x_7, x_8); +x_454 = lean_ctor_get(x_453, 0); +lean_inc(x_454); +x_455 = lean_ctor_get(x_453, 1); +lean_inc(x_455); +lean_dec(x_453); +x_456 = lean_ctor_get(x_6, 10); lean_inc(x_456); -x_457 = lean_ctor_get(x_455, 1); -lean_inc(x_457); -lean_dec(x_455); -x_458 = lean_ctor_get(x_6, 10); -lean_inc(x_458); lean_dec(x_6); -x_459 = lean_st_ref_get(x_7, x_457); +x_457 = lean_st_ref_get(x_7, x_455); lean_dec(x_7); -x_460 = lean_ctor_get(x_459, 0); -lean_inc(x_460); -x_461 = lean_ctor_get(x_459, 1); -lean_inc(x_461); -if (lean_is_exclusive(x_459)) { - lean_ctor_release(x_459, 0); - lean_ctor_release(x_459, 1); - x_462 = x_459; +x_458 = lean_ctor_get(x_457, 0); +lean_inc(x_458); +x_459 = lean_ctor_get(x_457, 1); +lean_inc(x_459); +if (lean_is_exclusive(x_457)) { + lean_ctor_release(x_457, 0); + lean_ctor_release(x_457, 1); + x_460 = x_457; } else { - lean_dec_ref(x_459); - x_462 = lean_box(0); + lean_dec_ref(x_457); + x_460 = lean_box(0); } -x_463 = lean_ctor_get(x_460, 0); -lean_inc(x_463); -lean_dec(x_460); -x_464 = lean_environment_main_module(x_463); -x_465 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__21; -lean_inc(x_456); -x_466 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_466, 0, x_456); -lean_ctor_set(x_466, 1, x_465); -x_467 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__58; -x_468 = l_Lean_addMacroScope(x_464, x_467, x_458); -x_469 = lean_box(0); -x_470 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__55; -x_471 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__50; -x_472 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_472, 0, x_456); -lean_ctor_set(x_472, 1, x_470); -lean_ctor_set(x_472, 2, x_468); -lean_ctor_set(x_472, 3, x_471); -x_473 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_474 = lean_array_push(x_473, x_466); -x_475 = lean_array_push(x_474, x_472); -x_476 = lean_box(2); -x_477 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__20; -x_478 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_478, 0, x_476); -lean_ctor_set(x_478, 1, x_477); -lean_ctor_set(x_478, 2, x_475); -x_479 = l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__7(x_388, x_469); -x_480 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8(x_479); -x_481 = l_Lean_Syntax_getAntiquotTerm(x_386); -lean_dec(x_386); -x_482 = lean_array_push(x_473, x_480); -x_483 = lean_array_push(x_482, x_481); -x_484 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_485 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_485, 0, x_476); -lean_ctor_set(x_485, 1, x_484); -lean_ctor_set(x_485, 2, x_483); -x_486 = lean_array_push(x_473, x_478); -x_487 = lean_array_push(x_486, x_485); -x_488 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_489 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_489, 0, x_476); -lean_ctor_set(x_489, 1, x_488); -lean_ctor_set(x_489, 2, x_487); -if (lean_is_scalar(x_462)) { - x_490 = lean_alloc_ctor(0, 2, 0); +x_461 = lean_ctor_get(x_458, 0); +lean_inc(x_461); +lean_dec(x_458); +x_462 = lean_environment_main_module(x_461); +x_463 = lean_box(0); +x_464 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__21; +lean_inc(x_454); +x_465 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_465, 0, x_454); +lean_ctor_set(x_465, 1, x_464); +x_466 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__58; +x_467 = l_Lean_addMacroScope(x_462, x_466, x_456); +x_468 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__55; +x_469 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__62; +x_470 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_470, 0, x_454); +lean_ctor_set(x_470, 1, x_468); +lean_ctor_set(x_470, 2, x_467); +lean_ctor_set(x_470, 3, x_469); +x_471 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_472 = lean_array_push(x_471, x_465); +x_473 = lean_array_push(x_472, x_470); +x_474 = lean_box(2); +x_475 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__20; +x_476 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_476, 0, x_474); +lean_ctor_set(x_476, 1, x_475); +lean_ctor_set(x_476, 2, x_473); +x_477 = l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__21(x_386, x_463); +x_478 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24(x_477); +x_479 = l_Lean_Syntax_getAntiquotTerm(x_384); +lean_dec(x_384); +x_480 = lean_array_push(x_471, x_478); +x_481 = lean_array_push(x_480, x_479); +x_482 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_483 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_483, 0, x_474); +lean_ctor_set(x_483, 1, x_482); +lean_ctor_set(x_483, 2, x_481); +x_484 = lean_array_push(x_471, x_476); +x_485 = lean_array_push(x_484, x_483); +x_486 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_487 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_487, 0, x_474); +lean_ctor_set(x_487, 1, x_486); +lean_ctor_set(x_487, 2, x_485); +if (lean_is_scalar(x_460)) { + x_488 = lean_alloc_ctor(0, 2, 0); } else { - x_490 = x_462; + x_488 = x_460; } -lean_ctor_set(x_490, 0, x_489); -lean_ctor_set(x_490, 1, x_461); -return x_490; +lean_ctor_set(x_488, 0, x_487); +lean_ctor_set(x_488, 1, x_459); +return x_488; } } else { -uint8_t x_491; -lean_dec(x_386); +uint8_t x_489; +lean_dec(x_384); lean_inc(x_1); -x_491 = l_Lean_Syntax_isTokenAntiquot(x_1); -if (x_491 == 0) +x_489 = l_Lean_Syntax_isTokenAntiquot(x_1); +if (x_489 == 0) { -lean_object* x_492; -x_492 = lean_box(0); -x_367 = x_492; -goto block_379; +lean_object* x_490; +x_490 = lean_box(0); +x_365 = x_490; +goto block_377; } else { -uint8_t x_493; -x_493 = l_Lean_Syntax_isEscapedAntiquot(x_1); -if (x_493 == 0) +uint8_t x_491; +x_491 = l_Lean_Syntax_isEscapedAntiquot(x_1); +if (x_491 == 0) { -lean_object* x_494; +lean_object* x_492; lean_dec(x_235); -x_494 = lean_box(0); -x_9 = x_494; +x_492 = lean_box(0); +x_9 = x_492; goto block_233; } else { -lean_object* x_495; -x_495 = lean_box(0); -x_367 = x_495; -goto block_379; +lean_object* x_493; +x_493 = lean_box(0); +x_365 = x_493; +goto block_377; } } } } -block_366: +block_364: { lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; uint8_t x_244; lean_object* x_245; lean_object* x_246; size_t x_247; size_t x_248; lean_object* x_249; lean_dec(x_236); @@ -10232,24 +11462,24 @@ lean_dec(x_246); x_248 = 0; if (x_244 == 0) { -lean_object* x_365; -x_365 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__46; -x_249 = x_365; -goto block_364; +lean_object* x_363; +x_363 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__46; +x_249 = x_363; +goto block_362; } else { x_249 = x_243; -goto block_364; +goto block_362; } -block_364: +block_362: { lean_object* x_250; lean_object* x_251; lean_object* x_252; x_250 = l_Lean_Elab_Term_Quotation_ArrayStxBuilder_push___closed__2; x_251 = l_Lean_Elab_Term_Quotation_ArrayStxBuilder_empty; lean_inc(x_7); lean_inc(x_6); -x_252 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9(x_235, x_237, x_250, x_248, x_249, x_245, x_247, x_248, x_251, x_2, x_3, x_4, x_5, x_6, x_7, x_241); +x_252 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25(x_235, x_237, x_250, x_248, x_249, x_245, x_247, x_248, x_251, x_2, x_3, x_4, x_5, x_6, x_7, x_241); lean_dec(x_245); lean_dec(x_237); if (lean_obj_tag(x_252) == 0) @@ -10306,14 +11536,14 @@ lean_ctor_set(x_274, 3, x_273); lean_inc(x_235); x_275 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_266, x_235); x_276 = l_Lean_Elab_Term_Quotation_ArrayStxBuilder_build(x_253); -x_277 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; +x_277 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; x_278 = lean_array_push(x_277, x_274); x_279 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; x_280 = lean_array_push(x_279, x_269); if (lean_obj_tag(x_275) == 0) { lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; -x_281 = l___private_Init_Meta_0__Lean_quoteNameMk(x_235); +x_281 = l_Lean_quoteNameMk(x_235); x_282 = lean_array_push(x_278, x_281); x_283 = lean_array_push(x_282, x_276); x_284 = lean_box(2); @@ -10337,839 +11567,834 @@ return x_259; } else { -lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; +lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_dec(x_235); x_290 = lean_ctor_get(x_275, 0); lean_inc(x_290); lean_dec(x_275); -x_291 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__3; +x_291 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; x_292 = l_String_intercalate(x_291, x_290); -x_293 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__4; +x_293 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4; x_294 = lean_string_append(x_293, x_292); lean_dec(x_292); -x_295 = l_Lean_nameLitKind; -x_296 = lean_box(2); -x_297 = l_Lean_Syntax_mkLit(x_295, x_294, x_296); -x_298 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; -x_299 = lean_array_push(x_298, x_297); -x_300 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__2; +x_295 = lean_box(2); +x_296 = l_Lean_Syntax_mkNameLit(x_294, x_295); +x_297 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; +x_298 = lean_array_push(x_297, x_296); +x_299 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2; if (lean_is_scalar(x_238)) { - x_301 = lean_alloc_ctor(1, 3, 0); + x_300 = lean_alloc_ctor(1, 3, 0); } else { - x_301 = x_238; + x_300 = x_238; } -lean_ctor_set(x_301, 0, x_296); -lean_ctor_set(x_301, 1, x_300); -lean_ctor_set(x_301, 2, x_299); -x_302 = lean_array_push(x_278, x_301); -x_303 = lean_array_push(x_302, x_276); -x_304 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_305 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_305, 0, x_296); -lean_ctor_set(x_305, 1, x_304); -lean_ctor_set(x_305, 2, x_303); -x_306 = lean_array_push(x_280, x_305); -x_307 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_308 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_308, 0, x_296); -lean_ctor_set(x_308, 1, x_307); -lean_ctor_set(x_308, 2, x_306); -lean_ctor_set(x_259, 0, x_308); +lean_ctor_set(x_300, 0, x_295); +lean_ctor_set(x_300, 1, x_299); +lean_ctor_set(x_300, 2, x_298); +x_301 = lean_array_push(x_278, x_300); +x_302 = lean_array_push(x_301, x_276); +x_303 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_304 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_304, 0, x_295); +lean_ctor_set(x_304, 1, x_303); +lean_ctor_set(x_304, 2, x_302); +x_305 = lean_array_push(x_280, x_304); +x_306 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_307 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_307, 0, x_295); +lean_ctor_set(x_307, 1, x_306); +lean_ctor_set(x_307, 2, x_305); +lean_ctor_set(x_259, 0, x_307); return x_259; } } else { -lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; -x_309 = lean_ctor_get(x_259, 0); -x_310 = lean_ctor_get(x_259, 1); -lean_inc(x_310); +lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; +x_308 = lean_ctor_get(x_259, 0); +x_309 = lean_ctor_get(x_259, 1); lean_inc(x_309); +lean_inc(x_308); lean_dec(x_259); -x_311 = lean_ctor_get(x_309, 0); -lean_inc(x_311); -lean_dec(x_309); -x_312 = lean_environment_main_module(x_311); -x_313 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__31; +x_310 = lean_ctor_get(x_308, 0); +lean_inc(x_310); +lean_dec(x_308); +x_311 = lean_environment_main_module(x_310); +x_312 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__31; lean_inc(x_258); -lean_inc(x_312); -x_314 = l_Lean_addMacroScope(x_312, x_313, x_258); -x_315 = lean_box(0); -x_316 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__29; -x_317 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__34; +lean_inc(x_311); +x_313 = l_Lean_addMacroScope(x_311, x_312, x_258); +x_314 = lean_box(0); +x_315 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__29; +x_316 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__34; lean_inc(x_256); -x_318 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_318, 0, x_256); -lean_ctor_set(x_318, 1, x_316); -lean_ctor_set(x_318, 2, x_314); -lean_ctor_set(x_318, 3, x_317); -x_319 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__40; -x_320 = l_Lean_addMacroScope(x_312, x_319, x_258); -x_321 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__37; -x_322 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__44; -x_323 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_323, 0, x_256); -lean_ctor_set(x_323, 1, x_321); -lean_ctor_set(x_323, 2, x_320); -lean_ctor_set(x_323, 3, x_322); +x_317 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_317, 0, x_256); +lean_ctor_set(x_317, 1, x_315); +lean_ctor_set(x_317, 2, x_313); +lean_ctor_set(x_317, 3, x_316); +x_318 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__40; +x_319 = l_Lean_addMacroScope(x_311, x_318, x_258); +x_320 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__37; +x_321 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__44; +x_322 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_322, 0, x_256); +lean_ctor_set(x_322, 1, x_320); +lean_ctor_set(x_322, 2, x_319); +lean_ctor_set(x_322, 3, x_321); lean_inc(x_235); -x_324 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_315, x_235); -x_325 = l_Lean_Elab_Term_Quotation_ArrayStxBuilder_build(x_253); -x_326 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; -x_327 = lean_array_push(x_326, x_323); -x_328 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_329 = lean_array_push(x_328, x_318); -if (lean_obj_tag(x_324) == 0) -{ -lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; -x_330 = l___private_Init_Meta_0__Lean_quoteNameMk(x_235); -x_331 = lean_array_push(x_327, x_330); -x_332 = lean_array_push(x_331, x_325); -x_333 = lean_box(2); -x_334 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_323 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_314, x_235); +x_324 = l_Lean_Elab_Term_Quotation_ArrayStxBuilder_build(x_253); +x_325 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; +x_326 = lean_array_push(x_325, x_322); +x_327 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_328 = lean_array_push(x_327, x_317); +if (lean_obj_tag(x_323) == 0) +{ +lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; +x_329 = l_Lean_quoteNameMk(x_235); +x_330 = lean_array_push(x_326, x_329); +x_331 = lean_array_push(x_330, x_324); +x_332 = lean_box(2); +x_333 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; if (lean_is_scalar(x_238)) { - x_335 = lean_alloc_ctor(1, 3, 0); + x_334 = lean_alloc_ctor(1, 3, 0); } else { - x_335 = x_238; + x_334 = x_238; } -lean_ctor_set(x_335, 0, x_333); -lean_ctor_set(x_335, 1, x_334); -lean_ctor_set(x_335, 2, x_332); -x_336 = lean_array_push(x_329, x_335); -x_337 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_338 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_338, 0, x_333); -lean_ctor_set(x_338, 1, x_337); -lean_ctor_set(x_338, 2, x_336); -x_339 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_339, 0, x_338); -lean_ctor_set(x_339, 1, x_310); -return x_339; +lean_ctor_set(x_334, 0, x_332); +lean_ctor_set(x_334, 1, x_333); +lean_ctor_set(x_334, 2, x_331); +x_335 = lean_array_push(x_328, x_334); +x_336 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_337 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_337, 0, x_332); +lean_ctor_set(x_337, 1, x_336); +lean_ctor_set(x_337, 2, x_335); +x_338 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_338, 0, x_337); +lean_ctor_set(x_338, 1, x_309); +return x_338; } else { -lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; +lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_dec(x_235); -x_340 = lean_ctor_get(x_324, 0); -lean_inc(x_340); -lean_dec(x_324); -x_341 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__3; -x_342 = l_String_intercalate(x_341, x_340); -x_343 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__4; -x_344 = lean_string_append(x_343, x_342); -lean_dec(x_342); -x_345 = l_Lean_nameLitKind; -x_346 = lean_box(2); -x_347 = l_Lean_Syntax_mkLit(x_345, x_344, x_346); -x_348 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; -x_349 = lean_array_push(x_348, x_347); -x_350 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__2; +x_339 = lean_ctor_get(x_323, 0); +lean_inc(x_339); +lean_dec(x_323); +x_340 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; +x_341 = l_String_intercalate(x_340, x_339); +x_342 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4; +x_343 = lean_string_append(x_342, x_341); +lean_dec(x_341); +x_344 = lean_box(2); +x_345 = l_Lean_Syntax_mkNameLit(x_343, x_344); +x_346 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; +x_347 = lean_array_push(x_346, x_345); +x_348 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2; if (lean_is_scalar(x_238)) { - x_351 = lean_alloc_ctor(1, 3, 0); -} else { - x_351 = x_238; -} -lean_ctor_set(x_351, 0, x_346); -lean_ctor_set(x_351, 1, x_350); -lean_ctor_set(x_351, 2, x_349); -x_352 = lean_array_push(x_327, x_351); -x_353 = lean_array_push(x_352, x_325); -x_354 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_355 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_355, 0, x_346); -lean_ctor_set(x_355, 1, x_354); -lean_ctor_set(x_355, 2, x_353); -x_356 = lean_array_push(x_329, x_355); -x_357 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_358 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_358, 0, x_346); -lean_ctor_set(x_358, 1, x_357); -lean_ctor_set(x_358, 2, x_356); -x_359 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_359, 0, x_358); -lean_ctor_set(x_359, 1, x_310); -return x_359; + x_349 = lean_alloc_ctor(1, 3, 0); +} else { + x_349 = x_238; +} +lean_ctor_set(x_349, 0, x_344); +lean_ctor_set(x_349, 1, x_348); +lean_ctor_set(x_349, 2, x_347); +x_350 = lean_array_push(x_326, x_349); +x_351 = lean_array_push(x_350, x_324); +x_352 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_353 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_353, 0, x_344); +lean_ctor_set(x_353, 1, x_352); +lean_ctor_set(x_353, 2, x_351); +x_354 = lean_array_push(x_328, x_353); +x_355 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_356 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_356, 0, x_344); +lean_ctor_set(x_356, 1, x_355); +lean_ctor_set(x_356, 2, x_354); +x_357 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_357, 0, x_356); +lean_ctor_set(x_357, 1, x_309); +return x_357; } } } else { -uint8_t x_360; +uint8_t x_358; lean_dec(x_238); lean_dec(x_235); lean_dec(x_7); lean_dec(x_6); -x_360 = !lean_is_exclusive(x_252); -if (x_360 == 0) +x_358 = !lean_is_exclusive(x_252); +if (x_358 == 0) { return x_252; } else { -lean_object* x_361; lean_object* x_362; lean_object* x_363; -x_361 = lean_ctor_get(x_252, 0); -x_362 = lean_ctor_get(x_252, 1); -lean_inc(x_362); -lean_inc(x_361); +lean_object* x_359; lean_object* x_360; lean_object* x_361; +x_359 = lean_ctor_get(x_252, 0); +x_360 = lean_ctor_get(x_252, 1); +lean_inc(x_360); +lean_inc(x_359); lean_dec(x_252); -x_363 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_363, 0, x_361); -lean_ctor_set(x_363, 1, x_362); -return x_363; +x_361 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_361, 0, x_359); +lean_ctor_set(x_361, 1, x_360); +return x_361; } } } } -block_379: +block_377: { -uint8_t x_368; -lean_dec(x_367); -x_368 = l_Lean_Syntax_isAntiquotSuffixSplice(x_1); -if (x_368 == 0) +uint8_t x_366; +lean_dec(x_365); +x_366 = l_Lean_Syntax_isAntiquotSuffixSplice(x_1); +if (x_366 == 0) { -uint8_t x_369; -x_369 = l_Lean_Syntax_isAntiquotSplice(x_1); -if (x_369 == 0) +uint8_t x_367; +x_367 = l_Lean_Syntax_isAntiquotSplice(x_1); +if (x_367 == 0) { -lean_object* x_370; -x_370 = lean_box(0); -x_236 = x_370; -goto block_366; +lean_object* x_368; +x_368 = lean_box(0); +x_236 = x_368; +goto block_364; } else { -uint8_t x_371; -x_371 = l_Lean_Syntax_isEscapedAntiquot(x_1); -if (x_371 == 0) +uint8_t x_369; +x_369 = l_Lean_Syntax_isEscapedAntiquot(x_1); +if (x_369 == 0) { -lean_object* x_372; lean_object* x_373; +lean_object* x_370; lean_object* x_371; lean_dec(x_235); -x_372 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__48; -x_373 = l_Lean_throwErrorAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__10(x_1, x_372, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_373; +x_370 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__48; +x_371 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__1(x_1, x_370, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_371; } else { -lean_object* x_374; -x_374 = lean_box(0); -x_236 = x_374; -goto block_366; +lean_object* x_372; +x_372 = lean_box(0); +x_236 = x_372; +goto block_364; } } } else { -uint8_t x_375; -x_375 = l_Lean_Syntax_isEscapedAntiquot(x_1); -if (x_375 == 0) +uint8_t x_373; +x_373 = l_Lean_Syntax_isEscapedAntiquot(x_1); +if (x_373 == 0) { -lean_object* x_376; lean_object* x_377; +lean_object* x_374; lean_object* x_375; lean_dec(x_235); -x_376 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__48; -x_377 = l_Lean_throwErrorAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__10(x_1, x_376, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_377; +x_374 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__48; +x_375 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__1(x_1, x_374, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_375; } else { -lean_object* x_378; -x_378 = lean_box(0); -x_236 = x_378; -goto block_366; +lean_object* x_376; +x_376 = lean_box(0); +x_236 = x_376; +goto block_364; } } } } case 2: { -lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; uint8_t x_502; +lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; uint8_t x_500; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_496 = lean_ctor_get(x_1, 1); -lean_inc(x_496); +x_494 = lean_ctor_get(x_1, 1); +lean_inc(x_494); lean_dec(x_1); lean_inc(x_6); -x_497 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_6, x_7, x_8); -x_498 = lean_ctor_get(x_497, 0); +x_495 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_6, x_7, x_8); +x_496 = lean_ctor_get(x_495, 0); +lean_inc(x_496); +x_497 = lean_ctor_get(x_495, 1); +lean_inc(x_497); +lean_dec(x_495); +x_498 = lean_ctor_get(x_6, 10); lean_inc(x_498); -x_499 = lean_ctor_get(x_497, 1); -lean_inc(x_499); -lean_dec(x_497); -x_500 = lean_ctor_get(x_6, 10); -lean_inc(x_500); lean_dec(x_6); -x_501 = lean_st_ref_get(x_7, x_499); +x_499 = lean_st_ref_get(x_7, x_497); lean_dec(x_7); -x_502 = !lean_is_exclusive(x_501); -if (x_502 == 0) +x_500 = !lean_is_exclusive(x_499); +if (x_500 == 0) { -lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; -x_503 = lean_ctor_get(x_501, 0); -x_504 = lean_ctor_get(x_503, 0); -lean_inc(x_504); -lean_dec(x_503); -x_505 = lean_environment_main_module(x_504); -x_506 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__7; -lean_inc(x_500); -lean_inc(x_505); -x_507 = l_Lean_addMacroScope(x_505, x_506, x_500); -x_508 = lean_box(0); -x_509 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__5; -x_510 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__10; +lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; +x_501 = lean_ctor_get(x_499, 0); +x_502 = lean_ctor_get(x_501, 0); +lean_inc(x_502); +lean_dec(x_501); +x_503 = lean_environment_main_module(x_502); +x_504 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__7; lean_inc(x_498); -x_511 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_511, 0, x_498); -lean_ctor_set(x_511, 1, x_509); -lean_ctor_set(x_511, 2, x_507); -lean_ctor_set(x_511, 3, x_510); -x_512 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__13; -x_513 = l_Lean_addMacroScope(x_505, x_512, x_500); -x_514 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__12; -x_515 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_515, 0, x_498); -lean_ctor_set(x_515, 1, x_514); -lean_ctor_set(x_515, 2, x_513); -lean_ctor_set(x_515, 3, x_508); -x_516 = lean_box(2); -x_517 = l_Lean_Syntax_mkStrLit(x_496, x_516); -lean_dec(x_496); -x_518 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_519 = lean_array_push(x_518, x_515); -x_520 = lean_array_push(x_519, x_517); -x_521 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_522 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_522, 0, x_516); -lean_ctor_set(x_522, 1, x_521); -lean_ctor_set(x_522, 2, x_520); -x_523 = lean_array_push(x_518, x_511); -x_524 = lean_array_push(x_523, x_522); -x_525 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_526 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_526, 0, x_516); -lean_ctor_set(x_526, 1, x_525); -lean_ctor_set(x_526, 2, x_524); -lean_ctor_set(x_501, 0, x_526); -return x_501; +lean_inc(x_503); +x_505 = l_Lean_addMacroScope(x_503, x_504, x_498); +x_506 = lean_box(0); +x_507 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__5; +x_508 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__10; +lean_inc(x_496); +x_509 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_509, 0, x_496); +lean_ctor_set(x_509, 1, x_507); +lean_ctor_set(x_509, 2, x_505); +lean_ctor_set(x_509, 3, x_508); +x_510 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__13; +x_511 = l_Lean_addMacroScope(x_503, x_510, x_498); +x_512 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__12; +x_513 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_513, 0, x_496); +lean_ctor_set(x_513, 1, x_512); +lean_ctor_set(x_513, 2, x_511); +lean_ctor_set(x_513, 3, x_506); +x_514 = lean_box(2); +x_515 = l_Lean_Syntax_mkStrLit(x_494, x_514); +lean_dec(x_494); +x_516 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_517 = lean_array_push(x_516, x_513); +x_518 = lean_array_push(x_517, x_515); +x_519 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_520 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_520, 0, x_514); +lean_ctor_set(x_520, 1, x_519); +lean_ctor_set(x_520, 2, x_518); +x_521 = lean_array_push(x_516, x_509); +x_522 = lean_array_push(x_521, x_520); +x_523 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_524 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_524, 0, x_514); +lean_ctor_set(x_524, 1, x_523); +lean_ctor_set(x_524, 2, x_522); +lean_ctor_set(x_499, 0, x_524); +return x_499; } else { -lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; -x_527 = lean_ctor_get(x_501, 0); -x_528 = lean_ctor_get(x_501, 1); -lean_inc(x_528); +lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; +x_525 = lean_ctor_get(x_499, 0); +x_526 = lean_ctor_get(x_499, 1); +lean_inc(x_526); +lean_inc(x_525); +lean_dec(x_499); +x_527 = lean_ctor_get(x_525, 0); lean_inc(x_527); -lean_dec(x_501); -x_529 = lean_ctor_get(x_527, 0); -lean_inc(x_529); -lean_dec(x_527); -x_530 = lean_environment_main_module(x_529); -x_531 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__7; -lean_inc(x_500); -lean_inc(x_530); -x_532 = l_Lean_addMacroScope(x_530, x_531, x_500); -x_533 = lean_box(0); -x_534 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__5; -x_535 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__10; +lean_dec(x_525); +x_528 = lean_environment_main_module(x_527); +x_529 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__7; lean_inc(x_498); -x_536 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_536, 0, x_498); -lean_ctor_set(x_536, 1, x_534); -lean_ctor_set(x_536, 2, x_532); -lean_ctor_set(x_536, 3, x_535); -x_537 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__13; -x_538 = l_Lean_addMacroScope(x_530, x_537, x_500); -x_539 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__12; -x_540 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_540, 0, x_498); -lean_ctor_set(x_540, 1, x_539); -lean_ctor_set(x_540, 2, x_538); -lean_ctor_set(x_540, 3, x_533); -x_541 = lean_box(2); -x_542 = l_Lean_Syntax_mkStrLit(x_496, x_541); -lean_dec(x_496); -x_543 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_544 = lean_array_push(x_543, x_540); -x_545 = lean_array_push(x_544, x_542); -x_546 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_547 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_547, 0, x_541); -lean_ctor_set(x_547, 1, x_546); -lean_ctor_set(x_547, 2, x_545); -x_548 = lean_array_push(x_543, x_536); -x_549 = lean_array_push(x_548, x_547); -x_550 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_551 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_551, 0, x_541); -lean_ctor_set(x_551, 1, x_550); -lean_ctor_set(x_551, 2, x_549); -x_552 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_552, 0, x_551); -lean_ctor_set(x_552, 1, x_528); -return x_552; +lean_inc(x_528); +x_530 = l_Lean_addMacroScope(x_528, x_529, x_498); +x_531 = lean_box(0); +x_532 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__5; +x_533 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__10; +lean_inc(x_496); +x_534 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_534, 0, x_496); +lean_ctor_set(x_534, 1, x_532); +lean_ctor_set(x_534, 2, x_530); +lean_ctor_set(x_534, 3, x_533); +x_535 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__13; +x_536 = l_Lean_addMacroScope(x_528, x_535, x_498); +x_537 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__12; +x_538 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_538, 0, x_496); +lean_ctor_set(x_538, 1, x_537); +lean_ctor_set(x_538, 2, x_536); +lean_ctor_set(x_538, 3, x_531); +x_539 = lean_box(2); +x_540 = l_Lean_Syntax_mkStrLit(x_494, x_539); +lean_dec(x_494); +x_541 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_542 = lean_array_push(x_541, x_538); +x_543 = lean_array_push(x_542, x_540); +x_544 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_545 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_545, 0, x_539); +lean_ctor_set(x_545, 1, x_544); +lean_ctor_set(x_545, 2, x_543); +x_546 = lean_array_push(x_541, x_534); +x_547 = lean_array_push(x_546, x_545); +x_548 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_549 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_549, 0, x_539); +lean_ctor_set(x_549, 1, x_548); +lean_ctor_set(x_549, 2, x_547); +x_550 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_550, 0, x_549); +lean_ctor_set(x_550, 1, x_526); +return x_550; } } default: { -uint8_t x_553; -x_553 = !lean_is_exclusive(x_1); -if (x_553 == 0) +uint8_t x_551; +x_551 = !lean_is_exclusive(x_1); +if (x_551 == 0) { -lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; uint8_t x_560; -x_554 = lean_ctor_get(x_1, 1); -x_555 = lean_ctor_get(x_1, 2); -x_556 = lean_ctor_get(x_1, 3); -x_557 = lean_ctor_get(x_1, 0); -lean_dec(x_557); -x_558 = lean_ctor_get(x_6, 2); -lean_inc(x_558); -x_559 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__51; -x_560 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_558, x_559); -lean_dec(x_558); -if (x_560 == 0) +lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; uint8_t x_558; +x_552 = lean_ctor_get(x_1, 1); +x_553 = lean_ctor_get(x_1, 2); +x_554 = lean_ctor_get(x_1, 3); +x_555 = lean_ctor_get(x_1, 0); +lean_dec(x_555); +x_556 = lean_ctor_get(x_6, 2); +lean_inc(x_556); +x_557 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__49; +x_558 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_556, x_557); +lean_dec(x_556); +if (x_558 == 0) { -lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; lean_object* x_565; uint8_t x_566; +lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; uint8_t x_564; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_inc(x_6); -x_561 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_6, x_7, x_8); -x_562 = lean_ctor_get(x_561, 0); +x_559 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_6, x_7, x_8); +x_560 = lean_ctor_get(x_559, 0); +lean_inc(x_560); +x_561 = lean_ctor_get(x_559, 1); +lean_inc(x_561); +lean_dec(x_559); +x_562 = lean_ctor_get(x_6, 10); lean_inc(x_562); -x_563 = lean_ctor_get(x_561, 1); -lean_inc(x_563); -lean_dec(x_561); -x_564 = lean_ctor_get(x_6, 10); -lean_inc(x_564); -lean_dec(x_6); -x_565 = lean_st_ref_get(x_7, x_563); -lean_dec(x_7); -x_566 = !lean_is_exclusive(x_565); -if (x_566 == 0) -{ -lean_object* x_567; lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; -x_567 = lean_ctor_get(x_565, 0); -x_568 = lean_ctor_get(x_567, 0); -lean_inc(x_568); -lean_dec(x_567); -x_569 = lean_environment_main_module(x_568); -x_570 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__6; -lean_inc(x_564); -lean_inc(x_569); -x_571 = l_Lean_addMacroScope(x_569, x_570, x_564); -x_572 = lean_box(0); -x_573 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__3; -x_574 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__9; +lean_dec(x_6); +x_563 = lean_st_ref_get(x_7, x_561); +lean_dec(x_7); +x_564 = !lean_is_exclusive(x_563); +if (x_564 == 0) +{ +lean_object* x_565; lean_object* x_566; lean_object* x_567; lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; +x_565 = lean_ctor_get(x_563, 0); +x_566 = lean_ctor_get(x_565, 0); +lean_inc(x_566); +lean_dec(x_565); +x_567 = lean_environment_main_module(x_566); +x_568 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__6; lean_inc(x_562); -lean_ctor_set(x_1, 3, x_574); -lean_ctor_set(x_1, 2, x_571); -lean_ctor_set(x_1, 1, x_573); -lean_ctor_set(x_1, 0, x_562); -x_575 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__13; -x_576 = l_Lean_addMacroScope(x_569, x_575, x_564); -x_577 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__12; -x_578 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_578, 0, x_562); -lean_ctor_set(x_578, 1, x_577); -lean_ctor_set(x_578, 2, x_576); -lean_ctor_set(x_578, 3, x_572); -lean_inc(x_555); -x_579 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_572, x_555); -x_580 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__10(x_556); -x_581 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; -x_582 = lean_array_push(x_581, x_578); -x_583 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_584 = lean_array_push(x_583, x_1); -x_585 = lean_ctor_get(x_554, 0); +lean_inc(x_567); +x_569 = l_Lean_addMacroScope(x_567, x_568, x_562); +x_570 = lean_box(0); +x_571 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__3; +x_572 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__9; +lean_inc(x_560); +lean_ctor_set(x_1, 3, x_572); +lean_ctor_set(x_1, 2, x_569); +lean_ctor_set(x_1, 1, x_571); +lean_ctor_set(x_1, 0, x_560); +x_573 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__13; +x_574 = l_Lean_addMacroScope(x_567, x_573, x_562); +x_575 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__12; +x_576 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_576, 0, x_560); +lean_ctor_set(x_576, 1, x_575); +lean_ctor_set(x_576, 2, x_574); +lean_ctor_set(x_576, 3, x_570); +lean_inc(x_553); +x_577 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_570, x_553); +x_578 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__26(x_554); +x_579 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; +x_580 = lean_array_push(x_579, x_576); +x_581 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_582 = lean_array_push(x_581, x_1); +x_583 = lean_ctor_get(x_552, 0); +lean_inc(x_583); +x_584 = lean_ctor_get(x_552, 1); +lean_inc(x_584); +x_585 = lean_ctor_get(x_552, 2); lean_inc(x_585); -x_586 = lean_ctor_get(x_554, 1); -lean_inc(x_586); -x_587 = lean_ctor_get(x_554, 2); -lean_inc(x_587); -lean_dec(x_554); -x_588 = lean_string_utf8_extract(x_585, x_586, x_587); -lean_dec(x_587); -lean_dec(x_586); +lean_dec(x_552); +x_586 = lean_string_utf8_extract(x_583, x_584, x_585); lean_dec(x_585); -x_589 = lean_box(2); -x_590 = l_Lean_Syntax_mkStrLit(x_588, x_589); -lean_dec(x_588); -x_591 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; -x_592 = lean_array_push(x_591, x_590); -x_593 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__17; -x_594 = l_Lean_Syntax_mkCApp(x_593, x_592); -x_595 = lean_array_push(x_582, x_594); -if (lean_obj_tag(x_579) == 0) -{ -lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; lean_object* x_603; -x_596 = l___private_Init_Meta_0__Lean_quoteNameMk(x_555); -x_597 = lean_array_push(x_595, x_596); -x_598 = lean_array_push(x_597, x_580); -x_599 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_600 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_600, 0, x_589); -lean_ctor_set(x_600, 1, x_599); -lean_ctor_set(x_600, 2, x_598); -x_601 = lean_array_push(x_584, x_600); -x_602 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_603 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_603, 0, x_589); -lean_ctor_set(x_603, 1, x_602); -lean_ctor_set(x_603, 2, x_601); -lean_ctor_set(x_565, 0, x_603); -return x_565; +lean_dec(x_584); +lean_dec(x_583); +x_587 = lean_box(2); +x_588 = l_Lean_Syntax_mkStrLit(x_586, x_587); +lean_dec(x_586); +x_589 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; +x_590 = lean_array_push(x_589, x_588); +x_591 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__17; +x_592 = l_Lean_Syntax_mkCApp(x_591, x_590); +x_593 = lean_array_push(x_580, x_592); +if (lean_obj_tag(x_577) == 0) +{ +lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; +x_594 = l_Lean_quoteNameMk(x_553); +x_595 = lean_array_push(x_593, x_594); +x_596 = lean_array_push(x_595, x_578); +x_597 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_598 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_598, 0, x_587); +lean_ctor_set(x_598, 1, x_597); +lean_ctor_set(x_598, 2, x_596); +x_599 = lean_array_push(x_582, x_598); +x_600 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_601 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_601, 0, x_587); +lean_ctor_set(x_601, 1, x_600); +lean_ctor_set(x_601, 2, x_599); +lean_ctor_set(x_563, 0, x_601); +return x_563; } else { -lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; lean_object* x_620; -lean_dec(x_555); -x_604 = lean_ctor_get(x_579, 0); -lean_inc(x_604); -lean_dec(x_579); -x_605 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__3; -x_606 = l_String_intercalate(x_605, x_604); -x_607 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__4; -x_608 = lean_string_append(x_607, x_606); -lean_dec(x_606); -x_609 = l_Lean_nameLitKind; -x_610 = l_Lean_Syntax_mkLit(x_609, x_608, x_589); -x_611 = lean_array_push(x_591, x_610); -x_612 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__2; -x_613 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_613, 0, x_589); -lean_ctor_set(x_613, 1, x_612); -lean_ctor_set(x_613, 2, x_611); -x_614 = lean_array_push(x_595, x_613); -x_615 = lean_array_push(x_614, x_580); -x_616 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; +lean_dec(x_553); +x_602 = lean_ctor_get(x_577, 0); +lean_inc(x_602); +lean_dec(x_577); +x_603 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; +x_604 = l_String_intercalate(x_603, x_602); +x_605 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4; +x_606 = lean_string_append(x_605, x_604); +lean_dec(x_604); +x_607 = l_Lean_Syntax_mkNameLit(x_606, x_587); +x_608 = lean_array_push(x_589, x_607); +x_609 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2; +x_610 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_610, 0, x_587); +lean_ctor_set(x_610, 1, x_609); +lean_ctor_set(x_610, 2, x_608); +x_611 = lean_array_push(x_593, x_610); +x_612 = lean_array_push(x_611, x_578); +x_613 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_614 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_614, 0, x_587); +lean_ctor_set(x_614, 1, x_613); +lean_ctor_set(x_614, 2, x_612); +x_615 = lean_array_push(x_582, x_614); +x_616 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; x_617 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_617, 0, x_589); +lean_ctor_set(x_617, 0, x_587); lean_ctor_set(x_617, 1, x_616); lean_ctor_set(x_617, 2, x_615); -x_618 = lean_array_push(x_584, x_617); -x_619 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_620 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_620, 0, x_589); -lean_ctor_set(x_620, 1, x_619); -lean_ctor_set(x_620, 2, x_618); -lean_ctor_set(x_565, 0, x_620); -return x_565; +lean_ctor_set(x_563, 0, x_617); +return x_563; } } else { -lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; -x_621 = lean_ctor_get(x_565, 0); -x_622 = lean_ctor_get(x_565, 1); -lean_inc(x_622); -lean_inc(x_621); -lean_dec(x_565); -x_623 = lean_ctor_get(x_621, 0); -lean_inc(x_623); -lean_dec(x_621); -x_624 = lean_environment_main_module(x_623); -x_625 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__6; -lean_inc(x_564); -lean_inc(x_624); -x_626 = l_Lean_addMacroScope(x_624, x_625, x_564); -x_627 = lean_box(0); -x_628 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__3; -x_629 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__9; +lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; +x_618 = lean_ctor_get(x_563, 0); +x_619 = lean_ctor_get(x_563, 1); +lean_inc(x_619); +lean_inc(x_618); +lean_dec(x_563); +x_620 = lean_ctor_get(x_618, 0); +lean_inc(x_620); +lean_dec(x_618); +x_621 = lean_environment_main_module(x_620); +x_622 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__6; lean_inc(x_562); -lean_ctor_set(x_1, 3, x_629); -lean_ctor_set(x_1, 2, x_626); -lean_ctor_set(x_1, 1, x_628); -lean_ctor_set(x_1, 0, x_562); -x_630 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__13; -x_631 = l_Lean_addMacroScope(x_624, x_630, x_564); -x_632 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__12; -x_633 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_633, 0, x_562); -lean_ctor_set(x_633, 1, x_632); -lean_ctor_set(x_633, 2, x_631); -lean_ctor_set(x_633, 3, x_627); -lean_inc(x_555); -x_634 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_627, x_555); -x_635 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__10(x_556); -x_636 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; -x_637 = lean_array_push(x_636, x_633); -x_638 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_639 = lean_array_push(x_638, x_1); -x_640 = lean_ctor_get(x_554, 0); -lean_inc(x_640); -x_641 = lean_ctor_get(x_554, 1); -lean_inc(x_641); -x_642 = lean_ctor_get(x_554, 2); -lean_inc(x_642); -lean_dec(x_554); -x_643 = lean_string_utf8_extract(x_640, x_641, x_642); -lean_dec(x_642); -lean_dec(x_641); +lean_inc(x_621); +x_623 = l_Lean_addMacroScope(x_621, x_622, x_562); +x_624 = lean_box(0); +x_625 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__3; +x_626 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__9; +lean_inc(x_560); +lean_ctor_set(x_1, 3, x_626); +lean_ctor_set(x_1, 2, x_623); +lean_ctor_set(x_1, 1, x_625); +lean_ctor_set(x_1, 0, x_560); +x_627 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__13; +x_628 = l_Lean_addMacroScope(x_621, x_627, x_562); +x_629 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__12; +x_630 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_630, 0, x_560); +lean_ctor_set(x_630, 1, x_629); +lean_ctor_set(x_630, 2, x_628); +lean_ctor_set(x_630, 3, x_624); +lean_inc(x_553); +x_631 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_624, x_553); +x_632 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__26(x_554); +x_633 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; +x_634 = lean_array_push(x_633, x_630); +x_635 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_636 = lean_array_push(x_635, x_1); +x_637 = lean_ctor_get(x_552, 0); +lean_inc(x_637); +x_638 = lean_ctor_get(x_552, 1); +lean_inc(x_638); +x_639 = lean_ctor_get(x_552, 2); +lean_inc(x_639); +lean_dec(x_552); +x_640 = lean_string_utf8_extract(x_637, x_638, x_639); +lean_dec(x_639); +lean_dec(x_638); +lean_dec(x_637); +x_641 = lean_box(2); +x_642 = l_Lean_Syntax_mkStrLit(x_640, x_641); lean_dec(x_640); -x_644 = lean_box(2); -x_645 = l_Lean_Syntax_mkStrLit(x_643, x_644); -lean_dec(x_643); -x_646 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; -x_647 = lean_array_push(x_646, x_645); -x_648 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__17; -x_649 = l_Lean_Syntax_mkCApp(x_648, x_647); -x_650 = lean_array_push(x_637, x_649); -if (lean_obj_tag(x_634) == 0) -{ -lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; lean_object* x_658; lean_object* x_659; -x_651 = l___private_Init_Meta_0__Lean_quoteNameMk(x_555); -x_652 = lean_array_push(x_650, x_651); -x_653 = lean_array_push(x_652, x_635); -x_654 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_643 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; +x_644 = lean_array_push(x_643, x_642); +x_645 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__17; +x_646 = l_Lean_Syntax_mkCApp(x_645, x_644); +x_647 = lean_array_push(x_634, x_646); +if (lean_obj_tag(x_631) == 0) +{ +lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; +x_648 = l_Lean_quoteNameMk(x_553); +x_649 = lean_array_push(x_647, x_648); +x_650 = lean_array_push(x_649, x_632); +x_651 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_652 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_652, 0, x_641); +lean_ctor_set(x_652, 1, x_651); +lean_ctor_set(x_652, 2, x_650); +x_653 = lean_array_push(x_636, x_652); +x_654 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; x_655 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_655, 0, x_644); +lean_ctor_set(x_655, 0, x_641); lean_ctor_set(x_655, 1, x_654); lean_ctor_set(x_655, 2, x_653); -x_656 = lean_array_push(x_639, x_655); -x_657 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_658 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_658, 0, x_644); -lean_ctor_set(x_658, 1, x_657); -lean_ctor_set(x_658, 2, x_656); -x_659 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_659, 0, x_658); -lean_ctor_set(x_659, 1, x_622); -return x_659; +x_656 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_656, 0, x_655); +lean_ctor_set(x_656, 1, x_619); +return x_656; } else { -lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; lean_object* x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; -lean_dec(x_555); -x_660 = lean_ctor_get(x_634, 0); -lean_inc(x_660); -lean_dec(x_634); -x_661 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__3; -x_662 = l_String_intercalate(x_661, x_660); -x_663 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__4; -x_664 = lean_string_append(x_663, x_662); -lean_dec(x_662); -x_665 = l_Lean_nameLitKind; -x_666 = l_Lean_Syntax_mkLit(x_665, x_664, x_644); -x_667 = lean_array_push(x_646, x_666); -x_668 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__2; +lean_object* x_657; lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; +lean_dec(x_553); +x_657 = lean_ctor_get(x_631, 0); +lean_inc(x_657); +lean_dec(x_631); +x_658 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; +x_659 = l_String_intercalate(x_658, x_657); +x_660 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4; +x_661 = lean_string_append(x_660, x_659); +lean_dec(x_659); +x_662 = l_Lean_Syntax_mkNameLit(x_661, x_641); +x_663 = lean_array_push(x_643, x_662); +x_664 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2; +x_665 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_665, 0, x_641); +lean_ctor_set(x_665, 1, x_664); +lean_ctor_set(x_665, 2, x_663); +x_666 = lean_array_push(x_647, x_665); +x_667 = lean_array_push(x_666, x_632); +x_668 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; x_669 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_669, 0, x_644); +lean_ctor_set(x_669, 0, x_641); lean_ctor_set(x_669, 1, x_668); lean_ctor_set(x_669, 2, x_667); -x_670 = lean_array_push(x_650, x_669); -x_671 = lean_array_push(x_670, x_635); -x_672 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_673 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_673, 0, x_644); -lean_ctor_set(x_673, 1, x_672); -lean_ctor_set(x_673, 2, x_671); -x_674 = lean_array_push(x_639, x_673); -x_675 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_676 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_676, 0, x_644); -lean_ctor_set(x_676, 1, x_675); -lean_ctor_set(x_676, 2, x_674); -x_677 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_677, 0, x_676); -lean_ctor_set(x_677, 1, x_622); -return x_677; +x_670 = lean_array_push(x_636, x_669); +x_671 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_672 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_672, 0, x_641); +lean_ctor_set(x_672, 1, x_671); +lean_ctor_set(x_672, 2, x_670); +x_673 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_673, 0, x_672); +lean_ctor_set(x_673, 1, x_619); +return x_673; } } } else { -lean_object* x_678; lean_object* x_679; +lean_object* x_674; lean_object* x_675; lean_free_object(x_1); -x_678 = lean_box(0); -x_679 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1(x_555, x_556, x_554, x_678, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_674 = lean_box(0); +x_675 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1(x_553, x_554, x_552, x_674, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_554); -return x_679; +lean_dec(x_552); +return x_675; } } else { -lean_object* x_680; lean_object* x_681; lean_object* x_682; lean_object* x_683; lean_object* x_684; uint8_t x_685; -x_680 = lean_ctor_get(x_1, 1); -x_681 = lean_ctor_get(x_1, 2); -x_682 = lean_ctor_get(x_1, 3); -lean_inc(x_682); -lean_inc(x_681); -lean_inc(x_680); +lean_object* x_676; lean_object* x_677; lean_object* x_678; lean_object* x_679; lean_object* x_680; uint8_t x_681; +x_676 = lean_ctor_get(x_1, 1); +x_677 = lean_ctor_get(x_1, 2); +x_678 = lean_ctor_get(x_1, 3); +lean_inc(x_678); +lean_inc(x_677); +lean_inc(x_676); lean_dec(x_1); -x_683 = lean_ctor_get(x_6, 2); -lean_inc(x_683); -x_684 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__51; -x_685 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_683, x_684); -lean_dec(x_683); -if (x_685 == 0) +x_679 = lean_ctor_get(x_6, 2); +lean_inc(x_679); +x_680 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__49; +x_681 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_679, x_680); +lean_dec(x_679); +if (x_681 == 0) { -lean_object* x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; lean_object* x_690; lean_object* x_691; lean_object* x_692; lean_object* x_693; lean_object* x_694; lean_object* x_695; lean_object* x_696; lean_object* x_697; lean_object* x_698; lean_object* x_699; lean_object* x_700; lean_object* x_701; lean_object* x_702; lean_object* x_703; lean_object* x_704; lean_object* x_705; lean_object* x_706; lean_object* x_707; lean_object* x_708; lean_object* x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; lean_object* x_718; lean_object* x_719; lean_object* x_720; lean_object* x_721; lean_object* x_722; +lean_object* x_682; lean_object* x_683; lean_object* x_684; lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; lean_object* x_690; lean_object* x_691; lean_object* x_692; lean_object* x_693; lean_object* x_694; lean_object* x_695; lean_object* x_696; lean_object* x_697; lean_object* x_698; lean_object* x_699; lean_object* x_700; lean_object* x_701; lean_object* x_702; lean_object* x_703; lean_object* x_704; lean_object* x_705; lean_object* x_706; lean_object* x_707; lean_object* x_708; lean_object* x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; lean_object* x_718; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_inc(x_6); -x_686 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_6, x_7, x_8); +x_682 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_6, x_7, x_8); +x_683 = lean_ctor_get(x_682, 0); +lean_inc(x_683); +x_684 = lean_ctor_get(x_682, 1); +lean_inc(x_684); +lean_dec(x_682); +x_685 = lean_ctor_get(x_6, 10); +lean_inc(x_685); +lean_dec(x_6); +x_686 = lean_st_ref_get(x_7, x_684); +lean_dec(x_7); x_687 = lean_ctor_get(x_686, 0); lean_inc(x_687); x_688 = lean_ctor_get(x_686, 1); lean_inc(x_688); -lean_dec(x_686); -x_689 = lean_ctor_get(x_6, 10); -lean_inc(x_689); -lean_dec(x_6); -x_690 = lean_st_ref_get(x_7, x_688); -lean_dec(x_7); -x_691 = lean_ctor_get(x_690, 0); +if (lean_is_exclusive(x_686)) { + lean_ctor_release(x_686, 0); + lean_ctor_release(x_686, 1); + x_689 = x_686; +} else { + lean_dec_ref(x_686); + x_689 = lean_box(0); +} +x_690 = lean_ctor_get(x_687, 0); +lean_inc(x_690); +lean_dec(x_687); +x_691 = lean_environment_main_module(x_690); +x_692 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__6; +lean_inc(x_685); lean_inc(x_691); -x_692 = lean_ctor_get(x_690, 1); -lean_inc(x_692); -if (lean_is_exclusive(x_690)) { - lean_ctor_release(x_690, 0); - lean_ctor_release(x_690, 1); - x_693 = x_690; -} else { - lean_dec_ref(x_690); - x_693 = lean_box(0); -} -x_694 = lean_ctor_get(x_691, 0); -lean_inc(x_694); -lean_dec(x_691); -x_695 = lean_environment_main_module(x_694); -x_696 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__6; -lean_inc(x_689); -lean_inc(x_695); -x_697 = l_Lean_addMacroScope(x_695, x_696, x_689); -x_698 = lean_box(0); -x_699 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__3; -x_700 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__9; -lean_inc(x_687); +x_693 = l_Lean_addMacroScope(x_691, x_692, x_685); +x_694 = lean_box(0); +x_695 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__3; +x_696 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__9; +lean_inc(x_683); +x_697 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_697, 0, x_683); +lean_ctor_set(x_697, 1, x_695); +lean_ctor_set(x_697, 2, x_693); +lean_ctor_set(x_697, 3, x_696); +x_698 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__13; +x_699 = l_Lean_addMacroScope(x_691, x_698, x_685); +x_700 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__12; x_701 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_701, 0, x_687); -lean_ctor_set(x_701, 1, x_699); -lean_ctor_set(x_701, 2, x_697); -lean_ctor_set(x_701, 3, x_700); -x_702 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__13; -x_703 = l_Lean_addMacroScope(x_695, x_702, x_689); -x_704 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__12; -x_705 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_705, 0, x_687); -lean_ctor_set(x_705, 1, x_704); -lean_ctor_set(x_705, 2, x_703); -lean_ctor_set(x_705, 3, x_698); -lean_inc(x_681); -x_706 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_698, x_681); -x_707 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__10(x_682); -x_708 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; -x_709 = lean_array_push(x_708, x_705); -x_710 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_711 = lean_array_push(x_710, x_701); -x_712 = lean_ctor_get(x_680, 0); -lean_inc(x_712); -x_713 = lean_ctor_get(x_680, 1); -lean_inc(x_713); -x_714 = lean_ctor_get(x_680, 2); -lean_inc(x_714); -lean_dec(x_680); -x_715 = lean_string_utf8_extract(x_712, x_713, x_714); -lean_dec(x_714); -lean_dec(x_713); -lean_dec(x_712); -x_716 = lean_box(2); -x_717 = l_Lean_Syntax_mkStrLit(x_715, x_716); -lean_dec(x_715); -x_718 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; -x_719 = lean_array_push(x_718, x_717); -x_720 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__17; -x_721 = l_Lean_Syntax_mkCApp(x_720, x_719); -x_722 = lean_array_push(x_709, x_721); -if (lean_obj_tag(x_706) == 0) -{ -lean_object* x_723; lean_object* x_724; lean_object* x_725; lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; lean_object* x_730; lean_object* x_731; -x_723 = l___private_Init_Meta_0__Lean_quoteNameMk(x_681); -x_724 = lean_array_push(x_722, x_723); -x_725 = lean_array_push(x_724, x_707); -x_726 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_727 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_727, 0, x_716); -lean_ctor_set(x_727, 1, x_726); -lean_ctor_set(x_727, 2, x_725); -x_728 = lean_array_push(x_711, x_727); -x_729 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_730 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_730, 0, x_716); -lean_ctor_set(x_730, 1, x_729); -lean_ctor_set(x_730, 2, x_728); -if (lean_is_scalar(x_693)) { - x_731 = lean_alloc_ctor(0, 2, 0); -} else { - x_731 = x_693; -} -lean_ctor_set(x_731, 0, x_730); -lean_ctor_set(x_731, 1, x_692); -return x_731; -} -else -{ -lean_object* x_732; lean_object* x_733; lean_object* x_734; lean_object* x_735; lean_object* x_736; lean_object* x_737; lean_object* x_738; lean_object* x_739; lean_object* x_740; lean_object* x_741; lean_object* x_742; lean_object* x_743; lean_object* x_744; lean_object* x_745; lean_object* x_746; lean_object* x_747; lean_object* x_748; lean_object* x_749; -lean_dec(x_681); -x_732 = lean_ctor_get(x_706, 0); -lean_inc(x_732); -lean_dec(x_706); -x_733 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__3; -x_734 = l_String_intercalate(x_733, x_732); -x_735 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__4; -x_736 = lean_string_append(x_735, x_734); -lean_dec(x_734); -x_737 = l_Lean_nameLitKind; -x_738 = l_Lean_Syntax_mkLit(x_737, x_736, x_716); -x_739 = lean_array_push(x_718, x_738); -x_740 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__2; -x_741 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_741, 0, x_716); -lean_ctor_set(x_741, 1, x_740); -lean_ctor_set(x_741, 2, x_739); -x_742 = lean_array_push(x_722, x_741); -x_743 = lean_array_push(x_742, x_707); -x_744 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_745 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_745, 0, x_716); -lean_ctor_set(x_745, 1, x_744); -lean_ctor_set(x_745, 2, x_743); -x_746 = lean_array_push(x_711, x_745); -x_747 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_748 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_748, 0, x_716); -lean_ctor_set(x_748, 1, x_747); -lean_ctor_set(x_748, 2, x_746); -if (lean_is_scalar(x_693)) { - x_749 = lean_alloc_ctor(0, 2, 0); -} else { - x_749 = x_693; -} -lean_ctor_set(x_749, 0, x_748); -lean_ctor_set(x_749, 1, x_692); -return x_749; -} -} -else -{ -lean_object* x_750; lean_object* x_751; -x_750 = lean_box(0); -x_751 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1(x_681, x_682, x_680, x_750, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_ctor_set(x_701, 0, x_683); +lean_ctor_set(x_701, 1, x_700); +lean_ctor_set(x_701, 2, x_699); +lean_ctor_set(x_701, 3, x_694); +lean_inc(x_677); +x_702 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_694, x_677); +x_703 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__26(x_678); +x_704 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; +x_705 = lean_array_push(x_704, x_701); +x_706 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_707 = lean_array_push(x_706, x_697); +x_708 = lean_ctor_get(x_676, 0); +lean_inc(x_708); +x_709 = lean_ctor_get(x_676, 1); +lean_inc(x_709); +x_710 = lean_ctor_get(x_676, 2); +lean_inc(x_710); +lean_dec(x_676); +x_711 = lean_string_utf8_extract(x_708, x_709, x_710); +lean_dec(x_710); +lean_dec(x_709); +lean_dec(x_708); +x_712 = lean_box(2); +x_713 = l_Lean_Syntax_mkStrLit(x_711, x_712); +lean_dec(x_711); +x_714 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; +x_715 = lean_array_push(x_714, x_713); +x_716 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__17; +x_717 = l_Lean_Syntax_mkCApp(x_716, x_715); +x_718 = lean_array_push(x_705, x_717); +if (lean_obj_tag(x_702) == 0) +{ +lean_object* x_719; lean_object* x_720; lean_object* x_721; lean_object* x_722; lean_object* x_723; lean_object* x_724; lean_object* x_725; lean_object* x_726; lean_object* x_727; +x_719 = l_Lean_quoteNameMk(x_677); +x_720 = lean_array_push(x_718, x_719); +x_721 = lean_array_push(x_720, x_703); +x_722 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_723 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_723, 0, x_712); +lean_ctor_set(x_723, 1, x_722); +lean_ctor_set(x_723, 2, x_721); +x_724 = lean_array_push(x_707, x_723); +x_725 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_726 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_726, 0, x_712); +lean_ctor_set(x_726, 1, x_725); +lean_ctor_set(x_726, 2, x_724); +if (lean_is_scalar(x_689)) { + x_727 = lean_alloc_ctor(0, 2, 0); +} else { + x_727 = x_689; +} +lean_ctor_set(x_727, 0, x_726); +lean_ctor_set(x_727, 1, x_688); +return x_727; +} +else +{ +lean_object* x_728; lean_object* x_729; lean_object* x_730; lean_object* x_731; lean_object* x_732; lean_object* x_733; lean_object* x_734; lean_object* x_735; lean_object* x_736; lean_object* x_737; lean_object* x_738; lean_object* x_739; lean_object* x_740; lean_object* x_741; lean_object* x_742; lean_object* x_743; lean_object* x_744; +lean_dec(x_677); +x_728 = lean_ctor_get(x_702, 0); +lean_inc(x_728); +lean_dec(x_702); +x_729 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; +x_730 = l_String_intercalate(x_729, x_728); +x_731 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4; +x_732 = lean_string_append(x_731, x_730); +lean_dec(x_730); +x_733 = l_Lean_Syntax_mkNameLit(x_732, x_712); +x_734 = lean_array_push(x_714, x_733); +x_735 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2; +x_736 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_736, 0, x_712); +lean_ctor_set(x_736, 1, x_735); +lean_ctor_set(x_736, 2, x_734); +x_737 = lean_array_push(x_718, x_736); +x_738 = lean_array_push(x_737, x_703); +x_739 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_740 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_740, 0, x_712); +lean_ctor_set(x_740, 1, x_739); +lean_ctor_set(x_740, 2, x_738); +x_741 = lean_array_push(x_707, x_740); +x_742 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_743 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_743, 0, x_712); +lean_ctor_set(x_743, 1, x_742); +lean_ctor_set(x_743, 2, x_741); +if (lean_is_scalar(x_689)) { + x_744 = lean_alloc_ctor(0, 2, 0); +} else { + x_744 = x_689; +} +lean_ctor_set(x_744, 0, x_743); +lean_ctor_set(x_744, 1, x_688); +return x_744; +} +} +else +{ +lean_object* x_745; lean_object* x_746; +x_745 = lean_box(0); +x_746 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1(x_677, x_678, x_676, x_745, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_680); -return x_751; +lean_dec(x_676); +return x_746; } } } @@ -11228,7 +12453,7 @@ lean_ctor_set(x_29, 0, x_16); lean_ctor_set(x_29, 1, x_27); lean_ctor_set(x_29, 2, x_25); lean_ctor_set(x_29, 3, x_28); -x_30 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; +x_30 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; lean_inc(x_16); lean_ctor_set(x_11, 1, x_30); lean_ctor_set(x_11, 0, x_16); @@ -11281,18 +12506,18 @@ x_55 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_55, 0, x_44); lean_ctor_set(x_55, 1, x_45); lean_ctor_set(x_55, 2, x_54); -x_56 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; +x_56 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; lean_inc(x_16); x_57 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_57, 0, x_16); lean_ctor_set(x_57, 1, x_56); -x_58 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; +x_58 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; x_59 = lean_array_push(x_58, x_11); lean_inc(x_59); x_60 = lean_array_push(x_59, x_55); lean_inc(x_57); x_61 = lean_array_push(x_60, x_57); -x_62 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__6; +x_62 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6; x_63 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_63, 0, x_44); lean_ctor_set(x_63, 1, x_62); @@ -11371,7 +12596,7 @@ lean_ctor_set(x_96, 0, x_16); lean_ctor_set(x_96, 1, x_94); lean_ctor_set(x_96, 2, x_92); lean_ctor_set(x_96, 3, x_95); -x_97 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; +x_97 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; lean_inc(x_16); lean_ctor_set(x_11, 1, x_97); lean_ctor_set(x_11, 0, x_16); @@ -11424,18 +12649,18 @@ x_122 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_122, 0, x_111); lean_ctor_set(x_122, 1, x_112); lean_ctor_set(x_122, 2, x_121); -x_123 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; +x_123 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; lean_inc(x_16); x_124 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_124, 0, x_16); lean_ctor_set(x_124, 1, x_123); -x_125 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; +x_125 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; x_126 = lean_array_push(x_125, x_11); lean_inc(x_126); x_127 = lean_array_push(x_126, x_122); lean_inc(x_124); x_128 = lean_array_push(x_127, x_124); -x_129 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__6; +x_129 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6; x_130 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_130, 0, x_111); lean_ctor_set(x_130, 1, x_129); @@ -11539,7 +12764,7 @@ lean_ctor_set(x_171, 0, x_157); lean_ctor_set(x_171, 1, x_169); lean_ctor_set(x_171, 2, x_167); lean_ctor_set(x_171, 3, x_170); -x_172 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; +x_172 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; lean_inc(x_157); x_173 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_173, 0, x_157); @@ -11593,18 +12818,18 @@ x_198 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_198, 0, x_187); lean_ctor_set(x_198, 1, x_188); lean_ctor_set(x_198, 2, x_197); -x_199 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; +x_199 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; lean_inc(x_157); x_200 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_200, 0, x_157); lean_ctor_set(x_200, 1, x_199); -x_201 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; +x_201 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; x_202 = lean_array_push(x_201, x_173); lean_inc(x_202); x_203 = lean_array_push(x_202, x_198); lean_inc(x_200); x_204 = lean_array_push(x_203, x_200); -x_205 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__6; +x_205 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6; x_206 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_206, 0, x_187); lean_ctor_set(x_206, 1, x_205); @@ -11670,13 +12895,40 @@ else lean_object* x_231; lean_object* x_232; lean_dec(x_11); x_231 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__2; -x_232 = l_Lean_throwErrorAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__10(x_1, x_231, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_232 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__1(x_1, x_231, x_2, x_3, x_4, x_5, x_6, x_7, x_8); return x_232; } } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__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, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__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, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_throwError___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___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) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___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_object* x_9, lean_object* x_10) { _start: { size_t x_11; size_t x_12; lean_object* x_13; @@ -11684,40 +12936,130 @@ x_11 = lean_unbox_usize(x_1); lean_dec(x_1); x_12 = lean_unbox_usize(x_2); lean_dec(x_2); -x_13 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__1(x_11, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_13 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4(x_11, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); return x_13; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__2___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__2(x_1, x_2); +x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5(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_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -size_t x_13; size_t x_14; lean_object* x_15; -x_13 = lean_unbox_usize(x_3); +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_15 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_16 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6(x_1, x_2, x_3, x_14, x_15, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_3); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__7___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +size_t x_14; size_t x_15; lean_object* x_16; x_14 = lean_unbox_usize(x_4); lean_dec(x_4); -x_15 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); +x_15 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_16 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__7(x_1, x_2, x_3, x_14, x_15, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_2); -return x_15; +lean_dec(x_3); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_15 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_16 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8(x_1, x_2, x_3, x_14, x_15, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +return x_16; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_15 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_16 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9(x_1, x_2, x_3, x_14, x_15, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__10___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_15 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_16 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__10(x_1, x_2, x_3, x_14, x_15, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__11___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_15 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_16 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__11(x_1, x_2, x_3, x_14, x_15, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { size_t x_6; size_t x_7; lean_object* x_8; @@ -11725,11 +13067,11 @@ x_6 = lean_unbox_usize(x_3); lean_dec(x_3); x_7 = lean_unbox_usize(x_4); lean_dec(x_4); -x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4(x_1, x_2, x_6, x_7, x_5); +x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__12(x_1, x_2, x_6, x_7, x_5); return x_8; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___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_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___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; @@ -11737,11 +13079,11 @@ x_9 = lean_unbox_usize(x_6); lean_dec(x_6); x_10 = lean_unbox_usize(x_7); lean_dec(x_7); -x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5(x_1, x_2, x_3, x_4, x_5, x_9, x_10, x_8); +x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13(x_1, x_2, x_3, x_4, x_5, x_9, x_10, x_8); return x_11; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___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_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__14___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: { size_t x_7; size_t x_8; lean_object* x_9; @@ -11749,11 +13091,132 @@ x_7 = lean_unbox_usize(x_4); lean_dec(x_4); x_8 = lean_unbox_usize(x_5); lean_dec(x_5); -x_9 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6(x_1, x_2, x_3, x_7, x_8, x_6); +x_9 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__14(x_1, x_2, x_3, x_7, x_8, x_6); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__15___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_15 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_16 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__15(x_1, x_2, x_3, x_14, x_15, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__16___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_15 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_16 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__16(x_1, x_2, x_3, x_14, x_15, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__17___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_15 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_16 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__17(x_1, x_2, x_3, x_14, x_15, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__18___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_15 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_16 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__18(x_1, x_2, x_3, x_14, x_15, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__19___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_15 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_16 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__19(x_1, x_2, x_3, x_14, x_15, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__20___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_15 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_16 = l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__20(x_1, x_2, x_3, x_14, x_15, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___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_throwError___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); return x_9; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___boxed(lean_object** _args) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; lean_object* x_3 = _args[2]; @@ -11771,16 +13234,17 @@ lean_object* x_14 = _args[13]; lean_object* x_15 = _args[14]; lean_object* x_16 = _args[15]; lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; _start: { -size_t x_18; lean_object* x_19; -x_18 = lean_unbox_usize(x_8); -lean_dec(x_8); -x_19 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_18, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -return x_19; +size_t x_19; lean_object* x_20; +x_19 = lean_unbox_usize(x_9); +lean_dec(x_9); +x_20 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_19, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +return x_20; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___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_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, lean_object* x_16) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___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_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, lean_object* x_16) { _start: { size_t x_17; size_t x_18; size_t x_19; lean_object* x_20; @@ -11790,18 +13254,18 @@ x_18 = lean_unbox_usize(x_7); lean_dec(x_7); x_19 = lean_unbox_usize(x_8); lean_dec(x_8); -x_20 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9(x_1, x_2, x_3, x_17, x_5, x_6, x_18, x_19, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +x_20 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25(x_1, x_2, x_3, x_17, x_5, x_6, x_18, x_19, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); lean_dec(x_6); lean_dec(x_2); lean_dec(x_1); return x_20; } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__11___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__27___boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__11(x_1); +x_2 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__27(x_1); lean_dec(x_1); return x_2; } @@ -11960,7 +13424,7 @@ x_12 = l_Lean_throwUnknownConstant___at_Lean_Elab_Term_Quotation_getQuotKind___s x_13 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_13, 0, x_12); lean_ctor_set(x_13, 1, x_11); -x_14 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__11; +x_14 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__11; x_15 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_15, 0, x_13); lean_ctor_set(x_15, 1, x_14); @@ -11992,7 +13456,7 @@ x_11 = lean_ctor_get(x_9, 1); lean_inc(x_11); lean_dec(x_9); x_12 = lean_box(0); -x_13 = l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(x_10, x_12); +x_13 = l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(x_10, x_12); x_14 = l_List_isEmpty___rarg(x_13); if (x_14 == 0) { @@ -20137,7 +21601,7 @@ static lean_object* _init_l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__49 _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__57; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__57; x_2 = l_Lean_Elab_Term_Quotation_mkTuple___closed__6; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -20147,7 +21611,7 @@ static lean_object* _init_l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__50 _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__59; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__59; x_2 = l_Lean_Elab_Term_Quotation_mkTuple___closed__6; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -20180,26 +21644,26 @@ return x_3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_stxQuot_expand(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_195; lean_object* x_196; uint8_t x_197; -x_195 = l_Lean_Syntax_getNumArgs(x_1); -x_196 = lean_unsigned_to_nat(1u); -x_197 = lean_nat_dec_eq(x_195, x_196); -lean_dec(x_195); -if (x_197 == 0) +lean_object* x_9; lean_object* x_194; lean_object* x_195; uint8_t x_196; +x_194 = l_Lean_Syntax_getNumArgs(x_1); +x_195 = lean_unsigned_to_nat(1u); +x_196 = lean_nat_dec_eq(x_194, x_195); +lean_dec(x_194); +if (x_196 == 0) { x_9 = x_1; -goto block_194; +goto block_193; } else { -lean_object* x_198; lean_object* x_199; -x_198 = lean_unsigned_to_nat(0u); -x_199 = l_Lean_Syntax_getArg(x_1, x_198); +lean_object* x_197; lean_object* x_198; +x_197 = lean_unsigned_to_nat(0u); +x_198 = l_Lean_Syntax_getArg(x_1, x_197); lean_dec(x_1); -x_9 = x_199; -goto block_194; +x_9 = x_198; +goto block_193; } -block_194: +block_193: { lean_object* x_10; lean_inc(x_7); @@ -20283,12 +21747,12 @@ lean_ctor_set(x_37, 0, x_18); lean_ctor_set(x_37, 1, x_35); lean_ctor_set(x_37, 2, x_34); lean_ctor_set(x_37, 3, x_36); -x_38 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; +x_38 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; lean_inc(x_18); x_39 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_39, 0, x_18); lean_ctor_set(x_39, 1, x_38); -x_40 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__8; +x_40 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__8; lean_inc(x_18); x_41 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_41, 0, x_18); @@ -20312,7 +21776,7 @@ x_50 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_50, 0, x_48); lean_ctor_set(x_50, 1, x_49); lean_ctor_set(x_50, 2, x_47); -x_51 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__12; +x_51 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__12; lean_inc(x_18); x_52 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_52, 0, x_18); @@ -20385,7 +21849,7 @@ lean_ctor_set(x_79, 0, x_18); lean_ctor_set(x_79, 1, x_77); lean_ctor_set(x_79, 2, x_76); lean_ctor_set(x_79, 3, x_78); -x_80 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__21; +x_80 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__21; lean_inc(x_18); x_81 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_81, 0, x_18); @@ -20403,7 +21867,7 @@ lean_ctor_set(x_86, 3, x_85); x_87 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; x_88 = lean_array_push(x_87, x_81); x_89 = lean_array_push(x_88, x_86); -x_90 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__20; +x_90 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__20; x_91 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_91, 0, x_48); lean_ctor_set(x_91, 1, x_90); @@ -20411,11 +21875,11 @@ lean_ctor_set(x_91, 2, x_89); lean_inc(x_11); x_92 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_29, x_11); x_93 = lean_array_push(x_87, x_91); -x_94 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; +x_94 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; x_95 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_95, 0, x_18); lean_ctor_set(x_95, 1, x_94); -x_96 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; +x_96 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; x_97 = lean_array_push(x_96, x_39); x_98 = lean_array_push(x_87, x_79); x_99 = lean_array_push(x_96, x_74); @@ -20434,31 +21898,30 @@ x_109 = lean_array_push(x_87, x_37); if (lean_obj_tag(x_92) == 0) { lean_object* x_175; -x_175 = l___private_Init_Meta_0__Lean_quoteNameMk(x_11); +x_175 = l_Lean_quoteNameMk(x_11); x_110 = x_175; goto block_174; } else { -lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; +lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_dec(x_11); x_176 = lean_ctor_get(x_92, 0); lean_inc(x_176); lean_dec(x_92); -x_177 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__3; +x_177 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; x_178 = l_String_intercalate(x_177, x_176); -x_179 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__4; +x_179 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4; x_180 = lean_string_append(x_179, x_178); lean_dec(x_178); -x_181 = l_Lean_nameLitKind; -x_182 = l_Lean_Syntax_mkLit(x_181, x_180, x_48); -x_183 = lean_array_push(x_46, x_182); -x_184 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__2; -x_185 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_185, 0, x_48); -lean_ctor_set(x_185, 1, x_184); -lean_ctor_set(x_185, 2, x_183); -x_110 = x_185; +x_181 = l_Lean_Syntax_mkNameLit(x_180, x_48); +x_182 = lean_array_push(x_46, x_181); +x_183 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2; +x_184 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_184, 0, x_48); +lean_ctor_set(x_184, 1, x_183); +lean_ctor_set(x_184, 2, x_182); +x_110 = x_184; goto block_174; } block_174: @@ -20487,7 +21950,7 @@ lean_inc(x_97); x_121 = lean_array_push(x_97, x_120); lean_inc(x_95); x_122 = lean_array_push(x_121, x_95); -x_123 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__6; +x_123 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6; x_124 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_124, 0, x_48); lean_ctor_set(x_124, 1, x_123); @@ -20503,14 +21966,14 @@ lean_ctor_set(x_128, 0, x_48); lean_ctor_set(x_128, 1, x_115); lean_ctor_set(x_128, 2, x_127); x_129 = lean_array_push(x_100, x_128); -x_130 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__11; +x_130 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__11; x_131 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_131, 0, x_48); lean_ctor_set(x_131, 1, x_130); lean_ctor_set(x_131, 2, x_129); lean_inc(x_101); x_132 = lean_array_push(x_101, x_131); -x_133 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__9; +x_133 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__9; x_134 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_134, 0, x_48); lean_ctor_set(x_134, 1, x_133); @@ -20620,33 +22083,33 @@ return x_173; } else { -uint8_t x_186; +uint8_t x_185; lean_dec(x_11); lean_dec(x_7); lean_dec(x_6); -x_186 = !lean_is_exclusive(x_14); -if (x_186 == 0) +x_185 = !lean_is_exclusive(x_14); +if (x_185 == 0) { return x_14; } else { -lean_object* x_187; lean_object* x_188; lean_object* x_189; -x_187 = lean_ctor_get(x_14, 0); -x_188 = lean_ctor_get(x_14, 1); -lean_inc(x_188); +lean_object* x_186; lean_object* x_187; lean_object* x_188; +x_186 = lean_ctor_get(x_14, 0); +x_187 = lean_ctor_get(x_14, 1); lean_inc(x_187); +lean_inc(x_186); lean_dec(x_14); -x_189 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_189, 0, x_187); -lean_ctor_set(x_189, 1, x_188); -return x_189; +x_188 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_188, 0, x_186); +lean_ctor_set(x_188, 1, x_187); +return x_188; } } } else { -uint8_t x_190; +uint8_t x_189; lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); @@ -20654,29 +22117,29 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_190 = !lean_is_exclusive(x_10); -if (x_190 == 0) +x_189 = !lean_is_exclusive(x_10); +if (x_189 == 0) { return x_10; } else { -lean_object* x_191; lean_object* x_192; lean_object* x_193; -x_191 = lean_ctor_get(x_10, 0); -x_192 = lean_ctor_get(x_10, 1); -lean_inc(x_192); +lean_object* x_190; lean_object* x_191; lean_object* x_192; +x_190 = lean_ctor_get(x_10, 0); +x_191 = lean_ctor_get(x_10, 1); lean_inc(x_191); +lean_inc(x_190); lean_dec(x_10); -x_193 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_193, 0, x_191); -lean_ctor_set(x_193, 1, x_192); -return x_193; +x_192 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_192, 0, x_190); +lean_ctor_set(x_192, 1, x_191); +return x_192; } } } } } -static lean_object* _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__1() { +static lean_object* _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__1() { _start: { lean_object* x_1; @@ -20684,27 +22147,27 @@ x_1 = lean_mk_string_from_bytes("Elab", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__2() { +static lean_object* _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__2; -x_2 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__1; +x_2 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__3() { +static lean_object* _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__2; +x_1 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__2; x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__4() { +static lean_object* _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__4() { _start: { lean_object* x_1; @@ -20712,35 +22175,35 @@ x_1 = lean_mk_string_from_bytes("Quotation", 9); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__5() { +static lean_object* _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__3; -x_2 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__4; +x_1 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__3; +x_2 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__6() { +static lean_object* _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__6() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("commandElab_stx_quot__", 22); +x_1 = lean_mk_string_from_bytes("commandElab_stx_quot_", 21); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__7() { +static lean_object* _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__5; -x_2 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__6; +x_1 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__5; +x_2 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__6; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__8() { +static lean_object* _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__8() { _start: { lean_object* x_1; @@ -20748,17 +22211,17 @@ x_1 = lean_mk_string_from_bytes("andthen", 7); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__9() { +static lean_object* _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__8; +x_2 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__8; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__10() { +static lean_object* _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__10() { _start: { lean_object* x_1; @@ -20766,17 +22229,17 @@ x_1 = lean_mk_string_from_bytes("elab_stx_quot", 13); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__11() { +static lean_object* _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__10; +x_1 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__10; x_2 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__12() { +static lean_object* _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -20786,23 +22249,23 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__13() { +static lean_object* _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__13() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__12; +x_1 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__12; x_2 = lean_alloc_ctor(8, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__14() { +static lean_object* _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__9; -x_2 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__11; -x_3 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__13; +x_1 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__9; +x_2 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__11; +x_3 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__13; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -20810,13 +22273,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__15() { +static lean_object* _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__7; +x_1 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__7; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__14; +x_3 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__14; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -20824,15 +22287,15 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot____() { +static lean_object* _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot__() { _start: { lean_object* x_1; -x_1 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__15; +x_1 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__15; return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__1() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -20842,7 +22305,7 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__2() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__2() { _start: { lean_object* x_1; @@ -20850,17 +22313,17 @@ x_1 = lean_mk_string_from_bytes("declaration", 11); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__3() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__1; -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__2; +x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__1; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__2; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__4() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__4() { _start: { lean_object* x_1; @@ -20868,17 +22331,17 @@ x_1 = lean_mk_string_from_bytes("declModifiers", 13); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__5() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__1; -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__4; +x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__1; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__6() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__6() { _start: { lean_object* x_1; @@ -20886,17 +22349,17 @@ x_1 = lean_mk_string_from_bytes("attributes", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__7() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__6; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__6; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__8() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__8() { _start: { lean_object* x_1; @@ -20904,7 +22367,7 @@ x_1 = lean_mk_string_from_bytes("@[", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__9() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__9() { _start: { lean_object* x_1; @@ -20912,17 +22375,17 @@ x_1 = lean_mk_string_from_bytes("attrInstance", 12); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__10() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__9; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__9; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__11() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__11() { _start: { lean_object* x_1; @@ -20930,17 +22393,17 @@ x_1 = lean_mk_string_from_bytes("attrKind", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__12() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__11; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__11; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__13() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -20950,13 +22413,13 @@ x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__14() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__12; -x_3 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__13; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__12; +x_3 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__13; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -20964,7 +22427,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__15() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__15() { _start: { lean_object* x_1; @@ -20972,17 +22435,17 @@ x_1 = lean_mk_string_from_bytes("Attr", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__16() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__4; -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__15; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__15; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__17() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__17() { _start: { lean_object* x_1; @@ -20990,17 +22453,17 @@ x_1 = lean_mk_string_from_bytes("simple", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__18() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__16; -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__17; +x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__16; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__17; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__19() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__19() { _start: { lean_object* x_1; @@ -21008,22 +22471,22 @@ x_1 = lean_mk_string_from_bytes("builtinTermElab", 15); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__20() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__20() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__19; +x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__19; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__21() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__19; +x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__19; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__20; +x_3 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__20; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -21031,27 +22494,27 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__22() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__19; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__19; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__23() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__14; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__14; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__24() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__24() { _start: { lean_object* x_1; @@ -21059,17 +22522,17 @@ x_1 = lean_mk_string_from_bytes("]", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__25() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__30; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__30; x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__26() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__26() { _start: { lean_object* x_1; @@ -21077,17 +22540,17 @@ x_1 = lean_mk_string_from_bytes("def", 3); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__27() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__27() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__1; -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__26; +x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__1; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__26; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__28() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__28() { _start: { lean_object* x_1; @@ -21095,17 +22558,17 @@ x_1 = lean_mk_string_from_bytes("declId", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__29() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__29() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__1; -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__28; +x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__1; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__28; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__30() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__30() { _start: { lean_object* x_1; @@ -21113,22 +22576,22 @@ x_1 = lean_mk_string_from_bytes("elabQuot", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__31() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__31() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__30; +x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__30; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__32() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__32() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__30; +x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__30; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__31; +x_3 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__31; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -21136,17 +22599,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__33() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__33() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__30; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__30; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__34() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__34() { _start: { lean_object* x_1; @@ -21154,17 +22617,17 @@ x_1 = lean_mk_string_from_bytes("optDeclSig", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__35() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__35() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__1; -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__34; +x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__1; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__34; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__36() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__36() { _start: { lean_object* x_1; @@ -21172,17 +22635,17 @@ x_1 = lean_mk_string_from_bytes("typeSpec", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__37() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__37() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__36; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__36; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__38() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__38() { _start: { lean_object* x_1; @@ -21190,7 +22653,7 @@ x_1 = lean_mk_string_from_bytes(":", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__39() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__39() { _start: { lean_object* x_1; @@ -21198,22 +22661,22 @@ x_1 = lean_mk_string_from_bytes("TermElab", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__40() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__40() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__39; +x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__39; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__41() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__41() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__39; +x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__39; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__40; +x_3 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__40; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -21221,51 +22684,51 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__42() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__42() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__39; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__39; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__43() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__43() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__3; -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__39; +x_1 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__3; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__39; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__44() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__44() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__43; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__43; 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_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__45() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__45() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__44; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__44; x_3 = lean_alloc_ctor(1, 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_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__46() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__46() { _start: { lean_object* x_1; @@ -21273,17 +22736,17 @@ x_1 = lean_mk_string_from_bytes("declValSimple", 13); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__47() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__47() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__1; -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__46; +x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__1; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__46; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__48() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__48() { _start: { lean_object* x_1; @@ -21291,22 +22754,22 @@ x_1 = lean_mk_string_from_bytes("adaptExpander", 13); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__49() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__49() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__48; +x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__48; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__50() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__50() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__48; +x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__48; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__49; +x_3 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__49; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -21314,51 +22777,51 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__51() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__51() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__48; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__48; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__52() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__52() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__3; -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__48; +x_1 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__3; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__48; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__53() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__53() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__52; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__52; 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_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__54() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__54() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__53; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__53; x_3 = lean_alloc_ctor(1, 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_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__55() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__55() { _start: { lean_object* x_1; @@ -21366,22 +22829,22 @@ x_1 = lean_mk_string_from_bytes("stxQuot.expand", 14); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__56() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__56() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__55; +x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__55; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__57() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__57() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__55; +x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__55; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__56; +x_3 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__56; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -21389,7 +22852,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__58() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__58() { _start: { lean_object* x_1; @@ -21397,17 +22860,17 @@ x_1 = lean_mk_string_from_bytes("stxQuot", 7); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__59() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__59() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__58; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__58; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__60() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__60() { _start: { lean_object* x_1; @@ -21415,61 +22878,61 @@ x_1 = lean_mk_string_from_bytes("expand", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__61() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__61() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__59; -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__60; +x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__59; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__60; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__62() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__62() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__5; -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__58; +x_1 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__5; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__58; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__63() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__63() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__62; -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__60; +x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__62; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__60; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__64() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__64() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__63; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__63; 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_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__65() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__65() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__64; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__64; x_3 = lean_alloc_ctor(1, 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_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__66() { +static lean_object* _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__66() { _start: { lean_object* x_1; lean_object* x_2; @@ -21478,11 +22941,11 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; uint8_t x_5; -x_4 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__7; +x_4 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__7; lean_inc(x_1); x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); if (x_5 == 0) @@ -21514,17 +22977,17 @@ lean_inc(x_13); x_14 = lean_ctor_get(x_2, 1); lean_inc(x_14); lean_dec(x_2); -x_15 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__8; +x_15 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__8; lean_inc(x_12); x_16 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_16, 0, x_12); lean_ctor_set(x_16, 1, x_15); -x_17 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__22; +x_17 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__22; lean_inc(x_13); lean_inc(x_14); x_18 = l_Lean_addMacroScope(x_14, x_17, x_13); x_19 = lean_box(0); -x_20 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__21; +x_20 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__21; lean_inc(x_12); x_21 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_21, 0, x_12); @@ -21542,14 +23005,14 @@ lean_ctor_set(x_26, 2, x_23); x_27 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; x_28 = lean_array_push(x_27, x_21); x_29 = lean_array_push(x_28, x_26); -x_30 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__18; +x_30 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__18; x_31 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_31, 0, x_24); lean_ctor_set(x_31, 1, x_30); lean_ctor_set(x_31, 2, x_29); -x_32 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__23; +x_32 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__23; x_33 = lean_array_push(x_32, x_31); -x_34 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__10; +x_34 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__10; x_35 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_35, 0, x_24); lean_ctor_set(x_35, 1, x_34); @@ -21559,16 +23022,16 @@ x_37 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_37, 0, x_24); lean_ctor_set(x_37, 1, x_25); lean_ctor_set(x_37, 2, x_36); -x_38 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__24; +x_38 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__24; lean_inc(x_12); x_39 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_39, 0, x_12); lean_ctor_set(x_39, 1, x_38); -x_40 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; +x_40 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; x_41 = lean_array_push(x_40, x_16); x_42 = lean_array_push(x_41, x_37); x_43 = lean_array_push(x_42, x_39); -x_44 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__7; +x_44 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__7; x_45 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_45, 0, x_24); lean_ctor_set(x_45, 1, x_44); @@ -21578,28 +23041,28 @@ x_47 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_47, 0, x_24); lean_ctor_set(x_47, 1, x_25); lean_ctor_set(x_47, 2, x_46); -x_48 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__25; +x_48 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__25; x_49 = lean_array_push(x_48, x_47); x_50 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; x_51 = lean_array_push(x_49, x_50); x_52 = lean_array_push(x_51, x_50); x_53 = lean_array_push(x_52, x_50); x_54 = lean_array_push(x_53, x_50); -x_55 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__5; +x_55 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__5; x_56 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_56, 0, x_24); lean_ctor_set(x_56, 1, x_55); lean_ctor_set(x_56, 2, x_54); -x_57 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__26; +x_57 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__26; lean_inc(x_12); x_58 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_58, 0, x_12); lean_ctor_set(x_58, 1, x_57); -x_59 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__33; +x_59 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__33; lean_inc(x_13); lean_inc(x_14); x_60 = l_Lean_addMacroScope(x_14, x_59, x_13); -x_61 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__32; +x_61 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__32; lean_inc(x_12); x_62 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_62, 0, x_12); @@ -21608,22 +23071,22 @@ lean_ctor_set(x_62, 2, x_60); lean_ctor_set(x_62, 3, x_19); x_63 = lean_array_push(x_27, x_62); x_64 = lean_array_push(x_63, x_50); -x_65 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__29; +x_65 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__29; x_66 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_66, 0, x_24); lean_ctor_set(x_66, 1, x_65); lean_ctor_set(x_66, 2, x_64); -x_67 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__38; +x_67 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__38; lean_inc(x_12); x_68 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_68, 0, x_12); lean_ctor_set(x_68, 1, x_67); -x_69 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__42; +x_69 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__42; lean_inc(x_13); lean_inc(x_14); x_70 = l_Lean_addMacroScope(x_14, x_69, x_13); -x_71 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__41; -x_72 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__45; +x_71 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__41; +x_72 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__45; lean_inc(x_12); x_73 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_73, 0, x_12); @@ -21632,7 +23095,7 @@ lean_ctor_set(x_73, 2, x_70); lean_ctor_set(x_73, 3, x_72); x_74 = lean_array_push(x_27, x_68); x_75 = lean_array_push(x_74, x_73); -x_76 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__37; +x_76 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__37; x_77 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_77, 0, x_24); lean_ctor_set(x_77, 1, x_76); @@ -21642,9 +23105,9 @@ x_79 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_79, 0, x_24); lean_ctor_set(x_79, 1, x_25); lean_ctor_set(x_79, 2, x_78); -x_80 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__63; +x_80 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__63; x_81 = lean_array_push(x_80, x_79); -x_82 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__35; +x_82 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__35; x_83 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_83, 0, x_24); lean_ctor_set(x_83, 1, x_82); @@ -21654,22 +23117,22 @@ lean_inc(x_12); x_85 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_85, 0, x_12); lean_ctor_set(x_85, 1, x_84); -x_86 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__51; +x_86 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__51; lean_inc(x_13); lean_inc(x_14); x_87 = l_Lean_addMacroScope(x_14, x_86, x_13); -x_88 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__50; -x_89 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__54; +x_88 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__50; +x_89 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__54; lean_inc(x_12); x_90 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_90, 0, x_12); lean_ctor_set(x_90, 1, x_88); lean_ctor_set(x_90, 2, x_87); lean_ctor_set(x_90, 3, x_89); -x_91 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__61; +x_91 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__61; x_92 = l_Lean_addMacroScope(x_14, x_91, x_13); -x_93 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__57; -x_94 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__65; +x_93 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__57; +x_94 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__65; x_95 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_95, 0, x_12); lean_ctor_set(x_95, 1, x_93); @@ -21690,12 +23153,12 @@ lean_ctor_set(x_101, 2, x_99); x_102 = lean_array_push(x_40, x_85); x_103 = lean_array_push(x_102, x_101); x_104 = lean_array_push(x_103, x_50); -x_105 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__47; +x_105 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__47; x_106 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_106, 0, x_24); lean_ctor_set(x_106, 1, x_105); lean_ctor_set(x_106, 2, x_104); -x_107 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__66; +x_107 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__66; x_108 = lean_array_push(x_107, x_58); x_109 = lean_array_push(x_108, x_66); x_110 = lean_array_push(x_109, x_83); @@ -21703,14 +23166,14 @@ x_111 = lean_array_push(x_110, x_106); x_112 = lean_array_push(x_111, x_50); x_113 = lean_array_push(x_112, x_50); x_114 = lean_array_push(x_113, x_50); -x_115 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__27; +x_115 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__27; x_116 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_116, 0, x_24); lean_ctor_set(x_116, 1, x_115); lean_ctor_set(x_116, 2, x_114); x_117 = lean_array_push(x_27, x_56); x_118 = lean_array_push(x_117, x_116); -x_119 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__3; +x_119 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__3; x_120 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_120, 0, x_24); lean_ctor_set(x_120, 1, x_119); @@ -21731,17 +23194,17 @@ lean_inc(x_123); x_124 = lean_ctor_get(x_2, 1); lean_inc(x_124); lean_dec(x_2); -x_125 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__8; +x_125 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__8; lean_inc(x_121); x_126 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_126, 0, x_121); lean_ctor_set(x_126, 1, x_125); -x_127 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__22; +x_127 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__22; lean_inc(x_123); lean_inc(x_124); x_128 = l_Lean_addMacroScope(x_124, x_127, x_123); x_129 = lean_box(0); -x_130 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__21; +x_130 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__21; lean_inc(x_121); x_131 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_131, 0, x_121); @@ -21759,14 +23222,14 @@ lean_ctor_set(x_136, 2, x_133); x_137 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; x_138 = lean_array_push(x_137, x_131); x_139 = lean_array_push(x_138, x_136); -x_140 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__18; +x_140 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__18; x_141 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_141, 0, x_134); lean_ctor_set(x_141, 1, x_140); lean_ctor_set(x_141, 2, x_139); -x_142 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__23; +x_142 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__23; x_143 = lean_array_push(x_142, x_141); -x_144 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__10; +x_144 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__10; x_145 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_145, 0, x_134); lean_ctor_set(x_145, 1, x_144); @@ -21776,16 +23239,16 @@ x_147 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_147, 0, x_134); lean_ctor_set(x_147, 1, x_135); lean_ctor_set(x_147, 2, x_146); -x_148 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__24; +x_148 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__24; lean_inc(x_121); x_149 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_149, 0, x_121); lean_ctor_set(x_149, 1, x_148); -x_150 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; +x_150 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; x_151 = lean_array_push(x_150, x_126); x_152 = lean_array_push(x_151, x_147); x_153 = lean_array_push(x_152, x_149); -x_154 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__7; +x_154 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__7; x_155 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_155, 0, x_134); lean_ctor_set(x_155, 1, x_154); @@ -21795,28 +23258,28 @@ x_157 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_157, 0, x_134); lean_ctor_set(x_157, 1, x_135); lean_ctor_set(x_157, 2, x_156); -x_158 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__25; +x_158 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__25; x_159 = lean_array_push(x_158, x_157); x_160 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; x_161 = lean_array_push(x_159, x_160); x_162 = lean_array_push(x_161, x_160); x_163 = lean_array_push(x_162, x_160); x_164 = lean_array_push(x_163, x_160); -x_165 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__5; +x_165 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__5; x_166 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_166, 0, x_134); lean_ctor_set(x_166, 1, x_165); lean_ctor_set(x_166, 2, x_164); -x_167 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__26; +x_167 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__26; lean_inc(x_121); x_168 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_168, 0, x_121); lean_ctor_set(x_168, 1, x_167); -x_169 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__33; +x_169 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__33; lean_inc(x_123); lean_inc(x_124); x_170 = l_Lean_addMacroScope(x_124, x_169, x_123); -x_171 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__32; +x_171 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__32; lean_inc(x_121); x_172 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_172, 0, x_121); @@ -21825,22 +23288,22 @@ lean_ctor_set(x_172, 2, x_170); lean_ctor_set(x_172, 3, x_129); x_173 = lean_array_push(x_137, x_172); x_174 = lean_array_push(x_173, x_160); -x_175 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__29; +x_175 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__29; x_176 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_176, 0, x_134); lean_ctor_set(x_176, 1, x_175); lean_ctor_set(x_176, 2, x_174); -x_177 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__38; +x_177 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__38; lean_inc(x_121); x_178 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_178, 0, x_121); lean_ctor_set(x_178, 1, x_177); -x_179 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__42; +x_179 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__42; lean_inc(x_123); lean_inc(x_124); x_180 = l_Lean_addMacroScope(x_124, x_179, x_123); -x_181 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__41; -x_182 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__45; +x_181 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__41; +x_182 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__45; lean_inc(x_121); x_183 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_183, 0, x_121); @@ -21849,7 +23312,7 @@ lean_ctor_set(x_183, 2, x_180); lean_ctor_set(x_183, 3, x_182); x_184 = lean_array_push(x_137, x_178); x_185 = lean_array_push(x_184, x_183); -x_186 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__37; +x_186 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__37; x_187 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_187, 0, x_134); lean_ctor_set(x_187, 1, x_186); @@ -21859,9 +23322,9 @@ x_189 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_189, 0, x_134); lean_ctor_set(x_189, 1, x_135); lean_ctor_set(x_189, 2, x_188); -x_190 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__63; +x_190 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__63; x_191 = lean_array_push(x_190, x_189); -x_192 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__35; +x_192 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__35; x_193 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_193, 0, x_134); lean_ctor_set(x_193, 1, x_192); @@ -21871,22 +23334,22 @@ lean_inc(x_121); x_195 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_195, 0, x_121); lean_ctor_set(x_195, 1, x_194); -x_196 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__51; +x_196 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__51; lean_inc(x_123); lean_inc(x_124); x_197 = l_Lean_addMacroScope(x_124, x_196, x_123); -x_198 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__50; -x_199 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__54; +x_198 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__50; +x_199 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__54; lean_inc(x_121); x_200 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_200, 0, x_121); lean_ctor_set(x_200, 1, x_198); lean_ctor_set(x_200, 2, x_197); lean_ctor_set(x_200, 3, x_199); -x_201 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__61; +x_201 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__61; x_202 = l_Lean_addMacroScope(x_124, x_201, x_123); -x_203 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__57; -x_204 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__65; +x_203 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__57; +x_204 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__65; x_205 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_205, 0, x_121); lean_ctor_set(x_205, 1, x_203); @@ -21907,12 +23370,12 @@ lean_ctor_set(x_211, 2, x_209); x_212 = lean_array_push(x_150, x_195); x_213 = lean_array_push(x_212, x_211); x_214 = lean_array_push(x_213, x_160); -x_215 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__47; +x_215 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__47; x_216 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_216, 0, x_134); lean_ctor_set(x_216, 1, x_215); lean_ctor_set(x_216, 2, x_214); -x_217 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__66; +x_217 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__66; x_218 = lean_array_push(x_217, x_168); x_219 = lean_array_push(x_218, x_176); x_220 = lean_array_push(x_219, x_193); @@ -21920,14 +23383,14 @@ x_221 = lean_array_push(x_220, x_216); x_222 = lean_array_push(x_221, x_160); x_223 = lean_array_push(x_222, x_160); x_224 = lean_array_push(x_223, x_160); -x_225 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__27; +x_225 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__27; x_226 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_226, 0, x_134); lean_ctor_set(x_226, 1, x_225); lean_ctor_set(x_226, 2, x_224); x_227 = lean_array_push(x_137, x_166); x_228 = lean_array_push(x_227, x_226); -x_229 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__3; +x_229 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__3; x_230 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_230, 0, x_134); lean_ctor_set(x_230, 1, x_229); @@ -21940,7 +23403,7 @@ return x_231; } } } -static lean_object* _init_l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__1() { +static lean_object* _init_l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__1() { _start: { lean_object* x_1; @@ -21948,16 +23411,16 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_stxQuot_expand), 8, return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633_(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_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874_(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_object* x_9) { _start: { lean_object* x_10; lean_object* x_11; -x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__1; +x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__1; x_11 = l_Lean_Elab_Term_adaptExpander(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -21967,27 +23430,27 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__1; x_2 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__5; -x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__30; +x_1 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__5; +x_2 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__30; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__4() { _start: { lean_object* x_1; @@ -21995,47 +23458,47 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__5() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__3; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__4; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__3; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__6() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__5; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__5; x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__7() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__6; -x_2 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__6; +x_2 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__8() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__7; -x_2 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__4; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__7; +x_2 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__9() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__9() { _start: { lean_object* x_1; @@ -22043,27 +23506,27 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__10() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__8; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__9; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__8; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__9; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__11() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__10; -x_2 = lean_unsigned_to_nat(4633u); +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__10; +x_2 = lean_unsigned_to_nat(4874u); x_3 = lean_name_mk_numeral(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__12() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__12() { _start: { lean_object* x_1; @@ -22071,31 +23534,31 @@ x_1 = l_Lean_Elab_Term_termElabAttribute; return x_1; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__13() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__13() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633_), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874_), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874_(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; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__12; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__2; -x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__11; -x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__13; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__12; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__2; +x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__11; +x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__13; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); return x_6; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633__declRange___closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874__declRange___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(219u); +x_1 = lean_unsigned_to_nat(222u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22103,11 +23566,11 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633__declRange___closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874__declRange___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(219u); +x_1 = lean_unsigned_to_nat(222u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22115,13 +23578,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633__declRange___closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874__declRange___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633__declRange___closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874__declRange___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633__declRange___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874__declRange___closed__2; x_4 = lean_unsigned_to_nat(31u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); @@ -22131,37 +23594,37 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633__declRange___closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874__declRange___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633__declRange___closed__3; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874__declRange___closed__3; x_2 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_2, 0, x_1); lean_ctor_set(x_2, 1, x_1); return x_2; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633__declRange(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874__declRange(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__11; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633__declRange___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__11; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874__declRange___closed__4; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639_(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_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880_(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_object* x_9) { _start: { lean_object* x_10; lean_object* x_11; -x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__1; +x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__1; x_11 = l_Lean_Elab_Term_adaptExpander(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -22171,41 +23634,41 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639____closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__10; -x_2 = lean_unsigned_to_nat(4639u); +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__10; +x_2 = lean_unsigned_to_nat(4880u); x_3 = lean_name_mk_numeral(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639____closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880____closed__3() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639_), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880_), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880_(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; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__12; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639____closed__1; -x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639____closed__2; -x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639____closed__3; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__12; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880____closed__1; +x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880____closed__2; +x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880____closed__3; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); return x_6; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639__declRange___closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880__declRange___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(220u); +x_1 = lean_unsigned_to_nat(223u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22213,11 +23676,11 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639__declRange___closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880__declRange___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(220u); +x_1 = lean_unsigned_to_nat(223u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22225,13 +23688,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639__declRange___closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880__declRange___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639__declRange___closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880__declRange___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639__declRange___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880__declRange___closed__2; x_4 = lean_unsigned_to_nat(30u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); @@ -22241,37 +23704,37 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639__declRange___closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880__declRange___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639__declRange___closed__3; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880__declRange___closed__3; x_2 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_2, 0, x_1); lean_ctor_set(x_2, 1, x_1); return x_2; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639__declRange(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880__declRange(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639____closed__2; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639__declRange___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880____closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880__declRange___closed__4; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645_(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_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886_(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_object* x_9) { _start: { lean_object* x_10; lean_object* x_11; -x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__1; +x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__1; x_11 = l_Lean_Elab_Term_adaptExpander(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886____closed__1() { _start: { lean_object* x_1; @@ -22279,61 +23742,61 @@ x_1 = lean_mk_string_from_bytes("funBinder", 9); return x_1; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645____closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645____closed__1; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645____closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645____closed__2; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886____closed__2; x_2 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645____closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__10; -x_2 = lean_unsigned_to_nat(4645u); +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__10; +x_2 = lean_unsigned_to_nat(4886u); x_3 = lean_name_mk_numeral(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645____closed__5() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886____closed__5() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645_), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886_), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886_(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; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__12; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645____closed__3; -x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645____closed__4; -x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645____closed__5; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__12; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886____closed__3; +x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886____closed__4; +x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886____closed__5; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); return x_6; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645__declRange___closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886__declRange___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(221u); +x_1 = lean_unsigned_to_nat(224u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22341,11 +23804,11 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645__declRange___closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886__declRange___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(221u); +x_1 = lean_unsigned_to_nat(224u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22353,13 +23816,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645__declRange___closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886__declRange___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645__declRange___closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886__declRange___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645__declRange___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886__declRange___closed__2; x_4 = lean_unsigned_to_nat(40u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); @@ -22369,37 +23832,37 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645__declRange___closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886__declRange___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645__declRange___closed__3; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886__declRange___closed__3; x_2 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_2, 0, x_1); lean_ctor_set(x_2, 1, x_1); return x_2; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645__declRange(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886__declRange(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645____closed__4; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645__declRange___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886____closed__4; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886__declRange___closed__4; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651_(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_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892_(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_object* x_9) { _start: { lean_object* x_10; lean_object* x_11; -x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__1; +x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__1; x_11 = l_Lean_Elab_Term_adaptExpander(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892____closed__1() { _start: { lean_object* x_1; @@ -22407,61 +23870,61 @@ x_1 = lean_mk_string_from_bytes("bracketedBinder", 15); return x_1; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651____closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651____closed__1; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651____closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651____closed__2; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892____closed__2; x_2 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651____closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__10; -x_2 = lean_unsigned_to_nat(4651u); +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__10; +x_2 = lean_unsigned_to_nat(4892u); x_3 = lean_name_mk_numeral(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651____closed__5() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892____closed__5() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651_), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892_), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892_(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; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__12; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651____closed__3; -x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651____closed__4; -x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651____closed__5; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__12; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892____closed__3; +x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892____closed__4; +x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892____closed__5; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); return x_6; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651__declRange___closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892__declRange___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(222u); +x_1 = lean_unsigned_to_nat(225u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22469,11 +23932,11 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651__declRange___closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892__declRange___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(222u); +x_1 = lean_unsigned_to_nat(225u); x_2 = lean_unsigned_to_nat(46u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22481,13 +23944,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651__declRange___closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892__declRange___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651__declRange___closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892__declRange___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651__declRange___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892__declRange___closed__2; x_4 = lean_unsigned_to_nat(46u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); @@ -22497,81 +23960,81 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651__declRange___closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892__declRange___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651__declRange___closed__3; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892__declRange___closed__3; x_2 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_2, 0, x_1); lean_ctor_set(x_2, 1, x_1); return x_2; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651__declRange(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892__declRange(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651____closed__4; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651__declRange___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892____closed__4; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892__declRange___closed__4; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657_(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_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898_(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_object* x_9) { _start: { lean_object* x_10; lean_object* x_11; -x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__1; +x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__1; x_11 = l_Lean_Elab_Term_adaptExpander(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__46; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__46; x_2 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657____closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__10; -x_2 = lean_unsigned_to_nat(4657u); +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__10; +x_2 = lean_unsigned_to_nat(4898u); x_3 = lean_name_mk_numeral(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657____closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898____closed__3() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657_), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898_), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898_(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; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__12; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657____closed__1; -x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657____closed__2; -x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657____closed__3; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__12; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898____closed__1; +x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898____closed__2; +x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898____closed__3; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); return x_6; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657__declRange___closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898__declRange___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(223u); +x_1 = lean_unsigned_to_nat(226u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22579,11 +24042,11 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657__declRange___closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898__declRange___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(223u); +x_1 = lean_unsigned_to_nat(226u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22591,13 +24054,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657__declRange___closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898__declRange___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657__declRange___closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898__declRange___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657__declRange___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898__declRange___closed__2; x_4 = lean_unsigned_to_nat(41u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); @@ -22607,37 +24070,37 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657__declRange___closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898__declRange___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657__declRange___closed__3; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898__declRange___closed__3; x_2 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_2, 0, x_1); lean_ctor_set(x_2, 1, x_1); return x_2; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657__declRange(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898__declRange(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657____closed__2; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657__declRange___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898____closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898__declRange___closed__4; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663_(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_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904_(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_object* x_9) { _start: { lean_object* x_10; lean_object* x_11; -x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__1; +x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__1; x_11 = l_Lean_Elab_Term_adaptExpander(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -22647,51 +24110,51 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663____closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663____closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904____closed__1; x_2 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663____closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__10; -x_2 = lean_unsigned_to_nat(4663u); +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__10; +x_2 = lean_unsigned_to_nat(4904u); x_3 = lean_name_mk_numeral(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663____closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904____closed__4() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663_), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904_), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904_(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; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__12; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663____closed__2; -x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663____closed__3; -x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663____closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__12; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904____closed__2; +x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904____closed__3; +x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904____closed__4; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); return x_6; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663__declRange___closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904__declRange___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(224u); +x_1 = lean_unsigned_to_nat(227u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22699,11 +24162,11 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663__declRange___closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904__declRange___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(224u); +x_1 = lean_unsigned_to_nat(227u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22711,13 +24174,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663__declRange___closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904__declRange___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663__declRange___closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904__declRange___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663__declRange___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904__declRange___closed__2; x_4 = lean_unsigned_to_nat(32u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); @@ -22727,81 +24190,81 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663__declRange___closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904__declRange___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663__declRange___closed__3; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904__declRange___closed__3; x_2 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_2, 0, x_1); lean_ctor_set(x_2, 1, x_1); return x_2; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663__declRange(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904__declRange(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663____closed__3; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663__declRange___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904____closed__3; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904__declRange___closed__4; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669_(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_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910_(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_object* x_9) { _start: { lean_object* x_10; lean_object* x_11; -x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__1; +x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__1; x_11 = l_Lean_Elab_Term_adaptExpander(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663____closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904____closed__1; x_2 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__8; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669____closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__10; -x_2 = lean_unsigned_to_nat(4669u); +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__10; +x_2 = lean_unsigned_to_nat(4910u); x_3 = lean_name_mk_numeral(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669____closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910____closed__3() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669_), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910_), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910_(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; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__12; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669____closed__1; -x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669____closed__2; -x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669____closed__3; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__12; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910____closed__1; +x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910____closed__2; +x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910____closed__3; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); return x_6; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669__declRange___closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910__declRange___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(225u); +x_1 = lean_unsigned_to_nat(228u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22809,11 +24272,11 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669__declRange___closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910__declRange___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(225u); +x_1 = lean_unsigned_to_nat(228u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22821,13 +24284,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669__declRange___closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910__declRange___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669__declRange___closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910__declRange___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669__declRange___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910__declRange___closed__2; x_4 = lean_unsigned_to_nat(35u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); @@ -22837,37 +24300,37 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669__declRange___closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910__declRange___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669__declRange___closed__3; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910__declRange___closed__3; x_2 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_2, 0, x_1); lean_ctor_set(x_2, 1, x_1); return x_2; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669__declRange(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910__declRange(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669____closed__2; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669__declRange___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910____closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910__declRange___closed__4; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675_(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_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916_(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_object* x_9) { _start: { lean_object* x_10; lean_object* x_11; -x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__1; +x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__1; x_11 = l_Lean_Elab_Term_adaptExpander(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -22877,51 +24340,51 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675____closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675____closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916____closed__1; x_2 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675____closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__10; -x_2 = lean_unsigned_to_nat(4675u); +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__10; +x_2 = lean_unsigned_to_nat(4916u); x_3 = lean_name_mk_numeral(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675____closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916____closed__4() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675_), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916_), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916_(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; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__12; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675____closed__2; -x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675____closed__3; -x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675____closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__12; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916____closed__2; +x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916____closed__3; +x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916____closed__4; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); return x_6; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675__declRange___closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916__declRange___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(226u); +x_1 = lean_unsigned_to_nat(229u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22929,11 +24392,11 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675__declRange___closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916__declRange___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(226u); +x_1 = lean_unsigned_to_nat(229u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22941,13 +24404,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675__declRange___closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916__declRange___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675__declRange___closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916__declRange___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675__declRange___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916__declRange___closed__2; x_4 = lean_unsigned_to_nat(34u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); @@ -22957,37 +24420,37 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675__declRange___closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916__declRange___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675__declRange___closed__3; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916__declRange___closed__3; x_2 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_2, 0, x_1); lean_ctor_set(x_2, 1, x_1); return x_2; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675__declRange(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916__declRange(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675____closed__3; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675__declRange___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916____closed__3; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916__declRange___closed__4; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681_(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_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922_(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_object* x_9) { _start: { lean_object* x_10; lean_object* x_11; -x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__1; +x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__1; x_11 = l_Lean_Elab_Term_adaptExpander(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -22997,51 +24460,51 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681____closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681____closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922____closed__1; x_2 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681____closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__10; -x_2 = lean_unsigned_to_nat(4681u); +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__10; +x_2 = lean_unsigned_to_nat(4922u); x_3 = lean_name_mk_numeral(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681____closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922____closed__4() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681_), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922_), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922_(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; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__12; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681____closed__2; -x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681____closed__3; -x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681____closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__12; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922____closed__2; +x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922____closed__3; +x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922____closed__4; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); return x_6; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681__declRange___closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922__declRange___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(227u); +x_1 = lean_unsigned_to_nat(230u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23049,11 +24512,11 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681__declRange___closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922__declRange___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(227u); +x_1 = lean_unsigned_to_nat(230u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23061,13 +24524,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681__declRange___closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922__declRange___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681__declRange___closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922__declRange___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681__declRange___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922__declRange___closed__2; x_4 = lean_unsigned_to_nat(35u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); @@ -23077,37 +24540,37 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681__declRange___closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922__declRange___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681__declRange___closed__3; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922__declRange___closed__3; x_2 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_2, 0, x_1); lean_ctor_set(x_2, 1, x_1); return x_2; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681__declRange(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922__declRange(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681____closed__3; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681__declRange___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922____closed__3; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922__declRange___closed__4; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687_(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_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928_(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_object* x_9) { _start: { lean_object* x_10; lean_object* x_11; -x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__1; +x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__1; x_11 = l_Lean_Elab_Term_adaptExpander(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -23117,51 +24580,51 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687____closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687____closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928____closed__1; x_2 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687____closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__10; -x_2 = lean_unsigned_to_nat(4687u); +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__10; +x_2 = lean_unsigned_to_nat(4928u); x_3 = lean_name_mk_numeral(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687____closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928____closed__4() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687_), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928_), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928_(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; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__12; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687____closed__2; -x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687____closed__3; -x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687____closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__12; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928____closed__2; +x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928____closed__3; +x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928____closed__4; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); return x_6; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687__declRange___closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928__declRange___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(228u); +x_1 = lean_unsigned_to_nat(231u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23169,11 +24632,11 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687__declRange___closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928__declRange___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(228u); +x_1 = lean_unsigned_to_nat(231u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23181,13 +24644,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687__declRange___closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928__declRange___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687__declRange___closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928__declRange___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687__declRange___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928__declRange___closed__2; x_4 = lean_unsigned_to_nat(35u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); @@ -23197,37 +24660,37 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687__declRange___closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928__declRange___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687__declRange___closed__3; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928__declRange___closed__3; x_2 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_2, 0, x_1); lean_ctor_set(x_2, 1, x_1); return x_2; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687__declRange(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928__declRange(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687____closed__3; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687__declRange___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928____closed__3; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928__declRange___closed__4; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693_(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_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934_(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_object* x_9) { _start: { lean_object* x_10; lean_object* x_11; -x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__1; +x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__1; x_11 = l_Lean_Elab_Term_adaptExpander(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -23237,51 +24700,51 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693____closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693____closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934____closed__1; x_2 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693____closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__10; -x_2 = lean_unsigned_to_nat(4693u); +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__10; +x_2 = lean_unsigned_to_nat(4934u); x_3 = lean_name_mk_numeral(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693____closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934____closed__4() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693_), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934_), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934_(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; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__12; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693____closed__2; -x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693____closed__3; -x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693____closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__12; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934____closed__2; +x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934____closed__3; +x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934____closed__4; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); return x_6; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693__declRange___closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934__declRange___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(229u); +x_1 = lean_unsigned_to_nat(232u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23289,11 +24752,11 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693__declRange___closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934__declRange___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(229u); +x_1 = lean_unsigned_to_nat(232u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23301,13 +24764,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693__declRange___closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934__declRange___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693__declRange___closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934__declRange___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693__declRange___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934__declRange___closed__2; x_4 = lean_unsigned_to_nat(35u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); @@ -23317,37 +24780,37 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693__declRange___closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934__declRange___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693__declRange___closed__3; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934__declRange___closed__3; x_2 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_2, 0, x_1); lean_ctor_set(x_2, 1, x_1); return x_2; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693__declRange(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934__declRange(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693____closed__3; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693__declRange___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934____closed__3; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934__declRange___closed__4; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699_(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_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940_(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_object* x_9) { _start: { lean_object* x_10; lean_object* x_11; -x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__1; +x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__1; x_11 = l_Lean_Elab_Term_adaptExpander(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -23357,51 +24820,51 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699____closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699____closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940____closed__1; x_2 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699____closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__10; -x_2 = lean_unsigned_to_nat(4699u); +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__10; +x_2 = lean_unsigned_to_nat(4940u); x_3 = lean_name_mk_numeral(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699____closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940____closed__4() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699_), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940_), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940_(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; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__12; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699____closed__2; -x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699____closed__3; -x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699____closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__12; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940____closed__2; +x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940____closed__3; +x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940____closed__4; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); return x_6; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699__declRange___closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940__declRange___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(230u); +x_1 = lean_unsigned_to_nat(233u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23409,11 +24872,11 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699__declRange___closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940__declRange___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(230u); +x_1 = lean_unsigned_to_nat(233u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23421,13 +24884,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699__declRange___closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940__declRange___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699__declRange___closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940__declRange___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699__declRange___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940__declRange___closed__2; x_4 = lean_unsigned_to_nat(37u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); @@ -23437,37 +24900,37 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699__declRange___closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940__declRange___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699__declRange___closed__3; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940__declRange___closed__3; x_2 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_2, 0, x_1); lean_ctor_set(x_2, 1, x_1); return x_2; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699__declRange(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940__declRange(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699____closed__3; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699__declRange___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940____closed__3; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940__declRange___closed__4; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705_(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_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946_(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_object* x_9) { _start: { lean_object* x_10; lean_object* x_11; -x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__1; +x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__1; x_11 = l_Lean_Elab_Term_adaptExpander(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -23477,41 +24940,41 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705____closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__10; -x_2 = lean_unsigned_to_nat(4705u); +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__10; +x_2 = lean_unsigned_to_nat(4946u); x_3 = lean_name_mk_numeral(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705____closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946____closed__3() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705_), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946_), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946_(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; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__12; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705____closed__1; -x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705____closed__2; -x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705____closed__3; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__12; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946____closed__1; +x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946____closed__2; +x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946____closed__3; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); return x_6; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705__declRange___closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946__declRange___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(231u); +x_1 = lean_unsigned_to_nat(234u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23519,11 +24982,11 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705__declRange___closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946__declRange___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(231u); +x_1 = lean_unsigned_to_nat(234u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23531,13 +24994,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705__declRange___closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946__declRange___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705__declRange___closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946__declRange___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705__declRange___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946__declRange___closed__2; x_4 = lean_unsigned_to_nat(37u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); @@ -23547,81 +25010,81 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705__declRange___closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946__declRange___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705__declRange___closed__3; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946__declRange___closed__3; x_2 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_2, 0, x_1); lean_ctor_set(x_2, 1, x_1); return x_2; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705__declRange(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946__declRange(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705____closed__2; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705__declRange___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946____closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946__declRange___closed__4; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711_(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_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952_(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_object* x_9) { _start: { lean_object* x_10; lean_object* x_11; -x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__1; +x_10 = l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__1; x_11 = l_Lean_Elab_Term_adaptExpander(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__1; +x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__1; x_2 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711____closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__10; -x_2 = lean_unsigned_to_nat(4711u); +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__10; +x_2 = lean_unsigned_to_nat(4952u); x_3 = lean_name_mk_numeral(x_1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711____closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952____closed__3() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711_), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952_), 9, 0); return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952_(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; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__12; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711____closed__1; -x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711____closed__2; -x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711____closed__3; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__12; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952____closed__1; +x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952____closed__2; +x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952____closed__3; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); return x_6; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711__declRange___closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952__declRange___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(232u); +x_1 = lean_unsigned_to_nat(235u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23629,11 +25092,11 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711__declRange___closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952__declRange___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(232u); +x_1 = lean_unsigned_to_nat(235u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23641,13 +25104,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711__declRange___closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952__declRange___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711__declRange___closed__1; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952__declRange___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711__declRange___closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952__declRange___closed__2; x_4 = lean_unsigned_to_nat(33u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); @@ -23657,23 +25120,23 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711__declRange___closed__4() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952__declRange___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711__declRange___closed__3; +x_1 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952__declRange___closed__3; x_2 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_2, 0, x_1); lean_ctor_set(x_2, 1, x_1); return x_2; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711__declRange(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952__declRange(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711____closed__2; -x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711__declRange___closed__4; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952____closed__2; +x_3 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952__declRange___closed__4; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } @@ -23683,7 +25146,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_2 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__1; +x_2 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__14___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } @@ -23692,7 +25155,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quota _start: { 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; lean_object* x_11; -x_4 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__2; +x_4 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__14___closed__2; x_5 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_4); @@ -24117,16 +25580,16 @@ lean_inc(x_2); x_6 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_6, 0, x_2); lean_ctor_set(x_6, 1, x_5); -x_7 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; +x_7 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; x_8 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_8, 0, x_2); lean_ctor_set(x_8, 1, x_7); -x_9 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; +x_9 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; x_10 = lean_array_push(x_9, x_6); x_11 = lean_array_push(x_10, x_1); x_12 = lean_array_push(x_11, x_8); x_13 = lean_box(2); -x_14 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639____closed__1; +x_14 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880____closed__1; x_15 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_15, 0, x_13); lean_ctor_set(x_15, 1, x_14); @@ -24137,33 +25600,33 @@ lean_ctor_set(x_16, 1, x_4); return x_16; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__3(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__3(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { -uint8_t x_4; -x_4 = lean_usize_dec_lt(x_2, x_1); -if (x_4 == 0) +uint8_t x_5; +x_5 = lean_usize_dec_lt(x_3, x_2); +if (x_5 == 0) { -return x_3; +return x_4; } else { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; lean_object* x_14; -x_5 = lean_array_uget(x_3, x_2); -x_6 = lean_unsigned_to_nat(0u); -x_7 = lean_array_uset(x_3, x_2, x_6); -x_8 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__3___lambda__1___boxed), 4, 1); -lean_closure_set(x_8, 0, x_5); -x_9 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__2; -x_10 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Unhygienic_instMonadQuotationUnhygienic___spec__2___rarg), 4, 2); -lean_closure_set(x_10, 0, x_9); -lean_closure_set(x_10, 1, x_8); -x_11 = l_Lean_Unhygienic_run___rarg(x_10); -x_12 = 1; -x_13 = lean_usize_add(x_2, x_12); -x_14 = lean_array_uset(x_7, x_2, x_11); -x_2 = x_13; +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; size_t x_13; size_t x_14; lean_object* x_15; +x_6 = lean_array_uget(x_4, x_3); +x_7 = lean_unsigned_to_nat(0u); +x_8 = lean_array_uset(x_4, x_3, x_7); +x_9 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__3___lambda__1___boxed), 4, 1); +lean_closure_set(x_9, 0, x_6); +x_10 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__2; +x_11 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Unhygienic_instMonadQuotationUnhygienic___spec__2___rarg), 4, 2); +lean_closure_set(x_11, 0, x_10); +lean_closure_set(x_11, 1, x_9); +x_12 = l_Lean_Unhygienic_run___rarg(x_11); +x_13 = 1; +x_14 = lean_usize_add(x_3, x_13); +x_15 = lean_array_uset(x_8, x_3, x_12); x_3 = x_14; +x_4 = x_15; goto _start; } } @@ -24221,7 +25684,7 @@ static lean_object* _init_l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__30; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__30; x_2 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -24292,96 +25755,95 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4(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_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4(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_object* x_9) { _start: { -if (lean_obj_tag(x_1) == 0) +if (lean_obj_tag(x_2) == 0) { -lean_object* x_9; lean_object* x_10; -lean_dec(x_6); -x_9 = lean_box(0); -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_10, 1, x_8); -return x_10; +lean_object* x_10; lean_object* x_11; +lean_dec(x_7); +x_10 = lean_box(0); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_9); +return x_11; } else { -uint8_t x_11; -x_11 = !lean_is_exclusive(x_1); -if (x_11 == 0) +uint8_t x_12; +x_12 = !lean_is_exclusive(x_2); +if (x_12 == 0) { -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; 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; -x_12 = lean_ctor_get(x_1, 0); -x_13 = lean_ctor_get(x_1, 1); -lean_inc(x_6); -x_14 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_6, x_7, x_8); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); +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; 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; +x_13 = lean_ctor_get(x_2, 0); +x_14 = lean_ctor_get(x_2, 1); +lean_inc(x_7); +x_15 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_7, x_8, x_9); +x_16 = lean_ctor_get(x_15, 0); lean_inc(x_16); -lean_dec(x_14); -x_17 = lean_ctor_get(x_6, 10); +x_17 = lean_ctor_get(x_15, 1); lean_inc(x_17); -x_18 = lean_st_ref_get(x_7, x_16); -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); +lean_dec(x_15); +x_18 = lean_ctor_get(x_7, 10); +lean_inc(x_18); +x_19 = lean_st_ref_get(x_8, x_17); +x_20 = lean_ctor_get(x_19, 0); lean_inc(x_20); -lean_dec(x_18); -x_21 = lean_ctor_get(x_19, 0); +x_21 = lean_ctor_get(x_19, 1); lean_inc(x_21); lean_dec(x_19); -x_22 = lean_environment_main_module(x_21); -x_23 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__5; -lean_inc(x_17); +x_22 = lean_ctor_get(x_20, 0); lean_inc(x_22); -x_24 = l_Lean_addMacroScope(x_22, x_23, x_17); -x_25 = lean_box(0); -x_26 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__3; -x_27 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__8; -lean_inc(x_15); -x_28 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_28, 0, x_15); -lean_ctor_set(x_28, 1, x_26); -lean_ctor_set(x_28, 2, x_24); -lean_ctor_set(x_28, 3, x_27); -x_29 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -x_30 = l_Lean_addMacroScope(x_22, x_29, x_17); -x_31 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; -x_32 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_32, 0, x_15); -lean_ctor_set(x_32, 1, x_31); -lean_ctor_set(x_32, 2, x_30); -lean_ctor_set(x_32, 3, x_25); -x_33 = l_Nat_repr(x_12); -x_34 = l_Lean_numLitKind; +lean_dec(x_20); +x_23 = lean_environment_main_module(x_22); +x_24 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__5; +lean_inc(x_18); +lean_inc(x_23); +x_25 = l_Lean_addMacroScope(x_23, x_24, x_18); +x_26 = lean_box(0); +x_27 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__3; +x_28 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__8; +lean_inc(x_16); +x_29 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_29, 0, x_16); +lean_ctor_set(x_29, 1, x_27); +lean_ctor_set(x_29, 2, x_25); +lean_ctor_set(x_29, 3, x_28); +x_30 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +x_31 = l_Lean_addMacroScope(x_23, x_30, x_18); +x_32 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; +x_33 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_33, 0, x_16); +lean_ctor_set(x_33, 1, x_32); +lean_ctor_set(x_33, 2, x_31); +lean_ctor_set(x_33, 3, x_26); +x_34 = l_Nat_repr(x_13); x_35 = lean_box(2); -x_36 = l_Lean_Syntax_mkLit(x_34, x_33, x_35); +x_36 = l_Lean_Syntax_mkNumLit(x_34, x_35); x_37 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_38 = lean_array_push(x_37, x_32); +x_38 = lean_array_push(x_37, x_33); x_39 = lean_array_push(x_38, x_36); x_40 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; x_41 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_41, 0, x_35); lean_ctor_set(x_41, 1, x_40); lean_ctor_set(x_41, 2, x_39); -x_42 = lean_array_push(x_37, x_28); +x_42 = lean_array_push(x_37, x_29); x_43 = lean_array_push(x_42, x_41); x_44 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; x_45 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_45, 0, x_35); lean_ctor_set(x_45, 1, x_44); lean_ctor_set(x_45, 2, x_43); -x_46 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4(x_13, x_2, x_3, x_4, x_5, x_6, x_7, x_20); +x_46 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4(x_1, x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_21); x_47 = !lean_is_exclusive(x_46); if (x_47 == 0) { lean_object* x_48; x_48 = lean_ctor_get(x_46, 0); -lean_ctor_set(x_1, 1, x_48); -lean_ctor_set(x_1, 0, x_45); -lean_ctor_set(x_46, 0, x_1); +lean_ctor_set(x_2, 1, x_48); +lean_ctor_set(x_2, 0, x_45); +lean_ctor_set(x_46, 0, x_2); return x_46; } else @@ -24392,32 +25854,32 @@ x_50 = lean_ctor_get(x_46, 1); lean_inc(x_50); lean_inc(x_49); lean_dec(x_46); -lean_ctor_set(x_1, 1, x_49); -lean_ctor_set(x_1, 0, x_45); +lean_ctor_set(x_2, 1, x_49); +lean_ctor_set(x_2, 0, x_45); x_51 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_51, 0, x_1); +lean_ctor_set(x_51, 0, x_2); lean_ctor_set(x_51, 1, x_50); return x_51; } } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; -x_52 = lean_ctor_get(x_1, 0); -x_53 = lean_ctor_get(x_1, 1); +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_52 = lean_ctor_get(x_2, 0); +x_53 = lean_ctor_get(x_2, 1); lean_inc(x_53); lean_inc(x_52); -lean_dec(x_1); -lean_inc(x_6); -x_54 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_6, x_7, x_8); +lean_dec(x_2); +lean_inc(x_7); +x_54 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_7, x_8, x_9); x_55 = lean_ctor_get(x_54, 0); lean_inc(x_55); x_56 = lean_ctor_get(x_54, 1); lean_inc(x_56); lean_dec(x_54); -x_57 = lean_ctor_get(x_6, 10); +x_57 = lean_ctor_get(x_7, 10); lean_inc(x_57); -x_58 = lean_st_ref_get(x_7, x_56); +x_58 = lean_st_ref_get(x_8, x_56); x_59 = lean_ctor_get(x_58, 0); lean_inc(x_59); x_60 = lean_ctor_get(x_58, 1); @@ -24449,48 +25911,47 @@ lean_ctor_set(x_72, 1, x_71); lean_ctor_set(x_72, 2, x_70); lean_ctor_set(x_72, 3, x_65); x_73 = l_Nat_repr(x_52); -x_74 = l_Lean_numLitKind; -x_75 = lean_box(2); -x_76 = l_Lean_Syntax_mkLit(x_74, x_73, x_75); -x_77 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_78 = lean_array_push(x_77, x_72); -x_79 = lean_array_push(x_78, x_76); -x_80 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_81 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_81, 0, x_75); -lean_ctor_set(x_81, 1, x_80); -lean_ctor_set(x_81, 2, x_79); -x_82 = lean_array_push(x_77, x_68); -x_83 = lean_array_push(x_82, x_81); -x_84 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_85 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_85, 0, x_75); -lean_ctor_set(x_85, 1, x_84); -lean_ctor_set(x_85, 2, x_83); -x_86 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4(x_53, x_2, x_3, x_4, x_5, x_6, x_7, x_60); -x_87 = lean_ctor_get(x_86, 0); +x_74 = lean_box(2); +x_75 = l_Lean_Syntax_mkNumLit(x_73, x_74); +x_76 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_77 = lean_array_push(x_76, x_72); +x_78 = lean_array_push(x_77, x_75); +x_79 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_80 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_80, 0, x_74); +lean_ctor_set(x_80, 1, x_79); +lean_ctor_set(x_80, 2, x_78); +x_81 = lean_array_push(x_76, x_68); +x_82 = lean_array_push(x_81, x_80); +x_83 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_84 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_84, 0, x_74); +lean_ctor_set(x_84, 1, x_83); +lean_ctor_set(x_84, 2, x_82); +x_85 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4(x_1, x_53, x_3, x_4, x_5, x_6, x_7, x_8, x_60); +x_86 = lean_ctor_get(x_85, 0); +lean_inc(x_86); +x_87 = lean_ctor_get(x_85, 1); lean_inc(x_87); -x_88 = lean_ctor_get(x_86, 1); -lean_inc(x_88); -if (lean_is_exclusive(x_86)) { - lean_ctor_release(x_86, 0); - lean_ctor_release(x_86, 1); - x_89 = x_86; +if (lean_is_exclusive(x_85)) { + lean_ctor_release(x_85, 0); + lean_ctor_release(x_85, 1); + x_88 = x_85; } else { - lean_dec_ref(x_86); - x_89 = lean_box(0); + lean_dec_ref(x_85); + x_88 = lean_box(0); } -x_90 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_90, 0, x_85); -lean_ctor_set(x_90, 1, x_87); -if (lean_is_scalar(x_89)) { - x_91 = lean_alloc_ctor(0, 2, 0); +x_89 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_89, 0, x_84); +lean_ctor_set(x_89, 1, x_86); +if (lean_is_scalar(x_88)) { + x_90 = lean_alloc_ctor(0, 2, 0); } else { - x_91 = x_89; + x_90 = x_88; } -lean_ctor_set(x_91, 0, x_90); -lean_ctor_set(x_91, 1, x_88); -return x_91; +lean_ctor_set(x_90, 0, x_89); +lean_ctor_set(x_90, 1, x_87); +return x_90; } } } @@ -24545,70 +26006,70 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__6(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_EXPORT lean_object* l_Array_mapIdxM_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__6(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_8; uint8_t x_9; -x_8 = lean_unsigned_to_nat(0u); -x_9 = lean_nat_dec_eq(x_4, x_8); -if (x_9 == 0) +lean_object* x_9; uint8_t x_10; +x_9 = lean_unsigned_to_nat(0u); +x_10 = lean_nat_dec_eq(x_5, x_9); +if (x_10 == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; -x_10 = lean_unsigned_to_nat(1u); -x_11 = lean_nat_sub(x_4, x_10); -lean_dec(x_4); -x_12 = lean_array_fget(x_3, x_5); -x_13 = lean_nat_dec_eq(x_5, x_2); -x_14 = lean_nat_add(x_5, x_10); +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_sub(x_5, x_11); lean_dec(x_5); -if (x_13 == 0) +x_13 = lean_array_fget(x_4, x_6); +x_14 = lean_nat_dec_eq(x_6, x_3); +x_15 = lean_nat_add(x_6, x_11); +lean_dec(x_6); +if (x_14 == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_15 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__3___lambda__1___boxed), 4, 1); -lean_closure_set(x_15, 0, x_12); -x_16 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__2; -x_17 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Unhygienic_instMonadQuotationUnhygienic___spec__2___rarg), 4, 2); -lean_closure_set(x_17, 0, x_16); -lean_closure_set(x_17, 1, x_15); -x_18 = l_Lean_Unhygienic_run___rarg(x_17); -x_19 = lean_array_push(x_7, x_18); -x_4 = x_11; -x_5 = x_14; -x_6 = lean_box(0); -x_7 = x_19; +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_16 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__3___lambda__1___boxed), 4, 1); +lean_closure_set(x_16, 0, x_13); +x_17 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__2; +x_18 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Unhygienic_instMonadQuotationUnhygienic___spec__2___rarg), 4, 2); +lean_closure_set(x_18, 0, x_17); +lean_closure_set(x_18, 1, x_16); +x_19 = l_Lean_Unhygienic_run___rarg(x_18); +x_20 = lean_array_push(x_8, x_19); +x_5 = x_12; +x_6 = x_15; +x_7 = lean_box(0); +x_8 = x_20; goto _start; } else { -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; lean_object* x_30; -x_21 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; -x_22 = lean_array_push(x_21, x_12); -x_23 = lean_box(2); -x_24 = l_Lean_nullKind; -x_25 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_25, 0, x_23); -lean_ctor_set(x_25, 1, x_24); -lean_ctor_set(x_25, 2, x_22); -x_26 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__3___lambda__1___boxed), 4, 1); -lean_closure_set(x_26, 0, x_25); -x_27 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__2; -x_28 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Unhygienic_instMonadQuotationUnhygienic___spec__2___rarg), 4, 2); -lean_closure_set(x_28, 0, x_27); -lean_closure_set(x_28, 1, x_26); -x_29 = l_Lean_Unhygienic_run___rarg(x_28); -x_30 = lean_array_push(x_7, x_29); -x_4 = x_11; -x_5 = x_14; -x_6 = lean_box(0); -x_7 = x_30; +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; lean_object* x_30; lean_object* x_31; +x_22 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; +x_23 = lean_array_push(x_22, x_13); +x_24 = lean_box(2); +x_25 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_26 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +lean_ctor_set(x_26, 2, x_23); +x_27 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__3___lambda__1___boxed), 4, 1); +lean_closure_set(x_27, 0, x_26); +x_28 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__2; +x_29 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Unhygienic_instMonadQuotationUnhygienic___spec__2___rarg), 4, 2); +lean_closure_set(x_29, 0, x_28); +lean_closure_set(x_29, 1, x_27); +x_30 = l_Lean_Unhygienic_run___rarg(x_29); +x_31 = lean_array_push(x_8, x_30); +x_5 = x_12; +x_6 = x_15; +x_7 = lean_box(0); +x_8 = x_31; goto _start; } } else { +lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -return x_7; +return x_8; } } } @@ -24620,18 +26081,19 @@ x_1 = lean_mk_string_from_bytes("-", 1); return x_1; } } -LEAN_EXPORT lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__7(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_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, 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_EXPORT lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__7(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_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, 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, lean_object* x_25) { _start: { -if (lean_obj_tag(x_16) == 0) +if (lean_obj_tag(x_18) == 0) { -lean_object* x_24; lean_object* x_25; -lean_dec(x_21); +lean_object* x_26; lean_object* x_27; +lean_dec(x_23); +lean_dec(x_17); +lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); -lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -24640,339 +26102,336 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -x_24 = lean_box(0); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_23); -return x_25; +x_26 = lean_box(0); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +return x_27; } else { -uint8_t x_26; -x_26 = !lean_is_exclusive(x_16); -if (x_26 == 0) +uint8_t x_28; +x_28 = !lean_is_exclusive(x_18); +if (x_28 == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; uint8_t x_85; -x_27 = lean_ctor_get(x_16, 0); -x_28 = lean_ctor_get(x_16, 1); -lean_inc(x_21); -x_29 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_21, x_22, x_23); -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_29, 1); -lean_inc(x_31); -lean_dec(x_29); -x_32 = lean_ctor_get(x_21, 10); +lean_object* x_29; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; uint8_t x_86; +x_29 = lean_ctor_get(x_18, 0); +x_30 = lean_ctor_get(x_18, 1); +lean_inc(x_23); +x_31 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_23, x_24, x_25); +x_32 = lean_ctor_get(x_31, 0); lean_inc(x_32); -x_33 = lean_st_ref_get(x_22, x_31); -x_34 = lean_ctor_get(x_33, 0); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec(x_31); +x_34 = lean_ctor_get(x_23, 10); lean_inc(x_34); -x_35 = lean_ctor_get(x_33, 1); -lean_inc(x_35); -lean_dec(x_33); -x_36 = lean_ctor_get(x_34, 0); +x_35 = lean_st_ref_get(x_24, x_33); +x_36 = lean_ctor_get(x_35, 0); lean_inc(x_36); -lean_dec(x_34); -x_37 = lean_environment_main_module(x_36); -x_38 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__1; -x_39 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__2; -lean_inc(x_4); -x_40 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_4); -lean_ctor_set(x_40, 2, x_39); -x_41 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__5; -lean_inc(x_32); +x_37 = lean_ctor_get(x_35, 1); lean_inc(x_37); -x_42 = l_Lean_addMacroScope(x_37, x_41, x_32); -x_43 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__29; -lean_inc(x_2); -x_44 = lean_name_mk_string(x_2, x_43); -x_45 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__4; -x_46 = lean_name_mk_string(x_44, x_45); +lean_dec(x_35); +x_38 = lean_ctor_get(x_36, 0); +lean_inc(x_38); +lean_dec(x_36); +x_39 = lean_environment_main_module(x_38); +x_40 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__1; +x_41 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__2; lean_inc(x_5); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_5); +x_42 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_5); +lean_ctor_set(x_42, 2, x_41); +x_43 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__5; +lean_inc(x_34); +lean_inc(x_39); +x_44 = l_Lean_addMacroScope(x_39, x_43, x_34); +x_45 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__29; +lean_inc(x_3); +x_46 = lean_name_mk_string(x_3, x_45); +x_47 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__4; +x_48 = lean_name_mk_string(x_46, x_47); lean_inc(x_6); -lean_ctor_set(x_16, 1, x_6); -lean_ctor_set(x_16, 0, x_47); -lean_inc(x_30); -x_48 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_48, 0, x_30); -lean_ctor_set(x_48, 1, x_40); -lean_ctor_set(x_48, 2, x_42); -lean_ctor_set(x_48, 3, x_16); -x_49 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__9; -x_50 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__10; -lean_inc(x_4); -x_51 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_51, 0, x_49); -lean_ctor_set(x_51, 1, x_4); -lean_ctor_set(x_51, 2, x_50); +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_6); +lean_inc(x_7); +lean_ctor_set(x_18, 1, x_7); +lean_ctor_set(x_18, 0, x_49); lean_inc(x_32); -lean_inc(x_9); -lean_inc(x_37); -x_52 = l_Lean_addMacroScope(x_37, x_9, x_32); -lean_inc(x_6); -lean_inc(x_30); -x_53 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_53, 0, x_30); -lean_ctor_set(x_53, 1, x_51); +x_50 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_50, 0, x_32); +lean_ctor_set(x_50, 1, x_42); +lean_ctor_set(x_50, 2, x_44); +lean_ctor_set(x_50, 3, x_18); +x_51 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__9; +x_52 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__10; +lean_inc(x_5); +x_53 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_5); lean_ctor_set(x_53, 2, x_52); -lean_ctor_set(x_53, 3, x_6); -x_54 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; -lean_inc(x_30); -x_55 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_55, 0, x_30); -lean_ctor_set(x_55, 1, x_54); -lean_inc(x_12); -x_56 = l_Lean_addMacroScope(x_37, x_12, x_32); -lean_inc(x_6); -lean_inc(x_11); -lean_inc(x_30); -x_57 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_57, 0, x_30); -lean_ctor_set(x_57, 1, x_11); -lean_ctor_set(x_57, 2, x_56); -lean_ctor_set(x_57, 3, x_6); -x_58 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__7___closed__1; -lean_inc(x_30); -x_59 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_59, 0, x_30); -lean_ctor_set(x_59, 1, x_58); -x_60 = lean_nat_sub(x_1, x_27); -lean_dec(x_27); -x_61 = l_Nat_repr(x_60); -x_62 = l_Lean_numLitKind; -x_63 = lean_box(2); -x_64 = l_Lean_Syntax_mkLit(x_62, x_61, x_63); -lean_inc(x_13); -x_65 = lean_array_push(x_13, x_57); -x_66 = lean_array_push(x_65, x_59); -x_67 = lean_array_push(x_66, x_64); +lean_inc(x_34); lean_inc(x_10); -x_68 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_68, 0, x_63); -lean_ctor_set(x_68, 1, x_10); -lean_ctor_set(x_68, 2, x_67); -lean_inc(x_15); -x_69 = lean_array_push(x_15, x_68); +lean_inc(x_39); +x_54 = l_Lean_addMacroScope(x_39, x_10, x_34); +lean_inc(x_7); +lean_inc(x_32); +x_55 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_55, 0, x_32); +lean_ctor_set(x_55, 1, x_53); +lean_ctor_set(x_55, 2, x_54); +lean_ctor_set(x_55, 3, x_7); +x_56 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; +lean_inc(x_32); +x_57 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_57, 0, x_32); +lean_ctor_set(x_57, 1, x_56); lean_inc(x_14); -x_70 = lean_array_push(x_69, x_14); +x_58 = l_Lean_addMacroScope(x_39, x_14, x_34); lean_inc(x_7); -x_71 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_71, 0, x_63); -lean_ctor_set(x_71, 1, x_7); -lean_ctor_set(x_71, 2, x_70); -x_72 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; -x_73 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_73, 0, x_30); -lean_ctor_set(x_73, 1, x_72); lean_inc(x_13); -x_74 = lean_array_push(x_13, x_55); -x_75 = lean_array_push(x_74, x_71); -x_76 = lean_array_push(x_75, x_73); -lean_inc(x_8); -x_77 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_77, 0, x_63); -lean_ctor_set(x_77, 1, x_8); -lean_ctor_set(x_77, 2, x_76); +lean_inc(x_32); +x_59 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_59, 0, x_32); +lean_ctor_set(x_59, 1, x_13); +lean_ctor_set(x_59, 2, x_58); +lean_ctor_set(x_59, 3, x_7); +x_60 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__7___closed__1; +lean_inc(x_32); +x_61 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_61, 0, x_32); +lean_ctor_set(x_61, 1, x_60); +x_62 = lean_nat_sub(x_2, x_29); +lean_dec(x_29); +x_63 = l_Nat_repr(x_62); +x_64 = lean_box(2); +x_65 = l_Lean_Syntax_mkNumLit(x_63, x_64); lean_inc(x_15); -x_78 = lean_array_push(x_15, x_53); -x_79 = lean_array_push(x_78, x_77); -lean_inc(x_7); -x_80 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_80, 0, x_63); -lean_ctor_set(x_80, 1, x_7); -lean_ctor_set(x_80, 2, x_79); +x_66 = lean_array_push(x_15, x_59); +x_67 = lean_array_push(x_66, x_61); +x_68 = lean_array_push(x_67, x_65); +lean_inc(x_12); +x_69 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_69, 0, x_64); +lean_ctor_set(x_69, 1, x_12); +lean_ctor_set(x_69, 2, x_68); +lean_inc(x_17); +x_70 = lean_array_push(x_17, x_69); +lean_inc(x_16); +x_71 = lean_array_push(x_70, x_16); +lean_inc(x_8); +x_72 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_72, 0, x_64); +lean_ctor_set(x_72, 1, x_8); +lean_ctor_set(x_72, 2, x_71); +x_73 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; +x_74 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_74, 0, x_32); +lean_ctor_set(x_74, 1, x_73); lean_inc(x_15); -x_81 = lean_array_push(x_15, x_48); -x_82 = lean_array_push(x_81, x_80); -lean_inc(x_3); -x_83 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_83, 0, x_63); -lean_ctor_set(x_83, 1, x_3); -lean_ctor_set(x_83, 2, x_82); -x_84 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_28, x_17, x_18, x_19, x_20, x_21, x_22, x_35); -x_85 = !lean_is_exclusive(x_84); -if (x_85 == 0) +x_75 = lean_array_push(x_15, x_57); +x_76 = lean_array_push(x_75, x_72); +x_77 = lean_array_push(x_76, x_74); +lean_inc(x_9); +x_78 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_78, 0, x_64); +lean_ctor_set(x_78, 1, x_9); +lean_ctor_set(x_78, 2, x_77); +lean_inc(x_17); +x_79 = lean_array_push(x_17, x_55); +x_80 = lean_array_push(x_79, x_78); +lean_inc(x_8); +x_81 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_81, 0, x_64); +lean_ctor_set(x_81, 1, x_8); +lean_ctor_set(x_81, 2, x_80); +lean_inc(x_17); +x_82 = lean_array_push(x_17, x_50); +x_83 = lean_array_push(x_82, x_81); +lean_inc(x_4); +x_84 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_84, 0, x_64); +lean_ctor_set(x_84, 1, x_4); +lean_ctor_set(x_84, 2, x_83); +x_85 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_30, x_19, x_20, x_21, x_22, x_23, x_24, x_37); +x_86 = !lean_is_exclusive(x_85); +if (x_86 == 0) { -lean_object* x_86; lean_object* x_87; -x_86 = lean_ctor_get(x_84, 0); -x_87 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_87, 0, x_83); -lean_ctor_set(x_87, 1, x_86); -lean_ctor_set(x_84, 0, x_87); -return x_84; +lean_object* x_87; lean_object* x_88; +x_87 = lean_ctor_get(x_85, 0); +x_88 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_88, 0, x_84); +lean_ctor_set(x_88, 1, x_87); +lean_ctor_set(x_85, 0, x_88); +return x_85; } else { -lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; -x_88 = lean_ctor_get(x_84, 0); -x_89 = lean_ctor_get(x_84, 1); +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_89 = lean_ctor_get(x_85, 0); +x_90 = lean_ctor_get(x_85, 1); +lean_inc(x_90); lean_inc(x_89); -lean_inc(x_88); -lean_dec(x_84); -x_90 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_90, 0, x_83); -lean_ctor_set(x_90, 1, x_88); -x_91 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_91, 0, x_90); +lean_dec(x_85); +x_91 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_91, 0, x_84); lean_ctor_set(x_91, 1, x_89); -return x_91; +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_91); +lean_ctor_set(x_92, 1, x_90); +return x_92; } } else { -lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; -x_92 = lean_ctor_get(x_16, 0); -x_93 = lean_ctor_get(x_16, 1); +lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; +x_93 = lean_ctor_get(x_18, 0); +x_94 = lean_ctor_get(x_18, 1); +lean_inc(x_94); lean_inc(x_93); -lean_inc(x_92); -lean_dec(x_16); -lean_inc(x_21); -x_94 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_21, x_22, x_23); -x_95 = lean_ctor_get(x_94, 0); -lean_inc(x_95); -x_96 = lean_ctor_get(x_94, 1); +lean_dec(x_18); +lean_inc(x_23); +x_95 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_23, x_24, x_25); +x_96 = lean_ctor_get(x_95, 0); lean_inc(x_96); -lean_dec(x_94); -x_97 = lean_ctor_get(x_21, 10); +x_97 = lean_ctor_get(x_95, 1); lean_inc(x_97); -x_98 = lean_st_ref_get(x_22, x_96); -x_99 = lean_ctor_get(x_98, 0); -lean_inc(x_99); -x_100 = lean_ctor_get(x_98, 1); +lean_dec(x_95); +x_98 = lean_ctor_get(x_23, 10); +lean_inc(x_98); +x_99 = lean_st_ref_get(x_24, x_97); +x_100 = lean_ctor_get(x_99, 0); lean_inc(x_100); -lean_dec(x_98); -x_101 = lean_ctor_get(x_99, 0); +x_101 = lean_ctor_get(x_99, 1); lean_inc(x_101); lean_dec(x_99); -x_102 = lean_environment_main_module(x_101); -x_103 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__1; -x_104 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__2; -lean_inc(x_4); -x_105 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_105, 0, x_103); -lean_ctor_set(x_105, 1, x_4); -lean_ctor_set(x_105, 2, x_104); -x_106 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__5; -lean_inc(x_97); +x_102 = lean_ctor_get(x_100, 0); lean_inc(x_102); -x_107 = l_Lean_addMacroScope(x_102, x_106, x_97); -x_108 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__29; -lean_inc(x_2); -x_109 = lean_name_mk_string(x_2, x_108); -x_110 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__4; -x_111 = lean_name_mk_string(x_109, x_110); +lean_dec(x_100); +x_103 = lean_environment_main_module(x_102); +x_104 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__1; +x_105 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__2; lean_inc(x_5); -x_112 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_112, 0, x_111); -lean_ctor_set(x_112, 1, x_5); +x_106 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_106, 0, x_104); +lean_ctor_set(x_106, 1, x_5); +lean_ctor_set(x_106, 2, x_105); +x_107 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__5; +lean_inc(x_98); +lean_inc(x_103); +x_108 = l_Lean_addMacroScope(x_103, x_107, x_98); +x_109 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__29; +lean_inc(x_3); +x_110 = lean_name_mk_string(x_3, x_109); +x_111 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__4; +x_112 = lean_name_mk_string(x_110, x_111); lean_inc(x_6); -x_113 = lean_alloc_ctor(1, 2, 0); +x_113 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_113, 0, x_112); lean_ctor_set(x_113, 1, x_6); -lean_inc(x_95); -x_114 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_114, 0, x_95); -lean_ctor_set(x_114, 1, x_105); -lean_ctor_set(x_114, 2, x_107); -lean_ctor_set(x_114, 3, x_113); -x_115 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__9; -x_116 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__10; -lean_inc(x_4); -x_117 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_117, 0, x_115); -lean_ctor_set(x_117, 1, x_4); -lean_ctor_set(x_117, 2, x_116); -lean_inc(x_97); -lean_inc(x_9); -lean_inc(x_102); -x_118 = l_Lean_addMacroScope(x_102, x_9, x_97); -lean_inc(x_6); -lean_inc(x_95); -x_119 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_119, 0, x_95); -lean_ctor_set(x_119, 1, x_117); -lean_ctor_set(x_119, 2, x_118); -lean_ctor_set(x_119, 3, x_6); -x_120 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; -lean_inc(x_95); -x_121 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_121, 0, x_95); -lean_ctor_set(x_121, 1, x_120); -lean_inc(x_12); -x_122 = l_Lean_addMacroScope(x_102, x_12, x_97); -lean_inc(x_6); -lean_inc(x_11); -lean_inc(x_95); -x_123 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_123, 0, x_95); -lean_ctor_set(x_123, 1, x_11); -lean_ctor_set(x_123, 2, x_122); -lean_ctor_set(x_123, 3, x_6); -x_124 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__7___closed__1; -lean_inc(x_95); -x_125 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_125, 0, x_95); -lean_ctor_set(x_125, 1, x_124); -x_126 = lean_nat_sub(x_1, x_92); -lean_dec(x_92); -x_127 = l_Nat_repr(x_126); -x_128 = l_Lean_numLitKind; -x_129 = lean_box(2); -x_130 = l_Lean_Syntax_mkLit(x_128, x_127, x_129); +lean_inc(x_7); +x_114 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_114, 0, x_113); +lean_ctor_set(x_114, 1, x_7); +lean_inc(x_96); +x_115 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_115, 0, x_96); +lean_ctor_set(x_115, 1, x_106); +lean_ctor_set(x_115, 2, x_108); +lean_ctor_set(x_115, 3, x_114); +x_116 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__9; +x_117 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__10; +lean_inc(x_5); +x_118 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_118, 0, x_116); +lean_ctor_set(x_118, 1, x_5); +lean_ctor_set(x_118, 2, x_117); +lean_inc(x_98); +lean_inc(x_10); +lean_inc(x_103); +x_119 = l_Lean_addMacroScope(x_103, x_10, x_98); +lean_inc(x_7); +lean_inc(x_96); +x_120 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_120, 0, x_96); +lean_ctor_set(x_120, 1, x_118); +lean_ctor_set(x_120, 2, x_119); +lean_ctor_set(x_120, 3, x_7); +x_121 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; +lean_inc(x_96); +x_122 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_122, 0, x_96); +lean_ctor_set(x_122, 1, x_121); +lean_inc(x_14); +x_123 = l_Lean_addMacroScope(x_103, x_14, x_98); +lean_inc(x_7); lean_inc(x_13); -x_131 = lean_array_push(x_13, x_123); -x_132 = lean_array_push(x_131, x_125); +lean_inc(x_96); +x_124 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_124, 0, x_96); +lean_ctor_set(x_124, 1, x_13); +lean_ctor_set(x_124, 2, x_123); +lean_ctor_set(x_124, 3, x_7); +x_125 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__7___closed__1; +lean_inc(x_96); +x_126 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_126, 0, x_96); +lean_ctor_set(x_126, 1, x_125); +x_127 = lean_nat_sub(x_2, x_93); +lean_dec(x_93); +x_128 = l_Nat_repr(x_127); +x_129 = lean_box(2); +x_130 = l_Lean_Syntax_mkNumLit(x_128, x_129); +lean_inc(x_15); +x_131 = lean_array_push(x_15, x_124); +x_132 = lean_array_push(x_131, x_126); x_133 = lean_array_push(x_132, x_130); -lean_inc(x_10); +lean_inc(x_12); x_134 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_134, 0, x_129); -lean_ctor_set(x_134, 1, x_10); +lean_ctor_set(x_134, 1, x_12); lean_ctor_set(x_134, 2, x_133); -lean_inc(x_15); -x_135 = lean_array_push(x_15, x_134); -lean_inc(x_14); -x_136 = lean_array_push(x_135, x_14); -lean_inc(x_7); +lean_inc(x_17); +x_135 = lean_array_push(x_17, x_134); +lean_inc(x_16); +x_136 = lean_array_push(x_135, x_16); +lean_inc(x_8); x_137 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_137, 0, x_129); -lean_ctor_set(x_137, 1, x_7); +lean_ctor_set(x_137, 1, x_8); lean_ctor_set(x_137, 2, x_136); -x_138 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; +x_138 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; x_139 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_139, 0, x_95); +lean_ctor_set(x_139, 0, x_96); lean_ctor_set(x_139, 1, x_138); -lean_inc(x_13); -x_140 = lean_array_push(x_13, x_121); +lean_inc(x_15); +x_140 = lean_array_push(x_15, x_122); x_141 = lean_array_push(x_140, x_137); x_142 = lean_array_push(x_141, x_139); -lean_inc(x_8); +lean_inc(x_9); x_143 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_143, 0, x_129); -lean_ctor_set(x_143, 1, x_8); +lean_ctor_set(x_143, 1, x_9); lean_ctor_set(x_143, 2, x_142); -lean_inc(x_15); -x_144 = lean_array_push(x_15, x_119); +lean_inc(x_17); +x_144 = lean_array_push(x_17, x_120); x_145 = lean_array_push(x_144, x_143); -lean_inc(x_7); +lean_inc(x_8); x_146 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_146, 0, x_129); -lean_ctor_set(x_146, 1, x_7); +lean_ctor_set(x_146, 1, x_8); lean_ctor_set(x_146, 2, x_145); -lean_inc(x_15); -x_147 = lean_array_push(x_15, x_114); +lean_inc(x_17); +x_147 = lean_array_push(x_17, x_115); x_148 = lean_array_push(x_147, x_146); -lean_inc(x_3); +lean_inc(x_4); x_149 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_149, 0, x_129); -lean_ctor_set(x_149, 1, x_3); +lean_ctor_set(x_149, 1, x_4); lean_ctor_set(x_149, 2, x_148); -x_150 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_93, x_17, x_18, x_19, x_20, x_21, x_22, x_100); +x_150 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_94, x_19, x_20, x_21, x_22, x_23, x_24, x_101); x_151 = lean_ctor_get(x_150, 0); lean_inc(x_151); x_152 = lean_ctor_get(x_150, 1); @@ -25040,20 +26499,21 @@ static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quot { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__4; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__4; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8(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_object* x_9, size_t x_10, size_t 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) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8(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_object* x_9, lean_object* x_10, size_t x_11, size_t 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, lean_object* x_20) { _start: { -uint8_t x_20; -x_20 = lean_usize_dec_lt(x_11, x_10); -if (x_20 == 0) +uint8_t x_21; +x_21 = lean_usize_dec_lt(x_12, x_11); +if (x_21 == 0) { -lean_object* x_21; -lean_dec(x_17); +lean_object* x_22; +lean_dec(x_18); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -25061,207 +26521,207 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_12); -lean_ctor_set(x_21, 1, x_19); -return x_21; +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_13); +lean_ctor_set(x_22, 1, x_20); +return x_22; } else { -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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; size_t x_102; size_t x_103; -x_22 = lean_array_uget(x_9, x_11); -lean_inc(x_17); -x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_17, x_18, x_19); -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); +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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; size_t x_103; size_t x_104; +x_23 = lean_array_uget(x_10, x_12); +lean_inc(x_18); +x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_18, x_19, x_20); +x_25 = lean_ctor_get(x_24, 0); lean_inc(x_25); -lean_dec(x_23); -x_26 = lean_ctor_get(x_17, 10); +x_26 = lean_ctor_get(x_24, 1); lean_inc(x_26); -x_27 = lean_st_ref_get(x_18, x_25); -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); +lean_dec(x_24); +x_27 = lean_ctor_get(x_18, 10); +lean_inc(x_27); +x_28 = lean_st_ref_get(x_19, x_26); +x_29 = lean_ctor_get(x_28, 0); lean_inc(x_29); -lean_dec(x_27); -x_30 = lean_ctor_get(x_28, 0); +x_30 = lean_ctor_get(x_28, 1); lean_inc(x_30); lean_dec(x_28); -x_31 = lean_environment_main_module(x_30); -x_32 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__7; -lean_inc(x_1); -x_33 = lean_name_mk_string(x_1, x_32); -lean_inc(x_24); -x_34 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_34, 0, x_24); -lean_ctor_set(x_34, 1, x_32); -x_35 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__9; -lean_inc(x_1); -x_36 = lean_name_mk_string(x_1, x_35); -x_37 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; -lean_inc(x_1); -x_38 = lean_name_mk_string(x_1, x_37); -x_39 = lean_box(2); -x_40 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; -lean_inc(x_5); -x_41 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_5); -lean_ctor_set(x_41, 2, x_40); -x_42 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__17; -lean_inc(x_24); -x_43 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_43, 0, x_24); -lean_ctor_set(x_43, 1, x_42); -x_44 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__1; -x_45 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__2; -lean_inc(x_3); -x_46 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_3); -lean_ctor_set(x_46, 2, x_45); -x_47 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__5; -x_48 = l_Lean_addMacroScope(x_31, x_47, x_26); -lean_inc(x_4); -lean_inc(x_24); -x_49 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_49, 0, x_24); -lean_ctor_set(x_49, 1, x_46); -lean_ctor_set(x_49, 2, x_48); -lean_ctor_set(x_49, 3, x_4); -x_50 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__5; -lean_inc(x_1); -x_51 = lean_name_mk_string(x_1, x_50); -x_52 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; -lean_inc(x_24); -x_53 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_53, 0, x_24); -lean_ctor_set(x_53, 1, x_52); -x_54 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__8; -lean_inc(x_1); -x_55 = lean_name_mk_string(x_1, x_54); -lean_inc(x_24); -x_56 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_56, 0, x_24); -lean_ctor_set(x_56, 1, x_54); -x_57 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__10; -lean_inc(x_1); -x_58 = lean_name_mk_string(x_1, x_57); -lean_inc(x_8); +x_31 = lean_ctor_get(x_29, 0); +lean_inc(x_31); +lean_dec(x_29); +x_32 = lean_environment_main_module(x_31); +x_33 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__7; +lean_inc(x_2); +x_34 = lean_name_mk_string(x_2, x_33); +lean_inc(x_25); +x_35 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_35, 0, x_25); +lean_ctor_set(x_35, 1, x_33); +x_36 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__9; +lean_inc(x_2); +x_37 = lean_name_mk_string(x_2, x_36); +x_38 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; +lean_inc(x_2); +x_39 = lean_name_mk_string(x_2, x_38); +x_40 = lean_box(2); +x_41 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; lean_inc(x_6); -x_59 = lean_array_push(x_6, x_8); +x_42 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_6); +lean_ctor_set(x_42, 2, x_41); +x_43 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__17; +lean_inc(x_25); +x_44 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_44, 0, x_25); +lean_ctor_set(x_44, 1, x_43); +x_45 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__1; +x_46 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__2; +lean_inc(x_4); +x_47 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_4); +lean_ctor_set(x_47, 2, x_46); +x_48 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__5; +x_49 = l_Lean_addMacroScope(x_32, x_48, x_27); lean_inc(x_5); -x_60 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_60, 0, x_39); -lean_ctor_set(x_60, 1, x_5); -lean_ctor_set(x_60, 2, x_59); -x_61 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__12; -lean_inc(x_24); -x_62 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_62, 0, x_24); -lean_ctor_set(x_62, 1, x_61); -x_63 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; -x_64 = lean_array_push(x_63, x_60); -x_65 = lean_array_push(x_64, x_62); -lean_inc(x_22); -x_66 = lean_array_push(x_65, x_22); -x_67 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_67, 0, x_39); -lean_ctor_set(x_67, 1, x_58); -lean_ctor_set(x_67, 2, x_66); -lean_inc(x_7); -x_68 = lean_array_push(x_7, x_56); -x_69 = lean_array_push(x_68, x_67); -x_70 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_70, 0, x_39); -lean_ctor_set(x_70, 1, x_55); -lean_ctor_set(x_70, 2, x_69); +lean_inc(x_25); +x_50 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_50, 0, x_25); +lean_ctor_set(x_50, 1, x_47); +lean_ctor_set(x_50, 2, x_49); +lean_ctor_set(x_50, 3, x_5); +x_51 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__5; +lean_inc(x_2); +x_52 = lean_name_mk_string(x_2, x_51); +x_53 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; +lean_inc(x_25); +x_54 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_54, 0, x_25); +lean_ctor_set(x_54, 1, x_53); +x_55 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__8; +lean_inc(x_2); +x_56 = lean_name_mk_string(x_2, x_55); +lean_inc(x_25); +x_57 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_57, 0, x_25); +lean_ctor_set(x_57, 1, x_55); +x_58 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__10; +lean_inc(x_2); +x_59 = lean_name_mk_string(x_2, x_58); +lean_inc(x_9); lean_inc(x_7); -x_71 = lean_array_push(x_7, x_70); -lean_inc(x_41); -x_72 = lean_array_push(x_71, x_41); -lean_inc(x_5); -x_73 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_73, 0, x_39); -lean_ctor_set(x_73, 1, x_5); -lean_ctor_set(x_73, 2, x_72); -x_74 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; -lean_inc(x_24); -x_75 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_75, 0, x_24); -lean_ctor_set(x_75, 1, x_74); -x_76 = lean_array_push(x_63, x_53); -x_77 = lean_array_push(x_76, x_73); -x_78 = lean_array_push(x_77, x_75); -x_79 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_79, 0, x_39); -lean_ctor_set(x_79, 1, x_51); -lean_ctor_set(x_79, 2, x_78); +x_60 = lean_array_push(x_7, x_9); lean_inc(x_6); -x_80 = lean_array_push(x_6, x_79); -lean_inc(x_5); -x_81 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_81, 0, x_39); -lean_ctor_set(x_81, 1, x_5); -lean_ctor_set(x_81, 2, x_80); +x_61 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_61, 0, x_40); +lean_ctor_set(x_61, 1, x_6); +lean_ctor_set(x_61, 2, x_60); +x_62 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__12; +lean_inc(x_25); +x_63 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_63, 0, x_25); +lean_ctor_set(x_63, 1, x_62); +x_64 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; +x_65 = lean_array_push(x_64, x_61); +x_66 = lean_array_push(x_65, x_63); +lean_inc(x_23); +x_67 = lean_array_push(x_66, x_23); +x_68 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_68, 0, x_40); +lean_ctor_set(x_68, 1, x_59); +lean_ctor_set(x_68, 2, x_67); +lean_inc(x_8); +x_69 = lean_array_push(x_8, x_57); +x_70 = lean_array_push(x_69, x_68); +x_71 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_71, 0, x_40); +lean_ctor_set(x_71, 1, x_56); +lean_ctor_set(x_71, 2, x_70); +lean_inc(x_8); +x_72 = lean_array_push(x_8, x_71); +lean_inc(x_42); +x_73 = lean_array_push(x_72, x_42); +lean_inc(x_6); +x_74 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_74, 0, x_40); +lean_ctor_set(x_74, 1, x_6); +lean_ctor_set(x_74, 2, x_73); +x_75 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; +lean_inc(x_25); +x_76 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_76, 0, x_25); +lean_ctor_set(x_76, 1, x_75); +x_77 = lean_array_push(x_64, x_54); +x_78 = lean_array_push(x_77, x_74); +x_79 = lean_array_push(x_78, x_76); +x_80 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_80, 0, x_40); +lean_ctor_set(x_80, 1, x_52); +lean_ctor_set(x_80, 2, x_79); lean_inc(x_7); -x_82 = lean_array_push(x_7, x_49); -x_83 = lean_array_push(x_82, x_81); -lean_inc(x_2); -x_84 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_84, 0, x_39); -lean_ctor_set(x_84, 1, x_2); -lean_ctor_set(x_84, 2, x_83); -x_85 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; -x_86 = lean_array_push(x_85, x_22); -lean_inc(x_41); -x_87 = lean_array_push(x_86, x_41); -x_88 = lean_array_push(x_87, x_41); -x_89 = lean_array_push(x_88, x_43); -x_90 = lean_array_push(x_89, x_84); -x_91 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_91, 0, x_39); -lean_ctor_set(x_91, 1, x_38); -lean_ctor_set(x_91, 2, x_90); +x_81 = lean_array_push(x_7, x_80); lean_inc(x_6); -x_92 = lean_array_push(x_6, x_91); -x_93 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_93, 0, x_39); -lean_ctor_set(x_93, 1, x_36); -lean_ctor_set(x_93, 2, x_92); -x_94 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__20; -x_95 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_95, 0, x_24); -lean_ctor_set(x_95, 1, x_94); -x_96 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; -x_97 = lean_array_push(x_96, x_34); -x_98 = lean_array_push(x_97, x_93); -x_99 = lean_array_push(x_98, x_95); -x_100 = lean_array_push(x_99, x_12); -x_101 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_101, 0, x_39); -lean_ctor_set(x_101, 1, x_33); -lean_ctor_set(x_101, 2, x_100); -x_102 = 1; -x_103 = lean_usize_add(x_11, x_102); -x_11 = x_103; -x_12 = x_101; -x_19 = x_29; +x_82 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_82, 0, x_40); +lean_ctor_set(x_82, 1, x_6); +lean_ctor_set(x_82, 2, x_81); +lean_inc(x_8); +x_83 = lean_array_push(x_8, x_50); +x_84 = lean_array_push(x_83, x_82); +lean_inc(x_3); +x_85 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_85, 0, x_40); +lean_ctor_set(x_85, 1, x_3); +lean_ctor_set(x_85, 2, x_84); +x_86 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; +x_87 = lean_array_push(x_86, x_23); +lean_inc(x_42); +x_88 = lean_array_push(x_87, x_42); +x_89 = lean_array_push(x_88, x_42); +x_90 = lean_array_push(x_89, x_44); +x_91 = lean_array_push(x_90, x_85); +x_92 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_92, 0, x_40); +lean_ctor_set(x_92, 1, x_39); +lean_ctor_set(x_92, 2, x_91); +lean_inc(x_7); +x_93 = lean_array_push(x_7, x_92); +x_94 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_94, 0, x_40); +lean_ctor_set(x_94, 1, x_37); +lean_ctor_set(x_94, 2, x_93); +x_95 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__20; +x_96 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_96, 0, x_25); +lean_ctor_set(x_96, 1, x_95); +x_97 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; +x_98 = lean_array_push(x_97, x_35); +x_99 = lean_array_push(x_98, x_94); +x_100 = lean_array_push(x_99, x_96); +x_101 = lean_array_push(x_100, x_13); +x_102 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_102, 0, x_40); +lean_ctor_set(x_102, 1, x_34); +lean_ctor_set(x_102, 2, x_101); +x_103 = 1; +x_104 = lean_usize_add(x_12, x_103); +x_12 = x_104; +x_13 = x_102; +x_20 = x_30; goto _start; } } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9(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_object* x_9, size_t x_10, size_t 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) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9(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_object* x_9, lean_object* x_10, size_t x_11, size_t 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, lean_object* x_20) { _start: { -uint8_t x_20; -x_20 = lean_usize_dec_lt(x_11, x_10); -if (x_20 == 0) +uint8_t x_21; +x_21 = lean_usize_dec_lt(x_12, x_11); +if (x_21 == 0) { -lean_object* x_21; -lean_dec(x_17); +lean_object* x_22; +lean_dec(x_18); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -25269,207 +26729,207 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_12); -lean_ctor_set(x_21, 1, x_19); -return x_21; +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_13); +lean_ctor_set(x_22, 1, x_20); +return x_22; } else { -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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; size_t x_102; size_t x_103; -x_22 = lean_array_uget(x_9, x_11); -lean_inc(x_17); -x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_17, x_18, x_19); -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); +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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; size_t x_103; size_t x_104; +x_23 = lean_array_uget(x_10, x_12); +lean_inc(x_18); +x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_18, x_19, x_20); +x_25 = lean_ctor_get(x_24, 0); lean_inc(x_25); -lean_dec(x_23); -x_26 = lean_ctor_get(x_17, 10); +x_26 = lean_ctor_get(x_24, 1); lean_inc(x_26); -x_27 = lean_st_ref_get(x_18, x_25); -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); +lean_dec(x_24); +x_27 = lean_ctor_get(x_18, 10); +lean_inc(x_27); +x_28 = lean_st_ref_get(x_19, x_26); +x_29 = lean_ctor_get(x_28, 0); lean_inc(x_29); -lean_dec(x_27); -x_30 = lean_ctor_get(x_28, 0); +x_30 = lean_ctor_get(x_28, 1); lean_inc(x_30); lean_dec(x_28); -x_31 = lean_environment_main_module(x_30); -x_32 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__7; -lean_inc(x_1); -x_33 = lean_name_mk_string(x_1, x_32); -lean_inc(x_24); -x_34 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_34, 0, x_24); -lean_ctor_set(x_34, 1, x_32); -x_35 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__9; -lean_inc(x_1); -x_36 = lean_name_mk_string(x_1, x_35); -x_37 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; -lean_inc(x_1); -x_38 = lean_name_mk_string(x_1, x_37); -x_39 = lean_box(2); -x_40 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; -lean_inc(x_5); -x_41 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_5); -lean_ctor_set(x_41, 2, x_40); -x_42 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__17; -lean_inc(x_24); -x_43 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_43, 0, x_24); -lean_ctor_set(x_43, 1, x_42); -x_44 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__1; -x_45 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__2; -lean_inc(x_3); -x_46 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_3); -lean_ctor_set(x_46, 2, x_45); -x_47 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__5; -x_48 = l_Lean_addMacroScope(x_31, x_47, x_26); -lean_inc(x_4); -lean_inc(x_24); -x_49 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_49, 0, x_24); -lean_ctor_set(x_49, 1, x_46); -lean_ctor_set(x_49, 2, x_48); -lean_ctor_set(x_49, 3, x_4); -x_50 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__5; -lean_inc(x_1); -x_51 = lean_name_mk_string(x_1, x_50); -x_52 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; -lean_inc(x_24); -x_53 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_53, 0, x_24); -lean_ctor_set(x_53, 1, x_52); -x_54 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__8; -lean_inc(x_1); -x_55 = lean_name_mk_string(x_1, x_54); -lean_inc(x_24); -x_56 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_56, 0, x_24); -lean_ctor_set(x_56, 1, x_54); -x_57 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__10; -lean_inc(x_1); -x_58 = lean_name_mk_string(x_1, x_57); -lean_inc(x_8); +x_31 = lean_ctor_get(x_29, 0); +lean_inc(x_31); +lean_dec(x_29); +x_32 = lean_environment_main_module(x_31); +x_33 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__7; +lean_inc(x_2); +x_34 = lean_name_mk_string(x_2, x_33); +lean_inc(x_25); +x_35 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_35, 0, x_25); +lean_ctor_set(x_35, 1, x_33); +x_36 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__9; +lean_inc(x_2); +x_37 = lean_name_mk_string(x_2, x_36); +x_38 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; +lean_inc(x_2); +x_39 = lean_name_mk_string(x_2, x_38); +x_40 = lean_box(2); +x_41 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; lean_inc(x_6); -x_59 = lean_array_push(x_6, x_8); +x_42 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_6); +lean_ctor_set(x_42, 2, x_41); +x_43 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__17; +lean_inc(x_25); +x_44 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_44, 0, x_25); +lean_ctor_set(x_44, 1, x_43); +x_45 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__1; +x_46 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__2; +lean_inc(x_4); +x_47 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_4); +lean_ctor_set(x_47, 2, x_46); +x_48 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__5; +x_49 = l_Lean_addMacroScope(x_32, x_48, x_27); lean_inc(x_5); -x_60 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_60, 0, x_39); -lean_ctor_set(x_60, 1, x_5); -lean_ctor_set(x_60, 2, x_59); -x_61 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__12; -lean_inc(x_24); -x_62 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_62, 0, x_24); -lean_ctor_set(x_62, 1, x_61); -x_63 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; -x_64 = lean_array_push(x_63, x_60); -x_65 = lean_array_push(x_64, x_62); -lean_inc(x_22); -x_66 = lean_array_push(x_65, x_22); -x_67 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_67, 0, x_39); -lean_ctor_set(x_67, 1, x_58); -lean_ctor_set(x_67, 2, x_66); -lean_inc(x_7); -x_68 = lean_array_push(x_7, x_56); -x_69 = lean_array_push(x_68, x_67); -x_70 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_70, 0, x_39); -lean_ctor_set(x_70, 1, x_55); -lean_ctor_set(x_70, 2, x_69); +lean_inc(x_25); +x_50 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_50, 0, x_25); +lean_ctor_set(x_50, 1, x_47); +lean_ctor_set(x_50, 2, x_49); +lean_ctor_set(x_50, 3, x_5); +x_51 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__5; +lean_inc(x_2); +x_52 = lean_name_mk_string(x_2, x_51); +x_53 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; +lean_inc(x_25); +x_54 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_54, 0, x_25); +lean_ctor_set(x_54, 1, x_53); +x_55 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__8; +lean_inc(x_2); +x_56 = lean_name_mk_string(x_2, x_55); +lean_inc(x_25); +x_57 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_57, 0, x_25); +lean_ctor_set(x_57, 1, x_55); +x_58 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__10; +lean_inc(x_2); +x_59 = lean_name_mk_string(x_2, x_58); +lean_inc(x_9); lean_inc(x_7); -x_71 = lean_array_push(x_7, x_70); -lean_inc(x_41); -x_72 = lean_array_push(x_71, x_41); -lean_inc(x_5); -x_73 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_73, 0, x_39); -lean_ctor_set(x_73, 1, x_5); -lean_ctor_set(x_73, 2, x_72); -x_74 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; -lean_inc(x_24); -x_75 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_75, 0, x_24); -lean_ctor_set(x_75, 1, x_74); -x_76 = lean_array_push(x_63, x_53); -x_77 = lean_array_push(x_76, x_73); -x_78 = lean_array_push(x_77, x_75); -x_79 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_79, 0, x_39); -lean_ctor_set(x_79, 1, x_51); -lean_ctor_set(x_79, 2, x_78); +x_60 = lean_array_push(x_7, x_9); lean_inc(x_6); -x_80 = lean_array_push(x_6, x_79); -lean_inc(x_5); -x_81 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_81, 0, x_39); -lean_ctor_set(x_81, 1, x_5); -lean_ctor_set(x_81, 2, x_80); +x_61 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_61, 0, x_40); +lean_ctor_set(x_61, 1, x_6); +lean_ctor_set(x_61, 2, x_60); +x_62 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__12; +lean_inc(x_25); +x_63 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_63, 0, x_25); +lean_ctor_set(x_63, 1, x_62); +x_64 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; +x_65 = lean_array_push(x_64, x_61); +x_66 = lean_array_push(x_65, x_63); +lean_inc(x_23); +x_67 = lean_array_push(x_66, x_23); +x_68 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_68, 0, x_40); +lean_ctor_set(x_68, 1, x_59); +lean_ctor_set(x_68, 2, x_67); +lean_inc(x_8); +x_69 = lean_array_push(x_8, x_57); +x_70 = lean_array_push(x_69, x_68); +x_71 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_71, 0, x_40); +lean_ctor_set(x_71, 1, x_56); +lean_ctor_set(x_71, 2, x_70); +lean_inc(x_8); +x_72 = lean_array_push(x_8, x_71); +lean_inc(x_42); +x_73 = lean_array_push(x_72, x_42); +lean_inc(x_6); +x_74 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_74, 0, x_40); +lean_ctor_set(x_74, 1, x_6); +lean_ctor_set(x_74, 2, x_73); +x_75 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; +lean_inc(x_25); +x_76 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_76, 0, x_25); +lean_ctor_set(x_76, 1, x_75); +x_77 = lean_array_push(x_64, x_54); +x_78 = lean_array_push(x_77, x_74); +x_79 = lean_array_push(x_78, x_76); +x_80 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_80, 0, x_40); +lean_ctor_set(x_80, 1, x_52); +lean_ctor_set(x_80, 2, x_79); lean_inc(x_7); -x_82 = lean_array_push(x_7, x_49); -x_83 = lean_array_push(x_82, x_81); -lean_inc(x_2); -x_84 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_84, 0, x_39); -lean_ctor_set(x_84, 1, x_2); -lean_ctor_set(x_84, 2, x_83); -x_85 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; -x_86 = lean_array_push(x_85, x_22); -lean_inc(x_41); -x_87 = lean_array_push(x_86, x_41); -x_88 = lean_array_push(x_87, x_41); -x_89 = lean_array_push(x_88, x_43); -x_90 = lean_array_push(x_89, x_84); -x_91 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_91, 0, x_39); -lean_ctor_set(x_91, 1, x_38); -lean_ctor_set(x_91, 2, x_90); +x_81 = lean_array_push(x_7, x_80); lean_inc(x_6); -x_92 = lean_array_push(x_6, x_91); -x_93 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_93, 0, x_39); -lean_ctor_set(x_93, 1, x_36); -lean_ctor_set(x_93, 2, x_92); -x_94 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__20; -x_95 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_95, 0, x_24); -lean_ctor_set(x_95, 1, x_94); -x_96 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; -x_97 = lean_array_push(x_96, x_34); -x_98 = lean_array_push(x_97, x_93); -x_99 = lean_array_push(x_98, x_95); -x_100 = lean_array_push(x_99, x_12); -x_101 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_101, 0, x_39); -lean_ctor_set(x_101, 1, x_33); -lean_ctor_set(x_101, 2, x_100); -x_102 = 1; -x_103 = lean_usize_add(x_11, x_102); -x_11 = x_103; -x_12 = x_101; -x_19 = x_29; +x_82 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_82, 0, x_40); +lean_ctor_set(x_82, 1, x_6); +lean_ctor_set(x_82, 2, x_81); +lean_inc(x_8); +x_83 = lean_array_push(x_8, x_50); +x_84 = lean_array_push(x_83, x_82); +lean_inc(x_3); +x_85 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_85, 0, x_40); +lean_ctor_set(x_85, 1, x_3); +lean_ctor_set(x_85, 2, x_84); +x_86 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; +x_87 = lean_array_push(x_86, x_23); +lean_inc(x_42); +x_88 = lean_array_push(x_87, x_42); +x_89 = lean_array_push(x_88, x_42); +x_90 = lean_array_push(x_89, x_44); +x_91 = lean_array_push(x_90, x_85); +x_92 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_92, 0, x_40); +lean_ctor_set(x_92, 1, x_39); +lean_ctor_set(x_92, 2, x_91); +lean_inc(x_7); +x_93 = lean_array_push(x_7, x_92); +x_94 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_94, 0, x_40); +lean_ctor_set(x_94, 1, x_37); +lean_ctor_set(x_94, 2, x_93); +x_95 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__20; +x_96 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_96, 0, x_25); +lean_ctor_set(x_96, 1, x_95); +x_97 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; +x_98 = lean_array_push(x_97, x_35); +x_99 = lean_array_push(x_98, x_94); +x_100 = lean_array_push(x_99, x_96); +x_101 = lean_array_push(x_100, x_13); +x_102 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_102, 0, x_40); +lean_ctor_set(x_102, 1, x_34); +lean_ctor_set(x_102, 2, x_101); +x_103 = 1; +x_104 = lean_usize_add(x_12, x_103); +x_12 = x_104; +x_13 = x_102; +x_20 = x_30; goto _start; } } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__10(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_object* x_9, size_t x_10, size_t 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) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__10(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_object* x_9, lean_object* x_10, size_t x_11, size_t 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, lean_object* x_20) { _start: { -uint8_t x_20; -x_20 = lean_usize_dec_lt(x_11, x_10); -if (x_20 == 0) +uint8_t x_21; +x_21 = lean_usize_dec_lt(x_12, x_11); +if (x_21 == 0) { -lean_object* x_21; -lean_dec(x_17); +lean_object* x_22; +lean_dec(x_18); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -25477,194 +26937,193 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_12); -lean_ctor_set(x_21, 1, x_19); -return x_21; +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_13); +lean_ctor_set(x_22, 1, x_20); +return x_22; } else { -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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; size_t x_102; size_t x_103; -x_22 = lean_array_uget(x_9, x_11); -lean_inc(x_17); -x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_17, x_18, x_19); -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); +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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; size_t x_103; size_t x_104; +x_23 = lean_array_uget(x_10, x_12); +lean_inc(x_18); +x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_18, x_19, x_20); +x_25 = lean_ctor_get(x_24, 0); lean_inc(x_25); -lean_dec(x_23); -x_26 = lean_ctor_get(x_17, 10); +x_26 = lean_ctor_get(x_24, 1); lean_inc(x_26); -x_27 = lean_st_ref_get(x_18, x_25); -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); +lean_dec(x_24); +x_27 = lean_ctor_get(x_18, 10); +lean_inc(x_27); +x_28 = lean_st_ref_get(x_19, x_26); +x_29 = lean_ctor_get(x_28, 0); lean_inc(x_29); -lean_dec(x_27); -x_30 = lean_ctor_get(x_28, 0); +x_30 = lean_ctor_get(x_28, 1); lean_inc(x_30); lean_dec(x_28); -x_31 = lean_environment_main_module(x_30); -x_32 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__7; -lean_inc(x_1); -x_33 = lean_name_mk_string(x_1, x_32); -lean_inc(x_24); -x_34 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_34, 0, x_24); -lean_ctor_set(x_34, 1, x_32); -x_35 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__9; -lean_inc(x_1); -x_36 = lean_name_mk_string(x_1, x_35); -x_37 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; -lean_inc(x_1); -x_38 = lean_name_mk_string(x_1, x_37); -x_39 = lean_box(2); -x_40 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; -lean_inc(x_5); -x_41 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_5); -lean_ctor_set(x_41, 2, x_40); -x_42 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__17; -lean_inc(x_24); -x_43 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_43, 0, x_24); -lean_ctor_set(x_43, 1, x_42); -x_44 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__1; -x_45 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__2; -lean_inc(x_3); -x_46 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_3); -lean_ctor_set(x_46, 2, x_45); -x_47 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__5; -x_48 = l_Lean_addMacroScope(x_31, x_47, x_26); -lean_inc(x_4); -lean_inc(x_24); -x_49 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_49, 0, x_24); -lean_ctor_set(x_49, 1, x_46); -lean_ctor_set(x_49, 2, x_48); -lean_ctor_set(x_49, 3, x_4); -x_50 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__5; -lean_inc(x_1); -x_51 = lean_name_mk_string(x_1, x_50); -x_52 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; -lean_inc(x_24); -x_53 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_53, 0, x_24); -lean_ctor_set(x_53, 1, x_52); -x_54 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__8; -lean_inc(x_1); -x_55 = lean_name_mk_string(x_1, x_54); -lean_inc(x_24); -x_56 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_56, 0, x_24); -lean_ctor_set(x_56, 1, x_54); -x_57 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__10; -lean_inc(x_1); -x_58 = lean_name_mk_string(x_1, x_57); -lean_inc(x_8); +x_31 = lean_ctor_get(x_29, 0); +lean_inc(x_31); +lean_dec(x_29); +x_32 = lean_environment_main_module(x_31); +x_33 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__7; +lean_inc(x_2); +x_34 = lean_name_mk_string(x_2, x_33); +lean_inc(x_25); +x_35 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_35, 0, x_25); +lean_ctor_set(x_35, 1, x_33); +x_36 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__9; +lean_inc(x_2); +x_37 = lean_name_mk_string(x_2, x_36); +x_38 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; +lean_inc(x_2); +x_39 = lean_name_mk_string(x_2, x_38); +x_40 = lean_box(2); +x_41 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; lean_inc(x_6); -x_59 = lean_array_push(x_6, x_8); +x_42 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_6); +lean_ctor_set(x_42, 2, x_41); +x_43 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__17; +lean_inc(x_25); +x_44 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_44, 0, x_25); +lean_ctor_set(x_44, 1, x_43); +x_45 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__1; +x_46 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__2; +lean_inc(x_4); +x_47 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_4); +lean_ctor_set(x_47, 2, x_46); +x_48 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__5; +x_49 = l_Lean_addMacroScope(x_32, x_48, x_27); lean_inc(x_5); -x_60 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_60, 0, x_39); -lean_ctor_set(x_60, 1, x_5); -lean_ctor_set(x_60, 2, x_59); -x_61 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__12; -lean_inc(x_24); -x_62 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_62, 0, x_24); -lean_ctor_set(x_62, 1, x_61); -x_63 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; -x_64 = lean_array_push(x_63, x_60); -x_65 = lean_array_push(x_64, x_62); -lean_inc(x_22); -x_66 = lean_array_push(x_65, x_22); -x_67 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_67, 0, x_39); -lean_ctor_set(x_67, 1, x_58); -lean_ctor_set(x_67, 2, x_66); -lean_inc(x_7); -x_68 = lean_array_push(x_7, x_56); -x_69 = lean_array_push(x_68, x_67); -x_70 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_70, 0, x_39); -lean_ctor_set(x_70, 1, x_55); -lean_ctor_set(x_70, 2, x_69); +lean_inc(x_25); +x_50 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_50, 0, x_25); +lean_ctor_set(x_50, 1, x_47); +lean_ctor_set(x_50, 2, x_49); +lean_ctor_set(x_50, 3, x_5); +x_51 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__5; +lean_inc(x_2); +x_52 = lean_name_mk_string(x_2, x_51); +x_53 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; +lean_inc(x_25); +x_54 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_54, 0, x_25); +lean_ctor_set(x_54, 1, x_53); +x_55 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__8; +lean_inc(x_2); +x_56 = lean_name_mk_string(x_2, x_55); +lean_inc(x_25); +x_57 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_57, 0, x_25); +lean_ctor_set(x_57, 1, x_55); +x_58 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__10; +lean_inc(x_2); +x_59 = lean_name_mk_string(x_2, x_58); +lean_inc(x_9); lean_inc(x_7); -x_71 = lean_array_push(x_7, x_70); -lean_inc(x_41); -x_72 = lean_array_push(x_71, x_41); -lean_inc(x_5); -x_73 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_73, 0, x_39); -lean_ctor_set(x_73, 1, x_5); -lean_ctor_set(x_73, 2, x_72); -x_74 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; -lean_inc(x_24); -x_75 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_75, 0, x_24); -lean_ctor_set(x_75, 1, x_74); -x_76 = lean_array_push(x_63, x_53); -x_77 = lean_array_push(x_76, x_73); -x_78 = lean_array_push(x_77, x_75); -x_79 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_79, 0, x_39); -lean_ctor_set(x_79, 1, x_51); -lean_ctor_set(x_79, 2, x_78); +x_60 = lean_array_push(x_7, x_9); lean_inc(x_6); -x_80 = lean_array_push(x_6, x_79); -lean_inc(x_5); -x_81 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_81, 0, x_39); -lean_ctor_set(x_81, 1, x_5); -lean_ctor_set(x_81, 2, x_80); +x_61 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_61, 0, x_40); +lean_ctor_set(x_61, 1, x_6); +lean_ctor_set(x_61, 2, x_60); +x_62 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__12; +lean_inc(x_25); +x_63 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_63, 0, x_25); +lean_ctor_set(x_63, 1, x_62); +x_64 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; +x_65 = lean_array_push(x_64, x_61); +x_66 = lean_array_push(x_65, x_63); +lean_inc(x_23); +x_67 = lean_array_push(x_66, x_23); +x_68 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_68, 0, x_40); +lean_ctor_set(x_68, 1, x_59); +lean_ctor_set(x_68, 2, x_67); +lean_inc(x_8); +x_69 = lean_array_push(x_8, x_57); +x_70 = lean_array_push(x_69, x_68); +x_71 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_71, 0, x_40); +lean_ctor_set(x_71, 1, x_56); +lean_ctor_set(x_71, 2, x_70); +lean_inc(x_8); +x_72 = lean_array_push(x_8, x_71); +lean_inc(x_42); +x_73 = lean_array_push(x_72, x_42); +lean_inc(x_6); +x_74 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_74, 0, x_40); +lean_ctor_set(x_74, 1, x_6); +lean_ctor_set(x_74, 2, x_73); +x_75 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; +lean_inc(x_25); +x_76 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_76, 0, x_25); +lean_ctor_set(x_76, 1, x_75); +x_77 = lean_array_push(x_64, x_54); +x_78 = lean_array_push(x_77, x_74); +x_79 = lean_array_push(x_78, x_76); +x_80 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_80, 0, x_40); +lean_ctor_set(x_80, 1, x_52); +lean_ctor_set(x_80, 2, x_79); lean_inc(x_7); -x_82 = lean_array_push(x_7, x_49); -x_83 = lean_array_push(x_82, x_81); -lean_inc(x_2); -x_84 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_84, 0, x_39); -lean_ctor_set(x_84, 1, x_2); -lean_ctor_set(x_84, 2, x_83); -x_85 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; -x_86 = lean_array_push(x_85, x_22); -lean_inc(x_41); -x_87 = lean_array_push(x_86, x_41); -x_88 = lean_array_push(x_87, x_41); -x_89 = lean_array_push(x_88, x_43); -x_90 = lean_array_push(x_89, x_84); -x_91 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_91, 0, x_39); -lean_ctor_set(x_91, 1, x_38); -lean_ctor_set(x_91, 2, x_90); +x_81 = lean_array_push(x_7, x_80); lean_inc(x_6); -x_92 = lean_array_push(x_6, x_91); -x_93 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_93, 0, x_39); -lean_ctor_set(x_93, 1, x_36); -lean_ctor_set(x_93, 2, x_92); -x_94 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__20; -x_95 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_95, 0, x_24); -lean_ctor_set(x_95, 1, x_94); -x_96 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; -x_97 = lean_array_push(x_96, x_34); -x_98 = lean_array_push(x_97, x_93); -x_99 = lean_array_push(x_98, x_95); -x_100 = lean_array_push(x_99, x_12); -x_101 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_101, 0, x_39); -lean_ctor_set(x_101, 1, x_33); -lean_ctor_set(x_101, 2, x_100); -x_102 = 1; -x_103 = lean_usize_add(x_11, x_102); -x_11 = x_103; -x_12 = x_101; -x_19 = x_29; +x_82 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_82, 0, x_40); +lean_ctor_set(x_82, 1, x_6); +lean_ctor_set(x_82, 2, x_81); +lean_inc(x_8); +x_83 = lean_array_push(x_8, x_50); +x_84 = lean_array_push(x_83, x_82); +lean_inc(x_3); +x_85 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_85, 0, x_40); +lean_ctor_set(x_85, 1, x_3); +lean_ctor_set(x_85, 2, x_84); +x_86 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; +x_87 = lean_array_push(x_86, x_23); +lean_inc(x_42); +x_88 = lean_array_push(x_87, x_42); +x_89 = lean_array_push(x_88, x_42); +x_90 = lean_array_push(x_89, x_44); +x_91 = lean_array_push(x_90, x_85); +x_92 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_92, 0, x_40); +lean_ctor_set(x_92, 1, x_39); +lean_ctor_set(x_92, 2, x_91); +lean_inc(x_7); +x_93 = lean_array_push(x_7, x_92); +x_94 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_94, 0, x_40); +lean_ctor_set(x_94, 1, x_37); +lean_ctor_set(x_94, 2, x_93); +x_95 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__20; +x_96 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_96, 0, x_25); +lean_ctor_set(x_96, 1, x_95); +x_97 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; +x_98 = lean_array_push(x_97, x_35); +x_99 = lean_array_push(x_98, x_94); +x_100 = lean_array_push(x_99, x_96); +x_101 = lean_array_push(x_100, x_13); +x_102 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_102, 0, x_40); +lean_ctor_set(x_102, 1, x_34); +lean_ctor_set(x_102, 2, x_101); +x_103 = 1; +x_104 = lean_usize_add(x_12, x_103); +x_12 = x_104; +x_13 = x_102; +x_20 = x_30; goto _start; } } @@ -25699,14 +27158,14 @@ lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean x_20 = lean_array_uget(x_18, x_17); x_21 = lean_unsigned_to_nat(0u); x_22 = lean_array_uset(x_18, x_17, x_21); -x_23 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__1; -x_24 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__2; +x_23 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__1; +x_24 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__2; lean_inc(x_2); x_25 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_25, 0, x_23); lean_ctor_set(x_25, 1, x_2); lean_ctor_set(x_25, 2, x_24); -x_26 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__4; +x_26 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__4; lean_inc(x_1); lean_inc(x_7); x_27 = l_Lean_addMacroScope(x_7, x_26, x_1); @@ -25769,223 +27228,224 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__12(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_object* x_9, size_t x_10, size_t 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) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__12(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_object* x_9, lean_object* x_10, size_t x_11, size_t 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, lean_object* x_20) { _start: { -uint8_t x_20; -x_20 = lean_usize_dec_lt(x_11, x_10); -if (x_20 == 0) -{ -lean_object* x_21; -lean_dec(x_17); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_12); -lean_ctor_set(x_21, 1, x_19); -return x_21; -} -else +uint8_t x_21; +x_21 = lean_usize_dec_lt(x_12, x_11); +if (x_21 == 0) { -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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; size_t x_102; size_t x_103; -x_22 = lean_array_uget(x_9, x_11); -lean_inc(x_17); -x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_17, x_18, x_19); -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); -lean_inc(x_25); -lean_dec(x_23); -x_26 = lean_ctor_get(x_17, 10); -lean_inc(x_26); -x_27 = lean_st_ref_get(x_18, x_25); -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); -lean_dec(x_27); -x_30 = lean_ctor_get(x_28, 0); -lean_inc(x_30); -lean_dec(x_28); -x_31 = lean_environment_main_module(x_30); -x_32 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__7; -lean_inc(x_1); -x_33 = lean_name_mk_string(x_1, x_32); -lean_inc(x_24); -x_34 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_34, 0, x_24); -lean_ctor_set(x_34, 1, x_32); -x_35 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__9; -lean_inc(x_1); -x_36 = lean_name_mk_string(x_1, x_35); -x_37 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; -lean_inc(x_1); -x_38 = lean_name_mk_string(x_1, x_37); -x_39 = lean_box(2); -x_40 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; -lean_inc(x_5); -x_41 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_5); -lean_ctor_set(x_41, 2, x_40); -x_42 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__17; -lean_inc(x_24); -x_43 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_43, 0, x_24); -lean_ctor_set(x_43, 1, x_42); -x_44 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__1; -x_45 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__2; -lean_inc(x_3); -x_46 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_3); -lean_ctor_set(x_46, 2, x_45); -x_47 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__5; -x_48 = l_Lean_addMacroScope(x_31, x_47, x_26); -lean_inc(x_4); -lean_inc(x_24); -x_49 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_49, 0, x_24); -lean_ctor_set(x_49, 1, x_46); -lean_ctor_set(x_49, 2, x_48); -lean_ctor_set(x_49, 3, x_4); -x_50 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__5; -lean_inc(x_1); -x_51 = lean_name_mk_string(x_1, x_50); -x_52 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; -lean_inc(x_24); -x_53 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_53, 0, x_24); -lean_ctor_set(x_53, 1, x_52); -x_54 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__8; -lean_inc(x_1); -x_55 = lean_name_mk_string(x_1, x_54); -lean_inc(x_24); -x_56 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_56, 0, x_24); -lean_ctor_set(x_56, 1, x_54); -x_57 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__10; -lean_inc(x_1); -x_58 = lean_name_mk_string(x_1, x_57); -lean_inc(x_8); +lean_object* x_22; +lean_dec(x_18); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_13); +lean_ctor_set(x_22, 1, x_20); +return x_22; +} +else +{ +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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; size_t x_103; size_t x_104; +x_23 = lean_array_uget(x_10, x_12); +lean_inc(x_18); +x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_18, x_19, x_20); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); +x_27 = lean_ctor_get(x_18, 10); +lean_inc(x_27); +x_28 = lean_st_ref_get(x_19, x_26); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = lean_ctor_get(x_29, 0); +lean_inc(x_31); +lean_dec(x_29); +x_32 = lean_environment_main_module(x_31); +x_33 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__7; +lean_inc(x_2); +x_34 = lean_name_mk_string(x_2, x_33); +lean_inc(x_25); +x_35 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_35, 0, x_25); +lean_ctor_set(x_35, 1, x_33); +x_36 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__9; +lean_inc(x_2); +x_37 = lean_name_mk_string(x_2, x_36); +x_38 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; +lean_inc(x_2); +x_39 = lean_name_mk_string(x_2, x_38); +x_40 = lean_box(2); +x_41 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; lean_inc(x_6); -x_59 = lean_array_push(x_6, x_8); +x_42 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_6); +lean_ctor_set(x_42, 2, x_41); +x_43 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__17; +lean_inc(x_25); +x_44 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_44, 0, x_25); +lean_ctor_set(x_44, 1, x_43); +x_45 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__1; +x_46 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__2; +lean_inc(x_4); +x_47 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_4); +lean_ctor_set(x_47, 2, x_46); +x_48 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__5; +x_49 = l_Lean_addMacroScope(x_32, x_48, x_27); lean_inc(x_5); -x_60 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_60, 0, x_39); -lean_ctor_set(x_60, 1, x_5); -lean_ctor_set(x_60, 2, x_59); -x_61 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__12; -lean_inc(x_24); -x_62 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_62, 0, x_24); -lean_ctor_set(x_62, 1, x_61); -x_63 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; -x_64 = lean_array_push(x_63, x_60); -x_65 = lean_array_push(x_64, x_62); -lean_inc(x_22); -x_66 = lean_array_push(x_65, x_22); -x_67 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_67, 0, x_39); -lean_ctor_set(x_67, 1, x_58); -lean_ctor_set(x_67, 2, x_66); -lean_inc(x_7); -x_68 = lean_array_push(x_7, x_56); -x_69 = lean_array_push(x_68, x_67); -x_70 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_70, 0, x_39); -lean_ctor_set(x_70, 1, x_55); -lean_ctor_set(x_70, 2, x_69); +lean_inc(x_25); +x_50 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_50, 0, x_25); +lean_ctor_set(x_50, 1, x_47); +lean_ctor_set(x_50, 2, x_49); +lean_ctor_set(x_50, 3, x_5); +x_51 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__5; +lean_inc(x_2); +x_52 = lean_name_mk_string(x_2, x_51); +x_53 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; +lean_inc(x_25); +x_54 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_54, 0, x_25); +lean_ctor_set(x_54, 1, x_53); +x_55 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__8; +lean_inc(x_2); +x_56 = lean_name_mk_string(x_2, x_55); +lean_inc(x_25); +x_57 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_57, 0, x_25); +lean_ctor_set(x_57, 1, x_55); +x_58 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__10; +lean_inc(x_2); +x_59 = lean_name_mk_string(x_2, x_58); +lean_inc(x_9); lean_inc(x_7); -x_71 = lean_array_push(x_7, x_70); -lean_inc(x_41); -x_72 = lean_array_push(x_71, x_41); -lean_inc(x_5); -x_73 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_73, 0, x_39); -lean_ctor_set(x_73, 1, x_5); -lean_ctor_set(x_73, 2, x_72); -x_74 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; -lean_inc(x_24); -x_75 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_75, 0, x_24); -lean_ctor_set(x_75, 1, x_74); -x_76 = lean_array_push(x_63, x_53); -x_77 = lean_array_push(x_76, x_73); -x_78 = lean_array_push(x_77, x_75); -x_79 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_79, 0, x_39); -lean_ctor_set(x_79, 1, x_51); -lean_ctor_set(x_79, 2, x_78); +x_60 = lean_array_push(x_7, x_9); lean_inc(x_6); -x_80 = lean_array_push(x_6, x_79); -lean_inc(x_5); -x_81 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_81, 0, x_39); -lean_ctor_set(x_81, 1, x_5); -lean_ctor_set(x_81, 2, x_80); +x_61 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_61, 0, x_40); +lean_ctor_set(x_61, 1, x_6); +lean_ctor_set(x_61, 2, x_60); +x_62 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__12; +lean_inc(x_25); +x_63 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_63, 0, x_25); +lean_ctor_set(x_63, 1, x_62); +x_64 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; +x_65 = lean_array_push(x_64, x_61); +x_66 = lean_array_push(x_65, x_63); +lean_inc(x_23); +x_67 = lean_array_push(x_66, x_23); +x_68 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_68, 0, x_40); +lean_ctor_set(x_68, 1, x_59); +lean_ctor_set(x_68, 2, x_67); +lean_inc(x_8); +x_69 = lean_array_push(x_8, x_57); +x_70 = lean_array_push(x_69, x_68); +x_71 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_71, 0, x_40); +lean_ctor_set(x_71, 1, x_56); +lean_ctor_set(x_71, 2, x_70); +lean_inc(x_8); +x_72 = lean_array_push(x_8, x_71); +lean_inc(x_42); +x_73 = lean_array_push(x_72, x_42); +lean_inc(x_6); +x_74 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_74, 0, x_40); +lean_ctor_set(x_74, 1, x_6); +lean_ctor_set(x_74, 2, x_73); +x_75 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; +lean_inc(x_25); +x_76 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_76, 0, x_25); +lean_ctor_set(x_76, 1, x_75); +x_77 = lean_array_push(x_64, x_54); +x_78 = lean_array_push(x_77, x_74); +x_79 = lean_array_push(x_78, x_76); +x_80 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_80, 0, x_40); +lean_ctor_set(x_80, 1, x_52); +lean_ctor_set(x_80, 2, x_79); lean_inc(x_7); -x_82 = lean_array_push(x_7, x_49); -x_83 = lean_array_push(x_82, x_81); -lean_inc(x_2); -x_84 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_84, 0, x_39); -lean_ctor_set(x_84, 1, x_2); -lean_ctor_set(x_84, 2, x_83); -x_85 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; -x_86 = lean_array_push(x_85, x_22); -lean_inc(x_41); -x_87 = lean_array_push(x_86, x_41); -x_88 = lean_array_push(x_87, x_41); -x_89 = lean_array_push(x_88, x_43); -x_90 = lean_array_push(x_89, x_84); -x_91 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_91, 0, x_39); -lean_ctor_set(x_91, 1, x_38); -lean_ctor_set(x_91, 2, x_90); +x_81 = lean_array_push(x_7, x_80); lean_inc(x_6); -x_92 = lean_array_push(x_6, x_91); -x_93 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_93, 0, x_39); -lean_ctor_set(x_93, 1, x_36); -lean_ctor_set(x_93, 2, x_92); -x_94 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__20; -x_95 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_95, 0, x_24); -lean_ctor_set(x_95, 1, x_94); -x_96 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; -x_97 = lean_array_push(x_96, x_34); -x_98 = lean_array_push(x_97, x_93); -x_99 = lean_array_push(x_98, x_95); -x_100 = lean_array_push(x_99, x_12); -x_101 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_101, 0, x_39); -lean_ctor_set(x_101, 1, x_33); -lean_ctor_set(x_101, 2, x_100); -x_102 = 1; -x_103 = lean_usize_add(x_11, x_102); -x_11 = x_103; -x_12 = x_101; -x_19 = x_29; +x_82 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_82, 0, x_40); +lean_ctor_set(x_82, 1, x_6); +lean_ctor_set(x_82, 2, x_81); +lean_inc(x_8); +x_83 = lean_array_push(x_8, x_50); +x_84 = lean_array_push(x_83, x_82); +lean_inc(x_3); +x_85 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_85, 0, x_40); +lean_ctor_set(x_85, 1, x_3); +lean_ctor_set(x_85, 2, x_84); +x_86 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; +x_87 = lean_array_push(x_86, x_23); +lean_inc(x_42); +x_88 = lean_array_push(x_87, x_42); +x_89 = lean_array_push(x_88, x_42); +x_90 = lean_array_push(x_89, x_44); +x_91 = lean_array_push(x_90, x_85); +x_92 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_92, 0, x_40); +lean_ctor_set(x_92, 1, x_39); +lean_ctor_set(x_92, 2, x_91); +lean_inc(x_7); +x_93 = lean_array_push(x_7, x_92); +x_94 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_94, 0, x_40); +lean_ctor_set(x_94, 1, x_37); +lean_ctor_set(x_94, 2, x_93); +x_95 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__20; +x_96 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_96, 0, x_25); +lean_ctor_set(x_96, 1, x_95); +x_97 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; +x_98 = lean_array_push(x_97, x_35); +x_99 = lean_array_push(x_98, x_94); +x_100 = lean_array_push(x_99, x_96); +x_101 = lean_array_push(x_100, x_13); +x_102 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_102, 0, x_40); +lean_ctor_set(x_102, 1, x_34); +lean_ctor_set(x_102, 2, x_101); +x_103 = 1; +x_104 = lean_usize_add(x_12, x_103); +x_12 = x_104; +x_13 = x_102; +x_20 = x_30; goto _start; } } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__13(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_object* x_9, size_t x_10, size_t 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) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__13(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_object* x_9, lean_object* x_10, size_t x_11, size_t 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, lean_object* x_20) { _start: { -uint8_t x_20; -x_20 = lean_usize_dec_lt(x_11, x_10); -if (x_20 == 0) +uint8_t x_21; +x_21 = lean_usize_dec_lt(x_12, x_11); +if (x_21 == 0) { -lean_object* x_21; -lean_dec(x_17); +lean_object* x_22; +lean_dec(x_18); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -25993,207 +27453,207 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_12); -lean_ctor_set(x_21, 1, x_19); -return x_21; +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_13); +lean_ctor_set(x_22, 1, x_20); +return x_22; } else { -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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; size_t x_102; size_t x_103; -x_22 = lean_array_uget(x_9, x_11); -lean_inc(x_17); -x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_17, x_18, x_19); -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); +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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; size_t x_103; size_t x_104; +x_23 = lean_array_uget(x_10, x_12); +lean_inc(x_18); +x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_18, x_19, x_20); +x_25 = lean_ctor_get(x_24, 0); lean_inc(x_25); -lean_dec(x_23); -x_26 = lean_ctor_get(x_17, 10); +x_26 = lean_ctor_get(x_24, 1); lean_inc(x_26); -x_27 = lean_st_ref_get(x_18, x_25); -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); +lean_dec(x_24); +x_27 = lean_ctor_get(x_18, 10); +lean_inc(x_27); +x_28 = lean_st_ref_get(x_19, x_26); +x_29 = lean_ctor_get(x_28, 0); lean_inc(x_29); -lean_dec(x_27); -x_30 = lean_ctor_get(x_28, 0); +x_30 = lean_ctor_get(x_28, 1); lean_inc(x_30); lean_dec(x_28); -x_31 = lean_environment_main_module(x_30); -x_32 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__7; -lean_inc(x_1); -x_33 = lean_name_mk_string(x_1, x_32); -lean_inc(x_24); -x_34 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_34, 0, x_24); -lean_ctor_set(x_34, 1, x_32); -x_35 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__9; -lean_inc(x_1); -x_36 = lean_name_mk_string(x_1, x_35); -x_37 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; -lean_inc(x_1); -x_38 = lean_name_mk_string(x_1, x_37); -x_39 = lean_box(2); -x_40 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; -lean_inc(x_5); -x_41 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_5); -lean_ctor_set(x_41, 2, x_40); -x_42 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__17; -lean_inc(x_24); -x_43 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_43, 0, x_24); -lean_ctor_set(x_43, 1, x_42); -x_44 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__1; -x_45 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__2; -lean_inc(x_3); -x_46 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_3); -lean_ctor_set(x_46, 2, x_45); -x_47 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__5; -x_48 = l_Lean_addMacroScope(x_31, x_47, x_26); -lean_inc(x_4); -lean_inc(x_24); -x_49 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_49, 0, x_24); -lean_ctor_set(x_49, 1, x_46); -lean_ctor_set(x_49, 2, x_48); -lean_ctor_set(x_49, 3, x_4); -x_50 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__5; -lean_inc(x_1); -x_51 = lean_name_mk_string(x_1, x_50); -x_52 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; -lean_inc(x_24); -x_53 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_53, 0, x_24); -lean_ctor_set(x_53, 1, x_52); -x_54 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__8; -lean_inc(x_1); -x_55 = lean_name_mk_string(x_1, x_54); -lean_inc(x_24); -x_56 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_56, 0, x_24); -lean_ctor_set(x_56, 1, x_54); -x_57 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__10; -lean_inc(x_1); -x_58 = lean_name_mk_string(x_1, x_57); -lean_inc(x_8); +x_31 = lean_ctor_get(x_29, 0); +lean_inc(x_31); +lean_dec(x_29); +x_32 = lean_environment_main_module(x_31); +x_33 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__7; +lean_inc(x_2); +x_34 = lean_name_mk_string(x_2, x_33); +lean_inc(x_25); +x_35 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_35, 0, x_25); +lean_ctor_set(x_35, 1, x_33); +x_36 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__9; +lean_inc(x_2); +x_37 = lean_name_mk_string(x_2, x_36); +x_38 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; +lean_inc(x_2); +x_39 = lean_name_mk_string(x_2, x_38); +x_40 = lean_box(2); +x_41 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; lean_inc(x_6); -x_59 = lean_array_push(x_6, x_8); +x_42 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_6); +lean_ctor_set(x_42, 2, x_41); +x_43 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__17; +lean_inc(x_25); +x_44 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_44, 0, x_25); +lean_ctor_set(x_44, 1, x_43); +x_45 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__1; +x_46 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__2; +lean_inc(x_4); +x_47 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_4); +lean_ctor_set(x_47, 2, x_46); +x_48 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__5; +x_49 = l_Lean_addMacroScope(x_32, x_48, x_27); lean_inc(x_5); -x_60 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_60, 0, x_39); -lean_ctor_set(x_60, 1, x_5); -lean_ctor_set(x_60, 2, x_59); -x_61 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__12; -lean_inc(x_24); -x_62 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_62, 0, x_24); -lean_ctor_set(x_62, 1, x_61); -x_63 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; -x_64 = lean_array_push(x_63, x_60); -x_65 = lean_array_push(x_64, x_62); -lean_inc(x_22); -x_66 = lean_array_push(x_65, x_22); -x_67 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_67, 0, x_39); -lean_ctor_set(x_67, 1, x_58); -lean_ctor_set(x_67, 2, x_66); -lean_inc(x_7); -x_68 = lean_array_push(x_7, x_56); -x_69 = lean_array_push(x_68, x_67); -x_70 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_70, 0, x_39); -lean_ctor_set(x_70, 1, x_55); -lean_ctor_set(x_70, 2, x_69); +lean_inc(x_25); +x_50 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_50, 0, x_25); +lean_ctor_set(x_50, 1, x_47); +lean_ctor_set(x_50, 2, x_49); +lean_ctor_set(x_50, 3, x_5); +x_51 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__5; +lean_inc(x_2); +x_52 = lean_name_mk_string(x_2, x_51); +x_53 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; +lean_inc(x_25); +x_54 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_54, 0, x_25); +lean_ctor_set(x_54, 1, x_53); +x_55 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__8; +lean_inc(x_2); +x_56 = lean_name_mk_string(x_2, x_55); +lean_inc(x_25); +x_57 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_57, 0, x_25); +lean_ctor_set(x_57, 1, x_55); +x_58 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__10; +lean_inc(x_2); +x_59 = lean_name_mk_string(x_2, x_58); +lean_inc(x_9); lean_inc(x_7); -x_71 = lean_array_push(x_7, x_70); -lean_inc(x_41); -x_72 = lean_array_push(x_71, x_41); -lean_inc(x_5); -x_73 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_73, 0, x_39); -lean_ctor_set(x_73, 1, x_5); -lean_ctor_set(x_73, 2, x_72); -x_74 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; -lean_inc(x_24); -x_75 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_75, 0, x_24); -lean_ctor_set(x_75, 1, x_74); -x_76 = lean_array_push(x_63, x_53); -x_77 = lean_array_push(x_76, x_73); -x_78 = lean_array_push(x_77, x_75); -x_79 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_79, 0, x_39); -lean_ctor_set(x_79, 1, x_51); -lean_ctor_set(x_79, 2, x_78); +x_60 = lean_array_push(x_7, x_9); lean_inc(x_6); -x_80 = lean_array_push(x_6, x_79); -lean_inc(x_5); -x_81 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_81, 0, x_39); -lean_ctor_set(x_81, 1, x_5); -lean_ctor_set(x_81, 2, x_80); +x_61 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_61, 0, x_40); +lean_ctor_set(x_61, 1, x_6); +lean_ctor_set(x_61, 2, x_60); +x_62 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__12; +lean_inc(x_25); +x_63 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_63, 0, x_25); +lean_ctor_set(x_63, 1, x_62); +x_64 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; +x_65 = lean_array_push(x_64, x_61); +x_66 = lean_array_push(x_65, x_63); +lean_inc(x_23); +x_67 = lean_array_push(x_66, x_23); +x_68 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_68, 0, x_40); +lean_ctor_set(x_68, 1, x_59); +lean_ctor_set(x_68, 2, x_67); +lean_inc(x_8); +x_69 = lean_array_push(x_8, x_57); +x_70 = lean_array_push(x_69, x_68); +x_71 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_71, 0, x_40); +lean_ctor_set(x_71, 1, x_56); +lean_ctor_set(x_71, 2, x_70); +lean_inc(x_8); +x_72 = lean_array_push(x_8, x_71); +lean_inc(x_42); +x_73 = lean_array_push(x_72, x_42); +lean_inc(x_6); +x_74 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_74, 0, x_40); +lean_ctor_set(x_74, 1, x_6); +lean_ctor_set(x_74, 2, x_73); +x_75 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; +lean_inc(x_25); +x_76 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_76, 0, x_25); +lean_ctor_set(x_76, 1, x_75); +x_77 = lean_array_push(x_64, x_54); +x_78 = lean_array_push(x_77, x_74); +x_79 = lean_array_push(x_78, x_76); +x_80 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_80, 0, x_40); +lean_ctor_set(x_80, 1, x_52); +lean_ctor_set(x_80, 2, x_79); lean_inc(x_7); -x_82 = lean_array_push(x_7, x_49); -x_83 = lean_array_push(x_82, x_81); -lean_inc(x_2); -x_84 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_84, 0, x_39); -lean_ctor_set(x_84, 1, x_2); -lean_ctor_set(x_84, 2, x_83); -x_85 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; -x_86 = lean_array_push(x_85, x_22); -lean_inc(x_41); -x_87 = lean_array_push(x_86, x_41); -x_88 = lean_array_push(x_87, x_41); -x_89 = lean_array_push(x_88, x_43); -x_90 = lean_array_push(x_89, x_84); -x_91 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_91, 0, x_39); -lean_ctor_set(x_91, 1, x_38); -lean_ctor_set(x_91, 2, x_90); +x_81 = lean_array_push(x_7, x_80); lean_inc(x_6); -x_92 = lean_array_push(x_6, x_91); -x_93 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_93, 0, x_39); -lean_ctor_set(x_93, 1, x_36); -lean_ctor_set(x_93, 2, x_92); -x_94 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__20; -x_95 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_95, 0, x_24); -lean_ctor_set(x_95, 1, x_94); -x_96 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; -x_97 = lean_array_push(x_96, x_34); -x_98 = lean_array_push(x_97, x_93); -x_99 = lean_array_push(x_98, x_95); -x_100 = lean_array_push(x_99, x_12); -x_101 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_101, 0, x_39); -lean_ctor_set(x_101, 1, x_33); -lean_ctor_set(x_101, 2, x_100); -x_102 = 1; -x_103 = lean_usize_add(x_11, x_102); -x_11 = x_103; -x_12 = x_101; -x_19 = x_29; +x_82 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_82, 0, x_40); +lean_ctor_set(x_82, 1, x_6); +lean_ctor_set(x_82, 2, x_81); +lean_inc(x_8); +x_83 = lean_array_push(x_8, x_50); +x_84 = lean_array_push(x_83, x_82); +lean_inc(x_3); +x_85 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_85, 0, x_40); +lean_ctor_set(x_85, 1, x_3); +lean_ctor_set(x_85, 2, x_84); +x_86 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; +x_87 = lean_array_push(x_86, x_23); +lean_inc(x_42); +x_88 = lean_array_push(x_87, x_42); +x_89 = lean_array_push(x_88, x_42); +x_90 = lean_array_push(x_89, x_44); +x_91 = lean_array_push(x_90, x_85); +x_92 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_92, 0, x_40); +lean_ctor_set(x_92, 1, x_39); +lean_ctor_set(x_92, 2, x_91); +lean_inc(x_7); +x_93 = lean_array_push(x_7, x_92); +x_94 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_94, 0, x_40); +lean_ctor_set(x_94, 1, x_37); +lean_ctor_set(x_94, 2, x_93); +x_95 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__20; +x_96 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_96, 0, x_25); +lean_ctor_set(x_96, 1, x_95); +x_97 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; +x_98 = lean_array_push(x_97, x_35); +x_99 = lean_array_push(x_98, x_94); +x_100 = lean_array_push(x_99, x_96); +x_101 = lean_array_push(x_100, x_13); +x_102 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_102, 0, x_40); +lean_ctor_set(x_102, 1, x_34); +lean_ctor_set(x_102, 2, x_101); +x_103 = 1; +x_104 = lean_usize_add(x_12, x_103); +x_12 = x_104; +x_13 = x_102; +x_20 = x_30; goto _start; } } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__14(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_object* x_9, size_t x_10, size_t 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) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__14(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_object* x_9, lean_object* x_10, size_t x_11, size_t 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, lean_object* x_20) { _start: { -uint8_t x_20; -x_20 = lean_usize_dec_lt(x_11, x_10); -if (x_20 == 0) +uint8_t x_21; +x_21 = lean_usize_dec_lt(x_12, x_11); +if (x_21 == 0) { -lean_object* x_21; -lean_dec(x_17); +lean_object* x_22; +lean_dec(x_18); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -26201,194 +27661,193 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_12); -lean_ctor_set(x_21, 1, x_19); -return x_21; +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_13); +lean_ctor_set(x_22, 1, x_20); +return x_22; } else { -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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; size_t x_102; size_t x_103; -x_22 = lean_array_uget(x_9, x_11); -lean_inc(x_17); -x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_17, x_18, x_19); -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); +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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; size_t x_103; size_t x_104; +x_23 = lean_array_uget(x_10, x_12); +lean_inc(x_18); +x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_18, x_19, x_20); +x_25 = lean_ctor_get(x_24, 0); lean_inc(x_25); -lean_dec(x_23); -x_26 = lean_ctor_get(x_17, 10); +x_26 = lean_ctor_get(x_24, 1); lean_inc(x_26); -x_27 = lean_st_ref_get(x_18, x_25); -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); +lean_dec(x_24); +x_27 = lean_ctor_get(x_18, 10); +lean_inc(x_27); +x_28 = lean_st_ref_get(x_19, x_26); +x_29 = lean_ctor_get(x_28, 0); lean_inc(x_29); -lean_dec(x_27); -x_30 = lean_ctor_get(x_28, 0); +x_30 = lean_ctor_get(x_28, 1); lean_inc(x_30); lean_dec(x_28); -x_31 = lean_environment_main_module(x_30); -x_32 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__7; -lean_inc(x_1); -x_33 = lean_name_mk_string(x_1, x_32); -lean_inc(x_24); -x_34 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_34, 0, x_24); -lean_ctor_set(x_34, 1, x_32); -x_35 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__9; -lean_inc(x_1); -x_36 = lean_name_mk_string(x_1, x_35); -x_37 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; -lean_inc(x_1); -x_38 = lean_name_mk_string(x_1, x_37); -x_39 = lean_box(2); -x_40 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; -lean_inc(x_5); -x_41 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_5); -lean_ctor_set(x_41, 2, x_40); -x_42 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__17; -lean_inc(x_24); -x_43 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_43, 0, x_24); -lean_ctor_set(x_43, 1, x_42); -x_44 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__1; -x_45 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__2; -lean_inc(x_3); -x_46 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_3); -lean_ctor_set(x_46, 2, x_45); -x_47 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__5; -x_48 = l_Lean_addMacroScope(x_31, x_47, x_26); -lean_inc(x_4); -lean_inc(x_24); -x_49 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_49, 0, x_24); -lean_ctor_set(x_49, 1, x_46); -lean_ctor_set(x_49, 2, x_48); -lean_ctor_set(x_49, 3, x_4); -x_50 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__5; -lean_inc(x_1); -x_51 = lean_name_mk_string(x_1, x_50); -x_52 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; -lean_inc(x_24); -x_53 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_53, 0, x_24); -lean_ctor_set(x_53, 1, x_52); -x_54 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__8; -lean_inc(x_1); -x_55 = lean_name_mk_string(x_1, x_54); -lean_inc(x_24); -x_56 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_56, 0, x_24); -lean_ctor_set(x_56, 1, x_54); -x_57 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__10; -lean_inc(x_1); -x_58 = lean_name_mk_string(x_1, x_57); -lean_inc(x_8); +x_31 = lean_ctor_get(x_29, 0); +lean_inc(x_31); +lean_dec(x_29); +x_32 = lean_environment_main_module(x_31); +x_33 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__7; +lean_inc(x_2); +x_34 = lean_name_mk_string(x_2, x_33); +lean_inc(x_25); +x_35 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_35, 0, x_25); +lean_ctor_set(x_35, 1, x_33); +x_36 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__9; +lean_inc(x_2); +x_37 = lean_name_mk_string(x_2, x_36); +x_38 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; +lean_inc(x_2); +x_39 = lean_name_mk_string(x_2, x_38); +x_40 = lean_box(2); +x_41 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; lean_inc(x_6); -x_59 = lean_array_push(x_6, x_8); +x_42 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_6); +lean_ctor_set(x_42, 2, x_41); +x_43 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__17; +lean_inc(x_25); +x_44 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_44, 0, x_25); +lean_ctor_set(x_44, 1, x_43); +x_45 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__1; +x_46 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__2; +lean_inc(x_4); +x_47 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_4); +lean_ctor_set(x_47, 2, x_46); +x_48 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__5; +x_49 = l_Lean_addMacroScope(x_32, x_48, x_27); lean_inc(x_5); -x_60 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_60, 0, x_39); -lean_ctor_set(x_60, 1, x_5); -lean_ctor_set(x_60, 2, x_59); -x_61 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__12; -lean_inc(x_24); -x_62 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_62, 0, x_24); -lean_ctor_set(x_62, 1, x_61); -x_63 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; -x_64 = lean_array_push(x_63, x_60); -x_65 = lean_array_push(x_64, x_62); -lean_inc(x_22); -x_66 = lean_array_push(x_65, x_22); -x_67 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_67, 0, x_39); -lean_ctor_set(x_67, 1, x_58); -lean_ctor_set(x_67, 2, x_66); -lean_inc(x_7); -x_68 = lean_array_push(x_7, x_56); -x_69 = lean_array_push(x_68, x_67); -x_70 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_70, 0, x_39); -lean_ctor_set(x_70, 1, x_55); -lean_ctor_set(x_70, 2, x_69); +lean_inc(x_25); +x_50 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_50, 0, x_25); +lean_ctor_set(x_50, 1, x_47); +lean_ctor_set(x_50, 2, x_49); +lean_ctor_set(x_50, 3, x_5); +x_51 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__5; +lean_inc(x_2); +x_52 = lean_name_mk_string(x_2, x_51); +x_53 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; +lean_inc(x_25); +x_54 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_54, 0, x_25); +lean_ctor_set(x_54, 1, x_53); +x_55 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__8; +lean_inc(x_2); +x_56 = lean_name_mk_string(x_2, x_55); +lean_inc(x_25); +x_57 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_57, 0, x_25); +lean_ctor_set(x_57, 1, x_55); +x_58 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__10; +lean_inc(x_2); +x_59 = lean_name_mk_string(x_2, x_58); +lean_inc(x_9); lean_inc(x_7); -x_71 = lean_array_push(x_7, x_70); -lean_inc(x_41); -x_72 = lean_array_push(x_71, x_41); -lean_inc(x_5); -x_73 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_73, 0, x_39); -lean_ctor_set(x_73, 1, x_5); -lean_ctor_set(x_73, 2, x_72); -x_74 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; -lean_inc(x_24); -x_75 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_75, 0, x_24); -lean_ctor_set(x_75, 1, x_74); -x_76 = lean_array_push(x_63, x_53); -x_77 = lean_array_push(x_76, x_73); -x_78 = lean_array_push(x_77, x_75); -x_79 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_79, 0, x_39); -lean_ctor_set(x_79, 1, x_51); -lean_ctor_set(x_79, 2, x_78); +x_60 = lean_array_push(x_7, x_9); lean_inc(x_6); -x_80 = lean_array_push(x_6, x_79); -lean_inc(x_5); -x_81 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_81, 0, x_39); -lean_ctor_set(x_81, 1, x_5); -lean_ctor_set(x_81, 2, x_80); +x_61 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_61, 0, x_40); +lean_ctor_set(x_61, 1, x_6); +lean_ctor_set(x_61, 2, x_60); +x_62 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__12; +lean_inc(x_25); +x_63 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_63, 0, x_25); +lean_ctor_set(x_63, 1, x_62); +x_64 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; +x_65 = lean_array_push(x_64, x_61); +x_66 = lean_array_push(x_65, x_63); +lean_inc(x_23); +x_67 = lean_array_push(x_66, x_23); +x_68 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_68, 0, x_40); +lean_ctor_set(x_68, 1, x_59); +lean_ctor_set(x_68, 2, x_67); +lean_inc(x_8); +x_69 = lean_array_push(x_8, x_57); +x_70 = lean_array_push(x_69, x_68); +x_71 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_71, 0, x_40); +lean_ctor_set(x_71, 1, x_56); +lean_ctor_set(x_71, 2, x_70); +lean_inc(x_8); +x_72 = lean_array_push(x_8, x_71); +lean_inc(x_42); +x_73 = lean_array_push(x_72, x_42); +lean_inc(x_6); +x_74 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_74, 0, x_40); +lean_ctor_set(x_74, 1, x_6); +lean_ctor_set(x_74, 2, x_73); +x_75 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; +lean_inc(x_25); +x_76 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_76, 0, x_25); +lean_ctor_set(x_76, 1, x_75); +x_77 = lean_array_push(x_64, x_54); +x_78 = lean_array_push(x_77, x_74); +x_79 = lean_array_push(x_78, x_76); +x_80 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_80, 0, x_40); +lean_ctor_set(x_80, 1, x_52); +lean_ctor_set(x_80, 2, x_79); lean_inc(x_7); -x_82 = lean_array_push(x_7, x_49); -x_83 = lean_array_push(x_82, x_81); -lean_inc(x_2); -x_84 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_84, 0, x_39); -lean_ctor_set(x_84, 1, x_2); -lean_ctor_set(x_84, 2, x_83); -x_85 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; -x_86 = lean_array_push(x_85, x_22); -lean_inc(x_41); -x_87 = lean_array_push(x_86, x_41); -x_88 = lean_array_push(x_87, x_41); -x_89 = lean_array_push(x_88, x_43); -x_90 = lean_array_push(x_89, x_84); -x_91 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_91, 0, x_39); -lean_ctor_set(x_91, 1, x_38); -lean_ctor_set(x_91, 2, x_90); +x_81 = lean_array_push(x_7, x_80); lean_inc(x_6); -x_92 = lean_array_push(x_6, x_91); -x_93 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_93, 0, x_39); -lean_ctor_set(x_93, 1, x_36); -lean_ctor_set(x_93, 2, x_92); -x_94 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__20; -x_95 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_95, 0, x_24); -lean_ctor_set(x_95, 1, x_94); -x_96 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; -x_97 = lean_array_push(x_96, x_34); -x_98 = lean_array_push(x_97, x_93); -x_99 = lean_array_push(x_98, x_95); -x_100 = lean_array_push(x_99, x_12); -x_101 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_101, 0, x_39); -lean_ctor_set(x_101, 1, x_33); -lean_ctor_set(x_101, 2, x_100); -x_102 = 1; -x_103 = lean_usize_add(x_11, x_102); -x_11 = x_103; -x_12 = x_101; -x_19 = x_29; +x_82 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_82, 0, x_40); +lean_ctor_set(x_82, 1, x_6); +lean_ctor_set(x_82, 2, x_81); +lean_inc(x_8); +x_83 = lean_array_push(x_8, x_50); +x_84 = lean_array_push(x_83, x_82); +lean_inc(x_3); +x_85 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_85, 0, x_40); +lean_ctor_set(x_85, 1, x_3); +lean_ctor_set(x_85, 2, x_84); +x_86 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; +x_87 = lean_array_push(x_86, x_23); +lean_inc(x_42); +x_88 = lean_array_push(x_87, x_42); +x_89 = lean_array_push(x_88, x_42); +x_90 = lean_array_push(x_89, x_44); +x_91 = lean_array_push(x_90, x_85); +x_92 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_92, 0, x_40); +lean_ctor_set(x_92, 1, x_39); +lean_ctor_set(x_92, 2, x_91); +lean_inc(x_7); +x_93 = lean_array_push(x_7, x_92); +x_94 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_94, 0, x_40); +lean_ctor_set(x_94, 1, x_37); +lean_ctor_set(x_94, 2, x_93); +x_95 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__20; +x_96 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_96, 0, x_25); +lean_ctor_set(x_96, 1, x_95); +x_97 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; +x_98 = lean_array_push(x_97, x_35); +x_99 = lean_array_push(x_98, x_94); +x_100 = lean_array_push(x_99, x_96); +x_101 = lean_array_push(x_100, x_13); +x_102 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_102, 0, x_40); +lean_ctor_set(x_102, 1, x_34); +lean_ctor_set(x_102, 2, x_101); +x_103 = 1; +x_104 = lean_usize_add(x_12, x_103); +x_12 = x_104; +x_13 = x_102; +x_20 = x_30; goto _start; } } @@ -26455,274 +27914,273 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15(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_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) { _start: { -if (lean_obj_tag(x_7) == 0) +if (lean_obj_tag(x_8) == 0) { -lean_object* x_15; -lean_dec(x_12); +lean_object* x_16; +lean_dec(x_13); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_6); -lean_ctor_set(x_15, 1, x_14); -return x_15; +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_7); +lean_ctor_set(x_16, 1, x_15); +return x_16; } else { -uint8_t x_16; -x_16 = !lean_is_exclusive(x_7); -if (x_16 == 0) +uint8_t x_17; +x_17 = !lean_is_exclusive(x_8); +if (x_17 == 0) { -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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; lean_object* x_67; lean_object* x_68; -x_17 = lean_ctor_get(x_7, 0); -x_18 = lean_ctor_get(x_7, 1); -lean_inc(x_12); -x_19 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_12, x_13, x_14); -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_19, 1); +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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_18 = lean_ctor_get(x_8, 0); +x_19 = lean_ctor_get(x_8, 1); +lean_inc(x_13); +x_20 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_13, x_14, x_15); +x_21 = lean_ctor_get(x_20, 0); lean_inc(x_21); -lean_dec(x_19); -x_22 = lean_ctor_get(x_12, 10); +x_22 = lean_ctor_get(x_20, 1); lean_inc(x_22); -x_23 = lean_st_ref_get(x_13, x_21); -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); +lean_dec(x_20); +x_23 = lean_ctor_get(x_13, 10); +lean_inc(x_23); +x_24 = lean_st_ref_get(x_14, x_22); +x_25 = lean_ctor_get(x_24, 0); lean_inc(x_25); -lean_dec(x_23); -x_26 = lean_ctor_get(x_24, 0); +x_26 = lean_ctor_get(x_24, 1); lean_inc(x_26); lean_dec(x_24); -x_27 = lean_environment_main_module(x_26); -x_28 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__3; -lean_inc(x_2); -x_29 = lean_name_mk_string(x_2, x_28); -x_30 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__1; -x_31 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__2; -lean_inc(x_3); -x_32 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_3); -lean_ctor_set(x_32, 2, x_31); -x_33 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__3; -lean_inc(x_22); +x_27 = lean_ctor_get(x_25, 0); lean_inc(x_27); -x_34 = l_Lean_addMacroScope(x_27, x_33, x_22); +lean_dec(x_25); +x_28 = lean_environment_main_module(x_27); +x_29 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__3; +lean_inc(x_3); +x_30 = lean_name_mk_string(x_3, x_29); +x_31 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__1; +x_32 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__2; lean_inc(x_4); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_4); +x_33 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_4); +lean_ctor_set(x_33, 2, x_32); +x_34 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__3; +lean_inc(x_23); +lean_inc(x_28); +x_35 = l_Lean_addMacroScope(x_28, x_34, x_23); lean_inc(x_5); -lean_ctor_set(x_7, 1, x_5); -lean_ctor_set(x_7, 0, x_35); -lean_inc(x_20); -x_36 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_36, 0, x_20); -lean_ctor_set(x_36, 1, x_32); -lean_ctor_set(x_36, 2, x_34); -lean_ctor_set(x_36, 3, x_7); -x_37 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__5; -lean_inc(x_2); -x_38 = lean_name_mk_string(x_2, x_37); -x_39 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; -lean_inc(x_20); -x_40 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_40, 0, x_20); -lean_ctor_set(x_40, 1, x_39); -x_41 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__4; -x_42 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__5; +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_5); +lean_inc(x_6); +lean_ctor_set(x_8, 1, x_6); +lean_ctor_set(x_8, 0, x_36); +lean_inc(x_21); +x_37 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_37, 0, x_21); +lean_ctor_set(x_37, 1, x_33); +lean_ctor_set(x_37, 2, x_35); +lean_ctor_set(x_37, 3, x_8); +x_38 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__5; lean_inc(x_3); -x_43 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_43, 0, x_41); -lean_ctor_set(x_43, 1, x_3); -lean_ctor_set(x_43, 2, x_42); -x_44 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__7; -lean_inc(x_22); -lean_inc(x_27); -x_45 = l_Lean_addMacroScope(x_27, x_44, x_22); -x_46 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__29; -lean_inc(x_1); -x_47 = lean_name_mk_string(x_1, x_46); -x_48 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__6; -x_49 = lean_name_mk_string(x_47, x_48); +x_39 = lean_name_mk_string(x_3, x_38); +x_40 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; +lean_inc(x_21); +x_41 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_41, 0, x_21); +lean_ctor_set(x_41, 1, x_40); +x_42 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__4; +x_43 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__5; lean_inc(x_4); -x_50 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_50, 0, x_49); -lean_ctor_set(x_50, 1, x_4); +x_44 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_44, 0, x_42); +lean_ctor_set(x_44, 1, x_4); +lean_ctor_set(x_44, 2, x_43); +x_45 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__7; +lean_inc(x_23); +lean_inc(x_28); +x_46 = l_Lean_addMacroScope(x_28, x_45, x_23); +x_47 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__29; +lean_inc(x_2); +x_48 = lean_name_mk_string(x_2, x_47); +x_49 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__6; +x_50 = lean_name_mk_string(x_48, x_49); lean_inc(x_5); -x_51 = lean_alloc_ctor(1, 2, 0); +x_51 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_51, 0, x_50); lean_ctor_set(x_51, 1, x_5); -lean_inc(x_20); -x_52 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_52, 0, x_20); -lean_ctor_set(x_52, 1, x_43); -lean_ctor_set(x_52, 2, x_45); -lean_ctor_set(x_52, 3, x_51); -x_53 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__9; -x_54 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__10; -lean_inc(x_3); -x_55 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_55, 0, x_53); -lean_ctor_set(x_55, 1, x_3); -lean_ctor_set(x_55, 2, x_54); -x_56 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -x_57 = l_Lean_addMacroScope(x_27, x_56, x_22); -lean_inc(x_5); -lean_inc(x_20); -x_58 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_58, 0, x_20); -lean_ctor_set(x_58, 1, x_55); -lean_ctor_set(x_58, 2, x_57); -lean_ctor_set(x_58, 3, x_5); -lean_inc(x_17); +lean_inc(x_6); +x_52 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_6); +lean_inc(x_21); +x_53 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_53, 0, x_21); +lean_ctor_set(x_53, 1, x_44); +lean_ctor_set(x_53, 2, x_46); +lean_ctor_set(x_53, 3, x_52); +x_54 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__9; +x_55 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__10; lean_inc(x_4); -x_59 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_4, x_17); -x_60 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_61 = lean_array_push(x_60, x_58); -x_62 = lean_array_push(x_60, x_52); -x_63 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; -x_64 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_64, 0, x_20); -lean_ctor_set(x_64, 1, x_63); -x_65 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; -x_66 = lean_array_push(x_65, x_40); -x_67 = lean_array_push(x_60, x_6); -x_68 = lean_array_push(x_60, x_36); -if (lean_obj_tag(x_59) == 0) -{ -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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; -x_69 = l___private_Init_Meta_0__Lean_quoteNameMk(x_17); -x_70 = lean_array_push(x_61, x_69); -x_71 = lean_box(2); -x_72 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_73 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_73, 0, x_71); -lean_ctor_set(x_73, 1, x_72); -lean_ctor_set(x_73, 2, x_70); -x_74 = lean_array_push(x_62, x_73); -lean_inc(x_29); -x_75 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_75, 0, x_71); -lean_ctor_set(x_75, 1, x_29); -lean_ctor_set(x_75, 2, x_74); -x_76 = lean_array_push(x_60, x_75); -x_77 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; -x_78 = lean_array_push(x_76, x_77); -x_79 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_79, 0, x_71); -lean_ctor_set(x_79, 1, x_72); -lean_ctor_set(x_79, 2, x_78); -x_80 = lean_array_push(x_66, x_79); -x_81 = lean_array_push(x_80, x_64); -x_82 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_82, 0, x_71); -lean_ctor_set(x_82, 1, x_38); -lean_ctor_set(x_82, 2, x_81); -x_83 = lean_array_push(x_67, x_82); -x_84 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_84, 0, x_71); -lean_ctor_set(x_84, 1, x_72); -lean_ctor_set(x_84, 2, x_83); -x_85 = lean_array_push(x_68, x_84); -x_86 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_86, 0, x_71); -lean_ctor_set(x_86, 1, x_29); -lean_ctor_set(x_86, 2, x_85); -x_6 = x_86; -x_7 = x_18; -x_14 = x_25; +x_56 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_56, 0, x_54); +lean_ctor_set(x_56, 1, x_4); +lean_ctor_set(x_56, 2, x_55); +x_57 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +x_58 = l_Lean_addMacroScope(x_28, x_57, x_23); +lean_inc(x_6); +lean_inc(x_21); +x_59 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_59, 0, x_21); +lean_ctor_set(x_59, 1, x_56); +lean_ctor_set(x_59, 2, x_58); +lean_ctor_set(x_59, 3, x_6); +lean_inc(x_18); +lean_inc(x_5); +x_60 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_5, x_18); +x_61 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_62 = lean_array_push(x_61, x_59); +x_63 = lean_array_push(x_61, x_53); +x_64 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; +x_65 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_65, 0, x_21); +lean_ctor_set(x_65, 1, x_64); +x_66 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; +x_67 = lean_array_push(x_66, x_41); +x_68 = lean_array_push(x_61, x_7); +x_69 = lean_array_push(x_61, x_37); +if (lean_obj_tag(x_60) == 0) +{ +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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_70 = l_Lean_quoteNameMk(x_18); +x_71 = lean_array_push(x_62, x_70); +x_72 = lean_box(2); +x_73 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_74 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_74, 0, x_72); +lean_ctor_set(x_74, 1, x_73); +lean_ctor_set(x_74, 2, x_71); +x_75 = lean_array_push(x_63, x_74); +lean_inc(x_30); +x_76 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_76, 0, x_72); +lean_ctor_set(x_76, 1, x_30); +lean_ctor_set(x_76, 2, x_75); +x_77 = lean_array_push(x_61, x_76); +x_78 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; +x_79 = lean_array_push(x_77, x_78); +x_80 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_80, 0, x_72); +lean_ctor_set(x_80, 1, x_73); +lean_ctor_set(x_80, 2, x_79); +x_81 = lean_array_push(x_67, x_80); +x_82 = lean_array_push(x_81, x_65); +x_83 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_83, 0, x_72); +lean_ctor_set(x_83, 1, x_39); +lean_ctor_set(x_83, 2, x_82); +x_84 = lean_array_push(x_68, x_83); +x_85 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_85, 0, x_72); +lean_ctor_set(x_85, 1, x_73); +lean_ctor_set(x_85, 2, x_84); +x_86 = lean_array_push(x_69, x_85); +x_87 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_87, 0, x_72); +lean_ctor_set(x_87, 1, x_30); +lean_ctor_set(x_87, 2, x_86); +x_7 = x_87; +x_8 = x_19; +x_15 = x_26; goto _start; } else { -lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; -lean_dec(x_17); -x_88 = lean_ctor_get(x_59, 0); -lean_inc(x_88); -lean_dec(x_59); -x_89 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__1; -lean_inc(x_2); -x_90 = lean_name_mk_string(x_2, x_89); -x_91 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__3; -x_92 = l_String_intercalate(x_91, x_88); -x_93 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__4; -x_94 = lean_string_append(x_93, x_92); -lean_dec(x_92); -x_95 = l_Lean_nameLitKind; +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_dec(x_18); +x_89 = lean_ctor_get(x_60, 0); +lean_inc(x_89); +lean_dec(x_60); +x_90 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__1; +lean_inc(x_3); +x_91 = lean_name_mk_string(x_3, x_90); +x_92 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; +x_93 = l_String_intercalate(x_92, x_89); +x_94 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4; +x_95 = lean_string_append(x_94, x_93); +lean_dec(x_93); x_96 = lean_box(2); -x_97 = l_Lean_Syntax_mkLit(x_95, x_94, x_96); +x_97 = l_Lean_Syntax_mkNameLit(x_95, x_96); x_98 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; x_99 = lean_array_push(x_98, x_97); x_100 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_100, 0, x_96); -lean_ctor_set(x_100, 1, x_90); +lean_ctor_set(x_100, 1, x_91); lean_ctor_set(x_100, 2, x_99); -x_101 = lean_array_push(x_61, x_100); +x_101 = lean_array_push(x_62, x_100); x_102 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; x_103 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_103, 0, x_96); lean_ctor_set(x_103, 1, x_102); lean_ctor_set(x_103, 2, x_101); -x_104 = lean_array_push(x_62, x_103); -lean_inc(x_29); +x_104 = lean_array_push(x_63, x_103); +lean_inc(x_30); x_105 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_105, 0, x_96); -lean_ctor_set(x_105, 1, x_29); +lean_ctor_set(x_105, 1, x_30); lean_ctor_set(x_105, 2, x_104); -x_106 = lean_array_push(x_60, x_105); +x_106 = lean_array_push(x_61, x_105); x_107 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; x_108 = lean_array_push(x_106, x_107); x_109 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_109, 0, x_96); lean_ctor_set(x_109, 1, x_102); lean_ctor_set(x_109, 2, x_108); -x_110 = lean_array_push(x_66, x_109); -x_111 = lean_array_push(x_110, x_64); +x_110 = lean_array_push(x_67, x_109); +x_111 = lean_array_push(x_110, x_65); x_112 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_112, 0, x_96); -lean_ctor_set(x_112, 1, x_38); +lean_ctor_set(x_112, 1, x_39); lean_ctor_set(x_112, 2, x_111); -x_113 = lean_array_push(x_67, x_112); +x_113 = lean_array_push(x_68, x_112); x_114 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_114, 0, x_96); lean_ctor_set(x_114, 1, x_102); lean_ctor_set(x_114, 2, x_113); -x_115 = lean_array_push(x_68, x_114); +x_115 = lean_array_push(x_69, x_114); x_116 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_116, 0, x_96); -lean_ctor_set(x_116, 1, x_29); +lean_ctor_set(x_116, 1, x_30); lean_ctor_set(x_116, 2, x_115); -x_6 = x_116; -x_7 = x_18; -x_14 = x_25; +x_7 = x_116; +x_8 = x_19; +x_15 = x_26; goto _start; } } else { lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; -x_118 = lean_ctor_get(x_7, 0); -x_119 = lean_ctor_get(x_7, 1); +x_118 = lean_ctor_get(x_8, 0); +x_119 = lean_ctor_get(x_8, 1); lean_inc(x_119); lean_inc(x_118); -lean_dec(x_7); -lean_inc(x_12); -x_120 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_12, x_13, x_14); +lean_dec(x_8); +lean_inc(x_13); +x_120 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_13, x_14, x_15); x_121 = lean_ctor_get(x_120, 0); lean_inc(x_121); x_122 = lean_ctor_get(x_120, 1); lean_inc(x_122); lean_dec(x_120); -x_123 = lean_ctor_get(x_12, 10); +x_123 = lean_ctor_get(x_13, 10); lean_inc(x_123); -x_124 = lean_st_ref_get(x_13, x_122); +x_124 = lean_st_ref_get(x_14, x_122); x_125 = lean_ctor_get(x_124, 0); lean_inc(x_125); x_126 = lean_ctor_get(x_124, 1); @@ -26733,65 +28191,65 @@ lean_inc(x_127); lean_dec(x_125); x_128 = lean_environment_main_module(x_127); x_129 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__3; -lean_inc(x_2); -x_130 = lean_name_mk_string(x_2, x_129); +lean_inc(x_3); +x_130 = lean_name_mk_string(x_3, x_129); x_131 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__1; x_132 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__2; -lean_inc(x_3); +lean_inc(x_4); x_133 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_133, 0, x_131); -lean_ctor_set(x_133, 1, x_3); +lean_ctor_set(x_133, 1, x_4); lean_ctor_set(x_133, 2, x_132); x_134 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__3; lean_inc(x_123); lean_inc(x_128); x_135 = l_Lean_addMacroScope(x_128, x_134, x_123); -lean_inc(x_4); +lean_inc(x_5); x_136 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_136, 0, x_134); -lean_ctor_set(x_136, 1, x_4); -lean_inc(x_5); +lean_ctor_set(x_136, 1, x_5); +lean_inc(x_6); x_137 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_137, 0, x_136); -lean_ctor_set(x_137, 1, x_5); +lean_ctor_set(x_137, 1, x_6); lean_inc(x_121); x_138 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_138, 0, x_121); lean_ctor_set(x_138, 1, x_133); lean_ctor_set(x_138, 2, x_135); lean_ctor_set(x_138, 3, x_137); -x_139 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__5; -lean_inc(x_2); -x_140 = lean_name_mk_string(x_2, x_139); -x_141 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; +x_139 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__5; +lean_inc(x_3); +x_140 = lean_name_mk_string(x_3, x_139); +x_141 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; lean_inc(x_121); x_142 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_142, 0, x_121); lean_ctor_set(x_142, 1, x_141); x_143 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__4; x_144 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__5; -lean_inc(x_3); +lean_inc(x_4); x_145 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_145, 0, x_143); -lean_ctor_set(x_145, 1, x_3); +lean_ctor_set(x_145, 1, x_4); lean_ctor_set(x_145, 2, x_144); x_146 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__7; lean_inc(x_123); lean_inc(x_128); x_147 = l_Lean_addMacroScope(x_128, x_146, x_123); -x_148 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__29; -lean_inc(x_1); -x_149 = lean_name_mk_string(x_1, x_148); +x_148 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__29; +lean_inc(x_2); +x_149 = lean_name_mk_string(x_2, x_148); x_150 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__6; x_151 = lean_name_mk_string(x_149, x_150); -lean_inc(x_4); +lean_inc(x_5); x_152 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_152, 0, x_151); -lean_ctor_set(x_152, 1, x_4); -lean_inc(x_5); +lean_ctor_set(x_152, 1, x_5); +lean_inc(x_6); x_153 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_153, 0, x_152); -lean_ctor_set(x_153, 1, x_5); +lean_ctor_set(x_153, 1, x_6); lean_inc(x_121); x_154 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_154, 0, x_121); @@ -26800,38 +28258,38 @@ lean_ctor_set(x_154, 2, x_147); lean_ctor_set(x_154, 3, x_153); x_155 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__9; x_156 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__10; -lean_inc(x_3); +lean_inc(x_4); x_157 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_157, 0, x_155); -lean_ctor_set(x_157, 1, x_3); +lean_ctor_set(x_157, 1, x_4); lean_ctor_set(x_157, 2, x_156); x_158 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; x_159 = l_Lean_addMacroScope(x_128, x_158, x_123); -lean_inc(x_5); +lean_inc(x_6); lean_inc(x_121); x_160 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_160, 0, x_121); lean_ctor_set(x_160, 1, x_157); lean_ctor_set(x_160, 2, x_159); -lean_ctor_set(x_160, 3, x_5); +lean_ctor_set(x_160, 3, x_6); lean_inc(x_118); -lean_inc(x_4); -x_161 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_4, x_118); +lean_inc(x_5); +x_161 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_5, x_118); x_162 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; x_163 = lean_array_push(x_162, x_160); x_164 = lean_array_push(x_162, x_154); -x_165 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; +x_165 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; x_166 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_166, 0, x_121); lean_ctor_set(x_166, 1, x_165); -x_167 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; +x_167 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; x_168 = lean_array_push(x_167, x_142); -x_169 = lean_array_push(x_162, x_6); +x_169 = lean_array_push(x_162, x_7); x_170 = lean_array_push(x_162, x_138); if (lean_obj_tag(x_161) == 0) { lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; -x_171 = l___private_Init_Meta_0__Lean_quoteNameMk(x_118); +x_171 = l_Lean_quoteNameMk(x_118); x_172 = lean_array_push(x_163, x_171); x_173 = lean_box(2); x_174 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; @@ -26868,172 +28326,170 @@ x_188 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_188, 0, x_173); lean_ctor_set(x_188, 1, x_130); lean_ctor_set(x_188, 2, x_187); -x_6 = x_188; -x_7 = x_119; -x_14 = x_126; +x_7 = x_188; +x_8 = x_119; +x_15 = x_126; goto _start; } else { -lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; +lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_dec(x_118); x_190 = lean_ctor_get(x_161, 0); lean_inc(x_190); lean_dec(x_161); -x_191 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__1; -lean_inc(x_2); -x_192 = lean_name_mk_string(x_2, x_191); -x_193 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__3; +x_191 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__1; +lean_inc(x_3); +x_192 = lean_name_mk_string(x_3, x_191); +x_193 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; x_194 = l_String_intercalate(x_193, x_190); -x_195 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__4; +x_195 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4; x_196 = lean_string_append(x_195, x_194); lean_dec(x_194); -x_197 = l_Lean_nameLitKind; -x_198 = lean_box(2); -x_199 = l_Lean_Syntax_mkLit(x_197, x_196, x_198); -x_200 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; -x_201 = lean_array_push(x_200, x_199); -x_202 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_202, 0, x_198); -lean_ctor_set(x_202, 1, x_192); -lean_ctor_set(x_202, 2, x_201); -x_203 = lean_array_push(x_163, x_202); -x_204 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_205 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_205, 0, x_198); -lean_ctor_set(x_205, 1, x_204); -lean_ctor_set(x_205, 2, x_203); -x_206 = lean_array_push(x_164, x_205); +x_197 = lean_box(2); +x_198 = l_Lean_Syntax_mkNameLit(x_196, x_197); +x_199 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; +x_200 = lean_array_push(x_199, x_198); +x_201 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_201, 0, x_197); +lean_ctor_set(x_201, 1, x_192); +lean_ctor_set(x_201, 2, x_200); +x_202 = lean_array_push(x_163, x_201); +x_203 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_204 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_204, 0, x_197); +lean_ctor_set(x_204, 1, x_203); +lean_ctor_set(x_204, 2, x_202); +x_205 = lean_array_push(x_164, x_204); lean_inc(x_130); -x_207 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_207, 0, x_198); -lean_ctor_set(x_207, 1, x_130); -lean_ctor_set(x_207, 2, x_206); -x_208 = lean_array_push(x_162, x_207); -x_209 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; -x_210 = lean_array_push(x_208, x_209); -x_211 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_211, 0, x_198); -lean_ctor_set(x_211, 1, x_204); -lean_ctor_set(x_211, 2, x_210); -x_212 = lean_array_push(x_168, x_211); -x_213 = lean_array_push(x_212, x_166); -x_214 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_214, 0, x_198); -lean_ctor_set(x_214, 1, x_140); -lean_ctor_set(x_214, 2, x_213); -x_215 = lean_array_push(x_169, x_214); -x_216 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_216, 0, x_198); -lean_ctor_set(x_216, 1, x_204); -lean_ctor_set(x_216, 2, x_215); -x_217 = lean_array_push(x_170, x_216); -x_218 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_218, 0, x_198); -lean_ctor_set(x_218, 1, x_130); -lean_ctor_set(x_218, 2, x_217); -x_6 = x_218; -x_7 = x_119; -x_14 = x_126; +x_206 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_206, 0, x_197); +lean_ctor_set(x_206, 1, x_130); +lean_ctor_set(x_206, 2, x_205); +x_207 = lean_array_push(x_162, x_206); +x_208 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; +x_209 = lean_array_push(x_207, x_208); +x_210 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_210, 0, x_197); +lean_ctor_set(x_210, 1, x_203); +lean_ctor_set(x_210, 2, x_209); +x_211 = lean_array_push(x_168, x_210); +x_212 = lean_array_push(x_211, x_166); +x_213 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_213, 0, x_197); +lean_ctor_set(x_213, 1, x_140); +lean_ctor_set(x_213, 2, x_212); +x_214 = lean_array_push(x_169, x_213); +x_215 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_215, 0, x_197); +lean_ctor_set(x_215, 1, x_203); +lean_ctor_set(x_215, 2, x_214); +x_216 = lean_array_push(x_170, x_215); +x_217 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_217, 0, x_197); +lean_ctor_set(x_217, 1, x_130); +lean_ctor_set(x_217, 2, x_216); +x_7 = x_217; +x_8 = x_119; +x_15 = x_126; goto _start; } } } } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___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, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___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, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -lean_inc(x_8); -x_11 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_8, x_9, x_10); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +lean_inc(x_9); +x_12 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_9, x_10, x_11); +x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); -lean_dec(x_11); -x_14 = lean_ctor_get(x_8, 10); +x_14 = lean_ctor_get(x_12, 1); lean_inc(x_14); -lean_dec(x_8); -x_15 = lean_st_ref_get(x_9, x_13); -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) +lean_dec(x_12); +x_15 = lean_ctor_get(x_9, 10); +lean_inc(x_15); +lean_dec(x_9); +x_16 = lean_st_ref_get(x_10, x_14); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) { -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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; -x_17 = lean_ctor_get(x_15, 0); -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -lean_dec(x_17); -x_19 = lean_environment_main_module(x_18); -x_20 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__7; +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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_18 = lean_ctor_get(x_16, 0); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +lean_dec(x_18); +x_20 = lean_environment_main_module(x_19); +x_21 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__7; lean_inc(x_1); -x_21 = lean_name_mk_string(x_1, x_20); -lean_inc(x_12); -x_22 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_22, 0, x_12); -lean_ctor_set(x_22, 1, x_20); -x_23 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__9; +x_22 = lean_name_mk_string(x_1, x_21); +lean_inc(x_13); +x_23 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_23, 0, x_13); +lean_ctor_set(x_23, 1, x_21); +x_24 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__9; lean_inc(x_1); -x_24 = lean_name_mk_string(x_1, x_23); -x_25 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; -x_26 = lean_name_mk_string(x_1, x_25); -x_27 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__17; -lean_inc(x_12); -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_12); -lean_ctor_set(x_28, 1, x_27); -x_29 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -x_30 = l_Lean_addMacroScope(x_19, x_29, x_14); -x_31 = lean_box(0); +x_25 = lean_name_mk_string(x_1, x_24); +x_26 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; +x_27 = lean_name_mk_string(x_1, x_26); +x_28 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__17; +lean_inc(x_13); +x_29 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_29, 0, x_13); +lean_ctor_set(x_29, 1, x_28); +x_30 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +x_31 = l_Lean_addMacroScope(x_20, x_30, x_15); x_32 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; -lean_inc(x_12); +lean_inc(x_13); x_33 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_33, 0, x_12); +lean_ctor_set(x_33, 0, x_13); lean_ctor_set(x_33, 1, x_32); -lean_ctor_set(x_33, 2, x_30); -lean_ctor_set(x_33, 3, x_31); +lean_ctor_set(x_33, 2, x_31); +lean_ctor_set(x_33, 3, x_2); x_34 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; -x_35 = lean_array_push(x_34, x_2); +x_35 = lean_array_push(x_34, x_3); x_36 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; x_37 = lean_array_push(x_35, x_36); x_38 = lean_array_push(x_37, x_36); -x_39 = lean_array_push(x_38, x_28); +x_39 = lean_array_push(x_38, x_29); x_40 = lean_array_push(x_39, x_33); x_41 = lean_box(2); x_42 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_26); +lean_ctor_set(x_42, 1, x_27); lean_ctor_set(x_42, 2, x_40); x_43 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; x_44 = lean_array_push(x_43, x_42); x_45 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_45, 0, x_41); -lean_ctor_set(x_45, 1, x_24); +lean_ctor_set(x_45, 1, x_25); lean_ctor_set(x_45, 2, x_44); x_46 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__20; x_47 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_47, 0, x_12); +lean_ctor_set(x_47, 0, x_13); lean_ctor_set(x_47, 1, x_46); x_48 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; -x_49 = lean_array_push(x_48, x_22); +x_49 = lean_array_push(x_48, x_23); x_50 = lean_array_push(x_49, x_45); x_51 = lean_array_push(x_50, x_47); -x_52 = lean_array_push(x_51, x_3); +x_52 = lean_array_push(x_51, x_4); x_53 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_53, 0, x_41); -lean_ctor_set(x_53, 1, x_21); +lean_ctor_set(x_53, 1, x_22); lean_ctor_set(x_53, 2, x_52); -lean_ctor_set(x_15, 0, x_53); -return x_15; +lean_ctor_set(x_16, 0, x_53); +return x_16; } else { -lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; -x_54 = lean_ctor_get(x_15, 0); -x_55 = lean_ctor_get(x_15, 1); +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_54 = lean_ctor_get(x_16, 0); +x_55 = lean_ctor_get(x_16, 1); lean_inc(x_55); lean_inc(x_54); -lean_dec(x_15); +lean_dec(x_16); x_56 = lean_ctor_get(x_54, 0); lean_inc(x_56); lean_dec(x_54); @@ -27041,9 +28497,9 @@ x_57 = lean_environment_main_module(x_56); x_58 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__7; lean_inc(x_1); x_59 = lean_name_mk_string(x_1, x_58); -lean_inc(x_12); +lean_inc(x_13); x_60 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_60, 0, x_12); +lean_ctor_set(x_60, 0, x_13); lean_ctor_set(x_60, 1, x_58); x_61 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__9; lean_inc(x_1); @@ -27051,159 +28507,162 @@ x_62 = lean_name_mk_string(x_1, x_61); x_63 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; x_64 = lean_name_mk_string(x_1, x_63); x_65 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__17; -lean_inc(x_12); +lean_inc(x_13); x_66 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_66, 0, x_12); +lean_ctor_set(x_66, 0, x_13); lean_ctor_set(x_66, 1, x_65); x_67 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -x_68 = l_Lean_addMacroScope(x_57, x_67, x_14); -x_69 = lean_box(0); -x_70 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; -lean_inc(x_12); -x_71 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_71, 0, x_12); -lean_ctor_set(x_71, 1, x_70); -lean_ctor_set(x_71, 2, x_68); -lean_ctor_set(x_71, 3, x_69); -x_72 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; -x_73 = lean_array_push(x_72, x_2); -x_74 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; -x_75 = lean_array_push(x_73, x_74); -x_76 = lean_array_push(x_75, x_74); -x_77 = lean_array_push(x_76, x_66); -x_78 = lean_array_push(x_77, x_71); -x_79 = lean_box(2); -x_80 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_80, 0, x_79); -lean_ctor_set(x_80, 1, x_64); -lean_ctor_set(x_80, 2, x_78); -x_81 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; -x_82 = lean_array_push(x_81, x_80); -x_83 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_83, 0, x_79); -lean_ctor_set(x_83, 1, x_62); -lean_ctor_set(x_83, 2, x_82); -x_84 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__20; -x_85 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_85, 0, x_12); -lean_ctor_set(x_85, 1, x_84); -x_86 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; -x_87 = lean_array_push(x_86, x_60); -x_88 = lean_array_push(x_87, x_83); -x_89 = lean_array_push(x_88, x_85); -x_90 = lean_array_push(x_89, x_3); -x_91 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_91, 0, x_79); -lean_ctor_set(x_91, 1, x_59); -lean_ctor_set(x_91, 2, x_90); -x_92 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_92, 0, x_91); -lean_ctor_set(x_92, 1, x_55); -return x_92; +x_68 = l_Lean_addMacroScope(x_57, x_67, x_15); +x_69 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; +lean_inc(x_13); +x_70 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_70, 0, x_13); +lean_ctor_set(x_70, 1, x_69); +lean_ctor_set(x_70, 2, x_68); +lean_ctor_set(x_70, 3, x_2); +x_71 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; +x_72 = lean_array_push(x_71, x_3); +x_73 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; +x_74 = lean_array_push(x_72, x_73); +x_75 = lean_array_push(x_74, x_73); +x_76 = lean_array_push(x_75, x_66); +x_77 = lean_array_push(x_76, x_70); +x_78 = lean_box(2); +x_79 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_79, 0, x_78); +lean_ctor_set(x_79, 1, x_64); +lean_ctor_set(x_79, 2, x_77); +x_80 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; +x_81 = lean_array_push(x_80, x_79); +x_82 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_82, 0, x_78); +lean_ctor_set(x_82, 1, x_62); +lean_ctor_set(x_82, 2, x_81); +x_83 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__20; +x_84 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_84, 0, x_13); +lean_ctor_set(x_84, 1, x_83); +x_85 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; +x_86 = lean_array_push(x_85, x_60); +x_87 = lean_array_push(x_86, x_82); +x_88 = lean_array_push(x_87, x_84); +x_89 = lean_array_push(x_88, x_4); +x_90 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_90, 0, x_78); +lean_ctor_set(x_90, 1, x_59); +lean_ctor_set(x_90, 2, x_89); +x_91 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_91, 0, x_90); +lean_ctor_set(x_91, 1, x_55); +return x_91; } } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___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_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_12; lean_object* x_13; -x_12 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__1___boxed), 10, 2); -lean_closure_set(x_12, 0, x_1); -lean_closure_set(x_12, 1, x_2); +lean_object* x_13; lean_object* x_14; +x_13 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__1___boxed), 11, 3); +lean_closure_set(x_13, 0, x_1); +lean_closure_set(x_13, 1, x_2); +lean_closure_set(x_13, 2, x_3); +lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_5); -x_13 = lean_apply_8(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_13) == 0) +x_14 = lean_apply_8(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); -lean_dec(x_13); -x_16 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_adaptRhs(x_12, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_15); -return x_16; +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_adaptRhs(x_13, x_15, x_6, x_7, x_8, x_9, x_10, x_11, x_16); +return x_17; } else { -uint8_t x_17; -lean_dec(x_12); +uint8_t x_18; +lean_dec(x_13); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -x_17 = !lean_is_exclusive(x_13); -if (x_17 == 0) +x_18 = !lean_is_exclusive(x_14); +if (x_18 == 0) { -return x_13; +return x_14; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_13, 0); -x_19 = lean_ctor_get(x_13, 1); +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_14, 0); +x_20 = lean_ctor_get(x_14, 1); +lean_inc(x_20); lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_13); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_18); -lean_ctor_set(x_20, 1, x_19); -return x_20; +lean_dec(x_14); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +return x_21; } } } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__3(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; -x_5 = lean_ctor_get(x_1, 1); -lean_inc(x_5); +lean_object* x_6; lean_object* x_7; +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); lean_dec(x_1); -x_6 = lean_apply_1(x_5, x_4); -if (lean_obj_tag(x_6) == 0) +x_7 = lean_apply_1(x_6, x_5); +if (lean_obj_tag(x_7) == 0) { -uint8_t x_7; -x_7 = !lean_is_exclusive(x_6); -if (x_7 == 0) +uint8_t x_8; +x_8 = !lean_is_exclusive(x_7); +if (x_8 == 0) { -lean_object* x_8; lean_object* x_9; -x_8 = lean_ctor_get(x_6, 0); -x_9 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__2), 11, 3); -lean_closure_set(x_9, 0, x_2); -lean_closure_set(x_9, 1, x_3); -lean_closure_set(x_9, 2, x_8); -lean_ctor_set(x_6, 0, x_9); -return x_6; +lean_object* x_9; lean_object* x_10; +x_9 = lean_ctor_get(x_7, 0); +x_10 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__2), 12, 4); +lean_closure_set(x_10, 0, x_2); +lean_closure_set(x_10, 1, x_3); +lean_closure_set(x_10, 2, x_4); +lean_closure_set(x_10, 3, x_9); +lean_ctor_set(x_7, 0, x_10); +return x_7; } else { -lean_object* x_10; uint8_t x_11; lean_object* x_12; lean_object* x_13; -x_10 = lean_ctor_get(x_6, 0); -x_11 = lean_ctor_get_uint8(x_6, sizeof(void*)*1); -lean_inc(x_10); -lean_dec(x_6); -x_12 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__2), 11, 3); -lean_closure_set(x_12, 0, x_2); -lean_closure_set(x_12, 1, x_3); -lean_closure_set(x_12, 2, x_10); -x_13 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set_uint8(x_13, sizeof(void*)*1, x_11); -return x_13; +lean_object* x_11; uint8_t x_12; lean_object* x_13; lean_object* x_14; +x_11 = lean_ctor_get(x_7, 0); +x_12 = lean_ctor_get_uint8(x_7, sizeof(void*)*1); +lean_inc(x_11); +lean_dec(x_7); +x_13 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__2), 12, 4); +lean_closure_set(x_13, 0, x_2); +lean_closure_set(x_13, 1, x_3); +lean_closure_set(x_13, 2, x_4); +lean_closure_set(x_13, 3, x_11); +x_14 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set_uint8(x_14, sizeof(void*)*1, x_12); +return x_14; } } else { +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_6; +return x_7; } } } @@ -27240,12 +28699,11 @@ return x_9; } } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__5(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_object* x_9) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__5(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_object* x_9, lean_object* x_10) { _start: { -lean_object* x_10; lean_object* x_11; -x_10 = lean_box(0); -x_11 = lean_apply_8(x_1, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_object* x_11; +x_11 = lean_apply_8(x_2, x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10); return x_11; } } @@ -27253,7 +28711,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__1; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__1; x_2 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_adaptRhs), 9, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -27354,51 +28812,68 @@ return x_3; static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__5() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__4; -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; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Eq", 2); +return x_1; } } static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__6() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__5; +x_2 = lean_string_utf8_byte_size(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__5; +x_2 = lean_unsigned_to_nat(0u); +x_3 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__6; +x_4 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__8() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__5; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); +x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__7() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__9() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Eq", 2); +x_1 = lean_mk_string_from_bytes("true", 4); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__8() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__7; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__9; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__9() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__7; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__9; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__8; +x_3 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__10; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -27406,463 +28881,664 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__10() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__7; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__9; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__11() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__13() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__10; -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; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Bool", 4); +return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__12() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__11; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__13; +x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__13() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__15() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("true", 4); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__14; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__9; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__14() { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8(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_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__13; -x_2 = lean_string_utf8_byte_size(x_1); -return x_2; -} +uint8_t x_12; +x_12 = !lean_is_exclusive(x_4); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_4, 0); +x_14 = lean_ctor_get(x_4, 1); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_15 = lean_apply_8(x_1, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +lean_inc(x_10); +lean_inc(x_9); +x_18 = lean_apply_7(x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_17); +if (lean_obj_tag(x_18) == 0) +{ +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; uint8_t x_26; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +lean_inc(x_9); +x_21 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_9, x_10, x_20); +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = lean_ctor_get(x_9, 10); +lean_inc(x_24); +lean_dec(x_9); +x_25 = lean_st_ref_get(x_10, x_23); +lean_dec(x_10); +x_26 = !lean_is_exclusive(x_25); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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_27 = lean_ctor_get(x_25, 0); +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +lean_dec(x_27); +x_29 = lean_environment_main_module(x_28); +x_30 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__4; +lean_inc(x_24); +lean_inc(x_29); +x_31 = l_Lean_addMacroScope(x_29, x_30, x_24); +lean_inc(x_3); +lean_ctor_set(x_4, 1, x_3); +lean_ctor_set(x_4, 0, x_30); +lean_inc(x_3); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_4); +lean_ctor_set(x_32, 1, x_3); +x_33 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__3; +lean_inc(x_22); +x_34 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_34, 0, x_22); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_34, 3, x_32); +x_35 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; +lean_inc(x_22); +x_36 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_36, 0, x_22); +lean_ctor_set(x_36, 1, x_35); +x_37 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__8; +lean_inc(x_24); +lean_inc(x_29); +x_38 = l_Lean_addMacroScope(x_29, x_37, x_24); +lean_inc(x_3); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_3); +lean_inc(x_3); +x_40 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_3); +x_41 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__7; +lean_inc(x_22); +x_42 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_42, 0, x_22); +lean_ctor_set(x_42, 1, x_41); +lean_ctor_set(x_42, 2, x_38); +lean_ctor_set(x_42, 3, x_40); +x_43 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__12; +x_44 = l_Lean_addMacroScope(x_29, x_43, x_24); +x_45 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__15; +lean_inc(x_3); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_3); +x_47 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_3); +x_48 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__11; +lean_inc(x_22); +x_49 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_49, 0, x_22); +lean_ctor_set(x_49, 1, x_48); +lean_ctor_set(x_49, 2, x_44); +lean_ctor_set(x_49, 3, x_47); +x_50 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_51 = lean_array_push(x_50, x_13); +x_52 = lean_array_push(x_51, x_49); +x_53 = lean_box(2); +x_54 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_55 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +lean_ctor_set(x_55, 2, x_52); +x_56 = lean_array_push(x_50, x_42); +x_57 = lean_array_push(x_56, x_55); +x_58 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_59 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_59, 0, x_53); +lean_ctor_set(x_59, 1, x_58); +lean_ctor_set(x_59, 2, x_57); +x_60 = lean_array_push(x_50, x_59); +x_61 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; +x_62 = lean_array_push(x_60, x_61); +x_63 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_63, 0, x_53); +lean_ctor_set(x_63, 1, x_54); +lean_ctor_set(x_63, 2, x_62); +x_64 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; +x_65 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_65, 0, x_22); +lean_ctor_set(x_65, 1, x_64); +x_66 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; +x_67 = lean_array_push(x_66, x_36); +x_68 = lean_array_push(x_67, x_63); +x_69 = lean_array_push(x_68, x_65); +x_70 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6; +x_71 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_71, 0, x_53); +lean_ctor_set(x_71, 1, x_70); +lean_ctor_set(x_71, 2, x_69); +x_72 = lean_array_push(x_66, x_71); +x_73 = lean_array_push(x_72, x_16); +x_74 = lean_array_push(x_73, x_19); +x_75 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_75, 0, x_53); +lean_ctor_set(x_75, 1, x_54); +lean_ctor_set(x_75, 2, x_74); +x_76 = lean_array_push(x_50, x_34); +x_77 = lean_array_push(x_76, x_75); +x_78 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_78, 0, x_53); +lean_ctor_set(x_78, 1, x_58); +lean_ctor_set(x_78, 2, x_77); +lean_ctor_set(x_25, 0, x_78); +return x_25; } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__15() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__13; -x_2 = lean_unsigned_to_nat(0u); -x_3 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__14; -x_4 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_4, 0, x_1); -lean_ctor_set(x_4, 1, x_2); -lean_ctor_set(x_4, 2, x_3); -return x_4; +lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; +x_79 = lean_ctor_get(x_25, 0); +x_80 = lean_ctor_get(x_25, 1); +lean_inc(x_80); +lean_inc(x_79); +lean_dec(x_25); +x_81 = lean_ctor_get(x_79, 0); +lean_inc(x_81); +lean_dec(x_79); +x_82 = lean_environment_main_module(x_81); +x_83 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__4; +lean_inc(x_24); +lean_inc(x_82); +x_84 = l_Lean_addMacroScope(x_82, x_83, x_24); +lean_inc(x_3); +lean_ctor_set(x_4, 1, x_3); +lean_ctor_set(x_4, 0, x_83); +lean_inc(x_3); +x_85 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_85, 0, x_4); +lean_ctor_set(x_85, 1, x_3); +x_86 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__3; +lean_inc(x_22); +x_87 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_87, 0, x_22); +lean_ctor_set(x_87, 1, x_86); +lean_ctor_set(x_87, 2, x_84); +lean_ctor_set(x_87, 3, x_85); +x_88 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; +lean_inc(x_22); +x_89 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_89, 0, x_22); +lean_ctor_set(x_89, 1, x_88); +x_90 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__8; +lean_inc(x_24); +lean_inc(x_82); +x_91 = l_Lean_addMacroScope(x_82, x_90, x_24); +lean_inc(x_3); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_3); +lean_inc(x_3); +x_93 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_93, 0, x_92); +lean_ctor_set(x_93, 1, x_3); +x_94 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__7; +lean_inc(x_22); +x_95 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_95, 0, x_22); +lean_ctor_set(x_95, 1, x_94); +lean_ctor_set(x_95, 2, x_91); +lean_ctor_set(x_95, 3, x_93); +x_96 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__12; +x_97 = l_Lean_addMacroScope(x_82, x_96, x_24); +x_98 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__15; +lean_inc(x_3); +x_99 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_99, 0, x_98); +lean_ctor_set(x_99, 1, x_3); +x_100 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_100, 0, x_99); +lean_ctor_set(x_100, 1, x_3); +x_101 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__11; +lean_inc(x_22); +x_102 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_102, 0, x_22); +lean_ctor_set(x_102, 1, x_101); +lean_ctor_set(x_102, 2, x_97); +lean_ctor_set(x_102, 3, x_100); +x_103 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_104 = lean_array_push(x_103, x_13); +x_105 = lean_array_push(x_104, x_102); +x_106 = lean_box(2); +x_107 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_108 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_108, 0, x_106); +lean_ctor_set(x_108, 1, x_107); +lean_ctor_set(x_108, 2, x_105); +x_109 = lean_array_push(x_103, x_95); +x_110 = lean_array_push(x_109, x_108); +x_111 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_112 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_112, 0, x_106); +lean_ctor_set(x_112, 1, x_111); +lean_ctor_set(x_112, 2, x_110); +x_113 = lean_array_push(x_103, x_112); +x_114 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; +x_115 = lean_array_push(x_113, x_114); +x_116 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_116, 0, x_106); +lean_ctor_set(x_116, 1, x_107); +lean_ctor_set(x_116, 2, x_115); +x_117 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; +x_118 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_118, 0, x_22); +lean_ctor_set(x_118, 1, x_117); +x_119 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; +x_120 = lean_array_push(x_119, x_89); +x_121 = lean_array_push(x_120, x_116); +x_122 = lean_array_push(x_121, x_118); +x_123 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6; +x_124 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_124, 0, x_106); +lean_ctor_set(x_124, 1, x_123); +lean_ctor_set(x_124, 2, x_122); +x_125 = lean_array_push(x_119, x_124); +x_126 = lean_array_push(x_125, x_16); +x_127 = lean_array_push(x_126, x_19); +x_128 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_128, 0, x_106); +lean_ctor_set(x_128, 1, x_107); +lean_ctor_set(x_128, 2, x_127); +x_129 = lean_array_push(x_103, x_87); +x_130 = lean_array_push(x_129, x_128); +x_131 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_131, 0, x_106); +lean_ctor_set(x_131, 1, x_111); +lean_ctor_set(x_131, 2, x_130); +x_132 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_132, 0, x_131); +lean_ctor_set(x_132, 1, x_80); +return x_132; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__16() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__13; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__17() { -_start: +uint8_t x_133; +lean_dec(x_16); +lean_free_object(x_4); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_3); +x_133 = !lean_is_exclusive(x_18); +if (x_133 == 0) { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Bool", 4); -return x_1; -} +return x_18; } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__18() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__17; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} +lean_object* x_134; lean_object* x_135; lean_object* x_136; +x_134 = lean_ctor_get(x_18, 0); +x_135 = lean_ctor_get(x_18, 1); +lean_inc(x_135); +lean_inc(x_134); +lean_dec(x_18); +x_136 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_136, 0, x_134); +lean_ctor_set(x_136, 1, x_135); +return x_136; } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__19() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__18; -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__13; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__20() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__19; -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; -} +uint8_t x_137; +lean_free_object(x_4); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_137 = !lean_is_exclusive(x_15); +if (x_137 == 0) +{ +return x_15; } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__21() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__20; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; +lean_object* x_138; lean_object* x_139; lean_object* x_140; +x_138 = lean_ctor_get(x_15, 0); +x_139 = lean_ctor_get(x_15, 1); +lean_inc(x_139); +lean_inc(x_138); +lean_dec(x_15); +x_140 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_140, 0, x_138); +lean_ctor_set(x_140, 1, x_139); +return x_140; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8(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_object* x_9, lean_object* x_10) { -_start: +} +else { -lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_11 = lean_ctor_get(x_3, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_3, 1); -lean_inc(x_12); -lean_dec(x_3); +lean_object* x_141; lean_object* x_142; lean_object* x_143; +x_141 = lean_ctor_get(x_4, 0); +x_142 = lean_ctor_get(x_4, 1); +lean_inc(x_142); +lean_inc(x_141); +lean_dec(x_4); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_4); -x_13 = lean_apply_8(x_1, x_12, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_13) == 0) +x_143 = lean_apply_8(x_1, x_142, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_143) == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); +lean_object* x_144; lean_object* x_145; lean_object* x_146; +x_144 = lean_ctor_get(x_143, 0); +lean_inc(x_144); +x_145 = lean_ctor_get(x_143, 1); +lean_inc(x_145); +lean_dec(x_143); +lean_inc(x_10); lean_inc(x_9); -lean_inc(x_8); -x_16 = lean_apply_7(x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_15); -if (lean_obj_tag(x_16) == 0) +x_146 = lean_apply_7(x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_145); +if (lean_obj_tag(x_146) == 0) { -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; uint8_t x_24; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec(x_16); -lean_inc(x_8); -x_19 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_8, x_9, x_18); -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_19, 1); -lean_inc(x_21); -lean_dec(x_19); -x_22 = lean_ctor_get(x_8, 10); -lean_inc(x_22); -lean_dec(x_8); -x_23 = lean_st_ref_get(x_9, x_21); +lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; +x_147 = lean_ctor_get(x_146, 0); +lean_inc(x_147); +x_148 = lean_ctor_get(x_146, 1); +lean_inc(x_148); +lean_dec(x_146); +lean_inc(x_9); +x_149 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_9, x_10, x_148); +x_150 = lean_ctor_get(x_149, 0); +lean_inc(x_150); +x_151 = lean_ctor_get(x_149, 1); +lean_inc(x_151); +lean_dec(x_149); +x_152 = lean_ctor_get(x_9, 10); +lean_inc(x_152); lean_dec(x_9); -x_24 = !lean_is_exclusive(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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; -x_25 = lean_ctor_get(x_23, 0); -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -lean_dec(x_25); -x_27 = lean_environment_main_module(x_26); -x_28 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__4; -lean_inc(x_22); -lean_inc(x_27); -x_29 = l_Lean_addMacroScope(x_27, x_28, x_22); -x_30 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__3; -x_31 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__6; -lean_inc(x_20); -x_32 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_32, 0, x_20); -lean_ctor_set(x_32, 1, x_30); -lean_ctor_set(x_32, 2, x_29); -lean_ctor_set(x_32, 3, x_31); -x_33 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; -lean_inc(x_20); -x_34 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_34, 0, x_20); -lean_ctor_set(x_34, 1, x_33); -x_35 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__10; -lean_inc(x_22); -lean_inc(x_27); -x_36 = l_Lean_addMacroScope(x_27, x_35, x_22); -x_37 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__9; -x_38 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__12; -lean_inc(x_20); -x_39 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_39, 0, x_20); -lean_ctor_set(x_39, 1, x_37); -lean_ctor_set(x_39, 2, x_36); -lean_ctor_set(x_39, 3, x_38); -x_40 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__16; -x_41 = l_Lean_addMacroScope(x_27, x_40, x_22); -x_42 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__15; -x_43 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__21; -lean_inc(x_20); -x_44 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_44, 0, x_20); -lean_ctor_set(x_44, 1, x_42); -lean_ctor_set(x_44, 2, x_41); -lean_ctor_set(x_44, 3, x_43); -x_45 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_46 = lean_array_push(x_45, x_11); -x_47 = lean_array_push(x_46, x_44); -x_48 = lean_box(2); -x_49 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_50 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_50, 0, x_48); -lean_ctor_set(x_50, 1, x_49); -lean_ctor_set(x_50, 2, x_47); -x_51 = lean_array_push(x_45, x_39); -x_52 = lean_array_push(x_51, x_50); -x_53 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_54 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_54, 0, x_48); -lean_ctor_set(x_54, 1, x_53); -lean_ctor_set(x_54, 2, x_52); -x_55 = lean_array_push(x_45, x_54); -x_56 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; -x_57 = lean_array_push(x_55, x_56); -x_58 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_58, 0, x_48); -lean_ctor_set(x_58, 1, x_49); -lean_ctor_set(x_58, 2, x_57); -x_59 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; -x_60 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_60, 0, x_20); -lean_ctor_set(x_60, 1, x_59); -x_61 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; -x_62 = lean_array_push(x_61, x_34); -x_63 = lean_array_push(x_62, x_58); -x_64 = lean_array_push(x_63, x_60); -x_65 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__6; -x_66 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_66, 0, x_48); -lean_ctor_set(x_66, 1, x_65); -lean_ctor_set(x_66, 2, x_64); -x_67 = lean_array_push(x_61, x_66); -x_68 = lean_array_push(x_67, x_14); -x_69 = lean_array_push(x_68, x_17); -x_70 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_70, 0, x_48); -lean_ctor_set(x_70, 1, x_49); -lean_ctor_set(x_70, 2, x_69); -x_71 = lean_array_push(x_45, x_32); -x_72 = lean_array_push(x_71, x_70); -x_73 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_73, 0, x_48); -lean_ctor_set(x_73, 1, x_53); -lean_ctor_set(x_73, 2, x_72); -lean_ctor_set(x_23, 0, x_73); -return x_23; +x_153 = lean_st_ref_get(x_10, x_151); +lean_dec(x_10); +x_154 = lean_ctor_get(x_153, 0); +lean_inc(x_154); +x_155 = lean_ctor_get(x_153, 1); +lean_inc(x_155); +if (lean_is_exclusive(x_153)) { + lean_ctor_release(x_153, 0); + lean_ctor_release(x_153, 1); + x_156 = x_153; +} else { + lean_dec_ref(x_153); + x_156 = lean_box(0); } -else -{ -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; -x_74 = lean_ctor_get(x_23, 0); -x_75 = lean_ctor_get(x_23, 1); -lean_inc(x_75); -lean_inc(x_74); -lean_dec(x_23); -x_76 = lean_ctor_get(x_74, 0); -lean_inc(x_76); -lean_dec(x_74); -x_77 = lean_environment_main_module(x_76); -x_78 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__4; -lean_inc(x_22); -lean_inc(x_77); -x_79 = l_Lean_addMacroScope(x_77, x_78, x_22); -x_80 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__3; -x_81 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__6; -lean_inc(x_20); -x_82 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_82, 0, x_20); -lean_ctor_set(x_82, 1, x_80); -lean_ctor_set(x_82, 2, x_79); -lean_ctor_set(x_82, 3, x_81); -x_83 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; -lean_inc(x_20); -x_84 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_84, 0, x_20); -lean_ctor_set(x_84, 1, x_83); -x_85 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__10; -lean_inc(x_22); -lean_inc(x_77); -x_86 = l_Lean_addMacroScope(x_77, x_85, x_22); -x_87 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__9; -x_88 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__12; -lean_inc(x_20); -x_89 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_89, 0, x_20); -lean_ctor_set(x_89, 1, x_87); -lean_ctor_set(x_89, 2, x_86); -lean_ctor_set(x_89, 3, x_88); -x_90 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__16; -x_91 = l_Lean_addMacroScope(x_77, x_90, x_22); -x_92 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__15; -x_93 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__21; -lean_inc(x_20); -x_94 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_94, 0, x_20); -lean_ctor_set(x_94, 1, x_92); -lean_ctor_set(x_94, 2, x_91); -lean_ctor_set(x_94, 3, x_93); -x_95 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_96 = lean_array_push(x_95, x_11); -x_97 = lean_array_push(x_96, x_94); -x_98 = lean_box(2); -x_99 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_100 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_100, 0, x_98); -lean_ctor_set(x_100, 1, x_99); -lean_ctor_set(x_100, 2, x_97); -x_101 = lean_array_push(x_95, x_89); -x_102 = lean_array_push(x_101, x_100); -x_103 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_104 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_104, 0, x_98); -lean_ctor_set(x_104, 1, x_103); -lean_ctor_set(x_104, 2, x_102); -x_105 = lean_array_push(x_95, x_104); -x_106 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; -x_107 = lean_array_push(x_105, x_106); -x_108 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_108, 0, x_98); -lean_ctor_set(x_108, 1, x_99); -lean_ctor_set(x_108, 2, x_107); -x_109 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; -x_110 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_110, 0, x_20); -lean_ctor_set(x_110, 1, x_109); -x_111 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; -x_112 = lean_array_push(x_111, x_84); -x_113 = lean_array_push(x_112, x_108); -x_114 = lean_array_push(x_113, x_110); -x_115 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__6; -x_116 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_116, 0, x_98); -lean_ctor_set(x_116, 1, x_115); -lean_ctor_set(x_116, 2, x_114); -x_117 = lean_array_push(x_111, x_116); -x_118 = lean_array_push(x_117, x_14); -x_119 = lean_array_push(x_118, x_17); -x_120 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_120, 0, x_98); -lean_ctor_set(x_120, 1, x_99); -lean_ctor_set(x_120, 2, x_119); -x_121 = lean_array_push(x_95, x_82); -x_122 = lean_array_push(x_121, x_120); -x_123 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_123, 0, x_98); -lean_ctor_set(x_123, 1, x_103); -lean_ctor_set(x_123, 2, x_122); -x_124 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_124, 0, x_123); -lean_ctor_set(x_124, 1, x_75); -return x_124; +x_157 = lean_ctor_get(x_154, 0); +lean_inc(x_157); +lean_dec(x_154); +x_158 = lean_environment_main_module(x_157); +x_159 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__4; +lean_inc(x_152); +lean_inc(x_158); +x_160 = l_Lean_addMacroScope(x_158, x_159, x_152); +lean_inc(x_3); +x_161 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_161, 0, x_159); +lean_ctor_set(x_161, 1, x_3); +lean_inc(x_3); +x_162 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_162, 0, x_161); +lean_ctor_set(x_162, 1, x_3); +x_163 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__3; +lean_inc(x_150); +x_164 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_164, 0, x_150); +lean_ctor_set(x_164, 1, x_163); +lean_ctor_set(x_164, 2, x_160); +lean_ctor_set(x_164, 3, x_162); +x_165 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; +lean_inc(x_150); +x_166 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_166, 0, x_150); +lean_ctor_set(x_166, 1, x_165); +x_167 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__8; +lean_inc(x_152); +lean_inc(x_158); +x_168 = l_Lean_addMacroScope(x_158, x_167, x_152); +lean_inc(x_3); +x_169 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_169, 0, x_167); +lean_ctor_set(x_169, 1, x_3); +lean_inc(x_3); +x_170 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_170, 0, x_169); +lean_ctor_set(x_170, 1, x_3); +x_171 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__7; +lean_inc(x_150); +x_172 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_172, 0, x_150); +lean_ctor_set(x_172, 1, x_171); +lean_ctor_set(x_172, 2, x_168); +lean_ctor_set(x_172, 3, x_170); +x_173 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__12; +x_174 = l_Lean_addMacroScope(x_158, x_173, x_152); +x_175 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__15; +lean_inc(x_3); +x_176 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_176, 0, x_175); +lean_ctor_set(x_176, 1, x_3); +x_177 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_177, 0, x_176); +lean_ctor_set(x_177, 1, x_3); +x_178 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__11; +lean_inc(x_150); +x_179 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_179, 0, x_150); +lean_ctor_set(x_179, 1, x_178); +lean_ctor_set(x_179, 2, x_174); +lean_ctor_set(x_179, 3, x_177); +x_180 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_181 = lean_array_push(x_180, x_141); +x_182 = lean_array_push(x_181, x_179); +x_183 = lean_box(2); +x_184 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_185 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_185, 0, x_183); +lean_ctor_set(x_185, 1, x_184); +lean_ctor_set(x_185, 2, x_182); +x_186 = lean_array_push(x_180, x_172); +x_187 = lean_array_push(x_186, x_185); +x_188 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_189 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_189, 0, x_183); +lean_ctor_set(x_189, 1, x_188); +lean_ctor_set(x_189, 2, x_187); +x_190 = lean_array_push(x_180, x_189); +x_191 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; +x_192 = lean_array_push(x_190, x_191); +x_193 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_193, 0, x_183); +lean_ctor_set(x_193, 1, x_184); +lean_ctor_set(x_193, 2, x_192); +x_194 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; +x_195 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_195, 0, x_150); +lean_ctor_set(x_195, 1, x_194); +x_196 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; +x_197 = lean_array_push(x_196, x_166); +x_198 = lean_array_push(x_197, x_193); +x_199 = lean_array_push(x_198, x_195); +x_200 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6; +x_201 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_201, 0, x_183); +lean_ctor_set(x_201, 1, x_200); +lean_ctor_set(x_201, 2, x_199); +x_202 = lean_array_push(x_196, x_201); +x_203 = lean_array_push(x_202, x_144); +x_204 = lean_array_push(x_203, x_147); +x_205 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_205, 0, x_183); +lean_ctor_set(x_205, 1, x_184); +lean_ctor_set(x_205, 2, x_204); +x_206 = lean_array_push(x_180, x_164); +x_207 = lean_array_push(x_206, x_205); +x_208 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_208, 0, x_183); +lean_ctor_set(x_208, 1, x_188); +lean_ctor_set(x_208, 2, x_207); +if (lean_is_scalar(x_156)) { + x_209 = lean_alloc_ctor(0, 2, 0); +} else { + x_209 = x_156; } +lean_ctor_set(x_209, 0, x_208); +lean_ctor_set(x_209, 1, x_155); +return x_209; } else { -uint8_t x_125; -lean_dec(x_14); -lean_dec(x_11); +lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; +lean_dec(x_144); +lean_dec(x_141); +lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); -x_125 = !lean_is_exclusive(x_16); -if (x_125 == 0) -{ -return x_16; +lean_dec(x_3); +x_210 = lean_ctor_get(x_146, 0); +lean_inc(x_210); +x_211 = lean_ctor_get(x_146, 1); +lean_inc(x_211); +if (lean_is_exclusive(x_146)) { + lean_ctor_release(x_146, 0); + lean_ctor_release(x_146, 1); + x_212 = x_146; +} else { + lean_dec_ref(x_146); + x_212 = lean_box(0); } -else -{ -lean_object* x_126; lean_object* x_127; lean_object* x_128; -x_126 = lean_ctor_get(x_16, 0); -x_127 = lean_ctor_get(x_16, 1); -lean_inc(x_127); -lean_inc(x_126); -lean_dec(x_16); -x_128 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_128, 0, x_126); -lean_ctor_set(x_128, 1, x_127); -return x_128; +if (lean_is_scalar(x_212)) { + x_213 = lean_alloc_ctor(1, 2, 0); +} else { + x_213 = x_212; } +lean_ctor_set(x_213, 0, x_210); +lean_ctor_set(x_213, 1, x_211); +return x_213; } } else { -uint8_t x_129; -lean_dec(x_11); +lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; +lean_dec(x_141); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); -x_129 = !lean_is_exclusive(x_13); -if (x_129 == 0) -{ -return x_13; +x_214 = lean_ctor_get(x_143, 0); +lean_inc(x_214); +x_215 = lean_ctor_get(x_143, 1); +lean_inc(x_215); +if (lean_is_exclusive(x_143)) { + lean_ctor_release(x_143, 0); + lean_ctor_release(x_143, 1); + x_216 = x_143; +} else { + lean_dec_ref(x_143); + x_216 = lean_box(0); } -else -{ -lean_object* x_130; lean_object* x_131; lean_object* x_132; -x_130 = lean_ctor_get(x_13, 0); -x_131 = lean_ctor_get(x_13, 1); -lean_inc(x_131); -lean_inc(x_130); -lean_dec(x_13); -x_132 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_132, 0, x_130); -lean_ctor_set(x_132, 1, x_131); -return x_132; +if (lean_is_scalar(x_216)) { + x_217 = lean_alloc_ctor(1, 2, 0); +} else { + x_217 = x_216; +} +lean_ctor_set(x_217, 0, x_214); +lean_ctor_set(x_217, 1, x_215); +return x_217; } } } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__9(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_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__9(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_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -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; -x_11 = lean_array_get_size(x_1); -x_12 = l_List_range(x_11); -lean_inc(x_8); -x_13 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4(x_12, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); +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_12 = lean_array_get_size(x_1); +x_13 = l_List_range(x_12); +lean_inc(x_9); +x_14 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4(x_2, x_13, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); -lean_dec(x_13); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_3); -lean_ctor_set(x_16, 1, x_14); -x_17 = lean_apply_8(x_2, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_15); -return x_17; +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_4); +lean_ctor_set(x_17, 1, x_15); +x_18 = lean_apply_8(x_3, x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_16); +return x_18; } } static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__1() { @@ -27883,7 +29559,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__30; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__30; x_2 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__6; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -27892,51 +29568,27 @@ return x_3; static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__3() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__3; -x_3 = lean_alloc_ctor(1, 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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__5() { -_start: -{ lean_object* x_1; x_1 = lean_mk_string_from_bytes("Syntax.matchesIdent", 19); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__6() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__5; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__3; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__7() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__5; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__3; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__6; +x_3 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__4; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -27944,7 +29596,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__8() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__6() { _start: { lean_object* x_1; @@ -27952,51 +29604,27 @@ x_1 = lean_mk_string_from_bytes("matchesIdent", 12); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__9() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__4; -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__8; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__6; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__10() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__30; -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__8; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__30; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__6; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__11() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__10; -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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__12() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__11; -x_3 = lean_alloc_ctor(1, 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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__13() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__9() { _start: { lean_object* x_1; @@ -28004,22 +29632,22 @@ x_1 = lean_mk_string_from_bytes("Syntax.matchesNull", 18); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__14() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__13; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__9; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__15() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__13; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__9; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__14; +x_3 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__10; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -28027,7 +29655,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__16() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__12() { _start: { lean_object* x_1; @@ -28035,51 +29663,27 @@ x_1 = lean_mk_string_from_bytes("matchesNull", 11); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__17() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__4; -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__16; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__12; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__18() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__30; -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__16; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__30; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__12; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__19() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__18; -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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__20() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__19; -x_3 = lean_alloc_ctor(1, 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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__21() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__15() { _start: { lean_object* x_1; @@ -28087,22 +29691,22 @@ x_1 = lean_mk_string_from_bytes("Syntax.matchesLit", 17); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__22() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__16() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__21; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__15; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__23() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__21; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__15; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__22; +x_3 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__16; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -28110,7 +29714,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__24() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__18() { _start: { lean_object* x_1; @@ -28118,51 +29722,27 @@ x_1 = lean_mk_string_from_bytes("matchesLit", 10); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__25() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__4; -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__24; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__18; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__26() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__30; -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__24; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__30; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__18; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__27() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__26; -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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__28() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__27; -x_3 = lean_alloc_ctor(1, 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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__29() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__21() { _start: { lean_object* x_1; @@ -28170,7 +29750,7 @@ x_1 = lean_mk_string_from_bytes("Init.Data.Option.BasicAux", 25); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__30() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__22() { _start: { lean_object* x_1; @@ -28178,7 +29758,7 @@ x_1 = lean_mk_string_from_bytes("Option.get!", 11); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__31() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__23() { _start: { lean_object* x_1; @@ -28186,743 +29766,755 @@ x_1 = lean_mk_string_from_bytes("value is none", 13); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__32() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__24() { _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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__29; -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__30; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__21; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__22; x_3 = lean_unsigned_to_nat(16u); x_4 = lean_unsigned_to_nat(14u); -x_5 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__31; +x_5 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__23; 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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10(uint8_t 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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10(lean_object* x_1, uint8_t 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, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { -if (x_1 == 0) -{ -lean_object* x_14; -x_14 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8), 10, 2); -lean_closure_set(x_14, 0, x_5); -lean_closure_set(x_14, 1, x_6); -if (lean_obj_tag(x_3) == 1) +if (x_2 == 0) { -lean_object* x_15; -x_15 = lean_ctor_get(x_3, 0); -lean_inc(x_15); -if (lean_obj_tag(x_15) == 0) +lean_object* x_16; +lean_inc(x_1); +x_16 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8), 11, 3); +lean_closure_set(x_16, 0, x_7); +lean_closure_set(x_16, 1, x_8); +lean_closure_set(x_16, 2, x_1); +if (lean_obj_tag(x_5) == 1) { -lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_16 = lean_ctor_get(x_3, 1); -lean_inc(x_16); -x_17 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; -x_18 = lean_string_dec_eq(x_16, x_17); -if (x_18 == 0) +lean_object* x_17; +x_17 = lean_ctor_get(x_5, 0); +lean_inc(x_17); +if (lean_obj_tag(x_17) == 0) { -lean_object* x_19; uint8_t x_20; -x_19 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__5; -x_20 = lean_string_dec_eq(x_16, x_19); -lean_dec(x_16); +lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_18 = lean_ctor_get(x_5, 1); +lean_inc(x_18); +x_19 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; +x_20 = lean_string_dec_eq(x_18, x_19); if (x_20 == 0) { -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; 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; -lean_dec(x_4); -lean_inc(x_11); -x_21 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_11, x_12, x_13); -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_21, 1); -lean_inc(x_23); -lean_dec(x_21); -x_24 = lean_ctor_get(x_11, 10); +lean_object* x_21; uint8_t x_22; +x_21 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__5; +x_22 = lean_string_dec_eq(x_18, x_21); +lean_dec(x_18); +if (x_22 == 0) +{ +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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +lean_dec(x_6); +lean_inc(x_13); +x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_13, x_14, x_15); +x_24 = lean_ctor_get(x_23, 0); lean_inc(x_24); -x_25 = lean_st_ref_get(x_12, x_23); -x_26 = lean_ctor_get(x_25, 0); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_ctor_get(x_13, 10); lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 1); -lean_inc(x_27); -lean_dec(x_25); -x_28 = lean_ctor_get(x_26, 0); +x_27 = lean_st_ref_get(x_14, x_25); +x_28 = lean_ctor_get(x_27, 0); lean_inc(x_28); -lean_dec(x_26); -x_29 = lean_environment_main_module(x_28); -x_30 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__7; -lean_inc(x_24); +x_29 = lean_ctor_get(x_27, 1); lean_inc(x_29); -x_31 = l_Lean_addMacroScope(x_29, x_30, x_24); -x_32 = lean_box(0); -x_33 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__1; -x_34 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__4; -lean_inc(x_22); -x_35 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_35, 0, x_22); -lean_ctor_set(x_35, 1, x_33); -lean_ctor_set(x_35, 2, x_31); -lean_ctor_set(x_35, 3, x_34); -x_36 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -x_37 = l_Lean_addMacroScope(x_29, x_36, x_24); -x_38 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; -x_39 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_39, 0, x_22); -lean_ctor_set(x_39, 1, x_38); -lean_ctor_set(x_39, 2, x_37); -lean_ctor_set(x_39, 3, x_32); -lean_inc(x_3); -x_40 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_32, x_3); -x_41 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_42 = lean_array_push(x_41, x_39); -x_43 = lean_array_push(x_41, x_35); -if (lean_obj_tag(x_40) == 0) +lean_dec(x_27); +x_30 = lean_ctor_get(x_28, 0); +lean_inc(x_30); +lean_dec(x_28); +x_31 = lean_environment_main_module(x_30); +x_32 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__7; +lean_inc(x_26); +lean_inc(x_31); +x_33 = l_Lean_addMacroScope(x_31, x_32, x_26); +x_34 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__2; +lean_inc(x_1); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_1); +lean_inc(x_1); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_1); +x_37 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__1; +lean_inc(x_24); +x_38 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_38, 0, x_24); +lean_ctor_set(x_38, 1, x_37); +lean_ctor_set(x_38, 2, x_33); +lean_ctor_set(x_38, 3, x_36); +x_39 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +x_40 = l_Lean_addMacroScope(x_31, x_39, x_26); +x_41 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; +lean_inc(x_1); +x_42 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_42, 0, x_24); +lean_ctor_set(x_42, 1, x_41); +lean_ctor_set(x_42, 2, x_40); +lean_ctor_set(x_42, 3, x_1); +lean_inc(x_5); +x_43 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_1, x_5); +x_44 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_45 = lean_array_push(x_44, x_42); +x_46 = lean_array_push(x_44, x_38); +if (lean_obj_tag(x_43) == 0) { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_44 = l___private_Init_Meta_0__Lean_quoteNameMk(x_3); -x_45 = lean_array_push(x_42, x_44); -x_46 = lean_box(2); -x_47 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_48 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_48, 0, x_46); -lean_ctor_set(x_48, 1, x_47); -lean_ctor_set(x_48, 2, x_45); -x_49 = lean_array_push(x_43, x_48); -x_50 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_47 = l_Lean_quoteNameMk(x_5); +x_48 = lean_array_push(x_45, x_47); +x_49 = lean_box(2); +x_50 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; x_51 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_51, 0, x_46); -lean_ctor_set(x_51, 1, x_50); -lean_ctor_set(x_51, 2, x_49); -x_52 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__9(x_2, x_14, x_51, x_7, x_8, x_9, x_10, x_11, x_12, x_27); -lean_dec(x_2); -return x_52; +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +lean_ctor_set(x_51, 2, x_48); +x_52 = lean_array_push(x_46, x_51); +x_53 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_54 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_54, 0, x_49); +lean_ctor_set(x_54, 1, x_53); +lean_ctor_set(x_54, 2, x_52); +x_55 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__9(x_3, x_4, x_16, x_54, x_9, x_10, x_11, x_12, x_13, x_14, x_29); +lean_dec(x_4); +lean_dec(x_3); +return x_55; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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_dec(x_5); +x_56 = lean_ctor_get(x_43, 0); +lean_inc(x_56); +lean_dec(x_43); +x_57 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; +x_58 = l_String_intercalate(x_57, x_56); +x_59 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4; +x_60 = lean_string_append(x_59, x_58); +lean_dec(x_58); +x_61 = lean_box(2); +x_62 = l_Lean_Syntax_mkNameLit(x_60, x_61); +x_63 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; +x_64 = lean_array_push(x_63, x_62); +x_65 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2; +x_66 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_66, 0, x_61); +lean_ctor_set(x_66, 1, x_65); +lean_ctor_set(x_66, 2, x_64); +x_67 = lean_array_push(x_45, x_66); +x_68 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_69 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_69, 0, x_61); +lean_ctor_set(x_69, 1, x_68); +lean_ctor_set(x_69, 2, x_67); +x_70 = lean_array_push(x_46, x_69); +x_71 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_72 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_72, 0, x_61); +lean_ctor_set(x_72, 1, x_71); +lean_ctor_set(x_72, 2, x_70); +x_73 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__9(x_3, x_4, x_16, x_72, x_9, x_10, x_11, x_12, x_13, x_14, x_29); +lean_dec(x_4); lean_dec(x_3); -x_53 = lean_ctor_get(x_40, 0); -lean_inc(x_53); -lean_dec(x_40); -x_54 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__3; -x_55 = l_String_intercalate(x_54, x_53); -x_56 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__4; -x_57 = lean_string_append(x_56, x_55); -lean_dec(x_55); -x_58 = l_Lean_nameLitKind; -x_59 = lean_box(2); -x_60 = l_Lean_Syntax_mkLit(x_58, x_57, x_59); -x_61 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; -x_62 = lean_array_push(x_61, x_60); -x_63 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__2; -x_64 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_64, 0, x_59); -lean_ctor_set(x_64, 1, x_63); -lean_ctor_set(x_64, 2, x_62); -x_65 = lean_array_push(x_42, x_64); -x_66 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_67 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_67, 0, x_59); -lean_ctor_set(x_67, 1, x_66); -lean_ctor_set(x_67, 2, x_65); -x_68 = lean_array_push(x_43, x_67); -x_69 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_70 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_70, 0, x_59); -lean_ctor_set(x_70, 1, x_69); -lean_ctor_set(x_70, 2, x_68); -x_71 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__9(x_2, x_14, x_70, x_7, x_8, x_9, x_10, x_11, x_12, x_27); -lean_dec(x_2); -return x_71; +return x_73; } } else { -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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; -lean_dec(x_3); -lean_inc(x_11); -x_72 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_11, x_12, x_13); -x_73 = lean_ctor_get(x_72, 0); -lean_inc(x_73); -x_74 = lean_ctor_get(x_72, 1); -lean_inc(x_74); -lean_dec(x_72); -x_75 = lean_ctor_get(x_11, 10); +lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; +lean_dec(x_5); +lean_inc(x_13); +x_74 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_13, x_14, x_15); +x_75 = lean_ctor_get(x_74, 0); lean_inc(x_75); -x_76 = lean_st_ref_get(x_12, x_74); -x_77 = lean_ctor_get(x_76, 0); +x_76 = lean_ctor_get(x_74, 1); +lean_inc(x_76); +lean_dec(x_74); +x_77 = lean_ctor_get(x_13, 10); lean_inc(x_77); -x_78 = lean_ctor_get(x_76, 1); -lean_inc(x_78); -lean_dec(x_76); -x_79 = lean_ctor_get(x_77, 0); +x_78 = lean_st_ref_get(x_14, x_76); +x_79 = lean_ctor_get(x_78, 0); lean_inc(x_79); -lean_dec(x_77); -x_80 = lean_environment_main_module(x_79); -x_81 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__9; -lean_inc(x_75); +x_80 = lean_ctor_get(x_78, 1); lean_inc(x_80); -x_82 = l_Lean_addMacroScope(x_80, x_81, x_75); -x_83 = lean_box(0); -x_84 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__7; -x_85 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__12; -lean_inc(x_73); -x_86 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_86, 0, x_73); -lean_ctor_set(x_86, 1, x_84); -lean_ctor_set(x_86, 2, x_82); -lean_ctor_set(x_86, 3, x_85); -x_87 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -x_88 = l_Lean_addMacroScope(x_80, x_87, x_75); -x_89 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; -x_90 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_90, 0, x_73); -lean_ctor_set(x_90, 1, x_89); -lean_ctor_set(x_90, 2, x_88); -lean_ctor_set(x_90, 3, x_83); -x_91 = l_Lean_Syntax_getId(x_4); -lean_dec(x_4); -lean_inc(x_91); -x_92 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_83, x_91); -x_93 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_94 = lean_array_push(x_93, x_90); -x_95 = lean_array_push(x_93, x_86); -if (lean_obj_tag(x_92) == 0) +lean_dec(x_78); +x_81 = lean_ctor_get(x_79, 0); +lean_inc(x_81); +lean_dec(x_79); +x_82 = lean_environment_main_module(x_81); +x_83 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__7; +lean_inc(x_77); +lean_inc(x_82); +x_84 = l_Lean_addMacroScope(x_82, x_83, x_77); +x_85 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__8; +lean_inc(x_1); +x_86 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_86, 0, x_85); +lean_ctor_set(x_86, 1, x_1); +lean_inc(x_1); +x_87 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_87, 0, x_86); +lean_ctor_set(x_87, 1, x_1); +x_88 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__5; +lean_inc(x_75); +x_89 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_89, 0, x_75); +lean_ctor_set(x_89, 1, x_88); +lean_ctor_set(x_89, 2, x_84); +lean_ctor_set(x_89, 3, x_87); +x_90 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +x_91 = l_Lean_addMacroScope(x_82, x_90, x_77); +x_92 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; +lean_inc(x_1); +x_93 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_93, 0, x_75); +lean_ctor_set(x_93, 1, x_92); +lean_ctor_set(x_93, 2, x_91); +lean_ctor_set(x_93, 3, x_1); +x_94 = l_Lean_Syntax_getId(x_6); +lean_dec(x_6); +lean_inc(x_94); +x_95 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_1, x_94); +x_96 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_97 = lean_array_push(x_96, x_93); +x_98 = lean_array_push(x_96, x_89); +if (lean_obj_tag(x_95) == 0) { -lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; -x_96 = l___private_Init_Meta_0__Lean_quoteNameMk(x_91); -x_97 = lean_array_push(x_94, x_96); -x_98 = lean_box(2); -x_99 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_100 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_100, 0, x_98); -lean_ctor_set(x_100, 1, x_99); -lean_ctor_set(x_100, 2, x_97); -x_101 = lean_array_push(x_95, x_100); -x_102 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_99 = l_Lean_quoteNameMk(x_94); +x_100 = lean_array_push(x_97, x_99); +x_101 = lean_box(2); +x_102 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; x_103 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_103, 0, x_98); +lean_ctor_set(x_103, 0, x_101); lean_ctor_set(x_103, 1, x_102); -lean_ctor_set(x_103, 2, x_101); -x_104 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__9(x_2, x_14, x_103, x_7, x_8, x_9, x_10, x_11, x_12, x_78); -lean_dec(x_2); -return x_104; +lean_ctor_set(x_103, 2, x_100); +x_104 = lean_array_push(x_98, x_103); +x_105 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_106 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_106, 0, x_101); +lean_ctor_set(x_106, 1, x_105); +lean_ctor_set(x_106, 2, x_104); +x_107 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__9(x_3, x_4, x_16, x_106, x_9, x_10, x_11, x_12, x_13, x_14, x_80); +lean_dec(x_4); +lean_dec(x_3); +return x_107; } else { -lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; -lean_dec(x_91); -x_105 = lean_ctor_get(x_92, 0); -lean_inc(x_105); -lean_dec(x_92); -x_106 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__3; -x_107 = l_String_intercalate(x_106, x_105); -x_108 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__4; -x_109 = lean_string_append(x_108, x_107); -lean_dec(x_107); -x_110 = l_Lean_nameLitKind; -x_111 = lean_box(2); -x_112 = l_Lean_Syntax_mkLit(x_110, x_109, x_111); -x_113 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; -x_114 = lean_array_push(x_113, x_112); -x_115 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__2; -x_116 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_116, 0, x_111); -lean_ctor_set(x_116, 1, x_115); -lean_ctor_set(x_116, 2, x_114); -x_117 = lean_array_push(x_94, x_116); -x_118 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_119 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_119, 0, x_111); -lean_ctor_set(x_119, 1, x_118); -lean_ctor_set(x_119, 2, x_117); -x_120 = lean_array_push(x_95, x_119); -x_121 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_122 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_122, 0, x_111); -lean_ctor_set(x_122, 1, x_121); -lean_ctor_set(x_122, 2, x_120); -x_123 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__9(x_2, x_14, x_122, x_7, x_8, x_9, x_10, x_11, x_12, x_78); -lean_dec(x_2); -return x_123; +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; +lean_dec(x_94); +x_108 = lean_ctor_get(x_95, 0); +lean_inc(x_108); +lean_dec(x_95); +x_109 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; +x_110 = l_String_intercalate(x_109, x_108); +x_111 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4; +x_112 = lean_string_append(x_111, x_110); +lean_dec(x_110); +x_113 = lean_box(2); +x_114 = l_Lean_Syntax_mkNameLit(x_112, x_113); +x_115 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; +x_116 = lean_array_push(x_115, x_114); +x_117 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2; +x_118 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_118, 0, x_113); +lean_ctor_set(x_118, 1, x_117); +lean_ctor_set(x_118, 2, x_116); +x_119 = lean_array_push(x_97, x_118); +x_120 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_121 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_121, 0, x_113); +lean_ctor_set(x_121, 1, x_120); +lean_ctor_set(x_121, 2, x_119); +x_122 = lean_array_push(x_98, x_121); +x_123 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_124 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_124, 0, x_113); +lean_ctor_set(x_124, 1, x_123); +lean_ctor_set(x_124, 2, x_122); +x_125 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__9(x_3, x_4, x_16, x_124, x_9, x_10, x_11, x_12, x_13, x_14, x_80); +lean_dec(x_4); +lean_dec(x_3); +return x_125; } } } else { -lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; -lean_dec(x_16); -lean_dec(x_4); -lean_dec(x_3); -lean_inc(x_11); -x_124 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_11, x_12, x_13); -x_125 = lean_ctor_get(x_124, 0); -lean_inc(x_125); -x_126 = lean_ctor_get(x_124, 1); -lean_inc(x_126); -lean_dec(x_124); -x_127 = lean_ctor_get(x_11, 10); +lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; +lean_dec(x_18); +lean_dec(x_6); +lean_dec(x_5); +lean_inc(x_13); +x_126 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_13, x_14, x_15); +x_127 = lean_ctor_get(x_126, 0); lean_inc(x_127); -x_128 = lean_st_ref_get(x_12, x_126); -x_129 = lean_ctor_get(x_128, 0); +x_128 = lean_ctor_get(x_126, 1); +lean_inc(x_128); +lean_dec(x_126); +x_129 = lean_ctor_get(x_13, 10); lean_inc(x_129); -x_130 = lean_ctor_get(x_128, 1); -lean_inc(x_130); -lean_dec(x_128); -x_131 = lean_ctor_get(x_129, 0); +x_130 = lean_st_ref_get(x_14, x_128); +x_131 = lean_ctor_get(x_130, 0); lean_inc(x_131); -lean_dec(x_129); -x_132 = lean_environment_main_module(x_131); -x_133 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__17; -lean_inc(x_127); +x_132 = lean_ctor_get(x_130, 1); lean_inc(x_132); -x_134 = l_Lean_addMacroScope(x_132, x_133, x_127); -x_135 = lean_box(0); -x_136 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__15; -x_137 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__20; -lean_inc(x_125); -x_138 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_138, 0, x_125); -lean_ctor_set(x_138, 1, x_136); -lean_ctor_set(x_138, 2, x_134); -lean_ctor_set(x_138, 3, x_137); -x_139 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -x_140 = l_Lean_addMacroScope(x_132, x_139, x_127); -x_141 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; -x_142 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_142, 0, x_125); -lean_ctor_set(x_142, 1, x_141); -lean_ctor_set(x_142, 2, x_140); -lean_ctor_set(x_142, 3, x_135); -x_143 = lean_array_get_size(x_2); -x_144 = l_Nat_repr(x_143); -x_145 = l_Lean_numLitKind; -x_146 = lean_box(2); -x_147 = l_Lean_Syntax_mkLit(x_145, x_144, x_146); -x_148 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_149 = lean_array_push(x_148, x_142); -x_150 = lean_array_push(x_149, x_147); -x_151 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_152 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_152, 0, x_146); -lean_ctor_set(x_152, 1, x_151); -lean_ctor_set(x_152, 2, x_150); -x_153 = lean_array_push(x_148, x_138); -x_154 = lean_array_push(x_153, x_152); -x_155 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_156 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_156, 0, x_146); -lean_ctor_set(x_156, 1, x_155); -lean_ctor_set(x_156, 2, x_154); -x_157 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__9(x_2, x_14, x_156, x_7, x_8, x_9, x_10, x_11, x_12, x_130); -lean_dec(x_2); -return x_157; +lean_dec(x_130); +x_133 = lean_ctor_get(x_131, 0); +lean_inc(x_133); +lean_dec(x_131); +x_134 = lean_environment_main_module(x_133); +x_135 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__13; +lean_inc(x_129); +lean_inc(x_134); +x_136 = l_Lean_addMacroScope(x_134, x_135, x_129); +x_137 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__14; +lean_inc(x_1); +x_138 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_138, 0, x_137); +lean_ctor_set(x_138, 1, x_1); +lean_inc(x_1); +x_139 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_139, 0, x_138); +lean_ctor_set(x_139, 1, x_1); +x_140 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__11; +lean_inc(x_127); +x_141 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_141, 0, x_127); +lean_ctor_set(x_141, 1, x_140); +lean_ctor_set(x_141, 2, x_136); +lean_ctor_set(x_141, 3, x_139); +x_142 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +x_143 = l_Lean_addMacroScope(x_134, x_142, x_129); +x_144 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; +x_145 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_145, 0, x_127); +lean_ctor_set(x_145, 1, x_144); +lean_ctor_set(x_145, 2, x_143); +lean_ctor_set(x_145, 3, x_1); +x_146 = lean_array_get_size(x_3); +x_147 = l_Nat_repr(x_146); +x_148 = lean_box(2); +x_149 = l_Lean_Syntax_mkNumLit(x_147, x_148); +x_150 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_151 = lean_array_push(x_150, x_145); +x_152 = lean_array_push(x_151, x_149); +x_153 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_154 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_154, 0, x_148); +lean_ctor_set(x_154, 1, x_153); +lean_ctor_set(x_154, 2, x_152); +x_155 = lean_array_push(x_150, x_141); +x_156 = lean_array_push(x_155, x_154); +x_157 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_158 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_158, 0, x_148); +lean_ctor_set(x_158, 1, x_157); +lean_ctor_set(x_158, 2, x_156); +x_159 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__9(x_3, x_4, x_16, x_158, x_9, x_10, x_11, x_12, x_13, x_14, x_132); +lean_dec(x_4); +lean_dec(x_3); +return x_159; } } else { -lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; -lean_dec(x_15); -lean_dec(x_4); -lean_inc(x_11); -x_158 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_11, x_12, x_13); -x_159 = lean_ctor_get(x_158, 0); -lean_inc(x_159); -x_160 = lean_ctor_get(x_158, 1); -lean_inc(x_160); -lean_dec(x_158); -x_161 = lean_ctor_get(x_11, 10); +lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; +lean_dec(x_17); +lean_dec(x_6); +lean_inc(x_13); +x_160 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_13, x_14, x_15); +x_161 = lean_ctor_get(x_160, 0); lean_inc(x_161); -x_162 = lean_st_ref_get(x_12, x_160); -x_163 = lean_ctor_get(x_162, 0); +x_162 = lean_ctor_get(x_160, 1); +lean_inc(x_162); +lean_dec(x_160); +x_163 = lean_ctor_get(x_13, 10); lean_inc(x_163); -x_164 = lean_ctor_get(x_162, 1); -lean_inc(x_164); -lean_dec(x_162); -x_165 = lean_ctor_get(x_163, 0); +x_164 = lean_st_ref_get(x_14, x_162); +x_165 = lean_ctor_get(x_164, 0); lean_inc(x_165); -lean_dec(x_163); -x_166 = lean_environment_main_module(x_165); -x_167 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__7; -lean_inc(x_161); +x_166 = lean_ctor_get(x_164, 1); lean_inc(x_166); -x_168 = l_Lean_addMacroScope(x_166, x_167, x_161); -x_169 = lean_box(0); -x_170 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__1; -x_171 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__4; -lean_inc(x_159); -x_172 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_172, 0, x_159); -lean_ctor_set(x_172, 1, x_170); -lean_ctor_set(x_172, 2, x_168); -lean_ctor_set(x_172, 3, x_171); -x_173 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -x_174 = l_Lean_addMacroScope(x_166, x_173, x_161); -x_175 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; -x_176 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_176, 0, x_159); -lean_ctor_set(x_176, 1, x_175); -lean_ctor_set(x_176, 2, x_174); -lean_ctor_set(x_176, 3, x_169); -lean_inc(x_3); -x_177 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_169, x_3); -x_178 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_179 = lean_array_push(x_178, x_176); -x_180 = lean_array_push(x_178, x_172); -if (lean_obj_tag(x_177) == 0) -{ -lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; -x_181 = l___private_Init_Meta_0__Lean_quoteNameMk(x_3); -x_182 = lean_array_push(x_179, x_181); -x_183 = lean_box(2); -x_184 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_185 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_185, 0, x_183); -lean_ctor_set(x_185, 1, x_184); -lean_ctor_set(x_185, 2, x_182); -x_186 = lean_array_push(x_180, x_185); -x_187 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +lean_dec(x_164); +x_167 = lean_ctor_get(x_165, 0); +lean_inc(x_167); +lean_dec(x_165); +x_168 = lean_environment_main_module(x_167); +x_169 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__7; +lean_inc(x_163); +lean_inc(x_168); +x_170 = l_Lean_addMacroScope(x_168, x_169, x_163); +x_171 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__2; +lean_inc(x_1); +x_172 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_172, 0, x_171); +lean_ctor_set(x_172, 1, x_1); +lean_inc(x_1); +x_173 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_173, 0, x_172); +lean_ctor_set(x_173, 1, x_1); +x_174 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__1; +lean_inc(x_161); +x_175 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_175, 0, x_161); +lean_ctor_set(x_175, 1, x_174); +lean_ctor_set(x_175, 2, x_170); +lean_ctor_set(x_175, 3, x_173); +x_176 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +x_177 = l_Lean_addMacroScope(x_168, x_176, x_163); +x_178 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; +lean_inc(x_1); +x_179 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_179, 0, x_161); +lean_ctor_set(x_179, 1, x_178); +lean_ctor_set(x_179, 2, x_177); +lean_ctor_set(x_179, 3, x_1); +lean_inc(x_5); +x_180 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_1, x_5); +x_181 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_182 = lean_array_push(x_181, x_179); +x_183 = lean_array_push(x_181, x_175); +if (lean_obj_tag(x_180) == 0) +{ +lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; +x_184 = l_Lean_quoteNameMk(x_5); +x_185 = lean_array_push(x_182, x_184); +x_186 = lean_box(2); +x_187 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; x_188 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_188, 0, x_183); +lean_ctor_set(x_188, 0, x_186); lean_ctor_set(x_188, 1, x_187); -lean_ctor_set(x_188, 2, x_186); -x_189 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__9(x_2, x_14, x_188, x_7, x_8, x_9, x_10, x_11, x_12, x_164); -lean_dec(x_2); -return x_189; +lean_ctor_set(x_188, 2, x_185); +x_189 = lean_array_push(x_183, x_188); +x_190 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_191 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_191, 0, x_186); +lean_ctor_set(x_191, 1, x_190); +lean_ctor_set(x_191, 2, x_189); +x_192 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__9(x_3, x_4, x_16, x_191, x_9, x_10, x_11, x_12, x_13, x_14, x_166); +lean_dec(x_4); +lean_dec(x_3); +return x_192; } else { -lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; +lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; +lean_dec(x_5); +x_193 = lean_ctor_get(x_180, 0); +lean_inc(x_193); +lean_dec(x_180); +x_194 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; +x_195 = l_String_intercalate(x_194, x_193); +x_196 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4; +x_197 = lean_string_append(x_196, x_195); +lean_dec(x_195); +x_198 = lean_box(2); +x_199 = l_Lean_Syntax_mkNameLit(x_197, x_198); +x_200 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; +x_201 = lean_array_push(x_200, x_199); +x_202 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2; +x_203 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_203, 0, x_198); +lean_ctor_set(x_203, 1, x_202); +lean_ctor_set(x_203, 2, x_201); +x_204 = lean_array_push(x_182, x_203); +x_205 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_206 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_206, 0, x_198); +lean_ctor_set(x_206, 1, x_205); +lean_ctor_set(x_206, 2, x_204); +x_207 = lean_array_push(x_183, x_206); +x_208 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_209 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_209, 0, x_198); +lean_ctor_set(x_209, 1, x_208); +lean_ctor_set(x_209, 2, x_207); +x_210 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__9(x_3, x_4, x_16, x_209, x_9, x_10, x_11, x_12, x_13, x_14, x_166); +lean_dec(x_4); lean_dec(x_3); -x_190 = lean_ctor_get(x_177, 0); -lean_inc(x_190); -lean_dec(x_177); -x_191 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__3; -x_192 = l_String_intercalate(x_191, x_190); -x_193 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__4; -x_194 = lean_string_append(x_193, x_192); -lean_dec(x_192); -x_195 = l_Lean_nameLitKind; -x_196 = lean_box(2); -x_197 = l_Lean_Syntax_mkLit(x_195, x_194, x_196); -x_198 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; -x_199 = lean_array_push(x_198, x_197); -x_200 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__2; -x_201 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_201, 0, x_196); -lean_ctor_set(x_201, 1, x_200); -lean_ctor_set(x_201, 2, x_199); -x_202 = lean_array_push(x_179, x_201); -x_203 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_204 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_204, 0, x_196); -lean_ctor_set(x_204, 1, x_203); -lean_ctor_set(x_204, 2, x_202); -x_205 = lean_array_push(x_180, x_204); -x_206 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_207 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_207, 0, x_196); -lean_ctor_set(x_207, 1, x_206); -lean_ctor_set(x_207, 2, x_205); -x_208 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__9(x_2, x_14, x_207, x_7, x_8, x_9, x_10, x_11, x_12, x_164); -lean_dec(x_2); -return x_208; +return x_210; } } } else { -lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; -lean_dec(x_4); -lean_inc(x_11); -x_209 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_11, x_12, x_13); -x_210 = lean_ctor_get(x_209, 0); -lean_inc(x_210); -x_211 = lean_ctor_get(x_209, 1); -lean_inc(x_211); -lean_dec(x_209); -x_212 = lean_ctor_get(x_11, 10); +lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; +lean_dec(x_6); +lean_inc(x_13); +x_211 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_13, x_14, x_15); +x_212 = lean_ctor_get(x_211, 0); lean_inc(x_212); -x_213 = lean_st_ref_get(x_12, x_211); -x_214 = lean_ctor_get(x_213, 0); +x_213 = lean_ctor_get(x_211, 1); +lean_inc(x_213); +lean_dec(x_211); +x_214 = lean_ctor_get(x_13, 10); lean_inc(x_214); -x_215 = lean_ctor_get(x_213, 1); -lean_inc(x_215); -lean_dec(x_213); -x_216 = lean_ctor_get(x_214, 0); +x_215 = lean_st_ref_get(x_14, x_213); +x_216 = lean_ctor_get(x_215, 0); lean_inc(x_216); -lean_dec(x_214); -x_217 = lean_environment_main_module(x_216); -x_218 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__7; -lean_inc(x_212); +x_217 = lean_ctor_get(x_215, 1); lean_inc(x_217); -x_219 = l_Lean_addMacroScope(x_217, x_218, x_212); -x_220 = lean_box(0); -x_221 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__1; -x_222 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__4; -lean_inc(x_210); -x_223 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_223, 0, x_210); -lean_ctor_set(x_223, 1, x_221); -lean_ctor_set(x_223, 2, x_219); -lean_ctor_set(x_223, 3, x_222); -x_224 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -x_225 = l_Lean_addMacroScope(x_217, x_224, x_212); -x_226 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; -x_227 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_227, 0, x_210); -lean_ctor_set(x_227, 1, x_226); -lean_ctor_set(x_227, 2, x_225); -lean_ctor_set(x_227, 3, x_220); -lean_inc(x_3); -x_228 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_220, x_3); -x_229 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_230 = lean_array_push(x_229, x_227); -x_231 = lean_array_push(x_229, x_223); -if (lean_obj_tag(x_228) == 0) -{ -lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; -x_232 = l___private_Init_Meta_0__Lean_quoteNameMk(x_3); -x_233 = lean_array_push(x_230, x_232); -x_234 = lean_box(2); -x_235 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_236 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_236, 0, x_234); -lean_ctor_set(x_236, 1, x_235); -lean_ctor_set(x_236, 2, x_233); -x_237 = lean_array_push(x_231, x_236); -x_238 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +lean_dec(x_215); +x_218 = lean_ctor_get(x_216, 0); +lean_inc(x_218); +lean_dec(x_216); +x_219 = lean_environment_main_module(x_218); +x_220 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___closed__7; +lean_inc(x_214); +lean_inc(x_219); +x_221 = l_Lean_addMacroScope(x_219, x_220, x_214); +x_222 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__2; +lean_inc(x_1); +x_223 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_223, 0, x_222); +lean_ctor_set(x_223, 1, x_1); +lean_inc(x_1); +x_224 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_224, 0, x_223); +lean_ctor_set(x_224, 1, x_1); +x_225 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__1; +lean_inc(x_212); +x_226 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_226, 0, x_212); +lean_ctor_set(x_226, 1, x_225); +lean_ctor_set(x_226, 2, x_221); +lean_ctor_set(x_226, 3, x_224); +x_227 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +x_228 = l_Lean_addMacroScope(x_219, x_227, x_214); +x_229 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; +lean_inc(x_1); +x_230 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_230, 0, x_212); +lean_ctor_set(x_230, 1, x_229); +lean_ctor_set(x_230, 2, x_228); +lean_ctor_set(x_230, 3, x_1); +lean_inc(x_5); +x_231 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_1, x_5); +x_232 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_233 = lean_array_push(x_232, x_230); +x_234 = lean_array_push(x_232, x_226); +if (lean_obj_tag(x_231) == 0) +{ +lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; +x_235 = l_Lean_quoteNameMk(x_5); +x_236 = lean_array_push(x_233, x_235); +x_237 = lean_box(2); +x_238 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; x_239 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_239, 0, x_234); +lean_ctor_set(x_239, 0, x_237); lean_ctor_set(x_239, 1, x_238); -lean_ctor_set(x_239, 2, x_237); -x_240 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__9(x_2, x_14, x_239, x_7, x_8, x_9, x_10, x_11, x_12, x_215); -lean_dec(x_2); -return x_240; +lean_ctor_set(x_239, 2, x_236); +x_240 = lean_array_push(x_234, x_239); +x_241 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_242 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_242, 0, x_237); +lean_ctor_set(x_242, 1, x_241); +lean_ctor_set(x_242, 2, x_240); +x_243 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__9(x_3, x_4, x_16, x_242, x_9, x_10, x_11, x_12, x_13, x_14, x_217); +lean_dec(x_4); +lean_dec(x_3); +return x_243; } else { -lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; +lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; +lean_dec(x_5); +x_244 = lean_ctor_get(x_231, 0); +lean_inc(x_244); +lean_dec(x_231); +x_245 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; +x_246 = l_String_intercalate(x_245, x_244); +x_247 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4; +x_248 = lean_string_append(x_247, x_246); +lean_dec(x_246); +x_249 = lean_box(2); +x_250 = l_Lean_Syntax_mkNameLit(x_248, x_249); +x_251 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; +x_252 = lean_array_push(x_251, x_250); +x_253 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2; +x_254 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_254, 0, x_249); +lean_ctor_set(x_254, 1, x_253); +lean_ctor_set(x_254, 2, x_252); +x_255 = lean_array_push(x_233, x_254); +x_256 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_257 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_257, 0, x_249); +lean_ctor_set(x_257, 1, x_256); +lean_ctor_set(x_257, 2, x_255); +x_258 = lean_array_push(x_234, x_257); +x_259 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_260 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_260, 0, x_249); +lean_ctor_set(x_260, 1, x_259); +lean_ctor_set(x_260, 2, x_258); +x_261 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__9(x_3, x_4, x_16, x_260, x_9, x_10, x_11, x_12, x_13, x_14, x_217); +lean_dec(x_4); lean_dec(x_3); -x_241 = lean_ctor_get(x_228, 0); -lean_inc(x_241); -lean_dec(x_228); -x_242 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__3; -x_243 = l_String_intercalate(x_242, x_241); -x_244 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__4; -x_245 = lean_string_append(x_244, x_243); -lean_dec(x_243); -x_246 = l_Lean_nameLitKind; -x_247 = lean_box(2); -x_248 = l_Lean_Syntax_mkLit(x_246, x_245, x_247); -x_249 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; -x_250 = lean_array_push(x_249, x_248); -x_251 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__2; -x_252 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_252, 0, x_247); -lean_ctor_set(x_252, 1, x_251); -lean_ctor_set(x_252, 2, x_250); -x_253 = lean_array_push(x_230, x_252); -x_254 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_255 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_255, 0, x_247); -lean_ctor_set(x_255, 1, x_254); -lean_ctor_set(x_255, 2, x_253); -x_256 = lean_array_push(x_231, x_255); -x_257 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_258 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_258, 0, x_247); -lean_ctor_set(x_258, 1, x_257); -lean_ctor_set(x_258, 2, x_256); -x_259 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__9(x_2, x_14, x_258, x_7, x_8, x_9, x_10, x_11, x_12, x_215); -lean_dec(x_2); -return x_259; +return x_261; } } } else { -lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; -lean_dec(x_2); -lean_inc(x_11); -x_260 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_11, x_12, x_13); -x_261 = lean_ctor_get(x_260, 0); -lean_inc(x_261); -x_262 = lean_ctor_get(x_260, 1); -lean_inc(x_262); -lean_dec(x_260); -x_263 = lean_ctor_get(x_11, 10); +lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; +lean_dec(x_4); +lean_dec(x_3); +lean_inc(x_13); +x_262 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_13, x_14, x_15); +x_263 = lean_ctor_get(x_262, 0); lean_inc(x_263); -x_264 = lean_st_ref_get(x_12, x_262); -x_265 = lean_ctor_get(x_264, 0); +x_264 = lean_ctor_get(x_262, 1); +lean_inc(x_264); +lean_dec(x_262); +x_265 = lean_ctor_get(x_13, 10); lean_inc(x_265); -x_266 = lean_ctor_get(x_264, 1); -lean_inc(x_266); -lean_dec(x_264); -x_267 = lean_ctor_get(x_265, 0); +x_266 = lean_st_ref_get(x_14, x_264); +x_267 = lean_ctor_get(x_266, 0); lean_inc(x_267); -lean_dec(x_265); -x_268 = lean_environment_main_module(x_267); -x_269 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__25; -lean_inc(x_263); +x_268 = lean_ctor_get(x_266, 1); lean_inc(x_268); -x_270 = l_Lean_addMacroScope(x_268, x_269, x_263); -x_271 = lean_box(0); -x_272 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__23; -x_273 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__28; -lean_inc(x_261); -x_274 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_274, 0, x_261); -lean_ctor_set(x_274, 1, x_272); -lean_ctor_set(x_274, 2, x_270); -lean_ctor_set(x_274, 3, x_273); -x_275 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -x_276 = l_Lean_addMacroScope(x_268, x_275, x_263); -x_277 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; -x_278 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_278, 0, x_261); -lean_ctor_set(x_278, 1, x_277); -lean_ctor_set(x_278, 2, x_276); -lean_ctor_set(x_278, 3, x_271); -lean_inc(x_3); -x_279 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_271, x_3); -x_280 = l_Lean_Syntax_isLit_x3f(x_3, x_4); -lean_dec(x_4); -x_281 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; -x_282 = lean_array_push(x_281, x_278); -x_283 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_284 = lean_array_push(x_283, x_274); -if (lean_obj_tag(x_279) == 0) -{ -lean_object* x_285; lean_object* x_286; -x_285 = l___private_Init_Meta_0__Lean_quoteNameMk(x_3); -x_286 = lean_array_push(x_282, x_285); -if (lean_obj_tag(x_280) == 0) -{ -lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; -x_287 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__32; -x_288 = l_panic___at_Lean_Name_getString_x21___spec__1(x_287); -x_289 = lean_box(2); -x_290 = l_Lean_Syntax_mkStrLit(x_288, x_289); -lean_dec(x_288); -x_291 = lean_array_push(x_286, x_290); -x_292 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_293 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_293, 0, x_289); -lean_ctor_set(x_293, 1, x_292); -lean_ctor_set(x_293, 2, x_291); -x_294 = lean_array_push(x_284, x_293); -x_295 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +lean_dec(x_266); +x_269 = lean_ctor_get(x_267, 0); +lean_inc(x_269); +lean_dec(x_267); +x_270 = lean_environment_main_module(x_269); +x_271 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__19; +lean_inc(x_265); +lean_inc(x_270); +x_272 = l_Lean_addMacroScope(x_270, x_271, x_265); +x_273 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__20; +lean_inc(x_1); +x_274 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_274, 0, x_273); +lean_ctor_set(x_274, 1, x_1); +lean_inc(x_1); +x_275 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_275, 0, x_274); +lean_ctor_set(x_275, 1, x_1); +x_276 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__17; +lean_inc(x_263); +x_277 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_277, 0, x_263); +lean_ctor_set(x_277, 1, x_276); +lean_ctor_set(x_277, 2, x_272); +lean_ctor_set(x_277, 3, x_275); +x_278 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +x_279 = l_Lean_addMacroScope(x_270, x_278, x_265); +x_280 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; +lean_inc(x_1); +x_281 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_281, 0, x_263); +lean_ctor_set(x_281, 1, x_280); +lean_ctor_set(x_281, 2, x_279); +lean_ctor_set(x_281, 3, x_1); +lean_inc(x_5); +lean_inc(x_1); +x_282 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_1, x_5); +x_283 = l_Lean_Syntax_isLit_x3f(x_5, x_6); +lean_dec(x_6); +x_284 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; +x_285 = lean_array_push(x_284, x_281); +x_286 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_287 = lean_array_push(x_286, x_277); +if (lean_obj_tag(x_282) == 0) +{ +lean_object* x_314; +x_314 = l_Lean_quoteNameMk(x_5); +x_288 = x_314; +goto block_313; +} +else +{ +lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; +lean_dec(x_5); +x_315 = lean_ctor_get(x_282, 0); +lean_inc(x_315); +lean_dec(x_282); +x_316 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; +x_317 = l_String_intercalate(x_316, x_315); +x_318 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4; +x_319 = lean_string_append(x_318, x_317); +lean_dec(x_317); +x_320 = lean_box(2); +x_321 = l_Lean_Syntax_mkNameLit(x_319, x_320); +x_322 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; +x_323 = lean_array_push(x_322, x_321); +x_324 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2; +x_325 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_325, 0, x_320); +lean_ctor_set(x_325, 1, x_324); +lean_ctor_set(x_325, 2, x_323); +x_288 = x_325; +goto block_313; +} +block_313: +{ +lean_object* x_289; +x_289 = lean_array_push(x_285, x_288); +if (lean_obj_tag(x_283) == 0) +{ +lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; +x_290 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__24; +x_291 = l_panic___at_Lean_TSyntax_getString___spec__1(x_290); +x_292 = lean_box(2); +x_293 = l_Lean_Syntax_mkStrLit(x_291, x_292); +lean_dec(x_291); +x_294 = lean_array_push(x_289, x_293); +x_295 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; x_296 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_296, 0, x_289); +lean_ctor_set(x_296, 0, x_292); lean_ctor_set(x_296, 1, x_295); lean_ctor_set(x_296, 2, x_294); -x_297 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_297, 0, x_296); -lean_ctor_set(x_297, 1, x_271); -x_298 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8(x_5, x_6, x_297, x_7, x_8, x_9, x_10, x_11, x_12, x_266); -return x_298; +x_297 = lean_array_push(x_287, x_296); +x_298 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_299 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_299, 0, x_292); +lean_ctor_set(x_299, 1, x_298); +lean_ctor_set(x_299, 2, x_297); +lean_inc(x_1); +x_300 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_300, 0, x_299); +lean_ctor_set(x_300, 1, x_1); +x_301 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8(x_7, x_8, x_1, x_300, x_9, x_10, x_11, x_12, x_13, x_14, x_268); +return x_301; } else { -lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; -x_299 = lean_ctor_get(x_280, 0); -lean_inc(x_299); -lean_dec(x_280); -x_300 = lean_box(2); -x_301 = l_Lean_Syntax_mkStrLit(x_299, x_300); -lean_dec(x_299); -x_302 = lean_array_push(x_286, x_301); -x_303 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_304 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_304, 0, x_300); -lean_ctor_set(x_304, 1, x_303); -lean_ctor_set(x_304, 2, x_302); -x_305 = lean_array_push(x_284, x_304); -x_306 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; +x_302 = lean_ctor_get(x_283, 0); +lean_inc(x_302); +lean_dec(x_283); +x_303 = lean_box(2); +x_304 = l_Lean_Syntax_mkStrLit(x_302, x_303); +lean_dec(x_302); +x_305 = lean_array_push(x_289, x_304); +x_306 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; x_307 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_307, 0, x_300); +lean_ctor_set(x_307, 0, x_303); lean_ctor_set(x_307, 1, x_306); lean_ctor_set(x_307, 2, x_305); -x_308 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_308, 0, x_307); -lean_ctor_set(x_308, 1, x_271); -x_309 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8(x_5, x_6, x_308, x_7, x_8, x_9, x_10, x_11, x_12, x_266); -return x_309; -} -} -else -{ -lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; -lean_dec(x_3); -x_310 = lean_ctor_get(x_279, 0); -lean_inc(x_310); -lean_dec(x_279); -x_311 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__3; -x_312 = l_String_intercalate(x_311, x_310); -x_313 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__4; -x_314 = lean_string_append(x_313, x_312); -lean_dec(x_312); -x_315 = l_Lean_nameLitKind; -x_316 = lean_box(2); -x_317 = l_Lean_Syntax_mkLit(x_315, x_314, x_316); -x_318 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; -x_319 = lean_array_push(x_318, x_317); -x_320 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__2; -x_321 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_321, 0, x_316); -lean_ctor_set(x_321, 1, x_320); -lean_ctor_set(x_321, 2, x_319); -x_322 = lean_array_push(x_282, x_321); -if (lean_obj_tag(x_280) == 0) -{ -lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; -x_323 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__32; -x_324 = l_panic___at_Lean_Name_getString_x21___spec__1(x_323); -x_325 = l_Lean_Syntax_mkStrLit(x_324, x_316); -lean_dec(x_324); -x_326 = lean_array_push(x_322, x_325); -x_327 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_328 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_328, 0, x_316); -lean_ctor_set(x_328, 1, x_327); -lean_ctor_set(x_328, 2, x_326); -x_329 = lean_array_push(x_284, x_328); -x_330 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_331 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_331, 0, x_316); -lean_ctor_set(x_331, 1, x_330); -lean_ctor_set(x_331, 2, x_329); -x_332 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_332, 0, x_331); -lean_ctor_set(x_332, 1, x_271); -x_333 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8(x_5, x_6, x_332, x_7, x_8, x_9, x_10, x_11, x_12, x_266); -return x_333; -} -else -{ -lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; -x_334 = lean_ctor_get(x_280, 0); -lean_inc(x_334); -lean_dec(x_280); -x_335 = l_Lean_Syntax_mkStrLit(x_334, x_316); -lean_dec(x_334); -x_336 = lean_array_push(x_322, x_335); -x_337 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_338 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_338, 0, x_316); -lean_ctor_set(x_338, 1, x_337); -lean_ctor_set(x_338, 2, x_336); -x_339 = lean_array_push(x_284, x_338); -x_340 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_341 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_341, 0, x_316); -lean_ctor_set(x_341, 1, x_340); -lean_ctor_set(x_341, 2, x_339); -x_342 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_342, 0, x_341); -lean_ctor_set(x_342, 1, x_271); -x_343 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8(x_5, x_6, x_342, x_7, x_8, x_9, x_10, x_11, x_12, x_266); -return x_343; +x_308 = lean_array_push(x_287, x_307); +x_309 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_310 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_310, 0, x_303); +lean_ctor_set(x_310, 1, x_309); +lean_ctor_set(x_310, 2, x_308); +lean_inc(x_1); +x_311 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_311, 0, x_310); +lean_ctor_set(x_311, 1, x_1); +x_312 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8(x_7, x_8, x_1, x_311, x_9, x_10, x_11, x_12, x_13, x_14, x_268); +return x_312; } } } @@ -28969,7 +30561,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__1; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__1; x_2 = 1; x_3 = lean_alloc_ctor(0, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -28977,47 +30569,46 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__12(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6) { _start: { -switch (lean_obj_tag(x_5)) { +switch (lean_obj_tag(x_6)) { case 1: { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; -lean_dec(x_4); -x_6 = lean_ctor_get(x_5, 0); -lean_inc(x_6); -x_7 = lean_ctor_get(x_5, 1); -lean_inc(x_7); +lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; lean_dec(x_5); -x_8 = lean_box(0); +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_6, 1); +lean_inc(x_8); +lean_dec(x_6); x_9 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_9, 0, x_1); -lean_ctor_set(x_9, 1, x_8); -x_10 = l_List_beq___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__5(x_6, x_9); +lean_ctor_set(x_9, 1, x_2); +x_10 = l_List_beq___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__5(x_7, x_9); lean_dec(x_9); -lean_dec(x_6); +lean_dec(x_7); if (x_10 == 0) { lean_object* x_11; -lean_dec(x_7); -lean_dec(x_2); +lean_dec(x_8); +lean_dec(x_3); x_11 = lean_box(1); return x_11; } else { lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_12 = lean_array_get_size(x_2); +x_12 = lean_array_get_size(x_3); x_13 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_13, 0, x_12); -x_14 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1(x_7, x_13); +x_14 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1(x_8, x_13); lean_dec(x_13); -lean_dec(x_7); +lean_dec(x_8); if (x_14 == 0) { lean_object* x_15; -lean_dec(x_2); +lean_dec(x_3); x_15 = lean_box(1); return x_15; } @@ -29025,7 +30616,7 @@ else { lean_object* x_16; uint8_t x_17; lean_object* x_18; x_16 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__11___boxed), 9, 1); -lean_closure_set(x_16, 0, x_2); +lean_closure_set(x_16, 0, x_3); x_17 = 1; x_18 = lean_alloc_ctor(0, 1, 1); lean_ctor_set(x_18, 0, x_16); @@ -29036,23 +30627,24 @@ return x_18; } case 3: { +lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -if (x_3 == 0) +if (x_4 == 0) { lean_object* x_19; +lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); x_19 = lean_box(1); return x_19; } else { lean_object* x_20; uint8_t x_21; -x_20 = lean_ctor_get(x_5, 0); +x_20 = lean_ctor_get(x_6, 0); lean_inc(x_20); -lean_dec(x_5); -x_21 = l_Lean_Syntax_structEq(x_4, x_20); +lean_dec(x_6); +x_21 = l_Lean_Syntax_structEq(x_5, x_20); if (x_21 == 0) { lean_object* x_22; @@ -29070,8 +30662,9 @@ return x_23; default: { lean_object* x_24; +lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); x_24 = lean_box(1); @@ -29080,63 +30673,63 @@ return x_24; } } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__13(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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__13(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: { -switch (lean_obj_tag(x_6)) { +switch (lean_obj_tag(x_7)) { case 2: { -lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_7 = lean_ctor_get(x_6, 0); -x_8 = lean_ctor_get(x_6, 1); -x_9 = lean_nat_dec_eq(x_7, x_1); -if (x_9 == 0) +lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_8 = lean_ctor_get(x_7, 0); +x_9 = lean_ctor_get(x_7, 1); +x_10 = lean_nat_dec_eq(x_8, x_1); +if (x_10 == 0) { -lean_object* x_10; +lean_object* x_11; lean_dec(x_3); -x_10 = lean_box(1); -return x_10; +x_11 = lean_box(1); +return x_11; } else { -uint8_t x_11; -x_11 = lean_nat_dec_eq(x_8, x_2); -if (x_11 == 0) +uint8_t x_12; +x_12 = lean_nat_dec_eq(x_9, x_2); +if (x_12 == 0) { -lean_object* x_12; +lean_object* x_13; lean_dec(x_3); -x_12 = lean_box(1); -return x_12; +x_13 = lean_box(1); +return x_13; } else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; -x_13 = lean_mk_empty_array_with_capacity(x_3); -x_14 = lean_unsigned_to_nat(0u); -x_15 = l_Array_mapIdxM_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__6(x_4, x_1, x_5, x_3, x_14, lean_box(0), x_13); -x_16 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__11___boxed), 9, 1); -lean_closure_set(x_16, 0, x_15); -x_17 = 1; -x_18 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_18, 0, x_16); -lean_ctor_set_uint8(x_18, sizeof(void*)*1, x_17); -return x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; +x_14 = lean_mk_empty_array_with_capacity(x_3); +x_15 = lean_unsigned_to_nat(0u); +x_16 = l_Array_mapIdxM_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__6(x_4, x_5, x_1, x_6, x_3, x_15, lean_box(0), x_14); +x_17 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__11___boxed), 9, 1); +lean_closure_set(x_17, 0, x_16); +x_18 = 1; +x_19 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set_uint8(x_19, sizeof(void*)*1, x_18); +return x_19; } } } case 3: { -lean_object* x_19; +lean_object* x_20; lean_dec(x_3); -x_19 = lean_box(2); -return x_19; +x_20 = lean_box(2); +return x_20; } default: { -lean_object* x_20; +lean_object* x_21; lean_dec(x_3); -x_20 = lean_box(1); -return x_20; +x_21 = lean_box(1); +return x_21; } } } @@ -29212,7 +30805,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("term_-_", 7); +x_1 = lean_mk_string_from_bytes("num", 3); return x_1; } } @@ -29230,26 +30823,44 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("discr.getNumArgs", 16); +x_1 = lean_mk_string_from_bytes("term_-_", 7); return x_1; } } static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__11() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__10; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("discr.getNumArgs", 16); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__13() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__10; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__12; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__12() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__10; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__12; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__11; +x_3 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__13; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -29257,7 +30868,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__13() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__15() { _start: { lean_object* x_1; @@ -29265,37 +30876,13 @@ x_1 = lean_mk_string_from_bytes("getNumArgs", 10); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__14() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__13; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__15() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__4; -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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); +x_1 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__15; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); +x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } @@ -29366,476 +30953,492 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__24() { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__23; -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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__25() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__24; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -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; 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; +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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_inc(x_1); -x_13 = l_List_range(x_1); -lean_inc(x_10); -x_14 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4(x_13, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); -lean_inc(x_16); -lean_dec(x_14); -lean_inc(x_10); -x_17 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_10, x_11, x_16); -x_18 = lean_ctor_get(x_17, 0); +x_15 = l_List_range(x_1); +lean_inc(x_12); +x_16 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4(x_2, x_15, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); -lean_inc(x_19); -lean_dec(x_17); -x_20 = lean_ctor_get(x_10, 10); +lean_dec(x_16); +lean_inc(x_12); +x_19 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_12, x_13, x_18); +x_20 = lean_ctor_get(x_19, 0); lean_inc(x_20); -x_21 = lean_st_ref_get(x_11, x_19); -x_22 = lean_ctor_get(x_21, 0); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = lean_ctor_get(x_12, 10); lean_inc(x_22); -x_23 = lean_ctor_get(x_21, 1); -lean_inc(x_23); -lean_dec(x_21); -x_24 = lean_ctor_get(x_22, 0); +x_23 = lean_st_ref_get(x_13, x_21); +x_24 = lean_ctor_get(x_23, 0); lean_inc(x_24); -lean_dec(x_22); -x_25 = lean_environment_main_module(x_24); -x_26 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__2___closed__4; -lean_inc(x_20); +x_25 = lean_ctor_get(x_23, 1); lean_inc(x_25); -x_27 = l_Lean_addMacroScope(x_25, x_26, x_20); -x_28 = lean_box(0); -x_29 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__2___closed__3; -x_30 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__2___closed__7; -lean_inc(x_18); -x_31 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_31, 0, x_18); -lean_ctor_set(x_31, 1, x_29); -lean_ctor_set(x_31, 2, x_27); -lean_ctor_set(x_31, 3, x_30); -x_32 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; -lean_inc(x_18); -x_33 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_33, 0, x_18); -lean_ctor_set(x_33, 1, x_32); -x_34 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__7; +lean_dec(x_23); +x_26 = lean_ctor_get(x_24, 0); +lean_inc(x_26); +lean_dec(x_24); +x_27 = lean_environment_main_module(x_26); +x_28 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__2___closed__4; +lean_inc(x_22); +lean_inc(x_27); +x_29 = l_Lean_addMacroScope(x_27, x_28, x_22); +x_30 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__2___closed__5; +lean_inc(x_3); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_3); +lean_inc(x_3); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_3); +x_33 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__2___closed__3; lean_inc(x_20); -lean_inc(x_25); -x_35 = l_Lean_addMacroScope(x_25, x_34, x_20); -x_36 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__3; -lean_inc(x_18); -x_37 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_37, 0, x_18); -lean_ctor_set(x_37, 1, x_36); -lean_ctor_set(x_37, 2, x_35); -lean_ctor_set(x_37, 3, x_28); -x_38 = l_Nat_repr(x_1); -x_39 = l_Lean_numLitKind; -x_40 = lean_box(2); -x_41 = l_Lean_Syntax_mkLit(x_39, x_38, x_40); -x_42 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__14; +x_34 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_34, 0, x_20); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_29); +lean_ctor_set(x_34, 3, x_32); +x_35 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; lean_inc(x_20); -x_43 = l_Lean_addMacroScope(x_25, x_42, x_20); -x_44 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__12; -lean_inc(x_18); -x_45 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_45, 0, x_18); -lean_ctor_set(x_45, 1, x_44); -lean_ctor_set(x_45, 2, x_43); -lean_ctor_set(x_45, 3, x_28); -x_46 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__7___closed__1; -lean_inc(x_18); -x_47 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_47, 0, x_18); +x_36 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_36, 0, x_20); +lean_ctor_set(x_36, 1, x_35); +x_37 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__7; +lean_inc(x_22); +lean_inc(x_27); +x_38 = l_Lean_addMacroScope(x_27, x_37, x_22); +x_39 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__3; +lean_inc(x_3); +lean_inc(x_20); +x_40 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_40, 0, x_20); +lean_ctor_set(x_40, 1, x_39); +lean_ctor_set(x_40, 2, x_38); +lean_ctor_set(x_40, 3, x_3); +x_41 = l_Nat_repr(x_1); +x_42 = lean_box(2); +x_43 = l_Lean_Syntax_mkNumLit(x_41, x_42); +x_44 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__16; +lean_inc(x_22); +x_45 = l_Lean_addMacroScope(x_27, x_44, x_22); +x_46 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__14; +lean_inc(x_3); +lean_inc(x_20); +x_47 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_47, 0, x_20); lean_ctor_set(x_47, 1, x_46); -lean_inc(x_2); -x_48 = l_Nat_repr(x_2); -x_49 = l_Lean_Syntax_mkLit(x_39, x_48, x_40); -x_50 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; -x_51 = lean_array_push(x_50, x_45); -x_52 = lean_array_push(x_51, x_47); -x_53 = lean_array_push(x_52, x_49); -x_54 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__9; -x_55 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_55, 0, x_40); -lean_ctor_set(x_55, 1, x_54); -lean_ctor_set(x_55, 2, x_53); -x_56 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_57 = lean_array_push(x_56, x_55); -x_58 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; -x_59 = lean_array_push(x_57, x_58); -x_60 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_61 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_61, 0, x_40); -lean_ctor_set(x_61, 1, x_60); -lean_ctor_set(x_61, 2, x_59); -x_62 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; -x_63 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_63, 0, x_18); +lean_ctor_set(x_47, 2, x_45); +lean_ctor_set(x_47, 3, x_3); +x_48 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__7___closed__1; +lean_inc(x_20); +x_49 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_49, 0, x_20); +lean_ctor_set(x_49, 1, x_48); +lean_inc(x_4); +x_50 = l_Nat_repr(x_4); +x_51 = l_Lean_Syntax_mkNumLit(x_50, x_42); +x_52 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; +x_53 = lean_array_push(x_52, x_47); +x_54 = lean_array_push(x_53, x_49); +x_55 = lean_array_push(x_54, x_51); +x_56 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__11; +x_57 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_57, 0, x_42); +lean_ctor_set(x_57, 1, x_56); +lean_ctor_set(x_57, 2, x_55); +x_58 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_59 = lean_array_push(x_58, x_57); +x_60 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; +x_61 = lean_array_push(x_59, x_60); +x_62 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_63 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_63, 0, x_42); lean_ctor_set(x_63, 1, x_62); -x_64 = lean_array_push(x_50, x_33); -lean_inc(x_64); -x_65 = lean_array_push(x_64, x_61); -lean_inc(x_63); -x_66 = lean_array_push(x_65, x_63); -x_67 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__6; -x_68 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_68, 0, x_40); -lean_ctor_set(x_68, 1, x_67); -lean_ctor_set(x_68, 2, x_66); -x_69 = lean_array_push(x_56, x_41); -x_70 = lean_array_push(x_69, x_68); -x_71 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_71, 0, x_40); -lean_ctor_set(x_71, 1, x_60); -lean_ctor_set(x_71, 2, x_70); -x_72 = lean_array_push(x_56, x_37); -x_73 = lean_array_push(x_72, x_71); -x_74 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_75 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_75, 0, x_40); -lean_ctor_set(x_75, 1, x_74); -lean_ctor_set(x_75, 2, x_73); -x_76 = lean_array_push(x_56, x_75); -x_77 = lean_array_push(x_76, x_58); -x_78 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_78, 0, x_40); -lean_ctor_set(x_78, 1, x_60); -lean_ctor_set(x_78, 2, x_77); -x_79 = lean_array_push(x_64, x_78); -x_80 = lean_array_push(x_79, x_63); -x_81 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_81, 0, x_40); -lean_ctor_set(x_81, 1, x_67); -lean_ctor_set(x_81, 2, x_80); -x_82 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; -x_83 = lean_array_push(x_82, x_81); -x_84 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_84, 0, x_40); -lean_ctor_set(x_84, 1, x_60); -lean_ctor_set(x_84, 2, x_83); -x_85 = lean_array_push(x_56, x_31); -x_86 = lean_array_push(x_85, x_84); -x_87 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_87, 0, x_40); -lean_ctor_set(x_87, 1, x_74); -lean_ctor_set(x_87, 2, x_86); -lean_inc(x_2); -x_88 = l_List_range(x_2); -x_89 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__2; -x_90 = lean_unsigned_to_nat(0u); -x_91 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -lean_inc(x_10); -x_92 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__7(x_2, x_89, x_74, x_90, x_28, x_28, x_60, x_67, x_91, x_54, x_44, x_42, x_50, x_58, x_56, x_88, x_6, x_7, x_8, x_9, x_10, x_11, x_23); -lean_dec(x_2); -x_93 = lean_ctor_get(x_92, 0); -lean_inc(x_93); -x_94 = lean_ctor_get(x_92, 1); -lean_inc(x_94); -lean_dec(x_92); -x_95 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_95, 0, x_87); -lean_ctor_set(x_95, 1, x_93); -x_96 = l_List_appendTR___rarg(x_15, x_95); +lean_ctor_set(x_63, 2, x_61); +x_64 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; +x_65 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_65, 0, x_20); +lean_ctor_set(x_65, 1, x_64); +x_66 = lean_array_push(x_52, x_36); +lean_inc(x_66); +x_67 = lean_array_push(x_66, x_63); +lean_inc(x_65); +x_68 = lean_array_push(x_67, x_65); +x_69 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6; +x_70 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_70, 0, x_42); +lean_ctor_set(x_70, 1, x_69); +lean_ctor_set(x_70, 2, x_68); +x_71 = lean_array_push(x_58, x_43); +x_72 = lean_array_push(x_71, x_70); +x_73 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_73, 0, x_42); +lean_ctor_set(x_73, 1, x_62); +lean_ctor_set(x_73, 2, x_72); +x_74 = lean_array_push(x_58, x_40); +x_75 = lean_array_push(x_74, x_73); +x_76 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_77 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_77, 0, x_42); +lean_ctor_set(x_77, 1, x_76); +lean_ctor_set(x_77, 2, x_75); +x_78 = lean_array_push(x_58, x_77); +x_79 = lean_array_push(x_78, x_60); +x_80 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_80, 0, x_42); +lean_ctor_set(x_80, 1, x_62); +lean_ctor_set(x_80, 2, x_79); +x_81 = lean_array_push(x_66, x_80); +x_82 = lean_array_push(x_81, x_65); +x_83 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_83, 0, x_42); +lean_ctor_set(x_83, 1, x_69); +lean_ctor_set(x_83, 2, x_82); +x_84 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; +x_85 = lean_array_push(x_84, x_83); +x_86 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_86, 0, x_42); +lean_ctor_set(x_86, 1, x_62); +lean_ctor_set(x_86, 2, x_85); +x_87 = lean_array_push(x_58, x_34); +x_88 = lean_array_push(x_87, x_86); +x_89 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_89, 0, x_42); +lean_ctor_set(x_89, 1, x_76); +lean_ctor_set(x_89, 2, x_88); +lean_inc(x_4); +x_90 = l_List_range(x_4); +x_91 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__2; +x_92 = lean_unsigned_to_nat(0u); +x_93 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +x_94 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__9; +lean_inc(x_12); +lean_inc_n(x_3, 2); +x_95 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__7(x_2, x_4, x_91, x_76, x_92, x_3, x_3, x_62, x_69, x_93, x_94, x_56, x_46, x_44, x_52, x_60, x_58, x_90, x_8, x_9, x_10, x_11, x_12, x_13, x_25); +lean_dec(x_4); +x_96 = lean_ctor_get(x_95, 0); +lean_inc(x_96); +x_97 = lean_ctor_get(x_95, 1); +lean_inc(x_97); +lean_dec(x_95); +x_98 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_98, 0, x_89); +lean_ctor_set(x_98, 1, x_96); +x_99 = l_List_appendTR___rarg(x_17, x_98); +lean_inc(x_13); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_97 = lean_apply_8(x_4, x_96, x_6, x_7, x_8, x_9, x_10, x_11, x_94); -if (lean_obj_tag(x_97) == 0) -{ -lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_98 = lean_ctor_get(x_97, 0); -lean_inc(x_98); -x_99 = lean_ctor_get(x_97, 1); -lean_inc(x_99); -lean_dec(x_97); -lean_inc(x_11); -lean_inc(x_10); -x_100 = lean_apply_7(x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_99); +x_100 = lean_apply_8(x_6, x_99, x_8, x_9, x_10, x_11, x_12, x_13, x_97); if (lean_obj_tag(x_100) == 0) { -lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; uint8_t x_107; +lean_object* x_101; lean_object* x_102; lean_object* x_103; x_101 = lean_ctor_get(x_100, 0); lean_inc(x_101); x_102 = lean_ctor_get(x_100, 1); lean_inc(x_102); lean_dec(x_100); -x_103 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_10, x_11, x_102); +lean_inc(x_13); +lean_inc(x_12); +x_103 = lean_apply_7(x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_102); +if (lean_obj_tag(x_103) == 0) +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; uint8_t x_110; x_104 = lean_ctor_get(x_103, 0); lean_inc(x_104); x_105 = lean_ctor_get(x_103, 1); lean_inc(x_105); lean_dec(x_103); -x_106 = lean_st_ref_get(x_11, x_105); -lean_dec(x_11); -x_107 = !lean_is_exclusive(x_106); -if (x_107 == 0) +x_106 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_12, x_13, x_105); +x_107 = lean_ctor_get(x_106, 0); +lean_inc(x_107); +x_108 = lean_ctor_get(x_106, 1); +lean_inc(x_108); +lean_dec(x_106); +x_109 = lean_st_ref_get(x_13, x_108); +lean_dec(x_13); +x_110 = !lean_is_exclusive(x_109); +if (x_110 == 0) { -lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; -x_108 = lean_ctor_get(x_106, 0); -x_109 = lean_ctor_get(x_108, 0); -lean_inc(x_109); -lean_dec(x_108); -x_110 = lean_environment_main_module(x_109); -x_111 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__4; -lean_inc(x_20); -lean_inc(x_110); -x_112 = l_Lean_addMacroScope(x_110, x_111, x_20); -x_113 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__3; -x_114 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__16; -lean_inc(x_104); -x_115 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_115, 0, x_104); -lean_ctor_set(x_115, 1, x_113); -lean_ctor_set(x_115, 2, x_112); -lean_ctor_set(x_115, 3, x_114); -lean_inc(x_104); -x_116 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_116, 0, x_104); -lean_ctor_set(x_116, 1, x_32); -x_117 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__23; -lean_inc(x_20); -lean_inc(x_110); -x_118 = l_Lean_addMacroScope(x_110, x_117, x_20); -x_119 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__19; -x_120 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__25; -lean_inc(x_104); -x_121 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_121, 0, x_104); -lean_ctor_set(x_121, 1, x_119); -lean_ctor_set(x_121, 2, x_118); -lean_ctor_set(x_121, 3, x_120); -x_122 = l_Lean_addMacroScope(x_110, x_42, x_20); -lean_inc(x_104); -x_123 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_123, 0, x_104); -lean_ctor_set(x_123, 1, x_44); -lean_ctor_set(x_123, 2, x_122); -lean_ctor_set(x_123, 3, x_28); -x_124 = l_Nat_repr(x_3); -x_125 = l_Lean_Syntax_mkLit(x_39, x_124, x_40); -x_126 = lean_array_push(x_56, x_123); -x_127 = lean_array_push(x_126, x_125); -x_128 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_128, 0, x_40); -lean_ctor_set(x_128, 1, x_60); +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; +x_111 = lean_ctor_get(x_109, 0); +x_112 = lean_ctor_get(x_111, 0); +lean_inc(x_112); +lean_dec(x_111); +x_113 = lean_environment_main_module(x_112); +x_114 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__4; +lean_inc(x_22); +lean_inc(x_113); +x_115 = l_Lean_addMacroScope(x_113, x_114, x_22); +lean_inc(x_3); +x_116 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_116, 0, x_114); +lean_ctor_set(x_116, 1, x_3); +lean_inc(x_3); +x_117 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_117, 0, x_116); +lean_ctor_set(x_117, 1, x_3); +x_118 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__3; +lean_inc(x_107); +x_119 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_119, 0, x_107); +lean_ctor_set(x_119, 1, x_118); +lean_ctor_set(x_119, 2, x_115); +lean_ctor_set(x_119, 3, x_117); +lean_inc(x_107); +x_120 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_120, 0, x_107); +lean_ctor_set(x_120, 1, x_35); +x_121 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__23; +lean_inc(x_22); +lean_inc(x_113); +x_122 = l_Lean_addMacroScope(x_113, x_121, x_22); +lean_inc(x_3); +x_123 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_123, 0, x_121); +lean_ctor_set(x_123, 1, x_3); +lean_inc(x_3); +x_124 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_124, 0, x_123); +lean_ctor_set(x_124, 1, x_3); +x_125 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__19; +lean_inc(x_107); +x_126 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_126, 0, x_107); +lean_ctor_set(x_126, 1, x_125); +lean_ctor_set(x_126, 2, x_122); +lean_ctor_set(x_126, 3, x_124); +x_127 = l_Lean_addMacroScope(x_113, x_44, x_22); +lean_inc(x_107); +x_128 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_128, 0, x_107); +lean_ctor_set(x_128, 1, x_46); lean_ctor_set(x_128, 2, x_127); -x_129 = lean_array_push(x_56, x_121); -x_130 = lean_array_push(x_129, x_128); -x_131 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_131, 0, x_40); -lean_ctor_set(x_131, 1, x_74); -lean_ctor_set(x_131, 2, x_130); -x_132 = lean_array_push(x_56, x_131); -x_133 = lean_array_push(x_132, x_58); -x_134 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_134, 0, x_40); -lean_ctor_set(x_134, 1, x_60); -lean_ctor_set(x_134, 2, x_133); -x_135 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_135, 0, x_104); -lean_ctor_set(x_135, 1, x_62); -x_136 = lean_array_push(x_50, x_116); -x_137 = lean_array_push(x_136, x_134); -x_138 = lean_array_push(x_137, x_135); +lean_ctor_set(x_128, 3, x_3); +x_129 = l_Nat_repr(x_5); +x_130 = l_Lean_Syntax_mkNumLit(x_129, x_42); +x_131 = lean_array_push(x_58, x_128); +x_132 = lean_array_push(x_131, x_130); +x_133 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_133, 0, x_42); +lean_ctor_set(x_133, 1, x_62); +lean_ctor_set(x_133, 2, x_132); +x_134 = lean_array_push(x_58, x_126); +x_135 = lean_array_push(x_134, x_133); +x_136 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_136, 0, x_42); +lean_ctor_set(x_136, 1, x_76); +lean_ctor_set(x_136, 2, x_135); +x_137 = lean_array_push(x_58, x_136); +x_138 = lean_array_push(x_137, x_60); x_139 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_139, 0, x_40); -lean_ctor_set(x_139, 1, x_67); +lean_ctor_set(x_139, 0, x_42); +lean_ctor_set(x_139, 1, x_62); lean_ctor_set(x_139, 2, x_138); -x_140 = lean_array_push(x_50, x_139); -x_141 = lean_array_push(x_140, x_98); -x_142 = lean_array_push(x_141, x_101); -x_143 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_143, 0, x_40); -lean_ctor_set(x_143, 1, x_60); -lean_ctor_set(x_143, 2, x_142); -x_144 = lean_array_push(x_56, x_115); -x_145 = lean_array_push(x_144, x_143); -x_146 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_146, 0, x_40); -lean_ctor_set(x_146, 1, x_74); -lean_ctor_set(x_146, 2, x_145); -lean_ctor_set(x_106, 0, x_146); -return x_106; +x_140 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_140, 0, x_107); +lean_ctor_set(x_140, 1, x_64); +x_141 = lean_array_push(x_52, x_120); +x_142 = lean_array_push(x_141, x_139); +x_143 = lean_array_push(x_142, x_140); +x_144 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_144, 0, x_42); +lean_ctor_set(x_144, 1, x_69); +lean_ctor_set(x_144, 2, x_143); +x_145 = lean_array_push(x_52, x_144); +x_146 = lean_array_push(x_145, x_101); +x_147 = lean_array_push(x_146, x_104); +x_148 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_148, 0, x_42); +lean_ctor_set(x_148, 1, x_62); +lean_ctor_set(x_148, 2, x_147); +x_149 = lean_array_push(x_58, x_119); +x_150 = lean_array_push(x_149, x_148); +x_151 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_151, 0, x_42); +lean_ctor_set(x_151, 1, x_76); +lean_ctor_set(x_151, 2, x_150); +lean_ctor_set(x_109, 0, x_151); +return x_109; } else { -lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; -x_147 = lean_ctor_get(x_106, 0); -x_148 = lean_ctor_get(x_106, 1); -lean_inc(x_148); -lean_inc(x_147); -lean_dec(x_106); -x_149 = lean_ctor_get(x_147, 0); -lean_inc(x_149); -lean_dec(x_147); -x_150 = lean_environment_main_module(x_149); -x_151 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__4; -lean_inc(x_20); -lean_inc(x_150); -x_152 = l_Lean_addMacroScope(x_150, x_151, x_20); -x_153 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__3; -x_154 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__16; -lean_inc(x_104); -x_155 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_155, 0, x_104); -lean_ctor_set(x_155, 1, x_153); -lean_ctor_set(x_155, 2, x_152); -lean_ctor_set(x_155, 3, x_154); -lean_inc(x_104); -x_156 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_156, 0, x_104); -lean_ctor_set(x_156, 1, x_32); -x_157 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__23; -lean_inc(x_20); -lean_inc(x_150); -x_158 = l_Lean_addMacroScope(x_150, x_157, x_20); -x_159 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__19; -x_160 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__25; -lean_inc(x_104); +lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; +x_152 = lean_ctor_get(x_109, 0); +x_153 = lean_ctor_get(x_109, 1); +lean_inc(x_153); +lean_inc(x_152); +lean_dec(x_109); +x_154 = lean_ctor_get(x_152, 0); +lean_inc(x_154); +lean_dec(x_152); +x_155 = lean_environment_main_module(x_154); +x_156 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__4; +lean_inc(x_22); +lean_inc(x_155); +x_157 = l_Lean_addMacroScope(x_155, x_156, x_22); +lean_inc(x_3); +x_158 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_158, 0, x_156); +lean_ctor_set(x_158, 1, x_3); +lean_inc(x_3); +x_159 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_159, 0, x_158); +lean_ctor_set(x_159, 1, x_3); +x_160 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__3; +lean_inc(x_107); x_161 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_161, 0, x_104); -lean_ctor_set(x_161, 1, x_159); -lean_ctor_set(x_161, 2, x_158); -lean_ctor_set(x_161, 3, x_160); -x_162 = l_Lean_addMacroScope(x_150, x_42, x_20); -lean_inc(x_104); -x_163 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_163, 0, x_104); -lean_ctor_set(x_163, 1, x_44); -lean_ctor_set(x_163, 2, x_162); -lean_ctor_set(x_163, 3, x_28); -x_164 = l_Nat_repr(x_3); -x_165 = l_Lean_Syntax_mkLit(x_39, x_164, x_40); -x_166 = lean_array_push(x_56, x_163); -x_167 = lean_array_push(x_166, x_165); -x_168 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_168, 0, x_40); -lean_ctor_set(x_168, 1, x_60); -lean_ctor_set(x_168, 2, x_167); -x_169 = lean_array_push(x_56, x_161); -x_170 = lean_array_push(x_169, x_168); -x_171 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_171, 0, x_40); -lean_ctor_set(x_171, 1, x_74); -lean_ctor_set(x_171, 2, x_170); -x_172 = lean_array_push(x_56, x_171); -x_173 = lean_array_push(x_172, x_58); -x_174 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_174, 0, x_40); -lean_ctor_set(x_174, 1, x_60); -lean_ctor_set(x_174, 2, x_173); -x_175 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_175, 0, x_104); +lean_ctor_set(x_161, 0, x_107); +lean_ctor_set(x_161, 1, x_160); +lean_ctor_set(x_161, 2, x_157); +lean_ctor_set(x_161, 3, x_159); +lean_inc(x_107); +x_162 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_162, 0, x_107); +lean_ctor_set(x_162, 1, x_35); +x_163 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__23; +lean_inc(x_22); +lean_inc(x_155); +x_164 = l_Lean_addMacroScope(x_155, x_163, x_22); +lean_inc(x_3); +x_165 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_165, 0, x_163); +lean_ctor_set(x_165, 1, x_3); +lean_inc(x_3); +x_166 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_166, 0, x_165); +lean_ctor_set(x_166, 1, x_3); +x_167 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__19; +lean_inc(x_107); +x_168 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_168, 0, x_107); +lean_ctor_set(x_168, 1, x_167); +lean_ctor_set(x_168, 2, x_164); +lean_ctor_set(x_168, 3, x_166); +x_169 = l_Lean_addMacroScope(x_155, x_44, x_22); +lean_inc(x_107); +x_170 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_170, 0, x_107); +lean_ctor_set(x_170, 1, x_46); +lean_ctor_set(x_170, 2, x_169); +lean_ctor_set(x_170, 3, x_3); +x_171 = l_Nat_repr(x_5); +x_172 = l_Lean_Syntax_mkNumLit(x_171, x_42); +x_173 = lean_array_push(x_58, x_170); +x_174 = lean_array_push(x_173, x_172); +x_175 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_175, 0, x_42); lean_ctor_set(x_175, 1, x_62); -x_176 = lean_array_push(x_50, x_156); -x_177 = lean_array_push(x_176, x_174); -x_178 = lean_array_push(x_177, x_175); -x_179 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_179, 0, x_40); -lean_ctor_set(x_179, 1, x_67); -lean_ctor_set(x_179, 2, x_178); -x_180 = lean_array_push(x_50, x_179); -x_181 = lean_array_push(x_180, x_98); -x_182 = lean_array_push(x_181, x_101); -x_183 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_183, 0, x_40); -lean_ctor_set(x_183, 1, x_60); -lean_ctor_set(x_183, 2, x_182); -x_184 = lean_array_push(x_56, x_155); -x_185 = lean_array_push(x_184, x_183); +lean_ctor_set(x_175, 2, x_174); +x_176 = lean_array_push(x_58, x_168); +x_177 = lean_array_push(x_176, x_175); +x_178 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_178, 0, x_42); +lean_ctor_set(x_178, 1, x_76); +lean_ctor_set(x_178, 2, x_177); +x_179 = lean_array_push(x_58, x_178); +x_180 = lean_array_push(x_179, x_60); +x_181 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_181, 0, x_42); +lean_ctor_set(x_181, 1, x_62); +lean_ctor_set(x_181, 2, x_180); +x_182 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_182, 0, x_107); +lean_ctor_set(x_182, 1, x_64); +x_183 = lean_array_push(x_52, x_162); +x_184 = lean_array_push(x_183, x_181); +x_185 = lean_array_push(x_184, x_182); x_186 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_186, 0, x_40); -lean_ctor_set(x_186, 1, x_74); +lean_ctor_set(x_186, 0, x_42); +lean_ctor_set(x_186, 1, x_69); lean_ctor_set(x_186, 2, x_185); -x_187 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_187, 0, x_186); -lean_ctor_set(x_187, 1, x_148); -return x_187; +x_187 = lean_array_push(x_52, x_186); +x_188 = lean_array_push(x_187, x_101); +x_189 = lean_array_push(x_188, x_104); +x_190 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_190, 0, x_42); +lean_ctor_set(x_190, 1, x_62); +lean_ctor_set(x_190, 2, x_189); +x_191 = lean_array_push(x_58, x_161); +x_192 = lean_array_push(x_191, x_190); +x_193 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_193, 0, x_42); +lean_ctor_set(x_193, 1, x_76); +lean_ctor_set(x_193, 2, x_192); +x_194 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_194, 0, x_193); +lean_ctor_set(x_194, 1, x_153); +return x_194; } } else { -uint8_t x_188; -lean_dec(x_98); -lean_dec(x_20); -lean_dec(x_11); -lean_dec(x_10); +uint8_t x_195; +lean_dec(x_101); +lean_dec(x_22); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_5); lean_dec(x_3); -x_188 = !lean_is_exclusive(x_100); -if (x_188 == 0) +x_195 = !lean_is_exclusive(x_103); +if (x_195 == 0) { -return x_100; +return x_103; } else { -lean_object* x_189; lean_object* x_190; lean_object* x_191; -x_189 = lean_ctor_get(x_100, 0); -x_190 = lean_ctor_get(x_100, 1); -lean_inc(x_190); -lean_inc(x_189); -lean_dec(x_100); -x_191 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_191, 0, x_189); -lean_ctor_set(x_191, 1, x_190); -return x_191; +lean_object* x_196; lean_object* x_197; lean_object* x_198; +x_196 = lean_ctor_get(x_103, 0); +x_197 = lean_ctor_get(x_103, 1); +lean_inc(x_197); +lean_inc(x_196); +lean_dec(x_103); +x_198 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_198, 0, x_196); +lean_ctor_set(x_198, 1, x_197); +return x_198; } } } else { -uint8_t x_192; -lean_dec(x_20); +uint8_t x_199; +lean_dec(x_22); +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_192 = !lean_is_exclusive(x_97); -if (x_192 == 0) +x_199 = !lean_is_exclusive(x_100); +if (x_199 == 0) { -return x_97; +return x_100; } else { -lean_object* x_193; lean_object* x_194; lean_object* x_195; -x_193 = lean_ctor_get(x_97, 0); -x_194 = lean_ctor_get(x_97, 1); -lean_inc(x_194); -lean_inc(x_193); -lean_dec(x_97); -x_195 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_195, 0, x_193); -lean_ctor_set(x_195, 1, x_194); -return x_195; +lean_object* x_200; lean_object* x_201; lean_object* x_202; +x_200 = lean_ctor_get(x_100, 0); +x_201 = lean_ctor_get(x_100, 1); +lean_inc(x_201); +lean_inc(x_200); +lean_dec(x_100); +x_202 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_202, 0, x_200); +lean_ctor_set(x_202, 1, x_201); +return x_202; } } } @@ -29934,38 +31537,38 @@ lean_inc(x_17); x_336 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_17, x_18, x_19); if (x_335 == 0) { -lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; +lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; x_337 = lean_ctor_get(x_336, 0); lean_inc(x_337); x_338 = lean_ctor_get(x_336, 1); lean_inc(x_338); lean_dec(x_336); x_339 = lean_box(2); -x_340 = l_Lean_nullKind; -x_341 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_341, 0, x_339); -lean_ctor_set(x_341, 1, x_340); -lean_ctor_set(x_341, 2, x_10); -x_20 = x_341; +lean_inc(x_2); +x_340 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_340, 0, x_339); +lean_ctor_set(x_340, 1, x_2); +lean_ctor_set(x_340, 2, x_10); +x_20 = x_340; x_21 = x_337; x_22 = x_338; goto block_332; } else { -lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; -x_342 = lean_ctor_get(x_336, 0); +lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; +x_341 = lean_ctor_get(x_336, 0); +lean_inc(x_341); +x_342 = lean_ctor_get(x_336, 1); lean_inc(x_342); -x_343 = lean_ctor_get(x_336, 1); -lean_inc(x_343); lean_dec(x_336); -x_344 = l_Lean_instInhabitedSyntax; -x_345 = lean_unsigned_to_nat(0u); -x_346 = lean_array_get(x_344, x_10, x_345); +x_343 = l_Lean_instInhabitedSyntax; +x_344 = lean_unsigned_to_nat(0u); +x_345 = lean_array_get(x_343, x_10, x_344); lean_dec(x_10); -x_20 = x_346; -x_21 = x_342; -x_22 = x_343; +x_20 = x_345; +x_21 = x_341; +x_22 = x_342; goto block_332; } block_332: @@ -29984,7 +31587,7 @@ x_27 = lean_ctor_get(x_26, 0); lean_inc(x_27); lean_dec(x_26); x_28 = lean_environment_main_module(x_27); -x_29 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__16; +x_29 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__16; lean_inc(x_1); x_30 = lean_name_mk_string(x_1, x_29); lean_inc(x_21); @@ -29998,13 +31601,13 @@ x_34 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_34, 0, x_32); lean_ctor_set(x_34, 1, x_2); lean_ctor_set(x_34, 2, x_33); -x_35 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__1; +x_35 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__12___closed__1; lean_inc(x_1); x_36 = lean_name_mk_string(x_1, x_35); -x_37 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__5; +x_37 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__5; lean_inc(x_1); x_38 = lean_name_mk_string(x_1, x_37); -x_39 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; +x_39 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; lean_inc(x_21); x_40 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_40, 0, x_21); @@ -30012,7 +31615,7 @@ lean_ctor_set(x_40, 1, x_39); x_41 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16___closed__1; lean_inc(x_1); x_42 = lean_name_mk_string(x_1, x_41); -x_43 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__3; +x_43 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; lean_inc(x_21); x_44 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_44, 0, x_21); @@ -30029,7 +31632,7 @@ lean_ctor_set(x_48, 0, x_21); lean_ctor_set(x_48, 1, x_47); lean_ctor_set(x_48, 2, x_46); lean_ctor_set(x_48, 3, x_3); -x_49 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; +x_49 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; x_50 = lean_array_push(x_49, x_4); x_51 = lean_array_push(x_50, x_44); x_52 = lean_array_push(x_51, x_48); @@ -30037,20 +31640,20 @@ x_53 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_53, 0, x_32); lean_ctor_set(x_53, 1, x_42); lean_ctor_set(x_53, 2, x_52); -x_54 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__8; +x_54 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__8; lean_inc(x_1); x_55 = lean_name_mk_string(x_1, x_54); lean_inc(x_21); x_56 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_56, 0, x_21); lean_ctor_set(x_56, 1, x_54); -x_57 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__21; +x_57 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__21; lean_inc(x_1); x_58 = lean_name_mk_string(x_1, x_57); -x_59 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__23; +x_59 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__23; lean_inc(x_1); x_60 = lean_name_mk_string(x_1, x_59); -x_61 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__25; +x_61 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__25; lean_inc(x_21); x_62 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_62, 0, x_21); @@ -30063,7 +31666,7 @@ lean_inc(x_21); x_66 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_66, 0, x_21); lean_ctor_set(x_66, 1, x_65); -x_67 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; +x_67 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; lean_inc(x_21); x_68 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_68, 0, x_21); @@ -30090,16 +31693,16 @@ x_76 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_76, 0, x_32); lean_ctor_set(x_76, 1, x_2); lean_ctor_set(x_76, 2, x_75); -x_77 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__12; +x_77 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__12; lean_inc(x_21); x_78 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_78, 0, x_21); lean_ctor_set(x_78, 1, x_77); -x_79 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__4; +x_79 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__4; lean_inc(x_23); lean_inc(x_28); x_80 = l_Lean_addMacroScope(x_28, x_79, x_23); -x_81 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__7; +x_81 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__7; lean_inc(x_3); x_82 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_82, 0, x_81); @@ -30108,7 +31711,7 @@ lean_inc(x_3); x_83 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_83, 0, x_82); lean_ctor_set(x_83, 1, x_3); -x_84 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__3; +x_84 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__3; lean_inc(x_21); x_85 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_85, 0, x_21); @@ -30143,9 +31746,9 @@ x_96 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_96, 0, x_32); lean_ctor_set(x_96, 1, x_60); lean_ctor_set(x_96, 2, x_95); -x_97 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__1; +x_97 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__14___closed__1; x_98 = lean_name_mk_string(x_1, x_97); -x_99 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__2; +x_99 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__14___closed__2; lean_inc(x_21); x_100 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_100, 0, x_21); @@ -30170,9 +31773,9 @@ x_106 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_106, 0, x_32); lean_ctor_set(x_106, 1, x_2); lean_ctor_set(x_106, 2, x_105); -x_107 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__76; +x_107 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__76; x_108 = l_Lean_addMacroScope(x_28, x_107, x_23); -x_109 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__77; +x_109 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__77; lean_inc(x_3); x_110 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_110, 0, x_109); @@ -30180,7 +31783,7 @@ lean_ctor_set(x_110, 1, x_3); x_111 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_111, 0, x_110); lean_ctor_set(x_111, 1, x_3); -x_112 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__75; +x_112 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__75; lean_inc(x_21); x_113 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_113, 0, x_21); @@ -30266,7 +31869,7 @@ x_142 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_142, 0, x_32); lean_ctor_set(x_142, 1, x_2); lean_ctor_set(x_142, 2, x_141); -x_143 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__20; +x_143 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__20; x_144 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_144, 0, x_21); lean_ctor_set(x_144, 1, x_143); @@ -30338,7 +31941,7 @@ x_169 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_169, 0, x_32); lean_ctor_set(x_169, 1, x_58); lean_ctor_set(x_169, 2, x_168); -x_170 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__30; +x_170 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__30; x_171 = lean_array_push(x_170, x_31); lean_inc(x_34); x_172 = lean_array_push(x_171, x_34); @@ -30365,7 +31968,7 @@ x_180 = lean_ctor_get(x_178, 0); lean_inc(x_180); lean_dec(x_178); x_181 = lean_environment_main_module(x_180); -x_182 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__16; +x_182 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__16; lean_inc(x_1); x_183 = lean_name_mk_string(x_1, x_182); lean_inc(x_21); @@ -30379,13 +31982,13 @@ x_187 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_187, 0, x_185); lean_ctor_set(x_187, 1, x_2); lean_ctor_set(x_187, 2, x_186); -x_188 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__1; +x_188 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__12___closed__1; lean_inc(x_1); x_189 = lean_name_mk_string(x_1, x_188); -x_190 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__5; +x_190 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__5; lean_inc(x_1); x_191 = lean_name_mk_string(x_1, x_190); -x_192 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; +x_192 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; lean_inc(x_21); x_193 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_193, 0, x_21); @@ -30393,7 +31996,7 @@ lean_ctor_set(x_193, 1, x_192); x_194 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16___closed__1; lean_inc(x_1); x_195 = lean_name_mk_string(x_1, x_194); -x_196 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__3; +x_196 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3; lean_inc(x_21); x_197 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_197, 0, x_21); @@ -30410,7 +32013,7 @@ lean_ctor_set(x_201, 0, x_21); lean_ctor_set(x_201, 1, x_200); lean_ctor_set(x_201, 2, x_199); lean_ctor_set(x_201, 3, x_3); -x_202 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; +x_202 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; x_203 = lean_array_push(x_202, x_4); x_204 = lean_array_push(x_203, x_197); x_205 = lean_array_push(x_204, x_201); @@ -30418,20 +32021,20 @@ x_206 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_206, 0, x_185); lean_ctor_set(x_206, 1, x_195); lean_ctor_set(x_206, 2, x_205); -x_207 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__8; +x_207 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__8; lean_inc(x_1); x_208 = lean_name_mk_string(x_1, x_207); lean_inc(x_21); x_209 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_209, 0, x_21); lean_ctor_set(x_209, 1, x_207); -x_210 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__21; +x_210 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__21; lean_inc(x_1); x_211 = lean_name_mk_string(x_1, x_210); -x_212 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__23; +x_212 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__23; lean_inc(x_1); x_213 = lean_name_mk_string(x_1, x_212); -x_214 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__25; +x_214 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__25; lean_inc(x_21); x_215 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_215, 0, x_21); @@ -30444,7 +32047,7 @@ lean_inc(x_21); x_219 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_219, 0, x_21); lean_ctor_set(x_219, 1, x_218); -x_220 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; +x_220 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; lean_inc(x_21); x_221 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_221, 0, x_21); @@ -30471,16 +32074,16 @@ x_229 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_229, 0, x_185); lean_ctor_set(x_229, 1, x_2); lean_ctor_set(x_229, 2, x_228); -x_230 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__12; +x_230 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__12; lean_inc(x_21); x_231 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_231, 0, x_21); lean_ctor_set(x_231, 1, x_230); -x_232 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__4; +x_232 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__4; lean_inc(x_23); lean_inc(x_181); x_233 = l_Lean_addMacroScope(x_181, x_232, x_23); -x_234 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__7; +x_234 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__7; lean_inc(x_3); x_235 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_235, 0, x_234); @@ -30489,7 +32092,7 @@ lean_inc(x_3); x_236 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_236, 0, x_235); lean_ctor_set(x_236, 1, x_3); -x_237 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__3; +x_237 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__3; lean_inc(x_21); x_238 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_238, 0, x_21); @@ -30524,9 +32127,9 @@ x_249 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_249, 0, x_185); lean_ctor_set(x_249, 1, x_213); lean_ctor_set(x_249, 2, x_248); -x_250 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__1; +x_250 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__14___closed__1; x_251 = lean_name_mk_string(x_1, x_250); -x_252 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__2; +x_252 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__14___closed__2; lean_inc(x_21); x_253 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_253, 0, x_21); @@ -30551,9 +32154,9 @@ x_259 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_259, 0, x_185); lean_ctor_set(x_259, 1, x_2); lean_ctor_set(x_259, 2, x_258); -x_260 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__76; +x_260 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__76; x_261 = l_Lean_addMacroScope(x_181, x_260, x_23); -x_262 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__77; +x_262 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__77; lean_inc(x_3); x_263 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_263, 0, x_262); @@ -30561,7 +32164,7 @@ lean_ctor_set(x_263, 1, x_3); x_264 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_264, 0, x_263); lean_ctor_set(x_264, 1, x_3); -x_265 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__75; +x_265 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__75; lean_inc(x_21); x_266 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_266, 0, x_21); @@ -30647,7 +32250,7 @@ x_295 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_295, 0, x_185); lean_ctor_set(x_295, 1, x_2); lean_ctor_set(x_295, 2, x_294); -x_296 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__20; +x_296 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__20; x_297 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_297, 0, x_21); lean_ctor_set(x_297, 1, x_296); @@ -30719,7 +32322,7 @@ x_322 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_322, 0, x_185); lean_ctor_set(x_322, 1, x_211); lean_ctor_set(x_322, 2, x_321); -x_323 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__30; +x_323 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__30; x_324 = lean_array_push(x_323, x_184); lean_inc(x_187); x_325 = lean_array_push(x_324, x_187); @@ -30762,489 +32365,561 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17(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_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, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17(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_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, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20) { _start: { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -lean_dec(x_12); -lean_inc(x_17); +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +lean_dec(x_13); +lean_inc(x_18); lean_inc(x_1); -x_20 = l_Lean_Elab_Term_Quotation_mkTuple(x_1, x_13, x_14, x_15, x_16, x_17, x_18, x_19); -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_20, 1); +x_21 = l_Lean_Elab_Term_Quotation_mkTuple(x_1, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +x_22 = lean_ctor_get(x_21, 0); lean_inc(x_22); -lean_dec(x_20); -x_23 = lean_array_get_size(x_1); -x_24 = lean_unsigned_to_nat(1u); -x_25 = lean_nat_dec_eq(x_23, x_24); -if (x_25 == 0) +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = lean_array_get_size(x_1); +x_25 = lean_unsigned_to_nat(1u); +x_26 = lean_nat_dec_eq(x_24, x_25); +if (x_26 == 0) { -size_t x_26; size_t x_27; lean_object* x_28; lean_object* x_29; 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; lean_object* x_44; lean_object* x_45; -x_26 = lean_usize_of_nat(x_23); -lean_dec(x_23); -x_27 = 0; -x_28 = lean_unsigned_to_nat(0u); -lean_inc(x_17); -lean_inc(x_21); +size_t x_27; size_t x_28; lean_object* x_29; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_27 = lean_usize_of_nat(x_24); +lean_dec(x_24); +x_28 = 0; +x_29 = lean_unsigned_to_nat(0u); +lean_inc(x_18); +lean_inc(x_22); lean_inc(x_6); lean_inc(x_5); lean_inc(x_3); lean_inc(x_4); lean_inc(x_7); lean_inc(x_2); -x_29 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8(x_2, x_7, x_28, x_4, x_3, x_5, x_6, x_21, x_1, x_26, x_27, x_10, x_13, x_14, x_15, x_16, x_17, x_18, x_22); +x_30 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8(x_10, x_2, x_7, x_29, x_4, x_3, x_5, x_6, x_22, x_1, x_27, x_28, x_11, x_14, x_15, x_16, x_17, x_18, x_19, x_23); lean_dec(x_1); -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_29, 1); +lean_dec(x_10); +x_31 = lean_ctor_get(x_30, 0); lean_inc(x_31); -lean_dec(x_29); -lean_inc(x_17); -x_32 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_17, x_18, x_31); -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_32, 1); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +lean_inc(x_18); +x_33 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_18, x_19, x_32); +x_34 = lean_ctor_get(x_33, 0); lean_inc(x_34); -lean_dec(x_32); -x_35 = lean_ctor_get(x_17, 10); +x_35 = lean_ctor_get(x_33, 1); lean_inc(x_35); -x_36 = lean_st_ref_get(x_18, x_34); -x_37 = lean_ctor_get(x_36, 0); -lean_inc(x_37); -x_38 = lean_ctor_get(x_36, 1); +lean_dec(x_33); +x_36 = lean_ctor_get(x_18, 10); +lean_inc(x_36); +x_37 = lean_st_ref_get(x_19, x_35); +x_38 = lean_ctor_get(x_37, 0); lean_inc(x_38); -lean_dec(x_36); -x_39 = lean_ctor_get(x_37, 0); +x_39 = lean_ctor_get(x_37, 1); lean_inc(x_39); lean_dec(x_37); -x_40 = lean_environment_main_module(x_39); -x_41 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__4; -x_42 = l_Lean_addMacroScope(x_40, x_41, x_35); -x_43 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17___closed__2; +x_40 = lean_ctor_get(x_38, 0); +lean_inc(x_40); +lean_dec(x_38); +x_41 = lean_environment_main_module(x_40); +x_42 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__4; +x_43 = l_Lean_addMacroScope(x_41, x_42, x_36); +x_44 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17___closed__2; lean_inc(x_4); -x_44 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_44, 0, x_33); -lean_ctor_set(x_44, 1, x_43); -lean_ctor_set(x_44, 2, x_42); -lean_ctor_set(x_44, 3, x_4); -x_45 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16(x_2, x_3, x_4, x_11, x_5, x_21, x_6, x_7, x_8, x_9, x_30, x_44, x_13, x_14, x_15, x_16, x_17, x_18, x_38); -return x_45; +x_45 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_45, 0, x_34); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_43); +lean_ctor_set(x_45, 3, x_4); +x_46 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16(x_2, x_3, x_4, x_12, x_5, x_22, x_6, x_7, x_8, x_9, x_31, x_45, x_14, x_15, x_16, x_17, x_18, x_19, x_39); +lean_dec(x_19); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +return x_46; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -lean_dec(x_23); -x_46 = lean_unsigned_to_nat(0u); -x_47 = lean_array_fget(x_1, x_46); +lean_object* x_47; lean_object* x_48; lean_object* x_49; +lean_dec(x_24); +lean_dec(x_10); +x_47 = lean_unsigned_to_nat(0u); +x_48 = lean_array_fget(x_1, x_47); lean_dec(x_1); -x_48 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16(x_2, x_3, x_4, x_11, x_5, x_21, x_6, x_7, x_8, x_9, x_10, x_47, x_13, x_14, x_15, x_16, x_17, x_18, x_22); -return x_48; +x_49 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16(x_2, x_3, x_4, x_12, x_5, x_22, x_6, x_7, x_8, x_9, x_11, x_48, x_14, x_15, x_16, x_17, x_18, x_19, x_23); +lean_dec(x_19); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +return x_49; } } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__18(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_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, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__18(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_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, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20) { _start: { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -lean_dec(x_12); -lean_inc(x_17); +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +lean_dec(x_13); +lean_inc(x_18); lean_inc(x_1); -x_20 = l_Lean_Elab_Term_Quotation_mkTuple(x_1, x_13, x_14, x_15, x_16, x_17, x_18, x_19); -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_20, 1); +x_21 = l_Lean_Elab_Term_Quotation_mkTuple(x_1, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +x_22 = lean_ctor_get(x_21, 0); lean_inc(x_22); -lean_dec(x_20); -x_23 = lean_array_get_size(x_1); -x_24 = lean_unsigned_to_nat(1u); -x_25 = lean_nat_dec_eq(x_23, x_24); -if (x_25 == 0) +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = lean_array_get_size(x_1); +x_25 = lean_unsigned_to_nat(1u); +x_26 = lean_nat_dec_eq(x_24, x_25); +if (x_26 == 0) { -size_t x_26; size_t x_27; lean_object* x_28; lean_object* x_29; 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; lean_object* x_44; lean_object* x_45; -x_26 = lean_usize_of_nat(x_23); -lean_dec(x_23); -x_27 = 0; -x_28 = lean_unsigned_to_nat(0u); -lean_inc(x_17); -lean_inc(x_21); +size_t x_27; size_t x_28; lean_object* x_29; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_27 = lean_usize_of_nat(x_24); +lean_dec(x_24); +x_28 = 0; +x_29 = lean_unsigned_to_nat(0u); +lean_inc(x_18); +lean_inc(x_22); lean_inc(x_6); lean_inc(x_5); lean_inc(x_3); lean_inc(x_4); lean_inc(x_7); lean_inc(x_2); -x_29 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9(x_2, x_7, x_28, x_4, x_3, x_5, x_6, x_21, x_1, x_26, x_27, x_10, x_13, x_14, x_15, x_16, x_17, x_18, x_22); +x_30 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9(x_10, x_2, x_7, x_29, x_4, x_3, x_5, x_6, x_22, x_1, x_27, x_28, x_11, x_14, x_15, x_16, x_17, x_18, x_19, x_23); lean_dec(x_1); -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_29, 1); +lean_dec(x_10); +x_31 = lean_ctor_get(x_30, 0); lean_inc(x_31); -lean_dec(x_29); -lean_inc(x_17); -x_32 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_17, x_18, x_31); -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_32, 1); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +lean_inc(x_18); +x_33 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_18, x_19, x_32); +x_34 = lean_ctor_get(x_33, 0); lean_inc(x_34); -lean_dec(x_32); -x_35 = lean_ctor_get(x_17, 10); +x_35 = lean_ctor_get(x_33, 1); lean_inc(x_35); -x_36 = lean_st_ref_get(x_18, x_34); -x_37 = lean_ctor_get(x_36, 0); -lean_inc(x_37); -x_38 = lean_ctor_get(x_36, 1); +lean_dec(x_33); +x_36 = lean_ctor_get(x_18, 10); +lean_inc(x_36); +x_37 = lean_st_ref_get(x_19, x_35); +x_38 = lean_ctor_get(x_37, 0); lean_inc(x_38); -lean_dec(x_36); -x_39 = lean_ctor_get(x_37, 0); +x_39 = lean_ctor_get(x_37, 1); lean_inc(x_39); lean_dec(x_37); -x_40 = lean_environment_main_module(x_39); -x_41 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__4; -x_42 = l_Lean_addMacroScope(x_40, x_41, x_35); -x_43 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17___closed__2; +x_40 = lean_ctor_get(x_38, 0); +lean_inc(x_40); +lean_dec(x_38); +x_41 = lean_environment_main_module(x_40); +x_42 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__4; +x_43 = l_Lean_addMacroScope(x_41, x_42, x_36); +x_44 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17___closed__2; lean_inc(x_4); -x_44 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_44, 0, x_33); -lean_ctor_set(x_44, 1, x_43); -lean_ctor_set(x_44, 2, x_42); -lean_ctor_set(x_44, 3, x_4); -x_45 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16(x_2, x_3, x_4, x_11, x_5, x_21, x_6, x_7, x_8, x_9, x_30, x_44, x_13, x_14, x_15, x_16, x_17, x_18, x_38); -return x_45; +x_45 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_45, 0, x_34); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_43); +lean_ctor_set(x_45, 3, x_4); +x_46 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16(x_2, x_3, x_4, x_12, x_5, x_22, x_6, x_7, x_8, x_9, x_31, x_45, x_14, x_15, x_16, x_17, x_18, x_19, x_39); +lean_dec(x_19); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +return x_46; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -lean_dec(x_23); -x_46 = lean_unsigned_to_nat(0u); -x_47 = lean_array_fget(x_1, x_46); +lean_object* x_47; lean_object* x_48; lean_object* x_49; +lean_dec(x_24); +lean_dec(x_10); +x_47 = lean_unsigned_to_nat(0u); +x_48 = lean_array_fget(x_1, x_47); lean_dec(x_1); -x_48 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16(x_2, x_3, x_4, x_11, x_5, x_21, x_6, x_7, x_8, x_9, x_10, x_47, x_13, x_14, x_15, x_16, x_17, x_18, x_22); -return x_48; +x_49 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16(x_2, x_3, x_4, x_12, x_5, x_22, x_6, x_7, x_8, x_9, x_11, x_48, x_14, x_15, x_16, x_17, x_18, x_19, x_23); +lean_dec(x_19); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +return x_49; } } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__19(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_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, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__19(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_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, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20) { _start: { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -lean_dec(x_12); -lean_inc(x_17); +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +lean_dec(x_13); +lean_inc(x_18); lean_inc(x_1); -x_20 = l_Lean_Elab_Term_Quotation_mkTuple(x_1, x_13, x_14, x_15, x_16, x_17, x_18, x_19); -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_20, 1); +x_21 = l_Lean_Elab_Term_Quotation_mkTuple(x_1, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +x_22 = lean_ctor_get(x_21, 0); lean_inc(x_22); -lean_dec(x_20); -x_23 = lean_array_get_size(x_1); -x_24 = lean_unsigned_to_nat(1u); -x_25 = lean_nat_dec_eq(x_23, x_24); -if (x_25 == 0) +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = lean_array_get_size(x_1); +x_25 = lean_unsigned_to_nat(1u); +x_26 = lean_nat_dec_eq(x_24, x_25); +if (x_26 == 0) { -size_t x_26; size_t x_27; lean_object* x_28; lean_object* x_29; 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; lean_object* x_44; lean_object* x_45; -x_26 = lean_usize_of_nat(x_23); -lean_dec(x_23); -x_27 = 0; -x_28 = lean_unsigned_to_nat(0u); -lean_inc(x_17); -lean_inc(x_21); +size_t x_27; size_t x_28; lean_object* x_29; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_27 = lean_usize_of_nat(x_24); +lean_dec(x_24); +x_28 = 0; +x_29 = lean_unsigned_to_nat(0u); +lean_inc(x_18); +lean_inc(x_22); lean_inc(x_6); lean_inc(x_5); lean_inc(x_3); lean_inc(x_4); lean_inc(x_7); lean_inc(x_2); -x_29 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__10(x_2, x_7, x_28, x_4, x_3, x_5, x_6, x_21, x_1, x_26, x_27, x_10, x_13, x_14, x_15, x_16, x_17, x_18, x_22); +x_30 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__10(x_10, x_2, x_7, x_29, x_4, x_3, x_5, x_6, x_22, x_1, x_27, x_28, x_11, x_14, x_15, x_16, x_17, x_18, x_19, x_23); lean_dec(x_1); -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_29, 1); +lean_dec(x_10); +x_31 = lean_ctor_get(x_30, 0); lean_inc(x_31); -lean_dec(x_29); -lean_inc(x_17); -x_32 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_17, x_18, x_31); -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_32, 1); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +lean_inc(x_18); +x_33 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_18, x_19, x_32); +x_34 = lean_ctor_get(x_33, 0); lean_inc(x_34); -lean_dec(x_32); -x_35 = lean_ctor_get(x_17, 10); +x_35 = lean_ctor_get(x_33, 1); lean_inc(x_35); -x_36 = lean_st_ref_get(x_18, x_34); -x_37 = lean_ctor_get(x_36, 0); -lean_inc(x_37); -x_38 = lean_ctor_get(x_36, 1); +lean_dec(x_33); +x_36 = lean_ctor_get(x_18, 10); +lean_inc(x_36); +x_37 = lean_st_ref_get(x_19, x_35); +x_38 = lean_ctor_get(x_37, 0); lean_inc(x_38); -lean_dec(x_36); -x_39 = lean_ctor_get(x_37, 0); +x_39 = lean_ctor_get(x_37, 1); lean_inc(x_39); lean_dec(x_37); -x_40 = lean_environment_main_module(x_39); -x_41 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__4; -x_42 = l_Lean_addMacroScope(x_40, x_41, x_35); -x_43 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17___closed__2; +x_40 = lean_ctor_get(x_38, 0); +lean_inc(x_40); +lean_dec(x_38); +x_41 = lean_environment_main_module(x_40); +x_42 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__4; +x_43 = l_Lean_addMacroScope(x_41, x_42, x_36); +x_44 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17___closed__2; lean_inc(x_4); -x_44 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_44, 0, x_33); -lean_ctor_set(x_44, 1, x_43); -lean_ctor_set(x_44, 2, x_42); -lean_ctor_set(x_44, 3, x_4); -x_45 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16(x_2, x_3, x_4, x_11, x_5, x_21, x_6, x_7, x_8, x_9, x_30, x_44, x_13, x_14, x_15, x_16, x_17, x_18, x_38); -return x_45; +x_45 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_45, 0, x_34); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_43); +lean_ctor_set(x_45, 3, x_4); +x_46 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16(x_2, x_3, x_4, x_12, x_5, x_22, x_6, x_7, x_8, x_9, x_31, x_45, x_14, x_15, x_16, x_17, x_18, x_19, x_39); +lean_dec(x_19); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +return x_46; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -lean_dec(x_23); -x_46 = lean_unsigned_to_nat(0u); -x_47 = lean_array_fget(x_1, x_46); +lean_object* x_47; lean_object* x_48; lean_object* x_49; +lean_dec(x_24); +lean_dec(x_10); +x_47 = lean_unsigned_to_nat(0u); +x_48 = lean_array_fget(x_1, x_47); lean_dec(x_1); -x_48 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16(x_2, x_3, x_4, x_11, x_5, x_21, x_6, x_7, x_8, x_9, x_10, x_47, x_13, x_14, x_15, x_16, x_17, x_18, x_22); -return x_48; +x_49 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16(x_2, x_3, x_4, x_12, x_5, x_22, x_6, x_7, x_8, x_9, x_11, x_48, x_14, x_15, x_16, x_17, x_18, x_19, x_23); +lean_dec(x_19); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +return x_49; } } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__20(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_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, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__20(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_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, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20) { _start: { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -lean_dec(x_12); -lean_inc(x_17); +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +lean_dec(x_13); +lean_inc(x_18); lean_inc(x_1); -x_20 = l_Lean_Elab_Term_Quotation_mkTuple(x_1, x_13, x_14, x_15, x_16, x_17, x_18, x_19); -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_20, 1); +x_21 = l_Lean_Elab_Term_Quotation_mkTuple(x_1, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +x_22 = lean_ctor_get(x_21, 0); lean_inc(x_22); -lean_dec(x_20); -x_23 = lean_array_get_size(x_1); -x_24 = lean_unsigned_to_nat(1u); -x_25 = lean_nat_dec_eq(x_23, x_24); -if (x_25 == 0) +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = lean_array_get_size(x_1); +x_25 = lean_unsigned_to_nat(1u); +x_26 = lean_nat_dec_eq(x_24, x_25); +if (x_26 == 0) { -size_t x_26; size_t x_27; lean_object* x_28; lean_object* x_29; 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; lean_object* x_44; lean_object* x_45; -x_26 = lean_usize_of_nat(x_23); -lean_dec(x_23); -x_27 = 0; -x_28 = lean_unsigned_to_nat(0u); -lean_inc(x_17); -lean_inc(x_21); +size_t x_27; size_t x_28; lean_object* x_29; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_27 = lean_usize_of_nat(x_24); +lean_dec(x_24); +x_28 = 0; +x_29 = lean_unsigned_to_nat(0u); +lean_inc(x_18); +lean_inc(x_22); lean_inc(x_6); lean_inc(x_5); lean_inc(x_3); lean_inc(x_4); lean_inc(x_7); lean_inc(x_2); -x_29 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__12(x_2, x_7, x_28, x_4, x_3, x_5, x_6, x_21, x_1, x_26, x_27, x_10, x_13, x_14, x_15, x_16, x_17, x_18, x_22); +x_30 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__12(x_10, x_2, x_7, x_29, x_4, x_3, x_5, x_6, x_22, x_1, x_27, x_28, x_11, x_14, x_15, x_16, x_17, x_18, x_19, x_23); lean_dec(x_1); -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_29, 1); +lean_dec(x_10); +x_31 = lean_ctor_get(x_30, 0); lean_inc(x_31); -lean_dec(x_29); -lean_inc(x_17); -x_32 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_17, x_18, x_31); -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_32, 1); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +lean_inc(x_18); +x_33 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_18, x_19, x_32); +x_34 = lean_ctor_get(x_33, 0); lean_inc(x_34); -lean_dec(x_32); -x_35 = lean_ctor_get(x_17, 10); +x_35 = lean_ctor_get(x_33, 1); lean_inc(x_35); -x_36 = lean_st_ref_get(x_18, x_34); -x_37 = lean_ctor_get(x_36, 0); -lean_inc(x_37); -x_38 = lean_ctor_get(x_36, 1); +lean_dec(x_33); +x_36 = lean_ctor_get(x_18, 10); +lean_inc(x_36); +x_37 = lean_st_ref_get(x_19, x_35); +x_38 = lean_ctor_get(x_37, 0); lean_inc(x_38); -lean_dec(x_36); -x_39 = lean_ctor_get(x_37, 0); +x_39 = lean_ctor_get(x_37, 1); lean_inc(x_39); lean_dec(x_37); -x_40 = lean_environment_main_module(x_39); -x_41 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__4; -x_42 = l_Lean_addMacroScope(x_40, x_41, x_35); -x_43 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17___closed__2; +x_40 = lean_ctor_get(x_38, 0); +lean_inc(x_40); +lean_dec(x_38); +x_41 = lean_environment_main_module(x_40); +x_42 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__4; +x_43 = l_Lean_addMacroScope(x_41, x_42, x_36); +x_44 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17___closed__2; lean_inc(x_4); -x_44 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_44, 0, x_33); -lean_ctor_set(x_44, 1, x_43); -lean_ctor_set(x_44, 2, x_42); -lean_ctor_set(x_44, 3, x_4); -x_45 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16(x_2, x_3, x_4, x_11, x_5, x_21, x_6, x_7, x_8, x_9, x_30, x_44, x_13, x_14, x_15, x_16, x_17, x_18, x_38); -return x_45; +x_45 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_45, 0, x_34); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_43); +lean_ctor_set(x_45, 3, x_4); +x_46 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16(x_2, x_3, x_4, x_12, x_5, x_22, x_6, x_7, x_8, x_9, x_31, x_45, x_14, x_15, x_16, x_17, x_18, x_19, x_39); +lean_dec(x_19); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +return x_46; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -lean_dec(x_23); -x_46 = lean_unsigned_to_nat(0u); -x_47 = lean_array_fget(x_1, x_46); +lean_object* x_47; lean_object* x_48; lean_object* x_49; +lean_dec(x_24); +lean_dec(x_10); +x_47 = lean_unsigned_to_nat(0u); +x_48 = lean_array_fget(x_1, x_47); lean_dec(x_1); -x_48 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16(x_2, x_3, x_4, x_11, x_5, x_21, x_6, x_7, x_8, x_9, x_10, x_47, x_13, x_14, x_15, x_16, x_17, x_18, x_22); -return x_48; +x_49 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16(x_2, x_3, x_4, x_12, x_5, x_22, x_6, x_7, x_8, x_9, x_11, x_48, x_14, x_15, x_16, x_17, x_18, x_19, x_23); +lean_dec(x_19); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +return x_49; } } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__21(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_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, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__21(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_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, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20) { _start: { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -lean_dec(x_12); -lean_inc(x_17); +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +lean_dec(x_13); +lean_inc(x_18); lean_inc(x_1); -x_20 = l_Lean_Elab_Term_Quotation_mkTuple(x_1, x_13, x_14, x_15, x_16, x_17, x_18, x_19); -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_20, 1); +x_21 = l_Lean_Elab_Term_Quotation_mkTuple(x_1, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +x_22 = lean_ctor_get(x_21, 0); lean_inc(x_22); -lean_dec(x_20); -x_23 = lean_array_get_size(x_1); -x_24 = lean_unsigned_to_nat(1u); -x_25 = lean_nat_dec_eq(x_23, x_24); -if (x_25 == 0) +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = lean_array_get_size(x_1); +x_25 = lean_unsigned_to_nat(1u); +x_26 = lean_nat_dec_eq(x_24, x_25); +if (x_26 == 0) { -size_t x_26; size_t x_27; lean_object* x_28; lean_object* x_29; 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; lean_object* x_44; lean_object* x_45; -x_26 = lean_usize_of_nat(x_23); -lean_dec(x_23); -x_27 = 0; -x_28 = lean_unsigned_to_nat(0u); -lean_inc(x_17); -lean_inc(x_21); +size_t x_27; size_t x_28; lean_object* x_29; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_27 = lean_usize_of_nat(x_24); +lean_dec(x_24); +x_28 = 0; +x_29 = lean_unsigned_to_nat(0u); +lean_inc(x_18); +lean_inc(x_22); lean_inc(x_6); lean_inc(x_5); lean_inc(x_3); lean_inc(x_4); lean_inc(x_7); lean_inc(x_2); -x_29 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__13(x_2, x_7, x_28, x_4, x_3, x_5, x_6, x_21, x_1, x_26, x_27, x_10, x_13, x_14, x_15, x_16, x_17, x_18, x_22); +x_30 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__13(x_10, x_2, x_7, x_29, x_4, x_3, x_5, x_6, x_22, x_1, x_27, x_28, x_11, x_14, x_15, x_16, x_17, x_18, x_19, x_23); lean_dec(x_1); -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_29, 1); +lean_dec(x_10); +x_31 = lean_ctor_get(x_30, 0); lean_inc(x_31); -lean_dec(x_29); -lean_inc(x_17); -x_32 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_17, x_18, x_31); -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_32, 1); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +lean_inc(x_18); +x_33 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_18, x_19, x_32); +x_34 = lean_ctor_get(x_33, 0); lean_inc(x_34); -lean_dec(x_32); -x_35 = lean_ctor_get(x_17, 10); +x_35 = lean_ctor_get(x_33, 1); lean_inc(x_35); -x_36 = lean_st_ref_get(x_18, x_34); -x_37 = lean_ctor_get(x_36, 0); -lean_inc(x_37); -x_38 = lean_ctor_get(x_36, 1); +lean_dec(x_33); +x_36 = lean_ctor_get(x_18, 10); +lean_inc(x_36); +x_37 = lean_st_ref_get(x_19, x_35); +x_38 = lean_ctor_get(x_37, 0); lean_inc(x_38); -lean_dec(x_36); -x_39 = lean_ctor_get(x_37, 0); +x_39 = lean_ctor_get(x_37, 1); lean_inc(x_39); lean_dec(x_37); -x_40 = lean_environment_main_module(x_39); -x_41 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__4; -x_42 = l_Lean_addMacroScope(x_40, x_41, x_35); -x_43 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17___closed__2; +x_40 = lean_ctor_get(x_38, 0); +lean_inc(x_40); +lean_dec(x_38); +x_41 = lean_environment_main_module(x_40); +x_42 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__4; +x_43 = l_Lean_addMacroScope(x_41, x_42, x_36); +x_44 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17___closed__2; lean_inc(x_4); -x_44 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_44, 0, x_33); -lean_ctor_set(x_44, 1, x_43); -lean_ctor_set(x_44, 2, x_42); -lean_ctor_set(x_44, 3, x_4); -x_45 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16(x_2, x_3, x_4, x_11, x_5, x_21, x_6, x_7, x_8, x_9, x_30, x_44, x_13, x_14, x_15, x_16, x_17, x_18, x_38); -return x_45; +x_45 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_45, 0, x_34); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_43); +lean_ctor_set(x_45, 3, x_4); +x_46 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16(x_2, x_3, x_4, x_12, x_5, x_22, x_6, x_7, x_8, x_9, x_31, x_45, x_14, x_15, x_16, x_17, x_18, x_19, x_39); +lean_dec(x_19); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +return x_46; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -lean_dec(x_23); -x_46 = lean_unsigned_to_nat(0u); -x_47 = lean_array_fget(x_1, x_46); +lean_object* x_47; lean_object* x_48; lean_object* x_49; +lean_dec(x_24); +lean_dec(x_10); +x_47 = lean_unsigned_to_nat(0u); +x_48 = lean_array_fget(x_1, x_47); lean_dec(x_1); -x_48 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16(x_2, x_3, x_4, x_11, x_5, x_21, x_6, x_7, x_8, x_9, x_10, x_47, x_13, x_14, x_15, x_16, x_17, x_18, x_22); -return x_48; +x_49 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16(x_2, x_3, x_4, x_12, x_5, x_22, x_6, x_7, x_8, x_9, x_11, x_48, x_14, x_15, x_16, x_17, x_18, x_19, x_23); +lean_dec(x_19); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +return x_49; } } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__22(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_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, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__22(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_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, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20) { _start: { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -lean_dec(x_12); -lean_inc(x_17); +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +lean_dec(x_13); +lean_inc(x_18); lean_inc(x_1); -x_20 = l_Lean_Elab_Term_Quotation_mkTuple(x_1, x_13, x_14, x_15, x_16, x_17, x_18, x_19); -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_20, 1); +x_21 = l_Lean_Elab_Term_Quotation_mkTuple(x_1, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +x_22 = lean_ctor_get(x_21, 0); lean_inc(x_22); -lean_dec(x_20); -x_23 = lean_array_get_size(x_1); -x_24 = lean_unsigned_to_nat(1u); -x_25 = lean_nat_dec_eq(x_23, x_24); -if (x_25 == 0) +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = lean_array_get_size(x_1); +x_25 = lean_unsigned_to_nat(1u); +x_26 = lean_nat_dec_eq(x_24, x_25); +if (x_26 == 0) { -size_t x_26; size_t x_27; lean_object* x_28; lean_object* x_29; 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; lean_object* x_44; lean_object* x_45; -x_26 = lean_usize_of_nat(x_23); -lean_dec(x_23); -x_27 = 0; -x_28 = lean_unsigned_to_nat(0u); -lean_inc(x_17); -lean_inc(x_21); +size_t x_27; size_t x_28; lean_object* x_29; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_27 = lean_usize_of_nat(x_24); +lean_dec(x_24); +x_28 = 0; +x_29 = lean_unsigned_to_nat(0u); +lean_inc(x_18); +lean_inc(x_22); lean_inc(x_6); lean_inc(x_5); lean_inc(x_3); lean_inc(x_4); lean_inc(x_7); lean_inc(x_2); -x_29 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__14(x_2, x_7, x_28, x_4, x_3, x_5, x_6, x_21, x_1, x_26, x_27, x_10, x_13, x_14, x_15, x_16, x_17, x_18, x_22); +x_30 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__14(x_10, x_2, x_7, x_29, x_4, x_3, x_5, x_6, x_22, x_1, x_27, x_28, x_11, x_14, x_15, x_16, x_17, x_18, x_19, x_23); lean_dec(x_1); -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_29, 1); +lean_dec(x_10); +x_31 = lean_ctor_get(x_30, 0); lean_inc(x_31); -lean_dec(x_29); -lean_inc(x_17); -x_32 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_17, x_18, x_31); -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_32, 1); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +lean_inc(x_18); +x_33 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_18, x_19, x_32); +x_34 = lean_ctor_get(x_33, 0); lean_inc(x_34); -lean_dec(x_32); -x_35 = lean_ctor_get(x_17, 10); +x_35 = lean_ctor_get(x_33, 1); lean_inc(x_35); -x_36 = lean_st_ref_get(x_18, x_34); -x_37 = lean_ctor_get(x_36, 0); -lean_inc(x_37); -x_38 = lean_ctor_get(x_36, 1); +lean_dec(x_33); +x_36 = lean_ctor_get(x_18, 10); +lean_inc(x_36); +x_37 = lean_st_ref_get(x_19, x_35); +x_38 = lean_ctor_get(x_37, 0); lean_inc(x_38); -lean_dec(x_36); -x_39 = lean_ctor_get(x_37, 0); +x_39 = lean_ctor_get(x_37, 1); lean_inc(x_39); lean_dec(x_37); -x_40 = lean_environment_main_module(x_39); -x_41 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__4; -x_42 = l_Lean_addMacroScope(x_40, x_41, x_35); -x_43 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17___closed__2; +x_40 = lean_ctor_get(x_38, 0); +lean_inc(x_40); +lean_dec(x_38); +x_41 = lean_environment_main_module(x_40); +x_42 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___closed__4; +x_43 = l_Lean_addMacroScope(x_41, x_42, x_36); +x_44 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17___closed__2; lean_inc(x_4); -x_44 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_44, 0, x_33); -lean_ctor_set(x_44, 1, x_43); -lean_ctor_set(x_44, 2, x_42); -lean_ctor_set(x_44, 3, x_4); -x_45 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16(x_2, x_3, x_4, x_11, x_5, x_21, x_6, x_7, x_8, x_9, x_30, x_44, x_13, x_14, x_15, x_16, x_17, x_18, x_38); -return x_45; +x_45 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_45, 0, x_34); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_43); +lean_ctor_set(x_45, 3, x_4); +x_46 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16(x_2, x_3, x_4, x_12, x_5, x_22, x_6, x_7, x_8, x_9, x_31, x_45, x_14, x_15, x_16, x_17, x_18, x_19, x_39); +lean_dec(x_19); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +return x_46; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -lean_dec(x_23); -x_46 = lean_unsigned_to_nat(0u); -x_47 = lean_array_fget(x_1, x_46); +lean_object* x_47; lean_object* x_48; lean_object* x_49; +lean_dec(x_24); +lean_dec(x_10); +x_47 = lean_unsigned_to_nat(0u); +x_48 = lean_array_fget(x_1, x_47); lean_dec(x_1); -x_48 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16(x_2, x_3, x_4, x_11, x_5, x_21, x_6, x_7, x_8, x_9, x_10, x_47, x_13, x_14, x_15, x_16, x_17, x_18, x_22); -return x_48; +x_49 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16(x_2, x_3, x_4, x_12, x_5, x_22, x_6, x_7, x_8, x_9, x_11, x_48, x_14, x_15, x_16, x_17, x_18, x_19, x_23); +lean_dec(x_19); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +return x_49; } } } @@ -31293,7 +32968,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__30; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__30; x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -31302,51 +32977,27 @@ return x_3; static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__6() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__5; -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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__7() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__6; -x_3 = lean_alloc_ctor(1, 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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__8() { -_start: -{ lean_object* x_1; x_1 = lean_mk_string_from_bytes("Array.getSepElems", 17); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__9() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__8; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__6; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__10() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__8; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__6; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__9; +x_3 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__7; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -31354,7 +33005,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__11() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__9() { _start: { lean_object* x_1; @@ -31362,41 +33013,17 @@ x_1 = lean_mk_string_from_bytes("getSepElems", 11); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__12() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Term_Quotation_ArrayStxBuilder_push___closed__2; -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__11; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__9; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__13() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__12; -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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__14() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__13; -x_3 = lean_alloc_ctor(1, 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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__15() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__11() { _start: { lean_object* x_1; @@ -31404,17 +33031,17 @@ x_1 = lean_mk_string_from_bytes("let_delayed", 11); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__16() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__15; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__11; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__17() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__13() { _start: { lean_object* x_1; @@ -31422,22 +33049,22 @@ x_1 = lean_mk_string_from_bytes("yes", 3); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__18() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__14() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__17; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__13; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__19() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__17; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__13; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__18; +x_3 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__14; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -31445,17 +33072,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__20() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__17; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__13; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__21() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__17() { _start: { lean_object* x_1; @@ -31463,17 +33090,17 @@ x_1 = lean_mk_string_from_bytes("simpleBinder", 12); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__22() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__21; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__17; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__23() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__19() { _start: { lean_object* x_1; @@ -31481,17 +33108,17 @@ x_1 = lean_mk_string_from_bytes("termIfThenElse", 14); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__24() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__23; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__19; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__25() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__21() { _start: { lean_object* x_1; @@ -31499,7 +33126,7 @@ x_1 = lean_mk_string_from_bytes("if", 2); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__26() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__22() { _start: { lean_object* x_1; @@ -31507,22 +33134,22 @@ x_1 = lean_mk_string_from_bytes("discr.isNone", 12); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__27() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__23() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__26; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__22; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__28() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__26; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__22; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__27; +x_3 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__23; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -31530,7 +33157,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__29() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__25() { _start: { lean_object* x_1; @@ -31538,17 +33165,17 @@ x_1 = lean_mk_string_from_bytes("isNone", 6); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__30() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__29; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__25; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__31() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__27() { _start: { lean_object* x_1; @@ -31556,7 +33183,7 @@ x_1 = lean_mk_string_from_bytes("then", 4); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__32() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__28() { _start: { lean_object* x_1; @@ -31564,1560 +33191,1611 @@ x_1 = lean_mk_string_from_bytes("else", 4); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23(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_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_11 = l_Lean_Syntax_antiquotSpliceKind_x3f(x_1); -x_12 = l_Lean_Syntax_getAntiquotSpliceContents(x_1); +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = l_Lean_Syntax_antiquotSpliceKind_x3f(x_1); +x_14 = l_Lean_Syntax_getAntiquotSpliceContents(x_1); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_13 = l_Lean_Elab_Term_Quotation_getAntiquotationIds(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_13) == 0) +x_15 = l_Lean_Elab_Term_Quotation_getAntiquotationIds(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = lean_box(0); +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_17 = lean_apply_8(x_2, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_15); -if (lean_obj_tag(x_17) == 0) +lean_inc(x_2); +x_18 = lean_apply_8(x_4, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_17); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_18, 0); lean_inc(x_19); -lean_dec(x_17); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_20 = lean_apply_7(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_19); -if (lean_obj_tag(x_20) == 0) +x_21 = lean_apply_7(x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_20); +if (lean_obj_tag(x_21) == 0) { -if (lean_obj_tag(x_11) == 0) +if (lean_obj_tag(x_13) == 0) { -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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_20, 1); +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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_22 = lean_ctor_get(x_21, 0); lean_inc(x_22); -lean_dec(x_20); -lean_inc(x_8); -x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_8, x_9, x_22); -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +lean_inc(x_10); +x_24 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_10, x_11, x_23); +x_25 = lean_ctor_get(x_24, 0); lean_inc(x_25); -lean_dec(x_23); -x_26 = lean_ctor_get(x_8, 10); +x_26 = lean_ctor_get(x_24, 1); lean_inc(x_26); -x_27 = lean_st_ref_get(x_9, x_25); -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); +lean_dec(x_24); +x_27 = lean_ctor_get(x_10, 10); +lean_inc(x_27); +x_28 = lean_st_ref_get(x_11, x_26); +x_29 = lean_ctor_get(x_28, 0); lean_inc(x_29); -lean_dec(x_27); -x_30 = lean_ctor_get(x_28, 0); +x_30 = lean_ctor_get(x_28, 1); lean_inc(x_30); lean_dec(x_28); -x_31 = lean_environment_main_module(x_30); -x_32 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__4; -lean_inc(x_26); +x_31 = lean_ctor_get(x_29, 0); lean_inc(x_31); -x_33 = l_Lean_addMacroScope(x_31, x_32, x_26); -x_34 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__3; -x_35 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__7; -lean_inc(x_24); -x_36 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_36, 0, x_24); -lean_ctor_set(x_36, 1, x_34); -lean_ctor_set(x_36, 2, x_33); -lean_ctor_set(x_36, 3, x_35); -x_37 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -lean_inc(x_26); -x_38 = l_Lean_addMacroScope(x_31, x_37, x_26); -x_39 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; -x_40 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_40, 0, x_24); -lean_ctor_set(x_40, 1, x_39); -lean_ctor_set(x_40, 2, x_38); -lean_ctor_set(x_40, 3, x_16); -x_41 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; -x_42 = lean_array_push(x_41, x_40); -x_43 = lean_box(2); -x_44 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_45 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_45, 0, x_43); -lean_ctor_set(x_45, 1, x_44); -lean_ctor_set(x_45, 2, x_42); -x_46 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_47 = lean_array_push(x_46, x_36); -x_48 = lean_array_push(x_47, x_45); -x_49 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_50 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_50, 0, x_43); -lean_ctor_set(x_50, 1, x_49); -lean_ctor_set(x_50, 2, x_48); -x_51 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__3; -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__2(x_11, x_51); -if (x_52 == 0) +lean_dec(x_29); +x_32 = lean_environment_main_module(x_31); +x_33 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__4; +lean_inc(x_27); +lean_inc(x_32); +x_34 = l_Lean_addMacroScope(x_32, x_33, x_27); +x_35 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__5; +lean_inc(x_2); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_2); +lean_inc(x_2); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_2); +x_38 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__3; +lean_inc(x_25); +x_39 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_39, 0, x_25); +lean_ctor_set(x_39, 1, x_38); +lean_ctor_set(x_39, 2, x_34); +lean_ctor_set(x_39, 3, x_37); +x_40 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +lean_inc(x_27); +x_41 = l_Lean_addMacroScope(x_32, x_40, x_27); +x_42 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; +lean_inc(x_2); +x_43 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_43, 0, x_25); +lean_ctor_set(x_43, 1, x_42); +lean_ctor_set(x_43, 2, x_41); +lean_ctor_set(x_43, 3, x_2); +x_44 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; +x_45 = lean_array_push(x_44, x_43); +x_46 = lean_box(2); +x_47 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_48 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_48, 0, x_46); +lean_ctor_set(x_48, 1, x_47); +lean_ctor_set(x_48, 2, x_45); +x_49 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_50 = lean_array_push(x_49, x_39); +x_51 = lean_array_push(x_50, x_48); +x_52 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_53 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_53, 0, x_46); +lean_ctor_set(x_53, 1, x_52); +lean_ctor_set(x_53, 2, x_51); +x_54 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__3; +x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5(x_13, x_54); +if (x_55 == 0) { -lean_object* x_53; lean_object* x_54; lean_object* x_55; -lean_dec(x_26); -x_53 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_54 = lean_box(0); -x_55 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17(x_14, x_53, x_44, x_16, x_41, x_46, x_49, x_21, x_12, x_18, x_50, x_54, x_4, x_5, x_6, x_7, x_8, x_9, x_29); -lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_55; +lean_object* x_56; lean_object* x_57; lean_object* x_58; +lean_dec(x_27); +x_56 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; +x_57 = lean_box(0); +x_58 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17(x_16, x_56, x_47, x_2, x_44, x_49, x_52, x_22, x_14, x_3, x_19, x_53, x_57, x_6, x_7, x_8, x_9, x_10, x_11, x_30); +return x_58; } else { -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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_inc(x_8); -x_56 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_8, x_9, x_29); -x_57 = lean_ctor_get(x_56, 0); -lean_inc(x_57); -x_58 = lean_ctor_get(x_56, 1); -lean_inc(x_58); -lean_dec(x_56); -x_59 = lean_st_ref_get(x_9, x_58); +lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; +lean_inc(x_10); +x_59 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_10, x_11, x_30); x_60 = lean_ctor_get(x_59, 0); lean_inc(x_60); x_61 = lean_ctor_get(x_59, 1); lean_inc(x_61); lean_dec(x_59); -x_62 = lean_ctor_get(x_60, 0); -lean_inc(x_62); -lean_dec(x_60); -x_63 = lean_environment_main_module(x_62); -x_64 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__12; -x_65 = l_Lean_addMacroScope(x_63, x_64, x_26); -x_66 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__10; -x_67 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__14; -x_68 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_68, 0, x_57); -lean_ctor_set(x_68, 1, x_66); -lean_ctor_set(x_68, 2, x_65); -lean_ctor_set(x_68, 3, x_67); -x_69 = lean_array_push(x_41, x_50); -x_70 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_70, 0, x_43); -lean_ctor_set(x_70, 1, x_44); -lean_ctor_set(x_70, 2, x_69); -x_71 = lean_array_push(x_46, x_68); -x_72 = lean_array_push(x_71, x_70); -x_73 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_73, 0, x_43); -lean_ctor_set(x_73, 1, x_49); -lean_ctor_set(x_73, 2, x_72); -x_74 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_75 = lean_box(0); -x_76 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17(x_14, x_74, x_44, x_16, x_41, x_46, x_49, x_21, x_12, x_18, x_73, x_75, x_4, x_5, x_6, x_7, x_8, x_9, x_61); -lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_76; +x_62 = lean_st_ref_get(x_11, x_61); +x_63 = lean_ctor_get(x_62, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_62, 1); +lean_inc(x_64); +lean_dec(x_62); +x_65 = lean_ctor_get(x_63, 0); +lean_inc(x_65); +lean_dec(x_63); +x_66 = lean_environment_main_module(x_65); +x_67 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__10; +x_68 = l_Lean_addMacroScope(x_66, x_67, x_27); +lean_inc(x_2); +x_69 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_69, 0, x_67); +lean_ctor_set(x_69, 1, x_2); +lean_inc(x_2); +x_70 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set(x_70, 1, x_2); +x_71 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__8; +x_72 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_72, 0, x_60); +lean_ctor_set(x_72, 1, x_71); +lean_ctor_set(x_72, 2, x_68); +lean_ctor_set(x_72, 3, x_70); +x_73 = lean_array_push(x_44, x_53); +x_74 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_74, 0, x_46); +lean_ctor_set(x_74, 1, x_47); +lean_ctor_set(x_74, 2, x_73); +x_75 = lean_array_push(x_49, x_72); +x_76 = lean_array_push(x_75, x_74); +x_77 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_77, 0, x_46); +lean_ctor_set(x_77, 1, x_52); +lean_ctor_set(x_77, 2, x_76); +x_78 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; +x_79 = lean_box(0); +x_80 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17(x_16, x_78, x_47, x_2, x_44, x_49, x_52, x_22, x_14, x_3, x_19, x_77, x_79, x_6, x_7, x_8, x_9, x_10, x_11, x_64); +return x_80; } } else { -lean_object* x_77; -x_77 = lean_ctor_get(x_11, 0); -lean_inc(x_77); -switch (lean_obj_tag(x_77)) { +lean_object* x_81; +x_81 = lean_ctor_get(x_13, 0); +lean_inc(x_81); +switch (lean_obj_tag(x_81)) { case 0: { -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; uint8_t x_109; -x_78 = lean_ctor_get(x_20, 0); -lean_inc(x_78); -x_79 = lean_ctor_get(x_20, 1); -lean_inc(x_79); -lean_dec(x_20); -lean_inc(x_8); -x_80 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_8, x_9, x_79); -x_81 = lean_ctor_get(x_80, 0); -lean_inc(x_81); -x_82 = lean_ctor_get(x_80, 1); +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; uint8_t x_115; +x_82 = lean_ctor_get(x_21, 0); lean_inc(x_82); -lean_dec(x_80); -x_83 = lean_ctor_get(x_8, 10); +x_83 = lean_ctor_get(x_21, 1); lean_inc(x_83); -x_84 = lean_st_ref_get(x_9, x_82); +lean_dec(x_21); +lean_inc(x_10); +x_84 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_10, x_11, x_83); x_85 = lean_ctor_get(x_84, 0); lean_inc(x_85); x_86 = lean_ctor_get(x_84, 1); lean_inc(x_86); lean_dec(x_84); -x_87 = lean_ctor_get(x_85, 0); +x_87 = lean_ctor_get(x_10, 10); lean_inc(x_87); -lean_dec(x_85); -x_88 = lean_environment_main_module(x_87); -x_89 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__4; -lean_inc(x_83); -lean_inc(x_88); -x_90 = l_Lean_addMacroScope(x_88, x_89, x_83); -x_91 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__3; -x_92 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__7; -lean_inc(x_81); -x_93 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_93, 0, x_81); -lean_ctor_set(x_93, 1, x_91); -lean_ctor_set(x_93, 2, x_90); -lean_ctor_set(x_93, 3, x_92); -x_94 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -lean_inc(x_83); -x_95 = l_Lean_addMacroScope(x_88, x_94, x_83); -x_96 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; -x_97 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_97, 0, x_81); -lean_ctor_set(x_97, 1, x_96); -lean_ctor_set(x_97, 2, x_95); -lean_ctor_set(x_97, 3, x_16); -x_98 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; -x_99 = lean_array_push(x_98, x_97); -x_100 = lean_box(2); -x_101 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_102 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_102, 0, x_100); -lean_ctor_set(x_102, 1, x_101); -lean_ctor_set(x_102, 2, x_99); -x_103 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_104 = lean_array_push(x_103, x_93); -x_105 = lean_array_push(x_104, x_102); -x_106 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_107 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_107, 0, x_100); -lean_ctor_set(x_107, 1, x_106); -lean_ctor_set(x_107, 2, x_105); -x_108 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__3; -x_109 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__2(x_11, x_108); -lean_dec(x_11); -if (x_109 == 0) +x_88 = lean_st_ref_get(x_11, x_86); +x_89 = lean_ctor_get(x_88, 0); +lean_inc(x_89); +x_90 = lean_ctor_get(x_88, 1); +lean_inc(x_90); +lean_dec(x_88); +x_91 = lean_ctor_get(x_89, 0); +lean_inc(x_91); +lean_dec(x_89); +x_92 = lean_environment_main_module(x_91); +x_93 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__4; +lean_inc(x_87); +lean_inc(x_92); +x_94 = l_Lean_addMacroScope(x_92, x_93, x_87); +x_95 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__5; +lean_inc(x_2); +x_96 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_96, 0, x_95); +lean_ctor_set(x_96, 1, x_2); +lean_inc(x_2); +x_97 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_97, 0, x_96); +lean_ctor_set(x_97, 1, x_2); +x_98 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__3; +lean_inc(x_85); +x_99 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_99, 0, x_85); +lean_ctor_set(x_99, 1, x_98); +lean_ctor_set(x_99, 2, x_94); +lean_ctor_set(x_99, 3, x_97); +x_100 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +lean_inc(x_87); +x_101 = l_Lean_addMacroScope(x_92, x_100, x_87); +x_102 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; +lean_inc(x_2); +x_103 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_103, 0, x_85); +lean_ctor_set(x_103, 1, x_102); +lean_ctor_set(x_103, 2, x_101); +lean_ctor_set(x_103, 3, x_2); +x_104 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; +x_105 = lean_array_push(x_104, x_103); +x_106 = lean_box(2); +x_107 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_108 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_108, 0, x_106); +lean_ctor_set(x_108, 1, x_107); +lean_ctor_set(x_108, 2, x_105); +x_109 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_110 = lean_array_push(x_109, x_99); +x_111 = lean_array_push(x_110, x_108); +x_112 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_113 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_113, 0, x_106); +lean_ctor_set(x_113, 1, x_112); +lean_ctor_set(x_113, 2, x_111); +x_114 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__3; +x_115 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5(x_13, x_114); +lean_dec(x_13); +if (x_115 == 0) { -lean_object* x_110; lean_object* x_111; lean_object* x_112; -lean_dec(x_83); -x_110 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_111 = lean_box(0); -x_112 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__18(x_14, x_110, x_101, x_16, x_98, x_103, x_106, x_78, x_12, x_18, x_107, x_111, x_4, x_5, x_6, x_7, x_8, x_9, x_86); -lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_112; +lean_object* x_116; lean_object* x_117; lean_object* x_118; +lean_dec(x_87); +x_116 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; +x_117 = lean_box(0); +x_118 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__18(x_16, x_116, x_107, x_2, x_104, x_109, x_112, x_82, x_14, x_3, x_19, x_113, x_117, x_6, x_7, x_8, x_9, x_10, x_11, x_90); +return x_118; } else { -lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; -lean_inc(x_8); -x_113 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_8, x_9, x_86); -x_114 = lean_ctor_get(x_113, 0); -lean_inc(x_114); -x_115 = lean_ctor_get(x_113, 1); -lean_inc(x_115); -lean_dec(x_113); -x_116 = lean_st_ref_get(x_9, x_115); -x_117 = lean_ctor_get(x_116, 0); -lean_inc(x_117); -x_118 = lean_ctor_get(x_116, 1); -lean_inc(x_118); -lean_dec(x_116); -x_119 = lean_ctor_get(x_117, 0); -lean_inc(x_119); -lean_dec(x_117); -x_120 = lean_environment_main_module(x_119); -x_121 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__12; -x_122 = l_Lean_addMacroScope(x_120, x_121, x_83); -x_123 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__10; -x_124 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__14; -x_125 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_125, 0, x_114); -lean_ctor_set(x_125, 1, x_123); -lean_ctor_set(x_125, 2, x_122); -lean_ctor_set(x_125, 3, x_124); -x_126 = lean_array_push(x_98, x_107); -x_127 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_127, 0, x_100); -lean_ctor_set(x_127, 1, x_101); -lean_ctor_set(x_127, 2, x_126); -x_128 = lean_array_push(x_103, x_125); -x_129 = lean_array_push(x_128, x_127); -x_130 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_130, 0, x_100); -lean_ctor_set(x_130, 1, x_106); -lean_ctor_set(x_130, 2, x_129); -x_131 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_132 = lean_box(0); -x_133 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__18(x_14, x_131, x_101, x_16, x_98, x_103, x_106, x_78, x_12, x_18, x_130, x_132, x_4, x_5, x_6, x_7, x_8, x_9, x_118); -lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_133; +lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; +lean_inc(x_10); +x_119 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_10, x_11, x_90); +x_120 = lean_ctor_get(x_119, 0); +lean_inc(x_120); +x_121 = lean_ctor_get(x_119, 1); +lean_inc(x_121); +lean_dec(x_119); +x_122 = lean_st_ref_get(x_11, x_121); +x_123 = lean_ctor_get(x_122, 0); +lean_inc(x_123); +x_124 = lean_ctor_get(x_122, 1); +lean_inc(x_124); +lean_dec(x_122); +x_125 = lean_ctor_get(x_123, 0); +lean_inc(x_125); +lean_dec(x_123); +x_126 = lean_environment_main_module(x_125); +x_127 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__10; +x_128 = l_Lean_addMacroScope(x_126, x_127, x_87); +lean_inc(x_2); +x_129 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_129, 0, x_127); +lean_ctor_set(x_129, 1, x_2); +lean_inc(x_2); +x_130 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_130, 0, x_129); +lean_ctor_set(x_130, 1, x_2); +x_131 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__8; +x_132 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_132, 0, x_120); +lean_ctor_set(x_132, 1, x_131); +lean_ctor_set(x_132, 2, x_128); +lean_ctor_set(x_132, 3, x_130); +x_133 = lean_array_push(x_104, x_113); +x_134 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_134, 0, x_106); +lean_ctor_set(x_134, 1, x_107); +lean_ctor_set(x_134, 2, x_133); +x_135 = lean_array_push(x_109, x_132); +x_136 = lean_array_push(x_135, x_134); +x_137 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_137, 0, x_106); +lean_ctor_set(x_137, 1, x_112); +lean_ctor_set(x_137, 2, x_136); +x_138 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; +x_139 = lean_box(0); +x_140 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__18(x_16, x_138, x_107, x_2, x_104, x_109, x_112, x_82, x_14, x_3, x_19, x_137, x_139, x_6, x_7, x_8, x_9, x_10, x_11, x_124); +return x_140; } } case 1: { -lean_object* x_134; -x_134 = lean_ctor_get(x_77, 0); -lean_inc(x_134); -switch (lean_obj_tag(x_134)) { +lean_object* x_141; +x_141 = lean_ctor_get(x_81, 0); +lean_inc(x_141); +switch (lean_obj_tag(x_141)) { case 0: { -lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; uint8_t x_139; -x_135 = lean_ctor_get(x_20, 0); -lean_inc(x_135); -x_136 = lean_ctor_get(x_20, 1); -lean_inc(x_136); -lean_dec(x_20); -x_137 = lean_ctor_get(x_77, 1); -lean_inc(x_137); -lean_dec(x_77); -x_138 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__15; -x_139 = lean_string_dec_eq(x_137, x_138); -lean_dec(x_137); -if (x_139 == 0) -{ -lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; uint8_t x_169; -lean_inc(x_8); -x_140 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_8, x_9, x_136); -x_141 = lean_ctor_get(x_140, 0); -lean_inc(x_141); -x_142 = lean_ctor_get(x_140, 1); +lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; uint8_t x_146; +x_142 = lean_ctor_get(x_21, 0); lean_inc(x_142); -lean_dec(x_140); -x_143 = lean_ctor_get(x_8, 10); +x_143 = lean_ctor_get(x_21, 1); lean_inc(x_143); -x_144 = lean_st_ref_get(x_9, x_142); -x_145 = lean_ctor_get(x_144, 0); -lean_inc(x_145); -x_146 = lean_ctor_get(x_144, 1); -lean_inc(x_146); +lean_dec(x_21); +x_144 = lean_ctor_get(x_81, 1); +lean_inc(x_144); +lean_dec(x_81); +x_145 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__15; +x_146 = lean_string_dec_eq(x_144, x_145); lean_dec(x_144); -x_147 = lean_ctor_get(x_145, 0); -lean_inc(x_147); -lean_dec(x_145); -x_148 = lean_environment_main_module(x_147); -x_149 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__4; -lean_inc(x_143); +if (x_146 == 0) +{ +lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; uint8_t x_178; +lean_inc(x_10); +x_147 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_10, x_11, x_143); +x_148 = lean_ctor_get(x_147, 0); lean_inc(x_148); -x_150 = l_Lean_addMacroScope(x_148, x_149, x_143); -x_151 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__3; -x_152 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__7; -lean_inc(x_141); -x_153 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_153, 0, x_141); -lean_ctor_set(x_153, 1, x_151); -lean_ctor_set(x_153, 2, x_150); -lean_ctor_set(x_153, 3, x_152); -x_154 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -lean_inc(x_143); -x_155 = l_Lean_addMacroScope(x_148, x_154, x_143); -x_156 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; -x_157 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_157, 0, x_141); -lean_ctor_set(x_157, 1, x_156); -lean_ctor_set(x_157, 2, x_155); -lean_ctor_set(x_157, 3, x_16); -x_158 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; -x_159 = lean_array_push(x_158, x_157); -x_160 = lean_box(2); -x_161 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_162 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_162, 0, x_160); +x_149 = lean_ctor_get(x_147, 1); +lean_inc(x_149); +lean_dec(x_147); +x_150 = lean_ctor_get(x_10, 10); +lean_inc(x_150); +x_151 = lean_st_ref_get(x_11, x_149); +x_152 = lean_ctor_get(x_151, 0); +lean_inc(x_152); +x_153 = lean_ctor_get(x_151, 1); +lean_inc(x_153); +lean_dec(x_151); +x_154 = lean_ctor_get(x_152, 0); +lean_inc(x_154); +lean_dec(x_152); +x_155 = lean_environment_main_module(x_154); +x_156 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__4; +lean_inc(x_150); +lean_inc(x_155); +x_157 = l_Lean_addMacroScope(x_155, x_156, x_150); +x_158 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__5; +lean_inc(x_2); +x_159 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_159, 0, x_158); +lean_ctor_set(x_159, 1, x_2); +lean_inc(x_2); +x_160 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_160, 0, x_159); +lean_ctor_set(x_160, 1, x_2); +x_161 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__3; +lean_inc(x_148); +x_162 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_162, 0, x_148); lean_ctor_set(x_162, 1, x_161); -lean_ctor_set(x_162, 2, x_159); -x_163 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_164 = lean_array_push(x_163, x_153); -x_165 = lean_array_push(x_164, x_162); -x_166 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_167 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_167, 0, x_160); -lean_ctor_set(x_167, 1, x_166); -lean_ctor_set(x_167, 2, x_165); -x_168 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__3; -x_169 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__2(x_11, x_168); -lean_dec(x_11); -if (x_169 == 0) +lean_ctor_set(x_162, 2, x_157); +lean_ctor_set(x_162, 3, x_160); +x_163 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +lean_inc(x_150); +x_164 = l_Lean_addMacroScope(x_155, x_163, x_150); +x_165 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; +lean_inc(x_2); +x_166 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_166, 0, x_148); +lean_ctor_set(x_166, 1, x_165); +lean_ctor_set(x_166, 2, x_164); +lean_ctor_set(x_166, 3, x_2); +x_167 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; +x_168 = lean_array_push(x_167, x_166); +x_169 = lean_box(2); +x_170 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_171 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_171, 0, x_169); +lean_ctor_set(x_171, 1, x_170); +lean_ctor_set(x_171, 2, x_168); +x_172 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_173 = lean_array_push(x_172, x_162); +x_174 = lean_array_push(x_173, x_171); +x_175 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_176 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_176, 0, x_169); +lean_ctor_set(x_176, 1, x_175); +lean_ctor_set(x_176, 2, x_174); +x_177 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__3; +x_178 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5(x_13, x_177); +lean_dec(x_13); +if (x_178 == 0) { -lean_object* x_170; lean_object* x_171; lean_object* x_172; -lean_dec(x_143); -x_170 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_171 = lean_box(0); -x_172 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__19(x_14, x_170, x_161, x_16, x_158, x_163, x_166, x_135, x_12, x_18, x_167, x_171, x_4, x_5, x_6, x_7, x_8, x_9, x_146); -lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_172; +lean_object* x_179; lean_object* x_180; lean_object* x_181; +lean_dec(x_150); +x_179 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; +x_180 = lean_box(0); +x_181 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__19(x_16, x_179, x_170, x_2, x_167, x_172, x_175, x_142, x_14, x_3, x_19, x_176, x_180, x_6, x_7, x_8, x_9, x_10, x_11, x_153); +return x_181; } else { -lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; -lean_inc(x_8); -x_173 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_8, x_9, x_146); -x_174 = lean_ctor_get(x_173, 0); -lean_inc(x_174); -x_175 = lean_ctor_get(x_173, 1); -lean_inc(x_175); -lean_dec(x_173); -x_176 = lean_st_ref_get(x_9, x_175); -x_177 = lean_ctor_get(x_176, 0); -lean_inc(x_177); -x_178 = lean_ctor_get(x_176, 1); -lean_inc(x_178); -lean_dec(x_176); -x_179 = lean_ctor_get(x_177, 0); -lean_inc(x_179); -lean_dec(x_177); -x_180 = lean_environment_main_module(x_179); -x_181 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__12; -x_182 = l_Lean_addMacroScope(x_180, x_181, x_143); -x_183 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__10; -x_184 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__14; -x_185 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_185, 0, x_174); -lean_ctor_set(x_185, 1, x_183); -lean_ctor_set(x_185, 2, x_182); -lean_ctor_set(x_185, 3, x_184); -x_186 = lean_array_push(x_158, x_167); -x_187 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_187, 0, x_160); -lean_ctor_set(x_187, 1, x_161); -lean_ctor_set(x_187, 2, x_186); -x_188 = lean_array_push(x_163, x_185); -x_189 = lean_array_push(x_188, x_187); -x_190 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_190, 0, x_160); -lean_ctor_set(x_190, 1, x_166); -lean_ctor_set(x_190, 2, x_189); -x_191 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_192 = lean_box(0); -x_193 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__19(x_14, x_191, x_161, x_16, x_158, x_163, x_166, x_135, x_12, x_18, x_190, x_192, x_4, x_5, x_6, x_7, x_8, x_9, x_178); -lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_193; +lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; +lean_inc(x_10); +x_182 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_10, x_11, x_153); +x_183 = lean_ctor_get(x_182, 0); +lean_inc(x_183); +x_184 = lean_ctor_get(x_182, 1); +lean_inc(x_184); +lean_dec(x_182); +x_185 = lean_st_ref_get(x_11, x_184); +x_186 = lean_ctor_get(x_185, 0); +lean_inc(x_186); +x_187 = lean_ctor_get(x_185, 1); +lean_inc(x_187); +lean_dec(x_185); +x_188 = lean_ctor_get(x_186, 0); +lean_inc(x_188); +lean_dec(x_186); +x_189 = lean_environment_main_module(x_188); +x_190 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__10; +x_191 = l_Lean_addMacroScope(x_189, x_190, x_150); +lean_inc(x_2); +x_192 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_192, 0, x_190); +lean_ctor_set(x_192, 1, x_2); +lean_inc(x_2); +x_193 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_193, 0, x_192); +lean_ctor_set(x_193, 1, x_2); +x_194 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__8; +x_195 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_195, 0, x_183); +lean_ctor_set(x_195, 1, x_194); +lean_ctor_set(x_195, 2, x_191); +lean_ctor_set(x_195, 3, x_193); +x_196 = lean_array_push(x_167, x_176); +x_197 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_197, 0, x_169); +lean_ctor_set(x_197, 1, x_170); +lean_ctor_set(x_197, 2, x_196); +x_198 = lean_array_push(x_172, x_195); +x_199 = lean_array_push(x_198, x_197); +x_200 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_200, 0, x_169); +lean_ctor_set(x_200, 1, x_175); +lean_ctor_set(x_200, 2, x_199); +x_201 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; +x_202 = lean_box(0); +x_203 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__19(x_16, x_201, x_170, x_2, x_167, x_172, x_175, x_142, x_14, x_3, x_19, x_200, x_202, x_6, x_7, x_8, x_9, x_10, x_11, x_187); +return x_203; } } else { -lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; uint8_t x_214; -lean_dec(x_11); +lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; uint8_t x_226; +lean_dec(x_13); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_inc(x_8); -x_194 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_8, x_9, x_136); -x_195 = lean_ctor_get(x_194, 0); -lean_inc(x_195); -x_196 = lean_ctor_get(x_194, 1); -lean_inc(x_196); -lean_dec(x_194); -x_197 = lean_ctor_get(x_8, 10); -lean_inc(x_197); -x_198 = lean_st_ref_get(x_9, x_196); -x_199 = lean_ctor_get(x_198, 0); -lean_inc(x_199); -x_200 = lean_ctor_get(x_198, 1); -lean_inc(x_200); -lean_dec(x_198); -x_201 = lean_ctor_get(x_199, 0); -lean_inc(x_201); -lean_dec(x_199); -x_202 = lean_environment_main_module(x_201); -x_203 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__76; -lean_inc(x_197); -x_204 = l_Lean_addMacroScope(x_202, x_203, x_197); -x_205 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__75; -x_206 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__79; -x_207 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_207, 0, x_195); -lean_ctor_set(x_207, 1, x_205); -lean_ctor_set(x_207, 2, x_204); -lean_ctor_set(x_207, 3, x_206); -x_208 = lean_array_get_size(x_14); -lean_inc(x_208); -x_209 = lean_mk_array(x_208, x_207); -x_210 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_8, x_9, x_200); -x_211 = lean_ctor_get(x_210, 0); -lean_inc(x_211); -x_212 = lean_ctor_get(x_210, 1); -lean_inc(x_212); -lean_dec(x_210); -x_213 = lean_st_ref_get(x_9, x_212); -lean_dec(x_9); -x_214 = !lean_is_exclusive(x_213); -if (x_214 == 0) -{ -lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; size_t x_278; size_t x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; size_t x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; -x_215 = lean_ctor_get(x_213, 0); -x_216 = lean_ctor_get(x_215, 0); -lean_inc(x_216); -lean_dec(x_215); -x_217 = lean_environment_main_module(x_216); -x_218 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__15; +lean_dec(x_3); +lean_inc(x_10); +x_204 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_10, x_11, x_143); +x_205 = lean_ctor_get(x_204, 0); +lean_inc(x_205); +x_206 = lean_ctor_get(x_204, 1); +lean_inc(x_206); +lean_dec(x_204); +x_207 = lean_ctor_get(x_10, 10); +lean_inc(x_207); +x_208 = lean_st_ref_get(x_11, x_206); +x_209 = lean_ctor_get(x_208, 0); +lean_inc(x_209); +x_210 = lean_ctor_get(x_208, 1); +lean_inc(x_210); +lean_dec(x_208); +x_211 = lean_ctor_get(x_209, 0); lean_inc(x_211); -x_219 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_219, 0, x_211); +lean_dec(x_209); +x_212 = lean_environment_main_module(x_211); +x_213 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__76; +lean_inc(x_207); +x_214 = l_Lean_addMacroScope(x_212, x_213, x_207); +x_215 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__77; +lean_inc(x_2); +x_216 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_216, 0, x_215); +lean_ctor_set(x_216, 1, x_2); +lean_inc(x_2); +x_217 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_217, 0, x_216); +lean_ctor_set(x_217, 1, x_2); +x_218 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__75; +x_219 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_219, 0, x_205); lean_ctor_set(x_219, 1, x_218); -x_220 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__20; -lean_inc(x_197); -lean_inc(x_217); -x_221 = l_Lean_addMacroScope(x_217, x_220, x_197); -x_222 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__19; -lean_inc(x_211); -x_223 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_223, 0, x_211); -lean_ctor_set(x_223, 1, x_222); -lean_ctor_set(x_223, 2, x_221); -lean_ctor_set(x_223, 3, x_16); -x_224 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__2; -lean_inc(x_211); -x_225 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_225, 0, x_211); -lean_ctor_set(x_225, 1, x_224); -x_226 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; -x_227 = lean_array_push(x_226, x_225); -x_228 = lean_box(2); -x_229 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_noOpMatchAdaptPats___lambda__1___closed__1; -x_230 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_230, 0, x_228); -lean_ctor_set(x_230, 1, x_229); -lean_ctor_set(x_230, 2, x_227); -x_231 = lean_array_push(x_226, x_230); -lean_inc(x_14); -lean_inc(x_231); -x_232 = l_Array_append___rarg(x_231, x_14); -x_233 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_234 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_234, 0, x_228); -lean_ctor_set(x_234, 1, x_233); -lean_ctor_set(x_234, 2, x_232); -x_235 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_236 = lean_array_push(x_235, x_234); -x_237 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; -x_238 = lean_array_push(x_236, x_237); -x_239 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__22; -x_240 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_240, 0, x_228); -lean_ctor_set(x_240, 1, x_239); -lean_ctor_set(x_240, 2, x_238); -x_241 = lean_array_push(x_226, x_240); -x_242 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_242, 0, x_228); -lean_ctor_set(x_242, 1, x_233); -lean_ctor_set(x_242, 2, x_241); -x_243 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__17; -lean_inc(x_211); -x_244 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_244, 0, x_211); -lean_ctor_set(x_244, 1, x_243); -x_245 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; +lean_ctor_set(x_219, 2, x_214); +lean_ctor_set(x_219, 3, x_217); +x_220 = lean_array_get_size(x_16); +lean_inc(x_220); +x_221 = lean_mk_array(x_220, x_219); +x_222 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_10, x_11, x_210); +x_223 = lean_ctor_get(x_222, 0); +lean_inc(x_223); +x_224 = lean_ctor_get(x_222, 1); +lean_inc(x_224); +lean_dec(x_222); +x_225 = lean_st_ref_get(x_11, x_224); +lean_dec(x_11); +x_226 = !lean_is_exclusive(x_225); +if (x_226 == 0) +{ +lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; size_t x_290; size_t x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; size_t x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; +x_227 = lean_ctor_get(x_225, 0); +x_228 = lean_ctor_get(x_227, 0); +lean_inc(x_228); +lean_dec(x_227); +x_229 = lean_environment_main_module(x_228); +x_230 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__11; +lean_inc(x_223); +x_231 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_231, 0, x_223); +lean_ctor_set(x_231, 1, x_230); +x_232 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__16; +lean_inc(x_207); +lean_inc(x_229); +x_233 = l_Lean_addMacroScope(x_229, x_232, x_207); +x_234 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__15; +lean_inc(x_2); +lean_inc(x_223); +x_235 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_235, 0, x_223); +lean_ctor_set(x_235, 1, x_234); +lean_ctor_set(x_235, 2, x_233); +lean_ctor_set(x_235, 3, x_2); +x_236 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__14___closed__2; lean_inc(x_223); -x_246 = lean_array_push(x_245, x_223); -x_247 = lean_array_push(x_246, x_242); -x_248 = lean_array_push(x_247, x_237); -x_249 = lean_array_push(x_248, x_244); -x_250 = lean_array_push(x_249, x_18); -x_251 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__12; +x_237 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_237, 0, x_223); +lean_ctor_set(x_237, 1, x_236); +x_238 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; +x_239 = lean_array_push(x_238, x_237); +x_240 = lean_box(2); +x_241 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_noOpMatchAdaptPats___lambda__1___closed__1; +x_242 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_242, 0, x_240); +lean_ctor_set(x_242, 1, x_241); +lean_ctor_set(x_242, 2, x_239); +x_243 = lean_array_push(x_238, x_242); +lean_inc(x_16); +lean_inc(x_243); +x_244 = l_Array_append___rarg(x_243, x_16); +x_245 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_246 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_246, 0, x_240); +lean_ctor_set(x_246, 1, x_245); +lean_ctor_set(x_246, 2, x_244); +x_247 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_248 = lean_array_push(x_247, x_246); +x_249 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; +x_250 = lean_array_push(x_248, x_249); +x_251 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__18; x_252 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_252, 0, x_228); +lean_ctor_set(x_252, 0, x_240); lean_ctor_set(x_252, 1, x_251); lean_ctor_set(x_252, 2, x_250); -x_253 = lean_array_push(x_226, x_252); -x_254 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__10; -x_255 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_255, 0, x_228); -lean_ctor_set(x_255, 1, x_254); -lean_ctor_set(x_255, 2, x_253); -x_256 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__20; -lean_inc(x_211); -x_257 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_257, 0, x_211); -lean_ctor_set(x_257, 1, x_256); -x_258 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__25; -lean_inc(x_211); -x_259 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_259, 0, x_211); -lean_ctor_set(x_259, 1, x_258); -x_260 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__30; -lean_inc(x_197); -lean_inc(x_217); -x_261 = l_Lean_addMacroScope(x_217, x_260, x_197); -x_262 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__28; -lean_inc(x_211); -x_263 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_263, 0, x_211); -lean_ctor_set(x_263, 1, x_262); -lean_ctor_set(x_263, 2, x_261); -lean_ctor_set(x_263, 3, x_16); -x_264 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__31; -lean_inc(x_211); -x_265 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_265, 0, x_211); -lean_ctor_set(x_265, 1, x_264); -x_266 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; -lean_inc(x_211); -x_267 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_267, 0, x_211); +x_253 = lean_array_push(x_238, x_252); +x_254 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_254, 0, x_240); +lean_ctor_set(x_254, 1, x_245); +lean_ctor_set(x_254, 2, x_253); +x_255 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__17; +lean_inc(x_223); +x_256 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_256, 0, x_223); +lean_ctor_set(x_256, 1, x_255); +x_257 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; +lean_inc(x_235); +x_258 = lean_array_push(x_257, x_235); +x_259 = lean_array_push(x_258, x_254); +x_260 = lean_array_push(x_259, x_249); +x_261 = lean_array_push(x_260, x_256); +x_262 = lean_array_push(x_261, x_19); +x_263 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__12; +x_264 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_264, 0, x_240); +lean_ctor_set(x_264, 1, x_263); +lean_ctor_set(x_264, 2, x_262); +x_265 = lean_array_push(x_238, x_264); +x_266 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__10; +x_267 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_267, 0, x_240); lean_ctor_set(x_267, 1, x_266); -x_268 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; -lean_inc(x_211); +lean_ctor_set(x_267, 2, x_265); +x_268 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__20; +lean_inc(x_223); x_269 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_269, 0, x_211); +lean_ctor_set(x_269, 0, x_223); lean_ctor_set(x_269, 1, x_268); -x_270 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; -x_271 = lean_array_push(x_270, x_267); -lean_inc(x_271); -x_272 = lean_array_push(x_271, x_237); -lean_inc(x_269); -x_273 = lean_array_push(x_272, x_269); -x_274 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__6; -x_275 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_275, 0, x_228); +x_270 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__21; +lean_inc(x_223); +x_271 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_271, 0, x_223); +lean_ctor_set(x_271, 1, x_270); +x_272 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__26; +lean_inc(x_207); +lean_inc(x_229); +x_273 = l_Lean_addMacroScope(x_229, x_272, x_207); +x_274 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__24; +lean_inc(x_2); +lean_inc(x_223); +x_275 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_275, 0, x_223); lean_ctor_set(x_275, 1, x_274); lean_ctor_set(x_275, 2, x_273); -x_276 = lean_array_push(x_226, x_275); -x_277 = lean_array_get_size(x_209); -x_278 = lean_usize_of_nat(x_277); -lean_dec(x_277); -x_279 = 0; -x_280 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_278, x_279, x_209); -lean_inc(x_276); -x_281 = l_Array_append___rarg(x_276, x_280); -x_282 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_282, 0, x_228); -lean_ctor_set(x_282, 1, x_233); -lean_ctor_set(x_282, 2, x_281); -x_283 = lean_array_push(x_235, x_223); +lean_ctor_set(x_275, 3, x_2); +x_276 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__27; +lean_inc(x_223); +x_277 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_277, 0, x_223); +lean_ctor_set(x_277, 1, x_276); +x_278 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; +lean_inc(x_223); +x_279 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_279, 0, x_223); +lean_ctor_set(x_279, 1, x_278); +x_280 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; +lean_inc(x_223); +x_281 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_281, 0, x_223); +lean_ctor_set(x_281, 1, x_280); +x_282 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; +x_283 = lean_array_push(x_282, x_279); lean_inc(x_283); -x_284 = lean_array_push(x_283, x_282); -x_285 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_286 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_286, 0, x_228); -lean_ctor_set(x_286, 1, x_285); -lean_ctor_set(x_286, 2, x_284); -x_287 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__32; -lean_inc(x_211); -x_288 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_288, 0, x_211); -lean_ctor_set(x_288, 1, x_287); -x_289 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__16; -lean_inc(x_211); -x_290 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_290, 0, x_211); -lean_ctor_set(x_290, 1, x_289); -x_291 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -lean_inc(x_197); -lean_inc(x_217); -x_292 = l_Lean_addMacroScope(x_217, x_291, x_197); -x_293 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; -lean_inc(x_211); -x_294 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_294, 0, x_211); -lean_ctor_set(x_294, 1, x_293); -lean_ctor_set(x_294, 2, x_292); -lean_ctor_set(x_294, 3, x_16); -x_295 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__63; +x_284 = lean_array_push(x_283, x_249); +lean_inc(x_281); +x_285 = lean_array_push(x_284, x_281); +x_286 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6; +x_287 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_287, 0, x_240); +lean_ctor_set(x_287, 1, x_286); +lean_ctor_set(x_287, 2, x_285); +x_288 = lean_array_push(x_238, x_287); +x_289 = lean_array_get_size(x_221); +x_290 = lean_usize_of_nat(x_289); +lean_dec(x_289); +x_291 = 0; +x_292 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_290, x_291, x_221); +lean_inc(x_288); +x_293 = l_Array_append___rarg(x_288, x_292); +x_294 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_294, 0, x_240); +lean_ctor_set(x_294, 1, x_245); +lean_ctor_set(x_294, 2, x_293); +x_295 = lean_array_push(x_247, x_235); +lean_inc(x_295); x_296 = lean_array_push(x_295, x_294); -x_297 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__46; +x_297 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; x_298 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_298, 0, x_228); +lean_ctor_set(x_298, 0, x_240); lean_ctor_set(x_298, 1, x_297); lean_ctor_set(x_298, 2, x_296); -x_299 = lean_array_push(x_226, x_298); -x_300 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_300, 0, x_228); -lean_ctor_set(x_300, 1, x_233); -lean_ctor_set(x_300, 2, x_299); -x_301 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__20; -lean_inc(x_211); +x_299 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__28; +lean_inc(x_223); +x_300 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_300, 0, x_223); +lean_ctor_set(x_300, 1, x_299); +x_301 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__16; +lean_inc(x_223); x_302 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_302, 0, x_211); +lean_ctor_set(x_302, 0, x_223); lean_ctor_set(x_302, 1, x_301); -x_303 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__25; -lean_inc(x_211); -x_304 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_304, 0, x_211); -lean_ctor_set(x_304, 1, x_303); -x_305 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__3___lambda__1___closed__1; -lean_inc(x_211); -x_306 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_306, 0, x_211); +x_303 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +lean_inc(x_207); +lean_inc(x_229); +x_304 = l_Lean_addMacroScope(x_229, x_303, x_207); +x_305 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; +lean_inc(x_2); +lean_inc(x_223); +x_306 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_306, 0, x_223); lean_ctor_set(x_306, 1, x_305); -x_307 = l_Lean_nullKind; -x_308 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_308, 0, x_228); -lean_ctor_set(x_308, 1, x_307); -lean_ctor_set(x_308, 2, x_12); -x_309 = lean_array_push(x_270, x_306); -x_310 = lean_array_push(x_309, x_308); -lean_inc(x_269); -x_311 = lean_array_push(x_310, x_269); -x_312 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639____closed__1; -x_313 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_313, 0, x_228); -lean_ctor_set(x_313, 1, x_312); -lean_ctor_set(x_313, 2, x_311); -x_314 = lean_array_push(x_226, x_313); -x_315 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_315, 0, x_228); -lean_ctor_set(x_315, 1, x_233); -lean_ctor_set(x_315, 2, x_314); -x_316 = lean_array_push(x_226, x_315); -x_317 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_317, 0, x_228); -lean_ctor_set(x_317, 1, x_233); -lean_ctor_set(x_317, 2, x_316); -x_318 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__12; -lean_inc(x_211); -x_319 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_319, 0, x_211); -lean_ctor_set(x_319, 1, x_318); -x_320 = lean_usize_of_nat(x_208); -lean_dec(x_208); -x_321 = lean_unsigned_to_nat(0u); -x_322 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__6; -x_323 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__11(x_197, x_321, x_322, x_16, x_16, x_211, x_217, x_233, x_226, x_237, x_235, x_285, x_274, x_269, x_271, x_320, x_279, x_14); -x_324 = l_Array_append___rarg(x_276, x_323); -x_325 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_325, 0, x_228); -lean_ctor_set(x_325, 1, x_233); -lean_ctor_set(x_325, 2, x_324); -x_326 = lean_array_push(x_283, x_325); -x_327 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_327, 0, x_228); -lean_ctor_set(x_327, 1, x_285); -lean_ctor_set(x_327, 2, x_326); -x_328 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; -x_329 = lean_array_push(x_328, x_304); -lean_inc(x_329); -x_330 = lean_array_push(x_329, x_317); -lean_inc(x_319); -x_331 = lean_array_push(x_330, x_319); -x_332 = lean_array_push(x_331, x_327); -x_333 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__24; -x_334 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_334, 0, x_228); -lean_ctor_set(x_334, 1, x_333); -lean_ctor_set(x_334, 2, x_332); -x_335 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_335, 0, x_228); -lean_ctor_set(x_335, 1, x_233); -lean_ctor_set(x_335, 2, x_231); -x_336 = lean_array_push(x_226, x_335); -x_337 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_337, 0, x_228); -lean_ctor_set(x_337, 1, x_233); -lean_ctor_set(x_337, 2, x_336); -x_338 = lean_array_push(x_329, x_337); -x_339 = lean_array_push(x_338, x_319); -x_340 = lean_array_push(x_339, x_135); -x_341 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_341, 0, x_228); -lean_ctor_set(x_341, 1, x_333); -lean_ctor_set(x_341, 2, x_340); -x_342 = lean_array_push(x_235, x_334); -x_343 = lean_array_push(x_342, x_341); -x_344 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_344, 0, x_228); -lean_ctor_set(x_344, 1, x_233); -lean_ctor_set(x_344, 2, x_343); -x_345 = lean_array_push(x_226, x_344); -x_346 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__22; -x_347 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_347, 0, x_228); -lean_ctor_set(x_347, 1, x_346); -lean_ctor_set(x_347, 2, x_345); -x_348 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__30; -x_349 = lean_array_push(x_348, x_290); -x_350 = lean_array_push(x_349, x_237); -x_351 = lean_array_push(x_350, x_237); -x_352 = lean_array_push(x_351, x_300); -x_353 = lean_array_push(x_352, x_302); -x_354 = lean_array_push(x_353, x_347); -x_355 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__17; -x_356 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_356, 0, x_228); -lean_ctor_set(x_356, 1, x_355); -lean_ctor_set(x_356, 2, x_354); -x_357 = lean_array_push(x_348, x_259); -x_358 = lean_array_push(x_357, x_263); -x_359 = lean_array_push(x_358, x_265); -x_360 = lean_array_push(x_359, x_286); -x_361 = lean_array_push(x_360, x_288); -x_362 = lean_array_push(x_361, x_356); -x_363 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__24; -x_364 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_364, 0, x_228); -lean_ctor_set(x_364, 1, x_363); -lean_ctor_set(x_364, 2, x_362); -x_365 = lean_array_push(x_328, x_219); -x_366 = lean_array_push(x_365, x_255); -x_367 = lean_array_push(x_366, x_257); -x_368 = lean_array_push(x_367, x_364); -x_369 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__16; -x_370 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_370, 0, x_228); -lean_ctor_set(x_370, 1, x_369); -lean_ctor_set(x_370, 2, x_368); -lean_ctor_set(x_213, 0, x_370); -return x_213; +lean_ctor_set(x_306, 2, x_304); +lean_ctor_set(x_306, 3, x_2); +x_307 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__63; +x_308 = lean_array_push(x_307, x_306); +x_309 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__46; +x_310 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_310, 0, x_240); +lean_ctor_set(x_310, 1, x_309); +lean_ctor_set(x_310, 2, x_308); +x_311 = lean_array_push(x_238, x_310); +x_312 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_312, 0, x_240); +lean_ctor_set(x_312, 1, x_245); +lean_ctor_set(x_312, 2, x_311); +x_313 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__20; +lean_inc(x_223); +x_314 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_314, 0, x_223); +lean_ctor_set(x_314, 1, x_313); +x_315 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__25; +lean_inc(x_223); +x_316 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_316, 0, x_223); +lean_ctor_set(x_316, 1, x_315); +x_317 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__3___lambda__1___closed__1; +lean_inc(x_223); +x_318 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_318, 0, x_223); +lean_ctor_set(x_318, 1, x_317); +x_319 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_319, 0, x_240); +lean_ctor_set(x_319, 1, x_245); +lean_ctor_set(x_319, 2, x_14); +x_320 = lean_array_push(x_282, x_318); +x_321 = lean_array_push(x_320, x_319); +lean_inc(x_281); +x_322 = lean_array_push(x_321, x_281); +x_323 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880____closed__1; +x_324 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_324, 0, x_240); +lean_ctor_set(x_324, 1, x_323); +lean_ctor_set(x_324, 2, x_322); +x_325 = lean_array_push(x_238, x_324); +x_326 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_326, 0, x_240); +lean_ctor_set(x_326, 1, x_245); +lean_ctor_set(x_326, 2, x_325); +x_327 = lean_array_push(x_238, x_326); +x_328 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_328, 0, x_240); +lean_ctor_set(x_328, 1, x_245); +lean_ctor_set(x_328, 2, x_327); +x_329 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__12; +lean_inc(x_223); +x_330 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_330, 0, x_223); +lean_ctor_set(x_330, 1, x_329); +x_331 = lean_usize_of_nat(x_220); +lean_dec(x_220); +x_332 = lean_unsigned_to_nat(0u); +x_333 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__6; +lean_inc(x_2); +x_334 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__11(x_207, x_332, x_333, x_2, x_2, x_223, x_229, x_245, x_238, x_249, x_247, x_297, x_286, x_281, x_283, x_331, x_291, x_16); +x_335 = l_Array_append___rarg(x_288, x_334); +x_336 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_336, 0, x_240); +lean_ctor_set(x_336, 1, x_245); +lean_ctor_set(x_336, 2, x_335); +x_337 = lean_array_push(x_295, x_336); +x_338 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_338, 0, x_240); +lean_ctor_set(x_338, 1, x_297); +lean_ctor_set(x_338, 2, x_337); +x_339 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; +x_340 = lean_array_push(x_339, x_316); +lean_inc(x_340); +x_341 = lean_array_push(x_340, x_328); +lean_inc(x_330); +x_342 = lean_array_push(x_341, x_330); +x_343 = lean_array_push(x_342, x_338); +x_344 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__24; +x_345 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_345, 0, x_240); +lean_ctor_set(x_345, 1, x_344); +lean_ctor_set(x_345, 2, x_343); +x_346 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_346, 0, x_240); +lean_ctor_set(x_346, 1, x_245); +lean_ctor_set(x_346, 2, x_243); +x_347 = lean_array_push(x_238, x_346); +x_348 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_348, 0, x_240); +lean_ctor_set(x_348, 1, x_245); +lean_ctor_set(x_348, 2, x_347); +x_349 = lean_array_push(x_340, x_348); +x_350 = lean_array_push(x_349, x_330); +x_351 = lean_array_push(x_350, x_142); +x_352 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_352, 0, x_240); +lean_ctor_set(x_352, 1, x_344); +lean_ctor_set(x_352, 2, x_351); +x_353 = lean_array_push(x_247, x_345); +x_354 = lean_array_push(x_353, x_352); +x_355 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_355, 0, x_240); +lean_ctor_set(x_355, 1, x_245); +lean_ctor_set(x_355, 2, x_354); +x_356 = lean_array_push(x_238, x_355); +x_357 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__22; +x_358 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_358, 0, x_240); +lean_ctor_set(x_358, 1, x_357); +lean_ctor_set(x_358, 2, x_356); +x_359 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__30; +x_360 = lean_array_push(x_359, x_302); +x_361 = lean_array_push(x_360, x_249); +x_362 = lean_array_push(x_361, x_249); +x_363 = lean_array_push(x_362, x_312); +x_364 = lean_array_push(x_363, x_314); +x_365 = lean_array_push(x_364, x_358); +x_366 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__17; +x_367 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_367, 0, x_240); +lean_ctor_set(x_367, 1, x_366); +lean_ctor_set(x_367, 2, x_365); +x_368 = lean_array_push(x_359, x_271); +x_369 = lean_array_push(x_368, x_275); +x_370 = lean_array_push(x_369, x_277); +x_371 = lean_array_push(x_370, x_298); +x_372 = lean_array_push(x_371, x_300); +x_373 = lean_array_push(x_372, x_367); +x_374 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__20; +x_375 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_375, 0, x_240); +lean_ctor_set(x_375, 1, x_374); +lean_ctor_set(x_375, 2, x_373); +x_376 = lean_array_push(x_339, x_231); +x_377 = lean_array_push(x_376, x_267); +x_378 = lean_array_push(x_377, x_269); +x_379 = lean_array_push(x_378, x_375); +x_380 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__12; +x_381 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_381, 0, x_240); +lean_ctor_set(x_381, 1, x_380); +lean_ctor_set(x_381, 2, x_379); +lean_ctor_set(x_225, 0, x_381); +return x_225; } else { -lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; size_t x_435; size_t x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; size_t x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; -x_371 = lean_ctor_get(x_213, 0); -x_372 = lean_ctor_get(x_213, 1); -lean_inc(x_372); -lean_inc(x_371); -lean_dec(x_213); -x_373 = lean_ctor_get(x_371, 0); -lean_inc(x_373); -lean_dec(x_371); -x_374 = lean_environment_main_module(x_373); -x_375 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__15; -lean_inc(x_211); -x_376 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_376, 0, x_211); -lean_ctor_set(x_376, 1, x_375); -x_377 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__20; -lean_inc(x_197); -lean_inc(x_374); -x_378 = l_Lean_addMacroScope(x_374, x_377, x_197); -x_379 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__19; -lean_inc(x_211); -x_380 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_380, 0, x_211); -lean_ctor_set(x_380, 1, x_379); -lean_ctor_set(x_380, 2, x_378); -lean_ctor_set(x_380, 3, x_16); -x_381 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__2; -lean_inc(x_211); -x_382 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_382, 0, x_211); -lean_ctor_set(x_382, 1, x_381); -x_383 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; -x_384 = lean_array_push(x_383, x_382); -x_385 = lean_box(2); -x_386 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_noOpMatchAdaptPats___lambda__1___closed__1; -x_387 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_387, 0, x_385); +lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; size_t x_446; size_t x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; size_t x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; +x_382 = lean_ctor_get(x_225, 0); +x_383 = lean_ctor_get(x_225, 1); +lean_inc(x_383); +lean_inc(x_382); +lean_dec(x_225); +x_384 = lean_ctor_get(x_382, 0); +lean_inc(x_384); +lean_dec(x_382); +x_385 = lean_environment_main_module(x_384); +x_386 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__11; +lean_inc(x_223); +x_387 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_387, 0, x_223); lean_ctor_set(x_387, 1, x_386); -lean_ctor_set(x_387, 2, x_384); -x_388 = lean_array_push(x_383, x_387); -lean_inc(x_14); -lean_inc(x_388); -x_389 = l_Array_append___rarg(x_388, x_14); -x_390 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_391 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_391, 0, x_385); +x_388 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__16; +lean_inc(x_207); +lean_inc(x_385); +x_389 = l_Lean_addMacroScope(x_385, x_388, x_207); +x_390 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__15; +lean_inc(x_2); +lean_inc(x_223); +x_391 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_391, 0, x_223); lean_ctor_set(x_391, 1, x_390); lean_ctor_set(x_391, 2, x_389); -x_392 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_393 = lean_array_push(x_392, x_391); -x_394 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; -x_395 = lean_array_push(x_393, x_394); -x_396 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__22; -x_397 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_397, 0, x_385); -lean_ctor_set(x_397, 1, x_396); -lean_ctor_set(x_397, 2, x_395); -x_398 = lean_array_push(x_383, x_397); -x_399 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_399, 0, x_385); -lean_ctor_set(x_399, 1, x_390); -lean_ctor_set(x_399, 2, x_398); -x_400 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__17; -lean_inc(x_211); -x_401 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_401, 0, x_211); -lean_ctor_set(x_401, 1, x_400); -x_402 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; -lean_inc(x_380); -x_403 = lean_array_push(x_402, x_380); -x_404 = lean_array_push(x_403, x_399); -x_405 = lean_array_push(x_404, x_394); -x_406 = lean_array_push(x_405, x_401); -x_407 = lean_array_push(x_406, x_18); -x_408 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__12; -x_409 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_409, 0, x_385); -lean_ctor_set(x_409, 1, x_408); -lean_ctor_set(x_409, 2, x_407); -x_410 = lean_array_push(x_383, x_409); -x_411 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__10; -x_412 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_412, 0, x_385); +lean_ctor_set(x_391, 3, x_2); +x_392 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__14___closed__2; +lean_inc(x_223); +x_393 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_393, 0, x_223); +lean_ctor_set(x_393, 1, x_392); +x_394 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; +x_395 = lean_array_push(x_394, x_393); +x_396 = lean_box(2); +x_397 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_noOpMatchAdaptPats___lambda__1___closed__1; +x_398 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_398, 0, x_396); +lean_ctor_set(x_398, 1, x_397); +lean_ctor_set(x_398, 2, x_395); +x_399 = lean_array_push(x_394, x_398); +lean_inc(x_16); +lean_inc(x_399); +x_400 = l_Array_append___rarg(x_399, x_16); +x_401 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_402 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_402, 0, x_396); +lean_ctor_set(x_402, 1, x_401); +lean_ctor_set(x_402, 2, x_400); +x_403 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_404 = lean_array_push(x_403, x_402); +x_405 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; +x_406 = lean_array_push(x_404, x_405); +x_407 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__18; +x_408 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_408, 0, x_396); +lean_ctor_set(x_408, 1, x_407); +lean_ctor_set(x_408, 2, x_406); +x_409 = lean_array_push(x_394, x_408); +x_410 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_410, 0, x_396); +lean_ctor_set(x_410, 1, x_401); +lean_ctor_set(x_410, 2, x_409); +x_411 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__17; +lean_inc(x_223); +x_412 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_412, 0, x_223); lean_ctor_set(x_412, 1, x_411); -lean_ctor_set(x_412, 2, x_410); -x_413 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__20; -lean_inc(x_211); -x_414 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_414, 0, x_211); -lean_ctor_set(x_414, 1, x_413); -x_415 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__25; -lean_inc(x_211); -x_416 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_416, 0, x_211); -lean_ctor_set(x_416, 1, x_415); -x_417 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__30; -lean_inc(x_197); -lean_inc(x_374); -x_418 = l_Lean_addMacroScope(x_374, x_417, x_197); -x_419 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__28; -lean_inc(x_211); -x_420 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_420, 0, x_211); +x_413 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; +lean_inc(x_391); +x_414 = lean_array_push(x_413, x_391); +x_415 = lean_array_push(x_414, x_410); +x_416 = lean_array_push(x_415, x_405); +x_417 = lean_array_push(x_416, x_412); +x_418 = lean_array_push(x_417, x_19); +x_419 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__12; +x_420 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_420, 0, x_396); lean_ctor_set(x_420, 1, x_419); lean_ctor_set(x_420, 2, x_418); -lean_ctor_set(x_420, 3, x_16); -x_421 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__31; -lean_inc(x_211); -x_422 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_422, 0, x_211); -lean_ctor_set(x_422, 1, x_421); -x_423 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; -lean_inc(x_211); -x_424 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_424, 0, x_211); -lean_ctor_set(x_424, 1, x_423); -x_425 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; -lean_inc(x_211); -x_426 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_426, 0, x_211); -lean_ctor_set(x_426, 1, x_425); -x_427 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; -x_428 = lean_array_push(x_427, x_424); -lean_inc(x_428); -x_429 = lean_array_push(x_428, x_394); -lean_inc(x_426); -x_430 = lean_array_push(x_429, x_426); -x_431 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__6; -x_432 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_432, 0, x_385); -lean_ctor_set(x_432, 1, x_431); -lean_ctor_set(x_432, 2, x_430); -x_433 = lean_array_push(x_383, x_432); -x_434 = lean_array_get_size(x_209); -x_435 = lean_usize_of_nat(x_434); -lean_dec(x_434); -x_436 = 0; -x_437 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_435, x_436, x_209); -lean_inc(x_433); -x_438 = l_Array_append___rarg(x_433, x_437); -x_439 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_439, 0, x_385); -lean_ctor_set(x_439, 1, x_390); -lean_ctor_set(x_439, 2, x_438); -x_440 = lean_array_push(x_392, x_380); -lean_inc(x_440); -x_441 = lean_array_push(x_440, x_439); -x_442 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_421 = lean_array_push(x_394, x_420); +x_422 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__10; +x_423 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_423, 0, x_396); +lean_ctor_set(x_423, 1, x_422); +lean_ctor_set(x_423, 2, x_421); +x_424 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__20; +lean_inc(x_223); +x_425 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_425, 0, x_223); +lean_ctor_set(x_425, 1, x_424); +x_426 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__21; +lean_inc(x_223); +x_427 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_427, 0, x_223); +lean_ctor_set(x_427, 1, x_426); +x_428 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__26; +lean_inc(x_207); +lean_inc(x_385); +x_429 = l_Lean_addMacroScope(x_385, x_428, x_207); +x_430 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__24; +lean_inc(x_2); +lean_inc(x_223); +x_431 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_431, 0, x_223); +lean_ctor_set(x_431, 1, x_430); +lean_ctor_set(x_431, 2, x_429); +lean_ctor_set(x_431, 3, x_2); +x_432 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__27; +lean_inc(x_223); +x_433 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_433, 0, x_223); +lean_ctor_set(x_433, 1, x_432); +x_434 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; +lean_inc(x_223); +x_435 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_435, 0, x_223); +lean_ctor_set(x_435, 1, x_434); +x_436 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; +lean_inc(x_223); +x_437 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_437, 0, x_223); +lean_ctor_set(x_437, 1, x_436); +x_438 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; +x_439 = lean_array_push(x_438, x_435); +lean_inc(x_439); +x_440 = lean_array_push(x_439, x_405); +lean_inc(x_437); +x_441 = lean_array_push(x_440, x_437); +x_442 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6; x_443 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_443, 0, x_385); +lean_ctor_set(x_443, 0, x_396); lean_ctor_set(x_443, 1, x_442); lean_ctor_set(x_443, 2, x_441); -x_444 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__32; -lean_inc(x_211); -x_445 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_445, 0, x_211); -lean_ctor_set(x_445, 1, x_444); -x_446 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__16; -lean_inc(x_211); -x_447 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_447, 0, x_211); -lean_ctor_set(x_447, 1, x_446); -x_448 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -lean_inc(x_197); -lean_inc(x_374); -x_449 = l_Lean_addMacroScope(x_374, x_448, x_197); -x_450 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; -lean_inc(x_211); -x_451 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_451, 0, x_211); -lean_ctor_set(x_451, 1, x_450); -lean_ctor_set(x_451, 2, x_449); -lean_ctor_set(x_451, 3, x_16); -x_452 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__63; -x_453 = lean_array_push(x_452, x_451); -x_454 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__46; -x_455 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_455, 0, x_385); -lean_ctor_set(x_455, 1, x_454); -lean_ctor_set(x_455, 2, x_453); -x_456 = lean_array_push(x_383, x_455); -x_457 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_457, 0, x_385); -lean_ctor_set(x_457, 1, x_390); -lean_ctor_set(x_457, 2, x_456); -x_458 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__20; -lean_inc(x_211); -x_459 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_459, 0, x_211); -lean_ctor_set(x_459, 1, x_458); -x_460 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__25; -lean_inc(x_211); -x_461 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_461, 0, x_211); -lean_ctor_set(x_461, 1, x_460); -x_462 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__3___lambda__1___closed__1; -lean_inc(x_211); -x_463 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_463, 0, x_211); -lean_ctor_set(x_463, 1, x_462); -x_464 = l_Lean_nullKind; -x_465 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_465, 0, x_385); -lean_ctor_set(x_465, 1, x_464); -lean_ctor_set(x_465, 2, x_12); -x_466 = lean_array_push(x_427, x_463); -x_467 = lean_array_push(x_466, x_465); -lean_inc(x_426); -x_468 = lean_array_push(x_467, x_426); -x_469 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639____closed__1; -x_470 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_470, 0, x_385); +x_444 = lean_array_push(x_394, x_443); +x_445 = lean_array_get_size(x_221); +x_446 = lean_usize_of_nat(x_445); +lean_dec(x_445); +x_447 = 0; +x_448 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_446, x_447, x_221); +lean_inc(x_444); +x_449 = l_Array_append___rarg(x_444, x_448); +x_450 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_450, 0, x_396); +lean_ctor_set(x_450, 1, x_401); +lean_ctor_set(x_450, 2, x_449); +x_451 = lean_array_push(x_403, x_391); +lean_inc(x_451); +x_452 = lean_array_push(x_451, x_450); +x_453 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_454 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_454, 0, x_396); +lean_ctor_set(x_454, 1, x_453); +lean_ctor_set(x_454, 2, x_452); +x_455 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__28; +lean_inc(x_223); +x_456 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_456, 0, x_223); +lean_ctor_set(x_456, 1, x_455); +x_457 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__16; +lean_inc(x_223); +x_458 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_458, 0, x_223); +lean_ctor_set(x_458, 1, x_457); +x_459 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +lean_inc(x_207); +lean_inc(x_385); +x_460 = l_Lean_addMacroScope(x_385, x_459, x_207); +x_461 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; +lean_inc(x_2); +lean_inc(x_223); +x_462 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_462, 0, x_223); +lean_ctor_set(x_462, 1, x_461); +lean_ctor_set(x_462, 2, x_460); +lean_ctor_set(x_462, 3, x_2); +x_463 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__63; +x_464 = lean_array_push(x_463, x_462); +x_465 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__46; +x_466 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_466, 0, x_396); +lean_ctor_set(x_466, 1, x_465); +lean_ctor_set(x_466, 2, x_464); +x_467 = lean_array_push(x_394, x_466); +x_468 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_468, 0, x_396); +lean_ctor_set(x_468, 1, x_401); +lean_ctor_set(x_468, 2, x_467); +x_469 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__20; +lean_inc(x_223); +x_470 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_470, 0, x_223); lean_ctor_set(x_470, 1, x_469); -lean_ctor_set(x_470, 2, x_468); -x_471 = lean_array_push(x_383, x_470); -x_472 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_472, 0, x_385); -lean_ctor_set(x_472, 1, x_390); -lean_ctor_set(x_472, 2, x_471); -x_473 = lean_array_push(x_383, x_472); -x_474 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_474, 0, x_385); -lean_ctor_set(x_474, 1, x_390); -lean_ctor_set(x_474, 2, x_473); -x_475 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__12; -lean_inc(x_211); -x_476 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_476, 0, x_211); -lean_ctor_set(x_476, 1, x_475); -x_477 = lean_usize_of_nat(x_208); -lean_dec(x_208); -x_478 = lean_unsigned_to_nat(0u); -x_479 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__6; -x_480 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__11(x_197, x_478, x_479, x_16, x_16, x_211, x_374, x_390, x_383, x_394, x_392, x_442, x_431, x_426, x_428, x_477, x_436, x_14); -x_481 = l_Array_append___rarg(x_433, x_480); +x_471 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__25; +lean_inc(x_223); +x_472 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_472, 0, x_223); +lean_ctor_set(x_472, 1, x_471); +x_473 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__3___lambda__1___closed__1; +lean_inc(x_223); +x_474 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_474, 0, x_223); +lean_ctor_set(x_474, 1, x_473); +x_475 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_475, 0, x_396); +lean_ctor_set(x_475, 1, x_401); +lean_ctor_set(x_475, 2, x_14); +x_476 = lean_array_push(x_438, x_474); +x_477 = lean_array_push(x_476, x_475); +lean_inc(x_437); +x_478 = lean_array_push(x_477, x_437); +x_479 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880____closed__1; +x_480 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_480, 0, x_396); +lean_ctor_set(x_480, 1, x_479); +lean_ctor_set(x_480, 2, x_478); +x_481 = lean_array_push(x_394, x_480); x_482 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_482, 0, x_385); -lean_ctor_set(x_482, 1, x_390); +lean_ctor_set(x_482, 0, x_396); +lean_ctor_set(x_482, 1, x_401); lean_ctor_set(x_482, 2, x_481); -x_483 = lean_array_push(x_440, x_482); +x_483 = lean_array_push(x_394, x_482); x_484 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_484, 0, x_385); -lean_ctor_set(x_484, 1, x_442); +lean_ctor_set(x_484, 0, x_396); +lean_ctor_set(x_484, 1, x_401); lean_ctor_set(x_484, 2, x_483); -x_485 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; -x_486 = lean_array_push(x_485, x_461); -lean_inc(x_486); -x_487 = lean_array_push(x_486, x_474); -lean_inc(x_476); -x_488 = lean_array_push(x_487, x_476); -x_489 = lean_array_push(x_488, x_484); -x_490 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__24; -x_491 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_491, 0, x_385); -lean_ctor_set(x_491, 1, x_490); -lean_ctor_set(x_491, 2, x_489); +x_485 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__12; +lean_inc(x_223); +x_486 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_486, 0, x_223); +lean_ctor_set(x_486, 1, x_485); +x_487 = lean_usize_of_nat(x_220); +lean_dec(x_220); +x_488 = lean_unsigned_to_nat(0u); +x_489 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__6; +lean_inc(x_2); +x_490 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__11(x_207, x_488, x_489, x_2, x_2, x_223, x_385, x_401, x_394, x_405, x_403, x_453, x_442, x_437, x_439, x_487, x_447, x_16); +x_491 = l_Array_append___rarg(x_444, x_490); x_492 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_492, 0, x_385); -lean_ctor_set(x_492, 1, x_390); -lean_ctor_set(x_492, 2, x_388); -x_493 = lean_array_push(x_383, x_492); +lean_ctor_set(x_492, 0, x_396); +lean_ctor_set(x_492, 1, x_401); +lean_ctor_set(x_492, 2, x_491); +x_493 = lean_array_push(x_451, x_492); x_494 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_494, 0, x_385); -lean_ctor_set(x_494, 1, x_390); +lean_ctor_set(x_494, 0, x_396); +lean_ctor_set(x_494, 1, x_453); lean_ctor_set(x_494, 2, x_493); -x_495 = lean_array_push(x_486, x_494); -x_496 = lean_array_push(x_495, x_476); -x_497 = lean_array_push(x_496, x_135); -x_498 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_498, 0, x_385); -lean_ctor_set(x_498, 1, x_490); -lean_ctor_set(x_498, 2, x_497); -x_499 = lean_array_push(x_392, x_491); -x_500 = lean_array_push(x_499, x_498); +x_495 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; +x_496 = lean_array_push(x_495, x_472); +lean_inc(x_496); +x_497 = lean_array_push(x_496, x_484); +lean_inc(x_486); +x_498 = lean_array_push(x_497, x_486); +x_499 = lean_array_push(x_498, x_494); +x_500 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__24; x_501 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_501, 0, x_385); -lean_ctor_set(x_501, 1, x_390); -lean_ctor_set(x_501, 2, x_500); -x_502 = lean_array_push(x_383, x_501); -x_503 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__22; +lean_ctor_set(x_501, 0, x_396); +lean_ctor_set(x_501, 1, x_500); +lean_ctor_set(x_501, 2, x_499); +x_502 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_502, 0, x_396); +lean_ctor_set(x_502, 1, x_401); +lean_ctor_set(x_502, 2, x_399); +x_503 = lean_array_push(x_394, x_502); x_504 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_504, 0, x_385); -lean_ctor_set(x_504, 1, x_503); -lean_ctor_set(x_504, 2, x_502); -x_505 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__30; -x_506 = lean_array_push(x_505, x_447); -x_507 = lean_array_push(x_506, x_394); -x_508 = lean_array_push(x_507, x_394); -x_509 = lean_array_push(x_508, x_457); -x_510 = lean_array_push(x_509, x_459); -x_511 = lean_array_push(x_510, x_504); -x_512 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__17; -x_513 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_513, 0, x_385); -lean_ctor_set(x_513, 1, x_512); -lean_ctor_set(x_513, 2, x_511); -x_514 = lean_array_push(x_505, x_416); -x_515 = lean_array_push(x_514, x_420); -x_516 = lean_array_push(x_515, x_422); -x_517 = lean_array_push(x_516, x_443); -x_518 = lean_array_push(x_517, x_445); -x_519 = lean_array_push(x_518, x_513); -x_520 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__24; -x_521 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_521, 0, x_385); -lean_ctor_set(x_521, 1, x_520); -lean_ctor_set(x_521, 2, x_519); -x_522 = lean_array_push(x_485, x_376); -x_523 = lean_array_push(x_522, x_412); -x_524 = lean_array_push(x_523, x_414); -x_525 = lean_array_push(x_524, x_521); -x_526 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__16; -x_527 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_527, 0, x_385); -lean_ctor_set(x_527, 1, x_526); -lean_ctor_set(x_527, 2, x_525); -x_528 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_528, 0, x_527); -lean_ctor_set(x_528, 1, x_372); -return x_528; +lean_ctor_set(x_504, 0, x_396); +lean_ctor_set(x_504, 1, x_401); +lean_ctor_set(x_504, 2, x_503); +x_505 = lean_array_push(x_496, x_504); +x_506 = lean_array_push(x_505, x_486); +x_507 = lean_array_push(x_506, x_142); +x_508 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_508, 0, x_396); +lean_ctor_set(x_508, 1, x_500); +lean_ctor_set(x_508, 2, x_507); +x_509 = lean_array_push(x_403, x_501); +x_510 = lean_array_push(x_509, x_508); +x_511 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_511, 0, x_396); +lean_ctor_set(x_511, 1, x_401); +lean_ctor_set(x_511, 2, x_510); +x_512 = lean_array_push(x_394, x_511); +x_513 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__22; +x_514 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_514, 0, x_396); +lean_ctor_set(x_514, 1, x_513); +lean_ctor_set(x_514, 2, x_512); +x_515 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__30; +x_516 = lean_array_push(x_515, x_458); +x_517 = lean_array_push(x_516, x_405); +x_518 = lean_array_push(x_517, x_405); +x_519 = lean_array_push(x_518, x_468); +x_520 = lean_array_push(x_519, x_470); +x_521 = lean_array_push(x_520, x_514); +x_522 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__17; +x_523 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_523, 0, x_396); +lean_ctor_set(x_523, 1, x_522); +lean_ctor_set(x_523, 2, x_521); +x_524 = lean_array_push(x_515, x_427); +x_525 = lean_array_push(x_524, x_431); +x_526 = lean_array_push(x_525, x_433); +x_527 = lean_array_push(x_526, x_454); +x_528 = lean_array_push(x_527, x_456); +x_529 = lean_array_push(x_528, x_523); +x_530 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__20; +x_531 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_531, 0, x_396); +lean_ctor_set(x_531, 1, x_530); +lean_ctor_set(x_531, 2, x_529); +x_532 = lean_array_push(x_495, x_387); +x_533 = lean_array_push(x_532, x_423); +x_534 = lean_array_push(x_533, x_425); +x_535 = lean_array_push(x_534, x_531); +x_536 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__12; +x_537 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_537, 0, x_396); +lean_ctor_set(x_537, 1, x_536); +lean_ctor_set(x_537, 2, x_535); +x_538 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_538, 0, x_537); +lean_ctor_set(x_538, 1, x_383); +return x_538; } } } case 1: { -lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; uint8_t x_560; -lean_dec(x_134); -lean_dec(x_77); -x_529 = lean_ctor_get(x_20, 0); -lean_inc(x_529); -x_530 = lean_ctor_get(x_20, 1); -lean_inc(x_530); -lean_dec(x_20); -lean_inc(x_8); -x_531 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_8, x_9, x_530); -x_532 = lean_ctor_get(x_531, 0); -lean_inc(x_532); -x_533 = lean_ctor_get(x_531, 1); -lean_inc(x_533); -lean_dec(x_531); -x_534 = lean_ctor_get(x_8, 10); -lean_inc(x_534); -x_535 = lean_st_ref_get(x_9, x_533); -x_536 = lean_ctor_get(x_535, 0); -lean_inc(x_536); -x_537 = lean_ctor_get(x_535, 1); -lean_inc(x_537); -lean_dec(x_535); -x_538 = lean_ctor_get(x_536, 0); -lean_inc(x_538); -lean_dec(x_536); -x_539 = lean_environment_main_module(x_538); -x_540 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__4; -lean_inc(x_534); +lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; lean_object* x_565; lean_object* x_566; lean_object* x_567; lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; uint8_t x_572; +lean_dec(x_141); +lean_dec(x_81); +x_539 = lean_ctor_get(x_21, 0); lean_inc(x_539); -x_541 = l_Lean_addMacroScope(x_539, x_540, x_534); -x_542 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__3; -x_543 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__7; -lean_inc(x_532); -x_544 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_544, 0, x_532); -lean_ctor_set(x_544, 1, x_542); -lean_ctor_set(x_544, 2, x_541); -lean_ctor_set(x_544, 3, x_543); -x_545 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -lean_inc(x_534); -x_546 = l_Lean_addMacroScope(x_539, x_545, x_534); -x_547 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; -x_548 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_548, 0, x_532); -lean_ctor_set(x_548, 1, x_547); -lean_ctor_set(x_548, 2, x_546); -lean_ctor_set(x_548, 3, x_16); -x_549 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; -x_550 = lean_array_push(x_549, x_548); -x_551 = lean_box(2); -x_552 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_553 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_553, 0, x_551); -lean_ctor_set(x_553, 1, x_552); -lean_ctor_set(x_553, 2, x_550); -x_554 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_555 = lean_array_push(x_554, x_544); -x_556 = lean_array_push(x_555, x_553); -x_557 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_558 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_558, 0, x_551); -lean_ctor_set(x_558, 1, x_557); -lean_ctor_set(x_558, 2, x_556); -x_559 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__3; -x_560 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__2(x_11, x_559); -lean_dec(x_11); -if (x_560 == 0) +x_540 = lean_ctor_get(x_21, 1); +lean_inc(x_540); +lean_dec(x_21); +lean_inc(x_10); +x_541 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_10, x_11, x_540); +x_542 = lean_ctor_get(x_541, 0); +lean_inc(x_542); +x_543 = lean_ctor_get(x_541, 1); +lean_inc(x_543); +lean_dec(x_541); +x_544 = lean_ctor_get(x_10, 10); +lean_inc(x_544); +x_545 = lean_st_ref_get(x_11, x_543); +x_546 = lean_ctor_get(x_545, 0); +lean_inc(x_546); +x_547 = lean_ctor_get(x_545, 1); +lean_inc(x_547); +lean_dec(x_545); +x_548 = lean_ctor_get(x_546, 0); +lean_inc(x_548); +lean_dec(x_546); +x_549 = lean_environment_main_module(x_548); +x_550 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__4; +lean_inc(x_544); +lean_inc(x_549); +x_551 = l_Lean_addMacroScope(x_549, x_550, x_544); +x_552 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__5; +lean_inc(x_2); +x_553 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_553, 0, x_552); +lean_ctor_set(x_553, 1, x_2); +lean_inc(x_2); +x_554 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_554, 0, x_553); +lean_ctor_set(x_554, 1, x_2); +x_555 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__3; +lean_inc(x_542); +x_556 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_556, 0, x_542); +lean_ctor_set(x_556, 1, x_555); +lean_ctor_set(x_556, 2, x_551); +lean_ctor_set(x_556, 3, x_554); +x_557 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +lean_inc(x_544); +x_558 = l_Lean_addMacroScope(x_549, x_557, x_544); +x_559 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; +lean_inc(x_2); +x_560 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_560, 0, x_542); +lean_ctor_set(x_560, 1, x_559); +lean_ctor_set(x_560, 2, x_558); +lean_ctor_set(x_560, 3, x_2); +x_561 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; +x_562 = lean_array_push(x_561, x_560); +x_563 = lean_box(2); +x_564 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_565 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_565, 0, x_563); +lean_ctor_set(x_565, 1, x_564); +lean_ctor_set(x_565, 2, x_562); +x_566 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_567 = lean_array_push(x_566, x_556); +x_568 = lean_array_push(x_567, x_565); +x_569 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_570 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_570, 0, x_563); +lean_ctor_set(x_570, 1, x_569); +lean_ctor_set(x_570, 2, x_568); +x_571 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__3; +x_572 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5(x_13, x_571); +lean_dec(x_13); +if (x_572 == 0) { -lean_object* x_561; lean_object* x_562; lean_object* x_563; -lean_dec(x_534); -x_561 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_562 = lean_box(0); -x_563 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__20(x_14, x_561, x_552, x_16, x_549, x_554, x_557, x_529, x_12, x_18, x_558, x_562, x_4, x_5, x_6, x_7, x_8, x_9, x_537); -lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_563; +lean_object* x_573; lean_object* x_574; lean_object* x_575; +lean_dec(x_544); +x_573 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; +x_574 = lean_box(0); +x_575 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__20(x_16, x_573, x_564, x_2, x_561, x_566, x_569, x_539, x_14, x_3, x_19, x_570, x_574, x_6, x_7, x_8, x_9, x_10, x_11, x_547); +return x_575; } else { -lean_object* x_564; lean_object* x_565; lean_object* x_566; lean_object* x_567; lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; -lean_inc(x_8); -x_564 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_8, x_9, x_537); -x_565 = lean_ctor_get(x_564, 0); -lean_inc(x_565); -x_566 = lean_ctor_get(x_564, 1); -lean_inc(x_566); -lean_dec(x_564); -x_567 = lean_st_ref_get(x_9, x_566); -x_568 = lean_ctor_get(x_567, 0); -lean_inc(x_568); -x_569 = lean_ctor_get(x_567, 1); -lean_inc(x_569); -lean_dec(x_567); -x_570 = lean_ctor_get(x_568, 0); -lean_inc(x_570); -lean_dec(x_568); -x_571 = lean_environment_main_module(x_570); -x_572 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__12; -x_573 = l_Lean_addMacroScope(x_571, x_572, x_534); -x_574 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__10; -x_575 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__14; -x_576 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_576, 0, x_565); -lean_ctor_set(x_576, 1, x_574); -lean_ctor_set(x_576, 2, x_573); -lean_ctor_set(x_576, 3, x_575); -x_577 = lean_array_push(x_549, x_558); -x_578 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_578, 0, x_551); -lean_ctor_set(x_578, 1, x_552); -lean_ctor_set(x_578, 2, x_577); -x_579 = lean_array_push(x_554, x_576); -x_580 = lean_array_push(x_579, x_578); -x_581 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_581, 0, x_551); -lean_ctor_set(x_581, 1, x_557); -lean_ctor_set(x_581, 2, x_580); -x_582 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_583 = lean_box(0); -x_584 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__20(x_14, x_582, x_552, x_16, x_549, x_554, x_557, x_529, x_12, x_18, x_581, x_583, x_4, x_5, x_6, x_7, x_8, x_9, x_569); -lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_584; +lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; +lean_inc(x_10); +x_576 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_10, x_11, x_547); +x_577 = lean_ctor_get(x_576, 0); +lean_inc(x_577); +x_578 = lean_ctor_get(x_576, 1); +lean_inc(x_578); +lean_dec(x_576); +x_579 = lean_st_ref_get(x_11, x_578); +x_580 = lean_ctor_get(x_579, 0); +lean_inc(x_580); +x_581 = lean_ctor_get(x_579, 1); +lean_inc(x_581); +lean_dec(x_579); +x_582 = lean_ctor_get(x_580, 0); +lean_inc(x_582); +lean_dec(x_580); +x_583 = lean_environment_main_module(x_582); +x_584 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__10; +x_585 = l_Lean_addMacroScope(x_583, x_584, x_544); +lean_inc(x_2); +x_586 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_586, 0, x_584); +lean_ctor_set(x_586, 1, x_2); +lean_inc(x_2); +x_587 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_587, 0, x_586); +lean_ctor_set(x_587, 1, x_2); +x_588 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__8; +x_589 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_589, 0, x_577); +lean_ctor_set(x_589, 1, x_588); +lean_ctor_set(x_589, 2, x_585); +lean_ctor_set(x_589, 3, x_587); +x_590 = lean_array_push(x_561, x_570); +x_591 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_591, 0, x_563); +lean_ctor_set(x_591, 1, x_564); +lean_ctor_set(x_591, 2, x_590); +x_592 = lean_array_push(x_566, x_589); +x_593 = lean_array_push(x_592, x_591); +x_594 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_594, 0, x_563); +lean_ctor_set(x_594, 1, x_569); +lean_ctor_set(x_594, 2, x_593); +x_595 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; +x_596 = lean_box(0); +x_597 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__20(x_16, x_595, x_564, x_2, x_561, x_566, x_569, x_539, x_14, x_3, x_19, x_594, x_596, x_6, x_7, x_8, x_9, x_10, x_11, x_581); +return x_597; } } default: { -lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; uint8_t x_616; -lean_dec(x_134); -lean_dec(x_77); -x_585 = lean_ctor_get(x_20, 0); -lean_inc(x_585); -x_586 = lean_ctor_get(x_20, 1); -lean_inc(x_586); -lean_dec(x_20); -lean_inc(x_8); -x_587 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_8, x_9, x_586); -x_588 = lean_ctor_get(x_587, 0); -lean_inc(x_588); -x_589 = lean_ctor_get(x_587, 1); -lean_inc(x_589); -lean_dec(x_587); -x_590 = lean_ctor_get(x_8, 10); -lean_inc(x_590); -x_591 = lean_st_ref_get(x_9, x_589); -x_592 = lean_ctor_get(x_591, 0); -lean_inc(x_592); -x_593 = lean_ctor_get(x_591, 1); -lean_inc(x_593); -lean_dec(x_591); -x_594 = lean_ctor_get(x_592, 0); -lean_inc(x_594); -lean_dec(x_592); -x_595 = lean_environment_main_module(x_594); -x_596 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__4; -lean_inc(x_590); -lean_inc(x_595); -x_597 = l_Lean_addMacroScope(x_595, x_596, x_590); -x_598 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__3; -x_599 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__7; -lean_inc(x_588); -x_600 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_600, 0, x_588); -lean_ctor_set(x_600, 1, x_598); -lean_ctor_set(x_600, 2, x_597); -lean_ctor_set(x_600, 3, x_599); -x_601 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -lean_inc(x_590); -x_602 = l_Lean_addMacroScope(x_595, x_601, x_590); -x_603 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; -x_604 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_604, 0, x_588); -lean_ctor_set(x_604, 1, x_603); -lean_ctor_set(x_604, 2, x_602); -lean_ctor_set(x_604, 3, x_16); -x_605 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; -x_606 = lean_array_push(x_605, x_604); -x_607 = lean_box(2); -x_608 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_609 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_609, 0, x_607); -lean_ctor_set(x_609, 1, x_608); -lean_ctor_set(x_609, 2, x_606); -x_610 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_611 = lean_array_push(x_610, x_600); -x_612 = lean_array_push(x_611, x_609); -x_613 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_614 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_614, 0, x_607); -lean_ctor_set(x_614, 1, x_613); -lean_ctor_set(x_614, 2, x_612); -x_615 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__3; -x_616 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__2(x_11, x_615); -lean_dec(x_11); -if (x_616 == 0) +lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; uint8_t x_631; +lean_dec(x_141); +lean_dec(x_81); +x_598 = lean_ctor_get(x_21, 0); +lean_inc(x_598); +x_599 = lean_ctor_get(x_21, 1); +lean_inc(x_599); +lean_dec(x_21); +lean_inc(x_10); +x_600 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_10, x_11, x_599); +x_601 = lean_ctor_get(x_600, 0); +lean_inc(x_601); +x_602 = lean_ctor_get(x_600, 1); +lean_inc(x_602); +lean_dec(x_600); +x_603 = lean_ctor_get(x_10, 10); +lean_inc(x_603); +x_604 = lean_st_ref_get(x_11, x_602); +x_605 = lean_ctor_get(x_604, 0); +lean_inc(x_605); +x_606 = lean_ctor_get(x_604, 1); +lean_inc(x_606); +lean_dec(x_604); +x_607 = lean_ctor_get(x_605, 0); +lean_inc(x_607); +lean_dec(x_605); +x_608 = lean_environment_main_module(x_607); +x_609 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__4; +lean_inc(x_603); +lean_inc(x_608); +x_610 = l_Lean_addMacroScope(x_608, x_609, x_603); +x_611 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__5; +lean_inc(x_2); +x_612 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_612, 0, x_611); +lean_ctor_set(x_612, 1, x_2); +lean_inc(x_2); +x_613 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_613, 0, x_612); +lean_ctor_set(x_613, 1, x_2); +x_614 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__3; +lean_inc(x_601); +x_615 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_615, 0, x_601); +lean_ctor_set(x_615, 1, x_614); +lean_ctor_set(x_615, 2, x_610); +lean_ctor_set(x_615, 3, x_613); +x_616 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +lean_inc(x_603); +x_617 = l_Lean_addMacroScope(x_608, x_616, x_603); +x_618 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; +lean_inc(x_2); +x_619 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_619, 0, x_601); +lean_ctor_set(x_619, 1, x_618); +lean_ctor_set(x_619, 2, x_617); +lean_ctor_set(x_619, 3, x_2); +x_620 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; +x_621 = lean_array_push(x_620, x_619); +x_622 = lean_box(2); +x_623 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_624 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_624, 0, x_622); +lean_ctor_set(x_624, 1, x_623); +lean_ctor_set(x_624, 2, x_621); +x_625 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_626 = lean_array_push(x_625, x_615); +x_627 = lean_array_push(x_626, x_624); +x_628 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_629 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_629, 0, x_622); +lean_ctor_set(x_629, 1, x_628); +lean_ctor_set(x_629, 2, x_627); +x_630 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__3; +x_631 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5(x_13, x_630); +lean_dec(x_13); +if (x_631 == 0) { -lean_object* x_617; lean_object* x_618; lean_object* x_619; -lean_dec(x_590); -x_617 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_618 = lean_box(0); -x_619 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__21(x_14, x_617, x_608, x_16, x_605, x_610, x_613, x_585, x_12, x_18, x_614, x_618, x_4, x_5, x_6, x_7, x_8, x_9, x_593); -lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_619; +lean_object* x_632; lean_object* x_633; lean_object* x_634; +lean_dec(x_603); +x_632 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; +x_633 = lean_box(0); +x_634 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__21(x_16, x_632, x_623, x_2, x_620, x_625, x_628, x_598, x_14, x_3, x_19, x_629, x_633, x_6, x_7, x_8, x_9, x_10, x_11, x_606); +return x_634; } else { -lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; -lean_inc(x_8); -x_620 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_8, x_9, x_593); -x_621 = lean_ctor_get(x_620, 0); -lean_inc(x_621); -x_622 = lean_ctor_get(x_620, 1); -lean_inc(x_622); -lean_dec(x_620); -x_623 = lean_st_ref_get(x_9, x_622); -x_624 = lean_ctor_get(x_623, 0); -lean_inc(x_624); -x_625 = lean_ctor_get(x_623, 1); -lean_inc(x_625); -lean_dec(x_623); -x_626 = lean_ctor_get(x_624, 0); -lean_inc(x_626); -lean_dec(x_624); -x_627 = lean_environment_main_module(x_626); -x_628 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__12; -x_629 = l_Lean_addMacroScope(x_627, x_628, x_590); -x_630 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__10; -x_631 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__14; -x_632 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_632, 0, x_621); -lean_ctor_set(x_632, 1, x_630); -lean_ctor_set(x_632, 2, x_629); -lean_ctor_set(x_632, 3, x_631); -x_633 = lean_array_push(x_605, x_614); -x_634 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_634, 0, x_607); -lean_ctor_set(x_634, 1, x_608); -lean_ctor_set(x_634, 2, x_633); -x_635 = lean_array_push(x_610, x_632); -x_636 = lean_array_push(x_635, x_634); -x_637 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_637, 0, x_607); -lean_ctor_set(x_637, 1, x_613); -lean_ctor_set(x_637, 2, x_636); -x_638 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_639 = lean_box(0); -x_640 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__21(x_14, x_638, x_608, x_16, x_605, x_610, x_613, x_585, x_12, x_18, x_637, x_639, x_4, x_5, x_6, x_7, x_8, x_9, x_625); -lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_640; +lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; +lean_inc(x_10); +x_635 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_10, x_11, x_606); +x_636 = lean_ctor_get(x_635, 0); +lean_inc(x_636); +x_637 = lean_ctor_get(x_635, 1); +lean_inc(x_637); +lean_dec(x_635); +x_638 = lean_st_ref_get(x_11, x_637); +x_639 = lean_ctor_get(x_638, 0); +lean_inc(x_639); +x_640 = lean_ctor_get(x_638, 1); +lean_inc(x_640); +lean_dec(x_638); +x_641 = lean_ctor_get(x_639, 0); +lean_inc(x_641); +lean_dec(x_639); +x_642 = lean_environment_main_module(x_641); +x_643 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__10; +x_644 = l_Lean_addMacroScope(x_642, x_643, x_603); +lean_inc(x_2); +x_645 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_645, 0, x_643); +lean_ctor_set(x_645, 1, x_2); +lean_inc(x_2); +x_646 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_646, 0, x_645); +lean_ctor_set(x_646, 1, x_2); +x_647 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__8; +x_648 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_648, 0, x_636); +lean_ctor_set(x_648, 1, x_647); +lean_ctor_set(x_648, 2, x_644); +lean_ctor_set(x_648, 3, x_646); +x_649 = lean_array_push(x_620, x_629); +x_650 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_650, 0, x_622); +lean_ctor_set(x_650, 1, x_623); +lean_ctor_set(x_650, 2, x_649); +x_651 = lean_array_push(x_625, x_648); +x_652 = lean_array_push(x_651, x_650); +x_653 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_653, 0, x_622); +lean_ctor_set(x_653, 1, x_628); +lean_ctor_set(x_653, 2, x_652); +x_654 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; +x_655 = lean_box(0); +x_656 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__21(x_16, x_654, x_623, x_2, x_620, x_625, x_628, x_598, x_14, x_3, x_19, x_653, x_655, x_6, x_7, x_8, x_9, x_10, x_11, x_640); +return x_656; } } } } default: { -lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; uint8_t x_672; -lean_dec(x_77); -x_641 = lean_ctor_get(x_20, 0); -lean_inc(x_641); -x_642 = lean_ctor_get(x_20, 1); -lean_inc(x_642); -lean_dec(x_20); -lean_inc(x_8); -x_643 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_8, x_9, x_642); -x_644 = lean_ctor_get(x_643, 0); -lean_inc(x_644); -x_645 = lean_ctor_get(x_643, 1); -lean_inc(x_645); -lean_dec(x_643); -x_646 = lean_ctor_get(x_8, 10); -lean_inc(x_646); -x_647 = lean_st_ref_get(x_9, x_645); -x_648 = lean_ctor_get(x_647, 0); -lean_inc(x_648); -x_649 = lean_ctor_get(x_647, 1); -lean_inc(x_649); -lean_dec(x_647); -x_650 = lean_ctor_get(x_648, 0); -lean_inc(x_650); -lean_dec(x_648); -x_651 = lean_environment_main_module(x_650); -x_652 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__4; -lean_inc(x_646); -lean_inc(x_651); -x_653 = l_Lean_addMacroScope(x_651, x_652, x_646); -x_654 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__3; -x_655 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__7; -lean_inc(x_644); -x_656 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_656, 0, x_644); -lean_ctor_set(x_656, 1, x_654); -lean_ctor_set(x_656, 2, x_653); -lean_ctor_set(x_656, 3, x_655); -x_657 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -lean_inc(x_646); -x_658 = l_Lean_addMacroScope(x_651, x_657, x_646); -x_659 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; -x_660 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_660, 0, x_644); -lean_ctor_set(x_660, 1, x_659); -lean_ctor_set(x_660, 2, x_658); -lean_ctor_set(x_660, 3, x_16); -x_661 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; -x_662 = lean_array_push(x_661, x_660); -x_663 = lean_box(2); -x_664 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_665 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_665, 0, x_663); -lean_ctor_set(x_665, 1, x_664); -lean_ctor_set(x_665, 2, x_662); -x_666 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_667 = lean_array_push(x_666, x_656); -x_668 = lean_array_push(x_667, x_665); -x_669 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; -x_670 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_670, 0, x_663); -lean_ctor_set(x_670, 1, x_669); -lean_ctor_set(x_670, 2, x_668); -x_671 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__3; -x_672 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__2(x_11, x_671); -lean_dec(x_11); -if (x_672 == 0) +lean_object* x_657; lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; lean_object* x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; lean_object* x_678; lean_object* x_679; lean_object* x_680; lean_object* x_681; lean_object* x_682; lean_object* x_683; lean_object* x_684; lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; uint8_t x_690; +lean_dec(x_81); +x_657 = lean_ctor_get(x_21, 0); +lean_inc(x_657); +x_658 = lean_ctor_get(x_21, 1); +lean_inc(x_658); +lean_dec(x_21); +lean_inc(x_10); +x_659 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_10, x_11, x_658); +x_660 = lean_ctor_get(x_659, 0); +lean_inc(x_660); +x_661 = lean_ctor_get(x_659, 1); +lean_inc(x_661); +lean_dec(x_659); +x_662 = lean_ctor_get(x_10, 10); +lean_inc(x_662); +x_663 = lean_st_ref_get(x_11, x_661); +x_664 = lean_ctor_get(x_663, 0); +lean_inc(x_664); +x_665 = lean_ctor_get(x_663, 1); +lean_inc(x_665); +lean_dec(x_663); +x_666 = lean_ctor_get(x_664, 0); +lean_inc(x_666); +lean_dec(x_664); +x_667 = lean_environment_main_module(x_666); +x_668 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__4; +lean_inc(x_662); +lean_inc(x_667); +x_669 = l_Lean_addMacroScope(x_667, x_668, x_662); +x_670 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__5; +lean_inc(x_2); +x_671 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_671, 0, x_670); +lean_ctor_set(x_671, 1, x_2); +lean_inc(x_2); +x_672 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_672, 0, x_671); +lean_ctor_set(x_672, 1, x_2); +x_673 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__3; +lean_inc(x_660); +x_674 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_674, 0, x_660); +lean_ctor_set(x_674, 1, x_673); +lean_ctor_set(x_674, 2, x_669); +lean_ctor_set(x_674, 3, x_672); +x_675 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +lean_inc(x_662); +x_676 = l_Lean_addMacroScope(x_667, x_675, x_662); +x_677 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; +lean_inc(x_2); +x_678 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_678, 0, x_660); +lean_ctor_set(x_678, 1, x_677); +lean_ctor_set(x_678, 2, x_676); +lean_ctor_set(x_678, 3, x_2); +x_679 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; +x_680 = lean_array_push(x_679, x_678); +x_681 = lean_box(2); +x_682 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_683 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_683, 0, x_681); +lean_ctor_set(x_683, 1, x_682); +lean_ctor_set(x_683, 2, x_680); +x_684 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_685 = lean_array_push(x_684, x_674); +x_686 = lean_array_push(x_685, x_683); +x_687 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__4; +x_688 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_688, 0, x_681); +lean_ctor_set(x_688, 1, x_687); +lean_ctor_set(x_688, 2, x_686); +x_689 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__3; +x_690 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5(x_13, x_689); +lean_dec(x_13); +if (x_690 == 0) { -lean_object* x_673; lean_object* x_674; lean_object* x_675; -lean_dec(x_646); -x_673 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_674 = lean_box(0); -x_675 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__22(x_14, x_673, x_664, x_16, x_661, x_666, x_669, x_641, x_12, x_18, x_670, x_674, x_4, x_5, x_6, x_7, x_8, x_9, x_649); -lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_675; +lean_object* x_691; lean_object* x_692; lean_object* x_693; +lean_dec(x_662); +x_691 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; +x_692 = lean_box(0); +x_693 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__22(x_16, x_691, x_682, x_2, x_679, x_684, x_687, x_657, x_14, x_3, x_19, x_688, x_692, x_6, x_7, x_8, x_9, x_10, x_11, x_665); +return x_693; } else { -lean_object* x_676; lean_object* x_677; lean_object* x_678; lean_object* x_679; lean_object* x_680; lean_object* x_681; lean_object* x_682; lean_object* x_683; lean_object* x_684; lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; lean_object* x_690; lean_object* x_691; lean_object* x_692; lean_object* x_693; lean_object* x_694; lean_object* x_695; lean_object* x_696; -lean_inc(x_8); -x_676 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_8, x_9, x_649); -x_677 = lean_ctor_get(x_676, 0); -lean_inc(x_677); -x_678 = lean_ctor_get(x_676, 1); -lean_inc(x_678); -lean_dec(x_676); -x_679 = lean_st_ref_get(x_9, x_678); -x_680 = lean_ctor_get(x_679, 0); -lean_inc(x_680); -x_681 = lean_ctor_get(x_679, 1); -lean_inc(x_681); -lean_dec(x_679); -x_682 = lean_ctor_get(x_680, 0); -lean_inc(x_682); -lean_dec(x_680); -x_683 = lean_environment_main_module(x_682); -x_684 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__12; -x_685 = l_Lean_addMacroScope(x_683, x_684, x_646); -x_686 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__10; -x_687 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__14; -x_688 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_688, 0, x_677); -lean_ctor_set(x_688, 1, x_686); -lean_ctor_set(x_688, 2, x_685); -lean_ctor_set(x_688, 3, x_687); -x_689 = lean_array_push(x_661, x_670); -x_690 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_690, 0, x_663); -lean_ctor_set(x_690, 1, x_664); -lean_ctor_set(x_690, 2, x_689); -x_691 = lean_array_push(x_666, x_688); -x_692 = lean_array_push(x_691, x_690); -x_693 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_693, 0, x_663); -lean_ctor_set(x_693, 1, x_669); -lean_ctor_set(x_693, 2, x_692); -x_694 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_695 = lean_box(0); -x_696 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__22(x_14, x_694, x_664, x_16, x_661, x_666, x_669, x_641, x_12, x_18, x_693, x_695, x_4, x_5, x_6, x_7, x_8, x_9, x_681); -lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_696; +lean_object* x_694; lean_object* x_695; lean_object* x_696; lean_object* x_697; lean_object* x_698; lean_object* x_699; lean_object* x_700; lean_object* x_701; lean_object* x_702; lean_object* x_703; lean_object* x_704; lean_object* x_705; lean_object* x_706; lean_object* x_707; lean_object* x_708; lean_object* x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; +lean_inc(x_10); +x_694 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_10, x_11, x_665); +x_695 = lean_ctor_get(x_694, 0); +lean_inc(x_695); +x_696 = lean_ctor_get(x_694, 1); +lean_inc(x_696); +lean_dec(x_694); +x_697 = lean_st_ref_get(x_11, x_696); +x_698 = lean_ctor_get(x_697, 0); +lean_inc(x_698); +x_699 = lean_ctor_get(x_697, 1); +lean_inc(x_699); +lean_dec(x_697); +x_700 = lean_ctor_get(x_698, 0); +lean_inc(x_700); +lean_dec(x_698); +x_701 = lean_environment_main_module(x_700); +x_702 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__10; +x_703 = l_Lean_addMacroScope(x_701, x_702, x_662); +lean_inc(x_2); +x_704 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_704, 0, x_702); +lean_ctor_set(x_704, 1, x_2); +lean_inc(x_2); +x_705 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_705, 0, x_704); +lean_ctor_set(x_705, 1, x_2); +x_706 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__8; +x_707 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_707, 0, x_695); +lean_ctor_set(x_707, 1, x_706); +lean_ctor_set(x_707, 2, x_703); +lean_ctor_set(x_707, 3, x_705); +x_708 = lean_array_push(x_679, x_688); +x_709 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_709, 0, x_681); +lean_ctor_set(x_709, 1, x_682); +lean_ctor_set(x_709, 2, x_708); +x_710 = lean_array_push(x_684, x_707); +x_711 = lean_array_push(x_710, x_709); +x_712 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_712, 0, x_681); +lean_ctor_set(x_712, 1, x_687); +lean_ctor_set(x_712, 2, x_711); +x_713 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; +x_714 = lean_box(0); +x_715 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__22(x_16, x_713, x_682, x_2, x_679, x_684, x_687, x_657, x_14, x_3, x_19, x_712, x_714, x_6, x_7, x_8, x_9, x_10, x_11, x_699); +return x_715; } } } @@ -33125,75 +34803,81 @@ return x_696; } else { -uint8_t x_697; -lean_dec(x_18); +uint8_t x_716; +lean_dec(x_19); +lean_dec(x_16); lean_dec(x_14); -lean_dec(x_12); +lean_dec(x_13); lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_697 = !lean_is_exclusive(x_20); -if (x_697 == 0) +lean_dec(x_3); +lean_dec(x_2); +x_716 = !lean_is_exclusive(x_21); +if (x_716 == 0) { -return x_20; +return x_21; } else { -lean_object* x_698; lean_object* x_699; lean_object* x_700; -x_698 = lean_ctor_get(x_20, 0); -x_699 = lean_ctor_get(x_20, 1); -lean_inc(x_699); -lean_inc(x_698); -lean_dec(x_20); -x_700 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_700, 0, x_698); -lean_ctor_set(x_700, 1, x_699); -return x_700; +lean_object* x_717; lean_object* x_718; lean_object* x_719; +x_717 = lean_ctor_get(x_21, 0); +x_718 = lean_ctor_get(x_21, 1); +lean_inc(x_718); +lean_inc(x_717); +lean_dec(x_21); +x_719 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_719, 0, x_717); +lean_ctor_set(x_719, 1, x_718); +return x_719; } } } else { -uint8_t x_701; +uint8_t x_720; +lean_dec(x_16); lean_dec(x_14); -lean_dec(x_12); +lean_dec(x_13); lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); lean_dec(x_3); -x_701 = !lean_is_exclusive(x_17); -if (x_701 == 0) +lean_dec(x_2); +x_720 = !lean_is_exclusive(x_18); +if (x_720 == 0) { -return x_17; +return x_18; } else { -lean_object* x_702; lean_object* x_703; lean_object* x_704; -x_702 = lean_ctor_get(x_17, 0); -x_703 = lean_ctor_get(x_17, 1); -lean_inc(x_703); -lean_inc(x_702); -lean_dec(x_17); -x_704 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_704, 0, x_702); -lean_ctor_set(x_704, 1, x_703); -return x_704; +lean_object* x_721; lean_object* x_722; lean_object* x_723; +x_721 = lean_ctor_get(x_18, 0); +x_722 = lean_ctor_get(x_18, 1); +lean_inc(x_722); +lean_inc(x_721); +lean_dec(x_18); +x_723 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_723, 0, x_721); +lean_ctor_set(x_723, 1, x_722); +return x_723; } } } else { -uint8_t x_705; -lean_dec(x_12); +uint8_t x_724; +lean_dec(x_14); +lean_dec(x_13); lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -33202,23 +34886,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_705 = !lean_is_exclusive(x_13); -if (x_705 == 0) +x_724 = !lean_is_exclusive(x_15); +if (x_724 == 0) { -return x_13; +return x_15; } else { -lean_object* x_706; lean_object* x_707; lean_object* x_708; -x_706 = lean_ctor_get(x_13, 0); -x_707 = lean_ctor_get(x_13, 1); -lean_inc(x_707); -lean_inc(x_706); -lean_dec(x_13); -x_708 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_708, 0, x_706); -lean_ctor_set(x_708, 1, x_707); -return x_708; +lean_object* x_725; lean_object* x_726; lean_object* x_727; +x_725 = lean_ctor_get(x_15, 0); +x_726 = lean_ctor_get(x_15, 1); +lean_inc(x_726); +lean_inc(x_725); +lean_dec(x_15); +x_727 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_727, 0, x_725); +lean_ctor_set(x_727, 1, x_726); +return x_727; } } } @@ -33258,7 +34942,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__26; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__26; x_2 = l_Lean_Elab_Term_Quotation_mkTuple___closed__6; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -33268,7 +34952,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__31; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__31; x_2 = l_Lean_Elab_Term_Quotation_mkTuple___closed__6; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -33309,7 +34993,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__39; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__39; x_2 = l_Lean_Elab_Term_Quotation_mkTuple___closed__6; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -33319,7 +35003,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__42; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__42; x_2 = l_Lean_Elab_Term_Quotation_mkTuple___closed__6; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -33378,7 +35062,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__30; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__30; x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__14; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -33396,8 +35080,8 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_15 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__12; -x_16 = l_Lean_throwErrorAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__10(x_2, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +x_15 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__12; +x_16 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__22(x_2, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); return x_16; } else @@ -33416,17 +35100,17 @@ if (lean_obj_tag(x_18) == 0) lean_object* x_19; lean_object* x_20; uint8_t x_21; x_19 = lean_ctor_get(x_17, 1); lean_inc(x_19); -x_20 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__15; +x_20 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__15; x_21 = lean_string_dec_eq(x_19, x_20); if (x_21 == 0) { lean_object* x_22; uint8_t x_23; -x_22 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__18; +x_22 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__18; x_23 = lean_string_dec_eq(x_19, x_22); if (x_23 == 0) { lean_object* x_24; uint8_t x_25; -x_24 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__1; +x_24 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__1; x_25 = lean_string_dec_eq(x_19, x_24); lean_dec(x_19); if (x_25 == 0) @@ -33438,23 +35122,23 @@ lean_dec(x_4); lean_dec(x_3); x_26 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_26, 0, x_17); -x_27 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__15; +x_27 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__15; x_28 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_28, 0, x_27); lean_ctor_set(x_28, 1, x_26); -x_29 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__17; +x_29 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__17; x_30 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_30, 0, x_28); lean_ctor_set(x_30, 1, x_29); -x_31 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__5; +x_31 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__5; x_32 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_32, 0, x_31); lean_ctor_set(x_32, 1, x_30); -x_33 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__11; +x_33 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__11; x_34 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_34, 0, x_32); lean_ctor_set(x_34, 1, x_33); -x_35 = l_Lean_throwErrorAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__10(x_2, x_34, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +x_35 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__22(x_2, x_34, x_7, x_8, x_9, x_10, x_11, x_12, x_13); return x_35; } else @@ -33497,7 +35181,7 @@ lean_inc(x_37); x_48 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_48, 0, x_37); lean_ctor_set(x_48, 1, x_47); -x_49 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__21; +x_49 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__21; lean_inc(x_37); x_50 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_50, 0, x_37); @@ -33526,16 +35210,16 @@ x_58 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFrom x_59 = lean_array_push(x_58, x_50); x_60 = lean_array_push(x_59, x_57); x_61 = lean_box(2); -x_62 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__20; +x_62 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__20; x_63 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_63, 0, x_61); lean_ctor_set(x_63, 1, x_62); lean_ctor_set(x_63, 2, x_60); -x_64 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8(x_4); +x_64 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24(x_4); x_65 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepFromSplice(x_1); x_66 = l_Lean_Syntax_mkStrLit(x_65, x_61); lean_dec(x_65); -x_67 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; +x_67 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; lean_inc(x_37); x_68 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_68, 0, x_37); @@ -33590,16 +35274,16 @@ x_91 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_91, 0, x_61); lean_ctor_set(x_91, 1, x_82); lean_ctor_set(x_91, 2, x_90); -x_92 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; +x_92 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; lean_inc(x_37); x_93 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_93, 0, x_37); lean_ctor_set(x_93, 1, x_92); -x_94 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; +x_94 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; x_95 = lean_array_push(x_94, x_68); x_96 = lean_array_push(x_95, x_91); x_97 = lean_array_push(x_96, x_93); -x_98 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__6; +x_98 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6; x_99 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_99, 0, x_61); lean_ctor_set(x_99, 1, x_98); @@ -33673,7 +35357,7 @@ lean_inc(x_37); x_134 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_134, 0, x_37); lean_ctor_set(x_134, 1, x_133); -x_135 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__21; +x_135 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__21; lean_inc(x_37); x_136 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_136, 0, x_37); @@ -33702,16 +35386,16 @@ x_144 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFro x_145 = lean_array_push(x_144, x_136); x_146 = lean_array_push(x_145, x_143); x_147 = lean_box(2); -x_148 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__20; +x_148 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__20; x_149 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_149, 0, x_147); lean_ctor_set(x_149, 1, x_148); lean_ctor_set(x_149, 2, x_146); -x_150 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8(x_4); +x_150 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24(x_4); x_151 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepFromSplice(x_1); x_152 = l_Lean_Syntax_mkStrLit(x_151, x_147); lean_dec(x_151); -x_153 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; +x_153 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; lean_inc(x_37); x_154 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_154, 0, x_37); @@ -33766,16 +35450,16 @@ x_177 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_177, 0, x_147); lean_ctor_set(x_177, 1, x_168); lean_ctor_set(x_177, 2, x_176); -x_178 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; +x_178 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; lean_inc(x_37); x_179 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_179, 0, x_37); lean_ctor_set(x_179, 1, x_178); -x_180 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; +x_180 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; x_181 = lean_array_push(x_180, x_154); x_182 = lean_array_push(x_181, x_177); x_183 = lean_array_push(x_182, x_179); -x_184 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__6; +x_184 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6; x_185 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_185, 0, x_147); lean_ctor_set(x_185, 1, x_184); @@ -33872,7 +35556,7 @@ lean_inc(x_215); x_226 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_226, 0, x_215); lean_ctor_set(x_226, 1, x_225); -x_227 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__21; +x_227 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__21; lean_inc(x_215); x_228 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_228, 0, x_215); @@ -33901,13 +35585,13 @@ x_236 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFro x_237 = lean_array_push(x_236, x_228); x_238 = lean_array_push(x_237, x_235); x_239 = lean_box(2); -x_240 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__20; +x_240 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__20; x_241 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_241, 0, x_239); lean_ctor_set(x_241, 1, x_240); lean_ctor_set(x_241, 2, x_238); -x_242 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8(x_4); -x_243 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; +x_242 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24(x_4); +x_243 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; lean_inc(x_215); x_244 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_244, 0, x_215); @@ -33962,16 +35646,16 @@ x_267 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_267, 0, x_239); lean_ctor_set(x_267, 1, x_258); lean_ctor_set(x_267, 2, x_266); -x_268 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; +x_268 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; lean_inc(x_215); x_269 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_269, 0, x_215); lean_ctor_set(x_269, 1, x_268); -x_270 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; +x_270 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; x_271 = lean_array_push(x_270, x_244); x_272 = lean_array_push(x_271, x_267); x_273 = lean_array_push(x_272, x_269); -x_274 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__6; +x_274 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6; x_275 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_275, 0, x_239); lean_ctor_set(x_275, 1, x_274); @@ -34044,7 +35728,7 @@ lean_inc(x_215); x_309 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_309, 0, x_215); lean_ctor_set(x_309, 1, x_308); -x_310 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__21; +x_310 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__21; lean_inc(x_215); x_311 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_311, 0, x_215); @@ -34073,13 +35757,13 @@ x_319 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFro x_320 = lean_array_push(x_319, x_311); x_321 = lean_array_push(x_320, x_318); x_322 = lean_box(2); -x_323 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__20; +x_323 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__20; x_324 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_324, 0, x_322); lean_ctor_set(x_324, 1, x_323); lean_ctor_set(x_324, 2, x_321); -x_325 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8(x_4); -x_326 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; +x_325 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24(x_4); +x_326 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; lean_inc(x_215); x_327 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_327, 0, x_215); @@ -34134,16 +35818,16 @@ x_350 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_350, 0, x_322); lean_ctor_set(x_350, 1, x_341); lean_ctor_set(x_350, 2, x_349); -x_351 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; +x_351 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; lean_inc(x_215); x_352 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_352, 0, x_215); lean_ctor_set(x_352, 1, x_351); -x_353 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; +x_353 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; x_354 = lean_array_push(x_353, x_327); x_355 = lean_array_push(x_354, x_350); x_356 = lean_array_push(x_355, x_352); -x_357 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__6; +x_357 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6; x_358 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_358, 0, x_322); lean_ctor_set(x_358, 1, x_357); @@ -34239,7 +35923,7 @@ lean_inc(x_387); x_398 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_398, 0, x_387); lean_ctor_set(x_398, 1, x_397); -x_399 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__50; +x_399 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__50; lean_inc(x_389); lean_inc(x_394); x_400 = l_Lean_addMacroScope(x_394, x_399, x_389); @@ -34251,19 +35935,19 @@ lean_inc(x_3); x_402 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_402, 0, x_401); lean_ctor_set(x_402, 1, x_3); -x_403 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__49; +x_403 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__49; lean_inc(x_387); x_404 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_404, 0, x_387); lean_ctor_set(x_404, 1, x_403); lean_ctor_set(x_404, 2, x_400); lean_ctor_set(x_404, 3, x_402); -x_405 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; +x_405 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; lean_inc(x_387); x_406 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_406, 0, x_387); lean_ctor_set(x_406, 1, x_405); -x_407 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__21; +x_407 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__21; lean_inc(x_387); x_408 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_408, 0, x_387); @@ -34292,12 +35976,12 @@ x_416 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFro x_417 = lean_array_push(x_416, x_408); x_418 = lean_array_push(x_417, x_415); x_419 = lean_box(2); -x_420 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__20; +x_420 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__20; x_421 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_421, 0, x_419); lean_ctor_set(x_421, 1, x_420); lean_ctor_set(x_421, 2, x_418); -x_422 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8(x_4); +x_422 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24(x_4); x_423 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; x_424 = lean_array_push(x_423, x_422); x_425 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; @@ -34319,18 +36003,18 @@ x_434 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_434, 0, x_419); lean_ctor_set(x_434, 1, x_425); lean_ctor_set(x_434, 2, x_433); -x_435 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; +x_435 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; lean_inc(x_387); x_436 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_436, 0, x_387); lean_ctor_set(x_436, 1, x_435); -x_437 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; +x_437 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; x_438 = lean_array_push(x_437, x_406); lean_inc(x_438); x_439 = lean_array_push(x_438, x_434); lean_inc(x_436); x_440 = lean_array_push(x_439, x_436); -x_441 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__6; +x_441 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6; x_442 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_442, 0, x_419); lean_ctor_set(x_442, 1, x_441); @@ -34455,7 +36139,7 @@ lean_inc(x_387); x_498 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_498, 0, x_387); lean_ctor_set(x_498, 1, x_497); -x_499 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__50; +x_499 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__50; lean_inc(x_389); lean_inc(x_494); x_500 = l_Lean_addMacroScope(x_494, x_499, x_389); @@ -34467,19 +36151,19 @@ lean_inc(x_3); x_502 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_502, 0, x_501); lean_ctor_set(x_502, 1, x_3); -x_503 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__49; +x_503 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__49; lean_inc(x_387); x_504 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_504, 0, x_387); lean_ctor_set(x_504, 1, x_503); lean_ctor_set(x_504, 2, x_500); lean_ctor_set(x_504, 3, x_502); -x_505 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; +x_505 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; lean_inc(x_387); x_506 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_506, 0, x_387); lean_ctor_set(x_506, 1, x_505); -x_507 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__21; +x_507 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__21; lean_inc(x_387); x_508 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_508, 0, x_387); @@ -34508,12 +36192,12 @@ x_516 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFro x_517 = lean_array_push(x_516, x_508); x_518 = lean_array_push(x_517, x_515); x_519 = lean_box(2); -x_520 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__20; +x_520 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__20; x_521 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_521, 0, x_519); lean_ctor_set(x_521, 1, x_520); lean_ctor_set(x_521, 2, x_518); -x_522 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8(x_4); +x_522 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24(x_4); x_523 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; x_524 = lean_array_push(x_523, x_522); x_525 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; @@ -34535,18 +36219,18 @@ x_534 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_534, 0, x_519); lean_ctor_set(x_534, 1, x_525); lean_ctor_set(x_534, 2, x_533); -x_535 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14; +x_535 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14; lean_inc(x_387); x_536 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_536, 0, x_387); lean_ctor_set(x_536, 1, x_535); -x_537 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; +x_537 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; x_538 = lean_array_push(x_537, x_506); lean_inc(x_538); x_539 = lean_array_push(x_538, x_534); lean_inc(x_536); x_540 = lean_array_push(x_539, x_536); -x_541 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__6; +x_541 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6; x_542 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_542, 0, x_519); lean_ctor_set(x_542, 1, x_541); @@ -34663,23 +36347,23 @@ lean_dec(x_4); lean_dec(x_3); x_592 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_592, 0, x_17); -x_593 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__15; +x_593 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__15; x_594 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_594, 0, x_593); lean_ctor_set(x_594, 1, x_592); -x_595 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__17; +x_595 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__17; x_596 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_596, 0, x_594); lean_ctor_set(x_596, 1, x_595); -x_597 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__5; +x_597 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__5; x_598 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_598, 0, x_597); lean_ctor_set(x_598, 1, x_596); -x_599 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__11; +x_599 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__11; x_600 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_600, 0, x_598); lean_ctor_set(x_600, 1, x_599); -x_601 = l_Lean_throwErrorAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__10(x_2, x_600, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +x_601 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__22(x_2, x_600, x_7, x_8, x_9, x_10, x_11, x_12, x_13); return x_601; } } @@ -34692,37 +36376,29 @@ lean_dec(x_4); lean_dec(x_3); x_602 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_602, 0, x_17); -x_603 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__15; +x_603 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__15; x_604 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_604, 0, x_603); lean_ctor_set(x_604, 1, x_602); -x_605 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__17; +x_605 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__17; x_606 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_606, 0, x_604); lean_ctor_set(x_606, 1, x_605); -x_607 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__5; +x_607 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__5; x_608 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_608, 0, x_607); lean_ctor_set(x_608, 1, x_606); -x_609 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__11; +x_609 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__11; x_610 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_610, 0, x_608); lean_ctor_set(x_610, 1, x_609); -x_611 = l_Lean_throwErrorAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__10(x_2, x_610, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +x_611 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__22(x_2, x_610, x_7, x_8, x_9, x_10, x_11, x_12, x_13); return x_611; } } } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25(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_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = lean_apply_8(x_2, x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_11; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { switch (lean_obj_tag(x_3)) { @@ -34794,7 +36470,7 @@ return x_16; } } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__1() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__1() { _start: { lean_object* x_1; @@ -34802,22 +36478,22 @@ x_1 = lean_mk_string_from_bytes("false", 5); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__2() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__1; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__1; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__3() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__1; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__2; +x_3 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__2; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -34825,51 +36501,27 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__4() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__1; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__5() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__18; -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__1; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__14; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__6() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__5; -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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__7() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__6; -x_3 = lean_alloc_ctor(1, 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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__8() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__6() { _start: { lean_object* x_1; @@ -34877,22 +36529,22 @@ x_1 = lean_mk_string_from_bytes("cond", 4); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__9() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__8; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__6; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__10() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__8; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__6; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__9; +x_3 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__7; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -34900,267 +36552,266 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__11() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__8; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__6; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__12() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__11; -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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__13() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__12; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -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; 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; lean_object* x_30; lean_object* x_31; lean_object* x_32; -lean_inc(x_10); -x_13 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_10, x_11, x_12); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = lean_ctor_get(x_10, 10); +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; 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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +lean_inc(x_12); +x_15 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_12, x_13, x_14); +x_16 = lean_ctor_get(x_15, 0); lean_inc(x_16); -x_17 = lean_st_ref_get(x_11, x_15); -x_18 = lean_ctor_get(x_17, 0); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_ctor_get(x_12, 10); lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); -lean_inc(x_19); -lean_dec(x_17); -x_20 = lean_ctor_get(x_18, 0); +x_19 = lean_st_ref_get(x_13, x_17); +x_20 = lean_ctor_get(x_19, 0); lean_inc(x_20); -lean_dec(x_18); -x_21 = lean_environment_main_module(x_20); -x_22 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__4; -lean_inc(x_16); -x_23 = l_Lean_addMacroScope(x_21, x_22, x_16); -x_24 = lean_box(0); -x_25 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__3; -x_26 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__7; -x_27 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_27, 0, x_14); -lean_ctor_set(x_27, 1, x_25); -lean_ctor_set(x_27, 2, x_23); -lean_ctor_set(x_27, 3, x_26); -x_28 = lean_unsigned_to_nat(0u); -lean_inc(x_10); -lean_inc(x_2); -x_29 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15(x_1, x_2, x_28, x_24, x_24, x_27, x_3, x_6, x_7, x_8, x_9, x_10, x_11, x_19); -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_29, 1); -lean_inc(x_31); -lean_dec(x_29); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_32 = lean_apply_8(x_4, x_24, x_6, x_7, x_8, x_9, x_10, x_11, x_31); -if (lean_obj_tag(x_32) == 0) -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = lean_ctor_get(x_20, 0); +lean_inc(x_22); +lean_dec(x_20); +x_23 = lean_environment_main_module(x_22); +x_24 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__4; +lean_inc(x_18); +x_25 = l_Lean_addMacroScope(x_23, x_24, x_18); +x_26 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__5; +lean_inc(x_1); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_1); +lean_inc(x_1); +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_1); +x_29 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__3; +x_30 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_30, 0, x_16); +lean_ctor_set(x_30, 1, x_29); +lean_ctor_set(x_30, 2, x_25); +lean_ctor_set(x_30, 3, x_28); +x_31 = lean_unsigned_to_nat(0u); +lean_inc(x_12); +lean_inc_n(x_1, 2); +lean_inc(x_4); +x_32 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15(x_2, x_3, x_4, x_31, x_1, x_1, x_30, x_5, x_8, x_9, x_10, x_11, x_12, x_13, x_21); x_33 = lean_ctor_get(x_32, 0); lean_inc(x_33); x_34 = lean_ctor_get(x_32, 1); lean_inc(x_34); lean_dec(x_32); +lean_inc(x_13); +lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); -x_35 = lean_apply_7(x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_34); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_35 = lean_apply_8(x_6, x_1, x_8, x_9, x_10, x_11, x_12, x_13, x_34); if (lean_obj_tag(x_35) == 0) { -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; +lean_object* x_36; lean_object* x_37; lean_object* x_38; x_36 = lean_ctor_get(x_35, 0); lean_inc(x_36); x_37 = lean_ctor_get(x_35, 1); lean_inc(x_37); lean_dec(x_35); -x_38 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_10, x_11, x_37); +lean_inc(x_13); +lean_inc(x_12); +x_38 = lean_apply_7(x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_37); +if (lean_obj_tag(x_38) == 0) +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; x_39 = lean_ctor_get(x_38, 0); lean_inc(x_39); x_40 = lean_ctor_get(x_38, 1); lean_inc(x_40); lean_dec(x_38); -x_41 = lean_st_ref_get(x_11, x_40); -lean_dec(x_11); -x_42 = !lean_is_exclusive(x_41); -if (x_42 == 0) +x_41 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_12, x_13, x_40); +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_41, 1); +lean_inc(x_43); +lean_dec(x_41); +x_44 = lean_st_ref_get(x_13, x_43); +lean_dec(x_13); +x_45 = !lean_is_exclusive(x_44); +if (x_45 == 0) { -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_43 = lean_ctor_get(x_41, 0); -x_44 = lean_ctor_get(x_43, 0); -lean_inc(x_44); -lean_dec(x_43); -x_45 = lean_environment_main_module(x_44); -x_46 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__3; -x_47 = lean_name_mk_string(x_2, x_46); -x_48 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__11; -x_49 = l_Lean_addMacroScope(x_45, x_48, x_16); -x_50 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__10; -x_51 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__13; -x_52 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_52, 0, x_39); -lean_ctor_set(x_52, 1, x_50); -lean_ctor_set(x_52, 2, x_49); -lean_ctor_set(x_52, 3, x_51); -x_53 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; -x_54 = lean_array_push(x_53, x_30); -x_55 = lean_array_push(x_54, x_33); -x_56 = lean_array_push(x_55, x_36); -x_57 = lean_box(2); -x_58 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_59 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_59, 0, x_57); -lean_ctor_set(x_59, 1, x_58); -lean_ctor_set(x_59, 2, x_56); -x_60 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_61 = lean_array_push(x_60, x_52); -x_62 = lean_array_push(x_61, x_59); +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; lean_object* x_67; +x_46 = lean_ctor_get(x_44, 0); +x_47 = lean_ctor_get(x_46, 0); +lean_inc(x_47); +lean_dec(x_46); +x_48 = lean_environment_main_module(x_47); +x_49 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__3; +x_50 = lean_name_mk_string(x_4, x_49); +x_51 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__9; +x_52 = l_Lean_addMacroScope(x_48, x_51, x_18); +lean_inc(x_1); +x_53 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_1); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_53); +lean_ctor_set(x_54, 1, x_1); +x_55 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__8; +x_56 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_56, 0, x_42); +lean_ctor_set(x_56, 1, x_55); +lean_ctor_set(x_56, 2, x_52); +lean_ctor_set(x_56, 3, x_54); +x_57 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; +x_58 = lean_array_push(x_57, x_33); +x_59 = lean_array_push(x_58, x_36); +x_60 = lean_array_push(x_59, x_39); +x_61 = lean_box(2); +x_62 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; x_63 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_63, 0, x_57); -lean_ctor_set(x_63, 1, x_47); -lean_ctor_set(x_63, 2, x_62); -lean_ctor_set(x_41, 0, x_63); -return x_41; +lean_ctor_set(x_63, 0, x_61); +lean_ctor_set(x_63, 1, x_62); +lean_ctor_set(x_63, 2, x_60); +x_64 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_65 = lean_array_push(x_64, x_56); +x_66 = lean_array_push(x_65, x_63); +x_67 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_67, 0, x_61); +lean_ctor_set(x_67, 1, x_50); +lean_ctor_set(x_67, 2, x_66); +lean_ctor_set(x_44, 0, x_67); +return x_44; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; -x_64 = lean_ctor_get(x_41, 0); -x_65 = lean_ctor_get(x_41, 1); -lean_inc(x_65); -lean_inc(x_64); -lean_dec(x_41); -x_66 = lean_ctor_get(x_64, 0); -lean_inc(x_66); -lean_dec(x_64); -x_67 = lean_environment_main_module(x_66); -x_68 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__3; -x_69 = lean_name_mk_string(x_2, x_68); -x_70 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__11; -x_71 = l_Lean_addMacroScope(x_67, x_70, x_16); -x_72 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__10; -x_73 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__13; -x_74 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_74, 0, x_39); -lean_ctor_set(x_74, 1, x_72); -lean_ctor_set(x_74, 2, x_71); -lean_ctor_set(x_74, 3, x_73); -x_75 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13; -x_76 = lean_array_push(x_75, x_30); -x_77 = lean_array_push(x_76, x_33); -x_78 = lean_array_push(x_77, x_36); -x_79 = lean_box(2); -x_80 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_81 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_81, 0, x_79); -lean_ctor_set(x_81, 1, x_80); -lean_ctor_set(x_81, 2, x_78); -x_82 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_83 = lean_array_push(x_82, x_74); -x_84 = lean_array_push(x_83, x_81); -x_85 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_85, 0, x_79); -lean_ctor_set(x_85, 1, x_69); -lean_ctor_set(x_85, 2, x_84); -x_86 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_86, 0, x_85); -lean_ctor_set(x_86, 1, x_65); -return x_86; +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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_68 = lean_ctor_get(x_44, 0); +x_69 = lean_ctor_get(x_44, 1); +lean_inc(x_69); +lean_inc(x_68); +lean_dec(x_44); +x_70 = lean_ctor_get(x_68, 0); +lean_inc(x_70); +lean_dec(x_68); +x_71 = lean_environment_main_module(x_70); +x_72 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__3; +x_73 = lean_name_mk_string(x_4, x_72); +x_74 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__9; +x_75 = l_Lean_addMacroScope(x_71, x_74, x_18); +lean_inc(x_1); +x_76 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_76, 0, x_74); +lean_ctor_set(x_76, 1, x_1); +x_77 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_77, 0, x_76); +lean_ctor_set(x_77, 1, x_1); +x_78 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__8; +x_79 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_79, 0, x_42); +lean_ctor_set(x_79, 1, x_78); +lean_ctor_set(x_79, 2, x_75); +lean_ctor_set(x_79, 3, x_77); +x_80 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13; +x_81 = lean_array_push(x_80, x_33); +x_82 = lean_array_push(x_81, x_36); +x_83 = lean_array_push(x_82, x_39); +x_84 = lean_box(2); +x_85 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_86 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_86, 0, x_84); +lean_ctor_set(x_86, 1, x_85); +lean_ctor_set(x_86, 2, x_83); +x_87 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_88 = lean_array_push(x_87, x_79); +x_89 = lean_array_push(x_88, x_86); +x_90 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_90, 0, x_84); +lean_ctor_set(x_90, 1, x_73); +lean_ctor_set(x_90, 2, x_89); +x_91 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_91, 0, x_90); +lean_ctor_set(x_91, 1, x_69); +return x_91; } } else { -uint8_t x_87; +uint8_t x_92; +lean_dec(x_36); lean_dec(x_33); -lean_dec(x_30); -lean_dec(x_16); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_2); -x_87 = !lean_is_exclusive(x_35); -if (x_87 == 0) +lean_dec(x_18); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_4); +lean_dec(x_1); +x_92 = !lean_is_exclusive(x_38); +if (x_92 == 0) { -return x_35; +return x_38; } else { -lean_object* x_88; lean_object* x_89; lean_object* x_90; -x_88 = lean_ctor_get(x_35, 0); -x_89 = lean_ctor_get(x_35, 1); -lean_inc(x_89); -lean_inc(x_88); -lean_dec(x_35); -x_90 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_90, 0, x_88); -lean_ctor_set(x_90, 1, x_89); -return x_90; +lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_93 = lean_ctor_get(x_38, 0); +x_94 = lean_ctor_get(x_38, 1); +lean_inc(x_94); +lean_inc(x_93); +lean_dec(x_38); +x_95 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_95, 0, x_93); +lean_ctor_set(x_95, 1, x_94); +return x_95; } } } else { -uint8_t x_91; -lean_dec(x_30); -lean_dec(x_16); +uint8_t x_96; +lean_dec(x_33); +lean_dec(x_18); +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_2); -x_91 = !lean_is_exclusive(x_32); -if (x_91 == 0) +lean_dec(x_4); +lean_dec(x_1); +x_96 = !lean_is_exclusive(x_35); +if (x_96 == 0) { -return x_32; +return x_35; } else { -lean_object* x_92; lean_object* x_93; lean_object* x_94; -x_92 = lean_ctor_get(x_32, 0); -x_93 = lean_ctor_get(x_32, 1); -lean_inc(x_93); -lean_inc(x_92); -lean_dec(x_32); -x_94 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_94, 0, x_92); -lean_ctor_set(x_94, 1, x_93); -return x_94; +lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_97 = lean_ctor_get(x_35, 0); +x_98 = lean_ctor_get(x_35, 1); +lean_inc(x_98); +lean_inc(x_97); +lean_dec(x_35); +x_99 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_99, 0, x_97); +lean_ctor_set(x_99, 1, x_98); +return x_99; } } } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__28___closed__1() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__1() { _start: { lean_object* x_1; @@ -35168,116 +36819,117 @@ x_1 = lean_mk_string_from_bytes("unsupported antiquotation kind in pattern", 41) return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__28___closed__2() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__28___closed__1; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__1; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__28(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_object* x_9) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27(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_object* x_9) { _start: { lean_object* x_10; lean_object* x_11; -x_10 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__28___closed__2; -x_11 = l_Lean_throwErrorAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__10(x_1, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_10 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__2; +x_11 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__22(x_1, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__29(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__28(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -lean_inc(x_10); -x_13 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_10, x_11, x_12); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +lean_inc(x_11); +x_14 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_11, x_12, x_13); +x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); -lean_dec(x_13); -x_16 = lean_ctor_get(x_10, 10); +x_16 = lean_ctor_get(x_14, 1); lean_inc(x_16); -lean_dec(x_10); -x_17 = lean_st_ref_get(x_11, x_15); -x_18 = !lean_is_exclusive(x_17); -if (x_18 == 0) +lean_dec(x_14); +x_17 = lean_ctor_get(x_11, 10); +lean_inc(x_17); +lean_dec(x_11); +x_18 = lean_st_ref_get(x_12, x_16); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) { -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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; -x_19 = lean_ctor_get(x_17, 0); -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -lean_dec(x_19); -x_21 = lean_environment_main_module(x_20); -x_22 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__7; +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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_20 = lean_ctor_get(x_18, 0); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +lean_dec(x_20); +x_22 = lean_environment_main_module(x_21); +x_23 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__7; lean_inc(x_1); -x_23 = lean_name_mk_string(x_1, x_22); -lean_inc(x_14); -x_24 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_24, 0, x_14); -lean_ctor_set(x_24, 1, x_22); -x_25 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__9; +x_24 = lean_name_mk_string(x_1, x_23); +lean_inc(x_15); +x_25 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_25, 0, x_15); +lean_ctor_set(x_25, 1, x_23); +x_26 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__9; lean_inc(x_1); -x_26 = lean_name_mk_string(x_1, x_25); -x_27 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; +x_27 = lean_name_mk_string(x_1, x_26); +x_28 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; lean_inc(x_1); -x_28 = lean_name_mk_string(x_1, x_27); -x_29 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__17; -lean_inc(x_14); -x_30 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_30, 0, x_14); -lean_ctor_set(x_30, 1, x_29); -x_31 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__3; +x_29 = lean_name_mk_string(x_1, x_28); +x_30 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__17; +lean_inc(x_15); +x_31 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_31, 0, x_15); +lean_ctor_set(x_31, 1, x_30); +x_32 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__3; lean_inc(x_1); -x_32 = lean_name_mk_string(x_1, x_31); -x_33 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__19; -x_34 = lean_name_mk_string(x_1, x_33); -x_35 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__21; -lean_inc(x_14); -x_36 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_36, 0, x_14); -lean_ctor_set(x_36, 1, x_35); -x_37 = l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__49; -lean_inc(x_16); -lean_inc(x_21); -x_38 = l_Lean_addMacroScope(x_21, x_37, x_16); -x_39 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__56; -x_40 = lean_name_mk_string(x_2, x_39); -x_41 = l_Lean_Elab_Term_Quotation_mkTuple___closed__6; -x_42 = lean_name_mk_string(x_40, x_41); -x_43 = lean_box(0); +x_33 = lean_name_mk_string(x_1, x_32); +x_34 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__19; +x_35 = lean_name_mk_string(x_1, x_34); +x_36 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__21; +lean_inc(x_15); +x_37 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_37, 0, x_15); +lean_ctor_set(x_37, 1, x_36); +x_38 = l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__49; +lean_inc(x_17); +lean_inc(x_22); +x_39 = l_Lean_addMacroScope(x_22, x_38, x_17); +x_40 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__56; +x_41 = lean_name_mk_string(x_2, x_40); +x_42 = l_Lean_Elab_Term_Quotation_mkTuple___closed__6; +x_43 = lean_name_mk_string(x_41, x_42); +lean_inc(x_3); x_44 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_44, 1, x_43); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_3); +lean_inc(x_3); x_45 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_45, 0, x_44); -lean_ctor_set(x_45, 1, x_43); +lean_ctor_set(x_45, 1, x_3); x_46 = l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__48; -lean_inc(x_14); +lean_inc(x_15); x_47 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_47, 0, x_14); +lean_ctor_set(x_47, 0, x_15); lean_ctor_set(x_47, 1, x_46); -lean_ctor_set(x_47, 2, x_38); +lean_ctor_set(x_47, 2, x_39); lean_ctor_set(x_47, 3, x_45); x_48 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_49 = lean_array_push(x_48, x_36); +x_49 = lean_array_push(x_48, x_37); x_50 = lean_array_push(x_49, x_47); x_51 = lean_box(2); x_52 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_52, 0, x_51); -lean_ctor_set(x_52, 1, x_34); +lean_ctor_set(x_52, 1, x_35); lean_ctor_set(x_52, 2, x_50); -x_53 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8(x_3); +x_53 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24(x_4); x_54 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -x_55 = l_Lean_addMacroScope(x_21, x_54, x_16); +x_55 = l_Lean_addMacroScope(x_22, x_54, x_17); x_56 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; -lean_inc(x_14); +lean_inc(x_15); x_57 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_57, 0, x_14); +lean_ctor_set(x_57, 0, x_15); lean_ctor_set(x_57, 1, x_56); lean_ctor_set(x_57, 2, x_55); -lean_ctor_set(x_57, 3, x_43); +lean_ctor_set(x_57, 3, x_3); x_58 = lean_array_push(x_48, x_53); x_59 = lean_array_push(x_58, x_57); x_60 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; @@ -35289,49 +36941,49 @@ x_62 = lean_array_push(x_48, x_52); x_63 = lean_array_push(x_62, x_61); x_64 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_64, 0, x_51); -lean_ctor_set(x_64, 1, x_32); +lean_ctor_set(x_64, 1, x_33); lean_ctor_set(x_64, 2, x_63); x_65 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; -x_66 = lean_array_push(x_65, x_4); +x_66 = lean_array_push(x_65, x_5); x_67 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; x_68 = lean_array_push(x_66, x_67); x_69 = lean_array_push(x_68, x_67); -x_70 = lean_array_push(x_69, x_30); +x_70 = lean_array_push(x_69, x_31); x_71 = lean_array_push(x_70, x_64); x_72 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_72, 0, x_51); -lean_ctor_set(x_72, 1, x_28); +lean_ctor_set(x_72, 1, x_29); lean_ctor_set(x_72, 2, x_71); x_73 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; x_74 = lean_array_push(x_73, x_72); x_75 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_75, 0, x_51); -lean_ctor_set(x_75, 1, x_26); +lean_ctor_set(x_75, 1, x_27); lean_ctor_set(x_75, 2, x_74); x_76 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__20; x_77 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_77, 0, x_14); +lean_ctor_set(x_77, 0, x_15); lean_ctor_set(x_77, 1, x_76); x_78 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; -x_79 = lean_array_push(x_78, x_24); +x_79 = lean_array_push(x_78, x_25); x_80 = lean_array_push(x_79, x_75); x_81 = lean_array_push(x_80, x_77); -x_82 = lean_array_push(x_81, x_5); +x_82 = lean_array_push(x_81, x_6); x_83 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_83, 0, x_51); -lean_ctor_set(x_83, 1, x_23); +lean_ctor_set(x_83, 1, x_24); lean_ctor_set(x_83, 2, x_82); -lean_ctor_set(x_17, 0, x_83); -return x_17; +lean_ctor_set(x_18, 0, x_83); +return x_18; } else { -lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; -x_84 = lean_ctor_get(x_17, 0); -x_85 = lean_ctor_get(x_17, 1); +lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; +x_84 = lean_ctor_get(x_18, 0); +x_85 = lean_ctor_get(x_18, 1); lean_inc(x_85); lean_inc(x_84); -lean_dec(x_17); +lean_dec(x_18); x_86 = lean_ctor_get(x_84, 0); lean_inc(x_86); lean_dec(x_84); @@ -35339,9 +36991,9 @@ x_87 = lean_environment_main_module(x_86); x_88 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__7; lean_inc(x_1); x_89 = lean_name_mk_string(x_1, x_88); -lean_inc(x_14); +lean_inc(x_15); x_90 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_90, 0, x_14); +lean_ctor_set(x_90, 0, x_15); lean_ctor_set(x_90, 1, x_88); x_91 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__9; lean_inc(x_1); @@ -35350,161 +37002,161 @@ x_93 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiq lean_inc(x_1); x_94 = lean_name_mk_string(x_1, x_93); x_95 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__17; -lean_inc(x_14); +lean_inc(x_15); x_96 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_96, 0, x_14); +lean_ctor_set(x_96, 0, x_15); lean_ctor_set(x_96, 1, x_95); x_97 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__3; lean_inc(x_1); x_98 = lean_name_mk_string(x_1, x_97); -x_99 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__19; +x_99 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__19; x_100 = lean_name_mk_string(x_1, x_99); -x_101 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__21; -lean_inc(x_14); +x_101 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__21; +lean_inc(x_15); x_102 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_102, 0, x_14); +lean_ctor_set(x_102, 0, x_15); lean_ctor_set(x_102, 1, x_101); x_103 = l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__49; -lean_inc(x_16); +lean_inc(x_17); lean_inc(x_87); -x_104 = l_Lean_addMacroScope(x_87, x_103, x_16); -x_105 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__56; +x_104 = l_Lean_addMacroScope(x_87, x_103, x_17); +x_105 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__56; x_106 = lean_name_mk_string(x_2, x_105); x_107 = l_Lean_Elab_Term_Quotation_mkTuple___closed__6; x_108 = lean_name_mk_string(x_106, x_107); -x_109 = lean_box(0); -x_110 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_110, 0, x_108); -lean_ctor_set(x_110, 1, x_109); -x_111 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_111, 0, x_110); -lean_ctor_set(x_111, 1, x_109); -x_112 = l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__48; -lean_inc(x_14); -x_113 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_113, 0, x_14); -lean_ctor_set(x_113, 1, x_112); -lean_ctor_set(x_113, 2, x_104); -lean_ctor_set(x_113, 3, x_111); -x_114 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; -x_115 = lean_array_push(x_114, x_102); -x_116 = lean_array_push(x_115, x_113); -x_117 = lean_box(2); -x_118 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_118, 0, x_117); -lean_ctor_set(x_118, 1, x_100); -lean_ctor_set(x_118, 2, x_116); -x_119 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8(x_3); -x_120 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -x_121 = l_Lean_addMacroScope(x_87, x_120, x_16); -x_122 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; -lean_inc(x_14); -x_123 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_123, 0, x_14); -lean_ctor_set(x_123, 1, x_122); -lean_ctor_set(x_123, 2, x_121); -lean_ctor_set(x_123, 3, x_109); -x_124 = lean_array_push(x_114, x_119); -x_125 = lean_array_push(x_124, x_123); -x_126 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -x_127 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_127, 0, x_117); -lean_ctor_set(x_127, 1, x_126); -lean_ctor_set(x_127, 2, x_125); -x_128 = lean_array_push(x_114, x_118); -x_129 = lean_array_push(x_128, x_127); -x_130 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_130, 0, x_117); -lean_ctor_set(x_130, 1, x_98); -lean_ctor_set(x_130, 2, x_129); -x_131 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; -x_132 = lean_array_push(x_131, x_4); -x_133 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; -x_134 = lean_array_push(x_132, x_133); -x_135 = lean_array_push(x_134, x_133); -x_136 = lean_array_push(x_135, x_96); -x_137 = lean_array_push(x_136, x_130); -x_138 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_138, 0, x_117); -lean_ctor_set(x_138, 1, x_94); -lean_ctor_set(x_138, 2, x_137); -x_139 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; -x_140 = lean_array_push(x_139, x_138); -x_141 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_141, 0, x_117); -lean_ctor_set(x_141, 1, x_92); -lean_ctor_set(x_141, 2, x_140); -x_142 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__20; -x_143 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_143, 0, x_14); -lean_ctor_set(x_143, 1, x_142); -x_144 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; -x_145 = lean_array_push(x_144, x_90); -x_146 = lean_array_push(x_145, x_141); -x_147 = lean_array_push(x_146, x_143); -x_148 = lean_array_push(x_147, x_5); -x_149 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_149, 0, x_117); -lean_ctor_set(x_149, 1, x_89); -lean_ctor_set(x_149, 2, x_148); -x_150 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_150, 0, x_149); -lean_ctor_set(x_150, 1, x_85); -return x_150; +lean_inc(x_3); +x_109 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_109, 0, x_108); +lean_ctor_set(x_109, 1, x_3); +lean_inc(x_3); +x_110 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_110, 0, x_109); +lean_ctor_set(x_110, 1, x_3); +x_111 = l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__48; +lean_inc(x_15); +x_112 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_112, 0, x_15); +lean_ctor_set(x_112, 1, x_111); +lean_ctor_set(x_112, 2, x_104); +lean_ctor_set(x_112, 3, x_110); +x_113 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___lambda__1___closed__12; +x_114 = lean_array_push(x_113, x_102); +x_115 = lean_array_push(x_114, x_112); +x_116 = lean_box(2); +x_117 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_117, 0, x_116); +lean_ctor_set(x_117, 1, x_100); +lean_ctor_set(x_117, 2, x_115); +x_118 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24(x_4); +x_119 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +x_120 = l_Lean_addMacroScope(x_87, x_119, x_17); +x_121 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; +lean_inc(x_15); +x_122 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_122, 0, x_15); +lean_ctor_set(x_122, 1, x_121); +lean_ctor_set(x_122, 2, x_120); +lean_ctor_set(x_122, 3, x_3); +x_123 = lean_array_push(x_113, x_118); +x_124 = lean_array_push(x_123, x_122); +x_125 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_126 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_126, 0, x_116); +lean_ctor_set(x_126, 1, x_125); +lean_ctor_set(x_126, 2, x_124); +x_127 = lean_array_push(x_113, x_117); +x_128 = lean_array_push(x_127, x_126); +x_129 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_129, 0, x_116); +lean_ctor_set(x_129, 1, x_98); +lean_ctor_set(x_129, 2, x_128); +x_130 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; +x_131 = lean_array_push(x_130, x_5); +x_132 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; +x_133 = lean_array_push(x_131, x_132); +x_134 = lean_array_push(x_133, x_132); +x_135 = lean_array_push(x_134, x_96); +x_136 = lean_array_push(x_135, x_129); +x_137 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_137, 0, x_116); +lean_ctor_set(x_137, 1, x_94); +lean_ctor_set(x_137, 2, x_136); +x_138 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; +x_139 = lean_array_push(x_138, x_137); +x_140 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_140, 0, x_116); +lean_ctor_set(x_140, 1, x_92); +lean_ctor_set(x_140, 2, x_139); +x_141 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__20; +x_142 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_142, 0, x_15); +lean_ctor_set(x_142, 1, x_141); +x_143 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; +x_144 = lean_array_push(x_143, x_90); +x_145 = lean_array_push(x_144, x_140); +x_146 = lean_array_push(x_145, x_142); +x_147 = lean_array_push(x_146, x_6); +x_148 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_148, 0, x_116); +lean_ctor_set(x_148, 1, x_89); +lean_ctor_set(x_148, 2, x_147); +x_149 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_149, 0, x_148); +lean_ctor_set(x_149, 1, x_85); +return x_149; } } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__30(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_object* x_9) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__29(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_object* x_9, lean_object* x_10) { _start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -lean_inc(x_7); -x_10 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_7, x_8, x_9); -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_10, 1); +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +lean_inc(x_8); +x_11 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___spec__3___rarg(x_8, x_9, x_10); +x_12 = lean_ctor_get(x_11, 0); lean_inc(x_12); -lean_dec(x_10); -x_13 = lean_ctor_get(x_7, 10); +x_13 = lean_ctor_get(x_11, 1); lean_inc(x_13); -lean_dec(x_7); -x_14 = lean_st_ref_get(x_8, x_12); -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) +lean_dec(x_11); +x_14 = lean_ctor_get(x_8, 10); +lean_inc(x_14); +lean_dec(x_8); +x_15 = lean_st_ref_get(x_9, x_13); +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) { -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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_16 = lean_ctor_get(x_14, 0); -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -x_18 = lean_environment_main_module(x_17); -x_19 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__7; -lean_inc(x_11); -x_20 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_20, 0, x_11); -lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_Syntax_getAntiquotTerm(x_1); -x_22 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__17; -lean_inc(x_11); -x_23 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_23, 0, x_11); -lean_ctor_set(x_23, 1, x_22); -x_24 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -x_25 = l_Lean_addMacroScope(x_18, x_24, x_13); -x_26 = lean_box(0); +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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_17 = lean_ctor_get(x_15, 0); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +lean_dec(x_17); +x_19 = lean_environment_main_module(x_18); +x_20 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__7; +lean_inc(x_12); +x_21 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_21, 0, x_12); +lean_ctor_set(x_21, 1, x_20); +x_22 = l_Lean_Syntax_getAntiquotTerm(x_1); +x_23 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__17; +lean_inc(x_12); +x_24 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_24, 0, x_12); +lean_ctor_set(x_24, 1, x_23); +x_25 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +x_26 = l_Lean_addMacroScope(x_19, x_25, x_14); x_27 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; -lean_inc(x_11); +lean_inc(x_12); x_28 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_28, 0, x_11); +lean_ctor_set(x_28, 0, x_12); lean_ctor_set(x_28, 1, x_27); -lean_ctor_set(x_28, 2, x_25); -lean_ctor_set(x_28, 3, x_26); +lean_ctor_set(x_28, 2, x_26); +lean_ctor_set(x_28, 3, x_2); x_29 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; -x_30 = lean_array_push(x_29, x_21); +x_30 = lean_array_push(x_29, x_22); x_31 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; x_32 = lean_array_push(x_30, x_31); x_33 = lean_array_push(x_32, x_31); -x_34 = lean_array_push(x_33, x_23); +x_34 = lean_array_push(x_33, x_24); x_35 = lean_array_push(x_34, x_28); x_36 = lean_box(2); x_37 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__12; @@ -35521,92 +37173,91 @@ lean_ctor_set(x_42, 1, x_41); lean_ctor_set(x_42, 2, x_40); x_43 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__20; x_44 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_44, 0, x_11); +lean_ctor_set(x_44, 0, x_12); lean_ctor_set(x_44, 1, x_43); x_45 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; -x_46 = lean_array_push(x_45, x_20); +x_46 = lean_array_push(x_45, x_21); x_47 = lean_array_push(x_46, x_42); x_48 = lean_array_push(x_47, x_44); -x_49 = lean_array_push(x_48, x_2); +x_49 = lean_array_push(x_48, x_3); x_50 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__8; x_51 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_51, 0, x_36); lean_ctor_set(x_51, 1, x_50); lean_ctor_set(x_51, 2, x_49); -lean_ctor_set(x_14, 0, x_51); -return x_14; +lean_ctor_set(x_15, 0, x_51); +return x_15; } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_52 = lean_ctor_get(x_14, 0); -x_53 = lean_ctor_get(x_14, 1); +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_52 = lean_ctor_get(x_15, 0); +x_53 = lean_ctor_get(x_15, 1); lean_inc(x_53); lean_inc(x_52); -lean_dec(x_14); +lean_dec(x_15); x_54 = lean_ctor_get(x_52, 0); lean_inc(x_54); lean_dec(x_52); x_55 = lean_environment_main_module(x_54); x_56 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__7; -lean_inc(x_11); +lean_inc(x_12); x_57 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_57, 0, x_11); +lean_ctor_set(x_57, 0, x_12); lean_ctor_set(x_57, 1, x_56); x_58 = l_Lean_Syntax_getAntiquotTerm(x_1); x_59 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__17; -lean_inc(x_11); +lean_inc(x_12); x_60 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_60, 0, x_11); +lean_ctor_set(x_60, 0, x_12); lean_ctor_set(x_60, 1, x_59); x_61 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -x_62 = l_Lean_addMacroScope(x_55, x_61, x_13); -x_63 = lean_box(0); -x_64 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; -lean_inc(x_11); -x_65 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_65, 0, x_11); -lean_ctor_set(x_65, 1, x_64); -lean_ctor_set(x_65, 2, x_62); -lean_ctor_set(x_65, 3, x_63); -x_66 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; -x_67 = lean_array_push(x_66, x_58); -x_68 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; -x_69 = lean_array_push(x_67, x_68); -x_70 = lean_array_push(x_69, x_68); -x_71 = lean_array_push(x_70, x_60); -x_72 = lean_array_push(x_71, x_65); -x_73 = lean_box(2); -x_74 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__12; -x_75 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_75, 0, x_73); -lean_ctor_set(x_75, 1, x_74); -lean_ctor_set(x_75, 2, x_72); -x_76 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; -x_77 = lean_array_push(x_76, x_75); -x_78 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__10; -x_79 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_79, 0, x_73); -lean_ctor_set(x_79, 1, x_78); -lean_ctor_set(x_79, 2, x_77); -x_80 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__20; -x_81 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_81, 0, x_11); -lean_ctor_set(x_81, 1, x_80); -x_82 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; -x_83 = lean_array_push(x_82, x_57); -x_84 = lean_array_push(x_83, x_79); -x_85 = lean_array_push(x_84, x_81); -x_86 = lean_array_push(x_85, x_2); -x_87 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__8; -x_88 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_88, 0, x_73); -lean_ctor_set(x_88, 1, x_87); -lean_ctor_set(x_88, 2, x_86); -x_89 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_89, 0, x_88); -lean_ctor_set(x_89, 1, x_53); -return x_89; +x_62 = l_Lean_addMacroScope(x_55, x_61, x_14); +x_63 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; +lean_inc(x_12); +x_64 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_64, 0, x_12); +lean_ctor_set(x_64, 1, x_63); +lean_ctor_set(x_64, 2, x_62); +lean_ctor_set(x_64, 3, x_2); +x_65 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; +x_66 = lean_array_push(x_65, x_58); +x_67 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; +x_68 = lean_array_push(x_66, x_67); +x_69 = lean_array_push(x_68, x_67); +x_70 = lean_array_push(x_69, x_60); +x_71 = lean_array_push(x_70, x_64); +x_72 = lean_box(2); +x_73 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__12; +x_74 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_74, 0, x_72); +lean_ctor_set(x_74, 1, x_73); +lean_ctor_set(x_74, 2, x_71); +x_75 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; +x_76 = lean_array_push(x_75, x_74); +x_77 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__10; +x_78 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_78, 0, x_72); +lean_ctor_set(x_78, 1, x_77); +lean_ctor_set(x_78, 2, x_76); +x_79 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__20; +x_80 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_80, 0, x_12); +lean_ctor_set(x_80, 1, x_79); +x_81 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__21; +x_82 = lean_array_push(x_81, x_57); +x_83 = lean_array_push(x_82, x_78); +x_84 = lean_array_push(x_83, x_80); +x_85 = lean_array_push(x_84, x_3); +x_86 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__8; +x_87 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_87, 0, x_72); +lean_ctor_set(x_87, 1, x_86); +lean_ctor_set(x_87, 2, x_85); +x_88 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_88, 0, x_87); +lean_ctor_set(x_88, 1, x_53); +return x_88; } } } @@ -35615,12 +37266,24 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__19; +x_3 = lean_alloc_ctor(1, 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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__2() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__3() { _start: { lean_object* x_1; @@ -35628,17 +37291,17 @@ x_1 = lean_mk_string_from_bytes("namedPattern", 12); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__3() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__2; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__4() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__5() { _start: { lean_object* x_1; @@ -35646,23 +37309,15 @@ x_1 = lean_mk_string_from_bytes("match (syntax) : unexpected pattern kind ", 41) return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__5() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__4; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__5; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__6() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__5___boxed), 9, 0); -return x_1; -} -} static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__7() { _start: { @@ -35674,20 +37329,6 @@ return x_1; static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__8() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__7; -x_3 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__6; -x_4 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_4, 0, x_1); -lean_ctor_set(x_4, 1, x_2); -lean_ctor_set(x_4, 2, x_3); -return x_4; -} -} -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__9() { -_start: -{ lean_object* x_1; x_1 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__7___boxed), 1, 0); return x_1; @@ -35696,100 +37337,100 @@ return x_1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo(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: { -uint8_t x_9; -x_9 = !lean_is_exclusive(x_1); -if (x_9 == 0) +lean_object* x_9; uint8_t x_10; +x_9 = lean_box(0); +x_10 = !lean_is_exclusive(x_1); +if (x_10 == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_10 = lean_ctor_get(x_1, 0); -x_11 = lean_ctor_get(x_1, 1); -x_12 = l_Lean_instInhabitedSyntax; -x_13 = l_List_head_x21___rarg(x_12, x_10); -x_14 = l_Lean_Syntax_isQuot(x_13); -if (x_14 == 0) +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_11 = lean_ctor_get(x_1, 0); +x_12 = lean_ctor_get(x_1, 1); +x_13 = lean_box(0); +x_14 = l_List_head_x21___rarg(x_13, x_11); +x_15 = l_Lean_Syntax_isQuot(x_14); +if (x_15 == 0) { -lean_object* x_15; uint8_t x_16; -x_15 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_noOpMatchAdaptPats___lambda__1___closed__1; -lean_inc(x_13); -x_16 = l_Lean_Syntax_isOfKind(x_13, x_15); -if (x_16 == 0) +lean_object* x_16; uint8_t x_17; +x_16 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_noOpMatchAdaptPats___lambda__1___closed__1; +lean_inc(x_14); +x_17 = l_Lean_Syntax_isOfKind(x_14, x_16); +if (x_17 == 0) { -lean_object* x_17; uint8_t x_18; -x_17 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__1; -lean_inc(x_13); -x_18 = l_Lean_Syntax_isOfKind(x_13, x_17); -if (x_18 == 0) +lean_object* x_18; uint8_t x_19; +x_18 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__2; +lean_inc(x_14); +x_19 = l_Lean_Syntax_isOfKind(x_14, x_18); +if (x_19 == 0) { -lean_object* x_19; uint8_t x_20; -x_19 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__3; -lean_inc(x_13); -x_20 = l_Lean_Syntax_isOfKind(x_13, x_19); -if (x_20 == 0) +lean_object* x_20; uint8_t x_21; +x_20 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__4; +lean_inc(x_14); +x_21 = l_Lean_Syntax_isOfKind(x_14, x_20); +if (x_21 == 0) { -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_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_free_object(x_1); +lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); -lean_inc(x_13); -x_21 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_21, 0, x_13); -x_22 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__5; -x_23 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_21); -x_24 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__3; -x_25 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_25, 0, x_23); -lean_ctor_set(x_25, 1, x_24); -x_26 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__1(x_13, x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_26; +lean_inc(x_14); +x_22 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_22, 0, x_14); +x_23 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__6; +x_24 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_22); +x_25 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__3; +x_26 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +x_27 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__1(x_14, x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_27; } else { -lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_27 = lean_unsigned_to_nat(0u); -x_28 = l_Lean_Syntax_getArg(x_13, x_27); -lean_inc(x_28); -x_29 = l_Lean_Syntax_isOfKind(x_28, x_17); -if (x_29 == 0) +lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_28 = lean_unsigned_to_nat(0u); +x_29 = l_Lean_Syntax_getArg(x_14, x_28); +lean_inc(x_29); +x_30 = l_Lean_Syntax_isOfKind(x_29, x_18); +if (x_30 == 0) { -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_dec(x_28); +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_dec(x_29); lean_free_object(x_1); +lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); -lean_inc(x_13); -x_30 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_30, 0, x_13); -x_31 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__5; -x_32 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_30); -x_33 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__3; -x_34 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_34, 0, x_32); -lean_ctor_set(x_34, 1, x_33); -x_35 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__1(x_13, x_34, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_35; +lean_inc(x_14); +x_31 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_31, 0, x_14); +x_32 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__6; +x_33 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_31); +x_34 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__3; +x_35 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +x_36 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__1(x_14, x_35, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_36; } else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; -x_36 = lean_unsigned_to_nat(2u); -x_37 = l_Lean_Syntax_getArg(x_13, x_36); -x_38 = l_Lean_nullKind; -x_39 = l_Lean_Syntax_isNodeOf(x_37, x_38, x_27); +lean_object* x_37; lean_object* x_38; uint8_t x_39; +x_37 = lean_unsigned_to_nat(2u); +x_38 = l_Lean_Syntax_getArg(x_14, x_37); +x_39 = l_Lean_Syntax_matchesNull(x_38, x_28); if (x_39 == 0) { lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -lean_dec(x_28); +lean_dec(x_29); lean_free_object(x_1); +lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); -lean_inc(x_13); +lean_inc(x_14); x_40 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_40, 0, x_13); -x_41 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__5; +lean_ctor_set(x_40, 0, x_14); +x_41 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__6; x_42 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_42, 0, x_41); lean_ctor_set(x_42, 1, x_40); @@ -35797,16 +37438,16 @@ x_43 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__3; x_44 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_44, 0, x_42); lean_ctor_set(x_44, 1, x_43); -x_45 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__1(x_13, x_44, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_45 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__1(x_14, x_44, x_2, x_3, x_4, x_5, x_6, x_7, x_8); return x_45; } else { lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; x_46 = lean_unsigned_to_nat(3u); -x_47 = l_Lean_Syntax_getArg(x_13, x_46); -lean_dec(x_13); -x_48 = l_List_tail_x21___rarg(x_10); +x_47 = l_Lean_Syntax_getArg(x_14, x_46); +lean_dec(x_14); +x_48 = l_List_tail_x21___rarg(x_11); x_49 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_49, 0, x_47); lean_ctor_set(x_49, 1, x_48); @@ -35826,10 +37467,11 @@ x_54 = lean_ctor_get(x_52, 2); lean_inc(x_54); x_55 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; lean_inc(x_52); -x_56 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__3), 4, 3); +x_56 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__3), 5, 4); lean_closure_set(x_56, 0, x_52); lean_closure_set(x_56, 1, x_55); -lean_closure_set(x_56, 2, x_28); +lean_closure_set(x_56, 2, x_9); +lean_closure_set(x_56, 3, x_29); x_57 = !lean_is_exclusive(x_52); if (x_57 == 0) { @@ -35869,10 +37511,11 @@ x_65 = lean_ctor_get(x_62, 2); lean_inc(x_65); x_66 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; lean_inc(x_62); -x_67 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__3), 4, 3); +x_67 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__3), 5, 4); lean_closure_set(x_67, 0, x_62); lean_closure_set(x_67, 1, x_66); -lean_closure_set(x_67, 2, x_28); +lean_closure_set(x_67, 2, x_9); +lean_closure_set(x_67, 3, x_29); if (lean_is_exclusive(x_62)) { lean_ctor_release(x_62, 0); lean_ctor_release(x_62, 1); @@ -35899,7 +37542,7 @@ return x_70; else { uint8_t x_71; -lean_dec(x_28); +lean_dec(x_29); x_71 = !lean_is_exclusive(x_50); if (x_71 == 0) { @@ -35927,8 +37570,8 @@ else { lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_free_object(x_1); +lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); @@ -35936,17 +37579,19 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); x_75 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_76 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__1___boxed), 10, 2); +x_76 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__1___boxed), 11, 3); lean_closure_set(x_76, 0, x_75); -lean_closure_set(x_76, 1, x_13); +lean_closure_set(x_76, 1, x_9); +lean_closure_set(x_76, 2, x_14); x_77 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__4), 2, 1); lean_closure_set(x_77, 0, x_76); -x_78 = lean_box(0); -x_79 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__6; +x_78 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__5___boxed), 10, 1); +lean_closure_set(x_78, 0, x_9); +x_79 = lean_box(0); x_80 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_80, 0, x_78); +lean_ctor_set(x_80, 0, x_79); lean_ctor_set(x_80, 1, x_77); -lean_ctor_set(x_80, 2, x_79); +lean_ctor_set(x_80, 2, x_78); x_81 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_81, 0, x_80); lean_ctor_set(x_81, 1, x_8); @@ -35955,537 +37600,568 @@ return x_81; } else { -lean_object* x_82; lean_object* x_83; -lean_dec(x_13); +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +lean_dec(x_14); lean_free_object(x_1); +lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_82 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__8; -x_83 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_83, 0, x_82); -lean_ctor_set(x_83, 1, x_8); -return x_83; +x_82 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__5___boxed), 10, 1); +lean_closure_set(x_82, 0, x_9); +x_83 = lean_box(0); +x_84 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__7; +x_85 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_85, 0, x_83); +lean_ctor_set(x_85, 1, x_84); +lean_ctor_set(x_85, 2, x_82); +x_86 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_86, 0, x_85); +lean_ctor_set(x_86, 1, x_8); +return x_86; } } else { -lean_object* x_84; lean_object* x_85; uint8_t x_156; +lean_object* x_87; lean_object* x_88; uint8_t x_160; lean_free_object(x_1); +lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); -lean_inc(x_13); -x_84 = l_Lean_Syntax_getQuotContent(x_13); -x_156 = l_Lean_Syntax_isAtom(x_84); -if (x_156 == 0) +lean_inc(x_14); +x_87 = l_Lean_Syntax_getQuotContent(x_14); +x_160 = l_Lean_Syntax_isAtom(x_87); +if (x_160 == 0) { -uint8_t x_157; -lean_inc(x_84); -x_157 = l_Lean_Syntax_isTokenAntiquot(x_84); -if (x_157 == 0) +uint8_t x_161; +lean_inc(x_87); +x_161 = l_Lean_Syntax_isTokenAntiquot(x_87); +if (x_161 == 0) { -uint8_t x_158; -lean_inc(x_84); -x_158 = l_Lean_Syntax_isAntiquots(x_84); -if (x_158 == 0) +uint8_t x_162; +lean_inc(x_87); +x_162 = l_Lean_Syntax_isAntiquots(x_87); +if (x_162 == 0) { -uint8_t x_159; -x_159 = l_Lean_Syntax_isAntiquotSuffixSplice(x_84); -if (x_159 == 0) +uint8_t x_163; +x_163 = l_Lean_Syntax_isAntiquotSuffixSplice(x_87); +if (x_163 == 0) { -uint8_t x_160; -x_160 = l_Lean_Syntax_isAntiquotSplice(x_84); -if (x_160 == 0) +uint8_t x_164; +x_164 = l_Lean_Syntax_isAntiquotSplice(x_87); +if (x_164 == 0) { -lean_object* x_161; +lean_object* x_165; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_161 = lean_box(0); -x_85 = x_161; -goto block_155; +x_165 = lean_box(0); +x_88 = x_165; +goto block_159; } else { -lean_object* x_162; lean_object* x_163; -lean_dec(x_13); -x_162 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__48; -x_163 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__1(x_84, x_162, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_163; +lean_object* x_166; lean_object* x_167; +lean_dec(x_14); +x_166 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__48; +x_167 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__1(x_87, x_166, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_167; } } else { -lean_object* x_164; lean_object* x_165; -lean_dec(x_13); -x_164 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__48; -x_165 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__1(x_84, x_164, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_165; +lean_object* x_168; lean_object* x_169; +lean_dec(x_14); +x_168 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__48; +x_169 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__1(x_87, x_168, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_169; } } else { -lean_object* x_166; uint8_t x_167; -lean_inc(x_84); -x_166 = l_Lean_Syntax_getCanonicalAntiquot(x_84); -x_167 = l_Lean_Syntax_isEscapedAntiquot(x_166); -if (x_167 == 0) +lean_object* x_170; uint8_t x_171; +lean_inc(x_87); +x_170 = l_Lean_Syntax_getCanonicalAntiquot(x_87); +x_171 = l_Lean_Syntax_isEscapedAntiquot(x_170); +if (x_171 == 0) { -lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; uint8_t x_174; lean_object* x_175; lean_object* x_176; uint8_t x_177; lean_object* x_178; -lean_dec(x_13); +lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; uint8_t x_178; lean_object* x_179; lean_object* x_180; uint8_t x_181; lean_object* x_182; +lean_dec(x_14); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_168 = l_Lean_Syntax_antiquotKinds(x_84); -x_169 = l_List_unzip___rarg(x_168); -x_170 = lean_ctor_get(x_169, 0); -lean_inc(x_170); -x_171 = lean_ctor_get(x_169, 1); -lean_inc(x_171); -lean_dec(x_169); -x_172 = l_Lean_Syntax_getAntiquotTerm(x_166); -x_173 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_noOpMatchAdaptPats___lambda__1___closed__1; -lean_inc(x_172); -x_174 = l_Lean_Syntax_isOfKind(x_172, x_173); -x_175 = lean_unsigned_to_nat(3u); -x_176 = l_Lean_Syntax_getArg(x_166, x_175); -lean_dec(x_166); -x_177 = l_Lean_Syntax_isNone(x_176); -lean_dec(x_176); -if (x_174 == 0) +x_172 = l_Lean_Syntax_antiquotKinds(x_87); +x_173 = l_List_unzip___rarg(x_172); +x_174 = lean_ctor_get(x_173, 0); +lean_inc(x_174); +x_175 = lean_ctor_get(x_173, 1); +lean_inc(x_175); +lean_dec(x_173); +x_176 = l_Lean_Syntax_getAntiquotTerm(x_170); +x_177 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_noOpMatchAdaptPats___lambda__1___closed__1; +lean_inc(x_176); +x_178 = l_Lean_Syntax_isOfKind(x_176, x_177); +x_179 = lean_unsigned_to_nat(3u); +x_180 = l_Lean_Syntax_getArg(x_170, x_179); +lean_dec(x_170); +x_181 = l_Lean_Syntax_isNone(x_180); +lean_dec(x_180); +if (x_178 == 0) { -lean_object* x_200; uint8_t x_201; -x_200 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__1; -lean_inc(x_172); -x_201 = l_Lean_Syntax_isOfKind(x_172, x_200); -if (x_201 == 0) +lean_object* x_205; uint8_t x_206; +x_205 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__2; +lean_inc(x_176); +x_206 = l_Lean_Syntax_isOfKind(x_176, x_205); +if (x_206 == 0) { -lean_object* x_202; -x_202 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__28___boxed), 9, 1); -lean_closure_set(x_202, 0, x_172); -x_178 = x_202; -goto block_199; +lean_object* x_207; +x_207 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___boxed), 9, 1); +lean_closure_set(x_207, 0, x_176); +x_182 = x_207; +goto block_204; } else { -lean_object* x_203; lean_object* x_204; lean_object* x_205; -x_203 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_204 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__2; -lean_inc(x_170); -x_205 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__29___boxed), 12, 4); -lean_closure_set(x_205, 0, x_203); -lean_closure_set(x_205, 1, x_204); -lean_closure_set(x_205, 2, x_170); -lean_closure_set(x_205, 3, x_172); -x_178 = x_205; -goto block_199; +lean_object* x_208; lean_object* x_209; lean_object* x_210; +x_208 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; +x_209 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__2; +lean_inc(x_174); +x_210 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__28___boxed), 13, 5); +lean_closure_set(x_210, 0, x_208); +lean_closure_set(x_210, 1, x_209); +lean_closure_set(x_210, 2, x_9); +lean_closure_set(x_210, 3, x_174); +lean_closure_set(x_210, 4, x_176); +x_182 = x_210; +goto block_204; } } else { -lean_object* x_206; -lean_dec(x_172); -x_206 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__1; -x_178 = x_206; -goto block_199; +lean_object* x_211; +lean_dec(x_176); +x_211 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__1; +x_182 = x_211; +goto block_204; } -block_199: +block_204: { -if (x_177 == 0) -{ -uint8_t x_179; uint8_t x_180; -x_179 = 1; -x_180 = l_List_foldr___at_List_and___spec__1(x_179, x_171); -lean_dec(x_171); -if (x_180 == 0) +if (x_181 == 0) { -lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; -x_181 = lean_box(0); -lean_inc(x_170); -x_182 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_182, 0, x_170); -lean_ctor_set(x_182, 1, x_181); -lean_inc(x_170); -x_183 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___boxed), 3, 2); -lean_closure_set(x_183, 0, x_170); -lean_closure_set(x_183, 1, x_178); -x_184 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__2; -x_185 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_186 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27), 12, 3); -lean_closure_set(x_186, 0, x_184); -lean_closure_set(x_186, 1, x_185); -lean_closure_set(x_186, 2, x_170); -x_187 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_187, 0, x_182); -lean_ctor_set(x_187, 1, x_183); -lean_ctor_set(x_187, 2, x_186); -x_188 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_188, 0, x_187); -lean_ctor_set(x_188, 1, x_8); -return x_188; -} -else +uint8_t x_183; uint8_t x_184; +x_183 = 1; +x_184 = l_List_foldr___at_List_and___spec__1(x_183, x_175); +lean_dec(x_175); +if (x_184 == 0) { -lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; -lean_dec(x_170); -x_189 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__4), 2, 1); -lean_closure_set(x_189, 0, x_178); -x_190 = lean_box(0); -x_191 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__6; +lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; +x_185 = lean_box(0); +lean_inc(x_174); +x_186 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_186, 0, x_174); +lean_ctor_set(x_186, 1, x_185); +lean_inc(x_174); +x_187 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___boxed), 3, 2); +lean_closure_set(x_187, 0, x_174); +lean_closure_set(x_187, 1, x_182); +x_188 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__1; +x_189 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__2; +x_190 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; +x_191 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___boxed), 14, 5); +lean_closure_set(x_191, 0, x_9); +lean_closure_set(x_191, 1, x_188); +lean_closure_set(x_191, 2, x_189); +lean_closure_set(x_191, 3, x_190); +lean_closure_set(x_191, 4, x_174); x_192 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_192, 0, x_190); -lean_ctor_set(x_192, 1, x_189); +lean_ctor_set(x_192, 0, x_186); +lean_ctor_set(x_192, 1, x_187); lean_ctor_set(x_192, 2, x_191); x_193 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_193, 0, x_192); lean_ctor_set(x_193, 1, x_8); return x_193; } -} else { lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; -lean_dec(x_171); -lean_dec(x_170); +lean_dec(x_174); x_194 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__4), 2, 1); -lean_closure_set(x_194, 0, x_178); -x_195 = lean_box(0); -x_196 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__6; +lean_closure_set(x_194, 0, x_182); +x_195 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__5___boxed), 10, 1); +lean_closure_set(x_195, 0, x_9); +x_196 = lean_box(0); x_197 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_197, 0, x_195); +lean_ctor_set(x_197, 0, x_196); lean_ctor_set(x_197, 1, x_194); -lean_ctor_set(x_197, 2, x_196); +lean_ctor_set(x_197, 2, x_195); x_198 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_198, 0, x_197); lean_ctor_set(x_198, 1, x_8); return x_198; } } +else +{ +lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; +lean_dec(x_175); +lean_dec(x_174); +x_199 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__4), 2, 1); +lean_closure_set(x_199, 0, x_182); +x_200 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__5___boxed), 10, 1); +lean_closure_set(x_200, 0, x_9); +x_201 = lean_box(0); +x_202 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_202, 0, x_201); +lean_ctor_set(x_202, 1, x_199); +lean_ctor_set(x_202, 2, x_200); +x_203 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_203, 0, x_202); +lean_ctor_set(x_203, 1, x_8); +return x_203; +} +} } else { -uint8_t x_207; -lean_dec(x_166); -x_207 = l_Lean_Syntax_isAntiquotSuffixSplice(x_84); -if (x_207 == 0) +uint8_t x_212; +lean_dec(x_170); +x_212 = l_Lean_Syntax_isAntiquotSuffixSplice(x_87); +if (x_212 == 0) { -uint8_t x_208; -x_208 = l_Lean_Syntax_isAntiquotSplice(x_84); -if (x_208 == 0) +uint8_t x_213; +x_213 = l_Lean_Syntax_isAntiquotSplice(x_87); +if (x_213 == 0) { -lean_object* x_209; +lean_object* x_214; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_209 = lean_box(0); -x_85 = x_209; -goto block_155; +x_214 = lean_box(0); +x_88 = x_214; +goto block_159; } else { -lean_object* x_210; lean_object* x_211; -lean_dec(x_13); -x_210 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__48; -x_211 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__1(x_84, x_210, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_211; +lean_object* x_215; lean_object* x_216; +lean_dec(x_14); +x_215 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__48; +x_216 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__1(x_87, x_215, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_216; } } else { -lean_object* x_212; lean_object* x_213; -lean_dec(x_13); -x_212 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__48; -x_213 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__1(x_84, x_212, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_213; +lean_object* x_217; lean_object* x_218; +lean_dec(x_14); +x_217 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__48; +x_218 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__1(x_87, x_217, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_218; } } } } else { -lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; -lean_dec(x_13); +lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; +lean_dec(x_14); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_214 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__30___boxed), 9, 1); -lean_closure_set(x_214, 0, x_84); -x_215 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__4), 2, 1); -lean_closure_set(x_215, 0, x_214); -x_216 = lean_box(0); -x_217 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__6; -x_218 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_218, 0, x_216); -lean_ctor_set(x_218, 1, x_215); -lean_ctor_set(x_218, 2, x_217); -x_219 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_219, 0, x_218); -lean_ctor_set(x_219, 1, x_8); -return x_219; +x_219 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__29___boxed), 10, 2); +lean_closure_set(x_219, 0, x_87); +lean_closure_set(x_219, 1, x_9); +x_220 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__4), 2, 1); +lean_closure_set(x_220, 0, x_219); +x_221 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__5___boxed), 10, 1); +lean_closure_set(x_221, 0, x_9); +x_222 = lean_box(0); +x_223 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_223, 0, x_222); +lean_ctor_set(x_223, 1, x_220); +lean_ctor_set(x_223, 2, x_221); +x_224 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_224, 0, x_223); +lean_ctor_set(x_224, 1, x_8); +return x_224; } } else { -lean_object* x_220; lean_object* x_221; -lean_dec(x_84); -lean_dec(x_13); +lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; +lean_dec(x_87); +lean_dec(x_14); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_220 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__8; -x_221 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_221, 0, x_220); -lean_ctor_set(x_221, 1, x_8); -return x_221; +x_225 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__5___boxed), 10, 1); +lean_closure_set(x_225, 0, x_9); +x_226 = lean_box(0); +x_227 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__7; +x_228 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_228, 0, x_226); +lean_ctor_set(x_228, 1, x_227); +lean_ctor_set(x_228, 2, x_225); +x_229 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_229, 0, x_228); +lean_ctor_set(x_229, 1, x_8); +return x_229; } -block_155: +block_159: { -lean_object* x_86; lean_object* x_87; lean_object* x_129; lean_object* x_130; uint8_t x_131; -lean_dec(x_85); -x_86 = l_Lean_Syntax_getArgs(x_84); -x_129 = lean_array_get_size(x_86); -x_130 = lean_unsigned_to_nat(1u); -x_131 = lean_nat_dec_eq(x_129, x_130); -lean_dec(x_129); -if (x_131 == 0) +lean_object* x_89; lean_object* x_90; lean_object* x_133; lean_object* x_134; uint8_t x_135; +lean_dec(x_88); +x_89 = l_Lean_Syntax_getArgs(x_87); +x_133 = lean_array_get_size(x_89); +x_134 = lean_unsigned_to_nat(1u); +x_135 = lean_nat_dec_eq(x_133, x_134); +lean_dec(x_133); +if (x_135 == 0) { -lean_object* x_132; -lean_dec(x_13); -x_132 = lean_box(0); -x_87 = x_132; -goto block_128; +lean_object* x_136; +lean_dec(x_14); +x_136 = lean_box(0); +x_90 = x_136; +goto block_132; } else { -lean_object* x_133; lean_object* x_134; uint8_t x_135; -x_133 = lean_unsigned_to_nat(0u); -x_134 = l_Lean_Syntax_getArg(x_84, x_133); -x_135 = l_Lean_Syntax_isAntiquotSuffixSplice(x_134); -if (x_135 == 0) +lean_object* x_137; lean_object* x_138; uint8_t x_139; +x_137 = lean_unsigned_to_nat(0u); +x_138 = l_Lean_Syntax_getArg(x_87, x_137); +x_139 = l_Lean_Syntax_isAntiquotSuffixSplice(x_138); +if (x_139 == 0) { -uint8_t x_136; -x_136 = l_Lean_Syntax_isAntiquotSplice(x_134); -if (x_136 == 0) +uint8_t x_140; +x_140 = l_Lean_Syntax_isAntiquotSplice(x_138); +if (x_140 == 0) { -lean_object* x_137; -lean_dec(x_134); -lean_dec(x_13); -x_137 = lean_box(0); -x_87 = x_137; -goto block_128; +lean_object* x_141; +lean_dec(x_138); +lean_dec(x_14); +x_141 = lean_box(0); +x_90 = x_141; +goto block_132; } else { -lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; -lean_dec(x_86); -lean_dec(x_84); -lean_inc(x_13); -x_138 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_138, 0, x_13); -x_139 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__15), 2, 1); -lean_closure_set(x_139, 0, x_13); -x_140 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23), 10, 1); -lean_closure_set(x_140, 0, x_134); -x_141 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_141, 0, x_138); -lean_ctor_set(x_141, 1, x_139); -lean_ctor_set(x_141, 2, x_140); -x_142 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_142, 0, x_141); -lean_ctor_set(x_142, 1, x_8); -return x_142; +lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; +lean_dec(x_89); +lean_dec(x_87); +lean_inc(x_14); +x_142 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_142, 0, x_14); +x_143 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__15), 2, 1); +lean_closure_set(x_143, 0, x_14); +x_144 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__1; +x_145 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23), 12, 3); +lean_closure_set(x_145, 0, x_138); +lean_closure_set(x_145, 1, x_9); +lean_closure_set(x_145, 2, x_144); +x_146 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_146, 0, x_142); +lean_ctor_set(x_146, 1, x_143); +lean_ctor_set(x_146, 2, x_145); +x_147 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_147, 0, x_146); +lean_ctor_set(x_147, 1, x_8); +return x_147; } } else { -lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; -lean_dec(x_86); -lean_dec(x_13); -x_143 = l_Lean_Syntax_getArg(x_134, x_133); -lean_inc(x_143); -x_144 = l_Lean_Syntax_getCanonicalAntiquot(x_143); -x_145 = l_Lean_Syntax_getAntiquotTerm(x_144); -lean_dec(x_144); -x_146 = l_Lean_Syntax_antiquotKinds(x_143); -x_147 = lean_box(0); -x_148 = l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__7(x_146, x_147); -x_149 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___boxed), 13, 5); -lean_closure_set(x_149, 0, x_134); -lean_closure_set(x_149, 1, x_84); -lean_closure_set(x_149, 2, x_147); -lean_closure_set(x_149, 3, x_148); -lean_closure_set(x_149, 4, x_145); -x_150 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__4), 2, 1); -lean_closure_set(x_150, 0, x_149); -x_151 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___boxed), 10, 1); -lean_closure_set(x_151, 0, x_147); -x_152 = lean_box(0); -x_153 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_153, 0, x_152); -lean_ctor_set(x_153, 1, x_150); -lean_ctor_set(x_153, 2, x_151); -x_154 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_154, 0, x_153); -lean_ctor_set(x_154, 1, x_8); -return x_154; +lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; +lean_dec(x_89); +lean_dec(x_14); +x_148 = l_Lean_Syntax_getArg(x_138, x_137); +lean_inc(x_148); +x_149 = l_Lean_Syntax_getCanonicalAntiquot(x_148); +x_150 = l_Lean_Syntax_getAntiquotTerm(x_149); +lean_dec(x_149); +x_151 = l_Lean_Syntax_antiquotKinds(x_148); +x_152 = l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__21(x_151, x_9); +x_153 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___boxed), 13, 5); +lean_closure_set(x_153, 0, x_138); +lean_closure_set(x_153, 1, x_87); +lean_closure_set(x_153, 2, x_9); +lean_closure_set(x_153, 3, x_152); +lean_closure_set(x_153, 4, x_150); +x_154 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__4), 2, 1); +lean_closure_set(x_154, 0, x_153); +x_155 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__5___boxed), 10, 1); +lean_closure_set(x_155, 0, x_9); +x_156 = lean_box(0); +x_157 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_157, 0, x_156); +lean_ctor_set(x_157, 1, x_154); +lean_ctor_set(x_157, 2, x_155); +x_158 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_158, 0, x_157); +lean_ctor_set(x_158, 1, x_8); +return x_158; } } -block_128: +block_132: { -lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; -lean_dec(x_87); -x_88 = lean_array_get_size(x_86); -x_89 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__9; -x_90 = lean_unsigned_to_nat(0u); -lean_inc(x_88); -x_91 = l_Array_findIdx_x3f_loop___rarg(x_86, x_89, x_88, x_90, lean_box(0)); -if (lean_obj_tag(x_91) == 0) +lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +lean_dec(x_90); +x_91 = lean_array_get_size(x_89); +x_92 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__8; +x_93 = lean_unsigned_to_nat(0u); +lean_inc(x_91); +x_94 = l_Array_findIdx_x3f_loop___rarg(x_89, x_92, x_91, x_93, lean_box(0)); +if (lean_obj_tag(x_94) == 0) { -lean_object* x_92; lean_object* x_93; uint8_t x_94; lean_object* x_95; lean_object* x_96; size_t x_97; size_t x_98; lean_object* x_99; uint8_t x_100; lean_object* x_101; lean_object* x_102; uint8_t x_103; -lean_dec(x_88); -lean_dec(x_86); -x_92 = l_Lean_Syntax_unescapeAntiquot(x_84); -lean_inc(x_92); -x_93 = l_Lean_Syntax_getKind(x_92); -x_94 = l_Lean_isLitKind(x_93); -x_95 = l_Lean_Syntax_getArgs(x_92); -x_96 = lean_array_get_size(x_95); -x_97 = lean_usize_of_nat(x_96); -lean_dec(x_96); -x_98 = 0; -x_99 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__3(x_97, x_98, x_95); -x_100 = l_Lean_Syntax_isIdent(x_92); -x_101 = lean_box(x_94); -lean_inc(x_92); -lean_inc(x_93); -lean_inc(x_99); -x_102 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___boxed), 13, 4); -lean_closure_set(x_102, 0, x_101); -lean_closure_set(x_102, 1, x_99); -lean_closure_set(x_102, 2, x_93); -lean_closure_set(x_102, 3, x_92); -if (x_100 == 0) +lean_object* x_95; lean_object* x_96; uint8_t x_97; lean_object* x_98; lean_object* x_99; size_t x_100; size_t x_101; lean_object* x_102; lean_object* x_103; uint8_t x_104; lean_object* x_105; lean_object* x_106; uint8_t x_107; +lean_dec(x_91); +lean_dec(x_89); +x_95 = l_Lean_Syntax_unescapeAntiquot(x_87); +lean_inc(x_95); +x_96 = l_Lean_Syntax_getKind(x_95); +x_97 = l_Lean_isLitKind(x_96); +x_98 = l_Lean_Syntax_getArgs(x_95); +x_99 = lean_array_get_size(x_98); +x_100 = lean_usize_of_nat(x_99); +lean_dec(x_99); +x_101 = 0; +x_102 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__1; +x_103 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__3(x_102, x_100, x_101, x_98); +x_104 = l_Lean_Syntax_isIdent(x_95); +x_105 = lean_box(x_97); +lean_inc(x_95); +lean_inc(x_96); +lean_inc(x_103); +x_106 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___boxed), 15, 6); +lean_closure_set(x_106, 0, x_9); +lean_closure_set(x_106, 1, x_105); +lean_closure_set(x_106, 2, x_103); +lean_closure_set(x_106, 3, x_102); +lean_closure_set(x_106, 4, x_96); +lean_closure_set(x_106, 5, x_95); +if (x_104 == 0) { -x_103 = x_94; -goto block_116; +x_107 = x_97; +goto block_119; } else { -uint8_t x_117; -x_117 = 1; -x_103 = x_117; -goto block_116; +uint8_t x_120; +x_120 = 1; +x_107 = x_120; +goto block_119; } -block_116: +block_119: { -lean_object* x_104; lean_object* x_105; -x_104 = lean_box(x_103); -lean_inc(x_92); -lean_inc(x_99); -lean_inc(x_93); -x_105 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__12___boxed), 5, 4); -lean_closure_set(x_105, 0, x_93); -lean_closure_set(x_105, 1, x_99); -lean_closure_set(x_105, 2, x_104); -lean_closure_set(x_105, 3, x_92); -if (x_103 == 0) -{ -lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; -lean_dec(x_92); -x_106 = lean_box(0); -x_107 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_107, 0, x_93); -lean_ctor_set(x_107, 1, x_106); -x_108 = lean_array_get_size(x_99); -lean_dec(x_99); -x_109 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_109, 0, x_108); +lean_object* x_108; lean_object* x_109; +x_108 = lean_box(x_107); +lean_inc(x_95); +lean_inc(x_103); +lean_inc(x_96); +x_109 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__12___boxed), 6, 5); +lean_closure_set(x_109, 0, x_96); +lean_closure_set(x_109, 1, x_9); +lean_closure_set(x_109, 2, x_103); +lean_closure_set(x_109, 3, x_108); +lean_closure_set(x_109, 4, x_95); +if (x_107 == 0) +{ +lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; +lean_dec(x_95); x_110 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_110, 0, x_107); -lean_ctor_set(x_110, 1, x_109); -x_111 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_111, 0, x_110); -lean_ctor_set(x_111, 1, x_105); -lean_ctor_set(x_111, 2, x_102); -x_112 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_96); +lean_ctor_set(x_110, 1, x_9); +x_111 = lean_array_get_size(x_103); +lean_dec(x_103); +x_112 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_112, 0, x_111); -lean_ctor_set(x_112, 1, x_8); -return x_112; -} -else -{ -lean_object* x_113; lean_object* x_114; lean_object* x_115; -lean_dec(x_99); -lean_dec(x_93); -x_113 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_113, 0, x_92); +x_113 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_113, 0, x_110); +lean_ctor_set(x_113, 1, x_112); x_114 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_114, 0, x_113); -lean_ctor_set(x_114, 1, x_105); -lean_ctor_set(x_114, 2, x_102); +lean_ctor_set(x_114, 1, x_109); +lean_ctor_set(x_114, 2, x_106); x_115 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_115, 0, x_114); lean_ctor_set(x_115, 1, x_8); return x_115; } +else +{ +lean_object* x_116; lean_object* x_117; lean_object* x_118; +lean_dec(x_103); +lean_dec(x_96); +x_116 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_116, 0, x_95); +x_117 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_117, 0, x_116); +lean_ctor_set(x_117, 1, x_109); +lean_ctor_set(x_117, 2, x_106); +x_118 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_118, 0, x_117); +lean_ctor_set(x_118, 1, x_8); +return x_118; +} } } else { -lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; -x_118 = lean_ctor_get(x_91, 0); -lean_inc(x_118); -lean_dec(x_91); -x_119 = l_Lean_Syntax_getNumArgs(x_84); -x_120 = lean_unsigned_to_nat(1u); -x_121 = lean_nat_sub(x_119, x_120); -lean_dec(x_119); -x_122 = lean_nat_sub(x_121, x_118); -lean_inc(x_122); -lean_inc(x_118); -x_123 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_123, 0, x_118); -lean_ctor_set(x_123, 1, x_122); -lean_inc(x_122); -lean_inc(x_118); -x_124 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__13___boxed), 6, 5); -lean_closure_set(x_124, 0, x_118); -lean_closure_set(x_124, 1, x_122); -lean_closure_set(x_124, 2, x_88); -lean_closure_set(x_124, 3, x_84); -lean_closure_set(x_124, 4, x_86); -x_125 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14), 12, 3); -lean_closure_set(x_125, 0, x_118); -lean_closure_set(x_125, 1, x_122); -lean_closure_set(x_125, 2, x_121); -x_126 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_126, 0, x_123); -lean_ctor_set(x_126, 1, x_124); -lean_ctor_set(x_126, 2, x_125); -x_127 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_127, 0, x_126); -lean_ctor_set(x_127, 1, x_8); -return x_127; +lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; +x_121 = lean_ctor_get(x_94, 0); +lean_inc(x_121); +lean_dec(x_94); +x_122 = l_Lean_Syntax_getNumArgs(x_87); +x_123 = lean_unsigned_to_nat(1u); +x_124 = lean_nat_sub(x_122, x_123); +lean_dec(x_122); +x_125 = lean_nat_sub(x_124, x_121); +lean_inc(x_125); +lean_inc(x_121); +x_126 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_126, 0, x_121); +lean_ctor_set(x_126, 1, x_125); +x_127 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__1; +lean_inc(x_125); +lean_inc(x_121); +x_128 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__13___boxed), 7, 6); +lean_closure_set(x_128, 0, x_121); +lean_closure_set(x_128, 1, x_125); +lean_closure_set(x_128, 2, x_91); +lean_closure_set(x_128, 3, x_127); +lean_closure_set(x_128, 4, x_87); +lean_closure_set(x_128, 5, x_89); +x_129 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___boxed), 14, 5); +lean_closure_set(x_129, 0, x_121); +lean_closure_set(x_129, 1, x_127); +lean_closure_set(x_129, 2, x_9); +lean_closure_set(x_129, 3, x_125); +lean_closure_set(x_129, 4, x_124); +x_130 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_130, 0, x_126); +lean_ctor_set(x_130, 1, x_128); +lean_ctor_set(x_130, 2, x_129); +x_131 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_131, 0, x_130); +lean_ctor_set(x_131, 1, x_8); +return x_131; } } } @@ -36493,195 +38169,195 @@ return x_127; } else { -lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; uint8_t x_226; -x_222 = lean_ctor_get(x_1, 0); -x_223 = lean_ctor_get(x_1, 1); -lean_inc(x_223); -lean_inc(x_222); +lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; uint8_t x_234; +x_230 = lean_ctor_get(x_1, 0); +x_231 = lean_ctor_get(x_1, 1); +lean_inc(x_231); +lean_inc(x_230); lean_dec(x_1); -x_224 = l_Lean_instInhabitedSyntax; -x_225 = l_List_head_x21___rarg(x_224, x_222); -x_226 = l_Lean_Syntax_isQuot(x_225); -if (x_226 == 0) -{ -lean_object* x_227; uint8_t x_228; -x_227 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_noOpMatchAdaptPats___lambda__1___closed__1; -lean_inc(x_225); -x_228 = l_Lean_Syntax_isOfKind(x_225, x_227); -if (x_228 == 0) -{ -lean_object* x_229; uint8_t x_230; -x_229 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__1; -lean_inc(x_225); -x_230 = l_Lean_Syntax_isOfKind(x_225, x_229); -if (x_230 == 0) -{ -lean_object* x_231; uint8_t x_232; -x_231 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__3; -lean_inc(x_225); -x_232 = l_Lean_Syntax_isOfKind(x_225, x_231); -if (x_232 == 0) -{ -lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; -lean_dec(x_223); -lean_dec(x_222); -lean_inc(x_225); -x_233 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_233, 0, x_225); -x_234 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__5; -x_235 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_235, 0, x_234); -lean_ctor_set(x_235, 1, x_233); -x_236 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__3; -x_237 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_237, 0, x_235); -lean_ctor_set(x_237, 1, x_236); -x_238 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__1(x_225, x_237, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_238; +x_232 = lean_box(0); +x_233 = l_List_head_x21___rarg(x_232, x_230); +x_234 = l_Lean_Syntax_isQuot(x_233); +if (x_234 == 0) +{ +lean_object* x_235; uint8_t x_236; +x_235 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_noOpMatchAdaptPats___lambda__1___closed__1; +lean_inc(x_233); +x_236 = l_Lean_Syntax_isOfKind(x_233, x_235); +if (x_236 == 0) +{ +lean_object* x_237; uint8_t x_238; +x_237 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__2; +lean_inc(x_233); +x_238 = l_Lean_Syntax_isOfKind(x_233, x_237); +if (x_238 == 0) +{ +lean_object* x_239; uint8_t x_240; +x_239 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__4; +lean_inc(x_233); +x_240 = l_Lean_Syntax_isOfKind(x_233, x_239); +if (x_240 == 0) +{ +lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; +lean_dec(x_231); +lean_dec(x_230); +lean_inc(x_233); +x_241 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_241, 0, x_233); +x_242 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__6; +x_243 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_243, 0, x_242); +lean_ctor_set(x_243, 1, x_241); +x_244 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__3; +x_245 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_245, 0, x_243); +lean_ctor_set(x_245, 1, x_244); +x_246 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__1(x_233, x_245, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_246; } else { -lean_object* x_239; lean_object* x_240; uint8_t x_241; -x_239 = lean_unsigned_to_nat(0u); -x_240 = l_Lean_Syntax_getArg(x_225, x_239); -lean_inc(x_240); -x_241 = l_Lean_Syntax_isOfKind(x_240, x_229); -if (x_241 == 0) +lean_object* x_247; lean_object* x_248; uint8_t x_249; +x_247 = lean_unsigned_to_nat(0u); +x_248 = l_Lean_Syntax_getArg(x_233, x_247); +lean_inc(x_248); +x_249 = l_Lean_Syntax_isOfKind(x_248, x_237); +if (x_249 == 0) { -lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; -lean_dec(x_240); -lean_dec(x_223); -lean_dec(x_222); -lean_inc(x_225); -x_242 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_242, 0, x_225); -x_243 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__5; -x_244 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_244, 0, x_243); -lean_ctor_set(x_244, 1, x_242); -x_245 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__3; -x_246 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_246, 0, x_244); -lean_ctor_set(x_246, 1, x_245); -x_247 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__1(x_225, x_246, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_247; +lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; +lean_dec(x_248); +lean_dec(x_231); +lean_dec(x_230); +lean_inc(x_233); +x_250 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_250, 0, x_233); +x_251 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__6; +x_252 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_252, 0, x_251); +lean_ctor_set(x_252, 1, x_250); +x_253 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__3; +x_254 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_254, 0, x_252); +lean_ctor_set(x_254, 1, x_253); +x_255 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__1(x_233, x_254, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_255; } else { -lean_object* x_248; lean_object* x_249; lean_object* x_250; uint8_t x_251; -x_248 = lean_unsigned_to_nat(2u); -x_249 = l_Lean_Syntax_getArg(x_225, x_248); -x_250 = l_Lean_nullKind; -x_251 = l_Lean_Syntax_isNodeOf(x_249, x_250, x_239); -if (x_251 == 0) +lean_object* x_256; lean_object* x_257; uint8_t x_258; +x_256 = lean_unsigned_to_nat(2u); +x_257 = l_Lean_Syntax_getArg(x_233, x_256); +x_258 = l_Lean_Syntax_matchesNull(x_257, x_247); +if (x_258 == 0) { -lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; -lean_dec(x_240); -lean_dec(x_223); -lean_dec(x_222); -lean_inc(x_225); -x_252 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_252, 0, x_225); -x_253 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__5; -x_254 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_254, 0, x_253); -lean_ctor_set(x_254, 1, x_252); -x_255 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__3; -x_256 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_256, 0, x_254); -lean_ctor_set(x_256, 1, x_255); -x_257 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__1(x_225, x_256, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_257; +lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; +lean_dec(x_248); +lean_dec(x_231); +lean_dec(x_230); +lean_inc(x_233); +x_259 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_259, 0, x_233); +x_260 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__6; +x_261 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_261, 0, x_260); +lean_ctor_set(x_261, 1, x_259); +x_262 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__3; +x_263 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_263, 0, x_261); +lean_ctor_set(x_263, 1, x_262); +x_264 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__1(x_233, x_263, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_264; } else { -lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; -x_258 = lean_unsigned_to_nat(3u); -x_259 = l_Lean_Syntax_getArg(x_225, x_258); -lean_dec(x_225); -x_260 = l_List_tail_x21___rarg(x_222); -x_261 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_261, 0, x_259); -lean_ctor_set(x_261, 1, x_260); -x_262 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_262, 0, x_261); -lean_ctor_set(x_262, 1, x_223); -x_263 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo(x_262, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_263) == 0) -{ -lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; -x_264 = lean_ctor_get(x_263, 0); -lean_inc(x_264); -x_265 = lean_ctor_get(x_263, 1); -lean_inc(x_265); -if (lean_is_exclusive(x_263)) { - lean_ctor_release(x_263, 0); - lean_ctor_release(x_263, 1); - x_266 = x_263; -} else { - lean_dec_ref(x_263); - x_266 = lean_box(0); -} -x_267 = lean_ctor_get(x_264, 0); -lean_inc(x_267); -x_268 = lean_ctor_get(x_264, 2); -lean_inc(x_268); -x_269 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -lean_inc(x_264); -x_270 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__3), 4, 3); -lean_closure_set(x_270, 0, x_264); -lean_closure_set(x_270, 1, x_269); -lean_closure_set(x_270, 2, x_240); -if (lean_is_exclusive(x_264)) { - lean_ctor_release(x_264, 0); - lean_ctor_release(x_264, 1); - lean_ctor_release(x_264, 2); - x_271 = x_264; -} else { - lean_dec_ref(x_264); - x_271 = lean_box(0); -} -if (lean_is_scalar(x_271)) { - x_272 = lean_alloc_ctor(0, 3, 0); -} else { - x_272 = x_271; -} -lean_ctor_set(x_272, 0, x_267); -lean_ctor_set(x_272, 1, x_270); -lean_ctor_set(x_272, 2, x_268); -if (lean_is_scalar(x_266)) { - x_273 = lean_alloc_ctor(0, 2, 0); +lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; +x_265 = lean_unsigned_to_nat(3u); +x_266 = l_Lean_Syntax_getArg(x_233, x_265); +lean_dec(x_233); +x_267 = l_List_tail_x21___rarg(x_230); +x_268 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_268, 0, x_266); +lean_ctor_set(x_268, 1, x_267); +x_269 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_269, 0, x_268); +lean_ctor_set(x_269, 1, x_231); +x_270 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo(x_269, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_270) == 0) +{ +lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; +x_271 = lean_ctor_get(x_270, 0); +lean_inc(x_271); +x_272 = lean_ctor_get(x_270, 1); +lean_inc(x_272); +if (lean_is_exclusive(x_270)) { + lean_ctor_release(x_270, 0); + lean_ctor_release(x_270, 1); + x_273 = x_270; } else { - x_273 = x_266; + lean_dec_ref(x_270); + x_273 = lean_box(0); } -lean_ctor_set(x_273, 0, x_272); -lean_ctor_set(x_273, 1, x_265); -return x_273; +x_274 = lean_ctor_get(x_271, 0); +lean_inc(x_274); +x_275 = lean_ctor_get(x_271, 2); +lean_inc(x_275); +x_276 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; +lean_inc(x_271); +x_277 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__3), 5, 4); +lean_closure_set(x_277, 0, x_271); +lean_closure_set(x_277, 1, x_276); +lean_closure_set(x_277, 2, x_9); +lean_closure_set(x_277, 3, x_248); +if (lean_is_exclusive(x_271)) { + lean_ctor_release(x_271, 0); + lean_ctor_release(x_271, 1); + lean_ctor_release(x_271, 2); + x_278 = x_271; +} else { + lean_dec_ref(x_271); + x_278 = lean_box(0); +} +if (lean_is_scalar(x_278)) { + x_279 = lean_alloc_ctor(0, 3, 0); +} else { + x_279 = x_278; +} +lean_ctor_set(x_279, 0, x_274); +lean_ctor_set(x_279, 1, x_277); +lean_ctor_set(x_279, 2, x_275); +if (lean_is_scalar(x_273)) { + x_280 = lean_alloc_ctor(0, 2, 0); +} else { + x_280 = x_273; +} +lean_ctor_set(x_280, 0, x_279); +lean_ctor_set(x_280, 1, x_272); +return x_280; } else { -lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; -lean_dec(x_240); -x_274 = lean_ctor_get(x_263, 0); -lean_inc(x_274); -x_275 = lean_ctor_get(x_263, 1); -lean_inc(x_275); -if (lean_is_exclusive(x_263)) { - lean_ctor_release(x_263, 0); - lean_ctor_release(x_263, 1); - x_276 = x_263; +lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; +lean_dec(x_248); +x_281 = lean_ctor_get(x_270, 0); +lean_inc(x_281); +x_282 = lean_ctor_get(x_270, 1); +lean_inc(x_282); +if (lean_is_exclusive(x_270)) { + lean_ctor_release(x_270, 0); + lean_ctor_release(x_270, 1); + x_283 = x_270; } else { - lean_dec_ref(x_263); - x_276 = lean_box(0); + lean_dec_ref(x_270); + x_283 = lean_box(0); } -if (lean_is_scalar(x_276)) { - x_277 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_283)) { + x_284 = lean_alloc_ctor(1, 2, 0); } else { - x_277 = x_276; + x_284 = x_283; } -lean_ctor_set(x_277, 0, x_274); -lean_ctor_set(x_277, 1, x_275); -return x_277; +lean_ctor_set(x_284, 0, x_281); +lean_ctor_set(x_284, 1, x_282); +return x_284; } } } @@ -36689,564 +38365,597 @@ return x_277; } else { -lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; -lean_dec(x_223); -lean_dec(x_222); +lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; +lean_dec(x_231); +lean_dec(x_230); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_278 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_279 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__1___boxed), 10, 2); -lean_closure_set(x_279, 0, x_278); -lean_closure_set(x_279, 1, x_225); -x_280 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__4), 2, 1); -lean_closure_set(x_280, 0, x_279); -x_281 = lean_box(0); -x_282 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__6; -x_283 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_283, 0, x_281); -lean_ctor_set(x_283, 1, x_280); -lean_ctor_set(x_283, 2, x_282); -x_284 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_284, 0, x_283); -lean_ctor_set(x_284, 1, x_8); -return x_284; +x_285 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; +x_286 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__1___boxed), 11, 3); +lean_closure_set(x_286, 0, x_285); +lean_closure_set(x_286, 1, x_9); +lean_closure_set(x_286, 2, x_233); +x_287 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__4), 2, 1); +lean_closure_set(x_287, 0, x_286); +x_288 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__5___boxed), 10, 1); +lean_closure_set(x_288, 0, x_9); +x_289 = lean_box(0); +x_290 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_290, 0, x_289); +lean_ctor_set(x_290, 1, x_287); +lean_ctor_set(x_290, 2, x_288); +x_291 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_291, 0, x_290); +lean_ctor_set(x_291, 1, x_8); +return x_291; } } else { -lean_object* x_285; lean_object* x_286; -lean_dec(x_225); -lean_dec(x_223); -lean_dec(x_222); +lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; +lean_dec(x_233); +lean_dec(x_231); +lean_dec(x_230); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_285 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__8; -x_286 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_286, 0, x_285); -lean_ctor_set(x_286, 1, x_8); -return x_286; +x_292 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__5___boxed), 10, 1); +lean_closure_set(x_292, 0, x_9); +x_293 = lean_box(0); +x_294 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__7; +x_295 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_295, 0, x_293); +lean_ctor_set(x_295, 1, x_294); +lean_ctor_set(x_295, 2, x_292); +x_296 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_296, 0, x_295); +lean_ctor_set(x_296, 1, x_8); +return x_296; } } else { -lean_object* x_287; lean_object* x_288; uint8_t x_359; -lean_dec(x_223); -lean_dec(x_222); -lean_inc(x_225); -x_287 = l_Lean_Syntax_getQuotContent(x_225); -x_359 = l_Lean_Syntax_isAtom(x_287); -if (x_359 == 0) +lean_object* x_297; lean_object* x_298; uint8_t x_370; +lean_dec(x_231); +lean_dec(x_230); +lean_inc(x_233); +x_297 = l_Lean_Syntax_getQuotContent(x_233); +x_370 = l_Lean_Syntax_isAtom(x_297); +if (x_370 == 0) { -uint8_t x_360; -lean_inc(x_287); -x_360 = l_Lean_Syntax_isTokenAntiquot(x_287); -if (x_360 == 0) +uint8_t x_371; +lean_inc(x_297); +x_371 = l_Lean_Syntax_isTokenAntiquot(x_297); +if (x_371 == 0) { -uint8_t x_361; -lean_inc(x_287); -x_361 = l_Lean_Syntax_isAntiquots(x_287); -if (x_361 == 0) +uint8_t x_372; +lean_inc(x_297); +x_372 = l_Lean_Syntax_isAntiquots(x_297); +if (x_372 == 0) { -uint8_t x_362; -x_362 = l_Lean_Syntax_isAntiquotSuffixSplice(x_287); -if (x_362 == 0) +uint8_t x_373; +x_373 = l_Lean_Syntax_isAntiquotSuffixSplice(x_297); +if (x_373 == 0) { -uint8_t x_363; -x_363 = l_Lean_Syntax_isAntiquotSplice(x_287); -if (x_363 == 0) +uint8_t x_374; +x_374 = l_Lean_Syntax_isAntiquotSplice(x_297); +if (x_374 == 0) { -lean_object* x_364; +lean_object* x_375; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_364 = lean_box(0); -x_288 = x_364; -goto block_358; +x_375 = lean_box(0); +x_298 = x_375; +goto block_369; } else { -lean_object* x_365; lean_object* x_366; -lean_dec(x_225); -x_365 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__48; -x_366 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__1(x_287, x_365, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_366; +lean_object* x_376; lean_object* x_377; +lean_dec(x_233); +x_376 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__48; +x_377 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__1(x_297, x_376, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_377; } } else { -lean_object* x_367; lean_object* x_368; -lean_dec(x_225); -x_367 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__48; -x_368 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__1(x_287, x_367, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_368; +lean_object* x_378; lean_object* x_379; +lean_dec(x_233); +x_378 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__48; +x_379 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__1(x_297, x_378, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_379; } } else { -lean_object* x_369; uint8_t x_370; -lean_inc(x_287); -x_369 = l_Lean_Syntax_getCanonicalAntiquot(x_287); -x_370 = l_Lean_Syntax_isEscapedAntiquot(x_369); -if (x_370 == 0) +lean_object* x_380; uint8_t x_381; +lean_inc(x_297); +x_380 = l_Lean_Syntax_getCanonicalAntiquot(x_297); +x_381 = l_Lean_Syntax_isEscapedAntiquot(x_380); +if (x_381 == 0) { -lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; uint8_t x_377; lean_object* x_378; lean_object* x_379; uint8_t x_380; lean_object* x_381; -lean_dec(x_225); +lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; uint8_t x_388; lean_object* x_389; lean_object* x_390; uint8_t x_391; lean_object* x_392; +lean_dec(x_233); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_371 = l_Lean_Syntax_antiquotKinds(x_287); -x_372 = l_List_unzip___rarg(x_371); -x_373 = lean_ctor_get(x_372, 0); -lean_inc(x_373); -x_374 = lean_ctor_get(x_372, 1); -lean_inc(x_374); -lean_dec(x_372); -x_375 = l_Lean_Syntax_getAntiquotTerm(x_369); -x_376 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_noOpMatchAdaptPats___lambda__1___closed__1; -lean_inc(x_375); -x_377 = l_Lean_Syntax_isOfKind(x_375, x_376); -x_378 = lean_unsigned_to_nat(3u); -x_379 = l_Lean_Syntax_getArg(x_369, x_378); -lean_dec(x_369); -x_380 = l_Lean_Syntax_isNone(x_379); -lean_dec(x_379); -if (x_377 == 0) +x_382 = l_Lean_Syntax_antiquotKinds(x_297); +x_383 = l_List_unzip___rarg(x_382); +x_384 = lean_ctor_get(x_383, 0); +lean_inc(x_384); +x_385 = lean_ctor_get(x_383, 1); +lean_inc(x_385); +lean_dec(x_383); +x_386 = l_Lean_Syntax_getAntiquotTerm(x_380); +x_387 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_noOpMatchAdaptPats___lambda__1___closed__1; +lean_inc(x_386); +x_388 = l_Lean_Syntax_isOfKind(x_386, x_387); +x_389 = lean_unsigned_to_nat(3u); +x_390 = l_Lean_Syntax_getArg(x_380, x_389); +lean_dec(x_380); +x_391 = l_Lean_Syntax_isNone(x_390); +lean_dec(x_390); +if (x_388 == 0) { -lean_object* x_403; uint8_t x_404; -x_403 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__1; -lean_inc(x_375); -x_404 = l_Lean_Syntax_isOfKind(x_375, x_403); -if (x_404 == 0) +lean_object* x_415; uint8_t x_416; +x_415 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__2; +lean_inc(x_386); +x_416 = l_Lean_Syntax_isOfKind(x_386, x_415); +if (x_416 == 0) { -lean_object* x_405; -x_405 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__28___boxed), 9, 1); -lean_closure_set(x_405, 0, x_375); -x_381 = x_405; -goto block_402; +lean_object* x_417; +x_417 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___boxed), 9, 1); +lean_closure_set(x_417, 0, x_386); +x_392 = x_417; +goto block_414; } else { -lean_object* x_406; lean_object* x_407; lean_object* x_408; -x_406 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_407 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__2; -lean_inc(x_373); -x_408 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__29___boxed), 12, 4); -lean_closure_set(x_408, 0, x_406); -lean_closure_set(x_408, 1, x_407); -lean_closure_set(x_408, 2, x_373); -lean_closure_set(x_408, 3, x_375); -x_381 = x_408; -goto block_402; +lean_object* x_418; lean_object* x_419; lean_object* x_420; +x_418 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; +x_419 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__2; +lean_inc(x_384); +x_420 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__28___boxed), 13, 5); +lean_closure_set(x_420, 0, x_418); +lean_closure_set(x_420, 1, x_419); +lean_closure_set(x_420, 2, x_9); +lean_closure_set(x_420, 3, x_384); +lean_closure_set(x_420, 4, x_386); +x_392 = x_420; +goto block_414; } } else { -lean_object* x_409; -lean_dec(x_375); -x_409 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__1; -x_381 = x_409; -goto block_402; +lean_object* x_421; +lean_dec(x_386); +x_421 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__1; +x_392 = x_421; +goto block_414; } -block_402: +block_414: { -if (x_380 == 0) +if (x_391 == 0) { -uint8_t x_382; uint8_t x_383; -x_382 = 1; -x_383 = l_List_foldr___at_List_and___spec__1(x_382, x_374); -lean_dec(x_374); -if (x_383 == 0) -{ -lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; -x_384 = lean_box(0); -lean_inc(x_373); -x_385 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_385, 0, x_373); -lean_ctor_set(x_385, 1, x_384); -lean_inc(x_373); -x_386 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___boxed), 3, 2); -lean_closure_set(x_386, 0, x_373); -lean_closure_set(x_386, 1, x_381); -x_387 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__2; -x_388 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_389 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27), 12, 3); -lean_closure_set(x_389, 0, x_387); -lean_closure_set(x_389, 1, x_388); -lean_closure_set(x_389, 2, x_373); -x_390 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_390, 0, x_385); -lean_ctor_set(x_390, 1, x_386); -lean_ctor_set(x_390, 2, x_389); -x_391 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_391, 0, x_390); -lean_ctor_set(x_391, 1, x_8); -return x_391; +uint8_t x_393; uint8_t x_394; +x_393 = 1; +x_394 = l_List_foldr___at_List_and___spec__1(x_393, x_385); +lean_dec(x_385); +if (x_394 == 0) +{ +lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; +x_395 = lean_box(0); +lean_inc(x_384); +x_396 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_396, 0, x_384); +lean_ctor_set(x_396, 1, x_395); +lean_inc(x_384); +x_397 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___boxed), 3, 2); +lean_closure_set(x_397, 0, x_384); +lean_closure_set(x_397, 1, x_392); +x_398 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__1; +x_399 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__2; +x_400 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; +x_401 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___boxed), 14, 5); +lean_closure_set(x_401, 0, x_9); +lean_closure_set(x_401, 1, x_398); +lean_closure_set(x_401, 2, x_399); +lean_closure_set(x_401, 3, x_400); +lean_closure_set(x_401, 4, x_384); +x_402 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_402, 0, x_396); +lean_ctor_set(x_402, 1, x_397); +lean_ctor_set(x_402, 2, x_401); +x_403 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_403, 0, x_402); +lean_ctor_set(x_403, 1, x_8); +return x_403; } else { -lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; -lean_dec(x_373); -x_392 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__4), 2, 1); -lean_closure_set(x_392, 0, x_381); -x_393 = lean_box(0); -x_394 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__6; -x_395 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_395, 0, x_393); -lean_ctor_set(x_395, 1, x_392); -lean_ctor_set(x_395, 2, x_394); -x_396 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_396, 0, x_395); -lean_ctor_set(x_396, 1, x_8); -return x_396; -} -} -else -{ -lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; -lean_dec(x_374); -lean_dec(x_373); -x_397 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__4), 2, 1); -lean_closure_set(x_397, 0, x_381); -x_398 = lean_box(0); -x_399 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__6; -x_400 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_400, 0, x_398); -lean_ctor_set(x_400, 1, x_397); -lean_ctor_set(x_400, 2, x_399); -x_401 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_401, 0, x_400); -lean_ctor_set(x_401, 1, x_8); -return x_401; +lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; +lean_dec(x_384); +x_404 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__4), 2, 1); +lean_closure_set(x_404, 0, x_392); +x_405 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__5___boxed), 10, 1); +lean_closure_set(x_405, 0, x_9); +x_406 = lean_box(0); +x_407 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_407, 0, x_406); +lean_ctor_set(x_407, 1, x_404); +lean_ctor_set(x_407, 2, x_405); +x_408 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_408, 0, x_407); +lean_ctor_set(x_408, 1, x_8); +return x_408; +} +} +else +{ +lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; +lean_dec(x_385); +lean_dec(x_384); +x_409 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__4), 2, 1); +lean_closure_set(x_409, 0, x_392); +x_410 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__5___boxed), 10, 1); +lean_closure_set(x_410, 0, x_9); +x_411 = lean_box(0); +x_412 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_412, 0, x_411); +lean_ctor_set(x_412, 1, x_409); +lean_ctor_set(x_412, 2, x_410); +x_413 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_413, 0, x_412); +lean_ctor_set(x_413, 1, x_8); +return x_413; } } } else { -uint8_t x_410; -lean_dec(x_369); -x_410 = l_Lean_Syntax_isAntiquotSuffixSplice(x_287); -if (x_410 == 0) +uint8_t x_422; +lean_dec(x_380); +x_422 = l_Lean_Syntax_isAntiquotSuffixSplice(x_297); +if (x_422 == 0) { -uint8_t x_411; -x_411 = l_Lean_Syntax_isAntiquotSplice(x_287); -if (x_411 == 0) +uint8_t x_423; +x_423 = l_Lean_Syntax_isAntiquotSplice(x_297); +if (x_423 == 0) { -lean_object* x_412; +lean_object* x_424; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_412 = lean_box(0); -x_288 = x_412; -goto block_358; +x_424 = lean_box(0); +x_298 = x_424; +goto block_369; } else { -lean_object* x_413; lean_object* x_414; -lean_dec(x_225); -x_413 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__48; -x_414 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__1(x_287, x_413, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_414; +lean_object* x_425; lean_object* x_426; +lean_dec(x_233); +x_425 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__48; +x_426 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__1(x_297, x_425, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_426; } } else { -lean_object* x_415; lean_object* x_416; -lean_dec(x_225); -x_415 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__48; -x_416 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__1(x_287, x_415, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_416; +lean_object* x_427; lean_object* x_428; +lean_dec(x_233); +x_427 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__48; +x_428 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__1(x_297, x_427, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_428; } } } } else { -lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; -lean_dec(x_225); +lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; +lean_dec(x_233); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_417 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__30___boxed), 9, 1); -lean_closure_set(x_417, 0, x_287); -x_418 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__4), 2, 1); -lean_closure_set(x_418, 0, x_417); -x_419 = lean_box(0); -x_420 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__6; -x_421 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_421, 0, x_419); -lean_ctor_set(x_421, 1, x_418); -lean_ctor_set(x_421, 2, x_420); -x_422 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_422, 0, x_421); -lean_ctor_set(x_422, 1, x_8); -return x_422; +x_429 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__29___boxed), 10, 2); +lean_closure_set(x_429, 0, x_297); +lean_closure_set(x_429, 1, x_9); +x_430 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__4), 2, 1); +lean_closure_set(x_430, 0, x_429); +x_431 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__5___boxed), 10, 1); +lean_closure_set(x_431, 0, x_9); +x_432 = lean_box(0); +x_433 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_433, 0, x_432); +lean_ctor_set(x_433, 1, x_430); +lean_ctor_set(x_433, 2, x_431); +x_434 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_434, 0, x_433); +lean_ctor_set(x_434, 1, x_8); +return x_434; } } else { -lean_object* x_423; lean_object* x_424; -lean_dec(x_287); -lean_dec(x_225); +lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; +lean_dec(x_297); +lean_dec(x_233); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_423 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__8; -x_424 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_424, 0, x_423); -lean_ctor_set(x_424, 1, x_8); -return x_424; +x_435 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__5___boxed), 10, 1); +lean_closure_set(x_435, 0, x_9); +x_436 = lean_box(0); +x_437 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__7; +x_438 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_438, 0, x_436); +lean_ctor_set(x_438, 1, x_437); +lean_ctor_set(x_438, 2, x_435); +x_439 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_439, 0, x_438); +lean_ctor_set(x_439, 1, x_8); +return x_439; } -block_358: +block_369: { -lean_object* x_289; lean_object* x_290; lean_object* x_332; lean_object* x_333; uint8_t x_334; -lean_dec(x_288); -x_289 = l_Lean_Syntax_getArgs(x_287); -x_332 = lean_array_get_size(x_289); -x_333 = lean_unsigned_to_nat(1u); -x_334 = lean_nat_dec_eq(x_332, x_333); -lean_dec(x_332); -if (x_334 == 0) +lean_object* x_299; lean_object* x_300; lean_object* x_343; lean_object* x_344; uint8_t x_345; +lean_dec(x_298); +x_299 = l_Lean_Syntax_getArgs(x_297); +x_343 = lean_array_get_size(x_299); +x_344 = lean_unsigned_to_nat(1u); +x_345 = lean_nat_dec_eq(x_343, x_344); +lean_dec(x_343); +if (x_345 == 0) { -lean_object* x_335; -lean_dec(x_225); -x_335 = lean_box(0); -x_290 = x_335; -goto block_331; +lean_object* x_346; +lean_dec(x_233); +x_346 = lean_box(0); +x_300 = x_346; +goto block_342; } else { -lean_object* x_336; lean_object* x_337; uint8_t x_338; -x_336 = lean_unsigned_to_nat(0u); -x_337 = l_Lean_Syntax_getArg(x_287, x_336); -x_338 = l_Lean_Syntax_isAntiquotSuffixSplice(x_337); -if (x_338 == 0) +lean_object* x_347; lean_object* x_348; uint8_t x_349; +x_347 = lean_unsigned_to_nat(0u); +x_348 = l_Lean_Syntax_getArg(x_297, x_347); +x_349 = l_Lean_Syntax_isAntiquotSuffixSplice(x_348); +if (x_349 == 0) { -uint8_t x_339; -x_339 = l_Lean_Syntax_isAntiquotSplice(x_337); -if (x_339 == 0) -{ -lean_object* x_340; -lean_dec(x_337); -lean_dec(x_225); -x_340 = lean_box(0); -x_290 = x_340; -goto block_331; -} -else +uint8_t x_350; +x_350 = l_Lean_Syntax_isAntiquotSplice(x_348); +if (x_350 == 0) { -lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; -lean_dec(x_289); -lean_dec(x_287); -lean_inc(x_225); -x_341 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_341, 0, x_225); -x_342 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__15), 2, 1); -lean_closure_set(x_342, 0, x_225); -x_343 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23), 10, 1); -lean_closure_set(x_343, 0, x_337); -x_344 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_344, 0, x_341); -lean_ctor_set(x_344, 1, x_342); -lean_ctor_set(x_344, 2, x_343); -x_345 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_345, 0, x_344); -lean_ctor_set(x_345, 1, x_8); -return x_345; -} +lean_object* x_351; +lean_dec(x_348); +lean_dec(x_233); +x_351 = lean_box(0); +x_300 = x_351; +goto block_342; } else { -lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; -lean_dec(x_289); -lean_dec(x_225); -x_346 = l_Lean_Syntax_getArg(x_337, x_336); -lean_inc(x_346); -x_347 = l_Lean_Syntax_getCanonicalAntiquot(x_346); -x_348 = l_Lean_Syntax_getAntiquotTerm(x_347); -lean_dec(x_347); -x_349 = l_Lean_Syntax_antiquotKinds(x_346); -x_350 = lean_box(0); -x_351 = l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__7(x_349, x_350); -x_352 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___boxed), 13, 5); -lean_closure_set(x_352, 0, x_337); -lean_closure_set(x_352, 1, x_287); -lean_closure_set(x_352, 2, x_350); -lean_closure_set(x_352, 3, x_351); -lean_closure_set(x_352, 4, x_348); -x_353 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__4), 2, 1); -lean_closure_set(x_353, 0, x_352); -x_354 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___boxed), 10, 1); -lean_closure_set(x_354, 0, x_350); -x_355 = lean_box(0); +lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; +lean_dec(x_299); +lean_dec(x_297); +lean_inc(x_233); +x_352 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_352, 0, x_233); +x_353 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__15), 2, 1); +lean_closure_set(x_353, 0, x_233); +x_354 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__1; +x_355 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23), 12, 3); +lean_closure_set(x_355, 0, x_348); +lean_closure_set(x_355, 1, x_9); +lean_closure_set(x_355, 2, x_354); x_356 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_356, 0, x_355); +lean_ctor_set(x_356, 0, x_352); lean_ctor_set(x_356, 1, x_353); -lean_ctor_set(x_356, 2, x_354); +lean_ctor_set(x_356, 2, x_355); x_357 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_357, 0, x_356); lean_ctor_set(x_357, 1, x_8); return x_357; } } -block_331: +else { -lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; -lean_dec(x_290); -x_291 = lean_array_get_size(x_289); -x_292 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__9; -x_293 = lean_unsigned_to_nat(0u); -lean_inc(x_291); -x_294 = l_Array_findIdx_x3f_loop___rarg(x_289, x_292, x_291, x_293, lean_box(0)); -if (lean_obj_tag(x_294) == 0) -{ -lean_object* x_295; lean_object* x_296; uint8_t x_297; lean_object* x_298; lean_object* x_299; size_t x_300; size_t x_301; lean_object* x_302; uint8_t x_303; lean_object* x_304; lean_object* x_305; uint8_t x_306; -lean_dec(x_291); -lean_dec(x_289); -x_295 = l_Lean_Syntax_unescapeAntiquot(x_287); -lean_inc(x_295); -x_296 = l_Lean_Syntax_getKind(x_295); -x_297 = l_Lean_isLitKind(x_296); -x_298 = l_Lean_Syntax_getArgs(x_295); -x_299 = lean_array_get_size(x_298); -x_300 = lean_usize_of_nat(x_299); +lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_dec(x_299); -x_301 = 0; -x_302 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__3(x_300, x_301, x_298); -x_303 = l_Lean_Syntax_isIdent(x_295); -x_304 = lean_box(x_297); -lean_inc(x_295); -lean_inc(x_296); -lean_inc(x_302); -x_305 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___boxed), 13, 4); -lean_closure_set(x_305, 0, x_304); -lean_closure_set(x_305, 1, x_302); -lean_closure_set(x_305, 2, x_296); -lean_closure_set(x_305, 3, x_295); -if (x_303 == 0) +lean_dec(x_233); +x_358 = l_Lean_Syntax_getArg(x_348, x_347); +lean_inc(x_358); +x_359 = l_Lean_Syntax_getCanonicalAntiquot(x_358); +x_360 = l_Lean_Syntax_getAntiquotTerm(x_359); +lean_dec(x_359); +x_361 = l_Lean_Syntax_antiquotKinds(x_358); +x_362 = l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__21(x_361, x_9); +x_363 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___boxed), 13, 5); +lean_closure_set(x_363, 0, x_348); +lean_closure_set(x_363, 1, x_297); +lean_closure_set(x_363, 2, x_9); +lean_closure_set(x_363, 3, x_362); +lean_closure_set(x_363, 4, x_360); +x_364 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__4), 2, 1); +lean_closure_set(x_364, 0, x_363); +x_365 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__5___boxed), 10, 1); +lean_closure_set(x_365, 0, x_9); +x_366 = lean_box(0); +x_367 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_367, 0, x_366); +lean_ctor_set(x_367, 1, x_364); +lean_ctor_set(x_367, 2, x_365); +x_368 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_368, 0, x_367); +lean_ctor_set(x_368, 1, x_8); +return x_368; +} +} +block_342: +{ +lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; +lean_dec(x_300); +x_301 = lean_array_get_size(x_299); +x_302 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__8; +x_303 = lean_unsigned_to_nat(0u); +lean_inc(x_301); +x_304 = l_Array_findIdx_x3f_loop___rarg(x_299, x_302, x_301, x_303, lean_box(0)); +if (lean_obj_tag(x_304) == 0) +{ +lean_object* x_305; lean_object* x_306; uint8_t x_307; lean_object* x_308; lean_object* x_309; size_t x_310; size_t x_311; lean_object* x_312; lean_object* x_313; uint8_t x_314; lean_object* x_315; lean_object* x_316; uint8_t x_317; +lean_dec(x_301); +lean_dec(x_299); +x_305 = l_Lean_Syntax_unescapeAntiquot(x_297); +lean_inc(x_305); +x_306 = l_Lean_Syntax_getKind(x_305); +x_307 = l_Lean_isLitKind(x_306); +x_308 = l_Lean_Syntax_getArgs(x_305); +x_309 = lean_array_get_size(x_308); +x_310 = lean_usize_of_nat(x_309); +lean_dec(x_309); +x_311 = 0; +x_312 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__1; +x_313 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__3(x_312, x_310, x_311, x_308); +x_314 = l_Lean_Syntax_isIdent(x_305); +x_315 = lean_box(x_307); +lean_inc(x_305); +lean_inc(x_306); +lean_inc(x_313); +x_316 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___boxed), 15, 6); +lean_closure_set(x_316, 0, x_9); +lean_closure_set(x_316, 1, x_315); +lean_closure_set(x_316, 2, x_313); +lean_closure_set(x_316, 3, x_312); +lean_closure_set(x_316, 4, x_306); +lean_closure_set(x_316, 5, x_305); +if (x_314 == 0) { -x_306 = x_297; -goto block_319; +x_317 = x_307; +goto block_329; } else { -uint8_t x_320; -x_320 = 1; -x_306 = x_320; -goto block_319; +uint8_t x_330; +x_330 = 1; +x_317 = x_330; +goto block_329; } -block_319: +block_329: { -lean_object* x_307; lean_object* x_308; -x_307 = lean_box(x_306); -lean_inc(x_295); -lean_inc(x_302); -lean_inc(x_296); -x_308 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__12___boxed), 5, 4); -lean_closure_set(x_308, 0, x_296); -lean_closure_set(x_308, 1, x_302); -lean_closure_set(x_308, 2, x_307); -lean_closure_set(x_308, 3, x_295); -if (x_306 == 0) -{ -lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; -lean_dec(x_295); -x_309 = lean_box(0); -x_310 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_310, 0, x_296); -lean_ctor_set(x_310, 1, x_309); -x_311 = lean_array_get_size(x_302); -lean_dec(x_302); -x_312 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_312, 0, x_311); -x_313 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_313, 0, x_310); -lean_ctor_set(x_313, 1, x_312); -x_314 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_314, 0, x_313); -lean_ctor_set(x_314, 1, x_308); -lean_ctor_set(x_314, 2, x_305); -x_315 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_315, 0, x_314); -lean_ctor_set(x_315, 1, x_8); -return x_315; +lean_object* x_318; lean_object* x_319; +x_318 = lean_box(x_317); +lean_inc(x_305); +lean_inc(x_313); +lean_inc(x_306); +x_319 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__12___boxed), 6, 5); +lean_closure_set(x_319, 0, x_306); +lean_closure_set(x_319, 1, x_9); +lean_closure_set(x_319, 2, x_313); +lean_closure_set(x_319, 3, x_318); +lean_closure_set(x_319, 4, x_305); +if (x_317 == 0) +{ +lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; +lean_dec(x_305); +x_320 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_320, 0, x_306); +lean_ctor_set(x_320, 1, x_9); +x_321 = lean_array_get_size(x_313); +lean_dec(x_313); +x_322 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_322, 0, x_321); +x_323 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_323, 0, x_320); +lean_ctor_set(x_323, 1, x_322); +x_324 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_324, 0, x_323); +lean_ctor_set(x_324, 1, x_319); +lean_ctor_set(x_324, 2, x_316); +x_325 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_325, 0, x_324); +lean_ctor_set(x_325, 1, x_8); +return x_325; } else { -lean_object* x_316; lean_object* x_317; lean_object* x_318; -lean_dec(x_302); -lean_dec(x_296); -x_316 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_316, 0, x_295); -x_317 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_317, 0, x_316); -lean_ctor_set(x_317, 1, x_308); -lean_ctor_set(x_317, 2, x_305); -x_318 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_318, 0, x_317); -lean_ctor_set(x_318, 1, x_8); -return x_318; +lean_object* x_326; lean_object* x_327; lean_object* x_328; +lean_dec(x_313); +lean_dec(x_306); +x_326 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_326, 0, x_305); +x_327 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_327, 0, x_326); +lean_ctor_set(x_327, 1, x_319); +lean_ctor_set(x_327, 2, x_316); +x_328 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_328, 0, x_327); +lean_ctor_set(x_328, 1, x_8); +return x_328; } } } else { -lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; -x_321 = lean_ctor_get(x_294, 0); -lean_inc(x_321); -lean_dec(x_294); -x_322 = l_Lean_Syntax_getNumArgs(x_287); -x_323 = lean_unsigned_to_nat(1u); -x_324 = lean_nat_sub(x_322, x_323); -lean_dec(x_322); -x_325 = lean_nat_sub(x_324, x_321); -lean_inc(x_325); -lean_inc(x_321); -x_326 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_326, 0, x_321); -lean_ctor_set(x_326, 1, x_325); -lean_inc(x_325); -lean_inc(x_321); -x_327 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__13___boxed), 6, 5); -lean_closure_set(x_327, 0, x_321); -lean_closure_set(x_327, 1, x_325); -lean_closure_set(x_327, 2, x_291); -lean_closure_set(x_327, 3, x_287); -lean_closure_set(x_327, 4, x_289); -x_328 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14), 12, 3); -lean_closure_set(x_328, 0, x_321); -lean_closure_set(x_328, 1, x_325); -lean_closure_set(x_328, 2, x_324); -x_329 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_329, 0, x_326); -lean_ctor_set(x_329, 1, x_327); -lean_ctor_set(x_329, 2, x_328); -x_330 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_330, 0, x_329); -lean_ctor_set(x_330, 1, x_8); -return x_330; +lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; +x_331 = lean_ctor_get(x_304, 0); +lean_inc(x_331); +lean_dec(x_304); +x_332 = l_Lean_Syntax_getNumArgs(x_297); +x_333 = lean_unsigned_to_nat(1u); +x_334 = lean_nat_sub(x_332, x_333); +lean_dec(x_332); +x_335 = lean_nat_sub(x_334, x_331); +lean_inc(x_335); +lean_inc(x_331); +x_336 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_336, 0, x_331); +lean_ctor_set(x_336, 1, x_335); +x_337 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__1; +lean_inc(x_335); +lean_inc(x_331); +x_338 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__13___boxed), 7, 6); +lean_closure_set(x_338, 0, x_331); +lean_closure_set(x_338, 1, x_335); +lean_closure_set(x_338, 2, x_301); +lean_closure_set(x_338, 3, x_337); +lean_closure_set(x_338, 4, x_297); +lean_closure_set(x_338, 5, x_299); +x_339 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___boxed), 14, 5); +lean_closure_set(x_339, 0, x_331); +lean_closure_set(x_339, 1, x_337); +lean_closure_set(x_339, 2, x_9); +lean_closure_set(x_339, 3, x_335); +lean_closure_set(x_339, 4, x_334); +x_340 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_340, 0, x_336); +lean_ctor_set(x_340, 1, x_338); +lean_ctor_set(x_340, 2, x_339); +x_341 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_341, 0, x_340); +lean_ctor_set(x_341, 1, x_8); +return x_341; } } } @@ -37276,29 +38985,31 @@ lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -size_t x_4; size_t x_5; lean_object* x_6; -x_4 = lean_unbox_usize(x_1); -lean_dec(x_1); +size_t x_5; size_t x_6; lean_object* x_7; x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__3(x_4, x_5, x_3); -return x_6; +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__3(x_1, x_5, x_6, x_4); +lean_dec(x_1); +return x_7; } } -LEAN_EXPORT lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___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_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___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_object* x_9) { _start: { -lean_object* x_9; -x_9 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); +lean_object* x_10; +x_10 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -return x_9; +lean_dec(x_1); +return x_10; } } LEAN_EXPORT lean_object* l_List_beq___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__5___boxed(lean_object* x_1, lean_object* x_2) { @@ -37312,15 +39023,16 @@ x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__6___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_EXPORT lean_object* l_Array_mapIdxM_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__6___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_8; -x_8 = l_Array_mapIdxM_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_object* x_9; +x_9 = l_Array_mapIdxM_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_8; +return x_9; } } LEAN_EXPORT lean_object* l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__7___boxed(lean_object** _args) { @@ -37347,17 +39059,21 @@ lean_object* x_20 = _args[19]; lean_object* x_21 = _args[20]; lean_object* x_22 = _args[21]; lean_object* x_23 = _args[22]; +lean_object* x_24 = _args[23]; +lean_object* x_25 = _args[24]; _start: { -lean_object* x_24; -x_24 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23); +lean_object* x_26; +x_26 = l_List_mapM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21, x_22, x_23, x_24, x_25); +lean_dec(x_24); lean_dec(x_22); +lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_2); lean_dec(x_1); -return x_24; +return x_26; } } LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8___boxed(lean_object** _args) { @@ -37380,21 +39096,23 @@ lean_object* x_16 = _args[15]; lean_object* x_17 = _args[16]; lean_object* x_18 = _args[17]; lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; _start: { -size_t x_20; size_t x_21; lean_object* x_22; -x_20 = lean_unbox_usize(x_10); -lean_dec(x_10); +size_t x_21; size_t x_22; lean_object* x_23; x_21 = lean_unbox_usize(x_11); lean_dec(x_11); -x_22 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_20, x_21, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); -lean_dec(x_18); +x_22 = lean_unbox_usize(x_12); +lean_dec(x_12); +x_23 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21, x_22, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +lean_dec(x_19); +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_9); -return x_22; +lean_dec(x_10); +lean_dec(x_1); +return x_23; } } LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___boxed(lean_object** _args) { @@ -37417,21 +39135,23 @@ lean_object* x_16 = _args[15]; lean_object* x_17 = _args[16]; lean_object* x_18 = _args[17]; lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; _start: { -size_t x_20; size_t x_21; lean_object* x_22; -x_20 = lean_unbox_usize(x_10); -lean_dec(x_10); +size_t x_21; size_t x_22; lean_object* x_23; x_21 = lean_unbox_usize(x_11); lean_dec(x_11); -x_22 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_20, x_21, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); -lean_dec(x_18); +x_22 = lean_unbox_usize(x_12); +lean_dec(x_12); +x_23 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21, x_22, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +lean_dec(x_19); +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_9); -return x_22; +lean_dec(x_10); +lean_dec(x_1); +return x_23; } } LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__10___boxed(lean_object** _args) { @@ -37454,21 +39174,23 @@ lean_object* x_16 = _args[15]; lean_object* x_17 = _args[16]; lean_object* x_18 = _args[17]; lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; _start: { -size_t x_20; size_t x_21; lean_object* x_22; -x_20 = lean_unbox_usize(x_10); -lean_dec(x_10); +size_t x_21; size_t x_22; lean_object* x_23; x_21 = lean_unbox_usize(x_11); lean_dec(x_11); -x_22 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__10(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_20, x_21, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); -lean_dec(x_18); +x_22 = lean_unbox_usize(x_12); +lean_dec(x_12); +x_23 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__10(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21, x_22, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +lean_dec(x_19); +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_9); -return x_22; +lean_dec(x_10); +lean_dec(x_1); +return x_23; } } LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__11___boxed(lean_object** _args) { @@ -37521,21 +39243,23 @@ lean_object* x_16 = _args[15]; lean_object* x_17 = _args[16]; lean_object* x_18 = _args[17]; lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; _start: { -size_t x_20; size_t x_21; lean_object* x_22; -x_20 = lean_unbox_usize(x_10); -lean_dec(x_10); +size_t x_21; size_t x_22; lean_object* x_23; x_21 = lean_unbox_usize(x_11); lean_dec(x_11); -x_22 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__12(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_20, x_21, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); -lean_dec(x_18); +x_22 = lean_unbox_usize(x_12); +lean_dec(x_12); +x_23 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__12(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21, x_22, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +lean_dec(x_19); +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_9); -return x_22; +lean_dec(x_10); +lean_dec(x_1); +return x_23; } } LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__13___boxed(lean_object** _args) { @@ -37558,21 +39282,23 @@ lean_object* x_16 = _args[15]; lean_object* x_17 = _args[16]; lean_object* x_18 = _args[17]; lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; _start: { -size_t x_20; size_t x_21; lean_object* x_22; -x_20 = lean_unbox_usize(x_10); -lean_dec(x_10); +size_t x_21; size_t x_22; lean_object* x_23; x_21 = lean_unbox_usize(x_11); lean_dec(x_11); -x_22 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__13(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_20, x_21, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); -lean_dec(x_18); +x_22 = lean_unbox_usize(x_12); +lean_dec(x_12); +x_23 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__13(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21, x_22, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +lean_dec(x_19); +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_9); -return x_22; +lean_dec(x_10); +lean_dec(x_1); +return x_23; } } LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__14___boxed(lean_object** _args) { @@ -37595,56 +39321,59 @@ lean_object* x_16 = _args[15]; lean_object* x_17 = _args[16]; lean_object* x_18 = _args[17]; lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; _start: { -size_t x_20; size_t x_21; lean_object* x_22; -x_20 = lean_unbox_usize(x_10); -lean_dec(x_10); +size_t x_21; size_t x_22; lean_object* x_23; x_21 = lean_unbox_usize(x_11); lean_dec(x_11); -x_22 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__14(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_20, x_21, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); -lean_dec(x_18); +x_22 = lean_unbox_usize(x_12); +lean_dec(x_12); +x_23 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__14(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21, x_22, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +lean_dec(x_19); +lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_9); -return x_22; +lean_dec(x_10); +lean_dec(x_1); +return x_23; } } -LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15___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_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) { _start: { -lean_object* x_15; -x_15 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_13); +lean_object* x_16; +x_16 = l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__15(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_14); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); -return x_15; +lean_dec(x_1); +return x_16; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___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, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___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, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_11; -x_11 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); +lean_object* x_12; +x_12 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -return x_11; +return x_12; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__5___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_object* x_9) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__5___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_object* x_9, lean_object* x_10) { _start: { -lean_object* x_10; -x_10 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_2); -return x_10; +lean_object* x_11; +x_11 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_3); +return x_11; } } LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__7___boxed(lean_object* x_1) { @@ -37657,23 +39386,24 @@ x_3 = lean_box(x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__9___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_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__9___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_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_11; -x_11 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_object* x_12; +x_12 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_2); lean_dec(x_1); -return x_11; +return x_12; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___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_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) { _start: { -uint8_t x_14; lean_object* x_15; -x_14 = lean_unbox(x_1); -lean_dec(x_1); -x_15 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10(x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_15; +uint8_t x_16; lean_object* x_17; +x_16 = lean_unbox(x_2); +lean_dec(x_2); +x_17 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10(x_1, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_17; } } LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__11___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_object* x_9) { @@ -37690,27 +39420,37 @@ lean_dec(x_3); return x_10; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__12___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: { -uint8_t x_6; lean_object* x_7; -x_6 = lean_unbox(x_3); -lean_dec(x_3); -x_7 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__12(x_1, x_2, x_6, x_4, x_5); -return x_7; +uint8_t x_7; lean_object* x_8; +x_7 = lean_unbox(x_4); +lean_dec(x_4); +x_8 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__12(x_1, x_2, x_3, x_7, x_5, x_6); +return x_8; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__13___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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__13___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: { -lean_object* x_7; -x_7 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__13(x_1, x_2, x_3, x_4, x_5, x_6); +lean_object* x_8; +x_8 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__13(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -return x_7; +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; +x_15 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_2); +return x_15; } } LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16___boxed(lean_object** _args) { @@ -37765,16 +39505,12 @@ lean_object* x_16 = _args[15]; lean_object* x_17 = _args[16]; lean_object* x_18 = _args[17]; lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; _start: { -lean_object* x_20; -x_20 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); -lean_dec(x_18); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -return x_20; +lean_object* x_21; +x_21 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__17(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +return x_21; } } LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__18___boxed(lean_object** _args) { @@ -37797,16 +39533,12 @@ lean_object* x_16 = _args[15]; lean_object* x_17 = _args[16]; lean_object* x_18 = _args[17]; lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; _start: { -lean_object* x_20; -x_20 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__18(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); -lean_dec(x_18); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -return x_20; +lean_object* x_21; +x_21 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__18(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +return x_21; } } LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__19___boxed(lean_object** _args) { @@ -37829,16 +39561,12 @@ lean_object* x_16 = _args[15]; lean_object* x_17 = _args[16]; lean_object* x_18 = _args[17]; lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; _start: { -lean_object* x_20; -x_20 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__19(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); -lean_dec(x_18); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -return x_20; +lean_object* x_21; +x_21 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__19(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +return x_21; } } LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__20___boxed(lean_object** _args) { @@ -37861,16 +39589,12 @@ lean_object* x_16 = _args[15]; lean_object* x_17 = _args[16]; lean_object* x_18 = _args[17]; lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; _start: { -lean_object* x_20; -x_20 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__20(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); -lean_dec(x_18); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -return x_20; +lean_object* x_21; +x_21 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__20(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +return x_21; } } LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__21___boxed(lean_object** _args) { @@ -37893,16 +39617,12 @@ lean_object* x_16 = _args[15]; lean_object* x_17 = _args[16]; lean_object* x_18 = _args[17]; lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; _start: { -lean_object* x_20; -x_20 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__21(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); -lean_dec(x_18); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -return x_20; +lean_object* x_21; +x_21 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__21(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +return x_21; } } LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__22___boxed(lean_object** _args) { @@ -37925,16 +39645,12 @@ lean_object* x_16 = _args[15]; lean_object* x_17 = _args[16]; lean_object* x_18 = _args[17]; lean_object* x_19 = _args[18]; +lean_object* x_20 = _args[19]; _start: { -lean_object* x_20; -x_20 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__22(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); -lean_dec(x_18); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -return x_20; +lean_object* x_21; +x_21 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__22(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +return x_21; } } LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { @@ -37946,58 +39662,58 @@ lean_dec(x_1); return x_14; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___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_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_11; -x_11 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_3); -return x_11; +lean_object* x_4; +x_4 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_4; -x_4 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26(x_1, x_2, x_3); -lean_dec(x_1); -return x_4; +lean_object* x_15; +x_15 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_2); +return x_15; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__28___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_object* x_9) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___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_object* x_9) { _start: { lean_object* x_10; -x_10 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__28(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_10 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_2); return x_10; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__29___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__28___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_13; -x_13 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__29(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); +lean_object* x_14; +x_14 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__28(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); -return x_13; +return x_14; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__30___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_object* x_9) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__29___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_object* x_9, lean_object* x_10) { _start: { -lean_object* x_10; -x_10 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__30(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_8); +lean_object* x_11; +x_11 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__29(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); lean_dec(x_1); -return x_10; +return x_11; } } LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { @@ -38074,9 +39790,9 @@ LEAN_EXPORT lean_object* l_Array_sequenceMap___at___private_Lean_Elab_Quotation_ { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_3 = lean_array_get_size(x_1); -x_4 = lean_unsigned_to_nat(0u); -x_5 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; -x_6 = l_Array_sequenceMap_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___spec__2(x_1, x_2, x_3, x_4, x_5); +x_4 = lean_mk_empty_array_with_capacity(x_3); +x_5 = lean_unsigned_to_nat(0u); +x_6 = l_Array_sequenceMap_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___spec__2(x_1, x_2, x_3, x_5, x_4); return x_6; } } @@ -38240,7 +39956,7 @@ lean_ctor_set(x_51, 3, x_29); x_52 = lean_array_push(x_37, x_36); x_53 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; x_54 = lean_array_push(x_52, x_53); -x_55 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__22; +x_55 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__18; x_56 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_56, 0, x_34); lean_ctor_set(x_56, 1, x_55); @@ -38304,7 +40020,7 @@ lean_ctor_set(x_81, 3, x_29); x_82 = lean_array_push(x_37, x_36); x_83 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; x_84 = lean_array_push(x_82, x_83); -x_85 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__22; +x_85 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__18; x_86 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_86, 0, x_34); lean_ctor_set(x_86, 1, x_85); @@ -38434,7 +40150,7 @@ lean_ctor_set(x_145, 0, x_137); lean_ctor_set(x_145, 1, x_119); lean_ctor_set(x_145, 2, x_144); lean_ctor_set(x_145, 3, x_118); -x_146 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__2; +x_146 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__14___closed__2; lean_inc(x_137); x_147 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_147, 0, x_137); @@ -38453,7 +40169,7 @@ lean_ctor_set(x_152, 2, x_151); x_153 = lean_array_push(x_131, x_152); x_154 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; x_155 = lean_array_push(x_153, x_154); -x_156 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__22; +x_156 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__18; x_157 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_157, 0, x_128); lean_ctor_set(x_157, 1, x_156); @@ -38513,7 +40229,7 @@ lean_ctor_set(x_181, 0, x_137); lean_ctor_set(x_181, 1, x_119); lean_ctor_set(x_181, 2, x_180); lean_ctor_set(x_181, 3, x_118); -x_182 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__2; +x_182 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__14___closed__2; lean_inc(x_137); x_183 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_183, 0, x_137); @@ -38532,7 +40248,7 @@ lean_ctor_set(x_188, 2, x_187); x_189 = lean_array_push(x_131, x_188); x_190 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__16; x_191 = lean_array_push(x_189, x_190); -x_192 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__22; +x_192 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__18; x_193 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_193, 0, x_128); lean_ctor_set(x_193, 1, x_192); @@ -38683,7 +40399,7 @@ else lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; x_17 = lean_unsigned_to_nat(0u); x_18 = l_Lean_Syntax_getArg(x_12, x_17); -x_19 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__1; +x_19 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__2; x_20 = l_Lean_Syntax_isOfKind(x_18, x_19); if (x_20 == 0) { @@ -38756,7 +40472,7 @@ else lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; x_38 = lean_unsigned_to_nat(0u); x_39 = l_Lean_Syntax_getArg(x_33, x_38); -x_40 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__1; +x_40 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__2; x_41 = l_Lean_Syntax_isOfKind(x_39, x_40); if (x_41 == 0) { @@ -41145,7 +42861,7 @@ x_20 = lean_st_ref_get(x_10, x_19); x_21 = lean_ctor_get(x_20, 1); lean_inc(x_21); lean_dec(x_20); -x_22 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__25; +x_22 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__25; lean_inc(x_18); x_23 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_23, 0, x_18); @@ -41156,8 +42872,8 @@ x_25 = l_List_redLength___rarg(x_24); x_26 = lean_mk_empty_array_with_capacity(x_25); lean_dec(x_25); x_27 = l_List_toArrayAux___rarg(x_24, x_26); -x_28 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__18; -x_29 = l_Lean_Syntax_SepArray_ofElems(x_28, x_27); +x_28 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__19; +x_29 = l_Lean_mkSepArray(x_27, x_28); lean_dec(x_27); lean_inc(x_1); x_30 = l_Array_append___rarg(x_1, x_29); @@ -41173,7 +42889,7 @@ x_36 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_36, 0, x_31); lean_ctor_set(x_36, 1, x_32); lean_ctor_set(x_36, 2, x_35); -x_37 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__12; +x_37 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__12; x_38 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_38, 0, x_18); lean_ctor_set(x_38, 1, x_37); @@ -41185,7 +42901,7 @@ x_41 = lean_array_push(x_40, x_23); x_42 = lean_array_push(x_41, x_36); x_43 = lean_array_push(x_42, x_38); x_44 = lean_array_push(x_43, x_39); -x_45 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__24; +x_45 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__24; x_46 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_46, 0, x_31); lean_ctor_set(x_46, 1, x_45); @@ -41263,7 +42979,7 @@ x_18 = lean_st_ref_get(x_10, x_17); x_19 = lean_ctor_get(x_18, 1); lean_inc(x_19); lean_dec(x_18); -x_20 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__15; +x_20 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__11; lean_inc(x_16); x_21 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_21, 0, x_16); @@ -41278,7 +42994,7 @@ x_26 = lean_array_push(x_25, x_14); x_27 = lean_array_push(x_26, x_23); x_28 = lean_array_push(x_27, x_4); x_29 = lean_box(2); -x_30 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__16; +x_30 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__12; x_31 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_31, 0, x_29); lean_ctor_set(x_31, 1, x_30); @@ -41292,51 +43008,51 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__8(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -if (lean_obj_tag(x_1) == 0) +if (lean_obj_tag(x_2) == 0) { -lean_object* x_3; -x_3 = l_List_reverse___rarg(x_2); -return x_3; +lean_object* x_4; +x_4 = l_List_reverse___rarg(x_3); +return x_4; } else { -uint8_t x_4; -x_4 = !lean_is_exclusive(x_1); -if (x_4 == 0) +uint8_t x_5; +x_5 = !lean_is_exclusive(x_2); +if (x_5 == 0) { -lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_5 = lean_ctor_get(x_1, 0); -x_6 = lean_ctor_get(x_1, 1); -x_7 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_7, 0, x_5); -lean_ctor_set(x_1, 1, x_2); -lean_ctor_set(x_1, 0, x_7); +lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_6 = lean_ctor_get(x_2, 0); +x_7 = lean_ctor_get(x_2, 1); +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_6); +lean_ctor_set(x_2, 1, x_3); +lean_ctor_set(x_2, 0, x_8); { -lean_object* _tmp_0 = x_6; -lean_object* _tmp_1 = x_1; -x_1 = _tmp_0; +lean_object* _tmp_1 = x_7; +lean_object* _tmp_2 = x_2; x_2 = _tmp_1; +x_3 = _tmp_2; } goto _start; } else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_ctor_get(x_1, 0); -x_10 = lean_ctor_get(x_1, 1); +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_ctor_get(x_2, 0); +x_11 = lean_ctor_get(x_2, 1); +lean_inc(x_11); lean_inc(x_10); -lean_inc(x_9); -lean_dec(x_1); -x_11 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_11, 0, x_9); -x_12 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_2); -x_1 = x_10; -x_2 = x_12; +lean_dec(x_2); +x_12 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_12, 0, x_10); +x_13 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_3); +x_2 = x_11; +x_3 = x_13; goto _start; } } @@ -41415,7 +43131,7 @@ static lean_object* _init_l_List_format___at___private_Lean_Elab_Quotation_0__Le _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__18; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__18; x_2 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; @@ -41473,7 +43189,7 @@ static lean_object* _init_l_List_format___at___private_Lean_Elab_Quotation_0__Le _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__24; +x_1 = l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__24; x_2 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; @@ -41573,7 +43289,7 @@ static lean_object* _init_l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__ _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; x_2 = lean_string_length(x_1); return x_2; } @@ -41591,7 +43307,7 @@ static lean_object* _init_l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__ _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7; x_2 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; @@ -41640,7 +43356,7 @@ x_19 = l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot x_20 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_20, 0, x_19); lean_ctor_set(x_20, 1, x_18); -x_21 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__16; +x_21 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__16; x_22 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_22, 0, x_20); lean_ctor_set(x_22, 1, x_21); @@ -41697,7 +43413,7 @@ x_43 = l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot x_44 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_44, 0, x_43); lean_ctor_set(x_44, 1, x_42); -x_45 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__16; +x_45 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__16; x_46 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_46, 0, x_44); lean_ctor_set(x_46, 1, x_45); @@ -41721,13 +43437,64 @@ goto _start; } } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___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, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___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, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_10) == 0) +{ +uint8_t x_11; +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +return x_10; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_10); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +return x_14; +} +} +else +{ +uint8_t x_15; +x_15 = !lean_is_exclusive(x_10); +if (x_15 == 0) +{ +return x_10; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_10, 0); +x_17 = lean_ctor_get(x_10, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_10); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +return x_18; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___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_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; x_12 = l_List_appendTR___rarg(x_1, x_2); x_13 = lean_array_to_list(lean_box(0), x_3); -x_14 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch), 9, 6); +x_14 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__1), 9, 6); lean_closure_set(x_14, 0, x_12); lean_closure_set(x_14, 1, x_13); lean_closure_set(x_14, 2, x_5); @@ -41738,7 +43505,7 @@ x_15 = l_Lean_Core_withFreshMacroScope___rarg(x_14, x_9, x_10, x_11); return x_15; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__3(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { uint8_t x_13; @@ -41795,7 +43562,7 @@ x_35 = lean_ctor_get(x_33, 0); lean_inc(x_35); lean_dec(x_33); x_36 = lean_environment_main_module(x_35); -x_37 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__16; +x_37 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__16; lean_inc(x_29); x_38 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_38, 0, x_29); @@ -41822,12 +43589,12 @@ lean_inc(x_41); x_48 = lean_array_push(x_47, x_41); lean_inc(x_48); x_49 = lean_array_push(x_48, x_46); -x_50 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__46; +x_50 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__46; x_51 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_51, 0, x_39); lean_ctor_set(x_51, 1, x_50); lean_ctor_set(x_51, 2, x_49); -x_52 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__18; +x_52 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__18; lean_inc(x_29); x_53 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_53, 0, x_29); @@ -41842,7 +43609,7 @@ x_59 = lean_array_get_size(x_58); x_60 = lean_usize_of_nat(x_59); lean_dec(x_59); x_61 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__6(x_50, x_48, x_60, x_24, x_58); -x_62 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__19; +x_62 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__19; x_63 = l_Lean_mkSepArray(x_61, x_62); lean_dec(x_61); x_64 = l_Array_append___rarg(x_55, x_63); @@ -41850,7 +43617,7 @@ x_65 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_65, 0, x_39); lean_ctor_set(x_65, 1, x_40); lean_ctor_set(x_65, 2, x_64); -x_66 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__20; +x_66 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__20; x_67 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_67, 0, x_29); lean_ctor_set(x_67, 1, x_66); @@ -41861,12 +43628,12 @@ lean_ctor_set(x_69, 1, x_40); lean_ctor_set(x_69, 2, x_68); x_70 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__19; x_71 = lean_array_push(x_70, x_69); -x_72 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__22; +x_72 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__22; x_73 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_73, 0, x_39); lean_ctor_set(x_73, 1, x_72); lean_ctor_set(x_73, 2, x_71); -x_74 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__30; +x_74 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__30; x_75 = lean_array_push(x_74, x_38); lean_inc(x_41); x_76 = lean_array_push(x_75, x_41); @@ -41874,7 +43641,7 @@ x_77 = lean_array_push(x_76, x_41); x_78 = lean_array_push(x_77, x_65); x_79 = lean_array_push(x_78, x_67); x_80 = lean_array_push(x_79, x_73); -x_81 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__17; +x_81 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__17; x_82 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_82, 0, x_39); lean_ctor_set(x_82, 1, x_81); @@ -41884,7 +43651,7 @@ lean_ctor_set(x_83, 0, x_21); lean_ctor_set(x_83, 1, x_82); x_84 = lean_array_push(x_4, x_83); x_85 = lean_box(0); -x_86 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__1(x_5, x_1, x_84, x_85, x_6, x_7, x_8, x_9, x_10, x_11, x_34); +x_86 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__2(x_5, x_1, x_84, x_85, x_6, x_7, x_8, x_9, x_10, x_11, x_34); return x_86; } else @@ -41927,16 +43694,16 @@ lean_object* x_91; lean_object* x_92; lean_dec(x_3); lean_dec(x_2); x_91 = lean_box(0); -x_92 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__1(x_5, x_1, x_4, x_91, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_92 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__2(x_5, x_1, x_4, x_91, x_6, x_7, x_8, x_9, x_10, x_11, x_12); return x_92; } } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__3(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_object* x_9) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4(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_object* x_9) { _start: { lean_object* x_10; lean_object* x_11; -x_10 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch), 9, 6); +x_10 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__1), 9, 6); lean_closure_set(x_10, 0, x_1); lean_closure_set(x_10, 1, x_2); lean_closure_set(x_10, 2, x_3); @@ -41947,7 +43714,7 @@ x_11 = l_Lean_Core_withFreshMacroScope___rarg(x_10, x_7, x_8, x_9); return x_11; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__1() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__1() { _start: { lean_object* x_1; @@ -41955,27 +43722,27 @@ x_1 = lean_mk_string_from_bytes("non-exhaustive 'match' (syntax)", 31); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__2() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__1; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___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_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__3() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__2; +x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__2; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__4() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__4() { _start: { lean_object* x_1; @@ -41983,20 +43750,20 @@ x_1 = lean_mk_string_from_bytes("_private.Lean.Elab.Quotation.0.Lean.Elab.Term.Q return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__5() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__5() { _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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepFromSplice___closed__1; -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__4; -x_3 = lean_unsigned_to_nat(541u); +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__4; +x_3 = lean_unsigned_to_nat(544u); x_4 = lean_unsigned_to_nat(12u); x_5 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepFromSplice___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); return x_6; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__6() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__6() { _start: { lean_object* x_1; lean_object* x_2; @@ -42007,31 +43774,31 @@ lean_ctor_set(x_2, 1, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__7() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__6; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__6; x_3 = lean_alloc_ctor(0, 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_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__8() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; -x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__7; +x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__7; x_3 = lean_alloc_ctor(0, 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_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__9() { +static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; @@ -42045,7 +43812,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4(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_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5(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_object* x_9, lean_object* x_10) { _start: { lean_dec(x_3); @@ -42054,7 +43821,7 @@ if (lean_obj_tag(x_1) == 0) if (lean_obj_tag(x_2) == 0) { lean_object* x_11; uint8_t x_12; lean_object* x_13; uint8_t x_14; -x_11 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__3; +x_11 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__3; x_12 = 2; x_13 = l_Lean_log___at_Lean_Elab_Term_traceAtCmdPos___spec__2(x_11, x_12, x_4, x_5, x_6, x_7, x_8, x_9, x_10); x_14 = !lean_is_exclusive(x_13); @@ -42110,7 +43877,7 @@ else lean_object* x_24; lean_object* x_25; lean_dec(x_21); lean_dec(x_20); -x_24 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__5; +x_24 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__5; x_25 = l_panic___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__1(x_24, x_4, x_5, x_6, x_7, x_8, x_9, x_10); return x_25; } @@ -42122,7 +43889,7 @@ if (lean_obj_tag(x_2) == 0) { lean_object* x_26; uint8_t x_27; lean_object* x_28; uint8_t x_29; lean_dec(x_1); -x_26 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__3; +x_26 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__3; x_27 = 2; x_28 = l_Lean_log___at_Lean_Elab_Term_traceAtCmdPos___spec__2(x_26, x_27, x_4, x_5, x_6, x_7, x_8, x_9, x_10); x_29 = !lean_is_exclusive(x_28); @@ -42195,7 +43962,7 @@ lean_inc(x_45); x_46 = lean_ctor_get(x_44, 1); lean_inc(x_46); lean_dec(x_44); -x_47 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__8; +x_47 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__8; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); @@ -42231,14 +43998,14 @@ lean_inc(x_57); lean_dec(x_42); x_58 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; lean_inc(x_38); -x_59 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__2), 12, 4); +x_59 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__3), 12, 4); lean_closure_set(x_59, 0, x_38); lean_closure_set(x_59, 1, x_55); lean_closure_set(x_59, 2, x_58); lean_closure_set(x_59, 3, x_56); lean_inc(x_37); x_60 = lean_array_to_list(lean_box(0), x_54); -x_61 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__3), 9, 2); +x_61 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4), 9, 2); lean_closure_set(x_61, 0, x_1); lean_closure_set(x_61, 1, x_60); lean_inc(x_9); @@ -42315,7 +44082,7 @@ lean_ctor_set(x_88, 0, x_72); lean_ctor_set(x_88, 1, x_87); x_89 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; x_90 = lean_array_push(x_89, x_86); -x_91 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__9; +x_91 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__9; x_92 = lean_array_push(x_90, x_91); x_93 = lean_array_push(x_92, x_91); x_94 = lean_array_push(x_93, x_88); @@ -42384,7 +44151,7 @@ lean_ctor_set(x_124, 0, x_72); lean_ctor_set(x_124, 1, x_123); x_125 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; x_126 = lean_array_push(x_125, x_122); -x_127 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__9; +x_127 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__9; x_128 = lean_array_push(x_126, x_127); x_129 = lean_array_push(x_128, x_127); x_130 = lean_array_push(x_129, x_124); @@ -42599,7 +44366,7 @@ lean_inc(x_174); x_175 = lean_ctor_get(x_173, 1); lean_inc(x_175); lean_dec(x_173); -x_176 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__8; +x_176 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__8; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); @@ -42635,14 +44402,14 @@ lean_inc(x_186); lean_dec(x_170); x_187 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; lean_inc(x_166); -x_188 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__2), 12, 4); +x_188 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__3), 12, 4); lean_closure_set(x_188, 0, x_166); lean_closure_set(x_188, 1, x_184); lean_closure_set(x_188, 2, x_187); lean_closure_set(x_188, 3, x_185); lean_inc(x_165); x_189 = lean_array_to_list(lean_box(0), x_183); -x_190 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__3), 9, 2); +x_190 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4), 9, 2); lean_closure_set(x_190, 0, x_1); lean_closure_set(x_190, 1, x_189); lean_inc(x_9); @@ -42726,7 +44493,7 @@ lean_ctor_set(x_218, 0, x_201); lean_ctor_set(x_218, 1, x_217); x_219 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; x_220 = lean_array_push(x_219, x_216); -x_221 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__9; +x_221 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__9; x_222 = lean_array_push(x_220, x_221); x_223 = lean_array_push(x_222, x_221); x_224 = lean_array_push(x_223, x_218); @@ -42966,7 +44733,7 @@ lean_inc(x_269); x_270 = lean_ctor_get(x_268, 1); lean_inc(x_270); lean_dec(x_268); -x_271 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__8; +x_271 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__8; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); @@ -43002,7 +44769,7 @@ lean_inc(x_281); lean_dec(x_265); x_282 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; lean_inc(x_260); -x_283 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__2), 12, 4); +x_283 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__3), 12, 4); lean_closure_set(x_283, 0, x_260); lean_closure_set(x_283, 1, x_279); lean_closure_set(x_283, 2, x_282); @@ -43012,7 +44779,7 @@ x_284 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_284, 0, x_259); lean_ctor_set(x_284, 1, x_260); x_285 = lean_array_to_list(lean_box(0), x_278); -x_286 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__3), 9, 2); +x_286 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4), 9, 2); lean_closure_set(x_286, 0, x_284); lean_closure_set(x_286, 1, x_285); lean_inc(x_9); @@ -43096,7 +44863,7 @@ lean_ctor_set(x_314, 0, x_297); lean_ctor_set(x_314, 1, x_313); x_315 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__18; x_316 = lean_array_push(x_315, x_312); -x_317 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__9; +x_317 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__9; x_318 = lean_array_push(x_316, x_317); x_319 = lean_array_push(x_318, x_317); x_320 = lean_array_push(x_319, x_314); @@ -43283,7 +45050,7 @@ static lean_object* _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quot { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__1; +x_2 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } @@ -43343,89 +45110,90 @@ return x_2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch(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_object* x_9) { _start: { -lean_object* x_10; uint8_t x_11; lean_object* x_12; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +lean_object* x_10; uint8_t x_11; lean_object* x_12; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; x_10 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___closed__3; -x_32 = lean_st_ref_get(x_8, x_9); -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_33, 3); +x_33 = lean_st_ref_get(x_8, x_9); +x_34 = lean_ctor_get(x_33, 0); lean_inc(x_34); -lean_dec(x_33); -x_35 = lean_ctor_get_uint8(x_34, sizeof(void*)*1); +x_35 = lean_ctor_get(x_34, 3); +lean_inc(x_35); lean_dec(x_34); -if (x_35 == 0) +x_36 = lean_ctor_get_uint8(x_35, sizeof(void*)*1); +lean_dec(x_35); +if (x_36 == 0) { -lean_object* x_36; uint8_t x_37; -x_36 = lean_ctor_get(x_32, 1); -lean_inc(x_36); -lean_dec(x_32); -x_37 = 0; -x_11 = x_37; -x_12 = x_36; -goto block_31; +lean_object* x_37; uint8_t x_38; +x_37 = lean_ctor_get(x_33, 1); +lean_inc(x_37); +lean_dec(x_33); +x_38 = 0; +x_11 = x_38; +x_12 = x_37; +goto block_32; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; -x_38 = lean_ctor_get(x_32, 1); -lean_inc(x_38); -lean_dec(x_32); -x_39 = l___private_Lean_Util_Trace_0__Lean_checkTraceOptionM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTerm___spec__2(x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_38); -x_40 = lean_ctor_get(x_39, 0); -lean_inc(x_40); -x_41 = lean_ctor_get(x_39, 1); +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; +x_39 = lean_ctor_get(x_33, 1); +lean_inc(x_39); +lean_dec(x_33); +x_40 = l___private_Lean_Util_Trace_0__Lean_checkTraceOptionM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTerm___spec__2(x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_39); +x_41 = lean_ctor_get(x_40, 0); lean_inc(x_41); -lean_dec(x_39); -x_42 = lean_unbox(x_40); +x_42 = lean_ctor_get(x_40, 1); +lean_inc(x_42); lean_dec(x_40); -x_11 = x_42; -x_12 = x_41; -goto block_31; +x_43 = lean_unbox(x_41); +lean_dec(x_41); +x_11 = x_43; +x_12 = x_42; +goto block_32; } -block_31: +block_32: { if (x_11 == 0) { lean_object* x_13; lean_object* x_14; x_13 = lean_box(0); -x_14 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4(x_1, x_2, x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_12); +x_14 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5(x_1, x_2, x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_12); return x_14; } else { -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; 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_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; 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; x_15 = lean_box(0); +x_16 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__1; lean_inc(x_1); -x_16 = l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__8(x_1, x_15); -x_17 = l_Lean_MessageData_ofList(x_16); -lean_dec(x_16); -x_18 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___closed__5; -x_19 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_17); -x_20 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___closed__7; -x_21 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); +x_17 = l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__8(x_16, x_1, x_15); +x_18 = l_Lean_MessageData_ofList(x_17); +lean_dec(x_17); +x_19 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___closed__5; +x_20 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_18); +x_21 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___closed__7; +x_22 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); lean_inc(x_2); -x_22 = l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__11___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__12(x_2, x_15); -x_23 = l_Lean_MessageData_ofList(x_22); -lean_dec(x_22); -x_24 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_24, 0, x_21); -lean_ctor_set(x_24, 1, x_23); -x_25 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__3; -x_26 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -x_27 = l_Lean_addTrace___at___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTerm___spec__1(x_10, x_26, x_3, x_4, x_5, x_6, x_7, x_8, x_12); -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); +x_23 = l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__11___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__12(x_2, x_15); +x_24 = l_Lean_MessageData_ofList(x_23); +lean_dec(x_23); +x_25 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_25, 0, x_22); +lean_ctor_set(x_25, 1, x_24); +x_26 = l_Lean_Elab_Term_Quotation_getQuotKind___closed__3; +x_27 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l_Lean_addTrace___at___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTerm___spec__1(x_10, x_27, x_3, x_4, x_5, x_6, x_7, x_8, x_12); +x_29 = lean_ctor_get(x_28, 0); lean_inc(x_29); -lean_dec(x_27); -x_30 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4(x_1, x_2, x_28, x_3, x_4, x_5, x_6, x_7, x_8, x_29); -return x_30; +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5(x_1, x_2, x_29, x_3, x_4, x_5, x_6, x_7, x_8, x_30); +return x_31; } } } @@ -43498,11 +45266,20 @@ lean_dec(x_1); return x_14; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___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, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__8(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___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, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; -x_12 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_12 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_4); return x_12; } @@ -43581,9 +45358,9 @@ LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Elab_Term_Quotation_match { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_3 = lean_array_get_size(x_1); -x_4 = lean_unsigned_to_nat(0u); -x_5 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; -x_6 = l_Array_sequenceMap_loop___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__2(x_1, x_2, x_3, x_4, x_5); +x_4 = lean_mk_empty_array_with_capacity(x_3); +x_5 = lean_unsigned_to_nat(0u); +x_6 = l_Array_sequenceMap_loop___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__2(x_1, x_2, x_3, x_5, x_4); return x_6; } } @@ -43666,31 +45443,11 @@ goto _start; } } } -static lean_object* _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__6___rarg___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Elab_unsupportedSyntaxExceptionId; -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__6___rarg___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__6___rarg___closed__1; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; -} -} LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__6___rarg(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__6___rarg___closed__2; +x_2 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___rarg___closed__2; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); @@ -43714,7 +45471,7 @@ if (x_5 == 0) { lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; x_6 = lean_array_uget(x_2, x_3); -x_7 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__2; +x_7 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__3; lean_inc(x_1); x_8 = lean_name_mk_string(x_1, x_7); lean_inc(x_6); @@ -43743,65 +45500,64 @@ return x_14; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; x_15 = lean_unsigned_to_nat(2u); x_16 = l_Lean_Syntax_getArg(x_6, x_15); -x_17 = l_Lean_nullKind; -x_18 = lean_unsigned_to_nat(0u); -x_19 = l_Lean_Syntax_isNodeOf(x_16, x_17, x_18); -if (x_19 == 0) +x_17 = lean_unsigned_to_nat(0u); +x_18 = l_Lean_Syntax_matchesNull(x_16, x_17); +if (x_18 == 0) { -uint8_t x_20; -x_20 = l_Lean_Syntax_isQuot(x_6); +uint8_t x_19; +x_19 = l_Lean_Syntax_isQuot(x_6); lean_dec(x_6); -if (x_20 == 0) +if (x_19 == 0) { -size_t x_21; size_t x_22; -x_21 = 1; -x_22 = lean_usize_add(x_3, x_21); -x_3 = x_22; +size_t x_20; size_t x_21; +x_20 = 1; +x_21 = lean_usize_add(x_3, x_20); +x_3 = x_21; goto _start; } else { -uint8_t x_24; +uint8_t x_23; lean_dec(x_1); -x_24 = 1; -return x_24; +x_23 = 1; +return x_23; } } else { -lean_object* x_25; lean_object* x_26; uint8_t x_27; -x_25 = lean_unsigned_to_nat(3u); -x_26 = l_Lean_Syntax_getArg(x_6, x_25); +lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_24 = lean_unsigned_to_nat(3u); +x_25 = l_Lean_Syntax_getArg(x_6, x_24); lean_dec(x_6); -x_27 = l_Lean_Syntax_isQuot(x_26); -lean_dec(x_26); -if (x_27 == 0) +x_26 = l_Lean_Syntax_isQuot(x_25); +lean_dec(x_25); +if (x_26 == 0) { -size_t x_28; size_t x_29; -x_28 = 1; -x_29 = lean_usize_add(x_3, x_28); -x_3 = x_29; +size_t x_27; size_t x_28; +x_27 = 1; +x_28 = lean_usize_add(x_3, x_27); +x_3 = x_28; goto _start; } else { -uint8_t x_31; +uint8_t x_30; lean_dec(x_1); -x_31 = 1; -return x_31; +x_30 = 1; +return x_30; } } } } else { -uint8_t x_32; +uint8_t x_31; lean_dec(x_1); -x_32 = 0; -return x_32; +x_31 = 0; +return x_31; } } } @@ -43880,7 +45636,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_match__syntax_expand___lambd _start: { lean_object* x_3; lean_object* x_4; uint8_t x_5; -x_3 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__1; +x_3 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__12___closed__1; x_4 = lean_name_mk_string(x_1, x_3); lean_inc(x_2); x_5 = l_Lean_Syntax_isOfKind(x_2, x_4); @@ -43894,27 +45650,26 @@ return x_6; } else { -lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +lean_object* x_7; lean_object* x_8; uint8_t x_9; x_7 = lean_unsigned_to_nat(0u); x_8 = l_Lean_Syntax_getArg(x_2, x_7); -x_9 = l_Lean_nullKind; -x_10 = l_Lean_Syntax_isNodeOf(x_8, x_9, x_7); -if (x_10 == 0) +x_9 = l_Lean_Syntax_matchesNull(x_8, x_7); +if (x_9 == 0) { -lean_object* x_11; +lean_object* x_10; lean_dec(x_2); -x_11 = lean_box(0); -return x_11; +x_10 = lean_box(0); +return x_10; } else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_unsigned_to_nat(1u); -x_13 = l_Lean_Syntax_getArg(x_2, x_12); +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_unsigned_to_nat(1u); +x_12 = l_Lean_Syntax_getArg(x_2, x_11); lean_dec(x_2); -x_14 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_14, 0, x_13); -return x_14; +x_13 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_13, 0, x_12); +return x_13; } } } @@ -43940,7 +45695,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_match__syntax_expand___lambd _start: { lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_5 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__23; +x_5 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__23; x_6 = lean_name_mk_string(x_1, x_5); lean_inc(x_4); x_7 = l_Lean_Syntax_isOfKind(x_4, x_6); @@ -43950,44 +45705,79 @@ if (x_7 == 0) lean_object* x_8; lean_dec(x_4); lean_dec(x_3); +lean_dec(x_2); x_8 = lean_box(0); return x_8; } else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +lean_object* x_9; lean_object* x_10; uint8_t x_11; x_9 = lean_unsigned_to_nat(1u); x_10 = l_Lean_Syntax_getArg(x_4, x_9); -x_11 = l_Lean_nullKind; lean_inc(x_10); -x_12 = l_Lean_Syntax_isNodeOf(x_10, x_11, x_9); -if (x_12 == 0) +x_11 = l_Lean_Syntax_matchesNull(x_10, x_9); +if (x_11 == 0) { -lean_object* x_13; +lean_object* x_12; lean_dec(x_10); lean_dec(x_4); lean_dec(x_3); -x_13 = lean_box(0); -return x_13; +lean_dec(x_2); +x_12 = lean_box(0); +return x_12; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_14 = lean_unsigned_to_nat(0u); -x_15 = l_Lean_Syntax_getArg(x_10, x_14); +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; +x_13 = lean_unsigned_to_nat(0u); +x_14 = l_Lean_Syntax_getArg(x_10, x_13); lean_dec(x_10); -x_16 = l_Lean_Syntax_getArgs(x_15); +x_15 = l_Lean_Syntax_getArgs(x_14); +lean_dec(x_14); +x_16 = lean_array_get_size(x_15); +x_17 = lean_nat_dec_lt(x_13, x_16); +if (x_17 == 0) +{ +lean_dec(x_16); lean_dec(x_15); -x_17 = lean_array_get_size(x_16); -x_18 = lean_nat_dec_lt(x_14, x_17); -if (x_18 == 0) +lean_dec(x_3); +x_18 = x_2; +goto block_32; +} +else +{ +uint8_t x_33; +x_33 = lean_nat_dec_le(x_16, x_16); +if (x_33 == 0) { -lean_object* x_19; lean_object* x_20; -lean_dec(x_17); lean_dec(x_16); +lean_dec(x_15); lean_dec(x_3); +x_18 = x_2; +goto block_32; +} +else +{ +size_t x_34; size_t x_35; lean_object* x_36; lean_object* x_37; +lean_dec(x_2); +x_34 = 0; +x_35 = lean_usize_of_nat(x_16); +lean_dec(x_16); +x_36 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(x_15, x_34, x_35, x_3); +lean_dec(x_15); +x_37 = lean_ctor_get(x_36, 1); +lean_inc(x_37); +lean_dec(x_36); +x_18 = x_37; +goto block_32; +} +} +block_32: +{ +lean_object* x_19; lean_object* x_20; x_19 = l_Lean_Elab_Term_Quotation_match__syntax_expand___lambda__3___closed__1; -x_20 = l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__1(x_2, x_19); +x_20 = l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__1(x_18, x_19); +lean_dec(x_18); if (lean_obj_tag(x_20) == 0) { lean_object* x_21; @@ -44030,117 +45820,6 @@ return x_31; } } } -else -{ -uint8_t x_32; -x_32 = lean_nat_dec_le(x_17, x_17); -if (x_32 == 0) -{ -lean_object* x_33; lean_object* x_34; -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_3); -x_33 = l_Lean_Elab_Term_Quotation_match__syntax_expand___lambda__3___closed__1; -x_34 = l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__1(x_2, x_33); -if (lean_obj_tag(x_34) == 0) -{ -lean_object* x_35; -lean_dec(x_4); -x_35 = lean_box(0); -return x_35; -} -else -{ -uint8_t x_36; -x_36 = !lean_is_exclusive(x_34); -if (x_36 == 0) -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_37 = lean_ctor_get(x_34, 0); -x_38 = lean_unsigned_to_nat(3u); -x_39 = l_Lean_Syntax_getArg(x_4, x_38); -lean_dec(x_4); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_37); -lean_ctor_set(x_40, 1, x_39); -lean_ctor_set(x_34, 0, x_40); -return x_34; -} -else -{ -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_41 = lean_ctor_get(x_34, 0); -lean_inc(x_41); -lean_dec(x_34); -x_42 = lean_unsigned_to_nat(3u); -x_43 = l_Lean_Syntax_getArg(x_4, x_42); -lean_dec(x_4); -x_44 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_44, 0, x_41); -lean_ctor_set(x_44, 1, x_43); -x_45 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_45, 0, x_44); -return x_45; -} -} -} -else -{ -size_t x_46; size_t x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_46 = 0; -x_47 = lean_usize_of_nat(x_17); -lean_dec(x_17); -x_48 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(x_16, x_46, x_47, x_3); -lean_dec(x_16); -x_49 = lean_ctor_get(x_48, 1); -lean_inc(x_49); -lean_dec(x_48); -x_50 = l_Lean_Elab_Term_Quotation_match__syntax_expand___lambda__3___closed__1; -x_51 = l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__1(x_49, x_50); -lean_dec(x_49); -if (lean_obj_tag(x_51) == 0) -{ -lean_object* x_52; -lean_dec(x_4); -x_52 = lean_box(0); -return x_52; -} -else -{ -uint8_t x_53; -x_53 = !lean_is_exclusive(x_51); -if (x_53 == 0) -{ -lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_54 = lean_ctor_get(x_51, 0); -x_55 = lean_unsigned_to_nat(3u); -x_56 = l_Lean_Syntax_getArg(x_4, x_55); -lean_dec(x_4); -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_54); -lean_ctor_set(x_57, 1, x_56); -lean_ctor_set(x_51, 0, x_57); -return x_51; -} -else -{ -lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_58 = lean_ctor_get(x_51, 0); -lean_inc(x_58); -lean_dec(x_51); -x_59 = lean_unsigned_to_nat(3u); -x_60 = l_Lean_Syntax_getArg(x_4, x_59); -lean_dec(x_4); -x_61 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_61, 0, x_58); -lean_ctor_set(x_61, 1, x_60); -x_62 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_62, 0, x_61); -return x_62; -} -} -} -} } } } @@ -44356,7 +46035,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_match__syntax_expand(lean_ob _start: { lean_object* x_9; uint8_t x_10; -x_9 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__17; +x_9 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__17; lean_inc(x_1); x_10 = l_Lean_Syntax_isOfKind(x_1, x_9); if (x_10 == 0) @@ -44374,15 +46053,14 @@ return x_11; } else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; x_12 = lean_unsigned_to_nat(1u); x_13 = l_Lean_Syntax_getArg(x_1, x_12); -x_14 = l_Lean_nullKind; -x_15 = lean_unsigned_to_nat(0u); -x_16 = l_Lean_Syntax_isNodeOf(x_13, x_14, x_15); -if (x_16 == 0) +x_14 = lean_unsigned_to_nat(0u); +x_15 = l_Lean_Syntax_matchesNull(x_13, x_14); +if (x_15 == 0) { -lean_object* x_17; +lean_object* x_16; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); @@ -44390,18 +46068,18 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_17 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13___rarg(x_8); -return x_17; +x_16 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13___rarg(x_8); +return x_16; } else { -lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_18 = lean_unsigned_to_nat(2u); -x_19 = l_Lean_Syntax_getArg(x_1, x_18); -x_20 = l_Lean_Syntax_isNodeOf(x_19, x_14, x_15); -if (x_20 == 0) +lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_17 = lean_unsigned_to_nat(2u); +x_18 = l_Lean_Syntax_getArg(x_1, x_17); +x_19 = l_Lean_Syntax_matchesNull(x_18, x_14); +if (x_19 == 0) { -lean_object* x_21; +lean_object* x_20; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); @@ -44409,67 +46087,67 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_21 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13___rarg(x_8); -return x_21; +x_20 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13___rarg(x_8); +return x_20; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_22 = lean_unsigned_to_nat(3u); -x_23 = l_Lean_Syntax_getArg(x_1, x_22); -x_24 = l_Lean_Syntax_getArgs(x_23); -lean_dec(x_23); -x_25 = lean_array_get_size(x_24); -x_26 = lean_nat_dec_lt(x_15, x_25); -x_27 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; -x_28 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_match__syntax_expand___lambda__1), 2, 1); -lean_closure_set(x_28, 0, x_27); -if (x_26 == 0) +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_21 = lean_unsigned_to_nat(3u); +x_22 = l_Lean_Syntax_getArg(x_1, x_21); +x_23 = l_Lean_Syntax_getArgs(x_22); +lean_dec(x_22); +x_24 = lean_array_get_size(x_23); +x_25 = lean_nat_dec_lt(x_14, x_24); +x_26 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__6; +x_27 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_match__syntax_expand___lambda__1), 2, 1); +lean_closure_set(x_27, 0, x_26); +if (x_25 == 0) { -lean_object* x_69; -lean_dec(x_25); +lean_object* x_68; lean_dec(x_24); -x_69 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; -x_29 = x_69; -goto block_68; +lean_dec(x_23); +x_68 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; +x_28 = x_68; +goto block_67; } else { -uint8_t x_70; -x_70 = lean_nat_dec_le(x_25, x_25); -if (x_70 == 0) +uint8_t x_69; +x_69 = lean_nat_dec_le(x_24, x_24); +if (x_69 == 0) { -lean_object* x_71; -lean_dec(x_25); +lean_object* x_70; lean_dec(x_24); -x_71 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; -x_29 = x_71; -goto block_68; +lean_dec(x_23); +x_70 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; +x_28 = x_70; +goto block_67; } else { -size_t x_72; size_t x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_72 = 0; -x_73 = lean_usize_of_nat(x_25); -lean_dec(x_25); -x_74 = l_Lean_Elab_Term_Quotation_match__syntax_expand___closed__1; -x_75 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(x_24, x_72, x_73, x_74); +size_t x_71; size_t x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_71 = 0; +x_72 = lean_usize_of_nat(x_24); lean_dec(x_24); -x_76 = lean_ctor_get(x_75, 1); -lean_inc(x_76); -lean_dec(x_75); -x_29 = x_76; -goto block_68; +x_73 = l_Lean_Elab_Term_Quotation_match__syntax_expand___closed__1; +x_74 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(x_23, x_71, x_72, x_73); +lean_dec(x_23); +x_75 = lean_ctor_get(x_74, 1); +lean_inc(x_75); +lean_dec(x_74); +x_28 = x_75; +goto block_67; } } -block_68: +block_67: { -lean_object* x_30; -x_30 = l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__1(x_29, x_28); -lean_dec(x_29); -if (lean_obj_tag(x_30) == 0) +lean_object* x_29; +x_29 = l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__1(x_28, x_27); +lean_dec(x_28); +if (lean_obj_tag(x_29) == 0) { -lean_object* x_31; +lean_object* x_30; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); @@ -44477,170 +46155,170 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_31 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13___rarg(x_8); -return x_31; +x_30 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13___rarg(x_8); +return x_30; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; -x_32 = lean_ctor_get(x_30, 0); -lean_inc(x_32); -lean_dec(x_30); -x_33 = lean_unsigned_to_nat(5u); -x_34 = l_Lean_Syntax_getArg(x_1, x_33); +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_31 = lean_ctor_get(x_29, 0); +lean_inc(x_31); +lean_dec(x_29); +x_32 = lean_unsigned_to_nat(5u); +x_33 = l_Lean_Syntax_getArg(x_1, x_32); lean_dec(x_1); -x_35 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__22; -lean_inc(x_34); -x_36 = l_Lean_Syntax_isOfKind(x_34, x_35); -if (x_36 == 0) +x_34 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__22; +lean_inc(x_33); +x_35 = l_Lean_Syntax_isOfKind(x_33, x_34); +if (x_35 == 0) { -lean_object* x_37; -lean_dec(x_34); -lean_dec(x_32); +lean_object* x_36; +lean_dec(x_33); +lean_dec(x_31); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_37 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13___rarg(x_8); -return x_37; +x_36 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13___rarg(x_8); +return x_36; } else { -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_38 = l_Lean_Syntax_getArg(x_34, x_15); -lean_dec(x_34); -x_39 = l_Lean_Syntax_getArgs(x_38); +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_37 = l_Lean_Syntax_getArg(x_33, x_14); +lean_dec(x_33); +x_38 = l_Lean_Syntax_getArgs(x_37); +lean_dec(x_37); +x_39 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; +x_40 = l_Lean_Elab_Term_Quotation_match__syntax_expand___closed__1; +x_41 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_match__syntax_expand___lambda__3), 4, 3); +lean_closure_set(x_41, 0, x_26); +lean_closure_set(x_41, 1, x_39); +lean_closure_set(x_41, 2, x_40); +x_42 = l_Array_sequenceMap___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__1(x_38, x_41); lean_dec(x_38); -x_40 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__15; -x_41 = l_Lean_Elab_Term_Quotation_match__syntax_expand___closed__1; -x_42 = lean_alloc_closure((void*)(l_Lean_Elab_Term_Quotation_match__syntax_expand___lambda__3___boxed), 4, 3); -lean_closure_set(x_42, 0, x_27); -lean_closure_set(x_42, 1, x_40); -lean_closure_set(x_42, 2, x_41); -x_43 = l_Array_sequenceMap___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__1(x_39, x_42); -lean_dec(x_39); -if (lean_obj_tag(x_43) == 0) +if (lean_obj_tag(x_42) == 0) { -lean_object* x_44; -lean_dec(x_32); +lean_object* x_43; +lean_dec(x_31); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_44 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13___rarg(x_8); -return x_44; +x_43 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13___rarg(x_8); +return x_43; } else { -lean_object* x_45; lean_object* x_46; size_t x_47; size_t x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_58; uint8_t x_59; -x_45 = lean_ctor_get(x_43, 0); -lean_inc(x_45); -lean_dec(x_43); -x_46 = lean_array_get_size(x_45); -x_47 = lean_usize_of_nat(x_46); -lean_dec(x_46); -x_48 = 0; -lean_inc(x_45); -x_49 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__3(x_47, x_48, x_45); -x_50 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__4(x_47, x_48, x_45); -x_58 = lean_array_get_size(x_50); -x_59 = lean_nat_dec_lt(x_15, x_58); -if (x_59 == 0) +lean_object* x_44; lean_object* x_45; size_t x_46; size_t x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_57; uint8_t x_58; +x_44 = lean_ctor_get(x_42, 0); +lean_inc(x_44); +lean_dec(x_42); +x_45 = lean_array_get_size(x_44); +x_46 = lean_usize_of_nat(x_45); +lean_dec(x_45); +x_47 = 0; +lean_inc(x_44); +x_48 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__3(x_46, x_47, x_44); +x_49 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__4(x_46, x_47, x_44); +x_57 = lean_array_get_size(x_49); +x_58 = lean_nat_dec_lt(x_14, x_57); +if (x_58 == 0) { -lean_object* x_60; -lean_dec(x_58); -lean_dec(x_50); +lean_object* x_59; +lean_dec(x_57); lean_dec(x_49); -lean_dec(x_32); +lean_dec(x_48); +lean_dec(x_31); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_60 = lean_box(0); -x_51 = x_60; -goto block_57; +x_59 = lean_box(0); +x_50 = x_59; +goto block_56; } else { -uint8_t x_61; -x_61 = lean_nat_dec_le(x_58, x_58); -if (x_61 == 0) +uint8_t x_60; +x_60 = lean_nat_dec_le(x_57, x_57); +if (x_60 == 0) { -lean_object* x_62; -lean_dec(x_58); -lean_dec(x_50); +lean_object* x_61; +lean_dec(x_57); lean_dec(x_49); -lean_dec(x_32); +lean_dec(x_48); +lean_dec(x_31); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_62 = lean_box(0); -x_51 = x_62; -goto block_57; +x_61 = lean_box(0); +x_50 = x_61; +goto block_56; } else { -size_t x_63; uint8_t x_64; -x_63 = lean_usize_of_nat(x_58); -lean_dec(x_58); -x_64 = l_Array_anyMUnsafe_any___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__8(x_27, x_48, x_50, x_48, x_63); -if (x_64 == 0) +size_t x_62; uint8_t x_63; +x_62 = lean_usize_of_nat(x_57); +lean_dec(x_57); +x_63 = l_Array_anyMUnsafe_any___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__8(x_26, x_47, x_49, x_47, x_62); +if (x_63 == 0) { -lean_object* x_65; -lean_dec(x_50); +lean_object* x_64; lean_dec(x_49); -lean_dec(x_32); +lean_dec(x_48); +lean_dec(x_31); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_65 = lean_box(0); -x_51 = x_65; -goto block_57; +x_64 = lean_box(0); +x_50 = x_64; +goto block_56; } else { -lean_object* x_66; lean_object* x_67; -x_66 = lean_box(0); -x_67 = l_Lean_Elab_Term_Quotation_match__syntax_expand___lambda__5(x_32, x_50, x_48, x_49, x_66, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_49); -return x_67; +lean_object* x_65; lean_object* x_66; +x_65 = lean_box(0); +x_66 = l_Lean_Elab_Term_Quotation_match__syntax_expand___lambda__5(x_31, x_49, x_47, x_48, x_65, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_48); +return x_66; } } } -block_57: +block_56: { -lean_object* x_52; uint8_t x_53; -lean_dec(x_51); -x_52 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__6___rarg(x_8); -x_53 = !lean_is_exclusive(x_52); -if (x_53 == 0) +lean_object* x_51; uint8_t x_52; +lean_dec(x_50); +x_51 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__6___rarg(x_8); +x_52 = !lean_is_exclusive(x_51); +if (x_52 == 0) { -return x_52; +return x_51; } else { -lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_54 = lean_ctor_get(x_52, 0); -x_55 = lean_ctor_get(x_52, 1); -lean_inc(x_55); +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_51, 0); +x_54 = lean_ctor_get(x_51, 1); lean_inc(x_54); -lean_dec(x_52); -x_56 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_56, 0, x_54); -lean_ctor_set(x_56, 1, x_55); -return x_56; +lean_inc(x_53); +lean_dec(x_51); +x_55 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +return x_55; } } } @@ -44750,15 +46428,6 @@ x_10 = lean_box(x_9); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_match__syntax_expand___lambda__3___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_Elab_Term_Quotation_match__syntax_expand___lambda__3(x_1, x_2, x_3, x_4); -lean_dec(x_2); -return x_5; -} -} LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_match__syntax_expand___lambda__4___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_object* x_9) { _start: { @@ -44815,7 +46484,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSynta _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__5; +x_1 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__5; x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -44833,8 +46502,8 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__12; -x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__17; +x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__12; +x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__17; x_4 = l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax___closed__2; x_5 = l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax___closed__3; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); @@ -44863,7 +46532,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSynta _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(573u); +x_1 = lean_unsigned_to_nat(576u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -44875,7 +46544,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSynta _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(574u); +x_1 = lean_unsigned_to_nat(577u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -44903,7 +46572,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSynta _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(573u); +x_1 = lean_unsigned_to_nat(576u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -44915,7 +46584,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSynta _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(573u); +x_1 = lean_unsigned_to_nat(576u); x_2 = lean_unsigned_to_nat(46u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -44961,7 +46630,7 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_13012_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_13744_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; @@ -45084,6 +46753,10 @@ l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerm lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__2___closed__3); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__2___closed__4 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__2___closed__4(); lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__2___closed__4); +l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__2___closed__5 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__2___closed__5(); +lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__2___closed__5); +l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__2___closed__6 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__2___closed__6(); +lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__2___closed__6); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepFromSplice___closed__1 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepFromSplice___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepFromSplice___closed__1); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepFromSplice___closed__2 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepFromSplice___closed__2(); @@ -45206,284 +46879,288 @@ l_Lean_Elab_Term_Quotation_ArrayStxBuilder_push___closed__3 = _init_l_Lean_Elab_ lean_mark_persistent(l_Lean_Elab_Term_Quotation_ArrayStxBuilder_push___closed__3); l_Lean_Elab_Term_Quotation_ArrayStxBuilder_push___closed__4 = _init_l_Lean_Elab_Term_Quotation_ArrayStxBuilder_push___closed__4(); lean_mark_persistent(l_Lean_Elab_Term_Quotation_ArrayStxBuilder_push___closed__4); -l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___closed__1 = _init_l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___closed__1(); -lean_mark_persistent(l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___closed__1); -l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___closed__2 = _init_l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___closed__2(); -lean_mark_persistent(l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___closed__2); -l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___closed__3 = _init_l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___closed__3(); -lean_mark_persistent(l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___closed__3); -l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___closed__4 = _init_l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___closed__4(); -lean_mark_persistent(l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___closed__4); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__1(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__1); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__1(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__1); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__2 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__2(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__2); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__3 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__3(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__3); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__4 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__4(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__4); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__5 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__5(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__5); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__6 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__6(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__6); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__7 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__7(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__7); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__8 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__8(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__8); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__9 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__9(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5___closed__9); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__1(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__1); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__2 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__2(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__2); -l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__1 = _init_l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__1(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__1); -l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__2 = _init_l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__2(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__2); -l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__3 = _init_l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__3(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__3); -l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__4 = _init_l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__4(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__8___closed__4); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__1 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__1(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__1); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__2 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__2(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__2); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__3 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__3(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__3); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__4 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__4(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__4); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__5 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__5(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__5); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__6 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__6(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__6); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__7 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__7(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__7); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__8 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__8(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__8); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__9 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__9(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__9); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__10 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__10(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__2___closed__10); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__1 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__1(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__1); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__2 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__2(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__2); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__3 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__3(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__3); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__4 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__4(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__4); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__5 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__5(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__5); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__6 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__6(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__6); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__7); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__8 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__8(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__8); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__9 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__9(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__9); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__10 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__10(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__10); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__11 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__11(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__11); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__12 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__12(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__12); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__13); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__14); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__15 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__15(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__15); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__16 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__16(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__16); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__17 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__17(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__17); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__18 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__18(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__18); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__19 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__19(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__19); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__20 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__20(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__20); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__21 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__21(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__21); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__22 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__22(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__22); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__23 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__23(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__23); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__24 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__24(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__24); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__25 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__25(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__25); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__26 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__26(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__26); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__27 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__27(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__27); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__28 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__28(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__28); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__29 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__29(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__29); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__30 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__30(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___lambda__3___closed__30); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__1 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__1(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__1); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__2 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__2(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__2); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__3 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__3(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__3); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__4 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__4(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__4); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__5 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__5(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__5); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__6 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__6(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__6); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__7 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__7(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__7); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__8 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__8(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__8); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__9 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__9(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__9); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__10 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__10(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__10); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__11 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__11(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__11); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__12 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__12(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__12); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__13 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__13(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__13); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__14 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__14(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__14); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__15 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__15(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__15); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__16 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__16(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__16); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__17 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__17(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__17); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__18 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__18(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__18); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__19 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__19(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__19); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__20 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__20(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__20); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__21 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__21(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__21); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__22 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__22(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__22); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__23 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__23(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__23); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__24 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__24(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__24); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__25 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__25(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__25); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__26 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__26(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__26); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__27 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__27(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__27); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__28 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__28(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__28); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__29 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__29(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__29); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__30 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__30(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__30); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__31 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__31(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__31); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__32 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__32(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__32); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__33 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__33(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__33); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__34 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__34(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__34); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__35 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__35(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__35); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__36 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__36(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__36); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__37 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__37(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__37); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__38 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__38(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__38); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__39 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__39(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__39); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__40 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__40(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__40); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__41 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__41(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__41); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__42 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__42(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__42); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__43 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__43(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__43); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__44 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__44(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__44); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__45 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__45(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__45); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__46 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__46(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__46); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__47 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__47(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__47); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__48 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__48(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__48); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__49 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__49(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__49); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__50 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__50(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__50); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__51 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__51(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__51); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__52 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__52(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__52); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__53 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__53(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__53); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__54 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__54(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__54); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__55 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__55(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__55); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__56 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__56(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__56); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__57 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__57(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__57); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__58 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__58(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__58); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__59 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__59(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__59); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__60 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__60(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__60); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__61 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__61(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__61); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__62 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__62(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__62); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__63 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__63(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__63); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__64 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__64(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__64); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__65 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__65(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__65); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__66 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__66(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__66); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__67 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__67(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__67); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__68 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__68(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__68); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__69 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__69(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__69); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__70 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__70(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__70); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__71 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__71(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__71); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__72 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__72(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__72); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__73 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__73(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__73); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__74 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__74(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__74); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__75 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__75(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__75); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__76 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__76(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__76); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__77 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__77(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__77); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__78 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__78(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__78); -l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__79 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__79(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__9___closed__79); +l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___rarg___closed__1 = _init_l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___rarg___closed__1(); +lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___rarg___closed__1); +l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___rarg___closed__2 = _init_l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___rarg___closed__2(); +lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3___rarg___closed__2); +l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__1 = _init_l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__1(); +lean_mark_persistent(l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__1); +l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__2 = _init_l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__2(); +lean_mark_persistent(l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__2); +l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__3 = _init_l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__3(); +lean_mark_persistent(l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__3); +l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__4 = _init_l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__4(); +lean_mark_persistent(l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__6___closed__4); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__12___closed__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__12___closed__1(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__12___closed__1); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__1(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__1); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__2 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__2(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__2); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__3 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__3(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__3); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__4 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__4(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__4); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__5 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__5(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__5); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__6 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__6(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__6); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__7 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__7(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__7); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__8 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__8(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__8); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__9 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__9(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__13___closed__9); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__14___closed__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__14___closed__1(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__14___closed__1); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__14___closed__2 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__14___closed__2(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__14___closed__2); +l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__1 = _init_l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__1(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__1); +l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2 = _init_l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__2); +l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3 = _init_l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__3); +l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4 = _init_l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__24___closed__4); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__1 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__1(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__1); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__2 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__2(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__2); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__3 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__3(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__3); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__4 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__4(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__4); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__5 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__5(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__5); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__6 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__6(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__6); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__7 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__7(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__7); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__8 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__8(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__8); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__9 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__9(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__9); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__10 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__10(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__2___closed__10); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__1 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__1(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__1); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__2 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__2(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__2); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__3 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__3(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__3); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__4 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__4(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__4); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__5 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__5(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__5); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__6); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__7); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__8 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__8(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__8); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__9 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__9(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__9); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__10 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__10(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__10); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__11 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__11(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__11); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__12 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__12(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__12); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__13); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__14); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__15 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__15(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__15); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__16 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__16(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__16); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__17 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__17(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__17); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__18 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__18(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__18); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__19 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__19(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__19); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__20 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__20(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__20); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__21 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__21(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__21); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__22 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__22(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__22); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__23 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__23(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__23); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__24 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__24(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__24); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__25 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__25(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__25); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__26 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__26(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__26); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__27 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__27(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__27); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__28 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__28(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__28); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__29 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__29(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__29); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__30 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__30(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___lambda__3___closed__30); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__1 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__1(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__1); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__2 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__2(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__2); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__3 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__3(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__3); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__4 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__4(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__4); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__5 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__5(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__5); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__6 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__6(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__6); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__7 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__7(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__7); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__8 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__8(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__8); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__9 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__9(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__9); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__10 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__10(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__10); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__11 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__11(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__11); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__12 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__12(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__12); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__13 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__13(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__13); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__14 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__14(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__14); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__15 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__15(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__15); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__16 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__16(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__16); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__17 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__17(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__17); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__18 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__18(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__18); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__19 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__19(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__19); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__20 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__20(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__20); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__21 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__21(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__21); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__22 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__22(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__22); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__23 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__23(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__23); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__24 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__24(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__24); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__25 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__25(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__25); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__26 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__26(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__26); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__27 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__27(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__27); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__28 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__28(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__28); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__29 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__29(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__29); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__30 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__30(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__30); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__31 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__31(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__31); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__32 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__32(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__32); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__33 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__33(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__33); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__34 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__34(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__34); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__35 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__35(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__35); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__36 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__36(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__36); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__37 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__37(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__37); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__38 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__38(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__38); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__39 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__39(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__39); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__40 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__40(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__40); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__41 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__41(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__41); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__42 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__42(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__42); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__43 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__43(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__43); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__44 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__44(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__44); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__45 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__45(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__45); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__46 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__46(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__46); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__47 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__47(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__47); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__48 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__48(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__48); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__49 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__49(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__49); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__50 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__50(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__50); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__51 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__51(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__51); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__52 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__52(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__52); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__53 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__53(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__53); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__54 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__54(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__54); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__55 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__55(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__55); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__56 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__56(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__56); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__57 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__57(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__57); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__58 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__58(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__58); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__59 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__59(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__59); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__60 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__60(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__60); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__61 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__61(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__61); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__62 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__62(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__62); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__63 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__63(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__63); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__64 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__64(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__64); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__65 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__65(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__65); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__66 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__66(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__66); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__67 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__67(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__67); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__68 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__68(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__68); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__69 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__69(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__69); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__70 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__70(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__70); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__71 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__71(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__71); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__72 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__72(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__72); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__73 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__73(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__73); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__74 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__74(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__74); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__75 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__75(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__75); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__76 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__76(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__76); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__77 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__77(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__77); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__78 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__78(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__78); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__79 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__79(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__25___closed__79); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__1 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__1); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__2 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__2(); @@ -45646,10 +47323,6 @@ l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__48); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__49 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__49(); lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__49); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__50 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__50(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__50); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__51 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__51(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__51); l_Lean_throwUnknownConstant___at_Lean_Elab_Term_Quotation_getQuotKind___spec__5___closed__1 = _init_l_Lean_throwUnknownConstant___at_Lean_Elab_Term_Quotation_getQuotKind___spec__5___closed__1(); lean_mark_persistent(l_Lean_throwUnknownConstant___at_Lean_Elab_Term_Quotation_getQuotKind___spec__5___closed__1); l_Lean_throwUnknownConstant___at_Lean_Elab_Term_Quotation_getQuotKind___spec__5___closed__2 = _init_l_Lean_throwUnknownConstant___at_Lean_Elab_Term_Quotation_getQuotKind___spec__5___closed__2(); @@ -45826,490 +47499,490 @@ l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__51 = _init_l_Lean_Elab_Term_ lean_mark_persistent(l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__51); l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__52 = _init_l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__52(); lean_mark_persistent(l_Lean_Elab_Term_Quotation_stxQuot_expand___closed__52); -l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__1 = _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__1); -l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__2 = _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__2); -l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__3 = _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__3(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__3); -l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__4 = _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__4(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__4); -l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__5 = _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__5(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__5); -l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__6 = _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__6(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__6); -l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__7 = _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__7(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__7); -l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__8 = _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__8(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__8); -l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__9 = _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__9(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__9); -l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__10 = _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__10(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__10); -l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__11 = _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__11(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__11); -l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__12 = _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__12(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__12); -l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__13 = _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__13(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__13); -l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__14 = _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__14(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__14); -l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__15 = _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__15(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation_commandElab__stx__quot_______closed__15); -l_Lean_Elab_Term_Quotation_commandElab__stx__quot____ = _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot____(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation_commandElab__stx__quot____); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__1 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__1); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__2 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__2); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__3 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__3(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__3); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__4 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__4(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__4); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__5 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__5(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__5); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__6 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__6(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__6); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__7 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__7(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__7); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__8 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__8(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__8); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__9 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__9(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__9); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__10 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__10(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__10); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__11 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__11(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__11); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__12 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__12(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__12); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__13 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__13(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__13); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__14 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__14(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__14); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__15 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__15(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__15); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__16 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__16(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__16); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__17 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__17(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__17); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__18 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__18(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__18); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__19 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__19(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__19); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__20 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__20(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__20); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__21 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__21(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__21); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__22 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__22(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__22); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__23 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__23(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__23); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__24 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__24(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__24); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__25 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__25(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__25); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__26 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__26(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__26); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__27 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__27(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__27); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__28 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__28(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__28); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__29 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__29(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__29); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__30 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__30(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__30); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__31 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__31(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__31); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__32 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__32(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__32); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__33 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__33(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__33); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__34 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__34(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__34); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__35 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__35(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__35); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__36 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__36(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__36); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__37 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__37(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__37); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__38 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__38(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__38); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__39 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__39(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__39); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__40 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__40(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__40); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__41 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__41(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__41); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__42 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__42(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__42); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__43 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__43(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__43); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__44 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__44(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__44); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__45 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__45(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__45); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__46 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__46(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__46); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__47 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__47(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__47); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__48 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__48(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__48); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__49 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__49(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__49); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__50 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__50(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__50); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__51 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__51(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__51); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__52 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__52(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__52); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__53 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__53(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__53); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__54 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__54(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__54); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__55 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__55(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__55); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__56 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__56(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__56); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__57 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__57(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__57); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__58 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__58(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__58); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__59 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__59(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__59); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__60 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__60(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__60); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__61 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__61(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__61); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__62 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__62(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__62); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__63 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__63(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__63); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__64 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__64(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__64); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__65 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__65(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__65); -l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__66 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__66(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot______1___closed__66); -l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__1 = _init_l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__4); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__5 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__5(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__5); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__6 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__6(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__6); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__7 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__7(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__7); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__8 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__8(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__8); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__9 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__9(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__9); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__10 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__10(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__10); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__11 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__11(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__11); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__12 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__12(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__12); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__13 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__13(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633____closed__13); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633_(lean_io_mk_world()); +l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__1 = _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__1); +l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__2 = _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__2); +l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__3 = _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__3); +l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__4 = _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__4); +l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__5 = _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__5); +l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__6 = _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__6(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__6); +l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__7 = _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__7(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__7); +l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__8 = _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__8(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__8); +l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__9 = _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__9(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__9); +l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__10 = _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__10(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__10); +l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__11 = _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__11(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__11); +l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__12 = _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__12(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__12); +l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__13 = _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__13(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__13); +l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__14 = _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__14(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__14); +l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__15 = _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__15(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__15); +l_Lean_Elab_Term_Quotation_commandElab__stx__quot__ = _init_l_Lean_Elab_Term_Quotation_commandElab__stx__quot__(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_commandElab__stx__quot__); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__1 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__1); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__2 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__2); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__3 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__3); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__4 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__4); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__5 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__5); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__6 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__6(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__6); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__7 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__7(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__7); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__8 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__8(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__8); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__9 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__9(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__9); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__10 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__10(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__10); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__11 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__11(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__11); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__12 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__12(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__12); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__13 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__13(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__13); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__14 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__14(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__14); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__15 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__15(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__15); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__16 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__16(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__16); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__17 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__17(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__17); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__18 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__18(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__18); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__19 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__19(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__19); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__20 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__20(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__20); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__21 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__21(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__21); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__22 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__22(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__22); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__23 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__23(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__23); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__24 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__24(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__24); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__25 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__25(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__25); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__26 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__26(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__26); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__27 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__27(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__27); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__28 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__28(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__28); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__29 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__29(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__29); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__30 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__30(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__30); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__31 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__31(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__31); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__32 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__32(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__32); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__33 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__33(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__33); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__34 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__34(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__34); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__35 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__35(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__35); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__36 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__36(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__36); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__37 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__37(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__37); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__38 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__38(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__38); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__39 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__39(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__39); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__40 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__40(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__40); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__41 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__41(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__41); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__42 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__42(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__42); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__43 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__43(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__43); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__44 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__44(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__44); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__45 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__45(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__45); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__46 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__46(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__46); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__47 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__47(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__47); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__48 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__48(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__48); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__49 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__49(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__49); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__50 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__50(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__50); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__51 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__51(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__51); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__52 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__52(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__52); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__53 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__53(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__53); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__54 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__54(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__54); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__55 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__55(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__55); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__56 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__56(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__56); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__57 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__57(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__57); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__58 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__58(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__58); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__59 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__59(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__59); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__60 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__60(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__60); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__61 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__61(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__61); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__62 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__62(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__62); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__63 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__63(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__63); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__64 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__64(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__64); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__65 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__65(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__65); +l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__66 = _init_l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__66(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__66); +l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__1 = _init_l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__4); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__5 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__5(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__5); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__6 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__6(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__6); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__7 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__7(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__7); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__8 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__8(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__8); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__9 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__9(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__9); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__10 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__10(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__10); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__11 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__11(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__11); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__12 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__12(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__12); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__13 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__13(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874____closed__13); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633__declRange___closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633__declRange___closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633__declRange___closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633__declRange___closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633__declRange___closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633__declRange___closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633__declRange___closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633__declRange___closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4633__declRange(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874__declRange___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874__declRange___closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874__declRange___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874__declRange___closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874__declRange___closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874__declRange___closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874__declRange___closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874__declRange___closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4874__declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639____closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639____closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639____closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639____closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639____closed__3); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639_(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880____closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880____closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880____closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880____closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880____closed__3); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639__declRange___closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639__declRange___closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639__declRange___closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639__declRange___closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639__declRange___closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639__declRange___closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639__declRange___closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639__declRange___closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4639__declRange(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880__declRange___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880__declRange___closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880__declRange___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880__declRange___closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880__declRange___closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880__declRange___closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880__declRange___closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880__declRange___closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4880__declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645____closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645____closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645____closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645____closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645____closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645____closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645____closed__4); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645____closed__5 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645____closed__5(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645____closed__5); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645_(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886____closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886____closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886____closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886____closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886____closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886____closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886____closed__4); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886____closed__5 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886____closed__5(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886____closed__5); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645__declRange___closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645__declRange___closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645__declRange___closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645__declRange___closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645__declRange___closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645__declRange___closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645__declRange___closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645__declRange___closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4645__declRange(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886__declRange___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886__declRange___closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886__declRange___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886__declRange___closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886__declRange___closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886__declRange___closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886__declRange___closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886__declRange___closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4886__declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651____closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651____closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651____closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651____closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651____closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651____closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651____closed__4); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651____closed__5 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651____closed__5(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651____closed__5); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651_(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892____closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892____closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892____closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892____closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892____closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892____closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892____closed__4); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892____closed__5 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892____closed__5(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892____closed__5); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651__declRange___closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651__declRange___closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651__declRange___closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651__declRange___closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651__declRange___closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651__declRange___closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651__declRange___closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651__declRange___closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4651__declRange(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892__declRange___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892__declRange___closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892__declRange___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892__declRange___closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892__declRange___closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892__declRange___closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892__declRange___closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892__declRange___closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4892__declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657____closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657____closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657____closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657____closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657____closed__3); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657_(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898____closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898____closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898____closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898____closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898____closed__3); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657__declRange___closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657__declRange___closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657__declRange___closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657__declRange___closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657__declRange___closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657__declRange___closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657__declRange___closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657__declRange___closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4657__declRange(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898__declRange___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898__declRange___closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898__declRange___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898__declRange___closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898__declRange___closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898__declRange___closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898__declRange___closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898__declRange___closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4898__declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663____closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663____closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663____closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663____closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663____closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663____closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663____closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663_(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904____closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904____closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904____closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904____closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904____closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904____closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904____closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663__declRange___closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663__declRange___closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663__declRange___closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663__declRange___closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663__declRange___closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663__declRange___closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663__declRange___closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663__declRange___closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4663__declRange(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904__declRange___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904__declRange___closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904__declRange___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904__declRange___closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904__declRange___closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904__declRange___closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904__declRange___closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904__declRange___closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4904__declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669____closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669____closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669____closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669____closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669____closed__3); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669_(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910____closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910____closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910____closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910____closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910____closed__3); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669__declRange___closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669__declRange___closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669__declRange___closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669__declRange___closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669__declRange___closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669__declRange___closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669__declRange___closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669__declRange___closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4669__declRange(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910__declRange___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910__declRange___closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910__declRange___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910__declRange___closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910__declRange___closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910__declRange___closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910__declRange___closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910__declRange___closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4910__declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675____closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675____closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675____closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675____closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675____closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675____closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675____closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675_(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916____closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916____closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916____closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916____closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916____closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916____closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916____closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675__declRange___closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675__declRange___closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675__declRange___closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675__declRange___closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675__declRange___closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675__declRange___closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675__declRange___closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675__declRange___closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4675__declRange(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916__declRange___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916__declRange___closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916__declRange___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916__declRange___closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916__declRange___closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916__declRange___closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916__declRange___closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916__declRange___closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4916__declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681____closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681____closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681____closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681____closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681____closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681____closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681____closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681_(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922____closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922____closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922____closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922____closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922____closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922____closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922____closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681__declRange___closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681__declRange___closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681__declRange___closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681__declRange___closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681__declRange___closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681__declRange___closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681__declRange___closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681__declRange___closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4681__declRange(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922__declRange___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922__declRange___closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922__declRange___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922__declRange___closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922__declRange___closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922__declRange___closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922__declRange___closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922__declRange___closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4922__declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687____closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687____closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687____closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687____closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687____closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687____closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687____closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687_(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928____closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928____closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928____closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928____closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928____closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928____closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928____closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687__declRange___closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687__declRange___closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687__declRange___closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687__declRange___closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687__declRange___closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687__declRange___closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687__declRange___closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687__declRange___closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4687__declRange(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928__declRange___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928__declRange___closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928__declRange___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928__declRange___closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928__declRange___closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928__declRange___closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928__declRange___closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928__declRange___closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4928__declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693____closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693____closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693____closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693____closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693____closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693____closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693____closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693_(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934____closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934____closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934____closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934____closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934____closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934____closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934____closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693__declRange___closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693__declRange___closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693__declRange___closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693__declRange___closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693__declRange___closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693__declRange___closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693__declRange___closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693__declRange___closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4693__declRange(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934__declRange___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934__declRange___closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934__declRange___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934__declRange___closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934__declRange___closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934__declRange___closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934__declRange___closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934__declRange___closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4934__declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699____closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699____closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699____closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699____closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699____closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699____closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699____closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699_(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940____closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940____closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940____closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940____closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940____closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940____closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940____closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940____closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699__declRange___closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699__declRange___closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699__declRange___closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699__declRange___closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699__declRange___closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699__declRange___closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699__declRange___closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699__declRange___closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4699__declRange(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940__declRange___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940__declRange___closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940__declRange___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940__declRange___closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940__declRange___closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940__declRange___closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940__declRange___closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940__declRange___closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4940__declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705____closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705____closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705____closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705____closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705____closed__3); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705_(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946____closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946____closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946____closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946____closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946____closed__3); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705__declRange___closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705__declRange___closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705__declRange___closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705__declRange___closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705__declRange___closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705__declRange___closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705__declRange___closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705__declRange___closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4705__declRange(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946__declRange___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946__declRange___closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946__declRange___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946__declRange___closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946__declRange___closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946__declRange___closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946__declRange___closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946__declRange___closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4946__declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711____closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711____closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711____closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711____closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711____closed__3); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711_(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952____closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952____closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952____closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952____closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952____closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952____closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952____closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952____closed__3); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711__declRange___closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711__declRange___closed__1); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711__declRange___closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711__declRange___closed__2); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711__declRange___closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711__declRange___closed__3); -l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711__declRange___closed__4(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711__declRange___closed__4); -res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4711__declRange(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952__declRange___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952__declRange___closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952__declRange___closed__1); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952__declRange___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952__declRange___closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952__declRange___closed__2); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952__declRange___closed__3 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952__declRange___closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952__declRange___closed__3); +l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952__declRange___closed__4 = _init_l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952__declRange___closed__4(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952__declRange___closed__4); +res = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_4952__declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_noOpMatchAdaptPats___lambda__1___closed__1 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_noOpMatchAdaptPats___lambda__1___closed__1(); @@ -46404,18 +48077,6 @@ l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__14); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__15 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__15(); lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__15); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__16 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__16(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__16); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__17 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__17(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__17); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__18 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__18(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__18); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__19 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__19(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__19); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__20 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__20(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__20); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__21 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__21(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__8___closed__21); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__1 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__1); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__2 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__2(); @@ -46464,22 +48125,6 @@ l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__23); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__24 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__24(); lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__24); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__25 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__25(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__25); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__26 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__26(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__26); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__27 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__27(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__27); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__28 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__28(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__28); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__29 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__29(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__29); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__30 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__30(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__30); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__31 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__31(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__31); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__32 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__32(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__32); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__12___closed__1 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__12___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__12___closed__1); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__1 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__1(); @@ -46528,10 +48173,6 @@ l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__22); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__23 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__23(); lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__23); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__24 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__24(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__24); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__25 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__25(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__25); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16___closed__1 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16___closed__1); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16___closed__2 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__16___closed__2(); @@ -46602,14 +48243,6 @@ l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__27); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__28 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__28(); lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__28); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__29 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__29(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__29); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__30 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__30(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__30); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__31 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__31(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__31); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__32 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__32(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__32); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__1 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__1); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__2 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__2(); @@ -46642,36 +48275,28 @@ l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__15); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__16 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__16(); lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__24___closed__16); +l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__1 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__1(); +lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__1); +l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__2 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__2(); +lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__2); +l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__3 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__3(); +lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__3); +l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__4 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__4(); +lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__4); +l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__5 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__5(); +lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__5); +l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__6 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__6(); +lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__6); +l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__7 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__7(); +lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__7); +l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__8 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__8(); +lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__8); +l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__9 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__9(); +lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__26___closed__9); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__1 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__1); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__2 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__2(); lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__2); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__3 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__3(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__3); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__4 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__4(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__4); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__5 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__5(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__5); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__6 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__6(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__6); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__7 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__7(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__7); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__8 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__8(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__8); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__9 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__9(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__9); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__10 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__10(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__10); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__11 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__11(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__11); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__12 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__12(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__12); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__13 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__13(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__13); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__28___closed__1 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__28___closed__1(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__28___closed__1); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__28___closed__2 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__28___closed__2(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__28___closed__2); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__1 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__1); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__2 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__2(); @@ -46688,8 +48313,6 @@ l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__7); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__8 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__8(); lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__8); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__9 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__9(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__9); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___lambda__1___closed__1 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___lambda__1___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___lambda__1___closed__1); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___lambda__1___closed__2 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___lambda__1___closed__2(); @@ -46732,24 +48355,24 @@ l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_c lean_mark_persistent(l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__11___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__12___closed__2); l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__11___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__12___closed__3 = _init_l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__11___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__12___closed__3(); lean_mark_persistent(l_List_mapTRAux___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__11___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__12___closed__3); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__1 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__1(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__1); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__2 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__2(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__2); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__3 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__3(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__3); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__4 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__4(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__4); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__5 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__5(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__5); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__6 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__6(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__6); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__7 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__7(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__7); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__8 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__8(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__8); -l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__9 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__9(); -lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4___closed__9); +l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__1 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__1(); +lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__1); +l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__2 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__2(); +lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__2); +l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__3 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__3(); +lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__3); +l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__4 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__4(); +lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__4); +l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__5 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__5(); +lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__5); +l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__6 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__6(); +lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__6); +l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__7 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__7(); +lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__7); +l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__8 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__8(); +lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__8); +l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__9 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__9(); +lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__9); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___closed__1 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___closed__1); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___closed__2 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___closed__2(); @@ -46764,10 +48387,6 @@ l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___cl lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___closed__6); l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___closed__7 = _init_l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___closed__7(); lean_mark_persistent(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___closed__7); -l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__6___rarg___closed__1 = _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__6___rarg___closed__1(); -lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__6___rarg___closed__1); -l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__6___rarg___closed__2 = _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__6___rarg___closed__2(); -lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__6___rarg___closed__2); l_Lean_Elab_Term_Quotation_match__syntax_expand___lambda__3___closed__1 = _init_l_Lean_Elab_Term_Quotation_match__syntax_expand___lambda__3___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_Quotation_match__syntax_expand___lambda__3___closed__1); l_Lean_Elab_Term_Quotation_match__syntax_expand___lambda__5___closed__1 = _init_l_Lean_Elab_Term_Quotation_match__syntax_expand___lambda__5___closed__1(); @@ -46809,7 +48428,7 @@ lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax_dec res = l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax_declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -res = l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_13012_(lean_io_mk_world()); +res = l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_13744_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/Quotation/Precheck.c b/stage0/stdlib/Lean/Elab/Quotation/Precheck.c index ca863c44d13..fceee12ed05 100644 --- a/stage0/stdlib/Lean/Elab/Quotation/Precheck.c +++ b/stage0/stdlib/Lean/Elab/Quotation/Precheck.c @@ -39,7 +39,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_precheck_hasQuotedIdent___bo static lean_object* l_Lean_Elab_Term_Quotation_mkPrecheckAttribute___closed__17; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_Quotation_precheck___spec__1___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_precheckChoice___closed__4; -extern lean_object* l_Lean_nullKind; static lean_object* l_Lean_addTrace___at_Lean_Elab_Term_Quotation_precheck___spec__3___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabPrecheckedQuot_declRange___closed__3; static lean_object* l_Lean_Elab_Term_Quotation_mkPrecheckAttribute___closed__8; @@ -100,7 +99,6 @@ LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Term_Quotation_precheck___sp static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabPrecheckedQuot___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_precheckChoice___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_precheck(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Syntax_SepArray_getElems___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_withNewLocal(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_precheck___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); @@ -111,6 +109,7 @@ static lean_object* l_Lean_addTrace___at_Lean_Elab_Term_Quotation_precheck___spe static lean_object* l_List_forIn_loop___at_Lean_Elab_Term_Quotation_precheckIdent___spec__3___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_quotPrecheck; static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_precheckIdent___closed__5; +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); lean_object* l_Lean_Elab_expandMacroImpl_x3f(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Quotation_precheckChoice___spec__5___closed__1; @@ -180,7 +179,7 @@ extern lean_object* l_Lean_Elab_Term_termElabAttribute; lean_object* l_Lean_Macro_expandMacro_x3f(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_Quotation_precheck___spec__7___closed__1; LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_Quotation_precheck___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_TSepArray_getElems___rarg(lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_mkPrecheckAttribute___closed__10; static lean_object* l_List_forIn_loop___at_Lean_Elab_Term_Quotation_precheckIdent___spec__3___closed__2; lean_object* l_Lean_Syntax_getQuotContent(lean_object*); @@ -3991,23 +3990,22 @@ return x_12; } else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; x_13 = lean_unsigned_to_nat(1u); x_14 = l_Lean_Syntax_getArg(x_1, x_13); lean_dec(x_1); -x_15 = l_Lean_nullKind; -x_16 = lean_unsigned_to_nat(0u); +x_15 = lean_unsigned_to_nat(0u); lean_inc(x_14); -x_17 = l_Lean_Syntax_isNodeOf(x_14, x_15, x_16); -if (x_17 == 0) +x_16 = l_Lean_Syntax_matchesNull(x_14, x_15); +if (x_16 == 0) { -lean_object* x_18; uint8_t x_19; -x_18 = lean_unsigned_to_nat(2u); +lean_object* x_17; uint8_t x_18; +x_17 = lean_unsigned_to_nat(2u); lean_inc(x_14); -x_19 = l_Lean_Syntax_isNodeOf(x_14, x_15, x_18); -if (x_19 == 0) +x_18 = l_Lean_Syntax_matchesNull(x_14, x_17); +if (x_18 == 0) { -lean_object* x_20; +lean_object* x_19; lean_dec(x_14); lean_dec(x_8); lean_dec(x_7); @@ -4016,25 +4014,25 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_20 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_precheckIdent___spec__1___rarg(x_9); -return x_20; +x_19 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_precheckIdent___spec__1___rarg(x_9); +return x_19; } else { -lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_21 = l_Lean_Syntax_getArg(x_14, x_16); -x_22 = l_Lean_Syntax_getArg(x_14, x_13); +lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_20 = l_Lean_Syntax_getArg(x_14, x_15); +x_21 = l_Lean_Syntax_getArg(x_14, x_13); lean_dec(x_14); -lean_inc(x_22); -x_23 = l_Lean_Syntax_isNodeOf(x_22, x_15, x_13); -if (x_23 == 0) +lean_inc(x_21); +x_22 = l_Lean_Syntax_matchesNull(x_21, x_13); +if (x_22 == 0) { -uint8_t x_24; -x_24 = l_Lean_Syntax_isNodeOf(x_22, x_15, x_16); -if (x_24 == 0) +uint8_t x_23; +x_23 = l_Lean_Syntax_matchesNull(x_21, x_15); +if (x_23 == 0) { -lean_object* x_25; -lean_dec(x_21); +lean_object* x_24; +lean_dec(x_20); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -4042,35 +4040,35 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_25 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_precheckIdent___spec__1___rarg(x_9); -return x_25; +x_24 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_precheckIdent___spec__1___rarg(x_9); +return x_24; } else { -lean_object* x_26; -x_26 = l_Lean_Elab_Term_Quotation_precheck(x_21, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_26; +lean_object* x_25; +x_25 = l_Lean_Elab_Term_Quotation_precheck(x_20, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_25; } } else { -lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_27 = l_Lean_Syntax_getArg(x_22, x_16); -lean_dec(x_22); -x_28 = l_Lean_Elab_Term_Quotation_precheckParen___closed__4; -lean_inc(x_27); -x_29 = l_Lean_Syntax_isOfKind(x_27, x_28); -if (x_29 == 0) +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = l_Lean_Syntax_getArg(x_21, x_15); +lean_dec(x_21); +x_27 = l_Lean_Elab_Term_Quotation_precheckParen___closed__4; +lean_inc(x_26); +x_28 = l_Lean_Syntax_isOfKind(x_26, x_27); +if (x_28 == 0) { -lean_object* x_30; uint8_t x_31; -x_30 = l_Lean_Elab_Term_Quotation_precheckParen___closed__6; -lean_inc(x_27); -x_31 = l_Lean_Syntax_isOfKind(x_27, x_30); -if (x_31 == 0) +lean_object* x_29; uint8_t x_30; +x_29 = l_Lean_Elab_Term_Quotation_precheckParen___closed__6; +lean_inc(x_26); +x_30 = l_Lean_Syntax_isOfKind(x_26, x_29); +if (x_30 == 0) { -lean_object* x_32; -lean_dec(x_27); -lean_dec(x_21); +lean_object* x_31; +lean_dec(x_26); +lean_dec(x_20); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -4078,16 +4076,16 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_32 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_precheckIdent___spec__1___rarg(x_9); -return x_32; +x_31 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_precheckIdent___spec__1___rarg(x_9); +return x_31; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = l_Lean_Syntax_getArg(x_27, x_13); -lean_dec(x_27); -x_34 = l_Lean_Syntax_getArgs(x_33); -lean_dec(x_33); +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = l_Lean_Syntax_getArg(x_26, x_13); +lean_dec(x_26); +x_33 = l_Lean_Syntax_getArgs(x_32); +lean_dec(x_32); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); @@ -4095,26 +4093,26 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_35 = l_Lean_Elab_Term_Quotation_precheck(x_21, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_35) == 0) +x_34 = l_Lean_Elab_Term_Quotation_precheck(x_20, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_34) == 0) { -uint8_t x_36; -x_36 = !lean_is_exclusive(x_35); -if (x_36 == 0) +uint8_t x_35; +x_35 = !lean_is_exclusive(x_34); +if (x_35 == 0) { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; -x_37 = lean_ctor_get(x_35, 1); -x_38 = lean_ctor_get(x_35, 0); -lean_dec(x_38); -x_39 = l_Lean_Syntax_SepArray_getElems___rarg(x_34); -lean_dec(x_34); -x_40 = lean_array_get_size(x_39); -x_41 = lean_nat_dec_lt(x_16, x_40); -if (x_41 == 0) +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; +x_36 = lean_ctor_get(x_34, 1); +x_37 = lean_ctor_get(x_34, 0); +lean_dec(x_37); +x_38 = l_Lean_Syntax_TSepArray_getElems___rarg(x_33); +lean_dec(x_33); +x_39 = lean_array_get_size(x_38); +x_40 = lean_nat_dec_lt(x_15, x_39); +if (x_40 == 0) { -lean_object* x_42; -lean_dec(x_40); +lean_object* x_41; lean_dec(x_39); +lean_dec(x_38); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -4122,19 +4120,19 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_42 = lean_box(0); -lean_ctor_set(x_35, 0, x_42); -return x_35; +x_41 = lean_box(0); +lean_ctor_set(x_34, 0, x_41); +return x_34; } else { -uint8_t x_43; -x_43 = lean_nat_dec_le(x_40, x_40); -if (x_43 == 0) +uint8_t x_42; +x_42 = lean_nat_dec_le(x_39, x_39); +if (x_42 == 0) { -lean_object* x_44; -lean_dec(x_40); +lean_object* x_43; lean_dec(x_39); +lean_dec(x_38); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -4142,39 +4140,39 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_44 = lean_box(0); -lean_ctor_set(x_35, 0, x_44); -return x_35; +x_43 = lean_box(0); +lean_ctor_set(x_34, 0, x_43); +return x_34; } else { -size_t x_45; size_t x_46; lean_object* x_47; lean_object* x_48; -lean_free_object(x_35); -x_45 = 0; -x_46 = lean_usize_of_nat(x_40); -lean_dec(x_40); -x_47 = lean_box(0); -x_48 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Quotation_precheckParen___spec__1(x_39, x_45, x_46, x_47, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_37); +size_t x_44; size_t x_45; lean_object* x_46; lean_object* x_47; +lean_free_object(x_34); +x_44 = 0; +x_45 = lean_usize_of_nat(x_39); lean_dec(x_39); -return x_48; +x_46 = lean_box(0); +x_47 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Quotation_precheckParen___spec__1(x_38, x_44, x_45, x_46, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_36); +lean_dec(x_38); +return x_47; } } } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; -x_49 = lean_ctor_get(x_35, 1); -lean_inc(x_49); -lean_dec(x_35); -x_50 = l_Lean_Syntax_SepArray_getElems___rarg(x_34); +lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; +x_48 = lean_ctor_get(x_34, 1); +lean_inc(x_48); lean_dec(x_34); -x_51 = lean_array_get_size(x_50); -x_52 = lean_nat_dec_lt(x_16, x_51); -if (x_52 == 0) +x_49 = l_Lean_Syntax_TSepArray_getElems___rarg(x_33); +lean_dec(x_33); +x_50 = lean_array_get_size(x_49); +x_51 = lean_nat_dec_lt(x_15, x_50); +if (x_51 == 0) { -lean_object* x_53; lean_object* x_54; -lean_dec(x_51); +lean_object* x_52; lean_object* x_53; lean_dec(x_50); +lean_dec(x_49); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -4182,21 +4180,21 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_53 = lean_box(0); -x_54 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_49); -return x_54; +x_52 = lean_box(0); +x_53 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_48); +return x_53; } else { -uint8_t x_55; -x_55 = lean_nat_dec_le(x_51, x_51); -if (x_55 == 0) +uint8_t x_54; +x_54 = lean_nat_dec_le(x_50, x_50); +if (x_54 == 0) { -lean_object* x_56; lean_object* x_57; -lean_dec(x_51); +lean_object* x_55; lean_object* x_56; lean_dec(x_50); +lean_dec(x_49); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -4204,30 +4202,30 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_56 = lean_box(0); -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_49); -return x_57; +x_55 = lean_box(0); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_48); +return x_56; } else { -size_t x_58; size_t x_59; lean_object* x_60; lean_object* x_61; -x_58 = 0; -x_59 = lean_usize_of_nat(x_51); -lean_dec(x_51); -x_60 = lean_box(0); -x_61 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Quotation_precheckParen___spec__1(x_50, x_58, x_59, x_60, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_49); +size_t x_57; size_t x_58; lean_object* x_59; lean_object* x_60; +x_57 = 0; +x_58 = lean_usize_of_nat(x_50); lean_dec(x_50); -return x_61; +x_59 = lean_box(0); +x_60 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Quotation_precheckParen___spec__1(x_49, x_57, x_58, x_59, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_48); +lean_dec(x_49); +return x_60; } } } } else { -uint8_t x_62; -lean_dec(x_34); +uint8_t x_61; +lean_dec(x_33); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -4235,32 +4233,32 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_62 = !lean_is_exclusive(x_35); -if (x_62 == 0) +x_61 = !lean_is_exclusive(x_34); +if (x_61 == 0) { -return x_35; +return x_34; } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_63 = lean_ctor_get(x_35, 0); -x_64 = lean_ctor_get(x_35, 1); -lean_inc(x_64); +lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_62 = lean_ctor_get(x_34, 0); +x_63 = lean_ctor_get(x_34, 1); lean_inc(x_63); -lean_dec(x_35); -x_65 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_65, 0, x_63); -lean_ctor_set(x_65, 1, x_64); -return x_65; +lean_inc(x_62); +lean_dec(x_34); +x_64 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_64, 0, x_62); +lean_ctor_set(x_64, 1, x_63); +return x_64; } } } } else { -lean_object* x_66; lean_object* x_67; -x_66 = l_Lean_Syntax_getArg(x_27, x_13); -lean_dec(x_27); +lean_object* x_65; lean_object* x_66; +x_65 = l_Lean_Syntax_getArg(x_26, x_13); +lean_dec(x_26); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); @@ -4268,20 +4266,20 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_67 = l_Lean_Elab_Term_Quotation_precheck(x_21, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_67) == 0) +x_66 = l_Lean_Elab_Term_Quotation_precheck(x_20, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_66) == 0) { -lean_object* x_68; lean_object* x_69; -x_68 = lean_ctor_get(x_67, 1); -lean_inc(x_68); -lean_dec(x_67); -x_69 = l_Lean_Elab_Term_Quotation_precheck(x_66, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_68); -return x_69; +lean_object* x_67; lean_object* x_68; +x_67 = lean_ctor_get(x_66, 1); +lean_inc(x_67); +lean_dec(x_66); +x_68 = l_Lean_Elab_Term_Quotation_precheck(x_65, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_67); +return x_68; } else { -uint8_t x_70; -lean_dec(x_66); +uint8_t x_69; +lean_dec(x_65); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -4289,23 +4287,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_70 = !lean_is_exclusive(x_67); -if (x_70 == 0) +x_69 = !lean_is_exclusive(x_66); +if (x_69 == 0) { -return x_67; +return x_66; } else { -lean_object* x_71; lean_object* x_72; lean_object* x_73; -x_71 = lean_ctor_get(x_67, 0); -x_72 = lean_ctor_get(x_67, 1); -lean_inc(x_72); +lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_70 = lean_ctor_get(x_66, 0); +x_71 = lean_ctor_get(x_66, 1); lean_inc(x_71); -lean_dec(x_67); -x_73 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_73, 0, x_71); -lean_ctor_set(x_73, 1, x_72); -return x_73; +lean_inc(x_70); +lean_dec(x_66); +x_72 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_72, 0, x_70); +lean_ctor_set(x_72, 1, x_71); +return x_72; } } } @@ -4314,7 +4312,7 @@ return x_73; } else { -lean_object* x_74; lean_object* x_75; +lean_object* x_73; lean_object* x_74; lean_dec(x_14); lean_dec(x_8); lean_dec(x_7); @@ -4323,11 +4321,11 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_74 = lean_box(0); -x_75 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_75, 0, x_74); -lean_ctor_set(x_75, 1, x_9); -return x_75; +x_73 = lean_box(0); +x_74 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_74, 0, x_73); +lean_ctor_set(x_74, 1, x_9); +return x_74; } } } diff --git a/stage0/stdlib/Lean/Elab/Quotation/Util.c b/stage0/stdlib/Lean/Elab/Quotation/Util.c index b9d798c35bf..9eefb9d7413 100644 --- a/stage0/stdlib/Lean/Elab/Quotation/Util.c +++ b/stage0/stdlib/Lean/Elab/Quotation/Util.c @@ -18,7 +18,6 @@ size_t lean_usize_add(size_t, size_t); uint8_t l_Lean_Syntax_isTokenAntiquot(lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); -extern lean_object* l_Lean_nullKind; lean_object* lean_name_mk_string(lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); lean_object* lean_array_uget(lean_object*, size_t); @@ -32,6 +31,7 @@ lean_object* lean_array_push(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); static lean_object* l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__7; lean_object* l_Lean_Syntax_getAntiquotTerm(lean_object*); +static lean_object* l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__12; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Quotation_getPatternsVars___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation_Util___hyg_7_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_getPatternsVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -42,6 +42,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_getPatternVars(lean_object*, static lean_object* l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__10; static lean_object* l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__11; static lean_object* l_Lean_Elab_Term_Quotation_getPatternVars___closed__5; +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_getPatternVars___closed__2; static lean_object* l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation_Util___hyg_7____closed__1; lean_object* l_Lean_replaceRef(lean_object*, lean_object*); @@ -52,7 +53,6 @@ static lean_object* l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Ter lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__4; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Quotation_getPatternsVars___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_choiceKind; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_getAntiquotationIds(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation_Util___hyg_7____closed__5; extern lean_object* l_Lean_instInhabitedSyntax; @@ -61,7 +61,6 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Ela size_t lean_usize_of_nat(lean_object*); uint8_t l_Lean_Syntax_isAntiquot(lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_getAntiquotationIds___closed__1; -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__8; static lean_object* l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__1; @@ -69,6 +68,7 @@ uint8_t lean_nat_dec_le(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_getPatternVars___closed__8; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__3(uint8_t, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_getPatternVars___closed__4; +static lean_object* l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__13; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_Quotation_getPatternVars___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -529,7 +529,7 @@ static lean_object* _init_l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_El _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Lean", 4); +x_1 = lean_mk_string_from_bytes("choice", 6); return x_1; } } @@ -547,7 +547,7 @@ static lean_object* _init_l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_El _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Parser", 6); +x_1 = lean_mk_string_from_bytes("Lean", 4); return x_1; } } @@ -555,7 +555,7 @@ static lean_object* _init_l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_El _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__3; +x_1 = lean_box(0); x_2 = l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -565,7 +565,7 @@ static lean_object* _init_l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_El _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Term", 4); +x_1 = lean_mk_string_from_bytes("Parser", 6); return x_1; } } @@ -583,7 +583,7 @@ static lean_object* _init_l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_El _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("hole", 4); +x_1 = lean_mk_string_from_bytes("Term", 4); return x_1; } } @@ -601,15 +601,33 @@ static lean_object* _init_l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_El _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("complex antiquotation not allowed here", 38); +x_1 = lean_mk_string_from_bytes("hole", 4); return x_1; } } static lean_object* _init_l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__11() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__9; +x_2 = l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__10; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("complex antiquotation not allowed here", 38); +return x_1; +} +} +static lean_object* _init_l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__13() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__10; +x_1 = l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__12; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } @@ -702,7 +720,7 @@ goto block_41; else { lean_object* x_43; uint8_t x_44; -x_43 = l_Lean_choiceKind; +x_43 = l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__3; x_44 = lean_name_eq(x_16, x_43); lean_dec(x_16); if (x_44 == 0) @@ -892,13 +910,13 @@ x_63 = l_Lean_Syntax_isIdent(x_62); if (x_63 == 0) { lean_object* x_64; uint8_t x_65; -x_64 = l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__9; +x_64 = l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__11; x_65 = l_Lean_Syntax_isOfKind(x_62, x_64); if (x_65 == 0) { lean_object* x_66; lean_object* x_67; uint8_t x_68; lean_dec(x_3); -x_66 = l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__11; +x_66 = l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__13; x_67 = l_Lean_throwErrorAt___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__1(x_2, x_66, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_8); @@ -1223,7 +1241,7 @@ static lean_object* _init_l_Lean_Elab_Term_Quotation_getPatternVars___closed__4( _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__7; +x_1 = l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__9; x_2 = l_Lean_Elab_Term_Quotation_getPatternVars___closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -1272,7 +1290,7 @@ x_9 = l_Lean_Syntax_isQuot(x_1); if (x_9 == 0) { lean_object* x_10; uint8_t x_11; -x_10 = l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__9; +x_10 = l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__11; lean_inc(x_1); x_11 = l_Lean_Syntax_isOfKind(x_1, x_10); if (x_11 == 0) @@ -1333,85 +1351,84 @@ return x_32; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +lean_object* x_33; lean_object* x_34; uint8_t x_35; x_33 = lean_unsigned_to_nat(2u); x_34 = l_Lean_Syntax_getArg(x_1, x_33); -x_35 = l_Lean_nullKind; -x_36 = l_Lean_Syntax_isNodeOf(x_34, x_35, x_23); -if (x_36 == 0) +x_35 = l_Lean_Syntax_matchesNull(x_34, x_23); +if (x_35 == 0) { -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; +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_dec(x_24); lean_inc(x_1); -x_37 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_37, 0, x_1); -x_38 = l_Lean_indentD(x_37); -x_39 = l_Lean_Elab_Term_Quotation_getPatternVars___closed__6; -x_40 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_38); -x_41 = l_Lean_Elab_Term_Quotation_getPatternVars___closed__7; -x_42 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_42, 0, x_40); -lean_ctor_set(x_42, 1, x_41); -x_43 = l_Lean_throwErrorAt___at_Lean_Elab_Term_Quotation_getPatternVars___spec__1(x_1, x_42, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_43; +x_36 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_36, 0, x_1); +x_37 = l_Lean_indentD(x_36); +x_38 = l_Lean_Elab_Term_Quotation_getPatternVars___closed__6; +x_39 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_37); +x_40 = l_Lean_Elab_Term_Quotation_getPatternVars___closed__7; +x_41 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +x_42 = l_Lean_throwErrorAt___at_Lean_Elab_Term_Quotation_getPatternVars___spec__1(x_1, x_41, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_42; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_unsigned_to_nat(3u); -x_45 = l_Lean_Syntax_getArg(x_1, x_44); +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_unsigned_to_nat(3u); +x_44 = l_Lean_Syntax_getArg(x_1, x_43); lean_dec(x_1); -x_46 = l_Lean_Elab_Term_Quotation_getPatternVars(x_45, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_46) == 0) +x_45 = l_Lean_Elab_Term_Quotation_getPatternVars(x_44, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_45) == 0) { -uint8_t x_47; -x_47 = !lean_is_exclusive(x_46); -if (x_47 == 0) +uint8_t x_46; +x_46 = !lean_is_exclusive(x_45); +if (x_46 == 0) { -lean_object* x_48; lean_object* x_49; -x_48 = lean_ctor_get(x_46, 0); -x_49 = lean_array_push(x_48, x_24); -lean_ctor_set(x_46, 0, x_49); -return x_46; +lean_object* x_47; lean_object* x_48; +x_47 = lean_ctor_get(x_45, 0); +x_48 = lean_array_push(x_47, x_24); +lean_ctor_set(x_45, 0, x_48); +return x_45; } else { -lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; -x_50 = lean_ctor_get(x_46, 0); -x_51 = lean_ctor_get(x_46, 1); -lean_inc(x_51); +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_49 = lean_ctor_get(x_45, 0); +x_50 = lean_ctor_get(x_45, 1); lean_inc(x_50); -lean_dec(x_46); -x_52 = lean_array_push(x_50, x_24); -x_53 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_53, 0, x_52); -lean_ctor_set(x_53, 1, x_51); -return x_53; +lean_inc(x_49); +lean_dec(x_45); +x_51 = lean_array_push(x_49, x_24); +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_50); +return x_52; } } else { -uint8_t x_54; +uint8_t x_53; lean_dec(x_24); -x_54 = !lean_is_exclusive(x_46); -if (x_54 == 0) +x_53 = !lean_is_exclusive(x_45); +if (x_53 == 0) { -return x_46; +return x_45; } else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_55 = lean_ctor_get(x_46, 0); -x_56 = lean_ctor_get(x_46, 1); -lean_inc(x_56); +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_ctor_get(x_45, 0); +x_55 = lean_ctor_get(x_45, 1); lean_inc(x_55); -lean_dec(x_46); -x_57 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_57, 0, x_55); -lean_ctor_set(x_57, 1, x_56); -return x_57; +lean_inc(x_54); +lean_dec(x_45); +x_56 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_56, 0, x_54); +lean_ctor_set(x_56, 1, x_55); +return x_56; } } } @@ -1420,24 +1437,24 @@ return x_57; } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_58 = l_Lean_Elab_Term_Quotation_getPatternVars___closed__8; -x_59 = lean_array_push(x_58, x_1); -x_60 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_60, 0, x_59); -lean_ctor_set(x_60, 1, x_8); -return x_60; +x_57 = l_Lean_Elab_Term_Quotation_getPatternVars___closed__8; +x_58 = lean_array_push(x_57, x_1); +x_59 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_8); +return x_59; } } else { -lean_object* x_61; lean_object* x_62; +lean_object* x_60; lean_object* x_61; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); @@ -1445,18 +1462,61 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_61 = l_Lean_Elab_Term_Quotation_getAntiquotationIds___closed__1; -x_62 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_62, 0, x_61); -lean_ctor_set(x_62, 1, x_8); +x_60 = l_Lean_Elab_Term_Quotation_getAntiquotationIds___closed__1; +x_61 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set(x_61, 1, x_8); +return x_61; +} +} +else +{ +lean_object* x_62; +x_62 = l_Lean_Elab_Term_Quotation_getAntiquotationIds(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_62) == 0) +{ +uint8_t x_63; +x_63 = !lean_is_exclusive(x_62); +if (x_63 == 0) +{ return x_62; } +else +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_62, 0); +x_65 = lean_ctor_get(x_62, 1); +lean_inc(x_65); +lean_inc(x_64); +lean_dec(x_62); +x_66 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +return x_66; +} } else { -lean_object* x_63; -x_63 = l_Lean_Elab_Term_Quotation_getAntiquotationIds(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_63; +uint8_t x_67; +x_67 = !lean_is_exclusive(x_62); +if (x_67 == 0) +{ +return x_62; +} +else +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_62, 0); +x_69 = lean_ctor_get(x_62, 1); +lean_inc(x_69); +lean_inc(x_68); +lean_dec(x_62); +x_70 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_70, 0, x_68); +lean_ctor_set(x_70, 1, x_69); +return x_70; +} +} } } } @@ -1679,6 +1739,10 @@ l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiq lean_mark_persistent(l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__10); l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__11 = _init_l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__11(); lean_mark_persistent(l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__11); +l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__12 = _init_l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__12(); +lean_mark_persistent(l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__12); +l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__13 = _init_l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__13(); +lean_mark_persistent(l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__2___closed__13); l_Lean_Elab_Term_Quotation_getAntiquotationIds___closed__1 = _init_l_Lean_Elab_Term_Quotation_getAntiquotationIds___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_Quotation_getAntiquotationIds___closed__1); l_Lean_Elab_Term_Quotation_getPatternVars___closed__1 = _init_l_Lean_Elab_Term_Quotation_getPatternVars___closed__1(); diff --git a/stage0/stdlib/Lean/Elab/SetOption.c b/stage0/stdlib/Lean/Elab/SetOption.c index 8c74671528c..c50dc87af0a 100644 --- a/stage0/stdlib/Lean/Elab/SetOption.c +++ b/stage0/stdlib/Lean/Elab/SetOption.c @@ -17,7 +17,6 @@ lean_object* lean_erase_macro_scopes(lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_elabSetOption_setOption___spec__1___rarg___lambda__1(lean_object*, lean_object*); lean_object* lean_io_error_to_string(lean_object*); -lean_object* l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_elabSetOption___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabSetOption_setOption___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabSetOption___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -26,8 +25,8 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_elabSetOption___spec__ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_elabSetOption_setOption___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabSetOption_setOption___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabSetOption___rarg___lambda__1___closed__5; +lean_object* l_Lean_Syntax_isNatLit_x3f(lean_object*); lean_object* l_Lean_getOptionDecl(lean_object*, lean_object*); -extern lean_object* l_Lean_numLitKind; LEAN_EXPORT lean_object* l_Lean_Elab_elabSetOption_setOption___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_isStrLit_x3f(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabSetOption___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -503,146 +502,145 @@ x_11 = lean_erase_macro_scopes(x_10); x_12 = l_Lean_Syntax_isStrLit_x3f(x_2); if (lean_obj_tag(x_12) == 0) { -lean_object* x_13; lean_object* x_14; -x_13 = l_Lean_numLitKind; -x_14 = l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(x_13, x_2); -if (lean_obj_tag(x_14) == 0) +lean_object* x_13; +x_13 = l_Lean_Syntax_isNatLit_x3f(x_2); +if (lean_obj_tag(x_13) == 0) { switch (lean_obj_tag(x_2)) { case 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_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_dec(x_11); lean_dec(x_8); lean_dec(x_7); -x_15 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_15, 0, x_2); -x_16 = l_Lean_Elab_elabSetOption___rarg___lambda__1___closed__2; -x_17 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_15); -x_18 = l_Lean_Elab_elabSetOption___rarg___lambda__1___closed__4; -x_19 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set(x_19, 1, x_18); -x_20 = l_Lean_throwError___at_Lean_Elab_elabSetOption___spec__1___rarg(x_3, x_4, x_5, x_6, x_19); -return x_20; +x_14 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, x_2); +x_15 = l_Lean_Elab_elabSetOption___rarg___lambda__1___closed__2; +x_16 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_14); +x_17 = l_Lean_Elab_elabSetOption___rarg___lambda__1___closed__4; +x_18 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +x_19 = l_Lean_throwError___at_Lean_Elab_elabSetOption___spec__1___rarg(x_3, x_4, x_5, x_6, x_18); +return x_19; } case 1: { -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_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_dec(x_11); lean_dec(x_8); lean_dec(x_7); -x_21 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_21, 0, x_2); -x_22 = l_Lean_Elab_elabSetOption___rarg___lambda__1___closed__2; -x_23 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_21); -x_24 = l_Lean_Elab_elabSetOption___rarg___lambda__1___closed__4; -x_25 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_25, 0, x_23); -lean_ctor_set(x_25, 1, x_24); -x_26 = l_Lean_throwError___at_Lean_Elab_elabSetOption___spec__2___rarg(x_3, x_4, x_5, x_6, x_25); -return x_26; +x_20 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_20, 0, x_2); +x_21 = l_Lean_Elab_elabSetOption___rarg___lambda__1___closed__2; +x_22 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_20); +x_23 = l_Lean_Elab_elabSetOption___rarg___lambda__1___closed__4; +x_24 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +x_25 = l_Lean_throwError___at_Lean_Elab_elabSetOption___spec__2___rarg(x_3, x_4, x_5, x_6, x_24); +return x_25; } case 2: { -lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_27 = lean_ctor_get(x_2, 1); -lean_inc(x_27); -x_28 = l_Lean_Elab_elabSetOption___rarg___lambda__1___closed__5; -x_29 = lean_string_dec_eq(x_27, x_28); -if (x_29 == 0) +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = lean_ctor_get(x_2, 1); +lean_inc(x_26); +x_27 = l_Lean_Elab_elabSetOption___rarg___lambda__1___closed__5; +x_28 = lean_string_dec_eq(x_26, x_27); +if (x_28 == 0) { -lean_object* x_30; uint8_t x_31; -x_30 = l_Lean_Elab_elabSetOption___rarg___lambda__1___closed__6; -x_31 = lean_string_dec_eq(x_27, x_30); -lean_dec(x_27); -if (x_31 == 0) +lean_object* x_29; uint8_t x_30; +x_29 = l_Lean_Elab_elabSetOption___rarg___lambda__1___closed__6; +x_30 = lean_string_dec_eq(x_26, x_29); +lean_dec(x_26); +if (x_30 == 0) { -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_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_dec(x_11); lean_dec(x_8); lean_dec(x_7); -x_32 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_32, 0, x_2); -x_33 = l_Lean_Elab_elabSetOption___rarg___lambda__1___closed__2; -x_34 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_32); -x_35 = l_Lean_Elab_elabSetOption___rarg___lambda__1___closed__4; -x_36 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -x_37 = l_Lean_throwError___at_Lean_Elab_elabSetOption___spec__3___rarg(x_3, x_4, x_5, x_6, x_36); -return x_37; +x_31 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_31, 0, x_2); +x_32 = l_Lean_Elab_elabSetOption___rarg___lambda__1___closed__2; +x_33 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_31); +x_34 = l_Lean_Elab_elabSetOption___rarg___lambda__1___closed__4; +x_35 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +x_36 = l_Lean_throwError___at_Lean_Elab_elabSetOption___spec__3___rarg(x_3, x_4, x_5, x_6, x_35); +return x_36; } else { -lean_object* x_38; lean_object* x_39; +lean_object* x_37; lean_object* x_38; lean_dec(x_2); -x_38 = l_Lean_Elab_elabSetOption___rarg___lambda__1___closed__7; -x_39 = l_Lean_Elab_elabSetOption_setOption___rarg(x_3, x_7, x_4, x_5, x_6, x_8, x_11, x_38); -return x_39; +x_37 = l_Lean_Elab_elabSetOption___rarg___lambda__1___closed__7; +x_38 = l_Lean_Elab_elabSetOption_setOption___rarg(x_3, x_7, x_4, x_5, x_6, x_8, x_11, x_37); +return x_38; } } else { -lean_object* x_40; lean_object* x_41; -lean_dec(x_27); +lean_object* x_39; lean_object* x_40; +lean_dec(x_26); lean_dec(x_2); -x_40 = l_Lean_Elab_elabSetOption___rarg___lambda__1___closed__8; -x_41 = l_Lean_Elab_elabSetOption_setOption___rarg(x_3, x_7, x_4, x_5, x_6, x_8, x_11, x_40); -return x_41; +x_39 = l_Lean_Elab_elabSetOption___rarg___lambda__1___closed__8; +x_40 = l_Lean_Elab_elabSetOption_setOption___rarg(x_3, x_7, x_4, x_5, x_6, x_8, x_11, x_39); +return x_40; } } default: { -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_dec(x_11); lean_dec(x_8); lean_dec(x_7); -x_42 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_42, 0, x_2); -x_43 = l_Lean_Elab_elabSetOption___rarg___lambda__1___closed__2; -x_44 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_44, 1, x_42); -x_45 = l_Lean_Elab_elabSetOption___rarg___lambda__1___closed__4; -x_46 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -x_47 = l_Lean_throwError___at_Lean_Elab_elabSetOption___spec__4___rarg(x_3, x_4, x_5, x_6, x_46); -return x_47; +x_41 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_41, 0, x_2); +x_42 = l_Lean_Elab_elabSetOption___rarg___lambda__1___closed__2; +x_43 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_41); +x_44 = l_Lean_Elab_elabSetOption___rarg___lambda__1___closed__4; +x_45 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +x_46 = l_Lean_throwError___at_Lean_Elab_elabSetOption___spec__4___rarg(x_3, x_4, x_5, x_6, x_45); +return x_46; } } } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_dec(x_2); -x_48 = lean_ctor_get(x_14, 0); -lean_inc(x_48); -lean_dec(x_14); -x_49 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_49, 0, x_48); -x_50 = l_Lean_Elab_elabSetOption_setOption___rarg(x_3, x_7, x_4, x_5, x_6, x_8, x_11, x_49); -return x_50; +x_47 = lean_ctor_get(x_13, 0); +lean_inc(x_47); +lean_dec(x_13); +x_48 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_48, 0, x_47); +x_49 = l_Lean_Elab_elabSetOption_setOption___rarg(x_3, x_7, x_4, x_5, x_6, x_8, x_11, x_48); +return x_49; } } else { -lean_object* x_51; lean_object* x_52; lean_object* x_53; +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_dec(x_2); -x_51 = lean_ctor_get(x_12, 0); -lean_inc(x_51); +x_50 = lean_ctor_get(x_12, 0); +lean_inc(x_50); lean_dec(x_12); -x_52 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_52, 0, x_51); -x_53 = l_Lean_Elab_elabSetOption_setOption___rarg(x_3, x_7, x_4, x_5, x_6, x_8, x_11, x_52); -return x_53; +x_51 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_51, 0, x_50); +x_52 = l_Lean_Elab_elabSetOption_setOption___rarg(x_3, x_7, x_4, x_5, x_6, x_8, x_11, x_51); +return x_52; } } } diff --git a/stage0/stdlib/Lean/Elab/StructInst.c b/stage0/stdlib/Lean/Elab/StructInst.c index 870853d32d6..ccb02146c72 100644 --- a/stage0/stdlib/Lean/Elab/StructInst.c +++ b/stage0/stdlib/Lean/Elab/StructInst.c @@ -26,7 +26,6 @@ LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Term_StructInst_DefaultF static lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__8; size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Term_StructInst_DefaultFields_propagateLoop___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_fieldIdxKind; lean_object* l_List_forM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_StructInst_expandStructInstExpectedType_declRange___closed__7; LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -65,8 +64,8 @@ LEAN_EXPORT uint8_t l_Std_AssocList_contains___at___private_Lean_Elab_StructInst static lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStructInstAux___spec__4___closed__1; static lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__4; LEAN_EXPORT lean_object* l_Array_filterMapM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_groupFields___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_nullKind; lean_object* l_Lean_Meta_whnfForall(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__11; static lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_addMissingFields___closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_StructInst_DefaultFields_propagateLoop___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_name_mk_string(lean_object*, lean_object*); @@ -77,7 +76,6 @@ uint8_t lean_usize_dec_eq(size_t, size_t); lean_object* lean_array_uget(lean_object*, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_throwAbortTerm___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources___spec__2___rarg(lean_object*); static lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__2; -static lean_object* l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__9; LEAN_EXPORT lean_object* l_List_mapTRAux___at_Lean_Elab_Term_StructInst_formatStruct___spec__1(lean_object*, lean_object*); lean_object* lean_expr_update_mdata(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__8; @@ -114,7 +112,7 @@ static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_StructInst_0 lean_object* l_Lean_Expr_getAutoParamTactic_x3f(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_reduce___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_getFieldValue_x3f(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructSource___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_propagateLoop___closed__6; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev_docString(lean_object*); @@ -125,11 +123,10 @@ LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_StructInst LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_markDefaultMissing(lean_object*); LEAN_EXPORT lean_object* l_List_mapTRAux___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_groupFields___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_mkDefaultValueAux_x3f___closed__1; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_10107_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_10239_(lean_object*); static lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructName___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_mkDefaultValueAux_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(lean_object*, lean_object*); uint8_t l_Lean_checkTraceOption(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__20; static lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkCtorHeaderAux___closed__1; @@ -150,6 +147,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_StructInst_elabStructInst static lean_object* l___regBuiltin_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev_declRange___closed__4; static lean_object* l_List_mapM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__3___closed__6; lean_object* l_Lean_mkIdentFrom(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__9; static lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructName___closed__5; static lean_object* l_List_format___at_Lean_Elab_Term_StructInst_formatStruct___spec__3___closed__3; lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(lean_object*, lean_object*); @@ -198,7 +196,7 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Stru lean_object* lean_string_utf8_byte_size(lean_object*); static lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructName___closed__1; lean_object* l_List_head_x21___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashMapImp_find_x3f___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkFieldMap___spec__4(lean_object*, lean_object*); lean_object* l_Lean_Expr_getRevArg_x21(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_tryToSynthesizeDefault___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -224,11 +222,13 @@ static lean_object* l_Lean_Elab_Term_StructInst_instToFormatFieldLHS___closed__3 extern lean_object* l_Lean_levelZero; static lean_object* l_List_format___at_Lean_Elab_Term_StructInst_formatStruct___spec__3___closed__4; LEAN_EXPORT lean_object* l_Array_filterMapM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_groupFields___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_StructInst_FieldLHS_toSyntax___closed__1; lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___closed__3; static lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__6; lean_object* l_Lean_Elab_Term_ensureHasType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_StructInst_FieldLHS_toSyntax___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_collectStructNames(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_StructInst_instInhabitedFieldLHS___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_tryToSynthesizeDefault_loop___lambda__1___boxed(lean_object*); @@ -237,7 +237,6 @@ static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expand LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_Field_isSimple___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkFieldMap___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Format_joinSep___at_Lean_Elab_Term_StructInst_formatField___spec__1(lean_object*, lean_object*); -lean_object* l_Lean_Syntax_SepArray_ofElems(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Elab_Term_StructInst_DefaultFields_getHierarchyDepth___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkCtorHeader(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_isSimpleField_x3f(lean_object*); @@ -267,8 +266,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_Struct_modifyFieldsM___at__ lean_object* l___private_Lean_Elab_Util_0__Lean_Elab_evalSyntaxConstantUnsafe(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_StructInst_throwFailedToElabField___spec__1(lean_object*); lean_object* l_Lean_Meta_project_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Syntax_SepArray_getElems___rarg(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashMap_toList___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_groupFields___spec__5(lean_object*); static lean_object* l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkFieldMap___spec__13___closed__3; @@ -289,6 +287,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_instToStringStruct; static lean_object* l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__17; lean_object* l_Lean_Meta_lambdaLetTelescope___at___private_Lean_Meta_InferType_0__Lean_Meta_inferLambdaType___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_groupFields(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_findField_x3f___lambda__1___boxed(lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_addTrace___at_Lean_Elab_Term_StructInst_DefaultFields_propagateLoop___spec__3___closed__4; @@ -299,6 +298,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_StructInst_expandStructInstExp static lean_object* l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__9; lean_object* lean_expr_consume_type_annotations(lean_object*); static lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructName_throwUnexpectedExpectedType___closed__3; +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); lean_object* l_Lean_Elab_expandMacroImpl_x3f(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStructInstAux___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); @@ -324,6 +324,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_Context_struc lean_object* l_Lean_Elab_Term_withMacroExpansion___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_throwAbortTerm___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources___spec__2___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_findDefaultMissing_x3f(lean_object*, lean_object*); +lean_object* l_Lean_Syntax_isFieldIdx_x3f(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_FieldLHS_toSyntax(uint8_t, lean_object*); static lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__17; LEAN_EXPORT lean_object* l_Std_Format_joinSep___at_Lean_Elab_Term_StructInst_formatStruct___spec__4(lean_object*, lean_object*); @@ -456,7 +457,6 @@ static lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructIns static lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__5; lean_object* lean_whnf(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources_go___lambda__1___closed__11; -static lean_object* l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__10; LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkCtorHeaderAux___closed__2; LEAN_EXPORT uint8_t l_Lean_Elab_Term_StructInst_Field_isSimple___rarg(lean_object*); @@ -500,14 +500,12 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_StructInst_expandStructInstExp static lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructName_throwUnknownExpectedType___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__10; lean_object* l_Lean_mkAtomFrom(lean_object*, lean_object*); -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructName_throwUnknownExpectedType___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources_go___lambda__1___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNumLitFields(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_TSepArray_getElems___rarg(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandCompositeFields___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_groupKind; -static lean_object* l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__7; LEAN_EXPORT uint8_t l_Lean_Elab_Term_StructInst_DefaultFields_tryToSynthesizeDefault_loop___lambda__1(lean_object*); uint8_t l_Lean_Syntax_isMissing(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwAbortTerm___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -537,6 +535,7 @@ static lean_object* l_Lean_Elab_Term_StructInst_throwFailedToElabField___rarg___ lean_object* l_Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOptionM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTerm___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_StructInst_formatField___closed__1; +static lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__10; lean_object* l_List_find_x3f___rarg(lean_object*, lean_object*); extern lean_object* l_Lean_Elab_Term_instMonadTermElabM; lean_object* l_Lean_mkHole(lean_object*); @@ -651,7 +650,6 @@ static lean_object* l_Lean_Elab_Term_StructInst_Field_toSyntax___closed__2; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); uint8_t l_Lean_Expr_hasLooseBVars(lean_object*); -static lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStructInstAux___spec__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Elab_Term_StructInst_DefaultFields_step___spec__2___closed__3; static lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources_go___lambda__1___closed__2; @@ -673,7 +671,6 @@ static lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructIns LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_groupFields___spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__16; static lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3___closed__1; -static lean_object* l_List_mapM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_groupFields___spec__8___closed__2; lean_object* l_Lean_Meta_getStructureName(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__11; static lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__6; @@ -709,7 +706,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_Stru static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_isModifyOp_x3f___spec__3___closed__3; lean_object* l_Lean_Expr_getAppFn(lean_object*); static lean_object* l_Lean_Elab_Term_StructInst_formatStruct___closed__2; -static lean_object* l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__8; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_groupFields___spec__7(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_trySynthStructInstance_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Elab_Term_StructInst_DefaultFields_step___spec__2___closed__1; @@ -815,40 +811,40 @@ return x_2; static lean_object* _init_l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = lean_box(2); -x_2 = l_Lean_nullKind; -x_3 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__1; -x_4 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_4, 0, x_1); -lean_ctor_set(x_4, 1, x_2); -lean_ctor_set(x_4, 2, x_3); -return x_4; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("null", 4); +return x_1; } } static lean_object* _init_l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__3() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Lean", 4); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__2; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; } } static lean_object* _init_l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__4() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(2); x_2 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__3; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; +x_3 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__1; +x_4 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; } } static lean_object* _init_l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__5() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Parser", 6); +x_1 = lean_mk_string_from_bytes("Lean", 4); return x_1; } } @@ -856,7 +852,7 @@ static lean_object* _init_l_Lean_Elab_Term_StructInst_expandStructInstExpectedTy _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__4; +x_1 = lean_box(0); x_2 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -866,7 +862,7 @@ static lean_object* _init_l_Lean_Elab_Term_StructInst_expandStructInstExpectedTy _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Term", 4); +x_1 = lean_mk_string_from_bytes("Parser", 6); return x_1; } } @@ -884,7 +880,7 @@ static lean_object* _init_l_Lean_Elab_Term_StructInst_expandStructInstExpectedTy _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("paren", 5); +x_1 = lean_mk_string_from_bytes("Term", 4); return x_1; } } @@ -902,26 +898,26 @@ static lean_object* _init_l_Lean_Elab_Term_StructInst_expandStructInstExpectedTy _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("(", 1); +x_1 = lean_mk_string_from_bytes("paren", 5); return x_1; } } static lean_object* _init_l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__12() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("null", 4); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__10; +x_2 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__11; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; } } static lean_object* _init_l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__13() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__12; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("(", 1); +return x_1; } } static lean_object* _init_l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__14() { @@ -936,7 +932,7 @@ static lean_object* _init_l_Lean_Elab_Term_StructInst_expandStructInstExpectedTy _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__8; +x_1 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__10; x_2 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__14; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -998,7 +994,7 @@ lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_ob x_7 = lean_unsigned_to_nat(1u); x_8 = l_Lean_Syntax_getArg(x_5, x_7); lean_dec(x_5); -x_9 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__2; +x_9 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__4; x_10 = l_Lean_Syntax_setArg(x_1, x_4, x_9); x_11 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); x_12 = !lean_is_exclusive(x_11); @@ -1006,7 +1002,7 @@ if (x_12 == 0) { 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; 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; 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; x_13 = lean_ctor_get(x_11, 0); -x_14 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__11; +x_14 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__13; lean_inc(x_13); x_15 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_15, 0, x_13); @@ -1027,7 +1023,7 @@ lean_ctor_set(x_23, 1, x_22); lean_ctor_set(x_23, 2, x_20); x_24 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__18; x_25 = lean_array_push(x_24, x_23); -x_26 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__13; +x_26 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__3; x_27 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_27, 0, x_21); lean_ctor_set(x_27, 1, x_26); @@ -1046,7 +1042,7 @@ x_33 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__20; x_34 = lean_array_push(x_33, x_15); x_35 = lean_array_push(x_34, x_30); x_36 = lean_array_push(x_35, x_32); -x_37 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__10; +x_37 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__12; x_38 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_38, 0, x_21); lean_ctor_set(x_38, 1, x_37); @@ -1062,7 +1058,7 @@ x_40 = lean_ctor_get(x_11, 1); lean_inc(x_40); lean_inc(x_39); lean_dec(x_11); -x_41 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__11; +x_41 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__13; lean_inc(x_39); x_42 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_42, 0, x_39); @@ -1083,7 +1079,7 @@ lean_ctor_set(x_50, 1, x_49); lean_ctor_set(x_50, 2, x_47); x_51 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__18; x_52 = lean_array_push(x_51, x_50); -x_53 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__13; +x_53 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__3; x_54 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_54, 0, x_48); lean_ctor_set(x_54, 1, x_53); @@ -1102,7 +1098,7 @@ x_60 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__20; x_61 = lean_array_push(x_60, x_42); x_62 = lean_array_push(x_61, x_57); x_63 = lean_array_push(x_62, x_59); -x_64 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__10; +x_64 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__12; x_65 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_65, 0, x_48); lean_ctor_set(x_65, 1, x_64); @@ -1139,7 +1135,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_StructInst_expandStructI _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__8; +x_1 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__10; x_2 = l___regBuiltin_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -1157,7 +1153,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_StructInst_expandStructI _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__4; +x_1 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__6; x_2 = l___regBuiltin_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -1168,7 +1164,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_StructInst_expandStructI { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__4; -x_2 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__7; +x_2 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__9; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } @@ -1259,7 +1255,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_StructInst_expandStructI _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(27u); +x_1 = lean_unsigned_to_nat(28u); x_2 = lean_unsigned_to_nat(44u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1271,7 +1267,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_StructInst_expandStructI _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(34u); +x_1 = lean_unsigned_to_nat(35u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1299,7 +1295,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_StructInst_expandStructI _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(27u); +x_1 = lean_unsigned_to_nat(28u); x_2 = lean_unsigned_to_nat(48u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1311,7 +1307,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_StructInst_expandStructI _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(27u); +x_1 = lean_unsigned_to_nat(28u); x_2 = lean_unsigned_to_nat(76u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1357,14 +1353,6 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -static lean_object* _init_l_Array_anyMUnsafe_any___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__1___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("structInstFieldAbbrev", 21); -return x_1; -} -} LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__1(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { _start: { @@ -1372,37 +1360,31 @@ uint8_t x_5; x_5 = lean_usize_dec_eq(x_3, x_4); if (x_5 == 0) { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +lean_object* x_6; lean_object* x_7; uint8_t x_8; x_6 = lean_array_uget(x_2, x_3); x_7 = l_Lean_Syntax_getKind(x_6); -x_8 = l_Array_anyMUnsafe_any___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__1___closed__1; -lean_inc(x_1); -x_9 = lean_name_mk_string(x_1, x_8); -x_10 = lean_name_eq(x_7, x_9); -lean_dec(x_9); +x_8 = lean_name_eq(x_7, x_1); lean_dec(x_7); -if (x_10 == 0) +if (x_8 == 0) { -size_t x_11; size_t x_12; -x_11 = 1; -x_12 = lean_usize_add(x_3, x_11); -x_3 = x_12; +size_t x_9; size_t x_10; +x_9 = 1; +x_10 = lean_usize_add(x_3, x_9); +x_3 = x_10; goto _start; } else { -uint8_t x_14; -lean_dec(x_1); -x_14 = 1; -return x_14; +uint8_t x_12; +x_12 = 1; +return x_12; } } else { -uint8_t x_15; -lean_dec(x_1); -x_15 = 0; -return x_15; +uint8_t x_13; +x_13 = 0; +return x_13; } } } @@ -1410,24 +1392,34 @@ static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_ _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("structInstField", 15); +x_1 = lean_mk_string_from_bytes("ident", 5); return x_1; } } static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__2() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__3() { +_start: +{ lean_object* x_1; x_1 = lean_mk_string_from_bytes("structInstLVal", 14); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__3() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__13; +x_2 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__3; x_3 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__1; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -1436,7 +1428,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__4() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__5() { _start: { lean_object* x_1; @@ -1444,101 +1436,112 @@ x_1 = lean_mk_string_from_bytes(":=", 2); return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2(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, lean_object* x_8, lean_object* x_9) { _start: { -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_3, x_2); -if (x_7 == 0) +uint8_t x_10; +x_10 = lean_usize_dec_lt(x_6, x_5); +if (x_10 == 0) { -lean_object* x_8; -lean_dec(x_5); +lean_object* x_11; +lean_dec(x_8); +lean_dec(x_3); lean_dec(x_1); -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_4); -lean_ctor_set(x_8, 1, x_6); -return x_8; +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_7); +lean_ctor_set(x_11, 1, x_9); +return x_11; } else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_9 = lean_array_uget(x_4, x_3); -x_10 = lean_unsigned_to_nat(0u); -x_11 = lean_array_uset(x_4, x_3, x_10); -lean_inc(x_9); -x_12 = l_Lean_Syntax_getKind(x_9); -x_13 = l_Array_anyMUnsafe_any___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__1___closed__1; -lean_inc(x_1); -x_14 = lean_name_mk_string(x_1, x_13); -x_15 = lean_name_eq(x_12, x_14); -lean_dec(x_14); -lean_dec(x_12); +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = lean_array_uget(x_7, x_6); +x_13 = lean_unsigned_to_nat(0u); +x_14 = lean_array_uset(x_7, x_6, x_13); +lean_inc(x_12); +x_15 = l_Lean_Syntax_isOfKind(x_12, x_2); if (x_15 == 0) { size_t x_16; size_t x_17; lean_object* x_18; x_16 = 1; -x_17 = lean_usize_add(x_3, x_16); -x_18 = lean_array_uset(x_11, x_3, x_9); -x_3 = x_17; -x_4 = x_18; +x_17 = lean_usize_add(x_6, x_16); +x_18 = lean_array_uset(x_14, x_6, x_12); +x_6 = x_17; +x_7 = x_18; goto _start; } else { -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; 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; size_t x_41; size_t x_42; lean_object* x_43; -x_20 = l_Lean_Syntax_getArg(x_9, x_10); -lean_dec(x_9); -lean_inc(x_5); -x_21 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_5, x_6); -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_21, 1); -lean_inc(x_23); -lean_dec(x_21); -x_24 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__1; -lean_inc(x_1); -x_25 = lean_name_mk_string(x_1, x_24); -x_26 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__2; -lean_inc(x_1); -x_27 = lean_name_mk_string(x_1, x_26); -x_28 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__17; +lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_20 = l_Lean_Syntax_getArg(x_12, x_13); +x_21 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__2; lean_inc(x_20); -x_29 = lean_array_push(x_28, x_20); +x_22 = l_Lean_Syntax_isOfKind(x_20, x_21); +if (x_22 == 0) +{ +size_t x_23; size_t x_24; lean_object* x_25; +lean_dec(x_20); +x_23 = 1; +x_24 = lean_usize_add(x_6, x_23); +x_25 = lean_array_uset(x_14, x_6, x_12); +x_6 = x_24; +x_7 = x_25; +goto _start; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; 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; lean_object* x_44; size_t x_45; size_t x_46; lean_object* x_47; +lean_dec(x_12); +lean_inc(x_8); +x_27 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_8, x_9); +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_27, 1); +lean_inc(x_29); +lean_dec(x_27); x_30 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__3; -x_31 = lean_array_push(x_29, x_30); -x_32 = lean_box(2); -x_33 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_27); -lean_ctor_set(x_33, 2, x_31); +lean_inc(x_1); +x_31 = lean_name_mk_string(x_1, x_30); +x_32 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__17; +lean_inc(x_20); +x_33 = lean_array_push(x_32, x_20); x_34 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__4; -x_35 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_35, 0, x_22); -lean_ctor_set(x_35, 1, x_34); -x_36 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__20; -x_37 = lean_array_push(x_36, x_33); -x_38 = lean_array_push(x_37, x_35); -x_39 = lean_array_push(x_38, x_20); -x_40 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_40, 0, x_32); -lean_ctor_set(x_40, 1, x_25); -lean_ctor_set(x_40, 2, x_39); -x_41 = 1; -x_42 = lean_usize_add(x_3, x_41); -x_43 = lean_array_uset(x_11, x_3, x_40); -x_3 = x_42; -x_4 = x_43; -x_6 = x_23; +x_35 = lean_array_push(x_33, x_34); +x_36 = lean_box(2); +x_37 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_31); +lean_ctor_set(x_37, 2, x_35); +x_38 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__5; +x_39 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_39, 0, x_28); +lean_ctor_set(x_39, 1, x_38); +x_40 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__20; +x_41 = lean_array_push(x_40, x_37); +x_42 = lean_array_push(x_41, x_39); +x_43 = lean_array_push(x_42, x_20); +lean_inc(x_3); +x_44 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_44, 0, x_36); +lean_ctor_set(x_44, 1, x_3); +lean_ctor_set(x_44, 2, x_43); +x_45 = 1; +x_46 = lean_usize_add(x_6, x_45); +x_47 = lean_array_uset(x_14, x_6, x_44); +x_6 = x_46; +x_7 = x_47; +x_9 = x_29; goto _start; } } } } +} static lean_object* _init_l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("{", 1); +x_1 = lean_mk_string_from_bytes("structInstFieldAbbrev", 21); return x_1; } } @@ -1546,7 +1549,7 @@ static lean_object* _init_l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbre _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes(",", 1); +x_1 = lean_mk_string_from_bytes("structInstField", 15); return x_1; } } @@ -1554,20 +1557,48 @@ static lean_object* _init_l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbre _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("}", 1); +x_1 = lean_mk_string_from_bytes("{", 1); return x_1; } } static lean_object* _init_l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__4() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(",", 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(2); +x_2 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__4; +x_3 = lean_alloc_ctor(2, 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_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("}", 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__7() { +_start: +{ lean_object* x_1; lean_object* x_2; x_1 = lean_unsigned_to_nat(6u); x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__5() { +static lean_object* _init_l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__8() { _start: { lean_object* x_1; lean_object* x_2; @@ -1576,13 +1607,13 @@ x_2 = l_Array_append___rarg(x_1, x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__6() { +static lean_object* _init_l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__13; -x_3 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__5; +x_2 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__3; +x_3 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__8; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1590,7 +1621,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__7() { +static lean_object* _init_l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__10() { _start: { lean_object* x_1; @@ -1598,7 +1629,7 @@ x_1 = lean_mk_string_from_bytes("..", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__8() { +static lean_object* _init_l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__11() { _start: { lean_object* x_1; @@ -1609,264 +1640,285 @@ return x_1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___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, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_11 = l_Lean_Syntax_getArgs(x_1); -x_12 = l_Lean_Syntax_SepArray_getElems___rarg(x_11); -lean_dec(x_11); -x_13 = lean_array_get_size(x_12); -x_14 = lean_unsigned_to_nat(0u); -x_15 = lean_nat_dec_lt(x_14, x_13); -if (x_15 == 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; lean_object* x_20; uint8_t x_21; +x_11 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__1; +lean_inc(x_1); +x_12 = lean_name_mk_string(x_1, x_11); +x_13 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__2; +lean_inc(x_1); +x_14 = lean_name_mk_string(x_1, x_13); +x_15 = lean_box(0); +lean_inc(x_14); +x_16 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +x_17 = l_Lean_Syntax_getArgs(x_2); +x_18 = l_Lean_Syntax_TSepArray_getElems___rarg(x_17); +lean_dec(x_17); +x_19 = lean_array_get_size(x_18); +x_20 = lean_unsigned_to_nat(0u); +x_21 = lean_nat_dec_lt(x_20, x_19); +if (x_21 == 0) { -lean_object* x_16; lean_object* x_17; -lean_dec(x_13); +lean_object* x_22; lean_object* x_23; +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_16); +lean_dec(x_14); lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -x_16 = lean_box(1); -x_17 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_10); -return x_17; +lean_dec(x_1); +x_22 = lean_box(1); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_10); +return x_23; } else { -uint8_t x_18; -x_18 = lean_nat_dec_le(x_13, x_13); -if (x_18 == 0) +uint8_t x_24; +x_24 = lean_nat_dec_le(x_19, x_19); +if (x_24 == 0) { -lean_object* x_19; lean_object* x_20; -lean_dec(x_13); +lean_object* x_25; lean_object* x_26; +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_16); +lean_dec(x_14); lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -x_19 = lean_box(1); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_10); -return x_20; +lean_dec(x_1); +x_25 = lean_box(1); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_10); +return x_26; } else { -size_t x_21; size_t x_22; uint8_t x_23; -x_21 = 0; -x_22 = lean_usize_of_nat(x_13); -lean_dec(x_13); -lean_inc(x_2); -x_23 = l_Array_anyMUnsafe_any___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__1(x_2, x_12, x_21, x_22); -if (x_23 == 0) +size_t x_27; size_t x_28; uint8_t x_29; +x_27 = 0; +x_28 = lean_usize_of_nat(x_19); +lean_dec(x_19); +x_29 = l_Array_anyMUnsafe_any___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__1(x_12, x_18, x_27, x_28); +if (x_29 == 0) { -lean_object* x_24; lean_object* x_25; +lean_object* x_30; lean_object* x_31; +lean_dec(x_18); +lean_dec(x_16); +lean_dec(x_14); lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -x_24 = lean_box(1); -x_25 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_10); -return x_25; +lean_dec(x_1); +x_30 = lean_box(1); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_10); +return x_31; } else { -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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_inc(x_9); -x_26 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2(x_2, x_22, x_21, x_12, x_9, x_10); -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -x_29 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_9, x_28); -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_29, 1); -lean_inc(x_31); -if (lean_is_exclusive(x_29)) { - lean_ctor_release(x_29, 0); - lean_ctor_release(x_29, 1); - x_32 = x_29; -} else { - lean_dec_ref(x_29); - x_32 = lean_box(0); -} -x_33 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__1; -lean_inc(x_30); -x_34 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_34, 0, x_30); -lean_ctor_set(x_34, 1, x_33); -x_35 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__2; -x_36 = l_Lean_Syntax_SepArray_ofElems(x_35, x_27); -lean_dec(x_27); -x_37 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__1; -x_38 = l_Array_append___rarg(x_37, x_36); -x_39 = lean_box(2); -x_40 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__13; -x_41 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_40); -lean_ctor_set(x_41, 2, x_38); -x_42 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__3; -lean_inc(x_30); -x_43 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_43, 0, x_30); -lean_ctor_set(x_43, 1, x_42); -x_44 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__4; -x_45 = lean_array_push(x_44, x_34); +x_32 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2(x_1, x_12, x_14, x_16, x_28, x_27, x_18, x_9, x_10); +lean_dec(x_16); +lean_dec(x_12); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +x_35 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_9, x_34); +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +if (lean_is_exclusive(x_35)) { + lean_ctor_release(x_35, 0); + lean_ctor_release(x_35, 1); + x_38 = x_35; +} else { + lean_dec_ref(x_35); + x_38 = lean_box(0); +} +x_39 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__3; +lean_inc(x_36); +x_40 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_40, 0, x_36); +lean_ctor_set(x_40, 1, x_39); +x_41 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__5; +x_42 = l_Lean_mkSepArray(x_33, x_41); +lean_dec(x_33); +x_43 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__1; +x_44 = l_Array_append___rarg(x_43, x_42); +x_45 = lean_box(2); +x_46 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__3; +x_47 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +lean_ctor_set(x_47, 2, x_44); +x_48 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__6; +lean_inc(x_36); +x_49 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_49, 0, x_36); +lean_ctor_set(x_49, 1, x_48); +x_50 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__7; +x_51 = lean_array_push(x_50, x_40); if (lean_obj_tag(x_6) == 0) { -x_46 = x_37; -goto block_87; +x_52 = x_43; +goto block_93; } else { -lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; -x_88 = lean_ctor_get(x_6, 0); -lean_inc(x_88); +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_94 = lean_ctor_get(x_6, 0); +lean_inc(x_94); lean_dec(x_6); -x_89 = l_Array_append___rarg(x_37, x_88); -x_90 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_90, 0, x_39); -lean_ctor_set(x_90, 1, x_40); -lean_ctor_set(x_90, 2, x_89); -x_91 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__8; -lean_inc(x_30); -x_92 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_92, 0, x_30); -lean_ctor_set(x_92, 1, x_91); -x_93 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__17; -x_94 = lean_array_push(x_93, x_90); -x_95 = lean_array_push(x_94, x_92); -x_46 = x_95; -goto block_87; -} -block_87: -{ -lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_47 = l_Array_append___rarg(x_37, x_46); -x_48 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_48, 0, x_39); -lean_ctor_set(x_48, 1, x_40); -lean_ctor_set(x_48, 2, x_47); -x_49 = lean_array_push(x_45, x_48); -x_50 = lean_array_push(x_49, x_41); +x_95 = l_Array_append___rarg(x_43, x_94); +x_96 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_96, 0, x_45); +lean_ctor_set(x_96, 1, x_46); +lean_ctor_set(x_96, 2, x_95); +x_97 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__11; +lean_inc(x_36); +x_98 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_98, 0, x_36); +lean_ctor_set(x_98, 1, x_97); +x_99 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__17; +x_100 = lean_array_push(x_99, x_96); +x_101 = lean_array_push(x_100, x_98); +x_52 = x_101; +goto block_93; +} +block_93: +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_53 = l_Array_append___rarg(x_43, x_52); +x_54 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_54, 0, x_45); +lean_ctor_set(x_54, 1, x_46); +lean_ctor_set(x_54, 2, x_53); +x_55 = lean_array_push(x_51, x_54); +x_56 = lean_array_push(x_55, x_47); if (lean_obj_tag(x_5) == 0) { -x_51 = x_37; -goto block_75; +x_57 = x_43; +goto block_81; } else { -lean_object* x_76; lean_object* x_77; -x_76 = lean_ctor_get(x_5, 0); -x_77 = l_Lean_Syntax_getHeadInfo_x3f(x_76); -if (lean_obj_tag(x_77) == 0) +lean_object* x_82; lean_object* x_83; +x_82 = lean_ctor_get(x_5, 0); +x_83 = l_Lean_Syntax_getHeadInfo_x3f(x_82); +if (lean_obj_tag(x_83) == 0) { -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_78 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__7; -lean_inc(x_30); -x_79 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_79, 0, x_30); -lean_ctor_set(x_79, 1, x_78); -x_80 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__18; -x_81 = lean_array_push(x_80, x_79); -x_51 = x_81; -goto block_75; +lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_84 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__10; +lean_inc(x_36); +x_85 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_85, 0, x_36); +lean_ctor_set(x_85, 1, x_84); +x_86 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__18; +x_87 = lean_array_push(x_86, x_85); +x_57 = x_87; +goto block_81; } else { -lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; -x_82 = lean_ctor_get(x_77, 0); -lean_inc(x_82); -lean_dec(x_77); -x_83 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__7; -x_84 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_84, 0, x_82); -lean_ctor_set(x_84, 1, x_83); -x_85 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__18; -x_86 = lean_array_push(x_85, x_84); -x_51 = x_86; -goto block_75; +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_88 = lean_ctor_get(x_83, 0); +lean_inc(x_88); +lean_dec(x_83); +x_89 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__10; +x_90 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_90, 0, x_88); +lean_ctor_set(x_90, 1, x_89); +x_91 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__18; +x_92 = lean_array_push(x_91, x_90); +x_57 = x_92; +goto block_81; } } -block_75: +block_81: { -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_52 = l_Array_append___rarg(x_37, x_51); -x_53 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_53, 0, x_39); -lean_ctor_set(x_53, 1, x_40); -lean_ctor_set(x_53, 2, x_52); -x_54 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__18; -x_55 = lean_array_push(x_54, x_53); -x_56 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_56, 0, x_39); -lean_ctor_set(x_56, 1, x_3); -lean_ctor_set(x_56, 2, x_55); -x_57 = lean_array_push(x_50, x_56); +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_58 = l_Array_append___rarg(x_43, x_57); +x_59 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_59, 0, x_45); +lean_ctor_set(x_59, 1, x_46); +lean_ctor_set(x_59, 2, x_58); +x_60 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__18; +x_61 = lean_array_push(x_60, x_59); +x_62 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_62, 0, x_45); +lean_ctor_set(x_62, 1, x_3); +lean_ctor_set(x_62, 2, x_61); +x_63 = lean_array_push(x_56, x_62); if (lean_obj_tag(x_8) == 0) { -lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; -lean_dec(x_30); -x_58 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__6; -x_59 = lean_array_push(x_57, x_58); -x_60 = lean_array_push(x_59, x_43); -x_61 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_61, 0, x_39); -lean_ctor_set(x_61, 1, x_4); -lean_ctor_set(x_61, 2, x_60); -if (lean_is_scalar(x_32)) { - x_62 = lean_alloc_ctor(0, 2, 0); +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +lean_dec(x_36); +x_64 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__9; +x_65 = lean_array_push(x_63, x_64); +x_66 = lean_array_push(x_65, x_49); +x_67 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_67, 0, x_45); +lean_ctor_set(x_67, 1, x_4); +lean_ctor_set(x_67, 2, x_66); +if (lean_is_scalar(x_38)) { + x_68 = lean_alloc_ctor(0, 2, 0); } else { - x_62 = x_32; + x_68 = x_38; } -lean_ctor_set(x_62, 0, x_61); -lean_ctor_set(x_62, 1, x_31); -return x_62; +lean_ctor_set(x_68, 0, x_67); +lean_ctor_set(x_68, 1, x_37); +return x_68; } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; 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; -x_63 = lean_ctor_get(x_8, 0); -lean_inc(x_63); +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; lean_object* x_79; lean_object* x_80; +x_69 = lean_ctor_get(x_8, 0); +lean_inc(x_69); lean_dec(x_8); -x_64 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__16; -x_65 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_65, 0, x_30); -lean_ctor_set(x_65, 1, x_64); -x_66 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__17; -x_67 = lean_array_push(x_66, x_65); -x_68 = lean_array_push(x_67, x_63); -x_69 = l_Array_append___rarg(x_37, x_68); -x_70 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_70, 0, x_39); -lean_ctor_set(x_70, 1, x_40); -lean_ctor_set(x_70, 2, x_69); -x_71 = lean_array_push(x_57, x_70); -x_72 = lean_array_push(x_71, x_43); -x_73 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_73, 0, x_39); -lean_ctor_set(x_73, 1, x_4); -lean_ctor_set(x_73, 2, x_72); -if (lean_is_scalar(x_32)) { - x_74 = lean_alloc_ctor(0, 2, 0); +x_70 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__16; +x_71 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_71, 0, x_36); +lean_ctor_set(x_71, 1, x_70); +x_72 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__17; +x_73 = lean_array_push(x_72, x_71); +x_74 = lean_array_push(x_73, x_69); +x_75 = l_Array_append___rarg(x_43, x_74); +x_76 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_76, 0, x_45); +lean_ctor_set(x_76, 1, x_46); +lean_ctor_set(x_76, 2, x_75); +x_77 = lean_array_push(x_63, x_76); +x_78 = lean_array_push(x_77, x_49); +x_79 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_79, 0, x_45); +lean_ctor_set(x_79, 1, x_4); +lean_ctor_set(x_79, 2, x_78); +if (lean_is_scalar(x_38)) { + x_80 = lean_alloc_ctor(0, 2, 0); } else { - x_74 = x_32; + x_80 = x_38; } -lean_ctor_set(x_74, 0, x_73); -lean_ctor_set(x_74, 1, x_31); -return x_74; +lean_ctor_set(x_80, 0, x_79); +lean_ctor_set(x_80, 1, x_37); +return x_80; } } } @@ -1886,14 +1938,13 @@ lean_dec(x_1); x_13 = l_Lean_Syntax_isNone(x_12); if (x_13 == 0) { -lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_14 = l_Lean_nullKind; -x_15 = lean_unsigned_to_nat(2u); +lean_object* x_14; uint8_t x_15; +x_14 = lean_unsigned_to_nat(2u); lean_inc(x_12); -x_16 = l_Lean_Syntax_isNodeOf(x_12, x_14, x_15); -if (x_16 == 0) +x_15 = l_Lean_Syntax_matchesNull(x_12, x_14); +if (x_15 == 0) { -lean_object* x_17; lean_object* x_18; +lean_object* x_16; lean_object* x_17; lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); @@ -1902,37 +1953,37 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_17 = lean_box(1); -x_18 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_10); -return x_18; +x_16 = lean_box(1); +x_17 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_10); +return x_17; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_19 = lean_unsigned_to_nat(1u); -x_20 = l_Lean_Syntax_getArg(x_12, x_19); +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_18 = lean_unsigned_to_nat(1u); +x_19 = l_Lean_Syntax_getArg(x_12, x_18); lean_dec(x_12); -x_21 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_21, 0, x_20); -x_22 = lean_box(0); -x_23 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1(x_2, x_3, x_4, x_5, x_8, x_6, x_22, x_21, x_9, x_10); +x_20 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_20, 0, x_19); +x_21 = lean_box(0); +x_22 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1(x_2, x_3, x_4, x_5, x_8, x_6, x_21, x_20, x_9, x_10); lean_dec(x_8); -lean_dec(x_2); -return x_23; +lean_dec(x_3); +return x_22; } } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_dec(x_12); +x_23 = lean_box(0); x_24 = lean_box(0); -x_25 = lean_box(0); -x_26 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1(x_2, x_3, x_4, x_5, x_8, x_6, x_25, x_24, x_9, x_10); +x_25 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1(x_2, x_3, x_4, x_5, x_8, x_6, x_24, x_23, x_9, x_10); lean_dec(x_8); -lean_dec(x_2); -return x_26; +lean_dec(x_3); +return x_25; } } } @@ -1984,14 +2035,13 @@ lean_dec(x_11); x_19 = l_Lean_Syntax_isNone(x_18); if (x_19 == 0) { -lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_20 = l_Lean_nullKind; -x_21 = lean_unsigned_to_nat(1u); +lean_object* x_20; uint8_t x_21; +x_20 = lean_unsigned_to_nat(1u); lean_inc(x_18); -x_22 = l_Lean_Syntax_isNodeOf(x_18, x_20, x_21); -if (x_22 == 0) +x_21 = l_Lean_Syntax_matchesNull(x_18, x_20); +if (x_21 == 0) { -lean_object* x_23; lean_object* x_24; +lean_object* x_22; lean_object* x_23; lean_dec(x_18); lean_dec(x_13); lean_dec(x_9); @@ -2000,32 +2050,32 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_23 = lean_box(1); -x_24 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_7); -return x_24; +x_22 = lean_box(1); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_7); +return x_23; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_25 = l_Lean_Syntax_getArg(x_18, x_17); +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_24 = l_Lean_Syntax_getArg(x_18, x_17); lean_dec(x_18); -x_26 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_26, 0, x_25); -x_27 = lean_box(0); -x_28 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__2(x_1, x_9, x_2, x_13, x_3, x_5, x_27, x_26, x_6, x_7); -return x_28; +x_25 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_25, 0, x_24); +x_26 = lean_box(0); +x_27 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__2(x_1, x_2, x_9, x_13, x_3, x_5, x_26, x_25, x_6, x_7); +return x_27; } } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_dec(x_18); +x_28 = lean_box(0); x_29 = lean_box(0); -x_30 = lean_box(0); -x_31 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__2(x_1, x_9, x_2, x_13, x_3, x_5, x_30, x_29, x_6, x_7); -return x_31; +x_30 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__2(x_1, x_2, x_9, x_13, x_3, x_5, x_29, x_28, x_6, x_7); +return x_30; } } } @@ -2056,48 +2106,47 @@ x_9 = l_Lean_Syntax_getArg(x_1, x_8); x_10 = l_Lean_Syntax_isNone(x_9); if (x_10 == 0) { -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = l_Lean_nullKind; -x_12 = lean_unsigned_to_nat(2u); +lean_object* x_11; uint8_t x_12; +x_11 = lean_unsigned_to_nat(2u); lean_inc(x_9); -x_13 = l_Lean_Syntax_isNodeOf(x_9, x_11, x_12); -if (x_13 == 0) +x_12 = l_Lean_Syntax_matchesNull(x_9, x_11); +if (x_12 == 0) { -lean_object* x_14; lean_object* x_15; +lean_object* x_13; lean_object* x_14; lean_dec(x_9); lean_dec(x_2); lean_dec(x_1); -x_14 = lean_box(1); -x_15 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_15, 1, x_3); -return x_15; +x_13 = lean_box(1); +x_14 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_3); +return x_14; } else { -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; -x_16 = lean_unsigned_to_nat(0u); -x_17 = l_Lean_Syntax_getArg(x_9, x_16); +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; +x_15 = lean_unsigned_to_nat(0u); +x_16 = l_Lean_Syntax_getArg(x_9, x_15); lean_dec(x_9); -x_18 = l_Lean_Syntax_getArgs(x_17); -lean_dec(x_17); -x_19 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_19, 0, x_18); -x_20 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__8; -x_21 = lean_box(0); -x_22 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3(x_1, x_20, x_4, x_21, x_19, x_2, x_3); -return x_22; +x_17 = l_Lean_Syntax_getArgs(x_16); +lean_dec(x_16); +x_18 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_18, 0, x_17); +x_19 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__10; +x_20 = lean_box(0); +x_21 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3(x_1, x_19, x_4, x_20, x_18, x_2, x_3); +return x_21; } } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_dec(x_9); -x_23 = lean_box(0); -x_24 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__8; -x_25 = lean_box(0); -x_26 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3(x_1, x_24, x_4, x_25, x_23, x_2, x_3); -return x_26; +x_22 = lean_box(0); +x_23 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__10; +x_24 = lean_box(0); +x_25 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3(x_1, x_23, x_4, x_24, x_22, x_2, x_3); +return x_25; } } } @@ -2112,20 +2161,23 @@ x_6 = lean_unbox_usize(x_4); lean_dec(x_4); x_7 = l_Array_anyMUnsafe_any___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__1(x_1, x_2, x_5, x_6); lean_dec(x_2); +lean_dec(x_1); x_8 = lean_box(x_7); return x_8; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__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) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__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, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_2); +size_t x_10; size_t x_11; lean_object* x_12; +x_10 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_11 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_12 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2(x_1, x_2, x_3, x_4, x_10, x_11, x_7, x_8, x_9); +lean_dec(x_4); lean_dec(x_2); -x_8 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_9 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2(x_1, x_7, x_8, x_4, x_5, x_6); -return x_9; +return x_12; } } LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___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, lean_object* x_9, lean_object* x_10) { @@ -2135,7 +2187,7 @@ lean_object* x_11; x_11 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_7); lean_dec(x_5); -lean_dec(x_1); +lean_dec(x_2); return x_11; } } @@ -2199,7 +2251,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_StructInst_expandStructI _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(37u); +x_1 = lean_unsigned_to_nat(38u); x_2 = lean_unsigned_to_nat(44u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2211,7 +2263,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_StructInst_expandStructI _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(49u); +x_1 = lean_unsigned_to_nat(48u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2239,7 +2291,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_StructInst_expandStructI _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(37u); +x_1 = lean_unsigned_to_nat(38u); x_2 = lean_unsigned_to_nat(48u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2251,7 +2303,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_StructInst_expandStructI _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(37u); +x_1 = lean_unsigned_to_nat(38u); x_2 = lean_unsigned_to_nat(75u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2350,7 +2402,7 @@ static lean_object* _init_l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_Str _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__8; +x_1 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__10; x_2 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources_go___lambda__1___closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -2368,7 +2420,7 @@ static lean_object* _init_l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_Str _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__8; +x_1 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__10; x_2 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources_go___lambda__1___closed__7; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -2386,7 +2438,7 @@ static lean_object* _init_l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_Str _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__8; +x_1 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__10; x_2 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources_go___lambda__1___closed__9; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -2492,14 +2544,14 @@ lean_ctor_set(x_41, 0, x_31); lean_ctor_set(x_41, 1, x_24); lean_ctor_set(x_41, 2, x_40); lean_ctor_set(x_41, 3, x_23); -x_42 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__4; +x_42 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__5; lean_inc(x_31); x_43 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_43, 0, x_31); lean_ctor_set(x_43, 1, x_42); x_44 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources_go___lambda__1___closed__11; x_45 = lean_array_push(x_44, x_41); -x_46 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__3; +x_46 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__4; x_47 = lean_array_push(x_45, x_46); x_48 = lean_array_push(x_47, x_46); x_49 = lean_array_push(x_48, x_43); @@ -2558,14 +2610,14 @@ lean_ctor_set(x_74, 0, x_31); lean_ctor_set(x_74, 1, x_24); lean_ctor_set(x_74, 2, x_73); lean_ctor_set(x_74, 3, x_23); -x_75 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__4; +x_75 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__5; lean_inc(x_31); x_76 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_76, 0, x_31); lean_ctor_set(x_76, 1, x_75); x_77 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources_go___lambda__1___closed__11; x_78 = lean_array_push(x_77, x_74); -x_79 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__3; +x_79 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__4; x_80 = lean_array_push(x_78, x_79); x_81 = lean_array_push(x_80, x_79); x_82 = lean_array_push(x_81, x_76); @@ -3300,7 +3352,7 @@ x_10 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__17; x_11 = lean_array_push(x_10, x_7); x_12 = lean_array_push(x_11, x_9); x_13 = lean_box(2); -x_14 = l_Lean_nullKind; +x_14 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__3; x_15 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_15, 0, x_13); lean_ctor_set(x_15, 1, x_14); @@ -3413,7 +3465,7 @@ static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Struct 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_Array_mapMUnsafe_map___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructSource___spec__3___closed__1; x_2 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructSource___spec__3___closed__2; -x_3 = lean_unsigned_to_nat(116u); +x_3 = lean_unsigned_to_nat(115u); x_4 = lean_unsigned_to_nat(43u); x_5 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructSource___spec__3___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -4142,8 +4194,8 @@ static lean_object* _init_l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Str _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__8; -x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__1; +x_1 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__10; +x_2 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__2; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } @@ -4160,7 +4212,7 @@ static lean_object* _init_l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Str _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__8; +x_1 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__10; x_2 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_isModifyOp_x3f___spec__3___closed__2; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -4738,266 +4790,272 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___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_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, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___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_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, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21) { _start: { -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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; uint8_t x_107; lean_object* x_108; lean_object* x_124; lean_object* x_125; lean_object* x_126; uint8_t x_127; -x_21 = lean_unsigned_to_nat(0u); -x_22 = l_Lean_Syntax_getArg(x_1, 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; uint8_t x_108; lean_object* x_109; lean_object* x_125; lean_object* x_126; lean_object* x_127; uint8_t x_128; +x_22 = lean_unsigned_to_nat(0u); +x_23 = l_Lean_Syntax_getArg(x_1, x_22); lean_dec(x_1); -x_23 = lean_unsigned_to_nat(1u); -x_24 = l_Lean_Syntax_getArg(x_22, x_23); -lean_dec(x_22); -x_25 = l_Lean_Elab_Term_StructInst_instInhabitedExplicitSourceInfo; -x_26 = lean_array_get(x_25, x_2, x_21); +x_24 = lean_unsigned_to_nat(1u); +x_25 = l_Lean_Syntax_getArg(x_23, x_24); +lean_dec(x_23); +x_26 = l_Lean_Elab_Term_StructInst_instInhabitedExplicitSourceInfo; +x_27 = lean_array_get(x_26, x_2, x_22); lean_dec(x_2); -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -lean_dec(x_26); -lean_inc(x_18); -x_28 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_18, x_19, x_20); -x_29 = lean_ctor_get(x_28, 0); -lean_inc(x_29); -x_30 = lean_ctor_get(x_28, 1); +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +lean_dec(x_27); +lean_inc(x_19); +x_29 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_19, x_20, x_21); +x_30 = lean_ctor_get(x_29, 0); lean_inc(x_30); -lean_dec(x_28); -x_31 = lean_ctor_get(x_18, 10); +x_31 = lean_ctor_get(x_29, 1); lean_inc(x_31); -x_32 = lean_st_ref_get(x_19, x_30); -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_32, 1); +lean_dec(x_29); +x_32 = lean_ctor_get(x_19, 10); +lean_inc(x_32); +x_33 = lean_st_ref_get(x_20, x_31); +x_34 = lean_ctor_get(x_33, 0); lean_inc(x_34); -lean_dec(x_32); -x_35 = lean_ctor_get(x_33, 0); +x_35 = lean_ctor_get(x_33, 1); lean_inc(x_35); lean_dec(x_33); -x_36 = lean_environment_main_module(x_35); -x_37 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__1; +x_36 = lean_ctor_get(x_34, 0); +lean_inc(x_36); +lean_dec(x_34); +x_37 = lean_environment_main_module(x_36); +x_38 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__1; lean_inc(x_3); -x_38 = lean_name_mk_string(x_3, x_37); -x_39 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__2; +x_39 = lean_name_mk_string(x_3, x_38); +x_40 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__2; lean_inc(x_3); -x_40 = lean_name_mk_string(x_3, x_39); -x_41 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__3; -lean_inc(x_29); -x_42 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_42, 0, x_29); -lean_ctor_set(x_42, 1, x_41); -x_43 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__7; -lean_inc(x_31); -lean_inc(x_36); -x_44 = l_Lean_addMacroScope(x_36, x_43, x_31); -x_45 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__6; +x_41 = lean_name_mk_string(x_3, x_40); +x_42 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__3; +lean_inc(x_30); +x_43 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_43, 0, x_30); +lean_ctor_set(x_43, 1, x_42); +x_44 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__7; +lean_inc(x_32); +lean_inc(x_37); +x_45 = l_Lean_addMacroScope(x_37, x_44, x_32); +x_46 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__6; lean_inc(x_4); -lean_inc(x_29); -x_46 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_46, 0, x_29); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_44); -lean_ctor_set(x_46, 3, x_4); -x_47 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__20; -x_48 = lean_array_push(x_47, x_27); -x_49 = lean_array_push(x_48, x_42); -x_50 = lean_array_push(x_49, x_46); -x_51 = lean_box(2); -x_52 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_52, 0, x_51); -lean_ctor_set(x_52, 1, x_40); -lean_ctor_set(x_52, 2, x_50); -x_53 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__8; +lean_inc(x_30); +x_47 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_47, 0, x_30); +lean_ctor_set(x_47, 1, x_46); +lean_ctor_set(x_47, 2, x_45); +lean_ctor_set(x_47, 3, x_4); +x_48 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__20; +x_49 = lean_array_push(x_48, x_28); +x_50 = lean_array_push(x_49, x_43); +x_51 = lean_array_push(x_50, x_47); +x_52 = lean_box(2); +x_53 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_41); +lean_ctor_set(x_53, 2, x_51); +x_54 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__8; lean_inc(x_3); -x_54 = lean_name_mk_string(x_3, x_53); -x_55 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__11; -lean_inc(x_29); -x_56 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_56, 0, x_29); -lean_ctor_set(x_56, 1, x_55); -x_57 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__12; -lean_inc(x_31); -lean_inc(x_36); -x_58 = l_Lean_addMacroScope(x_36, x_57, x_31); -x_59 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__11; +x_55 = lean_name_mk_string(x_3, x_54); +x_56 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__13; +lean_inc(x_30); +x_57 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_57, 0, x_30); +lean_ctor_set(x_57, 1, x_56); +x_58 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__12; +lean_inc(x_32); +lean_inc(x_37); +x_59 = l_Lean_addMacroScope(x_37, x_58, x_32); +x_60 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__11; lean_inc(x_4); -lean_inc(x_29); -x_60 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_60, 0, x_29); -lean_ctor_set(x_60, 1, x_59); -lean_ctor_set(x_60, 2, x_58); -lean_ctor_set(x_60, 3, x_4); -x_61 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__4; -lean_inc(x_29); -x_62 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_62, 0, x_29); -lean_ctor_set(x_62, 1, x_61); -x_63 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__19; -lean_inc(x_29); -x_64 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_64, 0, x_29); -lean_ctor_set(x_64, 1, x_63); -x_65 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources_go___lambda__1___closed__11; -lean_inc(x_56); -x_66 = lean_array_push(x_65, x_56); -x_67 = lean_array_push(x_66, x_60); -x_68 = lean_array_push(x_67, x_62); -x_69 = lean_array_push(x_68, x_24); -lean_inc(x_64); -x_70 = lean_array_push(x_69, x_64); -x_71 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_71, 0, x_51); -lean_ctor_set(x_71, 1, x_54); -lean_ctor_set(x_71, 2, x_70); -x_72 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__9; +lean_inc(x_30); +x_61 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_61, 0, x_30); +lean_ctor_set(x_61, 1, x_60); +lean_ctor_set(x_61, 2, x_59); +lean_ctor_set(x_61, 3, x_4); +x_62 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__5; +lean_inc(x_30); +x_63 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_63, 0, x_30); +lean_ctor_set(x_63, 1, x_62); +x_64 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__19; +lean_inc(x_30); +x_65 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_65, 0, x_30); +lean_ctor_set(x_65, 1, x_64); +x_66 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources_go___lambda__1___closed__11; +lean_inc(x_57); +x_67 = lean_array_push(x_66, x_57); +x_68 = lean_array_push(x_67, x_61); +x_69 = lean_array_push(x_68, x_63); +x_70 = lean_array_push(x_69, x_25); +lean_inc(x_65); +x_71 = lean_array_push(x_70, x_65); +x_72 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_72, 0, x_52); +lean_ctor_set(x_72, 1, x_55); +lean_ctor_set(x_72, 2, x_71); +x_73 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__11; lean_inc(x_3); -x_73 = lean_name_mk_string(x_3, x_72); -x_74 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__13; +x_74 = lean_name_mk_string(x_3, x_73); +x_75 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__13; lean_inc(x_3); -x_75 = lean_name_mk_string(x_3, x_74); -lean_inc(x_29); -x_76 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_76, 0, x_29); -lean_ctor_set(x_76, 1, x_74); -x_77 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__14; -x_78 = lean_name_mk_string(x_3, x_77); -x_79 = l_Lean_addMacroScope(x_36, x_5, x_31); -lean_inc(x_29); -x_80 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_80, 0, x_29); -lean_ctor_set(x_80, 1, x_6); -lean_ctor_set(x_80, 2, x_79); -lean_ctor_set(x_80, 3, x_4); -x_81 = lean_array_push(x_7, x_80); -x_82 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__13; +x_76 = lean_name_mk_string(x_3, x_75); +lean_inc(x_30); +x_77 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_77, 0, x_30); +lean_ctor_set(x_77, 1, x_75); +x_78 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__14; +x_79 = lean_name_mk_string(x_3, x_78); +x_80 = l_Lean_addMacroScope(x_37, x_5, x_32); +lean_inc(x_30); +x_81 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_81, 0, x_30); +lean_ctor_set(x_81, 1, x_6); +lean_ctor_set(x_81, 2, x_80); +lean_ctor_set(x_81, 3, x_4); +x_82 = lean_array_push(x_7, x_81); +lean_inc(x_8); x_83 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_83, 0, x_51); -lean_ctor_set(x_83, 1, x_82); -lean_ctor_set(x_83, 2, x_81); +lean_ctor_set(x_83, 0, x_52); +lean_ctor_set(x_83, 1, x_8); +lean_ctor_set(x_83, 2, x_82); x_84 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__15; x_85 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_85, 0, x_29); +lean_ctor_set(x_85, 0, x_30); lean_ctor_set(x_85, 1, x_84); -x_86 = lean_array_push(x_47, x_83); +x_86 = lean_array_push(x_48, x_83); x_87 = lean_array_push(x_86, x_85); -x_88 = lean_array_push(x_87, x_8); +x_88 = lean_array_push(x_87, x_9); x_89 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_89, 0, x_51); -lean_ctor_set(x_89, 1, x_78); +lean_ctor_set(x_89, 0, x_52); +lean_ctor_set(x_89, 1, x_79); lean_ctor_set(x_89, 2, x_88); -lean_inc(x_9); -x_90 = lean_array_push(x_9, x_76); -x_91 = lean_array_push(x_90, x_89); -x_92 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_92, 0, x_51); -lean_ctor_set(x_92, 1, x_75); -lean_ctor_set(x_92, 2, x_91); -lean_inc(x_9); -x_93 = lean_array_push(x_9, x_92); -x_94 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__3; -x_95 = lean_array_push(x_93, x_94); -x_96 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_96, 0, x_51); -lean_ctor_set(x_96, 1, x_82); -lean_ctor_set(x_96, 2, x_95); -x_97 = lean_array_push(x_47, x_56); -x_98 = lean_array_push(x_97, x_96); -x_99 = lean_array_push(x_98, x_64); -x_100 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_100, 0, x_51); -lean_ctor_set(x_100, 1, x_73); -lean_ctor_set(x_100, 2, x_99); -lean_inc(x_9); -x_101 = lean_array_push(x_9, x_71); -x_102 = lean_array_push(x_101, x_100); -x_103 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_103, 0, x_51); -lean_ctor_set(x_103, 1, x_82); -lean_ctor_set(x_103, 2, x_102); -x_104 = lean_array_push(x_9, x_52); -x_105 = lean_array_push(x_104, x_103); -x_106 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_106, 0, x_51); -lean_ctor_set(x_106, 1, x_38); -lean_ctor_set(x_106, 2, x_105); -x_124 = lean_st_ref_get(x_19, x_34); -x_125 = lean_ctor_get(x_124, 0); -lean_inc(x_125); -x_126 = lean_ctor_get(x_125, 3); +lean_inc(x_10); +x_90 = lean_array_push(x_10, x_77); +x_91 = lean_array_push(x_90, x_89); +x_92 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_92, 0, x_52); +lean_ctor_set(x_92, 1, x_76); +lean_ctor_set(x_92, 2, x_91); +x_93 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__1; +lean_inc(x_8); +x_94 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_94, 0, x_52); +lean_ctor_set(x_94, 1, x_8); +lean_ctor_set(x_94, 2, x_93); +lean_inc(x_10); +x_95 = lean_array_push(x_10, x_92); +x_96 = lean_array_push(x_95, x_94); +lean_inc(x_8); +x_97 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_97, 0, x_52); +lean_ctor_set(x_97, 1, x_8); +lean_ctor_set(x_97, 2, x_96); +x_98 = lean_array_push(x_48, x_57); +x_99 = lean_array_push(x_98, x_97); +x_100 = lean_array_push(x_99, x_65); +x_101 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_101, 0, x_52); +lean_ctor_set(x_101, 1, x_74); +lean_ctor_set(x_101, 2, x_100); +lean_inc(x_10); +x_102 = lean_array_push(x_10, x_72); +x_103 = lean_array_push(x_102, x_101); +x_104 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_104, 0, x_52); +lean_ctor_set(x_104, 1, x_8); +lean_ctor_set(x_104, 2, x_103); +x_105 = lean_array_push(x_10, x_53); +x_106 = lean_array_push(x_105, x_104); +x_107 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_107, 0, x_52); +lean_ctor_set(x_107, 1, x_39); +lean_ctor_set(x_107, 2, x_106); +x_125 = lean_st_ref_get(x_20, x_35); +x_126 = lean_ctor_get(x_125, 0); lean_inc(x_126); -lean_dec(x_125); -x_127 = lean_ctor_get_uint8(x_126, sizeof(void*)*1); +x_127 = lean_ctor_get(x_126, 3); +lean_inc(x_127); lean_dec(x_126); -if (x_127 == 0) +x_128 = lean_ctor_get_uint8(x_127, sizeof(void*)*1); +lean_dec(x_127); +if (x_128 == 0) { -lean_object* x_128; uint8_t x_129; -x_128 = lean_ctor_get(x_124, 1); -lean_inc(x_128); -lean_dec(x_124); -x_129 = 0; -x_107 = x_129; -x_108 = x_128; -goto block_123; +lean_object* x_129; uint8_t x_130; +x_129 = lean_ctor_get(x_125, 1); +lean_inc(x_129); +lean_dec(x_125); +x_130 = 0; +x_108 = x_130; +x_109 = x_129; +goto block_124; } else { -lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; uint8_t x_134; -x_130 = lean_ctor_get(x_124, 1); -lean_inc(x_130); -lean_dec(x_124); -lean_inc(x_12); -x_131 = l___private_Lean_Util_Trace_0__Lean_checkTraceOptionM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTerm___spec__2(x_12, x_14, x_15, x_16, x_17, x_18, x_19, x_130); -x_132 = lean_ctor_get(x_131, 0); -lean_inc(x_132); -x_133 = lean_ctor_get(x_131, 1); +lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; uint8_t x_135; +x_131 = lean_ctor_get(x_125, 1); +lean_inc(x_131); +lean_dec(x_125); +lean_inc(x_13); +x_132 = l___private_Lean_Util_Trace_0__Lean_checkTraceOptionM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTerm___spec__2(x_13, x_15, x_16, x_17, x_18, x_19, x_20, x_131); +x_133 = lean_ctor_get(x_132, 0); lean_inc(x_133); -lean_dec(x_131); -x_134 = lean_unbox(x_132); +x_134 = lean_ctor_get(x_132, 1); +lean_inc(x_134); lean_dec(x_132); -x_107 = x_134; -x_108 = x_133; -goto block_123; +x_135 = lean_unbox(x_133); +lean_dec(x_133); +x_108 = x_135; +x_109 = x_134; +goto block_124; } -block_123: +block_124: { -if (x_107 == 0) +if (x_108 == 0) { -lean_object* x_109; lean_object* x_110; -lean_dec(x_12); -x_109 = lean_box(0); -x_110 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__1(x_106, x_10, x_11, x_109, x_14, x_15, x_16, x_17, x_18, x_19, x_108); -return x_110; +lean_object* x_110; lean_object* x_111; +lean_dec(x_13); +x_110 = lean_box(0); +x_111 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__1(x_107, x_11, x_12, x_110, x_15, x_16, x_17, x_18, x_19, x_20, x_109); +return x_111; } else { -lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; -lean_inc(x_11); -x_111 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_111, 0, x_11); -x_112 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__17; -x_113 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_113, 0, x_112); -lean_ctor_set(x_113, 1, x_111); -x_114 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__19; -x_115 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_115, 0, x_113); -lean_ctor_set(x_115, 1, x_114); -lean_inc(x_106); -x_116 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_116, 0, x_106); -x_117 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_117, 0, x_115); -lean_ctor_set(x_117, 1, x_116); +lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; +lean_inc(x_12); +x_112 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_112, 0, x_12); +x_113 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__17; +x_114 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_114, 0, x_113); +lean_ctor_set(x_114, 1, x_112); +x_115 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__19; +x_116 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_116, 0, x_114); +lean_ctor_set(x_116, 1, x_115); +lean_inc(x_107); +x_117 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_117, 0, x_107); x_118 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_118, 0, x_117); -lean_ctor_set(x_118, 1, x_112); -x_119 = l_Lean_addTrace___at___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTerm___spec__1(x_12, x_118, x_14, x_15, x_16, x_17, x_18, x_19, x_108); -x_120 = lean_ctor_get(x_119, 0); -lean_inc(x_120); -x_121 = lean_ctor_get(x_119, 1); +lean_ctor_set(x_118, 0, x_116); +lean_ctor_set(x_118, 1, x_117); +x_119 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_119, 0, x_118); +lean_ctor_set(x_119, 1, x_113); +x_120 = l_Lean_addTrace___at___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTerm___spec__1(x_13, x_119, x_15, x_16, x_17, x_18, x_19, x_20, x_109); +x_121 = lean_ctor_get(x_120, 0); lean_inc(x_121); -lean_dec(x_119); -x_122 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__1(x_106, x_10, x_11, x_120, x_14, x_15, x_16, x_17, x_18, x_19, x_121); +x_122 = lean_ctor_get(x_120, 1); +lean_inc(x_122); lean_dec(x_120); -return x_122; +x_123 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__1(x_107, x_11, x_12, x_121, x_15, x_16, x_17, x_18, x_19, x_20, x_122); +lean_dec(x_121); +return x_123; } } } @@ -5047,8 +5105,8 @@ static lean_object* _init_l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_Str _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__8; -x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__2; +x_1 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__10; +x_2 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } @@ -5112,7 +5170,7 @@ static lean_object* _init_l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_Str _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__8; +x_1 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__10; x_2 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -5122,7 +5180,7 @@ static lean_object* _init_l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_Str _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__8; +x_1 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__10; x_2 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__2; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -5132,7 +5190,7 @@ static lean_object* _init_l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_Str _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__8; +x_1 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__10; x_2 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__8; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -5142,7 +5200,7 @@ static lean_object* _init_l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_Str _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__8; +x_1 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__10; x_2 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__13; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -5152,7 +5210,7 @@ static lean_object* _init_l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_Str _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__8; +x_1 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__10; x_2 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__14; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -5211,7 +5269,7 @@ x_37 = lean_array_get_size(x_36); x_38 = l_Array_toSubarray___rarg(x_36, x_15, x_37); x_39 = l_Array_ofSubarray___rarg(x_38); x_40 = lean_box(2); -x_41 = l_Lean_nullKind; +x_41 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__3; x_42 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_42, 0, x_40); lean_ctor_set(x_42, 1, x_41); @@ -5297,9 +5355,9 @@ goto block_78; if (x_60 == 0) { lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_62 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__8; +x_62 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__10; x_63 = lean_box(0); -x_64 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2(x_14, x_3, x_62, x_29, x_27, x_30, x_43, x_59, x_50, x_4, x_2, x_47, x_63, x_6, x_7, x_8, x_9, x_10, x_11, x_61); +x_64 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2(x_14, x_3, x_62, x_29, x_27, x_30, x_43, x_41, x_59, x_50, x_4, x_2, x_47, x_63, x_6, x_7, x_8, x_9, x_10, x_11, x_61); return x_64; } else @@ -5331,8 +5389,8 @@ lean_inc(x_74); x_75 = lean_ctor_get(x_73, 1); lean_inc(x_75); lean_dec(x_73); -x_76 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__8; -x_77 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2(x_14, x_3, x_76, x_29, x_27, x_30, x_43, x_59, x_50, x_4, x_2, x_47, x_74, x_6, x_7, x_8, x_9, x_10, x_11, x_75); +x_76 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__10; +x_77 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2(x_14, x_3, x_76, x_29, x_27, x_30, x_43, x_41, x_59, x_50, x_4, x_2, x_47, x_74, x_6, x_7, x_8, x_9, x_10, x_11, x_75); lean_dec(x_74); return x_77; } @@ -5402,7 +5460,7 @@ x_120 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_120, 0, x_118); lean_ctor_set(x_120, 1, x_119); lean_ctor_set(x_120, 2, x_117); -x_121 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__11; +x_121 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__13; lean_inc(x_99); x_122 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_122, 0, x_99); @@ -5418,7 +5476,7 @@ lean_ctor_set(x_126, 0, x_99); lean_ctor_set(x_126, 1, x_125); lean_ctor_set(x_126, 2, x_124); lean_ctor_set(x_126, 3, x_111); -x_127 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__4; +x_127 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__5; lean_inc(x_99); x_128 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_128, 0, x_99); @@ -5457,7 +5515,7 @@ lean_ctor_set(x_144, 2, x_142); lean_ctor_set(x_144, 3, x_111); x_145 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__18; x_146 = lean_array_push(x_145, x_144); -x_147 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__13; +x_147 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__3; x_148 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_148, 0, x_118); lean_ctor_set(x_148, 1, x_147); @@ -5483,7 +5541,7 @@ lean_ctor_set(x_160, 0, x_118); lean_ctor_set(x_160, 1, x_159); lean_ctor_set(x_160, 2, x_158); x_161 = lean_array_push(x_156, x_160); -x_162 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__3; +x_162 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__4; x_163 = lean_array_push(x_161, x_162); x_164 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_164, 0, x_118); @@ -5492,7 +5550,7 @@ lean_ctor_set(x_164, 2, x_163); x_165 = lean_array_push(x_114, x_122); x_166 = lean_array_push(x_165, x_164); x_167 = lean_array_push(x_166, x_130); -x_168 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__10; +x_168 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__12; x_169 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_169, 0, x_118); lean_ctor_set(x_169, 1, x_168); @@ -5691,12 +5749,13 @@ lean_object* x_17 = _args[16]; lean_object* x_18 = _args[17]; lean_object* x_19 = _args[18]; lean_object* x_20 = _args[19]; +lean_object* x_21 = _args[20]; _start: { -lean_object* x_21; -x_21 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); -lean_dec(x_13); -return x_21; +lean_object* x_22; +x_22 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20, x_21); +lean_dec(x_14); +return x_22; } } static lean_object* _init_l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructName_throwUnknownExpectedType___closed__1() { @@ -7136,7 +7195,7 @@ static lean_object* _init_l_List_format___at_Lean_Elab_Term_StructInst_formatStr _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__2; +x_1 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__4; x_2 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; @@ -7220,7 +7279,7 @@ static lean_object* _init_l_Lean_Elab_Term_StructInst_formatStruct___closed__2() _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__1; +x_1 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__3; x_2 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; @@ -7230,7 +7289,7 @@ static lean_object* _init_l_Lean_Elab_Term_StructInst_formatStruct___closed__3() _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__3; +x_1 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__6; x_2 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; @@ -7464,6 +7523,24 @@ x_1 = l_Lean_Elab_Term_StructInst_instToStringFieldStruct___closed__1; return x_1; } } +static lean_object* _init_l_Lean_Elab_Term_StructInst_FieldLHS_toSyntax___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("group", 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_StructInst_FieldLHS_toSyntax___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_StructInst_FieldLHS_toSyntax___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_FieldLHS_toSyntax(uint8_t x_1, lean_object* x_2) { _start: { @@ -7486,7 +7563,7 @@ x_8 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__17; x_9 = lean_array_push(x_8, x_6); x_10 = lean_array_push(x_9, x_7); x_11 = lean_box(2); -x_12 = l_Lean_groupKind; +x_12 = l_Lean_Elab_Term_StructInst_FieldLHS_toSyntax___closed__2; x_13 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_13, 0, x_11); lean_ctor_set(x_13, 1, x_12); @@ -7520,7 +7597,7 @@ x_20 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__17; x_21 = lean_array_push(x_20, x_19); x_22 = lean_array_push(x_21, x_17); x_23 = lean_box(2); -x_24 = l_Lean_groupKind; +x_24 = l_Lean_Elab_Term_StructInst_FieldLHS_toSyntax___closed__2; x_25 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_25, 0, x_23); lean_ctor_set(x_25, 1, x_24); @@ -7580,7 +7657,7 @@ static lean_object* _init_l_Lean_Elab_Term_StructInst_FieldVal_toSyntax___closed 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_Array_mapMUnsafe_map___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructSource___spec__3___closed__1; x_2 = l_Lean_Elab_Term_StructInst_FieldVal_toSyntax___closed__1; -x_3 = lean_unsigned_to_nat(322u); +x_3 = lean_unsigned_to_nat(321u); x_4 = lean_unsigned_to_nat(25u); x_5 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructSource___spec__3___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -7648,7 +7725,7 @@ static lean_object* _init_l_Lean_Elab_Term_StructInst_Field_toSyntax___closed__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; x_1 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructSource___spec__3___closed__1; x_2 = l_Lean_Elab_Term_StructInst_Field_toSyntax___closed__1; -x_3 = lean_unsigned_to_nat(330u); +x_3 = lean_unsigned_to_nat(329u); x_4 = lean_unsigned_to_nat(11u); x_5 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructSource___spec__3___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -7697,7 +7774,7 @@ lean_dec(x_17); x_19 = 0; x_20 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_Field_toSyntax___spec__1(x_18, x_19, x_16); x_21 = lean_box(2); -x_22 = l_Lean_nullKind; +x_22 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__3; x_23 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_23, 0, x_21); lean_ctor_set(x_23, 1, x_22); @@ -7746,7 +7823,7 @@ x_6 = lean_name_eq(x_4, x_5); if (x_6 == 0) { lean_object* x_7; uint8_t x_8; -x_7 = l_Lean_groupKind; +x_7 = l_Lean_Elab_Term_StructInst_FieldLHS_toSyntax___closed__2; x_8 = lean_name_eq(x_4, x_7); lean_dec(x_4); if (x_8 == 0) @@ -7755,109 +7832,107 @@ uint8_t x_9; x_9 = l_Lean_Syntax_isIdent(x_1); if (x_9 == 0) { -lean_object* x_10; lean_object* x_11; -x_10 = l_Lean_fieldIdxKind; -x_11 = l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(x_10, x_1); -if (lean_obj_tag(x_11) == 0) +lean_object* x_10; +x_10 = l_Lean_Syntax_isFieldIdx_x3f(x_1); +if (lean_obj_tag(x_10) == 0) { -lean_object* x_12; lean_object* x_13; +lean_object* x_11; lean_object* x_12; lean_dec(x_1); -x_12 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_toFieldLHS___closed__1; -x_13 = l_Lean_Macro_throwError___rarg(x_12, x_2, x_3); -return x_13; +x_11 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_toFieldLHS___closed__1; +x_12 = l_Lean_Macro_throwError___rarg(x_11, x_2, x_3); +return x_12; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_11, 0); -lean_inc(x_14); -lean_dec(x_11); -x_15 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_15, 0, x_1); -lean_ctor_set(x_15, 1, x_14); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_3); -return x_16; +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_10, 0); +lean_inc(x_13); +lean_dec(x_10); +x_14 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_14, 0, x_1); +lean_ctor_set(x_14, 1, x_13); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_3); +return x_15; } } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_17 = l_Lean_Syntax_getId(x_1); -x_18 = lean_erase_macro_scopes(x_17); +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = l_Lean_Syntax_getId(x_1); +x_17 = lean_erase_macro_scopes(x_16); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_1); +lean_ctor_set(x_18, 1, x_17); x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_1); -lean_ctor_set(x_19, 1, x_18); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); -return x_20; +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; } } else { -lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_21 = lean_unsigned_to_nat(1u); -x_22 = l_Lean_Syntax_getArg(x_1, x_21); +lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_20 = lean_unsigned_to_nat(1u); +x_21 = l_Lean_Syntax_getArg(x_1, x_20); lean_dec(x_1); -x_23 = l_Lean_Syntax_isIdent(x_22); -if (x_23 == 0) +x_22 = l_Lean_Syntax_isIdent(x_21); +if (x_22 == 0) { -lean_object* x_24; lean_object* x_25; -x_24 = l_Lean_fieldIdxKind; -x_25 = l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(x_24, x_22); -if (lean_obj_tag(x_25) == 0) +lean_object* x_23; +x_23 = l_Lean_Syntax_isFieldIdx_x3f(x_21); +if (lean_obj_tag(x_23) == 0) { -lean_object* x_26; lean_object* x_27; -lean_dec(x_22); -x_26 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_toFieldLHS___closed__1; -x_27 = l_Lean_Macro_throwError___rarg(x_26, x_2, x_3); -return x_27; +lean_object* x_24; lean_object* x_25; +lean_dec(x_21); +x_24 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_toFieldLHS___closed__1; +x_25 = l_Lean_Macro_throwError___rarg(x_24, x_2, x_3); +return x_25; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_25, 0); -lean_inc(x_28); -lean_dec(x_25); -x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_22); -lean_ctor_set(x_29, 1, x_28); -x_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_3); -return x_30; +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_23, 0); +lean_inc(x_26); +lean_dec(x_23); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_21); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_3); +return x_28; } } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_31 = l_Lean_Syntax_getId(x_22); -x_32 = lean_erase_macro_scopes(x_31); -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_22); -lean_ctor_set(x_33, 1, x_32); -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_3); -return x_34; +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_29 = l_Lean_Syntax_getId(x_21); +x_30 = lean_erase_macro_scopes(x_29); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_21); +lean_ctor_set(x_31, 1, x_30); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_3); +return x_32; } } } else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_dec(x_4); -x_35 = lean_unsigned_to_nat(1u); -x_36 = l_Lean_Syntax_getArg(x_1, x_35); -x_37 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_37, 0, x_1); -lean_ctor_set(x_37, 1, x_36); -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_37); -lean_ctor_set(x_38, 1, x_3); -return x_38; +x_33 = lean_unsigned_to_nat(1u); +x_34 = l_Lean_Syntax_getArg(x_1, x_33); +x_35 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_35, 0, x_1); +lean_ctor_set(x_35, 1, x_34); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_3); +return x_36; } } } @@ -10261,7 +10336,7 @@ static lean_object* _init_l_List_mapTRAux___at___private_Lean_Elab_StructInst_0_ 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_Array_mapMUnsafe_map___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructSource___spec__3___closed__1; x_2 = l_List_mapTRAux___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__2___closed__1; -x_3 = lean_unsigned_to_nat(429u); +x_3 = lean_unsigned_to_nat(428u); x_4 = lean_unsigned_to_nat(34u); x_5 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructSource___spec__3___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -11361,7 +11436,7 @@ static lean_object* _init_l_List_foldlM___at___private_Lean_Elab_StructInst_0__L 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_Array_mapMUnsafe_map___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructSource___spec__3___closed__1; x_2 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkFieldMap___spec__13___closed__1; -x_3 = lean_unsigned_to_nat(447u); +x_3 = lean_unsigned_to_nat(446u); x_4 = lean_unsigned_to_nat(11u); x_5 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructSource___spec__3___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -12775,18 +12850,6 @@ x_1 = l_Lean_Elab_Term_StructInst_instInhabitedField(lean_box(0)); return x_1; } } -static lean_object* _init_l_List_mapM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_groupFields___spec__8___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(2); -x_2 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__2; -x_3 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} -} LEAN_EXPORT lean_object* l_List_mapM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_groupFields___spec__8(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_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { @@ -12845,7 +12908,7 @@ if (lean_obj_tag(x_58) == 0) { lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; size_t x_66; size_t 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; x_59 = lean_unsigned_to_nat(4u); -x_60 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__2; +x_60 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__4; lean_inc(x_3); x_61 = l_Lean_Syntax_setArg(x_3, x_59, x_60); x_62 = l_List_redLength___rarg(x_36); @@ -12857,11 +12920,11 @@ x_66 = lean_usize_of_nat(x_65); lean_dec(x_65); x_67 = 0; x_68 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_groupFields___spec__2(x_66, x_67, x_64); -x_69 = l_List_mapM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_groupFields___spec__8___closed__2; +x_69 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__5; x_70 = l_Lean_mkSepArray(x_68, x_69); lean_dec(x_68); x_71 = lean_box(2); -x_72 = l_Lean_nullKind; +x_72 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__3; x_73 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_73, 0, x_71); lean_ctor_set(x_73, 1, x_72); @@ -15873,7 +15936,7 @@ static lean_object* _init_l_List_foldlM___at___private_Lean_Elab_StructInst_0__L _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__8; +x_1 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__10; x_2 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -15887,42 +15950,6 @@ x_1 = lean_mk_string_from_bytes("by", 2); return x_1; } } -static lean_object* _init_l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__7() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Tactic", 6); -return x_1; -} -} -static lean_object* _init_l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__6; -x_2 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__7; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__9() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("tacticSeq", 9); -return x_1; -} -} -static lean_object* _init_l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__10() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__8; -x_2 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__9; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___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, 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, lean_object* x_16, lean_object* x_17, lean_object* x_18) { _start: { @@ -17215,7 +17242,7 @@ return x_307; } else { -lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; uint8_t x_334; lean_object* x_335; +lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; uint8_t x_330; lean_object* x_331; x_308 = lean_ctor_get(x_303, 0); lean_inc(x_308); lean_dec(x_303); @@ -17234,138 +17261,131 @@ x_314 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_Stru x_315 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_315, 0, x_310); lean_ctor_set(x_315, 1, x_314); -x_316 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__18; -x_317 = lean_array_push(x_316, x_308); -x_318 = lean_box(2); -x_319 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__10; -x_320 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_320, 0, x_318); -lean_ctor_set(x_320, 1, x_319); -lean_ctor_set(x_320, 2, x_317); -x_321 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__17; -x_322 = lean_array_push(x_321, x_315); -x_323 = lean_array_push(x_322, x_320); -x_324 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__5; -x_325 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_325, 0, x_318); -lean_ctor_set(x_325, 1, x_324); -lean_ctor_set(x_325, 2, x_323); -x_326 = lean_unsigned_to_nat(0u); -x_327 = l_Lean_Expr_getAppNumArgsAux(x_251, x_326); -x_328 = lean_nat_sub(x_327, x_326); -lean_dec(x_327); -x_329 = lean_unsigned_to_nat(1u); -x_330 = lean_nat_sub(x_328, x_329); -lean_dec(x_328); -x_331 = l_Lean_Expr_getRevArg_x21(x_251, x_330); -x_332 = lean_expr_consume_type_annotations(x_331); -lean_ctor_set(x_253, 0, x_332); -x_333 = lean_box(0); -x_334 = 1; +x_316 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__17; +x_317 = lean_array_push(x_316, x_315); +x_318 = lean_array_push(x_317, x_308); +x_319 = lean_box(2); +x_320 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__5; +x_321 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_321, 0, x_319); +lean_ctor_set(x_321, 1, x_320); +lean_ctor_set(x_321, 2, x_318); +x_322 = lean_unsigned_to_nat(0u); +x_323 = l_Lean_Expr_getAppNumArgsAux(x_251, x_322); +x_324 = lean_nat_sub(x_323, x_322); +lean_dec(x_323); +x_325 = lean_unsigned_to_nat(1u); +x_326 = lean_nat_sub(x_324, x_325); +lean_dec(x_324); +x_327 = l_Lean_Expr_getRevArg_x21(x_251, x_326); +x_328 = lean_expr_consume_type_annotations(x_327); +lean_ctor_set(x_253, 0, x_328); +x_329 = lean_box(0); +x_330 = 1; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -x_335 = l_Lean_Elab_Term_elabTermEnsuringType(x_325, x_253, x_334, x_334, x_333, x_12, x_13, x_14, x_15, x_16, x_17, x_313); -if (lean_obj_tag(x_335) == 0) -{ -lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; uint8_t x_345; -x_336 = lean_ctor_get(x_335, 0); -lean_inc(x_336); -x_337 = lean_ctor_get(x_335, 1); -lean_inc(x_337); -lean_dec(x_335); -x_338 = lean_ctor_get(x_14, 1); -lean_inc(x_338); +x_331 = l_Lean_Elab_Term_elabTermEnsuringType(x_321, x_253, x_330, x_330, x_329, x_12, x_13, x_14, x_15, x_16, x_17, x_313); +if (lean_obj_tag(x_331) == 0) +{ +lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; uint8_t x_341; +x_332 = lean_ctor_get(x_331, 0); +lean_inc(x_332); +x_333 = lean_ctor_get(x_331, 1); +lean_inc(x_333); +lean_dec(x_331); +x_334 = lean_ctor_get(x_14, 1); +lean_inc(x_334); lean_inc(x_3); -x_339 = l_Lean_Name_append(x_4, x_3); +x_335 = l_Lean_Name_append(x_4, x_3); lean_dec(x_4); -lean_inc(x_336); -x_340 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_340, 0, x_339); -lean_ctor_set(x_340, 1, x_3); -lean_ctor_set(x_340, 2, x_338); -lean_ctor_set(x_340, 3, x_336); -lean_ctor_set(x_340, 4, x_6); -x_341 = lean_alloc_ctor(4, 1, 0); -lean_ctor_set(x_341, 0, x_340); -x_342 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__3; -x_343 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_343, 0, x_341); -lean_ctor_set(x_343, 1, x_342); -x_344 = l_Lean_Elab_pushInfoTree___at_Lean_Elab_Term_addDotCompletionInfo___spec__3(x_343, x_12, x_13, x_14, x_15, x_16, x_17, x_337); +lean_inc(x_332); +x_336 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_336, 0, x_335); +lean_ctor_set(x_336, 1, x_3); +lean_ctor_set(x_336, 2, x_334); +lean_ctor_set(x_336, 3, x_332); +lean_ctor_set(x_336, 4, x_6); +x_337 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_337, 0, x_336); +x_338 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__3; +x_339 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_339, 0, x_337); +lean_ctor_set(x_339, 1, x_338); +x_340 = l_Lean_Elab_pushInfoTree___at_Lean_Elab_Term_addDotCompletionInfo___spec__3(x_339, x_12, x_13, x_14, x_15, x_16, x_17, x_333); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); -x_345 = !lean_is_exclusive(x_344); -if (x_345 == 0) -{ -lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; -x_346 = lean_ctor_get(x_344, 0); -lean_dec(x_346); -lean_inc(x_336); -x_347 = l_Lean_mkApp(x_7, x_336); -x_348 = lean_expr_instantiate1(x_252, x_336); +x_341 = !lean_is_exclusive(x_340); +if (x_341 == 0) +{ +lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; +x_342 = lean_ctor_get(x_340, 0); +lean_dec(x_342); +lean_inc(x_332); +x_343 = l_Lean_mkApp(x_7, x_332); +x_344 = lean_expr_instantiate1(x_252, x_332); lean_dec(x_252); -x_349 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_349, 0, x_336); -x_350 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_350, 0, x_2); -lean_ctor_set(x_350, 1, x_8); -lean_ctor_set(x_350, 2, x_5); -lean_ctor_set(x_350, 3, x_349); -x_351 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_351, 0, x_350); -lean_ctor_set(x_351, 1, x_9); -x_352 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_352, 0, x_348); -lean_ctor_set(x_352, 1, x_351); -x_353 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_353, 0, x_347); -lean_ctor_set(x_353, 1, x_352); -lean_ctor_set(x_344, 0, x_353); -return x_344; -} -else -{ -lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; -x_354 = lean_ctor_get(x_344, 1); -lean_inc(x_354); -lean_dec(x_344); -lean_inc(x_336); -x_355 = l_Lean_mkApp(x_7, x_336); -x_356 = lean_expr_instantiate1(x_252, x_336); +x_345 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_345, 0, x_332); +x_346 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_346, 0, x_2); +lean_ctor_set(x_346, 1, x_8); +lean_ctor_set(x_346, 2, x_5); +lean_ctor_set(x_346, 3, x_345); +x_347 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_347, 0, x_346); +lean_ctor_set(x_347, 1, x_9); +x_348 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_348, 0, x_344); +lean_ctor_set(x_348, 1, x_347); +x_349 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_349, 0, x_343); +lean_ctor_set(x_349, 1, x_348); +lean_ctor_set(x_340, 0, x_349); +return x_340; +} +else +{ +lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; +x_350 = lean_ctor_get(x_340, 1); +lean_inc(x_350); +lean_dec(x_340); +lean_inc(x_332); +x_351 = l_Lean_mkApp(x_7, x_332); +x_352 = lean_expr_instantiate1(x_252, x_332); lean_dec(x_252); -x_357 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_357, 0, x_336); -x_358 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_358, 0, x_2); -lean_ctor_set(x_358, 1, x_8); -lean_ctor_set(x_358, 2, x_5); -lean_ctor_set(x_358, 3, x_357); -x_359 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_359, 0, x_358); -lean_ctor_set(x_359, 1, x_9); -x_360 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_360, 0, x_356); -lean_ctor_set(x_360, 1, x_359); -x_361 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_361, 0, x_355); -lean_ctor_set(x_361, 1, x_360); -x_362 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_362, 0, x_361); -lean_ctor_set(x_362, 1, x_354); -return x_362; -} -} -else -{ -uint8_t x_363; +x_353 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_353, 0, x_332); +x_354 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_354, 0, x_2); +lean_ctor_set(x_354, 1, x_8); +lean_ctor_set(x_354, 2, x_5); +lean_ctor_set(x_354, 3, x_353); +x_355 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_355, 0, x_354); +lean_ctor_set(x_355, 1, x_9); +x_356 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_356, 0, x_352); +lean_ctor_set(x_356, 1, x_355); +x_357 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_357, 0, x_351); +lean_ctor_set(x_357, 1, x_356); +x_358 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_358, 0, x_357); +lean_ctor_set(x_358, 1, x_350); +return x_358; +} +} +else +{ +uint8_t x_359; lean_dec(x_252); lean_dec(x_17); lean_dec(x_16); @@ -17380,186 +17400,186 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_363 = !lean_is_exclusive(x_335); -if (x_363 == 0) +x_359 = !lean_is_exclusive(x_331); +if (x_359 == 0) { -return x_335; +return x_331; } else { -lean_object* x_364; lean_object* x_365; lean_object* x_366; -x_364 = lean_ctor_get(x_335, 0); -x_365 = lean_ctor_get(x_335, 1); -lean_inc(x_365); -lean_inc(x_364); -lean_dec(x_335); -x_366 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_366, 0, x_364); -lean_ctor_set(x_366, 1, x_365); -return x_366; +lean_object* x_360; lean_object* x_361; lean_object* x_362; +x_360 = lean_ctor_get(x_331, 0); +x_361 = lean_ctor_get(x_331, 1); +lean_inc(x_361); +lean_inc(x_360); +lean_dec(x_331); +x_362 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_362, 0, x_360); +lean_ctor_set(x_362, 1, x_361); +return x_362; } } } } else { -lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; uint8_t x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; uint8_t x_393; +lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; uint8_t x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; uint8_t x_389; lean_dec(x_300); lean_dec(x_10); lean_ctor_set(x_253, 0, x_251); -x_367 = lean_ctor_get(x_16, 0); +x_363 = lean_ctor_get(x_16, 0); +lean_inc(x_363); +x_364 = lean_ctor_get(x_16, 1); +lean_inc(x_364); +x_365 = lean_ctor_get(x_16, 2); +lean_inc(x_365); +x_366 = lean_ctor_get(x_16, 3); +lean_inc(x_366); +x_367 = lean_ctor_get(x_16, 4); lean_inc(x_367); -x_368 = lean_ctor_get(x_16, 1); +x_368 = lean_ctor_get(x_16, 5); lean_inc(x_368); -x_369 = lean_ctor_get(x_16, 2); +x_369 = lean_ctor_get(x_16, 6); lean_inc(x_369); -x_370 = lean_ctor_get(x_16, 3); +x_370 = lean_ctor_get(x_16, 7); lean_inc(x_370); -x_371 = lean_ctor_get(x_16, 4); +x_371 = lean_ctor_get(x_16, 8); lean_inc(x_371); -x_372 = lean_ctor_get(x_16, 5); +x_372 = lean_ctor_get(x_16, 9); lean_inc(x_372); -x_373 = lean_ctor_get(x_16, 6); +x_373 = lean_ctor_get(x_16, 10); lean_inc(x_373); -x_374 = lean_ctor_get(x_16, 7); -lean_inc(x_374); -x_375 = lean_ctor_get(x_16, 8); -lean_inc(x_375); -x_376 = lean_ctor_get(x_16, 9); -lean_inc(x_376); -x_377 = lean_ctor_get(x_16, 10); -lean_inc(x_377); -x_378 = l_Lean_replaceRef(x_2, x_372); -lean_dec(x_372); -x_379 = lean_alloc_ctor(0, 11, 0); -lean_ctor_set(x_379, 0, x_367); -lean_ctor_set(x_379, 1, x_368); -lean_ctor_set(x_379, 2, x_369); -lean_ctor_set(x_379, 3, x_370); -lean_ctor_set(x_379, 4, x_371); -lean_ctor_set(x_379, 5, x_378); -lean_ctor_set(x_379, 6, x_373); -lean_ctor_set(x_379, 7, x_374); -lean_ctor_set(x_379, 8, x_375); -lean_ctor_set(x_379, 9, x_376); -lean_ctor_set(x_379, 10, x_377); -x_380 = 0; -x_381 = lean_box(0); +x_374 = l_Lean_replaceRef(x_2, x_368); +lean_dec(x_368); +x_375 = lean_alloc_ctor(0, 11, 0); +lean_ctor_set(x_375, 0, x_363); +lean_ctor_set(x_375, 1, x_364); +lean_ctor_set(x_375, 2, x_365); +lean_ctor_set(x_375, 3, x_366); +lean_ctor_set(x_375, 4, x_367); +lean_ctor_set(x_375, 5, x_374); +lean_ctor_set(x_375, 6, x_369); +lean_ctor_set(x_375, 7, x_370); +lean_ctor_set(x_375, 8, x_371); +lean_ctor_set(x_375, 9, x_372); +lean_ctor_set(x_375, 10, x_373); +x_376 = 0; +x_377 = lean_box(0); lean_inc(x_14); -x_382 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarImpl(x_253, x_380, x_381, x_14, x_15, x_379, x_17, x_18); -lean_dec(x_379); -x_383 = lean_ctor_get(x_382, 0); -lean_inc(x_383); -x_384 = lean_ctor_get(x_382, 1); -lean_inc(x_384); -lean_dec(x_382); -x_385 = l_Lean_Elab_Term_StructInst_markDefaultMissing(x_383); -x_386 = lean_ctor_get(x_14, 1); -lean_inc(x_386); +x_378 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarImpl(x_253, x_376, x_377, x_14, x_15, x_375, x_17, x_18); +lean_dec(x_375); +x_379 = lean_ctor_get(x_378, 0); +lean_inc(x_379); +x_380 = lean_ctor_get(x_378, 1); +lean_inc(x_380); +lean_dec(x_378); +x_381 = l_Lean_Elab_Term_StructInst_markDefaultMissing(x_379); +x_382 = lean_ctor_get(x_14, 1); +lean_inc(x_382); lean_inc(x_3); -x_387 = l_Lean_Name_append(x_4, x_3); +x_383 = l_Lean_Name_append(x_4, x_3); lean_dec(x_4); -lean_inc(x_385); -x_388 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_388, 0, x_387); -lean_ctor_set(x_388, 1, x_3); -lean_ctor_set(x_388, 2, x_386); -lean_ctor_set(x_388, 3, x_385); -lean_ctor_set(x_388, 4, x_6); -x_389 = lean_alloc_ctor(4, 1, 0); -lean_ctor_set(x_389, 0, x_388); -x_390 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__3; -x_391 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_391, 0, x_389); -lean_ctor_set(x_391, 1, x_390); -x_392 = l_Lean_Elab_pushInfoTree___at_Lean_Elab_Term_addDotCompletionInfo___spec__3(x_391, x_12, x_13, x_14, x_15, x_16, x_17, x_384); +lean_inc(x_381); +x_384 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_384, 0, x_383); +lean_ctor_set(x_384, 1, x_3); +lean_ctor_set(x_384, 2, x_382); +lean_ctor_set(x_384, 3, x_381); +lean_ctor_set(x_384, 4, x_6); +x_385 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_385, 0, x_384); +x_386 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__3; +x_387 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_387, 0, x_385); +lean_ctor_set(x_387, 1, x_386); +x_388 = l_Lean_Elab_pushInfoTree___at_Lean_Elab_Term_addDotCompletionInfo___spec__3(x_387, x_12, x_13, x_14, x_15, x_16, x_17, x_380); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); -x_393 = !lean_is_exclusive(x_392); -if (x_393 == 0) -{ -lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; -x_394 = lean_ctor_get(x_392, 0); -lean_dec(x_394); -lean_inc(x_385); -x_395 = l_Lean_mkApp(x_7, x_385); -x_396 = lean_expr_instantiate1(x_252, x_385); +x_389 = !lean_is_exclusive(x_388); +if (x_389 == 0) +{ +lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; +x_390 = lean_ctor_get(x_388, 0); +lean_dec(x_390); +lean_inc(x_381); +x_391 = l_Lean_mkApp(x_7, x_381); +x_392 = lean_expr_instantiate1(x_252, x_381); lean_dec(x_252); -x_397 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_397, 0, x_385); -x_398 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_398, 0, x_2); -lean_ctor_set(x_398, 1, x_8); -lean_ctor_set(x_398, 2, x_5); -lean_ctor_set(x_398, 3, x_397); -x_399 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_399, 0, x_398); -lean_ctor_set(x_399, 1, x_9); -x_400 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_400, 0, x_396); -lean_ctor_set(x_400, 1, x_399); -x_401 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_401, 0, x_395); -lean_ctor_set(x_401, 1, x_400); -lean_ctor_set(x_392, 0, x_401); -return x_392; -} -else -{ -lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; -x_402 = lean_ctor_get(x_392, 1); -lean_inc(x_402); -lean_dec(x_392); -lean_inc(x_385); -x_403 = l_Lean_mkApp(x_7, x_385); -x_404 = lean_expr_instantiate1(x_252, x_385); +x_393 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_393, 0, x_381); +x_394 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_394, 0, x_2); +lean_ctor_set(x_394, 1, x_8); +lean_ctor_set(x_394, 2, x_5); +lean_ctor_set(x_394, 3, x_393); +x_395 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_395, 0, x_394); +lean_ctor_set(x_395, 1, x_9); +x_396 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_396, 0, x_392); +lean_ctor_set(x_396, 1, x_395); +x_397 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_397, 0, x_391); +lean_ctor_set(x_397, 1, x_396); +lean_ctor_set(x_388, 0, x_397); +return x_388; +} +else +{ +lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; +x_398 = lean_ctor_get(x_388, 1); +lean_inc(x_398); +lean_dec(x_388); +lean_inc(x_381); +x_399 = l_Lean_mkApp(x_7, x_381); +x_400 = lean_expr_instantiate1(x_252, x_381); lean_dec(x_252); -x_405 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_405, 0, x_385); -x_406 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_406, 0, x_2); -lean_ctor_set(x_406, 1, x_8); -lean_ctor_set(x_406, 2, x_5); -lean_ctor_set(x_406, 3, x_405); -x_407 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_407, 0, x_406); -lean_ctor_set(x_407, 1, x_9); -x_408 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_408, 0, x_404); -lean_ctor_set(x_408, 1, x_407); -x_409 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_409, 0, x_403); -lean_ctor_set(x_409, 1, x_408); -x_410 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_410, 0, x_409); -lean_ctor_set(x_410, 1, x_402); -return x_410; -} -} -} -else -{ -lean_object* x_411; -x_411 = lean_ctor_get(x_253, 0); -lean_inc(x_411); +x_401 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_401, 0, x_381); +x_402 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_402, 0, x_2); +lean_ctor_set(x_402, 1, x_8); +lean_ctor_set(x_402, 2, x_5); +lean_ctor_set(x_402, 3, x_401); +x_403 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_403, 0, x_402); +lean_ctor_set(x_403, 1, x_9); +x_404 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_404, 0, x_400); +lean_ctor_set(x_404, 1, x_403); +x_405 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_405, 0, x_399); +lean_ctor_set(x_405, 1, x_404); +x_406 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_406, 0, x_405); +lean_ctor_set(x_406, 1, x_398); +return x_406; +} +} +} +else +{ +lean_object* x_407; +x_407 = lean_ctor_get(x_253, 0); +lean_inc(x_407); lean_dec(x_253); -if (lean_obj_tag(x_411) == 4) -{ -lean_object* x_412; lean_object* x_413; lean_object* x_414; -x_412 = lean_ctor_get(x_411, 0); -lean_inc(x_412); -lean_dec(x_411); -x_413 = lean_ctor_get(x_16, 2); -lean_inc(x_413); -x_414 = l___private_Lean_Elab_Util_0__Lean_Elab_evalSyntaxConstantUnsafe(x_10, x_413, x_412); -lean_dec(x_413); -if (lean_obj_tag(x_414) == 0) -{ -lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; +if (lean_obj_tag(x_407) == 4) +{ +lean_object* x_408; lean_object* x_409; lean_object* x_410; +x_408 = lean_ctor_get(x_407, 0); +lean_inc(x_408); +lean_dec(x_407); +x_409 = lean_ctor_get(x_16, 2); +lean_inc(x_409); +x_410 = l___private_Lean_Elab_Util_0__Lean_Elab_evalSyntaxConstantUnsafe(x_10, x_409, x_408); +lean_dec(x_409); +if (lean_obj_tag(x_410) == 0) +{ +lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_dec(x_252); lean_dec(x_251); lean_dec(x_9); @@ -17569,153 +17589,146 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_415 = lean_ctor_get(x_414, 0); -lean_inc(x_415); -lean_dec(x_414); -x_416 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_416, 0, x_415); -x_417 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_417, 0, x_416); -x_418 = l_Lean_throwError___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__2(x_417, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +x_411 = lean_ctor_get(x_410, 0); +lean_inc(x_411); +lean_dec(x_410); +x_412 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_412, 0, x_411); +x_413 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_413, 0, x_412); +x_414 = l_Lean_throwError___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__2(x_413, x_12, x_13, x_14, x_15, x_16, x_17, x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); -return x_418; +return x_414; } else { -lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; uint8_t x_446; lean_object* x_447; -x_419 = lean_ctor_get(x_414, 0); -lean_inc(x_419); -lean_dec(x_414); +lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; uint8_t x_438; lean_object* x_439; +x_415 = lean_ctor_get(x_410, 0); +lean_inc(x_415); +lean_dec(x_410); lean_inc(x_16); -x_420 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_16, x_17, x_18); -x_421 = lean_ctor_get(x_420, 0); -lean_inc(x_421); -x_422 = lean_ctor_get(x_420, 1); -lean_inc(x_422); -lean_dec(x_420); -x_423 = lean_st_ref_get(x_17, x_422); -x_424 = lean_ctor_get(x_423, 1); -lean_inc(x_424); -lean_dec(x_423); -x_425 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__6; -x_426 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_426, 0, x_421); -lean_ctor_set(x_426, 1, x_425); -x_427 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__18; -x_428 = lean_array_push(x_427, x_419); -x_429 = lean_box(2); -x_430 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__10; -x_431 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_431, 0, x_429); -lean_ctor_set(x_431, 1, x_430); -lean_ctor_set(x_431, 2, x_428); -x_432 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__17; -x_433 = lean_array_push(x_432, x_426); -x_434 = lean_array_push(x_433, x_431); -x_435 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__5; -x_436 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_436, 0, x_429); -lean_ctor_set(x_436, 1, x_435); -lean_ctor_set(x_436, 2, x_434); -x_437 = lean_unsigned_to_nat(0u); -x_438 = l_Lean_Expr_getAppNumArgsAux(x_251, x_437); -x_439 = lean_nat_sub(x_438, x_437); -lean_dec(x_438); -x_440 = lean_unsigned_to_nat(1u); -x_441 = lean_nat_sub(x_439, x_440); -lean_dec(x_439); -x_442 = l_Lean_Expr_getRevArg_x21(x_251, x_441); -x_443 = lean_expr_consume_type_annotations(x_442); -x_444 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_444, 0, x_443); -x_445 = lean_box(0); -x_446 = 1; +x_416 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_16, x_17, x_18); +x_417 = lean_ctor_get(x_416, 0); +lean_inc(x_417); +x_418 = lean_ctor_get(x_416, 1); +lean_inc(x_418); +lean_dec(x_416); +x_419 = lean_st_ref_get(x_17, x_418); +x_420 = lean_ctor_get(x_419, 1); +lean_inc(x_420); +lean_dec(x_419); +x_421 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__6; +x_422 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_422, 0, x_417); +lean_ctor_set(x_422, 1, x_421); +x_423 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__17; +x_424 = lean_array_push(x_423, x_422); +x_425 = lean_array_push(x_424, x_415); +x_426 = lean_box(2); +x_427 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__5; +x_428 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_428, 0, x_426); +lean_ctor_set(x_428, 1, x_427); +lean_ctor_set(x_428, 2, x_425); +x_429 = lean_unsigned_to_nat(0u); +x_430 = l_Lean_Expr_getAppNumArgsAux(x_251, x_429); +x_431 = lean_nat_sub(x_430, x_429); +lean_dec(x_430); +x_432 = lean_unsigned_to_nat(1u); +x_433 = lean_nat_sub(x_431, x_432); +lean_dec(x_431); +x_434 = l_Lean_Expr_getRevArg_x21(x_251, x_433); +x_435 = lean_expr_consume_type_annotations(x_434); +x_436 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_436, 0, x_435); +x_437 = lean_box(0); +x_438 = 1; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -x_447 = l_Lean_Elab_Term_elabTermEnsuringType(x_436, x_444, x_446, x_446, x_445, x_12, x_13, x_14, x_15, x_16, x_17, x_424); -if (lean_obj_tag(x_447) == 0) -{ -lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; -x_448 = lean_ctor_get(x_447, 0); -lean_inc(x_448); -x_449 = lean_ctor_get(x_447, 1); -lean_inc(x_449); -lean_dec(x_447); -x_450 = lean_ctor_get(x_14, 1); -lean_inc(x_450); +x_439 = l_Lean_Elab_Term_elabTermEnsuringType(x_428, x_436, x_438, x_438, x_437, x_12, x_13, x_14, x_15, x_16, x_17, x_420); +if (lean_obj_tag(x_439) == 0) +{ +lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; +x_440 = lean_ctor_get(x_439, 0); +lean_inc(x_440); +x_441 = lean_ctor_get(x_439, 1); +lean_inc(x_441); +lean_dec(x_439); +x_442 = lean_ctor_get(x_14, 1); +lean_inc(x_442); lean_inc(x_3); -x_451 = l_Lean_Name_append(x_4, x_3); +x_443 = l_Lean_Name_append(x_4, x_3); lean_dec(x_4); -lean_inc(x_448); -x_452 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_452, 0, x_451); -lean_ctor_set(x_452, 1, x_3); -lean_ctor_set(x_452, 2, x_450); -lean_ctor_set(x_452, 3, x_448); -lean_ctor_set(x_452, 4, x_6); -x_453 = lean_alloc_ctor(4, 1, 0); -lean_ctor_set(x_453, 0, x_452); -x_454 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__3; -x_455 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_455, 0, x_453); -lean_ctor_set(x_455, 1, x_454); -x_456 = l_Lean_Elab_pushInfoTree___at_Lean_Elab_Term_addDotCompletionInfo___spec__3(x_455, x_12, x_13, x_14, x_15, x_16, x_17, x_449); +lean_inc(x_440); +x_444 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_444, 0, x_443); +lean_ctor_set(x_444, 1, x_3); +lean_ctor_set(x_444, 2, x_442); +lean_ctor_set(x_444, 3, x_440); +lean_ctor_set(x_444, 4, x_6); +x_445 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_445, 0, x_444); +x_446 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__3; +x_447 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_447, 0, x_445); +lean_ctor_set(x_447, 1, x_446); +x_448 = l_Lean_Elab_pushInfoTree___at_Lean_Elab_Term_addDotCompletionInfo___spec__3(x_447, x_12, x_13, x_14, x_15, x_16, x_17, x_441); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); -x_457 = lean_ctor_get(x_456, 1); -lean_inc(x_457); -if (lean_is_exclusive(x_456)) { - lean_ctor_release(x_456, 0); - lean_ctor_release(x_456, 1); - x_458 = x_456; +x_449 = lean_ctor_get(x_448, 1); +lean_inc(x_449); +if (lean_is_exclusive(x_448)) { + lean_ctor_release(x_448, 0); + lean_ctor_release(x_448, 1); + x_450 = x_448; } else { - lean_dec_ref(x_456); - x_458 = lean_box(0); + lean_dec_ref(x_448); + x_450 = lean_box(0); } -lean_inc(x_448); -x_459 = l_Lean_mkApp(x_7, x_448); -x_460 = lean_expr_instantiate1(x_252, x_448); +lean_inc(x_440); +x_451 = l_Lean_mkApp(x_7, x_440); +x_452 = lean_expr_instantiate1(x_252, x_440); lean_dec(x_252); -x_461 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_461, 0, x_448); -x_462 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_462, 0, x_2); -lean_ctor_set(x_462, 1, x_8); -lean_ctor_set(x_462, 2, x_5); -lean_ctor_set(x_462, 3, x_461); -x_463 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_463, 0, x_462); -lean_ctor_set(x_463, 1, x_9); -x_464 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_464, 0, x_460); -lean_ctor_set(x_464, 1, x_463); -x_465 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_465, 0, x_459); -lean_ctor_set(x_465, 1, x_464); -if (lean_is_scalar(x_458)) { - x_466 = lean_alloc_ctor(0, 2, 0); +x_453 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_453, 0, x_440); +x_454 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_454, 0, x_2); +lean_ctor_set(x_454, 1, x_8); +lean_ctor_set(x_454, 2, x_5); +lean_ctor_set(x_454, 3, x_453); +x_455 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_455, 0, x_454); +lean_ctor_set(x_455, 1, x_9); +x_456 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_456, 0, x_452); +lean_ctor_set(x_456, 1, x_455); +x_457 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_457, 0, x_451); +lean_ctor_set(x_457, 1, x_456); +if (lean_is_scalar(x_450)) { + x_458 = lean_alloc_ctor(0, 2, 0); } else { - x_466 = x_458; + x_458 = x_450; } -lean_ctor_set(x_466, 0, x_465); -lean_ctor_set(x_466, 1, x_457); -return x_466; +lean_ctor_set(x_458, 0, x_457); +lean_ctor_set(x_458, 1, x_449); +return x_458; } else { -lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; +lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_dec(x_252); lean_dec(x_17); lean_dec(x_16); @@ -17730,146 +17743,146 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_467 = lean_ctor_get(x_447, 0); -lean_inc(x_467); -x_468 = lean_ctor_get(x_447, 1); -lean_inc(x_468); -if (lean_is_exclusive(x_447)) { - lean_ctor_release(x_447, 0); - lean_ctor_release(x_447, 1); - x_469 = x_447; +x_459 = lean_ctor_get(x_439, 0); +lean_inc(x_459); +x_460 = lean_ctor_get(x_439, 1); +lean_inc(x_460); +if (lean_is_exclusive(x_439)) { + lean_ctor_release(x_439, 0); + lean_ctor_release(x_439, 1); + x_461 = x_439; } else { - lean_dec_ref(x_447); - x_469 = lean_box(0); + lean_dec_ref(x_439); + x_461 = lean_box(0); } -if (lean_is_scalar(x_469)) { - x_470 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_461)) { + x_462 = lean_alloc_ctor(1, 2, 0); } else { - x_470 = x_469; + x_462 = x_461; } -lean_ctor_set(x_470, 0, x_467); -lean_ctor_set(x_470, 1, x_468); -return x_470; +lean_ctor_set(x_462, 0, x_459); +lean_ctor_set(x_462, 1, x_460); +return x_462; } } } else { -lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; uint8_t x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; -lean_dec(x_411); +lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; uint8_t x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; +lean_dec(x_407); lean_dec(x_10); -x_471 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_471, 0, x_251); -x_472 = lean_ctor_get(x_16, 0); +x_463 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_463, 0, x_251); +x_464 = lean_ctor_get(x_16, 0); +lean_inc(x_464); +x_465 = lean_ctor_get(x_16, 1); +lean_inc(x_465); +x_466 = lean_ctor_get(x_16, 2); +lean_inc(x_466); +x_467 = lean_ctor_get(x_16, 3); +lean_inc(x_467); +x_468 = lean_ctor_get(x_16, 4); +lean_inc(x_468); +x_469 = lean_ctor_get(x_16, 5); +lean_inc(x_469); +x_470 = lean_ctor_get(x_16, 6); +lean_inc(x_470); +x_471 = lean_ctor_get(x_16, 7); +lean_inc(x_471); +x_472 = lean_ctor_get(x_16, 8); lean_inc(x_472); -x_473 = lean_ctor_get(x_16, 1); +x_473 = lean_ctor_get(x_16, 9); lean_inc(x_473); -x_474 = lean_ctor_get(x_16, 2); +x_474 = lean_ctor_get(x_16, 10); lean_inc(x_474); -x_475 = lean_ctor_get(x_16, 3); -lean_inc(x_475); -x_476 = lean_ctor_get(x_16, 4); -lean_inc(x_476); -x_477 = lean_ctor_get(x_16, 5); -lean_inc(x_477); -x_478 = lean_ctor_get(x_16, 6); -lean_inc(x_478); -x_479 = lean_ctor_get(x_16, 7); -lean_inc(x_479); -x_480 = lean_ctor_get(x_16, 8); +x_475 = l_Lean_replaceRef(x_2, x_469); +lean_dec(x_469); +x_476 = lean_alloc_ctor(0, 11, 0); +lean_ctor_set(x_476, 0, x_464); +lean_ctor_set(x_476, 1, x_465); +lean_ctor_set(x_476, 2, x_466); +lean_ctor_set(x_476, 3, x_467); +lean_ctor_set(x_476, 4, x_468); +lean_ctor_set(x_476, 5, x_475); +lean_ctor_set(x_476, 6, x_470); +lean_ctor_set(x_476, 7, x_471); +lean_ctor_set(x_476, 8, x_472); +lean_ctor_set(x_476, 9, x_473); +lean_ctor_set(x_476, 10, x_474); +x_477 = 0; +x_478 = lean_box(0); +lean_inc(x_14); +x_479 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarImpl(x_463, x_477, x_478, x_14, x_15, x_476, x_17, x_18); +lean_dec(x_476); +x_480 = lean_ctor_get(x_479, 0); lean_inc(x_480); -x_481 = lean_ctor_get(x_16, 9); +x_481 = lean_ctor_get(x_479, 1); lean_inc(x_481); -x_482 = lean_ctor_get(x_16, 10); -lean_inc(x_482); -x_483 = l_Lean_replaceRef(x_2, x_477); -lean_dec(x_477); -x_484 = lean_alloc_ctor(0, 11, 0); -lean_ctor_set(x_484, 0, x_472); -lean_ctor_set(x_484, 1, x_473); -lean_ctor_set(x_484, 2, x_474); -lean_ctor_set(x_484, 3, x_475); -lean_ctor_set(x_484, 4, x_476); -lean_ctor_set(x_484, 5, x_483); -lean_ctor_set(x_484, 6, x_478); -lean_ctor_set(x_484, 7, x_479); -lean_ctor_set(x_484, 8, x_480); -lean_ctor_set(x_484, 9, x_481); -lean_ctor_set(x_484, 10, x_482); -x_485 = 0; -x_486 = lean_box(0); -lean_inc(x_14); -x_487 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarImpl(x_471, x_485, x_486, x_14, x_15, x_484, x_17, x_18); -lean_dec(x_484); -x_488 = lean_ctor_get(x_487, 0); -lean_inc(x_488); -x_489 = lean_ctor_get(x_487, 1); -lean_inc(x_489); -lean_dec(x_487); -x_490 = l_Lean_Elab_Term_StructInst_markDefaultMissing(x_488); -x_491 = lean_ctor_get(x_14, 1); -lean_inc(x_491); +lean_dec(x_479); +x_482 = l_Lean_Elab_Term_StructInst_markDefaultMissing(x_480); +x_483 = lean_ctor_get(x_14, 1); +lean_inc(x_483); lean_inc(x_3); -x_492 = l_Lean_Name_append(x_4, x_3); +x_484 = l_Lean_Name_append(x_4, x_3); lean_dec(x_4); -lean_inc(x_490); -x_493 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_493, 0, x_492); -lean_ctor_set(x_493, 1, x_3); -lean_ctor_set(x_493, 2, x_491); -lean_ctor_set(x_493, 3, x_490); -lean_ctor_set(x_493, 4, x_6); -x_494 = lean_alloc_ctor(4, 1, 0); -lean_ctor_set(x_494, 0, x_493); -x_495 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__3; -x_496 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_496, 0, x_494); -lean_ctor_set(x_496, 1, x_495); -x_497 = l_Lean_Elab_pushInfoTree___at_Lean_Elab_Term_addDotCompletionInfo___spec__3(x_496, x_12, x_13, x_14, x_15, x_16, x_17, x_489); +lean_inc(x_482); +x_485 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_485, 0, x_484); +lean_ctor_set(x_485, 1, x_3); +lean_ctor_set(x_485, 2, x_483); +lean_ctor_set(x_485, 3, x_482); +lean_ctor_set(x_485, 4, x_6); +x_486 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_486, 0, x_485); +x_487 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__3; +x_488 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_488, 0, x_486); +lean_ctor_set(x_488, 1, x_487); +x_489 = l_Lean_Elab_pushInfoTree___at_Lean_Elab_Term_addDotCompletionInfo___spec__3(x_488, x_12, x_13, x_14, x_15, x_16, x_17, x_481); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); -x_498 = lean_ctor_get(x_497, 1); -lean_inc(x_498); -if (lean_is_exclusive(x_497)) { - lean_ctor_release(x_497, 0); - lean_ctor_release(x_497, 1); - x_499 = x_497; +x_490 = lean_ctor_get(x_489, 1); +lean_inc(x_490); +if (lean_is_exclusive(x_489)) { + lean_ctor_release(x_489, 0); + lean_ctor_release(x_489, 1); + x_491 = x_489; } else { - lean_dec_ref(x_497); - x_499 = lean_box(0); + lean_dec_ref(x_489); + x_491 = lean_box(0); } -lean_inc(x_490); -x_500 = l_Lean_mkApp(x_7, x_490); -x_501 = lean_expr_instantiate1(x_252, x_490); +lean_inc(x_482); +x_492 = l_Lean_mkApp(x_7, x_482); +x_493 = lean_expr_instantiate1(x_252, x_482); lean_dec(x_252); -x_502 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_502, 0, x_490); -x_503 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_503, 0, x_2); -lean_ctor_set(x_503, 1, x_8); -lean_ctor_set(x_503, 2, x_5); -lean_ctor_set(x_503, 3, x_502); -x_504 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_504, 0, x_503); -lean_ctor_set(x_504, 1, x_9); -x_505 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_505, 0, x_501); -lean_ctor_set(x_505, 1, x_504); -x_506 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_506, 0, x_500); -lean_ctor_set(x_506, 1, x_505); -if (lean_is_scalar(x_499)) { - x_507 = lean_alloc_ctor(0, 2, 0); +x_494 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_494, 0, x_482); +x_495 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_495, 0, x_2); +lean_ctor_set(x_495, 1, x_8); +lean_ctor_set(x_495, 2, x_5); +lean_ctor_set(x_495, 3, x_494); +x_496 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_496, 0, x_495); +lean_ctor_set(x_496, 1, x_9); +x_497 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_497, 0, x_493); +lean_ctor_set(x_497, 1, x_496); +x_498 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_498, 0, x_492); +lean_ctor_set(x_498, 1, x_497); +if (lean_is_scalar(x_491)) { + x_499 = lean_alloc_ctor(0, 2, 0); } else { - x_507 = x_499; + x_499 = x_491; } -lean_ctor_set(x_507, 0, x_506); -lean_ctor_set(x_507, 1, x_498); -return x_507; +lean_ctor_set(x_499, 0, x_498); +lean_ctor_set(x_499, 1, x_490); +return x_499; } } } @@ -17878,83 +17891,83 @@ return x_507; } else { -lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; uint8_t x_513; +lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; uint8_t x_505; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_508 = l_Lean_indentExpr(x_1); -x_509 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkCtorHeaderAux___closed__2; -x_510 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_510, 0, x_509); -lean_ctor_set(x_510, 1, x_508); -x_511 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__17; -x_512 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_512, 0, x_510); -lean_ctor_set(x_512, 1, x_511); -x_513 = !lean_is_exclusive(x_16); -if (x_513 == 0) -{ -lean_object* x_514; lean_object* x_515; lean_object* x_516; -x_514 = lean_ctor_get(x_16, 5); -x_515 = l_Lean_replaceRef(x_2, x_514); -lean_dec(x_514); +x_500 = l_Lean_indentExpr(x_1); +x_501 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkCtorHeaderAux___closed__2; +x_502 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_502, 0, x_501); +lean_ctor_set(x_502, 1, x_500); +x_503 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__17; +x_504 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_504, 0, x_502); +lean_ctor_set(x_504, 1, x_503); +x_505 = !lean_is_exclusive(x_16); +if (x_505 == 0) +{ +lean_object* x_506; lean_object* x_507; lean_object* x_508; +x_506 = lean_ctor_get(x_16, 5); +x_507 = l_Lean_replaceRef(x_2, x_506); +lean_dec(x_506); lean_dec(x_2); -lean_ctor_set(x_16, 5, x_515); -x_516 = l_Lean_Elab_Term_StructInst_throwFailedToElabField___rarg(x_3, x_4, x_512, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_ctor_set(x_16, 5, x_507); +x_508 = l_Lean_Elab_Term_StructInst_throwFailedToElabField___rarg(x_3, x_4, x_504, x_12, x_13, x_14, x_15, x_16, x_17, x_18); lean_dec(x_17); lean_dec(x_15); lean_dec(x_13); -return x_516; -} -else -{ -lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; -x_517 = lean_ctor_get(x_16, 0); -x_518 = lean_ctor_get(x_16, 1); -x_519 = lean_ctor_get(x_16, 2); -x_520 = lean_ctor_get(x_16, 3); -x_521 = lean_ctor_get(x_16, 4); -x_522 = lean_ctor_get(x_16, 5); -x_523 = lean_ctor_get(x_16, 6); -x_524 = lean_ctor_get(x_16, 7); -x_525 = lean_ctor_get(x_16, 8); -x_526 = lean_ctor_get(x_16, 9); -x_527 = lean_ctor_get(x_16, 10); -lean_inc(x_527); -lean_inc(x_526); -lean_inc(x_525); -lean_inc(x_524); -lean_inc(x_523); -lean_inc(x_522); -lean_inc(x_521); -lean_inc(x_520); +return x_508; +} +else +{ +lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; +x_509 = lean_ctor_get(x_16, 0); +x_510 = lean_ctor_get(x_16, 1); +x_511 = lean_ctor_get(x_16, 2); +x_512 = lean_ctor_get(x_16, 3); +x_513 = lean_ctor_get(x_16, 4); +x_514 = lean_ctor_get(x_16, 5); +x_515 = lean_ctor_get(x_16, 6); +x_516 = lean_ctor_get(x_16, 7); +x_517 = lean_ctor_get(x_16, 8); +x_518 = lean_ctor_get(x_16, 9); +x_519 = lean_ctor_get(x_16, 10); lean_inc(x_519); lean_inc(x_518); lean_inc(x_517); +lean_inc(x_516); +lean_inc(x_515); +lean_inc(x_514); +lean_inc(x_513); +lean_inc(x_512); +lean_inc(x_511); +lean_inc(x_510); +lean_inc(x_509); lean_dec(x_16); -x_528 = l_Lean_replaceRef(x_2, x_522); -lean_dec(x_522); +x_520 = l_Lean_replaceRef(x_2, x_514); +lean_dec(x_514); lean_dec(x_2); -x_529 = lean_alloc_ctor(0, 11, 0); -lean_ctor_set(x_529, 0, x_517); -lean_ctor_set(x_529, 1, x_518); -lean_ctor_set(x_529, 2, x_519); -lean_ctor_set(x_529, 3, x_520); -lean_ctor_set(x_529, 4, x_521); -lean_ctor_set(x_529, 5, x_528); -lean_ctor_set(x_529, 6, x_523); -lean_ctor_set(x_529, 7, x_524); -lean_ctor_set(x_529, 8, x_525); -lean_ctor_set(x_529, 9, x_526); -lean_ctor_set(x_529, 10, x_527); -x_530 = l_Lean_Elab_Term_StructInst_throwFailedToElabField___rarg(x_3, x_4, x_512, x_12, x_13, x_14, x_15, x_529, x_17, x_18); +x_521 = lean_alloc_ctor(0, 11, 0); +lean_ctor_set(x_521, 0, x_509); +lean_ctor_set(x_521, 1, x_510); +lean_ctor_set(x_521, 2, x_511); +lean_ctor_set(x_521, 3, x_512); +lean_ctor_set(x_521, 4, x_513); +lean_ctor_set(x_521, 5, x_520); +lean_ctor_set(x_521, 6, x_515); +lean_ctor_set(x_521, 7, x_516); +lean_ctor_set(x_521, 8, x_517); +lean_ctor_set(x_521, 9, x_518); +lean_ctor_set(x_521, 10, x_519); +x_522 = l_Lean_Elab_Term_StructInst_throwFailedToElabField___rarg(x_3, x_4, x_504, x_12, x_13, x_14, x_15, x_521, x_17, x_18); lean_dec(x_17); lean_dec(x_15); lean_dec(x_13); -return x_530; +return x_522; } } } @@ -18919,7 +18932,7 @@ static lean_object* _init_l_Lean_Elab_Term_StructInst_DefaultFields_findDefaultM 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_Array_mapMUnsafe_map___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructSource___spec__3___closed__1; x_2 = l_Lean_Elab_Term_StructInst_DefaultFields_findDefaultMissing_x3f___lambda__1___closed__1; -x_3 = lean_unsigned_to_nat(707u); +x_3 = lean_unsigned_to_nat(706u); x_4 = lean_unsigned_to_nat(20u); x_5 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructSource___spec__3___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -19091,7 +19104,7 @@ static lean_object* _init_l_Lean_Elab_Term_StructInst_DefaultFields_getFieldName 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_Array_mapMUnsafe_map___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructSource___spec__3___closed__1; x_2 = l_Lean_Elab_Term_StructInst_DefaultFields_getFieldName___closed__1; -x_3 = lean_unsigned_to_nat(715u); +x_3 = lean_unsigned_to_nat(714u); x_4 = lean_unsigned_to_nat(9u); x_5 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructSource___spec__3___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -21994,7 +22007,7 @@ static lean_object* _init_l_List_forIn_loop___at_Lean_Elab_Term_StructInst_Defau 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_Array_mapMUnsafe_map___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructSource___spec__3___closed__1; x_2 = l_List_forIn_loop___at_Lean_Elab_Term_StructInst_DefaultFields_step___spec__2___closed__1; -x_3 = lean_unsigned_to_nat(831u); +x_3 = lean_unsigned_to_nat(830u); x_4 = lean_unsigned_to_nat(25u); x_5 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructSource___spec__3___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -25148,7 +25161,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_StructInst_elabStructIns _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(895u); +x_1 = lean_unsigned_to_nat(894u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -25160,7 +25173,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_StructInst_elabStructIns _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(903u); +x_1 = lean_unsigned_to_nat(902u); x_2 = lean_unsigned_to_nat(94u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -25188,7 +25201,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_StructInst_elabStructIns _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(895u); +x_1 = lean_unsigned_to_nat(894u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -25200,7 +25213,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_StructInst_elabStructIns _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(895u); +x_1 = lean_unsigned_to_nat(894u); x_2 = lean_unsigned_to_nat(48u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -25246,7 +25259,7 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_10107_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_10239_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; @@ -25371,8 +25384,6 @@ lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_StructInst_expandStructInstEx res = l___regBuiltin_Lean_Elab_Term_StructInst_expandStructInstExpectedType_declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Array_anyMUnsafe_any___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__1___closed__1 = _init_l_Array_anyMUnsafe_any___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__1___closed__1(); -lean_mark_persistent(l_Array_anyMUnsafe_any___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__1___closed__1); l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__1 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__1(); lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__1); l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__2 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__2(); @@ -25381,6 +25392,8 @@ l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbre lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__3); l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__4 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__4(); lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__4); +l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__5 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__5(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__5); l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__1 = _init_l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__1); l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__2 = _init_l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__2(); @@ -25397,6 +25410,12 @@ l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__7 lean_mark_persistent(l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__7); l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__8 = _init_l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__8(); lean_mark_persistent(l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__8); +l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__9 = _init_l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__9(); +lean_mark_persistent(l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__9); +l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__10 = _init_l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__10(); +lean_mark_persistent(l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__10); +l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__11 = _init_l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__11(); +lean_mark_persistent(l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__11); l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3___closed__1 = _init_l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3___closed__1); l___regBuiltin_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___closed__1(); @@ -25676,6 +25695,10 @@ l_Lean_Elab_Term_StructInst_instToStringFieldStruct___closed__1 = _init_l_Lean_E lean_mark_persistent(l_Lean_Elab_Term_StructInst_instToStringFieldStruct___closed__1); l_Lean_Elab_Term_StructInst_instToStringFieldStruct = _init_l_Lean_Elab_Term_StructInst_instToStringFieldStruct(); lean_mark_persistent(l_Lean_Elab_Term_StructInst_instToStringFieldStruct); +l_Lean_Elab_Term_StructInst_FieldLHS_toSyntax___closed__1 = _init_l_Lean_Elab_Term_StructInst_FieldLHS_toSyntax___closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_StructInst_FieldLHS_toSyntax___closed__1); +l_Lean_Elab_Term_StructInst_FieldLHS_toSyntax___closed__2 = _init_l_Lean_Elab_Term_StructInst_FieldLHS_toSyntax___closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_StructInst_FieldLHS_toSyntax___closed__2); l_Lean_Elab_Term_StructInst_FieldVal_toSyntax___closed__1 = _init_l_Lean_Elab_Term_StructInst_FieldVal_toSyntax___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_StructInst_FieldVal_toSyntax___closed__1); l_Lean_Elab_Term_StructInst_FieldVal_toSyntax___closed__2 = _init_l_Lean_Elab_Term_StructInst_FieldVal_toSyntax___closed__2(); @@ -25738,8 +25761,6 @@ l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getFieldIdx___clos lean_mark_persistent(l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getFieldIdx___closed__2); l_List_mapM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_groupFields___spec__8___closed__1 = _init_l_List_mapM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_groupFields___spec__8___closed__1(); lean_mark_persistent(l_List_mapM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_groupFields___spec__8___closed__1); -l_List_mapM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_groupFields___spec__8___closed__2 = _init_l_List_mapM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_groupFields___spec__8___closed__2(); -lean_mark_persistent(l_List_mapM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_groupFields___spec__8___closed__2); l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_addMissingFields___spec__3___closed__1 = _init_l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_addMissingFields___spec__3___closed__1(); lean_mark_persistent(l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_addMissingFields___spec__3___closed__1); l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_addMissingFields___closed__1 = _init_l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_addMissingFields___closed__1(); @@ -25774,14 +25795,6 @@ l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_e lean_mark_persistent(l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__5); l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__6 = _init_l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__6(); lean_mark_persistent(l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__6); -l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__7 = _init_l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__7(); -lean_mark_persistent(l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__7); -l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__8 = _init_l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__8(); -lean_mark_persistent(l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__8); -l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__9 = _init_l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__9(); -lean_mark_persistent(l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__9); -l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__10 = _init_l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__10(); -lean_mark_persistent(l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___lambda__1___closed__10); l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___closed__1 = _init_l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___closed__1(); lean_mark_persistent(l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___closed__1); l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___closed__2 = _init_l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__3___closed__2(); @@ -25900,7 +25913,7 @@ lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_StructInst_elabStructInst_dec res = l___regBuiltin_Lean_Elab_Term_StructInst_elabStructInst_declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -res = l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_10107_(lean_io_mk_world()); +res = l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_10239_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/Structure.c b/stage0/stdlib/Lean/Elab/Structure.c index 147bdf30c6c..065197ff827 100644 --- a/stage0/stdlib/Lean/Elab/Structure.c +++ b/stage0/stdlib/Lean/Elab/Structure.c @@ -67,7 +67,6 @@ static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStr lean_object* l_Lean_LocalDecl_userName(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_collectUniversesFromFields_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__6___closed__7; -extern lean_object* l_Lean_nullKind; lean_object* l_Lean_NameSet_contains___boxed(lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__3___closed__13; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParamFVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -137,11 +136,11 @@ lean_object* l_Lean_Expr_getAutoParamTactic_x3f(lean_object*); static lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__15___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withParents_go___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_522____closed__25; +uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_522____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_processSubfields_go___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__15___lambda__2___closed__1; lean_object* lean_environment_find(lean_object*, lean_object*); -lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___spec__4(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___closed__3; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_522____closed__1; @@ -262,7 +261,6 @@ static lean_object* l_Lean_Elab_Command_checkValidFieldModifier___lambda__3___cl LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabStructure___closed__13; -lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___spec__3(size_t, size_t, lean_object*); lean_object* l_Lean_Meta_mkAuxDefinition(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_327____closed__22; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -337,6 +335,7 @@ static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_registe lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_Options___hyg_6____spec__1(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkFreshId___at_Lean_Meta_mkFreshExprMVarAt___spec__2___rarg(lean_object*, lean_object*); lean_object* l_Lean_setReducibilityStatus___at_Lean_mkNoConfusionEnum_mkToCtorIdx___spec__14(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabImplicitLambda_loop___spec__1___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Elab_Structure_0__Lean_Elab_Command_validStructType(lean_object*); lean_object* l_Lean_Elab_expandMacroImpl_x3f(lean_object*, lean_object*, lean_object*, lean_object*); @@ -402,6 +401,7 @@ lean_object* l_Lean_replaceRef(lean_object*, lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ResolveName_resolveNamespace(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_copyFields_copy(lean_object*); +lean_object* l_Array_sequenceMap___at_Lean_Elab_elabDeriving___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withUsed(lean_object*); lean_object* l_Lean_Meta_mkLambdaFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_copyFields_copy___rarg___closed__1; @@ -442,6 +442,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_co LEAN_EXPORT lean_object* l_Lean_Elab_Command_StructFieldInfo_value_x3f___default; static lean_object* l_Lean_Elab_elabModifiers___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__1___lambda__3___closed__7; lean_object* l_Lean_Elab_Term_addLocalVarInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_elabDeriving___spec__4(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_getFieldType___lambda__2___closed__2; lean_object* l_Lean_mkConstWithLevelParams___at___private_Lean_Elab_PreDefinition_Basic_0__Lean_Elab_addNonRecAux___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_copyFields_copy___rarg___closed__2; @@ -512,6 +513,7 @@ static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure lean_object* l_Lean_Meta_reduceProjOf_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_checkValidFieldModifier___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___spec__4(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Elab_resolveGlobalConstNoOverloadWithInfo___at_Lean_Elab_elabDeriving___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__6___closed__2; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___closed__1; static lean_object* l_Lean_Elab_Command_elabStructure___closed__5; @@ -535,6 +537,7 @@ lean_object* l_Lean_mkNoConfusion(lean_object*, lean_object*, lean_object*, lean lean_object* l_Lean_Meta_getLocalInstances(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabStructure___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_panic___at_Lean_ResolveName_resolveNamespaceUsingScope_x3f___spec__1(lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2512____closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_StructFieldKind_toCtorIdx(uint8_t); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_327____closed__1; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_collectUniversesFromFields_go___spec__2___closed__10; @@ -545,6 +548,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_re static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabFieldTypeValue___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ResolveName_resolveGlobalName(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2512____closed__4; static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__4___closed__7; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkAuxConstructions___closed__5; lean_object* l_Lean_throwKernelException___at_Lean_Elab_Term_evalExpr___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -610,7 +614,6 @@ static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_copyFields_copy___rarg___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_522_(lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___spec__11___closed__2; -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalDecl_type(lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__8___closed__2; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_327____closed__18; @@ -628,7 +631,6 @@ static lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___private_Le LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___spec__7___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_mkCompositeField___spec__1___closed__2; lean_object* l_Lean_Name_getString_x21(lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2488____closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields___rarg(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_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_collectLevelParamsInStructure___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_collectUsed___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -655,7 +657,6 @@ static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStr static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___closed__6; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_environment_main_module(lean_object*); -lean_object* l_Lean_Elab_resolveGlobalConstNoOverloadWithInfo___at_Lean_Elab_elabDeriving___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_expr_eqv(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withParents_go___rarg___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, 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_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -675,7 +676,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_co static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__1___closed__4; lean_object* l_Lean_getDefaultFnForField_x3f(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__4___closed__8; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2488____closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withParents_go___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_327____closed__9; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___spec__5(lean_object*, size_t, size_t, lean_object*); @@ -685,6 +685,7 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabStr static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___lambda__4___closed__4; static lean_object* l_Lean_Elab_Command_elabStructure___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__4___lambda__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_elabDeriving___spec__5(size_t, size_t, lean_object*); lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandDeclId___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getProjectionFnInfo_x3f___at_Lean_Meta_unfoldProjInst_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_BinderInfo_isExplicit(uint8_t); @@ -749,7 +750,6 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Stru LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_collectUniversesFromFields_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_expandDeclSig(lean_object*); static lean_object* l_Lean_Elab_Command_checkValidFieldModifier___closed__2; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2488____closed__4; lean_object* l_Lean_Elab_Term_elabType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___spec__9___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -801,8 +801,8 @@ static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabFieldTypeValue___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyDefaultValue_x3f_go_x3f___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_copyFields_copy___rarg___lambda__5___boxed(lean_object**); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_11179_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2488_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_11219_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2512_(lean_object*); lean_object* l_Lean_Meta_instantiateMVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_TagDeclarationExtension_isTagged(lean_object*, lean_object*, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); @@ -832,7 +832,6 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_ lean_object* l_Lean_Meta_getStructureName(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__15___closed__1; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___spec__1___closed__4; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2488____closed__1; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_522____closed__7; lean_object* lean_string_length(lean_object*); lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_nameToPattern___spec__1___rarg(lean_object*, lean_object*, lean_object*); @@ -854,7 +853,6 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Stru uint8_t l_Lean_Elab_Modifiers_isPrivate(lean_object*); uint8_t l_Lean_NameSet_contains(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___rarg___lambda__4___closed__5; -lean_object* l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_327____closed__21; lean_object* l_Lean_Elab_Term_addAutoBoundImplicits(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_whnfD(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -934,7 +932,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_checkValidFieldModifier(lean_object LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_collectLevelParamsInFVars___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_withAutoBoundImplicit___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___spec__6(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkConst(lean_object*, lean_object*); @@ -957,6 +954,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_co LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_findFieldInfo_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_processOveriddenDefaultValues___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabModifiers___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__1___closed__2; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2512____closed__3; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_copyFields_copy___rarg___closed__7; static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__1___closed__3; lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(lean_object*, size_t, size_t, lean_object*); @@ -967,6 +965,7 @@ lean_object* l_Lean_throwError___at_Lean_Elab_Term_synthesizeInst___spec__1(lean static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_327____closed__13; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_getFieldType___lambda__5___closed__5; lean_object* l_Lean_setEnv___at_Lean_Elab_Term_evalExpr___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2512____closed__1; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_327____closed__12; LEAN_EXPORT uint8_t l___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabFieldTypeValue___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabStructure___lambda__3___boxed(lean_object**); @@ -6419,7 +6418,7 @@ static lean_object* _init_l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Str { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean_nullKind; +x_2 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__3___closed__6; x_3 = l_Lean_Elab_elabAttrs___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__3___closed__1; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -7677,7 +7676,7 @@ x_48 = l_Lean_Syntax_getArg(x_15, x_47); x_49 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___closed__4; x_50 = lean_array_push(x_49, x_48); x_51 = lean_box(2); -x_52 = l_Lean_nullKind; +x_52 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__3___closed__6; x_53 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_53, 0, x_51); lean_ctor_set(x_53, 1, x_52); @@ -7893,7 +7892,7 @@ x_48 = l_Lean_Syntax_getArg(x_15, x_47); x_49 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___closed__4; x_50 = lean_array_push(x_49, x_48); x_51 = lean_box(2); -x_52 = l_Lean_nullKind; +x_52 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__3___closed__6; x_53 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_53, 0, x_51); lean_ctor_set(x_53, 1, x_52); @@ -8637,7 +8636,7 @@ lean_dec(x_1); return x_8; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2488____closed__1() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2512____closed__1() { _start: { lean_object* x_1; @@ -8645,17 +8644,17 @@ x_1 = lean_mk_string_from_bytes("structureDiamondWarning", 23); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2488____closed__2() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2512____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2488____closed__1; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2512____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2488____closed__3() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2512____closed__3() { _start: { lean_object* x_1; @@ -8663,13 +8662,13 @@ x_1 = lean_mk_string_from_bytes("enable/disable warning messages for structure d return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2488____closed__4() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2512____closed__4() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = 0; x_2 = l_Lean_Elab_elabModifiers___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__1___closed__3; -x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2488____closed__3; +x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2512____closed__3; x_4 = lean_box(x_1); x_5 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_5, 0, x_4); @@ -8678,12 +8677,12 @@ lean_ctor_set(x_5, 2, x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2488_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2512_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2488____closed__2; -x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2488____closed__4; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2512____closed__2; +x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2512____closed__4; x_4 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_Options___hyg_6____spec__1(x_2, x_3, x_1); return x_4; } @@ -11491,7 +11490,7 @@ static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_c 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___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_mkCompositeField___closed__1; x_2 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_mkCompositeField___closed__2; -x_3 = lean_unsigned_to_nat(448u); +x_3 = lean_unsigned_to_nat(449u); x_4 = lean_unsigned_to_nat(70u); x_5 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_mkCompositeField___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -11966,7 +11965,7 @@ static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_c 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___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_mkCompositeField___closed__1; x_2 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_copyFields_copy___rarg___closed__1; -x_3 = lean_unsigned_to_nat(409u); +x_3 = lean_unsigned_to_nat(410u); x_4 = lean_unsigned_to_nat(86u); x_5 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_mkCompositeField___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -12836,6 +12835,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mk 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; uint8_t x_10; x_3 = lean_erase_macro_scopes(x_1); x_4 = l_Lean_Name_getString_x21(x_3); +lean_dec(x_3); x_5 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkToParentName___closed__1; x_6 = lean_string_append(x_5, x_4); lean_dec(x_4); @@ -17964,7 +17964,7 @@ LEAN_EXPORT uint8_t l___private_Lean_Elab_Structure_0__Lean_Elab_Command_levelMV lean_object* x_3; uint8_t x_4; x_3 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_3, 0, x_2); -x_4 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___spec__1(x_1, x_3); +x_4 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___spec__1(x_1, x_3); lean_dec(x_3); return x_4; } @@ -20812,7 +20812,7 @@ static lean_object* _init_l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Str 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___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_mkCompositeField___closed__1; x_2 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_registerStructure___spec__2___closed__1; -x_3 = lean_unsigned_to_nat(716u); +x_3 = lean_unsigned_to_nat(717u); x_4 = lean_unsigned_to_nat(21u); x_5 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_registerStructure___spec__2___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -20833,7 +20833,7 @@ static lean_object* _init_l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Str 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___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_mkCompositeField___closed__1; x_2 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_registerStructure___spec__2___closed__1; -x_3 = lean_unsigned_to_nat(715u); +x_3 = lean_unsigned_to_nat(716u); x_4 = lean_unsigned_to_nat(23u); x_5 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_registerStructure___spec__2___closed__4; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -22988,7 +22988,7 @@ static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Stru 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___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_mkCompositeField___closed__1; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___spec__1___closed__1; -x_3 = lean_unsigned_to_nat(763u); +x_3 = lean_unsigned_to_nat(764u); x_4 = lean_unsigned_to_nat(79u); x_5 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_mkCompositeField___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -23297,7 +23297,7 @@ static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_m 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___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_mkCompositeField___closed__1; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___spec__1___closed__1; -x_3 = lean_unsigned_to_nat(754u); +x_3 = lean_unsigned_to_nat(755u); x_4 = lean_unsigned_to_nat(72u); x_5 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_mkCompositeField___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -23925,7 +23925,7 @@ static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_m 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___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_mkCompositeField___closed__1; x_2 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___closed__1; -x_3 = lean_unsigned_to_nat(748u); +x_3 = lean_unsigned_to_nat(749u); x_4 = lean_unsigned_to_nat(68u); x_5 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_mkCompositeField___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -27507,7 +27507,7 @@ x_26 = lean_box(0); lean_inc(x_6); lean_inc(x_5); lean_inc(x_10); -x_27 = l_Lean_Elab_resolveGlobalConstNoOverloadWithInfo___at_Lean_Elab_elabDeriving___spec__2(x_10, x_26, x_5, x_6, x_7); +x_27 = l_Lean_Elab_resolveGlobalConstNoOverloadWithInfo___at_Lean_Elab_elabDeriving___spec__6(x_10, x_26, x_5, x_6, x_7); if (lean_obj_tag(x_27) == 0) { lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; size_t x_32; size_t x_33; @@ -27574,7 +27574,7 @@ x_43 = lean_box(0); lean_inc(x_6); lean_inc(x_5); lean_inc(x_10); -x_44 = l_Lean_Elab_resolveGlobalConstNoOverloadWithInfo___at_Lean_Elab_elabDeriving___spec__2(x_10, x_43, x_5, x_6, x_7); +x_44 = l_Lean_Elab_resolveGlobalConstNoOverloadWithInfo___at_Lean_Elab_elabDeriving___spec__6(x_10, x_43, x_5, x_6, x_7); if (lean_obj_tag(x_44) == 0) { lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; size_t x_49; size_t x_50; @@ -27690,7 +27690,7 @@ x_69 = lean_box(0); lean_inc(x_6); lean_inc(x_5); lean_inc(x_10); -x_70 = l_Lean_Elab_resolveGlobalConstNoOverloadWithInfo___at_Lean_Elab_elabDeriving___spec__2(x_10, x_69, x_5, x_6, x_7); +x_70 = l_Lean_Elab_resolveGlobalConstNoOverloadWithInfo___at_Lean_Elab_elabDeriving___spec__6(x_10, x_69, x_5, x_6, x_7); if (lean_obj_tag(x_70) == 0) { lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; size_t x_76; size_t x_77; @@ -27804,39 +27804,38 @@ lean_dec(x_1); x_9 = l_Lean_Syntax_isNone(x_8); if (x_9 == 0) { -lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l_Lean_nullKind; -x_11 = lean_unsigned_to_nat(2u); +lean_object* x_10; uint8_t x_11; +x_10 = lean_unsigned_to_nat(2u); lean_inc(x_8); -x_12 = l_Lean_Syntax_isNodeOf(x_8, x_10, x_11); -if (x_12 == 0) +x_11 = l_Lean_Syntax_matchesNull(x_8, x_10); +if (x_11 == 0) { -lean_object* x_13; +lean_object* x_12; lean_dec(x_8); lean_dec(x_6); -x_13 = lean_box(0); -return x_13; +x_12 = lean_box(0); +return x_12; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_14 = l_Lean_Syntax_getArg(x_8, x_7); +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = l_Lean_Syntax_getArg(x_8, x_7); lean_dec(x_8); -x_15 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_15, 0, x_14); -x_16 = lean_box(0); -x_17 = l_Lean_Elab_getOptDerivingClasses___at_Lean_Elab_Command_elabStructure___spec__2___lambda__1(x_6, x_16, x_15); -return x_17; +x_14 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, x_13); +x_15 = lean_box(0); +x_16 = l_Lean_Elab_getOptDerivingClasses___at_Lean_Elab_Command_elabStructure___spec__2___lambda__1(x_6, x_15, x_14); +return x_16; } } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_dec(x_8); +x_17 = lean_box(0); x_18 = lean_box(0); -x_19 = lean_box(0); -x_20 = l_Lean_Elab_getOptDerivingClasses___at_Lean_Elab_Command_elabStructure___spec__2___lambda__1(x_6, x_19, x_18); -return x_20; +x_19 = l_Lean_Elab_getOptDerivingClasses___at_Lean_Elab_Command_elabStructure___spec__2___lambda__1(x_6, x_18, x_17); +return x_19; } } } @@ -27901,166 +27900,165 @@ return x_8; } else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; x_9 = lean_unsigned_to_nat(0u); x_10 = l_Lean_Syntax_getArg(x_1, x_9); lean_dec(x_1); -x_11 = l_Lean_nullKind; -x_12 = lean_unsigned_to_nat(2u); +x_11 = lean_unsigned_to_nat(2u); lean_inc(x_10); -x_13 = l_Lean_Syntax_isNodeOf(x_10, x_11, x_12); -if (x_13 == 0) +x_12 = l_Lean_Syntax_matchesNull(x_10, x_11); +if (x_12 == 0) { -lean_object* x_14; lean_object* x_15; +lean_object* x_13; lean_object* x_14; lean_dec(x_10); lean_dec(x_3); lean_dec(x_2); -x_14 = l_Lean_Elab_elabAttrs___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__3___closed__1; -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_15, 1, x_4); -return x_15; +x_13 = l_Lean_Elab_elabAttrs___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__3___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_4); +return x_14; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; -x_16 = lean_unsigned_to_nat(1u); -x_17 = l_Lean_Syntax_getArg(x_10, x_16); +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_object* x_20; +x_15 = lean_unsigned_to_nat(1u); +x_16 = l_Lean_Syntax_getArg(x_10, x_15); lean_dec(x_10); -x_18 = l_Lean_Syntax_getArgs(x_17); -lean_dec(x_17); -x_19 = lean_array_get_size(x_18); -x_20 = lean_nat_dec_lt(x_9, x_19); -if (x_20 == 0) +x_17 = l_Lean_Syntax_getArgs(x_16); +lean_dec(x_16); +x_18 = lean_array_get_size(x_17); +x_19 = lean_nat_dec_lt(x_9, x_18); +if (x_19 == 0) { -lean_object* x_51; -lean_dec(x_19); +lean_object* x_50; lean_dec(x_18); -x_51 = l_Lean_Elab_elabAttrs___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__3___closed__1; -x_21 = x_51; -goto block_50; +lean_dec(x_17); +x_50 = l_Lean_Elab_elabAttrs___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__3___closed__1; +x_20 = x_50; +goto block_49; } else { -uint8_t x_52; -x_52 = lean_nat_dec_le(x_19, x_19); -if (x_52 == 0) +uint8_t x_51; +x_51 = lean_nat_dec_le(x_18, x_18); +if (x_51 == 0) { -lean_object* x_53; -lean_dec(x_19); +lean_object* x_52; lean_dec(x_18); -x_53 = l_Lean_Elab_elabAttrs___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__3___closed__1; -x_21 = x_53; -goto block_50; +lean_dec(x_17); +x_52 = l_Lean_Elab_elabAttrs___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__3___closed__1; +x_20 = x_52; +goto block_49; } else { -size_t x_54; size_t x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_54 = 0; -x_55 = lean_usize_of_nat(x_19); -lean_dec(x_19); -x_56 = l_Lean_Elab_getOptDerivingClasses___at_Lean_Elab_Command_elabStructure___spec__2___closed__3; -x_57 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(x_18, x_54, x_55, x_56); +size_t x_53; size_t x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_53 = 0; +x_54 = lean_usize_of_nat(x_18); lean_dec(x_18); -x_58 = lean_ctor_get(x_57, 1); -lean_inc(x_58); -lean_dec(x_57); -x_21 = x_58; -goto block_50; +x_55 = l_Lean_Elab_getOptDerivingClasses___at_Lean_Elab_Command_elabStructure___spec__2___closed__3; +x_56 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(x_17, x_53, x_54, x_55); +lean_dec(x_17); +x_57 = lean_ctor_get(x_56, 1); +lean_inc(x_57); +lean_dec(x_56); +x_20 = x_57; +goto block_49; } } -block_50: +block_49: { -lean_object* x_22; lean_object* x_23; -x_22 = l_Lean_Elab_getOptDerivingClasses___at_Lean_Elab_Command_elabStructure___spec__2___closed__4; -x_23 = l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___spec__1(x_21, x_22); -lean_dec(x_21); -if (lean_obj_tag(x_23) == 0) +lean_object* x_21; lean_object* x_22; +x_21 = l_Lean_Elab_getOptDerivingClasses___at_Lean_Elab_Command_elabStructure___spec__2___closed__4; +x_22 = l_Array_sequenceMap___at_Lean_Elab_elabDeriving___spec__2(x_20, x_21); +lean_dec(x_20); +if (lean_obj_tag(x_22) == 0) { -lean_object* x_24; lean_object* x_25; +lean_object* x_23; lean_object* x_24; lean_dec(x_3); lean_dec(x_2); -x_24 = l_Lean_Elab_elabAttrs___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__3___closed__1; -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_4); -return x_25; +x_23 = l_Lean_Elab_elabAttrs___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__3___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_4); +return x_24; } else { -lean_object* x_26; lean_object* x_27; size_t x_28; size_t x_29; 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; size_t x_37; lean_object* x_38; -x_26 = lean_ctor_get(x_23, 0); -lean_inc(x_26); -lean_dec(x_23); -x_27 = lean_array_get_size(x_26); -x_28 = lean_usize_of_nat(x_27); -lean_dec(x_27); -x_29 = 0; -lean_inc(x_26); -x_30 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___spec__3(x_28, x_29, x_26); -x_31 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__tactic_xb7_x2e_____x3b____1___spec__4(x_28, x_29, x_26); -x_32 = lean_array_get_size(x_30); -x_33 = l_Array_toSubarray___rarg(x_30, x_9, x_32); -x_34 = l_Lean_Elab_elabAttrs___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__3___closed__1; -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -x_36 = lean_array_get_size(x_31); -x_37 = lean_usize_of_nat(x_36); -lean_dec(x_36); -x_38 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabStructure___spec__3(x_31, x_37, x_29, x_35, x_2, x_3, x_4); -lean_dec(x_31); -if (lean_obj_tag(x_38) == 0) +lean_object* x_25; lean_object* x_26; size_t x_27; size_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; size_t x_36; lean_object* x_37; +x_25 = lean_ctor_get(x_22, 0); +lean_inc(x_25); +lean_dec(x_22); +x_26 = lean_array_get_size(x_25); +x_27 = lean_usize_of_nat(x_26); +lean_dec(x_26); +x_28 = 0; +lean_inc(x_25); +x_29 = l_Array_mapMUnsafe_map___at_Lean_Elab_elabDeriving___spec__4(x_27, x_28, x_25); +x_30 = l_Array_mapMUnsafe_map___at_Lean_Elab_elabDeriving___spec__5(x_27, x_28, x_25); +x_31 = lean_array_get_size(x_29); +x_32 = l_Array_toSubarray___rarg(x_29, x_9, x_31); +x_33 = l_Lean_Elab_elabAttrs___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__3___closed__1; +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +x_35 = lean_array_get_size(x_30); +x_36 = lean_usize_of_nat(x_35); +lean_dec(x_35); +x_37 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabStructure___spec__3(x_30, x_36, x_28, x_34, x_2, x_3, x_4); +lean_dec(x_30); +if (lean_obj_tag(x_37) == 0) { -uint8_t x_39; -x_39 = !lean_is_exclusive(x_38); -if (x_39 == 0) +uint8_t x_38; +x_38 = !lean_is_exclusive(x_37); +if (x_38 == 0) { -lean_object* x_40; lean_object* x_41; -x_40 = lean_ctor_get(x_38, 0); -x_41 = lean_ctor_get(x_40, 1); -lean_inc(x_41); -lean_dec(x_40); -lean_ctor_set(x_38, 0, x_41); -return x_38; +lean_object* x_39; lean_object* x_40; +x_39 = lean_ctor_get(x_37, 0); +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +lean_dec(x_39); +lean_ctor_set(x_37, 0, x_40); +return x_37; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_42 = lean_ctor_get(x_38, 0); -x_43 = lean_ctor_get(x_38, 1); -lean_inc(x_43); +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_41 = lean_ctor_get(x_37, 0); +x_42 = lean_ctor_get(x_37, 1); lean_inc(x_42); -lean_dec(x_38); -x_44 = lean_ctor_get(x_42, 1); -lean_inc(x_44); -lean_dec(x_42); -x_45 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_45, 0, x_44); -lean_ctor_set(x_45, 1, x_43); -return x_45; +lean_inc(x_41); +lean_dec(x_37); +x_43 = lean_ctor_get(x_41, 1); +lean_inc(x_43); +lean_dec(x_41); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_42); +return x_44; } } else { -uint8_t x_46; -x_46 = !lean_is_exclusive(x_38); -if (x_46 == 0) +uint8_t x_45; +x_45 = !lean_is_exclusive(x_37); +if (x_45 == 0) { -return x_38; +return x_37; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_38, 0); -x_48 = lean_ctor_get(x_38, 1); -lean_inc(x_48); +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_37, 0); +x_47 = lean_ctor_get(x_37, 1); lean_inc(x_47); -lean_dec(x_38); -x_49 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_49, 0, x_47); -lean_ctor_set(x_49, 1, x_48); -return x_49; +lean_inc(x_46); +lean_dec(x_37); +x_48 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_48, 0, x_46); +lean_ctor_set(x_48, 1, x_47); +return x_48; } } } @@ -29858,7 +29856,7 @@ lean_dec(x_8); return x_15; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_11179_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_11219_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; @@ -30255,15 +30253,15 @@ l___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___closed__3 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___closed__4 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___closed__4(); l___private_Lean_Elab_Structure_0__Lean_Elab_Command_findFieldInfo_x3f___closed__1 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_findFieldInfo_x3f___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_findFieldInfo_x3f___closed__1); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2488____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2488____closed__1(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2488____closed__1); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2488____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2488____closed__2(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2488____closed__2); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2488____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2488____closed__3(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2488____closed__3); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2488____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2488____closed__4(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2488____closed__4); -if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2488_(lean_io_mk_world()); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2512____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2512____closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2512____closed__1); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2512____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2512____closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2512____closed__2); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2512____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2512____closed__3(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2512____closed__3); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2512____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2512____closed__4(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2512____closed__4); +if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_2512_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Elab_Command_structureDiamondWarning = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Elab_Command_structureDiamondWarning); @@ -30566,7 +30564,7 @@ l_Lean_Elab_Command_elabStructure___closed__12 = _init_l_Lean_Elab_Command_elabS lean_mark_persistent(l_Lean_Elab_Command_elabStructure___closed__12); l_Lean_Elab_Command_elabStructure___closed__13 = _init_l_Lean_Elab_Command_elabStructure___closed__13(); lean_mark_persistent(l_Lean_Elab_Command_elabStructure___closed__13); -res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_11179_(lean_io_mk_world()); +res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_11219_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/Syntax.c b/stage0/stdlib/Lean/Elab/Syntax.c index c2698aaa939..80f399e238f 100644 --- a/stage0/stdlib/Lean/Elab/Syntax.c +++ b/stage0/stdlib/Lean/Elab/Syntax.c @@ -22,19 +22,17 @@ static lean_object* l_Lean_Elab_Term_toParserDescr_process___closed__9; lean_object* l_Lean_mkCIdentFrom(lean_object*, lean_object*); static lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___closed__3; static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__25; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processUnary(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_Term_toParserDescr_processNonReserved___spec__1___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_toParserDescr_process___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__9; static lean_object* l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__2; -LEAN_EXPORT lean_object* l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_inferMacroRulesAltKind___closed__1; size_t lean_usize_add(size_t, size_t); static lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat___closed__9; -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabSyntax___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_forM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_process___closed__4; static lean_object* l_Lean_Elab_Command_elabSyntax___lambda__5___closed__7; +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_toParserDescr_process___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_resolveSyntaxKind___closed__1; LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_toParserDescr_processNonReserved___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*); @@ -45,6 +43,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkLeftRec___lambda__2(lean_object*, static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__50; lean_object* l_Lean_stringToMessageData(lean_object*); static lean_object* l_Lean_Elab_Command_resolveSyntaxKind___closed__2; +static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__2___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processSeq___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__34; static lean_object* l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__7; @@ -56,11 +55,11 @@ static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_ LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandNoKindMacroRulesAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__1___boxed(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_Term_toParserDescr_process___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_nullKind; -static lean_object* l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__20; -lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabCommand___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSyntax___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Array_sequenceMap___at_Lean_Elab_Command_elabSyntax___spec__2___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSyntax___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_checkLeftRec___spec__10(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_Array_mapMUnsafe_map___at_Lean_Elab_Term_toParserDescr_processAlias___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__46; extern lean_object* l_Lean_Parser_leadPrec; lean_object* l_Lean_Elab_Command_liftTermElabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -73,19 +72,20 @@ lean_object* lean_name_mk_string(lean_object*, lean_object*); lean_object* l_Array_eraseIdx___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabSyntax___spec__6(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); static lean_object* l_Lean_Elab_Term_toParserDescr_process___closed__1; lean_object* lean_io_error_to_string(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat___closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_checkLeftRec___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___closed__3; static lean_object* l_Lean_Elab_Command_expandNoKindMacroRulesAux___lambda__1___closed__3; extern uint32_t l_Lean_idBeginEscape; static lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat_declRange___closed__6; -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_append___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_toParserDescr_processAlias___spec__4___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_mapTRAux___at_Lean_resolveGlobalConstCore___spec__2(lean_object*, lean_object*); -lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13___rarg(lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__1; static lean_object* l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__3; static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__13; @@ -103,7 +103,9 @@ static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_ lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_toParserDescr_process___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__13; +static lean_object* l_panic___at_Lean_Elab_Term_toParserDescr_processAlias___spec__2___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat___closed__2; +static lean_object* l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__6; uint8_t l_Lean_Parser_leadingIdentBehavior(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__60; LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandNoKindMacroRulesAux___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -111,10 +113,11 @@ static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyn static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__74; static lean_object* l_List_filterMap___at_Lean_Elab_Term_resolveParserName___spec__1___closed__2; static lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__2; +lean_object* l_Lean_Parser_getParserAliasInfo(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_ensureNoPrec(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_filterMap___at_Lean_Elab_Term_resolveParserName___spec__1___closed__1; +static lean_object* l_Lean_Elab_Term_toParserDescr_processAlias___closed__12; static lean_object* l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__6; -static lean_object* l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__16; lean_object* lean_environment_find(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__4; static lean_object* l_Lean_Elab_Command_checkRuleKind___closed__2; @@ -127,7 +130,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat___clos LEAN_EXPORT lean_object* l_Lean_Elab_resolveGlobalConstWithInfos___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__2(lean_object*, 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_resolveGlobalName___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__18; -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabSyntax___spec__9(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__76; static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__14; @@ -144,19 +146,25 @@ static lean_object* l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__ LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processNonReserved(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_expandNoKindMacroRulesAux___lambda__1___closed__6; static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__22; -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__6___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___closed__1; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_toParserDescr_processAlias___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Level_PP_Result_quote___spec__1(lean_object*, lean_object*); lean_object* l_Lean_mkIdentFrom(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processAlias(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_process___closed__15; static lean_object* l_Lean_Elab_Term_toParserDescr_processAtom___closed__2; static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__52; +extern lean_object* l_instInhabitedNat; +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Elab_Command_elabSyntax___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__3___closed__4; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_toParserDescr_processAlias___spec__5___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__29; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSyntax___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__8; lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_toParserDescr_processAlias___closed__5; lean_object* lean_array_get_size(lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -168,14 +176,14 @@ LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Comma LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_toParserDescr_processNonReserved___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSyntax___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33; -static lean_object* l_Lean_Elab_Command_elabSyntax___lambda__5___closed__14; uint8_t l_Char_isWhitespace(uint32_t); LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__6___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processSeq___lambda__1(lean_object*, 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_Term_toParserDescr_processAlias___lambda__1(lean_object*, 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_Term_toParserDescr_processSepBy1(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_liftMacroM___at_Lean_Elab_Command_elabSyntax___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_setEnv___at_Lean_Elab_Command_expandDeclId___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__10___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_withNotFirst(lean_object*); static lean_object* l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3___closed__2; static lean_object* l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__1; @@ -191,12 +199,13 @@ static lean_object* l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_ static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__6; lean_object* l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_mkNameFromParserSyntax_visit___closed__1; +static lean_object* l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__11; LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkNameFromParserSyntax(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_checkRuleKind___closed__1; uint8_t lean_usize_dec_lt(size_t, size_t); lean_object* l_liftExcept___at_Lean_Elab_liftMacroM___spec__1(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___closed__2; LEAN_EXPORT uint8_t l_Lean_Elab_Command_checkRuleKind(lean_object*, lean_object*); -extern lean_object* l_Lean_nameLitKind; LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkLeftRec___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processAtom___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_checkLeftRec___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -219,25 +228,32 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkNameFromParserSyntax_visit(lean_o LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabDeclareSyntaxCat___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat___closed__4; static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__22; +static lean_object* l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__4; static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__4; static lean_object* l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__4; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_toParserDescr_processAlias___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__10___boxed(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_Array_mapMUnsafe_map___at_Lean_Elab_Term_toParserDescr_processSeq___spec__2(size_t, size_t, 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_Command_resolveSyntaxKind(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__27; +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabSyntax___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabSyntax___lambda__5___closed__6; +static lean_object* l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__17; static lean_object* l_Lean_Elab_Term_checkLeftRec___closed__1; +static lean_object* l_Lean_Elab_Term_toParserDescr_processAlias___closed__8; extern lean_object* l_Lean_LocalContext_empty; -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_string_utf8_next(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__6; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_mkNameFromParserSyntax_visit___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__65; LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandOptPrecedence___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__39; +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabSyntax___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___closed__4; static lean_object* l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__9; static lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat___closed__8; lean_object* l_Lean_Elab_resolveGlobalConstWithInfos___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_8045____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveParserName___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveParserName___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__3___closed__5; @@ -256,24 +272,28 @@ static lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_p lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runTermElabM___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__3___closed__6; +static lean_object* l_Lean_Elab_Term_toParserDescr_processAlias___closed__9; LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveParserName___rarg___lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__11; LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__23; +static lean_object* l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveParserName___rarg___lambda__3___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Term_checkLeftRec___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__36; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_toParserDescr_process___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandNoKindMacroRulesAux___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_inferMacroRulesAltKind___spec__1(lean_object*, lean_object*); static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__13; static lean_object* l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__17; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_checkLeftRec___spec__1___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__8; +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Elab_Command_elabSyntax___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__19; lean_object* l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(lean_object*, lean_object*); +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__11; static lean_object* l_Lean_addTrace___at_Lean_Elab_Term_checkLeftRec___spec__3___closed__2; static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__51; @@ -282,7 +302,8 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabSyntaxAbbrev(lean_ lean_object* lean_st_ref_take(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_checkLeftRec___spec__8___boxed(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_elabSyntax___spec__1___boxed(lean_object*, lean_object*); -extern lean_object* l_Lean_numLitKind; +LEAN_EXPORT lean_object* l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(lean_object*, lean_object*); +lean_object* l_ReaderT_bind___at_Lean_Unhygienic_instMonadQuotationUnhygienic___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabSyntax_declRange___closed__3; static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__82; lean_object* lean_nat_sub(lean_object*, lean_object*); @@ -291,13 +312,15 @@ static lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat_declRa static lean_object* l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__8; LEAN_EXPORT lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_markAsTrailingParser___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_process___closed__14; +LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Elab_Command_elabSyntax___spec__2(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__9; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabSyntax(lean_object*); lean_object* l_Lean_Syntax_isStrLit_x3f(lean_object*); -static lean_object* l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__19; static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__35; -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(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_throwErrorAt___at_Lean_Elab_Term_checkLeftRec___spec__5(lean_object*, 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_Array_mapIdxM_map___at_Lean_Elab_Term_toParserDescr_processSeq___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_throwError___at_Lean_Elab_Term_toParserDescr_process___spec__7___boxed(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_Array_mapMUnsafe_map___at_Lean_Elab_Term_toParserDescr_processAlias___spec__5(size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Command_elabSyntax___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSyntax___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabSyntax___lambda__5___closed__9; @@ -306,7 +329,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_ static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__38; static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__21; static lean_object* l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__5; -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabSyntax___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__12; lean_object* l_Lean_Name_toString(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_resolveSyntaxKind___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -314,8 +336,10 @@ lean_object* l_Lean_replaceRef(lean_object*, lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ResolveName_resolveNamespace(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__2; +lean_object* l_Lean_Unhygienic_run___rarg(lean_object*); lean_object* l_String_capitalize(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_6964_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_ensureUnaryOutput(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_8045_(lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_process___closed__13; static lean_object* l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__7; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_expandNoKindMacroRulesAux___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -323,42 +347,49 @@ LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_checkLeftRec static lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__3; static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__44; LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveParserName___rarg___lambda__3(lean_object*, lean_object*); +lean_object* l_instInhabited___rarg(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat_declRange___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_strLitToPattern___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_logTrace___at_Lean_Elab_Command_elabCommand___spec__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__8; -LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandNoKindMacroRulesAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addTrace___at_Lean_Elab_Term_checkLeftRec___spec__3___closed__5; +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabSyntax___spec__11(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_ensureNoPrec___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabSyntaxAbbrev_declRange(lean_object*); lean_object* l_Nat_repr(lean_object*); static lean_object* l_Lean_addTrace___at_Lean_Elab_Term_checkLeftRec___spec__3___closed__8; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processAlias___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__45; LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkLeftRec___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__64; LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_ensureNoPrec___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addTrace___at_Lean_Elab_Term_checkLeftRec___spec__3___closed__1; +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabSyntax___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_withNestedParser(lean_object*); -static lean_object* l_Lean_Elab_Command_elabSyntax___lambda__5___closed__13; +LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_Term_toParserDescr_processAlias___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_panic___at_Lean_Elab_Term_toParserDescr_processAlias___spec__2___closed__1; +static lean_object* l_Lean_Elab_Term_toParserDescr_processAlias___closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_withNestedParser(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getId(lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__1; lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Init_Meta_0__Lean_quoteNameMk(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_resolveSyntaxKind___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_inferMacroRulesAltKind___spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__17; static lean_object* l___regBuiltin_Lean_Elab_Command_elabSyntax_declRange___closed__6; -extern lean_object* l_Lean_choiceKind; lean_object* l_Lean_Elab_Term_addAutoBoundImplicits_x27___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Parser_maxPrec; static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__57; +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__12(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabSyntax___lambda__5___closed__2; static lean_object* l_Lean_Elab_Term_checkLeftRec___lambda__2___closed__1; +static lean_object* l_Lean_Elab_Term_toParserDescr_processAlias___closed__10; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_checkLeftRec___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__83; 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_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__12___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_mkNameFromParserSyntax_visit___spec__1(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__7; static lean_object* l_Lean_Elab_Term_checkLeftRec___closed__2; @@ -367,6 +398,7 @@ uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__5; static lean_object* l_Lean_Elab_Term_toParserDescr_process___closed__6; static lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__3; +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__7; uint32_t lean_string_utf8_get(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__10; @@ -374,6 +406,7 @@ static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_ lean_object* l_Lean_Syntax_mkStrLit(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__2(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_throwError___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__12; static lean_object* l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__1; LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addTrace___at_Lean_Elab_Term_checkLeftRec___spec__3___closed__3; @@ -381,23 +414,28 @@ static lean_object* l_Lean_addTrace___at_Lean_Elab_Term_checkLeftRec___spec__3__ static lean_object* l___regBuiltin_Lean_Elab_Command_elabSyntaxAbbrev_declRange___closed__7; static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__12; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat_declRange(lean_object*); +static lean_object* l_Lean_Elab_Term_toParserDescr_process___closed__16; static lean_object* l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__3___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__68; static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__24; +lean_object* l_Array_unzip___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processSeq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Array_isEmpty___rarg(lean_object*); LEAN_EXPORT lean_object* l_List_filterMap___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__15(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_checkLeftRec___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__5; -extern lean_object* l_Lean_instInhabitedSyntax; LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabSyntax___closed__3; -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__6___boxed(lean_object*, lean_object*); lean_object* l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_resolveGlobalConstWithInfos___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_elabCommand(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_ensureUnaryOutput___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabSyntax___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_toParserDescr_processAlias___closed__6; LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_toParserDescr_processAlias___closed__4; lean_object* l_Lean_evalOptPrio(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__12; lean_object* l_Lean_ResolveName_resolveGlobalName(lean_object*, lean_object*, lean_object*, lean_object*); @@ -405,33 +443,33 @@ static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_ LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processAtom(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__2; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_checkLeftRec___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_toParserDescr_processAlias___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Command_elabSyntax_declRange___closed__4; static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__20; LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__9(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_resolveGlobalConstCore___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__6(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_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__10; size_t lean_usize_of_nat(lean_object*); static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__3; -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabSyntax___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSyntax___lambda__2(lean_object*, 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_pushInfoTree___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_type(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__2___rarg(lean_object*); static lean_object* l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__3___closed__7; LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_toParserDescr_processNonReserved___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processParserCategory___closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_expandNoKindMacroRulesAux___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabSyntax___spec__5(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Command_elabSyntaxAbbrev___closed__3; static lean_object* l_Lean_Elab_Term_toParserDescr_process___closed__10; static lean_object* l_Lean_Elab_Command_elabSyntax___lambda__5___closed__8; static lean_object* l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__1; -LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_checkLeftRec___spec__8___rarg___closed__2; +LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processNullaryOrCat(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceOptions(lean_object*); @@ -445,22 +483,20 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkLeftRec___lambda__1___boxed(lean_ lean_object* l_Lean_Macro_expandMacro_x3f(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_ensureUnaryParserAlias(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__43; -static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_checkLeftRec___spec__8___rarg___closed__1; static lean_object* l_Lean_Elab_Command_expandNoKindMacroRulesAux___lambda__1___closed__4; lean_object* l_Lean_mkAtomFrom(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSyntax___lambda__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabSyntaxAbbrev___closed__1; static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__15; -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabSyntax___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_mkNumLit(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__14(lean_object*, lean_object*, lean_object*, 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_Command_inferMacroRulesAltKind(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabSyntax___lambda__10___closed__1; -static lean_object* l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__8; lean_object* l_Lean_Parser_registerParserCategory(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); -lean_object* l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_toParserDescr_processAlias___spec__4(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__61; lean_object* l_Lean_Syntax_getQuotContent(lean_object*); +static lean_object* l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__15; LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkNameFromParserSyntax_appendCatName(lean_object*, lean_object*); uint8_t lean_uint32_dec_eq(uint32_t, uint32_t); LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_toParserDescr_processNonReserved___spec__2___rarg(lean_object*, lean_object*, lean_object*); @@ -478,32 +514,31 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_pro static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__69; static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__11; uint8_t l_String_isEmpty(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_toParserDescr_processAlias___spec__3(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__14; -static lean_object* l_Lean_Elab_Term_toParserDescr_processUnary___closed__3; +static lean_object* l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__3; lean_object* l_Lean_Syntax_getNumArgs(lean_object*); +extern lean_object* l_Lean_Elab_Term_instMonadTermElabM; static lean_object* l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__1; static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__59; static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31; static lean_object* l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__6; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_toParserDescr_process___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_toParserDescr_processUnary___closed__2; lean_object* l_Lean_Elab_Command_withMacroExpansion___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_environment_main_module(lean_object*); -static lean_object* l_Lean_Elab_Term_toParserDescr_processUnary___closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_checkLeftRec___spec__10___boxed(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__1___rarg(lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__5; LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Term_checkLeftRec___spec__3___boxed(lean_object*, 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___regBuiltin_Lean_Elab_Command_elabSyntax_declRange(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__6(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_checkLeftRec___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabSyntax_declRange___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_isAtomLikeSyntax___boxed(lean_object*); lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_markAsTrailingParser(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_Term_toParserDescr_processUnary___boxed(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_elabSyntax___spec__8(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveParserName___rarg___lambda__4(lean_object*, lean_object*); uint8_t l_Lean_Parser_isParserCategory(lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); @@ -514,8 +549,8 @@ static lean_object* l_Lean_Elab_Term_checkLeftRec___lambda__2___closed__4; lean_object* l_Lean_Parser_ensureConstantParserAlias(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_checkLeftRec___spec__1___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__15; -static lean_object* l_Lean_Elab_Term_toParserDescr_processUnary___closed__5; static lean_object* l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__3; lean_object* l_Lean_Syntax_getArgs(lean_object*); lean_object* l_Lean_Name_append(lean_object*, lean_object*); @@ -523,25 +558,25 @@ static lean_object* l_Lean_Elab_Term_checkLeftRec___closed__3; static lean_object* l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__3; lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandDeclId___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__26; +lean_object* l_instInhabitedReaderT___rarg___boxed(lean_object*, lean_object*); lean_object* l_Lean_Syntax_getKind(lean_object*); uint8_t l_Lean_Parser_isValidSyntaxNodeKind(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabSyntax___lambda__5___closed__11; 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_8845_(uint8_t, uint8_t); static lean_object* l_Lean_Elab_Command_elabSyntax___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_withNestedParser___rarg(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_Term_checkLeftRec(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_toParserDescr_processUnary___closed__7; +static lean_object* l_Lean_Elab_Term_toParserDescr_processAlias___closed__11; LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___boxed(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*); extern lean_object* l_Lean_Elab_Command_instInhabitedScope; static lean_object* l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__9; -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__58; static lean_object* l_Lean_Elab_Command_elabSyntaxAbbrev___closed__2; static lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__1; static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__75; -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__10(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_ensureUnaryOutput___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_process(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_checkSyntaxNodeKind___at_Lean_Elab_Command_resolveSyntaxKind___spec__2___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat___closed__7; @@ -552,15 +587,16 @@ lean_object* l_ReaderT_pure___at_Lean_Elab_liftMacroM___spec__2___rarg___boxed(l LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__11___boxed(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_throwErrorAt___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabSyntax___closed__2; -lean_object* l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__26; LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processParserCategory(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat_declRange___closed__5; +static lean_object* l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__14; static lean_object* l_Lean_addTrace___at_Lean_Elab_Term_checkLeftRec___spec__3___closed__4; lean_object* l_Lean_Elab_Command_getMainModule___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_process___closed__7; lean_object* l_Lean_Parser_isParserAlias(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_checkLeftRec___spec__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Command_elabSyntax_declRange___closed__1; static lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -570,15 +606,17 @@ LEAN_EXPORT lean_object* l_Lean_Elab_checkSyntaxNodeKind___at_Lean_Elab_Command_ LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkNameFromParserSyntax_appendCatName___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabSyntax___lambda__5___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_toParserDescr_process___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__17; +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__2; static lean_object* l_Lean_Elab_Term_toParserDescr_process___closed__12; lean_object* l_Lean_Elab_Command_getRef(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_mkNameLit(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoTree___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__42; lean_object* l_Lean_Name_getPrefix(lean_object*); static lean_object* l_Lean_Elab_Command_elabSyntax___lambda__11___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_checkLeftRec___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__16; static lean_object* l_Lean_Elab_Command_elabSyntax___lambda__8___closed__3; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addTrace___at_Lean_Elab_Term_checkLeftRec___spec__3___closed__6; @@ -589,15 +627,21 @@ static lean_object* l___regBuiltin_Lean_Elab_Command_elabSyntax___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Command_elabSyntaxAbbrev_declRange___closed__6; lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_process___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabSyntax___spec__9(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_checkLeftRec___spec__8___rarg(lean_object*); +static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__2___rarg___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat___closed__6; LEAN_EXPORT lean_object* l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3(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_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isNone(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_resolveGlobalConstWithInfos___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__53; +lean_object* lean_panic_fn(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__8___rarg(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabSyntaxAbbrev_declRange___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSyntaxAbbrev(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_String_trim(lean_object*); +static lean_object* l_Lean_Elab_Term_toParserDescr_processAlias___closed__7; static lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat_declRange___closed__3; static lean_object* l_Lean_Elab_Term_toParserDescr_processParserCategory___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandOptPrecedence(lean_object*, lean_object*, lean_object*); @@ -606,10 +650,8 @@ uint8_t l_Lean_isIdFirst(uint32_t); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__28; static lean_object* l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__16; -static lean_object* l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__18; static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__14; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processBinary(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__15; +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_process___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_withoutAutoBoundImplicit___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabSyntax___lambda__9___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveParserName(lean_object*); @@ -620,12 +662,12 @@ static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyn static lean_object* l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__5; static lean_object* l_Lean_Elab_Command_elabSyntax___lambda__4___closed__1; extern lean_object* l_Lean_Elab_unsupportedSyntaxExceptionId; -static lean_object* l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__14; +static lean_object* l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__2; static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__3; -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabSyntax___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_ReaderT_instMonadReaderT___rarg(lean_object*); lean_object* l_Lean_Parser_ensureBinaryParserAlias(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__6; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_6964____closed__1; +static lean_object* l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__7; static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__12; lean_object* lean_mk_syntax_ident(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat_declRange___closed__7; @@ -633,11 +675,11 @@ static lean_object* l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___clo LEAN_EXPORT uint8_t l_Lean_Elab_Term_toParserDescr_isValidAtom(lean_object*); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__3; -static lean_object* l_Lean_Elab_Term_toParserDescr_processUnary___closed__4; static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__37; static lean_object* l_Lean_Elab_Term_toParserDescr_process___closed__11; +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__8___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq(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_elabSyntax___spec__10___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__12___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__77; static lean_object* l_Lean_Elab_Command_expandNoKindMacroRulesAux___lambda__1___closed__1; static lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__2; @@ -645,6 +687,7 @@ lean_object* l_Lean_Core_resetMessageLog(lean_object*, lean_object*, lean_object LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveParserName___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_checkLeftRec___spec__7___closed__2; uint8_t l_List_isEmpty___rarg(lean_object*); +static lean_object* l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSyntax___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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___boxed(lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_getScope___rarg(lean_object*, lean_object*); @@ -652,13 +695,13 @@ LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_checkLeftRec LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__20; LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandNoKindMacroRulesAux___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_toParserDescr_processAlias___closed__2; static lean_object* l_Lean_Elab_Command_inferMacroRulesAltKind___closed__2; static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__79; LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12(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_List_forM___at_Lean_Elab_Term_checkLeftRec___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__32; static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__10; -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_toParserDescr_process___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Command_elabSyntaxAbbrev_declRange___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSyntax___lambda__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -667,12 +710,15 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_withNotF static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__56; static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__66; LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__85; static lean_object* l_Lean_Elab_Command_elabSyntax___lambda__5___closed__10; +static lean_object* l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Command_inferMacroRulesAltKind___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___boxed(lean_object*, 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_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___boxed(lean_object*, lean_object*, 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_mkConstWithLevelParams___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabSyntaxAbbrev_declRange___closed__4; static lean_object* l_Lean_Elab_Term_toParserDescr_ensureNoPrec___closed__2; +static lean_object* l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_checkLeftRec___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_withAutoBoundImplicit___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -680,7 +726,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_inferMacroRulesAltKind___lambda__1_ static lean_object* l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__2; uint8_t lean_string_utf8_at_end(lean_object*, lean_object*); lean_object* l_Lean_mkConst(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processBinary___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__24; static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__17; static lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat_declRange___closed__2; @@ -696,37 +741,35 @@ static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyn static lean_object* l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__12; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOptionM___at_Lean_Elab_Term_checkLeftRec___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__3___closed__2; -static lean_object* l_Lean_Elab_Term_toParserDescr_processUnary___closed__6; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__5___boxed(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_Command_elabSyntaxAbbrev___lambda__3(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_throwErrorAt___at_Lean_Elab_Command_expandNoKindMacroRulesAux___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabDeclareSyntaxCat(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__3___boxed(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___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabSyntaxAbbrev_declRange___closed__1; static lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__11; static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__16; +static lean_object* l_panic___at_Lean_Elab_Term_toParserDescr_processAlias___spec__2___closed__3; LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__71; static lean_object* l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__2; static lean_object* l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__1; -static lean_object* l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__9; static lean_object* l___regBuiltin_Lean_Elab_Command_elabSyntax___closed__1; LEAN_EXPORT lean_object* l_Lean_resolveGlobalName___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_expandNoKindMacroRulesAux___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSyntax___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_mkUnusedBaseName(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSyntax___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Command_elabSyntax___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_string_dec_eq(lean_object*, lean_object*); -lean_object* l_Lean_Syntax_mkLit(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSyntax___lambda__2___boxed(lean_object*, 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_Term_toParserDescr_processNonReserved___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__3___closed__1; -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabSyntax___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandNoKindMacroRulesAux___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandOptPrecedence(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: @@ -1087,98 +1130,195 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1(lean_object* x_1, size_t x_2, size_t 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, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -uint8_t x_12; -x_12 = lean_usize_dec_lt(x_3, x_2); -if (x_12 == 0) +uint8_t x_13; +x_13 = lean_usize_dec_lt(x_4, x_3); +if (x_13 == 0) { -lean_object* x_13; -lean_dec(x_9); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_4); -lean_ctor_set(x_13, 1, x_11); -return x_13; +lean_object* x_14; +lean_dec(x_10); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_5); +lean_ctor_set(x_14, 1, x_12); +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; 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; size_t x_51; size_t x_52; -x_14 = lean_ctor_get(x_1, 0); -x_15 = lean_array_uget(x_14, x_3); -lean_inc(x_9); -x_16 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_9, x_10, x_11); +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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; size_t x_58; size_t x_59; +x_15 = lean_ctor_get(x_2, 0); +x_16 = lean_array_uget(x_15, x_4); x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); x_18 = lean_ctor_get(x_16, 1); lean_inc(x_18); lean_dec(x_16); -x_19 = lean_ctor_get(x_9, 10); +x_19 = lean_ctor_get(x_5, 0); lean_inc(x_19); -x_20 = lean_st_ref_get(x_10, x_18); -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_20, 1); +x_20 = lean_ctor_get(x_5, 1); +lean_inc(x_20); +lean_dec(x_5); +lean_inc(x_10); +x_21 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_quoteAutoTactic___spec__1___rarg(x_10, x_11, x_12); +x_22 = lean_ctor_get(x_21, 0); lean_inc(x_22); -lean_dec(x_20); -x_23 = lean_ctor_get(x_21, 0); +x_23 = lean_ctor_get(x_21, 1); lean_inc(x_23); lean_dec(x_21); -x_24 = lean_environment_main_module(x_23); -x_25 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__15; -x_26 = l_Lean_addMacroScope(x_24, x_25, x_19); -x_27 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__11; -x_28 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__19; -lean_inc(x_17); -x_29 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_29, 0, x_17); -lean_ctor_set(x_29, 1, x_27); -lean_ctor_set(x_29, 2, x_26); -lean_ctor_set(x_29, 3, x_28); -x_30 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__26; -x_31 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_31, 0, x_17); -lean_ctor_set(x_31, 1, x_30); -x_32 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; -x_33 = lean_array_push(x_32, x_31); -x_34 = lean_box(2); -x_35 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__25; -x_36 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_36, 0, x_34); +x_24 = lean_ctor_get(x_10, 10); +lean_inc(x_24); +x_25 = lean_st_ref_get(x_11, x_23); +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +lean_dec(x_25); +x_28 = lean_ctor_get(x_26, 0); +lean_inc(x_28); +lean_dec(x_26); +x_29 = lean_environment_main_module(x_28); +x_30 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__15; +x_31 = l_Lean_addMacroScope(x_29, x_30, x_24); +x_32 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__11; +x_33 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__19; +lean_inc(x_22); +x_34 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_34, 0, x_22); +lean_ctor_set(x_34, 1, x_32); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_34, 3, x_33); +x_35 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__26; +x_36 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_36, 0, x_22); lean_ctor_set(x_36, 1, x_35); -lean_ctor_set(x_36, 2, x_33); -x_37 = lean_array_push(x_32, x_36); -x_38 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__23; -x_39 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_39, 0, x_34); -lean_ctor_set(x_39, 1, x_38); -lean_ctor_set(x_39, 2, x_37); -x_40 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__28; -x_41 = lean_array_push(x_40, x_39); -x_42 = lean_array_push(x_41, x_4); -x_43 = lean_array_push(x_42, x_15); -x_44 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_45 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_45, 0, x_34); -lean_ctor_set(x_45, 1, x_44); -lean_ctor_set(x_45, 2, x_43); -x_46 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; -x_47 = lean_array_push(x_46, x_29); -x_48 = lean_array_push(x_47, x_45); -x_49 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +x_37 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; +x_38 = lean_array_push(x_37, x_36); +x_39 = lean_box(2); +x_40 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__25; +x_41 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +lean_ctor_set(x_41, 2, x_38); +x_42 = lean_array_push(x_37, x_41); +x_43 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__23; +x_44 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_44, 0, x_39); +lean_ctor_set(x_44, 1, x_43); +lean_ctor_set(x_44, 2, x_42); +x_45 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__28; +x_46 = lean_array_push(x_45, x_44); +x_47 = lean_array_push(x_46, x_19); +x_48 = lean_array_push(x_47, x_17); +x_49 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; x_50 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_50, 0, x_34); +lean_ctor_set(x_50, 0, x_39); lean_ctor_set(x_50, 1, x_49); lean_ctor_set(x_50, 2, x_48); -x_51 = 1; -x_52 = lean_usize_add(x_3, x_51); -x_3 = x_52; -x_4 = x_50; -x_11 = x_22; +x_51 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; +x_52 = lean_array_push(x_51, x_34); +x_53 = lean_array_push(x_52, x_50); +x_54 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +x_55 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_55, 0, x_39); +lean_ctor_set(x_55, 1, x_54); +lean_ctor_set(x_55, 2, x_53); +x_56 = lean_nat_add(x_20, x_18); +lean_dec(x_18); +lean_dec(x_20); +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_56); +x_58 = 1; +x_59 = lean_usize_add(x_4, x_58); +x_4 = x_59; +x_5 = x_57; +x_12 = x_27; goto _start; } } } +static lean_object* _init_l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__2___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Elab_unsupportedSyntaxExceptionId; +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__2___rarg___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__2___rarg___closed__1; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__2___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__2___rarg___closed__2; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___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) { +_start: +{ +lean_object* x_7; +x_7 = lean_alloc_closure((void*)(l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__2___rarg), 1, 0); +return x_7; +} +} +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("term", 4); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___closed__2; +x_3 = lean_alloc_ctor(1, 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___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_instInhabitedNat; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq(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: { @@ -1193,80 +1333,175 @@ x_12 = lean_unsigned_to_nat(1u); x_13 = lean_nat_dec_eq(x_9, x_12); if (x_13 == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; size_t x_18; lean_object* x_19; size_t x_20; lean_object* x_21; uint8_t x_22; -x_14 = l_Lean_instInhabitedSyntax; +lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_14 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___closed__4; x_15 = lean_array_get(x_14, x_1, x_10); -x_16 = l_Array_toSubarray___rarg(x_1, x_12, x_9); -x_17 = lean_ctor_get(x_16, 2); -lean_inc(x_17); -x_18 = lean_usize_of_nat(x_17); -lean_dec(x_17); -x_19 = lean_ctor_get(x_16, 1); -lean_inc(x_19); -x_20 = lean_usize_of_nat(x_19); +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; size_t x_22; lean_object* x_23; size_t x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_17 = lean_ctor_get(x_15, 0); +x_18 = lean_ctor_get(x_15, 1); +x_19 = l_Array_toSubarray___rarg(x_1, x_12, x_9); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_17); +lean_ctor_set(x_20, 1, x_18); +x_21 = lean_ctor_get(x_19, 2); +lean_inc(x_21); +x_22 = lean_usize_of_nat(x_21); +lean_dec(x_21); +x_23 = lean_ctor_get(x_19, 1); +lean_inc(x_23); +x_24 = lean_usize_of_nat(x_23); +lean_dec(x_23); +x_25 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___closed__3; +x_26 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1(x_25, x_19, x_22, x_24, x_20, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_19); -x_21 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1(x_16, x_18, x_20, x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_16); -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) +x_27 = !lean_is_exclusive(x_26); +if (x_27 == 0) { -return x_21; +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_26, 0); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +lean_ctor_set(x_15, 1, x_30); +lean_ctor_set(x_15, 0, x_29); +lean_ctor_set(x_26, 0, x_15); +return x_26; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_21, 0); -x_24 = lean_ctor_get(x_21, 1); -lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_21); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_23); -lean_ctor_set(x_25, 1, x_24); -return x_25; +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_31 = lean_ctor_get(x_26, 0); +x_32 = lean_ctor_get(x_26, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_26); +x_33 = lean_ctor_get(x_31, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_31, 1); +lean_inc(x_34); +lean_dec(x_31); +lean_ctor_set(x_15, 1, x_34); +lean_ctor_set(x_15, 0, x_33); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_15); +lean_ctor_set(x_35, 1, x_32); +return x_35; } } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; size_t x_41; lean_object* x_42; size_t x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_36 = lean_ctor_get(x_15, 0); +x_37 = lean_ctor_get(x_15, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_15); +x_38 = l_Array_toSubarray___rarg(x_1, x_12, x_9); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_36); +lean_ctor_set(x_39, 1, x_37); +x_40 = lean_ctor_get(x_38, 2); +lean_inc(x_40); +x_41 = lean_usize_of_nat(x_40); +lean_dec(x_40); +x_42 = lean_ctor_get(x_38, 1); +lean_inc(x_42); +x_43 = lean_usize_of_nat(x_42); +lean_dec(x_42); +x_44 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___closed__3; +x_45 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1(x_44, x_38, x_41, x_43, x_39, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_38); +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_45, 1); +lean_inc(x_47); +if (lean_is_exclusive(x_45)) { + lean_ctor_release(x_45, 0); + lean_ctor_release(x_45, 1); + x_48 = x_45; +} else { + lean_dec_ref(x_45); + x_48 = lean_box(0); +} +x_49 = lean_ctor_get(x_46, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_46, 1); +lean_inc(x_50); +lean_dec(x_46); +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +if (lean_is_scalar(x_48)) { + x_52 = lean_alloc_ctor(0, 2, 0); +} else { + x_52 = x_48; +} +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_47); +return x_52; +} +} +else +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_dec(x_9); lean_dec(x_6); -x_26 = l_Lean_instInhabitedSyntax; -x_27 = lean_array_get(x_26, x_1, x_10); +x_53 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___closed__4; +x_54 = lean_array_get(x_53, x_1, x_10); lean_dec(x_1); -x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_27); -lean_ctor_set(x_28, 1, x_8); -return x_28; +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_8); +return x_55; } } else { -lean_object* x_29; +lean_object* x_56; lean_dec(x_9); lean_dec(x_6); lean_dec(x_1); -x_29 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13___rarg(x_8); -return x_29; +x_56 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__2___rarg(x_8); +return x_56; } } } -LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__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, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__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, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -size_t x_12; size_t x_13; lean_object* x_14; -x_12 = lean_unbox_usize(x_2); -lean_dec(x_2); +size_t x_13; size_t x_14; lean_object* x_15; x_13 = lean_unbox_usize(x_3); lean_dec(x_3); -x_14 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1(x_1, x_12, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); +x_14 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_15 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__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_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__2(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_6); lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -return x_14; +return x_7; } } LEAN_EXPORT lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___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) { @@ -1369,26 +1604,309 @@ x_2 = lean_alloc_closure((void*)(l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_ return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_withNestedParser___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, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +static lean_object* _init_l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__1() { _start: { -uint8_t x_11; -x_11 = !lean_is_exclusive(x_2); -if (x_11 == 0) +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("UnhygienicMain", 14); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__2() { +_start: { -uint8_t x_12; lean_object* x_13; -x_12 = 0; -lean_ctor_set_uint8(x_2, sizeof(void*)*1, x_12); -lean_ctor_set_uint8(x_2, sizeof(void*)*1 + 1, x_12); -x_13 = lean_apply_9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_13; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; } -else +} +static lean_object* _init_l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__3() { +_start: { -lean_object* x_14; uint8_t x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; -x_14 = lean_ctor_get(x_2, 0); -x_15 = lean_ctor_get_uint8(x_2, sizeof(void*)*1 + 2); -lean_inc(x_14); +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("ParserDescr.unary", 17); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__3; +x_2 = lean_string_utf8_byte_size(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__3; +x_2 = lean_unsigned_to_nat(0u); +x_3 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__4; +x_4 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("unary", 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__13; +x_2 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__6; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__16; +x_2 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__6; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__8; +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_Elab_Term_ensureUnaryOutput___lambda__1___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__9; +x_3 = lean_alloc_ctor(1, 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_Elab_Term_ensureUnaryOutput___lambda__1___closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("group", 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__11; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__12; +x_3 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__12; +x_2 = l_Lean_quoteNameMk(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; +x_2 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__14; +x_3 = lean_array_push(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__16() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(".", 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__17() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("`", 1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_ensureUnaryOutput___lambda__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; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_5 = lean_ctor_get(x_3, 1); +lean_inc(x_5); +lean_dec(x_3); +x_6 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__2; +x_7 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__7; +x_8 = l_Lean_addMacroScope(x_6, x_7, x_5); +x_9 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__5; +x_10 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__10; +x_11 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_11, 0, x_2); +lean_ctor_set(x_11, 1, x_9); +lean_ctor_set(x_11, 2, x_8); +lean_ctor_set(x_11, 3, x_10); +x_12 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; +x_13 = lean_array_push(x_12, x_11); +x_14 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__13; +if (lean_obj_tag(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 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__15; +x_16 = lean_array_push(x_15, x_1); +x_17 = lean_box(2); +x_18 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; +x_19 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +lean_ctor_set(x_19, 2, x_16); +x_20 = lean_array_push(x_13, x_19); +x_21 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +x_22 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_22, 0, x_17); +lean_ctor_set(x_22, 1, x_21); +lean_ctor_set(x_22, 2, x_20); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_4); +return x_23; +} +else +{ +lean_object* 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; 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; +x_24 = lean_ctor_get(x_14, 0); +lean_inc(x_24); +x_25 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__16; +x_26 = l_String_intercalate(x_25, x_24); +x_27 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__17; +x_28 = lean_string_append(x_27, x_26); +lean_dec(x_26); +x_29 = lean_box(2); +x_30 = l_Lean_Syntax_mkNameLit(x_28, x_29); +x_31 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; +x_32 = lean_array_push(x_31, x_30); +x_33 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__23; +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_29); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_32); +x_35 = lean_array_push(x_12, x_34); +x_36 = lean_array_push(x_35, x_1); +x_37 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; +x_38 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_38, 0, x_29); +lean_ctor_set(x_38, 1, x_37); +lean_ctor_set(x_38, 2, x_36); +x_39 = lean_array_push(x_13, x_38); +x_40 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +x_41 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_41, 0, x_29); +lean_ctor_set(x_41, 1, x_40); +lean_ctor_set(x_41, 2, x_39); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_4); +return x_42; +} +} +} +static lean_object* _init_l_Lean_Elab_Term_ensureUnaryOutput___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Level_PP_Result_quote___spec__1), 2, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_ensureUnaryOutput(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +lean_dec(x_1); +x_4 = lean_unsigned_to_nat(1u); +x_5 = lean_nat_dec_eq(x_3, x_4); +lean_dec(x_3); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_6 = lean_alloc_closure((void*)(l_Lean_Elab_Term_ensureUnaryOutput___lambda__1), 4, 1); +lean_closure_set(x_6, 0, x_2); +x_7 = l_Lean_Elab_Term_ensureUnaryOutput___closed__1; +x_8 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Unhygienic_instMonadQuotationUnhygienic___spec__2___rarg), 4, 2); +lean_closure_set(x_8, 0, x_7); +lean_closure_set(x_8, 1, x_6); +x_9 = l_Lean_Unhygienic_run___rarg(x_8); +return x_9; +} +else +{ +return x_2; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_withNestedParser(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_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; +x_11 = !lean_is_exclusive(x_2); +if (x_11 == 0) +{ +uint8_t x_12; lean_object* x_13; +x_12 = 0; +lean_ctor_set_uint8(x_2, sizeof(void*)*1, x_12); +lean_ctor_set_uint8(x_2, sizeof(void*)*1 + 1, x_12); +x_13 = lean_apply_9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_13; +} +else +{ +lean_object* x_14; uint8_t x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; +x_14 = lean_ctor_get(x_2, 0); +x_15 = lean_ctor_get_uint8(x_2, sizeof(void*)*1 + 2); +lean_inc(x_14); lean_dec(x_2); x_16 = 0; x_17 = lean_alloc_ctor(0, 1, 3); @@ -1401,14 +1919,6 @@ return x_18; } } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_withNestedParser(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_withNestedParser___rarg), 10, 0); -return x_2; -} -} LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOptionM___at_Lean_Elab_Term_checkLeftRec___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_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { @@ -1963,31 +2473,11 @@ lean_ctor_set(x_13, 1, x_10); return x_13; } } -static lean_object* _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_checkLeftRec___spec__8___rarg___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Elab_unsupportedSyntaxExceptionId; -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_checkLeftRec___spec__8___rarg___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_checkLeftRec___spec__8___rarg___closed__1; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; -} -} LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_checkLeftRec___spec__8___rarg(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_checkLeftRec___spec__8___rarg___closed__2; +x_2 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__2___rarg___closed__2; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); @@ -3747,7 +4237,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_ _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_checkLeftRec___spec__8___rarg___closed__2; +x_2 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__2___rarg___closed__2; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); @@ -3996,7 +4486,7 @@ x_20 = lean_st_ref_get(x_9, x_18); x_21 = !lean_is_exclusive(x_20); if (x_21 == 0) { -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; 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; lean_object* x_44; lean_object* x_45; +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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; x_22 = lean_ctor_get(x_20, 0); x_23 = lean_ctor_get(x_22, 0); lean_inc(x_23); @@ -4041,64 +4531,70 @@ x_45 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_45, 0, x_30); lean_ctor_set(x_45, 1, x_44); lean_ctor_set(x_45, 2, x_43); -lean_ctor_set(x_20, 0, x_45); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_11); +lean_ctor_set(x_20, 0, x_46); return x_20; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_46 = lean_ctor_get(x_20, 0); -x_47 = lean_ctor_get(x_20, 1); +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; +x_47 = lean_ctor_get(x_20, 0); +x_48 = lean_ctor_get(x_20, 1); +lean_inc(x_48); lean_inc(x_47); -lean_inc(x_46); lean_dec(x_20); -x_48 = lean_ctor_get(x_46, 0); -lean_inc(x_48); -lean_dec(x_46); -x_49 = lean_environment_main_module(x_48); -x_50 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__5; -lean_inc(x_19); +x_49 = lean_ctor_get(x_47, 0); lean_inc(x_49); -x_51 = l_Lean_addMacroScope(x_49, x_50, x_19); -x_52 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__3; -x_53 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__8; +lean_dec(x_47); +x_50 = lean_environment_main_module(x_49); +x_51 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__5; +lean_inc(x_19); +lean_inc(x_50); +x_52 = l_Lean_addMacroScope(x_50, x_51, x_19); +x_53 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__3; +x_54 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__8; lean_inc(x_17); -x_54 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_54, 0, x_17); -lean_ctor_set(x_54, 1, x_52); -lean_ctor_set(x_54, 2, x_51); -lean_ctor_set(x_54, 3, x_53); -x_55 = lean_box(2); -x_56 = l_Lean_Syntax_mkStrLit(x_15, x_55); +x_55 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_55, 0, x_17); +lean_ctor_set(x_55, 1, x_53); +lean_ctor_set(x_55, 2, x_52); +lean_ctor_set(x_55, 3, x_54); +x_56 = lean_box(2); +x_57 = l_Lean_Syntax_mkStrLit(x_15, x_56); lean_dec(x_15); -x_57 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__12; -x_58 = l_Lean_addMacroScope(x_49, x_57, x_19); -x_59 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__11; -x_60 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__17; -x_61 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_61, 0, x_17); -lean_ctor_set(x_61, 1, x_59); -lean_ctor_set(x_61, 2, x_58); -lean_ctor_set(x_61, 3, x_60); -x_62 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; -x_63 = lean_array_push(x_62, x_56); -x_64 = lean_array_push(x_63, x_61); -x_65 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_66 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_66, 0, x_55); -lean_ctor_set(x_66, 1, x_65); -lean_ctor_set(x_66, 2, x_64); -x_67 = lean_array_push(x_62, x_54); -x_68 = lean_array_push(x_67, x_66); -x_69 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_70 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_70, 0, x_55); -lean_ctor_set(x_70, 1, x_69); -lean_ctor_set(x_70, 2, x_68); -x_71 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_47); -return x_71; +x_58 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__12; +x_59 = l_Lean_addMacroScope(x_50, x_58, x_19); +x_60 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__11; +x_61 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__17; +x_62 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_62, 0, x_17); +lean_ctor_set(x_62, 1, x_60); +lean_ctor_set(x_62, 2, x_59); +lean_ctor_set(x_62, 3, x_61); +x_63 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; +x_64 = lean_array_push(x_63, x_57); +x_65 = lean_array_push(x_64, x_62); +x_66 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; +x_67 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_67, 0, x_56); +lean_ctor_set(x_67, 1, x_66); +lean_ctor_set(x_67, 2, x_65); +x_68 = lean_array_push(x_63, x_55); +x_69 = lean_array_push(x_68, x_67); +x_70 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +x_71 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_71, 0, x_56); +lean_ctor_set(x_71, 1, x_70); +lean_ctor_set(x_71, 2, x_69); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set(x_72, 1, x_11); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_72); +lean_ctor_set(x_73, 1, x_48); +return x_73; } } } @@ -4370,150 +4866,158 @@ return x_3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processAtom___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, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_12; uint8_t x_60; uint8_t x_61; uint8_t x_62; -x_60 = lean_ctor_get_uint8(x_3, sizeof(void*)*1 + 2); -x_61 = 0; -x_62 = l___private_Lean_Parser_Basic_0__Lean_Parser_beqLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8845_(x_60, x_61); -if (x_62 == 0) +lean_object* x_12; uint8_t x_64; uint8_t x_65; uint8_t x_66; +x_64 = lean_ctor_get_uint8(x_3, sizeof(void*)*1 + 2); +x_65 = 0; +x_66 = l___private_Lean_Parser_Basic_0__Lean_Parser_beqLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8845_(x_64, x_65); +if (x_66 == 0) { -uint8_t x_63; -x_63 = lean_ctor_get_uint8(x_3, sizeof(void*)*1); -if (x_63 == 0) +uint8_t x_67; +x_67 = lean_ctor_get_uint8(x_3, sizeof(void*)*1); +if (x_67 == 0) { -lean_object* x_64; -x_64 = lean_box(0); -x_12 = x_64; -goto block_59; +lean_object* x_68; +x_68 = lean_box(0); +x_12 = x_68; +goto block_63; } else { -lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; uint8_t x_70; +lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; uint8_t x_74; lean_inc(x_9); -x_65 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_toParserDescr_processNonReserved___spec__2___rarg(x_9, x_10, x_11); -x_66 = lean_ctor_get(x_65, 0); -lean_inc(x_66); -x_67 = lean_ctor_get(x_65, 1); -lean_inc(x_67); -lean_dec(x_65); -x_68 = lean_ctor_get(x_9, 10); -lean_inc(x_68); +x_69 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_toParserDescr_processNonReserved___spec__2___rarg(x_9, x_10, x_11); +x_70 = lean_ctor_get(x_69, 0); +lean_inc(x_70); +x_71 = lean_ctor_get(x_69, 1); +lean_inc(x_71); +lean_dec(x_69); +x_72 = lean_ctor_get(x_9, 10); +lean_inc(x_72); lean_dec(x_9); -x_69 = lean_st_ref_get(x_10, x_67); -x_70 = !lean_is_exclusive(x_69); -if (x_70 == 0) +x_73 = lean_st_ref_get(x_10, x_71); +x_74 = !lean_is_exclusive(x_73); +if (x_74 == 0) { -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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; -x_71 = lean_ctor_get(x_69, 0); -x_72 = lean_ctor_get(x_71, 0); +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; +x_75 = lean_ctor_get(x_73, 0); +x_76 = lean_ctor_get(x_75, 0); +lean_inc(x_76); +lean_dec(x_75); +x_77 = lean_environment_main_module(x_76); +x_78 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__5; lean_inc(x_72); -lean_dec(x_71); -x_73 = lean_environment_main_module(x_72); -x_74 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__5; -lean_inc(x_68); -lean_inc(x_73); -x_75 = l_Lean_addMacroScope(x_73, x_74, x_68); -x_76 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__3; -x_77 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__8; -lean_inc(x_66); -x_78 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_78, 0, x_66); -lean_ctor_set(x_78, 1, x_76); -lean_ctor_set(x_78, 2, x_75); -lean_ctor_set(x_78, 3, x_77); -x_79 = lean_box(2); -x_80 = l_Lean_Syntax_mkStrLit(x_1, x_79); -x_81 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__12; -x_82 = l_Lean_addMacroScope(x_73, x_81, x_68); -x_83 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__11; -x_84 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__17; -x_85 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_85, 0, x_66); -lean_ctor_set(x_85, 1, x_83); -lean_ctor_set(x_85, 2, x_82); -lean_ctor_set(x_85, 3, x_84); -x_86 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; -x_87 = lean_array_push(x_86, x_80); -x_88 = lean_array_push(x_87, x_85); -x_89 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_90 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_90, 0, x_79); -lean_ctor_set(x_90, 1, x_89); -lean_ctor_set(x_90, 2, x_88); -x_91 = lean_array_push(x_86, x_78); -x_92 = lean_array_push(x_91, x_90); -x_93 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +lean_inc(x_77); +x_79 = l_Lean_addMacroScope(x_77, x_78, x_72); +x_80 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__3; +x_81 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__8; +lean_inc(x_70); +x_82 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_82, 0, x_70); +lean_ctor_set(x_82, 1, x_80); +lean_ctor_set(x_82, 2, x_79); +lean_ctor_set(x_82, 3, x_81); +x_83 = lean_box(2); +x_84 = l_Lean_Syntax_mkStrLit(x_1, x_83); +x_85 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__12; +x_86 = l_Lean_addMacroScope(x_77, x_85, x_72); +x_87 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__11; +x_88 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__17; +x_89 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_89, 0, x_70); +lean_ctor_set(x_89, 1, x_87); +lean_ctor_set(x_89, 2, x_86); +lean_ctor_set(x_89, 3, x_88); +x_90 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; +x_91 = lean_array_push(x_90, x_84); +x_92 = lean_array_push(x_91, x_89); +x_93 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; x_94 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_94, 0, x_79); +lean_ctor_set(x_94, 0, x_83); lean_ctor_set(x_94, 1, x_93); lean_ctor_set(x_94, 2, x_92); -lean_ctor_set(x_69, 0, x_94); -return x_69; +x_95 = lean_array_push(x_90, x_82); +x_96 = lean_array_push(x_95, x_94); +x_97 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +x_98 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_98, 0, x_83); +lean_ctor_set(x_98, 1, x_97); +lean_ctor_set(x_98, 2, x_96); +x_99 = lean_unsigned_to_nat(1u); +x_100 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_100, 0, x_98); +lean_ctor_set(x_100, 1, x_99); +lean_ctor_set(x_73, 0, x_100); +return x_73; } else { -lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; -x_95 = lean_ctor_get(x_69, 0); -x_96 = lean_ctor_get(x_69, 1); -lean_inc(x_96); -lean_inc(x_95); -lean_dec(x_69); -x_97 = lean_ctor_get(x_95, 0); -lean_inc(x_97); -lean_dec(x_95); -x_98 = lean_environment_main_module(x_97); -x_99 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__5; -lean_inc(x_68); -lean_inc(x_98); -x_100 = l_Lean_addMacroScope(x_98, x_99, x_68); -x_101 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__3; -x_102 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__8; -lean_inc(x_66); -x_103 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_103, 0, x_66); -lean_ctor_set(x_103, 1, x_101); -lean_ctor_set(x_103, 2, x_100); -lean_ctor_set(x_103, 3, x_102); -x_104 = lean_box(2); -x_105 = l_Lean_Syntax_mkStrLit(x_1, x_104); -x_106 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__12; -x_107 = l_Lean_addMacroScope(x_98, x_106, x_68); -x_108 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__11; -x_109 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__17; -x_110 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_110, 0, x_66); -lean_ctor_set(x_110, 1, x_108); -lean_ctor_set(x_110, 2, x_107); -lean_ctor_set(x_110, 3, x_109); -x_111 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; -x_112 = lean_array_push(x_111, x_105); -x_113 = lean_array_push(x_112, x_110); -x_114 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_115 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_115, 0, x_104); -lean_ctor_set(x_115, 1, x_114); -lean_ctor_set(x_115, 2, x_113); -x_116 = lean_array_push(x_111, x_103); -x_117 = lean_array_push(x_116, x_115); -x_118 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_119 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_119, 0, x_104); -lean_ctor_set(x_119, 1, x_118); -lean_ctor_set(x_119, 2, x_117); -x_120 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_120, 0, x_119); -lean_ctor_set(x_120, 1, x_96); -return x_120; +lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; +x_101 = lean_ctor_get(x_73, 0); +x_102 = lean_ctor_get(x_73, 1); +lean_inc(x_102); +lean_inc(x_101); +lean_dec(x_73); +x_103 = lean_ctor_get(x_101, 0); +lean_inc(x_103); +lean_dec(x_101); +x_104 = lean_environment_main_module(x_103); +x_105 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__5; +lean_inc(x_72); +lean_inc(x_104); +x_106 = l_Lean_addMacroScope(x_104, x_105, x_72); +x_107 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__3; +x_108 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__8; +lean_inc(x_70); +x_109 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_109, 0, x_70); +lean_ctor_set(x_109, 1, x_107); +lean_ctor_set(x_109, 2, x_106); +lean_ctor_set(x_109, 3, x_108); +x_110 = lean_box(2); +x_111 = l_Lean_Syntax_mkStrLit(x_1, x_110); +x_112 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__12; +x_113 = l_Lean_addMacroScope(x_104, x_112, x_72); +x_114 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__11; +x_115 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__17; +x_116 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_116, 0, x_70); +lean_ctor_set(x_116, 1, x_114); +lean_ctor_set(x_116, 2, x_113); +lean_ctor_set(x_116, 3, x_115); +x_117 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; +x_118 = lean_array_push(x_117, x_111); +x_119 = lean_array_push(x_118, x_116); +x_120 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; +x_121 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_121, 0, x_110); +lean_ctor_set(x_121, 1, x_120); +lean_ctor_set(x_121, 2, x_119); +x_122 = lean_array_push(x_117, x_109); +x_123 = lean_array_push(x_122, x_121); +x_124 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +x_125 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_125, 0, x_110); +lean_ctor_set(x_125, 1, x_124); +lean_ctor_set(x_125, 2, x_123); +x_126 = lean_unsigned_to_nat(1u); +x_127 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +x_128 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_128, 0, x_127); +lean_ctor_set(x_128, 1, x_102); +return x_128; } } } else { -lean_object* x_121; -x_121 = lean_box(0); -x_12 = x_121; -goto block_59; +lean_object* x_129; +x_129 = lean_box(0); +x_12 = x_129; +goto block_63; } -block_59: +block_63: { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_dec(x_12); @@ -4531,7 +5035,7 @@ x_17 = lean_st_ref_get(x_10, x_15); x_18 = !lean_is_exclusive(x_17); if (x_18 == 0) { -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; 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_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; 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; x_19 = lean_ctor_get(x_17, 0); x_20 = lean_ctor_get(x_19, 0); lean_inc(x_20); @@ -4563,51 +5067,59 @@ x_37 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_37, 0, x_27); lean_ctor_set(x_37, 1, x_36); lean_ctor_set(x_37, 2, x_35); -lean_ctor_set(x_17, 0, x_37); +x_38 = lean_unsigned_to_nat(1u); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +lean_ctor_set(x_17, 0, x_39); return x_17; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_38 = lean_ctor_get(x_17, 0); -x_39 = lean_ctor_get(x_17, 1); -lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_17); -x_40 = lean_ctor_get(x_38, 0); +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_40 = lean_ctor_get(x_17, 0); +x_41 = lean_ctor_get(x_17, 1); +lean_inc(x_41); lean_inc(x_40); -lean_dec(x_38); -x_41 = lean_environment_main_module(x_40); -x_42 = l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__5; -x_43 = l_Lean_addMacroScope(x_41, x_42, x_16); -x_44 = l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__3; -x_45 = l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__8; -x_46 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_46, 0, x_14); -lean_ctor_set(x_46, 1, x_44); -lean_ctor_set(x_46, 2, x_43); -lean_ctor_set(x_46, 3, x_45); -x_47 = lean_box(2); -x_48 = l_Lean_Syntax_mkStrLit(x_1, x_47); -x_49 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; -x_50 = lean_array_push(x_49, x_48); -x_51 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_52 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_52, 0, x_47); -lean_ctor_set(x_52, 1, x_51); -lean_ctor_set(x_52, 2, x_50); -x_53 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; -x_54 = lean_array_push(x_53, x_46); -x_55 = lean_array_push(x_54, x_52); -x_56 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_57 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_57, 0, x_47); -lean_ctor_set(x_57, 1, x_56); -lean_ctor_set(x_57, 2, x_55); -x_58 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_58, 0, x_57); -lean_ctor_set(x_58, 1, x_39); -return x_58; +lean_dec(x_17); +x_42 = lean_ctor_get(x_40, 0); +lean_inc(x_42); +lean_dec(x_40); +x_43 = lean_environment_main_module(x_42); +x_44 = l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__5; +x_45 = l_Lean_addMacroScope(x_43, x_44, x_16); +x_46 = l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__3; +x_47 = l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__8; +x_48 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_48, 0, x_14); +lean_ctor_set(x_48, 1, x_46); +lean_ctor_set(x_48, 2, x_45); +lean_ctor_set(x_48, 3, x_47); +x_49 = lean_box(2); +x_50 = l_Lean_Syntax_mkStrLit(x_1, x_49); +x_51 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; +x_52 = lean_array_push(x_51, x_50); +x_53 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; +x_54 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_54, 0, x_49); +lean_ctor_set(x_54, 1, x_53); +lean_ctor_set(x_54, 2, x_52); +x_55 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; +x_56 = lean_array_push(x_55, x_48); +x_57 = lean_array_push(x_56, x_54); +x_58 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +x_59 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_59, 0, x_49); +lean_ctor_set(x_59, 1, x_58); +lean_ctor_set(x_59, 2, x_57); +x_60 = lean_unsigned_to_nat(1u); +x_61 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_61, 0, x_59); +lean_ctor_set(x_61, 1, x_60); +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_61); +lean_ctor_set(x_62, 1, x_41); +return x_62; } } } @@ -4795,22 +5307,6 @@ lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__8() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(".", 1); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__9() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("`", 1); -return x_1; -} -} LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processParserCategory___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, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { @@ -4899,7 +5395,7 @@ lean_dec(x_11); x_21 = !lean_is_exclusive(x_20); if (x_21 == 0) { -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; 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_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; 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; x_22 = lean_ctor_get(x_20, 0); x_23 = lean_ctor_get(x_22, 0); lean_inc(x_23); @@ -4918,28 +5414,30 @@ lean_ctor_set(x_30, 3, x_29); lean_inc(x_2); x_31 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_27, x_2); x_32 = l_Nat_repr(x_16); -x_33 = l_Lean_numLitKind; -x_34 = lean_box(2); -x_35 = l_Lean_Syntax_mkLit(x_33, x_32, x_34); -x_36 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; -x_37 = lean_array_push(x_36, x_30); +x_33 = lean_box(2); +x_34 = l_Lean_Syntax_mkNumLit(x_32, x_33); +x_35 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; +x_36 = lean_array_push(x_35, x_30); if (lean_obj_tag(x_31) == 0) { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_38 = l___private_Init_Meta_0__Lean_quoteNameMk(x_2); -x_39 = lean_array_push(x_36, x_38); -x_40 = lean_array_push(x_39, x_35); -x_41 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_42 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_42, 0, x_34); -lean_ctor_set(x_42, 1, x_41); -lean_ctor_set(x_42, 2, x_40); -x_43 = lean_array_push(x_37, x_42); -x_44 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_45 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_45, 0, x_34); -lean_ctor_set(x_45, 1, x_44); -lean_ctor_set(x_45, 2, x_43); +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; lean_object* x_44; lean_object* x_45; +x_37 = l_Lean_quoteNameMk(x_2); +x_38 = lean_array_push(x_35, x_37); +x_39 = lean_array_push(x_38, x_34); +x_40 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; +x_41 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_41, 0, x_33); +lean_ctor_set(x_41, 1, x_40); +lean_ctor_set(x_41, 2, x_39); +x_42 = lean_array_push(x_36, x_41); +x_43 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +x_44 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_44, 0, x_33); +lean_ctor_set(x_44, 1, x_43); +lean_ctor_set(x_44, 2, x_42); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_13); lean_ctor_set(x_20, 0, x_45); return x_20; } @@ -4950,40 +5448,42 @@ lean_dec(x_2); x_46 = lean_ctor_get(x_31, 0); lean_inc(x_46); lean_dec(x_31); -x_47 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__8; +x_47 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__16; x_48 = l_String_intercalate(x_47, x_46); -x_49 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__9; +x_49 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__17; x_50 = lean_string_append(x_49, x_48); lean_dec(x_48); -x_51 = l_Lean_nameLitKind; -x_52 = l_Lean_Syntax_mkLit(x_51, x_50, x_34); -x_53 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; -x_54 = lean_array_push(x_53, x_52); -x_55 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__23; -x_56 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_56, 0, x_34); -lean_ctor_set(x_56, 1, x_55); -lean_ctor_set(x_56, 2, x_54); -x_57 = lean_array_push(x_36, x_56); -x_58 = lean_array_push(x_57, x_35); -x_59 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_60 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_60, 0, x_34); -lean_ctor_set(x_60, 1, x_59); -lean_ctor_set(x_60, 2, x_58); -x_61 = lean_array_push(x_37, x_60); -x_62 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_63 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_63, 0, x_34); -lean_ctor_set(x_63, 1, x_62); -lean_ctor_set(x_63, 2, x_61); +x_51 = l_Lean_Syntax_mkNameLit(x_50, x_33); +x_52 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; +x_53 = lean_array_push(x_52, x_51); +x_54 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__23; +x_55 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_55, 0, x_33); +lean_ctor_set(x_55, 1, x_54); +lean_ctor_set(x_55, 2, x_53); +x_56 = lean_array_push(x_35, x_55); +x_57 = lean_array_push(x_56, x_34); +x_58 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; +x_59 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_59, 0, x_33); +lean_ctor_set(x_59, 1, x_58); +lean_ctor_set(x_59, 2, x_57); +x_60 = lean_array_push(x_36, x_59); +x_61 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +x_62 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_62, 0, x_33); +lean_ctor_set(x_62, 1, x_61); +lean_ctor_set(x_62, 2, x_60); +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_62); +lean_ctor_set(x_63, 1, x_13); lean_ctor_set(x_20, 0, x_63); return x_20; } } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; 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; lean_object* x_79; lean_object* x_80; +lean_object* x_64; lean_object* x_65; lean_object* x_66; 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; lean_object* x_79; x_64 = lean_ctor_get(x_20, 0); x_65 = lean_ctor_get(x_20, 1); lean_inc(x_65); @@ -5006,28 +5506,30 @@ lean_ctor_set(x_73, 3, x_72); lean_inc(x_2); x_74 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_70, x_2); x_75 = l_Nat_repr(x_16); -x_76 = l_Lean_numLitKind; -x_77 = lean_box(2); -x_78 = l_Lean_Syntax_mkLit(x_76, x_75, x_77); -x_79 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; -x_80 = lean_array_push(x_79, x_73); +x_76 = lean_box(2); +x_77 = l_Lean_Syntax_mkNumLit(x_75, x_76); +x_78 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; +x_79 = lean_array_push(x_78, x_73); if (lean_obj_tag(x_74) == 0) { -lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_81 = l___private_Init_Meta_0__Lean_quoteNameMk(x_2); -x_82 = lean_array_push(x_79, x_81); -x_83 = lean_array_push(x_82, x_78); -x_84 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_85 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_85, 0, x_77); -lean_ctor_set(x_85, 1, x_84); -lean_ctor_set(x_85, 2, x_83); -x_86 = lean_array_push(x_80, x_85); -x_87 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_88 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_88, 0, x_77); -lean_ctor_set(x_88, 1, x_87); -lean_ctor_set(x_88, 2, x_86); +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_80 = l_Lean_quoteNameMk(x_2); +x_81 = lean_array_push(x_78, x_80); +x_82 = lean_array_push(x_81, x_77); +x_83 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; +x_84 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_84, 0, x_76); +lean_ctor_set(x_84, 1, x_83); +lean_ctor_set(x_84, 2, x_82); +x_85 = lean_array_push(x_79, x_84); +x_86 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +x_87 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_87, 0, x_76); +lean_ctor_set(x_87, 1, x_86); +lean_ctor_set(x_87, 2, x_85); +x_88 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_88, 0, x_87); +lean_ctor_set(x_88, 1, x_13); x_89 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_89, 0, x_88); lean_ctor_set(x_89, 1, x_65); @@ -5040,33 +5542,35 @@ lean_dec(x_2); x_90 = lean_ctor_get(x_74, 0); lean_inc(x_90); lean_dec(x_74); -x_91 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__8; +x_91 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__16; x_92 = l_String_intercalate(x_91, x_90); -x_93 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__9; +x_93 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__17; x_94 = lean_string_append(x_93, x_92); lean_dec(x_92); -x_95 = l_Lean_nameLitKind; -x_96 = l_Lean_Syntax_mkLit(x_95, x_94, x_77); -x_97 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; -x_98 = lean_array_push(x_97, x_96); -x_99 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__23; -x_100 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_100, 0, x_77); -lean_ctor_set(x_100, 1, x_99); -lean_ctor_set(x_100, 2, x_98); -x_101 = lean_array_push(x_79, x_100); -x_102 = lean_array_push(x_101, x_78); -x_103 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_104 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_104, 0, x_77); -lean_ctor_set(x_104, 1, x_103); -lean_ctor_set(x_104, 2, x_102); -x_105 = lean_array_push(x_80, x_104); -x_106 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_107 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_107, 0, x_77); -lean_ctor_set(x_107, 1, x_106); -lean_ctor_set(x_107, 2, x_105); +x_95 = l_Lean_Syntax_mkNameLit(x_94, x_76); +x_96 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; +x_97 = lean_array_push(x_96, x_95); +x_98 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__23; +x_99 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_99, 0, x_76); +lean_ctor_set(x_99, 1, x_98); +lean_ctor_set(x_99, 2, x_97); +x_100 = lean_array_push(x_78, x_99); +x_101 = lean_array_push(x_100, x_77); +x_102 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; +x_103 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_103, 0, x_76); +lean_ctor_set(x_103, 1, x_102); +lean_ctor_set(x_103, 2, x_101); +x_104 = lean_array_push(x_79, x_103); +x_105 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +x_106 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_106, 0, x_76); +lean_ctor_set(x_106, 1, x_105); +lean_ctor_set(x_106, 2, x_104); +x_107 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_107, 0, x_106); +lean_ctor_set(x_107, 1, x_13); x_108 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_108, 0, x_107); lean_ctor_set(x_108, 1, x_65); @@ -5224,84 +5728,89 @@ lean_dec(x_1); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__5(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_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_process___spec__3(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_11; lean_object* x_12; uint8_t x_13; -x_11 = lean_ctor_get(x_8, 5); -x_12 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_6, x_7, x_8, x_9, x_10); -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) +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; uint8_t x_16; +x_9 = lean_ctor_get(x_6, 5); +x_10 = lean_ctor_get(x_2, 1); +lean_inc(x_10); +lean_inc(x_10); +x_11 = l_Lean_Elab_getBetterRef(x_9, x_10); +x_12 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_4, x_5, x_6, x_7, x_8); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(x_13, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_14); +lean_dec(x_2); +lean_dec(x_10); +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) { -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_12, 0); -lean_inc(x_11); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_11); -lean_ctor_set(x_15, 1, x_14); -lean_ctor_set_tag(x_12, 1); -lean_ctor_set(x_12, 0, x_15); -return x_12; +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_15, 0); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_11); +lean_ctor_set(x_18, 1, x_17); +lean_ctor_set_tag(x_15, 1); +lean_ctor_set(x_15, 0, x_18); +return x_15; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_16 = lean_ctor_get(x_12, 0); -x_17 = lean_ctor_get(x_12, 1); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_12); -lean_inc(x_11); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_11); -lean_ctor_set(x_18, 1, x_16); -x_19 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_17); -return x_19; +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_19 = lean_ctor_get(x_15, 0); +x_20 = lean_ctor_get(x_15, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_15); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_11); +lean_ctor_set(x_21, 1, x_19); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_20); +return x_22; } } } -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__4(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_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_toParserDescr_process___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_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -uint8_t x_12; -x_12 = !lean_is_exclusive(x_9); -if (x_12 == 0) +uint8_t x_10; +x_10 = !lean_is_exclusive(x_7); +if (x_10 == 0) { -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = lean_ctor_get(x_9, 5); -x_14 = l_Lean_replaceRef(x_1, x_13); -lean_dec(x_13); +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_ctor_get(x_7, 5); +x_12 = l_Lean_replaceRef(x_1, x_11); +lean_dec(x_11); lean_dec(x_1); -lean_ctor_set(x_9, 5, x_14); -x_15 = l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__5(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); +lean_ctor_set(x_7, 5, x_12); +x_13 = l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_process___spec__3(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -return x_15; +return x_13; } else { -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; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_16 = lean_ctor_get(x_9, 0); -x_17 = lean_ctor_get(x_9, 1); -x_18 = lean_ctor_get(x_9, 2); -x_19 = lean_ctor_get(x_9, 3); -x_20 = lean_ctor_get(x_9, 4); -x_21 = lean_ctor_get(x_9, 5); -x_22 = lean_ctor_get(x_9, 6); -x_23 = lean_ctor_get(x_9, 7); -x_24 = lean_ctor_get(x_9, 8); -x_25 = lean_ctor_get(x_9, 9); -x_26 = lean_ctor_get(x_9, 10); -lean_inc(x_26); -lean_inc(x_25); +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; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_14 = lean_ctor_get(x_7, 0); +x_15 = lean_ctor_get(x_7, 1); +x_16 = lean_ctor_get(x_7, 2); +x_17 = lean_ctor_get(x_7, 3); +x_18 = lean_ctor_get(x_7, 4); +x_19 = lean_ctor_get(x_7, 5); +x_20 = lean_ctor_get(x_7, 6); +x_21 = lean_ctor_get(x_7, 7); +x_22 = lean_ctor_get(x_7, 8); +x_23 = lean_ctor_get(x_7, 9); +x_24 = lean_ctor_get(x_7, 10); lean_inc(x_24); lean_inc(x_23); lean_inc(x_22); @@ -5311,341 +5820,312 @@ lean_inc(x_19); lean_inc(x_18); lean_inc(x_17); lean_inc(x_16); -lean_dec(x_9); -x_27 = l_Lean_replaceRef(x_1, x_21); -lean_dec(x_21); +lean_inc(x_15); +lean_inc(x_14); +lean_dec(x_7); +x_25 = l_Lean_replaceRef(x_1, x_19); +lean_dec(x_19); lean_dec(x_1); -x_28 = lean_alloc_ctor(0, 11, 0); -lean_ctor_set(x_28, 0, x_16); -lean_ctor_set(x_28, 1, x_17); -lean_ctor_set(x_28, 2, x_18); -lean_ctor_set(x_28, 3, x_19); -lean_ctor_set(x_28, 4, x_20); -lean_ctor_set(x_28, 5, x_27); -lean_ctor_set(x_28, 6, x_22); -lean_ctor_set(x_28, 7, x_23); -lean_ctor_set(x_28, 8, x_24); -lean_ctor_set(x_28, 9, x_25); -lean_ctor_set(x_28, 10, x_26); -x_29 = l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__5(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_28, x_10, x_11); -lean_dec(x_10); -lean_dec(x_28); +x_26 = lean_alloc_ctor(0, 11, 0); +lean_ctor_set(x_26, 0, x_14); +lean_ctor_set(x_26, 1, x_15); +lean_ctor_set(x_26, 2, x_16); +lean_ctor_set(x_26, 3, x_17); +lean_ctor_set(x_26, 4, x_18); +lean_ctor_set(x_26, 5, x_25); +lean_ctor_set(x_26, 6, x_20); +lean_ctor_set(x_26, 7, x_21); +lean_ctor_set(x_26, 8, x_22); +lean_ctor_set(x_26, 9, x_23); +lean_ctor_set(x_26, 10, x_24); +x_27 = l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_process___spec__3(x_2, x_3, x_4, x_5, x_6, x_26, x_8, x_9); lean_dec(x_8); -lean_dec(x_7); +lean_dec(x_26); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -return x_29; +return x_27; } } } -LEAN_EXPORT lean_object* l_Lean_resolveGlobalName___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__7(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_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_toParserDescr_process___spec__4(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_11; uint8_t x_12; -x_11 = lean_st_ref_get(x_9, x_10); -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_13 = lean_ctor_get(x_11, 0); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -lean_dec(x_13); -x_15 = lean_ctor_get(x_8, 6); -lean_inc(x_15); -x_16 = lean_ctor_get(x_8, 7); -lean_inc(x_16); -lean_dec(x_8); -x_17 = l_Lean_ResolveName_resolveGlobalName(x_14, x_15, x_16, x_1); -lean_ctor_set(x_11, 0, x_17); +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_checkLeftRec___spec__7___closed__2; +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_1); +lean_ctor_set(x_10, 1, x_9); +x_11 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_8); return x_11; } -else -{ -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_18 = lean_ctor_get(x_11, 0); -x_19 = lean_ctor_get(x_11, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_11); -x_20 = lean_ctor_get(x_18, 0); -lean_inc(x_20); -lean_dec(x_18); -x_21 = lean_ctor_get(x_8, 6); -lean_inc(x_21); -x_22 = lean_ctor_get(x_8, 7); -lean_inc(x_22); -lean_dec(x_8); -x_23 = l_Lean_ResolveName_resolveGlobalName(x_20, x_21, x_22, x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_19); -return x_24; -} -} -} -static lean_object* _init_l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("unknown constant '", 18); -return x_1; -} -} -static lean_object* _init_l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("'", 1); -return x_1; -} } -static lean_object* _init_l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__3; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8(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_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_toParserDescr_process___spec__5___rarg(lean_object* x_1) { _start: { -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_11 = lean_box(0); -x_12 = l_Lean_mkConst(x_1, x_11); -x_13 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_13, 0, x_12); -x_14 = l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__2; -x_15 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_15, 1, x_13); -x_16 = l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__4; -x_17 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_17, 0, x_15); -lean_ctor_set(x_17, 1, x_16); -x_18 = l_Lean_throwError___at_Lean_Elab_Term_checkLeftRec___spec__10(x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_18; +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__2___rarg___closed__2; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__6___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, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_toParserDescr_process___spec__5(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_13; lean_object* x_14; -x_13 = l_List_mapTRAux___at_Lean_resolveGlobalConstCore___spec__2(x_1, x_2); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_12); -return x_14; +lean_object* x_7; +x_7 = lean_alloc_closure((void*)(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_toParserDescr_process___spec__5___rarg), 1, 0); +return x_7; } } -LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__6(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_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_toParserDescr_process___spec__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_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -lean_inc(x_8); -lean_inc(x_1); -x_11 = l_Lean_resolveGlobalName___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -x_12 = lean_ctor_get(x_11, 0); +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; 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; 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_object* x_32; lean_object* x_33; +x_9 = lean_st_ref_get(x_7, x_8); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = lean_ctor_get(x_10, 0); lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); +lean_dec(x_10); +x_13 = lean_ctor_get(x_6, 3); lean_inc(x_13); -lean_dec(x_11); -x_14 = lean_box(0); -x_15 = l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(x_12, x_14); -x_16 = l_List_isEmpty___rarg(x_15); -if (x_16 == 0) +x_14 = lean_ctor_get(x_6, 4); +lean_inc(x_14); +x_15 = lean_ctor_get(x_6, 5); +lean_inc(x_15); +x_16 = lean_ctor_get(x_6, 6); +lean_inc(x_16); +x_17 = lean_ctor_get(x_6, 7); +lean_inc(x_17); +x_18 = lean_ctor_get(x_6, 10); +lean_inc(x_18); +lean_inc(x_12); +x_19 = lean_alloc_closure((void*)(l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_checkLeftRec___spec__1___lambda__1___boxed), 4, 1); +lean_closure_set(x_19, 0, x_12); +lean_inc(x_16); +x_20 = lean_alloc_closure((void*)(l_ReaderT_pure___at_Lean_Elab_liftMacroM___spec__2___rarg___boxed), 3, 1); +lean_closure_set(x_20, 0, x_16); +lean_inc(x_12); +x_21 = lean_alloc_closure((void*)(l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_checkLeftRec___spec__1___lambda__2___boxed), 4, 1); +lean_closure_set(x_21, 0, x_12); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_12); +x_22 = lean_alloc_closure((void*)(l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_checkLeftRec___spec__1___lambda__3___boxed), 6, 3); +lean_closure_set(x_22, 0, x_12); +lean_closure_set(x_22, 1, x_16); +lean_closure_set(x_22, 2, x_17); +lean_inc(x_12); +x_23 = lean_alloc_closure((void*)(l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_checkLeftRec___spec__1___lambda__4___boxed), 6, 3); +lean_closure_set(x_23, 0, x_12); +lean_closure_set(x_23, 1, x_16); +lean_closure_set(x_23, 2, x_17); +x_24 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_24, 0, x_19); +lean_ctor_set(x_24, 1, x_20); +lean_ctor_set(x_24, 2, x_21); +lean_ctor_set(x_24, 3, x_22); +lean_ctor_set(x_24, 4, x_23); +x_25 = lean_st_ref_get(x_7, x_11); +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +lean_dec(x_25); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +x_29 = lean_environment_main_module(x_12); +x_30 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_30, 0, x_24); +lean_ctor_set(x_30, 1, x_29); +lean_ctor_set(x_30, 2, x_18); +lean_ctor_set(x_30, 3, x_13); +lean_ctor_set(x_30, 4, x_14); +lean_ctor_set(x_30, 5, x_15); +x_31 = lean_box(0); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_28); +lean_ctor_set(x_32, 1, x_31); +x_33 = lean_apply_2(x_1, x_30, x_32); +if (lean_obj_tag(x_33) == 0) { -lean_object* x_17; lean_object* x_18; -lean_dec(x_1); -x_17 = lean_box(0); -x_18 = l_Lean_resolveGlobalConstCore___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__6___lambda__1(x_15, x_14, x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +x_37 = lean_st_ref_take(x_7, x_27); +x_38 = lean_ctor_get(x_37, 0); +lean_inc(x_38); +x_39 = lean_ctor_get(x_37, 1); +lean_inc(x_39); +lean_dec(x_37); +x_40 = !lean_is_exclusive(x_38); +if (x_40 == 0) +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; +x_41 = lean_ctor_get(x_38, 1); +lean_dec(x_41); +lean_ctor_set(x_38, 1, x_36); +x_42 = lean_st_ref_set(x_7, x_38, x_39); +x_43 = lean_ctor_get(x_42, 1); +lean_inc(x_43); +lean_dec(x_42); +x_44 = lean_ctor_get(x_35, 1); +lean_inc(x_44); +lean_dec(x_35); +x_45 = l_List_reverse___rarg(x_44); +x_46 = l_List_forM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__3(x_45, x_2, x_3, x_4, x_5, x_6, x_7, x_43); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_18; +x_47 = !lean_is_exclusive(x_46); +if (x_47 == 0) +{ +lean_object* x_48; +x_48 = lean_ctor_get(x_46, 0); +lean_dec(x_48); +lean_ctor_set(x_46, 0, x_34); +return x_46; } else { -lean_object* x_19; uint8_t x_20; -lean_dec(x_15); -x_19 = l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_49; lean_object* x_50; +x_49 = lean_ctor_get(x_46, 1); +lean_inc(x_49); +lean_dec(x_46); +x_50 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_50, 0, x_34); +lean_ctor_set(x_50, 1, x_49); +return x_50; +} +} +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_51 = lean_ctor_get(x_38, 0); +x_52 = lean_ctor_get(x_38, 2); +x_53 = lean_ctor_get(x_38, 3); +x_54 = lean_ctor_get(x_38, 4); +x_55 = lean_ctor_get(x_38, 5); +lean_inc(x_55); +lean_inc(x_54); +lean_inc(x_53); +lean_inc(x_52); +lean_inc(x_51); +lean_dec(x_38); +x_56 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_56, 0, x_51); +lean_ctor_set(x_56, 1, x_36); +lean_ctor_set(x_56, 2, x_52); +lean_ctor_set(x_56, 3, x_53); +lean_ctor_set(x_56, 4, x_54); +lean_ctor_set(x_56, 5, x_55); +x_57 = lean_st_ref_set(x_7, x_56, x_39); +x_58 = lean_ctor_get(x_57, 1); +lean_inc(x_58); +lean_dec(x_57); +x_59 = lean_ctor_get(x_35, 1); +lean_inc(x_59); +lean_dec(x_35); +x_60 = l_List_reverse___rarg(x_59); +x_61 = l_List_forM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__3(x_60, x_2, x_3, x_4, x_5, x_6, x_7, x_58); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_20 = !lean_is_exclusive(x_19); -if (x_20 == 0) -{ -return x_19; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_19, 0); -x_22 = lean_ctor_get(x_19, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_19); -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); -return x_23; +x_62 = lean_ctor_get(x_61, 1); +lean_inc(x_62); +if (lean_is_exclusive(x_61)) { + lean_ctor_release(x_61, 0); + lean_ctor_release(x_61, 1); + x_63 = x_61; +} else { + lean_dec_ref(x_61); + x_63 = lean_box(0); } +if (lean_is_scalar(x_63)) { + x_64 = lean_alloc_ctor(0, 2, 0); +} else { + x_64 = x_63; } +lean_ctor_set(x_64, 0, x_34); +lean_ctor_set(x_64, 1, x_62); +return x_64; } } -static lean_object* _init_l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("expected identifier", 19); -return x_1; -} -} -static lean_object* _init_l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3___closed__2() { -_start: +lean_object* x_65; +x_65 = lean_ctor_get(x_33, 0); +lean_inc(x_65); +lean_dec(x_33); +if (lean_obj_tag(x_65) == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3___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_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3___closed__3() { -_start: +lean_object* x_66; lean_object* x_67; lean_object* x_68; uint8_t x_69; +x_66 = lean_ctor_get(x_65, 0); +lean_inc(x_66); +x_67 = lean_ctor_get(x_65, 1); +lean_inc(x_67); +lean_dec(x_65); +x_68 = l_Lean_maxRecDepthErrorMessage; +x_69 = lean_string_dec_eq(x_67, x_68); +if (x_69 == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3___closed__2; -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} +lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_70 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_70, 0, x_67); +x_71 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_71, 0, x_70); +x_72 = l_Lean_throwErrorAt___at_Lean_Elab_Term_toParserDescr_process___spec__2(x_66, x_71, x_2, x_3, x_4, x_5, x_6, x_7, x_27); +return x_72; } -LEAN_EXPORT lean_object* l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3(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_object* x_9, lean_object* x_10) { -_start: -{ -if (lean_obj_tag(x_1) == 3) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_11 = lean_ctor_get(x_1, 2); -lean_inc(x_11); -x_12 = lean_ctor_get(x_1, 3); -lean_inc(x_12); -x_13 = l_List_filterMap___at_Lean_resolveGlobalConst___spec__1(x_12); -x_14 = l_List_isEmpty___rarg(x_13); -if (x_14 == 0) +else { -lean_object* x_15; -lean_dec(x_11); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_73; +lean_dec(x_67); +x_73 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_toParserDescr_process___spec__4(x_66, x_2, x_3, x_4, x_5, x_6, x_7, x_27); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_13); -lean_ctor_set(x_15, 1, x_10); -return x_15; +return x_73; } -else -{ -uint8_t x_16; -lean_dec(x_13); -x_16 = !lean_is_exclusive(x_8); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_8, 5); -x_18 = l_Lean_replaceRef(x_1, x_17); -lean_dec(x_17); -lean_dec(x_1); -lean_ctor_set(x_8, 5, x_18); -x_19 = l_Lean_resolveGlobalConstCore___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__6(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_19; } else { -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; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_20 = lean_ctor_get(x_8, 0); -x_21 = lean_ctor_get(x_8, 1); -x_22 = lean_ctor_get(x_8, 2); -x_23 = lean_ctor_get(x_8, 3); -x_24 = lean_ctor_get(x_8, 4); -x_25 = lean_ctor_get(x_8, 5); -x_26 = lean_ctor_get(x_8, 6); -x_27 = lean_ctor_get(x_8, 7); -x_28 = lean_ctor_get(x_8, 8); -x_29 = lean_ctor_get(x_8, 9); -x_30 = lean_ctor_get(x_8, 10); -lean_inc(x_30); -lean_inc(x_29); -lean_inc(x_28); -lean_inc(x_27); -lean_inc(x_26); -lean_inc(x_25); -lean_inc(x_24); -lean_inc(x_23); -lean_inc(x_22); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_8); -x_31 = l_Lean_replaceRef(x_1, x_25); -lean_dec(x_25); -lean_dec(x_1); -x_32 = lean_alloc_ctor(0, 11, 0); -lean_ctor_set(x_32, 0, x_20); -lean_ctor_set(x_32, 1, x_21); -lean_ctor_set(x_32, 2, x_22); -lean_ctor_set(x_32, 3, x_23); -lean_ctor_set(x_32, 4, x_24); -lean_ctor_set(x_32, 5, x_31); -lean_ctor_set(x_32, 6, x_26); -lean_ctor_set(x_32, 7, x_27); -lean_ctor_set(x_32, 8, x_28); -lean_ctor_set(x_32, 9, x_29); -lean_ctor_set(x_32, 10, x_30); -x_33 = l_Lean_resolveGlobalConstCore___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__6(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_32, x_9, x_10); -return x_33; -} -} +lean_object* x_74; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_74 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_toParserDescr_process___spec__5___rarg(x_27); +return x_74; } -else -{ -lean_object* x_34; lean_object* x_35; -x_34 = l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3___closed__3; -x_35 = l_Lean_throwErrorAt___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__4(x_1, x_34, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_35; } } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__11(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_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_process___spec__7(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_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; lean_object* x_12; uint8_t x_13; @@ -5683,1610 +6163,2262 @@ return x_19; } } } -LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__10(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_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_toParserDescr_process___spec__6(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_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_11; uint8_t x_12; -x_11 = lean_st_ref_get(x_9, x_10); -x_12 = !lean_is_exclusive(x_11); +uint8_t x_12; +x_12 = !lean_is_exclusive(x_9); if (x_12 == 0) { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_13 = lean_ctor_get(x_11, 0); -x_14 = lean_ctor_get(x_11, 1); -x_15 = lean_ctor_get(x_13, 0); -lean_inc(x_15); +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_9, 5); +x_14 = l_Lean_replaceRef(x_1, x_13); lean_dec(x_13); -lean_inc(x_1); -x_16 = lean_environment_find(x_15, x_1); -if (lean_obj_tag(x_16) == 0) -{ -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; -lean_free_object(x_11); -x_17 = lean_box(0); -x_18 = l_Lean_mkConst(x_1, x_17); -x_19 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_19, 0, x_18); -x_20 = l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__2; -x_21 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_19); -x_22 = l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__4; -x_23 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); -x_24 = l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__11(x_23, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); -return x_24; +lean_dec(x_1); +lean_ctor_set(x_9, 5, x_14); +x_15 = l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_process___spec__7(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_15; } else { -lean_object* x_25; -lean_dec(x_1); -x_25 = lean_ctor_get(x_16, 0); +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; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_16 = lean_ctor_get(x_9, 0); +x_17 = lean_ctor_get(x_9, 1); +x_18 = lean_ctor_get(x_9, 2); +x_19 = lean_ctor_get(x_9, 3); +x_20 = lean_ctor_get(x_9, 4); +x_21 = lean_ctor_get(x_9, 5); +x_22 = lean_ctor_get(x_9, 6); +x_23 = lean_ctor_get(x_9, 7); +x_24 = lean_ctor_get(x_9, 8); +x_25 = lean_ctor_get(x_9, 9); +x_26 = lean_ctor_get(x_9, 10); +lean_inc(x_26); lean_inc(x_25); -lean_dec(x_16); -lean_ctor_set(x_11, 0, x_25); -return x_11; +lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_9); +x_27 = l_Lean_replaceRef(x_1, x_21); +lean_dec(x_21); +lean_dec(x_1); +x_28 = lean_alloc_ctor(0, 11, 0); +lean_ctor_set(x_28, 0, x_16); +lean_ctor_set(x_28, 1, x_17); +lean_ctor_set(x_28, 2, x_18); +lean_ctor_set(x_28, 3, x_19); +lean_ctor_set(x_28, 4, x_20); +lean_ctor_set(x_28, 5, x_27); +lean_ctor_set(x_28, 6, x_22); +lean_ctor_set(x_28, 7, x_23); +lean_ctor_set(x_28, 8, x_24); +lean_ctor_set(x_28, 9, x_25); +lean_ctor_set(x_28, 10, x_26); +x_29 = l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_process___spec__7(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_28, x_10, x_11); +lean_dec(x_10); +lean_dec(x_28); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_29; } } -else +} +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_process___closed__1() { +_start: { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_26 = lean_ctor_get(x_11, 0); -x_27 = lean_ctor_get(x_11, 1); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_11); -x_28 = lean_ctor_get(x_26, 0); -lean_inc(x_28); -lean_dec(x_26); -lean_inc(x_1); -x_29 = lean_environment_find(x_28, x_1); -if (lean_obj_tag(x_29) == 0) +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("choice", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_process___closed__2() { +_start: { -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; -x_30 = lean_box(0); -x_31 = l_Lean_mkConst(x_1, x_30); -x_32 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_32, 0, x_31); -x_33 = l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__2; -x_34 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_32); -x_35 = l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__4; -x_36 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -x_37 = l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__11(x_36, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_27); -return x_37; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_toParserDescr_process___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; } -else +} +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_process___closed__3() { +_start: { -lean_object* x_38; lean_object* x_39; -lean_dec(x_1); -x_38 = lean_ctor_get(x_29, 0); -lean_inc(x_38); -lean_dec(x_29); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_27); -return x_39; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("paren", 5); +return x_1; } } +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_process___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Term_checkLeftRec___closed__2; +x_2 = l_Lean_Elab_Term_toParserDescr_process___closed__3; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__9(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_object* x_9, lean_object* x_10) { +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_process___closed__5() { _start: { -lean_object* x_11; -lean_inc(x_1); -x_11 = l_Lean_getConstInfo___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__10(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_11) == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Term_checkLeftRec___closed__2; +x_2 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__6; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_process___closed__6() { +_start: { -uint8_t x_12; -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Term_checkLeftRec___closed__2; +x_2 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__14; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_process___closed__7() { +_start: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_13 = lean_ctor_get(x_11, 0); -x_14 = l_Lean_ConstantInfo_levelParams(x_13); -lean_dec(x_13); -x_15 = lean_box(0); -x_16 = l_List_mapTRAux___at_Lean_mkConstWithLevelParams___spec__1(x_14, x_15); -x_17 = l_Lean_mkConst(x_1, x_16); -lean_ctor_set(x_11, 0, x_17); -return x_11; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("sepBy", 5); +return x_1; } -else +} +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_process___closed__8() { +_start: { -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_18 = lean_ctor_get(x_11, 0); -x_19 = lean_ctor_get(x_11, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_11); -x_20 = l_Lean_ConstantInfo_levelParams(x_18); -lean_dec(x_18); -x_21 = lean_box(0); -x_22 = l_List_mapTRAux___at_Lean_mkConstWithLevelParams___spec__1(x_20, x_21); -x_23 = l_Lean_mkConst(x_1, x_22); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_19); -return x_24; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Term_checkLeftRec___closed__2; +x_2 = l_Lean_Elab_Term_toParserDescr_process___closed__7; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; } } -else +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_process___closed__9() { +_start: { -uint8_t x_25; -lean_dec(x_1); -x_25 = !lean_is_exclusive(x_11); -if (x_25 == 0) +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("sepBy1", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_process___closed__10() { +_start: { -return x_11; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Term_checkLeftRec___closed__2; +x_2 = l_Lean_Elab_Term_toParserDescr_process___closed__9; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; } -else +} +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_process___closed__11() { +_start: { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_11, 0); -x_27 = lean_ctor_get(x_11, 1); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_11); -x_28 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -return x_28; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("atom", 4); +return x_1; } } +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_process___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Term_checkLeftRec___closed__2; +x_2 = l_Lean_Elab_Term_toParserDescr_process___closed__11; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoTree___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__13(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_object* x_9, lean_object* x_10) { +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_process___closed__13() { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_11 = lean_st_ref_get(x_9, x_10); -x_12 = lean_ctor_get(x_11, 1); -lean_inc(x_12); -lean_dec(x_11); -x_13 = lean_st_ref_get(x_5, x_12); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_14, 4); -lean_inc(x_15); -lean_dec(x_14); -x_16 = lean_ctor_get_uint8(x_15, sizeof(void*)*2); -lean_dec(x_15); -if (x_16 == 0) +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("nonReserved", 11); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_process___closed__14() { +_start: { -uint8_t x_17; -lean_dec(x_1); -x_17 = !lean_is_exclusive(x_13); -if (x_17 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Term_checkLeftRec___closed__2; +x_2 = l_Lean_Elab_Term_toParserDescr_process___closed__13; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_process___closed__15() { +_start: { -lean_object* x_18; lean_object* x_19; -x_18 = lean_ctor_get(x_13, 0); -lean_dec(x_18); -x_19 = lean_box(0); -lean_ctor_set(x_13, 0, x_19); -return x_13; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("unexpected syntax kind of category `syntax`: ", 45); +return x_1; } -else +} +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_process___closed__16() { +_start: { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_13, 1); -lean_inc(x_20); -lean_dec(x_13); -x_21 = lean_box(0); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_20); -return x_22; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Term_toParserDescr_process___closed__15; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -else +LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_process(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_object* x_9, lean_object* x_10) { +_start: { -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; uint8_t x_30; -x_23 = lean_ctor_get(x_13, 1); -lean_inc(x_23); -lean_dec(x_13); -x_24 = lean_st_ref_get(x_9, x_23); -x_25 = lean_ctor_get(x_24, 1); -lean_inc(x_25); -lean_dec(x_24); -x_26 = lean_st_ref_take(x_5, x_25); -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_27, 4); -lean_inc(x_28); -x_29 = lean_ctor_get(x_26, 1); -lean_inc(x_29); -lean_dec(x_26); -x_30 = !lean_is_exclusive(x_27); +lean_object* x_11; lean_object* x_12; uint8_t x_13; uint8_t x_14; +lean_inc(x_1); +x_11 = l_Lean_Syntax_getKind(x_1); +x_12 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; +x_13 = lean_name_eq(x_11, x_12); +x_14 = !lean_is_exclusive(x_8); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_8, 5); +x_16 = l_Lean_replaceRef(x_1, x_15); +lean_dec(x_15); +lean_ctor_set(x_8, 5, x_16); +if (x_13 == 0) +{ +lean_object* x_17; uint8_t x_18; +x_17 = l_Lean_Elab_Term_toParserDescr_process___closed__2; +x_18 = lean_name_eq(x_11, x_17); +if (x_18 == 0) +{ +lean_object* x_19; uint8_t x_20; +x_19 = l_Lean_Elab_Term_toParserDescr_process___closed__4; +x_20 = lean_name_eq(x_11, x_19); +if (x_20 == 0) +{ +lean_object* x_21; uint8_t x_22; +x_21 = l_Lean_Elab_Term_checkLeftRec___closed__4; +x_22 = lean_name_eq(x_11, x_21); +if (x_22 == 0) +{ +lean_object* x_23; uint8_t x_24; +x_23 = l_Lean_Elab_Term_toParserDescr_process___closed__5; +x_24 = lean_name_eq(x_11, x_23); +if (x_24 == 0) +{ +lean_object* x_25; uint8_t x_26; +x_25 = l_Lean_Elab_Term_toParserDescr_process___closed__6; +x_26 = lean_name_eq(x_11, x_25); +if (x_26 == 0) +{ +lean_object* x_27; uint8_t x_28; +x_27 = l_Lean_Elab_Term_toParserDescr_process___closed__8; +x_28 = lean_name_eq(x_11, x_27); +if (x_28 == 0) +{ +lean_object* x_29; uint8_t x_30; +x_29 = l_Lean_Elab_Term_toParserDescr_process___closed__10; +x_30 = lean_name_eq(x_11, x_29); if (x_30 == 0) { lean_object* x_31; uint8_t x_32; -x_31 = lean_ctor_get(x_27, 4); -lean_dec(x_31); -x_32 = !lean_is_exclusive(x_28); +x_31 = l_Lean_Elab_Term_toParserDescr_process___closed__12; +x_32 = lean_name_eq(x_11, x_31); if (x_32 == 0) { -lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; -x_33 = lean_ctor_get(x_28, 1); -x_34 = l_Std_PersistentArray_push___rarg(x_33, x_1); -lean_ctor_set(x_28, 1, x_34); -x_35 = lean_st_ref_set(x_5, x_27, x_29); -x_36 = !lean_is_exclusive(x_35); -if (x_36 == 0) +lean_object* x_33; uint8_t x_34; +x_33 = l_Lean_Elab_Term_toParserDescr_process___closed__14; +x_34 = lean_name_eq(x_11, x_33); +if (x_34 == 0) { -lean_object* x_37; lean_object* x_38; -x_37 = lean_ctor_get(x_35, 0); -lean_dec(x_37); -x_38 = lean_box(0); -lean_ctor_set(x_35, 0, x_38); -return x_35; -} -else +lean_object* x_35; lean_object* x_36; +lean_inc(x_1); +x_35 = lean_alloc_closure((void*)(l_Lean_Macro_expandMacro_x3f), 3, 1); +lean_closure_set(x_35, 0, x_1); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_36 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_toParserDescr_process___spec__1(x_35, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_36) == 0) { -lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_39 = lean_ctor_get(x_35, 1); -lean_inc(x_39); -lean_dec(x_35); -x_40 = lean_box(0); -x_41 = lean_alloc_ctor(0, 2, 0); +lean_object* x_37; +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +if (lean_obj_tag(x_37) == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +lean_dec(x_36); +x_39 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_39, 0, x_11); +x_40 = l_Lean_Elab_Term_toParserDescr_process___closed__16; +x_41 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_41, 0, x_40); lean_ctor_set(x_41, 1, x_39); -return x_41; -} +x_42 = l_Lean_addTrace___at_Lean_Elab_Term_checkLeftRec___spec__3___closed__8; +x_43 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +x_44 = l_Lean_throwErrorAt___at_Lean_Elab_Term_toParserDescr_process___spec__6(x_1, x_43, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_38); +return x_44; } else { -uint8_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_42 = lean_ctor_get_uint8(x_28, sizeof(void*)*2); -x_43 = lean_ctor_get(x_28, 0); -x_44 = lean_ctor_get(x_28, 1); -lean_inc(x_44); -lean_inc(x_43); -lean_dec(x_28); -x_45 = l_Std_PersistentArray_push___rarg(x_44, x_1); -x_46 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_46, 0, x_43); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set_uint8(x_46, sizeof(void*)*2, x_42); -lean_ctor_set(x_27, 4, x_46); -x_47 = lean_st_ref_set(x_5, x_27, x_29); -x_48 = lean_ctor_get(x_47, 1); -lean_inc(x_48); -if (lean_is_exclusive(x_47)) { - lean_ctor_release(x_47, 0); - lean_ctor_release(x_47, 1); - x_49 = x_47; -} else { - lean_dec_ref(x_47); - x_49 = lean_box(0); +lean_object* x_45; lean_object* x_46; +lean_dec(x_11); +lean_dec(x_1); +x_45 = lean_ctor_get(x_36, 1); +lean_inc(x_45); +lean_dec(x_36); +x_46 = lean_ctor_get(x_37, 0); +lean_inc(x_46); +lean_dec(x_37); +x_1 = x_46; +x_10 = x_45; +goto _start; } -x_50 = lean_box(0); -if (lean_is_scalar(x_49)) { - x_51 = lean_alloc_ctor(0, 2, 0); -} else { - x_51 = x_49; } -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_48); +else +{ +uint8_t x_48; +lean_dec(x_8); +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_48 = !lean_is_exclusive(x_36); +if (x_48 == 0) +{ +return x_36; +} +else +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_36, 0); +x_50 = lean_ctor_get(x_36, 1); +lean_inc(x_50); +lean_inc(x_49); +lean_dec(x_36); +x_51 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); return x_51; } } +} else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; lean_object* x_67; -x_52 = lean_ctor_get(x_27, 0); -x_53 = lean_ctor_get(x_27, 1); -x_54 = lean_ctor_get(x_27, 2); -x_55 = lean_ctor_get(x_27, 3); -lean_inc(x_55); -lean_inc(x_54); -lean_inc(x_53); -lean_inc(x_52); -lean_dec(x_27); -x_56 = lean_ctor_get_uint8(x_28, sizeof(void*)*2); -x_57 = lean_ctor_get(x_28, 0); -lean_inc(x_57); -x_58 = lean_ctor_get(x_28, 1); -lean_inc(x_58); -if (lean_is_exclusive(x_28)) { - lean_ctor_release(x_28, 0); - lean_ctor_release(x_28, 1); - x_59 = x_28; -} else { - lean_dec_ref(x_28); - x_59 = lean_box(0); +lean_object* x_52; +lean_dec(x_11); +x_52 = l_Lean_Elab_Term_toParserDescr_processNonReserved(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_52; } -x_60 = l_Std_PersistentArray_push___rarg(x_58, x_1); -if (lean_is_scalar(x_59)) { - x_61 = lean_alloc_ctor(0, 2, 1); -} else { - x_61 = x_59; } -lean_ctor_set(x_61, 0, x_57); -lean_ctor_set(x_61, 1, x_60); -lean_ctor_set_uint8(x_61, sizeof(void*)*2, x_56); -x_62 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_62, 0, x_52); -lean_ctor_set(x_62, 1, x_53); -lean_ctor_set(x_62, 2, x_54); -lean_ctor_set(x_62, 3, x_55); -lean_ctor_set(x_62, 4, x_61); -x_63 = lean_st_ref_set(x_5, x_62, x_29); -x_64 = lean_ctor_get(x_63, 1); -lean_inc(x_64); -if (lean_is_exclusive(x_63)) { - lean_ctor_release(x_63, 0); - lean_ctor_release(x_63, 1); - x_65 = x_63; -} else { - lean_dec_ref(x_63); - x_65 = lean_box(0); +else +{ +lean_object* x_53; +lean_dec(x_11); +x_53 = l_Lean_Elab_Term_toParserDescr_processAtom(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_53; } -x_66 = lean_box(0); -if (lean_is_scalar(x_65)) { - x_67 = lean_alloc_ctor(0, 2, 0); -} else { - x_67 = x_65; } -lean_ctor_set(x_67, 0, x_66); -lean_ctor_set(x_67, 1, x_64); -return x_67; +else +{ +lean_object* x_54; +lean_dec(x_11); +x_54 = l_Lean_Elab_Term_toParserDescr_processSepBy1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_54; } } +else +{ +lean_object* x_55; +lean_dec(x_11); +x_55 = l_Lean_Elab_Term_toParserDescr_processSepBy(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_55; } } -static lean_object* _init_l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___closed__1() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(32u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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_dec(x_11); +x_56 = lean_unsigned_to_nat(0u); +x_57 = l_Lean_Syntax_getArg(x_1, x_56); +x_58 = lean_unsigned_to_nat(2u); +x_59 = l_Lean_Syntax_getArg(x_1, x_58); +x_60 = lean_unsigned_to_nat(4u); +x_61 = l_Lean_Syntax_getArg(x_1, x_60); +lean_dec(x_1); +x_62 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; +x_63 = lean_array_push(x_62, x_59); +x_64 = lean_array_push(x_63, x_61); +x_65 = l_Lean_Elab_Term_toParserDescr_processAlias(x_57, x_64, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_65; } } -static lean_object* _init_l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___closed__1; -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_66; 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_dec(x_11); +x_66 = lean_unsigned_to_nat(0u); +x_67 = l_Lean_Syntax_getArg(x_1, x_66); +x_68 = lean_unsigned_to_nat(2u); +x_69 = l_Lean_Syntax_getArg(x_1, x_68); +lean_dec(x_1); +x_70 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; +x_71 = lean_array_push(x_70, x_69); +x_72 = l_Lean_Elab_Term_toParserDescr_processAlias(x_67, x_71, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_72; } } -static lean_object* _init_l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___closed__3() { -_start: +else { -size_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = 5; -x_2 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___closed__2; -x_3 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___closed__1; -x_4 = lean_unsigned_to_nat(0u); -x_5 = lean_alloc_ctor(0, 4, sizeof(size_t)*1); -lean_ctor_set(x_5, 0, x_2); -lean_ctor_set(x_5, 1, x_3); -lean_ctor_set(x_5, 2, x_4); -lean_ctor_set(x_5, 3, x_4); -lean_ctor_set_usize(x_5, 4, x_1); -return x_5; +lean_object* x_73; +lean_dec(x_11); +x_73 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_73; } } -LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12(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_object* x_9, lean_object* x_10) { -_start: +else { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_11 = lean_st_ref_get(x_9, x_10); -x_12 = lean_ctor_get(x_11, 1); -lean_inc(x_12); +lean_object* x_74; lean_object* x_75; lean_dec(x_11); -x_13 = lean_st_ref_get(x_5, x_12); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_14, 4); -lean_inc(x_15); -lean_dec(x_14); -x_16 = lean_ctor_get_uint8(x_15, sizeof(void*)*2); -lean_dec(x_15); -if (x_16 == 0) -{ -uint8_t x_17; +x_74 = lean_unsigned_to_nat(1u); +x_75 = l_Lean_Syntax_getArg(x_1, x_74); lean_dec(x_1); -x_17 = !lean_is_exclusive(x_13); -if (x_17 == 0) -{ -lean_object* x_18; lean_object* x_19; -x_18 = lean_ctor_get(x_13, 0); -lean_dec(x_18); -x_19 = lean_box(0); -lean_ctor_set(x_13, 0, x_19); -return x_13; +x_1 = x_75; +goto _start; +} } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_13, 1); -lean_inc(x_20); -lean_dec(x_13); -x_21 = lean_box(0); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_20); -return x_22; +lean_object* x_77; lean_object* x_78; +lean_dec(x_11); +x_77 = lean_unsigned_to_nat(0u); +x_78 = l_Lean_Syntax_getArg(x_1, x_77); +lean_dec(x_1); +x_1 = x_78; +goto _start; } } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_23 = lean_ctor_get(x_13, 1); -lean_inc(x_23); -lean_dec(x_13); -x_24 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___closed__3; -x_25 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_25, 0, x_1); -lean_ctor_set(x_25, 1, x_24); -x_26 = l_Lean_Elab_pushInfoTree___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__13(x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_23); -return x_26; -} +lean_object* x_80; +lean_dec(x_11); +x_80 = l_Lean_Elab_Term_toParserDescr_processSeq(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_80; } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__14(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: +else { -if (lean_obj_tag(x_3) == 0) +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_81 = lean_ctor_get(x_8, 0); +x_82 = lean_ctor_get(x_8, 1); +x_83 = lean_ctor_get(x_8, 2); +x_84 = lean_ctor_get(x_8, 3); +x_85 = lean_ctor_get(x_8, 4); +x_86 = lean_ctor_get(x_8, 5); +x_87 = lean_ctor_get(x_8, 6); +x_88 = lean_ctor_get(x_8, 7); +x_89 = lean_ctor_get(x_8, 8); +x_90 = lean_ctor_get(x_8, 9); +x_91 = lean_ctor_get(x_8, 10); +lean_inc(x_91); +lean_inc(x_90); +lean_inc(x_89); +lean_inc(x_88); +lean_inc(x_87); +lean_inc(x_86); +lean_inc(x_85); +lean_inc(x_84); +lean_inc(x_83); +lean_inc(x_82); +lean_inc(x_81); +lean_dec(x_8); +x_92 = l_Lean_replaceRef(x_1, x_86); +lean_dec(x_86); +x_93 = lean_alloc_ctor(0, 11, 0); +lean_ctor_set(x_93, 0, x_81); +lean_ctor_set(x_93, 1, x_82); +lean_ctor_set(x_93, 2, x_83); +lean_ctor_set(x_93, 3, x_84); +lean_ctor_set(x_93, 4, x_85); +lean_ctor_set(x_93, 5, x_92); +lean_ctor_set(x_93, 6, x_87); +lean_ctor_set(x_93, 7, x_88); +lean_ctor_set(x_93, 8, x_89); +lean_ctor_set(x_93, 9, x_90); +lean_ctor_set(x_93, 10, x_91); +if (x_13 == 0) { -lean_object* x_14; +lean_object* x_94; uint8_t x_95; +x_94 = l_Lean_Elab_Term_toParserDescr_process___closed__2; +x_95 = lean_name_eq(x_11, x_94); +if (x_95 == 0) +{ +lean_object* x_96; uint8_t x_97; +x_96 = l_Lean_Elab_Term_toParserDescr_process___closed__4; +x_97 = lean_name_eq(x_11, x_96); +if (x_97 == 0) +{ +lean_object* x_98; uint8_t x_99; +x_98 = l_Lean_Elab_Term_checkLeftRec___closed__4; +x_99 = lean_name_eq(x_11, x_98); +if (x_99 == 0) +{ +lean_object* x_100; uint8_t x_101; +x_100 = l_Lean_Elab_Term_toParserDescr_process___closed__5; +x_101 = lean_name_eq(x_11, x_100); +if (x_101 == 0) +{ +lean_object* x_102; uint8_t x_103; +x_102 = l_Lean_Elab_Term_toParserDescr_process___closed__6; +x_103 = lean_name_eq(x_11, x_102); +if (x_103 == 0) +{ +lean_object* x_104; uint8_t x_105; +x_104 = l_Lean_Elab_Term_toParserDescr_process___closed__8; +x_105 = lean_name_eq(x_11, x_104); +if (x_105 == 0) +{ +lean_object* x_106; uint8_t x_107; +x_106 = l_Lean_Elab_Term_toParserDescr_process___closed__10; +x_107 = lean_name_eq(x_11, x_106); +if (x_107 == 0) +{ +lean_object* x_108; uint8_t x_109; +x_108 = l_Lean_Elab_Term_toParserDescr_process___closed__12; +x_109 = lean_name_eq(x_11, x_108); +if (x_109 == 0) +{ +lean_object* x_110; uint8_t x_111; +x_110 = l_Lean_Elab_Term_toParserDescr_process___closed__14; +x_111 = lean_name_eq(x_11, x_110); +if (x_111 == 0) +{ +lean_object* x_112; lean_object* x_113; +lean_inc(x_1); +x_112 = lean_alloc_closure((void*)(l_Lean_Macro_expandMacro_x3f), 3, 1); +lean_closure_set(x_112, 0, x_1); +lean_inc(x_9); +lean_inc(x_93); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_113 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_toParserDescr_process___spec__1(x_112, x_4, x_5, x_6, x_7, x_93, x_9, x_10); +if (lean_obj_tag(x_113) == 0) +{ +lean_object* x_114; +x_114 = lean_ctor_get(x_113, 0); +lean_inc(x_114); +if (lean_obj_tag(x_114) == 0) +{ +lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; +x_115 = lean_ctor_get(x_113, 1); +lean_inc(x_115); +lean_dec(x_113); +x_116 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_116, 0, x_11); +x_117 = l_Lean_Elab_Term_toParserDescr_process___closed__16; +x_118 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_118, 0, x_117); +lean_ctor_set(x_118, 1, x_116); +x_119 = l_Lean_addTrace___at_Lean_Elab_Term_checkLeftRec___spec__3___closed__8; +x_120 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_120, 0, x_118); +lean_ctor_set(x_120, 1, x_119); +x_121 = l_Lean_throwErrorAt___at_Lean_Elab_Term_toParserDescr_process___spec__6(x_1, x_120, x_2, x_3, x_4, x_5, x_6, x_7, x_93, x_9, x_115); +return x_121; +} +else +{ +lean_object* x_122; lean_object* x_123; +lean_dec(x_11); +lean_dec(x_1); +x_122 = lean_ctor_get(x_113, 1); +lean_inc(x_122); +lean_dec(x_113); +x_123 = lean_ctor_get(x_114, 0); +lean_inc(x_123); +lean_dec(x_114); +x_1 = x_123; +x_8 = x_93; +x_10 = x_122; +goto _start; +} +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; +lean_dec(x_93); +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_4); -lean_ctor_set(x_14, 1, x_13); -return x_14; +x_125 = lean_ctor_get(x_113, 0); +lean_inc(x_125); +x_126 = lean_ctor_get(x_113, 1); +lean_inc(x_126); +if (lean_is_exclusive(x_113)) { + lean_ctor_release(x_113, 0); + lean_ctor_release(x_113, 1); + x_127 = x_113; +} else { + lean_dec_ref(x_113); + x_127 = lean_box(0); +} +if (lean_is_scalar(x_127)) { + x_128 = lean_alloc_ctor(1, 2, 0); +} else { + x_128 = x_127; +} +lean_ctor_set(x_128, 0, x_125); +lean_ctor_set(x_128, 1, x_126); +return x_128; +} } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; +lean_object* x_129; +lean_dec(x_11); +x_129 = l_Lean_Elab_Term_toParserDescr_processNonReserved(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_93, x_9, x_10); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); -x_15 = lean_ctor_get(x_3, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_3, 1); -lean_inc(x_16); lean_dec(x_3); -x_17 = l_Lean_mkConstWithLevelParams___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__9(x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_17) == 0) +lean_dec(x_2); +return x_129; +} +} +else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); -lean_inc(x_19); -lean_dec(x_17); -x_20 = lean_box(0); -lean_inc(x_1); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_1); -x_22 = l_Lean_LocalContext_empty; -x_23 = 0; -lean_inc(x_2); -x_24 = lean_alloc_ctor(0, 4, 1); -lean_ctor_set(x_24, 0, x_21); -lean_ctor_set(x_24, 1, x_22); -lean_ctor_set(x_24, 2, x_2); -lean_ctor_set(x_24, 3, x_18); -lean_ctor_set_uint8(x_24, sizeof(void*)*4, x_23); -x_25 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_25, 0, x_24); -x_26 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12(x_25, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_19); -x_27 = lean_ctor_get(x_26, 1); -lean_inc(x_27); -lean_dec(x_26); -x_28 = lean_box(0); -x_3 = x_16; -x_4 = x_28; -x_13 = x_27; -goto _start; +lean_object* x_130; +lean_dec(x_11); +x_130 = l_Lean_Elab_Term_toParserDescr_processAtom(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_93, x_9, x_10); +return x_130; } -else -{ -uint8_t x_30; -lean_dec(x_16); -lean_dec(x_2); -lean_dec(x_1); -x_30 = !lean_is_exclusive(x_17); -if (x_30 == 0) -{ -return x_17; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_17, 0); -x_32 = lean_ctor_get(x_17, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_17); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -return x_33; -} -} -} +lean_object* x_131; +lean_dec(x_11); +x_131 = l_Lean_Elab_Term_toParserDescr_processSepBy1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_93, x_9, x_10); +return x_131; } } -LEAN_EXPORT lean_object* l_Lean_Elab_resolveGlobalConstWithInfos___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__2___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, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: +else { -lean_object* x_12; -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_1); -lean_ctor_set(x_12, 1, x_11); -return x_12; +lean_object* x_132; +lean_dec(x_11); +x_132 = l_Lean_Elab_Term_toParserDescr_processSepBy(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_93, x_9, x_10); +return x_132; } } -LEAN_EXPORT lean_object* l_Lean_Elab_resolveGlobalConstWithInfos___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___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_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_1); -x_12 = l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_12) == 0) -{ -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; uint8_t x_20; -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec(x_12); -x_15 = lean_st_ref_get(x_10, x_14); -x_16 = lean_ctor_get(x_15, 1); -lean_inc(x_16); -lean_dec(x_15); -x_17 = lean_st_ref_get(x_6, x_16); -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_18, 4); -lean_inc(x_19); -lean_dec(x_18); -x_20 = lean_ctor_get_uint8(x_19, sizeof(void*)*2); -lean_dec(x_19); -if (x_20 == 0) +else { -uint8_t x_21; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); +lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; +lean_dec(x_11); +x_133 = lean_unsigned_to_nat(0u); +x_134 = l_Lean_Syntax_getArg(x_1, x_133); +x_135 = lean_unsigned_to_nat(2u); +x_136 = l_Lean_Syntax_getArg(x_1, x_135); +x_137 = lean_unsigned_to_nat(4u); +x_138 = l_Lean_Syntax_getArg(x_1, x_137); lean_dec(x_1); -x_21 = !lean_is_exclusive(x_17); -if (x_21 == 0) -{ -lean_object* x_22; -x_22 = lean_ctor_get(x_17, 0); -lean_dec(x_22); -lean_ctor_set(x_17, 0, x_13); -return x_17; +x_139 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; +x_140 = lean_array_push(x_139, x_136); +x_141 = lean_array_push(x_140, x_138); +x_142 = l_Lean_Elab_Term_toParserDescr_processAlias(x_134, x_141, x_2, x_3, x_4, x_5, x_6, x_7, x_93, x_9, x_10); +return x_142; +} } else { -lean_object* x_23; lean_object* x_24; -x_23 = lean_ctor_get(x_17, 1); -lean_inc(x_23); -lean_dec(x_17); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_13); -lean_ctor_set(x_24, 1, x_23); -return x_24; +lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; +lean_dec(x_11); +x_143 = lean_unsigned_to_nat(0u); +x_144 = l_Lean_Syntax_getArg(x_1, x_143); +x_145 = lean_unsigned_to_nat(2u); +x_146 = l_Lean_Syntax_getArg(x_1, x_145); +lean_dec(x_1); +x_147 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; +x_148 = lean_array_push(x_147, x_146); +x_149 = l_Lean_Elab_Term_toParserDescr_processAlias(x_144, x_148, x_2, x_3, x_4, x_5, x_6, x_7, x_93, x_9, x_10); +return x_149; } } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_17, 1); -lean_inc(x_25); -lean_dec(x_17); -x_26 = lean_box(0); -lean_inc(x_13); -x_27 = l_List_forIn_loop___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__14(x_1, x_2, x_13, x_26, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_25); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -if (lean_obj_tag(x_27) == 0) -{ -uint8_t x_28; -x_28 = !lean_is_exclusive(x_27); -if (x_28 == 0) -{ -lean_object* x_29; -x_29 = lean_ctor_get(x_27, 0); -lean_dec(x_29); -lean_ctor_set(x_27, 0, x_13); -return x_27; +lean_object* x_150; +lean_dec(x_11); +x_150 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_93, x_9, x_10); +return x_150; +} } else { -lean_object* x_30; lean_object* x_31; -x_30 = lean_ctor_get(x_27, 1); -lean_inc(x_30); -lean_dec(x_27); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_13); -lean_ctor_set(x_31, 1, x_30); -return x_31; +lean_object* x_151; lean_object* x_152; +lean_dec(x_11); +x_151 = lean_unsigned_to_nat(1u); +x_152 = l_Lean_Syntax_getArg(x_1, x_151); +lean_dec(x_1); +x_1 = x_152; +x_8 = x_93; +goto _start; } } else { -uint8_t x_32; -lean_dec(x_13); -x_32 = !lean_is_exclusive(x_27); -if (x_32 == 0) -{ -return x_27; +lean_object* x_154; lean_object* x_155; +lean_dec(x_11); +x_154 = lean_unsigned_to_nat(0u); +x_155 = l_Lean_Syntax_getArg(x_1, x_154); +lean_dec(x_1); +x_1 = x_155; +x_8 = x_93; +goto _start; +} } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_27, 0); -x_34 = lean_ctor_get(x_27, 1); -lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_27); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -return x_35; +lean_object* x_157; +lean_dec(x_11); +x_157 = l_Lean_Elab_Term_toParserDescr_processSeq(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_93, x_9, x_10); +return x_157; } } } } -else -{ -uint8_t x_36; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_36 = !lean_is_exclusive(x_12); -if (x_36 == 0) +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__1() { +_start: { -return x_12; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("ParserDescr.sepBy1", 18); +return x_1; } -else +} +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__2() { +_start: { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_12, 0); -x_38 = lean_ctor_get(x_12, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_12); -x_39 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_39, 0, x_37); -lean_ctor_set(x_39, 1, x_38); -return x_39; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__1; +x_2 = lean_string_utf8_byte_size(x_1); +return x_2; } } +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__1; +x_2 = lean_unsigned_to_nat(0u); +x_3 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__2; +x_4 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; } } -LEAN_EXPORT lean_object* l_List_filterMap___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__15(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__4() { _start: { -if (lean_obj_tag(x_2) == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__13; +x_2 = l_Lean_Elab_Term_toParserDescr_process___closed__9; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__5() { +_start: { -lean_object* x_3; -lean_dec(x_1); -x_3 = lean_box(0); +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__16; +x_2 = l_Lean_Elab_Term_toParserDescr_process___closed__9; +x_3 = lean_name_mk_string(x_1, x_2); return x_3; } -else +} +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__6() { +_start: { -uint8_t x_4; -x_4 = !lean_is_exclusive(x_2); -if (x_4 == 0) -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_5 = lean_ctor_get(x_2, 0); -x_6 = lean_ctor_get(x_2, 1); -lean_inc(x_5); -lean_inc(x_1); -x_7 = lean_environment_find(x_1, x_5); -if (lean_obj_tag(x_7) == 0) -{ -lean_free_object(x_2); -lean_dec(x_5); -x_2 = x_6; -goto _start; -} -else -{ -lean_object* x_9; lean_object* x_10; -x_9 = lean_ctor_get(x_7, 0); -lean_inc(x_9); -lean_dec(x_7); -x_10 = l_Lean_ConstantInfo_type(x_9); -lean_dec(x_9); -if (lean_obj_tag(x_10) == 4) -{ -lean_object* x_11; -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -lean_dec(x_10); -if (lean_obj_tag(x_11) == 1) -{ -lean_object* x_12; -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -if (lean_obj_tag(x_12) == 1) -{ -lean_object* x_13; -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -switch (lean_obj_tag(x_13)) { -case 0: -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_14 = lean_ctor_get(x_11, 1); -lean_inc(x_14); -lean_dec(x_11); -x_15 = lean_ctor_get(x_12, 1); -lean_inc(x_15); -lean_dec(x_12); -x_16 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__1; -x_17 = lean_string_dec_eq(x_15, x_16); -lean_dec(x_15); -if (x_17 == 0) -{ -lean_dec(x_14); -lean_free_object(x_2); -lean_dec(x_5); -x_2 = x_6; -goto _start; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__5; +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; } -else -{ -lean_object* x_19; uint8_t x_20; -x_19 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__12; -x_20 = lean_string_dec_eq(x_14, x_19); -if (x_20 == 0) -{ -lean_object* x_21; uint8_t x_22; -x_21 = l_List_filterMap___at_Lean_Elab_Term_resolveParserName___spec__1___closed__1; -x_22 = lean_string_dec_eq(x_14, x_21); -lean_dec(x_14); -if (x_22 == 0) -{ -lean_free_object(x_2); -lean_dec(x_5); -x_2 = x_6; -goto _start; } -else +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__7() { +_start: { -uint8_t x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_24 = 1; -x_25 = lean_box(x_24); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_5); -lean_ctor_set(x_26, 1, x_25); -x_27 = l_List_filterMap___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__15(x_1, x_6); -lean_ctor_set(x_2, 1, x_27); -lean_ctor_set(x_2, 0, x_26); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__6; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } -else +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__8() { +_start: { -uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -lean_dec(x_14); -x_28 = 1; -x_29 = lean_box(x_28); -x_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_5); -lean_ctor_set(x_30, 1, x_29); -x_31 = l_List_filterMap___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__15(x_1, x_6); -lean_ctor_set(x_2, 1, x_31); -lean_ctor_set(x_2, 0, x_30); +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(4u); +x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -} -case 1: -{ -lean_object* x_32; -x_32 = lean_ctor_get(x_13, 0); -lean_inc(x_32); -if (lean_obj_tag(x_32) == 0) -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_33 = lean_ctor_get(x_11, 1); -lean_inc(x_33); -lean_dec(x_11); -x_34 = lean_ctor_get(x_12, 1); -lean_inc(x_34); -lean_dec(x_12); -x_35 = lean_ctor_get(x_13, 1); -lean_inc(x_35); -lean_dec(x_13); -x_36 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__1; -x_37 = lean_string_dec_eq(x_35, x_36); -lean_dec(x_35); -if (x_37 == 0) +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__9() { +_start: { -lean_dec(x_34); -lean_dec(x_33); -lean_free_object(x_2); -lean_dec(x_5); -x_2 = x_6; -goto _start; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__15; +x_3 = l_Lean_mkCIdentFrom(x_1, x_2); +return x_3; } -else -{ -lean_object* x_39; uint8_t x_40; -x_39 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__3; -x_40 = lean_string_dec_eq(x_34, x_39); -lean_dec(x_34); -if (x_40 == 0) -{ -lean_dec(x_33); -lean_free_object(x_2); -lean_dec(x_5); -x_2 = x_6; -goto _start; } -else -{ -lean_object* x_42; uint8_t x_43; -x_42 = l_List_filterMap___at_Lean_Elab_Term_resolveParserName___spec__1___closed__2; -x_43 = lean_string_dec_eq(x_33, x_42); -if (x_43 == 0) -{ -uint8_t x_44; -x_44 = lean_string_dec_eq(x_33, x_39); -lean_dec(x_33); -if (x_44 == 0) +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__10() { +_start: { -lean_free_object(x_2); -lean_dec(x_5); -x_2 = x_6; -goto _start; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("true", 4); +return x_1; } -else +} +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__11() { +_start: { -uint8_t x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_46 = 0; -x_47 = lean_box(x_46); -x_48 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_48, 0, x_5); -lean_ctor_set(x_48, 1, x_47); -x_49 = l_List_filterMap___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__15(x_1, x_6); -lean_ctor_set(x_2, 1, x_49); -lean_ctor_set(x_2, 0, x_48); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__14; +x_2 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__10; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; } } -else +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__12() { +_start: { -uint8_t x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; -lean_dec(x_33); -x_50 = 0; -x_51 = lean_box(x_50); -x_52 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_52, 0, x_5); -lean_ctor_set(x_52, 1, x_51); -x_53 = l_List_filterMap___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__15(x_1, x_6); -lean_ctor_set(x_2, 1, x_53); -lean_ctor_set(x_2, 0, x_52); -return x_2; -} +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__11; +x_3 = l_Lean_mkCIdentFrom(x_1, x_2); +return x_3; } } +LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy1___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, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_92; lean_object* x_93; uint8_t x_94; lean_object* x_95; +x_92 = lean_unsigned_to_nat(5u); +x_93 = l_Lean_Syntax_getArg(x_3, x_92); +x_94 = l_Lean_Syntax_isNone(x_93); +lean_dec(x_93); +lean_inc(x_11); +x_95 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_toParserDescr_processNonReserved___spec__2___rarg(x_11, x_12, x_13); +if (x_94 == 0) +{ +lean_object* x_96; lean_object* x_97; uint8_t x_98; +x_96 = lean_ctor_get(x_95, 0); +lean_inc(x_96); +x_97 = lean_ctor_get(x_95, 1); +lean_inc(x_97); +lean_dec(x_95); +x_98 = 1; +x_14 = x_98; +x_15 = x_96; +x_16 = x_97; +goto block_91; } else { -lean_dec(x_32); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_free_object(x_2); -lean_dec(x_5); -x_2 = x_6; -goto _start; -} +lean_object* x_99; lean_object* x_100; uint8_t x_101; +x_99 = lean_ctor_get(x_95, 0); +lean_inc(x_99); +x_100 = lean_ctor_get(x_95, 1); +lean_inc(x_100); +lean_dec(x_95); +x_101 = 0; +x_14 = x_101; +x_15 = x_99; +x_16 = x_100; +goto block_91; } -default: +block_91: { -lean_dec(x_13); -lean_dec(x_12); +lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_17 = lean_ctor_get(x_11, 10); +lean_inc(x_17); lean_dec(x_11); -lean_free_object(x_2); -lean_dec(x_5); -x_2 = x_6; -goto _start; -} -} +x_18 = lean_st_ref_get(x_12, x_16); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) +{ +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; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_20 = lean_ctor_get(x_18, 0); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +lean_dec(x_20); +x_22 = lean_environment_main_module(x_21); +x_23 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__4; +x_24 = l_Lean_addMacroScope(x_22, x_23, x_17); +x_25 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__3; +x_26 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__7; +x_27 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_27, 0, x_15); +lean_ctor_set(x_27, 1, x_25); +lean_ctor_set(x_27, 2, x_24); +lean_ctor_set(x_27, 3, x_26); +x_28 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__8; +x_29 = lean_array_push(x_28, x_1); +x_30 = lean_array_push(x_29, x_2); +x_31 = lean_array_push(x_30, x_4); +x_32 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; +x_33 = lean_array_push(x_32, x_27); +if (x_14 == 0) +{ +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_34 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__9; +x_35 = lean_array_push(x_31, x_34); +x_36 = lean_box(2); +x_37 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; +x_38 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +lean_ctor_set(x_38, 2, x_35); +x_39 = lean_array_push(x_33, x_38); +x_40 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +x_41 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_41, 0, x_36); +lean_ctor_set(x_41, 1, x_40); +lean_ctor_set(x_41, 2, x_39); +x_42 = lean_unsigned_to_nat(1u); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +lean_ctor_set(x_18, 0, x_43); +return x_18; } else { -lean_dec(x_12); -lean_dec(x_11); -lean_free_object(x_2); -lean_dec(x_5); -x_2 = x_6; -goto _start; +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_44 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__12; +x_45 = lean_array_push(x_31, x_44); +x_46 = lean_box(2); +x_47 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; +x_48 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_48, 0, x_46); +lean_ctor_set(x_48, 1, x_47); +lean_ctor_set(x_48, 2, x_45); +x_49 = lean_array_push(x_33, x_48); +x_50 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +x_51 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_51, 0, x_46); +lean_ctor_set(x_51, 1, x_50); +lean_ctor_set(x_51, 2, x_49); +x_52 = lean_unsigned_to_nat(1u); +x_53 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_52); +lean_ctor_set(x_18, 0, x_53); +return x_18; } } else { -lean_dec(x_11); -lean_free_object(x_2); -lean_dec(x_5); -x_2 = x_6; -goto _start; -} +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; lean_object* x_67; lean_object* x_68; +x_54 = lean_ctor_get(x_18, 0); +x_55 = lean_ctor_get(x_18, 1); +lean_inc(x_55); +lean_inc(x_54); +lean_dec(x_18); +x_56 = lean_ctor_get(x_54, 0); +lean_inc(x_56); +lean_dec(x_54); +x_57 = lean_environment_main_module(x_56); +x_58 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__4; +x_59 = l_Lean_addMacroScope(x_57, x_58, x_17); +x_60 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__3; +x_61 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__7; +x_62 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_62, 0, x_15); +lean_ctor_set(x_62, 1, x_60); +lean_ctor_set(x_62, 2, x_59); +lean_ctor_set(x_62, 3, x_61); +x_63 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__8; +x_64 = lean_array_push(x_63, x_1); +x_65 = lean_array_push(x_64, x_2); +x_66 = lean_array_push(x_65, x_4); +x_67 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; +x_68 = lean_array_push(x_67, x_62); +if (x_14 == 0) +{ +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; lean_object* x_79; +x_69 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__9; +x_70 = lean_array_push(x_66, x_69); +x_71 = lean_box(2); +x_72 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; +x_73 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_73, 0, x_71); +lean_ctor_set(x_73, 1, x_72); +lean_ctor_set(x_73, 2, x_70); +x_74 = lean_array_push(x_68, x_73); +x_75 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +x_76 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_76, 0, x_71); +lean_ctor_set(x_76, 1, x_75); +lean_ctor_set(x_76, 2, x_74); +x_77 = lean_unsigned_to_nat(1u); +x_78 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_78, 0, x_76); +lean_ctor_set(x_78, 1, x_77); +x_79 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_79, 0, x_78); +lean_ctor_set(x_79, 1, x_55); +return x_79; } else { -lean_dec(x_10); -lean_free_object(x_2); -lean_dec(x_5); -x_2 = x_6; -goto _start; +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_80 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__12; +x_81 = lean_array_push(x_66, x_80); +x_82 = lean_box(2); +x_83 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; +x_84 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_84, 0, x_82); +lean_ctor_set(x_84, 1, x_83); +lean_ctor_set(x_84, 2, x_81); +x_85 = lean_array_push(x_68, x_84); +x_86 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +x_87 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_87, 0, x_82); +lean_ctor_set(x_87, 1, x_86); +lean_ctor_set(x_87, 2, x_85); +x_88 = lean_unsigned_to_nat(1u); +x_89 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_89, 0, x_87); +lean_ctor_set(x_89, 1, x_88); +x_90 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_90, 0, x_89); +lean_ctor_set(x_90, 1, x_55); +return x_90; } } } -else -{ -lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_59 = lean_ctor_get(x_2, 0); -x_60 = lean_ctor_get(x_2, 1); -lean_inc(x_60); -lean_inc(x_59); -lean_dec(x_2); -lean_inc(x_59); -lean_inc(x_1); -x_61 = lean_environment_find(x_1, x_59); -if (lean_obj_tag(x_61) == 0) -{ -lean_dec(x_59); -x_2 = x_60; -goto _start; } -else -{ -lean_object* x_63; lean_object* x_64; -x_63 = lean_ctor_get(x_61, 0); -lean_inc(x_63); -lean_dec(x_61); -x_64 = l_Lean_ConstantInfo_type(x_63); -lean_dec(x_63); -if (lean_obj_tag(x_64) == 4) -{ -lean_object* x_65; -x_65 = lean_ctor_get(x_64, 0); -lean_inc(x_65); -lean_dec(x_64); -if (lean_obj_tag(x_65) == 1) +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy1(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_object* x_9, lean_object* x_10) { +_start: { -lean_object* x_66; -x_66 = lean_ctor_get(x_65, 0); -lean_inc(x_66); -if (lean_obj_tag(x_66) == 1) +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; +x_11 = lean_unsigned_to_nat(1u); +x_12 = l_Lean_Syntax_getArg(x_1, x_11); +x_13 = lean_ctor_get(x_2, 0); +lean_inc(x_13); +x_14 = lean_ctor_get_uint8(x_2, sizeof(void*)*1 + 2); +x_15 = 0; +x_16 = lean_alloc_ctor(0, 1, 3); +lean_ctor_set(x_16, 0, x_13); +lean_ctor_set_uint8(x_16, sizeof(void*)*1, x_15); +lean_ctor_set_uint8(x_16, sizeof(void*)*1 + 1, x_15); +lean_ctor_set_uint8(x_16, sizeof(void*)*1 + 2, x_14); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_16); +x_17 = l_Lean_Elab_Term_toParserDescr_process(x_12, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_17) == 0) { -lean_object* x_67; -x_67 = lean_ctor_get(x_66, 0); -lean_inc(x_67); -switch (lean_obj_tag(x_67)) { -case 0: +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; uint8_t x_25; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = l_Lean_Elab_Term_ensureUnaryOutput(x_18); +x_21 = lean_unsigned_to_nat(3u); +x_22 = l_Lean_Syntax_getArg(x_1, x_21); +x_23 = lean_unsigned_to_nat(4u); +x_24 = l_Lean_Syntax_getArg(x_1, x_23); +x_25 = l_Lean_Syntax_isNone(x_24); +if (x_25 == 0) { -lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; -x_68 = lean_ctor_get(x_65, 1); -lean_inc(x_68); -lean_dec(x_65); -x_69 = lean_ctor_get(x_66, 1); -lean_inc(x_69); -lean_dec(x_66); -x_70 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__1; -x_71 = lean_string_dec_eq(x_69, x_70); -lean_dec(x_69); -if (x_71 == 0) +lean_object* x_26; lean_object* x_27; +x_26 = l_Lean_Syntax_getArg(x_24, x_11); +lean_dec(x_24); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_27 = l_Lean_Elab_Term_toParserDescr_process(x_26, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_19); +if (lean_obj_tag(x_27) == 0) { -lean_dec(x_68); -lean_dec(x_59); -x_2 = x_60; -goto _start; +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_27, 1); +lean_inc(x_29); +lean_dec(x_27); +x_30 = l_Lean_Elab_Term_ensureUnaryOutput(x_28); +x_31 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1(x_20, x_22, x_1, x_30, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_29); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_31; } else { -lean_object* x_73; uint8_t x_74; -x_73 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__12; -x_74 = lean_string_dec_eq(x_68, x_73); -if (x_74 == 0) -{ -lean_object* x_75; uint8_t x_76; -x_75 = l_List_filterMap___at_Lean_Elab_Term_resolveParserName___spec__1___closed__1; -x_76 = lean_string_dec_eq(x_68, x_75); -lean_dec(x_68); -if (x_76 == 0) +uint8_t x_32; +lean_dec(x_22); +lean_dec(x_20); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_32 = !lean_is_exclusive(x_27); +if (x_32 == 0) { -lean_dec(x_59); -x_2 = x_60; -goto _start; +return x_27; } else { -uint8_t x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; -x_78 = 1; -x_79 = lean_box(x_78); -x_80 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_80, 0, x_59); -lean_ctor_set(x_80, 1, x_79); -x_81 = l_List_filterMap___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__15(x_1, x_60); -x_82 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_82, 0, x_80); -lean_ctor_set(x_82, 1, x_81); -return x_82; +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_27, 0); +x_34 = lean_ctor_get(x_27, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_27); +x_35 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; +} } } else { -uint8_t x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; -lean_dec(x_68); -x_83 = 1; -x_84 = lean_box(x_83); -x_85 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_85, 0, x_59); -lean_ctor_set(x_85, 1, x_84); -x_86 = l_List_filterMap___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__15(x_1, x_60); -x_87 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_87, 0, x_85); -lean_ctor_set(x_87, 1, x_86); -return x_87; -} -} -} -case 1: -{ -lean_object* x_88; -x_88 = lean_ctor_get(x_67, 0); -lean_inc(x_88); -if (lean_obj_tag(x_88) == 0) -{ -lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; uint8_t x_93; -x_89 = lean_ctor_get(x_65, 1); -lean_inc(x_89); -lean_dec(x_65); -x_90 = lean_ctor_get(x_66, 1); -lean_inc(x_90); -lean_dec(x_66); -x_91 = lean_ctor_get(x_67, 1); -lean_inc(x_91); -lean_dec(x_67); -x_92 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__1; -x_93 = lean_string_dec_eq(x_91, x_92); -lean_dec(x_91); -if (x_93 == 0) -{ -lean_dec(x_90); -lean_dec(x_89); -lean_dec(x_59); -x_2 = x_60; -goto _start; +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +lean_dec(x_24); +lean_dec(x_16); +lean_inc(x_8); +x_36 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_toParserDescr_processNonReserved___spec__2___rarg(x_8, x_9, x_19); +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +lean_dec(x_36); +x_39 = lean_ctor_get(x_8, 10); +lean_inc(x_39); +x_40 = lean_st_ref_get(x_9, x_38); +x_41 = lean_ctor_get(x_40, 0); +lean_inc(x_41); +x_42 = lean_ctor_get(x_40, 1); +lean_inc(x_42); +lean_dec(x_40); +x_43 = lean_ctor_get(x_41, 0); +lean_inc(x_43); +lean_dec(x_41); +x_44 = lean_environment_main_module(x_43); +x_45 = l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__5; +x_46 = l_Lean_addMacroScope(x_44, x_45, x_39); +x_47 = l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__3; +x_48 = l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__8; +x_49 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_49, 0, x_37); +lean_ctor_set(x_49, 1, x_47); +lean_ctor_set(x_49, 2, x_46); +lean_ctor_set(x_49, 3, x_48); +x_50 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; +lean_inc(x_22); +x_51 = lean_array_push(x_50, x_22); +x_52 = lean_box(2); +x_53 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; +x_54 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +lean_ctor_set(x_54, 2, x_51); +x_55 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; +x_56 = lean_array_push(x_55, x_49); +x_57 = lean_array_push(x_56, x_54); +x_58 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +x_59 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_59, 0, x_52); +lean_ctor_set(x_59, 1, x_58); +lean_ctor_set(x_59, 2, x_57); +x_60 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1(x_20, x_22, x_1, x_59, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_42); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_60; } -else -{ -lean_object* x_95; uint8_t x_96; -x_95 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__3; -x_96 = lean_string_dec_eq(x_90, x_95); -lean_dec(x_90); -if (x_96 == 0) -{ -lean_dec(x_89); -lean_dec(x_59); -x_2 = x_60; -goto _start; } else { -lean_object* x_98; uint8_t x_99; -x_98 = l_List_filterMap___at_Lean_Elab_Term_resolveParserName___spec__1___closed__2; -x_99 = lean_string_dec_eq(x_89, x_98); -if (x_99 == 0) -{ -uint8_t x_100; -x_100 = lean_string_dec_eq(x_89, x_95); -lean_dec(x_89); -if (x_100 == 0) -{ -lean_dec(x_59); -x_2 = x_60; -goto _start; -} -else +uint8_t x_61; +lean_dec(x_16); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_61 = !lean_is_exclusive(x_17); +if (x_61 == 0) { -uint8_t x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; -x_102 = 0; -x_103 = lean_box(x_102); -x_104 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_104, 0, x_59); -lean_ctor_set(x_104, 1, x_103); -x_105 = l_List_filterMap___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__15(x_1, x_60); -x_106 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_106, 0, x_104); -lean_ctor_set(x_106, 1, x_105); -return x_106; -} +return x_17; } else { -uint8_t x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; -lean_dec(x_89); -x_107 = 0; -x_108 = lean_box(x_107); -x_109 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_109, 0, x_59); -lean_ctor_set(x_109, 1, x_108); -x_110 = l_List_filterMap___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__15(x_1, x_60); -x_111 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_111, 0, x_109); -lean_ctor_set(x_111, 1, x_110); -return x_111; +lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_62 = lean_ctor_get(x_17, 0); +x_63 = lean_ctor_get(x_17, 1); +lean_inc(x_63); +lean_inc(x_62); +lean_dec(x_17); +x_64 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_64, 0, x_62); +lean_ctor_set(x_64, 1, x_63); +return x_64; } } } } -else +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__1() { +_start: { -lean_dec(x_88); -lean_dec(x_67); -lean_dec(x_66); -lean_dec(x_65); -lean_dec(x_59); -x_2 = x_60; -goto _start; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("ParserDescr.sepBy", 17); +return x_1; } } -default: +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__2() { +_start: { -lean_dec(x_67); -lean_dec(x_66); -lean_dec(x_65); -lean_dec(x_59); -x_2 = x_60; -goto _start; -} +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__1; +x_2 = lean_string_utf8_byte_size(x_1); +return x_2; } } -else +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__3() { +_start: { -lean_dec(x_66); -lean_dec(x_65); -lean_dec(x_59); -x_2 = x_60; -goto _start; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__1; +x_2 = lean_unsigned_to_nat(0u); +x_3 = l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__2; +x_4 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; } } -else +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__4() { +_start: { -lean_dec(x_65); -lean_dec(x_59); -x_2 = x_60; -goto _start; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__13; +x_2 = l_Lean_Elab_Term_toParserDescr_process___closed__7; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; } } -else +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__5() { +_start: { -lean_dec(x_64); -lean_dec(x_59); -x_2 = x_60; -goto _start; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__16; +x_2 = l_Lean_Elab_Term_toParserDescr_process___closed__7; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; } } +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__5; +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_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__6; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveParserName___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__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, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy___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, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_11; lean_object* x_12; -x_11 = lean_box(0); -lean_inc(x_9); -x_12 = l_Lean_Elab_resolveGlobalConstWithInfos___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__2(x_1, x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_12) == 0) -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec(x_12); -x_15 = lean_st_ref_get(x_9, x_14); -lean_dec(x_9); -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) +uint8_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_92; lean_object* x_93; uint8_t x_94; lean_object* x_95; +x_92 = lean_unsigned_to_nat(5u); +x_93 = l_Lean_Syntax_getArg(x_3, x_92); +x_94 = l_Lean_Syntax_isNone(x_93); +lean_dec(x_93); +lean_inc(x_11); +x_95 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_toParserDescr_processNonReserved___spec__2___rarg(x_11, x_12, x_13); +if (x_94 == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_15, 0); -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -lean_dec(x_17); -x_19 = l_List_filterMap___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__15(x_18, x_13); -lean_ctor_set(x_15, 0, x_19); -return x_15; +lean_object* x_96; lean_object* x_97; uint8_t x_98; +x_96 = lean_ctor_get(x_95, 0); +lean_inc(x_96); +x_97 = lean_ctor_get(x_95, 1); +lean_inc(x_97); +lean_dec(x_95); +x_98 = 1; +x_14 = x_98; +x_15 = x_96; +x_16 = x_97; +goto block_91; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_20 = lean_ctor_get(x_15, 0); -x_21 = lean_ctor_get(x_15, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_15); -x_22 = lean_ctor_get(x_20, 0); -lean_inc(x_22); -lean_dec(x_20); -x_23 = l_List_filterMap___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__15(x_22, x_13); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_21); -return x_24; -} +lean_object* x_99; lean_object* x_100; uint8_t x_101; +x_99 = lean_ctor_get(x_95, 0); +lean_inc(x_99); +x_100 = lean_ctor_get(x_95, 1); +lean_inc(x_100); +lean_dec(x_95); +x_101 = 0; +x_14 = x_101; +x_15 = x_99; +x_16 = x_100; +goto block_91; } -else +block_91: { -uint8_t x_25; -lean_dec(x_9); -x_25 = !lean_is_exclusive(x_12); -if (x_25 == 0) +lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_17 = lean_ctor_get(x_11, 10); +lean_inc(x_17); +lean_dec(x_11); +x_18 = lean_st_ref_get(x_12, x_16); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) { -lean_object* x_26; lean_object* x_27; -x_26 = lean_ctor_get(x_12, 0); -lean_dec(x_26); -x_27 = lean_box(0); -lean_ctor_set_tag(x_12, 0); -lean_ctor_set(x_12, 0, x_27); -return x_12; +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; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_20 = lean_ctor_get(x_18, 0); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +lean_dec(x_20); +x_22 = lean_environment_main_module(x_21); +x_23 = l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__4; +x_24 = l_Lean_addMacroScope(x_22, x_23, x_17); +x_25 = l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__3; +x_26 = l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__7; +x_27 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_27, 0, x_15); +lean_ctor_set(x_27, 1, x_25); +lean_ctor_set(x_27, 2, x_24); +lean_ctor_set(x_27, 3, x_26); +x_28 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__8; +x_29 = lean_array_push(x_28, x_1); +x_30 = lean_array_push(x_29, x_2); +x_31 = lean_array_push(x_30, x_4); +x_32 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; +x_33 = lean_array_push(x_32, x_27); +if (x_14 == 0) +{ +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_34 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__9; +x_35 = lean_array_push(x_31, x_34); +x_36 = lean_box(2); +x_37 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; +x_38 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +lean_ctor_set(x_38, 2, x_35); +x_39 = lean_array_push(x_33, x_38); +x_40 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +x_41 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_41, 0, x_36); +lean_ctor_set(x_41, 1, x_40); +lean_ctor_set(x_41, 2, x_39); +x_42 = lean_unsigned_to_nat(1u); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +lean_ctor_set(x_18, 0, x_43); +return x_18; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_12, 1); -lean_inc(x_28); -lean_dec(x_12); -x_29 = lean_box(0); -x_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_28); -return x_30; -} -} +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_44 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__12; +x_45 = lean_array_push(x_31, x_44); +x_46 = lean_box(2); +x_47 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; +x_48 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_48, 0, x_46); +lean_ctor_set(x_48, 1, x_47); +lean_ctor_set(x_48, 2, x_45); +x_49 = lean_array_push(x_33, x_48); +x_50 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +x_51 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_51, 0, x_46); +lean_ctor_set(x_51, 1, x_50); +lean_ctor_set(x_51, 2, x_49); +x_52 = lean_unsigned_to_nat(1u); +x_53 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_52); +lean_ctor_set(x_18, 0, x_53); +return x_18; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(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_object* x_9, lean_object* x_10) { -_start: +else { -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = lean_ctor_get(x_8, 5); -x_12 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_6, x_7, x_8, x_9, x_10); -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; lean_object* x_67; lean_object* x_68; +x_54 = lean_ctor_get(x_18, 0); +x_55 = lean_ctor_get(x_18, 1); +lean_inc(x_55); +lean_inc(x_54); +lean_dec(x_18); +x_56 = lean_ctor_get(x_54, 0); +lean_inc(x_56); +lean_dec(x_54); +x_57 = lean_environment_main_module(x_56); +x_58 = l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__4; +x_59 = l_Lean_addMacroScope(x_57, x_58, x_17); +x_60 = l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__3; +x_61 = l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__7; +x_62 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_62, 0, x_15); +lean_ctor_set(x_62, 1, x_60); +lean_ctor_set(x_62, 2, x_59); +lean_ctor_set(x_62, 3, x_61); +x_63 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__8; +x_64 = lean_array_push(x_63, x_1); +x_65 = lean_array_push(x_64, x_2); +x_66 = lean_array_push(x_65, x_4); +x_67 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; +x_68 = lean_array_push(x_67, x_62); +if (x_14 == 0) { -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_12, 0); -lean_inc(x_11); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_11); -lean_ctor_set(x_15, 1, x_14); -lean_ctor_set_tag(x_12, 1); -lean_ctor_set(x_12, 0, x_15); -return x_12; +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; lean_object* x_79; +x_69 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__9; +x_70 = lean_array_push(x_66, x_69); +x_71 = lean_box(2); +x_72 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; +x_73 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_73, 0, x_71); +lean_ctor_set(x_73, 1, x_72); +lean_ctor_set(x_73, 2, x_70); +x_74 = lean_array_push(x_68, x_73); +x_75 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +x_76 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_76, 0, x_71); +lean_ctor_set(x_76, 1, x_75); +lean_ctor_set(x_76, 2, x_74); +x_77 = lean_unsigned_to_nat(1u); +x_78 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_78, 0, x_76); +lean_ctor_set(x_78, 1, x_77); +x_79 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_79, 0, x_78); +lean_ctor_set(x_79, 1, x_55); +return x_79; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_16 = lean_ctor_get(x_12, 0); -x_17 = lean_ctor_get(x_12, 1); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_12); -lean_inc(x_11); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_11); -lean_ctor_set(x_18, 1, x_16); -x_19 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_17); -return x_19; +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_80 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__12; +x_81 = lean_array_push(x_66, x_80); +x_82 = lean_box(2); +x_83 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; +x_84 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_84, 0, x_82); +lean_ctor_set(x_84, 1, x_83); +lean_ctor_set(x_84, 2, x_81); +x_85 = lean_array_push(x_68, x_84); +x_86 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +x_87 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_87, 0, x_82); +lean_ctor_set(x_87, 1, x_86); +lean_ctor_set(x_87, 2, x_85); +x_88 = lean_unsigned_to_nat(1u); +x_89 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_89, 0, x_87); +lean_ctor_set(x_89, 1, x_88); +x_90 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_90, 0, x_89); +lean_ctor_set(x_90, 1, x_55); +return x_90; +} } } } -LEAN_EXPORT lean_object* l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(lean_object* x_1, lean_object* x_2) { +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy(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_object* x_9, lean_object* x_10) { _start: { -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_3; -x_3 = l_List_reverse___rarg(x_2); -return x_3; -} -else +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; +x_11 = lean_unsigned_to_nat(1u); +x_12 = l_Lean_Syntax_getArg(x_1, x_11); +x_13 = lean_ctor_get(x_2, 0); +lean_inc(x_13); +x_14 = lean_ctor_get_uint8(x_2, sizeof(void*)*1 + 2); +x_15 = 0; +x_16 = lean_alloc_ctor(0, 1, 3); +lean_ctor_set(x_16, 0, x_13); +lean_ctor_set_uint8(x_16, sizeof(void*)*1, x_15); +lean_ctor_set_uint8(x_16, sizeof(void*)*1 + 1, x_15); +lean_ctor_set_uint8(x_16, sizeof(void*)*1 + 2, x_14); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_16); +x_17 = l_Lean_Elab_Term_toParserDescr_process(x_12, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_17) == 0) { -uint8_t x_4; -x_4 = !lean_is_exclusive(x_1); -if (x_4 == 0) +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; uint8_t x_25; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = l_Lean_Elab_Term_ensureUnaryOutput(x_18); +x_21 = lean_unsigned_to_nat(3u); +x_22 = l_Lean_Syntax_getArg(x_1, x_21); +x_23 = lean_unsigned_to_nat(4u); +x_24 = l_Lean_Syntax_getArg(x_1, x_23); +x_25 = l_Lean_Syntax_isNone(x_24); +if (x_25 == 0) { -lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_5 = lean_ctor_get(x_1, 0); -x_6 = lean_ctor_get(x_1, 1); -x_7 = lean_ctor_get(x_5, 0); +lean_object* x_26; lean_object* x_27; +x_26 = l_Lean_Syntax_getArg(x_24, x_11); +lean_dec(x_24); +lean_inc(x_9); +lean_inc(x_8); lean_inc(x_7); -lean_dec(x_5); -lean_ctor_set(x_1, 1, x_2); -lean_ctor_set(x_1, 0, x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_27 = l_Lean_Elab_Term_toParserDescr_process(x_26, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_19); +if (lean_obj_tag(x_27) == 0) { -lean_object* _tmp_0 = x_6; -lean_object* _tmp_1 = x_1; -x_1 = _tmp_0; -x_2 = _tmp_1; -} -goto _start; +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_27, 1); +lean_inc(x_29); +lean_dec(x_27); +x_30 = l_Lean_Elab_Term_ensureUnaryOutput(x_28); +x_31 = l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1(x_20, x_22, x_1, x_30, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_29); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_31; } else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_ctor_get(x_1, 0); -x_10 = lean_ctor_get(x_1, 1); -lean_inc(x_10); -lean_inc(x_9); -lean_dec(x_1); -x_11 = lean_ctor_get(x_9, 0); -lean_inc(x_11); +uint8_t x_32; +lean_dec(x_22); +lean_dec(x_20); lean_dec(x_9); -x_12 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_2); -x_1 = x_10; -x_2 = x_12; -goto _start; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_32 = !lean_is_exclusive(x_27); +if (x_32 == 0) +{ +return x_27; } +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_27, 0); +x_34 = lean_ctor_get(x_27, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_27); +x_35 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; } } } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__1() { +else +{ +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +lean_dec(x_24); +lean_dec(x_16); +lean_inc(x_8); +x_36 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_toParserDescr_processNonReserved___spec__2___rarg(x_8, x_9, x_19); +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +lean_dec(x_36); +x_39 = lean_ctor_get(x_8, 10); +lean_inc(x_39); +x_40 = lean_st_ref_get(x_9, x_38); +x_41 = lean_ctor_get(x_40, 0); +lean_inc(x_41); +x_42 = lean_ctor_get(x_40, 1); +lean_inc(x_42); +lean_dec(x_40); +x_43 = lean_ctor_get(x_41, 0); +lean_inc(x_43); +lean_dec(x_41); +x_44 = lean_environment_main_module(x_43); +x_45 = l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__5; +x_46 = l_Lean_addMacroScope(x_44, x_45, x_39); +x_47 = l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__3; +x_48 = l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__8; +x_49 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_49, 0, x_37); +lean_ctor_set(x_49, 1, x_47); +lean_ctor_set(x_49, 2, x_46); +lean_ctor_set(x_49, 3, x_48); +x_50 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; +lean_inc(x_22); +x_51 = lean_array_push(x_50, x_22); +x_52 = lean_box(2); +x_53 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; +x_54 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +lean_ctor_set(x_54, 2, x_51); +x_55 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; +x_56 = lean_array_push(x_55, x_49); +x_57 = lean_array_push(x_56, x_54); +x_58 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +x_59 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_59, 0, x_52); +lean_ctor_set(x_59, 1, x_58); +lean_ctor_set(x_59, 2, x_57); +x_60 = l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1(x_20, x_22, x_1, x_59, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_42); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_60; +} +} +else +{ +uint8_t x_61; +lean_dec(x_16); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_61 = !lean_is_exclusive(x_17); +if (x_61 == 0) +{ +return x_17; +} +else +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_62 = lean_ctor_get(x_17, 0); +x_63 = lean_ctor_get(x_17, 1); +lean_inc(x_63); +lean_inc(x_62); +lean_dec(x_17); +x_64 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_64, 0, x_62); +lean_ctor_set(x_64, 1, x_63); +return x_64; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_toParserDescr_processAlias___spec__1(size_t x_1, size_t 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, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("unknown parser declaration/category/alias '", 43); -return x_1; +uint8_t x_13; +x_13 = lean_usize_dec_lt(x_2, x_1); +if (x_13 == 0) +{ +lean_object* x_14; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_3); +lean_ctor_set(x_14, 1, x_12); +return x_14; } +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; +x_15 = lean_array_uget(x_3, x_2); +x_16 = lean_unsigned_to_nat(0u); +x_17 = lean_array_uset(x_3, x_2, x_16); +x_18 = lean_ctor_get(x_4, 0); +x_19 = lean_ctor_get_uint8(x_4, sizeof(void*)*1 + 2); +x_20 = 0; +lean_inc(x_18); +x_21 = lean_alloc_ctor(0, 1, 3); +lean_ctor_set(x_21, 0, x_18); +lean_ctor_set_uint8(x_21, sizeof(void*)*1, x_20); +lean_ctor_set_uint8(x_21, sizeof(void*)*1 + 1, x_20); +lean_ctor_set_uint8(x_21, sizeof(void*)*1 + 2, x_19); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_22 = l_Lean_Elab_Term_toParserDescr_process(x_15, x_21, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; lean_object* x_24; size_t x_25; size_t x_26; lean_object* x_27; +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = 1; +x_26 = lean_usize_add(x_2, x_25); +x_27 = lean_array_uset(x_17, x_2, x_23); +x_2 = x_26; +x_3 = x_27; +x_12 = x_24; +goto _start; } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__2() { +else +{ +uint8_t x_29; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_29 = !lean_is_exclusive(x_22); +if (x_29 == 0) +{ +return x_22; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_22, 0); +x_31 = lean_ctor_get(x_22, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_22); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +return x_32; +} +} +} +} +} +static lean_object* _init_l_panic___at_Lean_Elab_Term_toParserDescr_processAlias___spec__2___closed__1() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); +x_1 = l_Lean_Elab_Term_instMonadTermElabM; +x_2 = l_ReaderT_instMonadReaderT___rarg(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__3() { +static lean_object* _init_l_panic___at_Lean_Elab_Term_toParserDescr_processAlias___spec__2___closed__2() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("ParserDescr.const", 17); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_panic___at_Lean_Elab_Term_toParserDescr_processAlias___spec__2___closed__1; +x_2 = lean_box(0); +x_3 = l_instInhabited___rarg(x_1, x_2); +return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__4() { +static lean_object* _init_l_panic___at_Lean_Elab_Term_toParserDescr_processAlias___spec__2___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__3; -x_2 = lean_string_utf8_byte_size(x_1); +x_1 = l_panic___at_Lean_Elab_Term_toParserDescr_processAlias___spec__2___closed__2; +x_2 = lean_alloc_closure((void*)(l_instInhabitedReaderT___rarg___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__5() { +LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_Term_toParserDescr_processAlias___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_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__3; -x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__4; -x_4 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_4, 0, x_1); -lean_ctor_set(x_4, 1, x_2); -lean_ctor_set(x_4, 2, x_3); -return x_4; +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = l_panic___at_Lean_Elab_Term_toParserDescr_processAlias___spec__2___closed__3; +x_12 = lean_panic_fn(x_11, x_1); +x_13 = lean_apply_9(x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_13; } } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__6() { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_toParserDescr_processAlias___spec__3(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("const", 5); -return x_1; -} +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_2, x_3); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; +x_6 = lean_array_uget(x_1, x_2); +x_7 = lean_nat_add(x_4, x_6); +lean_dec(x_6); +lean_dec(x_4); +x_8 = 1; +x_9 = lean_usize_add(x_2, x_8); +x_2 = x_9; +x_4 = x_7; +goto _start; } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__7() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__13; -x_2 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__6; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; +return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__8() { +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_toParserDescr_processAlias___spec__4(size_t x_1, size_t x_2, lean_object* x_3) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__16; -x_2 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__6; -x_3 = lean_name_mk_string(x_1, x_2); +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ return x_3; } +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = lean_ctor_get(x_5, 0); +lean_inc(x_8); +lean_dec(x_5); +x_9 = 1; +x_10 = lean_usize_add(x_2, x_9); +x_11 = lean_array_uset(x_7, x_2, x_8); +x_2 = x_10; +x_3 = x_11; +goto _start; } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__9() { +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_toParserDescr_processAlias___spec__5(size_t x_1, size_t x_2, lean_object* x_3) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__8; -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ return x_3; } +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = l_Lean_Elab_Term_ensureUnaryOutput(x_5); +x_9 = 1; +x_10 = lean_usize_add(x_2, x_9); +x_11 = lean_array_uset(x_7, x_2, x_8); +x_2 = x_10; +x_3 = x_11; +goto _start; +} } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__10() { +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processAlias___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, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__9; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; +lean_object* x_12; lean_object* x_13; +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_2); +lean_ctor_set(x_12, 1, x_1); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_11); +return x_13; } } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__11() { +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processAlias___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("ParserDescr.parser", 18); +x_1 = lean_mk_string_from_bytes("Lean.Elab.Syntax", 16); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__12() { +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processAlias___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__11; -x_2 = lean_string_utf8_byte_size(x_1); -return x_2; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Lean.Elab.Term.toParserDescr.processAlias", 41); +return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__13() { +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processAlias___closed__3() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__11; -x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__12; -x_4 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_4, 0, x_1); -lean_ctor_set(x_4, 1, x_2); -lean_ctor_set(x_4, 2, x_3); -return x_4; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("unreachable code has been reached", 33); +return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__14() { +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processAlias___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_Elab_Term_toParserDescr_processAlias___closed__1; +x_2 = l_Lean_Elab_Term_toParserDescr_processAlias___closed__2; +x_3 = lean_unsigned_to_nat(168u); +x_4 = lean_unsigned_to_nat(21u); +x_5 = l_Lean_Elab_Term_toParserDescr_processAlias___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processAlias___closed__5() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("parser", 6); +x_1 = lean_mk_string_from_bytes("ParserDescr.const", 17); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processAlias___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Term_toParserDescr_processAlias___closed__5; +x_2 = lean_string_utf8_byte_size(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processAlias___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Elab_Term_toParserDescr_processAlias___closed__5; +x_2 = lean_unsigned_to_nat(0u); +x_3 = l_Lean_Elab_Term_toParserDescr_processAlias___closed__6; +x_4 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processAlias___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("const", 5); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__15() { +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processAlias___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__13; -x_2 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__14; +x_2 = l_Lean_Elab_Term_toParserDescr_processAlias___closed__8; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__16() { +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processAlias___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__16; -x_2 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__14; +x_2 = l_Lean_Elab_Term_toParserDescr_processAlias___closed__8; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__17() { +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processAlias___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__16; +x_2 = l_Lean_Elab_Term_toParserDescr_processAlias___closed__10; 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_Elab_Term_toParserDescr_processNullaryOrCat___closed__18() { +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processAlias___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__17; +x_2 = l_Lean_Elab_Term_toParserDescr_processAlias___closed__11; x_3 = lean_alloc_ctor(1, 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_Elab_Term_toParserDescr_processNullaryOrCat___closed__19() { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processAlias(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_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("ambiguous parser declaration ", 29); -return x_1; +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; lean_object* x_20; size_t x_21; size_t x_22; lean_object* x_23; +x_12 = l_Lean_Syntax_getId(x_1); +lean_dec(x_1); +x_13 = lean_erase_macro_scopes(x_12); +x_14 = lean_st_ref_get(x_10, x_11); +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +lean_dec(x_14); +x_16 = lean_ctor_get(x_9, 5); +lean_inc(x_16); +x_17 = l_Lean_Parser_getParserAliasInfo(x_13, x_15); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_array_get_size(x_2); +x_21 = lean_usize_of_nat(x_20); +lean_dec(x_20); +x_22 = 0; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_23 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_toParserDescr_processAlias___spec__1(x_21, x_22, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_19); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_243; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_243 = lean_ctor_get(x_18, 0); +lean_inc(x_243); +if (lean_obj_tag(x_243) == 0) +{ +lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; uint8_t x_249; +lean_dec(x_18); +x_244 = l_Array_unzip___rarg(x_24); +lean_dec(x_24); +x_245 = lean_ctor_get(x_244, 0); +lean_inc(x_245); +x_246 = lean_ctor_get(x_244, 1); +lean_inc(x_246); +lean_dec(x_244); +x_247 = lean_array_get_size(x_246); +x_248 = lean_unsigned_to_nat(0u); +x_249 = lean_nat_dec_lt(x_248, x_247); +if (x_249 == 0) +{ +lean_dec(x_247); +lean_dec(x_246); +x_26 = x_245; +x_27 = x_248; +goto block_242; } +else +{ +uint8_t x_250; +x_250 = lean_nat_dec_le(x_247, x_247); +if (x_250 == 0) +{ +lean_dec(x_247); +lean_dec(x_246); +x_26 = x_245; +x_27 = x_248; +goto block_242; } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__20() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__19; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +size_t x_251; lean_object* x_252; +x_251 = lean_usize_of_nat(x_247); +lean_dec(x_247); +x_252 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_toParserDescr_processAlias___spec__3(x_246, x_22, x_251, x_248); +lean_dec(x_246); +x_26 = x_245; +x_27 = x_252; +goto block_242; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processNullaryOrCat(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_object* x_9, lean_object* x_10) { -_start: +} +else { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_11 = lean_unsigned_to_nat(0u); -x_12 = l_Lean_Syntax_getArg(x_1, x_11); +uint8_t x_253; +x_253 = lean_ctor_get_uint8(x_18, sizeof(void*)*1); +lean_dec(x_18); +if (x_253 == 0) +{ +lean_object* x_254; lean_object* x_255; size_t x_256; lean_object* x_257; +x_254 = lean_ctor_get(x_243, 0); +lean_inc(x_254); +lean_dec(x_243); +x_255 = lean_array_get_size(x_24); +x_256 = lean_usize_of_nat(x_255); +lean_dec(x_255); +x_257 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_toParserDescr_processAlias___spec__4(x_256, x_22, x_24); +x_26 = x_257; +x_27 = x_254; +goto block_242; +} +else +{ +lean_object* x_258; lean_object* x_259; size_t x_260; lean_object* x_261; +x_258 = lean_ctor_get(x_243, 0); +lean_inc(x_258); +lean_dec(x_243); +x_259 = lean_array_get_size(x_24); +x_260 = lean_usize_of_nat(x_259); +lean_dec(x_259); +x_261 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_toParserDescr_processAlias___spec__5(x_260, x_22, x_24); +x_26 = x_261; +x_27 = x_258; +goto block_242; +} +} +block_242: +{ +lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_28 = lean_array_get_size(x_26); +x_29 = lean_unsigned_to_nat(0u); +x_30 = lean_nat_dec_eq(x_28, x_29); +if (x_30 == 0) +{ +lean_object* x_31; uint8_t x_32; +x_31 = lean_unsigned_to_nat(1u); +x_32 = lean_nat_dec_eq(x_28, x_31); +if (x_32 == 0) +{ +lean_object* x_33; uint8_t x_34; +x_33 = lean_unsigned_to_nat(2u); +x_34 = lean_nat_dec_eq(x_28, x_33); +lean_dec(x_28); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; +lean_dec(x_26); +lean_dec(x_16); +lean_dec(x_13); +x_35 = l_Lean_Elab_Term_toParserDescr_processAlias___closed__4; +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); @@ -7294,65 +8426,17 @@ lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_12); -x_13 = l_Lean_Elab_Term_resolveParserName___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__1(x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -if (lean_obj_tag(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; uint8_t x_22; -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = l_Lean_Syntax_getId(x_12); -lean_dec(x_12); -x_17 = lean_erase_macro_scopes(x_16); -x_18 = lean_st_ref_get(x_9, x_15); -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); -lean_inc(x_20); -lean_dec(x_18); -x_21 = lean_ctor_get(x_19, 0); -lean_inc(x_21); -lean_dec(x_19); -lean_inc(x_17); -x_22 = l_Lean_Parser_isParserCategory(x_21, x_17); -lean_dec(x_21); -if (x_22 == 0) -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; -x_23 = lean_st_ref_get(x_9, x_20); -x_24 = lean_ctor_get(x_23, 1); -lean_inc(x_24); -lean_dec(x_23); -x_25 = lean_ctor_get(x_8, 5); -lean_inc(x_25); -x_26 = l_Lean_Parser_isParserAlias(x_17, x_24); -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_unbox(x_27); -lean_dec(x_27); -if (x_28 == 0) +x_36 = l_panic___at_Lean_Elab_Term_toParserDescr_processAlias___spec__2(x_35, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_25); +if (lean_obj_tag(x_36) == 0) { -lean_object* x_29; 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_dec(x_25); -lean_dec(x_1); -x_29 = lean_ctor_get(x_26, 1); -lean_inc(x_29); -lean_dec(x_26); -x_30 = lean_alloc_ctor(4, 1, 0); -lean_ctor_set(x_30, 0, x_17); -x_31 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__2; -x_32 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_30); -x_33 = l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__4; -x_34 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_34, 0, x_32); -lean_ctor_set(x_34, 1, x_33); -x_35 = l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_34, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_29); +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +lean_dec(x_36); +x_39 = l_Lean_Elab_Term_toParserDescr_processAlias___lambda__1(x_27, x_37, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_38); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -7360,571 +8444,360 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -return x_35; +return x_39; } else { -lean_object* x_36; lean_object* x_37; -x_36 = lean_ctor_get(x_26, 1); -lean_inc(x_36); -lean_dec(x_26); -lean_inc(x_9); -lean_inc(x_8); -x_37 = l_Lean_Elab_Term_toParserDescr_ensureNoPrec(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_36); -lean_dec(x_1); -if (lean_obj_tag(x_37) == 0) -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_38 = lean_ctor_get(x_37, 1); -lean_inc(x_38); -lean_dec(x_37); -x_39 = lean_st_ref_get(x_9, x_38); -x_40 = lean_ctor_get(x_39, 1); -lean_inc(x_40); -lean_dec(x_39); -lean_inc(x_17); -x_41 = l_Lean_Parser_ensureConstantParserAlias(x_17, x_40); -if (lean_obj_tag(x_41) == 0) -{ -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; -lean_dec(x_25); -x_42 = lean_ctor_get(x_41, 1); -lean_inc(x_42); -lean_dec(x_41); -lean_inc(x_8); -x_43 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_toParserDescr_processNonReserved___spec__2___rarg(x_8, x_9, x_42); -x_44 = lean_ctor_get(x_43, 0); -lean_inc(x_44); -x_45 = lean_ctor_get(x_43, 1); -lean_inc(x_45); -lean_dec(x_43); -x_46 = lean_ctor_get(x_8, 10); -lean_inc(x_46); -lean_dec(x_8); -x_47 = lean_st_ref_get(x_9, x_45); +uint8_t x_40; +lean_dec(x_27); +lean_dec(x_10); lean_dec(x_9); -x_48 = !lean_is_exclusive(x_47); -if (x_48 == 0) +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_40 = !lean_is_exclusive(x_36); +if (x_40 == 0) { -lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_49 = lean_ctor_get(x_47, 0); -x_50 = lean_ctor_get(x_49, 0); -lean_inc(x_50); -lean_dec(x_49); -x_51 = lean_environment_main_module(x_50); -x_52 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__7; -x_53 = l_Lean_addMacroScope(x_51, x_52, x_46); -x_54 = lean_box(0); -x_55 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__5; -x_56 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__10; -x_57 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_57, 0, x_44); -lean_ctor_set(x_57, 1, x_55); -lean_ctor_set(x_57, 2, x_53); -lean_ctor_set(x_57, 3, x_56); -lean_inc(x_17); -x_58 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_54, x_17); -x_59 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; -x_60 = lean_array_push(x_59, x_57); -if (lean_obj_tag(x_58) == 0) -{ -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_61 = l___private_Init_Meta_0__Lean_quoteNameMk(x_17); -x_62 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; -x_63 = lean_array_push(x_62, x_61); -x_64 = lean_box(2); -x_65 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_66 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_66, 0, x_64); -lean_ctor_set(x_66, 1, x_65); -lean_ctor_set(x_66, 2, x_63); -x_67 = lean_array_push(x_60, x_66); -x_68 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_69 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_69, 0, x_64); -lean_ctor_set(x_69, 1, x_68); -lean_ctor_set(x_69, 2, x_67); -lean_ctor_set(x_47, 0, x_69); -return x_47; +return x_36; } else { -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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; -lean_dec(x_17); -x_70 = lean_ctor_get(x_58, 0); -lean_inc(x_70); -lean_dec(x_58); -x_71 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__8; -x_72 = l_String_intercalate(x_71, x_70); -x_73 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__9; -x_74 = lean_string_append(x_73, x_72); -lean_dec(x_72); -x_75 = l_Lean_nameLitKind; -x_76 = lean_box(2); -x_77 = l_Lean_Syntax_mkLit(x_75, x_74, x_76); -x_78 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; -x_79 = lean_array_push(x_78, x_77); -x_80 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__23; -x_81 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_81, 0, x_76); -lean_ctor_set(x_81, 1, x_80); -lean_ctor_set(x_81, 2, x_79); -x_82 = lean_array_push(x_78, x_81); -x_83 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_84 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_84, 0, x_76); -lean_ctor_set(x_84, 1, x_83); -lean_ctor_set(x_84, 2, x_82); -x_85 = lean_array_push(x_60, x_84); -x_86 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_87 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_87, 0, x_76); -lean_ctor_set(x_87, 1, x_86); -lean_ctor_set(x_87, 2, x_85); -lean_ctor_set(x_47, 0, x_87); -return x_47; +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_36, 0); +x_42 = lean_ctor_get(x_36, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_36); +x_43 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +return x_43; } } -else -{ -lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_88 = lean_ctor_get(x_47, 0); -x_89 = lean_ctor_get(x_47, 1); -lean_inc(x_89); -lean_inc(x_88); -lean_dec(x_47); -x_90 = lean_ctor_get(x_88, 0); -lean_inc(x_90); -lean_dec(x_88); -x_91 = lean_environment_main_module(x_90); -x_92 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__7; -x_93 = l_Lean_addMacroScope(x_91, x_92, x_46); -x_94 = lean_box(0); -x_95 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__5; -x_96 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__10; -x_97 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_97, 0, x_44); -lean_ctor_set(x_97, 1, x_95); -lean_ctor_set(x_97, 2, x_93); -lean_ctor_set(x_97, 3, x_96); -lean_inc(x_17); -x_98 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_94, x_17); -x_99 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; -x_100 = lean_array_push(x_99, x_97); -if (lean_obj_tag(x_98) == 0) -{ -lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_101 = l___private_Init_Meta_0__Lean_quoteNameMk(x_17); -x_102 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; -x_103 = lean_array_push(x_102, x_101); -x_104 = lean_box(2); -x_105 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_106 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_106, 0, x_104); -lean_ctor_set(x_106, 1, x_105); -lean_ctor_set(x_106, 2, x_103); -x_107 = lean_array_push(x_100, x_106); -x_108 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_109 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_109, 0, x_104); -lean_ctor_set(x_109, 1, x_108); -lean_ctor_set(x_109, 2, x_107); -x_110 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_110, 0, x_109); -lean_ctor_set(x_110, 1, x_89); -return x_110; } else { -lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; -lean_dec(x_17); -x_111 = lean_ctor_get(x_98, 0); -lean_inc(x_111); -lean_dec(x_98); -x_112 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__8; -x_113 = l_String_intercalate(x_112, x_111); -x_114 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__9; -x_115 = lean_string_append(x_114, x_113); -lean_dec(x_113); -x_116 = l_Lean_nameLitKind; -x_117 = lean_box(2); -x_118 = l_Lean_Syntax_mkLit(x_116, x_115, x_117); -x_119 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; -x_120 = lean_array_push(x_119, x_118); -x_121 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__23; -x_122 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_122, 0, x_117); -lean_ctor_set(x_122, 1, x_121); -lean_ctor_set(x_122, 2, x_120); -x_123 = lean_array_push(x_119, x_122); -x_124 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_125 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_125, 0, x_117); -lean_ctor_set(x_125, 1, x_124); -lean_ctor_set(x_125, 2, x_123); -x_126 = lean_array_push(x_100, x_125); -x_127 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_128 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_128, 0, x_117); -lean_ctor_set(x_128, 1, x_127); -lean_ctor_set(x_128, 2, x_126); -x_129 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_129, 0, x_128); -lean_ctor_set(x_129, 1, x_89); -return x_129; -} -} -} -else +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_44 = lean_array_fget(x_26, x_29); +x_45 = lean_array_fget(x_26, x_31); +lean_dec(x_26); +x_46 = lean_st_ref_get(x_10, x_25); +x_47 = lean_ctor_get(x_46, 1); +lean_inc(x_47); +lean_dec(x_46); +lean_inc(x_13); +x_48 = l_Lean_Parser_ensureBinaryParserAlias(x_13, x_47); +if (lean_obj_tag(x_48) == 0) { -uint8_t x_130; -lean_dec(x_17); +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; lean_object* x_67; +lean_dec(x_16); +x_49 = lean_ctor_get(x_48, 1); +lean_inc(x_49); +lean_dec(x_48); +lean_inc(x_9); +x_50 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_toParserDescr_processNonReserved___spec__2___rarg(x_9, x_10, x_49); +x_51 = lean_ctor_get(x_50, 0); +lean_inc(x_51); +x_52 = lean_ctor_get(x_50, 1); +lean_inc(x_52); +lean_dec(x_50); +x_53 = lean_ctor_get(x_9, 10); +lean_inc(x_53); +x_54 = lean_st_ref_get(x_10, x_52); +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_54, 1); +lean_inc(x_56); +lean_dec(x_54); +x_57 = lean_ctor_get(x_55, 0); +lean_inc(x_57); +lean_dec(x_55); +x_58 = lean_environment_main_module(x_57); +x_59 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__15; +x_60 = l_Lean_addMacroScope(x_58, x_59, x_53); +x_61 = lean_box(0); +x_62 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__11; +x_63 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__19; +x_64 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_64, 0, x_51); +lean_ctor_set(x_64, 1, x_62); +lean_ctor_set(x_64, 2, x_60); +lean_ctor_set(x_64, 3, x_63); +lean_inc(x_13); +x_65 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_61, x_13); +x_66 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; +x_67 = lean_array_push(x_66, x_64); +if (lean_obj_tag(x_65) == 0) +{ +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; lean_object* x_79; +x_68 = l_Lean_quoteNameMk(x_13); +x_69 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__28; +x_70 = lean_array_push(x_69, x_68); +x_71 = lean_array_push(x_70, x_44); +x_72 = lean_array_push(x_71, x_45); +x_73 = lean_box(2); +x_74 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; +x_75 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set(x_75, 1, x_74); +lean_ctor_set(x_75, 2, x_72); +x_76 = lean_array_push(x_67, x_75); +x_77 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +x_78 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_78, 0, x_73); +lean_ctor_set(x_78, 1, x_77); +lean_ctor_set(x_78, 2, x_76); +x_79 = l_Lean_Elab_Term_toParserDescr_processAlias___lambda__1(x_27, x_78, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_56); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -x_130 = !lean_is_exclusive(x_41); -if (x_130 == 0) -{ -lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; -x_131 = lean_ctor_get(x_41, 0); -x_132 = lean_io_error_to_string(x_131); -x_133 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_133, 0, x_132); -x_134 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_134, 0, x_133); -x_135 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_135, 0, x_25); -lean_ctor_set(x_135, 1, x_134); -lean_ctor_set(x_41, 0, x_135); -return x_41; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_79; } else { -lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; -x_136 = lean_ctor_get(x_41, 0); -x_137 = lean_ctor_get(x_41, 1); -lean_inc(x_137); -lean_inc(x_136); -lean_dec(x_41); -x_138 = lean_io_error_to_string(x_136); -x_139 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_139, 0, x_138); -x_140 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_140, 0, x_139); -x_141 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_141, 0, x_25); -lean_ctor_set(x_141, 1, x_140); -x_142 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_142, 0, x_141); -lean_ctor_set(x_142, 1, x_137); -return x_142; -} +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; +lean_dec(x_13); +x_80 = lean_ctor_get(x_65, 0); +lean_inc(x_80); +lean_dec(x_65); +x_81 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__16; +x_82 = l_String_intercalate(x_81, x_80); +x_83 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__17; +x_84 = lean_string_append(x_83, x_82); +lean_dec(x_82); +x_85 = lean_box(2); +x_86 = l_Lean_Syntax_mkNameLit(x_84, x_85); +x_87 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; +x_88 = lean_array_push(x_87, x_86); +x_89 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__23; +x_90 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_90, 0, x_85); +lean_ctor_set(x_90, 1, x_89); +lean_ctor_set(x_90, 2, x_88); +x_91 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__28; +x_92 = lean_array_push(x_91, x_90); +x_93 = lean_array_push(x_92, x_44); +x_94 = lean_array_push(x_93, x_45); +x_95 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; +x_96 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_96, 0, x_85); +lean_ctor_set(x_96, 1, x_95); +lean_ctor_set(x_96, 2, x_94); +x_97 = lean_array_push(x_67, x_96); +x_98 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +x_99 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_99, 0, x_85); +lean_ctor_set(x_99, 1, x_98); +lean_ctor_set(x_99, 2, x_97); +x_100 = l_Lean_Elab_Term_toParserDescr_processAlias___lambda__1(x_27, x_99, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_56); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_100; } } else { -uint8_t x_143; -lean_dec(x_25); -lean_dec(x_17); +uint8_t x_101; +lean_dec(x_45); +lean_dec(x_44); +lean_dec(x_27); +lean_dec(x_13); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -x_143 = !lean_is_exclusive(x_37); -if (x_143 == 0) -{ -return x_37; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_101 = !lean_is_exclusive(x_48); +if (x_101 == 0) +{ +lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_102 = lean_ctor_get(x_48, 0); +x_103 = lean_io_error_to_string(x_102); +x_104 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_104, 0, x_103); +x_105 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_105, 0, x_104); +x_106 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_106, 0, x_16); +lean_ctor_set(x_106, 1, x_105); +lean_ctor_set(x_48, 0, x_106); +return x_48; } else { -lean_object* x_144; lean_object* x_145; lean_object* x_146; -x_144 = lean_ctor_get(x_37, 0); -x_145 = lean_ctor_get(x_37, 1); -lean_inc(x_145); -lean_inc(x_144); -lean_dec(x_37); -x_146 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_146, 0, x_144); -lean_ctor_set(x_146, 1, x_145); -return x_146; -} -} +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; +x_107 = lean_ctor_get(x_48, 0); +x_108 = lean_ctor_get(x_48, 1); +lean_inc(x_108); +lean_inc(x_107); +lean_dec(x_48); +x_109 = lean_io_error_to_string(x_107); +x_110 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_110, 0, x_109); +x_111 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_111, 0, x_110); +x_112 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_112, 0, x_16); +lean_ctor_set(x_112, 1, x_111); +x_113 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_113, 0, x_112); +lean_ctor_set(x_113, 1, x_108); +return x_113; } } -else -{ -lean_object* x_147; -lean_dec(x_17); -x_147 = l_Lean_Elab_Term_toParserDescr_processParserCategory(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_20); -return x_147; } } else { -lean_object* x_148; lean_object* x_149; uint8_t x_150; -lean_dec(x_12); -x_148 = lean_ctor_get(x_14, 0); -lean_inc(x_148); -x_149 = lean_ctor_get(x_148, 1); -lean_inc(x_149); -x_150 = lean_unbox(x_149); -lean_dec(x_149); -if (x_150 == 0) -{ -lean_object* x_151; -x_151 = lean_ctor_get(x_14, 1); -lean_inc(x_151); -if (lean_obj_tag(x_151) == 0) +lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; +lean_dec(x_28); +x_114 = lean_array_fget(x_26, x_29); +lean_dec(x_26); +x_115 = lean_st_ref_get(x_10, x_25); +x_116 = lean_ctor_get(x_115, 1); +lean_inc(x_116); +lean_dec(x_115); +lean_inc(x_13); +x_117 = l_Lean_Parser_ensureUnaryParserAlias(x_13, x_116); +if (lean_obj_tag(x_117) == 0) { -lean_object* x_152; lean_object* x_153; lean_object* x_154; -lean_dec(x_14); -x_152 = lean_ctor_get(x_13, 1); -lean_inc(x_152); -lean_dec(x_13); -x_153 = lean_ctor_get(x_148, 0); -lean_inc(x_153); -lean_dec(x_148); +lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; +lean_dec(x_16); +x_118 = lean_ctor_get(x_117, 1); +lean_inc(x_118); +lean_dec(x_117); lean_inc(x_9); -lean_inc(x_8); -x_154 = l_Lean_Elab_Term_toParserDescr_ensureNoPrec(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_152); -lean_dec(x_1); -if (lean_obj_tag(x_154) == 0) -{ -lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; uint8_t x_161; -x_155 = lean_ctor_get(x_154, 1); -lean_inc(x_155); -lean_dec(x_154); -lean_inc(x_8); -x_156 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_toParserDescr_processNonReserved___spec__2___rarg(x_8, x_9, x_155); -x_157 = lean_ctor_get(x_156, 0); -lean_inc(x_157); -x_158 = lean_ctor_get(x_156, 1); -lean_inc(x_158); -lean_dec(x_156); -x_159 = lean_ctor_get(x_8, 10); -lean_inc(x_159); -lean_dec(x_8); -x_160 = lean_st_ref_get(x_9, x_158); +x_119 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_toParserDescr_processNonReserved___spec__2___rarg(x_9, x_10, x_118); +x_120 = lean_ctor_get(x_119, 0); +lean_inc(x_120); +x_121 = lean_ctor_get(x_119, 1); +lean_inc(x_121); +lean_dec(x_119); +x_122 = lean_ctor_get(x_9, 10); +lean_inc(x_122); +x_123 = lean_st_ref_get(x_10, x_121); +x_124 = lean_ctor_get(x_123, 0); +lean_inc(x_124); +x_125 = lean_ctor_get(x_123, 1); +lean_inc(x_125); +lean_dec(x_123); +x_126 = lean_ctor_get(x_124, 0); +lean_inc(x_126); +lean_dec(x_124); +x_127 = lean_environment_main_module(x_126); +x_128 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__7; +x_129 = l_Lean_addMacroScope(x_127, x_128, x_122); +x_130 = lean_box(0); +x_131 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__5; +x_132 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__10; +x_133 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_133, 0, x_120); +lean_ctor_set(x_133, 1, x_131); +lean_ctor_set(x_133, 2, x_129); +lean_ctor_set(x_133, 3, x_132); +lean_inc(x_13); +x_134 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_130, x_13); +x_135 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; +x_136 = lean_array_push(x_135, x_133); +if (lean_obj_tag(x_134) == 0) +{ +lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; +x_137 = l_Lean_quoteNameMk(x_13); +x_138 = lean_array_push(x_135, x_137); +x_139 = lean_array_push(x_138, x_114); +x_140 = lean_box(2); +x_141 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; +x_142 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_142, 0, x_140); +lean_ctor_set(x_142, 1, x_141); +lean_ctor_set(x_142, 2, x_139); +x_143 = lean_array_push(x_136, x_142); +x_144 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +x_145 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_145, 0, x_140); +lean_ctor_set(x_145, 1, x_144); +lean_ctor_set(x_145, 2, x_143); +x_146 = l_Lean_Elab_Term_toParserDescr_processAlias___lambda__1(x_27, x_145, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_125); +lean_dec(x_10); lean_dec(x_9); -x_161 = !lean_is_exclusive(x_160); -if (x_161 == 0) -{ -lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; -x_162 = lean_ctor_get(x_160, 0); -x_163 = lean_ctor_get(x_162, 0); -lean_inc(x_163); -lean_dec(x_162); -x_164 = lean_environment_main_module(x_163); -x_165 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__15; -x_166 = l_Lean_addMacroScope(x_164, x_165, x_159); -x_167 = lean_box(0); -x_168 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__13; -x_169 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__18; -x_170 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_170, 0, x_157); -lean_ctor_set(x_170, 1, x_168); -lean_ctor_set(x_170, 2, x_166); -lean_ctor_set(x_170, 3, x_169); -lean_inc(x_153); -x_171 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_167, x_153); -x_172 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; -x_173 = lean_array_push(x_172, x_170); -if (lean_obj_tag(x_171) == 0) -{ -lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; -x_174 = l___private_Init_Meta_0__Lean_quoteNameMk(x_153); -x_175 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; -x_176 = lean_array_push(x_175, x_174); -x_177 = lean_box(2); -x_178 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_179 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_179, 0, x_177); -lean_ctor_set(x_179, 1, x_178); -lean_ctor_set(x_179, 2, x_176); -x_180 = lean_array_push(x_173, x_179); -x_181 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_182 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_182, 0, x_177); -lean_ctor_set(x_182, 1, x_181); -lean_ctor_set(x_182, 2, x_180); -lean_ctor_set(x_160, 0, x_182); -return x_160; -} -else -{ -lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; -lean_dec(x_153); -x_183 = lean_ctor_get(x_171, 0); -lean_inc(x_183); -lean_dec(x_171); -x_184 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__8; -x_185 = l_String_intercalate(x_184, x_183); -x_186 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__9; -x_187 = lean_string_append(x_186, x_185); -lean_dec(x_185); -x_188 = l_Lean_nameLitKind; -x_189 = lean_box(2); -x_190 = l_Lean_Syntax_mkLit(x_188, x_187, x_189); -x_191 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; -x_192 = lean_array_push(x_191, x_190); -x_193 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__23; -x_194 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_194, 0, x_189); -lean_ctor_set(x_194, 1, x_193); -lean_ctor_set(x_194, 2, x_192); -x_195 = lean_array_push(x_191, x_194); -x_196 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_197 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_197, 0, x_189); -lean_ctor_set(x_197, 1, x_196); -lean_ctor_set(x_197, 2, x_195); -x_198 = lean_array_push(x_173, x_197); -x_199 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_200 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_200, 0, x_189); -lean_ctor_set(x_200, 1, x_199); -lean_ctor_set(x_200, 2, x_198); -lean_ctor_set(x_160, 0, x_200); -return x_160; -} -} -else -{ -lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; -x_201 = lean_ctor_get(x_160, 0); -x_202 = lean_ctor_get(x_160, 1); -lean_inc(x_202); -lean_inc(x_201); -lean_dec(x_160); -x_203 = lean_ctor_get(x_201, 0); -lean_inc(x_203); -lean_dec(x_201); -x_204 = lean_environment_main_module(x_203); -x_205 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__15; -x_206 = l_Lean_addMacroScope(x_204, x_205, x_159); -x_207 = lean_box(0); -x_208 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__13; -x_209 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__18; -x_210 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_210, 0, x_157); -lean_ctor_set(x_210, 1, x_208); -lean_ctor_set(x_210, 2, x_206); -lean_ctor_set(x_210, 3, x_209); -lean_inc(x_153); -x_211 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_207, x_153); -x_212 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; -x_213 = lean_array_push(x_212, x_210); -if (lean_obj_tag(x_211) == 0) -{ -lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; -x_214 = l___private_Init_Meta_0__Lean_quoteNameMk(x_153); -x_215 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; -x_216 = lean_array_push(x_215, x_214); -x_217 = lean_box(2); -x_218 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_219 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_219, 0, x_217); -lean_ctor_set(x_219, 1, x_218); -lean_ctor_set(x_219, 2, x_216); -x_220 = lean_array_push(x_213, x_219); -x_221 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_222 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_222, 0, x_217); -lean_ctor_set(x_222, 1, x_221); -lean_ctor_set(x_222, 2, x_220); -x_223 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_223, 0, x_222); -lean_ctor_set(x_223, 1, x_202); -return x_223; -} -else -{ -lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; -lean_dec(x_153); -x_224 = lean_ctor_get(x_211, 0); -lean_inc(x_224); -lean_dec(x_211); -x_225 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__8; -x_226 = l_String_intercalate(x_225, x_224); -x_227 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__9; -x_228 = lean_string_append(x_227, x_226); -lean_dec(x_226); -x_229 = l_Lean_nameLitKind; -x_230 = lean_box(2); -x_231 = l_Lean_Syntax_mkLit(x_229, x_228, x_230); -x_232 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; -x_233 = lean_array_push(x_232, x_231); -x_234 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__23; -x_235 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_235, 0, x_230); -lean_ctor_set(x_235, 1, x_234); -lean_ctor_set(x_235, 2, x_233); -x_236 = lean_array_push(x_232, x_235); -x_237 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_238 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_238, 0, x_230); -lean_ctor_set(x_238, 1, x_237); -lean_ctor_set(x_238, 2, x_236); -x_239 = lean_array_push(x_213, x_238); -x_240 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_241 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_241, 0, x_230); -lean_ctor_set(x_241, 1, x_240); -lean_ctor_set(x_241, 2, x_239); -x_242 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_242, 0, x_241); -lean_ctor_set(x_242, 1, x_202); -return x_242; -} -} +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_146; } else { -uint8_t x_243; -lean_dec(x_153); +lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; +lean_dec(x_13); +x_147 = lean_ctor_get(x_134, 0); +lean_inc(x_147); +lean_dec(x_134); +x_148 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__16; +x_149 = l_String_intercalate(x_148, x_147); +x_150 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__17; +x_151 = lean_string_append(x_150, x_149); +lean_dec(x_149); +x_152 = lean_box(2); +x_153 = l_Lean_Syntax_mkNameLit(x_151, x_152); +x_154 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; +x_155 = lean_array_push(x_154, x_153); +x_156 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__23; +x_157 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_157, 0, x_152); +lean_ctor_set(x_157, 1, x_156); +lean_ctor_set(x_157, 2, x_155); +x_158 = lean_array_push(x_135, x_157); +x_159 = lean_array_push(x_158, x_114); +x_160 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; +x_161 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_161, 0, x_152); +lean_ctor_set(x_161, 1, x_160); +lean_ctor_set(x_161, 2, x_159); +x_162 = lean_array_push(x_136, x_161); +x_163 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +x_164 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_164, 0, x_152); +lean_ctor_set(x_164, 1, x_163); +lean_ctor_set(x_164, 2, x_162); +x_165 = l_Lean_Elab_Term_toParserDescr_processAlias___lambda__1(x_27, x_164, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_125); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -x_243 = !lean_is_exclusive(x_154); -if (x_243 == 0) -{ -return x_154; -} -else -{ -lean_object* x_244; lean_object* x_245; lean_object* x_246; -x_244 = lean_ctor_get(x_154, 0); -x_245 = lean_ctor_get(x_154, 1); -lean_inc(x_245); -lean_inc(x_244); -lean_dec(x_154); -x_246 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_246, 0, x_244); -lean_ctor_set(x_246, 1, x_245); -return x_246; -} +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_165; } } else { -lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; -lean_dec(x_151); -lean_dec(x_148); -lean_dec(x_1); -x_247 = lean_ctor_get(x_13, 1); -lean_inc(x_247); +uint8_t x_166; +lean_dec(x_114); +lean_dec(x_27); lean_dec(x_13); -x_248 = lean_box(0); -x_249 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_14, x_248); -x_250 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_249, x_248); -x_251 = l_Lean_MessageData_ofList(x_250); -lean_dec(x_250); -x_252 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__20; -x_253 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_253, 0, x_252); -lean_ctor_set(x_253, 1, x_251); -x_254 = l_Lean_addTrace___at_Lean_Elab_Term_checkLeftRec___spec__3___closed__8; -x_255 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_255, 0, x_253); -lean_ctor_set(x_255, 1, x_254); -x_256 = l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_255, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_247); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -7932,100 +8805,117 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -return x_256; -} +x_166 = !lean_is_exclusive(x_117); +if (x_166 == 0) +{ +lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; +x_167 = lean_ctor_get(x_117, 0); +x_168 = lean_io_error_to_string(x_167); +x_169 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_169, 0, x_168); +x_170 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_170, 0, x_169); +x_171 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_171, 0, x_16); +lean_ctor_set(x_171, 1, x_170); +lean_ctor_set(x_117, 0, x_171); +return x_117; } else { -lean_object* x_257; -x_257 = lean_ctor_get(x_14, 1); -lean_inc(x_257); -if (lean_obj_tag(x_257) == 0) -{ -lean_object* x_258; lean_object* x_259; lean_object* x_260; -lean_dec(x_14); -x_258 = lean_ctor_get(x_13, 1); -lean_inc(x_258); -lean_dec(x_13); -x_259 = lean_ctor_get(x_148, 0); -lean_inc(x_259); -lean_dec(x_148); -x_260 = l_Lean_Elab_Term_toParserDescr_ensureNoPrec(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_258); -if (lean_obj_tag(x_260) == 0) -{ -uint8_t x_261; -x_261 = !lean_is_exclusive(x_260); -if (x_261 == 0) -{ -lean_object* x_262; lean_object* x_263; -x_262 = lean_ctor_get(x_260, 0); -lean_dec(x_262); -x_263 = l_Lean_mkIdentFrom(x_1, x_259); -lean_ctor_set(x_260, 0, x_263); -return x_260; +lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; +x_172 = lean_ctor_get(x_117, 0); +x_173 = lean_ctor_get(x_117, 1); +lean_inc(x_173); +lean_inc(x_172); +lean_dec(x_117); +x_174 = lean_io_error_to_string(x_172); +x_175 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_175, 0, x_174); +x_176 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_176, 0, x_175); +x_177 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_177, 0, x_16); +lean_ctor_set(x_177, 1, x_176); +x_178 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_178, 0, x_177); +lean_ctor_set(x_178, 1, x_173); +return x_178; } -else -{ -lean_object* x_264; lean_object* x_265; lean_object* x_266; -x_264 = lean_ctor_get(x_260, 1); -lean_inc(x_264); -lean_dec(x_260); -x_265 = l_Lean_mkIdentFrom(x_1, x_259); -x_266 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_266, 0, x_265); -lean_ctor_set(x_266, 1, x_264); -return x_266; } } -else -{ -uint8_t x_267; -lean_dec(x_259); -lean_dec(x_1); -x_267 = !lean_is_exclusive(x_260); -if (x_267 == 0) -{ -return x_260; } else { -lean_object* x_268; lean_object* x_269; lean_object* x_270; -x_268 = lean_ctor_get(x_260, 0); -x_269 = lean_ctor_get(x_260, 1); -lean_inc(x_269); -lean_inc(x_268); -lean_dec(x_260); -x_270 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_270, 0, x_268); -lean_ctor_set(x_270, 1, x_269); -return x_270; -} -} -} -else +lean_object* x_179; lean_object* x_180; lean_object* x_181; +lean_dec(x_28); +lean_dec(x_26); +x_179 = lean_st_ref_get(x_10, x_25); +x_180 = lean_ctor_get(x_179, 1); +lean_inc(x_180); +lean_dec(x_179); +lean_inc(x_13); +x_181 = l_Lean_Parser_ensureConstantParserAlias(x_13, x_180); +if (lean_obj_tag(x_181) == 0) { -lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; -lean_dec(x_257); -lean_dec(x_148); -lean_dec(x_1); -x_271 = lean_ctor_get(x_13, 1); -lean_inc(x_271); -lean_dec(x_13); -x_272 = lean_box(0); -x_273 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__17(x_14, x_272); -x_274 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_273, x_272); -x_275 = l_Lean_MessageData_ofList(x_274); -lean_dec(x_274); -x_276 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__20; -x_277 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_277, 0, x_276); -lean_ctor_set(x_277, 1, x_275); -x_278 = l_Lean_addTrace___at_Lean_Elab_Term_checkLeftRec___spec__3___closed__8; -x_279 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_279, 0, x_277); -lean_ctor_set(x_279, 1, x_278); -x_280 = l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_279, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_271); +lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; +lean_dec(x_16); +x_182 = lean_ctor_get(x_181, 1); +lean_inc(x_182); +lean_dec(x_181); +lean_inc(x_9); +x_183 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_toParserDescr_processNonReserved___spec__2___rarg(x_9, x_10, x_182); +x_184 = lean_ctor_get(x_183, 0); +lean_inc(x_184); +x_185 = lean_ctor_get(x_183, 1); +lean_inc(x_185); +lean_dec(x_183); +x_186 = lean_ctor_get(x_9, 10); +lean_inc(x_186); +x_187 = lean_st_ref_get(x_10, x_185); +x_188 = lean_ctor_get(x_187, 0); +lean_inc(x_188); +x_189 = lean_ctor_get(x_187, 1); +lean_inc(x_189); +lean_dec(x_187); +x_190 = lean_ctor_get(x_188, 0); +lean_inc(x_190); +lean_dec(x_188); +x_191 = lean_environment_main_module(x_190); +x_192 = l_Lean_Elab_Term_toParserDescr_processAlias___closed__9; +x_193 = l_Lean_addMacroScope(x_191, x_192, x_186); +x_194 = lean_box(0); +x_195 = l_Lean_Elab_Term_toParserDescr_processAlias___closed__7; +x_196 = l_Lean_Elab_Term_toParserDescr_processAlias___closed__12; +x_197 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_197, 0, x_184); +lean_ctor_set(x_197, 1, x_195); +lean_ctor_set(x_197, 2, x_193); +lean_ctor_set(x_197, 3, x_196); +lean_inc(x_13); +x_198 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_194, x_13); +x_199 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; +x_200 = lean_array_push(x_199, x_197); +if (lean_obj_tag(x_198) == 0) +{ +lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; +x_201 = l_Lean_quoteNameMk(x_13); +x_202 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; +x_203 = lean_array_push(x_202, x_201); +x_204 = lean_box(2); +x_205 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; +x_206 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_206, 0, x_204); +lean_ctor_set(x_206, 1, x_205); +lean_ctor_set(x_206, 2, x_203); +x_207 = lean_array_push(x_200, x_206); +x_208 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +x_209 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_209, 0, x_204); +lean_ctor_set(x_209, 1, x_208); +lean_ctor_set(x_209, 2, x_207); +x_210 = l_Lean_Elab_Term_toParserDescr_processAlias___lambda__1(x_27, x_209, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_189); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -8033,18 +8923,43 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -return x_280; +return x_210; } -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__5___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_object* x_9, lean_object* x_10) { -_start: +else { -lean_object* x_11; -x_11 = l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; +lean_dec(x_13); +x_211 = lean_ctor_get(x_198, 0); +lean_inc(x_211); +lean_dec(x_198); +x_212 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__16; +x_213 = l_String_intercalate(x_212, x_211); +x_214 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__17; +x_215 = lean_string_append(x_214, x_213); +lean_dec(x_213); +x_216 = lean_box(2); +x_217 = l_Lean_Syntax_mkNameLit(x_215, x_216); +x_218 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; +x_219 = lean_array_push(x_218, x_217); +x_220 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__23; +x_221 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_221, 0, x_216); +lean_ctor_set(x_221, 1, x_220); +lean_ctor_set(x_221, 2, x_219); +x_222 = lean_array_push(x_218, x_221); +x_223 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; +x_224 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_224, 0, x_216); +lean_ctor_set(x_224, 1, x_223); +lean_ctor_set(x_224, 2, x_222); +x_225 = lean_array_push(x_200, x_224); +x_226 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +x_227 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_227, 0, x_216); +lean_ctor_set(x_227, 1, x_226); +lean_ctor_set(x_227, 2, x_225); +x_228 = l_Lean_Elab_Term_toParserDescr_processAlias___lambda__1(x_27, x_227, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_189); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -8052,47 +8967,69 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -return x_11; +return x_228; } } -LEAN_EXPORT lean_object* l_Lean_resolveGlobalName___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__7___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_object* x_9, lean_object* x_10) { -_start: +else { -lean_object* x_11; -x_11 = l_Lean_resolveGlobalName___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +uint8_t x_229; +lean_dec(x_27); +lean_dec(x_13); +lean_dec(x_10); lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -return x_11; +x_229 = !lean_is_exclusive(x_181); +if (x_229 == 0) +{ +lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; +x_230 = lean_ctor_get(x_181, 0); +x_231 = lean_io_error_to_string(x_230); +x_232 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_232, 0, x_231); +x_233 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_233, 0, x_232); +x_234 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_234, 0, x_16); +lean_ctor_set(x_234, 1, x_233); +lean_ctor_set(x_181, 0, x_234); +return x_181; } +else +{ +lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; +x_235 = lean_ctor_get(x_181, 0); +x_236 = lean_ctor_get(x_181, 1); +lean_inc(x_236); +lean_inc(x_235); +lean_dec(x_181); +x_237 = lean_io_error_to_string(x_235); +x_238 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_238, 0, x_237); +x_239 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_239, 0, x_238); +x_240 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_240, 0, x_16); +lean_ctor_set(x_240, 1, x_239); +x_241 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_241, 0, x_240); +lean_ctor_set(x_241, 1, x_236); +return x_241; } -LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___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_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_11; } } -LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__6___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, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: +} +} +else { -lean_object* x_13; -x_13 = l_Lean_resolveGlobalConstCore___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__6___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); +uint8_t x_262; +lean_dec(x_18); +lean_dec(x_16); +lean_dec(x_13); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -8101,46 +9038,80 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_13; -} +x_262 = !lean_is_exclusive(x_23); +if (x_262 == 0) +{ +return x_23; } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__11___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_object* x_9, lean_object* x_10) { -_start: +else { -lean_object* x_11; -x_11 = l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__11(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_11; +lean_object* x_263; lean_object* x_264; lean_object* x_265; +x_263 = lean_ctor_get(x_23, 0); +x_264 = lean_ctor_get(x_23, 1); +lean_inc(x_264); +lean_inc(x_263); +lean_dec(x_23); +x_265 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_265, 0, x_263); +lean_ctor_set(x_265, 1, x_264); +return x_265; } } -LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__10___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_object* x_9, lean_object* x_10) { +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__5(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_object* x_9, lean_object* x_10) { _start: { -lean_object* x_11; -x_11 = l_Lean_getConstInfo___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__10(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_11; +lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_11 = lean_ctor_get(x_8, 5); +x_12 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_6, x_7, x_8, x_9, x_10); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_12, 0); +lean_inc(x_11); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_11); +lean_ctor_set(x_15, 1, x_14); +lean_ctor_set_tag(x_12, 1); +lean_ctor_set(x_12, 0, x_15); +return x_12; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_12, 0); +x_17 = lean_ctor_get(x_12, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_12); +lean_inc(x_11); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_11); +lean_ctor_set(x_18, 1, x_16); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +return x_19; } } -LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__9___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_object* x_9, lean_object* x_10) { +} +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__4(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_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_11; -x_11 = l_Lean_mkConstWithLevelParams___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +uint8_t x_12; +x_12 = !lean_is_exclusive(x_9); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_9, 5); +x_14 = l_Lean_replaceRef(x_1, x_13); +lean_dec(x_13); +lean_dec(x_1); +lean_ctor_set(x_9, 5, x_14); +x_15 = l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__5(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -8148,782 +9119,991 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -return x_11; -} +return x_15; } -LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoTree___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__13___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_object* x_9, lean_object* x_10) { -_start: +else { -lean_object* x_11; -x_11 = l_Lean_Elab_pushInfoTree___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__13(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +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; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_16 = lean_ctor_get(x_9, 0); +x_17 = lean_ctor_get(x_9, 1); +x_18 = lean_ctor_get(x_9, 2); +x_19 = lean_ctor_get(x_9, 3); +x_20 = lean_ctor_get(x_9, 4); +x_21 = lean_ctor_get(x_9, 5); +x_22 = lean_ctor_get(x_9, 6); +x_23 = lean_ctor_get(x_9, 7); +x_24 = lean_ctor_get(x_9, 8); +x_25 = lean_ctor_get(x_9, 9); +x_26 = lean_ctor_get(x_9, 10); +lean_inc(x_26); +lean_inc(x_25); +lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); lean_dec(x_9); +x_27 = l_Lean_replaceRef(x_1, x_21); +lean_dec(x_21); +lean_dec(x_1); +x_28 = lean_alloc_ctor(0, 11, 0); +lean_ctor_set(x_28, 0, x_16); +lean_ctor_set(x_28, 1, x_17); +lean_ctor_set(x_28, 2, x_18); +lean_ctor_set(x_28, 3, x_19); +lean_ctor_set(x_28, 4, x_20); +lean_ctor_set(x_28, 5, x_27); +lean_ctor_set(x_28, 6, x_22); +lean_ctor_set(x_28, 7, x_23); +lean_ctor_set(x_28, 8, x_24); +lean_ctor_set(x_28, 9, x_25); +lean_ctor_set(x_28, 10, x_26); +x_29 = l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__5(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_28, x_10, x_11); +lean_dec(x_10); +lean_dec(x_28); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -return x_11; +return x_29; } } -LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___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_object* x_9, lean_object* x_10) { +} +LEAN_EXPORT lean_object* l_Lean_resolveGlobalName___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__7(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_object* x_9, lean_object* x_10) { _start: { -lean_object* x_11; -x_11 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); +lean_object* x_11; uint8_t x_12; +x_11 = lean_st_ref_get(x_9, x_10); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +lean_dec(x_13); +x_15 = lean_ctor_get(x_8, 6); +lean_inc(x_15); +x_16 = lean_ctor_get(x_8, 7); +lean_inc(x_16); lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); +x_17 = l_Lean_ResolveName_resolveGlobalName(x_14, x_15, x_16, x_1); +lean_ctor_set(x_11, 0, x_17); return x_11; } -} -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__14___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: +else { -lean_object* x_14; -x_14 = l_List_forIn_loop___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__14(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_12); +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_18 = lean_ctor_get(x_11, 0); +x_19 = lean_ctor_get(x_11, 1); +lean_inc(x_19); +lean_inc(x_18); lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); +x_20 = lean_ctor_get(x_18, 0); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_ctor_get(x_8, 6); +lean_inc(x_21); +x_22 = lean_ctor_get(x_8, 7); +lean_inc(x_22); lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -return x_14; +x_23 = l_Lean_ResolveName_resolveGlobalName(x_20, x_21, x_22, x_1); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_19); +return x_24; } } -LEAN_EXPORT lean_object* l_Lean_Elab_resolveGlobalConstWithInfos___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__2___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, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +} +static lean_object* _init_l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__1() { _start: { -lean_object* x_12; -x_12 = l_Lean_Elab_resolveGlobalConstWithInfos___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__2___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_12; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("unknown constant '", 18); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16___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_object* x_9, lean_object* x_10) { +static lean_object* _init_l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__2() { _start: { -lean_object* x_11; -x_11 = l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_11; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_process___spec__3(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) { +static lean_object* _init_l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__3() { _start: { -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; uint8_t x_16; -x_9 = lean_ctor_get(x_6, 5); -x_10 = lean_ctor_get(x_2, 1); -lean_inc(x_10); -lean_inc(x_10); -x_11 = l_Lean_Elab_getBetterRef(x_9, x_10); -x_12 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_4, x_5, x_6, x_7, x_8); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec(x_12); -x_15 = l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(x_13, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_14); -lean_dec(x_2); -lean_dec(x_10); -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; -x_17 = lean_ctor_get(x_15, 0); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_11); -lean_ctor_set(x_18, 1, x_17); -lean_ctor_set_tag(x_15, 1); -lean_ctor_set(x_15, 0, x_18); -return x_15; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("'", 1); +return x_1; } -else +} +static lean_object* _init_l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__4() { +_start: { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_19 = lean_ctor_get(x_15, 0); -x_20 = lean_ctor_get(x_15, 1); -lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_15); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_11); -lean_ctor_set(x_21, 1, x_19); -x_22 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_20); -return x_22; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} } +LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8(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_object* x_9, lean_object* x_10) { +_start: +{ +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_11 = lean_box(0); +x_12 = l_Lean_mkConst(x_1, x_11); +x_13 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_13, 0, x_12); +x_14 = l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__2; +x_15 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_13); +x_16 = l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__4; +x_17 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +x_18 = l_Lean_throwError___at_Lean_Elab_Term_checkLeftRec___spec__10(x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_18; } } -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_toParserDescr_process___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_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__6___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, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -uint8_t x_10; -x_10 = !lean_is_exclusive(x_7); -if (x_10 == 0) +lean_object* x_13; lean_object* x_14; +x_13 = l_List_mapTRAux___at_Lean_resolveGlobalConstCore___spec__2(x_1, x_2); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_12); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__6(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_object* x_9, lean_object* x_10) { +_start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_11 = lean_ctor_get(x_7, 5); -x_12 = l_Lean_replaceRef(x_1, x_11); +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +lean_inc(x_8); +lean_inc(x_1); +x_11 = l_Lean_resolveGlobalName___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); lean_dec(x_11); +x_14 = lean_box(0); +x_15 = l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(x_12, x_14); +x_16 = l_List_isEmpty___rarg(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_dec(x_1); -lean_ctor_set(x_7, 5, x_12); -x_13 = l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_process___spec__3(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_17 = lean_box(0); +x_18 = l_Lean_resolveGlobalConstCore___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__6___lambda__1(x_15, x_14, x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -return x_13; +lean_dec(x_3); +lean_dec(x_2); +return x_18; } 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; 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; -x_14 = lean_ctor_get(x_7, 0); -x_15 = lean_ctor_get(x_7, 1); -x_16 = lean_ctor_get(x_7, 2); -x_17 = lean_ctor_get(x_7, 3); -x_18 = lean_ctor_get(x_7, 4); -x_19 = lean_ctor_get(x_7, 5); -x_20 = lean_ctor_get(x_7, 6); -x_21 = lean_ctor_get(x_7, 7); -x_22 = lean_ctor_get(x_7, 8); -x_23 = lean_ctor_get(x_7, 9); -x_24 = lean_ctor_get(x_7, 10); -lean_inc(x_24); -lean_inc(x_23); -lean_inc(x_22); -lean_inc(x_21); -lean_inc(x_20); -lean_inc(x_19); -lean_inc(x_18); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_14); -lean_dec(x_7); -x_25 = l_Lean_replaceRef(x_1, x_19); -lean_dec(x_19); -lean_dec(x_1); -x_26 = lean_alloc_ctor(0, 11, 0); -lean_ctor_set(x_26, 0, x_14); -lean_ctor_set(x_26, 1, x_15); -lean_ctor_set(x_26, 2, x_16); -lean_ctor_set(x_26, 3, x_17); -lean_ctor_set(x_26, 4, x_18); -lean_ctor_set(x_26, 5, x_25); -lean_ctor_set(x_26, 6, x_20); -lean_ctor_set(x_26, 7, x_21); -lean_ctor_set(x_26, 8, x_22); -lean_ctor_set(x_26, 9, x_23); -lean_ctor_set(x_26, 10, x_24); -x_27 = l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_process___spec__3(x_2, x_3, x_4, x_5, x_6, x_26, x_8, x_9); +lean_object* x_19; uint8_t x_20; +lean_dec(x_15); +x_19 = l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +lean_dec(x_9); lean_dec(x_8); -lean_dec(x_26); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -return x_27; +lean_dec(x_3); +lean_dec(x_2); +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +return x_19; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_19, 0); +x_22 = lean_ctor_get(x_19, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_19); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; } } } -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_toParserDescr_process___spec__4(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) { +} +static lean_object* _init_l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3___closed__1() { _start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_9 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_checkLeftRec___spec__7___closed__2; -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_1); -lean_ctor_set(x_10, 1, x_9); -x_11 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_11, 0, x_10); -lean_ctor_set(x_11, 1, x_8); -return x_11; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("expected identifier", 19); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_toParserDescr_process___spec__5___rarg(lean_object* x_1) { +static lean_object* _init_l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3___closed__2() { _start: { -lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_checkLeftRec___spec__8___rarg___closed__2; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3___closed__1; +x_2 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_toParserDescr_process___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +static lean_object* _init_l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3___closed__3() { _start: { -lean_object* x_7; -x_7 = lean_alloc_closure((void*)(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_toParserDescr_process___spec__5___rarg), 1, 0); -return x_7; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3___closed__2; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_toParserDescr_process___spec__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) { +LEAN_EXPORT lean_object* l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3(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_object* x_9, lean_object* x_10) { _start: { -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; 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; 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_object* x_32; lean_object* x_33; -x_9 = lean_st_ref_get(x_7, x_8); -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_9, 1); +if (lean_obj_tag(x_1) == 3) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_ctor_get(x_1, 2); lean_inc(x_11); -lean_dec(x_9); -x_12 = lean_ctor_get(x_10, 0); -lean_inc(x_12); -lean_dec(x_10); -x_13 = lean_ctor_get(x_6, 3); -lean_inc(x_13); -x_14 = lean_ctor_get(x_6, 4); -lean_inc(x_14); -x_15 = lean_ctor_get(x_6, 5); -lean_inc(x_15); -x_16 = lean_ctor_get(x_6, 6); -lean_inc(x_16); -x_17 = lean_ctor_get(x_6, 7); -lean_inc(x_17); -x_18 = lean_ctor_get(x_6, 10); -lean_inc(x_18); -lean_inc(x_12); -x_19 = lean_alloc_closure((void*)(l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_checkLeftRec___spec__1___lambda__1___boxed), 4, 1); -lean_closure_set(x_19, 0, x_12); -lean_inc(x_16); -x_20 = lean_alloc_closure((void*)(l_ReaderT_pure___at_Lean_Elab_liftMacroM___spec__2___rarg___boxed), 3, 1); -lean_closure_set(x_20, 0, x_16); -lean_inc(x_12); -x_21 = lean_alloc_closure((void*)(l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_checkLeftRec___spec__1___lambda__2___boxed), 4, 1); -lean_closure_set(x_21, 0, x_12); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_12); -x_22 = lean_alloc_closure((void*)(l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_checkLeftRec___spec__1___lambda__3___boxed), 6, 3); -lean_closure_set(x_22, 0, x_12); -lean_closure_set(x_22, 1, x_16); -lean_closure_set(x_22, 2, x_17); +x_12 = lean_ctor_get(x_1, 3); lean_inc(x_12); -x_23 = lean_alloc_closure((void*)(l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_checkLeftRec___spec__1___lambda__4___boxed), 6, 3); -lean_closure_set(x_23, 0, x_12); -lean_closure_set(x_23, 1, x_16); -lean_closure_set(x_23, 2, x_17); -x_24 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_24, 0, x_19); -lean_ctor_set(x_24, 1, x_20); -lean_ctor_set(x_24, 2, x_21); -lean_ctor_set(x_24, 3, x_22); -lean_ctor_set(x_24, 4, x_23); -x_25 = lean_st_ref_get(x_7, x_11); -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 1); -lean_inc(x_27); -lean_dec(x_25); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -x_29 = lean_environment_main_module(x_12); -x_30 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_30, 0, x_24); -lean_ctor_set(x_30, 1, x_29); -lean_ctor_set(x_30, 2, x_18); -lean_ctor_set(x_30, 3, x_13); -lean_ctor_set(x_30, 4, x_14); -lean_ctor_set(x_30, 5, x_15); -x_31 = lean_box(0); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_28); -lean_ctor_set(x_32, 1, x_31); -x_33 = lean_apply_2(x_1, x_30, x_32); -if (lean_obj_tag(x_33) == 0) -{ -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; -x_34 = lean_ctor_get(x_33, 0); -lean_inc(x_34); -x_35 = lean_ctor_get(x_33, 1); -lean_inc(x_35); -lean_dec(x_33); -x_36 = lean_ctor_get(x_35, 0); -lean_inc(x_36); -x_37 = lean_st_ref_take(x_7, x_27); -x_38 = lean_ctor_get(x_37, 0); -lean_inc(x_38); -x_39 = lean_ctor_get(x_37, 1); -lean_inc(x_39); -lean_dec(x_37); -x_40 = !lean_is_exclusive(x_38); -if (x_40 == 0) +x_13 = l_List_filterMap___at_Lean_resolveGlobalConst___spec__1(x_12); +x_14 = l_List_isEmpty___rarg(x_13); +if (x_14 == 0) { -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; -x_41 = lean_ctor_get(x_38, 1); -lean_dec(x_41); -lean_ctor_set(x_38, 1, x_36); -x_42 = lean_st_ref_set(x_7, x_38, x_39); -x_43 = lean_ctor_get(x_42, 1); -lean_inc(x_43); -lean_dec(x_42); -x_44 = lean_ctor_get(x_35, 1); -lean_inc(x_44); -lean_dec(x_35); -x_45 = l_List_reverse___rarg(x_44); -x_46 = l_List_forM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__3(x_45, x_2, x_3, x_4, x_5, x_6, x_7, x_43); +lean_object* x_15; +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_47 = !lean_is_exclusive(x_46); -if (x_47 == 0) +lean_dec(x_1); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_10); +return x_15; +} +else { -lean_object* x_48; -x_48 = lean_ctor_get(x_46, 0); -lean_dec(x_48); -lean_ctor_set(x_46, 0, x_34); -return x_46; +uint8_t x_16; +lean_dec(x_13); +x_16 = !lean_is_exclusive(x_8); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_8, 5); +x_18 = l_Lean_replaceRef(x_1, x_17); +lean_dec(x_17); +lean_dec(x_1); +lean_ctor_set(x_8, 5, x_18); +x_19 = l_Lean_resolveGlobalConstCore___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__6(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_19; +} +else +{ +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; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_20 = lean_ctor_get(x_8, 0); +x_21 = lean_ctor_get(x_8, 1); +x_22 = lean_ctor_get(x_8, 2); +x_23 = lean_ctor_get(x_8, 3); +x_24 = lean_ctor_get(x_8, 4); +x_25 = lean_ctor_get(x_8, 5); +x_26 = lean_ctor_get(x_8, 6); +x_27 = lean_ctor_get(x_8, 7); +x_28 = lean_ctor_get(x_8, 8); +x_29 = lean_ctor_get(x_8, 9); +x_30 = lean_ctor_get(x_8, 10); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_inc(x_27); +lean_inc(x_26); +lean_inc(x_25); +lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_8); +x_31 = l_Lean_replaceRef(x_1, x_25); +lean_dec(x_25); +lean_dec(x_1); +x_32 = lean_alloc_ctor(0, 11, 0); +lean_ctor_set(x_32, 0, x_20); +lean_ctor_set(x_32, 1, x_21); +lean_ctor_set(x_32, 2, x_22); +lean_ctor_set(x_32, 3, x_23); +lean_ctor_set(x_32, 4, x_24); +lean_ctor_set(x_32, 5, x_31); +lean_ctor_set(x_32, 6, x_26); +lean_ctor_set(x_32, 7, x_27); +lean_ctor_set(x_32, 8, x_28); +lean_ctor_set(x_32, 9, x_29); +lean_ctor_set(x_32, 10, x_30); +x_33 = l_Lean_resolveGlobalConstCore___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__6(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_32, x_9, x_10); +return x_33; +} +} +} +else +{ +lean_object* x_34; lean_object* x_35; +x_34 = l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3___closed__3; +x_35 = l_Lean_throwErrorAt___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__4(x_1, x_34, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_35; +} +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__11(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_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_11 = lean_ctor_get(x_8, 5); +x_12 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_6, x_7, x_8, x_9, x_10); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_12, 0); +lean_inc(x_11); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_11); +lean_ctor_set(x_15, 1, x_14); +lean_ctor_set_tag(x_12, 1); +lean_ctor_set(x_12, 0, x_15); +return x_12; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_12, 0); +x_17 = lean_ctor_get(x_12, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_12); +lean_inc(x_11); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_11); +lean_ctor_set(x_18, 1, x_16); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +return x_19; +} +} +} +LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__10(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_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; uint8_t x_12; +x_11 = lean_st_ref_get(x_9, x_10); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_11, 1); +x_15 = lean_ctor_get(x_13, 0); +lean_inc(x_15); +lean_dec(x_13); +lean_inc(x_1); +x_16 = lean_environment_find(x_15, x_1); +if (lean_obj_tag(x_16) == 0) +{ +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; +lean_free_object(x_11); +x_17 = lean_box(0); +x_18 = l_Lean_mkConst(x_1, x_17); +x_19 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_19, 0, x_18); +x_20 = l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__2; +x_21 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_19); +x_22 = l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__4; +x_23 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +x_24 = l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__11(x_23, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +return x_24; +} +else +{ +lean_object* x_25; +lean_dec(x_1); +x_25 = lean_ctor_get(x_16, 0); +lean_inc(x_25); +lean_dec(x_16); +lean_ctor_set(x_11, 0, x_25); +return x_11; +} +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_26 = lean_ctor_get(x_11, 0); +x_27 = lean_ctor_get(x_11, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_11); +x_28 = lean_ctor_get(x_26, 0); +lean_inc(x_28); +lean_dec(x_26); +lean_inc(x_1); +x_29 = lean_environment_find(x_28, x_1); +if (lean_obj_tag(x_29) == 0) +{ +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; +x_30 = lean_box(0); +x_31 = l_Lean_mkConst(x_1, x_30); +x_32 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_32, 0, x_31); +x_33 = l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__2; +x_34 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_32); +x_35 = l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__4; +x_36 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +x_37 = l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__11(x_36, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_27); +return x_37; +} +else +{ +lean_object* x_38; lean_object* x_39; +lean_dec(x_1); +x_38 = lean_ctor_get(x_29, 0); +lean_inc(x_38); +lean_dec(x_29); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_27); +return x_39; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__9(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_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +lean_inc(x_1); +x_11 = l_Lean_getConstInfo___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__10(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_13 = lean_ctor_get(x_11, 0); +x_14 = l_Lean_ConstantInfo_levelParams(x_13); +lean_dec(x_13); +x_15 = lean_box(0); +x_16 = l_List_mapTRAux___at_Lean_mkConstWithLevelParams___spec__1(x_14, x_15); +x_17 = l_Lean_mkConst(x_1, x_16); +lean_ctor_set(x_11, 0, x_17); +return x_11; } else { -lean_object* x_49; lean_object* x_50; -x_49 = lean_ctor_get(x_46, 1); -lean_inc(x_49); -lean_dec(x_46); -x_50 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_50, 0, x_34); -lean_ctor_set(x_50, 1, x_49); -return x_50; +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_18 = lean_ctor_get(x_11, 0); +x_19 = lean_ctor_get(x_11, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_11); +x_20 = l_Lean_ConstantInfo_levelParams(x_18); +lean_dec(x_18); +x_21 = lean_box(0); +x_22 = l_List_mapTRAux___at_Lean_mkConstWithLevelParams___spec__1(x_20, x_21); +x_23 = l_Lean_mkConst(x_1, x_22); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_19); +return x_24; } } else { -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_51 = lean_ctor_get(x_38, 0); -x_52 = lean_ctor_get(x_38, 2); -x_53 = lean_ctor_get(x_38, 3); -x_54 = lean_ctor_get(x_38, 4); -x_55 = lean_ctor_get(x_38, 5); -lean_inc(x_55); -lean_inc(x_54); -lean_inc(x_53); -lean_inc(x_52); -lean_inc(x_51); -lean_dec(x_38); -x_56 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_56, 0, x_51); -lean_ctor_set(x_56, 1, x_36); -lean_ctor_set(x_56, 2, x_52); -lean_ctor_set(x_56, 3, x_53); -lean_ctor_set(x_56, 4, x_54); -lean_ctor_set(x_56, 5, x_55); -x_57 = lean_st_ref_set(x_7, x_56, x_39); -x_58 = lean_ctor_get(x_57, 1); -lean_inc(x_58); -lean_dec(x_57); -x_59 = lean_ctor_get(x_35, 1); -lean_inc(x_59); -lean_dec(x_35); -x_60 = l_List_reverse___rarg(x_59); -x_61 = l_List_forM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__3(x_60, x_2, x_3, x_4, x_5, x_6, x_7, x_58); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_62 = lean_ctor_get(x_61, 1); -lean_inc(x_62); -if (lean_is_exclusive(x_61)) { - lean_ctor_release(x_61, 0); - lean_ctor_release(x_61, 1); - x_63 = x_61; -} else { - lean_dec_ref(x_61); - x_63 = lean_box(0); +uint8_t x_25; +lean_dec(x_1); +x_25 = !lean_is_exclusive(x_11); +if (x_25 == 0) +{ +return x_11; } -if (lean_is_scalar(x_63)) { - x_64 = lean_alloc_ctor(0, 2, 0); -} else { - x_64 = x_63; +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_11, 0); +x_27 = lean_ctor_get(x_11, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_11); +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +return x_28; } -lean_ctor_set(x_64, 0, x_34); -lean_ctor_set(x_64, 1, x_62); -return x_64; } } -else +} +LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoTree___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__13(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_object* x_9, lean_object* x_10) { +_start: { -lean_object* x_65; -x_65 = lean_ctor_get(x_33, 0); -lean_inc(x_65); -lean_dec(x_33); -if (lean_obj_tag(x_65) == 0) +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_11 = lean_st_ref_get(x_9, x_10); +x_12 = lean_ctor_get(x_11, 1); +lean_inc(x_12); +lean_dec(x_11); +x_13 = lean_st_ref_get(x_5, x_12); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_14, 4); +lean_inc(x_15); +lean_dec(x_14); +x_16 = lean_ctor_get_uint8(x_15, sizeof(void*)*2); +lean_dec(x_15); +if (x_16 == 0) { -lean_object* x_66; lean_object* x_67; lean_object* x_68; uint8_t x_69; -x_66 = lean_ctor_get(x_65, 0); -lean_inc(x_66); -x_67 = lean_ctor_get(x_65, 1); -lean_inc(x_67); -lean_dec(x_65); -x_68 = l_Lean_maxRecDepthErrorMessage; -x_69 = lean_string_dec_eq(x_67, x_68); -if (x_69 == 0) +uint8_t x_17; +lean_dec(x_1); +x_17 = !lean_is_exclusive(x_13); +if (x_17 == 0) { -lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_70 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_70, 0, x_67); -x_71 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_71, 0, x_70); -x_72 = l_Lean_throwErrorAt___at_Lean_Elab_Term_toParserDescr_process___spec__2(x_66, x_71, x_2, x_3, x_4, x_5, x_6, x_7, x_27); -return x_72; +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_13, 0); +lean_dec(x_18); +x_19 = lean_box(0); +lean_ctor_set(x_13, 0, x_19); +return x_13; } else { -lean_object* x_73; -lean_dec(x_67); -x_73 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_toParserDescr_process___spec__4(x_66, x_2, x_3, x_4, x_5, x_6, x_7, x_27); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_73; +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_13, 1); +lean_inc(x_20); +lean_dec(x_13); +x_21 = lean_box(0); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_20); +return x_22; } } else { -lean_object* x_74; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_74 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_toParserDescr_process___spec__5___rarg(x_27); -return x_74; +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; uint8_t x_30; +x_23 = lean_ctor_get(x_13, 1); +lean_inc(x_23); +lean_dec(x_13); +x_24 = lean_st_ref_get(x_9, x_23); +x_25 = lean_ctor_get(x_24, 1); +lean_inc(x_25); +lean_dec(x_24); +x_26 = lean_st_ref_take(x_5, x_25); +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_27, 4); +lean_inc(x_28); +x_29 = lean_ctor_get(x_26, 1); +lean_inc(x_29); +lean_dec(x_26); +x_30 = !lean_is_exclusive(x_27); +if (x_30 == 0) +{ +lean_object* x_31; uint8_t x_32; +x_31 = lean_ctor_get(x_27, 4); +lean_dec(x_31); +x_32 = !lean_is_exclusive(x_28); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_33 = lean_ctor_get(x_28, 1); +x_34 = l_Std_PersistentArray_push___rarg(x_33, x_1); +lean_ctor_set(x_28, 1, x_34); +x_35 = lean_st_ref_set(x_5, x_27, x_29); +x_36 = !lean_is_exclusive(x_35); +if (x_36 == 0) +{ +lean_object* x_37; lean_object* x_38; +x_37 = lean_ctor_get(x_35, 0); +lean_dec(x_37); +x_38 = lean_box(0); +lean_ctor_set(x_35, 0, x_38); +return x_35; +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_35, 1); +lean_inc(x_39); +lean_dec(x_35); +x_40 = lean_box(0); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_39); +return x_41; } } +else +{ +uint8_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_42 = lean_ctor_get_uint8(x_28, sizeof(void*)*2); +x_43 = lean_ctor_get(x_28, 0); +x_44 = lean_ctor_get(x_28, 1); +lean_inc(x_44); +lean_inc(x_43); +lean_dec(x_28); +x_45 = l_Std_PersistentArray_push___rarg(x_44, x_1); +x_46 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_46, 0, x_43); +lean_ctor_set(x_46, 1, x_45); +lean_ctor_set_uint8(x_46, sizeof(void*)*2, x_42); +lean_ctor_set(x_27, 4, x_46); +x_47 = lean_st_ref_set(x_5, x_27, x_29); +x_48 = lean_ctor_get(x_47, 1); +lean_inc(x_48); +if (lean_is_exclusive(x_47)) { + lean_ctor_release(x_47, 0); + lean_ctor_release(x_47, 1); + x_49 = x_47; +} else { + lean_dec_ref(x_47); + x_49 = lean_box(0); +} +x_50 = lean_box(0); +if (lean_is_scalar(x_49)) { + x_51 = lean_alloc_ctor(0, 2, 0); +} else { + x_51 = x_49; } +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_48); +return x_51; } -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_toParserDescr_process___spec__6(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_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -uint8_t x_12; -x_12 = !lean_is_exclusive(x_9); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = lean_ctor_get(x_9, 5); -x_14 = l_Lean_replaceRef(x_1, x_13); -lean_dec(x_13); -lean_ctor_set(x_9, 5, x_14); -x_15 = l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_9); -return x_15; } else { -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; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_16 = lean_ctor_get(x_9, 0); -x_17 = lean_ctor_get(x_9, 1); -x_18 = lean_ctor_get(x_9, 2); -x_19 = lean_ctor_get(x_9, 3); -x_20 = lean_ctor_get(x_9, 4); -x_21 = lean_ctor_get(x_9, 5); -x_22 = lean_ctor_get(x_9, 6); -x_23 = lean_ctor_get(x_9, 7); -x_24 = lean_ctor_get(x_9, 8); -x_25 = lean_ctor_get(x_9, 9); -x_26 = lean_ctor_get(x_9, 10); -lean_inc(x_26); -lean_inc(x_25); -lean_inc(x_24); -lean_inc(x_23); -lean_inc(x_22); -lean_inc(x_21); -lean_inc(x_20); -lean_inc(x_19); -lean_inc(x_18); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_9); -x_27 = l_Lean_replaceRef(x_1, x_21); -lean_dec(x_21); -x_28 = lean_alloc_ctor(0, 11, 0); -lean_ctor_set(x_28, 0, x_16); -lean_ctor_set(x_28, 1, x_17); -lean_ctor_set(x_28, 2, x_18); -lean_ctor_set(x_28, 3, x_19); -lean_ctor_set(x_28, 4, x_20); -lean_ctor_set(x_28, 5, x_27); -lean_ctor_set(x_28, 6, x_22); -lean_ctor_set(x_28, 7, x_23); -lean_ctor_set(x_28, 8, x_24); -lean_ctor_set(x_28, 9, x_25); -lean_ctor_set(x_28, 10, x_26); -x_29 = l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_28, x_10, x_11); -lean_dec(x_28); -return x_29; +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; lean_object* x_67; +x_52 = lean_ctor_get(x_27, 0); +x_53 = lean_ctor_get(x_27, 1); +x_54 = lean_ctor_get(x_27, 2); +x_55 = lean_ctor_get(x_27, 3); +lean_inc(x_55); +lean_inc(x_54); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_27); +x_56 = lean_ctor_get_uint8(x_28, sizeof(void*)*2); +x_57 = lean_ctor_get(x_28, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_28, 1); +lean_inc(x_58); +if (lean_is_exclusive(x_28)) { + lean_ctor_release(x_28, 0); + lean_ctor_release(x_28, 1); + x_59 = x_28; +} else { + lean_dec_ref(x_28); + x_59 = lean_box(0); +} +x_60 = l_Std_PersistentArray_push___rarg(x_58, x_1); +if (lean_is_scalar(x_59)) { + x_61 = lean_alloc_ctor(0, 2, 1); +} else { + x_61 = x_59; } +lean_ctor_set(x_61, 0, x_57); +lean_ctor_set(x_61, 1, x_60); +lean_ctor_set_uint8(x_61, sizeof(void*)*2, x_56); +x_62 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_62, 0, x_52); +lean_ctor_set(x_62, 1, x_53); +lean_ctor_set(x_62, 2, x_54); +lean_ctor_set(x_62, 3, x_55); +lean_ctor_set(x_62, 4, x_61); +x_63 = lean_st_ref_set(x_5, x_62, x_29); +x_64 = lean_ctor_get(x_63, 1); +lean_inc(x_64); +if (lean_is_exclusive(x_63)) { + lean_ctor_release(x_63, 0); + lean_ctor_release(x_63, 1); + x_65 = x_63; +} else { + lean_dec_ref(x_63); + x_65 = lean_box(0); } +x_66 = lean_box(0); +if (lean_is_scalar(x_65)) { + x_67 = lean_alloc_ctor(0, 2, 0); +} else { + x_67 = x_65; } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_process___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("paren", 5); -return x_1; +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_67, 1, x_64); +return x_67; } } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_process___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_checkLeftRec___closed__2; -x_2 = l_Lean_Elab_Term_toParserDescr_process___closed__1; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_process___closed__3() { +static lean_object* _init_l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___closed__1() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("unary", 5); -return x_1; +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(32u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_process___closed__4() { +static lean_object* _init_l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_checkLeftRec___closed__2; -x_2 = l_Lean_Elab_Term_toParserDescr_process___closed__3; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___closed__1; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_process___closed__5() { +static lean_object* _init_l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___closed__3() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_checkLeftRec___closed__2; -x_2 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__14; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; +size_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = 5; +x_2 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___closed__2; +x_3 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___closed__1; +x_4 = lean_unsigned_to_nat(0u); +x_5 = lean_alloc_ctor(0, 4, sizeof(size_t)*1); +lean_ctor_set(x_5, 0, x_2); +lean_ctor_set(x_5, 1, x_3); +lean_ctor_set(x_5, 2, x_4); +lean_ctor_set(x_5, 3, x_4); +lean_ctor_set_usize(x_5, 4, x_1); +return x_5; } } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_process___closed__6() { +LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12(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_object* x_9, lean_object* x_10) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("sepBy", 5); -return x_1; -} +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_11 = lean_st_ref_get(x_9, x_10); +x_12 = lean_ctor_get(x_11, 1); +lean_inc(x_12); +lean_dec(x_11); +x_13 = lean_st_ref_get(x_5, x_12); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_14, 4); +lean_inc(x_15); +lean_dec(x_14); +x_16 = lean_ctor_get_uint8(x_15, sizeof(void*)*2); +lean_dec(x_15); +if (x_16 == 0) +{ +uint8_t x_17; +lean_dec(x_1); +x_17 = !lean_is_exclusive(x_13); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_13, 0); +lean_dec(x_18); +x_19 = lean_box(0); +lean_ctor_set(x_13, 0, x_19); +return x_13; } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_process___closed__7() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_checkLeftRec___closed__2; -x_2 = l_Lean_Elab_Term_toParserDescr_process___closed__6; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_13, 1); +lean_inc(x_20); +lean_dec(x_13); +x_21 = lean_box(0); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_20); +return x_22; } } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_process___closed__8() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("sepBy1", 6); -return x_1; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_23 = lean_ctor_get(x_13, 1); +lean_inc(x_23); +lean_dec(x_13); +x_24 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___closed__3; +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_1); +lean_ctor_set(x_25, 1, x_24); +x_26 = l_Lean_Elab_pushInfoTree___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__13(x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_23); +return x_26; } } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_process___closed__9() { +} +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__14(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_checkLeftRec___closed__2; -x_2 = l_Lean_Elab_Term_toParserDescr_process___closed__8; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_14; +lean_dec(x_2); +lean_dec(x_1); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_4); +lean_ctor_set(x_14, 1, x_13); +return x_14; } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_process___closed__10() { -_start: +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +lean_dec(x_4); +x_15 = lean_ctor_get(x_3, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_3, 1); +lean_inc(x_16); +lean_dec(x_3); +x_17 = l_Lean_mkConstWithLevelParams___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__9(x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_17) == 0) { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("atom", 4); -return x_1; -} +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_box(0); +lean_inc(x_1); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_1); +x_22 = l_Lean_LocalContext_empty; +x_23 = 0; +lean_inc(x_2); +x_24 = lean_alloc_ctor(0, 4, 1); +lean_ctor_set(x_24, 0, x_21); +lean_ctor_set(x_24, 1, x_22); +lean_ctor_set(x_24, 2, x_2); +lean_ctor_set(x_24, 3, x_18); +lean_ctor_set_uint8(x_24, sizeof(void*)*4, x_23); +x_25 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_25, 0, x_24); +x_26 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12(x_25, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_19); +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); +lean_dec(x_26); +x_28 = lean_box(0); +x_3 = x_16; +x_4 = x_28; +x_13 = x_27; +goto _start; } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_process___closed__11() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_checkLeftRec___closed__2; -x_2 = l_Lean_Elab_Term_toParserDescr_process___closed__10; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_process___closed__12() { -_start: +uint8_t x_30; +lean_dec(x_16); +lean_dec(x_2); +lean_dec(x_1); +x_30 = !lean_is_exclusive(x_17); +if (x_30 == 0) { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("nonReserved", 11); -return x_1; -} +return x_17; } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_process___closed__13() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_checkLeftRec___closed__2; -x_2 = l_Lean_Elab_Term_toParserDescr_process___closed__12; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_17, 0); +x_32 = lean_ctor_get(x_17, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_17); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; } } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_process___closed__14() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("unexpected syntax kind of category `syntax`: ", 45); -return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_process___closed__15() { +} +LEAN_EXPORT lean_object* l_Lean_Elab_resolveGlobalConstWithInfos___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__2___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, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_toParserDescr_process___closed__14; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_12; +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_1); +lean_ctor_set(x_12, 1, x_11); +return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_process(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_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Elab_resolveGlobalConstWithInfos___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___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_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_11; lean_object* x_12; uint8_t x_13; uint8_t x_14; -lean_inc(x_1); -x_11 = l_Lean_Syntax_getKind(x_1); -x_12 = l_Lean_nullKind; -x_13 = lean_name_eq(x_11, x_12); -x_14 = !lean_is_exclusive(x_8); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; -x_15 = lean_ctor_get(x_8, 5); -x_16 = l_Lean_replaceRef(x_1, x_15); -lean_dec(x_15); -lean_ctor_set(x_8, 5, x_16); -if (x_13 == 0) -{ -lean_object* x_17; uint8_t x_18; -x_17 = l_Lean_choiceKind; -x_18 = lean_name_eq(x_11, x_17); -if (x_18 == 0) -{ -lean_object* x_19; uint8_t x_20; -x_19 = l_Lean_Elab_Term_toParserDescr_process___closed__2; -x_20 = lean_name_eq(x_11, x_19); -if (x_20 == 0) -{ -lean_object* x_21; uint8_t x_22; -x_21 = l_Lean_Elab_Term_checkLeftRec___closed__4; -x_22 = lean_name_eq(x_11, x_21); -if (x_22 == 0) -{ -lean_object* x_23; uint8_t x_24; -x_23 = l_Lean_Elab_Term_toParserDescr_process___closed__4; -x_24 = lean_name_eq(x_11, x_23); -if (x_24 == 0) -{ -lean_object* x_25; uint8_t x_26; -x_25 = l_Lean_Elab_Term_toParserDescr_process___closed__5; -x_26 = lean_name_eq(x_11, x_25); -if (x_26 == 0) -{ -lean_object* x_27; uint8_t x_28; -x_27 = l_Lean_Elab_Term_toParserDescr_process___closed__7; -x_28 = lean_name_eq(x_11, x_27); -if (x_28 == 0) -{ -lean_object* x_29; uint8_t x_30; -x_29 = l_Lean_Elab_Term_toParserDescr_process___closed__9; -x_30 = lean_name_eq(x_11, x_29); -if (x_30 == 0) -{ -lean_object* x_31; uint8_t x_32; -x_31 = l_Lean_Elab_Term_toParserDescr_process___closed__11; -x_32 = lean_name_eq(x_11, x_31); -if (x_32 == 0) -{ -lean_object* x_33; uint8_t x_34; -x_33 = l_Lean_Elab_Term_toParserDescr_process___closed__13; -x_34 = lean_name_eq(x_11, x_33); -if (x_34 == 0) -{ -lean_object* x_35; lean_object* x_36; -lean_inc(x_1); -x_35 = lean_alloc_closure((void*)(l_Lean_Macro_expandMacro_x3f), 3, 1); -lean_closure_set(x_35, 0, x_1); +lean_object* x_12; +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_36 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_toParserDescr_process___spec__1(x_35, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_36) == 0) +lean_inc(x_3); +lean_inc(x_1); +x_12 = l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_12) == 0) { -lean_object* x_37; -x_37 = lean_ctor_get(x_36, 0); -lean_inc(x_37); -if (lean_obj_tag(x_37) == 0) +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; uint8_t x_20; +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = lean_st_ref_get(x_10, x_14); +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +x_17 = lean_st_ref_get(x_6, x_16); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_18, 4); +lean_inc(x_19); +lean_dec(x_18); +x_20 = lean_ctor_get_uint8(x_19, sizeof(void*)*2); +lean_dec(x_19); +if (x_20 == 0) { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_38 = lean_ctor_get(x_36, 1); -lean_inc(x_38); -lean_dec(x_36); -x_39 = lean_alloc_ctor(4, 1, 0); -lean_ctor_set(x_39, 0, x_11); -x_40 = l_Lean_Elab_Term_toParserDescr_process___closed__15; -x_41 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set(x_41, 1, x_39); -x_42 = l_Lean_addTrace___at_Lean_Elab_Term_checkLeftRec___spec__3___closed__8; -x_43 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_43, 0, x_41); -lean_ctor_set(x_43, 1, x_42); -x_44 = l_Lean_throwErrorAt___at_Lean_Elab_Term_toParserDescr_process___spec__6(x_1, x_43, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_38); +uint8_t x_21; +lean_dec(x_10); lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); @@ -8931,961 +10111,823 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_44; +x_21 = !lean_is_exclusive(x_17); +if (x_21 == 0) +{ +lean_object* x_22; +x_22 = lean_ctor_get(x_17, 0); +lean_dec(x_22); +lean_ctor_set(x_17, 0, x_13); +return x_17; } else { -lean_object* x_45; lean_object* x_46; -lean_dec(x_11); -lean_dec(x_1); -x_45 = lean_ctor_get(x_36, 1); -lean_inc(x_45); -lean_dec(x_36); -x_46 = lean_ctor_get(x_37, 0); -lean_inc(x_46); -lean_dec(x_37); -x_1 = x_46; -x_10 = x_45; -goto _start; +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_17, 1); +lean_inc(x_23); +lean_dec(x_17); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_13); +lean_ctor_set(x_24, 1, x_23); +return x_24; +} +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_17, 1); +lean_inc(x_25); +lean_dec(x_17); +x_26 = lean_box(0); +lean_inc(x_13); +x_27 = l_List_forIn_loop___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__14(x_1, x_2, x_13, x_26, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_25); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +if (lean_obj_tag(x_27) == 0) +{ +uint8_t x_28; +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) +{ +lean_object* x_29; +x_29 = lean_ctor_get(x_27, 0); +lean_dec(x_29); +lean_ctor_set(x_27, 0, x_13); +return x_27; +} +else +{ +lean_object* x_30; lean_object* x_31; +x_30 = lean_ctor_get(x_27, 1); +lean_inc(x_30); +lean_dec(x_27); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_13); +lean_ctor_set(x_31, 1, x_30); +return x_31; } } else { -uint8_t x_48; -lean_dec(x_8); -lean_dec(x_11); -lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_48 = !lean_is_exclusive(x_36); -if (x_48 == 0) +uint8_t x_32; +lean_dec(x_13); +x_32 = !lean_is_exclusive(x_27); +if (x_32 == 0) { -return x_36; +return x_27; } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_36, 0); -x_50 = lean_ctor_get(x_36, 1); -lean_inc(x_50); -lean_inc(x_49); -lean_dec(x_36); -x_51 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_51, 0, x_49); -lean_ctor_set(x_51, 1, x_50); -return x_51; +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_27, 0); +x_34 = lean_ctor_get(x_27, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_27); +x_35 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; +} } } } else { -lean_object* x_52; -lean_dec(x_11); -x_52 = l_Lean_Elab_Term_toParserDescr_processNonReserved(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +uint8_t x_36; +lean_dec(x_10); lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_52; -} -} -else +lean_dec(x_1); +x_36 = !lean_is_exclusive(x_12); +if (x_36 == 0) { -lean_object* x_53; -lean_dec(x_11); -x_53 = l_Lean_Elab_Term_toParserDescr_processAtom(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_53; -} +return x_12; } else { -lean_object* x_54; -lean_dec(x_11); -x_54 = l_Lean_Elab_Term_toParserDescr_processSepBy1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_54; +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_12, 0); +x_38 = lean_ctor_get(x_12, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_12); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; } } -else -{ -lean_object* x_55; -lean_dec(x_11); -x_55 = l_Lean_Elab_Term_toParserDescr_processSepBy(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_55; } } -else +LEAN_EXPORT lean_object* l_List_filterMap___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__15(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_56; -lean_dec(x_11); -x_56 = l_Lean_Elab_Term_toParserDescr_processBinary(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_1); -return x_56; -} -} -else +if (lean_obj_tag(x_2) == 0) { -lean_object* x_57; -lean_dec(x_11); -x_57 = l_Lean_Elab_Term_toParserDescr_processUnary(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_object* x_3; lean_dec(x_1); -return x_57; -} +x_3 = lean_box(0); +return x_3; } else { -lean_object* x_58; -lean_dec(x_11); -x_58 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_58; -} +uint8_t x_4; +x_4 = !lean_is_exclusive(x_2); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_2, 0); +x_6 = lean_ctor_get(x_2, 1); +lean_inc(x_5); +lean_inc(x_1); +x_7 = lean_environment_find(x_1, x_5); +if (lean_obj_tag(x_7) == 0) +{ +lean_free_object(x_2); +lean_dec(x_5); +x_2 = x_6; +goto _start; } else { -lean_object* x_59; lean_object* x_60; +lean_object* x_9; lean_object* x_10; +x_9 = lean_ctor_get(x_7, 0); +lean_inc(x_9); +lean_dec(x_7); +x_10 = l_Lean_ConstantInfo_type(x_9); +lean_dec(x_9); +if (lean_obj_tag(x_10) == 4) +{ +lean_object* x_11; +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +lean_dec(x_10); +if (lean_obj_tag(x_11) == 1) +{ +lean_object* x_12; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +if (lean_obj_tag(x_12) == 1) +{ +lean_object* x_13; +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +switch (lean_obj_tag(x_13)) { +case 0: +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); lean_dec(x_11); -x_59 = lean_unsigned_to_nat(1u); -x_60 = l_Lean_Syntax_getArg(x_1, x_59); -lean_dec(x_1); -x_1 = x_60; +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_dec(x_12); +x_16 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__1; +x_17 = lean_string_dec_eq(x_15, x_16); +lean_dec(x_15); +if (x_17 == 0) +{ +lean_dec(x_14); +lean_free_object(x_2); +lean_dec(x_5); +x_2 = x_6; goto _start; } -} else { -lean_object* x_62; lean_object* x_63; -lean_dec(x_11); -x_62 = lean_unsigned_to_nat(0u); -x_63 = l_Lean_Syntax_getArg(x_1, x_62); -lean_dec(x_1); -x_1 = x_63; +lean_object* x_19; uint8_t x_20; +x_19 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__12; +x_20 = lean_string_dec_eq(x_14, x_19); +if (x_20 == 0) +{ +lean_object* x_21; uint8_t x_22; +x_21 = l_List_filterMap___at_Lean_Elab_Term_resolveParserName___spec__1___closed__1; +x_22 = lean_string_dec_eq(x_14, x_21); +lean_dec(x_14); +if (x_22 == 0) +{ +lean_free_object(x_2); +lean_dec(x_5); +x_2 = x_6; goto _start; } -} else { -lean_object* x_65; -lean_dec(x_11); -x_65 = l_Lean_Elab_Term_toParserDescr_processSeq(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_65; +uint8_t x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_24 = 1; +x_25 = lean_box(x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_5); +lean_ctor_set(x_26, 1, x_25); +x_27 = l_List_filterMap___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__15(x_1, x_6); +lean_ctor_set(x_2, 1, x_27); +lean_ctor_set(x_2, 0, x_26); +return x_2; } } else { -lean_object* x_66; 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_66 = lean_ctor_get(x_8, 0); -x_67 = lean_ctor_get(x_8, 1); -x_68 = lean_ctor_get(x_8, 2); -x_69 = lean_ctor_get(x_8, 3); -x_70 = lean_ctor_get(x_8, 4); -x_71 = lean_ctor_get(x_8, 5); -x_72 = lean_ctor_get(x_8, 6); -x_73 = lean_ctor_get(x_8, 7); -x_74 = lean_ctor_get(x_8, 8); -x_75 = lean_ctor_get(x_8, 9); -x_76 = lean_ctor_get(x_8, 10); -lean_inc(x_76); -lean_inc(x_75); -lean_inc(x_74); -lean_inc(x_73); -lean_inc(x_72); -lean_inc(x_71); -lean_inc(x_70); -lean_inc(x_69); -lean_inc(x_68); -lean_inc(x_67); -lean_inc(x_66); -lean_dec(x_8); -x_77 = l_Lean_replaceRef(x_1, x_71); -lean_dec(x_71); -x_78 = lean_alloc_ctor(0, 11, 0); -lean_ctor_set(x_78, 0, x_66); -lean_ctor_set(x_78, 1, x_67); -lean_ctor_set(x_78, 2, x_68); -lean_ctor_set(x_78, 3, x_69); -lean_ctor_set(x_78, 4, x_70); -lean_ctor_set(x_78, 5, x_77); -lean_ctor_set(x_78, 6, x_72); -lean_ctor_set(x_78, 7, x_73); -lean_ctor_set(x_78, 8, x_74); -lean_ctor_set(x_78, 9, x_75); -lean_ctor_set(x_78, 10, x_76); -if (x_13 == 0) -{ -lean_object* x_79; uint8_t x_80; -x_79 = l_Lean_choiceKind; -x_80 = lean_name_eq(x_11, x_79); -if (x_80 == 0) -{ -lean_object* x_81; uint8_t x_82; -x_81 = l_Lean_Elab_Term_toParserDescr_process___closed__2; -x_82 = lean_name_eq(x_11, x_81); -if (x_82 == 0) -{ -lean_object* x_83; uint8_t x_84; -x_83 = l_Lean_Elab_Term_checkLeftRec___closed__4; -x_84 = lean_name_eq(x_11, x_83); -if (x_84 == 0) -{ -lean_object* x_85; uint8_t x_86; -x_85 = l_Lean_Elab_Term_toParserDescr_process___closed__4; -x_86 = lean_name_eq(x_11, x_85); -if (x_86 == 0) -{ -lean_object* x_87; uint8_t x_88; -x_87 = l_Lean_Elab_Term_toParserDescr_process___closed__5; -x_88 = lean_name_eq(x_11, x_87); -if (x_88 == 0) -{ -lean_object* x_89; uint8_t x_90; -x_89 = l_Lean_Elab_Term_toParserDescr_process___closed__7; -x_90 = lean_name_eq(x_11, x_89); -if (x_90 == 0) -{ -lean_object* x_91; uint8_t x_92; -x_91 = l_Lean_Elab_Term_toParserDescr_process___closed__9; -x_92 = lean_name_eq(x_11, x_91); -if (x_92 == 0) -{ -lean_object* x_93; uint8_t x_94; -x_93 = l_Lean_Elab_Term_toParserDescr_process___closed__11; -x_94 = lean_name_eq(x_11, x_93); -if (x_94 == 0) -{ -lean_object* x_95; uint8_t x_96; -x_95 = l_Lean_Elab_Term_toParserDescr_process___closed__13; -x_96 = lean_name_eq(x_11, x_95); -if (x_96 == 0) -{ -lean_object* x_97; lean_object* x_98; -lean_inc(x_1); -x_97 = lean_alloc_closure((void*)(l_Lean_Macro_expandMacro_x3f), 3, 1); -lean_closure_set(x_97, 0, x_1); -lean_inc(x_9); -lean_inc(x_78); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_98 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_toParserDescr_process___spec__1(x_97, x_4, x_5, x_6, x_7, x_78, x_9, x_10); -if (lean_obj_tag(x_98) == 0) -{ -lean_object* x_99; -x_99 = lean_ctor_get(x_98, 0); -lean_inc(x_99); -if (lean_obj_tag(x_99) == 0) -{ -lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; -x_100 = lean_ctor_get(x_98, 1); -lean_inc(x_100); -lean_dec(x_98); -x_101 = lean_alloc_ctor(4, 1, 0); -lean_ctor_set(x_101, 0, x_11); -x_102 = l_Lean_Elab_Term_toParserDescr_process___closed__15; -x_103 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_103, 0, x_102); -lean_ctor_set(x_103, 1, x_101); -x_104 = l_Lean_addTrace___at_Lean_Elab_Term_checkLeftRec___spec__3___closed__8; -x_105 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_105, 0, x_103); -lean_ctor_set(x_105, 1, x_104); -x_106 = l_Lean_throwErrorAt___at_Lean_Elab_Term_toParserDescr_process___spec__6(x_1, x_105, x_2, x_3, x_4, x_5, x_6, x_7, x_78, x_9, x_100); -lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_106; +uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +lean_dec(x_14); +x_28 = 1; +x_29 = lean_box(x_28); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_5); +lean_ctor_set(x_30, 1, x_29); +x_31 = l_List_filterMap___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__15(x_1, x_6); +lean_ctor_set(x_2, 1, x_31); +lean_ctor_set(x_2, 0, x_30); +return x_2; } -else +} +} +case 1: +{ +lean_object* x_32; +x_32 = lean_ctor_get(x_13, 0); +lean_inc(x_32); +if (lean_obj_tag(x_32) == 0) { -lean_object* x_107; lean_object* x_108; +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_33 = lean_ctor_get(x_11, 1); +lean_inc(x_33); lean_dec(x_11); -lean_dec(x_1); -x_107 = lean_ctor_get(x_98, 1); -lean_inc(x_107); -lean_dec(x_98); -x_108 = lean_ctor_get(x_99, 0); -lean_inc(x_108); -lean_dec(x_99); -x_1 = x_108; -x_8 = x_78; -x_10 = x_107; +x_34 = lean_ctor_get(x_12, 1); +lean_inc(x_34); +lean_dec(x_12); +x_35 = lean_ctor_get(x_13, 1); +lean_inc(x_35); +lean_dec(x_13); +x_36 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__1; +x_37 = lean_string_dec_eq(x_35, x_36); +lean_dec(x_35); +if (x_37 == 0) +{ +lean_dec(x_34); +lean_dec(x_33); +lean_free_object(x_2); +lean_dec(x_5); +x_2 = x_6; goto _start; } -} else { -lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; -lean_dec(x_78); -lean_dec(x_11); -lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_6); +lean_object* x_39; uint8_t x_40; +x_39 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__3; +x_40 = lean_string_dec_eq(x_34, x_39); +lean_dec(x_34); +if (x_40 == 0) +{ +lean_dec(x_33); +lean_free_object(x_2); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_110 = lean_ctor_get(x_98, 0); -lean_inc(x_110); -x_111 = lean_ctor_get(x_98, 1); -lean_inc(x_111); -if (lean_is_exclusive(x_98)) { - lean_ctor_release(x_98, 0); - lean_ctor_release(x_98, 1); - x_112 = x_98; -} else { - lean_dec_ref(x_98); - x_112 = lean_box(0); -} -if (lean_is_scalar(x_112)) { - x_113 = lean_alloc_ctor(1, 2, 0); -} else { - x_113 = x_112; -} -lean_ctor_set(x_113, 0, x_110); -lean_ctor_set(x_113, 1, x_111); -return x_113; -} +x_2 = x_6; +goto _start; } else { -lean_object* x_114; -lean_dec(x_11); -x_114 = l_Lean_Elab_Term_toParserDescr_processNonReserved(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_78, x_9, x_10); -lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_6); +lean_object* x_42; uint8_t x_43; +x_42 = l_List_filterMap___at_Lean_Elab_Term_resolveParserName___spec__1___closed__2; +x_43 = lean_string_dec_eq(x_33, x_42); +if (x_43 == 0) +{ +uint8_t x_44; +x_44 = lean_string_dec_eq(x_33, x_39); +lean_dec(x_33); +if (x_44 == 0) +{ +lean_free_object(x_2); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_114; -} +x_2 = x_6; +goto _start; } else { -lean_object* x_115; -lean_dec(x_11); -x_115 = l_Lean_Elab_Term_toParserDescr_processAtom(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_78, x_9, x_10); -return x_115; +uint8_t x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_46 = 0; +x_47 = lean_box(x_46); +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_5); +lean_ctor_set(x_48, 1, x_47); +x_49 = l_List_filterMap___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__15(x_1, x_6); +lean_ctor_set(x_2, 1, x_49); +lean_ctor_set(x_2, 0, x_48); +return x_2; } } else { -lean_object* x_116; -lean_dec(x_11); -x_116 = l_Lean_Elab_Term_toParserDescr_processSepBy1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_78, x_9, x_10); -return x_116; +uint8_t x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +lean_dec(x_33); +x_50 = 0; +x_51 = lean_box(x_50); +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_5); +lean_ctor_set(x_52, 1, x_51); +x_53 = l_List_filterMap___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__15(x_1, x_6); +lean_ctor_set(x_2, 1, x_53); +lean_ctor_set(x_2, 0, x_52); +return x_2; } } -else -{ -lean_object* x_117; -lean_dec(x_11); -x_117 = l_Lean_Elab_Term_toParserDescr_processSepBy(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_78, x_9, x_10); -return x_117; } } else { -lean_object* x_118; +lean_dec(x_32); +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); -x_118 = l_Lean_Elab_Term_toParserDescr_processBinary(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_78, x_9, x_10); -lean_dec(x_1); -return x_118; +lean_free_object(x_2); +lean_dec(x_5); +x_2 = x_6; +goto _start; } } -else +default: { -lean_object* x_119; +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); -x_119 = l_Lean_Elab_Term_toParserDescr_processUnary(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_78, x_9, x_10); -lean_dec(x_1); -return x_119; -} +lean_free_object(x_2); +lean_dec(x_5); +x_2 = x_6; +goto _start; } -else -{ -lean_object* x_120; -lean_dec(x_11); -x_120 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_78, x_9, x_10); -return x_120; } } else { -lean_object* x_121; lean_object* x_122; +lean_dec(x_12); lean_dec(x_11); -x_121 = lean_unsigned_to_nat(1u); -x_122 = l_Lean_Syntax_getArg(x_1, x_121); -lean_dec(x_1); -x_1 = x_122; -x_8 = x_78; +lean_free_object(x_2); +lean_dec(x_5); +x_2 = x_6; goto _start; } } else { -lean_object* x_124; lean_object* x_125; lean_dec(x_11); -x_124 = lean_unsigned_to_nat(0u); -x_125 = l_Lean_Syntax_getArg(x_1, x_124); -lean_dec(x_1); -x_1 = x_125; -x_8 = x_78; +lean_free_object(x_2); +lean_dec(x_5); +x_2 = x_6; goto _start; } } else { -lean_object* x_127; -lean_dec(x_11); -x_127 = l_Lean_Elab_Term_toParserDescr_processSeq(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_78, x_9, x_10); -return x_127; -} -} -} +lean_dec(x_10); +lean_free_object(x_2); +lean_dec(x_5); +x_2 = x_6; +goto _start; } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("ParserDescr.sepBy1", 18); -return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__1; -x_2 = lean_string_utf8_byte_size(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__3() { -_start: +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_2, 0); +x_60 = lean_ctor_get(x_2, 1); +lean_inc(x_60); +lean_inc(x_59); +lean_dec(x_2); +lean_inc(x_59); +lean_inc(x_1); +x_61 = lean_environment_find(x_1, x_59); +if (lean_obj_tag(x_61) == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__1; -x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__2; -x_4 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_4, 0, x_1); -lean_ctor_set(x_4, 1, x_2); -lean_ctor_set(x_4, 2, x_3); -return x_4; -} +lean_dec(x_59); +x_2 = x_60; +goto _start; } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__4() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__13; -x_2 = l_Lean_Elab_Term_toParserDescr_process___closed__8; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__5() { -_start: +lean_object* x_63; lean_object* x_64; +x_63 = lean_ctor_get(x_61, 0); +lean_inc(x_63); +lean_dec(x_61); +x_64 = l_Lean_ConstantInfo_type(x_63); +lean_dec(x_63); +if (lean_obj_tag(x_64) == 4) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__16; -x_2 = l_Lean_Elab_Term_toParserDescr_process___closed__8; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__6() { -_start: +lean_object* x_65; +x_65 = lean_ctor_get(x_64, 0); +lean_inc(x_65); +lean_dec(x_64); +if (lean_obj_tag(x_65) == 1) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__5; -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_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__7() { -_start: +lean_object* x_66; +x_66 = lean_ctor_get(x_65, 0); +lean_inc(x_66); +if (lean_obj_tag(x_66) == 1) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__6; -x_3 = lean_alloc_ctor(1, 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_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__8() { -_start: +lean_object* x_67; +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +switch (lean_obj_tag(x_67)) { +case 0: { -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(4u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__9() { -_start: +lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; +x_68 = lean_ctor_get(x_65, 1); +lean_inc(x_68); +lean_dec(x_65); +x_69 = lean_ctor_get(x_66, 1); +lean_inc(x_69); +lean_dec(x_66); +x_70 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__1; +x_71 = lean_string_dec_eq(x_69, x_70); +lean_dec(x_69); +if (x_71 == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__15; -x_3 = l_Lean_mkCIdentFrom(x_1, x_2); -return x_3; -} +lean_dec(x_68); +lean_dec(x_59); +x_2 = x_60; +goto _start; } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__10() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("true", 4); -return x_1; -} +lean_object* x_73; uint8_t x_74; +x_73 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__12; +x_74 = lean_string_dec_eq(x_68, x_73); +if (x_74 == 0) +{ +lean_object* x_75; uint8_t x_76; +x_75 = l_List_filterMap___at_Lean_Elab_Term_resolveParserName___spec__1___closed__1; +x_76 = lean_string_dec_eq(x_68, x_75); +lean_dec(x_68); +if (x_76 == 0) +{ +lean_dec(x_59); +x_2 = x_60; +goto _start; } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__11() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__14; -x_2 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__10; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; +uint8_t x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_78 = 1; +x_79 = lean_box(x_78); +x_80 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_80, 0, x_59); +lean_ctor_set(x_80, 1, x_79); +x_81 = l_List_filterMap___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__15(x_1, x_60); +x_82 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_82, 0, x_80); +lean_ctor_set(x_82, 1, x_81); +return x_82; } } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__12() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__11; -x_3 = l_Lean_mkCIdentFrom(x_1, x_2); -return x_3; +uint8_t x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +lean_dec(x_68); +x_83 = 1; +x_84 = lean_box(x_83); +x_85 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_85, 0, x_59); +lean_ctor_set(x_85, 1, x_84); +x_86 = l_List_filterMap___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__15(x_1, x_60); +x_87 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_87, 0, x_85); +lean_ctor_set(x_87, 1, x_86); +return x_87; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy1___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, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: -{ -uint8_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_84; lean_object* x_85; uint8_t x_86; lean_object* x_87; -x_84 = lean_unsigned_to_nat(5u); -x_85 = l_Lean_Syntax_getArg(x_3, x_84); -x_86 = l_Lean_Syntax_isNone(x_85); -lean_dec(x_85); -lean_inc(x_11); -x_87 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_toParserDescr_processNonReserved___spec__2___rarg(x_11, x_12, x_13); -if (x_86 == 0) +} +case 1: { -lean_object* x_88; lean_object* x_89; uint8_t x_90; -x_88 = lean_ctor_get(x_87, 0); +lean_object* x_88; +x_88 = lean_ctor_get(x_67, 0); lean_inc(x_88); -x_89 = lean_ctor_get(x_87, 1); +if (lean_obj_tag(x_88) == 0) +{ +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; uint8_t x_93; +x_89 = lean_ctor_get(x_65, 1); lean_inc(x_89); -lean_dec(x_87); -x_90 = 1; -x_14 = x_90; -x_15 = x_88; -x_16 = x_89; -goto block_83; +lean_dec(x_65); +x_90 = lean_ctor_get(x_66, 1); +lean_inc(x_90); +lean_dec(x_66); +x_91 = lean_ctor_get(x_67, 1); +lean_inc(x_91); +lean_dec(x_67); +x_92 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__1; +x_93 = lean_string_dec_eq(x_91, x_92); +lean_dec(x_91); +if (x_93 == 0) +{ +lean_dec(x_90); +lean_dec(x_89); +lean_dec(x_59); +x_2 = x_60; +goto _start; } else { -lean_object* x_91; lean_object* x_92; uint8_t x_93; -x_91 = lean_ctor_get(x_87, 0); -lean_inc(x_91); -x_92 = lean_ctor_get(x_87, 1); -lean_inc(x_92); -lean_dec(x_87); -x_93 = 0; -x_14 = x_93; -x_15 = x_91; -x_16 = x_92; -goto block_83; +lean_object* x_95; uint8_t x_96; +x_95 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__3; +x_96 = lean_string_dec_eq(x_90, x_95); +lean_dec(x_90); +if (x_96 == 0) +{ +lean_dec(x_89); +lean_dec(x_59); +x_2 = x_60; +goto _start; } -block_83: +else { -lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_17 = lean_ctor_get(x_11, 10); -lean_inc(x_17); -lean_dec(x_11); -x_18 = lean_st_ref_get(x_12, x_16); -x_19 = !lean_is_exclusive(x_18); -if (x_19 == 0) +lean_object* x_98; uint8_t x_99; +x_98 = l_List_filterMap___at_Lean_Elab_Term_resolveParserName___spec__1___closed__2; +x_99 = lean_string_dec_eq(x_89, x_98); +if (x_99 == 0) { -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; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_20 = lean_ctor_get(x_18, 0); -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -lean_dec(x_20); -x_22 = lean_environment_main_module(x_21); -x_23 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__4; -x_24 = l_Lean_addMacroScope(x_22, x_23, x_17); -x_25 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__3; -x_26 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__7; -x_27 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_27, 0, x_15); -lean_ctor_set(x_27, 1, x_25); -lean_ctor_set(x_27, 2, x_24); -lean_ctor_set(x_27, 3, x_26); -x_28 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__8; -x_29 = lean_array_push(x_28, x_1); -x_30 = lean_array_push(x_29, x_2); -x_31 = lean_array_push(x_30, x_4); -x_32 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; -x_33 = lean_array_push(x_32, x_27); -if (x_14 == 0) +uint8_t x_100; +x_100 = lean_string_dec_eq(x_89, x_95); +lean_dec(x_89); +if (x_100 == 0) { -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; -x_34 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__9; -x_35 = lean_array_push(x_31, x_34); -x_36 = lean_box(2); -x_37 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_38 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -lean_ctor_set(x_38, 2, x_35); -x_39 = lean_array_push(x_33, x_38); -x_40 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_41 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_41, 0, x_36); -lean_ctor_set(x_41, 1, x_40); -lean_ctor_set(x_41, 2, x_39); -lean_ctor_set(x_18, 0, x_41); -return x_18; +lean_dec(x_59); +x_2 = x_60; +goto _start; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_42 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__12; -x_43 = lean_array_push(x_31, x_42); -x_44 = lean_box(2); -x_45 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_array_push(x_33, x_46); -x_48 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_49 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_49, 0, x_44); -lean_ctor_set(x_49, 1, x_48); -lean_ctor_set(x_49, 2, x_47); -lean_ctor_set(x_18, 0, x_49); -return x_18; +uint8_t x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_102 = 0; +x_103 = lean_box(x_102); +x_104 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_104, 0, x_59); +lean_ctor_set(x_104, 1, x_103); +x_105 = l_List_filterMap___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__15(x_1, x_60); +x_106 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_106, 0, x_104); +lean_ctor_set(x_106, 1, x_105); +return x_106; } } else { -lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_50 = lean_ctor_get(x_18, 0); -x_51 = lean_ctor_get(x_18, 1); -lean_inc(x_51); -lean_inc(x_50); -lean_dec(x_18); -x_52 = lean_ctor_get(x_50, 0); -lean_inc(x_52); -lean_dec(x_50); -x_53 = lean_environment_main_module(x_52); -x_54 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__4; -x_55 = l_Lean_addMacroScope(x_53, x_54, x_17); -x_56 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__3; -x_57 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__7; -x_58 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_58, 0, x_15); -lean_ctor_set(x_58, 1, x_56); -lean_ctor_set(x_58, 2, x_55); -lean_ctor_set(x_58, 3, x_57); -x_59 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__8; -x_60 = lean_array_push(x_59, x_1); -x_61 = lean_array_push(x_60, x_2); -x_62 = lean_array_push(x_61, x_4); -x_63 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; -x_64 = lean_array_push(x_63, x_58); -if (x_14 == 0) +uint8_t x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; +lean_dec(x_89); +x_107 = 0; +x_108 = lean_box(x_107); +x_109 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_109, 0, x_59); +lean_ctor_set(x_109, 1, x_108); +x_110 = l_List_filterMap___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__15(x_1, x_60); +x_111 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_111, 0, x_109); +lean_ctor_set(x_111, 1, x_110); +return x_111; +} +} +} +} +else { -lean_object* x_65; lean_object* x_66; 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; -x_65 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__9; -x_66 = lean_array_push(x_62, x_65); -x_67 = lean_box(2); -x_68 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_69 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_69, 0, x_67); -lean_ctor_set(x_69, 1, x_68); -lean_ctor_set(x_69, 2, x_66); -x_70 = lean_array_push(x_64, x_69); -x_71 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_72 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_72, 0, x_67); -lean_ctor_set(x_72, 1, x_71); -lean_ctor_set(x_72, 2, x_70); -x_73 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_51); -return x_73; +lean_dec(x_88); +lean_dec(x_67); +lean_dec(x_66); +lean_dec(x_65); +lean_dec(x_59); +x_2 = x_60; +goto _start; +} +} +default: +{ +lean_dec(x_67); +lean_dec(x_66); +lean_dec(x_65); +lean_dec(x_59); +x_2 = x_60; +goto _start; +} +} } else { -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; -x_74 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__12; -x_75 = lean_array_push(x_62, x_74); -x_76 = lean_box(2); -x_77 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_78 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_78, 0, x_76); -lean_ctor_set(x_78, 1, x_77); -lean_ctor_set(x_78, 2, x_75); -x_79 = lean_array_push(x_64, x_78); -x_80 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_81 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_81, 0, x_76); -lean_ctor_set(x_81, 1, x_80); -lean_ctor_set(x_81, 2, x_79); -x_82 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_82, 0, x_81); -lean_ctor_set(x_82, 1, x_51); -return x_82; +lean_dec(x_66); +lean_dec(x_65); +lean_dec(x_59); +x_2 = x_60; +goto _start; +} +} +else +{ +lean_dec(x_65); +lean_dec(x_59); +x_2 = x_60; +goto _start; } } +else +{ +lean_dec(x_64); +lean_dec(x_59); +x_2 = x_60; +goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy1(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_object* x_9, lean_object* x_10) { +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveParserName___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__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, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; -x_11 = lean_unsigned_to_nat(1u); -x_12 = l_Lean_Syntax_getArg(x_1, x_11); -x_13 = lean_ctor_get(x_2, 0); -lean_inc(x_13); -x_14 = lean_ctor_get_uint8(x_2, sizeof(void*)*1 + 2); -x_15 = 0; -x_16 = lean_alloc_ctor(0, 1, 3); -lean_ctor_set(x_16, 0, x_13); -lean_ctor_set_uint8(x_16, sizeof(void*)*1, x_15); -lean_ctor_set_uint8(x_16, sizeof(void*)*1 + 1, x_15); -lean_ctor_set_uint8(x_16, sizeof(void*)*1 + 2, x_14); +lean_object* x_11; lean_object* x_12; +x_11 = lean_box(0); lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_17 = l_Lean_Elab_Term_toParserDescr_process(x_12, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_17) == 0) +x_12 = l_Lean_Elab_resolveGlobalConstWithInfos___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__2(x_1, x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = lean_st_ref_get(x_9, x_14); +lean_dec(x_9); +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_15, 0); x_18 = lean_ctor_get(x_17, 0); lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); -lean_inc(x_19); lean_dec(x_17); -x_20 = lean_unsigned_to_nat(3u); -x_21 = l_Lean_Syntax_getArg(x_1, x_20); -x_22 = lean_unsigned_to_nat(4u); -x_23 = l_Lean_Syntax_getArg(x_1, x_22); -x_24 = l_Lean_Syntax_isNone(x_23); -if (x_24 == 0) -{ -lean_object* x_25; lean_object* x_26; -x_25 = l_Lean_Syntax_getArg(x_23, x_11); -lean_dec(x_23); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); -x_26 = l_Lean_Elab_Term_toParserDescr_process(x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_19); -if (lean_obj_tag(x_26) == 0) +x_19 = l_List_filterMap___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__15(x_18, x_13); +lean_ctor_set(x_15, 0, x_19); +return x_15; +} +else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -x_29 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1(x_18, x_21, x_1, x_27, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_28); -lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_29; +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_20 = lean_ctor_get(x_15, 0); +x_21 = lean_ctor_get(x_15, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_15); +x_22 = lean_ctor_get(x_20, 0); +lean_inc(x_22); +lean_dec(x_20); +x_23 = l_List_filterMap___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__15(x_22, x_13); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_21); +return x_24; +} } else { -uint8_t x_30; -lean_dec(x_21); -lean_dec(x_18); +uint8_t x_25; lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_30 = !lean_is_exclusive(x_26); -if (x_30 == 0) +x_25 = !lean_is_exclusive(x_12); +if (x_25 == 0) { -return x_26; +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_12, 0); +lean_dec(x_26); +x_27 = lean_box(0); +lean_ctor_set_tag(x_12, 0); +lean_ctor_set(x_12, 0, x_27); +return x_12; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_26, 0); -x_32 = lean_ctor_get(x_26, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_26); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -return x_33; +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_12, 1); +lean_inc(x_28); +lean_dec(x_12); +x_29 = lean_box(0); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_28); +return x_30; } } } +} +LEAN_EXPORT lean_object* l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; +x_3 = l_List_reverse___rarg(x_2); +return x_3; +} else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; -lean_dec(x_23); -lean_inc(x_8); -x_34 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_toParserDescr_processNonReserved___spec__2___rarg(x_8, x_9, x_19); -x_35 = lean_ctor_get(x_34, 0); -lean_inc(x_35); -x_36 = lean_ctor_get(x_34, 1); -lean_inc(x_36); -lean_dec(x_34); -x_37 = lean_ctor_get(x_8, 10); -lean_inc(x_37); -x_38 = lean_st_ref_get(x_9, x_36); -x_39 = lean_ctor_get(x_38, 0); -lean_inc(x_39); -x_40 = lean_ctor_get(x_38, 1); -lean_inc(x_40); -lean_dec(x_38); -x_41 = lean_ctor_get(x_39, 0); -lean_inc(x_41); -lean_dec(x_39); -x_42 = lean_environment_main_module(x_41); -x_43 = l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__5; -x_44 = l_Lean_addMacroScope(x_42, x_43, x_37); -x_45 = l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__3; -x_46 = l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__8; -x_47 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_47, 0, x_35); -lean_ctor_set(x_47, 1, x_45); -lean_ctor_set(x_47, 2, x_44); -lean_ctor_set(x_47, 3, x_46); -x_48 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; -lean_inc(x_21); -x_49 = lean_array_push(x_48, x_21); -x_50 = lean_box(2); -x_51 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_52 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_52, 0, x_50); -lean_ctor_set(x_52, 1, x_51); -lean_ctor_set(x_52, 2, x_49); -x_53 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; -x_54 = lean_array_push(x_53, x_47); -x_55 = lean_array_push(x_54, x_52); -x_56 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_57 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_57, 0, x_50); -lean_ctor_set(x_57, 1, x_56); -lean_ctor_set(x_57, 2, x_55); -x_58 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1(x_18, x_21, x_1, x_57, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_40); -lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_6); +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; +x_5 = lean_ctor_get(x_1, 0); +x_6 = lean_ctor_get(x_1, 1); +x_7 = lean_ctor_get(x_5, 0); +lean_inc(x_7); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_58; +lean_ctor_set(x_1, 1, x_2); +lean_ctor_set(x_1, 0, x_7); +{ +lean_object* _tmp_0 = x_6; +lean_object* _tmp_1 = x_1; +x_1 = _tmp_0; +x_2 = _tmp_1; } +goto _start; } else { -uint8_t x_59; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_1, 0); +x_10 = lean_ctor_get(x_1, 1); +lean_inc(x_10); +lean_inc(x_9); lean_dec(x_1); -x_59 = !lean_is_exclusive(x_17); -if (x_59 == 0) +x_11 = lean_ctor_get(x_9, 0); +lean_inc(x_11); +lean_dec(x_9); +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_2); +x_1 = x_10; +x_2 = x_12; +goto _start; +} +} +} +} +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__1() { +_start: { -return x_17; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("unknown parser declaration/category/alias '", 43); +return x_1; } -else +} +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__2() { +_start: { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_17, 0); -x_61 = lean_ctor_get(x_17, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_17); -x_62 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_62, 0, x_60); -lean_ctor_set(x_62, 1, x_61); -return x_62; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__1() { +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__4() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("ParserDescr.sepBy", 17); +x_1 = lean_mk_string_from_bytes("ParserDescr.parser", 18); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__2() { +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__1; +x_1 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__4; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__3() { +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__1; +x_1 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__4; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__2; +x_3 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__5; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -9893,276 +10935,81 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__4() { +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("parser", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__13; -x_2 = l_Lean_Elab_Term_toParserDescr_process___closed__6; +x_2 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__7; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__5() { +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__16; -x_2 = l_Lean_Elab_Term_toParserDescr_process___closed__6; +x_2 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__7; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__6() { +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__5; +x_2 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__9; 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_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__7() { +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__6; +x_2 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__10; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy___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, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__12() { _start: { -uint8_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_84; lean_object* x_85; uint8_t x_86; lean_object* x_87; -x_84 = lean_unsigned_to_nat(5u); -x_85 = l_Lean_Syntax_getArg(x_3, x_84); -x_86 = l_Lean_Syntax_isNone(x_85); -lean_dec(x_85); -lean_inc(x_11); -x_87 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_toParserDescr_processNonReserved___spec__2___rarg(x_11, x_12, x_13); -if (x_86 == 0) -{ -lean_object* x_88; lean_object* x_89; uint8_t x_90; -x_88 = lean_ctor_get(x_87, 0); -lean_inc(x_88); -x_89 = lean_ctor_get(x_87, 1); -lean_inc(x_89); -lean_dec(x_87); -x_90 = 1; -x_14 = x_90; -x_15 = x_88; -x_16 = x_89; -goto block_83; -} -else -{ -lean_object* x_91; lean_object* x_92; uint8_t x_93; -x_91 = lean_ctor_get(x_87, 0); -lean_inc(x_91); -x_92 = lean_ctor_get(x_87, 1); -lean_inc(x_92); -lean_dec(x_87); -x_93 = 0; -x_14 = x_93; -x_15 = x_91; -x_16 = x_92; -goto block_83; -} -block_83: -{ -lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_17 = lean_ctor_get(x_11, 10); -lean_inc(x_17); -lean_dec(x_11); -x_18 = lean_st_ref_get(x_12, x_16); -x_19 = !lean_is_exclusive(x_18); -if (x_19 == 0) -{ -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; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_20 = lean_ctor_get(x_18, 0); -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -lean_dec(x_20); -x_22 = lean_environment_main_module(x_21); -x_23 = l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__4; -x_24 = l_Lean_addMacroScope(x_22, x_23, x_17); -x_25 = l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__3; -x_26 = l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__7; -x_27 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_27, 0, x_15); -lean_ctor_set(x_27, 1, x_25); -lean_ctor_set(x_27, 2, x_24); -lean_ctor_set(x_27, 3, x_26); -x_28 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__8; -x_29 = lean_array_push(x_28, x_1); -x_30 = lean_array_push(x_29, x_2); -x_31 = lean_array_push(x_30, x_4); -x_32 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; -x_33 = lean_array_push(x_32, x_27); -if (x_14 == 0) -{ -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; -x_34 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__9; -x_35 = lean_array_push(x_31, x_34); -x_36 = lean_box(2); -x_37 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_38 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -lean_ctor_set(x_38, 2, x_35); -x_39 = lean_array_push(x_33, x_38); -x_40 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_41 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_41, 0, x_36); -lean_ctor_set(x_41, 1, x_40); -lean_ctor_set(x_41, 2, x_39); -lean_ctor_set(x_18, 0, x_41); -return x_18; -} -else -{ -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_42 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__12; -x_43 = lean_array_push(x_31, x_42); -x_44 = lean_box(2); -x_45 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = lean_array_push(x_33, x_46); -x_48 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_49 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_49, 0, x_44); -lean_ctor_set(x_49, 1, x_48); -lean_ctor_set(x_49, 2, x_47); -lean_ctor_set(x_18, 0, x_49); -return x_18; -} +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("ambiguous parser declaration ", 29); +return x_1; } -else -{ -lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_50 = lean_ctor_get(x_18, 0); -x_51 = lean_ctor_get(x_18, 1); -lean_inc(x_51); -lean_inc(x_50); -lean_dec(x_18); -x_52 = lean_ctor_get(x_50, 0); -lean_inc(x_52); -lean_dec(x_50); -x_53 = lean_environment_main_module(x_52); -x_54 = l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__4; -x_55 = l_Lean_addMacroScope(x_53, x_54, x_17); -x_56 = l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__3; -x_57 = l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__7; -x_58 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_58, 0, x_15); -lean_ctor_set(x_58, 1, x_56); -lean_ctor_set(x_58, 2, x_55); -lean_ctor_set(x_58, 3, x_57); -x_59 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__8; -x_60 = lean_array_push(x_59, x_1); -x_61 = lean_array_push(x_60, x_2); -x_62 = lean_array_push(x_61, x_4); -x_63 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; -x_64 = lean_array_push(x_63, x_58); -if (x_14 == 0) -{ -lean_object* x_65; lean_object* x_66; 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; -x_65 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__9; -x_66 = lean_array_push(x_62, x_65); -x_67 = lean_box(2); -x_68 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_69 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_69, 0, x_67); -lean_ctor_set(x_69, 1, x_68); -lean_ctor_set(x_69, 2, x_66); -x_70 = lean_array_push(x_64, x_69); -x_71 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_72 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_72, 0, x_67); -lean_ctor_set(x_72, 1, x_71); -lean_ctor_set(x_72, 2, x_70); -x_73 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_51); -return x_73; } -else +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__13() { +_start: { -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; -x_74 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__12; -x_75 = lean_array_push(x_62, x_74); -x_76 = lean_box(2); -x_77 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_78 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_78, 0, x_76); -lean_ctor_set(x_78, 1, x_77); -lean_ctor_set(x_78, 2, x_75); -x_79 = lean_array_push(x_64, x_78); -x_80 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_81 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_81, 0, x_76); -lean_ctor_set(x_81, 1, x_80); -lean_ctor_set(x_81, 2, x_79); -x_82 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_82, 0, x_81); -lean_ctor_set(x_82, 1, x_51); -return x_82; -} -} -} +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__12; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy(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_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processNullaryOrCat(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_object* x_9, lean_object* x_10) { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; -x_11 = lean_unsigned_to_nat(1u); +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = lean_unsigned_to_nat(0u); x_12 = l_Lean_Syntax_getArg(x_1, x_11); -x_13 = lean_ctor_get(x_2, 0); -lean_inc(x_13); -x_14 = lean_ctor_get_uint8(x_2, sizeof(void*)*1 + 2); -x_15 = 0; -x_16 = lean_alloc_ctor(0, 1, 3); -lean_ctor_set(x_16, 0, x_13); -lean_ctor_set_uint8(x_16, sizeof(void*)*1, x_15); -lean_ctor_set_uint8(x_16, sizeof(void*)*1 + 1, x_15); -lean_ctor_set_uint8(x_16, sizeof(void*)*1 + 2, x_14); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_17 = l_Lean_Elab_Term_toParserDescr_process(x_12, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_17) == 0) -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); -lean_inc(x_19); -lean_dec(x_17); -x_20 = lean_unsigned_to_nat(3u); -x_21 = l_Lean_Syntax_getArg(x_1, x_20); -x_22 = lean_unsigned_to_nat(4u); -x_23 = l_Lean_Syntax_getArg(x_1, x_22); -x_24 = l_Lean_Syntax_isNone(x_23); -if (x_24 == 0) -{ -lean_object* x_25; lean_object* x_26; -x_25 = l_Lean_Syntax_getArg(x_23, x_11); -lean_dec(x_23); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); @@ -10171,124 +11018,102 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_26 = l_Lean_Elab_Term_toParserDescr_process(x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_19); -if (lean_obj_tag(x_26) == 0) -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -x_29 = l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1(x_18, x_21, x_1, x_27, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_28); -lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_29; -} -else +lean_inc(x_12); +x_13 = l_Lean_Elab_Term_resolveParserName___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__1(x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +if (lean_obj_tag(x_14) == 0) { -uint8_t x_30; -lean_dec(x_21); +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; uint8_t x_22; +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_Syntax_getId(x_12); +x_17 = lean_erase_macro_scopes(x_16); +x_18 = lean_st_ref_get(x_9, x_15); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); lean_dec(x_18); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_30 = !lean_is_exclusive(x_26); -if (x_30 == 0) -{ -return x_26; -} -else -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_26, 0); -x_32 = lean_ctor_get(x_26, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_26); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -return x_33; -} -} -} -else +x_21 = lean_ctor_get(x_19, 0); +lean_inc(x_21); +lean_dec(x_19); +lean_inc(x_17); +x_22 = l_Lean_Parser_isParserCategory(x_21, x_17); +lean_dec(x_21); +if (x_22 == 0) { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_23 = lean_st_ref_get(x_9, x_20); +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); lean_dec(x_23); -lean_inc(x_8); -x_34 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_toParserDescr_processNonReserved___spec__2___rarg(x_8, x_9, x_19); -x_35 = lean_ctor_get(x_34, 0); -lean_inc(x_35); -x_36 = lean_ctor_get(x_34, 1); -lean_inc(x_36); -lean_dec(x_34); -x_37 = lean_ctor_get(x_8, 10); -lean_inc(x_37); -x_38 = lean_st_ref_get(x_9, x_36); -x_39 = lean_ctor_get(x_38, 0); -lean_inc(x_39); -x_40 = lean_ctor_get(x_38, 1); -lean_inc(x_40); -lean_dec(x_38); -x_41 = lean_ctor_get(x_39, 0); -lean_inc(x_41); -lean_dec(x_39); -x_42 = lean_environment_main_module(x_41); -x_43 = l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__5; -x_44 = l_Lean_addMacroScope(x_42, x_43, x_37); -x_45 = l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__3; -x_46 = l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__8; -x_47 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_47, 0, x_35); -lean_ctor_set(x_47, 1, x_45); -lean_ctor_set(x_47, 2, x_44); -lean_ctor_set(x_47, 3, x_46); -x_48 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; -lean_inc(x_21); -x_49 = lean_array_push(x_48, x_21); -x_50 = lean_box(2); -x_51 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_52 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_52, 0, x_50); -lean_ctor_set(x_52, 1, x_51); -lean_ctor_set(x_52, 2, x_49); -x_53 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; -x_54 = lean_array_push(x_53, x_47); -x_55 = lean_array_push(x_54, x_52); -x_56 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_57 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_57, 0, x_50); -lean_ctor_set(x_57, 1, x_56); -lean_ctor_set(x_57, 2, x_55); -x_58 = l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1(x_18, x_21, x_1, x_57, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_40); +x_25 = l_Lean_Parser_isParserAlias(x_17, x_24); +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +x_27 = lean_unbox(x_26); +lean_dec(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; lean_object* x_33; lean_object* x_34; +lean_dec(x_12); +lean_dec(x_1); +x_28 = lean_ctor_get(x_25, 1); +lean_inc(x_28); +lean_dec(x_25); +x_29 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_29, 0, x_17); +x_30 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__2; +x_31 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_29); +x_32 = l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__4; +x_33 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +x_34 = l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_process___spec__7(x_33, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_28); lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -return x_58; +return x_34; } +else +{ +lean_object* x_35; lean_object* x_36; +lean_dec(x_17); +x_35 = lean_ctor_get(x_25, 1); +lean_inc(x_35); +lean_dec(x_25); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_36 = l_Lean_Elab_Term_toParserDescr_ensureNoPrec(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_35); +lean_dec(x_1); +if (lean_obj_tag(x_36) == 0) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_36, 1); +lean_inc(x_37); +lean_dec(x_36); +x_38 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__3; +x_39 = l_Lean_Elab_Term_toParserDescr_processAlias(x_12, x_38, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_37); +return x_39; } else { -uint8_t x_59; +uint8_t x_40; +lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -10297,522 +11122,448 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_59 = !lean_is_exclusive(x_17); -if (x_59 == 0) +x_40 = !lean_is_exclusive(x_36); +if (x_40 == 0) { -return x_17; +return x_36; } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_17, 0); -x_61 = lean_ctor_get(x_17, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_17); -x_62 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_62, 0, x_60); -lean_ctor_set(x_62, 1, x_61); -return x_62; +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_36, 0); +x_42 = lean_ctor_get(x_36, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_36); +x_43 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +return x_43; } } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processBinary(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_object* x_9, lean_object* x_10) { -_start: +else { -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_11 = lean_unsigned_to_nat(0u); -x_12 = l_Lean_Syntax_getArg(x_1, x_11); -x_13 = l_Lean_Syntax_getId(x_12); +lean_object* x_44; +lean_dec(x_17); lean_dec(x_12); -x_14 = lean_erase_macro_scopes(x_13); -x_15 = lean_st_ref_get(x_9, x_10); -x_16 = lean_ctor_get(x_15, 1); -lean_inc(x_16); -lean_dec(x_15); -x_17 = lean_ctor_get(x_8, 5); -lean_inc(x_17); -lean_inc(x_14); -x_18 = l_Lean_Parser_ensureBinaryParserAlias(x_14, x_16); -if (lean_obj_tag(x_18) == 0) +x_44 = l_Lean_Elab_Term_toParserDescr_processParserCategory(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_20); +return x_44; +} +} +else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -lean_dec(x_17); -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -lean_dec(x_18); -x_20 = lean_unsigned_to_nat(2u); -x_21 = l_Lean_Syntax_getArg(x_1, x_20); -x_22 = !lean_is_exclusive(x_2); -if (x_22 == 0) +lean_object* x_45; lean_object* x_46; uint8_t x_47; +lean_dec(x_12); +x_45 = lean_ctor_get(x_14, 0); +lean_inc(x_45); +x_46 = lean_ctor_get(x_45, 1); +lean_inc(x_46); +x_47 = lean_unbox(x_46); +lean_dec(x_46); +if (x_47 == 0) { -uint8_t x_23; lean_object* x_24; -x_23 = 0; -lean_ctor_set_uint8(x_2, sizeof(void*)*1, x_23); -lean_ctor_set_uint8(x_2, sizeof(void*)*1 + 1, x_23); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); -x_24 = l_Lean_Elab_Term_toParserDescr_process(x_21, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_19); -if (lean_obj_tag(x_24) == 0) +lean_object* x_48; +x_48 = lean_ctor_get(x_14, 1); +lean_inc(x_48); +if (lean_obj_tag(x_48) == 0) { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -x_26 = lean_ctor_get(x_24, 1); -lean_inc(x_26); -lean_dec(x_24); -x_27 = lean_unsigned_to_nat(4u); -x_28 = l_Lean_Syntax_getArg(x_1, x_27); +lean_object* x_49; uint8_t x_50; +lean_dec(x_14); +x_49 = lean_ctor_get(x_13, 1); +lean_inc(x_49); +lean_dec(x_13); +x_50 = !lean_is_exclusive(x_45); +if (x_50 == 0) +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_45, 0); +x_52 = lean_ctor_get(x_45, 1); +lean_dec(x_52); lean_inc(x_9); lean_inc(x_8); -x_29 = l_Lean_Elab_Term_toParserDescr_process(x_28, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_26); -if (lean_obj_tag(x_29) == 0) +x_53 = l_Lean_Elab_Term_toParserDescr_ensureNoPrec(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_49); +lean_dec(x_1); +if (lean_obj_tag(x_53) == 0) { -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; uint8_t x_37; -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_29, 1); -lean_inc(x_31); -lean_dec(x_29); +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; uint8_t x_60; +x_54 = lean_ctor_get(x_53, 1); +lean_inc(x_54); +lean_dec(x_53); lean_inc(x_8); -x_32 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_toParserDescr_processNonReserved___spec__2___rarg(x_8, x_9, x_31); -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_32, 1); -lean_inc(x_34); -lean_dec(x_32); -x_35 = lean_ctor_get(x_8, 10); -lean_inc(x_35); +x_55 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_toParserDescr_processNonReserved___spec__2___rarg(x_8, x_9, x_54); +x_56 = lean_ctor_get(x_55, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_55, 1); +lean_inc(x_57); +lean_dec(x_55); +x_58 = lean_ctor_get(x_8, 10); +lean_inc(x_58); lean_dec(x_8); -x_36 = lean_st_ref_get(x_9, x_34); +x_59 = lean_st_ref_get(x_9, x_57); lean_dec(x_9); -x_37 = !lean_is_exclusive(x_36); -if (x_37 == 0) -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_38 = lean_ctor_get(x_36, 0); -x_39 = lean_ctor_get(x_38, 0); -lean_inc(x_39); -lean_dec(x_38); -x_40 = lean_environment_main_module(x_39); -x_41 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__15; -x_42 = l_Lean_addMacroScope(x_40, x_41, x_35); -x_43 = lean_box(0); -x_44 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__11; -x_45 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__19; -x_46 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_46, 0, x_33); -lean_ctor_set(x_46, 1, x_44); -lean_ctor_set(x_46, 2, x_42); -lean_ctor_set(x_46, 3, x_45); -lean_inc(x_14); -x_47 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_43, x_14); -x_48 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; -x_49 = lean_array_push(x_48, x_46); -if (lean_obj_tag(x_47) == 0) -{ -lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_50 = l___private_Init_Meta_0__Lean_quoteNameMk(x_14); -x_51 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__28; -x_52 = lean_array_push(x_51, x_50); -x_53 = lean_array_push(x_52, x_25); -x_54 = lean_array_push(x_53, x_30); -x_55 = lean_box(2); -x_56 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_57 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_57, 0, x_55); -lean_ctor_set(x_57, 1, x_56); -lean_ctor_set(x_57, 2, x_54); -x_58 = lean_array_push(x_49, x_57); -x_59 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_60 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_60, 0, x_55); -lean_ctor_set(x_60, 1, x_59); -lean_ctor_set(x_60, 2, x_58); -lean_ctor_set(x_36, 0, x_60); -return x_36; -} -else +x_60 = !lean_is_exclusive(x_59); +if (x_60 == 0) { -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; -lean_dec(x_14); -x_61 = lean_ctor_get(x_47, 0); -lean_inc(x_61); -lean_dec(x_47); -x_62 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__8; -x_63 = l_String_intercalate(x_62, x_61); -x_64 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__9; -x_65 = lean_string_append(x_64, x_63); -lean_dec(x_63); -x_66 = l_Lean_nameLitKind; -x_67 = lean_box(2); -x_68 = l_Lean_Syntax_mkLit(x_66, x_65, x_67); -x_69 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; -x_70 = lean_array_push(x_69, x_68); -x_71 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__23; -x_72 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_72, 0, x_67); -lean_ctor_set(x_72, 1, x_71); -lean_ctor_set(x_72, 2, x_70); -x_73 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__28; -x_74 = lean_array_push(x_73, x_72); -x_75 = lean_array_push(x_74, x_25); -x_76 = lean_array_push(x_75, x_30); +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_61 = lean_ctor_get(x_59, 0); +x_62 = lean_ctor_get(x_61, 0); +lean_inc(x_62); +lean_dec(x_61); +x_63 = lean_environment_main_module(x_62); +x_64 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__8; +x_65 = l_Lean_addMacroScope(x_63, x_64, x_58); +x_66 = lean_box(0); +x_67 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__6; +x_68 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__11; +x_69 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_69, 0, x_56); +lean_ctor_set(x_69, 1, x_67); +lean_ctor_set(x_69, 2, x_65); +lean_ctor_set(x_69, 3, x_68); +lean_inc(x_51); +x_70 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_66, x_51); +x_71 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; +x_72 = lean_array_push(x_71, x_69); +if (lean_obj_tag(x_70) == 0) +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_73 = l_Lean_quoteNameMk(x_51); +x_74 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; +x_75 = lean_array_push(x_74, x_73); +x_76 = lean_box(2); x_77 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; x_78 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_78, 0, x_67); +lean_ctor_set(x_78, 0, x_76); lean_ctor_set(x_78, 1, x_77); -lean_ctor_set(x_78, 2, x_76); -x_79 = lean_array_push(x_49, x_78); +lean_ctor_set(x_78, 2, x_75); +x_79 = lean_array_push(x_72, x_78); x_80 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; x_81 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_81, 0, x_67); +lean_ctor_set(x_81, 0, x_76); lean_ctor_set(x_81, 1, x_80); lean_ctor_set(x_81, 2, x_79); -lean_ctor_set(x_36, 0, x_81); -return x_36; -} +x_82 = lean_unsigned_to_nat(1u); +lean_ctor_set(x_45, 1, x_82); +lean_ctor_set(x_45, 0, x_81); +lean_ctor_set(x_59, 0, x_45); +return x_59; } else { -lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; -x_82 = lean_ctor_get(x_36, 0); -x_83 = lean_ctor_get(x_36, 1); +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; +lean_dec(x_51); +x_83 = lean_ctor_get(x_70, 0); lean_inc(x_83); -lean_inc(x_82); -lean_dec(x_36); -x_84 = lean_ctor_get(x_82, 0); -lean_inc(x_84); -lean_dec(x_82); -x_85 = lean_environment_main_module(x_84); -x_86 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__15; -x_87 = l_Lean_addMacroScope(x_85, x_86, x_35); -x_88 = lean_box(0); -x_89 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__11; -x_90 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__19; -x_91 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_91, 0, x_33); -lean_ctor_set(x_91, 1, x_89); -lean_ctor_set(x_91, 2, x_87); -lean_ctor_set(x_91, 3, x_90); -lean_inc(x_14); -x_92 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_88, x_14); -x_93 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; -x_94 = lean_array_push(x_93, x_91); -if (lean_obj_tag(x_92) == 0) -{ -lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; -x_95 = l___private_Init_Meta_0__Lean_quoteNameMk(x_14); -x_96 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__28; -x_97 = lean_array_push(x_96, x_95); -x_98 = lean_array_push(x_97, x_25); -x_99 = lean_array_push(x_98, x_30); -x_100 = lean_box(2); -x_101 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_102 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_102, 0, x_100); -lean_ctor_set(x_102, 1, x_101); -lean_ctor_set(x_102, 2, x_99); -x_103 = lean_array_push(x_94, x_102); -x_104 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_105 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_105, 0, x_100); -lean_ctor_set(x_105, 1, x_104); -lean_ctor_set(x_105, 2, x_103); -x_106 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_106, 0, x_105); -lean_ctor_set(x_106, 1, x_83); -return x_106; -} -else -{ -lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; -lean_dec(x_14); -x_107 = lean_ctor_get(x_92, 0); -lean_inc(x_107); -lean_dec(x_92); -x_108 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__8; -x_109 = l_String_intercalate(x_108, x_107); -x_110 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__9; -x_111 = lean_string_append(x_110, x_109); -lean_dec(x_109); -x_112 = l_Lean_nameLitKind; -x_113 = lean_box(2); -x_114 = l_Lean_Syntax_mkLit(x_112, x_111, x_113); -x_115 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; -x_116 = lean_array_push(x_115, x_114); -x_117 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__23; -x_118 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_118, 0, x_113); -lean_ctor_set(x_118, 1, x_117); -lean_ctor_set(x_118, 2, x_116); -x_119 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__28; -x_120 = lean_array_push(x_119, x_118); -x_121 = lean_array_push(x_120, x_25); -x_122 = lean_array_push(x_121, x_30); -x_123 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_124 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_124, 0, x_113); -lean_ctor_set(x_124, 1, x_123); -lean_ctor_set(x_124, 2, x_122); -x_125 = lean_array_push(x_94, x_124); -x_126 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_127 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_127, 0, x_113); -lean_ctor_set(x_127, 1, x_126); -lean_ctor_set(x_127, 2, x_125); -x_128 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_128, 0, x_127); -lean_ctor_set(x_128, 1, x_83); -return x_128; -} +lean_dec(x_70); +x_84 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__16; +x_85 = l_String_intercalate(x_84, x_83); +x_86 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__17; +x_87 = lean_string_append(x_86, x_85); +lean_dec(x_85); +x_88 = lean_box(2); +x_89 = l_Lean_Syntax_mkNameLit(x_87, x_88); +x_90 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; +x_91 = lean_array_push(x_90, x_89); +x_92 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__23; +x_93 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_93, 0, x_88); +lean_ctor_set(x_93, 1, x_92); +lean_ctor_set(x_93, 2, x_91); +x_94 = lean_array_push(x_90, x_93); +x_95 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; +x_96 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_96, 0, x_88); +lean_ctor_set(x_96, 1, x_95); +lean_ctor_set(x_96, 2, x_94); +x_97 = lean_array_push(x_72, x_96); +x_98 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +x_99 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_99, 0, x_88); +lean_ctor_set(x_99, 1, x_98); +lean_ctor_set(x_99, 2, x_97); +x_100 = lean_unsigned_to_nat(1u); +lean_ctor_set(x_45, 1, x_100); +lean_ctor_set(x_45, 0, x_99); +lean_ctor_set(x_59, 0, x_45); +return x_59; } } else { -uint8_t x_129; -lean_dec(x_25); -lean_dec(x_14); -lean_dec(x_9); -lean_dec(x_8); -x_129 = !lean_is_exclusive(x_29); -if (x_129 == 0) +lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; +x_101 = lean_ctor_get(x_59, 0); +x_102 = lean_ctor_get(x_59, 1); +lean_inc(x_102); +lean_inc(x_101); +lean_dec(x_59); +x_103 = lean_ctor_get(x_101, 0); +lean_inc(x_103); +lean_dec(x_101); +x_104 = lean_environment_main_module(x_103); +x_105 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__8; +x_106 = l_Lean_addMacroScope(x_104, x_105, x_58); +x_107 = lean_box(0); +x_108 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__6; +x_109 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__11; +x_110 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_110, 0, x_56); +lean_ctor_set(x_110, 1, x_108); +lean_ctor_set(x_110, 2, x_106); +lean_ctor_set(x_110, 3, x_109); +lean_inc(x_51); +x_111 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_107, x_51); +x_112 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; +x_113 = lean_array_push(x_112, x_110); +if (lean_obj_tag(x_111) == 0) { -return x_29; +lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; +x_114 = l_Lean_quoteNameMk(x_51); +x_115 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; +x_116 = lean_array_push(x_115, x_114); +x_117 = lean_box(2); +x_118 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; +x_119 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_119, 0, x_117); +lean_ctor_set(x_119, 1, x_118); +lean_ctor_set(x_119, 2, x_116); +x_120 = lean_array_push(x_113, x_119); +x_121 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +x_122 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_122, 0, x_117); +lean_ctor_set(x_122, 1, x_121); +lean_ctor_set(x_122, 2, x_120); +x_123 = lean_unsigned_to_nat(1u); +lean_ctor_set(x_45, 1, x_123); +lean_ctor_set(x_45, 0, x_122); +x_124 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_124, 0, x_45); +lean_ctor_set(x_124, 1, x_102); +return x_124; } else { -lean_object* x_130; lean_object* x_131; lean_object* x_132; -x_130 = lean_ctor_get(x_29, 0); -x_131 = lean_ctor_get(x_29, 1); -lean_inc(x_131); -lean_inc(x_130); -lean_dec(x_29); -x_132 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_132, 0, x_130); -lean_ctor_set(x_132, 1, x_131); -return x_132; +lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; +lean_dec(x_51); +x_125 = lean_ctor_get(x_111, 0); +lean_inc(x_125); +lean_dec(x_111); +x_126 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__16; +x_127 = l_String_intercalate(x_126, x_125); +x_128 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__17; +x_129 = lean_string_append(x_128, x_127); +lean_dec(x_127); +x_130 = lean_box(2); +x_131 = l_Lean_Syntax_mkNameLit(x_129, x_130); +x_132 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; +x_133 = lean_array_push(x_132, x_131); +x_134 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__23; +x_135 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_135, 0, x_130); +lean_ctor_set(x_135, 1, x_134); +lean_ctor_set(x_135, 2, x_133); +x_136 = lean_array_push(x_132, x_135); +x_137 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; +x_138 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_138, 0, x_130); +lean_ctor_set(x_138, 1, x_137); +lean_ctor_set(x_138, 2, x_136); +x_139 = lean_array_push(x_113, x_138); +x_140 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +x_141 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_141, 0, x_130); +lean_ctor_set(x_141, 1, x_140); +lean_ctor_set(x_141, 2, x_139); +x_142 = lean_unsigned_to_nat(1u); +lean_ctor_set(x_45, 1, x_142); +lean_ctor_set(x_45, 0, x_141); +x_143 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_143, 0, x_45); +lean_ctor_set(x_143, 1, x_102); +return x_143; } } } else { -uint8_t x_133; -lean_dec(x_2); -lean_dec(x_14); +uint8_t x_144; +lean_free_object(x_45); +lean_dec(x_51); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_133 = !lean_is_exclusive(x_24); -if (x_133 == 0) +x_144 = !lean_is_exclusive(x_53); +if (x_144 == 0) { -return x_24; +return x_53; } else { -lean_object* x_134; lean_object* x_135; lean_object* x_136; -x_134 = lean_ctor_get(x_24, 0); -x_135 = lean_ctor_get(x_24, 1); -lean_inc(x_135); -lean_inc(x_134); -lean_dec(x_24); -x_136 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_136, 0, x_134); -lean_ctor_set(x_136, 1, x_135); -return x_136; +lean_object* x_145; lean_object* x_146; lean_object* x_147; +x_145 = lean_ctor_get(x_53, 0); +x_146 = lean_ctor_get(x_53, 1); +lean_inc(x_146); +lean_inc(x_145); +lean_dec(x_53); +x_147 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_147, 0, x_145); +lean_ctor_set(x_147, 1, x_146); +return x_147; } } } else { -lean_object* x_137; uint8_t x_138; uint8_t x_139; lean_object* x_140; lean_object* x_141; -x_137 = lean_ctor_get(x_2, 0); -x_138 = lean_ctor_get_uint8(x_2, sizeof(void*)*1 + 2); -lean_inc(x_137); -lean_dec(x_2); -x_139 = 0; -x_140 = lean_alloc_ctor(0, 1, 3); -lean_ctor_set(x_140, 0, x_137); -lean_ctor_set_uint8(x_140, sizeof(void*)*1, x_139); -lean_ctor_set_uint8(x_140, sizeof(void*)*1 + 1, x_139); -lean_ctor_set_uint8(x_140, sizeof(void*)*1 + 2, x_138); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_140); -x_141 = l_Lean_Elab_Term_toParserDescr_process(x_21, x_140, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_19); -if (lean_obj_tag(x_141) == 0) -{ -lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; -x_142 = lean_ctor_get(x_141, 0); -lean_inc(x_142); -x_143 = lean_ctor_get(x_141, 1); -lean_inc(x_143); -lean_dec(x_141); -x_144 = lean_unsigned_to_nat(4u); -x_145 = l_Lean_Syntax_getArg(x_1, x_144); +lean_object* x_148; lean_object* x_149; +x_148 = lean_ctor_get(x_45, 0); +lean_inc(x_148); +lean_dec(x_45); lean_inc(x_9); lean_inc(x_8); -x_146 = l_Lean_Elab_Term_toParserDescr_process(x_145, x_140, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_143); -if (lean_obj_tag(x_146) == 0) +x_149 = l_Lean_Elab_Term_toParserDescr_ensureNoPrec(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_49); +lean_dec(x_1); +if (lean_obj_tag(x_149) == 0) { -lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; -x_147 = lean_ctor_get(x_146, 0); -lean_inc(x_147); -x_148 = lean_ctor_get(x_146, 1); -lean_inc(x_148); -lean_dec(x_146); -lean_inc(x_8); -x_149 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_toParserDescr_processNonReserved___spec__2___rarg(x_8, x_9, x_148); -x_150 = lean_ctor_get(x_149, 0); +lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; +x_150 = lean_ctor_get(x_149, 1); lean_inc(x_150); -x_151 = lean_ctor_get(x_149, 1); -lean_inc(x_151); lean_dec(x_149); -x_152 = lean_ctor_get(x_8, 10); +lean_inc(x_8); +x_151 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_toParserDescr_processNonReserved___spec__2___rarg(x_8, x_9, x_150); +x_152 = lean_ctor_get(x_151, 0); lean_inc(x_152); +x_153 = lean_ctor_get(x_151, 1); +lean_inc(x_153); +lean_dec(x_151); +x_154 = lean_ctor_get(x_8, 10); +lean_inc(x_154); lean_dec(x_8); -x_153 = lean_st_ref_get(x_9, x_151); +x_155 = lean_st_ref_get(x_9, x_153); lean_dec(x_9); -x_154 = lean_ctor_get(x_153, 0); -lean_inc(x_154); -x_155 = lean_ctor_get(x_153, 1); -lean_inc(x_155); -if (lean_is_exclusive(x_153)) { - lean_ctor_release(x_153, 0); - lean_ctor_release(x_153, 1); - x_156 = x_153; +x_156 = lean_ctor_get(x_155, 0); +lean_inc(x_156); +x_157 = lean_ctor_get(x_155, 1); +lean_inc(x_157); +if (lean_is_exclusive(x_155)) { + lean_ctor_release(x_155, 0); + lean_ctor_release(x_155, 1); + x_158 = x_155; } else { - lean_dec_ref(x_153); - x_156 = lean_box(0); + lean_dec_ref(x_155); + x_158 = lean_box(0); } -x_157 = lean_ctor_get(x_154, 0); -lean_inc(x_157); -lean_dec(x_154); -x_158 = lean_environment_main_module(x_157); -x_159 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__15; -x_160 = l_Lean_addMacroScope(x_158, x_159, x_152); -x_161 = lean_box(0); -x_162 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__11; -x_163 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__19; -x_164 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_164, 0, x_150); -lean_ctor_set(x_164, 1, x_162); -lean_ctor_set(x_164, 2, x_160); -lean_ctor_set(x_164, 3, x_163); -lean_inc(x_14); -x_165 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_161, x_14); -x_166 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; -x_167 = lean_array_push(x_166, x_164); -if (lean_obj_tag(x_165) == 0) -{ -lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; -x_168 = l___private_Init_Meta_0__Lean_quoteNameMk(x_14); -x_169 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__28; -x_170 = lean_array_push(x_169, x_168); -x_171 = lean_array_push(x_170, x_142); -x_172 = lean_array_push(x_171, x_147); +x_159 = lean_ctor_get(x_156, 0); +lean_inc(x_159); +lean_dec(x_156); +x_160 = lean_environment_main_module(x_159); +x_161 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__8; +x_162 = l_Lean_addMacroScope(x_160, x_161, x_154); +x_163 = lean_box(0); +x_164 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__6; +x_165 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__11; +x_166 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_166, 0, x_152); +lean_ctor_set(x_166, 1, x_164); +lean_ctor_set(x_166, 2, x_162); +lean_ctor_set(x_166, 3, x_165); +lean_inc(x_148); +x_167 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_163, x_148); +x_168 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; +x_169 = lean_array_push(x_168, x_166); +if (lean_obj_tag(x_167) == 0) +{ +lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; +x_170 = l_Lean_quoteNameMk(x_148); +x_171 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; +x_172 = lean_array_push(x_171, x_170); x_173 = lean_box(2); x_174 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; x_175 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_175, 0, x_173); lean_ctor_set(x_175, 1, x_174); lean_ctor_set(x_175, 2, x_172); -x_176 = lean_array_push(x_167, x_175); +x_176 = lean_array_push(x_169, x_175); x_177 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; x_178 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_178, 0, x_173); lean_ctor_set(x_178, 1, x_177); lean_ctor_set(x_178, 2, x_176); -if (lean_is_scalar(x_156)) { - x_179 = lean_alloc_ctor(0, 2, 0); +x_179 = lean_unsigned_to_nat(1u); +x_180 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_180, 0, x_178); +lean_ctor_set(x_180, 1, x_179); +if (lean_is_scalar(x_158)) { + x_181 = lean_alloc_ctor(0, 2, 0); } else { - x_179 = x_156; + x_181 = x_158; } -lean_ctor_set(x_179, 0, x_178); -lean_ctor_set(x_179, 1, x_155); -return x_179; +lean_ctor_set(x_181, 0, x_180); +lean_ctor_set(x_181, 1, x_157); +return x_181; } else { -lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; -lean_dec(x_14); -x_180 = lean_ctor_get(x_165, 0); -lean_inc(x_180); -lean_dec(x_165); -x_181 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__8; -x_182 = l_String_intercalate(x_181, x_180); -x_183 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__9; -x_184 = lean_string_append(x_183, x_182); -lean_dec(x_182); -x_185 = l_Lean_nameLitKind; -x_186 = lean_box(2); -x_187 = l_Lean_Syntax_mkLit(x_185, x_184, x_186); -x_188 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; -x_189 = lean_array_push(x_188, x_187); -x_190 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__23; -x_191 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_191, 0, x_186); -lean_ctor_set(x_191, 1, x_190); -lean_ctor_set(x_191, 2, x_189); -x_192 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__28; -x_193 = lean_array_push(x_192, x_191); -x_194 = lean_array_push(x_193, x_142); -x_195 = lean_array_push(x_194, x_147); -x_196 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_197 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_197, 0, x_186); -lean_ctor_set(x_197, 1, x_196); -lean_ctor_set(x_197, 2, x_195); -x_198 = lean_array_push(x_167, x_197); -x_199 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_200 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_200, 0, x_186); +lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; +lean_dec(x_148); +x_182 = lean_ctor_get(x_167, 0); +lean_inc(x_182); +lean_dec(x_167); +x_183 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__16; +x_184 = l_String_intercalate(x_183, x_182); +x_185 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__17; +x_186 = lean_string_append(x_185, x_184); +lean_dec(x_184); +x_187 = lean_box(2); +x_188 = l_Lean_Syntax_mkNameLit(x_186, x_187); +x_189 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; +x_190 = lean_array_push(x_189, x_188); +x_191 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__23; +x_192 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_192, 0, x_187); +lean_ctor_set(x_192, 1, x_191); +lean_ctor_set(x_192, 2, x_190); +x_193 = lean_array_push(x_189, x_192); +x_194 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; +x_195 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_195, 0, x_187); +lean_ctor_set(x_195, 1, x_194); +lean_ctor_set(x_195, 2, x_193); +x_196 = lean_array_push(x_169, x_195); +x_197 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +x_198 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_198, 0, x_187); +lean_ctor_set(x_198, 1, x_197); +lean_ctor_set(x_198, 2, x_196); +x_199 = lean_unsigned_to_nat(1u); +x_200 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_200, 0, x_198); lean_ctor_set(x_200, 1, x_199); -lean_ctor_set(x_200, 2, x_198); -if (lean_is_scalar(x_156)) { +if (lean_is_scalar(x_158)) { x_201 = lean_alloc_ctor(0, 2, 0); } else { - x_201 = x_156; + x_201 = x_158; } lean_ctor_set(x_201, 0, x_200); -lean_ctor_set(x_201, 1, x_155); +lean_ctor_set(x_201, 1, x_157); return x_201; } } else { lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; -lean_dec(x_142); -lean_dec(x_14); +lean_dec(x_148); lean_dec(x_9); lean_dec(x_8); -x_202 = lean_ctor_get(x_146, 0); +x_202 = lean_ctor_get(x_149, 0); lean_inc(x_202); -x_203 = lean_ctor_get(x_146, 1); +x_203 = lean_ctor_get(x_149, 1); lean_inc(x_203); -if (lean_is_exclusive(x_146)) { - lean_ctor_release(x_146, 0); - lean_ctor_release(x_146, 1); - x_204 = x_146; +if (lean_is_exclusive(x_149)) { + lean_ctor_release(x_149, 0); + lean_ctor_release(x_149, 1); + x_204 = x_149; } else { - lean_dec_ref(x_146); + lean_dec_ref(x_149); x_204 = lean_box(0); } if (lean_is_scalar(x_204)) { @@ -10825,595 +11576,204 @@ lean_ctor_set(x_205, 1, x_203); return x_205; } } -else -{ -lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; -lean_dec(x_140); -lean_dec(x_14); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_206 = lean_ctor_get(x_141, 0); -lean_inc(x_206); -x_207 = lean_ctor_get(x_141, 1); -lean_inc(x_207); -if (lean_is_exclusive(x_141)) { - lean_ctor_release(x_141, 0); - lean_ctor_release(x_141, 1); - x_208 = x_141; -} else { - lean_dec_ref(x_141); - x_208 = lean_box(0); -} -if (lean_is_scalar(x_208)) { - x_209 = lean_alloc_ctor(1, 2, 0); -} else { - x_209 = x_208; -} -lean_ctor_set(x_209, 0, x_206); -lean_ctor_set(x_209, 1, x_207); -return x_209; -} -} } else { -uint8_t x_210; -lean_dec(x_14); +lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; +lean_dec(x_48); +lean_dec(x_45); +lean_dec(x_1); +x_206 = lean_ctor_get(x_13, 1); +lean_inc(x_206); +lean_dec(x_13); +x_207 = lean_box(0); +x_208 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_14, x_207); +x_209 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_208, x_207); +x_210 = l_Lean_MessageData_ofList(x_209); +lean_dec(x_209); +x_211 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__13; +x_212 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_212, 0, x_211); +lean_ctor_set(x_212, 1, x_210); +x_213 = l_Lean_addTrace___at_Lean_Elab_Term_checkLeftRec___spec__3___closed__8; +x_214 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_214, 0, x_212); +lean_ctor_set(x_214, 1, x_213); +x_215 = l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_process___spec__7(x_214, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_206); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_210 = !lean_is_exclusive(x_18); -if (x_210 == 0) -{ -lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; -x_211 = lean_ctor_get(x_18, 0); -x_212 = lean_io_error_to_string(x_211); -x_213 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_213, 0, x_212); -x_214 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_214, 0, x_213); -x_215 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_215, 0, x_17); -lean_ctor_set(x_215, 1, x_214); -lean_ctor_set(x_18, 0, x_215); -return x_18; -} -else -{ -lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; -x_216 = lean_ctor_get(x_18, 0); -x_217 = lean_ctor_get(x_18, 1); -lean_inc(x_217); -lean_inc(x_216); -lean_dec(x_18); -x_218 = lean_io_error_to_string(x_216); -x_219 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_219, 0, x_218); -x_220 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_220, 0, x_219); -x_221 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_221, 0, x_17); -lean_ctor_set(x_221, 1, x_220); -x_222 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_222, 0, x_221); -lean_ctor_set(x_222, 1, x_217); -return x_222; -} -} -} -} -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processUnary___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("ParserDescr.unary", 17); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processUnary___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_toParserDescr_processUnary___closed__1; -x_2 = lean_string_utf8_byte_size(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processUnary___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Term_toParserDescr_processUnary___closed__1; -x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Term_toParserDescr_processUnary___closed__2; -x_4 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_4, 0, x_1); -lean_ctor_set(x_4, 1, x_2); -lean_ctor_set(x_4, 2, x_3); -return x_4; -} -} -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processUnary___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__13; -x_2 = l_Lean_Elab_Term_toParserDescr_process___closed__3; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processUnary___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__16; -x_2 = l_Lean_Elab_Term_toParserDescr_process___closed__3; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Elab_Term_toParserDescr_processUnary___closed__6() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_toParserDescr_processUnary___closed__5; -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_Elab_Term_toParserDescr_processUnary___closed__7() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_toParserDescr_processUnary___closed__6; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processUnary(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_object* x_9, lean_object* x_10) { -_start: -{ -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_11 = lean_unsigned_to_nat(0u); -x_12 = l_Lean_Syntax_getArg(x_1, x_11); -x_13 = l_Lean_Syntax_getId(x_12); -lean_dec(x_12); -x_14 = lean_erase_macro_scopes(x_13); -x_15 = lean_st_ref_get(x_9, x_10); -x_16 = lean_ctor_get(x_15, 1); -lean_inc(x_16); -lean_dec(x_15); -x_17 = lean_ctor_get(x_8, 5); -lean_inc(x_17); -lean_inc(x_14); -x_18 = l_Lean_Parser_ensureUnaryParserAlias(x_14, x_16); -if (lean_obj_tag(x_18) == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -lean_dec(x_17); -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -lean_dec(x_18); -x_20 = lean_unsigned_to_nat(2u); -x_21 = l_Lean_Syntax_getArg(x_1, x_20); -x_22 = !lean_is_exclusive(x_2); -if (x_22 == 0) -{ -uint8_t x_23; lean_object* x_24; -x_23 = 0; -lean_ctor_set_uint8(x_2, sizeof(void*)*1, x_23); -lean_ctor_set_uint8(x_2, sizeof(void*)*1 + 1, x_23); -lean_inc(x_9); -lean_inc(x_8); -x_24 = l_Lean_Elab_Term_toParserDescr_process(x_21, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_19); -if (lean_obj_tag(x_24) == 0) -{ -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_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -x_26 = lean_ctor_get(x_24, 1); -lean_inc(x_26); -lean_dec(x_24); -lean_inc(x_8); -x_27 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_toParserDescr_processNonReserved___spec__2___rarg(x_8, x_9, x_26); -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); -lean_dec(x_27); -x_30 = lean_ctor_get(x_8, 10); -lean_inc(x_30); -lean_dec(x_8); -x_31 = lean_st_ref_get(x_9, x_29); -lean_dec(x_9); -x_32 = !lean_is_exclusive(x_31); -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; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_33 = lean_ctor_get(x_31, 0); -x_34 = lean_ctor_get(x_33, 0); -lean_inc(x_34); -lean_dec(x_33); -x_35 = lean_environment_main_module(x_34); -x_36 = l_Lean_Elab_Term_toParserDescr_processUnary___closed__4; -x_37 = l_Lean_addMacroScope(x_35, x_36, x_30); -x_38 = lean_box(0); -x_39 = l_Lean_Elab_Term_toParserDescr_processUnary___closed__3; -x_40 = l_Lean_Elab_Term_toParserDescr_processUnary___closed__7; -x_41 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_41, 0, x_28); -lean_ctor_set(x_41, 1, x_39); -lean_ctor_set(x_41, 2, x_37); -lean_ctor_set(x_41, 3, x_40); -lean_inc(x_14); -x_42 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_38, x_14); -x_43 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; -x_44 = lean_array_push(x_43, x_41); -if (lean_obj_tag(x_42) == 0) -{ -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; -x_45 = l___private_Init_Meta_0__Lean_quoteNameMk(x_14); -x_46 = lean_array_push(x_43, x_45); -x_47 = lean_array_push(x_46, x_25); -x_48 = lean_box(2); -x_49 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_50 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_50, 0, x_48); -lean_ctor_set(x_50, 1, x_49); -lean_ctor_set(x_50, 2, x_47); -x_51 = lean_array_push(x_44, x_50); -x_52 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_53 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_53, 0, x_48); -lean_ctor_set(x_53, 1, x_52); -lean_ctor_set(x_53, 2, x_51); -lean_ctor_set(x_31, 0, x_53); -return x_31; -} -else -{ -lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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_dec(x_14); -x_54 = lean_ctor_get(x_42, 0); -lean_inc(x_54); -lean_dec(x_42); -x_55 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__8; -x_56 = l_String_intercalate(x_55, x_54); -x_57 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__9; -x_58 = lean_string_append(x_57, x_56); -lean_dec(x_56); -x_59 = l_Lean_nameLitKind; -x_60 = lean_box(2); -x_61 = l_Lean_Syntax_mkLit(x_59, x_58, x_60); -x_62 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; -x_63 = lean_array_push(x_62, x_61); -x_64 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__23; -x_65 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_65, 0, x_60); -lean_ctor_set(x_65, 1, x_64); -lean_ctor_set(x_65, 2, x_63); -x_66 = lean_array_push(x_43, x_65); -x_67 = lean_array_push(x_66, x_25); -x_68 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_69 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_69, 0, x_60); -lean_ctor_set(x_69, 1, x_68); -lean_ctor_set(x_69, 2, x_67); -x_70 = lean_array_push(x_44, x_69); -x_71 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_72 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_72, 0, x_60); -lean_ctor_set(x_72, 1, x_71); -lean_ctor_set(x_72, 2, x_70); -lean_ctor_set(x_31, 0, x_72); -return x_31; -} -} -else -{ -lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_73 = lean_ctor_get(x_31, 0); -x_74 = lean_ctor_get(x_31, 1); -lean_inc(x_74); -lean_inc(x_73); -lean_dec(x_31); -x_75 = lean_ctor_get(x_73, 0); -lean_inc(x_75); -lean_dec(x_73); -x_76 = lean_environment_main_module(x_75); -x_77 = l_Lean_Elab_Term_toParserDescr_processUnary___closed__4; -x_78 = l_Lean_addMacroScope(x_76, x_77, x_30); -x_79 = lean_box(0); -x_80 = l_Lean_Elab_Term_toParserDescr_processUnary___closed__3; -x_81 = l_Lean_Elab_Term_toParserDescr_processUnary___closed__7; -x_82 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_82, 0, x_28); -lean_ctor_set(x_82, 1, x_80); -lean_ctor_set(x_82, 2, x_78); -lean_ctor_set(x_82, 3, x_81); -lean_inc(x_14); -x_83 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_79, x_14); -x_84 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; -x_85 = lean_array_push(x_84, x_82); -if (lean_obj_tag(x_83) == 0) -{ -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; -x_86 = l___private_Init_Meta_0__Lean_quoteNameMk(x_14); -x_87 = lean_array_push(x_84, x_86); -x_88 = lean_array_push(x_87, x_25); -x_89 = lean_box(2); -x_90 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_91 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_91, 0, x_89); -lean_ctor_set(x_91, 1, x_90); -lean_ctor_set(x_91, 2, x_88); -x_92 = lean_array_push(x_85, x_91); -x_93 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_94 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_94, 0, x_89); -lean_ctor_set(x_94, 1, x_93); -lean_ctor_set(x_94, 2, x_92); -x_95 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_95, 0, x_94); -lean_ctor_set(x_95, 1, x_74); -return x_95; -} -else -{ -lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; -lean_dec(x_14); -x_96 = lean_ctor_get(x_83, 0); -lean_inc(x_96); -lean_dec(x_83); -x_97 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__8; -x_98 = l_String_intercalate(x_97, x_96); -x_99 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__9; -x_100 = lean_string_append(x_99, x_98); -lean_dec(x_98); -x_101 = l_Lean_nameLitKind; -x_102 = lean_box(2); -x_103 = l_Lean_Syntax_mkLit(x_101, x_100, x_102); -x_104 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; -x_105 = lean_array_push(x_104, x_103); -x_106 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__23; -x_107 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_107, 0, x_102); -lean_ctor_set(x_107, 1, x_106); -lean_ctor_set(x_107, 2, x_105); -x_108 = lean_array_push(x_84, x_107); -x_109 = lean_array_push(x_108, x_25); -x_110 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_111 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_111, 0, x_102); -lean_ctor_set(x_111, 1, x_110); -lean_ctor_set(x_111, 2, x_109); -x_112 = lean_array_push(x_85, x_111); -x_113 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_114 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_114, 0, x_102); -lean_ctor_set(x_114, 1, x_113); -lean_ctor_set(x_114, 2, x_112); -x_115 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_115, 0, x_114); -lean_ctor_set(x_115, 1, x_74); -return x_115; -} -} -} -else -{ -uint8_t x_116; -lean_dec(x_14); -lean_dec(x_9); -lean_dec(x_8); -x_116 = !lean_is_exclusive(x_24); -if (x_116 == 0) -{ -return x_24; -} -else -{ -lean_object* x_117; lean_object* x_118; lean_object* x_119; -x_117 = lean_ctor_get(x_24, 0); -x_118 = lean_ctor_get(x_24, 1); -lean_inc(x_118); -lean_inc(x_117); -lean_dec(x_24); -x_119 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_119, 0, x_117); -lean_ctor_set(x_119, 1, x_118); -return x_119; -} -} -} -else -{ -lean_object* x_120; uint8_t x_121; uint8_t x_122; lean_object* x_123; lean_object* x_124; -x_120 = lean_ctor_get(x_2, 0); -x_121 = lean_ctor_get_uint8(x_2, sizeof(void*)*1 + 2); -lean_inc(x_120); -lean_dec(x_2); -x_122 = 0; -x_123 = lean_alloc_ctor(0, 1, 3); -lean_ctor_set(x_123, 0, x_120); -lean_ctor_set_uint8(x_123, sizeof(void*)*1, x_122); -lean_ctor_set_uint8(x_123, sizeof(void*)*1 + 1, x_122); -lean_ctor_set_uint8(x_123, sizeof(void*)*1 + 2, x_121); -lean_inc(x_9); -lean_inc(x_8); -x_124 = l_Lean_Elab_Term_toParserDescr_process(x_21, x_123, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_19); -if (lean_obj_tag(x_124) == 0) -{ -lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; -x_125 = lean_ctor_get(x_124, 0); -lean_inc(x_125); -x_126 = lean_ctor_get(x_124, 1); -lean_inc(x_126); -lean_dec(x_124); -lean_inc(x_8); -x_127 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Term_toParserDescr_processNonReserved___spec__2___rarg(x_8, x_9, x_126); -x_128 = lean_ctor_get(x_127, 0); -lean_inc(x_128); -x_129 = lean_ctor_get(x_127, 1); -lean_inc(x_129); -lean_dec(x_127); -x_130 = lean_ctor_get(x_8, 10); -lean_inc(x_130); -lean_dec(x_8); -x_131 = lean_st_ref_get(x_9, x_129); -lean_dec(x_9); -x_132 = lean_ctor_get(x_131, 0); -lean_inc(x_132); -x_133 = lean_ctor_get(x_131, 1); -lean_inc(x_133); -if (lean_is_exclusive(x_131)) { - lean_ctor_release(x_131, 0); - lean_ctor_release(x_131, 1); - x_134 = x_131; -} else { - lean_dec_ref(x_131); - x_134 = lean_box(0); -} -x_135 = lean_ctor_get(x_132, 0); -lean_inc(x_135); -lean_dec(x_132); -x_136 = lean_environment_main_module(x_135); -x_137 = l_Lean_Elab_Term_toParserDescr_processUnary___closed__4; -x_138 = l_Lean_addMacroScope(x_136, x_137, x_130); -x_139 = lean_box(0); -x_140 = l_Lean_Elab_Term_toParserDescr_processUnary___closed__3; -x_141 = l_Lean_Elab_Term_toParserDescr_processUnary___closed__7; -x_142 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_142, 0, x_128); -lean_ctor_set(x_142, 1, x_140); -lean_ctor_set(x_142, 2, x_138); -lean_ctor_set(x_142, 3, x_141); -lean_inc(x_14); -x_143 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_139, x_14); -x_144 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; -x_145 = lean_array_push(x_144, x_142); -if (lean_obj_tag(x_143) == 0) -{ -lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; -x_146 = l___private_Init_Meta_0__Lean_quoteNameMk(x_14); -x_147 = lean_array_push(x_144, x_146); -x_148 = lean_array_push(x_147, x_125); -x_149 = lean_box(2); -x_150 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_151 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_151, 0, x_149); -lean_ctor_set(x_151, 1, x_150); -lean_ctor_set(x_151, 2, x_148); -x_152 = lean_array_push(x_145, x_151); -x_153 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_154 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_154, 0, x_149); -lean_ctor_set(x_154, 1, x_153); -lean_ctor_set(x_154, 2, x_152); -if (lean_is_scalar(x_134)) { - x_155 = lean_alloc_ctor(0, 2, 0); -} else { - x_155 = x_134; +lean_dec(x_3); +lean_dec(x_2); +return x_215; } -lean_ctor_set(x_155, 0, x_154); -lean_ctor_set(x_155, 1, x_133); -return x_155; } else { -lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; +lean_object* x_216; +x_216 = lean_ctor_get(x_14, 1); +lean_inc(x_216); +if (lean_obj_tag(x_216) == 0) +{ +lean_object* x_217; uint8_t x_218; lean_dec(x_14); -x_156 = lean_ctor_get(x_143, 0); -lean_inc(x_156); -lean_dec(x_143); -x_157 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__8; -x_158 = l_String_intercalate(x_157, x_156); -x_159 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__9; -x_160 = lean_string_append(x_159, x_158); -lean_dec(x_158); -x_161 = l_Lean_nameLitKind; -x_162 = lean_box(2); -x_163 = l_Lean_Syntax_mkLit(x_161, x_160, x_162); -x_164 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; -x_165 = lean_array_push(x_164, x_163); -x_166 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__23; -x_167 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_167, 0, x_162); -lean_ctor_set(x_167, 1, x_166); -lean_ctor_set(x_167, 2, x_165); -x_168 = lean_array_push(x_144, x_167); -x_169 = lean_array_push(x_168, x_125); -x_170 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_171 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_171, 0, x_162); -lean_ctor_set(x_171, 1, x_170); -lean_ctor_set(x_171, 2, x_169); -x_172 = lean_array_push(x_145, x_171); -x_173 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; -x_174 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_174, 0, x_162); -lean_ctor_set(x_174, 1, x_173); -lean_ctor_set(x_174, 2, x_172); -if (lean_is_scalar(x_134)) { - x_175 = lean_alloc_ctor(0, 2, 0); -} else { - x_175 = x_134; +x_217 = lean_ctor_get(x_13, 1); +lean_inc(x_217); +lean_dec(x_13); +x_218 = !lean_is_exclusive(x_45); +if (x_218 == 0) +{ +lean_object* x_219; lean_object* x_220; lean_object* x_221; +x_219 = lean_ctor_get(x_45, 0); +x_220 = lean_ctor_get(x_45, 1); +lean_dec(x_220); +x_221 = l_Lean_Elab_Term_toParserDescr_ensureNoPrec(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_217); +if (lean_obj_tag(x_221) == 0) +{ +uint8_t x_222; +x_222 = !lean_is_exclusive(x_221); +if (x_222 == 0) +{ +lean_object* x_223; lean_object* x_224; lean_object* x_225; +x_223 = lean_ctor_get(x_221, 0); +lean_dec(x_223); +x_224 = l_Lean_mkIdentFrom(x_1, x_219); +x_225 = lean_unsigned_to_nat(1u); +lean_ctor_set(x_45, 1, x_225); +lean_ctor_set(x_45, 0, x_224); +lean_ctor_set(x_221, 0, x_45); +return x_221; +} +else +{ +lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; +x_226 = lean_ctor_get(x_221, 1); +lean_inc(x_226); +lean_dec(x_221); +x_227 = l_Lean_mkIdentFrom(x_1, x_219); +x_228 = lean_unsigned_to_nat(1u); +lean_ctor_set(x_45, 1, x_228); +lean_ctor_set(x_45, 0, x_227); +x_229 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_229, 0, x_45); +lean_ctor_set(x_229, 1, x_226); +return x_229; +} +} +else +{ +uint8_t x_230; +lean_free_object(x_45); +lean_dec(x_219); +lean_dec(x_1); +x_230 = !lean_is_exclusive(x_221); +if (x_230 == 0) +{ +return x_221; +} +else +{ +lean_object* x_231; lean_object* x_232; lean_object* x_233; +x_231 = lean_ctor_get(x_221, 0); +x_232 = lean_ctor_get(x_221, 1); +lean_inc(x_232); +lean_inc(x_231); +lean_dec(x_221); +x_233 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_233, 0, x_231); +lean_ctor_set(x_233, 1, x_232); +return x_233; } -lean_ctor_set(x_175, 0, x_174); -lean_ctor_set(x_175, 1, x_133); -return x_175; } } else { -lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; -lean_dec(x_14); -lean_dec(x_9); -lean_dec(x_8); -x_176 = lean_ctor_get(x_124, 0); -lean_inc(x_176); -x_177 = lean_ctor_get(x_124, 1); -lean_inc(x_177); -if (lean_is_exclusive(x_124)) { - lean_ctor_release(x_124, 0); - lean_ctor_release(x_124, 1); - x_178 = x_124; +lean_object* x_234; lean_object* x_235; +x_234 = lean_ctor_get(x_45, 0); +lean_inc(x_234); +lean_dec(x_45); +x_235 = l_Lean_Elab_Term_toParserDescr_ensureNoPrec(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_217); +if (lean_obj_tag(x_235) == 0) +{ +lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; +x_236 = lean_ctor_get(x_235, 1); +lean_inc(x_236); +if (lean_is_exclusive(x_235)) { + lean_ctor_release(x_235, 0); + lean_ctor_release(x_235, 1); + x_237 = x_235; +} else { + lean_dec_ref(x_235); + x_237 = lean_box(0); +} +x_238 = l_Lean_mkIdentFrom(x_1, x_234); +x_239 = lean_unsigned_to_nat(1u); +x_240 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_240, 0, x_238); +lean_ctor_set(x_240, 1, x_239); +if (lean_is_scalar(x_237)) { + x_241 = lean_alloc_ctor(0, 2, 0); +} else { + x_241 = x_237; +} +lean_ctor_set(x_241, 0, x_240); +lean_ctor_set(x_241, 1, x_236); +return x_241; +} +else +{ +lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; +lean_dec(x_234); +lean_dec(x_1); +x_242 = lean_ctor_get(x_235, 0); +lean_inc(x_242); +x_243 = lean_ctor_get(x_235, 1); +lean_inc(x_243); +if (lean_is_exclusive(x_235)) { + lean_ctor_release(x_235, 0); + lean_ctor_release(x_235, 1); + x_244 = x_235; } else { - lean_dec_ref(x_124); - x_178 = lean_box(0); + lean_dec_ref(x_235); + x_244 = lean_box(0); } -if (lean_is_scalar(x_178)) { - x_179 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_244)) { + x_245 = lean_alloc_ctor(1, 2, 0); } else { - x_179 = x_178; + x_245 = x_244; } -lean_ctor_set(x_179, 0, x_176); -lean_ctor_set(x_179, 1, x_177); -return x_179; +lean_ctor_set(x_245, 0, x_242); +lean_ctor_set(x_245, 1, x_243); +return x_245; } } } else { -uint8_t x_180; -lean_dec(x_14); +lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; +lean_dec(x_216); +lean_dec(x_45); +lean_dec(x_1); +x_246 = lean_ctor_get(x_13, 1); +lean_inc(x_246); +lean_dec(x_13); +x_247 = lean_box(0); +x_248 = l_List_mapTRAux___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__16(x_14, x_247); +x_249 = l_List_mapTRAux___at_Lean_Meta_substCore___spec__6(x_248, x_247); +x_250 = l_Lean_MessageData_ofList(x_249); +lean_dec(x_249); +x_251 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__13; +x_252 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_252, 0, x_251); +lean_ctor_set(x_252, 1, x_250); +x_253 = l_Lean_addTrace___at_Lean_Elab_Term_checkLeftRec___spec__3___closed__8; +x_254 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_254, 0, x_252); +lean_ctor_set(x_254, 1, x_253); +x_255 = l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_process___spec__7(x_254, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_246); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -11422,42 +11782,8 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_180 = !lean_is_exclusive(x_18); -if (x_180 == 0) -{ -lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; -x_181 = lean_ctor_get(x_18, 0); -x_182 = lean_io_error_to_string(x_181); -x_183 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_183, 0, x_182); -x_184 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_184, 0, x_183); -x_185 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_185, 0, x_17); -lean_ctor_set(x_185, 1, x_184); -lean_ctor_set(x_18, 0, x_185); -return x_18; +return x_255; } -else -{ -lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; -x_186 = lean_ctor_get(x_18, 0); -x_187 = lean_ctor_get(x_18, 1); -lean_inc(x_187); -lean_inc(x_186); -lean_dec(x_18); -x_188 = lean_io_error_to_string(x_186); -x_189 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_189, 0, x_188); -x_190 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_190, 0, x_189); -x_191 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_191, 0, x_17); -lean_ctor_set(x_191, 1, x_190); -x_192 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_192, 0, x_191); -lean_ctor_set(x_192, 1, x_187); -return x_192; } } } @@ -11927,10 +12253,249 @@ return x_39; } } } -else +else +{ +uint8_t x_40; +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_40 = !lean_is_exclusive(x_14); +if (x_40 == 0) +{ +return x_14; +} +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_14, 0); +x_42 = lean_ctor_get(x_14, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_14); +x_43 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +return x_43; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_process___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) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_process___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_toParserDescr_process___spec__4___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_throwMaxRecDepthAt___at_Lean_Elab_Term_toParserDescr_process___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_toParserDescr_process___spec__5___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_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_toParserDescr_process___spec__5(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_process___spec__7___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_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_process___spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy1___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, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy___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, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_toParserDescr_processAlias___spec__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, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +size_t x_13; size_t x_14; lean_object* x_15; +x_13 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_14 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_15 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_toParserDescr_processAlias___spec__1(x_13, x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_4); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_toParserDescr_processAlias___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; size_t x_6; lean_object* x_7; +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_toParserDescr_processAlias___spec__3(x_1, x_5, x_6, x_4); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_toParserDescr_processAlias___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_toParserDescr_processAlias___spec__4(x_4, x_5, x_3); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_toParserDescr_processAlias___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_toParserDescr_processAlias___spec__5(x_4, x_5, x_3); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processAlias___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, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_Elab_Term_toParserDescr_processAlias___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__5___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_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_resolveGlobalName___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__7___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_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_resolveGlobalName___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___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_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__6___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, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: { -uint8_t x_40; +lean_object* x_13; +x_13 = l_Lean_resolveGlobalConstCore___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__6___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -11938,134 +12503,120 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_40 = !lean_is_exclusive(x_14); -if (x_40 == 0) -{ -return x_14; -} -else -{ -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_14, 0); -x_42 = lean_ctor_get(x_14, 1); -lean_inc(x_42); -lean_inc(x_41); -lean_dec(x_14); -x_43 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_43, 0, x_41); -lean_ctor_set(x_43, 1, x_42); -return x_43; -} -} +return x_13; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_process___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_Lean_throwError___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__11___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_object* x_9, lean_object* x_10) { _start: { -lean_object* x_9; -x_9 = l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_process___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_object* x_11; +x_11 = l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__11(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_9; +lean_dec(x_2); +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_toParserDescr_process___spec__4___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_Lean_getConstInfo___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__10___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_object* x_9, lean_object* x_10) { _start: { -lean_object* x_9; -x_9 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_toParserDescr_process___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_object* x_11; +x_11 = l_Lean_getConstInfo___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__10(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_9; +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_toParserDescr_process___spec__5___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_Lean_mkConstWithLevelParams___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__9___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_object* x_9, lean_object* x_10) { _start: { -lean_object* x_7; -x_7 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_toParserDescr_process___spec__5(x_1, x_2, x_3, x_4, x_5, x_6); +lean_object* x_11; +x_11 = l_Lean_mkConstWithLevelParams___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -return x_7; +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_toParserDescr_process___spec__6___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_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoTree___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__13___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_object* x_9, lean_object* x_10) { _start: { -lean_object* x_12; -x_12 = l_Lean_throwErrorAt___at_Lean_Elab_Term_toParserDescr_process___spec__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); +lean_object* x_11; +x_11 = l_Lean_Elab_pushInfoTree___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__13(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_1); -return x_12; +lean_dec(x_2); +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy1___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, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___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_object* x_9, lean_object* x_10) { _start: { -lean_object* x_14; -x_14 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_12); -lean_dec(x_10); +lean_object* x_11; +x_11 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -return x_14; +lean_dec(x_2); +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy___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, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__14___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { lean_object* x_14; -x_14 = l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +x_14 = l_List_forIn_loop___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__14(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_3); return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processBinary___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_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_Lean_Elab_Term_toParserDescr_processBinary(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_1); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processUnary___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_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Elab_resolveGlobalConstWithInfos___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__2___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, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_11; -x_11 = l_Lean_Elab_Term_toParserDescr_processUnary(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_1); -return x_11; +lean_object* x_12; +x_12 = l_Lean_Elab_resolveGlobalConstWithInfos___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__2___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_12; } } LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at_Lean_Elab_Term_toParserDescr_processSeq___spec__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, 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) { @@ -12328,19 +12879,10 @@ return x_3; static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__11() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(0u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; -} -} -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__12() { -_start: -{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); x_2 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_3 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__11; +x_3 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__3; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -12348,7 +12890,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__13() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__12() { _start: { lean_object* x_1; @@ -12356,17 +12898,17 @@ x_1 = lean_mk_string_from_bytes("attributes", 10); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__14() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__6; -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__13; +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__12; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__15() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__14() { _start: { lean_object* x_1; @@ -12374,7 +12916,7 @@ x_1 = lean_mk_string_from_bytes("@[", 2); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__16() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__15() { _start: { lean_object* x_1; @@ -12382,17 +12924,17 @@ x_1 = lean_mk_string_from_bytes("attrInstance", 12); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__17() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__6; -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__16; +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__15; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__18() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__17() { _start: { lean_object* x_1; @@ -12400,33 +12942,33 @@ x_1 = lean_mk_string_from_bytes("attrKind", 8); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__19() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__6; -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__18; +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__17; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__20() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__12; +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__11; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__21() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__19; -x_3 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__20; +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__18; +x_3 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__19; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -12434,7 +12976,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__22() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__21() { _start: { lean_object* x_1; @@ -12442,17 +12984,17 @@ x_1 = lean_mk_string_from_bytes("Attr", 4); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__23() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__4; -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__22; +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__21; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__24() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__23() { _start: { lean_object* x_1; @@ -12460,17 +13002,17 @@ x_1 = lean_mk_string_from_bytes("simple", 6); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__25() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__23; -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__24; +x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__22; +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__23; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__26() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__25() { _start: { lean_object* x_1; @@ -12478,22 +13020,22 @@ x_1 = lean_mk_string_from_bytes("termParser", 10); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__27() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__26() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__26; +x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__25; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__28() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__27() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__26; +x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__25; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__27; +x_3 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__26; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -12501,27 +13043,27 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__29() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__26; +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__25; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__30() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__29() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__21; +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__20; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__30() { _start: { lean_object* x_1; @@ -12529,7 +13071,7 @@ x_1 = lean_mk_string_from_bytes("]", 1); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__32() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31() { _start: { lean_object* x_1; lean_object* x_2; @@ -12538,17 +13080,17 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__32() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__32; -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__12; +x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31; +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__11; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__34() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33() { _start: { lean_object* x_1; @@ -12556,35 +13098,53 @@ x_1 = lean_mk_string_from_bytes("def", 3); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__35() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__34() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__6; -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__34; +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__35() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("declId", 6); +return x_1; +} +} static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__36() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__6; +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__35; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__37() { +_start: +{ lean_object* x_1; x_1 = lean_mk_string_from_bytes("optDeclSig", 10); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__37() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__38() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__6; -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__36; +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__37; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__38() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__39() { _start: { lean_object* x_1; @@ -12592,17 +13152,17 @@ x_1 = lean_mk_string_from_bytes("typeSpec", 8); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__39() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__40() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__6; -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__38; +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__39; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__40() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__41() { _start: { lean_object* x_1; @@ -12610,7 +13170,7 @@ x_1 = lean_mk_string_from_bytes(":", 1); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__41() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__42() { _start: { lean_object* x_1; @@ -12618,22 +13178,22 @@ x_1 = lean_mk_string_from_bytes("Lean.ParserDescr", 16); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__42() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__43() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__41; +x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__42; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__43() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__44() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__41; +x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__42; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__42; +x_3 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__43; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -12641,7 +13201,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__44() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__45() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -12653,29 +13213,29 @@ lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__45() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__46() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__44; +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__45; x_3 = lean_alloc_ctor(1, 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___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__46() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__47() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__12; +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__11; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__47() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__48() { _start: { lean_object* x_1; @@ -12683,17 +13243,17 @@ x_1 = lean_mk_string_from_bytes("declValSimple", 13); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__48() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__49() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__6; -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__47; +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__48; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__49() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__50() { _start: { lean_object* x_1; @@ -12701,7 +13261,7 @@ x_1 = lean_mk_string_from_bytes(":=", 2); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__50() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__51() { _start: { lean_object* x_1; @@ -12709,22 +13269,22 @@ x_1 = lean_mk_string_from_bytes("Lean.ParserDescr.node", 21); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__51() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__52() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__50; +x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__51; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__52() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__53() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__50; +x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__51; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__51; +x_3 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__52; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -12732,7 +13292,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__53() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__54() { _start: { lean_object* x_1; @@ -12740,41 +13300,41 @@ x_1 = lean_mk_string_from_bytes("node", 4); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__54() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__55() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__16; -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__53; +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__54; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__55() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__56() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__54; +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__55; 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___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__56() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__57() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__55; +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__56; x_3 = lean_alloc_ctor(1, 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___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__57() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__58() { _start: { lean_object* x_1; @@ -12782,7 +13342,25 @@ x_1 = lean_mk_string_from_bytes("`Lean.Parser.Term.quot", 22); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__58() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__59() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("num", 3); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__60() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__59; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__61() { _start: { lean_object* x_1; lean_object* x_2; @@ -12791,28 +13369,27 @@ x_2 = l_Nat_repr(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__59() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__62() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_numLitKind; -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__58; -x_3 = lean_box(2); -x_4 = l_Lean_Syntax_mkLit(x_1, x_2, x_3); -return x_4; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__61; +x_2 = lean_box(2); +x_3 = l_Lean_Syntax_mkNumLit(x_1, x_2); +return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__60() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__63() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__6; -x_2 = l_Lean_Elab_Term_toParserDescr_process___closed__1; +x_2 = l_Lean_Elab_Term_toParserDescr_process___closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__61() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__64() { _start: { lean_object* x_1; @@ -12820,7 +13397,7 @@ x_1 = lean_mk_string_from_bytes("(", 1); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__62() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__65() { _start: { lean_object* x_1; @@ -12828,22 +13405,22 @@ x_1 = lean_mk_string_from_bytes("Lean.ParserDescr.binary", 23); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__63() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__66() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__62; +x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__65; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__64() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__67() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__62; +x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__65; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__63; +x_3 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__66; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -12851,7 +13428,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__65() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__68() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -12863,19 +13440,19 @@ lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__66() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__69() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__65; +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__68; x_3 = lean_alloc_ctor(1, 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___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__67() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__70() { _start: { lean_object* x_1; @@ -12883,22 +13460,22 @@ x_1 = lean_mk_string_from_bytes("Lean.ParserDescr.symbol", 23); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__68() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__71() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__67; +x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__70; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__69() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__72() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__67; +x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__70; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__68; +x_3 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__71; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -12906,7 +13483,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__70() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__73() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -12918,19 +13495,37 @@ lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__71() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__74() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__70; +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__73; x_3 = lean_alloc_ctor(1, 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___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__72() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__75() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("str", 3); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__76() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__75; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__77() { _start: { lean_object* x_1; @@ -12938,7 +13533,7 @@ x_1 = lean_mk_string_from_bytes(")", 1); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__73() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__78() { _start: { lean_object* x_1; @@ -12946,22 +13541,22 @@ x_1 = lean_mk_string_from_bytes("Lean.ParserDescr.cat", 20); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__74() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__79() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__73; +x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__78; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__75() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__80() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__73; +x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__78; x_2 = lean_unsigned_to_nat(0u); -x_3 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__74; +x_3 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__79; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -12969,7 +13564,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__76() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__81() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -12981,37 +13576,19 @@ lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__77() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__82() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__76; +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__81; x_3 = lean_alloc_ctor(1, 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___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__78() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("num", 3); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__79() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__78; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__80() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__83() { _start: { lean_object* x_1; @@ -13019,25 +13596,7 @@ x_1 = lean_mk_string_from_bytes("0", 1); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__81() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("str", 3); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__82() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__81; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__83() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__84() { _start: { lean_object* x_1; @@ -13045,7 +13604,7 @@ x_1 = lean_mk_string_from_bytes("\")\"", 3); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__84() { +static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__85() { _start: { lean_object* x_1; lean_object* x_2; @@ -13059,7 +13618,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_decla { if (lean_obj_tag(x_1) == 1) { -lean_object* x_5; lean_object* x_6; lean_object* x_7; 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; 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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; +lean_object* x_5; lean_object* x_6; lean_object* x_7; 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; 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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; x_5 = lean_ctor_get(x_1, 1); lean_inc(x_5); x_6 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__1; @@ -13087,17 +13646,17 @@ lean_inc(x_19); x_20 = lean_ctor_get(x_18, 1); lean_inc(x_20); lean_dec(x_18); -x_21 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__15; +x_21 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__14; lean_inc(x_13); x_22 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_22, 0, x_13); lean_ctor_set(x_22, 1, x_21); -x_23 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__29; +x_23 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__28; lean_inc(x_16); lean_inc(x_19); x_24 = l_Lean_addMacroScope(x_19, x_23, x_16); x_25 = lean_box(0); -x_26 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__28; +x_26 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__27; lean_inc(x_13); x_27 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_27, 0, x_13); @@ -13106,17 +13665,17 @@ lean_ctor_set(x_27, 2, x_24); lean_ctor_set(x_27, 3, x_25); x_28 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; x_29 = lean_array_push(x_28, x_27); -x_30 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__12; +x_30 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__11; x_31 = lean_array_push(x_29, x_30); x_32 = lean_box(2); -x_33 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__25; +x_33 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__24; x_34 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_34, 0, x_32); lean_ctor_set(x_34, 1, x_33); lean_ctor_set(x_34, 2, x_31); -x_35 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__30; +x_35 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__29; x_36 = lean_array_push(x_35, x_34); -x_37 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__17; +x_37 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__16; x_38 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_38, 0, x_32); lean_ctor_set(x_38, 1, x_37); @@ -13128,7 +13687,7 @@ x_42 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_42, 0, x_32); lean_ctor_set(x_42, 1, x_41); lean_ctor_set(x_42, 2, x_40); -x_43 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31; +x_43 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__30; lean_inc(x_13); x_44 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_44, 0, x_13); @@ -13137,7 +13696,7 @@ x_45 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_ x_46 = lean_array_push(x_45, x_22); x_47 = lean_array_push(x_46, x_42); x_48 = lean_array_push(x_47, x_44); -x_49 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__14; +x_49 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__13; x_50 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_50, 0, x_32); lean_ctor_set(x_50, 1, x_49); @@ -13147,7 +13706,7 @@ x_52 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_52, 0, x_32); lean_ctor_set(x_52, 1, x_41); lean_ctor_set(x_52, 2, x_51); -x_53 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33; +x_53 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__32; x_54 = lean_array_push(x_53, x_52); x_55 = lean_array_push(x_54, x_30); x_56 = lean_array_push(x_55, x_30); @@ -13158,449 +13717,454 @@ x_60 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_60, 0, x_32); lean_ctor_set(x_60, 1, x_59); lean_ctor_set(x_60, 2, x_58); -x_61 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__34; +x_61 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33; lean_inc(x_13); x_62 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_62, 0, x_13); lean_ctor_set(x_62, 1, x_61); lean_inc(x_11); x_63 = lean_mk_syntax_ident(x_11); -x_64 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__40; +x_64 = lean_array_push(x_28, x_63); +x_65 = lean_array_push(x_64, x_30); +x_66 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__36; +x_67 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_67, 0, x_32); +lean_ctor_set(x_67, 1, x_66); +lean_ctor_set(x_67, 2, x_65); +x_68 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__41; lean_inc(x_13); -x_65 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_65, 0, x_13); -lean_ctor_set(x_65, 1, x_64); -x_66 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__16; +x_69 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_69, 0, x_13); +lean_ctor_set(x_69, 1, x_68); +x_70 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__16; lean_inc(x_16); lean_inc(x_19); -x_67 = l_Lean_addMacroScope(x_19, x_66, x_16); -x_68 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__43; -x_69 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__45; +x_71 = l_Lean_addMacroScope(x_19, x_70, x_16); +x_72 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__44; +x_73 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__46; lean_inc(x_13); -x_70 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_70, 0, x_13); -lean_ctor_set(x_70, 1, x_68); -lean_ctor_set(x_70, 2, x_67); -lean_ctor_set(x_70, 3, x_69); -x_71 = lean_array_push(x_28, x_65); -x_72 = lean_array_push(x_71, x_70); -x_73 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__39; -x_74 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_74, 0, x_32); -lean_ctor_set(x_74, 1, x_73); -lean_ctor_set(x_74, 2, x_72); -x_75 = lean_array_push(x_39, x_74); -x_76 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_76, 0, x_32); -lean_ctor_set(x_76, 1, x_41); -lean_ctor_set(x_76, 2, x_75); -x_77 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__46; -x_78 = lean_array_push(x_77, x_76); -x_79 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__37; +x_74 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_74, 0, x_13); +lean_ctor_set(x_74, 1, x_72); +lean_ctor_set(x_74, 2, x_71); +lean_ctor_set(x_74, 3, x_73); +x_75 = lean_array_push(x_28, x_69); +x_76 = lean_array_push(x_75, x_74); +x_77 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__40; +x_78 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_78, 0, x_32); +lean_ctor_set(x_78, 1, x_77); +lean_ctor_set(x_78, 2, x_76); +x_79 = lean_array_push(x_39, x_78); x_80 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_80, 0, x_32); -lean_ctor_set(x_80, 1, x_79); -lean_ctor_set(x_80, 2, x_78); -x_81 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__49; +lean_ctor_set(x_80, 1, x_41); +lean_ctor_set(x_80, 2, x_79); +x_81 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__47; +x_82 = lean_array_push(x_81, x_80); +x_83 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__38; +x_84 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_84, 0, x_32); +lean_ctor_set(x_84, 1, x_83); +lean_ctor_set(x_84, 2, x_82); +x_85 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__50; lean_inc(x_13); -x_82 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_82, 0, x_13); -lean_ctor_set(x_82, 1, x_81); -x_83 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__54; +x_86 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_86, 0, x_13); +lean_ctor_set(x_86, 1, x_85); +x_87 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__55; lean_inc(x_16); lean_inc(x_19); -x_84 = l_Lean_addMacroScope(x_19, x_83, x_16); -x_85 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__52; -x_86 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__56; +x_88 = l_Lean_addMacroScope(x_19, x_87, x_16); +x_89 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__53; +x_90 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__57; lean_inc(x_13); -x_87 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_87, 0, x_13); -lean_ctor_set(x_87, 1, x_85); -lean_ctor_set(x_87, 2, x_84); -lean_ctor_set(x_87, 3, x_86); -x_88 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__57; +x_91 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_91, 0, x_13); +lean_ctor_set(x_91, 1, x_89); +lean_ctor_set(x_91, 2, x_88); +lean_ctor_set(x_91, 3, x_90); +x_92 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__58; lean_inc(x_13); -x_89 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_89, 0, x_13); -lean_ctor_set(x_89, 1, x_88); -x_90 = lean_array_push(x_39, x_89); -x_91 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__25; -x_92 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_92, 0, x_32); -lean_ctor_set(x_92, 1, x_91); -lean_ctor_set(x_92, 2, x_90); -x_93 = lean_array_push(x_39, x_92); -x_94 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__23; -x_95 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_95, 0, x_32); -lean_ctor_set(x_95, 1, x_94); -lean_ctor_set(x_95, 2, x_93); -x_96 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__61; +x_93 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_93, 0, x_13); +lean_ctor_set(x_93, 1, x_92); +x_94 = lean_array_push(x_39, x_93); +x_95 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__25; +x_96 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_96, 0, x_32); +lean_ctor_set(x_96, 1, x_95); +lean_ctor_set(x_96, 2, x_94); +x_97 = lean_array_push(x_39, x_96); +x_98 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__23; +x_99 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_99, 0, x_32); +lean_ctor_set(x_99, 1, x_98); +lean_ctor_set(x_99, 2, x_97); +x_100 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__64; lean_inc(x_13); -x_97 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_97, 0, x_13); -lean_ctor_set(x_97, 1, x_96); +x_101 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_101, 0, x_13); +lean_ctor_set(x_101, 1, x_100); lean_inc(x_11); -x_98 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_25, x_11); -x_99 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__17; +x_102 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_25, x_11); +x_103 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__17; lean_inc(x_16); lean_inc(x_19); -x_100 = l_Lean_addMacroScope(x_19, x_99, x_16); -x_101 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__64; -x_102 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__66; +x_104 = l_Lean_addMacroScope(x_19, x_103, x_16); +x_105 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__67; +x_106 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__69; lean_inc(x_13); -x_103 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_103, 0, x_13); -lean_ctor_set(x_103, 1, x_101); -lean_ctor_set(x_103, 2, x_100); -lean_ctor_set(x_103, 3, x_102); -x_104 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__26; +x_107 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_107, 0, x_13); +lean_ctor_set(x_107, 1, x_105); +lean_ctor_set(x_107, 2, x_104); +lean_ctor_set(x_107, 3, x_106); +x_108 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__26; lean_inc(x_13); -x_105 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_105, 0, x_13); -lean_ctor_set(x_105, 1, x_104); -x_106 = lean_array_push(x_39, x_105); -x_107 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_107, 0, x_32); -lean_ctor_set(x_107, 1, x_91); -lean_ctor_set(x_107, 2, x_106); -x_108 = lean_array_push(x_39, x_107); -x_109 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_109, 0, x_32); -lean_ctor_set(x_109, 1, x_94); -lean_ctor_set(x_109, 2, x_108); -x_110 = l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__6; +x_109 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_109, 0, x_13); +lean_ctor_set(x_109, 1, x_108); +x_110 = lean_array_push(x_39, x_109); +x_111 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_111, 0, x_32); +lean_ctor_set(x_111, 1, x_95); +lean_ctor_set(x_111, 2, x_110); +x_112 = lean_array_push(x_39, x_111); +x_113 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_113, 0, x_32); +lean_ctor_set(x_113, 1, x_98); +lean_ctor_set(x_113, 2, x_112); +x_114 = l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__6; lean_inc(x_16); lean_inc(x_19); -x_111 = l_Lean_addMacroScope(x_19, x_110, x_16); -x_112 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__69; -x_113 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__71; +x_115 = l_Lean_addMacroScope(x_19, x_114, x_16); +x_116 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__72; +x_117 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__74; lean_inc(x_13); -x_114 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_114, 0, x_13); -lean_ctor_set(x_114, 1, x_112); -lean_ctor_set(x_114, 2, x_111); -lean_ctor_set(x_114, 3, x_113); -x_115 = l_Lean_Syntax_mkStrLit(x_9, x_32); +x_118 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_118, 0, x_13); +lean_ctor_set(x_118, 1, x_116); +lean_ctor_set(x_118, 2, x_115); +lean_ctor_set(x_118, 3, x_117); +x_119 = l_Lean_Syntax_mkStrLit(x_9, x_32); lean_dec(x_9); -x_116 = lean_array_push(x_39, x_115); -x_117 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_117, 0, x_32); -lean_ctor_set(x_117, 1, x_41); -lean_ctor_set(x_117, 2, x_116); -x_118 = lean_array_push(x_28, x_114); -lean_inc(x_118); -x_119 = lean_array_push(x_118, x_117); -x_120 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +x_120 = lean_array_push(x_39, x_119); x_121 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_121, 0, x_32); -lean_ctor_set(x_121, 1, x_120); -lean_ctor_set(x_121, 2, x_119); -x_122 = lean_array_push(x_28, x_121); -x_123 = lean_array_push(x_122, x_30); -x_124 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_124, 0, x_32); -lean_ctor_set(x_124, 1, x_41); -lean_ctor_set(x_124, 2, x_123); -x_125 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__72; +lean_ctor_set(x_121, 1, x_41); +lean_ctor_set(x_121, 2, x_120); +x_122 = lean_array_push(x_28, x_118); +lean_inc(x_122); +x_123 = lean_array_push(x_122, x_121); +x_124 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__8; +x_125 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_125, 0, x_32); +lean_ctor_set(x_125, 1, x_124); +lean_ctor_set(x_125, 2, x_123); +x_126 = lean_array_push(x_28, x_125); +x_127 = lean_array_push(x_126, x_30); +x_128 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_128, 0, x_32); +lean_ctor_set(x_128, 1, x_41); +lean_ctor_set(x_128, 2, x_127); +x_129 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__77; lean_inc(x_13); -x_126 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_126, 0, x_13); -lean_ctor_set(x_126, 1, x_125); -x_127 = lean_array_push(x_45, x_97); -lean_inc(x_127); -x_128 = lean_array_push(x_127, x_124); -lean_inc(x_126); -x_129 = lean_array_push(x_128, x_126); -x_130 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__60; -x_131 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_131, 0, x_32); -lean_ctor_set(x_131, 1, x_130); -lean_ctor_set(x_131, 2, x_129); -x_132 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__5; -x_133 = l_Lean_addMacroScope(x_19, x_132, x_16); -x_134 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__75; -x_135 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__77; +x_130 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_130, 0, x_13); +lean_ctor_set(x_130, 1, x_129); +x_131 = lean_array_push(x_45, x_101); +lean_inc(x_131); +x_132 = lean_array_push(x_131, x_128); +lean_inc(x_130); +x_133 = lean_array_push(x_132, x_130); +x_134 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__63; +x_135 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_135, 0, x_32); +lean_ctor_set(x_135, 1, x_134); +lean_ctor_set(x_135, 2, x_133); +x_136 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__5; +x_137 = l_Lean_addMacroScope(x_19, x_136, x_16); +x_138 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__80; +x_139 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__82; lean_inc(x_13); -x_136 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_136, 0, x_13); -lean_ctor_set(x_136, 1, x_134); -lean_ctor_set(x_136, 2, x_133); -lean_ctor_set(x_136, 3, x_135); +x_140 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_140, 0, x_13); +lean_ctor_set(x_140, 1, x_138); +lean_ctor_set(x_140, 2, x_137); +lean_ctor_set(x_140, 3, x_139); lean_inc(x_1); -x_137 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_25, x_1); -x_138 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__80; +x_141 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_25, x_1); +x_142 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__83; lean_inc(x_13); -x_139 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_139, 0, x_13); -lean_ctor_set(x_139, 1, x_138); -x_140 = lean_array_push(x_39, x_139); -x_141 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__79; -x_142 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_142, 0, x_32); -lean_ctor_set(x_142, 1, x_141); -lean_ctor_set(x_142, 2, x_140); -x_143 = lean_array_push(x_28, x_136); -x_144 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__83; -x_145 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_145, 0, x_13); -lean_ctor_set(x_145, 1, x_144); -x_146 = lean_array_push(x_39, x_145); -x_147 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__82; -x_148 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_148, 0, x_32); -lean_ctor_set(x_148, 1, x_147); -lean_ctor_set(x_148, 2, x_146); -x_149 = lean_array_push(x_39, x_148); -x_150 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_150, 0, x_32); -lean_ctor_set(x_150, 1, x_41); -lean_ctor_set(x_150, 2, x_149); -x_151 = lean_array_push(x_118, x_150); +x_143 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_143, 0, x_13); +lean_ctor_set(x_143, 1, x_142); +x_144 = lean_array_push(x_39, x_143); +x_145 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__60; +x_146 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_146, 0, x_32); +lean_ctor_set(x_146, 1, x_145); +lean_ctor_set(x_146, 2, x_144); +x_147 = lean_array_push(x_28, x_140); +x_148 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__84; +x_149 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_149, 0, x_13); +lean_ctor_set(x_149, 1, x_148); +x_150 = lean_array_push(x_39, x_149); +x_151 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__76; x_152 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_152, 0, x_32); -lean_ctor_set(x_152, 1, x_120); -lean_ctor_set(x_152, 2, x_151); -x_153 = lean_array_push(x_28, x_152); -x_154 = lean_array_push(x_153, x_30); -x_155 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_155, 0, x_32); -lean_ctor_set(x_155, 1, x_41); -lean_ctor_set(x_155, 2, x_154); -lean_inc(x_127); -x_156 = lean_array_push(x_127, x_155); -lean_inc(x_126); -x_157 = lean_array_push(x_156, x_126); -x_158 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_158, 0, x_32); -lean_ctor_set(x_158, 1, x_130); -lean_ctor_set(x_158, 2, x_157); -x_159 = lean_array_push(x_45, x_109); -x_160 = lean_array_push(x_28, x_103); -lean_inc(x_159); -x_161 = lean_array_push(x_159, x_131); -x_162 = lean_array_push(x_28, x_87); -x_163 = lean_array_push(x_45, x_95); -x_164 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__59; -x_165 = lean_array_push(x_163, x_164); -x_166 = lean_array_push(x_45, x_82); -x_167 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__84; -x_168 = lean_array_push(x_167, x_62); -x_169 = lean_array_push(x_168, x_63); -x_170 = lean_array_push(x_169, x_80); -x_171 = lean_array_push(x_28, x_60); -if (lean_obj_tag(x_98) == 0) -{ -lean_object* x_248; -x_248 = l___private_Init_Meta_0__Lean_quoteNameMk(x_11); -x_172 = x_248; -goto block_247; -} -else -{ -lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; +lean_ctor_set(x_152, 1, x_151); +lean_ctor_set(x_152, 2, x_150); +x_153 = lean_array_push(x_39, x_152); +x_154 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_154, 0, x_32); +lean_ctor_set(x_154, 1, x_41); +lean_ctor_set(x_154, 2, x_153); +x_155 = lean_array_push(x_122, x_154); +x_156 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_156, 0, x_32); +lean_ctor_set(x_156, 1, x_124); +lean_ctor_set(x_156, 2, x_155); +x_157 = lean_array_push(x_28, x_156); +x_158 = lean_array_push(x_157, x_30); +x_159 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_159, 0, x_32); +lean_ctor_set(x_159, 1, x_41); +lean_ctor_set(x_159, 2, x_158); +lean_inc(x_131); +x_160 = lean_array_push(x_131, x_159); +lean_inc(x_130); +x_161 = lean_array_push(x_160, x_130); +x_162 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_162, 0, x_32); +lean_ctor_set(x_162, 1, x_134); +lean_ctor_set(x_162, 2, x_161); +x_163 = lean_array_push(x_45, x_113); +x_164 = lean_array_push(x_28, x_107); +lean_inc(x_163); +x_165 = lean_array_push(x_163, x_135); +x_166 = lean_array_push(x_28, x_91); +x_167 = lean_array_push(x_45, x_99); +x_168 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__62; +x_169 = lean_array_push(x_167, x_168); +x_170 = lean_array_push(x_45, x_86); +x_171 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__85; +x_172 = lean_array_push(x_171, x_62); +x_173 = lean_array_push(x_172, x_67); +x_174 = lean_array_push(x_173, x_84); +x_175 = lean_array_push(x_28, x_60); +if (lean_obj_tag(x_102) == 0) +{ +lean_object* x_251; +x_251 = l_Lean_quoteNameMk(x_11); +x_176 = x_251; +goto block_250; +} +else +{ +lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_dec(x_11); -x_249 = lean_ctor_get(x_98, 0); -lean_inc(x_249); -lean_dec(x_98); -x_250 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__8; -x_251 = l_String_intercalate(x_250, x_249); -x_252 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__9; -x_253 = lean_string_append(x_252, x_251); -lean_dec(x_251); -x_254 = l_Lean_nameLitKind; -x_255 = l_Lean_Syntax_mkLit(x_254, x_253, x_32); -x_256 = lean_array_push(x_39, x_255); -x_257 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_257, 0, x_32); -lean_ctor_set(x_257, 1, x_94); -lean_ctor_set(x_257, 2, x_256); -x_172 = x_257; -goto block_247; -} -block_247: -{ -lean_object* x_173; lean_object* x_174; lean_object* x_175; -x_173 = lean_array_push(x_45, x_172); -x_174 = lean_array_push(x_173, x_164); -if (lean_obj_tag(x_137) == 0) +x_252 = lean_ctor_get(x_102, 0); +lean_inc(x_252); +lean_dec(x_102); +x_253 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__16; +x_254 = l_String_intercalate(x_253, x_252); +x_255 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__17; +x_256 = lean_string_append(x_255, x_254); +lean_dec(x_254); +x_257 = l_Lean_Syntax_mkNameLit(x_256, x_32); +x_258 = lean_array_push(x_39, x_257); +x_259 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_259, 0, x_32); +lean_ctor_set(x_259, 1, x_98); +lean_ctor_set(x_259, 2, x_258); +x_176 = x_259; +goto block_250; +} +block_250: +{ +lean_object* x_177; lean_object* x_178; lean_object* x_179; +x_177 = lean_array_push(x_45, x_176); +x_178 = lean_array_push(x_177, x_168); +if (lean_obj_tag(x_141) == 0) { -lean_object* x_237; -x_237 = l___private_Init_Meta_0__Lean_quoteNameMk(x_1); -x_175 = x_237; -goto block_236; +lean_object* x_241; +x_241 = l_Lean_quoteNameMk(x_1); +x_179 = x_241; +goto block_240; } else { -lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; +lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_dec(x_1); -x_238 = lean_ctor_get(x_137, 0); -lean_inc(x_238); -lean_dec(x_137); -x_239 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__8; -x_240 = l_String_intercalate(x_239, x_238); -x_241 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__9; -x_242 = lean_string_append(x_241, x_240); -lean_dec(x_240); -x_243 = l_Lean_nameLitKind; -x_244 = l_Lean_Syntax_mkLit(x_243, x_242, x_32); -x_245 = lean_array_push(x_39, x_244); -x_246 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_246, 0, x_32); -lean_ctor_set(x_246, 1, x_94); -lean_ctor_set(x_246, 2, x_245); -x_175 = x_246; -goto block_236; -} -block_236: -{ -lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; -x_176 = lean_array_push(x_28, x_175); -x_177 = lean_array_push(x_176, x_142); -x_178 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_178, 0, x_32); -lean_ctor_set(x_178, 1, x_41); -lean_ctor_set(x_178, 2, x_177); -x_179 = lean_array_push(x_143, x_178); -x_180 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_180, 0, x_32); -lean_ctor_set(x_180, 1, x_120); -lean_ctor_set(x_180, 2, x_179); -x_181 = lean_array_push(x_28, x_180); -x_182 = lean_array_push(x_181, x_30); -x_183 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_183, 0, x_32); -lean_ctor_set(x_183, 1, x_41); -lean_ctor_set(x_183, 2, x_182); -lean_inc(x_127); -x_184 = lean_array_push(x_127, x_183); -lean_inc(x_126); -x_185 = lean_array_push(x_184, x_126); -x_186 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_186, 0, x_32); -lean_ctor_set(x_186, 1, x_130); -lean_ctor_set(x_186, 2, x_185); -x_187 = lean_array_push(x_159, x_186); -x_188 = lean_array_push(x_187, x_158); -x_189 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_189, 0, x_32); -lean_ctor_set(x_189, 1, x_41); -lean_ctor_set(x_189, 2, x_188); -lean_inc(x_160); -x_190 = lean_array_push(x_160, x_189); -x_191 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_191, 0, x_32); -lean_ctor_set(x_191, 1, x_120); -lean_ctor_set(x_191, 2, x_190); -x_192 = lean_array_push(x_28, x_191); -x_193 = lean_array_push(x_192, x_30); -x_194 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_194, 0, x_32); -lean_ctor_set(x_194, 1, x_41); -lean_ctor_set(x_194, 2, x_193); -lean_inc(x_127); -x_195 = lean_array_push(x_127, x_194); -lean_inc(x_126); -x_196 = lean_array_push(x_195, x_126); -x_197 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_197, 0, x_32); -lean_ctor_set(x_197, 1, x_130); -lean_ctor_set(x_197, 2, x_196); -x_198 = lean_array_push(x_161, x_197); -x_199 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_199, 0, x_32); -lean_ctor_set(x_199, 1, x_41); -lean_ctor_set(x_199, 2, x_198); -x_200 = lean_array_push(x_160, x_199); +x_242 = lean_ctor_get(x_141, 0); +lean_inc(x_242); +lean_dec(x_141); +x_243 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__16; +x_244 = l_String_intercalate(x_243, x_242); +x_245 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__17; +x_246 = lean_string_append(x_245, x_244); +lean_dec(x_244); +x_247 = l_Lean_Syntax_mkNameLit(x_246, x_32); +x_248 = lean_array_push(x_39, x_247); +x_249 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_249, 0, x_32); +lean_ctor_set(x_249, 1, x_98); +lean_ctor_set(x_249, 2, x_248); +x_179 = x_249; +goto block_240; +} +block_240: +{ +lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; +x_180 = lean_array_push(x_28, x_179); +x_181 = lean_array_push(x_180, x_146); +x_182 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_182, 0, x_32); +lean_ctor_set(x_182, 1, x_41); +lean_ctor_set(x_182, 2, x_181); +x_183 = lean_array_push(x_147, x_182); +x_184 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_184, 0, x_32); +lean_ctor_set(x_184, 1, x_124); +lean_ctor_set(x_184, 2, x_183); +x_185 = lean_array_push(x_28, x_184); +x_186 = lean_array_push(x_185, x_30); +x_187 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_187, 0, x_32); +lean_ctor_set(x_187, 1, x_41); +lean_ctor_set(x_187, 2, x_186); +lean_inc(x_131); +x_188 = lean_array_push(x_131, x_187); +lean_inc(x_130); +x_189 = lean_array_push(x_188, x_130); +x_190 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_190, 0, x_32); +lean_ctor_set(x_190, 1, x_134); +lean_ctor_set(x_190, 2, x_189); +x_191 = lean_array_push(x_163, x_190); +x_192 = lean_array_push(x_191, x_162); +x_193 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_193, 0, x_32); +lean_ctor_set(x_193, 1, x_41); +lean_ctor_set(x_193, 2, x_192); +lean_inc(x_164); +x_194 = lean_array_push(x_164, x_193); +x_195 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_195, 0, x_32); +lean_ctor_set(x_195, 1, x_124); +lean_ctor_set(x_195, 2, x_194); +x_196 = lean_array_push(x_28, x_195); +x_197 = lean_array_push(x_196, x_30); +x_198 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_198, 0, x_32); +lean_ctor_set(x_198, 1, x_41); +lean_ctor_set(x_198, 2, x_197); +lean_inc(x_131); +x_199 = lean_array_push(x_131, x_198); +lean_inc(x_130); +x_200 = lean_array_push(x_199, x_130); x_201 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_201, 0, x_32); -lean_ctor_set(x_201, 1, x_120); +lean_ctor_set(x_201, 1, x_134); lean_ctor_set(x_201, 2, x_200); -x_202 = lean_array_push(x_28, x_201); -x_203 = lean_array_push(x_202, x_30); -x_204 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_204, 0, x_32); -lean_ctor_set(x_204, 1, x_41); -lean_ctor_set(x_204, 2, x_203); -lean_inc(x_127); -x_205 = lean_array_push(x_127, x_204); -lean_inc(x_126); -x_206 = lean_array_push(x_205, x_126); -x_207 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_207, 0, x_32); -lean_ctor_set(x_207, 1, x_130); -lean_ctor_set(x_207, 2, x_206); -x_208 = lean_array_push(x_174, x_207); -x_209 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_209, 0, x_32); -lean_ctor_set(x_209, 1, x_41); -lean_ctor_set(x_209, 2, x_208); -lean_inc(x_162); -x_210 = lean_array_push(x_162, x_209); +x_202 = lean_array_push(x_165, x_201); +x_203 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_203, 0, x_32); +lean_ctor_set(x_203, 1, x_41); +lean_ctor_set(x_203, 2, x_202); +x_204 = lean_array_push(x_164, x_203); +x_205 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_205, 0, x_32); +lean_ctor_set(x_205, 1, x_124); +lean_ctor_set(x_205, 2, x_204); +x_206 = lean_array_push(x_28, x_205); +x_207 = lean_array_push(x_206, x_30); +x_208 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_208, 0, x_32); +lean_ctor_set(x_208, 1, x_41); +lean_ctor_set(x_208, 2, x_207); +lean_inc(x_131); +x_209 = lean_array_push(x_131, x_208); +lean_inc(x_130); +x_210 = lean_array_push(x_209, x_130); x_211 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_211, 0, x_32); -lean_ctor_set(x_211, 1, x_120); +lean_ctor_set(x_211, 1, x_134); lean_ctor_set(x_211, 2, x_210); -x_212 = lean_array_push(x_28, x_211); -x_213 = lean_array_push(x_212, x_30); -x_214 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_214, 0, x_32); -lean_ctor_set(x_214, 1, x_41); -lean_ctor_set(x_214, 2, x_213); -x_215 = lean_array_push(x_127, x_214); -x_216 = lean_array_push(x_215, x_126); -x_217 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_217, 0, x_32); -lean_ctor_set(x_217, 1, x_130); -lean_ctor_set(x_217, 2, x_216); -x_218 = lean_array_push(x_165, x_217); -x_219 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_219, 0, x_32); -lean_ctor_set(x_219, 1, x_41); -lean_ctor_set(x_219, 2, x_218); -x_220 = lean_array_push(x_162, x_219); +x_212 = lean_array_push(x_178, x_211); +x_213 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_213, 0, x_32); +lean_ctor_set(x_213, 1, x_41); +lean_ctor_set(x_213, 2, x_212); +lean_inc(x_166); +x_214 = lean_array_push(x_166, x_213); +x_215 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_215, 0, x_32); +lean_ctor_set(x_215, 1, x_124); +lean_ctor_set(x_215, 2, x_214); +x_216 = lean_array_push(x_28, x_215); +x_217 = lean_array_push(x_216, x_30); +x_218 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_218, 0, x_32); +lean_ctor_set(x_218, 1, x_41); +lean_ctor_set(x_218, 2, x_217); +x_219 = lean_array_push(x_131, x_218); +x_220 = lean_array_push(x_219, x_130); x_221 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_221, 0, x_32); -lean_ctor_set(x_221, 1, x_120); +lean_ctor_set(x_221, 1, x_134); lean_ctor_set(x_221, 2, x_220); -x_222 = lean_array_push(x_166, x_221); -x_223 = lean_array_push(x_222, x_30); -x_224 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__48; +x_222 = lean_array_push(x_169, x_221); +x_223 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_223, 0, x_32); +lean_ctor_set(x_223, 1, x_41); +lean_ctor_set(x_223, 2, x_222); +x_224 = lean_array_push(x_166, x_223); x_225 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_225, 0, x_32); -lean_ctor_set(x_225, 1, x_224); -lean_ctor_set(x_225, 2, x_223); +lean_ctor_set(x_225, 1, x_124); +lean_ctor_set(x_225, 2, x_224); x_226 = lean_array_push(x_170, x_225); x_227 = lean_array_push(x_226, x_30); -x_228 = lean_array_push(x_227, x_30); -x_229 = lean_array_push(x_228, x_30); -x_230 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__35; -x_231 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_231, 0, x_32); -lean_ctor_set(x_231, 1, x_230); -lean_ctor_set(x_231, 2, x_229); -x_232 = lean_array_push(x_171, x_231); -x_233 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__8; -x_234 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_234, 0, x_32); -lean_ctor_set(x_234, 1, x_233); -lean_ctor_set(x_234, 2, x_232); -x_235 = l_Lean_Elab_Command_elabCommand(x_234, x_2, x_3, x_20); -return x_235; +x_228 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__49; +x_229 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_229, 0, x_32); +lean_ctor_set(x_229, 1, x_228); +lean_ctor_set(x_229, 2, x_227); +x_230 = lean_array_push(x_174, x_229); +x_231 = lean_array_push(x_230, x_30); +x_232 = lean_array_push(x_231, x_30); +x_233 = lean_array_push(x_232, x_30); +x_234 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__34; +x_235 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_235, 0, x_32); +lean_ctor_set(x_235, 1, x_234); +lean_ctor_set(x_235, 2, x_233); +x_236 = lean_array_push(x_175, x_235); +x_237 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__8; +x_238 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_238, 0, x_32); +lean_ctor_set(x_238, 1, x_237); +lean_ctor_set(x_238, 2, x_236); +x_239 = l_Lean_Elab_Command_elabCommand(x_238, x_2, x_3, x_20); +return x_239; } } } else { -lean_object* x_258; lean_object* x_259; +lean_object* x_260; lean_object* x_261; lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_258 = lean_box(0); -x_259 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_259, 0, x_258); -lean_ctor_set(x_259, 1, x_4); -return x_259; +x_260 = lean_box(0); +x_261 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_261, 0, x_260); +lean_ctor_set(x_261, 1, x_4); +return x_261; } } } @@ -13881,7 +14445,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(232u); +x_1 = lean_unsigned_to_nat(255u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13893,7 +14457,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(245u); +x_1 = lean_unsigned_to_nat(268u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13921,7 +14485,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(232u); +x_1 = lean_unsigned_to_nat(255u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13933,7 +14497,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(232u); +x_1 = lean_unsigned_to_nat(255u); x_2 = lean_unsigned_to_nat(56u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -14195,23 +14759,23 @@ LEAN_EXPORT uint8_t l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_isAtomLike lean_object* x_2; lean_object* x_3; uint8_t x_4; lean_inc(x_1); x_2 = l_Lean_Syntax_getKind(x_1); -x_3 = l_Lean_nullKind; +x_3 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; x_4 = lean_name_eq(x_2, x_3); if (x_4 == 0) { lean_object* x_5; uint8_t x_6; -x_5 = l_Lean_choiceKind; +x_5 = l_Lean_Elab_Term_toParserDescr_process___closed__2; x_6 = lean_name_eq(x_2, x_5); if (x_6 == 0) { lean_object* x_7; uint8_t x_8; -x_7 = l_Lean_Elab_Term_toParserDescr_process___closed__2; +x_7 = l_Lean_Elab_Term_toParserDescr_process___closed__4; x_8 = lean_name_eq(x_2, x_7); if (x_8 == 0) { lean_object* x_9; uint8_t x_10; lean_dec(x_1); -x_9 = l_Lean_Elab_Term_toParserDescr_process___closed__11; +x_9 = l_Lean_Elab_Term_toParserDescr_process___closed__12; x_10 = lean_name_eq(x_2, x_9); lean_dec(x_2); return x_10; @@ -14525,48 +15089,128 @@ lean_ctor_set(x_15, 0, x_13); lean_ctor_set(x_15, 1, x_14); x_16 = l_Lean_throwError___at_Lean_Elab_Command_resolveSyntaxKind___spec__3(x_15, x_2, x_3, x_10); lean_dec(x_3); -return x_16; -} -} +return x_16; +} +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_resolveSyntaxKind___spec__3___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_throwError___at_Lean_Elab_Command_resolveSyntaxKind___spec__3(x_1, x_2, x_3, x_4); +lean_dec(x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_checkSyntaxNodeKind___at_Lean_Elab_Command_resolveSyntaxKind___spec__2___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_Elab_checkSyntaxNodeKind___at_Lean_Elab_Command_resolveSyntaxKind___spec__2(x_1, x_2, x_3, x_4); +lean_dec(x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__2___rarg___closed__2; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg), 1, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Elab_Command_elabSyntax___spec__3(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; uint8_t x_7; +x_6 = lean_array_get_size(x_1); +x_7 = lean_nat_dec_lt(x_4, x_6); +lean_dec(x_6); +if (x_7 == 0) +{ +lean_object* x_8; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_5); +return x_8; +} +else +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_unsigned_to_nat(0u); +x_10 = lean_nat_dec_eq(x_3, x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_sub(x_3, x_11); +lean_dec(x_3); +x_13 = lean_array_fget(x_1, x_4); +lean_inc(x_2); +x_14 = lean_apply_1(x_2, x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; +lean_dec(x_12); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_15 = lean_box(0); +return x_15; } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_resolveSyntaxKind___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +else { -lean_object* x_5; -x_5 = l_Lean_throwError___at_Lean_Elab_Command_resolveSyntaxKind___spec__3(x_1, x_2, x_3, x_4); -lean_dec(x_3); -return x_5; +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_14, 0); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_nat_add(x_4, x_11); +lean_dec(x_4); +x_18 = lean_array_push(x_5, x_16); +x_3 = x_12; +x_4 = x_17; +x_5 = x_18; +goto _start; } } -LEAN_EXPORT lean_object* l_Lean_Elab_checkSyntaxNodeKind___at_Lean_Elab_Command_resolveSyntaxKind___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +else { -lean_object* x_5; -x_5 = l_Lean_Elab_checkSyntaxNodeKind___at_Lean_Elab_Command_resolveSyntaxKind___spec__2(x_1, x_2, x_3, x_4); +lean_object* x_20; +lean_dec(x_4); lean_dec(x_3); -return x_5; +lean_dec(x_2); +x_20 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_20, 0, x_5); +return x_20; } } -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_checkLeftRec___spec__8___rarg___closed__2; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Elab_Command_elabSyntax___spec__2(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg), 1, 0); -return x_3; +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_3 = lean_array_get_size(x_1); +x_4 = lean_mk_empty_array_with_capacity(x_3); +x_5 = lean_unsigned_to_nat(0u); +x_6 = l_Array_sequenceMap_loop___at_Lean_Elab_Command_elabSyntax___spec__3(x_1, x_2, x_3, x_5, x_4); +return x_6; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabSyntax___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabSyntax___spec__6(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; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; @@ -14620,7 +15264,7 @@ return x_20; } } } -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__5(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; lean_object* x_9; uint8_t x_10; @@ -14640,7 +15284,7 @@ lean_object* x_11; lean_object* x_12; x_11 = lean_ctor_get(x_3, 6); lean_dec(x_11); lean_ctor_set(x_3, 6, x_9); -x_12 = l_Lean_throwError___at_Lean_Elab_Command_elabSyntax___spec__4(x_2, x_3, x_4, x_8); +x_12 = l_Lean_throwError___at_Lean_Elab_Command_elabSyntax___spec__6(x_2, x_3, x_4, x_8); lean_dec(x_4); return x_12; } @@ -14671,13 +15315,13 @@ lean_ctor_set(x_20, 4, x_17); lean_ctor_set(x_20, 5, x_18); lean_ctor_set(x_20, 6, x_9); lean_ctor_set(x_20, 7, x_19); -x_21 = l_Lean_throwError___at_Lean_Elab_Command_elabSyntax___spec__4(x_2, x_20, x_4, x_8); +x_21 = l_Lean_throwError___at_Lean_Elab_Command_elabSyntax___spec__6(x_2, x_20, x_4, x_8); lean_dec(x_4); return x_21; } } } -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabSyntax___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabSyntax___spec__7(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; @@ -14691,26 +15335,26 @@ lean_ctor_set(x_7, 1, x_4); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__6___rarg(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__8___rarg(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_checkLeftRec___spec__8___rarg___closed__2; +x_2 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__2___rarg___closed__2; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__6(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__8(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__6___rarg), 1, 0); +x_3 = lean_alloc_closure((void*)(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__8___rarg), 1, 0); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabSyntax___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabSyntax___spec__4(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; 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; 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; 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; @@ -14952,14 +15596,14 @@ x_82 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_82, 0, x_79); x_83 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_83, 0, x_82); -x_84 = l_Lean_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__3(x_78, x_83, x_2, x_3, x_36); +x_84 = l_Lean_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__5(x_78, x_83, x_2, x_3, x_36); return x_84; } else { lean_object* x_85; lean_dec(x_79); -x_85 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabSyntax___spec__5(x_78, x_2, x_3, x_36); +x_85 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabSyntax___spec__7(x_78, x_2, x_3, x_36); lean_dec(x_3); lean_dec(x_2); return x_85; @@ -14970,13 +15614,13 @@ else lean_object* x_86; lean_dec(x_3); lean_dec(x_2); -x_86 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__6___rarg(x_36); +x_86 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__8___rarg(x_36); return x_86; } } } } -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__10(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; lean_object* x_9; uint8_t x_10; @@ -15030,7 +15674,7 @@ return x_21; } } } -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabSyntax___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabSyntax___spec__11(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; @@ -15044,26 +15688,26 @@ lean_ctor_set(x_7, 1, x_4); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__10___rarg(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__12___rarg(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_checkLeftRec___spec__8___rarg___closed__2; +x_2 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__2___rarg___closed__2; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__10(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__12(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__10___rarg), 1, 0); +x_3 = lean_alloc_closure((void*)(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__12___rarg), 1, 0); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabSyntax___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabSyntax___spec__9(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; 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; 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; 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; @@ -15305,7 +15949,7 @@ x_82 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_82, 0, x_79); x_83 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_83, 0, x_82); -x_84 = l_Lean_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__8(x_78, x_83, x_2, x_3, x_36); +x_84 = l_Lean_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__10(x_78, x_83, x_2, x_3, x_36); lean_dec(x_3); lean_dec(x_78); return x_84; @@ -15314,7 +15958,7 @@ else { lean_object* x_85; lean_dec(x_79); -x_85 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabSyntax___spec__9(x_78, x_2, x_3, x_36); +x_85 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabSyntax___spec__11(x_78, x_2, x_3, x_36); lean_dec(x_3); lean_dec(x_2); return x_85; @@ -15325,13 +15969,13 @@ else lean_object* x_86; lean_dec(x_3); lean_dec(x_2); -x_86 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__10___rarg(x_36); +x_86 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__12___rarg(x_36); return x_86; } } } } -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__13(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; lean_object* x_9; uint8_t x_10; @@ -15622,44 +16266,26 @@ static lean_object* _init_l_Lean_Elab_Command_elabSyntax___lambda__5___closed__1 _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("declId", 6); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Command_elabSyntax___lambda__5___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__12; -x_3 = lean_array_push(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Elab_Command_elabSyntax___lambda__5___closed__3() { -_start: -{ -lean_object* x_1; x_1 = lean_mk_string_from_bytes("ParserDescr.node", 16); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabSyntax___lambda__5___closed__4() { +static lean_object* _init_l_Lean_Elab_Command_elabSyntax___lambda__5___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__3; +x_1 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__1; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Command_elabSyntax___lambda__5___closed__5() { +static lean_object* _init_l_Lean_Elab_Command_elabSyntax___lambda__5___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__3; +x_1 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__1; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__4; +x_3 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__2; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -15667,17 +16293,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Command_elabSyntax___lambda__5___closed__6() { +static lean_object* _init_l_Lean_Elab_Command_elabSyntax___lambda__5___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__13; -x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__53; +x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__54; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_elabSyntax___lambda__5___closed__7() { +static lean_object* _init_l_Lean_Elab_Command_elabSyntax___lambda__5___closed__5() { _start: { lean_object* x_1; @@ -15685,22 +16311,22 @@ x_1 = lean_mk_string_from_bytes("Lean.TrailingParserDescr", 24); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabSyntax___lambda__5___closed__8() { +static lean_object* _init_l_Lean_Elab_Command_elabSyntax___lambda__5___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__7; +x_1 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__5; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Command_elabSyntax___lambda__5___closed__9() { +static lean_object* _init_l_Lean_Elab_Command_elabSyntax___lambda__5___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__7; +x_1 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__5; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__8; +x_3 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__6; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -15708,7 +16334,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Command_elabSyntax___lambda__5___closed__10() { +static lean_object* _init_l_Lean_Elab_Command_elabSyntax___lambda__5___closed__8() { _start: { lean_object* x_1; @@ -15716,22 +16342,22 @@ x_1 = lean_mk_string_from_bytes("ParserDescr.trailingNode", 24); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabSyntax___lambda__5___closed__11() { +static lean_object* _init_l_Lean_Elab_Command_elabSyntax___lambda__5___closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__10; +x_1 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__8; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Command_elabSyntax___lambda__5___closed__12() { +static lean_object* _init_l_Lean_Elab_Command_elabSyntax___lambda__5___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__10; +x_1 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__8; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__11; +x_3 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__9; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -15739,7 +16365,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Command_elabSyntax___lambda__5___closed__13() { +static lean_object* _init_l_Lean_Elab_Command_elabSyntax___lambda__5___closed__11() { _start: { lean_object* x_1; @@ -15747,429 +16373,454 @@ x_1 = lean_mk_string_from_bytes("trailingNode", 12); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabSyntax___lambda__5___closed__14() { +static lean_object* _init_l_Lean_Elab_Command_elabSyntax___lambda__5___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__13; -x_2 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__13; +x_2 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__11; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSyntax___lambda__5(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_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) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSyntax___lambda__5(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_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, lean_object* x_16) { _start: { -lean_object* x_16; lean_object* x_17; -x_16 = lean_alloc_closure((void*)(l_Lean_evalOptPrio), 3, 1); -lean_closure_set(x_16, 0, x_1); +lean_object* x_17; lean_object* x_18; +x_17 = lean_alloc_closure((void*)(l_Lean_evalOptPrio), 3, 1); +lean_closure_set(x_17, 0, x_1); +lean_inc(x_15); lean_inc(x_14); -lean_inc(x_13); -x_17 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabSyntax___spec__2(x_16, x_13, x_14, x_15); -if (lean_obj_tag(x_17) == 0) +x_18 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabSyntax___spec__4(x_17, x_14, x_15, x_16); +if (lean_obj_tag(x_18) == 0) { -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; 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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); +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; 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; +x_19 = lean_ctor_get(x_18, 0); lean_inc(x_19); -lean_dec(x_17); -x_20 = l_Lean_Elab_Command_getScope___rarg(x_14, x_19); -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_20, 1); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = l_Lean_Elab_Command_getScope___rarg(x_15, x_20); +x_22 = lean_ctor_get(x_21, 0); lean_inc(x_22); -lean_dec(x_20); -x_23 = lean_ctor_get(x_21, 2); +x_23 = lean_ctor_get(x_21, 1); lean_inc(x_23); lean_dec(x_21); -lean_inc(x_12); -x_24 = l_Lean_Name_append(x_23, x_12); -lean_dec(x_23); -x_25 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__3; +x_24 = lean_ctor_get(x_22, 2); +lean_inc(x_24); +lean_dec(x_22); +lean_inc(x_13); +x_25 = l_Lean_Name_append(x_24, x_13); +lean_dec(x_24); +x_26 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__3; lean_inc(x_2); -x_26 = lean_name_append_after(x_2, x_25); +x_27 = lean_name_append_after(x_2, x_26); lean_inc(x_3); -x_27 = l_Lean_mkIdentFrom(x_3, x_26); -x_28 = lean_box(0); -x_29 = l_Lean_Elab_Command_getScope___rarg(x_14, x_22); -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_29, 1); +x_28 = l_Lean_mkIdentFrom(x_3, x_27); +x_29 = lean_box(0); +x_30 = l_Lean_Elab_Command_getScope___rarg(x_15, x_23); +x_31 = lean_ctor_get(x_30, 0); lean_inc(x_31); -lean_dec(x_29); -x_32 = lean_ctor_get(x_30, 5); +x_32 = lean_ctor_get(x_30, 1); lean_inc(x_32); -x_33 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabSyntax___lambda__3___boxed), 11, 3); -lean_closure_set(x_33, 0, x_30); -lean_closure_set(x_33, 1, x_4); -lean_closure_set(x_33, 2, x_2); -x_34 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabBinders___rarg), 9, 2); -lean_closure_set(x_34, 0, x_32); -lean_closure_set(x_34, 1, x_33); -x_35 = lean_alloc_closure((void*)(l_Lean_Elab_Term_withAutoBoundImplicit___rarg), 8, 1); -lean_closure_set(x_35, 0, x_34); -lean_inc(x_13); -x_36 = l_Lean_Elab_Command_liftTermElabM___rarg(x_28, x_35, x_13, x_14, x_31); -if (lean_obj_tag(x_36) == 0) +lean_dec(x_30); +x_33 = lean_ctor_get(x_31, 5); +lean_inc(x_33); +x_34 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabSyntax___lambda__3___boxed), 11, 3); +lean_closure_set(x_34, 0, x_31); +lean_closure_set(x_34, 1, x_4); +lean_closure_set(x_34, 2, x_2); +x_35 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabBinders___rarg), 9, 2); +lean_closure_set(x_35, 0, x_33); +lean_closure_set(x_35, 1, x_34); +x_36 = lean_alloc_closure((void*)(l_Lean_Elab_Term_withAutoBoundImplicit___rarg), 8, 1); +lean_closure_set(x_36, 0, x_35); +lean_inc(x_14); +x_37 = l_Lean_Elab_Command_liftTermElabM___rarg(x_29, x_36, x_14, x_15, x_32); +if (lean_obj_tag(x_37) == 0) { -lean_object* x_37; lean_object* x_38; uint8_t x_39; -x_37 = lean_ctor_get(x_36, 0); -lean_inc(x_37); -x_38 = lean_ctor_get(x_36, 1); +lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_38 = lean_ctor_get(x_37, 0); lean_inc(x_38); -lean_dec(x_36); -x_39 = !lean_is_exclusive(x_37); -if (x_39 == 0) -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_37, 0); -x_41 = lean_ctor_get(x_37, 1); -lean_inc(x_3); -x_42 = l_Lean_mkIdentFrom(x_3, x_12); -if (lean_obj_tag(x_41) == 0) +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_37, 1); +lean_inc(x_40); +lean_dec(x_37); +x_41 = !lean_is_exclusive(x_38); +if (x_41 == 0) { -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; -x_43 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_13, x_14, x_38); -x_44 = lean_ctor_get(x_43, 0); -lean_inc(x_44); -x_45 = lean_ctor_get(x_43, 1); -lean_inc(x_45); +lean_object* x_42; lean_object* x_43; uint8_t x_44; +x_42 = lean_ctor_get(x_38, 1); +x_43 = lean_ctor_get(x_38, 0); lean_dec(x_43); -x_46 = l_Lean_Elab_Command_getCurrMacroScope(x_13, x_14, x_45); -x_47 = lean_ctor_get(x_46, 0); -lean_inc(x_47); -x_48 = lean_ctor_get(x_46, 1); -lean_inc(x_48); +x_44 = !lean_is_exclusive(x_39); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_39, 0); +x_46 = lean_ctor_get(x_39, 1); lean_dec(x_46); -x_49 = l_Lean_Elab_Command_getMainModule___rarg(x_14, x_48); -x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_3); +x_47 = l_Lean_mkIdentFrom(x_3, x_13); +if (lean_obj_tag(x_42) == 0) +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; +x_48 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_14, x_15, x_40); +x_49 = lean_ctor_get(x_48, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_48, 1); lean_inc(x_50); -x_51 = lean_ctor_get(x_49, 1); -lean_inc(x_51); -lean_dec(x_49); -x_52 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; +lean_dec(x_48); +x_51 = l_Lean_Elab_Command_getCurrMacroScope(x_14, x_15, x_50); +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +lean_dec(x_51); +x_54 = l_Lean_Elab_Command_getMainModule___rarg(x_15, x_53); +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_54, 1); +lean_inc(x_56); +lean_dec(x_54); +x_57 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; lean_inc(x_5); -x_53 = lean_name_mk_string(x_5, x_52); -x_54 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; +x_58 = lean_name_mk_string(x_5, x_57); +x_59 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; lean_inc(x_5); -x_55 = lean_name_mk_string(x_5, x_54); -x_56 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__13; +x_60 = lean_name_mk_string(x_5, x_59); +x_61 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__12; lean_inc(x_6); -x_57 = lean_name_mk_string(x_6, x_56); -x_58 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__15; -lean_inc(x_44); -x_59 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_59, 0, x_44); -lean_ctor_set(x_59, 1, x_58); -x_60 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__16; +x_62 = lean_name_mk_string(x_6, x_61); +x_63 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__14; +lean_inc(x_49); +x_64 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_64, 0, x_49); +lean_ctor_set(x_64, 1, x_63); +x_65 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__15; lean_inc(x_6); -x_61 = lean_name_mk_string(x_6, x_60); -x_62 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__22; -x_63 = lean_name_mk_string(x_7, x_62); -x_64 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__24; -x_65 = lean_name_mk_string(x_63, x_64); -x_66 = l_Nat_repr(x_18); -x_67 = l_Lean_numLitKind; -x_68 = lean_box(2); -x_69 = l_Lean_Syntax_mkLit(x_67, x_66, x_68); -x_70 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; -x_71 = lean_array_push(x_70, x_69); -x_72 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_73 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_73, 0, x_68); -lean_ctor_set(x_73, 1, x_72); -lean_ctor_set(x_73, 2, x_71); -x_74 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; -x_75 = lean_array_push(x_74, x_27); -x_76 = lean_array_push(x_75, x_73); -x_77 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_77, 0, x_68); -lean_ctor_set(x_77, 1, x_65); -lean_ctor_set(x_77, 2, x_76); -x_78 = lean_array_push(x_74, x_8); -x_79 = lean_array_push(x_78, x_77); +x_66 = lean_name_mk_string(x_6, x_65); +x_67 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__21; +x_68 = lean_name_mk_string(x_7, x_67); +x_69 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__23; +x_70 = lean_name_mk_string(x_68, x_69); +x_71 = l_Nat_repr(x_19); +x_72 = lean_box(2); +x_73 = l_Lean_Syntax_mkNumLit(x_71, x_72); +x_74 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; +x_75 = lean_array_push(x_74, x_73); +lean_inc(x_8); +x_76 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_76, 0, x_72); +lean_ctor_set(x_76, 1, x_8); +lean_ctor_set(x_76, 2, x_75); +x_77 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; +x_78 = lean_array_push(x_77, x_28); +x_79 = lean_array_push(x_78, x_76); x_80 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_80, 0, x_68); -lean_ctor_set(x_80, 1, x_61); +lean_ctor_set(x_80, 0, x_72); +lean_ctor_set(x_80, 1, x_70); lean_ctor_set(x_80, 2, x_79); -x_81 = lean_array_push(x_70, x_80); -x_82 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_82, 0, x_68); -lean_ctor_set(x_82, 1, x_72); -lean_ctor_set(x_82, 2, x_81); -x_83 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31; -lean_inc(x_44); -x_84 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_84, 0, x_44); -lean_ctor_set(x_84, 1, x_83); -x_85 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__28; -x_86 = lean_array_push(x_85, x_59); -x_87 = lean_array_push(x_86, x_82); -x_88 = lean_array_push(x_87, x_84); -x_89 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_89, 0, x_68); -lean_ctor_set(x_89, 1, x_57); -lean_ctor_set(x_89, 2, x_88); -x_90 = lean_array_push(x_70, x_89); -x_91 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_91, 0, x_68); -lean_ctor_set(x_91, 1, x_72); -lean_ctor_set(x_91, 2, x_90); -x_92 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__34; +x_81 = lean_array_push(x_77, x_9); +x_82 = lean_array_push(x_81, x_80); +x_83 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_83, 0, x_72); +lean_ctor_set(x_83, 1, x_66); +lean_ctor_set(x_83, 2, x_82); +x_84 = lean_array_push(x_74, x_83); +lean_inc(x_8); +x_85 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_85, 0, x_72); +lean_ctor_set(x_85, 1, x_8); +lean_ctor_set(x_85, 2, x_84); +x_86 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__30; +lean_inc(x_49); +x_87 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_87, 0, x_49); +lean_ctor_set(x_87, 1, x_86); +x_88 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__28; +x_89 = lean_array_push(x_88, x_64); +x_90 = lean_array_push(x_89, x_85); +x_91 = lean_array_push(x_90, x_87); +x_92 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_92, 0, x_72); +lean_ctor_set(x_92, 1, x_62); +lean_ctor_set(x_92, 2, x_91); +x_93 = lean_array_push(x_74, x_92); +lean_inc(x_8); +x_94 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_94, 0, x_72); +lean_ctor_set(x_94, 1, x_8); +lean_ctor_set(x_94, 2, x_93); +x_95 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__3; +lean_inc(x_8); +x_96 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_96, 0, x_72); +lean_ctor_set(x_96, 1, x_8); +lean_ctor_set(x_96, 2, x_95); +x_97 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33; lean_inc(x_5); -x_93 = lean_name_mk_string(x_5, x_92); -lean_inc(x_44); -x_94 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_94, 0, x_44); -lean_ctor_set(x_94, 1, x_92); -x_95 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__1; +x_98 = lean_name_mk_string(x_5, x_97); +lean_inc(x_49); +x_99 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_99, 0, x_49); +lean_ctor_set(x_99, 1, x_97); +x_100 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__35; lean_inc(x_5); -x_96 = lean_name_mk_string(x_5, x_95); -x_97 = lean_array_push(x_74, x_42); -x_98 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__12; -x_99 = lean_array_push(x_97, x_98); -x_100 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_100, 0, x_68); -lean_ctor_set(x_100, 1, x_96); -lean_ctor_set(x_100, 2, x_99); -x_101 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__36; +x_101 = lean_name_mk_string(x_5, x_100); +x_102 = lean_array_push(x_77, x_47); +lean_inc(x_96); +x_103 = lean_array_push(x_102, x_96); +x_104 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_104, 0, x_72); +lean_ctor_set(x_104, 1, x_101); +lean_ctor_set(x_104, 2, x_103); +x_105 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__37; lean_inc(x_5); -x_102 = lean_name_mk_string(x_5, x_101); -x_103 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__38; +x_106 = lean_name_mk_string(x_5, x_105); +x_107 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__39; lean_inc(x_6); -x_104 = lean_name_mk_string(x_6, x_103); -x_105 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__40; -lean_inc(x_44); -x_106 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_106, 0, x_44); -lean_ctor_set(x_106, 1, x_105); -x_107 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__12; -x_108 = lean_name_mk_string(x_9, x_107); -lean_inc(x_47); -lean_inc(x_108); -lean_inc(x_50); -x_109 = l_Lean_addMacroScope(x_50, x_108, x_47); -x_110 = lean_box(0); -lean_inc(x_108); -lean_ctor_set(x_37, 1, x_110); -lean_ctor_set(x_37, 0, x_108); -x_111 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_111, 0, x_37); -lean_ctor_set(x_111, 1, x_110); -x_112 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__43; -lean_inc(x_44); -x_113 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_113, 0, x_44); -lean_ctor_set(x_113, 1, x_112); -lean_ctor_set(x_113, 2, x_109); -lean_ctor_set(x_113, 3, x_111); -x_114 = lean_array_push(x_74, x_106); -x_115 = lean_array_push(x_114, x_113); -x_116 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_116, 0, x_68); -lean_ctor_set(x_116, 1, x_104); -lean_ctor_set(x_116, 2, x_115); -x_117 = lean_array_push(x_70, x_116); -x_118 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_118, 0, x_68); -lean_ctor_set(x_118, 1, x_72); -lean_ctor_set(x_118, 2, x_117); -x_119 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__2; -x_120 = lean_array_push(x_119, x_118); -x_121 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_121, 0, x_68); -lean_ctor_set(x_121, 1, x_102); -lean_ctor_set(x_121, 2, x_120); -x_122 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__47; -x_123 = lean_name_mk_string(x_5, x_122); -x_124 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__49; -lean_inc(x_44); -x_125 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_125, 0, x_44); -lean_ctor_set(x_125, 1, x_124); -x_126 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__7; +x_108 = lean_name_mk_string(x_6, x_107); +x_109 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__41; +lean_inc(x_49); +x_110 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_110, 0, x_49); +lean_ctor_set(x_110, 1, x_109); +x_111 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__12; +x_112 = lean_name_mk_string(x_10, x_111); +lean_inc(x_52); +lean_inc(x_112); +lean_inc(x_55); +x_113 = l_Lean_addMacroScope(x_55, x_112, x_52); +x_114 = lean_box(0); +lean_inc(x_112); +lean_ctor_set(x_39, 1, x_114); +lean_ctor_set(x_39, 0, x_112); +x_115 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_115, 0, x_39); +lean_ctor_set(x_115, 1, x_114); +x_116 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__44; +lean_inc(x_49); +x_117 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_117, 0, x_49); +lean_ctor_set(x_117, 1, x_116); +lean_ctor_set(x_117, 2, x_113); +lean_ctor_set(x_117, 3, x_115); +x_118 = lean_array_push(x_77, x_110); +x_119 = lean_array_push(x_118, x_117); +x_120 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_120, 0, x_72); +lean_ctor_set(x_120, 1, x_108); +lean_ctor_set(x_120, 2, x_119); +x_121 = lean_array_push(x_74, x_120); +lean_inc(x_8); +x_122 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_122, 0, x_72); +lean_ctor_set(x_122, 1, x_8); +lean_ctor_set(x_122, 2, x_121); +lean_inc(x_96); +x_123 = lean_array_push(x_77, x_96); +x_124 = lean_array_push(x_123, x_122); +x_125 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_125, 0, x_72); +lean_ctor_set(x_125, 1, x_106); +lean_ctor_set(x_125, 2, x_124); +x_126 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__48; +x_127 = lean_name_mk_string(x_5, x_126); +x_128 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__50; +lean_inc(x_49); +x_129 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_129, 0, x_49); +lean_ctor_set(x_129, 1, x_128); +x_130 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__7; lean_inc(x_6); -x_127 = lean_name_mk_string(x_6, x_126); -x_128 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__6; -x_129 = l_Lean_addMacroScope(x_50, x_128, x_47); -x_130 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__53; -x_131 = lean_name_mk_string(x_108, x_130); -x_132 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_132, 0, x_131); -lean_ctor_set(x_132, 1, x_110); -x_133 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_133, 0, x_132); -lean_ctor_set(x_133, 1, x_110); -x_134 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__5; -x_135 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_135, 0, x_44); -lean_ctor_set(x_135, 1, x_134); -lean_ctor_set(x_135, 2, x_129); -lean_ctor_set(x_135, 3, x_133); -lean_inc(x_24); -x_136 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_110, x_24); -x_137 = l_Nat_repr(x_10); -x_138 = l_Lean_Syntax_mkLit(x_67, x_137, x_68); -x_139 = lean_array_push(x_74, x_135); -x_140 = lean_array_push(x_85, x_125); -x_141 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__84; -x_142 = lean_array_push(x_141, x_94); -x_143 = lean_array_push(x_142, x_100); -x_144 = lean_array_push(x_143, x_121); -if (lean_obj_tag(x_11) == 0) +x_131 = lean_name_mk_string(x_6, x_130); +x_132 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__4; +x_133 = l_Lean_addMacroScope(x_55, x_132, x_52); +x_134 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__54; +x_135 = lean_name_mk_string(x_112, x_134); +lean_ctor_set(x_38, 1, x_114); +lean_ctor_set(x_38, 0, x_135); +x_136 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_136, 0, x_38); +lean_ctor_set(x_136, 1, x_114); +x_137 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__3; +x_138 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_138, 0, x_49); +lean_ctor_set(x_138, 1, x_137); +lean_ctor_set(x_138, 2, x_133); +lean_ctor_set(x_138, 3, x_136); +lean_inc(x_25); +x_139 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_114, x_25); +x_140 = l_Nat_repr(x_11); +x_141 = l_Lean_Syntax_mkNumLit(x_140, x_72); +x_142 = lean_array_push(x_77, x_138); +x_143 = lean_array_push(x_88, x_129); +x_144 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__85; +x_145 = lean_array_push(x_144, x_99); +x_146 = lean_array_push(x_145, x_104); +x_147 = lean_array_push(x_146, x_125); +if (lean_obj_tag(x_12) == 0) { -lean_object* x_205; -x_205 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__11; -x_145 = x_205; -goto block_204; +x_148 = x_95; +goto block_205; } else { lean_object* x_206; lean_object* x_207; -x_206 = lean_ctor_get(x_11, 0); +x_206 = lean_ctor_get(x_12, 0); lean_inc(x_206); -lean_dec(x_11); -x_207 = lean_array_push(x_70, x_206); -x_145 = x_207; -goto block_204; +lean_dec(x_12); +x_207 = lean_array_push(x_74, x_206); +x_148 = x_207; +goto block_205; } -block_204: +block_205: { -lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; -x_146 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__11; -x_147 = l_Array_append___rarg(x_146, x_145); -x_148 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_148, 0, x_68); -lean_ctor_set(x_148, 1, x_72); -lean_ctor_set(x_148, 2, x_147); -x_149 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__32; -x_150 = lean_array_push(x_149, x_148); -x_151 = lean_array_push(x_150, x_91); -x_152 = lean_array_push(x_151, x_98); -x_153 = lean_array_push(x_152, x_98); -x_154 = lean_array_push(x_153, x_98); -x_155 = lean_array_push(x_154, x_98); -x_156 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_156, 0, x_68); -lean_ctor_set(x_156, 1, x_55); -lean_ctor_set(x_156, 2, x_155); -x_157 = lean_array_push(x_74, x_156); -if (lean_obj_tag(x_136) == 0) +lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; +x_149 = l_Array_append___rarg(x_95, x_148); +lean_inc(x_8); +x_150 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_150, 0, x_72); +lean_ctor_set(x_150, 1, x_8); +lean_ctor_set(x_150, 2, x_149); +x_151 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31; +x_152 = lean_array_push(x_151, x_150); +x_153 = lean_array_push(x_152, x_94); +lean_inc(x_96); +x_154 = lean_array_push(x_153, x_96); +lean_inc(x_96); +x_155 = lean_array_push(x_154, x_96); +lean_inc(x_96); +x_156 = lean_array_push(x_155, x_96); +lean_inc(x_96); +x_157 = lean_array_push(x_156, x_96); +x_158 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_158, 0, x_72); +lean_ctor_set(x_158, 1, x_60); +lean_ctor_set(x_158, 2, x_157); +x_159 = lean_array_push(x_77, x_158); +if (lean_obj_tag(x_139) == 0) { -lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; +lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_dec(x_6); -x_158 = l___private_Init_Meta_0__Lean_quoteNameMk(x_24); -x_159 = lean_array_push(x_85, x_158); -x_160 = lean_array_push(x_159, x_138); -x_161 = lean_array_push(x_160, x_40); -x_162 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_162, 0, x_68); -lean_ctor_set(x_162, 1, x_72); -lean_ctor_set(x_162, 2, x_161); -x_163 = lean_array_push(x_139, x_162); +x_160 = l_Lean_quoteNameMk(x_25); +x_161 = lean_array_push(x_88, x_160); +x_162 = lean_array_push(x_161, x_141); +x_163 = lean_array_push(x_162, x_45); x_164 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_164, 0, x_68); -lean_ctor_set(x_164, 1, x_127); +lean_ctor_set(x_164, 0, x_72); +lean_ctor_set(x_164, 1, x_8); lean_ctor_set(x_164, 2, x_163); -x_165 = lean_array_push(x_140, x_164); -x_166 = lean_array_push(x_165, x_98); -x_167 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_167, 0, x_68); -lean_ctor_set(x_167, 1, x_123); -lean_ctor_set(x_167, 2, x_166); -x_168 = lean_array_push(x_144, x_167); -x_169 = lean_array_push(x_168, x_98); -x_170 = lean_array_push(x_169, x_98); -x_171 = lean_array_push(x_170, x_98); -x_172 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_172, 0, x_68); -lean_ctor_set(x_172, 1, x_93); -lean_ctor_set(x_172, 2, x_171); -x_173 = lean_array_push(x_157, x_172); +x_165 = lean_array_push(x_142, x_164); +x_166 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_166, 0, x_72); +lean_ctor_set(x_166, 1, x_131); +lean_ctor_set(x_166, 2, x_165); +x_167 = lean_array_push(x_143, x_166); +lean_inc(x_96); +x_168 = lean_array_push(x_167, x_96); +x_169 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_169, 0, x_72); +lean_ctor_set(x_169, 1, x_127); +lean_ctor_set(x_169, 2, x_168); +x_170 = lean_array_push(x_147, x_169); +lean_inc(x_96); +x_171 = lean_array_push(x_170, x_96); +lean_inc(x_96); +x_172 = lean_array_push(x_171, x_96); +x_173 = lean_array_push(x_172, x_96); x_174 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_174, 0, x_68); -lean_ctor_set(x_174, 1, x_53); +lean_ctor_set(x_174, 0, x_72); +lean_ctor_set(x_174, 1, x_98); lean_ctor_set(x_174, 2, x_173); -x_175 = l_Lean_Elab_Command_elabSyntax___lambda__4(x_3, x_174, x_13, x_14, x_51); -return x_175; +x_175 = lean_array_push(x_159, x_174); +x_176 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_176, 0, x_72); +lean_ctor_set(x_176, 1, x_58); +lean_ctor_set(x_176, 2, x_175); +x_177 = l_Lean_Elab_Command_elabSyntax___lambda__4(x_3, x_176, x_14, x_15, x_56); +return x_177; } else { -lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; -lean_dec(x_24); -x_176 = lean_ctor_get(x_136, 0); -lean_inc(x_176); -lean_dec(x_136); -x_177 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__22; -x_178 = lean_name_mk_string(x_6, x_177); -x_179 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__8; -x_180 = l_String_intercalate(x_179, x_176); -x_181 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__9; -x_182 = lean_string_append(x_181, x_180); -lean_dec(x_180); -x_183 = l_Lean_nameLitKind; -x_184 = l_Lean_Syntax_mkLit(x_183, x_182, x_68); -x_185 = lean_array_push(x_70, x_184); -x_186 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_186, 0, x_68); -lean_ctor_set(x_186, 1, x_178); -lean_ctor_set(x_186, 2, x_185); -x_187 = lean_array_push(x_85, x_186); -x_188 = lean_array_push(x_187, x_138); -x_189 = lean_array_push(x_188, x_40); -x_190 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_190, 0, x_68); -lean_ctor_set(x_190, 1, x_72); -lean_ctor_set(x_190, 2, x_189); -x_191 = lean_array_push(x_139, x_190); -x_192 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_192, 0, x_68); -lean_ctor_set(x_192, 1, x_127); -lean_ctor_set(x_192, 2, x_191); -x_193 = lean_array_push(x_140, x_192); -x_194 = lean_array_push(x_193, x_98); -x_195 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_195, 0, x_68); -lean_ctor_set(x_195, 1, x_123); -lean_ctor_set(x_195, 2, x_194); -x_196 = lean_array_push(x_144, x_195); -x_197 = lean_array_push(x_196, x_98); -x_198 = lean_array_push(x_197, x_98); -x_199 = lean_array_push(x_198, x_98); -x_200 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_200, 0, x_68); -lean_ctor_set(x_200, 1, x_93); -lean_ctor_set(x_200, 2, x_199); -x_201 = lean_array_push(x_157, x_200); -x_202 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_202, 0, x_68); -lean_ctor_set(x_202, 1, x_53); -lean_ctor_set(x_202, 2, x_201); -x_203 = l_Lean_Elab_Command_elabSyntax___lambda__4(x_3, x_202, x_13, x_14, x_51); -return x_203; +lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; +lean_dec(x_25); +x_178 = lean_ctor_get(x_139, 0); +lean_inc(x_178); +lean_dec(x_139); +x_179 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__22; +x_180 = lean_name_mk_string(x_6, x_179); +x_181 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__16; +x_182 = l_String_intercalate(x_181, x_178); +x_183 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__17; +x_184 = lean_string_append(x_183, x_182); +lean_dec(x_182); +x_185 = l_Lean_Syntax_mkNameLit(x_184, x_72); +x_186 = lean_array_push(x_74, x_185); +x_187 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_187, 0, x_72); +lean_ctor_set(x_187, 1, x_180); +lean_ctor_set(x_187, 2, x_186); +x_188 = lean_array_push(x_88, x_187); +x_189 = lean_array_push(x_188, x_141); +x_190 = lean_array_push(x_189, x_45); +x_191 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_191, 0, x_72); +lean_ctor_set(x_191, 1, x_8); +lean_ctor_set(x_191, 2, x_190); +x_192 = lean_array_push(x_142, x_191); +x_193 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_193, 0, x_72); +lean_ctor_set(x_193, 1, x_131); +lean_ctor_set(x_193, 2, x_192); +x_194 = lean_array_push(x_143, x_193); +lean_inc(x_96); +x_195 = lean_array_push(x_194, x_96); +x_196 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_196, 0, x_72); +lean_ctor_set(x_196, 1, x_127); +lean_ctor_set(x_196, 2, x_195); +x_197 = lean_array_push(x_147, x_196); +lean_inc(x_96); +x_198 = lean_array_push(x_197, x_96); +lean_inc(x_96); +x_199 = lean_array_push(x_198, x_96); +x_200 = lean_array_push(x_199, x_96); +x_201 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_201, 0, x_72); +lean_ctor_set(x_201, 1, x_98); +lean_ctor_set(x_201, 2, x_200); +x_202 = lean_array_push(x_159, x_201); +x_203 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_203, 0, x_72); +lean_ctor_set(x_203, 1, x_58); +lean_ctor_set(x_203, 2, x_202); +x_204 = l_Lean_Elab_Command_elabSyntax___lambda__4(x_3, x_203, x_14, x_15, x_56); +return x_204; } } } else { -lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; -x_208 = lean_ctor_get(x_41, 0); +lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; +x_208 = lean_ctor_get(x_42, 0); lean_inc(x_208); -lean_dec(x_41); -x_209 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_13, x_14, x_38); +lean_dec(x_42); +x_209 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_14, x_15, x_40); x_210 = lean_ctor_get(x_209, 0); lean_inc(x_210); x_211 = lean_ctor_get(x_209, 1); lean_inc(x_211); lean_dec(x_209); -x_212 = l_Lean_Elab_Command_getCurrMacroScope(x_13, x_14, x_211); +x_212 = l_Lean_Elab_Command_getCurrMacroScope(x_14, x_15, x_211); x_213 = lean_ctor_get(x_212, 0); lean_inc(x_213); x_214 = lean_ctor_get(x_212, 1); lean_inc(x_214); lean_dec(x_212); -x_215 = l_Lean_Elab_Command_getMainModule___rarg(x_14, x_214); +x_215 = l_Lean_Elab_Command_getMainModule___rarg(x_15, x_214); x_216 = lean_ctor_get(x_215, 0); lean_inc(x_216); x_217 = lean_ctor_get(x_215, 1); @@ -16181,991 +16832,1756 @@ x_219 = lean_name_mk_string(x_5, x_218); x_220 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; lean_inc(x_5); x_221 = lean_name_mk_string(x_5, x_220); -x_222 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__13; +x_222 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__12; lean_inc(x_6); x_223 = lean_name_mk_string(x_6, x_222); -x_224 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__15; +x_224 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__14; lean_inc(x_210); x_225 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_225, 0, x_210); lean_ctor_set(x_225, 1, x_224); -x_226 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__16; +x_226 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__15; lean_inc(x_6); x_227 = lean_name_mk_string(x_6, x_226); -x_228 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__22; +x_228 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__21; x_229 = lean_name_mk_string(x_7, x_228); -x_230 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__24; +x_230 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__23; x_231 = lean_name_mk_string(x_229, x_230); -x_232 = l_Nat_repr(x_18); -x_233 = l_Lean_numLitKind; -x_234 = lean_box(2); -x_235 = l_Lean_Syntax_mkLit(x_233, x_232, x_234); -x_236 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; -x_237 = lean_array_push(x_236, x_235); -x_238 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_239 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_239, 0, x_234); -lean_ctor_set(x_239, 1, x_238); -lean_ctor_set(x_239, 2, x_237); -x_240 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; -x_241 = lean_array_push(x_240, x_27); -x_242 = lean_array_push(x_241, x_239); -x_243 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_243, 0, x_234); -lean_ctor_set(x_243, 1, x_231); -lean_ctor_set(x_243, 2, x_242); -x_244 = lean_array_push(x_240, x_8); -x_245 = lean_array_push(x_244, x_243); +x_232 = l_Nat_repr(x_19); +x_233 = lean_box(2); +x_234 = l_Lean_Syntax_mkNumLit(x_232, x_233); +x_235 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; +x_236 = lean_array_push(x_235, x_234); +lean_inc(x_8); +x_237 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_237, 0, x_233); +lean_ctor_set(x_237, 1, x_8); +lean_ctor_set(x_237, 2, x_236); +x_238 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; +x_239 = lean_array_push(x_238, x_28); +x_240 = lean_array_push(x_239, x_237); +x_241 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_241, 0, x_233); +lean_ctor_set(x_241, 1, x_231); +lean_ctor_set(x_241, 2, x_240); +x_242 = lean_array_push(x_238, x_9); +x_243 = lean_array_push(x_242, x_241); +x_244 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_244, 0, x_233); +lean_ctor_set(x_244, 1, x_227); +lean_ctor_set(x_244, 2, x_243); +x_245 = lean_array_push(x_235, x_244); +lean_inc(x_8); x_246 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_246, 0, x_234); -lean_ctor_set(x_246, 1, x_227); +lean_ctor_set(x_246, 0, x_233); +lean_ctor_set(x_246, 1, x_8); lean_ctor_set(x_246, 2, x_245); -x_247 = lean_array_push(x_236, x_246); -x_248 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_248, 0, x_234); -lean_ctor_set(x_248, 1, x_238); -lean_ctor_set(x_248, 2, x_247); -x_249 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31; +x_247 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__30; lean_inc(x_210); -x_250 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_250, 0, x_210); -lean_ctor_set(x_250, 1, x_249); -x_251 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__28; -x_252 = lean_array_push(x_251, x_225); -x_253 = lean_array_push(x_252, x_248); -x_254 = lean_array_push(x_253, x_250); +x_248 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_248, 0, x_210); +lean_ctor_set(x_248, 1, x_247); +x_249 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__28; +x_250 = lean_array_push(x_249, x_225); +x_251 = lean_array_push(x_250, x_246); +x_252 = lean_array_push(x_251, x_248); +x_253 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_253, 0, x_233); +lean_ctor_set(x_253, 1, x_223); +lean_ctor_set(x_253, 2, x_252); +x_254 = lean_array_push(x_235, x_253); +lean_inc(x_8); x_255 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_255, 0, x_234); -lean_ctor_set(x_255, 1, x_223); +lean_ctor_set(x_255, 0, x_233); +lean_ctor_set(x_255, 1, x_8); lean_ctor_set(x_255, 2, x_254); -x_256 = lean_array_push(x_236, x_255); +x_256 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__3; +lean_inc(x_8); x_257 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_257, 0, x_234); -lean_ctor_set(x_257, 1, x_238); +lean_ctor_set(x_257, 0, x_233); +lean_ctor_set(x_257, 1, x_8); lean_ctor_set(x_257, 2, x_256); -x_258 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__34; +x_258 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33; lean_inc(x_5); x_259 = lean_name_mk_string(x_5, x_258); lean_inc(x_210); x_260 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_260, 0, x_210); lean_ctor_set(x_260, 1, x_258); -x_261 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__1; +x_261 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__35; lean_inc(x_5); x_262 = lean_name_mk_string(x_5, x_261); -x_263 = lean_array_push(x_240, x_42); -x_264 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__12; -x_265 = lean_array_push(x_263, x_264); -x_266 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_266, 0, x_234); -lean_ctor_set(x_266, 1, x_262); -lean_ctor_set(x_266, 2, x_265); -x_267 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__36; +x_263 = lean_array_push(x_238, x_47); +lean_inc(x_257); +x_264 = lean_array_push(x_263, x_257); +x_265 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_265, 0, x_233); +lean_ctor_set(x_265, 1, x_262); +lean_ctor_set(x_265, 2, x_264); +x_266 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__37; lean_inc(x_5); -x_268 = lean_name_mk_string(x_5, x_267); -x_269 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__38; +x_267 = lean_name_mk_string(x_5, x_266); +x_268 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__39; lean_inc(x_6); -x_270 = lean_name_mk_string(x_6, x_269); -x_271 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__40; +x_269 = lean_name_mk_string(x_6, x_268); +x_270 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__41; lean_inc(x_210); -x_272 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_272, 0, x_210); -lean_ctor_set(x_272, 1, x_271); -x_273 = l_List_filterMap___at_Lean_Elab_Term_resolveParserName___spec__1___closed__1; -lean_inc(x_9); -x_274 = lean_name_mk_string(x_9, x_273); +x_271 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_271, 0, x_210); +lean_ctor_set(x_271, 1, x_270); +x_272 = l_List_filterMap___at_Lean_Elab_Term_resolveParserName___spec__1___closed__1; +lean_inc(x_10); +x_273 = lean_name_mk_string(x_10, x_272); lean_inc(x_213); -lean_inc(x_274); +lean_inc(x_273); lean_inc(x_216); -x_275 = l_Lean_addMacroScope(x_216, x_274, x_213); -x_276 = lean_box(0); -lean_ctor_set(x_37, 1, x_276); -lean_ctor_set(x_37, 0, x_274); -x_277 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_277, 0, x_37); -lean_ctor_set(x_277, 1, x_276); -x_278 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__9; +x_274 = l_Lean_addMacroScope(x_216, x_273, x_213); +x_275 = lean_box(0); +lean_ctor_set(x_39, 1, x_275); +lean_ctor_set(x_39, 0, x_273); +x_276 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_276, 0, x_39); +lean_ctor_set(x_276, 1, x_275); +x_277 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__7; lean_inc(x_210); -x_279 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_279, 0, x_210); -lean_ctor_set(x_279, 1, x_278); -lean_ctor_set(x_279, 2, x_275); -lean_ctor_set(x_279, 3, x_277); -x_280 = lean_array_push(x_240, x_272); -x_281 = lean_array_push(x_280, x_279); -x_282 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_282, 0, x_234); -lean_ctor_set(x_282, 1, x_270); -lean_ctor_set(x_282, 2, x_281); -x_283 = lean_array_push(x_236, x_282); -x_284 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_284, 0, x_234); -lean_ctor_set(x_284, 1, x_238); -lean_ctor_set(x_284, 2, x_283); -x_285 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__2; -x_286 = lean_array_push(x_285, x_284); -x_287 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_287, 0, x_234); -lean_ctor_set(x_287, 1, x_268); -lean_ctor_set(x_287, 2, x_286); -x_288 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__47; -x_289 = lean_name_mk_string(x_5, x_288); -x_290 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__49; +x_278 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_278, 0, x_210); +lean_ctor_set(x_278, 1, x_277); +lean_ctor_set(x_278, 2, x_274); +lean_ctor_set(x_278, 3, x_276); +x_279 = lean_array_push(x_238, x_271); +x_280 = lean_array_push(x_279, x_278); +x_281 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_281, 0, x_233); +lean_ctor_set(x_281, 1, x_269); +lean_ctor_set(x_281, 2, x_280); +x_282 = lean_array_push(x_235, x_281); +lean_inc(x_8); +x_283 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_283, 0, x_233); +lean_ctor_set(x_283, 1, x_8); +lean_ctor_set(x_283, 2, x_282); +lean_inc(x_257); +x_284 = lean_array_push(x_238, x_257); +x_285 = lean_array_push(x_284, x_283); +x_286 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_286, 0, x_233); +lean_ctor_set(x_286, 1, x_267); +lean_ctor_set(x_286, 2, x_285); +x_287 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__48; +x_288 = lean_name_mk_string(x_5, x_287); +x_289 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__50; lean_inc(x_210); -x_291 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_291, 0, x_210); -lean_ctor_set(x_291, 1, x_290); -x_292 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__7; +x_290 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_290, 0, x_210); +lean_ctor_set(x_290, 1, x_289); +x_291 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__7; lean_inc(x_6); -x_293 = lean_name_mk_string(x_6, x_292); -x_294 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__14; -x_295 = l_Lean_addMacroScope(x_216, x_294, x_213); -x_296 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__12; -x_297 = lean_name_mk_string(x_9, x_296); -x_298 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__13; -x_299 = lean_name_mk_string(x_297, x_298); -x_300 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_300, 0, x_299); -lean_ctor_set(x_300, 1, x_276); -x_301 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_301, 0, x_300); -lean_ctor_set(x_301, 1, x_276); -x_302 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__12; -x_303 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_303, 0, x_210); -lean_ctor_set(x_303, 1, x_302); -lean_ctor_set(x_303, 2, x_295); -lean_ctor_set(x_303, 3, x_301); -lean_inc(x_24); -x_304 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_276, x_24); -x_305 = l_Nat_repr(x_10); -x_306 = l_Lean_Syntax_mkLit(x_233, x_305, x_234); -x_307 = l_Nat_repr(x_208); -x_308 = l_Lean_Syntax_mkLit(x_233, x_307, x_234); -x_309 = lean_array_push(x_240, x_303); -x_310 = lean_array_push(x_251, x_291); -x_311 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__84; -x_312 = lean_array_push(x_311, x_260); -x_313 = lean_array_push(x_312, x_266); -x_314 = lean_array_push(x_313, x_287); -if (lean_obj_tag(x_11) == 0) +x_292 = lean_name_mk_string(x_6, x_291); +x_293 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__12; +x_294 = l_Lean_addMacroScope(x_216, x_293, x_213); +x_295 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__12; +x_296 = lean_name_mk_string(x_10, x_295); +x_297 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__11; +x_298 = lean_name_mk_string(x_296, x_297); +lean_ctor_set(x_38, 1, x_275); +lean_ctor_set(x_38, 0, x_298); +x_299 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_299, 0, x_38); +lean_ctor_set(x_299, 1, x_275); +x_300 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__10; +x_301 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_301, 0, x_210); +lean_ctor_set(x_301, 1, x_300); +lean_ctor_set(x_301, 2, x_294); +lean_ctor_set(x_301, 3, x_299); +lean_inc(x_25); +x_302 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_275, x_25); +x_303 = l_Nat_repr(x_11); +x_304 = l_Lean_Syntax_mkNumLit(x_303, x_233); +x_305 = l_Nat_repr(x_208); +x_306 = l_Lean_Syntax_mkNumLit(x_305, x_233); +x_307 = lean_array_push(x_238, x_301); +x_308 = lean_array_push(x_249, x_290); +x_309 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__85; +x_310 = lean_array_push(x_309, x_260); +x_311 = lean_array_push(x_310, x_265); +x_312 = lean_array_push(x_311, x_286); +if (lean_obj_tag(x_12) == 0) { -lean_object* x_379; -x_379 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__11; -x_315 = x_379; -goto block_378; +x_313 = x_256; +goto block_374; } else { -lean_object* x_380; lean_object* x_381; -x_380 = lean_ctor_get(x_11, 0); -lean_inc(x_380); -lean_dec(x_11); -x_381 = lean_array_push(x_236, x_380); -x_315 = x_381; -goto block_378; -} -block_378: -{ -lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; -x_316 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__11; -x_317 = l_Array_append___rarg(x_316, x_315); -x_318 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_318, 0, x_234); -lean_ctor_set(x_318, 1, x_238); -lean_ctor_set(x_318, 2, x_317); -x_319 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__32; -x_320 = lean_array_push(x_319, x_318); +lean_object* x_375; lean_object* x_376; +x_375 = lean_ctor_get(x_12, 0); +lean_inc(x_375); +lean_dec(x_12); +x_376 = lean_array_push(x_235, x_375); +x_313 = x_376; +goto block_374; +} +block_374: +{ +lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; +x_314 = l_Array_append___rarg(x_256, x_313); +lean_inc(x_8); +x_315 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_315, 0, x_233); +lean_ctor_set(x_315, 1, x_8); +lean_ctor_set(x_315, 2, x_314); +x_316 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31; +x_317 = lean_array_push(x_316, x_315); +x_318 = lean_array_push(x_317, x_255); +lean_inc(x_257); +x_319 = lean_array_push(x_318, x_257); +lean_inc(x_257); +x_320 = lean_array_push(x_319, x_257); +lean_inc(x_257); x_321 = lean_array_push(x_320, x_257); -x_322 = lean_array_push(x_321, x_264); -x_323 = lean_array_push(x_322, x_264); -x_324 = lean_array_push(x_323, x_264); -x_325 = lean_array_push(x_324, x_264); -x_326 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_326, 0, x_234); -lean_ctor_set(x_326, 1, x_221); -lean_ctor_set(x_326, 2, x_325); -x_327 = lean_array_push(x_240, x_326); -if (lean_obj_tag(x_304) == 0) -{ -lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; +lean_inc(x_257); +x_322 = lean_array_push(x_321, x_257); +x_323 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_323, 0, x_233); +lean_ctor_set(x_323, 1, x_221); +lean_ctor_set(x_323, 2, x_322); +x_324 = lean_array_push(x_238, x_323); +if (lean_obj_tag(x_302) == 0) +{ +lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_dec(x_6); -x_328 = l___private_Init_Meta_0__Lean_quoteNameMk(x_24); -x_329 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__8; -x_330 = lean_array_push(x_329, x_328); -x_331 = lean_array_push(x_330, x_306); -x_332 = lean_array_push(x_331, x_308); -x_333 = lean_array_push(x_332, x_40); -x_334 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_334, 0, x_234); -lean_ctor_set(x_334, 1, x_238); -lean_ctor_set(x_334, 2, x_333); -x_335 = lean_array_push(x_309, x_334); +x_325 = l_Lean_quoteNameMk(x_25); +x_326 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__8; +x_327 = lean_array_push(x_326, x_325); +x_328 = lean_array_push(x_327, x_304); +x_329 = lean_array_push(x_328, x_306); +x_330 = lean_array_push(x_329, x_45); +x_331 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_331, 0, x_233); +lean_ctor_set(x_331, 1, x_8); +lean_ctor_set(x_331, 2, x_330); +x_332 = lean_array_push(x_307, x_331); +x_333 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_333, 0, x_233); +lean_ctor_set(x_333, 1, x_292); +lean_ctor_set(x_333, 2, x_332); +x_334 = lean_array_push(x_308, x_333); +lean_inc(x_257); +x_335 = lean_array_push(x_334, x_257); x_336 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_336, 0, x_234); -lean_ctor_set(x_336, 1, x_293); +lean_ctor_set(x_336, 0, x_233); +lean_ctor_set(x_336, 1, x_288); lean_ctor_set(x_336, 2, x_335); -x_337 = lean_array_push(x_310, x_336); -x_338 = lean_array_push(x_337, x_264); -x_339 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_339, 0, x_234); -lean_ctor_set(x_339, 1, x_289); -lean_ctor_set(x_339, 2, x_338); -x_340 = lean_array_push(x_314, x_339); -x_341 = lean_array_push(x_340, x_264); -x_342 = lean_array_push(x_341, x_264); -x_343 = lean_array_push(x_342, x_264); -x_344 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_344, 0, x_234); -lean_ctor_set(x_344, 1, x_259); -lean_ctor_set(x_344, 2, x_343); -x_345 = lean_array_push(x_327, x_344); -x_346 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_346, 0, x_234); -lean_ctor_set(x_346, 1, x_219); -lean_ctor_set(x_346, 2, x_345); -x_347 = l_Lean_Elab_Command_elabSyntax___lambda__4(x_3, x_346, x_13, x_14, x_217); -return x_347; -} -else -{ -lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; -lean_dec(x_24); -x_348 = lean_ctor_get(x_304, 0); -lean_inc(x_348); -lean_dec(x_304); -x_349 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__22; -x_350 = lean_name_mk_string(x_6, x_349); -x_351 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__8; -x_352 = l_String_intercalate(x_351, x_348); -x_353 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__9; -x_354 = lean_string_append(x_353, x_352); -lean_dec(x_352); -x_355 = l_Lean_nameLitKind; -x_356 = l_Lean_Syntax_mkLit(x_355, x_354, x_234); -x_357 = lean_array_push(x_236, x_356); -x_358 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_358, 0, x_234); -lean_ctor_set(x_358, 1, x_350); -lean_ctor_set(x_358, 2, x_357); -x_359 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__8; -x_360 = lean_array_push(x_359, x_358); -x_361 = lean_array_push(x_360, x_306); -x_362 = lean_array_push(x_361, x_308); -x_363 = lean_array_push(x_362, x_40); -x_364 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_364, 0, x_234); -lean_ctor_set(x_364, 1, x_238); -lean_ctor_set(x_364, 2, x_363); -x_365 = lean_array_push(x_309, x_364); -x_366 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_366, 0, x_234); -lean_ctor_set(x_366, 1, x_293); -lean_ctor_set(x_366, 2, x_365); -x_367 = lean_array_push(x_310, x_366); -x_368 = lean_array_push(x_367, x_264); -x_369 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_369, 0, x_234); -lean_ctor_set(x_369, 1, x_289); -lean_ctor_set(x_369, 2, x_368); -x_370 = lean_array_push(x_314, x_369); -x_371 = lean_array_push(x_370, x_264); -x_372 = lean_array_push(x_371, x_264); -x_373 = lean_array_push(x_372, x_264); -x_374 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_374, 0, x_234); -lean_ctor_set(x_374, 1, x_259); -lean_ctor_set(x_374, 2, x_373); -x_375 = lean_array_push(x_327, x_374); -x_376 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_376, 0, x_234); -lean_ctor_set(x_376, 1, x_219); -lean_ctor_set(x_376, 2, x_375); -x_377 = l_Lean_Elab_Command_elabSyntax___lambda__4(x_3, x_376, x_13, x_14, x_217); -return x_377; -} -} -} -} -else -{ -lean_object* x_382; lean_object* x_383; lean_object* x_384; -x_382 = lean_ctor_get(x_37, 0); -x_383 = lean_ctor_get(x_37, 1); -lean_inc(x_383); -lean_inc(x_382); -lean_dec(x_37); +x_337 = lean_array_push(x_312, x_336); +lean_inc(x_257); +x_338 = lean_array_push(x_337, x_257); +lean_inc(x_257); +x_339 = lean_array_push(x_338, x_257); +x_340 = lean_array_push(x_339, x_257); +x_341 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_341, 0, x_233); +lean_ctor_set(x_341, 1, x_259); +lean_ctor_set(x_341, 2, x_340); +x_342 = lean_array_push(x_324, x_341); +x_343 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_343, 0, x_233); +lean_ctor_set(x_343, 1, x_219); +lean_ctor_set(x_343, 2, x_342); +x_344 = l_Lean_Elab_Command_elabSyntax___lambda__4(x_3, x_343, x_14, x_15, x_217); +return x_344; +} +else +{ +lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; +lean_dec(x_25); +x_345 = lean_ctor_get(x_302, 0); +lean_inc(x_345); +lean_dec(x_302); +x_346 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__22; +x_347 = lean_name_mk_string(x_6, x_346); +x_348 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__16; +x_349 = l_String_intercalate(x_348, x_345); +x_350 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__17; +x_351 = lean_string_append(x_350, x_349); +lean_dec(x_349); +x_352 = l_Lean_Syntax_mkNameLit(x_351, x_233); +x_353 = lean_array_push(x_235, x_352); +x_354 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_354, 0, x_233); +lean_ctor_set(x_354, 1, x_347); +lean_ctor_set(x_354, 2, x_353); +x_355 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__8; +x_356 = lean_array_push(x_355, x_354); +x_357 = lean_array_push(x_356, x_304); +x_358 = lean_array_push(x_357, x_306); +x_359 = lean_array_push(x_358, x_45); +x_360 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_360, 0, x_233); +lean_ctor_set(x_360, 1, x_8); +lean_ctor_set(x_360, 2, x_359); +x_361 = lean_array_push(x_307, x_360); +x_362 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_362, 0, x_233); +lean_ctor_set(x_362, 1, x_292); +lean_ctor_set(x_362, 2, x_361); +x_363 = lean_array_push(x_308, x_362); +lean_inc(x_257); +x_364 = lean_array_push(x_363, x_257); +x_365 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_365, 0, x_233); +lean_ctor_set(x_365, 1, x_288); +lean_ctor_set(x_365, 2, x_364); +x_366 = lean_array_push(x_312, x_365); +lean_inc(x_257); +x_367 = lean_array_push(x_366, x_257); +lean_inc(x_257); +x_368 = lean_array_push(x_367, x_257); +x_369 = lean_array_push(x_368, x_257); +x_370 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_370, 0, x_233); +lean_ctor_set(x_370, 1, x_259); +lean_ctor_set(x_370, 2, x_369); +x_371 = lean_array_push(x_324, x_370); +x_372 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_372, 0, x_233); +lean_ctor_set(x_372, 1, x_219); +lean_ctor_set(x_372, 2, x_371); +x_373 = l_Lean_Elab_Command_elabSyntax___lambda__4(x_3, x_372, x_14, x_15, x_217); +return x_373; +} +} +} +} +else +{ +lean_object* x_377; lean_object* x_378; +x_377 = lean_ctor_get(x_39, 0); +lean_inc(x_377); +lean_dec(x_39); lean_inc(x_3); -x_384 = l_Lean_mkIdentFrom(x_3, x_12); -if (lean_obj_tag(x_383) == 0) +x_378 = l_Lean_mkIdentFrom(x_3, x_13); +if (lean_obj_tag(x_42) == 0) { -lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; -x_385 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_13, x_14, x_38); +lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; +x_379 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_14, x_15, x_40); +x_380 = lean_ctor_get(x_379, 0); +lean_inc(x_380); +x_381 = lean_ctor_get(x_379, 1); +lean_inc(x_381); +lean_dec(x_379); +x_382 = l_Lean_Elab_Command_getCurrMacroScope(x_14, x_15, x_381); +x_383 = lean_ctor_get(x_382, 0); +lean_inc(x_383); +x_384 = lean_ctor_get(x_382, 1); +lean_inc(x_384); +lean_dec(x_382); +x_385 = l_Lean_Elab_Command_getMainModule___rarg(x_15, x_384); x_386 = lean_ctor_get(x_385, 0); lean_inc(x_386); x_387 = lean_ctor_get(x_385, 1); lean_inc(x_387); lean_dec(x_385); -x_388 = l_Lean_Elab_Command_getCurrMacroScope(x_13, x_14, x_387); -x_389 = lean_ctor_get(x_388, 0); -lean_inc(x_389); -x_390 = lean_ctor_get(x_388, 1); -lean_inc(x_390); -lean_dec(x_388); -x_391 = l_Lean_Elab_Command_getMainModule___rarg(x_14, x_390); -x_392 = lean_ctor_get(x_391, 0); -lean_inc(x_392); -x_393 = lean_ctor_get(x_391, 1); -lean_inc(x_393); -lean_dec(x_391); -x_394 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; +x_388 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; lean_inc(x_5); -x_395 = lean_name_mk_string(x_5, x_394); -x_396 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; +x_389 = lean_name_mk_string(x_5, x_388); +x_390 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; lean_inc(x_5); -x_397 = lean_name_mk_string(x_5, x_396); -x_398 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__13; +x_391 = lean_name_mk_string(x_5, x_390); +x_392 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__12; lean_inc(x_6); -x_399 = lean_name_mk_string(x_6, x_398); -x_400 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__15; -lean_inc(x_386); -x_401 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_401, 0, x_386); -lean_ctor_set(x_401, 1, x_400); -x_402 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__16; +x_393 = lean_name_mk_string(x_6, x_392); +x_394 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__14; +lean_inc(x_380); +x_395 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_395, 0, x_380); +lean_ctor_set(x_395, 1, x_394); +x_396 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__15; lean_inc(x_6); -x_403 = lean_name_mk_string(x_6, x_402); -x_404 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__22; -x_405 = lean_name_mk_string(x_7, x_404); -x_406 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__24; -x_407 = lean_name_mk_string(x_405, x_406); -x_408 = l_Nat_repr(x_18); -x_409 = l_Lean_numLitKind; -x_410 = lean_box(2); -x_411 = l_Lean_Syntax_mkLit(x_409, x_408, x_410); -x_412 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; +x_397 = lean_name_mk_string(x_6, x_396); +x_398 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__21; +x_399 = lean_name_mk_string(x_7, x_398); +x_400 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__23; +x_401 = lean_name_mk_string(x_399, x_400); +x_402 = l_Nat_repr(x_19); +x_403 = lean_box(2); +x_404 = l_Lean_Syntax_mkNumLit(x_402, x_403); +x_405 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; +x_406 = lean_array_push(x_405, x_404); +lean_inc(x_8); +x_407 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_407, 0, x_403); +lean_ctor_set(x_407, 1, x_8); +lean_ctor_set(x_407, 2, x_406); +x_408 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; +x_409 = lean_array_push(x_408, x_28); +x_410 = lean_array_push(x_409, x_407); +x_411 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_411, 0, x_403); +lean_ctor_set(x_411, 1, x_401); +lean_ctor_set(x_411, 2, x_410); +x_412 = lean_array_push(x_408, x_9); x_413 = lean_array_push(x_412, x_411); -x_414 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_415 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_415, 0, x_410); -lean_ctor_set(x_415, 1, x_414); -lean_ctor_set(x_415, 2, x_413); -x_416 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; -x_417 = lean_array_push(x_416, x_27); -x_418 = lean_array_push(x_417, x_415); -x_419 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_419, 0, x_410); -lean_ctor_set(x_419, 1, x_407); -lean_ctor_set(x_419, 2, x_418); -x_420 = lean_array_push(x_416, x_8); -x_421 = lean_array_push(x_420, x_419); -x_422 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_422, 0, x_410); -lean_ctor_set(x_422, 1, x_403); -lean_ctor_set(x_422, 2, x_421); -x_423 = lean_array_push(x_412, x_422); -x_424 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_424, 0, x_410); -lean_ctor_set(x_424, 1, x_414); -lean_ctor_set(x_424, 2, x_423); -x_425 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31; -lean_inc(x_386); -x_426 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_426, 0, x_386); -lean_ctor_set(x_426, 1, x_425); -x_427 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__28; -x_428 = lean_array_push(x_427, x_401); -x_429 = lean_array_push(x_428, x_424); -x_430 = lean_array_push(x_429, x_426); -x_431 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_431, 0, x_410); -lean_ctor_set(x_431, 1, x_399); -lean_ctor_set(x_431, 2, x_430); -x_432 = lean_array_push(x_412, x_431); -x_433 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_433, 0, x_410); -lean_ctor_set(x_433, 1, x_414); -lean_ctor_set(x_433, 2, x_432); -x_434 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__34; +x_414 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_414, 0, x_403); +lean_ctor_set(x_414, 1, x_397); +lean_ctor_set(x_414, 2, x_413); +x_415 = lean_array_push(x_405, x_414); +lean_inc(x_8); +x_416 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_416, 0, x_403); +lean_ctor_set(x_416, 1, x_8); +lean_ctor_set(x_416, 2, x_415); +x_417 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__30; +lean_inc(x_380); +x_418 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_418, 0, x_380); +lean_ctor_set(x_418, 1, x_417); +x_419 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__28; +x_420 = lean_array_push(x_419, x_395); +x_421 = lean_array_push(x_420, x_416); +x_422 = lean_array_push(x_421, x_418); +x_423 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_423, 0, x_403); +lean_ctor_set(x_423, 1, x_393); +lean_ctor_set(x_423, 2, x_422); +x_424 = lean_array_push(x_405, x_423); +lean_inc(x_8); +x_425 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_425, 0, x_403); +lean_ctor_set(x_425, 1, x_8); +lean_ctor_set(x_425, 2, x_424); +x_426 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__3; +lean_inc(x_8); +x_427 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_427, 0, x_403); +lean_ctor_set(x_427, 1, x_8); +lean_ctor_set(x_427, 2, x_426); +x_428 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33; lean_inc(x_5); -x_435 = lean_name_mk_string(x_5, x_434); -lean_inc(x_386); -x_436 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_436, 0, x_386); -lean_ctor_set(x_436, 1, x_434); -x_437 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__1; +x_429 = lean_name_mk_string(x_5, x_428); +lean_inc(x_380); +x_430 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_430, 0, x_380); +lean_ctor_set(x_430, 1, x_428); +x_431 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__35; lean_inc(x_5); -x_438 = lean_name_mk_string(x_5, x_437); -x_439 = lean_array_push(x_416, x_384); -x_440 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__12; -x_441 = lean_array_push(x_439, x_440); -x_442 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_442, 0, x_410); -lean_ctor_set(x_442, 1, x_438); -lean_ctor_set(x_442, 2, x_441); -x_443 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__36; +x_432 = lean_name_mk_string(x_5, x_431); +x_433 = lean_array_push(x_408, x_378); +lean_inc(x_427); +x_434 = lean_array_push(x_433, x_427); +x_435 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_435, 0, x_403); +lean_ctor_set(x_435, 1, x_432); +lean_ctor_set(x_435, 2, x_434); +x_436 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__37; lean_inc(x_5); -x_444 = lean_name_mk_string(x_5, x_443); -x_445 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__38; +x_437 = lean_name_mk_string(x_5, x_436); +x_438 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__39; lean_inc(x_6); -x_446 = lean_name_mk_string(x_6, x_445); -x_447 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__40; -lean_inc(x_386); -x_448 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_448, 0, x_386); -lean_ctor_set(x_448, 1, x_447); -x_449 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__12; -x_450 = lean_name_mk_string(x_9, x_449); -lean_inc(x_389); -lean_inc(x_450); -lean_inc(x_392); -x_451 = l_Lean_addMacroScope(x_392, x_450, x_389); -x_452 = lean_box(0); -lean_inc(x_450); -x_453 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_453, 0, x_450); -lean_ctor_set(x_453, 1, x_452); -x_454 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_454, 0, x_453); -lean_ctor_set(x_454, 1, x_452); -x_455 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__43; -lean_inc(x_386); -x_456 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_456, 0, x_386); -lean_ctor_set(x_456, 1, x_455); -lean_ctor_set(x_456, 2, x_451); -lean_ctor_set(x_456, 3, x_454); -x_457 = lean_array_push(x_416, x_448); -x_458 = lean_array_push(x_457, x_456); -x_459 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_459, 0, x_410); -lean_ctor_set(x_459, 1, x_446); -lean_ctor_set(x_459, 2, x_458); -x_460 = lean_array_push(x_412, x_459); -x_461 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_461, 0, x_410); -lean_ctor_set(x_461, 1, x_414); -lean_ctor_set(x_461, 2, x_460); -x_462 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__2; -x_463 = lean_array_push(x_462, x_461); -x_464 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_464, 0, x_410); -lean_ctor_set(x_464, 1, x_444); -lean_ctor_set(x_464, 2, x_463); -x_465 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__47; -x_466 = lean_name_mk_string(x_5, x_465); -x_467 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__49; +x_439 = lean_name_mk_string(x_6, x_438); +x_440 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__41; +lean_inc(x_380); +x_441 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_441, 0, x_380); +lean_ctor_set(x_441, 1, x_440); +x_442 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__12; +x_443 = lean_name_mk_string(x_10, x_442); +lean_inc(x_383); +lean_inc(x_443); lean_inc(x_386); -x_468 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_468, 0, x_386); -lean_ctor_set(x_468, 1, x_467); -x_469 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__7; +x_444 = l_Lean_addMacroScope(x_386, x_443, x_383); +x_445 = lean_box(0); +lean_inc(x_443); +x_446 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_446, 0, x_443); +lean_ctor_set(x_446, 1, x_445); +x_447 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_447, 0, x_446); +lean_ctor_set(x_447, 1, x_445); +x_448 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__44; +lean_inc(x_380); +x_449 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_449, 0, x_380); +lean_ctor_set(x_449, 1, x_448); +lean_ctor_set(x_449, 2, x_444); +lean_ctor_set(x_449, 3, x_447); +x_450 = lean_array_push(x_408, x_441); +x_451 = lean_array_push(x_450, x_449); +x_452 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_452, 0, x_403); +lean_ctor_set(x_452, 1, x_439); +lean_ctor_set(x_452, 2, x_451); +x_453 = lean_array_push(x_405, x_452); +lean_inc(x_8); +x_454 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_454, 0, x_403); +lean_ctor_set(x_454, 1, x_8); +lean_ctor_set(x_454, 2, x_453); +lean_inc(x_427); +x_455 = lean_array_push(x_408, x_427); +x_456 = lean_array_push(x_455, x_454); +x_457 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_457, 0, x_403); +lean_ctor_set(x_457, 1, x_437); +lean_ctor_set(x_457, 2, x_456); +x_458 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__48; +x_459 = lean_name_mk_string(x_5, x_458); +x_460 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__50; +lean_inc(x_380); +x_461 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_461, 0, x_380); +lean_ctor_set(x_461, 1, x_460); +x_462 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__7; lean_inc(x_6); -x_470 = lean_name_mk_string(x_6, x_469); -x_471 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__6; -x_472 = l_Lean_addMacroScope(x_392, x_471, x_389); -x_473 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__53; -x_474 = lean_name_mk_string(x_450, x_473); -x_475 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_475, 0, x_474); -lean_ctor_set(x_475, 1, x_452); -x_476 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_476, 0, x_475); -lean_ctor_set(x_476, 1, x_452); -x_477 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__5; -x_478 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_478, 0, x_386); -lean_ctor_set(x_478, 1, x_477); -lean_ctor_set(x_478, 2, x_472); -lean_ctor_set(x_478, 3, x_476); -lean_inc(x_24); -x_479 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_452, x_24); -x_480 = l_Nat_repr(x_10); -x_481 = l_Lean_Syntax_mkLit(x_409, x_480, x_410); -x_482 = lean_array_push(x_416, x_478); -x_483 = lean_array_push(x_427, x_468); -x_484 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__84; -x_485 = lean_array_push(x_484, x_436); -x_486 = lean_array_push(x_485, x_442); -x_487 = lean_array_push(x_486, x_464); -if (lean_obj_tag(x_11) == 0) +x_463 = lean_name_mk_string(x_6, x_462); +x_464 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__4; +x_465 = l_Lean_addMacroScope(x_386, x_464, x_383); +x_466 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__54; +x_467 = lean_name_mk_string(x_443, x_466); +lean_ctor_set(x_38, 1, x_445); +lean_ctor_set(x_38, 0, x_467); +x_468 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_468, 0, x_38); +lean_ctor_set(x_468, 1, x_445); +x_469 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__3; +x_470 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_470, 0, x_380); +lean_ctor_set(x_470, 1, x_469); +lean_ctor_set(x_470, 2, x_465); +lean_ctor_set(x_470, 3, x_468); +lean_inc(x_25); +x_471 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_445, x_25); +x_472 = l_Nat_repr(x_11); +x_473 = l_Lean_Syntax_mkNumLit(x_472, x_403); +x_474 = lean_array_push(x_408, x_470); +x_475 = lean_array_push(x_419, x_461); +x_476 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__85; +x_477 = lean_array_push(x_476, x_430); +x_478 = lean_array_push(x_477, x_435); +x_479 = lean_array_push(x_478, x_457); +if (lean_obj_tag(x_12) == 0) { -lean_object* x_548; -x_548 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__11; -x_488 = x_548; -goto block_547; +x_480 = x_426; +goto block_537; } else { -lean_object* x_549; lean_object* x_550; -x_549 = lean_ctor_get(x_11, 0); -lean_inc(x_549); -lean_dec(x_11); -x_550 = lean_array_push(x_412, x_549); -x_488 = x_550; -goto block_547; -} -block_547: -{ -lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; -x_489 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__11; -x_490 = l_Array_append___rarg(x_489, x_488); -x_491 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_491, 0, x_410); -lean_ctor_set(x_491, 1, x_414); -lean_ctor_set(x_491, 2, x_490); -x_492 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__32; -x_493 = lean_array_push(x_492, x_491); -x_494 = lean_array_push(x_493, x_433); -x_495 = lean_array_push(x_494, x_440); -x_496 = lean_array_push(x_495, x_440); -x_497 = lean_array_push(x_496, x_440); -x_498 = lean_array_push(x_497, x_440); -x_499 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_499, 0, x_410); -lean_ctor_set(x_499, 1, x_397); -lean_ctor_set(x_499, 2, x_498); -x_500 = lean_array_push(x_416, x_499); -if (lean_obj_tag(x_479) == 0) -{ -lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; +lean_object* x_538; lean_object* x_539; +x_538 = lean_ctor_get(x_12, 0); +lean_inc(x_538); +lean_dec(x_12); +x_539 = lean_array_push(x_405, x_538); +x_480 = x_539; +goto block_537; +} +block_537: +{ +lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; +x_481 = l_Array_append___rarg(x_426, x_480); +lean_inc(x_8); +x_482 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_482, 0, x_403); +lean_ctor_set(x_482, 1, x_8); +lean_ctor_set(x_482, 2, x_481); +x_483 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31; +x_484 = lean_array_push(x_483, x_482); +x_485 = lean_array_push(x_484, x_425); +lean_inc(x_427); +x_486 = lean_array_push(x_485, x_427); +lean_inc(x_427); +x_487 = lean_array_push(x_486, x_427); +lean_inc(x_427); +x_488 = lean_array_push(x_487, x_427); +lean_inc(x_427); +x_489 = lean_array_push(x_488, x_427); +x_490 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_490, 0, x_403); +lean_ctor_set(x_490, 1, x_391); +lean_ctor_set(x_490, 2, x_489); +x_491 = lean_array_push(x_408, x_490); +if (lean_obj_tag(x_471) == 0) +{ +lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_dec(x_6); -x_501 = l___private_Init_Meta_0__Lean_quoteNameMk(x_24); -x_502 = lean_array_push(x_427, x_501); -x_503 = lean_array_push(x_502, x_481); -x_504 = lean_array_push(x_503, x_382); -x_505 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_505, 0, x_410); -lean_ctor_set(x_505, 1, x_414); -lean_ctor_set(x_505, 2, x_504); -x_506 = lean_array_push(x_482, x_505); -x_507 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_507, 0, x_410); -lean_ctor_set(x_507, 1, x_470); -lean_ctor_set(x_507, 2, x_506); -x_508 = lean_array_push(x_483, x_507); -x_509 = lean_array_push(x_508, x_440); -x_510 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_510, 0, x_410); -lean_ctor_set(x_510, 1, x_466); -lean_ctor_set(x_510, 2, x_509); -x_511 = lean_array_push(x_487, x_510); -x_512 = lean_array_push(x_511, x_440); -x_513 = lean_array_push(x_512, x_440); -x_514 = lean_array_push(x_513, x_440); -x_515 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_515, 0, x_410); -lean_ctor_set(x_515, 1, x_435); -lean_ctor_set(x_515, 2, x_514); -x_516 = lean_array_push(x_500, x_515); -x_517 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_517, 0, x_410); -lean_ctor_set(x_517, 1, x_395); -lean_ctor_set(x_517, 2, x_516); -x_518 = l_Lean_Elab_Command_elabSyntax___lambda__4(x_3, x_517, x_13, x_14, x_393); -return x_518; -} -else -{ -lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; -lean_dec(x_24); -x_519 = lean_ctor_get(x_479, 0); -lean_inc(x_519); -lean_dec(x_479); -x_520 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__22; -x_521 = lean_name_mk_string(x_6, x_520); -x_522 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__8; -x_523 = l_String_intercalate(x_522, x_519); -x_524 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__9; -x_525 = lean_string_append(x_524, x_523); -lean_dec(x_523); -x_526 = l_Lean_nameLitKind; -x_527 = l_Lean_Syntax_mkLit(x_526, x_525, x_410); -x_528 = lean_array_push(x_412, x_527); -x_529 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_529, 0, x_410); -lean_ctor_set(x_529, 1, x_521); -lean_ctor_set(x_529, 2, x_528); -x_530 = lean_array_push(x_427, x_529); -x_531 = lean_array_push(x_530, x_481); -x_532 = lean_array_push(x_531, x_382); +x_492 = l_Lean_quoteNameMk(x_25); +x_493 = lean_array_push(x_419, x_492); +x_494 = lean_array_push(x_493, x_473); +x_495 = lean_array_push(x_494, x_377); +x_496 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_496, 0, x_403); +lean_ctor_set(x_496, 1, x_8); +lean_ctor_set(x_496, 2, x_495); +x_497 = lean_array_push(x_474, x_496); +x_498 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_498, 0, x_403); +lean_ctor_set(x_498, 1, x_463); +lean_ctor_set(x_498, 2, x_497); +x_499 = lean_array_push(x_475, x_498); +lean_inc(x_427); +x_500 = lean_array_push(x_499, x_427); +x_501 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_501, 0, x_403); +lean_ctor_set(x_501, 1, x_459); +lean_ctor_set(x_501, 2, x_500); +x_502 = lean_array_push(x_479, x_501); +lean_inc(x_427); +x_503 = lean_array_push(x_502, x_427); +lean_inc(x_427); +x_504 = lean_array_push(x_503, x_427); +x_505 = lean_array_push(x_504, x_427); +x_506 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_506, 0, x_403); +lean_ctor_set(x_506, 1, x_429); +lean_ctor_set(x_506, 2, x_505); +x_507 = lean_array_push(x_491, x_506); +x_508 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_508, 0, x_403); +lean_ctor_set(x_508, 1, x_389); +lean_ctor_set(x_508, 2, x_507); +x_509 = l_Lean_Elab_Command_elabSyntax___lambda__4(x_3, x_508, x_14, x_15, x_387); +return x_509; +} +else +{ +lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; +lean_dec(x_25); +x_510 = lean_ctor_get(x_471, 0); +lean_inc(x_510); +lean_dec(x_471); +x_511 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__22; +x_512 = lean_name_mk_string(x_6, x_511); +x_513 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__16; +x_514 = l_String_intercalate(x_513, x_510); +x_515 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__17; +x_516 = lean_string_append(x_515, x_514); +lean_dec(x_514); +x_517 = l_Lean_Syntax_mkNameLit(x_516, x_403); +x_518 = lean_array_push(x_405, x_517); +x_519 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_519, 0, x_403); +lean_ctor_set(x_519, 1, x_512); +lean_ctor_set(x_519, 2, x_518); +x_520 = lean_array_push(x_419, x_519); +x_521 = lean_array_push(x_520, x_473); +x_522 = lean_array_push(x_521, x_377); +x_523 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_523, 0, x_403); +lean_ctor_set(x_523, 1, x_8); +lean_ctor_set(x_523, 2, x_522); +x_524 = lean_array_push(x_474, x_523); +x_525 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_525, 0, x_403); +lean_ctor_set(x_525, 1, x_463); +lean_ctor_set(x_525, 2, x_524); +x_526 = lean_array_push(x_475, x_525); +lean_inc(x_427); +x_527 = lean_array_push(x_526, x_427); +x_528 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_528, 0, x_403); +lean_ctor_set(x_528, 1, x_459); +lean_ctor_set(x_528, 2, x_527); +x_529 = lean_array_push(x_479, x_528); +lean_inc(x_427); +x_530 = lean_array_push(x_529, x_427); +lean_inc(x_427); +x_531 = lean_array_push(x_530, x_427); +x_532 = lean_array_push(x_531, x_427); x_533 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_533, 0, x_410); -lean_ctor_set(x_533, 1, x_414); +lean_ctor_set(x_533, 0, x_403); +lean_ctor_set(x_533, 1, x_429); lean_ctor_set(x_533, 2, x_532); -x_534 = lean_array_push(x_482, x_533); +x_534 = lean_array_push(x_491, x_533); x_535 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_535, 0, x_410); -lean_ctor_set(x_535, 1, x_470); +lean_ctor_set(x_535, 0, x_403); +lean_ctor_set(x_535, 1, x_389); lean_ctor_set(x_535, 2, x_534); -x_536 = lean_array_push(x_483, x_535); -x_537 = lean_array_push(x_536, x_440); -x_538 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_538, 0, x_410); -lean_ctor_set(x_538, 1, x_466); -lean_ctor_set(x_538, 2, x_537); -x_539 = lean_array_push(x_487, x_538); -x_540 = lean_array_push(x_539, x_440); -x_541 = lean_array_push(x_540, x_440); -x_542 = lean_array_push(x_541, x_440); -x_543 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_543, 0, x_410); -lean_ctor_set(x_543, 1, x_435); -lean_ctor_set(x_543, 2, x_542); -x_544 = lean_array_push(x_500, x_543); -x_545 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_545, 0, x_410); -lean_ctor_set(x_545, 1, x_395); -lean_ctor_set(x_545, 2, x_544); -x_546 = l_Lean_Elab_Command_elabSyntax___lambda__4(x_3, x_545, x_13, x_14, x_393); -return x_546; -} -} -} -else -{ -lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; lean_object* x_565; lean_object* x_566; lean_object* x_567; lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; lean_object* x_658; lean_object* x_659; -x_551 = lean_ctor_get(x_383, 0); -lean_inc(x_551); -lean_dec(x_383); -x_552 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_13, x_14, x_38); -x_553 = lean_ctor_get(x_552, 0); -lean_inc(x_553); -x_554 = lean_ctor_get(x_552, 1); -lean_inc(x_554); -lean_dec(x_552); -x_555 = l_Lean_Elab_Command_getCurrMacroScope(x_13, x_14, x_554); -x_556 = lean_ctor_get(x_555, 0); -lean_inc(x_556); -x_557 = lean_ctor_get(x_555, 1); -lean_inc(x_557); -lean_dec(x_555); -x_558 = l_Lean_Elab_Command_getMainModule___rarg(x_14, x_557); -x_559 = lean_ctor_get(x_558, 0); -lean_inc(x_559); -x_560 = lean_ctor_get(x_558, 1); -lean_inc(x_560); -lean_dec(x_558); -x_561 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; +x_536 = l_Lean_Elab_Command_elabSyntax___lambda__4(x_3, x_535, x_14, x_15, x_387); +return x_536; +} +} +} +else +{ +lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; lean_object* x_565; lean_object* x_566; lean_object* x_567; lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; +x_540 = lean_ctor_get(x_42, 0); +lean_inc(x_540); +lean_dec(x_42); +x_541 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_14, x_15, x_40); +x_542 = lean_ctor_get(x_541, 0); +lean_inc(x_542); +x_543 = lean_ctor_get(x_541, 1); +lean_inc(x_543); +lean_dec(x_541); +x_544 = l_Lean_Elab_Command_getCurrMacroScope(x_14, x_15, x_543); +x_545 = lean_ctor_get(x_544, 0); +lean_inc(x_545); +x_546 = lean_ctor_get(x_544, 1); +lean_inc(x_546); +lean_dec(x_544); +x_547 = l_Lean_Elab_Command_getMainModule___rarg(x_15, x_546); +x_548 = lean_ctor_get(x_547, 0); +lean_inc(x_548); +x_549 = lean_ctor_get(x_547, 1); +lean_inc(x_549); +lean_dec(x_547); +x_550 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; lean_inc(x_5); -x_562 = lean_name_mk_string(x_5, x_561); -x_563 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; +x_551 = lean_name_mk_string(x_5, x_550); +x_552 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; lean_inc(x_5); -x_564 = lean_name_mk_string(x_5, x_563); -x_565 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__13; +x_553 = lean_name_mk_string(x_5, x_552); +x_554 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__12; lean_inc(x_6); -x_566 = lean_name_mk_string(x_6, x_565); -x_567 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__15; -lean_inc(x_553); -x_568 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_568, 0, x_553); -lean_ctor_set(x_568, 1, x_567); -x_569 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__16; +x_555 = lean_name_mk_string(x_6, x_554); +x_556 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__14; +lean_inc(x_542); +x_557 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_557, 0, x_542); +lean_ctor_set(x_557, 1, x_556); +x_558 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__15; lean_inc(x_6); -x_570 = lean_name_mk_string(x_6, x_569); -x_571 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__22; -x_572 = lean_name_mk_string(x_7, x_571); -x_573 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__24; -x_574 = lean_name_mk_string(x_572, x_573); -x_575 = l_Nat_repr(x_18); -x_576 = l_Lean_numLitKind; -x_577 = lean_box(2); -x_578 = l_Lean_Syntax_mkLit(x_576, x_575, x_577); -x_579 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; -x_580 = lean_array_push(x_579, x_578); -x_581 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_582 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_582, 0, x_577); -lean_ctor_set(x_582, 1, x_581); -lean_ctor_set(x_582, 2, x_580); -x_583 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; -x_584 = lean_array_push(x_583, x_27); -x_585 = lean_array_push(x_584, x_582); -x_586 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_586, 0, x_577); -lean_ctor_set(x_586, 1, x_574); -lean_ctor_set(x_586, 2, x_585); -x_587 = lean_array_push(x_583, x_8); -x_588 = lean_array_push(x_587, x_586); +x_559 = lean_name_mk_string(x_6, x_558); +x_560 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__21; +x_561 = lean_name_mk_string(x_7, x_560); +x_562 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__23; +x_563 = lean_name_mk_string(x_561, x_562); +x_564 = l_Nat_repr(x_19); +x_565 = lean_box(2); +x_566 = l_Lean_Syntax_mkNumLit(x_564, x_565); +x_567 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; +x_568 = lean_array_push(x_567, x_566); +lean_inc(x_8); +x_569 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_569, 0, x_565); +lean_ctor_set(x_569, 1, x_8); +lean_ctor_set(x_569, 2, x_568); +x_570 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; +x_571 = lean_array_push(x_570, x_28); +x_572 = lean_array_push(x_571, x_569); +x_573 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_573, 0, x_565); +lean_ctor_set(x_573, 1, x_563); +lean_ctor_set(x_573, 2, x_572); +x_574 = lean_array_push(x_570, x_9); +x_575 = lean_array_push(x_574, x_573); +x_576 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_576, 0, x_565); +lean_ctor_set(x_576, 1, x_559); +lean_ctor_set(x_576, 2, x_575); +x_577 = lean_array_push(x_567, x_576); +lean_inc(x_8); +x_578 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_578, 0, x_565); +lean_ctor_set(x_578, 1, x_8); +lean_ctor_set(x_578, 2, x_577); +x_579 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__30; +lean_inc(x_542); +x_580 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_580, 0, x_542); +lean_ctor_set(x_580, 1, x_579); +x_581 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__28; +x_582 = lean_array_push(x_581, x_557); +x_583 = lean_array_push(x_582, x_578); +x_584 = lean_array_push(x_583, x_580); +x_585 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_585, 0, x_565); +lean_ctor_set(x_585, 1, x_555); +lean_ctor_set(x_585, 2, x_584); +x_586 = lean_array_push(x_567, x_585); +lean_inc(x_8); +x_587 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_587, 0, x_565); +lean_ctor_set(x_587, 1, x_8); +lean_ctor_set(x_587, 2, x_586); +x_588 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__3; +lean_inc(x_8); x_589 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_589, 0, x_577); -lean_ctor_set(x_589, 1, x_570); +lean_ctor_set(x_589, 0, x_565); +lean_ctor_set(x_589, 1, x_8); lean_ctor_set(x_589, 2, x_588); -x_590 = lean_array_push(x_579, x_589); -x_591 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_591, 0, x_577); -lean_ctor_set(x_591, 1, x_581); -lean_ctor_set(x_591, 2, x_590); -x_592 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31; -lean_inc(x_553); -x_593 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_593, 0, x_553); -lean_ctor_set(x_593, 1, x_592); -x_594 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__28; -x_595 = lean_array_push(x_594, x_568); -x_596 = lean_array_push(x_595, x_591); -x_597 = lean_array_push(x_596, x_593); -x_598 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_598, 0, x_577); -lean_ctor_set(x_598, 1, x_566); -lean_ctor_set(x_598, 2, x_597); -x_599 = lean_array_push(x_579, x_598); -x_600 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_600, 0, x_577); -lean_ctor_set(x_600, 1, x_581); -lean_ctor_set(x_600, 2, x_599); -x_601 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__34; +x_590 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33; lean_inc(x_5); -x_602 = lean_name_mk_string(x_5, x_601); -lean_inc(x_553); -x_603 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_603, 0, x_553); -lean_ctor_set(x_603, 1, x_601); -x_604 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__1; +x_591 = lean_name_mk_string(x_5, x_590); +lean_inc(x_542); +x_592 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_592, 0, x_542); +lean_ctor_set(x_592, 1, x_590); +x_593 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__35; lean_inc(x_5); -x_605 = lean_name_mk_string(x_5, x_604); -x_606 = lean_array_push(x_583, x_384); -x_607 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__12; -x_608 = lean_array_push(x_606, x_607); -x_609 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_609, 0, x_577); -lean_ctor_set(x_609, 1, x_605); -lean_ctor_set(x_609, 2, x_608); -x_610 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__36; +x_594 = lean_name_mk_string(x_5, x_593); +x_595 = lean_array_push(x_570, x_378); +lean_inc(x_589); +x_596 = lean_array_push(x_595, x_589); +x_597 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_597, 0, x_565); +lean_ctor_set(x_597, 1, x_594); +lean_ctor_set(x_597, 2, x_596); +x_598 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__37; lean_inc(x_5); -x_611 = lean_name_mk_string(x_5, x_610); -x_612 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__38; +x_599 = lean_name_mk_string(x_5, x_598); +x_600 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__39; lean_inc(x_6); -x_613 = lean_name_mk_string(x_6, x_612); -x_614 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__40; -lean_inc(x_553); -x_615 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_615, 0, x_553); -lean_ctor_set(x_615, 1, x_614); -x_616 = l_List_filterMap___at_Lean_Elab_Term_resolveParserName___spec__1___closed__1; -lean_inc(x_9); -x_617 = lean_name_mk_string(x_9, x_616); -lean_inc(x_556); -lean_inc(x_617); -lean_inc(x_559); -x_618 = l_Lean_addMacroScope(x_559, x_617, x_556); -x_619 = lean_box(0); -x_620 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_620, 0, x_617); -lean_ctor_set(x_620, 1, x_619); -x_621 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_621, 0, x_620); -lean_ctor_set(x_621, 1, x_619); -x_622 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__9; -lean_inc(x_553); -x_623 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_623, 0, x_553); +x_601 = lean_name_mk_string(x_6, x_600); +x_602 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__41; +lean_inc(x_542); +x_603 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_603, 0, x_542); +lean_ctor_set(x_603, 1, x_602); +x_604 = l_List_filterMap___at_Lean_Elab_Term_resolveParserName___spec__1___closed__1; +lean_inc(x_10); +x_605 = lean_name_mk_string(x_10, x_604); +lean_inc(x_545); +lean_inc(x_605); +lean_inc(x_548); +x_606 = l_Lean_addMacroScope(x_548, x_605, x_545); +x_607 = lean_box(0); +x_608 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_608, 0, x_605); +lean_ctor_set(x_608, 1, x_607); +x_609 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_609, 0, x_608); +lean_ctor_set(x_609, 1, x_607); +x_610 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__7; +lean_inc(x_542); +x_611 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_611, 0, x_542); +lean_ctor_set(x_611, 1, x_610); +lean_ctor_set(x_611, 2, x_606); +lean_ctor_set(x_611, 3, x_609); +x_612 = lean_array_push(x_570, x_603); +x_613 = lean_array_push(x_612, x_611); +x_614 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_614, 0, x_565); +lean_ctor_set(x_614, 1, x_601); +lean_ctor_set(x_614, 2, x_613); +x_615 = lean_array_push(x_567, x_614); +lean_inc(x_8); +x_616 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_616, 0, x_565); +lean_ctor_set(x_616, 1, x_8); +lean_ctor_set(x_616, 2, x_615); +lean_inc(x_589); +x_617 = lean_array_push(x_570, x_589); +x_618 = lean_array_push(x_617, x_616); +x_619 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_619, 0, x_565); +lean_ctor_set(x_619, 1, x_599); +lean_ctor_set(x_619, 2, x_618); +x_620 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__48; +x_621 = lean_name_mk_string(x_5, x_620); +x_622 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__50; +lean_inc(x_542); +x_623 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_623, 0, x_542); lean_ctor_set(x_623, 1, x_622); -lean_ctor_set(x_623, 2, x_618); -lean_ctor_set(x_623, 3, x_621); -x_624 = lean_array_push(x_583, x_615); -x_625 = lean_array_push(x_624, x_623); -x_626 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_626, 0, x_577); -lean_ctor_set(x_626, 1, x_613); -lean_ctor_set(x_626, 2, x_625); -x_627 = lean_array_push(x_579, x_626); -x_628 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_628, 0, x_577); -lean_ctor_set(x_628, 1, x_581); -lean_ctor_set(x_628, 2, x_627); -x_629 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__2; -x_630 = lean_array_push(x_629, x_628); -x_631 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_631, 0, x_577); -lean_ctor_set(x_631, 1, x_611); -lean_ctor_set(x_631, 2, x_630); -x_632 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__47; -x_633 = lean_name_mk_string(x_5, x_632); -x_634 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__49; -lean_inc(x_553); -x_635 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_635, 0, x_553); -lean_ctor_set(x_635, 1, x_634); -x_636 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__7; +x_624 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__7; lean_inc(x_6); -x_637 = lean_name_mk_string(x_6, x_636); -x_638 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__14; -x_639 = l_Lean_addMacroScope(x_559, x_638, x_556); -x_640 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__12; -x_641 = lean_name_mk_string(x_9, x_640); -x_642 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__13; -x_643 = lean_name_mk_string(x_641, x_642); -x_644 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_644, 0, x_643); -lean_ctor_set(x_644, 1, x_619); -x_645 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_645, 0, x_644); -lean_ctor_set(x_645, 1, x_619); -x_646 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__12; -x_647 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_647, 0, x_553); -lean_ctor_set(x_647, 1, x_646); -lean_ctor_set(x_647, 2, x_639); -lean_ctor_set(x_647, 3, x_645); -lean_inc(x_24); -x_648 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_619, x_24); -x_649 = l_Nat_repr(x_10); -x_650 = l_Lean_Syntax_mkLit(x_576, x_649, x_577); -x_651 = l_Nat_repr(x_551); -x_652 = l_Lean_Syntax_mkLit(x_576, x_651, x_577); -x_653 = lean_array_push(x_583, x_647); -x_654 = lean_array_push(x_594, x_635); -x_655 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__84; -x_656 = lean_array_push(x_655, x_603); -x_657 = lean_array_push(x_656, x_609); -x_658 = lean_array_push(x_657, x_631); -if (lean_obj_tag(x_11) == 0) +x_625 = lean_name_mk_string(x_6, x_624); +x_626 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__12; +x_627 = l_Lean_addMacroScope(x_548, x_626, x_545); +x_628 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__12; +x_629 = lean_name_mk_string(x_10, x_628); +x_630 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__11; +x_631 = lean_name_mk_string(x_629, x_630); +lean_ctor_set(x_38, 1, x_607); +lean_ctor_set(x_38, 0, x_631); +x_632 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_632, 0, x_38); +lean_ctor_set(x_632, 1, x_607); +x_633 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__10; +x_634 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_634, 0, x_542); +lean_ctor_set(x_634, 1, x_633); +lean_ctor_set(x_634, 2, x_627); +lean_ctor_set(x_634, 3, x_632); +lean_inc(x_25); +x_635 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_607, x_25); +x_636 = l_Nat_repr(x_11); +x_637 = l_Lean_Syntax_mkNumLit(x_636, x_565); +x_638 = l_Nat_repr(x_540); +x_639 = l_Lean_Syntax_mkNumLit(x_638, x_565); +x_640 = lean_array_push(x_570, x_634); +x_641 = lean_array_push(x_581, x_623); +x_642 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__85; +x_643 = lean_array_push(x_642, x_592); +x_644 = lean_array_push(x_643, x_597); +x_645 = lean_array_push(x_644, x_619); +if (lean_obj_tag(x_12) == 0) { -lean_object* x_723; -x_723 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__11; -x_659 = x_723; -goto block_722; +x_646 = x_588; +goto block_707; } else { -lean_object* x_724; lean_object* x_725; -x_724 = lean_ctor_get(x_11, 0); -lean_inc(x_724); -lean_dec(x_11); -x_725 = lean_array_push(x_579, x_724); -x_659 = x_725; -goto block_722; -} -block_722: -{ -lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; -x_660 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__11; -x_661 = l_Array_append___rarg(x_660, x_659); -x_662 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_662, 0, x_577); -lean_ctor_set(x_662, 1, x_581); -lean_ctor_set(x_662, 2, x_661); -x_663 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__32; -x_664 = lean_array_push(x_663, x_662); -x_665 = lean_array_push(x_664, x_600); -x_666 = lean_array_push(x_665, x_607); -x_667 = lean_array_push(x_666, x_607); -x_668 = lean_array_push(x_667, x_607); -x_669 = lean_array_push(x_668, x_607); -x_670 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_670, 0, x_577); -lean_ctor_set(x_670, 1, x_564); -lean_ctor_set(x_670, 2, x_669); -x_671 = lean_array_push(x_583, x_670); -if (lean_obj_tag(x_648) == 0) -{ -lean_object* x_672; lean_object* x_673; lean_object* x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; lean_object* x_678; lean_object* x_679; lean_object* x_680; lean_object* x_681; lean_object* x_682; lean_object* x_683; lean_object* x_684; lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; lean_object* x_690; lean_object* x_691; +lean_object* x_708; lean_object* x_709; +x_708 = lean_ctor_get(x_12, 0); +lean_inc(x_708); +lean_dec(x_12); +x_709 = lean_array_push(x_567, x_708); +x_646 = x_709; +goto block_707; +} +block_707: +{ +lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; +x_647 = l_Array_append___rarg(x_588, x_646); +lean_inc(x_8); +x_648 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_648, 0, x_565); +lean_ctor_set(x_648, 1, x_8); +lean_ctor_set(x_648, 2, x_647); +x_649 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31; +x_650 = lean_array_push(x_649, x_648); +x_651 = lean_array_push(x_650, x_587); +lean_inc(x_589); +x_652 = lean_array_push(x_651, x_589); +lean_inc(x_589); +x_653 = lean_array_push(x_652, x_589); +lean_inc(x_589); +x_654 = lean_array_push(x_653, x_589); +lean_inc(x_589); +x_655 = lean_array_push(x_654, x_589); +x_656 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_656, 0, x_565); +lean_ctor_set(x_656, 1, x_553); +lean_ctor_set(x_656, 2, x_655); +x_657 = lean_array_push(x_570, x_656); +if (lean_obj_tag(x_635) == 0) +{ +lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; lean_object* x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; lean_dec(x_6); -x_672 = l___private_Init_Meta_0__Lean_quoteNameMk(x_24); -x_673 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__8; -x_674 = lean_array_push(x_673, x_672); -x_675 = lean_array_push(x_674, x_650); -x_676 = lean_array_push(x_675, x_652); -x_677 = lean_array_push(x_676, x_382); -x_678 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_678, 0, x_577); -lean_ctor_set(x_678, 1, x_581); -lean_ctor_set(x_678, 2, x_677); -x_679 = lean_array_push(x_653, x_678); -x_680 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_680, 0, x_577); -lean_ctor_set(x_680, 1, x_637); -lean_ctor_set(x_680, 2, x_679); -x_681 = lean_array_push(x_654, x_680); -x_682 = lean_array_push(x_681, x_607); -x_683 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_683, 0, x_577); -lean_ctor_set(x_683, 1, x_633); -lean_ctor_set(x_683, 2, x_682); -x_684 = lean_array_push(x_658, x_683); -x_685 = lean_array_push(x_684, x_607); -x_686 = lean_array_push(x_685, x_607); -x_687 = lean_array_push(x_686, x_607); -x_688 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_688, 0, x_577); -lean_ctor_set(x_688, 1, x_602); -lean_ctor_set(x_688, 2, x_687); -x_689 = lean_array_push(x_671, x_688); -x_690 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_690, 0, x_577); -lean_ctor_set(x_690, 1, x_562); -lean_ctor_set(x_690, 2, x_689); -x_691 = l_Lean_Elab_Command_elabSyntax___lambda__4(x_3, x_690, x_13, x_14, x_560); -return x_691; -} -else -{ -lean_object* x_692; lean_object* x_693; lean_object* x_694; lean_object* x_695; lean_object* x_696; lean_object* x_697; lean_object* x_698; lean_object* x_699; lean_object* x_700; lean_object* x_701; lean_object* x_702; lean_object* x_703; lean_object* x_704; lean_object* x_705; lean_object* x_706; lean_object* x_707; lean_object* x_708; lean_object* x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; lean_object* x_718; lean_object* x_719; lean_object* x_720; lean_object* x_721; -lean_dec(x_24); -x_692 = lean_ctor_get(x_648, 0); -lean_inc(x_692); -lean_dec(x_648); -x_693 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__22; -x_694 = lean_name_mk_string(x_6, x_693); -x_695 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__8; -x_696 = l_String_intercalate(x_695, x_692); -x_697 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__9; -x_698 = lean_string_append(x_697, x_696); -lean_dec(x_696); -x_699 = l_Lean_nameLitKind; -x_700 = l_Lean_Syntax_mkLit(x_699, x_698, x_577); -x_701 = lean_array_push(x_579, x_700); -x_702 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_702, 0, x_577); -lean_ctor_set(x_702, 1, x_694); -lean_ctor_set(x_702, 2, x_701); -x_703 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__8; -x_704 = lean_array_push(x_703, x_702); -x_705 = lean_array_push(x_704, x_650); -x_706 = lean_array_push(x_705, x_652); -x_707 = lean_array_push(x_706, x_382); -x_708 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_708, 0, x_577); -lean_ctor_set(x_708, 1, x_581); -lean_ctor_set(x_708, 2, x_707); -x_709 = lean_array_push(x_653, x_708); -x_710 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_710, 0, x_577); -lean_ctor_set(x_710, 1, x_637); -lean_ctor_set(x_710, 2, x_709); -x_711 = lean_array_push(x_654, x_710); -x_712 = lean_array_push(x_711, x_607); -x_713 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_713, 0, x_577); -lean_ctor_set(x_713, 1, x_633); -lean_ctor_set(x_713, 2, x_712); -x_714 = lean_array_push(x_658, x_713); -x_715 = lean_array_push(x_714, x_607); -x_716 = lean_array_push(x_715, x_607); -x_717 = lean_array_push(x_716, x_607); -x_718 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_718, 0, x_577); -lean_ctor_set(x_718, 1, x_602); -lean_ctor_set(x_718, 2, x_717); -x_719 = lean_array_push(x_671, x_718); -x_720 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_720, 0, x_577); -lean_ctor_set(x_720, 1, x_562); -lean_ctor_set(x_720, 2, x_719); -x_721 = l_Lean_Elab_Command_elabSyntax___lambda__4(x_3, x_720, x_13, x_14, x_560); -return x_721; -} -} -} -} -} -else -{ -uint8_t x_726; -lean_dec(x_27); -lean_dec(x_24); -lean_dec(x_18); +x_658 = l_Lean_quoteNameMk(x_25); +x_659 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__8; +x_660 = lean_array_push(x_659, x_658); +x_661 = lean_array_push(x_660, x_637); +x_662 = lean_array_push(x_661, x_639); +x_663 = lean_array_push(x_662, x_377); +x_664 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_664, 0, x_565); +lean_ctor_set(x_664, 1, x_8); +lean_ctor_set(x_664, 2, x_663); +x_665 = lean_array_push(x_640, x_664); +x_666 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_666, 0, x_565); +lean_ctor_set(x_666, 1, x_625); +lean_ctor_set(x_666, 2, x_665); +x_667 = lean_array_push(x_641, x_666); +lean_inc(x_589); +x_668 = lean_array_push(x_667, x_589); +x_669 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_669, 0, x_565); +lean_ctor_set(x_669, 1, x_621); +lean_ctor_set(x_669, 2, x_668); +x_670 = lean_array_push(x_645, x_669); +lean_inc(x_589); +x_671 = lean_array_push(x_670, x_589); +lean_inc(x_589); +x_672 = lean_array_push(x_671, x_589); +x_673 = lean_array_push(x_672, x_589); +x_674 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_674, 0, x_565); +lean_ctor_set(x_674, 1, x_591); +lean_ctor_set(x_674, 2, x_673); +x_675 = lean_array_push(x_657, x_674); +x_676 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_676, 0, x_565); +lean_ctor_set(x_676, 1, x_551); +lean_ctor_set(x_676, 2, x_675); +x_677 = l_Lean_Elab_Command_elabSyntax___lambda__4(x_3, x_676, x_14, x_15, x_549); +return x_677; +} +else +{ +lean_object* x_678; lean_object* x_679; lean_object* x_680; lean_object* x_681; lean_object* x_682; lean_object* x_683; lean_object* x_684; lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; lean_object* x_690; lean_object* x_691; lean_object* x_692; lean_object* x_693; lean_object* x_694; lean_object* x_695; lean_object* x_696; lean_object* x_697; lean_object* x_698; lean_object* x_699; lean_object* x_700; lean_object* x_701; lean_object* x_702; lean_object* x_703; lean_object* x_704; lean_object* x_705; lean_object* x_706; +lean_dec(x_25); +x_678 = lean_ctor_get(x_635, 0); +lean_inc(x_678); +lean_dec(x_635); +x_679 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__22; +x_680 = lean_name_mk_string(x_6, x_679); +x_681 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__16; +x_682 = l_String_intercalate(x_681, x_678); +x_683 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__17; +x_684 = lean_string_append(x_683, x_682); +lean_dec(x_682); +x_685 = l_Lean_Syntax_mkNameLit(x_684, x_565); +x_686 = lean_array_push(x_567, x_685); +x_687 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_687, 0, x_565); +lean_ctor_set(x_687, 1, x_680); +lean_ctor_set(x_687, 2, x_686); +x_688 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__8; +x_689 = lean_array_push(x_688, x_687); +x_690 = lean_array_push(x_689, x_637); +x_691 = lean_array_push(x_690, x_639); +x_692 = lean_array_push(x_691, x_377); +x_693 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_693, 0, x_565); +lean_ctor_set(x_693, 1, x_8); +lean_ctor_set(x_693, 2, x_692); +x_694 = lean_array_push(x_640, x_693); +x_695 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_695, 0, x_565); +lean_ctor_set(x_695, 1, x_625); +lean_ctor_set(x_695, 2, x_694); +x_696 = lean_array_push(x_641, x_695); +lean_inc(x_589); +x_697 = lean_array_push(x_696, x_589); +x_698 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_698, 0, x_565); +lean_ctor_set(x_698, 1, x_621); +lean_ctor_set(x_698, 2, x_697); +x_699 = lean_array_push(x_645, x_698); +lean_inc(x_589); +x_700 = lean_array_push(x_699, x_589); +lean_inc(x_589); +x_701 = lean_array_push(x_700, x_589); +x_702 = lean_array_push(x_701, x_589); +x_703 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_703, 0, x_565); +lean_ctor_set(x_703, 1, x_591); +lean_ctor_set(x_703, 2, x_702); +x_704 = lean_array_push(x_657, x_703); +x_705 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_705, 0, x_565); +lean_ctor_set(x_705, 1, x_551); +lean_ctor_set(x_705, 2, x_704); +x_706 = l_Lean_Elab_Command_elabSyntax___lambda__4(x_3, x_705, x_14, x_15, x_549); +return x_706; +} +} +} +} +} +else +{ +lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; +x_710 = lean_ctor_get(x_38, 1); +lean_inc(x_710); +lean_dec(x_38); +x_711 = lean_ctor_get(x_39, 0); +lean_inc(x_711); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_712 = x_39; +} else { + lean_dec_ref(x_39); + x_712 = lean_box(0); +} +lean_inc(x_3); +x_713 = l_Lean_mkIdentFrom(x_3, x_13); +if (lean_obj_tag(x_710) == 0) +{ +lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; lean_object* x_718; lean_object* x_719; lean_object* x_720; lean_object* x_721; lean_object* x_722; lean_object* x_723; lean_object* x_724; lean_object* x_725; lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; lean_object* x_730; lean_object* x_731; lean_object* x_732; lean_object* x_733; lean_object* x_734; lean_object* x_735; lean_object* x_736; lean_object* x_737; lean_object* x_738; lean_object* x_739; lean_object* x_740; lean_object* x_741; lean_object* x_742; lean_object* x_743; lean_object* x_744; lean_object* x_745; lean_object* x_746; lean_object* x_747; lean_object* x_748; lean_object* x_749; lean_object* x_750; lean_object* x_751; lean_object* x_752; lean_object* x_753; lean_object* x_754; lean_object* x_755; lean_object* x_756; lean_object* x_757; lean_object* x_758; lean_object* x_759; lean_object* x_760; lean_object* x_761; lean_object* x_762; lean_object* x_763; lean_object* x_764; lean_object* x_765; lean_object* x_766; lean_object* x_767; lean_object* x_768; lean_object* x_769; lean_object* x_770; lean_object* x_771; lean_object* x_772; lean_object* x_773; lean_object* x_774; lean_object* x_775; lean_object* x_776; lean_object* x_777; lean_object* x_778; lean_object* x_779; lean_object* x_780; lean_object* x_781; lean_object* x_782; lean_object* x_783; lean_object* x_784; lean_object* x_785; lean_object* x_786; lean_object* x_787; lean_object* x_788; lean_object* x_789; lean_object* x_790; lean_object* x_791; lean_object* x_792; lean_object* x_793; lean_object* x_794; lean_object* x_795; lean_object* x_796; lean_object* x_797; lean_object* x_798; lean_object* x_799; lean_object* x_800; lean_object* x_801; lean_object* x_802; lean_object* x_803; lean_object* x_804; lean_object* x_805; lean_object* x_806; lean_object* x_807; lean_object* x_808; lean_object* x_809; lean_object* x_810; lean_object* x_811; lean_object* x_812; lean_object* x_813; lean_object* x_814; lean_object* x_815; lean_object* x_816; +x_714 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_14, x_15, x_40); +x_715 = lean_ctor_get(x_714, 0); +lean_inc(x_715); +x_716 = lean_ctor_get(x_714, 1); +lean_inc(x_716); +lean_dec(x_714); +x_717 = l_Lean_Elab_Command_getCurrMacroScope(x_14, x_15, x_716); +x_718 = lean_ctor_get(x_717, 0); +lean_inc(x_718); +x_719 = lean_ctor_get(x_717, 1); +lean_inc(x_719); +lean_dec(x_717); +x_720 = l_Lean_Elab_Command_getMainModule___rarg(x_15, x_719); +x_721 = lean_ctor_get(x_720, 0); +lean_inc(x_721); +x_722 = lean_ctor_get(x_720, 1); +lean_inc(x_722); +lean_dec(x_720); +x_723 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; +lean_inc(x_5); +x_724 = lean_name_mk_string(x_5, x_723); +x_725 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; +lean_inc(x_5); +x_726 = lean_name_mk_string(x_5, x_725); +x_727 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__12; +lean_inc(x_6); +x_728 = lean_name_mk_string(x_6, x_727); +x_729 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__14; +lean_inc(x_715); +x_730 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_730, 0, x_715); +lean_ctor_set(x_730, 1, x_729); +x_731 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__15; +lean_inc(x_6); +x_732 = lean_name_mk_string(x_6, x_731); +x_733 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__21; +x_734 = lean_name_mk_string(x_7, x_733); +x_735 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__23; +x_736 = lean_name_mk_string(x_734, x_735); +x_737 = l_Nat_repr(x_19); +x_738 = lean_box(2); +x_739 = l_Lean_Syntax_mkNumLit(x_737, x_738); +x_740 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; +x_741 = lean_array_push(x_740, x_739); +lean_inc(x_8); +x_742 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_742, 0, x_738); +lean_ctor_set(x_742, 1, x_8); +lean_ctor_set(x_742, 2, x_741); +x_743 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; +x_744 = lean_array_push(x_743, x_28); +x_745 = lean_array_push(x_744, x_742); +x_746 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_746, 0, x_738); +lean_ctor_set(x_746, 1, x_736); +lean_ctor_set(x_746, 2, x_745); +x_747 = lean_array_push(x_743, x_9); +x_748 = lean_array_push(x_747, x_746); +x_749 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_749, 0, x_738); +lean_ctor_set(x_749, 1, x_732); +lean_ctor_set(x_749, 2, x_748); +x_750 = lean_array_push(x_740, x_749); +lean_inc(x_8); +x_751 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_751, 0, x_738); +lean_ctor_set(x_751, 1, x_8); +lean_ctor_set(x_751, 2, x_750); +x_752 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__30; +lean_inc(x_715); +x_753 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_753, 0, x_715); +lean_ctor_set(x_753, 1, x_752); +x_754 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__28; +x_755 = lean_array_push(x_754, x_730); +x_756 = lean_array_push(x_755, x_751); +x_757 = lean_array_push(x_756, x_753); +x_758 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_758, 0, x_738); +lean_ctor_set(x_758, 1, x_728); +lean_ctor_set(x_758, 2, x_757); +x_759 = lean_array_push(x_740, x_758); +lean_inc(x_8); +x_760 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_760, 0, x_738); +lean_ctor_set(x_760, 1, x_8); +lean_ctor_set(x_760, 2, x_759); +x_761 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__3; +lean_inc(x_8); +x_762 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_762, 0, x_738); +lean_ctor_set(x_762, 1, x_8); +lean_ctor_set(x_762, 2, x_761); +x_763 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33; +lean_inc(x_5); +x_764 = lean_name_mk_string(x_5, x_763); +lean_inc(x_715); +x_765 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_765, 0, x_715); +lean_ctor_set(x_765, 1, x_763); +x_766 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__35; +lean_inc(x_5); +x_767 = lean_name_mk_string(x_5, x_766); +x_768 = lean_array_push(x_743, x_713); +lean_inc(x_762); +x_769 = lean_array_push(x_768, x_762); +x_770 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_770, 0, x_738); +lean_ctor_set(x_770, 1, x_767); +lean_ctor_set(x_770, 2, x_769); +x_771 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__37; +lean_inc(x_5); +x_772 = lean_name_mk_string(x_5, x_771); +x_773 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__39; +lean_inc(x_6); +x_774 = lean_name_mk_string(x_6, x_773); +x_775 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__41; +lean_inc(x_715); +x_776 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_776, 0, x_715); +lean_ctor_set(x_776, 1, x_775); +x_777 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__12; +x_778 = lean_name_mk_string(x_10, x_777); +lean_inc(x_718); +lean_inc(x_778); +lean_inc(x_721); +x_779 = l_Lean_addMacroScope(x_721, x_778, x_718); +x_780 = lean_box(0); +lean_inc(x_778); +if (lean_is_scalar(x_712)) { + x_781 = lean_alloc_ctor(0, 2, 0); +} else { + x_781 = x_712; +} +lean_ctor_set(x_781, 0, x_778); +lean_ctor_set(x_781, 1, x_780); +x_782 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_782, 0, x_781); +lean_ctor_set(x_782, 1, x_780); +x_783 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__44; +lean_inc(x_715); +x_784 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_784, 0, x_715); +lean_ctor_set(x_784, 1, x_783); +lean_ctor_set(x_784, 2, x_779); +lean_ctor_set(x_784, 3, x_782); +x_785 = lean_array_push(x_743, x_776); +x_786 = lean_array_push(x_785, x_784); +x_787 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_787, 0, x_738); +lean_ctor_set(x_787, 1, x_774); +lean_ctor_set(x_787, 2, x_786); +x_788 = lean_array_push(x_740, x_787); +lean_inc(x_8); +x_789 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_789, 0, x_738); +lean_ctor_set(x_789, 1, x_8); +lean_ctor_set(x_789, 2, x_788); +lean_inc(x_762); +x_790 = lean_array_push(x_743, x_762); +x_791 = lean_array_push(x_790, x_789); +x_792 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_792, 0, x_738); +lean_ctor_set(x_792, 1, x_772); +lean_ctor_set(x_792, 2, x_791); +x_793 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__48; +x_794 = lean_name_mk_string(x_5, x_793); +x_795 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__50; +lean_inc(x_715); +x_796 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_796, 0, x_715); +lean_ctor_set(x_796, 1, x_795); +x_797 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__7; +lean_inc(x_6); +x_798 = lean_name_mk_string(x_6, x_797); +x_799 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__4; +x_800 = l_Lean_addMacroScope(x_721, x_799, x_718); +x_801 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__54; +x_802 = lean_name_mk_string(x_778, x_801); +x_803 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_803, 0, x_802); +lean_ctor_set(x_803, 1, x_780); +x_804 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_804, 0, x_803); +lean_ctor_set(x_804, 1, x_780); +x_805 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__3; +x_806 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_806, 0, x_715); +lean_ctor_set(x_806, 1, x_805); +lean_ctor_set(x_806, 2, x_800); +lean_ctor_set(x_806, 3, x_804); +lean_inc(x_25); +x_807 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_780, x_25); +x_808 = l_Nat_repr(x_11); +x_809 = l_Lean_Syntax_mkNumLit(x_808, x_738); +x_810 = lean_array_push(x_743, x_806); +x_811 = lean_array_push(x_754, x_796); +x_812 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__85; +x_813 = lean_array_push(x_812, x_765); +x_814 = lean_array_push(x_813, x_770); +x_815 = lean_array_push(x_814, x_792); +if (lean_obj_tag(x_12) == 0) +{ +x_816 = x_761; +goto block_873; +} +else +{ +lean_object* x_874; lean_object* x_875; +x_874 = lean_ctor_get(x_12, 0); +lean_inc(x_874); +lean_dec(x_12); +x_875 = lean_array_push(x_740, x_874); +x_816 = x_875; +goto block_873; +} +block_873: +{ +lean_object* x_817; lean_object* x_818; lean_object* x_819; lean_object* x_820; lean_object* x_821; lean_object* x_822; lean_object* x_823; lean_object* x_824; lean_object* x_825; lean_object* x_826; lean_object* x_827; +x_817 = l_Array_append___rarg(x_761, x_816); +lean_inc(x_8); +x_818 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_818, 0, x_738); +lean_ctor_set(x_818, 1, x_8); +lean_ctor_set(x_818, 2, x_817); +x_819 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31; +x_820 = lean_array_push(x_819, x_818); +x_821 = lean_array_push(x_820, x_760); +lean_inc(x_762); +x_822 = lean_array_push(x_821, x_762); +lean_inc(x_762); +x_823 = lean_array_push(x_822, x_762); +lean_inc(x_762); +x_824 = lean_array_push(x_823, x_762); +lean_inc(x_762); +x_825 = lean_array_push(x_824, x_762); +x_826 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_826, 0, x_738); +lean_ctor_set(x_826, 1, x_726); +lean_ctor_set(x_826, 2, x_825); +x_827 = lean_array_push(x_743, x_826); +if (lean_obj_tag(x_807) == 0) +{ +lean_object* x_828; lean_object* x_829; lean_object* x_830; lean_object* x_831; lean_object* x_832; lean_object* x_833; lean_object* x_834; lean_object* x_835; lean_object* x_836; lean_object* x_837; lean_object* x_838; lean_object* x_839; lean_object* x_840; lean_object* x_841; lean_object* x_842; lean_object* x_843; lean_object* x_844; lean_object* x_845; +lean_dec(x_6); +x_828 = l_Lean_quoteNameMk(x_25); +x_829 = lean_array_push(x_754, x_828); +x_830 = lean_array_push(x_829, x_809); +x_831 = lean_array_push(x_830, x_711); +x_832 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_832, 0, x_738); +lean_ctor_set(x_832, 1, x_8); +lean_ctor_set(x_832, 2, x_831); +x_833 = lean_array_push(x_810, x_832); +x_834 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_834, 0, x_738); +lean_ctor_set(x_834, 1, x_798); +lean_ctor_set(x_834, 2, x_833); +x_835 = lean_array_push(x_811, x_834); +lean_inc(x_762); +x_836 = lean_array_push(x_835, x_762); +x_837 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_837, 0, x_738); +lean_ctor_set(x_837, 1, x_794); +lean_ctor_set(x_837, 2, x_836); +x_838 = lean_array_push(x_815, x_837); +lean_inc(x_762); +x_839 = lean_array_push(x_838, x_762); +lean_inc(x_762); +x_840 = lean_array_push(x_839, x_762); +x_841 = lean_array_push(x_840, x_762); +x_842 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_842, 0, x_738); +lean_ctor_set(x_842, 1, x_764); +lean_ctor_set(x_842, 2, x_841); +x_843 = lean_array_push(x_827, x_842); +x_844 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_844, 0, x_738); +lean_ctor_set(x_844, 1, x_724); +lean_ctor_set(x_844, 2, x_843); +x_845 = l_Lean_Elab_Command_elabSyntax___lambda__4(x_3, x_844, x_14, x_15, x_722); +return x_845; +} +else +{ +lean_object* x_846; lean_object* x_847; lean_object* x_848; lean_object* x_849; lean_object* x_850; lean_object* x_851; lean_object* x_852; lean_object* x_853; lean_object* x_854; lean_object* x_855; lean_object* x_856; lean_object* x_857; lean_object* x_858; lean_object* x_859; lean_object* x_860; lean_object* x_861; lean_object* x_862; lean_object* x_863; lean_object* x_864; lean_object* x_865; lean_object* x_866; lean_object* x_867; lean_object* x_868; lean_object* x_869; lean_object* x_870; lean_object* x_871; lean_object* x_872; +lean_dec(x_25); +x_846 = lean_ctor_get(x_807, 0); +lean_inc(x_846); +lean_dec(x_807); +x_847 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__22; +x_848 = lean_name_mk_string(x_6, x_847); +x_849 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__16; +x_850 = l_String_intercalate(x_849, x_846); +x_851 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__17; +x_852 = lean_string_append(x_851, x_850); +lean_dec(x_850); +x_853 = l_Lean_Syntax_mkNameLit(x_852, x_738); +x_854 = lean_array_push(x_740, x_853); +x_855 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_855, 0, x_738); +lean_ctor_set(x_855, 1, x_848); +lean_ctor_set(x_855, 2, x_854); +x_856 = lean_array_push(x_754, x_855); +x_857 = lean_array_push(x_856, x_809); +x_858 = lean_array_push(x_857, x_711); +x_859 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_859, 0, x_738); +lean_ctor_set(x_859, 1, x_8); +lean_ctor_set(x_859, 2, x_858); +x_860 = lean_array_push(x_810, x_859); +x_861 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_861, 0, x_738); +lean_ctor_set(x_861, 1, x_798); +lean_ctor_set(x_861, 2, x_860); +x_862 = lean_array_push(x_811, x_861); +lean_inc(x_762); +x_863 = lean_array_push(x_862, x_762); +x_864 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_864, 0, x_738); +lean_ctor_set(x_864, 1, x_794); +lean_ctor_set(x_864, 2, x_863); +x_865 = lean_array_push(x_815, x_864); +lean_inc(x_762); +x_866 = lean_array_push(x_865, x_762); +lean_inc(x_762); +x_867 = lean_array_push(x_866, x_762); +x_868 = lean_array_push(x_867, x_762); +x_869 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_869, 0, x_738); +lean_ctor_set(x_869, 1, x_764); +lean_ctor_set(x_869, 2, x_868); +x_870 = lean_array_push(x_827, x_869); +x_871 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_871, 0, x_738); +lean_ctor_set(x_871, 1, x_724); +lean_ctor_set(x_871, 2, x_870); +x_872 = l_Lean_Elab_Command_elabSyntax___lambda__4(x_3, x_871, x_14, x_15, x_722); +return x_872; +} +} +} +else +{ +lean_object* x_876; lean_object* x_877; lean_object* x_878; lean_object* x_879; lean_object* x_880; lean_object* x_881; lean_object* x_882; lean_object* x_883; lean_object* x_884; lean_object* x_885; lean_object* x_886; lean_object* x_887; lean_object* x_888; lean_object* x_889; lean_object* x_890; lean_object* x_891; lean_object* x_892; lean_object* x_893; lean_object* x_894; lean_object* x_895; lean_object* x_896; lean_object* x_897; lean_object* x_898; lean_object* x_899; lean_object* x_900; lean_object* x_901; lean_object* x_902; lean_object* x_903; lean_object* x_904; lean_object* x_905; lean_object* x_906; lean_object* x_907; lean_object* x_908; lean_object* x_909; lean_object* x_910; lean_object* x_911; lean_object* x_912; lean_object* x_913; lean_object* x_914; lean_object* x_915; lean_object* x_916; lean_object* x_917; lean_object* x_918; lean_object* x_919; lean_object* x_920; lean_object* x_921; lean_object* x_922; lean_object* x_923; lean_object* x_924; lean_object* x_925; lean_object* x_926; lean_object* x_927; lean_object* x_928; lean_object* x_929; lean_object* x_930; lean_object* x_931; lean_object* x_932; lean_object* x_933; lean_object* x_934; lean_object* x_935; lean_object* x_936; lean_object* x_937; lean_object* x_938; lean_object* x_939; lean_object* x_940; lean_object* x_941; lean_object* x_942; lean_object* x_943; lean_object* x_944; lean_object* x_945; lean_object* x_946; lean_object* x_947; lean_object* x_948; lean_object* x_949; lean_object* x_950; lean_object* x_951; lean_object* x_952; lean_object* x_953; lean_object* x_954; lean_object* x_955; lean_object* x_956; lean_object* x_957; lean_object* x_958; lean_object* x_959; lean_object* x_960; lean_object* x_961; lean_object* x_962; lean_object* x_963; lean_object* x_964; lean_object* x_965; lean_object* x_966; lean_object* x_967; lean_object* x_968; lean_object* x_969; lean_object* x_970; lean_object* x_971; lean_object* x_972; lean_object* x_973; lean_object* x_974; lean_object* x_975; lean_object* x_976; lean_object* x_977; lean_object* x_978; lean_object* x_979; lean_object* x_980; lean_object* x_981; lean_object* x_982; lean_object* x_983; +x_876 = lean_ctor_get(x_710, 0); +lean_inc(x_876); +lean_dec(x_710); +x_877 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_14, x_15, x_40); +x_878 = lean_ctor_get(x_877, 0); +lean_inc(x_878); +x_879 = lean_ctor_get(x_877, 1); +lean_inc(x_879); +lean_dec(x_877); +x_880 = l_Lean_Elab_Command_getCurrMacroScope(x_14, x_15, x_879); +x_881 = lean_ctor_get(x_880, 0); +lean_inc(x_881); +x_882 = lean_ctor_get(x_880, 1); +lean_inc(x_882); +lean_dec(x_880); +x_883 = l_Lean_Elab_Command_getMainModule___rarg(x_15, x_882); +x_884 = lean_ctor_get(x_883, 0); +lean_inc(x_884); +x_885 = lean_ctor_get(x_883, 1); +lean_inc(x_885); +lean_dec(x_883); +x_886 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; +lean_inc(x_5); +x_887 = lean_name_mk_string(x_5, x_886); +x_888 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; +lean_inc(x_5); +x_889 = lean_name_mk_string(x_5, x_888); +x_890 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__12; +lean_inc(x_6); +x_891 = lean_name_mk_string(x_6, x_890); +x_892 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__14; +lean_inc(x_878); +x_893 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_893, 0, x_878); +lean_ctor_set(x_893, 1, x_892); +x_894 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__15; +lean_inc(x_6); +x_895 = lean_name_mk_string(x_6, x_894); +x_896 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__21; +x_897 = lean_name_mk_string(x_7, x_896); +x_898 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__23; +x_899 = lean_name_mk_string(x_897, x_898); +x_900 = l_Nat_repr(x_19); +x_901 = lean_box(2); +x_902 = l_Lean_Syntax_mkNumLit(x_900, x_901); +x_903 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; +x_904 = lean_array_push(x_903, x_902); +lean_inc(x_8); +x_905 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_905, 0, x_901); +lean_ctor_set(x_905, 1, x_8); +lean_ctor_set(x_905, 2, x_904); +x_906 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; +x_907 = lean_array_push(x_906, x_28); +x_908 = lean_array_push(x_907, x_905); +x_909 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_909, 0, x_901); +lean_ctor_set(x_909, 1, x_899); +lean_ctor_set(x_909, 2, x_908); +x_910 = lean_array_push(x_906, x_9); +x_911 = lean_array_push(x_910, x_909); +x_912 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_912, 0, x_901); +lean_ctor_set(x_912, 1, x_895); +lean_ctor_set(x_912, 2, x_911); +x_913 = lean_array_push(x_903, x_912); +lean_inc(x_8); +x_914 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_914, 0, x_901); +lean_ctor_set(x_914, 1, x_8); +lean_ctor_set(x_914, 2, x_913); +x_915 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__30; +lean_inc(x_878); +x_916 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_916, 0, x_878); +lean_ctor_set(x_916, 1, x_915); +x_917 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__28; +x_918 = lean_array_push(x_917, x_893); +x_919 = lean_array_push(x_918, x_914); +x_920 = lean_array_push(x_919, x_916); +x_921 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_921, 0, x_901); +lean_ctor_set(x_921, 1, x_891); +lean_ctor_set(x_921, 2, x_920); +x_922 = lean_array_push(x_903, x_921); +lean_inc(x_8); +x_923 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_923, 0, x_901); +lean_ctor_set(x_923, 1, x_8); +lean_ctor_set(x_923, 2, x_922); +x_924 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__3; +lean_inc(x_8); +x_925 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_925, 0, x_901); +lean_ctor_set(x_925, 1, x_8); +lean_ctor_set(x_925, 2, x_924); +x_926 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33; +lean_inc(x_5); +x_927 = lean_name_mk_string(x_5, x_926); +lean_inc(x_878); +x_928 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_928, 0, x_878); +lean_ctor_set(x_928, 1, x_926); +x_929 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__35; +lean_inc(x_5); +x_930 = lean_name_mk_string(x_5, x_929); +x_931 = lean_array_push(x_906, x_713); +lean_inc(x_925); +x_932 = lean_array_push(x_931, x_925); +x_933 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_933, 0, x_901); +lean_ctor_set(x_933, 1, x_930); +lean_ctor_set(x_933, 2, x_932); +x_934 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__37; +lean_inc(x_5); +x_935 = lean_name_mk_string(x_5, x_934); +x_936 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__39; +lean_inc(x_6); +x_937 = lean_name_mk_string(x_6, x_936); +x_938 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__41; +lean_inc(x_878); +x_939 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_939, 0, x_878); +lean_ctor_set(x_939, 1, x_938); +x_940 = l_List_filterMap___at_Lean_Elab_Term_resolveParserName___spec__1___closed__1; +lean_inc(x_10); +x_941 = lean_name_mk_string(x_10, x_940); +lean_inc(x_881); +lean_inc(x_941); +lean_inc(x_884); +x_942 = l_Lean_addMacroScope(x_884, x_941, x_881); +x_943 = lean_box(0); +if (lean_is_scalar(x_712)) { + x_944 = lean_alloc_ctor(0, 2, 0); +} else { + x_944 = x_712; +} +lean_ctor_set(x_944, 0, x_941); +lean_ctor_set(x_944, 1, x_943); +x_945 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_945, 0, x_944); +lean_ctor_set(x_945, 1, x_943); +x_946 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__7; +lean_inc(x_878); +x_947 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_947, 0, x_878); +lean_ctor_set(x_947, 1, x_946); +lean_ctor_set(x_947, 2, x_942); +lean_ctor_set(x_947, 3, x_945); +x_948 = lean_array_push(x_906, x_939); +x_949 = lean_array_push(x_948, x_947); +x_950 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_950, 0, x_901); +lean_ctor_set(x_950, 1, x_937); +lean_ctor_set(x_950, 2, x_949); +x_951 = lean_array_push(x_903, x_950); +lean_inc(x_8); +x_952 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_952, 0, x_901); +lean_ctor_set(x_952, 1, x_8); +lean_ctor_set(x_952, 2, x_951); +lean_inc(x_925); +x_953 = lean_array_push(x_906, x_925); +x_954 = lean_array_push(x_953, x_952); +x_955 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_955, 0, x_901); +lean_ctor_set(x_955, 1, x_935); +lean_ctor_set(x_955, 2, x_954); +x_956 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__48; +x_957 = lean_name_mk_string(x_5, x_956); +x_958 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__50; +lean_inc(x_878); +x_959 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_959, 0, x_878); +lean_ctor_set(x_959, 1, x_958); +x_960 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__7; +lean_inc(x_6); +x_961 = lean_name_mk_string(x_6, x_960); +x_962 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__12; +x_963 = l_Lean_addMacroScope(x_884, x_962, x_881); +x_964 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__12; +x_965 = lean_name_mk_string(x_10, x_964); +x_966 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__11; +x_967 = lean_name_mk_string(x_965, x_966); +x_968 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_968, 0, x_967); +lean_ctor_set(x_968, 1, x_943); +x_969 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_969, 0, x_968); +lean_ctor_set(x_969, 1, x_943); +x_970 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__10; +x_971 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_971, 0, x_878); +lean_ctor_set(x_971, 1, x_970); +lean_ctor_set(x_971, 2, x_963); +lean_ctor_set(x_971, 3, x_969); +lean_inc(x_25); +x_972 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_943, x_25); +x_973 = l_Nat_repr(x_11); +x_974 = l_Lean_Syntax_mkNumLit(x_973, x_901); +x_975 = l_Nat_repr(x_876); +x_976 = l_Lean_Syntax_mkNumLit(x_975, x_901); +x_977 = lean_array_push(x_906, x_971); +x_978 = lean_array_push(x_917, x_959); +x_979 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__85; +x_980 = lean_array_push(x_979, x_928); +x_981 = lean_array_push(x_980, x_933); +x_982 = lean_array_push(x_981, x_955); +if (lean_obj_tag(x_12) == 0) +{ +x_983 = x_924; +goto block_1044; +} +else +{ +lean_object* x_1045; lean_object* x_1046; +x_1045 = lean_ctor_get(x_12, 0); +lean_inc(x_1045); +lean_dec(x_12); +x_1046 = lean_array_push(x_903, x_1045); +x_983 = x_1046; +goto block_1044; +} +block_1044: +{ +lean_object* x_984; lean_object* x_985; lean_object* x_986; lean_object* x_987; lean_object* x_988; lean_object* x_989; lean_object* x_990; lean_object* x_991; lean_object* x_992; lean_object* x_993; lean_object* x_994; +x_984 = l_Array_append___rarg(x_924, x_983); +lean_inc(x_8); +x_985 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_985, 0, x_901); +lean_ctor_set(x_985, 1, x_8); +lean_ctor_set(x_985, 2, x_984); +x_986 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31; +x_987 = lean_array_push(x_986, x_985); +x_988 = lean_array_push(x_987, x_923); +lean_inc(x_925); +x_989 = lean_array_push(x_988, x_925); +lean_inc(x_925); +x_990 = lean_array_push(x_989, x_925); +lean_inc(x_925); +x_991 = lean_array_push(x_990, x_925); +lean_inc(x_925); +x_992 = lean_array_push(x_991, x_925); +x_993 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_993, 0, x_901); +lean_ctor_set(x_993, 1, x_889); +lean_ctor_set(x_993, 2, x_992); +x_994 = lean_array_push(x_906, x_993); +if (lean_obj_tag(x_972) == 0) +{ +lean_object* x_995; lean_object* x_996; lean_object* x_997; lean_object* x_998; lean_object* x_999; lean_object* x_1000; lean_object* x_1001; lean_object* x_1002; lean_object* x_1003; lean_object* x_1004; lean_object* x_1005; lean_object* x_1006; lean_object* x_1007; lean_object* x_1008; lean_object* x_1009; lean_object* x_1010; lean_object* x_1011; lean_object* x_1012; lean_object* x_1013; lean_object* x_1014; +lean_dec(x_6); +x_995 = l_Lean_quoteNameMk(x_25); +x_996 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__8; +x_997 = lean_array_push(x_996, x_995); +x_998 = lean_array_push(x_997, x_974); +x_999 = lean_array_push(x_998, x_976); +x_1000 = lean_array_push(x_999, x_711); +x_1001 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1001, 0, x_901); +lean_ctor_set(x_1001, 1, x_8); +lean_ctor_set(x_1001, 2, x_1000); +x_1002 = lean_array_push(x_977, x_1001); +x_1003 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1003, 0, x_901); +lean_ctor_set(x_1003, 1, x_961); +lean_ctor_set(x_1003, 2, x_1002); +x_1004 = lean_array_push(x_978, x_1003); +lean_inc(x_925); +x_1005 = lean_array_push(x_1004, x_925); +x_1006 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1006, 0, x_901); +lean_ctor_set(x_1006, 1, x_957); +lean_ctor_set(x_1006, 2, x_1005); +x_1007 = lean_array_push(x_982, x_1006); +lean_inc(x_925); +x_1008 = lean_array_push(x_1007, x_925); +lean_inc(x_925); +x_1009 = lean_array_push(x_1008, x_925); +x_1010 = lean_array_push(x_1009, x_925); +x_1011 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1011, 0, x_901); +lean_ctor_set(x_1011, 1, x_927); +lean_ctor_set(x_1011, 2, x_1010); +x_1012 = lean_array_push(x_994, x_1011); +x_1013 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1013, 0, x_901); +lean_ctor_set(x_1013, 1, x_887); +lean_ctor_set(x_1013, 2, x_1012); +x_1014 = l_Lean_Elab_Command_elabSyntax___lambda__4(x_3, x_1013, x_14, x_15, x_885); +return x_1014; +} +else +{ +lean_object* x_1015; lean_object* x_1016; lean_object* x_1017; lean_object* x_1018; lean_object* x_1019; lean_object* x_1020; lean_object* x_1021; lean_object* x_1022; lean_object* x_1023; lean_object* x_1024; lean_object* x_1025; lean_object* x_1026; lean_object* x_1027; lean_object* x_1028; lean_object* x_1029; lean_object* x_1030; lean_object* x_1031; lean_object* x_1032; lean_object* x_1033; lean_object* x_1034; lean_object* x_1035; lean_object* x_1036; lean_object* x_1037; lean_object* x_1038; lean_object* x_1039; lean_object* x_1040; lean_object* x_1041; lean_object* x_1042; lean_object* x_1043; +lean_dec(x_25); +x_1015 = lean_ctor_get(x_972, 0); +lean_inc(x_1015); +lean_dec(x_972); +x_1016 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__22; +x_1017 = lean_name_mk_string(x_6, x_1016); +x_1018 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__16; +x_1019 = l_String_intercalate(x_1018, x_1015); +x_1020 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__17; +x_1021 = lean_string_append(x_1020, x_1019); +lean_dec(x_1019); +x_1022 = l_Lean_Syntax_mkNameLit(x_1021, x_901); +x_1023 = lean_array_push(x_903, x_1022); +x_1024 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1024, 0, x_901); +lean_ctor_set(x_1024, 1, x_1017); +lean_ctor_set(x_1024, 2, x_1023); +x_1025 = l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__8; +x_1026 = lean_array_push(x_1025, x_1024); +x_1027 = lean_array_push(x_1026, x_974); +x_1028 = lean_array_push(x_1027, x_976); +x_1029 = lean_array_push(x_1028, x_711); +x_1030 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1030, 0, x_901); +lean_ctor_set(x_1030, 1, x_8); +lean_ctor_set(x_1030, 2, x_1029); +x_1031 = lean_array_push(x_977, x_1030); +x_1032 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1032, 0, x_901); +lean_ctor_set(x_1032, 1, x_961); +lean_ctor_set(x_1032, 2, x_1031); +x_1033 = lean_array_push(x_978, x_1032); +lean_inc(x_925); +x_1034 = lean_array_push(x_1033, x_925); +x_1035 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1035, 0, x_901); +lean_ctor_set(x_1035, 1, x_957); +lean_ctor_set(x_1035, 2, x_1034); +x_1036 = lean_array_push(x_982, x_1035); +lean_inc(x_925); +x_1037 = lean_array_push(x_1036, x_925); +lean_inc(x_925); +x_1038 = lean_array_push(x_1037, x_925); +x_1039 = lean_array_push(x_1038, x_925); +x_1040 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1040, 0, x_901); +lean_ctor_set(x_1040, 1, x_927); +lean_ctor_set(x_1040, 2, x_1039); +x_1041 = lean_array_push(x_994, x_1040); +x_1042 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1042, 0, x_901); +lean_ctor_set(x_1042, 1, x_887); +lean_ctor_set(x_1042, 2, x_1041); +x_1043 = l_Lean_Elab_Command_elabSyntax___lambda__4(x_3, x_1042, x_14, x_15, x_885); +return x_1043; +} +} +} +} +} +else +{ +uint8_t x_1047; +lean_dec(x_28); +lean_dec(x_25); +lean_dec(x_19); +lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); @@ -17177,29 +18593,30 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_726 = !lean_is_exclusive(x_36); -if (x_726 == 0) +x_1047 = !lean_is_exclusive(x_37); +if (x_1047 == 0) { -return x_36; +return x_37; } else { -lean_object* x_727; lean_object* x_728; lean_object* x_729; -x_727 = lean_ctor_get(x_36, 0); -x_728 = lean_ctor_get(x_36, 1); -lean_inc(x_728); -lean_inc(x_727); -lean_dec(x_36); -x_729 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_729, 0, x_727); -lean_ctor_set(x_729, 1, x_728); -return x_729; +lean_object* x_1048; lean_object* x_1049; lean_object* x_1050; +x_1048 = lean_ctor_get(x_37, 0); +x_1049 = lean_ctor_get(x_37, 1); +lean_inc(x_1049); +lean_inc(x_1048); +lean_dec(x_37); +x_1050 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1050, 0, x_1048); +lean_ctor_set(x_1050, 1, x_1049); +return x_1050; } } } else { -uint8_t x_730; +uint8_t x_1051; +lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); @@ -17213,58 +18630,59 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_730 = !lean_is_exclusive(x_17); -if (x_730 == 0) +x_1051 = !lean_is_exclusive(x_18); +if (x_1051 == 0) { -return x_17; +return x_18; } else { -lean_object* x_731; lean_object* x_732; lean_object* x_733; -x_731 = lean_ctor_get(x_17, 0); -x_732 = lean_ctor_get(x_17, 1); -lean_inc(x_732); -lean_inc(x_731); -lean_dec(x_17); -x_733 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_733, 0, x_731); -lean_ctor_set(x_733, 1, x_732); -return x_733; +lean_object* x_1052; lean_object* x_1053; lean_object* x_1054; +x_1052 = lean_ctor_get(x_18, 0); +x_1053 = lean_ctor_get(x_18, 1); +lean_inc(x_1053); +lean_inc(x_1052); +lean_dec(x_18); +x_1054 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1054, 0, x_1052); +lean_ctor_set(x_1054, 1, x_1053); +return x_1054; } } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSyntax___lambda__6(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_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) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSyntax___lambda__6(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_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, lean_object* x_16) { _start: { -if (lean_obj_tag(x_11) == 0) +if (lean_obj_tag(x_12) == 0) { -lean_object* x_16; lean_object* x_17; +lean_object* x_17; lean_object* x_18; lean_inc(x_4); lean_inc(x_2); -x_16 = lean_alloc_closure((void*)(l_Lean_Elab_Command_mkNameFromParserSyntax), 4, 2); -lean_closure_set(x_16, 0, x_2); -lean_closure_set(x_16, 1, x_4); +x_17 = lean_alloc_closure((void*)(l_Lean_Elab_Command_mkNameFromParserSyntax), 4, 2); +lean_closure_set(x_17, 0, x_2); +lean_closure_set(x_17, 1, x_4); +lean_inc(x_15); lean_inc(x_14); -lean_inc(x_13); -x_17 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabSyntax___spec__7(x_16, x_13, x_14, x_15); -if (lean_obj_tag(x_17) == 0) +x_18 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabSyntax___spec__9(x_17, x_14, x_15, x_16); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_18, 0); lean_inc(x_19); -lean_dec(x_17); -x_20 = l_Lean_Elab_Command_elabSyntax___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_12, x_10, x_18, x_13, x_14, x_19); -return x_20; +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = l_Lean_Elab_Command_elabSyntax___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_13, x_11, x_19, x_14, x_15, x_20); +return x_21; } else { -uint8_t x_21; +uint8_t x_22; +lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); -lean_dec(x_12); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -17275,36 +18693,36 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_21 = !lean_is_exclusive(x_17); -if (x_21 == 0) +x_22 = !lean_is_exclusive(x_18); +if (x_22 == 0) { -return x_17; +return x_18; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_17, 0); -x_23 = lean_ctor_get(x_17, 1); +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_18, 0); +x_24 = lean_ctor_get(x_18, 1); +lean_inc(x_24); lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_17); -x_24 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_23); -return x_24; +lean_dec(x_18); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; } } } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_11, 0); -lean_inc(x_25); -lean_dec(x_11); -x_26 = l_Lean_Syntax_getId(x_25); -lean_dec(x_25); -x_27 = l_Lean_Elab_Command_elabSyntax___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_12, x_10, x_26, x_13, x_14, x_15); -return x_27; +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_12, 0); +lean_inc(x_26); +lean_dec(x_12); +x_27 = l_Lean_Syntax_getId(x_26); +lean_dec(x_26); +x_28 = l_Lean_Elab_Command_elabSyntax___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_13, x_11, x_27, x_14, x_15, x_16); +return x_28; } } } @@ -17314,7 +18732,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSyntax___lambda__7(lean_object* lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_dec(x_13); x_17 = lean_box(2); -x_18 = l_Lean_nullKind; +x_18 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; x_19 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_19, 0, x_17); lean_ctor_set(x_19, 1, x_18); @@ -17327,7 +18745,7 @@ if (lean_obj_tag(x_12) == 0) { lean_object* x_21; lean_object* x_22; x_21 = l_Lean_Parser_leadPrec; -x_22 = l_Lean_Elab_Command_elabSyntax___lambda__6(x_2, x_3, x_4, x_19, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_21, x_14, x_15, x_16); +x_22 = l_Lean_Elab_Command_elabSyntax___lambda__6(x_2, x_3, x_4, x_19, x_5, x_6, x_7, x_18, x_8, x_9, x_10, x_11, x_21, x_14, x_15, x_16); return x_22; } else @@ -17340,7 +18758,7 @@ x_24 = lean_alloc_closure((void*)(l_Lean_evalPrec), 3, 1); lean_closure_set(x_24, 0, x_23); lean_inc(x_15); lean_inc(x_14); -x_25 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabSyntax___spec__2(x_24, x_14, x_15, x_16); +x_25 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabSyntax___spec__4(x_24, x_14, x_15, x_16); if (lean_obj_tag(x_25) == 0) { lean_object* x_26; lean_object* x_27; lean_object* x_28; @@ -17349,7 +18767,7 @@ lean_inc(x_26); x_27 = lean_ctor_get(x_25, 1); lean_inc(x_27); lean_dec(x_25); -x_28 = l_Lean_Elab_Command_elabSyntax___lambda__6(x_2, x_3, x_4, x_19, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_26, x_14, x_15, x_27); +x_28 = l_Lean_Elab_Command_elabSyntax___lambda__6(x_2, x_3, x_4, x_19, x_5, x_6, x_7, x_18, x_8, x_9, x_10, x_11, x_26, x_14, x_15, x_27); return x_28; } else @@ -17395,7 +18813,7 @@ if (lean_obj_tag(x_12) == 0) { lean_object* x_33; lean_object* x_34; x_33 = l_Lean_Parser_maxPrec; -x_34 = l_Lean_Elab_Command_elabSyntax___lambda__6(x_2, x_3, x_4, x_19, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_33, x_14, x_15, x_16); +x_34 = l_Lean_Elab_Command_elabSyntax___lambda__6(x_2, x_3, x_4, x_19, x_5, x_6, x_7, x_18, x_8, x_9, x_10, x_11, x_33, x_14, x_15, x_16); return x_34; } else @@ -17408,7 +18826,7 @@ x_36 = lean_alloc_closure((void*)(l_Lean_evalPrec), 3, 1); lean_closure_set(x_36, 0, x_35); lean_inc(x_15); lean_inc(x_14); -x_37 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabSyntax___spec__2(x_36, x_14, x_15, x_16); +x_37 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabSyntax___spec__4(x_36, x_14, x_15, x_16); if (lean_obj_tag(x_37) == 0) { lean_object* x_38; lean_object* x_39; lean_object* x_40; @@ -17417,7 +18835,7 @@ lean_inc(x_38); x_39 = lean_ctor_get(x_37, 1); lean_inc(x_39); lean_dec(x_37); -x_40 = l_Lean_Elab_Command_elabSyntax___lambda__6(x_2, x_3, x_4, x_19, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_38, x_14, x_15, x_39); +x_40 = l_Lean_Elab_Command_elabSyntax___lambda__6(x_2, x_3, x_4, x_19, x_5, x_6, x_7, x_18, x_8, x_9, x_10, x_11, x_38, x_14, x_15, x_39); return x_40; } else @@ -17494,7 +18912,7 @@ x_17 = l_Lean_Syntax_getArg(x_1, x_16); x_18 = l_Lean_Syntax_getArgs(x_17); lean_dec(x_17); x_19 = l_Lean_Elab_Command_elabSyntax___lambda__8___closed__1; -x_20 = l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__1(x_18, x_19); +x_20 = l_Array_sequenceMap___at_Lean_Elab_Command_elabSyntax___spec__2(x_18, x_19); lean_dec(x_18); if (lean_obj_tag(x_20) == 0) { @@ -17562,7 +18980,7 @@ x_35 = l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNull x_36 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_36, 0, x_34); lean_ctor_set(x_36, 1, x_35); -x_37 = l_Lean_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__11(x_24, x_36, x_13, x_14, x_29); +x_37 = l_Lean_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__13(x_24, x_36, x_13, x_14, x_29); lean_dec(x_14); lean_dec(x_24); x_38 = !lean_is_exclusive(x_37); @@ -17613,14 +19031,13 @@ x_16 = l_Lean_Syntax_getArg(x_1, x_15); x_17 = l_Lean_Syntax_isNone(x_16); if (x_17 == 0) { -lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_18 = l_Lean_nullKind; -x_19 = lean_unsigned_to_nat(1u); +lean_object* x_18; uint8_t x_19; +x_18 = lean_unsigned_to_nat(1u); lean_inc(x_16); -x_20 = l_Lean_Syntax_isNodeOf(x_16, x_18, x_19); -if (x_20 == 0) +x_19 = l_Lean_Syntax_matchesNull(x_16, x_18); +if (x_19 == 0) { -lean_object* x_21; +lean_object* x_20; lean_dec(x_16); lean_dec(x_13); lean_dec(x_12); @@ -17634,25 +19051,25 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_21 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_14); -return x_21; +x_20 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_14); +return x_20; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; -x_22 = lean_unsigned_to_nat(0u); -x_23 = l_Lean_Syntax_getArg(x_16, x_22); +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_21 = lean_unsigned_to_nat(0u); +x_22 = l_Lean_Syntax_getArg(x_16, x_21); lean_dec(x_16); -x_24 = l_Lean_Elab_Command_elabSyntax___lambda__9___closed__1; +x_23 = l_Lean_Elab_Command_elabSyntax___lambda__9___closed__1; lean_inc(x_3); -x_25 = lean_name_mk_string(x_3, x_24); -lean_inc(x_23); -x_26 = l_Lean_Syntax_isOfKind(x_23, x_25); -lean_dec(x_25); -if (x_26 == 0) +x_24 = lean_name_mk_string(x_3, x_23); +lean_inc(x_22); +x_25 = l_Lean_Syntax_isOfKind(x_22, x_24); +lean_dec(x_24); +if (x_25 == 0) { -lean_object* x_27; -lean_dec(x_23); +lean_object* x_26; +lean_dec(x_22); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -17665,31 +19082,31 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_27 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_14); -return x_27; +x_26 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_14); +return x_26; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_28 = lean_unsigned_to_nat(3u); -x_29 = l_Lean_Syntax_getArg(x_23, x_28); -lean_dec(x_23); -x_30 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_30, 0, x_29); -x_31 = lean_box(0); -x_32 = l_Lean_Elab_Command_elabSyntax___lambda__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_11, x_9, x_31, x_30, x_12, x_13, x_14); -return x_32; +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_27 = lean_unsigned_to_nat(3u); +x_28 = l_Lean_Syntax_getArg(x_22, x_27); +lean_dec(x_22); +x_29 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_29, 0, x_28); +x_30 = lean_box(0); +x_31 = l_Lean_Elab_Command_elabSyntax___lambda__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_11, x_9, x_30, x_29, x_12, x_13, x_14); +return x_31; } } } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_dec(x_16); +x_32 = lean_box(0); x_33 = lean_box(0); -x_34 = lean_box(0); -x_35 = l_Lean_Elab_Command_elabSyntax___lambda__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_11, x_9, x_34, x_33, x_12, x_13, x_14); -return x_35; +x_34 = l_Lean_Elab_Command_elabSyntax___lambda__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_11, x_9, x_33, x_32, x_12, x_13, x_14); +return x_34; } } } @@ -17711,14 +19128,13 @@ x_15 = l_Lean_Syntax_getArg(x_1, x_14); x_16 = l_Lean_Syntax_isNone(x_15); if (x_16 == 0) { -lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_17 = l_Lean_nullKind; -x_18 = lean_unsigned_to_nat(1u); +lean_object* x_17; uint8_t x_18; +x_17 = lean_unsigned_to_nat(1u); lean_inc(x_15); -x_19 = l_Lean_Syntax_isNodeOf(x_15, x_17, x_18); -if (x_19 == 0) +x_18 = l_Lean_Syntax_matchesNull(x_15, x_17); +if (x_18 == 0) { -lean_object* x_20; +lean_object* x_19; lean_dec(x_15); lean_dec(x_12); lean_dec(x_11); @@ -17731,25 +19147,25 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_20 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_13); -return x_20; +x_19 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_13); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_21 = lean_unsigned_to_nat(0u); -x_22 = l_Lean_Syntax_getArg(x_15, x_21); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = lean_unsigned_to_nat(0u); +x_21 = l_Lean_Syntax_getArg(x_15, x_20); lean_dec(x_15); -x_23 = l_Lean_Elab_Command_elabSyntax___lambda__10___closed__1; +x_22 = l_Lean_Elab_Command_elabSyntax___lambda__10___closed__1; lean_inc(x_3); -x_24 = lean_name_mk_string(x_3, x_23); -lean_inc(x_22); -x_25 = l_Lean_Syntax_isOfKind(x_22, x_24); -lean_dec(x_24); -if (x_25 == 0) +x_23 = lean_name_mk_string(x_3, x_22); +lean_inc(x_21); +x_24 = l_Lean_Syntax_isOfKind(x_21, x_23); +lean_dec(x_23); +if (x_24 == 0) { -lean_object* x_26; -lean_dec(x_22); +lean_object* x_25; +lean_dec(x_21); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -17761,31 +19177,31 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_26 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_13); -return x_26; +x_25 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_13); +return x_25; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_27 = lean_unsigned_to_nat(3u); -x_28 = l_Lean_Syntax_getArg(x_22, x_27); -lean_dec(x_22); -x_29 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_29, 0, x_28); -x_30 = lean_box(0); -x_31 = l_Lean_Elab_Command_elabSyntax___lambda__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_10, x_30, x_29, x_11, x_12, x_13); -return x_31; +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_26 = lean_unsigned_to_nat(3u); +x_27 = l_Lean_Syntax_getArg(x_21, x_26); +lean_dec(x_21); +x_28 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_28, 0, x_27); +x_29 = lean_box(0); +x_30 = l_Lean_Elab_Command_elabSyntax___lambda__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_10, x_29, x_28, x_11, x_12, x_13); +return x_30; } } } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_dec(x_15); +x_31 = lean_box(0); x_32 = lean_box(0); -x_33 = lean_box(0); -x_34 = l_Lean_Elab_Command_elabSyntax___lambda__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_10, x_33, x_32, x_11, x_12, x_13); -return x_34; +x_33 = l_Lean_Elab_Command_elabSyntax___lambda__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_10, x_32, x_31, x_11, x_12, x_13); +return x_33; } } } @@ -17807,7 +19223,7 @@ x_12 = l_Lean_Syntax_getArg(x_1, x_11); x_13 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__5; lean_inc(x_2); x_14 = lean_name_mk_string(x_2, x_13); -x_15 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__18; +x_15 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__17; lean_inc(x_14); x_16 = lean_name_mk_string(x_14, x_15); lean_inc(x_12); @@ -17837,13 +19253,12 @@ x_20 = l_Lean_Syntax_getArg(x_1, x_19); x_21 = l_Lean_Syntax_isNone(x_20); if (x_21 == 0) { -lean_object* x_22; uint8_t x_23; -x_22 = l_Lean_nullKind; +uint8_t x_22; lean_inc(x_20); -x_23 = l_Lean_Syntax_isNodeOf(x_20, x_22, x_11); -if (x_23 == 0) +x_22 = l_Lean_Syntax_matchesNull(x_20, x_11); +if (x_22 == 0) { -lean_object* x_24; +lean_object* x_23; lean_dec(x_20); lean_dec(x_14); lean_dec(x_12); @@ -17855,25 +19270,25 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_24 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_10); -return x_24; +x_23 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_10); +return x_23; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_25 = lean_unsigned_to_nat(0u); -x_26 = l_Lean_Syntax_getArg(x_20, x_25); +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_24 = lean_unsigned_to_nat(0u); +x_25 = l_Lean_Syntax_getArg(x_20, x_24); lean_dec(x_20); -x_27 = l_Lean_Elab_Command_elabSyntax___lambda__11___closed__1; +x_26 = l_Lean_Elab_Command_elabSyntax___lambda__11___closed__1; lean_inc(x_2); -x_28 = lean_name_mk_string(x_2, x_27); -lean_inc(x_26); -x_29 = l_Lean_Syntax_isOfKind(x_26, x_28); -lean_dec(x_28); -if (x_29 == 0) +x_27 = lean_name_mk_string(x_2, x_26); +lean_inc(x_25); +x_28 = l_Lean_Syntax_isOfKind(x_25, x_27); +lean_dec(x_27); +if (x_28 == 0) { -lean_object* x_30; -lean_dec(x_26); +lean_object* x_29; +lean_dec(x_25); lean_dec(x_14); lean_dec(x_12); lean_dec(x_9); @@ -17884,30 +19299,30 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_30 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_10); -return x_30; +x_29 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_10); +return x_29; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_31 = l_Lean_Syntax_getArg(x_26, x_11); -lean_dec(x_26); -x_32 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_32, 0, x_31); -x_33 = lean_box(0); -x_34 = l_Lean_Elab_Command_elabSyntax___lambda__10(x_1, x_3, x_4, x_14, x_2, x_12, x_5, x_7, x_33, x_32, x_8, x_9, x_10); -return x_34; +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_30 = l_Lean_Syntax_getArg(x_25, x_11); +lean_dec(x_25); +x_31 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_31, 0, x_30); +x_32 = lean_box(0); +x_33 = l_Lean_Elab_Command_elabSyntax___lambda__10(x_1, x_3, x_4, x_14, x_2, x_12, x_5, x_7, x_32, x_31, x_8, x_9, x_10); +return x_33; } } } else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_dec(x_20); +x_34 = lean_box(0); x_35 = lean_box(0); -x_36 = lean_box(0); -x_37 = l_Lean_Elab_Command_elabSyntax___lambda__10(x_1, x_3, x_4, x_14, x_2, x_12, x_5, x_7, x_36, x_35, x_8, x_9, x_10); -return x_37; +x_36 = l_Lean_Elab_Command_elabSyntax___lambda__10(x_1, x_3, x_4, x_14, x_2, x_12, x_5, x_7, x_35, x_34, x_8, x_9, x_10); +return x_36; } } } @@ -17972,66 +19387,65 @@ x_9 = l_Lean_Syntax_getArg(x_1, x_8); x_10 = l_Lean_Syntax_isNone(x_9); if (x_10 == 0) { -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = l_Lean_nullKind; -x_12 = lean_unsigned_to_nat(1u); +lean_object* x_11; uint8_t x_12; +x_11 = lean_unsigned_to_nat(1u); lean_inc(x_9); -x_13 = l_Lean_Syntax_isNodeOf(x_9, x_11, x_12); -if (x_13 == 0) +x_12 = l_Lean_Syntax_matchesNull(x_9, x_11); +if (x_12 == 0) { -lean_object* x_14; +lean_object* x_13; lean_dec(x_9); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_14 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_4); -return x_14; +x_13 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_4); +return x_13; } else { -lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_15 = l_Lean_Syntax_getArg(x_9, x_8); +lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_14 = l_Lean_Syntax_getArg(x_9, x_8); lean_dec(x_9); -x_16 = l_Lean_Elab_Command_elabSyntax___closed__4; -lean_inc(x_15); -x_17 = l_Lean_Syntax_isOfKind(x_15, x_16); -if (x_17 == 0) +x_15 = l_Lean_Elab_Command_elabSyntax___closed__4; +lean_inc(x_14); +x_16 = l_Lean_Syntax_isOfKind(x_14, x_15); +if (x_16 == 0) { -lean_object* x_18; -lean_dec(x_15); +lean_object* x_17; +lean_dec(x_14); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_18 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_4); -return x_18; +x_17 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_4); +return x_17; } 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; -x_19 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_19, 0, x_15); -x_20 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__4; -x_21 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__6; -x_22 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__2; -x_23 = lean_box(0); +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_18 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_18, 0, x_14); +x_19 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__4; +x_20 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__6; +x_21 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__2; +x_22 = lean_box(0); lean_inc(x_1); -x_24 = l_Lean_Elab_Command_elabSyntax___lambda__11(x_1, x_20, x_1, x_21, x_22, x_23, x_19, x_2, x_3, x_4); -return x_24; +x_23 = l_Lean_Elab_Command_elabSyntax___lambda__11(x_1, x_19, x_1, x_20, x_21, x_22, x_18, x_2, x_3, x_4); +return x_23; } } } 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_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_dec(x_9); -x_25 = lean_box(0); -x_26 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__4; -x_27 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__6; -x_28 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__2; -x_29 = lean_box(0); +x_24 = lean_box(0); +x_25 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__4; +x_26 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__6; +x_27 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__2; +x_28 = lean_box(0); lean_inc(x_1); -x_30 = l_Lean_Elab_Command_elabSyntax___lambda__11(x_1, x_26, x_1, x_27, x_28, x_29, x_25, x_2, x_3, x_4); -return x_30; +x_29 = l_Lean_Elab_Command_elabSyntax___lambda__11(x_1, x_25, x_1, x_26, x_27, x_28, x_24, x_2, x_3, x_4); +return x_29; } } } @@ -18046,70 +19460,88 @@ lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabSyntax___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Elab_Command_elabSyntax___spec__3___boxed(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; +x_6 = l_Array_sequenceMap_loop___at_Lean_Elab_Command_elabSyntax___spec__3(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Elab_Command_elabSyntax___spec__2___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Array_sequenceMap___at_Lean_Elab_Command_elabSyntax___spec__2(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabSyntax___spec__6___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_throwError___at_Lean_Elab_Command_elabSyntax___spec__4(x_1, x_2, x_3, x_4); +x_5 = l_Lean_throwError___at_Lean_Elab_Command_elabSyntax___spec__6(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabSyntax___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabSyntax___spec__7___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_throwMaxRecDepthAt___at_Lean_Elab_Command_elabSyntax___spec__5(x_1, x_2, x_3, x_4); +x_5 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabSyntax___spec__7(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_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__6___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__8___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__6(x_1, x_2); +x_3 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__8(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__10___boxed(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; -x_6 = l_Lean_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__8(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__10(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); lean_dec(x_1); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabSyntax___spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabSyntax___spec__11___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_throwMaxRecDepthAt___at_Lean_Elab_Command_elabSyntax___spec__9(x_1, x_2, x_3, x_4); +x_5 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabSyntax___spec__11(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_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__10___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__12___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__10(x_1, x_2); +x_3 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__12(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__13___boxed(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; -x_6 = l_Lean_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__11(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__13(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); lean_dec(x_1); return x_6; @@ -18176,7 +19608,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabSyntax_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(303u); +x_1 = lean_unsigned_to_nat(326u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -18188,7 +19620,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabSyntax_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(330u); +x_1 = lean_unsigned_to_nat(353u); x_2 = lean_unsigned_to_nat(43u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -18216,7 +19648,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabSyntax_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(303u); +x_1 = lean_unsigned_to_nat(326u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -18228,7 +19660,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabSyntax_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(303u); +x_1 = lean_unsigned_to_nat(326u); x_2 = lean_unsigned_to_nat(45u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -18279,7 +19711,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__1(lean_o { lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; x_11 = lean_box(2); -x_12 = l_Lean_nullKind; +x_12 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; x_13 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_13, 0, x_11); lean_ctor_set(x_13, 1, x_12); @@ -18530,7 +19962,7 @@ x_17 = l_Lean_Syntax_getArg(x_1, x_16); x_18 = l_Lean_Syntax_getArgs(x_17); lean_dec(x_17); x_19 = l_Lean_Elab_Command_elabSyntax___lambda__8___closed__1; -x_20 = l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__1(x_18, x_19); +x_20 = l_Array_sequenceMap___at_Lean_Elab_Command_elabSyntax___spec__2(x_18, x_19); lean_dec(x_18); if (lean_obj_tag(x_20) == 0) { @@ -18573,602 +20005,915 @@ lean_inc(x_8); x_31 = l_Lean_Elab_Command_liftTermElabM___rarg(x_23, x_30, x_8, x_9, x_26); if (lean_obj_tag(x_31) == 0) { -lean_object* x_32; lean_object* x_33; uint8_t x_34; +lean_object* x_32; uint8_t x_33; x_32 = lean_ctor_get(x_31, 0); lean_inc(x_32); -x_33 = lean_ctor_get(x_31, 1); -lean_inc(x_33); +x_33 = !lean_is_exclusive(x_32); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_34 = lean_ctor_get(x_32, 0); +x_35 = lean_ctor_get(x_32, 1); +lean_dec(x_35); +x_36 = lean_ctor_get(x_31, 1); +lean_inc(x_36); lean_dec(x_31); -x_34 = !lean_is_exclusive(x_32); -if (x_34 == 0) +x_37 = !lean_is_exclusive(x_34); +if (x_37 == 0) { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; uint8_t x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; -x_35 = lean_ctor_get(x_32, 0); -x_36 = lean_ctor_get(x_32, 1); -lean_dec(x_36); -x_37 = l_Lean_Elab_Command_getScope___rarg(x_9, x_33); -x_38 = lean_ctor_get(x_37, 0); -lean_inc(x_38); -x_39 = lean_ctor_get(x_37, 1); -lean_inc(x_39); -lean_dec(x_37); -x_40 = lean_ctor_get(x_38, 2); -lean_inc(x_40); -lean_dec(x_38); -x_41 = l_Lean_Syntax_getId(x_12); +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; uint8_t x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; +x_38 = lean_ctor_get(x_34, 0); +x_39 = lean_ctor_get(x_34, 1); +lean_dec(x_39); +x_40 = l_Lean_Elab_Command_getScope___rarg(x_9, x_36); +x_41 = lean_ctor_get(x_40, 0); lean_inc(x_41); -x_42 = l_Lean_Name_append(x_40, x_41); +x_42 = lean_ctor_get(x_40, 1); +lean_inc(x_42); lean_dec(x_40); -x_43 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_8, x_9, x_39); -x_44 = lean_ctor_get(x_43, 0); +x_43 = lean_ctor_get(x_41, 2); +lean_inc(x_43); +lean_dec(x_41); +x_44 = l_Lean_Syntax_getId(x_12); lean_inc(x_44); -x_45 = lean_ctor_get(x_43, 1); -lean_inc(x_45); +x_45 = l_Lean_Name_append(x_43, x_44); lean_dec(x_43); -x_46 = l_Lean_Elab_Command_getCurrMacroScope(x_8, x_9, x_45); +x_46 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_8, x_9, x_42); x_47 = lean_ctor_get(x_46, 0); lean_inc(x_47); x_48 = lean_ctor_get(x_46, 1); lean_inc(x_48); lean_dec(x_46); -x_49 = l_Lean_Elab_Command_getMainModule___rarg(x_9, x_48); +x_49 = l_Lean_Elab_Command_getCurrMacroScope(x_8, x_9, x_48); x_50 = lean_ctor_get(x_49, 0); lean_inc(x_50); x_51 = lean_ctor_get(x_49, 1); lean_inc(x_51); lean_dec(x_49); -x_52 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; -lean_inc(x_2); -x_53 = lean_name_mk_string(x_2, x_52); -x_54 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; +x_52 = l_Lean_Elab_Command_getMainModule___rarg(x_9, x_51); +x_53 = lean_ctor_get(x_52, 0); +lean_inc(x_53); +x_54 = lean_ctor_get(x_52, 1); +lean_inc(x_54); +lean_dec(x_52); +x_55 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; lean_inc(x_2); -x_55 = lean_name_mk_string(x_2, x_54); -x_56 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__34; +x_56 = lean_name_mk_string(x_2, x_55); +x_57 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; lean_inc(x_2); -x_57 = lean_name_mk_string(x_2, x_56); -lean_inc(x_44); -x_58 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_58, 0, x_44); -lean_ctor_set(x_58, 1, x_56); -x_59 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__1; +x_58 = lean_name_mk_string(x_2, x_57); +x_59 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33; lean_inc(x_2); x_60 = lean_name_mk_string(x_2, x_59); -x_61 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; -x_62 = lean_array_push(x_61, x_12); -x_63 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__12; -x_64 = lean_array_push(x_62, x_63); -x_65 = lean_box(2); -x_66 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_66, 0, x_65); -lean_ctor_set(x_66, 1, x_60); -lean_ctor_set(x_66, 2, x_64); -x_67 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__36; +lean_inc(x_47); +x_61 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_61, 0, x_47); +lean_ctor_set(x_61, 1, x_59); +x_62 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__35; lean_inc(x_2); -x_68 = lean_name_mk_string(x_2, x_67); -x_69 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__5; -x_70 = lean_name_mk_string(x_3, x_69); -x_71 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__38; -lean_inc(x_70); -x_72 = lean_name_mk_string(x_70, x_71); -x_73 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__40; -lean_inc(x_44); -x_74 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_74, 0, x_44); -lean_ctor_set(x_74, 1, x_73); -x_75 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__12; -x_76 = lean_name_mk_string(x_4, x_75); +x_63 = lean_name_mk_string(x_2, x_62); +x_64 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; +x_65 = lean_array_push(x_64, x_12); +x_66 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__11; +x_67 = lean_array_push(x_65, x_66); +x_68 = lean_box(2); +x_69 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_69, 0, x_68); +lean_ctor_set(x_69, 1, x_63); +lean_ctor_set(x_69, 2, x_67); +x_70 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__37; +lean_inc(x_2); +x_71 = lean_name_mk_string(x_2, x_70); +x_72 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__5; +x_73 = lean_name_mk_string(x_3, x_72); +x_74 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__39; +lean_inc(x_73); +x_75 = lean_name_mk_string(x_73, x_74); +x_76 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__41; lean_inc(x_47); -lean_inc(x_76); +x_77 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_77, 0, x_47); +lean_ctor_set(x_77, 1, x_76); +x_78 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__12; +x_79 = lean_name_mk_string(x_4, x_78); lean_inc(x_50); -x_77 = l_Lean_addMacroScope(x_50, x_76, x_47); -x_78 = lean_box(0); -lean_inc(x_76); -lean_ctor_set(x_32, 1, x_78); -lean_ctor_set(x_32, 0, x_76); -x_79 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_79, 0, x_32); -lean_ctor_set(x_79, 1, x_78); -x_80 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__43; -lean_inc(x_44); -x_81 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_81, 0, x_44); -lean_ctor_set(x_81, 1, x_80); -lean_ctor_set(x_81, 2, x_77); -lean_ctor_set(x_81, 3, x_79); -x_82 = lean_array_push(x_61, x_74); -x_83 = lean_array_push(x_82, x_81); -x_84 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_84, 0, x_65); -lean_ctor_set(x_84, 1, x_72); -lean_ctor_set(x_84, 2, x_83); -x_85 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; +lean_inc(x_79); +lean_inc(x_53); +x_80 = l_Lean_addMacroScope(x_53, x_79, x_50); +x_81 = lean_box(0); +lean_inc(x_79); +lean_ctor_set(x_34, 1, x_81); +lean_ctor_set(x_34, 0, x_79); +x_82 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_82, 0, x_34); +lean_ctor_set(x_82, 1, x_81); +x_83 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__44; +lean_inc(x_47); +x_84 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_84, 0, x_47); +lean_ctor_set(x_84, 1, x_83); +lean_ctor_set(x_84, 2, x_80); +lean_ctor_set(x_84, 3, x_82); +x_85 = lean_array_push(x_64, x_77); x_86 = lean_array_push(x_85, x_84); -x_87 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_88 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_88, 0, x_65); -lean_ctor_set(x_88, 1, x_87); -lean_ctor_set(x_88, 2, x_86); -x_89 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__46; -x_90 = lean_array_push(x_89, x_88); +x_87 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_87, 0, x_68); +lean_ctor_set(x_87, 1, x_75); +lean_ctor_set(x_87, 2, x_86); +x_88 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; +x_89 = lean_array_push(x_88, x_87); +x_90 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; x_91 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_91, 0, x_65); -lean_ctor_set(x_91, 1, x_68); -lean_ctor_set(x_91, 2, x_90); +lean_ctor_set(x_91, 0, x_68); +lean_ctor_set(x_91, 1, x_90); +lean_ctor_set(x_91, 2, x_89); x_92 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__47; -x_93 = lean_name_mk_string(x_2, x_92); -x_94 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__49; -lean_inc(x_44); -x_95 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_95, 0, x_44); -lean_ctor_set(x_95, 1, x_94); -x_96 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__7; -lean_inc(x_70); -x_97 = lean_name_mk_string(x_70, x_96); -x_98 = l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__3___closed__7; -x_99 = l_Lean_addMacroScope(x_50, x_98, x_47); -x_100 = l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__3___closed__6; -x_101 = lean_name_mk_string(x_76, x_100); -x_102 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_102, 0, x_101); -lean_ctor_set(x_102, 1, x_78); -x_103 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_103, 0, x_102); -lean_ctor_set(x_103, 1, x_78); -x_104 = l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__3___closed__5; -x_105 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_105, 0, x_44); -lean_ctor_set(x_105, 1, x_104); -lean_ctor_set(x_105, 2, x_99); -lean_ctor_set(x_105, 3, x_103); -x_106 = 1; -x_107 = l_Lean_Name_toString(x_41, x_106); -x_108 = l_Lean_Syntax_mkStrLit(x_107, x_65); -lean_dec(x_107); -lean_inc(x_42); -x_109 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_78, x_42); -x_110 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__28; -x_111 = lean_array_push(x_110, x_108); -x_112 = lean_array_push(x_61, x_105); -x_113 = lean_array_push(x_110, x_95); -x_114 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__84; -x_115 = lean_array_push(x_114, x_58); -x_116 = lean_array_push(x_115, x_66); -x_117 = lean_array_push(x_116, x_91); +x_93 = lean_array_push(x_92, x_91); +x_94 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_94, 0, x_68); +lean_ctor_set(x_94, 1, x_71); +lean_ctor_set(x_94, 2, x_93); +x_95 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__48; +x_96 = lean_name_mk_string(x_2, x_95); +x_97 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__50; +lean_inc(x_47); +x_98 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_98, 0, x_47); +lean_ctor_set(x_98, 1, x_97); +x_99 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__7; +lean_inc(x_73); +x_100 = lean_name_mk_string(x_73, x_99); +x_101 = l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__3___closed__7; +x_102 = l_Lean_addMacroScope(x_53, x_101, x_50); +x_103 = l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__3___closed__6; +x_104 = lean_name_mk_string(x_79, x_103); +lean_ctor_set(x_32, 1, x_81); +lean_ctor_set(x_32, 0, x_104); +x_105 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_105, 0, x_32); +lean_ctor_set(x_105, 1, x_81); +x_106 = l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__3___closed__5; +x_107 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_107, 0, x_47); +lean_ctor_set(x_107, 1, x_106); +lean_ctor_set(x_107, 2, x_102); +lean_ctor_set(x_107, 3, x_105); +x_108 = 1; +x_109 = l_Lean_Name_toString(x_44, x_108); +x_110 = l_Lean_Syntax_mkStrLit(x_109, x_68); +lean_dec(x_109); +lean_inc(x_45); +x_111 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_81, x_45); +x_112 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__28; +x_113 = lean_array_push(x_112, x_110); +x_114 = lean_array_push(x_64, x_107); +x_115 = lean_array_push(x_112, x_98); +x_116 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__85; +x_117 = lean_array_push(x_116, x_61); +x_118 = lean_array_push(x_117, x_69); +x_119 = lean_array_push(x_118, x_94); if (lean_obj_tag(x_7) == 0) { -lean_object* x_178; -x_178 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__11; -x_118 = x_178; -goto block_177; +lean_object* x_179; +x_179 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__3; +x_120 = x_179; +goto block_178; } else { -lean_object* x_179; lean_object* x_180; -x_179 = lean_ctor_get(x_7, 0); -lean_inc(x_179); +lean_object* x_180; lean_object* x_181; +x_180 = lean_ctor_get(x_7, 0); +lean_inc(x_180); lean_dec(x_7); -x_180 = lean_array_push(x_85, x_179); -x_118 = x_180; -goto block_177; -} -block_177: +x_181 = lean_array_push(x_88, x_180); +x_120 = x_181; +goto block_178; +} +block_178: +{ +lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; +x_121 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__3; +x_122 = l_Array_append___rarg(x_121, x_120); +x_123 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_123, 0, x_68); +lean_ctor_set(x_123, 1, x_90); +lean_ctor_set(x_123, 2, x_122); +x_124 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31; +x_125 = lean_array_push(x_124, x_123); +x_126 = lean_array_push(x_125, x_66); +x_127 = lean_array_push(x_126, x_66); +x_128 = lean_array_push(x_127, x_66); +x_129 = lean_array_push(x_128, x_66); +x_130 = lean_array_push(x_129, x_66); +x_131 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_131, 0, x_68); +lean_ctor_set(x_131, 1, x_58); +lean_ctor_set(x_131, 2, x_130); +x_132 = lean_array_push(x_64, x_131); +if (lean_obj_tag(x_111) == 0) { -lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; -x_119 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__11; -x_120 = l_Array_append___rarg(x_119, x_118); -x_121 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_121, 0, x_65); -lean_ctor_set(x_121, 1, x_87); -lean_ctor_set(x_121, 2, x_120); -x_122 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__32; -x_123 = lean_array_push(x_122, x_121); -x_124 = lean_array_push(x_123, x_63); -x_125 = lean_array_push(x_124, x_63); -x_126 = lean_array_push(x_125, x_63); -x_127 = lean_array_push(x_126, x_63); -x_128 = lean_array_push(x_127, x_63); -x_129 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_129, 0, x_65); -lean_ctor_set(x_129, 1, x_55); -lean_ctor_set(x_129, 2, x_128); -x_130 = lean_array_push(x_61, x_129); -if (lean_obj_tag(x_109) == 0) -{ -lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; -lean_dec(x_70); -x_131 = l___private_Init_Meta_0__Lean_quoteNameMk(x_42); -x_132 = lean_array_push(x_111, x_131); -x_133 = lean_array_push(x_132, x_35); -x_134 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_134, 0, x_65); -lean_ctor_set(x_134, 1, x_87); -lean_ctor_set(x_134, 2, x_133); -x_135 = lean_array_push(x_112, x_134); +lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; +lean_dec(x_73); +x_133 = l_Lean_quoteNameMk(x_45); +x_134 = lean_array_push(x_113, x_133); +x_135 = lean_array_push(x_134, x_38); x_136 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_136, 0, x_65); -lean_ctor_set(x_136, 1, x_97); +lean_ctor_set(x_136, 0, x_68); +lean_ctor_set(x_136, 1, x_90); lean_ctor_set(x_136, 2, x_135); -x_137 = lean_array_push(x_113, x_136); -x_138 = lean_array_push(x_137, x_63); -x_139 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_139, 0, x_65); -lean_ctor_set(x_139, 1, x_93); -lean_ctor_set(x_139, 2, x_138); -x_140 = lean_array_push(x_117, x_139); -x_141 = lean_array_push(x_140, x_63); -x_142 = lean_array_push(x_141, x_63); -x_143 = lean_array_push(x_142, x_63); -x_144 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_144, 0, x_65); -lean_ctor_set(x_144, 1, x_57); -lean_ctor_set(x_144, 2, x_143); -x_145 = lean_array_push(x_130, x_144); +x_137 = lean_array_push(x_114, x_136); +x_138 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_138, 0, x_68); +lean_ctor_set(x_138, 1, x_100); +lean_ctor_set(x_138, 2, x_137); +x_139 = lean_array_push(x_115, x_138); +x_140 = lean_array_push(x_139, x_66); +x_141 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_141, 0, x_68); +lean_ctor_set(x_141, 1, x_96); +lean_ctor_set(x_141, 2, x_140); +x_142 = lean_array_push(x_119, x_141); +x_143 = lean_array_push(x_142, x_66); +x_144 = lean_array_push(x_143, x_66); +x_145 = lean_array_push(x_144, x_66); x_146 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_146, 0, x_65); -lean_ctor_set(x_146, 1, x_53); +lean_ctor_set(x_146, 0, x_68); +lean_ctor_set(x_146, 1, x_60); lean_ctor_set(x_146, 2, x_145); -lean_inc(x_146); -x_147 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_147, 0, x_146); -x_148 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_5, x_146, x_147, x_8, x_9, x_51); -return x_148; +x_147 = lean_array_push(x_132, x_146); +x_148 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_148, 0, x_68); +lean_ctor_set(x_148, 1, x_56); +lean_ctor_set(x_148, 2, x_147); +lean_inc(x_148); +x_149 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_149, 0, x_148); +x_150 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_5, x_148, x_149, x_8, x_9, x_54); +return x_150; } -else -{ -lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; -lean_dec(x_42); -x_149 = lean_ctor_get(x_109, 0); -lean_inc(x_149); -lean_dec(x_109); -x_150 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__22; -x_151 = lean_name_mk_string(x_70, x_150); -x_152 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__8; -x_153 = l_String_intercalate(x_152, x_149); -x_154 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__9; -x_155 = lean_string_append(x_154, x_153); -lean_dec(x_153); -x_156 = l_Lean_nameLitKind; -x_157 = l_Lean_Syntax_mkLit(x_156, x_155, x_65); -x_158 = lean_array_push(x_85, x_157); -x_159 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_159, 0, x_65); -lean_ctor_set(x_159, 1, x_151); -lean_ctor_set(x_159, 2, x_158); -x_160 = lean_array_push(x_111, x_159); -x_161 = lean_array_push(x_160, x_35); -x_162 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_162, 0, x_65); -lean_ctor_set(x_162, 1, x_87); -lean_ctor_set(x_162, 2, x_161); -x_163 = lean_array_push(x_112, x_162); -x_164 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_164, 0, x_65); -lean_ctor_set(x_164, 1, x_97); -lean_ctor_set(x_164, 2, x_163); -x_165 = lean_array_push(x_113, x_164); -x_166 = lean_array_push(x_165, x_63); -x_167 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_167, 0, x_65); -lean_ctor_set(x_167, 1, x_93); -lean_ctor_set(x_167, 2, x_166); -x_168 = lean_array_push(x_117, x_167); -x_169 = lean_array_push(x_168, x_63); -x_170 = lean_array_push(x_169, x_63); -x_171 = lean_array_push(x_170, x_63); -x_172 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_172, 0, x_65); -lean_ctor_set(x_172, 1, x_57); -lean_ctor_set(x_172, 2, x_171); -x_173 = lean_array_push(x_130, x_172); -x_174 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_174, 0, x_65); -lean_ctor_set(x_174, 1, x_53); -lean_ctor_set(x_174, 2, x_173); -lean_inc(x_174); -x_175 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_175, 0, x_174); -x_176 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_5, x_174, x_175, x_8, x_9, x_51); -return x_176; +else +{ +lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; +lean_dec(x_45); +x_151 = lean_ctor_get(x_111, 0); +lean_inc(x_151); +lean_dec(x_111); +x_152 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__22; +x_153 = lean_name_mk_string(x_73, x_152); +x_154 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__16; +x_155 = l_String_intercalate(x_154, x_151); +x_156 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__17; +x_157 = lean_string_append(x_156, x_155); +lean_dec(x_155); +x_158 = l_Lean_Syntax_mkNameLit(x_157, x_68); +x_159 = lean_array_push(x_88, x_158); +x_160 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_160, 0, x_68); +lean_ctor_set(x_160, 1, x_153); +lean_ctor_set(x_160, 2, x_159); +x_161 = lean_array_push(x_113, x_160); +x_162 = lean_array_push(x_161, x_38); +x_163 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_163, 0, x_68); +lean_ctor_set(x_163, 1, x_90); +lean_ctor_set(x_163, 2, x_162); +x_164 = lean_array_push(x_114, x_163); +x_165 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_165, 0, x_68); +lean_ctor_set(x_165, 1, x_100); +lean_ctor_set(x_165, 2, x_164); +x_166 = lean_array_push(x_115, x_165); +x_167 = lean_array_push(x_166, x_66); +x_168 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_168, 0, x_68); +lean_ctor_set(x_168, 1, x_96); +lean_ctor_set(x_168, 2, x_167); +x_169 = lean_array_push(x_119, x_168); +x_170 = lean_array_push(x_169, x_66); +x_171 = lean_array_push(x_170, x_66); +x_172 = lean_array_push(x_171, x_66); +x_173 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_173, 0, x_68); +lean_ctor_set(x_173, 1, x_60); +lean_ctor_set(x_173, 2, x_172); +x_174 = lean_array_push(x_132, x_173); +x_175 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_175, 0, x_68); +lean_ctor_set(x_175, 1, x_56); +lean_ctor_set(x_175, 2, x_174); +lean_inc(x_175); +x_176 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_176, 0, x_175); +x_177 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_5, x_175, x_176, x_8, x_9, x_54); +return x_177; } } } else { -lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; uint8_t x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; -x_181 = lean_ctor_get(x_32, 0); -lean_inc(x_181); -lean_dec(x_32); -x_182 = l_Lean_Elab_Command_getScope___rarg(x_9, x_33); -x_183 = lean_ctor_get(x_182, 0); -lean_inc(x_183); -x_184 = lean_ctor_get(x_182, 1); +lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; uint8_t x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; +x_182 = lean_ctor_get(x_34, 0); +lean_inc(x_182); +lean_dec(x_34); +x_183 = l_Lean_Elab_Command_getScope___rarg(x_9, x_36); +x_184 = lean_ctor_get(x_183, 0); lean_inc(x_184); -lean_dec(x_182); -x_185 = lean_ctor_get(x_183, 2); +x_185 = lean_ctor_get(x_183, 1); lean_inc(x_185); lean_dec(x_183); -x_186 = l_Lean_Syntax_getId(x_12); +x_186 = lean_ctor_get(x_184, 2); lean_inc(x_186); -x_187 = l_Lean_Name_append(x_185, x_186); -lean_dec(x_185); -x_188 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_8, x_9, x_184); -x_189 = lean_ctor_get(x_188, 0); -lean_inc(x_189); -x_190 = lean_ctor_get(x_188, 1); +lean_dec(x_184); +x_187 = l_Lean_Syntax_getId(x_12); +lean_inc(x_187); +x_188 = l_Lean_Name_append(x_186, x_187); +lean_dec(x_186); +x_189 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_8, x_9, x_185); +x_190 = lean_ctor_get(x_189, 0); lean_inc(x_190); -lean_dec(x_188); -x_191 = l_Lean_Elab_Command_getCurrMacroScope(x_8, x_9, x_190); -x_192 = lean_ctor_get(x_191, 0); -lean_inc(x_192); -x_193 = lean_ctor_get(x_191, 1); +x_191 = lean_ctor_get(x_189, 1); +lean_inc(x_191); +lean_dec(x_189); +x_192 = l_Lean_Elab_Command_getCurrMacroScope(x_8, x_9, x_191); +x_193 = lean_ctor_get(x_192, 0); lean_inc(x_193); -lean_dec(x_191); -x_194 = l_Lean_Elab_Command_getMainModule___rarg(x_9, x_193); -x_195 = lean_ctor_get(x_194, 0); -lean_inc(x_195); -x_196 = lean_ctor_get(x_194, 1); +x_194 = lean_ctor_get(x_192, 1); +lean_inc(x_194); +lean_dec(x_192); +x_195 = l_Lean_Elab_Command_getMainModule___rarg(x_9, x_194); +x_196 = lean_ctor_get(x_195, 0); lean_inc(x_196); -lean_dec(x_194); -x_197 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; +x_197 = lean_ctor_get(x_195, 1); +lean_inc(x_197); +lean_dec(x_195); +x_198 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; lean_inc(x_2); -x_198 = lean_name_mk_string(x_2, x_197); -x_199 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; +x_199 = lean_name_mk_string(x_2, x_198); +x_200 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; lean_inc(x_2); -x_200 = lean_name_mk_string(x_2, x_199); -x_201 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__34; +x_201 = lean_name_mk_string(x_2, x_200); +x_202 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33; lean_inc(x_2); -x_202 = lean_name_mk_string(x_2, x_201); -lean_inc(x_189); -x_203 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_203, 0, x_189); -lean_ctor_set(x_203, 1, x_201); -x_204 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__1; +x_203 = lean_name_mk_string(x_2, x_202); +lean_inc(x_190); +x_204 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_204, 0, x_190); +lean_ctor_set(x_204, 1, x_202); +x_205 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__35; lean_inc(x_2); -x_205 = lean_name_mk_string(x_2, x_204); -x_206 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; -x_207 = lean_array_push(x_206, x_12); -x_208 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__12; -x_209 = lean_array_push(x_207, x_208); -x_210 = lean_box(2); -x_211 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_211, 0, x_210); -lean_ctor_set(x_211, 1, x_205); -lean_ctor_set(x_211, 2, x_209); -x_212 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__36; +x_206 = lean_name_mk_string(x_2, x_205); +x_207 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; +x_208 = lean_array_push(x_207, x_12); +x_209 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__11; +x_210 = lean_array_push(x_208, x_209); +x_211 = lean_box(2); +x_212 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_212, 0, x_211); +lean_ctor_set(x_212, 1, x_206); +lean_ctor_set(x_212, 2, x_210); +x_213 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__37; lean_inc(x_2); -x_213 = lean_name_mk_string(x_2, x_212); -x_214 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__5; -x_215 = lean_name_mk_string(x_3, x_214); -x_216 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__38; -lean_inc(x_215); -x_217 = lean_name_mk_string(x_215, x_216); -x_218 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__40; -lean_inc(x_189); -x_219 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_219, 0, x_189); -lean_ctor_set(x_219, 1, x_218); -x_220 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__12; -x_221 = lean_name_mk_string(x_4, x_220); -lean_inc(x_192); -lean_inc(x_221); -lean_inc(x_195); -x_222 = l_Lean_addMacroScope(x_195, x_221, x_192); -x_223 = lean_box(0); -lean_inc(x_221); -x_224 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_224, 0, x_221); -lean_ctor_set(x_224, 1, x_223); -x_225 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_225, 0, x_224); -lean_ctor_set(x_225, 1, x_223); -x_226 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__43; -lean_inc(x_189); -x_227 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_227, 0, x_189); -lean_ctor_set(x_227, 1, x_226); -lean_ctor_set(x_227, 2, x_222); -lean_ctor_set(x_227, 3, x_225); -x_228 = lean_array_push(x_206, x_219); -x_229 = lean_array_push(x_228, x_227); -x_230 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_230, 0, x_210); -lean_ctor_set(x_230, 1, x_217); -lean_ctor_set(x_230, 2, x_229); -x_231 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; -x_232 = lean_array_push(x_231, x_230); -x_233 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; -x_234 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_234, 0, x_210); -lean_ctor_set(x_234, 1, x_233); -lean_ctor_set(x_234, 2, x_232); -x_235 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__46; -x_236 = lean_array_push(x_235, x_234); -x_237 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_237, 0, x_210); -lean_ctor_set(x_237, 1, x_213); -lean_ctor_set(x_237, 2, x_236); -x_238 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__47; -x_239 = lean_name_mk_string(x_2, x_238); -x_240 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__49; -lean_inc(x_189); -x_241 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_241, 0, x_189); -lean_ctor_set(x_241, 1, x_240); -x_242 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__7; -lean_inc(x_215); -x_243 = lean_name_mk_string(x_215, x_242); -x_244 = l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__3___closed__7; -x_245 = l_Lean_addMacroScope(x_195, x_244, x_192); -x_246 = l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__3___closed__6; -x_247 = lean_name_mk_string(x_221, x_246); -x_248 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_248, 0, x_247); -lean_ctor_set(x_248, 1, x_223); +x_214 = lean_name_mk_string(x_2, x_213); +x_215 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__5; +x_216 = lean_name_mk_string(x_3, x_215); +x_217 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__39; +lean_inc(x_216); +x_218 = lean_name_mk_string(x_216, x_217); +x_219 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__41; +lean_inc(x_190); +x_220 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_220, 0, x_190); +lean_ctor_set(x_220, 1, x_219); +x_221 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__12; +x_222 = lean_name_mk_string(x_4, x_221); +lean_inc(x_193); +lean_inc(x_222); +lean_inc(x_196); +x_223 = l_Lean_addMacroScope(x_196, x_222, x_193); +x_224 = lean_box(0); +lean_inc(x_222); +x_225 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_225, 0, x_222); +lean_ctor_set(x_225, 1, x_224); +x_226 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_226, 0, x_225); +lean_ctor_set(x_226, 1, x_224); +x_227 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__44; +lean_inc(x_190); +x_228 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_228, 0, x_190); +lean_ctor_set(x_228, 1, x_227); +lean_ctor_set(x_228, 2, x_223); +lean_ctor_set(x_228, 3, x_226); +x_229 = lean_array_push(x_207, x_220); +x_230 = lean_array_push(x_229, x_228); +x_231 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_231, 0, x_211); +lean_ctor_set(x_231, 1, x_218); +lean_ctor_set(x_231, 2, x_230); +x_232 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; +x_233 = lean_array_push(x_232, x_231); +x_234 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; +x_235 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_235, 0, x_211); +lean_ctor_set(x_235, 1, x_234); +lean_ctor_set(x_235, 2, x_233); +x_236 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__47; +x_237 = lean_array_push(x_236, x_235); +x_238 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_238, 0, x_211); +lean_ctor_set(x_238, 1, x_214); +lean_ctor_set(x_238, 2, x_237); +x_239 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__48; +x_240 = lean_name_mk_string(x_2, x_239); +x_241 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__50; +lean_inc(x_190); +x_242 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_242, 0, x_190); +lean_ctor_set(x_242, 1, x_241); +x_243 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__7; +lean_inc(x_216); +x_244 = lean_name_mk_string(x_216, x_243); +x_245 = l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__3___closed__7; +x_246 = l_Lean_addMacroScope(x_196, x_245, x_193); +x_247 = l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__3___closed__6; +x_248 = lean_name_mk_string(x_222, x_247); +lean_ctor_set(x_32, 1, x_224); +lean_ctor_set(x_32, 0, x_248); x_249 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_249, 0, x_248); -lean_ctor_set(x_249, 1, x_223); +lean_ctor_set(x_249, 0, x_32); +lean_ctor_set(x_249, 1, x_224); x_250 = l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__3___closed__5; x_251 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_251, 0, x_189); +lean_ctor_set(x_251, 0, x_190); lean_ctor_set(x_251, 1, x_250); -lean_ctor_set(x_251, 2, x_245); +lean_ctor_set(x_251, 2, x_246); lean_ctor_set(x_251, 3, x_249); x_252 = 1; -x_253 = l_Lean_Name_toString(x_186, x_252); -x_254 = l_Lean_Syntax_mkStrLit(x_253, x_210); +x_253 = l_Lean_Name_toString(x_187, x_252); +x_254 = l_Lean_Syntax_mkStrLit(x_253, x_211); lean_dec(x_253); -lean_inc(x_187); -x_255 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_223, x_187); +lean_inc(x_188); +x_255 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_224, x_188); x_256 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__28; x_257 = lean_array_push(x_256, x_254); -x_258 = lean_array_push(x_206, x_251); -x_259 = lean_array_push(x_256, x_241); -x_260 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__84; -x_261 = lean_array_push(x_260, x_203); -x_262 = lean_array_push(x_261, x_211); -x_263 = lean_array_push(x_262, x_237); +x_258 = lean_array_push(x_207, x_251); +x_259 = lean_array_push(x_256, x_242); +x_260 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__85; +x_261 = lean_array_push(x_260, x_204); +x_262 = lean_array_push(x_261, x_212); +x_263 = lean_array_push(x_262, x_238); if (lean_obj_tag(x_7) == 0) { -lean_object* x_324; -x_324 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__11; -x_264 = x_324; -goto block_323; +lean_object* x_323; +x_323 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__3; +x_264 = x_323; +goto block_322; } else { -lean_object* x_325; lean_object* x_326; -x_325 = lean_ctor_get(x_7, 0); -lean_inc(x_325); +lean_object* x_324; lean_object* x_325; +x_324 = lean_ctor_get(x_7, 0); +lean_inc(x_324); lean_dec(x_7); -x_326 = lean_array_push(x_231, x_325); -x_264 = x_326; -goto block_323; +x_325 = lean_array_push(x_232, x_324); +x_264 = x_325; +goto block_322; } -block_323: +block_322: { lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; -x_265 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__11; +x_265 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__3; x_266 = l_Array_append___rarg(x_265, x_264); x_267 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_267, 0, x_210); -lean_ctor_set(x_267, 1, x_233); +lean_ctor_set(x_267, 0, x_211); +lean_ctor_set(x_267, 1, x_234); lean_ctor_set(x_267, 2, x_266); -x_268 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__32; +x_268 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31; x_269 = lean_array_push(x_268, x_267); -x_270 = lean_array_push(x_269, x_208); -x_271 = lean_array_push(x_270, x_208); -x_272 = lean_array_push(x_271, x_208); -x_273 = lean_array_push(x_272, x_208); -x_274 = lean_array_push(x_273, x_208); +x_270 = lean_array_push(x_269, x_209); +x_271 = lean_array_push(x_270, x_209); +x_272 = lean_array_push(x_271, x_209); +x_273 = lean_array_push(x_272, x_209); +x_274 = lean_array_push(x_273, x_209); x_275 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_275, 0, x_210); -lean_ctor_set(x_275, 1, x_200); +lean_ctor_set(x_275, 0, x_211); +lean_ctor_set(x_275, 1, x_201); lean_ctor_set(x_275, 2, x_274); -x_276 = lean_array_push(x_206, x_275); +x_276 = lean_array_push(x_207, x_275); if (lean_obj_tag(x_255) == 0) { lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; -lean_dec(x_215); -x_277 = l___private_Init_Meta_0__Lean_quoteNameMk(x_187); +lean_dec(x_216); +x_277 = l_Lean_quoteNameMk(x_188); x_278 = lean_array_push(x_257, x_277); -x_279 = lean_array_push(x_278, x_181); +x_279 = lean_array_push(x_278, x_182); x_280 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_280, 0, x_210); -lean_ctor_set(x_280, 1, x_233); +lean_ctor_set(x_280, 0, x_211); +lean_ctor_set(x_280, 1, x_234); lean_ctor_set(x_280, 2, x_279); x_281 = lean_array_push(x_258, x_280); x_282 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_282, 0, x_210); -lean_ctor_set(x_282, 1, x_243); +lean_ctor_set(x_282, 0, x_211); +lean_ctor_set(x_282, 1, x_244); lean_ctor_set(x_282, 2, x_281); x_283 = lean_array_push(x_259, x_282); -x_284 = lean_array_push(x_283, x_208); +x_284 = lean_array_push(x_283, x_209); x_285 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_285, 0, x_210); -lean_ctor_set(x_285, 1, x_239); +lean_ctor_set(x_285, 0, x_211); +lean_ctor_set(x_285, 1, x_240); lean_ctor_set(x_285, 2, x_284); x_286 = lean_array_push(x_263, x_285); -x_287 = lean_array_push(x_286, x_208); -x_288 = lean_array_push(x_287, x_208); -x_289 = lean_array_push(x_288, x_208); +x_287 = lean_array_push(x_286, x_209); +x_288 = lean_array_push(x_287, x_209); +x_289 = lean_array_push(x_288, x_209); x_290 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_290, 0, x_210); -lean_ctor_set(x_290, 1, x_202); +lean_ctor_set(x_290, 0, x_211); +lean_ctor_set(x_290, 1, x_203); lean_ctor_set(x_290, 2, x_289); x_291 = lean_array_push(x_276, x_290); x_292 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_292, 0, x_210); -lean_ctor_set(x_292, 1, x_198); +lean_ctor_set(x_292, 0, x_211); +lean_ctor_set(x_292, 1, x_199); lean_ctor_set(x_292, 2, x_291); lean_inc(x_292); x_293 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); lean_closure_set(x_293, 0, x_292); -x_294 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_5, x_292, x_293, x_8, x_9, x_196); +x_294 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_5, x_292, x_293, x_8, x_9, x_197); return x_294; } else { -lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; -lean_dec(x_187); +lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; +lean_dec(x_188); x_295 = lean_ctor_get(x_255, 0); lean_inc(x_295); lean_dec(x_255); x_296 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__22; -x_297 = lean_name_mk_string(x_215, x_296); -x_298 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__8; +x_297 = lean_name_mk_string(x_216, x_296); +x_298 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__16; x_299 = l_String_intercalate(x_298, x_295); -x_300 = l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__9; +x_300 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__17; x_301 = lean_string_append(x_300, x_299); lean_dec(x_299); -x_302 = l_Lean_nameLitKind; -x_303 = l_Lean_Syntax_mkLit(x_302, x_301, x_210); -x_304 = lean_array_push(x_231, x_303); -x_305 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_305, 0, x_210); -lean_ctor_set(x_305, 1, x_297); -lean_ctor_set(x_305, 2, x_304); -x_306 = lean_array_push(x_257, x_305); -x_307 = lean_array_push(x_306, x_181); -x_308 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_308, 0, x_210); -lean_ctor_set(x_308, 1, x_233); -lean_ctor_set(x_308, 2, x_307); -x_309 = lean_array_push(x_258, x_308); -x_310 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_310, 0, x_210); -lean_ctor_set(x_310, 1, x_243); -lean_ctor_set(x_310, 2, x_309); -x_311 = lean_array_push(x_259, x_310); -x_312 = lean_array_push(x_311, x_208); -x_313 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_313, 0, x_210); -lean_ctor_set(x_313, 1, x_239); -lean_ctor_set(x_313, 2, x_312); -x_314 = lean_array_push(x_263, x_313); -x_315 = lean_array_push(x_314, x_208); -x_316 = lean_array_push(x_315, x_208); -x_317 = lean_array_push(x_316, x_208); -x_318 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_318, 0, x_210); -lean_ctor_set(x_318, 1, x_202); -lean_ctor_set(x_318, 2, x_317); -x_319 = lean_array_push(x_276, x_318); -x_320 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_320, 0, x_210); -lean_ctor_set(x_320, 1, x_198); -lean_ctor_set(x_320, 2, x_319); -lean_inc(x_320); -x_321 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_321, 0, x_320); -x_322 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_5, x_320, x_321, x_8, x_9, x_196); -return x_322; -} -} -} -} -else -{ -uint8_t x_327; +x_302 = l_Lean_Syntax_mkNameLit(x_301, x_211); +x_303 = lean_array_push(x_232, x_302); +x_304 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_304, 0, x_211); +lean_ctor_set(x_304, 1, x_297); +lean_ctor_set(x_304, 2, x_303); +x_305 = lean_array_push(x_257, x_304); +x_306 = lean_array_push(x_305, x_182); +x_307 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_307, 0, x_211); +lean_ctor_set(x_307, 1, x_234); +lean_ctor_set(x_307, 2, x_306); +x_308 = lean_array_push(x_258, x_307); +x_309 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_309, 0, x_211); +lean_ctor_set(x_309, 1, x_244); +lean_ctor_set(x_309, 2, x_308); +x_310 = lean_array_push(x_259, x_309); +x_311 = lean_array_push(x_310, x_209); +x_312 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_312, 0, x_211); +lean_ctor_set(x_312, 1, x_240); +lean_ctor_set(x_312, 2, x_311); +x_313 = lean_array_push(x_263, x_312); +x_314 = lean_array_push(x_313, x_209); +x_315 = lean_array_push(x_314, x_209); +x_316 = lean_array_push(x_315, x_209); +x_317 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_317, 0, x_211); +lean_ctor_set(x_317, 1, x_203); +lean_ctor_set(x_317, 2, x_316); +x_318 = lean_array_push(x_276, x_317); +x_319 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_319, 0, x_211); +lean_ctor_set(x_319, 1, x_199); +lean_ctor_set(x_319, 2, x_318); +lean_inc(x_319); +x_320 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_320, 0, x_319); +x_321 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_5, x_319, x_320, x_8, x_9, x_197); +return x_321; +} +} +} +} +else +{ +lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; uint8_t x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; +x_326 = lean_ctor_get(x_32, 0); +lean_inc(x_326); +lean_dec(x_32); +x_327 = lean_ctor_get(x_31, 1); +lean_inc(x_327); +lean_dec(x_31); +x_328 = lean_ctor_get(x_326, 0); +lean_inc(x_328); +if (lean_is_exclusive(x_326)) { + lean_ctor_release(x_326, 0); + lean_ctor_release(x_326, 1); + x_329 = x_326; +} else { + lean_dec_ref(x_326); + x_329 = lean_box(0); +} +x_330 = l_Lean_Elab_Command_getScope___rarg(x_9, x_327); +x_331 = lean_ctor_get(x_330, 0); +lean_inc(x_331); +x_332 = lean_ctor_get(x_330, 1); +lean_inc(x_332); +lean_dec(x_330); +x_333 = lean_ctor_get(x_331, 2); +lean_inc(x_333); +lean_dec(x_331); +x_334 = l_Lean_Syntax_getId(x_12); +lean_inc(x_334); +x_335 = l_Lean_Name_append(x_333, x_334); +lean_dec(x_333); +x_336 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_8, x_9, x_332); +x_337 = lean_ctor_get(x_336, 0); +lean_inc(x_337); +x_338 = lean_ctor_get(x_336, 1); +lean_inc(x_338); +lean_dec(x_336); +x_339 = l_Lean_Elab_Command_getCurrMacroScope(x_8, x_9, x_338); +x_340 = lean_ctor_get(x_339, 0); +lean_inc(x_340); +x_341 = lean_ctor_get(x_339, 1); +lean_inc(x_341); +lean_dec(x_339); +x_342 = l_Lean_Elab_Command_getMainModule___rarg(x_9, x_341); +x_343 = lean_ctor_get(x_342, 0); +lean_inc(x_343); +x_344 = lean_ctor_get(x_342, 1); +lean_inc(x_344); +lean_dec(x_342); +x_345 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; +lean_inc(x_2); +x_346 = lean_name_mk_string(x_2, x_345); +x_347 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; +lean_inc(x_2); +x_348 = lean_name_mk_string(x_2, x_347); +x_349 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33; +lean_inc(x_2); +x_350 = lean_name_mk_string(x_2, x_349); +lean_inc(x_337); +x_351 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_351, 0, x_337); +lean_ctor_set(x_351, 1, x_349); +x_352 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__35; +lean_inc(x_2); +x_353 = lean_name_mk_string(x_2, x_352); +x_354 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; +x_355 = lean_array_push(x_354, x_12); +x_356 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__11; +x_357 = lean_array_push(x_355, x_356); +x_358 = lean_box(2); +x_359 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_359, 0, x_358); +lean_ctor_set(x_359, 1, x_353); +lean_ctor_set(x_359, 2, x_357); +x_360 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__37; +lean_inc(x_2); +x_361 = lean_name_mk_string(x_2, x_360); +x_362 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__5; +x_363 = lean_name_mk_string(x_3, x_362); +x_364 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__39; +lean_inc(x_363); +x_365 = lean_name_mk_string(x_363, x_364); +x_366 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__41; +lean_inc(x_337); +x_367 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_367, 0, x_337); +lean_ctor_set(x_367, 1, x_366); +x_368 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__12; +x_369 = lean_name_mk_string(x_4, x_368); +lean_inc(x_340); +lean_inc(x_369); +lean_inc(x_343); +x_370 = l_Lean_addMacroScope(x_343, x_369, x_340); +x_371 = lean_box(0); +lean_inc(x_369); +if (lean_is_scalar(x_329)) { + x_372 = lean_alloc_ctor(0, 2, 0); +} else { + x_372 = x_329; +} +lean_ctor_set(x_372, 0, x_369); +lean_ctor_set(x_372, 1, x_371); +x_373 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_373, 0, x_372); +lean_ctor_set(x_373, 1, x_371); +x_374 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__44; +lean_inc(x_337); +x_375 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_375, 0, x_337); +lean_ctor_set(x_375, 1, x_374); +lean_ctor_set(x_375, 2, x_370); +lean_ctor_set(x_375, 3, x_373); +x_376 = lean_array_push(x_354, x_367); +x_377 = lean_array_push(x_376, x_375); +x_378 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_378, 0, x_358); +lean_ctor_set(x_378, 1, x_365); +lean_ctor_set(x_378, 2, x_377); +x_379 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__27; +x_380 = lean_array_push(x_379, x_378); +x_381 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; +x_382 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_382, 0, x_358); +lean_ctor_set(x_382, 1, x_381); +lean_ctor_set(x_382, 2, x_380); +x_383 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__47; +x_384 = lean_array_push(x_383, x_382); +x_385 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_385, 0, x_358); +lean_ctor_set(x_385, 1, x_361); +lean_ctor_set(x_385, 2, x_384); +x_386 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__48; +x_387 = lean_name_mk_string(x_2, x_386); +x_388 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__50; +lean_inc(x_337); +x_389 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_389, 0, x_337); +lean_ctor_set(x_389, 1, x_388); +x_390 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__7; +lean_inc(x_363); +x_391 = lean_name_mk_string(x_363, x_390); +x_392 = l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__3___closed__7; +x_393 = l_Lean_addMacroScope(x_343, x_392, x_340); +x_394 = l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__3___closed__6; +x_395 = lean_name_mk_string(x_369, x_394); +x_396 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_396, 0, x_395); +lean_ctor_set(x_396, 1, x_371); +x_397 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_397, 0, x_396); +lean_ctor_set(x_397, 1, x_371); +x_398 = l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__3___closed__5; +x_399 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_399, 0, x_337); +lean_ctor_set(x_399, 1, x_398); +lean_ctor_set(x_399, 2, x_393); +lean_ctor_set(x_399, 3, x_397); +x_400 = 1; +x_401 = l_Lean_Name_toString(x_334, x_400); +x_402 = l_Lean_Syntax_mkStrLit(x_401, x_358); +lean_dec(x_401); +lean_inc(x_335); +x_403 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_371, x_335); +x_404 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__28; +x_405 = lean_array_push(x_404, x_402); +x_406 = lean_array_push(x_354, x_399); +x_407 = lean_array_push(x_404, x_389); +x_408 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__85; +x_409 = lean_array_push(x_408, x_351); +x_410 = lean_array_push(x_409, x_359); +x_411 = lean_array_push(x_410, x_385); +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_471; +x_471 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__3; +x_412 = x_471; +goto block_470; +} +else +{ +lean_object* x_472; lean_object* x_473; +x_472 = lean_ctor_get(x_7, 0); +lean_inc(x_472); +lean_dec(x_7); +x_473 = lean_array_push(x_379, x_472); +x_412 = x_473; +goto block_470; +} +block_470: +{ +lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; +x_413 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__3; +x_414 = l_Array_append___rarg(x_413, x_412); +x_415 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_415, 0, x_358); +lean_ctor_set(x_415, 1, x_381); +lean_ctor_set(x_415, 2, x_414); +x_416 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31; +x_417 = lean_array_push(x_416, x_415); +x_418 = lean_array_push(x_417, x_356); +x_419 = lean_array_push(x_418, x_356); +x_420 = lean_array_push(x_419, x_356); +x_421 = lean_array_push(x_420, x_356); +x_422 = lean_array_push(x_421, x_356); +x_423 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_423, 0, x_358); +lean_ctor_set(x_423, 1, x_348); +lean_ctor_set(x_423, 2, x_422); +x_424 = lean_array_push(x_354, x_423); +if (lean_obj_tag(x_403) == 0) +{ +lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; +lean_dec(x_363); +x_425 = l_Lean_quoteNameMk(x_335); +x_426 = lean_array_push(x_405, x_425); +x_427 = lean_array_push(x_426, x_328); +x_428 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_428, 0, x_358); +lean_ctor_set(x_428, 1, x_381); +lean_ctor_set(x_428, 2, x_427); +x_429 = lean_array_push(x_406, x_428); +x_430 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_430, 0, x_358); +lean_ctor_set(x_430, 1, x_391); +lean_ctor_set(x_430, 2, x_429); +x_431 = lean_array_push(x_407, x_430); +x_432 = lean_array_push(x_431, x_356); +x_433 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_433, 0, x_358); +lean_ctor_set(x_433, 1, x_387); +lean_ctor_set(x_433, 2, x_432); +x_434 = lean_array_push(x_411, x_433); +x_435 = lean_array_push(x_434, x_356); +x_436 = lean_array_push(x_435, x_356); +x_437 = lean_array_push(x_436, x_356); +x_438 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_438, 0, x_358); +lean_ctor_set(x_438, 1, x_350); +lean_ctor_set(x_438, 2, x_437); +x_439 = lean_array_push(x_424, x_438); +x_440 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_440, 0, x_358); +lean_ctor_set(x_440, 1, x_346); +lean_ctor_set(x_440, 2, x_439); +lean_inc(x_440); +x_441 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_441, 0, x_440); +x_442 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_5, x_440, x_441, x_8, x_9, x_344); +return x_442; +} +else +{ +lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; +lean_dec(x_335); +x_443 = lean_ctor_get(x_403, 0); +lean_inc(x_443); +lean_dec(x_403); +x_444 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__22; +x_445 = lean_name_mk_string(x_363, x_444); +x_446 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__16; +x_447 = l_String_intercalate(x_446, x_443); +x_448 = l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__17; +x_449 = lean_string_append(x_448, x_447); +lean_dec(x_447); +x_450 = l_Lean_Syntax_mkNameLit(x_449, x_358); +x_451 = lean_array_push(x_379, x_450); +x_452 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_452, 0, x_358); +lean_ctor_set(x_452, 1, x_445); +lean_ctor_set(x_452, 2, x_451); +x_453 = lean_array_push(x_405, x_452); +x_454 = lean_array_push(x_453, x_328); +x_455 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_455, 0, x_358); +lean_ctor_set(x_455, 1, x_381); +lean_ctor_set(x_455, 2, x_454); +x_456 = lean_array_push(x_406, x_455); +x_457 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_457, 0, x_358); +lean_ctor_set(x_457, 1, x_391); +lean_ctor_set(x_457, 2, x_456); +x_458 = lean_array_push(x_407, x_457); +x_459 = lean_array_push(x_458, x_356); +x_460 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_460, 0, x_358); +lean_ctor_set(x_460, 1, x_387); +lean_ctor_set(x_460, 2, x_459); +x_461 = lean_array_push(x_411, x_460); +x_462 = lean_array_push(x_461, x_356); +x_463 = lean_array_push(x_462, x_356); +x_464 = lean_array_push(x_463, x_356); +x_465 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_465, 0, x_358); +lean_ctor_set(x_465, 1, x_350); +lean_ctor_set(x_465, 2, x_464); +x_466 = lean_array_push(x_424, x_465); +x_467 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_467, 0, x_358); +lean_ctor_set(x_467, 1, x_346); +lean_ctor_set(x_467, 2, x_466); +lean_inc(x_467); +x_468 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_468, 0, x_467); +x_469 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_5, x_467, x_468, x_8, x_9, x_344); +return x_469; +} +} +} +} +else +{ +uint8_t x_474; lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); @@ -19177,23 +20922,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_327 = !lean_is_exclusive(x_31); -if (x_327 == 0) +x_474 = !lean_is_exclusive(x_31); +if (x_474 == 0) { return x_31; } else { -lean_object* x_328; lean_object* x_329; lean_object* x_330; -x_328 = lean_ctor_get(x_31, 0); -x_329 = lean_ctor_get(x_31, 1); -lean_inc(x_329); -lean_inc(x_328); +lean_object* x_475; lean_object* x_476; lean_object* x_477; +x_475 = lean_ctor_get(x_31, 0); +x_476 = lean_ctor_get(x_31, 1); +lean_inc(x_476); +lean_inc(x_475); lean_dec(x_31); -x_330 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_330, 0, x_328); -lean_ctor_set(x_330, 1, x_329); -return x_330; +x_477 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_477, 0, x_475); +lean_ctor_set(x_477, 1, x_476); +return x_477; } } } @@ -19242,68 +20987,67 @@ x_9 = l_Lean_Syntax_getArg(x_1, x_8); x_10 = l_Lean_Syntax_isNone(x_9); if (x_10 == 0) { -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = l_Lean_nullKind; -x_12 = lean_unsigned_to_nat(1u); +lean_object* x_11; uint8_t x_12; +x_11 = lean_unsigned_to_nat(1u); lean_inc(x_9); -x_13 = l_Lean_Syntax_isNodeOf(x_9, x_11, x_12); -if (x_13 == 0) +x_12 = l_Lean_Syntax_matchesNull(x_9, x_11); +if (x_12 == 0) { -lean_object* x_14; +lean_object* x_13; lean_dec(x_9); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_14 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_4); -return x_14; +x_13 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_4); +return x_13; } else { -lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_15 = l_Lean_Syntax_getArg(x_9, x_8); +lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_14 = l_Lean_Syntax_getArg(x_9, x_8); lean_dec(x_9); -x_16 = l_Lean_Elab_Command_elabSyntax___closed__4; -lean_inc(x_15); -x_17 = l_Lean_Syntax_isOfKind(x_15, x_16); -if (x_17 == 0) +x_15 = l_Lean_Elab_Command_elabSyntax___closed__4; +lean_inc(x_14); +x_16 = l_Lean_Syntax_isOfKind(x_14, x_15); +if (x_16 == 0) { -lean_object* x_18; -lean_dec(x_15); +lean_object* x_17; +lean_dec(x_14); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_18 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_4); -return x_18; +x_17 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___rarg(x_4); +return x_17; } 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; -x_19 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_19, 0, x_15); -x_20 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__6; -x_21 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__4; -x_22 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__2; -x_23 = lean_box(0); +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_18 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_18, 0, x_14); +x_19 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__6; +x_20 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__4; +x_21 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__2; +x_22 = lean_box(0); lean_inc(x_1); -x_24 = l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__3(x_1, x_20, x_21, x_22, x_1, x_23, x_19, x_2, x_3, x_4); +x_23 = l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__3(x_1, x_19, x_20, x_21, x_1, x_22, x_18, x_2, x_3, x_4); lean_dec(x_1); -return x_24; +return x_23; } } } 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_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_dec(x_9); -x_25 = lean_box(0); -x_26 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__6; -x_27 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__4; -x_28 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__2; -x_29 = lean_box(0); +x_24 = lean_box(0); +x_25 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__6; +x_26 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__4; +x_27 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__2; +x_28 = lean_box(0); lean_inc(x_1); -x_30 = l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__3(x_1, x_26, x_27, x_28, x_1, x_29, x_25, x_2, x_3, x_4); +x_29 = l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__3(x_1, x_25, x_26, x_27, x_1, x_28, x_24, x_2, x_3, x_4); lean_dec(x_1); -return x_30; +return x_29; } } } @@ -19379,7 +21123,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabSyntaxAbbrev_decl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(332u); +x_1 = lean_unsigned_to_nat(355u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19391,7 +21135,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabSyntaxAbbrev_decl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(338u); +x_1 = lean_unsigned_to_nat(361u); x_2 = lean_unsigned_to_nat(49u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19419,7 +21163,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabSyntaxAbbrev_decl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(332u); +x_1 = lean_unsigned_to_nat(355u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19431,7 +21175,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Command_elabSyntaxAbbrev_decl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(332u); +x_1 = lean_unsigned_to_nat(355u); x_2 = lean_unsigned_to_nat(57u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19532,7 +21276,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Comma _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_checkLeftRec___spec__8___rarg___closed__2; +x_2 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__2___rarg___closed__2; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); @@ -19551,7 +21295,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Comma _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_checkLeftRec___spec__8___rarg___closed__2; +x_2 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__2___rarg___closed__2; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); @@ -19614,79 +21358,78 @@ return x_7; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +lean_object* x_8; lean_object* x_9; uint8_t x_10; x_8 = lean_unsigned_to_nat(1u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); lean_dec(x_1); -x_10 = l_Lean_nullKind; lean_inc(x_9); -x_11 = l_Lean_Syntax_isNodeOf(x_9, x_10, x_8); -if (x_11 == 0) +x_10 = l_Lean_Syntax_matchesNull(x_9, x_8); +if (x_10 == 0) { -lean_object* x_12; +lean_object* x_11; lean_dec(x_9); lean_dec(x_3); lean_dec(x_2); -x_12 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_inferMacroRulesAltKind___spec__1___rarg(x_4); -return x_12; +x_11 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_inferMacroRulesAltKind___spec__1___rarg(x_4); +return x_11; } else { -lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_13 = lean_unsigned_to_nat(0u); -x_14 = l_Lean_Syntax_getArg(x_9, x_13); +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = lean_unsigned_to_nat(0u); +x_13 = l_Lean_Syntax_getArg(x_9, x_12); lean_dec(x_9); -lean_inc(x_14); -x_15 = l_Lean_Syntax_isNodeOf(x_14, x_10, x_8); -if (x_15 == 0) +lean_inc(x_13); +x_14 = l_Lean_Syntax_matchesNull(x_13, x_8); +if (x_14 == 0) { -lean_object* x_16; -lean_dec(x_14); +lean_object* x_15; +lean_dec(x_13); lean_dec(x_3); lean_dec(x_2); -x_16 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_inferMacroRulesAltKind___spec__1___rarg(x_4); -return x_16; +x_15 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_inferMacroRulesAltKind___spec__1___rarg(x_4); +return x_15; } else { -lean_object* x_17; uint8_t x_18; -x_17 = l_Lean_Syntax_getArg(x_14, x_13); -lean_dec(x_14); -x_18 = l_Lean_Syntax_isQuot(x_17); -if (x_18 == 0) +lean_object* x_16; uint8_t x_17; +x_16 = l_Lean_Syntax_getArg(x_13, x_12); +lean_dec(x_13); +x_17 = l_Lean_Syntax_isQuot(x_16); +if (x_17 == 0) { -lean_object* x_19; uint8_t x_20; -lean_dec(x_17); +lean_object* x_18; uint8_t x_19; +lean_dec(x_16); lean_dec(x_3); lean_dec(x_2); -x_19 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_inferMacroRulesAltKind___spec__2___rarg(x_4); -x_20 = !lean_is_exclusive(x_19); -if (x_20 == 0) +x_18 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_inferMacroRulesAltKind___spec__2___rarg(x_4); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) { -return x_19; +return x_18; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_19, 0); -x_22 = lean_ctor_get(x_19, 1); -lean_inc(x_22); +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_18, 0); +x_21 = lean_ctor_get(x_18, 1); lean_inc(x_21); -lean_dec(x_19); -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); -return x_23; +lean_inc(x_20); +lean_dec(x_18); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; } } else { -lean_object* x_24; lean_object* x_25; -x_24 = lean_box(0); -x_25 = l_Lean_Elab_Command_inferMacroRulesAltKind___lambda__1(x_17, x_24, x_2, x_3, x_4); +lean_object* x_23; lean_object* x_24; +x_23 = lean_box(0); +x_24 = l_Lean_Elab_Command_inferMacroRulesAltKind___lambda__1(x_16, x_23, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_2); -return x_25; +return x_24; } } } @@ -19892,6 +21635,117 @@ return x_26; } } } +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandNoKindMacroRulesAux___spec__4(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; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_5 = l_Lean_Elab_Command_getRef(x_2, x_3, x_4); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_5, 1); +lean_inc(x_7); +lean_dec(x_5); +x_8 = lean_ctor_get(x_2, 4); +lean_inc(x_8); +lean_inc(x_8); +x_9 = l_Lean_Elab_getBetterRef(x_6, x_8); +lean_dec(x_6); +x_10 = l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1(x_1, x_2, x_3, x_7); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1(x_11, x_8, x_2, x_3, x_12); +lean_dec(x_2); +lean_dec(x_8); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_9); +lean_ctor_set(x_16, 1, x_15); +lean_ctor_set_tag(x_13, 1); +lean_ctor_set(x_13, 0, x_16); +return x_13; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_17 = lean_ctor_get(x_13, 0); +x_18 = lean_ctor_get(x_13, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_13); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_9); +lean_ctor_set(x_19, 1, x_17); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_18); +return x_20; +} +} +} +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_expandNoKindMacroRulesAux___spec__3(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; lean_object* x_9; uint8_t x_10; +x_6 = l_Lean_Elab_Command_getRef(x_3, x_4, x_5); +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_6, 1); +lean_inc(x_8); +lean_dec(x_6); +x_9 = l_Lean_replaceRef(x_1, x_7); +lean_dec(x_7); +lean_dec(x_1); +x_10 = !lean_is_exclusive(x_3); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_3, 6); +lean_dec(x_11); +lean_ctor_set(x_3, 6, x_9); +x_12 = l_Lean_throwError___at_Lean_Elab_Command_expandNoKindMacroRulesAux___spec__4(x_2, x_3, x_4, x_8); +lean_dec(x_4); +return x_12; +} +else +{ +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; lean_object* x_20; lean_object* x_21; +x_13 = lean_ctor_get(x_3, 0); +x_14 = lean_ctor_get(x_3, 1); +x_15 = lean_ctor_get(x_3, 2); +x_16 = lean_ctor_get(x_3, 3); +x_17 = lean_ctor_get(x_3, 4); +x_18 = lean_ctor_get(x_3, 5); +x_19 = lean_ctor_get(x_3, 7); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_3); +x_20 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_20, 0, x_13); +lean_ctor_set(x_20, 1, x_14); +lean_ctor_set(x_20, 2, x_15); +lean_ctor_set(x_20, 3, x_16); +lean_ctor_set(x_20, 4, x_17); +lean_ctor_set(x_20, 5, x_18); +lean_ctor_set(x_20, 6, x_9); +lean_ctor_set(x_20, 7, x_19); +x_21 = l_Lean_throwError___at_Lean_Elab_Command_expandNoKindMacroRulesAux___spec__4(x_2, x_20, x_4, x_8); +lean_dec(x_4); +return x_21; +} +} +} static lean_object* _init_l_Lean_Elab_Command_expandNoKindMacroRulesAux___lambda__1___closed__1() { _start: { @@ -19947,162 +21801,167 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandNoKindMacroRulesAux___lambda_ _start: { lean_object* x_10; uint8_t x_11; -x_10 = l_Lean_choiceKind; +x_10 = l_Lean_Elab_Term_toParserDescr_process___closed__2; x_11 = lean_name_eq(x_5, x_10); if (x_11 == 0) { lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; lean_object* x_16; lean_dec(x_4); +lean_dec(x_3); x_12 = lean_array_get_size(x_1); x_13 = lean_unsigned_to_nat(0u); x_14 = lean_nat_dec_lt(x_13, x_12); if (x_14 == 0) { -lean_object* x_68; -x_68 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__11; -x_15 = x_68; +lean_object* x_74; +x_74 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__3; +x_15 = x_74; x_16 = x_9; -goto block_67; +goto block_73; } else { -uint8_t x_69; -x_69 = lean_nat_dec_le(x_12, x_12); -if (x_69 == 0) +uint8_t x_75; +x_75 = lean_nat_dec_le(x_12, x_12); +if (x_75 == 0) { -lean_object* x_70; -x_70 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__11; -x_15 = x_70; +lean_object* x_76; +x_76 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__3; +x_15 = x_76; x_16 = x_9; -goto block_67; +goto block_73; } else { -size_t x_71; size_t x_72; lean_object* x_73; lean_object* x_74; -x_71 = 0; -x_72 = lean_usize_of_nat(x_12); -x_73 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__11; +size_t x_77; size_t x_78; lean_object* x_79; lean_object* x_80; +x_77 = 0; +x_78 = lean_usize_of_nat(x_12); +x_79 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__3; lean_inc(x_8); lean_inc(x_7); -x_74 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_expandNoKindMacroRulesAux___spec__2(x_5, x_1, x_71, x_72, x_73, x_7, x_8, x_9); -if (lean_obj_tag(x_74) == 0) +x_80 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_expandNoKindMacroRulesAux___spec__2(x_5, x_1, x_77, x_78, x_79, x_7, x_8, x_9); +if (lean_obj_tag(x_80) == 0) { -lean_object* x_75; lean_object* x_76; -x_75 = lean_ctor_get(x_74, 0); -lean_inc(x_75); -x_76 = lean_ctor_get(x_74, 1); -lean_inc(x_76); -lean_dec(x_74); -x_15 = x_75; -x_16 = x_76; -goto block_67; +lean_object* x_81; lean_object* x_82; +x_81 = lean_ctor_get(x_80, 0); +lean_inc(x_81); +x_82 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +lean_dec(x_80); +x_15 = x_81; +x_16 = x_82; +goto block_73; } else { -uint8_t x_77; +uint8_t x_83; lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); lean_dec(x_2); -x_77 = !lean_is_exclusive(x_74); -if (x_77 == 0) +lean_dec(x_1); +x_83 = !lean_is_exclusive(x_80); +if (x_83 == 0) { -return x_74; +return x_80; } else { -lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_78 = lean_ctor_get(x_74, 0); -x_79 = lean_ctor_get(x_74, 1); -lean_inc(x_79); -lean_inc(x_78); -lean_dec(x_74); -x_80 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_80, 0, x_78); -lean_ctor_set(x_80, 1, x_79); -return x_80; +lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_84 = lean_ctor_get(x_80, 0); +x_85 = lean_ctor_get(x_80, 1); +lean_inc(x_85); +lean_inc(x_84); +lean_dec(x_80); +x_86 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_86, 0, x_84); +lean_ctor_set(x_86, 1, x_85); +return x_86; } } } } -block_67: +block_73: { lean_object* x_17; lean_object* x_18; if (x_14 == 0) { -lean_object* x_54; +lean_object* x_60; lean_dec(x_12); -x_54 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__11; -x_17 = x_54; +lean_dec(x_1); +x_60 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__3; +x_17 = x_60; x_18 = x_16; -goto block_53; +goto block_59; } else { -uint8_t x_55; -x_55 = lean_nat_dec_le(x_12, x_12); -if (x_55 == 0) +uint8_t x_61; +x_61 = lean_nat_dec_le(x_12, x_12); +if (x_61 == 0) { -lean_object* x_56; +lean_object* x_62; lean_dec(x_12); -x_56 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__11; -x_17 = x_56; +lean_dec(x_1); +x_62 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__3; +x_17 = x_62; x_18 = x_16; -goto block_53; +goto block_59; } else { -size_t x_57; size_t x_58; lean_object* x_59; lean_object* x_60; -x_57 = 0; -x_58 = lean_usize_of_nat(x_12); +size_t x_63; size_t x_64; lean_object* x_65; lean_object* x_66; +x_63 = 0; +x_64 = lean_usize_of_nat(x_12); lean_dec(x_12); -x_59 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__11; +x_65 = l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__3; lean_inc(x_8); lean_inc(x_7); -x_60 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_expandNoKindMacroRulesAux___spec__1(x_5, x_1, x_57, x_58, x_59, x_7, x_8, x_16); -if (lean_obj_tag(x_60) == 0) +x_66 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_expandNoKindMacroRulesAux___spec__1(x_5, x_1, x_63, x_64, x_65, x_7, x_8, x_16); +lean_dec(x_1); +if (lean_obj_tag(x_66) == 0) { -lean_object* x_61; lean_object* x_62; -x_61 = lean_ctor_get(x_60, 0); -lean_inc(x_61); -x_62 = lean_ctor_get(x_60, 1); -lean_inc(x_62); -lean_dec(x_60); -x_17 = x_61; -x_18 = x_62; -goto block_53; +lean_object* x_67; lean_object* x_68; +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +lean_dec(x_66); +x_17 = x_67; +x_18 = x_68; +goto block_59; } else { -uint8_t x_63; +uint8_t x_69; lean_dec(x_15); lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); lean_dec(x_2); -x_63 = !lean_is_exclusive(x_60); -if (x_63 == 0) +x_69 = !lean_is_exclusive(x_66); +if (x_69 == 0) { -return x_60; +return x_66; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_ctor_get(x_60, 0); -x_65 = lean_ctor_get(x_60, 1); -lean_inc(x_65); -lean_inc(x_64); -lean_dec(x_60); -x_66 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_66, 0, x_64); -lean_ctor_set(x_66, 1, x_65); -return x_66; +lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_70 = lean_ctor_get(x_66, 0); +x_71 = lean_ctor_get(x_66, 1); +lean_inc(x_71); +lean_inc(x_70); +lean_dec(x_66); +x_72 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_72, 0, x_70); +lean_ctor_set(x_72, 1, x_71); +return x_72; } } } } -block_53: +block_59: { uint8_t x_19; x_19 = l_Array_isEmpty___rarg(x_17); @@ -20124,137 +21983,158 @@ x_23 = lean_ctor_get(x_21, 1); lean_inc(x_23); lean_dec(x_21); x_24 = lean_box(0); +lean_inc(x_8); +lean_inc(x_7); x_25 = lean_apply_5(x_2, x_24, x_17, x_7, x_8, x_23); if (lean_obj_tag(x_25) == 0) { -uint8_t x_26; -x_26 = !lean_is_exclusive(x_25); -if (x_26 == 0) +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_object* x_32; uint8_t x_33; +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +lean_dec(x_25); +x_28 = l_Lean_MonadRef_mkInfoFromRefPos___at___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___spec__1(x_7, x_8, x_27); +x_29 = lean_ctor_get(x_28, 1); +lean_inc(x_29); +lean_dec(x_28); +x_30 = l_Lean_Elab_Command_getCurrMacroScope(x_7, x_8, x_29); +lean_dec(x_7); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_32 = l_Lean_Elab_Command_getMainModule___rarg(x_8, x_31); +lean_dec(x_8); +x_33 = !lean_is_exclusive(x_32); +if (x_33 == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_27 = lean_ctor_get(x_25, 0); -x_28 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; -x_29 = lean_array_push(x_28, x_22); -x_30 = lean_array_push(x_29, x_27); -x_31 = lean_box(2); -x_32 = l_Lean_nullKind; -x_33 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -lean_ctor_set(x_33, 2, x_30); -lean_ctor_set(x_25, 0, x_33); -return x_25; +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; +x_34 = lean_ctor_get(x_32, 0); +lean_dec(x_34); +x_35 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; +x_36 = lean_array_push(x_35, x_22); +x_37 = lean_array_push(x_36, x_26); +x_38 = lean_box(2); +x_39 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; +x_40 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_40, 0, x_38); +lean_ctor_set(x_40, 1, x_39); +lean_ctor_set(x_40, 2, x_37); +lean_ctor_set(x_32, 0, x_40); +return x_32; } else { -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; -x_34 = lean_ctor_get(x_25, 0); -x_35 = lean_ctor_get(x_25, 1); -lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_25); -x_36 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; -x_37 = lean_array_push(x_36, x_22); -x_38 = lean_array_push(x_37, x_34); -x_39 = lean_box(2); -x_40 = l_Lean_nullKind; -x_41 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_40); -lean_ctor_set(x_41, 2, x_38); -x_42 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_35); -return x_42; +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_41 = lean_ctor_get(x_32, 1); +lean_inc(x_41); +lean_dec(x_32); +x_42 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29; +x_43 = lean_array_push(x_42, x_22); +x_44 = lean_array_push(x_43, x_26); +x_45 = lean_box(2); +x_46 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__21; +x_47 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +lean_ctor_set(x_47, 2, x_44); +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_41); +return x_48; } } else { -uint8_t x_43; +uint8_t x_49; lean_dec(x_22); -x_43 = !lean_is_exclusive(x_25); -if (x_43 == 0) +lean_dec(x_8); +lean_dec(x_7); +x_49 = !lean_is_exclusive(x_25); +if (x_49 == 0) { return x_25; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_25, 0); -x_45 = lean_ctor_get(x_25, 1); -lean_inc(x_45); -lean_inc(x_44); +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_25, 0); +x_51 = lean_ctor_get(x_25, 1); +lean_inc(x_51); +lean_inc(x_50); lean_dec(x_25); -x_46 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -return x_46; +x_52 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_50); +lean_ctor_set(x_52, 1, x_51); +return x_52; } } } else { -uint8_t x_47; +uint8_t x_53; lean_dec(x_17); lean_dec(x_8); lean_dec(x_7); lean_dec(x_2); -x_47 = !lean_is_exclusive(x_21); -if (x_47 == 0) +x_53 = !lean_is_exclusive(x_21); +if (x_53 == 0) { return x_21; } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_48 = lean_ctor_get(x_21, 0); -x_49 = lean_ctor_get(x_21, 1); -lean_inc(x_49); -lean_inc(x_48); +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_ctor_get(x_21, 0); +x_55 = lean_ctor_get(x_21, 1); +lean_inc(x_55); +lean_inc(x_54); lean_dec(x_21); -x_50 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_50, 0, x_48); -lean_ctor_set(x_50, 1, x_49); -return x_50; +x_56 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_56, 0, x_54); +lean_ctor_set(x_56, 1, x_55); +return x_56; } } } else { -lean_object* x_51; lean_object* x_52; +lean_object* x_57; lean_object* x_58; lean_dec(x_17); -x_51 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_51, 0, x_5); -x_52 = lean_apply_5(x_2, x_51, x_15, x_7, x_8, x_18); -return x_52; +x_57 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_57, 0, x_5); +x_58 = lean_apply_5(x_2, x_57, x_15, x_7, x_8, x_18); +return x_58; } } } } else { -lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_dec(x_5); lean_dec(x_2); -x_81 = l_Lean_stringToMessageData(x_3); -x_82 = l_Lean_Elab_Command_expandNoKindMacroRulesAux___lambda__1___closed__2; -lean_inc(x_81); -x_83 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_83, 0, x_82); -lean_ctor_set(x_83, 1, x_81); -x_84 = l_Lean_Elab_Command_expandNoKindMacroRulesAux___lambda__1___closed__4; -x_85 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_85, 0, x_83); -lean_ctor_set(x_85, 1, x_84); -x_86 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_86, 0, x_85); -lean_ctor_set(x_86, 1, x_81); -x_87 = l_Lean_Elab_Command_expandNoKindMacroRulesAux___lambda__1___closed__6; -x_88 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_88, 0, x_86); -lean_ctor_set(x_88, 1, x_87); -x_89 = l_Lean_throwErrorAt___at_Lean_Elab_Command_elabCommand___spec__11(x_4, x_88, x_7, x_8, x_9); -return x_89; +lean_dec(x_1); +x_87 = l_Lean_stringToMessageData(x_3); +lean_dec(x_3); +x_88 = l_Lean_Elab_Command_expandNoKindMacroRulesAux___lambda__1___closed__2; +lean_inc(x_87); +x_89 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_89, 0, x_88); +lean_ctor_set(x_89, 1, x_87); +x_90 = l_Lean_Elab_Command_expandNoKindMacroRulesAux___lambda__1___closed__4; +x_91 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_91, 0, x_89); +lean_ctor_set(x_91, 1, x_90); +x_92 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_92, 0, x_91); +lean_ctor_set(x_92, 1, x_87); +x_93 = l_Lean_Elab_Command_expandNoKindMacroRulesAux___lambda__1___closed__6; +x_94 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_94, 0, x_92); +lean_ctor_set(x_94, 1, x_93); +x_95 = l_Lean_throwErrorAt___at_Lean_Elab_Command_expandNoKindMacroRulesAux___spec__3(x_4, x_94, x_7, x_8, x_9); +return x_95; } } } @@ -20262,7 +22142,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandNoKindMacroRulesAux(lean_obje _start: { lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_7 = l_Lean_instInhabitedSyntax; +x_7 = lean_box(0); x_8 = lean_unsigned_to_nat(0u); x_9 = lean_array_get(x_7, x_1, x_8); lean_inc(x_5); @@ -20283,13 +22163,11 @@ if (x_13 == 0) lean_object* x_14; lean_object* x_15; x_14 = lean_box(0); x_15 = l_Lean_Elab_Command_expandNoKindMacroRulesAux___lambda__1(x_1, x_3, x_2, x_9, x_11, x_14, x_4, x_5, x_12); -lean_dec(x_2); return x_15; } else { lean_object* x_16; lean_object* x_17; uint8_t x_18; -lean_inc(x_11); x_16 = l_Lean_Name_getString_x21(x_11); x_17 = l_Lean_Elab_Command_checkRuleKind___closed__1; x_18 = lean_string_dec_eq(x_16, x_17); @@ -20299,7 +22177,6 @@ if (x_18 == 0) lean_object* x_19; lean_object* x_20; x_19 = lean_box(0); x_20 = l_Lean_Elab_Command_expandNoKindMacroRulesAux___lambda__1(x_1, x_3, x_2, x_9, x_11, x_19, x_4, x_5, x_12); -lean_dec(x_2); return x_20; } else @@ -20309,7 +22186,6 @@ x_21 = l_Lean_Name_getPrefix(x_11); lean_dec(x_11); x_22 = lean_box(0); x_23 = l_Lean_Elab_Command_expandNoKindMacroRulesAux___lambda__1(x_1, x_3, x_2, x_9, x_21, x_22, x_4, x_5, x_12); -lean_dec(x_2); return x_23; } } @@ -20322,6 +22198,7 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); +lean_dec(x_1); x_24 = !lean_is_exclusive(x_10); if (x_24 == 0) { @@ -20371,24 +22248,22 @@ lean_dec(x_1); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandNoKindMacroRulesAux___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, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandNoKindMacroRulesAux___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_10; -x_10 = l_Lean_Elab_Command_expandNoKindMacroRulesAux___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_6); +lean_object* x_5; +x_5 = l_Lean_throwError___at_Lean_Elab_Command_expandNoKindMacroRulesAux___spec__4(x_1, x_2, x_3, x_4); lean_dec(x_3); -lean_dec(x_1); -return x_10; +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandNoKindMacroRulesAux___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_Lean_Elab_Command_expandNoKindMacroRulesAux___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, lean_object* x_9) { _start: { -lean_object* x_7; -x_7 = l_Lean_Elab_Command_expandNoKindMacroRulesAux(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_1); -return x_7; +lean_object* x_10; +x_10 = l_Lean_Elab_Command_expandNoKindMacroRulesAux___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_6); +return x_10; } } LEAN_EXPORT lean_object* l_Lean_Elab_Command_strLitToPattern(lean_object* x_1, lean_object* x_2, lean_object* x_3) { @@ -20429,7 +22304,7 @@ lean_dec(x_2); return x_4; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_6964____closed__1() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_8045____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -20439,11 +22314,11 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_6964_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_8045_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_6964____closed__1; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_8045____closed__1; x_3 = l_Lean_registerTraceClass(x_2, x_1); return x_3; } @@ -20527,6 +22402,54 @@ l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mk lean_mark_persistent(l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__28); l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29 = _init_l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29(); lean_mark_persistent(l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__29); +l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__2___rarg___closed__1 = _init_l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__2___rarg___closed__1(); +lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__2___rarg___closed__1); +l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__2___rarg___closed__2 = _init_l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__2___rarg___closed__2(); +lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__2___rarg___closed__2); +l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___closed__1 = _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___closed__1(); +lean_mark_persistent(l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___closed__1); +l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___closed__2 = _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___closed__2(); +lean_mark_persistent(l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___closed__2); +l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___closed__3 = _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___closed__3(); +lean_mark_persistent(l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___closed__3); +l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___closed__4 = _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___closed__4(); +lean_mark_persistent(l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___closed__4); +l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__1 = _init_l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__1); +l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__2 = _init_l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__2); +l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__3 = _init_l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__3); +l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__4 = _init_l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__4); +l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__5 = _init_l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__5); +l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__6 = _init_l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__6(); +lean_mark_persistent(l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__6); +l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__7 = _init_l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__7(); +lean_mark_persistent(l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__7); +l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__8 = _init_l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__8(); +lean_mark_persistent(l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__8); +l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__9 = _init_l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__9(); +lean_mark_persistent(l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__9); +l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__10 = _init_l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__10(); +lean_mark_persistent(l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__10); +l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__11 = _init_l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__11(); +lean_mark_persistent(l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__11); +l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__12 = _init_l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__12(); +lean_mark_persistent(l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__12); +l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__13 = _init_l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__13(); +lean_mark_persistent(l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__13); +l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__14 = _init_l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__14(); +lean_mark_persistent(l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__14); +l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__15 = _init_l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__15(); +lean_mark_persistent(l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__15); +l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__16 = _init_l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__16(); +lean_mark_persistent(l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__16); +l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__17 = _init_l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__17(); +lean_mark_persistent(l_Lean_Elab_Term_ensureUnaryOutput___lambda__1___closed__17); +l_Lean_Elab_Term_ensureUnaryOutput___closed__1 = _init_l_Lean_Elab_Term_ensureUnaryOutput___closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_ensureUnaryOutput___closed__1); l_Lean_addTrace___at_Lean_Elab_Term_checkLeftRec___spec__3___closed__1 = _init_l_Lean_addTrace___at_Lean_Elab_Term_checkLeftRec___spec__3___closed__1(); lean_mark_persistent(l_Lean_addTrace___at_Lean_Elab_Term_checkLeftRec___spec__3___closed__1); l_Lean_addTrace___at_Lean_Elab_Term_checkLeftRec___spec__3___closed__2 = _init_l_Lean_addTrace___at_Lean_Elab_Term_checkLeftRec___spec__3___closed__2(); @@ -20547,10 +22470,6 @@ l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_checkLeftRec___spec__7___closed__1 lean_mark_persistent(l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_checkLeftRec___spec__7___closed__1); l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_checkLeftRec___spec__7___closed__2 = _init_l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_checkLeftRec___spec__7___closed__2(); lean_mark_persistent(l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_checkLeftRec___spec__7___closed__2); -l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_checkLeftRec___spec__8___rarg___closed__1 = _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_checkLeftRec___spec__8___rarg___closed__1(); -lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_checkLeftRec___spec__8___rarg___closed__1); -l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_checkLeftRec___spec__8___rarg___closed__2 = _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_checkLeftRec___spec__8___rarg___closed__2(); -lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_checkLeftRec___spec__8___rarg___closed__2); l_Lean_Elab_Term_checkLeftRec___lambda__2___closed__1 = _init_l_Lean_Elab_Term_checkLeftRec___lambda__2___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_checkLeftRec___lambda__2___closed__1); l_Lean_Elab_Term_checkLeftRec___lambda__2___closed__2 = _init_l_Lean_Elab_Term_checkLeftRec___lambda__2___closed__2(); @@ -20639,10 +22558,6 @@ l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__6 = _ lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__6); l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__7 = _init_l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__7(); lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__7); -l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__8 = _init_l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__8(); -lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__8); -l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__9 = _init_l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__9(); -lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__9); l_Lean_Elab_Term_toParserDescr_processParserCategory___closed__1 = _init_l_Lean_Elab_Term_toParserDescr_processParserCategory___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processParserCategory___closed__1); l_Lean_Elab_Term_toParserDescr_processParserCategory___closed__2 = _init_l_Lean_Elab_Term_toParserDescr_processParserCategory___closed__2(); @@ -20651,66 +22566,6 @@ l_Lean_Elab_Term_toParserDescr_ensureNoPrec___closed__1 = _init_l_Lean_Elab_Term lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_ensureNoPrec___closed__1); l_Lean_Elab_Term_toParserDescr_ensureNoPrec___closed__2 = _init_l_Lean_Elab_Term_toParserDescr_ensureNoPrec___closed__2(); lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_ensureNoPrec___closed__2); -l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__1 = _init_l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__1(); -lean_mark_persistent(l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__1); -l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__2 = _init_l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__2(); -lean_mark_persistent(l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__2); -l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__3 = _init_l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__3(); -lean_mark_persistent(l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__3); -l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__4 = _init_l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__4(); -lean_mark_persistent(l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__4); -l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3___closed__1 = _init_l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3___closed__1(); -lean_mark_persistent(l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3___closed__1); -l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3___closed__2 = _init_l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3___closed__2(); -lean_mark_persistent(l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3___closed__2); -l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3___closed__3 = _init_l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3___closed__3(); -lean_mark_persistent(l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3___closed__3); -l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___closed__1 = _init_l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___closed__1(); -lean_mark_persistent(l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___closed__1); -l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___closed__2 = _init_l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___closed__2(); -lean_mark_persistent(l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___closed__2); -l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___closed__3 = _init_l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___closed__3(); -lean_mark_persistent(l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___closed__3); -l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__1 = _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__1); -l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__2 = _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__2); -l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__3 = _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__3(); -lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__3); -l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__4 = _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__4(); -lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__4); -l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__5 = _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__5(); -lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__5); -l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__6 = _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__6(); -lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__6); -l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__7 = _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__7(); -lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__7); -l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__8 = _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__8(); -lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__8); -l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__9 = _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__9(); -lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__9); -l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__10 = _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__10(); -lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__10); -l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__11 = _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__11(); -lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__11); -l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__12 = _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__12(); -lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__12); -l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__13 = _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__13(); -lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__13); -l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__14 = _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__14(); -lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__14); -l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__15 = _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__15(); -lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__15); -l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__16 = _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__16(); -lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__16); -l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__17 = _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__17(); -lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__17); -l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__18 = _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__18(); -lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__18); -l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__19 = _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__19(); -lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__19); -l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__20 = _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__20(); -lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__20); l_Lean_Elab_Term_toParserDescr_process___closed__1 = _init_l_Lean_Elab_Term_toParserDescr_process___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_process___closed__1); l_Lean_Elab_Term_toParserDescr_process___closed__2 = _init_l_Lean_Elab_Term_toParserDescr_process___closed__2(); @@ -20741,6 +22596,8 @@ l_Lean_Elab_Term_toParserDescr_process___closed__14 = _init_l_Lean_Elab_Term_toP lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_process___closed__14); l_Lean_Elab_Term_toParserDescr_process___closed__15 = _init_l_Lean_Elab_Term_toParserDescr_process___closed__15(); lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_process___closed__15); +l_Lean_Elab_Term_toParserDescr_process___closed__16 = _init_l_Lean_Elab_Term_toParserDescr_process___closed__16(); +lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_process___closed__16); l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__1 = _init_l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__1); l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__2 = _init_l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__2(); @@ -20779,20 +22636,82 @@ l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__6 = _init_l_Le lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__6); l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__7 = _init_l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__7(); lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__7); -l_Lean_Elab_Term_toParserDescr_processUnary___closed__1 = _init_l_Lean_Elab_Term_toParserDescr_processUnary___closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processUnary___closed__1); -l_Lean_Elab_Term_toParserDescr_processUnary___closed__2 = _init_l_Lean_Elab_Term_toParserDescr_processUnary___closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processUnary___closed__2); -l_Lean_Elab_Term_toParserDescr_processUnary___closed__3 = _init_l_Lean_Elab_Term_toParserDescr_processUnary___closed__3(); -lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processUnary___closed__3); -l_Lean_Elab_Term_toParserDescr_processUnary___closed__4 = _init_l_Lean_Elab_Term_toParserDescr_processUnary___closed__4(); -lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processUnary___closed__4); -l_Lean_Elab_Term_toParserDescr_processUnary___closed__5 = _init_l_Lean_Elab_Term_toParserDescr_processUnary___closed__5(); -lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processUnary___closed__5); -l_Lean_Elab_Term_toParserDescr_processUnary___closed__6 = _init_l_Lean_Elab_Term_toParserDescr_processUnary___closed__6(); -lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processUnary___closed__6); -l_Lean_Elab_Term_toParserDescr_processUnary___closed__7 = _init_l_Lean_Elab_Term_toParserDescr_processUnary___closed__7(); -lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processUnary___closed__7); +l_panic___at_Lean_Elab_Term_toParserDescr_processAlias___spec__2___closed__1 = _init_l_panic___at_Lean_Elab_Term_toParserDescr_processAlias___spec__2___closed__1(); +lean_mark_persistent(l_panic___at_Lean_Elab_Term_toParserDescr_processAlias___spec__2___closed__1); +l_panic___at_Lean_Elab_Term_toParserDescr_processAlias___spec__2___closed__2 = _init_l_panic___at_Lean_Elab_Term_toParserDescr_processAlias___spec__2___closed__2(); +lean_mark_persistent(l_panic___at_Lean_Elab_Term_toParserDescr_processAlias___spec__2___closed__2); +l_panic___at_Lean_Elab_Term_toParserDescr_processAlias___spec__2___closed__3 = _init_l_panic___at_Lean_Elab_Term_toParserDescr_processAlias___spec__2___closed__3(); +lean_mark_persistent(l_panic___at_Lean_Elab_Term_toParserDescr_processAlias___spec__2___closed__3); +l_Lean_Elab_Term_toParserDescr_processAlias___closed__1 = _init_l_Lean_Elab_Term_toParserDescr_processAlias___closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processAlias___closed__1); +l_Lean_Elab_Term_toParserDescr_processAlias___closed__2 = _init_l_Lean_Elab_Term_toParserDescr_processAlias___closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processAlias___closed__2); +l_Lean_Elab_Term_toParserDescr_processAlias___closed__3 = _init_l_Lean_Elab_Term_toParserDescr_processAlias___closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processAlias___closed__3); +l_Lean_Elab_Term_toParserDescr_processAlias___closed__4 = _init_l_Lean_Elab_Term_toParserDescr_processAlias___closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processAlias___closed__4); +l_Lean_Elab_Term_toParserDescr_processAlias___closed__5 = _init_l_Lean_Elab_Term_toParserDescr_processAlias___closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processAlias___closed__5); +l_Lean_Elab_Term_toParserDescr_processAlias___closed__6 = _init_l_Lean_Elab_Term_toParserDescr_processAlias___closed__6(); +lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processAlias___closed__6); +l_Lean_Elab_Term_toParserDescr_processAlias___closed__7 = _init_l_Lean_Elab_Term_toParserDescr_processAlias___closed__7(); +lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processAlias___closed__7); +l_Lean_Elab_Term_toParserDescr_processAlias___closed__8 = _init_l_Lean_Elab_Term_toParserDescr_processAlias___closed__8(); +lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processAlias___closed__8); +l_Lean_Elab_Term_toParserDescr_processAlias___closed__9 = _init_l_Lean_Elab_Term_toParserDescr_processAlias___closed__9(); +lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processAlias___closed__9); +l_Lean_Elab_Term_toParserDescr_processAlias___closed__10 = _init_l_Lean_Elab_Term_toParserDescr_processAlias___closed__10(); +lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processAlias___closed__10); +l_Lean_Elab_Term_toParserDescr_processAlias___closed__11 = _init_l_Lean_Elab_Term_toParserDescr_processAlias___closed__11(); +lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processAlias___closed__11); +l_Lean_Elab_Term_toParserDescr_processAlias___closed__12 = _init_l_Lean_Elab_Term_toParserDescr_processAlias___closed__12(); +lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processAlias___closed__12); +l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__1 = _init_l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__1(); +lean_mark_persistent(l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__1); +l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__2 = _init_l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__2(); +lean_mark_persistent(l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__2); +l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__3 = _init_l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__3(); +lean_mark_persistent(l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__3); +l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__4 = _init_l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__4(); +lean_mark_persistent(l_Lean_throwUnknownConstant___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__8___closed__4); +l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3___closed__1 = _init_l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3___closed__1(); +lean_mark_persistent(l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3___closed__1); +l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3___closed__2 = _init_l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3___closed__2(); +lean_mark_persistent(l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3___closed__2); +l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3___closed__3 = _init_l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3___closed__3(); +lean_mark_persistent(l_Lean_resolveGlobalConst___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__3___closed__3); +l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___closed__1 = _init_l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___closed__1(); +lean_mark_persistent(l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___closed__1); +l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___closed__2 = _init_l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___closed__2(); +lean_mark_persistent(l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___closed__2); +l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___closed__3 = _init_l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___closed__3(); +lean_mark_persistent(l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_toParserDescr_processNullaryOrCat___spec__12___closed__3); +l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__1 = _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__1); +l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__2 = _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__2); +l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__3 = _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__3); +l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__4 = _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__4); +l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__5 = _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__5); +l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__6 = _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__6(); +lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__6); +l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__7 = _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__7(); +lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__7); +l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__8 = _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__8(); +lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__8); +l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__9 = _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__9(); +lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__9); +l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__10 = _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__10(); +lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__10); +l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__11 = _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__11(); +lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__11); +l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__12 = _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__12(); +lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__12); +l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__13 = _init_l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__13(); +lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__13); l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__1 = _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__1); l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__2 = _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__2(); @@ -20961,6 +22880,8 @@ l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___c lean_mark_persistent(l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__83); l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__84 = _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__84(); lean_mark_persistent(l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__84); +l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__85 = _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__85(); +lean_mark_persistent(l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__85); l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__1 = _init_l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__1); l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__2 = _init_l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__2(); @@ -21041,10 +22962,6 @@ l_Lean_Elab_Command_elabSyntax___lambda__5___closed__11 = _init_l_Lean_Elab_Comm lean_mark_persistent(l_Lean_Elab_Command_elabSyntax___lambda__5___closed__11); l_Lean_Elab_Command_elabSyntax___lambda__5___closed__12 = _init_l_Lean_Elab_Command_elabSyntax___lambda__5___closed__12(); lean_mark_persistent(l_Lean_Elab_Command_elabSyntax___lambda__5___closed__12); -l_Lean_Elab_Command_elabSyntax___lambda__5___closed__13 = _init_l_Lean_Elab_Command_elabSyntax___lambda__5___closed__13(); -lean_mark_persistent(l_Lean_Elab_Command_elabSyntax___lambda__5___closed__13); -l_Lean_Elab_Command_elabSyntax___lambda__5___closed__14 = _init_l_Lean_Elab_Command_elabSyntax___lambda__5___closed__14(); -lean_mark_persistent(l_Lean_Elab_Command_elabSyntax___lambda__5___closed__14); l_Lean_Elab_Command_elabSyntax___lambda__8___closed__1 = _init_l_Lean_Elab_Command_elabSyntax___lambda__8___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_elabSyntax___lambda__8___closed__1); l_Lean_Elab_Command_elabSyntax___lambda__8___closed__2 = _init_l_Lean_Elab_Command_elabSyntax___lambda__8___closed__2(); @@ -21155,9 +23072,9 @@ l_Lean_Elab_Command_expandNoKindMacroRulesAux___lambda__1___closed__5 = _init_l_ lean_mark_persistent(l_Lean_Elab_Command_expandNoKindMacroRulesAux___lambda__1___closed__5); l_Lean_Elab_Command_expandNoKindMacroRulesAux___lambda__1___closed__6 = _init_l_Lean_Elab_Command_expandNoKindMacroRulesAux___lambda__1___closed__6(); lean_mark_persistent(l_Lean_Elab_Command_expandNoKindMacroRulesAux___lambda__1___closed__6); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_6964____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_6964____closed__1(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_6964____closed__1); -res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_6964_(lean_io_mk_world()); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_8045____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_8045____closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_8045____closed__1); +res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_8045_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Basic.c b/stage0/stdlib/Lean/Elab/Tactic/Basic.c index d9577807759..ce7f9b041f2 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Basic.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Basic.c @@ -20,6 +20,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Tactic_wit static lean_object* l_Lean_Elab_Tactic_getNameOfIdent_x27___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_withCaseRef(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTacticAux___boxed__const__1; +static lean_object* l_Lean_Elab_Tactic_evalTacticAux___closed__8; static lean_object* l_Lean_Elab_Tactic_evalTacticAux___closed__4; static lean_object* l_Lean_Elab_Tactic_expandTacticMacroFns_loop___closed__1; size_t lean_usize_add(size_t, size_t); @@ -42,7 +43,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_expandTact LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_throwNoGoalsToBeSolved___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_withCaseRef___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_throwAbortTactic___at_Lean_Elab_Tactic_done___spec__1___rarg___closed__2; -extern lean_object* l_Lean_nullKind; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_adaptExpander___lambda__1(lean_object*, 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_getResetInfoTrees___at_Lean_Elab_Tactic_withTacticInfoContext___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOptionM___at_Lean_Elab_Tactic_expandTacticMacroFns_loop___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -237,6 +237,7 @@ LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Elab_Tactic_saveTacticInfoForT LEAN_EXPORT lean_object* l_Lean_Elab_withInfoTreeContext___at_Lean_Elab_Tactic_withMacroExpansion___spec__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Tactic_expandTacticMacroFns_loop___spec__3___boxed(lean_object*, 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_Tactic_liftMetaTactic___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_evalTacticAux___closed__7; static lean_object* l_Lean_addTrace___at_Lean_Elab_Tactic_expandTacticMacroFns_loop___spec__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getNameOfIdent_x27___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_liftMetaTactic1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -5872,6 +5873,24 @@ return x_1; static lean_object* _init_l_Lean_Elab_Tactic_evalTacticAux___closed__2() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("null", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_evalTacticAux___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Tactic_evalTacticAux___closed__2; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_evalTacticAux___closed__4() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); x_2 = l_Lean_Elab_Tactic_mkTacticAttribute___closed__3; @@ -5879,7 +5898,7 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalTacticAux___closed__3() { +static lean_object* _init_l_Lean_Elab_Tactic_evalTacticAux___closed__5() { _start: { lean_object* x_1; @@ -5887,17 +5906,17 @@ x_1 = lean_mk_string_from_bytes("step", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalTacticAux___closed__4() { +static lean_object* _init_l_Lean_Elab_Tactic_evalTacticAux___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_evalTacticAux___closed__2; -x_2 = l_Lean_Elab_Tactic_evalTacticAux___closed__3; +x_1 = l_Lean_Elab_Tactic_evalTacticAux___closed__4; +x_2 = l_Lean_Elab_Tactic_evalTacticAux___closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalTacticAux___closed__5() { +static lean_object* _init_l_Lean_Elab_Tactic_evalTacticAux___closed__7() { _start: { lean_object* x_1; @@ -5905,11 +5924,11 @@ x_1 = lean_mk_string_from_bytes("unexpected tactic", 17); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalTacticAux___closed__6() { +static lean_object* _init_l_Lean_Elab_Tactic_evalTacticAux___closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic_evalTacticAux___closed__5; +x_1 = l_Lean_Elab_Tactic_evalTacticAux___closed__7; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } @@ -5995,13 +6014,13 @@ case 1: lean_object* x_30; lean_object* x_31; uint8_t x_32; x_30 = lean_ctor_get(x_1, 1); lean_inc(x_30); -x_31 = l_Lean_nullKind; +x_31 = l_Lean_Elab_Tactic_evalTacticAux___closed__3; x_32 = lean_name_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; -x_33 = l_Lean_Elab_Tactic_evalTacticAux___closed__4; +x_33 = l_Lean_Elab_Tactic_evalTacticAux___closed__6; x_34 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_evalTacticAux___lambda__3), 11, 8); lean_closure_set(x_34, 0, x_1); lean_closure_set(x_34, 1, x_2); @@ -6087,7 +6106,7 @@ lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean x_51 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_51, 0, x_1); x_52 = l_Lean_indentD(x_51); -x_53 = l_Lean_Elab_Tactic_evalTacticAux___closed__6; +x_53 = l_Lean_Elab_Tactic_evalTacticAux___closed__8; x_54 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_54, 0, x_53); lean_ctor_set(x_54, 1, x_52); @@ -6224,13 +6243,13 @@ case 1: lean_object* x_79; lean_object* x_80; uint8_t x_81; x_79 = lean_ctor_get(x_1, 1); lean_inc(x_79); -x_80 = l_Lean_nullKind; +x_80 = l_Lean_Elab_Tactic_evalTacticAux___closed__3; x_81 = lean_name_eq(x_79, x_80); lean_dec(x_79); if (x_81 == 0) { lean_object* x_82; lean_object* x_83; lean_object* x_84; -x_82 = l_Lean_Elab_Tactic_evalTacticAux___closed__4; +x_82 = l_Lean_Elab_Tactic_evalTacticAux___closed__6; x_83 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_evalTacticAux___lambda__3), 11, 8); lean_closure_set(x_83, 0, x_1); lean_closure_set(x_83, 1, x_2); @@ -6316,7 +6335,7 @@ lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; x_100 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_100, 0, x_1); x_101 = l_Lean_indentD(x_100); -x_102 = l_Lean_Elab_Tactic_evalTacticAux___closed__6; +x_102 = l_Lean_Elab_Tactic_evalTacticAux___closed__8; x_103 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_103, 0, x_102); lean_ctor_set(x_103, 1, x_101); @@ -12210,7 +12229,7 @@ x_6 = l_Lean_Elab_Tactic_withCaseRef___rarg___closed__1; x_7 = lean_array_push(x_6, x_3); x_8 = lean_array_push(x_7, x_4); x_9 = lean_box(2); -x_10 = l_Lean_nullKind; +x_10 = l_Lean_Elab_Tactic_evalTacticAux___closed__3; x_11 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_11, 0, x_9); lean_ctor_set(x_11, 1, x_10); @@ -12250,7 +12269,7 @@ static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basi _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_evalTacticAux___closed__2; +x_1 = l_Lean_Elab_Tactic_evalTacticAux___closed__4; x_2 = l_Lean_Elab_Tactic_mkTacticAttribute___closed__10; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -12452,6 +12471,10 @@ l_Lean_Elab_Tactic_evalTacticAux___closed__5 = _init_l_Lean_Elab_Tactic_evalTact lean_mark_persistent(l_Lean_Elab_Tactic_evalTacticAux___closed__5); l_Lean_Elab_Tactic_evalTacticAux___closed__6 = _init_l_Lean_Elab_Tactic_evalTacticAux___closed__6(); lean_mark_persistent(l_Lean_Elab_Tactic_evalTacticAux___closed__6); +l_Lean_Elab_Tactic_evalTacticAux___closed__7 = _init_l_Lean_Elab_Tactic_evalTacticAux___closed__7(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalTacticAux___closed__7); +l_Lean_Elab_Tactic_evalTacticAux___closed__8 = _init_l_Lean_Elab_Tactic_evalTacticAux___closed__8(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalTacticAux___closed__8); l_Lean_Elab_Tactic_evalTacticAux___boxed__const__1 = _init_l_Lean_Elab_Tactic_evalTacticAux___boxed__const__1(); lean_mark_persistent(l_Lean_Elab_Tactic_evalTacticAux___boxed__const__1); l_Lean_Elab_Tactic_expandTacticMacro___closed__1 = _init_l_Lean_Elab_Tactic_expandTacticMacro___closed__1(); diff --git a/stage0/stdlib/Lean/Elab/Tactic/BuiltinTactic.c b/stage0/stdlib/Lean/Elab/Tactic/BuiltinTactic.c index 5052d4ac4c0..1838ad50889 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/BuiltinTactic.c +++ b/stage0/stdlib/Lean/Elab/Tactic/BuiltinTactic.c @@ -63,6 +63,7 @@ static lean_object* l_Lean_Elab_Tactic_evalIntro___closed__20; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_evalOpen___spec__29___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSkip_declRange___closed__7; lean_object* lean_mk_empty_array_with_capacity(lean_object*); +static lean_object* l_Lean_Elab_Tactic_addCheckpoints_push___closed__9; static lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Tactic_evalOpen___spec__13___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalIntros_declRange___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalRotateRight(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -81,7 +82,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalOpen_declRange___closed_ lean_object* l_Lean_LocalDecl_userName(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalIntro_declRange___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalAnyGoals_declRange___closed__3; -extern lean_object* l_Lean_nullKind; +lean_object* l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_findM_x3f___at___private_Lean_Elab_Tactic_BuiltinTactic_0__Lean_Elab_Tactic_findTag_x3f___spec__3___boxed(lean_object*, 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_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_evalOpen___spec__31___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSleep_declRange___closed__2; @@ -107,7 +108,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSubstVars_declRange___cl LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalChoiceAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_append___rarg(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalDbgTrace_declRange___closed__1; -static lean_object* l_Lean_Elab_Tactic_evalIntro___closed__41; static lean_object* l_Lean_Elab_Tactic_evalIntro___closed__13; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalDbgTrace___closed__1; lean_object* l_List_mapTRAux___at_Lean_resolveGlobalConstCore___spec__2(lean_object*, lean_object*); @@ -154,7 +154,6 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_renameInac lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalExact___spec__1___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalUnknown(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalRotateRight_declRange___closed__1; -static lean_object* l_Lean_Elab_Tactic_evalIntro___closed__43; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_evalDbgTrace(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalDone___closed__5; static lean_object* l_Lean_Elab_Tactic_evalCase___closed__1; @@ -170,10 +169,8 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalUnknown___boxed(lean_object*, le static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalIntro_declRange___closed__3; static lean_object* l___private_Lean_Elab_Tactic_BuiltinTactic_0__Lean_Elab_Tactic_getCaseGoals___closed__1; static lean_object* l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Tactic_evalOpen___spec__3___closed__2; -static lean_object* l_Lean_Elab_Tactic_evalIntro___closed__40; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalSubstVars___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSubstVars_declRange(lean_object*); -lean_object* l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_evalAllGoals(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_forEachVar(lean_object*, 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___regBuiltin_Lean_Elab_Tactic_elabSetOption_declRange(lean_object*); @@ -392,6 +389,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Tactic_elabSetOption_declRange___cl lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalOpen___boxed(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___private_Lean_Elab_Tactic_BuiltinTactic_0__Lean_Elab_Tactic_getCaseGoals___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_isNatLit_x3f(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSubstVars_declRange___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalTraceMessage___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalOpen___closed__1; @@ -407,6 +405,7 @@ lean_object* l_Lean_Elab_Term_expandMatchAltsIntoMatchTactic___boxed(lean_object static lean_object* l_Lean_Elab_Tactic_evalAssumption___rarg___closed__1; static lean_object* l_Lean_Elab_Tactic_evalIntro___closed__15; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalFailIfSuccess_declRange___closed__4; +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_evalIntro___closed__18; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_renameInaccessibles___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalFirst___closed__4; @@ -416,7 +415,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalIntros_declRange___close lean_object* lean_st_ref_take(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentArray_mapM___at_Lean_Elab_Tactic_renameInaccessibles___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getOptionDecl(lean_object*, lean_object*); -extern lean_object* l_Lean_numLitKind; +static lean_object* l_Lean_Elab_Tactic_evalFail___closed__10; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalCase___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_evalTraceState(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalFocus___closed__3; @@ -445,6 +444,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalFirst___boxed(lean_object*, lean static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSleep_declRange___closed__3; lean_object* l_Lean_Syntax_isStrLit_x3f(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalIntro_declRange___closed__4; +static lean_object* l_Lean_Elab_Tactic_addCheckpoints_push___closed__12; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalAnyGoals___lambda__1___boxed(lean_object*, 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___regBuiltin_Lean_Elab_Tactic_evalTacticSeq(lean_object*); static lean_object* l_Lean_Elab_Tactic_evalIntro___closed__11; @@ -650,7 +650,6 @@ static lean_object* l_Lean_Elab_Tactic_evalAnyGoals___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_evalRevert_declRange(lean_object*); lean_object* l_Lean_Meta_introNCore(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_BuiltinTactic_0__Lean_Elab_Tactic_getCaseGoals(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_evalIntro___closed__44; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalWithAnnotateState___lambda__1(lean_object*, 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_Tactic_evalFirst___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalCase_declRange___closed__1; @@ -721,6 +720,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalAssumption___boxed(lean_object*) LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalFirst_loop___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabSetOption___at_Lean_Elab_Tactic_elabSetOption___spec__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_forEachVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_addCheckpoints_push___closed__10; lean_object* l_ReaderT_pure___at_Lean_Elab_Tactic_saveTacticInfoForToken___spec__1___rarg___boxed(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_Tactic_evalRevert(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_Tactic_evalFocus(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -733,17 +733,14 @@ lean_object* l_Lean_mkAtomFrom(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalIntros___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalFocus_declRange___closed__5; static lean_object* l_Lean_resolveNamespace___at_Lean_Elab_Tactic_evalOpen___spec__6___closed__1; -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabSetOption___at_Lean_Elab_Tactic_elabSetOption___spec__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Open_0__Lean_Elab_OpenDecl_elabOpenHiding___at_Lean_Elab_Tactic_evalOpen___spec__17(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkFreshExprMVarAt(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_evalIntros___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalParen_declRange___closed__2; -extern lean_object* l_Lean_groupKind; extern lean_object* l_Lean_instInhabitedException; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalFocus_declRange___closed__3; static lean_object* l_Lean_Elab_Tactic_evalAnyGoals___closed__2; -lean_object* l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Meta_substVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_saveState___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_evalOpen___spec__16(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -760,6 +757,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalIntroMatch___lambda__1(lean_obje static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalFailIfSuccess_declRange___closed__5; lean_object* l_Lean_Elab_goalsToMessageData(lean_object*); lean_object* l_Std_PersistentArray_push___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_addCheckpoints_push___closed__11; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSleep(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalUnknown___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalFailIfSuccess___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -812,7 +810,6 @@ static lean_object* l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Tactic_eval lean_object* l_Lean_Meta_sortFVarIds(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalTacticSeq_declRange___closed__5; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_evalOpen___spec__2___closed__2; -static lean_object* l_Lean_Elab_Tactic_evalIntro___closed__42; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSleep___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalAnyGoals_declRange___closed__7; static lean_object* l_Lean_Elab_Tactic_evalIntro___closed__4; @@ -860,6 +857,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalOpen___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSkip___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalTacticSeqBracketed_declRange___closed__1; lean_object* l_Lean_KVMap_insertCore(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_evalFail___closed__9; static lean_object* l_Lean_Elab_Tactic_renameInaccessibles___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalWithAnnotateState___closed__5; LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Tactic_evalIntro___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*); @@ -960,6 +958,7 @@ LEAN_EXPORT lean_object* l_Lean_resolveUniqueNamespace___at_Lean_Elab_Tactic_eva static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalChoice_declRange___closed__3; uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTraceState___boxed(lean_object*); +lean_object* l_Lean_TSyntax_getString(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalDone(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_evalIntro(lean_object*); extern lean_object* l_Lean_Expr_instBEqExpr; @@ -2372,13 +2371,31 @@ return x_3; static lean_object* _init_l_Lean_Elab_Tactic_addCheckpoints_push___closed__3() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("null", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_addCheckpoints_push___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__3; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_addCheckpoints_push___closed__5() { +_start: +{ lean_object* x_1; lean_object* x_2; x_1 = lean_unsigned_to_nat(1u); x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic_addCheckpoints_push___closed__4() { +static lean_object* _init_l_Lean_Elab_Tactic_addCheckpoints_push___closed__6() { _start: { lean_object* x_1; @@ -2386,17 +2403,17 @@ x_1 = lean_mk_string_from_bytes("checkpoint", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_addCheckpoints_push___closed__5() { +static lean_object* _init_l_Lean_Elab_Tactic_addCheckpoints_push___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Tactic_evalWithAnnotateState___closed__6; -x_2 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__4; +x_2 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__6; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_addCheckpoints_push___closed__6() { +static lean_object* _init_l_Lean_Elab_Tactic_addCheckpoints_push___closed__8() { _start: { lean_object* x_1; lean_object* x_2; @@ -2405,7 +2422,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic_addCheckpoints_push___closed__7() { +static lean_object* _init_l_Lean_Elab_Tactic_addCheckpoints_push___closed__9() { _start: { lean_object* x_1; lean_object* x_2; @@ -2414,13 +2431,13 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic_addCheckpoints_push___closed__8() { +static lean_object* _init_l_Lean_Elab_Tactic_addCheckpoints_push___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean_nullKind; -x_3 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__7; +x_2 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__4; +x_3 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__9; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2428,6 +2445,24 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } +static lean_object* _init_l_Lean_Elab_Tactic_addCheckpoints_push___closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("group", 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_addCheckpoints_push___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__11; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_addCheckpoints_push(lean_object* x_1, lean_object* x_2) { _start: { @@ -2439,32 +2474,32 @@ lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_obj x_4 = l_Lean_instInhabitedSyntax; x_5 = l_Array_back___rarg(x_4, x_1); x_6 = lean_box(2); -x_7 = l_Lean_nullKind; +x_7 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__4; x_8 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_8, 0, x_6); lean_ctor_set(x_8, 1, x_7); lean_ctor_set(x_8, 2, x_1); -x_9 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__3; +x_9 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__5; x_10 = lean_array_push(x_9, x_8); x_11 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__2; x_12 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_12, 0, x_6); lean_ctor_set(x_12, 1, x_11); lean_ctor_set(x_12, 2, x_10); -x_13 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__4; +x_13 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__6; x_14 = l_Lean_mkAtomFrom(x_5, x_13); -x_15 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__6; +x_15 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__8; x_16 = lean_array_push(x_15, x_14); x_17 = lean_array_push(x_16, x_12); -x_18 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__5; +x_18 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__7; x_19 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_19, 0, x_6); lean_ctor_set(x_19, 1, x_18); lean_ctor_set(x_19, 2, x_17); x_20 = lean_array_push(x_15, x_19); -x_21 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__8; +x_21 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__10; x_22 = lean_array_push(x_20, x_21); -x_23 = l_Lean_groupKind; +x_23 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__12; x_24 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_24, 0, x_6); lean_ctor_set(x_24, 1, x_23); @@ -2530,10 +2565,10 @@ lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean x_29 = lean_ctor_get(x_21, 1); lean_inc(x_29); lean_dec(x_21); -x_30 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__3; +x_30 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__5; x_31 = lean_array_push(x_30, x_20); x_32 = lean_box(2); -x_33 = l_Lean_nullKind; +x_33 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__4; x_34 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_34, 0, x_32); lean_ctor_set(x_34, 1, x_33); @@ -2544,7 +2579,7 @@ x_37 = lean_unsigned_to_nat(1u); x_38 = lean_nat_add(x_2, x_37); lean_dec(x_2); x_39 = l_Lean_Elab_Tactic_addCheckpoints_push(x_36, x_4); -x_40 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__7; +x_40 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__9; x_2 = x_38; x_3 = x_40; x_4 = x_39; @@ -2721,7 +2756,7 @@ lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; x_26 = lean_ctor_get(x_19, 1); lean_inc(x_26); lean_dec(x_19); -x_27 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__7; +x_27 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__9; x_28 = l_Lean_Elab_Tactic_addCheckpoints_go(x_1, x_12, x_27, x_27, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_26); lean_dec(x_1); x_29 = !lean_is_exclusive(x_28); @@ -5044,7 +5079,7 @@ x_14 = lean_ctor_get(x_12, 1); lean_inc(x_14); lean_dec(x_12); x_15 = lean_box(0); -x_16 = l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(x_13, x_15); +x_16 = l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(x_13, x_15); x_17 = l_List_isEmpty___rarg(x_16); if (x_17 == 0) { @@ -6638,7 +6673,7 @@ static lean_object* _init_l___private_Lean_Elab_Open_0__Lean_Elab_OpenDecl_resol _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__7; +x_1 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__9; x_2 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_2, 0, x_1); lean_ctor_set(x_2, 1, x_1); @@ -10348,39 +10383,38 @@ x_23 = lean_erase_macro_scopes(x_22); x_24 = l_Lean_Syntax_isStrLit_x3f(x_2); if (lean_obj_tag(x_24) == 0) { -lean_object* x_25; lean_object* x_26; -x_25 = l_Lean_numLitKind; -x_26 = l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(x_25, x_2); -if (lean_obj_tag(x_26) == 0) +lean_object* x_25; +x_25 = l_Lean_Syntax_isNatLit_x3f(x_2); +if (lean_obj_tag(x_25) == 0) { if (lean_obj_tag(x_2) == 2) { -lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_27 = lean_ctor_get(x_2, 1); -lean_inc(x_27); -x_28 = l_Lean_Elab_elabSetOption___at_Lean_Elab_Tactic_elabSetOption___spec__1___closed__3; -x_29 = lean_string_dec_eq(x_27, x_28); -if (x_29 == 0) +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = lean_ctor_get(x_2, 1); +lean_inc(x_26); +x_27 = l_Lean_Elab_elabSetOption___at_Lean_Elab_Tactic_elabSetOption___spec__1___closed__3; +x_28 = lean_string_dec_eq(x_26, x_27); +if (x_28 == 0) { -lean_object* x_30; uint8_t x_31; -x_30 = l_Lean_Elab_elabSetOption___at_Lean_Elab_Tactic_elabSetOption___spec__1___closed__4; -x_31 = lean_string_dec_eq(x_27, x_30); -lean_dec(x_27); -if (x_31 == 0) +lean_object* x_29; uint8_t x_30; +x_29 = l_Lean_Elab_elabSetOption___at_Lean_Elab_Tactic_elabSetOption___spec__1___closed__4; +x_30 = lean_string_dec_eq(x_26, x_29); +lean_dec(x_26); +if (x_30 == 0) { -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_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_dec(x_23); -x_32 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_32, 0, x_2); -x_33 = l_Lean_Elab_elabSetOption___at_Lean_Elab_Tactic_elabSetOption___spec__1___closed__2; -x_34 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_32); -x_35 = l_Lean_Elab_nestedExceptionToMessageData___at_Lean_Elab_Tactic_evalOpen___spec__23___closed__1; -x_36 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -x_37 = l_Lean_throwError___at_Lean_Elab_Tactic_elabSetOption___spec__5(x_36, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21); +x_31 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_31, 0, x_2); +x_32 = l_Lean_Elab_elabSetOption___at_Lean_Elab_Tactic_elabSetOption___spec__1___closed__2; +x_33 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_31); +x_34 = l_Lean_Elab_nestedExceptionToMessageData___at_Lean_Elab_Tactic_evalOpen___spec__23___closed__1; +x_35 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +x_36 = l_Lean_throwError___at_Lean_Elab_Tactic_elabSetOption___spec__5(x_35, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -10389,42 +10423,42 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_37; +return x_36; } else { -lean_object* x_38; lean_object* x_39; +lean_object* x_37; lean_object* x_38; lean_dec(x_2); -x_38 = l_Lean_Elab_elabSetOption___at_Lean_Elab_Tactic_elabSetOption___spec__1___closed__5; -x_39 = l_Lean_Elab_elabSetOption_setOption___at_Lean_Elab_Tactic_elabSetOption___spec__6(x_23, x_38, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21); -return x_39; +x_37 = l_Lean_Elab_elabSetOption___at_Lean_Elab_Tactic_elabSetOption___spec__1___closed__5; +x_38 = l_Lean_Elab_elabSetOption_setOption___at_Lean_Elab_Tactic_elabSetOption___spec__6(x_23, x_37, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21); +return x_38; } } else { -lean_object* x_40; lean_object* x_41; -lean_dec(x_27); +lean_object* x_39; lean_object* x_40; +lean_dec(x_26); lean_dec(x_2); -x_40 = l_Lean_Elab_elabSetOption___at_Lean_Elab_Tactic_elabSetOption___spec__1___closed__6; -x_41 = l_Lean_Elab_elabSetOption_setOption___at_Lean_Elab_Tactic_elabSetOption___spec__6(x_23, x_40, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21); -return x_41; +x_39 = l_Lean_Elab_elabSetOption___at_Lean_Elab_Tactic_elabSetOption___spec__1___closed__6; +x_40 = l_Lean_Elab_elabSetOption_setOption___at_Lean_Elab_Tactic_elabSetOption___spec__6(x_23, x_39, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21); +return x_40; } } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_dec(x_23); -x_42 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_42, 0, x_2); -x_43 = l_Lean_Elab_elabSetOption___at_Lean_Elab_Tactic_elabSetOption___spec__1___closed__2; -x_44 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_44, 1, x_42); -x_45 = l_Lean_Elab_nestedExceptionToMessageData___at_Lean_Elab_Tactic_evalOpen___spec__23___closed__1; -x_46 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -x_47 = l_Lean_throwError___at_Lean_Elab_Tactic_elabSetOption___spec__5(x_46, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21); +x_41 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_41, 0, x_2); +x_42 = l_Lean_Elab_elabSetOption___at_Lean_Elab_Tactic_elabSetOption___spec__1___closed__2; +x_43 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_41); +x_44 = l_Lean_Elab_nestedExceptionToMessageData___at_Lean_Elab_Tactic_evalOpen___spec__23___closed__1; +x_45 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +x_46 = l_Lean_throwError___at_Lean_Elab_Tactic_elabSetOption___spec__5(x_45, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -10433,33 +10467,33 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_47; +return x_46; } } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_dec(x_2); -x_48 = lean_ctor_get(x_26, 0); -lean_inc(x_48); -lean_dec(x_26); -x_49 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_49, 0, x_48); -x_50 = l_Lean_Elab_elabSetOption_setOption___at_Lean_Elab_Tactic_elabSetOption___spec__6(x_23, x_49, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21); -return x_50; +x_47 = lean_ctor_get(x_25, 0); +lean_inc(x_47); +lean_dec(x_25); +x_48 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_48, 0, x_47); +x_49 = l_Lean_Elab_elabSetOption_setOption___at_Lean_Elab_Tactic_elabSetOption___spec__6(x_23, x_48, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21); +return x_49; } } else { -lean_object* x_51; lean_object* x_52; lean_object* x_53; +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_dec(x_2); -x_51 = lean_ctor_get(x_24, 0); -lean_inc(x_51); +x_50 = lean_ctor_get(x_24, 0); +lean_inc(x_50); lean_dec(x_24); -x_52 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_52, 0, x_51); -x_53 = l_Lean_Elab_elabSetOption_setOption___at_Lean_Elab_Tactic_elabSetOption___spec__6(x_23, x_52, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21); -return x_53; +x_51 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_51, 0, x_50); +x_52 = l_Lean_Elab_elabSetOption_setOption___at_Lean_Elab_Tactic_elabSetOption___spec__6(x_23, x_51, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21); +return x_52; } } } @@ -11258,7 +11292,7 @@ x_13 = lean_ctor_get(x_11, 1); lean_inc(x_13); lean_dec(x_11); x_14 = lean_box(0); -x_15 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__7; +x_15 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__9; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); @@ -11958,7 +11992,7 @@ static lean_object* _init_l_Lean_Elab_Tactic_evalAnyGoals___closed__1() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__7; +x_1 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__9; x_2 = 0; x_3 = lean_box(x_2); x_4 = lean_alloc_ctor(0, 2, 0); @@ -15508,29 +15542,11 @@ static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("null", 4); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic_evalIntro___closed__3; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__5() { -_start: -{ -lean_object* x_1; x_1 = lean_mk_string_from_bytes(";", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__6() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__4() { _start: { lean_object* x_1; lean_object* x_2; @@ -15539,7 +15555,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__7() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__5() { _start: { lean_object* x_1; @@ -15547,17 +15563,17 @@ x_1 = lean_mk_string_from_bytes("ident", 5); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__8() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic_evalIntro___closed__7; +x_2 = l_Lean_Elab_Tactic_evalIntro___closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__9() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__7() { _start: { lean_object* x_1; @@ -15565,17 +15581,17 @@ x_1 = lean_mk_string_from_bytes("Term", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__10() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Tactic_evalWithAnnotateState___closed__4; -x_2 = l_Lean_Elab_Tactic_evalIntro___closed__9; +x_2 = l_Lean_Elab_Tactic_evalIntro___closed__7; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__11() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__9() { _start: { lean_object* x_1; @@ -15583,17 +15599,17 @@ x_1 = lean_mk_string_from_bytes("hole", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__12() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_evalIntro___closed__10; -x_2 = l_Lean_Elab_Tactic_evalIntro___closed__11; +x_1 = l_Lean_Elab_Tactic_evalIntro___closed__8; +x_2 = l_Lean_Elab_Tactic_evalIntro___closed__9; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__13() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__11() { _start: { lean_object* x_1; @@ -15601,22 +15617,22 @@ x_1 = lean_mk_string_from_bytes("h", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__14() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic_evalIntro___closed__13; +x_1 = l_Lean_Elab_Tactic_evalIntro___closed__11; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__15() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Tactic_evalIntro___closed__13; +x_1 = l_Lean_Elab_Tactic_evalIntro___closed__11; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Tactic_evalIntro___closed__14; +x_3 = l_Lean_Elab_Tactic_evalIntro___closed__12; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -15624,17 +15640,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__16() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic_evalIntro___closed__13; +x_2 = l_Lean_Elab_Tactic_evalIntro___closed__11; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__17() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__15() { _start: { lean_object* x_1; @@ -15642,31 +15658,17 @@ x_1 = lean_mk_string_from_bytes("match", 5); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__18() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Tactic_evalWithAnnotateState___closed__6; -x_2 = l_Lean_Elab_Tactic_evalIntro___closed__17; +x_2 = l_Lean_Elab_Tactic_evalIntro___closed__15; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__19() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = lean_box(2); -x_2 = l_Lean_Elab_Tactic_evalIntro___closed__4; -x_3 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__7; -x_4 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_4, 0, x_1); -lean_ctor_set(x_4, 1, x_2); -lean_ctor_set(x_4, 2, x_3); -return x_4; -} -} -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__20() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__17() { _start: { lean_object* x_1; @@ -15674,27 +15676,27 @@ x_1 = lean_mk_string_from_bytes("matchDiscr", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__21() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_evalIntro___closed__10; -x_2 = l_Lean_Elab_Tactic_evalIntro___closed__20; +x_1 = l_Lean_Elab_Tactic_evalIntro___closed__8; +x_2 = l_Lean_Elab_Tactic_evalIntro___closed__17; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__22() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__6; -x_2 = l_Lean_Elab_Tactic_evalIntro___closed__19; +x_1 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__8; +x_2 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__10; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__23() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__20() { _start: { lean_object* x_1; @@ -15702,7 +15704,7 @@ x_1 = lean_mk_string_from_bytes("with", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__24() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__21() { _start: { lean_object* x_1; @@ -15710,17 +15712,17 @@ x_1 = lean_mk_string_from_bytes("matchAlts", 9); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__25() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_evalIntro___closed__10; -x_2 = l_Lean_Elab_Tactic_evalIntro___closed__24; +x_1 = l_Lean_Elab_Tactic_evalIntro___closed__8; +x_2 = l_Lean_Elab_Tactic_evalIntro___closed__21; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__26() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__23() { _start: { lean_object* x_1; @@ -15728,17 +15730,17 @@ x_1 = lean_mk_string_from_bytes("matchAlt", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__27() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_evalIntro___closed__10; -x_2 = l_Lean_Elab_Tactic_evalIntro___closed__26; +x_1 = l_Lean_Elab_Tactic_evalIntro___closed__8; +x_2 = l_Lean_Elab_Tactic_evalIntro___closed__23; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__28() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__25() { _start: { lean_object* x_1; @@ -15746,7 +15748,7 @@ x_1 = lean_mk_string_from_bytes("|", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__29() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__26() { _start: { lean_object* x_1; @@ -15754,7 +15756,7 @@ x_1 = lean_mk_string_from_bytes("=>", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__30() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__27() { _start: { lean_object* x_1; @@ -15762,17 +15764,17 @@ x_1 = lean_mk_string_from_bytes("syntheticHole", 13); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__31() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_evalIntro___closed__10; -x_2 = l_Lean_Elab_Tactic_evalIntro___closed__30; +x_1 = l_Lean_Elab_Tactic_evalIntro___closed__8; +x_2 = l_Lean_Elab_Tactic_evalIntro___closed__27; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__32() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__29() { _start: { lean_object* x_1; @@ -15780,7 +15782,7 @@ x_1 = lean_mk_string_from_bytes("?", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__33() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__30() { _start: { lean_object* x_1; @@ -15788,7 +15790,7 @@ x_1 = lean_mk_string_from_bytes("_", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__34() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__31() { _start: { lean_object* x_1; lean_object* x_2; @@ -15797,7 +15799,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__35() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__32() { _start: { lean_object* x_1; lean_object* x_2; @@ -15806,7 +15808,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__36() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__33() { _start: { lean_object* x_1; @@ -15814,17 +15816,17 @@ x_1 = lean_mk_string_from_bytes("tacticTry_", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__37() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__34() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Tactic_evalWithAnnotateState___closed__6; -x_2 = l_Lean_Elab_Tactic_evalIntro___closed__36; +x_2 = l_Lean_Elab_Tactic_evalIntro___closed__33; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__38() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__35() { _start: { lean_object* x_1; @@ -15832,25 +15834,7 @@ x_1 = lean_mk_string_from_bytes("try", 3); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__39() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("group", 5); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__40() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic_evalIntro___closed__39; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__41() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__36() { _start: { lean_object* x_1; @@ -15858,17 +15842,17 @@ x_1 = lean_mk_string_from_bytes("clear", 5); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__42() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__37() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Tactic_evalWithAnnotateState___closed__6; -x_2 = l_Lean_Elab_Tactic_evalIntro___closed__41; +x_2 = l_Lean_Elab_Tactic_evalIntro___closed__36; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__43() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__38() { _start: { lean_object* x_1; lean_object* x_2; @@ -15877,12 +15861,12 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__44() { +static lean_object* _init_l_Lean_Elab_Tactic_evalIntro___closed__39() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic_evalIntro___closed__33; +x_2 = l_Lean_Elab_Tactic_evalIntro___closed__30; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } @@ -15911,28 +15895,27 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(0u); +x_16 = lean_unsigned_to_nat(0u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -uint8_t x_19; +uint8_t x_18; lean_inc(x_15); -x_19 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_14); -if (x_19 == 0) +x_18 = l_Lean_Syntax_matchesNull(x_15, x_14); +if (x_18 == 0) { -lean_object* x_20; uint8_t x_21; -x_20 = l_Lean_Syntax_getNumArgs(x_15); -x_21 = lean_nat_dec_le(x_14, x_20); -if (x_21 == 0) +lean_object* x_19; uint8_t x_20; +x_19 = l_Lean_Syntax_getNumArgs(x_15); +x_20 = lean_nat_dec_le(x_14, x_19); +if (x_20 == 0) { -lean_object* x_22; -lean_dec(x_20); +lean_object* x_21; +lean_dec(x_19); lean_dec(x_15); lean_dec(x_9); lean_dec(x_8); @@ -15942,23 +15925,24 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_22 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalExact___spec__1___rarg(x_10); -return x_22; +x_21 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalExact___spec__1___rarg(x_10); +return x_21; } else { -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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_23 = l_Lean_Syntax_getArg(x_15, x_17); -x_24 = l_Lean_Syntax_getArgs(x_15); +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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_22 = l_Lean_Syntax_getArg(x_15, x_16); +x_23 = l_Lean_Syntax_getArgs(x_15); lean_dec(x_15); -x_25 = lean_nat_sub(x_20, x_17); -lean_dec(x_20); -x_26 = l_Array_extract___rarg(x_24, x_14, x_25); -x_27 = lean_box(2); +x_24 = lean_nat_sub(x_19, x_16); +lean_dec(x_19); +x_25 = l_Array_extract___rarg(x_23, x_14, x_24); +x_26 = lean_box(2); +x_27 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__4; x_28 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_28, 0, x_27); -lean_ctor_set(x_28, 1, x_16); -lean_ctor_set(x_28, 2, x_26); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +lean_ctor_set(x_28, 2, x_25); x_29 = l_Lean_Syntax_getArgs(x_28); lean_dec(x_28); lean_inc(x_8); @@ -15977,325 +15961,324 @@ lean_inc(x_31); x_36 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_36, 0, x_31); lean_ctor_set(x_36, 1, x_35); -x_37 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__3; -x_38 = lean_array_push(x_37, x_23); -x_39 = l_Lean_Elab_Tactic_evalIntro___closed__4; -x_40 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_40, 0, x_27); -lean_ctor_set(x_40, 1, x_39); -lean_ctor_set(x_40, 2, x_38); -x_41 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__6; -x_42 = lean_array_push(x_41, x_36); -lean_inc(x_42); -x_43 = lean_array_push(x_42, x_40); -x_44 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_44, 0, x_27); -lean_ctor_set(x_44, 1, x_11); -lean_ctor_set(x_44, 2, x_43); -x_45 = l_Lean_Elab_Tactic_evalIntro___closed__5; -x_46 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_46, 0, x_31); -lean_ctor_set(x_46, 1, x_45); -x_47 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__7; -x_48 = l_Array_append___rarg(x_47, x_29); -x_49 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_49, 0, x_27); -lean_ctor_set(x_49, 1, x_39); -lean_ctor_set(x_49, 2, x_48); -x_50 = lean_array_push(x_42, x_49); -x_51 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_51, 0, x_27); -lean_ctor_set(x_51, 1, x_11); -lean_ctor_set(x_51, 2, x_50); -x_52 = l_Lean_Elab_Tactic_evalIntro___closed__6; -x_53 = lean_array_push(x_52, x_44); -x_54 = lean_array_push(x_53, x_46); -x_55 = lean_array_push(x_54, x_51); -x_56 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_56, 0, x_27); -lean_ctor_set(x_56, 1, x_39); -lean_ctor_set(x_56, 2, x_55); -x_57 = lean_array_push(x_37, x_56); -x_58 = l___regBuiltin_Lean_Elab_Tactic_evalSeq1___closed__2; -x_59 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_59, 0, x_27); -lean_ctor_set(x_59, 1, x_58); -lean_ctor_set(x_59, 2, x_57); -x_60 = l_Lean_Elab_Tactic_evalTacticAux(x_59, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_34); -return x_60; -} -} +x_37 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__5; +x_38 = lean_array_push(x_37, x_22); +x_39 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_39, 0, x_26); +lean_ctor_set(x_39, 1, x_27); +lean_ctor_set(x_39, 2, x_38); +x_40 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__8; +x_41 = lean_array_push(x_40, x_36); +lean_inc(x_41); +x_42 = lean_array_push(x_41, x_39); +x_43 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_43, 0, x_26); +lean_ctor_set(x_43, 1, x_11); +lean_ctor_set(x_43, 2, x_42); +x_44 = l_Lean_Elab_Tactic_evalIntro___closed__3; +x_45 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_45, 0, x_31); +lean_ctor_set(x_45, 1, x_44); +x_46 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__9; +x_47 = l_Array_append___rarg(x_46, x_29); +x_48 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_48, 0, x_26); +lean_ctor_set(x_48, 1, x_27); +lean_ctor_set(x_48, 2, x_47); +x_49 = lean_array_push(x_41, x_48); +x_50 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_50, 0, x_26); +lean_ctor_set(x_50, 1, x_11); +lean_ctor_set(x_50, 2, x_49); +x_51 = l_Lean_Elab_Tactic_evalIntro___closed__4; +x_52 = lean_array_push(x_51, x_43); +x_53 = lean_array_push(x_52, x_45); +x_54 = lean_array_push(x_53, x_50); +x_55 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_55, 0, x_26); +lean_ctor_set(x_55, 1, x_27); +lean_ctor_set(x_55, 2, x_54); +x_56 = lean_array_push(x_37, x_55); +x_57 = l___regBuiltin_Lean_Elab_Tactic_evalSeq1___closed__2; +x_58 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_58, 0, x_26); +lean_ctor_set(x_58, 1, x_57); +lean_ctor_set(x_58, 2, x_56); +x_59 = l_Lean_Elab_Tactic_evalTacticAux(x_58, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_34); +return x_59; +} +} else { -lean_object* x_61; lean_object* x_62; uint8_t x_63; -x_61 = l_Lean_Syntax_getArg(x_15, x_17); +lean_object* x_60; lean_object* x_61; uint8_t x_62; +x_60 = l_Lean_Syntax_getArg(x_15, x_16); lean_dec(x_15); -x_62 = l_Lean_Elab_Tactic_evalIntro___closed__8; -lean_inc(x_61); -x_63 = l_Lean_Syntax_isOfKind(x_61, x_62); -if (x_63 == 0) +x_61 = l_Lean_Elab_Tactic_evalIntro___closed__6; +lean_inc(x_60); +x_62 = l_Lean_Syntax_isOfKind(x_60, x_61); +if (x_62 == 0) { -lean_object* x_64; uint8_t x_65; -x_64 = l_Lean_Elab_Tactic_evalIntro___closed__12; -lean_inc(x_61); -x_65 = l_Lean_Syntax_isOfKind(x_61, x_64); -if (x_65 == 0) +lean_object* x_63; uint8_t x_64; +x_63 = l_Lean_Elab_Tactic_evalIntro___closed__10; +lean_inc(x_60); +x_64 = l_Lean_Syntax_isOfKind(x_60, x_63); +if (x_64 == 0) { -lean_object* x_66; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; +lean_object* x_65; lean_object* x_66; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_inc(x_8); -x_66 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Tactic_evalIntro___spec__1___rarg(x_8, x_9, x_10); -x_67 = lean_ctor_get(x_66, 0); +x_65 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Tactic_evalIntro___spec__1___rarg(x_8, x_9, x_10); +x_66 = lean_ctor_get(x_65, 0); +lean_inc(x_66); +x_67 = lean_ctor_get(x_65, 1); lean_inc(x_67); -x_68 = lean_ctor_get(x_66, 1); +lean_dec(x_65); +x_68 = lean_ctor_get(x_8, 10); lean_inc(x_68); -lean_dec(x_66); -x_69 = lean_ctor_get(x_8, 10); -lean_inc(x_69); -x_70 = lean_st_ref_get(x_9, x_68); -x_71 = lean_ctor_get(x_70, 0); +x_69 = lean_st_ref_get(x_9, x_67); +x_70 = lean_ctor_get(x_69, 0); +lean_inc(x_70); +x_71 = lean_ctor_get(x_69, 1); lean_inc(x_71); -x_72 = lean_ctor_get(x_70, 1); +lean_dec(x_69); +x_72 = lean_ctor_get(x_70, 0); lean_inc(x_72); lean_dec(x_70); -x_73 = lean_ctor_get(x_71, 0); -lean_inc(x_73); -lean_dec(x_71); -x_74 = lean_environment_main_module(x_73); -x_75 = l_Lean_Elab_Tactic_evalIntro___closed__1; -lean_inc(x_67); -x_76 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_76, 0, x_67); -lean_ctor_set(x_76, 1, x_75); -x_77 = l_Lean_Elab_Tactic_evalIntro___closed__16; -x_78 = l_Lean_addMacroScope(x_74, x_77, x_69); -x_79 = lean_box(0); -x_80 = l_Lean_Elab_Tactic_evalIntro___closed__15; -lean_inc(x_67); -x_81 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_81, 0, x_67); -lean_ctor_set(x_81, 1, x_80); -lean_ctor_set(x_81, 2, x_78); -lean_ctor_set(x_81, 3, x_79); -x_82 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__3; -lean_inc(x_81); -x_83 = lean_array_push(x_82, x_81); -x_84 = lean_box(2); -x_85 = l_Lean_Elab_Tactic_evalIntro___closed__4; -x_86 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_86, 0, x_84); -lean_ctor_set(x_86, 1, x_85); -lean_ctor_set(x_86, 2, x_83); -x_87 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__6; -x_88 = lean_array_push(x_87, x_76); -lean_inc(x_86); -x_89 = lean_array_push(x_88, x_86); -x_90 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_90, 0, x_84); -lean_ctor_set(x_90, 1, x_11); -lean_ctor_set(x_90, 2, x_89); -x_91 = l_Lean_Elab_Tactic_evalIntro___closed__5; -lean_inc(x_67); -x_92 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_92, 0, x_67); -lean_ctor_set(x_92, 1, x_91); -x_93 = l_Lean_Elab_Tactic_evalIntro___closed__17; -lean_inc(x_67); -x_94 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_94, 0, x_67); -lean_ctor_set(x_94, 1, x_93); -x_95 = l_Lean_Elab_Tactic_evalIntro___closed__22; -x_96 = lean_array_push(x_95, x_81); -x_97 = l_Lean_Elab_Tactic_evalIntro___closed__21; -x_98 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_98, 0, x_84); -lean_ctor_set(x_98, 1, x_97); -lean_ctor_set(x_98, 2, x_96); -x_99 = lean_array_push(x_82, x_98); -x_100 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_100, 0, x_84); -lean_ctor_set(x_100, 1, x_85); -lean_ctor_set(x_100, 2, x_99); -x_101 = l_Lean_Elab_Tactic_evalIntro___closed__23; -lean_inc(x_67); -x_102 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_102, 0, x_67); -lean_ctor_set(x_102, 1, x_101); -x_103 = l_Lean_Elab_Tactic_evalIntro___closed__28; -lean_inc(x_67); -x_104 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_104, 0, x_67); -lean_ctor_set(x_104, 1, x_103); -x_105 = lean_array_push(x_82, x_61); -x_106 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_106, 0, x_84); -lean_ctor_set(x_106, 1, x_85); -lean_ctor_set(x_106, 2, x_105); -x_107 = lean_array_push(x_82, x_106); -x_108 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_108, 0, x_84); -lean_ctor_set(x_108, 1, x_85); -lean_ctor_set(x_108, 2, x_107); -x_109 = l_Lean_Elab_Tactic_evalIntro___closed__29; -lean_inc(x_67); -x_110 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_110, 0, x_67); -lean_ctor_set(x_110, 1, x_109); -x_111 = l_Lean_Elab_Tactic_evalIntro___closed__32; -lean_inc(x_67); -x_112 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_112, 0, x_67); -lean_ctor_set(x_112, 1, x_111); -x_113 = l_Lean_Elab_Tactic_evalIntro___closed__33; -lean_inc(x_67); -x_114 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_114, 0, x_67); -lean_ctor_set(x_114, 1, x_113); -x_115 = lean_array_push(x_82, x_114); -x_116 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_116, 0, x_84); -lean_ctor_set(x_116, 1, x_64); -lean_ctor_set(x_116, 2, x_115); -x_117 = lean_array_push(x_87, x_112); -x_118 = lean_array_push(x_117, x_116); -x_119 = l_Lean_Elab_Tactic_evalIntro___closed__31; -x_120 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_120, 0, x_84); -lean_ctor_set(x_120, 1, x_119); -lean_ctor_set(x_120, 2, x_118); -x_121 = l_Lean_Elab_Tactic_evalIntro___closed__34; -x_122 = lean_array_push(x_121, x_104); -x_123 = lean_array_push(x_122, x_108); -x_124 = lean_array_push(x_123, x_110); -x_125 = lean_array_push(x_124, x_120); -x_126 = l_Lean_Elab_Tactic_evalIntro___closed__27; -x_127 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_127, 0, x_84); -lean_ctor_set(x_127, 1, x_126); -lean_ctor_set(x_127, 2, x_125); -x_128 = lean_array_push(x_82, x_127); -x_129 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_129, 0, x_84); -lean_ctor_set(x_129, 1, x_85); -lean_ctor_set(x_129, 2, x_128); -x_130 = lean_array_push(x_82, x_129); -x_131 = l_Lean_Elab_Tactic_evalIntro___closed__25; -x_132 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_132, 0, x_84); -lean_ctor_set(x_132, 1, x_131); -lean_ctor_set(x_132, 2, x_130); -x_133 = l_Lean_Elab_Tactic_evalIntro___closed__35; -x_134 = lean_array_push(x_133, x_94); -x_135 = l_Lean_Elab_Tactic_evalIntro___closed__19; -x_136 = lean_array_push(x_134, x_135); -x_137 = lean_array_push(x_136, x_135); -x_138 = lean_array_push(x_137, x_100); -x_139 = lean_array_push(x_138, x_102); -x_140 = lean_array_push(x_139, x_132); -x_141 = l_Lean_Elab_Tactic_evalIntro___closed__18; -x_142 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_142, 0, x_84); -lean_ctor_set(x_142, 1, x_141); -lean_ctor_set(x_142, 2, x_140); -x_143 = l_Lean_Elab_Tactic_evalIntro___closed__38; -lean_inc(x_67); -x_144 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_144, 0, x_67); -lean_ctor_set(x_144, 1, x_143); -x_145 = l_Lean_Elab_Tactic_evalIntro___closed__41; -x_146 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_146, 0, x_67); -lean_ctor_set(x_146, 1, x_145); -x_147 = lean_array_push(x_87, x_146); -x_148 = lean_array_push(x_147, x_86); -x_149 = l_Lean_Elab_Tactic_evalIntro___closed__42; -x_150 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_150, 0, x_84); -lean_ctor_set(x_150, 1, x_149); -lean_ctor_set(x_150, 2, x_148); -x_151 = lean_array_push(x_87, x_150); -x_152 = lean_array_push(x_151, x_135); -x_153 = l_Lean_Elab_Tactic_evalIntro___closed__40; -x_154 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_154, 0, x_84); -lean_ctor_set(x_154, 1, x_153); -lean_ctor_set(x_154, 2, x_152); -x_155 = lean_array_push(x_82, x_154); -x_156 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_156, 0, x_84); -lean_ctor_set(x_156, 1, x_85); -lean_ctor_set(x_156, 2, x_155); -x_157 = lean_array_push(x_82, x_156); -x_158 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__2; -x_159 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_159, 0, x_84); -lean_ctor_set(x_159, 1, x_158); -lean_ctor_set(x_159, 2, x_157); -x_160 = lean_array_push(x_82, x_159); -x_161 = l___regBuiltin_Lean_Elab_Tactic_evalTacticSeq___closed__2; -x_162 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_162, 0, x_84); -lean_ctor_set(x_162, 1, x_161); -lean_ctor_set(x_162, 2, x_160); -x_163 = lean_array_push(x_87, x_144); -x_164 = lean_array_push(x_163, x_162); -x_165 = l_Lean_Elab_Tactic_evalIntro___closed__37; -x_166 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_166, 0, x_84); -lean_ctor_set(x_166, 1, x_165); -lean_ctor_set(x_166, 2, x_164); -x_167 = l_Lean_Elab_Tactic_evalIntro___closed__43; -x_168 = lean_array_push(x_167, x_90); -lean_inc(x_92); -x_169 = lean_array_push(x_168, x_92); -x_170 = lean_array_push(x_169, x_142); -x_171 = lean_array_push(x_170, x_92); -x_172 = lean_array_push(x_171, x_166); -x_173 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_173, 0, x_84); -lean_ctor_set(x_173, 1, x_85); -lean_ctor_set(x_173, 2, x_172); -x_174 = lean_array_push(x_82, x_173); -x_175 = l___regBuiltin_Lean_Elab_Tactic_evalSeq1___closed__2; -x_176 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_176, 0, x_84); -lean_ctor_set(x_176, 1, x_175); -lean_ctor_set(x_176, 2, x_174); -x_177 = l_Lean_Elab_Tactic_evalTacticAux(x_176, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_72); -return x_177; -} -else -{ -lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; -x_178 = l_Lean_Syntax_getArg(x_61, x_17); -lean_dec(x_61); -x_179 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_179, 0, x_178); -x_180 = l_Lean_Elab_Tactic_evalIntro___closed__44; -x_181 = l_Lean_Elab_Tactic_evalIntro_introStep(x_179, x_180, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_181; +x_73 = lean_environment_main_module(x_72); +x_74 = l_Lean_Elab_Tactic_evalIntro___closed__1; +lean_inc(x_66); +x_75 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_75, 0, x_66); +lean_ctor_set(x_75, 1, x_74); +x_76 = l_Lean_Elab_Tactic_evalIntro___closed__14; +x_77 = l_Lean_addMacroScope(x_73, x_76, x_68); +x_78 = lean_box(0); +x_79 = l_Lean_Elab_Tactic_evalIntro___closed__13; +lean_inc(x_66); +x_80 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_80, 0, x_66); +lean_ctor_set(x_80, 1, x_79); +lean_ctor_set(x_80, 2, x_77); +lean_ctor_set(x_80, 3, x_78); +x_81 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__5; +lean_inc(x_80); +x_82 = lean_array_push(x_81, x_80); +x_83 = lean_box(2); +x_84 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__4; +x_85 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_85, 0, x_83); +lean_ctor_set(x_85, 1, x_84); +lean_ctor_set(x_85, 2, x_82); +x_86 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__8; +x_87 = lean_array_push(x_86, x_75); +lean_inc(x_85); +x_88 = lean_array_push(x_87, x_85); +x_89 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_89, 0, x_83); +lean_ctor_set(x_89, 1, x_11); +lean_ctor_set(x_89, 2, x_88); +x_90 = l_Lean_Elab_Tactic_evalIntro___closed__3; +lean_inc(x_66); +x_91 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_91, 0, x_66); +lean_ctor_set(x_91, 1, x_90); +x_92 = l_Lean_Elab_Tactic_evalIntro___closed__15; +lean_inc(x_66); +x_93 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_93, 0, x_66); +lean_ctor_set(x_93, 1, x_92); +x_94 = l_Lean_Elab_Tactic_evalIntro___closed__19; +x_95 = lean_array_push(x_94, x_80); +x_96 = l_Lean_Elab_Tactic_evalIntro___closed__18; +x_97 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_97, 0, x_83); +lean_ctor_set(x_97, 1, x_96); +lean_ctor_set(x_97, 2, x_95); +x_98 = lean_array_push(x_81, x_97); +x_99 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_99, 0, x_83); +lean_ctor_set(x_99, 1, x_84); +lean_ctor_set(x_99, 2, x_98); +x_100 = l_Lean_Elab_Tactic_evalIntro___closed__20; +lean_inc(x_66); +x_101 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_101, 0, x_66); +lean_ctor_set(x_101, 1, x_100); +x_102 = l_Lean_Elab_Tactic_evalIntro___closed__25; +lean_inc(x_66); +x_103 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_103, 0, x_66); +lean_ctor_set(x_103, 1, x_102); +x_104 = lean_array_push(x_81, x_60); +x_105 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_105, 0, x_83); +lean_ctor_set(x_105, 1, x_84); +lean_ctor_set(x_105, 2, x_104); +x_106 = lean_array_push(x_81, x_105); +x_107 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_107, 0, x_83); +lean_ctor_set(x_107, 1, x_84); +lean_ctor_set(x_107, 2, x_106); +x_108 = l_Lean_Elab_Tactic_evalIntro___closed__26; +lean_inc(x_66); +x_109 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_109, 0, x_66); +lean_ctor_set(x_109, 1, x_108); +x_110 = l_Lean_Elab_Tactic_evalIntro___closed__29; +lean_inc(x_66); +x_111 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_111, 0, x_66); +lean_ctor_set(x_111, 1, x_110); +x_112 = l_Lean_Elab_Tactic_evalIntro___closed__30; +lean_inc(x_66); +x_113 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_113, 0, x_66); +lean_ctor_set(x_113, 1, x_112); +x_114 = lean_array_push(x_81, x_113); +x_115 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_115, 0, x_83); +lean_ctor_set(x_115, 1, x_63); +lean_ctor_set(x_115, 2, x_114); +x_116 = lean_array_push(x_86, x_111); +x_117 = lean_array_push(x_116, x_115); +x_118 = l_Lean_Elab_Tactic_evalIntro___closed__28; +x_119 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_119, 0, x_83); +lean_ctor_set(x_119, 1, x_118); +lean_ctor_set(x_119, 2, x_117); +x_120 = l_Lean_Elab_Tactic_evalIntro___closed__31; +x_121 = lean_array_push(x_120, x_103); +x_122 = lean_array_push(x_121, x_107); +x_123 = lean_array_push(x_122, x_109); +x_124 = lean_array_push(x_123, x_119); +x_125 = l_Lean_Elab_Tactic_evalIntro___closed__24; +x_126 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_126, 0, x_83); +lean_ctor_set(x_126, 1, x_125); +lean_ctor_set(x_126, 2, x_124); +x_127 = lean_array_push(x_81, x_126); +x_128 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_128, 0, x_83); +lean_ctor_set(x_128, 1, x_84); +lean_ctor_set(x_128, 2, x_127); +x_129 = lean_array_push(x_81, x_128); +x_130 = l_Lean_Elab_Tactic_evalIntro___closed__22; +x_131 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_131, 0, x_83); +lean_ctor_set(x_131, 1, x_130); +lean_ctor_set(x_131, 2, x_129); +x_132 = l_Lean_Elab_Tactic_evalIntro___closed__32; +x_133 = lean_array_push(x_132, x_93); +x_134 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__10; +x_135 = lean_array_push(x_133, x_134); +x_136 = lean_array_push(x_135, x_134); +x_137 = lean_array_push(x_136, x_99); +x_138 = lean_array_push(x_137, x_101); +x_139 = lean_array_push(x_138, x_131); +x_140 = l_Lean_Elab_Tactic_evalIntro___closed__16; +x_141 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_141, 0, x_83); +lean_ctor_set(x_141, 1, x_140); +lean_ctor_set(x_141, 2, x_139); +x_142 = l_Lean_Elab_Tactic_evalIntro___closed__35; +lean_inc(x_66); +x_143 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_143, 0, x_66); +lean_ctor_set(x_143, 1, x_142); +x_144 = l_Lean_Elab_Tactic_evalIntro___closed__36; +x_145 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_145, 0, x_66); +lean_ctor_set(x_145, 1, x_144); +x_146 = lean_array_push(x_86, x_145); +x_147 = lean_array_push(x_146, x_85); +x_148 = l_Lean_Elab_Tactic_evalIntro___closed__37; +x_149 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_149, 0, x_83); +lean_ctor_set(x_149, 1, x_148); +lean_ctor_set(x_149, 2, x_147); +x_150 = lean_array_push(x_86, x_149); +x_151 = lean_array_push(x_150, x_134); +x_152 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__12; +x_153 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_153, 0, x_83); +lean_ctor_set(x_153, 1, x_152); +lean_ctor_set(x_153, 2, x_151); +x_154 = lean_array_push(x_81, x_153); +x_155 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_155, 0, x_83); +lean_ctor_set(x_155, 1, x_84); +lean_ctor_set(x_155, 2, x_154); +x_156 = lean_array_push(x_81, x_155); +x_157 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__2; +x_158 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_158, 0, x_83); +lean_ctor_set(x_158, 1, x_157); +lean_ctor_set(x_158, 2, x_156); +x_159 = lean_array_push(x_81, x_158); +x_160 = l___regBuiltin_Lean_Elab_Tactic_evalTacticSeq___closed__2; +x_161 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_161, 0, x_83); +lean_ctor_set(x_161, 1, x_160); +lean_ctor_set(x_161, 2, x_159); +x_162 = lean_array_push(x_86, x_143); +x_163 = lean_array_push(x_162, x_161); +x_164 = l_Lean_Elab_Tactic_evalIntro___closed__34; +x_165 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_165, 0, x_83); +lean_ctor_set(x_165, 1, x_164); +lean_ctor_set(x_165, 2, x_163); +x_166 = l_Lean_Elab_Tactic_evalIntro___closed__38; +x_167 = lean_array_push(x_166, x_89); +lean_inc(x_91); +x_168 = lean_array_push(x_167, x_91); +x_169 = lean_array_push(x_168, x_141); +x_170 = lean_array_push(x_169, x_91); +x_171 = lean_array_push(x_170, x_165); +x_172 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_172, 0, x_83); +lean_ctor_set(x_172, 1, x_84); +lean_ctor_set(x_172, 2, x_171); +x_173 = lean_array_push(x_81, x_172); +x_174 = l___regBuiltin_Lean_Elab_Tactic_evalSeq1___closed__2; +x_175 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_175, 0, x_83); +lean_ctor_set(x_175, 1, x_174); +lean_ctor_set(x_175, 2, x_173); +x_176 = l_Lean_Elab_Tactic_evalTacticAux(x_175, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_71); +return x_176; +} +else +{ +lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; +x_177 = l_Lean_Syntax_getArg(x_60, x_16); +lean_dec(x_60); +x_178 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_178, 0, x_177); +x_179 = l_Lean_Elab_Tactic_evalIntro___closed__39; +x_180 = l_Lean_Elab_Tactic_evalIntro_introStep(x_178, x_179, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_180; } } else { -lean_object* x_182; lean_object* x_183; lean_object* x_184; -lean_inc(x_61); -x_182 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_182, 0, x_61); -x_183 = l_Lean_Syntax_getId(x_61); -lean_dec(x_61); -x_184 = l_Lean_Elab_Tactic_evalIntro_introStep(x_182, x_183, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_184; +lean_object* x_181; lean_object* x_182; lean_object* x_183; +lean_inc(x_60); +x_181 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_181, 0, x_60); +x_182 = l_Lean_Syntax_getId(x_60); +lean_dec(x_60); +x_183 = l_Lean_Elab_Tactic_evalIntro_introStep(x_181, x_182, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_183; } } } else { -lean_object* x_185; lean_object* x_186; lean_object* x_187; +lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_dec(x_15); -x_185 = lean_box(0); -x_186 = l_Lean_Elab_Tactic_evalIntro___closed__44; -x_187 = l_Lean_Elab_Tactic_evalIntro_introStep(x_185, x_186, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_187; +x_184 = lean_box(0); +x_185 = l_Lean_Elab_Tactic_evalIntro___closed__39; +x_186 = l_Lean_Elab_Tactic_evalIntro_introStep(x_184, x_185, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_186; } } } @@ -17404,22 +17387,21 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_unsigned_to_nat(1u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); lean_dec(x_1); -x_16 = l_Lean_nullKind; -x_17 = lean_unsigned_to_nat(0u); +x_16 = lean_unsigned_to_nat(0u); lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_16, x_17); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_16); +if (x_17 == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = l_Lean_Syntax_getArgs(x_15); +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = l_Lean_Syntax_getArgs(x_15); lean_dec(x_15); -lean_inc(x_19); -x_20 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_evalIntros___lambda__1), 10, 1); -lean_closure_set(x_20, 0, x_19); +lean_inc(x_18); +x_19 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_evalIntros___lambda__1), 10, 1); +lean_closure_set(x_19, 0, x_18); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); @@ -17428,34 +17410,34 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_21 = l_Lean_Elab_Tactic_withMainContext___rarg(x_20, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_21) == 0) +x_20 = l_Lean_Elab_Tactic_withMainContext___rarg(x_19, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; size_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_22 = lean_ctor_get(x_21, 0); +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; size_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); lean_inc(x_22); -x_23 = lean_ctor_get(x_21, 1); -lean_inc(x_23); -lean_dec(x_21); -x_24 = lean_array_get_size(x_22); -x_25 = l_Array_toSubarray___rarg(x_22, x_17, x_24); -x_26 = lean_array_get_size(x_19); -x_27 = lean_usize_of_nat(x_26); -lean_dec(x_26); -x_28 = lean_box_usize(x_27); -x_29 = l_Lean_Elab_Tactic_evalIntros___boxed__const__1; -x_30 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_evalIntros___lambda__2___boxed), 13, 4); -lean_closure_set(x_30, 0, x_19); -lean_closure_set(x_30, 1, x_28); -lean_closure_set(x_30, 2, x_29); -lean_closure_set(x_30, 3, x_25); -x_31 = l_Lean_Elab_Tactic_withMainContext___rarg(x_30, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_23); -return x_31; +lean_dec(x_20); +x_23 = lean_array_get_size(x_21); +x_24 = l_Array_toSubarray___rarg(x_21, x_16, x_23); +x_25 = lean_array_get_size(x_18); +x_26 = lean_usize_of_nat(x_25); +lean_dec(x_25); +x_27 = lean_box_usize(x_26); +x_28 = l_Lean_Elab_Tactic_evalIntros___boxed__const__1; +x_29 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_evalIntros___lambda__2___boxed), 13, 4); +lean_closure_set(x_29, 0, x_18); +lean_closure_set(x_29, 1, x_27); +lean_closure_set(x_29, 2, x_28); +lean_closure_set(x_29, 3, x_24); +x_30 = l_Lean_Elab_Tactic_withMainContext___rarg(x_29, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_22); +return x_30; } else { -uint8_t x_32; -lean_dec(x_19); +uint8_t x_31; +lean_dec(x_18); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -17464,33 +17446,33 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_32 = !lean_is_exclusive(x_21); -if (x_32 == 0) +x_31 = !lean_is_exclusive(x_20); +if (x_31 == 0) { -return x_21; +return x_20; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_21, 0); -x_34 = lean_ctor_get(x_21, 1); -lean_inc(x_34); +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_20, 0); +x_33 = lean_ctor_get(x_20, 1); lean_inc(x_33); -lean_dec(x_21); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -return x_35; +lean_inc(x_32); +lean_dec(x_20); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; } } } else { -lean_object* x_36; lean_object* x_37; +lean_object* x_35; lean_object* x_36; lean_dec(x_15); -x_36 = l_Lean_Elab_Tactic_evalIntros___closed__3; -x_37 = l_Lean_Elab_Tactic_withMainContext___rarg(x_36, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_37; +x_35 = l_Lean_Elab_Tactic_evalIntros___closed__3; +x_36 = l_Lean_Elab_Tactic_withMainContext___rarg(x_35, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_36; } } } @@ -18221,7 +18203,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalClear(lean_object* x_1, lean_obj _start: { lean_object* x_11; uint8_t x_12; -x_11 = l_Lean_Elab_Tactic_evalIntro___closed__42; +x_11 = l_Lean_Elab_Tactic_evalIntro___closed__37; lean_inc(x_1); x_12 = l_Lean_Syntax_isOfKind(x_1, x_11); if (x_12 == 0) @@ -18461,7 +18443,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_evalClear(lean_object* { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_2 = l___regBuiltin_Lean_Elab_Tactic_evalWithAnnotateState___closed__6; -x_3 = l_Lean_Elab_Tactic_evalIntro___closed__42; +x_3 = l_Lean_Elab_Tactic_evalIntro___closed__37; x_4 = l___regBuiltin_Lean_Elab_Tactic_evalClear___closed__2; x_5 = l___regBuiltin_Lean_Elab_Tactic_evalClear___closed__3; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); @@ -21634,7 +21616,7 @@ static lean_object* _init_l_Lean_Elab_ContextInfo_saveNoFileMap___at_Lean_Elab_T { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_resolveGlobalConstNoOverloadCore___at_Lean_Elab_Tactic_evalOpen___spec__10___closed__3; -x_2 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__7; +x_2 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__9; x_3 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -23197,7 +23179,7 @@ lean_inc(x_19); lean_dec(x_14); lean_inc(x_17); x_20 = lean_local_ctx_num_indices(x_17); -x_21 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__7; +x_21 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__9; x_22 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_22, 0, x_21); lean_ctor_set(x_22, 1, x_17); @@ -24016,11 +23998,11 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_ev _start: { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_13 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__6; +x_13 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__8; x_14 = lean_array_push(x_13, x_1); x_15 = lean_array_push(x_14, x_2); x_16 = lean_box(2); -x_17 = l_Lean_nullKind; +x_17 = l_Lean_Elab_Tactic_addCheckpoints_push___closed__4; x_18 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_18, 0, x_16); lean_ctor_set(x_18, 1, x_17); @@ -25730,20 +25712,38 @@ static lean_object* _init_l_Lean_Elab_Tactic_evalFail___closed__5() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("\n", 1); +x_1 = lean_mk_string_from_bytes("str", 3); return x_1; } } static lean_object* _init_l_Lean_Elab_Tactic_evalFail___closed__6() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Tactic_evalFail___closed__5; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_evalFail___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("\n", 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_evalFail___closed__8() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic_evalFail___closed__5; +x_1 = l_Lean_Elab_Tactic_evalFail___closed__7; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalFail___closed__7() { +static lean_object* _init_l_Lean_Elab_Tactic_evalFail___closed__9() { _start: { lean_object* x_1; @@ -25751,11 +25751,11 @@ x_1 = lean_mk_string_from_bytes("tactic 'fail' failed\n", 21); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalFail___closed__8() { +static lean_object* _init_l_Lean_Elab_Tactic_evalFail___closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic_evalFail___closed__7; +x_1 = l_Lean_Elab_Tactic_evalFail___closed__9; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } @@ -25788,54 +25788,53 @@ return x_20; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; x_21 = lean_unsigned_to_nat(1u); x_22 = l_Lean_Syntax_getArg(x_1, x_21); lean_dec(x_1); -x_23 = l_Lean_nullKind; -x_24 = lean_unsigned_to_nat(0u); +x_23 = lean_unsigned_to_nat(0u); lean_inc(x_22); -x_25 = l_Lean_Syntax_isNodeOf(x_22, x_23, x_24); -if (x_25 == 0) +x_24 = l_Lean_Syntax_matchesNull(x_22, x_23); +if (x_24 == 0) { -uint8_t x_26; +uint8_t x_25; lean_inc(x_22); -x_26 = l_Lean_Syntax_isNodeOf(x_22, x_23, x_21); -if (x_26 == 0) +x_25 = l_Lean_Syntax_matchesNull(x_22, x_21); +if (x_25 == 0) { -lean_object* x_27; +lean_object* x_26; lean_dec(x_22); lean_dec(x_17); -x_27 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalExact___spec__1___rarg(x_13); -return x_27; +x_26 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalExact___spec__1___rarg(x_13); +return x_26; } else { -lean_object* x_28; lean_object* x_29; -x_28 = l_Lean_Syntax_getArg(x_22, x_24); +lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_27 = l_Lean_Syntax_getArg(x_22, x_23); lean_dec(x_22); -x_29 = l_Lean_Syntax_isStrLit_x3f(x_28); -lean_dec(x_28); -if (lean_obj_tag(x_29) == 0) +x_28 = l_Lean_Elab_Tactic_evalFail___closed__6; +lean_inc(x_27); +x_29 = l_Lean_Syntax_isOfKind(x_27, x_28); +if (x_29 == 0) { lean_object* x_30; +lean_dec(x_27); lean_dec(x_17); -x_30 = l_Lean_Elab_throwIllFormedSyntax___at_Lean_Elab_Tactic_evalTraceMessage___spec__1(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +x_30 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalExact___spec__1___rarg(x_13); return x_30; } else { 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; -x_31 = lean_ctor_get(x_29, 0); -lean_inc(x_31); -lean_dec(x_29); +x_31 = l_Lean_TSyntax_getString(x_27); x_32 = l_Lean_stringToMessageData(x_31); lean_dec(x_31); x_33 = l_Lean_Elab_nestedExceptionToMessageData___at_Lean_Elab_Tactic_evalOpen___spec__23___closed__1; x_34 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_34, 0, x_33); lean_ctor_set(x_34, 1, x_32); -x_35 = l_Lean_Elab_Tactic_evalFail___closed__6; +x_35 = l_Lean_Elab_Tactic_evalFail___closed__8; x_36 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_36, 0, x_34); lean_ctor_set(x_36, 1, x_35); @@ -25854,7 +25853,7 @@ else { lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_dec(x_22); -x_40 = l_Lean_Elab_Tactic_evalFail___closed__8; +x_40 = l_Lean_Elab_Tactic_evalFail___closed__10; x_41 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_41, 0, x_40); lean_ctor_set(x_41, 1, x_17); @@ -25938,7 +25937,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalFail_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(373u); +x_1 = lean_unsigned_to_nat(370u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26164,7 +26163,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalDbgTrace_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(375u); +x_1 = lean_unsigned_to_nat(372u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26176,7 +26175,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalDbgTrace_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(378u); +x_1 = lean_unsigned_to_nat(375u); x_2 = lean_unsigned_to_nat(29u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26204,7 +26203,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalDbgTrace_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(375u); +x_1 = lean_unsigned_to_nat(372u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26216,7 +26215,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalDbgTrace_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(375u); +x_1 = lean_unsigned_to_nat(372u); x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26265,95 +26264,94 @@ return x_4; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalSleep(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_object* x_9, lean_object* x_10) { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +lean_object* x_11; lean_object* x_12; lean_object* x_13; x_11 = lean_unsigned_to_nat(1u); x_12 = l_Lean_Syntax_getArg(x_1, x_11); -x_13 = l_Lean_numLitKind; -x_14 = l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(x_13, x_12); +x_13 = l_Lean_Syntax_isNatLit_x3f(x_12); lean_dec(x_12); -if (lean_obj_tag(x_14) == 0) +if (lean_obj_tag(x_13) == 0) { -lean_object* x_15; -x_15 = l_Lean_Elab_throwIllFormedSyntax___at_Lean_Elab_Tactic_evalTraceMessage___spec__1(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_15; +lean_object* x_14; +x_14 = l_Lean_Elab_throwIllFormedSyntax___at_Lean_Elab_Tactic_evalTraceMessage___spec__1(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_14; } else { -lean_object* x_16; uint32_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_16 = lean_ctor_get(x_14, 0); -lean_inc(x_16); -lean_dec(x_14); -x_17 = lean_uint32_of_nat(x_16); -lean_dec(x_16); -x_18 = lean_st_ref_get(x_9, x_10); -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -lean_dec(x_18); -x_20 = lean_ctor_get(x_8, 5); -x_21 = l_IO_sleep(x_17, x_19); -if (lean_obj_tag(x_21) == 0) +lean_object* x_15; uint32_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_15 = lean_ctor_get(x_13, 0); +lean_inc(x_15); +lean_dec(x_13); +x_16 = lean_uint32_of_nat(x_15); +lean_dec(x_15); +x_17 = lean_st_ref_get(x_9, x_10); +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); +lean_dec(x_17); +x_19 = lean_ctor_get(x_8, 5); +x_20 = l_IO_sleep(x_16, x_18); +if (lean_obj_tag(x_20) == 0) { -uint8_t x_22; -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) +uint8_t x_21; +x_21 = !lean_is_exclusive(x_20); +if (x_21 == 0) { -return x_21; +return x_20; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_21, 0); -x_24 = lean_ctor_get(x_21, 1); -lean_inc(x_24); +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_20, 0); +x_23 = lean_ctor_get(x_20, 1); lean_inc(x_23); -lean_dec(x_21); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_23); -lean_ctor_set(x_25, 1, x_24); -return x_25; +lean_inc(x_22); +lean_dec(x_20); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +return x_24; } } else { -uint8_t x_26; -x_26 = !lean_is_exclusive(x_21); -if (x_26 == 0) +uint8_t x_25; +x_25 = !lean_is_exclusive(x_20); +if (x_25 == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_27 = lean_ctor_get(x_21, 0); -x_28 = lean_io_error_to_string(x_27); -x_29 = lean_alloc_ctor(2, 1, 0); +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_26 = lean_ctor_get(x_20, 0); +x_27 = lean_io_error_to_string(x_26); +x_28 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_28, 0, x_27); +x_29 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_29, 0, x_28); -x_30 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_30, 0, x_29); -lean_inc(x_20); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_20); -lean_ctor_set(x_31, 1, x_30); -lean_ctor_set(x_21, 0, x_31); -return x_21; +lean_inc(x_19); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_19); +lean_ctor_set(x_30, 1, x_29); +lean_ctor_set(x_20, 0, x_30); +return x_20; } else { -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; -x_32 = lean_ctor_get(x_21, 0); -x_33 = lean_ctor_get(x_21, 1); -lean_inc(x_33); +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; +x_31 = lean_ctor_get(x_20, 0); +x_32 = lean_ctor_get(x_20, 1); lean_inc(x_32); -lean_dec(x_21); -x_34 = lean_io_error_to_string(x_32); -x_35 = lean_alloc_ctor(2, 1, 0); +lean_inc(x_31); +lean_dec(x_20); +x_33 = lean_io_error_to_string(x_31); +x_34 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_34, 0, x_33); +x_35 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_35, 0, x_34); -x_36 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_36, 0, x_35); -lean_inc(x_20); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_20); -lean_ctor_set(x_37, 1, x_36); -x_38 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_38, 0, x_37); -lean_ctor_set(x_38, 1, x_33); -return x_38; +lean_inc(x_19); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_19); +lean_ctor_set(x_36, 1, x_35); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_32); +return x_37; } } } @@ -26436,7 +26434,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSleep_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(380u); +x_1 = lean_unsigned_to_nat(377u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26448,7 +26446,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSleep_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(383u); +x_1 = lean_unsigned_to_nat(380u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26476,7 +26474,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSleep_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(380u); +x_1 = lean_unsigned_to_nat(377u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26488,7 +26486,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSleep_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(380u); +x_1 = lean_unsigned_to_nat(377u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26715,6 +26713,14 @@ l_Lean_Elab_Tactic_addCheckpoints_push___closed__7 = _init_l_Lean_Elab_Tactic_ad lean_mark_persistent(l_Lean_Elab_Tactic_addCheckpoints_push___closed__7); l_Lean_Elab_Tactic_addCheckpoints_push___closed__8 = _init_l_Lean_Elab_Tactic_addCheckpoints_push___closed__8(); lean_mark_persistent(l_Lean_Elab_Tactic_addCheckpoints_push___closed__8); +l_Lean_Elab_Tactic_addCheckpoints_push___closed__9 = _init_l_Lean_Elab_Tactic_addCheckpoints_push___closed__9(); +lean_mark_persistent(l_Lean_Elab_Tactic_addCheckpoints_push___closed__9); +l_Lean_Elab_Tactic_addCheckpoints_push___closed__10 = _init_l_Lean_Elab_Tactic_addCheckpoints_push___closed__10(); +lean_mark_persistent(l_Lean_Elab_Tactic_addCheckpoints_push___closed__10); +l_Lean_Elab_Tactic_addCheckpoints_push___closed__11 = _init_l_Lean_Elab_Tactic_addCheckpoints_push___closed__11(); +lean_mark_persistent(l_Lean_Elab_Tactic_addCheckpoints_push___closed__11); +l_Lean_Elab_Tactic_addCheckpoints_push___closed__12 = _init_l_Lean_Elab_Tactic_addCheckpoints_push___closed__12(); +lean_mark_persistent(l_Lean_Elab_Tactic_addCheckpoints_push___closed__12); l___regBuiltin_Lean_Elab_Tactic_evalTacticSeq1Indented___closed__1 = _init_l___regBuiltin_Lean_Elab_Tactic_evalTacticSeq1Indented___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_evalTacticSeq1Indented___closed__1); l___regBuiltin_Lean_Elab_Tactic_evalTacticSeq1Indented___closed__2 = _init_l___regBuiltin_Lean_Elab_Tactic_evalTacticSeq1Indented___closed__2(); @@ -27503,16 +27509,6 @@ l_Lean_Elab_Tactic_evalIntro___closed__38 = _init_l_Lean_Elab_Tactic_evalIntro__ lean_mark_persistent(l_Lean_Elab_Tactic_evalIntro___closed__38); l_Lean_Elab_Tactic_evalIntro___closed__39 = _init_l_Lean_Elab_Tactic_evalIntro___closed__39(); lean_mark_persistent(l_Lean_Elab_Tactic_evalIntro___closed__39); -l_Lean_Elab_Tactic_evalIntro___closed__40 = _init_l_Lean_Elab_Tactic_evalIntro___closed__40(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalIntro___closed__40); -l_Lean_Elab_Tactic_evalIntro___closed__41 = _init_l_Lean_Elab_Tactic_evalIntro___closed__41(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalIntro___closed__41); -l_Lean_Elab_Tactic_evalIntro___closed__42 = _init_l_Lean_Elab_Tactic_evalIntro___closed__42(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalIntro___closed__42); -l_Lean_Elab_Tactic_evalIntro___closed__43 = _init_l_Lean_Elab_Tactic_evalIntro___closed__43(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalIntro___closed__43); -l_Lean_Elab_Tactic_evalIntro___closed__44 = _init_l_Lean_Elab_Tactic_evalIntro___closed__44(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalIntro___closed__44); l___regBuiltin_Lean_Elab_Tactic_evalIntro___closed__1 = _init_l___regBuiltin_Lean_Elab_Tactic_evalIntro___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_evalIntro___closed__1); l___regBuiltin_Lean_Elab_Tactic_evalIntro___closed__2 = _init_l___regBuiltin_Lean_Elab_Tactic_evalIntro___closed__2(); @@ -27877,6 +27873,10 @@ l_Lean_Elab_Tactic_evalFail___closed__7 = _init_l_Lean_Elab_Tactic_evalFail___cl lean_mark_persistent(l_Lean_Elab_Tactic_evalFail___closed__7); l_Lean_Elab_Tactic_evalFail___closed__8 = _init_l_Lean_Elab_Tactic_evalFail___closed__8(); lean_mark_persistent(l_Lean_Elab_Tactic_evalFail___closed__8); +l_Lean_Elab_Tactic_evalFail___closed__9 = _init_l_Lean_Elab_Tactic_evalFail___closed__9(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalFail___closed__9); +l_Lean_Elab_Tactic_evalFail___closed__10 = _init_l_Lean_Elab_Tactic_evalFail___closed__10(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalFail___closed__10); l___regBuiltin_Lean_Elab_Tactic_evalFail___closed__1 = _init_l___regBuiltin_Lean_Elab_Tactic_evalFail___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_evalFail___closed__1); l___regBuiltin_Lean_Elab_Tactic_evalFail___closed__2 = _init_l___regBuiltin_Lean_Elab_Tactic_evalFail___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Config.c b/stage0/stdlib/Lean/Elab/Tactic/Config.c index 62c0f0ba0a1..98c0b1e0dfa 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Config.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Config.c @@ -13,275 +13,275 @@ #ifdef __cplusplus extern "C" { #endif -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__43; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__215; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__212; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__98; -static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__17; -static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__8; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__219; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__217; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__207; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__36; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__19; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__206; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__146; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__81; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__219; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__212; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__211; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__172; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__20; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__238; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__228; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__10; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__174; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__32; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__215; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__98; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__73; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__207; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__43; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__120; lean_object* lean_mk_empty_array_with_capacity(lean_object*); -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__39; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__65; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__140; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__109; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__118; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__211; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__172; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab________; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__177; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__106; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__158; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__71; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__7; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__202; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__24; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__190; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__63; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab____; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__123; lean_object* lean_name_mk_string(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__70; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__151; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__183; -static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__2; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__69; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__237; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__114; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__188; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__26; -static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__12; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__83; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__168; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__75; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__173; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__204; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__169; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__90; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__68; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__16; -static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__11; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__81; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__73; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__115; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__228; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__16; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__49; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__164; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__105; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__115; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__85; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__186; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__124; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__134; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__87; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__138; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__143; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__68; lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__56; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__105; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__85; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__95; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__79; lean_object* lean_array_push(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__186; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__66; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__165; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__234; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__77; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__49; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__164; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__177; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__10; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__174; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__32; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__176; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__101; +static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__14; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__214; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__167; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__160; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__232; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__92; lean_object* lean_string_utf8_byte_size(lean_object*); -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__3; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__213; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__101; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__238; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__20; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__120; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__117; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__224; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__185; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__149; -static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__4; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__11; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__135; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__159; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__112; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__205; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__187; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__113; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__226; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__41; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__218; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__34; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__99; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__180; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__61; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__195; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__233; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__48; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__162; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__220; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__170; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__21; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__38; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__107; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__222; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__50; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__13; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__88; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__131; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__200; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__110; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__154; -static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__9; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__80; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__6; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__145; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__35; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__189; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__104; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__194; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__31; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__161; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__127; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__209; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__59; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__156; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__45; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__231; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__128; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__12; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__116; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__192; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__58; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__196; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__148; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__160; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__203; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__232; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__82; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__102; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__136; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__150; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__97; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__1; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__171; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__142; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__166; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__125; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__53; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__229; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__144; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__197; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__181; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__221; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__42; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__5; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__210; -static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__10; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__205; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__61; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__136; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__33; +static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__18; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__137; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__84; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__93; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__22; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__121; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__132; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__78; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__153; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__103; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__91; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__46; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__110; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__86; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__178; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__198; +static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__7; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__193; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__25; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__119; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__152; +static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__15; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__29; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__62; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__191; +static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__5; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__37; +static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__3; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__240; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__100; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__42; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__53; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__208; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__230; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__55; +static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__17; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__126; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__108; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__133; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__5; +static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__9; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__28; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__14; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__210; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__15; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__217; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__18; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__76; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__30; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__225; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__184; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__142; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__3; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__163; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__213; +static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__8; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__125; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__176; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__181; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__166; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__221; +static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__4; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__229; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__149; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__77; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__224; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__185; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__117; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__165; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__56; +static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__1; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__234; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__66; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__150; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__1; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__2; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__64; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__47; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__236; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__89; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__45; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__60; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__157; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__74; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__40; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__175; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__67; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__129; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__189; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__201; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__4; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__27; +static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__6; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__9; lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__30; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__225; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__184; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__18; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__15; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__163; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__76; -static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__6; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__7; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__71; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__133; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__100; -static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__15; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__239; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__223; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__141; -static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__13; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__111; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__155; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__182; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__57; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__130; -static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__14; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__64; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__17; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__94; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__44; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__139; -static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__3; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__8; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__67; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__4; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__235; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__175; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__40; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__201; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__122; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__27; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__2; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__51; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__89; -static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__18; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__236; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__60; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__47; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__129; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__74; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__157; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__9; -static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__5; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__103; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__132; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__14; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__54; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__137; -static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__16; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__84; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__33; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__216; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__72; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__93; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__91; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__138; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__87; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__134; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__92; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__28; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__62; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__17; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__122; +static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__13; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__57; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__131; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__139; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__220; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__94; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__8; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__144; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__48; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__162; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__34; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__233; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__195; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__226; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__223; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__141; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__239; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__111; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__200; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__182; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__155; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__99; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__50; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__6; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__80; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__154; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__130; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__170; +static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__12; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__107; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__222; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__38; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__21; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__179; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__199; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__88; +static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__11; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__23; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__227; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__72; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__216; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__96; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__41; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__218; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__13; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__180; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__54; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__147; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__52; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__194; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__112; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__113; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__209; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__31; +static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__2; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__59; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__187; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__159; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__82; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__135; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__203; +static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__10; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__90; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__102; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__169; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__116; uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__124; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__167; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__95; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__123; -static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__7; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__192; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__128; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__231; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__12; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__58; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__148; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__26; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__114; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__196; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__75; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__156; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__161; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__127; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__143; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__79; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__214; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__190; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__63; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__37; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__158; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__106; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__202; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__108; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__24; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__55; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__240; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__208; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__126; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__230; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__193; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__29; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__52; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__22; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__152; -static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__1; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__147; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__25; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__119; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__86; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__191; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__178; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__199; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__179; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__23; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__96; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__227; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__78; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__121; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__198; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__46; -static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__153; -static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__1() { +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__35; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__70; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__183; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__109; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__65; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__39; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__36; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__173; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__188; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__104; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__145; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__83; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__168; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__11; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__151; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__204; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__69; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__237; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__19; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__51; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__206; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__146; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__118; +static lean_object* l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__16; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__140; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__197; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__171; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__235; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__97; +static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__44; +static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__1() { _start: { lean_object* x_1; @@ -289,17 +289,17 @@ x_1 = lean_mk_string_from_bytes("Lean", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__2() { +static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__1; +x_2 = l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__3() { +static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__3() { _start: { lean_object* x_1; @@ -307,17 +307,17 @@ x_1 = lean_mk_string_from_bytes("Elab", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__4() { +static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__2; -x_2 = l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__3; +x_1 = l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__2; +x_2 = l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__5() { +static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__5() { _start: { lean_object* x_1; @@ -325,35 +325,35 @@ x_1 = lean_mk_string_from_bytes("Tactic", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__6() { +static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__4; -x_2 = l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__5; +x_1 = l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__4; +x_2 = l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__7() { +static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__7() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("commandDeclare_config_elab____", 30); +x_1 = lean_mk_string_from_bytes("commandDeclare_config_elab__", 28); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__8() { +static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__6; -x_2 = l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__7; +x_1 = l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__6; +x_2 = l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__7; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__9() { +static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__9() { _start: { lean_object* x_1; @@ -361,17 +361,17 @@ x_1 = lean_mk_string_from_bytes("andthen", 7); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__10() { +static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__9; +x_2 = l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__9; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__11() { +static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__11() { _start: { lean_object* x_1; @@ -379,17 +379,17 @@ x_1 = lean_mk_string_from_bytes("declare_config_elab", 19); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__12() { +static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__11; +x_1 = l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__11; x_2 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__13() { +static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__13() { _start: { lean_object* x_1; @@ -397,33 +397,33 @@ x_1 = lean_mk_string_from_bytes("ident", 5); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__14() { +static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__13; +x_2 = l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__13; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__15() { +static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__15() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__14; +x_1 = l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__14; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__16() { +static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__10; -x_2 = l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__12; -x_3 = l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__15; +x_1 = l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__10; +x_2 = l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__12; +x_3 = l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__15; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -431,13 +431,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__17() { +static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__10; -x_2 = l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__16; -x_3 = l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__15; +x_1 = l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__10; +x_2 = l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__16; +x_3 = l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__15; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -445,13 +445,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__18() { +static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__8; +x_1 = l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__8; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__17; +x_3 = l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__17; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -459,15 +459,15 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab________() { +static lean_object* _init_l_Lean_Elab_Tactic_commandDeclare__config__elab____() { _start: { lean_object* x_1; -x_1 = l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__18; +x_1 = l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__18; return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__1() { _start: { lean_object* x_1; @@ -475,17 +475,17 @@ x_1 = lean_mk_string_from_bytes("null", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__2() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__1; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__3() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__3() { _start: { lean_object* x_1; @@ -493,17 +493,17 @@ x_1 = lean_mk_string_from_bytes("Parser", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__4() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__2; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__3; +x_1 = l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__2; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__5() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__5() { _start: { lean_object* x_1; @@ -511,17 +511,17 @@ x_1 = lean_mk_string_from_bytes("Command", 7); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__6() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__4; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__5; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__4; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__7() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__7() { _start: { lean_object* x_1; @@ -529,17 +529,17 @@ x_1 = lean_mk_string_from_bytes("declaration", 11); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__8() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__6; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__7; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__6; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__7; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__9() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__9() { _start: { lean_object* x_1; @@ -547,17 +547,17 @@ x_1 = lean_mk_string_from_bytes("declModifiers", 13); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__10() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__6; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__9; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__6; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__9; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__11() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__11() { _start: { lean_object* x_1; lean_object* x_2; @@ -566,13 +566,13 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__12() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__2; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__11; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__2; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__11; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -580,7 +580,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__13() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__13() { _start: { lean_object* x_1; @@ -588,17 +588,17 @@ x_1 = lean_mk_string_from_bytes("unsafe", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__14() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__6; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__13; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__6; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__13; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__15() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__15() { _start: { lean_object* x_1; lean_object* x_2; @@ -607,7 +607,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__16() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__16() { _start: { lean_object* x_1; lean_object* x_2; @@ -616,47 +616,47 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__17() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__16; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__12; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__16; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__12; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__18() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__17; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__12; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__17; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__12; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__19() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__18; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__12; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__18; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__12; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__20() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__19; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__12; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__19; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__12; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__21() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__21() { _start: { lean_object* x_1; @@ -664,17 +664,17 @@ x_1 = lean_mk_string_from_bytes("def", 3); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__22() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__6; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__21; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__6; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__21; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__23() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__23() { _start: { lean_object* x_1; @@ -682,17 +682,17 @@ x_1 = lean_mk_string_from_bytes("declId", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__24() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__6; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__23; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__6; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__23; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__25() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__25() { _start: { lean_object* x_1; @@ -700,22 +700,22 @@ x_1 = lean_mk_string_from_bytes("evalUnsafe", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__26() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__26() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__25; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__25; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__27() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__27() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__25; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__25; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__26; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__26; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -723,17 +723,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__28() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__25; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__25; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__29() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__29() { _start: { lean_object* x_1; lean_object* x_2; @@ -742,7 +742,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__30() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__30() { _start: { lean_object* x_1; @@ -750,17 +750,17 @@ x_1 = lean_mk_string_from_bytes("optDeclSig", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__31() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__31() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__6; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__30; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__6; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__30; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__32() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__32() { _start: { lean_object* x_1; @@ -768,17 +768,17 @@ x_1 = lean_mk_string_from_bytes("Term", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__33() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__33() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__4; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__32; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__4; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__32; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__34() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__34() { _start: { lean_object* x_1; @@ -786,17 +786,17 @@ x_1 = lean_mk_string_from_bytes("explicitBinder", 14); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__35() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__35() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__33; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__34; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__33; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__34; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__36() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__36() { _start: { lean_object* x_1; @@ -804,7 +804,7 @@ x_1 = lean_mk_string_from_bytes("(", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__37() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__37() { _start: { lean_object* x_1; @@ -812,22 +812,22 @@ x_1 = lean_mk_string_from_bytes("e", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__38() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__38() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__37; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__37; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__39() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__39() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__37; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__37; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__38; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__38; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -835,17 +835,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__40() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__40() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__37; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__37; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__41() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__41() { _start: { lean_object* x_1; @@ -853,7 +853,7 @@ x_1 = lean_mk_string_from_bytes(":", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__42() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__42() { _start: { lean_object* x_1; @@ -861,22 +861,22 @@ x_1 = lean_mk_string_from_bytes("Expr", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__43() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__43() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__42; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__42; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__44() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__44() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__42; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__42; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__43; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__43; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -884,51 +884,51 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__45() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__45() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__42; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__42; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__46() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__46() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__2; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__42; +x_1 = l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__2; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__42; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__47() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__47() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__46; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__46; 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__48() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__48() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__47; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__47; x_3 = lean_alloc_ctor(1, 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__49() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__49() { _start: { lean_object* x_1; @@ -936,7 +936,7 @@ x_1 = lean_mk_string_from_bytes(")", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__50() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__50() { _start: { lean_object* x_1; lean_object* x_2; @@ -945,7 +945,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__51() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__51() { _start: { lean_object* x_1; @@ -953,17 +953,17 @@ x_1 = lean_mk_string_from_bytes("typeSpec", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__52() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__52() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__33; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__51; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__33; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__51; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__53() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__53() { _start: { lean_object* x_1; @@ -971,17 +971,17 @@ x_1 = lean_mk_string_from_bytes("app", 3); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__54() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__54() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__33; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__53; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__33; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__53; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__55() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__55() { _start: { lean_object* x_1; @@ -989,22 +989,22 @@ x_1 = lean_mk_string_from_bytes("TermElabM", 9); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__56() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__56() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__55; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__55; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__57() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__57() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__55; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__55; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__56; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__56; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1012,61 +1012,61 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__58() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__58() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__55; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__55; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__59() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__59() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__4; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__32; +x_1 = l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__4; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__32; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__60() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__60() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__59; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__55; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__59; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__55; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__61() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__61() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__60; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__60; 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__62() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__62() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__61; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__61; x_3 = lean_alloc_ctor(1, 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__63() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__63() { _start: { lean_object* x_1; @@ -1074,17 +1074,17 @@ x_1 = lean_mk_string_from_bytes("declValSimple", 13); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__64() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__64() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__6; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__63; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__6; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__63; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__65() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__65() { _start: { lean_object* x_1; @@ -1092,7 +1092,7 @@ x_1 = lean_mk_string_from_bytes(":=", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__66() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__66() { _start: { lean_object* x_1; @@ -1100,22 +1100,22 @@ x_1 = lean_mk_string_from_bytes("Term.evalExpr", 13); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__67() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__67() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__66; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__66; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__68() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__68() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__66; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__66; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__67; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__67; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1123,17 +1123,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__69() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__69() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__32; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__32; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__70() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__70() { _start: { lean_object* x_1; @@ -1141,51 +1141,51 @@ x_1 = lean_mk_string_from_bytes("evalExpr", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__71() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__71() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__69; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__70; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__69; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__70; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__72() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__72() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__59; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__70; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__59; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__70; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__73() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__73() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__72; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__72; 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__74() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__74() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__73; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__73; x_3 = lean_alloc_ctor(1, 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__75() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__75() { _start: { lean_object* x_1; @@ -1193,17 +1193,17 @@ x_1 = lean_mk_string_from_bytes("doubleQuotedName", 16); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__76() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__76() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__33; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__75; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__33; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__75; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__77() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__77() { _start: { lean_object* x_1; @@ -1211,7 +1211,7 @@ x_1 = lean_mk_string_from_bytes("`", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__78() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__78() { _start: { lean_object* x_1; lean_object* x_2; @@ -1220,7 +1220,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__79() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__79() { _start: { lean_object* x_1; lean_object* x_2; @@ -1229,7 +1229,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__80() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__80() { _start: { lean_object* x_1; @@ -1237,17 +1237,17 @@ x_1 = lean_mk_string_from_bytes("attributes", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__81() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__81() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__33; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__80; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__33; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__80; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__82() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__82() { _start: { lean_object* x_1; @@ -1255,7 +1255,7 @@ x_1 = lean_mk_string_from_bytes("@[", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__83() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__83() { _start: { lean_object* x_1; @@ -1263,17 +1263,17 @@ x_1 = lean_mk_string_from_bytes("attrInstance", 12); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__84() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__84() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__33; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__83; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__33; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__83; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__85() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__85() { _start: { lean_object* x_1; @@ -1281,33 +1281,33 @@ x_1 = lean_mk_string_from_bytes("attrKind", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__86() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__86() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__33; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__85; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__33; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__85; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__87() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__87() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__15; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__12; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__15; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__12; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__88() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__88() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__86; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__87; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__86; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__87; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1315,7 +1315,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__89() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__89() { _start: { lean_object* x_1; @@ -1323,17 +1323,17 @@ x_1 = lean_mk_string_from_bytes("Attr", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__90() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__90() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__4; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__89; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__4; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__89; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__91() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__91() { _start: { lean_object* x_1; @@ -1341,17 +1341,17 @@ x_1 = lean_mk_string_from_bytes("simple", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__92() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__92() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__90; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__91; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__90; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__91; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__93() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__93() { _start: { lean_object* x_1; @@ -1359,22 +1359,22 @@ x_1 = lean_mk_string_from_bytes("implementedBy", 13); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__94() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__94() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__93; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__93; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__95() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__95() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__93; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__93; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__94; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__94; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1382,27 +1382,27 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__96() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__96() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__93; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__93; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__97() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__97() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__29; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__88; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__29; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__88; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__98() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__98() { _start: { lean_object* x_1; @@ -1410,7 +1410,7 @@ x_1 = lean_mk_string_from_bytes("]", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__99() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__99() { _start: { lean_object* x_1; @@ -1418,17 +1418,17 @@ x_1 = lean_mk_string_from_bytes("opaque", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__100() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__100() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__6; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__99; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__6; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__99; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__101() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__101() { _start: { lean_object* x_1; @@ -1436,22 +1436,22 @@ x_1 = lean_mk_string_from_bytes("eval", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__102() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__102() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__101; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__101; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__103() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__103() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__101; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__101; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__102; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__102; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1459,17 +1459,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__104() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__104() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__101; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__101; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__105() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__105() { _start: { lean_object* x_1; @@ -1477,17 +1477,17 @@ x_1 = lean_mk_string_from_bytes("declSig", 7); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__106() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__106() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__6; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__105; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__6; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__105; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__107() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__107() { _start: { lean_object* x_1; lean_object* x_2; @@ -1496,33 +1496,33 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__108() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__108() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__20; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__12; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__20; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__12; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__109() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__109() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__108; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__12; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__108; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__12; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__110() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__110() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__10; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__109; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__10; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__109; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1530,7 +1530,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__111() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__111() { _start: { lean_object* x_1; @@ -1538,22 +1538,22 @@ x_1 = lean_mk_string_from_bytes("optConfig", 9); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__112() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__112() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__111; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__111; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__113() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__113() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__111; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__111; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__112; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__112; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1561,17 +1561,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__114() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__114() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__111; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__111; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__115() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__115() { _start: { lean_object* x_1; @@ -1579,22 +1579,22 @@ x_1 = lean_mk_string_from_bytes("Syntax", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__116() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__116() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__115; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__115; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__117() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__117() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__115; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__115; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__116; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__116; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1602,51 +1602,51 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__118() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__118() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__115; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__115; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__119() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__119() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__2; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__115; +x_1 = l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__2; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__115; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__120() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__120() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__119; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__119; 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__121() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__121() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__120; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__120; x_3 = lean_alloc_ctor(1, 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__122() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__122() { _start: { lean_object* x_1; @@ -1654,17 +1654,17 @@ x_1 = lean_mk_string_from_bytes("do", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__123() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__123() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__33; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__122; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__33; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__122; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__124() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__124() { _start: { lean_object* x_1; @@ -1672,17 +1672,17 @@ x_1 = lean_mk_string_from_bytes("doSeqIndent", 11); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__125() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__125() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__33; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__124; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__33; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__124; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__126() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__126() { _start: { lean_object* x_1; @@ -1690,17 +1690,17 @@ x_1 = lean_mk_string_from_bytes("doSeqItem", 9); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__127() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__127() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__33; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__126; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__33; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__126; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__128() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__128() { _start: { lean_object* x_1; @@ -1708,17 +1708,17 @@ x_1 = lean_mk_string_from_bytes("doIf", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__129() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__129() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__33; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__128; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__33; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__128; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__130() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__130() { _start: { lean_object* x_1; @@ -1726,7 +1726,7 @@ x_1 = lean_mk_string_from_bytes("if", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__131() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__131() { _start: { lean_object* x_1; @@ -1734,17 +1734,17 @@ x_1 = lean_mk_string_from_bytes("doIfProp", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__132() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__132() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__33; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__131; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__33; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__131; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__133() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__133() { _start: { lean_object* x_1; @@ -1752,22 +1752,22 @@ x_1 = lean_mk_string_from_bytes("optConfig.isNone", 16); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__134() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__134() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__133; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__133; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__135() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__135() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__133; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__133; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__134; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__134; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1775,7 +1775,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__136() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__136() { _start: { lean_object* x_1; @@ -1783,27 +1783,27 @@ x_1 = lean_mk_string_from_bytes("isNone", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__137() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__137() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__114; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__136; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__114; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__136; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__138() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__138() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__29; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__12; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__29; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__12; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__139() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__139() { _start: { lean_object* x_1; @@ -1811,7 +1811,7 @@ x_1 = lean_mk_string_from_bytes("then", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__140() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__140() { _start: { lean_object* x_1; @@ -1819,17 +1819,17 @@ x_1 = lean_mk_string_from_bytes("doReturn", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__141() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__141() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__33; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__140; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__33; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__140; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__142() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__142() { _start: { lean_object* x_1; @@ -1837,7 +1837,7 @@ x_1 = lean_mk_string_from_bytes("return", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__143() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__143() { _start: { lean_object* x_1; @@ -1845,17 +1845,17 @@ x_1 = lean_mk_string_from_bytes("structInst", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__144() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__144() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__33; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__143; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__33; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__143; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__145() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__145() { _start: { lean_object* x_1; @@ -1863,7 +1863,7 @@ x_1 = lean_mk_string_from_bytes("{", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__146() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__146() { _start: { lean_object* x_1; @@ -1871,23 +1871,23 @@ x_1 = lean_mk_string_from_bytes("optEllipsis", 11); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__147() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__147() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__33; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__146; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__33; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__146; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__148() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__148() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__147; -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__87; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__147; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__87; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1895,7 +1895,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__149() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__149() { _start: { lean_object* x_1; @@ -1903,7 +1903,7 @@ x_1 = lean_mk_string_from_bytes("}", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__150() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__150() { _start: { lean_object* x_1; @@ -1911,7 +1911,7 @@ x_1 = lean_mk_string_from_bytes("else", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__151() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__151() { _start: { lean_object* x_1; @@ -1919,17 +1919,17 @@ x_1 = lean_mk_string_from_bytes("doLetArrow", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__152() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__152() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__33; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__151; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__33; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__151; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__153() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__153() { _start: { lean_object* x_1; @@ -1937,7 +1937,7 @@ x_1 = lean_mk_string_from_bytes("let", 3); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__154() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__154() { _start: { lean_object* x_1; @@ -1945,17 +1945,17 @@ x_1 = lean_mk_string_from_bytes("doIdDecl", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__155() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__155() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__33; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__154; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__33; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__154; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__156() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__156() { _start: { lean_object* x_1; @@ -1963,22 +1963,22 @@ x_1 = lean_mk_string_from_bytes("c", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__157() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__157() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__156; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__156; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__158() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__158() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__156; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__156; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__157; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__157; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1986,17 +1986,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__159() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__159() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__156; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__156; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__160() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__160() { _start: { lean_object* x_1; @@ -2004,7 +2004,7 @@ x_1 = lean_mk_string_from_bytes("←", 3); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__161() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__161() { _start: { lean_object* x_1; @@ -2012,17 +2012,17 @@ x_1 = lean_mk_string_from_bytes("doExpr", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__162() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__162() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__33; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__161; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__33; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__161; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__163() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__163() { _start: { lean_object* x_1; @@ -2030,17 +2030,17 @@ x_1 = lean_mk_string_from_bytes("term_<|_", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__164() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__164() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__163; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__163; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__165() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__165() { _start: { lean_object* x_1; @@ -2048,22 +2048,22 @@ x_1 = lean_mk_string_from_bytes("withoutModifyingState", 21); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__166() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__166() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__165; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__165; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__167() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__167() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__165; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__165; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__166; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__166; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2071,51 +2071,51 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__168() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__168() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__165; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__165; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__169() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__169() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__2; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__165; +x_1 = l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__2; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__165; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__170() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__170() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__169; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__169; 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__171() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__171() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__170; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__170; x_3 = lean_alloc_ctor(1, 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__172() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__172() { _start: { lean_object* x_1; @@ -2123,7 +2123,7 @@ x_1 = lean_mk_string_from_bytes("<|", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__173() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__173() { _start: { lean_object* x_1; @@ -2131,22 +2131,22 @@ x_1 = lean_mk_string_from_bytes("withLCtx", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__174() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__174() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__173; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__173; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__175() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__175() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__173; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__173; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__174; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__174; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2154,17 +2154,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__176() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__176() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__173; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__173; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__177() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__177() { _start: { lean_object* x_1; @@ -2172,51 +2172,51 @@ x_1 = lean_mk_string_from_bytes("Meta", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__178() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__178() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__2; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__177; +x_1 = l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__2; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__177; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__179() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__179() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__178; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__173; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__178; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__173; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__180() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__180() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__179; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__179; 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__181() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__181() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__180; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__180; x_3 = lean_alloc_ctor(1, 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__182() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__182() { _start: { lean_object* x_1; @@ -2224,17 +2224,17 @@ x_1 = lean_mk_string_from_bytes("choice", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__183() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__183() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__182; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__182; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__184() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__184() { _start: { lean_object* x_1; @@ -2242,17 +2242,17 @@ x_1 = lean_mk_string_from_bytes("term{}", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__185() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__185() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__184; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__184; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__186() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__186() { _start: { lean_object* x_1; @@ -2260,22 +2260,22 @@ x_1 = lean_mk_string_from_bytes("withSaveInfoContext", 19); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__187() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__187() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__186; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__186; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__188() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__188() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__186; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__186; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__187; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__187; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2283,51 +2283,51 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__189() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__189() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__186; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__186; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__190() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__190() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__4; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__186; +x_1 = l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__4; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__186; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__191() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__191() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__190; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__190; 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__192() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__192() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__191; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__191; x_3 = lean_alloc_ctor(1, 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__193() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__193() { _start: { lean_object* x_1; @@ -2335,22 +2335,22 @@ x_1 = lean_mk_string_from_bytes("Term.withSynthesize", 19); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__194() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__194() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__193; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__193; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__195() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__195() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__193; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__193; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__194; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__194; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2358,7 +2358,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__196() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__196() { _start: { lean_object* x_1; @@ -2366,51 +2366,51 @@ x_1 = lean_mk_string_from_bytes("withSynthesize", 14); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__197() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__197() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__69; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__196; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__69; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__196; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__198() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__198() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__59; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__196; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__59; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__196; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__199() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__199() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__198; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__198; 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__200() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__200() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__199; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__199; x_3 = lean_alloc_ctor(1, 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__201() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__201() { _start: { lean_object* x_1; @@ -2418,22 +2418,22 @@ x_1 = lean_mk_string_from_bytes("Term.elabTermEnsuringType", 25); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__202() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__202() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__201; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__201; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__203() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__203() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__201; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__201; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__202; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__202; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2441,7 +2441,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__204() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__204() { _start: { lean_object* x_1; @@ -2449,51 +2449,51 @@ x_1 = lean_mk_string_from_bytes("elabTermEnsuringType", 20); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__205() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__205() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__69; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__204; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__69; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__204; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__206() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__206() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__59; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__204; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__59; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__204; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__207() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__207() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__206; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__206; 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__208() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__208() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__207; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__207; x_3 = lean_alloc_ctor(1, 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__209() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__209() { _start: { lean_object* x_1; @@ -2501,17 +2501,17 @@ x_1 = lean_mk_string_from_bytes("arrayRef", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__210() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__210() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__33; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__209; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__33; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__209; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__211() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__211() { _start: { lean_object* x_1; @@ -2519,7 +2519,7 @@ x_1 = lean_mk_string_from_bytes("[", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__212() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__212() { _start: { lean_object* x_1; @@ -2527,17 +2527,17 @@ x_1 = lean_mk_string_from_bytes("num", 3); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__213() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__213() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__212; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__212; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__214() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__214() { _start: { lean_object* x_1; @@ -2545,7 +2545,7 @@ x_1 = lean_mk_string_from_bytes("0", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__215() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__215() { _start: { lean_object* x_1; @@ -2553,7 +2553,7 @@ x_1 = lean_mk_string_from_bytes("3", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__216() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__216() { _start: { lean_object* x_1; @@ -2561,17 +2561,17 @@ x_1 = lean_mk_string_from_bytes("paren", 5); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__217() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__217() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__33; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__216; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__33; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__216; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__218() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__218() { _start: { lean_object* x_1; @@ -2579,22 +2579,22 @@ x_1 = lean_mk_string_from_bytes("Lean.mkConst", 12); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__219() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__219() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__218; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__218; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__220() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__220() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__218; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__218; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__219; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__219; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2602,7 +2602,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__221() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__221() { _start: { lean_object* x_1; @@ -2610,41 +2610,41 @@ x_1 = lean_mk_string_from_bytes("mkConst", 7); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__222() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__222() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__2; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__221; +x_1 = l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__2; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__221; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__223() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__223() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__222; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__222; 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__224() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__224() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__223; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__223; x_3 = lean_alloc_ctor(1, 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__225() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__225() { _start: { lean_object* x_1; @@ -2652,22 +2652,22 @@ x_1 = lean_mk_string_from_bytes("Term.synthesizeSyntheticMVarsNoPostponing", 41) return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__226() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__226() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__225; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__225; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__227() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__227() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__225; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__225; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__226; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__226; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2675,7 +2675,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__228() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__228() { _start: { lean_object* x_1; @@ -2683,51 +2683,51 @@ x_1 = lean_mk_string_from_bytes("synthesizeSyntheticMVarsNoPostponing", 36); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__229() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__229() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__69; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__228; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__69; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__228; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__230() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__230() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__59; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__228; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__59; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__228; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__231() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__231() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__230; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__230; 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__232() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__232() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__231; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__231; x_3 = lean_alloc_ctor(1, 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__233() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__233() { _start: { lean_object* x_1; @@ -2735,22 +2735,22 @@ x_1 = lean_mk_string_from_bytes("instantiateMVars", 16); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__234() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__234() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__233; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__233; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__235() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__235() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__233; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__233; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__234; +x_3 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__234; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2758,65 +2758,65 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__236() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__236() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__233; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__233; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__237() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__237() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__178; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__233; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__178; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__233; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__238() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__238() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__237; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__237; 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__239() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__239() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__238; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__238; x_3 = lean_alloc_ctor(1, 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_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__240() { +static lean_object* _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__240() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__29; -x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__110; +x_1 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__29; +x_2 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__110; x_3 = lean_array_push(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; uint8_t x_5; -x_4 = l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__8; +x_4 = l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__8; lean_inc(x_1); x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); if (x_5 == 0) @@ -2843,77 +2843,77 @@ x_12 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRul x_13 = !lean_is_exclusive(x_12); 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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; +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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; x_14 = lean_ctor_get(x_12, 0); x_15 = lean_ctor_get(x_2, 2); lean_inc(x_15); x_16 = lean_ctor_get(x_2, 1); lean_inc(x_16); lean_dec(x_2); -x_17 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__13; +x_17 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__13; lean_inc(x_14); x_18 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_18, 0, x_14); lean_ctor_set(x_18, 1, x_17); -x_19 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__15; +x_19 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__15; x_20 = lean_array_push(x_19, x_18); x_21 = lean_box(2); -x_22 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__14; +x_22 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__14; x_23 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_23, 0, x_21); lean_ctor_set(x_23, 1, x_22); lean_ctor_set(x_23, 2, x_20); x_24 = lean_array_push(x_19, x_23); -x_25 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__2; +x_25 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__2; x_26 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_26, 0, x_21); lean_ctor_set(x_26, 1, x_25); lean_ctor_set(x_26, 2, x_24); -x_27 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__20; +x_27 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__20; x_28 = lean_array_push(x_27, x_26); -x_29 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__12; +x_29 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__12; x_30 = lean_array_push(x_28, x_29); -x_31 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__10; +x_31 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__10; x_32 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_32, 0, x_21); lean_ctor_set(x_32, 1, x_31); lean_ctor_set(x_32, 2, x_30); -x_33 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__21; +x_33 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__21; lean_inc(x_14); x_34 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_34, 0, x_14); lean_ctor_set(x_34, 1, x_33); -x_35 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__28; +x_35 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__28; lean_inc(x_15); lean_inc(x_16); x_36 = l_Lean_addMacroScope(x_16, x_35, x_15); x_37 = lean_box(0); -x_38 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__27; +x_38 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__27; lean_inc(x_14); x_39 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_39, 0, x_14); lean_ctor_set(x_39, 1, x_38); lean_ctor_set(x_39, 2, x_36); lean_ctor_set(x_39, 3, x_37); -x_40 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__29; +x_40 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__29; lean_inc(x_39); x_41 = lean_array_push(x_40, x_39); x_42 = lean_array_push(x_41, x_29); -x_43 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__24; +x_43 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__24; x_44 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_44, 0, x_21); lean_ctor_set(x_44, 1, x_43); lean_ctor_set(x_44, 2, x_42); -x_45 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__36; +x_45 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__36; lean_inc(x_14); x_46 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_46, 0, x_14); lean_ctor_set(x_46, 1, x_45); -x_47 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__40; +x_47 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__40; lean_inc(x_15); lean_inc(x_16); x_48 = l_Lean_addMacroScope(x_16, x_47, x_15); -x_49 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__39; +x_49 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__39; lean_inc(x_14); x_50 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_50, 0, x_14); @@ -2926,17 +2926,17 @@ x_52 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_52, 0, x_21); lean_ctor_set(x_52, 1, x_25); lean_ctor_set(x_52, 2, x_51); -x_53 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__41; +x_53 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__41; lean_inc(x_14); x_54 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_54, 0, x_14); lean_ctor_set(x_54, 1, x_53); -x_55 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__45; +x_55 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__45; lean_inc(x_15); lean_inc(x_16); x_56 = l_Lean_addMacroScope(x_16, x_55, x_15); -x_57 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__44; -x_58 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__48; +x_57 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__44; +x_58 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__48; lean_inc(x_14); x_59 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_59, 0, x_14); @@ -2950,12 +2950,12 @@ x_62 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_62, 0, x_21); lean_ctor_set(x_62, 1, x_25); lean_ctor_set(x_62, 2, x_61); -x_63 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__49; +x_63 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__49; lean_inc(x_14); x_64 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_64, 0, x_14); lean_ctor_set(x_64, 1, x_63); -x_65 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__50; +x_65 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__50; lean_inc(x_46); x_66 = lean_array_push(x_65, x_46); lean_inc(x_66); @@ -2964,7 +2964,7 @@ x_68 = lean_array_push(x_67, x_62); x_69 = lean_array_push(x_68, x_29); lean_inc(x_64); x_70 = lean_array_push(x_69, x_64); -x_71 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__35; +x_71 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__35; x_72 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_72, 0, x_21); lean_ctor_set(x_72, 1, x_71); @@ -2974,12 +2974,12 @@ x_74 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_74, 0, x_21); lean_ctor_set(x_74, 1, x_25); lean_ctor_set(x_74, 2, x_73); -x_75 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__58; +x_75 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__58; lean_inc(x_15); lean_inc(x_16); x_76 = l_Lean_addMacroScope(x_16, x_75, x_15); -x_77 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__57; -x_78 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__62; +x_77 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__57; +x_78 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__62; lean_inc(x_14); x_79 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_79, 0, x_14); @@ -2994,14 +2994,14 @@ lean_ctor_set(x_81, 1, x_25); lean_ctor_set(x_81, 2, x_80); x_82 = lean_array_push(x_40, x_79); x_83 = lean_array_push(x_82, x_81); -x_84 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__54; +x_84 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__54; x_85 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_85, 0, x_21); lean_ctor_set(x_85, 1, x_84); lean_ctor_set(x_85, 2, x_83); lean_inc(x_60); x_86 = lean_array_push(x_60, x_85); -x_87 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__52; +x_87 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__52; x_88 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_88, 0, x_21); lean_ctor_set(x_88, 1, x_87); @@ -3016,40 +3016,40 @@ x_91 = lean_array_push(x_40, x_74); lean_inc(x_90); lean_inc(x_91); x_92 = lean_array_push(x_91, x_90); -x_93 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__31; +x_93 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__31; x_94 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_94, 0, x_21); lean_ctor_set(x_94, 1, x_93); lean_ctor_set(x_94, 2, x_92); -x_95 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__65; +x_95 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__65; lean_inc(x_14); x_96 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_96, 0, x_14); lean_ctor_set(x_96, 1, x_95); -x_97 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__71; +x_97 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__71; lean_inc(x_15); lean_inc(x_16); x_98 = l_Lean_addMacroScope(x_16, x_97, x_15); -x_99 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__68; -x_100 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__74; +x_99 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__68; +x_100 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__74; lean_inc(x_14); x_101 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_101, 0, x_14); lean_ctor_set(x_101, 1, x_99); lean_ctor_set(x_101, 2, x_98); lean_ctor_set(x_101, 3, x_100); -x_102 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__77; +x_102 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__77; lean_inc(x_14); x_103 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_103, 0, x_14); lean_ctor_set(x_103, 1, x_102); -x_104 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__78; +x_104 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__78; lean_inc(x_103); x_105 = lean_array_push(x_104, x_103); x_106 = lean_array_push(x_105, x_103); lean_inc(x_11); x_107 = lean_array_push(x_106, x_11); -x_108 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__76; +x_108 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__76; x_109 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_109, 0, x_21); lean_ctor_set(x_109, 1, x_108); @@ -3073,12 +3073,12 @@ x_117 = lean_array_push(x_104, x_96); lean_inc(x_117); x_118 = lean_array_push(x_117, x_116); x_119 = lean_array_push(x_118, x_29); -x_120 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__64; +x_120 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__64; x_121 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_121, 0, x_21); lean_ctor_set(x_121, 1, x_120); lean_ctor_set(x_121, 2, x_119); -x_122 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__79; +x_122 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__79; x_123 = lean_array_push(x_122, x_34); lean_inc(x_123); x_124 = lean_array_push(x_123, x_44); @@ -3087,28 +3087,28 @@ x_126 = lean_array_push(x_125, x_121); x_127 = lean_array_push(x_126, x_29); x_128 = lean_array_push(x_127, x_29); x_129 = lean_array_push(x_128, x_29); -x_130 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__22; +x_130 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__22; x_131 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_131, 0, x_21); lean_ctor_set(x_131, 1, x_130); lean_ctor_set(x_131, 2, x_129); x_132 = lean_array_push(x_40, x_32); x_133 = lean_array_push(x_132, x_131); -x_134 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__8; +x_134 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__8; x_135 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_135, 0, x_21); lean_ctor_set(x_135, 1, x_134); lean_ctor_set(x_135, 2, x_133); -x_136 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__82; +x_136 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__82; lean_inc(x_14); x_137 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_137, 0, x_14); lean_ctor_set(x_137, 1, x_136); -x_138 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__96; +x_138 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__96; lean_inc(x_15); lean_inc(x_16); x_139 = l_Lean_addMacroScope(x_16, x_138, x_15); -x_140 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__95; +x_140 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__95; lean_inc(x_14); x_141 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_141, 0, x_14); @@ -3122,14 +3122,14 @@ lean_ctor_set(x_143, 1, x_25); lean_ctor_set(x_143, 2, x_142); x_144 = lean_array_push(x_40, x_141); x_145 = lean_array_push(x_144, x_143); -x_146 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__92; +x_146 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__92; x_147 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_147, 0, x_21); lean_ctor_set(x_147, 1, x_146); lean_ctor_set(x_147, 2, x_145); -x_148 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__97; +x_148 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__97; x_149 = lean_array_push(x_148, x_147); -x_150 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__84; +x_150 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__84; x_151 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_151, 0, x_21); lean_ctor_set(x_151, 1, x_150); @@ -3139,7 +3139,7 @@ x_153 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_153, 0, x_21); lean_ctor_set(x_153, 1, x_25); lean_ctor_set(x_153, 2, x_152); -x_154 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__98; +x_154 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__98; lean_inc(x_14); x_155 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_155, 0, x_14); @@ -3148,7 +3148,7 @@ x_156 = lean_array_push(x_104, x_137); x_157 = lean_array_push(x_156, x_153); lean_inc(x_155); x_158 = lean_array_push(x_157, x_155); -x_159 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__81; +x_159 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__81; x_160 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_160, 0, x_21); lean_ctor_set(x_160, 1, x_159); @@ -3158,7 +3158,7 @@ x_162 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_162, 0, x_21); lean_ctor_set(x_162, 1, x_25); lean_ctor_set(x_162, 2, x_161); -x_163 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__17; +x_163 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__17; x_164 = lean_array_push(x_163, x_162); x_165 = lean_array_push(x_164, x_29); x_166 = lean_array_push(x_165, x_29); @@ -3168,16 +3168,16 @@ x_169 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_169, 0, x_21); lean_ctor_set(x_169, 1, x_31); lean_ctor_set(x_169, 2, x_168); -x_170 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__99; +x_170 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__99; lean_inc(x_14); x_171 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_171, 0, x_14); lean_ctor_set(x_171, 1, x_170); -x_172 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__104; +x_172 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__104; lean_inc(x_15); lean_inc(x_16); x_173 = l_Lean_addMacroScope(x_16, x_172, x_15); -x_174 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__103; +x_174 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__103; lean_inc(x_14); x_175 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_175, 0, x_14); @@ -3192,17 +3192,17 @@ lean_ctor_set(x_178, 0, x_21); lean_ctor_set(x_178, 1, x_43); lean_ctor_set(x_178, 2, x_177); x_179 = lean_array_push(x_91, x_88); -x_180 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__106; +x_180 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__106; x_181 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_181, 0, x_21); lean_ctor_set(x_181, 1, x_180); lean_ctor_set(x_181, 2, x_179); -x_182 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__107; +x_182 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__107; x_183 = lean_array_push(x_182, x_171); x_184 = lean_array_push(x_183, x_178); x_185 = lean_array_push(x_184, x_181); x_186 = lean_array_push(x_185, x_29); -x_187 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__100; +x_187 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__100; x_188 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_188, 0, x_21); lean_ctor_set(x_188, 1, x_187); @@ -3213,1649 +3213,1661 @@ x_191 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_191, 0, x_21); lean_ctor_set(x_191, 1, x_134); lean_ctor_set(x_191, 2, x_190); -x_192 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__114; +x_192 = lean_array_push(x_40, x_9); +x_193 = lean_array_push(x_192, x_29); +x_194 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_194, 0, x_21); +lean_ctor_set(x_194, 1, x_43); +lean_ctor_set(x_194, 2, x_193); +x_195 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__114; lean_inc(x_15); lean_inc(x_16); -x_193 = l_Lean_addMacroScope(x_16, x_192, x_15); -x_194 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__113; +x_196 = l_Lean_addMacroScope(x_16, x_195, x_15); +x_197 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__113; lean_inc(x_14); -x_195 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_195, 0, x_14); -lean_ctor_set(x_195, 1, x_194); -lean_ctor_set(x_195, 2, x_193); -lean_ctor_set(x_195, 3, x_37); -lean_inc(x_195); -x_196 = lean_array_push(x_19, x_195); -x_197 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_197, 0, x_21); -lean_ctor_set(x_197, 1, x_25); -lean_ctor_set(x_197, 2, x_196); -x_198 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__118; +x_198 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_198, 0, x_14); +lean_ctor_set(x_198, 1, x_197); +lean_ctor_set(x_198, 2, x_196); +lean_ctor_set(x_198, 3, x_37); +lean_inc(x_198); +x_199 = lean_array_push(x_19, x_198); +x_200 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_200, 0, x_21); +lean_ctor_set(x_200, 1, x_25); +lean_ctor_set(x_200, 2, x_199); +x_201 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__118; lean_inc(x_15); lean_inc(x_16); -x_199 = l_Lean_addMacroScope(x_16, x_198, x_15); -x_200 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__117; -x_201 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__121; +x_202 = l_Lean_addMacroScope(x_16, x_201, x_15); +x_203 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__117; +x_204 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__121; lean_inc(x_14); -x_202 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_202, 0, x_14); -lean_ctor_set(x_202, 1, x_200); -lean_ctor_set(x_202, 2, x_199); -lean_ctor_set(x_202, 3, x_201); +x_205 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_205, 0, x_14); +lean_ctor_set(x_205, 1, x_203); +lean_ctor_set(x_205, 2, x_202); +lean_ctor_set(x_205, 3, x_204); lean_inc(x_60); -x_203 = lean_array_push(x_60, x_202); -x_204 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_204, 0, x_21); -lean_ctor_set(x_204, 1, x_25); -lean_ctor_set(x_204, 2, x_203); -x_205 = lean_array_push(x_66, x_197); -x_206 = lean_array_push(x_205, x_204); -x_207 = lean_array_push(x_206, x_29); +x_206 = lean_array_push(x_60, x_205); +x_207 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_207, 0, x_21); +lean_ctor_set(x_207, 1, x_25); +lean_ctor_set(x_207, 2, x_206); +x_208 = lean_array_push(x_66, x_200); +x_209 = lean_array_push(x_208, x_207); +x_210 = lean_array_push(x_209, x_29); lean_inc(x_64); -x_208 = lean_array_push(x_207, x_64); -x_209 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_209, 0, x_21); -lean_ctor_set(x_209, 1, x_71); -lean_ctor_set(x_209, 2, x_208); -x_210 = lean_array_push(x_19, x_209); -x_211 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_211, 0, x_21); -lean_ctor_set(x_211, 1, x_25); -lean_ctor_set(x_211, 2, x_210); -x_212 = lean_array_push(x_40, x_211); -x_213 = lean_array_push(x_212, x_90); +x_211 = lean_array_push(x_210, x_64); +x_212 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_212, 0, x_21); +lean_ctor_set(x_212, 1, x_71); +lean_ctor_set(x_212, 2, x_211); +x_213 = lean_array_push(x_19, x_212); x_214 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_214, 0, x_21); -lean_ctor_set(x_214, 1, x_93); +lean_ctor_set(x_214, 1, x_25); lean_ctor_set(x_214, 2, x_213); -x_215 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__122; +x_215 = lean_array_push(x_40, x_214); +x_216 = lean_array_push(x_215, x_90); +x_217 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_217, 0, x_21); +lean_ctor_set(x_217, 1, x_93); +lean_ctor_set(x_217, 2, x_216); +x_218 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__122; lean_inc(x_14); -x_216 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_216, 0, x_14); -lean_ctor_set(x_216, 1, x_215); -x_217 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__130; +x_219 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_219, 0, x_14); +lean_ctor_set(x_219, 1, x_218); +x_220 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__130; lean_inc(x_14); -x_218 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_218, 0, x_14); -lean_ctor_set(x_218, 1, x_217); -x_219 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__137; +x_221 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_221, 0, x_14); +lean_ctor_set(x_221, 1, x_220); +x_222 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__137; lean_inc(x_15); lean_inc(x_16); -x_220 = l_Lean_addMacroScope(x_16, x_219, x_15); -x_221 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__135; +x_223 = l_Lean_addMacroScope(x_16, x_222, x_15); +x_224 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__135; lean_inc(x_14); -x_222 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_222, 0, x_14); -lean_ctor_set(x_222, 1, x_221); -lean_ctor_set(x_222, 2, x_220); -lean_ctor_set(x_222, 3, x_37); -x_223 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__138; -x_224 = lean_array_push(x_223, x_222); -x_225 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__132; -x_226 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_226, 0, x_21); -lean_ctor_set(x_226, 1, x_225); -lean_ctor_set(x_226, 2, x_224); -x_227 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__139; +x_225 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_225, 0, x_14); +lean_ctor_set(x_225, 1, x_224); +lean_ctor_set(x_225, 2, x_223); +lean_ctor_set(x_225, 3, x_37); +x_226 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__138; +x_227 = lean_array_push(x_226, x_225); +x_228 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__132; +x_229 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_229, 0, x_21); +lean_ctor_set(x_229, 1, x_228); +lean_ctor_set(x_229, 2, x_227); +x_230 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__139; lean_inc(x_14); -x_228 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_228, 0, x_14); -lean_ctor_set(x_228, 1, x_227); -x_229 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__142; +x_231 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_231, 0, x_14); +lean_ctor_set(x_231, 1, x_230); +x_232 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__142; lean_inc(x_14); -x_230 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_230, 0, x_14); -lean_ctor_set(x_230, 1, x_229); -x_231 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__145; +x_233 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_233, 0, x_14); +lean_ctor_set(x_233, 1, x_232); +x_234 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__145; lean_inc(x_14); -x_232 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_232, 0, x_14); -lean_ctor_set(x_232, 1, x_231); -x_233 = lean_array_push(x_60, x_11); -x_234 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_234, 0, x_21); -lean_ctor_set(x_234, 1, x_25); -lean_ctor_set(x_234, 2, x_233); -x_235 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__149; +x_235 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_235, 0, x_14); +lean_ctor_set(x_235, 1, x_234); +x_236 = lean_array_push(x_60, x_11); +x_237 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_237, 0, x_21); +lean_ctor_set(x_237, 1, x_25); +lean_ctor_set(x_237, 2, x_236); +x_238 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__149; lean_inc(x_14); -x_236 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_236, 0, x_14); -lean_ctor_set(x_236, 1, x_235); -x_237 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__16; -lean_inc(x_232); -x_238 = lean_array_push(x_237, x_232); -x_239 = lean_array_push(x_238, x_29); -x_240 = lean_array_push(x_239, x_29); -x_241 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__148; -x_242 = lean_array_push(x_240, x_241); -lean_inc(x_242); -x_243 = lean_array_push(x_242, x_234); -lean_inc(x_236); -x_244 = lean_array_push(x_243, x_236); -x_245 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__144; -x_246 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_246, 0, x_21); -lean_ctor_set(x_246, 1, x_245); -lean_ctor_set(x_246, 2, x_244); -x_247 = lean_array_push(x_19, x_246); -x_248 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_248, 0, x_21); -lean_ctor_set(x_248, 1, x_25); -lean_ctor_set(x_248, 2, x_247); -x_249 = lean_array_push(x_40, x_230); -x_250 = lean_array_push(x_249, x_248); -x_251 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__141; -x_252 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_252, 0, x_21); -lean_ctor_set(x_252, 1, x_251); -lean_ctor_set(x_252, 2, x_250); -x_253 = lean_array_push(x_40, x_252); -x_254 = lean_array_push(x_253, x_29); -x_255 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__127; -x_256 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_256, 0, x_21); -lean_ctor_set(x_256, 1, x_255); -lean_ctor_set(x_256, 2, x_254); -x_257 = lean_array_push(x_19, x_256); -x_258 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_258, 0, x_21); -lean_ctor_set(x_258, 1, x_25); -lean_ctor_set(x_258, 2, x_257); -x_259 = lean_array_push(x_19, x_258); -x_260 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__125; +x_239 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_239, 0, x_14); +lean_ctor_set(x_239, 1, x_238); +x_240 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__16; +lean_inc(x_235); +x_241 = lean_array_push(x_240, x_235); +x_242 = lean_array_push(x_241, x_29); +x_243 = lean_array_push(x_242, x_29); +x_244 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__148; +x_245 = lean_array_push(x_243, x_244); +lean_inc(x_245); +x_246 = lean_array_push(x_245, x_237); +lean_inc(x_239); +x_247 = lean_array_push(x_246, x_239); +x_248 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__144; +x_249 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_249, 0, x_21); +lean_ctor_set(x_249, 1, x_248); +lean_ctor_set(x_249, 2, x_247); +x_250 = lean_array_push(x_19, x_249); +x_251 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_251, 0, x_21); +lean_ctor_set(x_251, 1, x_25); +lean_ctor_set(x_251, 2, x_250); +x_252 = lean_array_push(x_40, x_233); +x_253 = lean_array_push(x_252, x_251); +x_254 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__141; +x_255 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_255, 0, x_21); +lean_ctor_set(x_255, 1, x_254); +lean_ctor_set(x_255, 2, x_253); +x_256 = lean_array_push(x_40, x_255); +x_257 = lean_array_push(x_256, x_29); +x_258 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__127; +x_259 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_259, 0, x_21); +lean_ctor_set(x_259, 1, x_258); +lean_ctor_set(x_259, 2, x_257); +x_260 = lean_array_push(x_19, x_259); x_261 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_261, 0, x_21); -lean_ctor_set(x_261, 1, x_260); -lean_ctor_set(x_261, 2, x_259); -x_262 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__150; +lean_ctor_set(x_261, 1, x_25); +lean_ctor_set(x_261, 2, x_260); +x_262 = lean_array_push(x_19, x_261); +x_263 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__125; +x_264 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_264, 0, x_21); +lean_ctor_set(x_264, 1, x_263); +lean_ctor_set(x_264, 2, x_262); +x_265 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__150; lean_inc(x_14); -x_263 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_263, 0, x_14); -lean_ctor_set(x_263, 1, x_262); -x_264 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__153; +x_266 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_266, 0, x_14); +lean_ctor_set(x_266, 1, x_265); +x_267 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__153; lean_inc(x_14); -x_265 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_265, 0, x_14); -lean_ctor_set(x_265, 1, x_264); -x_266 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__159; +x_268 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_268, 0, x_14); +lean_ctor_set(x_268, 1, x_267); +x_269 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__159; lean_inc(x_15); lean_inc(x_16); -x_267 = l_Lean_addMacroScope(x_16, x_266, x_15); -x_268 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__158; +x_270 = l_Lean_addMacroScope(x_16, x_269, x_15); +x_271 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__158; lean_inc(x_14); -x_269 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_269, 0, x_14); -lean_ctor_set(x_269, 1, x_268); -lean_ctor_set(x_269, 2, x_267); -lean_ctor_set(x_269, 3, x_37); -x_270 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__160; +x_272 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_272, 0, x_14); +lean_ctor_set(x_272, 1, x_271); +lean_ctor_set(x_272, 2, x_270); +lean_ctor_set(x_272, 3, x_37); +x_273 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__160; lean_inc(x_14); -x_271 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_271, 0, x_14); -lean_ctor_set(x_271, 1, x_270); -x_272 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__168; +x_274 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_274, 0, x_14); +lean_ctor_set(x_274, 1, x_273); +x_275 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__168; lean_inc(x_15); lean_inc(x_16); -x_273 = l_Lean_addMacroScope(x_16, x_272, x_15); -x_274 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__167; -x_275 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__171; +x_276 = l_Lean_addMacroScope(x_16, x_275, x_15); +x_277 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__167; +x_278 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__171; lean_inc(x_14); -x_276 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_276, 0, x_14); -lean_ctor_set(x_276, 1, x_274); -lean_ctor_set(x_276, 2, x_273); -lean_ctor_set(x_276, 3, x_275); -x_277 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__172; +x_279 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_279, 0, x_14); +lean_ctor_set(x_279, 1, x_277); +lean_ctor_set(x_279, 2, x_276); +lean_ctor_set(x_279, 3, x_278); +x_280 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__172; lean_inc(x_14); -x_278 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_278, 0, x_14); -lean_ctor_set(x_278, 1, x_277); -x_279 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__176; +x_281 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_281, 0, x_14); +lean_ctor_set(x_281, 1, x_280); +x_282 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__176; lean_inc(x_15); lean_inc(x_16); -x_280 = l_Lean_addMacroScope(x_16, x_279, x_15); -x_281 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__175; -x_282 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__181; +x_283 = l_Lean_addMacroScope(x_16, x_282, x_15); +x_284 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__175; +x_285 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__181; lean_inc(x_14); -x_283 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_283, 0, x_14); -lean_ctor_set(x_283, 1, x_281); -lean_ctor_set(x_283, 2, x_280); -lean_ctor_set(x_283, 3, x_282); -x_284 = lean_array_push(x_40, x_232); -lean_inc(x_236); -x_285 = lean_array_push(x_284, x_236); -x_286 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__185; -x_287 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_287, 0, x_21); -lean_ctor_set(x_287, 1, x_286); -lean_ctor_set(x_287, 2, x_285); -x_288 = lean_array_push(x_242, x_29); -x_289 = lean_array_push(x_288, x_236); +x_286 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_286, 0, x_14); +lean_ctor_set(x_286, 1, x_284); +lean_ctor_set(x_286, 2, x_283); +lean_ctor_set(x_286, 3, x_285); +x_287 = lean_array_push(x_40, x_235); +lean_inc(x_239); +x_288 = lean_array_push(x_287, x_239); +x_289 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__185; x_290 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_290, 0, x_21); -lean_ctor_set(x_290, 1, x_245); -lean_ctor_set(x_290, 2, x_289); -x_291 = lean_array_push(x_40, x_287); -x_292 = lean_array_push(x_291, x_290); -x_293 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__183; -x_294 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_294, 0, x_21); -lean_ctor_set(x_294, 1, x_293); -lean_ctor_set(x_294, 2, x_292); -lean_inc(x_294); -x_295 = lean_array_push(x_40, x_294); -x_296 = lean_array_push(x_295, x_294); +lean_ctor_set(x_290, 1, x_289); +lean_ctor_set(x_290, 2, x_288); +x_291 = lean_array_push(x_245, x_29); +x_292 = lean_array_push(x_291, x_239); +x_293 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_293, 0, x_21); +lean_ctor_set(x_293, 1, x_248); +lean_ctor_set(x_293, 2, x_292); +x_294 = lean_array_push(x_40, x_290); +x_295 = lean_array_push(x_294, x_293); +x_296 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__183; x_297 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_297, 0, x_21); -lean_ctor_set(x_297, 1, x_25); -lean_ctor_set(x_297, 2, x_296); -x_298 = lean_array_push(x_40, x_283); +lean_ctor_set(x_297, 1, x_296); +lean_ctor_set(x_297, 2, x_295); +lean_inc(x_297); +x_298 = lean_array_push(x_40, x_297); x_299 = lean_array_push(x_298, x_297); x_300 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_300, 0, x_21); -lean_ctor_set(x_300, 1, x_84); +lean_ctor_set(x_300, 1, x_25); lean_ctor_set(x_300, 2, x_299); -x_301 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__189; +x_301 = lean_array_push(x_40, x_286); +x_302 = lean_array_push(x_301, x_300); +x_303 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_303, 0, x_21); +lean_ctor_set(x_303, 1, x_84); +lean_ctor_set(x_303, 2, x_302); +x_304 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__189; lean_inc(x_15); lean_inc(x_16); -x_302 = l_Lean_addMacroScope(x_16, x_301, x_15); -x_303 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__188; -x_304 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__192; +x_305 = l_Lean_addMacroScope(x_16, x_304, x_15); +x_306 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__188; +x_307 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__192; lean_inc(x_14); -x_305 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_305, 0, x_14); -lean_ctor_set(x_305, 1, x_303); -lean_ctor_set(x_305, 2, x_302); -lean_ctor_set(x_305, 3, x_304); -x_306 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__197; +x_308 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_308, 0, x_14); +lean_ctor_set(x_308, 1, x_306); +lean_ctor_set(x_308, 2, x_305); +lean_ctor_set(x_308, 3, x_307); +x_309 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__197; lean_inc(x_15); lean_inc(x_16); -x_307 = l_Lean_addMacroScope(x_16, x_306, x_15); -x_308 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__195; -x_309 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__200; +x_310 = l_Lean_addMacroScope(x_16, x_309, x_15); +x_311 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__195; +x_312 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__200; lean_inc(x_14); -x_310 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_310, 0, x_14); -lean_ctor_set(x_310, 1, x_308); -lean_ctor_set(x_310, 2, x_307); -lean_ctor_set(x_310, 3, x_309); -x_311 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__205; +x_313 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_313, 0, x_14); +lean_ctor_set(x_313, 1, x_311); +lean_ctor_set(x_313, 2, x_310); +lean_ctor_set(x_313, 3, x_312); +x_314 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__205; lean_inc(x_15); lean_inc(x_16); -x_312 = l_Lean_addMacroScope(x_16, x_311, x_15); -x_313 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__203; -x_314 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__208; +x_315 = l_Lean_addMacroScope(x_16, x_314, x_15); +x_316 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__203; +x_317 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__208; lean_inc(x_14); -x_315 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_315, 0, x_14); -lean_ctor_set(x_315, 1, x_313); -lean_ctor_set(x_315, 2, x_312); -lean_ctor_set(x_315, 3, x_314); -x_316 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__211; +x_318 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_318, 0, x_14); +lean_ctor_set(x_318, 1, x_316); +lean_ctor_set(x_318, 2, x_315); +lean_ctor_set(x_318, 3, x_317); +x_319 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__211; lean_inc(x_14); -x_317 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_317, 0, x_14); -lean_ctor_set(x_317, 1, x_316); -x_318 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__214; +x_320 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_320, 0, x_14); +lean_ctor_set(x_320, 1, x_319); +x_321 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__214; lean_inc(x_14); -x_319 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_319, 0, x_14); -lean_ctor_set(x_319, 1, x_318); -x_320 = lean_array_push(x_19, x_319); -x_321 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__213; -x_322 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_322, 0, x_21); +x_322 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_322, 0, x_14); lean_ctor_set(x_322, 1, x_321); -lean_ctor_set(x_322, 2, x_320); -x_323 = lean_array_push(x_182, x_195); -lean_inc(x_317); -x_324 = lean_array_push(x_323, x_317); -x_325 = lean_array_push(x_324, x_322); +x_323 = lean_array_push(x_19, x_322); +x_324 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__213; +x_325 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_325, 0, x_21); +lean_ctor_set(x_325, 1, x_324); +lean_ctor_set(x_325, 2, x_323); +x_326 = lean_array_push(x_182, x_198); +lean_inc(x_320); +x_327 = lean_array_push(x_326, x_320); +x_328 = lean_array_push(x_327, x_325); lean_inc(x_155); -x_326 = lean_array_push(x_325, x_155); -x_327 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__210; -x_328 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_328, 0, x_21); -lean_ctor_set(x_328, 1, x_327); -lean_ctor_set(x_328, 2, x_326); -x_329 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__215; +x_329 = lean_array_push(x_328, x_155); +x_330 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__210; +x_331 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_331, 0, x_21); +lean_ctor_set(x_331, 1, x_330); +lean_ctor_set(x_331, 2, x_329); +x_332 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__215; lean_inc(x_14); -x_330 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_330, 0, x_14); -lean_ctor_set(x_330, 1, x_329); -x_331 = lean_array_push(x_19, x_330); -x_332 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_332, 0, x_21); -lean_ctor_set(x_332, 1, x_321); -lean_ctor_set(x_332, 2, x_331); -x_333 = lean_array_push(x_182, x_328); -x_334 = lean_array_push(x_333, x_317); -x_335 = lean_array_push(x_334, x_332); -x_336 = lean_array_push(x_335, x_155); -x_337 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_337, 0, x_21); -lean_ctor_set(x_337, 1, x_327); -lean_ctor_set(x_337, 2, x_336); -x_338 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__222; +x_333 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_333, 0, x_14); +lean_ctor_set(x_333, 1, x_332); +x_334 = lean_array_push(x_19, x_333); +x_335 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_335, 0, x_21); +lean_ctor_set(x_335, 1, x_324); +lean_ctor_set(x_335, 2, x_334); +x_336 = lean_array_push(x_182, x_331); +x_337 = lean_array_push(x_336, x_320); +x_338 = lean_array_push(x_337, x_335); +x_339 = lean_array_push(x_338, x_155); +x_340 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_340, 0, x_21); +lean_ctor_set(x_340, 1, x_330); +lean_ctor_set(x_340, 2, x_339); +x_341 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__222; lean_inc(x_15); lean_inc(x_16); -x_339 = l_Lean_addMacroScope(x_16, x_338, x_15); -x_340 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__220; -x_341 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__224; +x_342 = l_Lean_addMacroScope(x_16, x_341, x_15); +x_343 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__220; +x_344 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__224; lean_inc(x_14); -x_342 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_342, 0, x_14); -lean_ctor_set(x_342, 1, x_340); -lean_ctor_set(x_342, 2, x_339); -lean_ctor_set(x_342, 3, x_341); -x_343 = lean_array_push(x_19, x_109); -x_344 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_344, 0, x_21); -lean_ctor_set(x_344, 1, x_25); -lean_ctor_set(x_344, 2, x_343); -x_345 = lean_array_push(x_40, x_342); -x_346 = lean_array_push(x_345, x_344); +x_345 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_345, 0, x_14); +lean_ctor_set(x_345, 1, x_343); +lean_ctor_set(x_345, 2, x_342); +lean_ctor_set(x_345, 3, x_344); +x_346 = lean_array_push(x_19, x_109); x_347 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_347, 0, x_21); -lean_ctor_set(x_347, 1, x_84); +lean_ctor_set(x_347, 1, x_25); lean_ctor_set(x_347, 2, x_346); -x_348 = lean_array_push(x_40, x_347); -x_349 = lean_array_push(x_348, x_29); +x_348 = lean_array_push(x_40, x_345); +x_349 = lean_array_push(x_348, x_347); x_350 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_350, 0, x_21); -lean_ctor_set(x_350, 1, x_25); +lean_ctor_set(x_350, 1, x_84); lean_ctor_set(x_350, 2, x_349); -x_351 = lean_array_push(x_104, x_46); -x_352 = lean_array_push(x_351, x_350); -x_353 = lean_array_push(x_352, x_64); -x_354 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__217; -x_355 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_355, 0, x_21); -lean_ctor_set(x_355, 1, x_354); -lean_ctor_set(x_355, 2, x_353); -x_356 = lean_array_push(x_40, x_337); -x_357 = lean_array_push(x_356, x_355); +x_351 = lean_array_push(x_40, x_350); +x_352 = lean_array_push(x_351, x_29); +x_353 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_353, 0, x_21); +lean_ctor_set(x_353, 1, x_25); +lean_ctor_set(x_353, 2, x_352); +x_354 = lean_array_push(x_104, x_46); +x_355 = lean_array_push(x_354, x_353); +x_356 = lean_array_push(x_355, x_64); +x_357 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__217; x_358 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_358, 0, x_21); -lean_ctor_set(x_358, 1, x_25); -lean_ctor_set(x_358, 2, x_357); -x_359 = lean_array_push(x_40, x_315); +lean_ctor_set(x_358, 1, x_357); +lean_ctor_set(x_358, 2, x_356); +x_359 = lean_array_push(x_40, x_340); x_360 = lean_array_push(x_359, x_358); x_361 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_361, 0, x_21); -lean_ctor_set(x_361, 1, x_84); +lean_ctor_set(x_361, 1, x_25); lean_ctor_set(x_361, 2, x_360); -x_362 = lean_array_push(x_19, x_361); -x_363 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__162; +x_362 = lean_array_push(x_40, x_318); +x_363 = lean_array_push(x_362, x_361); x_364 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_364, 0, x_21); -lean_ctor_set(x_364, 1, x_363); -lean_ctor_set(x_364, 2, x_362); -lean_inc(x_269); -x_365 = lean_array_push(x_182, x_269); -x_366 = lean_array_push(x_365, x_29); -x_367 = lean_array_push(x_366, x_271); -lean_inc(x_367); -x_368 = lean_array_push(x_367, x_364); -x_369 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__155; -x_370 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_370, 0, x_21); -lean_ctor_set(x_370, 1, x_369); -lean_ctor_set(x_370, 2, x_368); -x_371 = lean_array_push(x_104, x_265); -x_372 = lean_array_push(x_371, x_29); -lean_inc(x_372); -x_373 = lean_array_push(x_372, x_370); -x_374 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__152; -x_375 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_375, 0, x_21); -lean_ctor_set(x_375, 1, x_374); -lean_ctor_set(x_375, 2, x_373); -x_376 = lean_array_push(x_40, x_375); -x_377 = lean_array_push(x_376, x_29); +lean_ctor_set(x_364, 1, x_84); +lean_ctor_set(x_364, 2, x_363); +x_365 = lean_array_push(x_19, x_364); +x_366 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__162; +x_367 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_367, 0, x_21); +lean_ctor_set(x_367, 1, x_366); +lean_ctor_set(x_367, 2, x_365); +lean_inc(x_272); +x_368 = lean_array_push(x_182, x_272); +x_369 = lean_array_push(x_368, x_29); +x_370 = lean_array_push(x_369, x_274); +lean_inc(x_370); +x_371 = lean_array_push(x_370, x_367); +x_372 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__155; +x_373 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_373, 0, x_21); +lean_ctor_set(x_373, 1, x_372); +lean_ctor_set(x_373, 2, x_371); +x_374 = lean_array_push(x_104, x_268); +x_375 = lean_array_push(x_374, x_29); +lean_inc(x_375); +x_376 = lean_array_push(x_375, x_373); +x_377 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__152; x_378 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_378, 0, x_21); -lean_ctor_set(x_378, 1, x_255); -lean_ctor_set(x_378, 2, x_377); -x_379 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__229; +lean_ctor_set(x_378, 1, x_377); +lean_ctor_set(x_378, 2, x_376); +x_379 = lean_array_push(x_40, x_378); +x_380 = lean_array_push(x_379, x_29); +x_381 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_381, 0, x_21); +lean_ctor_set(x_381, 1, x_258); +lean_ctor_set(x_381, 2, x_380); +x_382 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__229; lean_inc(x_15); lean_inc(x_16); -x_380 = l_Lean_addMacroScope(x_16, x_379, x_15); -x_381 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__227; -x_382 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__232; +x_383 = l_Lean_addMacroScope(x_16, x_382, x_15); +x_384 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__227; +x_385 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__232; lean_inc(x_14); -x_383 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_383, 0, x_14); -lean_ctor_set(x_383, 1, x_381); -lean_ctor_set(x_383, 2, x_380); -lean_ctor_set(x_383, 3, x_382); -x_384 = lean_array_push(x_19, x_383); -x_385 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_385, 0, x_21); -lean_ctor_set(x_385, 1, x_363); -lean_ctor_set(x_385, 2, x_384); -x_386 = lean_array_push(x_40, x_385); -x_387 = lean_array_push(x_386, x_29); +x_386 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_386, 0, x_14); +lean_ctor_set(x_386, 1, x_384); +lean_ctor_set(x_386, 2, x_383); +lean_ctor_set(x_386, 3, x_385); +x_387 = lean_array_push(x_19, x_386); x_388 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_388, 0, x_21); -lean_ctor_set(x_388, 1, x_255); +lean_ctor_set(x_388, 1, x_366); lean_ctor_set(x_388, 2, x_387); -x_389 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__236; -x_390 = l_Lean_addMacroScope(x_16, x_389, x_15); -x_391 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__235; -x_392 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__239; -x_393 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_393, 0, x_14); -lean_ctor_set(x_393, 1, x_391); -lean_ctor_set(x_393, 2, x_390); -lean_ctor_set(x_393, 3, x_392); -x_394 = lean_array_push(x_19, x_269); -x_395 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_395, 0, x_21); -lean_ctor_set(x_395, 1, x_25); -lean_ctor_set(x_395, 2, x_394); -x_396 = lean_array_push(x_40, x_393); -lean_inc(x_395); -x_397 = lean_array_push(x_396, x_395); +x_389 = lean_array_push(x_40, x_388); +x_390 = lean_array_push(x_389, x_29); +x_391 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_391, 0, x_21); +lean_ctor_set(x_391, 1, x_258); +lean_ctor_set(x_391, 2, x_390); +x_392 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__236; +x_393 = l_Lean_addMacroScope(x_16, x_392, x_15); +x_394 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__235; +x_395 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__239; +x_396 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_396, 0, x_14); +lean_ctor_set(x_396, 1, x_394); +lean_ctor_set(x_396, 2, x_393); +lean_ctor_set(x_396, 3, x_395); +x_397 = lean_array_push(x_19, x_272); x_398 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_398, 0, x_21); -lean_ctor_set(x_398, 1, x_84); +lean_ctor_set(x_398, 1, x_25); lean_ctor_set(x_398, 2, x_397); -x_399 = lean_array_push(x_19, x_398); -x_400 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_400, 0, x_21); -lean_ctor_set(x_400, 1, x_363); -lean_ctor_set(x_400, 2, x_399); -x_401 = lean_array_push(x_40, x_400); -x_402 = lean_array_push(x_401, x_29); +x_399 = lean_array_push(x_40, x_396); +lean_inc(x_398); +x_400 = lean_array_push(x_399, x_398); +x_401 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_401, 0, x_21); +lean_ctor_set(x_401, 1, x_84); +lean_ctor_set(x_401, 2, x_400); +x_402 = lean_array_push(x_19, x_401); x_403 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_403, 0, x_21); -lean_ctor_set(x_403, 1, x_255); +lean_ctor_set(x_403, 1, x_366); lean_ctor_set(x_403, 2, x_402); -x_404 = lean_array_push(x_104, x_378); -x_405 = lean_array_push(x_404, x_388); -x_406 = lean_array_push(x_405, x_403); -x_407 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_407, 0, x_21); -lean_ctor_set(x_407, 1, x_25); -lean_ctor_set(x_407, 2, x_406); -x_408 = lean_array_push(x_19, x_407); -x_409 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_409, 0, x_21); -lean_ctor_set(x_409, 1, x_260); -lean_ctor_set(x_409, 2, x_408); -x_410 = lean_array_push(x_40, x_216); -lean_inc(x_410); -x_411 = lean_array_push(x_410, x_409); -x_412 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__123; -x_413 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_413, 0, x_21); -lean_ctor_set(x_413, 1, x_412); -lean_ctor_set(x_413, 2, x_411); -x_414 = lean_array_push(x_19, x_413); -x_415 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_415, 0, x_21); -lean_ctor_set(x_415, 1, x_25); -lean_ctor_set(x_415, 2, x_414); -x_416 = lean_array_push(x_40, x_310); -x_417 = lean_array_push(x_416, x_415); +x_404 = lean_array_push(x_40, x_403); +x_405 = lean_array_push(x_404, x_29); +x_406 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_406, 0, x_21); +lean_ctor_set(x_406, 1, x_258); +lean_ctor_set(x_406, 2, x_405); +x_407 = lean_array_push(x_104, x_381); +x_408 = lean_array_push(x_407, x_391); +x_409 = lean_array_push(x_408, x_406); +x_410 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_410, 0, x_21); +lean_ctor_set(x_410, 1, x_25); +lean_ctor_set(x_410, 2, x_409); +x_411 = lean_array_push(x_19, x_410); +x_412 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_412, 0, x_21); +lean_ctor_set(x_412, 1, x_263); +lean_ctor_set(x_412, 2, x_411); +x_413 = lean_array_push(x_40, x_219); +lean_inc(x_413); +x_414 = lean_array_push(x_413, x_412); +x_415 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__123; +x_416 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_416, 0, x_21); +lean_ctor_set(x_416, 1, x_415); +lean_ctor_set(x_416, 2, x_414); +x_417 = lean_array_push(x_19, x_416); x_418 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_418, 0, x_21); -lean_ctor_set(x_418, 1, x_84); +lean_ctor_set(x_418, 1, x_25); lean_ctor_set(x_418, 2, x_417); -x_419 = lean_array_push(x_104, x_305); -lean_inc(x_278); -x_420 = lean_array_push(x_419, x_278); -x_421 = lean_array_push(x_420, x_418); -x_422 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__164; -x_423 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_423, 0, x_21); -lean_ctor_set(x_423, 1, x_422); -lean_ctor_set(x_423, 2, x_421); -x_424 = lean_array_push(x_104, x_300); -lean_inc(x_278); -x_425 = lean_array_push(x_424, x_278); -x_426 = lean_array_push(x_425, x_423); -x_427 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_427, 0, x_21); -lean_ctor_set(x_427, 1, x_422); -lean_ctor_set(x_427, 2, x_426); -x_428 = lean_array_push(x_104, x_276); -x_429 = lean_array_push(x_428, x_278); -x_430 = lean_array_push(x_429, x_427); -x_431 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_431, 0, x_21); -lean_ctor_set(x_431, 1, x_422); -lean_ctor_set(x_431, 2, x_430); -x_432 = lean_array_push(x_19, x_431); -x_433 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_433, 0, x_21); -lean_ctor_set(x_433, 1, x_363); -lean_ctor_set(x_433, 2, x_432); -x_434 = lean_array_push(x_367, x_433); -x_435 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_435, 0, x_21); -lean_ctor_set(x_435, 1, x_369); -lean_ctor_set(x_435, 2, x_434); -x_436 = lean_array_push(x_372, x_435); -x_437 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_437, 0, x_21); -lean_ctor_set(x_437, 1, x_374); -lean_ctor_set(x_437, 2, x_436); -x_438 = lean_array_push(x_40, x_437); -x_439 = lean_array_push(x_438, x_29); +x_419 = lean_array_push(x_40, x_313); +x_420 = lean_array_push(x_419, x_418); +x_421 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_421, 0, x_21); +lean_ctor_set(x_421, 1, x_84); +lean_ctor_set(x_421, 2, x_420); +x_422 = lean_array_push(x_104, x_308); +lean_inc(x_281); +x_423 = lean_array_push(x_422, x_281); +x_424 = lean_array_push(x_423, x_421); +x_425 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__164; +x_426 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_426, 0, x_21); +lean_ctor_set(x_426, 1, x_425); +lean_ctor_set(x_426, 2, x_424); +x_427 = lean_array_push(x_104, x_303); +lean_inc(x_281); +x_428 = lean_array_push(x_427, x_281); +x_429 = lean_array_push(x_428, x_426); +x_430 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_430, 0, x_21); +lean_ctor_set(x_430, 1, x_425); +lean_ctor_set(x_430, 2, x_429); +x_431 = lean_array_push(x_104, x_279); +x_432 = lean_array_push(x_431, x_281); +x_433 = lean_array_push(x_432, x_430); +x_434 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_434, 0, x_21); +lean_ctor_set(x_434, 1, x_425); +lean_ctor_set(x_434, 2, x_433); +x_435 = lean_array_push(x_19, x_434); +x_436 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_436, 0, x_21); +lean_ctor_set(x_436, 1, x_366); +lean_ctor_set(x_436, 2, x_435); +x_437 = lean_array_push(x_370, x_436); +x_438 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_438, 0, x_21); +lean_ctor_set(x_438, 1, x_372); +lean_ctor_set(x_438, 2, x_437); +x_439 = lean_array_push(x_375, x_438); x_440 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_440, 0, x_21); -lean_ctor_set(x_440, 1, x_255); +lean_ctor_set(x_440, 1, x_377); lean_ctor_set(x_440, 2, x_439); -x_441 = lean_array_push(x_176, x_395); -x_442 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_442, 0, x_21); -lean_ctor_set(x_442, 1, x_84); -lean_ctor_set(x_442, 2, x_441); -x_443 = lean_array_push(x_19, x_442); -x_444 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_444, 0, x_21); -lean_ctor_set(x_444, 1, x_363); -lean_ctor_set(x_444, 2, x_443); -x_445 = lean_array_push(x_40, x_444); -x_446 = lean_array_push(x_445, x_29); +x_441 = lean_array_push(x_40, x_440); +x_442 = lean_array_push(x_441, x_29); +x_443 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_443, 0, x_21); +lean_ctor_set(x_443, 1, x_258); +lean_ctor_set(x_443, 2, x_442); +x_444 = lean_array_push(x_176, x_398); +x_445 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_445, 0, x_21); +lean_ctor_set(x_445, 1, x_84); +lean_ctor_set(x_445, 2, x_444); +x_446 = lean_array_push(x_19, x_445); x_447 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_447, 0, x_21); -lean_ctor_set(x_447, 1, x_255); +lean_ctor_set(x_447, 1, x_366); lean_ctor_set(x_447, 2, x_446); -x_448 = lean_array_push(x_40, x_440); -x_449 = lean_array_push(x_448, x_447); +x_448 = lean_array_push(x_40, x_447); +x_449 = lean_array_push(x_448, x_29); x_450 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_450, 0, x_21); -lean_ctor_set(x_450, 1, x_25); +lean_ctor_set(x_450, 1, x_258); lean_ctor_set(x_450, 2, x_449); -x_451 = lean_array_push(x_19, x_450); -x_452 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_452, 0, x_21); -lean_ctor_set(x_452, 1, x_260); -lean_ctor_set(x_452, 2, x_451); -x_453 = lean_array_push(x_40, x_263); -x_454 = lean_array_push(x_453, x_452); +x_451 = lean_array_push(x_40, x_443); +x_452 = lean_array_push(x_451, x_450); +x_453 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_453, 0, x_21); +lean_ctor_set(x_453, 1, x_25); +lean_ctor_set(x_453, 2, x_452); +x_454 = lean_array_push(x_19, x_453); x_455 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_455, 0, x_21); -lean_ctor_set(x_455, 1, x_25); +lean_ctor_set(x_455, 1, x_263); lean_ctor_set(x_455, 2, x_454); -x_456 = lean_array_push(x_237, x_218); -x_457 = lean_array_push(x_456, x_226); -x_458 = lean_array_push(x_457, x_228); -x_459 = lean_array_push(x_458, x_261); -x_460 = lean_array_push(x_459, x_29); -x_461 = lean_array_push(x_460, x_455); -x_462 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__129; -x_463 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_463, 0, x_21); -lean_ctor_set(x_463, 1, x_462); -lean_ctor_set(x_463, 2, x_461); -x_464 = lean_array_push(x_40, x_463); -x_465 = lean_array_push(x_464, x_29); +x_456 = lean_array_push(x_40, x_266); +x_457 = lean_array_push(x_456, x_455); +x_458 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_458, 0, x_21); +lean_ctor_set(x_458, 1, x_25); +lean_ctor_set(x_458, 2, x_457); +x_459 = lean_array_push(x_240, x_221); +x_460 = lean_array_push(x_459, x_229); +x_461 = lean_array_push(x_460, x_231); +x_462 = lean_array_push(x_461, x_264); +x_463 = lean_array_push(x_462, x_29); +x_464 = lean_array_push(x_463, x_458); +x_465 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__129; x_466 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_466, 0, x_21); -lean_ctor_set(x_466, 1, x_255); -lean_ctor_set(x_466, 2, x_465); -x_467 = lean_array_push(x_19, x_466); -x_468 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_468, 0, x_21); -lean_ctor_set(x_468, 1, x_25); -lean_ctor_set(x_468, 2, x_467); -x_469 = lean_array_push(x_19, x_468); -x_470 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_470, 0, x_21); -lean_ctor_set(x_470, 1, x_260); -lean_ctor_set(x_470, 2, x_469); -x_471 = lean_array_push(x_410, x_470); -x_472 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_472, 0, x_21); -lean_ctor_set(x_472, 1, x_412); -lean_ctor_set(x_472, 2, x_471); -x_473 = lean_array_push(x_117, x_472); -x_474 = lean_array_push(x_473, x_29); +lean_ctor_set(x_466, 1, x_465); +lean_ctor_set(x_466, 2, x_464); +x_467 = lean_array_push(x_40, x_466); +x_468 = lean_array_push(x_467, x_29); +x_469 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_469, 0, x_21); +lean_ctor_set(x_469, 1, x_258); +lean_ctor_set(x_469, 2, x_468); +x_470 = lean_array_push(x_19, x_469); +x_471 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_471, 0, x_21); +lean_ctor_set(x_471, 1, x_25); +lean_ctor_set(x_471, 2, x_470); +x_472 = lean_array_push(x_19, x_471); +x_473 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_473, 0, x_21); +lean_ctor_set(x_473, 1, x_263); +lean_ctor_set(x_473, 2, x_472); +x_474 = lean_array_push(x_413, x_473); x_475 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_475, 0, x_21); -lean_ctor_set(x_475, 1, x_120); +lean_ctor_set(x_475, 1, x_415); lean_ctor_set(x_475, 2, x_474); -x_476 = lean_array_push(x_123, x_9); -x_477 = lean_array_push(x_476, x_214); -x_478 = lean_array_push(x_477, x_475); -x_479 = lean_array_push(x_478, x_29); -x_480 = lean_array_push(x_479, x_29); -x_481 = lean_array_push(x_480, x_29); -x_482 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_482, 0, x_21); -lean_ctor_set(x_482, 1, x_130); -lean_ctor_set(x_482, 2, x_481); -x_483 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__240; -x_484 = lean_array_push(x_483, x_482); +x_476 = lean_array_push(x_117, x_475); +x_477 = lean_array_push(x_476, x_29); +x_478 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_478, 0, x_21); +lean_ctor_set(x_478, 1, x_120); +lean_ctor_set(x_478, 2, x_477); +x_479 = lean_array_push(x_123, x_194); +x_480 = lean_array_push(x_479, x_217); +x_481 = lean_array_push(x_480, x_478); +x_482 = lean_array_push(x_481, x_29); +x_483 = lean_array_push(x_482, x_29); +x_484 = lean_array_push(x_483, x_29); x_485 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_485, 0, x_21); -lean_ctor_set(x_485, 1, x_134); +lean_ctor_set(x_485, 1, x_130); lean_ctor_set(x_485, 2, x_484); -x_486 = lean_array_push(x_104, x_135); -x_487 = lean_array_push(x_486, x_191); -x_488 = lean_array_push(x_487, x_485); -x_489 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_489, 0, x_21); -lean_ctor_set(x_489, 1, x_25); -lean_ctor_set(x_489, 2, x_488); -lean_ctor_set(x_12, 0, x_489); +x_486 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__240; +x_487 = lean_array_push(x_486, x_485); +x_488 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_488, 0, x_21); +lean_ctor_set(x_488, 1, x_134); +lean_ctor_set(x_488, 2, x_487); +x_489 = lean_array_push(x_104, x_135); +x_490 = lean_array_push(x_489, x_191); +x_491 = lean_array_push(x_490, x_488); +x_492 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_492, 0, x_21); +lean_ctor_set(x_492, 1, x_25); +lean_ctor_set(x_492, 2, x_491); +lean_ctor_set(x_12, 0, x_492); return x_12; } else { -lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; lean_object* x_565; lean_object* x_566; lean_object* x_567; lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; lean_object* x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; lean_object* x_678; lean_object* x_679; lean_object* x_680; lean_object* x_681; lean_object* x_682; lean_object* x_683; lean_object* x_684; lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; lean_object* x_690; lean_object* x_691; lean_object* x_692; lean_object* x_693; lean_object* x_694; lean_object* x_695; lean_object* x_696; lean_object* x_697; lean_object* x_698; lean_object* x_699; lean_object* x_700; lean_object* x_701; lean_object* x_702; lean_object* x_703; lean_object* x_704; lean_object* x_705; lean_object* x_706; lean_object* x_707; lean_object* x_708; lean_object* x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; lean_object* x_718; lean_object* x_719; lean_object* x_720; lean_object* x_721; lean_object* x_722; lean_object* x_723; lean_object* x_724; lean_object* x_725; lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; lean_object* x_730; lean_object* x_731; lean_object* x_732; lean_object* x_733; lean_object* x_734; lean_object* x_735; lean_object* x_736; lean_object* x_737; lean_object* x_738; lean_object* x_739; lean_object* x_740; lean_object* x_741; lean_object* x_742; lean_object* x_743; lean_object* x_744; lean_object* x_745; lean_object* x_746; lean_object* x_747; lean_object* x_748; lean_object* x_749; lean_object* x_750; lean_object* x_751; lean_object* x_752; lean_object* x_753; lean_object* x_754; lean_object* x_755; lean_object* x_756; lean_object* x_757; lean_object* x_758; lean_object* x_759; lean_object* x_760; lean_object* x_761; lean_object* x_762; lean_object* x_763; lean_object* x_764; lean_object* x_765; lean_object* x_766; lean_object* x_767; lean_object* x_768; lean_object* x_769; lean_object* x_770; lean_object* x_771; lean_object* x_772; lean_object* x_773; lean_object* x_774; lean_object* x_775; lean_object* x_776; lean_object* x_777; lean_object* x_778; lean_object* x_779; lean_object* x_780; lean_object* x_781; lean_object* x_782; lean_object* x_783; lean_object* x_784; lean_object* x_785; lean_object* x_786; lean_object* x_787; lean_object* x_788; lean_object* x_789; lean_object* x_790; lean_object* x_791; lean_object* x_792; lean_object* x_793; lean_object* x_794; lean_object* x_795; lean_object* x_796; lean_object* x_797; lean_object* x_798; lean_object* x_799; lean_object* x_800; lean_object* x_801; lean_object* x_802; lean_object* x_803; lean_object* x_804; lean_object* x_805; lean_object* x_806; lean_object* x_807; lean_object* x_808; lean_object* x_809; lean_object* x_810; lean_object* x_811; lean_object* x_812; lean_object* x_813; lean_object* x_814; lean_object* x_815; lean_object* x_816; lean_object* x_817; lean_object* x_818; lean_object* x_819; lean_object* x_820; lean_object* x_821; lean_object* x_822; lean_object* x_823; lean_object* x_824; lean_object* x_825; lean_object* x_826; lean_object* x_827; lean_object* x_828; lean_object* x_829; lean_object* x_830; lean_object* x_831; lean_object* x_832; lean_object* x_833; lean_object* x_834; lean_object* x_835; lean_object* x_836; lean_object* x_837; lean_object* x_838; lean_object* x_839; lean_object* x_840; lean_object* x_841; lean_object* x_842; lean_object* x_843; lean_object* x_844; lean_object* x_845; lean_object* x_846; lean_object* x_847; lean_object* x_848; lean_object* x_849; lean_object* x_850; lean_object* x_851; lean_object* x_852; lean_object* x_853; lean_object* x_854; lean_object* x_855; lean_object* x_856; lean_object* x_857; lean_object* x_858; lean_object* x_859; lean_object* x_860; lean_object* x_861; lean_object* x_862; lean_object* x_863; lean_object* x_864; lean_object* x_865; lean_object* x_866; lean_object* x_867; lean_object* x_868; lean_object* x_869; lean_object* x_870; lean_object* x_871; lean_object* x_872; lean_object* x_873; lean_object* x_874; lean_object* x_875; lean_object* x_876; lean_object* x_877; lean_object* x_878; lean_object* x_879; lean_object* x_880; lean_object* x_881; lean_object* x_882; lean_object* x_883; lean_object* x_884; lean_object* x_885; lean_object* x_886; lean_object* x_887; lean_object* x_888; lean_object* x_889; lean_object* x_890; lean_object* x_891; lean_object* x_892; lean_object* x_893; lean_object* x_894; lean_object* x_895; lean_object* x_896; lean_object* x_897; lean_object* x_898; lean_object* x_899; lean_object* x_900; lean_object* x_901; lean_object* x_902; lean_object* x_903; lean_object* x_904; lean_object* x_905; lean_object* x_906; lean_object* x_907; lean_object* x_908; lean_object* x_909; lean_object* x_910; lean_object* x_911; lean_object* x_912; lean_object* x_913; lean_object* x_914; lean_object* x_915; lean_object* x_916; lean_object* x_917; lean_object* x_918; lean_object* x_919; lean_object* x_920; lean_object* x_921; lean_object* x_922; lean_object* x_923; lean_object* x_924; lean_object* x_925; lean_object* x_926; lean_object* x_927; lean_object* x_928; lean_object* x_929; lean_object* x_930; lean_object* x_931; lean_object* x_932; lean_object* x_933; lean_object* x_934; lean_object* x_935; lean_object* x_936; lean_object* x_937; lean_object* x_938; lean_object* x_939; lean_object* x_940; lean_object* x_941; lean_object* x_942; lean_object* x_943; lean_object* x_944; lean_object* x_945; lean_object* x_946; lean_object* x_947; lean_object* x_948; lean_object* x_949; lean_object* x_950; lean_object* x_951; lean_object* x_952; lean_object* x_953; lean_object* x_954; lean_object* x_955; lean_object* x_956; lean_object* x_957; lean_object* x_958; lean_object* x_959; lean_object* x_960; lean_object* x_961; lean_object* x_962; lean_object* x_963; lean_object* x_964; lean_object* x_965; lean_object* x_966; lean_object* x_967; -x_490 = lean_ctor_get(x_12, 0); -x_491 = lean_ctor_get(x_12, 1); -lean_inc(x_491); -lean_inc(x_490); -lean_dec(x_12); -x_492 = lean_ctor_get(x_2, 2); -lean_inc(x_492); -x_493 = lean_ctor_get(x_2, 1); +lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; lean_object* x_565; lean_object* x_566; lean_object* x_567; lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; lean_object* x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; lean_object* x_678; lean_object* x_679; lean_object* x_680; lean_object* x_681; lean_object* x_682; lean_object* x_683; lean_object* x_684; lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; lean_object* x_690; lean_object* x_691; lean_object* x_692; lean_object* x_693; lean_object* x_694; lean_object* x_695; lean_object* x_696; lean_object* x_697; lean_object* x_698; lean_object* x_699; lean_object* x_700; lean_object* x_701; lean_object* x_702; lean_object* x_703; lean_object* x_704; lean_object* x_705; lean_object* x_706; lean_object* x_707; lean_object* x_708; lean_object* x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; lean_object* x_718; lean_object* x_719; lean_object* x_720; lean_object* x_721; lean_object* x_722; lean_object* x_723; lean_object* x_724; lean_object* x_725; lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; lean_object* x_730; lean_object* x_731; lean_object* x_732; lean_object* x_733; lean_object* x_734; lean_object* x_735; lean_object* x_736; lean_object* x_737; lean_object* x_738; lean_object* x_739; lean_object* x_740; lean_object* x_741; lean_object* x_742; lean_object* x_743; lean_object* x_744; lean_object* x_745; lean_object* x_746; lean_object* x_747; lean_object* x_748; lean_object* x_749; lean_object* x_750; lean_object* x_751; lean_object* x_752; lean_object* x_753; lean_object* x_754; lean_object* x_755; lean_object* x_756; lean_object* x_757; lean_object* x_758; lean_object* x_759; lean_object* x_760; lean_object* x_761; lean_object* x_762; lean_object* x_763; lean_object* x_764; lean_object* x_765; lean_object* x_766; lean_object* x_767; lean_object* x_768; lean_object* x_769; lean_object* x_770; lean_object* x_771; lean_object* x_772; lean_object* x_773; lean_object* x_774; lean_object* x_775; lean_object* x_776; lean_object* x_777; lean_object* x_778; lean_object* x_779; lean_object* x_780; lean_object* x_781; lean_object* x_782; lean_object* x_783; lean_object* x_784; lean_object* x_785; lean_object* x_786; lean_object* x_787; lean_object* x_788; lean_object* x_789; lean_object* x_790; lean_object* x_791; lean_object* x_792; lean_object* x_793; lean_object* x_794; lean_object* x_795; lean_object* x_796; lean_object* x_797; lean_object* x_798; lean_object* x_799; lean_object* x_800; lean_object* x_801; lean_object* x_802; lean_object* x_803; lean_object* x_804; lean_object* x_805; lean_object* x_806; lean_object* x_807; lean_object* x_808; lean_object* x_809; lean_object* x_810; lean_object* x_811; lean_object* x_812; lean_object* x_813; lean_object* x_814; lean_object* x_815; lean_object* x_816; lean_object* x_817; lean_object* x_818; lean_object* x_819; lean_object* x_820; lean_object* x_821; lean_object* x_822; lean_object* x_823; lean_object* x_824; lean_object* x_825; lean_object* x_826; lean_object* x_827; lean_object* x_828; lean_object* x_829; lean_object* x_830; lean_object* x_831; lean_object* x_832; lean_object* x_833; lean_object* x_834; lean_object* x_835; lean_object* x_836; lean_object* x_837; lean_object* x_838; lean_object* x_839; lean_object* x_840; lean_object* x_841; lean_object* x_842; lean_object* x_843; lean_object* x_844; lean_object* x_845; lean_object* x_846; lean_object* x_847; lean_object* x_848; lean_object* x_849; lean_object* x_850; lean_object* x_851; lean_object* x_852; lean_object* x_853; lean_object* x_854; lean_object* x_855; lean_object* x_856; lean_object* x_857; lean_object* x_858; lean_object* x_859; lean_object* x_860; lean_object* x_861; lean_object* x_862; lean_object* x_863; lean_object* x_864; lean_object* x_865; lean_object* x_866; lean_object* x_867; lean_object* x_868; lean_object* x_869; lean_object* x_870; lean_object* x_871; lean_object* x_872; lean_object* x_873; lean_object* x_874; lean_object* x_875; lean_object* x_876; lean_object* x_877; lean_object* x_878; lean_object* x_879; lean_object* x_880; lean_object* x_881; lean_object* x_882; lean_object* x_883; lean_object* x_884; lean_object* x_885; lean_object* x_886; lean_object* x_887; lean_object* x_888; lean_object* x_889; lean_object* x_890; lean_object* x_891; lean_object* x_892; lean_object* x_893; lean_object* x_894; lean_object* x_895; lean_object* x_896; lean_object* x_897; lean_object* x_898; lean_object* x_899; lean_object* x_900; lean_object* x_901; lean_object* x_902; lean_object* x_903; lean_object* x_904; lean_object* x_905; lean_object* x_906; lean_object* x_907; lean_object* x_908; lean_object* x_909; lean_object* x_910; lean_object* x_911; lean_object* x_912; lean_object* x_913; lean_object* x_914; lean_object* x_915; lean_object* x_916; lean_object* x_917; lean_object* x_918; lean_object* x_919; lean_object* x_920; lean_object* x_921; lean_object* x_922; lean_object* x_923; lean_object* x_924; lean_object* x_925; lean_object* x_926; lean_object* x_927; lean_object* x_928; lean_object* x_929; lean_object* x_930; lean_object* x_931; lean_object* x_932; lean_object* x_933; lean_object* x_934; lean_object* x_935; lean_object* x_936; lean_object* x_937; lean_object* x_938; lean_object* x_939; lean_object* x_940; lean_object* x_941; lean_object* x_942; lean_object* x_943; lean_object* x_944; lean_object* x_945; lean_object* x_946; lean_object* x_947; lean_object* x_948; lean_object* x_949; lean_object* x_950; lean_object* x_951; lean_object* x_952; lean_object* x_953; lean_object* x_954; lean_object* x_955; lean_object* x_956; lean_object* x_957; lean_object* x_958; lean_object* x_959; lean_object* x_960; lean_object* x_961; lean_object* x_962; lean_object* x_963; lean_object* x_964; lean_object* x_965; lean_object* x_966; lean_object* x_967; lean_object* x_968; lean_object* x_969; lean_object* x_970; lean_object* x_971; lean_object* x_972; lean_object* x_973; +x_493 = lean_ctor_get(x_12, 0); +x_494 = lean_ctor_get(x_12, 1); +lean_inc(x_494); lean_inc(x_493); +lean_dec(x_12); +x_495 = lean_ctor_get(x_2, 2); +lean_inc(x_495); +x_496 = lean_ctor_get(x_2, 1); +lean_inc(x_496); lean_dec(x_2); -x_494 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__13; -lean_inc(x_490); -x_495 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_495, 0, x_490); -lean_ctor_set(x_495, 1, x_494); -x_496 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__15; -x_497 = lean_array_push(x_496, x_495); -x_498 = lean_box(2); -x_499 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__14; -x_500 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_500, 0, x_498); -lean_ctor_set(x_500, 1, x_499); -lean_ctor_set(x_500, 2, x_497); -x_501 = lean_array_push(x_496, x_500); -x_502 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__2; +x_497 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__13; +lean_inc(x_493); +x_498 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_498, 0, x_493); +lean_ctor_set(x_498, 1, x_497); +x_499 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__15; +x_500 = lean_array_push(x_499, x_498); +x_501 = lean_box(2); +x_502 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__14; x_503 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_503, 0, x_498); +lean_ctor_set(x_503, 0, x_501); lean_ctor_set(x_503, 1, x_502); -lean_ctor_set(x_503, 2, x_501); -x_504 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__20; -x_505 = lean_array_push(x_504, x_503); -x_506 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__12; -x_507 = lean_array_push(x_505, x_506); -x_508 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__10; -x_509 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_509, 0, x_498); -lean_ctor_set(x_509, 1, x_508); -lean_ctor_set(x_509, 2, x_507); -x_510 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__21; -lean_inc(x_490); -x_511 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_511, 0, x_490); -lean_ctor_set(x_511, 1, x_510); -x_512 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__28; -lean_inc(x_492); +lean_ctor_set(x_503, 2, x_500); +x_504 = lean_array_push(x_499, x_503); +x_505 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__2; +x_506 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_506, 0, x_501); +lean_ctor_set(x_506, 1, x_505); +lean_ctor_set(x_506, 2, x_504); +x_507 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__20; +x_508 = lean_array_push(x_507, x_506); +x_509 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__12; +x_510 = lean_array_push(x_508, x_509); +x_511 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__10; +x_512 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_512, 0, x_501); +lean_ctor_set(x_512, 1, x_511); +lean_ctor_set(x_512, 2, x_510); +x_513 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__21; +lean_inc(x_493); +x_514 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_514, 0, x_493); +lean_ctor_set(x_514, 1, x_513); +x_515 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__28; +lean_inc(x_495); +lean_inc(x_496); +x_516 = l_Lean_addMacroScope(x_496, x_515, x_495); +x_517 = lean_box(0); +x_518 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__27; +lean_inc(x_493); +x_519 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_519, 0, x_493); +lean_ctor_set(x_519, 1, x_518); +lean_ctor_set(x_519, 2, x_516); +lean_ctor_set(x_519, 3, x_517); +x_520 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__29; +lean_inc(x_519); +x_521 = lean_array_push(x_520, x_519); +x_522 = lean_array_push(x_521, x_509); +x_523 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__24; +x_524 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_524, 0, x_501); +lean_ctor_set(x_524, 1, x_523); +lean_ctor_set(x_524, 2, x_522); +x_525 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__36; +lean_inc(x_493); +x_526 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_526, 0, x_493); +lean_ctor_set(x_526, 1, x_525); +x_527 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__40; +lean_inc(x_495); +lean_inc(x_496); +x_528 = l_Lean_addMacroScope(x_496, x_527, x_495); +x_529 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__39; +lean_inc(x_493); +x_530 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_530, 0, x_493); +lean_ctor_set(x_530, 1, x_529); +lean_ctor_set(x_530, 2, x_528); +lean_ctor_set(x_530, 3, x_517); +lean_inc(x_530); +x_531 = lean_array_push(x_499, x_530); +x_532 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_532, 0, x_501); +lean_ctor_set(x_532, 1, x_505); +lean_ctor_set(x_532, 2, x_531); +x_533 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__41; lean_inc(x_493); -x_513 = l_Lean_addMacroScope(x_493, x_512, x_492); -x_514 = lean_box(0); -x_515 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__27; -lean_inc(x_490); -x_516 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_516, 0, x_490); -lean_ctor_set(x_516, 1, x_515); -lean_ctor_set(x_516, 2, x_513); -lean_ctor_set(x_516, 3, x_514); -x_517 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__29; -lean_inc(x_516); -x_518 = lean_array_push(x_517, x_516); -x_519 = lean_array_push(x_518, x_506); -x_520 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__24; -x_521 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_521, 0, x_498); -lean_ctor_set(x_521, 1, x_520); -lean_ctor_set(x_521, 2, x_519); -x_522 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__36; -lean_inc(x_490); -x_523 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_523, 0, x_490); -lean_ctor_set(x_523, 1, x_522); -x_524 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__40; -lean_inc(x_492); +x_534 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_534, 0, x_493); +lean_ctor_set(x_534, 1, x_533); +x_535 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__45; +lean_inc(x_495); +lean_inc(x_496); +x_536 = l_Lean_addMacroScope(x_496, x_535, x_495); +x_537 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__44; +x_538 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__48; lean_inc(x_493); -x_525 = l_Lean_addMacroScope(x_493, x_524, x_492); -x_526 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__39; -lean_inc(x_490); -x_527 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_527, 0, x_490); -lean_ctor_set(x_527, 1, x_526); -lean_ctor_set(x_527, 2, x_525); -lean_ctor_set(x_527, 3, x_514); -lean_inc(x_527); -x_528 = lean_array_push(x_496, x_527); -x_529 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_529, 0, x_498); -lean_ctor_set(x_529, 1, x_502); -lean_ctor_set(x_529, 2, x_528); -x_530 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__41; -lean_inc(x_490); -x_531 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_531, 0, x_490); -lean_ctor_set(x_531, 1, x_530); -x_532 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__45; -lean_inc(x_492); +x_539 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_539, 0, x_493); +lean_ctor_set(x_539, 1, x_537); +lean_ctor_set(x_539, 2, x_536); +lean_ctor_set(x_539, 3, x_538); +x_540 = lean_array_push(x_520, x_534); +lean_inc(x_540); +x_541 = lean_array_push(x_540, x_539); +x_542 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_542, 0, x_501); +lean_ctor_set(x_542, 1, x_505); +lean_ctor_set(x_542, 2, x_541); +x_543 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__49; lean_inc(x_493); -x_533 = l_Lean_addMacroScope(x_493, x_532, x_492); -x_534 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__44; -x_535 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__48; -lean_inc(x_490); -x_536 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_536, 0, x_490); -lean_ctor_set(x_536, 1, x_534); -lean_ctor_set(x_536, 2, x_533); -lean_ctor_set(x_536, 3, x_535); -x_537 = lean_array_push(x_517, x_531); -lean_inc(x_537); -x_538 = lean_array_push(x_537, x_536); -x_539 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_539, 0, x_498); -lean_ctor_set(x_539, 1, x_502); -lean_ctor_set(x_539, 2, x_538); -x_540 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__49; -lean_inc(x_490); -x_541 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_541, 0, x_490); -lean_ctor_set(x_541, 1, x_540); -x_542 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__50; -lean_inc(x_523); -x_543 = lean_array_push(x_542, x_523); -lean_inc(x_543); -x_544 = lean_array_push(x_543, x_529); -x_545 = lean_array_push(x_544, x_539); -x_546 = lean_array_push(x_545, x_506); -lean_inc(x_541); -x_547 = lean_array_push(x_546, x_541); -x_548 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__35; -x_549 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_549, 0, x_498); -lean_ctor_set(x_549, 1, x_548); -lean_ctor_set(x_549, 2, x_547); -x_550 = lean_array_push(x_496, x_549); -x_551 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_551, 0, x_498); -lean_ctor_set(x_551, 1, x_502); -lean_ctor_set(x_551, 2, x_550); -x_552 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__58; -lean_inc(x_492); +x_544 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_544, 0, x_493); +lean_ctor_set(x_544, 1, x_543); +x_545 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__50; +lean_inc(x_526); +x_546 = lean_array_push(x_545, x_526); +lean_inc(x_546); +x_547 = lean_array_push(x_546, x_532); +x_548 = lean_array_push(x_547, x_542); +x_549 = lean_array_push(x_548, x_509); +lean_inc(x_544); +x_550 = lean_array_push(x_549, x_544); +x_551 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__35; +x_552 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_552, 0, x_501); +lean_ctor_set(x_552, 1, x_551); +lean_ctor_set(x_552, 2, x_550); +x_553 = lean_array_push(x_499, x_552); +x_554 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_554, 0, x_501); +lean_ctor_set(x_554, 1, x_505); +lean_ctor_set(x_554, 2, x_553); +x_555 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__58; +lean_inc(x_495); +lean_inc(x_496); +x_556 = l_Lean_addMacroScope(x_496, x_555, x_495); +x_557 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__57; +x_558 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__62; lean_inc(x_493); -x_553 = l_Lean_addMacroScope(x_493, x_552, x_492); -x_554 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__57; -x_555 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__62; -lean_inc(x_490); -x_556 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_556, 0, x_490); -lean_ctor_set(x_556, 1, x_554); -lean_ctor_set(x_556, 2, x_553); -lean_ctor_set(x_556, 3, x_555); +x_559 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_559, 0, x_493); +lean_ctor_set(x_559, 1, x_557); +lean_ctor_set(x_559, 2, x_556); +lean_ctor_set(x_559, 3, x_558); lean_inc(x_11); -x_557 = lean_array_push(x_496, x_11); -x_558 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_558, 0, x_498); -lean_ctor_set(x_558, 1, x_502); -lean_ctor_set(x_558, 2, x_557); -x_559 = lean_array_push(x_517, x_556); -x_560 = lean_array_push(x_559, x_558); -x_561 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__54; -x_562 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_562, 0, x_498); -lean_ctor_set(x_562, 1, x_561); -lean_ctor_set(x_562, 2, x_560); -lean_inc(x_537); -x_563 = lean_array_push(x_537, x_562); -x_564 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__52; +x_560 = lean_array_push(x_499, x_11); +x_561 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_561, 0, x_501); +lean_ctor_set(x_561, 1, x_505); +lean_ctor_set(x_561, 2, x_560); +x_562 = lean_array_push(x_520, x_559); +x_563 = lean_array_push(x_562, x_561); +x_564 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__54; x_565 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_565, 0, x_498); +lean_ctor_set(x_565, 0, x_501); lean_ctor_set(x_565, 1, x_564); lean_ctor_set(x_565, 2, x_563); -lean_inc(x_565); -x_566 = lean_array_push(x_496, x_565); -x_567 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_567, 0, x_498); -lean_ctor_set(x_567, 1, x_502); -lean_ctor_set(x_567, 2, x_566); -x_568 = lean_array_push(x_517, x_551); -lean_inc(x_567); +lean_inc(x_540); +x_566 = lean_array_push(x_540, x_565); +x_567 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__52; +x_568 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_568, 0, x_501); +lean_ctor_set(x_568, 1, x_567); +lean_ctor_set(x_568, 2, x_566); lean_inc(x_568); -x_569 = lean_array_push(x_568, x_567); -x_570 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__31; -x_571 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_571, 0, x_498); -lean_ctor_set(x_571, 1, x_570); -lean_ctor_set(x_571, 2, x_569); -x_572 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__65; -lean_inc(x_490); -x_573 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_573, 0, x_490); -lean_ctor_set(x_573, 1, x_572); -x_574 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__71; -lean_inc(x_492); +x_569 = lean_array_push(x_499, x_568); +x_570 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_570, 0, x_501); +lean_ctor_set(x_570, 1, x_505); +lean_ctor_set(x_570, 2, x_569); +x_571 = lean_array_push(x_520, x_554); +lean_inc(x_570); +lean_inc(x_571); +x_572 = lean_array_push(x_571, x_570); +x_573 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__31; +x_574 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_574, 0, x_501); +lean_ctor_set(x_574, 1, x_573); +lean_ctor_set(x_574, 2, x_572); +x_575 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__65; lean_inc(x_493); -x_575 = l_Lean_addMacroScope(x_493, x_574, x_492); -x_576 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__68; -x_577 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__74; -lean_inc(x_490); -x_578 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_578, 0, x_490); -lean_ctor_set(x_578, 1, x_576); -lean_ctor_set(x_578, 2, x_575); -lean_ctor_set(x_578, 3, x_577); -x_579 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__77; -lean_inc(x_490); -x_580 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_580, 0, x_490); -lean_ctor_set(x_580, 1, x_579); -x_581 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__78; -lean_inc(x_580); -x_582 = lean_array_push(x_581, x_580); -x_583 = lean_array_push(x_582, x_580); +x_576 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_576, 0, x_493); +lean_ctor_set(x_576, 1, x_575); +x_577 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__71; +lean_inc(x_495); +lean_inc(x_496); +x_578 = l_Lean_addMacroScope(x_496, x_577, x_495); +x_579 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__68; +x_580 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__74; +lean_inc(x_493); +x_581 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_581, 0, x_493); +lean_ctor_set(x_581, 1, x_579); +lean_ctor_set(x_581, 2, x_578); +lean_ctor_set(x_581, 3, x_580); +x_582 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__77; +lean_inc(x_493); +x_583 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_583, 0, x_493); +lean_ctor_set(x_583, 1, x_582); +x_584 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__78; +lean_inc(x_583); +x_585 = lean_array_push(x_584, x_583); +x_586 = lean_array_push(x_585, x_583); lean_inc(x_11); -x_584 = lean_array_push(x_583, x_11); -x_585 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__76; -x_586 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_586, 0, x_498); -lean_ctor_set(x_586, 1, x_585); -lean_ctor_set(x_586, 2, x_584); +x_587 = lean_array_push(x_586, x_11); +x_588 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__76; +x_589 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_589, 0, x_501); +lean_ctor_set(x_589, 1, x_588); +lean_ctor_set(x_589, 2, x_587); lean_inc(x_11); -x_587 = lean_array_push(x_581, x_11); -lean_inc(x_586); -x_588 = lean_array_push(x_587, x_586); -x_589 = lean_array_push(x_588, x_527); -x_590 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_590, 0, x_498); -lean_ctor_set(x_590, 1, x_502); -lean_ctor_set(x_590, 2, x_589); -x_591 = lean_array_push(x_517, x_578); -x_592 = lean_array_push(x_591, x_590); +x_590 = lean_array_push(x_584, x_11); +lean_inc(x_589); +x_591 = lean_array_push(x_590, x_589); +x_592 = lean_array_push(x_591, x_530); x_593 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_593, 0, x_498); -lean_ctor_set(x_593, 1, x_561); +lean_ctor_set(x_593, 0, x_501); +lean_ctor_set(x_593, 1, x_505); lean_ctor_set(x_593, 2, x_592); -x_594 = lean_array_push(x_581, x_573); -lean_inc(x_594); +x_594 = lean_array_push(x_520, x_581); x_595 = lean_array_push(x_594, x_593); -x_596 = lean_array_push(x_595, x_506); -x_597 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__64; -x_598 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_598, 0, x_498); -lean_ctor_set(x_598, 1, x_597); -lean_ctor_set(x_598, 2, x_596); -x_599 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__79; -x_600 = lean_array_push(x_599, x_511); -lean_inc(x_600); -x_601 = lean_array_push(x_600, x_521); -x_602 = lean_array_push(x_601, x_571); -x_603 = lean_array_push(x_602, x_598); -x_604 = lean_array_push(x_603, x_506); -x_605 = lean_array_push(x_604, x_506); -x_606 = lean_array_push(x_605, x_506); -x_607 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__22; -x_608 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_608, 0, x_498); -lean_ctor_set(x_608, 1, x_607); -lean_ctor_set(x_608, 2, x_606); -x_609 = lean_array_push(x_517, x_509); -x_610 = lean_array_push(x_609, x_608); -x_611 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__8; -x_612 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_612, 0, x_498); -lean_ctor_set(x_612, 1, x_611); -lean_ctor_set(x_612, 2, x_610); -x_613 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__82; -lean_inc(x_490); -x_614 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_614, 0, x_490); -lean_ctor_set(x_614, 1, x_613); -x_615 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__96; -lean_inc(x_492); +x_596 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_596, 0, x_501); +lean_ctor_set(x_596, 1, x_564); +lean_ctor_set(x_596, 2, x_595); +x_597 = lean_array_push(x_584, x_576); +lean_inc(x_597); +x_598 = lean_array_push(x_597, x_596); +x_599 = lean_array_push(x_598, x_509); +x_600 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__64; +x_601 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_601, 0, x_501); +lean_ctor_set(x_601, 1, x_600); +lean_ctor_set(x_601, 2, x_599); +x_602 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__79; +x_603 = lean_array_push(x_602, x_514); +lean_inc(x_603); +x_604 = lean_array_push(x_603, x_524); +x_605 = lean_array_push(x_604, x_574); +x_606 = lean_array_push(x_605, x_601); +x_607 = lean_array_push(x_606, x_509); +x_608 = lean_array_push(x_607, x_509); +x_609 = lean_array_push(x_608, x_509); +x_610 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__22; +x_611 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_611, 0, x_501); +lean_ctor_set(x_611, 1, x_610); +lean_ctor_set(x_611, 2, x_609); +x_612 = lean_array_push(x_520, x_512); +x_613 = lean_array_push(x_612, x_611); +x_614 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__8; +x_615 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_615, 0, x_501); +lean_ctor_set(x_615, 1, x_614); +lean_ctor_set(x_615, 2, x_613); +x_616 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__82; +lean_inc(x_493); +x_617 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_617, 0, x_493); +lean_ctor_set(x_617, 1, x_616); +x_618 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__96; +lean_inc(x_495); +lean_inc(x_496); +x_619 = l_Lean_addMacroScope(x_496, x_618, x_495); +x_620 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__95; +lean_inc(x_493); +x_621 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_621, 0, x_493); +lean_ctor_set(x_621, 1, x_620); +lean_ctor_set(x_621, 2, x_619); +lean_ctor_set(x_621, 3, x_517); +x_622 = lean_array_push(x_499, x_519); +x_623 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_623, 0, x_501); +lean_ctor_set(x_623, 1, x_505); +lean_ctor_set(x_623, 2, x_622); +x_624 = lean_array_push(x_520, x_621); +x_625 = lean_array_push(x_624, x_623); +x_626 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__92; +x_627 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_627, 0, x_501); +lean_ctor_set(x_627, 1, x_626); +lean_ctor_set(x_627, 2, x_625); +x_628 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__97; +x_629 = lean_array_push(x_628, x_627); +x_630 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__84; +x_631 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_631, 0, x_501); +lean_ctor_set(x_631, 1, x_630); +lean_ctor_set(x_631, 2, x_629); +x_632 = lean_array_push(x_499, x_631); +x_633 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_633, 0, x_501); +lean_ctor_set(x_633, 1, x_505); +lean_ctor_set(x_633, 2, x_632); +x_634 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__98; lean_inc(x_493); -x_616 = l_Lean_addMacroScope(x_493, x_615, x_492); -x_617 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__95; -lean_inc(x_490); -x_618 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_618, 0, x_490); -lean_ctor_set(x_618, 1, x_617); -lean_ctor_set(x_618, 2, x_616); -lean_ctor_set(x_618, 3, x_514); -x_619 = lean_array_push(x_496, x_516); -x_620 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_620, 0, x_498); -lean_ctor_set(x_620, 1, x_502); -lean_ctor_set(x_620, 2, x_619); -x_621 = lean_array_push(x_517, x_618); -x_622 = lean_array_push(x_621, x_620); -x_623 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__92; -x_624 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_624, 0, x_498); -lean_ctor_set(x_624, 1, x_623); -lean_ctor_set(x_624, 2, x_622); -x_625 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__97; -x_626 = lean_array_push(x_625, x_624); -x_627 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__84; -x_628 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_628, 0, x_498); -lean_ctor_set(x_628, 1, x_627); -lean_ctor_set(x_628, 2, x_626); -x_629 = lean_array_push(x_496, x_628); -x_630 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_630, 0, x_498); -lean_ctor_set(x_630, 1, x_502); -lean_ctor_set(x_630, 2, x_629); -x_631 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__98; -lean_inc(x_490); -x_632 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_632, 0, x_490); -lean_ctor_set(x_632, 1, x_631); -x_633 = lean_array_push(x_581, x_614); -x_634 = lean_array_push(x_633, x_630); -lean_inc(x_632); -x_635 = lean_array_push(x_634, x_632); -x_636 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__81; -x_637 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_637, 0, x_498); -lean_ctor_set(x_637, 1, x_636); -lean_ctor_set(x_637, 2, x_635); -x_638 = lean_array_push(x_496, x_637); -x_639 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_639, 0, x_498); -lean_ctor_set(x_639, 1, x_502); -lean_ctor_set(x_639, 2, x_638); -x_640 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__17; -x_641 = lean_array_push(x_640, x_639); -x_642 = lean_array_push(x_641, x_506); -x_643 = lean_array_push(x_642, x_506); -x_644 = lean_array_push(x_643, x_506); -x_645 = lean_array_push(x_644, x_506); -x_646 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_646, 0, x_498); -lean_ctor_set(x_646, 1, x_508); -lean_ctor_set(x_646, 2, x_645); -x_647 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__99; -lean_inc(x_490); -x_648 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_648, 0, x_490); -lean_ctor_set(x_648, 1, x_647); -x_649 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__104; -lean_inc(x_492); +x_635 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_635, 0, x_493); +lean_ctor_set(x_635, 1, x_634); +x_636 = lean_array_push(x_584, x_617); +x_637 = lean_array_push(x_636, x_633); +lean_inc(x_635); +x_638 = lean_array_push(x_637, x_635); +x_639 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__81; +x_640 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_640, 0, x_501); +lean_ctor_set(x_640, 1, x_639); +lean_ctor_set(x_640, 2, x_638); +x_641 = lean_array_push(x_499, x_640); +x_642 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_642, 0, x_501); +lean_ctor_set(x_642, 1, x_505); +lean_ctor_set(x_642, 2, x_641); +x_643 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__17; +x_644 = lean_array_push(x_643, x_642); +x_645 = lean_array_push(x_644, x_509); +x_646 = lean_array_push(x_645, x_509); +x_647 = lean_array_push(x_646, x_509); +x_648 = lean_array_push(x_647, x_509); +x_649 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_649, 0, x_501); +lean_ctor_set(x_649, 1, x_511); +lean_ctor_set(x_649, 2, x_648); +x_650 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__99; lean_inc(x_493); -x_650 = l_Lean_addMacroScope(x_493, x_649, x_492); -x_651 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__103; -lean_inc(x_490); -x_652 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_652, 0, x_490); -lean_ctor_set(x_652, 1, x_651); -lean_ctor_set(x_652, 2, x_650); -lean_ctor_set(x_652, 3, x_514); -x_653 = lean_array_push(x_517, x_652); -lean_inc(x_653); -x_654 = lean_array_push(x_653, x_506); -x_655 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_655, 0, x_498); -lean_ctor_set(x_655, 1, x_520); -lean_ctor_set(x_655, 2, x_654); -x_656 = lean_array_push(x_568, x_565); -x_657 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__106; +x_651 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_651, 0, x_493); +lean_ctor_set(x_651, 1, x_650); +x_652 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__104; +lean_inc(x_495); +lean_inc(x_496); +x_653 = l_Lean_addMacroScope(x_496, x_652, x_495); +x_654 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__103; +lean_inc(x_493); +x_655 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_655, 0, x_493); +lean_ctor_set(x_655, 1, x_654); +lean_ctor_set(x_655, 2, x_653); +lean_ctor_set(x_655, 3, x_517); +x_656 = lean_array_push(x_520, x_655); +lean_inc(x_656); +x_657 = lean_array_push(x_656, x_509); x_658 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_658, 0, x_498); -lean_ctor_set(x_658, 1, x_657); -lean_ctor_set(x_658, 2, x_656); -x_659 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__107; -x_660 = lean_array_push(x_659, x_648); -x_661 = lean_array_push(x_660, x_655); -x_662 = lean_array_push(x_661, x_658); -x_663 = lean_array_push(x_662, x_506); -x_664 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__100; -x_665 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_665, 0, x_498); -lean_ctor_set(x_665, 1, x_664); -lean_ctor_set(x_665, 2, x_663); -x_666 = lean_array_push(x_517, x_646); -x_667 = lean_array_push(x_666, x_665); +lean_ctor_set(x_658, 0, x_501); +lean_ctor_set(x_658, 1, x_523); +lean_ctor_set(x_658, 2, x_657); +x_659 = lean_array_push(x_571, x_568); +x_660 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__106; +x_661 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_661, 0, x_501); +lean_ctor_set(x_661, 1, x_660); +lean_ctor_set(x_661, 2, x_659); +x_662 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__107; +x_663 = lean_array_push(x_662, x_651); +x_664 = lean_array_push(x_663, x_658); +x_665 = lean_array_push(x_664, x_661); +x_666 = lean_array_push(x_665, x_509); +x_667 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__100; x_668 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_668, 0, x_498); -lean_ctor_set(x_668, 1, x_611); -lean_ctor_set(x_668, 2, x_667); -x_669 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__114; -lean_inc(x_492); -lean_inc(x_493); -x_670 = l_Lean_addMacroScope(x_493, x_669, x_492); -x_671 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__113; -lean_inc(x_490); -x_672 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_672, 0, x_490); -lean_ctor_set(x_672, 1, x_671); -lean_ctor_set(x_672, 2, x_670); -lean_ctor_set(x_672, 3, x_514); -lean_inc(x_672); -x_673 = lean_array_push(x_496, x_672); +lean_ctor_set(x_668, 0, x_501); +lean_ctor_set(x_668, 1, x_667); +lean_ctor_set(x_668, 2, x_666); +x_669 = lean_array_push(x_520, x_649); +x_670 = lean_array_push(x_669, x_668); +x_671 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_671, 0, x_501); +lean_ctor_set(x_671, 1, x_614); +lean_ctor_set(x_671, 2, x_670); +x_672 = lean_array_push(x_520, x_9); +x_673 = lean_array_push(x_672, x_509); x_674 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_674, 0, x_498); -lean_ctor_set(x_674, 1, x_502); +lean_ctor_set(x_674, 0, x_501); +lean_ctor_set(x_674, 1, x_523); lean_ctor_set(x_674, 2, x_673); -x_675 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__118; -lean_inc(x_492); +x_675 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__114; +lean_inc(x_495); +lean_inc(x_496); +x_676 = l_Lean_addMacroScope(x_496, x_675, x_495); +x_677 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__113; +lean_inc(x_493); +x_678 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_678, 0, x_493); +lean_ctor_set(x_678, 1, x_677); +lean_ctor_set(x_678, 2, x_676); +lean_ctor_set(x_678, 3, x_517); +lean_inc(x_678); +x_679 = lean_array_push(x_499, x_678); +x_680 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_680, 0, x_501); +lean_ctor_set(x_680, 1, x_505); +lean_ctor_set(x_680, 2, x_679); +x_681 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__118; +lean_inc(x_495); +lean_inc(x_496); +x_682 = l_Lean_addMacroScope(x_496, x_681, x_495); +x_683 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__117; +x_684 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__121; lean_inc(x_493); -x_676 = l_Lean_addMacroScope(x_493, x_675, x_492); -x_677 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__117; -x_678 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__121; -lean_inc(x_490); -x_679 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_679, 0, x_490); -lean_ctor_set(x_679, 1, x_677); -lean_ctor_set(x_679, 2, x_676); -lean_ctor_set(x_679, 3, x_678); -lean_inc(x_537); -x_680 = lean_array_push(x_537, x_679); -x_681 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_681, 0, x_498); -lean_ctor_set(x_681, 1, x_502); -lean_ctor_set(x_681, 2, x_680); -x_682 = lean_array_push(x_543, x_674); -x_683 = lean_array_push(x_682, x_681); -x_684 = lean_array_push(x_683, x_506); -lean_inc(x_541); -x_685 = lean_array_push(x_684, x_541); -x_686 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_686, 0, x_498); -lean_ctor_set(x_686, 1, x_548); -lean_ctor_set(x_686, 2, x_685); -x_687 = lean_array_push(x_496, x_686); -x_688 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_688, 0, x_498); -lean_ctor_set(x_688, 1, x_502); -lean_ctor_set(x_688, 2, x_687); -x_689 = lean_array_push(x_517, x_688); -x_690 = lean_array_push(x_689, x_567); -x_691 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_691, 0, x_498); -lean_ctor_set(x_691, 1, x_570); -lean_ctor_set(x_691, 2, x_690); -x_692 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__122; -lean_inc(x_490); -x_693 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_693, 0, x_490); -lean_ctor_set(x_693, 1, x_692); -x_694 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__130; -lean_inc(x_490); -x_695 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_695, 0, x_490); -lean_ctor_set(x_695, 1, x_694); -x_696 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__137; -lean_inc(x_492); +x_685 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_685, 0, x_493); +lean_ctor_set(x_685, 1, x_683); +lean_ctor_set(x_685, 2, x_682); +lean_ctor_set(x_685, 3, x_684); +lean_inc(x_540); +x_686 = lean_array_push(x_540, x_685); +x_687 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_687, 0, x_501); +lean_ctor_set(x_687, 1, x_505); +lean_ctor_set(x_687, 2, x_686); +x_688 = lean_array_push(x_546, x_680); +x_689 = lean_array_push(x_688, x_687); +x_690 = lean_array_push(x_689, x_509); +lean_inc(x_544); +x_691 = lean_array_push(x_690, x_544); +x_692 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_692, 0, x_501); +lean_ctor_set(x_692, 1, x_551); +lean_ctor_set(x_692, 2, x_691); +x_693 = lean_array_push(x_499, x_692); +x_694 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_694, 0, x_501); +lean_ctor_set(x_694, 1, x_505); +lean_ctor_set(x_694, 2, x_693); +x_695 = lean_array_push(x_520, x_694); +x_696 = lean_array_push(x_695, x_570); +x_697 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_697, 0, x_501); +lean_ctor_set(x_697, 1, x_573); +lean_ctor_set(x_697, 2, x_696); +x_698 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__122; lean_inc(x_493); -x_697 = l_Lean_addMacroScope(x_493, x_696, x_492); -x_698 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__135; -lean_inc(x_490); -x_699 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_699, 0, x_490); +x_699 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_699, 0, x_493); lean_ctor_set(x_699, 1, x_698); -lean_ctor_set(x_699, 2, x_697); -lean_ctor_set(x_699, 3, x_514); -x_700 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__138; -x_701 = lean_array_push(x_700, x_699); -x_702 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__132; -x_703 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_703, 0, x_498); -lean_ctor_set(x_703, 1, x_702); -lean_ctor_set(x_703, 2, x_701); -x_704 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__139; -lean_inc(x_490); -x_705 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_705, 0, x_490); +x_700 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__130; +lean_inc(x_493); +x_701 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_701, 0, x_493); +lean_ctor_set(x_701, 1, x_700); +x_702 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__137; +lean_inc(x_495); +lean_inc(x_496); +x_703 = l_Lean_addMacroScope(x_496, x_702, x_495); +x_704 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__135; +lean_inc(x_493); +x_705 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_705, 0, x_493); lean_ctor_set(x_705, 1, x_704); -x_706 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__142; -lean_inc(x_490); -x_707 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_707, 0, x_490); -lean_ctor_set(x_707, 1, x_706); -x_708 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__145; -lean_inc(x_490); -x_709 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_709, 0, x_490); +lean_ctor_set(x_705, 2, x_703); +lean_ctor_set(x_705, 3, x_517); +x_706 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__138; +x_707 = lean_array_push(x_706, x_705); +x_708 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__132; +x_709 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_709, 0, x_501); lean_ctor_set(x_709, 1, x_708); -x_710 = lean_array_push(x_537, x_11); -x_711 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_711, 0, x_498); -lean_ctor_set(x_711, 1, x_502); -lean_ctor_set(x_711, 2, x_710); -x_712 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__149; -lean_inc(x_490); +lean_ctor_set(x_709, 2, x_707); +x_710 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__139; +lean_inc(x_493); +x_711 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_711, 0, x_493); +lean_ctor_set(x_711, 1, x_710); +x_712 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__142; +lean_inc(x_493); x_713 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_713, 0, x_490); +lean_ctor_set(x_713, 0, x_493); lean_ctor_set(x_713, 1, x_712); -x_714 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__16; -lean_inc(x_709); -x_715 = lean_array_push(x_714, x_709); -x_716 = lean_array_push(x_715, x_506); -x_717 = lean_array_push(x_716, x_506); -x_718 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__148; -x_719 = lean_array_push(x_717, x_718); +x_714 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__145; +lean_inc(x_493); +x_715 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_715, 0, x_493); +lean_ctor_set(x_715, 1, x_714); +x_716 = lean_array_push(x_540, x_11); +x_717 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_717, 0, x_501); +lean_ctor_set(x_717, 1, x_505); +lean_ctor_set(x_717, 2, x_716); +x_718 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__149; +lean_inc(x_493); +x_719 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_719, 0, x_493); +lean_ctor_set(x_719, 1, x_718); +x_720 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__16; +lean_inc(x_715); +x_721 = lean_array_push(x_720, x_715); +x_722 = lean_array_push(x_721, x_509); +x_723 = lean_array_push(x_722, x_509); +x_724 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__148; +x_725 = lean_array_push(x_723, x_724); +lean_inc(x_725); +x_726 = lean_array_push(x_725, x_717); lean_inc(x_719); -x_720 = lean_array_push(x_719, x_711); -lean_inc(x_713); -x_721 = lean_array_push(x_720, x_713); -x_722 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__144; -x_723 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_723, 0, x_498); -lean_ctor_set(x_723, 1, x_722); -lean_ctor_set(x_723, 2, x_721); -x_724 = lean_array_push(x_496, x_723); -x_725 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_725, 0, x_498); -lean_ctor_set(x_725, 1, x_502); -lean_ctor_set(x_725, 2, x_724); -x_726 = lean_array_push(x_517, x_707); -x_727 = lean_array_push(x_726, x_725); -x_728 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__141; +x_727 = lean_array_push(x_726, x_719); +x_728 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__144; x_729 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_729, 0, x_498); +lean_ctor_set(x_729, 0, x_501); lean_ctor_set(x_729, 1, x_728); lean_ctor_set(x_729, 2, x_727); -x_730 = lean_array_push(x_517, x_729); -x_731 = lean_array_push(x_730, x_506); -x_732 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__127; -x_733 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_733, 0, x_498); -lean_ctor_set(x_733, 1, x_732); -lean_ctor_set(x_733, 2, x_731); -x_734 = lean_array_push(x_496, x_733); +x_730 = lean_array_push(x_499, x_729); +x_731 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_731, 0, x_501); +lean_ctor_set(x_731, 1, x_505); +lean_ctor_set(x_731, 2, x_730); +x_732 = lean_array_push(x_520, x_713); +x_733 = lean_array_push(x_732, x_731); +x_734 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__141; x_735 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_735, 0, x_498); -lean_ctor_set(x_735, 1, x_502); -lean_ctor_set(x_735, 2, x_734); -x_736 = lean_array_push(x_496, x_735); -x_737 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__125; -x_738 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_738, 0, x_498); -lean_ctor_set(x_738, 1, x_737); -lean_ctor_set(x_738, 2, x_736); -x_739 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__150; -lean_inc(x_490); -x_740 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_740, 0, x_490); -lean_ctor_set(x_740, 1, x_739); -x_741 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__153; -lean_inc(x_490); -x_742 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_742, 0, x_490); -lean_ctor_set(x_742, 1, x_741); -x_743 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__159; -lean_inc(x_492); +lean_ctor_set(x_735, 0, x_501); +lean_ctor_set(x_735, 1, x_734); +lean_ctor_set(x_735, 2, x_733); +x_736 = lean_array_push(x_520, x_735); +x_737 = lean_array_push(x_736, x_509); +x_738 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__127; +x_739 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_739, 0, x_501); +lean_ctor_set(x_739, 1, x_738); +lean_ctor_set(x_739, 2, x_737); +x_740 = lean_array_push(x_499, x_739); +x_741 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_741, 0, x_501); +lean_ctor_set(x_741, 1, x_505); +lean_ctor_set(x_741, 2, x_740); +x_742 = lean_array_push(x_499, x_741); +x_743 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__125; +x_744 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_744, 0, x_501); +lean_ctor_set(x_744, 1, x_743); +lean_ctor_set(x_744, 2, x_742); +x_745 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__150; lean_inc(x_493); -x_744 = l_Lean_addMacroScope(x_493, x_743, x_492); -x_745 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__158; -lean_inc(x_490); -x_746 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_746, 0, x_490); +x_746 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_746, 0, x_493); lean_ctor_set(x_746, 1, x_745); -lean_ctor_set(x_746, 2, x_744); -lean_ctor_set(x_746, 3, x_514); -x_747 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__160; -lean_inc(x_490); +x_747 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__153; +lean_inc(x_493); x_748 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_748, 0, x_490); +lean_ctor_set(x_748, 0, x_493); lean_ctor_set(x_748, 1, x_747); -x_749 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__168; -lean_inc(x_492); +x_749 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__159; +lean_inc(x_495); +lean_inc(x_496); +x_750 = l_Lean_addMacroScope(x_496, x_749, x_495); +x_751 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__158; lean_inc(x_493); -x_750 = l_Lean_addMacroScope(x_493, x_749, x_492); -x_751 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__167; -x_752 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__171; -lean_inc(x_490); -x_753 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_753, 0, x_490); -lean_ctor_set(x_753, 1, x_751); -lean_ctor_set(x_753, 2, x_750); -lean_ctor_set(x_753, 3, x_752); -x_754 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__172; -lean_inc(x_490); -x_755 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_755, 0, x_490); -lean_ctor_set(x_755, 1, x_754); -x_756 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__176; -lean_inc(x_492); +x_752 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_752, 0, x_493); +lean_ctor_set(x_752, 1, x_751); +lean_ctor_set(x_752, 2, x_750); +lean_ctor_set(x_752, 3, x_517); +x_753 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__160; lean_inc(x_493); -x_757 = l_Lean_addMacroScope(x_493, x_756, x_492); -x_758 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__175; -x_759 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__181; -lean_inc(x_490); -x_760 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_760, 0, x_490); -lean_ctor_set(x_760, 1, x_758); -lean_ctor_set(x_760, 2, x_757); -lean_ctor_set(x_760, 3, x_759); -x_761 = lean_array_push(x_517, x_709); -lean_inc(x_713); -x_762 = lean_array_push(x_761, x_713); -x_763 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__185; -x_764 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_764, 0, x_498); -lean_ctor_set(x_764, 1, x_763); -lean_ctor_set(x_764, 2, x_762); -x_765 = lean_array_push(x_719, x_506); -x_766 = lean_array_push(x_765, x_713); -x_767 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_767, 0, x_498); -lean_ctor_set(x_767, 1, x_722); -lean_ctor_set(x_767, 2, x_766); -x_768 = lean_array_push(x_517, x_764); -x_769 = lean_array_push(x_768, x_767); -x_770 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__183; -x_771 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_771, 0, x_498); -lean_ctor_set(x_771, 1, x_770); -lean_ctor_set(x_771, 2, x_769); -lean_inc(x_771); -x_772 = lean_array_push(x_517, x_771); -x_773 = lean_array_push(x_772, x_771); -x_774 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_774, 0, x_498); -lean_ctor_set(x_774, 1, x_502); -lean_ctor_set(x_774, 2, x_773); -x_775 = lean_array_push(x_517, x_760); -x_776 = lean_array_push(x_775, x_774); +x_754 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_754, 0, x_493); +lean_ctor_set(x_754, 1, x_753); +x_755 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__168; +lean_inc(x_495); +lean_inc(x_496); +x_756 = l_Lean_addMacroScope(x_496, x_755, x_495); +x_757 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__167; +x_758 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__171; +lean_inc(x_493); +x_759 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_759, 0, x_493); +lean_ctor_set(x_759, 1, x_757); +lean_ctor_set(x_759, 2, x_756); +lean_ctor_set(x_759, 3, x_758); +x_760 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__172; +lean_inc(x_493); +x_761 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_761, 0, x_493); +lean_ctor_set(x_761, 1, x_760); +x_762 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__176; +lean_inc(x_495); +lean_inc(x_496); +x_763 = l_Lean_addMacroScope(x_496, x_762, x_495); +x_764 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__175; +x_765 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__181; +lean_inc(x_493); +x_766 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_766, 0, x_493); +lean_ctor_set(x_766, 1, x_764); +lean_ctor_set(x_766, 2, x_763); +lean_ctor_set(x_766, 3, x_765); +x_767 = lean_array_push(x_520, x_715); +lean_inc(x_719); +x_768 = lean_array_push(x_767, x_719); +x_769 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__185; +x_770 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_770, 0, x_501); +lean_ctor_set(x_770, 1, x_769); +lean_ctor_set(x_770, 2, x_768); +x_771 = lean_array_push(x_725, x_509); +x_772 = lean_array_push(x_771, x_719); +x_773 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_773, 0, x_501); +lean_ctor_set(x_773, 1, x_728); +lean_ctor_set(x_773, 2, x_772); +x_774 = lean_array_push(x_520, x_770); +x_775 = lean_array_push(x_774, x_773); +x_776 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__183; x_777 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_777, 0, x_498); -lean_ctor_set(x_777, 1, x_561); -lean_ctor_set(x_777, 2, x_776); -x_778 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__189; -lean_inc(x_492); +lean_ctor_set(x_777, 0, x_501); +lean_ctor_set(x_777, 1, x_776); +lean_ctor_set(x_777, 2, x_775); +lean_inc(x_777); +x_778 = lean_array_push(x_520, x_777); +x_779 = lean_array_push(x_778, x_777); +x_780 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_780, 0, x_501); +lean_ctor_set(x_780, 1, x_505); +lean_ctor_set(x_780, 2, x_779); +x_781 = lean_array_push(x_520, x_766); +x_782 = lean_array_push(x_781, x_780); +x_783 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_783, 0, x_501); +lean_ctor_set(x_783, 1, x_564); +lean_ctor_set(x_783, 2, x_782); +x_784 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__189; +lean_inc(x_495); +lean_inc(x_496); +x_785 = l_Lean_addMacroScope(x_496, x_784, x_495); +x_786 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__188; +x_787 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__192; lean_inc(x_493); -x_779 = l_Lean_addMacroScope(x_493, x_778, x_492); -x_780 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__188; -x_781 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__192; -lean_inc(x_490); -x_782 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_782, 0, x_490); -lean_ctor_set(x_782, 1, x_780); -lean_ctor_set(x_782, 2, x_779); -lean_ctor_set(x_782, 3, x_781); -x_783 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__197; -lean_inc(x_492); +x_788 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_788, 0, x_493); +lean_ctor_set(x_788, 1, x_786); +lean_ctor_set(x_788, 2, x_785); +lean_ctor_set(x_788, 3, x_787); +x_789 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__197; +lean_inc(x_495); +lean_inc(x_496); +x_790 = l_Lean_addMacroScope(x_496, x_789, x_495); +x_791 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__195; +x_792 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__200; lean_inc(x_493); -x_784 = l_Lean_addMacroScope(x_493, x_783, x_492); -x_785 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__195; -x_786 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__200; -lean_inc(x_490); -x_787 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_787, 0, x_490); -lean_ctor_set(x_787, 1, x_785); -lean_ctor_set(x_787, 2, x_784); -lean_ctor_set(x_787, 3, x_786); -x_788 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__205; -lean_inc(x_492); +x_793 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_793, 0, x_493); +lean_ctor_set(x_793, 1, x_791); +lean_ctor_set(x_793, 2, x_790); +lean_ctor_set(x_793, 3, x_792); +x_794 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__205; +lean_inc(x_495); +lean_inc(x_496); +x_795 = l_Lean_addMacroScope(x_496, x_794, x_495); +x_796 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__203; +x_797 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__208; lean_inc(x_493); -x_789 = l_Lean_addMacroScope(x_493, x_788, x_492); -x_790 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__203; -x_791 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__208; -lean_inc(x_490); -x_792 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_792, 0, x_490); -lean_ctor_set(x_792, 1, x_790); -lean_ctor_set(x_792, 2, x_789); -lean_ctor_set(x_792, 3, x_791); -x_793 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__211; -lean_inc(x_490); -x_794 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_794, 0, x_490); -lean_ctor_set(x_794, 1, x_793); -x_795 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__214; -lean_inc(x_490); -x_796 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_796, 0, x_490); -lean_ctor_set(x_796, 1, x_795); -x_797 = lean_array_push(x_496, x_796); -x_798 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__213; -x_799 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_799, 0, x_498); -lean_ctor_set(x_799, 1, x_798); -lean_ctor_set(x_799, 2, x_797); -x_800 = lean_array_push(x_659, x_672); -lean_inc(x_794); -x_801 = lean_array_push(x_800, x_794); -x_802 = lean_array_push(x_801, x_799); -lean_inc(x_632); -x_803 = lean_array_push(x_802, x_632); -x_804 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__210; +x_798 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_798, 0, x_493); +lean_ctor_set(x_798, 1, x_796); +lean_ctor_set(x_798, 2, x_795); +lean_ctor_set(x_798, 3, x_797); +x_799 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__211; +lean_inc(x_493); +x_800 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_800, 0, x_493); +lean_ctor_set(x_800, 1, x_799); +x_801 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__214; +lean_inc(x_493); +x_802 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_802, 0, x_493); +lean_ctor_set(x_802, 1, x_801); +x_803 = lean_array_push(x_499, x_802); +x_804 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__213; x_805 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_805, 0, x_498); +lean_ctor_set(x_805, 0, x_501); lean_ctor_set(x_805, 1, x_804); lean_ctor_set(x_805, 2, x_803); -x_806 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__215; -lean_inc(x_490); -x_807 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_807, 0, x_490); -lean_ctor_set(x_807, 1, x_806); -x_808 = lean_array_push(x_496, x_807); -x_809 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_809, 0, x_498); -lean_ctor_set(x_809, 1, x_798); -lean_ctor_set(x_809, 2, x_808); -x_810 = lean_array_push(x_659, x_805); -x_811 = lean_array_push(x_810, x_794); -x_812 = lean_array_push(x_811, x_809); -x_813 = lean_array_push(x_812, x_632); -x_814 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_814, 0, x_498); -lean_ctor_set(x_814, 1, x_804); -lean_ctor_set(x_814, 2, x_813); -x_815 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__222; -lean_inc(x_492); +x_806 = lean_array_push(x_662, x_678); +lean_inc(x_800); +x_807 = lean_array_push(x_806, x_800); +x_808 = lean_array_push(x_807, x_805); +lean_inc(x_635); +x_809 = lean_array_push(x_808, x_635); +x_810 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__210; +x_811 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_811, 0, x_501); +lean_ctor_set(x_811, 1, x_810); +lean_ctor_set(x_811, 2, x_809); +x_812 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__215; +lean_inc(x_493); +x_813 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_813, 0, x_493); +lean_ctor_set(x_813, 1, x_812); +x_814 = lean_array_push(x_499, x_813); +x_815 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_815, 0, x_501); +lean_ctor_set(x_815, 1, x_804); +lean_ctor_set(x_815, 2, x_814); +x_816 = lean_array_push(x_662, x_811); +x_817 = lean_array_push(x_816, x_800); +x_818 = lean_array_push(x_817, x_815); +x_819 = lean_array_push(x_818, x_635); +x_820 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_820, 0, x_501); +lean_ctor_set(x_820, 1, x_810); +lean_ctor_set(x_820, 2, x_819); +x_821 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__222; +lean_inc(x_495); +lean_inc(x_496); +x_822 = l_Lean_addMacroScope(x_496, x_821, x_495); +x_823 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__220; +x_824 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__224; lean_inc(x_493); -x_816 = l_Lean_addMacroScope(x_493, x_815, x_492); -x_817 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__220; -x_818 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__224; -lean_inc(x_490); -x_819 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_819, 0, x_490); -lean_ctor_set(x_819, 1, x_817); -lean_ctor_set(x_819, 2, x_816); -lean_ctor_set(x_819, 3, x_818); -x_820 = lean_array_push(x_496, x_586); -x_821 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_821, 0, x_498); -lean_ctor_set(x_821, 1, x_502); -lean_ctor_set(x_821, 2, x_820); -x_822 = lean_array_push(x_517, x_819); -x_823 = lean_array_push(x_822, x_821); -x_824 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_824, 0, x_498); -lean_ctor_set(x_824, 1, x_561); -lean_ctor_set(x_824, 2, x_823); -x_825 = lean_array_push(x_517, x_824); -x_826 = lean_array_push(x_825, x_506); +x_825 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_825, 0, x_493); +lean_ctor_set(x_825, 1, x_823); +lean_ctor_set(x_825, 2, x_822); +lean_ctor_set(x_825, 3, x_824); +x_826 = lean_array_push(x_499, x_589); x_827 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_827, 0, x_498); -lean_ctor_set(x_827, 1, x_502); +lean_ctor_set(x_827, 0, x_501); +lean_ctor_set(x_827, 1, x_505); lean_ctor_set(x_827, 2, x_826); -x_828 = lean_array_push(x_581, x_523); +x_828 = lean_array_push(x_520, x_825); x_829 = lean_array_push(x_828, x_827); -x_830 = lean_array_push(x_829, x_541); -x_831 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__217; -x_832 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_832, 0, x_498); -lean_ctor_set(x_832, 1, x_831); -lean_ctor_set(x_832, 2, x_830); -x_833 = lean_array_push(x_517, x_814); -x_834 = lean_array_push(x_833, x_832); -x_835 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_835, 0, x_498); -lean_ctor_set(x_835, 1, x_502); -lean_ctor_set(x_835, 2, x_834); -x_836 = lean_array_push(x_517, x_792); -x_837 = lean_array_push(x_836, x_835); +x_830 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_830, 0, x_501); +lean_ctor_set(x_830, 1, x_564); +lean_ctor_set(x_830, 2, x_829); +x_831 = lean_array_push(x_520, x_830); +x_832 = lean_array_push(x_831, x_509); +x_833 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_833, 0, x_501); +lean_ctor_set(x_833, 1, x_505); +lean_ctor_set(x_833, 2, x_832); +x_834 = lean_array_push(x_584, x_526); +x_835 = lean_array_push(x_834, x_833); +x_836 = lean_array_push(x_835, x_544); +x_837 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__217; x_838 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_838, 0, x_498); -lean_ctor_set(x_838, 1, x_561); -lean_ctor_set(x_838, 2, x_837); -x_839 = lean_array_push(x_496, x_838); -x_840 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__162; +lean_ctor_set(x_838, 0, x_501); +lean_ctor_set(x_838, 1, x_837); +lean_ctor_set(x_838, 2, x_836); +x_839 = lean_array_push(x_520, x_820); +x_840 = lean_array_push(x_839, x_838); x_841 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_841, 0, x_498); -lean_ctor_set(x_841, 1, x_840); -lean_ctor_set(x_841, 2, x_839); -lean_inc(x_746); -x_842 = lean_array_push(x_659, x_746); -x_843 = lean_array_push(x_842, x_506); -x_844 = lean_array_push(x_843, x_748); -lean_inc(x_844); -x_845 = lean_array_push(x_844, x_841); -x_846 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__155; +lean_ctor_set(x_841, 0, x_501); +lean_ctor_set(x_841, 1, x_505); +lean_ctor_set(x_841, 2, x_840); +x_842 = lean_array_push(x_520, x_798); +x_843 = lean_array_push(x_842, x_841); +x_844 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_844, 0, x_501); +lean_ctor_set(x_844, 1, x_564); +lean_ctor_set(x_844, 2, x_843); +x_845 = lean_array_push(x_499, x_844); +x_846 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__162; x_847 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_847, 0, x_498); +lean_ctor_set(x_847, 0, x_501); lean_ctor_set(x_847, 1, x_846); lean_ctor_set(x_847, 2, x_845); -x_848 = lean_array_push(x_581, x_742); -x_849 = lean_array_push(x_848, x_506); -lean_inc(x_849); -x_850 = lean_array_push(x_849, x_847); -x_851 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__152; -x_852 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_852, 0, x_498); -lean_ctor_set(x_852, 1, x_851); -lean_ctor_set(x_852, 2, x_850); -x_853 = lean_array_push(x_517, x_852); -x_854 = lean_array_push(x_853, x_506); -x_855 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_855, 0, x_498); -lean_ctor_set(x_855, 1, x_732); -lean_ctor_set(x_855, 2, x_854); -x_856 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__229; -lean_inc(x_492); +lean_inc(x_752); +x_848 = lean_array_push(x_662, x_752); +x_849 = lean_array_push(x_848, x_509); +x_850 = lean_array_push(x_849, x_754); +lean_inc(x_850); +x_851 = lean_array_push(x_850, x_847); +x_852 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__155; +x_853 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_853, 0, x_501); +lean_ctor_set(x_853, 1, x_852); +lean_ctor_set(x_853, 2, x_851); +x_854 = lean_array_push(x_584, x_748); +x_855 = lean_array_push(x_854, x_509); +lean_inc(x_855); +x_856 = lean_array_push(x_855, x_853); +x_857 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__152; +x_858 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_858, 0, x_501); +lean_ctor_set(x_858, 1, x_857); +lean_ctor_set(x_858, 2, x_856); +x_859 = lean_array_push(x_520, x_858); +x_860 = lean_array_push(x_859, x_509); +x_861 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_861, 0, x_501); +lean_ctor_set(x_861, 1, x_738); +lean_ctor_set(x_861, 2, x_860); +x_862 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__229; +lean_inc(x_495); +lean_inc(x_496); +x_863 = l_Lean_addMacroScope(x_496, x_862, x_495); +x_864 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__227; +x_865 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__232; lean_inc(x_493); -x_857 = l_Lean_addMacroScope(x_493, x_856, x_492); -x_858 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__227; -x_859 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__232; -lean_inc(x_490); -x_860 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_860, 0, x_490); -lean_ctor_set(x_860, 1, x_858); -lean_ctor_set(x_860, 2, x_857); -lean_ctor_set(x_860, 3, x_859); -x_861 = lean_array_push(x_496, x_860); -x_862 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_862, 0, x_498); -lean_ctor_set(x_862, 1, x_840); -lean_ctor_set(x_862, 2, x_861); -x_863 = lean_array_push(x_517, x_862); -x_864 = lean_array_push(x_863, x_506); -x_865 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_865, 0, x_498); -lean_ctor_set(x_865, 1, x_732); -lean_ctor_set(x_865, 2, x_864); -x_866 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__236; -x_867 = l_Lean_addMacroScope(x_493, x_866, x_492); -x_868 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__235; -x_869 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__239; -x_870 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_870, 0, x_490); -lean_ctor_set(x_870, 1, x_868); -lean_ctor_set(x_870, 2, x_867); -lean_ctor_set(x_870, 3, x_869); -x_871 = lean_array_push(x_496, x_746); -x_872 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_872, 0, x_498); -lean_ctor_set(x_872, 1, x_502); -lean_ctor_set(x_872, 2, x_871); -x_873 = lean_array_push(x_517, x_870); -lean_inc(x_872); -x_874 = lean_array_push(x_873, x_872); -x_875 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_875, 0, x_498); -lean_ctor_set(x_875, 1, x_561); -lean_ctor_set(x_875, 2, x_874); -x_876 = lean_array_push(x_496, x_875); -x_877 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_877, 0, x_498); -lean_ctor_set(x_877, 1, x_840); -lean_ctor_set(x_877, 2, x_876); -x_878 = lean_array_push(x_517, x_877); -x_879 = lean_array_push(x_878, x_506); -x_880 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_880, 0, x_498); -lean_ctor_set(x_880, 1, x_732); -lean_ctor_set(x_880, 2, x_879); -x_881 = lean_array_push(x_581, x_855); -x_882 = lean_array_push(x_881, x_865); -x_883 = lean_array_push(x_882, x_880); -x_884 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_884, 0, x_498); -lean_ctor_set(x_884, 1, x_502); -lean_ctor_set(x_884, 2, x_883); -x_885 = lean_array_push(x_496, x_884); +x_866 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_866, 0, x_493); +lean_ctor_set(x_866, 1, x_864); +lean_ctor_set(x_866, 2, x_863); +lean_ctor_set(x_866, 3, x_865); +x_867 = lean_array_push(x_499, x_866); +x_868 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_868, 0, x_501); +lean_ctor_set(x_868, 1, x_846); +lean_ctor_set(x_868, 2, x_867); +x_869 = lean_array_push(x_520, x_868); +x_870 = lean_array_push(x_869, x_509); +x_871 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_871, 0, x_501); +lean_ctor_set(x_871, 1, x_738); +lean_ctor_set(x_871, 2, x_870); +x_872 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__236; +x_873 = l_Lean_addMacroScope(x_496, x_872, x_495); +x_874 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__235; +x_875 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__239; +x_876 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_876, 0, x_493); +lean_ctor_set(x_876, 1, x_874); +lean_ctor_set(x_876, 2, x_873); +lean_ctor_set(x_876, 3, x_875); +x_877 = lean_array_push(x_499, x_752); +x_878 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_878, 0, x_501); +lean_ctor_set(x_878, 1, x_505); +lean_ctor_set(x_878, 2, x_877); +x_879 = lean_array_push(x_520, x_876); +lean_inc(x_878); +x_880 = lean_array_push(x_879, x_878); +x_881 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_881, 0, x_501); +lean_ctor_set(x_881, 1, x_564); +lean_ctor_set(x_881, 2, x_880); +x_882 = lean_array_push(x_499, x_881); +x_883 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_883, 0, x_501); +lean_ctor_set(x_883, 1, x_846); +lean_ctor_set(x_883, 2, x_882); +x_884 = lean_array_push(x_520, x_883); +x_885 = lean_array_push(x_884, x_509); x_886 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_886, 0, x_498); -lean_ctor_set(x_886, 1, x_737); +lean_ctor_set(x_886, 0, x_501); +lean_ctor_set(x_886, 1, x_738); lean_ctor_set(x_886, 2, x_885); -x_887 = lean_array_push(x_517, x_693); -lean_inc(x_887); -x_888 = lean_array_push(x_887, x_886); -x_889 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__123; +x_887 = lean_array_push(x_584, x_861); +x_888 = lean_array_push(x_887, x_871); +x_889 = lean_array_push(x_888, x_886); x_890 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_890, 0, x_498); -lean_ctor_set(x_890, 1, x_889); -lean_ctor_set(x_890, 2, x_888); -x_891 = lean_array_push(x_496, x_890); +lean_ctor_set(x_890, 0, x_501); +lean_ctor_set(x_890, 1, x_505); +lean_ctor_set(x_890, 2, x_889); +x_891 = lean_array_push(x_499, x_890); x_892 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_892, 0, x_498); -lean_ctor_set(x_892, 1, x_502); +lean_ctor_set(x_892, 0, x_501); +lean_ctor_set(x_892, 1, x_743); lean_ctor_set(x_892, 2, x_891); -x_893 = lean_array_push(x_517, x_787); +x_893 = lean_array_push(x_520, x_699); +lean_inc(x_893); x_894 = lean_array_push(x_893, x_892); -x_895 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_895, 0, x_498); -lean_ctor_set(x_895, 1, x_561); -lean_ctor_set(x_895, 2, x_894); -x_896 = lean_array_push(x_581, x_782); -lean_inc(x_755); -x_897 = lean_array_push(x_896, x_755); -x_898 = lean_array_push(x_897, x_895); -x_899 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__164; -x_900 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_900, 0, x_498); -lean_ctor_set(x_900, 1, x_899); -lean_ctor_set(x_900, 2, x_898); -x_901 = lean_array_push(x_581, x_777); -lean_inc(x_755); -x_902 = lean_array_push(x_901, x_755); -x_903 = lean_array_push(x_902, x_900); -x_904 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_904, 0, x_498); -lean_ctor_set(x_904, 1, x_899); -lean_ctor_set(x_904, 2, x_903); -x_905 = lean_array_push(x_581, x_753); -x_906 = lean_array_push(x_905, x_755); -x_907 = lean_array_push(x_906, x_904); -x_908 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_908, 0, x_498); -lean_ctor_set(x_908, 1, x_899); -lean_ctor_set(x_908, 2, x_907); -x_909 = lean_array_push(x_496, x_908); +x_895 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__123; +x_896 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_896, 0, x_501); +lean_ctor_set(x_896, 1, x_895); +lean_ctor_set(x_896, 2, x_894); +x_897 = lean_array_push(x_499, x_896); +x_898 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_898, 0, x_501); +lean_ctor_set(x_898, 1, x_505); +lean_ctor_set(x_898, 2, x_897); +x_899 = lean_array_push(x_520, x_793); +x_900 = lean_array_push(x_899, x_898); +x_901 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_901, 0, x_501); +lean_ctor_set(x_901, 1, x_564); +lean_ctor_set(x_901, 2, x_900); +x_902 = lean_array_push(x_584, x_788); +lean_inc(x_761); +x_903 = lean_array_push(x_902, x_761); +x_904 = lean_array_push(x_903, x_901); +x_905 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__164; +x_906 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_906, 0, x_501); +lean_ctor_set(x_906, 1, x_905); +lean_ctor_set(x_906, 2, x_904); +x_907 = lean_array_push(x_584, x_783); +lean_inc(x_761); +x_908 = lean_array_push(x_907, x_761); +x_909 = lean_array_push(x_908, x_906); x_910 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_910, 0, x_498); -lean_ctor_set(x_910, 1, x_840); +lean_ctor_set(x_910, 0, x_501); +lean_ctor_set(x_910, 1, x_905); lean_ctor_set(x_910, 2, x_909); -x_911 = lean_array_push(x_844, x_910); -x_912 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_912, 0, x_498); -lean_ctor_set(x_912, 1, x_846); -lean_ctor_set(x_912, 2, x_911); -x_913 = lean_array_push(x_849, x_912); +x_911 = lean_array_push(x_584, x_759); +x_912 = lean_array_push(x_911, x_761); +x_913 = lean_array_push(x_912, x_910); x_914 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_914, 0, x_498); -lean_ctor_set(x_914, 1, x_851); +lean_ctor_set(x_914, 0, x_501); +lean_ctor_set(x_914, 1, x_905); lean_ctor_set(x_914, 2, x_913); -x_915 = lean_array_push(x_517, x_914); -x_916 = lean_array_push(x_915, x_506); -x_917 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_917, 0, x_498); -lean_ctor_set(x_917, 1, x_732); -lean_ctor_set(x_917, 2, x_916); -x_918 = lean_array_push(x_653, x_872); -x_919 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_919, 0, x_498); -lean_ctor_set(x_919, 1, x_561); -lean_ctor_set(x_919, 2, x_918); -x_920 = lean_array_push(x_496, x_919); -x_921 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_921, 0, x_498); -lean_ctor_set(x_921, 1, x_840); -lean_ctor_set(x_921, 2, x_920); -x_922 = lean_array_push(x_517, x_921); -x_923 = lean_array_push(x_922, x_506); -x_924 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_924, 0, x_498); -lean_ctor_set(x_924, 1, x_732); -lean_ctor_set(x_924, 2, x_923); -x_925 = lean_array_push(x_517, x_917); -x_926 = lean_array_push(x_925, x_924); +x_915 = lean_array_push(x_499, x_914); +x_916 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_916, 0, x_501); +lean_ctor_set(x_916, 1, x_846); +lean_ctor_set(x_916, 2, x_915); +x_917 = lean_array_push(x_850, x_916); +x_918 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_918, 0, x_501); +lean_ctor_set(x_918, 1, x_852); +lean_ctor_set(x_918, 2, x_917); +x_919 = lean_array_push(x_855, x_918); +x_920 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_920, 0, x_501); +lean_ctor_set(x_920, 1, x_857); +lean_ctor_set(x_920, 2, x_919); +x_921 = lean_array_push(x_520, x_920); +x_922 = lean_array_push(x_921, x_509); +x_923 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_923, 0, x_501); +lean_ctor_set(x_923, 1, x_738); +lean_ctor_set(x_923, 2, x_922); +x_924 = lean_array_push(x_656, x_878); +x_925 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_925, 0, x_501); +lean_ctor_set(x_925, 1, x_564); +lean_ctor_set(x_925, 2, x_924); +x_926 = lean_array_push(x_499, x_925); x_927 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_927, 0, x_498); -lean_ctor_set(x_927, 1, x_502); +lean_ctor_set(x_927, 0, x_501); +lean_ctor_set(x_927, 1, x_846); lean_ctor_set(x_927, 2, x_926); -x_928 = lean_array_push(x_496, x_927); -x_929 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_929, 0, x_498); -lean_ctor_set(x_929, 1, x_737); -lean_ctor_set(x_929, 2, x_928); -x_930 = lean_array_push(x_517, x_740); -x_931 = lean_array_push(x_930, x_929); -x_932 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_932, 0, x_498); -lean_ctor_set(x_932, 1, x_502); -lean_ctor_set(x_932, 2, x_931); -x_933 = lean_array_push(x_714, x_695); -x_934 = lean_array_push(x_933, x_703); -x_935 = lean_array_push(x_934, x_705); -x_936 = lean_array_push(x_935, x_738); -x_937 = lean_array_push(x_936, x_506); -x_938 = lean_array_push(x_937, x_932); -x_939 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__129; -x_940 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_940, 0, x_498); -lean_ctor_set(x_940, 1, x_939); -lean_ctor_set(x_940, 2, x_938); -x_941 = lean_array_push(x_517, x_940); -x_942 = lean_array_push(x_941, x_506); -x_943 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_943, 0, x_498); -lean_ctor_set(x_943, 1, x_732); -lean_ctor_set(x_943, 2, x_942); -x_944 = lean_array_push(x_496, x_943); -x_945 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_945, 0, x_498); -lean_ctor_set(x_945, 1, x_502); -lean_ctor_set(x_945, 2, x_944); -x_946 = lean_array_push(x_496, x_945); -x_947 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_947, 0, x_498); -lean_ctor_set(x_947, 1, x_737); -lean_ctor_set(x_947, 2, x_946); -x_948 = lean_array_push(x_887, x_947); +x_928 = lean_array_push(x_520, x_927); +x_929 = lean_array_push(x_928, x_509); +x_930 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_930, 0, x_501); +lean_ctor_set(x_930, 1, x_738); +lean_ctor_set(x_930, 2, x_929); +x_931 = lean_array_push(x_520, x_923); +x_932 = lean_array_push(x_931, x_930); +x_933 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_933, 0, x_501); +lean_ctor_set(x_933, 1, x_505); +lean_ctor_set(x_933, 2, x_932); +x_934 = lean_array_push(x_499, x_933); +x_935 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_935, 0, x_501); +lean_ctor_set(x_935, 1, x_743); +lean_ctor_set(x_935, 2, x_934); +x_936 = lean_array_push(x_520, x_746); +x_937 = lean_array_push(x_936, x_935); +x_938 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_938, 0, x_501); +lean_ctor_set(x_938, 1, x_505); +lean_ctor_set(x_938, 2, x_937); +x_939 = lean_array_push(x_720, x_701); +x_940 = lean_array_push(x_939, x_709); +x_941 = lean_array_push(x_940, x_711); +x_942 = lean_array_push(x_941, x_744); +x_943 = lean_array_push(x_942, x_509); +x_944 = lean_array_push(x_943, x_938); +x_945 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__129; +x_946 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_946, 0, x_501); +lean_ctor_set(x_946, 1, x_945); +lean_ctor_set(x_946, 2, x_944); +x_947 = lean_array_push(x_520, x_946); +x_948 = lean_array_push(x_947, x_509); x_949 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_949, 0, x_498); -lean_ctor_set(x_949, 1, x_889); +lean_ctor_set(x_949, 0, x_501); +lean_ctor_set(x_949, 1, x_738); lean_ctor_set(x_949, 2, x_948); -x_950 = lean_array_push(x_594, x_949); -x_951 = lean_array_push(x_950, x_506); -x_952 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_952, 0, x_498); -lean_ctor_set(x_952, 1, x_597); -lean_ctor_set(x_952, 2, x_951); -x_953 = lean_array_push(x_600, x_9); -x_954 = lean_array_push(x_953, x_691); -x_955 = lean_array_push(x_954, x_952); -x_956 = lean_array_push(x_955, x_506); -x_957 = lean_array_push(x_956, x_506); -x_958 = lean_array_push(x_957, x_506); -x_959 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_959, 0, x_498); -lean_ctor_set(x_959, 1, x_607); -lean_ctor_set(x_959, 2, x_958); -x_960 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__240; -x_961 = lean_array_push(x_960, x_959); -x_962 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_962, 0, x_498); -lean_ctor_set(x_962, 1, x_611); -lean_ctor_set(x_962, 2, x_961); -x_963 = lean_array_push(x_581, x_612); -x_964 = lean_array_push(x_963, x_668); -x_965 = lean_array_push(x_964, x_962); -x_966 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_966, 0, x_498); -lean_ctor_set(x_966, 1, x_502); -lean_ctor_set(x_966, 2, x_965); -x_967 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_967, 0, x_966); -lean_ctor_set(x_967, 1, x_491); -return x_967; +x_950 = lean_array_push(x_499, x_949); +x_951 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_951, 0, x_501); +lean_ctor_set(x_951, 1, x_505); +lean_ctor_set(x_951, 2, x_950); +x_952 = lean_array_push(x_499, x_951); +x_953 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_953, 0, x_501); +lean_ctor_set(x_953, 1, x_743); +lean_ctor_set(x_953, 2, x_952); +x_954 = lean_array_push(x_893, x_953); +x_955 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_955, 0, x_501); +lean_ctor_set(x_955, 1, x_895); +lean_ctor_set(x_955, 2, x_954); +x_956 = lean_array_push(x_597, x_955); +x_957 = lean_array_push(x_956, x_509); +x_958 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_958, 0, x_501); +lean_ctor_set(x_958, 1, x_600); +lean_ctor_set(x_958, 2, x_957); +x_959 = lean_array_push(x_603, x_674); +x_960 = lean_array_push(x_959, x_697); +x_961 = lean_array_push(x_960, x_958); +x_962 = lean_array_push(x_961, x_509); +x_963 = lean_array_push(x_962, x_509); +x_964 = lean_array_push(x_963, x_509); +x_965 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_965, 0, x_501); +lean_ctor_set(x_965, 1, x_610); +lean_ctor_set(x_965, 2, x_964); +x_966 = l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__240; +x_967 = lean_array_push(x_966, x_965); +x_968 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_968, 0, x_501); +lean_ctor_set(x_968, 1, x_614); +lean_ctor_set(x_968, 2, x_967); +x_969 = lean_array_push(x_584, x_615); +x_970 = lean_array_push(x_969, x_671); +x_971 = lean_array_push(x_970, x_968); +x_972 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_972, 0, x_501); +lean_ctor_set(x_972, 1, x_505); +lean_ctor_set(x_972, 2, x_971); +x_973 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_973, 0, x_972); +lean_ctor_set(x_973, 1, x_494); +return x_973; } } } @@ -4877,524 +4889,524 @@ lean_dec_ref(res); res = initialize_Lean_Elab_SyntheticMVars(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__1 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__1); -l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__2 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__2(); -lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__2); -l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__3 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__3(); -lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__3); -l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__4 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__4(); -lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__4); -l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__5 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__5(); -lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__5); -l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__6 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__6(); -lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__6); -l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__7 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__7(); -lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__7); -l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__8 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__8(); -lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__8); -l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__9 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__9(); -lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__9); -l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__10 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__10(); -lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__10); -l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__11 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__11(); -lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__11); -l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__12 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__12(); -lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__12); -l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__13 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__13(); -lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__13); -l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__14 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__14(); -lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__14); -l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__15 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__15(); -lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__15); -l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__16 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__16(); -lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__16); -l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__17 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__17(); -lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__17); -l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__18 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__18(); -lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab___________closed__18); -l_Lean_Elab_Tactic_commandDeclare__config__elab________ = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab________(); -lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab________); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__1 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__1); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__2 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__2(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__2); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__3 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__3(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__3); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__4 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__4(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__4); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__5 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__5(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__5); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__6 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__6(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__6); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__7 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__7(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__7); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__8 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__8(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__8); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__9 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__9(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__9); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__10 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__10(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__10); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__11 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__11(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__11); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__12 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__12(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__12); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__13 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__13(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__13); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__14 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__14(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__14); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__15 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__15(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__15); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__16 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__16(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__16); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__17 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__17(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__17); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__18 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__18(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__18); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__19 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__19(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__19); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__20 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__20(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__20); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__21 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__21(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__21); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__22 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__22(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__22); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__23 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__23(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__23); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__24 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__24(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__24); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__25 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__25(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__25); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__26 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__26(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__26); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__27 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__27(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__27); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__28 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__28(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__28); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__29 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__29(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__29); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__30 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__30(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__30); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__31 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__31(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__31); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__32 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__32(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__32); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__33 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__33(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__33); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__34 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__34(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__34); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__35 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__35(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__35); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__36 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__36(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__36); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__37 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__37(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__37); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__38 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__38(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__38); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__39 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__39(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__39); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__40 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__40(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__40); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__41 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__41(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__41); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__42 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__42(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__42); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__43 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__43(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__43); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__44 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__44(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__44); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__45 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__45(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__45); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__46 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__46(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__46); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__47 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__47(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__47); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__48 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__48(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__48); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__49 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__49(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__49); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__50 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__50(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__50); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__51 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__51(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__51); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__52 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__52(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__52); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__53 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__53(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__53); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__54 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__54(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__54); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__55 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__55(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__55); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__56 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__56(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__56); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__57 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__57(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__57); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__58 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__58(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__58); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__59 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__59(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__59); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__60 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__60(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__60); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__61 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__61(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__61); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__62 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__62(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__62); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__63 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__63(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__63); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__64 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__64(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__64); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__65 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__65(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__65); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__66 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__66(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__66); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__67 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__67(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__67); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__68 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__68(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__68); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__69 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__69(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__69); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__70 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__70(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__70); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__71 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__71(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__71); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__72 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__72(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__72); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__73 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__73(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__73); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__74 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__74(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__74); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__75 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__75(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__75); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__76 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__76(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__76); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__77 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__77(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__77); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__78 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__78(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__78); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__79 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__79(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__79); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__80 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__80(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__80); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__81 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__81(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__81); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__82 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__82(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__82); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__83 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__83(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__83); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__84 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__84(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__84); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__85 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__85(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__85); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__86 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__86(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__86); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__87 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__87(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__87); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__88 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__88(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__88); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__89 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__89(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__89); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__90 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__90(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__90); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__91 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__91(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__91); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__92 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__92(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__92); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__93 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__93(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__93); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__94 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__94(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__94); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__95 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__95(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__95); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__96 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__96(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__96); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__97 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__97(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__97); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__98 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__98(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__98); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__99 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__99(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__99); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__100 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__100(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__100); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__101 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__101(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__101); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__102 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__102(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__102); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__103 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__103(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__103); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__104 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__104(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__104); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__105 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__105(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__105); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__106 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__106(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__106); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__107 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__107(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__107); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__108 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__108(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__108); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__109 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__109(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__109); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__110 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__110(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__110); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__111 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__111(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__111); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__112 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__112(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__112); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__113 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__113(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__113); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__114 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__114(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__114); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__115 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__115(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__115); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__116 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__116(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__116); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__117 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__117(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__117); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__118 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__118(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__118); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__119 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__119(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__119); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__120 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__120(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__120); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__121 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__121(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__121); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__122 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__122(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__122); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__123 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__123(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__123); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__124 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__124(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__124); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__125 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__125(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__125); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__126 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__126(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__126); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__127 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__127(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__127); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__128 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__128(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__128); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__129 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__129(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__129); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__130 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__130(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__130); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__131 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__131(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__131); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__132 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__132(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__132); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__133 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__133(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__133); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__134 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__134(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__134); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__135 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__135(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__135); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__136 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__136(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__136); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__137 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__137(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__137); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__138 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__138(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__138); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__139 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__139(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__139); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__140 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__140(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__140); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__141 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__141(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__141); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__142 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__142(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__142); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__143 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__143(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__143); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__144 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__144(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__144); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__145 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__145(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__145); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__146 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__146(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__146); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__147 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__147(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__147); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__148 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__148(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__148); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__149 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__149(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__149); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__150 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__150(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__150); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__151 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__151(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__151); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__152 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__152(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__152); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__153 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__153(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__153); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__154 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__154(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__154); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__155 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__155(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__155); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__156 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__156(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__156); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__157 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__157(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__157); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__158 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__158(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__158); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__159 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__159(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__159); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__160 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__160(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__160); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__161 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__161(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__161); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__162 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__162(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__162); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__163 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__163(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__163); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__164 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__164(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__164); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__165 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__165(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__165); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__166 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__166(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__166); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__167 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__167(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__167); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__168 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__168(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__168); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__169 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__169(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__169); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__170 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__170(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__170); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__171 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__171(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__171); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__172 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__172(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__172); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__173 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__173(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__173); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__174 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__174(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__174); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__175 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__175(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__175); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__176 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__176(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__176); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__177 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__177(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__177); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__178 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__178(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__178); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__179 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__179(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__179); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__180 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__180(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__180); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__181 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__181(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__181); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__182 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__182(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__182); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__183 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__183(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__183); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__184 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__184(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__184); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__185 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__185(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__185); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__186 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__186(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__186); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__187 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__187(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__187); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__188 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__188(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__188); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__189 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__189(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__189); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__190 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__190(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__190); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__191 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__191(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__191); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__192 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__192(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__192); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__193 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__193(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__193); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__194 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__194(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__194); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__195 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__195(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__195); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__196 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__196(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__196); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__197 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__197(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__197); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__198 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__198(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__198); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__199 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__199(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__199); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__200 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__200(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__200); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__201 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__201(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__201); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__202 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__202(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__202); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__203 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__203(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__203); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__204 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__204(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__204); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__205 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__205(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__205); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__206 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__206(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__206); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__207 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__207(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__207); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__208 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__208(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__208); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__209 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__209(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__209); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__210 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__210(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__210); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__211 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__211(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__211); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__212 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__212(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__212); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__213 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__213(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__213); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__214 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__214(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__214); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__215 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__215(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__215); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__216 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__216(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__216); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__217 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__217(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__217); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__218 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__218(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__218); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__219 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__219(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__219); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__220 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__220(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__220); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__221 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__221(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__221); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__222 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__222(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__222); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__223 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__223(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__223); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__224 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__224(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__224); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__225 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__225(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__225); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__226 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__226(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__226); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__227 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__227(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__227); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__228 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__228(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__228); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__229 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__229(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__229); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__230 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__230(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__230); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__231 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__231(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__231); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__232 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__232(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__232); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__233 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__233(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__233); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__234 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__234(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__234); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__235 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__235(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__235); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__236 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__236(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__236); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__237 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__237(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__237); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__238 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__238(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__238); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__239 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__239(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__239); -l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__240 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__240(); -lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab__________1___closed__240); +l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__1 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__1); +l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__2 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__2); +l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__3 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__3); +l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__4 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__4(); +lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__4); +l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__5 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__5(); +lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__5); +l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__6 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__6(); +lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__6); +l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__7 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__7(); +lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__7); +l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__8 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__8(); +lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__8); +l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__9 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__9(); +lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__9); +l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__10 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__10(); +lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__10); +l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__11 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__11(); +lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__11); +l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__12 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__12(); +lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__12); +l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__13 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__13(); +lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__13); +l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__14 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__14(); +lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__14); +l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__15 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__15(); +lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__15); +l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__16 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__16(); +lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__16); +l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__17 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__17(); +lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__17); +l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__18 = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__18(); +lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab_______closed__18); +l_Lean_Elab_Tactic_commandDeclare__config__elab____ = _init_l_Lean_Elab_Tactic_commandDeclare__config__elab____(); +lean_mark_persistent(l_Lean_Elab_Tactic_commandDeclare__config__elab____); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__1 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__1); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__2 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__2); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__3 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__3); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__4 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__4(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__4); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__5 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__5(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__5); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__6 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__6(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__6); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__7 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__7(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__7); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__8 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__8(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__8); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__9 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__9(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__9); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__10 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__10(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__10); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__11 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__11(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__11); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__12 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__12(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__12); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__13 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__13(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__13); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__14 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__14(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__14); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__15 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__15(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__15); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__16 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__16(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__16); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__17 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__17(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__17); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__18 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__18(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__18); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__19 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__19(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__19); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__20 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__20(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__20); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__21 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__21(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__21); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__22 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__22(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__22); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__23 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__23(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__23); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__24 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__24(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__24); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__25 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__25(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__25); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__26 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__26(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__26); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__27 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__27(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__27); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__28 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__28(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__28); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__29 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__29(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__29); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__30 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__30(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__30); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__31 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__31(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__31); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__32 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__32(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__32); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__33 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__33(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__33); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__34 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__34(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__34); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__35 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__35(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__35); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__36 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__36(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__36); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__37 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__37(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__37); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__38 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__38(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__38); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__39 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__39(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__39); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__40 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__40(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__40); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__41 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__41(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__41); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__42 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__42(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__42); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__43 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__43(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__43); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__44 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__44(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__44); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__45 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__45(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__45); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__46 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__46(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__46); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__47 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__47(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__47); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__48 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__48(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__48); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__49 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__49(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__49); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__50 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__50(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__50); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__51 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__51(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__51); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__52 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__52(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__52); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__53 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__53(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__53); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__54 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__54(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__54); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__55 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__55(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__55); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__56 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__56(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__56); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__57 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__57(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__57); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__58 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__58(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__58); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__59 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__59(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__59); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__60 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__60(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__60); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__61 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__61(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__61); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__62 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__62(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__62); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__63 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__63(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__63); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__64 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__64(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__64); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__65 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__65(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__65); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__66 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__66(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__66); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__67 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__67(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__67); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__68 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__68(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__68); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__69 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__69(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__69); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__70 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__70(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__70); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__71 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__71(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__71); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__72 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__72(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__72); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__73 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__73(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__73); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__74 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__74(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__74); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__75 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__75(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__75); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__76 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__76(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__76); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__77 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__77(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__77); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__78 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__78(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__78); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__79 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__79(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__79); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__80 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__80(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__80); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__81 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__81(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__81); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__82 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__82(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__82); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__83 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__83(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__83); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__84 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__84(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__84); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__85 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__85(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__85); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__86 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__86(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__86); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__87 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__87(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__87); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__88 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__88(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__88); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__89 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__89(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__89); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__90 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__90(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__90); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__91 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__91(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__91); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__92 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__92(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__92); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__93 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__93(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__93); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__94 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__94(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__94); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__95 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__95(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__95); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__96 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__96(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__96); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__97 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__97(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__97); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__98 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__98(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__98); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__99 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__99(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__99); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__100 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__100(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__100); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__101 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__101(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__101); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__102 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__102(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__102); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__103 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__103(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__103); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__104 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__104(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__104); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__105 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__105(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__105); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__106 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__106(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__106); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__107 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__107(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__107); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__108 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__108(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__108); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__109 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__109(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__109); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__110 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__110(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__110); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__111 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__111(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__111); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__112 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__112(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__112); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__113 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__113(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__113); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__114 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__114(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__114); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__115 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__115(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__115); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__116 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__116(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__116); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__117 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__117(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__117); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__118 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__118(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__118); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__119 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__119(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__119); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__120 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__120(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__120); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__121 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__121(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__121); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__122 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__122(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__122); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__123 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__123(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__123); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__124 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__124(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__124); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__125 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__125(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__125); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__126 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__126(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__126); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__127 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__127(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__127); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__128 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__128(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__128); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__129 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__129(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__129); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__130 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__130(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__130); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__131 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__131(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__131); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__132 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__132(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__132); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__133 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__133(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__133); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__134 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__134(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__134); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__135 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__135(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__135); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__136 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__136(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__136); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__137 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__137(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__137); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__138 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__138(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__138); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__139 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__139(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__139); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__140 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__140(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__140); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__141 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__141(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__141); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__142 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__142(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__142); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__143 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__143(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__143); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__144 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__144(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__144); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__145 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__145(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__145); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__146 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__146(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__146); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__147 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__147(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__147); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__148 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__148(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__148); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__149 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__149(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__149); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__150 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__150(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__150); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__151 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__151(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__151); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__152 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__152(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__152); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__153 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__153(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__153); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__154 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__154(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__154); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__155 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__155(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__155); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__156 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__156(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__156); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__157 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__157(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__157); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__158 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__158(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__158); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__159 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__159(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__159); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__160 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__160(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__160); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__161 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__161(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__161); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__162 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__162(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__162); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__163 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__163(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__163); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__164 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__164(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__164); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__165 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__165(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__165); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__166 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__166(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__166); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__167 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__167(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__167); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__168 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__168(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__168); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__169 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__169(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__169); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__170 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__170(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__170); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__171 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__171(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__171); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__172 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__172(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__172); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__173 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__173(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__173); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__174 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__174(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__174); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__175 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__175(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__175); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__176 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__176(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__176); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__177 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__177(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__177); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__178 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__178(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__178); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__179 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__179(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__179); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__180 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__180(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__180); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__181 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__181(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__181); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__182 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__182(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__182); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__183 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__183(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__183); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__184 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__184(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__184); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__185 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__185(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__185); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__186 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__186(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__186); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__187 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__187(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__187); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__188 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__188(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__188); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__189 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__189(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__189); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__190 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__190(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__190); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__191 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__191(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__191); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__192 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__192(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__192); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__193 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__193(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__193); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__194 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__194(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__194); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__195 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__195(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__195); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__196 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__196(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__196); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__197 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__197(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__197); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__198 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__198(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__198); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__199 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__199(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__199); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__200 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__200(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__200); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__201 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__201(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__201); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__202 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__202(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__202); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__203 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__203(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__203); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__204 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__204(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__204); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__205 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__205(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__205); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__206 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__206(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__206); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__207 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__207(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__207); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__208 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__208(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__208); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__209 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__209(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__209); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__210 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__210(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__210); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__211 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__211(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__211); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__212 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__212(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__212); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__213 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__213(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__213); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__214 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__214(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__214); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__215 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__215(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__215); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__216 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__216(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__216); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__217 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__217(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__217); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__218 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__218(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__218); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__219 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__219(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__219); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__220 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__220(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__220); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__221 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__221(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__221); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__222 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__222(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__222); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__223 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__223(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__223); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__224 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__224(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__224); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__225 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__225(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__225); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__226 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__226(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__226); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__227 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__227(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__227); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__228 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__228(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__228); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__229 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__229(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__229); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__230 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__230(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__230); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__231 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__231(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__231); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__232 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__232(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__232); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__233 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__233(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__233); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__234 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__234(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__234); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__235 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__235(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__235); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__236 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__236(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__236); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__237 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__237(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__237); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__238 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__238(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__238); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__239 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__239(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__239); +l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__240 = _init_l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__240(); +lean_mark_persistent(l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandDeclare__config__elab______1___closed__240); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Elab/Tactic/Conv/Basic.c b/stage0/stdlib/Lean/Elab/Tactic/Conv/Basic.c index b21ecd6160c..18edbac586b 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Conv/Basic.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Conv/Basic.c @@ -40,7 +40,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalNestedConv_declRang static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalWhnf___closed__14; lean_object* l_Lean_Elab_Tactic_evalTacticSeq1Indented(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___regBuiltin_Lean_Elab_Tactic_Conv_evalConvSeq_declRange(lean_object*); -extern lean_object* l_Lean_nullKind; lean_object* l_Lean_Elab_Tactic_SavedState_restore(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_List_mapM___at_Lean_Elab_Tactic_Conv_remarkAsConvGoal___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalReduce___closed__5; @@ -143,6 +142,7 @@ lean_object* l_Lean_Meta_replaceTargetEq(lean_object*, lean_object*, lean_object static lean_object* l_Lean_Elab_Tactic_Conv_evalConvSeqBracketed___lambda__2___closed__4; lean_object* l_Lean_Meta_getMVarType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalConvSeqBracketed_declRange(lean_object*); +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); lean_object* l_Lean_Meta_applyRefl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalNestedTactic_declRange___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalConvConvSeq_declRange___closed__7; @@ -256,7 +256,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalConvSeq1Indent lean_object* l_Lean_Meta_matchEq_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalNestedTactic___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalConvSeqBracketed(lean_object*); -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalNestedTactic___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalDecl_type(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalConvConvSeq_declRange___closed__6; @@ -6686,22 +6685,21 @@ return x_1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalConv___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, 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) { _start: { -lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +lean_object* x_16; lean_object* x_17; uint8_t x_18; x_16 = lean_unsigned_to_nat(2u); x_17 = l_Lean_Syntax_getArg(x_1, x_16); -x_18 = l_Lean_nullKind; lean_inc(x_17); -x_19 = l_Lean_Syntax_isNodeOf(x_17, x_18, x_16); -if (x_19 == 0) +x_18 = l_Lean_Syntax_matchesNull(x_17, x_16); +if (x_18 == 0) { -lean_object* x_20; uint8_t x_21; +lean_object* x_19; uint8_t x_20; lean_dec(x_4); lean_dec(x_3); -x_20 = lean_unsigned_to_nat(0u); -x_21 = l_Lean_Syntax_isNodeOf(x_17, x_18, x_20); -if (x_21 == 0) +x_19 = lean_unsigned_to_nat(0u); +x_20 = l_Lean_Syntax_matchesNull(x_17, x_19); +if (x_20 == 0) { -lean_object* x_22; +lean_object* x_21; lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); @@ -6712,24 +6710,25 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_2); -x_22 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalExact___spec__1___rarg(x_15); -return x_22; +x_21 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalExact___spec__1___rarg(x_15); +return x_21; } else { -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; lean_object* x_30; lean_object* x_31; uint8_t x_32; -x_23 = lean_unsigned_to_nat(3u); -x_24 = l_Lean_Syntax_getArg(x_1, x_23); -x_25 = lean_unsigned_to_nat(4u); -x_26 = l_Lean_Syntax_getArg(x_1, x_25); -x_27 = l_Lean_Elab_Tactic_Conv_evalConvSeqBracketed___lambda__2___closed__24; -x_28 = lean_array_push(x_27, x_2); -x_29 = lean_array_push(x_28, x_24); -x_30 = lean_box(2); +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; lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_22 = lean_unsigned_to_nat(3u); +x_23 = l_Lean_Syntax_getArg(x_1, x_22); +x_24 = lean_unsigned_to_nat(4u); +x_25 = l_Lean_Syntax_getArg(x_1, x_24); +x_26 = l_Lean_Elab_Tactic_Conv_evalConvSeqBracketed___lambda__2___closed__24; +x_27 = lean_array_push(x_26, x_2); +x_28 = lean_array_push(x_27, x_23); +x_29 = lean_box(2); +x_30 = l_Lean_Elab_Tactic_Conv_evalConvSeqBracketed___lambda__2___closed__9; x_31 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_18); -lean_ctor_set(x_31, 2, x_29); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +lean_ctor_set(x_31, 2, x_28); x_32 = !lean_is_exclusive(x_13); if (x_32 == 0) { @@ -6742,7 +6741,7 @@ lean_ctor_set(x_13, 5, x_34); if (lean_obj_tag(x_6) == 0) { lean_object* x_35; -x_35 = l___private_Lean_Elab_Tactic_Conv_Basic_0__Lean_Elab_Tactic_Conv_convTarget(x_26, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_35 = l___private_Lean_Elab_Tactic_Conv_Basic_0__Lean_Elab_Tactic_Conv_convTarget(x_25, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); return x_35; } else @@ -6753,7 +6752,7 @@ lean_inc(x_36); lean_dec(x_6); x_37 = l_Lean_Syntax_getId(x_36); lean_dec(x_36); -x_38 = l___private_Lean_Elab_Tactic_Conv_Basic_0__Lean_Elab_Tactic_Conv_convLocalDecl(x_26, x_37, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_38 = l___private_Lean_Elab_Tactic_Conv_Basic_0__Lean_Elab_Tactic_Conv_convLocalDecl(x_25, x_37, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); return x_38; } } @@ -6801,7 +6800,7 @@ lean_ctor_set(x_51, 10, x_49); if (lean_obj_tag(x_6) == 0) { lean_object* x_52; -x_52 = l___private_Lean_Elab_Tactic_Conv_Basic_0__Lean_Elab_Tactic_Conv_convTarget(x_26, x_7, x_8, x_9, x_10, x_11, x_12, x_51, x_14, x_15); +x_52 = l___private_Lean_Elab_Tactic_Conv_Basic_0__Lean_Elab_Tactic_Conv_convTarget(x_25, x_7, x_8, x_9, x_10, x_11, x_12, x_51, x_14, x_15); return x_52; } else @@ -6812,7 +6811,7 @@ lean_inc(x_53); lean_dec(x_6); x_54 = l_Lean_Syntax_getId(x_53); lean_dec(x_53); -x_55 = l___private_Lean_Elab_Tactic_Conv_Basic_0__Lean_Elab_Tactic_Conv_convLocalDecl(x_26, x_54, x_7, x_8, x_9, x_10, x_11, x_12, x_51, x_14, x_15); +x_55 = l___private_Lean_Elab_Tactic_Conv_Basic_0__Lean_Elab_Tactic_Conv_convLocalDecl(x_25, x_54, x_7, x_8, x_9, x_10, x_11, x_12, x_51, x_14, x_15); return x_55; } } @@ -6828,33 +6827,33 @@ x_58 = lean_unsigned_to_nat(3u); x_59 = l_Lean_Syntax_getArg(x_1, x_58); x_60 = lean_unsigned_to_nat(4u); x_61 = l_Lean_Syntax_getArg(x_1, x_60); +x_62 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalConvSeq___closed__1; +lean_inc(x_3); +x_63 = lean_name_mk_string(x_3, x_62); lean_inc(x_13); -x_62 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Tactic_evalIntro___spec__1___rarg(x_13, x_14, x_15); -x_63 = lean_ctor_get(x_62, 0); -lean_inc(x_63); -x_64 = lean_ctor_get(x_62, 1); -lean_inc(x_64); -lean_dec(x_62); -x_65 = lean_st_ref_get(x_14, x_64); -x_66 = lean_ctor_get(x_65, 1); +x_64 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Tactic_evalIntro___spec__1___rarg(x_13, x_14, x_15); +x_65 = lean_ctor_get(x_64, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_64, 1); lean_inc(x_66); -lean_dec(x_65); -x_67 = l_Lean_Syntax_getHeadInfo_x3f(x_2); +lean_dec(x_64); +x_67 = lean_st_ref_get(x_14, x_66); +x_68 = lean_ctor_get(x_67, 1); +lean_inc(x_68); +lean_dec(x_67); +x_69 = l_Lean_Syntax_getHeadInfo_x3f(x_2); lean_dec(x_2); -x_68 = l_Lean_Syntax_getHeadInfo_x3f(x_59); +x_70 = l_Lean_Syntax_getHeadInfo_x3f(x_59); lean_dec(x_59); -x_69 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalConvSeq___closed__1; -lean_inc(x_3); -x_70 = lean_name_mk_string(x_3, x_69); x_71 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalConvSeq1Indented___closed__1; lean_inc(x_3); x_72 = lean_name_mk_string(x_3, x_71); x_73 = l_Lean_Elab_Tactic_Conv_evalConv___lambda__1___closed__1; lean_inc(x_3); x_74 = lean_name_mk_string(x_3, x_73); -lean_inc(x_63); +lean_inc(x_65); x_75 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_75, 0, x_63); +lean_ctor_set(x_75, 0, x_65); lean_ctor_set(x_75, 1, x_73); x_76 = l_Lean_Elab_Tactic_Conv_evalConvSeqBracketed___lambda__2___closed__24; x_77 = lean_array_push(x_76, x_75); @@ -6865,9 +6864,9 @@ lean_ctor_set(x_80, 0, x_79); lean_ctor_set(x_80, 1, x_74); lean_ctor_set(x_80, 2, x_78); x_81 = l_Lean_Elab_Tactic_Conv_evalConv___lambda__1___closed__2; -lean_inc(x_63); +lean_inc(x_65); x_82 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_82, 0, x_63); +lean_ctor_set(x_82, 0, x_65); lean_ctor_set(x_82, 1, x_81); x_83 = l_Lean_Elab_Tactic_Conv_evalConvSeqBracketed___lambda__2___closed__21; x_84 = lean_array_push(x_83, x_82); @@ -6886,14 +6885,14 @@ lean_ctor_set(x_90, 2, x_88); x_91 = l_Lean_Elab_Tactic_Conv_evalConvSeqBracketed___lambda__2___closed__12; x_92 = lean_name_mk_string(x_3, x_91); x_93 = l_Lean_Elab_Tactic_Conv_evalConvSeqBracketed___lambda__2___closed__14; -lean_inc(x_63); +lean_inc(x_65); x_94 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_94, 0, x_63); +lean_ctor_set(x_94, 0, x_65); lean_ctor_set(x_94, 1, x_93); x_95 = l_Lean_Elab_Tactic_Conv_evalConvSeqBracketed___lambda__2___closed__25; -lean_inc(x_63); +lean_inc(x_65); x_96 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_96, 0, x_63); +lean_ctor_set(x_96, 0, x_65); lean_ctor_set(x_96, 1, x_95); x_97 = l_Lean_Elab_Tactic_Conv_evalConvSeqBracketed___lambda__2___closed__26; x_98 = lean_array_push(x_97, x_94); @@ -6924,20 +6923,20 @@ lean_ctor_set(x_110, 2, x_109); x_111 = lean_array_push(x_83, x_110); x_112 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_112, 0, x_79); -lean_ctor_set(x_112, 1, x_70); +lean_ctor_set(x_112, 1, x_63); lean_ctor_set(x_112, 2, x_111); -if (lean_obj_tag(x_67) == 0) +if (lean_obj_tag(x_69) == 0) { -lean_inc(x_63); -x_113 = x_63; +lean_inc(x_65); +x_113 = x_65; goto block_157; } else { lean_object* x_158; -x_158 = lean_ctor_get(x_67, 0); +x_158 = lean_ctor_get(x_69, 0); lean_inc(x_158); -lean_dec(x_67); +lean_dec(x_69); x_113 = x_158; goto block_157; } @@ -6956,12 +6955,12 @@ lean_object* x_118; lean_object* x_119; lean_object* x_120; x_118 = l_Lean_Elab_Tactic_Conv_evalConv___lambda__1___closed__6; x_119 = lean_array_push(x_117, x_118); x_120 = lean_array_push(x_119, x_103); -if (lean_obj_tag(x_68) == 0) +if (lean_obj_tag(x_70) == 0) { lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; x_121 = l_Lean_Elab_Tactic_Conv_evalConv___lambda__1___closed__7; x_122 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_122, 0, x_63); +lean_ctor_set(x_122, 0, x_65); lean_ctor_set(x_122, 1, x_121); x_123 = lean_array_push(x_120, x_122); x_124 = lean_array_push(x_123, x_112); @@ -6969,16 +6968,16 @@ x_125 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_125, 0, x_79); lean_ctor_set(x_125, 1, x_4); lean_ctor_set(x_125, 2, x_124); -x_126 = l_Lean_Elab_Tactic_evalTacticAux(x_125, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_66); +x_126 = l_Lean_Elab_Tactic_evalTacticAux(x_125, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_68); return x_126; } else { lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; -lean_dec(x_63); -x_127 = lean_ctor_get(x_68, 0); +lean_dec(x_65); +x_127 = lean_ctor_get(x_70, 0); lean_inc(x_127); -lean_dec(x_68); +lean_dec(x_70); x_128 = l_Lean_Elab_Tactic_Conv_evalConv___lambda__1___closed__7; x_129 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_129, 0, x_127); @@ -6989,7 +6988,7 @@ x_132 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_132, 0, x_79); lean_ctor_set(x_132, 1, x_4); lean_ctor_set(x_132, 2, x_131); -x_133 = l_Lean_Elab_Tactic_evalTacticAux(x_132, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_66); +x_133 = l_Lean_Elab_Tactic_evalTacticAux(x_132, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_68); return x_133; } } @@ -7000,9 +6999,9 @@ x_134 = lean_ctor_get(x_6, 0); lean_inc(x_134); lean_dec(x_6); x_135 = l_Lean_Elab_Tactic_Conv_evalConv___lambda__1___closed__8; -lean_inc(x_63); +lean_inc(x_65); x_136 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_136, 0, x_63); +lean_ctor_set(x_136, 0, x_65); lean_ctor_set(x_136, 1, x_135); x_137 = lean_array_push(x_76, x_136); x_138 = lean_array_push(x_137, x_134); @@ -7014,12 +7013,12 @@ lean_ctor_set(x_141, 1, x_85); lean_ctor_set(x_141, 2, x_140); x_142 = lean_array_push(x_117, x_141); x_143 = lean_array_push(x_142, x_103); -if (lean_obj_tag(x_68) == 0) +if (lean_obj_tag(x_70) == 0) { lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; x_144 = l_Lean_Elab_Tactic_Conv_evalConv___lambda__1___closed__7; x_145 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_145, 0, x_63); +lean_ctor_set(x_145, 0, x_65); lean_ctor_set(x_145, 1, x_144); x_146 = lean_array_push(x_143, x_145); x_147 = lean_array_push(x_146, x_112); @@ -7027,16 +7026,16 @@ x_148 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_148, 0, x_79); lean_ctor_set(x_148, 1, x_4); lean_ctor_set(x_148, 2, x_147); -x_149 = l_Lean_Elab_Tactic_evalTacticAux(x_148, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_66); +x_149 = l_Lean_Elab_Tactic_evalTacticAux(x_148, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_68); return x_149; } else { lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; -lean_dec(x_63); -x_150 = lean_ctor_get(x_68, 0); +lean_dec(x_65); +x_150 = lean_ctor_get(x_70, 0); lean_inc(x_150); -lean_dec(x_68); +lean_dec(x_70); x_151 = l_Lean_Elab_Tactic_Conv_evalConv___lambda__1___closed__7; x_152 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_152, 0, x_150); @@ -7047,7 +7046,7 @@ x_155 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_155, 0, x_79); lean_ctor_set(x_155, 1, x_4); lean_ctor_set(x_155, 2, x_154); -x_156 = l_Lean_Elab_Tactic_evalTacticAux(x_155, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_66); +x_156 = l_Lean_Elab_Tactic_evalTacticAux(x_155, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_68); return x_156; } } @@ -7097,14 +7096,13 @@ x_17 = l_Lean_Syntax_getArg(x_1, x_16); x_18 = l_Lean_Syntax_isNone(x_17); if (x_18 == 0) { -lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_19 = l_Lean_nullKind; -x_20 = lean_unsigned_to_nat(2u); +lean_object* x_19; uint8_t x_20; +x_19 = lean_unsigned_to_nat(2u); lean_inc(x_17); -x_21 = l_Lean_Syntax_isNodeOf(x_17, x_19, x_20); -if (x_21 == 0) +x_20 = l_Lean_Syntax_matchesNull(x_17, x_19); +if (x_20 == 0) { -lean_object* x_22; +lean_object* x_21; lean_dec(x_17); lean_dec(x_15); lean_dec(x_9); @@ -7116,33 +7114,33 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_22 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalExact___spec__1___rarg(x_10); -return x_22; +x_21 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalExact___spec__1___rarg(x_10); +return x_21; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_23 = l_Lean_Syntax_getArg(x_17, x_16); +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_22 = l_Lean_Syntax_getArg(x_17, x_16); lean_dec(x_17); -x_24 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_24, 0, x_23); -x_25 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalWhnf___closed__8; -x_26 = lean_box(0); -x_27 = l_Lean_Elab_Tactic_Conv_evalConv___lambda__1(x_1, x_15, x_25, x_11, x_26, x_24, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_23 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_23, 0, x_22); +x_24 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalWhnf___closed__8; +x_25 = lean_box(0); +x_26 = l_Lean_Elab_Tactic_Conv_evalConv___lambda__1(x_1, x_15, x_24, x_11, x_25, x_23, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_1); -return x_27; +return x_26; } } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_dec(x_17); -x_28 = lean_box(0); -x_29 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalWhnf___closed__8; -x_30 = lean_box(0); -x_31 = l_Lean_Elab_Tactic_Conv_evalConv___lambda__1(x_1, x_15, x_29, x_11, x_30, x_28, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_27 = lean_box(0); +x_28 = l___regBuiltin_Lean_Elab_Tactic_Conv_evalWhnf___closed__8; +x_29 = lean_box(0); +x_30 = l_Lean_Elab_Tactic_Conv_evalConv___lambda__1(x_1, x_15, x_28, x_11, x_29, x_27, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_1); -return x_31; +return x_30; } } } diff --git a/stage0/stdlib/Lean/Elab/Tactic/Conv/Congr.c b/stage0/stdlib/Lean/Elab/Tactic/Conv/Congr.c index f0aa27c6f7b..bf36d66e608 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Conv/Congr.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Conv/Congr.c @@ -31,7 +31,6 @@ lean_object* lean_mk_empty_array_with_capacity(lean_object*); lean_object* l_Lean_mkSort(lean_object*); lean_object* l_List_get___rarg(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalExt_declRange___closed__7; -extern lean_object* l_Lean_nullKind; lean_object* l_Lean_mkApp6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Conv_Congr_0__Lean_Elab_Tactic_Conv_congrImplies___closed__2; static lean_object* l_Lean_Elab_Tactic_Conv_evalArg___lambda__2___closed__4; @@ -55,11 +54,11 @@ lean_object* l_Lean_throwError___at___private_Lean_Meta_Basic_0__Lean_Meta_proce static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalCongr_declRange___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalLhs_declRange___closed__6; lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalExact___spec__1___rarg(lean_object*); +lean_object* l_Lean_TSyntax_getNat(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalArg___closed__1; static lean_object* l___private_Lean_Elab_Tactic_Conv_Congr_0__Lean_Elab_Tactic_Conv_extCore___lambda__2___closed__2; static lean_object* l_Lean_Elab_Tactic_Conv_evalArg___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Elab_Tactic_Conv_Congr_0__Lean_Elab_Tactic_Conv_congrApp___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalExt_declRange___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_congr___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -119,9 +118,9 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalRhs___rarg(lean_object*, le static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalRhs_declRange___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalRhs_declRange___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalArg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); lean_object* l_Lean_Meta_applyRefl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Conv_Congr_0__Lean_Elab_Tactic_Conv_selectIdx___lambda__1___closed__4; -extern lean_object* l_Lean_numLitKind; lean_object* lean_nat_sub(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalCongr___closed__10; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalLhs(lean_object*); @@ -192,7 +191,6 @@ lean_object* l_Int_toNat(lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Conv_Congr_0__Lean_Elab_Tactic_Conv_extCore___lambda__2___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalLhs_declRange___closed__5; LEAN_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Elab_Tactic_Conv_Congr_0__Lean_Elab_Tactic_Conv_selectIdx___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalExt(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalRhs___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalLhs(lean_object*); @@ -4635,52 +4633,52 @@ static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalArg___lambda__2___closed__ _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("invalid 'arg' conv tactic, index must be greater than 0", 55); +x_1 = lean_mk_string_from_bytes("num", 3); return x_1; } } static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalArg___lambda__2___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic_Conv_evalArg___lambda__2___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Tactic_Conv_evalArg___lambda__2___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; } } static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalArg___lambda__2___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("num", 3); +x_1 = lean_mk_string_from_bytes("invalid 'arg' conv tactic, index must be greater than 0", 55); return x_1; } } static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalArg___lambda__2___closed__4() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic_Conv_evalArg___lambda__2___closed__3; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Tactic_Conv_evalArg___lambda__2___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalArg___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_13; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_dec(x_2); -x_25 = lean_unsigned_to_nat(2u); -x_26 = l_Lean_Syntax_getArg(x_1, x_25); +x_13 = lean_unsigned_to_nat(2u); +x_14 = l_Lean_Syntax_getArg(x_1, x_13); lean_dec(x_1); -x_27 = l_Lean_Elab_Tactic_Conv_evalArg___lambda__2___closed__4; -lean_inc(x_26); -x_28 = l_Lean_Syntax_isOfKind(x_26, x_27); -if (x_28 == 0) +x_15 = l_Lean_Elab_Tactic_Conv_evalArg___lambda__2___closed__2; +lean_inc(x_14); +x_16 = l_Lean_Syntax_isOfKind(x_14, x_15); +if (x_16 == 0) { -lean_object* x_29; -lean_dec(x_26); +lean_object* x_17; +lean_dec(x_14); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -4690,53 +4688,31 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_29 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalExact___spec__1___rarg(x_12); -return x_29; -} -else -{ -lean_object* x_30; lean_object* x_31; -x_30 = l_Lean_numLitKind; -x_31 = l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(x_30, x_26); -lean_dec(x_26); -if (lean_obj_tag(x_31) == 0) -{ -lean_object* x_32; -x_32 = lean_unsigned_to_nat(0u); -x_13 = x_32; -goto block_24; +x_17 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalExact___spec__1___rarg(x_12); +return x_17; } else { -lean_object* x_33; -x_33 = lean_ctor_get(x_31, 0); -lean_inc(x_33); -lean_dec(x_31); -x_13 = x_33; -goto block_24; -} -} -block_24: -{ -lean_object* x_14; uint8_t x_15; -x_14 = lean_unsigned_to_nat(0u); -x_15 = lean_nat_dec_eq(x_13, x_14); -if (x_15 == 0) +lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_18 = l_Lean_TSyntax_getNat(x_14); +x_19 = lean_unsigned_to_nat(0u); +x_20 = lean_nat_dec_eq(x_18, x_19); +if (x_20 == 0) { -lean_object* x_16; lean_object* x_17; -x_16 = lean_box(0); -x_17 = l_Lean_Elab_Tactic_Conv_evalArg___lambda__1(x_13, x_3, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_object* x_21; lean_object* x_22; +x_21 = lean_box(0); +x_22 = l_Lean_Elab_Tactic_Conv_evalArg___lambda__1(x_18, x_3, x_21, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_3); -lean_dec(x_13); -return x_17; +lean_dec(x_18); +return x_22; } else { -lean_object* x_18; lean_object* x_19; uint8_t x_20; -lean_dec(x_13); +lean_object* x_23; lean_object* x_24; uint8_t x_25; +lean_dec(x_18); lean_dec(x_3); -x_18 = l_Lean_Elab_Tactic_Conv_evalArg___lambda__2___closed__2; -x_19 = l_Lean_throwError___at_Lean_Elab_Tactic_refineCore___spec__1(x_18, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_23 = l_Lean_Elab_Tactic_Conv_evalArg___lambda__2___closed__4; +x_24 = l_Lean_throwError___at_Lean_Elab_Tactic_refineCore___spec__1(x_23, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -4745,23 +4721,23 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_20 = !lean_is_exclusive(x_19); -if (x_20 == 0) +x_25 = !lean_is_exclusive(x_24); +if (x_25 == 0) { -return x_19; +return x_24; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_19, 0); -x_22 = lean_ctor_get(x_19, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_19); -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); -return x_23; +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_24, 0); +x_27 = lean_ctor_get(x_24, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_24); +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +return x_28; } } } @@ -4807,13 +4783,12 @@ x_15 = l_Lean_Syntax_getArg(x_1, x_14); x_16 = l_Lean_Syntax_isNone(x_15); if (x_16 == 0) { -lean_object* x_17; uint8_t x_18; -x_17 = l_Lean_nullKind; +uint8_t x_17; lean_inc(x_15); -x_18 = l_Lean_Syntax_isNodeOf(x_15, x_17, x_14); -if (x_18 == 0) +x_17 = l_Lean_Syntax_matchesNull(x_15, x_14); +if (x_17 == 0) { -lean_object* x_19; +lean_object* x_18; lean_dec(x_15); lean_dec(x_9); lean_dec(x_8); @@ -4824,30 +4799,30 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_19 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalExact___spec__1___rarg(x_10); -return x_19; +x_18 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalExact___spec__1___rarg(x_10); +return x_18; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_20 = lean_unsigned_to_nat(0u); -x_21 = l_Lean_Syntax_getArg(x_15, x_20); +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_19 = lean_unsigned_to_nat(0u); +x_20 = l_Lean_Syntax_getArg(x_15, x_19); lean_dec(x_15); -x_22 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_22, 0, x_21); -x_23 = lean_box(0); -x_24 = l_Lean_Elab_Tactic_Conv_evalArg___lambda__2(x_1, x_23, x_22, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_24; +x_21 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_21, 0, x_20); +x_22 = lean_box(0); +x_23 = l_Lean_Elab_Tactic_Conv_evalArg___lambda__2(x_1, x_22, x_21, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_23; } } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_dec(x_15); +x_24 = lean_box(0); x_25 = lean_box(0); -x_26 = lean_box(0); -x_27 = l_Lean_Elab_Tactic_Conv_evalArg___lambda__2(x_1, x_26, x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_27; +x_26 = l_Lean_Elab_Tactic_Conv_evalArg___lambda__2(x_1, x_25, x_24, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_26; } } } diff --git a/stage0/stdlib/Lean/Elab/Tactic/Delta.c b/stage0/stdlib/Lean/Elab/Tactic/Delta.c index 3f66022b9aa..147bf97d481 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Delta.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Delta.c @@ -22,6 +22,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalDelta___lambda__1(lean_object*, lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalDelta(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalDecl_userName(lean_object*); +lean_object* l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); lean_object* lean_name_mk_string(lean_object*, lean_object*); lean_object* l_List_mapTRAux___at_Lean_resolveGlobalConstCore___spec__2(lean_object*, lean_object*); extern lean_object* l_Std_Format_defWidth; @@ -79,7 +80,6 @@ lean_object* l_Lean_LocalDecl_type(lean_object*); static lean_object* l_Lean_Elab_Tactic_deltaLocalDecl___closed__1; lean_object* l_Lean_Meta_throwTacticEx___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_deltaLocalDecl___lambda__1___boxed(lean_object*, lean_object*); -lean_object* l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_Tactic_evalDelta___spec__1___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_evalDelta_declRange(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_evalDelta___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -981,7 +981,7 @@ x_13 = lean_ctor_get(x_11, 1); lean_inc(x_13); lean_dec(x_11); x_14 = lean_box(0); -x_15 = l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(x_12, x_14); +x_15 = l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(x_12, x_14); x_16 = l_List_isEmpty___rarg(x_15); if (x_16 == 0) { diff --git a/stage0/stdlib/Lean/Elab/Tactic/Induction.c b/stage0/stdlib/Lean/Elab/Tactic/Induction.c index e8792276c8c..13fa652fbad 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Induction.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Induction.c @@ -51,11 +51,11 @@ LEAN_EXPORT lean_object* l_Lean_Elab_withSaveInfoContext___at___private_Lean_Ela static lean_object* l_Lean_Elab_Tactic_ElimApp_setMotiveArg___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalInduction___closed__4; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_go___spec__5___lambda__2___closed__2; -extern lean_object* l_Lean_nullKind; LEAN_EXPORT lean_object* l_Std_PersistentArray_mapM___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_saveAltVarsInfo___spec__10___boxed(lean_object*, 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_Tactic_evalInduction_checkTargets___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_whnfForall(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalAlt___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOptionM___at_Lean_Elab_Tactic_expandTacticMacroFns_loop___spec__2(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_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_go___spec__5___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -188,6 +188,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabCasesTargets___boxed__const__1; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_generalizeTargets___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_ElimApp_evalAlts_go___spec__5___lambda__3___closed__2; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_checkAltsOfOptInductionAlts___spec__1___closed__2; +static lean_object* l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_evalAlt___spec__1___closed__3; LEAN_EXPORT lean_object* l_Std_PersistentArray_mapMAux___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_saveAltVarsInfo___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalInduction___closed__6; static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getElimNameInfo___lambda__2___closed__3; @@ -391,7 +392,6 @@ static lean_object* l_Lean_Elab_Tactic_ElimApp_mkElimApp_loop___closed__1; lean_object* l_Lean_Meta_throwTacticEx___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_getNumExplicitFields___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_appendGoals(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_getElimNameInfo___lambda__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_assignExprMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_ElimApp_evalAlts___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -547,6 +547,7 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Tact LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_generalizeVars___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_getAltNumFields___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_evalAlt___spec__1___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_getBindingName___boxed(lean_object*); lean_object* l_Array_findIdx_x3f_loop___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_ElimApp_getAltNumFields___lambda__1___closed__2; @@ -971,6 +972,24 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } +static lean_object* _init_l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_evalAlt___spec__1___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("null", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_evalAlt___spec__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_evalAlt___spec__1___closed__2; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_evalAlt___spec__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, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { @@ -979,7 +998,7 @@ x_13 = l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_evalAlt___spec__1___ x_14 = lean_array_push(x_13, x_1); x_15 = lean_array_push(x_14, x_2); x_16 = lean_box(2); -x_17 = l_Lean_nullKind; +x_17 = l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_evalAlt___spec__1___closed__3; x_18 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_18, 0, x_16); lean_ctor_set(x_18, 1, x_17); @@ -13063,7 +13082,7 @@ static lean_object* _init_l___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Ta { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean_nullKind; +x_2 = l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_evalAlt___spec__1___closed__3; x_3 = l_Lean_Elab_Tactic_ElimApp_State_alts___default___closed__1; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -13156,7 +13175,7 @@ x_8 = lean_array_uset(x_4, x_3, x_7); x_9 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_expandMultiAlt_x3f___spec__1___closed__1; x_10 = lean_array_push(x_9, x_6); x_11 = lean_box(2); -x_12 = l_Lean_nullKind; +x_12 = l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_evalAlt___spec__1___closed__3; x_13 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_13, 0, x_11); lean_ctor_set(x_13, 1, x_12); @@ -13342,7 +13361,7 @@ x_13 = l_Lean_Elab_Tactic_ElimApp_State_alts___default___closed__1; x_14 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_expandInductionAlts_x3f___spec__2(x_2, x_10, x_9, x_13); lean_dec(x_2); x_15 = lean_box(2); -x_16 = l_Lean_nullKind; +x_16 = l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_evalAlt___spec__1___closed__3; x_17 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_17, 0, x_15); lean_ctor_set(x_17, 1, x_16); @@ -13416,7 +13435,7 @@ x_10 = lean_ctor_get(x_7, 0); x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_expandMultiAlt_x3f___spec__1___closed__1; x_12 = lean_array_push(x_11, x_10); x_13 = lean_box(2); -x_14 = l_Lean_nullKind; +x_14 = l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_evalAlt___spec__1___closed__3; x_15 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_15, 0, x_13); lean_ctor_set(x_15, 1, x_14); @@ -13434,7 +13453,7 @@ lean_dec(x_7); x_18 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_expandMultiAlt_x3f___spec__1___closed__1; x_19 = lean_array_push(x_18, x_17); x_20 = lean_box(2); -x_21 = l_Lean_nullKind; +x_21 = l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_evalAlt___spec__1___closed__3; x_22 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_22, 0, x_20); lean_ctor_set(x_22, 1, x_21); @@ -13488,7 +13507,7 @@ x_10 = lean_ctor_get(x_7, 0); x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_expandMultiAlt_x3f___spec__1___closed__1; x_12 = lean_array_push(x_11, x_10); x_13 = lean_box(2); -x_14 = l_Lean_nullKind; +x_14 = l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_evalAlt___spec__1___closed__3; x_15 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_15, 0, x_13); lean_ctor_set(x_15, 1, x_14); @@ -13506,7 +13525,7 @@ lean_dec(x_7); x_18 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Induction_0__Lean_Elab_Tactic_expandMultiAlt_x3f___spec__1___closed__1; x_19 = lean_array_push(x_18, x_17); x_20 = lean_box(2); -x_21 = l_Lean_nullKind; +x_21 = l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_evalAlt___spec__1___closed__3; x_22 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_22, 0, x_20); lean_ctor_set(x_22, 1, x_21); @@ -14522,7 +14541,7 @@ x_13 = lean_ctor_get(x_11, 1); lean_inc(x_13); lean_dec(x_11); x_14 = lean_box(0); -x_15 = l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(x_12, x_14); +x_15 = l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(x_12, x_14); x_16 = l_List_isEmpty___rarg(x_15); if (x_16 == 0) { @@ -21242,6 +21261,10 @@ l_Lean_Elab_Tactic_isHoleRHS___closed__10 = _init_l_Lean_Elab_Tactic_isHoleRHS__ lean_mark_persistent(l_Lean_Elab_Tactic_isHoleRHS___closed__10); l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_evalAlt___spec__1___closed__1 = _init_l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_evalAlt___spec__1___closed__1(); lean_mark_persistent(l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_evalAlt___spec__1___closed__1); +l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_evalAlt___spec__1___closed__2 = _init_l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_evalAlt___spec__1___closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_evalAlt___spec__1___closed__2); +l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_evalAlt___spec__1___closed__3 = _init_l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_evalAlt___spec__1___closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic_withCaseRef___at_Lean_Elab_Tactic_evalAlt___spec__1___closed__3); l_Lean_Elab_Tactic_evalAlt___lambda__2___closed__1 = _init_l_Lean_Elab_Tactic_evalAlt___lambda__2___closed__1(); lean_mark_persistent(l_Lean_Elab_Tactic_evalAlt___lambda__2___closed__1); l_Lean_Elab_Tactic_evalAlt___lambda__2___closed__2 = _init_l_Lean_Elab_Tactic_evalAlt___lambda__2___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Match.c b/stage0/stdlib/Lean/Elab/Tactic/Match.c index 91cc67837f6..c6ce0a531a5 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Match.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Match.c @@ -26,7 +26,6 @@ lean_object* lean_mk_empty_array_with_capacity(lean_object*); lean_object* l_Lean_Syntax_getHeadInfo_x3f(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__2___closed__2; static lean_object* l_Lean_Elab_Tactic_evalMatch___closed__2; -extern lean_object* l_Lean_nullKind; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_name_mk_string(lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); @@ -45,7 +44,7 @@ lean_object* l_Lean_Elab_Tactic_evalTacticAux(lean_object*, lean_object*, lean_o static lean_object* l_Lean_Elab_Tactic_evalMatch___closed__4; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__1; lean_object* lean_st_ref_get(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___boxed(lean_object*, lean_object*, lean_object*, 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__2___closed__6; lean_object* l_Lean_mkIdentFrom(lean_object*, lean_object*); lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(lean_object*, lean_object*); @@ -55,6 +54,7 @@ lean_object* lean_array_push(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__11; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__9; +static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__27; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__4; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__23; lean_object* lean_string_utf8_byte_size(lean_object*); @@ -117,7 +117,7 @@ LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Tactic_ extern lean_object* l_Lean_Elab_Tactic_tacticElabAttribute; lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__20; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__26; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalMatch___spec__5___rarg(lean_object*); lean_object* l_Lean_Syntax_getSepArgs(lean_object*); @@ -191,19 +191,37 @@ static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tact _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___lambda__1___boxed), 7, 0); +x_1 = lean_mk_string_from_bytes("null", 4); return x_1; } } static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__3() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__2; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___lambda__1___boxed), 7, 0); +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__5() { +_start: +{ lean_object* x_1; x_1 = lean_mk_string_from_bytes("syntheticHole", 13); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__4() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__6() { _start: { lean_object* x_1; @@ -211,7 +229,7 @@ x_1 = lean_mk_string_from_bytes("hole", 4); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__5() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__7() { _start: { lean_object* x_1; @@ -219,7 +237,7 @@ x_1 = lean_mk_string_from_bytes("?", 1); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__6() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__8() { _start: { lean_object* x_1; @@ -227,22 +245,22 @@ x_1 = lean_mk_string_from_bytes("rhs", 3); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__7() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__6; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__8; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__8() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__6; +x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__8; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__7; +x_3 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__9; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -250,17 +268,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__9() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__6; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__8; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__10() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__12() { _start: { lean_object* x_1; lean_object* x_2; @@ -269,7 +287,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__11() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__13() { _start: { lean_object* x_1; @@ -277,7 +295,7 @@ x_1 = lean_mk_string_from_bytes("Tactic", 6); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__12() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__14() { _start: { lean_object* x_1; @@ -285,25 +303,15 @@ x_1 = lean_mk_string_from_bytes("case", 4); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__13() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__15() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("null", 4); +x_1 = lean_mk_string_from_bytes("binderIdent", 11); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__14() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__13; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__15() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__16() { _start: { lean_object* x_1; @@ -311,7 +319,7 @@ x_1 = lean_mk_string_from_bytes("tacticSeq", 9); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__16() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__17() { _start: { lean_object* x_1; @@ -319,7 +327,7 @@ x_1 = lean_mk_string_from_bytes("tacticSeq1Indented", 18); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__17() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__18() { _start: { lean_object* x_1; @@ -327,17 +335,17 @@ x_1 = lean_mk_string_from_bytes("group", 5); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__18() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__17; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__18; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__19() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__20() { _start: { lean_object* x_1; @@ -345,7 +353,7 @@ x_1 = lean_mk_string_from_bytes("withAnnotateState", 17); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__20() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__21() { _start: { lean_object* x_1; @@ -353,7 +361,7 @@ x_1 = lean_mk_string_from_bytes("with_annotate_state", 19); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__21() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__22() { _start: { lean_object* x_1; @@ -361,7 +369,7 @@ x_1 = lean_mk_string_from_bytes("skip", 4); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__22() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__23() { _start: { lean_object* x_1; lean_object* x_2; @@ -370,7 +378,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__23() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__24() { _start: { lean_object* x_1; lean_object* x_2; @@ -379,7 +387,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__24() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__25() { _start: { lean_object* x_1; @@ -387,7 +395,7 @@ x_1 = lean_mk_string_from_bytes("=>", 2); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__25() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__26() { _start: { lean_object* x_1; @@ -395,807 +403,828 @@ x_1 = lean_mk_string_from_bytes("match", 5); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__26() { +static lean_object* _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__27() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__25; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__26; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__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, size_t x_9, size_t x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__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, lean_object* x_9, size_t x_10, size_t x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_14; lean_object* x_15; uint8_t x_23; -x_23 = lean_usize_dec_lt(x_10, x_9); -if (x_23 == 0) +lean_object* x_15; lean_object* x_16; uint8_t x_24; +x_24 = lean_usize_dec_lt(x_11, x_10); +if (x_24 == 0) { -lean_object* x_24; -lean_dec(x_12); +lean_object* x_25; +lean_dec(x_13); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_11); -lean_ctor_set(x_24, 1, x_13); -return x_24; +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_12); +lean_ctor_set(x_25, 1, x_14); +return x_25; } 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_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; uint8_t x_40; -x_25 = lean_array_uget(x_8, x_10); -x_26 = lean_ctor_get(x_11, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_11, 1); +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_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; uint8_t x_41; +x_26 = lean_array_uget(x_9, x_11); +x_27 = lean_ctor_get(x_12, 0); lean_inc(x_27); -lean_dec(x_11); -x_28 = lean_ctor_get(x_27, 0); +x_28 = lean_ctor_get(x_12, 1); lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); +lean_dec(x_12); +x_29 = lean_ctor_get(x_28, 0); lean_inc(x_29); -lean_dec(x_27); -x_30 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__1; -x_31 = lean_array_push(x_30, x_25); -x_32 = lean_box(2); -x_33 = l_Lean_nullKind; -x_34 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_34, 0, x_32); -lean_ctor_set(x_34, 1, x_33); -lean_ctor_set(x_34, 2, x_31); -x_35 = lean_unsigned_to_nat(1u); -lean_inc(x_6); -x_36 = l_Lean_Syntax_setArg(x_6, x_35, x_34); -x_37 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__2; -x_38 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__3; -lean_inc(x_5); -x_39 = lean_name_mk_string(x_5, x_38); -lean_inc(x_7); -x_40 = l_Lean_Syntax_isOfKind(x_7, x_39); -if (x_40 == 0) -{ -lean_object* x_41; lean_object* x_42; uint8_t x_43; -x_41 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__4; -lean_inc(x_5); -x_42 = lean_name_mk_string(x_5, x_41); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__1; +x_32 = lean_array_push(x_31, x_26); +x_33 = lean_box(2); +x_34 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__3; +x_35 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +lean_ctor_set(x_35, 2, x_32); +x_36 = lean_unsigned_to_nat(1u); lean_inc(x_7); -x_43 = l_Lean_Syntax_isOfKind(x_7, x_42); -lean_dec(x_42); -if (x_43 == 0) +x_37 = l_Lean_Syntax_setArg(x_7, x_36, x_35); +x_38 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__4; +x_39 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__5; +lean_inc(x_6); +x_40 = lean_name_mk_string(x_6, x_39); +lean_inc(x_8); +x_41 = l_Lean_Syntax_isOfKind(x_8, x_40); +if (x_41 == 0) { -uint8_t x_44; -x_44 = !lean_is_exclusive(x_13); +lean_object* x_42; lean_object* x_43; uint8_t x_44; +x_42 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__6; +lean_inc(x_6); +x_43 = lean_name_mk_string(x_6, x_42); +lean_inc(x_8); +x_44 = l_Lean_Syntax_isOfKind(x_8, x_43); +lean_dec(x_43); if (x_44 == 0) { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; -x_45 = lean_ctor_get(x_13, 0); -x_46 = lean_nat_add(x_45, x_35); -lean_ctor_set(x_13, 0, x_46); -x_47 = lean_ctor_get(x_12, 0); -lean_inc(x_47); -x_48 = lean_ctor_get(x_12, 1); +uint8_t x_45; +x_45 = !lean_is_exclusive(x_14); +if (x_45 == 0) +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; +x_46 = lean_ctor_get(x_14, 0); +x_47 = lean_nat_add(x_46, x_36); +lean_ctor_set(x_14, 0, x_47); +x_48 = lean_ctor_get(x_13, 0); lean_inc(x_48); -x_49 = lean_ctor_get(x_12, 3); +x_49 = lean_ctor_get(x_13, 1); lean_inc(x_49); -x_50 = lean_ctor_get(x_12, 4); +x_50 = lean_ctor_get(x_13, 3); lean_inc(x_50); -x_51 = lean_ctor_get(x_12, 5); +x_51 = lean_ctor_get(x_13, 4); lean_inc(x_51); -lean_inc(x_45); -lean_inc(x_48); -x_52 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_52, 0, x_47); -lean_ctor_set(x_52, 1, x_48); -lean_ctor_set(x_52, 2, x_45); -lean_ctor_set(x_52, 3, x_49); -lean_ctor_set(x_52, 4, x_50); -lean_ctor_set(x_52, 5, x_51); -x_53 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_52, x_13); -x_54 = lean_ctor_get(x_53, 0); -lean_inc(x_54); -x_55 = lean_ctor_get(x_53, 1); +x_52 = lean_ctor_get(x_13, 5); +lean_inc(x_52); +lean_inc(x_46); +lean_inc(x_49); +x_53 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_53, 0, x_48); +lean_ctor_set(x_53, 1, x_49); +lean_ctor_set(x_53, 2, x_46); +lean_ctor_set(x_53, 3, x_50); +lean_ctor_set(x_53, 4, x_51); +lean_ctor_set(x_53, 5, x_52); +x_54 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_53, x_14); +x_55 = lean_ctor_get(x_54, 0); lean_inc(x_55); -lean_dec(x_53); -x_56 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__5; -lean_inc(x_54); -x_57 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_57, 0, x_54); -lean_ctor_set(x_57, 1, x_56); -x_58 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__9; -x_59 = l_Lean_addMacroScope(x_48, x_58, x_45); -x_60 = lean_box(0); -x_61 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__8; -x_62 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_62, 0, x_54); -lean_ctor_set(x_62, 1, x_61); -lean_ctor_set(x_62, 2, x_59); -lean_ctor_set(x_62, 3, x_60); -x_63 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__10; -x_64 = lean_array_push(x_63, x_57); -x_65 = lean_array_push(x_64, x_62); -x_66 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_66, 0, x_32); -lean_ctor_set(x_66, 1, x_39); -lean_ctor_set(x_66, 2, x_65); -x_67 = l_Lean_Syntax_getArg(x_66, x_35); -lean_inc(x_12); -x_68 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_12, x_55); -x_69 = lean_ctor_get(x_68, 0); -lean_inc(x_69); -x_70 = lean_ctor_get(x_68, 1); +x_56 = lean_ctor_get(x_54, 1); +lean_inc(x_56); +lean_dec(x_54); +x_57 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__7; +lean_inc(x_55); +x_58 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_58, 0, x_55); +lean_ctor_set(x_58, 1, x_57); +x_59 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__11; +x_60 = l_Lean_addMacroScope(x_49, x_59, x_46); +x_61 = lean_box(0); +x_62 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__10; +x_63 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_63, 0, x_55); +lean_ctor_set(x_63, 1, x_62); +lean_ctor_set(x_63, 2, x_60); +lean_ctor_set(x_63, 3, x_61); +x_64 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__12; +x_65 = lean_array_push(x_64, x_58); +x_66 = lean_array_push(x_65, x_63); +x_67 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_67, 0, x_33); +lean_ctor_set(x_67, 1, x_40); +lean_ctor_set(x_67, 2, x_66); +x_68 = l_Lean_Syntax_getArg(x_67, x_36); +lean_inc(x_13); +x_69 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_13, x_56); +x_70 = lean_ctor_get(x_69, 0); lean_inc(x_70); -lean_dec(x_68); -x_71 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__11; +x_71 = lean_ctor_get(x_69, 1); +lean_inc(x_71); +lean_dec(x_69); +x_72 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__13; +lean_inc(x_5); +x_73 = lean_name_mk_string(x_5, x_72); +x_74 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__14; +lean_inc(x_73); +x_75 = lean_name_mk_string(x_73, x_74); +lean_inc(x_70); +x_76 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_76, 0, x_70); +lean_ctor_set(x_76, 1, x_74); +x_77 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__15; lean_inc(x_4); -x_72 = lean_name_mk_string(x_4, x_71); -x_73 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__12; -lean_inc(x_72); -x_74 = lean_name_mk_string(x_72, x_73); -lean_inc(x_69); -x_75 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_75, 0, x_69); -lean_ctor_set(x_75, 1, x_73); -x_76 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__14; +x_78 = lean_name_mk_string(x_4, x_77); +x_79 = lean_array_push(x_31, x_68); +x_80 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_80, 0, x_33); +lean_ctor_set(x_80, 1, x_78); +lean_ctor_set(x_80, 2, x_79); lean_inc(x_3); -x_77 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_77, 0, x_32); -lean_ctor_set(x_77, 1, x_76); -lean_ctor_set(x_77, 2, x_3); -x_78 = lean_unsigned_to_nat(2u); -x_79 = l_Lean_Syntax_getArg(x_36, x_78); -x_80 = l_Lean_Syntax_getHeadInfo_x3f(x_79); -x_81 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__15; -lean_inc(x_72); -x_82 = lean_name_mk_string(x_72, x_81); -x_83 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__16; -lean_inc(x_72); -x_84 = lean_name_mk_string(x_72, x_83); -x_85 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__19; -lean_inc(x_72); -x_86 = lean_name_mk_string(x_72, x_85); -x_87 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__20; -lean_inc(x_69); -x_88 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_88, 0, x_69); -lean_ctor_set(x_88, 1, x_87); -x_89 = lean_unsigned_to_nat(0u); -x_90 = l_Lean_Syntax_getArg(x_36, x_89); -x_91 = lean_array_push(x_63, x_90); -x_92 = lean_array_push(x_91, x_79); -x_93 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_93, 0, x_32); -lean_ctor_set(x_93, 1, x_33); -lean_ctor_set(x_93, 2, x_92); -x_94 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__21; -x_95 = lean_name_mk_string(x_72, x_94); -lean_inc(x_69); -x_96 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_96, 0, x_69); -lean_ctor_set(x_96, 1, x_94); -x_97 = lean_array_push(x_30, x_96); -x_98 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_98, 0, x_32); -lean_ctor_set(x_98, 1, x_95); -lean_ctor_set(x_98, 2, x_97); -x_99 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__22; -x_100 = lean_array_push(x_99, x_88); -x_101 = lean_array_push(x_100, x_93); -x_102 = lean_array_push(x_101, x_98); -x_103 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_103, 0, x_32); -lean_ctor_set(x_103, 1, x_86); -lean_ctor_set(x_103, 2, x_102); -x_104 = lean_array_push(x_63, x_103); -lean_inc(x_77); -x_105 = lean_array_push(x_104, x_77); -x_106 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__18; +x_81 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_81, 0, x_33); +lean_ctor_set(x_81, 1, x_34); +lean_ctor_set(x_81, 2, x_3); +x_82 = lean_unsigned_to_nat(2u); +x_83 = l_Lean_Syntax_getArg(x_37, x_82); +x_84 = l_Lean_Syntax_getHeadInfo_x3f(x_83); +x_85 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__16; +lean_inc(x_73); +x_86 = lean_name_mk_string(x_73, x_85); +x_87 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__17; +lean_inc(x_73); +x_88 = lean_name_mk_string(x_73, x_87); +x_89 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__20; +lean_inc(x_73); +x_90 = lean_name_mk_string(x_73, x_89); +x_91 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__21; +lean_inc(x_70); +x_92 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_92, 0, x_70); +lean_ctor_set(x_92, 1, x_91); +x_93 = lean_unsigned_to_nat(0u); +x_94 = l_Lean_Syntax_getArg(x_37, x_93); +x_95 = lean_array_push(x_64, x_94); +x_96 = lean_array_push(x_95, x_83); +x_97 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_97, 0, x_33); +lean_ctor_set(x_97, 1, x_34); +lean_ctor_set(x_97, 2, x_96); +x_98 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__22; +x_99 = lean_name_mk_string(x_73, x_98); +lean_inc(x_70); +x_100 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_100, 0, x_70); +lean_ctor_set(x_100, 1, x_98); +x_101 = lean_array_push(x_31, x_100); +x_102 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_102, 0, x_33); +lean_ctor_set(x_102, 1, x_99); +lean_ctor_set(x_102, 2, x_101); +x_103 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__23; +x_104 = lean_array_push(x_103, x_92); +x_105 = lean_array_push(x_104, x_97); +x_106 = lean_array_push(x_105, x_102); x_107 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_107, 0, x_32); -lean_ctor_set(x_107, 1, x_106); -lean_ctor_set(x_107, 2, x_105); -lean_inc(x_7); -x_108 = lean_array_push(x_63, x_7); -lean_inc(x_77); -x_109 = lean_array_push(x_108, x_77); -x_110 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_110, 0, x_32); -lean_ctor_set(x_110, 1, x_106); -lean_ctor_set(x_110, 2, x_109); -x_111 = lean_array_push(x_63, x_107); -x_112 = lean_array_push(x_111, x_110); -x_113 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_113, 0, x_32); -lean_ctor_set(x_113, 1, x_76); -lean_ctor_set(x_113, 2, x_112); -x_114 = lean_array_push(x_30, x_113); -x_115 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_115, 0, x_32); -lean_ctor_set(x_115, 1, x_84); -lean_ctor_set(x_115, 2, x_114); -x_116 = lean_array_push(x_30, x_115); +lean_ctor_set(x_107, 0, x_33); +lean_ctor_set(x_107, 1, x_90); +lean_ctor_set(x_107, 2, x_106); +x_108 = lean_array_push(x_64, x_107); +lean_inc(x_81); +x_109 = lean_array_push(x_108, x_81); +x_110 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__19; +x_111 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_111, 0, x_33); +lean_ctor_set(x_111, 1, x_110); +lean_ctor_set(x_111, 2, x_109); +lean_inc(x_8); +x_112 = lean_array_push(x_64, x_8); +lean_inc(x_81); +x_113 = lean_array_push(x_112, x_81); +x_114 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_114, 0, x_33); +lean_ctor_set(x_114, 1, x_110); +lean_ctor_set(x_114, 2, x_113); +x_115 = lean_array_push(x_64, x_111); +x_116 = lean_array_push(x_115, x_114); x_117 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_117, 0, x_32); -lean_ctor_set(x_117, 1, x_82); +lean_ctor_set(x_117, 0, x_33); +lean_ctor_set(x_117, 1, x_34); lean_ctor_set(x_117, 2, x_116); -x_118 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__23; -x_119 = lean_array_push(x_118, x_75); -x_120 = lean_array_push(x_119, x_67); -x_121 = lean_array_push(x_120, x_77); -if (lean_obj_tag(x_80) == 0) -{ -lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; +x_118 = lean_array_push(x_31, x_117); +x_119 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_119, 0, x_33); +lean_ctor_set(x_119, 1, x_88); +lean_ctor_set(x_119, 2, x_118); +x_120 = lean_array_push(x_31, x_119); +x_121 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_121, 0, x_33); +lean_ctor_set(x_121, 1, x_86); +lean_ctor_set(x_121, 2, x_120); x_122 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__24; -x_123 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_123, 0, x_69); -lean_ctor_set(x_123, 1, x_122); -x_124 = lean_array_push(x_121, x_123); -x_125 = lean_array_push(x_124, x_117); -x_126 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_126, 0, x_32); -lean_ctor_set(x_126, 1, x_74); -lean_ctor_set(x_126, 2, x_125); -x_127 = lean_array_push(x_28, x_126); -x_128 = lean_unsigned_to_nat(3u); -x_129 = l_Lean_Syntax_setArg(x_36, x_128, x_66); -x_130 = lean_box(0); -lean_inc(x_12); -x_131 = lean_apply_7(x_37, x_129, x_26, x_127, x_29, x_130, x_12, x_70); -if (lean_obj_tag(x_131) == 0) -{ -lean_object* x_132; lean_object* x_133; -x_132 = lean_ctor_get(x_131, 0); -lean_inc(x_132); -x_133 = lean_ctor_get(x_131, 1); -lean_inc(x_133); -lean_dec(x_131); -x_14 = x_132; -x_15 = x_133; -goto block_22; +x_123 = lean_array_push(x_122, x_76); +x_124 = lean_array_push(x_123, x_80); +x_125 = lean_array_push(x_124, x_81); +if (lean_obj_tag(x_84) == 0) +{ +lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; +x_126 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__25; +x_127 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_127, 0, x_70); +lean_ctor_set(x_127, 1, x_126); +x_128 = lean_array_push(x_125, x_127); +x_129 = lean_array_push(x_128, x_121); +x_130 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_130, 0, x_33); +lean_ctor_set(x_130, 1, x_75); +lean_ctor_set(x_130, 2, x_129); +x_131 = lean_array_push(x_29, x_130); +x_132 = lean_unsigned_to_nat(3u); +x_133 = l_Lean_Syntax_setArg(x_37, x_132, x_67); +x_134 = lean_box(0); +lean_inc(x_13); +x_135 = lean_apply_7(x_38, x_133, x_27, x_131, x_30, x_134, x_13, x_71); +if (lean_obj_tag(x_135) == 0) +{ +lean_object* x_136; lean_object* x_137; +x_136 = lean_ctor_get(x_135, 0); +lean_inc(x_136); +x_137 = lean_ctor_get(x_135, 1); +lean_inc(x_137); +lean_dec(x_135); +x_15 = x_136; +x_16 = x_137; +goto block_23; } else { -uint8_t x_134; -lean_dec(x_12); +uint8_t x_138; +lean_dec(x_13); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_134 = !lean_is_exclusive(x_131); -if (x_134 == 0) +x_138 = !lean_is_exclusive(x_135); +if (x_138 == 0) { -return x_131; +return x_135; } else { -lean_object* x_135; lean_object* x_136; lean_object* x_137; -x_135 = lean_ctor_get(x_131, 0); -x_136 = lean_ctor_get(x_131, 1); -lean_inc(x_136); -lean_inc(x_135); -lean_dec(x_131); -x_137 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_137, 0, x_135); -lean_ctor_set(x_137, 1, x_136); -return x_137; +lean_object* x_139; lean_object* x_140; lean_object* x_141; +x_139 = lean_ctor_get(x_135, 0); +x_140 = lean_ctor_get(x_135, 1); +lean_inc(x_140); +lean_inc(x_139); +lean_dec(x_135); +x_141 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_141, 0, x_139); +lean_ctor_set(x_141, 1, x_140); +return x_141; } } } else { -lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; -lean_dec(x_69); -x_138 = lean_ctor_get(x_80, 0); -lean_inc(x_138); -lean_dec(x_80); -x_139 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__24; -x_140 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_140, 0, x_138); -lean_ctor_set(x_140, 1, x_139); -x_141 = lean_array_push(x_121, x_140); -x_142 = lean_array_push(x_141, x_117); -x_143 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_143, 0, x_32); -lean_ctor_set(x_143, 1, x_74); -lean_ctor_set(x_143, 2, x_142); -x_144 = lean_array_push(x_28, x_143); -x_145 = lean_unsigned_to_nat(3u); -x_146 = l_Lean_Syntax_setArg(x_36, x_145, x_66); -x_147 = lean_box(0); -lean_inc(x_12); -x_148 = lean_apply_7(x_37, x_146, x_26, x_144, x_29, x_147, x_12, x_70); -if (lean_obj_tag(x_148) == 0) -{ -lean_object* x_149; lean_object* x_150; -x_149 = lean_ctor_get(x_148, 0); -lean_inc(x_149); -x_150 = lean_ctor_get(x_148, 1); -lean_inc(x_150); -lean_dec(x_148); -x_14 = x_149; -x_15 = x_150; -goto block_22; +lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; +lean_dec(x_70); +x_142 = lean_ctor_get(x_84, 0); +lean_inc(x_142); +lean_dec(x_84); +x_143 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__25; +x_144 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_144, 0, x_142); +lean_ctor_set(x_144, 1, x_143); +x_145 = lean_array_push(x_125, x_144); +x_146 = lean_array_push(x_145, x_121); +x_147 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_147, 0, x_33); +lean_ctor_set(x_147, 1, x_75); +lean_ctor_set(x_147, 2, x_146); +x_148 = lean_array_push(x_29, x_147); +x_149 = lean_unsigned_to_nat(3u); +x_150 = l_Lean_Syntax_setArg(x_37, x_149, x_67); +x_151 = lean_box(0); +lean_inc(x_13); +x_152 = lean_apply_7(x_38, x_150, x_27, x_148, x_30, x_151, x_13, x_71); +if (lean_obj_tag(x_152) == 0) +{ +lean_object* x_153; lean_object* x_154; +x_153 = lean_ctor_get(x_152, 0); +lean_inc(x_153); +x_154 = lean_ctor_get(x_152, 1); +lean_inc(x_154); +lean_dec(x_152); +x_15 = x_153; +x_16 = x_154; +goto block_23; } else { -uint8_t x_151; -lean_dec(x_12); +uint8_t x_155; +lean_dec(x_13); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_151 = !lean_is_exclusive(x_148); -if (x_151 == 0) +x_155 = !lean_is_exclusive(x_152); +if (x_155 == 0) { -return x_148; +return x_152; } else { -lean_object* x_152; lean_object* x_153; lean_object* x_154; -x_152 = lean_ctor_get(x_148, 0); -x_153 = lean_ctor_get(x_148, 1); -lean_inc(x_153); -lean_inc(x_152); -lean_dec(x_148); -x_154 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_154, 0, x_152); -lean_ctor_set(x_154, 1, x_153); -return x_154; +lean_object* x_156; lean_object* x_157; lean_object* x_158; +x_156 = lean_ctor_get(x_152, 0); +x_157 = lean_ctor_get(x_152, 1); +lean_inc(x_157); +lean_inc(x_156); +lean_dec(x_152); +x_158 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_158, 0, x_156); +lean_ctor_set(x_158, 1, x_157); +return x_158; } } } } else { -lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; -x_155 = lean_ctor_get(x_13, 0); -x_156 = lean_ctor_get(x_13, 1); -lean_inc(x_156); -lean_inc(x_155); -lean_dec(x_13); -x_157 = lean_nat_add(x_155, x_35); -x_158 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_158, 0, x_157); -lean_ctor_set(x_158, 1, x_156); -x_159 = lean_ctor_get(x_12, 0); -lean_inc(x_159); -x_160 = lean_ctor_get(x_12, 1); +lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; +x_159 = lean_ctor_get(x_14, 0); +x_160 = lean_ctor_get(x_14, 1); lean_inc(x_160); -x_161 = lean_ctor_get(x_12, 3); -lean_inc(x_161); -x_162 = lean_ctor_get(x_12, 4); -lean_inc(x_162); -x_163 = lean_ctor_get(x_12, 5); +lean_inc(x_159); +lean_dec(x_14); +x_161 = lean_nat_add(x_159, x_36); +x_162 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_162, 0, x_161); +lean_ctor_set(x_162, 1, x_160); +x_163 = lean_ctor_get(x_13, 0); lean_inc(x_163); -lean_inc(x_155); -lean_inc(x_160); -x_164 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_164, 0, x_159); -lean_ctor_set(x_164, 1, x_160); -lean_ctor_set(x_164, 2, x_155); -lean_ctor_set(x_164, 3, x_161); -lean_ctor_set(x_164, 4, x_162); -lean_ctor_set(x_164, 5, x_163); -x_165 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_164, x_158); -x_166 = lean_ctor_get(x_165, 0); +x_164 = lean_ctor_get(x_13, 1); +lean_inc(x_164); +x_165 = lean_ctor_get(x_13, 3); +lean_inc(x_165); +x_166 = lean_ctor_get(x_13, 4); lean_inc(x_166); -x_167 = lean_ctor_get(x_165, 1); +x_167 = lean_ctor_get(x_13, 5); lean_inc(x_167); -lean_dec(x_165); -x_168 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__5; -lean_inc(x_166); -x_169 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_169, 0, x_166); -lean_ctor_set(x_169, 1, x_168); -x_170 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__9; -x_171 = l_Lean_addMacroScope(x_160, x_170, x_155); -x_172 = lean_box(0); -x_173 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__8; -x_174 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_174, 0, x_166); -lean_ctor_set(x_174, 1, x_173); -lean_ctor_set(x_174, 2, x_171); -lean_ctor_set(x_174, 3, x_172); -x_175 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__10; -x_176 = lean_array_push(x_175, x_169); -x_177 = lean_array_push(x_176, x_174); -x_178 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_178, 0, x_32); -lean_ctor_set(x_178, 1, x_39); -lean_ctor_set(x_178, 2, x_177); -x_179 = l_Lean_Syntax_getArg(x_178, x_35); -lean_inc(x_12); -x_180 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_12, x_167); -x_181 = lean_ctor_get(x_180, 0); -lean_inc(x_181); -x_182 = lean_ctor_get(x_180, 1); -lean_inc(x_182); -lean_dec(x_180); -x_183 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__11; +lean_inc(x_159); +lean_inc(x_164); +x_168 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_168, 0, x_163); +lean_ctor_set(x_168, 1, x_164); +lean_ctor_set(x_168, 2, x_159); +lean_ctor_set(x_168, 3, x_165); +lean_ctor_set(x_168, 4, x_166); +lean_ctor_set(x_168, 5, x_167); +x_169 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_168, x_162); +x_170 = lean_ctor_get(x_169, 0); +lean_inc(x_170); +x_171 = lean_ctor_get(x_169, 1); +lean_inc(x_171); +lean_dec(x_169); +x_172 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__7; +lean_inc(x_170); +x_173 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_173, 0, x_170); +lean_ctor_set(x_173, 1, x_172); +x_174 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__11; +x_175 = l_Lean_addMacroScope(x_164, x_174, x_159); +x_176 = lean_box(0); +x_177 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__10; +x_178 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_178, 0, x_170); +lean_ctor_set(x_178, 1, x_177); +lean_ctor_set(x_178, 2, x_175); +lean_ctor_set(x_178, 3, x_176); +x_179 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__12; +x_180 = lean_array_push(x_179, x_173); +x_181 = lean_array_push(x_180, x_178); +x_182 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_182, 0, x_33); +lean_ctor_set(x_182, 1, x_40); +lean_ctor_set(x_182, 2, x_181); +x_183 = l_Lean_Syntax_getArg(x_182, x_36); +lean_inc(x_13); +x_184 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_13, x_171); +x_185 = lean_ctor_get(x_184, 0); +lean_inc(x_185); +x_186 = lean_ctor_get(x_184, 1); +lean_inc(x_186); +lean_dec(x_184); +x_187 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__13; +lean_inc(x_5); +x_188 = lean_name_mk_string(x_5, x_187); +x_189 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__14; +lean_inc(x_188); +x_190 = lean_name_mk_string(x_188, x_189); +lean_inc(x_185); +x_191 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_191, 0, x_185); +lean_ctor_set(x_191, 1, x_189); +x_192 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__15; lean_inc(x_4); -x_184 = lean_name_mk_string(x_4, x_183); -x_185 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__12; -lean_inc(x_184); -x_186 = lean_name_mk_string(x_184, x_185); -lean_inc(x_181); -x_187 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_187, 0, x_181); -lean_ctor_set(x_187, 1, x_185); -x_188 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__14; +x_193 = lean_name_mk_string(x_4, x_192); +x_194 = lean_array_push(x_31, x_183); +x_195 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_195, 0, x_33); +lean_ctor_set(x_195, 1, x_193); +lean_ctor_set(x_195, 2, x_194); lean_inc(x_3); -x_189 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_189, 0, x_32); -lean_ctor_set(x_189, 1, x_188); -lean_ctor_set(x_189, 2, x_3); -x_190 = lean_unsigned_to_nat(2u); -x_191 = l_Lean_Syntax_getArg(x_36, x_190); -x_192 = l_Lean_Syntax_getHeadInfo_x3f(x_191); -x_193 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__15; -lean_inc(x_184); -x_194 = lean_name_mk_string(x_184, x_193); -x_195 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__16; -lean_inc(x_184); -x_196 = lean_name_mk_string(x_184, x_195); -x_197 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__19; -lean_inc(x_184); -x_198 = lean_name_mk_string(x_184, x_197); -x_199 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__20; -lean_inc(x_181); -x_200 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_200, 0, x_181); -lean_ctor_set(x_200, 1, x_199); -x_201 = lean_unsigned_to_nat(0u); -x_202 = l_Lean_Syntax_getArg(x_36, x_201); -x_203 = lean_array_push(x_175, x_202); -x_204 = lean_array_push(x_203, x_191); -x_205 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_205, 0, x_32); -lean_ctor_set(x_205, 1, x_33); -lean_ctor_set(x_205, 2, x_204); +x_196 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_196, 0, x_33); +lean_ctor_set(x_196, 1, x_34); +lean_ctor_set(x_196, 2, x_3); +x_197 = lean_unsigned_to_nat(2u); +x_198 = l_Lean_Syntax_getArg(x_37, x_197); +x_199 = l_Lean_Syntax_getHeadInfo_x3f(x_198); +x_200 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__16; +lean_inc(x_188); +x_201 = lean_name_mk_string(x_188, x_200); +x_202 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__17; +lean_inc(x_188); +x_203 = lean_name_mk_string(x_188, x_202); +x_204 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__20; +lean_inc(x_188); +x_205 = lean_name_mk_string(x_188, x_204); x_206 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__21; -x_207 = lean_name_mk_string(x_184, x_206); -lean_inc(x_181); -x_208 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_208, 0, x_181); -lean_ctor_set(x_208, 1, x_206); -x_209 = lean_array_push(x_30, x_208); -x_210 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_210, 0, x_32); -lean_ctor_set(x_210, 1, x_207); -lean_ctor_set(x_210, 2, x_209); -x_211 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__22; -x_212 = lean_array_push(x_211, x_200); -x_213 = lean_array_push(x_212, x_205); -x_214 = lean_array_push(x_213, x_210); -x_215 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_215, 0, x_32); -lean_ctor_set(x_215, 1, x_198); -lean_ctor_set(x_215, 2, x_214); -x_216 = lean_array_push(x_175, x_215); -lean_inc(x_189); -x_217 = lean_array_push(x_216, x_189); -x_218 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__18; -x_219 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_219, 0, x_32); -lean_ctor_set(x_219, 1, x_218); -lean_ctor_set(x_219, 2, x_217); -lean_inc(x_7); -x_220 = lean_array_push(x_175, x_7); -lean_inc(x_189); -x_221 = lean_array_push(x_220, x_189); +lean_inc(x_185); +x_207 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_207, 0, x_185); +lean_ctor_set(x_207, 1, x_206); +x_208 = lean_unsigned_to_nat(0u); +x_209 = l_Lean_Syntax_getArg(x_37, x_208); +x_210 = lean_array_push(x_179, x_209); +x_211 = lean_array_push(x_210, x_198); +x_212 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_212, 0, x_33); +lean_ctor_set(x_212, 1, x_34); +lean_ctor_set(x_212, 2, x_211); +x_213 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__22; +x_214 = lean_name_mk_string(x_188, x_213); +lean_inc(x_185); +x_215 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_215, 0, x_185); +lean_ctor_set(x_215, 1, x_213); +x_216 = lean_array_push(x_31, x_215); +x_217 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_217, 0, x_33); +lean_ctor_set(x_217, 1, x_214); +lean_ctor_set(x_217, 2, x_216); +x_218 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__23; +x_219 = lean_array_push(x_218, x_207); +x_220 = lean_array_push(x_219, x_212); +x_221 = lean_array_push(x_220, x_217); x_222 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_222, 0, x_32); -lean_ctor_set(x_222, 1, x_218); +lean_ctor_set(x_222, 0, x_33); +lean_ctor_set(x_222, 1, x_205); lean_ctor_set(x_222, 2, x_221); -x_223 = lean_array_push(x_175, x_219); -x_224 = lean_array_push(x_223, x_222); -x_225 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_225, 0, x_32); -lean_ctor_set(x_225, 1, x_188); -lean_ctor_set(x_225, 2, x_224); -x_226 = lean_array_push(x_30, x_225); -x_227 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_227, 0, x_32); -lean_ctor_set(x_227, 1, x_196); -lean_ctor_set(x_227, 2, x_226); -x_228 = lean_array_push(x_30, x_227); +x_223 = lean_array_push(x_179, x_222); +lean_inc(x_196); +x_224 = lean_array_push(x_223, x_196); +x_225 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__19; +x_226 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_226, 0, x_33); +lean_ctor_set(x_226, 1, x_225); +lean_ctor_set(x_226, 2, x_224); +lean_inc(x_8); +x_227 = lean_array_push(x_179, x_8); +lean_inc(x_196); +x_228 = lean_array_push(x_227, x_196); x_229 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_229, 0, x_32); -lean_ctor_set(x_229, 1, x_194); +lean_ctor_set(x_229, 0, x_33); +lean_ctor_set(x_229, 1, x_225); lean_ctor_set(x_229, 2, x_228); -x_230 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__23; -x_231 = lean_array_push(x_230, x_187); -x_232 = lean_array_push(x_231, x_179); -x_233 = lean_array_push(x_232, x_189); -if (lean_obj_tag(x_192) == 0) -{ -lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; -x_234 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__24; -x_235 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_235, 0, x_181); -lean_ctor_set(x_235, 1, x_234); -x_236 = lean_array_push(x_233, x_235); -x_237 = lean_array_push(x_236, x_229); -x_238 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_238, 0, x_32); -lean_ctor_set(x_238, 1, x_186); -lean_ctor_set(x_238, 2, x_237); -x_239 = lean_array_push(x_28, x_238); -x_240 = lean_unsigned_to_nat(3u); -x_241 = l_Lean_Syntax_setArg(x_36, x_240, x_178); -x_242 = lean_box(0); -lean_inc(x_12); -x_243 = lean_apply_7(x_37, x_241, x_26, x_239, x_29, x_242, x_12, x_182); -if (lean_obj_tag(x_243) == 0) -{ -lean_object* x_244; lean_object* x_245; -x_244 = lean_ctor_get(x_243, 0); -lean_inc(x_244); -x_245 = lean_ctor_get(x_243, 1); -lean_inc(x_245); -lean_dec(x_243); -x_14 = x_244; -x_15 = x_245; -goto block_22; +x_230 = lean_array_push(x_179, x_226); +x_231 = lean_array_push(x_230, x_229); +x_232 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_232, 0, x_33); +lean_ctor_set(x_232, 1, x_34); +lean_ctor_set(x_232, 2, x_231); +x_233 = lean_array_push(x_31, x_232); +x_234 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_234, 0, x_33); +lean_ctor_set(x_234, 1, x_203); +lean_ctor_set(x_234, 2, x_233); +x_235 = lean_array_push(x_31, x_234); +x_236 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_236, 0, x_33); +lean_ctor_set(x_236, 1, x_201); +lean_ctor_set(x_236, 2, x_235); +x_237 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__24; +x_238 = lean_array_push(x_237, x_191); +x_239 = lean_array_push(x_238, x_195); +x_240 = lean_array_push(x_239, x_196); +if (lean_obj_tag(x_199) == 0) +{ +lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; +x_241 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__25; +x_242 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_242, 0, x_185); +lean_ctor_set(x_242, 1, x_241); +x_243 = lean_array_push(x_240, x_242); +x_244 = lean_array_push(x_243, x_236); +x_245 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_245, 0, x_33); +lean_ctor_set(x_245, 1, x_190); +lean_ctor_set(x_245, 2, x_244); +x_246 = lean_array_push(x_29, x_245); +x_247 = lean_unsigned_to_nat(3u); +x_248 = l_Lean_Syntax_setArg(x_37, x_247, x_182); +x_249 = lean_box(0); +lean_inc(x_13); +x_250 = lean_apply_7(x_38, x_248, x_27, x_246, x_30, x_249, x_13, x_186); +if (lean_obj_tag(x_250) == 0) +{ +lean_object* x_251; lean_object* x_252; +x_251 = lean_ctor_get(x_250, 0); +lean_inc(x_251); +x_252 = lean_ctor_get(x_250, 1); +lean_inc(x_252); +lean_dec(x_250); +x_15 = x_251; +x_16 = x_252; +goto block_23; } else { -lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; -lean_dec(x_12); +lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; +lean_dec(x_13); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_246 = lean_ctor_get(x_243, 0); -lean_inc(x_246); -x_247 = lean_ctor_get(x_243, 1); -lean_inc(x_247); -if (lean_is_exclusive(x_243)) { - lean_ctor_release(x_243, 0); - lean_ctor_release(x_243, 1); - x_248 = x_243; +x_253 = lean_ctor_get(x_250, 0); +lean_inc(x_253); +x_254 = lean_ctor_get(x_250, 1); +lean_inc(x_254); +if (lean_is_exclusive(x_250)) { + lean_ctor_release(x_250, 0); + lean_ctor_release(x_250, 1); + x_255 = x_250; } else { - lean_dec_ref(x_243); - x_248 = lean_box(0); + lean_dec_ref(x_250); + x_255 = lean_box(0); } -if (lean_is_scalar(x_248)) { - x_249 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_255)) { + x_256 = lean_alloc_ctor(1, 2, 0); } else { - x_249 = x_248; + x_256 = x_255; } -lean_ctor_set(x_249, 0, x_246); -lean_ctor_set(x_249, 1, x_247); -return x_249; +lean_ctor_set(x_256, 0, x_253); +lean_ctor_set(x_256, 1, x_254); +return x_256; } } else { -lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; -lean_dec(x_181); -x_250 = lean_ctor_get(x_192, 0); -lean_inc(x_250); -lean_dec(x_192); -x_251 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__24; -x_252 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_252, 0, x_250); -lean_ctor_set(x_252, 1, x_251); -x_253 = lean_array_push(x_233, x_252); -x_254 = lean_array_push(x_253, x_229); -x_255 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_255, 0, x_32); -lean_ctor_set(x_255, 1, x_186); -lean_ctor_set(x_255, 2, x_254); -x_256 = lean_array_push(x_28, x_255); -x_257 = lean_unsigned_to_nat(3u); -x_258 = l_Lean_Syntax_setArg(x_36, x_257, x_178); -x_259 = lean_box(0); -lean_inc(x_12); -x_260 = lean_apply_7(x_37, x_258, x_26, x_256, x_29, x_259, x_12, x_182); -if (lean_obj_tag(x_260) == 0) -{ -lean_object* x_261; lean_object* x_262; -x_261 = lean_ctor_get(x_260, 0); -lean_inc(x_261); -x_262 = lean_ctor_get(x_260, 1); -lean_inc(x_262); -lean_dec(x_260); -x_14 = x_261; -x_15 = x_262; -goto block_22; +lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; +lean_dec(x_185); +x_257 = lean_ctor_get(x_199, 0); +lean_inc(x_257); +lean_dec(x_199); +x_258 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__25; +x_259 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_259, 0, x_257); +lean_ctor_set(x_259, 1, x_258); +x_260 = lean_array_push(x_240, x_259); +x_261 = lean_array_push(x_260, x_236); +x_262 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_262, 0, x_33); +lean_ctor_set(x_262, 1, x_190); +lean_ctor_set(x_262, 2, x_261); +x_263 = lean_array_push(x_29, x_262); +x_264 = lean_unsigned_to_nat(3u); +x_265 = l_Lean_Syntax_setArg(x_37, x_264, x_182); +x_266 = lean_box(0); +lean_inc(x_13); +x_267 = lean_apply_7(x_38, x_265, x_27, x_263, x_30, x_266, x_13, x_186); +if (lean_obj_tag(x_267) == 0) +{ +lean_object* x_268; lean_object* x_269; +x_268 = lean_ctor_get(x_267, 0); +lean_inc(x_268); +x_269 = lean_ctor_get(x_267, 1); +lean_inc(x_269); +lean_dec(x_267); +x_15 = x_268; +x_16 = x_269; +goto block_23; } else { -lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; -lean_dec(x_12); +lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; +lean_dec(x_13); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_263 = lean_ctor_get(x_260, 0); -lean_inc(x_263); -x_264 = lean_ctor_get(x_260, 1); -lean_inc(x_264); -if (lean_is_exclusive(x_260)) { - lean_ctor_release(x_260, 0); - lean_ctor_release(x_260, 1); - x_265 = x_260; +x_270 = lean_ctor_get(x_267, 0); +lean_inc(x_270); +x_271 = lean_ctor_get(x_267, 1); +lean_inc(x_271); +if (lean_is_exclusive(x_267)) { + lean_ctor_release(x_267, 0); + lean_ctor_release(x_267, 1); + x_272 = x_267; } else { - lean_dec_ref(x_260); - x_265 = lean_box(0); + lean_dec_ref(x_267); + x_272 = lean_box(0); } -if (lean_is_scalar(x_265)) { - x_266 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_272)) { + x_273 = lean_alloc_ctor(1, 2, 0); } else { - x_266 = x_265; + x_273 = x_272; } -lean_ctor_set(x_266, 0, x_263); -lean_ctor_set(x_266, 1, x_264); -return x_266; +lean_ctor_set(x_273, 0, x_270); +lean_ctor_set(x_273, 1, x_271); +return x_273; } } } } else { -lean_object* x_267; uint8_t x_268; lean_object* x_269; -x_267 = lean_array_get_size(x_2); -x_268 = lean_nat_dec_lt(x_35, x_267); -lean_dec(x_267); -lean_inc(x_12); -x_269 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_12, x_13); -if (x_268 == 0) -{ -lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; -x_270 = lean_ctor_get(x_269, 0); -lean_inc(x_270); -x_271 = lean_ctor_get(x_269, 1); -lean_inc(x_271); -lean_dec(x_269); +lean_object* x_274; uint8_t x_275; lean_object* x_276; +x_274 = lean_array_get_size(x_2); +x_275 = lean_nat_dec_lt(x_36, x_274); +lean_dec(x_274); +lean_inc(x_13); +x_276 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_13, x_14); +if (x_275 == 0) +{ +lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; +x_277 = lean_ctor_get(x_276, 0); +lean_inc(x_277); +x_278 = lean_ctor_get(x_276, 1); +lean_inc(x_278); +lean_dec(x_276); lean_inc(x_1); -lean_inc(x_7); -x_272 = l_Lean_mkIdentFrom(x_7, x_1); -x_273 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__5; -x_274 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_274, 0, x_270); -lean_ctor_set(x_274, 1, x_273); -x_275 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__10; -x_276 = lean_array_push(x_275, x_274); -x_277 = lean_array_push(x_276, x_272); -x_278 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_278, 0, x_32); -lean_ctor_set(x_278, 1, x_39); -lean_ctor_set(x_278, 2, x_277); -x_279 = lean_nat_add(x_29, x_35); -lean_dec(x_29); -x_280 = lean_unsigned_to_nat(3u); -x_281 = l_Lean_Syntax_setArg(x_36, x_280, x_278); -x_282 = lean_box(0); -lean_inc(x_12); -x_283 = lean_apply_7(x_37, x_281, x_26, x_28, x_279, x_282, x_12, x_271); -if (lean_obj_tag(x_283) == 0) -{ -lean_object* x_284; lean_object* x_285; -x_284 = lean_ctor_get(x_283, 0); -lean_inc(x_284); -x_285 = lean_ctor_get(x_283, 1); -lean_inc(x_285); -lean_dec(x_283); -x_14 = x_284; -x_15 = x_285; -goto block_22; +lean_inc(x_8); +x_279 = l_Lean_mkIdentFrom(x_8, x_1); +x_280 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__7; +x_281 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_281, 0, x_277); +lean_ctor_set(x_281, 1, x_280); +x_282 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__12; +x_283 = lean_array_push(x_282, x_281); +x_284 = lean_array_push(x_283, x_279); +x_285 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_285, 0, x_33); +lean_ctor_set(x_285, 1, x_40); +lean_ctor_set(x_285, 2, x_284); +x_286 = lean_nat_add(x_30, x_36); +lean_dec(x_30); +x_287 = lean_unsigned_to_nat(3u); +x_288 = l_Lean_Syntax_setArg(x_37, x_287, x_285); +x_289 = lean_box(0); +lean_inc(x_13); +x_290 = lean_apply_7(x_38, x_288, x_27, x_29, x_286, x_289, x_13, x_278); +if (lean_obj_tag(x_290) == 0) +{ +lean_object* x_291; lean_object* x_292; +x_291 = lean_ctor_get(x_290, 0); +lean_inc(x_291); +x_292 = lean_ctor_get(x_290, 1); +lean_inc(x_292); +lean_dec(x_290); +x_15 = x_291; +x_16 = x_292; +goto block_23; } else { -uint8_t x_286; -lean_dec(x_12); +uint8_t x_293; +lean_dec(x_13); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_286 = !lean_is_exclusive(x_283); -if (x_286 == 0) +x_293 = !lean_is_exclusive(x_290); +if (x_293 == 0) { -return x_283; +return x_290; } else { -lean_object* x_287; lean_object* x_288; lean_object* x_289; -x_287 = lean_ctor_get(x_283, 0); -x_288 = lean_ctor_get(x_283, 1); -lean_inc(x_288); -lean_inc(x_287); -lean_dec(x_283); -x_289 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_289, 0, x_287); -lean_ctor_set(x_289, 1, x_288); -return x_289; +lean_object* x_294; lean_object* x_295; lean_object* x_296; +x_294 = lean_ctor_get(x_290, 0); +x_295 = lean_ctor_get(x_290, 1); +lean_inc(x_295); +lean_inc(x_294); +lean_dec(x_290); +x_296 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_296, 0, x_294); +lean_ctor_set(x_296, 1, x_295); +return x_296; } } } else { -lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; -x_290 = lean_ctor_get(x_269, 0); -lean_inc(x_290); -x_291 = lean_ctor_get(x_269, 1); -lean_inc(x_291); -lean_dec(x_269); -x_292 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__26; -lean_inc(x_29); -x_293 = lean_name_append_index_after(x_292, x_29); -x_294 = l_Lean_Name_append(x_1, x_293); -lean_inc(x_7); -x_295 = l_Lean_mkIdentFrom(x_7, x_294); -x_296 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__5; -x_297 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_297, 0, x_290); -lean_ctor_set(x_297, 1, x_296); -x_298 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__10; -x_299 = lean_array_push(x_298, x_297); -x_300 = lean_array_push(x_299, x_295); -x_301 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_301, 0, x_32); -lean_ctor_set(x_301, 1, x_39); -lean_ctor_set(x_301, 2, x_300); -x_302 = lean_nat_add(x_29, x_35); -lean_dec(x_29); -x_303 = lean_unsigned_to_nat(3u); -x_304 = l_Lean_Syntax_setArg(x_36, x_303, x_301); -x_305 = lean_box(0); -lean_inc(x_12); -x_306 = lean_apply_7(x_37, x_304, x_26, x_28, x_302, x_305, x_12, x_291); -if (lean_obj_tag(x_306) == 0) -{ -lean_object* x_307; lean_object* x_308; -x_307 = lean_ctor_get(x_306, 0); -lean_inc(x_307); -x_308 = lean_ctor_get(x_306, 1); -lean_inc(x_308); -lean_dec(x_306); -x_14 = x_307; -x_15 = x_308; -goto block_22; +lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; +x_297 = lean_ctor_get(x_276, 0); +lean_inc(x_297); +x_298 = lean_ctor_get(x_276, 1); +lean_inc(x_298); +lean_dec(x_276); +x_299 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__27; +lean_inc(x_30); +x_300 = lean_name_append_index_after(x_299, x_30); +x_301 = l_Lean_Name_append(x_1, x_300); +lean_inc(x_8); +x_302 = l_Lean_mkIdentFrom(x_8, x_301); +x_303 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__7; +x_304 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_304, 0, x_297); +lean_ctor_set(x_304, 1, x_303); +x_305 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__12; +x_306 = lean_array_push(x_305, x_304); +x_307 = lean_array_push(x_306, x_302); +x_308 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_308, 0, x_33); +lean_ctor_set(x_308, 1, x_40); +lean_ctor_set(x_308, 2, x_307); +x_309 = lean_nat_add(x_30, x_36); +lean_dec(x_30); +x_310 = lean_unsigned_to_nat(3u); +x_311 = l_Lean_Syntax_setArg(x_37, x_310, x_308); +x_312 = lean_box(0); +lean_inc(x_13); +x_313 = lean_apply_7(x_38, x_311, x_27, x_29, x_309, x_312, x_13, x_298); +if (lean_obj_tag(x_313) == 0) +{ +lean_object* x_314; lean_object* x_315; +x_314 = lean_ctor_get(x_313, 0); +lean_inc(x_314); +x_315 = lean_ctor_get(x_313, 1); +lean_inc(x_315); +lean_dec(x_313); +x_15 = x_314; +x_16 = x_315; +goto block_23; } else { -uint8_t x_309; -lean_dec(x_12); +uint8_t x_316; +lean_dec(x_13); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_309 = !lean_is_exclusive(x_306); -if (x_309 == 0) +x_316 = !lean_is_exclusive(x_313); +if (x_316 == 0) { -return x_306; +return x_313; } else { -lean_object* x_310; lean_object* x_311; lean_object* x_312; -x_310 = lean_ctor_get(x_306, 0); -x_311 = lean_ctor_get(x_306, 1); -lean_inc(x_311); -lean_inc(x_310); -lean_dec(x_306); -x_312 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_312, 0, x_310); -lean_ctor_set(x_312, 1, x_311); -return x_312; +lean_object* x_317; lean_object* x_318; lean_object* x_319; +x_317 = lean_ctor_get(x_313, 0); +x_318 = lean_ctor_get(x_313, 1); +lean_inc(x_318); +lean_inc(x_317); +lean_dec(x_313); +x_319 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_319, 0, x_317); +lean_ctor_set(x_319, 1, x_318); +return x_319; } } } @@ -1203,85 +1232,87 @@ return x_312; } else { -lean_object* x_313; lean_object* x_314; -lean_dec(x_39); -x_313 = lean_box(0); -lean_inc(x_12); -x_314 = lean_apply_7(x_37, x_36, x_26, x_28, x_29, x_313, x_12, x_13); -if (lean_obj_tag(x_314) == 0) -{ -lean_object* x_315; lean_object* x_316; -x_315 = lean_ctor_get(x_314, 0); -lean_inc(x_315); -x_316 = lean_ctor_get(x_314, 1); -lean_inc(x_316); -lean_dec(x_314); -x_14 = x_315; -x_15 = x_316; -goto block_22; +lean_object* x_320; lean_object* x_321; +lean_dec(x_40); +x_320 = lean_box(0); +lean_inc(x_13); +x_321 = lean_apply_7(x_38, x_37, x_27, x_29, x_30, x_320, x_13, x_14); +if (lean_obj_tag(x_321) == 0) +{ +lean_object* x_322; lean_object* x_323; +x_322 = lean_ctor_get(x_321, 0); +lean_inc(x_322); +x_323 = lean_ctor_get(x_321, 1); +lean_inc(x_323); +lean_dec(x_321); +x_15 = x_322; +x_16 = x_323; +goto block_23; } else { -uint8_t x_317; -lean_dec(x_12); +uint8_t x_324; +lean_dec(x_13); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_317 = !lean_is_exclusive(x_314); -if (x_317 == 0) +x_324 = !lean_is_exclusive(x_321); +if (x_324 == 0) { -return x_314; +return x_321; } else { -lean_object* x_318; lean_object* x_319; lean_object* x_320; -x_318 = lean_ctor_get(x_314, 0); -x_319 = lean_ctor_get(x_314, 1); -lean_inc(x_319); -lean_inc(x_318); -lean_dec(x_314); -x_320 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_320, 0, x_318); -lean_ctor_set(x_320, 1, x_319); -return x_320; +lean_object* x_325; lean_object* x_326; lean_object* x_327; +x_325 = lean_ctor_get(x_321, 0); +x_326 = lean_ctor_get(x_321, 1); +lean_inc(x_326); +lean_inc(x_325); +lean_dec(x_321); +x_327 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_327, 0, x_325); +lean_ctor_set(x_327, 1, x_326); +return x_327; } } } } -block_22: +block_23: { -if (lean_obj_tag(x_14) == 0) +if (lean_obj_tag(x_15) == 0) { -lean_object* x_16; lean_object* x_17; -lean_dec(x_12); +lean_object* x_17; lean_object* x_18; +lean_dec(x_13); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_16 = lean_ctor_get(x_14, 0); -lean_inc(x_16); -lean_dec(x_14); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_15); -return x_17; +x_17 = lean_ctor_get(x_15, 0); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_16); +return x_18; } else { -lean_object* x_18; size_t x_19; size_t x_20; -x_18 = lean_ctor_get(x_14, 0); -lean_inc(x_18); -lean_dec(x_14); -x_19 = 1; -x_20 = lean_usize_add(x_10, x_19); -x_10 = x_20; -x_11 = x_18; -x_13 = x_15; +lean_object* x_19; size_t x_20; size_t x_21; +x_19 = lean_ctor_get(x_15, 0); +lean_inc(x_19); +lean_dec(x_15); +x_20 = 1; +x_21 = lean_usize_add(x_11, x_20); +x_11 = x_21; +x_12 = x_19; +x_14 = x_16; goto _start; } } @@ -1387,7 +1418,7 @@ x_14 = lean_ctor_get(x_7, 1); x_15 = !lean_is_exclusive(x_14); if (x_15 == 0) { -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; size_t x_24; size_t x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +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; size_t x_24; size_t x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; x_16 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__2___closed__8; x_17 = l_Lean_Syntax_setKind(x_12, x_16); x_18 = lean_unsigned_to_nat(3u); @@ -1400,363 +1431,366 @@ x_23 = lean_array_get_size(x_22); x_24 = lean_usize_of_nat(x_23); lean_dec(x_23); x_25 = 0; -x_26 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__2___closed__4; -x_27 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__2___closed__6; +x_26 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__2___closed__2; +x_27 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__2___closed__4; +x_28 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__2___closed__6; lean_inc(x_8); lean_inc(x_3); lean_inc(x_1); -x_28 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1(x_1, x_2, x_3, x_26, x_27, x_17, x_19, x_22, x_24, x_25, x_7, x_8, x_9); +x_29 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1(x_1, x_2, x_3, x_26, x_27, x_28, x_17, x_19, x_22, x_24, x_25, x_7, x_8, x_9); lean_dec(x_22); -if (lean_obj_tag(x_28) == 0) +if (lean_obj_tag(x_29) == 0) { -lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; -x_29 = lean_ctor_get(x_28, 0); -lean_inc(x_29); -x_30 = lean_ctor_get(x_29, 1); +lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_30 = lean_ctor_get(x_29, 0); lean_inc(x_30); -x_31 = lean_ctor_get(x_28, 1); +x_31 = lean_ctor_get(x_30, 1); lean_inc(x_31); -lean_dec(x_28); -x_32 = !lean_is_exclusive(x_29); -if (x_32 == 0) -{ -lean_object* x_33; uint8_t x_34; -x_33 = lean_ctor_get(x_29, 1); -lean_dec(x_33); -x_34 = !lean_is_exclusive(x_30); -if (x_34 == 0) -{ -size_t x_35; size_t x_36; -x_35 = 1; -x_36 = lean_usize_add(x_6, x_35); -x_6 = x_36; -x_7 = x_29; -x_9 = x_31; +x_32 = lean_ctor_get(x_29, 1); +lean_inc(x_32); +lean_dec(x_29); +x_33 = !lean_is_exclusive(x_30); +if (x_33 == 0) +{ +lean_object* x_34; uint8_t x_35; +x_34 = lean_ctor_get(x_30, 1); +lean_dec(x_34); +x_35 = !lean_is_exclusive(x_31); +if (x_35 == 0) +{ +size_t x_36; size_t x_37; +x_36 = 1; +x_37 = lean_usize_add(x_6, x_36); +x_6 = x_37; +x_7 = x_30; +x_9 = x_32; goto _start; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; size_t x_41; size_t x_42; -x_38 = lean_ctor_get(x_30, 0); -x_39 = lean_ctor_get(x_30, 1); +lean_object* x_39; lean_object* x_40; lean_object* x_41; size_t x_42; size_t x_43; +x_39 = lean_ctor_get(x_31, 0); +x_40 = lean_ctor_get(x_31, 1); +lean_inc(x_40); lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_30); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -lean_ctor_set(x_29, 1, x_40); -x_41 = 1; -x_42 = lean_usize_add(x_6, x_41); -x_6 = x_42; -x_7 = x_29; -x_9 = x_31; +lean_dec(x_31); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +lean_ctor_set(x_30, 1, x_41); +x_42 = 1; +x_43 = lean_usize_add(x_6, x_42); +x_6 = x_43; +x_7 = x_30; +x_9 = x_32; goto _start; } } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; size_t x_50; size_t x_51; -x_44 = lean_ctor_get(x_29, 0); -lean_inc(x_44); -lean_dec(x_29); +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; size_t x_51; size_t x_52; x_45 = lean_ctor_get(x_30, 0); lean_inc(x_45); -x_46 = lean_ctor_get(x_30, 1); +lean_dec(x_30); +x_46 = lean_ctor_get(x_31, 0); lean_inc(x_46); -if (lean_is_exclusive(x_30)) { - lean_ctor_release(x_30, 0); - lean_ctor_release(x_30, 1); - x_47 = x_30; +x_47 = lean_ctor_get(x_31, 1); +lean_inc(x_47); +if (lean_is_exclusive(x_31)) { + lean_ctor_release(x_31, 0); + lean_ctor_release(x_31, 1); + x_48 = x_31; } else { - lean_dec_ref(x_30); - x_47 = lean_box(0); + lean_dec_ref(x_31); + x_48 = lean_box(0); } -if (lean_is_scalar(x_47)) { - x_48 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_48)) { + x_49 = lean_alloc_ctor(0, 2, 0); } else { - x_48 = x_47; -} -lean_ctor_set(x_48, 0, x_45); -lean_ctor_set(x_48, 1, x_46); -x_49 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_49, 0, x_44); -lean_ctor_set(x_49, 1, x_48); -x_50 = 1; -x_51 = lean_usize_add(x_6, x_50); -x_6 = x_51; -x_7 = x_49; -x_9 = x_31; + x_49 = x_48; +} +lean_ctor_set(x_49, 0, x_46); +lean_ctor_set(x_49, 1, x_47); +x_50 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_50, 0, x_45); +lean_ctor_set(x_50, 1, x_49); +x_51 = 1; +x_52 = lean_usize_add(x_6, x_51); +x_6 = x_52; +x_7 = x_50; +x_9 = x_32; goto _start; } } else { -uint8_t x_53; +uint8_t x_54; lean_dec(x_8); lean_dec(x_3); lean_dec(x_1); -x_53 = !lean_is_exclusive(x_28); -if (x_53 == 0) +x_54 = !lean_is_exclusive(x_29); +if (x_54 == 0) { -return x_28; +return x_29; } else { -lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_54 = lean_ctor_get(x_28, 0); -x_55 = lean_ctor_get(x_28, 1); +lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_29, 0); +x_56 = lean_ctor_get(x_29, 1); +lean_inc(x_56); lean_inc(x_55); -lean_inc(x_54); -lean_dec(x_28); -x_56 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_56, 0, x_54); -lean_ctor_set(x_56, 1, x_55); -return x_56; +lean_dec(x_29); +x_57 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_56); +return x_57; } } } else { -lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; lean_object* x_67; size_t x_68; size_t x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_57 = lean_ctor_get(x_14, 0); -x_58 = lean_ctor_get(x_14, 1); +lean_object* x_58; lean_object* x_59; 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; lean_object* x_67; lean_object* x_68; size_t x_69; size_t x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_58 = lean_ctor_get(x_14, 0); +x_59 = lean_ctor_get(x_14, 1); +lean_inc(x_59); lean_inc(x_58); -lean_inc(x_57); lean_dec(x_14); -x_59 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__2___closed__8; -x_60 = l_Lean_Syntax_setKind(x_12, x_59); -x_61 = lean_unsigned_to_nat(3u); -x_62 = l_Lean_Syntax_getArg(x_60, x_61); -x_63 = lean_unsigned_to_nat(1u); -x_64 = l_Lean_Syntax_getArg(x_60, x_63); -x_65 = l_Lean_Syntax_getSepArgs(x_64); -lean_dec(x_64); -x_66 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_66, 0, x_57); -lean_ctor_set(x_66, 1, x_58); -lean_ctor_set(x_7, 1, x_66); -x_67 = lean_array_get_size(x_65); -x_68 = lean_usize_of_nat(x_67); -lean_dec(x_67); -x_69 = 0; -x_70 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__2___closed__4; -x_71 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__2___closed__6; +x_60 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__2___closed__8; +x_61 = l_Lean_Syntax_setKind(x_12, x_60); +x_62 = lean_unsigned_to_nat(3u); +x_63 = l_Lean_Syntax_getArg(x_61, x_62); +x_64 = lean_unsigned_to_nat(1u); +x_65 = l_Lean_Syntax_getArg(x_61, x_64); +x_66 = l_Lean_Syntax_getSepArgs(x_65); +lean_dec(x_65); +x_67 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_67, 0, x_58); +lean_ctor_set(x_67, 1, x_59); +lean_ctor_set(x_7, 1, x_67); +x_68 = lean_array_get_size(x_66); +x_69 = lean_usize_of_nat(x_68); +lean_dec(x_68); +x_70 = 0; +x_71 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__2___closed__2; +x_72 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__2___closed__4; +x_73 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__2___closed__6; lean_inc(x_8); lean_inc(x_3); lean_inc(x_1); -x_72 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1(x_1, x_2, x_3, x_70, x_71, x_60, x_62, x_65, x_68, x_69, x_7, x_8, x_9); -lean_dec(x_65); -if (lean_obj_tag(x_72) == 0) +x_74 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1(x_1, x_2, x_3, x_71, x_72, x_73, x_61, x_63, x_66, x_69, x_70, x_7, x_8, x_9); +lean_dec(x_66); +if (lean_obj_tag(x_74) == 0) { -lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; size_t x_83; size_t x_84; -x_73 = lean_ctor_get(x_72, 0); -lean_inc(x_73); -x_74 = lean_ctor_get(x_73, 1); -lean_inc(x_74); -x_75 = lean_ctor_get(x_72, 1); +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; size_t x_85; size_t x_86; +x_75 = lean_ctor_get(x_74, 0); lean_inc(x_75); -lean_dec(x_72); -x_76 = lean_ctor_get(x_73, 0); +x_76 = lean_ctor_get(x_75, 1); lean_inc(x_76); -if (lean_is_exclusive(x_73)) { - lean_ctor_release(x_73, 0); - lean_ctor_release(x_73, 1); - x_77 = x_73; -} else { - lean_dec_ref(x_73); - x_77 = lean_box(0); -} -x_78 = lean_ctor_get(x_74, 0); +x_77 = lean_ctor_get(x_74, 1); +lean_inc(x_77); +lean_dec(x_74); +x_78 = lean_ctor_get(x_75, 0); lean_inc(x_78); -x_79 = lean_ctor_get(x_74, 1); -lean_inc(x_79); -if (lean_is_exclusive(x_74)) { - lean_ctor_release(x_74, 0); - lean_ctor_release(x_74, 1); - x_80 = x_74; +if (lean_is_exclusive(x_75)) { + lean_ctor_release(x_75, 0); + lean_ctor_release(x_75, 1); + x_79 = x_75; } else { - lean_dec_ref(x_74); - x_80 = lean_box(0); + lean_dec_ref(x_75); + x_79 = lean_box(0); +} +x_80 = lean_ctor_get(x_76, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_76, 1); +lean_inc(x_81); +if (lean_is_exclusive(x_76)) { + lean_ctor_release(x_76, 0); + lean_ctor_release(x_76, 1); + x_82 = x_76; +} else { + lean_dec_ref(x_76); + x_82 = lean_box(0); } -if (lean_is_scalar(x_80)) { - x_81 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_82)) { + x_83 = lean_alloc_ctor(0, 2, 0); } else { - x_81 = x_80; + x_83 = x_82; } -lean_ctor_set(x_81, 0, x_78); -lean_ctor_set(x_81, 1, x_79); -if (lean_is_scalar(x_77)) { - x_82 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_83, 0, x_80); +lean_ctor_set(x_83, 1, x_81); +if (lean_is_scalar(x_79)) { + x_84 = lean_alloc_ctor(0, 2, 0); } else { - x_82 = x_77; -} -lean_ctor_set(x_82, 0, x_76); -lean_ctor_set(x_82, 1, x_81); -x_83 = 1; -x_84 = lean_usize_add(x_6, x_83); -x_6 = x_84; -x_7 = x_82; -x_9 = x_75; + x_84 = x_79; +} +lean_ctor_set(x_84, 0, x_78); +lean_ctor_set(x_84, 1, x_83); +x_85 = 1; +x_86 = lean_usize_add(x_6, x_85); +x_6 = x_86; +x_7 = x_84; +x_9 = x_77; goto _start; } else { -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_dec(x_8); lean_dec(x_3); lean_dec(x_1); -x_86 = lean_ctor_get(x_72, 0); -lean_inc(x_86); -x_87 = lean_ctor_get(x_72, 1); -lean_inc(x_87); -if (lean_is_exclusive(x_72)) { - lean_ctor_release(x_72, 0); - lean_ctor_release(x_72, 1); - x_88 = x_72; +x_88 = lean_ctor_get(x_74, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_74, 1); +lean_inc(x_89); +if (lean_is_exclusive(x_74)) { + lean_ctor_release(x_74, 0); + lean_ctor_release(x_74, 1); + x_90 = x_74; } else { - lean_dec_ref(x_72); - x_88 = lean_box(0); + lean_dec_ref(x_74); + x_90 = lean_box(0); } -if (lean_is_scalar(x_88)) { - x_89 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_90)) { + x_91 = lean_alloc_ctor(1, 2, 0); } else { - x_89 = x_88; + x_91 = x_90; } -lean_ctor_set(x_89, 0, x_86); -lean_ctor_set(x_89, 1, x_87); -return x_89; +lean_ctor_set(x_91, 0, x_88); +lean_ctor_set(x_91, 1, x_89); +return x_91; } } } else { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; size_t x_105; size_t x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; -x_90 = lean_ctor_get(x_7, 1); -x_91 = lean_ctor_get(x_7, 0); -lean_inc(x_90); -lean_inc(x_91); -lean_dec(x_7); -x_92 = lean_ctor_get(x_90, 0); +lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; size_t x_107; size_t x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_92 = lean_ctor_get(x_7, 1); +x_93 = lean_ctor_get(x_7, 0); lean_inc(x_92); -x_93 = lean_ctor_get(x_90, 1); lean_inc(x_93); -if (lean_is_exclusive(x_90)) { - lean_ctor_release(x_90, 0); - lean_ctor_release(x_90, 1); - x_94 = x_90; +lean_dec(x_7); +x_94 = lean_ctor_get(x_92, 0); +lean_inc(x_94); +x_95 = lean_ctor_get(x_92, 1); +lean_inc(x_95); +if (lean_is_exclusive(x_92)) { + lean_ctor_release(x_92, 0); + lean_ctor_release(x_92, 1); + x_96 = x_92; } else { - lean_dec_ref(x_90); - x_94 = lean_box(0); -} -x_95 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__2___closed__8; -x_96 = l_Lean_Syntax_setKind(x_12, x_95); -x_97 = lean_unsigned_to_nat(3u); -x_98 = l_Lean_Syntax_getArg(x_96, x_97); -x_99 = lean_unsigned_to_nat(1u); -x_100 = l_Lean_Syntax_getArg(x_96, x_99); -x_101 = l_Lean_Syntax_getSepArgs(x_100); -lean_dec(x_100); -if (lean_is_scalar(x_94)) { - x_102 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_92); + x_96 = lean_box(0); +} +x_97 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__2___closed__8; +x_98 = l_Lean_Syntax_setKind(x_12, x_97); +x_99 = lean_unsigned_to_nat(3u); +x_100 = l_Lean_Syntax_getArg(x_98, x_99); +x_101 = lean_unsigned_to_nat(1u); +x_102 = l_Lean_Syntax_getArg(x_98, x_101); +x_103 = l_Lean_Syntax_getSepArgs(x_102); +lean_dec(x_102); +if (lean_is_scalar(x_96)) { + x_104 = lean_alloc_ctor(0, 2, 0); } else { - x_102 = x_94; -} -lean_ctor_set(x_102, 0, x_92); -lean_ctor_set(x_102, 1, x_93); -x_103 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_103, 0, x_91); -lean_ctor_set(x_103, 1, x_102); -x_104 = lean_array_get_size(x_101); -x_105 = lean_usize_of_nat(x_104); -lean_dec(x_104); -x_106 = 0; -x_107 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__2___closed__4; -x_108 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__2___closed__6; + x_104 = x_96; +} +lean_ctor_set(x_104, 0, x_94); +lean_ctor_set(x_104, 1, x_95); +x_105 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_105, 0, x_93); +lean_ctor_set(x_105, 1, x_104); +x_106 = lean_array_get_size(x_103); +x_107 = lean_usize_of_nat(x_106); +lean_dec(x_106); +x_108 = 0; +x_109 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__2___closed__2; +x_110 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__2___closed__4; +x_111 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__2___closed__6; lean_inc(x_8); lean_inc(x_3); lean_inc(x_1); -x_109 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1(x_1, x_2, x_3, x_107, x_108, x_96, x_98, x_101, x_105, x_106, x_103, x_8, x_9); -lean_dec(x_101); -if (lean_obj_tag(x_109) == 0) -{ -lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; size_t x_120; size_t x_121; -x_110 = lean_ctor_get(x_109, 0); -lean_inc(x_110); -x_111 = lean_ctor_get(x_110, 1); -lean_inc(x_111); -x_112 = lean_ctor_get(x_109, 1); -lean_inc(x_112); -lean_dec(x_109); -x_113 = lean_ctor_get(x_110, 0); +x_112 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1(x_1, x_2, x_3, x_109, x_110, x_111, x_98, x_100, x_103, x_107, x_108, x_105, x_8, x_9); +lean_dec(x_103); +if (lean_obj_tag(x_112) == 0) +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; size_t x_123; size_t x_124; +x_113 = lean_ctor_get(x_112, 0); lean_inc(x_113); -if (lean_is_exclusive(x_110)) { - lean_ctor_release(x_110, 0); - lean_ctor_release(x_110, 1); - x_114 = x_110; -} else { - lean_dec_ref(x_110); - x_114 = lean_box(0); -} -x_115 = lean_ctor_get(x_111, 0); +x_114 = lean_ctor_get(x_113, 1); +lean_inc(x_114); +x_115 = lean_ctor_get(x_112, 1); lean_inc(x_115); -x_116 = lean_ctor_get(x_111, 1); +lean_dec(x_112); +x_116 = lean_ctor_get(x_113, 0); lean_inc(x_116); -if (lean_is_exclusive(x_111)) { - lean_ctor_release(x_111, 0); - lean_ctor_release(x_111, 1); - x_117 = x_111; +if (lean_is_exclusive(x_113)) { + lean_ctor_release(x_113, 0); + lean_ctor_release(x_113, 1); + x_117 = x_113; } else { - lean_dec_ref(x_111); + lean_dec_ref(x_113); x_117 = lean_box(0); } -if (lean_is_scalar(x_117)) { - x_118 = lean_alloc_ctor(0, 2, 0); +x_118 = lean_ctor_get(x_114, 0); +lean_inc(x_118); +x_119 = lean_ctor_get(x_114, 1); +lean_inc(x_119); +if (lean_is_exclusive(x_114)) { + lean_ctor_release(x_114, 0); + lean_ctor_release(x_114, 1); + x_120 = x_114; +} else { + lean_dec_ref(x_114); + x_120 = lean_box(0); +} +if (lean_is_scalar(x_120)) { + x_121 = lean_alloc_ctor(0, 2, 0); } else { - x_118 = x_117; + x_121 = x_120; } -lean_ctor_set(x_118, 0, x_115); -lean_ctor_set(x_118, 1, x_116); -if (lean_is_scalar(x_114)) { - x_119 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_121, 0, x_118); +lean_ctor_set(x_121, 1, x_119); +if (lean_is_scalar(x_117)) { + x_122 = lean_alloc_ctor(0, 2, 0); } else { - x_119 = x_114; -} -lean_ctor_set(x_119, 0, x_113); -lean_ctor_set(x_119, 1, x_118); -x_120 = 1; -x_121 = lean_usize_add(x_6, x_120); -x_6 = x_121; -x_7 = x_119; -x_9 = x_112; + x_122 = x_117; +} +lean_ctor_set(x_122, 0, x_116); +lean_ctor_set(x_122, 1, x_121); +x_123 = 1; +x_124 = lean_usize_add(x_6, x_123); +x_6 = x_124; +x_7 = x_122; +x_9 = x_115; goto _start; } else { -lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; +lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_dec(x_8); lean_dec(x_3); lean_dec(x_1); -x_123 = lean_ctor_get(x_109, 0); -lean_inc(x_123); -x_124 = lean_ctor_get(x_109, 1); -lean_inc(x_124); -if (lean_is_exclusive(x_109)) { - lean_ctor_release(x_109, 0); - lean_ctor_release(x_109, 1); - x_125 = x_109; +x_126 = lean_ctor_get(x_112, 0); +lean_inc(x_126); +x_127 = lean_ctor_get(x_112, 1); +lean_inc(x_127); +if (lean_is_exclusive(x_112)) { + lean_ctor_release(x_112, 0); + lean_ctor_release(x_112, 1); + x_128 = x_112; } else { - lean_dec_ref(x_109); - x_125 = lean_box(0); + lean_dec_ref(x_112); + x_128 = lean_box(0); } -if (lean_is_scalar(x_125)) { - x_126 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_128)) { + x_129 = lean_alloc_ctor(1, 2, 0); } else { - x_126 = x_125; + x_129 = x_128; } -lean_ctor_set(x_126, 0, x_123); -lean_ctor_set(x_126, 1, x_124); -return x_126; +lean_ctor_set(x_129, 0, x_126); +lean_ctor_set(x_129, 1, x_127); +return x_129; } } } @@ -1800,7 +1834,7 @@ static lean_object* _init_l___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__2___closed__6; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__25; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__26; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } @@ -1864,7 +1898,7 @@ lean_dec(x_17); x_22 = l___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___closed__4; x_23 = l_Lean_Syntax_setKind(x_2, x_22); x_24 = lean_box(2); -x_25 = l_Lean_nullKind; +x_25 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__3; x_26 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_26, 0, x_24); lean_ctor_set(x_26, 1, x_25); @@ -1898,7 +1932,7 @@ lean_dec(x_17); x_36 = l___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___closed__4; x_37 = l_Lean_Syntax_setKind(x_2, x_36); x_38 = lean_box(2); -x_39 = l_Lean_nullKind; +x_39 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__3; x_40 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_40, 0, x_38); lean_ctor_set(x_40, 1, x_39); @@ -1955,18 +1989,18 @@ lean_dec(x_5); return x_8; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__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, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__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, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -size_t x_14; size_t x_15; lean_object* x_16; -x_14 = lean_unbox_usize(x_9); -lean_dec(x_9); +size_t x_15; size_t x_16; lean_object* x_17; x_15 = lean_unbox_usize(x_10); lean_dec(x_10); -x_16 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_14, x_15, x_11, x_12, x_13); -lean_dec(x_8); +x_16 = lean_unbox_usize(x_11); +lean_dec(x_11); +x_17 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_15, x_16, x_12, x_13, x_14); +lean_dec(x_9); lean_dec(x_2); -return x_16; +return x_17; } } LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__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, lean_object* x_7, lean_object* x_8, lean_object* x_9) { @@ -2728,7 +2762,7 @@ static lean_object* _init_l_Lean_Elab_Tactic_evalMatch___closed__1() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__2___closed__4; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__11; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__13; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } @@ -2844,7 +2878,7 @@ x_27 = l_Lean_Elab_Tactic_evalMatch___closed__6; x_28 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_28, 0, x_21); lean_ctor_set(x_28, 1, x_27); -x_29 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__10; +x_29 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__12; x_30 = lean_array_push(x_29, x_28); x_31 = lean_array_push(x_30, x_18); x_32 = lean_box(2); @@ -2863,7 +2897,7 @@ lean_ctor_set(x_38, 2, x_36); x_39 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__1; x_40 = lean_array_push(x_39, x_38); x_41 = l_Array_append___rarg(x_40, x_19); -x_42 = l_Lean_nullKind; +x_42 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__3; x_43 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_43, 0, x_32); lean_ctor_set(x_43, 1, x_42); @@ -3054,7 +3088,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalMatch___closed__1( { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Tactic_evalMatch___closed__1; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__25; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__26; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } @@ -3082,7 +3116,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalMatch___closed__4( { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Elab_Tactic_evalMatch___closed__3; -x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__11; +x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__13; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } @@ -3137,7 +3171,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalMatch_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(52u); +x_1 = lean_unsigned_to_nat(53u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -3149,7 +3183,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalMatch_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(57u); +x_1 = lean_unsigned_to_nat(58u); x_2 = lean_unsigned_to_nat(52u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -3177,7 +3211,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalMatch_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(52u); +x_1 = lean_unsigned_to_nat(53u); x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -3189,7 +3223,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalMatch_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(52u); +x_1 = lean_unsigned_to_nat(53u); x_2 = lean_unsigned_to_nat(13u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -3312,6 +3346,8 @@ l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tact lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__25); l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__26 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__26(); lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__26); +l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__27 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__27(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__1___closed__27); l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__2___closed__1 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__2___closed__1(); lean_mark_persistent(l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__2___closed__1); l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__2___closed__2 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Match_0__Lean_Elab_Tactic_mkAuxiliaryMatchTerm___spec__2___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Rewrite.c b/stage0/stdlib/Lean/Elab/Tactic/Rewrite.c index aa5bbcd24c9..a03ee9957d8 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Rewrite.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Rewrite.c @@ -29,8 +29,8 @@ lean_object* lean_nat_div(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentArray_mapM___at_Lean_Elab_Tactic_elabRewriteConfig___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentArray_mapM___at_Lean_Elab_Tactic_elabRewriteConfig___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_withRWRulesSeq___spec__10___lambda__2___closed__7; -extern lean_object* l_Lean_nullKind; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_rewriteTarget___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentArray_mapMAux___at_Lean_Elab_Tactic_elabRewriteConfig___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_SavedState_restore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_withRWRulesSeq___spec__10___closed__1; @@ -54,6 +54,7 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_withRWRule static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalRewriteSeq___closed__10; lean_object* l_Std_PersistentArray_append___rarg(lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_withRWRulesSeq___closed__4; lean_object* l_Lean_mkIdentFrom(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabRewriteConfig___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); @@ -114,7 +115,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabRewrit static lean_object* l_Lean_withoutModifyingState___at_Lean_Elab_Tactic_elabRewriteConfig___spec__12___at_Lean_Elab_Tactic_elabRewriteConfig___spec__13___closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_rewriteTarget___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalRewriteSeq_declRange___closed__6; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820_(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832_(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwError___at___private_Lean_Elab_Tactic_Basic_0__Lean_Elab_Tactic_evalTacticUsing_loop___spec__2(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_throwError___at_Lean_Elab_Tactic_withRWRulesSeq___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_format_pretty(lean_object*, lean_object*); @@ -134,39 +135,38 @@ static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalRewriteSeq_declRange___c LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_withRWRulesSeq___spec__10___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_rewrite(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_rewriteLocalDecl(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__6; LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_Tactic_withRWRulesSeq___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_instInhabitedSyntax; lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ResolveName_resolveGlobalName(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_rewriteLocalDecl___lambda__1(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__4; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_withRWRulesSeq___spec__10___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, 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_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabRewriteConfig___spec__11(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentArray_mapMAux___at_Lean_Elab_Tactic_elabRewriteConfig___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); extern lean_object* l_Lean_Elab_Tactic_tacticElabAttribute; static lean_object* l_Lean_Elab_Tactic_evalRewriteSeq___lambda__2___closed__4; -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_withRWRulesSeq___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__6; +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_withRWRulesSeq___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_withRWRulesSeq___spec__10___lambda__2___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalRewriteSeq___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_pure___at_Lean_Elab_Tactic_saveTacticInfoForToken___spec__1___rarg___boxed(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_Std_Range_forIn_loop___at_Lean_Elab_Tactic_withRWRulesSeq___spec__10___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_withRWRulesSeq___spec__10___lambda__2___closed__8; +static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__4; lean_object* l_Lean_LocalDecl_type(lean_object*); lean_object* l_Lean_Meta_throwTacticEx___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_withRWRulesSeq___spec__10___lambda__2___closed__3; lean_object* l_Lean_Elab_Tactic_saveState___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__1; +static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__5; LEAN_EXPORT lean_object* l_Lean_resolveGlobalName___at_Lean_Elab_Tactic_withRWRulesSeq___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalRewriteSeq___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabRewriteConfig___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabRewriteConfig(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabRewriteConfig___spec__5(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_getLocalDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_getEqnsFor_x3f(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__2; LEAN_EXPORT lean_object* l_Lean_withoutModifyingState___at_Lean_Elab_Tactic_elabRewriteConfig___spec__12___at_Lean_Elab_Tactic_elabRewriteConfig___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_Tactic_withRWRulesSeq___spec__1___closed__3; uint8_t lean_nat_dec_le(lean_object*, lean_object*); @@ -175,14 +175,14 @@ static lean_object* l_Lean_resolveGlobalConst___at_Lean_Elab_Tactic_withRWRulesS lean_object* l_Lean_Elab_Term_SavedState_restore(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArgs(lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentArray_mapMAux___at_Lean_Elab_Tactic_elabRewriteConfig___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__5; +static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_rewriteLocalDecl___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, 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_Std_Range_forIn_loop___at_Lean_Elab_Tactic_withRWRulesSeq___spec__10___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__2; static lean_object* l_Lean_withoutModifyingState___at_Lean_Elab_Tactic_elabRewriteConfig___spec__12___at_Lean_Elab_Tactic_elabRewriteConfig___spec__13___closed__6; lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_withRWRulesSeq___spec__10___lambda__2___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalRewriteSeq___closed__9; -static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalRewriteSeq___boxed(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___regBuiltin_Lean_Elab_Tactic_evalRewriteSeq(lean_object*); lean_object* lean_nat_mul(lean_object*, lean_object*); @@ -192,6 +192,7 @@ lean_object* l_Lean_Elab_Tactic_elabTerm(lean_object*, lean_object*, uint8_t, le uint8_t l_Lean_Syntax_isNone(lean_object*); lean_object* l_Lean_Elab_Term_evalExpr___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_elabRewriteConfig___closed__1; +static lean_object* l_Lean_Elab_Tactic_withRWRulesSeq___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalRewriteSeq___closed__1; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_withRWRulesSeq___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_withLocation(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1290,7 +1291,7 @@ x_13 = lean_ctor_get(x_11, 1); lean_inc(x_13); lean_dec(x_11); x_14 = lean_box(0); -x_15 = l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(x_12, x_14); +x_15 = l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(x_12, x_14); x_16 = l_List_isEmpty___rarg(x_15); if (x_16 == 0) { @@ -2500,68 +2501,68 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_withRWRulesSeq___spec__10(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_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, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_withRWRulesSeq___spec__10(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_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, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19) { _start: { -uint8_t x_19; -x_19 = lean_nat_dec_le(x_7, x_6); -if (x_19 == 0) +uint8_t x_20; +x_20 = lean_nat_dec_le(x_8, x_7); +if (x_20 == 0) { -lean_object* x_20; uint8_t x_21; -x_20 = lean_unsigned_to_nat(0u); -x_21 = lean_nat_dec_eq(x_5, x_20); -if (x_21 == 0) +lean_object* x_21; uint8_t x_22; +x_21 = lean_unsigned_to_nat(0u); +x_22 = lean_nat_dec_eq(x_6, x_21); +if (x_22 == 0) { -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; uint8_t x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; lean_object* x_36; -lean_dec(x_9); -x_22 = lean_unsigned_to_nat(1u); -x_23 = lean_nat_sub(x_5, x_22); -lean_dec(x_5); -x_24 = lean_unsigned_to_nat(2u); -x_25 = lean_nat_mul(x_6, x_24); -x_26 = l_Lean_instInhabitedSyntax; -x_27 = lean_array_get(x_26, x_2, x_25); -x_28 = lean_nat_add(x_25, x_22); -lean_dec(x_25); -x_29 = lean_nat_dec_lt(x_28, x_4); -lean_inc(x_27); +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; uint8_t x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; lean_object* x_37; +lean_dec(x_10); +x_23 = lean_unsigned_to_nat(1u); +x_24 = lean_nat_sub(x_6, x_23); +lean_dec(x_6); +x_25 = lean_unsigned_to_nat(2u); +x_26 = lean_nat_mul(x_7, x_25); +x_27 = l_Lean_instInhabitedSyntax; +x_28 = lean_array_get(x_27, x_2, x_26); +x_29 = lean_nat_add(x_26, x_23); +lean_dec(x_26); +x_30 = lean_nat_dec_lt(x_29, x_5); +lean_inc(x_28); lean_inc(x_3); -x_30 = lean_array_push(x_3, x_27); -x_31 = l_Lean_Syntax_getArg(x_27, x_20); -x_32 = l_Lean_Syntax_isNone(x_31); -lean_dec(x_31); -x_33 = l_Lean_Syntax_getArg(x_27, x_22); -x_34 = l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_withRWRulesSeq___spec__10___closed__2; -lean_inc(x_33); -x_35 = l_Lean_Syntax_isOfKind(x_33, x_34); -if (x_29 == 0) +x_31 = lean_array_push(x_3, x_28); +x_32 = l_Lean_Syntax_getArg(x_28, x_21); +x_33 = l_Lean_Syntax_isNone(x_32); +lean_dec(x_32); +x_34 = l_Lean_Syntax_getArg(x_28, x_23); +x_35 = l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_withRWRulesSeq___spec__10___closed__2; +lean_inc(x_34); +x_36 = l_Lean_Syntax_isOfKind(x_34, x_35); +if (x_30 == 0) { lean_object* x_62; -lean_dec(x_28); +lean_dec(x_29); x_62 = lean_box(0); -x_36 = x_62; +x_37 = x_62; goto block_61; } else { lean_object* x_63; -x_63 = lean_array_fget(x_2, x_28); -lean_dec(x_28); -x_36 = x_63; +x_63 = lean_array_fget(x_2, x_29); +lean_dec(x_29); +x_37 = x_63; goto block_61; } block_61: { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; -x_37 = lean_array_push(x_30, x_36); -x_38 = lean_box(2); -x_39 = l_Lean_nullKind; +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_38 = lean_array_push(x_31, x_37); +x_39 = lean_box(2); +lean_inc(x_4); x_40 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -lean_ctor_set(x_40, 2, x_37); -x_41 = l_Lean_Elab_Tactic_mkInitialTacticInfo(x_40, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); -if (x_32 == 0) +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_4); +lean_ctor_set(x_40, 2, x_38); +x_41 = l_Lean_Elab_Tactic_mkInitialTacticInfo(x_40, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); +if (x_33 == 0) { uint8_t x_59; x_59 = 1; @@ -2583,20 +2584,21 @@ lean_inc(x_43); x_44 = lean_ctor_get(x_41, 1); lean_inc(x_44); lean_dec(x_41); -x_45 = lean_box(x_35); +x_45 = lean_box(x_36); x_46 = lean_box(x_42); lean_inc(x_1); -lean_inc(x_33); +lean_inc(x_34); x_47 = lean_alloc_closure((void*)(l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_withRWRulesSeq___spec__10___lambda__2___boxed), 16, 7); -lean_closure_set(x_47, 0, x_27); +lean_closure_set(x_47, 0, x_28); lean_closure_set(x_47, 1, x_45); -lean_closure_set(x_47, 2, x_33); +lean_closure_set(x_47, 2, x_34); lean_closure_set(x_47, 3, x_1); lean_closure_set(x_47, 4, x_46); -lean_closure_set(x_47, 5, x_33); -lean_closure_set(x_47, 6, x_34); +lean_closure_set(x_47, 5, x_34); +lean_closure_set(x_47, 6, x_35); x_48 = lean_alloc_closure((void*)(l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_withRWRulesSeq___spec__10___lambda__3), 11, 1); lean_closure_set(x_48, 0, x_43); +lean_inc(x_18); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); @@ -2604,27 +2606,27 @@ lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); -lean_inc(x_10); -x_49 = l_Lean_Elab_withInfoTreeContext___at___private_Lean_Elab_Tactic_Basic_0__Lean_Elab_Tactic_evalTacticUsing_loop___spec__3(x_47, x_48, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_44); +x_49 = l_Lean_Elab_withInfoTreeContext___at___private_Lean_Elab_Tactic_Basic_0__Lean_Elab_Tactic_evalTacticUsing_loop___spec__3(x_47, x_48, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_44); if (lean_obj_tag(x_49) == 0) { lean_object* x_50; lean_object* x_51; lean_object* x_52; x_50 = lean_ctor_get(x_49, 1); lean_inc(x_50); lean_dec(x_49); -x_51 = lean_nat_add(x_6, x_8); -lean_dec(x_6); +x_51 = lean_nat_add(x_7, x_9); +lean_dec(x_7); x_52 = lean_box(0); -x_5 = x_23; -x_6 = x_51; -x_9 = x_52; -x_18 = x_50; +x_6 = x_24; +x_7 = x_51; +x_10 = x_52; +x_19 = x_50; goto _start; } else { uint8_t x_54; -lean_dec(x_23); +lean_dec(x_24); +lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -2632,8 +2634,8 @@ lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_6); +lean_dec(x_7); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); x_54 = !lean_is_exclusive(x_49); @@ -2661,6 +2663,7 @@ return x_57; else { lean_object* x_64; +lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -2668,20 +2671,21 @@ lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); x_64 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_64, 0, x_9); -lean_ctor_set(x_64, 1, x_18); +lean_ctor_set(x_64, 0, x_10); +lean_ctor_set(x_64, 1, x_19); return x_64; } } else { lean_object* x_65; +lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -2689,14 +2693,14 @@ lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); x_65 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_65, 0, x_9); -lean_ctor_set(x_65, 1, x_18); +lean_ctor_set(x_65, 0, x_10); +lean_ctor_set(x_65, 1, x_19); return x_65; } } @@ -2713,6 +2717,24 @@ return x_2; static lean_object* _init_l_Lean_Elab_Tactic_withRWRulesSeq___closed__2() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("null", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_withRWRulesSeq___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Tactic_withRWRulesSeq___closed__2; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_withRWRulesSeq___closed__4() { +_start: +{ lean_object* x_1; lean_object* x_2; x_1 = lean_box(0); x_2 = lean_alloc_closure((void*)(l_ReaderT_pure___at_Lean_Elab_Tactic_saveTacticInfoForToken___spec__1___rarg___boxed), 10, 1); @@ -2734,7 +2756,7 @@ x_18 = l_Lean_Elab_Tactic_withRWRulesSeq___closed__1; x_19 = lean_array_push(x_18, x_1); x_20 = lean_array_push(x_19, x_14); x_21 = lean_box(2); -x_22 = l_Lean_nullKind; +x_22 = l_Lean_Elab_Tactic_withRWRulesSeq___closed__3; x_23 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_23, 0, x_21); lean_ctor_set(x_23, 1, x_22); @@ -2747,7 +2769,7 @@ lean_inc(x_26); lean_dec(x_24); x_27 = lean_alloc_closure((void*)(l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_withRWRulesSeq___spec__10___lambda__3), 11, 1); lean_closure_set(x_27, 0, x_25); -x_28 = l_Lean_Elab_Tactic_withRWRulesSeq___closed__2; +x_28 = l_Lean_Elab_Tactic_withRWRulesSeq___closed__4; lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -2770,7 +2792,7 @@ x_34 = lean_nat_div(x_32, x_33); lean_dec(x_32); x_35 = lean_box(0); lean_inc(x_34); -x_36 = l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_withRWRulesSeq___spec__10(x_3, x_17, x_18, x_31, x_34, x_13, x_34, x_15, x_35, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_30); +x_36 = l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_withRWRulesSeq___spec__10(x_3, x_17, x_18, x_22, x_31, x_34, x_13, x_34, x_15, x_35, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_30); lean_dec(x_34); lean_dec(x_31); lean_dec(x_17); @@ -2976,15 +2998,16 @@ lean_object* x_15 = _args[14]; lean_object* x_16 = _args[15]; lean_object* x_17 = _args[16]; lean_object* x_18 = _args[17]; +lean_object* x_19 = _args[18]; _start: { -lean_object* x_19; -x_19 = l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_withRWRulesSeq___spec__10(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_object* x_20; +x_20 = l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_withRWRulesSeq___spec__10(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); +lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_4); +lean_dec(x_5); lean_dec(x_2); -return x_19; +return x_20; } } LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_withRWRulesSeq___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { @@ -2996,7 +3019,7 @@ lean_dec(x_2); return x_13; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__1() { _start: { lean_object* x_1; @@ -3004,17 +3027,17 @@ x_1 = lean_mk_string_from_bytes("Meta", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__2() { +static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_withRWRulesSeq___spec__10___lambda__2___closed__2; -x_2 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__1; +x_2 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__3() { +static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__3() { _start: { lean_object* x_1; @@ -3022,17 +3045,17 @@ x_1 = lean_mk_string_from_bytes("Rewrite", 7); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__4() { +static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__2; -x_2 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__3; +x_1 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__2; +x_2 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__5() { +static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__5() { _start: { lean_object* x_1; @@ -3040,21 +3063,21 @@ x_1 = lean_mk_string_from_bytes("Config", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__6() { +static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__4; -x_2 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__5; +x_1 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__4; +x_2 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820_(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_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832_(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; -x_9 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__6; +x_9 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__6; x_10 = l_Lean_Elab_Term_evalExpr___rarg(x_9, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); return x_10; } @@ -4599,7 +4622,7 @@ static lean_object* _init_l_Lean_Elab_Tactic_elabRewriteConfig___closed__1() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__6; +x_2 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__6; x_3 = l_Lean_mkConst(x_2, x_1); return x_3; } @@ -4656,7 +4679,7 @@ lean_inc(x_17); x_18 = lean_ctor_get(x_16, 1); lean_inc(x_18); lean_dec(x_16); -x_19 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820_(x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_18); +x_19 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832_(x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_18); return x_19; } else @@ -5377,18 +5400,22 @@ l_Lean_Elab_Tactic_withRWRulesSeq___closed__1 = _init_l_Lean_Elab_Tactic_withRWR lean_mark_persistent(l_Lean_Elab_Tactic_withRWRulesSeq___closed__1); l_Lean_Elab_Tactic_withRWRulesSeq___closed__2 = _init_l_Lean_Elab_Tactic_withRWRulesSeq___closed__2(); lean_mark_persistent(l_Lean_Elab_Tactic_withRWRulesSeq___closed__2); -l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__1 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__1); -l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__2 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__2(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__2); -l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__3 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__3(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__3); -l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__4 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__4(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__4); -l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__5 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__5(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__5); -l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__6 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__6(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_820____closed__6); +l_Lean_Elab_Tactic_withRWRulesSeq___closed__3 = _init_l_Lean_Elab_Tactic_withRWRulesSeq___closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic_withRWRulesSeq___closed__3); +l_Lean_Elab_Tactic_withRWRulesSeq___closed__4 = _init_l_Lean_Elab_Tactic_withRWRulesSeq___closed__4(); +lean_mark_persistent(l_Lean_Elab_Tactic_withRWRulesSeq___closed__4); +l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__1 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__1); +l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__2 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__2); +l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__3 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__3); +l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__4 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__4(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__4); +l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__5 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__5(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__5); +l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__6 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__6(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_832____closed__6); l_Lean_withoutModifyingState___at_Lean_Elab_Tactic_elabRewriteConfig___spec__12___at_Lean_Elab_Tactic_elabRewriteConfig___spec__13___closed__1 = _init_l_Lean_withoutModifyingState___at_Lean_Elab_Tactic_elabRewriteConfig___spec__12___at_Lean_Elab_Tactic_elabRewriteConfig___spec__13___closed__1(); lean_mark_persistent(l_Lean_withoutModifyingState___at_Lean_Elab_Tactic_elabRewriteConfig___spec__12___at_Lean_Elab_Tactic_elabRewriteConfig___spec__13___closed__1); l_Lean_withoutModifyingState___at_Lean_Elab_Tactic_elabRewriteConfig___spec__12___at_Lean_Elab_Tactic_elabRewriteConfig___spec__13___closed__2 = _init_l_Lean_withoutModifyingState___at_Lean_Elab_Tactic_elabRewriteConfig___spec__12___at_Lean_Elab_Tactic_elabRewriteConfig___spec__13___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Simp.c b/stage0/stdlib/Lean/Elab/Tactic/Simp.c index 7499145e8a0..69392b77c05 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Simp.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Simp.c @@ -41,6 +41,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_m lean_object* l_Lean_LocalDecl_userName(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalDSimp_declRange___closed__5; LEAN_EXPORT lean_object* l_Std_PersistentArray_mapM___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_withoutModifyingState___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__12___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__13___closed__5; lean_object* l_Lean_Elab_Tactic_SavedState_restore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_elabSimpConfigCore___closed__2; @@ -51,13 +52,13 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_SimpKind_toCtorIdx(uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_tacticToDischarge___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_SimpExtension_getTheorems(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_beqSimpKind____x40_Lean_Elab_Tactic_Simp___hyg_370_(uint8_t, uint8_t); +LEAN_EXPORT uint8_t l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_beqSimpKind____x40_Lean_Elab_Tactic_Simp___hyg_371_(uint8_t, uint8_t); lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_withoutModifyingStateWithInfoAndMessagesImpl___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_mapTRAux___at_Lean_resolveGlobalConstCore___spec__2(lean_object*, lean_object*); extern lean_object* l_Std_Format_defWidth; static lean_object* l_Lean_Elab_Tactic_tacticToDischarge___closed__23; -static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__10(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__2; lean_object* l_Lean_SourceInfo_fromRef(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSimpAll___closed__3; lean_object* l_List_filterMap___at_Lean_resolveGlobalConst___spec__1(lean_object*); @@ -88,10 +89,10 @@ static lean_object* l_Lean_Elab_Tactic_tacticToDischarge___lambda__3___closed__1 uint8_t lean_name_eq(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_mkSimpContext___spec__3___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalDSimp_declRange___closed__6; -static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__6; static lean_object* l_Lean_withoutModifyingState___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__12___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__13___closed__7; static lean_object* l_Lean_throwUnknownConstant___at___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs___spec__9___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSimpAll_declRange___closed__1; +static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__6; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_addDeclToUnfoldOrTheorem(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_SimpTheorems_add(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs___spec__18___closed__6; @@ -124,11 +125,12 @@ static lean_object* l_Lean_Elab_Tactic_tacticToDischarge___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabSimpConfigCore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Simp_DischargeWrapper_with(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_mkSimpContext___lambda__2(lean_object*, uint8_t, uint8_t, uint8_t, 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___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_beqSimpKind____x40_Lean_Elab_Tactic_Simp___hyg_371____boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__1; static lean_object* l_Lean_Elab_Tactic_tacticToDischarge___closed__20; static lean_object* l_Lean_Elab_Tactic_mkSimpContext___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoTree___at___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs___spec__16(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_Std_PersistentArray_mapMAux___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_mkSimpContext___boxed(lean_object*, lean_object*, lean_object*, 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___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs___spec__1___rarg(lean_object*); extern lean_object* l_Lean_LocalContext_empty; @@ -136,7 +138,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_evalDSimp(lean_object*) LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalSimpAll(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___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_addSimpTheorem___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_mapTRAux___at_Lean_resolveGlobalConstNoOverload___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_beqSimpKind____x40_Lean_Elab_Tactic_Simp___hyg_370____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withoutModifyingState___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__12___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Std_PersistentHashMap_contains___at_Lean_Meta_SimpTheorems_erase___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs___spec__1___rarg___closed__2; @@ -151,9 +152,9 @@ LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Tactic_ lean_object* l_Std_RBNode_setBlack___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorems_erase___at___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs___spec__17___lambda__1(lean_object*, lean_object*, 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_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__5(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__4; lean_object* l_Lean_Elab_Term_elabTermEnsuringType(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__4; static lean_object* l_Lean_resolveGlobalConst___at___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs___spec__4___closed__3; static lean_object* l_Lean_withoutModifyingState___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__12___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__13___closed__2; static lean_object* l_Lean_Elab_Tactic_tacticToDischarge___closed__3; @@ -166,8 +167,8 @@ static lean_object* l_Lean_Elab_Tactic_mkSimpContext___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_mkSimpContext___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, 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_Tactic_elabSimpConfigCtxCore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__8; -static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_123____closed__2; +static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__8; +static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_124____closed__2; lean_object* l_Lean_Elab_Tactic_expandOptLocation(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_addSimpTheorem___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Meta_SimpTheorems_isLemma(lean_object*, lean_object*); @@ -200,7 +201,7 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_mkSimpCo LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalDSimp_declRange___closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabSimpConfig(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_123____closed__1; +static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_124____closed__1; uint8_t l_Lean_Name_quickCmp(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_mkSimpContext___lambda__2___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_resolveGlobalConstNoOverloadWithInfo___at___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -249,8 +250,8 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabSimpConfigCore___boxed(lean_obje lean_object* l_Lean_ResolveName_resolveGlobalName(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_SimpTheorems_eraseCore(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_mkSimpContext___spec__3___at_Lean_Elab_Tactic_mkSimpContext___spec__4(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_241____closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_tacticToDischarge___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_242____closed__3; extern lean_object* l_Lean_Meta_instInhabitedSimpTheorems; static lean_object* l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_addDeclToUnfoldOrTheorem___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Simp_DischargeWrapper_with___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -280,7 +281,6 @@ static lean_object* l_Lean_Elab_Tactic_tacticToDischarge___closed__11; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs___spec__6___boxed(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_Std_PersistentArray_mapMAux___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__3(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_pushInfoTree___at___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Tactic_tacticToDischarge___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_saveState___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSimp___closed__2; @@ -306,6 +306,7 @@ LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Tactic_ uint8_t l_Std_RBNode_isRed___rarg(lean_object*); lean_object* l_Lean_Elab_Term_SavedState_restore(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getKind(lean_object*); +static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_242____closed__2; lean_object* l_Lean_Meta_dsimpGoal(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_simpAll(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSimp___closed__4; @@ -313,18 +314,18 @@ static lean_object* l_Lean_Elab_Tactic_tacticToDischarge___closed__16; static lean_object* l_Lean_throwUnknownConstant___at___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs___spec__9___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSimp___closed__8; static lean_object* l_Lean_resolveGlobalConst___at___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs___spec__4___closed__1; -static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_241____closed__2; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSimpAll_declRange___closed__7; -static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_241____closed__1; static lean_object* l_Lean_Elab_Tactic_tacticToDischarge___closed__12; +static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_242____closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs_resolveSimpIdTheorem_x3f(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_Tactic_evalDSimp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__3; +static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__3; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSimpAll___closed__2; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_mapTRAux___at_Lean_mkConstWithLevelParams___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Meta_getSimpExtension_x3f(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__5; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwUnknownConstant___at___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs___spec__9___closed__2; static lean_object* l_Lean_Elab_Tactic_elabSimpConfigCore___closed__3; @@ -333,6 +334,7 @@ uint8_t lean_nat_dec_eq(lean_object*, lean_object*); lean_object* l_Lean_Meta_simpGoal(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_withoutModifyingState___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__12___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__13___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_resolveGlobalConstNoOverloadWithInfo___at___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs_resolveSimpIdTheorem_x3f___closed__3; lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorems_erase___at___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs___spec__17___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isNone(lean_object*); @@ -343,11 +345,11 @@ static lean_object* l_Lean_Elab_Tactic_tacticToDischarge___closed__13; LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs___spec__12(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_pushInfoLeaf___at___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs___spec__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_elabSimpConfigCtxCore___closed__2; -static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__5; lean_object* l_Lean_Meta_getPropHyps(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs___spec__18___closed__1; lean_object* l_Lean_Meta_instantiateMVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_dsimpLocation_go(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs_resolveSimpIdTheorem_x3f___boxed(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_Tactic_tacticToDischarge(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSimpAll_declRange___closed__3; @@ -370,6 +372,7 @@ static lean_object* l_Lean_Elab_Tactic_tacticToDischarge___closed__2; lean_object* l_Lean_Meta_SimpTheoremsArray_addTheorem(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_SimpTheorems_addConst(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSimp___closed__5; +static lean_object* l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs_resolveSimpIdTheorem_x3f___closed__4; static lean_object* l_Lean_Elab_Tactic_mkSimpContext___closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_mkSimpContext___spec__3___at_Lean_Elab_Tactic_mkSimpContext___spec__4___boxed(lean_object*, lean_object*, lean_object*, 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___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_addSimpTheorem(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -380,9 +383,9 @@ static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSimp___closed__6; static lean_object* l_Lean_resolveGlobalConstNoOverload___at___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs___spec__3___closed__2; uint8_t l_List_isEmpty___rarg(lean_object*); static lean_object* l_Lean_Elab_Tactic_elabSimpConfigCtxCore___closed__1; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_241_(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_123_(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5_(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_242_(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_124_(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6_(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_addDeclToUnfoldOrTheorem___lambda__1___boxed(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_Tactic_dsimpLocation_go___boxed(lean_object*, lean_object*, 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_Tactic_elabDSimpConfigCore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -392,7 +395,6 @@ lean_object* l_List_toString___at_Lean_resolveGlobalConstNoOverloadCore___spec__ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_SimpKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Meta_SimpTheoremsArray_isErased(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentArray_mapM___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__7; LEAN_EXPORT lean_object* l_Std_PersistentArray_mapM___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_tacticToDischarge___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -402,6 +404,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instBEqSimpKind; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalDSimp___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs___boxed__const__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_mkSimpContext(lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__7; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSimp_declRange___closed__2; static lean_object* l_Lean_Elab_Tactic_SimpKind_noConfusion___rarg___closed__1; lean_object* l_Lean_Elab_Term_isLocalIdent_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -425,8 +428,7 @@ LEAN_EXPORT lean_object* l_Lean_resolveGlobalName___at___private_Lean_Elab_Tacti lean_object* l_Lean_Meta_getSimpCongrTheorems___rarg(lean_object*, lean_object*); lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___boxed(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___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_addSimpTheorem___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Syntax_isIdent(lean_object*); -static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__1() { _start: { lean_object* x_1; @@ -434,17 +436,17 @@ x_1 = lean_mk_string_from_bytes("Lean", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__2() { +static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__1; +x_2 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__3() { +static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__3() { _start: { lean_object* x_1; @@ -452,17 +454,17 @@ x_1 = lean_mk_string_from_bytes("Meta", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__4() { +static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__2; -x_2 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__3; +x_1 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__2; +x_2 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__5() { +static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__5() { _start: { lean_object* x_1; @@ -470,17 +472,17 @@ x_1 = lean_mk_string_from_bytes("Simp", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__6() { +static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__4; -x_2 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__5; +x_1 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__4; +x_2 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__7() { +static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__7() { _start: { lean_object* x_1; @@ -488,21 +490,21 @@ x_1 = lean_mk_string_from_bytes("Config", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__8() { +static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__6; -x_2 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__7; +x_1 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__6; +x_2 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__7; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5_(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_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6_(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; -x_9 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__8; +x_9 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__8; x_10 = l_Lean_Elab_Term_evalExpr___rarg(x_9, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); return x_10; } @@ -2047,7 +2049,7 @@ static lean_object* _init_l_Lean_Elab_Tactic_elabSimpConfigCore___closed__1() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__8; +x_2 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__8; x_3 = l_Lean_mkConst(x_2, x_1); return x_3; } @@ -2119,7 +2121,7 @@ lean_inc(x_17); x_18 = lean_ctor_get(x_16, 1); lean_inc(x_18); lean_dec(x_16); -x_19 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5_(x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_18); +x_19 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6_(x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_18); return x_19; } else @@ -2341,7 +2343,7 @@ lean_dec(x_1); return x_9; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_123____closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_124____closed__1() { _start: { lean_object* x_1; @@ -2349,21 +2351,21 @@ x_1 = lean_mk_string_from_bytes("ConfigCtx", 9); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_123____closed__2() { +static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_124____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__6; -x_2 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_123____closed__1; +x_1 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__6; +x_2 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_124____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_123_(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_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_124_(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; -x_9 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_123____closed__2; +x_9 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_124____closed__2; x_10 = l_Lean_Elab_Term_evalExpr___rarg(x_9, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); return x_10; } @@ -2373,7 +2375,7 @@ static lean_object* _init_l_Lean_Elab_Tactic_elabSimpConfigCtxCore___closed__1() { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_123____closed__2; +x_2 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_124____closed__2; x_3 = l_Lean_mkConst(x_2, x_1); return x_3; } @@ -2465,7 +2467,7 @@ lean_inc(x_27); x_28 = lean_ctor_get(x_26, 1); lean_inc(x_28); lean_dec(x_26); -x_29 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_123_(x_27, x_2, x_3, x_4, x_5, x_6, x_7, x_28); +x_29 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_124_(x_27, x_2, x_3, x_4, x_5, x_6, x_7, x_28); return x_29; } else @@ -2523,7 +2525,7 @@ lean_dec(x_1); return x_9; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_241____closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_242____closed__1() { _start: { lean_object* x_1; @@ -2531,31 +2533,31 @@ x_1 = lean_mk_string_from_bytes("DSimp", 5); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_241____closed__2() { +static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_242____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__4; -x_2 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_241____closed__1; +x_1 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__4; +x_2 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_242____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_241____closed__3() { +static lean_object* _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_242____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_241____closed__2; -x_2 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__7; +x_1 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_242____closed__2; +x_2 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__7; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_241_(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_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_242_(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; -x_9 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_241____closed__3; +x_9 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_242____closed__3; x_10 = l_Lean_Elab_Term_evalExpr___rarg(x_9, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); return x_10; } @@ -2565,7 +2567,7 @@ static lean_object* _init_l_Lean_Elab_Tactic_elabDSimpConfigCore___closed__1() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_241____closed__3; +x_2 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_242____closed__3; x_3 = l_Lean_mkConst(x_2, x_1); return x_3; } @@ -2649,7 +2651,7 @@ lean_inc(x_27); x_28 = lean_ctor_get(x_26, 1); lean_inc(x_28); lean_dec(x_26); -x_29 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_241_(x_27, x_2, x_3, x_4, x_5, x_6, x_7, x_28); +x_29 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_242_(x_27, x_2, x_3, x_4, x_5, x_6, x_7, x_28); return x_29; } else @@ -2802,7 +2804,7 @@ x_1 = 0; return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_beqSimpKind____x40_Lean_Elab_Tactic_Simp___hyg_370_(uint8_t x_1, uint8_t x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_beqSimpKind____x40_Lean_Elab_Tactic_Simp___hyg_371_(uint8_t x_1, uint8_t x_2) { _start: { lean_object* x_3; lean_object* x_4; uint8_t x_5; @@ -2814,7 +2816,7 @@ lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_beqSimpKind____x40_Lean_Elab_Tactic_Simp___hyg_370____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_beqSimpKind____x40_Lean_Elab_Tactic_Simp___hyg_371____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; @@ -2822,7 +2824,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_Elab_Tactic_Simp_0__Lean_Elab_Tactic_beqSimpKind____x40_Lean_Elab_Tactic_Simp___hyg_370_(x_3, x_4); +x_5 = l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_beqSimpKind____x40_Lean_Elab_Tactic_Simp___hyg_371_(x_3, x_4); x_6 = lean_box(x_5); return x_6; } @@ -2831,7 +2833,7 @@ static lean_object* _init_l_Lean_Elab_Tactic_instBEqSimpKind___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_beqSimpKind____x40_Lean_Elab_Tactic_Simp___hyg_370____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_beqSimpKind____x40_Lean_Elab_Tactic_Simp___hyg_371____boxed), 2, 0); return x_1; } } @@ -3222,7 +3224,7 @@ static lean_object* _init_l_Lean_Elab_Tactic_tacticToDischarge___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__2; +x_1 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__2; x_2 = l_Lean_Elab_Tactic_tacticToDischarge___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -4080,7 +4082,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_a { uint8_t x_10; uint8_t x_11; x_10 = 2; -x_11 = l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_beqSimpKind____x40_Lean_Elab_Tactic_Simp___hyg_370_(x_1, x_10); +x_11 = l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_beqSimpKind____x40_Lean_Elab_Tactic_Simp___hyg_371_(x_1, x_10); if (x_11 == 0) { lean_object* x_12; @@ -4709,6 +4711,24 @@ return x_1; static lean_object* _init_l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs_resolveSimpIdTheorem_x3f___closed__1() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("ident", 5); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs_resolveSimpIdTheorem_x3f___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs_resolveSimpIdTheorem_x3f___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs_resolveSimpIdTheorem_x3f___closed__3() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); x_2 = lean_box(0); @@ -4718,7 +4738,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs_resolveSimpIdTheorem_x3f___closed__2() { +static lean_object* _init_l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs_resolveSimpIdTheorem_x3f___closed__4() { _start: { lean_object* x_1; @@ -4729,112 +4749,114 @@ return x_1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs_resolveSimpIdTheorem_x3f(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_object* x_9, lean_object* x_10) { _start: { -lean_object* x_11; lean_object* x_12; uint8_t x_16; -x_16 = l_Lean_Syntax_isIdent(x_1); -if (x_16 == 0) -{ -lean_object* x_17; -x_17 = l_Lean_Elab_Term_elabCDotFunctionAlias_x3f(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_17) == 0) +lean_object* x_11; lean_object* x_12; lean_object* x_16; uint8_t x_17; +x_16 = l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs_resolveSimpIdTheorem_x3f___closed__2; +lean_inc(x_1); +x_17 = l_Lean_Syntax_isOfKind(x_1, x_16); +if (x_17 == 0) { lean_object* x_18; -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); +x_18 = l_Lean_Elab_Term_elabCDotFunctionAlias_x3f(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10); if (lean_obj_tag(x_18) == 0) { -uint8_t x_19; -x_19 = !lean_is_exclusive(x_17); -if (x_19 == 0) +lean_object* x_19; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +if (lean_obj_tag(x_19) == 0) { -lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_17, 0); -lean_dec(x_20); -x_21 = lean_box(0); -lean_ctor_set(x_17, 0, x_21); -return x_17; +uint8_t x_20; +x_20 = !lean_is_exclusive(x_18); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_18, 0); +lean_dec(x_21); +x_22 = lean_box(0); +lean_ctor_set(x_18, 0, x_22); +return x_18; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_17, 1); -lean_inc(x_22); -lean_dec(x_17); -x_23 = lean_box(0); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_22); -return x_24; +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_18, 1); +lean_inc(x_23); +lean_dec(x_18); +x_24 = lean_box(0); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_23); +return x_25; } } else { -uint8_t x_25; -x_25 = !lean_is_exclusive(x_17); -if (x_25 == 0) +uint8_t x_26; +x_26 = !lean_is_exclusive(x_18); +if (x_26 == 0) { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_17, 0); -lean_dec(x_26); +lean_object* x_27; lean_object* x_28; lean_object* x_29; x_27 = lean_ctor_get(x_18, 0); -lean_inc(x_27); -lean_dec(x_18); -x_28 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_28, 0, x_27); -lean_ctor_set(x_17, 0, x_28); -return x_17; +lean_dec(x_27); +x_28 = lean_ctor_get(x_19, 0); +lean_inc(x_28); +lean_dec(x_19); +x_29 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_18, 0, x_29); +return x_18; } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_29 = lean_ctor_get(x_17, 1); -lean_inc(x_29); -lean_dec(x_17); -x_30 = lean_ctor_get(x_18, 0); +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_30 = lean_ctor_get(x_18, 1); lean_inc(x_30); lean_dec(x_18); -x_31 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_31, 0, x_30); -x_32 = lean_alloc_ctor(0, 2, 0); +x_31 = lean_ctor_get(x_19, 0); +lean_inc(x_31); +lean_dec(x_19); +x_32 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_29); -return x_32; +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_30); +return x_33; } } } else { -uint8_t x_33; -x_33 = !lean_is_exclusive(x_17); -if (x_33 == 0) +uint8_t x_34; +x_34 = !lean_is_exclusive(x_18); +if (x_34 == 0) { -return x_17; +return x_18; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_17, 0); -x_35 = lean_ctor_get(x_17, 1); +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_18, 0); +x_36 = lean_ctor_get(x_18, 1); +lean_inc(x_36); lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_17); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -return x_36; +lean_dec(x_18); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +return x_37; } } } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_57; uint8_t x_58; lean_object* x_59; -x_37 = l_Lean_Elab_Tactic_saveState___rarg(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -x_38 = lean_ctor_get(x_37, 0); -lean_inc(x_38); -x_39 = lean_ctor_get(x_37, 1); +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_58; uint8_t x_59; lean_object* x_60; +x_38 = l_Lean_Elab_Tactic_saveState___rarg(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_39 = lean_ctor_get(x_38, 0); lean_inc(x_39); -lean_dec(x_37); -x_57 = l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs_resolveSimpIdTheorem_x3f___closed__2; -x_58 = 1; +x_40 = lean_ctor_get(x_38, 1); +lean_inc(x_40); +lean_dec(x_38); +x_58 = l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs_resolveSimpIdTheorem_x3f___closed__4; +x_59 = 1; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); @@ -4842,148 +4864,148 @@ lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_1); -x_59 = l_Lean_Elab_Term_resolveId_x3f(x_1, x_57, x_58, x_4, x_5, x_6, x_7, x_8, x_9, x_39); -if (lean_obj_tag(x_59) == 0) +x_60 = l_Lean_Elab_Term_resolveId_x3f(x_1, x_58, x_59, x_4, x_5, x_6, x_7, x_8, x_9, x_40); +if (lean_obj_tag(x_60) == 0) { -lean_object* x_60; -lean_dec(x_38); +lean_object* x_61; +lean_dec(x_39); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_60 = lean_ctor_get(x_59, 0); -lean_inc(x_60); -if (lean_obj_tag(x_60) == 0) -{ -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_61 = lean_ctor_get(x_59, 1); +x_61 = lean_ctor_get(x_60, 0); lean_inc(x_61); -lean_dec(x_59); -x_62 = l_Lean_Syntax_getId(x_1); +if (lean_obj_tag(x_61) == 0) +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_62 = lean_ctor_get(x_60, 1); +lean_inc(x_62); +lean_dec(x_60); +x_63 = l_Lean_Syntax_getId(x_1); lean_dec(x_1); -x_63 = lean_erase_macro_scopes(x_62); -x_64 = lean_st_ref_get(x_9, x_61); +x_64 = lean_erase_macro_scopes(x_63); +x_65 = lean_st_ref_get(x_9, x_62); lean_dec(x_9); -x_65 = lean_ctor_get(x_64, 1); -lean_inc(x_65); -lean_dec(x_64); -x_66 = l_Lean_Meta_getSimpExtension_x3f(x_63, x_65); -x_67 = lean_ctor_get(x_66, 0); -lean_inc(x_67); -if (lean_obj_tag(x_67) == 0) -{ -lean_object* x_68; lean_object* x_69; -x_68 = lean_ctor_get(x_66, 1); +x_66 = lean_ctor_get(x_65, 1); +lean_inc(x_66); +lean_dec(x_65); +x_67 = l_Lean_Meta_getSimpExtension_x3f(x_64, x_66); +x_68 = lean_ctor_get(x_67, 0); lean_inc(x_68); -lean_dec(x_66); -x_69 = l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs_resolveSimpIdTheorem_x3f___closed__1; -x_11 = x_69; -x_12 = x_68; +if (lean_obj_tag(x_68) == 0) +{ +lean_object* x_69; lean_object* x_70; +x_69 = lean_ctor_get(x_67, 1); +lean_inc(x_69); +lean_dec(x_67); +x_70 = l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs_resolveSimpIdTheorem_x3f___closed__3; +x_11 = x_70; +x_12 = x_69; goto block_15; } else { -lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_70 = lean_ctor_get(x_66, 1); -lean_inc(x_70); -lean_dec(x_66); -x_71 = lean_ctor_get(x_67, 0); +lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_71 = lean_ctor_get(x_67, 1); lean_inc(x_71); lean_dec(x_67); -x_72 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_72, 0, x_71); -x_73 = lean_box(0); -x_74 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_74, 0, x_72); -lean_ctor_set(x_74, 1, x_73); -x_11 = x_74; -x_12 = x_70; +x_72 = lean_ctor_get(x_68, 0); +lean_inc(x_72); +lean_dec(x_68); +x_73 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_73, 0, x_72); +x_74 = lean_box(0); +x_75 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set(x_75, 1, x_74); +x_11 = x_75; +x_12 = x_71; goto block_15; } } else { -lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_dec(x_9); lean_dec(x_1); -x_75 = lean_ctor_get(x_59, 1); -lean_inc(x_75); -lean_dec(x_59); -x_76 = lean_ctor_get(x_60, 0); +x_76 = lean_ctor_get(x_60, 1); lean_inc(x_76); lean_dec(x_60); -x_77 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_77, 0, x_76); -x_78 = lean_box(0); -x_79 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_79, 0, x_77); -lean_ctor_set(x_79, 1, x_78); -x_11 = x_79; -x_12 = x_75; +x_77 = lean_ctor_get(x_61, 0); +lean_inc(x_77); +lean_dec(x_61); +x_78 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_78, 0, x_77); +x_79 = lean_box(0); +x_80 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_80, 0, x_78); +lean_ctor_set(x_80, 1, x_79); +x_11 = x_80; +x_12 = x_76; goto block_15; } } else { -lean_object* x_80; -x_80 = lean_ctor_get(x_59, 1); -lean_inc(x_80); -lean_dec(x_59); -x_40 = x_80; -goto block_56; +lean_object* x_81; +x_81 = lean_ctor_get(x_60, 1); +lean_inc(x_81); +lean_dec(x_60); +x_41 = x_81; +goto block_57; } -block_56: +block_57: { -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_41 = l_Lean_Elab_Tactic_SavedState_restore(x_38, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_40); +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_42 = l_Lean_Elab_Tactic_SavedState_restore(x_39, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_41); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_42 = lean_ctor_get(x_41, 1); -lean_inc(x_42); -lean_dec(x_41); -x_43 = l_Lean_Syntax_getId(x_1); +x_43 = lean_ctor_get(x_42, 1); +lean_inc(x_43); +lean_dec(x_42); +x_44 = l_Lean_Syntax_getId(x_1); lean_dec(x_1); -x_44 = lean_erase_macro_scopes(x_43); -x_45 = lean_st_ref_get(x_9, x_42); +x_45 = lean_erase_macro_scopes(x_44); +x_46 = lean_st_ref_get(x_9, x_43); lean_dec(x_9); -x_46 = lean_ctor_get(x_45, 1); -lean_inc(x_46); -lean_dec(x_45); -x_47 = l_Lean_Meta_getSimpExtension_x3f(x_44, x_46); -x_48 = lean_ctor_get(x_47, 0); -lean_inc(x_48); -if (lean_obj_tag(x_48) == 0) -{ -lean_object* x_49; lean_object* x_50; -x_49 = lean_ctor_get(x_47, 1); +x_47 = lean_ctor_get(x_46, 1); +lean_inc(x_47); +lean_dec(x_46); +x_48 = l_Lean_Meta_getSimpExtension_x3f(x_45, x_47); +x_49 = lean_ctor_get(x_48, 0); lean_inc(x_49); -lean_dec(x_47); -x_50 = l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs_resolveSimpIdTheorem_x3f___closed__1; -x_11 = x_50; -x_12 = x_49; +if (lean_obj_tag(x_49) == 0) +{ +lean_object* x_50; lean_object* x_51; +x_50 = lean_ctor_get(x_48, 1); +lean_inc(x_50); +lean_dec(x_48); +x_51 = l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs_resolveSimpIdTheorem_x3f___closed__3; +x_11 = x_51; +x_12 = x_50; goto block_15; } else { -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_51 = lean_ctor_get(x_47, 1); -lean_inc(x_51); -lean_dec(x_47); -x_52 = lean_ctor_get(x_48, 0); +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_52 = lean_ctor_get(x_48, 1); lean_inc(x_52); lean_dec(x_48); -x_53 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_53, 0, x_52); -x_54 = lean_box(0); -x_55 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_55, 0, x_53); -lean_ctor_set(x_55, 1, x_54); -x_11 = x_55; -x_12 = x_51; +x_53 = lean_ctor_get(x_49, 0); +lean_inc(x_53); +lean_dec(x_49); +x_54 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_54, 0, x_53); +x_55 = lean_box(0); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_54); +lean_ctor_set(x_56, 1, x_55); +x_11 = x_56; +x_12 = x_52; goto block_15; } } @@ -5290,7 +5312,7 @@ x_13 = lean_ctor_get(x_11, 1); lean_inc(x_13); lean_dec(x_11); x_14 = lean_box(0); -x_15 = l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(x_12, x_14); +x_15 = l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(x_12, x_14); x_16 = l_List_isEmpty___rarg(x_15); if (x_16 == 0) { @@ -11536,7 +11558,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_mkSimpContext___lambda__3(uint8_t x_ { uint8_t x_13; uint8_t x_14; x_13 = 2; -x_14 = l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_beqSimpKind____x40_Lean_Elab_Tactic_Simp___hyg_370_(x_1, x_13); +x_14 = l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_beqSimpKind____x40_Lean_Elab_Tactic_Simp___hyg_371_(x_1, x_13); if (x_14 == 0) { lean_object* x_15; lean_object* x_16; @@ -11618,7 +11640,7 @@ if (x_20 == 0) uint8_t x_21; uint8_t x_22; lean_dec(x_1); x_21 = 1; -x_22 = l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_beqSimpKind____x40_Lean_Elab_Tactic_Simp___hyg_370_(x_3, x_21); +x_22 = l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_beqSimpKind____x40_Lean_Elab_Tactic_Simp___hyg_371_(x_3, x_21); if (x_22 == 0) { lean_object* x_23; lean_object* x_24; @@ -12230,7 +12252,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSimp___closed__3() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__2; +x_1 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__2; x_2 = l___regBuiltin_Lean_Elab_Tactic_evalSimp___closed__2; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -12296,7 +12318,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSimp_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(278u); +x_1 = lean_unsigned_to_nat(281u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -12308,7 +12330,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSimp_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(281u); +x_1 = lean_unsigned_to_nat(284u); x_2 = lean_unsigned_to_nat(74u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -12336,7 +12358,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSimp_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(278u); +x_1 = lean_unsigned_to_nat(281u); x_2 = lean_unsigned_to_nat(45u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -12348,7 +12370,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSimp_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(278u); +x_1 = lean_unsigned_to_nat(281u); x_2 = lean_unsigned_to_nat(53u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -12629,7 +12651,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSimpAll_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(283u); +x_1 = lean_unsigned_to_nat(286u); x_2 = lean_unsigned_to_nat(44u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -12641,7 +12663,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSimpAll_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(287u); +x_1 = lean_unsigned_to_nat(290u); x_2 = lean_unsigned_to_nat(43u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -12669,7 +12691,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSimpAll_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(283u); +x_1 = lean_unsigned_to_nat(286u); x_2 = lean_unsigned_to_nat(48u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -12681,7 +12703,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSimpAll_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(283u); +x_1 = lean_unsigned_to_nat(286u); x_2 = lean_unsigned_to_nat(59u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13194,7 +13216,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalDSimp_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(306u); +x_1 = lean_unsigned_to_nat(309u); x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13206,7 +13228,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalDSimp_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(308u); +x_1 = lean_unsigned_to_nat(311u); x_2 = lean_unsigned_to_nat(46u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13234,7 +13256,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalDSimp_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(306u); +x_1 = lean_unsigned_to_nat(309u); x_2 = lean_unsigned_to_nat(46u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13246,7 +13268,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalDSimp_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(306u); +x_1 = lean_unsigned_to_nat(309u); x_2 = lean_unsigned_to_nat(55u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13329,22 +13351,22 @@ lean_dec_ref(res); res = initialize_Lean_Elab_Tactic_Config(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__1 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__1); -l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__2 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__2(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__2); -l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__3 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__3(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__3); -l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__4 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__4(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__4); -l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__5 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__5(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__5); -l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__6 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__6(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__6); -l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__7 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__7(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__7); -l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__8 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__8(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_5____closed__8); +l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__1 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__1); +l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__2 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__2); +l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__3 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__3); +l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__4 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__4(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__4); +l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__5 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__5(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__5); +l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__6 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__6(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__6); +l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__7 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__7(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__7); +l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__8 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__8(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6____closed__8); l_Lean_withoutModifyingState___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__12___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__13___closed__1 = _init_l_Lean_withoutModifyingState___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__12___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__13___closed__1(); lean_mark_persistent(l_Lean_withoutModifyingState___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__12___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__13___closed__1); l_Lean_withoutModifyingState___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__12___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__13___closed__2 = _init_l_Lean_withoutModifyingState___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__12___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__13___closed__2(); @@ -13367,22 +13389,22 @@ l_Lean_Elab_Tactic_elabSimpConfigCore___closed__2 = _init_l_Lean_Elab_Tactic_ela lean_mark_persistent(l_Lean_Elab_Tactic_elabSimpConfigCore___closed__2); l_Lean_Elab_Tactic_elabSimpConfigCore___closed__3 = _init_l_Lean_Elab_Tactic_elabSimpConfigCore___closed__3(); lean_mark_persistent(l_Lean_Elab_Tactic_elabSimpConfigCore___closed__3); -l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_123____closed__1 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_123____closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_123____closed__1); -l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_123____closed__2 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_123____closed__2(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_123____closed__2); +l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_124____closed__1 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_124____closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_124____closed__1); +l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_124____closed__2 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_124____closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_124____closed__2); l_Lean_Elab_Tactic_elabSimpConfigCtxCore___closed__1 = _init_l_Lean_Elab_Tactic_elabSimpConfigCtxCore___closed__1(); lean_mark_persistent(l_Lean_Elab_Tactic_elabSimpConfigCtxCore___closed__1); l_Lean_Elab_Tactic_elabSimpConfigCtxCore___closed__2 = _init_l_Lean_Elab_Tactic_elabSimpConfigCtxCore___closed__2(); lean_mark_persistent(l_Lean_Elab_Tactic_elabSimpConfigCtxCore___closed__2); l_Lean_Elab_Tactic_elabSimpConfigCtxCore___closed__3 = _init_l_Lean_Elab_Tactic_elabSimpConfigCtxCore___closed__3(); lean_mark_persistent(l_Lean_Elab_Tactic_elabSimpConfigCtxCore___closed__3); -l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_241____closed__1 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_241____closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_241____closed__1); -l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_241____closed__2 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_241____closed__2(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_241____closed__2); -l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_241____closed__3 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_241____closed__3(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_241____closed__3); +l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_242____closed__1 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_242____closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_242____closed__1); +l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_242____closed__2 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_242____closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_242____closed__2); +l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_242____closed__3 = _init_l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_242____closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_242____closed__3); l_Lean_Elab_Tactic_elabDSimpConfigCore___closed__1 = _init_l_Lean_Elab_Tactic_elabDSimpConfigCore___closed__1(); lean_mark_persistent(l_Lean_Elab_Tactic_elabDSimpConfigCore___closed__1); l_Lean_Elab_Tactic_elabDSimpConfigCore___closed__2 = _init_l_Lean_Elab_Tactic_elabDSimpConfigCore___closed__2(); @@ -13463,6 +13485,10 @@ l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs_resolveSimpId lean_mark_persistent(l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs_resolveSimpIdTheorem_x3f___closed__1); l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs_resolveSimpIdTheorem_x3f___closed__2 = _init_l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs_resolveSimpIdTheorem_x3f___closed__2(); lean_mark_persistent(l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs_resolveSimpIdTheorem_x3f___closed__2); +l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs_resolveSimpIdTheorem_x3f___closed__3 = _init_l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs_resolveSimpIdTheorem_x3f___closed__3(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs_resolveSimpIdTheorem_x3f___closed__3); +l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs_resolveSimpIdTheorem_x3f___closed__4 = _init_l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs_resolveSimpIdTheorem_x3f___closed__4(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs_resolveSimpIdTheorem_x3f___closed__4); l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs___spec__1___rarg___closed__1 = _init_l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs___spec__1___rarg___closed__1(); lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs___spec__1___rarg___closed__1); l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs___spec__1___rarg___closed__2 = _init_l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Tactic_Simp_0__Lean_Elab_Tactic_elabSimpArgs___spec__1___rarg___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Unfold.c b/stage0/stdlib/Lean/Elab/Tactic/Unfold.c index e7378d23409..9992bab46cf 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Unfold.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Unfold.c @@ -18,6 +18,7 @@ size_t lean_usize_add(size_t, size_t); lean_object* l_Lean_stringToMessageData(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalUnfold_declRange___closed__5; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_evalUnfold_go___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); lean_object* lean_name_mk_string(lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); LEAN_EXPORT lean_object* l_Lean_resolveGlobalConst___at_Lean_Elab_Tactic_evalUnfold_go___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -77,7 +78,6 @@ extern lean_object* l_Lean_Elab_Tactic_tacticElabAttribute; lean_object* l_Lean_Meta_throwTacticEx___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_unfoldLocalDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at_Lean_Elab_Tactic_evalUnfold_go___spec__5___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Syntax_getSepArgs(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_evalUnfold_go___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_resolveGlobalConstNoOverload___at_Lean_Elab_Tactic_evalUnfold_go___spec__1___closed__3; @@ -564,7 +564,7 @@ x_13 = lean_ctor_get(x_11, 1); lean_inc(x_13); lean_dec(x_11); x_14 = lean_box(0); -x_15 = l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(x_12, x_14); +x_15 = l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(x_12, x_14); x_16 = l_List_isEmpty___rarg(x_15); if (x_16 == 0) { diff --git a/stage0/stdlib/Lean/Elab/Term.c b/stage0/stdlib/Lean/Elab/Term.c index c801675bc77..0065c67b163 100644 --- a/stage0/stdlib/Lean/Elab/Term.c +++ b/stage0/stdlib/Lean/Elab/Term.c @@ -94,7 +94,6 @@ static lean_object* l_Lean_Elab_Term_instMetaEvalTermElabM___rarg___closed__16; static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_tryLiftAndCoe___closed__9; LEAN_EXPORT lean_object* l_Lean_Elab_Term_autoLift; LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutPending(lean_object*); -extern lean_object* l_Lean_nullKind; LEAN_EXPORT lean_object* l_Lean_Elab_Term_instAddErrorMessageContextTermElabM(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_State_infoState___default___closed__3; lean_object* l_Lean_Meta_whnfForall(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -372,6 +371,7 @@ static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_tryCoeSort___la LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_tryLiftAndCoe___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__15; static lean_object* l_Lean_Elab_Term_instInhabitedMVarErrorKind___closed__2; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15536____closed__3; static lean_object* l_Lean_Elab_Term_MVarErrorInfo_logError___closed__7; static lean_object* l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_observing___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -443,6 +443,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_LVal_isFieldName___boxed(lean_object*) LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLevel(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkTermElabAttributeUnsafe(lean_object*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_tryLiftAndCoe___closed__15; +static lean_object* l_Lean_Elab_Term_mkExplicitBinder___closed__16; LEAN_EXPORT lean_object* l_Lean_Elab_Term_withAutoBoundImplicit_loop___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__5___boxed(lean_object*, lean_object*, 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_Term_elabTermEnsuringType(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -475,12 +476,14 @@ static lean_object* l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__11; lean_object* l_Lean_MessageData_toString(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_tryCoeSort___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Term_observing(lean_object*); +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabImplicitLambda_loop___spec__1___rarg(lean_object*, uint8_t, 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_Term_withoutMacroStackAtErr(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__14(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_instMonadMacroAdapterTermElabM___closed__4; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_Term_0__Lean_Elab_Term_isLambdaWithImplicit___spec__1(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withMacroExpansion(lean_object*); +LEAN_EXPORT lean_object* l_List_filterTRAux___at_Lean_Elab_Term_resolveId_x3f___spec__2(lean_object*, lean_object*); lean_object* l_Lean_Elab_expandMacroImpl_x3f(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_decorateErrorMessageWithLambdaImplicitVars___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -554,6 +557,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkCoe(lean_object*, lean_object*, lean lean_object* l_Lean_mkAnnotation(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__7___boxed(lean_object*, 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_log___at_Lean_Elab_Term_traceAtCmdPos___spec__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_mkExplicitBinder___closed__15; static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__4___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__4___boxed(lean_object*, lean_object*, 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_Term_withAutoBoundImplicit(lean_object*); @@ -621,7 +625,7 @@ lean_object* l_Lean_Meta_getDelayedAssignment_x3f(lean_object*, lean_object*, le lean_object* lean_st_mk_ref(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_5964_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_5941_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13238_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13252_(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkAuxName___at_Lean_Elab_Term_mkAuxName___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_setElabConfig(lean_object*); static lean_object* l_Lean_Elab_Term_mkExplicitBinder___closed__14; @@ -636,7 +640,6 @@ static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_tryCoeSort___la LEAN_EXPORT lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Term_expandDeclId___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getId(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15522____closed__1; static lean_object* l_Lean_Elab_Tactic_instInhabitedSnapshot___closed__22; static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_tryCoeSort___lambda__1___closed__9; extern lean_object* l_Lean_Expr_instHashableExpr; @@ -685,6 +688,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Term_0_ static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_tryCoeSort___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_getSyntheticMVarDecl_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t lean_uint64_of_nat(lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13252____closed__2; static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isHole___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveId_x3f(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_setEnv___at_Lean_Elab_Term_evalExpr___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -827,10 +831,10 @@ LEAN_EXPORT lean_object* l_Lean_evalConst___at_Lean_Elab_Term_evalExpr___spec__1 lean_object* l_Std_PersistentArray_toList___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__7___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_getDeclName_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__12(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_instMonadInfoTreeTermElabM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveId_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13252____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instInhabitedCacheKey; LEAN_EXPORT lean_object* l_Lean_Elab_Term_levelMVarToParam_x27(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_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_instMetaEvalTermElabM___spec__6(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -912,6 +916,7 @@ static lean_object* l_Lean_Elab_Term_instMetaEvalTermElabM___rarg___closed__6; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Term_0__Lean_Elab_Term_withoutModifyingStateWithInfoAndMessagesImpl___spec__6(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_tryLiftAndCoe___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_instInhabitedSnapshot___closed__15; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15536____closed__2; static lean_object* l_Lean_Elab_Term_logUnassignedUsingErrorInfos___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_addDotCompletionInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); @@ -1074,6 +1079,7 @@ LEAN_EXPORT uint8_t l_Lean_Elab_Term_Context_ignoreTCFailures___default; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_expandDeclId___spec__4___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_getLetRecsToLift___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveName_process___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15536____closed__1; lean_object* lean_panic_fn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Term_expandDeclId___spec__2___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__6___closed__2; @@ -1117,7 +1123,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_throwTypeMismatchError___rarg(lean_obj LEAN_EXPORT lean_object* l_Lean_Elab_Term_addDotCompletionInfo___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Expr_instBEqExpr; LEAN_EXPORT lean_object* l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__1(lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15522____closed__2; static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_ContainsPendingMVar_visit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_mkExplicitBinder___closed__4; @@ -1165,12 +1170,11 @@ lean_object* l_IO_println___at_Lean_instEval___spec__1(lean_object*, lean_object static lean_object* l_Lean_Elab_Tactic_instInhabitedSnapshot___closed__17; LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Elab_Term_addTermInfo___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_throwErrorIfErrors(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15522_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15536_(lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentArray_mapM___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_evalExpr___spec__13(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_addAutoBoundImplicits(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_whnfD(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13238____closed__2; static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_tryLiftAndCoe___closed__11; lean_object* l_Std_RBNode_findCore___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimApp___spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Tactic_beqCacheKey____x40_Lean_Elab_Term___hyg_623____boxed(lean_object*, lean_object*); @@ -1198,7 +1202,6 @@ static lean_object* l_Lean_Elab_Term_MVarErrorInfo_logError___closed__9; static lean_object* l_Lean_Elab_Term_MVarErrorInfo_logError_addArgName___closed__1; static lean_object* l_Lean_Elab_Term_expandDeclId___closed__1; LEAN_EXPORT lean_object* l_List_mapM___at_Lean_Elab_Term_resolveName_x27___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13238____closed__1; static lean_object* l_Lean_Elab_Term_instInhabitedLetRecToLift___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Term_expandDeclId___spec__2___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_compile_decl(lean_object*, lean_object*, lean_object*); @@ -1341,10 +1344,8 @@ static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_tryLiftAndCoe__ static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingElabFns___closed__2; static lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Term_expandDeclId___spec__2___closed__3; uint8_t l_Lean_isStructure(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_filterAux___at_Lean_Elab_Term_resolveId_x3f___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_addTermInfo___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, 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_Term_TermElabM_toIO(lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15522____closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_applyAttributes___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_dropTermParens(lean_object*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_mkSyntheticSorryFor___closed__1; @@ -16743,7 +16744,7 @@ static lean_object* _init_l_Lean_Elab_Term_mkExplicitBinder___closed__6() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes(":", 1); +x_1 = lean_mk_string_from_bytes("null", 4); return x_1; } } @@ -16751,15 +16752,33 @@ static lean_object* _init_l_Lean_Elab_Term_mkExplicitBinder___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(2); +x_1 = lean_box(0); x_2 = l_Lean_Elab_Term_mkExplicitBinder___closed__6; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_mkExplicitBinder___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(":", 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_mkExplicitBinder___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(2); +x_2 = l_Lean_Elab_Term_mkExplicitBinder___closed__8; x_3 = lean_alloc_ctor(2, 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_Lean_Elab_Term_mkExplicitBinder___closed__8() { +static lean_object* _init_l_Lean_Elab_Term_mkExplicitBinder___closed__10() { _start: { lean_object* x_1; lean_object* x_2; @@ -16768,22 +16787,22 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_mkExplicitBinder___closed__9() { +static lean_object* _init_l_Lean_Elab_Term_mkExplicitBinder___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_mkExplicitBinder___closed__8; -x_2 = l_Lean_Elab_Term_mkExplicitBinder___closed__7; +x_1 = l_Lean_Elab_Term_mkExplicitBinder___closed__10; +x_2 = l_Lean_Elab_Term_mkExplicitBinder___closed__9; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_mkExplicitBinder___closed__10() { +static lean_object* _init_l_Lean_Elab_Term_mkExplicitBinder___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean_nullKind; +x_2 = l_Lean_Elab_Term_mkExplicitBinder___closed__7; x_3 = l_Lean_Elab_Term_instInhabitedLetRecToLift___closed__1; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -16792,7 +16811,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_mkExplicitBinder___closed__11() { +static lean_object* _init_l_Lean_Elab_Term_mkExplicitBinder___closed__13() { _start: { lean_object* x_1; @@ -16800,19 +16819,19 @@ x_1 = lean_mk_string_from_bytes(")", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_mkExplicitBinder___closed__12() { +static lean_object* _init_l_Lean_Elab_Term_mkExplicitBinder___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(2); -x_2 = l_Lean_Elab_Term_mkExplicitBinder___closed__11; +x_2 = l_Lean_Elab_Term_mkExplicitBinder___closed__13; x_3 = lean_alloc_ctor(2, 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_Lean_Elab_Term_mkExplicitBinder___closed__13() { +static lean_object* _init_l_Lean_Elab_Term_mkExplicitBinder___closed__15() { _start: { lean_object* x_1; lean_object* x_2; @@ -16821,11 +16840,11 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_mkExplicitBinder___closed__14() { +static lean_object* _init_l_Lean_Elab_Term_mkExplicitBinder___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_mkExplicitBinder___closed__13; +x_1 = l_Lean_Elab_Term_mkExplicitBinder___closed__15; x_2 = l_Lean_Elab_Term_mkExplicitBinder___closed__4; x_3 = lean_array_push(x_1, x_2); return x_3; @@ -16838,23 +16857,23 @@ lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_obj x_3 = l_Lean_Elab_Term_mkExplicitBinder___closed__5; x_4 = lean_array_push(x_3, x_1); x_5 = lean_box(2); -x_6 = l_Lean_nullKind; +x_6 = l_Lean_Elab_Term_mkExplicitBinder___closed__7; x_7 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_7, 0, x_5); lean_ctor_set(x_7, 1, x_6); lean_ctor_set(x_7, 2, x_4); -x_8 = l_Lean_Elab_Term_mkExplicitBinder___closed__9; +x_8 = l_Lean_Elab_Term_mkExplicitBinder___closed__11; x_9 = lean_array_push(x_8, x_2); x_10 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_10, 0, x_5); lean_ctor_set(x_10, 1, x_6); lean_ctor_set(x_10, 2, x_9); -x_11 = l_Lean_Elab_Term_mkExplicitBinder___closed__14; +x_11 = l_Lean_Elab_Term_mkExplicitBinder___closed__16; x_12 = lean_array_push(x_11, x_7); x_13 = lean_array_push(x_12, x_10); -x_14 = l_Lean_Elab_Term_mkExplicitBinder___closed__10; +x_14 = l_Lean_Elab_Term_mkExplicitBinder___closed__12; x_15 = lean_array_push(x_13, x_14); -x_16 = l_Lean_Elab_Term_mkExplicitBinder___closed__12; +x_16 = l_Lean_Elab_Term_mkExplicitBinder___closed__14; x_17 = lean_array_push(x_15, x_16); x_18 = l_Lean_Elab_Term_mkExplicitBinder___closed__2; x_19 = lean_alloc_ctor(1, 3, 0); @@ -22833,7 +22852,7 @@ return x_44; else { lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_59; lean_object* x_60; -x_45 = l_Lean_Elab_Term_mkExplicitBinder___closed__8; +x_45 = l_Lean_Elab_Term_mkExplicitBinder___closed__10; lean_inc(x_35); x_46 = lean_array_push(x_45, x_35); lean_inc(x_26); @@ -22921,7 +22940,7 @@ lean_ctor_set(x_78, 1, x_77); x_79 = l___private_Lean_Elab_Term_0__Lean_Elab_Term_tryLiftAndCoe___closed__5; lean_inc(x_78); x_80 = l_Lean_mkConst(x_79, x_78); -x_81 = l_Lean_Elab_Term_mkExplicitBinder___closed__13; +x_81 = l_Lean_Elab_Term_mkExplicitBinder___closed__15; lean_inc(x_35); x_82 = lean_array_push(x_81, x_35); lean_inc(x_26); @@ -23808,7 +23827,7 @@ return x_243; else { lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_258; lean_object* x_259; -x_244 = l_Lean_Elab_Term_mkExplicitBinder___closed__8; +x_244 = l_Lean_Elab_Term_mkExplicitBinder___closed__10; lean_inc(x_234); x_245 = lean_array_push(x_244, x_234); lean_inc(x_26); @@ -23896,7 +23915,7 @@ lean_ctor_set(x_277, 1, x_276); x_278 = l___private_Lean_Elab_Term_0__Lean_Elab_Term_tryLiftAndCoe___closed__5; lean_inc(x_277); x_279 = l_Lean_mkConst(x_278, x_277); -x_280 = l_Lean_Elab_Term_mkExplicitBinder___closed__13; +x_280 = l_Lean_Elab_Term_mkExplicitBinder___closed__15; lean_inc(x_234); x_281 = lean_array_push(x_280, x_234); lean_inc(x_26); @@ -24796,7 +24815,7 @@ return x_433; else { lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_448; lean_object* x_449; -x_434 = l_Lean_Elab_Term_mkExplicitBinder___closed__8; +x_434 = l_Lean_Elab_Term_mkExplicitBinder___closed__10; lean_inc(x_424); x_435 = lean_array_push(x_434, x_424); lean_inc(x_415); @@ -24884,7 +24903,7 @@ lean_ctor_set(x_467, 1, x_466); x_468 = l___private_Lean_Elab_Term_0__Lean_Elab_Term_tryLiftAndCoe___closed__5; lean_inc(x_467); x_469 = l_Lean_mkConst(x_468, x_467); -x_470 = l_Lean_Elab_Term_mkExplicitBinder___closed__13; +x_470 = l_Lean_Elab_Term_mkExplicitBinder___closed__15; lean_inc(x_424); x_471 = lean_array_push(x_470, x_424); lean_inc(x_415); @@ -31397,35 +31416,34 @@ return x_1; } else { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; x_4 = lean_unsigned_to_nat(1u); x_5 = l_Lean_Syntax_getArg(x_1, x_4); -x_6 = l_Lean_nullKind; -x_7 = lean_unsigned_to_nat(2u); +x_6 = lean_unsigned_to_nat(2u); lean_inc(x_5); -x_8 = l_Lean_Syntax_isNodeOf(x_5, x_6, x_7); -if (x_8 == 0) +x_7 = l_Lean_Syntax_matchesNull(x_5, x_6); +if (x_7 == 0) { lean_dec(x_5); return x_1; } else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_9 = lean_unsigned_to_nat(0u); -x_10 = l_Lean_Syntax_getArg(x_5, x_9); -x_11 = l_Lean_Syntax_getArg(x_5, x_4); +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_unsigned_to_nat(0u); +x_9 = l_Lean_Syntax_getArg(x_5, x_8); +x_10 = l_Lean_Syntax_getArg(x_5, x_4); lean_dec(x_5); -x_12 = l_Lean_Syntax_isNodeOf(x_11, x_6, x_9); -if (x_12 == 0) +x_11 = l_Lean_Syntax_matchesNull(x_10, x_8); +if (x_11 == 0) { -lean_dec(x_10); +lean_dec(x_9); return x_1; } else { lean_dec(x_1); -x_1 = x_10; +x_1 = x_9; goto _start; } } @@ -31699,44 +31717,43 @@ return x_4; } else { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; x_5 = lean_unsigned_to_nat(1u); x_6 = l_Lean_Syntax_getArg(x_1, x_5); lean_dec(x_1); -x_7 = l_Lean_nullKind; -x_8 = lean_unsigned_to_nat(2u); +x_7 = lean_unsigned_to_nat(2u); lean_inc(x_6); -x_9 = l_Lean_Syntax_isNodeOf(x_6, x_7, x_8); -if (x_9 == 0) +x_8 = l_Lean_Syntax_matchesNull(x_6, x_7); +if (x_8 == 0) { -uint8_t x_10; +uint8_t x_9; lean_dec(x_6); -x_10 = 0; -return x_10; +x_9 = 0; +return x_9; } else { -lean_object* x_11; uint8_t x_12; -x_11 = l_Lean_Syntax_getArg(x_6, x_5); +lean_object* x_10; uint8_t x_11; +x_10 = l_Lean_Syntax_getArg(x_6, x_5); lean_dec(x_6); -lean_inc(x_11); -x_12 = l_Lean_Syntax_isNodeOf(x_11, x_7, x_5); -if (x_12 == 0) +lean_inc(x_10); +x_11 = l_Lean_Syntax_matchesNull(x_10, x_5); +if (x_11 == 0) { -uint8_t x_13; -lean_dec(x_11); -x_13 = 0; -return x_13; +uint8_t x_12; +lean_dec(x_10); +x_12 = 0; +return x_12; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_14 = lean_unsigned_to_nat(0u); -x_15 = l_Lean_Syntax_getArg(x_11, x_14); -lean_dec(x_11); -x_16 = l___private_Lean_Elab_Term_0__Lean_Elab_Term_isTypeAscription___closed__2; -x_17 = l_Lean_Syntax_isOfKind(x_15, x_16); -return x_17; +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_13 = lean_unsigned_to_nat(0u); +x_14 = l_Lean_Syntax_getArg(x_10, x_13); +lean_dec(x_10); +x_15 = l___private_Lean_Elab_Term_0__Lean_Elab_Term_isTypeAscription___closed__2; +x_16 = l_Lean_Syntax_isOfKind(x_14, x_15); +return x_16; } } } @@ -38618,7 +38635,7 @@ lean_ctor_set(x_23, 1, x_22); x_24 = l___private_Lean_Elab_Term_0__Lean_Elab_Term_tryCoeSort___closed__2; lean_inc(x_23); x_25 = l_Lean_mkConst(x_24, x_23); -x_26 = l_Lean_Elab_Term_mkExplicitBinder___closed__8; +x_26 = l_Lean_Elab_Term_mkExplicitBinder___closed__10; lean_inc(x_1); x_27 = lean_array_push(x_26, x_1); lean_inc(x_13); @@ -41163,7 +41180,7 @@ lean_dec(x_3); return x_10; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13238____closed__1() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13252____closed__1() { _start: { lean_object* x_1; @@ -41171,21 +41188,21 @@ x_1 = lean_mk_string_from_bytes("letrec", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13238____closed__2() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13252____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Term_MVarErrorInfo_logError___closed__11; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13238____closed__1; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13252____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13238_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13252_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13238____closed__2; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13252____closed__2; x_3 = l_Lean_registerTraceClass(x_2, x_1); return x_3; } @@ -41482,7 +41499,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_isLetRecAuxMVar(lean_object* x_1, lean _start: { lean_object* x_9; uint8_t x_10; lean_object* x_11; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; -x_9 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13238____closed__2; +x_9 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13252____closed__2; x_38 = lean_st_ref_get(x_7, x_8); x_39 = lean_ctor_get(x_38, 0); lean_inc(x_39); @@ -43708,7 +43725,7 @@ return x_22; } } } -LEAN_EXPORT lean_object* l_List_filterAux___at_Lean_Elab_Term_resolveId_x3f___spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_List_filterTRAux___at_Lean_Elab_Term_resolveId_x3f___spec__2(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -43892,7 +43909,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveId_x3f___lambda__2(uint8_t x_1, { lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_box(0); -x_13 = l_List_filterAux___at_Lean_Elab_Term_resolveId_x3f___spec__2(x_4, x_12); +x_13 = l_List_filterTRAux___at_Lean_Elab_Term_resolveId_x3f___spec__2(x_4, x_12); x_14 = l_List_mapTRAux___at_Lean_Elab_Term_resolveId_x3f___spec__3(x_13, x_12); if (lean_obj_tag(x_14) == 0) { @@ -51673,7 +51690,7 @@ x_3 = lean_alloc_closure((void*)(l_Lean_Elab_withoutModifyingStateWithInfoAndMes return x_3; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15522____closed__1() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15536____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -51683,7 +51700,7 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15522____closed__2() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15536____closed__2() { _start: { lean_object* x_1; @@ -51691,17 +51708,17 @@ x_1 = lean_mk_string_from_bytes("debug", 5); return x_1; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15522____closed__3() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15536____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Term_MVarErrorInfo_logError___closed__11; -x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15522____closed__2; +x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15536____closed__2; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15522_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15536_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; @@ -51713,7 +51730,7 @@ lean_object* x_4; lean_object* x_5; lean_object* x_6; x_4 = lean_ctor_get(x_3, 1); lean_inc(x_4); lean_dec(x_3); -x_5 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15522____closed__1; +x_5 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15536____closed__1; x_6 = l_Lean_registerTraceClass(x_5, x_4); if (lean_obj_tag(x_6) == 0) { @@ -51721,7 +51738,7 @@ lean_object* x_7; lean_object* x_8; lean_object* x_9; x_7 = lean_ctor_get(x_6, 1); lean_inc(x_7); lean_dec(x_6); -x_8 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15522____closed__3; +x_8 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15536____closed__3; x_9 = l_Lean_registerTraceClass(x_8, x_7); return x_9; } @@ -52213,6 +52230,10 @@ l_Lean_Elab_Term_mkExplicitBinder___closed__13 = _init_l_Lean_Elab_Term_mkExplic lean_mark_persistent(l_Lean_Elab_Term_mkExplicitBinder___closed__13); l_Lean_Elab_Term_mkExplicitBinder___closed__14 = _init_l_Lean_Elab_Term_mkExplicitBinder___closed__14(); lean_mark_persistent(l_Lean_Elab_Term_mkExplicitBinder___closed__14); +l_Lean_Elab_Term_mkExplicitBinder___closed__15 = _init_l_Lean_Elab_Term_mkExplicitBinder___closed__15(); +lean_mark_persistent(l_Lean_Elab_Term_mkExplicitBinder___closed__15); +l_Lean_Elab_Term_mkExplicitBinder___closed__16 = _init_l_Lean_Elab_Term_mkExplicitBinder___closed__16(); +lean_mark_persistent(l_Lean_Elab_Term_mkExplicitBinder___closed__16); l_Lean_Elab_Term_levelMVarToParam___closed__1 = _init_l_Lean_Elab_Term_levelMVarToParam___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_levelMVarToParam___closed__1); l_Lean_Elab_Term_levelMVarToParam___closed__2 = _init_l_Lean_Elab_Term_levelMVarToParam___closed__2(); @@ -52552,11 +52573,11 @@ l_Lean_Elab_Term_mkAuxName___closed__1 = _init_l_Lean_Elab_Term_mkAuxName___clos lean_mark_persistent(l_Lean_Elab_Term_mkAuxName___closed__1); l_Lean_Elab_Term_mkAuxName___closed__2 = _init_l_Lean_Elab_Term_mkAuxName___closed__2(); lean_mark_persistent(l_Lean_Elab_Term_mkAuxName___closed__2); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13238____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13238____closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13238____closed__1); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13238____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13238____closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13238____closed__2); -res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13238_(lean_io_mk_world()); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13252____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13252____closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13252____closed__1); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13252____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13252____closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13252____closed__2); +res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13252_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); l_Lean_Elab_Term_isLetRecAuxMVar___lambda__2___closed__1 = _init_l_Lean_Elab_Term_isLetRecAuxMVar___lambda__2___closed__1(); @@ -52729,13 +52750,13 @@ l_Lean_Elab_Term_expandDeclId___closed__1 = _init_l_Lean_Elab_Term_expandDeclId_ lean_mark_persistent(l_Lean_Elab_Term_expandDeclId___closed__1); l_Lean_Elab_Term_expandDeclId___closed__2 = _init_l_Lean_Elab_Term_expandDeclId___closed__2(); lean_mark_persistent(l_Lean_Elab_Term_expandDeclId___closed__2); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15522____closed__1 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15522____closed__1(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15522____closed__1); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15522____closed__2 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15522____closed__2(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15522____closed__2); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15522____closed__3 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15522____closed__3(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15522____closed__3); -res = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15522_(lean_io_mk_world()); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15536____closed__1 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15536____closed__1(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15536____closed__1); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15536____closed__2 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15536____closed__2(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15536____closed__2); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15536____closed__3 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15536____closed__3(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15536____closed__3); +res = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_15536_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/Util.c b/stage0/stdlib/Lean/Elab/Util.c index c1ec572d2b1..8d475210718 100644 --- a/stage0/stdlib/Lean/Elab/Util.c +++ b/stage0/stdlib/Lean/Elab/Util.c @@ -14,9 +14,7 @@ extern "C" { #endif lean_object* l_List_reverse___rarg(lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__2; lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__4; size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_nestedExceptionToMessageData___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_liftMacroM___spec__3(lean_object*); @@ -28,6 +26,7 @@ lean_object* l_Lean_throwMaxRecDepthAt___rarg(lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_mkElabAttribute___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Elab_addMacroStack___spec__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__2; static lean_object* l_Lean_Elab_mkElabAttribute___rarg___lambda__3___closed__3; static lean_object* l_Lean_Elab_mkMacroAttributeUnsafe___closed__7; lean_object* l_Lean_Macro_getCurrNamespace(lean_object*, lean_object*); @@ -60,12 +59,12 @@ LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___rarg___lambda__2(lean_object*, uint8_t lean_name_eq(lean_object*, lean_object*); static lean_object* l_Lean_Elab_logException___rarg___lambda__1___closed__3; static lean_object* l_Lean_Elab_mkElabAttribute___rarg___lambda__2___closed__15; -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Elab_getBetterRef___spec__1___boxed(lean_object*, lean_object*); lean_object* l_Lean_Elab_throwUnsupportedSyntax___rarg(lean_object*); static lean_object* l_Lean_Elab_mkElabAttribute___rarg___lambda__2___closed__5; static lean_object* l_Lean_Elab_mkElabAttribute___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_throwErrorWithNestedErrors___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Elab_getBetterRef___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_mkElabAttribute___rarg___lambda__2___closed__14; lean_object* lean_array_get_size(lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); @@ -79,7 +78,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_instMonadMacroAdapter___rarg___lambda__1(le LEAN_EXPORT lean_object* l_Lean_Elab_expandMacroImpl_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_checkSyntaxNodeKind___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_expandMacroImpl_x3f___closed__2; -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Elab_getBetterRef___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_logException___rarg___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_logException(lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); @@ -87,6 +85,7 @@ LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Elab_liftMacroM___spec__1(lean_o static lean_object* l_Lean_Elab_nestedExceptionToMessageData___rarg___lambda__1___closed__2; lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOptionM___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_liftMacroM___spec__3___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_addMacroStack___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___rarg___lambda__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -110,6 +109,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_checkSyntaxNodeKindAtNamespaces(lean_object LEAN_EXPORT lean_object* l_Lean_Elab_expandMacroImpl_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Elab_getBetterRef___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_expandMacroImpl_x3f(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Elab_getBetterRef___spec__1___boxed(lean_object*, lean_object*); lean_object* l_Lean_addTrace___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Elab_liftMacroM___spec__2___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_logException___rarg___lambda__1___closed__4; @@ -137,9 +137,9 @@ lean_object* l_Lean_Name_toExprAux(lean_object*); static lean_object* l_Lean_Elab_mkElabAttribute___rarg___lambda__2___closed__17; lean_object* l_Nat_repr(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_expandMacroImpl_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2386_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_864_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2390_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_868_(lean_object*); static lean_object* l_Lean_Elab_mkElabAttribute___rarg___lambda__2___closed__8; lean_object* l_Lean_logTrace___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_registerInitAttrUnsafe___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); @@ -176,17 +176,18 @@ static lean_object* l_Lean_Elab_syntaxNodeKindOfAttrParam___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_getBetterRef___lambda__1___boxed(lean_object*); static lean_object* l_Lean_Elab_mkElabAttribute___rarg___lambda__2___closed__19; size_t lean_usize_of_nat(lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___rarg___lambda__5___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Util_0__Lean_Elab_evalSyntaxConstantUnsafe___closed__1; lean_object* l_Lean_Environment_evalConstCheck___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_addMacroStack___rarg___lambda__1___closed__3; static lean_object* l_Lean_Elab_throwErrorWithNestedErrors___rarg___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_logDbgTrace(lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__3; LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Elab_liftMacroM___spec__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_unsetTrailing(lean_object*); static lean_object* l_Lean_Elab_syntaxNodeKindOfAttrParam___closed__1; lean_object* l_Lean_Syntax_getPos_x3f(lean_object*, uint8_t); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__6; static lean_object* l_Lean_Elab_addMacroStack___rarg___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___rarg___lambda__6___boxed(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_trace___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -200,7 +201,6 @@ static lean_object* l_List_foldl___at_Lean_Elab_addMacroStack___spec__1___closed lean_object* lean_environment_main_module(lean_object*); static lean_object* l_Lean_Elab_checkSyntaxNodeKind___rarg___lambda__1___closed__2; static lean_object* l_Lean_Elab_expandOptNamedPrio___closed__6; -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__6; LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Elab_addMacroStack___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_mkElabAttribute(lean_object*); static lean_object* l_Lean_Elab_mkElabAttribute___rarg___lambda__2___closed__11; @@ -254,11 +254,9 @@ LEAN_EXPORT lean_object* l_Lean_Elab_nestedExceptionToMessageData___rarg___lambd lean_object* l_Lean_mkNatLit(lean_object*); lean_object* l_Lean_mkStrLit(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_checkSyntaxNodeKindAtCurrentNamespaces(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2386____closed__2; static lean_object* l_Lean_Elab_mkMacroAttributeUnsafe___closed__2; lean_object* l_Lean_log___rarg(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_liftMacroM___spec__3___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2386____closed__1; LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_expandMacroImpl_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_expandOptNamedPrio___closed__4; uint8_t l_List_isEmpty___rarg(lean_object*); @@ -270,11 +268,13 @@ LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___rarg___lambda__1(lean_object*, LEAN_EXPORT lean_object* l_Lean_Elab_addMacroStack___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_addMacroStack___rarg___lambda__1___closed__1; static lean_object* l_Lean_Elab_addMacroStack___rarg___lambda__1___closed__2; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2390____closed__1; lean_object* l_Lean_mkConst(lean_object*, lean_object*); static lean_object* l_Lean_Elab_syntaxNodeKindOfAttrParam___closed__4; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2390____closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_nestedExceptionToMessageData___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Util_0__Lean_Elab_evalSyntaxConstantUnsafe___closed__2; -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__1; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_syntaxNodeKindOfAttrParam(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_evalPrio(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_mkMacroAttributeUnsafe___closed__4; @@ -284,8 +284,8 @@ lean_object* l_Lean_Syntax_reprint(lean_object*); static lean_object* l_Lean_Elab_mkElabAttribute___rarg___lambda__2___closed__16; LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_liftMacroM___spec__3___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwErrorWithNestedErrors___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__5; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_throwErrorWithNestedErrors___spec__1(lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_logDbgTrace___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Exception_toMessageData(lean_object*); static lean_object* l_Lean_Elab_addMacroStack___rarg___lambda__1___closed__4; @@ -525,7 +525,7 @@ lean_dec(x_1); return x_4; } } -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Elab_getBetterRef___spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Elab_getBetterRef___spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -570,7 +570,7 @@ x_2 = lean_ctor_get(x_1, 0); x_3 = 0; x_4 = l_Lean_Syntax_getPos_x3f(x_2, x_3); x_5 = lean_box(0); -x_6 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Elab_getBetterRef___spec__1(x_4, x_5); +x_6 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Elab_getBetterRef___spec__1(x_4, x_5); lean_dec(x_4); if (x_6 == 0) { @@ -631,11 +631,11 @@ return x_1; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Elab_getBetterRef___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Elab_getBetterRef___spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Elab_getBetterRef___spec__1(x_1, x_2); +x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Elab_getBetterRef___spec__1(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -661,7 +661,7 @@ lean_dec(x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__1() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__1() { _start: { lean_object* x_1; @@ -669,17 +669,17 @@ x_1 = lean_mk_string_from_bytes("pp", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__2() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__1; +x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__3() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__3() { _start: { lean_object* x_1; @@ -687,17 +687,17 @@ x_1 = lean_mk_string_from_bytes("macroStack", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__4() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__2; -x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__3; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__2; +x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__5() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__5() { _start: { lean_object* x_1; @@ -705,13 +705,13 @@ x_1 = lean_mk_string_from_bytes("dispaly macro expansion stack", 29); return x_1; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__6() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__6() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = 0; -x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__1; -x_3 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__5; +x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__1; +x_3 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__5; x_4 = lean_box(x_1); x_5 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_5, 0, x_4); @@ -720,12 +720,12 @@ lean_ctor_set(x_5, 2, x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__4; -x_3 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__6; +x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__4; +x_3 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__6; x_4 = l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_54____spec__1(x_2, x_3, x_1); return x_4; } @@ -2109,7 +2109,7 @@ x_8 = l_Lean_Elab_mkElabAttribute___rarg(x_2, x_3, x_4, x_5, x_6, x_7, x_1); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_864_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_868_(lean_object* x_1) { _start: { lean_object* x_2; @@ -4368,7 +4368,7 @@ x_8 = l_Array_mapMUnsafe_map___at_Lean_Elab_throwErrorWithNestedErrors___spec__1 return x_8; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2386____closed__1() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2390____closed__1() { _start: { lean_object* x_1; @@ -4376,17 +4376,17 @@ x_1 = lean_mk_string_from_bytes("step", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2386____closed__2() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2390____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_logDbgTrace___rarg___closed__1; -x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2386____closed__1; +x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2390____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2386_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2390_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; @@ -4398,7 +4398,7 @@ lean_object* x_4; lean_object* x_5; lean_object* x_6; x_4 = lean_ctor_get(x_3, 1); lean_inc(x_4); lean_dec(x_3); -x_5 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2386____closed__2; +x_5 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2390____closed__2; x_6 = l_Lean_registerTraceClass(x_5, x_4); return x_6; } @@ -4489,19 +4489,19 @@ l_Lean_Elab_expandOptNamedPrio___closed__8 = _init_l_Lean_Elab_expandOptNamedPri lean_mark_persistent(l_Lean_Elab_expandOptNamedPrio___closed__8); l_Lean_Elab_getBetterRef___closed__1 = _init_l_Lean_Elab_getBetterRef___closed__1(); lean_mark_persistent(l_Lean_Elab_getBetterRef___closed__1); -l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__1 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__1(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__1); -l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__2 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__2(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__2); -l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__3 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__3(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__3); -l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__4 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__4(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__4); -l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__5 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__5(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__5); -l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__6 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__6(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236____closed__6); -if (builtin) {res = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_236_(lean_io_mk_world()); +l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__1 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__1(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__1); +l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__2 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__2(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__2); +l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__3 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__3(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__3); +l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__4 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__4(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__4); +l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__5 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__5(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__5); +l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__6 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__6(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240____closed__6); +if (builtin) {res = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_240_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Elab_pp_macroStack = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Elab_pp_macroStack); @@ -4612,7 +4612,7 @@ l_Lean_Elab_mkMacroAttributeUnsafe___closed__9 = _init_l_Lean_Elab_mkMacroAttrib lean_mark_persistent(l_Lean_Elab_mkMacroAttributeUnsafe___closed__9); l_Lean_Elab_mkMacroAttributeUnsafe___closed__10 = _init_l_Lean_Elab_mkMacroAttributeUnsafe___closed__10(); lean_mark_persistent(l_Lean_Elab_mkMacroAttributeUnsafe___closed__10); -if (builtin) {res = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_864_(lean_io_mk_world()); +if (builtin) {res = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_868_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Elab_macroAttribute = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Elab_macroAttribute); @@ -4647,11 +4647,11 @@ l_Lean_Elab_throwErrorWithNestedErrors___rarg___lambda__1___closed__1 = _init_l_ lean_mark_persistent(l_Lean_Elab_throwErrorWithNestedErrors___rarg___lambda__1___closed__1); l_Lean_Elab_throwErrorWithNestedErrors___rarg___lambda__1___closed__2 = _init_l_Lean_Elab_throwErrorWithNestedErrors___rarg___lambda__1___closed__2(); lean_mark_persistent(l_Lean_Elab_throwErrorWithNestedErrors___rarg___lambda__1___closed__2); -l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2386____closed__1 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2386____closed__1(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2386____closed__1); -l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2386____closed__2 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2386____closed__2(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2386____closed__2); -res = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2386_(lean_io_mk_world()); +l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2390____closed__1 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2390____closed__1(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2390____closed__1); +l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2390____closed__2 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2390____closed__2(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2390____closed__2); +res = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2390_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Exception.c b/stage0/stdlib/Lean/Exception.c index 00dcec0d353..bab4fe449c8 100644 --- a/stage0/stdlib/Lean/Exception.c +++ b/stage0/stdlib/Lean/Exception.c @@ -35,7 +35,6 @@ static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThr static lean_object* l_Lean_termThrowError_______closed__17; static lean_object* l_Lean_termThrowError_______closed__1; LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___rarg(lean_object*, lean_object*, lean_object*); -uint8_t lean_name_eq(lean_object*, lean_object*); static lean_object* l_Lean_termThrowErrorAt_________closed__8; static lean_object* l_Lean_throwMaxRecDepthAt___rarg___closed__2; lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(lean_object*, lean_object*); @@ -44,7 +43,6 @@ LEAN_EXPORT lean_object* l_Lean_throwErrorAt___rarg___lambda__1(lean_object*, le LEAN_EXPORT lean_object* l_Lean_termThrowError____; lean_object* lean_array_push(lean_object*, lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___rarg___closed__1; -extern lean_object* l_Lean_interpolatedStrKind; static lean_object* l_Lean_termThrowError_______closed__4; LEAN_EXPORT lean_object* l_Lean_throwKernelException(lean_object*, lean_object*); static lean_object* l_Lean_termThrowErrorAt_________closed__1; @@ -109,12 +107,12 @@ static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThr LEAN_EXPORT lean_object* l_Lean_throwError(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedException; static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__2; +static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__29; static lean_object* l_Lean_termThrowError_______closed__10; static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__22; LEAN_EXPORT lean_object* l_Lean_withIncRecDepth(lean_object*, lean_object*); static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__3; static lean_object* l_Lean_termThrowError_______closed__7; -lean_object* l_Lean_Syntax_getKind(lean_object*); static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__6; LEAN_EXPORT lean_object* l_Lean_throwKernelException___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___rarg___lambda__1___boxed(lean_object*, lean_object*); @@ -137,6 +135,7 @@ LEAN_EXPORT lean_object* l_Lean_throwError___rarg___lambda__2(lean_object*, lean LEAN_EXPORT lean_object* l_Lean_instMonadRecDepthStateRefT_x27___rarg(lean_object*); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__14; +static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__28; static lean_object* l_Lean_termThrowErrorAt_________closed__4; static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__16; LEAN_EXPORT lean_object* l_Lean_ofExcept(lean_object*, lean_object*, lean_object*); @@ -1161,7 +1160,7 @@ static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__t _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Parser", 6); +x_1 = lean_mk_string_from_bytes("interpolatedStrKind", 19); return x_1; } } @@ -1169,7 +1168,7 @@ static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__t _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_termThrowError_______closed__2; +x_1 = lean_box(0); x_2 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -1179,7 +1178,7 @@ static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__t _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Term", 4); +x_1 = lean_mk_string_from_bytes("Parser", 6); return x_1; } } @@ -1187,7 +1186,7 @@ static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__t _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__2; +x_1 = l_Lean_termThrowError_______closed__2; x_2 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -1197,7 +1196,7 @@ static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__t _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("app", 3); +x_1 = lean_mk_string_from_bytes("Term", 4); return x_1; } } @@ -1215,26 +1214,44 @@ static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__t _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Lean.throwError", 15); +x_1 = lean_mk_string_from_bytes("app", 3); return x_1; } } static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__8() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__6; +x_2 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__7; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Lean.throwError", 15); +return x_1; +} +} +static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__10() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__7; +x_1 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__9; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__9() { +static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__7; +x_1 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__9; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__8; +x_3 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__10; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1242,7 +1259,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__10() { +static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__12() { _start: { lean_object* x_1; @@ -1250,41 +1267,41 @@ x_1 = lean_mk_string_from_bytes("throwError", 10); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__11() { +static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_termThrowError_______closed__2; -x_2 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__10; +x_2 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__12; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__12() { +static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__11; +x_2 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__13; 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___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__13() { +static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__12; +x_2 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__14; x_3 = lean_alloc_ctor(1, 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___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__14() { +static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__16() { _start: { lean_object* x_1; @@ -1292,17 +1309,17 @@ x_1 = lean_mk_string_from_bytes("null", 4); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__15() { +static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__14; +x_2 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__16; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__16() { +static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__18() { _start: { lean_object* x_1; lean_object* x_2; @@ -1311,7 +1328,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__17() { +static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__19() { _start: { lean_object* x_1; lean_object* x_2; @@ -1320,7 +1337,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__18() { +static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__20() { _start: { lean_object* x_1; @@ -1328,17 +1345,17 @@ x_1 = lean_mk_string_from_bytes("paren", 5); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__19() { +static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__4; -x_2 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__18; +x_1 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__6; +x_2 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__20; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__20() { +static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__22() { _start: { lean_object* x_1; @@ -1346,7 +1363,7 @@ x_1 = lean_mk_string_from_bytes("(", 1); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__21() { +static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__23() { _start: { lean_object* x_1; @@ -1354,17 +1371,17 @@ x_1 = lean_mk_string_from_bytes("termM!_", 7); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__22() { +static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_termThrowError_______closed__2; -x_2 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__21; +x_2 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__23; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__23() { +static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__25() { _start: { lean_object* x_1; @@ -1372,7 +1389,7 @@ x_1 = lean_mk_string_from_bytes("m!", 2); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__24() { +static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__26() { _start: { lean_object* x_1; lean_object* x_2; @@ -1381,13 +1398,13 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__25() { +static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__27() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__15; -x_3 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__24; +x_2 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__17; +x_3 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__26; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1395,7 +1412,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__26() { +static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__28() { _start: { lean_object* x_1; @@ -1403,7 +1420,7 @@ x_1 = lean_mk_string_from_bytes(")", 1); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__27() { +static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__29() { _start: { lean_object* x_1; lean_object* x_2; @@ -1432,263 +1449,261 @@ return x_7; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; x_8 = lean_unsigned_to_nat(1u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); lean_dec(x_1); +x_10 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__2; lean_inc(x_9); -x_10 = l_Lean_Syntax_getKind(x_9); -x_11 = l_Lean_interpolatedStrKind; -x_12 = lean_name_eq(x_10, x_11); -lean_dec(x_10); -if (x_12 == 0) +x_11 = l_Lean_Syntax_isOfKind(x_9, x_10); +if (x_11 == 0) { -lean_object* x_13; uint8_t x_14; +lean_object* x_12; uint8_t x_13; lean_inc(x_2); -x_13 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_14 = !lean_is_exclusive(x_13); -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; lean_object* 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; lean_object* x_32; -x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_ctor_get(x_2, 2); +x_12 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_13 = !lean_is_exclusive(x_12); +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; 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; +x_14 = lean_ctor_get(x_12, 0); +x_15 = lean_ctor_get(x_2, 2); +lean_inc(x_15); +x_16 = lean_ctor_get(x_2, 1); lean_inc(x_16); -x_17 = lean_ctor_get(x_2, 1); -lean_inc(x_17); lean_dec(x_2); -x_18 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__11; -x_19 = l_Lean_addMacroScope(x_17, x_18, x_16); -x_20 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__9; -x_21 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__13; -x_22 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_22, 0, x_15); -lean_ctor_set(x_22, 1, x_20); -lean_ctor_set(x_22, 2, x_19); -lean_ctor_set(x_22, 3, x_21); -x_23 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__16; -x_24 = lean_array_push(x_23, x_9); -x_25 = lean_box(2); -x_26 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__15; -x_27 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_27, 0, x_25); -lean_ctor_set(x_27, 1, x_26); -lean_ctor_set(x_27, 2, x_24); -x_28 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__17; -x_29 = lean_array_push(x_28, x_22); -x_30 = lean_array_push(x_29, x_27); -x_31 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__6; -x_32 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_32, 0, x_25); -lean_ctor_set(x_32, 1, x_31); -lean_ctor_set(x_32, 2, x_30); -lean_ctor_set(x_13, 0, x_32); -return x_13; +x_17 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__13; +x_18 = l_Lean_addMacroScope(x_16, x_17, x_15); +x_19 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__11; +x_20 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__15; +x_21 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_21, 0, x_14); +lean_ctor_set(x_21, 1, x_19); +lean_ctor_set(x_21, 2, x_18); +lean_ctor_set(x_21, 3, x_20); +x_22 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__18; +x_23 = lean_array_push(x_22, x_9); +x_24 = lean_box(2); +x_25 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__17; +x_26 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +lean_ctor_set(x_26, 2, x_23); +x_27 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__19; +x_28 = lean_array_push(x_27, x_21); +x_29 = lean_array_push(x_28, x_26); +x_30 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__8; +x_31 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_31, 0, x_24); +lean_ctor_set(x_31, 1, x_30); +lean_ctor_set(x_31, 2, x_29); +lean_ctor_set(x_12, 0, x_31); +return x_12; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_33 = lean_ctor_get(x_13, 0); -x_34 = lean_ctor_get(x_13, 1); -lean_inc(x_34); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_32 = lean_ctor_get(x_12, 0); +x_33 = lean_ctor_get(x_12, 1); lean_inc(x_33); -lean_dec(x_13); -x_35 = lean_ctor_get(x_2, 2); +lean_inc(x_32); +lean_dec(x_12); +x_34 = lean_ctor_get(x_2, 2); +lean_inc(x_34); +x_35 = lean_ctor_get(x_2, 1); lean_inc(x_35); -x_36 = lean_ctor_get(x_2, 1); -lean_inc(x_36); lean_dec(x_2); -x_37 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__11; -x_38 = l_Lean_addMacroScope(x_36, x_37, x_35); -x_39 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__9; -x_40 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__13; -x_41 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_41, 0, x_33); -lean_ctor_set(x_41, 1, x_39); -lean_ctor_set(x_41, 2, x_38); -lean_ctor_set(x_41, 3, x_40); -x_42 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__16; -x_43 = lean_array_push(x_42, x_9); -x_44 = lean_box(2); -x_45 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__15; -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_43); -x_47 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__17; -x_48 = lean_array_push(x_47, x_41); -x_49 = lean_array_push(x_48, x_46); -x_50 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__6; -x_51 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_51, 0, x_44); -lean_ctor_set(x_51, 1, x_50); -lean_ctor_set(x_51, 2, x_49); -x_52 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_52, 0, x_51); -lean_ctor_set(x_52, 1, x_34); -return x_52; +x_36 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__13; +x_37 = l_Lean_addMacroScope(x_35, x_36, x_34); +x_38 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__11; +x_39 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__15; +x_40 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_40, 0, x_32); +lean_ctor_set(x_40, 1, x_38); +lean_ctor_set(x_40, 2, x_37); +lean_ctor_set(x_40, 3, x_39); +x_41 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__18; +x_42 = lean_array_push(x_41, x_9); +x_43 = lean_box(2); +x_44 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__17; +x_45 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_42); +x_46 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__19; +x_47 = lean_array_push(x_46, x_40); +x_48 = lean_array_push(x_47, x_45); +x_49 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__8; +x_50 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_50, 0, x_43); +lean_ctor_set(x_50, 1, x_49); +lean_ctor_set(x_50, 2, x_48); +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_33); +return x_51; } } else { -lean_object* x_53; uint8_t x_54; +lean_object* x_52; uint8_t x_53; lean_inc(x_2); -x_53 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_54 = !lean_is_exclusive(x_53); -if (x_54 == 0) +x_52 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_53 = !lean_is_exclusive(x_52); +if (x_53 == 0) { -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; -x_55 = lean_ctor_get(x_53, 0); -x_56 = lean_ctor_get(x_2, 2); +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_54 = lean_ctor_get(x_52, 0); +x_55 = lean_ctor_get(x_2, 2); +lean_inc(x_55); +x_56 = lean_ctor_get(x_2, 1); lean_inc(x_56); -x_57 = lean_ctor_get(x_2, 1); -lean_inc(x_57); lean_dec(x_2); -x_58 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__11; -x_59 = l_Lean_addMacroScope(x_57, x_58, x_56); -x_60 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__9; -x_61 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__13; -lean_inc(x_55); -x_62 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_62, 0, x_55); -lean_ctor_set(x_62, 1, x_60); -lean_ctor_set(x_62, 2, x_59); -lean_ctor_set(x_62, 3, x_61); -x_63 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__20; -lean_inc(x_55); -x_64 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_64, 0, x_55); -lean_ctor_set(x_64, 1, x_63); -x_65 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__23; -lean_inc(x_55); -x_66 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_66, 0, x_55); -lean_ctor_set(x_66, 1, x_65); -x_67 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__17; -x_68 = lean_array_push(x_67, x_66); -x_69 = lean_array_push(x_68, x_9); -x_70 = lean_box(2); -x_71 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__22; -x_72 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_72, 0, x_70); -lean_ctor_set(x_72, 1, x_71); -lean_ctor_set(x_72, 2, x_69); -x_73 = lean_array_push(x_67, x_72); -x_74 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__25; -x_75 = lean_array_push(x_73, x_74); -x_76 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__15; -x_77 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_77, 0, x_70); -lean_ctor_set(x_77, 1, x_76); -lean_ctor_set(x_77, 2, x_75); -x_78 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__26; -x_79 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_79, 0, x_55); -lean_ctor_set(x_79, 1, x_78); -x_80 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__27; -x_81 = lean_array_push(x_80, x_64); -x_82 = lean_array_push(x_81, x_77); -x_83 = lean_array_push(x_82, x_79); -x_84 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__19; -x_85 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_85, 0, x_70); -lean_ctor_set(x_85, 1, x_84); -lean_ctor_set(x_85, 2, x_83); -x_86 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__16; -x_87 = lean_array_push(x_86, x_85); -x_88 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_88, 0, x_70); -lean_ctor_set(x_88, 1, x_76); -lean_ctor_set(x_88, 2, x_87); -x_89 = lean_array_push(x_67, x_62); -x_90 = lean_array_push(x_89, x_88); -x_91 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__6; -x_92 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_92, 0, x_70); -lean_ctor_set(x_92, 1, x_91); -lean_ctor_set(x_92, 2, x_90); -lean_ctor_set(x_53, 0, x_92); -return x_53; +x_57 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__13; +x_58 = l_Lean_addMacroScope(x_56, x_57, x_55); +x_59 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__11; +x_60 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__15; +lean_inc(x_54); +x_61 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_61, 0, x_54); +lean_ctor_set(x_61, 1, x_59); +lean_ctor_set(x_61, 2, x_58); +lean_ctor_set(x_61, 3, x_60); +x_62 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__22; +lean_inc(x_54); +x_63 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_63, 0, x_54); +lean_ctor_set(x_63, 1, x_62); +x_64 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__25; +lean_inc(x_54); +x_65 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_65, 0, x_54); +lean_ctor_set(x_65, 1, x_64); +x_66 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__19; +x_67 = lean_array_push(x_66, x_65); +x_68 = lean_array_push(x_67, x_9); +x_69 = lean_box(2); +x_70 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__24; +x_71 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_71, 0, x_69); +lean_ctor_set(x_71, 1, x_70); +lean_ctor_set(x_71, 2, x_68); +x_72 = lean_array_push(x_66, x_71); +x_73 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__27; +x_74 = lean_array_push(x_72, x_73); +x_75 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__17; +x_76 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_76, 0, x_69); +lean_ctor_set(x_76, 1, x_75); +lean_ctor_set(x_76, 2, x_74); +x_77 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__28; +x_78 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_78, 0, x_54); +lean_ctor_set(x_78, 1, x_77); +x_79 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__29; +x_80 = lean_array_push(x_79, x_63); +x_81 = lean_array_push(x_80, x_76); +x_82 = lean_array_push(x_81, x_78); +x_83 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__21; +x_84 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_84, 0, x_69); +lean_ctor_set(x_84, 1, x_83); +lean_ctor_set(x_84, 2, x_82); +x_85 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__18; +x_86 = lean_array_push(x_85, x_84); +x_87 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_87, 0, x_69); +lean_ctor_set(x_87, 1, x_75); +lean_ctor_set(x_87, 2, x_86); +x_88 = lean_array_push(x_66, x_61); +x_89 = lean_array_push(x_88, x_87); +x_90 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__8; +x_91 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_91, 0, x_69); +lean_ctor_set(x_91, 1, x_90); +lean_ctor_set(x_91, 2, x_89); +lean_ctor_set(x_52, 0, x_91); +return x_52; } else { -lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; -x_93 = lean_ctor_get(x_53, 0); -x_94 = lean_ctor_get(x_53, 1); -lean_inc(x_94); +lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; +x_92 = lean_ctor_get(x_52, 0); +x_93 = lean_ctor_get(x_52, 1); lean_inc(x_93); -lean_dec(x_53); -x_95 = lean_ctor_get(x_2, 2); +lean_inc(x_92); +lean_dec(x_52); +x_94 = lean_ctor_get(x_2, 2); +lean_inc(x_94); +x_95 = lean_ctor_get(x_2, 1); lean_inc(x_95); -x_96 = lean_ctor_get(x_2, 1); -lean_inc(x_96); lean_dec(x_2); -x_97 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__11; -x_98 = l_Lean_addMacroScope(x_96, x_97, x_95); -x_99 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__9; -x_100 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__13; -lean_inc(x_93); -x_101 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_101, 0, x_93); -lean_ctor_set(x_101, 1, x_99); -lean_ctor_set(x_101, 2, x_98); -lean_ctor_set(x_101, 3, x_100); -x_102 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__20; -lean_inc(x_93); -x_103 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_103, 0, x_93); -lean_ctor_set(x_103, 1, x_102); -x_104 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__23; -lean_inc(x_93); -x_105 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_105, 0, x_93); -lean_ctor_set(x_105, 1, x_104); -x_106 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__17; -x_107 = lean_array_push(x_106, x_105); -x_108 = lean_array_push(x_107, x_9); -x_109 = lean_box(2); -x_110 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__22; -x_111 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_111, 0, x_109); -lean_ctor_set(x_111, 1, x_110); -lean_ctor_set(x_111, 2, x_108); -x_112 = lean_array_push(x_106, x_111); -x_113 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__25; -x_114 = lean_array_push(x_112, x_113); -x_115 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__15; -x_116 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_116, 0, x_109); -lean_ctor_set(x_116, 1, x_115); -lean_ctor_set(x_116, 2, x_114); -x_117 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__26; -x_118 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_118, 0, x_93); -lean_ctor_set(x_118, 1, x_117); -x_119 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__27; -x_120 = lean_array_push(x_119, x_103); -x_121 = lean_array_push(x_120, x_116); -x_122 = lean_array_push(x_121, x_118); -x_123 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__19; -x_124 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_124, 0, x_109); -lean_ctor_set(x_124, 1, x_123); -lean_ctor_set(x_124, 2, x_122); -x_125 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__16; -x_126 = lean_array_push(x_125, x_124); -x_127 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_127, 0, x_109); -lean_ctor_set(x_127, 1, x_115); -lean_ctor_set(x_127, 2, x_126); -x_128 = lean_array_push(x_106, x_101); -x_129 = lean_array_push(x_128, x_127); -x_130 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__6; -x_131 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_131, 0, x_109); -lean_ctor_set(x_131, 1, x_130); -lean_ctor_set(x_131, 2, x_129); -x_132 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_132, 0, x_131); -lean_ctor_set(x_132, 1, x_94); -return x_132; +x_96 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__13; +x_97 = l_Lean_addMacroScope(x_95, x_96, x_94); +x_98 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__11; +x_99 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__15; +lean_inc(x_92); +x_100 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_100, 0, x_92); +lean_ctor_set(x_100, 1, x_98); +lean_ctor_set(x_100, 2, x_97); +lean_ctor_set(x_100, 3, x_99); +x_101 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__22; +lean_inc(x_92); +x_102 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_102, 0, x_92); +lean_ctor_set(x_102, 1, x_101); +x_103 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__25; +lean_inc(x_92); +x_104 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_104, 0, x_92); +lean_ctor_set(x_104, 1, x_103); +x_105 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__19; +x_106 = lean_array_push(x_105, x_104); +x_107 = lean_array_push(x_106, x_9); +x_108 = lean_box(2); +x_109 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__24; +x_110 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_110, 0, x_108); +lean_ctor_set(x_110, 1, x_109); +lean_ctor_set(x_110, 2, x_107); +x_111 = lean_array_push(x_105, x_110); +x_112 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__27; +x_113 = lean_array_push(x_111, x_112); +x_114 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__17; +x_115 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_115, 0, x_108); +lean_ctor_set(x_115, 1, x_114); +lean_ctor_set(x_115, 2, x_113); +x_116 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__28; +x_117 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_117, 0, x_92); +lean_ctor_set(x_117, 1, x_116); +x_118 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__29; +x_119 = lean_array_push(x_118, x_102); +x_120 = lean_array_push(x_119, x_115); +x_121 = lean_array_push(x_120, x_117); +x_122 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__21; +x_123 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_123, 0, x_108); +lean_ctor_set(x_123, 1, x_122); +lean_ctor_set(x_123, 2, x_121); +x_124 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__18; +x_125 = lean_array_push(x_124, x_123); +x_126 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_126, 0, x_108); +lean_ctor_set(x_126, 1, x_114); +lean_ctor_set(x_126, 2, x_125); +x_127 = lean_array_push(x_105, x_100); +x_128 = lean_array_push(x_127, x_126); +x_129 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__8; +x_130 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_130, 0, x_108); +lean_ctor_set(x_130, 1, x_129); +lean_ctor_set(x_130, 2, x_128); +x_131 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_131, 0, x_130); +lean_ctor_set(x_131, 1, x_93); +return x_131; } } } @@ -1787,265 +1802,263 @@ return x_7; } 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; uint8_t x_14; +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; x_8 = lean_unsigned_to_nat(1u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); x_10 = lean_unsigned_to_nat(2u); x_11 = l_Lean_Syntax_getArg(x_1, x_10); lean_dec(x_1); +x_12 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__2; lean_inc(x_11); -x_12 = l_Lean_Syntax_getKind(x_11); -x_13 = l_Lean_interpolatedStrKind; -x_14 = lean_name_eq(x_12, x_13); -lean_dec(x_12); -if (x_14 == 0) +x_13 = l_Lean_Syntax_isOfKind(x_11, x_12); +if (x_13 == 0) { -lean_object* x_15; uint8_t x_16; +lean_object* x_14; uint8_t x_15; lean_inc(x_2); -x_15 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) +x_14 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) { -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; 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_object* x_32; lean_object* x_33; lean_object* x_34; -x_17 = lean_ctor_get(x_15, 0); -x_18 = lean_ctor_get(x_2, 2); +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; 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_object* x_32; lean_object* x_33; +x_16 = lean_ctor_get(x_14, 0); +x_17 = lean_ctor_get(x_2, 2); +lean_inc(x_17); +x_18 = lean_ctor_get(x_2, 1); lean_inc(x_18); -x_19 = lean_ctor_get(x_2, 1); -lean_inc(x_19); lean_dec(x_2); -x_20 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__5; -x_21 = l_Lean_addMacroScope(x_19, x_20, x_18); -x_22 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__3; -x_23 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__7; -x_24 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_24, 0, x_17); -lean_ctor_set(x_24, 1, x_22); -lean_ctor_set(x_24, 2, x_21); -lean_ctor_set(x_24, 3, x_23); -x_25 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__17; -x_26 = lean_array_push(x_25, x_9); -x_27 = lean_array_push(x_26, x_11); -x_28 = lean_box(2); -x_29 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__15; -x_30 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -lean_ctor_set(x_30, 2, x_27); -x_31 = lean_array_push(x_25, x_24); -x_32 = lean_array_push(x_31, x_30); -x_33 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__6; -x_34 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_34, 0, x_28); -lean_ctor_set(x_34, 1, x_33); -lean_ctor_set(x_34, 2, x_32); -lean_ctor_set(x_15, 0, x_34); -return x_15; +x_19 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__5; +x_20 = l_Lean_addMacroScope(x_18, x_19, x_17); +x_21 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__3; +x_22 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__7; +x_23 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_23, 0, x_16); +lean_ctor_set(x_23, 1, x_21); +lean_ctor_set(x_23, 2, x_20); +lean_ctor_set(x_23, 3, x_22); +x_24 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__19; +x_25 = lean_array_push(x_24, x_9); +x_26 = lean_array_push(x_25, x_11); +x_27 = lean_box(2); +x_28 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__17; +x_29 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +lean_ctor_set(x_29, 2, x_26); +x_30 = lean_array_push(x_24, x_23); +x_31 = lean_array_push(x_30, x_29); +x_32 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__8; +x_33 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_33, 0, x_27); +lean_ctor_set(x_33, 1, x_32); +lean_ctor_set(x_33, 2, x_31); +lean_ctor_set(x_14, 0, x_33); +return x_14; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_35 = lean_ctor_get(x_15, 0); -x_36 = lean_ctor_get(x_15, 1); -lean_inc(x_36); +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_34 = lean_ctor_get(x_14, 0); +x_35 = lean_ctor_get(x_14, 1); lean_inc(x_35); -lean_dec(x_15); -x_37 = lean_ctor_get(x_2, 2); +lean_inc(x_34); +lean_dec(x_14); +x_36 = lean_ctor_get(x_2, 2); +lean_inc(x_36); +x_37 = lean_ctor_get(x_2, 1); lean_inc(x_37); -x_38 = lean_ctor_get(x_2, 1); -lean_inc(x_38); lean_dec(x_2); -x_39 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__5; -x_40 = l_Lean_addMacroScope(x_38, x_39, x_37); -x_41 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__3; -x_42 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__7; -x_43 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_43, 0, x_35); -lean_ctor_set(x_43, 1, x_41); -lean_ctor_set(x_43, 2, x_40); -lean_ctor_set(x_43, 3, x_42); -x_44 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__17; -x_45 = lean_array_push(x_44, x_9); -x_46 = lean_array_push(x_45, x_11); -x_47 = lean_box(2); -x_48 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__15; -x_49 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_49, 0, x_47); -lean_ctor_set(x_49, 1, x_48); -lean_ctor_set(x_49, 2, x_46); -x_50 = lean_array_push(x_44, x_43); -x_51 = lean_array_push(x_50, x_49); -x_52 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__6; -x_53 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_53, 0, x_47); -lean_ctor_set(x_53, 1, x_52); -lean_ctor_set(x_53, 2, x_51); -x_54 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_36); -return x_54; +x_38 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__5; +x_39 = l_Lean_addMacroScope(x_37, x_38, x_36); +x_40 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__3; +x_41 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__7; +x_42 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_42, 0, x_34); +lean_ctor_set(x_42, 1, x_40); +lean_ctor_set(x_42, 2, x_39); +lean_ctor_set(x_42, 3, x_41); +x_43 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__19; +x_44 = lean_array_push(x_43, x_9); +x_45 = lean_array_push(x_44, x_11); +x_46 = lean_box(2); +x_47 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__17; +x_48 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_48, 0, x_46); +lean_ctor_set(x_48, 1, x_47); +lean_ctor_set(x_48, 2, x_45); +x_49 = lean_array_push(x_43, x_42); +x_50 = lean_array_push(x_49, x_48); +x_51 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__8; +x_52 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_52, 0, x_46); +lean_ctor_set(x_52, 1, x_51); +lean_ctor_set(x_52, 2, x_50); +x_53 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_35); +return x_53; } } else { -lean_object* x_55; uint8_t x_56; +lean_object* x_54; uint8_t x_55; lean_inc(x_2); -x_55 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_56 = !lean_is_exclusive(x_55); -if (x_56 == 0) +x_54 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); +x_55 = !lean_is_exclusive(x_54); +if (x_55 == 0) { -lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; -x_57 = lean_ctor_get(x_55, 0); -x_58 = lean_ctor_get(x_2, 2); +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_56 = lean_ctor_get(x_54, 0); +x_57 = lean_ctor_get(x_2, 2); +lean_inc(x_57); +x_58 = lean_ctor_get(x_2, 1); lean_inc(x_58); -x_59 = lean_ctor_get(x_2, 1); -lean_inc(x_59); lean_dec(x_2); -x_60 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__5; -x_61 = l_Lean_addMacroScope(x_59, x_60, x_58); -x_62 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__3; -x_63 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__7; -lean_inc(x_57); -x_64 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_64, 0, x_57); -lean_ctor_set(x_64, 1, x_62); -lean_ctor_set(x_64, 2, x_61); -lean_ctor_set(x_64, 3, x_63); -x_65 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__20; -lean_inc(x_57); -x_66 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_66, 0, x_57); -lean_ctor_set(x_66, 1, x_65); -x_67 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__23; -lean_inc(x_57); -x_68 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_68, 0, x_57); -lean_ctor_set(x_68, 1, x_67); -x_69 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__17; -x_70 = lean_array_push(x_69, x_68); -x_71 = lean_array_push(x_70, x_11); -x_72 = lean_box(2); -x_73 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__22; -x_74 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_74, 0, x_72); -lean_ctor_set(x_74, 1, x_73); -lean_ctor_set(x_74, 2, x_71); -x_75 = lean_array_push(x_69, x_74); -x_76 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__25; -x_77 = lean_array_push(x_75, x_76); -x_78 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__15; -x_79 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_79, 0, x_72); -lean_ctor_set(x_79, 1, x_78); -lean_ctor_set(x_79, 2, x_77); -x_80 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__26; -x_81 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_81, 0, x_57); -lean_ctor_set(x_81, 1, x_80); -x_82 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__27; -x_83 = lean_array_push(x_82, x_66); -x_84 = lean_array_push(x_83, x_79); -x_85 = lean_array_push(x_84, x_81); -x_86 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__19; -x_87 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_87, 0, x_72); -lean_ctor_set(x_87, 1, x_86); -lean_ctor_set(x_87, 2, x_85); -x_88 = lean_array_push(x_69, x_9); -x_89 = lean_array_push(x_88, x_87); -x_90 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_90, 0, x_72); -lean_ctor_set(x_90, 1, x_78); -lean_ctor_set(x_90, 2, x_89); -x_91 = lean_array_push(x_69, x_64); -x_92 = lean_array_push(x_91, x_90); -x_93 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__6; -x_94 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_94, 0, x_72); -lean_ctor_set(x_94, 1, x_93); -lean_ctor_set(x_94, 2, x_92); -lean_ctor_set(x_55, 0, x_94); -return x_55; +x_59 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__5; +x_60 = l_Lean_addMacroScope(x_58, x_59, x_57); +x_61 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__3; +x_62 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__7; +lean_inc(x_56); +x_63 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_63, 0, x_56); +lean_ctor_set(x_63, 1, x_61); +lean_ctor_set(x_63, 2, x_60); +lean_ctor_set(x_63, 3, x_62); +x_64 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__22; +lean_inc(x_56); +x_65 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_65, 0, x_56); +lean_ctor_set(x_65, 1, x_64); +x_66 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__25; +lean_inc(x_56); +x_67 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_67, 0, x_56); +lean_ctor_set(x_67, 1, x_66); +x_68 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__19; +x_69 = lean_array_push(x_68, x_67); +x_70 = lean_array_push(x_69, x_11); +x_71 = lean_box(2); +x_72 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__24; +x_73 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_73, 0, x_71); +lean_ctor_set(x_73, 1, x_72); +lean_ctor_set(x_73, 2, x_70); +x_74 = lean_array_push(x_68, x_73); +x_75 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__27; +x_76 = lean_array_push(x_74, x_75); +x_77 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__17; +x_78 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_78, 0, x_71); +lean_ctor_set(x_78, 1, x_77); +lean_ctor_set(x_78, 2, x_76); +x_79 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__28; +x_80 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_80, 0, x_56); +lean_ctor_set(x_80, 1, x_79); +x_81 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__29; +x_82 = lean_array_push(x_81, x_65); +x_83 = lean_array_push(x_82, x_78); +x_84 = lean_array_push(x_83, x_80); +x_85 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__21; +x_86 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_86, 0, x_71); +lean_ctor_set(x_86, 1, x_85); +lean_ctor_set(x_86, 2, x_84); +x_87 = lean_array_push(x_68, x_9); +x_88 = lean_array_push(x_87, x_86); +x_89 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_89, 0, x_71); +lean_ctor_set(x_89, 1, x_77); +lean_ctor_set(x_89, 2, x_88); +x_90 = lean_array_push(x_68, x_63); +x_91 = lean_array_push(x_90, x_89); +x_92 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__8; +x_93 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_93, 0, x_71); +lean_ctor_set(x_93, 1, x_92); +lean_ctor_set(x_93, 2, x_91); +lean_ctor_set(x_54, 0, x_93); +return x_54; } else { -lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; -x_95 = lean_ctor_get(x_55, 0); -x_96 = lean_ctor_get(x_55, 1); -lean_inc(x_96); +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; +x_94 = lean_ctor_get(x_54, 0); +x_95 = lean_ctor_get(x_54, 1); lean_inc(x_95); -lean_dec(x_55); -x_97 = lean_ctor_get(x_2, 2); +lean_inc(x_94); +lean_dec(x_54); +x_96 = lean_ctor_get(x_2, 2); +lean_inc(x_96); +x_97 = lean_ctor_get(x_2, 1); lean_inc(x_97); -x_98 = lean_ctor_get(x_2, 1); -lean_inc(x_98); lean_dec(x_2); -x_99 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__5; -x_100 = l_Lean_addMacroScope(x_98, x_99, x_97); -x_101 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__3; -x_102 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__7; -lean_inc(x_95); -x_103 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_103, 0, x_95); -lean_ctor_set(x_103, 1, x_101); -lean_ctor_set(x_103, 2, x_100); -lean_ctor_set(x_103, 3, x_102); -x_104 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__20; -lean_inc(x_95); -x_105 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_105, 0, x_95); -lean_ctor_set(x_105, 1, x_104); -x_106 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__23; -lean_inc(x_95); -x_107 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_107, 0, x_95); -lean_ctor_set(x_107, 1, x_106); -x_108 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__17; -x_109 = lean_array_push(x_108, x_107); -x_110 = lean_array_push(x_109, x_11); -x_111 = lean_box(2); -x_112 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__22; -x_113 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_113, 0, x_111); -lean_ctor_set(x_113, 1, x_112); -lean_ctor_set(x_113, 2, x_110); -x_114 = lean_array_push(x_108, x_113); -x_115 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__25; -x_116 = lean_array_push(x_114, x_115); -x_117 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__15; -x_118 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_118, 0, x_111); -lean_ctor_set(x_118, 1, x_117); -lean_ctor_set(x_118, 2, x_116); -x_119 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__26; -x_120 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_120, 0, x_95); -lean_ctor_set(x_120, 1, x_119); -x_121 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__27; -x_122 = lean_array_push(x_121, x_105); -x_123 = lean_array_push(x_122, x_118); -x_124 = lean_array_push(x_123, x_120); -x_125 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__19; -x_126 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_126, 0, x_111); -lean_ctor_set(x_126, 1, x_125); -lean_ctor_set(x_126, 2, x_124); -x_127 = lean_array_push(x_108, x_9); -x_128 = lean_array_push(x_127, x_126); -x_129 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_129, 0, x_111); -lean_ctor_set(x_129, 1, x_117); -lean_ctor_set(x_129, 2, x_128); -x_130 = lean_array_push(x_108, x_103); -x_131 = lean_array_push(x_130, x_129); -x_132 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__6; -x_133 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_133, 0, x_111); -lean_ctor_set(x_133, 1, x_132); -lean_ctor_set(x_133, 2, x_131); -x_134 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_134, 0, x_133); -lean_ctor_set(x_134, 1, x_96); -return x_134; +x_98 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__5; +x_99 = l_Lean_addMacroScope(x_97, x_98, x_96); +x_100 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__3; +x_101 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__7; +lean_inc(x_94); +x_102 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_102, 0, x_94); +lean_ctor_set(x_102, 1, x_100); +lean_ctor_set(x_102, 2, x_99); +lean_ctor_set(x_102, 3, x_101); +x_103 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__22; +lean_inc(x_94); +x_104 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_104, 0, x_94); +lean_ctor_set(x_104, 1, x_103); +x_105 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__25; +lean_inc(x_94); +x_106 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_106, 0, x_94); +lean_ctor_set(x_106, 1, x_105); +x_107 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__19; +x_108 = lean_array_push(x_107, x_106); +x_109 = lean_array_push(x_108, x_11); +x_110 = lean_box(2); +x_111 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__24; +x_112 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_112, 0, x_110); +lean_ctor_set(x_112, 1, x_111); +lean_ctor_set(x_112, 2, x_109); +x_113 = lean_array_push(x_107, x_112); +x_114 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__27; +x_115 = lean_array_push(x_113, x_114); +x_116 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__17; +x_117 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_117, 0, x_110); +lean_ctor_set(x_117, 1, x_116); +lean_ctor_set(x_117, 2, x_115); +x_118 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__28; +x_119 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_119, 0, x_94); +lean_ctor_set(x_119, 1, x_118); +x_120 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__29; +x_121 = lean_array_push(x_120, x_104); +x_122 = lean_array_push(x_121, x_117); +x_123 = lean_array_push(x_122, x_119); +x_124 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__21; +x_125 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_125, 0, x_110); +lean_ctor_set(x_125, 1, x_124); +lean_ctor_set(x_125, 2, x_123); +x_126 = lean_array_push(x_107, x_9); +x_127 = lean_array_push(x_126, x_125); +x_128 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_128, 0, x_110); +lean_ctor_set(x_128, 1, x_116); +lean_ctor_set(x_128, 2, x_127); +x_129 = lean_array_push(x_107, x_102); +x_130 = lean_array_push(x_129, x_128); +x_131 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__8; +x_132 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_132, 0, x_110); +lean_ctor_set(x_132, 1, x_131); +lean_ctor_set(x_132, 2, x_130); +x_133 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_133, 0, x_132); +lean_ctor_set(x_133, 1, x_95); +return x_133; } } } @@ -2206,6 +2219,10 @@ l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___clo lean_mark_persistent(l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__26); l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__27 = _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__27(); lean_mark_persistent(l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__27); +l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__28 = _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__28(); +lean_mark_persistent(l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__28); +l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__29 = _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__29(); +lean_mark_persistent(l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__29); l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__1 = _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__1(); lean_mark_persistent(l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__1); l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__2 = _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__2(); diff --git a/stage0/stdlib/Lean/Expr.c b/stage0/stdlib/Lean/Expr.c index d1c3619e717..e4293c066c5 100644 --- a/stage0/stdlib/Lean/Expr.c +++ b/stage0/stdlib/Lean/Expr.c @@ -456,7 +456,6 @@ lean_object* lean_expr_abstract_range(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Expr_constLevels_x21___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_getArg_x21___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Expr_bindingDomain_x21___closed__1; -lean_object* l_panic___at_String_toNat_x21___spec__1(lean_object*); static lean_object* l___private_Lean_Expr_0__Lean_reprExpr____x40_Lean_Expr___hyg_2336____closed__41; static lean_object* l_Lean_mkInaccessible___closed__2; LEAN_EXPORT lean_object* l_Lean_isLHSGoal_x3f___boxed(lean_object*); @@ -749,6 +748,7 @@ lean_object* lean_expr_update_lambda(lean_object*, uint8_t, lean_object*, lean_o static lean_object* l___private_Lean_Expr_0__Lean_reprExpr____x40_Lean_Expr___hyg_2336____closed__27; static lean_object* l_Lean_Expr_mkData___closed__4; static lean_object* l___private_Lean_Expr_0__Lean_reprExpr____x40_Lean_Expr___hyg_2336____closed__9; +lean_object* l_panic___at_Lean_TSyntax_getNat___spec__1(lean_object*); LEAN_EXPORT lean_object* lean_expr_mk_sort(lean_object*); LEAN_EXPORT lean_object* l_Lean_Literal_hash___boxed(lean_object*); lean_object* lean_level_update_succ(lean_object*, lean_object*); @@ -10936,7 +10936,7 @@ else { lean_object* x_3; lean_object* x_4; x_3 = l_Lean_Expr_bvarIdx_x21___closed__3; -x_4 = l_panic___at_String_toNat_x21___spec__1(x_3); +x_4 = l_panic___at_Lean_TSyntax_getNat___spec__1(x_3); return x_4; } } @@ -11691,7 +11691,7 @@ else { lean_object* x_3; lean_object* x_4; x_3 = l_Lean_Expr_projIdx_x21___closed__2; -x_4 = l_panic___at_String_toNat_x21___spec__1(x_3); +x_4 = l_panic___at_Lean_TSyntax_getNat___spec__1(x_3); return x_4; } } diff --git a/stage0/stdlib/Lean/Hygiene.c b/stage0/stdlib/Lean/Hygiene.c index b2bbb04635e..91faa6c94c2 100644 --- a/stage0/stdlib/Lean/Hygiene.c +++ b/stage0/stdlib/Lean/Hygiene.c @@ -24,6 +24,7 @@ lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Hygiene_0__Lean_mkInaccessibleUserName(uint8_t, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Hygiene___hyg_332____closed__5; LEAN_EXPORT uint8_t l_Lean_getSanitizeNames(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Hygiene_0__Lean_sanitizeSyntaxAux___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_sanitizeName(lean_object*, lean_object*); lean_object* l_Lean_mkIdentFrom(lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Hygiene___hyg_332____closed__3; @@ -1073,6 +1074,17 @@ goto _start; } } } +LEAN_EXPORT lean_object* l___private_Lean_Hygiene_0__Lean_sanitizeSyntaxAux___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; +x_4 = l_Lean_mkIdentFrom(x_1, x_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___private_Lean_Hygiene_0__Lean_sanitizeSyntaxAux(lean_object* x_1, lean_object* x_2) { _start: { @@ -1170,63 +1182,41 @@ uint8_t x_29; x_29 = l_Lean_Name_hasMacroScopes(x_26); if (x_29 == 0) { -lean_object* x_30; lean_object* x_31; -x_30 = l_Lean_mkIdentFrom(x_1, x_26); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_2); -return x_31; -} -else -{ -lean_object* x_32; uint8_t x_33; -x_32 = l_Lean_sanitizeName(x_26, x_2); -x_33 = !lean_is_exclusive(x_32); -if (x_33 == 0) -{ -lean_object* x_34; lean_object* x_35; -x_34 = lean_ctor_get(x_32, 0); -x_35 = l_Lean_mkIdentFrom(x_1, x_34); -lean_ctor_set(x_32, 0, x_35); -return x_32; +lean_object* x_30; +x_30 = l___private_Lean_Hygiene_0__Lean_sanitizeSyntaxAux___lambda__1(x_1, x_26, x_2); +return x_30; } else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_36 = lean_ctor_get(x_32, 0); -x_37 = lean_ctor_get(x_32, 1); -lean_inc(x_37); -lean_inc(x_36); -lean_dec(x_32); -x_38 = l_Lean_mkIdentFrom(x_1, x_36); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_37); -return x_39; -} +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_31 = l_Lean_sanitizeName(x_26, x_2); +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec(x_31); +x_34 = l___private_Lean_Hygiene_0__Lean_sanitizeSyntaxAux___lambda__1(x_1, x_32, x_33); +return x_34; } } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; +lean_object* x_35; lean_object* x_36; lean_dec(x_26); -x_40 = lean_ctor_get(x_28, 0); -lean_inc(x_40); +x_35 = lean_ctor_get(x_28, 0); +lean_inc(x_35); lean_dec(x_28); -x_41 = l_Lean_mkIdentFrom(x_1, x_40); -x_42 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_2); -return x_42; +x_36 = l___private_Lean_Hygiene_0__Lean_sanitizeSyntaxAux___lambda__1(x_1, x_35, x_2); +return x_36; } } default: { -lean_object* x_43; -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_1); -lean_ctor_set(x_43, 1, x_2); -return x_43; +lean_object* x_37; +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_1); +lean_ctor_set(x_37, 1, x_2); +return x_37; } } } diff --git a/stage0/stdlib/Lean/KeyedDeclsAttribute.c b/stage0/stdlib/Lean/KeyedDeclsAttribute.c index 89de986d51a..9ae323efaaa 100644 --- a/stage0/stdlib/Lean/KeyedDeclsAttribute.c +++ b/stage0/stdlib/Lean/KeyedDeclsAttribute.c @@ -82,7 +82,6 @@ LEAN_EXPORT lean_object* l_Std_AssocList_find_x3f___at_Lean_KeyedDeclsAttribute_ static lean_object* l_Lean_KeyedDeclsAttribute_init___rarg___lambda__11___closed__2; static lean_object* l_Lean_KeyedDeclsAttribute_init___rarg___lambda__10___closed__3; lean_object* l_id___rarg___boxed(lean_object*); -LEAN_EXPORT lean_object* l_List_filterAux___at_Lean_KeyedDeclsAttribute_getEntries___spec__12(lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_KeyedDeclsAttribute_init___rarg___lambda__4___boxed(lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); @@ -90,6 +89,7 @@ static lean_object* l_Lean_KeyedDeclsAttribute_init___rarg___lambda__9___closed_ LEAN_EXPORT lean_object* l_Std_PersistentHashMap_findAux___at___private_Lean_KeyedDeclsAttribute_0__Lean_KeyedDeclsAttribute_Table_insert___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_mkHashMap___at_Lean_KeyedDeclsAttribute_instInhabitedExtensionState___spec__1___rarg___boxed(lean_object*); static lean_object* l_Lean_KeyedDeclsAttribute_init___rarg___lambda__9___closed__6; +LEAN_EXPORT lean_object* l_List_filterTRAux___at_Lean_KeyedDeclsAttribute_getEntries___spec__13___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_KeyedDeclsAttribute_init___rarg___lambda__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_KeyedDeclsAttribute_init___rarg___lambda__11___closed__1; lean_object* l_Lean_getConstInfo___at_Lean_registerInitAttrUnsafe___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); @@ -104,12 +104,10 @@ lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashMapImp_find_x3f___at_Lean_KeyedDeclsAttribute_getEntries___spec__7(lean_object*); LEAN_EXPORT lean_object* l_Lean_KeyedDeclsAttribute_getEntries___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_KeyedDeclsAttribute_init(lean_object*); -LEAN_EXPORT lean_object* l_List_filterAux___at_Lean_KeyedDeclsAttribute_getEntries___spec__12___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_evalConstCheck___at_Lean_KeyedDeclsAttribute_init___spec__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_ofExcept___at_Lean_KeyedDeclsAttribute_init___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_eraseAux___at_Lean_KeyedDeclsAttribute_ExtensionState_insert___spec__2(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_findAtAux___at___private_Lean_KeyedDeclsAttribute_0__Lean_KeyedDeclsAttribute_Table_insert___spec__4(lean_object*); -LEAN_EXPORT lean_object* l_List_filterAux___at_Lean_KeyedDeclsAttribute_getEntries___spec__13(lean_object*); static lean_object* l_Lean_KeyedDeclsAttribute_init___rarg___lambda__10___closed__4; size_t lean_usize_shift_right(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_KeyedDeclsAttribute_instInhabitedOLeanEntry; @@ -121,6 +119,7 @@ LEAN_EXPORT lean_object* l_Lean_KeyedDeclsAttribute_init___rarg___lambda__9(lean LEAN_EXPORT lean_object* l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SMap_fold___at_Lean_KeyedDeclsAttribute_mkStateOfTable___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_findAux___at___private_Lean_KeyedDeclsAttribute_0__Lean_KeyedDeclsAttribute_Table_insert___spec__3___rarg(lean_object*, size_t, lean_object*); +LEAN_EXPORT lean_object* l_List_filterTRAux___at_Lean_KeyedDeclsAttribute_getEntries___spec__12___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_KeyedDeclsAttribute_init___rarg___lambda__6___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_mkHashMap___at_Lean_KeyedDeclsAttribute_ExtensionState_table___default___spec__1___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_KeyedDeclsAttribute_mkStateOfTable___spec__9___rarg(lean_object*, lean_object*, size_t, size_t, lean_object*); @@ -146,7 +145,6 @@ LEAN_EXPORT lean_object* l_Lean_KeyedDeclsAttribute_mkStateOfTable___rarg___lamb LEAN_EXPORT lean_object* l_Lean_KeyedDeclsAttribute_init___rarg___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_KeyedDeclsAttribute_init___rarg___lambda__9___closed__11; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAtCollisionNodeAux___at___private_Lean_KeyedDeclsAttribute_0__Lean_KeyedDeclsAttribute_Table_insert___spec__26___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_filterAux___at_Lean_KeyedDeclsAttribute_getEntries___spec__13___rarg(lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_132_(uint8_t, uint8_t); static lean_object* l_Lean_KeyedDeclsAttribute_instInhabitedDef___lambda__1___closed__1; @@ -288,6 +286,7 @@ lean_object* l_Array_eraseIdx_x27___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_findAtAux___at___private_Lean_KeyedDeclsAttribute_0__Lean_KeyedDeclsAttribute_Table_insert___spec__4___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_mul(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAux___at___private_Lean_KeyedDeclsAttribute_0__Lean_KeyedDeclsAttribute_Table_insert___spec__13(lean_object*); +LEAN_EXPORT lean_object* l_List_filterTRAux___at_Lean_KeyedDeclsAttribute_getEntries___spec__13(lean_object*); LEAN_EXPORT lean_object* l_Lean_KeyedDeclsAttribute_init___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_KeyedDeclsAttribute_instInhabitedDef___closed__2; @@ -296,6 +295,7 @@ LEAN_EXPORT lean_object* l_Std_HashMapImp_find_x3f___at___private_Lean_KeyedDecl LEAN_EXPORT lean_object* l_Lean_KeyedDeclsAttribute_ExtensionState_erase(lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashMapImp_find_x3f___at_Lean_KeyedDeclsAttribute_getEntries___spec__5___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_filterTRAux___at_Lean_KeyedDeclsAttribute_getEntries___spec__12(lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_foldlMAux___at_Lean_KeyedDeclsAttribute_mkStateOfTable___spec__8___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_KeyedDeclsAttribute_mkStateOfTable___spec__13___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_KeyedDeclsAttribute_instInhabitedDef___closed__1; @@ -7248,7 +7248,7 @@ x_4 = lean_nat_dec_eq(x_2, x_3); return x_4; } } -LEAN_EXPORT lean_object* l_List_filterAux___at_Lean_KeyedDeclsAttribute_getEntries___spec__12___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_filterTRAux___at_Lean_KeyedDeclsAttribute_getEntries___spec__12___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_2) == 0) @@ -7328,15 +7328,15 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_List_filterAux___at_Lean_KeyedDeclsAttribute_getEntries___spec__12(lean_object* x_1) { +LEAN_EXPORT lean_object* l_List_filterTRAux___at_Lean_KeyedDeclsAttribute_getEntries___spec__12(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_List_filterAux___at_Lean_KeyedDeclsAttribute_getEntries___spec__12___rarg), 3, 0); +x_2 = lean_alloc_closure((void*)(l_List_filterTRAux___at_Lean_KeyedDeclsAttribute_getEntries___spec__12___rarg), 3, 0); return x_2; } } -LEAN_EXPORT lean_object* l_List_filterAux___at_Lean_KeyedDeclsAttribute_getEntries___spec__13___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_filterTRAux___at_Lean_KeyedDeclsAttribute_getEntries___spec__13___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_2) == 0) @@ -7416,11 +7416,11 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_List_filterAux___at_Lean_KeyedDeclsAttribute_getEntries___spec__13(lean_object* x_1) { +LEAN_EXPORT lean_object* l_List_filterTRAux___at_Lean_KeyedDeclsAttribute_getEntries___spec__13(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_List_filterAux___at_Lean_KeyedDeclsAttribute_getEntries___spec__13___rarg), 3, 0); +x_2 = lean_alloc_closure((void*)(l_List_filterTRAux___at_Lean_KeyedDeclsAttribute_getEntries___spec__13___rarg), 3, 0); return x_2; } } @@ -7448,7 +7448,7 @@ if (lean_obj_tag(x_9) == 0) if (x_11 == 0) { lean_object* x_12; -x_12 = l_List_filterAux___at_Lean_KeyedDeclsAttribute_getEntries___spec__12___rarg(x_10, x_8, x_8); +x_12 = l_List_filterTRAux___at_Lean_KeyedDeclsAttribute_getEntries___spec__12___rarg(x_10, x_8, x_8); return x_12; } else @@ -7465,7 +7465,7 @@ lean_object* x_13; lean_object* x_14; x_13 = lean_ctor_get(x_9, 0); lean_inc(x_13); lean_dec(x_9); -x_14 = l_List_filterAux___at_Lean_KeyedDeclsAttribute_getEntries___spec__13___rarg(x_10, x_13, x_8); +x_14 = l_List_filterTRAux___at_Lean_KeyedDeclsAttribute_getEntries___spec__13___rarg(x_10, x_13, x_8); return x_14; } else diff --git a/stage0/stdlib/Lean/Level.c b/stage0/stdlib/Lean/Level.c index 79295812ac3..3f8dce68419 100644 --- a/stage0/stdlib/Lean/Level.c +++ b/stage0/stdlib/Lean/Level.c @@ -47,6 +47,7 @@ static lean_object* l_Lean_Level_mkData___closed__3; extern lean_object* l_Std_Format_defWidth; lean_object* l_Std_RBMap_instForInRBMapProd___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instReprLevel___closed__1; +LEAN_EXPORT lean_object* l_Lean_Level_instQuoteLevelMkStrAnonymous(lean_object*); uint64_t lean_uint64_add(uint64_t, uint64_t); lean_object* l_Lean_SourceInfo_fromRef(lean_object*); uint64_t lean_bool_to_uint64(uint8_t); @@ -140,7 +141,6 @@ LEAN_EXPORT uint8_t l_Lean_Level_isParam(lean_object*); static lean_object* l_Lean_Level_PP_Result_format___closed__2; LEAN_EXPORT lean_object* l_Lean_Level_updateSucc_x21(lean_object*, lean_object*); static lean_object* l_Lean_Level_updateMax_x21___closed__1; -extern lean_object* l_Lean_numLitKind; lean_object* l_ReaderT_bind___at_Lean_Unhygienic_instMonadQuotationUnhygienic___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Level_isParam___boxed(lean_object*); static lean_object* l_Lean_Level_PP_toResult___closed__2; @@ -249,6 +249,7 @@ static lean_object* l_Lean_Level_PP_Result_format___closed__3; static lean_object* l_Lean_Level_PP_Result_quote___lambda__3___closed__5; LEAN_EXPORT lean_object* l_Lean_Level_PP_Result_imax(lean_object*, lean_object*); static lean_object* l___private_Lean_Level_0__Lean_reprLevel____x40_Lean_Level___hyg_800____closed__20; +lean_object* l_Lean_Syntax_mkNumLit(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instReprData___lambda__3(lean_object*, uint64_t, lean_object*, lean_object*); static lean_object* l_Lean_Level_PP_Result_format___closed__5; static lean_object* l_Lean_Level_PP_Result_quote___lambda__3___closed__6; @@ -271,7 +272,6 @@ LEAN_EXPORT lean_object* l_Lean_mkLevelSucc(lean_object*); static lean_object* l_Lean_Level_PP_Result_quote___lambda__3___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Level_0__Lean_Level_getMaxArgsAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Level_PP_toResult___closed__3; -LEAN_EXPORT lean_object* l_Lean_Level_instQuoteLevel(lean_object*); LEAN_EXPORT lean_object* l_Lean_Level_find_x3f(lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); uint8_t l_Std_RBNode_isRed___rarg(lean_object*); @@ -390,7 +390,6 @@ LEAN_EXPORT lean_object* l_Nat_toLevel___boxed(lean_object*); LEAN_EXPORT uint8_t l_Lean_Level_isMax(lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Level_0__Lean_Level_isExplicitSubsumed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedLevel; -lean_object* l_Lean_Syntax_mkLit(lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Level_0__Lean_Level_mkMaxAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Level_0__Lean_reprMVarId____x40_Lean_Level___hyg_503____closed__8; @@ -5622,15 +5621,14 @@ return x_6; LEAN_EXPORT lean_object* l_Lean_Level_PP_Result_quote___lambda__2(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_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; x_5 = l_Nat_repr(x_1); -x_6 = l_Lean_numLitKind; -x_7 = lean_box(2); -x_8 = l_Lean_Syntax_mkLit(x_6, x_5, x_7); -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_8); -lean_ctor_set(x_9, 1, x_4); -return x_9; +x_6 = lean_box(2); +x_7 = l_Lean_Syntax_mkNumLit(x_5, x_6); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_4); +return x_8; } } static lean_object* _init_l_Lean_Level_PP_Result_quote___lambda__3___closed__1() { @@ -5717,7 +5715,7 @@ return x_2; LEAN_EXPORT lean_object* l_Lean_Level_PP_Result_quote___lambda__3(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; 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; 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_6; lean_object* x_7; 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; 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; x_6 = lean_unsigned_to_nat(65u); x_7 = l_Lean_Level_PP_Result_quote(x_1, x_6); x_8 = l_Lean_Level_PP_Result_format___closed__1; @@ -5727,22 +5725,21 @@ lean_ctor_set(x_9, 1, x_8); x_10 = lean_unsigned_to_nat(1u); x_11 = lean_nat_add(x_2, x_10); x_12 = l_Nat_repr(x_11); -x_13 = l_Lean_numLitKind; -x_14 = lean_box(2); -x_15 = l_Lean_Syntax_mkLit(x_13, x_12, x_14); -x_16 = l_Lean_Level_PP_Result_quote___lambda__3___closed__9; -x_17 = lean_array_push(x_16, x_7); -x_18 = lean_array_push(x_17, x_9); -x_19 = lean_array_push(x_18, x_15); -x_20 = l_Lean_Level_PP_Result_quote___lambda__3___closed__8; -x_21 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_21, 0, x_14); -lean_ctor_set(x_21, 1, x_20); -lean_ctor_set(x_21, 2, x_19); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_5); -return x_22; +x_13 = lean_box(2); +x_14 = l_Lean_Syntax_mkNumLit(x_12, x_13); +x_15 = l_Lean_Level_PP_Result_quote___lambda__3___closed__9; +x_16 = lean_array_push(x_15, x_7); +x_17 = lean_array_push(x_16, x_9); +x_18 = lean_array_push(x_17, x_14); +x_19 = l_Lean_Level_PP_Result_quote___lambda__3___closed__8; +x_20 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_20, 0, x_13); +lean_ctor_set(x_20, 1, x_19); +lean_ctor_set(x_20, 2, x_18); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_5); +return x_21; } } static lean_object* _init_l_Lean_Level_PP_Result_quote___lambda__4___closed__1() { @@ -6196,7 +6193,7 @@ lean_dec(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Level_instQuoteLevel(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Level_instQuoteLevelMkStrAnonymous(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; diff --git a/stage0/stdlib/Lean/Message.c b/stage0/stdlib/Lean/Message.c index 5d988ddab6a..5e073675ebc 100644 --- a/stage0/stdlib/Lean/Message.c +++ b/stage0/stdlib/Lean/Message.c @@ -142,7 +142,6 @@ LEAN_EXPORT lean_object* l_Std_PersistentArray_foldlM___at_Lean_MessageLog_getIn LEAN_EXPORT lean_object* l_Lean_MessageData_toString(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Message_caption___default; static lean_object* l_Lean_MessageData_ofList___closed__6; -lean_object* l_Lean_Syntax_expandInterpolatedStr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); static lean_object* l_Lean_Message_toString___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_MessageData_instAppendMessageData(lean_object*, lean_object*); @@ -154,6 +153,7 @@ LEAN_EXPORT lean_object* l_Lean_stringToMessageData___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_KernelException_toMessageData(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MessageLog_instAppendMessageLog; LEAN_EXPORT lean_object* l_Lean_instToMessageDataArray(lean_object*); +lean_object* l_Lean_TSyntax_expandInterpolatedStr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Lean_MessageSeverity_toCtorIdx(uint8_t); LEAN_EXPORT lean_object* l_Lean_instInhabitedMessageData; @@ -199,6 +199,7 @@ static lean_object* l_Lean_MessageData_instantiateMVars___closed__3; static lean_object* l_Lean_KernelException_toMessageData___closed__14; static lean_object* l_Lean_KernelException_toMessageData___closed__28; LEAN_EXPORT lean_object* l_Lean___aux__Lean__Message______macroRules__Lean__termM_x21____1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_instToMessageDataTSyntax___boxed(lean_object*); static lean_object* l_Lean___aux__Lean__Message______macroRules__Lean__termM_x21____1___closed__15; static lean_object* l_Lean_Message_toString___closed__1; lean_object* l_Function_comp___rarg(lean_object*, lean_object*, lean_object*); @@ -220,6 +221,7 @@ LEAN_EXPORT lean_object* l_Lean_MessageData_instCoeLevelMessageData(lean_object* LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_MessageLog_hasErrors___spec__3___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedMessage___closed__1; static lean_object* l_Lean_termM_x21_____closed__10; +LEAN_EXPORT lean_object* l_Lean_instToMessageDataTSyntax___rarg(lean_object*); static lean_object* l_Lean_addMessageContextPartial___rarg___lambda__1___closed__1; lean_object* l_Lean_ppExpr(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_KernelException_toMessageData___closed__15; @@ -315,6 +317,7 @@ static lean_object* l_Lean_KernelException_toMessageData___closed__35; static lean_object* l_Lean_termM_x21_____closed__5; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); lean_object* l_Lean_ppTerm(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_instToMessageDataTSyntax(lean_object*); LEAN_EXPORT lean_object* l_Lean_MessageData_isNest___boxed(lean_object*); static lean_object* l_Lean_KernelException_toMessageData___closed__23; LEAN_EXPORT lean_object* l_Lean_instToMessageDataList___rarg(lean_object*, lean_object*); @@ -6082,6 +6085,32 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } +LEAN_EXPORT lean_object* l_Lean_instToMessageDataTSyntax___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_instToMessageDataTSyntax(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_instToMessageDataTSyntax___rarg), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_instToMessageDataTSyntax___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_instToMessageDataTSyntax(x_1); +lean_dec(x_1); +return x_2; +} +} LEAN_EXPORT lean_object* l_Lean_instToMessageDataFormat(lean_object* x_1) { _start: { @@ -6716,10 +6745,53 @@ lean_ctor_set(x_27, 0, x_21); lean_ctor_set(x_27, 1, x_25); lean_ctor_set(x_27, 2, x_24); lean_ctor_set(x_27, 3, x_26); -x_28 = l_Lean_Syntax_expandInterpolatedStr(x_9, x_19, x_27, x_2, x_22); +x_28 = l_Lean_TSyntax_expandInterpolatedStr(x_9, x_19, x_27, x_2, x_22); lean_dec(x_9); +if (lean_obj_tag(x_28) == 0) +{ +uint8_t x_29; +x_29 = !lean_is_exclusive(x_28); +if (x_29 == 0) +{ return x_28; } +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_28, 0); +x_31 = lean_ctor_get(x_28, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_28); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +return x_32; +} +} +else +{ +uint8_t x_33; +x_33 = !lean_is_exclusive(x_28); +if (x_33 == 0) +{ +return x_28; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_28, 0); +x_35 = lean_ctor_get(x_28, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_28); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +} +} } } static lean_object* _init_l_Lean_toMessageList___closed__1() { diff --git a/stage0/stdlib/Lean/Meta/Basic.c b/stage0/stdlib/Lean/Meta/Basic.c index ad8c74540c4..392aeb62aa3 100644 --- a/stage0/stdlib/Lean/Meta/Basic.c +++ b/stage0/stdlib/Lean/Meta/Basic.c @@ -251,7 +251,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_forallMetaBoundedTelescope(lean_object*, le LEAN_EXPORT lean_object* l_Lean_Meta_instMonadEnvMetaM___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAuxAux_process___rarg___boxed(lean_object*, lean_object*, 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___private_Lean_Meta_Basic_0__Lean_Meta_isClassExpensive_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_8567_(uint8_t, uint8_t); +uint8_t l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_9542_(uint8_t, uint8_t); static lean_object* l_Lean_Meta_instAlternativeMetaM___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_setInlineAttribute___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withConfig___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -377,7 +377,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_instMonadBacktrackSavedStateMetaM; LEAN_EXPORT lean_object* l_Lean_Meta_getZetaFVarIds___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalInstances(lean_object*); lean_object* l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Meta_instInhabitedSavedState___closed__5; LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_withIncRecDepth___spec__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); @@ -388,6 +387,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_resetZetaFVarIds___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_instMonadEnvMetaM___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentArray_forIn___at___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Meta_Basic_0__Lean_Meta_fvarsSizeLtMaxFVars(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Meta_instMetaEvalMetaM___rarg___closed__9; LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Meta_withLocalInstancesImp___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_swap(lean_object*, lean_object*, lean_object*); @@ -775,7 +775,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_resettingSynth LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_getNumPostponed___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getExprMVarAssignment_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclD___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_assignLevelMVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withReducible___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -972,6 +971,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withoutProofIrrelevance___rarg(lean_object* LEAN_EXPORT lean_object* l_Lean_Meta_orElse(lean_object*); LEAN_EXPORT uint8_t l_Array_qsort_sort___at_Lean_Meta_sortFVarIds___spec__1___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Meta_instMetaEvalMetaM___rarg___closed__11; static lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___closed__5; lean_object* l_Lean_mkConst(lean_object*, lean_object*); @@ -1397,7 +1397,7 @@ x_1 = l_Lean_Meta_instInhabitedInfoCacheKey___closed__3; return x_1; } } -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -1444,7 +1444,7 @@ x_5 = lean_ctor_get(x_1, 1); x_6 = lean_ctor_get_uint8(x_2, sizeof(void*)*2); x_7 = lean_ctor_get(x_2, 0); x_8 = lean_ctor_get(x_2, 1); -x_9 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_8567_(x_3, x_6); +x_9 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_9542_(x_3, x_6); if (x_9 == 0) { uint8_t x_10; @@ -1464,17 +1464,17 @@ return x_12; else { uint8_t x_13; -x_13 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1(x_5, x_8); +x_13 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1(x_5, x_8); return x_13; } } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1(x_1, x_2); +x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -7122,7 +7122,7 @@ x_8 = lean_ctor_get(x_6, 0); x_9 = 0; x_10 = lean_unbox(x_8); lean_dec(x_8); -x_11 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_8567_(x_10, x_9); +x_11 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_9542_(x_10, x_9); x_12 = lean_box(x_11); lean_ctor_set(x_6, 0, x_12); return x_6; @@ -7138,7 +7138,7 @@ lean_dec(x_6); x_15 = 0; x_16 = lean_unbox(x_13); lean_dec(x_13); -x_17 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_8567_(x_16, x_15); +x_17 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_9542_(x_16, x_15); x_18 = lean_box(x_17); x_19 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_19, 0, x_18); @@ -7172,7 +7172,7 @@ x_8 = lean_ctor_get(x_6, 0); x_9 = 2; x_10 = lean_unbox(x_8); lean_dec(x_8); -x_11 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_8567_(x_10, x_9); +x_11 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_9542_(x_10, x_9); x_12 = lean_box(x_11); lean_ctor_set(x_6, 0, x_12); return x_6; @@ -7188,7 +7188,7 @@ lean_dec(x_6); x_15 = 2; x_16 = lean_unbox(x_13); lean_dec(x_13); -x_17 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_8567_(x_16, x_15); +x_17 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_9542_(x_16, x_15); x_18 = lean_box(x_17); x_19 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_19, 0, x_18); diff --git a/stage0/stdlib/Lean/Meta/GetConst.c b/stage0/stdlib/Lean/Meta/GetConst.c index 36e210d25ff..8acbe9a86f9 100644 --- a/stage0/stdlib/Lean/Meta/GetConst.c +++ b/stage0/stdlib/Lean/Meta/GetConst.c @@ -21,7 +21,7 @@ lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getConstNoEx_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isReducible___at___private_Lean_Meta_GetConst_0__Lean_Meta_canUnfoldDefault___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getConst_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_8567_(uint8_t, uint8_t); +uint8_t l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_9542_(uint8_t, uint8_t); lean_object* l_Lean_ConstantInfo_name(lean_object*); LEAN_EXPORT lean_object* l_Lean_getReducibilityStatus___at___private_Lean_Meta_GetConst_0__Lean_Meta_canUnfoldDefault___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_GetConst_0__Lean_Meta_canUnfoldDefault(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -315,7 +315,7 @@ x_39 = lean_ctor_get(x_38, 0); lean_inc(x_39); lean_dec(x_38); x_40 = 3; -x_41 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_8567_(x_6, x_40); +x_41 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_9542_(x_6, x_40); if (x_41 == 0) { uint8_t x_42; lean_object* x_43; @@ -360,7 +360,7 @@ x_51 = lean_ctor_get(x_49, 0); lean_inc(x_51); lean_dec(x_49); x_52 = 3; -x_53 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_8567_(x_6, x_52); +x_53 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_9542_(x_6, x_52); if (x_53 == 0) { uint8_t x_54; lean_object* x_55; lean_object* x_56; diff --git a/stage0/stdlib/Lean/Meta/IndPredBelow.c b/stage0/stdlib/Lean/Meta/IndPredBelow.c index 5bd21337670..f4cb6038bad 100644 --- a/stage0/stdlib/Lean/Meta/IndPredBelow.c +++ b/stage0/stdlib/Lean/Meta/IndPredBelow.c @@ -115,7 +115,6 @@ lean_object* l_List_zipWith___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_IndPredBelow_mkBelowMatcher___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_IndPredBelow_mkBelowMatcher_transformFields(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_IndPredBelow_mkCtorType_rebuild___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_IR_IRType_beq___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Meta_IndPredBelow_mkCtorType_checkCount___spec__9(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_IndPredBelow_instInhabitedVariables; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Meta_IndPredBelow_mkContext_addMotives___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -186,6 +185,7 @@ uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_IndPredBelow_mkCtorType_rebuild___spec__1___closed__1; lean_object* l_Lean_Meta_getMVarType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withNewBinderInfos___at_Lean_Meta_IndPredBelow_mkContext_mkHeader___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_IR_IRType_beq___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_IndPredBelow_mkCtorType_mkBelowBinder___lambda__1(lean_object*, 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_Expr_withAppAux___at_Lean_Meta_IndPredBelow_findBelowIdx___spec__2___lambda__2(lean_object*, 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_Meta_withLocalDecl___at_Lean_Meta_IndPredBelow_mkCtorType_checkCount___spec__8(lean_object*, lean_object*); @@ -13737,7 +13737,7 @@ lean_dec(x_10); lean_inc(x_2); x_25 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_25, 0, x_2); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_IR_IRType_beq___spec__1(x_24, x_25); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_IR_IRType_beq___spec__1(x_24, x_25); lean_dec(x_25); lean_dec(x_24); if (x_26 == 0) diff --git a/stage0/stdlib/Lean/Meta/Match/Basic.c b/stage0/stdlib/Lean/Meta/Match/Basic.c index b18a0400195..352139fc5f0 100644 --- a/stage0/stdlib/Lean/Meta/Match/Basic.c +++ b/stage0/stdlib/Lean/Meta/Match/Basic.c @@ -41,7 +41,6 @@ lean_object* l_Lean_Meta_mkAppM(lean_object*, lean_object*, lean_object*, lean_o lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* l_Lean_throwError___at___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTRAux___at_Lean_Meta_Match_Example_applyFVarSubst___spec__2___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_filterAux___at_Lean_Meta_Match_Alt_replaceFVarId___spec__1(lean_object*, lean_object*, lean_object*); lean_object* l_Array_extract___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Match_Pattern_toMessageData___closed__10; static lean_object* l_List_foldl___at_Lean_Meta_Match_Pattern_toMessageData___spec__1___closed__3; @@ -111,7 +110,6 @@ LEAN_EXPORT lean_object* l_List_mapM___at_Lean_Meta_Match_Pattern_toExpr_visit__ uint8_t l_Lean_Expr_hasExprMVar(lean_object*); LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Match_Pattern_collectFVars___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_replaceRef(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_filterAux___at_Lean_Meta_Match_Alt_replaceFVarId___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Meta_Match_Example_toMessageData___spec__1(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_fvarId_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Match_Pattern_collectFVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -231,6 +229,7 @@ lean_object* l_Lean_Meta_instantiateMVars(lean_object*, lean_object*, lean_objec LEAN_EXPORT lean_object* l_List_mapTRAux___at_Lean_Meta_Match_Alt_toMessageData___spec__1(lean_object*, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withExistingLocalDecls___at_Lean_Meta_Match_Alt_toMessageData___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_filterTRAux___at_Lean_Meta_Match_Alt_replaceFVarId___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Meta_Match_Alt_checkAndReplaceFVarId___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_indentD(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Match_examplesToMessageData(lean_object*); @@ -246,6 +245,7 @@ LEAN_EXPORT lean_object* l_List_mapTRAux___at_Lean_Meta_Match_Alt_applyFVarSubst static lean_object* l_Lean_Meta_Match_Problem_toMessageData___lambda__1___closed__6; static lean_object* l_List_foldl___at_Lean_Meta_Match_Pattern_toMessageData___spec__1___closed__2; lean_object* l_Lean_indentExpr(lean_object*); +LEAN_EXPORT lean_object* l_List_filterTRAux___at_Lean_Meta_Match_Alt_replaceFVarId___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Match_Pattern_toMessageData___closed__16; LEAN_EXPORT lean_object* l_Lean_Meta_Match_instInhabitedProblem; static lean_object* l_List_foldl___at_Lean_Meta_Match_Pattern_toMessageData___spec__1___closed__1; @@ -3790,7 +3790,7 @@ return x_20; } } } -LEAN_EXPORT lean_object* l_List_filterAux___at_Lean_Meta_Match_Alt_replaceFVarId___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_filterTRAux___at_Lean_Meta_Match_Alt_replaceFVarId___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_2) == 0) @@ -3981,7 +3981,7 @@ lean_inc(x_1); x_8 = l_Lean_Expr_replaceFVarId(x_5, x_1, x_2); lean_dec(x_5); x_9 = lean_box(0); -x_10 = l_List_filterAux___at_Lean_Meta_Match_Alt_replaceFVarId___spec__1(x_1, x_6, x_9); +x_10 = l_List_filterTRAux___at_Lean_Meta_Match_Alt_replaceFVarId___spec__1(x_1, x_6, x_9); lean_inc(x_1); x_11 = l_List_mapTRAux___at_Lean_Meta_Match_Alt_replaceFVarId___spec__2(x_1, x_2, x_10, x_9); x_12 = l_List_mapTRAux___at_Lean_Meta_Match_Alt_replaceFVarId___spec__3(x_1, x_2, x_7, x_9); @@ -4008,7 +4008,7 @@ lean_inc(x_1); x_18 = l_Lean_Expr_replaceFVarId(x_15, x_1, x_2); lean_dec(x_15); x_19 = lean_box(0); -x_20 = l_List_filterAux___at_Lean_Meta_Match_Alt_replaceFVarId___spec__1(x_1, x_16, x_19); +x_20 = l_List_filterTRAux___at_Lean_Meta_Match_Alt_replaceFVarId___spec__1(x_1, x_16, x_19); lean_inc(x_1); x_21 = l_List_mapTRAux___at_Lean_Meta_Match_Alt_replaceFVarId___spec__2(x_1, x_2, x_20, x_19); x_22 = l_List_mapTRAux___at_Lean_Meta_Match_Alt_replaceFVarId___spec__3(x_1, x_2, x_17, x_19); @@ -4022,11 +4022,11 @@ return x_23; } } } -LEAN_EXPORT lean_object* l_List_filterAux___at_Lean_Meta_Match_Alt_replaceFVarId___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_filterTRAux___at_Lean_Meta_Match_Alt_replaceFVarId___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_List_filterAux___at_Lean_Meta_Match_Alt_replaceFVarId___spec__1(x_1, x_2, x_3); +x_4 = l_List_filterTRAux___at_Lean_Meta_Match_Alt_replaceFVarId___spec__1(x_1, x_2, x_3); lean_dec(x_1); return x_4; } diff --git a/stage0/stdlib/Lean/Meta/Match/Match.c b/stage0/stdlib/Lean/Meta/Match/Match.c index 95b04df5e1d..e9188c35899 100644 --- a/stage0/stdlib/Lean/Meta/Match/Match.c +++ b/stage0/stdlib/Lean/Meta/Match/Match.c @@ -13,7 +13,6 @@ #ifdef __cplusplus extern "C" { #endif -LEAN_EXPORT lean_object* l_List_filterAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processConstructor___spec__8(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Match_mkMatcherAuxDefinition___lambda__2___closed__10; lean_object* l_List_reverse___rarg(lean_object*); static lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_List_moveToFront_loop___rarg___closed__1; @@ -170,6 +169,7 @@ static lean_object* l_Lean_Meta_Match_mkMatcher___lambda__15___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processConstructor___spec__11(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_checkNumPatterns(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_filterTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processConstructor___spec__8(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_constructorApp_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Match_mkMatcher___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Match_mkMatcher___lambda__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -188,7 +188,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_wi lean_object* lean_expr_instantiate1(lean_object*, lean_object*); lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_collectValues(lean_object*); -LEAN_EXPORT lean_object* l_List_filterAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processArrayLit___spec__6___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_isNextVar(lean_object*); lean_object* lean_array_get_size(lean_object*); @@ -237,7 +236,6 @@ static lean_object* l_Lean_Meta_Match_getMkMatcherInputInContext___closed__1; LEAN_EXPORT lean_object* l_Lean_LocalContext_foldrM___at_Lean_Meta_Match_withCleanLCtxFor___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_foldAux___at_Lean_Meta_Match_mkMatcher___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_checkNextPatternTypes___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_filterAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processValue___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Match_mkMatcher___lambda__14___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_Meta_Match_mkMatcher(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Meta_Match_withMkMatcherInput___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -315,6 +313,7 @@ LEAN_EXPORT uint8_t l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_isArray LEAN_EXPORT uint8_t l_List_foldr___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_hasCtorPattern___spec__1(uint8_t, lean_object*); static lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_throwCasesException___rarg___closed__16; lean_object* l_Lean_ConstantInfo_levelParams(lean_object*); +LEAN_EXPORT lean_object* l_List_filterTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_expandVarIntoCtor_x3f___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_mkHashSet___at_Lean_Meta_Match_State_used___default___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Meta_Match_withCleanLCtxFor___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_mapTRAux___at_Lean_MessageData_instCoeListExprMessageData___spec__1(lean_object*, lean_object*); @@ -327,7 +326,6 @@ static lean_object* l_Lean_Meta_Match_processInaccessibleAsCtor___closed__1; lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_Match_withMkMatcherInput___spec__3(lean_object*); static lean_object* l_List_mapTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_expandNatValuePattern___spec__1___closed__2; -LEAN_EXPORT lean_object* l_List_filterAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processConstructor___spec__8___boxed(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Match_MkMatcherInput_collectFVars___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Meta_Match_withCleanLCtxFor___spec__6(lean_object*, lean_object*, size_t, size_t, lean_object*); @@ -390,7 +388,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_ex static lean_object* l_Lean_Meta_Match_processInaccessibleAsCtor___closed__4; lean_object* l_Lean_Meta_FVarSubst_apply(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_replace___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processLeaf___spec__8(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_filterAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processValue___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldr___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_isNatValueTransition___spec__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_throwCasesException___rarg___closed__5; lean_object* l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_54____spec__1(lean_object*, lean_object*, lean_object*); @@ -472,6 +469,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Match_mkMatcher___lambda__17___boxed(lean_o lean_object* l_Lean_LocalContext_getFVar_x21(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_withEqs_go___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_process_tryToProcess___closed__4; +LEAN_EXPORT lean_object* l_List_filterTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_expandVarIntoCtor_x3f___spec__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldr___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_isConstructorTransition___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Match_mkMatcherAuxDefinition___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_left(size_t, size_t); @@ -621,6 +619,7 @@ static lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_throwNo extern lean_object* l_Lean_Meta_Match_instInhabitedPattern; extern uint8_t l_instInhabitedBool; LEAN_EXPORT lean_object* l_Lean_Meta_Match_MkMatcherInput_collectFVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_filterTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processValue___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentArray_foldrM___at_Lean_Meta_Match_withCleanLCtxFor___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_updateAlts(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_withAlts_mkMinorType___boxed__const__1; @@ -691,7 +690,6 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Match_mkMatcher__ LEAN_EXPORT lean_object* l_List_foldr___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_checkNumPatterns___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processLeaf___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_updateAlts___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_filterAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_expandVarIntoCtor_x3f___spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_throwCasesException___rarg___closed__10; LEAN_EXPORT lean_object* l_Array_indexOfAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_getUElimPos_x3f___spec__1(lean_object*, lean_object*, lean_object*); lean_object* l_Array_ofSubarray___rarg(lean_object*); @@ -699,6 +697,7 @@ LEAN_EXPORT lean_object* l_List_mapM___at___private_Lean_Meta_Match_Match_0__Lea lean_object* l_Lean_Meta_getLevel(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_MatcherApp_addArg___lambda__4___closed__2; LEAN_EXPORT lean_object* l_Nat_foldAux___at_Lean_Meta_Match_mkMatcher___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_filterTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processConstructor___spec__8___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_nat_mul(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_traceState(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processArrayLit___spec__5(lean_object*, lean_object*, lean_object*); @@ -719,7 +718,6 @@ uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Match_mkMatcher___lambda__5___boxed(lean_object**); static lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_throwCasesException___rarg___closed__9; static lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_throwNonSupported___lambda__1___closed__1; -LEAN_EXPORT lean_object* l_List_filterAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_expandVarIntoCtor_x3f___spec__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Match_mkMatcher___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_List_foldr___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_isNatValueTransition___spec__1(uint8_t, lean_object*); static lean_object* l_Lean_Meta_Match_Unify_assign___closed__5; @@ -731,7 +729,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_pr LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Match_mkMatcher___spec__12___boxed(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_List_foldr___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_hasArrayLitPattern___spec__1___boxed(lean_object*, lean_object*); lean_object* l_Lean_Meta_mkArrow(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_filterAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processValue___spec__4___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Match_getMkMatcherInputInContext___lambda__2(lean_object*, 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_Array_foldlMUnsafe_fold___at_Lean_Meta_Match_getMkMatcherInputInContext___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_List_foldr___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_hasArrayLitPattern___spec__1(uint8_t, lean_object*); @@ -741,6 +738,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Match_mkMatcher___lambda__18(lean_object*, lean_object* l_instBEq___rarg(lean_object*, lean_object*, lean_object*); lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTRAux___at_Lean_Meta_Match_mkMatcher___spec__8___at_Lean_Meta_Match_mkMatcher___spec__9___at_Lean_Meta_Match_mkMatcher___spec__10(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_filterTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processValue___spec__4___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_List_mapTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processValue___spec__6___closed__1; lean_object* l_Lean_Meta_isExprDefEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarImpl(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -748,6 +746,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_is static lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_updateAlts___lambda__3___closed__1; lean_object* l_Lean_Meta_Match_withGoalOf___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Meta_Match_withCleanLCtxFor___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_filterTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processArrayLit___spec__6___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTRAux___at_Lean_Meta_Match_mkMatcher___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_reduceMatcher_x3f___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -771,7 +770,6 @@ LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Meta_Match_withClea lean_object* lean_mk_array(lean_object*, lean_object*); lean_object* l_Lean_Meta_withExistingLocalDecls___at_Lean_Meta_Match_Alt_toMessageData___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkEqHEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_filterAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processArrayLit___spec__6(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_traceStep___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_unify_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_panic___at_Lean_Meta_ACLt_lt_lexSameCtor___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -807,6 +805,7 @@ lean_object* l_Lean_indentD(lean_object*); static lean_object* l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_Match___hyg_9662____closed__5; LEAN_EXPORT lean_object* l_List_filterMapM_loop___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processConstructor___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_withEqs___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_filterTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processArrayLit___spec__6(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Match_mkMatcherAuxDefinition___lambda__3(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Match_mkMatcher___lambda__6___closed__3; lean_object* l_Lean_Meta_Match_examplesToMessageData(lean_object*); @@ -958,6 +957,7 @@ uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_updateAlts___lambda__2___closed__1; static lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_updateAlts___closed__1; static lean_object* l_Lean_Meta_Match_withMkMatcherInput___rarg___lambda__1___closed__2; +LEAN_EXPORT lean_object* l_List_filterTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processValue___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_List_foldr___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_checkNumPatterns___spec__1(lean_object* x_1, uint8_t x_2, lean_object* x_3) { _start: { @@ -8287,7 +8287,7 @@ return x_25; } } } -LEAN_EXPORT lean_object* l_List_filterAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_expandVarIntoCtor_x3f___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_filterTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_expandVarIntoCtor_x3f___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_2) == 0) @@ -8675,7 +8675,7 @@ if (x_33 == 0) 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; x_34 = lean_ctor_get(x_24, 0); x_35 = lean_box(0); -x_36 = l_List_filterAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_expandVarIntoCtor_x3f___spec__2(x_34, x_22, x_35); +x_36 = l_List_filterTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_expandVarIntoCtor_x3f___spec__2(x_34, x_22, x_35); lean_inc(x_34); x_37 = l_List_mapTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_expandVarIntoCtor_x3f___spec__3(x_34, x_36, x_35); x_38 = lean_ctor_get(x_13, 4); @@ -8712,7 +8712,7 @@ x_48 = lean_ctor_get(x_24, 0); lean_inc(x_48); lean_dec(x_24); x_49 = lean_box(0); -x_50 = l_List_filterAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_expandVarIntoCtor_x3f___spec__2(x_48, x_22, x_49); +x_50 = l_List_filterTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_expandVarIntoCtor_x3f___spec__2(x_48, x_22, x_49); lean_inc(x_48); x_51 = l_List_mapTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_expandVarIntoCtor_x3f___spec__3(x_48, x_50, x_49); x_52 = lean_ctor_get(x_13, 4); @@ -8761,7 +8761,7 @@ if (lean_is_exclusive(x_24)) { x_65 = lean_box(0); } x_66 = lean_box(0); -x_67 = l_List_filterAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_expandVarIntoCtor_x3f___spec__2(x_64, x_22, x_66); +x_67 = l_List_filterTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_expandVarIntoCtor_x3f___spec__2(x_64, x_22, x_66); lean_inc(x_64); x_68 = l_List_mapTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_expandVarIntoCtor_x3f___spec__3(x_64, x_67, x_66); x_69 = lean_ctor_get(x_13, 4); @@ -9089,11 +9089,11 @@ lean_dec(x_5); return x_11; } } -LEAN_EXPORT lean_object* l_List_filterAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_expandVarIntoCtor_x3f___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_filterTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_expandVarIntoCtor_x3f___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_List_filterAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_expandVarIntoCtor_x3f___spec__2(x_1, x_2, x_3); +x_4 = l_List_filterTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_expandVarIntoCtor_x3f___spec__2(x_1, x_2, x_3); lean_dec(x_1); return x_4; } @@ -11347,7 +11347,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_List_filterAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processConstructor___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_filterTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processConstructor___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_2) == 0) @@ -12156,7 +12156,7 @@ lean_dec(x_27); x_30 = l_List_mapTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processConstructor___spec__7(x_20, x_29, x_23); x_31 = lean_ctor_get(x_1, 2); lean_inc(x_31); -x_32 = l_List_filterAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processConstructor___spec__8(x_25, x_31, x_23); +x_32 = l_List_filterTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processConstructor___spec__8(x_25, x_31, x_23); x_33 = l_List_mapTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processConstructor___spec__9(x_20, x_32, x_23); x_34 = l_List_reverse___rarg(x_33); lean_inc(x_18); @@ -12677,11 +12677,11 @@ lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_List_filterAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processConstructor___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_filterTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processConstructor___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_List_filterAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processConstructor___spec__8(x_1, x_2, x_3); +x_4 = l_List_filterTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processConstructor___spec__8(x_1, x_2, x_3); lean_dec(x_1); return x_4; } @@ -13907,7 +13907,7 @@ x_3 = lean_box(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_List_filterAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processValue___spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_List_filterTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processValue___spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -14070,7 +14070,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_List_filterAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processValue___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_filterTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processValue___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_2) == 0) @@ -14771,7 +14771,7 @@ x_22 = lean_box(0); x_23 = l_List_mapTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processValue___spec__2(x_3, x_19, x_21, x_22); lean_dec(x_19); x_24 = l_List_mapTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processValue___spec__3(x_4, x_23, x_22); -x_25 = l_List_filterAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processValue___spec__4(x_1, x_20, x_22); +x_25 = l_List_filterTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processValue___spec__4(x_1, x_20, x_22); lean_inc(x_4); x_26 = l_List_mapTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processValue___spec__5(x_2, x_5, x_6, x_7, x_8, x_9, lean_box(0), x_4, x_10, x_13, x_25, x_22); x_27 = l_List_mapTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processValue___spec__6(x_2, x_11, x_5, x_6, x_7, x_8, x_9, lean_box(0), x_1, x_10, x_13, x_26, x_22); @@ -14882,7 +14882,7 @@ lean_dec(x_21); x_22 = lean_ctor_get(x_3, 0); lean_dec(x_22); x_23 = lean_box(0); -x_24 = l_List_filterAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processValue___spec__1(x_20, x_23); +x_24 = l_List_filterTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processValue___spec__1(x_20, x_23); x_25 = lean_ctor_get(x_4, 0); lean_inc(x_25); lean_dec(x_4); @@ -14906,7 +14906,7 @@ lean_inc(x_29); lean_inc(x_28); lean_dec(x_3); x_30 = lean_box(0); -x_31 = l_List_filterAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processValue___spec__1(x_28, x_30); +x_31 = l_List_filterTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processValue___spec__1(x_28, x_30); x_32 = lean_ctor_get(x_4, 0); lean_inc(x_32); lean_dec(x_4); @@ -15425,11 +15425,11 @@ lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_List_filterAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processValue___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_filterTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processValue___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_List_filterAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processValue___spec__4(x_1, x_2, x_3); +x_4 = l_List_filterTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processValue___spec__4(x_1, x_2, x_3); lean_dec(x_1); return x_4; } @@ -16463,7 +16463,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_List_filterAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processArrayLit___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_filterTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processArrayLit___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_2) == 0) @@ -17256,7 +17256,7 @@ lean_inc(x_24); x_25 = lean_ctor_get(x_1, 3); lean_inc(x_25); x_26 = lean_box(0); -x_27 = l_List_filterAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processValue___spec__1(x_24, x_26); +x_27 = l_List_filterTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processValue___spec__1(x_24, x_26); x_28 = lean_ctor_get(x_21, 0); lean_inc(x_28); lean_dec(x_21); @@ -17309,7 +17309,7 @@ lean_inc(x_47); x_48 = l_List_mapTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processArrayLit___spec__4(x_5, x_45, x_47, x_40); lean_dec(x_45); x_49 = l_List_mapTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processArrayLit___spec__5(x_38, x_48, x_40); -x_50 = l_List_filterAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processArrayLit___spec__6(x_35, x_46, x_40); +x_50 = l_List_filterTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processArrayLit___spec__6(x_35, x_46, x_40); lean_inc(x_38); x_51 = l_List_mapTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processArrayLit___spec__7(x_38, x_50, x_40); lean_inc(x_15); @@ -17620,11 +17620,11 @@ lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_List_filterAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processArrayLit___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_filterTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processArrayLit___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_List_filterAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processArrayLit___spec__6(x_1, x_2, x_3); +x_4 = l_List_filterTRAux___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processArrayLit___spec__6(x_1, x_2, x_3); lean_dec(x_1); return x_4; } diff --git a/stage0/stdlib/Lean/Meta/PPGoal.c b/stage0/stdlib/Lean/Meta/PPGoal.c index 0b3e69e0e63..4c3bb0222ab 100644 --- a/stage0/stdlib/Lean/Meta/PPGoal.c +++ b/stage0/stdlib/Lean/Meta/PPGoal.c @@ -62,7 +62,7 @@ LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_ToHide_hasInaccessible LEAN_EXPORT lean_object* l_Std_HashMapImp_find_x3f___at_Lean_Meta_ToHide_visitVisibleExpr_visit___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_pp_inaccessibleNames; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Meta_ppGoal_ppVars___spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Meta_ppGoal_ppVars___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_ToHide_isMarked___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_pp_auxDecls; LEAN_EXPORT lean_object* l_Std_RBNode_findCore___at_Lean_Meta_ToHide_isMarked___spec__1(lean_object*, lean_object*); @@ -116,6 +116,7 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_PPGo LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_MetavarContext_DependsOn_dep_visitMain___at_Lean_Meta_ToHide_hasVisibleDep___spec__34(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_PPGoal_0__Lean_Meta_ToHide_getInitialHiddenInaccessible___spec__12___boxed(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___private_Lean_MetavarContext_0__Lean_MetavarContext_DependsOn_dep_visitMain___at_Lean_Meta_ToHide_hasVisibleDep___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Meta_ppGoal_ppVars___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_ToHide_hasVisibleDep___spec__22___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_PPGoal_0__Lean_Meta_ToHide_getInitialHiddenInaccessible___spec__13___boxed(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___private_Lean_MetavarContext_0__Lean_MetavarContext_DependsOn_dep_visitApp___at_Lean_Meta_ToHide_hasInaccessibleNameDep___spec__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -308,7 +309,6 @@ uint8_t l_Lean_Expr_hasMVar(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_PPGoal_0__Lean_Meta_ToHide_getInitialHiddenInaccessible___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Std_RBNode_isRed___rarg(lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_ToHide_hasVisibleDep___spec__31(lean_object*, lean_object*, size_t, size_t); -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Meta_ppGoal_ppVars___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withPPInaccessibleNamesImp(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_ppGoal___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_ToHide_hasVisibleDep___spec__8(lean_object*, lean_object*, size_t, size_t); @@ -19736,7 +19736,7 @@ lean_dec(x_5); return x_10; } } -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Meta_ppGoal_ppVars___spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Meta_ppGoal_ppVars___spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -19854,13 +19854,13 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; x_19 = lean_ctor_get(x_17, 0); x_20 = lean_ctor_get(x_17, 1); x_21 = lean_box(0); -x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Meta_ppGoal_ppVars___spec__1(x_4, x_21); +x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Meta_ppGoal_ppVars___spec__1(x_4, x_21); if (x_22 == 0) { lean_object* x_23; uint8_t x_24; x_23 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_23, 0, x_19); -x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Meta_ppGoal_ppVars___spec__1(x_4, x_23); +x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Meta_ppGoal_ppVars___spec__1(x_4, x_23); if (x_24 == 0) { lean_object* x_25; @@ -19986,13 +19986,13 @@ lean_inc(x_51); lean_inc(x_50); lean_dec(x_17); x_52 = lean_box(0); -x_53 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Meta_ppGoal_ppVars___spec__1(x_4, x_52); +x_53 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Meta_ppGoal_ppVars___spec__1(x_4, x_52); if (x_53 == 0) { lean_object* x_54; uint8_t x_55; x_54 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_54, 0, x_50); -x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Meta_ppGoal_ppVars___spec__1(x_4, x_54); +x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Meta_ppGoal_ppVars___spec__1(x_4, x_54); if (x_55 == 0) { lean_object* x_56; @@ -20471,11 +20471,11 @@ return x_186; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Meta_ppGoal_ppVars___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Meta_ppGoal_ppVars___spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Meta_ppGoal_ppVars___spec__1(x_1, x_2); +x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Meta_ppGoal_ppVars___spec__1(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); diff --git a/stage0/stdlib/Lean/Meta/RecursorInfo.c b/stage0/stdlib/Lean/Meta/RecursorInfo.c index 3a6787dcaee..ccd9d0e4ad1 100644 --- a/stage0/stdlib/Lean/Meta/RecursorInfo.c +++ b/stage0/stdlib/Lean/Meta/RecursorInfo.c @@ -61,7 +61,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Attribute_Recursor_getMajorPos___lambda__2_ static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_RecursorInfo___hyg_2711____lambda__2___closed__14; lean_object* l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_RecursorInfo___hyg_2711____lambda__3___boxed(lean_object*, lean_object*, lean_object*); -lean_object* l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_RecursorInfo_0__Lean_Meta_getProduceMotiveAndRecursive___spec__3___lambda__3___boxed(lean_object*, lean_object*, lean_object*, 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___private_Lean_Meta_RecursorInfo_0__Lean_Meta_checkMotive___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); @@ -130,6 +129,7 @@ lean_object* l_Lean_RecursorVal_getMajorIdx(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_RecursorInfo_0__Lean_Meta_getMajorPosIfAuxRecursor_x3f___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_RecursorInfo_numMinors(lean_object*); +lean_object* l_Lean_Syntax_isNatLit_x3f(lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_RecursorInfo___hyg_2711____lambda__2___closed__2; static lean_object* l___private_Lean_Meta_RecursorInfo_0__Lean_Meta_getMajorPosDepElim___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_getMajorPos_x3f(lean_object*, lean_object*); @@ -142,7 +142,6 @@ static lean_object* l_Lean_Meta_instToStringRecursorUnivLevelPos___closed__1; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_RecursorInfo___hyg_2711____closed__5; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_RecursorInfo_0__Lean_Meta_mkRecursorInfoAux___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_toStringAux___at_Lean_Meta_RecursorInfo_instToStringRecursorInfo___spec__4___closed__2; -extern lean_object* l_Lean_numLitKind; lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_RecursorInfo_0__Lean_Meta_checkMotive(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_RecursorInfo___hyg_2711____lambda__2___closed__1; @@ -9157,82 +9156,81 @@ return x_9; } else { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +lean_object* x_10; lean_object* x_11; lean_object* x_12; x_10 = lean_unsigned_to_nat(1u); x_11 = l_Lean_Syntax_getArg(x_1, x_10); -x_12 = l_Lean_numLitKind; -x_13 = l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(x_12, x_11); +x_12 = l_Lean_Syntax_isNatLit_x3f(x_11); lean_dec(x_11); -if (lean_obj_tag(x_13) == 0) +if (lean_obj_tag(x_12) == 0) { -lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_14 = l_Lean_Meta_Attribute_Recursor_getMajorPos___closed__13; -x_15 = l_Lean_throwErrorAt___at_Lean_Meta_Attribute_Recursor_getMajorPos___spec__1(x_1, x_14, x_2, x_3, x_4); +lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_13 = l_Lean_Meta_Attribute_Recursor_getMajorPos___closed__13; +x_14 = l_Lean_throwErrorAt___at_Lean_Meta_Attribute_Recursor_getMajorPos___spec__1(x_1, x_13, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_1); -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) { -return x_15; +return x_14; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_15, 0); -x_18 = lean_ctor_get(x_15, 1); -lean_inc(x_18); +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_14, 0); +x_17 = lean_ctor_get(x_14, 1); lean_inc(x_17); -lean_dec(x_15); -x_19 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set(x_19, 1, x_18); -return x_19; +lean_inc(x_16); +lean_dec(x_14); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +return x_18; } } else { -lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_20 = lean_ctor_get(x_13, 0); -lean_inc(x_20); -lean_dec(x_13); -x_21 = lean_unsigned_to_nat(0u); -x_22 = lean_nat_dec_eq(x_20, x_21); -if (x_22 == 0) +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = lean_ctor_get(x_12, 0); +lean_inc(x_19); +lean_dec(x_12); +x_20 = lean_unsigned_to_nat(0u); +x_21 = lean_nat_dec_eq(x_19, x_20); +if (x_21 == 0) { -lean_object* x_23; lean_object* x_24; +lean_object* x_22; lean_object* x_23; lean_dec(x_1); -x_23 = lean_box(0); -x_24 = l_Lean_Meta_Attribute_Recursor_getMajorPos___lambda__2(x_20, x_23, x_2, x_3, x_4); +x_22 = lean_box(0); +x_23 = l_Lean_Meta_Attribute_Recursor_getMajorPos___lambda__2(x_19, x_22, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_20); -return x_24; +lean_dec(x_19); +return x_23; } else { -lean_object* x_25; lean_object* x_26; uint8_t x_27; -lean_dec(x_20); -x_25 = l_Lean_Meta_Attribute_Recursor_getMajorPos___closed__13; -x_26 = l_Lean_throwErrorAt___at_Lean_Meta_Attribute_Recursor_getMajorPos___spec__1(x_1, x_25, x_2, x_3, x_4); +lean_object* x_24; lean_object* x_25; uint8_t x_26; +lean_dec(x_19); +x_24 = l_Lean_Meta_Attribute_Recursor_getMajorPos___closed__13; +x_25 = l_Lean_throwErrorAt___at_Lean_Meta_Attribute_Recursor_getMajorPos___spec__1(x_1, x_24, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_1); -x_27 = !lean_is_exclusive(x_26); -if (x_27 == 0) +x_26 = !lean_is_exclusive(x_25); +if (x_26 == 0) { -return x_26; +return x_25; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_26, 0); -x_29 = lean_ctor_get(x_26, 1); -lean_inc(x_29); +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_25, 0); +x_28 = lean_ctor_get(x_25, 1); lean_inc(x_28); -lean_dec(x_26); -x_30 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -return x_30; +lean_inc(x_27); +lean_dec(x_25); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; } } } diff --git a/stage0/stdlib/Lean/Meta/Reduce.c b/stage0/stdlib/Lean/Meta/Reduce.c index 0f978b227e6..6483d2957c6 100644 --- a/stage0/stdlib/Lean/Meta/Reduce.c +++ b/stage0/stdlib/Lean/Meta/Reduce.c @@ -64,7 +64,6 @@ static lean_object* l_Lean_Meta_reduce_visit___lambda__4___closed__4; lean_object* lean_st_mk_ref(lean_object*, lean_object*); uint64_t l_Lean_Expr_hash(lean_object*); static lean_object* l_Lean_Meta_reduce_visit___lambda__4___closed__9; -lean_object* l_panic___at_String_toNat_x21___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_reduce_visit___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_natLit_x3f(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at_Lean_Meta_reduce_visit___spec__6___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -99,6 +98,7 @@ lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_reduce_visit___lambda__2(uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_reduce_visit___lambda__3(uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_panic___at_Lean_TSyntax_getNat___spec__1(lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Meta_reduce_visit___spec__4___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_reduce_visit___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1663,7 +1663,7 @@ if (lean_obj_tag(x_45) == 0) { lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; x_46 = l_Lean_Meta_reduce_visit___lambda__4___closed__9; -x_47 = l_panic___at_String_toNat_x21___spec__1(x_46); +x_47 = l_panic___at_Lean_TSyntax_getNat___spec__1(x_46); x_48 = lean_nat_add(x_47, x_28); lean_dec(x_47); x_49 = l_Lean_mkRawNatLit(x_48); @@ -1747,7 +1747,7 @@ if (lean_obj_tag(x_68) == 0) { lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; x_69 = l_Lean_Meta_reduce_visit___lambda__4___closed__9; -x_70 = l_panic___at_String_toNat_x21___spec__1(x_69); +x_70 = l_panic___at_Lean_TSyntax_getNat___spec__1(x_69); x_71 = lean_nat_add(x_70, x_28); lean_dec(x_70); x_72 = l_Lean_mkRawNatLit(x_71); diff --git a/stage0/stdlib/Lean/Meta/Tactic/LinearArith/Solver.c b/stage0/stdlib/Lean/Meta/Tactic/LinearArith/Solver.c index a4c8da26e4f..97c6eafd81c 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/LinearArith/Solver.c +++ b/stage0/stdlib/Lean/Meta/Tactic/LinearArith/Solver.c @@ -21,14 +21,14 @@ LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Lean_Meta_Tactic_Li static lean_object* l_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_341____spec__1___closed__3; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_341____closed__10; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_148____closed__6; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__2; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__2; LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqAssumptionId____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_2773_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_get___boxed(lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__19; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__11; uint8_t lean_usize_dec_eq(size_t, size_t); lean_object* lean_array_uget(lean_object*, size_t); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__11; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__19; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_AssumptionId_id___default; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_CnstrKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprAssumptionId____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_2849____boxed(lean_object*, lean_object*); @@ -45,13 +45,11 @@ static lean_object* l_Lean_Meta_Linear_CnstrKind_noConfusion___rarg___closed__1; lean_object* l_Lean_Rat_floor(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_CnstrKind_ofNat___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_148_(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__8; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__9; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__6; static lean_object* l_Lean_Meta_Linear_instReprCnstr___closed__1; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__6; -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4110_(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4109_(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instOrdVar; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__17; static lean_object* l_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_341____spec__1___closed__5; lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); @@ -61,6 +59,7 @@ LEAN_EXPORT uint8_t l_Lean_Meta_Linear_instDecidableEqVar(lean_object*, lean_obj static lean_object* l_Lean_Meta_Linear_instReprJustification___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_72____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_CnstrKind_noConfusion___rarg(uint8_t, uint8_t, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__17; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_eval_x3f___lambda__1(lean_object*, lean_object*); lean_object* l_Std_Format_joinSep___at_instReprProd___spec__1(lean_object*, lean_object*); lean_object* l_Array_shrink___rarg(lean_object*, lean_object*); @@ -69,14 +68,15 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instReprCnstr; static lean_object* l_Lean_Meta_Linear_Poly_eval_x3f___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_341_(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_148____closed__11; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__8; LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at_Lean_Meta_Linear_Cnstr_getBound___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_getBestBound_x3f(lean_object*, lean_object*, uint8_t, uint8_t); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__9; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Assignment_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Assignment_push(lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l_Lean_Meta_Linear_instReprCnstrKind___closed__1; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_341____closed__12; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__4; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprAssumptionId____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_2849_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instDecidableEqVar___boxed(lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); @@ -109,10 +109,12 @@ lean_object* lean_int_mul(lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_341____closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instInhabitedVar; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__12; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_341____closed__15; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_341____closed__14; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Cnstr_isStrict___boxed(lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Linear_Cnstr_isUnsat___spec__1(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__7; static lean_object* l_Lean_Meta_Linear_Cnstr_isUnsat___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_eval_x3f___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instDecidableEqCnstrKind___boxed(lean_object*, lean_object*); @@ -122,31 +124,29 @@ static lean_object* l_Lean_Meta_Linear_pickAssignment_x3f___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instReprJustification; LEAN_EXPORT uint8_t l_Lean_Meta_Linear_instDecidableEqCnstr(lean_object*, lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__7; static lean_object* l_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_341____spec__1___closed__4; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_341____closed__7; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__10; lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3308_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_add_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__20; static lean_object* l_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_341____spec__1___closed__1; uint8_t l_Lean_Rat_instDecidableLeRatInstLERat(lean_object*, lean_object*); lean_object* l_Nat_repr(lean_object*); static lean_object* l_Lean_Meta_Linear_instBEqCnstr___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3675____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__10; LEAN_EXPORT uint8_t l_Lean_Meta_Linear_instDecidableEqAssumptionId(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instLTVar; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__10; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3674____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__10; uint8_t l___private_Lean_Data_Rat_0__Lean_beqRat____x40_Lean_Data_Rat___hyg_34_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Cnstr_getBound___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_eval_x3f(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__20; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3436____closed__2; LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_ordVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_23_(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_341____closed__16; lean_object* l_Lean_Rat_add(lean_object*, lean_object*); lean_object* l_Array_back___rarg(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__12; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_combine_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Linear_instReprPoly___closed__1; lean_object* lean_array_to_list(lean_object*, lean_object*); @@ -158,22 +158,23 @@ lean_object* l_Lean_Rat_sub(lean_object*, lean_object*); extern lean_object* l_Int_instInhabitedInt; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3436____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_eval_x3f___lambda__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__3; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3436____closed__1; -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3875_(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3874_(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_370____lambda__1___closed__1; static lean_object* l_Lean_Meta_Linear_Assignment_val___default___closed__1; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__3; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_370____lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_solve(lean_object*, lean_object*); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__15; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instInhabitedCnstr; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_148____closed__12; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Cnstr_isUnsat___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_CnstrKind_noConfusion___rarg___lambda__1___boxed(lean_object*); lean_object* l_instDecidableEqProd___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Linear_instDecidableEqCnstrKind(uint8_t, uint8_t); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__5; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__18; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__14; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__22; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_341____closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_ordVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_23____boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3436____closed__5; @@ -187,65 +188,68 @@ size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Cnstr_getBound(lean_object*, lean_object*); static lean_object* l_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_341____spec__1___closed__8; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3436____closed__4; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__14; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_State_getBestUpperBound_x3f(lean_object*); lean_object* l_Lean_Rat_inv(lean_object*); static lean_object* l_Lean_Meta_Linear_pickAssignment_x3f___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instReprPoly; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__22; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__18; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instBEqCnstr; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_pickAssignment_x3f(lean_object*, uint8_t, lean_object*, uint8_t); -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3675_(uint8_t, uint8_t); +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3674_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Assignment_get_x3f___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3875____boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__15; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instReprVar; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instInhabitedAssignment; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instReprAssumptionId; lean_object* lean_int_neg(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3874____boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_341____closed__13; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Linear_Poly_eval_x3f___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Meta_Linear_instInhabitedJustification___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instReprCnstrKind; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213_(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__16; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212_(lean_object*, lean_object*); extern uint8_t l_instInhabitedBool; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4110____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__24; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3436____closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4109____boxed(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Linear_Cnstr_isUnsat___closed__4; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_341____closed__6; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_148____closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_combine___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instBEqJustification; uint8_t l_Lean_Rat_lt(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__13; static lean_object* l_Lean_Meta_Linear_resolve___closed__1; static lean_object* l_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_341____spec__1___closed__7; LEAN_EXPORT uint8_t l_Lean_Meta_Linear_instDecidableEqJustification(lean_object*, lean_object*); uint8_t l_Lean_Rat_isInt(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__13; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_370____lambda__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__7; uint8_t lean_int_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_add_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_370_(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__16; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instDecidableEqCnstr___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Linear_Poly_add_go___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_get(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Linear_instInhabitedCnstrKind; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__24; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_148____closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_resolve___boxed(lean_object*, lean_object*, lean_object*); uint8_t l_Array_instDecidableEqArray___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__3; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__23; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_148____boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__5; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__5; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__3; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__2; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_72_(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_148____closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_CnstrKind_noConfusion___rarg___lambda__1(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instDecidableEqPoly___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_add(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_148____closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_State_getNumVars(lean_object*); @@ -261,24 +265,19 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instInhabitedAssumptionId; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Assignment_size___boxed(lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Linear_Cnstr_isStrict(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3308____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__21; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Assignment_shrink(lean_object*, lean_object*); lean_object* lean_string_length(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__23; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__21; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3436____closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instBEqCnstrKind; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Linear_getBestBound_x3f___spec__1(lean_object*, uint8_t, uint8_t, lean_object*, size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____boxed(lean_object*, lean_object*); lean_object* lean_int_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_scale___boxed(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Linear_Poly_eval_x3f___spec__1___closed__1; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Linear_Poly_eval_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Linear_instReprAssumptionId___closed__1; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__4; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__9; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_341____closed__11; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__9; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_State_getNumVars___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_size___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_370____boxed(lean_object*, lean_object*); @@ -288,20 +287,20 @@ static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta LEAN_EXPORT lean_object* l_Lean_Meta_Linear_CnstrKind_noConfusion(lean_object*); uint8_t lean_int_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instInhabitedState; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_combine(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_341____spec__1___closed__6; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691_(uint8_t, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690_(uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_getMaxVarCoeff___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Linear_Poly_scale___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____boxed(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Linear_Poly_getMaxVarCoeff___closed__1; lean_object* lean_nat_to_int(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__1; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__1; static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_341____closed__8; static lean_object* l_Lean_Meta_Linear_pickAssignment_x3f___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instDecidableEqAssumptionId___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Assignment_shrink___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_State_assignment___default; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_pickAssignment_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_getMaxVar___boxed(lean_object*); @@ -309,9 +308,10 @@ static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta LEAN_EXPORT lean_object* l_Lean_Meta_Linear_State_getBestLowerBound_x3f___boxed(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_341____closed__17; LEAN_EXPORT uint8_t l_Lean_Meta_Linear_CnstrKind_ofNat(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__4; lean_object* l_Nat_decEq___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__8; lean_object* l_Lean_Rat_mul(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__8; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_341____closed__4; lean_object* l_Repr_addAppParen(lean_object*, lean_object*); @@ -3106,7 +3106,7 @@ x_6 = lean_box(x_5); return x_6; } } -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3675_(uint8_t x_1, uint8_t x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3674_(uint8_t x_1, uint8_t x_2) { _start: { lean_object* x_3; lean_object* x_4; uint8_t x_5; @@ -3118,7 +3118,7 @@ lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3675____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3674____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; @@ -3126,7 +3126,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_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3675_(x_3, x_4); +x_5 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3674_(x_3, x_4); x_6 = lean_box(x_5); return x_6; } @@ -3135,7 +3135,7 @@ static lean_object* _init_l_Lean_Meta_Linear_instBEqCnstrKind___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3675____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3674____boxed), 2, 0); return x_1; } } @@ -3147,7 +3147,7 @@ x_1 = l_Lean_Meta_Linear_instBEqCnstrKind___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__1() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__1() { _start: { lean_object* x_1; @@ -3155,33 +3155,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Meta.Linear.CnstrKind.eq", 29); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__2() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__1; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____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_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__3() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3436____closed__4; -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__2; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____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_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__4() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__3; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__3; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3189,23 +3189,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__5() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3436____closed__5; -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__2; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____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_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__6() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__5; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__5; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3213,7 +3213,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__7() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__7() { _start: { lean_object* x_1; @@ -3221,33 +3221,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Meta.Linear.CnstrKind.div", 30); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__8() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__7; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__7; 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_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__9() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3436____closed__4; -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__8; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__8; 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_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__10() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__10() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__9; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__9; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3255,23 +3255,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__11() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3436____closed__5; -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__8; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__8; 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_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__12() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__12() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__11; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__11; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3279,7 +3279,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__13() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__13() { _start: { lean_object* x_1; @@ -3287,33 +3287,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Meta.Linear.CnstrKind.lt", 29); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__14() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__14() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__13; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__13; 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_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__15() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3436____closed__4; -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__14; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__14; 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_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__16() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__16() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__15; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__15; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3321,23 +3321,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__17() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3436____closed__5; -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__14; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__14; 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_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__18() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__18() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__17; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__17; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3345,7 +3345,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__19() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__19() { _start: { lean_object* x_1; @@ -3353,33 +3353,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Meta.Linear.CnstrKind.le", 29); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__20() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__20() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__19; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__19; 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_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__21() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3436____closed__4; -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__20; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__20; 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_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__22() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__22() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__21; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__21; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3387,23 +3387,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__23() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3436____closed__5; -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__20; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__20; 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_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__24() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__24() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__23; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__23; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3411,7 +3411,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691_(uint8_t x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690_(uint8_t x_1, lean_object* x_2) { _start: { switch (x_1) { @@ -3423,14 +3423,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_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__4; +x_5 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__4; 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_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__6; +x_7 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__6; x_8 = l_Repr_addAppParen(x_7, x_2); return x_8; } @@ -3443,14 +3443,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_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__10; +x_11 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__10; 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_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__12; +x_13 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__12; x_14 = l_Repr_addAppParen(x_13, x_2); return x_14; } @@ -3463,14 +3463,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_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__16; +x_17 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__16; 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_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__18; +x_19 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__18; x_20 = l_Repr_addAppParen(x_19, x_2); return x_20; } @@ -3483,14 +3483,14 @@ x_22 = lean_nat_dec_le(x_21, x_2); if (x_22 == 0) { lean_object* x_23; lean_object* x_24; -x_23 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__22; +x_23 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__22; x_24 = l_Repr_addAppParen(x_23, x_2); return x_24; } else { lean_object* x_25; lean_object* x_26; -x_25 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__24; +x_25 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__24; x_26 = l_Repr_addAppParen(x_25, x_2); return x_26; } @@ -3498,13 +3498,13 @@ return x_26; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____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_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691_(x_3, x_2); +x_4 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690_(x_3, x_2); lean_dec(x_2); return x_4; } @@ -3513,7 +3513,7 @@ static lean_object* _init_l_Lean_Meta_Linear_instReprCnstrKind___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____boxed), 2, 0); return x_1; } } @@ -3549,7 +3549,7 @@ x_1 = l_Lean_Meta_Linear_instInhabitedCnstr___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3875_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3874_(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; @@ -3599,11 +3599,11 @@ return x_18; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3875____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3874____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3875_(x_1, x_2); +x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3874_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -3614,7 +3614,7 @@ LEAN_EXPORT uint8_t l_Lean_Meta_Linear_instDecidableEqCnstr(lean_object* x_1, le _start: { uint8_t x_3; -x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3875_(x_1, x_2); +x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3874_(x_1, x_2); return x_3; } } @@ -3629,7 +3629,7 @@ x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4110_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4109_(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; @@ -3641,7 +3641,7 @@ x_7 = lean_ctor_get_uint8(x_2, sizeof(void*)*3); x_8 = lean_ctor_get(x_2, 0); x_9 = lean_ctor_get(x_2, 1); x_10 = lean_ctor_get(x_2, 2); -x_11 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3675_(x_3, x_7); +x_11 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3674_(x_3, x_7); if (x_11 == 0) { uint8_t x_12; @@ -3679,11 +3679,11 @@ return x_18; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4110____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4109____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4110_(x_1, x_2); +x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4109_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -3694,7 +3694,7 @@ static lean_object* _init_l_Lean_Meta_Linear_instBEqCnstr___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4110____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4109____boxed), 2, 0); return x_1; } } @@ -3706,7 +3706,7 @@ x_1 = l_Lean_Meta_Linear_instBEqCnstr___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__1() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__1() { _start: { lean_object* x_1; @@ -3714,33 +3714,33 @@ x_1 = lean_mk_string_from_bytes("kind", 4); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__2() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__1; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____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_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__3() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__2; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__2; x_3 = lean_alloc_ctor(4, 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_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__4() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__3; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__3; x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_148____closed__5; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -3748,7 +3748,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__5() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__5() { _start: { lean_object* x_1; @@ -3756,17 +3756,17 @@ x_1 = lean_mk_string_from_bytes("lhs", 3); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__6() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__5; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__5; 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_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__7() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__7() { _start: { lean_object* x_1; @@ -3774,17 +3774,17 @@ x_1 = lean_mk_string_from_bytes("rhs", 3); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__8() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__7; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__7; 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_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__9() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__9() { _start: { lean_object* x_1; @@ -3792,24 +3792,24 @@ x_1 = lean_mk_string_from_bytes("jst", 3); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__10() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__9; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__9; x_2 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212_(lean_object* x_1, lean_object* x_2) { _start: { uint8_t 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; 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; 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; 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; uint8_t x_42; lean_object* x_43; x_3 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); x_4 = lean_unsigned_to_nat(0u); -x_5 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691_(x_3, x_4); -x_6 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__4; +x_5 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690_(x_3, x_4); +x_6 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__4; x_7 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); @@ -3821,7 +3821,7 @@ x_10 = lean_box(1); x_11 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_11, 0, x_9); lean_ctor_set(x_11, 1, x_10); -x_12 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__6; +x_12 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__6; x_13 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_13, 0, x_11); lean_ctor_set(x_13, 1, x_12); @@ -3841,7 +3841,7 @@ lean_ctor_set(x_19, 1, x_8); x_20 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_20, 0, x_19); lean_ctor_set(x_20, 1, x_10); -x_21 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__8; +x_21 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__8; x_22 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_22, 0, x_20); lean_ctor_set(x_22, 1, x_21); @@ -3863,7 +3863,7 @@ lean_ctor_set(x_28, 1, x_8); x_29 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_29, 0, x_28); lean_ctor_set(x_29, 1, x_10); -x_30 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__10; +x_30 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__10; x_31 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_31, 0, x_29); lean_ctor_set(x_31, 1, x_30); @@ -3896,11 +3896,11 @@ lean_ctor_set_uint8(x_43, sizeof(void*)*1, x_42); return x_43; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213_(x_1, x_2); +x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212_(x_1, x_2); lean_dec(x_2); return x_3; } @@ -3909,7 +3909,7 @@ static lean_object* _init_l_Lean_Meta_Linear_instReprCnstr___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____boxed), 2, 0); return x_1; } } @@ -5827,54 +5827,54 @@ l_Lean_Meta_Linear_instBEqCnstrKind___closed__1 = _init_l_Lean_Meta_Linear_instB lean_mark_persistent(l_Lean_Meta_Linear_instBEqCnstrKind___closed__1); l_Lean_Meta_Linear_instBEqCnstrKind = _init_l_Lean_Meta_Linear_instBEqCnstrKind(); lean_mark_persistent(l_Lean_Meta_Linear_instBEqCnstrKind); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__1 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__1); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__2 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__2(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__2); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__3 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__3(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__3); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__4 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__4(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__4); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__5 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__5(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__5); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__6 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__6(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__6); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__7 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__7(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__7); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__8 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__8(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__8); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__9 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__9(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__9); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__10 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__10(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__10); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__11 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__11(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__11); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__12 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__12(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__12); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__13 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__13(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__13); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__14 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__14(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__14); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__15 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__15(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__15); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__16 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__16(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__16); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__17 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__17(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__17); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__18 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__18(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__18); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__19 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__19(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__19); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__20 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__20(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__20); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__21 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__21(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__21); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__22 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__22(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__22); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__23 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__23(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__23); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__24 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__24(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3691____closed__24); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__1 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__1); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__2 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__2); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__3 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__3); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__4 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__4); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__5 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__5); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__6 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__6); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__7 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__7); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__8 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__8); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__9 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__9(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__9); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__10 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__10(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__10); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__11 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__11(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__11); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__12 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__12(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__12); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__13 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__13(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__13); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__14 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__14(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__14); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__15 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__15(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__15); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__16 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__16(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__16); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__17 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__17(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__17); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__18 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__18(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__18); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__19 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__19(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__19); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__20 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__20(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__20); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__21 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__21(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__21); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__22 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__22(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__22); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__23 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__23(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__23); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__24 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__24(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3690____closed__24); l_Lean_Meta_Linear_instReprCnstrKind___closed__1 = _init_l_Lean_Meta_Linear_instReprCnstrKind___closed__1(); lean_mark_persistent(l_Lean_Meta_Linear_instReprCnstrKind___closed__1); l_Lean_Meta_Linear_instReprCnstrKind = _init_l_Lean_Meta_Linear_instReprCnstrKind(); @@ -5887,26 +5887,26 @@ l_Lean_Meta_Linear_instBEqCnstr___closed__1 = _init_l_Lean_Meta_Linear_instBEqCn lean_mark_persistent(l_Lean_Meta_Linear_instBEqCnstr___closed__1); l_Lean_Meta_Linear_instBEqCnstr = _init_l_Lean_Meta_Linear_instBEqCnstr(); lean_mark_persistent(l_Lean_Meta_Linear_instBEqCnstr); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__1 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__1); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__2 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__2(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__2); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__3 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__3(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__3); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__4 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__4(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__4); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__5 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__5(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__5); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__6 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__6(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__6); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__7 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__7(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__7); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__8 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__8(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__8); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__9 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__9(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__9); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__10 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__10(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4213____closed__10); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__1 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__1); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__2 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__2); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__3 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__3); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__4 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__4); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__5 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__5); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__6 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__6); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__7 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__7); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__8 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__8); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__9 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__9(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__9); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__10 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__10(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4212____closed__10); l_Lean_Meta_Linear_instReprCnstr___closed__1 = _init_l_Lean_Meta_Linear_instReprCnstr___closed__1(); lean_mark_persistent(l_Lean_Meta_Linear_instReprCnstr___closed__1); l_Lean_Meta_Linear_instReprCnstr = _init_l_Lean_Meta_Linear_instReprCnstr(); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Simp/SimpTheorems.c b/stage0/stdlib/Lean/Meta/Tactic/Simp/SimpTheorems.c index 627f404fed6..10ceeda1b4c 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Simp/SimpTheorems.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Simp/SimpTheorems.c @@ -15,52 +15,64 @@ extern "C" { #endif LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorem_getName___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_SimpTheorems_addDeclToUnfold___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_SimpTheorems_erase___rarg___closed__2; static lean_object* l_Lean_Meta_mkSimpAttr___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorems_pre___default; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__12; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__4; lean_object* l_Lean_ScopedEnvExtension_modifyState___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__99; lean_object* l_Lean_Meta_mkPropExt(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_SimpTheorems_toUnfoldThms___default___closed__2; size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorems_isLemma___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__50; LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorems_erase___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorems_erase___at_Lean_Meta_mkSimpAttr___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insert___at_Lean_Meta_addSimpTheoremEntry___spec__5(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__80; LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheoremsArray_isDeclToUnfold___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getSimpTheorems(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__33; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__94; static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__3___closed__7; lean_object* l_Lean_stringToMessageData(lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__78; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorems_addDeclToUnfold___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_commandRegister__simp__attr_________closed__2; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__91; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_SimpTheorems_eraseCore___spec__1(lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__27; lean_object* lean_mk_empty_array_with_capacity(lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__67; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_SimpTheorems_eraseCore___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__40; lean_object* lean_nat_div(lean_object*, lean_object*); lean_object* l_Lean_Meta_mkForallFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Std_Data_HashMap_0__Std_numBucketsForCapacity(lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__45; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_SimpTheorems___hyg_4550____closed__4; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__9; lean_object* l_Lean_LocalDecl_userName(lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__22; lean_object* l_Lean_throwError___at_Lean_registerTagAttribute___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__92; lean_object* lean_name_mk_string(lean_object*, lean_object*); static lean_object* l_panic___at_Lean_Meta_addSimpTheoremEntry___spec__15___closed__1; uint8_t lean_usize_dec_eq(size_t, size_t); lean_object* lean_array_uget(lean_object*, size_t); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_mkSimpTheoremCore___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkSimpAttr___lambda__1___closed__13; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__41; LEAN_EXPORT lean_object* l_Lean_Meta_SimpExtension_getTheorems(lean_object*, lean_object*, lean_object*, lean_object*); uint64_t lean_uint64_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_mkSimpAttr___spec__1(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__88; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__13; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__96; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__72; +static lean_object* l_Lean_Meta_commandRegister__simp__attr_______closed__9; static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_addSimpTheoremEntry_updateLemmaNames(lean_object*, lean_object*); lean_object* l_Lean_Meta_mkAppM(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__6; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_isRflTheorem___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_SimpTheorems_erase___rarg___closed__1; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); @@ -68,24 +80,20 @@ lean_object* l_Lean_throwError___at___private_Lean_Meta_Basic_0__Lean_Meta_proce lean_object* l_Lean_registerSimpleScopedEnvExtension___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_addSimpTheoremEntry___spec__9(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_commandRegister__simp__attr_______closed__8; static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addSimpTheorem___spec__1___closed__9; -static lean_object* l_Lean_Meta_commandRegister__simp__attr_________closed__4; static lean_object* l_Lean_Meta_mkSimpAttr___lambda__1___closed__6; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__75; LEAN_EXPORT lean_object* l_Std_mkHashMap___at_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_SimpTheorems___hyg_4455____spec__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorems_erase___at_Lean_Meta_mkSimpAttr___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addSimpTheorem___spec__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorem_getName(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertVal___at_Lean_Meta_addSimpTheoremEntry___spec__10(lean_object*, lean_object*); size_t lean_usize_sub(size_t, size_t); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__82; lean_object* lean_environment_find(lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__69; LEAN_EXPORT lean_object* l_Lean_Meta_addSimpTheorem(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_SimpTheorem_keys___default___closed__1; LEAN_EXPORT uint8_t l_Lean_Meta_instBEqSimpTheorem(lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__61; LEAN_EXPORT uint8_t l_Lean_Meta_SimpTheoremsArray_isDeclToUnfold(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_mkSimpTheoremCore___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorem_priority___default; @@ -94,25 +102,25 @@ lean_object* l_id___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Std_AssocList_replace___at_Lean_Meta_registerSimpAttr___spec__6(lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addSimpTheorem___spec__1___closed__8; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__11; LEAN_EXPORT uint8_t l_Lean_Meta_SimpTheorems_isDeclToUnfold(lean_object*, lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__59; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__7___boxed(lean_object*, 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_panic___at_Lean_Meta_addSimpTheoremEntry___spec__15(lean_object*); -static lean_object* l_Lean_Meta_commandRegister__simp__attr_________closed__9; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_isPerm(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__31; uint8_t lean_name_eq(lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__36; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__34; uint8_t l_Lean_Expr_isApp(lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_SimpTheorems___hyg_4550____closed__1; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_addSimpTheoremEntry___spec__12(lean_object*, lean_object*, size_t, size_t); +static lean_object* l_Lean_Meta_commandRegister__simp__attr_______closed__11; lean_object* l_Lean_mkIdentFrom(lean_object*, lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__38; static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___closed__2; lean_object* l_Lean_Expr_appFn_x21(lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__26; lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkSimpAttr___lambda__1___closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorems_add(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__21; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_findAtAux___at_Lean_Meta_addSimpTheoremEntry___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_DiscrTree_instInhabitedDiscrTree(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__3___closed__5; @@ -121,16 +129,20 @@ LEAN_EXPORT uint8_t l_Lean_Meta_SimpTheorem_perm___default; LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorems_lemmaNames___default; lean_object* lean_expr_instantiate1(lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__23; lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorems_registerDeclToUnfoldThms(lean_object*, lean_object*, lean_object*); lean_object* l_Std_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_SimpTheoremsArray_eraseTheorem___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkSimpAttr___lambda__1___closed__17; +static lean_object* l_Lean_Meta_commandRegister__simp__attr_______closed__4; +static lean_object* l_Lean_Meta_commandRegister__simp__attr_______closed__17; static lean_object* l_Lean_Meta_mkSimpAttr___lambda__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__58; LEAN_EXPORT lean_object* l_Lean_Meta_mkSimpAttr___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__58; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__75; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_SimpTheoremsArray_isDeclToUnfold___spec__1(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorems_addDeclToUnfold(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_right(size_t, size_t); @@ -142,160 +154,150 @@ static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_ static lean_object* l_Lean_Meta_isRflProofCore___closed__8; lean_object* lean_string_utf8_byte_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_isRflProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__82; static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_shouldPreprocess___closed__1; LEAN_EXPORT lean_object* l_Std_AssocList_contains___at_Lean_Meta_registerSimpAttr___spec__2___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__61; static lean_object* l_Lean_Meta_instInhabitedSimpTheorems___closed__2; uint8_t lean_usize_dec_lt(size_t, size_t); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__40; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__9; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_checkTypeIsProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_nameLitKind; -static lean_object* l_Lean_Meta_commandRegister__simp__attr_________closed__15; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__67; LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorems_addDeclToUnfoldCore(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___closed__4; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__71; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__7; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__31; lean_object* lean_nat_add(lean_object*, lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__12; static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___closed__1; LEAN_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_mkSimpTheoremsFromConst___spec__1___boxed(lean_object*, lean_object*, lean_object*, 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___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_mkSimpTheoremsFromConst___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorems_erase(lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__63; LEAN_EXPORT lean_object* l_List_mapM___at___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__24; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__83; LEAN_EXPORT uint8_t l_Array_contains___at_Lean_Meta_addSimpTheoremEntry___spec__11(lean_object*, lean_object*); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); size_t lean_uint64_to_usize(uint64_t); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__97; uint8_t l___private_Lean_Meta_DiscrTreeTypes_0__Lean_Meta_DiscrTree_beqKey____x40_Lean_Meta_DiscrTreeTypes___hyg_89_(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_commandRegister__simp__attr_________closed__11; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__74; LEAN_EXPORT uint8_t l_Std_PersistentHashMap_contains___at_Lean_Meta_SimpTheorems_erase___spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__37; LEAN_EXPORT lean_object* l_Std_HashMap_insert___at_Lean_Meta_registerSimpAttr___spec__1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__64; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__102; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__90; lean_object* l_Lean_getAttrParamOptPrio(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__3___closed__12; LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorem_keys___default; lean_object* l_Lean_ConstantInfo_levelParams(lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__26; static lean_object* l_Lean_Meta_instInhabitedSimpTheorems___closed__1; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__55; static lean_object* l_Lean_Meta_mkSimpAttr___lambda__1___closed__19; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_SimpTheorems___hyg_4550____closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_isRflTheorem(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); lean_object* l_Lean_mkProj(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_addSimpTheoremEntry___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__79; lean_object* l_Lean_Meta_mkEqFalse(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Std_PersistentHashMap_containsAtAux___at_Lean_Meta_SimpTheorems_erase___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__19; static lean_object* l_Lean_Meta_instToFormatSimpTheorem___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorems_toUnfoldThms___default; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__51; LEAN_EXPORT lean_object* l_Lean_Meta_instBEqSimpTheorem___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__42; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__53; +static lean_object* l_Lean_Meta_commandRegister__simp__attr_______closed__12; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_SimpTheorems___hyg_4550____closed__5; static lean_object* l_Lean_Meta_SimpTheorems_toUnfoldThms___default___closed__3; static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___closed__7; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__100; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__3; lean_object* l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_SimpTheorems_addConst___spec__1(lean_object*, size_t, size_t, lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_forallMetaTelescopeReducingAux(lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); lean_object* l_Lean_ScopedEnvExtension_addLocalEntry___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__49; lean_object* lean_nat_sub(lean_object*, lean_object*); lean_object* l_Std_PersistentHashMap_insert___at_Lean_NameSSet_insert___spec__2(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__95; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__35; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocessProof(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashMapImp_moveEntries___at_Lean_Meta_registerSimpAttr___spec__4(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__16; LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorems_addDeclToUnfold___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_createNodes___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_addSimpTheoremEntry___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorems_add_getName_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__36; static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___closed__3; static lean_object* l_Lean_Meta_SimpTheorems_addDeclToUnfold___closed__1; LEAN_EXPORT uint8_t l_Lean_Meta_SimpTheorems_isLemma(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_mkSimpTheoremCore___lambda__2___boxed(lean_object*, lean_object*, 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___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_mkSimpTheoremCore___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__43; lean_object* l_Lean_Meta_mkEqSymm(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorems_add___boxed(lean_object*, lean_object*, 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_Std_PersistentHashMap_find_x3f___at_Lean_Meta_addSimpTheoremEntry___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_simpExtension; static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_mkSimpTheoremCore___lambda__2___closed__1; static size_t l_Std_PersistentHashMap_findAux___at_Lean_Meta_addSimpTheoremEntry___spec__3___closed__1; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__20; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__20; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__73; static lean_object* l_Lean_Meta_instToFormatSimpTheorem___closed__6; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__101; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__32; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__10; static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_mkSimpTheoremCore___lambda__2___closed__2; lean_object* l_Lean_Name_toString(lean_object*, uint8_t); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__101; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_mkSimpTheorems___spec__1(lean_object*, uint8_t, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__43; lean_object* l_Lean_Meta_mkLambdaFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_AssocList_find_x3f___at_Lean_Meta_getSimpExtension_x3f___spec__2___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_commandRegister__simp__attr_________closed__12; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addSimpTheoremEntry___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__68; static size_t l_Std_PersistentHashMap_findAux___at_Lean_Meta_addSimpTheoremEntry___spec__3___closed__2; lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__39; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__65; lean_object* l_Std_mkHashMapImp___rarg(lean_object*); LEAN_EXPORT lean_object* l_Std_AssocList_find_x3f___at_Lean_Meta_getSimpExtension_x3f___spec__2(lean_object*, lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__69; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__11; lean_object* l_List_mapM___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_mkFun___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_panic___at_Lean_Expr_getRevArg_x21___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_contains___at_Lean_Meta_SimpTheorems_erase___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_mkHashMap___at_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_SimpTheorems___hyg_4455____spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getSimpTheorems___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__87; static lean_object* l_Lean_Meta_mkSimpExt___closed__1; static lean_object* l_Lean_Meta_isRflProofCore___closed__6; static lean_object* l_Lean_Meta_getSimpTheorems___closed__1; uint64_t l_Lean_Name_hash(lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__85; lean_object* l_Nat_repr(lean_object*); static lean_object* l_Lean_Meta_SimpTheorem_getName___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_commandRegister__simp__attr____; static lean_object* l_Lean_Meta_SimpTheorem_getName___closed__2; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__102; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__16; lean_object* lean_st_mk_ref(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Meta_isRflTheorem___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkSimpExt___lambda__1(lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__66; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__56; lean_object* l_Lean_Syntax_getId(lean_object*); -lean_object* l___private_Init_Meta_0__Lean_quoteNameMk(lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__56; extern lean_object* l_Lean_Expr_instHashableExpr; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__77; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__66; static lean_object* l_Lean_Meta_mkSimpAttr___lambda__1___closed__11; -static lean_object* l_Lean_Meta_commandRegister__simp__attr_________closed__1; lean_object* l_Std_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__2___closed__2; static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___closed__5; -static lean_object* l_Lean_Meta_commandRegister__simp__attr_________closed__17; lean_object* l_Array_back___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Meta_SimpTheorems_toUnfoldThms___default___closed__1; uint8_t l_Lean_ConstantInfo_hasValue(lean_object*); size_t lean_usize_shift_left(size_t, size_t); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__90; lean_object* l_Lean_Meta_mkAuxLemma(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_InfoCacheKey_instHashableInfoCacheKey___boxed(lean_object*); -static lean_object* l_Lean_Meta_commandRegister__simp__attr_________closed__8; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__77; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__98; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__81; LEAN_EXPORT uint8_t l_Std_PersistentHashMap_containsAux___at_Lean_Meta_SimpTheorems_erase___spec__2(lean_object*, size_t, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__80; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorems_erase___at_Lean_Meta_mkSimpAttr___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_mkStrLit(lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkSimpAttr___lambda__1___closed__20; uint8_t l_Lean_Expr_isAppOfArity(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_commandRegister__simp__attr_______closed__10; static lean_object* l_Lean_Meta_isRflProofCore___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_isPerm___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__48; static lean_object* l_Lean_Meta_isRflProofCore___closed__2; static lean_object* l_Lean_Meta_SimpTheorem_getValue___closed__4; size_t lean_usize_modn(size_t, lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__68; static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__3___closed__13; lean_object* l_Lean_registerBuiltinAttribute(lean_object*, lean_object*); lean_object* l_Lean_Meta_mkEqTrue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -306,15 +308,17 @@ static lean_object* l_Lean_Meta_instToFormatSimpTheorem___closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorems_isDeclToUnfold___boxed(lean_object*, lean_object*); uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SimpExtension_getTheorems___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__85; LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_SimpTheorems___hyg_4550_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_SimpTheorems___hyg_4455_(lean_object*); +static lean_object* l_Lean_Meta_commandRegister__simp__attr_______closed__7; size_t lean_usize_mul(size_t, size_t); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__23; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__49; static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__3___closed__9; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__24; LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorems_eraseCore(lean_object*, lean_object*); uint8_t l_Lean_Expr_isForall(lean_object*); LEAN_EXPORT lean_object* l_Array_contains___at_Lean_Meta_addSimpTheoremEntry___spec__11___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__59; lean_object* lean_whnf(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_containsAtAux___at_Lean_Meta_SimpTheorems_erase___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_instInhabitedSimpTheorems; @@ -334,59 +338,53 @@ static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_ lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_SimpTheoremsArray_isErased___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__34; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__88; lean_object* l_Lean_ScopedEnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkSimpExt(lean_object*, lean_object*); size_t lean_usize_land(size_t, size_t); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__92; LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_isPerm___spec__1(lean_object*); -static lean_object* l_Lean_Meta_commandRegister__simp__attr_________closed__3; static lean_object* l_Lean_Meta_SimpTheorem_getValue___closed__2; static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addSimpTheorem___spec__1___closed__7; lean_object* l_Lean_Meta_DiscrTree_mkPath(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_shouldPreprocess(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__86; LEAN_EXPORT lean_object* l___private_Init_Data_Array_BinSearch_0__Array_binInsertAux___at_Lean_Meta_addSimpTheoremEntry___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_checkTypeIsProp___closed__1; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__93; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__5; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__25; LEAN_EXPORT lean_object* l_List_mapM___at___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkSimpAttr___lambda__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_mkSimpTheoremCore___lambda__1(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorems_erased___default; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__87; static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_checkTypeIsProp___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__3___closed__4; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__30; LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorem_getValue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__76; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__78; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__91; lean_object* l_String_intercalate(lean_object*, lean_object*); lean_object* l_List_redLength___rarg(lean_object*); -static lean_object* l_Lean_Meta_commandRegister__simp__attr_________closed__16; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_mkSimpTheoremCore___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__93; static lean_object* l_Lean_getConstInfo___at_Lean_Meta_isRflTheorem___spec__1___closed__3; static lean_object* l_Lean_Meta_mkSimpAttr___lambda__1___closed__4; lean_object* lean_name_append_after(lean_object*, lean_object*); extern lean_object* l_Lean_Meta_DiscrTree_instInhabitedKey; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__14; static lean_object* l_Lean_Meta_mkSimpAttr___lambda__1___closed__9; static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__3___closed__8; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__28; static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheoremsArray_eraseTheorem(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_SimpTheoremsArray_isErased___spec__1(lean_object*, lean_object*, size_t, size_t); lean_object* l_Lean_Meta_whnfR(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_addSimpTheoremEntry___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_commandRegister__simp__attr_______closed__14; static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__3___closed__11; -static lean_object* l_Lean_Meta_commandRegister__simp__attr_________closed__7; uint8_t lean_expr_eqv(lean_object*, lean_object*); lean_object* l_Lean_Meta_getEqnsFor_x3f(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_commandRegister__simp__attr_______closed__5; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__95; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__8(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__54; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__62; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__79; uint8_t lean_nat_dec_le(lean_object*, lean_object*); uint8_t lean_usize_dec_le(size_t, size_t); static lean_object* l_Lean_Meta_isRflProofCore___closed__7; @@ -394,23 +392,26 @@ static lean_object* l_Lean_Meta_mkSimpAttr___lambda__1___closed__10; LEAN_EXPORT lean_object* l_Lean_Meta_registerSimpAttr(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___closed__10; LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_isPerm___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__52; lean_object* l_Lean_Name_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorem_name_x3f___default; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__84; +static lean_object* l_Lean_Meta_commandRegister__simp__attr_______closed__16; lean_object* l_Lean_Syntax_getKind(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__2___closed__3; static lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_addSimpTheoremEntry___spec__9___closed__1; +lean_object* l_Lean_quoteNameMk(lean_object*); static lean_object* l_Lean_Meta_SimpTheorems_pre___default___closed__1; static lean_object* l_Lean_Meta_mkSimpExt___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorems_erase___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_checkBadRewrite___closed__1; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__48; +static lean_object* l_Lean_Meta_commandRegister__simp__attr_______closed__15; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_findAux___at_Lean_Meta_addSimpTheoremEntry___spec__3(lean_object*, size_t, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__29; static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_checkBadRewrite___closed__2; lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__2___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__15; static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__1___closed__2; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__84; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__62; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__54; static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addSimpTheorem___spec__1___closed__5; static lean_object* l_Lean_Meta_mkSimpExt___closed__2; uint8_t l_Lean_Meta_DiscrTree_Key_lt(lean_object*, lean_object*); @@ -419,18 +420,17 @@ static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_ lean_object* l_List_mapTRAux___at_Lean_mkConstWithLevelParams___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocessProof___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getSimpExtension_x3f(lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__18; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__22; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__33; lean_object* lean_nat_mul(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_instToMessageDataSimpTheorem(lean_object*); +lean_object* l_Lean_Syntax_mkNameLit(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashMapImp_expand___at_Lean_Meta_registerSimpAttr___spec__3(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_commandRegister__simp__attr_________closed__5; lean_object* l_Lean_throwError___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addSimpTheorem___spec__1___closed__3; static lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_addSimpTheoremEntry___spec__9___closed__3; static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__3___closed__3; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__27; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__70; LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorems_add_getName_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isFVar(lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); @@ -438,49 +438,45 @@ static lean_object* l_Lean_Meta_mkSimpAttr___lambda__1___closed__18; static lean_object* l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_addSimpTheoremEntry___spec__1___closed__1; lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkSimpAttr___lambda__1___closed__15; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__45; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_isRflTheorem___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__63; static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_checkBadRewrite___closed__3; lean_object* l_List_toArrayAux___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_SimpTheoremsArray_isDeclToUnfold___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isNone(lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__83; static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__3___closed__6; lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__81; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__29; static lean_object* l_Lean_Meta_mkSimpAttr___lambda__1___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go(uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isExprDefEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__52; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_addSimpTheorem___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_addSimpTheoremEntry___spec__1___closed__2; +static lean_object* l_Lean_Meta_commandRegister__simp__attr_______closed__6; static lean_object* l_Lean_Meta_isRflProofCore___closed__5; LEAN_EXPORT lean_object* l_Array_binInsertM___at_Lean_Meta_addSimpTheoremEntry___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_commandRegister__simp__attr_________closed__14; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__51; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__46; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__86; LEAN_EXPORT lean_object* l_Lean_Meta_mkSimpAttr___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_mkSimpTheoremCore___boxed(lean_object*, 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___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_shouldPreprocess___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__35; static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addSimpTheorem___spec__1___closed__1; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__98; LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorems_post___default; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__73; lean_object* l_instHashableProd___rarg___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__32; LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addSimpTheorem___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__10; LEAN_EXPORT lean_object* l_Lean_Meta_mkSimpAttr___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_instantiateMVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__25; LEAN_EXPORT lean_object* l_Lean_Meta_mkSimpAttr___lambda__1(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_SimpTheorems_addConst___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_instInhabitedSimpTheorem; lean_object* lean_mk_array(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocessProof___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_instToFormatSimpTheorem___closed__2; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__39; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__55; uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorems_toUnfold___default; static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__2___closed__1; @@ -490,9 +486,8 @@ extern lean_object* l_Lean_Expr_instBEqExpr; static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___closed__9; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAux___at_Lean_Meta_addSimpTheoremEntry___spec__6(lean_object*, size_t, size_t, lean_object*, lean_object*); static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addSimpTheorem___spec__1___closed__2; +static lean_object* l_Lean_Meta_commandRegister__simp__attr_______closed__3; lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_openAbstractMVarsResult___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__19; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__65; LEAN_EXPORT uint8_t l_Std_AssocList_contains___at_Lean_Meta_registerSimpAttr___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocessProof___spec__1(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_mkSimpTheoremsFromConst(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -501,117 +496,121 @@ LEAN_EXPORT lean_object* l___private_Init_Data_Array_BinSearch_0__Array_binInser LEAN_EXPORT lean_object* l_Std_HashMapImp_find_x3f___at_Lean_Meta_getSimpExtension_x3f___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addSimpTheorem___spec__1___closed__6; lean_object* l_Lean_Expr_instantiateLevelParamsCore_visit___at_Lean_Expr_instantiateLevelParamsArray___spec__1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorems_addConst___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__94; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__7; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__71; LEAN_EXPORT uint8_t l_Lean_Meta_SimpTheorem_post___default; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__6(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__42; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__100; static lean_object* l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_addSimpTheoremEntry___spec__1___closed__3; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__8; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__37; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); static lean_object* l_Lean_Meta_instToFormatSimpTheorem___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_isRflProofCore___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheoremsArray_addTheorem(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkSimpAttr___lambda__1___closed__7; -static lean_object* l_Lean_Meta_commandRegister__simp__attr_________closed__13; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_mkSimpAttr___spec__1___boxed(lean_object*, lean_object*, lean_object*, 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_Meta_SimpTheorems_addConst(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__57; static lean_object* l_Lean_Meta_mkSimpAttr___lambda__1___closed__16; lean_object* l_Lean_Expr_getAppFn(lean_object*); static lean_object* l_Lean_getConstInfo___at_Lean_Meta_isRflTheorem___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorems_erase___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_simpExtensionMapRef; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__30; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__18; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__76; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_checkBadRewrite(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_registerSimpAttr___closed__1; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_SimpTheorems___hyg_4550____closed__2; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAux___at_Lean_Meta_addSimpTheoremEntry___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_SimpTheorem_getValue___closed__3; static lean_object* l_Std_PersistentHashMap_insertAux___at_Lean_Meta_addSimpTheoremEntry___spec__6___closed__1; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__60; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__74; static lean_object* l_Lean_Meta_mkSimpAttr___lambda__1___closed__12; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__47; LEAN_EXPORT lean_object* l_Std_AssocList_foldlM___at_Lean_Meta_registerSimpAttr___spec__5(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_mkSimpTheorems___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfo___at_Lean_Meta_isRflTheorem___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_instToFormatSimpTheorem(lean_object*); uint8_t l_List_isEmpty___rarg(lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__2; -static lean_object* l_Lean_Meta_commandRegister__simp__attr_________closed__6; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__2; +static lean_object* l_Lean_Meta_commandRegister__simp__attr_______closed__1; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__89; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_instInhabitedSimpTheorem___closed__3; LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Meta_isRflTheorem___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__47; LEAN_EXPORT lean_object* l_Lean_Meta_addSimpTheorem___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__60; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__89; lean_object* lean_usize_to_nat(size_t); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLetDeclImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_PersistentHashMap_insert___at_Lean_Meta_getEqnsFor_x3f___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_SimpTheoremsArray_isErased(lean_object*, lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__53; lean_object* l_Lean_Meta_mkEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_WHNF_0__Lean_Meta_toCtorWhenK___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_PersistentHashMap_find_x3f___at_Lean_Meta_getEqnsFor_x3f___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Meta_isProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_mkSimpTheoremCore(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_indentExpr(lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__14; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__28; LEAN_EXPORT lean_object* l_Lean_Meta_addSimpTheoremEntry(lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__50; +static lean_object* l_Lean_Meta_commandRegister__simp__attr_______closed__2; LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addSimpTheorem___spec__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__15; uint8_t l_Lean_Meta_hasSmartUnfoldingDecl(lean_object*, lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__3; uint8_t l_Std_PersistentHashMap_contains___at_Lean_NameSSet_contains___spec__3(lean_object*, lean_object*); lean_object* l_Lean_mkConst(lean_object*, lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__1; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_findAtAux___at_Lean_Meta_addSimpTheoremEntry___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addSimpTheoremEntry___spec__7(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_constName_x21(lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__1; lean_object* l_Array_insertAt___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkSimpTheorems(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_commandRegister__simp__attr_________closed__10; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__8___boxed(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_Meta_mkSimpAttr(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___closed__11; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__17; lean_object* l_Std_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__17; lean_object* l_List_appendTR___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Meta_instInhabitedSimpTheorem___closed__2; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__44; lean_object* l_Std_PersistentHashMap_erase___at_Lean_Meta_Instances_eraseCore___spec__1(lean_object*, lean_object*); uint64_t l_Lean_Meta_DiscrTree_Key_hash(lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_containsAux___at_Lean_Meta_SimpTheorems_erase___spec__2___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__57; lean_object* lean_expr_update_const(lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkSimpAttr___lambda__1___closed__14; static lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_addSimpTheoremEntry___spec__9___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_instInhabitedSimpEntry; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__44; static lean_object* l_Lean_Meta_instInhabitedSimpEntry___closed__1; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__70; static lean_object* l_Lean_getConstInfo___at_Lean_Meta_isRflTheorem___spec__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_mkSimpAttr___lambda__2___boxed(lean_object*, lean_object*); lean_object* l_Lean_Meta_DiscrTree_instInhabitedTrie(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_shouldPreprocess___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__13; static lean_object* l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_addSimpTheoremEntry___spec__1___closed__4; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__21; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_addSimpTheorem___spec__2(lean_object*, uint8_t, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__41; lean_object* l_Lean_Meta_DiscrTree_empty(lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__99; LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheoremsArray_isErased___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorem_levelParams___default; -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__46; lean_object* l_Std_PersistentHashMap_mkCollisionNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_mkSimpTheoremsFromConst___spec__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static uint64_t l_Lean_Meta_instInhabitedSimpTheorem___closed__1; -lean_object* l_Lean_Syntax_mkLit(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__72; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__64; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__96; +static lean_object* l_Lean_Meta_commandRegister__simp__attr_______closed__13; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__7(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_getFVarLocalDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__38; -LEAN_EXPORT lean_object* l_Lean_Meta_commandRegister__simp__attr______; static lean_object* l_Lean_Meta_mkSimpAttr___lambda__1___closed__3; +static lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__97; static lean_object* _init_l_Lean_Meta_SimpTheorem_keys___default___closed__1() { _start: { @@ -15366,7 +15365,7 @@ x_4 = lean_box(x_3); return x_4; } } -static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr_________closed__1() { +static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr_______closed__1() { _start: { lean_object* x_1; @@ -15374,35 +15373,35 @@ x_1 = lean_mk_string_from_bytes("Meta", 4); return x_1; } } -static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr_________closed__2() { +static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr_______closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Meta_mkSimpAttr___lambda__1___closed__14; -x_2 = l_Lean_Meta_commandRegister__simp__attr_________closed__1; +x_2 = l_Lean_Meta_commandRegister__simp__attr_______closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr_________closed__3() { +static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr_______closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("commandRegister_simp_attr___", 28); +x_1 = lean_mk_string_from_bytes("commandRegister_simp_attr__", 27); return x_1; } } -static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr_________closed__4() { +static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr_______closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_commandRegister__simp__attr_________closed__2; -x_2 = l_Lean_Meta_commandRegister__simp__attr_________closed__3; +x_1 = l_Lean_Meta_commandRegister__simp__attr_______closed__2; +x_2 = l_Lean_Meta_commandRegister__simp__attr_______closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr_________closed__5() { +static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr_______closed__5() { _start: { lean_object* x_1; @@ -15410,17 +15409,17 @@ x_1 = lean_mk_string_from_bytes("andthen", 7); return x_1; } } -static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr_________closed__6() { +static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr_______closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta_commandRegister__simp__attr_________closed__5; +x_2 = l_Lean_Meta_commandRegister__simp__attr_______closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr_________closed__7() { +static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr_______closed__7() { _start: { lean_object* x_1; @@ -15428,17 +15427,17 @@ x_1 = lean_mk_string_from_bytes("register_simp_attr", 18); return x_1; } } -static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr_________closed__8() { +static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr_______closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_commandRegister__simp__attr_________closed__7; +x_1 = l_Lean_Meta_commandRegister__simp__attr_______closed__7; x_2 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr_________closed__9() { +static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr_______closed__9() { _start: { lean_object* x_1; @@ -15446,33 +15445,33 @@ x_1 = lean_mk_string_from_bytes("ident", 5); return x_1; } } -static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr_________closed__10() { +static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr_______closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta_commandRegister__simp__attr_________closed__9; +x_2 = l_Lean_Meta_commandRegister__simp__attr_______closed__9; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr_________closed__11() { +static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr_______closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_commandRegister__simp__attr_________closed__10; +x_1 = l_Lean_Meta_commandRegister__simp__attr_______closed__10; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr_________closed__12() { +static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr_______closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Meta_commandRegister__simp__attr_________closed__6; -x_2 = l_Lean_Meta_commandRegister__simp__attr_________closed__8; -x_3 = l_Lean_Meta_commandRegister__simp__attr_________closed__11; +x_1 = l_Lean_Meta_commandRegister__simp__attr_______closed__6; +x_2 = l_Lean_Meta_commandRegister__simp__attr_______closed__8; +x_3 = l_Lean_Meta_commandRegister__simp__attr_______closed__11; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -15480,7 +15479,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr_________closed__13() { +static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr_______closed__13() { _start: { lean_object* x_1; @@ -15488,33 +15487,33 @@ x_1 = lean_mk_string_from_bytes("str", 3); return x_1; } } -static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr_________closed__14() { +static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr_______closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta_commandRegister__simp__attr_________closed__13; +x_2 = l_Lean_Meta_commandRegister__simp__attr_______closed__13; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr_________closed__15() { +static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr_______closed__15() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_commandRegister__simp__attr_________closed__14; +x_1 = l_Lean_Meta_commandRegister__simp__attr_______closed__14; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr_________closed__16() { +static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr_______closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Meta_commandRegister__simp__attr_________closed__6; -x_2 = l_Lean_Meta_commandRegister__simp__attr_________closed__12; -x_3 = l_Lean_Meta_commandRegister__simp__attr_________closed__15; +x_1 = l_Lean_Meta_commandRegister__simp__attr_______closed__6; +x_2 = l_Lean_Meta_commandRegister__simp__attr_______closed__12; +x_3 = l_Lean_Meta_commandRegister__simp__attr_______closed__15; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -15522,13 +15521,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr_________closed__17() { +static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr_______closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Meta_commandRegister__simp__attr_________closed__4; +x_1 = l_Lean_Meta_commandRegister__simp__attr_______closed__4; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Meta_commandRegister__simp__attr_________closed__16; +x_3 = l_Lean_Meta_commandRegister__simp__attr_______closed__16; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -15536,15 +15535,15 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr______() { +static lean_object* _init_l_Lean_Meta_commandRegister__simp__attr____() { _start: { lean_object* x_1; -x_1 = l_Lean_Meta_commandRegister__simp__attr_________closed__17; +x_1 = l_Lean_Meta_commandRegister__simp__attr_______closed__17; return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__1() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -15554,7 +15553,7 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__2() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__2() { _start: { lean_object* x_1; @@ -15562,17 +15561,17 @@ x_1 = lean_mk_string_from_bytes("Attr", 4); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__3() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__1; -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__2; +x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__1; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__2; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__4() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__4() { _start: { lean_object* x_1; @@ -15580,17 +15579,17 @@ x_1 = lean_mk_string_from_bytes("null", 4); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__5() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__4; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__4; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__6() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__6() { _start: { lean_object* x_1; @@ -15598,17 +15597,17 @@ x_1 = lean_mk_string_from_bytes("Command", 7); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__7() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Meta_mkSimpAttr___lambda__1___closed__16; -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__6; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__6; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__8() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__8() { _start: { lean_object* x_1; @@ -15616,22 +15615,22 @@ x_1 = lean_mk_string_from_bytes("initialize", 10); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__9() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__7; -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__8; +x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__7; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__8; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__10() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__5; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__5; x_3 = l_Lean_Meta_SimpTheorem_keys___default___closed__1; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -15640,7 +15639,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__11() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__11() { _start: { lean_object* x_1; @@ -15648,22 +15647,22 @@ x_1 = lean_mk_string_from_bytes("ext", 3); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__12() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__11; +x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__11; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__13() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__11; +x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__11; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__12; +x_3 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__12; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -15671,17 +15670,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__14() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__11; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__11; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__15() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__15() { _start: { lean_object* x_1; @@ -15689,17 +15688,17 @@ x_1 = lean_mk_string_from_bytes("Term", 4); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__16() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Meta_mkSimpAttr___lambda__1___closed__16; -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__15; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__15; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__17() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__17() { _start: { lean_object* x_1; @@ -15707,17 +15706,17 @@ x_1 = lean_mk_string_from_bytes("typeSpec", 8); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__18() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__16; -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__17; +x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__16; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__17; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__19() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__19() { _start: { lean_object* x_1; @@ -15725,22 +15724,22 @@ x_1 = lean_mk_string_from_bytes("SimpExtension", 13); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__20() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__20() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__19; +x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__19; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__21() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__19; +x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__19; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__20; +x_3 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__20; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -15748,51 +15747,51 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__22() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__19; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__19; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__23() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_commandRegister__simp__attr_________closed__2; -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__19; +x_1 = l_Lean_Meta_commandRegister__simp__attr_______closed__2; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__19; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__24() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__23; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__23; 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_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__25() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__24; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__24; x_3 = lean_alloc_ctor(1, 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_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__26() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__26() { _start: { lean_object* x_1; lean_object* x_2; @@ -15801,7 +15800,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__27() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__27() { _start: { lean_object* x_1; @@ -15809,7 +15808,7 @@ x_1 = lean_mk_string_from_bytes("←", 3); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__28() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__28() { _start: { lean_object* x_1; lean_object* x_2; @@ -15818,7 +15817,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__29() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__29() { _start: { lean_object* x_1; @@ -15826,17 +15825,17 @@ x_1 = lean_mk_string_from_bytes("doSeqIndent", 11); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__30() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__30() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__16; -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__29; +x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__16; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__29; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__31() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__31() { _start: { lean_object* x_1; @@ -15844,17 +15843,17 @@ x_1 = lean_mk_string_from_bytes("doSeqItem", 9); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__32() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__32() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__16; -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__31; +x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__16; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__31; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__33() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__33() { _start: { lean_object* x_1; @@ -15862,17 +15861,17 @@ x_1 = lean_mk_string_from_bytes("doExpr", 6); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__34() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__34() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__16; -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__33; +x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__16; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__33; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__35() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__35() { _start: { lean_object* x_1; @@ -15880,17 +15879,17 @@ x_1 = lean_mk_string_from_bytes("app", 3); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__36() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__36() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__16; -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__35; +x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__16; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__35; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__37() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__37() { _start: { lean_object* x_1; @@ -15898,22 +15897,22 @@ x_1 = lean_mk_string_from_bytes("registerSimpAttr", 16); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__38() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__38() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__37; +x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__37; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__39() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__39() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__37; +x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__37; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__38; +x_3 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__38; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -15921,51 +15920,51 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__40() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__40() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__37; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__37; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__41() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__41() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_commandRegister__simp__attr_________closed__2; -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__37; +x_1 = l_Lean_Meta_commandRegister__simp__attr_______closed__2; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__37; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__42() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__42() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__41; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__41; 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_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__43() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__43() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__42; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__42; x_3 = lean_alloc_ctor(1, 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_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__44() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__44() { _start: { lean_object* x_1; lean_object* x_2; @@ -15974,17 +15973,17 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__45() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__45() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__44; -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__10; +x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__44; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__10; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__46() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__46() { _start: { lean_object* x_1; @@ -15992,17 +15991,17 @@ x_1 = lean_mk_string_from_bytes("syntax", 6); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__47() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__47() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__7; -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__46; +x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__7; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__46; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__48() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__48() { _start: { lean_object* x_1; @@ -16010,33 +16009,33 @@ x_1 = lean_mk_string_from_bytes("attrKind", 8); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__49() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__49() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__16; -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__48; +x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__16; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__48; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__50() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__50() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__3___closed__9; -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__10; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__10; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__51() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__51() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__49; -x_3 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__50; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__49; +x_3 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__50; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -16044,7 +16043,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__52() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__52() { _start: { lean_object* x_1; @@ -16052,17 +16051,17 @@ x_1 = lean_mk_string_from_bytes("namedName", 9); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__53() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__53() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__7; -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__52; +x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__7; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__52; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__54() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__54() { _start: { lean_object* x_1; @@ -16070,7 +16069,7 @@ x_1 = lean_mk_string_from_bytes("(", 1); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__55() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__55() { _start: { lean_object* x_1; @@ -16078,7 +16077,7 @@ x_1 = lean_mk_string_from_bytes("name", 4); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__56() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__56() { _start: { lean_object* x_1; @@ -16086,7 +16085,7 @@ x_1 = lean_mk_string_from_bytes(":=", 2); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__57() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__57() { _start: { lean_object* x_1; @@ -16094,7 +16093,7 @@ x_1 = lean_mk_string_from_bytes(")", 1); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__58() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__58() { _start: { lean_object* x_1; lean_object* x_2; @@ -16103,7 +16102,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__59() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__59() { _start: { lean_object* x_1; @@ -16111,17 +16110,17 @@ x_1 = lean_mk_string_from_bytes("Syntax", 6); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__60() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__60() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Meta_mkSimpAttr___lambda__1___closed__16; -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__59; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__59; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__61() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__61() { _start: { lean_object* x_1; @@ -16129,17 +16128,17 @@ x_1 = lean_mk_string_from_bytes("atom", 4); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__62() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__62() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__60; -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__61; +x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__60; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__61; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__63() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__63() { _start: { lean_object* x_1; @@ -16147,17 +16146,17 @@ x_1 = lean_mk_string_from_bytes("stx_?", 5); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__64() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__64() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__63; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__63; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__65() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__65() { _start: { lean_object* x_1; @@ -16165,17 +16164,17 @@ x_1 = lean_mk_string_from_bytes("paren", 5); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__66() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__66() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__60; -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__65; +x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__60; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__65; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__67() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__67() { _start: { lean_object* x_1; @@ -16183,17 +16182,17 @@ x_1 = lean_mk_string_from_bytes("stx_<|>_", 8); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__68() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__68() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__67; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__67; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__69() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__69() { _start: { lean_object* x_1; @@ -16201,17 +16200,17 @@ x_1 = lean_mk_string_from_bytes("cat", 3); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__70() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__70() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__60; -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__69; +x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__60; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__69; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__71() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__71() { _start: { lean_object* x_1; @@ -16219,22 +16218,22 @@ x_1 = lean_mk_string_from_bytes("Parser.Tactic.simpPre", 21); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__72() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__72() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__71; +x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__71; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__73() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__73() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__71; +x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__71; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__72; +x_3 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__72; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -16242,17 +16241,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__74() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__74() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__1; +x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__1; x_2 = l_Lean_Meta_mkSimpAttr___lambda__1___closed__17; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__75() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__75() { _start: { lean_object* x_1; @@ -16260,51 +16259,51 @@ x_1 = lean_mk_string_from_bytes("simpPre", 7); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__76() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__76() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__74; -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__75; +x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__74; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__75; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__77() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__77() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Meta_mkSimpAttr___lambda__1___closed__18; -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__75; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__75; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__78() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__78() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__77; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__77; 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_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__79() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__79() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__78; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__78; x_3 = lean_alloc_ctor(1, 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_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__80() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__80() { _start: { lean_object* x_1; @@ -16312,7 +16311,7 @@ x_1 = lean_mk_string_from_bytes("<|>", 3); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__81() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__81() { _start: { lean_object* x_1; @@ -16320,22 +16319,22 @@ x_1 = lean_mk_string_from_bytes("Parser.Tactic.simpPost", 22); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__82() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__82() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__81; +x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__81; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__83() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__83() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__81; +x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__81; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__82; +x_3 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__82; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -16343,17 +16342,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__84() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__84() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__74; +x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__74; x_2 = l_Lean_Meta_mkSimpAttr___lambda__1___closed__19; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__85() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__85() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -16365,19 +16364,19 @@ lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__86() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__86() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__85; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__85; x_3 = lean_alloc_ctor(1, 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_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__87() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__87() { _start: { lean_object* x_1; @@ -16385,7 +16384,7 @@ x_1 = lean_mk_string_from_bytes("?", 1); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__88() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__88() { _start: { lean_object* x_1; @@ -16393,22 +16392,22 @@ x_1 = lean_mk_string_from_bytes("prio", 4); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__89() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__89() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__88; +x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__88; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__90() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__90() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__88; +x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__88; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__89; +x_3 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__89; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -16416,17 +16415,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__91() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__91() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__88; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__88; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__92() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__92() { _start: { lean_object* x_1; @@ -16434,22 +16433,22 @@ x_1 = lean_mk_string_from_bytes("attr", 4); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__93() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__93() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__92; +x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__92; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__94() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__94() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__92; +x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__92; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__93; +x_3 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__93; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -16457,17 +16456,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__95() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__95() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__92; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__92; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__96() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__96() { _start: { lean_object* x_1; lean_object* x_2; @@ -16476,27 +16475,27 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__97() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__97() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__96; -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__10; +x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__96; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__10; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__98() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__98() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__97; -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__51; +x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__97; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__51; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__99() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__99() { _start: { lean_object* x_1; @@ -16504,17 +16503,17 @@ x_1 = lean_mk_string_from_bytes("quotedName", 10); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__100() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__100() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__16; -x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__99; +x_1 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__16; +x_2 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__99; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__101() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__101() { _start: { lean_object* x_1; @@ -16522,7 +16521,7 @@ x_1 = lean_mk_string_from_bytes(".", 1); return x_1; } } -static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__102() { +static lean_object* _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__102() { _start: { lean_object* x_1; @@ -16530,11 +16529,11 @@ x_1 = lean_mk_string_from_bytes("`", 1); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; uint8_t x_5; -x_4 = l_Lean_Meta_commandRegister__simp__attr_________closed__4; +x_4 = l_Lean_Meta_commandRegister__simp__attr_______closed__4; lean_inc(x_1); x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); if (x_5 == 0) @@ -16550,7 +16549,7 @@ return x_7; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t 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; 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; x_8 = lean_unsigned_to_nat(1u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); x_10 = lean_unsigned_to_nat(2u); @@ -16560,830 +16559,386 @@ x_12 = l_Lean_Syntax_getId(x_9); x_13 = 1; lean_inc(x_12); x_14 = l_Lean_Name_toString(x_12, x_13); -x_15 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__3; +x_15 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__3; lean_inc(x_12); x_16 = l_Lean_Name_append(x_15, x_12); x_17 = l_Lean_mkIdentFrom(x_9, x_16); lean_inc(x_2); x_18 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_3); -x_19 = !lean_is_exclusive(x_18); -if (x_19 == 0) -{ -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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; -x_20 = lean_ctor_get(x_18, 0); -x_21 = lean_ctor_get(x_2, 2); -lean_inc(x_21); -x_22 = lean_ctor_get(x_2, 1); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +if (lean_is_exclusive(x_18)) { + lean_ctor_release(x_18, 0); + lean_ctor_release(x_18, 1); + x_21 = x_18; +} else { + lean_dec_ref(x_18); + x_21 = lean_box(0); +} +x_22 = lean_ctor_get(x_2, 2); lean_inc(x_22); +x_23 = lean_ctor_get(x_2, 1); +lean_inc(x_23); lean_dec(x_2); -x_23 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__8; -lean_inc(x_20); -x_24 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_24, 0, x_20); -lean_ctor_set(x_24, 1, x_23); -x_25 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__14; -lean_inc(x_21); +x_24 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__8; +lean_inc(x_19); +x_25 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_25, 0, x_19); +lean_ctor_set(x_25, 1, x_24); +x_26 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__14; lean_inc(x_22); -x_26 = l_Lean_addMacroScope(x_22, x_25, x_21); -x_27 = lean_box(0); -x_28 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__13; -lean_inc(x_20); -x_29 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_29, 0, x_20); -lean_ctor_set(x_29, 1, x_28); -lean_ctor_set(x_29, 2, x_26); -lean_ctor_set(x_29, 3, x_27); -x_30 = l_Lean_Meta_instToFormatSimpTheorem___closed__1; -lean_inc(x_20); -x_31 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_31, 0, x_20); -lean_ctor_set(x_31, 1, x_30); -x_32 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__22; -lean_inc(x_21); +lean_inc(x_23); +x_27 = l_Lean_addMacroScope(x_23, x_26, x_22); +x_28 = lean_box(0); +x_29 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__13; +lean_inc(x_19); +x_30 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_30, 0, x_19); +lean_ctor_set(x_30, 1, x_29); +lean_ctor_set(x_30, 2, x_27); +lean_ctor_set(x_30, 3, x_28); +x_31 = l_Lean_Meta_instToFormatSimpTheorem___closed__1; +lean_inc(x_19); +x_32 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_32, 0, x_19); +lean_ctor_set(x_32, 1, x_31); +x_33 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__22; lean_inc(x_22); -x_33 = l_Lean_addMacroScope(x_22, x_32, x_21); -x_34 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__21; -x_35 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__25; -lean_inc(x_20); -x_36 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_36, 0, x_20); -lean_ctor_set(x_36, 1, x_34); -lean_ctor_set(x_36, 2, x_33); -lean_ctor_set(x_36, 3, x_35); -x_37 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__26; -lean_inc(x_31); -x_38 = lean_array_push(x_37, x_31); -x_39 = lean_array_push(x_38, x_36); -x_40 = lean_box(2); -x_41 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__18; -x_42 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_42, 0, x_40); -lean_ctor_set(x_42, 1, x_41); -lean_ctor_set(x_42, 2, x_39); -x_43 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__27; -lean_inc(x_20); -x_44 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_44, 0, x_20); -lean_ctor_set(x_44, 1, x_43); -x_45 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__28; -x_46 = lean_array_push(x_45, x_29); -x_47 = lean_array_push(x_46, x_42); -x_48 = lean_array_push(x_47, x_44); -x_49 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__5; -x_50 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_50, 0, x_40); -lean_ctor_set(x_50, 1, x_49); -lean_ctor_set(x_50, 2, x_48); -x_51 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__40; -lean_inc(x_21); +lean_inc(x_23); +x_34 = l_Lean_addMacroScope(x_23, x_33, x_22); +x_35 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__21; +x_36 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__25; +lean_inc(x_19); +x_37 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_37, 0, x_19); +lean_ctor_set(x_37, 1, x_35); +lean_ctor_set(x_37, 2, x_34); +lean_ctor_set(x_37, 3, x_36); +x_38 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__26; +lean_inc(x_32); +x_39 = lean_array_push(x_38, x_32); +x_40 = lean_array_push(x_39, x_37); +x_41 = lean_box(2); +x_42 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__18; +x_43 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +lean_ctor_set(x_43, 2, x_40); +x_44 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__27; +lean_inc(x_19); +x_45 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_45, 0, x_19); +lean_ctor_set(x_45, 1, x_44); +x_46 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__28; +x_47 = lean_array_push(x_46, x_30); +x_48 = lean_array_push(x_47, x_43); +x_49 = lean_array_push(x_48, x_45); +x_50 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__5; +x_51 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_51, 0, x_41); +lean_ctor_set(x_51, 1, x_50); +lean_ctor_set(x_51, 2, x_49); +x_52 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__40; lean_inc(x_22); -x_52 = l_Lean_addMacroScope(x_22, x_51, x_21); -x_53 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__39; -x_54 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__43; -lean_inc(x_20); -x_55 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_55, 0, x_20); -lean_ctor_set(x_55, 1, x_53); -lean_ctor_set(x_55, 2, x_52); -lean_ctor_set(x_55, 3, x_54); +lean_inc(x_23); +x_53 = l_Lean_addMacroScope(x_23, x_52, x_22); +x_54 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__39; +x_55 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__43; +lean_inc(x_19); +x_56 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_56, 0, x_19); +lean_ctor_set(x_56, 1, x_54); +lean_ctor_set(x_56, 2, x_53); +lean_ctor_set(x_56, 3, x_55); lean_inc(x_12); -x_56 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_27, x_12); -x_57 = lean_array_push(x_37, x_55); -x_58 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__45; -x_59 = lean_array_push(x_58, x_24); -x_60 = lean_array_push(x_59, x_50); -x_61 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__46; -lean_inc(x_20); -x_62 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_62, 0, x_20); -lean_ctor_set(x_62, 1, x_61); -x_63 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__54; -lean_inc(x_20); -x_64 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_64, 0, x_20); -lean_ctor_set(x_64, 1, x_63); -x_65 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__55; -lean_inc(x_20); -x_66 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_66, 0, x_20); -lean_ctor_set(x_66, 1, x_65); -x_67 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__56; -lean_inc(x_20); -x_68 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_68, 0, x_20); -lean_ctor_set(x_68, 1, x_67); -x_69 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__57; -lean_inc(x_20); -x_70 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_70, 0, x_20); -lean_ctor_set(x_70, 1, x_69); -x_71 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__58; -lean_inc(x_64); -x_72 = lean_array_push(x_71, x_64); -x_73 = lean_array_push(x_72, x_66); -x_74 = lean_array_push(x_73, x_68); -x_75 = lean_array_push(x_74, x_17); -lean_inc(x_70); -x_76 = lean_array_push(x_75, x_70); -x_77 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__53; -x_78 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_78, 0, x_40); -lean_ctor_set(x_78, 1, x_77); -lean_ctor_set(x_78, 2, x_76); -x_79 = l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__3___closed__9; -x_80 = lean_array_push(x_79, x_78); -x_81 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_81, 0, x_40); -lean_ctor_set(x_81, 1, x_49); -lean_ctor_set(x_81, 2, x_80); -x_82 = l_Lean_Syntax_mkStrLit(x_14, x_40); +x_57 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_28, x_12); +x_58 = lean_array_push(x_38, x_56); +x_59 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__45; +x_60 = lean_array_push(x_59, x_25); +x_61 = lean_array_push(x_60, x_51); +x_62 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__46; +lean_inc(x_19); +x_63 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_63, 0, x_19); +lean_ctor_set(x_63, 1, x_62); +x_64 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__54; +lean_inc(x_19); +x_65 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_65, 0, x_19); +lean_ctor_set(x_65, 1, x_64); +x_66 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__55; +lean_inc(x_19); +x_67 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_67, 0, x_19); +lean_ctor_set(x_67, 1, x_66); +x_68 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__56; +lean_inc(x_19); +x_69 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_69, 0, x_19); +lean_ctor_set(x_69, 1, x_68); +x_70 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__57; +lean_inc(x_19); +x_71 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_71, 0, x_19); +lean_ctor_set(x_71, 1, x_70); +x_72 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__58; +lean_inc(x_65); +x_73 = lean_array_push(x_72, x_65); +x_74 = lean_array_push(x_73, x_67); +x_75 = lean_array_push(x_74, x_69); +x_76 = lean_array_push(x_75, x_17); +lean_inc(x_71); +x_77 = lean_array_push(x_76, x_71); +x_78 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__53; +x_79 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_79, 0, x_41); +lean_ctor_set(x_79, 1, x_78); +lean_ctor_set(x_79, 2, x_77); +x_80 = l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__3___closed__9; +x_81 = lean_array_push(x_80, x_79); +x_82 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_82, 0, x_41); +lean_ctor_set(x_82, 1, x_50); +lean_ctor_set(x_82, 2, x_81); +x_83 = l_Lean_Syntax_mkStrLit(x_14, x_41); lean_dec(x_14); -x_83 = lean_array_push(x_79, x_82); -x_84 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__62; -x_85 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_85, 0, x_40); -lean_ctor_set(x_85, 1, x_84); -lean_ctor_set(x_85, 2, x_83); -x_86 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__76; -lean_inc(x_21); +x_84 = lean_array_push(x_80, x_83); +x_85 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__62; +x_86 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_86, 0, x_41); +lean_ctor_set(x_86, 1, x_85); +lean_ctor_set(x_86, 2, x_84); +x_87 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__76; lean_inc(x_22); -x_87 = l_Lean_addMacroScope(x_22, x_86, x_21); -x_88 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__73; -x_89 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__79; -lean_inc(x_20); -x_90 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_90, 0, x_20); -lean_ctor_set(x_90, 1, x_88); -lean_ctor_set(x_90, 2, x_87); -lean_ctor_set(x_90, 3, x_89); -x_91 = lean_array_push(x_37, x_90); -x_92 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__10; -x_93 = lean_array_push(x_91, x_92); -x_94 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__70; -x_95 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_95, 0, x_40); -lean_ctor_set(x_95, 1, x_94); -lean_ctor_set(x_95, 2, x_93); -x_96 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__80; -lean_inc(x_20); -x_97 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_97, 0, x_20); -lean_ctor_set(x_97, 1, x_96); -x_98 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__84; -lean_inc(x_21); +lean_inc(x_23); +x_88 = l_Lean_addMacroScope(x_23, x_87, x_22); +x_89 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__73; +x_90 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__79; +lean_inc(x_19); +x_91 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_91, 0, x_19); +lean_ctor_set(x_91, 1, x_89); +lean_ctor_set(x_91, 2, x_88); +lean_ctor_set(x_91, 3, x_90); +x_92 = lean_array_push(x_38, x_91); +x_93 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__10; +x_94 = lean_array_push(x_92, x_93); +x_95 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__70; +x_96 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_96, 0, x_41); +lean_ctor_set(x_96, 1, x_95); +lean_ctor_set(x_96, 2, x_94); +x_97 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__80; +lean_inc(x_19); +x_98 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_98, 0, x_19); +lean_ctor_set(x_98, 1, x_97); +x_99 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__84; lean_inc(x_22); -x_99 = l_Lean_addMacroScope(x_22, x_98, x_21); -x_100 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__83; -x_101 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__86; -lean_inc(x_20); -x_102 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_102, 0, x_20); -lean_ctor_set(x_102, 1, x_100); -lean_ctor_set(x_102, 2, x_99); -lean_ctor_set(x_102, 3, x_101); -x_103 = lean_array_push(x_37, x_102); -x_104 = lean_array_push(x_103, x_92); -x_105 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_105, 0, x_40); -lean_ctor_set(x_105, 1, x_94); -lean_ctor_set(x_105, 2, x_104); -x_106 = lean_array_push(x_45, x_95); -x_107 = lean_array_push(x_106, x_97); -x_108 = lean_array_push(x_107, x_105); -x_109 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__68; -x_110 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_110, 0, x_40); -lean_ctor_set(x_110, 1, x_109); -lean_ctor_set(x_110, 2, x_108); -x_111 = lean_array_push(x_79, x_110); -x_112 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_112, 0, x_40); -lean_ctor_set(x_112, 1, x_49); -lean_ctor_set(x_112, 2, x_111); -x_113 = lean_array_push(x_45, x_64); -lean_inc(x_113); -x_114 = lean_array_push(x_113, x_112); -lean_inc(x_70); -x_115 = lean_array_push(x_114, x_70); -x_116 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__66; -x_117 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_117, 0, x_40); -lean_ctor_set(x_117, 1, x_116); -lean_ctor_set(x_117, 2, x_115); -x_118 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__87; -lean_inc(x_20); -x_119 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_119, 0, x_20); -lean_ctor_set(x_119, 1, x_118); -x_120 = lean_array_push(x_37, x_117); -lean_inc(x_119); -x_121 = lean_array_push(x_120, x_119); -x_122 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__64; -x_123 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_123, 0, x_40); -lean_ctor_set(x_123, 1, x_122); -lean_ctor_set(x_123, 2, x_121); -x_124 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__91; -lean_inc(x_21); +lean_inc(x_23); +x_100 = l_Lean_addMacroScope(x_23, x_99, x_22); +x_101 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__83; +x_102 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__86; +lean_inc(x_19); +x_103 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_103, 0, x_19); +lean_ctor_set(x_103, 1, x_101); +lean_ctor_set(x_103, 2, x_100); +lean_ctor_set(x_103, 3, x_102); +x_104 = lean_array_push(x_38, x_103); +x_105 = lean_array_push(x_104, x_93); +x_106 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_106, 0, x_41); +lean_ctor_set(x_106, 1, x_95); +lean_ctor_set(x_106, 2, x_105); +x_107 = lean_array_push(x_46, x_96); +x_108 = lean_array_push(x_107, x_98); +x_109 = lean_array_push(x_108, x_106); +x_110 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__68; +x_111 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_111, 0, x_41); +lean_ctor_set(x_111, 1, x_110); +lean_ctor_set(x_111, 2, x_109); +x_112 = lean_array_push(x_80, x_111); +x_113 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_113, 0, x_41); +lean_ctor_set(x_113, 1, x_50); +lean_ctor_set(x_113, 2, x_112); +x_114 = lean_array_push(x_46, x_65); +lean_inc(x_114); +x_115 = lean_array_push(x_114, x_113); +lean_inc(x_71); +x_116 = lean_array_push(x_115, x_71); +x_117 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__66; +x_118 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_118, 0, x_41); +lean_ctor_set(x_118, 1, x_117); +lean_ctor_set(x_118, 2, x_116); +x_119 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__87; +lean_inc(x_19); +x_120 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_120, 0, x_19); +lean_ctor_set(x_120, 1, x_119); +x_121 = lean_array_push(x_38, x_118); +lean_inc(x_120); +x_122 = lean_array_push(x_121, x_120); +x_123 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__64; +x_124 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_124, 0, x_41); +lean_ctor_set(x_124, 1, x_123); +lean_ctor_set(x_124, 2, x_122); +x_125 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__91; lean_inc(x_22); -x_125 = l_Lean_addMacroScope(x_22, x_124, x_21); -x_126 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__90; -lean_inc(x_20); -x_127 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_127, 0, x_20); -lean_ctor_set(x_127, 1, x_126); -lean_ctor_set(x_127, 2, x_125); -lean_ctor_set(x_127, 3, x_27); -x_128 = lean_array_push(x_37, x_127); -x_129 = lean_array_push(x_128, x_92); -x_130 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_130, 0, x_40); -lean_ctor_set(x_130, 1, x_94); -lean_ctor_set(x_130, 2, x_129); -x_131 = lean_array_push(x_79, x_130); -x_132 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_132, 0, x_40); -lean_ctor_set(x_132, 1, x_49); -lean_ctor_set(x_132, 2, x_131); -x_133 = lean_array_push(x_113, x_132); -x_134 = lean_array_push(x_133, x_70); -x_135 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_135, 0, x_40); -lean_ctor_set(x_135, 1, x_116); -lean_ctor_set(x_135, 2, x_134); -x_136 = lean_array_push(x_37, x_135); -x_137 = lean_array_push(x_136, x_119); -x_138 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_138, 0, x_40); -lean_ctor_set(x_138, 1, x_122); -lean_ctor_set(x_138, 2, x_137); -x_139 = lean_array_push(x_45, x_85); -x_140 = lean_array_push(x_139, x_123); -x_141 = lean_array_push(x_140, x_138); -x_142 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_142, 0, x_40); -lean_ctor_set(x_142, 1, x_49); -lean_ctor_set(x_142, 2, x_141); -x_143 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__95; -x_144 = l_Lean_addMacroScope(x_22, x_143, x_21); -x_145 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__94; -x_146 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_146, 0, x_20); -lean_ctor_set(x_146, 1, x_145); -lean_ctor_set(x_146, 2, x_144); -lean_ctor_set(x_146, 3, x_27); -x_147 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__98; -x_148 = lean_array_push(x_147, x_62); -x_149 = lean_array_push(x_148, x_92); -x_150 = lean_array_push(x_149, x_81); -x_151 = lean_array_push(x_150, x_92); -x_152 = lean_array_push(x_151, x_142); -x_153 = lean_array_push(x_152, x_31); -x_154 = lean_array_push(x_153, x_146); -x_155 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__47; -x_156 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_156, 0, x_40); -lean_ctor_set(x_156, 1, x_155); -lean_ctor_set(x_156, 2, x_154); -if (lean_obj_tag(x_56) == 0) +lean_inc(x_23); +x_126 = l_Lean_addMacroScope(x_23, x_125, x_22); +x_127 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__90; +lean_inc(x_19); +x_128 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_128, 0, x_19); +lean_ctor_set(x_128, 1, x_127); +lean_ctor_set(x_128, 2, x_126); +lean_ctor_set(x_128, 3, x_28); +x_129 = lean_array_push(x_38, x_128); +x_130 = lean_array_push(x_129, x_93); +x_131 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_131, 0, x_41); +lean_ctor_set(x_131, 1, x_95); +lean_ctor_set(x_131, 2, x_130); +x_132 = lean_array_push(x_80, x_131); +x_133 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_133, 0, x_41); +lean_ctor_set(x_133, 1, x_50); +lean_ctor_set(x_133, 2, x_132); +x_134 = lean_array_push(x_114, x_133); +x_135 = lean_array_push(x_134, x_71); +x_136 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_136, 0, x_41); +lean_ctor_set(x_136, 1, x_117); +lean_ctor_set(x_136, 2, x_135); +x_137 = lean_array_push(x_38, x_136); +x_138 = lean_array_push(x_137, x_120); +x_139 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_139, 0, x_41); +lean_ctor_set(x_139, 1, x_123); +lean_ctor_set(x_139, 2, x_138); +x_140 = lean_array_push(x_46, x_86); +x_141 = lean_array_push(x_140, x_124); +x_142 = lean_array_push(x_141, x_139); +x_143 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_143, 0, x_41); +lean_ctor_set(x_143, 1, x_50); +lean_ctor_set(x_143, 2, x_142); +x_144 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__95; +x_145 = l_Lean_addMacroScope(x_23, x_144, x_22); +x_146 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__94; +x_147 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_147, 0, x_19); +lean_ctor_set(x_147, 1, x_146); +lean_ctor_set(x_147, 2, x_145); +lean_ctor_set(x_147, 3, x_28); +x_148 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__98; +x_149 = lean_array_push(x_148, x_63); +x_150 = lean_array_push(x_149, x_93); +x_151 = lean_array_push(x_150, x_82); +x_152 = lean_array_push(x_151, x_93); +x_153 = lean_array_push(x_152, x_143); +x_154 = lean_array_push(x_153, x_32); +x_155 = lean_array_push(x_154, x_147); +x_156 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__47; +x_157 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_157, 0, x_41); +lean_ctor_set(x_157, 1, x_156); +lean_ctor_set(x_157, 2, x_155); +if (lean_obj_tag(x_57) == 0) { -lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; -x_157 = l___private_Init_Meta_0__Lean_quoteNameMk(x_12); -x_158 = lean_array_push(x_37, x_157); -x_159 = lean_array_push(x_158, x_11); -x_160 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_160, 0, x_40); -lean_ctor_set(x_160, 1, x_49); -lean_ctor_set(x_160, 2, x_159); -x_161 = lean_array_push(x_57, x_160); -x_162 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__36; -x_163 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_163, 0, x_40); -lean_ctor_set(x_163, 1, x_162); -lean_ctor_set(x_163, 2, x_161); -x_164 = lean_array_push(x_79, x_163); -x_165 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__34; -x_166 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_166, 0, x_40); -lean_ctor_set(x_166, 1, x_165); -lean_ctor_set(x_166, 2, x_164); -x_167 = lean_array_push(x_37, x_166); -x_168 = lean_array_push(x_167, x_92); -x_169 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__32; -x_170 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_170, 0, x_40); -lean_ctor_set(x_170, 1, x_169); -lean_ctor_set(x_170, 2, x_168); -x_171 = lean_array_push(x_79, x_170); -x_172 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_172, 0, x_40); -lean_ctor_set(x_172, 1, x_49); -lean_ctor_set(x_172, 2, x_171); -x_173 = lean_array_push(x_79, x_172); -x_174 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__30; -x_175 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_175, 0, x_40); -lean_ctor_set(x_175, 1, x_174); -lean_ctor_set(x_175, 2, x_173); -x_176 = lean_array_push(x_60, x_175); -x_177 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__9; -x_178 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_178, 0, x_40); -lean_ctor_set(x_178, 1, x_177); -lean_ctor_set(x_178, 2, x_176); -x_179 = lean_array_push(x_37, x_178); -x_180 = lean_array_push(x_179, x_156); -x_181 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_181, 0, x_40); -lean_ctor_set(x_181, 1, x_49); -lean_ctor_set(x_181, 2, x_180); -lean_ctor_set(x_18, 0, x_181); -return x_18; +lean_object* x_185; +x_185 = l_Lean_quoteNameMk(x_12); +x_158 = x_185; +goto block_184; } else { -lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; +lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_dec(x_12); -x_182 = lean_ctor_get(x_56, 0); -lean_inc(x_182); -lean_dec(x_56); -x_183 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__101; -x_184 = l_String_intercalate(x_183, x_182); -x_185 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__102; -x_186 = lean_string_append(x_185, x_184); -lean_dec(x_184); -x_187 = l_Lean_nameLitKind; -x_188 = l_Lean_Syntax_mkLit(x_187, x_186, x_40); -x_189 = lean_array_push(x_79, x_188); -x_190 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__100; -x_191 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_191, 0, x_40); -lean_ctor_set(x_191, 1, x_190); -lean_ctor_set(x_191, 2, x_189); -x_192 = lean_array_push(x_37, x_191); -x_193 = lean_array_push(x_192, x_11); +x_186 = lean_ctor_get(x_57, 0); +lean_inc(x_186); +lean_dec(x_57); +x_187 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__101; +x_188 = l_String_intercalate(x_187, x_186); +x_189 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__102; +x_190 = lean_string_append(x_189, x_188); +lean_dec(x_188); +x_191 = l_Lean_Syntax_mkNameLit(x_190, x_41); +x_192 = lean_array_push(x_80, x_191); +x_193 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__100; x_194 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_194, 0, x_40); -lean_ctor_set(x_194, 1, x_49); -lean_ctor_set(x_194, 2, x_193); -x_195 = lean_array_push(x_57, x_194); -x_196 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__36; -x_197 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_197, 0, x_40); -lean_ctor_set(x_197, 1, x_196); -lean_ctor_set(x_197, 2, x_195); -x_198 = lean_array_push(x_79, x_197); -x_199 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__34; -x_200 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_200, 0, x_40); -lean_ctor_set(x_200, 1, x_199); -lean_ctor_set(x_200, 2, x_198); -x_201 = lean_array_push(x_37, x_200); -x_202 = lean_array_push(x_201, x_92); -x_203 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__32; -x_204 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_204, 0, x_40); -lean_ctor_set(x_204, 1, x_203); -lean_ctor_set(x_204, 2, x_202); -x_205 = lean_array_push(x_79, x_204); -x_206 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_206, 0, x_40); -lean_ctor_set(x_206, 1, x_49); -lean_ctor_set(x_206, 2, x_205); -x_207 = lean_array_push(x_79, x_206); -x_208 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__30; -x_209 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_209, 0, x_40); -lean_ctor_set(x_209, 1, x_208); -lean_ctor_set(x_209, 2, x_207); -x_210 = lean_array_push(x_60, x_209); -x_211 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__9; -x_212 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_212, 0, x_40); -lean_ctor_set(x_212, 1, x_211); -lean_ctor_set(x_212, 2, x_210); -x_213 = lean_array_push(x_37, x_212); -x_214 = lean_array_push(x_213, x_156); -x_215 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_215, 0, x_40); -lean_ctor_set(x_215, 1, x_49); -lean_ctor_set(x_215, 2, x_214); -lean_ctor_set(x_18, 0, x_215); -return x_18; -} -} -else -{ -lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; -x_216 = lean_ctor_get(x_18, 0); -x_217 = lean_ctor_get(x_18, 1); -lean_inc(x_217); -lean_inc(x_216); -lean_dec(x_18); -x_218 = lean_ctor_get(x_2, 2); -lean_inc(x_218); -x_219 = lean_ctor_get(x_2, 1); -lean_inc(x_219); -lean_dec(x_2); -x_220 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__8; -lean_inc(x_216); -x_221 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_221, 0, x_216); -lean_ctor_set(x_221, 1, x_220); -x_222 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__14; -lean_inc(x_218); -lean_inc(x_219); -x_223 = l_Lean_addMacroScope(x_219, x_222, x_218); -x_224 = lean_box(0); -x_225 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__13; -lean_inc(x_216); -x_226 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_226, 0, x_216); -lean_ctor_set(x_226, 1, x_225); -lean_ctor_set(x_226, 2, x_223); -lean_ctor_set(x_226, 3, x_224); -x_227 = l_Lean_Meta_instToFormatSimpTheorem___closed__1; -lean_inc(x_216); -x_228 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_228, 0, x_216); -lean_ctor_set(x_228, 1, x_227); -x_229 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__22; -lean_inc(x_218); -lean_inc(x_219); -x_230 = l_Lean_addMacroScope(x_219, x_229, x_218); -x_231 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__21; -x_232 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__25; -lean_inc(x_216); -x_233 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_233, 0, x_216); -lean_ctor_set(x_233, 1, x_231); -lean_ctor_set(x_233, 2, x_230); -lean_ctor_set(x_233, 3, x_232); -x_234 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__26; -lean_inc(x_228); -x_235 = lean_array_push(x_234, x_228); -x_236 = lean_array_push(x_235, x_233); -x_237 = lean_box(2); -x_238 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__18; -x_239 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_239, 0, x_237); -lean_ctor_set(x_239, 1, x_238); -lean_ctor_set(x_239, 2, x_236); -x_240 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__27; -lean_inc(x_216); -x_241 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_241, 0, x_216); -lean_ctor_set(x_241, 1, x_240); -x_242 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__28; -x_243 = lean_array_push(x_242, x_226); -x_244 = lean_array_push(x_243, x_239); -x_245 = lean_array_push(x_244, x_241); -x_246 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__5; -x_247 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_247, 0, x_237); -lean_ctor_set(x_247, 1, x_246); -lean_ctor_set(x_247, 2, x_245); -x_248 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__40; -lean_inc(x_218); -lean_inc(x_219); -x_249 = l_Lean_addMacroScope(x_219, x_248, x_218); -x_250 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__39; -x_251 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__43; -lean_inc(x_216); -x_252 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_252, 0, x_216); -lean_ctor_set(x_252, 1, x_250); -lean_ctor_set(x_252, 2, x_249); -lean_ctor_set(x_252, 3, x_251); -lean_inc(x_12); -x_253 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_224, x_12); -x_254 = lean_array_push(x_234, x_252); -x_255 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__45; -x_256 = lean_array_push(x_255, x_221); -x_257 = lean_array_push(x_256, x_247); -x_258 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__46; -lean_inc(x_216); -x_259 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_259, 0, x_216); -lean_ctor_set(x_259, 1, x_258); -x_260 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__54; -lean_inc(x_216); -x_261 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_261, 0, x_216); -lean_ctor_set(x_261, 1, x_260); -x_262 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__55; -lean_inc(x_216); -x_263 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_263, 0, x_216); -lean_ctor_set(x_263, 1, x_262); -x_264 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__56; -lean_inc(x_216); -x_265 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_265, 0, x_216); -lean_ctor_set(x_265, 1, x_264); -x_266 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__57; -lean_inc(x_216); -x_267 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_267, 0, x_216); -lean_ctor_set(x_267, 1, x_266); -x_268 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__58; -lean_inc(x_261); -x_269 = lean_array_push(x_268, x_261); -x_270 = lean_array_push(x_269, x_263); -x_271 = lean_array_push(x_270, x_265); -x_272 = lean_array_push(x_271, x_17); -lean_inc(x_267); -x_273 = lean_array_push(x_272, x_267); -x_274 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__53; -x_275 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_275, 0, x_237); -lean_ctor_set(x_275, 1, x_274); -lean_ctor_set(x_275, 2, x_273); -x_276 = l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__3___closed__9; -x_277 = lean_array_push(x_276, x_275); -x_278 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_278, 0, x_237); -lean_ctor_set(x_278, 1, x_246); -lean_ctor_set(x_278, 2, x_277); -x_279 = l_Lean_Syntax_mkStrLit(x_14, x_237); -lean_dec(x_14); -x_280 = lean_array_push(x_276, x_279); -x_281 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__62; -x_282 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_282, 0, x_237); -lean_ctor_set(x_282, 1, x_281); -lean_ctor_set(x_282, 2, x_280); -x_283 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__76; -lean_inc(x_218); -lean_inc(x_219); -x_284 = l_Lean_addMacroScope(x_219, x_283, x_218); -x_285 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__73; -x_286 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__79; -lean_inc(x_216); -x_287 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_287, 0, x_216); -lean_ctor_set(x_287, 1, x_285); -lean_ctor_set(x_287, 2, x_284); -lean_ctor_set(x_287, 3, x_286); -x_288 = lean_array_push(x_234, x_287); -x_289 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__10; -x_290 = lean_array_push(x_288, x_289); -x_291 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__70; -x_292 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_292, 0, x_237); -lean_ctor_set(x_292, 1, x_291); -lean_ctor_set(x_292, 2, x_290); -x_293 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__80; -lean_inc(x_216); -x_294 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_294, 0, x_216); -lean_ctor_set(x_294, 1, x_293); -x_295 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__84; -lean_inc(x_218); -lean_inc(x_219); -x_296 = l_Lean_addMacroScope(x_219, x_295, x_218); -x_297 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__83; -x_298 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__86; -lean_inc(x_216); -x_299 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_299, 0, x_216); -lean_ctor_set(x_299, 1, x_297); -lean_ctor_set(x_299, 2, x_296); -lean_ctor_set(x_299, 3, x_298); -x_300 = lean_array_push(x_234, x_299); -x_301 = lean_array_push(x_300, x_289); -x_302 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_302, 0, x_237); -lean_ctor_set(x_302, 1, x_291); -lean_ctor_set(x_302, 2, x_301); -x_303 = lean_array_push(x_242, x_292); -x_304 = lean_array_push(x_303, x_294); -x_305 = lean_array_push(x_304, x_302); -x_306 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__68; -x_307 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_307, 0, x_237); -lean_ctor_set(x_307, 1, x_306); -lean_ctor_set(x_307, 2, x_305); -x_308 = lean_array_push(x_276, x_307); -x_309 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_309, 0, x_237); -lean_ctor_set(x_309, 1, x_246); -lean_ctor_set(x_309, 2, x_308); -x_310 = lean_array_push(x_242, x_261); -lean_inc(x_310); -x_311 = lean_array_push(x_310, x_309); -lean_inc(x_267); -x_312 = lean_array_push(x_311, x_267); -x_313 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__66; -x_314 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_314, 0, x_237); -lean_ctor_set(x_314, 1, x_313); -lean_ctor_set(x_314, 2, x_312); -x_315 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__87; -lean_inc(x_216); -x_316 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_316, 0, x_216); -lean_ctor_set(x_316, 1, x_315); -x_317 = lean_array_push(x_234, x_314); -lean_inc(x_316); -x_318 = lean_array_push(x_317, x_316); -x_319 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__64; -x_320 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_320, 0, x_237); -lean_ctor_set(x_320, 1, x_319); -lean_ctor_set(x_320, 2, x_318); -x_321 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__91; -lean_inc(x_218); -lean_inc(x_219); -x_322 = l_Lean_addMacroScope(x_219, x_321, x_218); -x_323 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__90; -lean_inc(x_216); -x_324 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_324, 0, x_216); -lean_ctor_set(x_324, 1, x_323); -lean_ctor_set(x_324, 2, x_322); -lean_ctor_set(x_324, 3, x_224); -x_325 = lean_array_push(x_234, x_324); -x_326 = lean_array_push(x_325, x_289); -x_327 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_327, 0, x_237); -lean_ctor_set(x_327, 1, x_291); -lean_ctor_set(x_327, 2, x_326); -x_328 = lean_array_push(x_276, x_327); -x_329 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_329, 0, x_237); -lean_ctor_set(x_329, 1, x_246); -lean_ctor_set(x_329, 2, x_328); -x_330 = lean_array_push(x_310, x_329); -x_331 = lean_array_push(x_330, x_267); -x_332 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_332, 0, x_237); -lean_ctor_set(x_332, 1, x_313); -lean_ctor_set(x_332, 2, x_331); -x_333 = lean_array_push(x_234, x_332); -x_334 = lean_array_push(x_333, x_316); -x_335 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_335, 0, x_237); -lean_ctor_set(x_335, 1, x_319); -lean_ctor_set(x_335, 2, x_334); -x_336 = lean_array_push(x_242, x_282); -x_337 = lean_array_push(x_336, x_320); -x_338 = lean_array_push(x_337, x_335); -x_339 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_339, 0, x_237); -lean_ctor_set(x_339, 1, x_246); -lean_ctor_set(x_339, 2, x_338); -x_340 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__95; -x_341 = l_Lean_addMacroScope(x_219, x_340, x_218); -x_342 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__94; -x_343 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_343, 0, x_216); -lean_ctor_set(x_343, 1, x_342); -lean_ctor_set(x_343, 2, x_341); -lean_ctor_set(x_343, 3, x_224); -x_344 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__98; -x_345 = lean_array_push(x_344, x_259); -x_346 = lean_array_push(x_345, x_289); -x_347 = lean_array_push(x_346, x_278); -x_348 = lean_array_push(x_347, x_289); -x_349 = lean_array_push(x_348, x_339); -x_350 = lean_array_push(x_349, x_228); -x_351 = lean_array_push(x_350, x_343); -x_352 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__47; -x_353 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_353, 0, x_237); -lean_ctor_set(x_353, 1, x_352); -lean_ctor_set(x_353, 2, x_351); -if (lean_obj_tag(x_253) == 0) -{ -lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; -x_354 = l___private_Init_Meta_0__Lean_quoteNameMk(x_12); -x_355 = lean_array_push(x_234, x_354); -x_356 = lean_array_push(x_355, x_11); -x_357 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_357, 0, x_237); -lean_ctor_set(x_357, 1, x_246); -lean_ctor_set(x_357, 2, x_356); -x_358 = lean_array_push(x_254, x_357); -x_359 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__36; -x_360 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_360, 0, x_237); -lean_ctor_set(x_360, 1, x_359); -lean_ctor_set(x_360, 2, x_358); -x_361 = lean_array_push(x_276, x_360); -x_362 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__34; -x_363 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_363, 0, x_237); -lean_ctor_set(x_363, 1, x_362); -lean_ctor_set(x_363, 2, x_361); -x_364 = lean_array_push(x_234, x_363); -x_365 = lean_array_push(x_364, x_289); -x_366 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__32; -x_367 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_367, 0, x_237); -lean_ctor_set(x_367, 1, x_366); -lean_ctor_set(x_367, 2, x_365); -x_368 = lean_array_push(x_276, x_367); -x_369 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_369, 0, x_237); -lean_ctor_set(x_369, 1, x_246); -lean_ctor_set(x_369, 2, x_368); -x_370 = lean_array_push(x_276, x_369); -x_371 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__30; -x_372 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_372, 0, x_237); -lean_ctor_set(x_372, 1, x_371); -lean_ctor_set(x_372, 2, x_370); -x_373 = lean_array_push(x_257, x_372); -x_374 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__9; -x_375 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_375, 0, x_237); -lean_ctor_set(x_375, 1, x_374); -lean_ctor_set(x_375, 2, x_373); -x_376 = lean_array_push(x_234, x_375); -x_377 = lean_array_push(x_376, x_353); -x_378 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_378, 0, x_237); -lean_ctor_set(x_378, 1, x_246); -lean_ctor_set(x_378, 2, x_377); -x_379 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_379, 0, x_378); -lean_ctor_set(x_379, 1, x_217); -return x_379; -} -else -{ -lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; -lean_dec(x_12); -x_380 = lean_ctor_get(x_253, 0); -lean_inc(x_380); -lean_dec(x_253); -x_381 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__101; -x_382 = l_String_intercalate(x_381, x_380); -x_383 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__102; -x_384 = lean_string_append(x_383, x_382); -lean_dec(x_382); -x_385 = l_Lean_nameLitKind; -x_386 = l_Lean_Syntax_mkLit(x_385, x_384, x_237); -x_387 = lean_array_push(x_276, x_386); -x_388 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__100; -x_389 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_389, 0, x_237); -lean_ctor_set(x_389, 1, x_388); -lean_ctor_set(x_389, 2, x_387); -x_390 = lean_array_push(x_234, x_389); -x_391 = lean_array_push(x_390, x_11); -x_392 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_392, 0, x_237); -lean_ctor_set(x_392, 1, x_246); -lean_ctor_set(x_392, 2, x_391); -x_393 = lean_array_push(x_254, x_392); -x_394 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__36; -x_395 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_395, 0, x_237); -lean_ctor_set(x_395, 1, x_394); -lean_ctor_set(x_395, 2, x_393); -x_396 = lean_array_push(x_276, x_395); -x_397 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__34; -x_398 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_398, 0, x_237); -lean_ctor_set(x_398, 1, x_397); -lean_ctor_set(x_398, 2, x_396); -x_399 = lean_array_push(x_234, x_398); -x_400 = lean_array_push(x_399, x_289); -x_401 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__32; -x_402 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_402, 0, x_237); -lean_ctor_set(x_402, 1, x_401); -lean_ctor_set(x_402, 2, x_400); -x_403 = lean_array_push(x_276, x_402); -x_404 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_404, 0, x_237); -lean_ctor_set(x_404, 1, x_246); -lean_ctor_set(x_404, 2, x_403); -x_405 = lean_array_push(x_276, x_404); -x_406 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__30; -x_407 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_407, 0, x_237); -lean_ctor_set(x_407, 1, x_406); -lean_ctor_set(x_407, 2, x_405); -x_408 = lean_array_push(x_257, x_407); -x_409 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__9; -x_410 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_410, 0, x_237); -lean_ctor_set(x_410, 1, x_409); -lean_ctor_set(x_410, 2, x_408); -x_411 = lean_array_push(x_234, x_410); -x_412 = lean_array_push(x_411, x_353); -x_413 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_413, 0, x_237); -lean_ctor_set(x_413, 1, x_246); -lean_ctor_set(x_413, 2, x_412); -x_414 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_414, 0, x_413); -lean_ctor_set(x_414, 1, x_217); -return x_414; +lean_ctor_set(x_194, 0, x_41); +lean_ctor_set(x_194, 1, x_193); +lean_ctor_set(x_194, 2, x_192); +x_158 = x_194; +goto block_184; +} +block_184: +{ +lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; +x_159 = lean_array_push(x_38, x_158); +x_160 = lean_array_push(x_159, x_11); +x_161 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_161, 0, x_41); +lean_ctor_set(x_161, 1, x_50); +lean_ctor_set(x_161, 2, x_160); +x_162 = lean_array_push(x_58, x_161); +x_163 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__36; +x_164 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_164, 0, x_41); +lean_ctor_set(x_164, 1, x_163); +lean_ctor_set(x_164, 2, x_162); +x_165 = lean_array_push(x_80, x_164); +x_166 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__34; +x_167 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_167, 0, x_41); +lean_ctor_set(x_167, 1, x_166); +lean_ctor_set(x_167, 2, x_165); +x_168 = lean_array_push(x_38, x_167); +x_169 = lean_array_push(x_168, x_93); +x_170 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__32; +x_171 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_171, 0, x_41); +lean_ctor_set(x_171, 1, x_170); +lean_ctor_set(x_171, 2, x_169); +x_172 = lean_array_push(x_80, x_171); +x_173 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_173, 0, x_41); +lean_ctor_set(x_173, 1, x_50); +lean_ctor_set(x_173, 2, x_172); +x_174 = lean_array_push(x_80, x_173); +x_175 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__30; +x_176 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_176, 0, x_41); +lean_ctor_set(x_176, 1, x_175); +lean_ctor_set(x_176, 2, x_174); +x_177 = lean_array_push(x_61, x_176); +x_178 = l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__9; +x_179 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_179, 0, x_41); +lean_ctor_set(x_179, 1, x_178); +lean_ctor_set(x_179, 2, x_177); +x_180 = lean_array_push(x_38, x_179); +x_181 = lean_array_push(x_180, x_157); +x_182 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_182, 0, x_41); +lean_ctor_set(x_182, 1, x_50); +lean_ctor_set(x_182, 2, x_181); +if (lean_is_scalar(x_21)) { + x_183 = lean_alloc_ctor(0, 2, 0); +} else { + x_183 = x_21; } +lean_ctor_set(x_183, 0, x_182); +lean_ctor_set(x_183, 1, x_20); +return x_183; } } } @@ -17710,246 +17265,246 @@ l_Lean_Meta_SimpTheorem_getValue___closed__4 = _init_l_Lean_Meta_SimpTheorem_get lean_mark_persistent(l_Lean_Meta_SimpTheorem_getValue___closed__4); l_Lean_Meta_SimpTheorems_addDeclToUnfold___closed__1 = _init_l_Lean_Meta_SimpTheorems_addDeclToUnfold___closed__1(); lean_mark_persistent(l_Lean_Meta_SimpTheorems_addDeclToUnfold___closed__1); -l_Lean_Meta_commandRegister__simp__attr_________closed__1 = _init_l_Lean_Meta_commandRegister__simp__attr_________closed__1(); -lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr_________closed__1); -l_Lean_Meta_commandRegister__simp__attr_________closed__2 = _init_l_Lean_Meta_commandRegister__simp__attr_________closed__2(); -lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr_________closed__2); -l_Lean_Meta_commandRegister__simp__attr_________closed__3 = _init_l_Lean_Meta_commandRegister__simp__attr_________closed__3(); -lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr_________closed__3); -l_Lean_Meta_commandRegister__simp__attr_________closed__4 = _init_l_Lean_Meta_commandRegister__simp__attr_________closed__4(); -lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr_________closed__4); -l_Lean_Meta_commandRegister__simp__attr_________closed__5 = _init_l_Lean_Meta_commandRegister__simp__attr_________closed__5(); -lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr_________closed__5); -l_Lean_Meta_commandRegister__simp__attr_________closed__6 = _init_l_Lean_Meta_commandRegister__simp__attr_________closed__6(); -lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr_________closed__6); -l_Lean_Meta_commandRegister__simp__attr_________closed__7 = _init_l_Lean_Meta_commandRegister__simp__attr_________closed__7(); -lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr_________closed__7); -l_Lean_Meta_commandRegister__simp__attr_________closed__8 = _init_l_Lean_Meta_commandRegister__simp__attr_________closed__8(); -lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr_________closed__8); -l_Lean_Meta_commandRegister__simp__attr_________closed__9 = _init_l_Lean_Meta_commandRegister__simp__attr_________closed__9(); -lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr_________closed__9); -l_Lean_Meta_commandRegister__simp__attr_________closed__10 = _init_l_Lean_Meta_commandRegister__simp__attr_________closed__10(); -lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr_________closed__10); -l_Lean_Meta_commandRegister__simp__attr_________closed__11 = _init_l_Lean_Meta_commandRegister__simp__attr_________closed__11(); -lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr_________closed__11); -l_Lean_Meta_commandRegister__simp__attr_________closed__12 = _init_l_Lean_Meta_commandRegister__simp__attr_________closed__12(); -lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr_________closed__12); -l_Lean_Meta_commandRegister__simp__attr_________closed__13 = _init_l_Lean_Meta_commandRegister__simp__attr_________closed__13(); -lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr_________closed__13); -l_Lean_Meta_commandRegister__simp__attr_________closed__14 = _init_l_Lean_Meta_commandRegister__simp__attr_________closed__14(); -lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr_________closed__14); -l_Lean_Meta_commandRegister__simp__attr_________closed__15 = _init_l_Lean_Meta_commandRegister__simp__attr_________closed__15(); -lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr_________closed__15); -l_Lean_Meta_commandRegister__simp__attr_________closed__16 = _init_l_Lean_Meta_commandRegister__simp__attr_________closed__16(); -lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr_________closed__16); -l_Lean_Meta_commandRegister__simp__attr_________closed__17 = _init_l_Lean_Meta_commandRegister__simp__attr_________closed__17(); -lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr_________closed__17); -l_Lean_Meta_commandRegister__simp__attr______ = _init_l_Lean_Meta_commandRegister__simp__attr______(); -lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr______); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__1 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__1(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__1); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__2 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__2(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__2); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__3 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__3(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__3); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__4 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__4(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__4); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__5 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__5(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__5); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__6 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__6(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__6); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__7 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__7(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__7); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__8 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__8(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__8); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__9 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__9(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__9); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__10 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__10(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__10); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__11 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__11(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__11); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__12 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__12(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__12); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__13 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__13(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__13); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__14 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__14(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__14); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__15 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__15(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__15); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__16 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__16(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__16); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__17 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__17(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__17); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__18 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__18(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__18); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__19 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__19(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__19); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__20 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__20(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__20); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__21 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__21(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__21); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__22 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__22(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__22); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__23 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__23(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__23); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__24 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__24(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__24); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__25 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__25(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__25); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__26 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__26(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__26); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__27 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__27(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__27); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__28 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__28(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__28); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__29 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__29(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__29); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__30 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__30(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__30); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__31 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__31(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__31); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__32 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__32(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__32); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__33 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__33(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__33); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__34 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__34(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__34); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__35 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__35(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__35); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__36 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__36(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__36); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__37 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__37(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__37); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__38 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__38(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__38); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__39 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__39(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__39); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__40 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__40(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__40); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__41 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__41(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__41); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__42 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__42(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__42); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__43 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__43(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__43); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__44 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__44(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__44); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__45 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__45(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__45); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__46 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__46(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__46); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__47 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__47(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__47); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__48 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__48(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__48); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__49 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__49(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__49); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__50 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__50(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__50); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__51 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__51(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__51); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__52 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__52(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__52); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__53 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__53(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__53); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__54 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__54(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__54); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__55 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__55(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__55); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__56 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__56(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__56); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__57 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__57(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__57); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__58 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__58(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__58); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__59 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__59(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__59); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__60 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__60(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__60); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__61 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__61(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__61); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__62 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__62(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__62); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__63 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__63(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__63); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__64 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__64(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__64); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__65 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__65(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__65); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__66 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__66(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__66); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__67 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__67(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__67); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__68 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__68(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__68); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__69 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__69(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__69); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__70 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__70(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__70); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__71 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__71(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__71); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__72 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__72(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__72); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__73 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__73(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__73); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__74 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__74(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__74); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__75 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__75(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__75); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__76 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__76(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__76); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__77 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__77(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__77); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__78 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__78(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__78); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__79 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__79(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__79); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__80 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__80(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__80); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__81 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__81(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__81); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__82 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__82(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__82); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__83 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__83(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__83); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__84 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__84(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__84); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__85 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__85(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__85); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__86 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__86(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__86); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__87 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__87(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__87); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__88 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__88(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__88); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__89 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__89(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__89); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__90 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__90(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__90); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__91 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__91(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__91); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__92 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__92(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__92); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__93 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__93(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__93); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__94 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__94(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__94); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__95 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__95(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__95); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__96 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__96(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__96); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__97 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__97(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__97); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__98 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__98(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__98); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__99 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__99(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__99); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__100 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__100(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__100); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__101 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__101(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__101); -l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__102 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__102(); -lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr________1___closed__102); +l_Lean_Meta_commandRegister__simp__attr_______closed__1 = _init_l_Lean_Meta_commandRegister__simp__attr_______closed__1(); +lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr_______closed__1); +l_Lean_Meta_commandRegister__simp__attr_______closed__2 = _init_l_Lean_Meta_commandRegister__simp__attr_______closed__2(); +lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr_______closed__2); +l_Lean_Meta_commandRegister__simp__attr_______closed__3 = _init_l_Lean_Meta_commandRegister__simp__attr_______closed__3(); +lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr_______closed__3); +l_Lean_Meta_commandRegister__simp__attr_______closed__4 = _init_l_Lean_Meta_commandRegister__simp__attr_______closed__4(); +lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr_______closed__4); +l_Lean_Meta_commandRegister__simp__attr_______closed__5 = _init_l_Lean_Meta_commandRegister__simp__attr_______closed__5(); +lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr_______closed__5); +l_Lean_Meta_commandRegister__simp__attr_______closed__6 = _init_l_Lean_Meta_commandRegister__simp__attr_______closed__6(); +lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr_______closed__6); +l_Lean_Meta_commandRegister__simp__attr_______closed__7 = _init_l_Lean_Meta_commandRegister__simp__attr_______closed__7(); +lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr_______closed__7); +l_Lean_Meta_commandRegister__simp__attr_______closed__8 = _init_l_Lean_Meta_commandRegister__simp__attr_______closed__8(); +lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr_______closed__8); +l_Lean_Meta_commandRegister__simp__attr_______closed__9 = _init_l_Lean_Meta_commandRegister__simp__attr_______closed__9(); +lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr_______closed__9); +l_Lean_Meta_commandRegister__simp__attr_______closed__10 = _init_l_Lean_Meta_commandRegister__simp__attr_______closed__10(); +lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr_______closed__10); +l_Lean_Meta_commandRegister__simp__attr_______closed__11 = _init_l_Lean_Meta_commandRegister__simp__attr_______closed__11(); +lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr_______closed__11); +l_Lean_Meta_commandRegister__simp__attr_______closed__12 = _init_l_Lean_Meta_commandRegister__simp__attr_______closed__12(); +lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr_______closed__12); +l_Lean_Meta_commandRegister__simp__attr_______closed__13 = _init_l_Lean_Meta_commandRegister__simp__attr_______closed__13(); +lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr_______closed__13); +l_Lean_Meta_commandRegister__simp__attr_______closed__14 = _init_l_Lean_Meta_commandRegister__simp__attr_______closed__14(); +lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr_______closed__14); +l_Lean_Meta_commandRegister__simp__attr_______closed__15 = _init_l_Lean_Meta_commandRegister__simp__attr_______closed__15(); +lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr_______closed__15); +l_Lean_Meta_commandRegister__simp__attr_______closed__16 = _init_l_Lean_Meta_commandRegister__simp__attr_______closed__16(); +lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr_______closed__16); +l_Lean_Meta_commandRegister__simp__attr_______closed__17 = _init_l_Lean_Meta_commandRegister__simp__attr_______closed__17(); +lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr_______closed__17); +l_Lean_Meta_commandRegister__simp__attr____ = _init_l_Lean_Meta_commandRegister__simp__attr____(); +lean_mark_persistent(l_Lean_Meta_commandRegister__simp__attr____); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__1 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__1(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__1); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__2 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__2(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__2); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__3 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__3(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__3); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__4 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__4(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__4); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__5 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__5(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__5); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__6 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__6(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__6); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__7 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__7(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__7); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__8 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__8(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__8); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__9 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__9(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__9); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__10 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__10(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__10); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__11 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__11(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__11); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__12 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__12(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__12); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__13 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__13(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__13); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__14 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__14(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__14); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__15 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__15(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__15); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__16 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__16(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__16); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__17 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__17(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__17); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__18 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__18(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__18); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__19 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__19(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__19); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__20 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__20(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__20); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__21 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__21(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__21); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__22 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__22(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__22); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__23 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__23(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__23); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__24 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__24(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__24); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__25 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__25(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__25); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__26 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__26(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__26); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__27 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__27(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__27); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__28 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__28(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__28); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__29 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__29(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__29); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__30 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__30(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__30); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__31 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__31(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__31); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__32 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__32(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__32); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__33 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__33(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__33); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__34 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__34(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__34); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__35 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__35(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__35); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__36 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__36(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__36); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__37 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__37(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__37); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__38 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__38(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__38); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__39 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__39(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__39); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__40 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__40(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__40); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__41 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__41(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__41); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__42 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__42(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__42); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__43 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__43(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__43); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__44 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__44(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__44); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__45 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__45(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__45); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__46 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__46(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__46); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__47 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__47(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__47); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__48 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__48(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__48); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__49 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__49(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__49); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__50 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__50(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__50); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__51 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__51(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__51); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__52 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__52(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__52); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__53 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__53(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__53); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__54 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__54(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__54); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__55 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__55(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__55); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__56 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__56(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__56); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__57 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__57(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__57); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__58 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__58(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__58); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__59 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__59(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__59); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__60 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__60(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__60); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__61 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__61(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__61); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__62 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__62(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__62); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__63 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__63(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__63); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__64 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__64(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__64); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__65 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__65(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__65); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__66 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__66(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__66); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__67 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__67(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__67); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__68 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__68(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__68); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__69 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__69(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__69); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__70 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__70(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__70); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__71 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__71(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__71); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__72 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__72(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__72); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__73 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__73(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__73); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__74 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__74(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__74); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__75 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__75(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__75); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__76 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__76(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__76); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__77 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__77(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__77); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__78 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__78(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__78); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__79 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__79(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__79); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__80 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__80(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__80); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__81 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__81(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__81); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__82 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__82(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__82); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__83 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__83(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__83); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__84 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__84(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__84); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__85 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__85(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__85); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__86 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__86(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__86); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__87 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__87(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__87); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__88 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__88(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__88); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__89 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__89(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__89); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__90 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__90(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__90); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__91 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__91(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__91); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__92 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__92(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__92); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__93 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__93(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__93); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__94 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__94(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__94); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__95 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__95(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__95); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__96 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__96(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__96); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__97 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__97(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__97); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__98 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__98(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__98); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__99 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__99(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__99); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__100 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__100(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__100); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__101 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__101(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__101); +l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__102 = _init_l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__102(); +lean_mark_persistent(l_Lean_Meta___aux__Lean__Meta__Tactic__Simp__SimpTheorems______macroRules__Lean__Meta__commandRegister__simp__attr______1___closed__102); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Meta/WHNF.c b/stage0/stdlib/Lean/Meta/WHNF.c index 06da487ee4e..2be23d766e4 100644 --- a/stage0/stdlib/Lean/Meta/WHNF.c +++ b/stage0/stdlib/Lean/Meta/WHNF.c @@ -129,7 +129,7 @@ size_t lean_usize_shift_right(size_t, size_t); static lean_object* l_Lean_Meta_reduceNative_x3f___closed__14; LEAN_EXPORT lean_object* l_Lean_Meta_smartUnfoldingMatch_x3f(lean_object*); lean_object* l_Lean_Expr_getRevArg_x21(lean_object*, lean_object*); -uint8_t l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_8567_(uint8_t, uint8_t); +uint8_t l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_9542_(uint8_t, uint8_t); static lean_object* l_Lean_Meta_toCtorIfLit___closed__16; static lean_object* l_Lean_Meta_canUnfoldAtMatcher___closed__10; uint8_t lean_usize_dec_lt(size_t, size_t); @@ -10020,7 +10020,7 @@ lean_inc(x_9); lean_dec(x_7); x_10 = 2; x_11 = lean_unbox(x_8); -x_12 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_8567_(x_11, x_10); +x_12 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_9542_(x_11, x_10); if (x_12 == 0) { lean_object* x_13; uint8_t x_14; lean_object* x_15; @@ -20464,7 +20464,7 @@ x_10 = lean_ctor_get(x_7, 1); x_11 = 3; x_12 = lean_unbox(x_9); lean_dec(x_9); -x_13 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_8567_(x_12, x_11); +x_13 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_9542_(x_12, x_11); if (x_13 == 0) { lean_object* x_14; @@ -20496,7 +20496,7 @@ lean_dec(x_7); x_18 = 3; x_19 = lean_unbox(x_16); lean_dec(x_16); -x_20 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_8567_(x_19, x_18); +x_20 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_9542_(x_19, x_18); if (x_20 == 0) { lean_object* x_21; lean_object* x_22; diff --git a/stage0/stdlib/Lean/MetavarContext.c b/stage0/stdlib/Lean/MetavarContext.c index b0da6f87d97..efd6c9bcfd9 100644 --- a/stage0/stdlib/Lean/MetavarContext.c +++ b/stage0/stdlib/Lean/MetavarContext.c @@ -662,7 +662,6 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_MetavarContext_MkBindi lean_object* l_Std_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentArray_anyMAux___at_Lean_MetavarContext_MkBinding_collectForwardDeps___spec__22___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashSetImp_expand___at___private_Lean_MetavarContext_0__Lean_MetavarContext_DependsOn_shouldVisit___spec__4(lean_object*, lean_object*); -lean_object* l_panic___at_String_toNat_x21___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Std_HashSetImp_moveEntries___at___private_Lean_MetavarContext_0__Lean_MetavarContext_DependsOn_shouldVisit___spec__5(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_MetavarContext_instantiateLCtxMVars___spec__6(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_MetavarContext_DependsOn_dep_visitMain___at_Lean_MetavarContext_localDeclDependsOn___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); @@ -1089,6 +1088,7 @@ LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_MetavarContext_ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_MetavarContext_exprDependsOn_x27___spec__7___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_expr_update_lambda(lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_MetavarContext_instantiateExprMVars___spec__27___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_panic___at_Lean_TSyntax_getNat___spec__1(lean_object*); lean_object* lean_name_mk_numeral(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_elimApp___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_level_update_succ(lean_object*, lean_object*); @@ -6047,7 +6047,7 @@ if (lean_obj_tag(x_3) == 0) { lean_object* x_4; lean_object* x_5; x_4 = l_Lean_MetavarContext_getLevelDepth___closed__2; -x_5 = l_panic___at_String_toNat_x21___spec__1(x_4); +x_5 = l_panic___at_Lean_TSyntax_getNat___spec__1(x_4); return x_5; } else diff --git a/stage0/stdlib/Lean/Parser.c b/stage0/stdlib/Lean/Parser.c index c1a590088f6..a30077e709b 100644 --- a/stage0/stdlib/Lean/Parser.c +++ b/stage0/stdlib/Lean/Parser.c @@ -83,11 +83,11 @@ static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__9 LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_strLit_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__160; lean_object* l_Lean_Parser_optional_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__232; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_ident_parenthesizer___closed__8; static lean_object* l_Lean_PrettyPrinter_Parenthesizer_interpretParserDescr___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_mkAntiquot_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Formatter_checkWsBefore_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_interpolatedStrKind; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Formatter_numLit_formatter___closed__3; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_ident_parenthesizer___closed__4; lean_object* l_Lean_Parser_withPosition(lean_object*); @@ -106,7 +106,6 @@ lean_object* l_Lean_Parser_many_formatter(lean_object*, lean_object*, lean_objec static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__201; lean_object* l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_numLit_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_nameLitKind; lean_object* l_Lean_Parser_Term_str_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_ident_parenthesizer___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_interpretParserDescr___elambda__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -161,12 +160,13 @@ static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__1 static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__149; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__117; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__101; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__234; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__66; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__20; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__238; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_numLit_parenthesizer___closed__4; LEAN_EXPORT lean_object* lean_mk_antiquot_parenthesizer(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_interpretParserDescr___elambda__6(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_numLitKind; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__228; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__207; lean_object* l_Lean_Parser_charLit_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -174,7 +174,6 @@ static lean_object* l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_charLit_pare lean_object* l_Lean_Parser_checkNoWsBefore(lean_object*); lean_object* l_Lean_Parser_scientificLit_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__151; -extern lean_object* l_Lean_strLitKind; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__32; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__10; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__174; @@ -188,6 +187,7 @@ lean_object* l_Lean_Parser_manyIndent_formatter(lean_object*, lean_object*, lean static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__5; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_ident_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Formatter_charLit_formatter(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__237; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__69; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__206; lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkWsBefore_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -228,7 +228,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_interpretParserDescr___e static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__39; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__109; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__65; -extern lean_object* l_Lean_charLitKind; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__173; extern lean_object* l_Lean_PrettyPrinter_combinatorFormatterAttribute; lean_object* l_Lean_Parser_strLit_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -261,7 +260,7 @@ static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__1 static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__75; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__159; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__205; -lean_object* l_Lean_Parser_registerAlias(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Parser_registerAlias(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_interpretParserDescr___elambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__61; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__112; @@ -292,7 +291,6 @@ LEAN_EXPORT lean_object* lean_pretty_printer_parenthesizer_interpret_parser_desc static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__187; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_interpretParserDescr___elambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__95; -extern lean_object* l_Lean_identKind; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Formatter_ident_formatter___closed__7; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__79; lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkNoWsBefore_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -305,7 +303,6 @@ lean_object* l_Lean_Parser_getUnaryAlias___rarg(lean_object*, lean_object*, lean static lean_object* l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_numLit_parenthesizer___closed__1; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__88; extern lean_object* l_Lean_PrettyPrinter_Formatter_formatterAliasesRef; -extern lean_object* l_Lean_scientificLitKind; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__23; lean_object* l_Lean_Parser_Term_ident_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__180; @@ -323,6 +320,7 @@ static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__2 static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__48; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__162; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_interpretParserDescr___elambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__239; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__141; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_numLit_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__111; @@ -333,6 +331,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_interpretParserDescr___e static lean_object* l___regBuiltin_Lean_PrettyPrinter_Formatter_strLit_formatter___closed__2; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__34; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__220; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__233; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__195; lean_object* l_Lean_Parser_charLit_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__200; @@ -387,11 +386,13 @@ LEAN_EXPORT lean_object* lean_pretty_printer_formatter_interpret_parser_descr(le static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__74; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__60; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__47; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__235; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Formatter_ident_formatter___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_mkAntiquot_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Parenthesizer_withPosition_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__2; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_scientificLit_parenthesizer___closed__4; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__236; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__157; lean_object* l_Lean_Parser_many_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__15; @@ -600,22 +601,34 @@ return x_2; static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__13() { _start: { +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = 0; +x_3 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__14() { +_start: +{ lean_object* x_1; x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_checkWsBefore_formatter___boxed), 5, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__14() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__15() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__13; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__14; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__15() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__16() { _start: { lean_object* x_1; @@ -623,7 +636,7 @@ x_1 = l_Lean_PrettyPrinter_Formatter_formatterAliasesRef; return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__16() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__17() { _start: { lean_object* x_1; @@ -631,17 +644,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_checkWsBefor return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__17() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__18() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__16; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__17; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__18() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__19() { _start: { lean_object* x_1; @@ -649,7 +662,7 @@ x_1 = l_Lean_PrettyPrinter_Parenthesizer_parenthesizerAliasesRef; return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__19() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__20() { _start: { lean_object* x_1; @@ -657,17 +670,17 @@ x_1 = lean_mk_string_from_bytes("noWs", 4); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__20() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__19; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__20; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__21() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__22() { _start: { lean_object* x_1; @@ -675,26 +688,26 @@ x_1 = lean_mk_string_from_bytes("no space before", 15); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__22() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__23() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__21; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__22; x_2 = l_Lean_Parser_checkNoWsBefore(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__23() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__24() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__22; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__23; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__24() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__25() { _start: { lean_object* x_1; @@ -702,27 +715,27 @@ x_1 = lean_mk_string_from_bytes("checkNoWsBefore", 15); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__25() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__9; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__24; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__25; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__26() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__27() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__25; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__26; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__27() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__28() { _start: { lean_object* x_1; @@ -730,17 +743,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_checkNoWsBefore_ return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__28() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__29() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__27; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__28; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__29() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__30() { _start: { lean_object* x_1; @@ -748,17 +761,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_checkNoWsBef return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__30() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__31() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__29; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__30; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__31() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__32() { _start: { lean_object* x_1; @@ -766,17 +779,17 @@ x_1 = lean_mk_string_from_bytes("linebreak", 9); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__32() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__33() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__31; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__32; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__33() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__34() { _start: { lean_object* x_1; @@ -784,26 +797,26 @@ x_1 = lean_mk_string_from_bytes("line break", 10); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__34() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__35() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__33; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__34; x_2 = l_Lean_Parser_checkLinebreakBefore(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__35() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__36() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__34; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__35; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__36() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__37() { _start: { lean_object* x_1; @@ -811,27 +824,27 @@ x_1 = lean_mk_string_from_bytes("checkLinebreakBefore", 20); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__37() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__38() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__9; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__36; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__37; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__38() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__39() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__37; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__38; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__39() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__40() { _start: { lean_object* x_1; @@ -839,17 +852,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_checkLinebreakBe return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__40() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__41() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__39; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__40; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__41() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__42() { _start: { lean_object* x_1; @@ -857,17 +870,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_checkLinebre return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__42() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__43() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__41; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__42; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__43() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__44() { _start: { lean_object* x_1; @@ -875,17 +888,17 @@ x_1 = lean_mk_string_from_bytes("num", 3); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__44() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__45() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__43; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__44; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__45() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__46() { _start: { lean_object* x_1; lean_object* x_2; @@ -895,17 +908,39 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__46() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__47() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_numLitKind; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__45; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__47() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__48() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__49() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__48; +x_2 = 1; +x_3 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__50() { _start: { lean_object* x_1; @@ -913,17 +948,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_numLit_formatter), 5, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__48() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__51() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__47; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__50; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__49() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__52() { _start: { lean_object* x_1; @@ -931,17 +966,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_numLit_parenthesizer), 5, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__50() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__53() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__49; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__52; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__51() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__54() { _start: { lean_object* x_1; @@ -949,17 +984,17 @@ x_1 = lean_mk_string_from_bytes("str", 3); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__52() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__55() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__51; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__54; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__53() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__56() { _start: { lean_object* x_1; lean_object* x_2; @@ -969,17 +1004,17 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__54() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__57() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_strLitKind; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__55; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__55() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__58() { _start: { lean_object* x_1; @@ -987,17 +1022,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_strLit_formatter), 5, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__56() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__59() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__55; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__58; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__57() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__60() { _start: { lean_object* x_1; @@ -1005,17 +1040,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_strLit_parenthesizer), 5, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__58() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__61() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__57; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__60; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__59() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__62() { _start: { lean_object* x_1; @@ -1023,17 +1058,17 @@ x_1 = lean_mk_string_from_bytes("char", 4); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__60() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__63() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__59; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__62; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__61() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__64() { _start: { lean_object* x_1; lean_object* x_2; @@ -1043,17 +1078,17 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__62() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__65() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_charLitKind; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__63; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__63() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__66() { _start: { lean_object* x_1; @@ -1061,17 +1096,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_charLit_formatter), 5, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__64() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__67() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__63; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__66; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__65() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__68() { _start: { lean_object* x_1; @@ -1079,17 +1114,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_charLit_parenthesizer), 5, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__66() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__69() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__65; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__68; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__67() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__70() { _start: { lean_object* x_1; @@ -1097,17 +1132,17 @@ x_1 = lean_mk_string_from_bytes("name", 4); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__68() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__71() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__67; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__70; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__69() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__72() { _start: { lean_object* x_1; lean_object* x_2; @@ -1117,17 +1152,17 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__70() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__73() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_nameLitKind; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__71; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__71() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__74() { _start: { lean_object* x_1; @@ -1135,17 +1170,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_nameLit_formatter), 5, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__72() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__75() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__71; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__74; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__73() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__76() { _start: { lean_object* x_1; @@ -1153,17 +1188,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_nameLit_parenthesizer), 5, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__74() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__77() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__73; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__76; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__75() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__78() { _start: { lean_object* x_1; @@ -1171,17 +1206,17 @@ x_1 = lean_mk_string_from_bytes("scientific", 10); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__76() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__79() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__75; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__78; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__77() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__80() { _start: { lean_object* x_1; lean_object* x_2; @@ -1191,17 +1226,17 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__78() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__81() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_scientificLitKind; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__79; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__79() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__82() { _start: { lean_object* x_1; @@ -1209,17 +1244,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_scientificLit_formatter), 5, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__80() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__83() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__79; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__82; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__81() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__84() { _start: { lean_object* x_1; @@ -1227,17 +1262,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_scientificLit_parenthesizer), 5, return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__82() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__85() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__81; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__84; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__83() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__86() { _start: { lean_object* x_1; @@ -1245,17 +1280,17 @@ x_1 = lean_mk_string_from_bytes("ident", 5); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__84() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__87() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__83; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__86; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__85() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__88() { _start: { lean_object* x_1; lean_object* x_2; @@ -1265,17 +1300,17 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__86() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__89() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_identKind; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__87; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__87() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__90() { _start: { lean_object* x_1; @@ -1283,17 +1318,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_ident_formatter), 5, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__88() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__91() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__87; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__90; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__89() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__92() { _start: { lean_object* x_1; @@ -1301,17 +1336,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_ident_parenthesizer), 5, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__90() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__93() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__89; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__92; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__91() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__94() { _start: { lean_object* x_1; @@ -1319,17 +1354,17 @@ x_1 = lean_mk_string_from_bytes("colGt", 5); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__92() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__95() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__91; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__94; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__93() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__96() { _start: { lean_object* x_1; @@ -1337,11 +1372,11 @@ x_1 = lean_alloc_closure((void*)(l_id___rarg___boxed), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__94() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__97() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__93; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__96; x_2 = lean_box(1); x_3 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_3, 0, x_1); @@ -1350,7 +1385,7 @@ lean_ctor_set(x_3, 2, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__95() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__98() { _start: { lean_object* x_1; @@ -1358,59 +1393,59 @@ x_1 = lean_mk_string_from_bytes("checkColGt", 10); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__96() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__99() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__95; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__98; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_checkColGtFn___boxed), 3, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__97() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__100() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__94; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__96; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__97; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__99; x_3 = lean_alloc_ctor(0, 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_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__98() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__101() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__97; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__100; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__99() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__102() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__9; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__95; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__98; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__100() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__103() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__99; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__102; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__101() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__104() { _start: { lean_object* x_1; @@ -1418,17 +1453,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_checkColGt_forma return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__102() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__105() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__101; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__104; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__103() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__106() { _start: { lean_object* x_1; @@ -1436,17 +1471,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_checkColGt_p return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__104() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__107() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__103; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__106; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__105() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__108() { _start: { lean_object* x_1; @@ -1454,17 +1489,17 @@ x_1 = lean_mk_string_from_bytes("colGe", 5); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__106() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__109() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__105; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__108; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__107() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__110() { _start: { lean_object* x_1; @@ -1472,59 +1507,59 @@ x_1 = lean_mk_string_from_bytes("checkColGe", 10); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__108() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__111() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__107; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__110; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_checkColGeFn___boxed), 3, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__109() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__112() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__94; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__108; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__97; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__111; x_3 = lean_alloc_ctor(0, 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_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__110() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__113() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__109; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__112; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__111() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__114() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__9; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__107; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__110; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__112() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__115() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__111; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__114; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__113() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__116() { _start: { lean_object* x_1; @@ -1532,17 +1567,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_checkColGe_forma return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__114() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__117() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__113; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__116; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__115() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__118() { _start: { lean_object* x_1; @@ -1550,17 +1585,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_checkColGe_p return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__116() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__119() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__115; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__118; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__117() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__120() { _start: { lean_object* x_1; @@ -1568,17 +1603,17 @@ x_1 = lean_mk_string_from_bytes("lookahead", 9); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__118() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__121() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__117; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__120; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__119() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__122() { _start: { lean_object* x_1; @@ -1586,37 +1621,59 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_lookahead), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__120() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__123() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__119; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__122; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__121() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__124() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__9; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__117; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__120; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__122() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__125() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__121; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__124; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__123() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__126() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__127() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__126; +x_2 = 1; +x_3 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__128() { _start: { lean_object* x_1; @@ -1624,17 +1681,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_lookahead_format return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__124() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__129() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__123; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__128; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__125() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__130() { _start: { lean_object* x_1; @@ -1642,17 +1699,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_lookahead_pa return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__126() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__131() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__125; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__130; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__127() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__132() { _start: { lean_object* x_1; @@ -1660,17 +1717,17 @@ x_1 = lean_mk_string_from_bytes("atomic", 6); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__128() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__133() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__127; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__132; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__129() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__134() { _start: { lean_object* x_1; @@ -1678,37 +1735,37 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_atomic), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__130() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__135() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__129; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__134; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__131() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__136() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__9; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__127; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__132; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__132() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__137() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__131; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__136; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__133() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__138() { _start: { lean_object* x_1; @@ -1716,17 +1773,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_atomic_formatter return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__134() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__139() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__133; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__138; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__135() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__140() { _start: { lean_object* x_1; @@ -1734,17 +1791,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_atomic_paren return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__136() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__141() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__135; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__140; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__137() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__142() { _start: { lean_object* x_1; @@ -1752,17 +1809,17 @@ x_1 = lean_mk_string_from_bytes("many", 4); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__138() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__143() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__137; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__142; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__139() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__144() { _start: { lean_object* x_1; @@ -1770,37 +1827,37 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_many), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__140() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__145() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__139; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__144; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__141() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__146() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__9; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__137; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__142; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__142() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__147() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__141; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__146; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__143() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__148() { _start: { lean_object* x_1; @@ -1808,17 +1865,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_many_formatter), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__144() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__149() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__143; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__148; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__145() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__150() { _start: { lean_object* x_1; @@ -1826,17 +1883,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_many_parenthesizer), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__146() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__151() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__145; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__150; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__147() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__152() { _start: { lean_object* x_1; @@ -1844,17 +1901,17 @@ x_1 = lean_mk_string_from_bytes("many1", 5); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__148() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__153() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__147; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__152; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__149() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__154() { _start: { lean_object* x_1; @@ -1862,37 +1919,37 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_many1), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__150() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__155() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__149; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__154; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__151() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__156() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__9; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__147; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__152; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__152() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__157() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__151; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__156; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__153() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__158() { _start: { lean_object* x_1; @@ -1900,17 +1957,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_many1_formatter), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__154() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__159() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__153; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__158; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__155() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__160() { _start: { lean_object* x_1; @@ -1918,17 +1975,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_many1_parenthesizer), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__156() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__161() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__155; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__160; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__157() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__162() { _start: { lean_object* x_1; @@ -1936,17 +1993,17 @@ x_1 = lean_mk_string_from_bytes("manyIndent", 10); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__158() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__163() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__157; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__162; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__159() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__164() { _start: { lean_object* x_1; @@ -1954,37 +2011,37 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_manyIndent), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__160() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__165() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__159; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__164; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__161() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__166() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__9; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__157; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__162; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__162() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__167() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__161; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__166; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__163() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__168() { _start: { lean_object* x_1; @@ -1992,17 +2049,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_manyIndent_formatter), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__164() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__169() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__163; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__168; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__165() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__170() { _start: { lean_object* x_1; @@ -2010,17 +2067,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_manyIndent_parenthesizer), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__166() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__171() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__165; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__170; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__167() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__172() { _start: { lean_object* x_1; @@ -2028,17 +2085,17 @@ x_1 = lean_mk_string_from_bytes("many1Indent", 11); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__168() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__173() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__167; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__172; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__169() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__174() { _start: { lean_object* x_1; @@ -2046,37 +2103,37 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_many1Indent), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__170() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__175() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__169; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__174; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__171() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__176() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__9; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__167; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__172; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__172() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__177() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__171; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__176; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__173() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__178() { _start: { lean_object* x_1; @@ -2084,17 +2141,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_many1Indent_formatter), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__174() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__179() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__173; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__178; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__175() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__180() { _start: { lean_object* x_1; @@ -2102,17 +2159,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_many1Indent_parenthesizer), 6, 0) return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__176() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__181() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__175; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__180; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__177() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__182() { _start: { lean_object* x_1; @@ -2120,17 +2177,17 @@ x_1 = lean_mk_string_from_bytes("optional", 8); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__178() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__183() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__177; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__182; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__179() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__184() { _start: { lean_object* x_1; @@ -2138,37 +2195,49 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_optional), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__180() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__185() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__179; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__184; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__181() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__186() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__9; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__177; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__182; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__182() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__187() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__181; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__186; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__183() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__188() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__48; +x_2 = 0; +x_3 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__189() { _start: { lean_object* x_1; @@ -2176,17 +2245,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_optional_formatter), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__184() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__190() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__183; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__189; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__185() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__191() { _start: { lean_object* x_1; @@ -2194,17 +2263,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_optional_parenthesizer), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__186() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__192() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__185; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__191; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__187() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__193() { _start: { lean_object* x_1; @@ -2212,17 +2281,17 @@ x_1 = lean_mk_string_from_bytes("withPosition", 12); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__188() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__194() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__187; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__193; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__189() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__195() { _start: { lean_object* x_1; @@ -2230,37 +2299,37 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_withPosition), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__190() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__196() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__189; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__195; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__191() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__197() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__9; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__187; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__193; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__192() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__198() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__191; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__197; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__193() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__199() { _start: { lean_object* x_1; @@ -2268,17 +2337,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_withPosition_for return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__194() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__200() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__193; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__199; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__195() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__201() { _start: { lean_object* x_1; @@ -2286,17 +2355,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_withPosition return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__196() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__202() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__195; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__201; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__197() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__203() { _start: { lean_object* x_1; @@ -2304,17 +2373,17 @@ x_1 = lean_mk_string_from_bytes("interpolatedStr", 15); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__198() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__204() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__197; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__203; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__199() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__205() { _start: { lean_object* x_1; @@ -2322,27 +2391,45 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_interpolatedStr), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__200() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__206() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__199; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__205; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__201() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__207() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("interpolatedStrKind", 19); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__208() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__207; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__209() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_interpolatedStrKind; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__208; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__202() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__210() { _start: { lean_object* x_1; @@ -2350,17 +2437,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_interpolatedStr_ return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__203() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__211() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__202; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__210; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__204() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__212() { _start: { lean_object* x_1; @@ -2368,17 +2455,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_interpolated return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__205() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__213() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__204; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__212; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__206() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__214() { _start: { lean_object* x_1; @@ -2386,17 +2473,17 @@ x_1 = lean_mk_string_from_bytes("orelse", 6); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__207() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__215() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__206; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__214; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__208() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__216() { _start: { lean_object* x_1; @@ -2404,37 +2491,37 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_orelse), 2, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__209() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__217() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__208; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__216; x_2 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__210() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__218() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__9; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__206; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__214; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__211() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__219() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__210; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__218; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__212() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__220() { _start: { lean_object* x_1; @@ -2442,17 +2529,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_orelse_formatter return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__213() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__221() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__212; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__220; x_2 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__214() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__222() { _start: { lean_object* x_1; @@ -2460,17 +2547,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_orelse_paren return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__215() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__223() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__214; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__222; x_2 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__216() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__224() { _start: { lean_object* x_1; @@ -2478,17 +2565,17 @@ x_1 = lean_mk_string_from_bytes("andthen", 7); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__217() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__225() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__216; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__224; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__218() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__226() { _start: { lean_object* x_1; @@ -2496,37 +2583,37 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_andthen), 2, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__219() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__227() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__218; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__226; x_2 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__220() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__228() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__9; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__216; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__224; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__221() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__229() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__220; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__228; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__222() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__230() { _start: { lean_object* x_1; @@ -2534,17 +2621,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatte return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__223() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__231() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__222; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__230; x_2 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__224() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__232() { _start: { lean_object* x_1; @@ -2552,17 +2639,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_pare return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__225() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__233() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__224; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__232; x_2 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__226() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__234() { _start: { lean_object* x_1; @@ -2570,17 +2657,17 @@ x_1 = lean_mk_string_from_bytes("notFollowedBy", 13); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__227() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__235() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__226; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__234; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__228() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__236() { _start: { lean_object* x_1; @@ -2588,17 +2675,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_notFollowedB return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__229() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__237() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__228; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__236; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__230() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__238() { _start: { lean_object* x_1; @@ -2606,11 +2693,11 @@ x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_notFollowedBy_fo return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__231() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__239() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__230; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__238; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; @@ -2619,646 +2706,627 @@ return x_2; LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8_(lean_object* x_1) { _start: { -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__2; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__5; -x_4 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__12; -x_5 = l_Lean_Parser_registerAlias(x_2, x_3, x_4, x_1); -if (lean_obj_tag(x_5) == 0) -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_6 = lean_ctor_get(x_5, 1); -lean_inc(x_6); -lean_dec(x_5); -x_7 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__15; -x_8 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__14; -x_9 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_2, x_8, x_6); -if (lean_obj_tag(x_9) == 0) -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_10 = lean_ctor_get(x_9, 1); -lean_inc(x_10); -lean_dec(x_9); -x_11 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__18; -x_12 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__17; -x_13 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_2, x_12, x_10); -if (lean_obj_tag(x_13) == 0) -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -lean_dec(x_13); -x_15 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__20; -x_16 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__23; -x_17 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__26; -x_18 = l_Lean_Parser_registerAlias(x_15, x_16, x_17, x_14); -if (lean_obj_tag(x_18) == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -lean_dec(x_18); -x_20 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__28; -x_21 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_15, x_20, x_19); -if (lean_obj_tag(x_21) == 0) -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_21, 1); -lean_inc(x_22); -lean_dec(x_21); -x_23 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__30; -x_24 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_15, x_23, x_22); -if (lean_obj_tag(x_24) == 0) -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_25 = lean_ctor_get(x_24, 1); -lean_inc(x_25); -lean_dec(x_24); -x_26 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__32; -x_27 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__35; -x_28 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__38; -x_29 = l_Lean_Parser_registerAlias(x_26, x_27, x_28, x_25); -if (lean_obj_tag(x_29) == 0) -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_29, 1); -lean_inc(x_30); -lean_dec(x_29); -x_31 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__40; -x_32 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_26, x_31, x_30); -if (lean_obj_tag(x_32) == 0) +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_2 = lean_box(0); +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__2; +x_4 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__5; +x_5 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__12; +x_6 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__13; +x_7 = l_Lean_Parser_registerAlias(x_3, x_4, x_5, x_6, x_1); +if (lean_obj_tag(x_7) == 0) { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_32, 1); -lean_inc(x_33); -lean_dec(x_32); -x_34 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__42; -x_35 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_26, x_34, x_33); -if (lean_obj_tag(x_35) == 0) +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_7, 1); +lean_inc(x_8); +lean_dec(x_7); +x_9 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__16; +x_10 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__15; +x_11 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_3, x_10, x_8); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = lean_ctor_get(x_11, 1); +lean_inc(x_12); +lean_dec(x_11); +x_13 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__19; +x_14 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__18; +x_15 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_3, x_14, x_12); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_36 = lean_ctor_get(x_35, 1); -lean_inc(x_36); -lean_dec(x_35); -x_37 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__44; -x_38 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__45; -x_39 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__46; -x_40 = l_Lean_Parser_registerAlias(x_37, x_38, x_39, x_36); -if (lean_obj_tag(x_40) == 0) -{ -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_40, 1); -lean_inc(x_41); -lean_dec(x_40); -x_42 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__48; -x_43 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_37, x_42, x_41); +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +x_17 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__21; +x_18 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__24; +x_19 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__27; +x_20 = l_Lean_Parser_registerAlias(x_17, x_18, x_19, x_6, x_16); +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_20, 1); +lean_inc(x_21); +lean_dec(x_20); +x_22 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__29; +x_23 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_17, x_22, x_21); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); +x_25 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__31; +x_26 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_17, x_25, x_24); +if (lean_obj_tag(x_26) == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); +lean_dec(x_26); +x_28 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__33; +x_29 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__36; +x_30 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__39; +x_31 = l_Lean_Parser_registerAlias(x_28, x_29, x_30, x_6, x_27); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +lean_dec(x_31); +x_33 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__41; +x_34 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_28, x_33, x_32); +if (lean_obj_tag(x_34) == 0) +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_34, 1); +lean_inc(x_35); +lean_dec(x_34); +x_36 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__43; +x_37 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_28, x_36, x_35); +if (lean_obj_tag(x_37) == 0) +{ +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_38 = lean_ctor_get(x_37, 1); +lean_inc(x_38); +lean_dec(x_37); +x_39 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__45; +x_40 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__46; +x_41 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__47; +x_42 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__49; +x_43 = l_Lean_Parser_registerAlias(x_39, x_40, x_41, x_42, x_38); if (lean_obj_tag(x_43) == 0) { lean_object* x_44; lean_object* x_45; lean_object* x_46; x_44 = lean_ctor_get(x_43, 1); lean_inc(x_44); lean_dec(x_43); -x_45 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__50; -x_46 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_37, x_45, x_44); +x_45 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__51; +x_46 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_39, x_45, x_44); if (lean_obj_tag(x_46) == 0) { -lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +lean_object* x_47; lean_object* x_48; lean_object* x_49; x_47 = lean_ctor_get(x_46, 1); lean_inc(x_47); lean_dec(x_46); -x_48 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__52; -x_49 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__53; -x_50 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__54; -x_51 = l_Lean_Parser_registerAlias(x_48, x_49, x_50, x_47); -if (lean_obj_tag(x_51) == 0) -{ -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_51, 1); -lean_inc(x_52); -lean_dec(x_51); -x_53 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__56; -x_54 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_48, x_53, x_52); +x_48 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__53; +x_49 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_39, x_48, x_47); +if (lean_obj_tag(x_49) == 0) +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_50 = lean_ctor_get(x_49, 1); +lean_inc(x_50); +lean_dec(x_49); +x_51 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__55; +x_52 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__56; +x_53 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__57; +x_54 = l_Lean_Parser_registerAlias(x_51, x_52, x_53, x_42, x_50); if (lean_obj_tag(x_54) == 0) { lean_object* x_55; lean_object* x_56; lean_object* x_57; x_55 = lean_ctor_get(x_54, 1); lean_inc(x_55); lean_dec(x_54); -x_56 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__58; -x_57 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_48, x_56, x_55); +x_56 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__59; +x_57 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_51, x_56, x_55); if (lean_obj_tag(x_57) == 0) { -lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; +lean_object* x_58; lean_object* x_59; lean_object* x_60; x_58 = lean_ctor_get(x_57, 1); lean_inc(x_58); lean_dec(x_57); -x_59 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__60; -x_60 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__61; -x_61 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__62; -x_62 = l_Lean_Parser_registerAlias(x_59, x_60, x_61, x_58); -if (lean_obj_tag(x_62) == 0) -{ -lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_63 = lean_ctor_get(x_62, 1); -lean_inc(x_63); -lean_dec(x_62); -x_64 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__64; -x_65 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_59, x_64, x_63); +x_59 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__61; +x_60 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_51, x_59, x_58); +if (lean_obj_tag(x_60) == 0) +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_61 = lean_ctor_get(x_60, 1); +lean_inc(x_61); +lean_dec(x_60); +x_62 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__63; +x_63 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__64; +x_64 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__65; +x_65 = l_Lean_Parser_registerAlias(x_62, x_63, x_64, x_42, x_61); if (lean_obj_tag(x_65) == 0) { lean_object* x_66; lean_object* x_67; lean_object* x_68; x_66 = lean_ctor_get(x_65, 1); lean_inc(x_66); lean_dec(x_65); -x_67 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__66; -x_68 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_59, x_67, x_66); +x_67 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__67; +x_68 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_62, x_67, x_66); if (lean_obj_tag(x_68) == 0) { -lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; +lean_object* x_69; lean_object* x_70; lean_object* x_71; x_69 = lean_ctor_get(x_68, 1); lean_inc(x_69); lean_dec(x_68); -x_70 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__68; -x_71 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__69; -x_72 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__70; -x_73 = l_Lean_Parser_registerAlias(x_70, x_71, x_72, x_69); -if (lean_obj_tag(x_73) == 0) -{ -lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_74 = lean_ctor_get(x_73, 1); -lean_inc(x_74); -lean_dec(x_73); -x_75 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__72; -x_76 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_70, x_75, x_74); +x_70 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__69; +x_71 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_62, x_70, x_69); +if (lean_obj_tag(x_71) == 0) +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_72 = lean_ctor_get(x_71, 1); +lean_inc(x_72); +lean_dec(x_71); +x_73 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__71; +x_74 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__72; +x_75 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__73; +x_76 = l_Lean_Parser_registerAlias(x_73, x_74, x_75, x_42, x_72); if (lean_obj_tag(x_76) == 0) { lean_object* x_77; lean_object* x_78; lean_object* x_79; x_77 = lean_ctor_get(x_76, 1); lean_inc(x_77); lean_dec(x_76); -x_78 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__74; -x_79 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_70, x_78, x_77); +x_78 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__75; +x_79 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_73, x_78, x_77); if (lean_obj_tag(x_79) == 0) { -lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; +lean_object* x_80; lean_object* x_81; lean_object* x_82; x_80 = lean_ctor_get(x_79, 1); lean_inc(x_80); lean_dec(x_79); -x_81 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__76; -x_82 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__77; -x_83 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__78; -x_84 = l_Lean_Parser_registerAlias(x_81, x_82, x_83, x_80); -if (lean_obj_tag(x_84) == 0) -{ -lean_object* x_85; lean_object* x_86; lean_object* x_87; -x_85 = lean_ctor_get(x_84, 1); -lean_inc(x_85); -lean_dec(x_84); -x_86 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__80; -x_87 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_81, x_86, x_85); +x_81 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__77; +x_82 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_73, x_81, x_80); +if (lean_obj_tag(x_82) == 0) +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_83 = lean_ctor_get(x_82, 1); +lean_inc(x_83); +lean_dec(x_82); +x_84 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__79; +x_85 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__80; +x_86 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__81; +x_87 = l_Lean_Parser_registerAlias(x_84, x_85, x_86, x_42, x_83); if (lean_obj_tag(x_87) == 0) { lean_object* x_88; lean_object* x_89; lean_object* x_90; x_88 = lean_ctor_get(x_87, 1); lean_inc(x_88); lean_dec(x_87); -x_89 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__82; -x_90 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_81, x_89, x_88); +x_89 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__83; +x_90 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_84, x_89, x_88); if (lean_obj_tag(x_90) == 0) { -lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; +lean_object* x_91; lean_object* x_92; lean_object* x_93; x_91 = lean_ctor_get(x_90, 1); lean_inc(x_91); lean_dec(x_90); -x_92 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__84; -x_93 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__85; -x_94 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__86; -x_95 = l_Lean_Parser_registerAlias(x_92, x_93, x_94, x_91); -if (lean_obj_tag(x_95) == 0) -{ -lean_object* x_96; lean_object* x_97; lean_object* x_98; -x_96 = lean_ctor_get(x_95, 1); -lean_inc(x_96); -lean_dec(x_95); -x_97 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__88; -x_98 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_92, x_97, x_96); +x_92 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__85; +x_93 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_84, x_92, x_91); +if (lean_obj_tag(x_93) == 0) +{ +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_94 = lean_ctor_get(x_93, 1); +lean_inc(x_94); +lean_dec(x_93); +x_95 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__87; +x_96 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__88; +x_97 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__89; +x_98 = l_Lean_Parser_registerAlias(x_95, x_96, x_97, x_42, x_94); if (lean_obj_tag(x_98) == 0) { lean_object* x_99; lean_object* x_100; lean_object* x_101; x_99 = lean_ctor_get(x_98, 1); lean_inc(x_99); lean_dec(x_98); -x_100 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__90; -x_101 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_92, x_100, x_99); +x_100 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__91; +x_101 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_95, x_100, x_99); if (lean_obj_tag(x_101) == 0) { -lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +lean_object* x_102; lean_object* x_103; lean_object* x_104; x_102 = lean_ctor_get(x_101, 1); lean_inc(x_102); lean_dec(x_101); -x_103 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__92; -x_104 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__98; -x_105 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__100; -x_106 = l_Lean_Parser_registerAlias(x_103, x_104, x_105, x_102); -if (lean_obj_tag(x_106) == 0) +x_103 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__93; +x_104 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_95, x_103, x_102); +if (lean_obj_tag(x_104) == 0) { -lean_object* x_107; lean_object* x_108; lean_object* x_109; -x_107 = lean_ctor_get(x_106, 1); -lean_inc(x_107); -lean_dec(x_106); -x_108 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__102; -x_109 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_103, x_108, x_107); +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; +x_105 = lean_ctor_get(x_104, 1); +lean_inc(x_105); +lean_dec(x_104); +x_106 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__95; +x_107 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__101; +x_108 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__103; +x_109 = l_Lean_Parser_registerAlias(x_106, x_107, x_108, x_6, x_105); if (lean_obj_tag(x_109) == 0) { lean_object* x_110; lean_object* x_111; lean_object* x_112; x_110 = lean_ctor_get(x_109, 1); lean_inc(x_110); lean_dec(x_109); -x_111 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__104; -x_112 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_103, x_111, x_110); +x_111 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__105; +x_112 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_106, x_111, x_110); if (lean_obj_tag(x_112) == 0) { -lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; +lean_object* x_113; lean_object* x_114; lean_object* x_115; x_113 = lean_ctor_get(x_112, 1); lean_inc(x_113); lean_dec(x_112); -x_114 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__106; -x_115 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__110; -x_116 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__112; -x_117 = l_Lean_Parser_registerAlias(x_114, x_115, x_116, x_113); -if (lean_obj_tag(x_117) == 0) -{ -lean_object* x_118; lean_object* x_119; lean_object* x_120; -x_118 = lean_ctor_get(x_117, 1); -lean_inc(x_118); -lean_dec(x_117); -x_119 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__114; -x_120 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_114, x_119, x_118); +x_114 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__107; +x_115 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_106, x_114, x_113); +if (lean_obj_tag(x_115) == 0) +{ +lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; +x_116 = lean_ctor_get(x_115, 1); +lean_inc(x_116); +lean_dec(x_115); +x_117 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__109; +x_118 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__113; +x_119 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__115; +x_120 = l_Lean_Parser_registerAlias(x_117, x_118, x_119, x_6, x_116); if (lean_obj_tag(x_120) == 0) { lean_object* x_121; lean_object* x_122; lean_object* x_123; x_121 = lean_ctor_get(x_120, 1); lean_inc(x_121); lean_dec(x_120); -x_122 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__116; -x_123 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_114, x_122, x_121); +x_122 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__117; +x_123 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_117, x_122, x_121); if (lean_obj_tag(x_123) == 0) { -lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; +lean_object* x_124; lean_object* x_125; lean_object* x_126; x_124 = lean_ctor_get(x_123, 1); lean_inc(x_124); lean_dec(x_123); -x_125 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__118; -x_126 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__120; -x_127 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__122; -x_128 = l_Lean_Parser_registerAlias(x_125, x_126, x_127, x_124); -if (lean_obj_tag(x_128) == 0) -{ -lean_object* x_129; lean_object* x_130; lean_object* x_131; -x_129 = lean_ctor_get(x_128, 1); -lean_inc(x_129); -lean_dec(x_128); -x_130 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__124; -x_131 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_125, x_130, x_129); -if (lean_obj_tag(x_131) == 0) +x_125 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__119; +x_126 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_117, x_125, x_124); +if (lean_obj_tag(x_126) == 0) { -lean_object* x_132; lean_object* x_133; lean_object* x_134; -x_132 = lean_ctor_get(x_131, 1); -lean_inc(x_132); -lean_dec(x_131); -x_133 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__126; -x_134 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_125, x_133, x_132); -if (lean_obj_tag(x_134) == 0) -{ -lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; -x_135 = lean_ctor_get(x_134, 1); -lean_inc(x_135); -lean_dec(x_134); -x_136 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__128; -x_137 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__130; -x_138 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__132; -x_139 = l_Lean_Parser_registerAlias(x_136, x_137, x_138, x_135); -if (lean_obj_tag(x_139) == 0) -{ -lean_object* x_140; lean_object* x_141; lean_object* x_142; -x_140 = lean_ctor_get(x_139, 1); -lean_inc(x_140); -lean_dec(x_139); -x_141 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__134; -x_142 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_136, x_141, x_140); -if (lean_obj_tag(x_142) == 0) -{ -lean_object* x_143; lean_object* x_144; lean_object* x_145; -x_143 = lean_ctor_get(x_142, 1); -lean_inc(x_143); -lean_dec(x_142); -x_144 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__136; -x_145 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_136, x_144, x_143); -if (lean_obj_tag(x_145) == 0) -{ -lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; -x_146 = lean_ctor_get(x_145, 1); -lean_inc(x_146); -lean_dec(x_145); -x_147 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__138; -x_148 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__140; -x_149 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__142; -x_150 = l_Lean_Parser_registerAlias(x_147, x_148, x_149, x_146); -if (lean_obj_tag(x_150) == 0) -{ -lean_object* x_151; lean_object* x_152; lean_object* x_153; -x_151 = lean_ctor_get(x_150, 1); -lean_inc(x_151); -lean_dec(x_150); -x_152 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__144; -x_153 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_147, x_152, x_151); -if (lean_obj_tag(x_153) == 0) -{ -lean_object* x_154; lean_object* x_155; lean_object* x_156; -x_154 = lean_ctor_get(x_153, 1); -lean_inc(x_154); -lean_dec(x_153); -x_155 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__146; -x_156 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_147, x_155, x_154); -if (lean_obj_tag(x_156) == 0) -{ -lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; -x_157 = lean_ctor_get(x_156, 1); -lean_inc(x_157); -lean_dec(x_156); -x_158 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__148; -x_159 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__150; -x_160 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__152; -x_161 = l_Lean_Parser_registerAlias(x_158, x_159, x_160, x_157); -if (lean_obj_tag(x_161) == 0) -{ -lean_object* x_162; lean_object* x_163; lean_object* x_164; -x_162 = lean_ctor_get(x_161, 1); -lean_inc(x_162); -lean_dec(x_161); -x_163 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__154; -x_164 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_158, x_163, x_162); -if (lean_obj_tag(x_164) == 0) +lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; +x_127 = lean_ctor_get(x_126, 1); +lean_inc(x_127); +lean_dec(x_126); +x_128 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__121; +x_129 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__123; +x_130 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__125; +x_131 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__127; +x_132 = l_Lean_Parser_registerAlias(x_128, x_129, x_130, x_131, x_127); +if (lean_obj_tag(x_132) == 0) +{ +lean_object* x_133; lean_object* x_134; lean_object* x_135; +x_133 = lean_ctor_get(x_132, 1); +lean_inc(x_133); +lean_dec(x_132); +x_134 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__129; +x_135 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_128, x_134, x_133); +if (lean_obj_tag(x_135) == 0) +{ +lean_object* x_136; lean_object* x_137; lean_object* x_138; +x_136 = lean_ctor_get(x_135, 1); +lean_inc(x_136); +lean_dec(x_135); +x_137 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__131; +x_138 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_128, x_137, x_136); +if (lean_obj_tag(x_138) == 0) +{ +lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; +x_139 = lean_ctor_get(x_138, 1); +lean_inc(x_139); +lean_dec(x_138); +x_140 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__133; +x_141 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__135; +x_142 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__137; +x_143 = l_Lean_Parser_registerAlias(x_140, x_141, x_142, x_6, x_139); +if (lean_obj_tag(x_143) == 0) +{ +lean_object* x_144; lean_object* x_145; lean_object* x_146; +x_144 = lean_ctor_get(x_143, 1); +lean_inc(x_144); +lean_dec(x_143); +x_145 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__139; +x_146 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_140, x_145, x_144); +if (lean_obj_tag(x_146) == 0) { -lean_object* x_165; lean_object* x_166; lean_object* x_167; -x_165 = lean_ctor_get(x_164, 1); -lean_inc(x_165); -lean_dec(x_164); -x_166 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__156; -x_167 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_158, x_166, x_165); -if (lean_obj_tag(x_167) == 0) +lean_object* x_147; lean_object* x_148; lean_object* x_149; +x_147 = lean_ctor_get(x_146, 1); +lean_inc(x_147); +lean_dec(x_146); +x_148 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__141; +x_149 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_140, x_148, x_147); +if (lean_obj_tag(x_149) == 0) { -lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; -x_168 = lean_ctor_get(x_167, 1); -lean_inc(x_168); -lean_dec(x_167); -x_169 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__158; -x_170 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__160; -x_171 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__162; -x_172 = l_Lean_Parser_registerAlias(x_169, x_170, x_171, x_168); -if (lean_obj_tag(x_172) == 0) +lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; +x_150 = lean_ctor_get(x_149, 1); +lean_inc(x_150); +lean_dec(x_149); +x_151 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__143; +x_152 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__145; +x_153 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__147; +x_154 = l_Lean_Parser_registerAlias(x_151, x_152, x_153, x_42, x_150); +if (lean_obj_tag(x_154) == 0) +{ +lean_object* x_155; lean_object* x_156; lean_object* x_157; +x_155 = lean_ctor_get(x_154, 1); +lean_inc(x_155); +lean_dec(x_154); +x_156 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__149; +x_157 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_151, x_156, x_155); +if (lean_obj_tag(x_157) == 0) +{ +lean_object* x_158; lean_object* x_159; lean_object* x_160; +x_158 = lean_ctor_get(x_157, 1); +lean_inc(x_158); +lean_dec(x_157); +x_159 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__151; +x_160 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_151, x_159, x_158); +if (lean_obj_tag(x_160) == 0) { -lean_object* x_173; lean_object* x_174; lean_object* x_175; -x_173 = lean_ctor_get(x_172, 1); -lean_inc(x_173); -lean_dec(x_172); -x_174 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__164; -x_175 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_169, x_174, x_173); -if (lean_obj_tag(x_175) == 0) +lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; +x_161 = lean_ctor_get(x_160, 1); +lean_inc(x_161); +lean_dec(x_160); +x_162 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__153; +x_163 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__155; +x_164 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__157; +x_165 = l_Lean_Parser_registerAlias(x_162, x_163, x_164, x_42, x_161); +if (lean_obj_tag(x_165) == 0) +{ +lean_object* x_166; lean_object* x_167; lean_object* x_168; +x_166 = lean_ctor_get(x_165, 1); +lean_inc(x_166); +lean_dec(x_165); +x_167 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__159; +x_168 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_162, x_167, x_166); +if (lean_obj_tag(x_168) == 0) { -lean_object* x_176; lean_object* x_177; lean_object* x_178; -x_176 = lean_ctor_get(x_175, 1); -lean_inc(x_176); -lean_dec(x_175); -x_177 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__166; -x_178 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_169, x_177, x_176); -if (lean_obj_tag(x_178) == 0) +lean_object* x_169; lean_object* x_170; lean_object* x_171; +x_169 = lean_ctor_get(x_168, 1); +lean_inc(x_169); +lean_dec(x_168); +x_170 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__161; +x_171 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_162, x_170, x_169); +if (lean_obj_tag(x_171) == 0) { -lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; -x_179 = lean_ctor_get(x_178, 1); -lean_inc(x_179); -lean_dec(x_178); -x_180 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__168; -x_181 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__170; -x_182 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__172; -x_183 = l_Lean_Parser_registerAlias(x_180, x_181, x_182, x_179); -if (lean_obj_tag(x_183) == 0) -{ -lean_object* x_184; lean_object* x_185; lean_object* x_186; -x_184 = lean_ctor_get(x_183, 1); -lean_inc(x_184); -lean_dec(x_183); -x_185 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__174; -x_186 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_180, x_185, x_184); -if (lean_obj_tag(x_186) == 0) -{ -lean_object* x_187; lean_object* x_188; lean_object* x_189; -x_187 = lean_ctor_get(x_186, 1); -lean_inc(x_187); -lean_dec(x_186); -x_188 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__176; -x_189 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_180, x_188, x_187); -if (lean_obj_tag(x_189) == 0) +lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; +x_172 = lean_ctor_get(x_171, 1); +lean_inc(x_172); +lean_dec(x_171); +x_173 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__163; +x_174 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__165; +x_175 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__167; +x_176 = l_Lean_Parser_registerAlias(x_173, x_174, x_175, x_42, x_172); +if (lean_obj_tag(x_176) == 0) +{ +lean_object* x_177; lean_object* x_178; lean_object* x_179; +x_177 = lean_ctor_get(x_176, 1); +lean_inc(x_177); +lean_dec(x_176); +x_178 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__169; +x_179 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_173, x_178, x_177); +if (lean_obj_tag(x_179) == 0) +{ +lean_object* x_180; lean_object* x_181; lean_object* x_182; +x_180 = lean_ctor_get(x_179, 1); +lean_inc(x_180); +lean_dec(x_179); +x_181 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__171; +x_182 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_173, x_181, x_180); +if (lean_obj_tag(x_182) == 0) +{ +lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; +x_183 = lean_ctor_get(x_182, 1); +lean_inc(x_183); +lean_dec(x_182); +x_184 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__173; +x_185 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__175; +x_186 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__177; +x_187 = l_Lean_Parser_registerAlias(x_184, x_185, x_186, x_42, x_183); +if (lean_obj_tag(x_187) == 0) +{ +lean_object* x_188; lean_object* x_189; lean_object* x_190; +x_188 = lean_ctor_get(x_187, 1); +lean_inc(x_188); +lean_dec(x_187); +x_189 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__179; +x_190 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_184, x_189, x_188); +if (lean_obj_tag(x_190) == 0) { -lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; -x_190 = lean_ctor_get(x_189, 1); -lean_inc(x_190); -lean_dec(x_189); -x_191 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__178; -x_192 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__180; -x_193 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__182; -x_194 = l_Lean_Parser_registerAlias(x_191, x_192, x_193, x_190); -if (lean_obj_tag(x_194) == 0) -{ -lean_object* x_195; lean_object* x_196; lean_object* x_197; -x_195 = lean_ctor_get(x_194, 1); -lean_inc(x_195); -lean_dec(x_194); -x_196 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__184; -x_197 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_191, x_196, x_195); -if (lean_obj_tag(x_197) == 0) +lean_object* x_191; lean_object* x_192; lean_object* x_193; +x_191 = lean_ctor_get(x_190, 1); +lean_inc(x_191); +lean_dec(x_190); +x_192 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__181; +x_193 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_184, x_192, x_191); +if (lean_obj_tag(x_193) == 0) { -lean_object* x_198; lean_object* x_199; lean_object* x_200; -x_198 = lean_ctor_get(x_197, 1); -lean_inc(x_198); -lean_dec(x_197); -x_199 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__186; -x_200 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_191, x_199, x_198); -if (lean_obj_tag(x_200) == 0) +lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; +x_194 = lean_ctor_get(x_193, 1); +lean_inc(x_194); +lean_dec(x_193); +x_195 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__183; +x_196 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__185; +x_197 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__187; +x_198 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__188; +x_199 = l_Lean_Parser_registerAlias(x_195, x_196, x_197, x_198, x_194); +if (lean_obj_tag(x_199) == 0) +{ +lean_object* x_200; lean_object* x_201; lean_object* x_202; +x_200 = lean_ctor_get(x_199, 1); +lean_inc(x_200); +lean_dec(x_199); +x_201 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__190; +x_202 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_195, x_201, x_200); +if (lean_obj_tag(x_202) == 0) { -lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; -x_201 = lean_ctor_get(x_200, 1); -lean_inc(x_201); -lean_dec(x_200); -x_202 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__188; -x_203 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__190; +lean_object* x_203; lean_object* x_204; lean_object* x_205; +x_203 = lean_ctor_get(x_202, 1); +lean_inc(x_203); +lean_dec(x_202); x_204 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__192; -x_205 = l_Lean_Parser_registerAlias(x_202, x_203, x_204, x_201); +x_205 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_195, x_204, x_203); if (lean_obj_tag(x_205) == 0) { -lean_object* x_206; lean_object* x_207; lean_object* x_208; +lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; x_206 = lean_ctor_get(x_205, 1); lean_inc(x_206); lean_dec(x_205); x_207 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__194; -x_208 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_202, x_207, x_206); -if (lean_obj_tag(x_208) == 0) -{ -lean_object* x_209; lean_object* x_210; lean_object* x_211; -x_209 = lean_ctor_get(x_208, 1); -lean_inc(x_209); -lean_dec(x_208); -x_210 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__196; -x_211 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_202, x_210, x_209); -if (lean_obj_tag(x_211) == 0) -{ -lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; -x_212 = lean_ctor_get(x_211, 1); -lean_inc(x_212); -lean_dec(x_211); -x_213 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__198; -x_214 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__200; -x_215 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__201; -x_216 = l_Lean_Parser_registerAlias(x_213, x_214, x_215, x_212); +x_208 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__196; +x_209 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__198; +x_210 = l_Lean_Parser_registerAlias(x_207, x_208, x_209, x_6, x_206); +if (lean_obj_tag(x_210) == 0) +{ +lean_object* x_211; lean_object* x_212; lean_object* x_213; +x_211 = lean_ctor_get(x_210, 1); +lean_inc(x_211); +lean_dec(x_210); +x_212 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__200; +x_213 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_207, x_212, x_211); +if (lean_obj_tag(x_213) == 0) +{ +lean_object* x_214; lean_object* x_215; lean_object* x_216; +x_214 = lean_ctor_get(x_213, 1); +lean_inc(x_214); +lean_dec(x_213); +x_215 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__202; +x_216 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_207, x_215, x_214); if (lean_obj_tag(x_216) == 0) { -lean_object* x_217; lean_object* x_218; lean_object* x_219; +lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; x_217 = lean_ctor_get(x_216, 1); lean_inc(x_217); lean_dec(x_216); -x_218 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__203; -x_219 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_213, x_218, x_217); -if (lean_obj_tag(x_219) == 0) -{ -lean_object* x_220; lean_object* x_221; lean_object* x_222; -x_220 = lean_ctor_get(x_219, 1); -lean_inc(x_220); -lean_dec(x_219); -x_221 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__205; -x_222 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_213, x_221, x_220); -if (lean_obj_tag(x_222) == 0) -{ -lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; -x_223 = lean_ctor_get(x_222, 1); -lean_inc(x_223); -lean_dec(x_222); -x_224 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__207; -x_225 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__209; -x_226 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__211; -x_227 = l_Lean_Parser_registerAlias(x_224, x_225, x_226, x_223); +x_218 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__204; +x_219 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__206; +x_220 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__209; +x_221 = l_Lean_Parser_registerAlias(x_218, x_219, x_220, x_42, x_217); +if (lean_obj_tag(x_221) == 0) +{ +lean_object* x_222; lean_object* x_223; lean_object* x_224; +x_222 = lean_ctor_get(x_221, 1); +lean_inc(x_222); +lean_dec(x_221); +x_223 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__211; +x_224 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_218, x_223, x_222); +if (lean_obj_tag(x_224) == 0) +{ +lean_object* x_225; lean_object* x_226; lean_object* x_227; +x_225 = lean_ctor_get(x_224, 1); +lean_inc(x_225); +lean_dec(x_224); +x_226 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__213; +x_227 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_218, x_226, x_225); if (lean_obj_tag(x_227) == 0) { -lean_object* x_228; lean_object* x_229; lean_object* x_230; +lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; x_228 = lean_ctor_get(x_227, 1); lean_inc(x_228); lean_dec(x_227); -x_229 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__213; -x_230 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_224, x_229, x_228); -if (lean_obj_tag(x_230) == 0) -{ -lean_object* x_231; lean_object* x_232; lean_object* x_233; -x_231 = lean_ctor_get(x_230, 1); -lean_inc(x_231); -lean_dec(x_230); -x_232 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__215; -x_233 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_224, x_232, x_231); -if (lean_obj_tag(x_233) == 0) -{ -lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; -x_234 = lean_ctor_get(x_233, 1); -lean_inc(x_234); -lean_dec(x_233); -x_235 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__217; -x_236 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__219; -x_237 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__221; -x_238 = l_Lean_Parser_registerAlias(x_235, x_236, x_237, x_234); +x_229 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__215; +x_230 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__217; +x_231 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__219; +x_232 = l_Lean_Parser_registerAlias(x_229, x_230, x_231, x_42, x_228); +if (lean_obj_tag(x_232) == 0) +{ +lean_object* x_233; lean_object* x_234; lean_object* x_235; +x_233 = lean_ctor_get(x_232, 1); +lean_inc(x_233); +lean_dec(x_232); +x_234 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__221; +x_235 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_229, x_234, x_233); +if (lean_obj_tag(x_235) == 0) +{ +lean_object* x_236; lean_object* x_237; lean_object* x_238; +x_236 = lean_ctor_get(x_235, 1); +lean_inc(x_236); +lean_dec(x_235); +x_237 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__223; +x_238 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_229, x_237, x_236); if (lean_obj_tag(x_238) == 0) { -lean_object* x_239; lean_object* x_240; lean_object* x_241; +lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; x_239 = lean_ctor_get(x_238, 1); lean_inc(x_239); lean_dec(x_238); -x_240 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__223; -x_241 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_235, x_240, x_239); -if (lean_obj_tag(x_241) == 0) -{ -lean_object* x_242; lean_object* x_243; lean_object* x_244; -x_242 = lean_ctor_get(x_241, 1); -lean_inc(x_242); -lean_dec(x_241); -x_243 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__225; -x_244 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_235, x_243, x_242); -if (lean_obj_tag(x_244) == 0) -{ -lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; -x_245 = lean_ctor_get(x_244, 1); -lean_inc(x_245); -lean_dec(x_244); -x_246 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__94; -x_247 = lean_alloc_closure((void*)(l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____lambda__1), 2, 1); -lean_closure_set(x_247, 0, x_246); -x_248 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_248, 0, x_247); -x_249 = lean_box(0); -x_250 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__227; -x_251 = l_Lean_Parser_registerAlias(x_250, x_248, x_249, x_245); -if (lean_obj_tag(x_251) == 0) -{ -lean_object* x_252; lean_object* x_253; lean_object* x_254; -x_252 = lean_ctor_get(x_251, 1); -lean_inc(x_252); -lean_dec(x_251); -x_253 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__229; -x_254 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_250, x_253, x_252); -if (lean_obj_tag(x_254) == 0) -{ -lean_object* x_255; lean_object* x_256; lean_object* x_257; -x_255 = lean_ctor_get(x_254, 1); -lean_inc(x_255); -lean_dec(x_254); -x_256 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__231; -x_257 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_250, x_256, x_255); -return x_257; -} -else -{ -uint8_t x_258; -x_258 = !lean_is_exclusive(x_254); -if (x_258 == 0) -{ -return x_254; -} -else +x_240 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__225; +x_241 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__227; +x_242 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__229; +x_243 = l_Lean_Parser_registerAlias(x_240, x_241, x_242, x_6, x_239); +if (lean_obj_tag(x_243) == 0) +{ +lean_object* x_244; lean_object* x_245; lean_object* x_246; +x_244 = lean_ctor_get(x_243, 1); +lean_inc(x_244); +lean_dec(x_243); +x_245 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__231; +x_246 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_240, x_245, x_244); +if (lean_obj_tag(x_246) == 0) +{ +lean_object* x_247; lean_object* x_248; lean_object* x_249; +x_247 = lean_ctor_get(x_246, 1); +lean_inc(x_247); +lean_dec(x_246); +x_248 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__233; +x_249 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_240, x_248, x_247); +if (lean_obj_tag(x_249) == 0) +{ +lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; +x_250 = lean_ctor_get(x_249, 1); +lean_inc(x_250); +lean_dec(x_249); +x_251 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__97; +x_252 = lean_alloc_closure((void*)(l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____lambda__1), 2, 1); +lean_closure_set(x_252, 0, x_251); +x_253 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_253, 0, x_252); +x_254 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__235; +x_255 = l_Lean_Parser_registerAlias(x_254, x_253, x_2, x_42, x_250); +if (lean_obj_tag(x_255) == 0) +{ +lean_object* x_256; lean_object* x_257; lean_object* x_258; +x_256 = lean_ctor_get(x_255, 1); +lean_inc(x_256); +lean_dec(x_255); +x_257 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__237; +x_258 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_254, x_257, x_256); +if (lean_obj_tag(x_258) == 0) { lean_object* x_259; lean_object* x_260; lean_object* x_261; -x_259 = lean_ctor_get(x_254, 0); -x_260 = lean_ctor_get(x_254, 1); -lean_inc(x_260); +x_259 = lean_ctor_get(x_258, 1); lean_inc(x_259); -lean_dec(x_254); -x_261 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_261, 0, x_259); -lean_ctor_set(x_261, 1, x_260); +lean_dec(x_258); +x_260 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__239; +x_261 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_254, x_260, x_259); return x_261; } -} -} else { uint8_t x_262; -x_262 = !lean_is_exclusive(x_251); +x_262 = !lean_is_exclusive(x_258); if (x_262 == 0) { -return x_251; +return x_258; } else { lean_object* x_263; lean_object* x_264; lean_object* x_265; -x_263 = lean_ctor_get(x_251, 0); -x_264 = lean_ctor_get(x_251, 1); +x_263 = lean_ctor_get(x_258, 0); +x_264 = lean_ctor_get(x_258, 1); lean_inc(x_264); lean_inc(x_263); -lean_dec(x_251); +lean_dec(x_258); x_265 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_265, 0, x_263); lean_ctor_set(x_265, 1, x_264); @@ -3269,19 +3337,19 @@ return x_265; else { uint8_t x_266; -x_266 = !lean_is_exclusive(x_244); +x_266 = !lean_is_exclusive(x_255); if (x_266 == 0) { -return x_244; +return x_255; } else { lean_object* x_267; lean_object* x_268; lean_object* x_269; -x_267 = lean_ctor_get(x_244, 0); -x_268 = lean_ctor_get(x_244, 1); +x_267 = lean_ctor_get(x_255, 0); +x_268 = lean_ctor_get(x_255, 1); lean_inc(x_268); lean_inc(x_267); -lean_dec(x_244); +lean_dec(x_255); x_269 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_269, 0, x_267); lean_ctor_set(x_269, 1, x_268); @@ -3292,19 +3360,19 @@ return x_269; else { uint8_t x_270; -x_270 = !lean_is_exclusive(x_241); +x_270 = !lean_is_exclusive(x_249); if (x_270 == 0) { -return x_241; +return x_249; } else { lean_object* x_271; lean_object* x_272; lean_object* x_273; -x_271 = lean_ctor_get(x_241, 0); -x_272 = lean_ctor_get(x_241, 1); +x_271 = lean_ctor_get(x_249, 0); +x_272 = lean_ctor_get(x_249, 1); lean_inc(x_272); lean_inc(x_271); -lean_dec(x_241); +lean_dec(x_249); x_273 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_273, 0, x_271); lean_ctor_set(x_273, 1, x_272); @@ -3315,19 +3383,19 @@ return x_273; else { uint8_t x_274; -x_274 = !lean_is_exclusive(x_238); +x_274 = !lean_is_exclusive(x_246); if (x_274 == 0) { -return x_238; +return x_246; } else { lean_object* x_275; lean_object* x_276; lean_object* x_277; -x_275 = lean_ctor_get(x_238, 0); -x_276 = lean_ctor_get(x_238, 1); +x_275 = lean_ctor_get(x_246, 0); +x_276 = lean_ctor_get(x_246, 1); lean_inc(x_276); lean_inc(x_275); -lean_dec(x_238); +lean_dec(x_246); x_277 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_277, 0, x_275); lean_ctor_set(x_277, 1, x_276); @@ -3338,19 +3406,19 @@ return x_277; else { uint8_t x_278; -x_278 = !lean_is_exclusive(x_233); +x_278 = !lean_is_exclusive(x_243); if (x_278 == 0) { -return x_233; +return x_243; } else { lean_object* x_279; lean_object* x_280; lean_object* x_281; -x_279 = lean_ctor_get(x_233, 0); -x_280 = lean_ctor_get(x_233, 1); +x_279 = lean_ctor_get(x_243, 0); +x_280 = lean_ctor_get(x_243, 1); lean_inc(x_280); lean_inc(x_279); -lean_dec(x_233); +lean_dec(x_243); x_281 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_281, 0, x_279); lean_ctor_set(x_281, 1, x_280); @@ -3361,19 +3429,19 @@ return x_281; else { uint8_t x_282; -x_282 = !lean_is_exclusive(x_230); +x_282 = !lean_is_exclusive(x_238); if (x_282 == 0) { -return x_230; +return x_238; } else { lean_object* x_283; lean_object* x_284; lean_object* x_285; -x_283 = lean_ctor_get(x_230, 0); -x_284 = lean_ctor_get(x_230, 1); +x_283 = lean_ctor_get(x_238, 0); +x_284 = lean_ctor_get(x_238, 1); lean_inc(x_284); lean_inc(x_283); -lean_dec(x_230); +lean_dec(x_238); x_285 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_285, 0, x_283); lean_ctor_set(x_285, 1, x_284); @@ -3384,19 +3452,19 @@ return x_285; else { uint8_t x_286; -x_286 = !lean_is_exclusive(x_227); +x_286 = !lean_is_exclusive(x_235); if (x_286 == 0) { -return x_227; +return x_235; } else { lean_object* x_287; lean_object* x_288; lean_object* x_289; -x_287 = lean_ctor_get(x_227, 0); -x_288 = lean_ctor_get(x_227, 1); +x_287 = lean_ctor_get(x_235, 0); +x_288 = lean_ctor_get(x_235, 1); lean_inc(x_288); lean_inc(x_287); -lean_dec(x_227); +lean_dec(x_235); x_289 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_289, 0, x_287); lean_ctor_set(x_289, 1, x_288); @@ -3407,19 +3475,19 @@ return x_289; else { uint8_t x_290; -x_290 = !lean_is_exclusive(x_222); +x_290 = !lean_is_exclusive(x_232); if (x_290 == 0) { -return x_222; +return x_232; } else { lean_object* x_291; lean_object* x_292; lean_object* x_293; -x_291 = lean_ctor_get(x_222, 0); -x_292 = lean_ctor_get(x_222, 1); +x_291 = lean_ctor_get(x_232, 0); +x_292 = lean_ctor_get(x_232, 1); lean_inc(x_292); lean_inc(x_291); -lean_dec(x_222); +lean_dec(x_232); x_293 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_293, 0, x_291); lean_ctor_set(x_293, 1, x_292); @@ -3430,19 +3498,19 @@ return x_293; else { uint8_t x_294; -x_294 = !lean_is_exclusive(x_219); +x_294 = !lean_is_exclusive(x_227); if (x_294 == 0) { -return x_219; +return x_227; } else { lean_object* x_295; lean_object* x_296; lean_object* x_297; -x_295 = lean_ctor_get(x_219, 0); -x_296 = lean_ctor_get(x_219, 1); +x_295 = lean_ctor_get(x_227, 0); +x_296 = lean_ctor_get(x_227, 1); lean_inc(x_296); lean_inc(x_295); -lean_dec(x_219); +lean_dec(x_227); x_297 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_297, 0, x_295); lean_ctor_set(x_297, 1, x_296); @@ -3453,19 +3521,19 @@ return x_297; else { uint8_t x_298; -x_298 = !lean_is_exclusive(x_216); +x_298 = !lean_is_exclusive(x_224); if (x_298 == 0) { -return x_216; +return x_224; } else { lean_object* x_299; lean_object* x_300; lean_object* x_301; -x_299 = lean_ctor_get(x_216, 0); -x_300 = lean_ctor_get(x_216, 1); +x_299 = lean_ctor_get(x_224, 0); +x_300 = lean_ctor_get(x_224, 1); lean_inc(x_300); lean_inc(x_299); -lean_dec(x_216); +lean_dec(x_224); x_301 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_301, 0, x_299); lean_ctor_set(x_301, 1, x_300); @@ -3476,19 +3544,19 @@ return x_301; else { uint8_t x_302; -x_302 = !lean_is_exclusive(x_211); +x_302 = !lean_is_exclusive(x_221); if (x_302 == 0) { -return x_211; +return x_221; } else { lean_object* x_303; lean_object* x_304; lean_object* x_305; -x_303 = lean_ctor_get(x_211, 0); -x_304 = lean_ctor_get(x_211, 1); +x_303 = lean_ctor_get(x_221, 0); +x_304 = lean_ctor_get(x_221, 1); lean_inc(x_304); lean_inc(x_303); -lean_dec(x_211); +lean_dec(x_221); x_305 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_305, 0, x_303); lean_ctor_set(x_305, 1, x_304); @@ -3499,19 +3567,19 @@ return x_305; else { uint8_t x_306; -x_306 = !lean_is_exclusive(x_208); +x_306 = !lean_is_exclusive(x_216); if (x_306 == 0) { -return x_208; +return x_216; } else { lean_object* x_307; lean_object* x_308; lean_object* x_309; -x_307 = lean_ctor_get(x_208, 0); -x_308 = lean_ctor_get(x_208, 1); +x_307 = lean_ctor_get(x_216, 0); +x_308 = lean_ctor_get(x_216, 1); lean_inc(x_308); lean_inc(x_307); -lean_dec(x_208); +lean_dec(x_216); x_309 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_309, 0, x_307); lean_ctor_set(x_309, 1, x_308); @@ -3522,19 +3590,19 @@ return x_309; else { uint8_t x_310; -x_310 = !lean_is_exclusive(x_205); +x_310 = !lean_is_exclusive(x_213); if (x_310 == 0) { -return x_205; +return x_213; } else { lean_object* x_311; lean_object* x_312; lean_object* x_313; -x_311 = lean_ctor_get(x_205, 0); -x_312 = lean_ctor_get(x_205, 1); +x_311 = lean_ctor_get(x_213, 0); +x_312 = lean_ctor_get(x_213, 1); lean_inc(x_312); lean_inc(x_311); -lean_dec(x_205); +lean_dec(x_213); x_313 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_313, 0, x_311); lean_ctor_set(x_313, 1, x_312); @@ -3545,19 +3613,19 @@ return x_313; else { uint8_t x_314; -x_314 = !lean_is_exclusive(x_200); +x_314 = !lean_is_exclusive(x_210); if (x_314 == 0) { -return x_200; +return x_210; } else { lean_object* x_315; lean_object* x_316; lean_object* x_317; -x_315 = lean_ctor_get(x_200, 0); -x_316 = lean_ctor_get(x_200, 1); +x_315 = lean_ctor_get(x_210, 0); +x_316 = lean_ctor_get(x_210, 1); lean_inc(x_316); lean_inc(x_315); -lean_dec(x_200); +lean_dec(x_210); x_317 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_317, 0, x_315); lean_ctor_set(x_317, 1, x_316); @@ -3568,19 +3636,19 @@ return x_317; else { uint8_t x_318; -x_318 = !lean_is_exclusive(x_197); +x_318 = !lean_is_exclusive(x_205); if (x_318 == 0) { -return x_197; +return x_205; } else { lean_object* x_319; lean_object* x_320; lean_object* x_321; -x_319 = lean_ctor_get(x_197, 0); -x_320 = lean_ctor_get(x_197, 1); +x_319 = lean_ctor_get(x_205, 0); +x_320 = lean_ctor_get(x_205, 1); lean_inc(x_320); lean_inc(x_319); -lean_dec(x_197); +lean_dec(x_205); x_321 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_321, 0, x_319); lean_ctor_set(x_321, 1, x_320); @@ -3591,19 +3659,19 @@ return x_321; else { uint8_t x_322; -x_322 = !lean_is_exclusive(x_194); +x_322 = !lean_is_exclusive(x_202); if (x_322 == 0) { -return x_194; +return x_202; } else { lean_object* x_323; lean_object* x_324; lean_object* x_325; -x_323 = lean_ctor_get(x_194, 0); -x_324 = lean_ctor_get(x_194, 1); +x_323 = lean_ctor_get(x_202, 0); +x_324 = lean_ctor_get(x_202, 1); lean_inc(x_324); lean_inc(x_323); -lean_dec(x_194); +lean_dec(x_202); x_325 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_325, 0, x_323); lean_ctor_set(x_325, 1, x_324); @@ -3614,19 +3682,19 @@ return x_325; else { uint8_t x_326; -x_326 = !lean_is_exclusive(x_189); +x_326 = !lean_is_exclusive(x_199); if (x_326 == 0) { -return x_189; +return x_199; } else { lean_object* x_327; lean_object* x_328; lean_object* x_329; -x_327 = lean_ctor_get(x_189, 0); -x_328 = lean_ctor_get(x_189, 1); +x_327 = lean_ctor_get(x_199, 0); +x_328 = lean_ctor_get(x_199, 1); lean_inc(x_328); lean_inc(x_327); -lean_dec(x_189); +lean_dec(x_199); x_329 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_329, 0, x_327); lean_ctor_set(x_329, 1, x_328); @@ -3637,19 +3705,19 @@ return x_329; else { uint8_t x_330; -x_330 = !lean_is_exclusive(x_186); +x_330 = !lean_is_exclusive(x_193); if (x_330 == 0) { -return x_186; +return x_193; } else { lean_object* x_331; lean_object* x_332; lean_object* x_333; -x_331 = lean_ctor_get(x_186, 0); -x_332 = lean_ctor_get(x_186, 1); +x_331 = lean_ctor_get(x_193, 0); +x_332 = lean_ctor_get(x_193, 1); lean_inc(x_332); lean_inc(x_331); -lean_dec(x_186); +lean_dec(x_193); x_333 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_333, 0, x_331); lean_ctor_set(x_333, 1, x_332); @@ -3660,19 +3728,19 @@ return x_333; else { uint8_t x_334; -x_334 = !lean_is_exclusive(x_183); +x_334 = !lean_is_exclusive(x_190); if (x_334 == 0) { -return x_183; +return x_190; } else { lean_object* x_335; lean_object* x_336; lean_object* x_337; -x_335 = lean_ctor_get(x_183, 0); -x_336 = lean_ctor_get(x_183, 1); +x_335 = lean_ctor_get(x_190, 0); +x_336 = lean_ctor_get(x_190, 1); lean_inc(x_336); lean_inc(x_335); -lean_dec(x_183); +lean_dec(x_190); x_337 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_337, 0, x_335); lean_ctor_set(x_337, 1, x_336); @@ -3683,19 +3751,19 @@ return x_337; else { uint8_t x_338; -x_338 = !lean_is_exclusive(x_178); +x_338 = !lean_is_exclusive(x_187); if (x_338 == 0) { -return x_178; +return x_187; } else { lean_object* x_339; lean_object* x_340; lean_object* x_341; -x_339 = lean_ctor_get(x_178, 0); -x_340 = lean_ctor_get(x_178, 1); +x_339 = lean_ctor_get(x_187, 0); +x_340 = lean_ctor_get(x_187, 1); lean_inc(x_340); lean_inc(x_339); -lean_dec(x_178); +lean_dec(x_187); x_341 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_341, 0, x_339); lean_ctor_set(x_341, 1, x_340); @@ -3706,19 +3774,19 @@ return x_341; else { uint8_t x_342; -x_342 = !lean_is_exclusive(x_175); +x_342 = !lean_is_exclusive(x_182); if (x_342 == 0) { -return x_175; +return x_182; } else { lean_object* x_343; lean_object* x_344; lean_object* x_345; -x_343 = lean_ctor_get(x_175, 0); -x_344 = lean_ctor_get(x_175, 1); +x_343 = lean_ctor_get(x_182, 0); +x_344 = lean_ctor_get(x_182, 1); lean_inc(x_344); lean_inc(x_343); -lean_dec(x_175); +lean_dec(x_182); x_345 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_345, 0, x_343); lean_ctor_set(x_345, 1, x_344); @@ -3729,19 +3797,19 @@ return x_345; else { uint8_t x_346; -x_346 = !lean_is_exclusive(x_172); +x_346 = !lean_is_exclusive(x_179); if (x_346 == 0) { -return x_172; +return x_179; } else { lean_object* x_347; lean_object* x_348; lean_object* x_349; -x_347 = lean_ctor_get(x_172, 0); -x_348 = lean_ctor_get(x_172, 1); +x_347 = lean_ctor_get(x_179, 0); +x_348 = lean_ctor_get(x_179, 1); lean_inc(x_348); lean_inc(x_347); -lean_dec(x_172); +lean_dec(x_179); x_349 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_349, 0, x_347); lean_ctor_set(x_349, 1, x_348); @@ -3752,19 +3820,19 @@ return x_349; else { uint8_t x_350; -x_350 = !lean_is_exclusive(x_167); +x_350 = !lean_is_exclusive(x_176); if (x_350 == 0) { -return x_167; +return x_176; } else { lean_object* x_351; lean_object* x_352; lean_object* x_353; -x_351 = lean_ctor_get(x_167, 0); -x_352 = lean_ctor_get(x_167, 1); +x_351 = lean_ctor_get(x_176, 0); +x_352 = lean_ctor_get(x_176, 1); lean_inc(x_352); lean_inc(x_351); -lean_dec(x_167); +lean_dec(x_176); x_353 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_353, 0, x_351); lean_ctor_set(x_353, 1, x_352); @@ -3775,19 +3843,19 @@ return x_353; else { uint8_t x_354; -x_354 = !lean_is_exclusive(x_164); +x_354 = !lean_is_exclusive(x_171); if (x_354 == 0) { -return x_164; +return x_171; } else { lean_object* x_355; lean_object* x_356; lean_object* x_357; -x_355 = lean_ctor_get(x_164, 0); -x_356 = lean_ctor_get(x_164, 1); +x_355 = lean_ctor_get(x_171, 0); +x_356 = lean_ctor_get(x_171, 1); lean_inc(x_356); lean_inc(x_355); -lean_dec(x_164); +lean_dec(x_171); x_357 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_357, 0, x_355); lean_ctor_set(x_357, 1, x_356); @@ -3798,19 +3866,19 @@ return x_357; else { uint8_t x_358; -x_358 = !lean_is_exclusive(x_161); +x_358 = !lean_is_exclusive(x_168); if (x_358 == 0) { -return x_161; +return x_168; } else { lean_object* x_359; lean_object* x_360; lean_object* x_361; -x_359 = lean_ctor_get(x_161, 0); -x_360 = lean_ctor_get(x_161, 1); +x_359 = lean_ctor_get(x_168, 0); +x_360 = lean_ctor_get(x_168, 1); lean_inc(x_360); lean_inc(x_359); -lean_dec(x_161); +lean_dec(x_168); x_361 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_361, 0, x_359); lean_ctor_set(x_361, 1, x_360); @@ -3821,19 +3889,19 @@ return x_361; else { uint8_t x_362; -x_362 = !lean_is_exclusive(x_156); +x_362 = !lean_is_exclusive(x_165); if (x_362 == 0) { -return x_156; +return x_165; } else { lean_object* x_363; lean_object* x_364; lean_object* x_365; -x_363 = lean_ctor_get(x_156, 0); -x_364 = lean_ctor_get(x_156, 1); +x_363 = lean_ctor_get(x_165, 0); +x_364 = lean_ctor_get(x_165, 1); lean_inc(x_364); lean_inc(x_363); -lean_dec(x_156); +lean_dec(x_165); x_365 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_365, 0, x_363); lean_ctor_set(x_365, 1, x_364); @@ -3844,19 +3912,19 @@ return x_365; else { uint8_t x_366; -x_366 = !lean_is_exclusive(x_153); +x_366 = !lean_is_exclusive(x_160); if (x_366 == 0) { -return x_153; +return x_160; } else { lean_object* x_367; lean_object* x_368; lean_object* x_369; -x_367 = lean_ctor_get(x_153, 0); -x_368 = lean_ctor_get(x_153, 1); +x_367 = lean_ctor_get(x_160, 0); +x_368 = lean_ctor_get(x_160, 1); lean_inc(x_368); lean_inc(x_367); -lean_dec(x_153); +lean_dec(x_160); x_369 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_369, 0, x_367); lean_ctor_set(x_369, 1, x_368); @@ -3867,19 +3935,19 @@ return x_369; else { uint8_t x_370; -x_370 = !lean_is_exclusive(x_150); +x_370 = !lean_is_exclusive(x_157); if (x_370 == 0) { -return x_150; +return x_157; } else { lean_object* x_371; lean_object* x_372; lean_object* x_373; -x_371 = lean_ctor_get(x_150, 0); -x_372 = lean_ctor_get(x_150, 1); +x_371 = lean_ctor_get(x_157, 0); +x_372 = lean_ctor_get(x_157, 1); lean_inc(x_372); lean_inc(x_371); -lean_dec(x_150); +lean_dec(x_157); x_373 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_373, 0, x_371); lean_ctor_set(x_373, 1, x_372); @@ -3890,19 +3958,19 @@ return x_373; else { uint8_t x_374; -x_374 = !lean_is_exclusive(x_145); +x_374 = !lean_is_exclusive(x_154); if (x_374 == 0) { -return x_145; +return x_154; } else { lean_object* x_375; lean_object* x_376; lean_object* x_377; -x_375 = lean_ctor_get(x_145, 0); -x_376 = lean_ctor_get(x_145, 1); +x_375 = lean_ctor_get(x_154, 0); +x_376 = lean_ctor_get(x_154, 1); lean_inc(x_376); lean_inc(x_375); -lean_dec(x_145); +lean_dec(x_154); x_377 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_377, 0, x_375); lean_ctor_set(x_377, 1, x_376); @@ -3913,19 +3981,19 @@ return x_377; else { uint8_t x_378; -x_378 = !lean_is_exclusive(x_142); +x_378 = !lean_is_exclusive(x_149); if (x_378 == 0) { -return x_142; +return x_149; } else { lean_object* x_379; lean_object* x_380; lean_object* x_381; -x_379 = lean_ctor_get(x_142, 0); -x_380 = lean_ctor_get(x_142, 1); +x_379 = lean_ctor_get(x_149, 0); +x_380 = lean_ctor_get(x_149, 1); lean_inc(x_380); lean_inc(x_379); -lean_dec(x_142); +lean_dec(x_149); x_381 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_381, 0, x_379); lean_ctor_set(x_381, 1, x_380); @@ -3936,19 +4004,19 @@ return x_381; else { uint8_t x_382; -x_382 = !lean_is_exclusive(x_139); +x_382 = !lean_is_exclusive(x_146); if (x_382 == 0) { -return x_139; +return x_146; } else { lean_object* x_383; lean_object* x_384; lean_object* x_385; -x_383 = lean_ctor_get(x_139, 0); -x_384 = lean_ctor_get(x_139, 1); +x_383 = lean_ctor_get(x_146, 0); +x_384 = lean_ctor_get(x_146, 1); lean_inc(x_384); lean_inc(x_383); -lean_dec(x_139); +lean_dec(x_146); x_385 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_385, 0, x_383); lean_ctor_set(x_385, 1, x_384); @@ -3959,19 +4027,19 @@ return x_385; else { uint8_t x_386; -x_386 = !lean_is_exclusive(x_134); +x_386 = !lean_is_exclusive(x_143); if (x_386 == 0) { -return x_134; +return x_143; } else { lean_object* x_387; lean_object* x_388; lean_object* x_389; -x_387 = lean_ctor_get(x_134, 0); -x_388 = lean_ctor_get(x_134, 1); +x_387 = lean_ctor_get(x_143, 0); +x_388 = lean_ctor_get(x_143, 1); lean_inc(x_388); lean_inc(x_387); -lean_dec(x_134); +lean_dec(x_143); x_389 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_389, 0, x_387); lean_ctor_set(x_389, 1, x_388); @@ -3982,19 +4050,19 @@ return x_389; else { uint8_t x_390; -x_390 = !lean_is_exclusive(x_131); +x_390 = !lean_is_exclusive(x_138); if (x_390 == 0) { -return x_131; +return x_138; } else { lean_object* x_391; lean_object* x_392; lean_object* x_393; -x_391 = lean_ctor_get(x_131, 0); -x_392 = lean_ctor_get(x_131, 1); +x_391 = lean_ctor_get(x_138, 0); +x_392 = lean_ctor_get(x_138, 1); lean_inc(x_392); lean_inc(x_391); -lean_dec(x_131); +lean_dec(x_138); x_393 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_393, 0, x_391); lean_ctor_set(x_393, 1, x_392); @@ -4005,19 +4073,19 @@ return x_393; else { uint8_t x_394; -x_394 = !lean_is_exclusive(x_128); +x_394 = !lean_is_exclusive(x_135); if (x_394 == 0) { -return x_128; +return x_135; } else { lean_object* x_395; lean_object* x_396; lean_object* x_397; -x_395 = lean_ctor_get(x_128, 0); -x_396 = lean_ctor_get(x_128, 1); +x_395 = lean_ctor_get(x_135, 0); +x_396 = lean_ctor_get(x_135, 1); lean_inc(x_396); lean_inc(x_395); -lean_dec(x_128); +lean_dec(x_135); x_397 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_397, 0, x_395); lean_ctor_set(x_397, 1, x_396); @@ -4028,19 +4096,19 @@ return x_397; else { uint8_t x_398; -x_398 = !lean_is_exclusive(x_123); +x_398 = !lean_is_exclusive(x_132); if (x_398 == 0) { -return x_123; +return x_132; } else { lean_object* x_399; lean_object* x_400; lean_object* x_401; -x_399 = lean_ctor_get(x_123, 0); -x_400 = lean_ctor_get(x_123, 1); +x_399 = lean_ctor_get(x_132, 0); +x_400 = lean_ctor_get(x_132, 1); lean_inc(x_400); lean_inc(x_399); -lean_dec(x_123); +lean_dec(x_132); x_401 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_401, 0, x_399); lean_ctor_set(x_401, 1, x_400); @@ -4051,19 +4119,19 @@ return x_401; else { uint8_t x_402; -x_402 = !lean_is_exclusive(x_120); +x_402 = !lean_is_exclusive(x_126); if (x_402 == 0) { -return x_120; +return x_126; } else { lean_object* x_403; lean_object* x_404; lean_object* x_405; -x_403 = lean_ctor_get(x_120, 0); -x_404 = lean_ctor_get(x_120, 1); +x_403 = lean_ctor_get(x_126, 0); +x_404 = lean_ctor_get(x_126, 1); lean_inc(x_404); lean_inc(x_403); -lean_dec(x_120); +lean_dec(x_126); x_405 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_405, 0, x_403); lean_ctor_set(x_405, 1, x_404); @@ -4074,19 +4142,19 @@ return x_405; else { uint8_t x_406; -x_406 = !lean_is_exclusive(x_117); +x_406 = !lean_is_exclusive(x_123); if (x_406 == 0) { -return x_117; +return x_123; } else { lean_object* x_407; lean_object* x_408; lean_object* x_409; -x_407 = lean_ctor_get(x_117, 0); -x_408 = lean_ctor_get(x_117, 1); +x_407 = lean_ctor_get(x_123, 0); +x_408 = lean_ctor_get(x_123, 1); lean_inc(x_408); lean_inc(x_407); -lean_dec(x_117); +lean_dec(x_123); x_409 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_409, 0, x_407); lean_ctor_set(x_409, 1, x_408); @@ -4097,19 +4165,19 @@ return x_409; else { uint8_t x_410; -x_410 = !lean_is_exclusive(x_112); +x_410 = !lean_is_exclusive(x_120); if (x_410 == 0) { -return x_112; +return x_120; } else { lean_object* x_411; lean_object* x_412; lean_object* x_413; -x_411 = lean_ctor_get(x_112, 0); -x_412 = lean_ctor_get(x_112, 1); +x_411 = lean_ctor_get(x_120, 0); +x_412 = lean_ctor_get(x_120, 1); lean_inc(x_412); lean_inc(x_411); -lean_dec(x_112); +lean_dec(x_120); x_413 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_413, 0, x_411); lean_ctor_set(x_413, 1, x_412); @@ -4120,19 +4188,19 @@ return x_413; else { uint8_t x_414; -x_414 = !lean_is_exclusive(x_109); +x_414 = !lean_is_exclusive(x_115); if (x_414 == 0) { -return x_109; +return x_115; } else { lean_object* x_415; lean_object* x_416; lean_object* x_417; -x_415 = lean_ctor_get(x_109, 0); -x_416 = lean_ctor_get(x_109, 1); +x_415 = lean_ctor_get(x_115, 0); +x_416 = lean_ctor_get(x_115, 1); lean_inc(x_416); lean_inc(x_415); -lean_dec(x_109); +lean_dec(x_115); x_417 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_417, 0, x_415); lean_ctor_set(x_417, 1, x_416); @@ -4143,19 +4211,19 @@ return x_417; else { uint8_t x_418; -x_418 = !lean_is_exclusive(x_106); +x_418 = !lean_is_exclusive(x_112); if (x_418 == 0) { -return x_106; +return x_112; } else { lean_object* x_419; lean_object* x_420; lean_object* x_421; -x_419 = lean_ctor_get(x_106, 0); -x_420 = lean_ctor_get(x_106, 1); +x_419 = lean_ctor_get(x_112, 0); +x_420 = lean_ctor_get(x_112, 1); lean_inc(x_420); lean_inc(x_419); -lean_dec(x_106); +lean_dec(x_112); x_421 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_421, 0, x_419); lean_ctor_set(x_421, 1, x_420); @@ -4166,19 +4234,19 @@ return x_421; else { uint8_t x_422; -x_422 = !lean_is_exclusive(x_101); +x_422 = !lean_is_exclusive(x_109); if (x_422 == 0) { -return x_101; +return x_109; } else { lean_object* x_423; lean_object* x_424; lean_object* x_425; -x_423 = lean_ctor_get(x_101, 0); -x_424 = lean_ctor_get(x_101, 1); +x_423 = lean_ctor_get(x_109, 0); +x_424 = lean_ctor_get(x_109, 1); lean_inc(x_424); lean_inc(x_423); -lean_dec(x_101); +lean_dec(x_109); x_425 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_425, 0, x_423); lean_ctor_set(x_425, 1, x_424); @@ -4189,19 +4257,19 @@ return x_425; else { uint8_t x_426; -x_426 = !lean_is_exclusive(x_98); +x_426 = !lean_is_exclusive(x_104); if (x_426 == 0) { -return x_98; +return x_104; } else { lean_object* x_427; lean_object* x_428; lean_object* x_429; -x_427 = lean_ctor_get(x_98, 0); -x_428 = lean_ctor_get(x_98, 1); +x_427 = lean_ctor_get(x_104, 0); +x_428 = lean_ctor_get(x_104, 1); lean_inc(x_428); lean_inc(x_427); -lean_dec(x_98); +lean_dec(x_104); x_429 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_429, 0, x_427); lean_ctor_set(x_429, 1, x_428); @@ -4212,19 +4280,19 @@ return x_429; else { uint8_t x_430; -x_430 = !lean_is_exclusive(x_95); +x_430 = !lean_is_exclusive(x_101); if (x_430 == 0) { -return x_95; +return x_101; } else { lean_object* x_431; lean_object* x_432; lean_object* x_433; -x_431 = lean_ctor_get(x_95, 0); -x_432 = lean_ctor_get(x_95, 1); +x_431 = lean_ctor_get(x_101, 0); +x_432 = lean_ctor_get(x_101, 1); lean_inc(x_432); lean_inc(x_431); -lean_dec(x_95); +lean_dec(x_101); x_433 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_433, 0, x_431); lean_ctor_set(x_433, 1, x_432); @@ -4235,19 +4303,19 @@ return x_433; else { uint8_t x_434; -x_434 = !lean_is_exclusive(x_90); +x_434 = !lean_is_exclusive(x_98); if (x_434 == 0) { -return x_90; +return x_98; } else { lean_object* x_435; lean_object* x_436; lean_object* x_437; -x_435 = lean_ctor_get(x_90, 0); -x_436 = lean_ctor_get(x_90, 1); +x_435 = lean_ctor_get(x_98, 0); +x_436 = lean_ctor_get(x_98, 1); lean_inc(x_436); lean_inc(x_435); -lean_dec(x_90); +lean_dec(x_98); x_437 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_437, 0, x_435); lean_ctor_set(x_437, 1, x_436); @@ -4258,19 +4326,19 @@ return x_437; else { uint8_t x_438; -x_438 = !lean_is_exclusive(x_87); +x_438 = !lean_is_exclusive(x_93); if (x_438 == 0) { -return x_87; +return x_93; } else { lean_object* x_439; lean_object* x_440; lean_object* x_441; -x_439 = lean_ctor_get(x_87, 0); -x_440 = lean_ctor_get(x_87, 1); +x_439 = lean_ctor_get(x_93, 0); +x_440 = lean_ctor_get(x_93, 1); lean_inc(x_440); lean_inc(x_439); -lean_dec(x_87); +lean_dec(x_93); x_441 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_441, 0, x_439); lean_ctor_set(x_441, 1, x_440); @@ -4281,19 +4349,19 @@ return x_441; else { uint8_t x_442; -x_442 = !lean_is_exclusive(x_84); +x_442 = !lean_is_exclusive(x_90); if (x_442 == 0) { -return x_84; +return x_90; } else { lean_object* x_443; lean_object* x_444; lean_object* x_445; -x_443 = lean_ctor_get(x_84, 0); -x_444 = lean_ctor_get(x_84, 1); +x_443 = lean_ctor_get(x_90, 0); +x_444 = lean_ctor_get(x_90, 1); lean_inc(x_444); lean_inc(x_443); -lean_dec(x_84); +lean_dec(x_90); x_445 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_445, 0, x_443); lean_ctor_set(x_445, 1, x_444); @@ -4304,19 +4372,19 @@ return x_445; else { uint8_t x_446; -x_446 = !lean_is_exclusive(x_79); +x_446 = !lean_is_exclusive(x_87); if (x_446 == 0) { -return x_79; +return x_87; } else { lean_object* x_447; lean_object* x_448; lean_object* x_449; -x_447 = lean_ctor_get(x_79, 0); -x_448 = lean_ctor_get(x_79, 1); +x_447 = lean_ctor_get(x_87, 0); +x_448 = lean_ctor_get(x_87, 1); lean_inc(x_448); lean_inc(x_447); -lean_dec(x_79); +lean_dec(x_87); x_449 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_449, 0, x_447); lean_ctor_set(x_449, 1, x_448); @@ -4327,19 +4395,19 @@ return x_449; else { uint8_t x_450; -x_450 = !lean_is_exclusive(x_76); +x_450 = !lean_is_exclusive(x_82); if (x_450 == 0) { -return x_76; +return x_82; } else { lean_object* x_451; lean_object* x_452; lean_object* x_453; -x_451 = lean_ctor_get(x_76, 0); -x_452 = lean_ctor_get(x_76, 1); +x_451 = lean_ctor_get(x_82, 0); +x_452 = lean_ctor_get(x_82, 1); lean_inc(x_452); lean_inc(x_451); -lean_dec(x_76); +lean_dec(x_82); x_453 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_453, 0, x_451); lean_ctor_set(x_453, 1, x_452); @@ -4350,19 +4418,19 @@ return x_453; else { uint8_t x_454; -x_454 = !lean_is_exclusive(x_73); +x_454 = !lean_is_exclusive(x_79); if (x_454 == 0) { -return x_73; +return x_79; } else { lean_object* x_455; lean_object* x_456; lean_object* x_457; -x_455 = lean_ctor_get(x_73, 0); -x_456 = lean_ctor_get(x_73, 1); +x_455 = lean_ctor_get(x_79, 0); +x_456 = lean_ctor_get(x_79, 1); lean_inc(x_456); lean_inc(x_455); -lean_dec(x_73); +lean_dec(x_79); x_457 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_457, 0, x_455); lean_ctor_set(x_457, 1, x_456); @@ -4373,19 +4441,19 @@ return x_457; else { uint8_t x_458; -x_458 = !lean_is_exclusive(x_68); +x_458 = !lean_is_exclusive(x_76); if (x_458 == 0) { -return x_68; +return x_76; } else { lean_object* x_459; lean_object* x_460; lean_object* x_461; -x_459 = lean_ctor_get(x_68, 0); -x_460 = lean_ctor_get(x_68, 1); +x_459 = lean_ctor_get(x_76, 0); +x_460 = lean_ctor_get(x_76, 1); lean_inc(x_460); lean_inc(x_459); -lean_dec(x_68); +lean_dec(x_76); x_461 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_461, 0, x_459); lean_ctor_set(x_461, 1, x_460); @@ -4396,19 +4464,19 @@ return x_461; else { uint8_t x_462; -x_462 = !lean_is_exclusive(x_65); +x_462 = !lean_is_exclusive(x_71); if (x_462 == 0) { -return x_65; +return x_71; } else { lean_object* x_463; lean_object* x_464; lean_object* x_465; -x_463 = lean_ctor_get(x_65, 0); -x_464 = lean_ctor_get(x_65, 1); +x_463 = lean_ctor_get(x_71, 0); +x_464 = lean_ctor_get(x_71, 1); lean_inc(x_464); lean_inc(x_463); -lean_dec(x_65); +lean_dec(x_71); x_465 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_465, 0, x_463); lean_ctor_set(x_465, 1, x_464); @@ -4419,19 +4487,19 @@ return x_465; else { uint8_t x_466; -x_466 = !lean_is_exclusive(x_62); +x_466 = !lean_is_exclusive(x_68); if (x_466 == 0) { -return x_62; +return x_68; } else { lean_object* x_467; lean_object* x_468; lean_object* x_469; -x_467 = lean_ctor_get(x_62, 0); -x_468 = lean_ctor_get(x_62, 1); +x_467 = lean_ctor_get(x_68, 0); +x_468 = lean_ctor_get(x_68, 1); lean_inc(x_468); lean_inc(x_467); -lean_dec(x_62); +lean_dec(x_68); x_469 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_469, 0, x_467); lean_ctor_set(x_469, 1, x_468); @@ -4442,19 +4510,19 @@ return x_469; else { uint8_t x_470; -x_470 = !lean_is_exclusive(x_57); +x_470 = !lean_is_exclusive(x_65); if (x_470 == 0) { -return x_57; +return x_65; } else { lean_object* x_471; lean_object* x_472; lean_object* x_473; -x_471 = lean_ctor_get(x_57, 0); -x_472 = lean_ctor_get(x_57, 1); +x_471 = lean_ctor_get(x_65, 0); +x_472 = lean_ctor_get(x_65, 1); lean_inc(x_472); lean_inc(x_471); -lean_dec(x_57); +lean_dec(x_65); x_473 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_473, 0, x_471); lean_ctor_set(x_473, 1, x_472); @@ -4465,19 +4533,19 @@ return x_473; else { uint8_t x_474; -x_474 = !lean_is_exclusive(x_54); +x_474 = !lean_is_exclusive(x_60); if (x_474 == 0) { -return x_54; +return x_60; } else { lean_object* x_475; lean_object* x_476; lean_object* x_477; -x_475 = lean_ctor_get(x_54, 0); -x_476 = lean_ctor_get(x_54, 1); +x_475 = lean_ctor_get(x_60, 0); +x_476 = lean_ctor_get(x_60, 1); lean_inc(x_476); lean_inc(x_475); -lean_dec(x_54); +lean_dec(x_60); x_477 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_477, 0, x_475); lean_ctor_set(x_477, 1, x_476); @@ -4488,19 +4556,19 @@ return x_477; else { uint8_t x_478; -x_478 = !lean_is_exclusive(x_51); +x_478 = !lean_is_exclusive(x_57); if (x_478 == 0) { -return x_51; +return x_57; } else { lean_object* x_479; lean_object* x_480; lean_object* x_481; -x_479 = lean_ctor_get(x_51, 0); -x_480 = lean_ctor_get(x_51, 1); +x_479 = lean_ctor_get(x_57, 0); +x_480 = lean_ctor_get(x_57, 1); lean_inc(x_480); lean_inc(x_479); -lean_dec(x_51); +lean_dec(x_57); x_481 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_481, 0, x_479); lean_ctor_set(x_481, 1, x_480); @@ -4511,19 +4579,19 @@ return x_481; else { uint8_t x_482; -x_482 = !lean_is_exclusive(x_46); +x_482 = !lean_is_exclusive(x_54); if (x_482 == 0) { -return x_46; +return x_54; } else { lean_object* x_483; lean_object* x_484; lean_object* x_485; -x_483 = lean_ctor_get(x_46, 0); -x_484 = lean_ctor_get(x_46, 1); +x_483 = lean_ctor_get(x_54, 0); +x_484 = lean_ctor_get(x_54, 1); lean_inc(x_484); lean_inc(x_483); -lean_dec(x_46); +lean_dec(x_54); x_485 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_485, 0, x_483); lean_ctor_set(x_485, 1, x_484); @@ -4534,19 +4602,19 @@ return x_485; else { uint8_t x_486; -x_486 = !lean_is_exclusive(x_43); +x_486 = !lean_is_exclusive(x_49); if (x_486 == 0) { -return x_43; +return x_49; } else { lean_object* x_487; lean_object* x_488; lean_object* x_489; -x_487 = lean_ctor_get(x_43, 0); -x_488 = lean_ctor_get(x_43, 1); +x_487 = lean_ctor_get(x_49, 0); +x_488 = lean_ctor_get(x_49, 1); lean_inc(x_488); lean_inc(x_487); -lean_dec(x_43); +lean_dec(x_49); x_489 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_489, 0, x_487); lean_ctor_set(x_489, 1, x_488); @@ -4557,19 +4625,19 @@ return x_489; else { uint8_t x_490; -x_490 = !lean_is_exclusive(x_40); +x_490 = !lean_is_exclusive(x_46); if (x_490 == 0) { -return x_40; +return x_46; } else { lean_object* x_491; lean_object* x_492; lean_object* x_493; -x_491 = lean_ctor_get(x_40, 0); -x_492 = lean_ctor_get(x_40, 1); +x_491 = lean_ctor_get(x_46, 0); +x_492 = lean_ctor_get(x_46, 1); lean_inc(x_492); lean_inc(x_491); -lean_dec(x_40); +lean_dec(x_46); x_493 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_493, 0, x_491); lean_ctor_set(x_493, 1, x_492); @@ -4580,19 +4648,19 @@ return x_493; else { uint8_t x_494; -x_494 = !lean_is_exclusive(x_35); +x_494 = !lean_is_exclusive(x_43); if (x_494 == 0) { -return x_35; +return x_43; } else { lean_object* x_495; lean_object* x_496; lean_object* x_497; -x_495 = lean_ctor_get(x_35, 0); -x_496 = lean_ctor_get(x_35, 1); +x_495 = lean_ctor_get(x_43, 0); +x_496 = lean_ctor_get(x_43, 1); lean_inc(x_496); lean_inc(x_495); -lean_dec(x_35); +lean_dec(x_43); x_497 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_497, 0, x_495); lean_ctor_set(x_497, 1, x_496); @@ -4603,19 +4671,19 @@ return x_497; else { uint8_t x_498; -x_498 = !lean_is_exclusive(x_32); +x_498 = !lean_is_exclusive(x_37); if (x_498 == 0) { -return x_32; +return x_37; } else { lean_object* x_499; lean_object* x_500; lean_object* x_501; -x_499 = lean_ctor_get(x_32, 0); -x_500 = lean_ctor_get(x_32, 1); +x_499 = lean_ctor_get(x_37, 0); +x_500 = lean_ctor_get(x_37, 1); lean_inc(x_500); lean_inc(x_499); -lean_dec(x_32); +lean_dec(x_37); x_501 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_501, 0, x_499); lean_ctor_set(x_501, 1, x_500); @@ -4626,19 +4694,19 @@ return x_501; else { uint8_t x_502; -x_502 = !lean_is_exclusive(x_29); +x_502 = !lean_is_exclusive(x_34); if (x_502 == 0) { -return x_29; +return x_34; } else { lean_object* x_503; lean_object* x_504; lean_object* x_505; -x_503 = lean_ctor_get(x_29, 0); -x_504 = lean_ctor_get(x_29, 1); +x_503 = lean_ctor_get(x_34, 0); +x_504 = lean_ctor_get(x_34, 1); lean_inc(x_504); lean_inc(x_503); -lean_dec(x_29); +lean_dec(x_34); x_505 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_505, 0, x_503); lean_ctor_set(x_505, 1, x_504); @@ -4649,19 +4717,19 @@ return x_505; else { uint8_t x_506; -x_506 = !lean_is_exclusive(x_24); +x_506 = !lean_is_exclusive(x_31); if (x_506 == 0) { -return x_24; +return x_31; } else { lean_object* x_507; lean_object* x_508; lean_object* x_509; -x_507 = lean_ctor_get(x_24, 0); -x_508 = lean_ctor_get(x_24, 1); +x_507 = lean_ctor_get(x_31, 0); +x_508 = lean_ctor_get(x_31, 1); lean_inc(x_508); lean_inc(x_507); -lean_dec(x_24); +lean_dec(x_31); x_509 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_509, 0, x_507); lean_ctor_set(x_509, 1, x_508); @@ -4672,19 +4740,19 @@ return x_509; else { uint8_t x_510; -x_510 = !lean_is_exclusive(x_21); +x_510 = !lean_is_exclusive(x_26); if (x_510 == 0) { -return x_21; +return x_26; } else { lean_object* x_511; lean_object* x_512; lean_object* x_513; -x_511 = lean_ctor_get(x_21, 0); -x_512 = lean_ctor_get(x_21, 1); +x_511 = lean_ctor_get(x_26, 0); +x_512 = lean_ctor_get(x_26, 1); lean_inc(x_512); lean_inc(x_511); -lean_dec(x_21); +lean_dec(x_26); x_513 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_513, 0, x_511); lean_ctor_set(x_513, 1, x_512); @@ -4695,19 +4763,19 @@ return x_513; else { uint8_t x_514; -x_514 = !lean_is_exclusive(x_18); +x_514 = !lean_is_exclusive(x_23); if (x_514 == 0) { -return x_18; +return x_23; } else { lean_object* x_515; lean_object* x_516; lean_object* x_517; -x_515 = lean_ctor_get(x_18, 0); -x_516 = lean_ctor_get(x_18, 1); +x_515 = lean_ctor_get(x_23, 0); +x_516 = lean_ctor_get(x_23, 1); lean_inc(x_516); lean_inc(x_515); -lean_dec(x_18); +lean_dec(x_23); x_517 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_517, 0, x_515); lean_ctor_set(x_517, 1, x_516); @@ -4718,19 +4786,19 @@ return x_517; else { uint8_t x_518; -x_518 = !lean_is_exclusive(x_13); +x_518 = !lean_is_exclusive(x_20); if (x_518 == 0) { -return x_13; +return x_20; } else { lean_object* x_519; lean_object* x_520; lean_object* x_521; -x_519 = lean_ctor_get(x_13, 0); -x_520 = lean_ctor_get(x_13, 1); +x_519 = lean_ctor_get(x_20, 0); +x_520 = lean_ctor_get(x_20, 1); lean_inc(x_520); lean_inc(x_519); -lean_dec(x_13); +lean_dec(x_20); x_521 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_521, 0, x_519); lean_ctor_set(x_521, 1, x_520); @@ -4741,19 +4809,19 @@ return x_521; else { uint8_t x_522; -x_522 = !lean_is_exclusive(x_9); +x_522 = !lean_is_exclusive(x_15); if (x_522 == 0) { -return x_9; +return x_15; } else { lean_object* x_523; lean_object* x_524; lean_object* x_525; -x_523 = lean_ctor_get(x_9, 0); -x_524 = lean_ctor_get(x_9, 1); +x_523 = lean_ctor_get(x_15, 0); +x_524 = lean_ctor_get(x_15, 1); lean_inc(x_524); lean_inc(x_523); -lean_dec(x_9); +lean_dec(x_15); x_525 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_525, 0, x_523); lean_ctor_set(x_525, 1, x_524); @@ -4764,19 +4832,19 @@ return x_525; else { uint8_t x_526; -x_526 = !lean_is_exclusive(x_5); +x_526 = !lean_is_exclusive(x_11); if (x_526 == 0) { -return x_5; +return x_11; } else { lean_object* x_527; lean_object* x_528; lean_object* x_529; -x_527 = lean_ctor_get(x_5, 0); -x_528 = lean_ctor_get(x_5, 1); +x_527 = lean_ctor_get(x_11, 0); +x_528 = lean_ctor_get(x_11, 1); lean_inc(x_528); lean_inc(x_527); -lean_dec(x_5); +lean_dec(x_11); x_529 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_529, 0, x_527); lean_ctor_set(x_529, 1, x_528); @@ -4784,6 +4852,29 @@ return x_529; } } } +else +{ +uint8_t x_530; +x_530 = !lean_is_exclusive(x_7); +if (x_530 == 0) +{ +return x_7; +} +else +{ +lean_object* x_531; lean_object* x_532; lean_object* x_533; +x_531 = lean_ctor_get(x_7, 0); +x_532 = lean_ctor_get(x_7, 1); +lean_inc(x_532); +lean_inc(x_531); +lean_dec(x_7); +x_533 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_533, 0, x_531); +lean_ctor_set(x_533, 1, x_532); +return x_533; +} +} +} } LEAN_EXPORT lean_object* lean_mk_antiquot_parenthesizer(lean_object* x_1, lean_object* x_2, uint8_t x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: @@ -4855,7 +4946,7 @@ static lean_object* _init_l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_ident_ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_ident_parenthesizer___closed__4; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__83; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__86; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } @@ -4899,7 +4990,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_ident_p { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_2 = l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_ident_parenthesizer___closed__8; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__84; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__87; x_4 = l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_ident_parenthesizer___closed__7; x_5 = l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_ident_parenthesizer___closed__9; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); @@ -4955,7 +5046,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_numLit_ { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_2 = l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_ident_parenthesizer___closed__8; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__44; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__45; x_4 = l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_numLit_parenthesizer___closed__3; x_5 = l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_numLit_parenthesizer___closed__4; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); @@ -5011,7 +5102,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_scienti { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_2 = l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_ident_parenthesizer___closed__8; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__76; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__79; x_4 = l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_scientificLit_parenthesizer___closed__3; x_5 = l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_scientificLit_parenthesizer___closed__4; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); @@ -5067,7 +5158,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_charLit { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_2 = l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_ident_parenthesizer___closed__8; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__60; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__63; x_4 = l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_charLit_parenthesizer___closed__3; x_5 = l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_charLit_parenthesizer___closed__4; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); @@ -5123,7 +5214,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_strLit_ { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_2 = l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_ident_parenthesizer___closed__8; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__52; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__55; x_4 = l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_strLit_parenthesizer___closed__3; x_5 = l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_strLit_parenthesizer___closed__4; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); @@ -5236,7 +5327,7 @@ lean_dec(x_6); x_8 = lean_ctor_get(x_2, 5); lean_inc(x_8); lean_dec(x_2); -x_9 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__18; +x_9 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__19; x_10 = l_Lean_Parser_getConstAlias___rarg(x_9, x_5, x_7); if (lean_obj_tag(x_10) == 0) { @@ -5317,7 +5408,7 @@ lean_inc(x_31); lean_dec(x_30); x_32 = lean_ctor_get(x_2, 5); lean_inc(x_32); -x_33 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__18; +x_33 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__19; x_34 = l_Lean_Parser_getUnaryAlias___rarg(x_33, x_28, x_31); if (lean_obj_tag(x_34) == 0) { @@ -5445,7 +5536,7 @@ lean_inc(x_66); lean_dec(x_65); x_67 = lean_ctor_get(x_2, 5); lean_inc(x_67); -x_68 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__18; +x_68 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__19; x_69 = l_Lean_Parser_getBinaryAlias___rarg(x_68, x_62, x_66); if (lean_obj_tag(x_69) == 0) { @@ -6167,7 +6258,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_ident_formatter(lean_obj { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_PrettyPrinter_Formatter_ident_formatter___closed__1; -x_7 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__87; +x_7 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__90; x_8 = l_Lean_PrettyPrinter_Formatter_andthen_formatter(x_6, x_7, x_1, x_2, x_3, x_4, x_5); return x_8; } @@ -6195,7 +6286,7 @@ static lean_object* _init_l___regBuiltin_Lean_PrettyPrinter_Formatter_ident_form { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_PrettyPrinter_Formatter_ident_formatter___closed__2; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__83; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__86; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } @@ -6239,7 +6330,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Formatter_ident_forma { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_2 = l___regBuiltin_Lean_PrettyPrinter_Formatter_ident_formatter___closed__6; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__84; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__87; x_4 = l___regBuiltin_Lean_PrettyPrinter_Formatter_ident_formatter___closed__5; x_5 = l___regBuiltin_Lean_PrettyPrinter_Formatter_ident_formatter___closed__7; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); @@ -6251,7 +6342,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_numLit_formatter(lean_ob { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_PrettyPrinter_Formatter_ident_formatter___closed__1; -x_7 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__47; +x_7 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__50; x_8 = l_Lean_PrettyPrinter_Formatter_andthen_formatter(x_6, x_7, x_1, x_2, x_3, x_4, x_5); return x_8; } @@ -6289,7 +6380,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Formatter_numLit_form { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_2 = l___regBuiltin_Lean_PrettyPrinter_Formatter_ident_formatter___closed__6; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__44; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__45; x_4 = l___regBuiltin_Lean_PrettyPrinter_Formatter_numLit_formatter___closed__2; x_5 = l___regBuiltin_Lean_PrettyPrinter_Formatter_numLit_formatter___closed__3; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); @@ -6301,7 +6392,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_scientificLit_formatter( { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_PrettyPrinter_Formatter_ident_formatter___closed__1; -x_7 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__79; +x_7 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__82; x_8 = l_Lean_PrettyPrinter_Formatter_andthen_formatter(x_6, x_7, x_1, x_2, x_3, x_4, x_5); return x_8; } @@ -6339,7 +6430,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Formatter_scientificL { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_2 = l___regBuiltin_Lean_PrettyPrinter_Formatter_ident_formatter___closed__6; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__76; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__79; x_4 = l___regBuiltin_Lean_PrettyPrinter_Formatter_scientificLit_formatter___closed__2; x_5 = l___regBuiltin_Lean_PrettyPrinter_Formatter_scientificLit_formatter___closed__3; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); @@ -6351,7 +6442,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_charLit_formatter(lean_o { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_PrettyPrinter_Formatter_ident_formatter___closed__1; -x_7 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__63; +x_7 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__66; x_8 = l_Lean_PrettyPrinter_Formatter_andthen_formatter(x_6, x_7, x_1, x_2, x_3, x_4, x_5); return x_8; } @@ -6389,7 +6480,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Formatter_charLit_for { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_2 = l___regBuiltin_Lean_PrettyPrinter_Formatter_ident_formatter___closed__6; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__60; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__63; x_4 = l___regBuiltin_Lean_PrettyPrinter_Formatter_charLit_formatter___closed__2; x_5 = l___regBuiltin_Lean_PrettyPrinter_Formatter_charLit_formatter___closed__3; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); @@ -6401,7 +6492,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_strLit_formatter(lean_ob { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_PrettyPrinter_Formatter_ident_formatter___closed__1; -x_7 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__55; +x_7 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__58; x_8 = l_Lean_PrettyPrinter_Formatter_andthen_formatter(x_6, x_7, x_1, x_2, x_3, x_4, x_5); return x_8; } @@ -6439,7 +6530,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Formatter_strLit_form { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_2 = l___regBuiltin_Lean_PrettyPrinter_Formatter_ident_formatter___closed__6; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__52; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__55; x_4 = l___regBuiltin_Lean_PrettyPrinter_Formatter_strLit_formatter___closed__2; x_5 = l___regBuiltin_Lean_PrettyPrinter_Formatter_strLit_formatter___closed__3; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); @@ -6553,7 +6644,7 @@ lean_dec(x_6); x_8 = lean_ctor_get(x_2, 5); lean_inc(x_8); lean_dec(x_2); -x_9 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__15; +x_9 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__16; x_10 = l_Lean_Parser_getConstAlias___rarg(x_9, x_5, x_7); if (lean_obj_tag(x_10) == 0) { @@ -6634,7 +6725,7 @@ lean_inc(x_31); lean_dec(x_30); x_32 = lean_ctor_get(x_2, 5); lean_inc(x_32); -x_33 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__15; +x_33 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__16; x_34 = l_Lean_Parser_getUnaryAlias___rarg(x_33, x_28, x_31); if (lean_obj_tag(x_34) == 0) { @@ -6762,7 +6853,7 @@ lean_inc(x_66); lean_dec(x_65); x_67 = lean_ctor_get(x_2, 5); lean_inc(x_67); -x_68 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__15; +x_68 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__16; x_69 = l_Lean_Parser_getBinaryAlias___rarg(x_68, x_62, x_66); if (lean_obj_tag(x_69) == 0) { @@ -7935,6 +8026,22 @@ l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__230 = _init_l_Lean_Pa lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__230); l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__231 = _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__231(); lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__231); +l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__232 = _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__232(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__232); +l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__233 = _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__233(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__233); +l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__234 = _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__234(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__234); +l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__235 = _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__235(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__235); +l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__236 = _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__236(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__236); +l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__237 = _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__237(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__237); +l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__238 = _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__238(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__238); +l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__239 = _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__239(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__239); res = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); diff --git a/stage0/stdlib/Lean/Parser/Attr.c b/stage0/stdlib/Lean/Parser/Attr.c index 880e94b27f5..85abe830d92 100644 --- a/stage0/stdlib/Lean/Parser/Attr.c +++ b/stage0/stdlib/Lean/Parser/Attr.c @@ -182,6 +182,7 @@ static lean_object* l_Lean_Parser_Attr_defaultInstance___elambda__1___closed__2; static lean_object* l_Lean_Parser_Attr_simple___elambda__1___closed__17; static lean_object* l_Lean_Parser_Attr_class___elambda__1___closed__10; static lean_object* l_Lean_Parser_Attr_externEntry_formatter___closed__1; +uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Attr_macro_formatter(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Attr_extern_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Attr_externEntry___closed__4; @@ -538,7 +539,6 @@ static lean_object* l_Lean_Parser_Attr_simple___elambda__1___closed__11; static lean_object* l___regBuiltin_Lean_Parser_Attr_instance_formatter___closed__1; static lean_object* l_Lean_Parser_Attr_export___elambda__1___closed__4; static lean_object* l_Lean_Parser_Attr_export___closed__4; -uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Attr_export___elambda__1___closed__8; lean_object* l_Lean_Parser_ident_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Attr_instance_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -868,7 +868,7 @@ x_4 = l_Lean_Parser_checkPrecFn(x_3, x_1, x_2); x_5 = lean_ctor_get(x_4, 4); lean_inc(x_5); x_6 = lean_box(0); -x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); +x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); lean_dec(x_5); if (x_7 == 0) { @@ -1390,7 +1390,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -1409,7 +1409,7 @@ lean_inc(x_1); x_20 = l_Lean_Parser_ident___elambda__1(x_1, x_14); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); -x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_16); +x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_16); lean_dec(x_21); if (x_22 == 0) { @@ -1419,7 +1419,7 @@ x_23 = l_Lean_Parser_Attr_simple___elambda__1___closed__4; x_24 = l_Lean_Parser_ParserState_mkNode(x_20, x_23, x_19); x_25 = lean_ctor_get(x_24, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_16); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_16); lean_dec(x_25); if (x_26 == 0) { @@ -1443,7 +1443,7 @@ x_29 = l_Lean_Parser_Attr_simple___elambda__1___closed__4; x_30 = l_Lean_Parser_ParserState_mkNode(x_28, x_29, x_19); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); lean_dec(x_31); if (x_32 == 0) { @@ -2104,7 +2104,7 @@ x_22 = l_Lean_Parser_checkPrecFn(x_21, x_1, x_2); x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); x_24 = lean_box(0); -x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_24); +x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_24); lean_dec(x_23); if (x_25 == 0) { @@ -2134,7 +2134,7 @@ if (x_33 == 0) lean_object* x_34; uint8_t x_35; x_34 = lean_ctor_get(x_29, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_24); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_24); lean_dec(x_34); if (x_35 == 0) { @@ -2142,7 +2142,7 @@ lean_object* x_36; lean_object* x_37; uint8_t x_38; x_36 = l_Lean_Parser_ParserState_mkNode(x_29, x_7, x_27); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_24); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_24); lean_dec(x_37); if (x_38 == 0) { @@ -2165,7 +2165,7 @@ x_40 = l_Lean_Parser_ident___elambda__1(x_1, x_29); x_41 = l_Lean_Parser_ParserState_mkNode(x_40, x_7, x_27); x_42 = lean_ctor_get(x_41, 4); lean_inc(x_42); -x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_24); +x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_24); lean_dec(x_42); if (x_43 == 0) { @@ -2188,7 +2188,7 @@ lean_inc(x_1); x_45 = l_Lean_Parser_tokenAntiquotFn(x_1, x_29); x_46 = lean_ctor_get(x_45, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_24); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_24); lean_dec(x_46); if (x_47 == 0) { @@ -2196,7 +2196,7 @@ lean_object* x_48; lean_object* x_49; uint8_t x_50; x_48 = l_Lean_Parser_ParserState_mkNode(x_45, x_7, x_27); x_49 = lean_ctor_get(x_48, 4); lean_inc(x_49); -x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_24); +x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_24); lean_dec(x_49); if (x_50 == 0) { @@ -2219,7 +2219,7 @@ x_52 = l_Lean_Parser_ident___elambda__1(x_1, x_45); x_53 = l_Lean_Parser_ParserState_mkNode(x_52, x_7, x_27); x_54 = lean_ctor_get(x_53, 4); lean_inc(x_54); -x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_24); +x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_24); lean_dec(x_54); if (x_55 == 0) { @@ -2802,7 +2802,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -2833,7 +2833,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_20, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -2842,7 +2842,7 @@ x_27 = l_Lean_Parser_Attr_export___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_20, x_27, x_17); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -2866,7 +2866,7 @@ x_33 = l_Lean_Parser_Attr_export___elambda__1___closed__2; x_34 = l_Lean_Parser_ParserState_mkNode(x_32, x_33, x_17); x_35 = lean_ctor_get(x_34, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_14); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_14); lean_dec(x_35); if (x_36 == 0) { @@ -2889,7 +2889,7 @@ lean_inc(x_1); x_38 = l_Lean_Parser_tokenAntiquotFn(x_1, x_20); x_39 = lean_ctor_get(x_38, 4); lean_inc(x_39); -x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); +x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); lean_dec(x_39); if (x_40 == 0) { @@ -2898,7 +2898,7 @@ x_41 = l_Lean_Parser_Attr_export___elambda__1___closed__2; x_42 = l_Lean_Parser_ParserState_mkNode(x_38, x_41, x_17); x_43 = lean_ctor_get(x_42, 4); lean_inc(x_43); -x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_14); +x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_14); lean_dec(x_43); if (x_44 == 0) { @@ -2922,7 +2922,7 @@ x_47 = l_Lean_Parser_Attr_export___elambda__1___closed__2; x_48 = l_Lean_Parser_ParserState_mkNode(x_46, x_47, x_17); x_49 = lean_ctor_get(x_48, 4); lean_inc(x_49); -x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_14); +x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_14); lean_dec(x_49); if (x_50 == 0) { @@ -3501,7 +3501,7 @@ x_22 = l_Lean_Parser_checkPrecFn(x_21, x_1, x_2); x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); x_24 = lean_box(0); -x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_24); +x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_24); lean_dec(x_23); if (x_25 == 0) { @@ -3531,7 +3531,7 @@ if (x_33 == 0) lean_object* x_34; uint8_t x_35; x_34 = lean_ctor_get(x_29, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_24); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_24); lean_dec(x_34); if (x_35 == 0) { @@ -3539,7 +3539,7 @@ lean_object* x_36; lean_object* x_37; uint8_t x_38; x_36 = l_Lean_Parser_ParserState_mkNode(x_29, x_7, x_27); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_24); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_24); lean_dec(x_37); if (x_38 == 0) { @@ -3562,7 +3562,7 @@ x_40 = l_Lean_Parser_numLit___elambda__1(x_1, x_29); x_41 = l_Lean_Parser_ParserState_mkNode(x_40, x_7, x_27); x_42 = lean_ctor_get(x_41, 4); lean_inc(x_42); -x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_24); +x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_24); lean_dec(x_42); if (x_43 == 0) { @@ -3585,7 +3585,7 @@ lean_inc(x_1); x_45 = l_Lean_Parser_tokenAntiquotFn(x_1, x_29); x_46 = lean_ctor_get(x_45, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_24); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_24); lean_dec(x_46); if (x_47 == 0) { @@ -3593,7 +3593,7 @@ lean_object* x_48; lean_object* x_49; uint8_t x_50; x_48 = l_Lean_Parser_ParserState_mkNode(x_45, x_7, x_27); x_49 = lean_ctor_get(x_48, 4); lean_inc(x_49); -x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_24); +x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_24); lean_dec(x_49); if (x_50 == 0) { @@ -3616,7 +3616,7 @@ x_52 = l_Lean_Parser_numLit___elambda__1(x_1, x_45); x_53 = l_Lean_Parser_ParserState_mkNode(x_52, x_7, x_27); x_54 = lean_ctor_get(x_53, 4); lean_inc(x_54); -x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_24); +x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_24); lean_dec(x_54); if (x_55 == 0) { @@ -4177,7 +4177,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -4210,7 +4210,7 @@ x_25 = l_Lean_Parser_Attr_class___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_20, x_25, x_17); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); lean_dec(x_27); if (x_28 == 0) { @@ -4234,7 +4234,7 @@ x_31 = l_Lean_Parser_Attr_class___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_17); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -4778,7 +4778,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -4810,7 +4810,7 @@ if (x_26 == 0) lean_object* x_27; uint8_t x_28; x_27 = lean_ctor_get(x_22, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); lean_dec(x_27); if (x_28 == 0) { @@ -4820,7 +4820,7 @@ x_29 = l_Lean_Parser_Attr_instance___elambda__1___closed__2; x_30 = l_Lean_Parser_ParserState_mkNode(x_22, x_29, x_19); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); lean_dec(x_31); if (x_32 == 0) { @@ -4844,7 +4844,7 @@ x_35 = l_Lean_Parser_Attr_instance___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_19); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_16); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_16); lean_dec(x_37); if (x_38 == 0) { @@ -4867,7 +4867,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_22); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); lean_dec(x_41); if (x_42 == 0) { @@ -4877,7 +4877,7 @@ x_43 = l_Lean_Parser_Attr_instance___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_19); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); lean_dec(x_45); if (x_46 == 0) { @@ -4901,7 +4901,7 @@ x_49 = l_Lean_Parser_Attr_instance___elambda__1___closed__2; x_50 = l_Lean_Parser_ParserState_mkNode(x_48, x_49, x_19); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); lean_dec(x_51); if (x_52 == 0) { @@ -5502,7 +5502,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -5534,7 +5534,7 @@ if (x_26 == 0) lean_object* x_27; uint8_t x_28; x_27 = lean_ctor_get(x_22, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); lean_dec(x_27); if (x_28 == 0) { @@ -5544,7 +5544,7 @@ x_29 = l_Lean_Parser_Attr_defaultInstance___elambda__1___closed__2; x_30 = l_Lean_Parser_ParserState_mkNode(x_22, x_29, x_19); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); lean_dec(x_31); if (x_32 == 0) { @@ -5568,7 +5568,7 @@ x_35 = l_Lean_Parser_Attr_defaultInstance___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_19); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_16); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_16); lean_dec(x_37); if (x_38 == 0) { @@ -5591,7 +5591,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_22); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); lean_dec(x_41); if (x_42 == 0) { @@ -5601,7 +5601,7 @@ x_43 = l_Lean_Parser_Attr_defaultInstance___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_19); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); lean_dec(x_45); if (x_46 == 0) { @@ -5625,7 +5625,7 @@ x_49 = l_Lean_Parser_Attr_defaultInstance___elambda__1___closed__2; x_50 = l_Lean_Parser_ParserState_mkNode(x_48, x_49, x_19); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); lean_dec(x_51); if (x_52 == 0) { @@ -6259,7 +6259,7 @@ x_16 = l_Lean_Parser_checkPrecFn(x_15, x_1, x_2); x_17 = lean_ctor_get(x_16, 4); lean_inc(x_17); x_18 = lean_box(0); -x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); +x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); lean_dec(x_17); if (x_19 == 0) { @@ -6279,7 +6279,7 @@ lean_inc(x_1); x_22 = lean_apply_2(x_6, x_1, x_16); x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); -x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_18); +x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_18); lean_dec(x_23); if (x_24 == 0) { @@ -6289,7 +6289,7 @@ x_25 = l_Lean_Parser_Attr_externEntry___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_22, x_25, x_21); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_18); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_18); lean_dec(x_27); if (x_28 == 0) { @@ -6311,7 +6311,7 @@ lean_inc(x_1); x_30 = lean_apply_2(x_4, x_1, x_22); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_18); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_18); lean_dec(x_31); if (x_32 == 0) { @@ -6320,7 +6320,7 @@ x_33 = l_Lean_Parser_Attr_externEntry___elambda__1___closed__2; x_34 = l_Lean_Parser_ParserState_mkNode(x_30, x_33, x_21); x_35 = lean_ctor_get(x_34, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_18); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_18); lean_dec(x_35); if (x_36 == 0) { @@ -6344,7 +6344,7 @@ x_39 = l_Lean_Parser_Attr_externEntry___elambda__1___closed__2; x_40 = l_Lean_Parser_ParserState_mkNode(x_38, x_39, x_21); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_18); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_18); lean_dec(x_41); if (x_42 == 0) { @@ -6662,7 +6662,7 @@ x_16 = l_Lean_Parser_checkPrecFn(x_15, x_1, x_2); x_17 = lean_ctor_get(x_16, 4); lean_inc(x_17); x_18 = lean_box(0); -x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); +x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); lean_dec(x_17); if (x_19 == 0) { @@ -6708,7 +6708,7 @@ goto block_51; lean_object* x_30; uint8_t x_31; x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_18); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_18); lean_dec(x_30); if (x_31 == 0) { @@ -6719,7 +6719,7 @@ x_32 = l_Lean_Parser_Attr_extern___elambda__1___closed__2; x_33 = l_Lean_Parser_ParserState_mkNode(x_29, x_32, x_21); x_34 = lean_ctor_get(x_33, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_18); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_18); lean_dec(x_34); if (x_35 == 0) { @@ -6741,7 +6741,7 @@ lean_inc(x_1); x_37 = lean_apply_2(x_4, x_1, x_29); x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_18); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_18); lean_dec(x_38); if (x_39 == 0) { @@ -6751,7 +6751,7 @@ x_40 = l_Lean_Parser_Attr_extern___elambda__1___closed__2; x_41 = l_Lean_Parser_ParserState_mkNode(x_37, x_40, x_21); x_42 = lean_ctor_get(x_41, 4); lean_inc(x_42); -x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_18); +x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_18); lean_dec(x_42); if (x_43 == 0) { @@ -6775,7 +6775,7 @@ x_46 = l_Lean_Parser_Attr_extern___elambda__1___closed__2; x_47 = l_Lean_Parser_ParserState_mkNode(x_45, x_46, x_21); x_48 = lean_ctor_get(x_47, 4); lean_inc(x_48); -x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_18); +x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_18); lean_dec(x_48); if (x_49 == 0) { diff --git a/stage0/stdlib/Lean/Parser/Basic.c b/stage0/stdlib/Lean/Parser/Basic.c index 93be35ad5ac..ccb27d5757d 100644 --- a/stage0/stdlib/Lean/Parser/Basic.c +++ b/stage0/stdlib/Lean/Parser/Basic.c @@ -33,7 +33,6 @@ LEAN_EXPORT lean_object* l_Std_RBNode_find___at_Lean_Parser_indexed___spec__3___ LEAN_EXPORT lean_object* l_Std_RBNode_find___at_Lean_Parser_indexed___spec__5___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_nameLitAux(lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); -extern lean_object* l_Lean_fieldIdxKind; static lean_object* l_Lean_Parser_categoryParserFn___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_forArgsM___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_checkLinebreakBeforeFn(lean_object*, lean_object*, lean_object*); @@ -71,10 +70,10 @@ LEAN_EXPORT lean_object* l_Lean_Parser_manyNoAntiquot(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_decimalNumberFn___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_RBNode_find___at_Lean_Parser_indexed___spec__8___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Parser_notFollowedByFn___closed__1; -extern lean_object* l_Lean_nullKind; static lean_object* l_Lean_Parser_Parser_info___default___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_leadPrec; LEAN_EXPORT lean_object* l_Lean_Parser_notFollowedByFn___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_decimalNumberFn___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_strLitNoAntiquot; static lean_object* l_Lean_Parser_octalNumberFn___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_quotedCharCoreFn(lean_object*, lean_object*, lean_object*); @@ -109,6 +108,7 @@ LEAN_EXPORT uint8_t l_Lean_Parser_isQuotableCharDefault(uint32_t); LEAN_EXPORT lean_object* l_Lean_Parser_hexNumberFn___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_RBNode_find___at_Lean_Parser_indexed___spec__2(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_forArgsM___spec__1___at_Lean_Syntax_forArgsM___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_mkTokenAndFixPos___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_longestMatchFn(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_fieldIdx___closed__6; static lean_object* l_Lean_Parser_errorAtSavedPos___closed__1; @@ -140,6 +140,7 @@ static lean_object* l_Lean_Parser_dbgTraceStateFn___closed__1; static lean_object* l_Lean_Parser_checkNoImmediateColon___closed__1; lean_object* l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_dbgTraceStateFn___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_optionalFn___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Parser_info___default___elambda__1___boxed(lean_object*); LEAN_EXPORT uint8_t l_Lean_Parser_ParserContext_suppressInsideQuot___default; LEAN_EXPORT lean_object* l_Lean_Parser_ParserInfo_collectKinds___default(lean_object*); @@ -174,6 +175,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_antiquotNestedExpr___elambda__1(lean_obje 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; static lean_object* l_Lean_Parser_mkAntiquot___closed__9; +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Parser_octalNumberFn___lambda__1(uint32_t); LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_sepByFnAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_isIdRest___boxed(lean_object*); @@ -184,7 +186,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_orelseFn(lean_object*, lean_object*, lean LEAN_EXPORT lean_object* l_Std_RBNode_find___at_Lean_Parser_indexed___spec__4(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 uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_mkTokenAndFixPos___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9280____lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_mkTrailingNode(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_RBNode_ins___at_Lean_Parser_TokenMap_insert___spec__4(lean_object*); @@ -215,7 +216,6 @@ static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIden LEAN_EXPORT lean_object* l_Lean_Parser_mkNodeToken(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_qpartition_loop___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_runLongestMatchParser___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_nameLitNoAntiquot___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Error_unexpected___default; size_t lean_usize_shift_right(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Parser_symbolInfo___elambda__1(lean_object*); @@ -240,7 +240,6 @@ static lean_object* l_Lean_Parser_instInhabitedParserInfo___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_nameLitFn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_instInhabitedParserInfo___lambda__2___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_strLitFnAux___boxed(lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_nameLitKind; LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_sepByFnAux_parse___lambda__2(lean_object*, lean_object*, uint8_t, 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_8861____closed__9; static lean_object* l_Lean_Parser_decimalNumberFn_parseOptDot___closed__1; @@ -275,9 +274,9 @@ static lean_object* l_Lean_Parser_scientificLitFn___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_LeadingIdentBehavior_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_antiquotNestedExpr___elambda__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_instInhabitedParserCategory; -static lean_object* l_Lean_Parser_fieldIdx___closed__8; static lean_object* l_Lean_Parser_LeadingIdentBehavior_noConfusion___rarg___closed__1; static lean_object* l_Lean_Parser_scientificLitFn___closed__1; +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_rawIdentFn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_mkAntiquotSplice___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insert___at_Lean_Parser_SyntaxNodeKindSet_insert___spec__1(lean_object*, lean_object*, lean_object*); @@ -292,6 +291,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_isIdFirstO static lean_object* l_Lean_Parser_mkAntiquot___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_nonReservedSymbolNoAntiquot(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Syntax_forArgsM(lean_object*); +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_decimalNumberFn(lean_object*, lean_object*, lean_object*); @@ -345,7 +345,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_nonReservedSymbolFnAux(lean_object*, lean static lean_object* l_Lean_Parser_errorAtSavedPos___closed__2; static lean_object* l_Lean_Parser_dbgTraceStateFn___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_sepByElemParser(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_numLitNoAntiquot___closed__4; 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_8861_(uint8_t, lean_object*); @@ -364,7 +363,6 @@ static lean_object* l_Lean_Parser_rawIdentNoAntiquot___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_nodeInfo(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_checkLhsPrecFn(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_symbolFn___boxed(lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_numLitKind; LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_withAntiquotSuffixSpliceFn___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_LeadingIdentBehavior_noConfusion___rarg(uint8_t, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_LeadingIdentBehavior_noConfusion___rarg___lambda__1___boxed(lean_object*); @@ -384,7 +382,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_noFirstTokenInfo(lean_object*); static lean_object* l_Lean_Parser_antiquotExpr___elambda__1___closed__3; lean_object* lean_array_swap(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_toStringAux___at_Lean_Parser_dbgTraceStateFn___spec__2___boxed(lean_object*, lean_object*); -extern lean_object* l_Lean_strLitKind; LEAN_EXPORT lean_object* l_Lean_Parser_nonReservedSymbolInfo(lean_object*, uint8_t); static lean_object* l_Lean_Parser_tokenAntiquotFn___lambda__2___closed__8; LEAN_EXPORT lean_object* l_Lean_Parser_unicodeSymbol___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); @@ -426,6 +423,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_isQuotableCharDefault___boxed(lean_object LEAN_EXPORT lean_object* l_List_toStringAux___at_Lean_Parser_dbgTraceStateFn___spec__2(uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_orelseFnCore___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_checkLineEqFn(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_optionalFn___closed__2; LEAN_EXPORT lean_object* l_Std_RBNode_find___at_Lean_Parser_indexed___spec__5(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_symbolInfo___elambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_withPositionAfterLinebreak(lean_object*); @@ -506,15 +504,13 @@ LEAN_EXPORT lean_object* l_Lean_Parser_rawCh___boxed(lean_object*, lean_object*) 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; -extern lean_object* l_Lean_choiceKind; -extern lean_object* l_Lean_charLitKind; lean_object* l_Std_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); 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_Lean_Parser_withForbidden___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_foldArgsM___spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_strLitNoAntiquot___closed__4; lean_object* l_Array_back___rarg(lean_object*, lean_object*); +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_Std_RBNode_ins___at_Lean_Parser_TokenMap_insert___spec__6(lean_object*); @@ -539,6 +535,7 @@ static lean_object* l_Lean_Parser_ParserState_mkEOIError___closed__1; LEAN_EXPORT uint8_t l_Lean_Parser_isIdCont(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_mergeOrElseErrors___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_checkNoImmediateColon___elambda__1(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_decimalNumberFn___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_nonReservedSymbolInfo___elambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_TokenMap_instInhabitedTokenMap(lean_object*); uint32_t lean_string_utf8_get(lean_object*, lean_object*); @@ -578,6 +575,7 @@ static lean_object* l_Lean_Parser_nameLitNoAntiquot___closed__3; LEAN_EXPORT lean_object* l_Std_RBNode_find___at_Lean_Parser_indexed___spec__7___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_longestMatchMkResult(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_RBNode_find___at_Lean_Parser_indexed___spec__7___rarg___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_orelseFnCore___lambda__1___closed__1; size_t lean_usize_mul(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Parser_LeadingIdentBehavior_noConfusion___rarg___lambda__1(lean_object*); static lean_object* l_Lean_Parser_antiquotExpr___elambda__1___closed__2; @@ -594,7 +592,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_tokenAntiquotFn___lambda__2(lean_object*, static lean_object* l_Lean_Parser_quotedCharCoreFn___closed__1; 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; -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkAntiquot___closed__21; LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_restore(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*); @@ -603,10 +600,10 @@ LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_pickNonNon LEAN_EXPORT lean_object* l_Lean_Parser_sepBy1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Parser_Basic_0__Lean_Parser_isIdFirstOrBeginEscape(uint32_t); LEAN_EXPORT lean_object* l_Lean_Parser_binNumberFn___lambda__1___boxed(lean_object*); -static lean_object* l_Lean_Parser_charLitNoAntiquot___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_identFnAux_parse(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_popSyntax(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_termParser(lean_object*); +static lean_object* l_Lean_Parser_decimalNumberFn___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_sepByFn___boxed(lean_object*, lean_object*, lean_object*, 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*); @@ -657,6 +654,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_sepBy1Info(lean_object*, lean_object*); static lean_object* l_Lean_Parser_quotedCharFn___closed__1; static lean_object* l_Lean_Parser_charLitFnAux___closed__1; static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8861____closed__14; +static lean_object* l_Lean_Parser_orelseFnCore___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_hexDigitFn___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_mergeErrors___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_Error_expectedToString(lean_object*); @@ -679,14 +677,12 @@ LEAN_EXPORT lean_object* l_Lean_Parser_checkColGeFn(lean_object*, lean_object*, LEAN_EXPORT lean_object* l_Lean_Parser_charLitFnAux___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_prattParser___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_nodeWithAntiquot(lean_object*, lean_object*, lean_object*, uint8_t); -extern lean_object* l_Lean_identKind; LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_mkErrorAt(lean_object*, lean_object*, lean_object*, lean_object*); 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_Lean_Parser_symbol___boxed(lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9280____closed__1; -static lean_object* l_Lean_Parser_scientificLitNoAntiquot___closed__4; 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*); @@ -709,7 +705,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_mkTrailingNode___boxed(lean_o static lean_object* l_Lean_Parser_mkAntiquotSplice___closed__6; lean_object* l_Lean_Syntax_getNumArgs(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_forArgsM___spec__1___at_Lean_Syntax_forArgsM___spec__2___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, size_t, lean_object*); -extern lean_object* l_Lean_scientificLitKind; +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_trailingLoop___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_sepByNoAntiquot___boxed(lean_object*, lean_object*, lean_object*); @@ -742,6 +738,7 @@ 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*); 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*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_foldArgsM___spec__1___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_chFn___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -775,12 +772,12 @@ LEAN_EXPORT lean_object* l_Lean_Parser_orelseFnCore(lean_object*, lean_object*, static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8861____closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_epsilonInfo; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAux_traverse___at_Lean_Parser_SyntaxNodeKindSet_insert___spec__3(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_strLitFnAux___closed__2; LEAN_EXPORT lean_object* l_List_toString___at_Lean_Parser_FirstTokens_toStr___spec__1(lean_object*); static lean_object* l_Lean_Parser_epsilonInfo___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_nonReservedSymbolInfo___elambda__2___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_instInhabitedParser; LEAN_EXPORT lean_object* l_Lean_Parser_categoryParser___elambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_mkTokenAndFixPos___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_pushNone___elambda__1(lean_object*); 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*); @@ -791,6 +788,7 @@ LEAN_EXPORT lean_object* l_Std_RBNode_find___at_Lean_Parser_indexed___spec__4___ LEAN_EXPORT lean_object* l_Lean_Parser_instBEqError; static lean_object* l_Lean_Parser_mkAntiquot___closed__19; LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_sepByFnAux_parse___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_mkNameLit(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_rawAux(lean_object*, uint8_t, lean_object*, lean_object*); lean_object* l_UInt32_decEq___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_decimalNumberFn_parseOptDot___boxed(lean_object*, lean_object*); @@ -829,6 +827,7 @@ static lean_object* l_Lean_Parser_checkNoImmediateColon___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_sepBy1NoAntiquot(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Parser_instInhabitedParserFn___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_symbol___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_nameLitFn___closed__4; 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*); @@ -847,6 +846,7 @@ static lean_object* l_Lean_Parser_nameLitFn___closed__1; LEAN_EXPORT lean_object* l_Std_RBNode_find___at_Lean_Parser_indexed___spec__1___rarg(lean_object*, lean_object*); 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 uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_mkTokenAndFixPos___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_nonReservedSymbolInfo___closed__4; static lean_object* l_Lean_Parser_instInhabitedInputContext___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_strAux_parse___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -870,7 +870,6 @@ static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9301____ LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9280_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9301_(lean_object*); LEAN_EXPORT uint8_t l_Lean_Parser_instInhabitedLeadingIdentBehavior; -static lean_object* l_Lean_Parser_fieldIdx___closed__7; 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*); @@ -952,13 +951,15 @@ lean_object* l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg(lean_object*, le LEAN_EXPORT lean_object* l_Lean_Parser_indexed___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); 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; +static lean_object* l_Lean_Parser_charLitFnAux___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_keepPrevError(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkAntiquot___closed__25; LEAN_EXPORT lean_object* l_Lean_Parser_checkWsBefore___elambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_sepBy1___elambda__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_sepByFnAux_parse___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(lean_object*, lean_object*); 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___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8861____closed__1; static lean_object* l_Lean_Parser_unicodeSymbolInfo___closed__1; @@ -987,6 +988,7 @@ static lean_object* l_Lean_Parser_antiquotExpr___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_charLitNoAntiquot; static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8861____closed__17; lean_object* l_Lean_mkErrorStringWithPos(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_fieldIdxFn___closed__2; 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*); @@ -2549,7 +2551,7 @@ x_1 = lean_box(0); return x_1; } } -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -2592,7 +2594,7 @@ LEAN_EXPORT uint8_t l_Lean_Parser_ParserState_hasError(lean_object* x_1) { lean_object* x_2; lean_object* x_3; uint8_t x_4; x_2 = lean_ctor_get(x_1, 4); x_3 = lean_box(0); -x_4 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_2, x_3); +x_4 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_2, x_3); if (x_4 == 0) { uint8_t x_5; @@ -2607,11 +2609,11 @@ return x_6; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_1, x_2); +x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -2992,7 +2994,7 @@ lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; x_5 = lean_ctor_get(x_1, 0); x_6 = lean_ctor_get(x_1, 4); x_7 = lean_box(0); -x_8 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_6, x_7); +x_8 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_6, x_7); if (x_8 == 0) { lean_object* x_9; uint8_t x_10; @@ -3061,7 +3063,7 @@ lean_inc(x_25); lean_inc(x_24); lean_dec(x_1); x_29 = lean_box(0); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_29); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_29); if (x_30 == 0) { lean_object* x_31; uint8_t x_32; @@ -5018,7 +5020,7 @@ x_5 = lean_apply_2(x_1, x_3, x_4); x_6 = lean_ctor_get(x_5, 4); lean_inc(x_6); x_7 = lean_box(0); -x_8 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_6, x_7); +x_8 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_6, x_7); lean_dec(x_6); if (x_8 == 0) { @@ -5673,7 +5675,7 @@ lean_inc(x_6); x_7 = lean_ctor_get(x_3, 4); lean_inc(x_7); x_8 = lean_box(0); -x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); +x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); if (x_9 == 0) { lean_dec(x_7); @@ -6247,12 +6249,30 @@ lean_dec(x_3); return x_6; } } +static lean_object* _init_l_Lean_Parser_orelseFnCore___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("choice", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_orelseFnCore___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_orelseFnCore___lambda__1___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Parser_orelseFnCore___lambda__1(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; lean_object* x_9; lean_object* x_10; x_6 = l_Lean_Parser_ParserState_pushSyntax(x_4, x_1); -x_7 = l_Lean_choiceKind; +x_7 = l_Lean_Parser_orelseFnCore___lambda__1___closed__2; x_8 = l_Lean_Parser_ParserState_mkNode(x_6, x_7, x_2); x_9 = lean_box(0); x_10 = lean_apply_2(x_3, x_8, x_9); @@ -6329,7 +6349,7 @@ x_19 = lean_apply_2(x_2, x_4, x_5); x_20 = lean_ctor_get(x_19, 4); lean_inc(x_20); x_21 = lean_box(0); -x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_20, x_21); +x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_20, x_21); lean_dec(x_20); if (x_22 == 0) { @@ -6365,7 +6385,7 @@ return x_9; else { lean_object* x_26; uint8_t x_27; -x_26 = l_Lean_choiceKind; +x_26 = l_Lean_Parser_orelseFnCore___lambda__1___closed__2; lean_inc(x_16); x_27 = l_Lean_Syntax_isOfKind(x_16, x_26); if (x_27 == 0) @@ -6743,6 +6763,24 @@ return x_8; } } } +static lean_object* _init_l_Lean_Parser_optionalFn___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("null", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_optionalFn___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_optionalFn___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Parser_optionalFn(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { @@ -6757,7 +6795,7 @@ x_7 = lean_apply_2(x_1, x_2, x_3); x_8 = lean_ctor_get(x_7, 4); lean_inc(x_8); x_9 = lean_box(0); -x_10 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_8, x_9); +x_10 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_8, x_9); lean_dec(x_8); if (x_10 == 0) { @@ -6770,7 +6808,7 @@ if (x_12 == 0) { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l_Lean_nullKind; +x_13 = l_Lean_Parser_optionalFn___closed__2; x_14 = l_Lean_Parser_ParserState_mkNode(x_7, x_13, x_5); return x_14; } @@ -6778,7 +6816,7 @@ else { lean_object* x_15; lean_object* x_16; lean_object* x_17; x_15 = l_Lean_Parser_ParserState_restore(x_7, x_5, x_6); -x_16 = l_Lean_nullKind; +x_16 = l_Lean_Parser_optionalFn___closed__2; x_17 = l_Lean_Parser_ParserState_mkNode(x_15, x_16, x_5); return x_17; } @@ -6787,7 +6825,7 @@ else { lean_object* x_18; lean_object* x_19; lean_dec(x_6); -x_18 = l_Lean_nullKind; +x_18 = l_Lean_Parser_optionalFn___closed__2; x_19 = l_Lean_Parser_ParserState_mkNode(x_7, x_18, x_5); return x_19; } @@ -6868,7 +6906,7 @@ x_7 = lean_apply_2(x_1, x_2, x_3); x_8 = lean_ctor_get(x_7, 4); lean_inc(x_8); x_9 = lean_box(0); -x_10 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_8, x_9); +x_10 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_8, x_9); lean_dec(x_8); if (x_10 == 0) { @@ -6938,7 +6976,7 @@ x_8 = lean_apply_2(x_1, x_3, x_4); x_9 = lean_ctor_get(x_8, 4); lean_inc(x_9); x_10 = lean_box(0); -x_11 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_9, x_10); +x_11 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_9, x_10); lean_dec(x_9); if (x_11 == 0) { @@ -7019,7 +7057,7 @@ return x_11; else { lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = l_Lean_nullKind; +x_12 = l_Lean_Parser_optionalFn___closed__2; x_13 = l_Lean_Parser_ParserState_mkNode(x_4, x_12, x_3); x_14 = l_Lean_Parser_manyAux(x_1, x_2, x_13); return x_14; @@ -7079,7 +7117,7 @@ x_7 = lean_apply_2(x_1, x_2, x_3); x_8 = lean_ctor_get(x_7, 4); lean_inc(x_8); x_9 = lean_box(0); -x_10 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_8, x_9); +x_10 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_8, x_9); lean_dec(x_8); if (x_10 == 0) { @@ -7149,7 +7187,7 @@ lean_inc(x_4); x_5 = lean_array_get_size(x_4); lean_dec(x_4); x_6 = l_Lean_Parser_manyAux(x_1, x_2, x_3); -x_7 = l_Lean_nullKind; +x_7 = l_Lean_Parser_optionalFn___closed__2; x_8 = l_Lean_Parser_ParserState_mkNode(x_6, x_7, x_5); return x_8; } @@ -7186,14 +7224,14 @@ x_6 = lean_apply_2(x_1, x_2, x_3); x_7 = lean_ctor_get(x_6, 4); lean_inc(x_7); x_8 = lean_box(0); -x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); +x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); lean_dec(x_7); if (x_9 == 0) { lean_object* x_10; lean_object* x_11; lean_dec(x_2); lean_dec(x_1); -x_10 = l_Lean_nullKind; +x_10 = l_Lean_Parser_optionalFn___closed__2; x_11 = l_Lean_Parser_ParserState_mkNode(x_6, x_10, x_5); return x_11; } @@ -7201,7 +7239,7 @@ else { lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = l_Lean_Parser_manyAux(x_1, x_2, x_6); -x_13 = l_Lean_nullKind; +x_13 = l_Lean_Parser_optionalFn___closed__2; x_14 = l_Lean_Parser_ParserState_mkNode(x_12, x_13, x_5); return x_14; } @@ -7270,7 +7308,7 @@ return x_14; else { lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = l_Lean_nullKind; +x_15 = l_Lean_Parser_optionalFn___closed__2; x_16 = l_Lean_Parser_ParserState_mkNode(x_7, x_15, x_6); x_17 = l___private_Lean_Parser_Basic_0__Lean_Parser_sepByFnAux_parse(x_1, x_2, x_3, x_4, x_3, x_5, x_16); return x_17; @@ -7294,7 +7332,7 @@ x_11 = lean_apply_2(x_1, x_2, x_6); x_12 = lean_ctor_get(x_11, 4); lean_inc(x_12); x_13 = lean_box(0); -x_14 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_12, x_13); +x_14 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_12, x_13); lean_dec(x_12); if (x_14 == 0) { @@ -7304,7 +7342,7 @@ lean_dec(x_2); lean_dec(x_1); x_15 = l_Lean_Parser_ParserState_restore(x_11, x_9, x_10); lean_dec(x_9); -x_16 = l_Lean_nullKind; +x_16 = l_Lean_Parser_optionalFn___closed__2; x_17 = l_Lean_Parser_ParserState_mkNode(x_15, x_16, x_5); return x_17; } @@ -7342,7 +7380,7 @@ return x_15; else { lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_16 = l_Lean_nullKind; +x_16 = l_Lean_Parser_optionalFn___closed__2; x_17 = l_Lean_Parser_ParserState_mkNode(x_7, x_16, x_6); x_18 = lean_box(0); x_19 = l___private_Lean_Parser_Basic_0__Lean_Parser_sepByFnAux_parse___lambda__3(x_1, x_2, x_3, x_4, x_5, x_17, x_18); @@ -7366,7 +7404,7 @@ x_11 = lean_apply_2(x_1, x_6, x_7); x_12 = lean_ctor_get(x_11, 4); lean_inc(x_12); x_13 = lean_box(0); -x_14 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_12, x_13); +x_14 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_12, x_13); lean_dec(x_12); if (x_14 == 0) { @@ -7387,7 +7425,7 @@ lean_dec(x_10); lean_dec(x_9); x_17 = lean_box(0); x_18 = l_Lean_Parser_ParserState_pushSyntax(x_11, x_17); -x_19 = l_Lean_nullKind; +x_19 = l_Lean_Parser_optionalFn___closed__2; x_20 = l_Lean_Parser_ParserState_mkNode(x_18, x_19, x_4); return x_20; } @@ -7396,7 +7434,7 @@ else lean_object* x_21; lean_object* x_22; lean_object* x_23; x_21 = l_Lean_Parser_ParserState_restore(x_11, x_9, x_10); lean_dec(x_9); -x_22 = l_Lean_nullKind; +x_22 = l_Lean_Parser_optionalFn___closed__2; x_23 = l_Lean_Parser_ParserState_mkNode(x_21, x_22, x_4); return x_23; } @@ -7406,7 +7444,7 @@ else lean_object* x_24; lean_object* x_25; lean_dec(x_10); lean_dec(x_9); -x_24 = l_Lean_nullKind; +x_24 = l_Lean_Parser_optionalFn___closed__2; x_25 = l_Lean_Parser_ParserState_mkNode(x_11, x_24, x_4); return x_25; } @@ -7739,7 +7777,7 @@ x_5 = lean_apply_2(x_1, x_3, x_4); x_6 = lean_ctor_get(x_5, 4); lean_inc(x_6); x_7 = lean_box(0); -x_8 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_6, x_7); +x_8 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_6, x_7); lean_dec(x_6); if (x_8 == 0) { @@ -8047,7 +8085,7 @@ x_5 = l_Lean_Parser_satisfyFn(x_1, x_2, x_3, x_4); x_6 = lean_ctor_get(x_5, 4); lean_inc(x_6); x_7 = lean_box(0); -x_8 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_6, x_7); +x_8 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_6, x_7); lean_dec(x_6); if (x_8 == 0) { @@ -8323,7 +8361,7 @@ x_25 = l_Lean_Parser_finishCommentBlock(x_24, x_1, x_23); x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); x_27 = lean_box(0); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_27); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_27); lean_dec(x_26); if (x_28 == 0) { @@ -8371,7 +8409,7 @@ x_35 = l_Lean_Parser_takeUntilFn(x_34, x_1, x_33); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); x_37 = lean_box(0); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_37); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_37); lean_dec(x_36); if (x_38 == 0) { @@ -8530,7 +8568,7 @@ x_6 = lean_apply_2(x_1, x_3, x_4); x_7 = lean_ctor_get(x_6, 4); lean_inc(x_7); x_8 = lean_box(0); -x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); +x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); lean_dec(x_7); if (x_9 == 0) { @@ -8584,7 +8622,7 @@ x_13 = l_Lean_Parser_satisfyFn(x_6, x_11, x_3, x_4); x_14 = lean_ctor_get(x_13, 4); lean_inc(x_14); x_15 = lean_box(0); -x_16 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_14, x_15); +x_16 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_14, x_15); lean_dec(x_14); if (x_16 == 0) { @@ -8631,7 +8669,7 @@ x_13 = l_Lean_Parser_satisfyFn(x_6, x_11, x_3, x_4); x_14 = lean_ctor_get(x_13, 4); lean_inc(x_14); x_15 = lean_box(0); -x_16 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_14, x_15); +x_16 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_14, x_15); lean_dec(x_14); if (x_16 == 0) { @@ -8860,7 +8898,7 @@ x_20 = l_Lean_Parser_hexDigitFn(x_2, x_19); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); x_22 = lean_box(0); -x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_22); +x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_22); lean_dec(x_21); if (x_23 == 0) { @@ -8872,7 +8910,7 @@ lean_object* x_24; lean_object* x_25; uint8_t x_26; x_24 = l_Lean_Parser_hexDigitFn(x_2, x_20); x_25 = lean_ctor_get(x_24, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_22); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_22); lean_dec(x_25); if (x_26 == 0) { @@ -8884,7 +8922,7 @@ lean_object* x_27; lean_object* x_28; uint8_t x_29; x_27 = l_Lean_Parser_hexDigitFn(x_2, x_24); x_28 = lean_ctor_get(x_27, 4); lean_inc(x_28); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_22); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_22); lean_dec(x_28); if (x_29 == 0) { @@ -8909,7 +8947,7 @@ x_32 = l_Lean_Parser_hexDigitFn(x_2, x_31); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); x_34 = lean_box(0); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_34); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_34); lean_dec(x_33); if (x_35 == 0) { @@ -9109,6 +9147,24 @@ x_1 = lean_mk_string_from_bytes("missing end of character literal", 32); return x_1; } } +static lean_object* _init_l_Lean_Parser_charLitFnAux___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("char", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_charLitFnAux___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_charLitFnAux___closed__2; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Parser_charLitFnAux(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { @@ -9136,7 +9192,7 @@ lean_object* x_14; uint8_t x_15; x_14 = lean_ctor_get(x_3, 4); lean_inc(x_14); lean_dec(x_3); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_14, x_13); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_14, x_13); lean_dec(x_14); if (x_15 == 0) { @@ -9165,7 +9221,7 @@ return x_23; else { lean_object* x_24; lean_object* x_25; -x_24 = l_Lean_charLitKind; +x_24 = l_Lean_Parser_charLitFnAux___closed__3; x_25 = l_Lean_Parser_mkNodeToken(x_24, x_1, x_2, x_18); return x_25; } @@ -9180,7 +9236,7 @@ x_26 = l_Lean_Parser_quotedCharFn___closed__1; x_27 = l_Lean_Parser_quotedCharCoreFn(x_26, x_2, x_10); x_28 = lean_ctor_get(x_27, 4); lean_inc(x_28); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_13); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_13); lean_dec(x_28); if (x_29 == 0) { @@ -9210,7 +9266,7 @@ return x_38; else { lean_object* x_39; lean_object* x_40; -x_39 = l_Lean_charLitKind; +x_39 = l_Lean_Parser_charLitFnAux___closed__3; x_40 = l_Lean_Parser_mkNodeToken(x_39, x_1, x_2, x_33); return x_40; } @@ -9242,6 +9298,24 @@ static lean_object* _init_l_Lean_Parser_strLitFnAux___closed__1() { _start: { lean_object* x_1; +x_1 = lean_mk_string_from_bytes("str", 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_strLitFnAux___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_strLitFnAux___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_strLitFnAux___closed__3() { +_start: +{ +lean_object* x_1; x_1 = lean_mk_string_from_bytes("unterminated string literal", 27); return x_1; } @@ -9282,7 +9356,7 @@ x_17 = l_Lean_Parser_quotedCharCoreFn(x_16, x_2, x_10); x_18 = lean_ctor_get(x_17, 4); lean_inc(x_18); x_19 = lean_box(0); -x_20 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_18, x_19); +x_20 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_18, x_19); lean_dec(x_18); if (x_20 == 0) { @@ -9299,7 +9373,7 @@ goto _start; else { lean_object* x_22; lean_object* x_23; -x_22 = l_Lean_strLitKind; +x_22 = l_Lean_Parser_strLitFnAux___closed__2; x_23 = l_Lean_Parser_mkNodeToken(x_22, x_1, x_2, x_10); return x_23; } @@ -9308,7 +9382,7 @@ else { lean_object* x_24; lean_object* x_25; lean_dec(x_6); -x_24 = l_Lean_Parser_strLitFnAux___closed__1; +x_24 = l_Lean_Parser_strLitFnAux___closed__3; x_25 = l_Lean_Parser_ParserState_mkUnexpectedErrorAt(x_3, x_24, x_1); return x_25; } @@ -9478,6 +9552,42 @@ lean_dec(x_1); return x_3; } } +static lean_object* _init_l_Lean_Parser_decimalNumberFn___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("num", 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_decimalNumberFn___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_decimalNumberFn___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_decimalNumberFn___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("scientific", 10); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_decimalNumberFn___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_decimalNumberFn___closed__3; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Parser_decimalNumberFn(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { @@ -9505,7 +9615,7 @@ x_15 = lean_uint32_dec_eq(x_9, x_14); if (x_15 == 0) { lean_object* x_16; lean_object* x_17; -x_16 = l_Lean_numLitKind; +x_16 = l_Lean_Parser_decimalNumberFn___closed__2; x_17 = l_Lean_Parser_mkNodeToken(x_16, x_1, x_2, x_5); return x_17; } @@ -9514,7 +9624,7 @@ else lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; x_18 = l_Lean_Parser_decimalNumberFn_parseOptDot(x_2, x_5); x_19 = l_Lean_Parser_decimalNumberFn_parseOptExp(x_2, x_18); -x_20 = l_Lean_scientificLitKind; +x_20 = l_Lean_Parser_decimalNumberFn___closed__4; x_21 = l_Lean_Parser_mkNodeToken(x_20, x_1, x_2, x_19); return x_21; } @@ -9524,7 +9634,7 @@ else lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; x_22 = l_Lean_Parser_decimalNumberFn_parseOptDot(x_2, x_5); x_23 = l_Lean_Parser_decimalNumberFn_parseOptExp(x_2, x_22); -x_24 = l_Lean_scientificLitKind; +x_24 = l_Lean_Parser_decimalNumberFn___closed__4; x_25 = l_Lean_Parser_mkNodeToken(x_24, x_1, x_2, x_23); return x_25; } @@ -9534,7 +9644,7 @@ else lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; x_26 = l_Lean_Parser_decimalNumberFn_parseOptDot(x_2, x_5); x_27 = l_Lean_Parser_decimalNumberFn_parseOptExp(x_2, x_26); -x_28 = l_Lean_scientificLitKind; +x_28 = l_Lean_Parser_decimalNumberFn___closed__4; x_29 = l_Lean_Parser_mkNodeToken(x_28, x_1, x_2, x_27); return x_29; } @@ -9596,12 +9706,12 @@ x_6 = l_Lean_Parser_satisfyFn(x_4, x_5, x_2, x_3); x_7 = lean_ctor_get(x_6, 4); lean_inc(x_7); x_8 = lean_box(0); -x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); +x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); lean_dec(x_7); if (x_9 == 0) { lean_object* x_10; lean_object* x_11; -x_10 = l_Lean_numLitKind; +x_10 = l_Lean_Parser_decimalNumberFn___closed__2; x_11 = l_Lean_Parser_mkNodeToken(x_10, x_1, x_2, x_6); return x_11; } @@ -9609,7 +9719,7 @@ else { lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = l_Lean_Parser_takeWhileFn(x_4, x_2, x_6); -x_13 = l_Lean_numLitKind; +x_13 = l_Lean_Parser_decimalNumberFn___closed__2; x_14 = l_Lean_Parser_mkNodeToken(x_13, x_1, x_2, x_12); return x_14; } @@ -9682,12 +9792,12 @@ x_6 = l_Lean_Parser_satisfyFn(x_4, x_5, x_2, x_3); x_7 = lean_ctor_get(x_6, 4); lean_inc(x_7); x_8 = lean_box(0); -x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); +x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); lean_dec(x_7); if (x_9 == 0) { lean_object* x_10; lean_object* x_11; -x_10 = l_Lean_numLitKind; +x_10 = l_Lean_Parser_decimalNumberFn___closed__2; x_11 = l_Lean_Parser_mkNodeToken(x_10, x_1, x_2, x_6); return x_11; } @@ -9695,7 +9805,7 @@ else { lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = l_Lean_Parser_takeWhileFn(x_4, x_2, x_6); -x_13 = l_Lean_numLitKind; +x_13 = l_Lean_Parser_decimalNumberFn___closed__2; x_14 = l_Lean_Parser_mkNodeToken(x_13, x_1, x_2, x_12); return x_14; } @@ -9838,12 +9948,12 @@ x_6 = l_Lean_Parser_satisfyFn(x_4, x_5, x_2, x_3); x_7 = lean_ctor_get(x_6, 4); lean_inc(x_7); x_8 = lean_box(0); -x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); +x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); lean_dec(x_7); if (x_9 == 0) { lean_object* x_10; lean_object* x_11; -x_10 = l_Lean_numLitKind; +x_10 = l_Lean_Parser_decimalNumberFn___closed__2; x_11 = l_Lean_Parser_mkNodeToken(x_10, x_1, x_2, x_6); return x_11; } @@ -9851,7 +9961,7 @@ else { lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = l_Lean_Parser_takeWhileFn(x_4, x_2, x_6); -x_13 = l_Lean_numLitKind; +x_13 = l_Lean_Parser_decimalNumberFn___closed__2; x_14 = l_Lean_Parser_mkNodeToken(x_13, x_1, x_2, x_12); return x_14; } @@ -10128,7 +10238,7 @@ x_5 = lean_box(x_4); return x_5; } } -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_mkTokenAndFixPos___spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_mkTokenAndFixPos___spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -10202,7 +10312,7 @@ lean_object* x_9; lean_object* x_10; uint8_t x_11; x_9 = lean_ctor_get(x_2, 0); x_10 = lean_ctor_get(x_3, 6); lean_inc(x_9); -x_11 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_mkTokenAndFixPos___spec__1(x_10, x_2); +x_11 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_mkTokenAndFixPos___spec__1(x_10, x_2); lean_dec(x_2); if (x_11 == 0) { @@ -10260,7 +10370,7 @@ x_28 = lean_ctor_get(x_3, 6); lean_inc(x_27); x_29 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_29, 0, x_27); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_mkTokenAndFixPos___spec__1(x_28, x_29); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_mkTokenAndFixPos___spec__1(x_28, x_29); lean_dec(x_29); if (x_30 == 0) { @@ -10311,11 +10421,11 @@ return x_45; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_mkTokenAndFixPos___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_mkTokenAndFixPos___spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_mkTokenAndFixPos___spec__1(x_1, x_2); +x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_mkTokenAndFixPos___spec__1(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -10652,7 +10762,7 @@ x_8 = lean_box(0); x_9 = l_Lean_Parser_identFnAux_parse(x_1, x_6, x_8, x_2, x_7); x_10 = lean_ctor_get(x_9, 4); lean_inc(x_10); -x_11 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_10, x_6); +x_11 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_10, x_6); lean_dec(x_10); if (x_11 == 0) { @@ -10668,7 +10778,7 @@ x_14 = l_Array_back___rarg(x_13, x_12); lean_dec(x_12); if (lean_obj_tag(x_14) == 3) { -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; +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_14, 0); lean_inc(x_15); x_16 = lean_ctor_get(x_14, 1); @@ -10686,18 +10796,17 @@ x_21 = lean_string_utf8_extract(x_18, x_19, x_20); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); -x_22 = l_Lean_nameLitKind; -x_23 = l_Lean_Syntax_mkLit(x_22, x_21, x_15); -x_24 = l_Lean_Parser_ParserState_pushSyntax(x_17, x_23); -return x_24; +x_22 = l_Lean_Syntax_mkNameLit(x_21, x_15); +x_23 = l_Lean_Parser_ParserState_pushSyntax(x_17, x_22); +return x_23; } else { -lean_object* x_25; lean_object* x_26; +lean_object* x_24; lean_object* x_25; lean_dec(x_14); -x_25 = l___private_Lean_Parser_Basic_0__Lean_Parser_nameLitAux___closed__1; -x_26 = l_Lean_Parser_ParserState_mkError(x_9, x_25); -return x_26; +x_24 = l___private_Lean_Parser_Basic_0__Lean_Parser_nameLitAux___closed__1; +x_25 = l_Lean_Parser_ParserState_mkError(x_9, x_24); +return x_25; } } } @@ -11079,7 +11188,7 @@ x_8 = l_Lean_Parser_tokenFn(x_2, x_3, x_4); x_9 = lean_ctor_get(x_8, 4); lean_inc(x_9); x_10 = lean_box(0); -x_11 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_9, x_10); +x_11 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_9, x_10); lean_dec(x_9); if (x_11 == 0) { @@ -11154,7 +11263,7 @@ x_10 = l_Lean_Parser_tokenFn(x_6, x_3, x_4); x_11 = lean_ctor_get(x_10, 4); lean_inc(x_11); x_12 = lean_box(0); -x_13 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_11, x_12); +x_13 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_11, x_12); lean_dec(x_11); if (x_13 == 0) { @@ -11374,7 +11483,7 @@ x_10 = l_Lean_Parser_tokenFn(x_9, x_3, x_4); x_11 = lean_ctor_get(x_10, 4); lean_inc(x_11); x_12 = lean_box(0); -x_13 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_11, x_12); +x_13 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_11, x_12); lean_dec(x_11); if (x_13 == 0) { @@ -12071,7 +12180,7 @@ x_9 = l_Lean_Parser_tokenFn(x_3, x_4, x_5); x_10 = lean_ctor_get(x_9, 4); lean_inc(x_10); x_11 = lean_box(0); -x_12 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_10, x_11); +x_12 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_10, x_11); lean_dec(x_10); if (x_12 == 0) { @@ -12368,7 +12477,7 @@ x_7 = l_Lean_Parser_tokenFn(x_6, x_1, x_2); x_8 = lean_ctor_get(x_7, 4); lean_inc(x_8); x_9 = lean_box(0); -x_10 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_8, x_9); +x_10 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_8, x_9); lean_dec(x_8); if (x_10 == 0) { @@ -12384,7 +12493,7 @@ lean_inc(x_11); x_12 = l_Lean_instInhabitedSyntax; x_13 = l_Array_back___rarg(x_12, x_11); lean_dec(x_11); -x_14 = l_Lean_numLitKind; +x_14 = l_Lean_Parser_decimalNumberFn___closed__2; x_15 = l_Lean_Syntax_isOfKind(x_13, x_14); if (x_15 == 0) { @@ -12407,21 +12516,13 @@ return x_7; static lean_object* _init_l_Lean_Parser_numLitNoAntiquot___closed__1() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("num", 3); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_numLitNoAntiquot___closed__2() { -_start: -{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_numLitNoAntiquot___closed__1; +x_1 = l_Lean_Parser_decimalNumberFn___closed__1; x_2 = l_Lean_Parser_mkAtomicInfo(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_numLitNoAntiquot___closed__3() { +static lean_object* _init_l_Lean_Parser_numLitNoAntiquot___closed__2() { _start: { lean_object* x_1; @@ -12429,12 +12530,12 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_numLitFn), 2, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_numLitNoAntiquot___closed__4() { +static lean_object* _init_l_Lean_Parser_numLitNoAntiquot___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_numLitNoAntiquot___closed__2; -x_2 = l_Lean_Parser_numLitNoAntiquot___closed__3; +x_1 = l_Lean_Parser_numLitNoAntiquot___closed__1; +x_2 = l_Lean_Parser_numLitNoAntiquot___closed__2; x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -12445,7 +12546,7 @@ static lean_object* _init_l_Lean_Parser_numLitNoAntiquot() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_numLitNoAntiquot___closed__4; +x_1 = l_Lean_Parser_numLitNoAntiquot___closed__3; return x_1; } } @@ -12484,7 +12585,7 @@ x_7 = l_Lean_Parser_tokenFn(x_6, x_1, x_2); x_8 = lean_ctor_get(x_7, 4); lean_inc(x_8); x_9 = lean_box(0); -x_10 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_8, x_9); +x_10 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_8, x_9); lean_dec(x_8); if (x_10 == 0) { @@ -12500,7 +12601,7 @@ lean_inc(x_11); x_12 = l_Lean_instInhabitedSyntax; x_13 = l_Array_back___rarg(x_12, x_11); lean_dec(x_11); -x_14 = l_Lean_scientificLitKind; +x_14 = l_Lean_Parser_decimalNumberFn___closed__4; x_15 = l_Lean_Syntax_isOfKind(x_13, x_14); if (x_15 == 0) { @@ -12523,21 +12624,13 @@ return x_7; static lean_object* _init_l_Lean_Parser_scientificLitNoAntiquot___closed__1() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("scientific", 10); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_scientificLitNoAntiquot___closed__2() { -_start: -{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_scientificLitNoAntiquot___closed__1; +x_1 = l_Lean_Parser_decimalNumberFn___closed__3; x_2 = l_Lean_Parser_mkAtomicInfo(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_scientificLitNoAntiquot___closed__3() { +static lean_object* _init_l_Lean_Parser_scientificLitNoAntiquot___closed__2() { _start: { lean_object* x_1; @@ -12545,12 +12638,12 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_scientificLitFn), 2, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_scientificLitNoAntiquot___closed__4() { +static lean_object* _init_l_Lean_Parser_scientificLitNoAntiquot___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_scientificLitNoAntiquot___closed__2; -x_2 = l_Lean_Parser_scientificLitNoAntiquot___closed__3; +x_1 = l_Lean_Parser_scientificLitNoAntiquot___closed__1; +x_2 = l_Lean_Parser_scientificLitNoAntiquot___closed__2; x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -12561,7 +12654,7 @@ static lean_object* _init_l_Lean_Parser_scientificLitNoAntiquot() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_scientificLitNoAntiquot___closed__4; +x_1 = l_Lean_Parser_scientificLitNoAntiquot___closed__3; return x_1; } } @@ -12600,7 +12693,7 @@ x_7 = l_Lean_Parser_tokenFn(x_6, x_1, x_2); x_8 = lean_ctor_get(x_7, 4); lean_inc(x_8); x_9 = lean_box(0); -x_10 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_8, x_9); +x_10 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_8, x_9); lean_dec(x_8); if (x_10 == 0) { @@ -12616,7 +12709,7 @@ lean_inc(x_11); x_12 = l_Lean_instInhabitedSyntax; x_13 = l_Array_back___rarg(x_12, x_11); lean_dec(x_11); -x_14 = l_Lean_strLitKind; +x_14 = l_Lean_Parser_strLitFnAux___closed__2; x_15 = l_Lean_Syntax_isOfKind(x_13, x_14); if (x_15 == 0) { @@ -12639,21 +12732,13 @@ return x_7; static lean_object* _init_l_Lean_Parser_strLitNoAntiquot___closed__1() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("str", 3); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_strLitNoAntiquot___closed__2() { -_start: -{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_strLitNoAntiquot___closed__1; +x_1 = l_Lean_Parser_strLitFnAux___closed__1; x_2 = l_Lean_Parser_mkAtomicInfo(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_strLitNoAntiquot___closed__3() { +static lean_object* _init_l_Lean_Parser_strLitNoAntiquot___closed__2() { _start: { lean_object* x_1; @@ -12661,12 +12746,12 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_strLitFn), 2, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_strLitNoAntiquot___closed__4() { +static lean_object* _init_l_Lean_Parser_strLitNoAntiquot___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_strLitNoAntiquot___closed__2; -x_2 = l_Lean_Parser_strLitNoAntiquot___closed__3; +x_1 = l_Lean_Parser_strLitNoAntiquot___closed__1; +x_2 = l_Lean_Parser_strLitNoAntiquot___closed__2; x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -12677,7 +12762,7 @@ static lean_object* _init_l_Lean_Parser_strLitNoAntiquot() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_strLitNoAntiquot___closed__4; +x_1 = l_Lean_Parser_strLitNoAntiquot___closed__3; return x_1; } } @@ -12724,7 +12809,7 @@ x_7 = l_Lean_Parser_tokenFn(x_6, x_1, x_2); x_8 = lean_ctor_get(x_7, 4); lean_inc(x_8); x_9 = lean_box(0); -x_10 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_8, x_9); +x_10 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_8, x_9); lean_dec(x_8); if (x_10 == 0) { @@ -12740,7 +12825,7 @@ lean_inc(x_11); x_12 = l_Lean_instInhabitedSyntax; x_13 = l_Array_back___rarg(x_12, x_11); lean_dec(x_11); -x_14 = l_Lean_charLitKind; +x_14 = l_Lean_Parser_charLitFnAux___closed__3; x_15 = l_Lean_Syntax_isOfKind(x_13, x_14); if (x_15 == 0) { @@ -12763,21 +12848,13 @@ return x_7; static lean_object* _init_l_Lean_Parser_charLitNoAntiquot___closed__1() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("char", 4); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_charLitNoAntiquot___closed__2() { -_start: -{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_charLitNoAntiquot___closed__1; +x_1 = l_Lean_Parser_charLitFnAux___closed__2; x_2 = l_Lean_Parser_mkAtomicInfo(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_charLitNoAntiquot___closed__3() { +static lean_object* _init_l_Lean_Parser_charLitNoAntiquot___closed__2() { _start: { lean_object* x_1; @@ -12785,12 +12862,12 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_charLitFn), 2, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_charLitNoAntiquot___closed__4() { +static lean_object* _init_l_Lean_Parser_charLitNoAntiquot___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_charLitNoAntiquot___closed__2; -x_2 = l_Lean_Parser_charLitNoAntiquot___closed__3; +x_1 = l_Lean_Parser_charLitNoAntiquot___closed__1; +x_2 = l_Lean_Parser_charLitNoAntiquot___closed__2; x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -12801,7 +12878,7 @@ static lean_object* _init_l_Lean_Parser_charLitNoAntiquot() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_charLitNoAntiquot___closed__4; +x_1 = l_Lean_Parser_charLitNoAntiquot___closed__3; return x_1; } } @@ -12825,6 +12902,24 @@ lean_ctor_set(x_3, 1, x_1); return x_3; } } +static lean_object* _init_l_Lean_Parser_nameLitFn___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("name", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_nameLitFn___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_nameLitFn___closed__3; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Parser_nameLitFn(lean_object* x_1, lean_object* x_2) { _start: { @@ -12840,7 +12935,7 @@ x_7 = l_Lean_Parser_tokenFn(x_6, x_1, x_2); x_8 = lean_ctor_get(x_7, 4); lean_inc(x_8); x_9 = lean_box(0); -x_10 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_8, x_9); +x_10 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_8, x_9); lean_dec(x_8); if (x_10 == 0) { @@ -12856,7 +12951,7 @@ lean_inc(x_11); x_12 = l_Lean_instInhabitedSyntax; x_13 = l_Array_back___rarg(x_12, x_11); lean_dec(x_11); -x_14 = l_Lean_nameLitKind; +x_14 = l_Lean_Parser_nameLitFn___closed__4; x_15 = l_Lean_Syntax_isOfKind(x_13, x_14); if (x_15 == 0) { @@ -12879,21 +12974,13 @@ return x_7; static lean_object* _init_l_Lean_Parser_nameLitNoAntiquot___closed__1() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("name", 4); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_nameLitNoAntiquot___closed__2() { -_start: -{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_nameLitNoAntiquot___closed__1; +x_1 = l_Lean_Parser_nameLitFn___closed__3; x_2 = l_Lean_Parser_mkAtomicInfo(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_nameLitNoAntiquot___closed__3() { +static lean_object* _init_l_Lean_Parser_nameLitNoAntiquot___closed__2() { _start: { lean_object* x_1; @@ -12901,12 +12988,12 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_nameLitFn), 2, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_nameLitNoAntiquot___closed__4() { +static lean_object* _init_l_Lean_Parser_nameLitNoAntiquot___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_nameLitNoAntiquot___closed__2; -x_2 = l_Lean_Parser_nameLitNoAntiquot___closed__3; +x_1 = l_Lean_Parser_nameLitNoAntiquot___closed__1; +x_2 = l_Lean_Parser_nameLitNoAntiquot___closed__2; x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -12917,7 +13004,7 @@ static lean_object* _init_l_Lean_Parser_nameLitNoAntiquot() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_nameLitNoAntiquot___closed__4; +x_1 = l_Lean_Parser_nameLitNoAntiquot___closed__3; return x_1; } } @@ -12956,7 +13043,7 @@ x_7 = l_Lean_Parser_tokenFn(x_6, x_1, x_2); x_8 = lean_ctor_get(x_7, 4); lean_inc(x_8); x_9 = lean_box(0); -x_10 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_8, x_9); +x_10 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_8, x_9); lean_dec(x_8); if (x_10 == 0) { @@ -13080,7 +13167,7 @@ x_8 = l_Lean_Parser_tokenFn(x_7, x_2, x_3); x_9 = lean_ctor_get(x_8, 4); lean_inc(x_9); x_10 = lean_box(0); -x_11 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_9, x_10); +x_11 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_9, x_10); lean_dec(x_9); if (x_11 == 0) { @@ -13514,7 +13601,7 @@ lean_object* x_12; lean_object* x_13; uint8_t x_14; x_12 = lean_ctor_get(x_6, 4); lean_inc(x_12); x_13 = lean_box(0); -x_14 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_12, x_13); +x_14 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_12, x_13); lean_dec(x_12); if (x_14 == 0) { @@ -14183,7 +14270,7 @@ return x_2; else { lean_object* x_8; lean_object* x_9; -x_8 = l_Lean_choiceKind; +x_8 = l_Lean_Parser_orelseFnCore___lambda__1___closed__2; x_9 = l_Lean_Parser_ParserState_mkNode(x_2, x_8, x_1); return x_9; } @@ -25972,6 +26059,16 @@ x_2 = lean_alloc_closure((void*)(l_Std_RBNode_find___at_Lean_Parser_indexed___sp return x_2; } } +static lean_object* _init_l_Lean_Parser_indexed___rarg___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_nonReservedSymbolInfo___closed__3; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Parser_indexed___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4) { _start: { @@ -26190,7 +26287,7 @@ if (x_52 == 0) lean_object* x_53; lean_object* x_54; lean_object* x_55; x_53 = lean_ctor_get(x_5, 1); lean_dec(x_53); -x_54 = l_Lean_identKind; +x_54 = l_Lean_Parser_indexed___rarg___closed__1; x_55 = l_Std_RBNode_find___at_Lean_Parser_indexed___spec__3___rarg(x_1, x_54); if (lean_obj_tag(x_55) == 0) { @@ -26215,7 +26312,7 @@ lean_object* x_58; lean_object* x_59; lean_object* x_60; x_58 = lean_ctor_get(x_5, 0); lean_inc(x_58); lean_dec(x_5); -x_59 = l_Lean_identKind; +x_59 = l_Lean_Parser_indexed___rarg___closed__1; x_60 = l_Std_RBNode_find___at_Lean_Parser_indexed___spec__3___rarg(x_1, x_59); if (lean_obj_tag(x_60) == 0) { @@ -26256,7 +26353,7 @@ lean_dec(x_67); if (lean_obj_tag(x_68) == 0) { lean_object* x_69; lean_object* x_70; -x_69 = l_Lean_identKind; +x_69 = l_Lean_Parser_indexed___rarg___closed__1; x_70 = l_Std_RBNode_find___at_Lean_Parser_indexed___spec__5___rarg(x_1, x_69); if (lean_obj_tag(x_70) == 0) { @@ -26299,7 +26396,7 @@ lean_dec(x_75); if (lean_obj_tag(x_76) == 0) { lean_object* x_77; lean_object* x_78; -x_77 = l_Lean_identKind; +x_77 = l_Lean_Parser_indexed___rarg___closed__1; x_78 = l_Std_RBNode_find___at_Lean_Parser_indexed___spec__5___rarg(x_1, x_77); if (lean_obj_tag(x_78) == 0) { @@ -26352,7 +26449,7 @@ if (lean_obj_tag(x_88) == 0) { lean_object* x_89; lean_object* x_90; lean_dec(x_87); -x_89 = l_Lean_identKind; +x_89 = l_Lean_Parser_indexed___rarg___closed__1; x_90 = l_Std_RBNode_find___at_Lean_Parser_indexed___spec__7___rarg(x_1, x_89); if (lean_obj_tag(x_90) == 0) { @@ -26377,7 +26474,7 @@ lean_object* x_93; lean_object* x_94; uint8_t x_95; x_93 = lean_ctor_get(x_88, 0); lean_inc(x_93); lean_dec(x_88); -x_94 = l_Lean_identKind; +x_94 = l_Lean_Parser_indexed___rarg___closed__1; x_95 = lean_name_eq(x_87, x_94); lean_dec(x_87); if (x_95 == 0) @@ -26421,7 +26518,7 @@ if (lean_obj_tag(x_101) == 0) { lean_object* x_102; lean_object* x_103; lean_dec(x_100); -x_102 = l_Lean_identKind; +x_102 = l_Lean_Parser_indexed___rarg___closed__1; x_103 = l_Std_RBNode_find___at_Lean_Parser_indexed___spec__7___rarg(x_1, x_102); if (lean_obj_tag(x_103) == 0) { @@ -26450,7 +26547,7 @@ lean_object* x_108; lean_object* x_109; uint8_t x_110; x_108 = lean_ctor_get(x_101, 0); lean_inc(x_108); lean_dec(x_101); -x_109 = l_Lean_identKind; +x_109 = l_Lean_Parser_indexed___rarg___closed__1; x_110 = lean_name_eq(x_100, x_109); lean_dec(x_100); if (x_110 == 0) @@ -27140,7 +27237,7 @@ static lean_object* _init_l_Lean_Parser_pushNone___elambda__1___rarg___closed__1 { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean_nullKind; +x_2 = l_Lean_Parser_optionalFn___closed__2; x_3 = l_Lean_Parser_instInhabitedInputContext___closed__1; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -27314,7 +27411,7 @@ x_7 = l_Lean_Parser_symbolFnAux(x_5, x_6, x_1, x_2); x_8 = lean_ctor_get(x_7, 4); lean_inc(x_8); x_9 = lean_box(0); -x_10 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_8, x_9); +x_10 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_8, x_9); lean_dec(x_8); if (x_10 == 0) { @@ -27333,7 +27430,7 @@ lean_inc(x_1); x_15 = l___private_Lean_Parser_Basic_0__Lean_Parser_addQuotDepthFn(x_13, x_14, x_1, x_7); x_16 = lean_ctor_get(x_15, 4); lean_inc(x_16); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_16, x_9); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_16, x_9); lean_dec(x_16); if (x_17 == 0) { @@ -27680,7 +27777,7 @@ x_8 = l_Lean_Parser_checkNoWsBeforeFn(x_7, x_2, x_1); x_9 = lean_ctor_get(x_8, 4); lean_inc(x_9); x_10 = lean_box(0); -x_11 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_9, x_10); +x_11 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_9, x_10); lean_dec(x_9); if (x_11 == 0) { @@ -27697,7 +27794,7 @@ lean_inc(x_2); x_21 = l_Lean_Parser_symbolFnAux(x_19, x_20, x_2, x_8); x_22 = lean_ctor_get(x_21, 4); lean_inc(x_22); -x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_22, x_10); +x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_22, x_10); lean_dec(x_22); if (x_23 == 0) { @@ -27714,7 +27811,7 @@ lean_inc(x_2); x_26 = l_Lean_Parser_symbolFnAux(x_24, x_25, x_2, x_21); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_10); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_10); lean_dec(x_27); if (x_28 == 0) { @@ -27728,7 +27825,7 @@ lean_object* x_29; lean_object* x_30; uint8_t x_31; x_29 = l_Lean_Parser_checkNoWsBeforeFn(x_7, x_2, x_26); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_10); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_10); lean_dec(x_30); if (x_31 == 0) { @@ -27751,7 +27848,7 @@ goto block_18; lean_object* x_13; uint8_t x_14; x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); -x_14 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_10); +x_14 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_10); lean_dec(x_13); if (x_14 == 0) { @@ -27779,7 +27876,7 @@ lean_object* x_3; lean_object* x_4; uint8_t x_5; x_3 = lean_ctor_get(x_2, 4); lean_inc(x_3); x_4 = lean_box(0); -x_5 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_3, x_4); +x_5 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_3, x_4); lean_dec(x_3); if (x_5 == 0) { @@ -28109,7 +28206,7 @@ x_5 = l_Lean_Parser_checkPrecFn(x_4, x_2, x_3); x_6 = lean_ctor_get(x_5, 4); lean_inc(x_6); x_7 = lean_box(0); -x_8 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_6, x_7); +x_8 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_6, x_7); lean_dec(x_6); if (x_8 == 0) { @@ -28134,7 +28231,7 @@ x_5 = l_Lean_Parser_checkPrecFn(x_4, x_2, x_3); x_6 = lean_ctor_get(x_5, 4); lean_inc(x_6); x_7 = lean_box(0); -x_8 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_6, x_7); +x_8 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_6, x_7); lean_dec(x_6); if (x_8 == 0) { @@ -28653,7 +28750,7 @@ x_5 = l_Lean_Parser_checkPrecFn(x_4, x_2, x_3); x_6 = lean_ctor_get(x_5, 4); lean_inc(x_6); x_7 = lean_box(0); -x_8 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_6, x_7); +x_8 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_6, x_7); lean_dec(x_6); if (x_8 == 0) { @@ -28751,7 +28848,7 @@ x_4 = l_Lean_Parser_mkAntiquotSplice___closed__2; x_5 = l_Lean_Name_append(x_1, x_4); x_6 = lean_ctor_get(x_2, 0); lean_inc(x_6); -x_7 = l_Lean_nullKind; +x_7 = l_Lean_Parser_optionalFn___closed__2; x_8 = l_Lean_Parser_nodeInfo(x_7, x_6); x_9 = lean_ctor_get(x_2, 1); lean_inc(x_9); @@ -28879,7 +28976,7 @@ x_8 = lean_apply_2(x_2, x_3, x_4); x_9 = lean_ctor_get(x_8, 4); lean_inc(x_9); x_10 = lean_box(0); -x_11 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_9, x_10); +x_11 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_9, x_10); lean_dec(x_9); if (x_11 == 0) { @@ -28930,7 +29027,7 @@ x_7 = lean_apply_2(x_6, x_4, x_5); x_8 = lean_ctor_get(x_7, 4); lean_inc(x_8); x_9 = lean_box(0); -x_10 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_8, x_9); +x_10 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_8, x_9); lean_dec(x_8); if (x_10 == 0) { @@ -29463,7 +29560,7 @@ lean_dec(x_4); if (x_7 == 0) { lean_object* x_8; lean_object* x_9; -x_8 = l_Lean_nullKind; +x_8 = l_Lean_Parser_optionalFn___closed__2; x_9 = l_Lean_Parser_ParserState_mkNode(x_1, x_8, x_2); return x_9; } @@ -29535,7 +29632,7 @@ lean_dec(x_9); x_12 = lean_ctor_get(x_10, 4); lean_inc(x_12); x_13 = lean_box(0); -x_14 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_12, x_13); +x_14 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_12, x_13); lean_dec(x_12); if (x_14 == 0) { @@ -29674,7 +29771,7 @@ x_12 = l_Lean_Parser_trailingLoopStep(x_2, x_10, x_3, x_4, x_11); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -29774,7 +29871,7 @@ lean_dec(x_9); x_12 = lean_ctor_get(x_10, 4); lean_inc(x_12); x_13 = lean_box(0); -x_14 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_12, x_13); +x_14 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_12, x_13); lean_dec(x_12); if (x_14 == 0) { @@ -29858,7 +29955,7 @@ lean_inc(x_2); x_16 = l_Lean_Parser_leadingParserAux(x_1, x_2, x_3, x_5, x_6); x_17 = lean_ctor_get(x_16, 4); lean_inc(x_17); -x_18 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_15); +x_18 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_15); lean_dec(x_17); if (x_18 == 0) { @@ -29882,7 +29979,7 @@ lean_inc(x_5); x_21 = l_Lean_Parser_orelseFnCore(x_4, x_8, x_20, x_5, x_6); x_22 = lean_ctor_get(x_21, 4); lean_inc(x_22); -x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_22, x_15); +x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_22, x_15); lean_dec(x_22); if (x_23 == 0) { @@ -29917,6 +30014,24 @@ x_1 = lean_mk_string_from_bytes("field index", 11); return x_1; } } +static lean_object* _init_l_Lean_Parser_fieldIdxFn___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("fieldIdx", 8); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_fieldIdxFn___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_fieldIdxFn___closed__2; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Parser_fieldIdxFn(lean_object* x_1, lean_object* x_2) { _start: { @@ -29951,7 +30066,7 @@ lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_dec(x_4); x_15 = l_Lean_Parser_decimalNumberFn_parseOptDot___closed__1; x_16 = l_Lean_Parser_takeWhileFn(x_15, x_1, x_2); -x_17 = l_Lean_fieldIdxKind; +x_17 = l_Lean_Parser_fieldIdxFn___closed__3; x_18 = l_Lean_Parser_mkNodeToken(x_17, x_5, x_1, x_16); return x_18; } @@ -29979,55 +30094,37 @@ return x_3; static lean_object* _init_l_Lean_Parser_fieldIdx___closed__1() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("fieldIdx", 8); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_fieldIdx___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Parser_fieldIdx___closed__1; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_fieldIdx___closed__3() { -_start: -{ lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; -x_1 = l_Lean_Parser_fieldIdx___closed__1; -x_2 = l_Lean_Parser_fieldIdx___closed__2; +x_1 = l_Lean_Parser_fieldIdxFn___closed__2; +x_2 = l_Lean_Parser_fieldIdxFn___closed__3; x_3 = 1; x_4 = 0; x_5 = l_Lean_Parser_mkAntiquot(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Parser_fieldIdx___closed__4() { +static lean_object* _init_l_Lean_Parser_fieldIdx___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_fieldIdx___closed__1; +x_1 = l_Lean_Parser_fieldIdxFn___closed__2; x_2 = l_Lean_Parser_mkAtomicInfo(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_fieldIdx___closed__5() { +static lean_object* _init_l_Lean_Parser_fieldIdx___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_fieldIdx___closed__3; +x_1 = l_Lean_Parser_fieldIdx___closed__1; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l_Lean_Parser_fieldIdx___closed__4; +x_3 = l_Lean_Parser_fieldIdx___closed__2; x_4 = l_Lean_Parser_orelseInfo(x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_fieldIdx___closed__6() { +static lean_object* _init_l_Lean_Parser_fieldIdx___closed__4() { _start: { lean_object* x_1; @@ -30035,26 +30132,26 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_fieldIdxFn___boxed), 2, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_fieldIdx___closed__7() { +static lean_object* _init_l_Lean_Parser_fieldIdx___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_fieldIdx___closed__3; +x_1 = l_Lean_Parser_fieldIdx___closed__1; x_2 = lean_ctor_get(x_1, 1); lean_inc(x_2); -x_3 = l_Lean_Parser_fieldIdx___closed__6; +x_3 = l_Lean_Parser_fieldIdx___closed__4; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_withAntiquotFn), 4, 2); lean_closure_set(x_4, 0, x_2); lean_closure_set(x_4, 1, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_fieldIdx___closed__8() { +static lean_object* _init_l_Lean_Parser_fieldIdx___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_fieldIdx___closed__5; -x_2 = l_Lean_Parser_fieldIdx___closed__7; +x_1 = l_Lean_Parser_fieldIdx___closed__3; +x_2 = l_Lean_Parser_fieldIdx___closed__5; x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -30065,7 +30162,7 @@ static lean_object* _init_l_Lean_Parser_fieldIdx() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_fieldIdx___closed__8; +x_1 = l_Lean_Parser_fieldIdx___closed__6; return x_1; } } @@ -30764,8 +30861,16 @@ l_Lean_Parser_incQuotDepth___closed__1 = _init_l_Lean_Parser_incQuotDepth___clos lean_mark_persistent(l_Lean_Parser_incQuotDepth___closed__1); l_Lean_Parser_decQuotDepth___closed__1 = _init_l_Lean_Parser_decQuotDepth___closed__1(); lean_mark_persistent(l_Lean_Parser_decQuotDepth___closed__1); +l_Lean_Parser_orelseFnCore___lambda__1___closed__1 = _init_l_Lean_Parser_orelseFnCore___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Parser_orelseFnCore___lambda__1___closed__1); +l_Lean_Parser_orelseFnCore___lambda__1___closed__2 = _init_l_Lean_Parser_orelseFnCore___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Parser_orelseFnCore___lambda__1___closed__2); l_Lean_Parser_orelseFnCore___closed__1 = _init_l_Lean_Parser_orelseFnCore___closed__1(); lean_mark_persistent(l_Lean_Parser_orelseFnCore___closed__1); +l_Lean_Parser_optionalFn___closed__1 = _init_l_Lean_Parser_optionalFn___closed__1(); +lean_mark_persistent(l_Lean_Parser_optionalFn___closed__1); +l_Lean_Parser_optionalFn___closed__2 = _init_l_Lean_Parser_optionalFn___closed__2(); +lean_mark_persistent(l_Lean_Parser_optionalFn___closed__2); l_Lean_Parser_notFollowedByFn___closed__1 = _init_l_Lean_Parser_notFollowedByFn___closed__1(); lean_mark_persistent(l_Lean_Parser_notFollowedByFn___closed__1); l_Lean_Parser_manyAux___lambda__3___closed__1 = _init_l_Lean_Parser_manyAux___lambda__3___closed__1(); @@ -30788,10 +30893,26 @@ l_Lean_Parser_quotedCharFn___closed__1 = _init_l_Lean_Parser_quotedCharFn___clos lean_mark_persistent(l_Lean_Parser_quotedCharFn___closed__1); l_Lean_Parser_charLitFnAux___closed__1 = _init_l_Lean_Parser_charLitFnAux___closed__1(); lean_mark_persistent(l_Lean_Parser_charLitFnAux___closed__1); +l_Lean_Parser_charLitFnAux___closed__2 = _init_l_Lean_Parser_charLitFnAux___closed__2(); +lean_mark_persistent(l_Lean_Parser_charLitFnAux___closed__2); +l_Lean_Parser_charLitFnAux___closed__3 = _init_l_Lean_Parser_charLitFnAux___closed__3(); +lean_mark_persistent(l_Lean_Parser_charLitFnAux___closed__3); l_Lean_Parser_strLitFnAux___closed__1 = _init_l_Lean_Parser_strLitFnAux___closed__1(); lean_mark_persistent(l_Lean_Parser_strLitFnAux___closed__1); +l_Lean_Parser_strLitFnAux___closed__2 = _init_l_Lean_Parser_strLitFnAux___closed__2(); +lean_mark_persistent(l_Lean_Parser_strLitFnAux___closed__2); +l_Lean_Parser_strLitFnAux___closed__3 = _init_l_Lean_Parser_strLitFnAux___closed__3(); +lean_mark_persistent(l_Lean_Parser_strLitFnAux___closed__3); l_Lean_Parser_decimalNumberFn_parseOptDot___closed__1 = _init_l_Lean_Parser_decimalNumberFn_parseOptDot___closed__1(); lean_mark_persistent(l_Lean_Parser_decimalNumberFn_parseOptDot___closed__1); +l_Lean_Parser_decimalNumberFn___closed__1 = _init_l_Lean_Parser_decimalNumberFn___closed__1(); +lean_mark_persistent(l_Lean_Parser_decimalNumberFn___closed__1); +l_Lean_Parser_decimalNumberFn___closed__2 = _init_l_Lean_Parser_decimalNumberFn___closed__2(); +lean_mark_persistent(l_Lean_Parser_decimalNumberFn___closed__2); +l_Lean_Parser_decimalNumberFn___closed__3 = _init_l_Lean_Parser_decimalNumberFn___closed__3(); +lean_mark_persistent(l_Lean_Parser_decimalNumberFn___closed__3); +l_Lean_Parser_decimalNumberFn___closed__4 = _init_l_Lean_Parser_decimalNumberFn___closed__4(); +lean_mark_persistent(l_Lean_Parser_decimalNumberFn___closed__4); l_Lean_Parser_binNumberFn___closed__1 = _init_l_Lean_Parser_binNumberFn___closed__1(); lean_mark_persistent(l_Lean_Parser_binNumberFn___closed__1); l_Lean_Parser_binNumberFn___closed__2 = _init_l_Lean_Parser_binNumberFn___closed__2(); @@ -30844,8 +30965,6 @@ l_Lean_Parser_numLitNoAntiquot___closed__2 = _init_l_Lean_Parser_numLitNoAntiquo lean_mark_persistent(l_Lean_Parser_numLitNoAntiquot___closed__2); l_Lean_Parser_numLitNoAntiquot___closed__3 = _init_l_Lean_Parser_numLitNoAntiquot___closed__3(); lean_mark_persistent(l_Lean_Parser_numLitNoAntiquot___closed__3); -l_Lean_Parser_numLitNoAntiquot___closed__4 = _init_l_Lean_Parser_numLitNoAntiquot___closed__4(); -lean_mark_persistent(l_Lean_Parser_numLitNoAntiquot___closed__4); l_Lean_Parser_numLitNoAntiquot = _init_l_Lean_Parser_numLitNoAntiquot(); lean_mark_persistent(l_Lean_Parser_numLitNoAntiquot); l_Lean_Parser_scientificLitFn___closed__1 = _init_l_Lean_Parser_scientificLitFn___closed__1(); @@ -30858,8 +30977,6 @@ l_Lean_Parser_scientificLitNoAntiquot___closed__2 = _init_l_Lean_Parser_scientif lean_mark_persistent(l_Lean_Parser_scientificLitNoAntiquot___closed__2); l_Lean_Parser_scientificLitNoAntiquot___closed__3 = _init_l_Lean_Parser_scientificLitNoAntiquot___closed__3(); lean_mark_persistent(l_Lean_Parser_scientificLitNoAntiquot___closed__3); -l_Lean_Parser_scientificLitNoAntiquot___closed__4 = _init_l_Lean_Parser_scientificLitNoAntiquot___closed__4(); -lean_mark_persistent(l_Lean_Parser_scientificLitNoAntiquot___closed__4); l_Lean_Parser_scientificLitNoAntiquot = _init_l_Lean_Parser_scientificLitNoAntiquot(); lean_mark_persistent(l_Lean_Parser_scientificLitNoAntiquot); l_Lean_Parser_strLitFn___closed__1 = _init_l_Lean_Parser_strLitFn___closed__1(); @@ -30872,8 +30989,6 @@ l_Lean_Parser_strLitNoAntiquot___closed__2 = _init_l_Lean_Parser_strLitNoAntiquo lean_mark_persistent(l_Lean_Parser_strLitNoAntiquot___closed__2); l_Lean_Parser_strLitNoAntiquot___closed__3 = _init_l_Lean_Parser_strLitNoAntiquot___closed__3(); lean_mark_persistent(l_Lean_Parser_strLitNoAntiquot___closed__3); -l_Lean_Parser_strLitNoAntiquot___closed__4 = _init_l_Lean_Parser_strLitNoAntiquot___closed__4(); -lean_mark_persistent(l_Lean_Parser_strLitNoAntiquot___closed__4); l_Lean_Parser_strLitNoAntiquot = _init_l_Lean_Parser_strLitNoAntiquot(); lean_mark_persistent(l_Lean_Parser_strLitNoAntiquot); l_Lean_Parser_charLitFn___closed__1 = _init_l_Lean_Parser_charLitFn___closed__1(); @@ -30888,22 +31003,22 @@ l_Lean_Parser_charLitNoAntiquot___closed__2 = _init_l_Lean_Parser_charLitNoAntiq lean_mark_persistent(l_Lean_Parser_charLitNoAntiquot___closed__2); l_Lean_Parser_charLitNoAntiquot___closed__3 = _init_l_Lean_Parser_charLitNoAntiquot___closed__3(); lean_mark_persistent(l_Lean_Parser_charLitNoAntiquot___closed__3); -l_Lean_Parser_charLitNoAntiquot___closed__4 = _init_l_Lean_Parser_charLitNoAntiquot___closed__4(); -lean_mark_persistent(l_Lean_Parser_charLitNoAntiquot___closed__4); l_Lean_Parser_charLitNoAntiquot = _init_l_Lean_Parser_charLitNoAntiquot(); lean_mark_persistent(l_Lean_Parser_charLitNoAntiquot); l_Lean_Parser_nameLitFn___closed__1 = _init_l_Lean_Parser_nameLitFn___closed__1(); lean_mark_persistent(l_Lean_Parser_nameLitFn___closed__1); l_Lean_Parser_nameLitFn___closed__2 = _init_l_Lean_Parser_nameLitFn___closed__2(); lean_mark_persistent(l_Lean_Parser_nameLitFn___closed__2); +l_Lean_Parser_nameLitFn___closed__3 = _init_l_Lean_Parser_nameLitFn___closed__3(); +lean_mark_persistent(l_Lean_Parser_nameLitFn___closed__3); +l_Lean_Parser_nameLitFn___closed__4 = _init_l_Lean_Parser_nameLitFn___closed__4(); +lean_mark_persistent(l_Lean_Parser_nameLitFn___closed__4); l_Lean_Parser_nameLitNoAntiquot___closed__1 = _init_l_Lean_Parser_nameLitNoAntiquot___closed__1(); lean_mark_persistent(l_Lean_Parser_nameLitNoAntiquot___closed__1); l_Lean_Parser_nameLitNoAntiquot___closed__2 = _init_l_Lean_Parser_nameLitNoAntiquot___closed__2(); lean_mark_persistent(l_Lean_Parser_nameLitNoAntiquot___closed__2); l_Lean_Parser_nameLitNoAntiquot___closed__3 = _init_l_Lean_Parser_nameLitNoAntiquot___closed__3(); lean_mark_persistent(l_Lean_Parser_nameLitNoAntiquot___closed__3); -l_Lean_Parser_nameLitNoAntiquot___closed__4 = _init_l_Lean_Parser_nameLitNoAntiquot___closed__4(); -lean_mark_persistent(l_Lean_Parser_nameLitNoAntiquot___closed__4); l_Lean_Parser_nameLitNoAntiquot = _init_l_Lean_Parser_nameLitNoAntiquot(); lean_mark_persistent(l_Lean_Parser_nameLitNoAntiquot); l_Lean_Parser_identFn___closed__1 = _init_l_Lean_Parser_identFn___closed__1(); @@ -31005,6 +31120,8 @@ l_Lean_Parser_instInhabitedParserCategory___closed__1 = _init_l_Lean_Parser_inst lean_mark_persistent(l_Lean_Parser_instInhabitedParserCategory___closed__1); l_Lean_Parser_instInhabitedParserCategory = _init_l_Lean_Parser_instInhabitedParserCategory(); 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_9280____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9280____closed__1(); lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9280____closed__1); if (builtin) {res = l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9280_(lean_io_mk_world()); @@ -31211,6 +31328,10 @@ l_Lean_Parser_categoryParserOfStackFn___closed__2 = _init_l_Lean_Parser_category lean_mark_persistent(l_Lean_Parser_categoryParserOfStackFn___closed__2); l_Lean_Parser_fieldIdxFn___closed__1 = _init_l_Lean_Parser_fieldIdxFn___closed__1(); lean_mark_persistent(l_Lean_Parser_fieldIdxFn___closed__1); +l_Lean_Parser_fieldIdxFn___closed__2 = _init_l_Lean_Parser_fieldIdxFn___closed__2(); +lean_mark_persistent(l_Lean_Parser_fieldIdxFn___closed__2); +l_Lean_Parser_fieldIdxFn___closed__3 = _init_l_Lean_Parser_fieldIdxFn___closed__3(); +lean_mark_persistent(l_Lean_Parser_fieldIdxFn___closed__3); l_Lean_Parser_fieldIdx___closed__1 = _init_l_Lean_Parser_fieldIdx___closed__1(); lean_mark_persistent(l_Lean_Parser_fieldIdx___closed__1); l_Lean_Parser_fieldIdx___closed__2 = _init_l_Lean_Parser_fieldIdx___closed__2(); @@ -31223,10 +31344,6 @@ l_Lean_Parser_fieldIdx___closed__5 = _init_l_Lean_Parser_fieldIdx___closed__5(); lean_mark_persistent(l_Lean_Parser_fieldIdx___closed__5); l_Lean_Parser_fieldIdx___closed__6 = _init_l_Lean_Parser_fieldIdx___closed__6(); lean_mark_persistent(l_Lean_Parser_fieldIdx___closed__6); -l_Lean_Parser_fieldIdx___closed__7 = _init_l_Lean_Parser_fieldIdx___closed__7(); -lean_mark_persistent(l_Lean_Parser_fieldIdx___closed__7); -l_Lean_Parser_fieldIdx___closed__8 = _init_l_Lean_Parser_fieldIdx___closed__8(); -lean_mark_persistent(l_Lean_Parser_fieldIdx___closed__8); l_Lean_Parser_fieldIdx = _init_l_Lean_Parser_fieldIdx(); lean_mark_persistent(l_Lean_Parser_fieldIdx); l_Lean_Parser_skip___closed__1 = _init_l_Lean_Parser_skip___closed__1(); diff --git a/stage0/stdlib/Lean/Parser/Command.c b/stage0/stdlib/Lean/Parser/Command.c index 39036a3080f..85ba299cb17 100644 --- a/stage0/stdlib/Lean/Parser/Command.c +++ b/stage0/stdlib/Lean/Parser/Command.c @@ -41,7 +41,6 @@ static lean_object* l_Lean_Parser_Command_theorem___elambda__1___closed__11; static lean_object* l_Lean_Parser_Command_classInductive_formatter___closed__5; static lean_object* l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Term_quot_docString___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__15; static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__20; static lean_object* l_Lean_Parser_Command_classTk___elambda__1___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Command_structure_parenthesizer___closed__2; @@ -122,7 +121,6 @@ static lean_object* l_Lean_Parser_Command_extends___elambda__1___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Command_quot_formatter___closed__1; static lean_object* l_Lean_Parser_Command_optionValue_formatter___closed__5; static lean_object* l_Lean_Parser_Command_openScoped_parenthesizer___closed__3; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__5; static lean_object* l_Lean_Parser_Term_quot___closed__11; LEAN_EXPORT lean_object* l_Lean_Parser_Command_ctor; static lean_object* l_Lean_Parser_Command_structure___closed__3; @@ -145,7 +143,6 @@ static lean_object* l_Lean_Parser_Command_mutual_parenthesizer___closed__12; static lean_object* l___regBuiltin_Lean_Parser_Command_openRenaming_parenthesizer___closed__2; lean_object* l_Lean_Parser_tokenAntiquotFn(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_synth_declRange___closed__3; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__18; static lean_object* l_Lean_Parser_Command_declSig___elambda__1___closed__3; static lean_object* l_Lean_Parser_Command_openRenamingItem___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Command_nonrec_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -157,6 +154,7 @@ static lean_object* l_Lean_Parser_Command_axiom___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Command_terminationByCore_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_noncomputable_formatter___closed__1; static lean_object* l_Lean_Parser_Command_section___closed__3; +static lean_object* l_Lean_Parser_Command_declVal___elambda__1___closed__5; static lean_object* l_Lean_Parser_Command_set__option_formatter___closed__6; static lean_object* l_Lean_Parser_Command_builtin__initialize___elambda__1___closed__3; static lean_object* l_Lean_Parser_Command_export___elambda__1___closed__14; @@ -167,6 +165,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_deriving___elambda__1(lean_object static lean_object* l_Lean_Parser_Command_terminationSuffix___elambda__1___closed__5; static lean_object* l_Lean_Parser_Command_export_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Command_structCtor___elambda__1___closed__11; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__17; static lean_object* l_Lean_Parser_Command_check_formatter___closed__1; static lean_object* l_Lean_Parser_Command_structCtor_formatter___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_structCtor_formatter(lean_object*); @@ -207,6 +206,7 @@ static lean_object* l_Lean_Parser_Command_openScoped_formatter___closed__2; static lean_object* l_Lean_Parser_Term_open___closed__6; static lean_object* l_Lean_Parser_Command_opaque_parenthesizer___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_open_formatter(lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__8; LEAN_EXPORT lean_object* l_Lean_Parser_Command_visibility; static lean_object* l___regBuiltin_Lean_Parser_Command_example_parenthesizer___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_declId_formatter(lean_object*); @@ -245,7 +245,6 @@ static lean_object* l_Lean_Parser_Command_variable___elambda__1___closed__10; static lean_object* l___regBuiltin_Lean_Parser_Command_private_parenthesizer___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_terminationBy; static lean_object* l_Lean_Parser_Command_structImplicitBinder___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__42; static lean_object* l_Lean_Parser_Command_extends___elambda__1___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Command_printAxioms___elambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_end___elambda__1(lean_object*, lean_object*); @@ -267,6 +266,7 @@ static lean_object* l_Lean_Parser_Command_print___elambda__1___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Command_noncomputableSection_declRange___closed__6; static lean_object* l_Lean_Parser_Command_extends_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_genInjectiveTheorems___elambda__1___closed__10; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__44; static lean_object* l_Lean_Parser_Command_example___closed__3; lean_object* l_Lean_Parser_ppGroup_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_Term_attrKind_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -320,6 +320,7 @@ static lean_object* l_Lean_Parser_Command_private___elambda__1___closed__6; lean_object* l_Lean_Parser_notFollowedByFn___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_variable_declRange___closed__4; static lean_object* l_Lean_Parser_Command_declModifiers___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__9; static lean_object* l_Lean_Parser_Command_openRenaming___elambda__1___closed__3; static lean_object* l_Lean_Parser_Command_terminationByCore_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_structCtor___closed__4; @@ -328,15 +329,18 @@ static lean_object* l_Lean_Parser_Command_decreasingBy___elambda__1___closed__7; static lean_object* l_Lean_Parser_Command_noncomputableSection___elambda__1___closed__8; static lean_object* l_Lean_Parser_Command_terminationByCore_formatter___closed__5; lean_object* l_Lean_Parser_many(lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__40; static lean_object* l_Lean_Parser_Command_declModifiers_formatter___closed__3; static lean_object* l_Lean_Parser_Command_extends_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_set__option_declRange___closed__4; static lean_object* l_Lean_Parser_Command_ctor_parenthesizer___closed__3; lean_object* l_Lean_Parser_Term_bracketedBinder(uint8_t); static lean_object* l_Lean_Parser_Command_deriving___elambda__1___closed__12; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__45; lean_object* lean_name_mk_string(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_deriving_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Command_resolve__name___elambda__1___closed__1; +static lean_object* l_Lean_Parser_Command_terminationHintMany___closed__22; static lean_object* l___regBuiltin_Lean_Parser_Command_openRenaming_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_print_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_print_declRange___closed__4; @@ -354,11 +358,11 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_end_parenthesizer(lean_object*, l static lean_object* l_Lean_Parser_Command_optDefDeriving_formatter___closed__3; static lean_object* l_Lean_Parser_Term_set__option___closed__1; static lean_object* l_Lean_Parser_Command_abbrev___closed__10; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__27; static lean_object* l_Lean_Parser_Command_noncomputable___elambda__1___closed__11; static lean_object* l___regBuiltin_Lean_Parser_Command_namespace_declRange___closed__5; static lean_object* l_Lean_Parser_Command_whereStructField___elambda__1___closed__6; static lean_object* l_Lean_Parser_Command_opaque___elambda__1___closed__5; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__3; static lean_object* l_Lean_Parser_Term_quot___elambda__1___closed__22; static lean_object* l_Lean_Parser_Command_declValSimple___elambda__1___closed__13; static lean_object* l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__2; @@ -398,6 +402,7 @@ static lean_object* l_Lean_Parser_Command_ctor___elambda__1___closed__5; lean_object* l_Lean_Parser_sepBy1_parenthesizer(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_axiom_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_structImplicitBinder___elambda__1___closed__6; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__4; static lean_object* l_Lean_Parser_Command_def___elambda__1___closed__8; static lean_object* l_Lean_Parser_Command_nonrec___closed__3; static lean_object* l_Lean_Parser_Command_namedPrio___elambda__1___closed__6; @@ -435,6 +440,7 @@ static lean_object* l_Lean_Parser_Command_extends___elambda__1___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_eval(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Tactic_set__option_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_whereStructInst_parenthesizer___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__2; static lean_object* l_Lean_Parser_Command_opaque___closed__3; static lean_object* l_Lean_Parser_Command_deriving_formatter___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Command_universe_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -448,7 +454,6 @@ static lean_object* l_Lean_Parser_Command_openScoped___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Command_declValSimple_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_open___elambda__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_example___closed__5; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__28; static lean_object* l_Lean_Parser_Command_mutual___elambda__1___closed__19; static lean_object* l_Lean_Parser_Command_openScoped___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_openOnly; @@ -505,7 +510,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_printAxioms_declRange___c static lean_object* l_Lean_Parser_Command_inductive___elambda__1___closed__18; static lean_object* l_Lean_Parser_Command_ctor_parenthesizer___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Command_structFields_formatter___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__30; static lean_object* l_Lean_Parser_Command_declSig___elambda__1___closed__15; static lean_object* l___regBuiltin_Lean_Parser_Command_in_declRange___closed__7; lean_object* l_Lean_Parser_group_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -666,6 +670,7 @@ static lean_object* l_Lean_Parser_Command_declValEqns___elambda__1___closed__4; static lean_object* l_Lean_Parser_Command_def___elambda__1___closed__14; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_quot_declRange(lean_object*); static lean_object* l_Lean_Parser_Command_theorem___closed__5; +static lean_object* l_Lean_Parser_Command_declVal___elambda__1___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_in_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_variable_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_structImplicitBinder_formatter___closed__7; @@ -673,6 +678,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_set__option_declRange___clos lean_object* l_id___rarg___boxed(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_attribute_formatter___closed__1; lean_object* l_Lean_Parser_Term_bracketedBinder_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__33; static lean_object* l___regBuiltin_Lean_Parser_Command_end_formatter___closed__1; static lean_object* l_Lean_Parser_Command_abbrev___closed__5; static lean_object* l_Lean_Parser_Command_section___closed__1; @@ -696,7 +702,6 @@ static lean_object* l_Lean_Parser_Command_namespace___closed__8; static lean_object* l_Lean_Parser_Command_declModifiers_formatter___closed__25; static lean_object* l_Lean_Parser_Command_structCtor_formatter___closed__3; static lean_object* l_Lean_Parser_Command_deriving_formatter___closed__8; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__46; lean_object* l_Lean_Parser_ParserState_mkNode(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_terminationByCore___elambda__1___closed__10; static lean_object* l_Lean_Parser_Tactic_open___closed__3; @@ -725,6 +730,7 @@ static lean_object* l_Lean_Parser_Command_instance___elambda__1___closed__15; static lean_object* l_Lean_Parser_Command_terminationHintMany___closed__16; static lean_object* l_Lean_Parser_Command_theorem_formatter___closed__5; static lean_object* l_Lean_Parser_Command_inductive___closed__5; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__23; static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__14; static lean_object* l_Lean_Parser_Command_quot_formatter___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_instance_parenthesizer(lean_object*); @@ -756,6 +762,7 @@ static lean_object* l_Lean_Parser_Term_set__option___closed__6; static lean_object* l_Lean_Parser_Command_initialize___elambda__1___closed__17; static lean_object* l_Lean_Parser_Command_structure_formatter___closed__4; static lean_object* l_Lean_Parser_Command_structure___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__41; static lean_object* l_Lean_Parser_Command_structureTk___elambda__1___closed__1; static lean_object* l_Lean_Parser_Command_eraseAttr_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_optDeriving_parenthesizer___closed__5; @@ -766,6 +773,7 @@ static lean_object* l_Lean_Parser_Command_namedPrio___closed__14; static lean_object* l_Lean_Parser_Command_eraseAttr___elambda__1___closed__4; static lean_object* l_Lean_Parser_Command_ctor___closed__3; static lean_object* l_Lean_Parser_Command_namedPrio___elambda__1___closed__17; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__13; static lean_object* l_Lean_Parser_Command_mutual___elambda__1___closed__21; lean_object* l_Lean_Parser_Term_attributes_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_resolve__name_parenthesizer___closed__1; @@ -832,11 +840,11 @@ static lean_object* l_Lean_Parser_Command_openScoped_formatter___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_set__option_declRange___closed__5; static lean_object* l_Lean_Parser_Command_namedPrio___elambda__1___closed__15; static lean_object* l_Lean_Parser_Command_structExplicitBinder___elambda__1___closed__18; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__34; static lean_object* l_Lean_Parser_Term_precheckedQuot_formatter___closed__2; static lean_object* l_Lean_Parser_Command_init__quot_formatter___closed__1; static lean_object* l_Lean_Parser_Command_openRenaming_formatter___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_classInductive_parenthesizer(lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__2; static lean_object* l_Lean_Parser_Command_structCtor___elambda__1___closed__1; static lean_object* l_Lean_Parser_Command_eraseAttr___elambda__1___closed__2; static lean_object* l_Lean_Parser_Command_classInductive___elambda__1___closed__7; @@ -864,6 +872,7 @@ static lean_object* l_Lean_Parser_Command_openHiding___elambda__1___closed__11; lean_object* l_Lean_Parser_addBuiltinParser(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_init__quot_declRange___closed__6; static lean_object* l_Lean_Parser_Command_classInductive_parenthesizer___closed__6; +static lean_object* l_Lean_Parser_Command_declVal___elambda__1___closed__3; static lean_object* l_Lean_Parser_Command_theorem___elambda__1___closed__13; static lean_object* l_Lean_Parser_Command_openSimple___elambda__1___closed__5; static lean_object* l_Lean_Parser_Command_export___closed__12; @@ -875,7 +884,6 @@ lean_object* lean_array_get_size(lean_object*); static lean_object* l_Lean_Parser_Command_declId_formatter___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_structCtor_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Command_quot_parenthesizer___closed__7; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__47; LEAN_EXPORT lean_object* l_Lean_Parser_Command_structFields; static lean_object* l_Lean_Parser_Command_abbrev___closed__3; static lean_object* l_Lean_Parser_Command_abbrev_formatter___closed__6; @@ -909,7 +917,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_abbrev_formatter(lean_object*, le static lean_object* l_Lean_Parser_Command_initialize_formatter___closed__4; static lean_object* l_Lean_Parser_Command_genInjectiveTheorems___elambda__1___closed__11; static lean_object* l_Lean_Parser_Command_inductive___elambda__1___closed__10; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__25; static lean_object* l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__9; static lean_object* l_Lean_Parser_Command_printAxioms___elambda__1___closed__3; static lean_object* l_Lean_Parser_Term_set__option___elambda__1___closed__1; @@ -946,6 +953,8 @@ lean_object* l_Lean_Parser_unicodeSymbol_formatter(lean_object*, lean_object*, l static lean_object* l_Lean_Parser_Command_terminationByElement_formatter___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Command_structSimpleBinder_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_quot_formatter___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__38; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__21; static lean_object* l_Lean_Parser_Command_declSig___elambda__1___closed__10; static lean_object* l_Lean_Parser_Command_set__option_formatter___closed__1; static lean_object* l_Lean_Parser_Command_structFields___elambda__1___closed__8; @@ -1003,9 +1012,9 @@ static lean_object* l_Lean_Parser_Command_terminationBy_formatter___closed__5; static lean_object* l_Lean_Parser_Command_mutual_formatter___closed__4; static lean_object* l_Lean_Parser_Command_end_formatter___closed__4; static lean_object* l_Lean_Parser_Command_optDeriving_formatter___closed__5; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__37; static lean_object* l___regBuiltin_Lean_Parser_Command_noncomputableSection_declRange___closed__4; static lean_object* l_Lean_Parser_Command_print___elambda__1___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__50; static lean_object* l___regBuiltin_Lean_Parser_Command_attribute_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_universe_declRange___closed__6; lean_object* l_Lean_Parser_checkPrecFn___boxed(lean_object*, lean_object*, lean_object*); @@ -1061,7 +1070,6 @@ static lean_object* l_Lean_Parser_Command_unsafe_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_set__option_declRange___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Command_instance___elambda__1(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_check__failure_parenthesizer___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__6; static lean_object* l_Lean_Parser_Command_structExplicitBinder_formatter___closed__2; static lean_object* l_Lean_Parser_Term_precheckedQuot_formatter___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_classInductive_formatter(lean_object*); @@ -1187,7 +1195,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_openScoped_formatter___cl static lean_object* l___regBuiltin_Lean_Parser_Command_end_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_declValSimple___elambda__1___closed__14; static lean_object* l_Lean_Parser_Command_eraseAttr___elambda__1___closed__10; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__29; static lean_object* l_Lean_Parser_Command_ctor_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_openSimple___elambda__1___closed__1; static lean_object* l_Lean_Parser_Command_inductive___elambda__1___closed__20; @@ -1200,6 +1207,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_axiom_formatter___closed_ static lean_object* l_Lean_Parser_Command_openScoped___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Command_initialize___elambda__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_example___elambda__1___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__48; static lean_object* l___regBuiltin_Lean_Parser_Command_in_declRange___closed__5; static lean_object* l_Lean_Parser_Tactic_open___elambda__1___closed__2; static lean_object* l_Lean_Parser_Command_deriving___elambda__1___closed__10; @@ -1221,7 +1229,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_def_formatter___closed__2 static lean_object* l_Lean_Parser_Command_deriving___closed__6; static lean_object* l_Lean_Parser_Command_mutual___closed__6; static lean_object* l_Lean_Parser_Command_private___closed__5; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__50; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__37; static lean_object* l_Lean_Parser_Command_export_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_initialize___elambda__1___closed__5; static lean_object* l_Lean_Parser_Command_namedPrio_formatter___closed__2; @@ -1325,6 +1333,7 @@ static lean_object* l_Lean_Parser_Command_structure___elambda__1___closed__20; static lean_object* l_Lean_Parser_Command_mutual_parenthesizer___closed__9; LEAN_EXPORT lean_object* l_Lean_Parser_Command_eraseAttr___elambda__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_reduce___closed__5; +static lean_object* l_Lean_Parser_Command_declVal_formatter___closed__3; static lean_object* l_Lean_Parser_Command_declValSimple_formatter___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Command_private_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_classInductive_formatter___closed__6; @@ -1333,16 +1342,15 @@ static lean_object* l_Lean_Parser_Command_attribute_formatter___closed__2; static lean_object* l_Lean_Parser_Command_inductive___elambda__1___closed__13; static lean_object* l_Lean_Parser_Command_classInductive_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_structSimpleBinder___closed__3; +uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_end_parenthesizer___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_axiom_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_optionValue___elambda__1___closed__8; static lean_object* l_Lean_Parser_Command_optDeclSig_parenthesizer___closed__10; static lean_object* l_Lean_Parser_Command_openScoped___closed__7; static lean_object* l_Lean_Parser_Command_whereStructField___elambda__1___closed__5; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__21; static lean_object* l_Lean_Parser_Command_optDefDeriving_formatter___closed__1; static lean_object* l_Lean_Parser_Command_print_parenthesizer___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__38; static lean_object* l_Lean_Parser_Command_axiom___elambda__1___closed__11; static lean_object* l_Lean_Parser_Command_structSimpleBinder_formatter___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Command_section_declRange___closed__2; @@ -1383,7 +1391,6 @@ static lean_object* l_Lean_Parser_Command_structSimpleBinder___closed__6; static lean_object* l_Lean_Parser_Command_structFields_formatter___closed__10; static lean_object* l_Lean_Parser_Command_structCtor_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Term_quot___elambda__1___closed__6; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__34; lean_object* l_Lean_Parser_Term_typeSpec_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_set__option_declRange___closed__6; static lean_object* l_Lean_Parser_Command_initialize_parenthesizer___closed__10; @@ -1406,7 +1413,6 @@ lean_object* l_Lean_PrettyPrinter_Formatter_trailingNode_formatter(lean_object*, static lean_object* l_Lean_Parser_Command_openRenaming_formatter___closed__3; static lean_object* l_Lean_Parser_Command_check___elambda__1___closed__1; static lean_object* l_Lean_Parser_Command_set__option___elambda__1___closed__4; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__48; static lean_object* l_Lean_Parser_Command_section_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_set__option___elambda__1___closed__8; static lean_object* l_Lean_Parser_Command_declValSimple___elambda__1___closed__1; @@ -1415,6 +1421,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_structCtor_parenthesizer_ static lean_object* l_Lean_Parser_Command_axiom_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_mutual___elambda__1___closed__12; static lean_object* l___regBuiltin_Lean_Parser_Command_eval_declRange___closed__6; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__29; static lean_object* l_Lean_Parser_Command_decreasingBy___closed__5; static lean_object* l_Lean_Parser_Command_structSimpleBinder_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Tactic_set__option___elambda__1___closed__1; @@ -1485,7 +1492,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_terminationByCore_parenthesizer(l static lean_object* l_Lean_Parser_Command_structFields_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Command_openSimple_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_decreasingBy_formatter___closed__2; -LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424_(lean_object*); static lean_object* l_Lean_Parser_Command_reduce___elambda__1___closed__8; static lean_object* l_Lean_Parser_Command_initialize___closed__4; static lean_object* l_Lean_Parser_Command_def___elambda__1___closed__4; @@ -1498,7 +1505,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_in_parenthesizer___closed LEAN_EXPORT lean_object* l_Lean_Parser_Command_resolve__name; static lean_object* l_Lean_Parser_Command_declValSimple___closed__10; LEAN_EXPORT lean_object* l_Lean_Parser_Command_declValEqns___elambda__1(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__45; static lean_object* l_Lean_Parser_Command_structExplicitBinder___closed__9; static lean_object* l_Lean_Parser_Command_example_parenthesizer___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Command_check__failure_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1506,6 +1512,7 @@ static lean_object* l_Lean_Parser_Command_attribute___closed__6; static lean_object* l_Lean_Parser_Command_optionValue_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_optDeclSig_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_precheckedQuot_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__46; LEAN_EXPORT lean_object* l_Lean_Parser_Command_in; static lean_object* l_Lean_Parser_Command_opaque___elambda__1___closed__10; static lean_object* l_Lean_Parser_Command_openHiding___closed__8; @@ -1529,7 +1536,6 @@ static lean_object* l_Lean_Parser_Command_whereStructInst___closed__9; static lean_object* l_Lean_Parser_Command_genInjectiveTheorems___elambda__1___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Command_optNamedPrio_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_printAxioms_declRange(lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__33; static lean_object* l_Lean_Parser_Command_check__failure_formatter___closed__4; static lean_object* l_Lean_Parser_Command_exit___elambda__1___closed__2; static lean_object* l_Lean_Parser_Term_open___elambda__1___closed__12; @@ -1551,7 +1557,6 @@ static lean_object* l_Lean_Parser_Command_classInductive___closed__7; static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__19; static lean_object* l_Lean_Parser_Command_optionValue___elambda__1___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_terminationBy_formatter___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__23; static lean_object* l_Lean_Parser_Command_structInstBinder___elambda__1___closed__4; static lean_object* l_Lean_Parser_Command_terminationByElement___elambda__1___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Command_ctor_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1588,6 +1593,7 @@ static lean_object* l_Lean_Parser_Command_opaque_formatter___closed__2; static lean_object* l_Lean_Parser_Command_whereStructField___elambda__1___closed__7; static lean_object* l_Lean_Parser_Command_in_parenthesizer___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_attribute; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__25; static lean_object* l___regBuiltin_Lean_Parser_Command_open_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_terminationHintMany___closed__20; static lean_object* l_Lean_Parser_Command_structImplicitBinder___elambda__1___closed__7; @@ -1596,6 +1602,7 @@ static lean_object* l_Lean_Parser_Command_mutual_formatter___closed__3; static lean_object* l_Lean_Parser_Command_nonrec_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_structure___closed__10; static lean_object* l_Lean_Parser_Command_decreasingBy___closed__7; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__22; static lean_object* l_Lean_Parser_Command_attribute_formatter___closed__8; lean_object* l_Lean_Parser_nonReservedSymbolFnAux(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_private___elambda__1___closed__10; @@ -1689,6 +1696,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_quot_declRange___closed__ static lean_object* l_Lean_Parser_Command_declModifiers___closed__5; static lean_object* l_Lean_Parser_Command_structureTk___closed__1; static lean_object* l_Lean_Parser_Command_attribute_formatter___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__47; static lean_object* l_Lean_Parser_Command_moduleDoc___elambda__1___closed__14; static lean_object* l_Lean_Parser_Command_optDeriving_formatter___closed__4; static lean_object* l_Lean_Parser_Command_init__quot___closed__1; @@ -1714,13 +1722,11 @@ static lean_object* l_Lean_Parser_Command_terminationByCore_formatter___closed__ static lean_object* l_Lean_Parser_Command_export_parenthesizer___closed__8; static lean_object* l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Command_opaque_formatter___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__13; static lean_object* l_Lean_Parser_Command_openDecl___closed__6; static lean_object* l_Lean_Parser_Command_protected_formatter___closed__1; static lean_object* l_Lean_Parser_Term_quot___elambda__1___closed__9; static lean_object* l_Lean_Parser_Command_inductive_formatter___closed__11; static lean_object* l_Lean_Parser_Command_def_parenthesizer___closed__7; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__41; static lean_object* l_Lean_Parser_Command_export___closed__9; static lean_object* l_Lean_Parser_Command_reduce___elambda__1___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Command_check__failure___elambda__1(lean_object*, lean_object*); @@ -1813,7 +1819,6 @@ lean_object* l_Lean_Parser_Term_whereDecls_formatter(lean_object*, lean_object*, static lean_object* l_Lean_Parser_Command_mutual_formatter___closed__5; static lean_object* l_Lean_Parser_Command_export___elambda__1___closed__12; static lean_object* l_Lean_Parser_Command_openRenamingItem___elambda__1___closed__3; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__31; static lean_object* l_Lean_Parser_Command_def___elambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Term_open___elambda__1(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_quot_declRange___closed__5; @@ -1846,7 +1851,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_open_formatter(lean_obj static lean_object* l_Lean_Parser_Command_protected___elambda__1___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_quot_declRange(lean_object*); static lean_object* l_Lean_Parser_Command_genInjectiveTheorems___closed__7; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__9; static lean_object* l_Lean_Parser_Command_structFields___elambda__1___closed__2; static lean_object* l_Lean_Parser_Command_print___closed__3; static lean_object* l_Lean_Parser_Command_structInstBinder___elambda__1___closed__18; @@ -1870,7 +1874,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_precheckedQuot_parenthesizer static lean_object* l_Lean_Parser_Term_quot___elambda__1___closed__17; LEAN_EXPORT lean_object* l_Lean_Parser_Command_axiom_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_openDecl_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__40; static lean_object* l_Lean_Parser_Command_axiom___closed__7; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_eraseAttr_formatter(lean_object*); static lean_object* l_Lean_Parser_Command_eval___elambda__1___closed__2; @@ -1889,6 +1892,7 @@ static lean_object* l_Lean_Parser_Command_structureTk___elambda__1___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_check__failure(lean_object*); static lean_object* l_Lean_Parser_Command_terminationBy_formatter___closed__3; static lean_object* l_Lean_Parser_Command_structFields_parenthesizer___closed__4; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_terminationSuffix___elambda__1(lean_object*, lean_object*); lean_object* l_Lean_Parser_many1_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_Term_attrInstance___elambda__1(lean_object*, lean_object*); @@ -1911,7 +1915,6 @@ static lean_object* l_Lean_Parser_Term_precheckedQuot___closed__4; static lean_object* l_Lean_Parser_Command_structSimpleBinder___elambda__1___closed__4; static lean_object* l_Lean_Parser_Command_def___elambda__1___closed__1; static lean_object* l_Lean_Parser_Command_universe_parenthesizer___closed__3; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__27; LEAN_EXPORT lean_object* l_Lean_Parser_Command_example_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_classTk___elambda__1___closed__1; static lean_object* l_Lean_Parser_Command_variable___closed__1; @@ -1925,7 +1928,6 @@ static lean_object* l_Lean_Parser_Command_openDecl___closed__5; static lean_object* l_Lean_Parser_Command_classInductive_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_namedPrio_formatter___closed__9; static lean_object* l_Lean_Parser_Command_reduce___elambda__1___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__44; static lean_object* l___regBuiltin_Lean_Parser_Command_genInjectiveTheorems_declRange___closed__7; static lean_object* l_Lean_Parser_Command_section___elambda__1___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Command_universe_declRange___closed__5; @@ -1951,7 +1953,6 @@ static lean_object* l_Lean_Parser_Command_synth_parenthesizer___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_quot_declRange___closed__3; static lean_object* l_Lean_Parser_Command_openRenamingItem___closed__2; static lean_object* l_Lean_Parser_Command_structExplicitBinder___elambda__1___closed__13; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__1; lean_object* l_Lean_Parser_Term_structInst___elambda__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_structure_parenthesizer___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_open_parenthesizer(lean_object*); @@ -1977,6 +1978,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_terminationBy_formatter__ static lean_object* l_Lean_Parser_Command_mutual_parenthesizer___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_extends___elambda__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_noncomputableSection___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__28; static lean_object* l_Lean_Parser_Command_instance___elambda__1___closed__13; static lean_object* l___regBuiltin_Lean_Parser_Term_open_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_namespace___elambda__1___closed__2; @@ -2006,6 +2008,7 @@ static lean_object* l_Lean_Parser_Command_quot_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Command_deriving___elambda__1___closed__19; static lean_object* l_Lean_Parser_Command_resolve__name_formatter___closed__3; static lean_object* l_Lean_Parser_Command_decreasingBy___closed__3; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__5; static lean_object* l_Lean_Parser_Command_whereStructInst___closed__6; static lean_object* l_Lean_Parser_Tactic_open___closed__9; static lean_object* l_Lean_Parser_Command_namedPrio___closed__8; @@ -2028,7 +2031,6 @@ static lean_object* l_Lean_Parser_Command_structFields_parenthesizer___closed__1 LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_in_declRange(lean_object*); static lean_object* l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_openScoped_parenthesizer___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Command_decreasingBy; static lean_object* l_Lean_Parser_Command_optDeclSig___elambda__1___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Command_declValEqns_parenthesizer___closed__2; @@ -2044,6 +2046,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_end_declRange___closed__3 static lean_object* l_Lean_Parser_Command_instance___elambda__1___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Command_private_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_declModifiers___closed__3; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__14; static lean_object* l_Lean_Parser_Command_optDefDeriving___closed__7; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_printAxioms_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Command_end___closed__7; @@ -2091,6 +2094,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_set__option; static lean_object* l_Lean_Parser_Command_deriving_formatter___closed__6; static lean_object* l_Lean_Parser_Command_declaration_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Command_openScoped___elambda__1___closed__10; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__18; static lean_object* l_Lean_Parser_Command_set__option___elambda__1___closed__14; static lean_object* l_Lean_Parser_Command_end_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_optDeclSig_formatter___closed__1; @@ -2107,11 +2111,13 @@ static lean_object* l_Lean_Parser_Tactic_set__option_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Command_attribute_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_synth___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Command_printAxioms_declRange___closed__6; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__15; lean_object* l_Lean_Parser_orelseInfo(lean_object*, lean_object*); lean_object* l_Lean_Parser_termParser_formatter___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_structCtor___closed__2; static lean_object* l_Lean_Parser_Command_example_formatter___closed__4; static lean_object* l_Lean_Parser_Command_init__quot___closed__7; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__30; static lean_object* l_Lean_Parser_Command_set__option___elambda__1___closed__5; static lean_object* l_Lean_Parser_Command_openOnly_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_terminationSuffix_formatter___closed__1; @@ -2235,6 +2241,7 @@ static lean_object* l_Lean_Parser_Command_declVal___closed__2; static lean_object* l_Lean_Parser_Command_declValSimple_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_theorem___elambda__1___closed__10; static lean_object* l___regBuiltin_Lean_Parser_Command_namedPrio_formatter___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__6; static lean_object* l_Lean_Parser_Command_terminationBy___elambda__1___closed__9; static lean_object* l___regBuiltin_Lean_Parser_Command_exit_declRange___closed__3; static lean_object* l_Lean_Parser_Command_print___elambda__1___closed__12; @@ -2265,13 +2272,13 @@ static lean_object* l_Lean_Parser_Command_theorem_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_structExplicitBinder_formatter___closed__6; static lean_object* l_Lean_Parser_Command_structInstBinder___closed__4; static lean_object* l_Lean_Parser_Term_precheckedQuot___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__17; LEAN_EXPORT lean_object* l_Lean_Parser_Command_namedPrio; static lean_object* l_Lean_Parser_Command_inductive___closed__3; static lean_object* l_Lean_Parser_Command_declId___elambda__1___closed__10; static lean_object* l_Lean_Parser_Command_noncomputableSection_formatter___closed__1; static lean_object* l_Lean_Parser_Command_quot_formatter___closed__7; static lean_object* l_Lean_Parser_Command_mutual_formatter___closed__7; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__42; static lean_object* l_Lean_Parser_Command_attribute_parenthesizer___closed__8; static lean_object* l_Lean_Parser_Command_namedPrio___closed__6; static lean_object* l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__9; @@ -2316,7 +2323,6 @@ static lean_object* l_Lean_Parser_Command_structCtor___elambda__1___closed__12; static lean_object* l_Lean_Parser_Command_structFields___closed__6; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_open_declRange(lean_object*); static lean_object* l_Lean_Parser_Command_structCtor_parenthesizer___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__8; LEAN_EXPORT lean_object* l_Lean_Parser_Command_inductive; static lean_object* l___regBuiltin_Lean_Parser_Command_optDeriving_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_optNamedPrio; @@ -2325,6 +2331,7 @@ static lean_object* l_Lean_Parser_Command_declId___elambda__1___closed__20; static lean_object* l___regBuiltin_Lean_Parser_Command_mutual_declRange___closed__6; extern lean_object* l_Lean_PrettyPrinter_formatterAttribute; static lean_object* l___regBuiltin_Lean_Parser_Command_resolve__name_parenthesizer___closed__1; +static lean_object* l_Lean_Parser_Command_declVal___elambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_def_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_quot_declRange___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_mutual_declRange___closed__7; @@ -2431,6 +2438,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_eval_declRange___closed__ static lean_object* l___regBuiltin_Lean_Parser_Command_init__quot_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_namespace_formatter___closed__2; lean_object* l_Lean_PrettyPrinter_Parenthesizer_many1Unbox_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__16; static lean_object* l_Lean_Parser_Command_openRenamingItem_formatter___closed__3; static lean_object* l_Lean_Parser_Command_structInstBinder_parenthesizer___closed__9; lean_object* l_Lean_FileMap_toPosition(lean_object*, lean_object*); @@ -2473,7 +2481,6 @@ static lean_object* l_Lean_Parser_Command_namedPrio_formatter___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Command_structSimpleBinder_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_namespace_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_set__option___closed__10; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__39; static lean_object* l_Lean_Parser_Command_moduleDoc___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Command_terminationByElement_formatter___closed__1; static lean_object* l_Lean_Parser_Command_structure_parenthesizer___closed__12; @@ -2574,6 +2581,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_exit_parenthesizer(l static lean_object* l_Lean_Parser_Command_openScoped_formatter___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_classTk___elambda__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_classInductive___closed__9; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__24; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_variable_formatter(lean_object*); static lean_object* l_Lean_Parser_Command_eval___elambda__1___closed__9; static lean_object* l___regBuiltin_Lean_Parser_Command_openSimple_formatter___closed__2; @@ -2618,7 +2626,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_synth_declRange___closed_ static lean_object* l_Lean_Parser_Command_structure_formatter___closed__7; static lean_object* l_Lean_Parser_Command_structFields___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_universe_parenthesizer___closed__2; -lean_object* l_Lean_Parser_registerAlias(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Parser_registerAlias(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_terminationByElement___closed__2; static lean_object* l_Lean_Parser_Command_declId___elambda__1___closed__26; static lean_object* l___regBuiltin_Lean_Parser_Command_resolve__name_declRange___closed__5; @@ -2650,6 +2658,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_eval_formatter___closed__ static lean_object* l_Lean_Parser_Command_set__option___elambda__1___closed__12; static lean_object* l_Lean_Parser_Command_namedPrio___elambda__1___closed__12; static lean_object* l_Lean_Parser_Command_openDecl___closed__3; +static lean_object* l_Lean_Parser_Command_declVal___closed__5; static lean_object* l_Lean_Parser_Command_declValSimple___closed__2; static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__13; static lean_object* l_Lean_Parser_Command_printAxioms___closed__8; @@ -2659,6 +2668,7 @@ static lean_object* l_Lean_Parser_Command_initialize___elambda__1___closed__22; static lean_object* l___regBuiltin_Lean_Parser_Command_check_parenthesizer___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_export; LEAN_EXPORT lean_object* l_Lean_Parser_Command_instance_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__49; static lean_object* l_Lean_Parser_Command_openRenaming___closed__3; static lean_object* l_Lean_Parser_Command_moduleDoc___closed__3; static lean_object* l_Lean_Parser_Command_structInstBinder_formatter___closed__5; @@ -2704,7 +2714,9 @@ static lean_object* l_Lean_Parser_Command_optionValue___closed__1; static lean_object* l_Lean_Parser_Command_init__quot___elambda__1___closed__7; static lean_object* l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__10; lean_object* l_Lean_PrettyPrinter_Formatter_withPosition_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__10; static lean_object* l___regBuiltin_Lean_Parser_Tactic_open_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__32; LEAN_EXPORT lean_object* l_Lean_Parser_Command_terminationHint(lean_object*); lean_object* l_Lean_Parser_withResultOfFn(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_genInjectiveTheorems___elambda__1___closed__4; @@ -2734,6 +2746,7 @@ static lean_object* l_Lean_Parser_Tactic_open___closed__2; static lean_object* l_Lean_Parser_Command_declaration_parenthesizer___closed__12; static lean_object* l_Lean_Parser_Command_private___elambda__1___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_variable_declRange___closed__7; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__20; static lean_object* l_Lean_Parser_Command_resolve__name___closed__3; static lean_object* l_Lean_Parser_Command_structImplicitBinder___closed__6; static lean_object* l_Lean_Parser_Command_instance_formatter___closed__6; @@ -2816,19 +2829,18 @@ static lean_object* l_Lean_Parser_Command_terminationByElement_parenthesizer___c static lean_object* l_Lean_Parser_Command_terminationByElement___closed__1; static lean_object* l_Lean_Parser_Command_builtin__initialize___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__11; static lean_object* l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Command_private; static lean_object* l_Lean_Parser_Command_check___closed__6; static lean_object* l_Lean_Parser_Command_openRenamingItem___elambda__1___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Command_openDecl_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_decreasingBy___elambda__1___closed__9; +static lean_object* l_Lean_Parser_Command_declVal_parenthesizer___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_structSimpleBinder; lean_object* l_Lean_Parser_mkAntiquot_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_extends___closed__4; static lean_object* l_Lean_Parser_Term_quot_formatter___closed__3; lean_object* l_Lean_Parser_sepBy1(lean_object*, lean_object*, lean_object*, uint8_t); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__19; static lean_object* l_Lean_Parser_Command_noncomputable___elambda__1___closed__5; static lean_object* l_Lean_Parser_Command_check___elambda__1___closed__9; static lean_object* l_Lean_Parser_Command_def___closed__11; @@ -2854,7 +2866,6 @@ static lean_object* l_Lean_Parser_Command_declValSimple___elambda__1___closed__1 static lean_object* l_Lean_Parser_Term_precheckedQuot___elambda__1___closed__2; static lean_object* l_Lean_Parser_Command_terminationBy___elambda__1___closed__1; static lean_object* l_Lean_Parser_Command_ctor___elambda__1___closed__10; -extern lean_object* l_Lean_groupKind; static lean_object* l_Lean_Parser_Command_terminationByElement_formatter___closed__2; static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_noncomputableSection___elambda__1___closed__7; @@ -2964,7 +2975,6 @@ static lean_object* l_Lean_Parser_Command_declaration_formatter___closed__8; static lean_object* l_Lean_Parser_Command_inductive___closed__11; LEAN_EXPORT lean_object* l_Lean_Parser_Command_whereStructField___elambda__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_set__option_formatter___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__20; extern lean_object* l_Lean_Parser_numLit; static lean_object* l_Lean_Parser_Command_terminationHintMany___elambda__1___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_end_formatter(lean_object*); @@ -2995,7 +3005,6 @@ static lean_object* l_Lean_Parser_Command_openSimple___closed__4; static lean_object* l_Lean_Parser_Command_instance_parenthesizer___closed__7; lean_object* l_Lean_Parser_Term_ident___elambda__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_terminationBy___elambda__1___closed__17; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__12; static lean_object* l___regBuiltin_Lean_Parser_Command_synth_formatter___closed__1; static lean_object* l_Lean_Parser_Command_inductive___elambda__1___closed__8; static lean_object* l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__3; @@ -3059,6 +3068,7 @@ static lean_object* l_Lean_Parser_Command_genInjectiveTheorems_formatter___close static lean_object* l___regBuiltin_Lean_Parser_Command_noncomputable_parenthesizer___closed__2; lean_object* l_Lean_Syntax_getNumArgs(lean_object*); static lean_object* l_Lean_Parser_Command_quot___elambda__1___closed__5; +static lean_object* l_Lean_Parser_Command_terminationHintMany___closed__21; LEAN_EXPORT lean_object* l_Lean_Parser_Command_structImplicitBinder; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_reduce(lean_object*); static lean_object* l_Lean_Parser_Command_openRenaming_formatter___closed__5; @@ -3170,7 +3180,6 @@ static lean_object* l_Lean_Parser_Command_opaque_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_whereStructField___elambda__1___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_deriving_declRange___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Command_nonrec_parenthesizer___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__43; static lean_object* l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_openDecl___elambda__1___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_resolve__name_declRange___closed__4; @@ -3250,7 +3259,6 @@ static lean_object* l_Lean_Parser_Command_noncomputable_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_section_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_classTk_formatter___closed__2; static lean_object* l_Lean_Parser_Command_decreasingBy___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__26; static lean_object* l_Lean_Parser_Command_instance___closed__6; static lean_object* l_Lean_Parser_Command_namespace___elambda__1___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Command_structImplicitBinder_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -3319,7 +3327,6 @@ static lean_object* l_Lean_Parser_Command_openOnly___elambda__1___closed__3; static lean_object* l_Lean_Parser_Command_terminationByElement___elambda__1___closed__8; static lean_object* l_Lean_Parser_Command_check__failure___elambda__1___closed__8; LEAN_EXPORT lean_object* l_Lean_Parser_Command_declModifiersF_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__35; static lean_object* l_Lean_Parser_Command_structInstBinder___elambda__1___closed__1; static lean_object* l_Lean_Parser_Command_protected___elambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Command_declId_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -3331,7 +3338,6 @@ static lean_object* l_Lean_Parser_Command_structFields_formatter___closed__6; static lean_object* l_Lean_Parser_Command_declSig___elambda__1___closed__1; static lean_object* l_Lean_Parser_Command_inductive___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_mutual_parenthesizer___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__36; static lean_object* l___regBuiltin_Lean_Parser_Command_attribute_declRange___closed__5; static lean_object* l_Lean_Parser_Command_terminationByElement___elambda__1___closed__17; static lean_object* l_Lean_Parser_Command_declValSimple___closed__4; @@ -3421,6 +3427,7 @@ static lean_object* l_Lean_Parser_Command_printAxioms___closed__2; static lean_object* l_Lean_Parser_Command_in___closed__8; static lean_object* l_Lean_Parser_Tactic_open___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_openScoped_formatter(lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__7; static lean_object* l_Lean_Parser_Command_namespace_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_deriving___elambda__1___closed__21; static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__21; @@ -3501,6 +3508,7 @@ static lean_object* l_Lean_Parser_Command_example_formatter___closed__2; static lean_object* l_Lean_Parser_Command_def_parenthesizer___closed__9; static lean_object* l_Lean_Parser_Command_terminationBy___closed__5; static lean_object* l_Lean_Parser_Command_initialize_formatter___closed__10; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__35; static lean_object* l_Lean_Parser_Command_structFields___elambda__1___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Command_attribute___elambda__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_optDeclSig_parenthesizer___closed__1; @@ -3553,6 +3561,7 @@ static lean_object* l_Lean_Parser_Command_declId___elambda__1___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Command_abbrev_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_structExplicitBinder_formatter___closed__1; static lean_object* l_Lean_Parser_Command_printAxioms___elambda__1___closed__6; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__36; static lean_object* l_Lean_Parser_Command_declValSimple___closed__1; static lean_object* l_Lean_Parser_Command_structure___elambda__1___closed__2; lean_object* l_Lean_Parser_ident_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -3599,6 +3608,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_decreasingBy_parenthesize static lean_object* l_Lean_Parser_Command_resolve__name_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_builtin__initialize_declRange___closed__1; static lean_object* l_Lean_Parser_Command_openHiding___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__26; static lean_object* l_Lean_Parser_Command_check__failure___elambda__1___closed__6; static lean_object* l_Lean_Parser_Command_moduleDoc_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_print_declRange___closed__1; @@ -3623,6 +3633,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_openHiding_parenthes static lean_object* l_Lean_Parser_Command_noncomputableSection___elambda__1___closed__11; static lean_object* l_Lean_Parser_Command_terminationBy_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_declModifiers_formatter___closed__2; +static lean_object* l_Lean_Parser_Command_declVal_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_partial_formatter___closed__1; static lean_object* l_Lean_Parser_Command_set__option_formatter___closed__7; static lean_object* l_Lean_Parser_Command_nonrec___closed__6; @@ -3770,6 +3781,7 @@ static lean_object* l_Lean_Parser_Term_open___elambda__1___closed__6; static lean_object* l_Lean_Parser_Command_terminationSuffix___elambda__1___closed__4; static lean_object* l_Lean_Parser_Command_classInductive_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_def_formatter___closed__8; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__12; static lean_object* l_Lean_Parser_Command_set__option___closed__1; static lean_object* l_Lean_Parser_Command_terminationBy___closed__3; static lean_object* l_Lean_Parser_Command_declId___elambda__1___closed__8; @@ -3850,6 +3862,7 @@ static lean_object* l_Lean_Parser_Command_attribute___elambda__1___closed__8; static lean_object* l_Lean_Parser_Command_declModifiers_formatter___closed__13; static lean_object* l_Lean_Parser_Command_opaque_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Command_optDeriving___elambda__1___closed__12; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__31; static lean_object* l_Lean_Parser_Command_builtin__initialize___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_open(lean_object*); static lean_object* l_Lean_Parser_Command_instance_formatter___closed__9; @@ -3936,6 +3949,7 @@ static lean_object* l_Lean_Parser_Command_optDefDeriving_parenthesizer___closed_ LEAN_EXPORT lean_object* l_Lean_Parser_Command_partial; static lean_object* l_Lean_Parser_Command_declId___elambda__1___closed__4; static lean_object* l_Lean_Parser_Command_initialize___elambda__1___closed__11; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__43; static lean_object* l_Lean_Parser_Command_declModifiers___closed__11; static lean_object* l_Lean_Parser_Command_mutual___elambda__1___closed__13; static lean_object* l_Lean_Parser_Command_builtin__initialize___closed__2; @@ -4083,7 +4097,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_structFields_parenthesize static lean_object* l_Lean_Parser_Command_classInductive_formatter___closed__4; static lean_object* l_Lean_Parser_Term_set__option___elambda__1___closed__4; static lean_object* l_Lean_Parser_Command_instance_parenthesizer___closed__6; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__22; extern lean_object* l_Lean_Parser_Term_matchAltsWhereDecls; static lean_object* l___regBuiltin_Lean_Parser_Command_structFields_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_section_declRange___closed__6; @@ -4109,8 +4122,6 @@ static lean_object* l_Lean_Parser_Command_printAxioms_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Command_builtin__initialize_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Tactic_set__option___elambda__1___closed__8; static lean_object* l_Lean_Parser_Command_openHiding_parenthesizer___closed__3; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__10; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__32; static lean_object* l_Lean_Parser_Tactic_open___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Command_printAxioms; static lean_object* l_Lean_Parser_Command_whereStructInst___elambda__1___closed__29; @@ -4127,10 +4138,10 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_whereStructInst___elambda__1(lean static lean_object* l_Lean_Parser_Term_precheckedQuot___elambda__1___closed__9; static lean_object* l_Lean_Parser_Command_extends___closed__6; static lean_object* l_Lean_Parser_Command_namedPrio_formatter___closed__7; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_precheckedQuot_formatter___closed__2; static lean_object* l_Lean_Parser_Command_declId___elambda__1___closed__18; static lean_object* l_Lean_Parser_Command_openHiding___elambda__1___closed__17; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__16; static lean_object* l_Lean_Parser_Command_structFields___elambda__1___closed__20; lean_object* l_Lean_Parser_unicodeSymbolInfo(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_openDecl___closed__1; @@ -4149,7 +4160,6 @@ static lean_object* l_Lean_Parser_Command_openHiding___elambda__1___closed__4; static lean_object* l_Lean_Parser_Command_derivingClasses_formatter___closed__6; static lean_object* l_Lean_Parser_Command_structInstBinder_parenthesizer___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Tactic_open_declRange___closed__3; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__49; static lean_object* l___regBuiltin_Lean_Parser_Command_initialize_declRange___closed__1; static lean_object* l_Lean_Parser_Command_protected___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Command_axiom; @@ -4192,6 +4202,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_open___elambda__1___lambda__1(lea static lean_object* l_Lean_Parser_Command_structExplicitBinder_formatter___closed__12; static lean_object* l_Lean_Parser_Command_optionValue_parenthesizer___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_abbrev_formatter(lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__19; static lean_object* l_Lean_Parser_Command_openRenaming_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Command_whereStructInst___elambda__1___closed__33; static lean_object* l_Lean_Parser_Command_mutual___elambda__1___closed__14; @@ -4274,6 +4285,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_structInstBinder_parenthe static lean_object* l___regBuiltin_Lean_Parser_Term_quot_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_declValEqns___elambda__1___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Command_mutual_formatter___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__39; static lean_object* l_Lean_Parser_Command_mutual_formatter___closed__12; static lean_object* l_Lean_Parser_Command_end___closed__1; static lean_object* l_Lean_Parser_Command_structure___elambda__1___closed__6; @@ -4303,6 +4315,7 @@ static lean_object* l_Lean_Parser_Command_whereStructInst___elambda__1___closed_ static lean_object* l_Lean_Parser_Command_theorem___elambda__1___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Command_declId_formatter___closed__2; static lean_object* l_Lean_Parser_Command_builtin__initialize___elambda__1___closed__4; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__11; static lean_object* l_Lean_Parser_Command_opaque___elambda__1___closed__8; static lean_object* l_Lean_Parser_Command_partial___elambda__1___closed__4; static lean_object* l_Lean_Parser_Command_resolve__name___elambda__1___closed__10; @@ -4325,7 +4338,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_genInjectiveTheorems_form static lean_object* l_Lean_Parser_Command_instance_parenthesizer___closed__9; static lean_object* l_Lean_Parser_Command_synth___elambda__1___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_noncomputableSection_formatter___closed__2; -uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_optDefDeriving___closed__8; static lean_object* l_Lean_Parser_Command_structCtor_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_structureTk_parenthesizer___closed__1; @@ -4348,7 +4360,6 @@ static lean_object* l_Lean_Parser_Command_namespace___closed__6; static lean_object* l_Lean_Parser_Command_theorem_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_initialize_declRange___closed__6; static lean_object* l_Lean_Parser_Command_classInductive___elambda__1___closed__12; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__14; static lean_object* l_Lean_Parser_Command_init__quot___closed__4; static lean_object* l_Lean_Parser_Command_terminationByCore___closed__7; static lean_object* l_Lean_Parser_Term_precheckedQuot___closed__7; @@ -4376,6 +4387,7 @@ static lean_object* l_Lean_Parser_Command_mutual___closed__5; static lean_object* l_Lean_Parser_Tactic_set__option___closed__8; static lean_object* l_Lean_Parser_Command_derivingClasses_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_declSig_formatter___closed__3; +static lean_object* l_Lean_Parser_Command_declVal_formatter___closed__2; static lean_object* l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__12; static lean_object* l___regBuiltin_Lean_Parser_Command_check__failure_declRange___closed__6; static lean_object* l_Lean_Parser_Command_openScoped___elambda__1___closed__2; @@ -4522,8 +4534,6 @@ static lean_object* l_Lean_Parser_Command_attribute_parenthesizer___closed__10; static lean_object* l_Lean_Parser_Command_check___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_precheckedQuot(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_eval_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__24; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__7; extern lean_object* l_Lean_Parser_strLit; static lean_object* l_Lean_Parser_Command_mutual___elambda__1___closed__25; static lean_object* l_Lean_Parser_Command_eraseAttr_formatter___closed__3; @@ -5001,7 +5011,7 @@ x_25 = l_Lean_Parser_checkPrecFn(x_24, x_1, x_2); x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); x_27 = lean_box(0); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_27); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_27); lean_dec(x_26); if (x_28 == 0) { @@ -5043,7 +5053,7 @@ goto block_67; lean_object* x_38; uint8_t x_39; x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_27); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_27); lean_dec(x_38); if (x_39 == 0) { @@ -5052,7 +5062,7 @@ lean_dec(x_19); x_40 = l_Lean_Parser_ParserState_mkNode(x_37, x_10, x_30); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_27); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_27); lean_dec(x_41); if (x_42 == 0) { @@ -5076,7 +5086,7 @@ lean_inc(x_1); x_46 = l___private_Lean_Parser_Basic_0__Lean_Parser_addQuotDepthFn(x_44, x_45, x_1, x_37); x_47 = lean_ctor_get(x_46, 4); lean_inc(x_47); -x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_27); +x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_27); lean_dec(x_47); if (x_48 == 0) { @@ -5085,7 +5095,7 @@ lean_dec(x_19); x_49 = l_Lean_Parser_ParserState_mkNode(x_46, x_10, x_30); x_50 = lean_ctor_get(x_49, 4); lean_inc(x_50); -x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_27); +x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_27); lean_dec(x_50); if (x_51 == 0) { @@ -5118,7 +5128,7 @@ lean_object* x_58; lean_object* x_59; uint8_t x_60; x_58 = l_Lean_Parser_ParserState_mkNode(x_54, x_10, x_30); x_59 = lean_ctor_get(x_58, 4); lean_inc(x_59); -x_60 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_59, x_27); +x_60 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_59, x_27); lean_dec(x_59); if (x_60 == 0) { @@ -5141,7 +5151,7 @@ x_62 = l_Lean_Parser_tokenAntiquotFn(x_1, x_54); x_63 = l_Lean_Parser_ParserState_mkNode(x_62, x_10, x_30); x_64 = lean_ctor_get(x_63, 4); lean_inc(x_64); -x_65 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_64, x_27); +x_65 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_64, x_27); lean_dec(x_64); if (x_65 == 0) { @@ -5887,7 +5897,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -5918,7 +5928,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_20, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -5927,7 +5937,7 @@ x_27 = l_Lean_Parser_Term_precheckedQuot___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_20, x_27, x_17); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -5951,7 +5961,7 @@ x_33 = l_Lean_Parser_Term_precheckedQuot___elambda__1___closed__2; x_34 = l_Lean_Parser_ParserState_mkNode(x_32, x_33, x_17); x_35 = lean_ctor_get(x_34, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_14); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_14); lean_dec(x_35); if (x_36 == 0) { @@ -5974,7 +5984,7 @@ lean_inc(x_1); x_38 = l_Lean_Parser_tokenAntiquotFn(x_1, x_20); x_39 = lean_ctor_get(x_38, 4); lean_inc(x_39); -x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); +x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); lean_dec(x_39); if (x_40 == 0) { @@ -5983,7 +5993,7 @@ x_41 = l_Lean_Parser_Term_precheckedQuot___elambda__1___closed__2; x_42 = l_Lean_Parser_ParserState_mkNode(x_38, x_41, x_17); x_43 = lean_ctor_get(x_42, 4); lean_inc(x_43); -x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_14); +x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_14); lean_dec(x_43); if (x_44 == 0) { @@ -6007,7 +6017,7 @@ x_47 = l_Lean_Parser_Term_precheckedQuot___elambda__1___closed__2; x_48 = l_Lean_Parser_ParserState_mkNode(x_46, x_47, x_17); x_49 = lean_ctor_get(x_48, 4); lean_inc(x_49); -x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_14); +x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_14); lean_dec(x_49); if (x_50 == 0) { @@ -6671,7 +6681,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -6714,7 +6724,7 @@ goto block_60; lean_object* x_26; uint8_t x_27; x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); lean_dec(x_26); if (x_27 == 0) { @@ -6724,7 +6734,7 @@ x_28 = l_Lean_Parser_Command_quot___elambda__1___closed__3; x_29 = l_Lean_Parser_ParserState_mkNode(x_25, x_28, x_17); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); lean_dec(x_30); if (x_31 == 0) { @@ -6748,7 +6758,7 @@ lean_inc(x_1); x_35 = l___private_Lean_Parser_Basic_0__Lean_Parser_addQuotDepthFn(x_33, x_34, x_1, x_25); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); lean_dec(x_36); if (x_37 == 0) { @@ -6758,7 +6768,7 @@ x_38 = l_Lean_Parser_Command_quot___elambda__1___closed__3; x_39 = l_Lean_Parser_ParserState_mkNode(x_35, x_38, x_17); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); lean_dec(x_40); if (x_41 == 0) { @@ -6793,7 +6803,7 @@ x_49 = l_Lean_Parser_Command_quot___elambda__1___closed__3; x_50 = l_Lean_Parser_ParserState_mkNode(x_45, x_49, x_17); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_14); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_14); lean_dec(x_51); if (x_52 == 0) { @@ -6817,7 +6827,7 @@ x_55 = l_Lean_Parser_Command_quot___elambda__1___closed__3; x_56 = l_Lean_Parser_ParserState_mkNode(x_54, x_55, x_17); x_57 = lean_ctor_get(x_56, 4); lean_inc(x_57); -x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_14); +x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_14); lean_dec(x_57); if (x_58 == 0) { @@ -7636,7 +7646,7 @@ static lean_object* _init_l_Lean_Parser_Command_terminationHintMany___closed__17 _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_id___rarg___boxed), 1, 0); +x_1 = lean_mk_string_from_bytes("group", 5); return x_1; } } @@ -7644,7 +7654,25 @@ static lean_object* _init_l_Lean_Parser_Command_terminationHintMany___closed__18 _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_terminationHintMany___closed__17; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_Command_terminationHintMany___closed__17; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Command_terminationHintMany___closed__19() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_id___rarg___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Command_terminationHintMany___closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_terminationHintMany___closed__19; x_2 = lean_box(1); x_3 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_3, 0, x_1); @@ -7653,7 +7681,7 @@ lean_ctor_set(x_3, 2, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_terminationHintMany___closed__19() { +static lean_object* _init_l_Lean_Parser_Command_terminationHintMany___closed__21() { _start: { lean_object* x_1; @@ -7661,11 +7689,11 @@ x_1 = lean_mk_string_from_bytes("irrelevant", 10); return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_terminationHintMany___closed__20() { +static lean_object* _init_l_Lean_Parser_Command_terminationHintMany___closed__22() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_terminationHintMany___closed__19; +x_1 = l_Lean_Parser_Command_terminationHintMany___closed__21; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_checkColGeFn___boxed), 3, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -7709,14 +7737,14 @@ x_20 = l_Lean_Parser_Command_terminationHintMany___closed__16; x_21 = lean_alloc_closure((void*)(l_Lean_Parser_andthenFn), 4, 2); lean_closure_set(x_21, 0, x_20); lean_closure_set(x_21, 1, x_17); -x_22 = l_Lean_groupKind; +x_22 = l_Lean_Parser_Command_terminationHintMany___closed__18; x_23 = l_Lean_Parser_nodeInfo(x_22, x_19); x_24 = lean_alloc_closure((void*)(l_Lean_Parser_nodeFn), 4, 2); lean_closure_set(x_24, 0, x_22); lean_closure_set(x_24, 1, x_21); -x_25 = l_Lean_Parser_Command_terminationHintMany___closed__18; +x_25 = l_Lean_Parser_Command_terminationHintMany___closed__20; x_26 = l_Lean_Parser_andthenInfo(x_25, x_23); -x_27 = l_Lean_Parser_Command_terminationHintMany___closed__20; +x_27 = l_Lean_Parser_Command_terminationHintMany___closed__22; x_28 = lean_alloc_closure((void*)(l_Lean_Parser_andthenFn), 4, 2); lean_closure_set(x_28, 0, x_27); lean_closure_set(x_28, 1, x_24); @@ -8075,7 +8103,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -8107,7 +8135,7 @@ if (x_26 == 0) lean_object* x_27; uint8_t x_28; x_27 = lean_ctor_get(x_22, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); lean_dec(x_27); if (x_28 == 0) { @@ -8117,7 +8145,7 @@ x_29 = l_Lean_Parser_Command_terminationByCore___elambda__1___closed__2; x_30 = l_Lean_Parser_ParserState_mkNode(x_22, x_29, x_19); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); lean_dec(x_31); if (x_32 == 0) { @@ -8141,7 +8169,7 @@ x_35 = l_Lean_Parser_Command_terminationByCore___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_19); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_16); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_16); lean_dec(x_37); if (x_38 == 0) { @@ -8164,7 +8192,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_22); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); lean_dec(x_41); if (x_42 == 0) { @@ -8174,7 +8202,7 @@ x_43 = l_Lean_Parser_Command_terminationByCore___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_19); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); lean_dec(x_45); if (x_46 == 0) { @@ -8198,7 +8226,7 @@ x_49 = l_Lean_Parser_Command_terminationByCore___elambda__1___closed__2; x_50 = l_Lean_Parser_ParserState_mkNode(x_48, x_49, x_19); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); lean_dec(x_51); if (x_52 == 0) { @@ -8485,7 +8513,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -8517,7 +8545,7 @@ if (x_26 == 0) lean_object* x_27; uint8_t x_28; x_27 = lean_ctor_get(x_22, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); lean_dec(x_27); if (x_28 == 0) { @@ -8527,7 +8555,7 @@ x_29 = l_Lean_Parser_Command_decreasingBy___elambda__1___closed__2; x_30 = l_Lean_Parser_ParserState_mkNode(x_22, x_29, x_19); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); lean_dec(x_31); if (x_32 == 0) { @@ -8551,7 +8579,7 @@ x_35 = l_Lean_Parser_Command_decreasingBy___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_19); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_16); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_16); lean_dec(x_37); if (x_38 == 0) { @@ -8574,7 +8602,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_22); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); lean_dec(x_41); if (x_42 == 0) { @@ -8584,7 +8612,7 @@ x_43 = l_Lean_Parser_Command_decreasingBy___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_19); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); lean_dec(x_45); if (x_46 == 0) { @@ -8608,7 +8636,7 @@ x_49 = l_Lean_Parser_Command_decreasingBy___elambda__1___closed__2; x_50 = l_Lean_Parser_ParserState_mkNode(x_48, x_49, x_19); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); lean_dec(x_51); if (x_52 == 0) { @@ -8993,7 +9021,7 @@ x_16 = l_Lean_Parser_checkPrecFn(x_15, x_1, x_2); x_17 = lean_ctor_get(x_16, 4); lean_inc(x_17); x_18 = lean_box(0); -x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); +x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); lean_dec(x_17); if (x_19 == 0) { @@ -9017,7 +9045,7 @@ lean_inc(x_1); x_32 = l_Lean_Parser_orelseFnCore(x_29, x_30, x_31, x_1, x_16); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_18); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_18); lean_dec(x_33); if (x_34 == 0) { @@ -9034,7 +9062,7 @@ lean_inc(x_1); x_35 = lean_apply_2(x_6, x_1, x_32); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_18); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_18); lean_dec(x_36); if (x_37 == 0) { @@ -9062,7 +9090,7 @@ if (x_44 == 0) lean_object* x_45; uint8_t x_46; x_45 = lean_ctor_get(x_40, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_18); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_18); lean_dec(x_45); if (x_46 == 0) { @@ -9079,7 +9107,7 @@ lean_inc(x_1); x_49 = l_Lean_Parser_categoryParser___elambda__1(x_47, x_48, x_1, x_40); x_50 = lean_ctor_get(x_49, 4); lean_inc(x_50); -x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_18); +x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_18); lean_dec(x_50); if (x_51 == 0) { @@ -9104,7 +9132,7 @@ lean_inc(x_1); x_53 = l_Lean_Parser_tokenAntiquotFn(x_1, x_40); x_54 = lean_ctor_get(x_53, 4); lean_inc(x_54); -x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_18); +x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_18); lean_dec(x_54); if (x_55 == 0) { @@ -9121,7 +9149,7 @@ lean_inc(x_1); x_58 = l_Lean_Parser_categoryParser___elambda__1(x_56, x_57, x_1, x_53); x_59 = lean_ctor_get(x_58, 4); lean_inc(x_59); -x_60 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_59, x_18); +x_60 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_59, x_18); lean_dec(x_59); if (x_60 == 0) { @@ -9148,7 +9176,7 @@ x_23 = l_Lean_Parser_Command_terminationByElement___elambda__1___closed__2; x_24 = l_Lean_Parser_ParserState_mkNode(x_22, x_23, x_21); x_25 = lean_ctor_get(x_24, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_18); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_18); lean_dec(x_25); if (x_26 == 0) { @@ -9368,7 +9396,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_terminationByElement; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l_Lean_Parser_Command_terminationHintMany___closed__18; +x_3 = l_Lean_Parser_Command_terminationHintMany___closed__20; x_4 = l_Lean_Parser_andthenInfo(x_3, x_2); return x_4; } @@ -9377,7 +9405,7 @@ static lean_object* _init_l_Lean_Parser_Command_terminationBy___elambda__1___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_terminationHintMany___closed__20; +x_1 = l_Lean_Parser_Command_terminationHintMany___closed__22; x_2 = l_Lean_Parser_Command_terminationByElement___closed__10; x_3 = lean_alloc_closure((void*)(l_Lean_Parser_andthenFn), 4, 2); lean_closure_set(x_3, 0, x_1); @@ -9533,7 +9561,7 @@ x_18 = l_Lean_Parser_checkPrecFn(x_17, x_1, x_2); x_19 = lean_ctor_get(x_18, 4); lean_inc(x_19); x_20 = lean_box(0); -x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_20); +x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_20); lean_dec(x_19); if (x_21 == 0) { @@ -9569,7 +9597,7 @@ if (x_37 == 0) lean_object* x_38; uint8_t x_39; x_38 = lean_ctor_get(x_33, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_20); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_20); lean_dec(x_38); if (x_39 == 0) { @@ -9613,7 +9641,7 @@ lean_inc(x_1); x_45 = l_Lean_Parser_tokenAntiquotFn(x_1, x_33); x_46 = lean_ctor_get(x_45, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_20); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_20); lean_dec(x_46); if (x_47 == 0) { @@ -9657,7 +9685,7 @@ x_25 = l_Lean_Parser_Command_terminationBy___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_24, x_25, x_23); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_20); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_20); lean_dec(x_27); if (x_28 == 0) { @@ -9863,7 +9891,7 @@ x_7 = lean_apply_2(x_4, x_1, x_2); x_8 = lean_ctor_get(x_7, 4); lean_inc(x_8); x_9 = lean_box(0); -x_10 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_8, x_9); +x_10 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_8, x_9); lean_dec(x_8); if (x_10 == 0) { @@ -10093,7 +10121,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -10124,7 +10152,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_20, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -10133,7 +10161,7 @@ x_27 = l_Lean_Parser_Command_moduleDoc___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_20, x_27, x_17); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -10156,7 +10184,7 @@ x_33 = l_Lean_Parser_Command_moduleDoc___elambda__1___closed__2; x_34 = l_Lean_Parser_ParserState_mkNode(x_32, x_33, x_17); x_35 = lean_ctor_get(x_34, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_14); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_14); lean_dec(x_35); if (x_36 == 0) { @@ -10179,7 +10207,7 @@ lean_inc(x_1); x_38 = l_Lean_Parser_tokenAntiquotFn(x_1, x_20); x_39 = lean_ctor_get(x_38, 4); lean_inc(x_39); -x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); +x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); lean_dec(x_39); if (x_40 == 0) { @@ -10188,7 +10216,7 @@ x_41 = l_Lean_Parser_Command_moduleDoc___elambda__1___closed__2; x_42 = l_Lean_Parser_ParserState_mkNode(x_38, x_41, x_17); x_43 = lean_ctor_get(x_42, 4); lean_inc(x_43); -x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_14); +x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_14); lean_dec(x_43); if (x_44 == 0) { @@ -10211,7 +10239,7 @@ x_47 = l_Lean_Parser_Command_moduleDoc___elambda__1___closed__2; x_48 = l_Lean_Parser_ParserState_mkNode(x_46, x_47, x_17); x_49 = lean_ctor_get(x_48, 4); lean_inc(x_49); -x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_14); +x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_14); lean_dec(x_49); if (x_50 == 0) { @@ -10983,7 +11011,7 @@ x_25 = l_Lean_Parser_checkPrecFn(x_24, x_1, x_2); x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); x_27 = lean_box(0); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_27); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_27); lean_dec(x_26); if (x_28 == 0) { @@ -11003,7 +11031,7 @@ lean_inc(x_1); x_31 = l_Lean_Parser_atomicFn(x_6, x_1, x_25); x_32 = lean_ctor_get(x_31, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_27); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_27); lean_dec(x_32); if (x_33 == 0) { @@ -11012,7 +11040,7 @@ lean_dec(x_19); x_34 = l_Lean_Parser_ParserState_mkNode(x_31, x_10, x_30); x_35 = lean_ctor_get(x_34, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_27); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_27); lean_dec(x_35); if (x_36 == 0) { @@ -11058,7 +11086,7 @@ goto block_76; lean_object* x_46; uint8_t x_47; x_46 = lean_ctor_get(x_45, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_27); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_27); lean_dec(x_46); if (x_47 == 0) { @@ -11067,7 +11095,7 @@ lean_dec(x_19); x_48 = l_Lean_Parser_ParserState_mkNode(x_45, x_10, x_30); x_49 = lean_ctor_get(x_48, 4); lean_inc(x_49); -x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_27); +x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_27); lean_dec(x_49); if (x_50 == 0) { @@ -11091,7 +11119,7 @@ lean_inc(x_1); x_54 = l_Lean_Parser_categoryParser___elambda__1(x_52, x_53, x_1, x_45); x_55 = lean_ctor_get(x_54, 4); lean_inc(x_55); -x_56 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_55, x_27); +x_56 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_55, x_27); lean_dec(x_55); if (x_56 == 0) { @@ -11100,7 +11128,7 @@ lean_dec(x_19); x_57 = l_Lean_Parser_ParserState_mkNode(x_54, x_10, x_30); x_58 = lean_ctor_get(x_57, 4); lean_inc(x_58); -x_59 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_58, x_27); +x_59 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_58, x_27); lean_dec(x_58); if (x_59 == 0) { @@ -11134,7 +11162,7 @@ lean_object* x_67; lean_object* x_68; uint8_t x_69; x_67 = l_Lean_Parser_ParserState_mkNode(x_63, x_10, x_30); x_68 = lean_ctor_get(x_67, 4); lean_inc(x_68); -x_69 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_68, x_27); +x_69 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_68, x_27); lean_dec(x_68); if (x_69 == 0) { @@ -11157,7 +11185,7 @@ x_71 = l_Lean_Parser_tokenAntiquotFn(x_1, x_63); x_72 = l_Lean_Parser_ParserState_mkNode(x_71, x_10, x_30); x_73 = lean_ctor_get(x_72, 4); lean_inc(x_73); -x_74 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_73, x_27); +x_74 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_73, x_27); lean_dec(x_73); if (x_74 == 0) { @@ -11532,7 +11560,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -11565,7 +11593,7 @@ x_25 = l_Lean_Parser_Command_private___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_20, x_25, x_17); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); lean_dec(x_27); if (x_28 == 0) { @@ -11589,7 +11617,7 @@ x_31 = l_Lean_Parser_Command_private___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_17); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -11836,7 +11864,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -11869,7 +11897,7 @@ x_25 = l_Lean_Parser_Command_protected___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_20, x_25, x_17); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); lean_dec(x_27); if (x_28 == 0) { @@ -11893,7 +11921,7 @@ x_31 = l_Lean_Parser_Command_protected___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_17); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -12193,7 +12221,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -12226,7 +12254,7 @@ x_25 = l_Lean_Parser_Command_noncomputable___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_20, x_25, x_17); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); lean_dec(x_27); if (x_28 == 0) { @@ -12250,7 +12278,7 @@ x_31 = l_Lean_Parser_Command_noncomputable___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_17); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -12497,7 +12525,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -12530,7 +12558,7 @@ x_25 = l_Lean_Parser_Command_unsafe___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_20, x_25, x_17); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); lean_dec(x_27); if (x_28 == 0) { @@ -12554,7 +12582,7 @@ x_31 = l_Lean_Parser_Command_unsafe___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_17); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -12801,7 +12829,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -12834,7 +12862,7 @@ x_25 = l_Lean_Parser_Command_partial___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_20, x_25, x_17); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); lean_dec(x_27); if (x_28 == 0) { @@ -12858,7 +12886,7 @@ x_31 = l_Lean_Parser_Command_partial___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_17); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -13105,7 +13133,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -13138,7 +13166,7 @@ x_25 = l_Lean_Parser_Command_nonrec___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_20, x_25, x_17); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); lean_dec(x_27); if (x_28 == 0) { @@ -13162,7 +13190,7 @@ x_31 = l_Lean_Parser_Command_nonrec___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_17); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -13612,7 +13640,7 @@ x_24 = l_Lean_Parser_checkPrecFn(x_23, x_1, x_2); x_25 = lean_ctor_get(x_24, 4); lean_inc(x_25); x_26 = lean_box(0); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_26); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_26); lean_dec(x_25); if (x_27 == 0) { @@ -13636,7 +13664,7 @@ lean_inc(x_1); x_30 = lean_apply_2(x_14, x_1, x_24); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_26); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_26); lean_dec(x_31); if (x_32 == 0) { @@ -13655,7 +13683,7 @@ lean_inc(x_1); x_40 = lean_apply_2(x_12, x_1, x_30); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_26); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_26); lean_dec(x_41); if (x_42 == 0) { @@ -13673,7 +13701,7 @@ lean_inc(x_1); x_43 = lean_apply_2(x_10, x_1, x_40); x_44 = lean_ctor_get(x_43, 4); lean_inc(x_44); -x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_26); +x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_26); lean_dec(x_44); if (x_45 == 0) { @@ -13690,7 +13718,7 @@ lean_inc(x_1); x_46 = lean_apply_2(x_8, x_1, x_43); x_47 = lean_ctor_get(x_46, 4); lean_inc(x_47); -x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_26); +x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_26); lean_dec(x_47); if (x_48 == 0) { @@ -13706,7 +13734,7 @@ lean_inc(x_1); x_49 = lean_apply_2(x_4, x_1, x_46); x_50 = lean_ctor_get(x_49, 4); lean_inc(x_50); -x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_26); +x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_26); lean_dec(x_50); if (x_51 == 0) { @@ -13733,7 +13761,7 @@ x_34 = l_Lean_Parser_Command_declModifiers___elambda__2___closed__2; x_35 = l_Lean_Parser_ParserState_mkNode(x_33, x_34, x_29); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_26); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_26); lean_dec(x_36); if (x_37 == 0) { @@ -14238,7 +14266,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -14257,7 +14285,7 @@ lean_inc(x_1); x_20 = l_Lean_Parser_ident___elambda__1(x_1, x_14); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); -x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_16); +x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_16); lean_dec(x_21); if (x_22 == 0) { @@ -14267,7 +14295,7 @@ x_23 = l_Lean_Parser_Command_declId___elambda__1___closed__2; x_24 = l_Lean_Parser_ParserState_mkNode(x_20, x_23, x_19); x_25 = lean_ctor_get(x_24, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_16); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_16); lean_dec(x_25); if (x_26 == 0) { @@ -14291,7 +14319,7 @@ x_29 = l_Lean_Parser_Command_declId___elambda__1___closed__2; x_30 = l_Lean_Parser_ParserState_mkNode(x_28, x_29, x_19); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); lean_dec(x_31); if (x_32 == 0) { @@ -14604,7 +14632,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -14623,7 +14651,7 @@ lean_inc(x_1); x_20 = lean_apply_2(x_4, x_1, x_14); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); -x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_16); +x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_16); lean_dec(x_21); if (x_22 == 0) { @@ -14632,7 +14660,7 @@ x_23 = l_Lean_Parser_Command_declSig___elambda__1___closed__2; x_24 = l_Lean_Parser_ParserState_mkNode(x_20, x_23, x_19); x_25 = lean_ctor_get(x_24, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_16); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_16); lean_dec(x_25); if (x_26 == 0) { @@ -14656,7 +14684,7 @@ x_29 = l_Lean_Parser_Command_declSig___elambda__1___closed__2; x_30 = l_Lean_Parser_ParserState_mkNode(x_28, x_29, x_19); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); lean_dec(x_31); if (x_32 == 0) { @@ -14884,7 +14912,7 @@ x_16 = l_Lean_Parser_checkPrecFn(x_15, x_1, x_2); x_17 = lean_ctor_get(x_16, 4); lean_inc(x_17); x_18 = lean_box(0); -x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); +x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); lean_dec(x_17); if (x_19 == 0) { @@ -14904,7 +14932,7 @@ lean_inc(x_1); x_22 = lean_apply_2(x_4, x_1, x_16); x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); -x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_18); +x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_18); lean_dec(x_23); if (x_24 == 0) { @@ -14914,7 +14942,7 @@ x_25 = l_Lean_Parser_Command_optDeclSig___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_22, x_25, x_21); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_18); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_18); lean_dec(x_27); if (x_28 == 0) { @@ -14938,7 +14966,7 @@ x_31 = l_Lean_Parser_Command_optDeclSig___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_21); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_18); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_18); lean_dec(x_33); if (x_34 == 0) { @@ -15241,7 +15269,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -15286,7 +15314,7 @@ goto block_51; lean_object* x_28; uint8_t x_29; x_28 = lean_ctor_get(x_27, 4); lean_inc(x_28); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_16); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_16); lean_dec(x_28); if (x_29 == 0) { @@ -15296,7 +15324,7 @@ x_30 = l_Lean_Parser_Command_declValSimple___elambda__1___closed__2; x_31 = l_Lean_Parser_ParserState_mkNode(x_27, x_30, x_19); x_32 = lean_ctor_get(x_31, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_16); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_16); lean_dec(x_32); if (x_33 == 0) { @@ -15320,7 +15348,7 @@ lean_inc(x_1); x_37 = l_Lean_Parser_categoryParser___elambda__1(x_35, x_36, x_1, x_27); x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_16); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_16); lean_dec(x_38); if (x_39 == 0) { @@ -15330,7 +15358,7 @@ x_40 = l_Lean_Parser_Command_declValSimple___elambda__1___closed__2; x_41 = l_Lean_Parser_ParserState_mkNode(x_37, x_40, x_19); x_42 = lean_ctor_get(x_41, 4); lean_inc(x_42); -x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_16); +x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_16); lean_dec(x_42); if (x_43 == 0) { @@ -15354,7 +15382,7 @@ x_46 = l_Lean_Parser_Command_declValSimple___elambda__1___closed__2; x_47 = l_Lean_Parser_ParserState_mkNode(x_45, x_46, x_19); x_48 = lean_ctor_get(x_47, 4); lean_inc(x_48); -x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_16); +x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_16); lean_dec(x_48); if (x_49 == 0) { @@ -15600,7 +15628,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -15620,7 +15648,7 @@ x_19 = l_Lean_Parser_Command_declValEqns___elambda__1___closed__2; x_20 = l_Lean_Parser_ParserState_mkNode(x_18, x_19, x_17); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); -x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_14); +x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_14); lean_dec(x_21); if (x_22 == 0) { @@ -15820,7 +15848,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -15840,7 +15868,7 @@ x_19 = l_Lean_Parser_Command_whereStructField___elambda__1___closed__2; x_20 = l_Lean_Parser_ParserState_mkNode(x_18, x_19, x_17); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); -x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_14); +x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_14); lean_dec(x_21); if (x_22 == 0) { @@ -16115,7 +16143,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_whereStructInst___elambda__1___closed__18; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l_Lean_Parser_Command_terminationHintMany___closed__18; +x_3 = l_Lean_Parser_Command_terminationHintMany___closed__20; x_4 = l_Lean_Parser_andthenInfo(x_3, x_2); return x_4; } @@ -16127,7 +16155,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_whereStructInst___elambda__1___closed__18; x_2 = lean_ctor_get(x_1, 1); lean_inc(x_2); -x_3 = l_Lean_Parser_Command_terminationHintMany___closed__20; +x_3 = l_Lean_Parser_Command_terminationHintMany___closed__22; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_andthenFn), 4, 2); lean_closure_set(x_4, 0, x_3); lean_closure_set(x_4, 1, x_2); @@ -16374,7 +16402,7 @@ x_20 = l_Lean_Parser_checkPrecFn(x_19, x_1, x_2); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); x_22 = lean_box(0); -x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_22); +x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_22); lean_dec(x_21); if (x_23 == 0) { @@ -16425,7 +16453,7 @@ goto block_60; lean_object* x_34; uint8_t x_35; x_34 = lean_ctor_get(x_33, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_22); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_22); lean_dec(x_34); if (x_35 == 0) { @@ -16441,7 +16469,7 @@ x_36 = l_Lean_Parser_Command_whereStructInst___elambda__1___closed__2; x_37 = l_Lean_Parser_ParserState_mkNode(x_33, x_36, x_25); x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_22); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_22); lean_dec(x_38); if (x_39 == 0) { @@ -16478,7 +16506,7 @@ lean_ctor_set_uint8(x_45, sizeof(void*)*7, x_12); x_46 = lean_apply_2(x_42, x_45, x_33); x_47 = lean_ctor_get(x_46, 4); lean_inc(x_47); -x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_22); +x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_22); lean_dec(x_47); if (x_48 == 0) { @@ -16488,7 +16516,7 @@ x_49 = l_Lean_Parser_Command_whereStructInst___elambda__1___closed__2; x_50 = l_Lean_Parser_ParserState_mkNode(x_46, x_49, x_25); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_22); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_22); lean_dec(x_51); if (x_52 == 0) { @@ -16512,7 +16540,7 @@ x_55 = l_Lean_Parser_Command_whereStructInst___elambda__1___closed__2; x_56 = l_Lean_Parser_ParserState_mkNode(x_54, x_55, x_25); x_57 = lean_ctor_get(x_56, 4); lean_inc(x_57); -x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_22); +x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_22); lean_dec(x_57); if (x_58 == 0) { @@ -16655,6 +16683,35 @@ return x_1; static lean_object* _init_l_Lean_Parser_Command_declVal___elambda__1___closed__1() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("declVal", 7); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Command_declVal___elambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_quot___elambda__1___closed__2; +x_2 = l_Lean_Parser_Command_declVal___elambda__1___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Command_declVal___elambda__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Command_declVal___elambda__1___closed__1; +x_2 = l_Lean_Parser_Command_declVal___elambda__1___closed__2; +x_3 = 1; +x_4 = l_Lean_Parser_mkAntiquot(x_1, x_2, x_3, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Command_declVal___elambda__1___closed__4() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_declValEqns___closed__5; x_2 = l_Lean_Parser_Command_whereStructInst___closed__8; @@ -16664,15 +16721,55 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } +static lean_object* _init_l_Lean_Parser_Command_declVal___elambda__1___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_declValSimple___closed__9; +x_2 = l_Lean_Parser_Command_declVal___elambda__1___closed__4; +x_3 = lean_alloc_closure((void*)(l_Lean_Parser_orelseFn), 4, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Parser_Command_declVal___elambda__1(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; -x_3 = l_Lean_Parser_Command_declValSimple___closed__9; -x_4 = l_Lean_Parser_Command_declVal___elambda__1___closed__1; -x_5 = 1; -x_6 = l_Lean_Parser_orelseFnCore(x_3, x_4, x_5, x_1, x_2); -return x_6; +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint32_t x_8; uint32_t x_9; uint8_t x_10; +x_3 = l_Lean_Parser_Command_declVal___elambda__1___closed__3; +x_4 = lean_ctor_get(x_3, 1); +lean_inc(x_4); +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +lean_dec(x_5); +x_7 = lean_ctor_get(x_2, 2); +lean_inc(x_7); +x_8 = lean_string_utf8_get(x_6, x_7); +lean_dec(x_7); +lean_dec(x_6); +x_9 = 36; +x_10 = lean_uint32_dec_eq(x_8, x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; +lean_dec(x_4); +x_11 = l_Lean_Parser_Command_declValSimple___closed__9; +x_12 = l_Lean_Parser_Command_declVal___elambda__1___closed__4; +x_13 = 1; +x_14 = l_Lean_Parser_orelseFnCore(x_11, x_12, x_13, x_1, x_2); +return x_14; +} +else +{ +lean_object* x_15; uint8_t x_16; lean_object* x_17; +x_15 = l_Lean_Parser_Command_declVal___elambda__1___closed__5; +x_16 = 0; +x_17 = l_Lean_Parser_orelseFnCore(x_4, x_15, x_16, x_1, x_2); +return x_17; +} } } static lean_object* _init_l_Lean_Parser_Command_declVal___closed__1() { @@ -16704,17 +16801,29 @@ return x_4; static lean_object* _init_l_Lean_Parser_Command_declVal___closed__3() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Command_declVal___elambda__1___closed__3; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_3 = l_Lean_Parser_Command_declVal___closed__2; +x_4 = l_Lean_Parser_orelseInfo(x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Command_declVal___closed__4() { +_start: +{ lean_object* x_1; x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Command_declVal___elambda__1), 2, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_declVal___closed__4() { +static lean_object* _init_l_Lean_Parser_Command_declVal___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_declVal___closed__2; -x_2 = l_Lean_Parser_Command_declVal___closed__3; +x_1 = l_Lean_Parser_Command_declVal___closed__3; +x_2 = l_Lean_Parser_Command_declVal___closed__4; x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -16725,7 +16834,7 @@ static lean_object* _init_l_Lean_Parser_Command_declVal() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Command_declVal___closed__4; +x_1 = l_Lean_Parser_Command_declVal___closed__5; return x_1; } } @@ -16791,7 +16900,7 @@ static lean_object* _init_l_Lean_Parser_Command_abbrev___elambda__1___closed__7( { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_optDeclSig___closed__6; -x_2 = l_Lean_Parser_Command_declVal___closed__3; +x_2 = l_Lean_Parser_Command_declVal___closed__4; x_3 = lean_alloc_closure((void*)(l_Lean_Parser_andthenFn), 4, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); @@ -16905,7 +17014,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -16949,7 +17058,7 @@ goto block_55; lean_object* x_26; uint8_t x_27; x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); lean_dec(x_26); if (x_27 == 0) { @@ -16958,7 +17067,7 @@ x_28 = l_Lean_Parser_Command_abbrev___elambda__1___closed__2; x_29 = l_Lean_Parser_ParserState_mkNode(x_25, x_28, x_17); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); lean_dec(x_30); if (x_31 == 0) { @@ -16980,7 +17089,7 @@ lean_inc(x_1); x_33 = l_Lean_Parser_Command_declId___elambda__1(x_1, x_25); x_34 = lean_ctor_get(x_33, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_14); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_14); lean_dec(x_34); if (x_35 == 0) { @@ -16989,7 +17098,7 @@ x_36 = l_Lean_Parser_Command_abbrev___elambda__1___closed__2; x_37 = l_Lean_Parser_ParserState_mkNode(x_33, x_36, x_17); x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_14); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_14); lean_dec(x_38); if (x_39 == 0) { @@ -17011,7 +17120,7 @@ lean_inc(x_1); x_41 = l_Lean_Parser_Command_optDeclSig___elambda__1(x_1, x_33); x_42 = lean_ctor_get(x_41, 4); lean_inc(x_42); -x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_14); +x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_14); lean_dec(x_42); if (x_43 == 0) { @@ -17020,7 +17129,7 @@ x_44 = l_Lean_Parser_Command_abbrev___elambda__1___closed__2; x_45 = l_Lean_Parser_ParserState_mkNode(x_41, x_44, x_17); x_46 = lean_ctor_get(x_45, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_14); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_14); lean_dec(x_46); if (x_47 == 0) { @@ -17044,7 +17153,7 @@ x_50 = l_Lean_Parser_Command_abbrev___elambda__1___closed__2; x_51 = l_Lean_Parser_ParserState_mkNode(x_49, x_50, x_17); x_52 = lean_ctor_get(x_51, 4); lean_inc(x_52); -x_53 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_52, x_14); +x_53 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_52, x_14); lean_dec(x_52); if (x_53 == 0) { @@ -17270,7 +17379,7 @@ static lean_object* _init_l_Lean_Parser_Command_optDefDeriving___closed__9() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_optDefDeriving___closed__3; -x_2 = l_Lean_Parser_Command_terminationHintMany___closed__18; +x_2 = l_Lean_Parser_Command_terminationHintMany___closed__20; x_3 = l_Lean_Parser_andthenInfo(x_1, x_2); return x_3; } @@ -17427,7 +17536,7 @@ static lean_object* _init_l_Lean_Parser_Command_def___elambda__1___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_declVal___closed__3; +x_1 = l_Lean_Parser_Command_declVal___closed__4; x_2 = l_Lean_Parser_Command_def___elambda__1___closed__7; x_3 = lean_alloc_closure((void*)(l_Lean_Parser_andthenFn), 4, 2); lean_closure_set(x_3, 0, x_1); @@ -17557,7 +17666,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -17604,7 +17713,7 @@ x_28 = l_Lean_Parser_Command_def___elambda__1___closed__2; x_29 = l_Lean_Parser_ParserState_mkNode(x_27, x_28, x_19); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_16); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_16); lean_dec(x_30); if (x_31 == 0) { @@ -17624,7 +17733,7 @@ return x_32; lean_object* x_35; uint8_t x_36; x_35 = lean_ctor_get(x_34, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_16); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_16); lean_dec(x_35); if (x_36 == 0) { @@ -17639,7 +17748,7 @@ lean_inc(x_1); x_37 = l_Lean_Parser_Command_declId___elambda__1(x_1, x_34); x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_16); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_16); lean_dec(x_38); if (x_39 == 0) { @@ -17654,7 +17763,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_Command_optDeclSig___elambda__1(x_1, x_37); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); lean_dec(x_41); if (x_42 == 0) { @@ -17669,7 +17778,7 @@ lean_inc(x_1); x_43 = l_Lean_Parser_Command_declVal___elambda__1(x_1, x_40); x_44 = lean_ctor_get(x_43, 4); lean_inc(x_44); -x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_16); +x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_16); lean_dec(x_44); if (x_45 == 0) { @@ -17684,7 +17793,7 @@ lean_inc(x_1); x_46 = lean_apply_2(x_4, x_1, x_43); x_47 = lean_ctor_get(x_46, 4); lean_inc(x_47); -x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_16); +x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_16); lean_dec(x_47); if (x_48 == 0) { @@ -17918,7 +18027,7 @@ static lean_object* _init_l_Lean_Parser_Command_theorem___elambda__1___closed__7 _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_declVal___closed__3; +x_1 = l_Lean_Parser_Command_declVal___closed__4; x_2 = l_Lean_Parser_Command_terminationSuffix___closed__2; x_3 = lean_alloc_closure((void*)(l_Lean_Parser_andthenFn), 4, 2); lean_closure_set(x_3, 0, x_1); @@ -18045,7 +18154,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -18091,7 +18200,7 @@ x_26 = l_Lean_Parser_Command_theorem___elambda__1___closed__2; x_27 = l_Lean_Parser_ParserState_mkNode(x_25, x_26, x_17); x_28 = lean_ctor_get(x_27, 4); lean_inc(x_28); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_14); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_14); lean_dec(x_28); if (x_29 == 0) { @@ -18111,7 +18220,7 @@ return x_30; lean_object* x_33; uint8_t x_34; x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -18125,7 +18234,7 @@ lean_inc(x_1); x_35 = l_Lean_Parser_Command_declId___elambda__1(x_1, x_32); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); lean_dec(x_36); if (x_37 == 0) { @@ -18139,7 +18248,7 @@ lean_inc(x_1); x_38 = l_Lean_Parser_Command_declSig___elambda__1(x_1, x_35); x_39 = lean_ctor_get(x_38, 4); lean_inc(x_39); -x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); +x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); lean_dec(x_39); if (x_40 == 0) { @@ -18153,7 +18262,7 @@ lean_inc(x_1); x_41 = l_Lean_Parser_Command_declVal___elambda__1(x_1, x_38); x_42 = lean_ctor_get(x_41, 4); lean_inc(x_42); -x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_14); +x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_14); lean_dec(x_42); if (x_43 == 0) { @@ -18502,7 +18611,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -18547,7 +18656,7 @@ goto block_57; lean_object* x_28; uint8_t x_29; x_28 = lean_ctor_get(x_27, 4); lean_inc(x_28); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_16); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_16); lean_dec(x_28); if (x_29 == 0) { @@ -18557,7 +18666,7 @@ x_30 = l_Lean_Parser_Command_opaque___elambda__1___closed__2; x_31 = l_Lean_Parser_ParserState_mkNode(x_27, x_30, x_19); x_32 = lean_ctor_get(x_31, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_16); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_16); lean_dec(x_32); if (x_33 == 0) { @@ -18579,7 +18688,7 @@ lean_inc(x_1); x_35 = l_Lean_Parser_Command_declId___elambda__1(x_1, x_27); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_16); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_16); lean_dec(x_36); if (x_37 == 0) { @@ -18589,7 +18698,7 @@ x_38 = l_Lean_Parser_Command_opaque___elambda__1___closed__2; x_39 = l_Lean_Parser_ParserState_mkNode(x_35, x_38, x_19); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_16); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_16); lean_dec(x_40); if (x_41 == 0) { @@ -18611,7 +18720,7 @@ lean_inc(x_1); x_43 = l_Lean_Parser_Command_declSig___elambda__1(x_1, x_35); x_44 = lean_ctor_get(x_43, 4); lean_inc(x_44); -x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_16); +x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_16); lean_dec(x_44); if (x_45 == 0) { @@ -18621,7 +18730,7 @@ x_46 = l_Lean_Parser_Command_opaque___elambda__1___closed__2; x_47 = l_Lean_Parser_ParserState_mkNode(x_43, x_46, x_19); x_48 = lean_ctor_get(x_47, 4); lean_inc(x_48); -x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_16); +x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_16); lean_dec(x_48); if (x_49 == 0) { @@ -18645,7 +18754,7 @@ x_52 = l_Lean_Parser_Command_opaque___elambda__1___closed__2; x_53 = l_Lean_Parser_ParserState_mkNode(x_51, x_52, x_19); x_54 = lean_ctor_get(x_53, 4); lean_inc(x_54); -x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_16); +x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_16); lean_dec(x_54); if (x_55 == 0) { @@ -19008,7 +19117,7 @@ x_16 = l_Lean_Parser_checkPrecFn(x_15, x_1, x_2); x_17 = lean_ctor_get(x_16, 4); lean_inc(x_17); x_18 = lean_box(0); -x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); +x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); lean_dec(x_17); if (x_19 == 0) { @@ -19029,7 +19138,7 @@ lean_inc(x_1); x_22 = l_Lean_Parser_Term_attrKind___elambda__1(x_1, x_16); x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); -x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_18); +x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_18); lean_dec(x_23); if (x_24 == 0) { @@ -19041,7 +19150,7 @@ x_32 = l_Lean_Parser_Command_instance___elambda__1___closed__1; x_33 = l_Lean_Parser_ParserState_mkNode(x_22, x_32, x_21); x_34 = lean_ctor_get(x_33, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_18); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_18); lean_dec(x_34); if (x_35 == 0) { @@ -19088,7 +19197,7 @@ goto block_60; lean_object* x_45; uint8_t x_46; x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_18); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_18); lean_dec(x_45); if (x_46 == 0) { @@ -19104,7 +19213,7 @@ lean_inc(x_1); x_47 = lean_apply_2(x_6, x_1, x_44); x_48 = lean_ctor_get(x_47, 4); lean_inc(x_48); -x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_18); +x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_18); lean_dec(x_48); if (x_49 == 0) { @@ -19119,7 +19228,7 @@ lean_inc(x_1); x_50 = lean_apply_2(x_4, x_1, x_47); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_18); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_18); lean_dec(x_51); if (x_52 == 0) { @@ -19133,7 +19242,7 @@ lean_inc(x_1); x_53 = l_Lean_Parser_Command_declSig___elambda__1(x_1, x_50); x_54 = lean_ctor_get(x_53, 4); lean_inc(x_54); -x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_18); +x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_18); lean_dec(x_54); if (x_55 == 0) { @@ -19147,7 +19256,7 @@ lean_inc(x_1); x_56 = l_Lean_Parser_Command_declVal___elambda__1(x_1, x_53); x_57 = lean_ctor_get(x_56, 4); lean_inc(x_57); -x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_18); +x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_18); lean_dec(x_57); if (x_58 == 0) { @@ -19175,7 +19284,7 @@ x_26 = l_Lean_Parser_Command_instance___elambda__1___closed__1; x_27 = l_Lean_Parser_ParserState_mkNode(x_25, x_26, x_21); x_28 = lean_ctor_get(x_27, 4); lean_inc(x_28); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_18); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_18); lean_dec(x_28); if (x_29 == 0) { @@ -19494,7 +19603,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -19538,7 +19647,7 @@ goto block_47; lean_object* x_26; uint8_t x_27; x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); lean_dec(x_26); if (x_27 == 0) { @@ -19547,7 +19656,7 @@ x_28 = l_Lean_Parser_Command_axiom___elambda__1___closed__2; x_29 = l_Lean_Parser_ParserState_mkNode(x_25, x_28, x_17); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); lean_dec(x_30); if (x_31 == 0) { @@ -19569,7 +19678,7 @@ lean_inc(x_1); x_33 = l_Lean_Parser_Command_declId___elambda__1(x_1, x_25); x_34 = lean_ctor_get(x_33, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_14); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_14); lean_dec(x_34); if (x_35 == 0) { @@ -19578,7 +19687,7 @@ x_36 = l_Lean_Parser_Command_axiom___elambda__1___closed__2; x_37 = l_Lean_Parser_ParserState_mkNode(x_33, x_36, x_17); x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_14); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_14); lean_dec(x_38); if (x_39 == 0) { @@ -19602,7 +19711,7 @@ x_42 = l_Lean_Parser_Command_axiom___elambda__1___closed__2; x_43 = l_Lean_Parser_ParserState_mkNode(x_41, x_42, x_17); x_44 = lean_ctor_get(x_43, 4); lean_inc(x_44); -x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_14); +x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_14); lean_dec(x_44); if (x_45 == 0) { @@ -19789,7 +19898,7 @@ static lean_object* _init_l_Lean_Parser_Command_example___elambda__1___closed__6 { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_declSig___closed__6; -x_2 = l_Lean_Parser_Command_declVal___closed__3; +x_2 = l_Lean_Parser_Command_declVal___closed__4; x_3 = lean_alloc_closure((void*)(l_Lean_Parser_andthenFn), 4, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); @@ -19891,7 +20000,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -19935,7 +20044,7 @@ goto block_47; lean_object* x_26; uint8_t x_27; x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); lean_dec(x_26); if (x_27 == 0) { @@ -19944,7 +20053,7 @@ x_28 = l_Lean_Parser_Command_example___elambda__1___closed__2; x_29 = l_Lean_Parser_ParserState_mkNode(x_25, x_28, x_17); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); lean_dec(x_30); if (x_31 == 0) { @@ -19966,7 +20075,7 @@ lean_inc(x_1); x_33 = l_Lean_Parser_Command_declSig___elambda__1(x_1, x_25); x_34 = lean_ctor_get(x_33, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_14); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_14); lean_dec(x_34); if (x_35 == 0) { @@ -19975,7 +20084,7 @@ x_36 = l_Lean_Parser_Command_example___elambda__1___closed__2; x_37 = l_Lean_Parser_ParserState_mkNode(x_33, x_36, x_17); x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_14); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_14); lean_dec(x_38); if (x_39 == 0) { @@ -19999,7 +20108,7 @@ x_42 = l_Lean_Parser_Command_example___elambda__1___closed__2; x_43 = l_Lean_Parser_ParserState_mkNode(x_41, x_42, x_17); x_44 = lean_ctor_get(x_43, 4); lean_inc(x_44); -x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_14); +x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_14); lean_dec(x_44); if (x_45 == 0) { @@ -20308,7 +20417,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -20352,7 +20461,7 @@ goto block_55; lean_object* x_26; uint8_t x_27; x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); lean_dec(x_26); if (x_27 == 0) { @@ -20361,7 +20470,7 @@ x_28 = l_Lean_Parser_Command_ctor___elambda__1___closed__2; x_29 = l_Lean_Parser_ParserState_mkNode(x_25, x_28, x_17); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); lean_dec(x_30); if (x_31 == 0) { @@ -20383,7 +20492,7 @@ lean_inc(x_1); x_33 = l_Lean_Parser_Command_declModifiers___elambda__2(x_1, x_25); x_34 = lean_ctor_get(x_33, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_14); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_14); lean_dec(x_34); if (x_35 == 0) { @@ -20392,7 +20501,7 @@ x_36 = l_Lean_Parser_Command_ctor___elambda__1___closed__2; x_37 = l_Lean_Parser_ParserState_mkNode(x_33, x_36, x_17); x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_14); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_14); lean_dec(x_38); if (x_39 == 0) { @@ -20414,7 +20523,7 @@ lean_inc(x_1); x_41 = l_Lean_Parser_ident___elambda__1(x_1, x_33); x_42 = lean_ctor_get(x_41, 4); lean_inc(x_42); -x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_14); +x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_14); lean_dec(x_42); if (x_43 == 0) { @@ -20423,7 +20532,7 @@ x_44 = l_Lean_Parser_Command_ctor___elambda__1___closed__2; x_45 = l_Lean_Parser_ParserState_mkNode(x_41, x_44, x_17); x_46 = lean_ctor_get(x_45, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_14); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_14); lean_dec(x_46); if (x_47 == 0) { @@ -20447,7 +20556,7 @@ x_50 = l_Lean_Parser_Command_ctor___elambda__1___closed__2; x_51 = l_Lean_Parser_ParserState_mkNode(x_49, x_50, x_17); x_52 = lean_ctor_get(x_51, 4); lean_inc(x_52); -x_53 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_52, x_14); +x_53 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_52, x_14); lean_dec(x_52); if (x_53 == 0) { @@ -20723,7 +20832,7 @@ static lean_object* _init_l_Lean_Parser_Command_derivingClasses___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_groupKind; +x_1 = l_Lean_Parser_Command_terminationHintMany___closed__18; x_2 = l_Lean_Parser_Command_derivingClasses___closed__10; x_3 = l_Lean_Parser_nodeInfo(x_1, x_2); return x_3; @@ -20733,7 +20842,7 @@ static lean_object* _init_l_Lean_Parser_Command_derivingClasses___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_groupKind; +x_1 = l_Lean_Parser_Command_terminationHintMany___closed__18; x_2 = l_Lean_Parser_Command_derivingClasses___closed__11; x_3 = lean_alloc_closure((void*)(l_Lean_Parser_nodeFn), 4, 2); lean_closure_set(x_3, 0, x_1); @@ -20941,7 +21050,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -20962,7 +21071,7 @@ x_21 = l_Lean_Parser_Command_optDeriving___elambda__1___closed__2; x_22 = l_Lean_Parser_ParserState_mkNode(x_20, x_21, x_19); x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); -x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_16); +x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_16); lean_dec(x_23); if (x_24 == 0) { @@ -21323,7 +21432,7 @@ x_16 = l_Lean_Parser_checkPrecFn(x_15, x_1, x_2); x_17 = lean_ctor_get(x_16, 4); lean_inc(x_17); x_18 = lean_box(0); -x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); +x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); lean_dec(x_17); if (x_19 == 0) { @@ -21371,7 +21480,7 @@ x_30 = l_Lean_Parser_Command_inductive___elambda__1___closed__2; x_31 = l_Lean_Parser_ParserState_mkNode(x_29, x_30, x_21); x_32 = lean_ctor_get(x_31, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_18); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_18); lean_dec(x_32); if (x_33 == 0) { @@ -21391,7 +21500,7 @@ return x_34; lean_object* x_37; uint8_t x_38; x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_18); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_18); lean_dec(x_37); if (x_38 == 0) { @@ -21407,7 +21516,7 @@ lean_inc(x_1); x_39 = l_Lean_Parser_Command_declId___elambda__1(x_1, x_36); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_18); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_18); lean_dec(x_40); if (x_41 == 0) { @@ -21423,7 +21532,7 @@ lean_inc(x_1); x_42 = l_Lean_Parser_Command_optDeclSig___elambda__1(x_1, x_39); x_43 = lean_ctor_get(x_42, 4); lean_inc(x_43); -x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_18); +x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_18); lean_dec(x_43); if (x_44 == 0) { @@ -21439,7 +21548,7 @@ lean_inc(x_1); x_45 = lean_apply_2(x_6, x_1, x_42); x_46 = lean_ctor_get(x_45, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_18); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_18); lean_dec(x_46); if (x_47 == 0) { @@ -21454,7 +21563,7 @@ lean_inc(x_1); x_48 = lean_apply_2(x_4, x_1, x_45); x_49 = lean_ctor_get(x_48, 4); lean_inc(x_49); -x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_18); +x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_18); lean_dec(x_49); if (x_50 == 0) { @@ -21701,7 +21810,7 @@ static lean_object* _init_l_Lean_Parser_Command_classInductive___elambda__1___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_groupKind; +x_1 = l_Lean_Parser_Command_terminationHintMany___closed__18; x_2 = l_Lean_Parser_Command_classInductive___elambda__1___closed__7; x_3 = lean_alloc_closure((void*)(l_Lean_Parser_nodeFn), 4, 2); lean_closure_set(x_3, 0, x_1); @@ -21801,7 +21910,7 @@ x_16 = l_Lean_Parser_checkPrecFn(x_15, x_1, x_2); x_17 = lean_ctor_get(x_16, 4); lean_inc(x_17); x_18 = lean_box(0); -x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); +x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); lean_dec(x_17); if (x_19 == 0) { @@ -21822,7 +21931,7 @@ lean_inc(x_1); x_23 = l_Lean_Parser_atomicFn(x_22, x_1, x_16); x_24 = lean_ctor_get(x_23, 4); lean_inc(x_24); -x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_18); +x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_18); lean_dec(x_24); if (x_25 == 0) { @@ -21838,7 +21947,7 @@ lean_inc(x_1); x_33 = l_Lean_Parser_Command_declId___elambda__1(x_1, x_23); x_34 = lean_ctor_get(x_33, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_18); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_18); lean_dec(x_34); if (x_35 == 0) { @@ -21854,7 +21963,7 @@ lean_inc(x_1); x_36 = l_Lean_Parser_Command_optDeclSig___elambda__1(x_1, x_33); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_18); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_18); lean_dec(x_37); if (x_38 == 0) { @@ -21870,7 +21979,7 @@ lean_inc(x_1); x_39 = lean_apply_2(x_6, x_1, x_36); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_18); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_18); lean_dec(x_40); if (x_41 == 0) { @@ -21885,7 +21994,7 @@ lean_inc(x_1); x_42 = lean_apply_2(x_4, x_1, x_39); x_43 = lean_ctor_get(x_42, 4); lean_inc(x_43); -x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_18); +x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_18); lean_dec(x_43); if (x_44 == 0) { @@ -21911,7 +22020,7 @@ x_27 = l_Lean_Parser_Command_classInductive___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_26, x_27, x_21); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_18); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_18); lean_dec(x_29); if (x_30 == 0) { @@ -21963,7 +22072,7 @@ static lean_object* _init_l_Lean_Parser_Command_classInductive___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_groupKind; +x_1 = l_Lean_Parser_Command_terminationHintMany___closed__18; x_2 = l_Lean_Parser_Command_classInductive___closed__2; x_3 = l_Lean_Parser_nodeInfo(x_1, x_2); return x_3; @@ -22294,7 +22403,7 @@ x_16 = l_Lean_Parser_checkPrecFn(x_15, x_1, x_2); x_17 = lean_ctor_get(x_16, 4); lean_inc(x_17); x_18 = lean_box(0); -x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); +x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); lean_dec(x_17); if (x_19 == 0) { @@ -22316,7 +22425,7 @@ lean_inc(x_1); x_23 = l_Lean_Parser_atomicFn(x_22, x_1, x_16); x_24 = lean_ctor_get(x_23, 4); lean_inc(x_24); -x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_18); +x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_18); lean_dec(x_24); if (x_25 == 0) { @@ -22333,7 +22442,7 @@ lean_inc(x_1); x_33 = lean_apply_2(x_6, x_1, x_23); x_34 = lean_ctor_get(x_33, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_18); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_18); lean_dec(x_34); if (x_35 == 0) { @@ -22349,7 +22458,7 @@ lean_inc(x_1); x_36 = l_Lean_Parser_Command_optDeclSig___elambda__1(x_1, x_33); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_18); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_18); lean_dec(x_37); if (x_38 == 0) { @@ -22365,7 +22474,7 @@ lean_inc(x_1); x_39 = lean_apply_2(x_4, x_1, x_36); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_18); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_18); lean_dec(x_40); if (x_41 == 0) { @@ -22411,7 +22520,7 @@ x_27 = l_Lean_Parser_Command_structExplicitBinder___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_26, x_27, x_21); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_18); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_18); lean_dec(x_29); if (x_30 == 0) { @@ -22772,7 +22881,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -22793,7 +22902,7 @@ lean_inc(x_1); x_21 = l_Lean_Parser_atomicFn(x_20, x_1, x_14); x_22 = lean_ctor_get(x_21, 4); lean_inc(x_22); -x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_22, x_16); +x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_22, x_16); lean_dec(x_22); if (x_23 == 0) { @@ -22809,7 +22918,7 @@ lean_inc(x_1); x_31 = lean_apply_2(x_4, x_1, x_21); x_32 = lean_ctor_get(x_31, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_16); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_16); lean_dec(x_32); if (x_33 == 0) { @@ -22824,7 +22933,7 @@ lean_inc(x_1); x_34 = l_Lean_Parser_Command_declSig___elambda__1(x_1, x_31); x_35 = lean_ctor_get(x_34, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_16); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_16); lean_dec(x_35); if (x_36 == 0) { @@ -22869,7 +22978,7 @@ x_25 = l_Lean_Parser_Command_structImplicitBinder___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_24, x_25, x_19); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); lean_dec(x_27); if (x_28 == 0) { @@ -23253,7 +23362,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -23274,7 +23383,7 @@ lean_inc(x_1); x_21 = l_Lean_Parser_atomicFn(x_20, x_1, x_14); x_22 = lean_ctor_get(x_21, 4); lean_inc(x_22); -x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_22, x_16); +x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_22, x_16); lean_dec(x_22); if (x_23 == 0) { @@ -23290,7 +23399,7 @@ lean_inc(x_1); x_31 = lean_apply_2(x_4, x_1, x_21); x_32 = lean_ctor_get(x_31, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_16); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_16); lean_dec(x_32); if (x_33 == 0) { @@ -23305,7 +23414,7 @@ lean_inc(x_1); x_34 = l_Lean_Parser_Command_declSig___elambda__1(x_1, x_31); x_35 = lean_ctor_get(x_34, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_16); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_16); lean_dec(x_35); if (x_36 == 0) { @@ -23350,7 +23459,7 @@ x_25 = l_Lean_Parser_Command_structInstBinder___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_24, x_25, x_19); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); lean_dec(x_27); if (x_28 == 0) { @@ -23658,7 +23767,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -23678,7 +23787,7 @@ lean_inc(x_1); x_21 = l_Lean_Parser_atomicFn(x_20, x_1, x_14); x_22 = lean_ctor_get(x_21, 4); lean_inc(x_22); -x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_22, x_16); +x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_22, x_16); lean_dec(x_22); if (x_23 == 0) { @@ -23688,7 +23797,7 @@ x_24 = l_Lean_Parser_Command_structSimpleBinder___elambda__1___closed__2; x_25 = l_Lean_Parser_ParserState_mkNode(x_21, x_24, x_19); x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_16); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_16); lean_dec(x_26); if (x_27 == 0) { @@ -23710,7 +23819,7 @@ lean_inc(x_1); x_29 = l_Lean_Parser_Command_optDeclSig___elambda__1(x_1, x_21); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_16); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_16); lean_dec(x_30); if (x_31 == 0) { @@ -23720,7 +23829,7 @@ x_32 = l_Lean_Parser_Command_structSimpleBinder___elambda__1___closed__2; x_33 = l_Lean_Parser_ParserState_mkNode(x_29, x_32, x_19); x_34 = lean_ctor_get(x_33, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_16); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_16); lean_dec(x_34); if (x_35 == 0) { @@ -23744,7 +23853,7 @@ x_38 = l_Lean_Parser_Command_structSimpleBinder___elambda__1___closed__2; x_39 = l_Lean_Parser_ParserState_mkNode(x_37, x_38, x_19); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_16); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_16); lean_dec(x_40); if (x_41 == 0) { @@ -24007,7 +24116,7 @@ static lean_object* _init_l_Lean_Parser_Command_structFields___elambda__1___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_terminationHintMany___closed__18; +x_1 = l_Lean_Parser_Command_terminationHintMany___closed__20; x_2 = l_Lean_Parser_Command_structFields___elambda__1___closed__10; x_3 = l_Lean_Parser_andthenInfo(x_1, x_2); return x_3; @@ -24051,7 +24160,7 @@ static lean_object* _init_l_Lean_Parser_Command_structFields___elambda__1___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_terminationHintMany___closed__18; +x_1 = l_Lean_Parser_Command_terminationHintMany___closed__20; x_2 = l_Lean_Parser_Command_structFields___elambda__1___closed__14; x_3 = l_Lean_Parser_andthenInfo(x_1, x_2); return x_3; @@ -24061,7 +24170,7 @@ static lean_object* _init_l_Lean_Parser_Command_structFields___elambda__1___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_terminationHintMany___closed__20; +x_1 = l_Lean_Parser_Command_terminationHintMany___closed__22; x_2 = l_Lean_Parser_Command_structFields___elambda__1___closed__15; x_3 = lean_alloc_closure((void*)(l_Lean_Parser_andthenFn), 4, 2); lean_closure_set(x_3, 0, x_1); @@ -24174,7 +24283,7 @@ x_18 = l_Lean_Parser_checkPrecFn(x_17, x_1, x_2); x_19 = lean_ctor_get(x_18, 4); lean_inc(x_19); x_20 = lean_box(0); -x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_20); +x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_20); lean_dec(x_19); if (x_21 == 0) { @@ -24215,7 +24324,7 @@ x_30 = l_Lean_Parser_Command_structFields___elambda__1___closed__2; x_31 = l_Lean_Parser_ParserState_mkNode(x_29, x_30, x_23); x_32 = lean_ctor_get(x_31, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_20); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_20); lean_dec(x_32); if (x_33 == 0) { @@ -24474,7 +24583,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -24495,7 +24604,7 @@ x_20 = l_Lean_Parser_Command_structCtor___elambda__1___closed__2; x_21 = l_Lean_Parser_ParserState_mkNode(x_19, x_20, x_17); x_22 = lean_ctor_get(x_21, 4); lean_inc(x_22); -x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_22, x_14); +x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_22, x_14); lean_dec(x_22); if (x_23 == 0) { @@ -24764,7 +24873,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -24797,7 +24906,7 @@ x_25 = l_Lean_Parser_Command_structureTk___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_20, x_25, x_17); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); lean_dec(x_27); if (x_28 == 0) { @@ -24821,7 +24930,7 @@ x_31 = l_Lean_Parser_Command_structureTk___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_17); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -25041,7 +25150,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -25074,7 +25183,7 @@ x_25 = l_Lean_Parser_Command_classTk___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_20, x_25, x_17); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); lean_dec(x_27); if (x_28 == 0) { @@ -25098,7 +25207,7 @@ x_31 = l_Lean_Parser_Command_classTk___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_17); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -25365,7 +25474,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -25397,7 +25506,7 @@ if (x_26 == 0) lean_object* x_27; uint8_t x_28; x_27 = lean_ctor_get(x_22, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); lean_dec(x_27); if (x_28 == 0) { @@ -25407,7 +25516,7 @@ x_29 = l_Lean_Parser_Command_extends___elambda__1___closed__2; x_30 = l_Lean_Parser_ParserState_mkNode(x_22, x_29, x_19); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); lean_dec(x_31); if (x_32 == 0) { @@ -25431,7 +25540,7 @@ x_35 = l_Lean_Parser_Command_extends___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_19); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_16); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_16); lean_dec(x_37); if (x_38 == 0) { @@ -25454,7 +25563,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_22); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); lean_dec(x_41); if (x_42 == 0) { @@ -25464,7 +25573,7 @@ x_43 = l_Lean_Parser_Command_extends___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_19); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); lean_dec(x_45); if (x_46 == 0) { @@ -25488,7 +25597,7 @@ x_49 = l_Lean_Parser_Command_extends___elambda__1___closed__2; x_50 = l_Lean_Parser_ParserState_mkNode(x_48, x_49, x_19); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); lean_dec(x_51); if (x_52 == 0) { @@ -26001,7 +26110,7 @@ x_20 = l_Lean_Parser_checkPrecFn(x_19, x_1, x_2); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); x_22 = lean_box(0); -x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_22); +x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_22); lean_dec(x_21); if (x_23 == 0) { @@ -26026,7 +26135,7 @@ lean_inc(x_1); x_29 = l_Lean_Parser_orelseFnCore(x_26, x_27, x_28, x_1, x_20); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_22); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_22); lean_dec(x_30); if (x_31 == 0) { @@ -26044,7 +26153,7 @@ lean_inc(x_1); x_39 = l_Lean_Parser_Command_declId___elambda__1(x_1, x_29); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_22); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_22); lean_dec(x_40); if (x_41 == 0) { @@ -26062,7 +26171,7 @@ lean_inc(x_1); x_42 = lean_apply_2(x_10, x_1, x_39); x_43 = lean_ctor_get(x_42, 4); lean_inc(x_43); -x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_22); +x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_22); lean_dec(x_43); if (x_44 == 0) { @@ -26079,7 +26188,7 @@ lean_inc(x_1); x_45 = lean_apply_2(x_8, x_1, x_42); x_46 = lean_ctor_get(x_45, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_22); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_22); lean_dec(x_46); if (x_47 == 0) { @@ -26095,7 +26204,7 @@ lean_inc(x_1); x_48 = lean_apply_2(x_6, x_1, x_45); x_49 = lean_ctor_get(x_48, 4); lean_inc(x_49); -x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_22); +x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_22); lean_dec(x_49); if (x_50 == 0) { @@ -26110,7 +26219,7 @@ lean_inc(x_1); x_51 = lean_apply_2(x_4, x_1, x_48); x_52 = lean_ctor_get(x_51, 4); lean_inc(x_52); -x_53 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_52, x_22); +x_53 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_52, x_22); lean_dec(x_52); if (x_53 == 0) { @@ -26137,7 +26246,7 @@ x_33 = l_Lean_Parser_Command_structure___elambda__1___closed__2; x_34 = l_Lean_Parser_ParserState_mkNode(x_32, x_33, x_25); x_35 = lean_ctor_get(x_34, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_22); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_22); lean_dec(x_35); if (x_36 == 0) { @@ -26538,7 +26647,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -26556,7 +26665,7 @@ lean_inc(x_1); x_18 = l_Lean_Parser_Command_declModifiers___elambda__2(x_1, x_12); x_19 = lean_ctor_get(x_18, 4); lean_inc(x_19); -x_20 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_14); +x_20 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_14); lean_dec(x_19); if (x_20 == 0) { @@ -26565,7 +26674,7 @@ x_21 = l_Lean_Parser_Command_declaration___elambda__1___closed__2; x_22 = l_Lean_Parser_ParserState_mkNode(x_18, x_21, x_17); x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); -x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_14); +x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_14); lean_dec(x_23); if (x_24 == 0) { @@ -26592,7 +26701,7 @@ x_30 = l_Lean_Parser_Command_declaration___elambda__1___closed__2; x_31 = l_Lean_Parser_ParserState_mkNode(x_29, x_30, x_17); x_32 = lean_ctor_get(x_31, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_14); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_14); lean_dec(x_32); if (x_33 == 0) { @@ -26837,7 +26946,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_declaration_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(97u); +x_1 = lean_unsigned_to_nat(98u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26849,7 +26958,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_declaration_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(98u); +x_1 = lean_unsigned_to_nat(99u); x_2 = lean_unsigned_to_nat(162u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26877,7 +26986,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_declaration_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(97u); +x_1 = lean_unsigned_to_nat(98u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26889,7 +26998,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_declaration_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(97u); +x_1 = lean_unsigned_to_nat(98u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28519,6 +28628,23 @@ return x_6; static lean_object* _init_l_Lean_Parser_Command_declVal_formatter___closed__1() { _start: { +lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_Parser_Command_declVal___elambda__1___closed__1; +x_2 = l_Lean_Parser_Command_declVal___elambda__1___closed__2; +x_3 = 1; +x_4 = lean_box(x_3); +x_5 = lean_box(x_3); +x_6 = lean_alloc_closure((void*)(l_Lean_Parser_mkAntiquot_formatter___boxed), 9, 4); +lean_closure_set(x_6, 0, x_1); +lean_closure_set(x_6, 1, x_2); +lean_closure_set(x_6, 2, x_4); +lean_closure_set(x_6, 3, x_5); +return x_6; +} +} +static lean_object* _init_l_Lean_Parser_Command_declVal_formatter___closed__2() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Parser_Command_declValEqns_formatter___closed__2; x_2 = l___regBuiltin_Lean_Parser_Command_whereStructInst_formatter___closed__2; @@ -28528,12 +28654,24 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } +static lean_object* _init_l_Lean_Parser_Command_declVal_formatter___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___regBuiltin_Lean_Parser_Command_declValSimple_formatter___closed__2; +x_2 = l_Lean_Parser_Command_declVal_formatter___closed__2; +x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_orelse_formatter), 7, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Parser_Command_declVal_formatter(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; -x_6 = l___regBuiltin_Lean_Parser_Command_declValSimple_formatter___closed__2; -x_7 = l_Lean_Parser_Command_declVal_formatter___closed__1; +x_6 = l_Lean_Parser_Command_declVal_formatter___closed__1; +x_7 = l_Lean_Parser_Command_declVal_formatter___closed__3; x_8 = l_Lean_PrettyPrinter_Formatter_orelse_formatter(x_6, x_7, x_1, x_2, x_3, x_4, x_5); return x_8; } @@ -34364,6 +34502,23 @@ return x_6; static lean_object* _init_l_Lean_Parser_Command_declVal_parenthesizer___closed__1() { _start: { +lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_Parser_Command_declVal___elambda__1___closed__1; +x_2 = l_Lean_Parser_Command_declVal___elambda__1___closed__2; +x_3 = 1; +x_4 = lean_box(x_3); +x_5 = lean_box(x_3); +x_6 = lean_alloc_closure((void*)(l_Lean_Parser_mkAntiquot_parenthesizer___boxed), 9, 4); +lean_closure_set(x_6, 0, x_1); +lean_closure_set(x_6, 1, x_2); +lean_closure_set(x_6, 2, x_4); +lean_closure_set(x_6, 3, x_5); +return x_6; +} +} +static lean_object* _init_l_Lean_Parser_Command_declVal_parenthesizer___closed__2() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Parser_Command_declValEqns_parenthesizer___closed__2; x_2 = l___regBuiltin_Lean_Parser_Command_whereStructInst_parenthesizer___closed__2; @@ -34373,13 +34528,25 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } +static lean_object* _init_l_Lean_Parser_Command_declVal_parenthesizer___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___regBuiltin_Lean_Parser_Command_declValSimple_parenthesizer___closed__2; +x_2 = l_Lean_Parser_Command_declVal_parenthesizer___closed__2; +x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_orelse_parenthesizer), 7, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Parser_Command_declVal_parenthesizer(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; -x_6 = l___regBuiltin_Lean_Parser_Command_declValSimple_parenthesizer___closed__2; -x_7 = l_Lean_Parser_Command_declVal_parenthesizer___closed__1; -x_8 = l_Lean_PrettyPrinter_Parenthesizer_orelse_parenthesizer(x_6, x_7, x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Parser_Command_declVal_parenthesizer___closed__1; +x_7 = l_Lean_Parser_Command_declVal_parenthesizer___closed__3; +x_8 = l_Lean_PrettyPrinter_Parenthesizer_withAntiquot_parenthesizer(x_6, x_7, x_1, x_2, x_3, x_4, x_5); return x_8; } } @@ -38889,7 +39056,7 @@ x_16 = l_Lean_Parser_checkPrecFn(x_15, x_1, x_2); x_17 = lean_ctor_get(x_16, 4); lean_inc(x_17); x_18 = lean_box(0); -x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); +x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); lean_dec(x_17); if (x_19 == 0) { @@ -38936,7 +39103,7 @@ x_30 = l_Lean_Parser_Command_deriving___elambda__1___closed__2; x_31 = l_Lean_Parser_ParserState_mkNode(x_29, x_30, x_21); x_32 = lean_ctor_get(x_31, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_18); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_18); lean_dec(x_32); if (x_33 == 0) { @@ -38956,7 +39123,7 @@ return x_34; lean_object* x_37; uint8_t x_38; x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_18); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_18); lean_dec(x_37); if (x_38 == 0) { @@ -38968,7 +39135,7 @@ x_39 = l_Lean_Parser_Command_deriving___elambda__1___closed__2; x_40 = l_Lean_Parser_ParserState_mkNode(x_36, x_39, x_21); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_18); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_18); lean_dec(x_41); if (x_42 == 0) { @@ -39013,7 +39180,7 @@ goto block_69; lean_object* x_51; uint8_t x_52; x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_18); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_18); lean_dec(x_51); if (x_52 == 0) { @@ -39030,7 +39197,7 @@ lean_inc(x_1); x_53 = lean_apply_2(x_6, x_1, x_50); x_54 = lean_ctor_get(x_53, 4); lean_inc(x_54); -x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_18); +x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_18); lean_dec(x_54); if (x_55 == 0) { @@ -39057,7 +39224,7 @@ if (x_61 == 0) lean_object* x_62; uint8_t x_63; x_62 = lean_ctor_get(x_58, 4); lean_inc(x_62); -x_63 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_62, x_18); +x_63 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_62, x_18); lean_dec(x_62); if (x_63 == 0) { @@ -39081,7 +39248,7 @@ lean_inc(x_1); x_65 = l_Lean_Parser_tokenAntiquotFn(x_1, x_58); x_66 = lean_ctor_get(x_65, 4); lean_inc(x_66); -x_67 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_66, x_18); +x_67 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_66, x_18); lean_dec(x_66); if (x_67 == 0) { @@ -39267,7 +39434,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_deriving_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(99u); +x_1 = lean_unsigned_to_nat(100u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -39279,7 +39446,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_deriving_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(99u); +x_1 = lean_unsigned_to_nat(100u); x_2 = lean_unsigned_to_nat(138u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -39307,7 +39474,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_deriving_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(99u); +x_1 = lean_unsigned_to_nat(100u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -39319,7 +39486,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_deriving_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(99u); +x_1 = lean_unsigned_to_nat(100u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -39823,7 +39990,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -39869,7 +40036,7 @@ x_28 = l_Lean_Parser_Command_noncomputableSection___elambda__1___closed__2; x_29 = l_Lean_Parser_ParserState_mkNode(x_27, x_28, x_19); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_16); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_16); lean_dec(x_30); if (x_31 == 0) { @@ -39889,7 +40056,7 @@ return x_32; lean_object* x_35; uint8_t x_36; x_35 = lean_ctor_get(x_34, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_16); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_16); lean_dec(x_35); if (x_36 == 0) { @@ -39916,7 +40083,7 @@ if (x_42 == 0) lean_object* x_43; uint8_t x_44; x_43 = lean_ctor_get(x_39, 4); lean_inc(x_43); -x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_16); +x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_16); lean_dec(x_43); if (x_44 == 0) { @@ -39940,7 +40107,7 @@ lean_inc(x_1); x_46 = l_Lean_Parser_tokenAntiquotFn(x_1, x_39); x_47 = lean_ctor_get(x_46, 4); lean_inc(x_47); -x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_16); +x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_16); lean_dec(x_47); if (x_48 == 0) { @@ -40091,7 +40258,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_noncomputableSectio _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(100u); +x_1 = lean_unsigned_to_nat(101u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40103,7 +40270,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_noncomputableSectio _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(100u); +x_1 = lean_unsigned_to_nat(101u); x_2 = lean_unsigned_to_nat(115u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40131,7 +40298,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_noncomputableSectio _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(100u); +x_1 = lean_unsigned_to_nat(101u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40143,7 +40310,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_noncomputableSectio _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(100u); +x_1 = lean_unsigned_to_nat(101u); x_2 = lean_unsigned_to_nat(48u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40517,7 +40684,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -40549,7 +40716,7 @@ if (x_26 == 0) lean_object* x_27; uint8_t x_28; x_27 = lean_ctor_get(x_22, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); lean_dec(x_27); if (x_28 == 0) { @@ -40559,7 +40726,7 @@ x_29 = l_Lean_Parser_Command_section___elambda__1___closed__2; x_30 = l_Lean_Parser_ParserState_mkNode(x_22, x_29, x_19); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); lean_dec(x_31); if (x_32 == 0) { @@ -40583,7 +40750,7 @@ x_35 = l_Lean_Parser_Command_section___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_19); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_16); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_16); lean_dec(x_37); if (x_38 == 0) { @@ -40606,7 +40773,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_22); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); lean_dec(x_41); if (x_42 == 0) { @@ -40616,7 +40783,7 @@ x_43 = l_Lean_Parser_Command_section___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_19); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); lean_dec(x_45); if (x_46 == 0) { @@ -40640,7 +40807,7 @@ x_49 = l_Lean_Parser_Command_section___elambda__1___closed__2; x_50 = l_Lean_Parser_ParserState_mkNode(x_48, x_49, x_19); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); lean_dec(x_51); if (x_52 == 0) { @@ -40757,7 +40924,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_section_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(101u); +x_1 = lean_unsigned_to_nat(102u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40769,7 +40936,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_section_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(101u); +x_1 = lean_unsigned_to_nat(102u); x_2 = lean_unsigned_to_nat(89u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40797,7 +40964,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_section_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(101u); +x_1 = lean_unsigned_to_nat(102u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40809,7 +40976,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_section_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(101u); +x_1 = lean_unsigned_to_nat(102u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -41151,7 +41318,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -41182,7 +41349,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_20, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -41191,7 +41358,7 @@ x_27 = l_Lean_Parser_Command_namespace___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_20, x_27, x_17); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -41215,7 +41382,7 @@ x_33 = l_Lean_Parser_Command_namespace___elambda__1___closed__2; x_34 = l_Lean_Parser_ParserState_mkNode(x_32, x_33, x_17); x_35 = lean_ctor_get(x_34, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_14); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_14); lean_dec(x_35); if (x_36 == 0) { @@ -41238,7 +41405,7 @@ lean_inc(x_1); x_38 = l_Lean_Parser_tokenAntiquotFn(x_1, x_20); x_39 = lean_ctor_get(x_38, 4); lean_inc(x_39); -x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); +x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); lean_dec(x_39); if (x_40 == 0) { @@ -41247,7 +41414,7 @@ x_41 = l_Lean_Parser_Command_namespace___elambda__1___closed__2; x_42 = l_Lean_Parser_ParserState_mkNode(x_38, x_41, x_17); x_43 = lean_ctor_get(x_42, 4); lean_inc(x_43); -x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_14); +x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_14); lean_dec(x_43); if (x_44 == 0) { @@ -41271,7 +41438,7 @@ x_47 = l_Lean_Parser_Command_namespace___elambda__1___closed__2; x_48 = l_Lean_Parser_ParserState_mkNode(x_46, x_47, x_17); x_49 = lean_ctor_get(x_48, 4); lean_inc(x_49); -x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_14); +x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_14); lean_dec(x_49); if (x_50 == 0) { @@ -41408,7 +41575,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_namespace_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(102u); +x_1 = lean_unsigned_to_nat(103u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -41420,7 +41587,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_namespace_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(102u); +x_1 = lean_unsigned_to_nat(103u); x_2 = lean_unsigned_to_nat(82u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -41448,7 +41615,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_namespace_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(102u); +x_1 = lean_unsigned_to_nat(103u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -41460,7 +41627,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_namespace_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(102u); +x_1 = lean_unsigned_to_nat(103u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -41851,7 +42018,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -41883,7 +42050,7 @@ if (x_26 == 0) lean_object* x_27; uint8_t x_28; x_27 = lean_ctor_get(x_22, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); lean_dec(x_27); if (x_28 == 0) { @@ -41893,7 +42060,7 @@ x_29 = l_Lean_Parser_Command_end___elambda__1___closed__2; x_30 = l_Lean_Parser_ParserState_mkNode(x_22, x_29, x_19); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); lean_dec(x_31); if (x_32 == 0) { @@ -41917,7 +42084,7 @@ x_35 = l_Lean_Parser_Command_end___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_19); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_16); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_16); lean_dec(x_37); if (x_38 == 0) { @@ -41940,7 +42107,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_22); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); lean_dec(x_41); if (x_42 == 0) { @@ -41950,7 +42117,7 @@ x_43 = l_Lean_Parser_Command_end___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_19); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); lean_dec(x_45); if (x_46 == 0) { @@ -41974,7 +42141,7 @@ x_49 = l_Lean_Parser_Command_end___elambda__1___closed__2; x_50 = l_Lean_Parser_ParserState_mkNode(x_48, x_49, x_19); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); lean_dec(x_51); if (x_52 == 0) { @@ -42112,7 +42279,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_end_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(103u); +x_1 = lean_unsigned_to_nat(104u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -42124,7 +42291,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_end_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(103u); +x_1 = lean_unsigned_to_nat(104u); x_2 = lean_unsigned_to_nat(85u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -42152,7 +42319,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_end_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(103u); +x_1 = lean_unsigned_to_nat(104u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -42164,7 +42331,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_end_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(103u); +x_1 = lean_unsigned_to_nat(104u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -42556,7 +42723,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -42588,7 +42755,7 @@ if (x_26 == 0) lean_object* x_27; uint8_t x_28; x_27 = lean_ctor_get(x_22, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); lean_dec(x_27); if (x_28 == 0) { @@ -42598,7 +42765,7 @@ x_29 = l_Lean_Parser_Command_variable___elambda__1___closed__2; x_30 = l_Lean_Parser_ParserState_mkNode(x_22, x_29, x_19); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); lean_dec(x_31); if (x_32 == 0) { @@ -42622,7 +42789,7 @@ x_35 = l_Lean_Parser_Command_variable___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_19); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_16); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_16); lean_dec(x_37); if (x_38 == 0) { @@ -42645,7 +42812,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_22); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); lean_dec(x_41); if (x_42 == 0) { @@ -42655,7 +42822,7 @@ x_43 = l_Lean_Parser_Command_variable___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_19); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); lean_dec(x_45); if (x_46 == 0) { @@ -42679,7 +42846,7 @@ x_49 = l_Lean_Parser_Command_variable___elambda__1___closed__2; x_50 = l_Lean_Parser_ParserState_mkNode(x_48, x_49, x_19); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); lean_dec(x_51); if (x_52 == 0) { @@ -42817,7 +42984,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_variable_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(104u); +x_1 = lean_unsigned_to_nat(105u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -42829,7 +42996,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_variable_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(104u); +x_1 = lean_unsigned_to_nat(105u); x_2 = lean_unsigned_to_nat(114u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -42857,7 +43024,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_variable_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(104u); +x_1 = lean_unsigned_to_nat(105u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -42869,7 +43036,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_variable_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(104u); +x_1 = lean_unsigned_to_nat(105u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -43280,7 +43447,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -43312,7 +43479,7 @@ if (x_26 == 0) lean_object* x_27; uint8_t x_28; x_27 = lean_ctor_get(x_22, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); lean_dec(x_27); if (x_28 == 0) { @@ -43322,7 +43489,7 @@ x_29 = l_Lean_Parser_Command_universe___elambda__1___closed__2; x_30 = l_Lean_Parser_ParserState_mkNode(x_22, x_29, x_19); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); lean_dec(x_31); if (x_32 == 0) { @@ -43346,7 +43513,7 @@ x_35 = l_Lean_Parser_Command_universe___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_19); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_16); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_16); lean_dec(x_37); if (x_38 == 0) { @@ -43369,7 +43536,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_22); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); lean_dec(x_41); if (x_42 == 0) { @@ -43379,7 +43546,7 @@ x_43 = l_Lean_Parser_Command_universe___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_19); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); lean_dec(x_45); if (x_46 == 0) { @@ -43403,7 +43570,7 @@ x_49 = l_Lean_Parser_Command_universe___elambda__1___closed__2; x_50 = l_Lean_Parser_ParserState_mkNode(x_48, x_49, x_19); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); lean_dec(x_51); if (x_52 == 0) { @@ -43541,7 +43708,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_universe_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(105u); +x_1 = lean_unsigned_to_nat(106u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -43553,7 +43720,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_universe_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(105u); +x_1 = lean_unsigned_to_nat(106u); x_2 = lean_unsigned_to_nat(87u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -43581,7 +43748,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_universe_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(105u); +x_1 = lean_unsigned_to_nat(106u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -43593,7 +43760,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_universe_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(105u); +x_1 = lean_unsigned_to_nat(106u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -43979,7 +44146,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -44010,7 +44177,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_20, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -44019,7 +44186,7 @@ x_27 = l_Lean_Parser_Command_check___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_20, x_27, x_17); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -44045,7 +44212,7 @@ x_35 = l_Lean_Parser_Command_check___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_17); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); lean_dec(x_37); if (x_38 == 0) { @@ -44068,7 +44235,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_20); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); lean_dec(x_41); if (x_42 == 0) { @@ -44077,7 +44244,7 @@ x_43 = l_Lean_Parser_Command_check___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_17); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); lean_dec(x_45); if (x_46 == 0) { @@ -44103,7 +44270,7 @@ x_51 = l_Lean_Parser_Command_check___elambda__1___closed__2; x_52 = l_Lean_Parser_ParserState_mkNode(x_50, x_51, x_17); x_53 = lean_ctor_get(x_52, 4); lean_inc(x_53); -x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); +x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); lean_dec(x_53); if (x_54 == 0) { @@ -44240,7 +44407,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_check_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(106u); +x_1 = lean_unsigned_to_nat(107u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -44252,7 +44419,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_check_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(106u); +x_1 = lean_unsigned_to_nat(107u); x_2 = lean_unsigned_to_nat(84u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -44280,7 +44447,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_check_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(106u); +x_1 = lean_unsigned_to_nat(107u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -44292,7 +44459,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_check_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(106u); +x_1 = lean_unsigned_to_nat(107u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -44678,7 +44845,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -44709,7 +44876,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_20, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -44718,7 +44885,7 @@ x_27 = l_Lean_Parser_Command_check__failure___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_20, x_27, x_17); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -44744,7 +44911,7 @@ x_35 = l_Lean_Parser_Command_check__failure___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_17); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); lean_dec(x_37); if (x_38 == 0) { @@ -44767,7 +44934,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_20); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); lean_dec(x_41); if (x_42 == 0) { @@ -44776,7 +44943,7 @@ x_43 = l_Lean_Parser_Command_check__failure___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_17); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); lean_dec(x_45); if (x_46 == 0) { @@ -44802,7 +44969,7 @@ x_51 = l_Lean_Parser_Command_check__failure___elambda__1___closed__2; x_52 = l_Lean_Parser_ParserState_mkNode(x_50, x_51, x_17); x_53 = lean_ctor_get(x_52, 4); lean_inc(x_53); -x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); +x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); lean_dec(x_53); if (x_54 == 0) { @@ -44939,7 +45106,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_check__failure_decl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(107u); +x_1 = lean_unsigned_to_nat(108u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -44951,7 +45118,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_check__failure_decl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(107u); +x_1 = lean_unsigned_to_nat(108u); x_2 = lean_unsigned_to_nat(92u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -44979,7 +45146,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_check__failure_decl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(107u); +x_1 = lean_unsigned_to_nat(108u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -44991,7 +45158,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_check__failure_decl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(107u); +x_1 = lean_unsigned_to_nat(108u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -45377,7 +45544,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -45408,7 +45575,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_20, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -45417,7 +45584,7 @@ x_27 = l_Lean_Parser_Command_reduce___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_20, x_27, x_17); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -45443,7 +45610,7 @@ x_35 = l_Lean_Parser_Command_reduce___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_17); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); lean_dec(x_37); if (x_38 == 0) { @@ -45466,7 +45633,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_20); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); lean_dec(x_41); if (x_42 == 0) { @@ -45475,7 +45642,7 @@ x_43 = l_Lean_Parser_Command_reduce___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_17); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); lean_dec(x_45); if (x_46 == 0) { @@ -45501,7 +45668,7 @@ x_51 = l_Lean_Parser_Command_reduce___elambda__1___closed__2; x_52 = l_Lean_Parser_ParserState_mkNode(x_50, x_51, x_17); x_53 = lean_ctor_get(x_52, 4); lean_inc(x_53); -x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); +x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); lean_dec(x_53); if (x_54 == 0) { @@ -45638,7 +45805,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_reduce_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(108u); +x_1 = lean_unsigned_to_nat(109u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -45650,7 +45817,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_reduce_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(108u); +x_1 = lean_unsigned_to_nat(109u); x_2 = lean_unsigned_to_nat(85u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -45678,7 +45845,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_reduce_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(108u); +x_1 = lean_unsigned_to_nat(109u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -45690,7 +45857,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_reduce_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(108u); +x_1 = lean_unsigned_to_nat(109u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -46076,7 +46243,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -46107,7 +46274,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_20, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -46116,7 +46283,7 @@ x_27 = l_Lean_Parser_Command_eval___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_20, x_27, x_17); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -46142,7 +46309,7 @@ x_35 = l_Lean_Parser_Command_eval___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_17); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); lean_dec(x_37); if (x_38 == 0) { @@ -46165,7 +46332,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_20); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); lean_dec(x_41); if (x_42 == 0) { @@ -46174,7 +46341,7 @@ x_43 = l_Lean_Parser_Command_eval___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_17); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); lean_dec(x_45); if (x_46 == 0) { @@ -46200,7 +46367,7 @@ x_51 = l_Lean_Parser_Command_eval___elambda__1___closed__2; x_52 = l_Lean_Parser_ParserState_mkNode(x_50, x_51, x_17); x_53 = lean_ctor_get(x_52, 4); lean_inc(x_53); -x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); +x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); lean_dec(x_53); if (x_54 == 0) { @@ -46337,7 +46504,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_eval_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(109u); +x_1 = lean_unsigned_to_nat(110u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -46349,7 +46516,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_eval_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(109u); +x_1 = lean_unsigned_to_nat(110u); x_2 = lean_unsigned_to_nat(83u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -46377,7 +46544,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_eval_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(109u); +x_1 = lean_unsigned_to_nat(110u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -46389,7 +46556,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_eval_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(109u); +x_1 = lean_unsigned_to_nat(110u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -46775,7 +46942,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -46806,7 +46973,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_20, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -46815,7 +46982,7 @@ x_27 = l_Lean_Parser_Command_synth___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_20, x_27, x_17); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -46841,7 +47008,7 @@ x_35 = l_Lean_Parser_Command_synth___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_17); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); lean_dec(x_37); if (x_38 == 0) { @@ -46864,7 +47031,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_20); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); lean_dec(x_41); if (x_42 == 0) { @@ -46873,7 +47040,7 @@ x_43 = l_Lean_Parser_Command_synth___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_17); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); lean_dec(x_45); if (x_46 == 0) { @@ -46899,7 +47066,7 @@ x_51 = l_Lean_Parser_Command_synth___elambda__1___closed__2; x_52 = l_Lean_Parser_ParserState_mkNode(x_50, x_51, x_17); x_53 = lean_ctor_get(x_52, 4); lean_inc(x_53); -x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); +x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); lean_dec(x_53); if (x_54 == 0) { @@ -47036,7 +47203,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_synth_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(110u); +x_1 = lean_unsigned_to_nat(111u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -47048,7 +47215,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_synth_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(110u); +x_1 = lean_unsigned_to_nat(111u); x_2 = lean_unsigned_to_nat(84u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -47076,7 +47243,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_synth_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(110u); +x_1 = lean_unsigned_to_nat(111u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -47088,7 +47255,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_synth_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(110u); +x_1 = lean_unsigned_to_nat(111u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -47462,7 +47629,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -47495,7 +47662,7 @@ x_25 = l_Lean_Parser_Command_exit___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_20, x_25, x_17); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); lean_dec(x_27); if (x_28 == 0) { @@ -47519,7 +47686,7 @@ x_31 = l_Lean_Parser_Command_exit___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_17); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -47643,7 +47810,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_exit_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(111u); +x_1 = lean_unsigned_to_nat(112u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -47655,7 +47822,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_exit_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(111u); +x_1 = lean_unsigned_to_nat(112u); x_2 = lean_unsigned_to_nat(68u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -47683,7 +47850,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_exit_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(111u); +x_1 = lean_unsigned_to_nat(112u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -47695,7 +47862,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_exit_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(111u); +x_1 = lean_unsigned_to_nat(112u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -48077,7 +48244,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -48108,7 +48275,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_20, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -48117,7 +48284,7 @@ x_27 = l_Lean_Parser_Command_print___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_20, x_27, x_17); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -48144,7 +48311,7 @@ x_36 = l_Lean_Parser_Command_print___elambda__1___closed__2; x_37 = l_Lean_Parser_ParserState_mkNode(x_35, x_36, x_17); x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_14); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_14); lean_dec(x_38); if (x_39 == 0) { @@ -48167,7 +48334,7 @@ lean_inc(x_1); x_41 = l_Lean_Parser_tokenAntiquotFn(x_1, x_20); x_42 = lean_ctor_get(x_41, 4); lean_inc(x_42); -x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_14); +x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_14); lean_dec(x_42); if (x_43 == 0) { @@ -48176,7 +48343,7 @@ x_44 = l_Lean_Parser_Command_print___elambda__1___closed__2; x_45 = l_Lean_Parser_ParserState_mkNode(x_41, x_44, x_17); x_46 = lean_ctor_get(x_45, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_14); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_14); lean_dec(x_46); if (x_47 == 0) { @@ -48203,7 +48370,7 @@ x_53 = l_Lean_Parser_Command_print___elambda__1___closed__2; x_54 = l_Lean_Parser_ParserState_mkNode(x_52, x_53, x_17); x_55 = lean_ctor_get(x_54, 4); lean_inc(x_55); -x_56 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_55, x_14); +x_56 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_55, x_14); lean_dec(x_55); if (x_56 == 0) { @@ -48352,7 +48519,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_print_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(112u); +x_1 = lean_unsigned_to_nat(113u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -48364,7 +48531,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_print_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(112u); +x_1 = lean_unsigned_to_nat(113u); x_2 = lean_unsigned_to_nat(92u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -48392,7 +48559,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_print_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(112u); +x_1 = lean_unsigned_to_nat(113u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -48404,7 +48571,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_print_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(112u); +x_1 = lean_unsigned_to_nat(113u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -48842,7 +49009,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -48887,7 +49054,7 @@ x_26 = l_Lean_Parser_Command_printAxioms___elambda__1___closed__2; x_27 = l_Lean_Parser_ParserState_mkNode(x_25, x_26, x_17); x_28 = lean_ctor_get(x_27, 4); lean_inc(x_28); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_14); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_14); lean_dec(x_28); if (x_29 == 0) { @@ -48907,7 +49074,7 @@ return x_30; lean_object* x_33; uint8_t x_34; x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -48933,7 +49100,7 @@ if (x_40 == 0) lean_object* x_41; uint8_t x_42; x_41 = lean_ctor_get(x_37, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); lean_dec(x_41); if (x_42 == 0) { @@ -48956,7 +49123,7 @@ lean_inc(x_1); x_44 = l_Lean_Parser_tokenAntiquotFn(x_1, x_37); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); lean_dec(x_45); if (x_46 == 0) { @@ -49106,7 +49273,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_printAxioms_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(113u); +x_1 = lean_unsigned_to_nat(114u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -49118,7 +49285,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_printAxioms_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(113u); +x_1 = lean_unsigned_to_nat(114u); x_2 = lean_unsigned_to_nat(110u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -49146,7 +49313,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_printAxioms_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(113u); +x_1 = lean_unsigned_to_nat(114u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -49158,7 +49325,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_printAxioms_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(113u); +x_1 = lean_unsigned_to_nat(114u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -49574,7 +49741,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -49605,7 +49772,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_20, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -49614,7 +49781,7 @@ x_27 = l_Lean_Parser_Command_resolve__name___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_20, x_27, x_17); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -49638,7 +49805,7 @@ x_33 = l_Lean_Parser_Command_resolve__name___elambda__1___closed__2; x_34 = l_Lean_Parser_ParserState_mkNode(x_32, x_33, x_17); x_35 = lean_ctor_get(x_34, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_14); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_14); lean_dec(x_35); if (x_36 == 0) { @@ -49661,7 +49828,7 @@ lean_inc(x_1); x_38 = l_Lean_Parser_tokenAntiquotFn(x_1, x_20); x_39 = lean_ctor_get(x_38, 4); lean_inc(x_39); -x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); +x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); lean_dec(x_39); if (x_40 == 0) { @@ -49670,7 +49837,7 @@ x_41 = l_Lean_Parser_Command_resolve__name___elambda__1___closed__2; x_42 = l_Lean_Parser_ParserState_mkNode(x_38, x_41, x_17); x_43 = lean_ctor_get(x_42, 4); lean_inc(x_43); -x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_14); +x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_14); lean_dec(x_43); if (x_44 == 0) { @@ -49694,7 +49861,7 @@ x_47 = l_Lean_Parser_Command_resolve__name___elambda__1___closed__2; x_48 = l_Lean_Parser_ParserState_mkNode(x_46, x_47, x_17); x_49 = lean_ctor_get(x_48, 4); lean_inc(x_49); -x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_14); +x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_14); lean_dec(x_49); if (x_50 == 0) { @@ -49831,7 +49998,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_resolve__name_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(114u); +x_1 = lean_unsigned_to_nat(115u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -49843,7 +50010,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_resolve__name_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(114u); +x_1 = lean_unsigned_to_nat(115u); x_2 = lean_unsigned_to_nat(86u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -49871,7 +50038,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_resolve__name_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(114u); +x_1 = lean_unsigned_to_nat(115u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -49883,7 +50050,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_resolve__name_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(114u); +x_1 = lean_unsigned_to_nat(115u); x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -50249,7 +50416,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -50282,7 +50449,7 @@ x_25 = l_Lean_Parser_Command_init__quot___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_20, x_25, x_17); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); lean_dec(x_27); if (x_28 == 0) { @@ -50306,7 +50473,7 @@ x_31 = l_Lean_Parser_Command_init__quot___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_17); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -50430,7 +50597,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_init__quot_declRang _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(115u); +x_1 = lean_unsigned_to_nat(116u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -50442,7 +50609,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_init__quot_declRang _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(115u); +x_1 = lean_unsigned_to_nat(116u); x_2 = lean_unsigned_to_nat(72u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -50470,7 +50637,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_init__quot_declRang _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(115u); +x_1 = lean_unsigned_to_nat(116u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -50482,7 +50649,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_init__quot_declRang _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(115u); +x_1 = lean_unsigned_to_nat(116u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -51047,7 +51214,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -51091,7 +51258,7 @@ goto block_47; lean_object* x_26; uint8_t x_27; x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); lean_dec(x_26); if (x_27 == 0) { @@ -51100,7 +51267,7 @@ x_28 = l_Lean_Parser_Command_set__option___elambda__1___closed__2; x_29 = l_Lean_Parser_ParserState_mkNode(x_25, x_28, x_17); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); lean_dec(x_30); if (x_31 == 0) { @@ -51122,7 +51289,7 @@ lean_inc(x_1); x_33 = l_Lean_Parser_ident___elambda__1(x_1, x_25); x_34 = lean_ctor_get(x_33, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_14); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_14); lean_dec(x_34); if (x_35 == 0) { @@ -51131,7 +51298,7 @@ x_36 = l_Lean_Parser_Command_set__option___elambda__1___closed__2; x_37 = l_Lean_Parser_ParserState_mkNode(x_33, x_36, x_17); x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_14); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_14); lean_dec(x_38); if (x_39 == 0) { @@ -51155,7 +51322,7 @@ x_42 = l_Lean_Parser_Command_set__option___elambda__1___closed__2; x_43 = l_Lean_Parser_ParserState_mkNode(x_41, x_42, x_17); x_44 = lean_ctor_get(x_43, 4); lean_inc(x_44); -x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_14); +x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_14); lean_dec(x_44); if (x_45 == 0) { @@ -51315,7 +51482,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_set__option_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(117u); +x_1 = lean_unsigned_to_nat(118u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -51327,7 +51494,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_set__option_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(117u); +x_1 = lean_unsigned_to_nat(118u); x_2 = lean_unsigned_to_nat(109u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -51355,7 +51522,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_set__option_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(117u); +x_1 = lean_unsigned_to_nat(118u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -51367,7 +51534,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_set__option_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(117u); +x_1 = lean_unsigned_to_nat(118u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -51961,7 +52128,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -51992,7 +52159,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_20, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -52001,7 +52168,7 @@ x_27 = l_Lean_Parser_Command_eraseAttr___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_20, x_27, x_17); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -52025,7 +52192,7 @@ x_33 = l_Lean_Parser_Command_eraseAttr___elambda__1___closed__2; x_34 = l_Lean_Parser_ParserState_mkNode(x_32, x_33, x_17); x_35 = lean_ctor_get(x_34, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_14); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_14); lean_dec(x_35); if (x_36 == 0) { @@ -52048,7 +52215,7 @@ lean_inc(x_1); x_38 = l_Lean_Parser_tokenAntiquotFn(x_1, x_20); x_39 = lean_ctor_get(x_38, 4); lean_inc(x_39); -x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); +x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); lean_dec(x_39); if (x_40 == 0) { @@ -52057,7 +52224,7 @@ x_41 = l_Lean_Parser_Command_eraseAttr___elambda__1___closed__2; x_42 = l_Lean_Parser_ParserState_mkNode(x_38, x_41, x_17); x_43 = lean_ctor_get(x_42, 4); lean_inc(x_43); -x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_14); +x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_14); lean_dec(x_43); if (x_44 == 0) { @@ -52081,7 +52248,7 @@ x_47 = l_Lean_Parser_Command_eraseAttr___elambda__1___closed__2; x_48 = l_Lean_Parser_ParserState_mkNode(x_46, x_47, x_17); x_49 = lean_ctor_get(x_48, 4); lean_inc(x_49); -x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_14); +x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_14); lean_dec(x_49); if (x_50 == 0) { @@ -52524,7 +52691,7 @@ x_16 = l_Lean_Parser_checkPrecFn(x_15, x_1, x_2); x_17 = lean_ctor_get(x_16, 4); lean_inc(x_17); x_18 = lean_box(0); -x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); +x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); lean_dec(x_17); if (x_19 == 0) { @@ -52571,7 +52738,7 @@ x_30 = l_Lean_Parser_Command_attribute___elambda__1___closed__2; x_31 = l_Lean_Parser_ParserState_mkNode(x_29, x_30, x_21); x_32 = lean_ctor_get(x_31, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_18); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_18); lean_dec(x_32); if (x_33 == 0) { @@ -52591,7 +52758,7 @@ return x_34; lean_object* x_37; uint8_t x_38; x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_18); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_18); lean_dec(x_37); if (x_38 == 0) { @@ -52603,7 +52770,7 @@ x_39 = l_Lean_Parser_Command_attribute___elambda__1___closed__2; x_40 = l_Lean_Parser_ParserState_mkNode(x_36, x_39, x_21); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_18); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_18); lean_dec(x_41); if (x_42 == 0) { @@ -52648,7 +52815,7 @@ goto block_69; lean_object* x_51; uint8_t x_52; x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_18); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_18); lean_dec(x_51); if (x_52 == 0) { @@ -52665,7 +52832,7 @@ lean_inc(x_1); x_53 = lean_apply_2(x_6, x_1, x_50); x_54 = lean_ctor_get(x_53, 4); lean_inc(x_54); -x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_18); +x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_18); lean_dec(x_54); if (x_55 == 0) { @@ -52692,7 +52859,7 @@ if (x_61 == 0) lean_object* x_62; uint8_t x_63; x_62 = lean_ctor_get(x_58, 4); lean_inc(x_62); -x_63 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_62, x_18); +x_63 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_62, x_18); lean_dec(x_62); if (x_63 == 0) { @@ -52716,7 +52883,7 @@ lean_inc(x_1); x_65 = l_Lean_Parser_tokenAntiquotFn(x_1, x_58); x_66 = lean_ctor_get(x_65, 4); lean_inc(x_66); -x_67 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_66, x_18); +x_67 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_66, x_18); lean_dec(x_66); if (x_67 == 0) { @@ -52902,7 +53069,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_attribute_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(119u); +x_1 = lean_unsigned_to_nat(120u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -52914,7 +53081,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_attribute_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(119u); +x_1 = lean_unsigned_to_nat(120u); x_2 = lean_unsigned_to_nat(152u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -52942,7 +53109,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_attribute_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(119u); +x_1 = lean_unsigned_to_nat(120u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -52954,7 +53121,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_attribute_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(119u); +x_1 = lean_unsigned_to_nat(120u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -53798,7 +53965,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -53842,7 +54009,7 @@ goto block_84; lean_object* x_28; uint8_t x_29; x_28 = lean_ctor_get(x_27, 4); lean_inc(x_28); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_16); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_16); lean_dec(x_28); if (x_29 == 0) { @@ -53853,7 +54020,7 @@ x_30 = l_Lean_Parser_Command_export___elambda__1___closed__2; x_31 = l_Lean_Parser_ParserState_mkNode(x_27, x_30, x_19); x_32 = lean_ctor_get(x_31, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_16); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_16); lean_dec(x_32); if (x_33 == 0) { @@ -53875,7 +54042,7 @@ lean_inc(x_1); x_35 = l_Lean_Parser_ident___elambda__1(x_1, x_27); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_16); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_16); lean_dec(x_36); if (x_37 == 0) { @@ -53886,7 +54053,7 @@ x_38 = l_Lean_Parser_Command_export___elambda__1___closed__2; x_39 = l_Lean_Parser_ParserState_mkNode(x_35, x_38, x_19); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_16); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_16); lean_dec(x_40); if (x_41 == 0) { @@ -53931,7 +54098,7 @@ goto block_82; lean_object* x_50; uint8_t x_51; x_50 = lean_ctor_get(x_49, 4); lean_inc(x_50); -x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_16); +x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_16); lean_dec(x_50); if (x_51 == 0) { @@ -53942,7 +54109,7 @@ x_52 = l_Lean_Parser_Command_export___elambda__1___closed__2; x_53 = l_Lean_Parser_ParserState_mkNode(x_49, x_52, x_19); x_54 = lean_ctor_get(x_53, 4); lean_inc(x_54); -x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_16); +x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_16); lean_dec(x_54); if (x_55 == 0) { @@ -53964,7 +54131,7 @@ lean_inc(x_1); x_57 = lean_apply_2(x_4, x_1, x_49); x_58 = lean_ctor_get(x_57, 4); lean_inc(x_58); -x_59 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_58, x_16); +x_59 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_58, x_16); lean_dec(x_58); if (x_59 == 0) { @@ -53974,7 +54141,7 @@ x_60 = l_Lean_Parser_Command_export___elambda__1___closed__2; x_61 = l_Lean_Parser_ParserState_mkNode(x_57, x_60, x_19); x_62 = lean_ctor_get(x_61, 4); lean_inc(x_62); -x_63 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_62, x_16); +x_63 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_62, x_16); lean_dec(x_62); if (x_63 == 0) { @@ -54009,7 +54176,7 @@ x_71 = l_Lean_Parser_Command_export___elambda__1___closed__2; x_72 = l_Lean_Parser_ParserState_mkNode(x_67, x_71, x_19); x_73 = lean_ctor_get(x_72, 4); lean_inc(x_73); -x_74 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_73, x_16); +x_74 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_73, x_16); lean_dec(x_73); if (x_74 == 0) { @@ -54033,7 +54200,7 @@ x_77 = l_Lean_Parser_Command_export___elambda__1___closed__2; x_78 = l_Lean_Parser_ParserState_mkNode(x_76, x_77, x_19); x_79 = lean_ctor_get(x_78, 4); lean_inc(x_79); -x_80 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_79, x_16); +x_80 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_79, x_16); lean_dec(x_79); if (x_80 == 0) { @@ -54217,7 +54384,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_export_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(120u); +x_1 = lean_unsigned_to_nat(121u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -54229,7 +54396,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_export_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(120u); +x_1 = lean_unsigned_to_nat(121u); x_2 = lean_unsigned_to_nat(109u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -54257,7 +54424,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_export_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(120u); +x_1 = lean_unsigned_to_nat(121u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -54269,7 +54436,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_export_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(120u); +x_1 = lean_unsigned_to_nat(121u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -54699,7 +54866,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_ident; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l_Lean_Parser_Command_terminationHintMany___closed__18; +x_3 = l_Lean_Parser_Command_terminationHintMany___closed__20; x_4 = l_Lean_Parser_andthenInfo(x_3, x_2); return x_4; } @@ -54818,7 +54985,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -54838,7 +55005,7 @@ lean_inc(x_1); x_21 = l_Lean_Parser_atomicFn(x_20, x_1, x_14); x_22 = lean_ctor_get(x_21, 4); lean_inc(x_22); -x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_22, x_16); +x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_22, x_16); lean_dec(x_22); if (x_23 == 0) { @@ -54848,7 +55015,7 @@ x_24 = l_Lean_Parser_Command_openHiding___elambda__1___closed__2; x_25 = l_Lean_Parser_ParserState_mkNode(x_21, x_24, x_19); x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_16); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_16); lean_dec(x_26); if (x_27 == 0) { @@ -54872,7 +55039,7 @@ x_30 = l_Lean_Parser_Command_openHiding___elambda__1___closed__2; x_31 = l_Lean_Parser_ParserState_mkNode(x_29, x_30, x_19); x_32 = lean_ctor_get(x_31, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_16); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_16); lean_dec(x_32); if (x_33 == 0) { @@ -55220,7 +55387,7 @@ x_25 = l_Lean_Parser_checkPrecFn(x_24, x_1, x_2); x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); x_27 = lean_box(0); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_27); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_27); lean_dec(x_26); if (x_28 == 0) { @@ -55240,7 +55407,7 @@ lean_inc(x_1); x_31 = l_Lean_Parser_ident___elambda__1(x_1, x_25); x_32 = lean_ctor_get(x_31, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_27); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_27); lean_dec(x_32); if (x_33 == 0) { @@ -55250,7 +55417,7 @@ lean_dec(x_18); x_40 = l_Lean_Parser_ParserState_mkNode(x_31, x_10, x_30); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_27); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_27); lean_dec(x_41); if (x_42 == 0) { @@ -55296,7 +55463,7 @@ goto block_69; lean_object* x_51; uint8_t x_52; x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_27); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_27); lean_dec(x_51); if (x_52 == 0) { @@ -55350,7 +55517,7 @@ x_63 = l_Lean_Parser_Command_openHiding___elambda__1___closed__9; x_64 = l_Lean_Parser_ParserState_mkError(x_50, x_63); x_65 = lean_ctor_get(x_64, 4); lean_inc(x_65); -x_66 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_65, x_27); +x_66 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_65, x_27); lean_dec(x_65); if (x_66 == 0) { @@ -55384,7 +55551,7 @@ lean_object* x_35; lean_object* x_36; uint8_t x_37; x_35 = l_Lean_Parser_ParserState_mkNode(x_34, x_10, x_30); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_27); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_27); lean_dec(x_36); if (x_37 == 0) { @@ -55429,7 +55596,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_ident; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l_Lean_Parser_Command_terminationHintMany___closed__18; +x_3 = l_Lean_Parser_Command_terminationHintMany___closed__20; x_4 = l_Lean_Parser_andthenInfo(x_3, x_2); return x_4; } @@ -55708,7 +55875,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -55728,7 +55895,7 @@ lean_inc(x_1); x_21 = l_Lean_Parser_atomicFn(x_20, x_1, x_14); x_22 = lean_ctor_get(x_21, 4); lean_inc(x_22); -x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_22, x_16); +x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_22, x_16); lean_dec(x_22); if (x_23 == 0) { @@ -55738,7 +55905,7 @@ x_24 = l_Lean_Parser_Command_openRenaming___elambda__1___closed__2; x_25 = l_Lean_Parser_ParserState_mkNode(x_21, x_24, x_19); x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_16); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_16); lean_dec(x_26); if (x_27 == 0) { @@ -55762,7 +55929,7 @@ x_30 = l_Lean_Parser_Command_openRenaming___elambda__1___closed__2; x_31 = l_Lean_Parser_ParserState_mkNode(x_29, x_30, x_19); x_32 = lean_ctor_get(x_31, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_16); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_16); lean_dec(x_32); if (x_33 == 0) { @@ -56023,7 +56190,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -56044,7 +56211,7 @@ lean_inc(x_1); x_21 = l_Lean_Parser_atomicFn(x_20, x_1, x_14); x_22 = lean_ctor_get(x_21, 4); lean_inc(x_22); -x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_22, x_16); +x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_22, x_16); lean_dec(x_22); if (x_23 == 0) { @@ -56055,7 +56222,7 @@ x_24 = l_Lean_Parser_Command_openOnly___elambda__1___closed__2; x_25 = l_Lean_Parser_ParserState_mkNode(x_21, x_24, x_19); x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_16); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_16); lean_dec(x_26); if (x_27 == 0) { @@ -56077,7 +56244,7 @@ lean_inc(x_1); x_29 = lean_apply_2(x_4, x_1, x_21); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_16); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_16); lean_dec(x_30); if (x_31 == 0) { @@ -56087,7 +56254,7 @@ x_32 = l_Lean_Parser_Command_openOnly___elambda__1___closed__2; x_33 = l_Lean_Parser_ParserState_mkNode(x_29, x_32, x_19); x_34 = lean_ctor_get(x_33, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_16); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_16); lean_dec(x_34); if (x_35 == 0) { @@ -56123,7 +56290,7 @@ x_44 = l_Lean_Parser_Command_openOnly___elambda__1___closed__2; x_45 = l_Lean_Parser_ParserState_mkNode(x_39, x_44, x_19); x_46 = lean_ctor_get(x_45, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_16); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_16); lean_dec(x_46); if (x_47 == 0) { @@ -56147,7 +56314,7 @@ x_50 = l_Lean_Parser_Command_openOnly___elambda__1___closed__2; x_51 = l_Lean_Parser_ParserState_mkNode(x_49, x_50, x_19); x_52 = lean_ctor_get(x_51, 4); lean_inc(x_52); -x_53 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_52, x_16); +x_53 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_52, x_16); lean_dec(x_52); if (x_53 == 0) { @@ -56390,7 +56557,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -56411,7 +56578,7 @@ x_21 = l_Lean_Parser_Command_openSimple___elambda__1___closed__2; x_22 = l_Lean_Parser_ParserState_mkNode(x_20, x_21, x_19); x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); -x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_16); +x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_16); lean_dec(x_23); if (x_24 == 0) { @@ -56667,7 +56834,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -56699,7 +56866,7 @@ if (x_26 == 0) lean_object* x_27; uint8_t x_28; x_27 = lean_ctor_get(x_22, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); lean_dec(x_27); if (x_28 == 0) { @@ -56709,7 +56876,7 @@ x_29 = l_Lean_Parser_Command_openScoped___elambda__1___closed__2; x_30 = l_Lean_Parser_ParserState_mkNode(x_22, x_29, x_19); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); lean_dec(x_31); if (x_32 == 0) { @@ -56733,7 +56900,7 @@ x_35 = l_Lean_Parser_Command_openScoped___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_19); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_16); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_16); lean_dec(x_37); if (x_38 == 0) { @@ -56756,7 +56923,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_22); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); lean_dec(x_41); if (x_42 == 0) { @@ -56766,7 +56933,7 @@ x_43 = l_Lean_Parser_Command_openScoped___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_19); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); lean_dec(x_45); if (x_46 == 0) { @@ -56790,7 +56957,7 @@ x_49 = l_Lean_Parser_Command_openScoped___elambda__1___closed__2; x_50 = l_Lean_Parser_ParserState_mkNode(x_48, x_49, x_19); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); lean_dec(x_51); if (x_52 == 0) { @@ -57074,7 +57241,7 @@ if (x_17 == 0) lean_object* x_19; uint8_t x_20; x_19 = lean_ctor_get(x_12, 4); lean_inc(x_19); -x_20 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_18); +x_20 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_18); lean_dec(x_19); if (x_20 == 0) { @@ -57095,7 +57262,7 @@ lean_inc(x_2); x_22 = l_Lean_Parser_tokenAntiquotFn(x_2, x_12); x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); -x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_18); +x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_18); lean_dec(x_23); if (x_24 == 0) { @@ -57162,7 +57329,7 @@ if (x_44 == 0) lean_object* x_46; uint8_t x_47; x_46 = lean_ctor_get(x_39, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_45); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_45); lean_dec(x_46); if (x_47 == 0) { @@ -57183,7 +57350,7 @@ lean_inc(x_35); x_49 = l_Lean_Parser_tokenAntiquotFn(x_35, x_39); x_50 = lean_ctor_get(x_49, 4); lean_inc(x_50); -x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_45); +x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_45); lean_dec(x_50); if (x_51 == 0) { @@ -57320,7 +57487,7 @@ x_26 = l_Lean_Parser_checkPrecFn(x_25, x_1, x_2); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); x_28 = lean_box(0); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_28); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_28); lean_dec(x_27); if (x_29 == 0) { @@ -57369,7 +57536,7 @@ if (x_40 == 0) lean_object* x_41; uint8_t x_42; x_41 = lean_ctor_get(x_36, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_28); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_28); lean_dec(x_41); if (x_42 == 0) { @@ -57378,7 +57545,7 @@ lean_dec(x_34); x_43 = l_Lean_Parser_ParserState_mkNode(x_36, x_5, x_31); x_44 = lean_ctor_get(x_43, 4); lean_inc(x_44); -x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_28); +x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_28); lean_dec(x_44); if (x_45 == 0) { @@ -57400,7 +57567,7 @@ x_47 = l_Lean_Parser_Command_openDecl___elambda__1(x_34, x_36); x_48 = l_Lean_Parser_ParserState_mkNode(x_47, x_5, x_31); x_49 = lean_ctor_get(x_48, 4); lean_inc(x_49); -x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_28); +x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_28); lean_dec(x_49); if (x_50 == 0) { @@ -57423,7 +57590,7 @@ lean_inc(x_34); x_52 = l_Lean_Parser_tokenAntiquotFn(x_34, x_36); x_53 = lean_ctor_get(x_52, 4); lean_inc(x_53); -x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_28); +x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_28); lean_dec(x_53); if (x_54 == 0) { @@ -57432,7 +57599,7 @@ lean_dec(x_34); x_55 = l_Lean_Parser_ParserState_mkNode(x_52, x_5, x_31); x_56 = lean_ctor_get(x_55, 4); lean_inc(x_56); -x_57 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_56, x_28); +x_57 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_56, x_28); lean_dec(x_56); if (x_57 == 0) { @@ -57454,7 +57621,7 @@ x_59 = l_Lean_Parser_Command_openDecl___elambda__1(x_34, x_52); x_60 = l_Lean_Parser_ParserState_mkNode(x_59, x_5, x_31); x_61 = lean_ctor_get(x_60, 4); lean_inc(x_61); -x_62 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_61, x_28); +x_62 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_61, x_28); lean_dec(x_61); if (x_62 == 0) { @@ -57605,7 +57772,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_open_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(128u); +x_1 = lean_unsigned_to_nat(129u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -57617,7 +57784,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_open_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(128u); +x_1 = lean_unsigned_to_nat(129u); x_2 = lean_unsigned_to_nat(90u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -57645,7 +57812,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_open_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(128u); +x_1 = lean_unsigned_to_nat(129u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -57657,7 +57824,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_open_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(128u); +x_1 = lean_unsigned_to_nat(129u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -59432,7 +59599,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_quot___closed__1; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l_Lean_Parser_Command_terminationHintMany___closed__18; +x_3 = l_Lean_Parser_Command_terminationHintMany___closed__20; x_4 = l_Lean_Parser_andthenInfo(x_3, x_2); return x_4; } @@ -59648,7 +59815,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -59694,7 +59861,7 @@ x_28 = l_Lean_Parser_Command_mutual___elambda__1___closed__2; x_29 = l_Lean_Parser_ParserState_mkNode(x_27, x_28, x_19); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_16); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_16); lean_dec(x_30); if (x_31 == 0) { @@ -59714,7 +59881,7 @@ return x_32; lean_object* x_35; uint8_t x_36; x_35 = lean_ctor_get(x_34, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_16); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_16); lean_dec(x_35); if (x_36 == 0) { @@ -59730,7 +59897,7 @@ lean_inc(x_1); x_37 = lean_apply_2(x_4, x_1, x_34); x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_16); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_16); lean_dec(x_38); if (x_39 == 0) { @@ -59756,7 +59923,7 @@ if (x_45 == 0) lean_object* x_46; uint8_t x_47; x_46 = lean_ctor_get(x_42, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_16); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_16); lean_dec(x_46); if (x_47 == 0) { @@ -59779,7 +59946,7 @@ lean_inc(x_1); x_49 = l_Lean_Parser_tokenAntiquotFn(x_1, x_42); x_50 = lean_ctor_get(x_49, 4); lean_inc(x_50); -x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_16); +x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_16); lean_dec(x_50); if (x_51 == 0) { @@ -59961,7 +60128,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_mutual_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(130u); +x_1 = lean_unsigned_to_nat(131u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -59973,7 +60140,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_mutual_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(130u); +x_1 = lean_unsigned_to_nat(131u); x_2 = lean_unsigned_to_nat(169u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -60001,7 +60168,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_mutual_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(130u); +x_1 = lean_unsigned_to_nat(131u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -60013,7 +60180,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_mutual_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(130u); +x_1 = lean_unsigned_to_nat(131u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -60730,7 +60897,7 @@ x_16 = l_Lean_Parser_checkPrecFn(x_15, x_1, x_2); x_17 = lean_ctor_get(x_16, 4); lean_inc(x_17); x_18 = lean_box(0); -x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); +x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); lean_dec(x_17); if (x_19 == 0) { @@ -60751,7 +60918,7 @@ lean_inc(x_1); x_22 = lean_apply_2(x_6, x_1, x_16); x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); -x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_18); +x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_18); lean_dec(x_23); if (x_24 == 0) { @@ -60779,7 +60946,7 @@ if (x_38 == 0) lean_object* x_39; uint8_t x_40; x_39 = lean_ctor_get(x_34, 4); lean_inc(x_39); -x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_18); +x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_18); lean_dec(x_39); if (x_40 == 0) { @@ -60794,7 +60961,7 @@ lean_inc(x_1); x_41 = lean_apply_2(x_4, x_1, x_34); x_42 = lean_ctor_get(x_41, 4); lean_inc(x_42); -x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_18); +x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_18); lean_dec(x_42); if (x_43 == 0) { @@ -60818,7 +60985,7 @@ lean_inc(x_1); x_45 = l_Lean_Parser_tokenAntiquotFn(x_1, x_34); x_46 = lean_ctor_get(x_45, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_18); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_18); lean_dec(x_46); if (x_47 == 0) { @@ -60833,7 +61000,7 @@ lean_inc(x_1); x_48 = lean_apply_2(x_4, x_1, x_45); x_49 = lean_ctor_get(x_48, 4); lean_inc(x_49); -x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_18); +x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_18); lean_dec(x_49); if (x_50 == 0) { @@ -60858,7 +61025,7 @@ x_26 = l_Lean_Parser_Command_initialize___elambda__1___closed__2; x_27 = l_Lean_Parser_ParserState_mkNode(x_25, x_26, x_21); x_28 = lean_ctor_get(x_27, 4); lean_inc(x_28); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_18); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_18); lean_dec(x_28); if (x_29 == 0) { @@ -61020,7 +61187,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_initialize_declRang _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(131u); +x_1 = lean_unsigned_to_nat(132u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -61032,7 +61199,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_initialize_declRang _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(131u); +x_1 = lean_unsigned_to_nat(132u); x_2 = lean_unsigned_to_nat(173u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -61060,7 +61227,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_initialize_declRang _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(131u); +x_1 = lean_unsigned_to_nat(132u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -61072,7 +61239,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_initialize_declRang _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(131u); +x_1 = lean_unsigned_to_nat(132u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -61646,7 +61813,7 @@ x_16 = l_Lean_Parser_checkPrecFn(x_15, x_1, x_2); x_17 = lean_ctor_get(x_16, 4); lean_inc(x_17); x_18 = lean_box(0); -x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); +x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); lean_dec(x_17); if (x_19 == 0) { @@ -61667,7 +61834,7 @@ lean_inc(x_1); x_22 = lean_apply_2(x_6, x_1, x_16); x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); -x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_18); +x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_18); lean_dec(x_23); if (x_24 == 0) { @@ -61695,7 +61862,7 @@ if (x_38 == 0) lean_object* x_39; uint8_t x_40; x_39 = lean_ctor_get(x_34, 4); lean_inc(x_39); -x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_18); +x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_18); lean_dec(x_39); if (x_40 == 0) { @@ -61710,7 +61877,7 @@ lean_inc(x_1); x_41 = lean_apply_2(x_4, x_1, x_34); x_42 = lean_ctor_get(x_41, 4); lean_inc(x_42); -x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_18); +x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_18); lean_dec(x_42); if (x_43 == 0) { @@ -61734,7 +61901,7 @@ lean_inc(x_1); x_45 = l_Lean_Parser_tokenAntiquotFn(x_1, x_34); x_46 = lean_ctor_get(x_45, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_18); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_18); lean_dec(x_46); if (x_47 == 0) { @@ -61749,7 +61916,7 @@ lean_inc(x_1); x_48 = lean_apply_2(x_4, x_1, x_45); x_49 = lean_ctor_get(x_48, 4); lean_inc(x_49); -x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_18); +x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_18); lean_dec(x_49); if (x_50 == 0) { @@ -61774,7 +61941,7 @@ x_26 = l_Lean_Parser_Command_builtin__initialize___elambda__1___closed__2; x_27 = l_Lean_Parser_ParserState_mkNode(x_25, x_26, x_21); x_28 = lean_ctor_get(x_27, 4); lean_inc(x_28); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_18); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_18); lean_dec(x_28); if (x_29 == 0) { @@ -61922,7 +62089,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_builtin__initialize _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(132u); +x_1 = lean_unsigned_to_nat(133u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -61934,7 +62101,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_builtin__initialize _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(132u); +x_1 = lean_unsigned_to_nat(133u); x_2 = lean_unsigned_to_nat(189u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -61962,7 +62129,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_builtin__initialize _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(132u); +x_1 = lean_unsigned_to_nat(133u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -61974,7 +62141,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_builtin__initialize _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(132u); +x_1 = lean_unsigned_to_nat(133u); x_2 = lean_unsigned_to_nat(48u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -62298,7 +62465,7 @@ x_4 = l_Lean_Parser_checkPrecFn(x_3, x_1, x_2); x_5 = lean_ctor_get(x_4, 4); lean_inc(x_5); x_6 = lean_box(0); -x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); +x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); lean_dec(x_5); if (x_7 == 0) { @@ -62312,7 +62479,7 @@ x_8 = lean_unsigned_to_nat(0u); x_9 = l_Lean_Parser_checkLhsPrecFn(x_8, x_1, x_4); x_10 = lean_ctor_get(x_9, 4); lean_inc(x_10); -x_11 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_10, x_6); +x_11 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_10, x_6); lean_dec(x_10); if (x_11 == 0) { @@ -62334,7 +62501,7 @@ x_17 = l_Lean_Parser_ParserState_mkTrailingNode(x_15, x_16, x_13); lean_dec(x_13); x_18 = lean_ctor_get(x_17, 4); lean_inc(x_18); -x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_18, x_6); +x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_18, x_6); lean_dec(x_18); if (x_19 == 0) { @@ -62458,7 +62625,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_in_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(134u); +x_1 = lean_unsigned_to_nat(135u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -62470,7 +62637,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_in_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(134u); +x_1 = lean_unsigned_to_nat(135u); x_2 = lean_unsigned_to_nat(87u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -62498,7 +62665,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_in_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(134u); +x_1 = lean_unsigned_to_nat(135u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -62510,7 +62677,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_in_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(134u); +x_1 = lean_unsigned_to_nat(135u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -62856,7 +63023,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -62887,7 +63054,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_20, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -62896,7 +63063,7 @@ x_27 = l_Lean_Parser_Command_genInjectiveTheorems___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_20, x_27, x_17); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -62920,7 +63087,7 @@ x_33 = l_Lean_Parser_Command_genInjectiveTheorems___elambda__1___closed__2; x_34 = l_Lean_Parser_ParserState_mkNode(x_32, x_33, x_17); x_35 = lean_ctor_get(x_34, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_14); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_14); lean_dec(x_35); if (x_36 == 0) { @@ -62943,7 +63110,7 @@ lean_inc(x_1); x_38 = l_Lean_Parser_tokenAntiquotFn(x_1, x_20); x_39 = lean_ctor_get(x_38, 4); lean_inc(x_39); -x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); +x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); lean_dec(x_39); if (x_40 == 0) { @@ -62952,7 +63119,7 @@ x_41 = l_Lean_Parser_Command_genInjectiveTheorems___elambda__1___closed__2; x_42 = l_Lean_Parser_ParserState_mkNode(x_38, x_41, x_17); x_43 = lean_ctor_get(x_42, 4); lean_inc(x_43); -x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_14); +x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_14); lean_dec(x_43); if (x_44 == 0) { @@ -62976,7 +63143,7 @@ x_47 = l_Lean_Parser_Command_genInjectiveTheorems___elambda__1___closed__2; x_48 = l_Lean_Parser_ParserState_mkNode(x_46, x_47, x_17); x_49 = lean_ctor_get(x_48, 4); lean_inc(x_49); -x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_14); +x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_14); lean_dec(x_49); if (x_50 == 0) { @@ -63113,7 +63280,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_genInjectiveTheorem _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(139u); +x_1 = lean_unsigned_to_nat(140u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -63125,7 +63292,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_genInjectiveTheorem _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(139u); +x_1 = lean_unsigned_to_nat(140u); x_2 = lean_unsigned_to_nat(102u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -63153,7 +63320,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_genInjectiveTheorem _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(139u); +x_1 = lean_unsigned_to_nat(140u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -63165,7 +63332,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_genInjectiveTheorem _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(139u); +x_1 = lean_unsigned_to_nat(140u); x_2 = lean_unsigned_to_nat(48u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -63451,7 +63618,7 @@ x_1 = l_Lean_Parser_Command_ctor___closed__2; return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__1() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -63461,7 +63628,7 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__2() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__2() { _start: { lean_object* x_1; lean_object* x_2; @@ -63471,7 +63638,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__3() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__3() { _start: { lean_object* x_1; @@ -63479,27 +63646,49 @@ x_1 = lean_mk_string_from_bytes("declModifiersF", 14); return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__4() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_quot___elambda__1___closed__2; -x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__3; +x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__5() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__4; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__4; +x_1 = lean_unsigned_to_nat(1u); x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__6() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__7() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__6; +x_2 = 1; +x_3 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__8() { _start: { lean_object* x_1; @@ -63507,17 +63696,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Command_declModifiersF_formatter) return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__7() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__6; +x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__8; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__8() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__10() { _start: { lean_object* x_1; @@ -63525,7 +63714,7 @@ x_1 = l_Lean_PrettyPrinter_Formatter_formatterAliasesRef; return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__9() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__11() { _start: { lean_object* x_1; @@ -63533,17 +63722,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Command_declModifiersF_parenthesi return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__10() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__9; +x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__11; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__11() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__13() { _start: { lean_object* x_1; @@ -63551,7 +63740,7 @@ x_1 = l_Lean_PrettyPrinter_Parenthesizer_parenthesizerAliasesRef; return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__12() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__14() { _start: { lean_object* x_1; @@ -63559,17 +63748,17 @@ x_1 = lean_mk_string_from_bytes("nestedDeclModifiers", 19); return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__13() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__12; +x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__14; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__14() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__16() { _start: { lean_object* x_1; lean_object* x_2; @@ -63579,7 +63768,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__15() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__17() { _start: { lean_object* x_1; @@ -63587,27 +63776,27 @@ x_1 = lean_mk_string_from_bytes("declModifiersT", 14); return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__16() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_quot___elambda__1___closed__2; -x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__15; +x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__17; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__17() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__19() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__16; +x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__18; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__18() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__20() { _start: { lean_object* x_1; @@ -63615,17 +63804,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Command_declModifiersT_formatter) return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__19() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__21() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__18; +x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__20; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__20() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__22() { _start: { lean_object* x_1; @@ -63633,17 +63822,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Command_declModifiersT_parenthesi return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__21() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__23() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__20; +x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__22; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__22() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -63653,7 +63842,7 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__23() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__25() { _start: { lean_object* x_1; lean_object* x_2; @@ -63663,7 +63852,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__24() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__26() { _start: { lean_object* x_1; lean_object* x_2; @@ -63673,7 +63862,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__25() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__27() { _start: { lean_object* x_1; lean_object* x_2; @@ -63683,7 +63872,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__26() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__28() { _start: { lean_object* x_1; lean_object* x_2; @@ -63693,7 +63882,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__27() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__29() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -63703,7 +63892,7 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__28() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__30() { _start: { lean_object* x_1; lean_object* x_2; @@ -63713,7 +63902,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__29() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__31() { _start: { lean_object* x_1; lean_object* x_2; @@ -63723,7 +63912,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__30() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__32() { _start: { lean_object* x_1; lean_object* x_2; @@ -63733,7 +63922,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__31() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__33() { _start: { lean_object* x_1; lean_object* x_2; @@ -63743,25 +63932,17 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__32() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("declVal", 7); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__33() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__34() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__32; +x_2 = l_Lean_Parser_Command_declVal___elambda__1___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__34() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__35() { _start: { lean_object* x_1; lean_object* x_2; @@ -63771,27 +63952,17 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__35() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_quot___elambda__1___closed__2; -x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__32; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__36() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__36() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__35; +x_1 = l_Lean_Parser_Command_declVal___elambda__1___closed__2; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__37() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__37() { _start: { lean_object* x_1; lean_object* x_2; @@ -63801,7 +63972,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__38() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__38() { _start: { lean_object* x_1; lean_object* x_2; @@ -63811,7 +63982,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__39() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__39() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -63821,7 +63992,7 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__40() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__40() { _start: { lean_object* x_1; lean_object* x_2; @@ -63831,7 +64002,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__41() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__41() { _start: { lean_object* x_1; lean_object* x_2; @@ -63841,7 +64012,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__42() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__42() { _start: { lean_object* x_1; lean_object* x_2; @@ -63851,7 +64022,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__43() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__43() { _start: { lean_object* x_1; lean_object* x_2; @@ -63861,7 +64032,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__44() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__44() { _start: { lean_object* x_1; @@ -63869,17 +64040,17 @@ x_1 = lean_mk_string_from_bytes("openDecl", 8); return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__45() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__45() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__44; +x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__44; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__46() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__46() { _start: { lean_object* x_1; lean_object* x_2; @@ -63889,27 +64060,27 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__47() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__47() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_quot___elambda__1___closed__2; -x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__44; +x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__44; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__48() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__48() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__47; +x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__47; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__49() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__49() { _start: { lean_object* x_1; lean_object* x_2; @@ -63919,7 +64090,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__50() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__50() { _start: { lean_object* x_1; lean_object* x_2; @@ -63929,647 +64100,648 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424_(lean_object* x_1) { _start: { -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__1; -x_3 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__2; -x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__5; -x_5 = l_Lean_Parser_registerAlias(x_2, x_3, x_4, x_1); -if (lean_obj_tag(x_5) == 0) -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_6 = lean_ctor_get(x_5, 1); -lean_inc(x_6); -lean_dec(x_5); -x_7 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__8; -x_8 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__7; -x_9 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_2, x_8, x_6); -if (lean_obj_tag(x_9) == 0) -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_10 = lean_ctor_get(x_9, 1); -lean_inc(x_10); -lean_dec(x_9); -x_11 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__11; -x_12 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__10; -x_13 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_2, x_12, x_10); -if (lean_obj_tag(x_13) == 0) -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -lean_dec(x_13); -x_15 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__13; -x_16 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__14; -x_17 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__17; -x_18 = l_Lean_Parser_registerAlias(x_15, x_16, x_17, x_14); -if (lean_obj_tag(x_18) == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -lean_dec(x_18); -x_20 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__19; -x_21 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_15, x_20, x_19); -if (lean_obj_tag(x_21) == 0) -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_21, 1); -lean_inc(x_22); -lean_dec(x_21); -x_23 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__21; -x_24 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_15, x_23, x_22); -if (lean_obj_tag(x_24) == 0) -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_25 = lean_ctor_get(x_24, 1); -lean_inc(x_25); -lean_dec(x_24); -x_26 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__22; -x_27 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__23; -x_28 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__24; -x_29 = l_Lean_Parser_registerAlias(x_26, x_27, x_28, x_25); -if (lean_obj_tag(x_29) == 0) -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_29, 1); -lean_inc(x_30); -lean_dec(x_29); -x_31 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__25; -x_32 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_26, x_31, x_30); -if (lean_obj_tag(x_32) == 0) -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_32, 1); -lean_inc(x_33); -lean_dec(x_32); -x_34 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__26; -x_35 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_26, x_34, x_33); -if (lean_obj_tag(x_35) == 0) +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__1; +x_3 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__2; +x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__5; +x_5 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__7; +x_6 = l_Lean_Parser_registerAlias(x_2, x_3, x_4, x_5, x_1); +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_7 = lean_ctor_get(x_6, 1); +lean_inc(x_7); +lean_dec(x_6); +x_8 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__10; +x_9 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__9; +x_10 = l_Lean_Parser_registerAliasCore___rarg(x_8, x_2, x_9, x_7); +if (lean_obj_tag(x_10) == 0) { -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_36 = lean_ctor_get(x_35, 1); -lean_inc(x_36); -lean_dec(x_35); -x_37 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__27; -x_38 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__28; -x_39 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__29; -x_40 = l_Lean_Parser_registerAlias(x_37, x_38, x_39, x_36); -if (lean_obj_tag(x_40) == 0) +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +lean_dec(x_10); +x_12 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__13; +x_13 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__12; +x_14 = l_Lean_Parser_registerAliasCore___rarg(x_12, x_2, x_13, x_11); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_40, 1); -lean_inc(x_41); -lean_dec(x_40); -x_42 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__30; -x_43 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_37, x_42, x_41); -if (lean_obj_tag(x_43) == 0) +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +lean_dec(x_14); +x_16 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__15; +x_17 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__16; +x_18 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__19; +x_19 = l_Lean_Parser_registerAlias(x_16, x_17, x_18, x_5, x_15); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_19, 1); +lean_inc(x_20); +lean_dec(x_19); +x_21 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__21; +x_22 = l_Lean_Parser_registerAliasCore___rarg(x_8, x_16, x_21, x_20); +if (lean_obj_tag(x_22) == 0) { -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_43, 1); -lean_inc(x_44); -lean_dec(x_43); -x_45 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__31; -x_46 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_37, x_45, x_44); -if (lean_obj_tag(x_46) == 0) +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +lean_dec(x_22); +x_24 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__23; +x_25 = l_Lean_Parser_registerAliasCore___rarg(x_12, x_16, x_24, x_23); +if (lean_obj_tag(x_25) == 0) { -lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_47 = lean_ctor_get(x_46, 1); -lean_inc(x_47); -lean_dec(x_46); -x_48 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__33; -x_49 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__34; -x_50 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__36; -x_51 = l_Lean_Parser_registerAlias(x_48, x_49, x_50, x_47); -if (lean_obj_tag(x_51) == 0) -{ -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_51, 1); -lean_inc(x_52); -lean_dec(x_51); -x_53 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__37; -x_54 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_48, x_53, x_52); -if (lean_obj_tag(x_54) == 0) +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +x_27 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__24; +x_28 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__25; +x_29 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__26; +x_30 = l_Lean_Parser_registerAlias(x_27, x_28, x_29, x_5, x_26); +if (lean_obj_tag(x_30) == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_32 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__27; +x_33 = l_Lean_Parser_registerAliasCore___rarg(x_8, x_27, x_32, x_31); +if (lean_obj_tag(x_33) == 0) { -lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_55 = lean_ctor_get(x_54, 1); -lean_inc(x_55); -lean_dec(x_54); -x_56 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__38; -x_57 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_48, x_56, x_55); -if (lean_obj_tag(x_57) == 0) +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_33, 1); +lean_inc(x_34); +lean_dec(x_33); +x_35 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__28; +x_36 = l_Lean_Parser_registerAliasCore___rarg(x_12, x_27, x_35, x_34); +if (lean_obj_tag(x_36) == 0) { -lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_58 = lean_ctor_get(x_57, 1); -lean_inc(x_58); -lean_dec(x_57); -x_59 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__39; -x_60 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__40; -x_61 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__41; -x_62 = l_Lean_Parser_registerAlias(x_59, x_60, x_61, x_58); -if (lean_obj_tag(x_62) == 0) -{ -lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_63 = lean_ctor_get(x_62, 1); -lean_inc(x_63); -lean_dec(x_62); -x_64 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__42; -x_65 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_59, x_64, x_63); -if (lean_obj_tag(x_65) == 0) +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_37 = lean_ctor_get(x_36, 1); +lean_inc(x_37); +lean_dec(x_36); +x_38 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__29; +x_39 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__30; +x_40 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__31; +x_41 = l_Lean_Parser_registerAlias(x_38, x_39, x_40, x_5, x_37); +if (lean_obj_tag(x_41) == 0) { -lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_66 = lean_ctor_get(x_65, 1); -lean_inc(x_66); -lean_dec(x_65); -x_67 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__43; -x_68 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_59, x_67, x_66); -if (lean_obj_tag(x_68) == 0) +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_41, 1); +lean_inc(x_42); +lean_dec(x_41); +x_43 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__32; +x_44 = l_Lean_Parser_registerAliasCore___rarg(x_8, x_38, x_43, x_42); +if (lean_obj_tag(x_44) == 0) { -lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; -x_69 = lean_ctor_get(x_68, 1); -lean_inc(x_69); -lean_dec(x_68); -x_70 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__45; -x_71 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__46; -x_72 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__48; -x_73 = l_Lean_Parser_registerAlias(x_70, x_71, x_72, x_69); -if (lean_obj_tag(x_73) == 0) -{ -lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_74 = lean_ctor_get(x_73, 1); -lean_inc(x_74); -lean_dec(x_73); -x_75 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__49; -x_76 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_70, x_75, x_74); -if (lean_obj_tag(x_76) == 0) -{ -lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_77 = lean_ctor_get(x_76, 1); -lean_inc(x_77); -lean_dec(x_76); -x_78 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__50; -x_79 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_70, x_78, x_77); -return x_79; -} -else +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_44, 1); +lean_inc(x_45); +lean_dec(x_44); +x_46 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__33; +x_47 = l_Lean_Parser_registerAliasCore___rarg(x_12, x_38, x_46, x_45); +if (lean_obj_tag(x_47) == 0) { -uint8_t x_80; -x_80 = !lean_is_exclusive(x_76); -if (x_80 == 0) +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_48 = lean_ctor_get(x_47, 1); +lean_inc(x_48); +lean_dec(x_47); +x_49 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__34; +x_50 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__35; +x_51 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__36; +x_52 = l_Lean_Parser_registerAlias(x_49, x_50, x_51, x_5, x_48); +if (lean_obj_tag(x_52) == 0) +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_52, 1); +lean_inc(x_53); +lean_dec(x_52); +x_54 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__37; +x_55 = l_Lean_Parser_registerAliasCore___rarg(x_8, x_49, x_54, x_53); +if (lean_obj_tag(x_55) == 0) { -return x_76; -} -else +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_55, 1); +lean_inc(x_56); +lean_dec(x_55); +x_57 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__38; +x_58 = l_Lean_Parser_registerAliasCore___rarg(x_12, x_49, x_57, x_56); +if (lean_obj_tag(x_58) == 0) { -lean_object* x_81; lean_object* x_82; lean_object* x_83; -x_81 = lean_ctor_get(x_76, 0); -x_82 = lean_ctor_get(x_76, 1); +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_59 = lean_ctor_get(x_58, 1); +lean_inc(x_59); +lean_dec(x_58); +x_60 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__39; +x_61 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__40; +x_62 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__41; +x_63 = l_Lean_Parser_registerAlias(x_60, x_61, x_62, x_5, x_59); +if (lean_obj_tag(x_63) == 0) +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_63, 1); +lean_inc(x_64); +lean_dec(x_63); +x_65 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__42; +x_66 = l_Lean_Parser_registerAliasCore___rarg(x_8, x_60, x_65, x_64); +if (lean_obj_tag(x_66) == 0) +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_66, 1); +lean_inc(x_67); +lean_dec(x_66); +x_68 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__43; +x_69 = l_Lean_Parser_registerAliasCore___rarg(x_12, x_60, x_68, x_67); +if (lean_obj_tag(x_69) == 0) +{ +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_70 = lean_ctor_get(x_69, 1); +lean_inc(x_70); +lean_dec(x_69); +x_71 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__45; +x_72 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__46; +x_73 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__48; +x_74 = l_Lean_Parser_registerAlias(x_71, x_72, x_73, x_5, x_70); +if (lean_obj_tag(x_74) == 0) +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_75 = lean_ctor_get(x_74, 1); +lean_inc(x_75); +lean_dec(x_74); +x_76 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__49; +x_77 = l_Lean_Parser_registerAliasCore___rarg(x_8, x_71, x_76, x_75); +if (lean_obj_tag(x_77) == 0) +{ +lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_78 = lean_ctor_get(x_77, 1); +lean_inc(x_78); +lean_dec(x_77); +x_79 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__50; +x_80 = l_Lean_Parser_registerAliasCore___rarg(x_12, x_71, x_79, x_78); +return x_80; +} +else +{ +uint8_t x_81; +x_81 = !lean_is_exclusive(x_77); +if (x_81 == 0) +{ +return x_77; +} +else +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_82 = lean_ctor_get(x_77, 0); +x_83 = lean_ctor_get(x_77, 1); +lean_inc(x_83); lean_inc(x_82); -lean_inc(x_81); -lean_dec(x_76); -x_83 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_83, 0, x_81); -lean_ctor_set(x_83, 1, x_82); -return x_83; +lean_dec(x_77); +x_84 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_84, 0, x_82); +lean_ctor_set(x_84, 1, x_83); +return x_84; } } } else { -uint8_t x_84; -x_84 = !lean_is_exclusive(x_73); -if (x_84 == 0) +uint8_t x_85; +x_85 = !lean_is_exclusive(x_74); +if (x_85 == 0) { -return x_73; +return x_74; } else { -lean_object* x_85; lean_object* x_86; lean_object* x_87; -x_85 = lean_ctor_get(x_73, 0); -x_86 = lean_ctor_get(x_73, 1); +lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_86 = lean_ctor_get(x_74, 0); +x_87 = lean_ctor_get(x_74, 1); +lean_inc(x_87); lean_inc(x_86); -lean_inc(x_85); -lean_dec(x_73); -x_87 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_87, 0, x_85); -lean_ctor_set(x_87, 1, x_86); -return x_87; +lean_dec(x_74); +x_88 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_88, 0, x_86); +lean_ctor_set(x_88, 1, x_87); +return x_88; } } } else { -uint8_t x_88; -x_88 = !lean_is_exclusive(x_68); -if (x_88 == 0) +uint8_t x_89; +x_89 = !lean_is_exclusive(x_69); +if (x_89 == 0) { -return x_68; +return x_69; } else { -lean_object* x_89; lean_object* x_90; lean_object* x_91; -x_89 = lean_ctor_get(x_68, 0); -x_90 = lean_ctor_get(x_68, 1); +lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_90 = lean_ctor_get(x_69, 0); +x_91 = lean_ctor_get(x_69, 1); +lean_inc(x_91); lean_inc(x_90); -lean_inc(x_89); -lean_dec(x_68); -x_91 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_91, 0, x_89); -lean_ctor_set(x_91, 1, x_90); -return x_91; +lean_dec(x_69); +x_92 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_91); +return x_92; } } } else { -uint8_t x_92; -x_92 = !lean_is_exclusive(x_65); -if (x_92 == 0) +uint8_t x_93; +x_93 = !lean_is_exclusive(x_66); +if (x_93 == 0) { -return x_65; +return x_66; } else { -lean_object* x_93; lean_object* x_94; lean_object* x_95; -x_93 = lean_ctor_get(x_65, 0); -x_94 = lean_ctor_get(x_65, 1); +lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_94 = lean_ctor_get(x_66, 0); +x_95 = lean_ctor_get(x_66, 1); +lean_inc(x_95); lean_inc(x_94); -lean_inc(x_93); -lean_dec(x_65); -x_95 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_95, 0, x_93); -lean_ctor_set(x_95, 1, x_94); -return x_95; +lean_dec(x_66); +x_96 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_96, 0, x_94); +lean_ctor_set(x_96, 1, x_95); +return x_96; } } } else { -uint8_t x_96; -x_96 = !lean_is_exclusive(x_62); -if (x_96 == 0) +uint8_t x_97; +x_97 = !lean_is_exclusive(x_63); +if (x_97 == 0) { -return x_62; +return x_63; } else { -lean_object* x_97; lean_object* x_98; lean_object* x_99; -x_97 = lean_ctor_get(x_62, 0); -x_98 = lean_ctor_get(x_62, 1); +lean_object* x_98; lean_object* x_99; lean_object* x_100; +x_98 = lean_ctor_get(x_63, 0); +x_99 = lean_ctor_get(x_63, 1); +lean_inc(x_99); lean_inc(x_98); -lean_inc(x_97); -lean_dec(x_62); -x_99 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_99, 0, x_97); -lean_ctor_set(x_99, 1, x_98); -return x_99; +lean_dec(x_63); +x_100 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_100, 0, x_98); +lean_ctor_set(x_100, 1, x_99); +return x_100; } } } else { -uint8_t x_100; -x_100 = !lean_is_exclusive(x_57); -if (x_100 == 0) +uint8_t x_101; +x_101 = !lean_is_exclusive(x_58); +if (x_101 == 0) { -return x_57; +return x_58; } else { -lean_object* x_101; lean_object* x_102; lean_object* x_103; -x_101 = lean_ctor_get(x_57, 0); -x_102 = lean_ctor_get(x_57, 1); +lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_102 = lean_ctor_get(x_58, 0); +x_103 = lean_ctor_get(x_58, 1); +lean_inc(x_103); lean_inc(x_102); -lean_inc(x_101); -lean_dec(x_57); -x_103 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_103, 0, x_101); -lean_ctor_set(x_103, 1, x_102); -return x_103; +lean_dec(x_58); +x_104 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_104, 0, x_102); +lean_ctor_set(x_104, 1, x_103); +return x_104; } } } else { -uint8_t x_104; -x_104 = !lean_is_exclusive(x_54); -if (x_104 == 0) +uint8_t x_105; +x_105 = !lean_is_exclusive(x_55); +if (x_105 == 0) { -return x_54; +return x_55; } else { -lean_object* x_105; lean_object* x_106; lean_object* x_107; -x_105 = lean_ctor_get(x_54, 0); -x_106 = lean_ctor_get(x_54, 1); +lean_object* x_106; lean_object* x_107; lean_object* x_108; +x_106 = lean_ctor_get(x_55, 0); +x_107 = lean_ctor_get(x_55, 1); +lean_inc(x_107); lean_inc(x_106); -lean_inc(x_105); -lean_dec(x_54); -x_107 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_107, 0, x_105); -lean_ctor_set(x_107, 1, x_106); -return x_107; +lean_dec(x_55); +x_108 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_108, 0, x_106); +lean_ctor_set(x_108, 1, x_107); +return x_108; } } } else { -uint8_t x_108; -x_108 = !lean_is_exclusive(x_51); -if (x_108 == 0) +uint8_t x_109; +x_109 = !lean_is_exclusive(x_52); +if (x_109 == 0) { -return x_51; +return x_52; } else { -lean_object* x_109; lean_object* x_110; lean_object* x_111; -x_109 = lean_ctor_get(x_51, 0); -x_110 = lean_ctor_get(x_51, 1); +lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_110 = lean_ctor_get(x_52, 0); +x_111 = lean_ctor_get(x_52, 1); +lean_inc(x_111); lean_inc(x_110); -lean_inc(x_109); -lean_dec(x_51); -x_111 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_111, 0, x_109); -lean_ctor_set(x_111, 1, x_110); -return x_111; +lean_dec(x_52); +x_112 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_112, 0, x_110); +lean_ctor_set(x_112, 1, x_111); +return x_112; } } } else { -uint8_t x_112; -x_112 = !lean_is_exclusive(x_46); -if (x_112 == 0) +uint8_t x_113; +x_113 = !lean_is_exclusive(x_47); +if (x_113 == 0) { -return x_46; +return x_47; } else { -lean_object* x_113; lean_object* x_114; lean_object* x_115; -x_113 = lean_ctor_get(x_46, 0); -x_114 = lean_ctor_get(x_46, 1); +lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_114 = lean_ctor_get(x_47, 0); +x_115 = lean_ctor_get(x_47, 1); +lean_inc(x_115); lean_inc(x_114); -lean_inc(x_113); -lean_dec(x_46); -x_115 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_115, 0, x_113); -lean_ctor_set(x_115, 1, x_114); -return x_115; +lean_dec(x_47); +x_116 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_116, 0, x_114); +lean_ctor_set(x_116, 1, x_115); +return x_116; } } } else { -uint8_t x_116; -x_116 = !lean_is_exclusive(x_43); -if (x_116 == 0) +uint8_t x_117; +x_117 = !lean_is_exclusive(x_44); +if (x_117 == 0) { -return x_43; +return x_44; } else { -lean_object* x_117; lean_object* x_118; lean_object* x_119; -x_117 = lean_ctor_get(x_43, 0); -x_118 = lean_ctor_get(x_43, 1); +lean_object* x_118; lean_object* x_119; lean_object* x_120; +x_118 = lean_ctor_get(x_44, 0); +x_119 = lean_ctor_get(x_44, 1); +lean_inc(x_119); lean_inc(x_118); -lean_inc(x_117); -lean_dec(x_43); -x_119 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_119, 0, x_117); -lean_ctor_set(x_119, 1, x_118); -return x_119; +lean_dec(x_44); +x_120 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_120, 0, x_118); +lean_ctor_set(x_120, 1, x_119); +return x_120; } } } else { -uint8_t x_120; -x_120 = !lean_is_exclusive(x_40); -if (x_120 == 0) +uint8_t x_121; +x_121 = !lean_is_exclusive(x_41); +if (x_121 == 0) { -return x_40; +return x_41; } else { -lean_object* x_121; lean_object* x_122; lean_object* x_123; -x_121 = lean_ctor_get(x_40, 0); -x_122 = lean_ctor_get(x_40, 1); +lean_object* x_122; lean_object* x_123; lean_object* x_124; +x_122 = lean_ctor_get(x_41, 0); +x_123 = lean_ctor_get(x_41, 1); +lean_inc(x_123); lean_inc(x_122); -lean_inc(x_121); -lean_dec(x_40); -x_123 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_123, 0, x_121); -lean_ctor_set(x_123, 1, x_122); -return x_123; +lean_dec(x_41); +x_124 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_124, 0, x_122); +lean_ctor_set(x_124, 1, x_123); +return x_124; } } } else { -uint8_t x_124; -x_124 = !lean_is_exclusive(x_35); -if (x_124 == 0) +uint8_t x_125; +x_125 = !lean_is_exclusive(x_36); +if (x_125 == 0) { -return x_35; +return x_36; } else { -lean_object* x_125; lean_object* x_126; lean_object* x_127; -x_125 = lean_ctor_get(x_35, 0); -x_126 = lean_ctor_get(x_35, 1); +lean_object* x_126; lean_object* x_127; lean_object* x_128; +x_126 = lean_ctor_get(x_36, 0); +x_127 = lean_ctor_get(x_36, 1); +lean_inc(x_127); lean_inc(x_126); -lean_inc(x_125); -lean_dec(x_35); -x_127 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_127, 0, x_125); -lean_ctor_set(x_127, 1, x_126); -return x_127; +lean_dec(x_36); +x_128 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_128, 0, x_126); +lean_ctor_set(x_128, 1, x_127); +return x_128; } } } else { -uint8_t x_128; -x_128 = !lean_is_exclusive(x_32); -if (x_128 == 0) +uint8_t x_129; +x_129 = !lean_is_exclusive(x_33); +if (x_129 == 0) { -return x_32; +return x_33; } else { -lean_object* x_129; lean_object* x_130; lean_object* x_131; -x_129 = lean_ctor_get(x_32, 0); -x_130 = lean_ctor_get(x_32, 1); +lean_object* x_130; lean_object* x_131; lean_object* x_132; +x_130 = lean_ctor_get(x_33, 0); +x_131 = lean_ctor_get(x_33, 1); +lean_inc(x_131); lean_inc(x_130); -lean_inc(x_129); -lean_dec(x_32); -x_131 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_131, 0, x_129); -lean_ctor_set(x_131, 1, x_130); -return x_131; +lean_dec(x_33); +x_132 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_132, 0, x_130); +lean_ctor_set(x_132, 1, x_131); +return x_132; } } } else { -uint8_t x_132; -x_132 = !lean_is_exclusive(x_29); -if (x_132 == 0) +uint8_t x_133; +x_133 = !lean_is_exclusive(x_30); +if (x_133 == 0) { -return x_29; +return x_30; } else { -lean_object* x_133; lean_object* x_134; lean_object* x_135; -x_133 = lean_ctor_get(x_29, 0); -x_134 = lean_ctor_get(x_29, 1); +lean_object* x_134; lean_object* x_135; lean_object* x_136; +x_134 = lean_ctor_get(x_30, 0); +x_135 = lean_ctor_get(x_30, 1); +lean_inc(x_135); lean_inc(x_134); -lean_inc(x_133); -lean_dec(x_29); -x_135 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_135, 0, x_133); -lean_ctor_set(x_135, 1, x_134); -return x_135; +lean_dec(x_30); +x_136 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_136, 0, x_134); +lean_ctor_set(x_136, 1, x_135); +return x_136; } } } else { -uint8_t x_136; -x_136 = !lean_is_exclusive(x_24); -if (x_136 == 0) +uint8_t x_137; +x_137 = !lean_is_exclusive(x_25); +if (x_137 == 0) { -return x_24; +return x_25; } else { -lean_object* x_137; lean_object* x_138; lean_object* x_139; -x_137 = lean_ctor_get(x_24, 0); -x_138 = lean_ctor_get(x_24, 1); +lean_object* x_138; lean_object* x_139; lean_object* x_140; +x_138 = lean_ctor_get(x_25, 0); +x_139 = lean_ctor_get(x_25, 1); +lean_inc(x_139); lean_inc(x_138); -lean_inc(x_137); -lean_dec(x_24); -x_139 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_139, 0, x_137); -lean_ctor_set(x_139, 1, x_138); -return x_139; +lean_dec(x_25); +x_140 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_140, 0, x_138); +lean_ctor_set(x_140, 1, x_139); +return x_140; } } } else { -uint8_t x_140; -x_140 = !lean_is_exclusive(x_21); -if (x_140 == 0) +uint8_t x_141; +x_141 = !lean_is_exclusive(x_22); +if (x_141 == 0) { -return x_21; +return x_22; } else { -lean_object* x_141; lean_object* x_142; lean_object* x_143; -x_141 = lean_ctor_get(x_21, 0); -x_142 = lean_ctor_get(x_21, 1); +lean_object* x_142; lean_object* x_143; lean_object* x_144; +x_142 = lean_ctor_get(x_22, 0); +x_143 = lean_ctor_get(x_22, 1); +lean_inc(x_143); lean_inc(x_142); -lean_inc(x_141); -lean_dec(x_21); -x_143 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_143, 0, x_141); -lean_ctor_set(x_143, 1, x_142); -return x_143; +lean_dec(x_22); +x_144 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_144, 0, x_142); +lean_ctor_set(x_144, 1, x_143); +return x_144; } } } else { -uint8_t x_144; -x_144 = !lean_is_exclusive(x_18); -if (x_144 == 0) +uint8_t x_145; +x_145 = !lean_is_exclusive(x_19); +if (x_145 == 0) { -return x_18; +return x_19; } else { -lean_object* x_145; lean_object* x_146; lean_object* x_147; -x_145 = lean_ctor_get(x_18, 0); -x_146 = lean_ctor_get(x_18, 1); +lean_object* x_146; lean_object* x_147; lean_object* x_148; +x_146 = lean_ctor_get(x_19, 0); +x_147 = lean_ctor_get(x_19, 1); +lean_inc(x_147); lean_inc(x_146); -lean_inc(x_145); -lean_dec(x_18); -x_147 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_147, 0, x_145); -lean_ctor_set(x_147, 1, x_146); -return x_147; +lean_dec(x_19); +x_148 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_148, 0, x_146); +lean_ctor_set(x_148, 1, x_147); +return x_148; } } } else { -uint8_t x_148; -x_148 = !lean_is_exclusive(x_13); -if (x_148 == 0) +uint8_t x_149; +x_149 = !lean_is_exclusive(x_14); +if (x_149 == 0) { -return x_13; +return x_14; } else { -lean_object* x_149; lean_object* x_150; lean_object* x_151; -x_149 = lean_ctor_get(x_13, 0); -x_150 = lean_ctor_get(x_13, 1); +lean_object* x_150; lean_object* x_151; lean_object* x_152; +x_150 = lean_ctor_get(x_14, 0); +x_151 = lean_ctor_get(x_14, 1); +lean_inc(x_151); lean_inc(x_150); -lean_inc(x_149); -lean_dec(x_13); -x_151 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_151, 0, x_149); -lean_ctor_set(x_151, 1, x_150); -return x_151; +lean_dec(x_14); +x_152 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_152, 0, x_150); +lean_ctor_set(x_152, 1, x_151); +return x_152; } } } else { -uint8_t x_152; -x_152 = !lean_is_exclusive(x_9); -if (x_152 == 0) +uint8_t x_153; +x_153 = !lean_is_exclusive(x_10); +if (x_153 == 0) { -return x_9; +return x_10; } else { -lean_object* x_153; lean_object* x_154; lean_object* x_155; -x_153 = lean_ctor_get(x_9, 0); -x_154 = lean_ctor_get(x_9, 1); +lean_object* x_154; lean_object* x_155; lean_object* x_156; +x_154 = lean_ctor_get(x_10, 0); +x_155 = lean_ctor_get(x_10, 1); +lean_inc(x_155); lean_inc(x_154); -lean_inc(x_153); -lean_dec(x_9); -x_155 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_155, 0, x_153); -lean_ctor_set(x_155, 1, x_154); -return x_155; +lean_dec(x_10); +x_156 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_156, 0, x_154); +lean_ctor_set(x_156, 1, x_155); +return x_156; } } } else { -uint8_t x_156; -x_156 = !lean_is_exclusive(x_5); -if (x_156 == 0) +uint8_t x_157; +x_157 = !lean_is_exclusive(x_6); +if (x_157 == 0) { -return x_5; +return x_6; } else { -lean_object* x_157; lean_object* x_158; lean_object* x_159; -x_157 = lean_ctor_get(x_5, 0); -x_158 = lean_ctor_get(x_5, 1); +lean_object* x_158; lean_object* x_159; lean_object* x_160; +x_158 = lean_ctor_get(x_6, 0); +x_159 = lean_ctor_get(x_6, 1); +lean_inc(x_159); lean_inc(x_158); -lean_inc(x_157); -lean_dec(x_5); -x_159 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_159, 0, x_157); -lean_ctor_set(x_159, 1, x_158); -return x_159; +lean_dec(x_6); +x_160 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_160, 0, x_158); +lean_ctor_set(x_160, 1, x_159); +return x_160; } } } @@ -64735,7 +64907,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -64779,7 +64951,7 @@ goto block_48; lean_object* x_26; uint8_t x_27; x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); lean_dec(x_26); if (x_27 == 0) { @@ -64788,7 +64960,7 @@ x_28 = l_Lean_Parser_Term_open___elambda__1___closed__1; x_29 = l_Lean_Parser_ParserState_mkNode(x_25, x_28, x_17); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); lean_dec(x_30); if (x_31 == 0) { @@ -64810,7 +64982,7 @@ lean_inc(x_1); x_33 = l_Lean_Parser_Command_openDecl___elambda__1(x_1, x_25); x_34 = lean_ctor_get(x_33, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_14); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_14); lean_dec(x_34); if (x_35 == 0) { @@ -64819,7 +64991,7 @@ x_36 = l_Lean_Parser_Term_open___elambda__1___closed__1; x_37 = l_Lean_Parser_ParserState_mkNode(x_33, x_36, x_17); x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_14); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_14); lean_dec(x_38); if (x_39 == 0) { @@ -64844,7 +65016,7 @@ x_43 = l_Lean_Parser_Term_open___elambda__1___closed__1; x_44 = l_Lean_Parser_ParserState_mkNode(x_42, x_43, x_17); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); lean_dec(x_45); if (x_46 == 0) { @@ -64995,7 +65167,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_open_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(156u); +x_1 = lean_unsigned_to_nat(157u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -65007,7 +65179,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_open_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(156u); +x_1 = lean_unsigned_to_nat(157u); x_2 = lean_unsigned_to_nat(125u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -65035,7 +65207,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_open_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(156u); +x_1 = lean_unsigned_to_nat(157u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -65047,7 +65219,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_open_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(156u); +x_1 = lean_unsigned_to_nat(157u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -65490,7 +65662,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -65535,7 +65707,7 @@ x_26 = l_Lean_Parser_Term_set__option___elambda__1___closed__1; x_27 = l_Lean_Parser_ParserState_mkNode(x_25, x_26, x_17); x_28 = lean_ctor_get(x_27, 4); lean_inc(x_28); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_14); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_14); lean_dec(x_28); if (x_29 == 0) { @@ -65555,7 +65727,7 @@ return x_30; lean_object* x_33; uint8_t x_34; x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -65570,7 +65742,7 @@ lean_inc(x_1); x_35 = l_Lean_Parser_Term_ident___elambda__1(x_1, x_32); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); lean_dec(x_36); if (x_37 == 0) { @@ -65585,7 +65757,7 @@ lean_inc(x_1); x_38 = l_Lean_Parser_Command_optionValue___elambda__1(x_1, x_35); x_39 = lean_ctor_get(x_38, 4); lean_inc(x_39); -x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); +x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); lean_dec(x_39); if (x_40 == 0) { @@ -65611,7 +65783,7 @@ if (x_46 == 0) lean_object* x_47; uint8_t x_48; x_47 = lean_ctor_get(x_43, 4); lean_inc(x_47); -x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_14); +x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_14); lean_dec(x_47); if (x_48 == 0) { @@ -65636,7 +65808,7 @@ lean_inc(x_1); x_52 = l_Lean_Parser_tokenAntiquotFn(x_1, x_43); x_53 = lean_ctor_get(x_52, 4); lean_inc(x_53); -x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); +x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); lean_dec(x_53); if (x_54 == 0) { @@ -65802,7 +65974,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_set__option_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(157u); +x_1 = lean_unsigned_to_nat(158u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -65814,7 +65986,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_set__option_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(157u); +x_1 = lean_unsigned_to_nat(158u); x_2 = lean_unsigned_to_nat(145u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -65842,7 +66014,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_set__option_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(157u); +x_1 = lean_unsigned_to_nat(158u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -65854,7 +66026,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_set__option_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(157u); +x_1 = lean_unsigned_to_nat(158u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -66313,7 +66485,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -66357,7 +66529,7 @@ goto block_48; lean_object* x_26; uint8_t x_27; x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); lean_dec(x_26); if (x_27 == 0) { @@ -66366,7 +66538,7 @@ x_28 = l_Lean_Parser_Tactic_open___elambda__1___closed__3; x_29 = l_Lean_Parser_ParserState_mkNode(x_25, x_28, x_17); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); lean_dec(x_30); if (x_31 == 0) { @@ -66388,7 +66560,7 @@ lean_inc(x_1); x_33 = l_Lean_Parser_Command_openDecl___elambda__1(x_1, x_25); x_34 = lean_ctor_get(x_33, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_14); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_14); lean_dec(x_34); if (x_35 == 0) { @@ -66397,7 +66569,7 @@ x_36 = l_Lean_Parser_Tactic_open___elambda__1___closed__3; x_37 = l_Lean_Parser_ParserState_mkNode(x_33, x_36, x_17); x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_14); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_14); lean_dec(x_38); if (x_39 == 0) { @@ -66422,7 +66594,7 @@ x_43 = l_Lean_Parser_Tactic_open___elambda__1___closed__3; x_44 = l_Lean_Parser_ParserState_mkNode(x_42, x_43, x_17); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); lean_dec(x_45); if (x_46 == 0) { @@ -66591,7 +66763,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_open_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(161u); +x_1 = lean_unsigned_to_nat(162u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -66603,7 +66775,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_open_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(161u); +x_1 = lean_unsigned_to_nat(162u); x_2 = lean_unsigned_to_nat(126u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -66631,7 +66803,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_open_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(161u); +x_1 = lean_unsigned_to_nat(162u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -66643,7 +66815,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_open_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(161u); +x_1 = lean_unsigned_to_nat(162u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -67058,7 +67230,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -67103,7 +67275,7 @@ x_26 = l_Lean_Parser_Tactic_set__option___elambda__1___closed__1; x_27 = l_Lean_Parser_ParserState_mkNode(x_25, x_26, x_17); x_28 = lean_ctor_get(x_27, 4); lean_inc(x_28); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_14); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_14); lean_dec(x_28); if (x_29 == 0) { @@ -67123,7 +67295,7 @@ return x_30; lean_object* x_33; uint8_t x_34; x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -67138,7 +67310,7 @@ lean_inc(x_1); x_35 = l_Lean_Parser_ident___elambda__1(x_1, x_32); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); lean_dec(x_36); if (x_37 == 0) { @@ -67153,7 +67325,7 @@ lean_inc(x_1); x_38 = l_Lean_Parser_Command_optionValue___elambda__1(x_1, x_35); x_39 = lean_ctor_get(x_38, 4); lean_inc(x_39); -x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); +x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); lean_dec(x_39); if (x_40 == 0) { @@ -67179,7 +67351,7 @@ if (x_46 == 0) lean_object* x_47; uint8_t x_48; x_47 = lean_ctor_get(x_43, 4); lean_inc(x_47); -x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_14); +x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_14); lean_dec(x_47); if (x_48 == 0) { @@ -67202,7 +67374,7 @@ lean_inc(x_1); x_50 = l_Lean_Parser_tokenAntiquotFn(x_1, x_43); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_14); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_14); lean_dec(x_51); if (x_52 == 0) { @@ -67366,7 +67538,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_set__option_declRang _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(162u); +x_1 = lean_unsigned_to_nat(163u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -67378,7 +67550,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_set__option_declRang _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(162u); +x_1 = lean_unsigned_to_nat(163u); x_2 = lean_unsigned_to_nat(146u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -67406,7 +67578,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_set__option_declRang _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(162u); +x_1 = lean_unsigned_to_nat(163u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -67418,7 +67590,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_set__option_declRang _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(162u); +x_1 = lean_unsigned_to_nat(163u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -68137,6 +68309,10 @@ l_Lean_Parser_Command_terminationHintMany___closed__19 = _init_l_Lean_Parser_Com lean_mark_persistent(l_Lean_Parser_Command_terminationHintMany___closed__19); l_Lean_Parser_Command_terminationHintMany___closed__20 = _init_l_Lean_Parser_Command_terminationHintMany___closed__20(); lean_mark_persistent(l_Lean_Parser_Command_terminationHintMany___closed__20); +l_Lean_Parser_Command_terminationHintMany___closed__21 = _init_l_Lean_Parser_Command_terminationHintMany___closed__21(); +lean_mark_persistent(l_Lean_Parser_Command_terminationHintMany___closed__21); +l_Lean_Parser_Command_terminationHintMany___closed__22 = _init_l_Lean_Parser_Command_terminationHintMany___closed__22(); +lean_mark_persistent(l_Lean_Parser_Command_terminationHintMany___closed__22); l_Lean_Parser_Command_terminationHint1___elambda__1___closed__1 = _init_l_Lean_Parser_Command_terminationHint1___elambda__1___closed__1(); lean_mark_persistent(l_Lean_Parser_Command_terminationHint1___elambda__1___closed__1); l_Lean_Parser_Command_terminationHint1___elambda__1___closed__2 = _init_l_Lean_Parser_Command_terminationHint1___elambda__1___closed__2(); @@ -69221,6 +69397,14 @@ l_Lean_Parser_Command_whereStructInst = _init_l_Lean_Parser_Command_whereStructI lean_mark_persistent(l_Lean_Parser_Command_whereStructInst); l_Lean_Parser_Command_declVal___elambda__1___closed__1 = _init_l_Lean_Parser_Command_declVal___elambda__1___closed__1(); lean_mark_persistent(l_Lean_Parser_Command_declVal___elambda__1___closed__1); +l_Lean_Parser_Command_declVal___elambda__1___closed__2 = _init_l_Lean_Parser_Command_declVal___elambda__1___closed__2(); +lean_mark_persistent(l_Lean_Parser_Command_declVal___elambda__1___closed__2); +l_Lean_Parser_Command_declVal___elambda__1___closed__3 = _init_l_Lean_Parser_Command_declVal___elambda__1___closed__3(); +lean_mark_persistent(l_Lean_Parser_Command_declVal___elambda__1___closed__3); +l_Lean_Parser_Command_declVal___elambda__1___closed__4 = _init_l_Lean_Parser_Command_declVal___elambda__1___closed__4(); +lean_mark_persistent(l_Lean_Parser_Command_declVal___elambda__1___closed__4); +l_Lean_Parser_Command_declVal___elambda__1___closed__5 = _init_l_Lean_Parser_Command_declVal___elambda__1___closed__5(); +lean_mark_persistent(l_Lean_Parser_Command_declVal___elambda__1___closed__5); l_Lean_Parser_Command_declVal___closed__1 = _init_l_Lean_Parser_Command_declVal___closed__1(); lean_mark_persistent(l_Lean_Parser_Command_declVal___closed__1); l_Lean_Parser_Command_declVal___closed__2 = _init_l_Lean_Parser_Command_declVal___closed__2(); @@ -69229,6 +69413,8 @@ l_Lean_Parser_Command_declVal___closed__3 = _init_l_Lean_Parser_Command_declVal_ lean_mark_persistent(l_Lean_Parser_Command_declVal___closed__3); l_Lean_Parser_Command_declVal___closed__4 = _init_l_Lean_Parser_Command_declVal___closed__4(); lean_mark_persistent(l_Lean_Parser_Command_declVal___closed__4); +l_Lean_Parser_Command_declVal___closed__5 = _init_l_Lean_Parser_Command_declVal___closed__5(); +lean_mark_persistent(l_Lean_Parser_Command_declVal___closed__5); l_Lean_Parser_Command_declVal = _init_l_Lean_Parser_Command_declVal(); lean_mark_persistent(l_Lean_Parser_Command_declVal); l_Lean_Parser_Command_abbrev___elambda__1___closed__1 = _init_l_Lean_Parser_Command_abbrev___elambda__1___closed__1(); @@ -70737,6 +70923,10 @@ if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); l_Lean_Parser_Command_declVal_formatter___closed__1 = _init_l_Lean_Parser_Command_declVal_formatter___closed__1(); lean_mark_persistent(l_Lean_Parser_Command_declVal_formatter___closed__1); +l_Lean_Parser_Command_declVal_formatter___closed__2 = _init_l_Lean_Parser_Command_declVal_formatter___closed__2(); +lean_mark_persistent(l_Lean_Parser_Command_declVal_formatter___closed__2); +l_Lean_Parser_Command_declVal_formatter___closed__3 = _init_l_Lean_Parser_Command_declVal_formatter___closed__3(); +lean_mark_persistent(l_Lean_Parser_Command_declVal_formatter___closed__3); l_Lean_Parser_Command_abbrev_formatter___closed__1 = _init_l_Lean_Parser_Command_abbrev_formatter___closed__1(); lean_mark_persistent(l_Lean_Parser_Command_abbrev_formatter___closed__1); l_Lean_Parser_Command_abbrev_formatter___closed__2 = _init_l_Lean_Parser_Command_abbrev_formatter___closed__2(); @@ -71685,6 +71875,10 @@ if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); l_Lean_Parser_Command_declVal_parenthesizer___closed__1 = _init_l_Lean_Parser_Command_declVal_parenthesizer___closed__1(); lean_mark_persistent(l_Lean_Parser_Command_declVal_parenthesizer___closed__1); +l_Lean_Parser_Command_declVal_parenthesizer___closed__2 = _init_l_Lean_Parser_Command_declVal_parenthesizer___closed__2(); +lean_mark_persistent(l_Lean_Parser_Command_declVal_parenthesizer___closed__2); +l_Lean_Parser_Command_declVal_parenthesizer___closed__3 = _init_l_Lean_Parser_Command_declVal_parenthesizer___closed__3(); +lean_mark_persistent(l_Lean_Parser_Command_declVal_parenthesizer___closed__3); l_Lean_Parser_Command_abbrev_parenthesizer___closed__1 = _init_l_Lean_Parser_Command_abbrev_parenthesizer___closed__1(); lean_mark_persistent(l_Lean_Parser_Command_abbrev_parenthesizer___closed__1); l_Lean_Parser_Command_abbrev_parenthesizer___closed__2 = _init_l_Lean_Parser_Command_abbrev_parenthesizer___closed__2(); @@ -75691,107 +75885,107 @@ l_Lean_Parser_Command_declModifiersF = _init_l_Lean_Parser_Command_declModifiers lean_mark_persistent(l_Lean_Parser_Command_declModifiersF); l_Lean_Parser_Command_declModifiersT = _init_l_Lean_Parser_Command_declModifiersT(); lean_mark_persistent(l_Lean_Parser_Command_declModifiersT); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__1 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__1(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__1); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__2 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__2(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__2); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__3 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__3(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__3); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__4 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__4(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__4); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__5 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__5(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__5); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__6 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__6(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__6); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__7 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__7(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__7); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__8 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__8(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__8); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__9 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__9(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__9); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__10 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__10(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__10); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__11 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__11(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__11); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__12 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__12(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__12); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__13 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__13(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__13); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__14 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__14(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__14); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__15 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__15(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__15); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__16 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__16(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__16); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__17 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__17(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__17); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__18 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__18(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__18); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__19 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__19(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__19); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__20 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__20(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__20); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__21 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__21(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__21); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__22 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__22(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__22); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__23 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__23(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__23); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__24 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__24(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__24); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__25 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__25(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__25); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__26 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__26(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__26); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__27 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__27(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__27); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__28 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__28(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__28); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__29 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__29(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__29); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__30 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__30(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__30); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__31 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__31(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__31); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__32 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__32(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__32); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__33 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__33(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__33); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__34 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__34(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__34); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__35 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__35(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__35); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__36 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__36(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__36); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__37 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__37(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__37); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__38 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__38(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__38); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__39 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__39(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__39); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__40 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__40(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__40); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__41 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__41(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__41); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__42 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__42(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__42); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__43 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__43(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__43); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__44 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__44(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__44); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__45 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__45(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__45); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__46 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__46(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__46); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__47 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__47(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__47); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__48 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__48(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__48); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__49 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__49(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__49); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__50 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__50(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415____closed__50); -res = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2415_(lean_io_mk_world()); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__1 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__1(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__1); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__2 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__2(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__2); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__3 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__3(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__3); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__4 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__4(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__4); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__5 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__5(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__5); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__6 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__6(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__6); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__7 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__7(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__7); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__8 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__8(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__8); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__9 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__9(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__9); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__10 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__10(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__10); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__11 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__11(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__11); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__12 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__12(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__12); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__13 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__13(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__13); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__14 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__14(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__14); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__15 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__15(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__15); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__16 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__16(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__16); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__17 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__17(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__17); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__18 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__18(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__18); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__19 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__19(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__19); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__20 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__20(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__20); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__21 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__21(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__21); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__22 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__22(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__22); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__23 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__23(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__23); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__24 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__24(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__24); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__25 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__25(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__25); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__26 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__26(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__26); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__27 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__27(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__27); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__28 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__28(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__28); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__29 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__29(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__29); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__30 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__30(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__30); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__31 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__31(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__31); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__32 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__32(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__32); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__33 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__33(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__33); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__34 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__34(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__34); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__35 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__35(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__35); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__36 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__36(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__36); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__37 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__37(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__37); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__38 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__38(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__38); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__39 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__39(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__39); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__40 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__40(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__40); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__41 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__41(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__41); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__42 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__42(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__42); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__43 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__43(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__43); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__44 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__44(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__44); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__45 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__45(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__45); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__46 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__46(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__46); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__47 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__47(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__47); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__48 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__48(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__48); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__49 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__49(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__49); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__50 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__50(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424____closed__50); +res = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_2424_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); l_Lean_Parser_Term_open___elambda__1___closed__1 = _init_l_Lean_Parser_Term_open___elambda__1___closed__1(); diff --git a/stage0/stdlib/Lean/Parser/Do.c b/stage0/stdlib/Lean/Parser/Do.c index 1f0fe6a2d86..f406437990e 100644 --- a/stage0/stdlib/Lean/Parser/Do.c +++ b/stage0/stdlib/Lean/Parser/Do.c @@ -256,6 +256,7 @@ static lean_object* l_Lean_Parser_Term_doExpr___closed__8; static lean_object* l_Lean_Parser_Term_doMatch_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Term_doPatDecl___elambda__1___closed__5; static lean_object* l_Lean_Parser_Term_doForDecl___elambda__1___closed__23; +static lean_object* l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__19; static lean_object* l_Lean_Parser_Term_doIf___closed__6; static lean_object* l_Lean_Parser_Term_doIfProp_formatter___closed__1; static lean_object* l_Lean_Parser_Term_doReturn_parenthesizer___closed__4; @@ -669,6 +670,7 @@ static lean_object* l_Lean_Parser_Term_doIf___closed__5; static lean_object* l_Lean_Parser_Term_doHave___elambda__1___closed__10; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_doLetRec_parenthesizer(lean_object*); lean_object* l_Lean_Parser_Term_letRecDecls_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_doIfProp___elambda__1___closed__2; static lean_object* l_Lean_Parser_Term_doContinue_formatter___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Term_doReassignArrow___elambda__1(lean_object*, lean_object*); @@ -1160,6 +1162,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_termReturn_declRange___close static lean_object* l_Lean_Parser_Term_doElem_quot_formatter___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Term_doIfLet___elambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_doSeqIndent_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Term_doLetRec___elambda__1___closed__10; static lean_object* l_Lean_Parser_Term_termUnless___elambda__1___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_doDbgTrace_declRange___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Term_termUnless___elambda__1(lean_object*, lean_object*); @@ -1352,7 +1355,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_doIfLetPure___elambda__1(lean_object static lean_object* l_Lean_Parser_Term_doExpr___elambda__1___closed__15; static lean_object* l___regBuiltin_Lean_Parser_Term_doUnless_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_doSeqItem___elambda__1___closed__1; -lean_object* l_Lean_Parser_registerAlias(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Parser_registerAlias(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_doElem_quot_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Term_doUnless___elambda__1___closed__13; static lean_object* l_Lean_Parser_Term_doDbgTrace_parenthesizer___closed__3; @@ -1477,7 +1480,6 @@ static lean_object* l_Lean_Parser_Term_doCatch_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_doMatch_declRange___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_doIf_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_doSeqIndent___elambda__1___closed__11; -extern lean_object* l_Lean_groupKind; static lean_object* l_Lean_Parser_Term_leftArrow___elambda__1___closed__9; static lean_object* l_Lean_Parser_Term_doLet_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Term_doSeqIndent___closed__6; @@ -1673,6 +1675,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_doReturn_formatter(lean_object*, lea LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_termTry_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Term_doMatch___closed__12; LEAN_EXPORT lean_object* l_Lean_Parser_Term_termUnless; +static lean_object* l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__20; static lean_object* l_Lean_Parser_Term_doIf___elambda__1___closed__31; static lean_object* l_Lean_Parser_Term_doLet___elambda__1___closed__4; static lean_object* l_Lean_Parser_Term_doFor___elambda__1___closed__19; @@ -2205,7 +2208,6 @@ static lean_object* l_Lean_Parser_Term_doIf___elambda__1___closed__7; static lean_object* l_Lean_Parser_Term_doUnless_formatter___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_doTry_declRange___closed__6; static lean_object* l_Lean_Parser_Term_doReassignArrow___closed__6; -uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_doLetElse___elambda__1___closed__23; LEAN_EXPORT lean_object* l_Lean_Parser_Term_doReassign_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_doNested___elambda__1___closed__2; @@ -2227,6 +2229,7 @@ static lean_object* l_Lean_Parser_Term_doForDecl___elambda__1___closed__11; static lean_object* l_Lean_Parser_Term_notFollowedByRedefinedTermToken___elambda__1___closed__4; lean_object* l_Lean_PrettyPrinter_Formatter_notFollowedBy_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_doHave_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Term_doLetRec___elambda__1___closed__9; static lean_object* l_Lean_Parser_Term_doIdDecl_parenthesizer___closed__10; LEAN_EXPORT lean_object* l_Lean_Parser_Term_do___elambda__1(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_doIf_declRange___closed__5; @@ -2877,7 +2880,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -2895,7 +2898,7 @@ lean_inc(x_1); x_18 = l_Lean_Parser_Term_leftArrow___elambda__1(x_1, x_12); x_19 = lean_ctor_get(x_18, 4); lean_inc(x_19); -x_20 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_14); +x_20 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_14); lean_dec(x_19); if (x_20 == 0) { @@ -2904,7 +2907,7 @@ x_21 = l_Lean_Parser_Term_liftMethod___elambda__1___closed__8; x_22 = l_Lean_Parser_ParserState_mkNode(x_18, x_21, x_17); x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); -x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_14); +x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_14); lean_dec(x_23); if (x_24 == 0) { @@ -2930,7 +2933,7 @@ x_29 = l_Lean_Parser_Term_liftMethod___elambda__1___closed__8; x_30 = l_Lean_Parser_ParserState_mkNode(x_28, x_29, x_17); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_14); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_14); lean_dec(x_31); if (x_32 == 0) { @@ -3618,7 +3621,7 @@ x_28 = l_Lean_Parser_checkPrecFn(x_27, x_1, x_2); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); x_30 = lean_box(0); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_30); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_30); lean_dec(x_29); if (x_31 == 0) { @@ -3639,7 +3642,7 @@ lean_inc(x_1); x_36 = l_Lean_Parser_categoryParser___elambda__1(x_34, x_35, x_1, x_28); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_30); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_30); lean_dec(x_37); if (x_38 == 0) { @@ -3648,7 +3651,7 @@ lean_dec(x_8); x_39 = l_Lean_Parser_ParserState_mkNode(x_36, x_13, x_33); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_30); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_30); lean_dec(x_40); if (x_41 == 0) { @@ -3671,7 +3674,7 @@ x_43 = lean_apply_2(x_8, x_1, x_36); x_44 = l_Lean_Parser_ParserState_mkNode(x_43, x_13, x_33); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_30); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_30); lean_dec(x_45); if (x_46 == 0) { @@ -3995,7 +3998,7 @@ x_26 = l_Lean_Parser_checkPrecFn(x_25, x_1, x_2); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); x_28 = lean_box(0); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_28); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_28); lean_dec(x_27); if (x_29 == 0) { @@ -4034,7 +4037,7 @@ x_36 = lean_apply_2(x_32, x_35, x_26); x_37 = l_Lean_Parser_ParserState_mkNode(x_36, x_5, x_31); x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_28); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_28); lean_dec(x_38); if (x_39 == 0) { @@ -4393,7 +4396,7 @@ x_30 = l_Lean_Parser_checkPrecFn(x_29, x_1, x_2); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); x_32 = lean_box(0); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_32); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_32); lean_dec(x_31); if (x_33 == 0) { @@ -4443,7 +4446,7 @@ lean_object* x_44; lean_object* x_45; uint8_t x_46; x_44 = l_Lean_Parser_ParserState_mkNode(x_43, x_9, x_35); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_32); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_32); lean_dec(x_45); if (x_46 == 0) { @@ -4463,7 +4466,7 @@ return x_47; lean_object* x_50; uint8_t x_51; x_50 = lean_ctor_get(x_49, 4); lean_inc(x_50); -x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_32); +x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_32); lean_dec(x_50); if (x_51 == 0) { @@ -4494,7 +4497,7 @@ lean_ctor_set_uint8(x_53, sizeof(void*)*7, x_22); x_54 = lean_apply_2(x_52, x_53, x_49); x_55 = lean_ctor_get(x_54, 4); lean_inc(x_55); -x_56 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_55, x_32); +x_56 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_55, x_32); lean_dec(x_55); if (x_56 == 0) { @@ -5806,22 +5809,44 @@ return x_2; static lean_object* _init_l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__4() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__5() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__4; +x_2 = 1; +x_3 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__6() { +_start: +{ lean_object* x_1; x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_doSeq_formatter), 5, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__5() { +static lean_object* _init_l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__4; +x_1 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__6; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__6() { +static lean_object* _init_l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__8() { _start: { lean_object* x_1; @@ -5829,7 +5854,7 @@ x_1 = l_Lean_PrettyPrinter_Formatter_formatterAliasesRef; return x_1; } } -static lean_object* _init_l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__7() { +static lean_object* _init_l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__9() { _start: { lean_object* x_1; @@ -5837,17 +5862,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_doSeq_parenthesizer), 5, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__8() { +static lean_object* _init_l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__7; +x_1 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__9; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__9() { +static lean_object* _init_l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__11() { _start: { lean_object* x_1; @@ -5855,7 +5880,7 @@ x_1 = l_Lean_PrettyPrinter_Parenthesizer_parenthesizerAliasesRef; return x_1; } } -static lean_object* _init_l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__10() { +static lean_object* _init_l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__12() { _start: { lean_object* x_1; @@ -5863,17 +5888,17 @@ x_1 = lean_mk_string_from_bytes("termBeforeDo", 12); return x_1; } } -static lean_object* _init_l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__11() { +static lean_object* _init_l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__10; +x_2 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__12; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__12() { +static lean_object* _init_l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__14() { _start: { lean_object* x_1; lean_object* x_2; @@ -5883,27 +5908,27 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__13() { +static lean_object* _init_l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_liftMethod___elambda__1___closed__6; -x_2 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__10; +x_2 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__12; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__14() { +static lean_object* _init_l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__16() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__13; +x_1 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__15; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__15() { +static lean_object* _init_l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__17() { _start: { lean_object* x_1; @@ -5911,17 +5936,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_termBeforeDo_formatter), 5, return x_1; } } -static lean_object* _init_l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__16() { +static lean_object* _init_l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__18() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__15; +x_1 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__17; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__17() { +static lean_object* _init_l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__19() { _start: { lean_object* x_1; @@ -5929,11 +5954,11 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_termBeforeDo_parenthesizer), return x_1; } } -static lean_object* _init_l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__18() { +static lean_object* _init_l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__20() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__17; +x_1 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__19; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; @@ -5942,169 +5967,170 @@ return x_2; LEAN_EXPORT lean_object* l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157_(lean_object* x_1) { _start: { -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_2 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__1; x_3 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__2; x_4 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__3; -x_5 = l_Lean_Parser_registerAlias(x_2, x_3, x_4, x_1); -if (lean_obj_tag(x_5) == 0) +x_5 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__5; +x_6 = l_Lean_Parser_registerAlias(x_2, x_3, x_4, x_5, x_1); +if (lean_obj_tag(x_6) == 0) { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_6 = lean_ctor_get(x_5, 1); -lean_inc(x_6); -lean_dec(x_5); -x_7 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__6; -x_8 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__5; -x_9 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_2, x_8, x_6); -if (lean_obj_tag(x_9) == 0) +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_7 = lean_ctor_get(x_6, 1); +lean_inc(x_7); +lean_dec(x_6); +x_8 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__8; +x_9 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__7; +x_10 = l_Lean_Parser_registerAliasCore___rarg(x_8, x_2, x_9, x_7); +if (lean_obj_tag(x_10) == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_10 = lean_ctor_get(x_9, 1); -lean_inc(x_10); -lean_dec(x_9); -x_11 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__9; -x_12 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__8; -x_13 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_2, x_12, x_10); -if (lean_obj_tag(x_13) == 0) +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +lean_dec(x_10); +x_12 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__11; +x_13 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__10; +x_14 = l_Lean_Parser_registerAliasCore___rarg(x_12, x_2, x_13, x_11); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -lean_dec(x_13); -x_15 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__11; -x_16 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__12; +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +lean_dec(x_14); +x_16 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__13; x_17 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__14; -x_18 = l_Lean_Parser_registerAlias(x_15, x_16, x_17, x_14); -if (lean_obj_tag(x_18) == 0) +x_18 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__16; +x_19 = l_Lean_Parser_registerAlias(x_16, x_17, x_18, x_5, x_15); +if (lean_obj_tag(x_19) == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -lean_dec(x_18); -x_20 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__16; -x_21 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_15, x_20, x_19); -if (lean_obj_tag(x_21) == 0) +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_19, 1); +lean_inc(x_20); +lean_dec(x_19); +x_21 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__18; +x_22 = l_Lean_Parser_registerAliasCore___rarg(x_8, x_16, x_21, x_20); +if (lean_obj_tag(x_22) == 0) { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_21, 1); -lean_inc(x_22); -lean_dec(x_21); -x_23 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__18; -x_24 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_15, x_23, x_22); -return x_24; +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +lean_dec(x_22); +x_24 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__20; +x_25 = l_Lean_Parser_registerAliasCore___rarg(x_12, x_16, x_24, x_23); +return x_25; } else { -uint8_t x_25; -x_25 = !lean_is_exclusive(x_21); -if (x_25 == 0) +uint8_t x_26; +x_26 = !lean_is_exclusive(x_22); +if (x_26 == 0) { -return x_21; +return x_22; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_21, 0); -x_27 = lean_ctor_get(x_21, 1); +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_22, 0); +x_28 = lean_ctor_get(x_22, 1); +lean_inc(x_28); lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_21); -x_28 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -return x_28; +lean_dec(x_22); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; } } } else { -uint8_t x_29; -x_29 = !lean_is_exclusive(x_18); -if (x_29 == 0) +uint8_t x_30; +x_30 = !lean_is_exclusive(x_19); +if (x_30 == 0) { -return x_18; +return x_19; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_18, 0); -x_31 = lean_ctor_get(x_18, 1); +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_19, 0); +x_32 = lean_ctor_get(x_19, 1); +lean_inc(x_32); lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_18); -x_32 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -return x_32; +lean_dec(x_19); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; } } } else { -uint8_t x_33; -x_33 = !lean_is_exclusive(x_13); -if (x_33 == 0) +uint8_t x_34; +x_34 = !lean_is_exclusive(x_14); +if (x_34 == 0) { -return x_13; +return x_14; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_13, 0); -x_35 = lean_ctor_get(x_13, 1); +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_14, 0); +x_36 = lean_ctor_get(x_14, 1); +lean_inc(x_36); lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_13); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -return x_36; +lean_dec(x_14); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +return x_37; } } } else { -uint8_t x_37; -x_37 = !lean_is_exclusive(x_9); -if (x_37 == 0) +uint8_t x_38; +x_38 = !lean_is_exclusive(x_10); +if (x_38 == 0) { -return x_9; +return x_10; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_9, 0); -x_39 = lean_ctor_get(x_9, 1); +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_10, 0); +x_40 = lean_ctor_get(x_10, 1); +lean_inc(x_40); lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_9); -x_40 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -return x_40; +lean_dec(x_10); +x_41 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +return x_41; } } } else { -uint8_t x_41; -x_41 = !lean_is_exclusive(x_5); -if (x_41 == 0) +uint8_t x_42; +x_42 = !lean_is_exclusive(x_6); +if (x_42 == 0) { -return x_5; +return x_6; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_5, 0); -x_43 = lean_ctor_get(x_5, 1); +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_6, 0); +x_44 = lean_ctor_get(x_6, 1); +lean_inc(x_44); lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_5); -x_44 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_44, 1, x_43); -return x_44; +lean_dec(x_6); +x_45 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +return x_45; } } } @@ -6876,7 +6902,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -6921,7 +6947,7 @@ goto block_49; lean_object* x_28; uint8_t x_29; x_28 = lean_ctor_get(x_27, 4); lean_inc(x_28); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_16); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_16); lean_dec(x_28); if (x_29 == 0) { @@ -6931,7 +6957,7 @@ x_30 = l_Lean_Parser_Term_doLet___elambda__1___closed__2; x_31 = l_Lean_Parser_ParserState_mkNode(x_27, x_30, x_19); x_32 = lean_ctor_get(x_31, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_16); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_16); lean_dec(x_32); if (x_33 == 0) { @@ -6953,7 +6979,7 @@ lean_inc(x_1); x_35 = lean_apply_2(x_4, x_1, x_27); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_16); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_16); lean_dec(x_36); if (x_37 == 0) { @@ -6962,7 +6988,7 @@ x_38 = l_Lean_Parser_Term_doLet___elambda__1___closed__2; x_39 = l_Lean_Parser_ParserState_mkNode(x_35, x_38, x_19); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_16); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_16); lean_dec(x_40); if (x_41 == 0) { @@ -6986,7 +7012,7 @@ x_44 = l_Lean_Parser_Term_doLet___elambda__1___closed__2; x_45 = l_Lean_Parser_ParserState_mkNode(x_43, x_44, x_19); x_46 = lean_ctor_get(x_45, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_16); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_16); lean_dec(x_46); if (x_47 == 0) { @@ -7779,7 +7805,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -7825,7 +7851,7 @@ x_26 = l_Lean_Parser_Term_doLetElse___elambda__1___closed__2; x_27 = l_Lean_Parser_ParserState_mkNode(x_25, x_26, x_17); x_28 = lean_ctor_get(x_27, 4); lean_inc(x_28); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_14); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_14); lean_dec(x_28); if (x_29 == 0) { @@ -7845,7 +7871,7 @@ return x_30; lean_object* x_33; uint8_t x_34; x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -7856,7 +7882,7 @@ x_35 = l_Lean_Parser_Term_doLetElse___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_32, x_35, x_17); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); lean_dec(x_37); if (x_38 == 0) { @@ -7880,7 +7906,7 @@ lean_inc(x_1); x_42 = l_Lean_Parser_categoryParser___elambda__1(x_40, x_41, x_1, x_32); x_43 = lean_ctor_get(x_42, 4); lean_inc(x_43); -x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_14); +x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_14); lean_dec(x_43); if (x_44 == 0) { @@ -7891,7 +7917,7 @@ x_45 = l_Lean_Parser_Term_doLetElse___elambda__1___closed__2; x_46 = l_Lean_Parser_ParserState_mkNode(x_42, x_45, x_17); x_47 = lean_ctor_get(x_46, 4); lean_inc(x_47); -x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_14); +x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_14); lean_dec(x_47); if (x_48 == 0) { @@ -7936,7 +7962,7 @@ goto block_92; lean_object* x_57; uint8_t x_58; x_57 = lean_ctor_get(x_56, 4); lean_inc(x_57); -x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_14); +x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_14); lean_dec(x_57); if (x_58 == 0) { @@ -7952,7 +7978,7 @@ lean_inc(x_1); x_59 = l_Lean_Parser_categoryParser___elambda__1(x_40, x_41, x_1, x_56); x_60 = lean_ctor_get(x_59, 4); lean_inc(x_60); -x_61 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_60, x_14); +x_61 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_60, x_14); lean_dec(x_60); if (x_61 == 0) { @@ -8016,7 +8042,7 @@ goto block_81; lean_object* x_64; uint8_t x_65; x_64 = lean_ctor_get(x_63, 4); lean_inc(x_64); -x_65 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_64, x_14); +x_65 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_64, x_14); lean_dec(x_64); if (x_65 == 0) { @@ -8042,7 +8068,7 @@ if (x_71 == 0) lean_object* x_72; uint8_t x_73; x_72 = lean_ctor_get(x_68, 4); lean_inc(x_72); -x_73 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_72, x_14); +x_73 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_72, x_14); lean_dec(x_72); if (x_73 == 0) { @@ -8066,7 +8092,7 @@ lean_inc(x_1); x_76 = l_Lean_Parser_tokenAntiquotFn(x_1, x_68); x_77 = lean_ctor_get(x_76, 4); lean_inc(x_77); -x_78 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_77, x_14); +x_78 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_77, x_14); lean_dec(x_77); if (x_78 == 0) { @@ -8803,7 +8829,7 @@ static lean_object* _init_l_Lean_Parser_Term_doLetRec___elambda__1___closed__6() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_letRecDecls___elambda__1), 2, 0); +x_1 = lean_mk_string_from_bytes("group", 5); return x_1; } } @@ -8811,17 +8837,35 @@ static lean_object* _init_l_Lean_Parser_Term_doLetRec___elambda__1___closed__7() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_Term_doLetRec___elambda__1___closed__6; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Term_doLetRec___elambda__1___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_letRecDecls___elambda__1), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Term_doLetRec___elambda__1___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_leftArrow___elambda__1___closed__5; x_2 = l_Lean_Parser_Term_doLetRec___elambda__1___closed__5; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_doLetRec___elambda__1___closed__8() { +static lean_object* _init_l_Lean_Parser_Term_doLetRec___elambda__1___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_doLetRec___elambda__1___closed__7; +x_1 = l_Lean_Parser_Term_doLetRec___elambda__1___closed__9; x_2 = l_Lean_Parser_Term_leftArrow___elambda__1___closed__5; x_3 = lean_string_append(x_1, x_2); return x_3; @@ -8838,11 +8882,11 @@ x_5 = l_Lean_Parser_Term_doLet___elambda__1___closed__6; x_6 = lean_alloc_closure((void*)(l_Lean_Parser_andthenFn), 4, 2); lean_closure_set(x_6, 0, x_5); lean_closure_set(x_6, 1, x_4); -x_7 = l_Lean_groupKind; +x_7 = l_Lean_Parser_Term_doLetRec___elambda__1___closed__7; x_8 = lean_alloc_closure((void*)(l_Lean_Parser_nodeFn), 4, 2); lean_closure_set(x_8, 0, x_7); lean_closure_set(x_8, 1, x_6); -x_9 = l_Lean_Parser_Term_doLetRec___elambda__1___closed__6; +x_9 = l_Lean_Parser_Term_doLetRec___elambda__1___closed__8; x_10 = lean_alloc_closure((void*)(l_Lean_Parser_andthenFn), 4, 2); lean_closure_set(x_10, 0, x_8); lean_closure_set(x_10, 1, x_9); @@ -8882,7 +8926,7 @@ x_26 = l_Lean_Parser_checkPrecFn(x_25, x_1, x_2); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); x_28 = lean_box(0); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_28); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_28); lean_dec(x_27); if (x_29 == 0) { @@ -8912,7 +8956,7 @@ if (x_38 == 0) lean_object* x_53; uint8_t x_54; x_53 = lean_ctor_get(x_34, 4); lean_inc(x_53); -x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_28); +x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_28); lean_dec(x_53); if (x_54 == 0) { @@ -8923,7 +8967,7 @@ goto block_52; else { lean_object* x_55; lean_object* x_56; lean_object* x_57; uint32_t x_58; uint8_t x_59; -x_55 = l_Lean_Parser_Term_doLetRec___elambda__1___closed__8; +x_55 = l_Lean_Parser_Term_doLetRec___elambda__1___closed__10; lean_inc(x_1); x_56 = l_Lean_Parser_nonReservedSymbolFnAux(x_3, x_55, x_1, x_34); x_57 = lean_ctor_get(x_56, 2); @@ -8954,7 +8998,7 @@ lean_inc(x_1); x_61 = l_Lean_Parser_tokenAntiquotFn(x_1, x_34); x_62 = lean_ctor_get(x_61, 4); lean_inc(x_62); -x_63 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_62, x_28); +x_63 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_62, x_28); lean_dec(x_62); if (x_63 == 0) { @@ -8965,7 +9009,7 @@ goto block_52; else { lean_object* x_64; lean_object* x_65; lean_object* x_66; uint32_t x_67; uint8_t x_68; -x_64 = l_Lean_Parser_Term_doLetRec___elambda__1___closed__8; +x_64 = l_Lean_Parser_Term_doLetRec___elambda__1___closed__10; lean_inc(x_1); x_65 = l_Lean_Parser_nonReservedSymbolFnAux(x_3, x_64, x_1, x_61); x_66 = lean_ctor_get(x_65, 2); @@ -8996,7 +9040,7 @@ lean_inc(x_31); x_40 = l_Lean_Parser_ParserState_mkNode(x_39, x_7, x_31); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_28); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_28); lean_dec(x_41); if (x_42 == 0) { @@ -9004,7 +9048,7 @@ lean_object* x_43; lean_object* x_44; uint8_t x_45; x_43 = l_Lean_Parser_ParserState_mkNode(x_40, x_11, x_31); x_44 = lean_ctor_get(x_43, 4); lean_inc(x_44); -x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_28); +x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_28); lean_dec(x_44); if (x_45 == 0) { @@ -9027,7 +9071,7 @@ x_47 = l_Lean_Parser_Term_letRecDecls___elambda__1(x_1, x_40); x_48 = l_Lean_Parser_ParserState_mkNode(x_47, x_11, x_31); x_49 = lean_ctor_get(x_48, 4); lean_inc(x_49); -x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_28); +x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_28); lean_dec(x_49); if (x_50 == 0) { @@ -9079,7 +9123,7 @@ static lean_object* _init_l_Lean_Parser_Term_doLetRec___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_groupKind; +x_1 = l_Lean_Parser_Term_doLetRec___elambda__1___closed__7; x_2 = l_Lean_Parser_Term_doLetRec___closed__2; x_3 = l_Lean_Parser_nodeInfo(x_1, x_2); return x_3; @@ -9698,7 +9742,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -9717,7 +9761,7 @@ lean_inc(x_1); x_19 = l_Lean_Parser_atomicFn(x_18, x_1, x_12); x_20 = lean_ctor_get(x_19, 4); lean_inc(x_20); -x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_20, x_14); +x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_20, x_14); lean_dec(x_20); if (x_21 == 0) { @@ -9726,7 +9770,7 @@ x_22 = l_Lean_Parser_Term_doIdDecl___elambda__1___closed__2; x_23 = l_Lean_Parser_ParserState_mkNode(x_19, x_22, x_17); x_24 = lean_ctor_get(x_23, 4); lean_inc(x_24); -x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_14); +x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_14); lean_dec(x_24); if (x_25 == 0) { @@ -9752,7 +9796,7 @@ x_30 = l_Lean_Parser_Term_doIdDecl___elambda__1___closed__2; x_31 = l_Lean_Parser_ParserState_mkNode(x_29, x_30, x_17); x_32 = lean_ctor_get(x_31, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_14); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_14); lean_dec(x_32); if (x_33 == 0) { @@ -10085,7 +10129,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -10105,7 +10149,7 @@ lean_inc(x_1); x_21 = l_Lean_Parser_atomicFn(x_20, x_1, x_14); x_22 = lean_ctor_get(x_21, 4); lean_inc(x_22); -x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_22, x_16); +x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_22, x_16); lean_dec(x_22); if (x_23 == 0) { @@ -10115,7 +10159,7 @@ x_24 = l_Lean_Parser_Term_doPatDecl___elambda__1___closed__2; x_25 = l_Lean_Parser_ParserState_mkNode(x_21, x_24, x_19); x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_16); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_16); lean_dec(x_26); if (x_27 == 0) { @@ -10139,7 +10183,7 @@ lean_inc(x_1); x_31 = l_Lean_Parser_categoryParser___elambda__1(x_29, x_30, x_1, x_21); x_32 = lean_ctor_get(x_31, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_16); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_16); lean_dec(x_32); if (x_33 == 0) { @@ -10149,7 +10193,7 @@ x_34 = l_Lean_Parser_Term_doPatDecl___elambda__1___closed__2; x_35 = l_Lean_Parser_ParserState_mkNode(x_31, x_34, x_19); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_16); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_16); lean_dec(x_36); if (x_37 == 0) { @@ -10173,7 +10217,7 @@ x_40 = l_Lean_Parser_Term_doPatDecl___elambda__1___closed__2; x_41 = l_Lean_Parser_ParserState_mkNode(x_39, x_40, x_19); x_42 = lean_ctor_get(x_41, 4); lean_inc(x_42); -x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_16); +x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_16); lean_dec(x_42); if (x_43 == 0) { @@ -10346,7 +10390,7 @@ if (x_18 == 0) lean_object* x_20; uint8_t x_21; x_20 = lean_ctor_get(x_13, 4); lean_inc(x_20); -x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_20, x_19); +x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_20, x_19); lean_dec(x_20); if (x_21 == 0) { @@ -10361,7 +10405,7 @@ lean_inc(x_3); x_22 = lean_apply_2(x_2, x_3, x_13); x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); -x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_19); +x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_19); lean_dec(x_23); if (x_24 == 0) { @@ -10386,7 +10430,7 @@ lean_inc(x_3); x_29 = l_Lean_Parser_tokenAntiquotFn(x_3, x_13); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_19); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_19); lean_dec(x_30); if (x_31 == 0) { @@ -10401,7 +10445,7 @@ lean_inc(x_3); x_32 = lean_apply_2(x_2, x_3, x_29); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_19); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_19); lean_dec(x_33); if (x_34 == 0) { @@ -10472,7 +10516,7 @@ if (x_57 == 0) lean_object* x_59; uint8_t x_60; x_59 = lean_ctor_get(x_52, 4); lean_inc(x_59); -x_60 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_59, x_58); +x_60 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_59, x_58); lean_dec(x_59); if (x_60 == 0) { @@ -10487,7 +10531,7 @@ lean_inc(x_48); x_61 = lean_apply_2(x_2, x_48, x_52); x_62 = lean_ctor_get(x_61, 4); lean_inc(x_62); -x_63 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_62, x_58); +x_63 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_62, x_58); lean_dec(x_62); if (x_63 == 0) { @@ -10512,7 +10556,7 @@ lean_inc(x_48); x_68 = l_Lean_Parser_tokenAntiquotFn(x_48, x_52); x_69 = lean_ctor_get(x_68, 4); lean_inc(x_69); -x_70 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_69, x_58); +x_70 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_69, x_58); lean_dec(x_69); if (x_70 == 0) { @@ -10527,7 +10571,7 @@ lean_inc(x_48); x_71 = lean_apply_2(x_2, x_48, x_68); x_72 = lean_ctor_get(x_71, 4); lean_inc(x_72); -x_73 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_72, x_58); +x_73 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_72, x_58); lean_dec(x_72); if (x_73 == 0) { @@ -10636,7 +10680,7 @@ x_28 = l_Lean_Parser_checkPrecFn(x_27, x_1, x_2); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); x_30 = lean_box(0); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_30); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_30); lean_dec(x_29); if (x_31 == 0) { @@ -10686,7 +10730,7 @@ if (x_42 == 0) lean_object* x_49; uint8_t x_50; x_49 = lean_ctor_get(x_38, 4); lean_inc(x_49); -x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_30); +x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_30); lean_dec(x_49); if (x_50 == 0) { @@ -10702,7 +10746,7 @@ lean_inc(x_36); x_51 = lean_apply_2(x_4, x_36, x_38); x_52 = lean_ctor_get(x_51, 4); lean_inc(x_52); -x_53 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_52, x_30); +x_53 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_52, x_30); lean_dec(x_52); if (x_53 == 0) { @@ -10729,7 +10773,7 @@ lean_inc(x_36); x_58 = l_Lean_Parser_tokenAntiquotFn(x_36, x_38); x_59 = lean_ctor_get(x_58, 4); lean_inc(x_59); -x_60 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_59, x_30); +x_60 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_59, x_30); lean_dec(x_59); if (x_60 == 0) { @@ -10745,7 +10789,7 @@ lean_inc(x_36); x_61 = lean_apply_2(x_4, x_36, x_58); x_62 = lean_ctor_get(x_61, 4); lean_inc(x_62); -x_63 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_62, x_30); +x_63 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_62, x_30); lean_dec(x_62); if (x_63 == 0) { @@ -10771,7 +10815,7 @@ lean_object* x_44; lean_object* x_45; uint8_t x_46; x_44 = l_Lean_Parser_ParserState_mkNode(x_43, x_7, x_33); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_30); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_30); lean_dec(x_45); if (x_46 == 0) { @@ -11937,7 +11981,7 @@ x_6 = l_Lean_Parser_atomicFn(x_5, x_1, x_2); x_7 = lean_ctor_get(x_6, 4); lean_inc(x_7); x_8 = lean_box(0); -x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); +x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); lean_dec(x_7); if (x_9 == 0) { @@ -12171,7 +12215,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -12189,7 +12233,7 @@ lean_inc(x_1); x_18 = l_Lean_Parser_Term_notFollowedByRedefinedTermToken___elambda__1(x_1, x_12); x_19 = lean_ctor_get(x_18, 4); lean_inc(x_19); -x_20 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_14); +x_20 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_14); lean_dec(x_19); if (x_20 == 0) { @@ -12198,7 +12242,7 @@ x_21 = l_Lean_Parser_Term_doReassign___elambda__1___closed__2; x_22 = l_Lean_Parser_ParserState_mkNode(x_18, x_21, x_17); x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); -x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_14); +x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_14); lean_dec(x_23); if (x_24 == 0) { @@ -12225,7 +12269,7 @@ x_30 = l_Lean_Parser_Term_doReassign___elambda__1___closed__2; x_31 = l_Lean_Parser_ParserState_mkNode(x_29, x_30, x_17); x_32 = lean_ctor_get(x_31, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_14); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_14); lean_dec(x_32); if (x_33 == 0) { @@ -13162,7 +13206,7 @@ x_18 = l_Lean_Parser_checkPrecFn(x_17, x_1, x_2); x_19 = lean_ctor_get(x_18, 4); lean_inc(x_19); x_20 = lean_box(0); -x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_20); +x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_20); lean_dec(x_19); if (x_21 == 0) { @@ -13186,7 +13230,7 @@ lean_inc(x_1); x_24 = l_Lean_Parser_Term_notFollowedByRedefinedTermToken___elambda__1(x_1, x_18); x_25 = lean_ctor_get(x_24, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_20); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_20); lean_dec(x_25); if (x_26 == 0) { @@ -13201,7 +13245,7 @@ x_27 = l_Lean_Parser_Term_doReassignArrow___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_24, x_27, x_23); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_20); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_20); lean_dec(x_29); if (x_30 == 0) { @@ -13240,7 +13284,7 @@ x_39 = l_Lean_Parser_Term_doReassignArrow___elambda__1___closed__2; x_40 = l_Lean_Parser_ParserState_mkNode(x_38, x_39, x_23); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_20); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_20); lean_dec(x_41); if (x_42 == 0) { @@ -13818,7 +13862,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -13849,7 +13893,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_20, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -13858,7 +13902,7 @@ x_27 = l_Lean_Parser_Term_doHave___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_20, x_27, x_17); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -13882,7 +13926,7 @@ x_33 = l_Lean_Parser_Term_doHave___elambda__1___closed__2; x_34 = l_Lean_Parser_ParserState_mkNode(x_32, x_33, x_17); x_35 = lean_ctor_get(x_34, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_14); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_14); lean_dec(x_35); if (x_36 == 0) { @@ -13905,7 +13949,7 @@ lean_inc(x_1); x_38 = l_Lean_Parser_tokenAntiquotFn(x_1, x_20); x_39 = lean_ctor_get(x_38, 4); lean_inc(x_39); -x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); +x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); lean_dec(x_39); if (x_40 == 0) { @@ -13914,7 +13958,7 @@ x_41 = l_Lean_Parser_Term_doHave___elambda__1___closed__2; x_42 = l_Lean_Parser_ParserState_mkNode(x_38, x_41, x_17); x_43 = lean_ctor_get(x_42, 4); lean_inc(x_43); -x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_14); +x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_14); lean_dec(x_43); if (x_44 == 0) { @@ -13938,7 +13982,7 @@ x_47 = l_Lean_Parser_Term_doHave___elambda__1___closed__2; x_48 = l_Lean_Parser_ParserState_mkNode(x_46, x_47, x_17); x_49 = lean_ctor_get(x_48, 4); lean_inc(x_49); -x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_14); +x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_14); lean_dec(x_49); if (x_50 == 0) { @@ -14435,7 +14479,7 @@ goto block_48; lean_object* x_21; uint8_t x_22; x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); -x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_19); +x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_19); lean_dec(x_21); if (x_22 == 0) { @@ -14474,7 +14518,7 @@ x_30 = l_Lean_Parser_Term_elseIf___elambda__1___lambda__1___closed__1; x_31 = l_Lean_Parser_ParserState_mkError(x_20, x_30); x_32 = lean_ctor_get(x_31, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_19); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_19); lean_dec(x_32); if (x_33 == 0) { @@ -14600,7 +14644,7 @@ goto block_98; lean_object* x_71; uint8_t x_72; x_71 = lean_ctor_get(x_70, 4); lean_inc(x_71); -x_72 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_71, x_69); +x_72 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_71, x_69); lean_dec(x_71); if (x_72 == 0) { @@ -14639,7 +14683,7 @@ x_80 = l_Lean_Parser_Term_elseIf___elambda__1___lambda__1___closed__1; x_81 = l_Lean_Parser_ParserState_mkError(x_70, x_80); x_82 = lean_ctor_get(x_81, 4); lean_inc(x_82); -x_83 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_82, x_69); +x_83 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_82, x_69); lean_dec(x_82); if (x_83 == 0) { @@ -14746,7 +14790,7 @@ x_4 = l_Lean_Parser_Term_elseIf___elambda__1___closed__4; x_5 = lean_alloc_closure((void*)(l_Lean_Parser_Term_elseIf___elambda__1___lambda__1___boxed), 4, 2); lean_closure_set(x_5, 0, x_3); lean_closure_set(x_5, 1, x_4); -x_6 = l_Lean_groupKind; +x_6 = l_Lean_Parser_Term_doLetRec___elambda__1___closed__7; x_7 = lean_alloc_closure((void*)(l_Lean_Parser_nodeFn), 4, 2); lean_closure_set(x_7, 0, x_6); lean_closure_set(x_7, 1, x_5); @@ -14796,7 +14840,7 @@ static lean_object* _init_l_Lean_Parser_Term_elseIf___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_groupKind; +x_1 = l_Lean_Parser_Term_doLetRec___elambda__1___closed__7; x_2 = l_Lean_Parser_Term_elseIf___closed__4; x_3 = l_Lean_Parser_nodeInfo(x_1, x_2); return x_3; @@ -14945,7 +14989,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -14976,7 +15020,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_20, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -14985,7 +15029,7 @@ x_27 = l_Lean_Parser_Term_doIfLetPure___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_20, x_27, x_17); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -15011,7 +15055,7 @@ x_35 = l_Lean_Parser_Term_doIfLetPure___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_17); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); lean_dec(x_37); if (x_38 == 0) { @@ -15034,7 +15078,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_20); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); lean_dec(x_41); if (x_42 == 0) { @@ -15043,7 +15087,7 @@ x_43 = l_Lean_Parser_Term_doIfLetPure___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_17); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); lean_dec(x_45); if (x_46 == 0) { @@ -15069,7 +15113,7 @@ x_51 = l_Lean_Parser_Term_doIfLetPure___elambda__1___closed__2; x_52 = l_Lean_Parser_ParserState_mkNode(x_50, x_51, x_17); x_53 = lean_ctor_get(x_52, 4); lean_inc(x_53); -x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); +x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); lean_dec(x_53); if (x_54 == 0) { @@ -15332,7 +15376,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -15363,7 +15407,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_20, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -15372,7 +15416,7 @@ x_27 = l_Lean_Parser_Term_doIfLetBind___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_20, x_27, x_17); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -15398,7 +15442,7 @@ x_35 = l_Lean_Parser_Term_doIfLetBind___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_17); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); lean_dec(x_37); if (x_38 == 0) { @@ -15421,7 +15465,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_20); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); lean_dec(x_41); if (x_42 == 0) { @@ -15430,7 +15474,7 @@ x_43 = l_Lean_Parser_Term_doIfLetBind___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_17); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); lean_dec(x_45); if (x_46 == 0) { @@ -15456,7 +15500,7 @@ x_51 = l_Lean_Parser_Term_doIfLetBind___elambda__1___closed__2; x_52 = l_Lean_Parser_ParserState_mkNode(x_50, x_51, x_17); x_53 = lean_ctor_get(x_52, 4); lean_inc(x_53); -x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); +x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); lean_dec(x_53); if (x_54 == 0) { @@ -15704,7 +15748,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -15748,7 +15792,7 @@ goto block_52; lean_object* x_26; uint8_t x_27; x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); lean_dec(x_26); if (x_27 == 0) { @@ -15757,7 +15801,7 @@ x_28 = l_Lean_Parser_Term_doIfLet___elambda__1___closed__2; x_29 = l_Lean_Parser_ParserState_mkNode(x_25, x_28, x_17); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); lean_dec(x_30); if (x_31 == 0) { @@ -15781,7 +15825,7 @@ lean_inc(x_1); x_35 = l_Lean_Parser_categoryParser___elambda__1(x_33, x_34, x_1, x_25); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); lean_dec(x_36); if (x_37 == 0) { @@ -15790,7 +15834,7 @@ x_38 = l_Lean_Parser_Term_doIfLet___elambda__1___closed__2; x_39 = l_Lean_Parser_ParserState_mkNode(x_35, x_38, x_17); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); lean_dec(x_40); if (x_41 == 0) { @@ -15817,7 +15861,7 @@ x_47 = l_Lean_Parser_Term_doIfLet___elambda__1___closed__2; x_48 = l_Lean_Parser_ParserState_mkNode(x_46, x_47, x_17); x_49 = lean_ctor_get(x_48, 4); lean_inc(x_49); -x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_14); +x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_14); lean_dec(x_49); if (x_50 == 0) { @@ -16063,7 +16107,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -16082,7 +16126,7 @@ lean_inc(x_1); x_20 = lean_apply_2(x_4, x_1, x_14); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); -x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_16); +x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_16); lean_dec(x_21); if (x_22 == 0) { @@ -16091,7 +16135,7 @@ x_23 = l_Lean_Parser_Term_doIfProp___elambda__1___closed__2; x_24 = l_Lean_Parser_ParserState_mkNode(x_20, x_23, x_19); x_25 = lean_ctor_get(x_24, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_16); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_16); lean_dec(x_25); if (x_26 == 0) { @@ -16117,7 +16161,7 @@ x_31 = l_Lean_Parser_Term_doIfProp___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_19); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_16); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_16); lean_dec(x_33); if (x_34 == 0) { @@ -16413,7 +16457,7 @@ goto block_53; lean_object* x_23; uint8_t x_24; x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); -x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_21); +x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_21); lean_dec(x_23); if (x_24 == 0) { @@ -16430,7 +16474,7 @@ lean_inc(x_5); x_25 = l_Lean_Parser_Term_doIfCond___elambda__1(x_5, x_22); x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_21); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_21); lean_dec(x_26); if (x_27 == 0) { @@ -16458,7 +16502,7 @@ if (x_33 == 0) lean_object* x_34; uint8_t x_35; x_34 = lean_ctor_get(x_30, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_21); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_21); lean_dec(x_34); if (x_35 == 0) { @@ -16474,7 +16518,7 @@ lean_inc(x_5); x_36 = l_Lean_Parser_Term_doSeq___elambda__1(x_5, x_30); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_21); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_21); lean_dec(x_37); if (x_38 == 0) { @@ -16490,7 +16534,7 @@ lean_inc(x_5); x_39 = lean_apply_2(x_3, x_5, x_36); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_21); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_21); lean_dec(x_40); if (x_41 == 0) { @@ -16514,7 +16558,7 @@ lean_inc(x_5); x_43 = l_Lean_Parser_tokenAntiquotFn(x_5, x_30); x_44 = lean_ctor_get(x_43, 4); lean_inc(x_44); -x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_21); +x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_21); lean_dec(x_44); if (x_45 == 0) { @@ -16530,7 +16574,7 @@ lean_inc(x_5); x_46 = l_Lean_Parser_Term_doSeq___elambda__1(x_5, x_43); x_47 = lean_ctor_get(x_46, 4); lean_inc(x_47); -x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_21); +x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_21); lean_dec(x_47); if (x_48 == 0) { @@ -16546,7 +16590,7 @@ lean_inc(x_5); x_49 = lean_apply_2(x_3, x_5, x_46); x_50 = lean_ctor_get(x_49, 4); lean_inc(x_50); -x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_21); +x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_21); lean_dec(x_50); if (x_51 == 0) { @@ -16616,7 +16660,7 @@ goto block_101; lean_object* x_71; uint8_t x_72; x_71 = lean_ctor_get(x_70, 4); lean_inc(x_71); -x_72 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_71, x_69); +x_72 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_71, x_69); lean_dec(x_71); if (x_72 == 0) { @@ -16633,7 +16677,7 @@ lean_inc(x_5); x_73 = l_Lean_Parser_Term_doIfCond___elambda__1(x_5, x_70); x_74 = lean_ctor_get(x_73, 4); lean_inc(x_74); -x_75 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_74, x_69); +x_75 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_74, x_69); lean_dec(x_74); if (x_75 == 0) { @@ -16661,7 +16705,7 @@ if (x_81 == 0) lean_object* x_82; uint8_t x_83; x_82 = lean_ctor_get(x_78, 4); lean_inc(x_82); -x_83 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_82, x_69); +x_83 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_82, x_69); lean_dec(x_82); if (x_83 == 0) { @@ -16677,7 +16721,7 @@ lean_inc(x_5); x_84 = l_Lean_Parser_Term_doSeq___elambda__1(x_5, x_78); x_85 = lean_ctor_get(x_84, 4); lean_inc(x_85); -x_86 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_85, x_69); +x_86 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_85, x_69); lean_dec(x_85); if (x_86 == 0) { @@ -16693,7 +16737,7 @@ lean_inc(x_5); x_87 = lean_apply_2(x_3, x_5, x_84); x_88 = lean_ctor_get(x_87, 4); lean_inc(x_88); -x_89 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_88, x_69); +x_89 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_88, x_69); lean_dec(x_88); if (x_89 == 0) { @@ -16717,7 +16761,7 @@ lean_inc(x_5); x_91 = l_Lean_Parser_tokenAntiquotFn(x_5, x_78); x_92 = lean_ctor_get(x_91, 4); lean_inc(x_92); -x_93 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_92, x_69); +x_93 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_92, x_69); lean_dec(x_92); if (x_93 == 0) { @@ -16733,7 +16777,7 @@ lean_inc(x_5); x_94 = l_Lean_Parser_Term_doSeq___elambda__1(x_5, x_91); x_95 = lean_ctor_get(x_94, 4); lean_inc(x_95); -x_96 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_95, x_69); +x_96 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_95, x_69); lean_dec(x_95); if (x_96 == 0) { @@ -16749,7 +16793,7 @@ lean_inc(x_5); x_97 = lean_apply_2(x_3, x_5, x_94); x_98 = lean_ctor_get(x_97, 4); lean_inc(x_98); -x_99 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_98, x_69); +x_99 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_98, x_69); lean_dec(x_98); if (x_99 == 0) { @@ -16834,7 +16878,7 @@ goto block_154; lean_object* x_124; uint8_t x_125; x_124 = lean_ctor_get(x_123, 4); lean_inc(x_124); -x_125 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_124, x_122); +x_125 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_124, x_122); lean_dec(x_124); if (x_125 == 0) { @@ -16851,7 +16895,7 @@ lean_inc(x_112); x_126 = l_Lean_Parser_Term_doIfCond___elambda__1(x_112, x_123); x_127 = lean_ctor_get(x_126, 4); lean_inc(x_127); -x_128 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_127, x_122); +x_128 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_127, x_122); lean_dec(x_127); if (x_128 == 0) { @@ -16879,7 +16923,7 @@ if (x_134 == 0) lean_object* x_135; uint8_t x_136; x_135 = lean_ctor_get(x_131, 4); lean_inc(x_135); -x_136 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_135, x_122); +x_136 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_135, x_122); lean_dec(x_135); if (x_136 == 0) { @@ -16895,7 +16939,7 @@ lean_inc(x_112); x_137 = l_Lean_Parser_Term_doSeq___elambda__1(x_112, x_131); x_138 = lean_ctor_get(x_137, 4); lean_inc(x_138); -x_139 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_138, x_122); +x_139 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_138, x_122); lean_dec(x_138); if (x_139 == 0) { @@ -16911,7 +16955,7 @@ lean_inc(x_112); x_140 = lean_apply_2(x_3, x_112, x_137); x_141 = lean_ctor_get(x_140, 4); lean_inc(x_141); -x_142 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_141, x_122); +x_142 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_141, x_122); lean_dec(x_141); if (x_142 == 0) { @@ -16935,7 +16979,7 @@ lean_inc(x_112); x_144 = l_Lean_Parser_tokenAntiquotFn(x_112, x_131); x_145 = lean_ctor_get(x_144, 4); lean_inc(x_145); -x_146 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_145, x_122); +x_146 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_145, x_122); lean_dec(x_145); if (x_146 == 0) { @@ -16951,7 +16995,7 @@ lean_inc(x_112); x_147 = l_Lean_Parser_Term_doSeq___elambda__1(x_112, x_144); x_148 = lean_ctor_get(x_147, 4); lean_inc(x_148); -x_149 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_148, x_122); +x_149 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_148, x_122); lean_dec(x_148); if (x_149 == 0) { @@ -16967,7 +17011,7 @@ lean_inc(x_112); x_150 = lean_apply_2(x_3, x_112, x_147); x_151 = lean_ctor_get(x_150, 4); lean_inc(x_151); -x_152 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_151, x_122); +x_152 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_151, x_122); lean_dec(x_151); if (x_152 == 0) { @@ -17168,7 +17212,7 @@ static lean_object* _init_l_Lean_Parser_Term_doIf___elambda__1___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_groupKind; +x_1 = l_Lean_Parser_Term_doLetRec___elambda__1___closed__7; x_2 = l_Lean_Parser_Term_doIf___elambda__1___closed__16; x_3 = l_Lean_Parser_nodeInfo(x_1, x_2); return x_3; @@ -17178,7 +17222,7 @@ static lean_object* _init_l_Lean_Parser_Term_doIf___elambda__1___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_groupKind; +x_1 = l_Lean_Parser_Term_doLetRec___elambda__1___closed__7; x_2 = l_Lean_Parser_Term_doIf___elambda__1___closed__17; x_3 = lean_alloc_closure((void*)(l_Lean_Parser_nodeFn), 4, 2); lean_closure_set(x_3, 0, x_1); @@ -17429,7 +17473,7 @@ x_31 = l_Lean_Parser_checkPrecFn(x_30, x_1, x_2); x_32 = lean_ctor_get(x_31, 4); lean_inc(x_32); x_33 = lean_box(0); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_33); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_33); lean_dec(x_32); if (x_34 == 0) { @@ -17492,7 +17536,7 @@ goto block_82; lean_object* x_53; uint8_t x_54; x_53 = lean_ctor_get(x_52, 4); lean_inc(x_53); -x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_33); +x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_33); lean_dec(x_53); if (x_54 == 0) { @@ -17509,7 +17553,7 @@ lean_inc(x_1); x_55 = l_Lean_Parser_Term_doIfCond___elambda__1(x_1, x_52); x_56 = lean_ctor_get(x_55, 4); lean_inc(x_56); -x_57 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_56, x_33); +x_57 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_56, x_33); lean_dec(x_56); if (x_57 == 0) { @@ -17536,7 +17580,7 @@ if (x_62 == 0) lean_object* x_63; uint8_t x_64; x_63 = lean_ctor_get(x_59, 4); lean_inc(x_63); -x_64 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_63, x_33); +x_64 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_63, x_33); lean_dec(x_63); if (x_64 == 0) { @@ -17552,7 +17596,7 @@ lean_inc(x_1); x_65 = l_Lean_Parser_Term_doSeq___elambda__1(x_1, x_59); x_66 = lean_ctor_get(x_65, 4); lean_inc(x_66); -x_67 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_66, x_33); +x_67 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_66, x_33); lean_dec(x_66); if (x_67 == 0) { @@ -17568,7 +17612,7 @@ lean_inc(x_1); x_68 = lean_apply_2(x_4, x_1, x_65); x_69 = lean_ctor_get(x_68, 4); lean_inc(x_69); -x_70 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_69, x_33); +x_70 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_69, x_33); lean_dec(x_69); if (x_70 == 0) { @@ -17594,7 +17638,7 @@ lean_inc(x_1); x_72 = l_Lean_Parser_tokenAntiquotFn(x_1, x_59); x_73 = lean_ctor_get(x_72, 4); lean_inc(x_73); -x_74 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_73, x_33); +x_74 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_73, x_33); lean_dec(x_73); if (x_74 == 0) { @@ -17610,7 +17654,7 @@ lean_inc(x_1); x_75 = l_Lean_Parser_Term_doSeq___elambda__1(x_1, x_72); x_76 = lean_ctor_get(x_75, 4); lean_inc(x_76); -x_77 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_76, x_33); +x_77 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_76, x_33); lean_dec(x_76); if (x_77 == 0) { @@ -17626,7 +17670,7 @@ lean_inc(x_1); x_78 = lean_apply_2(x_4, x_1, x_75); x_79 = lean_ctor_get(x_78, 4); lean_inc(x_79); -x_80 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_79, x_33); +x_80 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_79, x_33); lean_dec(x_79); if (x_80 == 0) { @@ -17692,7 +17736,7 @@ goto block_123; lean_object* x_94; uint8_t x_95; x_94 = lean_ctor_get(x_93, 4); lean_inc(x_94); -x_95 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_94, x_33); +x_95 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_94, x_33); lean_dec(x_94); if (x_95 == 0) { @@ -17710,7 +17754,7 @@ lean_inc(x_86); x_96 = l_Lean_Parser_Term_doIfCond___elambda__1(x_86, x_93); x_97 = lean_ctor_get(x_96, 4); lean_inc(x_97); -x_98 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_97, x_33); +x_98 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_97, x_33); lean_dec(x_97); if (x_98 == 0) { @@ -17738,7 +17782,7 @@ if (x_103 == 0) lean_object* x_104; uint8_t x_105; x_104 = lean_ctor_get(x_100, 4); lean_inc(x_104); -x_105 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_104, x_33); +x_105 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_104, x_33); lean_dec(x_104); if (x_105 == 0) { @@ -17755,7 +17799,7 @@ lean_inc(x_86); x_106 = l_Lean_Parser_Term_doSeq___elambda__1(x_86, x_100); x_107 = lean_ctor_get(x_106, 4); lean_inc(x_107); -x_108 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_107, x_33); +x_108 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_107, x_33); lean_dec(x_107); if (x_108 == 0) { @@ -17772,7 +17816,7 @@ lean_inc(x_86); x_109 = lean_apply_2(x_4, x_86, x_106); x_110 = lean_ctor_get(x_109, 4); lean_inc(x_110); -x_111 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_110, x_33); +x_111 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_110, x_33); lean_dec(x_110); if (x_111 == 0) { @@ -17798,7 +17842,7 @@ lean_inc(x_86); x_113 = l_Lean_Parser_tokenAntiquotFn(x_86, x_100); x_114 = lean_ctor_get(x_113, 4); lean_inc(x_114); -x_115 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_114, x_33); +x_115 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_114, x_33); lean_dec(x_114); if (x_115 == 0) { @@ -17815,7 +17859,7 @@ lean_inc(x_86); x_116 = l_Lean_Parser_Term_doSeq___elambda__1(x_86, x_113); x_117 = lean_ctor_get(x_116, 4); lean_inc(x_117); -x_118 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_117, x_33); +x_118 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_117, x_33); lean_dec(x_117); if (x_118 == 0) { @@ -17832,7 +17876,7 @@ lean_inc(x_86); x_119 = lean_apply_2(x_4, x_86, x_116); x_120 = lean_ctor_get(x_119, 4); lean_inc(x_120); -x_121 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_120, x_33); +x_121 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_120, x_33); lean_dec(x_120); if (x_121 == 0) { @@ -17861,7 +17905,7 @@ lean_object* x_41; lean_object* x_42; uint8_t x_43; x_41 = l_Lean_Parser_ParserState_mkNode(x_40, x_10, x_36); x_42 = lean_ctor_get(x_41, 4); lean_inc(x_42); -x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_33); +x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_33); lean_dec(x_42); if (x_43 == 0) { @@ -18639,11 +18683,10 @@ return x_2; LEAN_EXPORT lean_object* l_Lean_Parser_Term_elseIf_formatter(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; -x_6 = l_Lean_groupKind; -x_7 = l_Lean_Parser_Term_elseIf_formatter___closed__6; -x_8 = l_Lean_PrettyPrinter_Formatter_node_formatter(x_6, x_7, x_1, x_2, x_3, x_4, x_5); -return x_8; +lean_object* x_6; lean_object* x_7; +x_6 = l_Lean_Parser_Term_elseIf_formatter___closed__6; +x_7 = l_Lean_Parser_group_formatter(x_6, x_1, x_2, x_3, x_4, x_5); +return x_7; } } static lean_object* _init_l_Lean_Parser_Term_doIf_formatter___closed__1() { @@ -18689,7 +18732,7 @@ static lean_object* _init_l_Lean_Parser_Term_doIf_formatter___closed__4() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_doIf_formatter___closed__3; -x_2 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__4; +x_2 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__6; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); @@ -18781,7 +18824,7 @@ static lean_object* _init_l_Lean_Parser_Term_doIf_formatter___closed__13() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_elseIf_formatter___closed__1; -x_2 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__4; +x_2 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__6; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); @@ -18826,7 +18869,7 @@ static lean_object* _init_l_Lean_Parser_Term_doIf_formatter___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__4; +x_1 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__6; x_2 = l_Lean_Parser_Term_doIf_formatter___closed__16; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -19415,11 +19458,10 @@ return x_2; LEAN_EXPORT lean_object* l_Lean_Parser_Term_elseIf_parenthesizer(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; -x_6 = l_Lean_groupKind; -x_7 = l_Lean_Parser_Term_elseIf_parenthesizer___closed__6; -x_8 = l_Lean_PrettyPrinter_Parenthesizer_node_parenthesizer(x_6, x_7, x_1, x_2, x_3, x_4, x_5); -return x_8; +lean_object* x_6; lean_object* x_7; +x_6 = l_Lean_Parser_Term_elseIf_parenthesizer___closed__6; +x_7 = l_Lean_Parser_group_parenthesizer(x_6, x_1, x_2, x_3, x_4, x_5); +return x_7; } } static lean_object* _init_l_Lean_Parser_Term_doIf_parenthesizer___closed__1() { @@ -19465,7 +19507,7 @@ static lean_object* _init_l_Lean_Parser_Term_doIf_parenthesizer___closed__4() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_doIf_parenthesizer___closed__3; -x_2 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__7; +x_2 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__9; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); @@ -19557,7 +19599,7 @@ static lean_object* _init_l_Lean_Parser_Term_doIf_parenthesizer___closed__13() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_elseIf_parenthesizer___closed__1; -x_2 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__7; +x_2 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__9; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); @@ -19602,7 +19644,7 @@ static lean_object* _init_l_Lean_Parser_Term_doIf_parenthesizer___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__7; +x_1 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__9; x_2 = l_Lean_Parser_Term_doIf_parenthesizer___closed__16; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -19956,7 +19998,7 @@ x_30 = l_Lean_Parser_checkPrecFn(x_29, x_1, x_2); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); x_32 = lean_box(0); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_32); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_32); lean_dec(x_31); if (x_33 == 0) { @@ -20006,7 +20048,7 @@ lean_object* x_44; lean_object* x_45; uint8_t x_46; x_44 = l_Lean_Parser_ParserState_mkNode(x_43, x_9, x_35); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_32); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_32); lean_dec(x_45); if (x_46 == 0) { @@ -20026,7 +20068,7 @@ return x_47; lean_object* x_50; uint8_t x_51; x_50 = lean_ctor_get(x_49, 4); lean_inc(x_50); -x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_32); +x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_32); lean_dec(x_50); if (x_51 == 0) { @@ -20057,7 +20099,7 @@ x_54 = lean_unsigned_to_nat(0u); x_55 = l_Lean_Parser_categoryParser___elambda__1(x_3, x_54, x_53, x_49); x_56 = lean_ctor_get(x_55, 4); lean_inc(x_56); -x_57 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_56, x_32); +x_57 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_56, x_32); lean_dec(x_56); if (x_57 == 0) { @@ -20083,7 +20125,7 @@ if (x_63 == 0) lean_object* x_64; uint8_t x_65; x_64 = lean_ctor_get(x_60, 4); lean_inc(x_64); -x_65 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_64, x_32); +x_65 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_64, x_32); lean_dec(x_64); if (x_65 == 0) { @@ -20106,7 +20148,7 @@ lean_inc(x_1); x_67 = l_Lean_Parser_tokenAntiquotFn(x_1, x_60); x_68 = lean_ctor_get(x_67, 4); lean_inc(x_68); -x_69 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_68, x_32); +x_69 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_68, x_32); lean_dec(x_68); if (x_69 == 0) { @@ -20433,7 +20475,7 @@ static lean_object* _init_l_Lean_Parser_Term_doUnless_formatter___closed__5() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_doUnless_formatter___closed__4; -x_2 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__4; +x_2 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__6; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); @@ -20571,7 +20613,7 @@ static lean_object* _init_l_Lean_Parser_Term_doUnless_parenthesizer___closed__5( { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_doUnless_parenthesizer___closed__4; -x_2 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__7; +x_2 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__9; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); @@ -20948,7 +20990,7 @@ x_20 = l_Lean_Parser_checkPrecFn(x_19, x_1, x_2); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); x_22 = lean_box(0); -x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_22); +x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_22); lean_dec(x_21); if (x_23 == 0) { @@ -20974,7 +21016,7 @@ lean_inc(x_1); x_26 = lean_apply_2(x_4, x_1, x_20); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_22); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_22); lean_dec(x_27); if (x_28 == 0) { @@ -20997,7 +21039,7 @@ lean_inc(x_1); x_38 = l_Lean_Parser_categoryParser___elambda__1(x_36, x_37, x_1, x_26); x_39 = lean_ctor_get(x_38, 4); lean_inc(x_39); -x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_22); +x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_22); lean_dec(x_39); if (x_40 == 0) { @@ -21030,7 +21072,7 @@ if (x_47 == 0) lean_object* x_48; uint8_t x_49; x_48 = lean_ctor_get(x_43, 4); lean_inc(x_48); -x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_22); +x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_22); lean_dec(x_48); if (x_49 == 0) { @@ -21068,7 +21110,7 @@ lean_inc(x_1); x_53 = l_Lean_Parser_tokenAntiquotFn(x_1, x_43); x_54 = lean_ctor_get(x_53, 4); lean_inc(x_54); -x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_22); +x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_22); lean_dec(x_54); if (x_55 == 0) { @@ -21108,7 +21150,7 @@ x_30 = l_Lean_Parser_Term_doForDecl___elambda__1___closed__2; x_31 = l_Lean_Parser_ParserState_mkNode(x_29, x_30, x_25); x_32 = lean_ctor_get(x_31, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_22); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_22); lean_dec(x_32); if (x_33 == 0) { @@ -21487,7 +21529,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -21533,7 +21575,7 @@ x_28 = l_Lean_Parser_Term_doFor___elambda__1___closed__2; x_29 = l_Lean_Parser_ParserState_mkNode(x_27, x_28, x_19); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_16); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_16); lean_dec(x_30); if (x_31 == 0) { @@ -21553,7 +21595,7 @@ return x_32; lean_object* x_35; uint8_t x_36; x_35 = lean_ctor_get(x_34, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_16); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_16); lean_dec(x_35); if (x_36 == 0) { @@ -21569,7 +21611,7 @@ lean_inc(x_1); x_37 = lean_apply_2(x_4, x_1, x_34); x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_16); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_16); lean_dec(x_38); if (x_39 == 0) { @@ -21595,7 +21637,7 @@ if (x_45 == 0) lean_object* x_46; uint8_t x_47; x_46 = lean_ctor_get(x_42, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_16); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_16); lean_dec(x_46); if (x_47 == 0) { @@ -21618,7 +21660,7 @@ lean_inc(x_1); x_49 = l_Lean_Parser_tokenAntiquotFn(x_1, x_42); x_50 = lean_ctor_get(x_49, 4); lean_inc(x_50); -x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_16); +x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_16); lean_dec(x_50); if (x_51 == 0) { @@ -22787,7 +22829,7 @@ x_20 = l_Lean_Parser_checkPrecFn(x_19, x_1, x_2); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); x_22 = lean_box(0); -x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_22); +x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_22); lean_dec(x_21); if (x_23 == 0) { @@ -22836,7 +22878,7 @@ x_34 = l_Lean_Parser_Term_doMatch___elambda__1___closed__2; x_35 = l_Lean_Parser_ParserState_mkNode(x_33, x_34, x_25); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_22); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_22); lean_dec(x_36); if (x_37 == 0) { @@ -22856,7 +22898,7 @@ return x_38; lean_object* x_41; uint8_t x_42; x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_22); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_22); lean_dec(x_41); if (x_42 == 0) { @@ -22875,7 +22917,7 @@ lean_inc(x_1); x_43 = lean_apply_2(x_10, x_1, x_40); x_44 = lean_ctor_get(x_43, 4); lean_inc(x_44); -x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_22); +x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_22); lean_dec(x_44); if (x_45 == 0) { @@ -22893,7 +22935,7 @@ lean_inc(x_1); x_46 = lean_apply_2(x_8, x_1, x_43); x_47 = lean_ctor_get(x_46, 4); lean_inc(x_47); -x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_22); +x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_22); lean_dec(x_47); if (x_48 == 0) { @@ -22910,7 +22952,7 @@ lean_inc(x_1); x_49 = lean_apply_2(x_6, x_1, x_46); x_50 = lean_ctor_get(x_49, 4); lean_inc(x_50); -x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_22); +x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_22); lean_dec(x_50); if (x_51 == 0) { @@ -22937,7 +22979,7 @@ if (x_57 == 0) lean_object* x_58; uint8_t x_59; x_58 = lean_ctor_get(x_54, 4); lean_inc(x_58); -x_59 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_58, x_22); +x_59 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_58, x_22); lean_dec(x_58); if (x_59 == 0) { @@ -22961,7 +23003,7 @@ lean_inc(x_1); x_61 = l_Lean_Parser_tokenAntiquotFn(x_1, x_54); x_62 = lean_ctor_get(x_61, 4); lean_inc(x_62); -x_63 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_62, x_22); +x_63 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_62, x_22); lean_dec(x_62); if (x_63 == 0) { @@ -23265,7 +23307,7 @@ static lean_object* _init_l_Lean_Parser_Term_doMatchAlts_formatter___closed__1() _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__4; +x_1 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__6; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_Term_matchAlts_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -23505,7 +23547,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_doMatchAlts_parenthesizer(lean_objec _start: { lean_object* x_6; lean_object* x_7; -x_6 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__7; +x_6 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__9; x_7 = l_Lean_Parser_Term_matchAlts_parenthesizer(x_6, x_1, x_2, x_3, x_4, x_5); return x_7; } @@ -23976,7 +24018,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -23996,7 +24038,7 @@ lean_inc(x_1); x_21 = l_Lean_Parser_atomicFn(x_20, x_1, x_14); x_22 = lean_ctor_get(x_21, 4); lean_inc(x_22); -x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_22, x_16); +x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_22, x_16); lean_dec(x_22); if (x_23 == 0) { @@ -24006,7 +24048,7 @@ x_24 = l_Lean_Parser_Term_doCatch___elambda__1___closed__2; x_25 = l_Lean_Parser_ParserState_mkNode(x_21, x_24, x_19); x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_16); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_16); lean_dec(x_26); if (x_27 == 0) { @@ -24028,7 +24070,7 @@ lean_inc(x_1); x_29 = lean_apply_2(x_4, x_1, x_21); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_16); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_16); lean_dec(x_30); if (x_31 == 0) { @@ -24037,7 +24079,7 @@ x_32 = l_Lean_Parser_Term_doCatch___elambda__1___closed__2; x_33 = l_Lean_Parser_ParserState_mkNode(x_29, x_32, x_19); x_34 = lean_ctor_get(x_33, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_16); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_16); lean_dec(x_34); if (x_35 == 0) { @@ -24059,7 +24101,7 @@ lean_inc(x_1); x_37 = l_Lean_Parser_darrow___elambda__1(x_1, x_29); x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_16); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_16); lean_dec(x_38); if (x_39 == 0) { @@ -24068,7 +24110,7 @@ x_40 = l_Lean_Parser_Term_doCatch___elambda__1___closed__2; x_41 = l_Lean_Parser_ParserState_mkNode(x_37, x_40, x_19); x_42 = lean_ctor_get(x_41, 4); lean_inc(x_42); -x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_16); +x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_16); lean_dec(x_42); if (x_43 == 0) { @@ -24092,7 +24134,7 @@ x_46 = l_Lean_Parser_Term_doCatch___elambda__1___closed__2; x_47 = l_Lean_Parser_ParserState_mkNode(x_45, x_46, x_19); x_48 = lean_ctor_get(x_47, 4); lean_inc(x_48); -x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_16); +x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_16); lean_dec(x_48); if (x_49 == 0) { @@ -24379,7 +24421,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -24411,7 +24453,7 @@ if (x_26 == 0) lean_object* x_27; uint8_t x_28; x_27 = lean_ctor_get(x_22, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); lean_dec(x_27); if (x_28 == 0) { @@ -24421,7 +24463,7 @@ x_29 = l_Lean_Parser_Term_doCatchMatch___elambda__1___closed__2; x_30 = l_Lean_Parser_ParserState_mkNode(x_22, x_29, x_19); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); lean_dec(x_31); if (x_32 == 0) { @@ -24445,7 +24487,7 @@ x_35 = l_Lean_Parser_Term_doCatchMatch___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_19); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_16); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_16); lean_dec(x_37); if (x_38 == 0) { @@ -24468,7 +24510,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_22); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); lean_dec(x_41); if (x_42 == 0) { @@ -24478,7 +24520,7 @@ x_43 = l_Lean_Parser_Term_doCatchMatch___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_19); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); lean_dec(x_45); if (x_46 == 0) { @@ -24502,7 +24544,7 @@ x_49 = l_Lean_Parser_Term_doCatchMatch___elambda__1___closed__2; x_50 = l_Lean_Parser_ParserState_mkNode(x_48, x_49, x_19); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); lean_dec(x_51); if (x_52 == 0) { @@ -24766,7 +24808,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -24797,7 +24839,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_20, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -24806,7 +24848,7 @@ x_27 = l_Lean_Parser_Term_doFinally___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_20, x_27, x_17); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -24830,7 +24872,7 @@ x_33 = l_Lean_Parser_Term_doFinally___elambda__1___closed__2; x_34 = l_Lean_Parser_ParserState_mkNode(x_32, x_33, x_17); x_35 = lean_ctor_get(x_34, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_14); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_14); lean_dec(x_35); if (x_36 == 0) { @@ -24853,7 +24895,7 @@ lean_inc(x_1); x_38 = l_Lean_Parser_tokenAntiquotFn(x_1, x_20); x_39 = lean_ctor_get(x_38, 4); lean_inc(x_39); -x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); +x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); lean_dec(x_39); if (x_40 == 0) { @@ -24862,7 +24904,7 @@ x_41 = l_Lean_Parser_Term_doFinally___elambda__1___closed__2; x_42 = l_Lean_Parser_ParserState_mkNode(x_38, x_41, x_17); x_43 = lean_ctor_get(x_42, 4); lean_inc(x_43); -x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_14); +x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_14); lean_dec(x_43); if (x_44 == 0) { @@ -24886,7 +24928,7 @@ x_47 = l_Lean_Parser_Term_doFinally___elambda__1___closed__2; x_48 = l_Lean_Parser_ParserState_mkNode(x_46, x_47, x_17); x_49 = lean_ctor_get(x_48, 4); lean_inc(x_49); -x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_14); +x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_14); lean_dec(x_49); if (x_50 == 0) { @@ -25248,7 +25290,7 @@ x_16 = l_Lean_Parser_checkPrecFn(x_15, x_1, x_2); x_17 = lean_ctor_get(x_16, 4); lean_inc(x_17); x_18 = lean_box(0); -x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); +x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); lean_dec(x_17); if (x_19 == 0) { @@ -25294,7 +25336,7 @@ goto block_59; lean_object* x_30; uint8_t x_31; x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_18); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_18); lean_dec(x_30); if (x_31 == 0) { @@ -25305,7 +25347,7 @@ x_32 = l_Lean_Parser_Term_doTry___elambda__1___closed__2; x_33 = l_Lean_Parser_ParserState_mkNode(x_29, x_32, x_21); x_34 = lean_ctor_get(x_33, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_18); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_18); lean_dec(x_34); if (x_35 == 0) { @@ -25327,7 +25369,7 @@ lean_inc(x_1); x_37 = l_Lean_Parser_Term_doSeq___elambda__1(x_1, x_29); x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_18); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_18); lean_dec(x_38); if (x_39 == 0) { @@ -25338,7 +25380,7 @@ x_40 = l_Lean_Parser_Term_doTry___elambda__1___closed__2; x_41 = l_Lean_Parser_ParserState_mkNode(x_37, x_40, x_21); x_42 = lean_ctor_get(x_41, 4); lean_inc(x_42); -x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_18); +x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_18); lean_dec(x_42); if (x_43 == 0) { @@ -25360,7 +25402,7 @@ lean_inc(x_1); x_45 = lean_apply_2(x_4, x_1, x_37); x_46 = lean_ctor_get(x_45, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_18); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_18); lean_dec(x_46); if (x_47 == 0) { @@ -25370,7 +25412,7 @@ x_48 = l_Lean_Parser_Term_doTry___elambda__1___closed__2; x_49 = l_Lean_Parser_ParserState_mkNode(x_45, x_48, x_21); x_50 = lean_ctor_get(x_49, 4); lean_inc(x_50); -x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_18); +x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_18); lean_dec(x_50); if (x_51 == 0) { @@ -25394,7 +25436,7 @@ x_54 = l_Lean_Parser_Term_doTry___elambda__1___closed__2; x_55 = l_Lean_Parser_ParserState_mkNode(x_53, x_54, x_21); x_56 = lean_ctor_get(x_55, 4); lean_inc(x_56); -x_57 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_56, x_18); +x_57 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_56, x_18); lean_dec(x_56); if (x_57 == 0) { @@ -25750,7 +25792,7 @@ static lean_object* _init_l_Lean_Parser_Term_doCatch_formatter___closed__9() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_doCatch_formatter___closed__8; -x_2 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__4; +x_2 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__6; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); @@ -25952,7 +25994,7 @@ static lean_object* _init_l_Lean_Parser_Term_doFinally_formatter___closed__3() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_doFinally_formatter___closed__2; -x_2 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__4; +x_2 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__6; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); @@ -26089,7 +26131,7 @@ static lean_object* _init_l_Lean_Parser_Term_doTry_formatter___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__4; +x_1 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__6; x_2 = l_Lean_Parser_Term_doTry_formatter___closed__6; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -26256,7 +26298,7 @@ static lean_object* _init_l_Lean_Parser_Term_doCatch_parenthesizer___closed__9() { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_doCatch_parenthesizer___closed__8; -x_2 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__7; +x_2 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__9; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); @@ -26458,7 +26500,7 @@ static lean_object* _init_l_Lean_Parser_Term_doFinally_parenthesizer___closed__3 { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_doFinally_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__7; +x_2 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__9; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); @@ -26595,7 +26637,7 @@ static lean_object* _init_l_Lean_Parser_Term_doTry_parenthesizer___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__7; +x_1 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__9; x_2 = l_Lean_Parser_Term_doTry_parenthesizer___closed__6; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -26809,7 +26851,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -26842,7 +26884,7 @@ x_25 = l_Lean_Parser_Term_doBreak___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_20, x_25, x_17); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); lean_dec(x_27); if (x_28 == 0) { @@ -26866,7 +26908,7 @@ x_31 = l_Lean_Parser_Term_doBreak___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_17); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -27392,7 +27434,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -27425,7 +27467,7 @@ x_25 = l_Lean_Parser_Term_doContinue___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_20, x_25, x_17); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); lean_dec(x_27); if (x_28 == 0) { @@ -27449,7 +27491,7 @@ x_31 = l_Lean_Parser_Term_doContinue___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_17); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -27873,7 +27915,7 @@ if (x_18 == 0) lean_object* x_20; uint8_t x_21; x_20 = lean_ctor_get(x_13, 4); lean_inc(x_20); -x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_20, x_19); +x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_20, x_19); lean_dec(x_20); if (x_21 == 0) { @@ -27895,7 +27937,7 @@ lean_inc(x_3); x_23 = l_Lean_Parser_tokenAntiquotFn(x_3, x_13); x_24 = lean_ctor_get(x_23, 4); lean_inc(x_24); -x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_19); +x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_19); lean_dec(x_24); if (x_25 == 0) { @@ -27963,7 +28005,7 @@ if (x_45 == 0) lean_object* x_47; uint8_t x_48; x_47 = lean_ctor_get(x_40, 4); lean_inc(x_47); -x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_46); +x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_46); lean_dec(x_47); if (x_48 == 0) { @@ -27985,7 +28027,7 @@ lean_inc(x_36); x_50 = l_Lean_Parser_tokenAntiquotFn(x_36, x_40); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_46); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_46); lean_dec(x_51); if (x_52 == 0) { @@ -28183,7 +28225,7 @@ x_28 = l_Lean_Parser_checkPrecFn(x_27, x_1, x_2); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); x_30 = lean_box(0); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_30); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_30); lean_dec(x_29); if (x_31 == 0) { @@ -28233,7 +28275,7 @@ if (x_42 == 0) lean_object* x_43; uint8_t x_44; x_43 = lean_ctor_get(x_38, 4); lean_inc(x_43); -x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_30); +x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_30); lean_dec(x_43); if (x_44 == 0) { @@ -28243,7 +28285,7 @@ lean_dec(x_4); x_45 = l_Lean_Parser_ParserState_mkNode(x_38, x_7, x_33); x_46 = lean_ctor_get(x_45, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_30); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_30); lean_dec(x_46); if (x_47 == 0) { @@ -28265,7 +28307,7 @@ x_49 = lean_apply_2(x_4, x_36, x_38); x_50 = l_Lean_Parser_ParserState_mkNode(x_49, x_7, x_33); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_30); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_30); lean_dec(x_51); if (x_52 == 0) { @@ -28288,7 +28330,7 @@ lean_inc(x_36); x_54 = l_Lean_Parser_tokenAntiquotFn(x_36, x_38); x_55 = lean_ctor_get(x_54, 4); lean_inc(x_55); -x_56 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_55, x_30); +x_56 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_55, x_30); lean_dec(x_55); if (x_56 == 0) { @@ -28298,7 +28340,7 @@ lean_dec(x_4); x_57 = l_Lean_Parser_ParserState_mkNode(x_54, x_7, x_33); x_58 = lean_ctor_get(x_57, 4); lean_inc(x_58); -x_59 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_58, x_30); +x_59 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_58, x_30); lean_dec(x_58); if (x_59 == 0) { @@ -28320,7 +28362,7 @@ x_61 = lean_apply_2(x_4, x_36, x_54); x_62 = l_Lean_Parser_ParserState_mkNode(x_61, x_7, x_33); x_63 = lean_ctor_get(x_62, 4); lean_inc(x_63); -x_64 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_63, x_30); +x_64 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_63, x_30); lean_dec(x_63); if (x_64 == 0) { @@ -29000,7 +29042,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -29032,7 +29074,7 @@ if (x_26 == 0) lean_object* x_27; uint8_t x_28; x_27 = lean_ctor_get(x_22, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); lean_dec(x_27); if (x_28 == 0) { @@ -29042,7 +29084,7 @@ x_29 = l_Lean_Parser_Term_doDbgTrace___elambda__1___closed__2; x_30 = l_Lean_Parser_ParserState_mkNode(x_22, x_29, x_19); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); lean_dec(x_31); if (x_32 == 0) { @@ -29068,7 +29110,7 @@ x_37 = l_Lean_Parser_Term_doDbgTrace___elambda__1___closed__2; x_38 = l_Lean_Parser_ParserState_mkNode(x_36, x_37, x_19); x_39 = lean_ctor_get(x_38, 4); lean_inc(x_39); -x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_16); +x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_16); lean_dec(x_39); if (x_40 == 0) { @@ -29091,7 +29133,7 @@ lean_inc(x_1); x_42 = l_Lean_Parser_tokenAntiquotFn(x_1, x_22); x_43 = lean_ctor_get(x_42, 4); lean_inc(x_43); -x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_16); +x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_16); lean_dec(x_43); if (x_44 == 0) { @@ -29101,7 +29143,7 @@ x_45 = l_Lean_Parser_Term_doDbgTrace___elambda__1___closed__2; x_46 = l_Lean_Parser_ParserState_mkNode(x_42, x_45, x_19); x_47 = lean_ctor_get(x_46, 4); lean_inc(x_47); -x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_16); +x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_16); lean_dec(x_47); if (x_48 == 0) { @@ -29127,7 +29169,7 @@ x_53 = l_Lean_Parser_Term_doDbgTrace___elambda__1___closed__2; x_54 = l_Lean_Parser_ParserState_mkNode(x_52, x_53, x_19); x_55 = lean_ctor_get(x_54, 4); lean_inc(x_55); -x_56 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_55, x_16); +x_56 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_55, x_16); lean_dec(x_55); if (x_56 == 0) { @@ -29759,7 +29801,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -29790,7 +29832,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_20, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -29799,7 +29841,7 @@ x_27 = l_Lean_Parser_Term_doAssert___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_20, x_27, x_17); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -29825,7 +29867,7 @@ x_35 = l_Lean_Parser_Term_doAssert___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_17); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); lean_dec(x_37); if (x_38 == 0) { @@ -29848,7 +29890,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_20); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); lean_dec(x_41); if (x_42 == 0) { @@ -29857,7 +29899,7 @@ x_43 = l_Lean_Parser_Term_doAssert___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_17); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); lean_dec(x_45); if (x_46 == 0) { @@ -29883,7 +29925,7 @@ x_51 = l_Lean_Parser_Term_doAssert___elambda__1___closed__2; x_52 = l_Lean_Parser_ParserState_mkNode(x_50, x_51, x_17); x_53 = lean_ctor_get(x_52, 4); lean_inc(x_53); -x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); +x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); lean_dec(x_53); if (x_54 == 0) { @@ -30549,7 +30591,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -30567,7 +30609,7 @@ lean_inc(x_1); x_18 = l_Lean_Parser_Term_notFollowedByRedefinedTermToken___elambda__1(x_1, x_12); x_19 = lean_ctor_get(x_18, 4); lean_inc(x_19); -x_20 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_14); +x_20 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_14); lean_dec(x_19); if (x_20 == 0) { @@ -30576,7 +30618,7 @@ x_21 = l_Lean_Parser_Term_doExpr___elambda__1___closed__2; x_22 = l_Lean_Parser_ParserState_mkNode(x_18, x_21, x_17); x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); -x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_14); +x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_14); lean_dec(x_23); if (x_24 == 0) { @@ -30600,7 +30642,7 @@ lean_inc(x_1); x_28 = l_Lean_Parser_categoryParser___elambda__1(x_26, x_27, x_1, x_18); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -30609,7 +30651,7 @@ x_31 = l_Lean_Parser_Term_doExpr___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_28, x_31, x_17); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -30635,7 +30677,7 @@ x_39 = l_Lean_Parser_Term_doExpr___elambda__1___closed__2; x_40 = l_Lean_Parser_ParserState_mkNode(x_38, x_39, x_17); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); lean_dec(x_41); if (x_42 == 0) { @@ -31285,7 +31327,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -31316,7 +31358,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_20, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -31325,7 +31367,7 @@ x_27 = l_Lean_Parser_Term_doNested___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_20, x_27, x_17); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -31349,7 +31391,7 @@ x_33 = l_Lean_Parser_Term_doNested___elambda__1___closed__2; x_34 = l_Lean_Parser_ParserState_mkNode(x_32, x_33, x_17); x_35 = lean_ctor_get(x_34, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_14); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_14); lean_dec(x_35); if (x_36 == 0) { @@ -31372,7 +31414,7 @@ lean_inc(x_1); x_38 = l_Lean_Parser_tokenAntiquotFn(x_1, x_20); x_39 = lean_ctor_get(x_38, 4); lean_inc(x_39); -x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); +x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); lean_dec(x_39); if (x_40 == 0) { @@ -31381,7 +31423,7 @@ x_41 = l_Lean_Parser_Term_doNested___elambda__1___closed__2; x_42 = l_Lean_Parser_ParserState_mkNode(x_38, x_41, x_17); x_43 = lean_ctor_get(x_42, 4); lean_inc(x_43); -x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_14); +x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_14); lean_dec(x_43); if (x_44 == 0) { @@ -31405,7 +31447,7 @@ x_47 = l_Lean_Parser_Term_doNested___elambda__1___closed__2; x_48 = l_Lean_Parser_ParserState_mkNode(x_46, x_47, x_17); x_49 = lean_ctor_get(x_48, 4); lean_inc(x_49); -x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_14); +x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_14); lean_dec(x_49); if (x_50 == 0) { @@ -31880,7 +31922,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -31911,7 +31953,7 @@ if (x_31 == 0) lean_object* x_32; uint8_t x_33; x_32 = lean_ctor_get(x_27, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_14); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_14); lean_dec(x_32); if (x_33 == 0) { @@ -31934,7 +31976,7 @@ lean_inc(x_1); x_35 = l_Lean_Parser_tokenAntiquotFn(x_1, x_27); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); lean_dec(x_36); if (x_37 == 0) { @@ -31957,7 +31999,7 @@ x_19 = l_Lean_Parser_Term_do___elambda__1___closed__1; x_20 = l_Lean_Parser_ParserState_mkNode(x_18, x_19, x_17); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); -x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_14); +x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_14); lean_dec(x_21); if (x_22 == 0) { @@ -32606,7 +32648,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -32649,7 +32691,7 @@ goto block_60; lean_object* x_26; uint8_t x_27; x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); lean_dec(x_26); if (x_27 == 0) { @@ -32659,7 +32701,7 @@ x_28 = l_Lean_Parser_Term_doElem_quot___elambda__1___closed__3; x_29 = l_Lean_Parser_ParserState_mkNode(x_25, x_28, x_17); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); lean_dec(x_30); if (x_31 == 0) { @@ -32683,7 +32725,7 @@ lean_inc(x_1); x_35 = l___private_Lean_Parser_Basic_0__Lean_Parser_addQuotDepthFn(x_33, x_34, x_1, x_25); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); lean_dec(x_36); if (x_37 == 0) { @@ -32693,7 +32735,7 @@ x_38 = l_Lean_Parser_Term_doElem_quot___elambda__1___closed__3; x_39 = l_Lean_Parser_ParserState_mkNode(x_35, x_38, x_17); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); lean_dec(x_40); if (x_41 == 0) { @@ -32728,7 +32770,7 @@ x_49 = l_Lean_Parser_Term_doElem_quot___elambda__1___closed__3; x_50 = l_Lean_Parser_ParserState_mkNode(x_45, x_49, x_17); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_14); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_14); lean_dec(x_51); if (x_52 == 0) { @@ -32752,7 +32794,7 @@ x_55 = l_Lean_Parser_Term_doElem_quot___elambda__1___closed__3; x_56 = l_Lean_Parser_ParserState_mkNode(x_54, x_55, x_17); x_57 = lean_ctor_get(x_56, 4); lean_inc(x_57); -x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_14); +x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_14); lean_dec(x_57); if (x_58 == 0) { @@ -33387,7 +33429,7 @@ x_18 = l_Lean_Parser_checkPrecFn(x_17, x_1, x_2); x_19 = lean_ctor_get(x_18, 4); lean_inc(x_19); x_20 = lean_box(0); -x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_20); +x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_20); lean_dec(x_19); if (x_21 == 0) { @@ -33438,7 +33480,7 @@ x_32 = l_Lean_Parser_Term_termUnless___elambda__1___closed__2; x_33 = l_Lean_Parser_ParserState_mkNode(x_31, x_32, x_23); x_34 = lean_ctor_get(x_33, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_20); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_20); lean_dec(x_34); if (x_35 == 0) { @@ -33458,7 +33500,7 @@ return x_36; lean_object* x_39; uint8_t x_40; x_39 = lean_ctor_get(x_38, 4); lean_inc(x_39); -x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_20); +x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_20); lean_dec(x_39); if (x_40 == 0) { @@ -33490,7 +33532,7 @@ x_44 = lean_unsigned_to_nat(0u); x_45 = l_Lean_Parser_categoryParser___elambda__1(x_43, x_44, x_42, x_38); x_46 = lean_ctor_get(x_45, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_20); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_20); lean_dec(x_46); if (x_47 == 0) { @@ -33516,7 +33558,7 @@ if (x_53 == 0) lean_object* x_54; uint8_t x_55; x_54 = lean_ctor_get(x_50, 4); lean_inc(x_54); -x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_20); +x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_20); lean_dec(x_54); if (x_55 == 0) { @@ -33539,7 +33581,7 @@ lean_inc(x_1); x_57 = l_Lean_Parser_tokenAntiquotFn(x_1, x_50); x_58 = lean_ctor_get(x_57, 4); lean_inc(x_58); -x_59 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_58, x_20); +x_59 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_58, x_20); lean_dec(x_58); if (x_59 == 0) { @@ -34002,7 +34044,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -34048,7 +34090,7 @@ x_28 = l_Lean_Parser_Term_termFor___elambda__1___closed__2; x_29 = l_Lean_Parser_ParserState_mkNode(x_27, x_28, x_19); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_16); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_16); lean_dec(x_30); if (x_31 == 0) { @@ -34068,7 +34110,7 @@ return x_32; lean_object* x_35; uint8_t x_36; x_35 = lean_ctor_get(x_34, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_16); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_16); lean_dec(x_35); if (x_36 == 0) { @@ -34084,7 +34126,7 @@ lean_inc(x_1); x_37 = lean_apply_2(x_4, x_1, x_34); x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_16); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_16); lean_dec(x_38); if (x_39 == 0) { @@ -34110,7 +34152,7 @@ if (x_45 == 0) lean_object* x_46; uint8_t x_47; x_46 = lean_ctor_get(x_42, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_16); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_16); lean_dec(x_46); if (x_47 == 0) { @@ -34133,7 +34175,7 @@ lean_inc(x_1); x_49 = l_Lean_Parser_tokenAntiquotFn(x_1, x_42); x_50 = lean_ctor_get(x_49, 4); lean_inc(x_50); -x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_16); +x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_16); lean_dec(x_50); if (x_51 == 0) { @@ -34594,7 +34636,7 @@ x_16 = l_Lean_Parser_checkPrecFn(x_15, x_1, x_2); x_17 = lean_ctor_get(x_16, 4); lean_inc(x_17); x_18 = lean_box(0); -x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); +x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); lean_dec(x_17); if (x_19 == 0) { @@ -34640,7 +34682,7 @@ goto block_59; lean_object* x_30; uint8_t x_31; x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_18); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_18); lean_dec(x_30); if (x_31 == 0) { @@ -34651,7 +34693,7 @@ x_32 = l_Lean_Parser_Term_termTry___elambda__1___closed__2; x_33 = l_Lean_Parser_ParserState_mkNode(x_29, x_32, x_21); x_34 = lean_ctor_get(x_33, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_18); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_18); lean_dec(x_34); if (x_35 == 0) { @@ -34673,7 +34715,7 @@ lean_inc(x_1); x_37 = l_Lean_Parser_Term_doSeq___elambda__1(x_1, x_29); x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_18); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_18); lean_dec(x_38); if (x_39 == 0) { @@ -34684,7 +34726,7 @@ x_40 = l_Lean_Parser_Term_termTry___elambda__1___closed__2; x_41 = l_Lean_Parser_ParserState_mkNode(x_37, x_40, x_21); x_42 = lean_ctor_get(x_41, 4); lean_inc(x_42); -x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_18); +x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_18); lean_dec(x_42); if (x_43 == 0) { @@ -34706,7 +34748,7 @@ lean_inc(x_1); x_45 = lean_apply_2(x_4, x_1, x_37); x_46 = lean_ctor_get(x_45, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_18); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_18); lean_dec(x_46); if (x_47 == 0) { @@ -34716,7 +34758,7 @@ x_48 = l_Lean_Parser_Term_termTry___elambda__1___closed__2; x_49 = l_Lean_Parser_ParserState_mkNode(x_45, x_48, x_21); x_50 = lean_ctor_get(x_49, 4); lean_inc(x_50); -x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_18); +x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_18); lean_dec(x_50); if (x_51 == 0) { @@ -34740,7 +34782,7 @@ x_54 = l_Lean_Parser_Term_termTry___elambda__1___closed__2; x_55 = l_Lean_Parser_ParserState_mkNode(x_53, x_54, x_21); x_56 = lean_ctor_get(x_55, 4); lean_inc(x_56); -x_57 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_56, x_18); +x_57 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_56, x_18); lean_dec(x_56); if (x_57 == 0) { @@ -35222,7 +35264,7 @@ x_20 = l_Lean_Parser_checkPrecFn(x_19, x_1, x_2); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); x_22 = lean_box(0); -x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_22); +x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_22); lean_dec(x_21); if (x_23 == 0) { @@ -35273,7 +35315,7 @@ if (x_35 == 0) lean_object* x_36; uint8_t x_37; x_36 = lean_ctor_get(x_31, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_22); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_22); lean_dec(x_36); if (x_37 == 0) { @@ -35284,7 +35326,7 @@ x_38 = l_Lean_Parser_Term_termReturn___elambda__1___closed__2; x_39 = l_Lean_Parser_ParserState_mkNode(x_31, x_38, x_25); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_22); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_22); lean_dec(x_40); if (x_41 == 0) { @@ -35307,7 +35349,7 @@ x_44 = l_Lean_Parser_Term_termReturn___elambda__1___closed__2; x_45 = l_Lean_Parser_ParserState_mkNode(x_43, x_44, x_25); x_46 = lean_ctor_get(x_45, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_22); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_22); lean_dec(x_46); if (x_47 == 0) { @@ -35330,7 +35372,7 @@ lean_inc(x_28); x_49 = l_Lean_Parser_tokenAntiquotFn(x_28, x_31); x_50 = lean_ctor_get(x_49, 4); lean_inc(x_50); -x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_22); +x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_22); lean_dec(x_50); if (x_51 == 0) { @@ -35341,7 +35383,7 @@ x_52 = l_Lean_Parser_Term_termReturn___elambda__1___closed__2; x_53 = l_Lean_Parser_ParserState_mkNode(x_49, x_52, x_25); x_54 = lean_ctor_get(x_53, 4); lean_inc(x_54); -x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_22); +x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_22); lean_dec(x_54); if (x_55 == 0) { @@ -35364,7 +35406,7 @@ x_58 = l_Lean_Parser_Term_termReturn___elambda__1___closed__2; x_59 = l_Lean_Parser_ParserState_mkNode(x_57, x_58, x_25); x_60 = lean_ctor_get(x_59, 4); lean_inc(x_60); -x_61 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_60, x_22); +x_61 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_60, x_22); lean_dec(x_60); if (x_61 == 0) { @@ -36216,6 +36258,10 @@ l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__17 = _init_ lean_mark_persistent(l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__17); l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__18 = _init_l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__18(); lean_mark_persistent(l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__18); +l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__19 = _init_l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__19(); +lean_mark_persistent(l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__19); +l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__20 = _init_l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__20(); +lean_mark_persistent(l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157____closed__20); res = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_157_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); @@ -36625,6 +36671,10 @@ l_Lean_Parser_Term_doLetRec___elambda__1___closed__7 = _init_l_Lean_Parser_Term_ lean_mark_persistent(l_Lean_Parser_Term_doLetRec___elambda__1___closed__7); l_Lean_Parser_Term_doLetRec___elambda__1___closed__8 = _init_l_Lean_Parser_Term_doLetRec___elambda__1___closed__8(); lean_mark_persistent(l_Lean_Parser_Term_doLetRec___elambda__1___closed__8); +l_Lean_Parser_Term_doLetRec___elambda__1___closed__9 = _init_l_Lean_Parser_Term_doLetRec___elambda__1___closed__9(); +lean_mark_persistent(l_Lean_Parser_Term_doLetRec___elambda__1___closed__9); +l_Lean_Parser_Term_doLetRec___elambda__1___closed__10 = _init_l_Lean_Parser_Term_doLetRec___elambda__1___closed__10(); +lean_mark_persistent(l_Lean_Parser_Term_doLetRec___elambda__1___closed__10); l_Lean_Parser_Term_doLetRec___closed__1 = _init_l_Lean_Parser_Term_doLetRec___closed__1(); lean_mark_persistent(l_Lean_Parser_Term_doLetRec___closed__1); l_Lean_Parser_Term_doLetRec___closed__2 = _init_l_Lean_Parser_Term_doLetRec___closed__2(); diff --git a/stage0/stdlib/Lean/Parser/Extension.c b/stage0/stdlib/Lean/Parser/Extension.c index c895a440578..72b7e52e981 100644 --- a/stage0/stdlib/Lean/Parser/Extension.c +++ b/stage0/stdlib/Lean/Parser/Extension.c @@ -14,27 +14,29 @@ extern "C" { #endif lean_object* l_List_reverse___rarg(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5431____closed__2; uint8_t l_Lean_isRecCore(lean_object*, lean_object*); static lean_object* l_Lean_Parser_declareLeadingBuiltinParser___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_builtinTokenTable; static lean_object* l_Lean_Parser_parserOfStackFn___lambda__2___closed__5; lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_KVMap_setBool(lean_object*, lean_object*, uint8_t); +static lean_object* l_Lean_Parser_getParserAliasInfo___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__2; size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_addTokenConfig(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__5; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_foldlMAux___at___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add___spec__4___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_tokenAntiquotFn(lean_object*, lean_object*); static lean_object* l_Lean_Parser_declareLeadingBuiltinParser___closed__1; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_containsAtAux___at_Lean_Parser_isValidSyntaxNodeKind___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___lambda__2___closed__9; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__2; lean_object* l_Lean_stringToMessageData(lean_object*); lean_object* l_Lean_Parser_andthenInfo(lean_object*, lean_object*); lean_object* l_Lean_Parser_ParserState_mkError(lean_object*, lean_object*); static lean_object* l_Lean_Parser_ParserExtension_instInhabitedOLeanEntry___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_evalParserConstUnsafe(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_mkParserContext(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insert___at___private_Lean_Parser_Extension_0__Lean_Parser_addParserCategoryCore___spec__4(lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); @@ -53,60 +55,66 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Parser_addToken___spec__2__ lean_object* lean_array_uget(lean_object*, size_t); LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_io_error_to_string(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_getBinaryAlias___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_35____closed__3; LEAN_EXPORT lean_object* l_IO_ofExcept___at___private_Lean_Parser_Extension_0__Lean_Parser_addBuiltinParserCategory___spec__1___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_registerParserAttributeHook(lean_object*, lean_object*); extern lean_object* l_Lean_noConfusionExt; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5420____closed__2; +LEAN_EXPORT lean_object* l_Std_RBNode_find___at_Lean_Parser_getParserAliasInfo___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Parser_addParserTokens___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___closed__4; static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___lambda__3___closed__4; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__2; static lean_object* l_Lean_Parser_registerBuiltinParserAttribute___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_mkParserAttributeImpl___elambda__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_getTokenTable___boxed(lean_object*); LEAN_EXPORT uint8_t l_Lean_Parser_leadingIdentBehavior(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3895____closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_addParser(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_getParserAliasInfo(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__1; extern lean_object* l_Lean_Parser_categoryParserFnRef; static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___lambda__2___closed__12; LEAN_EXPORT lean_object* l_Lean_Parser_withOpenDeclFn(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_addBuiltinParserCategory___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_addBuiltinParserCategory___closed__2; -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_compileParserDescr_visit___lambda__2(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_getConstAlias___rarg___boxed(lean_object*, lean_object*, lean_object*); size_t lean_usize_sub(size_t, size_t); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5442____closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_builtinParserCategoriesRef; LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_withNamespaces(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); lean_object* lean_environment_find(lean_object*, lean_object*); static lean_object* l_List_forM___at___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add___spec__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_parserExtension; -lean_object* l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Parser_Extension_0__Lean_Parser_withNamespaces___spec__2(uint8_t, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_isRec___at___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_PersistentHashMap_foldlMAux_traverse___at_Lean_MetavarContext_getExprAssignmentDomain___spec__4___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_id___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_ParserAliasInfo_stackSz_x3f___default; lean_object* lean_st_ref_get(lean_object*, lean_object*); lean_object* l_Lean_Parser_Trie_empty(lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_35____closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_mkParserState(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_addToken___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__1; uint8_t lean_name_eq(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__3; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_6____closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5442____closed__2; extern lean_object* l_Lean_declRangeExt; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__4; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__4; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__11; static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___lambda__4___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_registerBuiltinDynamicParserAttribute(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_getConstAlias___rarg___closed__2; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_forM___at___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_Parser_ParserAliasInfo_autoGroupArgs___default(lean_object*); static lean_object* l_Lean_Parser_parserOfStackFn___lambda__2___closed__4; lean_object* l_Lean_getConstInfo___at_Lean_registerInitAttrUnsafe___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_mkInputContext(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_addBuiltinParser___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_addBuiltinParser(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); @@ -115,6 +123,7 @@ lean_object* lean_array_push(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); lean_object* l_Std_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___lambda__2___closed__11; +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_runParserAttributeHooks___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_findDeclarationRangesCore_x3f___at___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___spec__3___closed__1; @@ -126,23 +135,23 @@ static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinPars LEAN_EXPORT lean_object* l_Lean_Parser_categoryParserFnImpl(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkParserAttributeImpl___closed__2; LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_Parser_mkParserOfConstantUnsafe___spec__2___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__5; lean_object* l_Lean_Parser_checkPrecFn___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_leadingParserAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_right(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_Parser_getParserAliasInfo___boxed(lean_object*, lean_object*); lean_object* l_Lean_ScopedEnvExtension_addScopedEntry___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkParserState___closed__1; static lean_object* l_Lean_Parser_parserOfStackFn___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_isParserCategory___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__3; +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_foldlMAux___at_Lean_Parser_getSyntaxNodeKinds___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_evalInsideQuot___elambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_findDeclarationRangesCore_x3f___at___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); uint8_t l_Lean_NameMap_contains___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Parser_parserOfStackFn___lambda__2___closed__1; -extern lean_object* l_Lean_nameLitKind; +static lean_object* l_Lean_Parser_getParserAliasInfo___closed__1; static lean_object* l_Lean_Parser_getConstAlias___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_findDeclarationRangesCore_x3f___at___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_mkCategoryAntiquotParser(lean_object*); @@ -150,17 +159,21 @@ lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Parser_addToken___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ofExcept___at_Lean_Parser_addToken___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_addParserCategory(lean_object*, lean_object*, uint8_t); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__4; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_findAtAux___at_Lean_Parser_getCategory___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Parser_addToken___spec__3___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__6; static lean_object* l_Lean_Parser_declareTrailingBuiltinParser___closed__1; static lean_object* l_Lean_Parser_mkParserOfConstantUnsafe___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_getUnaryAlias(lean_object*); +uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__3; LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_registerAliasCore___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_PersistentHashMap_insert___at_Lean_Parser_SyntaxNodeKindSet_insert___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_throwParserCategoryAlreadyDefined(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__8; lean_object* l_Lean_mkAppN(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__6; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_foldlM___at_Lean_Parser_getSyntaxNodeKinds___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_ParserExtension_addEntryImpl___closed__1; size_t lean_uint64_to_usize(uint64_t); @@ -184,12 +197,13 @@ LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_Parser_mkParserOfConstantUnsafe LEAN_EXPORT lean_object* l_Lean_Parser_compileParserDescr_visit___lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_List_forM___at___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add___spec__1___closed__2; static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___lambda__2___closed__17; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__1; static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___lambda__4___closed__1; static lean_object* l_Lean_Parser_withOpenDeclFnCore___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_mkParserAttributeImpl___elambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_internal_parseQuotWithCurrentStage; lean_object* lean_array_fget(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__1; +lean_object* l_Lean_Syntax_isNatLit_x3f(lean_object*); uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_132_(uint8_t, uint8_t); @@ -198,16 +212,17 @@ lean_object* l_Lean_Parser_nonReservedSymbolFnAux(lean_object*, lean_object*, le static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_updateBuiltinTokens___closed__2; static lean_object* l_Lean_Parser_declareLeadingBuiltinParser___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_ParserExtension_addEntryImpl___closed__4; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__13; static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___lambda__3___closed__2; LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_Parser_mkParserOfConstantUnsafe___spec__2(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____closed__2; lean_object* lean_st_ref_take(lean_object*, lean_object*); lean_object* l_Lean_Parser_nodeInfo(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ScopedEnvExtension_addLocalEntry___rarg(lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_numLitKind; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____closed__1; static lean_object* l_Lean_Parser_categoryParserFnImpl___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5294____closed__3; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_findAux___at_Lean_Parser_getCategory___spec__2(lean_object*, size_t, lean_object*); LEAN_EXPORT uint8_t l_Std_PersistentHashMap_containsAtAux___at___private_Lean_Parser_Extension_0__Lean_Parser_addParserCategoryCore___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_getBinaryAlias___rarg___closed__1; @@ -220,7 +235,6 @@ LEAN_EXPORT lean_object* l_Std_RBNode_find___at_Lean_Parser_getAlias___spec__1__ LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAux___at___private_Lean_Parser_Extension_0__Lean_Parser_addParserCategoryCore___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_PersistentHashMap_insertAux___at___private_Lean_Parser_Extension_0__Lean_Parser_addParserCategoryCore___spec__5___closed__1; static lean_object* l_Lean_Parser_mkParserOfConstantUnsafe___closed__1; -extern lean_object* l_Lean_strLitKind; lean_object* l_Lean_Parser_nonReservedSymbolInfo(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Parser_ParserExtension_instInhabitedState; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAux___at___private_Lean_Parser_Extension_0__Lean_Parser_addParserCategoryCore___spec__5(lean_object*, size_t, size_t, lean_object*, lean_object*); @@ -229,61 +243,63 @@ LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Parser_Extension_0__Lea lean_object* l_Lean_Parser_ParserContext_resolveName(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_mkParserOfConstantUnsafe(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_registerAttributeOfBuilder(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__9; static lean_object* l_Lean_Parser_registerAlias___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_commandParser(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__3; lean_object* l_Lean_Parser_sepBy(lean_object*, lean_object*, lean_object*, uint8_t); static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___lambda__2___closed__2; -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3165_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_1928_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_1963_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3895_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5294_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5305_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5283_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5272_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_1951_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3313_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_4043_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_2021_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_1986_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5453_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5431_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5420_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5442_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_180_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_6_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_35_(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__5; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_containsAtAux___at___private_Lean_Parser_Extension_0__Lean_Parser_addParserCategoryCore___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Parser_getSyntaxNodeKinds___spec__3(lean_object*, size_t, size_t, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5294____closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__3; static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___lambda__2___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_ParserExtension_addEntryImpl(lean_object*, lean_object*); lean_object* l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_54____spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Std_PersistentHashMap_contains___at___private_Lean_Parser_Extension_0__Lean_Parser_addParserCategoryCore___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ParserExtension_State_categories___default; -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Parser_Extension_0__Lean_Parser_addTrailingParserAux___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_toString(lean_object*, uint8_t); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____lambda__2___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__5; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__8; lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___lambda__2___closed__8; lean_object* l_Lean_ResolveName_resolveNamespace(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__5; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____closed__3; LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__2; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_foldlMAux___at_Lean_Parser_getSyntaxNodeKinds___spec__2___lambda__1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5294____closed__2; +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Internal_isStage0(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_foldlMAux___at_Lean_Parser_getSyntaxNodeKinds___spec__2(lean_object*, lean_object*); lean_object* l_Lean_Parser_checkLhsPrecFn___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_throwUnknownParserCategory___rarg(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5305____closed__2; LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_ParserExtension_OLeanEntry_toEntry(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Name_quickCmp(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_instCoeForAllParserParserAliasValue(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__1; lean_object* l_Lean_Name_toExprAux(lean_object*); static lean_object* l_Lean_Parser_mkParserOfConstantUnsafe___closed__4; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__3; lean_object* l_Std_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___lambda__2___closed__19; LEAN_EXPORT lean_object* l_Lean_Parser_parserOfStackFn___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -292,6 +308,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_addToken(lean_object*, uint8_t, lean_obje LEAN_EXPORT lean_object* l_Lean_ofExcept___at_Lean_Parser_addToken___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_categoryParserFnImpl___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_registerBuiltinParserAttribute___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_registerAlias___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Attribute_Builtin_getPrio(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_runParserCategory___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_parserAliasesRef; @@ -301,12 +318,11 @@ static lean_object* l_Lean_Parser_mkParserOfConstantUnsafe___closed__8; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_find_x3f___at_Lean_Parser_getCategory___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_withNamespaces___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Parser_Extension_0__Lean_Parser_withNamespaces___spec__1(uint8_t, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5305____closed__1; LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Parser_Extension_0__Lean_Parser_addTrailingParserAux___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__2; static lean_object* l_Lean_Parser_mkParserAttributeImpl___closed__1; -extern lean_object* l_Lean_choiceKind; -extern lean_object* l_Lean_charLitKind; lean_object* l_Std_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_registerAlias___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_findAtAux___at_Lean_Parser_getCategory___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___lambda__3___closed__3; static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_throwParserCategoryAlreadyDefined___rarg___closed__2; @@ -315,24 +331,24 @@ LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_Parser static lean_object* l_Lean_Parser_ParserExtension_addEntryImpl___closed__2; static lean_object* l_Lean_Parser_registerParserAttributeHook___closed__1; static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___lambda__4___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5294____closed__4; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__4; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__6; static lean_object* l_Lean_Parser_ParserExtension_instInhabitedState___closed__2; size_t lean_usize_shift_left(size_t, size_t); static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___lambda__3___closed__5; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__5; lean_object* lean_eval_const(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_mkParserAttributeImpl(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_parserOfStackFn___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_ParserExtension_mkInitial(lean_object*); uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); static lean_object* l_Lean_Parser_ParserExtension_instInhabitedState___closed__1; +LEAN_EXPORT lean_object* l_Lean_Parser_parserAliases2infoRef; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAux_traverse___at___private_Lean_Parser_Extension_0__Lean_Parser_addParserCategoryCore___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_compileParserDescr(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_declareTrailingBuiltinParser(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Parser_withOpenDeclFnCore___spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_withOpenDeclFnCore___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_addParserTokens(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__7; uint32_t lean_string_utf8_get(lean_object*, lean_object*); static lean_object* l_Lean_findDeclarationRanges_x3f___at___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___spec__1___closed__1; static uint8_t l_Lean_Parser_isValidSyntaxNodeKind___closed__1; @@ -345,11 +361,9 @@ LEAN_EXPORT lean_object* l_Std_PersistentHashMap_foldlMAux_traverse___at___priva LEAN_EXPORT lean_object* l_Lean_Parser_evalInsideQuot___elambda__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_Trie_insert_loop___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_compileParserDescr_visit___lambda__1___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3895____closed__2; lean_object* l_Lean_registerBuiltinAttribute(lean_object*, lean_object*); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_parserOfStackFn___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__2; lean_object* l_Lean_throwError___at_Lean_AttributeImpl_erase___default___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at_Lean_Parser_runParserAttributeHooks___spec__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_withOpen(lean_object*); @@ -360,34 +374,36 @@ static lean_object* l_Lean_Parser_parserOfStackFn___lambda__1___closed__1; static lean_object* l_Lean_Parser_withOpenDeclFnCore___closed__6; static lean_object* l_Lean_Parser_mkParserOfConstantUnsafe___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_parserAlias2kindRef; +LEAN_EXPORT lean_object* l_Std_RBNode_find___at_Lean_Parser_getParserAliasInfo___spec__1(lean_object*, lean_object*); size_t lean_usize_mul(size_t, size_t); lean_object* l_Lean_FileMap_ofString(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____lambda__1___closed__1; lean_object* l_Lean_Parser_whitespace(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_mkParserAttributeImpl___elambda__2___boxed(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__7; static lean_object* l_Lean_Parser_registerAliasCore___rarg___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_registerAliasCore___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_getTokenTable(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_mkCategoryAntiquotParserFn(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_getParserPriority___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__1; LEAN_EXPORT uint8_t l_Std_PersistentHashMap_containsAtAux___at_Lean_Parser_isValidSyntaxNodeKind___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_isValidSyntaxNodeKind___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Parser_Extension_0__Lean_Parser_withNamespaces___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ParserExtension_instInhabitedEntry; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__2; -LEAN_EXPORT lean_object* l_Lean_Parser_registerAlias(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__5; +LEAN_EXPORT lean_object* l_Lean_Parser_registerAlias(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_isParserAlias___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Parser_addLeadingParser___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_addTrailingParser(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_type(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_mkParserAttributeImpl___elambda__2___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__3; lean_object* l_Lean_Parser_trailingNodeFn(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___lambda__2___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_addBuiltinLeadingParser(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5272____closed__4; lean_object* l_Lean_ScopedEnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*); size_t lean_usize_land(size_t, size_t); static lean_object* l_Lean_Parser_withOpenDeclFnCore___closed__5; @@ -397,7 +413,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_getUnaryAlias___rarg___boxed(lean_object* LEAN_EXPORT lean_object* l_Lean_Parser_getParserPriority___boxed(lean_object*); static lean_object* l_Lean_Parser_getParserPriority___closed__1; static lean_object* l_Lean_Parser_getUnaryAlias___rarg___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_mkParserAttributeImpl___elambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_findDeclarationRanges_x3f___at___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ensureUnaryParserAlias(lean_object*, lean_object*); @@ -407,15 +422,14 @@ LEAN_EXPORT lean_object* l_Lean_Parser_parserOfStackFn___lambda__2(lean_object*, lean_object* l_Lean_Parser_sepBy1(lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Parser_getConstAlias___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_declareLeadingBuiltinParser(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_compileParserDescr_visit(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_registerParserCategory(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); lean_object* l_Lean_Parser_nodeWithAntiquot(lean_object*, lean_object*, lean_object*, uint8_t); -extern lean_object* l_Lean_identKind; LEAN_EXPORT lean_object* l_Lean_Parser_getAlias___rarg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____lambda__1___closed__2; lean_object* l_Lean_Parser_trailingLoop(lean_object*, lean_object*, lean_object*); uint8_t lean_uint32_dec_eq(uint32_t, uint32_t); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__7; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____lambda__2___closed__1; static lean_object* l_Lean_Parser_throwUnknownParserCategory___rarg___closed__2; LEAN_EXPORT uint8_t l_Std_PersistentHashMap_containsAux___at___private_Lean_Parser_Extension_0__Lean_Parser_addParserCategoryCore___spec__2(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAux_traverse___at___private_Lean_Parser_Extension_0__Lean_Parser_addParserCategoryCore___spec__6(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -429,16 +443,15 @@ LEAN_EXPORT lean_object* l_Lean_Parser_registerAliasCore___rarg___lambda__1___bo LEAN_EXPORT lean_object* l_Lean_Parser_mkParserAttributeImpl___elambda__1___rarg___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getNumArgs(lean_object*); static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___lambda__2___closed__5; -extern lean_object* l_Lean_scientificLitKind; LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____lambda__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_throwParserCategoryAlreadyDefined___rarg(lean_object*); static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___lambda__3___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__14; LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____lambda__1(lean_object*); static lean_object* l_List_forM___at___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add___spec__1___closed__1; static lean_object* l_List_forM___at___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add___spec__1___closed__4; static lean_object* l_Lean_Parser_withOpenDeclFnCore___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_ParserExtension_instInhabitedOLeanEntry; LEAN_EXPORT uint8_t l_Lean_Parser_isParserCategory(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -448,10 +461,10 @@ LEAN_EXPORT lean_object* l_IO_ofExcept___at___private_Lean_Parser_Extension_0__L lean_object* l_Lean_Parser_categoryParser(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ensureConstantParserAlias(lean_object*, lean_object*); static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___lambda__2___closed__14; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____lambda__2___closed__2; lean_object* l_Lean_Syntax_getArgs(lean_object*); lean_object* l_Lean_Name_append(lean_object*, lean_object*); static lean_object* l_Lean_Parser_categoryParserFnImpl___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5272____closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_mkParserAttributeImpl___elambda__1___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Std_PersistentHashMap_containsAux___at_Lean_Parser_isValidSyntaxNodeKind___spec__2(lean_object*, size_t, lean_object*); lean_object* l_Lean_Syntax_getKind(lean_object*); @@ -460,6 +473,7 @@ lean_object* l_Lean_findDocString_x3f(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_ParserExtension_addEntryImpl___closed__5; static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_updateBuiltinTokens___closed__1; lean_object* l_Lean_registerScopedEnvExtensionUnsafe___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_runParserCategory(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_addLeadingParser(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_ParserExtension_addEntryImpl___closed__3; @@ -467,46 +481,48 @@ static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinPars uint8_t lean_is_aux_recursor(lean_object*, lean_object*); lean_object* l_Lean_Parser_symbolInfo(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_parserAttributeHooks; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5272____closed__1; +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Parser_withOpenDeclFnCore___spec__1(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_updateBuiltinTokens(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_ParserAliasInfo_autoGroupArgs___default___boxed(lean_object*); lean_object* l_Lean_Parser_orelseFnCore(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); extern lean_object* l_Lean_instInhabitedDeclarationRanges; LEAN_EXPORT lean_object* l_Lean_Parser_ParserExtension_State_kinds___default; extern lean_object* l_Lean_Parser_epsilonInfo; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__6; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__6; LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_Parser_mkParserOfConstantUnsafe___spec__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_getParserPriority(lean_object*); +static lean_object* l_Lean_Parser_ParserAliasInfo_stackSz_x3f___default___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_builtinSyntaxNodeKindSetRef; LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_foldlMAux___at___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add___spec__4___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_registerBuiltinNodeKind(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_isParserAlias(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_getSyntaxKindOfParserAlias_x3f___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_mkParserState___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ParserExtension_Entry_toOLeanEntry___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_getBinaryAlias(lean_object*); lean_object* l_Lean_declareBuiltin(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_findAux___at_Lean_Parser_getCategory___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isRec___at___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5283____closed__2; lean_object* l_Lean_Name_getPrefix(lean_object*); static lean_object* l_Std_PersistentHashMap_foldlMAux___at___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add___spec__4___closed__1; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_findDeclarationRanges_x3f___at___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_runParserAttributeHooks(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5420____closed__3; static lean_object* l_Lean_Parser_ParserExtension_instInhabitedEntry___closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add___spec__5(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_withOpenFn___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____closed__3; static lean_object* l_Lean_Parser_mkParserOfConstantUnsafe___closed__7; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Parser_Extension_0__Lean_Parser_withNamespaces___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_addTokenConfig___closed__1; LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Parser_addToken___spec__3(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_declareLeadingBuiltinParser___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5283____closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_getSyntaxKindOfParserAlias_x3f(lean_object*, lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); lean_object* l_String_trim(lean_object*); @@ -524,10 +540,10 @@ LEAN_EXPORT lean_object* l_Lean_Parser_mkParserOfConstant(lean_object*, lean_obj LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_addTrailingParserAux(lean_object*, lean_object*, lean_object*); lean_object* lean_io_initializing(lean_object*); LEAN_EXPORT lean_object* l_Std_RBNode_find___at_Lean_Parser_getSyntaxKindOfParserAlias_x3f___spec__1___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5442____closed__4; static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___lambda__2___closed__7; static lean_object* l_Lean_Parser_getConstAlias___rarg___closed__3; static lean_object* l_Lean_Parser_withOpenDeclFnCore___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__1; lean_object* l_Lean_ScopedEnvExtension_addEntry___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_getSyntaxNodeKinds(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_addBuiltinTrailingParser(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -538,28 +554,28 @@ LEAN_EXPORT lean_object* l_Lean_Parser_ensureBinaryParserAlias(lean_object*, lea LEAN_EXPORT lean_object* l_Lean_Parser_ParserExtension_State_tokens___default; lean_object* l_Lean_Parser_ParserState_toErrorMsg(lean_object*, lean_object*); lean_object* l_Lean_Attribute_Builtin_ensureNoArgs(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_Parser_mkParserOfConstantUnsafe___spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_contains___at___private_Lean_Parser_Extension_0__Lean_Parser_addParserCategoryCore___spec__1___boxed(lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_parserOfStackFn(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__8; LEAN_EXPORT lean_object* l_Lean_Parser_registerAliasCore___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_Parser_mkParserOfConstantUnsafe___spec__3___boxed(lean_object*, lean_object*); static lean_object* l_Std_PersistentHashMap_foldlMAux___at_Lean_Parser_getSyntaxNodeKinds___spec__2___closed__1; static lean_object* l_Lean_Parser_getParserPriority___closed__3; lean_object* l_Lean_mkNatLit(lean_object*); lean_object* l_Lean_mkStrLit(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__6; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_leadingIdentBehavior___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___lambda__2___closed__10; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5420____closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_registerAliasCore___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_parserOfStackFn___lambda__2___closed__6; static lean_object* l_Lean_Parser_getParserPriority___closed__5; LEAN_EXPORT lean_object* l_Std_RBNode_find___at_Lean_Parser_getSyntaxKindOfParserAlias_x3f___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5453____closed__2; static lean_object* l_Lean_Parser_registerAliasCore___rarg___lambda__2___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_addParserCategoryCore(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_4043____closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_registerAliasCore___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_ParserExtension_mkInitial___closed__1; lean_object* lean_usize_to_nat(size_t); @@ -567,43 +583,48 @@ LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAtCollisionNodeAux___at__ lean_object* l_Lean_registerAttributeImplBuilder(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_ParserState_mkUnexpectedError(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_symbolFnAux(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5420____closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5453____closed__1; LEAN_EXPORT lean_object* l_List_mapTRAux___at_Lean_Parser_addLeadingParser___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_withOpenFn(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkParserOfConstantUnsafe___closed__3; -uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__12; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_containsAux___at_Lean_Parser_isValidSyntaxNodeKind___spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkParserOfConstantUnsafe___closed__5; extern lean_object* l_Lean_builtinDeclRanges; uint8_t lean_string_utf8_at_end(lean_object*, lean_object*); lean_object* l_Lean_mkConst(lean_object*, lean_object*); static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___lambda__2___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____closed__2; LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____lambda__1(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5442____closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_getAlias___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Std_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ParserExtension_Entry_toOLeanEntry(lean_object*); lean_object* l_Lean_ScopedEnvExtension_activateScoped___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__10; static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_withOpenDeclFnCore(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5272____closed__3; lean_object* l_Lean_Parser_Trie_find_x3f_loop___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_getCategory(lean_object*, lean_object*); static lean_object* l_Lean_Parser_getConstAlias___rarg___closed__4; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5431____closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_instCoeForAllParserParserAliasValue__1(lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_foldlMAux_traverse___at___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add___spec__6___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__3; static lean_object* l_Lean_Parser_registerBuiltinNodeKind___closed__1; static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___lambda__2___closed__15; static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___lambda__2___closed__18; lean_object* l_Std_PersistentHashMap_mkCollisionNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_4043____closed__1; LEAN_EXPORT lean_object* l_Std_RBNode_find___at_Lean_Parser_getAlias___spec__1(lean_object*); uint8_t lean_string_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_evalInsideQuot(lean_object*, lean_object*); static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_addTokenConfig___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__7; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_contains___at_Lean_Parser_isValidSyntaxNodeKind___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_registerParserCategory___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__7; lean_object* l_Lean_Parser_setLhsPrecFn___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_andthenFn(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_6____closed__1() { @@ -738,41 +759,167 @@ return x_13; } } } +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("choice", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("ident", 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__3; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("str", 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__5; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("num", 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__7; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("scientific", 10); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__9; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("char", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__11; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__13() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("name", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__13; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74_(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_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; 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; -x_2 = l_Lean_choiceKind; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__2; x_3 = l_Lean_Parser_registerBuiltinNodeKind(x_2, x_1); x_4 = lean_ctor_get(x_3, 1); lean_inc(x_4); lean_dec(x_3); -x_5 = l_Lean_identKind; +x_5 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__4; x_6 = l_Lean_Parser_registerBuiltinNodeKind(x_5, x_4); x_7 = lean_ctor_get(x_6, 1); lean_inc(x_7); lean_dec(x_6); -x_8 = l_Lean_strLitKind; +x_8 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__6; x_9 = l_Lean_Parser_registerBuiltinNodeKind(x_8, x_7); x_10 = lean_ctor_get(x_9, 1); lean_inc(x_10); lean_dec(x_9); -x_11 = l_Lean_numLitKind; +x_11 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__8; x_12 = l_Lean_Parser_registerBuiltinNodeKind(x_11, x_10); x_13 = lean_ctor_get(x_12, 1); lean_inc(x_13); lean_dec(x_12); -x_14 = l_Lean_scientificLitKind; +x_14 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__10; x_15 = l_Lean_Parser_registerBuiltinNodeKind(x_14, x_13); x_16 = lean_ctor_get(x_15, 1); lean_inc(x_16); lean_dec(x_15); -x_17 = l_Lean_charLitKind; +x_17 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__12; x_18 = l_Lean_Parser_registerBuiltinNodeKind(x_17, x_16); x_19 = lean_ctor_get(x_18, 1); lean_inc(x_19); lean_dec(x_18); -x_20 = l_Lean_nameLitKind; +x_20 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__14; x_21 = l_Lean_Parser_registerBuiltinNodeKind(x_20, x_19); return x_21; } @@ -4319,7 +4466,52 @@ lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_1928_(lean_object* x_1) { +static lean_object* _init_l_Lean_Parser_ParserAliasInfo_stackSz_x3f___default___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_ParserAliasInfo_stackSz_x3f___default() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Parser_ParserAliasInfo_stackSz_x3f___default___closed__1; +return x_1; +} +} +LEAN_EXPORT uint8_t l_Lean_Parser_ParserAliasInfo_autoGroupArgs___default(lean_object* x_1) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +uint8_t x_2; +x_2 = 0; +return x_2; +} +else +{ +uint8_t x_3; +x_3 = 1; +return x_3; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Parser_ParserAliasInfo_autoGroupArgs___default___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Lean_Parser_ParserAliasInfo_autoGroupArgs___default(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_1951_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; @@ -4345,7 +4537,7 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_1963_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_1986_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; @@ -4371,122 +4563,307 @@ return x_7; } } } -static lean_object* _init_l_Lean_Parser_registerAlias___closed__1() { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_2021_(lean_object* x_1) { _start: { -lean_object* x_1; -x_1 = l_Lean_Parser_parserAliasesRef; -return x_1; +lean_object* x_2; lean_object* x_3; uint8_t x_4; +x_2 = lean_box(0); +x_3 = lean_st_mk_ref(x_2, x_1); +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_3, 1); +lean_inc(x_6); +lean_inc(x_5); +lean_dec(x_3); +x_7 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_7, 0, x_5); +lean_ctor_set(x_7, 1, x_6); +return x_7; } } -static lean_object* _init_l_Lean_Parser_registerAlias___closed__2() { +} +LEAN_EXPORT lean_object* l_Std_RBNode_find___at_Lean_Parser_getParserAliasInfo___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 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; lean_object* x_7; uint8_t x_8; +x_4 = lean_ctor_get(x_1, 0); +x_5 = lean_ctor_get(x_1, 1); +x_6 = lean_ctor_get(x_1, 2); +x_7 = lean_ctor_get(x_1, 3); +x_8 = l_Lean_Name_quickCmp(x_2, x_5); +switch (x_8) { +case 0: +{ +x_1 = x_4; +goto _start; +} +case 1: +{ +lean_object* x_10; +lean_inc(x_6); +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_6); +return x_10; +} +default: +{ +x_1 = x_7; +goto _start; +} +} +} +} +} +static lean_object* _init_l_Lean_Parser_getParserAliasInfo___closed__1() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_parserAlias2kindRef; +x_1 = l_Lean_Parser_parserAliases2infoRef; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_registerAlias(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l_Lean_Parser_getParserAliasInfo___closed__2() { _start: { -lean_object* x_5; lean_object* x_6; -x_5 = l_Lean_Parser_registerAlias___closed__1; -lean_inc(x_1); -x_6 = l_Lean_Parser_registerAliasCore___rarg(x_5, x_1, x_2, x_4); -if (lean_obj_tag(x_6) == 0) +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l_Lean_Parser_ParserAliasInfo_stackSz_x3f___default___closed__1; +x_2 = 1; +x_3 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Parser_getParserAliasInfo(lean_object* x_1, lean_object* x_2) { +_start: { -if (lean_obj_tag(x_3) == 0) +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = l_Lean_Parser_getParserAliasInfo___closed__1; +x_4 = lean_st_ref_get(x_3, x_2); +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) { -uint8_t x_7; -lean_dec(x_1); -x_7 = !lean_is_exclusive(x_6); -if (x_7 == 0) +lean_object* x_6; lean_object* x_7; +x_6 = lean_ctor_get(x_4, 0); +x_7 = l_Std_RBNode_find___at_Lean_Parser_getParserAliasInfo___spec__1(x_6, x_1); +lean_dec(x_6); +if (lean_obj_tag(x_7) == 0) { -lean_object* x_8; lean_object* x_9; -x_8 = lean_ctor_get(x_6, 0); -lean_dec(x_8); -x_9 = lean_box(0); -lean_ctor_set(x_6, 0, x_9); -return x_6; +lean_object* x_8; +x_8 = l_Lean_Parser_getParserAliasInfo___closed__2; +lean_ctor_set(x_4, 0, x_8); +return x_4; +} +else +{ +lean_object* x_9; +x_9 = lean_ctor_get(x_7, 0); +lean_inc(x_9); +lean_dec(x_7); +lean_ctor_set(x_4, 0, x_9); +return x_4; +} } else { lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_ctor_get(x_6, 1); +x_10 = lean_ctor_get(x_4, 0); +x_11 = lean_ctor_get(x_4, 1); +lean_inc(x_11); lean_inc(x_10); -lean_dec(x_6); -x_11 = lean_box(0); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_10); -return x_12; -} +lean_dec(x_4); +x_12 = l_Std_RBNode_find___at_Lean_Parser_getParserAliasInfo___spec__1(x_10, x_1); +lean_dec(x_10); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; lean_object* x_14; +x_13 = l_Lean_Parser_getParserAliasInfo___closed__2; +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_11); +return x_14; } else { -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; lean_object* x_20; uint8_t x_21; -x_13 = lean_ctor_get(x_6, 1); -lean_inc(x_13); +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_12, 0); +lean_inc(x_15); +lean_dec(x_12); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_11); +return x_16; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_RBNode_find___at_Lean_Parser_getParserAliasInfo___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_RBNode_find___at_Lean_Parser_getParserAliasInfo___spec__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Parser_getParserAliasInfo___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Parser_getParserAliasInfo(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Parser_registerAlias___lambda__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; uint8_t x_11; +x_5 = l_Lean_Parser_getParserAliasInfo___closed__1; +x_6 = lean_st_ref_take(x_5, x_4); +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_6, 1); +lean_inc(x_8); lean_dec(x_6); -x_14 = lean_ctor_get(x_3, 0); -lean_inc(x_14); -lean_dec(x_3); -x_15 = l_Lean_Parser_registerAlias___closed__2; -x_16 = lean_st_ref_take(x_15, x_13); -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec(x_16); -x_19 = l_Std_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(x_17, x_1, x_14); -x_20 = lean_st_ref_set(x_15, x_19, x_18); -x_21 = !lean_is_exclusive(x_20); -if (x_21 == 0) +x_9 = l_Std_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(x_7, x_1, x_2); +x_10 = lean_st_ref_set(x_5, x_9, x_8); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) { -return x_20; +return x_10; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_20, 0); -x_23 = lean_ctor_get(x_20, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_20); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_23); -return x_24; +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_10); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +return x_14; } } } +static lean_object* _init_l_Lean_Parser_registerAlias___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Parser_parserAliasesRef; +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_registerAlias___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Parser_parserAlias2kindRef; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Parser_registerAlias(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; +x_6 = l_Lean_Parser_registerAlias___closed__1; +lean_inc(x_1); +x_7 = l_Lean_Parser_registerAliasCore___rarg(x_6, x_1, x_2, x_5); +if (lean_obj_tag(x_7) == 0) +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_ctor_get(x_7, 1); +lean_inc(x_8); +lean_dec(x_7); +x_9 = lean_box(0); +x_10 = l_Lean_Parser_registerAlias___lambda__1(x_1, x_4, x_9, x_8); +return x_10; +} else { -uint8_t x_25; +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; lean_object* x_20; lean_object* x_21; +x_11 = lean_ctor_get(x_7, 1); +lean_inc(x_11); +lean_dec(x_7); +x_12 = lean_ctor_get(x_3, 0); +lean_inc(x_12); +lean_dec(x_3); +x_13 = l_Lean_Parser_registerAlias___closed__2; +x_14 = lean_st_ref_take(x_13, x_11); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +lean_inc(x_1); +x_17 = l_Std_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(x_15, x_1, x_12); +x_18 = lean_st_ref_set(x_13, x_17, x_16); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = l_Lean_Parser_registerAlias___lambda__1(x_1, x_4, x_19, x_20); +lean_dec(x_19); +return x_21; +} +} +else +{ +uint8_t x_22; +lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_25 = !lean_is_exclusive(x_6); -if (x_25 == 0) +x_22 = !lean_is_exclusive(x_7); +if (x_22 == 0) { -return x_6; +return x_7; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_6, 0); -x_27 = lean_ctor_get(x_6, 1); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_6); -x_28 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -return x_28; +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_7, 0); +x_24 = lean_ctor_get(x_7, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_7); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; } } } } +LEAN_EXPORT lean_object* l_Lean_Parser_registerAlias___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_registerAlias___lambda__1(x_1, x_2, x_3, x_4); +lean_dec(x_3); +return x_5; +} +} LEAN_EXPORT lean_object* l_Lean_Parser_instCoeParserParserAliasValue(lean_object* x_1) { _start: { @@ -6717,7 +7094,7 @@ x_6 = l_Lean_Parser_mkParserOfConstantUnsafe(x_2, x_5, x_3, x_4); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3165_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3313_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; @@ -6895,7 +7272,7 @@ x_8 = l_Lean_Parser_runParserAttributeHooks(x_1, x_2, x_7, x_4, x_5, x_6); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____lambda__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____lambda__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_7; @@ -6940,7 +7317,7 @@ return x_15; } } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____lambda__2___closed__1() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____lambda__2___closed__1() { _start: { lean_object* x_1; @@ -6948,25 +7325,25 @@ x_1 = lean_mk_string_from_bytes("attribute cannot be erased", 26); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____lambda__2___closed__2() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____lambda__2___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____lambda__2___closed__1; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____lambda__2___closed__1; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____lambda__2(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; -x_5 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____lambda__2___closed__2; +x_5 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____lambda__2___closed__2; x_6 = l_Lean_throwError___at_Lean_AttributeImpl_erase___default___spec__1(x_5, x_2, x_3, x_4); return x_6; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__1() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__1() { _start: { lean_object* x_1; @@ -6974,17 +7351,17 @@ x_1 = lean_mk_string_from_bytes("runBuiltinParserAttributeHooks", 30); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__2() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__1; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__3() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__3() { _start: { lean_object* x_1; @@ -6992,12 +7369,12 @@ x_1 = lean_mk_string_from_bytes("explicitly run hooks normally activated by buil return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__4() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__4() { _start: { lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__2; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__3; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__2; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__3; x_3 = 0; x_4 = lean_alloc_ctor(0, 2, 1); lean_ctor_set(x_4, 0, x_1); @@ -7006,29 +7383,29 @@ lean_ctor_set_uint8(x_4, sizeof(void*)*2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__5() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__5() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____lambda__1___boxed), 6, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____lambda__1___boxed), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__6() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__6() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____lambda__2___boxed), 4, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____lambda__2___boxed), 4, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__7() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__4; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__5; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__6; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__4; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__5; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__6; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -7036,37 +7413,37 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__7; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__7; x_3 = l_Lean_registerBuiltinAttribute(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____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_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____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) { _start: { uint8_t x_7; lean_object* x_8; x_7 = lean_unbox(x_3); lean_dec(x_3); -x_8 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____lambda__1(x_1, x_2, x_7, x_4, x_5, x_6); +x_8 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____lambda__1(x_1, x_2, x_7, x_4, x_5, x_6); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____lambda__2___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_initFn____x40_Lean_Parser_Extension___hyg_3246____lambda__2(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____lambda__2(x_1, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____lambda__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____lambda__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_7; @@ -7111,7 +7488,7 @@ return x_15; } } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__1() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__1() { _start: { lean_object* x_1; @@ -7119,17 +7496,17 @@ x_1 = lean_mk_string_from_bytes("runParserAttributeHooks", 23); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__2() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__1; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__3() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__3() { _start: { lean_object* x_1; @@ -7137,12 +7514,12 @@ x_1 = lean_mk_string_from_bytes("explicitly run hooks normally activated by pars return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__4() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__4() { _start: { lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__2; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__3; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__2; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__3; x_3 = 0; x_4 = lean_alloc_ctor(0, 2, 1); lean_ctor_set(x_4, 0, x_1); @@ -7151,21 +7528,21 @@ lean_ctor_set_uint8(x_4, sizeof(void*)*2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__5() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__5() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____lambda__1___boxed), 6, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____lambda__1___boxed), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__6() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__4; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__5; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__6; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__4; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__5; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__6; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -7173,22 +7550,22 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__6; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__6; x_3 = l_Lean_registerBuiltinAttribute(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____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_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____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) { _start: { uint8_t x_7; lean_object* x_8; x_7 = lean_unbox(x_3); lean_dec(x_3); -x_8 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____lambda__1(x_1, x_2, x_7, x_4, x_5, x_6); +x_8 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____lambda__1(x_1, x_2, x_7, x_4, x_5, x_6); return x_8; } } @@ -7338,7 +7715,7 @@ return x_36; } } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__1() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__1() { _start: { lean_object* x_1; @@ -7346,17 +7723,17 @@ x_1 = lean_mk_string_from_bytes("parserExt", 9); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__2() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__1; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__3() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__3() { _start: { lean_object* x_1; @@ -7364,7 +7741,7 @@ x_1 = lean_alloc_closure((void*)(l___private_Lean_Parser_Extension_0__Lean_Parse return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__4() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__4() { _start: { lean_object* x_1; @@ -7372,7 +7749,7 @@ x_1 = lean_alloc_closure((void*)(l___private_Lean_Parser_Extension_0__Lean_Parse return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__5() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__5() { _start: { lean_object* x_1; @@ -7380,7 +7757,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_ParserExtension_Entry_toOLeanEntr return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__6() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__6() { _start: { lean_object* x_1; @@ -7388,7 +7765,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_ParserExtension_addEntryImpl), 2, return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__7() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__7() { _start: { lean_object* x_1; @@ -7396,16 +7773,16 @@ x_1 = lean_alloc_closure((void*)(l_id___rarg___boxed), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__8() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__8() { _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; lean_object* x_7; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__2; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__3; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__4; -x_4 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__5; -x_5 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__6; -x_6 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__7; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__2; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__3; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__4; +x_4 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__5; +x_5 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__6; +x_6 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__7; x_7 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_7, 0, x_1); lean_ctor_set(x_7, 1, x_2); @@ -7416,11 +7793,11 @@ lean_ctor_set(x_7, 5, x_6); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__8; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__8; x_3 = l_Lean_registerScopedEnvExtensionUnsafe___rarg(x_2, x_1); return x_3; } @@ -7584,7 +7961,7 @@ return x_21; } } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__1() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__1() { _start: { lean_object* x_1; @@ -7592,17 +7969,17 @@ x_1 = lean_mk_string_from_bytes("internal", 8); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__2() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__1; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__3() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__3() { _start: { lean_object* x_1; @@ -7610,17 +7987,17 @@ x_1 = lean_mk_string_from_bytes("parseQuotWithCurrentStage", 25); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__4() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__2; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__3; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__2; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__5() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__5() { _start: { lean_object* x_1; @@ -7628,13 +8005,13 @@ x_1 = lean_mk_string_from_bytes("(Lean bootstrapping) use parsers from the curre return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__6() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__6() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = 0; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__1; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__5; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__1; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__5; x_4 = lean_box(x_1); x_5 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_5, 0, x_4); @@ -7643,12 +8020,12 @@ lean_ctor_set(x_5, 2, x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__4; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__6; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__4; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__6; x_4 = l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_54____spec__1(x_2, x_3, x_1); return x_4; } @@ -8035,7 +8412,7 @@ lean_inc(x_23); x_35 = l_Lean_Parser_leadingParserAux(x_12, x_23, x_24, x_2, x_3); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_34); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_34); lean_dec(x_36); if (x_37 == 0) { @@ -8059,7 +8436,7 @@ lean_inc(x_2); x_40 = l_Lean_Parser_orelseFnCore(x_25, x_27, x_39, x_2, x_3); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_34); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_34); lean_dec(x_41); if (x_42 == 0) { @@ -8078,7 +8455,7 @@ return x_43; } } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3895____closed__1() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_4043____closed__1() { _start: { lean_object* x_1; @@ -8086,7 +8463,7 @@ x_1 = l_Lean_Parser_categoryParserFnRef; return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3895____closed__2() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_4043____closed__2() { _start: { lean_object* x_1; @@ -8094,12 +8471,12 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_categoryParserFnImpl), 3, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3895_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_4043_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; uint8_t x_5; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3895____closed__1; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3895____closed__2; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_4043____closed__1; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_4043____closed__2; x_4 = lean_st_ref_set(x_2, x_3, x_1); x_5 = !lean_is_exclusive(x_4); if (x_5 == 0) @@ -9096,7 +9473,7 @@ x_12 = l_Lean_Parser_categoryParserFnImpl(x_2, x_9, x_11); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -9316,35 +9693,34 @@ return x_7; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; +lean_object* x_8; lean_object* x_9; x_8 = l_Lean_Syntax_getArg(x_1, x_3); -x_9 = l_Lean_numLitKind; -x_10 = l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(x_9, x_8); +x_9 = l_Lean_Syntax_isNatLit_x3f(x_8); lean_dec(x_8); -if (lean_obj_tag(x_10) == 0) +if (lean_obj_tag(x_9) == 0) { -lean_object* x_11; -x_11 = l_Lean_Parser_getParserPriority___closed__4; -return x_11; +lean_object* x_10; +x_10 = l_Lean_Parser_getParserPriority___closed__4; +return x_10; } else { -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_10, 0); -lean_inc(x_12); -lean_dec(x_10); -x_13 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_13, 0, x_12); -return x_13; +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_9, 0); +lean_inc(x_11); +lean_dec(x_9); +x_12 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_12, 0, x_11); +return x_12; } } } else { -lean_object* x_14; +lean_object* x_13; lean_dec(x_2); -x_14 = l_Lean_Parser_getParserPriority___closed__5; -return x_14; +x_13 = l_Lean_Parser_getParserPriority___closed__5; +return x_13; } } } @@ -10855,7 +11231,7 @@ lean_ctor_set_uint8(x_9, sizeof(void*)*2, x_8); x_10 = lean_alloc_closure((void*)(l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___boxed), 8, 2); lean_closure_set(x_10, 0, x_1); lean_closure_set(x_10, 1, x_2); -x_11 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__6; +x_11 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__6; x_12 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_12, 0, x_9); lean_ctor_set(x_12, 1, x_10); @@ -11869,7 +12245,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_mkParserAttributeImpl___elambda__1___rarg _start: { lean_object* x_4; lean_object* x_5; -x_4 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____lambda__2___closed__2; +x_4 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____lambda__2___closed__2; x_5 = l_Lean_throwError___at_Lean_AttributeImpl_erase___default___spec__1(x_4, x_1, x_2, x_3); return x_5; } @@ -11981,7 +12357,7 @@ x_5 = l_Lean_registerBuiltinAttribute(x_4, x_3); return x_5; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____lambda__1___closed__1() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____lambda__1___closed__1() { _start: { lean_object* x_1; @@ -11989,23 +12365,23 @@ x_1 = lean_mk_string_from_bytes("invalid parser attribute implementation builder return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____lambda__1___closed__2() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____lambda__1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____lambda__1___closed__1; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____lambda__1___closed__1; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____lambda__1(lean_object* x_1) { _start: { if (lean_obj_tag(x_1) == 0) { lean_object* x_2; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____lambda__1___closed__2; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____lambda__1___closed__2; return x_2; } else @@ -12023,7 +12399,7 @@ if (lean_obj_tag(x_4) == 0) { lean_object* x_5; lean_dec(x_3); -x_5 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____lambda__1___closed__2; +x_5 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____lambda__1___closed__2; return x_5; } else @@ -12057,7 +12433,7 @@ lean_object* x_12; lean_dec(x_7); lean_dec(x_6); lean_dec(x_3); -x_12 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____lambda__1___closed__2; +x_12 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____lambda__1___closed__2; return x_12; } } @@ -12067,7 +12443,7 @@ lean_object* x_13; lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); -x_13 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____lambda__1___closed__2; +x_13 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____lambda__1___closed__2; return x_13; } } @@ -12077,13 +12453,13 @@ else lean_object* x_14; lean_dec(x_3); lean_dec(x_1); -x_14 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____lambda__1___closed__2; +x_14 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____lambda__1___closed__2; return x_14; } } } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____closed__1() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____closed__1() { _start: { lean_object* x_1; @@ -12091,30 +12467,30 @@ x_1 = lean_mk_string_from_bytes("parserAttr", 10); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____closed__2() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____closed__1; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____closed__3() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____closed__3() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____closed__2; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____closed__3; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____closed__2; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____closed__3; x_4 = l_Lean_registerAttributeImplBuilder(x_2, x_3, x_1); return x_4; } @@ -12146,7 +12522,7 @@ lean_ctor_set(x_13, 1, x_12); x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_10); lean_ctor_set(x_14, 1, x_13); -x_15 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____closed__2; +x_15 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____closed__2; x_16 = l_Lean_registerAttributeOfBuilder(x_8, x_15, x_14, x_9); return x_16; } @@ -12186,7 +12562,7 @@ x_7 = l_Lean_Parser_registerParserCategory(x_1, x_2, x_3, x_6, x_5); return x_7; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5272____closed__1() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5420____closed__1() { _start: { lean_object* x_1; @@ -12194,17 +12570,17 @@ x_1 = lean_mk_string_from_bytes("builtinTermParser", 17); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5272____closed__2() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5420____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5272____closed__1; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5420____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5272____closed__3() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5420____closed__3() { _start: { lean_object* x_1; @@ -12212,28 +12588,28 @@ x_1 = lean_mk_string_from_bytes("term", 4); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5272____closed__4() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5420____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5272____closed__3; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5420____closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5272_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5420_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; lean_object* x_5; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5272____closed__2; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5272____closed__4; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5420____closed__2; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5420____closed__4; x_4 = 0; x_5 = l_Lean_Parser_registerBuiltinParserAttribute(x_2, x_3, x_4, x_1); return x_5; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5283____closed__1() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5431____closed__1() { _start: { lean_object* x_1; @@ -12241,27 +12617,27 @@ x_1 = lean_mk_string_from_bytes("termParser", 10); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5283____closed__2() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5431____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5283____closed__1; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5431____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5283_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5431_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5283____closed__2; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5272____closed__4; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5431____closed__2; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5420____closed__4; x_4 = l_Lean_Parser_registerBuiltinDynamicParserAttribute(x_2, x_3, x_1); return x_4; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5294____closed__1() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5442____closed__1() { _start: { lean_object* x_1; @@ -12269,17 +12645,17 @@ x_1 = lean_mk_string_from_bytes("builtinCommandParser", 20); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5294____closed__2() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5442____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5294____closed__1; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5442____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5294____closed__3() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5442____closed__3() { _start: { lean_object* x_1; @@ -12287,28 +12663,28 @@ x_1 = lean_mk_string_from_bytes("command", 7); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5294____closed__4() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5442____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5294____closed__3; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5442____closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5294_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5442_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; lean_object* x_5; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5294____closed__2; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5294____closed__4; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5442____closed__2; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5442____closed__4; x_4 = 0; x_5 = l_Lean_Parser_registerBuiltinParserAttribute(x_2, x_3, x_4, x_1); return x_5; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5305____closed__1() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5453____closed__1() { _start: { lean_object* x_1; @@ -12316,22 +12692,22 @@ x_1 = lean_mk_string_from_bytes("commandParser", 13); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5305____closed__2() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5453____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5305____closed__1; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5453____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5305_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5453_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5305____closed__2; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5294____closed__4; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5453____closed__2; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5442____closed__4; x_4 = l_Lean_Parser_registerBuiltinDynamicParserAttribute(x_2, x_3, x_1); return x_4; } @@ -12340,7 +12716,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_commandParser(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5294____closed__4; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5442____closed__4; x_3 = l_Lean_Parser_categoryParser(x_2, x_1); return x_3; } @@ -13113,7 +13489,7 @@ x_6 = l_Lean_Parser_evalParserConstUnsafe(x_1, x_4, x_2); x_7 = lean_ctor_get(x_6, 4); lean_inc(x_7); x_8 = lean_box(0); -x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); +x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); lean_dec(x_7); if (x_9 == 0) { @@ -13613,7 +13989,7 @@ static lean_object* _init_l_Lean_Parser_parserOfStack___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__7; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__7; x_2 = lean_box(1); x_3 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_3, 0, x_1); @@ -13693,6 +14069,34 @@ lean_mark_persistent(l_Lean_Parser_builtinSyntaxNodeKindSetRef); lean_dec_ref(res); }l_Lean_Parser_registerBuiltinNodeKind___closed__1 = _init_l_Lean_Parser_registerBuiltinNodeKind___closed__1(); lean_mark_persistent(l_Lean_Parser_registerBuiltinNodeKind___closed__1); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__1(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__1); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__2(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__2); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__3(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__3); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__4 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__4(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__4); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__5 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__5(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__5); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__6 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__6(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__6); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__7 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__7(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__7); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__8 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__8(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__8); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__9 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__9(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__9); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__10 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__10(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__10); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__11 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__11(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__11); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__12 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__12(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__12); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__13 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__13(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__13); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__14 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__14(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74____closed__14); res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_74_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); @@ -13779,17 +14183,30 @@ l_Lean_Parser_getUnaryAlias___rarg___closed__1 = _init_l_Lean_Parser_getUnaryAli lean_mark_persistent(l_Lean_Parser_getUnaryAlias___rarg___closed__1); l_Lean_Parser_getBinaryAlias___rarg___closed__1 = _init_l_Lean_Parser_getBinaryAlias___rarg___closed__1(); lean_mark_persistent(l_Lean_Parser_getBinaryAlias___rarg___closed__1); -if (builtin) {res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_1928_(lean_io_mk_world()); +l_Lean_Parser_ParserAliasInfo_stackSz_x3f___default___closed__1 = _init_l_Lean_Parser_ParserAliasInfo_stackSz_x3f___default___closed__1(); +lean_mark_persistent(l_Lean_Parser_ParserAliasInfo_stackSz_x3f___default___closed__1); +l_Lean_Parser_ParserAliasInfo_stackSz_x3f___default = _init_l_Lean_Parser_ParserAliasInfo_stackSz_x3f___default(); +lean_mark_persistent(l_Lean_Parser_ParserAliasInfo_stackSz_x3f___default); +if (builtin) {res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_1951_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Parser_parserAliasesRef = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Parser_parserAliasesRef); lean_dec_ref(res); -}if (builtin) {res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_1963_(lean_io_mk_world()); +}if (builtin) {res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_1986_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Parser_parserAlias2kindRef = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Parser_parserAlias2kindRef); lean_dec_ref(res); -}l_Lean_Parser_registerAlias___closed__1 = _init_l_Lean_Parser_registerAlias___closed__1(); +}if (builtin) {res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_2021_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +l_Lean_Parser_parserAliases2infoRef = lean_io_result_get_value(res); +lean_mark_persistent(l_Lean_Parser_parserAliases2infoRef); +lean_dec_ref(res); +}l_Lean_Parser_getParserAliasInfo___closed__1 = _init_l_Lean_Parser_getParserAliasInfo___closed__1(); +lean_mark_persistent(l_Lean_Parser_getParserAliasInfo___closed__1); +l_Lean_Parser_getParserAliasInfo___closed__2 = _init_l_Lean_Parser_getParserAliasInfo___closed__2(); +lean_mark_persistent(l_Lean_Parser_getParserAliasInfo___closed__2); +l_Lean_Parser_registerAlias___closed__1 = _init_l_Lean_Parser_registerAlias___closed__1(); lean_mark_persistent(l_Lean_Parser_registerAlias___closed__1); l_Lean_Parser_registerAlias___closed__2 = _init_l_Lean_Parser_registerAlias___closed__2(); lean_mark_persistent(l_Lean_Parser_registerAlias___closed__2); @@ -13809,85 +14226,85 @@ l_Lean_Parser_mkParserOfConstantUnsafe___closed__7 = _init_l_Lean_Parser_mkParse lean_mark_persistent(l_Lean_Parser_mkParserOfConstantUnsafe___closed__7); l_Lean_Parser_mkParserOfConstantUnsafe___closed__8 = _init_l_Lean_Parser_mkParserOfConstantUnsafe___closed__8(); lean_mark_persistent(l_Lean_Parser_mkParserOfConstantUnsafe___closed__8); -if (builtin) {res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3165_(lean_io_mk_world()); +if (builtin) {res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3313_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Parser_parserAttributeHooks = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Parser_parserAttributeHooks); lean_dec_ref(res); }l_Lean_Parser_registerParserAttributeHook___closed__1 = _init_l_Lean_Parser_registerParserAttributeHook___closed__1(); lean_mark_persistent(l_Lean_Parser_registerParserAttributeHook___closed__1); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____lambda__2___closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____lambda__2___closed__1); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____lambda__2___closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____lambda__2___closed__2(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____lambda__2___closed__2); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__1(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__1); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__2(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__2); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__3(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__3); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__4 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__4(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__4); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__5 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__5(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__5); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__6 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__6(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__6); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__7 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__7(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246____closed__7); -res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3246_(lean_io_mk_world()); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____lambda__2___closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____lambda__2___closed__1); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____lambda__2___closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____lambda__2___closed__2); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__1(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__1); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__2(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__2); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__3(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__3); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__4 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__4(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__4); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__5 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__5(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__5); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__6 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__6(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__6); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__7 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__7(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394____closed__7); +res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3394_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__1(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__1); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__2(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__2); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__3(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__3); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__4 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__4(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__4); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__5 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__5(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__5); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__6 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__6(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290____closed__6); -res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3290_(lean_io_mk_world()); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__1(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__1); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__2(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__2); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__3(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__3); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__4 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__4(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__4); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__5 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__5(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__5); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__6 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__6(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438____closed__6); +res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3438_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__1(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__1); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__2(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__2); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__3(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__3); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__4 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__4(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__4); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__5 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__5(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__5); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__6 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__6(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__6); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__7 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__7(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__7); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__8 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__8(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446____closed__8); -if (builtin) {res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3446_(lean_io_mk_world()); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__1(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__1); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__2(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__2); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__3(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__3); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__4 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__4(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__4); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__5 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__5(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__5); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__6 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__6(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__6); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__7 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__7(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__7); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__8 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__8(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594____closed__8); +if (builtin) {res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3594_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Parser_parserExtension = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Parser_parserExtension); lean_dec_ref(res); }l_Lean_Parser_isParserCategory___closed__1 = _init_l_Lean_Parser_isParserCategory___closed__1(); lean_mark_persistent(l_Lean_Parser_isParserCategory___closed__1); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__1(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__1); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__2(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__2); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__3(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__3); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__4 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__4(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__4); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__5 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__5(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__5); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__6 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__6(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626____closed__6); -if (builtin) {res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3626_(lean_io_mk_world()); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__1(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__1); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__2(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__2); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__3(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__3); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__4 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__4(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__4); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__5 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__5(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__5); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__6 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__6(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774____closed__6); +if (builtin) {res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3774_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Parser_internal_parseQuotWithCurrentStage = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Parser_internal_parseQuotWithCurrentStage); @@ -13902,11 +14319,11 @@ l_Lean_Parser_categoryParserFnImpl___closed__3 = _init_l_Lean_Parser_categoryPar lean_mark_persistent(l_Lean_Parser_categoryParserFnImpl___closed__3); l_Lean_Parser_categoryParserFnImpl___closed__4 = _init_l_Lean_Parser_categoryParserFnImpl___closed__4(); lean_mark_persistent(l_Lean_Parser_categoryParserFnImpl___closed__4); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3895____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3895____closed__1(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3895____closed__1); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3895____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3895____closed__2(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3895____closed__2); -res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3895_(lean_io_mk_world()); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_4043____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_4043____closed__1(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_4043____closed__1); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_4043____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_4043____closed__2(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_4043____closed__2); +res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_4043_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); l_Lean_ScopedEnvExtension_add___at_Lean_Parser_addToken___spec__3___closed__1 = _init_l_Lean_ScopedEnvExtension_add___at_Lean_Parser_addToken___spec__3___closed__1(); @@ -14026,53 +14443,53 @@ l_Lean_Parser_mkParserAttributeImpl___closed__1 = _init_l_Lean_Parser_mkParserAt lean_mark_persistent(l_Lean_Parser_mkParserAttributeImpl___closed__1); l_Lean_Parser_mkParserAttributeImpl___closed__2 = _init_l_Lean_Parser_mkParserAttributeImpl___closed__2(); lean_mark_persistent(l_Lean_Parser_mkParserAttributeImpl___closed__2); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____lambda__1___closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____lambda__1___closed__1); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____lambda__1___closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____lambda__1___closed__2(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____lambda__1___closed__2); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____closed__1(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____closed__1); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____closed__2(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____closed__2); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____closed__3(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166____closed__3); -res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5166_(lean_io_mk_world()); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____lambda__1___closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____lambda__1___closed__1); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____lambda__1___closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____lambda__1___closed__2); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____closed__1(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____closed__1); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____closed__2(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____closed__2); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____closed__3(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314____closed__3); +res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5314_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5272____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5272____closed__1(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5272____closed__1); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5272____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5272____closed__2(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5272____closed__2); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5272____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5272____closed__3(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5272____closed__3); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5272____closed__4 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5272____closed__4(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5272____closed__4); -res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5272_(lean_io_mk_world()); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5420____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5420____closed__1(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5420____closed__1); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5420____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5420____closed__2(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5420____closed__2); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5420____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5420____closed__3(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5420____closed__3); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5420____closed__4 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5420____closed__4(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5420____closed__4); +res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5420_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5283____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5283____closed__1(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5283____closed__1); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5283____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5283____closed__2(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5283____closed__2); -res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5283_(lean_io_mk_world()); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5431____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5431____closed__1(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5431____closed__1); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5431____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5431____closed__2(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5431____closed__2); +res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5431_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5294____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5294____closed__1(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5294____closed__1); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5294____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5294____closed__2(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5294____closed__2); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5294____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5294____closed__3(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5294____closed__3); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5294____closed__4 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5294____closed__4(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5294____closed__4); -res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5294_(lean_io_mk_world()); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5442____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5442____closed__1(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5442____closed__1); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5442____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5442____closed__2(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5442____closed__2); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5442____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5442____closed__3(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5442____closed__3); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5442____closed__4 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5442____closed__4(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5442____closed__4); +res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5442_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5305____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5305____closed__1(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5305____closed__1); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5305____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5305____closed__2(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5305____closed__2); -res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5305_(lean_io_mk_world()); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5453____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5453____closed__1(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5453____closed__1); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5453____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5453____closed__2(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5453____closed__2); +res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5453_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); l_Lean_Parser_withOpenDeclFnCore___closed__1 = _init_l_Lean_Parser_withOpenDeclFnCore___closed__1(); diff --git a/stage0/stdlib/Lean/Parser/Extra.c b/stage0/stdlib/Lean/Parser/Extra.c index 80b28db3f99..9815deb1ffe 100644 --- a/stage0/stdlib/Lean/Parser/Extra.c +++ b/stage0/stdlib/Lean/Parser/Extra.c @@ -14,21 +14,17 @@ extern "C" { #endif lean_object* l_List_reverse___rarg(lean_object*); -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__5; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__97; lean_object* l_Lean_PrettyPrinter_Formatter_rawIdentNoAntiquot_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkAntiquot_formatter___closed__2; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__22; lean_object* l_Lean_PrettyPrinter_Parenthesizer_withoutInfo_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_many___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__102; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__50; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__18; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__57; LEAN_EXPORT lean_object* l_Lean_Parser_many1Indent_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_charLit___closed__1; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__29; LEAN_EXPORT lean_object* l_Lean_Parser_nonReservedSymbol_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__26; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__30; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__6; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__48; static lean_object* l_List_forIn_loop___at_Lean_Parser_sepByIndent_formatter___spec__2___closed__1; lean_object* l_Lean_PrettyPrinter_Formatter_visitArgs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); @@ -36,13 +32,19 @@ static lean_object* l_Lean_Parser_numLit___elambda__1___closed__2; static lean_object* l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_ppRealGroup_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_tokenAntiquotFn(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__28; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__38; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__21; LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Parser_sepByIndent_formatter___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__14; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__107; lean_object* l_Lean_PrettyPrinter_Parenthesizer_categoryParser_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__39; static lean_object* l_Lean_Parser_mkAntiquot_formatter___closed__14; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__8; lean_object* l_Lean_Parser_andthenInfo(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__89; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__28; lean_object* l_Lean_PrettyPrinter_Formatter_numLitNoAntiquot_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__51; LEAN_EXPORT lean_object* l_Lean_Parser_ppHardSpace_parenthesizer___rarg(lean_object*); static lean_object* l_Lean_Parser_optional___closed__1; static lean_object* l_Lean_Parser_commandParser_formatter___rarg___closed__2; @@ -52,21 +54,22 @@ static lean_object* l_Lean_Parser_mkAntiquot_formatter___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_mkAntiquot_parenthesizer(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ppSpace_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__19; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__30; lean_object* l_Lean_PrettyPrinter_Formatter_identNoAntiquot_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__15; static lean_object* l_Lean_Parser_antiquotNestedExpr_formatter___closed__10; LEAN_EXPORT lean_object* l_Lean_ppHardSpace_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ppGroup_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Parenthesizer_rawIdentNoAntiquot_parenthesizer___boxed(lean_object*); static lean_object* l_Lean_Parser_ppHardSpace___closed__1; lean_object* l_Lean_PrettyPrinter_Parenthesizer_withAntiquot_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__58; -extern lean_object* l_Lean_nullKind; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__75; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__94; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__17; static lean_object* l_Lean_Parser_mkAntiquotSplice_formatter___closed__6; static lean_object* l_Lean_Parser_ppHardSpace___closed__2; LEAN_EXPORT lean_object* l_Lean_ppHardLineUnlessUngrouped_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_notFollowedByFn___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__5; lean_object* l_Lean_PrettyPrinter_Formatter_sepByNoAntiquot_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_many(lean_object*); LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at_Lean_Parser_sepByIndent_formatter___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -75,137 +78,130 @@ lean_object* lean_name_mk_string(lean_object*, lean_object*); lean_object* l_Lean_Syntax_getOptional_x3f(lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); static lean_object* l_Lean_Parser_numLit___closed__4; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__98; lean_object* lean_array_uget(lean_object*, size_t); -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__57; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__4; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__16; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__50; static lean_object* l_Lean_Parser_antiquotExpr_formatter___closed__2; lean_object* l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ppSpace_parenthesizer___rarg(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__32; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__7; lean_object* l_Lean_PrettyPrinter_Parenthesizer_identNoAntiquot_parenthesizer___boxed(lean_object*); static lean_object* l_Lean_Parser_many1Indent_parenthesizer___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_sepBy1_parenthesizer(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkAntiquot_formatter___closed__3; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__11; LEAN_EXPORT lean_object* l_Lean_Parser_sepBy_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__10; LEAN_EXPORT lean_object* l_Lean_Parser_ppGroup___boxed(lean_object*); -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__3; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__8; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__44; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_symbol_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_charLit; static lean_object* l_Lean_Parser_mkAntiquot_formatter___closed__15; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__73; static lean_object* l_Lean_Parser_numLit___elambda__1___closed__1; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__67; lean_object* l_Lean_PrettyPrinter_Formatter_indent(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Parenthesizer_orelse_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__12; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__44; LEAN_EXPORT lean_object* l_Lean_Parser_ppAllowUngrouped_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_ident_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_sepBy_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__31; lean_object* l_Lean_PrettyPrinter_Formatter_atomic_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__90; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__4; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__18; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__80; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__6; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__15; LEAN_EXPORT lean_object* l_Lean_Parser_group_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Formatter_manyNoAntiquot_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_sepByIndent_parenthesizer(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_nameLit_formatter___closed__3; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__27; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__99; LEAN_EXPORT lean_object* l_Lean_Parser_ppDedentIfGrouped___boxed(lean_object*); extern lean_object* l_Lean_Parser_pushNone; static lean_object* l_Lean_Parser_ident_formatter___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__20; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__101; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__17; static lean_object* l_Lean_Parser_many___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_scientificLit; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__9; LEAN_EXPORT lean_object* l_Lean_ppSpace_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_sepBy1Indent_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__59; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__31; static lean_object* l_Lean_Parser_many_parenthesizer___closed__1; lean_object* l_Lean_PrettyPrinter_Formatter_orelse_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_group(lean_object*); -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__27; LEAN_EXPORT lean_object* l_Lean_Parser_ident; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__97; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__23; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__7; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__15; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__3; lean_object* l_id___rarg___boxed(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__24; lean_object* l_Lean_Parser_checkLinebreakBefore___elambda__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__72; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__96; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__61; LEAN_EXPORT lean_object* l_Lean_Parser_ppDedentIfGrouped(lean_object*); lean_object* l_Lean_Parser_ParserState_mkNode(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_sepByElemParser_formatter___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_withAntiquotSpliceAndSuffix_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_many1Indent___lambda__1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__11; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__9; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__38; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__49; static lean_object* l_Lean_Parser_antiquotNestedExpr_parenthesizer___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__21; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__107; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__40; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__108; LEAN_EXPORT lean_object* l_Lean_Parser_sepByIndent___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__17; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__24; lean_object* l_Lean_Parser_pushNone___elambda__1___boxed(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__34; +lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Level_PP_Result_quote___spec__1(lean_object*, lean_object*); lean_object* l_Lean_mkIdentFrom(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__43; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__8; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__7; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__71; LEAN_EXPORT lean_object* l_Lean_Parser_scientificLit___elambda__1(lean_object*, lean_object*); lean_object* l_Lean_Parser_orelseFn(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__100; lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__33; static lean_object* l_Lean_Parser_strLit_formatter___closed__1; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__4; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__42; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__10; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__32; static lean_object* l_Lean_Parser_mkAntiquot_formatter___closed__8; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__13; lean_object* l_Lean_PrettyPrinter_Parenthesizer_nameLitNoAntiquot_parenthesizer___boxed(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__41; lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_optional_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__88; LEAN_EXPORT lean_object* l_Lean_Parser_ppIndent_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_optional_formatter___closed__4; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__20; lean_object* lean_string_append(lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkAntiquot_formatter___closed__9; lean_object* l_List_range(lean_object*); static lean_object* l_Lean_Parser_strLit_formatter___closed__2; static lean_object* l_Lean_Parser_mkAntiquot_formatter___closed__17; lean_object* l_Lean_Parser_mkAtomicInfo(lean_object*); +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__35; static lean_object* l_Lean_Parser_charLit___elambda__1___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__104; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__72; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__73; LEAN_EXPORT lean_object* l_Lean_Parser_unicodeSymbol_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__41; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__10; static lean_object* l_Lean_Parser_numLit___closed__3; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__32; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__36; static lean_object* l_Lean_Parser_optional_formatter___closed__2; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__1; static lean_object* l_Lean_Parser_charLit___elambda__1___closed__2; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__13; static lean_object* l_Lean_Parser_mkAntiquot_formatter___closed__1; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__13; lean_object* l_Lean_PrettyPrinter_Formatter_nonReservedSymbolNoAntiquot_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__83; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__45; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__21; LEAN_EXPORT lean_object* l_Lean_Parser_sepBy1_formatter(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__37; lean_object* l_Lean_PrettyPrinter_Parenthesizer_numLitNoAntiquot_parenthesizer___boxed(lean_object*); static lean_object* l_Lean_Parser_antiquotNestedExpr_formatter___closed__5; lean_object* l_Lean_PrettyPrinter_Formatter_unicodeSymbolNoAntiquot_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__55; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__23; static lean_object* l_Lean_Parser_antiquotNestedExpr_parenthesizer___closed__5; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__34; static lean_object* l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__8; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__24; LEAN_EXPORT lean_object* l_Lean_Parser_antiquotNestedExpr_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_group_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Parenthesizer_scientificLitNoAntiquot_parenthesizer___boxed(lean_object*); static lean_object* l_Lean_Parser_antiquotNestedExpr_parenthesizer___closed__3; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__2; lean_object* lean_string_utf8_byte_size(lean_object*); -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__38; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__53; LEAN_EXPORT lean_object* l_Lean_Parser_mkAntiquot_formatter___boxed(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_leadingNode_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_antiquotNestedExpr_parenthesizer___closed__1; @@ -216,238 +212,243 @@ lean_object* l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(lean_object*, lean_obj static lean_object* l_Lean_Parser_mkAntiquot_formatter___closed__4; static lean_object* l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_many_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__20; lean_object* l_Lean_Parser_nameLitFn(lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__45; LEAN_EXPORT lean_object* l_Lean_Parser_ppLine; -extern lean_object* l_Lean_nameLitKind; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__16; static lean_object* l___regBuiltin_Lean_Parser_antiquotNestedExpr_parenthesizer___closed__2; static lean_object* l_Lean_Parser_strLit___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_rawIdent_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__21; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__49; static lean_object* l_Lean_Parser_strLit_formatter___closed__3; static lean_object* l_Lean_Parser_leadingNode_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_optional_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_antiquotNestedExpr_formatter___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_sepByIndent_formatter___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Parenthesizer_sepByNoAntiquot_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__17; LEAN_EXPORT lean_object* l_Lean_Parser_sepBy1Indent_parenthesizer(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29________; static lean_object* l_Lean_Parser_ident___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__94; lean_object* l_Lean_PrettyPrinter_Parenthesizer_atomic_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__36; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__32; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__10; LEAN_EXPORT lean_object* l_Lean_Parser_ppAllowUngrouped_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_rawIdent___closed__3; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__26; LEAN_EXPORT lean_object* l_Lean_ppAllowUngrouped_formatter___boxed(lean_object*); lean_object* l_Lean_PrettyPrinter_Parenthesizer_unicodeSymbolNoAntiquot_parenthesizer___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__114; LEAN_EXPORT lean_object* l_Lean_Parser_optional(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__4; lean_object* l_Lean_PrettyPrinter_Formatter_fill(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__57; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__70; static lean_object* l_Lean_Parser_antiquotExpr_formatter___closed__3; static lean_object* l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__7; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__9; static lean_object* l_Lean_Parser_sepByIndent___closed__5; +uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_numLit___closed__1; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__45; lean_object* l_Lean_Parser_rawIdentFn(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__40; lean_object* l_Lean_Parser_registerAliasCore___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__67; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__69; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__20; LEAN_EXPORT lean_object* l_Lean_Parser_strLit___elambda__1(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__35; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__44; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__21; lean_object* l_Lean_PrettyPrinter_Formatter_checkColGe_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ppIndent_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__70; lean_object* l_Lean_PrettyPrinter_Formatter_symbolNoAntiquot_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_PrettyPrinter_parenthesizerAttribute; LEAN_EXPORT lean_object* l_Lean_Parser_ppRealFill_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__36; lean_object* l_Lean_PrettyPrinter_Parenthesizer_setExpected_parenthesizer___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_termParser_formatter___boxed(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__75; +static lean_object* l_Lean_Parser_ident_formatter___closed__3; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Parser_sepByIndent_formatter___spec__3(lean_object*, size_t, size_t); -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__23; LEAN_EXPORT lean_object* l_Lean_ppHardLineUnlessUngrouped_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_numLit___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__58; static lean_object* l_Lean_Parser_antiquotNestedExpr_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_mkAntiquotSplice_formatter___closed__7; lean_object* l_Lean_Syntax_MonadTraverser_getCur___at_Lean_PrettyPrinter_Formatter_visitArgs___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__13; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__11; lean_object* l_Lean_PrettyPrinter_Formatter_tokenWithAntiquot_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_ident___closed__4; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__75; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__27; lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ppSpace_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__61; static lean_object* l_Lean_Parser_mkAntiquotSplice_parenthesizer___closed__1; static lean_object* l_Lean_Parser_ident___closed__2; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_sepByElemParser_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_group_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_ppHardSpace_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_antiquotNestedExpr_formatter___closed__7; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__23; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__12; LEAN_EXPORT lean_object* l_Lean_ppRealFill_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_charLit_formatter___closed__2; lean_object* l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_mkAntiquotSplice_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__26; lean_object* l_Lean_PrettyPrinter_Formatter_setExpected_formatter___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Formatter_strLitNoAntiquot_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__112; static lean_object* l_Lean_Parser_strLit_parenthesizer___closed__2; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__35; lean_object* lean_st_ref_take(lean_object*, lean_object*); lean_object* l_Lean_Parser_nodeInfo(lean_object*, lean_object*); static lean_object* l_Lean_Parser_rawIdent___closed__4; -extern lean_object* l_Lean_numLitKind; static lean_object* l_Lean_Parser_strLit_parenthesizer___closed__1; +lean_object* l_ReaderT_bind___at_Lean_Unhygienic_instMonadQuotationUnhygienic___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__43; LEAN_EXPORT lean_object* l_Lean_Parser_charLit_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__8; +static lean_object* l_Lean_Parser_charLit_formatter___closed__4; LEAN_EXPORT lean_object* l_Lean_ppAllowUngrouped_formatter___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__82; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__45; lean_object* l_Lean_PrettyPrinter_Parenthesizer_nonReservedSymbolNoAntiquot_parenthesizer___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_scientificLit_formatter___closed__3; lean_object* l_Lean_Parser_noFirstTokenInfo(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_scientificLit_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__12; -extern lean_object* l_Lean_strLitKind; static lean_object* l_Lean_Parser_mkAntiquotSplice_parenthesizer___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__59; LEAN_EXPORT lean_object* l_Lean_Parser_strLit_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__31; lean_object* l_Lean_PrettyPrinter_Formatter_checkPrec_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__11; +LEAN_EXPORT lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Formatter_decQuotDepth_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Parenthesizer_visitArgs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__58; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__25; LEAN_EXPORT lean_object* l_Lean_Parser_nodeWithAntiquot_formatter(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__20; lean_object* l_Lean_Parser_sepBy(lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_antiquotNestedExpr_parenthesizer(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_manyIndent_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_numLit_formatter___closed__1; static lean_object* l_Lean_Parser_sepByElemParser_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_notSymbol_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__32; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__10; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__12; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__3; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__43; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__15; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__6; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__72; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__23; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__96; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__46; lean_object* l_Lean_Name_toString(lean_object*, uint8_t); +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__46; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__8; LEAN_EXPORT lean_object* l_Lean_Parser_ppHardSpace_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_nameLit___closed__3; static lean_object* l_Lean_Parser_many___closed__1; lean_object* l_Lean_Parser_numLitFn(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_many1_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Unhygienic_run___rarg(lean_object*); lean_object* l_Lean_Parser_optionaInfo(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__18; +LEAN_EXPORT lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__86; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__52; static lean_object* l_Lean_Parser_strLit___elambda__1___closed__2; static lean_object* l_Lean_Parser_mkAntiquotSplice_formatter___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_sepByIndent_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__53; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__110; lean_object* l_Lean_PrettyPrinter_Parenthesizer_manyNoAntiquot_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_nameLit; lean_object* l_Lean_PrettyPrinter_Formatter_node_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_sepByIndent_formatter___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__48; static lean_object* l_Lean_Parser_many1Indent_formatter___closed__1; LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Parser_sepByIndent_formatter___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__61; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Parser_sepByIndent_formatter___spec__3___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__26; lean_object* l_Lean_Parser_charLitFn(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__100; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__1; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__29; +static lean_object* l_Lean_Parser_strLit_formatter___closed__4; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__88; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__13; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__33; static lean_object* l_Lean_Parser_mkAntiquot_formatter___closed__6; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__73; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__41; +LEAN_EXPORT lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_strLitFn(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_antiquotNestedExpr_parenthesizer___closed__3; static lean_object* l_Lean_Parser_charLit___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__42; lean_object* l_Lean_PrettyPrinter_Parenthesizer_charLitNoAntiquot_parenthesizer___boxed(lean_object*); +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__14; LEAN_EXPORT lean_object* l_Lean_Parser_many1___elambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_manyIndent(lean_object*); static lean_object* l_Lean_Parser_nameLit___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_antiquotNestedExpr_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__76; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__3; +static lean_object* l_Lean_Parser_numLit_formatter___closed__4; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__33; LEAN_EXPORT lean_object* l_Lean_Parser_ppSpace; LEAN_EXPORT lean_object* l_Lean_Parser_ppHardSpace; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__99; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__39; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__2; lean_object* l_Lean_Parser_orelseInfo(lean_object*, lean_object*); static lean_object* l_Lean_Parser_nameLit___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_termParser_formatter___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__26; static lean_object* l_Lean_Parser_sepByElemParser_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_ppIndent(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_scientificLit_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__50; LEAN_EXPORT lean_object* l_Lean_Parser_ppRealFill(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_commandParser_formatter___boxed(lean_object*); static lean_object* l_Lean_Parser_mkAntiquotSplice_formatter___closed__3; lean_object* l_Lean_Syntax_getId(lean_object*); -lean_object* l___private_Init_Meta_0__Lean_quoteNameMk(lean_object*); -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__65; -extern lean_object* l_Lean_charLitKind; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__111; static lean_object* l_Lean_Parser_nameLit___closed__4; extern lean_object* l_Lean_Parser_maxPrec; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__34; LEAN_EXPORT lean_object* l_Lean_Parser_strLit_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__110; LEAN_EXPORT lean_object* l_Lean_Parser_rawIdent; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__69; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__82; static lean_object* l_Lean_Parser_antiquotExpr_parenthesizer___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_ppDedent(lean_object*); static lean_object* l_Lean_Parser_sepByIndent_parenthesizer___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__30; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__11; static lean_object* l_Lean_Parser_strLit___elambda__1___closed__1; static lean_object* l_Lean_Parser_rawIdent___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_many1Indent_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_scientificLitFn(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__80; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__6; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__93; LEAN_EXPORT lean_object* l_Lean_Parser_nodeWithAntiquot_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__22; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__91; LEAN_EXPORT lean_object* l_Lean_Parser_sepByIndent(lean_object*, lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_Parser_charLit___closed__2; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__12; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__78; LEAN_EXPORT lean_object* l_Lean_Parser_sepByIndent_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Formatter_scientificLitNoAntiquot_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_PrettyPrinter_formatterAttribute; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__19; uint32_t lean_string_utf8_get(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__81; LEAN_EXPORT lean_object* l_Lean_Parser_ppLine_parenthesizer___rarg(lean_object*); lean_object* l_Lean_Syntax_mkStrLit(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__31; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__30; static lean_object* l_Lean_Parser_mkAntiquot_formatter___closed__10; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__25; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__68; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__98; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__16; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__28; static lean_object* l_Lean_Parser_sepByIndent___closed__1; static lean_object* l_Lean_Parser_antiquotExpr_parenthesizer___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__95; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__18; static lean_object* l_Lean_Parser_optional_formatter___closed__3; -LEAN_EXPORT lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29______; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__31; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__46; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__79; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__5; static lean_object* l_Lean_Parser_sepByIndent___closed__4; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__11; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__42; LEAN_EXPORT lean_object* l_Lean_Parser_ppDedentIfGrouped_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Parenthesizer_pushNone_parenthesizer___boxed(lean_object*); static lean_object* l_Lean_Parser_mkAntiquot_parenthesizer___closed__8; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__86; static lean_object* l_Lean_Parser_rawIdent___elambda__1___closed__1; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__19; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__43; static lean_object* l_Lean_Parser_numLit_formatter___closed__3; static lean_object* l___regBuiltin_Lean_Parser_antiquotNestedExpr_parenthesizer___closed__4; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__101; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__85; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__66; static lean_object* l_Lean_Parser_scientificLit___closed__3; static lean_object* l_Lean_Parser_mkAntiquot_parenthesizer___closed__9; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__111; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__48; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__77; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__3; static lean_object* l_Lean_Parser_numLit_parenthesizer___closed__2; static lean_object* l_Lean_ppHardSpace_formatter___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__52; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__105; LEAN_EXPORT lean_object* l_Lean_Parser_ppIndent___boxed(lean_object*); -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__51; static lean_object* l_Lean_Parser_optional_parenthesizer___closed__1; static lean_object* l_Lean_Parser_antiquotNestedExpr_formatter___closed__1; static lean_object* l_Lean_Parser_mkAntiquot_formatter___closed__11; @@ -457,16 +458,12 @@ static lean_object* l_Lean_Parser_mkAntiquot_parenthesizer___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_sepBy1_parenthesizer___boxed(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_commandParser_formatter(lean_object*); static lean_object* l_Lean_Parser_rawIdent_parenthesizer___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__29; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__2; -lean_object* l_Lean_Parser_registerAlias(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Parser_registerAlias(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_strLit___closed__3; static lean_object* l_Lean_Parser_optional___lambda__1___closed__1; lean_object* l_Lean_Parser_manyAux(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_charLit___closed__4; size_t lean_usize_of_nat(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__64; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__16; static lean_object* l_Lean_Parser_antiquotNestedExpr_parenthesizer___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_ppRealGroup(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_optional___lambda__1___boxed(lean_object*, lean_object*, lean_object*); @@ -476,156 +473,171 @@ lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkAntiquot_parenthesizer___closed__2; lean_object* l_Lean_PrettyPrinter_Formatter_nameLitNoAntiquot_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_rawIdent___closed__2; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__27; static lean_object* l_Lean_Parser_antiquotNestedExpr_formatter___closed__6; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__89; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__14; LEAN_EXPORT lean_object* l_Lean_Parser_mkAntiquot_formatter(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ppHardSpace_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_many1_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__17; static lean_object* l_Lean_Parser_nameLit___elambda__1___closed__1; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__66; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__56; LEAN_EXPORT lean_object* l_Lean_Parser_ppAllowUngrouped; LEAN_EXPORT lean_object* l_Lean_Parser_notSymbol(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_sepByElemParser_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_antiquotNestedExpr_formatter___closed__2; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__24; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__56; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__22; LEAN_EXPORT lean_object* l_Lean_Parser_termParser_formatter(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__87; LEAN_EXPORT lean_object* l_Lean_Parser_mkAntiquot_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_sepBy1(lean_object*, lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_Parser_sepByIndent___closed__2; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__49; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__15; LEAN_EXPORT lean_object* l_Lean_Parser_ppLine_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_groupKind; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__68; LEAN_EXPORT lean_object* l_Lean_Parser_sepBy1Indent_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_ident___elambda__1___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__60; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__47; static lean_object* l_Lean_Parser_mkAntiquot_parenthesizer___closed__1; -extern lean_object* l_Lean_identKind; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__36; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__4; static lean_object* l_Lean_Parser_sepByIndent_parenthesizer___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_commandParser_formatter___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_many_formatter___closed__3; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__22; LEAN_EXPORT lean_object* l_Lean_Parser_ident___elambda__1(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__79; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__104; lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkNoWsBefore_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_charLit___elambda__1(lean_object*, lean_object*); uint8_t lean_uint32_dec_eq(uint32_t, uint32_t); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__81; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__3; static lean_object* l_Lean_Parser_nameLit_parenthesizer___closed__1; lean_object* l_String_intercalate(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_numLit; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__51; static lean_object* l_Lean_Parser_many___closed__4; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__27; static lean_object* l_Lean_Parser_charLit_formatter___closed__1; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__27; static lean_object* l_Lean_Parser_antiquotExpr_parenthesizer___closed__3; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__9; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__40; lean_object* l_Lean_Macro_throwError___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__67; LEAN_EXPORT lean_object* l_Lean_Parser_ppHardSpace___lambda__1___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__40; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__9; LEAN_EXPORT lean_object* l_Lean_Parser_ppHardLineUnlessUngrouped; LEAN_EXPORT lean_object* l_Lean_Parser_numLit___elambda__1(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__24; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__14; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__95; lean_object* l_Lean_PrettyPrinter_Parenthesizer_decQuotDepth_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ppDedent___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304_(lean_object*); +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__44; +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610_(lean_object*); extern lean_object* l_Lean_PrettyPrinter_Formatter_formatterAliasesRef; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__1; -extern lean_object* l_Lean_scientificLitKind; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__83; static lean_object* l_Lean_Parser_scientificLit___elambda__1___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__74; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__33; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__92; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__109; +LEAN_EXPORT lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__70; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__8; lean_object* l_Lean_PrettyPrinter_Formatter_checkNoWsBefore_formatter___boxed(lean_object*); +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__17; static lean_object* l_Lean_Parser_sepByIndent___closed__3; static lean_object* l_List_forIn_loop___at_Lean_Parser_sepByIndent_formatter___spec__2___closed__2; LEAN_EXPORT lean_object* l_Lean_ppDedentIfGrouped_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__16; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__19; static lean_object* l_Lean_Parser_termParser_formatter___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_ppGroup_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__9; LEAN_EXPORT lean_object* l_Lean_Parser_commandParser_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_scientificLit_parenthesizer___closed__2; static lean_object* l_Lean_Parser_mkAntiquot_formatter___closed__16; +static lean_object* l_Lean_Parser_group_formatter___closed__1; lean_object* l_Lean_Syntax_getArgs(lean_object*); lean_object* l_Lean_Name_append(lean_object*, lean_object*); lean_object* l_Lean_Macro_resolveGlobalName(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__37; static lean_object* l_Lean_Parser_scientificLit___elambda__1___closed__2; static lean_object* l_Lean_Parser_termParser_formatter___rarg___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__64; static lean_object* l_Lean_Parser_mkAntiquot_formatter___closed__7; +lean_object* l_Lean_quoteNameMk(lean_object*); static lean_object* l_Lean_Parser_nameLit_parenthesizer___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_many___elambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ppAllowUngrouped_formatter(lean_object*); -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__54; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__35; static lean_object* l_Lean_Parser_charLit_parenthesizer___closed__1; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__62; lean_object* l_Lean_Parser_symbolInfo(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_charLit_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_optional___lambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__7; static lean_object* l_Lean_Parser_mkAntiquot_parenthesizer___closed__6; static lean_object* l_Lean_Parser_charLit_parenthesizer___closed__2; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__20; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__74; static lean_object* l_Lean_Parser_charLit_formatter___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__71; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__65; lean_object* l_Lean_Parser_orelseFnCore(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__39; extern lean_object* l_Lean_Parser_epsilonInfo; LEAN_EXPORT lean_object* l_Lean_Parser_notSymbol_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ppLine_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__9; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__4; static lean_object* l_Lean_Parser_optional___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_many1Indent(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__7; lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkNoImmediateColon_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_ppDedent_formatter___closed__1; lean_object* l_Lean_PrettyPrinter_Formatter_pushNone_formatter___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_sepByIndent_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__10; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__4; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__2; lean_object* l_Std_Format_getIndent(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__47; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__35; static lean_object* l_Lean_Parser_ident___elambda__1___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__63; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__60; LEAN_EXPORT lean_object* l_Lean_Parser_notSymbol_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_sepBy_formatter(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_antiquotNestedExpr_formatter___closed__9; +lean_object* l_Lean_Syntax_mkNameLit(lean_object*, lean_object*); lean_object* l_Lean_Parser_withAntiquotSpliceAndSuffix(lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_leadingNode_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkAntiquotSplice_formatter___closed__2; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__7; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__71; LEAN_EXPORT lean_object* l_Lean_Parser_sepByElemParser_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_nameLit_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_antiquotExpr_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ident_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_optional_formatter___closed__1; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__16; lean_object* l_Lean_PrettyPrinter_Formatter_checkNoImmediateColon_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_sepBy1Indent(lean_object*, lean_object*, lean_object*, uint8_t); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_mkAntiquotSplice_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkAntiquotSplice_formatter___closed__1; lean_object* l_Lean_PrettyPrinter_Parenthesizer_symbolNoAntiquot_parenthesizer___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_numLit_formatter___closed__2; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__63; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__24; LEAN_EXPORT lean_object* l_Lean_Parser_sepBy1_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__18; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__76; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__2; +uint8_t l_Lean_Syntax_isNone(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__30; LEAN_EXPORT lean_object* l_Lean_Parser_numLit_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__37; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__103; static lean_object* l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__66; LEAN_EXPORT lean_object* l_Lean_ppSpace_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__55; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__108; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__55; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__34; +static lean_object* l_Lean_Parser_scientificLit_formatter___closed__4; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__41; lean_object* l_String_trim(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__85; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__23; lean_object* l_Lean_PrettyPrinter_Formatter_fold(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__77; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__50; lean_object* l_Lean_Parser_nodeFn(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__13; static lean_object* l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__4; lean_object* l_Lean_Parser_optionalFn(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__105; lean_object* l_Lean_PrettyPrinter_Formatter_withoutInfo_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__56; static lean_object* l_Lean_Parser_many_formatter___closed__2; lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkLinebreakBefore_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_nameLit_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -633,74 +645,74 @@ lean_object* l_Lean_PrettyPrinter_Formatter_concat___lambda__1___boxed(lean_obje LEAN_EXPORT lean_object* l_Lean_Parser_unicodeSymbol_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_rawIdent_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__29; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__1; static lean_object* l_Lean_Parser_scientificLit_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__9; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__25; static lean_object* l_List_forIn_loop___at_Lean_Parser_sepByIndent_formatter___spec__2___closed__3; extern lean_object* l_Lean_PrettyPrinter_Parenthesizer_parenthesizerAliasesRef; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__49; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__52; lean_object* l_Lean_PrettyPrinter_Parenthesizer_withAntiquotSuffixSplice_parenthesizer___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__33; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__6; LEAN_EXPORT lean_object* l_Lean_ppLine_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_int_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__62; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__54; LEAN_EXPORT lean_object* l_Lean_Parser_nameLit___elambda__1(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__84; LEAN_EXPORT lean_object* l_Lean_Parser_ppHardLineUnlessUngrouped_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_rawIdent___closed__1; static lean_object* l_Lean_Parser_many_formatter___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__84; static lean_object* l_Lean_Parser_scientificLit_formatter___closed__2; static lean_object* l_Lean_Parser_strLit___closed__4; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__62; static lean_object* l_Lean_Parser_mkAntiquotSplice_formatter___closed__4; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__37; lean_object* l_Lean_PrettyPrinter_Formatter_charLitNoAntiquot_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkAntiquot_formatter___closed__13; LEAN_EXPORT lean_object* l_Lean_Parser_withAntiquotSpliceAndSuffix_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_mkAntiquot(lean_object*, lean_object*, uint8_t, uint8_t); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__29; lean_object* l_Lean_PrettyPrinter_Parenthesizer_withPosition_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__52; static lean_object* l_Lean_Parser_nameLit_formatter___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__93; LEAN_EXPORT lean_object* l_Lean_Parser_ppGroup(lean_object*); static lean_object* l_Lean_Parser_antiquotNestedExpr_parenthesizer___closed__4; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__106; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__5; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__63; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__106; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Formatter_pushWhitespace(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__1; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_many_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__28; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__47; static lean_object* l_Lean_Parser_antiquotNestedExpr_formatter___closed__4; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__91; static lean_object* l_Lean_Parser_ident_parenthesizer___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__78; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__22; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__8; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__38; LEAN_EXPORT lean_object* l_Lean_Parser_ppRealFill___boxed(lean_object*); lean_object* lean_nat_mod(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter(lean_object*); -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__25; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__14; LEAN_EXPORT lean_object* l_Lean_Parser_sepBy1Indent___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ppHardLineUnlessUngrouped_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__21; LEAN_EXPORT lean_object* l_Lean_Parser_ppAllowUngrouped_parenthesizer___rarg(lean_object*); static lean_object* l_Lean_Parser_mkAntiquot_parenthesizer___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_ppHardLineUnlessUngrouped_parenthesizer___rarg(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__19; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__51; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__47; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__26; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__34; lean_object* l_Lean_PrettyPrinter_Parenthesizer_strLitNoAntiquot_parenthesizer___boxed(lean_object*); -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__60; static lean_object* l_Lean_Parser_rawIdent_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_ppHardSpace___lambda__1(lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkColGe_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__25; static lean_object* l_Lean_Parser_mkAntiquot_formatter___closed__12; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__74; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__90; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__102; static lean_object* l_Lean_Parser_antiquotExpr_formatter___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_mkAntiquotSplice_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_notSymbol_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__12; LEAN_EXPORT lean_object* l_Lean_Parser_nonReservedSymbol_formatter(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__46; LEAN_EXPORT lean_object* l_Lean_Parser_antiquotExpr_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__54; LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at_Lean_Parser_sepByIndent_formatter___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ppRealGroup___boxed(lean_object*); static lean_object* l_Lean_Parser_antiquotNestedExpr_formatter___closed__8; @@ -708,30 +720,26 @@ extern lean_object* l_Lean_PrettyPrinter_backtrackExceptionId; LEAN_EXPORT lean_object* l_Lean_ppRealGroup_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_numLit_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_optional___elambda__1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__109; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__39; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__65; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__92; LEAN_EXPORT lean_object* l_Lean_Parser_sepBy_parenthesizer(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_symbolFnAux(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__103; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__1; static lean_object* l_Lean_Parser_ident___closed__3; static lean_object* l_Lean_Parser_mkAntiquot_parenthesizer___closed__5; -uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_nameLit_formatter___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_sepByIndent_formatter___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__64; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__2; lean_object* l_Lean_PrettyPrinter_Formatter_withAntiquotSuffixSplice_formatter___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Formatter_setLhsPrec_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ident_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__11; static lean_object* l_List_forIn_loop___at_Lean_Parser_sepByIndent_formatter___spec__2___closed__4; static lean_object* l_Lean_Parser_scientificLit___closed__2; static lean_object* l_Lean_Parser_many1Indent___closed__1; static lean_object* l_Lean_ppDedentIfGrouped_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__5; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__29; LEAN_EXPORT lean_object* l_Lean_Parser_symbol_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__3; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__10; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__48; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__31; lean_object* l_Lean_PrettyPrinter_Parenthesizer_tokenWithAntiquot_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Formatter_group(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_identFn(lean_object*, lean_object*); @@ -739,47 +747,45 @@ LEAN_EXPORT lean_object* l_Lean_Parser_nameLit_parenthesizer(lean_object*, lean_ lean_object* lean_nat_to_int(lean_object*); static lean_object* l_Lean_Parser_scientificLit_parenthesizer___closed__1; LEAN_EXPORT lean_object* l_Lean_ppAllowUngrouped_formatter___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__42; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__53; static lean_object* l_Lean_Parser_mkAntiquot_parenthesizer___closed__7; static lean_object* l_Lean_Parser_nameLit___elambda__1___closed__2; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__19; LEAN_EXPORT lean_object* l_Lean_Parser_nonReservedSymbol_parenthesizer(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_notSymbol_parenthesizer___rarg(lean_object*); +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_rawIdent___elambda__1(lean_object*, lean_object*); -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__5; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_strLit; LEAN_EXPORT lean_object* l_Lean_Parser_many1(lean_object*); -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__18; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__25; lean_object* l_Lean_PrettyPrinter_Formatter_andthen_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_scientificLit___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_mkAntiquotSplice_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__28; LEAN_EXPORT lean_object* l_Lean_Parser_sepByElemParser_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_nodeWithAntiquot_parenthesizer(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__28; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__30; static lean_object* l_Lean_Parser_strLit___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_ppDedent_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__69; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__14; LEAN_EXPORT lean_object* l_Lean_Parser_sepBy1Indent_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__15; LEAN_EXPORT lean_object* l_Lean_Parser_notSymbol_formatter___rarg(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__3; static lean_object* l_Lean_Parser_antiquotExpr_formatter___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__59; static lean_object* l_Lean_Parser_leadingNode_formatter___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__113; LEAN_EXPORT lean_object* l_Lean_Parser_termParser_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__68; LEAN_EXPORT lean_object* l_Lean_ppDedent_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_manyIndent_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Syntax_mkLit(lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); lean_object* l_ReaderT_bind___at_Lean_PrettyPrinter_Formatter_categoryFormatterCore___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Parenthesizer_node_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__87; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__22; LEAN_EXPORT lean_object* l_Lean_Parser_ppLine_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_mkAntiquotSplice_formatter___closed__8; lean_object* l_Lean_PrettyPrinter_Formatter_categoryParser_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_andthenFn(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_many1Indent___closed__2; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__12; static lean_object* _init_l_Lean_Parser_leadingNode_formatter___closed__1() { _start: { @@ -2011,15 +2017,33 @@ static lean_object* _init_l_Lean_Parser_mkAntiquotSplice_formatter___closed__5() _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("]", 1); +x_1 = lean_mk_string_from_bytes("null", 4); return x_1; } } static lean_object* _init_l_Lean_Parser_mkAntiquotSplice_formatter___closed__6() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_mkAntiquotSplice_formatter___closed__5; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_mkAntiquotSplice_formatter___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("]", 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_mkAntiquotSplice_formatter___closed__8() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_mkAntiquotSplice_formatter___closed__5; +x_1 = l_Lean_Parser_mkAntiquotSplice_formatter___closed__7; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -2031,11 +2055,11 @@ LEAN_EXPORT lean_object* l_Lean_Parser_mkAntiquotSplice_formatter(lean_object* x 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; 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; lean_object* x_25; lean_object* x_26; x_9 = l_Lean_Parser_mkAntiquotSplice_formatter___closed__2; x_10 = l_Lean_Name_append(x_1, x_9); -x_11 = l_Lean_nullKind; +x_11 = l_Lean_Parser_mkAntiquotSplice_formatter___closed__6; x_12 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_node_formatter), 7, 2); lean_closure_set(x_12, 0, x_11); lean_closure_set(x_12, 1, x_2); -x_13 = l_Lean_Parser_mkAntiquotSplice_formatter___closed__6; +x_13 = l_Lean_Parser_mkAntiquotSplice_formatter___closed__8; x_14 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_14, 0, x_13); lean_closure_set(x_14, 1, x_3); @@ -2177,7 +2201,7 @@ static lean_object* _init_l_Lean_Parser_mkAntiquotSplice_parenthesizer___closed_ _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_mkAntiquotSplice_formatter___closed__5; +x_1 = l_Lean_Parser_mkAntiquotSplice_formatter___closed__7; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -2189,7 +2213,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_mkAntiquotSplice_parenthesizer(lean_objec 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; 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; lean_object* x_25; lean_object* x_26; x_9 = l_Lean_Parser_mkAntiquotSplice_formatter___closed__2; x_10 = l_Lean_Name_append(x_1, x_9); -x_11 = l_Lean_nullKind; +x_11 = l_Lean_Parser_mkAntiquotSplice_formatter___closed__6; x_12 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_node_parenthesizer), 7, 2); lean_closure_set(x_12, 0, x_11); lean_closure_set(x_12, 1, x_2); @@ -2612,7 +2636,7 @@ lean_inc(x_4); x_5 = lean_array_get_size(x_4); lean_dec(x_4); x_6 = l_Lean_Parser_manyAux(x_1, x_2, x_3); -x_7 = l_Lean_nullKind; +x_7 = l_Lean_Parser_mkAntiquotSplice_formatter___closed__6; x_8 = l_Lean_Parser_ParserState_mkNode(x_6, x_7, x_5); return x_8; } @@ -2720,14 +2744,14 @@ x_6 = lean_apply_2(x_1, x_2, x_3); x_7 = lean_ctor_get(x_6, 4); lean_inc(x_7); x_8 = lean_box(0); -x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); +x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); lean_dec(x_7); if (x_9 == 0) { lean_object* x_10; lean_object* x_11; lean_dec(x_2); lean_dec(x_1); -x_10 = l_Lean_nullKind; +x_10 = l_Lean_Parser_mkAntiquotSplice_formatter___closed__6; x_11 = l_Lean_Parser_ParserState_mkNode(x_6, x_10, x_5); return x_11; } @@ -2735,7 +2759,7 @@ else { lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = l_Lean_Parser_manyAux(x_1, x_2, x_6); -x_13 = l_Lean_nullKind; +x_13 = l_Lean_Parser_mkAntiquotSplice_formatter___closed__6; x_14 = l_Lean_Parser_ParserState_mkNode(x_12, x_13, x_5); return x_14; } @@ -2786,9 +2810,19 @@ return x_1; static lean_object* _init_l_Lean_Parser_ident_formatter___closed__2() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_ident_formatter___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_ident_formatter___closed__3() { +_start: +{ lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; x_1 = l_Lean_Parser_ident_formatter___closed__1; -x_2 = l_Lean_identKind; +x_2 = l_Lean_Parser_ident_formatter___closed__2; x_3 = 1; x_4 = 0; x_5 = lean_box(x_3); @@ -2805,7 +2839,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_ident_formatter(lean_object* x_1, lean_ob _start: { lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_6 = l_Lean_Parser_ident_formatter___closed__2; +x_6 = l_Lean_Parser_ident_formatter___closed__3; x_7 = l_Lean_Parser_antiquotExpr_formatter___closed__4; x_8 = l_Lean_PrettyPrinter_Formatter_orelse_formatter(x_6, x_7, x_1, x_2, x_3, x_4, x_5); return x_8; @@ -2816,7 +2850,7 @@ static lean_object* _init_l_Lean_Parser_ident_parenthesizer___closed__1() { { lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; x_1 = l_Lean_Parser_ident_formatter___closed__1; -x_2 = l_Lean_identKind; +x_2 = l_Lean_Parser_ident_formatter___closed__2; x_3 = 1; x_4 = 0; x_5 = lean_box(x_3); @@ -2844,7 +2878,7 @@ static lean_object* _init_l_Lean_Parser_ident___elambda__1___closed__1() { { lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; x_1 = l_Lean_Parser_ident_formatter___closed__1; -x_2 = l_Lean_identKind; +x_2 = l_Lean_Parser_ident_formatter___closed__2; x_3 = 1; x_4 = 0; x_5 = l_Lean_Parser_mkAntiquot(x_1, x_2, x_3, x_4); @@ -2956,7 +2990,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_rawIdent_formatter(lean_object* x_1, lean _start: { lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_6 = l_Lean_Parser_ident_formatter___closed__2; +x_6 = l_Lean_Parser_ident_formatter___closed__3; x_7 = l_Lean_Parser_rawIdent_formatter___closed__1; x_8 = l_Lean_PrettyPrinter_Formatter_orelse_formatter(x_6, x_7, x_1, x_2, x_3, x_4, x_5); return x_8; @@ -3096,9 +3130,19 @@ return x_1; static lean_object* _init_l_Lean_Parser_numLit_formatter___closed__2() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_numLit_formatter___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_numLit_formatter___closed__3() { +_start: +{ lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; x_1 = l_Lean_Parser_numLit_formatter___closed__1; -x_2 = l_Lean_numLitKind; +x_2 = l_Lean_Parser_numLit_formatter___closed__2; x_3 = 1; x_4 = 0; x_5 = lean_box(x_3); @@ -3111,7 +3155,7 @@ lean_closure_set(x_7, 3, x_6); return x_7; } } -static lean_object* _init_l_Lean_Parser_numLit_formatter___closed__3() { +static lean_object* _init_l_Lean_Parser_numLit_formatter___closed__4() { _start: { lean_object* x_1; @@ -3123,8 +3167,8 @@ LEAN_EXPORT lean_object* l_Lean_Parser_numLit_formatter(lean_object* x_1, lean_o _start: { lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_6 = l_Lean_Parser_numLit_formatter___closed__2; -x_7 = l_Lean_Parser_numLit_formatter___closed__3; +x_6 = l_Lean_Parser_numLit_formatter___closed__3; +x_7 = l_Lean_Parser_numLit_formatter___closed__4; x_8 = l_Lean_PrettyPrinter_Formatter_orelse_formatter(x_6, x_7, x_1, x_2, x_3, x_4, x_5); return x_8; } @@ -3134,7 +3178,7 @@ static lean_object* _init_l_Lean_Parser_numLit_parenthesizer___closed__1() { { lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; x_1 = l_Lean_Parser_numLit_formatter___closed__1; -x_2 = l_Lean_numLitKind; +x_2 = l_Lean_Parser_numLit_formatter___closed__2; x_3 = 1; x_4 = 0; x_5 = lean_box(x_3); @@ -3170,7 +3214,7 @@ static lean_object* _init_l_Lean_Parser_numLit___elambda__1___closed__1() { { lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; x_1 = l_Lean_Parser_numLit_formatter___closed__1; -x_2 = l_Lean_numLitKind; +x_2 = l_Lean_Parser_numLit_formatter___closed__2; x_3 = 1; x_4 = 0; x_5 = l_Lean_Parser_mkAntiquot(x_1, x_2, x_3, x_4); @@ -3281,9 +3325,19 @@ return x_1; static lean_object* _init_l_Lean_Parser_scientificLit_formatter___closed__2() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_scientificLit_formatter___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_scientificLit_formatter___closed__3() { +_start: +{ lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; x_1 = l_Lean_Parser_scientificLit_formatter___closed__1; -x_2 = l_Lean_scientificLitKind; +x_2 = l_Lean_Parser_scientificLit_formatter___closed__2; x_3 = 1; x_4 = 0; x_5 = lean_box(x_3); @@ -3296,7 +3350,7 @@ lean_closure_set(x_7, 3, x_6); return x_7; } } -static lean_object* _init_l_Lean_Parser_scientificLit_formatter___closed__3() { +static lean_object* _init_l_Lean_Parser_scientificLit_formatter___closed__4() { _start: { lean_object* x_1; @@ -3308,8 +3362,8 @@ LEAN_EXPORT lean_object* l_Lean_Parser_scientificLit_formatter(lean_object* x_1, _start: { lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_6 = l_Lean_Parser_scientificLit_formatter___closed__2; -x_7 = l_Lean_Parser_scientificLit_formatter___closed__3; +x_6 = l_Lean_Parser_scientificLit_formatter___closed__3; +x_7 = l_Lean_Parser_scientificLit_formatter___closed__4; x_8 = l_Lean_PrettyPrinter_Formatter_orelse_formatter(x_6, x_7, x_1, x_2, x_3, x_4, x_5); return x_8; } @@ -3319,7 +3373,7 @@ static lean_object* _init_l_Lean_Parser_scientificLit_parenthesizer___closed__1( { lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; x_1 = l_Lean_Parser_scientificLit_formatter___closed__1; -x_2 = l_Lean_scientificLitKind; +x_2 = l_Lean_Parser_scientificLit_formatter___closed__2; x_3 = 1; x_4 = 0; x_5 = lean_box(x_3); @@ -3355,7 +3409,7 @@ static lean_object* _init_l_Lean_Parser_scientificLit___elambda__1___closed__1() { lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; x_1 = l_Lean_Parser_scientificLit_formatter___closed__1; -x_2 = l_Lean_scientificLitKind; +x_2 = l_Lean_Parser_scientificLit_formatter___closed__2; x_3 = 1; x_4 = 0; x_5 = l_Lean_Parser_mkAntiquot(x_1, x_2, x_3, x_4); @@ -3466,9 +3520,19 @@ return x_1; static lean_object* _init_l_Lean_Parser_strLit_formatter___closed__2() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_strLit_formatter___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_strLit_formatter___closed__3() { +_start: +{ lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; x_1 = l_Lean_Parser_strLit_formatter___closed__1; -x_2 = l_Lean_strLitKind; +x_2 = l_Lean_Parser_strLit_formatter___closed__2; x_3 = 1; x_4 = 0; x_5 = lean_box(x_3); @@ -3481,7 +3545,7 @@ lean_closure_set(x_7, 3, x_6); return x_7; } } -static lean_object* _init_l_Lean_Parser_strLit_formatter___closed__3() { +static lean_object* _init_l_Lean_Parser_strLit_formatter___closed__4() { _start: { lean_object* x_1; @@ -3493,8 +3557,8 @@ LEAN_EXPORT lean_object* l_Lean_Parser_strLit_formatter(lean_object* x_1, lean_o _start: { lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_6 = l_Lean_Parser_strLit_formatter___closed__2; -x_7 = l_Lean_Parser_strLit_formatter___closed__3; +x_6 = l_Lean_Parser_strLit_formatter___closed__3; +x_7 = l_Lean_Parser_strLit_formatter___closed__4; x_8 = l_Lean_PrettyPrinter_Formatter_orelse_formatter(x_6, x_7, x_1, x_2, x_3, x_4, x_5); return x_8; } @@ -3504,7 +3568,7 @@ static lean_object* _init_l_Lean_Parser_strLit_parenthesizer___closed__1() { { lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; x_1 = l_Lean_Parser_strLit_formatter___closed__1; -x_2 = l_Lean_strLitKind; +x_2 = l_Lean_Parser_strLit_formatter___closed__2; x_3 = 1; x_4 = 0; x_5 = lean_box(x_3); @@ -3540,7 +3604,7 @@ static lean_object* _init_l_Lean_Parser_strLit___elambda__1___closed__1() { { lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; x_1 = l_Lean_Parser_strLit_formatter___closed__1; -x_2 = l_Lean_strLitKind; +x_2 = l_Lean_Parser_strLit_formatter___closed__2; x_3 = 1; x_4 = 0; x_5 = l_Lean_Parser_mkAntiquot(x_1, x_2, x_3, x_4); @@ -3651,9 +3715,19 @@ return x_1; static lean_object* _init_l_Lean_Parser_charLit_formatter___closed__2() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_charLit_formatter___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_charLit_formatter___closed__3() { +_start: +{ lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; x_1 = l_Lean_Parser_charLit_formatter___closed__1; -x_2 = l_Lean_charLitKind; +x_2 = l_Lean_Parser_charLit_formatter___closed__2; x_3 = 1; x_4 = 0; x_5 = lean_box(x_3); @@ -3666,7 +3740,7 @@ lean_closure_set(x_7, 3, x_6); return x_7; } } -static lean_object* _init_l_Lean_Parser_charLit_formatter___closed__3() { +static lean_object* _init_l_Lean_Parser_charLit_formatter___closed__4() { _start: { lean_object* x_1; @@ -3678,8 +3752,8 @@ LEAN_EXPORT lean_object* l_Lean_Parser_charLit_formatter(lean_object* x_1, lean_ _start: { lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_6 = l_Lean_Parser_charLit_formatter___closed__2; -x_7 = l_Lean_Parser_charLit_formatter___closed__3; +x_6 = l_Lean_Parser_charLit_formatter___closed__3; +x_7 = l_Lean_Parser_charLit_formatter___closed__4; x_8 = l_Lean_PrettyPrinter_Formatter_orelse_formatter(x_6, x_7, x_1, x_2, x_3, x_4, x_5); return x_8; } @@ -3689,7 +3763,7 @@ static lean_object* _init_l_Lean_Parser_charLit_parenthesizer___closed__1() { { lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; x_1 = l_Lean_Parser_charLit_formatter___closed__1; -x_2 = l_Lean_charLitKind; +x_2 = l_Lean_Parser_charLit_formatter___closed__2; x_3 = 1; x_4 = 0; x_5 = lean_box(x_3); @@ -3725,7 +3799,7 @@ static lean_object* _init_l_Lean_Parser_charLit___elambda__1___closed__1() { { lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; x_1 = l_Lean_Parser_charLit_formatter___closed__1; -x_2 = l_Lean_charLitKind; +x_2 = l_Lean_Parser_charLit_formatter___closed__2; x_3 = 1; x_4 = 0; x_5 = l_Lean_Parser_mkAntiquot(x_1, x_2, x_3, x_4); @@ -3836,9 +3910,19 @@ return x_1; static lean_object* _init_l_Lean_Parser_nameLit_formatter___closed__2() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_nameLit_formatter___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_nameLit_formatter___closed__3() { +_start: +{ lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; x_1 = l_Lean_Parser_nameLit_formatter___closed__1; -x_2 = l_Lean_nameLitKind; +x_2 = l_Lean_Parser_nameLit_formatter___closed__2; x_3 = 1; x_4 = 0; x_5 = lean_box(x_3); @@ -3851,7 +3935,7 @@ lean_closure_set(x_7, 3, x_6); return x_7; } } -static lean_object* _init_l_Lean_Parser_nameLit_formatter___closed__3() { +static lean_object* _init_l_Lean_Parser_nameLit_formatter___closed__4() { _start: { lean_object* x_1; @@ -3863,8 +3947,8 @@ LEAN_EXPORT lean_object* l_Lean_Parser_nameLit_formatter(lean_object* x_1, lean_ _start: { lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_6 = l_Lean_Parser_nameLit_formatter___closed__2; -x_7 = l_Lean_Parser_nameLit_formatter___closed__3; +x_6 = l_Lean_Parser_nameLit_formatter___closed__3; +x_7 = l_Lean_Parser_nameLit_formatter___closed__4; x_8 = l_Lean_PrettyPrinter_Formatter_orelse_formatter(x_6, x_7, x_1, x_2, x_3, x_4, x_5); return x_8; } @@ -3874,7 +3958,7 @@ static lean_object* _init_l_Lean_Parser_nameLit_parenthesizer___closed__1() { { lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; x_1 = l_Lean_Parser_nameLit_formatter___closed__1; -x_2 = l_Lean_nameLitKind; +x_2 = l_Lean_Parser_nameLit_formatter___closed__2; x_3 = 1; x_4 = 0; x_5 = lean_box(x_3); @@ -3910,7 +3994,7 @@ static lean_object* _init_l_Lean_Parser_nameLit___elambda__1___closed__1() { { lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; x_1 = l_Lean_Parser_nameLit_formatter___closed__1; -x_2 = l_Lean_nameLitKind; +x_2 = l_Lean_Parser_nameLit_formatter___closed__2; x_3 = 1; x_4 = 0; x_5 = l_Lean_Parser_mkAntiquot(x_1, x_2, x_3, x_4); @@ -4010,11 +4094,29 @@ x_1 = l_Lean_Parser_nameLit___closed__4; return x_1; } } +static lean_object* _init_l_Lean_Parser_group_formatter___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("group", 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_group_formatter___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_group_formatter___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Parser_group_formatter(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; -x_7 = l_Lean_groupKind; +x_7 = l_Lean_Parser_group_formatter___closed__2; x_8 = l_Lean_PrettyPrinter_Formatter_node_formatter(x_7, x_1, x_2, x_3, x_4, x_5, x_6); return x_8; } @@ -4023,7 +4125,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_group_parenthesizer(lean_object* x_1, lea _start: { lean_object* x_7; lean_object* x_8; -x_7 = l_Lean_groupKind; +x_7 = l_Lean_Parser_group_formatter___closed__2; x_8 = l_Lean_PrettyPrinter_Parenthesizer_node_parenthesizer(x_7, x_1, x_2, x_3, x_4, x_5, x_6); return x_8; } @@ -4034,7 +4136,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_group(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; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l_Lean_groupKind; +x_3 = l_Lean_Parser_group_formatter___closed__2; x_4 = l_Lean_Parser_nodeInfo(x_3, x_2); x_5 = lean_ctor_get(x_1, 1); lean_inc(x_5); @@ -4514,15 +4616,14 @@ goto _start; } else { -lean_object* x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; -x_21 = l_Lean_nullKind; -x_22 = l_Lean_Syntax_isNodeOf(x_12, x_21, x_8); -x_23 = lean_box(x_22); -x_24 = lean_array_push(x_7, x_23); +uint8_t x_21; lean_object* x_22; lean_object* x_23; +x_21 = l_Lean_Syntax_matchesNull(x_12, x_8); +x_22 = lean_box(x_21); +x_23 = lean_array_push(x_7, x_22); x_4 = x_11; x_5 = x_16; x_6 = lean_box(0); -x_7 = x_24; +x_7 = x_23; goto _start; } } @@ -6386,25 +6487,25 @@ lean_dec(x_1); return x_5; } } -static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__1() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("termRegister_parser_alias(Kind:=_)___", 37); +x_1 = lean_mk_string_from_bytes("termRegister_parser_alias(Kind:=_)____", 38); return x_1; } } -static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__2() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__4; -x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__1; +x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__3() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__3() { _start: { lean_object* x_1; @@ -6412,17 +6513,17 @@ x_1 = lean_mk_string_from_bytes("andthen", 7); return x_1; } } -static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__4() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__3; +x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__5() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__5() { _start: { lean_object* x_1; @@ -6430,35 +6531,17 @@ x_1 = lean_mk_string_from_bytes("register_parser_alias", 21); return x_1; } } -static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__6() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__5; +x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__5; x_2 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__7() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("group", 5); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__7; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__9() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__7() { _start: { lean_object* x_1; lean_object* x_2; @@ -6468,7 +6551,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__10() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__8() { _start: { lean_object* x_1; @@ -6476,11 +6559,11 @@ x_1 = lean_mk_string_from_bytes("kind", 4); return x_1; } } -static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__11() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__9() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__10; +x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__8; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -6488,13 +6571,13 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__12() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__4; -x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__9; -x_3 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__11; +x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__4; +x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__7; +x_3 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__9; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -6502,7 +6585,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__13() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__11() { _start: { lean_object* x_1; @@ -6510,23 +6593,23 @@ x_1 = lean_mk_string_from_bytes(" := ", 4); return x_1; } } -static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__14() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__13; +x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__11; x_2 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__15() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__4; -x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__12; -x_3 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__14; +x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__4; +x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__10; +x_3 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__12; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -6534,7 +6617,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__16() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -6546,13 +6629,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__17() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__4; -x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__15; -x_3 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__16; +x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__4; +x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__13; +x_3 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__14; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -6560,7 +6643,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__18() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__16() { _start: { lean_object* x_1; lean_object* x_2; @@ -6570,13 +6653,13 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__19() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__4; -x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__17; -x_3 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__18; +x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__4; +x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__15; +x_3 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__16; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -6584,37 +6667,37 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__20() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__8; -x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__19; +x_1 = l_Lean_Parser_group_formatter___closed__2; +x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__17; x_3 = lean_alloc_ctor(1, 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_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__21() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_optional_formatter___closed__2; -x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__20; +x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__18; x_3 = lean_alloc_ctor(1, 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_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__22() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__4; -x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__6; -x_3 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__21; +x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__4; +x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__6; +x_3 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__19; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -6622,7 +6705,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__23() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__21() { _start: { lean_object* x_1; @@ -6630,45 +6713,45 @@ x_1 = lean_mk_string_from_bytes("strLit", 6); return x_1; } } -static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__24() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__4; -x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__23; +x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__21; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__25() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__23() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__24; +x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__22; x_2 = lean_alloc_ctor(8, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__26() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_optional_formatter___closed__2; -x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__25; +x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__23; x_3 = lean_alloc_ctor(1, 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_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__27() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__4; -x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__22; -x_3 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__26; +x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__4; +x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__20; +x_3 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__24; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -6676,7 +6759,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__28() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -6686,23 +6769,91 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__29() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__27() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__28; +x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__26; x_2 = lean_alloc_ctor(8, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__30() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__28() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__4; +x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__25; +x_3 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__27; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__29() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("colGt", 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__30() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__29; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__31() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__30; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__32() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__4; +x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__31; +x_3 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__14; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__33() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_optional_formatter___closed__2; +x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__32; +x_3 = lean_alloc_ctor(1, 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_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__34() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__4; -x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__27; -x_3 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__29; +x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__4; +x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__28; +x_3 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__33; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -6710,13 +6861,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__31() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__35() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__2; +x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__30; +x_3 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__34; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -6724,149 +6875,242 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29______() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29________() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__31; +x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__35; return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__1() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("expected non-overloaded constant name", 37); +x_1 = lean_mk_string_from_bytes("choice", 6); return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__2() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Term", 4); +x_1 = lean_mk_string_from_bytes("term{}", 6); return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__3() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__4; -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__2; +x_1 = lean_box(0); +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__4() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__5() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("do", 2); +x_1 = lean_mk_string_from_bytes("{", 1); return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__5() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__6() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__3; -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__4; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("}", 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(2u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__6() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__8() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("doSeqIndent", 11); +x_1 = lean_mk_string_from_bytes("structInst", 10); return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__7() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__9() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__3; -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__6; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__8() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__10() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("null", 4); +x_1 = lean_mk_string_from_bytes("optEllipsis", 11); return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__9() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__11() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__8; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(6u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__10() { +LEAN_EXPORT lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1(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; 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; 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; 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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_6 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__5; +lean_inc(x_3); +x_7 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_7, 0, x_3); +lean_ctor_set(x_7, 1, x_6); +x_8 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__6; +x_9 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_9, 0, x_3); +lean_ctor_set(x_9, 1, x_8); +x_10 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__7; +lean_inc(x_7); +x_11 = lean_array_push(x_10, x_7); +lean_inc(x_9); +x_12 = lean_array_push(x_11, x_9); +x_13 = lean_box(2); +x_14 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__4; +x_15 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_14); +lean_ctor_set(x_15, 2, x_12); +x_16 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__8; +lean_inc(x_1); +x_17 = lean_name_mk_string(x_1, x_16); +x_18 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__9; +x_19 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_19, 0, x_13); +lean_ctor_set(x_19, 1, x_2); +lean_ctor_set(x_19, 2, x_18); +x_20 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__10; +x_21 = lean_name_mk_string(x_1, x_20); +x_22 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__11; +lean_inc(x_19); +x_23 = lean_array_push(x_22, x_19); +x_24 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_24, 0, x_13); +lean_ctor_set(x_24, 1, x_21); +lean_ctor_set(x_24, 2, x_23); +x_25 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__12; +x_26 = lean_array_push(x_25, x_7); +lean_inc(x_19); +x_27 = lean_array_push(x_26, x_19); +lean_inc(x_19); +x_28 = lean_array_push(x_27, x_19); +x_29 = lean_array_push(x_28, x_24); +x_30 = lean_array_push(x_29, x_19); +x_31 = lean_array_push(x_30, x_9); +x_32 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_32, 0, x_13); +lean_ctor_set(x_32, 1, x_17); +lean_ctor_set(x_32, 2, x_31); +x_33 = lean_array_push(x_10, x_15); +x_34 = lean_array_push(x_33, x_32); +x_35 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__2; +x_36 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_36, 0, x_13); +lean_ctor_set(x_36, 1, x_35); +lean_ctor_set(x_36, 2, x_34); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_5); +return x_37; +} +} +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("doSeqItem", 9); +x_1 = lean_mk_string_from_bytes("expected non-overloaded constant name", 37); return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__11() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__3; -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__10; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Term", 4); +return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__12() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("doExpr", 6); +x_1 = lean_mk_string_from_bytes("do", 2); return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__13() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__4() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__3; -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__12; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("doSeqIndent", 11); +return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__14() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__5() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("app", 3); +x_1 = lean_mk_string_from_bytes("doSeqItem", 9); return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__15() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__6() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__3; -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__14; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("doExpr", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("app", 3); +return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__16() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__8() { _start: { lean_object* x_1; @@ -6874,22 +7118,22 @@ x_1 = lean_mk_string_from_bytes("Parser.registerAlias", 20); return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__17() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__16; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__8; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__18() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__16; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__8; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__17; +x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__9; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -6897,7 +7141,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__19() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -6907,7 +7151,7 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__20() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__12() { _start: { lean_object* x_1; @@ -6915,51 +7159,25 @@ x_1 = lean_mk_string_from_bytes("registerAlias", 13); return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__21() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__19; -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__20; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__11; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__12; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__22() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__14() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__4; -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__20; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__23() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__22; -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___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__24() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__23; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Level_PP_Result_quote___spec__1), 2, 0); +return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__25() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__15() { _start: { lean_object* x_1; @@ -6967,17 +7185,7 @@ x_1 = lean_mk_string_from_bytes("namedArgument", 13); return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__26() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__3; -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__25; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__27() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__16() { _start: { lean_object* x_1; @@ -6985,22 +7193,22 @@ x_1 = lean_mk_string_from_bytes("kind?", 5); return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__28() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__17() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__27; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__16; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__29() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__27; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__16; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__28; +x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__17; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -7008,17 +7216,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__30() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__27; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__16; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__31() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__20() { _start: { lean_object* x_1; @@ -7026,7 +7234,7 @@ x_1 = lean_mk_string_from_bytes(":=", 2); return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__32() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__21() { _start: { lean_object* x_1; @@ -7034,22 +7242,22 @@ x_1 = lean_mk_string_from_bytes("some", 4); return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__33() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__22() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__32; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__21; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__34() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__32; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__21; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__33; +x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__22; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -7057,17 +7265,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__35() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__32; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__21; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__36() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__25() { _start: { lean_object* x_1; @@ -7075,69 +7283,51 @@ x_1 = lean_mk_string_from_bytes("Option", 6); return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__37() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__36; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__25; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__38() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__27() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__37; -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__32; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__26; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__21; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__39() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__38; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__27; 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___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__40() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__29() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__39; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__28; x_3 = lean_alloc_ctor(1, 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___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__41() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(1u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__42() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(2u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__43() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__30() { _start: { lean_object* x_1; lean_object* x_2; @@ -7146,31 +7336,22 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__44() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(3u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__45() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__31() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(0u); +x_1 = lean_unsigned_to_nat(4u); x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__46() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__32() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__9; -x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__45; +x_2 = l_Lean_Parser_mkAntiquotSplice_formatter___closed__6; +x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__9; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -7178,7 +7359,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__47() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__33() { _start: { lean_object* x_1; @@ -7186,22 +7367,22 @@ x_1 = lean_mk_string_from_bytes("PrettyPrinter.Formatter.registerAlias", 37); return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__48() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__34() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__47; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__33; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__49() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__35() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__47; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__33; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__48; +x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__34; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -7209,7 +7390,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__50() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__36() { _start: { lean_object* x_1; @@ -7217,17 +7398,17 @@ x_1 = lean_mk_string_from_bytes("PrettyPrinter", 13); return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__51() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__37() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__50; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__36; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__52() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__38() { _start: { lean_object* x_1; @@ -7235,81 +7416,27 @@ x_1 = lean_mk_string_from_bytes("Formatter", 9); return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__53() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__51; -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__52; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__54() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__53; -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__20; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__55() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__39() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__2; -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__50; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__37; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__38; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__56() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__40() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__55; -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__52; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__39; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__12; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__57() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__56; -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__20; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__58() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__57; -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___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__59() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__58; -x_3 = lean_alloc_ctor(1, 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___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__60() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__41() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -7319,7 +7446,7 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__61() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__42() { _start: { lean_object* x_1; @@ -7327,22 +7454,22 @@ x_1 = lean_mk_string_from_bytes("PrettyPrinter.Parenthesizer.registerAlias", 41) return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__62() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__43() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__61; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__42; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__63() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__44() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__61; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__42; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__62; +x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__43; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -7350,7 +7477,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__64() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__45() { _start: { lean_object* x_1; @@ -7358,81 +7485,46 @@ x_1 = lean_mk_string_from_bytes("Parenthesizer", 13); return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__65() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__51; -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__64; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__66() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__65; -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__20; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__67() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__46() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__55; -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__64; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__37; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__45; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__68() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__47() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__67; -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__20; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__46; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__12; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__69() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__48() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__68; -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___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__70() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__69; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); +x_2 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_parenthesizer___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__71() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__49() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_parenthesizer___closed__1; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(3u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__72() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__50() { _start: { lean_object* x_1; @@ -7440,17 +7532,7 @@ x_1 = lean_mk_string_from_bytes("quotedName", 10); return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__73() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__3; -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__72; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__74() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__51() { _start: { lean_object* x_1; @@ -7458,7 +7540,7 @@ x_1 = lean_mk_string_from_bytes(".", 1); return x_1; } } -static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__75() { +static lean_object* _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__52() { _start: { lean_object* x_1; @@ -7466,458 +7548,787 @@ x_1 = lean_mk_string_from_bytes("`", 1); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________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, lean_object* x_7) { _start: { -lean_object* x_4; uint8_t x_5; -x_4 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__2; -lean_inc(x_1); -x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); -if (x_5 == 0) -{ -lean_object* x_6; lean_object* x_7; -lean_dec(x_2); -lean_dec(x_1); -x_6 = lean_box(1); -x_7 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_7, 0, x_6); -lean_ctor_set(x_7, 1, x_3); -return x_7; -} -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; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_8 = lean_unsigned_to_nat(1u); +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; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_8 = lean_unsigned_to_nat(2u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); -x_10 = lean_unsigned_to_nat(2u); +x_10 = lean_unsigned_to_nat(3u); x_11 = l_Lean_Syntax_getArg(x_1, x_10); -x_12 = lean_unsigned_to_nat(3u); +x_12 = lean_unsigned_to_nat(4u); x_13 = l_Lean_Syntax_getArg(x_1, x_12); -lean_dec(x_1); -x_14 = l_Lean_Syntax_getOptional_x3f(x_11); -lean_dec(x_11); +x_14 = l_Lean_Syntax_getOptional_x3f(x_13); +lean_dec(x_13); x_15 = l_Lean_Syntax_getOptional_x3f(x_9); lean_dec(x_9); -x_16 = l_Lean_Syntax_getId(x_13); -lean_inc(x_2); +x_16 = l_Lean_Syntax_getId(x_11); +lean_inc(x_6); lean_inc(x_16); -x_17 = l_Lean_Macro_resolveGlobalName(x_16, x_2, x_3); -if (lean_obj_tag(x_17) == 0) +x_17 = l_Lean_Macro_resolveGlobalName(x_16, x_6, x_7); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_18; -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -if (lean_obj_tag(x_18) == 0) +lean_object* x_218; +x_218 = lean_box(0); +x_18 = x_218; +goto block_217; +} +else +{ +uint8_t x_219; +x_219 = !lean_is_exclusive(x_14); +if (x_219 == 0) +{ +x_18 = x_14; +goto block_217; +} +else +{ +lean_object* x_220; lean_object* x_221; +x_220 = lean_ctor_get(x_14, 0); +lean_inc(x_220); +lean_dec(x_14); +x_221 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_221, 0, x_220); +x_18 = x_221; +goto block_217; +} +} +block_217: { lean_object* x_19; lean_object* x_20; lean_object* x_21; +if (lean_obj_tag(x_15) == 0) +{ +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_199; lean_object* x_200; lean_object* x_201; +x_199 = lean_ctor_get(x_17, 0); +lean_inc(x_199); +x_200 = lean_ctor_get(x_17, 1); +lean_inc(x_200); +lean_dec(x_17); +x_201 = lean_box(0); +x_19 = x_201; +x_20 = x_199; +x_21 = x_200; +goto block_198; +} +else +{ +uint8_t x_202; +lean_dec(x_18); lean_dec(x_16); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_202 = !lean_is_exclusive(x_17); +if (x_202 == 0) +{ +return x_17; +} +else +{ +lean_object* x_203; lean_object* x_204; lean_object* x_205; +x_203 = lean_ctor_get(x_17, 0); +x_204 = lean_ctor_get(x_17, 1); +lean_inc(x_204); +lean_inc(x_203); +lean_dec(x_17); +x_205 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_205, 0, x_203); +lean_ctor_set(x_205, 1, x_204); +return x_205; +} +} +} +else +{ +if (lean_obj_tag(x_17) == 0) +{ +uint8_t x_206; +x_206 = !lean_is_exclusive(x_15); +if (x_206 == 0) +{ +lean_object* x_207; lean_object* x_208; +x_207 = lean_ctor_get(x_17, 0); +lean_inc(x_207); +x_208 = lean_ctor_get(x_17, 1); +lean_inc(x_208); +lean_dec(x_17); +x_19 = x_15; +x_20 = x_207; +x_21 = x_208; +goto block_198; +} +else +{ +lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; +x_209 = lean_ctor_get(x_15, 0); +lean_inc(x_209); lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -x_19 = lean_ctor_get(x_17, 1); -lean_inc(x_19); +x_210 = lean_ctor_get(x_17, 0); +lean_inc(x_210); +x_211 = lean_ctor_get(x_17, 1); +lean_inc(x_211); lean_dec(x_17); -x_20 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__1; -x_21 = l_Lean_Macro_throwError___rarg(x_20, x_2, x_19); +x_212 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_212, 0, x_209); +x_19 = x_212; +x_20 = x_210; +x_21 = x_211; +goto block_198; +} +} +else +{ +uint8_t x_213; +lean_dec(x_18); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); lean_dec(x_2); -return x_21; +x_213 = !lean_is_exclusive(x_17); +if (x_213 == 0) +{ +return x_17; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_22 = lean_ctor_get(x_18, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_17, 1); -lean_inc(x_23); +lean_object* x_214; lean_object* x_215; lean_object* x_216; +x_214 = lean_ctor_get(x_17, 0); +x_215 = lean_ctor_get(x_17, 1); +lean_inc(x_215); +lean_inc(x_214); lean_dec(x_17); -x_24 = lean_ctor_get(x_18, 1); -lean_inc(x_24); +x_216 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_216, 0, x_214); +lean_ctor_set(x_216, 1, x_215); +return x_216; +} +} +} +block_198: +{ +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_22; lean_object* x_23; +lean_dec(x_19); lean_dec(x_18); -x_25 = lean_ctor_get(x_22, 0); +lean_dec(x_16); +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_22 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__1; +x_23 = l_Lean_Macro_throwError___rarg(x_22, x_6, x_21); +lean_dec(x_6); +return x_23; +} +else +{ +lean_object* 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; +x_24 = lean_ctor_get(x_20, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_20, 1); lean_inc(x_25); -x_26 = lean_ctor_get(x_22, 1); -lean_inc(x_26); -lean_dec(x_22); -if (lean_obj_tag(x_26) == 0) +if (lean_is_exclusive(x_20)) { + lean_ctor_release(x_20, 0); + lean_ctor_release(x_20, 1); + x_26 = x_20; +} else { + lean_dec_ref(x_20); + x_26 = lean_box(0); +} +x_27 = lean_ctor_get(x_24, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_24, 1); +lean_inc(x_28); +if (lean_is_exclusive(x_24)) { + lean_ctor_release(x_24, 0); + lean_ctor_release(x_24, 1); + x_29 = x_24; +} else { + lean_dec_ref(x_24); + x_29 = lean_box(0); +} +if (lean_obj_tag(x_28) == 0) { -if (lean_obj_tag(x_24) == 0) +if (lean_obj_tag(x_25) == 0) { -if (lean_obj_tag(x_14) == 0) +if (lean_obj_tag(x_19) == 0) { -uint8_t x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; -x_157 = 1; +uint8_t x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; +x_189 = 1; lean_inc(x_16); -x_158 = l_Lean_Name_toString(x_16, x_157); -x_159 = lean_box(2); -x_160 = l_Lean_Syntax_mkStrLit(x_158, x_159); -lean_dec(x_158); -x_27 = x_160; -goto block_156; +x_190 = l_Lean_Name_toString(x_16, x_189); +x_191 = lean_box(2); +x_192 = l_Lean_Syntax_mkStrLit(x_190, x_191); +lean_dec(x_190); +x_30 = x_192; +goto block_188; } else { -lean_object* x_161; -x_161 = lean_ctor_get(x_14, 0); -lean_inc(x_161); -lean_dec(x_14); -x_27 = x_161; -goto block_156; +lean_object* x_193; +x_193 = lean_ctor_get(x_19, 0); +lean_inc(x_193); +lean_dec(x_19); +x_30 = x_193; +goto block_188; } } else { -lean_object* x_162; lean_object* x_163; +lean_object* x_194; lean_object* x_195; +lean_dec(x_29); +lean_dec(x_27); +lean_dec(x_26); lean_dec(x_25); -lean_dec(x_24); +lean_dec(x_19); +lean_dec(x_18); lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -x_162 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__1; -x_163 = l_Lean_Macro_throwError___rarg(x_162, x_2, x_23); +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_3); lean_dec(x_2); -return x_163; +x_194 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__1; +x_195 = l_Lean_Macro_throwError___rarg(x_194, x_6, x_21); +lean_dec(x_6); +return x_195; } } else { -lean_object* x_164; lean_object* x_165; +lean_object* x_196; lean_object* x_197; +lean_dec(x_29); +lean_dec(x_28); +lean_dec(x_27); lean_dec(x_26); lean_dec(x_25); -lean_dec(x_24); +lean_dec(x_19); +lean_dec(x_18); lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -x_164 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__1; -x_165 = l_Lean_Macro_throwError___rarg(x_164, x_2, x_23); +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_3); lean_dec(x_2); -return x_165; +x_196 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__1; +x_197 = l_Lean_Macro_throwError___rarg(x_196, x_6, x_21); +lean_dec(x_6); +return x_197; } -block_156: +block_188: { -lean_object* x_28; lean_object* x_29; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; +lean_inc(x_6); +x_31 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_6, x_21); +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +if (lean_is_exclusive(x_31)) { + lean_ctor_release(x_31, 0); + lean_ctor_release(x_31, 1); + x_34 = x_31; +} else { + lean_dec_ref(x_31); + x_34 = lean_box(0); +} +x_35 = lean_ctor_get(x_6, 2); +lean_inc(x_35); +x_36 = lean_ctor_get(x_6, 1); +lean_inc(x_36); +lean_dec(x_6); +x_37 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__2; lean_inc(x_2); -x_28 = l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(x_2, x_23); -x_29 = lean_ctor_get(x_28, 0); -lean_inc(x_29); -x_30 = lean_ctor_get(x_28, 1); -lean_inc(x_30); -if (lean_is_exclusive(x_28)) { - lean_ctor_release(x_28, 0); - lean_ctor_release(x_28, 1); - x_31 = x_28; +x_38 = lean_name_mk_string(x_2, x_37); +x_39 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__3; +lean_inc(x_38); +x_40 = lean_name_mk_string(x_38, x_39); +lean_inc(x_32); +x_41 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_41, 0, x_32); +lean_ctor_set(x_41, 1, x_39); +x_42 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__4; +lean_inc(x_38); +x_43 = lean_name_mk_string(x_38, x_42); +x_44 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__5; +lean_inc(x_38); +x_45 = lean_name_mk_string(x_38, x_44); +x_46 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__6; +lean_inc(x_38); +x_47 = lean_name_mk_string(x_38, x_46); +x_48 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__7; +lean_inc(x_38); +x_49 = lean_name_mk_string(x_38, x_48); +x_50 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__13; +lean_inc(x_35); +lean_inc(x_36); +x_51 = l_Lean_addMacroScope(x_36, x_50, x_35); +x_52 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__12; +x_53 = lean_name_mk_string(x_2, x_52); +x_54 = lean_box(0); +if (lean_is_scalar(x_29)) { + x_55 = lean_alloc_ctor(0, 2, 0); +} else { + x_55 = x_29; +} +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +if (lean_is_scalar(x_26)) { + x_56 = lean_alloc_ctor(1, 2, 0); } else { - lean_dec_ref(x_28); - x_31 = lean_box(0); + x_56 = x_26; } -x_32 = lean_ctor_get(x_2, 2); +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_54); +x_57 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__10; lean_inc(x_32); -x_33 = lean_ctor_get(x_2, 1); -lean_inc(x_33); -lean_dec(x_2); -x_34 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__4; -lean_inc(x_29); -x_35 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_35, 0, x_29); -lean_ctor_set(x_35, 1, x_34); -x_36 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__21; +x_58 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_58, 0, x_32); +lean_ctor_set(x_58, 1, x_57); +lean_ctor_set(x_58, 2, x_51); +lean_ctor_set(x_58, 3, x_56); +x_59 = l_Lean_Parser_mkAntiquotSplice_formatter___closed__6; +lean_inc(x_38); +x_60 = lean_alloc_closure((void*)(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___boxed), 5, 2); +lean_closure_set(x_60, 0, x_38); +lean_closure_set(x_60, 1, x_59); +x_61 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__14; +x_62 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Unhygienic_instMonadQuotationUnhygienic___spec__2___rarg), 4, 2); +lean_closure_set(x_62, 0, x_61); +lean_closure_set(x_62, 1, x_60); +x_63 = l_Lean_Unhygienic_run___rarg(x_62); +x_64 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__15; +lean_inc(x_38); +x_65 = lean_name_mk_string(x_38, x_64); +x_66 = l_Lean_Parser_antiquotNestedExpr_formatter___closed__3; lean_inc(x_32); -lean_inc(x_33); -x_37 = l_Lean_addMacroScope(x_33, x_36, x_32); -x_38 = lean_box(0); -x_39 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__18; -x_40 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__24; -lean_inc(x_29); -x_41 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_41, 0, x_29); -lean_ctor_set(x_41, 1, x_39); -lean_ctor_set(x_41, 2, x_37); -lean_ctor_set(x_41, 3, x_40); -x_42 = l_Lean_Parser_antiquotNestedExpr_formatter___closed__3; -lean_inc(x_29); -x_43 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_43, 0, x_29); -lean_ctor_set(x_43, 1, x_42); -x_44 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__30; +x_67 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_67, 0, x_32); +lean_ctor_set(x_67, 1, x_66); +x_68 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__19; +lean_inc(x_35); +lean_inc(x_36); +x_69 = l_Lean_addMacroScope(x_36, x_68, x_35); +x_70 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__18; lean_inc(x_32); -lean_inc(x_33); -x_45 = l_Lean_addMacroScope(x_33, x_44, x_32); -x_46 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__29; -lean_inc(x_29); -x_47 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_47, 0, x_29); -lean_ctor_set(x_47, 1, x_46); -lean_ctor_set(x_47, 2, x_45); -lean_ctor_set(x_47, 3, x_38); -x_48 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__31; -lean_inc(x_29); -x_49 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_49, 0, x_29); -lean_ctor_set(x_49, 1, x_48); -x_50 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__35; +x_71 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_71, 0, x_32); +lean_ctor_set(x_71, 1, x_70); +lean_ctor_set(x_71, 2, x_69); +lean_ctor_set(x_71, 3, x_54); +x_72 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__20; lean_inc(x_32); -lean_inc(x_33); -x_51 = l_Lean_addMacroScope(x_33, x_50, x_32); -x_52 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__34; -x_53 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__40; -lean_inc(x_29); -x_54 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_54, 0, x_29); -lean_ctor_set(x_54, 1, x_52); -lean_ctor_set(x_54, 2, x_51); -lean_ctor_set(x_54, 3, x_53); -lean_inc(x_25); -x_55 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_38, x_25); -x_56 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__42; -x_57 = lean_array_push(x_56, x_54); -x_58 = l_Lean_Parser_antiquotNestedExpr_formatter___closed__7; -lean_inc(x_29); -x_59 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_59, 0, x_29); -lean_ctor_set(x_59, 1, x_58); -x_60 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__43; -x_61 = lean_array_push(x_60, x_43); -x_62 = lean_array_push(x_61, x_47); -x_63 = lean_array_push(x_62, x_49); -x_64 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__44; +x_73 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_73, 0, x_32); +lean_ctor_set(x_73, 1, x_72); +x_74 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__24; +lean_inc(x_35); +lean_inc(x_36); +x_75 = l_Lean_addMacroScope(x_36, x_74, x_35); +x_76 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__23; +x_77 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__29; +lean_inc(x_32); +x_78 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_78, 0, x_32); +lean_ctor_set(x_78, 1, x_76); +lean_ctor_set(x_78, 2, x_75); +lean_ctor_set(x_78, 3, x_77); lean_inc(x_27); -x_65 = lean_array_push(x_64, x_27); -lean_inc(x_13); -x_66 = lean_array_push(x_65, x_13); -x_67 = lean_array_push(x_56, x_41); -x_68 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__54; +x_79 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_54, x_27); +x_80 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__7; +x_81 = lean_array_push(x_80, x_78); +x_82 = l_Lean_Parser_antiquotNestedExpr_formatter___closed__7; lean_inc(x_32); -lean_inc(x_33); -x_69 = l_Lean_addMacroScope(x_33, x_68, x_32); -x_70 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__49; -x_71 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__59; -lean_inc(x_29); -x_72 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_72, 0, x_29); -lean_ctor_set(x_72, 1, x_70); -lean_ctor_set(x_72, 2, x_69); -lean_ctor_set(x_72, 3, x_71); -x_73 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__60; -x_74 = l_Lean_Name_append(x_16, x_73); -lean_inc(x_13); -x_75 = l_Lean_mkIdentFrom(x_13, x_74); -x_76 = lean_array_push(x_56, x_27); -lean_inc(x_76); -x_77 = lean_array_push(x_76, x_75); -x_78 = lean_box(2); -x_79 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__9; -x_80 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_80, 0, x_78); -lean_ctor_set(x_80, 1, x_79); -lean_ctor_set(x_80, 2, x_77); -x_81 = lean_array_push(x_56, x_72); -x_82 = lean_array_push(x_81, x_80); -x_83 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__15; -x_84 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_84, 0, x_78); -lean_ctor_set(x_84, 1, x_83); -lean_ctor_set(x_84, 2, x_82); -x_85 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__41; -x_86 = lean_array_push(x_85, x_84); -x_87 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__13; -x_88 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_88, 0, x_78); -lean_ctor_set(x_88, 1, x_87); -lean_ctor_set(x_88, 2, x_86); -x_89 = lean_array_push(x_56, x_88); -x_90 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__46; -x_91 = lean_array_push(x_89, x_90); -x_92 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__11; -x_93 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_93, 0, x_78); -lean_ctor_set(x_93, 1, x_92); -lean_ctor_set(x_93, 2, x_91); -x_94 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__66; -x_95 = l_Lean_addMacroScope(x_33, x_94, x_32); -x_96 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__63; -x_97 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__70; -x_98 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_98, 0, x_29); -lean_ctor_set(x_98, 1, x_96); -lean_ctor_set(x_98, 2, x_95); -lean_ctor_set(x_98, 3, x_97); -x_99 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__71; -x_100 = l_Lean_Name_append(x_16, x_99); +x_83 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_83, 0, x_32); +lean_ctor_set(x_83, 1, x_82); +x_84 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__30; +x_85 = lean_array_push(x_84, x_67); +x_86 = lean_array_push(x_85, x_71); +x_87 = lean_array_push(x_86, x_73); +x_88 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__31; +lean_inc(x_30); +x_89 = lean_array_push(x_88, x_30); +lean_inc(x_11); +x_90 = lean_array_push(x_89, x_11); +x_91 = lean_array_push(x_80, x_58); +x_92 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__40; +lean_inc(x_35); +lean_inc(x_36); +x_93 = l_Lean_addMacroScope(x_36, x_92, x_35); +x_94 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__36; +x_95 = lean_name_mk_string(x_3, x_94); +x_96 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__38; +lean_inc(x_95); +x_97 = lean_name_mk_string(x_95, x_96); +x_98 = lean_name_mk_string(x_97, x_52); +x_99 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_99, 0, x_98); +lean_ctor_set(x_99, 1, x_54); +x_100 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_100, 0, x_99); +lean_ctor_set(x_100, 1, x_54); +x_101 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__35; +lean_inc(x_32); +x_102 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_102, 0, x_32); +lean_ctor_set(x_102, 1, x_101); +lean_ctor_set(x_102, 2, x_93); +lean_ctor_set(x_102, 3, x_100); +x_103 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__41; +x_104 = l_Lean_Name_append(x_16, x_103); +lean_inc(x_11); +x_105 = l_Lean_mkIdentFrom(x_11, x_104); +x_106 = lean_array_push(x_80, x_30); +lean_inc(x_106); +x_107 = lean_array_push(x_106, x_105); +x_108 = lean_box(2); +x_109 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_109, 0, x_108); +lean_ctor_set(x_109, 1, x_59); +lean_ctor_set(x_109, 2, x_107); +x_110 = lean_array_push(x_80, x_102); +x_111 = lean_array_push(x_110, x_109); +lean_inc(x_49); +x_112 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_112, 0, x_108); +lean_ctor_set(x_112, 1, x_49); +lean_ctor_set(x_112, 2, x_111); +x_113 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__11; +x_114 = lean_array_push(x_113, x_112); +lean_inc(x_47); +x_115 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_115, 0, x_108); +lean_ctor_set(x_115, 1, x_47); +lean_ctor_set(x_115, 2, x_114); +x_116 = lean_array_push(x_80, x_115); +x_117 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__32; +x_118 = lean_array_push(x_116, x_117); +lean_inc(x_45); +x_119 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_119, 0, x_108); +lean_ctor_set(x_119, 1, x_45); +lean_ctor_set(x_119, 2, x_118); +x_120 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__47; +x_121 = l_Lean_addMacroScope(x_36, x_120, x_35); +x_122 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__45; +x_123 = lean_name_mk_string(x_95, x_122); +x_124 = lean_name_mk_string(x_123, x_52); +x_125 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_125, 0, x_124); +lean_ctor_set(x_125, 1, x_54); +x_126 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_126, 0, x_125); +lean_ctor_set(x_126, 1, x_54); +x_127 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__44; +x_128 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_128, 0, x_32); +lean_ctor_set(x_128, 1, x_127); +lean_ctor_set(x_128, 2, x_121); +lean_ctor_set(x_128, 3, x_126); +x_129 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__48; +x_130 = l_Lean_Name_append(x_16, x_129); lean_dec(x_16); -x_101 = l_Lean_mkIdentFrom(x_13, x_100); -x_102 = lean_array_push(x_76, x_101); -x_103 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_103, 0, x_78); -lean_ctor_set(x_103, 1, x_79); -lean_ctor_set(x_103, 2, x_102); -x_104 = lean_array_push(x_56, x_98); -x_105 = lean_array_push(x_104, x_103); -x_106 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_106, 0, x_78); -lean_ctor_set(x_106, 1, x_83); -lean_ctor_set(x_106, 2, x_105); -x_107 = lean_array_push(x_85, x_106); -x_108 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_108, 0, x_78); -lean_ctor_set(x_108, 1, x_87); -lean_ctor_set(x_108, 2, x_107); -x_109 = lean_array_push(x_56, x_108); -x_110 = lean_array_push(x_109, x_90); -x_111 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_111, 0, x_78); -lean_ctor_set(x_111, 1, x_92); -lean_ctor_set(x_111, 2, x_110); -x_112 = lean_array_push(x_56, x_35); -if (lean_obj_tag(x_15) == 0) +x_131 = l_Lean_mkIdentFrom(x_11, x_130); +x_132 = lean_array_push(x_106, x_131); +x_133 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_133, 0, x_108); +lean_ctor_set(x_133, 1, x_59); +lean_ctor_set(x_133, 2, x_132); +x_134 = lean_array_push(x_80, x_128); +x_135 = lean_array_push(x_134, x_133); +lean_inc(x_49); +x_136 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_136, 0, x_108); +lean_ctor_set(x_136, 1, x_49); +lean_ctor_set(x_136, 2, x_135); +x_137 = lean_array_push(x_113, x_136); +lean_inc(x_47); +x_138 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_138, 0, x_108); +lean_ctor_set(x_138, 1, x_47); +lean_ctor_set(x_138, 2, x_137); +x_139 = lean_array_push(x_80, x_138); +x_140 = lean_array_push(x_139, x_117); +lean_inc(x_45); +x_141 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_141, 0, x_108); +lean_ctor_set(x_141, 1, x_45); +lean_ctor_set(x_141, 2, x_140); +x_142 = lean_array_push(x_80, x_41); +if (lean_obj_tag(x_18) == 0) +{ +x_143 = x_63; +goto block_186; +} +else +{ +lean_object* x_187; +lean_dec(x_63); +x_187 = lean_ctor_get(x_18, 0); +lean_inc(x_187); +lean_dec(x_18); +x_143 = x_187; +goto block_186; +} +block_186: +{ +lean_object* x_144; lean_object* x_145; +x_144 = lean_array_push(x_90, x_143); +if (lean_obj_tag(x_79) == 0) { -if (lean_obj_tag(x_55) == 0) +lean_dec(x_38); +if (lean_obj_tag(x_5) == 0) { -lean_object* x_143; -x_143 = l___private_Init_Meta_0__Lean_quoteNameMk(x_25); -x_113 = x_143; -goto block_142; +lean_object* x_173; +x_173 = l_Lean_quoteNameMk(x_27); +x_145 = x_173; +goto block_172; } else { -lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; -lean_dec(x_25); -x_144 = lean_ctor_get(x_55, 0); -lean_inc(x_144); -lean_dec(x_55); -x_145 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__74; -x_146 = l_String_intercalate(x_145, x_144); -x_147 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__75; -x_148 = lean_string_append(x_147, x_146); -lean_dec(x_146); -x_149 = l_Lean_nameLitKind; -x_150 = l_Lean_Syntax_mkLit(x_149, x_148, x_78); -x_151 = lean_array_push(x_85, x_150); -x_152 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__73; -x_153 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_153, 0, x_78); -lean_ctor_set(x_153, 1, x_152); -lean_ctor_set(x_153, 2, x_151); -x_113 = x_153; -goto block_142; -} -} -else -{ -lean_object* x_154; lean_object* x_155; -lean_dec(x_55); -lean_dec(x_25); -x_154 = lean_ctor_get(x_15, 0); -lean_inc(x_154); -lean_dec(x_15); -x_155 = l_Lean_Syntax_getArg(x_154, x_12); -lean_dec(x_154); -x_113 = x_155; -goto block_142; +lean_object* x_174; +lean_dec(x_27); +x_174 = lean_ctor_get(x_5, 0); +lean_inc(x_174); +lean_dec(x_5); +x_145 = x_174; +goto block_172; +} } -block_142: +else { -lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; -x_114 = lean_array_push(x_85, x_113); -x_115 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_115, 0, x_78); -lean_ctor_set(x_115, 1, x_79); -lean_ctor_set(x_115, 2, x_114); -x_116 = lean_array_push(x_57, x_115); -x_117 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_117, 0, x_78); -lean_ctor_set(x_117, 1, x_83); -lean_ctor_set(x_117, 2, x_116); -x_118 = lean_array_push(x_63, x_117); -x_119 = lean_array_push(x_118, x_59); -x_120 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__26; -x_121 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_121, 0, x_78); -lean_ctor_set(x_121, 1, x_120); -lean_ctor_set(x_121, 2, x_119); -x_122 = lean_array_push(x_66, x_121); -x_123 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_123, 0, x_78); -lean_ctor_set(x_123, 1, x_79); -lean_ctor_set(x_123, 2, x_122); -x_124 = lean_array_push(x_67, x_123); -x_125 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_125, 0, x_78); -lean_ctor_set(x_125, 1, x_83); -lean_ctor_set(x_125, 2, x_124); -x_126 = lean_array_push(x_85, x_125); -x_127 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_127, 0, x_78); -lean_ctor_set(x_127, 1, x_87); -lean_ctor_set(x_127, 2, x_126); -x_128 = lean_array_push(x_56, x_127); -x_129 = lean_array_push(x_128, x_90); -x_130 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_130, 0, x_78); -lean_ctor_set(x_130, 1, x_92); -lean_ctor_set(x_130, 2, x_129); -x_131 = lean_array_push(x_64, x_130); -x_132 = lean_array_push(x_131, x_93); -x_133 = lean_array_push(x_132, x_111); -x_134 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_134, 0, x_78); -lean_ctor_set(x_134, 1, x_79); -lean_ctor_set(x_134, 2, x_133); -x_135 = lean_array_push(x_85, x_134); -x_136 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__7; -x_137 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_137, 0, x_78); -lean_ctor_set(x_137, 1, x_136); -lean_ctor_set(x_137, 2, x_135); -x_138 = lean_array_push(x_112, x_137); -x_139 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__5; -x_140 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_140, 0, x_78); -lean_ctor_set(x_140, 1, x_139); -lean_ctor_set(x_140, 2, x_138); -if (lean_is_scalar(x_31)) { - x_141 = lean_alloc_ctor(0, 2, 0); +lean_dec(x_27); +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; +x_175 = lean_ctor_get(x_79, 0); +lean_inc(x_175); +lean_dec(x_79); +x_176 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__50; +x_177 = lean_name_mk_string(x_38, x_176); +x_178 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__51; +x_179 = l_String_intercalate(x_178, x_175); +x_180 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__52; +x_181 = lean_string_append(x_180, x_179); +lean_dec(x_179); +x_182 = l_Lean_Syntax_mkNameLit(x_181, x_108); +x_183 = lean_array_push(x_113, x_182); +x_184 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_184, 0, x_108); +lean_ctor_set(x_184, 1, x_177); +lean_ctor_set(x_184, 2, x_183); +x_145 = x_184; +goto block_172; +} +else +{ +lean_object* x_185; +lean_dec(x_79); +lean_dec(x_38); +x_185 = lean_ctor_get(x_5, 0); +lean_inc(x_185); +lean_dec(x_5); +x_145 = x_185; +goto block_172; +} +} +block_172: +{ +lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; +x_146 = lean_array_push(x_113, x_145); +x_147 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_147, 0, x_108); +lean_ctor_set(x_147, 1, x_59); +lean_ctor_set(x_147, 2, x_146); +x_148 = lean_array_push(x_81, x_147); +lean_inc(x_49); +x_149 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_149, 0, x_108); +lean_ctor_set(x_149, 1, x_49); +lean_ctor_set(x_149, 2, x_148); +x_150 = lean_array_push(x_87, x_149); +x_151 = lean_array_push(x_150, x_83); +x_152 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_152, 0, x_108); +lean_ctor_set(x_152, 1, x_65); +lean_ctor_set(x_152, 2, x_151); +x_153 = lean_array_push(x_144, x_152); +x_154 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_154, 0, x_108); +lean_ctor_set(x_154, 1, x_59); +lean_ctor_set(x_154, 2, x_153); +x_155 = lean_array_push(x_91, x_154); +x_156 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_156, 0, x_108); +lean_ctor_set(x_156, 1, x_49); +lean_ctor_set(x_156, 2, x_155); +x_157 = lean_array_push(x_113, x_156); +x_158 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_158, 0, x_108); +lean_ctor_set(x_158, 1, x_47); +lean_ctor_set(x_158, 2, x_157); +x_159 = lean_array_push(x_80, x_158); +x_160 = lean_array_push(x_159, x_117); +x_161 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_161, 0, x_108); +lean_ctor_set(x_161, 1, x_45); +lean_ctor_set(x_161, 2, x_160); +x_162 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__49; +x_163 = lean_array_push(x_162, x_161); +x_164 = lean_array_push(x_163, x_119); +x_165 = lean_array_push(x_164, x_141); +x_166 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_166, 0, x_108); +lean_ctor_set(x_166, 1, x_59); +lean_ctor_set(x_166, 2, x_165); +x_167 = lean_array_push(x_113, x_166); +x_168 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_168, 0, x_108); +lean_ctor_set(x_168, 1, x_43); +lean_ctor_set(x_168, 2, x_167); +x_169 = lean_array_push(x_142, x_168); +x_170 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_170, 0, x_108); +lean_ctor_set(x_170, 1, x_40); +lean_ctor_set(x_170, 2, x_169); +if (lean_is_scalar(x_34)) { + x_171 = lean_alloc_ctor(0, 2, 0); } else { - x_141 = x_31; + x_171 = x_34; } -lean_ctor_set(x_141, 0, x_140); -lean_ctor_set(x_141, 1, x_30); -return x_141; +lean_ctor_set(x_171, 0, x_170); +lean_ctor_set(x_171, 1, x_33); +return x_171; } } } } +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__2; +lean_inc(x_1); +x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; +lean_dec(x_2); +lean_dec(x_1); +x_6 = lean_box(1); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_3); +return x_7; +} else { -uint8_t x_166; -lean_dec(x_16); +lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_8 = lean_unsigned_to_nat(1u); +x_9 = l_Lean_Syntax_getArg(x_1, x_8); +x_10 = l_Lean_Syntax_isNone(x_9); +if (x_10 == 0) +{ +uint8_t x_11; +lean_inc(x_9); +x_11 = l_Lean_Syntax_matchesNull(x_9, x_8); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_9); +lean_dec(x_2); +lean_dec(x_1); +x_12 = lean_box(1); +x_13 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_3); +return x_13; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_14 = lean_unsigned_to_nat(0u); +x_15 = l_Lean_Syntax_getArg(x_9, x_14); +lean_dec(x_9); +x_16 = l_Lean_Parser_group_formatter___closed__2; +lean_inc(x_15); +x_17 = l_Lean_Syntax_isOfKind(x_15, x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); lean_dec(x_2); -x_166 = !lean_is_exclusive(x_17); -if (x_166 == 0) +lean_dec(x_1); +x_18 = lean_box(1); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_3); +return x_19; +} +else { -return x_17; +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; +x_20 = lean_unsigned_to_nat(3u); +x_21 = l_Lean_Syntax_getArg(x_15, x_20); +lean_dec(x_15); +x_22 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_22, 0, x_21); +x_23 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__4; +x_24 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__2; +x_25 = lean_box(0); +x_26 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2(x_1, x_23, x_24, x_25, x_22, x_2, x_3); +lean_dec(x_1); +return x_26; +} +} } else { -lean_object* x_167; lean_object* x_168; lean_object* x_169; -x_167 = lean_ctor_get(x_17, 0); -x_168 = lean_ctor_get(x_17, 1); -lean_inc(x_168); -lean_inc(x_167); -lean_dec(x_17); -x_169 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_169, 0, x_167); -lean_ctor_set(x_169, 1, x_168); -return x_169; +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +lean_dec(x_9); +x_27 = lean_box(0); +x_28 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__4; +x_29 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__2; +x_30 = lean_box(0); +x_31 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2(x_1, x_28, x_29, x_30, x_27, x_2, x_3); +lean_dec(x_1); +return x_31; +} } } } +LEAN_EXPORT lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___boxed(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; +x_6 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_4); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________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, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_4); +lean_dec(x_1); +return x_8; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__1() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__1() { _start: { lean_object* x_1; @@ -7925,37 +8336,59 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_group), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__2() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__1; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__1; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__3() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__4; -x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__7; +x_2 = l_Lean_Parser_group_formatter___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__4() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__3; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__3; +x_1 = lean_unsigned_to_nat(1u); x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__5() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__6() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__5; +x_2 = 0; +x_3 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__7() { _start: { lean_object* x_1; @@ -7963,17 +8396,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_group_formatter), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__6() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__5; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__7; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__7() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__9() { _start: { lean_object* x_1; @@ -7981,7 +8414,7 @@ x_1 = l_Lean_PrettyPrinter_Formatter_formatterAliasesRef; return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__8() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__10() { _start: { lean_object* x_1; @@ -7989,17 +8422,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_group_parenthesizer), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__9() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__8; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__10; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__10() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__12() { _start: { lean_object* x_1; @@ -8007,7 +8440,7 @@ x_1 = l_Lean_PrettyPrinter_Parenthesizer_parenthesizerAliasesRef; return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__11() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__13() { _start: { lean_object* x_1; @@ -8015,17 +8448,17 @@ x_1 = lean_mk_string_from_bytes("ppHardSpace", 11); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__12() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__11; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__13; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__13() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__15() { _start: { lean_object* x_1; lean_object* x_2; @@ -8035,27 +8468,39 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__14() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__4; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__11; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__13; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__15() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__17() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__14; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__16; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__16() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__18() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = 0; +x_3 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__19() { _start: { lean_object* x_1; @@ -8063,17 +8508,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_ppHardSpace_formatter___boxed), 5, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__17() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__20() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__16; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__19; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__18() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__21() { _start: { lean_object* x_1; @@ -8081,17 +8526,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_ppHardSpace_parenthesizer___boxed return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__19() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__22() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__18; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__21; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__20() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__23() { _start: { lean_object* x_1; @@ -8099,37 +8544,37 @@ x_1 = lean_mk_string_from_bytes("ppSpace", 7); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__21() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__20; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__23; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__22() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__4; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__20; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__23; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__23() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__26() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__22; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__25; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__24() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__27() { _start: { lean_object* x_1; @@ -8137,17 +8582,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_ppSpace_formatter___boxed), 5, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__25() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__28() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__24; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__27; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__26() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__29() { _start: { lean_object* x_1; @@ -8155,17 +8600,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_ppSpace_parenthesizer___boxed), 4 return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__27() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__30() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__26; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__29; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__28() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__31() { _start: { lean_object* x_1; @@ -8173,37 +8618,37 @@ x_1 = lean_mk_string_from_bytes("ppLine", 6); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__29() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__32() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__28; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__31; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__30() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__33() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__4; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__28; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__31; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__31() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__34() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__30; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__33; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__32() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__35() { _start: { lean_object* x_1; @@ -8211,17 +8656,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_ppLine_formatter___boxed), 5, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__33() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__36() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__32; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__35; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__34() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__37() { _start: { lean_object* x_1; @@ -8229,17 +8674,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_ppLine_parenthesizer___boxed), 4, return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__35() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__38() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__34; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__37; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__36() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__39() { _start: { lean_object* x_1; @@ -8247,17 +8692,17 @@ x_1 = lean_mk_string_from_bytes("ppGroup", 7); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__37() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__40() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__36; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__39; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__38() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__41() { _start: { lean_object* x_1; @@ -8265,37 +8710,37 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_ppGroup___boxed), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__39() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__42() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__38; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__41; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__40() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__43() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__4; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__36; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__39; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__41() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__44() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__40; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__43; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__42() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__45() { _start: { lean_object* x_1; @@ -8303,17 +8748,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_ppGroup_formatter), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__43() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__46() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__42; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__45; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__44() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__47() { _start: { lean_object* x_1; @@ -8321,17 +8766,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_ppGroup_parenthesizer), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__45() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__48() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__44; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__47; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__46() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__49() { _start: { lean_object* x_1; @@ -8339,17 +8784,17 @@ x_1 = lean_mk_string_from_bytes("ppRealGroup", 11); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__47() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__50() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__46; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__49; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__48() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__51() { _start: { lean_object* x_1; @@ -8357,37 +8802,37 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_ppRealGroup___boxed), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__49() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__52() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__48; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__51; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__50() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__53() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__4; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__46; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__49; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__51() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__54() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__50; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__53; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__52() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__55() { _start: { lean_object* x_1; @@ -8395,17 +8840,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_ppRealGroup_formatter), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__53() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__56() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__52; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__55; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__54() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__57() { _start: { lean_object* x_1; @@ -8413,17 +8858,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_ppRealGroup_parenthesizer), 6, 0) return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__55() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__58() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__54; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__57; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__56() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__59() { _start: { lean_object* x_1; @@ -8431,17 +8876,17 @@ x_1 = lean_mk_string_from_bytes("ppRealFill", 10); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__57() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__60() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__56; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__59; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__58() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__61() { _start: { lean_object* x_1; @@ -8449,37 +8894,37 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_ppRealFill___boxed), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__59() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__62() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__58; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__61; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__60() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__63() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__4; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__56; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__59; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__61() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__64() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__60; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__63; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__62() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__65() { _start: { lean_object* x_1; @@ -8487,17 +8932,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_ppRealFill_formatter), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__63() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__66() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__62; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__65; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__64() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__67() { _start: { lean_object* x_1; @@ -8505,17 +8950,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_ppRealFill_parenthesizer), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__65() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__68() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__64; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__67; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__66() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__69() { _start: { lean_object* x_1; @@ -8523,17 +8968,17 @@ x_1 = lean_mk_string_from_bytes("ppIndent", 8); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__67() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__70() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__66; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__69; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__68() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__71() { _start: { lean_object* x_1; @@ -8541,37 +8986,37 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_ppIndent___boxed), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__69() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__72() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__68; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__71; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__70() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__73() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__4; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__66; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__69; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__71() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__74() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__70; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__73; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__72() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__75() { _start: { lean_object* x_1; @@ -8579,17 +9024,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_ppIndent_formatter), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__73() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__76() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__72; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__75; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__74() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__77() { _start: { lean_object* x_1; @@ -8597,17 +9042,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_ppIndent_parenthesizer), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__75() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__78() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__74; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__77; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__76() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__79() { _start: { lean_object* x_1; @@ -8615,17 +9060,17 @@ x_1 = lean_mk_string_from_bytes("ppDedent", 8); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__77() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__80() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__76; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__79; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__78() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__81() { _start: { lean_object* x_1; @@ -8633,37 +9078,37 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_ppDedent___boxed), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__79() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__82() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__78; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__81; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__80() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__83() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__4; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__76; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__79; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__81() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__84() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__80; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__83; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__82() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__85() { _start: { lean_object* x_1; @@ -8671,17 +9116,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_ppDedent_formatter), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__83() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__86() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__82; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__85; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__84() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__87() { _start: { lean_object* x_1; @@ -8689,17 +9134,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_ppDedent_parenthesizer), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__85() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__88() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__84; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__87; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__86() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__89() { _start: { lean_object* x_1; @@ -8707,37 +9152,37 @@ x_1 = lean_mk_string_from_bytes("ppAllowUngrouped", 16); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__87() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__90() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__86; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__89; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__88() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__91() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__4; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__86; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__89; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__89() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__92() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__88; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__91; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__90() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__93() { _start: { lean_object* x_1; @@ -8745,17 +9190,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_ppAllowUngrouped_formatter___boxed), 1, return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__91() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__94() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__90; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__93; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__92() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__95() { _start: { lean_object* x_1; @@ -8763,17 +9208,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_ppAllowUngrouped_parenthesizer___ return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__93() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__96() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__92; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__95; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__94() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__97() { _start: { lean_object* x_1; @@ -8781,17 +9226,17 @@ x_1 = lean_mk_string_from_bytes("ppDedentIfGrouped", 17); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__95() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__98() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__94; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__97; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__96() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__99() { _start: { lean_object* x_1; @@ -8799,37 +9244,37 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_ppDedentIfGrouped___boxed), 1, 0) return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__97() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__100() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__96; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__99; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__98() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__101() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__4; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__94; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__97; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__99() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__102() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__98; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__101; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__100() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__103() { _start: { lean_object* x_1; @@ -8837,17 +9282,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_ppDedentIfGrouped_formatter), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__101() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__104() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__100; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__103; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__102() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__105() { _start: { lean_object* x_1; @@ -8855,17 +9300,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_ppDedentIfGrouped_parenthesizer), return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__103() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__106() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__102; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__105; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__104() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__107() { _start: { lean_object* x_1; @@ -8873,37 +9318,37 @@ x_1 = lean_mk_string_from_bytes("ppHardLineUnlessUngrouped", 25); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__105() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__108() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__104; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__107; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__106() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__109() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__4; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__104; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__107; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__107() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__110() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__106; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__109; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__108() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__111() { _start: { lean_object* x_1; @@ -8911,17 +9356,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_ppHardLineUnlessUngrouped_formatter___bo return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__109() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__112() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__108; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__111; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__110() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__113() { _start: { lean_object* x_1; @@ -8929,1128 +9374,1130 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_ppHardLineUnlessUngrouped_parenth return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__111() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__114() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__110; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__113; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610_(lean_object* x_1) { _start: { -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__8; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__2; -x_4 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__4; -x_5 = l_Lean_Parser_registerAlias(x_2, x_3, x_4, x_1); -if (lean_obj_tag(x_5) == 0) -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_6 = lean_ctor_get(x_5, 1); -lean_inc(x_6); -lean_dec(x_5); -x_7 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__7; -x_8 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__6; -x_9 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_2, x_8, x_6); -if (lean_obj_tag(x_9) == 0) -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_10 = lean_ctor_get(x_9, 1); -lean_inc(x_10); -lean_dec(x_9); -x_11 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__10; -x_12 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__9; -x_13 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_2, x_12, x_10); -if (lean_obj_tag(x_13) == 0) -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -lean_dec(x_13); -x_15 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__12; -x_16 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__13; -x_17 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__15; -x_18 = l_Lean_Parser_registerAlias(x_15, x_16, x_17, x_14); -if (lean_obj_tag(x_18) == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -lean_dec(x_18); -x_20 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__17; -x_21 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_15, x_20, x_19); -if (lean_obj_tag(x_21) == 0) -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_21, 1); -lean_inc(x_22); -lean_dec(x_21); -x_23 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__19; -x_24 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_15, x_23, x_22); -if (lean_obj_tag(x_24) == 0) -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_25 = lean_ctor_get(x_24, 1); -lean_inc(x_25); -lean_dec(x_24); -x_26 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__21; -x_27 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__23; -x_28 = l_Lean_Parser_registerAlias(x_26, x_16, x_27, x_25); -if (lean_obj_tag(x_28) == 0) -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_28, 1); -lean_inc(x_29); -lean_dec(x_28); -x_30 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__25; -x_31 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_26, x_30, x_29); -if (lean_obj_tag(x_31) == 0) -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_31, 1); -lean_inc(x_32); -lean_dec(x_31); -x_33 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__27; -x_34 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_26, x_33, x_32); -if (lean_obj_tag(x_34) == 0) +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_2 = l_Lean_Parser_group_formatter___closed__2; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__2; +x_4 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__4; +x_5 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__6; +x_6 = l_Lean_Parser_registerAlias(x_2, x_3, x_4, x_5, x_1); +if (lean_obj_tag(x_6) == 0) { -lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_35 = lean_ctor_get(x_34, 1); -lean_inc(x_35); -lean_dec(x_34); -x_36 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__29; -x_37 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__31; -x_38 = l_Lean_Parser_registerAlias(x_36, x_16, x_37, x_35); -if (lean_obj_tag(x_38) == 0) -{ -lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_39 = lean_ctor_get(x_38, 1); -lean_inc(x_39); -lean_dec(x_38); -x_40 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__33; -x_41 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_36, x_40, x_39); -if (lean_obj_tag(x_41) == 0) +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_7 = lean_ctor_get(x_6, 1); +lean_inc(x_7); +lean_dec(x_6); +x_8 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__9; +x_9 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__8; +x_10 = l_Lean_Parser_registerAliasCore___rarg(x_8, x_2, x_9, x_7); +if (lean_obj_tag(x_10) == 0) { -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_41, 1); -lean_inc(x_42); -lean_dec(x_41); -x_43 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__35; -x_44 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_36, x_43, x_42); -if (lean_obj_tag(x_44) == 0) +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +lean_dec(x_10); +x_12 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__12; +x_13 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__11; +x_14 = l_Lean_Parser_registerAliasCore___rarg(x_12, x_2, x_13, x_11); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_45 = lean_ctor_get(x_44, 1); -lean_inc(x_45); -lean_dec(x_44); -x_46 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__37; -x_47 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__39; -x_48 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__41; -x_49 = l_Lean_Parser_registerAlias(x_46, x_47, x_48, x_45); -if (lean_obj_tag(x_49) == 0) -{ -lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_50 = lean_ctor_get(x_49, 1); -lean_inc(x_50); -lean_dec(x_49); -x_51 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__43; -x_52 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_46, x_51, x_50); -if (lean_obj_tag(x_52) == 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; +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +lean_dec(x_14); +x_16 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__14; +x_17 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__15; +x_18 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__17; +x_19 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__18; +x_20 = l_Lean_Parser_registerAlias(x_16, x_17, x_18, x_19, x_15); +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_20, 1); +lean_inc(x_21); +lean_dec(x_20); +x_22 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__20; +x_23 = l_Lean_Parser_registerAliasCore___rarg(x_8, x_16, x_22, x_21); +if (lean_obj_tag(x_23) == 0) { -lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_53 = lean_ctor_get(x_52, 1); -lean_inc(x_53); -lean_dec(x_52); -x_54 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__45; -x_55 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_46, x_54, x_53); -if (lean_obj_tag(x_55) == 0) -{ -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_56 = lean_ctor_get(x_55, 1); -lean_inc(x_56); -lean_dec(x_55); -x_57 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__47; -x_58 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__49; -x_59 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__51; -x_60 = l_Lean_Parser_registerAlias(x_57, x_58, x_59, x_56); -if (lean_obj_tag(x_60) == 0) -{ -lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_61 = lean_ctor_get(x_60, 1); -lean_inc(x_61); -lean_dec(x_60); -x_62 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__53; -x_63 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_57, x_62, x_61); -if (lean_obj_tag(x_63) == 0) -{ -lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_ctor_get(x_63, 1); -lean_inc(x_64); -lean_dec(x_63); -x_65 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__55; -x_66 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_57, x_65, x_64); -if (lean_obj_tag(x_66) == 0) +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); +x_25 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__22; +x_26 = l_Lean_Parser_registerAliasCore___rarg(x_12, x_16, x_25, x_24); +if (lean_obj_tag(x_26) == 0) { -lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_67 = lean_ctor_get(x_66, 1); -lean_inc(x_67); -lean_dec(x_66); -x_68 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__57; -x_69 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__59; -x_70 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__61; -x_71 = l_Lean_Parser_registerAlias(x_68, x_69, x_70, x_67); -if (lean_obj_tag(x_71) == 0) -{ -lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_72 = lean_ctor_get(x_71, 1); -lean_inc(x_72); -lean_dec(x_71); -x_73 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__63; -x_74 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_68, x_73, x_72); -if (lean_obj_tag(x_74) == 0) -{ -lean_object* x_75; lean_object* x_76; lean_object* x_77; -x_75 = lean_ctor_get(x_74, 1); -lean_inc(x_75); -lean_dec(x_74); -x_76 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__65; -x_77 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_68, x_76, x_75); -if (lean_obj_tag(x_77) == 0) -{ -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; -x_78 = lean_ctor_get(x_77, 1); -lean_inc(x_78); -lean_dec(x_77); -x_79 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__67; -x_80 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__69; -x_81 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__71; -x_82 = l_Lean_Parser_registerAlias(x_79, x_80, x_81, x_78); -if (lean_obj_tag(x_82) == 0) -{ -lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_83 = lean_ctor_get(x_82, 1); -lean_inc(x_83); -lean_dec(x_82); -x_84 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__73; -x_85 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_79, x_84, x_83); -if (lean_obj_tag(x_85) == 0) -{ -lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_86 = lean_ctor_get(x_85, 1); -lean_inc(x_86); -lean_dec(x_85); -x_87 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__75; -x_88 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_79, x_87, x_86); -if (lean_obj_tag(x_88) == 0) -{ -lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_89 = lean_ctor_get(x_88, 1); -lean_inc(x_89); -lean_dec(x_88); -x_90 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__77; -x_91 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__79; -x_92 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__81; -x_93 = l_Lean_Parser_registerAlias(x_90, x_91, x_92, x_89); -if (lean_obj_tag(x_93) == 0) -{ -lean_object* x_94; lean_object* x_95; lean_object* x_96; -x_94 = lean_ctor_get(x_93, 1); -lean_inc(x_94); -lean_dec(x_93); -x_95 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__83; -x_96 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_90, x_95, x_94); -if (lean_obj_tag(x_96) == 0) -{ -lean_object* x_97; lean_object* x_98; lean_object* x_99; -x_97 = lean_ctor_get(x_96, 1); -lean_inc(x_97); -lean_dec(x_96); -x_98 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__85; -x_99 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_90, x_98, x_97); -if (lean_obj_tag(x_99) == 0) -{ -lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; -x_100 = lean_ctor_get(x_99, 1); -lean_inc(x_100); -lean_dec(x_99); -x_101 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__87; -x_102 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__89; -x_103 = l_Lean_Parser_registerAlias(x_101, x_16, x_102, x_100); -if (lean_obj_tag(x_103) == 0) -{ -lean_object* x_104; lean_object* x_105; lean_object* x_106; -x_104 = lean_ctor_get(x_103, 1); -lean_inc(x_104); -lean_dec(x_103); -x_105 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__91; -x_106 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_101, x_105, x_104); -if (lean_obj_tag(x_106) == 0) -{ -lean_object* x_107; lean_object* x_108; lean_object* x_109; -x_107 = lean_ctor_get(x_106, 1); -lean_inc(x_107); -lean_dec(x_106); -x_108 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__93; -x_109 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_101, x_108, x_107); -if (lean_obj_tag(x_109) == 0) -{ -lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; -x_110 = lean_ctor_get(x_109, 1); -lean_inc(x_110); -lean_dec(x_109); -x_111 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__95; -x_112 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__97; -x_113 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__99; -x_114 = l_Lean_Parser_registerAlias(x_111, x_112, x_113, x_110); -if (lean_obj_tag(x_114) == 0) -{ -lean_object* x_115; lean_object* x_116; lean_object* x_117; -x_115 = lean_ctor_get(x_114, 1); -lean_inc(x_115); -lean_dec(x_114); -x_116 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__101; -x_117 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_111, x_116, x_115); -if (lean_obj_tag(x_117) == 0) -{ -lean_object* x_118; lean_object* x_119; lean_object* x_120; -x_118 = lean_ctor_get(x_117, 1); -lean_inc(x_118); -lean_dec(x_117); -x_119 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__103; -x_120 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_111, x_119, x_118); -if (lean_obj_tag(x_120) == 0) -{ -lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; -x_121 = lean_ctor_get(x_120, 1); -lean_inc(x_121); -lean_dec(x_120); -x_122 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__105; -x_123 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__107; -x_124 = l_Lean_Parser_registerAlias(x_122, x_16, x_123, x_121); -if (lean_obj_tag(x_124) == 0) -{ -lean_object* x_125; lean_object* x_126; lean_object* x_127; -x_125 = lean_ctor_get(x_124, 1); -lean_inc(x_125); -lean_dec(x_124); -x_126 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__109; -x_127 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_122, x_126, x_125); -if (lean_obj_tag(x_127) == 0) -{ -lean_object* x_128; lean_object* x_129; lean_object* x_130; -x_128 = lean_ctor_get(x_127, 1); -lean_inc(x_128); -lean_dec(x_127); -x_129 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__111; -x_130 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_122, x_129, x_128); -return x_130; -} -else +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); +lean_dec(x_26); +x_28 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__24; +x_29 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__26; +x_30 = l_Lean_Parser_registerAlias(x_28, x_17, x_29, x_19, x_27); +if (lean_obj_tag(x_30) == 0) { -uint8_t x_131; -x_131 = !lean_is_exclusive(x_127); -if (x_131 == 0) +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_32 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__28; +x_33 = l_Lean_Parser_registerAliasCore___rarg(x_8, x_28, x_32, x_31); +if (lean_obj_tag(x_33) == 0) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_33, 1); +lean_inc(x_34); +lean_dec(x_33); +x_35 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__30; +x_36 = l_Lean_Parser_registerAliasCore___rarg(x_12, x_28, x_35, x_34); +if (lean_obj_tag(x_36) == 0) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_37 = lean_ctor_get(x_36, 1); +lean_inc(x_37); +lean_dec(x_36); +x_38 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__32; +x_39 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__34; +x_40 = l_Lean_Parser_registerAlias(x_38, x_17, x_39, x_19, x_37); +if (lean_obj_tag(x_40) == 0) +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_40, 1); +lean_inc(x_41); +lean_dec(x_40); +x_42 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__36; +x_43 = l_Lean_Parser_registerAliasCore___rarg(x_8, x_38, x_42, x_41); +if (lean_obj_tag(x_43) == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_43, 1); +lean_inc(x_44); +lean_dec(x_43); +x_45 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__38; +x_46 = l_Lean_Parser_registerAliasCore___rarg(x_12, x_38, x_45, x_44); +if (lean_obj_tag(x_46) == 0) { -return x_127; -} -else +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_47 = lean_ctor_get(x_46, 1); +lean_inc(x_47); +lean_dec(x_46); +x_48 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__40; +x_49 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__42; +x_50 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__44; +x_51 = l_Lean_Parser_registerAlias(x_48, x_49, x_50, x_19, x_47); +if (lean_obj_tag(x_51) == 0) +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_51, 1); +lean_inc(x_52); +lean_dec(x_51); +x_53 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__46; +x_54 = l_Lean_Parser_registerAliasCore___rarg(x_8, x_48, x_53, x_52); +if (lean_obj_tag(x_54) == 0) +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_54, 1); +lean_inc(x_55); +lean_dec(x_54); +x_56 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__48; +x_57 = l_Lean_Parser_registerAliasCore___rarg(x_12, x_48, x_56, x_55); +if (lean_obj_tag(x_57) == 0) +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_58 = lean_ctor_get(x_57, 1); +lean_inc(x_58); +lean_dec(x_57); +x_59 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__50; +x_60 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__52; +x_61 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__54; +x_62 = l_Lean_Parser_registerAlias(x_59, x_60, x_61, x_19, x_58); +if (lean_obj_tag(x_62) == 0) +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_62, 1); +lean_inc(x_63); +lean_dec(x_62); +x_64 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__56; +x_65 = l_Lean_Parser_registerAliasCore___rarg(x_8, x_59, x_64, x_63); +if (lean_obj_tag(x_65) == 0) { -lean_object* x_132; lean_object* x_133; lean_object* x_134; -x_132 = lean_ctor_get(x_127, 0); -x_133 = lean_ctor_get(x_127, 1); -lean_inc(x_133); -lean_inc(x_132); -lean_dec(x_127); -x_134 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_134, 0, x_132); -lean_ctor_set(x_134, 1, x_133); -return x_134; -} -} -} -else +lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_66 = lean_ctor_get(x_65, 1); +lean_inc(x_66); +lean_dec(x_65); +x_67 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__58; +x_68 = l_Lean_Parser_registerAliasCore___rarg(x_12, x_59, x_67, x_66); +if (lean_obj_tag(x_68) == 0) { -uint8_t x_135; -x_135 = !lean_is_exclusive(x_124); -if (x_135 == 0) +lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_69 = lean_ctor_get(x_68, 1); +lean_inc(x_69); +lean_dec(x_68); +x_70 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__60; +x_71 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__62; +x_72 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__64; +x_73 = l_Lean_Parser_registerAlias(x_70, x_71, x_72, x_19, x_69); +if (lean_obj_tag(x_73) == 0) +{ +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_73, 1); +lean_inc(x_74); +lean_dec(x_73); +x_75 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__66; +x_76 = l_Lean_Parser_registerAliasCore___rarg(x_8, x_70, x_75, x_74); +if (lean_obj_tag(x_76) == 0) { -return x_124; -} -else +lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_77 = lean_ctor_get(x_76, 1); +lean_inc(x_77); +lean_dec(x_76); +x_78 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__68; +x_79 = l_Lean_Parser_registerAliasCore___rarg(x_12, x_70, x_78, x_77); +if (lean_obj_tag(x_79) == 0) { -lean_object* x_136; lean_object* x_137; lean_object* x_138; -x_136 = lean_ctor_get(x_124, 0); -x_137 = lean_ctor_get(x_124, 1); -lean_inc(x_137); -lean_inc(x_136); -lean_dec(x_124); -x_138 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_138, 0, x_136); -lean_ctor_set(x_138, 1, x_137); -return x_138; +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_80 = lean_ctor_get(x_79, 1); +lean_inc(x_80); +lean_dec(x_79); +x_81 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__70; +x_82 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__72; +x_83 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__74; +x_84 = l_Lean_Parser_registerAlias(x_81, x_82, x_83, x_19, x_80); +if (lean_obj_tag(x_84) == 0) +{ +lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_85 = lean_ctor_get(x_84, 1); +lean_inc(x_85); +lean_dec(x_84); +x_86 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__76; +x_87 = l_Lean_Parser_registerAliasCore___rarg(x_8, x_81, x_86, x_85); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_88 = lean_ctor_get(x_87, 1); +lean_inc(x_88); +lean_dec(x_87); +x_89 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__78; +x_90 = l_Lean_Parser_registerAliasCore___rarg(x_12, x_81, x_89, x_88); +if (lean_obj_tag(x_90) == 0) +{ +lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_91 = lean_ctor_get(x_90, 1); +lean_inc(x_91); +lean_dec(x_90); +x_92 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__80; +x_93 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__82; +x_94 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__84; +x_95 = l_Lean_Parser_registerAlias(x_92, x_93, x_94, x_19, x_91); +if (lean_obj_tag(x_95) == 0) +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_95, 1); +lean_inc(x_96); +lean_dec(x_95); +x_97 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__86; +x_98 = l_Lean_Parser_registerAliasCore___rarg(x_8, x_92, x_97, x_96); +if (lean_obj_tag(x_98) == 0) +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_99 = lean_ctor_get(x_98, 1); +lean_inc(x_99); +lean_dec(x_98); +x_100 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__88; +x_101 = l_Lean_Parser_registerAliasCore___rarg(x_12, x_92, x_100, x_99); +if (lean_obj_tag(x_101) == 0) +{ +lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; +x_102 = lean_ctor_get(x_101, 1); +lean_inc(x_102); +lean_dec(x_101); +x_103 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__90; +x_104 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__92; +x_105 = l_Lean_Parser_registerAlias(x_103, x_17, x_104, x_19, x_102); +if (lean_obj_tag(x_105) == 0) +{ +lean_object* x_106; lean_object* x_107; lean_object* x_108; +x_106 = lean_ctor_get(x_105, 1); +lean_inc(x_106); +lean_dec(x_105); +x_107 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__94; +x_108 = l_Lean_Parser_registerAliasCore___rarg(x_8, x_103, x_107, x_106); +if (lean_obj_tag(x_108) == 0) +{ +lean_object* x_109; lean_object* x_110; lean_object* x_111; +x_109 = lean_ctor_get(x_108, 1); +lean_inc(x_109); +lean_dec(x_108); +x_110 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__96; +x_111 = l_Lean_Parser_registerAliasCore___rarg(x_12, x_103, x_110, x_109); +if (lean_obj_tag(x_111) == 0) +{ +lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_112 = lean_ctor_get(x_111, 1); +lean_inc(x_112); +lean_dec(x_111); +x_113 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__98; +x_114 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__100; +x_115 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__102; +x_116 = l_Lean_Parser_registerAlias(x_113, x_114, x_115, x_19, x_112); +if (lean_obj_tag(x_116) == 0) +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; +x_117 = lean_ctor_get(x_116, 1); +lean_inc(x_117); +lean_dec(x_116); +x_118 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__104; +x_119 = l_Lean_Parser_registerAliasCore___rarg(x_8, x_113, x_118, x_117); +if (lean_obj_tag(x_119) == 0) +{ +lean_object* x_120; lean_object* x_121; lean_object* x_122; +x_120 = lean_ctor_get(x_119, 1); +lean_inc(x_120); +lean_dec(x_119); +x_121 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__106; +x_122 = l_Lean_Parser_registerAliasCore___rarg(x_12, x_113, x_121, x_120); +if (lean_obj_tag(x_122) == 0) +{ +lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; +x_123 = lean_ctor_get(x_122, 1); +lean_inc(x_123); +lean_dec(x_122); +x_124 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__108; +x_125 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__110; +x_126 = l_Lean_Parser_registerAlias(x_124, x_17, x_125, x_19, x_123); +if (lean_obj_tag(x_126) == 0) +{ +lean_object* x_127; lean_object* x_128; lean_object* x_129; +x_127 = lean_ctor_get(x_126, 1); +lean_inc(x_127); +lean_dec(x_126); +x_128 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__112; +x_129 = l_Lean_Parser_registerAliasCore___rarg(x_8, x_124, x_128, x_127); +if (lean_obj_tag(x_129) == 0) +{ +lean_object* x_130; lean_object* x_131; lean_object* x_132; +x_130 = lean_ctor_get(x_129, 1); +lean_inc(x_130); +lean_dec(x_129); +x_131 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__114; +x_132 = l_Lean_Parser_registerAliasCore___rarg(x_12, x_124, x_131, x_130); +return x_132; +} +else +{ +uint8_t x_133; +x_133 = !lean_is_exclusive(x_129); +if (x_133 == 0) +{ +return x_129; +} +else +{ +lean_object* x_134; lean_object* x_135; lean_object* x_136; +x_134 = lean_ctor_get(x_129, 0); +x_135 = lean_ctor_get(x_129, 1); +lean_inc(x_135); +lean_inc(x_134); +lean_dec(x_129); +x_136 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_136, 0, x_134); +lean_ctor_set(x_136, 1, x_135); +return x_136; +} +} +} +else +{ +uint8_t x_137; +x_137 = !lean_is_exclusive(x_126); +if (x_137 == 0) +{ +return x_126; +} +else +{ +lean_object* x_138; lean_object* x_139; lean_object* x_140; +x_138 = lean_ctor_get(x_126, 0); +x_139 = lean_ctor_get(x_126, 1); +lean_inc(x_139); +lean_inc(x_138); +lean_dec(x_126); +x_140 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_140, 0, x_138); +lean_ctor_set(x_140, 1, x_139); +return x_140; } } } else { -uint8_t x_139; -x_139 = !lean_is_exclusive(x_120); -if (x_139 == 0) +uint8_t x_141; +x_141 = !lean_is_exclusive(x_122); +if (x_141 == 0) { -return x_120; +return x_122; } else { -lean_object* x_140; lean_object* x_141; lean_object* x_142; -x_140 = lean_ctor_get(x_120, 0); -x_141 = lean_ctor_get(x_120, 1); -lean_inc(x_141); -lean_inc(x_140); -lean_dec(x_120); -x_142 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_142, 0, x_140); -lean_ctor_set(x_142, 1, x_141); -return x_142; +lean_object* x_142; lean_object* x_143; lean_object* x_144; +x_142 = lean_ctor_get(x_122, 0); +x_143 = lean_ctor_get(x_122, 1); +lean_inc(x_143); +lean_inc(x_142); +lean_dec(x_122); +x_144 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_144, 0, x_142); +lean_ctor_set(x_144, 1, x_143); +return x_144; } } } else { -uint8_t x_143; -x_143 = !lean_is_exclusive(x_117); -if (x_143 == 0) +uint8_t x_145; +x_145 = !lean_is_exclusive(x_119); +if (x_145 == 0) { -return x_117; +return x_119; } else { -lean_object* x_144; lean_object* x_145; lean_object* x_146; -x_144 = lean_ctor_get(x_117, 0); -x_145 = lean_ctor_get(x_117, 1); -lean_inc(x_145); -lean_inc(x_144); -lean_dec(x_117); -x_146 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_146, 0, x_144); -lean_ctor_set(x_146, 1, x_145); -return x_146; +lean_object* x_146; lean_object* x_147; lean_object* x_148; +x_146 = lean_ctor_get(x_119, 0); +x_147 = lean_ctor_get(x_119, 1); +lean_inc(x_147); +lean_inc(x_146); +lean_dec(x_119); +x_148 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_148, 0, x_146); +lean_ctor_set(x_148, 1, x_147); +return x_148; } } } else { -uint8_t x_147; -x_147 = !lean_is_exclusive(x_114); -if (x_147 == 0) +uint8_t x_149; +x_149 = !lean_is_exclusive(x_116); +if (x_149 == 0) { -return x_114; +return x_116; } else { -lean_object* x_148; lean_object* x_149; lean_object* x_150; -x_148 = lean_ctor_get(x_114, 0); -x_149 = lean_ctor_get(x_114, 1); -lean_inc(x_149); -lean_inc(x_148); -lean_dec(x_114); -x_150 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_150, 0, x_148); -lean_ctor_set(x_150, 1, x_149); -return x_150; +lean_object* x_150; lean_object* x_151; lean_object* x_152; +x_150 = lean_ctor_get(x_116, 0); +x_151 = lean_ctor_get(x_116, 1); +lean_inc(x_151); +lean_inc(x_150); +lean_dec(x_116); +x_152 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_152, 0, x_150); +lean_ctor_set(x_152, 1, x_151); +return x_152; } } } else { -uint8_t x_151; -x_151 = !lean_is_exclusive(x_109); -if (x_151 == 0) +uint8_t x_153; +x_153 = !lean_is_exclusive(x_111); +if (x_153 == 0) { -return x_109; +return x_111; } else { -lean_object* x_152; lean_object* x_153; lean_object* x_154; -x_152 = lean_ctor_get(x_109, 0); -x_153 = lean_ctor_get(x_109, 1); -lean_inc(x_153); -lean_inc(x_152); -lean_dec(x_109); -x_154 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_154, 0, x_152); -lean_ctor_set(x_154, 1, x_153); -return x_154; +lean_object* x_154; lean_object* x_155; lean_object* x_156; +x_154 = lean_ctor_get(x_111, 0); +x_155 = lean_ctor_get(x_111, 1); +lean_inc(x_155); +lean_inc(x_154); +lean_dec(x_111); +x_156 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_156, 0, x_154); +lean_ctor_set(x_156, 1, x_155); +return x_156; } } } else { -uint8_t x_155; -x_155 = !lean_is_exclusive(x_106); -if (x_155 == 0) +uint8_t x_157; +x_157 = !lean_is_exclusive(x_108); +if (x_157 == 0) { -return x_106; +return x_108; } else { -lean_object* x_156; lean_object* x_157; lean_object* x_158; -x_156 = lean_ctor_get(x_106, 0); -x_157 = lean_ctor_get(x_106, 1); -lean_inc(x_157); -lean_inc(x_156); -lean_dec(x_106); -x_158 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_158, 0, x_156); -lean_ctor_set(x_158, 1, x_157); -return x_158; +lean_object* x_158; lean_object* x_159; lean_object* x_160; +x_158 = lean_ctor_get(x_108, 0); +x_159 = lean_ctor_get(x_108, 1); +lean_inc(x_159); +lean_inc(x_158); +lean_dec(x_108); +x_160 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_160, 0, x_158); +lean_ctor_set(x_160, 1, x_159); +return x_160; } } } else { -uint8_t x_159; -x_159 = !lean_is_exclusive(x_103); -if (x_159 == 0) +uint8_t x_161; +x_161 = !lean_is_exclusive(x_105); +if (x_161 == 0) { -return x_103; +return x_105; } else { -lean_object* x_160; lean_object* x_161; lean_object* x_162; -x_160 = lean_ctor_get(x_103, 0); -x_161 = lean_ctor_get(x_103, 1); -lean_inc(x_161); -lean_inc(x_160); -lean_dec(x_103); -x_162 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_162, 0, x_160); -lean_ctor_set(x_162, 1, x_161); -return x_162; +lean_object* x_162; lean_object* x_163; lean_object* x_164; +x_162 = lean_ctor_get(x_105, 0); +x_163 = lean_ctor_get(x_105, 1); +lean_inc(x_163); +lean_inc(x_162); +lean_dec(x_105); +x_164 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_164, 0, x_162); +lean_ctor_set(x_164, 1, x_163); +return x_164; } } } else { -uint8_t x_163; -x_163 = !lean_is_exclusive(x_99); -if (x_163 == 0) +uint8_t x_165; +x_165 = !lean_is_exclusive(x_101); +if (x_165 == 0) { -return x_99; +return x_101; } else { -lean_object* x_164; lean_object* x_165; lean_object* x_166; -x_164 = lean_ctor_get(x_99, 0); -x_165 = lean_ctor_get(x_99, 1); -lean_inc(x_165); -lean_inc(x_164); -lean_dec(x_99); -x_166 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_166, 0, x_164); -lean_ctor_set(x_166, 1, x_165); -return x_166; +lean_object* x_166; lean_object* x_167; lean_object* x_168; +x_166 = lean_ctor_get(x_101, 0); +x_167 = lean_ctor_get(x_101, 1); +lean_inc(x_167); +lean_inc(x_166); +lean_dec(x_101); +x_168 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_168, 0, x_166); +lean_ctor_set(x_168, 1, x_167); +return x_168; } } } else { -uint8_t x_167; -x_167 = !lean_is_exclusive(x_96); -if (x_167 == 0) +uint8_t x_169; +x_169 = !lean_is_exclusive(x_98); +if (x_169 == 0) { -return x_96; +return x_98; } else { -lean_object* x_168; lean_object* x_169; lean_object* x_170; -x_168 = lean_ctor_get(x_96, 0); -x_169 = lean_ctor_get(x_96, 1); -lean_inc(x_169); -lean_inc(x_168); -lean_dec(x_96); -x_170 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_170, 0, x_168); -lean_ctor_set(x_170, 1, x_169); -return x_170; +lean_object* x_170; lean_object* x_171; lean_object* x_172; +x_170 = lean_ctor_get(x_98, 0); +x_171 = lean_ctor_get(x_98, 1); +lean_inc(x_171); +lean_inc(x_170); +lean_dec(x_98); +x_172 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_172, 0, x_170); +lean_ctor_set(x_172, 1, x_171); +return x_172; } } } else { -uint8_t x_171; -x_171 = !lean_is_exclusive(x_93); -if (x_171 == 0) +uint8_t x_173; +x_173 = !lean_is_exclusive(x_95); +if (x_173 == 0) { -return x_93; +return x_95; } else { -lean_object* x_172; lean_object* x_173; lean_object* x_174; -x_172 = lean_ctor_get(x_93, 0); -x_173 = lean_ctor_get(x_93, 1); -lean_inc(x_173); -lean_inc(x_172); -lean_dec(x_93); -x_174 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_174, 0, x_172); -lean_ctor_set(x_174, 1, x_173); -return x_174; +lean_object* x_174; lean_object* x_175; lean_object* x_176; +x_174 = lean_ctor_get(x_95, 0); +x_175 = lean_ctor_get(x_95, 1); +lean_inc(x_175); +lean_inc(x_174); +lean_dec(x_95); +x_176 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_176, 0, x_174); +lean_ctor_set(x_176, 1, x_175); +return x_176; } } } else { -uint8_t x_175; -x_175 = !lean_is_exclusive(x_88); -if (x_175 == 0) +uint8_t x_177; +x_177 = !lean_is_exclusive(x_90); +if (x_177 == 0) { -return x_88; +return x_90; } else { -lean_object* x_176; lean_object* x_177; lean_object* x_178; -x_176 = lean_ctor_get(x_88, 0); -x_177 = lean_ctor_get(x_88, 1); -lean_inc(x_177); -lean_inc(x_176); -lean_dec(x_88); -x_178 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_178, 0, x_176); -lean_ctor_set(x_178, 1, x_177); -return x_178; +lean_object* x_178; lean_object* x_179; lean_object* x_180; +x_178 = lean_ctor_get(x_90, 0); +x_179 = lean_ctor_get(x_90, 1); +lean_inc(x_179); +lean_inc(x_178); +lean_dec(x_90); +x_180 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_180, 0, x_178); +lean_ctor_set(x_180, 1, x_179); +return x_180; } } } else { -uint8_t x_179; -x_179 = !lean_is_exclusive(x_85); -if (x_179 == 0) +uint8_t x_181; +x_181 = !lean_is_exclusive(x_87); +if (x_181 == 0) { -return x_85; +return x_87; } else { -lean_object* x_180; lean_object* x_181; lean_object* x_182; -x_180 = lean_ctor_get(x_85, 0); -x_181 = lean_ctor_get(x_85, 1); -lean_inc(x_181); -lean_inc(x_180); -lean_dec(x_85); -x_182 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_182, 0, x_180); -lean_ctor_set(x_182, 1, x_181); -return x_182; +lean_object* x_182; lean_object* x_183; lean_object* x_184; +x_182 = lean_ctor_get(x_87, 0); +x_183 = lean_ctor_get(x_87, 1); +lean_inc(x_183); +lean_inc(x_182); +lean_dec(x_87); +x_184 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_184, 0, x_182); +lean_ctor_set(x_184, 1, x_183); +return x_184; } } } else { -uint8_t x_183; -x_183 = !lean_is_exclusive(x_82); -if (x_183 == 0) +uint8_t x_185; +x_185 = !lean_is_exclusive(x_84); +if (x_185 == 0) { -return x_82; +return x_84; } else { -lean_object* x_184; lean_object* x_185; lean_object* x_186; -x_184 = lean_ctor_get(x_82, 0); -x_185 = lean_ctor_get(x_82, 1); -lean_inc(x_185); -lean_inc(x_184); -lean_dec(x_82); -x_186 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_186, 0, x_184); -lean_ctor_set(x_186, 1, x_185); -return x_186; +lean_object* x_186; lean_object* x_187; lean_object* x_188; +x_186 = lean_ctor_get(x_84, 0); +x_187 = lean_ctor_get(x_84, 1); +lean_inc(x_187); +lean_inc(x_186); +lean_dec(x_84); +x_188 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_188, 0, x_186); +lean_ctor_set(x_188, 1, x_187); +return x_188; } } } else { -uint8_t x_187; -x_187 = !lean_is_exclusive(x_77); -if (x_187 == 0) +uint8_t x_189; +x_189 = !lean_is_exclusive(x_79); +if (x_189 == 0) { -return x_77; +return x_79; } else { -lean_object* x_188; lean_object* x_189; lean_object* x_190; -x_188 = lean_ctor_get(x_77, 0); -x_189 = lean_ctor_get(x_77, 1); -lean_inc(x_189); -lean_inc(x_188); -lean_dec(x_77); -x_190 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_190, 0, x_188); -lean_ctor_set(x_190, 1, x_189); -return x_190; +lean_object* x_190; lean_object* x_191; lean_object* x_192; +x_190 = lean_ctor_get(x_79, 0); +x_191 = lean_ctor_get(x_79, 1); +lean_inc(x_191); +lean_inc(x_190); +lean_dec(x_79); +x_192 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_192, 0, x_190); +lean_ctor_set(x_192, 1, x_191); +return x_192; } } } else { -uint8_t x_191; -x_191 = !lean_is_exclusive(x_74); -if (x_191 == 0) +uint8_t x_193; +x_193 = !lean_is_exclusive(x_76); +if (x_193 == 0) { -return x_74; +return x_76; } else { -lean_object* x_192; lean_object* x_193; lean_object* x_194; -x_192 = lean_ctor_get(x_74, 0); -x_193 = lean_ctor_get(x_74, 1); -lean_inc(x_193); -lean_inc(x_192); -lean_dec(x_74); -x_194 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_194, 0, x_192); -lean_ctor_set(x_194, 1, x_193); -return x_194; +lean_object* x_194; lean_object* x_195; lean_object* x_196; +x_194 = lean_ctor_get(x_76, 0); +x_195 = lean_ctor_get(x_76, 1); +lean_inc(x_195); +lean_inc(x_194); +lean_dec(x_76); +x_196 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_196, 0, x_194); +lean_ctor_set(x_196, 1, x_195); +return x_196; } } } else { -uint8_t x_195; -x_195 = !lean_is_exclusive(x_71); -if (x_195 == 0) +uint8_t x_197; +x_197 = !lean_is_exclusive(x_73); +if (x_197 == 0) { -return x_71; +return x_73; } else { -lean_object* x_196; lean_object* x_197; lean_object* x_198; -x_196 = lean_ctor_get(x_71, 0); -x_197 = lean_ctor_get(x_71, 1); -lean_inc(x_197); -lean_inc(x_196); -lean_dec(x_71); -x_198 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_198, 0, x_196); -lean_ctor_set(x_198, 1, x_197); -return x_198; +lean_object* x_198; lean_object* x_199; lean_object* x_200; +x_198 = lean_ctor_get(x_73, 0); +x_199 = lean_ctor_get(x_73, 1); +lean_inc(x_199); +lean_inc(x_198); +lean_dec(x_73); +x_200 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_200, 0, x_198); +lean_ctor_set(x_200, 1, x_199); +return x_200; } } } else { -uint8_t x_199; -x_199 = !lean_is_exclusive(x_66); -if (x_199 == 0) +uint8_t x_201; +x_201 = !lean_is_exclusive(x_68); +if (x_201 == 0) { -return x_66; +return x_68; } else { -lean_object* x_200; lean_object* x_201; lean_object* x_202; -x_200 = lean_ctor_get(x_66, 0); -x_201 = lean_ctor_get(x_66, 1); -lean_inc(x_201); -lean_inc(x_200); -lean_dec(x_66); -x_202 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_202, 0, x_200); -lean_ctor_set(x_202, 1, x_201); -return x_202; +lean_object* x_202; lean_object* x_203; lean_object* x_204; +x_202 = lean_ctor_get(x_68, 0); +x_203 = lean_ctor_get(x_68, 1); +lean_inc(x_203); +lean_inc(x_202); +lean_dec(x_68); +x_204 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_204, 0, x_202); +lean_ctor_set(x_204, 1, x_203); +return x_204; } } } else { -uint8_t x_203; -x_203 = !lean_is_exclusive(x_63); -if (x_203 == 0) +uint8_t x_205; +x_205 = !lean_is_exclusive(x_65); +if (x_205 == 0) { -return x_63; +return x_65; } else { -lean_object* x_204; lean_object* x_205; lean_object* x_206; -x_204 = lean_ctor_get(x_63, 0); -x_205 = lean_ctor_get(x_63, 1); -lean_inc(x_205); -lean_inc(x_204); -lean_dec(x_63); -x_206 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_206, 0, x_204); -lean_ctor_set(x_206, 1, x_205); -return x_206; +lean_object* x_206; lean_object* x_207; lean_object* x_208; +x_206 = lean_ctor_get(x_65, 0); +x_207 = lean_ctor_get(x_65, 1); +lean_inc(x_207); +lean_inc(x_206); +lean_dec(x_65); +x_208 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_208, 0, x_206); +lean_ctor_set(x_208, 1, x_207); +return x_208; } } } else { -uint8_t x_207; -x_207 = !lean_is_exclusive(x_60); -if (x_207 == 0) +uint8_t x_209; +x_209 = !lean_is_exclusive(x_62); +if (x_209 == 0) { -return x_60; +return x_62; } else { -lean_object* x_208; lean_object* x_209; lean_object* x_210; -x_208 = lean_ctor_get(x_60, 0); -x_209 = lean_ctor_get(x_60, 1); -lean_inc(x_209); -lean_inc(x_208); -lean_dec(x_60); -x_210 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_210, 0, x_208); -lean_ctor_set(x_210, 1, x_209); -return x_210; +lean_object* x_210; lean_object* x_211; lean_object* x_212; +x_210 = lean_ctor_get(x_62, 0); +x_211 = lean_ctor_get(x_62, 1); +lean_inc(x_211); +lean_inc(x_210); +lean_dec(x_62); +x_212 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_212, 0, x_210); +lean_ctor_set(x_212, 1, x_211); +return x_212; } } } else { -uint8_t x_211; -x_211 = !lean_is_exclusive(x_55); -if (x_211 == 0) +uint8_t x_213; +x_213 = !lean_is_exclusive(x_57); +if (x_213 == 0) { -return x_55; +return x_57; } else { -lean_object* x_212; lean_object* x_213; lean_object* x_214; -x_212 = lean_ctor_get(x_55, 0); -x_213 = lean_ctor_get(x_55, 1); -lean_inc(x_213); -lean_inc(x_212); -lean_dec(x_55); -x_214 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_214, 0, x_212); -lean_ctor_set(x_214, 1, x_213); -return x_214; +lean_object* x_214; lean_object* x_215; lean_object* x_216; +x_214 = lean_ctor_get(x_57, 0); +x_215 = lean_ctor_get(x_57, 1); +lean_inc(x_215); +lean_inc(x_214); +lean_dec(x_57); +x_216 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_216, 0, x_214); +lean_ctor_set(x_216, 1, x_215); +return x_216; } } } else { -uint8_t x_215; -x_215 = !lean_is_exclusive(x_52); -if (x_215 == 0) +uint8_t x_217; +x_217 = !lean_is_exclusive(x_54); +if (x_217 == 0) { -return x_52; +return x_54; } else { -lean_object* x_216; lean_object* x_217; lean_object* x_218; -x_216 = lean_ctor_get(x_52, 0); -x_217 = lean_ctor_get(x_52, 1); -lean_inc(x_217); -lean_inc(x_216); -lean_dec(x_52); -x_218 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_218, 0, x_216); -lean_ctor_set(x_218, 1, x_217); -return x_218; +lean_object* x_218; lean_object* x_219; lean_object* x_220; +x_218 = lean_ctor_get(x_54, 0); +x_219 = lean_ctor_get(x_54, 1); +lean_inc(x_219); +lean_inc(x_218); +lean_dec(x_54); +x_220 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_220, 0, x_218); +lean_ctor_set(x_220, 1, x_219); +return x_220; } } } else { -uint8_t x_219; -x_219 = !lean_is_exclusive(x_49); -if (x_219 == 0) +uint8_t x_221; +x_221 = !lean_is_exclusive(x_51); +if (x_221 == 0) { -return x_49; +return x_51; } else { -lean_object* x_220; lean_object* x_221; lean_object* x_222; -x_220 = lean_ctor_get(x_49, 0); -x_221 = lean_ctor_get(x_49, 1); -lean_inc(x_221); -lean_inc(x_220); -lean_dec(x_49); -x_222 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_222, 0, x_220); -lean_ctor_set(x_222, 1, x_221); -return x_222; +lean_object* x_222; lean_object* x_223; lean_object* x_224; +x_222 = lean_ctor_get(x_51, 0); +x_223 = lean_ctor_get(x_51, 1); +lean_inc(x_223); +lean_inc(x_222); +lean_dec(x_51); +x_224 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_224, 0, x_222); +lean_ctor_set(x_224, 1, x_223); +return x_224; } } } else { -uint8_t x_223; -x_223 = !lean_is_exclusive(x_44); -if (x_223 == 0) +uint8_t x_225; +x_225 = !lean_is_exclusive(x_46); +if (x_225 == 0) { -return x_44; +return x_46; } else { -lean_object* x_224; lean_object* x_225; lean_object* x_226; -x_224 = lean_ctor_get(x_44, 0); -x_225 = lean_ctor_get(x_44, 1); -lean_inc(x_225); -lean_inc(x_224); -lean_dec(x_44); -x_226 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_226, 0, x_224); -lean_ctor_set(x_226, 1, x_225); -return x_226; +lean_object* x_226; lean_object* x_227; lean_object* x_228; +x_226 = lean_ctor_get(x_46, 0); +x_227 = lean_ctor_get(x_46, 1); +lean_inc(x_227); +lean_inc(x_226); +lean_dec(x_46); +x_228 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_228, 0, x_226); +lean_ctor_set(x_228, 1, x_227); +return x_228; } } } else { -uint8_t x_227; -x_227 = !lean_is_exclusive(x_41); -if (x_227 == 0) +uint8_t x_229; +x_229 = !lean_is_exclusive(x_43); +if (x_229 == 0) { -return x_41; +return x_43; } else { -lean_object* x_228; lean_object* x_229; lean_object* x_230; -x_228 = lean_ctor_get(x_41, 0); -x_229 = lean_ctor_get(x_41, 1); -lean_inc(x_229); -lean_inc(x_228); -lean_dec(x_41); -x_230 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_230, 0, x_228); -lean_ctor_set(x_230, 1, x_229); -return x_230; +lean_object* x_230; lean_object* x_231; lean_object* x_232; +x_230 = lean_ctor_get(x_43, 0); +x_231 = lean_ctor_get(x_43, 1); +lean_inc(x_231); +lean_inc(x_230); +lean_dec(x_43); +x_232 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_232, 0, x_230); +lean_ctor_set(x_232, 1, x_231); +return x_232; } } } else { -uint8_t x_231; -x_231 = !lean_is_exclusive(x_38); -if (x_231 == 0) +uint8_t x_233; +x_233 = !lean_is_exclusive(x_40); +if (x_233 == 0) { -return x_38; +return x_40; } else { -lean_object* x_232; lean_object* x_233; lean_object* x_234; -x_232 = lean_ctor_get(x_38, 0); -x_233 = lean_ctor_get(x_38, 1); -lean_inc(x_233); -lean_inc(x_232); -lean_dec(x_38); -x_234 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_234, 0, x_232); -lean_ctor_set(x_234, 1, x_233); -return x_234; +lean_object* x_234; lean_object* x_235; lean_object* x_236; +x_234 = lean_ctor_get(x_40, 0); +x_235 = lean_ctor_get(x_40, 1); +lean_inc(x_235); +lean_inc(x_234); +lean_dec(x_40); +x_236 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_236, 0, x_234); +lean_ctor_set(x_236, 1, x_235); +return x_236; } } } else { -uint8_t x_235; -x_235 = !lean_is_exclusive(x_34); -if (x_235 == 0) +uint8_t x_237; +x_237 = !lean_is_exclusive(x_36); +if (x_237 == 0) { -return x_34; +return x_36; } else { -lean_object* x_236; lean_object* x_237; lean_object* x_238; -x_236 = lean_ctor_get(x_34, 0); -x_237 = lean_ctor_get(x_34, 1); -lean_inc(x_237); -lean_inc(x_236); -lean_dec(x_34); -x_238 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_238, 0, x_236); -lean_ctor_set(x_238, 1, x_237); -return x_238; +lean_object* x_238; lean_object* x_239; lean_object* x_240; +x_238 = lean_ctor_get(x_36, 0); +x_239 = lean_ctor_get(x_36, 1); +lean_inc(x_239); +lean_inc(x_238); +lean_dec(x_36); +x_240 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_240, 0, x_238); +lean_ctor_set(x_240, 1, x_239); +return x_240; } } } else { -uint8_t x_239; -x_239 = !lean_is_exclusive(x_31); -if (x_239 == 0) +uint8_t x_241; +x_241 = !lean_is_exclusive(x_33); +if (x_241 == 0) { -return x_31; +return x_33; } else { -lean_object* x_240; lean_object* x_241; lean_object* x_242; -x_240 = lean_ctor_get(x_31, 0); -x_241 = lean_ctor_get(x_31, 1); -lean_inc(x_241); -lean_inc(x_240); -lean_dec(x_31); -x_242 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_242, 0, x_240); -lean_ctor_set(x_242, 1, x_241); -return x_242; +lean_object* x_242; lean_object* x_243; lean_object* x_244; +x_242 = lean_ctor_get(x_33, 0); +x_243 = lean_ctor_get(x_33, 1); +lean_inc(x_243); +lean_inc(x_242); +lean_dec(x_33); +x_244 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_244, 0, x_242); +lean_ctor_set(x_244, 1, x_243); +return x_244; } } } else { -uint8_t x_243; -x_243 = !lean_is_exclusive(x_28); -if (x_243 == 0) +uint8_t x_245; +x_245 = !lean_is_exclusive(x_30); +if (x_245 == 0) { -return x_28; +return x_30; } else { -lean_object* x_244; lean_object* x_245; lean_object* x_246; -x_244 = lean_ctor_get(x_28, 0); -x_245 = lean_ctor_get(x_28, 1); -lean_inc(x_245); -lean_inc(x_244); -lean_dec(x_28); -x_246 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_246, 0, x_244); -lean_ctor_set(x_246, 1, x_245); -return x_246; +lean_object* x_246; lean_object* x_247; lean_object* x_248; +x_246 = lean_ctor_get(x_30, 0); +x_247 = lean_ctor_get(x_30, 1); +lean_inc(x_247); +lean_inc(x_246); +lean_dec(x_30); +x_248 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_248, 0, x_246); +lean_ctor_set(x_248, 1, x_247); +return x_248; } } } else { -uint8_t x_247; -x_247 = !lean_is_exclusive(x_24); -if (x_247 == 0) +uint8_t x_249; +x_249 = !lean_is_exclusive(x_26); +if (x_249 == 0) { -return x_24; +return x_26; } else { -lean_object* x_248; lean_object* x_249; lean_object* x_250; -x_248 = lean_ctor_get(x_24, 0); -x_249 = lean_ctor_get(x_24, 1); -lean_inc(x_249); -lean_inc(x_248); -lean_dec(x_24); -x_250 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_250, 0, x_248); -lean_ctor_set(x_250, 1, x_249); -return x_250; +lean_object* x_250; lean_object* x_251; lean_object* x_252; +x_250 = lean_ctor_get(x_26, 0); +x_251 = lean_ctor_get(x_26, 1); +lean_inc(x_251); +lean_inc(x_250); +lean_dec(x_26); +x_252 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_252, 0, x_250); +lean_ctor_set(x_252, 1, x_251); +return x_252; } } } else { -uint8_t x_251; -x_251 = !lean_is_exclusive(x_21); -if (x_251 == 0) +uint8_t x_253; +x_253 = !lean_is_exclusive(x_23); +if (x_253 == 0) { -return x_21; +return x_23; } else { -lean_object* x_252; lean_object* x_253; lean_object* x_254; -x_252 = lean_ctor_get(x_21, 0); -x_253 = lean_ctor_get(x_21, 1); -lean_inc(x_253); -lean_inc(x_252); -lean_dec(x_21); -x_254 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_254, 0, x_252); -lean_ctor_set(x_254, 1, x_253); -return x_254; +lean_object* x_254; lean_object* x_255; lean_object* x_256; +x_254 = lean_ctor_get(x_23, 0); +x_255 = lean_ctor_get(x_23, 1); +lean_inc(x_255); +lean_inc(x_254); +lean_dec(x_23); +x_256 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_256, 0, x_254); +lean_ctor_set(x_256, 1, x_255); +return x_256; } } } else { -uint8_t x_255; -x_255 = !lean_is_exclusive(x_18); -if (x_255 == 0) +uint8_t x_257; +x_257 = !lean_is_exclusive(x_20); +if (x_257 == 0) { -return x_18; +return x_20; } else { -lean_object* x_256; lean_object* x_257; lean_object* x_258; -x_256 = lean_ctor_get(x_18, 0); -x_257 = lean_ctor_get(x_18, 1); -lean_inc(x_257); -lean_inc(x_256); -lean_dec(x_18); -x_258 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_258, 0, x_256); -lean_ctor_set(x_258, 1, x_257); -return x_258; +lean_object* x_258; lean_object* x_259; lean_object* x_260; +x_258 = lean_ctor_get(x_20, 0); +x_259 = lean_ctor_get(x_20, 1); +lean_inc(x_259); +lean_inc(x_258); +lean_dec(x_20); +x_260 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_260, 0, x_258); +lean_ctor_set(x_260, 1, x_259); +return x_260; } } } else { -uint8_t x_259; -x_259 = !lean_is_exclusive(x_13); -if (x_259 == 0) +uint8_t x_261; +x_261 = !lean_is_exclusive(x_14); +if (x_261 == 0) { -return x_13; +return x_14; } else { -lean_object* x_260; lean_object* x_261; lean_object* x_262; -x_260 = lean_ctor_get(x_13, 0); -x_261 = lean_ctor_get(x_13, 1); -lean_inc(x_261); -lean_inc(x_260); -lean_dec(x_13); -x_262 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_262, 0, x_260); -lean_ctor_set(x_262, 1, x_261); -return x_262; +lean_object* x_262; lean_object* x_263; lean_object* x_264; +x_262 = lean_ctor_get(x_14, 0); +x_263 = lean_ctor_get(x_14, 1); +lean_inc(x_263); +lean_inc(x_262); +lean_dec(x_14); +x_264 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_264, 0, x_262); +lean_ctor_set(x_264, 1, x_263); +return x_264; } } } else { -uint8_t x_263; -x_263 = !lean_is_exclusive(x_9); -if (x_263 == 0) +uint8_t x_265; +x_265 = !lean_is_exclusive(x_10); +if (x_265 == 0) { -return x_9; +return x_10; } else { -lean_object* x_264; lean_object* x_265; lean_object* x_266; -x_264 = lean_ctor_get(x_9, 0); -x_265 = lean_ctor_get(x_9, 1); -lean_inc(x_265); -lean_inc(x_264); -lean_dec(x_9); -x_266 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_266, 0, x_264); -lean_ctor_set(x_266, 1, x_265); -return x_266; +lean_object* x_266; lean_object* x_267; lean_object* x_268; +x_266 = lean_ctor_get(x_10, 0); +x_267 = lean_ctor_get(x_10, 1); +lean_inc(x_267); +lean_inc(x_266); +lean_dec(x_10); +x_268 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_268, 0, x_266); +lean_ctor_set(x_268, 1, x_267); +return x_268; } } } else { -uint8_t x_267; -x_267 = !lean_is_exclusive(x_5); -if (x_267 == 0) +uint8_t x_269; +x_269 = !lean_is_exclusive(x_6); +if (x_269 == 0) { -return x_5; +return x_6; } else { -lean_object* x_268; lean_object* x_269; lean_object* x_270; -x_268 = lean_ctor_get(x_5, 0); -x_269 = lean_ctor_get(x_5, 1); -lean_inc(x_269); -lean_inc(x_268); -lean_dec(x_5); -x_270 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_270, 0, x_268); -lean_ctor_set(x_270, 1, x_269); -return x_270; +lean_object* x_270; lean_object* x_271; lean_object* x_272; +x_270 = lean_ctor_get(x_6, 0); +x_271 = lean_ctor_get(x_6, 1); +lean_inc(x_271); +lean_inc(x_270); +lean_dec(x_6); +x_272 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_272, 0, x_270); +lean_ctor_set(x_272, 1, x_271); +return x_272; } } } @@ -10234,6 +10681,10 @@ l_Lean_Parser_mkAntiquotSplice_formatter___closed__5 = _init_l_Lean_Parser_mkAnt lean_mark_persistent(l_Lean_Parser_mkAntiquotSplice_formatter___closed__5); l_Lean_Parser_mkAntiquotSplice_formatter___closed__6 = _init_l_Lean_Parser_mkAntiquotSplice_formatter___closed__6(); lean_mark_persistent(l_Lean_Parser_mkAntiquotSplice_formatter___closed__6); +l_Lean_Parser_mkAntiquotSplice_formatter___closed__7 = _init_l_Lean_Parser_mkAntiquotSplice_formatter___closed__7(); +lean_mark_persistent(l_Lean_Parser_mkAntiquotSplice_formatter___closed__7); +l_Lean_Parser_mkAntiquotSplice_formatter___closed__8 = _init_l_Lean_Parser_mkAntiquotSplice_formatter___closed__8(); +lean_mark_persistent(l_Lean_Parser_mkAntiquotSplice_formatter___closed__8); l_Lean_Parser_sepByElemParser_formatter___closed__1 = _init_l_Lean_Parser_sepByElemParser_formatter___closed__1(); lean_mark_persistent(l_Lean_Parser_sepByElemParser_formatter___closed__1); l_Lean_Parser_sepByElemParser_formatter___closed__2 = _init_l_Lean_Parser_sepByElemParser_formatter___closed__2(); @@ -10280,6 +10731,8 @@ l_Lean_Parser_ident_formatter___closed__1 = _init_l_Lean_Parser_ident_formatter_ lean_mark_persistent(l_Lean_Parser_ident_formatter___closed__1); l_Lean_Parser_ident_formatter___closed__2 = _init_l_Lean_Parser_ident_formatter___closed__2(); lean_mark_persistent(l_Lean_Parser_ident_formatter___closed__2); +l_Lean_Parser_ident_formatter___closed__3 = _init_l_Lean_Parser_ident_formatter___closed__3(); +lean_mark_persistent(l_Lean_Parser_ident_formatter___closed__3); l_Lean_Parser_ident_parenthesizer___closed__1 = _init_l_Lean_Parser_ident_parenthesizer___closed__1(); lean_mark_persistent(l_Lean_Parser_ident_parenthesizer___closed__1); l_Lean_Parser_ident___elambda__1___closed__1 = _init_l_Lean_Parser_ident___elambda__1___closed__1(); @@ -10320,6 +10773,8 @@ l_Lean_Parser_numLit_formatter___closed__2 = _init_l_Lean_Parser_numLit_formatte lean_mark_persistent(l_Lean_Parser_numLit_formatter___closed__2); l_Lean_Parser_numLit_formatter___closed__3 = _init_l_Lean_Parser_numLit_formatter___closed__3(); lean_mark_persistent(l_Lean_Parser_numLit_formatter___closed__3); +l_Lean_Parser_numLit_formatter___closed__4 = _init_l_Lean_Parser_numLit_formatter___closed__4(); +lean_mark_persistent(l_Lean_Parser_numLit_formatter___closed__4); l_Lean_Parser_numLit_parenthesizer___closed__1 = _init_l_Lean_Parser_numLit_parenthesizer___closed__1(); lean_mark_persistent(l_Lean_Parser_numLit_parenthesizer___closed__1); l_Lean_Parser_numLit_parenthesizer___closed__2 = _init_l_Lean_Parser_numLit_parenthesizer___closed__2(); @@ -10344,6 +10799,8 @@ l_Lean_Parser_scientificLit_formatter___closed__2 = _init_l_Lean_Parser_scientif lean_mark_persistent(l_Lean_Parser_scientificLit_formatter___closed__2); l_Lean_Parser_scientificLit_formatter___closed__3 = _init_l_Lean_Parser_scientificLit_formatter___closed__3(); lean_mark_persistent(l_Lean_Parser_scientificLit_formatter___closed__3); +l_Lean_Parser_scientificLit_formatter___closed__4 = _init_l_Lean_Parser_scientificLit_formatter___closed__4(); +lean_mark_persistent(l_Lean_Parser_scientificLit_formatter___closed__4); l_Lean_Parser_scientificLit_parenthesizer___closed__1 = _init_l_Lean_Parser_scientificLit_parenthesizer___closed__1(); lean_mark_persistent(l_Lean_Parser_scientificLit_parenthesizer___closed__1); l_Lean_Parser_scientificLit_parenthesizer___closed__2 = _init_l_Lean_Parser_scientificLit_parenthesizer___closed__2(); @@ -10368,6 +10825,8 @@ l_Lean_Parser_strLit_formatter___closed__2 = _init_l_Lean_Parser_strLit_formatte lean_mark_persistent(l_Lean_Parser_strLit_formatter___closed__2); l_Lean_Parser_strLit_formatter___closed__3 = _init_l_Lean_Parser_strLit_formatter___closed__3(); lean_mark_persistent(l_Lean_Parser_strLit_formatter___closed__3); +l_Lean_Parser_strLit_formatter___closed__4 = _init_l_Lean_Parser_strLit_formatter___closed__4(); +lean_mark_persistent(l_Lean_Parser_strLit_formatter___closed__4); l_Lean_Parser_strLit_parenthesizer___closed__1 = _init_l_Lean_Parser_strLit_parenthesizer___closed__1(); lean_mark_persistent(l_Lean_Parser_strLit_parenthesizer___closed__1); l_Lean_Parser_strLit_parenthesizer___closed__2 = _init_l_Lean_Parser_strLit_parenthesizer___closed__2(); @@ -10392,6 +10851,8 @@ l_Lean_Parser_charLit_formatter___closed__2 = _init_l_Lean_Parser_charLit_format lean_mark_persistent(l_Lean_Parser_charLit_formatter___closed__2); l_Lean_Parser_charLit_formatter___closed__3 = _init_l_Lean_Parser_charLit_formatter___closed__3(); lean_mark_persistent(l_Lean_Parser_charLit_formatter___closed__3); +l_Lean_Parser_charLit_formatter___closed__4 = _init_l_Lean_Parser_charLit_formatter___closed__4(); +lean_mark_persistent(l_Lean_Parser_charLit_formatter___closed__4); l_Lean_Parser_charLit_parenthesizer___closed__1 = _init_l_Lean_Parser_charLit_parenthesizer___closed__1(); lean_mark_persistent(l_Lean_Parser_charLit_parenthesizer___closed__1); l_Lean_Parser_charLit_parenthesizer___closed__2 = _init_l_Lean_Parser_charLit_parenthesizer___closed__2(); @@ -10416,6 +10877,8 @@ l_Lean_Parser_nameLit_formatter___closed__2 = _init_l_Lean_Parser_nameLit_format lean_mark_persistent(l_Lean_Parser_nameLit_formatter___closed__2); l_Lean_Parser_nameLit_formatter___closed__3 = _init_l_Lean_Parser_nameLit_formatter___closed__3(); lean_mark_persistent(l_Lean_Parser_nameLit_formatter___closed__3); +l_Lean_Parser_nameLit_formatter___closed__4 = _init_l_Lean_Parser_nameLit_formatter___closed__4(); +lean_mark_persistent(l_Lean_Parser_nameLit_formatter___closed__4); l_Lean_Parser_nameLit_parenthesizer___closed__1 = _init_l_Lean_Parser_nameLit_parenthesizer___closed__1(); lean_mark_persistent(l_Lean_Parser_nameLit_parenthesizer___closed__1); l_Lean_Parser_nameLit_parenthesizer___closed__2 = _init_l_Lean_Parser_nameLit_parenthesizer___closed__2(); @@ -10434,6 +10897,10 @@ l_Lean_Parser_nameLit___closed__4 = _init_l_Lean_Parser_nameLit___closed__4(); lean_mark_persistent(l_Lean_Parser_nameLit___closed__4); l_Lean_Parser_nameLit = _init_l_Lean_Parser_nameLit(); lean_mark_persistent(l_Lean_Parser_nameLit); +l_Lean_Parser_group_formatter___closed__1 = _init_l_Lean_Parser_group_formatter___closed__1(); +lean_mark_persistent(l_Lean_Parser_group_formatter___closed__1); +l_Lean_Parser_group_formatter___closed__2 = _init_l_Lean_Parser_group_formatter___closed__2(); +lean_mark_persistent(l_Lean_Parser_group_formatter___closed__2); l_Lean_Parser_many1Indent_formatter___closed__1 = _init_l_Lean_Parser_many1Indent_formatter___closed__1(); lean_mark_persistent(l_Lean_Parser_many1Indent_formatter___closed__1); l_Lean_Parser_many1Indent_parenthesizer___closed__1 = _init_l_Lean_Parser_many1Indent_parenthesizer___closed__1(); @@ -10488,443 +10955,435 @@ l_Lean_ppDedent_formatter___closed__1 = _init_l_Lean_ppDedent_formatter___closed lean_mark_persistent(l_Lean_ppDedent_formatter___closed__1); l_Lean_ppDedentIfGrouped_formatter___closed__1 = _init_l_Lean_ppDedentIfGrouped_formatter___closed__1(); lean_mark_persistent(l_Lean_ppDedentIfGrouped_formatter___closed__1); -l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__1 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__1(); -lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__1); -l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__2 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__2(); -lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__2); -l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__3 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__3(); -lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__3); -l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__4 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__4(); -lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__4); -l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__5 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__5(); -lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__5); -l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__6 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__6(); -lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__6); -l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__7 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__7(); -lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__7); -l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__8 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__8(); -lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__8); -l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__9 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__9(); -lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__9); -l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__10 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__10(); -lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__10); -l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__11 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__11(); -lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__11); -l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__12 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__12(); -lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__12); -l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__13 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__13(); -lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__13); -l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__14 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__14(); -lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__14); -l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__15 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__15(); -lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__15); -l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__16 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__16(); -lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__16); -l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__17 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__17(); -lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__17); -l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__18 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__18(); -lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__18); -l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__19 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__19(); -lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__19); -l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__20 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__20(); -lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__20); -l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__21 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__21(); -lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__21); -l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__22 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__22(); -lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__22); -l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__23 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__23(); -lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__23); -l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__24 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__24(); -lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__24); -l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__25 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__25(); -lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__25); -l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__26 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__26(); -lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__26); -l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__27 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__27(); -lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__27); -l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__28 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__28(); -lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__28); -l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__29 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__29(); -lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__29); -l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__30 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__30(); -lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__30); -l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__31 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__31(); -lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_________closed__31); -l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29______ = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29______(); -lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29______); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__1 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__1(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__1); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__2 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__2(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__2); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__3 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__3(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__3); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__4 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__4(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__4); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__5 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__5(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__5); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__6 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__6(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__6); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__7 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__7(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__7); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__8 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__8(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__8); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__9 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__9(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__9); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__10 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__10(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__10); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__11 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__11(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__11); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__12 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__12(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__12); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__13 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__13(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__13); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__14 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__14(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__14); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__15 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__15(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__15); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__16 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__16(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__16); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__17 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__17(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__17); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__18 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__18(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__18); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__19 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__19(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__19); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__20 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__20(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__20); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__21 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__21(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__21); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__22 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__22(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__22); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__23 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__23(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__23); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__24 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__24(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__24); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__25 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__25(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__25); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__26 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__26(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__26); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__27 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__27(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__27); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__28 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__28(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__28); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__29 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__29(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__29); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__30 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__30(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__30); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__31 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__31(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__31); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__32 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__32(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__32); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__33 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__33(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__33); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__34 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__34(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__34); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__35 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__35(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__35); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__36 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__36(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__36); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__37 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__37(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__37); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__38 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__38(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__38); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__39 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__39(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__39); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__40 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__40(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__40); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__41 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__41(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__41); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__42 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__42(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__42); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__43 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__43(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__43); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__44 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__44(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__44); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__45 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__45(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__45); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__46 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__46(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__46); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__47 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__47(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__47); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__48 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__48(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__48); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__49 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__49(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__49); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__50 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__50(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__50); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__51 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__51(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__51); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__52 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__52(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__52); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__53 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__53(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__53); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__54 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__54(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__54); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__55 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__55(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__55); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__56 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__56(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__56); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__57 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__57(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__57); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__58 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__58(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__58); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__59 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__59(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__59); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__60 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__60(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__60); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__61 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__61(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__61); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__62 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__62(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__62); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__63 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__63(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__63); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__64 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__64(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__64); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__65 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__65(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__65); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__66 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__66(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__66); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__67 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__67(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__67); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__68 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__68(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__68); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__69 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__69(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__69); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__70 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__70(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__70); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__71 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__71(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__71); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__72 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__72(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__72); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__73 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__73(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__73); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__74 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__74(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__74); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__75 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__75(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29________1___closed__75); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__1(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__1); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__2(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__2); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__3(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__3); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__4 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__4(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__4); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__5 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__5(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__5); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__6 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__6(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__6); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__7 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__7(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__7); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__8 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__8(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__8); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__9 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__9(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__9); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__10 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__10(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__10); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__11 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__11(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__11); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__12 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__12(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__12); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__13 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__13(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__13); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__14 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__14(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__14); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__15 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__15(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__15); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__16 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__16(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__16); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__17 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__17(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__17); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__18 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__18(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__18); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__19 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__19(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__19); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__20 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__20(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__20); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__21 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__21(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__21); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__22 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__22(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__22); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__23 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__23(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__23); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__24 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__24(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__24); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__25 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__25(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__25); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__26 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__26(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__26); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__27 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__27(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__27); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__28 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__28(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__28); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__29 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__29(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__29); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__30 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__30(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__30); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__31 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__31(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__31); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__32 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__32(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__32); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__33 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__33(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__33); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__34 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__34(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__34); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__35 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__35(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__35); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__36 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__36(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__36); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__37 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__37(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__37); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__38 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__38(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__38); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__39 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__39(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__39); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__40 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__40(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__40); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__41 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__41(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__41); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__42 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__42(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__42); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__43 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__43(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__43); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__44 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__44(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__44); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__45 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__45(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__45); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__46 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__46(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__46); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__47 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__47(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__47); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__48 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__48(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__48); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__49 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__49(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__49); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__50 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__50(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__50); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__51 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__51(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__51); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__52 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__52(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__52); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__53 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__53(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__53); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__54 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__54(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__54); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__55 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__55(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__55); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__56 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__56(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__56); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__57 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__57(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__57); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__58 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__58(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__58); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__59 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__59(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__59); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__60 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__60(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__60); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__61 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__61(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__61); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__62 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__62(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__62); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__63 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__63(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__63); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__64 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__64(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__64); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__65 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__65(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__65); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__66 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__66(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__66); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__67 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__67(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__67); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__68 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__68(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__68); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__69 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__69(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__69); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__70 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__70(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__70); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__71 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__71(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__71); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__72 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__72(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__72); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__73 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__73(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__73); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__74 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__74(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__74); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__75 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__75(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__75); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__76 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__76(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__76); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__77 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__77(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__77); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__78 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__78(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__78); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__79 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__79(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__79); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__80 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__80(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__80); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__81 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__81(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__81); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__82 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__82(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__82); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__83 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__83(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__83); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__84 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__84(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__84); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__85 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__85(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__85); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__86 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__86(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__86); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__87 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__87(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__87); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__88 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__88(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__88); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__89 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__89(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__89); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__90 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__90(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__90); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__91 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__91(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__91); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__92 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__92(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__92); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__93 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__93(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__93); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__94 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__94(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__94); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__95 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__95(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__95); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__96 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__96(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__96); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__97 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__97(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__97); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__98 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__98(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__98); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__99 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__99(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__99); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__100 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__100(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__100); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__101 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__101(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__101); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__102 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__102(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__102); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__103 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__103(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__103); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__104 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__104(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__104); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__105 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__105(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__105); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__106 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__106(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__106); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__107 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__107(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__107); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__108 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__108(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__108); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__109 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__109(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__109); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__110 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__110(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__110); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__111 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__111(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304____closed__111); -res = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1304_(lean_io_mk_world()); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__1 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__1(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__1); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__2 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__2(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__2); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__3 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__3(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__3); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__4 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__4(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__4); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__5 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__5(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__5); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__6 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__6(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__6); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__7 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__7(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__7); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__8 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__8(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__8); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__9 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__9(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__9); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__10 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__10(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__10); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__11 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__11(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__11); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__12 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__12(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__12); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__13 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__13(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__13); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__14 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__14(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__14); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__15 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__15(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__15); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__16 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__16(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__16); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__17 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__17(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__17); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__18 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__18(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__18); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__19 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__19(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__19); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__20 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__20(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__20); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__21 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__21(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__21); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__22 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__22(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__22); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__23 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__23(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__23); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__24 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__24(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__24); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__25 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__25(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__25); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__26 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__26(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__26); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__27 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__27(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__27); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__28 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__28(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__28); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__29 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__29(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__29); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__30 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__30(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__30); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__31 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__31(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__31); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__32 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__32(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__32); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__33 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__33(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__33); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__34 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__34(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__34); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__35 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__35(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__35); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29________ = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29________(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29________); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__1 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__1); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__2 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__2); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__3 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__3(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__3); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__4 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__4(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__4); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__5 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__5(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__5); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__6 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__6(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__6); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__7 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__7(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__7); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__8 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__8(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__8); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__9 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__9(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__9); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__10 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__10(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__10); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__11 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__11(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__11); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__12 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__12(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__12); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__1 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__1); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__2 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__2); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__3 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__3(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__3); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__4 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__4(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__4); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__5 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__5(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__5); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__6 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__6(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__6); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__7 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__7(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__7); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__8 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__8(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__8); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__9 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__9(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__9); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__10 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__10(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__10); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__11 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__11(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__11); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__12 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__12(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__12); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__13 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__13(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__13); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__14 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__14(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__14); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__15 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__15(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__15); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__16 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__16(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__16); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__17 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__17(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__17); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__18 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__18(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__18); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__19 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__19(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__19); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__20 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__20(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__20); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__21 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__21(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__21); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__22 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__22(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__22); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__23 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__23(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__23); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__24 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__24(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__24); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__25 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__25(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__25); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__26 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__26(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__26); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__27 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__27(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__27); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__28 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__28(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__28); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__29 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__29(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__29); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__30 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__30(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__30); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__31 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__31(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__31); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__32 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__32(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__32); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__33 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__33(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__33); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__34 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__34(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__34); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__35 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__35(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__35); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__36 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__36(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__36); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__37 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__37(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__37); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__38 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__38(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__38); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__39 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__39(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__39); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__40 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__40(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__40); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__41 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__41(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__41); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__42 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__42(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__42); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__43 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__43(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__43); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__44 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__44(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__44); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__45 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__45(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__45); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__46 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__46(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__46); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__47 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__47(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__47); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__48 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__48(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__48); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__49 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__49(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__49); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__50 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__50(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__50); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__51 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__51(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__51); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__52 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__52(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__2___closed__52); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__1(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__1); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__2(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__2); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__3(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__3); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__4 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__4(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__4); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__5 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__5(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__5); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__6 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__6(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__6); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__7 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__7(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__7); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__8 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__8(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__8); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__9 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__9(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__9); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__10 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__10(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__10); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__11 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__11(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__11); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__12 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__12(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__12); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__13 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__13(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__13); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__14 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__14(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__14); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__15 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__15(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__15); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__16 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__16(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__16); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__17 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__17(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__17); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__18 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__18(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__18); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__19 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__19(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__19); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__20 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__20(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__20); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__21 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__21(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__21); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__22 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__22(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__22); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__23 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__23(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__23); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__24 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__24(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__24); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__25 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__25(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__25); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__26 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__26(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__26); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__27 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__27(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__27); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__28 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__28(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__28); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__29 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__29(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__29); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__30 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__30(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__30); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__31 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__31(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__31); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__32 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__32(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__32); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__33 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__33(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__33); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__34 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__34(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__34); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__35 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__35(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__35); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__36 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__36(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__36); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__37 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__37(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__37); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__38 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__38(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__38); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__39 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__39(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__39); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__40 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__40(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__40); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__41 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__41(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__41); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__42 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__42(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__42); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__43 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__43(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__43); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__44 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__44(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__44); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__45 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__45(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__45); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__46 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__46(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__46); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__47 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__47(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__47); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__48 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__48(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__48); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__49 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__49(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__49); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__50 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__50(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__50); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__51 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__51(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__51); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__52 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__52(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__52); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__53 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__53(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__53); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__54 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__54(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__54); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__55 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__55(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__55); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__56 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__56(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__56); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__57 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__57(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__57); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__58 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__58(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__58); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__59 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__59(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__59); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__60 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__60(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__60); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__61 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__61(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__61); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__62 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__62(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__62); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__63 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__63(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__63); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__64 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__64(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__64); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__65 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__65(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__65); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__66 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__66(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__66); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__67 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__67(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__67); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__68 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__68(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__68); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__69 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__69(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__69); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__70 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__70(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__70); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__71 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__71(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__71); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__72 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__72(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__72); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__73 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__73(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__73); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__74 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__74(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__74); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__75 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__75(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__75); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__76 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__76(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__76); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__77 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__77(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__77); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__78 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__78(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__78); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__79 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__79(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__79); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__80 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__80(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__80); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__81 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__81(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__81); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__82 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__82(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__82); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__83 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__83(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__83); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__84 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__84(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__84); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__85 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__85(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__85); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__86 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__86(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__86); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__87 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__87(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__87); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__88 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__88(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__88); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__89 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__89(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__89); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__90 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__90(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__90); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__91 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__91(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__91); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__92 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__92(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__92); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__93 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__93(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__93); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__94 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__94(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__94); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__95 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__95(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__95); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__96 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__96(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__96); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__97 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__97(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__97); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__98 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__98(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__98); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__99 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__99(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__99); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__100 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__100(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__100); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__101 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__101(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__101); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__102 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__102(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__102); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__103 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__103(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__103); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__104 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__104(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__104); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__105 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__105(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__105); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__106 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__106(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__106); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__107 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__107(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__107); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__108 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__108(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__108); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__109 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__109(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__109); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__110 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__110(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__110); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__111 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__111(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__111); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__112 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__112(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__112); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__113 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__113(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__113); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__114 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__114(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610____closed__114); +res = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_1610_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Parser/Level.c b/stage0/stdlib/Lean/Parser/Level.c index 3edbf3b6b8b..39bd2d2b5b1 100644 --- a/stage0/stdlib/Lean/Parser/Level.c +++ b/stage0/stdlib/Lean/Parser/Level.c @@ -125,6 +125,7 @@ lean_object* l_Lean_Parser_checkPrecFn(lean_object*, lean_object*, lean_object*) LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Level_ident_declRange(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Level_paren_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Level_max___elambda__1___lambda__2(lean_object*, lean_object*); +uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Level_addLit_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Level_num_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Level_max___closed__3; @@ -355,7 +356,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Level_hole_formatter(lean_object*, lean_o static lean_object* l_Lean_Parser_Level_hole___closed__7; static lean_object* l_Lean_Parser_Level_imax___closed__4; static lean_object* l_Lean_Parser_Level_paren_parenthesizer___closed__1; -uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Level_paren_formatter___closed__2; lean_object* l_Lean_Parser_ident_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_symbol_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -728,7 +728,7 @@ x_25 = l_Lean_Parser_checkPrecFn(x_24, x_1, x_2); x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); x_27 = lean_box(0); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_27); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_27); lean_dec(x_26); if (x_28 == 0) { @@ -770,7 +770,7 @@ goto block_67; lean_object* x_38; uint8_t x_39; x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_27); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_27); lean_dec(x_38); if (x_39 == 0) { @@ -779,7 +779,7 @@ lean_dec(x_19); x_40 = l_Lean_Parser_ParserState_mkNode(x_37, x_10, x_30); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_27); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_27); lean_dec(x_41); if (x_42 == 0) { @@ -803,7 +803,7 @@ lean_inc(x_1); x_46 = l_Lean_Parser_categoryParser___elambda__1(x_44, x_45, x_1, x_37); x_47 = lean_ctor_get(x_46, 4); lean_inc(x_47); -x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_27); +x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_27); lean_dec(x_47); if (x_48 == 0) { @@ -812,7 +812,7 @@ lean_dec(x_19); x_49 = l_Lean_Parser_ParserState_mkNode(x_46, x_10, x_30); x_50 = lean_ctor_get(x_49, 4); lean_inc(x_50); -x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_27); +x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_27); lean_dec(x_50); if (x_51 == 0) { @@ -845,7 +845,7 @@ lean_object* x_58; lean_object* x_59; uint8_t x_60; x_58 = l_Lean_Parser_ParserState_mkNode(x_54, x_10, x_30); x_59 = lean_ctor_get(x_58, 4); lean_inc(x_59); -x_60 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_59, x_27); +x_60 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_59, x_27); lean_dec(x_59); if (x_60 == 0) { @@ -868,7 +868,7 @@ x_62 = l_Lean_Parser_tokenAntiquotFn(x_1, x_54); x_63 = l_Lean_Parser_ParserState_mkNode(x_62, x_10, x_30); x_64 = lean_ctor_get(x_63, 4); lean_inc(x_64); -x_65 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_64, x_27); +x_65 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_64, x_27); lean_dec(x_64); if (x_65 == 0) { @@ -1684,7 +1684,7 @@ x_23 = l_Lean_Parser_checkPrecFn(x_22, x_1, x_2); x_24 = lean_ctor_get(x_23, 4); lean_inc(x_24); x_25 = lean_box(0); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_25); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_25); lean_dec(x_24); if (x_26 == 0) { @@ -1715,7 +1715,7 @@ if (x_34 == 0) lean_object* x_35; uint8_t x_36; x_35 = lean_ctor_get(x_30, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_25); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_25); lean_dec(x_35); if (x_36 == 0) { @@ -1724,7 +1724,7 @@ lean_dec(x_6); x_37 = l_Lean_Parser_ParserState_mkNode(x_30, x_8, x_28); x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_25); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_25); lean_dec(x_38); if (x_39 == 0) { @@ -1747,7 +1747,7 @@ x_41 = lean_apply_2(x_6, x_1, x_30); x_42 = l_Lean_Parser_ParserState_mkNode(x_41, x_8, x_28); x_43 = lean_ctor_get(x_42, 4); lean_inc(x_43); -x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_25); +x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_25); lean_dec(x_43); if (x_44 == 0) { @@ -1770,7 +1770,7 @@ lean_inc(x_1); x_46 = l_Lean_Parser_tokenAntiquotFn(x_1, x_30); x_47 = lean_ctor_get(x_46, 4); lean_inc(x_47); -x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_25); +x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_25); lean_dec(x_47); if (x_48 == 0) { @@ -1779,7 +1779,7 @@ lean_dec(x_6); x_49 = l_Lean_Parser_ParserState_mkNode(x_46, x_8, x_28); x_50 = lean_ctor_get(x_49, 4); lean_inc(x_50); -x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_25); +x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_25); lean_dec(x_50); if (x_51 == 0) { @@ -1802,7 +1802,7 @@ x_53 = lean_apply_2(x_6, x_1, x_46); x_54 = l_Lean_Parser_ParserState_mkNode(x_53, x_8, x_28); x_55 = lean_ctor_get(x_54, 4); lean_inc(x_55); -x_56 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_55, x_25); +x_56 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_55, x_25); lean_dec(x_55); if (x_56 == 0) { @@ -2461,7 +2461,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -2493,7 +2493,7 @@ if (x_26 == 0) lean_object* x_27; uint8_t x_28; x_27 = lean_ctor_get(x_22, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); lean_dec(x_27); if (x_28 == 0) { @@ -2503,7 +2503,7 @@ x_29 = l_Lean_Parser_Level_imax___elambda__1___closed__2; x_30 = l_Lean_Parser_ParserState_mkNode(x_22, x_29, x_19); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); lean_dec(x_31); if (x_32 == 0) { @@ -2527,7 +2527,7 @@ x_35 = l_Lean_Parser_Level_imax___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_19); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_16); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_16); lean_dec(x_37); if (x_38 == 0) { @@ -2550,7 +2550,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_22); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); lean_dec(x_41); if (x_42 == 0) { @@ -2560,7 +2560,7 @@ x_43 = l_Lean_Parser_Level_imax___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_19); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); lean_dec(x_45); if (x_46 == 0) { @@ -2584,7 +2584,7 @@ x_49 = l_Lean_Parser_Level_imax___elambda__1___closed__2; x_50 = l_Lean_Parser_ParserState_mkNode(x_48, x_49, x_19); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); lean_dec(x_51); if (x_52 == 0) { @@ -3155,7 +3155,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -3188,7 +3188,7 @@ x_25 = l_Lean_Parser_Level_hole___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_20, x_25, x_17); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); lean_dec(x_27); if (x_28 == 0) { @@ -3212,7 +3212,7 @@ x_31 = l_Lean_Parser_Level_hole___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_17); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -3606,7 +3606,7 @@ x_4 = l_Lean_Parser_checkPrecFn(x_3, x_1, x_2); x_5 = lean_ctor_get(x_4, 4); lean_inc(x_5); x_6 = lean_box(0); -x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); +x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); lean_dec(x_5); if (x_7 == 0) { @@ -3857,7 +3857,7 @@ x_4 = l_Lean_Parser_checkPrecFn(x_3, x_1, x_2); x_5 = lean_ctor_get(x_4, 4); lean_inc(x_5); x_6 = lean_box(0); -x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); +x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); lean_dec(x_5); if (x_7 == 0) { @@ -4145,7 +4145,7 @@ x_4 = l_Lean_Parser_checkPrecFn(x_3, x_1, x_2); x_5 = lean_ctor_get(x_4, 4); lean_inc(x_5); x_6 = lean_box(0); -x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); +x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); lean_dec(x_5); if (x_7 == 0) { @@ -4159,7 +4159,7 @@ x_8 = lean_unsigned_to_nat(0u); x_9 = l_Lean_Parser_checkLhsPrecFn(x_8, x_1, x_4); x_10 = lean_ctor_get(x_9, 4); lean_inc(x_10); -x_11 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_10, x_6); +x_11 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_10, x_6); lean_dec(x_10); if (x_11 == 0) { @@ -4194,7 +4194,7 @@ if (x_22 == 0) lean_object* x_23; uint8_t x_24; x_23 = lean_ctor_get(x_16, 4); lean_inc(x_23); -x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_6); +x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_6); lean_dec(x_23); if (x_24 == 0) { @@ -4204,7 +4204,7 @@ x_26 = l_Lean_Parser_ParserState_mkTrailingNode(x_16, x_25, x_13); lean_dec(x_13); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_6); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_6); lean_dec(x_27); if (x_28 == 0) { @@ -4229,7 +4229,7 @@ x_32 = l_Lean_Parser_ParserState_mkTrailingNode(x_30, x_31, x_13); lean_dec(x_13); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_6); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_6); lean_dec(x_33); if (x_34 == 0) { @@ -4252,7 +4252,7 @@ lean_inc(x_1); x_36 = l_Lean_Parser_tokenAntiquotFn(x_1, x_16); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_6); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_6); lean_dec(x_37); if (x_38 == 0) { @@ -4262,7 +4262,7 @@ x_40 = l_Lean_Parser_ParserState_mkTrailingNode(x_36, x_39, x_13); lean_dec(x_13); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_6); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_6); lean_dec(x_41); if (x_42 == 0) { @@ -4287,7 +4287,7 @@ x_46 = l_Lean_Parser_ParserState_mkTrailingNode(x_44, x_45, x_13); lean_dec(x_13); x_47 = lean_ctor_get(x_46, 4); lean_inc(x_47); -x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_6); +x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_6); lean_dec(x_47); if (x_48 == 0) { diff --git a/stage0/stdlib/Lean/Parser/Module.c b/stage0/stdlib/Lean/Parser/Module.c index 332d73f2ce0..f9c766d61a5 100644 --- a/stage0/stdlib/Lean/Parser/Module.c +++ b/stage0/stdlib/Lean/Parser/Module.c @@ -43,7 +43,6 @@ static lean_object* l_Lean_Parser_Module_import___closed__4; lean_object* l_Lean_PrettyPrinter_Parenthesizer_withAntiquot_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Module_module_formatter___closed__9; static lean_object* l_Lean_Parser_parseHeader___closed__2; -extern lean_object* l_Lean_nullKind; lean_object* l_Lean_Parser_many(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Module_0__Lean_Parser_mkEOI(lean_object*); lean_object* lean_name_mk_string(lean_object*, lean_object*); @@ -140,6 +139,7 @@ static lean_object* l_Lean_Parser_Module_module___closed__4; lean_object* l_Lean_Parser_optional(lean_object*); static lean_object* l_Lean_Parser_Module_module_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Module_prelude_parenthesizer___closed__3; +uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Module_header___elambda__1___closed__1; static lean_object* l_Lean_Parser_Module_module___elambda__1___closed__8; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Module_prelude_parenthesizer(lean_object*); @@ -304,6 +304,7 @@ lean_object* l_Lean_Parser_nodeFn(lean_object*, lean_object*, lean_object*, lean LEAN_EXPORT lean_object* l_Lean_Parser_testParseModuleAux_parse(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Module_import___elambda__1___closed__16; uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); +lean_object* l_Lean_mkListNode(lean_object*); static lean_object* l_Lean_Parser_Module_import___closed__5; static lean_object* l_Lean_Parser_Module_module_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Module_header_formatter___closed__3; @@ -341,7 +342,6 @@ static lean_object* l_Lean_Parser_Module_prelude___elambda__1___closed__14; static lean_object* l_Lean_Parser_Module_prelude___elambda__1___closed__8; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Module_import_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Module_import___elambda__1___closed__3; -uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Parser_ident_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at_Lean_Parser_parseCommand___spec__1___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_string_utf8_at_end(lean_object*, lean_object*); @@ -588,7 +588,7 @@ x_20 = l_Lean_Parser_checkPrecFn(x_19, x_1, x_2); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); x_22 = lean_box(0); -x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_22); +x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_22); lean_dec(x_21); if (x_23 == 0) { @@ -619,7 +619,7 @@ lean_object* x_32; lean_object* x_33; uint8_t x_34; x_32 = l_Lean_Parser_ParserState_mkNode(x_27, x_5, x_25); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_22); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_22); lean_dec(x_33); if (x_34 == 0) { @@ -642,7 +642,7 @@ x_36 = l_Lean_Parser_tokenAntiquotFn(x_1, x_27); x_37 = l_Lean_Parser_ParserState_mkNode(x_36, x_5, x_25); x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_22); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_22); lean_dec(x_38); if (x_39 == 0) { @@ -991,7 +991,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -1036,7 +1036,7 @@ goto block_49; lean_object* x_28; uint8_t x_29; x_28 = lean_ctor_get(x_27, 4); lean_inc(x_28); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_16); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_16); lean_dec(x_28); if (x_29 == 0) { @@ -1046,7 +1046,7 @@ x_30 = l_Lean_Parser_Module_import___elambda__1___closed__2; x_31 = l_Lean_Parser_ParserState_mkNode(x_27, x_30, x_19); x_32 = lean_ctor_get(x_31, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_16); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_16); lean_dec(x_32); if (x_33 == 0) { @@ -1068,7 +1068,7 @@ lean_inc(x_1); x_35 = lean_apply_2(x_4, x_1, x_27); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_16); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_16); lean_dec(x_36); if (x_37 == 0) { @@ -1077,7 +1077,7 @@ x_38 = l_Lean_Parser_Module_import___elambda__1___closed__2; x_39 = l_Lean_Parser_ParserState_mkNode(x_35, x_38, x_19); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_16); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_16); lean_dec(x_40); if (x_41 == 0) { @@ -1101,7 +1101,7 @@ x_44 = l_Lean_Parser_Module_import___elambda__1___closed__2; x_45 = l_Lean_Parser_ParserState_mkNode(x_43, x_44, x_19); x_46 = lean_ctor_get(x_45, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_16); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_16); lean_dec(x_46); if (x_47 == 0) { @@ -1468,7 +1468,7 @@ x_16 = l_Lean_Parser_checkPrecFn(x_15, x_1, x_2); x_17 = lean_ctor_get(x_16, 4); lean_inc(x_17); x_18 = lean_box(0); -x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); +x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); lean_dec(x_17); if (x_19 == 0) { @@ -1488,7 +1488,7 @@ lean_inc(x_1); x_22 = lean_apply_2(x_6, x_1, x_16); x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); -x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_18); +x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_18); lean_dec(x_23); if (x_24 == 0) { @@ -1498,7 +1498,7 @@ x_25 = l_Lean_Parser_Module_header___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_22, x_25, x_21); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_18); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_18); lean_dec(x_27); if (x_28 == 0) { @@ -1522,7 +1522,7 @@ x_31 = l_Lean_Parser_Module_header___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_21); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_18); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_18); lean_dec(x_33); if (x_34 == 0) { @@ -2883,7 +2883,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -2902,7 +2902,7 @@ lean_inc(x_1); x_20 = l_Lean_Parser_Module_header___elambda__1(x_1, x_14); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); -x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_16); +x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_16); lean_dec(x_21); if (x_22 == 0) { @@ -2912,7 +2912,7 @@ x_23 = l_Lean_Parser_Module_module_formatter___closed__2; x_24 = l_Lean_Parser_ParserState_mkNode(x_20, x_23, x_19); x_25 = lean_ctor_get(x_24, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_16); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_16); lean_dec(x_25); if (x_26 == 0) { @@ -2936,7 +2936,7 @@ x_29 = l_Lean_Parser_Module_module_formatter___closed__2; x_30 = l_Lean_Parser_ParserState_mkNode(x_28, x_29, x_19); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); lean_dec(x_31); if (x_32 == 0) { @@ -5169,101 +5169,93 @@ uint8_t x_15; x_15 = !lean_is_exclusive(x_14); if (x_15 == 0) { -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; lean_object* x_25; +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_16 = lean_ctor_get(x_14, 0); -x_17 = lean_box(2); -x_18 = l_Lean_nullKind; -x_19 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set(x_19, 1, x_18); -lean_ctor_set(x_19, 2, x_16); -x_20 = l_Lean_Parser_testParseModule___closed__1; -x_21 = lean_array_push(x_20, x_10); -x_22 = lean_array_push(x_21, x_19); -x_23 = l_Lean_Parser_Module_module_formatter___closed__2; -x_24 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_24, 0, x_17); -lean_ctor_set(x_24, 1, x_23); -lean_ctor_set(x_24, 2, x_22); -x_25 = l_Lean_Syntax_updateLeading(x_24); -lean_ctor_set(x_14, 0, x_25); +x_17 = l_Lean_mkListNode(x_16); +x_18 = l_Lean_Parser_testParseModule___closed__1; +x_19 = lean_array_push(x_18, x_10); +x_20 = lean_array_push(x_19, x_17); +x_21 = lean_box(2); +x_22 = l_Lean_Parser_Module_module_formatter___closed__2; +x_23 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +lean_ctor_set(x_23, 2, x_20); +x_24 = l_Lean_Syntax_updateLeading(x_23); +lean_ctor_set(x_14, 0, x_24); return x_14; } else { -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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_26 = lean_ctor_get(x_14, 0); -x_27 = lean_ctor_get(x_14, 1); -lean_inc(x_27); +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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_25 = lean_ctor_get(x_14, 0); +x_26 = lean_ctor_get(x_14, 1); lean_inc(x_26); +lean_inc(x_25); lean_dec(x_14); -x_28 = lean_box(2); -x_29 = l_Lean_nullKind; -x_30 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -lean_ctor_set(x_30, 2, x_26); -x_31 = l_Lean_Parser_testParseModule___closed__1; -x_32 = lean_array_push(x_31, x_10); -x_33 = lean_array_push(x_32, x_30); -x_34 = l_Lean_Parser_Module_module_formatter___closed__2; -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_28); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_35, 2, x_33); -x_36 = l_Lean_Syntax_updateLeading(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_27); -return x_37; +x_27 = l_Lean_mkListNode(x_25); +x_28 = l_Lean_Parser_testParseModule___closed__1; +x_29 = lean_array_push(x_28, x_10); +x_30 = lean_array_push(x_29, x_27); +x_31 = lean_box(2); +x_32 = l_Lean_Parser_Module_module_formatter___closed__2; +x_33 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +lean_ctor_set(x_33, 2, x_30); +x_34 = l_Lean_Syntax_updateLeading(x_33); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_26); +return x_35; } } else { -uint8_t x_38; +uint8_t x_36; lean_dec(x_10); -x_38 = !lean_is_exclusive(x_14); -if (x_38 == 0) +x_36 = !lean_is_exclusive(x_14); +if (x_36 == 0) { return x_14; } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_39 = lean_ctor_get(x_14, 0); -x_40 = lean_ctor_get(x_14, 1); -lean_inc(x_40); -lean_inc(x_39); +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_14, 0); +x_38 = lean_ctor_get(x_14, 1); +lean_inc(x_38); +lean_inc(x_37); lean_dec(x_14); -x_41 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_40); -return x_41; +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; } } } else { -uint8_t x_42; +uint8_t x_40; lean_dec(x_5); lean_dec(x_1); -x_42 = !lean_is_exclusive(x_6); -if (x_42 == 0) +x_40 = !lean_is_exclusive(x_6); +if (x_40 == 0) { return x_6; } else { -lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_43 = lean_ctor_get(x_6, 0); -x_44 = lean_ctor_get(x_6, 1); -lean_inc(x_44); -lean_inc(x_43); +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_6, 0); +x_42 = lean_ctor_get(x_6, 1); +lean_inc(x_42); +lean_inc(x_41); lean_dec(x_6); -x_45 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_45, 0, x_43); -lean_ctor_set(x_45, 1, x_44); -return x_45; +x_43 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +return x_43; } } } diff --git a/stage0/stdlib/Lean/Parser/StrInterpolation.c b/stage0/stdlib/Lean/Parser/StrInterpolation.c index df10352ce7b..7d798e6dbc9 100644 --- a/stage0/stdlib/Lean/Parser/StrInterpolation.c +++ b/stage0/stdlib/Lean/Parser/StrInterpolation.c @@ -15,15 +15,19 @@ extern "C" { #endif lean_object* l_Lean_Parser_ParserState_mkError(lean_object*, lean_object*); lean_object* l_Lean_Parser_quotedCharCoreFn(lean_object*, lean_object*, lean_object*); +lean_object* lean_name_mk_string(lean_object*, lean_object*); uint8_t l_Lean_Parser_isQuotableCharDefault(uint32_t); +static lean_object* l_Lean_Parser_interpolatedStrFn_parse___closed__5; lean_object* l_Lean_Parser_ParserState_next(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_interpolatedStrNoAntiquot___closed__1; lean_object* l_Lean_Parser_ParserState_pushSyntax(lean_object*, lean_object*); lean_object* l_Lean_Parser_ParserState_mkNode(lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); -extern lean_object* l_Lean_interpolatedStrKind; lean_object* l_Lean_Parser_mkAtomicInfo(lean_object*); lean_object* l_Lean_Parser_mkNodeToken(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_interpolatedStrFn_parse___closed__7; +static lean_object* l_Lean_Parser_interpolatedStrFn_parse___closed__4; +uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(lean_object*, lean_object*); lean_object* lean_string_utf8_next(lean_object*, lean_object*); static lean_object* l_Lean_Parser_interpolatedStr___elambda__1___closed__1; lean_object* l_Lean_Parser_ParserState_setPos(lean_object*, lean_object*); @@ -35,6 +39,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_interpolatedStrNoAntiquot(lean_object*); static lean_object* l_Lean_Parser_interpolatedStrFn_parse___closed__3; lean_object* l_Lean_Parser_orelseInfo(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Parser_isQuotableCharForStrInterpolant(uint32_t); +static lean_object* l_Lean_Parser_interpolatedStrFn_parse___closed__6; uint32_t lean_string_utf8_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_interpolatedStr(lean_object*); static lean_object* l_Lean_Parser_interpolatedStrFn___closed__2; @@ -49,9 +54,7 @@ lean_object* l_Lean_Parser_mkAntiquot(lean_object*, lean_object*, uint8_t, uint8 LEAN_EXPORT lean_object* l_Lean_Parser_isQuotableCharForStrInterpolant___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_interpolatedStrFn(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_ParserState_mkUnexpectedError(lean_object*, lean_object*, lean_object*); -uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(lean_object*, lean_object*); uint8_t lean_string_utf8_at_end(lean_object*, lean_object*); -extern lean_object* l_Lean_interpolatedStrLitKind; LEAN_EXPORT uint8_t l_Lean_Parser_isQuotableCharForStrInterpolant(uint32_t x_1) { _start: { @@ -87,19 +90,55 @@ static lean_object* _init_l_Lean_Parser_interpolatedStrFn_parse___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("'}'", 3); +x_1 = lean_mk_string_from_bytes("interpolatedStrLitKind", 22); return x_1; } } static lean_object* _init_l_Lean_Parser_interpolatedStrFn_parse___closed__2() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_interpolatedStrFn_parse___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_interpolatedStrFn_parse___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("interpolatedStrKind", 19); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_interpolatedStrFn_parse___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_interpolatedStrFn_parse___closed__3; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_interpolatedStrFn_parse___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("'}'", 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_interpolatedStrFn_parse___closed__6() { +_start: +{ lean_object* x_1; x_1 = lean_alloc_closure((void*)(l_Lean_Parser_isQuotableCharForStrInterpolant___boxed), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_interpolatedStrFn_parse___closed__3() { +static lean_object* _init_l_Lean_Parser_interpolatedStrFn_parse___closed__7() { _start: { lean_object* x_1; @@ -141,7 +180,7 @@ goto _start; else { lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_19 = l_Lean_interpolatedStrLitKind; +x_19 = l_Lean_Parser_interpolatedStrFn_parse___closed__2; x_20 = l_Lean_Parser_mkNodeToken(x_19, x_4, x_5, x_11); lean_inc(x_1); lean_inc(x_5); @@ -149,7 +188,7 @@ x_21 = lean_apply_2(x_1, x_5, x_20); x_22 = lean_ctor_get(x_21, 4); lean_inc(x_22); x_23 = lean_box(0); -x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_22, x_23); +x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_22, x_23); lean_dec(x_22); if (x_24 == 0) { @@ -174,9 +213,9 @@ lean_dec(x_5); lean_dec(x_1); x_29 = lean_box(0); x_30 = l_Lean_Parser_ParserState_pushSyntax(x_21, x_29); -x_31 = l_Lean_interpolatedStrKind; +x_31 = l_Lean_Parser_interpolatedStrFn_parse___closed__4; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_3); -x_33 = l_Lean_Parser_interpolatedStrFn_parse___closed__1; +x_33 = l_Lean_Parser_interpolatedStrFn_parse___closed__5; x_34 = l_Lean_Parser_ParserState_setError(x_32, x_33); return x_34; } @@ -195,12 +234,12 @@ goto _start; else { lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; -x_38 = l_Lean_Parser_interpolatedStrFn_parse___closed__2; +x_38 = l_Lean_Parser_interpolatedStrFn_parse___closed__6; x_39 = l_Lean_Parser_quotedCharCoreFn(x_38, x_5, x_11); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); x_41 = lean_box(0); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_41); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_41); lean_dec(x_40); if (x_42 == 0) { @@ -221,10 +260,10 @@ else { lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_dec(x_1); -x_44 = l_Lean_interpolatedStrLitKind; +x_44 = l_Lean_Parser_interpolatedStrFn_parse___closed__2; x_45 = l_Lean_Parser_mkNodeToken(x_44, x_4, x_5, x_11); lean_dec(x_5); -x_46 = l_Lean_interpolatedStrKind; +x_46 = l_Lean_Parser_interpolatedStrFn_parse___closed__4; x_47 = l_Lean_Parser_ParserState_mkNode(x_45, x_46, x_3); return x_47; } @@ -238,9 +277,9 @@ lean_dec(x_4); lean_dec(x_1); x_48 = lean_box(0); x_49 = l_Lean_Parser_ParserState_pushSyntax(x_6, x_48); -x_50 = l_Lean_interpolatedStrKind; +x_50 = l_Lean_Parser_interpolatedStrFn_parse___closed__4; x_51 = l_Lean_Parser_ParserState_mkNode(x_49, x_50, x_3); -x_52 = l_Lean_Parser_interpolatedStrFn_parse___closed__3; +x_52 = l_Lean_Parser_interpolatedStrFn_parse___closed__7; x_53 = l_Lean_Parser_ParserState_setError(x_51, x_52); return x_53; } @@ -367,7 +406,7 @@ static lean_object* _init_l_Lean_Parser_interpolatedStr___elambda__1___closed__1 { lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; x_1 = l_Lean_Parser_interpolatedStrNoAntiquot___closed__1; -x_2 = l_Lean_interpolatedStrKind; +x_2 = l_Lean_Parser_interpolatedStrFn_parse___closed__4; x_3 = 1; x_4 = 0; x_5 = l_Lean_Parser_mkAntiquot(x_1, x_2, x_3, x_4); @@ -458,6 +497,14 @@ l_Lean_Parser_interpolatedStrFn_parse___closed__2 = _init_l_Lean_Parser_interpol lean_mark_persistent(l_Lean_Parser_interpolatedStrFn_parse___closed__2); l_Lean_Parser_interpolatedStrFn_parse___closed__3 = _init_l_Lean_Parser_interpolatedStrFn_parse___closed__3(); lean_mark_persistent(l_Lean_Parser_interpolatedStrFn_parse___closed__3); +l_Lean_Parser_interpolatedStrFn_parse___closed__4 = _init_l_Lean_Parser_interpolatedStrFn_parse___closed__4(); +lean_mark_persistent(l_Lean_Parser_interpolatedStrFn_parse___closed__4); +l_Lean_Parser_interpolatedStrFn_parse___closed__5 = _init_l_Lean_Parser_interpolatedStrFn_parse___closed__5(); +lean_mark_persistent(l_Lean_Parser_interpolatedStrFn_parse___closed__5); +l_Lean_Parser_interpolatedStrFn_parse___closed__6 = _init_l_Lean_Parser_interpolatedStrFn_parse___closed__6(); +lean_mark_persistent(l_Lean_Parser_interpolatedStrFn_parse___closed__6); +l_Lean_Parser_interpolatedStrFn_parse___closed__7 = _init_l_Lean_Parser_interpolatedStrFn_parse___closed__7(); +lean_mark_persistent(l_Lean_Parser_interpolatedStrFn_parse___closed__7); l_Lean_Parser_interpolatedStrFn___closed__1 = _init_l_Lean_Parser_interpolatedStrFn___closed__1(); lean_mark_persistent(l_Lean_Parser_interpolatedStrFn___closed__1); l_Lean_Parser_interpolatedStrFn___closed__2 = _init_l_Lean_Parser_interpolatedStrFn___closed__2(); diff --git a/stage0/stdlib/Lean/Parser/Syntax.c b/stage0/stdlib/Lean/Parser/Syntax.c index e5bc22570c8..d5475a34462 100644 --- a/stage0/stdlib/Lean/Parser/Syntax.c +++ b/stage0/stdlib/Lean/Parser/Syntax.c @@ -595,6 +595,7 @@ static lean_object* l_Lean_Parser_Command_elab_formatter___closed__13; LEAN_EXPORT lean_object* l_Lean_Parser_Command_elab; static lean_object* l_Lean_Parser_Command_elab___closed__7; static lean_object* l_Lean_Parser_Command_identPrec___closed__5; +uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_prio_quot_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Syntax_sepBy___elambda__1___closed__6; static lean_object* l_Lean_Parser_Command_notation___closed__6; @@ -2000,7 +2001,6 @@ static lean_object* l_Lean_Parser_Syntax_cat___elambda__1___closed__1; static lean_object* l_Lean_Parser_Command_syntaxCat___elambda__1___closed__13; static lean_object* l_Lean_Parser_Command_syntax_formatter___closed__5; static lean_object* l_Lean_Parser_Command_syntax_parenthesizer___closed__7; -uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_elab___closed__8; static lean_object* l_Lean_Parser_Command_elab__rules_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Command_macroTail_formatter___closed__7; @@ -2607,7 +2607,7 @@ x_22 = l_Lean_Parser_checkPrecFn(x_21, x_1, x_2); x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); x_24 = lean_box(0); -x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_24); +x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_24); lean_dec(x_23); if (x_25 == 0) { @@ -2637,7 +2637,7 @@ if (x_33 == 0) lean_object* x_34; uint8_t x_35; x_34 = lean_ctor_get(x_29, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_24); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_24); lean_dec(x_34); if (x_35 == 0) { @@ -2645,7 +2645,7 @@ lean_object* x_36; lean_object* x_37; uint8_t x_38; x_36 = l_Lean_Parser_ParserState_mkNode(x_29, x_7, x_27); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_24); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_24); lean_dec(x_37); if (x_38 == 0) { @@ -2670,7 +2670,7 @@ x_42 = l_Lean_Parser_categoryParser___elambda__1(x_40, x_41, x_1, x_29); x_43 = l_Lean_Parser_ParserState_mkNode(x_42, x_7, x_27); x_44 = lean_ctor_get(x_43, 4); lean_inc(x_44); -x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_24); +x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_24); lean_dec(x_44); if (x_45 == 0) { @@ -2693,7 +2693,7 @@ lean_inc(x_1); x_47 = l_Lean_Parser_tokenAntiquotFn(x_1, x_29); x_48 = lean_ctor_get(x_47, 4); lean_inc(x_48); -x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_24); +x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_24); lean_dec(x_48); if (x_49 == 0) { @@ -2701,7 +2701,7 @@ lean_object* x_50; lean_object* x_51; uint8_t x_52; x_50 = l_Lean_Parser_ParserState_mkNode(x_47, x_7, x_27); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_24); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_24); lean_dec(x_51); if (x_52 == 0) { @@ -2726,7 +2726,7 @@ x_56 = l_Lean_Parser_categoryParser___elambda__1(x_54, x_55, x_1, x_47); x_57 = l_Lean_Parser_ParserState_mkNode(x_56, x_7, x_27); x_58 = lean_ctor_get(x_57, 4); lean_inc(x_58); -x_59 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_58, x_24); +x_59 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_58, x_24); lean_dec(x_58); if (x_59 == 0) { @@ -2914,7 +2914,7 @@ x_4 = l_Lean_Parser_checkPrecFn(x_3, x_1, x_2); x_5 = lean_ctor_get(x_4, 4); lean_inc(x_5); x_6 = lean_box(0); -x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); +x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); lean_dec(x_5); if (x_7 == 0) { @@ -3409,7 +3409,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -3453,7 +3453,7 @@ goto block_60; lean_object* x_28; uint8_t x_29; x_28 = lean_ctor_get(x_27, 4); lean_inc(x_28); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_16); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_16); lean_dec(x_28); if (x_29 == 0) { @@ -3464,7 +3464,7 @@ x_30 = l_Lean_Parser_Syntax_paren___elambda__1___closed__2; x_31 = l_Lean_Parser_ParserState_mkNode(x_27, x_30, x_19); x_32 = lean_ctor_get(x_31, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_16); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_16); lean_dec(x_32); if (x_33 == 0) { @@ -3486,7 +3486,7 @@ lean_inc(x_1); x_35 = lean_apply_2(x_4, x_1, x_27); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_16); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_16); lean_dec(x_36); if (x_37 == 0) { @@ -3496,7 +3496,7 @@ x_38 = l_Lean_Parser_Syntax_paren___elambda__1___closed__2; x_39 = l_Lean_Parser_ParserState_mkNode(x_35, x_38, x_19); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_16); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_16); lean_dec(x_40); if (x_41 == 0) { @@ -3531,7 +3531,7 @@ x_49 = l_Lean_Parser_Syntax_paren___elambda__1___closed__2; x_50 = l_Lean_Parser_ParserState_mkNode(x_45, x_49, x_19); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); lean_dec(x_51); if (x_52 == 0) { @@ -3555,7 +3555,7 @@ x_55 = l_Lean_Parser_Syntax_paren___elambda__1___closed__2; x_56 = l_Lean_Parser_ParserState_mkNode(x_54, x_55, x_19); x_57 = lean_ctor_get(x_56, 4); lean_inc(x_57); -x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_16); +x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_16); lean_dec(x_57); if (x_58 == 0) { @@ -4267,7 +4267,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -4286,7 +4286,7 @@ lean_inc(x_1); x_20 = l_Lean_Parser_ident___elambda__1(x_1, x_14); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); -x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_16); +x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_16); lean_dec(x_21); if (x_22 == 0) { @@ -4296,7 +4296,7 @@ x_23 = l_Lean_Parser_Syntax_cat___elambda__1___closed__2; x_24 = l_Lean_Parser_ParserState_mkNode(x_20, x_23, x_19); x_25 = lean_ctor_get(x_24, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_16); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_16); lean_dec(x_25); if (x_26 == 0) { @@ -4320,7 +4320,7 @@ x_29 = l_Lean_Parser_Syntax_cat___elambda__1___closed__2; x_30 = l_Lean_Parser_ParserState_mkNode(x_28, x_29, x_19); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); lean_dec(x_31); if (x_32 == 0) { @@ -5164,7 +5164,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -5184,7 +5184,7 @@ lean_inc(x_1); x_20 = l_Lean_Parser_ident___elambda__1(x_1, x_14); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); -x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_16); +x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_16); lean_dec(x_21); if (x_22 == 0) { @@ -5195,7 +5195,7 @@ x_23 = l_Lean_Parser_Syntax_unary___elambda__1___closed__2; x_24 = l_Lean_Parser_ParserState_mkNode(x_20, x_23, x_19); x_25 = lean_ctor_get(x_24, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_16); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_16); lean_dec(x_25); if (x_26 == 0) { @@ -5217,7 +5217,7 @@ x_28 = l_Lean_Parser_Syntax_unary___elambda__1___closed__4; x_29 = l_Lean_Parser_checkNoWsBeforeFn(x_28, x_1, x_20); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_16); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_16); lean_dec(x_30); if (x_31 == 0) { @@ -5228,7 +5228,7 @@ x_32 = l_Lean_Parser_Syntax_unary___elambda__1___closed__2; x_33 = l_Lean_Parser_ParserState_mkNode(x_29, x_32, x_19); x_34 = lean_ctor_get(x_33, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_16); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_16); lean_dec(x_34); if (x_35 == 0) { @@ -5274,7 +5274,7 @@ goto block_77; lean_object* x_45; uint8_t x_46; x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); lean_dec(x_45); if (x_46 == 0) { @@ -5285,7 +5285,7 @@ x_47 = l_Lean_Parser_Syntax_unary___elambda__1___closed__2; x_48 = l_Lean_Parser_ParserState_mkNode(x_44, x_47, x_19); x_49 = lean_ctor_get(x_48, 4); lean_inc(x_49); -x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_16); +x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_16); lean_dec(x_49); if (x_50 == 0) { @@ -5307,7 +5307,7 @@ lean_inc(x_1); x_52 = lean_apply_2(x_4, x_1, x_44); x_53 = lean_ctor_get(x_52, 4); lean_inc(x_53); -x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_16); +x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_16); lean_dec(x_53); if (x_54 == 0) { @@ -5317,7 +5317,7 @@ x_55 = l_Lean_Parser_Syntax_unary___elambda__1___closed__2; x_56 = l_Lean_Parser_ParserState_mkNode(x_52, x_55, x_19); x_57 = lean_ctor_get(x_56, 4); lean_inc(x_57); -x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_16); +x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_16); lean_dec(x_57); if (x_58 == 0) { @@ -5352,7 +5352,7 @@ x_66 = l_Lean_Parser_Syntax_unary___elambda__1___closed__2; x_67 = l_Lean_Parser_ParserState_mkNode(x_62, x_66, x_19); x_68 = lean_ctor_get(x_67, 4); lean_inc(x_68); -x_69 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_68, x_16); +x_69 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_68, x_16); lean_dec(x_68); if (x_69 == 0) { @@ -5376,7 +5376,7 @@ x_72 = l_Lean_Parser_Syntax_unary___elambda__1___closed__2; x_73 = l_Lean_Parser_ParserState_mkNode(x_71, x_72, x_19); x_74 = lean_ctor_get(x_73, 4); lean_inc(x_74); -x_75 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_74, x_16); +x_75 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_74, x_16); lean_dec(x_74); if (x_75 == 0) { @@ -6030,7 +6030,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -6050,7 +6050,7 @@ lean_inc(x_1); x_20 = l_Lean_Parser_ident___elambda__1(x_1, x_14); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); -x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_16); +x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_16); lean_dec(x_21); if (x_22 == 0) { @@ -6061,7 +6061,7 @@ x_23 = l_Lean_Parser_Syntax_binary___elambda__1___closed__2; x_24 = l_Lean_Parser_ParserState_mkNode(x_20, x_23, x_19); x_25 = lean_ctor_get(x_24, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_16); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_16); lean_dec(x_25); if (x_26 == 0) { @@ -6083,7 +6083,7 @@ x_28 = l_Lean_Parser_Syntax_unary___elambda__1___closed__4; x_29 = l_Lean_Parser_checkNoWsBeforeFn(x_28, x_1, x_20); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_16); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_16); lean_dec(x_30); if (x_31 == 0) { @@ -6094,7 +6094,7 @@ x_32 = l_Lean_Parser_Syntax_binary___elambda__1___closed__2; x_33 = l_Lean_Parser_ParserState_mkNode(x_29, x_32, x_19); x_34 = lean_ctor_get(x_33, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_16); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_16); lean_dec(x_34); if (x_35 == 0) { @@ -6140,7 +6140,7 @@ goto block_101; lean_object* x_45; uint8_t x_46; x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); lean_dec(x_45); if (x_46 == 0) { @@ -6151,7 +6151,7 @@ x_47 = l_Lean_Parser_Syntax_binary___elambda__1___closed__2; x_48 = l_Lean_Parser_ParserState_mkNode(x_44, x_47, x_19); x_49 = lean_ctor_get(x_48, 4); lean_inc(x_49); -x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_16); +x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_16); lean_dec(x_49); if (x_50 == 0) { @@ -6174,7 +6174,7 @@ lean_inc(x_1); x_52 = lean_apply_2(x_4, x_1, x_44); x_53 = lean_ctor_get(x_52, 4); lean_inc(x_53); -x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_16); +x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_16); lean_dec(x_53); if (x_54 == 0) { @@ -6185,7 +6185,7 @@ x_55 = l_Lean_Parser_Syntax_binary___elambda__1___closed__2; x_56 = l_Lean_Parser_ParserState_mkNode(x_52, x_55, x_19); x_57 = lean_ctor_get(x_56, 4); lean_inc(x_57); -x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_16); +x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_16); lean_dec(x_57); if (x_58 == 0) { @@ -6230,7 +6230,7 @@ goto block_99; lean_object* x_67; uint8_t x_68; x_67 = lean_ctor_get(x_66, 4); lean_inc(x_67); -x_68 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_67, x_16); +x_68 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_67, x_16); lean_dec(x_67); if (x_68 == 0) { @@ -6241,7 +6241,7 @@ x_69 = l_Lean_Parser_Syntax_binary___elambda__1___closed__2; x_70 = l_Lean_Parser_ParserState_mkNode(x_66, x_69, x_19); x_71 = lean_ctor_get(x_70, 4); lean_inc(x_71); -x_72 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_71, x_16); +x_72 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_71, x_16); lean_dec(x_71); if (x_72 == 0) { @@ -6263,7 +6263,7 @@ lean_inc(x_1); x_74 = lean_apply_2(x_4, x_1, x_66); x_75 = lean_ctor_get(x_74, 4); lean_inc(x_75); -x_76 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_75, x_16); +x_76 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_75, x_16); lean_dec(x_75); if (x_76 == 0) { @@ -6273,7 +6273,7 @@ x_77 = l_Lean_Parser_Syntax_binary___elambda__1___closed__2; x_78 = l_Lean_Parser_ParserState_mkNode(x_74, x_77, x_19); x_79 = lean_ctor_get(x_78, 4); lean_inc(x_79); -x_80 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_79, x_16); +x_80 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_79, x_16); lean_dec(x_79); if (x_80 == 0) { @@ -6308,7 +6308,7 @@ x_88 = l_Lean_Parser_Syntax_binary___elambda__1___closed__2; x_89 = l_Lean_Parser_ParserState_mkNode(x_84, x_88, x_19); x_90 = lean_ctor_get(x_89, 4); lean_inc(x_90); -x_91 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_90, x_16); +x_91 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_90, x_16); lean_dec(x_90); if (x_91 == 0) { @@ -6332,7 +6332,7 @@ x_94 = l_Lean_Parser_Syntax_binary___elambda__1___closed__2; x_95 = l_Lean_Parser_ParserState_mkNode(x_93, x_94, x_19); x_96 = lean_ctor_get(x_95, 4); lean_inc(x_96); -x_97 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_96, x_16); +x_97 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_96, x_16); lean_dec(x_96); if (x_97 == 0) { @@ -7189,7 +7189,7 @@ x_39 = l_Lean_Parser_checkPrecFn(x_38, x_1, x_2); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); x_41 = lean_box(0); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_41); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_41); lean_dec(x_40); if (x_42 == 0) { @@ -7236,7 +7236,7 @@ lean_object* x_53; lean_object* x_54; uint8_t x_55; x_53 = l_Lean_Parser_ParserState_mkNode(x_52, x_24, x_44); x_54 = lean_ctor_get(x_53, 4); lean_inc(x_54); -x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_41); +x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_41); lean_dec(x_54); if (x_55 == 0) { @@ -7256,7 +7256,7 @@ return x_56; lean_object* x_59; uint8_t x_60; x_59 = lean_ctor_get(x_58, 4); lean_inc(x_59); -x_60 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_59, x_41); +x_60 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_59, x_41); lean_dec(x_59); if (x_60 == 0) { @@ -7268,7 +7268,7 @@ lean_dec(x_4); x_61 = l_Lean_Parser_ParserState_mkNode(x_58, x_24, x_44); x_62 = lean_ctor_get(x_61, 4); lean_inc(x_62); -x_63 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_62, x_41); +x_63 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_62, x_41); lean_dec(x_62); if (x_63 == 0) { @@ -7290,7 +7290,7 @@ lean_inc(x_1); x_65 = lean_apply_2(x_4, x_1, x_58); x_66 = lean_ctor_get(x_65, 4); lean_inc(x_66); -x_67 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_66, x_41); +x_67 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_66, x_41); lean_dec(x_66); if (x_67 == 0) { @@ -7301,7 +7301,7 @@ lean_dec(x_12); x_68 = l_Lean_Parser_ParserState_mkNode(x_65, x_24, x_44); x_69 = lean_ctor_get(x_68, 4); lean_inc(x_69); -x_70 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_69, x_41); +x_70 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_69, x_41); lean_dec(x_69); if (x_70 == 0) { @@ -7346,7 +7346,7 @@ goto block_97; lean_object* x_79; uint8_t x_80; x_79 = lean_ctor_get(x_78, 4); lean_inc(x_79); -x_80 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_79, x_41); +x_80 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_79, x_41); lean_dec(x_79); if (x_80 == 0) { @@ -7363,7 +7363,7 @@ lean_inc(x_1); x_81 = l_Lean_Parser_strLit___elambda__1(x_1, x_78); x_82 = lean_ctor_get(x_81, 4); lean_inc(x_82); -x_83 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_82, x_41); +x_83 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_82, x_41); lean_dec(x_82); if (x_83 == 0) { @@ -7380,7 +7380,7 @@ lean_inc(x_1); x_84 = lean_apply_2(x_16, x_1, x_81); x_85 = lean_ctor_get(x_84, 4); lean_inc(x_85); -x_86 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_85, x_41); +x_86 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_85, x_41); lean_dec(x_85); if (x_86 == 0) { @@ -7396,7 +7396,7 @@ lean_inc(x_1); x_87 = lean_apply_2(x_12, x_1, x_84); x_88 = lean_ctor_get(x_87, 4); lean_inc(x_88); -x_89 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_88, x_41); +x_89 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_88, x_41); lean_dec(x_88); if (x_89 == 0) { @@ -8346,7 +8346,7 @@ x_18 = l_Lean_Parser_checkPrecFn(x_17, x_1, x_2); x_19 = lean_ctor_get(x_18, 4); lean_inc(x_19); x_20 = lean_box(0); -x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_20); +x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_20); lean_dec(x_19); if (x_21 == 0) { @@ -8394,7 +8394,7 @@ x_32 = l_Lean_Parser_Syntax_sepBy1___elambda__1___closed__2; x_33 = l_Lean_Parser_ParserState_mkNode(x_31, x_32, x_23); x_34 = lean_ctor_get(x_33, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_20); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_20); lean_dec(x_34); if (x_35 == 0) { @@ -8414,7 +8414,7 @@ return x_36; lean_object* x_39; uint8_t x_40; x_39 = lean_ctor_get(x_38, 4); lean_inc(x_39); -x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_20); +x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_20); lean_dec(x_39); if (x_40 == 0) { @@ -8427,7 +8427,7 @@ x_41 = l_Lean_Parser_Syntax_sepBy1___elambda__1___closed__2; x_42 = l_Lean_Parser_ParserState_mkNode(x_38, x_41, x_23); x_43 = lean_ctor_get(x_42, 4); lean_inc(x_43); -x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_20); +x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_20); lean_dec(x_43); if (x_44 == 0) { @@ -8449,7 +8449,7 @@ lean_inc(x_1); x_46 = lean_apply_2(x_4, x_1, x_38); x_47 = lean_ctor_get(x_46, 4); lean_inc(x_47); -x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_20); +x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_20); lean_dec(x_47); if (x_48 == 0) { @@ -8461,7 +8461,7 @@ x_49 = l_Lean_Parser_Syntax_sepBy1___elambda__1___closed__2; x_50 = l_Lean_Parser_ParserState_mkNode(x_46, x_49, x_23); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_20); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_20); lean_dec(x_51); if (x_52 == 0) { @@ -8506,7 +8506,7 @@ goto block_79; lean_object* x_61; uint8_t x_62; x_61 = lean_ctor_get(x_60, 4); lean_inc(x_61); -x_62 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_61, x_20); +x_62 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_61, x_20); lean_dec(x_61); if (x_62 == 0) { @@ -8523,7 +8523,7 @@ lean_inc(x_1); x_63 = l_Lean_Parser_strLit___elambda__1(x_1, x_60); x_64 = lean_ctor_get(x_63, 4); lean_inc(x_64); -x_65 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_64, x_20); +x_65 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_64, x_20); lean_dec(x_64); if (x_65 == 0) { @@ -8540,7 +8540,7 @@ lean_inc(x_1); x_66 = lean_apply_2(x_8, x_1, x_63); x_67 = lean_ctor_get(x_66, 4); lean_inc(x_67); -x_68 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_67, x_20); +x_68 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_67, x_20); lean_dec(x_67); if (x_68 == 0) { @@ -8556,7 +8556,7 @@ lean_inc(x_1); x_69 = lean_apply_2(x_6, x_1, x_66); x_70 = lean_ctor_get(x_69, 4); lean_inc(x_70); -x_71 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_70, x_20); +x_71 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_70, x_20); lean_dec(x_70); if (x_71 == 0) { @@ -9158,7 +9158,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -9178,7 +9178,7 @@ x_19 = l_Lean_Parser_Syntax_atom___elambda__1___closed__2; x_20 = l_Lean_Parser_ParserState_mkNode(x_18, x_19, x_17); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); -x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_14); +x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_14); lean_dec(x_21); if (x_22 == 0) { @@ -9687,7 +9687,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -9718,7 +9718,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_20, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -9727,7 +9727,7 @@ x_27 = l_Lean_Parser_Syntax_nonReserved___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_20, x_27, x_17); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -9751,7 +9751,7 @@ x_33 = l_Lean_Parser_Syntax_nonReserved___elambda__1___closed__2; x_34 = l_Lean_Parser_ParserState_mkNode(x_32, x_33, x_17); x_35 = lean_ctor_get(x_34, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_14); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_14); lean_dec(x_35); if (x_36 == 0) { @@ -9774,7 +9774,7 @@ lean_inc(x_1); x_38 = l_Lean_Parser_tokenAntiquotFn(x_1, x_20); x_39 = lean_ctor_get(x_38, 4); lean_inc(x_39); -x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); +x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); lean_dec(x_39); if (x_40 == 0) { @@ -9783,7 +9783,7 @@ x_41 = l_Lean_Parser_Syntax_nonReserved___elambda__1___closed__2; x_42 = l_Lean_Parser_ParserState_mkNode(x_38, x_41, x_17); x_43 = lean_ctor_get(x_42, 4); lean_inc(x_43); -x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_14); +x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_14); lean_dec(x_43); if (x_44 == 0) { @@ -9807,7 +9807,7 @@ x_47 = l_Lean_Parser_Syntax_nonReserved___elambda__1___closed__2; x_48 = l_Lean_Parser_ParserState_mkNode(x_46, x_47, x_17); x_49 = lean_ctor_get(x_48, 4); lean_inc(x_49); -x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_14); +x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_14); lean_dec(x_49); if (x_50 == 0) { @@ -10455,7 +10455,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -10498,7 +10498,7 @@ goto block_60; lean_object* x_26; uint8_t x_27; x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); lean_dec(x_26); if (x_27 == 0) { @@ -10508,7 +10508,7 @@ x_28 = l_Lean_Parser_Term_stx_quot___elambda__1___closed__5; x_29 = l_Lean_Parser_ParserState_mkNode(x_25, x_28, x_17); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); lean_dec(x_30); if (x_31 == 0) { @@ -10532,7 +10532,7 @@ lean_inc(x_1); x_35 = l___private_Lean_Parser_Basic_0__Lean_Parser_addQuotDepthFn(x_33, x_34, x_1, x_25); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); lean_dec(x_36); if (x_37 == 0) { @@ -10542,7 +10542,7 @@ x_38 = l_Lean_Parser_Term_stx_quot___elambda__1___closed__5; x_39 = l_Lean_Parser_ParserState_mkNode(x_35, x_38, x_17); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); lean_dec(x_40); if (x_41 == 0) { @@ -10577,7 +10577,7 @@ x_49 = l_Lean_Parser_Term_stx_quot___elambda__1___closed__5; x_50 = l_Lean_Parser_ParserState_mkNode(x_45, x_49, x_17); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_14); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_14); lean_dec(x_51); if (x_52 == 0) { @@ -10601,7 +10601,7 @@ x_55 = l_Lean_Parser_Term_stx_quot___elambda__1___closed__5; x_56 = l_Lean_Parser_ParserState_mkNode(x_54, x_55, x_17); x_57 = lean_ctor_get(x_56, 4); lean_inc(x_57); -x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_14); +x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_14); lean_dec(x_57); if (x_58 == 0) { @@ -11288,7 +11288,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -11331,7 +11331,7 @@ goto block_60; lean_object* x_26; uint8_t x_27; x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); lean_dec(x_26); if (x_27 == 0) { @@ -11341,7 +11341,7 @@ x_28 = l_Lean_Parser_Term_prec_quot___elambda__1___closed__2; x_29 = l_Lean_Parser_ParserState_mkNode(x_25, x_28, x_17); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); lean_dec(x_30); if (x_31 == 0) { @@ -11365,7 +11365,7 @@ lean_inc(x_1); x_35 = l___private_Lean_Parser_Basic_0__Lean_Parser_addQuotDepthFn(x_33, x_34, x_1, x_25); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); lean_dec(x_36); if (x_37 == 0) { @@ -11375,7 +11375,7 @@ x_38 = l_Lean_Parser_Term_prec_quot___elambda__1___closed__2; x_39 = l_Lean_Parser_ParserState_mkNode(x_35, x_38, x_17); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); lean_dec(x_40); if (x_41 == 0) { @@ -11410,7 +11410,7 @@ x_49 = l_Lean_Parser_Term_prec_quot___elambda__1___closed__2; x_50 = l_Lean_Parser_ParserState_mkNode(x_45, x_49, x_17); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_14); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_14); lean_dec(x_51); if (x_52 == 0) { @@ -11434,7 +11434,7 @@ x_55 = l_Lean_Parser_Term_prec_quot___elambda__1___closed__2; x_56 = l_Lean_Parser_ParserState_mkNode(x_54, x_55, x_17); x_57 = lean_ctor_get(x_56, 4); lean_inc(x_57); -x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_14); +x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_14); lean_dec(x_57); if (x_58 == 0) { @@ -12141,7 +12141,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -12184,7 +12184,7 @@ goto block_60; lean_object* x_26; uint8_t x_27; x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); lean_dec(x_26); if (x_27 == 0) { @@ -12194,7 +12194,7 @@ x_28 = l_Lean_Parser_Term_prio_quot___elambda__1___closed__3; x_29 = l_Lean_Parser_ParserState_mkNode(x_25, x_28, x_17); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); lean_dec(x_30); if (x_31 == 0) { @@ -12218,7 +12218,7 @@ lean_inc(x_1); x_35 = l___private_Lean_Parser_Basic_0__Lean_Parser_addQuotDepthFn(x_33, x_34, x_1, x_25); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); lean_dec(x_36); if (x_37 == 0) { @@ -12228,7 +12228,7 @@ x_38 = l_Lean_Parser_Term_prio_quot___elambda__1___closed__3; x_39 = l_Lean_Parser_ParserState_mkNode(x_35, x_38, x_17); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); lean_dec(x_40); if (x_41 == 0) { @@ -12263,7 +12263,7 @@ x_49 = l_Lean_Parser_Term_prio_quot___elambda__1___closed__3; x_50 = l_Lean_Parser_ParserState_mkNode(x_45, x_49, x_17); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_14); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_14); lean_dec(x_51); if (x_52 == 0) { @@ -12287,7 +12287,7 @@ x_55 = l_Lean_Parser_Term_prio_quot___elambda__1___closed__3; x_56 = l_Lean_Parser_ParserState_mkNode(x_54, x_55, x_17); x_57 = lean_ctor_get(x_56, 4); lean_inc(x_57); -x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_14); +x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_14); lean_dec(x_57); if (x_58 == 0) { @@ -13037,7 +13037,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -13057,7 +13057,7 @@ lean_inc(x_1); x_19 = l_Lean_Parser_atomicFn(x_18, x_1, x_12); x_20 = lean_ctor_get(x_19, 4); lean_inc(x_20); -x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_20, x_14); +x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_20, x_14); lean_dec(x_20); if (x_21 == 0) { @@ -13067,7 +13067,7 @@ x_22 = l_Lean_Parser_Command_namedName___elambda__1___closed__4; x_23 = l_Lean_Parser_ParserState_mkNode(x_19, x_22, x_17); x_24 = lean_ctor_get(x_23, 4); lean_inc(x_24); -x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_14); +x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_14); lean_dec(x_24); if (x_25 == 0) { @@ -13113,7 +13113,7 @@ goto block_67; lean_object* x_35; uint8_t x_36; x_35 = lean_ctor_get(x_34, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_14); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_14); lean_dec(x_35); if (x_36 == 0) { @@ -13123,7 +13123,7 @@ x_37 = l_Lean_Parser_Command_namedName___elambda__1___closed__4; x_38 = l_Lean_Parser_ParserState_mkNode(x_34, x_37, x_17); x_39 = lean_ctor_get(x_38, 4); lean_inc(x_39); -x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); +x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); lean_dec(x_39); if (x_40 == 0) { @@ -13145,7 +13145,7 @@ lean_inc(x_1); x_42 = l_Lean_Parser_ident___elambda__1(x_1, x_34); x_43 = lean_ctor_get(x_42, 4); lean_inc(x_43); -x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_14); +x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_14); lean_dec(x_43); if (x_44 == 0) { @@ -13155,7 +13155,7 @@ x_45 = l_Lean_Parser_Command_namedName___elambda__1___closed__4; x_46 = l_Lean_Parser_ParserState_mkNode(x_42, x_45, x_17); x_47 = lean_ctor_get(x_46, 4); lean_inc(x_47); -x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_14); +x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_14); lean_dec(x_47); if (x_48 == 0) { @@ -13190,7 +13190,7 @@ x_56 = l_Lean_Parser_Command_namedName___elambda__1___closed__4; x_57 = l_Lean_Parser_ParserState_mkNode(x_52, x_56, x_17); x_58 = lean_ctor_get(x_57, 4); lean_inc(x_58); -x_59 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_58, x_14); +x_59 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_58, x_14); lean_dec(x_58); if (x_59 == 0) { @@ -13214,7 +13214,7 @@ x_62 = l_Lean_Parser_Command_namedName___elambda__1___closed__4; x_63 = l_Lean_Parser_ParserState_mkNode(x_61, x_62, x_17); x_64 = lean_ctor_get(x_63, 4); lean_inc(x_64); -x_65 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_64, x_14); +x_65 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_64, x_14); lean_dec(x_64); if (x_65 == 0) { @@ -13526,7 +13526,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -13559,7 +13559,7 @@ x_25 = l_Lean_Parser_Command_prefix___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_20, x_25, x_17); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); lean_dec(x_27); if (x_28 == 0) { @@ -13583,7 +13583,7 @@ x_31 = l_Lean_Parser_Command_prefix___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_17); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -13822,7 +13822,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -13855,7 +13855,7 @@ x_25 = l_Lean_Parser_Command_infix___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_20, x_25, x_17); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); lean_dec(x_27); if (x_28 == 0) { @@ -13879,7 +13879,7 @@ x_31 = l_Lean_Parser_Command_infix___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_17); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -14118,7 +14118,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -14151,7 +14151,7 @@ x_25 = l_Lean_Parser_Command_infixl___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_20, x_25, x_17); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); lean_dec(x_27); if (x_28 == 0) { @@ -14175,7 +14175,7 @@ x_31 = l_Lean_Parser_Command_infixl___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_17); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -14414,7 +14414,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -14447,7 +14447,7 @@ x_25 = l_Lean_Parser_Command_infixr___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_20, x_25, x_17); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); lean_dec(x_27); if (x_28 == 0) { @@ -14471,7 +14471,7 @@ x_31 = l_Lean_Parser_Command_infixr___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_17); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -14710,7 +14710,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -14743,7 +14743,7 @@ x_25 = l_Lean_Parser_Command_postfix___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_20, x_25, x_17); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); lean_dec(x_27); if (x_28 == 0) { @@ -14767,7 +14767,7 @@ x_31 = l_Lean_Parser_Command_postfix___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_17); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -15242,7 +15242,7 @@ x_16 = l_Lean_Parser_checkPrecFn(x_15, x_1, x_2); x_17 = lean_ctor_get(x_16, 4); lean_inc(x_17); x_18 = lean_box(0); -x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); +x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); lean_dec(x_17); if (x_19 == 0) { @@ -15262,7 +15262,7 @@ lean_inc(x_1); x_22 = l_Lean_Parser_Term_attrKind___elambda__1(x_1, x_16); x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); -x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_18); +x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_18); lean_dec(x_23); if (x_24 == 0) { @@ -15278,7 +15278,7 @@ lean_inc(x_1); x_32 = l_Lean_Parser_Command_mixfixKind___elambda__1(x_1, x_22); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_18); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_18); lean_dec(x_33); if (x_34 == 0) { @@ -15294,7 +15294,7 @@ lean_inc(x_1); x_35 = l_Lean_Parser_precedence___elambda__1(x_1, x_32); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_18); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_18); lean_dec(x_36); if (x_37 == 0) { @@ -15310,7 +15310,7 @@ lean_inc(x_1); x_38 = lean_apply_2(x_6, x_1, x_35); x_39 = lean_ctor_get(x_38, 4); lean_inc(x_39); -x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_18); +x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_18); lean_dec(x_39); if (x_40 == 0) { @@ -15325,7 +15325,7 @@ lean_inc(x_1); x_41 = lean_apply_2(x_4, x_1, x_38); x_42 = lean_ctor_get(x_41, 4); lean_inc(x_42); -x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_18); +x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_18); lean_dec(x_42); if (x_43 == 0) { @@ -15339,7 +15339,7 @@ lean_inc(x_1); x_44 = l_Lean_Parser_strLit___elambda__1(x_1, x_41); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_18); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_18); lean_dec(x_45); if (x_46 == 0) { @@ -15353,7 +15353,7 @@ lean_inc(x_1); x_47 = l_Lean_Parser_darrow___elambda__1(x_1, x_44); x_48 = lean_ctor_get(x_47, 4); lean_inc(x_48); -x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_18); +x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_18); lean_dec(x_48); if (x_49 == 0) { @@ -15383,7 +15383,7 @@ x_26 = l_Lean_Parser_Command_mixfix___elambda__1___closed__2; x_27 = l_Lean_Parser_ParserState_mkNode(x_25, x_26, x_21); x_28 = lean_ctor_get(x_27, 4); lean_inc(x_28); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_18); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_18); lean_dec(x_28); if (x_29 == 0) { @@ -17514,7 +17514,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -17533,7 +17533,7 @@ lean_inc(x_1); x_20 = l_Lean_Parser_ident___elambda__1(x_1, x_14); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); -x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_16); +x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_16); lean_dec(x_21); if (x_22 == 0) { @@ -17543,7 +17543,7 @@ x_23 = l_Lean_Parser_Command_identPrec___elambda__1___closed__2; x_24 = l_Lean_Parser_ParserState_mkNode(x_20, x_23, x_19); x_25 = lean_ctor_get(x_24, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_16); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_16); lean_dec(x_25); if (x_26 == 0) { @@ -17567,7 +17567,7 @@ x_29 = l_Lean_Parser_Command_identPrec___elambda__1___closed__2; x_30 = l_Lean_Parser_ParserState_mkNode(x_28, x_29, x_19); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); lean_dec(x_31); if (x_32 == 0) { @@ -17885,7 +17885,7 @@ lean_inc(x_4); x_5 = lean_ctor_get(x_2, 4); lean_inc(x_5); x_6 = lean_box(0); -x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); +x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); lean_dec(x_5); if (x_7 == 0) { @@ -18226,7 +18226,7 @@ x_20 = l_Lean_Parser_checkPrecFn(x_19, x_1, x_2); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); x_22 = lean_box(0); -x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_22); +x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_22); lean_dec(x_21); if (x_23 == 0) { @@ -18249,7 +18249,7 @@ lean_inc(x_1); x_26 = l_Lean_Parser_Term_attrKind___elambda__1(x_1, x_20); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_22); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_22); lean_dec(x_27); if (x_28 == 0) { @@ -18263,7 +18263,7 @@ x_36 = l_Lean_Parser_Command_notation___elambda__1___closed__2; x_37 = l_Lean_Parser_ParserState_mkNode(x_26, x_36, x_25); x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_22); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_22); lean_dec(x_38); if (x_39 == 0) { @@ -18310,7 +18310,7 @@ goto block_69; lean_object* x_49; uint8_t x_50; x_49 = lean_ctor_get(x_48, 4); lean_inc(x_49); -x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_22); +x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_22); lean_dec(x_49); if (x_50 == 0) { @@ -18328,7 +18328,7 @@ lean_inc(x_1); x_51 = lean_apply_2(x_10, x_1, x_48); x_52 = lean_ctor_get(x_51, 4); lean_inc(x_52); -x_53 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_52, x_22); +x_53 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_52, x_22); lean_dec(x_52); if (x_53 == 0) { @@ -18345,7 +18345,7 @@ lean_inc(x_1); x_54 = lean_apply_2(x_8, x_1, x_51); x_55 = lean_ctor_get(x_54, 4); lean_inc(x_55); -x_56 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_55, x_22); +x_56 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_55, x_22); lean_dec(x_55); if (x_56 == 0) { @@ -18361,7 +18361,7 @@ lean_inc(x_1); x_57 = lean_apply_2(x_6, x_1, x_54); x_58 = lean_ctor_get(x_57, 4); lean_inc(x_58); -x_59 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_58, x_22); +x_59 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_58, x_22); lean_dec(x_58); if (x_59 == 0) { @@ -18376,7 +18376,7 @@ lean_inc(x_1); x_60 = lean_apply_2(x_4, x_1, x_57); x_61 = lean_ctor_get(x_60, 4); lean_inc(x_61); -x_62 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_61, x_22); +x_62 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_61, x_22); lean_dec(x_61); if (x_62 == 0) { @@ -18390,7 +18390,7 @@ lean_inc(x_1); x_63 = l_Lean_Parser_darrow___elambda__1(x_1, x_60); x_64 = lean_ctor_get(x_63, 4); lean_inc(x_64); -x_65 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_64, x_22); +x_65 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_64, x_22); lean_dec(x_64); if (x_65 == 0) { @@ -18421,7 +18421,7 @@ x_30 = l_Lean_Parser_Command_notation___elambda__1___closed__2; x_31 = l_Lean_Parser_ParserState_mkNode(x_29, x_30, x_25); x_32 = lean_ctor_get(x_31, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_22); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_22); lean_dec(x_32); if (x_33 == 0) { @@ -20466,7 +20466,7 @@ x_22 = l_Lean_Parser_checkPrecFn(x_21, x_1, x_2); x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); x_24 = lean_box(0); -x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_24); +x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_24); lean_dec(x_23); if (x_25 == 0) { @@ -20490,7 +20490,7 @@ lean_inc(x_1); x_28 = lean_apply_2(x_12, x_1, x_22); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_24); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_24); lean_dec(x_29); if (x_30 == 0) { @@ -20504,7 +20504,7 @@ x_38 = l_Lean_Parser_Command_syntax___elambda__1___closed__2; x_39 = l_Lean_Parser_ParserState_mkNode(x_28, x_38, x_27); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_24); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_24); lean_dec(x_40); if (x_41 == 0) { @@ -20526,7 +20526,7 @@ lean_inc(x_1); x_43 = l_Lean_Parser_Term_attrKind___elambda__1(x_1, x_28); x_44 = lean_ctor_get(x_43, 4); lean_inc(x_44); -x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_24); +x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_24); lean_dec(x_44); if (x_45 == 0) { @@ -20540,7 +20540,7 @@ x_46 = l_Lean_Parser_Command_syntax___elambda__1___closed__2; x_47 = l_Lean_Parser_ParserState_mkNode(x_43, x_46, x_27); x_48 = lean_ctor_get(x_47, 4); lean_inc(x_48); -x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_24); +x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_24); lean_dec(x_48); if (x_49 == 0) { @@ -20586,7 +20586,7 @@ goto block_86; lean_object* x_59; uint8_t x_60; x_59 = lean_ctor_get(x_58, 4); lean_inc(x_59); -x_60 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_59, x_24); +x_60 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_59, x_24); lean_dec(x_59); if (x_60 == 0) { @@ -20605,7 +20605,7 @@ lean_inc(x_1); x_61 = lean_apply_2(x_10, x_1, x_58); x_62 = lean_ctor_get(x_61, 4); lean_inc(x_62); -x_63 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_62, x_24); +x_63 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_62, x_24); lean_dec(x_62); if (x_63 == 0) { @@ -20623,7 +20623,7 @@ lean_inc(x_1); x_64 = lean_apply_2(x_8, x_1, x_61); x_65 = lean_ctor_get(x_64, 4); lean_inc(x_65); -x_66 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_65, x_24); +x_66 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_65, x_24); lean_dec(x_65); if (x_66 == 0) { @@ -20640,7 +20640,7 @@ lean_inc(x_1); x_67 = lean_apply_2(x_6, x_1, x_64); x_68 = lean_ctor_get(x_67, 4); lean_inc(x_68); -x_69 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_68, x_24); +x_69 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_68, x_24); lean_dec(x_68); if (x_69 == 0) { @@ -20656,7 +20656,7 @@ lean_inc(x_1); x_70 = lean_apply_2(x_4, x_1, x_67); x_71 = lean_ctor_get(x_70, 4); lean_inc(x_71); -x_72 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_71, x_24); +x_72 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_71, x_24); lean_dec(x_71); if (x_72 == 0) { @@ -20682,7 +20682,7 @@ if (x_78 == 0) lean_object* x_79; uint8_t x_80; x_79 = lean_ctor_get(x_75, 4); lean_inc(x_79); -x_80 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_79, x_24); +x_80 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_79, x_24); lean_dec(x_79); if (x_80 == 0) { @@ -20705,7 +20705,7 @@ lean_inc(x_1); x_82 = l_Lean_Parser_tokenAntiquotFn(x_1, x_75); x_83 = lean_ctor_get(x_82, 4); lean_inc(x_83); -x_84 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_83, x_24); +x_84 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_83, x_24); lean_dec(x_83); if (x_84 == 0) { @@ -20736,7 +20736,7 @@ x_32 = l_Lean_Parser_Command_syntax___elambda__1___closed__2; x_33 = l_Lean_Parser_ParserState_mkNode(x_31, x_32, x_27); x_34 = lean_ctor_get(x_33, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_24); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_24); lean_dec(x_34); if (x_35 == 0) { @@ -21613,7 +21613,7 @@ x_16 = l_Lean_Parser_checkPrecFn(x_15, x_1, x_2); x_17 = lean_ctor_get(x_16, 4); lean_inc(x_17); x_18 = lean_box(0); -x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); +x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); lean_dec(x_17); if (x_19 == 0) { @@ -21634,7 +21634,7 @@ lean_inc(x_1); x_22 = lean_apply_2(x_6, x_1, x_16); x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); -x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_18); +x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_18); lean_dec(x_23); if (x_24 == 0) { @@ -21645,7 +21645,7 @@ x_32 = l_Lean_Parser_Command_syntaxAbbrev___elambda__1___closed__2; x_33 = l_Lean_Parser_ParserState_mkNode(x_22, x_32, x_21); x_34 = lean_ctor_get(x_33, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_18); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_18); lean_dec(x_34); if (x_35 == 0) { @@ -21691,7 +21691,7 @@ goto block_63; lean_object* x_45; uint8_t x_46; x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_18); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_18); lean_dec(x_45); if (x_46 == 0) { @@ -21707,7 +21707,7 @@ lean_inc(x_1); x_47 = l_Lean_Parser_ident___elambda__1(x_1, x_44); x_48 = lean_ctor_get(x_47, 4); lean_inc(x_48); -x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_18); +x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_18); lean_dec(x_48); if (x_49 == 0) { @@ -21734,7 +21734,7 @@ if (x_55 == 0) lean_object* x_56; uint8_t x_57; x_56 = lean_ctor_get(x_52, 4); lean_inc(x_56); -x_57 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_56, x_18); +x_57 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_56, x_18); lean_dec(x_56); if (x_57 == 0) { @@ -21758,7 +21758,7 @@ lean_inc(x_1); x_59 = l_Lean_Parser_tokenAntiquotFn(x_1, x_52); x_60 = lean_ctor_get(x_59, 4); lean_inc(x_60); -x_61 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_60, x_18); +x_61 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_60, x_18); lean_dec(x_60); if (x_61 == 0) { @@ -21786,7 +21786,7 @@ x_26 = l_Lean_Parser_Command_syntaxAbbrev___elambda__1___closed__2; x_27 = l_Lean_Parser_ParserState_mkNode(x_25, x_26, x_21); x_28 = lean_ctor_get(x_27, 4); lean_inc(x_28); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_18); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_18); lean_dec(x_28); if (x_29 == 0) { @@ -22427,7 +22427,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -22460,7 +22460,7 @@ x_25 = l_Lean_Parser_Command_catBehaviorBoth___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_20, x_25, x_17); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); lean_dec(x_27); if (x_28 == 0) { @@ -22484,7 +22484,7 @@ x_31 = l_Lean_Parser_Command_catBehaviorBoth___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_17); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -22732,7 +22732,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -22765,7 +22765,7 @@ x_25 = l_Lean_Parser_Command_catBehaviorSymbol___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_20, x_25, x_17); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); lean_dec(x_27); if (x_28 == 0) { @@ -22789,7 +22789,7 @@ x_31 = l_Lean_Parser_Command_catBehaviorSymbol___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_17); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -23246,7 +23246,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -23291,7 +23291,7 @@ goto block_49; lean_object* x_28; uint8_t x_29; x_28 = lean_ctor_get(x_27, 4); lean_inc(x_28); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_16); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_16); lean_dec(x_28); if (x_29 == 0) { @@ -23301,7 +23301,7 @@ x_30 = l_Lean_Parser_Command_syntaxCat___elambda__1___closed__2; x_31 = l_Lean_Parser_ParserState_mkNode(x_27, x_30, x_19); x_32 = lean_ctor_get(x_31, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_16); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_16); lean_dec(x_32); if (x_33 == 0) { @@ -23323,7 +23323,7 @@ lean_inc(x_1); x_35 = l_Lean_Parser_ident___elambda__1(x_1, x_27); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_16); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_16); lean_dec(x_36); if (x_37 == 0) { @@ -23333,7 +23333,7 @@ x_38 = l_Lean_Parser_Command_syntaxCat___elambda__1___closed__2; x_39 = l_Lean_Parser_ParserState_mkNode(x_35, x_38, x_19); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_16); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_16); lean_dec(x_40); if (x_41 == 0) { @@ -23357,7 +23357,7 @@ x_44 = l_Lean_Parser_Command_syntaxCat___elambda__1___closed__2; x_45 = l_Lean_Parser_ParserState_mkNode(x_43, x_44, x_19); x_46 = lean_ctor_get(x_45, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_16); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_16); lean_dec(x_46); if (x_47 == 0) { @@ -24566,7 +24566,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -24585,7 +24585,7 @@ lean_inc(x_1); x_20 = lean_apply_2(x_4, x_1, x_14); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); -x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_16); +x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_16); lean_dec(x_21); if (x_22 == 0) { @@ -24594,7 +24594,7 @@ x_23 = l_Lean_Parser_Command_macroArg___elambda__1___closed__2; x_24 = l_Lean_Parser_ParserState_mkNode(x_20, x_23, x_19); x_25 = lean_ctor_get(x_24, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_16); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_16); lean_dec(x_25); if (x_26 == 0) { @@ -24620,7 +24620,7 @@ x_31 = l_Lean_Parser_Command_macroArg___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_19); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_16); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_16); lean_dec(x_33); if (x_34 == 0) { @@ -25013,7 +25013,7 @@ x_6 = l_Lean_Parser_atomicFn(x_5, x_1, x_2); x_7 = lean_ctor_get(x_6, 4); lean_inc(x_7); x_8 = lean_box(0); -x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); +x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); lean_dec(x_7); if (x_9 == 0) { @@ -25028,7 +25028,7 @@ lean_inc(x_1); x_10 = l_Lean_Parser_darrow___elambda__1(x_1, x_6); x_11 = lean_ctor_get(x_10, 4); lean_inc(x_11); -x_12 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_11, x_8); +x_12 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_11, x_8); lean_dec(x_11); if (x_12 == 0) { @@ -25265,7 +25265,7 @@ x_6 = l_Lean_Parser_atomicFn(x_5, x_1, x_2); x_7 = lean_ctor_get(x_6, 4); lean_inc(x_7); x_8 = lean_box(0); -x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); +x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); lean_dec(x_7); if (x_9 == 0) { @@ -25280,7 +25280,7 @@ lean_inc(x_1); x_10 = l_Lean_Parser_darrow___elambda__1(x_1, x_6); x_11 = lean_ctor_get(x_10, 4); lean_inc(x_11); -x_12 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_11, x_8); +x_12 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_11, x_8); lean_dec(x_11); if (x_12 == 0) { @@ -25390,7 +25390,7 @@ x_6 = l_Lean_Parser_atomicFn(x_5, x_1, x_2); x_7 = lean_ctor_get(x_6, 4); lean_inc(x_7); x_8 = lean_box(0); -x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); +x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); lean_dec(x_7); if (x_9 == 0) { @@ -25405,7 +25405,7 @@ lean_inc(x_1); x_10 = l_Lean_Parser_darrow___elambda__1(x_1, x_6); x_11 = lean_ctor_get(x_10, 4); lean_inc(x_11); -x_12 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_11, x_8); +x_12 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_11, x_8); lean_dec(x_11); if (x_12 == 0) { @@ -25592,7 +25592,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -25615,7 +25615,7 @@ x_22 = l_Lean_Parser_Command_macroTail___elambda__1___closed__2; x_23 = l_Lean_Parser_ParserState_mkNode(x_21, x_22, x_17); x_24 = lean_ctor_get(x_23, 4); lean_inc(x_24); -x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_14); +x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_14); lean_dec(x_24); if (x_25 == 0) { @@ -25971,7 +25971,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -25992,7 +25992,7 @@ x_20 = l_Lean_Parser_Command_macro___elambda__1___closed__2; x_21 = l_Lean_Parser_ParserState_mkNode(x_19, x_20, x_17); x_22 = lean_ctor_get(x_21, 4); lean_inc(x_22); -x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_22, x_14); +x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_22, x_14); lean_dec(x_22); if (x_23 == 0) { @@ -27987,7 +27987,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -28008,7 +28008,7 @@ x_20 = l_Lean_Parser_Command_elab__rules___elambda__1___closed__2; x_21 = l_Lean_Parser_ParserState_mkNode(x_19, x_20, x_17); x_22 = lean_ctor_get(x_21, 4); lean_inc(x_22); -x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_22, x_14); +x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_22, x_14); lean_dec(x_22); if (x_23 == 0) { @@ -28927,7 +28927,7 @@ x_18 = l_Lean_Parser_checkPrecFn(x_17, x_1, x_2); x_19 = lean_ctor_get(x_18, 4); lean_inc(x_19); x_20 = lean_box(0); -x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_20); +x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_20); lean_dec(x_19); if (x_21 == 0) { @@ -28952,7 +28952,7 @@ lean_inc(x_1); x_25 = l_Lean_Parser_atomicFn(x_24, x_1, x_18); x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_20); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_20); lean_dec(x_26); if (x_27 == 0) { @@ -28967,7 +28967,7 @@ x_28 = l_Lean_Parser_Command_elabTail___elambda__1___closed__2; x_29 = l_Lean_Parser_ParserState_mkNode(x_25, x_28, x_23); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_20); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_20); lean_dec(x_30); if (x_31 == 0) { @@ -28989,7 +28989,7 @@ lean_inc(x_1); x_33 = l_Lean_Parser_darrow___elambda__1(x_1, x_25); x_34 = lean_ctor_get(x_33, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_20); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_20); lean_dec(x_34); if (x_35 == 0) { @@ -29004,7 +29004,7 @@ x_36 = l_Lean_Parser_Command_elabTail___elambda__1___closed__2; x_37 = l_Lean_Parser_ParserState_mkNode(x_33, x_36, x_23); x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_20); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_20); lean_dec(x_38); if (x_39 == 0) { @@ -29042,7 +29042,7 @@ x_47 = l_Lean_Parser_Command_elabTail___elambda__1___closed__2; x_48 = l_Lean_Parser_ParserState_mkNode(x_46, x_47, x_23); x_49 = lean_ctor_get(x_48, 4); lean_inc(x_49); -x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_20); +x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_20); lean_dec(x_49); if (x_50 == 0) { @@ -29428,7 +29428,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -29449,7 +29449,7 @@ x_20 = l_Lean_Parser_Command_elab___elambda__1___closed__2; x_21 = l_Lean_Parser_ParserState_mkNode(x_19, x_20, x_17); x_22 = lean_ctor_get(x_21, 4); lean_inc(x_22); -x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_22, x_14); +x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_22, x_14); lean_dec(x_22); if (x_23 == 0) { diff --git a/stage0/stdlib/Lean/Parser/Tactic.c b/stage0/stdlib/Lean/Parser/Tactic.c index 038f815fe5e..b86e6db4524 100644 --- a/stage0/stdlib/Lean/Parser/Tactic.c +++ b/stage0/stdlib/Lean/Parser/Tactic.c @@ -70,6 +70,7 @@ static lean_object* l_Lean_Parser_Tactic_unknown___elambda__1___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Tactic_nativeDecide_formatter(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_introMatch___elambda__1___lambda__1(lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Parser_Tactic_tacticSeqBracketed; +static lean_object* l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__19; static lean_object* l_Lean_Parser_Tactic_matchAlts_parenthesizer___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Tactic_match_declRange(lean_object*); static lean_object* l_Lean_Parser_Tactic_nativeDecide___elambda__1___closed__8; @@ -148,6 +149,7 @@ static lean_object* l_Lean_Parser_Tactic_decide_formatter___closed__1; static lean_object* l_Lean_Parser_Tactic_nativeDecide___closed__1; lean_object* l_Lean_Parser_Term_matchAlts_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Tactic_nestedTactic___closed__1; +uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Parser_registerAliasCore___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_introMatch_formatter___closed__1; static lean_object* l_Lean_Parser_Tactic_introMatch_parenthesizer___closed__3; @@ -256,7 +258,7 @@ lean_object* l_Lean_Parser_sepBy1_parenthesizer___boxed(lean_object*, lean_objec static lean_object* l_Lean_Parser_Tactic_match___elambda__1___closed__17; static lean_object* l_Lean_Parser_Tactic_unknown_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Tactic_match___closed__2; -lean_object* l_Lean_Parser_registerAlias(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Parser_registerAlias(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_Term_generalizingParam_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_unknown_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_unknown___closed__7; @@ -320,6 +322,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_match_formatter(lean_object*, lean LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_decide; static lean_object* l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__15; static lean_object* l___regBuiltin_Lean_Parser_Tactic_match_declRange___closed__7; +static lean_object* l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__18; lean_object* l_Lean_Parser_symbolInfo(lean_object*); static lean_object* l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__5; extern lean_object* l_Lean_Parser_Tactic_tacticSeq; @@ -398,7 +401,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_decide_formatter(lean_object*, lea static lean_object* l_Lean_Parser_Tactic_decide___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Tactic_decide_formatter___closed__2; static lean_object* l_Lean_Parser_Tactic_match___elambda__1___lambda__1___closed__1; -uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Parser_ident_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_unknown___elambda__1___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_introMatch___elambda__1(lean_object*, lean_object*); @@ -533,22 +535,44 @@ return x_2; static lean_object* _init_l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__12() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__13() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__12; +x_2 = 1; +x_3 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__14() { +_start: +{ lean_object* x_1; x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Tactic_tacticSeq_formatter), 5, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__13() { +static lean_object* _init_l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__15() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__12; +x_1 = l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__14; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__14() { +static lean_object* _init_l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__16() { _start: { lean_object* x_1; @@ -556,7 +580,7 @@ x_1 = l_Lean_PrettyPrinter_Formatter_formatterAliasesRef; return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__15() { +static lean_object* _init_l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__17() { _start: { lean_object* x_1; @@ -564,17 +588,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Tactic_tacticSeq_parenthesizer), return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__16() { +static lean_object* _init_l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__18() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__15; +x_1 = l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__17; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__17() { +static lean_object* _init_l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__19() { _start: { lean_object* x_1; @@ -585,74 +609,75 @@ return x_1; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6_(lean_object* x_1) { _start: { -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_2 = l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__2; x_3 = l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__3; x_4 = l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__11; -x_5 = l_Lean_Parser_registerAlias(x_2, x_3, x_4, x_1); -if (lean_obj_tag(x_5) == 0) +x_5 = l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__13; +x_6 = l_Lean_Parser_registerAlias(x_2, x_3, x_4, x_5, x_1); +if (lean_obj_tag(x_6) == 0) { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_6 = lean_ctor_get(x_5, 1); -lean_inc(x_6); -lean_dec(x_5); -x_7 = l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__14; -x_8 = l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__13; -x_9 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_2, x_8, x_6); -if (lean_obj_tag(x_9) == 0) +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_7 = lean_ctor_get(x_6, 1); +lean_inc(x_7); +lean_dec(x_6); +x_8 = l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__16; +x_9 = l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__15; +x_10 = l_Lean_Parser_registerAliasCore___rarg(x_8, x_2, x_9, x_7); +if (lean_obj_tag(x_10) == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_10 = lean_ctor_get(x_9, 1); -lean_inc(x_10); -lean_dec(x_9); -x_11 = l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__17; -x_12 = l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__16; -x_13 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_2, x_12, x_10); -return x_13; +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +lean_dec(x_10); +x_12 = l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__19; +x_13 = l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__18; +x_14 = l_Lean_Parser_registerAliasCore___rarg(x_12, x_2, x_13, x_11); +return x_14; } else { -uint8_t x_14; -x_14 = !lean_is_exclusive(x_9); -if (x_14 == 0) +uint8_t x_15; +x_15 = !lean_is_exclusive(x_10); +if (x_15 == 0) { -return x_9; +return x_10; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_9, 0); -x_16 = lean_ctor_get(x_9, 1); +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_10, 0); +x_17 = lean_ctor_get(x_10, 1); +lean_inc(x_17); lean_inc(x_16); -lean_inc(x_15); -lean_dec(x_9); -x_17 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_17, 0, x_15); -lean_ctor_set(x_17, 1, x_16); -return x_17; +lean_dec(x_10); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +return x_18; } } } else { -uint8_t x_18; -x_18 = !lean_is_exclusive(x_5); -if (x_18 == 0) +uint8_t x_19; +x_19 = !lean_is_exclusive(x_6); +if (x_19 == 0) { -return x_5; +return x_6; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_5, 0); -x_20 = lean_ctor_get(x_5, 1); +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_6, 0); +x_21 = lean_ctor_get(x_6, 1); +lean_inc(x_21); lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_5); -x_21 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -return x_21; +lean_dec(x_6); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; } } } @@ -685,7 +710,7 @@ x_7 = l_Lean_Parser_ident___elambda__1(x_1, x_2); x_8 = lean_ctor_get(x_7, 4); lean_inc(x_8); x_9 = lean_box(0); -x_10 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_8, x_9); +x_10 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_8, x_9); lean_dec(x_8); if (x_10 == 0) { @@ -736,7 +761,7 @@ x_24 = l_Lean_Parser_ident___elambda__1(x_23, x_2); x_25 = lean_ctor_get(x_24, 4); lean_inc(x_25); x_26 = lean_box(0); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_26); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_26); lean_dec(x_25); if (x_27 == 0) { @@ -886,7 +911,7 @@ x_18 = l_Lean_Parser_checkPrecFn(x_17, x_1, x_2); x_19 = lean_ctor_get(x_18, 4); lean_inc(x_19); x_20 = lean_box(0); -x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_20); +x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_20); lean_dec(x_19); if (x_21 == 0) { @@ -923,7 +948,7 @@ lean_inc(x_26); x_27 = l_Lean_Parser_ident___elambda__1(x_26, x_18); x_28 = lean_ctor_get(x_27, 4); lean_inc(x_28); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_20); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_20); lean_dec(x_28); if (x_29 == 0) { @@ -933,7 +958,7 @@ x_30 = l_Lean_Parser_Tactic_unknown___elambda__1___closed__2; x_31 = l_Lean_Parser_ParserState_mkNode(x_27, x_30, x_23); x_32 = lean_ctor_get(x_31, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_20); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_20); lean_dec(x_32); if (x_33 == 0) { @@ -958,7 +983,7 @@ x_38 = l_Lean_Parser_Tactic_unknown___elambda__1___closed__2; x_39 = l_Lean_Parser_ParserState_mkNode(x_37, x_38, x_23); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_20); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_20); lean_dec(x_40); if (x_41 == 0) { @@ -2062,7 +2087,7 @@ x_42 = l_Lean_Parser_checkPrecFn(x_41, x_1, x_2); x_43 = lean_ctor_get(x_42, 4); lean_inc(x_43); x_44 = lean_box(0); -x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_44); +x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_44); lean_dec(x_43); if (x_45 == 0) { @@ -2109,7 +2134,7 @@ lean_object* x_55; lean_object* x_56; uint8_t x_57; x_55 = l_Lean_Parser_ParserState_mkNode(x_54, x_27, x_47); x_56 = lean_ctor_get(x_55, 4); lean_inc(x_56); -x_57 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_56, x_44); +x_57 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_56, x_44); lean_dec(x_56); if (x_57 == 0) { @@ -2129,7 +2154,7 @@ return x_58; lean_object* x_61; uint8_t x_62; x_61 = lean_ctor_get(x_60, 4); lean_inc(x_61); -x_62 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_61, x_44); +x_62 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_61, x_44); lean_dec(x_61); if (x_62 == 0) { @@ -2148,7 +2173,7 @@ lean_inc(x_1); x_63 = lean_apply_2(x_24, x_1, x_60); x_64 = lean_ctor_get(x_63, 4); lean_inc(x_64); -x_65 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_64, x_44); +x_65 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_64, x_44); lean_dec(x_64); if (x_65 == 0) { @@ -2166,7 +2191,7 @@ lean_inc(x_1); x_66 = lean_apply_2(x_21, x_1, x_63); x_67 = lean_ctor_get(x_66, 4); lean_inc(x_67); -x_68 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_67, x_44); +x_68 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_67, x_44); lean_dec(x_67); if (x_68 == 0) { @@ -2183,7 +2208,7 @@ lean_inc(x_1); x_69 = lean_apply_2(x_18, x_1, x_66); x_70 = lean_ctor_get(x_69, 4); lean_inc(x_70); -x_71 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_70, x_44); +x_71 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_70, x_44); lean_dec(x_70); if (x_71 == 0) { @@ -2209,7 +2234,7 @@ if (x_76 == 0) lean_object* x_77; uint8_t x_78; x_77 = lean_ctor_get(x_73, 4); lean_inc(x_77); -x_78 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_77, x_44); +x_78 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_77, x_44); lean_dec(x_77); if (x_78 == 0) { @@ -2233,7 +2258,7 @@ lean_inc(x_1); x_80 = l_Lean_Parser_tokenAntiquotFn(x_1, x_73); x_81 = lean_ctor_get(x_80, 4); lean_inc(x_81); -x_82 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_81, x_44); +x_82 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_81, x_44); lean_dec(x_81); if (x_82 == 0) { @@ -2491,7 +2516,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_matchRhs_formatter___closed__2() { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_matchRhs_formatter___closed__1; -x_2 = l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__12; +x_2 = l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__14; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_orelse_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); @@ -2787,7 +2812,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_matchRhs_parenthesizer___closed__ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_matchRhs_parenthesizer___closed__1; -x_2 = l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__15; +x_2 = l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__17; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_orelse_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); @@ -3221,7 +3246,7 @@ x_23 = l_Lean_Parser_checkPrecFn(x_22, x_1, x_2); x_24 = lean_ctor_get(x_23, 4); lean_inc(x_24); x_25 = lean_box(0); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_25); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_25); lean_dec(x_24); if (x_26 == 0) { @@ -3252,7 +3277,7 @@ if (x_34 == 0) lean_object* x_35; uint8_t x_36; x_35 = lean_ctor_get(x_30, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_25); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_25); lean_dec(x_35); if (x_36 == 0) { @@ -3261,7 +3286,7 @@ lean_dec(x_6); x_37 = l_Lean_Parser_ParserState_mkNode(x_30, x_8, x_28); x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_25); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_25); lean_dec(x_38); if (x_39 == 0) { @@ -3284,7 +3309,7 @@ x_41 = lean_apply_2(x_6, x_1, x_30); x_42 = l_Lean_Parser_ParserState_mkNode(x_41, x_8, x_28); x_43 = lean_ctor_get(x_42, 4); lean_inc(x_43); -x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_25); +x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_25); lean_dec(x_43); if (x_44 == 0) { @@ -3307,7 +3332,7 @@ lean_inc(x_1); x_46 = l_Lean_Parser_tokenAntiquotFn(x_1, x_30); x_47 = lean_ctor_get(x_46, 4); lean_inc(x_47); -x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_25); +x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_25); lean_dec(x_47); if (x_48 == 0) { @@ -3316,7 +3341,7 @@ lean_dec(x_6); x_49 = l_Lean_Parser_ParserState_mkNode(x_46, x_8, x_28); x_50 = lean_ctor_get(x_49, 4); lean_inc(x_50); -x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_25); +x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_25); lean_dec(x_50); if (x_51 == 0) { @@ -3339,7 +3364,7 @@ x_53 = lean_apply_2(x_6, x_1, x_46); x_54 = l_Lean_Parser_ParserState_mkNode(x_53, x_8, x_28); x_55 = lean_ctor_get(x_54, 4); lean_inc(x_55); -x_56 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_55, x_25); +x_56 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_55, x_25); lean_dec(x_55); if (x_56 == 0) { @@ -3901,7 +3926,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -3934,7 +3959,7 @@ x_25 = l_Lean_Parser_Tactic_decide___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_20, x_25, x_17); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); lean_dec(x_27); if (x_28 == 0) { @@ -3958,7 +3983,7 @@ x_31 = l_Lean_Parser_Tactic_decide___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_17); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -4490,7 +4515,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -4523,7 +4548,7 @@ x_25 = l_Lean_Parser_Tactic_nativeDecide___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_20, x_25, x_17); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); lean_dec(x_27); if (x_28 == 0) { @@ -4547,7 +4572,7 @@ x_31 = l_Lean_Parser_Tactic_nativeDecide___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_17); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -4987,6 +5012,10 @@ l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__16 = _i lean_mark_persistent(l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__16); l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__17 = _init_l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__17(); lean_mark_persistent(l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__17); +l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__18 = _init_l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__18(); +lean_mark_persistent(l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__18); +l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__19 = _init_l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__19(); +lean_mark_persistent(l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6____closed__19); res = l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Tactic___hyg_6_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); diff --git a/stage0/stdlib/Lean/Parser/Term.c b/stage0/stdlib/Lean/Parser/Term.c index a0a7e700858..87d5b922a82 100644 --- a/stage0/stdlib/Lean/Parser/Term.c +++ b/stage0/stdlib/Lean/Parser/Term.c @@ -35,7 +35,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_dbgTrace_declRange___closed_ static lean_object* l___regBuiltin_Lean_Parser_Term_forInMacro_x27_declRange___closed__7; lean_object* l_Lean_PrettyPrinter_Parenthesizer_visitToken___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_scoped___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__34; static lean_object* l_Lean_Parser_Term_explicitUniv___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Term_arrow_declRange___closed__1; static lean_object* l_Lean_Parser_Term_letIdLhs_parenthesizer___closed__7; @@ -52,6 +51,7 @@ static lean_object* l_Lean_Parser_Term_tupleTail_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Tactic_seq1___closed__3; static lean_object* l_Lean_Parser_Term_matchAlt___closed__5; static lean_object* l_Lean_Parser_Term_instBinder_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_Term_letIdBinder___elambda__1___closed__2; static lean_object* l_Lean_Parser_Term_fromTerm___elambda__1___closed__12; static lean_object* l_Lean_Parser_Term_tupleTail___elambda__1___closed__5; static lean_object* l_Lean_Parser_Term_let__tmp___closed__8; @@ -224,7 +224,6 @@ static lean_object* l_Lean_Parser_Term_typeOf___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Term_showRhs; static lean_object* l_Lean_Parser_Term_matchDiscr___elambda__1___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_sorry_formatter(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__19; static lean_object* l_Lean_Parser_Term_anonymousCtor_formatter___closed__2; static lean_object* l_Lean_Parser_Command_docComment___elambda__1___closed__8; static lean_object* l_Lean_Parser_Command_docComment___closed__3; @@ -255,6 +254,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_docComment___elambda__1___lambda_ static lean_object* l_Lean_Parser_Tactic_tacticSeqBracketed___closed__9; LEAN_EXPORT lean_object* l_Lean_Parser_Term_pipeProj_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_showRhs_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__44; static lean_object* l___regBuiltin_Lean_Parser_Tactic_quotSeq_declRange___closed__3; static lean_object* l_Lean_Parser_Term_letrec_formatter___closed__7; static lean_object* l_Lean_Parser_Term_let__tmp_formatter___closed__2; @@ -291,6 +291,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_namedPattern_parenthesi static lean_object* l_Lean_Parser_Term_ident_formatter___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_haveDecl_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Term_letRecDecls___elambda__1___closed__6; +static lean_object* l_Lean_Parser_Term_funImplicitBinder___elambda__1___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_binrel_declRange___closed__4; static lean_object* l_Lean_Parser_Term_whereDecls___elambda__1___closed__17; LEAN_EXPORT lean_object* l_Lean_Parser_Term_attributes_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -336,6 +337,7 @@ static lean_object* l_Lean_Parser_Term_noImplicitLambda_formatter___closed__1; static lean_object* l_Lean_Parser_Term_strictImplicitLeftBracket___closed__2; static lean_object* l_Lean_Parser_Term_paren___elambda__1___closed__16; static lean_object* l_Lean_Parser_Term_haveEqnsDecl_parenthesizer___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__8; static lean_object* l___regBuiltin_Lean_Parser_Term_structInstArrayRef_formatter___closed__1; static lean_object* l_Lean_Parser_Term_byTactic___closed__4; static lean_object* l_Lean_Parser_Term_structInstField___elambda__1___closed__7; @@ -361,6 +363,7 @@ static lean_object* l_Lean_Parser_Term_forInMacro___closed__3; static lean_object* l_Lean_Parser_Term_trailing__parser___elambda__1___closed__8; lean_object* l_Lean_PrettyPrinter_Parenthesizer_withAntiquot_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_explicitUniv_parenthesizer___closed__1; +static lean_object* l_Lean_Parser_Term_binderType___elambda__1___closed__1; static lean_object* l_Lean_Parser_Term_anonymousCtor_formatter___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Term_binderType_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_optIdent___closed__5; @@ -380,7 +383,6 @@ static lean_object* l_Lean_Parser_Term_explicitUniv_formatter___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_match_parenthesizer(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_attr_quot_formatter___closed__1; static lean_object* l_Lean_Parser_Term_assert_formatter___closed__2; -extern lean_object* l_Lean_nullKind; LEAN_EXPORT lean_object* l_Lean_Parser_Term_instBinder; static lean_object* l_Lean_Parser_Tactic_tacticSeqBracketed___closed__4; static lean_object* l_Lean_Parser_Term_fun___closed__7; @@ -397,6 +399,7 @@ static lean_object* l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___close static lean_object* l_Lean_Parser_Term_inaccessible___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_semicolonOrLinebreak_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_strictImplicitLeftBracket_formatter___closed__3; +static lean_object* l_Lean_Parser_Term_funBinder___elambda__1___closed__7; static lean_object* l_Lean_Parser_Term_letrec___elambda__1___closed__6; static lean_object* l_Lean_Parser_Term_strictImplicitLeftBracket___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_forInMacro_x27(lean_object*); @@ -454,6 +457,7 @@ static lean_object* l_Lean_Parser_Term_letEqnsDecl___closed__2; static lean_object* l_Lean_Parser_Term_type___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Term_attributes___elambda__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_fun_formatter___closed__3; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__7; static lean_object* l_Lean_Parser_Term_forInMacro_x27_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_hole___elambda__1___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Term_dotIdent_declRange___closed__5; @@ -485,13 +489,13 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_forall___elambda__1(lean_object*, le lean_object* l_Lean_Parser_setLhsPrecFn(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_completion_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_letMVar_formatter___closed__7; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__47; static lean_object* l___regBuiltin_Lean_Parser_Term_letMVar_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_trueVal_formatter___closed__1; lean_object* l_Lean_PrettyPrinter_Parenthesizer_lookahead_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_argument___closed__7; static lean_object* l_Lean_Parser_Term_inaccessible___closed__9; static lean_object* l_Lean_Parser_Term_char___closed__3; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__49; static lean_object* l_Lean_Parser_Term_strictImplicitRightBracket_formatter___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_pipeProj_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Tactic_tacticSeqBracketed_parenthesizer___closed__2; @@ -516,7 +520,6 @@ static lean_object* l_Lean_Parser_Term_pipeProj_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_sorry_parenthesizer___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_binop_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Term_trueVal_parenthesizer___closed__2; -static lean_object* l_Lean_Parser_Term_letIdLhs___elambda__1___closed__9; static lean_object* l_Lean_Parser_Term_letMVar___elambda__1___closed__14; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_ensureTypeOf_formatter(lean_object*); static lean_object* l_Lean_Parser_Term_type___elambda__1___closed__21; @@ -582,6 +585,7 @@ static lean_object* l_Lean_Parser_Term_macroDollarArg_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Term_anonymousCtor_formatter___closed__3; static lean_object* l_Lean_Parser_Term_implicitBinder___elambda__1___closed__3; static lean_object* l_Lean_Parser_Term_forall___elambda__1___closed__17; +static lean_object* l_Lean_Parser_Term_funImplicitBinder_formatter___closed__7; static lean_object* l_Lean_Parser_Term_trailing__parser___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_waitIfTypeMVar_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_semicolonOrLinebreak_formatter___closed__1; @@ -602,6 +606,7 @@ static lean_object* l_Lean_Parser_Term_match___elambda__1___closed__21; static lean_object* l___regBuiltin_Lean_Parser_Term_arrayRef_declRange___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Term_funBinder_quot_declRange___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_cdot_parenthesizer___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__24; static lean_object* l_Lean_Parser_Term_inaccessible___closed__4; static lean_object* l_Lean_Parser_Term_match_formatter___closed__3; static lean_object* l_Lean_Parser_Term_subst_parenthesizer___closed__2; @@ -623,7 +628,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_let_declRange___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Term_forall_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_namedPattern_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_darrow___elambda__1___closed__4; -LEAN_EXPORT lean_object* l_Lean_Parser_Term_letrec___elambda__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_Term_letrec___elambda__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_attributes_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Term_macroDollarArg_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_prop_formatter___closed__3; @@ -643,10 +648,11 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_matchDiscr_quot_formatter(lean_objec static lean_object* l_Lean_Parser_Term_dynamicQuot___elambda__1___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_nomatch_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_match_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Term_binderType___elambda__1___closed__2; static lean_object* l_Lean_Parser_Term_simpleBinder_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_structInstFieldAbbrev___elambda__1___closed__10; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__33; static lean_object* l_Lean_Parser_Term_letIdLhs_parenthesizer___closed__4; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__2; static lean_object* l_Lean_Parser_Term_letPatDecl___elambda__1___closed__11; static lean_object* l_Lean_Parser_Term_whereDecls___elambda__1___closed__16; static lean_object* l_Lean_Parser_Term_suffices_parenthesizer___closed__7; @@ -702,6 +708,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_hole___elambda__1(lean_object*, lean static lean_object* l_Lean_Parser_Term_forall___elambda__1___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_whereDecls_formatter(lean_object*); static lean_object* l_Lean_Parser_Term_structInst_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_Term_funBinder_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Term_withAnonymousAntiquot___elambda__1___closed__3; static lean_object* l_Lean_Parser_Term_binderTactic___elambda__1___closed__3; static lean_object* l_Lean_Parser_Term_dotIdent___elambda__1___closed__6; @@ -756,6 +763,7 @@ static lean_object* l_Lean_Parser_Term_forall___elambda__1___closed__21; static lean_object* l_Lean_Parser_Term_implicitBinder_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_structInstFieldAbbrev_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Term_explicitUniv_parenthesizer___closed__4; +static lean_object* l_Lean_Parser_Term_funBinder_formatter___closed__7; static lean_object* l_Lean_Parser_Term_dbgTrace_parenthesizer___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Term_explicitUniv_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_waitIfTypeContainsMVar___elambda__1___closed__2; @@ -835,10 +843,10 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_instBinder___elambda__1(lean_object* static lean_object* l_Lean_Parser_Term_anonymousCtor___elambda__1___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_stateRefT_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_haveIdLhs___elambda__1___closed__2; +static lean_object* l_Lean_Parser_Term_letIdBinder_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_withAnonymousAntiquot___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_assert_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_bracketedBinder_quot_declRange___closed__4; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__44; LEAN_EXPORT lean_object* l_Lean_Parser_Term_byTactic; static lean_object* l___regBuiltin_Lean_Parser_Term_depArrow_declRange___closed__2; static lean_object* l_Lean_Parser_Term_falseVal___elambda__1___closed__3; @@ -858,7 +866,6 @@ static lean_object* l_Lean_Parser_Term_trailing__parser___elambda__1___closed__1 static lean_object* l___regBuiltin_Lean_Parser_Term_let__tmp_declRange___closed__2; static lean_object* l_Lean_Parser_Term_cdot_parenthesizer___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_type(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__8; LEAN_EXPORT lean_object* l_Lean_Parser_Term_bracketedBinder_quot_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_anonymousCtor___closed__8; static lean_object* l_Lean_Parser_Term_proj___elambda__1___closed__4; @@ -896,6 +903,7 @@ static lean_object* l_Lean_Parser_Term_explicitUniv_formatter___closed__7; static lean_object* l_Lean_Parser_Term_doubleQuotedName___elambda__1___closed__5; static lean_object* l_Lean_Parser_Term_pipeCompletion___elambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Term_motive; +static lean_object* l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__27; static lean_object* l___regBuiltin_Lean_Parser_Term_app_declRange___closed__1; static lean_object* l_Lean_Parser_Term_show___elambda__1___closed__6; static lean_object* l_Lean_Parser_Term_letRecDecls___elambda__1___closed__3; @@ -1000,6 +1008,7 @@ static lean_object* l_Lean_Parser_Term_funImplicitBinder___elambda__1___closed__ static lean_object* l_Lean_Parser_Term_let__tmp___elambda__1___closed__1; static lean_object* l_Lean_Parser_Term_sorry_formatter___closed__3; static lean_object* l_Lean_Parser_Term_trailing__parser___elambda__1___closed__5; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__48; static lean_object* l_Lean_Parser_Term_app_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_explicitBinder_formatter___closed__1; static lean_object* l_Lean_Parser_Term_structInstLVal___closed__7; @@ -1036,7 +1045,6 @@ static lean_object* l_Lean_Parser_Term_letPatDecl___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_trailing__parser_declRange___closed__4; static lean_object* l_Lean_Parser_Term_strictImplicitRightBracket_parenthesizer___closed__3; lean_object* l_Lean_PrettyPrinter_Parenthesizer_withoutPosition_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__25; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_let__delayed_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Term_motive___closed__10; LEAN_EXPORT lean_object* l_Lean_Parser_Term_attributes_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1050,7 +1058,6 @@ static lean_object* l_Lean_Parser_Term_binop___closed__2; static lean_object* l_Lean_Parser_Term_letPatDecl_parenthesizer___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_macroDollarArg_formatter___closed__2; lean_object* l_Lean_Parser_pushNone___elambda__1___boxed(lean_object*); -static lean_object* l_Lean_Parser_Term_letIdLhs___elambda__1___closed__8; lean_object* l_Lean_Parser_notFollowedByFn(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_funSimpleBinder_formatter___closed__3; static lean_object* l_Lean_Parser_Term_structInst___elambda__1___closed__12; @@ -1181,6 +1188,7 @@ lean_object* l_Lean_Parser_ppIndent_parenthesizer(lean_object*, lean_object*, le LEAN_EXPORT lean_object* l_Lean_Parser_Term_syntheticHole_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_docComment_formatter___closed__2; static lean_object* l_Lean_Parser_Term_waitIfTypeMVar___elambda__1___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_seq1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_forInMacro_x27_declRange(lean_object*); static lean_object* l_Lean_Parser_Term_suffices___closed__8; @@ -1221,7 +1229,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_quotedName_declRange___close static lean_object* l_Lean_Parser_Term_strictImplicitRightBracket___closed__1; static lean_object* l_Lean_Parser_Term_hole_formatter___closed__2; static lean_object* l_Lean_Parser_Term_binrel__no__prop_formatter___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__45; static lean_object* l_Lean_Parser_Term_borrowed_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_completion_formatter___closed__1; static lean_object* l_Lean_Parser_Command_docComment___elambda__1___closed__12; @@ -1252,6 +1259,7 @@ static lean_object* l_Lean_Parser_Term_optEllipsis___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_fun_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Term_depArrow___elambda__1___closed__4; static lean_object* l_Lean_Parser_Term_optType___closed__1; +static lean_object* l_Lean_Parser_Term_funBinder___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Term_depArrow_declRange___closed__1; static lean_object* l_Lean_Parser_Term_local___closed__4; static lean_object* l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__1; @@ -1260,6 +1268,7 @@ static lean_object* l_Lean_Parser_Term_structInst_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_explicitUniv___elambda__1___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Tactic_quotSeq_formatter___closed__2; lean_object* l_Lean_Parser_unicodeSymbol_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__15; static lean_object* l_Lean_Parser_Term_forall___elambda__1___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Term_attr_quot_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_instBinder___closed__3; @@ -1267,6 +1276,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_doubleQuotedName___elambda__1(lean_o LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_quotedName(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_app_declRange(lean_object*); static lean_object* l_Lean_Parser_Term_haveIdLhs___closed__3; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__18; static lean_object* l___regBuiltin_Lean_Parser_Term_doubleQuotedName_formatter___closed__2; static lean_object* l_Lean_Parser_Term_instBinder_parenthesizer___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_proj_parenthesizer___closed__1; @@ -1285,7 +1295,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_defaultOrOfNonempty_declRang static lean_object* l_Lean_Parser_Term_funBinder_quot___closed__2; static lean_object* l_Lean_Parser_Term_assert_parenthesizer___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_tacticSeq___elambda__1(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__40; static lean_object* l_Lean_Parser_Term_suffices___elambda__1___closed__3; static lean_object* l_Lean_Parser_Term_scoped_parenthesizer___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Tactic_tacticSeqBracketed_formatter___closed__3; @@ -1294,7 +1303,6 @@ static lean_object* l_Lean_Parser_Term_argument___closed__5; static lean_object* l_Lean_Parser_Term_haveDecl___elambda__1___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_forall(lean_object*); static lean_object* l_Lean_Parser_Term_optExprPrecedence_formatter___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__9; static lean_object* l_Lean_Parser_Term_let__tmp___elambda__1___closed__9; static lean_object* l_Lean_Parser_Term_sorry_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_binderType_parenthesizer(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1310,6 +1318,7 @@ static lean_object* l_Lean_Parser_Term_namedPattern___elambda__1___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_dotIdent_declRange___closed__4; static lean_object* l_Lean_Parser_Term_whereDecls___elambda__1___closed__5; static lean_object* l_Lean_Parser_Term_attr_quot___elambda__1___closed__9; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__30; static lean_object* l_Lean_Parser_Term_show___elambda__1___closed__3; static lean_object* l_Lean_Parser_Term_generalizingParam_formatter___closed__1; static lean_object* l_Lean_Parser_Term_tupleTail___elambda__1___closed__8; @@ -1375,6 +1384,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_doubleQuotedName_declRa static lean_object* l_Lean_Parser_Term_forall_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_waitIfContainsMVar___elambda__1___closed__12; static lean_object* l_Lean_Parser_Term_attrKind___elambda__1___closed__6; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__5; static lean_object* l_Lean_Parser_Term_letMVar___elambda__1___closed__7; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_haveDecl_formatter(lean_object*); static lean_object* l_Lean_Parser_Term_letRecDecls___closed__5; @@ -1393,7 +1403,6 @@ static lean_object* l_Lean_Parser_Tactic_seq1___elambda__1___closed__5; static lean_object* l_Lean_Parser_Term_whereDecls___elambda__1___closed__15; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Tactic_quot(lean_object*); static lean_object* l_Lean_Parser_Term_paren_formatter___closed__6; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__36; static lean_object* l_Lean_Parser_Term_haveIdDecl___elambda__1___closed__2; static lean_object* l_Lean_Parser_Term_proj___elambda__1___closed__1; lean_object* l_Lean_Parser_group_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1454,6 +1463,7 @@ lean_object* l_Lean_PrettyPrinter_Parenthesizer_rawCh_parenthesizer___boxed(lean LEAN_EXPORT lean_object* l_Lean_Parser_Term_pipeCompletion___elambda__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_inaccessible___elambda__1___closed__13; LEAN_EXPORT lean_object* l_Lean_Parser_Term_binrel__no__prop_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Term_funImplicitBinder___closed__7; static lean_object* l_Lean_Parser_Term_letDecl___closed__2; static lean_object* l_Lean_Parser_Term_binrel___closed__9; static lean_object* l___regBuiltin_Lean_Parser_Term_sorry_declRange___closed__3; @@ -1516,6 +1526,7 @@ static lean_object* l_Lean_Parser_Term_let_formatter___closed__1; static lean_object* l_Lean_Parser_Term_anonymousCtor___elambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_tacticParser_formatter___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_attributes___closed__8; +LEAN_EXPORT lean_object* l_Lean_Parser_Term_letIdBinder; static lean_object* l_Lean_Parser_Term_paren___elambda__1___closed__10; static lean_object* l___regBuiltin_Lean_Parser_Term_ident_declRange___closed__5; static lean_object* l_Lean_Parser_Term_typeOf___closed__7; @@ -1561,6 +1572,7 @@ static lean_object* l_Lean_Parser_Term_unreachable___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_app_formatter___closed__1; static lean_object* l_Lean_Parser_Term_strictImplicitLeftBracket___elambda__1___closed__1; static lean_object* l_Lean_Parser_Term_namedArgument___elambda__1___closed__4; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__22; static lean_object* l_Lean_Parser_Term_binrel_formatter___closed__6; static lean_object* l_Lean_Parser_Term_whereDecls___elambda__1___closed__7; lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkPrec_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1636,13 +1648,13 @@ static lean_object* l_Lean_Parser_Term_noindex___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_forInMacro_x27_declRange___closed__1; static lean_object* l_Lean_Parser_Term_matchDiscr___elambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Term_scientific_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Term_funBinder___elambda__1___closed__6; static lean_object* l_Lean_Parser_Term_match___closed__3; static lean_object* l_Lean_Parser_Term_sort___elambda__1___closed__5; static lean_object* l_Lean_Parser_Term_fun___elambda__1___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_char_declRange___closed__7; static lean_object* l_Lean_Parser_Term_attrInstance___closed__5; static lean_object* l_Lean_Parser_Term_waitIfTypeContainsMVar___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__35; static lean_object* l_Lean_Parser_Term_sorry___elambda__1___closed__4; static lean_object* l_Lean_Parser_Term_structInst_formatter___closed__17; static lean_object* l_Lean_Parser_Term_suffices___elambda__1___closed__2; @@ -1735,6 +1747,7 @@ static lean_object* l_Lean_Parser_Term_hole___closed__6; static lean_object* l_Lean_Parser_Term_withAnonymousAntiquot___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Term_letRecDecls_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_forall_parenthesizer___closed__3; +uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_quot_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_simpleBinderWithoutType___closed__4; static lean_object* l_Lean_Parser_Term_optEllipsis___elambda__1___closed__2; @@ -1790,6 +1803,7 @@ static lean_object* l_Lean_Parser_Term_byTactic_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Term_assert_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_haveDecl___elambda__1___closed__8; static lean_object* l_Lean_Parser_Term_structInstLVal___elambda__1___closed__4; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__14; static lean_object* l_Lean_Parser_Term_haveIdDecl_formatter___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_doubleQuotedName_declRange___closed__1; static lean_object* l_Lean_Parser_Term_completion___closed__3; @@ -1803,6 +1817,7 @@ static lean_object* l_Lean_Parser_Term_defaultOrOfNonempty___elambda__1___closed static lean_object* l___regBuiltin_Lean_Parser_Tactic_quot_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_let_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_trueVal___elambda__1___closed__5; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__28; static lean_object* l_Lean_Parser_Term_motive___elambda__1___closed__7; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_panic(lean_object*); static lean_object* l_Lean_Parser_Term_structInst_parenthesizer___closed__8; @@ -1883,6 +1898,7 @@ extern lean_object* l_Lean_PrettyPrinter_parenthesizerAttribute; static lean_object* l_Lean_Parser_Term_funStrictImplicitBinder___closed__1; static lean_object* l_Lean_Parser_Term_suffices___closed__4; static lean_object* l_Lean_Parser_Term_tupleTail_parenthesizer___closed__5; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__29; static lean_object* l_Lean_Parser_Term_sorry___closed__7; static lean_object* l_Lean_Parser_Term_attr_quot___elambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Term_let__fun_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1922,6 +1938,7 @@ static lean_object* l_Lean_Parser_Term_haveIdDecl_formatter___closed__3; static lean_object* l_Lean_Parser_Term_hole___closed__4; static lean_object* l_Lean_Parser_Term_matchDiscr_quot_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Term_structInstLVal___closed__3; +static lean_object* l_Lean_Parser_Term_letIdBinder___elambda__1___closed__4; static lean_object* l_Lean_Parser_Term_fun___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_quotedName_formatter___closed__2; static lean_object* l_Lean_Parser_Term_basicFun___elambda__1___closed__6; @@ -1976,7 +1993,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_fun_declRange___closed__1; static lean_object* l_Lean_Parser_Term_sort___closed__2; static lean_object* l_Lean_Parser_Term_scoped___elambda__1___closed__6; static lean_object* l_Lean_Parser_Term_let__delayed___closed__8; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__7; static lean_object* l_Lean_Parser_Term_let__tmp_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Term_let___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_optEllipsis_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -2018,7 +2034,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_pipeProj_declRange___closed_ static lean_object* l___regBuiltin_Lean_Parser_Term_dbgTrace_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_quotedName___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Term_namedPattern; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__11; static lean_object* l___regBuiltin_Lean_Parser_Term_sort_declRange___closed__4; static lean_object* l_Lean_Parser_Term_binrel___elambda__1___closed__9; static lean_object* l_Lean_Parser_Term_motive_parenthesizer___closed__3; @@ -2090,6 +2105,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_simpleBinder_formatter( static lean_object* l_Lean_Parser_Term_ensureExpectedType___elambda__1___closed__11; static lean_object* l_Lean_Parser_Term_structInstLVal_formatter___closed__9; static lean_object* l_Lean_Parser_Term_binrel___elambda__1___closed__4; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__12; static lean_object* l_Lean_Parser_Term_generalizingParam___elambda__1___closed__13; static lean_object* l_Lean_Parser_Term_namedArgument_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Command_docComment___closed__4; @@ -2282,6 +2298,7 @@ static lean_object* l_Lean_Parser_Term_whereDecls___closed__4; static lean_object* l_Lean_Parser_Term_typeOf___elambda__1___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_attr_quot_formatter(lean_object*); static lean_object* l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__4; +static lean_object* l_Lean_Parser_Term_funSimpleBinder___elambda__1___closed__4; static lean_object* l_Lean_Parser_Term_typeSpec___closed__2; static lean_object* l_Lean_Parser_Term_attributes___elambda__1___closed__7; static lean_object* l_Lean_Parser_Term_dynamicQuot___closed__7; @@ -2302,7 +2319,6 @@ static lean_object* l_Lean_Parser_Term_quotedName_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_dynamicQuot_formatter___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Term_simpleBinderWithoutType_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_ensureTypeOf___elambda__1___closed__4; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__31; static lean_object* l_Lean_Parser_Term_falseVal___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Term_letDecl; lean_object* l_Lean_Parser_scientificLit_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -2333,7 +2349,6 @@ static lean_object* l_Lean_Parser_Term_namedArgument_formatter___closed__4; static lean_object* l_Lean_Parser_Term_typeOf_formatter___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_byTactic_parenthesizer(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_sort; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__39; lean_object* l_Lean_Parser_strLit_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_let__tmp_formatter___closed__6; static lean_object* l_Lean_Parser_Term_explicitUniv___elambda__1___closed__6; @@ -2503,9 +2518,9 @@ static lean_object* l_Lean_Parser_Term_quotedName___elambda__1___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Term_defaultOrOfNonempty___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_panic_declRange___closed__7; +static lean_object* l_Lean_Parser_Term_letIdBinder___closed__4; static lean_object* l_Lean_Parser_Term_structInstFieldAbbrev___elambda__1___closed__2; static lean_object* l_Lean_Parser_Term_stateRefT___closed__4; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_letIdDecl___elambda__1(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Level_quot_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_let__tmp_declRange___closed__4; @@ -2537,6 +2552,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_have_parenthesizer(lean static lean_object* l_Lean_Parser_Term_suffices___elambda__1___closed__4; static lean_object* l_Lean_Parser_Term_dynamicQuot___closed__12; static lean_object* l___regBuiltin_Lean_Parser_Term_binrel_parenthesizer___closed__1; +static lean_object* l_Lean_Parser_Term_letIdBinder___closed__1; static lean_object* l_Lean_Parser_Term_structInst_formatter___closed__13; static lean_object* l_Lean_Parser_Term_panic___closed__8; static lean_object* l_Lean_Parser_Term_let__delayed___closed__1; @@ -2560,7 +2576,6 @@ static lean_object* l_Lean_Parser_Term_stateRefT___closed__8; static lean_object* l_Lean_Parser_Term_dynamicQuot___elambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_attrInstance___elambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_dotIdent; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Term_quotedName; static lean_object* l_Lean_Parser_Term_explicitUniv_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Term_fromTerm___elambda__1___closed__11; @@ -2634,7 +2649,6 @@ static lean_object* l_Lean_Parser_Term_fun___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_type_declRange___closed__3; static lean_object* l_Lean_Parser_Term_sorry___closed__5; static lean_object* l_Lean_Parser_Term_borrowed___elambda__1___closed__11; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__27; static lean_object* l___regBuiltin_Lean_Parser_Term_binrel_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_binop__lazy_declRange___closed__1; static lean_object* l_Lean_Parser_Term_fromTerm___elambda__1___closed__3; @@ -2645,6 +2659,7 @@ static lean_object* l_Lean_Parser_Term_structInstLVal___elambda__1___closed__15; static lean_object* l_Lean_Parser_Term_have_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_attrInstance___closed__4; static lean_object* l_Lean_Parser_Term_attr_quot_parenthesizer___closed__6; +static lean_object* l_Lean_Parser_Term_funBinder___elambda__1___closed__5; static lean_object* l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__19; static lean_object* l___regBuiltin_Lean_Parser_Term_matchDiscr_quot_declRange___closed__7; static lean_object* l_Lean_Parser_Term_cdot___elambda__1___closed__12; @@ -2687,7 +2702,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_strictImplicitLeftBracket_parenthesi static lean_object* l_Lean_Parser_Term_let__fun_formatter___closed__7; static lean_object* l_Lean_Parser_Term_defaultOrOfNonempty___closed__7; static lean_object* l_Lean_Parser_Term_depArrow___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__17; lean_object* l_Lean_PrettyPrinter_Parenthesizer_interpolatedStr_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_let__tmp_formatter___closed__1; static lean_object* l_Lean_Parser_Term_panic___closed__4; @@ -2794,6 +2808,7 @@ static lean_object* l_Lean_Parser_Command_docComment___elambda__1___closed__9; static lean_object* l___regBuiltin_Lean_Parser_Term_app_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_anonymousCtor_formatter___closed__5; static lean_object* l_Lean_Parser_Term_ensureTypeOf___elambda__1___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__10; static lean_object* l_Lean_Parser_Tactic_quot___elambda__1___closed__1; static lean_object* l_Lean_Parser_Term_sort___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Term_sufficesDecl_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -2810,6 +2825,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_dynamicQuot_declRange___clos static lean_object* l_Lean_Parser_Term_subst___elambda__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Term_borrowed_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_show_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__32; LEAN_EXPORT lean_object* l_Lean_Parser_Term_showRhs_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_forInMacro_x27_parenthesizer___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Term_whereDecls; @@ -2835,6 +2851,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_forInMacro_parenthesizer(lean_object LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_unreachable_declRange(lean_object*); static lean_object* l_Lean_Parser_Term_binderDefault___elambda__1___closed__5; lean_object* l_Lean_Parser_checkNoWsBeforeFn(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Term_letIdBinder___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_dbgTrace_formatter(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_arrayRef_declRange___closed__5; static lean_object* l_Lean_Parser_Term_structInstLVal_formatter___closed__11; @@ -2945,6 +2962,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_syntheticHole_parenthesizer_ static lean_object* l_Lean_Parser_Term_structInst___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Term_attributes; static lean_object* l_Lean_Parser_Term_inaccessible___elambda__1___closed__3; +static lean_object* l_Lean_Parser_Term_funSimpleBinder_parenthesizer___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_letEqnsDecl_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_withAnonymousAntiquot___elambda__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_scoped___closed__3; @@ -2995,9 +3013,11 @@ static lean_object* l_Lean_Parser_Term_ensureExpectedType_formatter___closed__3; static lean_object* l_Lean_Parser_Level_quot___closed__2; static lean_object* l_Lean_Parser_Term_leading__parser_parenthesizer___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_defaultOrOfNonempty_parenthesizer(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__42; extern lean_object* l_Lean_Parser_maxPrec; static lean_object* l_Lean_Parser_Term_doubleQuotedName___elambda__1___closed__16; static lean_object* l_Lean_Parser_Term_local_formatter___closed__1; +static lean_object* l_Lean_Parser_Term_letIdBinder___elambda__1___closed__1; static lean_object* l_Lean_Parser_Term_haveEqnsDecl_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Term_strictImplicitLeftBracket___elambda__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_argument___elambda__1___closed__1; @@ -3044,6 +3064,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_assert_declRange___closed__6 static lean_object* l_Lean_Parser_Term_forInMacro_formatter___closed__2; static lean_object* l_Lean_Parser_Term_dynamicQuot___closed__10; lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_addQuotDepthFn___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__43; static lean_object* l_Lean_Parser_Term_letEqnsDecl_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Term_local___elambda__1___closed__4; static lean_object* l_Lean_Parser_Term_letPatDecl___elambda__1___closed__8; @@ -3092,6 +3113,7 @@ static lean_object* l_Lean_Parser_Term_binderDefault___closed__2; static lean_object* l_Lean_Parser_Term_byTactic___elambda__1___closed__13; LEAN_EXPORT lean_object* l_Lean_Parser_Term_fromTerm_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_structInstField___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__37; static lean_object* l_Lean_Parser_Term_binderTactic___elambda__1___closed__9; static lean_object* l_Lean_Parser_Term_structInstLVal___closed__6; static lean_object* l_Lean_Parser_Term_binop___closed__4; @@ -3129,6 +3151,7 @@ static lean_object* l_Lean_Parser_Term_defaultOrOfNonempty___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_structInst_declRange___closed__6; static lean_object* l_Lean_Parser_Term_suffices___elambda__1___closed__6; static lean_object* l_Lean_Parser_Term_sort___elambda__1___closed__7; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__20; static lean_object* l_Lean_Parser_Term_structInstField___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_sorry_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Term_pipeProj_parenthesizer___closed__5; @@ -3172,6 +3195,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_noImplicitLambda_declRange__ static lean_object* l___regBuiltin_Lean_Parser_Term_byTactic_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_nomatch_formatter___closed__1; static lean_object* l_Lean_Parser_Term_haveDecl___closed__4; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__16; static lean_object* l_Lean_Parser_Term_let__delayed_formatter___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Term_unreachable; LEAN_EXPORT lean_object* l_Lean_Parser_Term_letPatDecl_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -3344,11 +3368,11 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_sort(lean_object*); static lean_object* l_Lean_Parser_Term_explicit___closed__7; static lean_object* l_Lean_Parser_Term_completion___closed__1; static lean_object* l_Lean_Parser_Term_namedPattern___closed__5; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__15; static lean_object* l_Lean_Parser_Term_bracketedBinder_quot_parenthesizer___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_show_formatter___closed__2; static lean_object* l_Lean_Parser_Term_structInstLVal___elambda__1___closed__12; lean_object* l_Lean_PrettyPrinter_Parenthesizer_pushNone_parenthesizer___boxed(lean_object*); +static lean_object* l_Lean_Parser_Term_funSimpleBinder_formatter___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Term_app_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_explicitBinder(uint8_t); static lean_object* l_Lean_Parser_Term_type___closed__1; @@ -3425,7 +3449,6 @@ static lean_object* l_Lean_Parser_Term_structInstField___closed__9; static lean_object* l_Lean_Parser_Term_dynamicQuot_formatter___closed__2; static lean_object* l_Lean_Parser_Term_binrel__no__prop___elambda__1___closed__6; static lean_object* l_Lean_Parser_Term_letMVar___closed__5; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__18; static lean_object* l_Lean_Parser_Term_structInst___elambda__1___closed__21; LEAN_EXPORT lean_object* l_Lean_Parser_Term_isIdent___boxed(lean_object*); static lean_object* l_Lean_Parser_Term_dynamicQuot___closed__3; @@ -3455,6 +3478,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_letEqnsDecl; static lean_object* l_Lean_Parser_Term_dbgTrace___elambda__1___closed__1; static lean_object* l_Lean_Parser_Term_generalizingParam_formatter___closed__5; static lean_object* l_Lean_Parser_Term_motive_formatter___closed__5; +static lean_object* l_Lean_Parser_Term_letIdBinder_formatter___closed__1; lean_object* l_Lean_Parser_rawCh(uint32_t, uint8_t); static lean_object* l_Lean_Parser_Term_funStrictImplicitBinder___elambda__1___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_nomatch(lean_object*); @@ -3468,6 +3492,7 @@ static lean_object* l_Lean_Parser_Term_binderDefault___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_explicit_formatter___closed__1; static lean_object* l_Lean_Parser_Term_letEqnsDecl_formatter___closed__2; static lean_object* l_Lean_Parser_Term_have___closed__8; +static lean_object* l_Lean_Parser_Term_letIdBinder_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_anonymousCtor_declRange___closed__3; static lean_object* l_Lean_Parser_Term_structInstLVal_formatter___closed__2; static lean_object* l_Lean_Parser_Term_sorry___elambda__1___closed__6; @@ -3481,12 +3506,10 @@ static lean_object* l_Lean_Parser_Term_forall_formatter___closed__8; static lean_object* l_Lean_Parser_Term_quotedName___elambda__1___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Term_letEqnsDecl___elambda__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_structInstLVal___elambda__1___closed__26; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__3; static lean_object* l_Lean_Parser_Tactic_tacticSeqBracketed___elambda__1___closed__9; static lean_object* l_Lean_Parser_Term_defaultOrOfNonempty_parenthesizer___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Term_scoped_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_matchDiscr_quot_parenthesizer___closed__6; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__6; static lean_object* l_Lean_Parser_Term_matchAlt___closed__11; static lean_object* l_Lean_Parser_Term_ensureExpectedType___elambda__1___closed__10; LEAN_EXPORT lean_object* l_Lean_Parser_Term_funBinder_quot; @@ -3570,7 +3593,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_binop_parenthesizer(lean_object*, le static lean_object* l_Lean_Parser_Term_optExprPrecedence___closed__5; static lean_object* l_Lean_Parser_Term_haveDecl_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_matchDiscr_quot___closed__4; -lean_object* l_Lean_Parser_registerAlias(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Parser_registerAlias(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_quotSeq___elambda__1___closed__4; static lean_object* l_Lean_Parser_Term_pipeProj___elambda__1___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Term_haveIdDecl_parenthesizer___closed__2; @@ -3644,7 +3667,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_quotedName_parenthesize LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_typeOf(lean_object*); static lean_object* l_Lean_Parser_Term_haveIdDecl___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_arrow_declRange___closed__5; -static lean_object* l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__18; static lean_object* l_Lean_Parser_Term_borrowed___closed__7; static lean_object* l_Lean_Parser_Term_generalizingParam___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Term_binderIdent_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -3655,6 +3677,7 @@ static lean_object* l_Lean_Parser_Term_arrayRef___elambda__1___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_str_declRange___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Term_noindex; static lean_object* l_Lean_Parser_Term_generalizingParam___elambda__1___closed__3; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__27; static lean_object* l_Lean_Parser_Term_letEqnsDecl_formatter___closed__1; static lean_object* l_Lean_Parser_Term_defaultOrOfNonempty___elambda__1___closed__6; static lean_object* l_Lean_Parser_Term_cdot___elambda__1___closed__2; @@ -3689,6 +3712,7 @@ static lean_object* l_Lean_Parser_Term_waitIfTypeContainsMVar___elambda__1___clo static lean_object* l_Lean_Parser_Term_structInstField___elambda__1___closed__2; static lean_object* l_Lean_Parser_Command_docComment_formatter___closed__5; static lean_object* l_Lean_Parser_Term_typeOf_parenthesizer___closed__4; +LEAN_EXPORT lean_object* l_Lean_Parser_Term_letIdBinder___elambda__1(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_str___closed__1; static lean_object* l_Lean_Parser_Command_docComment___elambda__1___closed__5; static lean_object* l_Lean_Parser_Term_macroDollarArg___elambda__1___closed__12; @@ -3737,6 +3761,7 @@ static lean_object* l_Lean_Parser_Term_prop___closed__5; lean_object* l_Lean_ppHardSpace_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_letMVar_formatter___closed__9; lean_object* l_Lean_Parser_many1_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__40; static lean_object* l_Lean_Parser_Term_explicit___elambda__1___closed__5; static lean_object* l_Lean_Parser_Term_funImplicitBinder___closed__5; static lean_object* l_Lean_Parser_Term_let__tmp_parenthesizer___closed__5; @@ -3819,7 +3844,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_leading__parser_declRange___ static lean_object* l_Lean_Parser_Term_haveIdDecl_formatter___closed__1; static lean_object* l_Lean_Parser_Term_syntheticHole_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Term_fun___elambda__1___closed__7; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__20; static lean_object* l_Lean_Parser_Term_fun___closed__9; static lean_object* l_Lean_Parser_Term_let__delayed___elambda__1___closed__1; static lean_object* l_Lean_Parser_Tactic_tacticSeqBracketed_parenthesizer___closed__16; @@ -3830,6 +3854,7 @@ static lean_object* l_Lean_Parser_Term_binop__lazy_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_noindex___closed__6; static lean_object* l_Lean_Parser_Tactic_quotSeq___elambda__1___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Tactic_tacticSeq1Indented_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_Term_letIdBinder_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_letEqnsDecl_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_borrowed___elambda__1___closed__8; static lean_object* l_Lean_Parser_Term_typeAscription___closed__7; @@ -3866,14 +3891,12 @@ static lean_object* l_Lean_Parser_Term_ensureExpectedType___closed__3; static lean_object* l_Lean_Parser_Term_dbgTrace___elambda__1___closed__2; static lean_object* l_Lean_Parser_Term_scoped___elambda__1___closed__3; static lean_object* l_Lean_Parser_Term_explicit___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__10; static lean_object* l_Lean_Parser_Term_binderType___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Term_str_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_anonymousCtor___elambda__1___closed__15; static lean_object* l_Lean_Parser_Tactic_quot_formatter___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Term_byTactic_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_structInstArrayRef_formatter___closed__5; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__32; static lean_object* l_Lean_Parser_Tactic_tacticSeqBracketed_formatter___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Term_simpleBinder_formatter___closed__2; static lean_object* l_Lean_Parser_Term_macroDollarArg___closed__7; @@ -3918,7 +3941,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_letMVar_declRange___closed__ static lean_object* l_Lean_Parser_Term_depArrow_parenthesizer___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_arrayRef_formatter(lean_object*); static lean_object* l_Lean_Parser_Term_arrow___elambda__1___closed__1; -extern lean_object* l_Lean_groupKind; static lean_object* l_Lean_Parser_Term_binop__lazy___elambda__1___closed__1; static lean_object* l_Lean_Parser_Term_ensureTypeOf_formatter___closed__3; static lean_object* l_Lean_Parser_Term_have___elambda__1___closed__2; @@ -3936,6 +3958,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_let_declRange___closed__1; static lean_object* l_Lean_Parser_Tactic_seq1___elambda__1___closed__7; static lean_object* l_Lean_Parser_Term_namedArgument___closed__8; LEAN_EXPORT lean_object* l_Lean_Parser_Term_fromTerm_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__26; static lean_object* l___regBuiltin_Lean_Parser_Tactic_quotSeq_declRange___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_arrayRef_declRange___closed__2; static lean_object* l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__3; @@ -3997,6 +4020,7 @@ static lean_object* l_Lean_Parser_Term_byTactic___elambda__1___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Term_binderIdent; static lean_object* l_Lean_Parser_Term_attributes___elambda__1___closed__11; LEAN_EXPORT lean_object* l_Lean_Parser_Term_simpleBinder___elambda__1(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__39; static lean_object* l_Lean_Parser_Term_falseVal_formatter___closed__1; static lean_object* l_Lean_Parser_Term_binop__lazy___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_typeOf_declRange___closed__7; @@ -4162,6 +4186,7 @@ static lean_object* l_Lean_Parser_Level_quot___elambda__1___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Term_match_formatter___closed__2; static lean_object* l_Lean_Parser_Term_match___elambda__1___closed__4; static lean_object* l_Lean_Parser_Term_explicitUniv_parenthesizer___closed__3; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__45; LEAN_EXPORT lean_object* l_Lean_Parser_Term_let__fun_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_PrettyPrinter_Formatter_formatterAliasesRef; LEAN_EXPORT lean_object* l_Lean_Parser_Term_attrKind; @@ -4183,6 +4208,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_letIdLhs_parenthesizer(lean_object*, static lean_object* l___regBuiltin_Lean_Parser_Tactic_seq1_formatter___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_sort_declRange(lean_object*); static lean_object* l_Lean_Parser_Term_letMVar___elambda__1___closed__9; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_pipeCompletion_declRange___closed__2; static lean_object* l_Lean_Parser_Term_ensureTypeOf___elambda__1___closed__11; static lean_object* l_Lean_Parser_Term_depArrow_formatter___closed__3; @@ -4196,7 +4222,6 @@ static lean_object* l_Lean_Parser_Term_letPatDecl_formatter___closed__7; static lean_object* l_Lean_Parser_Term_optEllipsis_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_binderDefault___elambda__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_nomatch___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__12; static lean_object* l_Lean_Parser_Term_haveDecl___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_ident_declRange(lean_object*); static lean_object* l_Lean_Parser_Tactic_quot___elambda__1___closed__13; @@ -4290,6 +4315,8 @@ static lean_object* l_Lean_Parser_Term_subst___closed__3; static lean_object* l_Lean_Parser_Term_structInstFieldAbbrev___elambda__1___closed__8; static lean_object* l_Lean_Parser_Term_scoped___elambda__1___closed__10; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_waitIfContainsMVar_formatter(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_Term_instCoeTSyntaxConsSyntaxNodeKindMkStrAnonymousNil__1___boxed(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__9; static lean_object* l_Lean_Parser_Term_motive___closed__1; static lean_object* l_Lean_Parser_Term_type___elambda__1___closed__8; static lean_object* l_Lean_Parser_Term_let__fun_formatter___closed__4; @@ -4299,7 +4326,6 @@ static lean_object* l_Lean_Parser_Term_sufficesDecl___closed__2; static lean_object* l_Lean_Parser_Term_anonymousCtor___closed__7; static lean_object* l_Lean_Parser_Command_docComment___closed__8; static lean_object* l_Lean_Parser_Term_structInstArrayRef___elambda__1___closed__6; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__43; LEAN_EXPORT lean_object* l_Lean_Parser_Term_letDecl_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_structInstLVal___elambda__1___closed__13; static lean_object* l___regBuiltin_Lean_Parser_Term_letrec_declRange___closed__3; @@ -4313,7 +4339,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_leading__parser_declRange___ static lean_object* l___regBuiltin_Lean_Parser_Term_anonymousCtor_declRange___closed__4; static lean_object* l_Lean_Parser_Term_sufficesDecl___elambda__1___closed__2; static lean_object* l_Lean_Parser_darrow___closed__1; -LEAN_EXPORT lean_object* l_Lean_Parser_Term_letrec___elambda__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_Term_letrec___elambda__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_dynamicQuot_formatter___closed__4; static lean_object* l_Lean_Parser_Term_binop_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_waitIfTypeContainsMVar_declRange___closed__1; @@ -4328,7 +4354,6 @@ static lean_object* l_Lean_Parser_Term_paren___elambda__1___closed__2; lean_object* l_Lean_Parser_categoryParser(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_attr_quot_parenthesizer___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_num_declRange___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__26; static lean_object* l_Lean_Parser_Term_withAnonymousAntiquot___closed__10; static lean_object* l_Lean_Parser_Tactic_tacticSeq___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Term_letIdLhs___elambda__1(lean_object*, lean_object*); @@ -4380,6 +4405,7 @@ static lean_object* l_Lean_Parser_Term_letRecDecl___elambda__1___closed__5; static lean_object* l_Lean_Parser_Term_arrow_parenthesizer___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_anonymousCtor; static lean_object* l_Lean_Parser_Term_stateRefT___closed__9; +static lean_object* l_Lean_Parser_Term_funBinder_formatter___closed__8; static lean_object* l_Lean_Parser_Term_forall___closed__9; static lean_object* l___regBuiltin_Lean_Parser_Level_quot_declRange___closed__7; static lean_object* l_Lean_Parser_Level_quot___elambda__1___closed__4; @@ -4488,7 +4514,6 @@ static lean_object* l_Lean_Parser_Term_structInst___closed__13; static lean_object* l_Lean_Parser_Term_byTactic_x27___closed__2; static lean_object* l_Lean_Parser_Term_ellipsis___elambda__1___closed__2; static lean_object* l_Lean_Parser_Term_forInMacro_x27_parenthesizer___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__24; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_funBinder_quot_declRange(lean_object*); static lean_object* l_Lean_Parser_Term_matchDiscr_quot___closed__2; static lean_object* l_Lean_Parser_Term_dynamicQuot___elambda__1___closed__5; @@ -4602,6 +4627,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_doubleQuotedName_declRange__ static lean_object* l_Lean_Parser_Term_waitIfTypeMVar___closed__11; static lean_object* l___regBuiltin_Lean_Parser_Term_pipeProj_declRange___closed__3; lean_object* l_Lean_Parser_symbolInfo(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__11; static lean_object* l_Lean_Parser_Tactic_quot___closed__9; LEAN_EXPORT lean_object* l_Lean_Parser_Term_optIdent_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_haveEqnsDecl___closed__6; @@ -4649,6 +4675,7 @@ static lean_object* l_Lean_Parser_Term_let__fun_formatter___closed__2; static lean_object* l_Lean_Parser_Term_argument_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_let__tmp___elambda__1___closed__10; static lean_object* l_Lean_Parser_Term_matchAltsWhereDecls_formatter___closed__4; +static lean_object* l_Lean_Parser_Term_funSimpleBinder___elambda__1___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_whereDecls_parenthesizer(lean_object*); lean_object* l_Lean_Parser_orelseFnCore(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_let__delayed_declRange___closed__6; @@ -4816,6 +4843,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_ellipsis_parenthesizer( static lean_object* l_Lean_Parser_Term_let___elambda__1___closed__4; static lean_object* l_Lean_Parser_Term_attributes_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_haveIdDecl___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__31; static lean_object* l_Lean_Parser_Term_letIdDecl_formatter___closed__6; static lean_object* l_Lean_Parser_Term_trueVal_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_trailing__parser_parenthesizer___closed__2; @@ -4849,7 +4877,6 @@ static lean_object* l_Lean_Parser_Term_tupleTail___closed__2; static lean_object* l_Lean_Parser_Term_dbgTrace_formatter___closed__6; static lean_object* l_Lean_Parser_Term_assert___elambda__1___closed__4; static lean_object* l_Lean_Parser_Term_attributes___elambda__1___closed__9; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__42; static lean_object* l___regBuiltin_Lean_Parser_Term_defaultOrOfNonempty_declRange___closed__2; static lean_object* l_Lean_Parser_Term_typeAscription_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Term_depArrow_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -4891,6 +4918,7 @@ static lean_object* l_Lean_Parser_Term_letMVar___elambda__1___closed__3; static lean_object* l_Lean_Parser_Term_strictImplicitRightBracket___elambda__1___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_proj_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_motive_formatter___closed__1; +static lean_object* l_Lean_Parser_Term_letIdBinder___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Term_strictImplicitBinder_parenthesizer(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_binderDefault___elambda__1___closed__1; static lean_object* l_Lean_Parser_Term_letRecDecl_parenthesizer___closed__1; @@ -4909,9 +4937,9 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_prop(lean_object*); static lean_object* l_Lean_Parser_Term_attr_quot___elambda__1___closed__14; static lean_object* l_Lean_Parser_Tactic_tacticSeq1Indented___closed__6; static lean_object* l_Lean_Parser_Term_dbgTrace___elambda__1___closed__5; +LEAN_EXPORT lean_object* l_Lean_Parser_Term_instCoeTSyntaxConsSyntaxNodeKindMkStrAnonymousNil__1(lean_object*); static lean_object* l_Lean_Parser_Term_letIdDecl___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_binrel_declRange___closed__6; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__37; static lean_object* l_Lean_Parser_Term_fun___elambda__1___closed__16; static lean_object* l_Lean_Parser_Term_syntheticHole___elambda__1___closed__11; static lean_object* l___regBuiltin_Lean_Parser_Term_suffices_formatter___closed__2; @@ -4941,10 +4969,12 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_forall_declRange___closed__5 static lean_object* l_Lean_Parser_Term_have_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Term_attributes___elambda__1___closed__4; static lean_object* l_Lean_Parser_Term_sort_parenthesizer___closed__4; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_optType_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Tactic_tacticSeq_formatter___closed__1; static lean_object* l_Lean_Parser_Term_noImplicitLambda_formatter___closed__3; static lean_object* l_Lean_Parser_Term_borrowed___elambda__1___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__17; static lean_object* l_Lean_Parser_Term_fun_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_stateRefT_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_let__fun_declRange___closed__6; @@ -5045,9 +5075,7 @@ static lean_object* l_Lean_Parser_Term_stateRefT___elambda__1___closed__13; static lean_object* l_Lean_Parser_Term_unreachable_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_fromTerm_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_structInstFieldAbbrev_formatter___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__21; static lean_object* l_Lean_Parser_Term_panic_formatter___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__38; static lean_object* l___regBuiltin_Lean_Parser_Term_typeAscription_formatter___closed__1; lean_object* l_Lean_Parser_levelParser_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_attrInstance_formatter___closed__1; @@ -5060,6 +5088,7 @@ static lean_object* l_Lean_Parser_Term_binrel___elambda__1___closed__8; static lean_object* l_Lean_Parser_Term_doubleQuotedName___closed__9; static lean_object* l_Lean_Parser_Term_letrec___elambda__1___closed__3; static lean_object* l_Lean_Parser_Term_forall___elambda__1___closed__3; +static lean_object* l_Lean_Parser_Term_funImplicitBinder___elambda__1___closed__7; static lean_object* l_Lean_Parser_Term_letRecDecl_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_structInstFieldAbbrev___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_namedArgument_formatter___closed__1; @@ -5120,7 +5149,6 @@ static lean_object* l_Lean_Parser_Term_unreachable___elambda__1___closed__8; static lean_object* l_Lean_Parser_Term_waitIfTypeContainsMVar___closed__7; static lean_object* l_Lean_Parser_Term_funBinder___closed__1; static lean_object* l_Lean_Parser_Term_typeAscription___elambda__1___closed__11; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__13; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_syntheticHole_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Term_binrel_formatter___closed__3; static lean_object* l_Lean_Parser_Term_type_formatter___closed__3; @@ -5153,7 +5181,6 @@ static lean_object* l_Lean_Parser_Term_let__delayed___elambda__1___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_simpleBinderWithoutType_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Term_match___closed__11; LEAN_EXPORT lean_object* l_Lean_Parser_Term_arrow_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__41; static lean_object* l_Lean_Parser_Term_ellipsis_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_forInMacro_declRange___closed__1; static lean_object* l_Lean_Parser_Term_forInMacro___elambda__1___closed__12; @@ -5186,6 +5213,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_binop_declRange___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_ensureExpectedType_declRange(lean_object*); static lean_object* l_Lean_Parser_Term_panic___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Term_nomatch_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_Term_letIdBinder_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_namedPattern_formatter___closed__1; static lean_object* l_Lean_Parser_Term_generalizingParam_formatter___closed__8; static lean_object* l_Lean_Parser_Term_letIdLhs_formatter___closed__4; @@ -5211,7 +5239,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_explicit_declRange___closed_ static lean_object* l_Lean_Parser_Term_matchAlt_formatter___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_borrowed_declRange___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_cdot_declRange___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__23; static lean_object* l_Lean_Parser_Term_forInMacro___closed__1; static lean_object* l_Lean_Parser_Term_namedPattern_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Term_macroDollarArg___closed__8; @@ -5232,6 +5259,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_app_declRange___closed__6; lean_object* l_Lean_Parser_nameLit_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_proj_parenthesizer___closed__2; static lean_object* l_Lean_Parser_semicolonOrLinebreak_formatter___closed__3; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__26; static lean_object* l_Lean_Parser_Term_structInstFieldAbbrev___elambda__1___closed__15; static lean_object* l_Lean_Parser_Term_forall___elambda__1___closed__18; static lean_object* l___regBuiltin_Lean_Parser_Term_structInstLVal_formatter___closed__2; @@ -5242,7 +5270,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_namedArgument_parenthesizer_ lean_object* l_Lean_Parser_unicodeSymbol_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_let__fun_formatter___closed__1; static lean_object* l_Lean_Parser_Term_panic___elambda__1___closed__8; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__14; static lean_object* l_Lean_Parser_Term_forall___elambda__1___closed__14; static lean_object* l_Lean_Parser_Term_typeAscription_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Term_doubleQuotedName___closed__1; @@ -5268,6 +5295,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_ensureExpectedType_form static lean_object* l_Lean_Parser_Term_let__delayed___closed__5; static lean_object* l_Lean_Parser_Term_borrowed___closed__6; static lean_object* l_Lean_Parser_Term_binop___closed__5; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__6; static lean_object* l_Lean_Parser_Tactic_tacticSeq___elambda__1___closed__6; static lean_object* l_Lean_Parser_Term_byTactic_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_letDecl_parenthesizer___closed__1; @@ -5283,15 +5311,14 @@ static lean_object* l_Lean_Parser_Term_noImplicitLambda___closed__2; static lean_object* l_Lean_Parser_Term_funSimpleBinder_parenthesizer___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_ensureExpectedType_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_tupleTail_formatter___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__25; static lean_object* l___regBuiltin_Lean_Parser_Term_binop__lazy_declRange___closed__6; static lean_object* l_Lean_Parser_Term_trailing__parser___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_waitIfContainsMVar_declRange___closed__3; static lean_object* l_Lean_Parser_Term_letRecDecl___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_sufficesDecl_formatter___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__30; static lean_object* l_Lean_Parser_Term_structInst___elambda__1___closed__8; LEAN_EXPORT lean_object* l_Lean_Parser_Term_assert_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_let_declRange___closed__3; extern lean_object* l_Lean_PrettyPrinter_Parenthesizer_parenthesizerAliasesRef; static lean_object* l___regBuiltin_Lean_Parser_Term_defaultOrOfNonempty_formatter___closed__2; @@ -5379,7 +5406,6 @@ static lean_object* l_Lean_Parser_Tactic_tacticSeq1Indented_parenthesizer___clos static lean_object* l_Lean_Parser_Term_strictImplicitRightBracket___elambda__1___closed__2; static lean_object* l_Lean_Parser_Tactic_tacticSeqBracketed_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Term_trailing__parser___elambda__1___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__28; static lean_object* l___regBuiltin_Lean_Parser_Term_suffices_declRange___closed__1; static lean_object* l_Lean_Parser_Term_defaultOrOfNonempty___elambda__1___closed__7; static lean_object* l_Lean_Parser_Term_sorry___elambda__1___closed__10; @@ -5406,6 +5432,7 @@ static lean_object* l_Lean_Parser_Term_waitIfTypeMVar___closed__6; static lean_object* l_Lean_Parser_Term_optExprPrecedence_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_noImplicitLambda___closed__1; static lean_object* l_Lean_Parser_Term_sufficesDecl_formatter___closed__5; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__46; static lean_object* l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__24; static lean_object* l_Lean_Parser_Term_structInstField_parenthesizer___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_syntheticHole___elambda__1(lean_object*, lean_object*); @@ -5435,7 +5462,6 @@ static lean_object* l_Lean_Parser_Term_macroDollarArg___elambda__1___closed__9; static lean_object* l_Lean_Parser_Term_haveEqnsDecl___elambda__1___closed__3; static lean_object* l_Lean_Parser_Term_structInstArrayRef___closed__6; static lean_object* l_Lean_Parser_Term_letIdDecl___closed__4; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__33; static lean_object* l___regBuiltin_Lean_Parser_Term_binderDefault_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_borrowed_declRange___closed__4; static lean_object* l_Lean_Parser_Term_binop___elambda__1___closed__1; @@ -5559,7 +5585,6 @@ static lean_object* l_Lean_Parser_Term_leading__parser___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_namedPattern_declRange___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_defaultOrOfNonempty_declRange___closed__1; static lean_object* l_Lean_Parser_Term_noindex___closed__7; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__22; static lean_object* l_Lean_Parser_Term_simpleBinder___elambda__1___closed__3; static lean_object* l_Lean_Parser_Term_simpleBinderWithoutType_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_arrayRef_formatter___closed__1; @@ -5576,6 +5601,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_num_declRange___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_str_declRange___closed__3; static lean_object* l_Lean_Parser_Term_funBinder_quot___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Term_strictImplicitBinder_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Term_letIdBinder___elambda__1___closed__3; static lean_object* l_Lean_Parser_Term_optIdent_parenthesizer___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_structInst(lean_object*); static lean_object* l_Lean_Parser_Term_macroDollarArg___closed__6; @@ -5654,7 +5680,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_binderDefault_formatter(lean_object* static lean_object* l_Lean_Parser_Term_binderDefault_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_bracketedBinder_quot___elambda__1___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_macroDollarArg_parenthesizer___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__29; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_match_declRange(lean_object*); static lean_object* l_Lean_Parser_Term_macroLastArg___closed__3; static lean_object* l_Lean_Parser_Term_nomatch___closed__5; @@ -5672,6 +5697,7 @@ static lean_object* l_Lean_Parser_Term_forInMacro_x27___elambda__1___closed__8; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_attr_quot(lean_object*); static lean_object* l_Lean_Parser_Term_stateRefT_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Term_paren___elambda__1___closed__15; +static lean_object* l_Lean_Parser_Term_funSimpleBinder___closed__5; static lean_object* l_Lean_Parser_Term_parenSpecial_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_local___elambda__1___closed__7; static lean_object* l_Lean_Parser_Term_paren_parenthesizer___closed__4; @@ -5707,6 +5733,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_binop_declRange___closed__2; static lean_object* l_Lean_Parser_Term_noindex_formatter___closed__1; static lean_object* l_Lean_Parser_Term_letMVar___elambda__1___closed__17; static lean_object* l_Lean_Parser_Term_funSimpleBinder_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__36; static lean_object* l___regBuiltin_Lean_Parser_Term_forInMacro_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_haveEqnsDecl_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_prop_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -5748,7 +5775,6 @@ static lean_object* l_Lean_Parser_Term_argument___closed__1; lean_object* l_Lean_PrettyPrinter_Parenthesizer_trailingNode_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_generalizingParam_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_doubleQuotedName___elambda__1___closed__2; -static lean_object* l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__17; static lean_object* l_Lean_Parser_Term_trueVal_parenthesizer___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Term_scientific___elambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_letMVar_parenthesizer(lean_object*); @@ -5775,6 +5801,7 @@ static lean_object* l_Lean_Parser_Term_anonymousCtor_formatter___closed__7; static lean_object* l_Lean_Parser_Term_structInstLVal___elambda__1___closed__11; static lean_object* l_Lean_Parser_Term_borrowed___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_suffices_parenthesizer___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__19; LEAN_EXPORT lean_object* l_Lean_Parser_Term_forInMacro_x27; static lean_object* l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__25; static lean_object* l_Lean_Parser_Term_subst_parenthesizer___closed__3; @@ -5793,6 +5820,7 @@ static lean_object* l_Lean_Parser_Term_letEqnsDecl_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_stateRefT_declRange___closed__1; static lean_object* l_Lean_Parser_Term_bracketedBinder_quot___elambda__1___closed__9; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_waitIfTypeContainsMVar_declRange(lean_object*); +static lean_object* l_Lean_Parser_Term_funBinder___elambda__1___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_binrel_declRange___closed__5; static lean_object* l_Lean_Parser_Term_optEllipsis_parenthesizer___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_attr_quot_declRange___closed__4; @@ -5842,6 +5870,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_leading__parser_declRange___ static lean_object* l_Lean_Parser_Tactic_tacticSeqBracketed_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_noImplicitLambda___elambda__1___closed__8; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_waitIfTypeMVar_parenthesizer(lean_object*); +static lean_object* l_Lean_Parser_Term_funBinder_parenthesizer___closed__8; LEAN_EXPORT lean_object* l_Lean_Parser_Term_inaccessible___elambda__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_matchDiscr___elambda__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Term_matchAlts_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -5891,11 +5920,11 @@ static lean_object* l_Lean_Parser_Term_generalizingParam_parenthesizer___closed_ lean_object* l_Lean_Parser_ParserState_mkUnexpectedError(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_inaccessible_formatter___closed__1; static lean_object* l_Lean_Parser_Term_syntheticHole_parenthesizer___closed__5; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__16; static lean_object* l_Lean_Parser_Term_binderTactic_formatter___closed__6; static lean_object* l_Lean_Parser_Term_forInMacro___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Level_quot(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__34; static lean_object* l_Lean_Parser_Term_nomatch___closed__3; static lean_object* l_Lean_Parser_Term_argument___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_arrayRef_declRange(lean_object*); @@ -5943,6 +5972,7 @@ static lean_object* l_Lean_Parser_Term_let__fun___elambda__1___closed__6; static lean_object* l_Lean_Parser_Term_whereDecls___elambda__1___closed__13; static lean_object* l_Lean_Parser_Term_optEllipsis___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_let__tmp_formatter___closed__2; +LEAN_EXPORT lean_object* l_Lean_Parser_Term_letIdBinder_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_borrowed___elambda__1___closed__6; static lean_object* l_Lean_Parser_Term_letRecDecl___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_trueVal_formatter___closed__2; @@ -5985,7 +6015,7 @@ static lean_object* l_Lean_Parser_Term_show_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_trueVal___elambda__1___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_sorry_declRange___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_63_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883_(lean_object*); static lean_object* l_Lean_Parser_Term_defaultOrOfNonempty___elambda__1___closed__4; static lean_object* l_Lean_Parser_Term_noindex___elambda__1___closed__5; static lean_object* l_Lean_Parser_Term_typeAscription___elambda__1___closed__9; @@ -5999,8 +6029,10 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_let(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_match_formatter(lean_object*); static lean_object* l_Lean_Parser_Term_let__fun_parenthesizer___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_doubleQuotedName_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__13; static lean_object* l_Lean_Parser_Term_letDecl___elambda__1___closed__4; static lean_object* l_Lean_Parser_Term_ellipsis___elambda__1___closed__3; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__41; static lean_object* l_Lean_Parser_Term_haveIdDecl_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_depArrow_formatter___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_arrayRef_formatter___closed__1; @@ -6009,7 +6041,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_dbgTrace_formatter(lean_object*, lea LEAN_EXPORT lean_object* l_Lean_Parser_Term_binop__lazy; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_pipeCompletion_formatter(lean_object*); static lean_object* l_Lean_Parser_Term_tupleTail_formatter___closed__1; -uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_structInstFieldAbbrev___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Tactic_quot_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_matchAlt___closed__6; @@ -6039,6 +6070,7 @@ static lean_object* l_Lean_Parser_Term_forall___elambda__1___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_pipeProj(lean_object*); static lean_object* l_Lean_Parser_Term_binderDefault___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_typeAscription_formatter(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_Term_instCoeTSyntaxConsSyntaxNodeKindMkStrAnonymousNil(lean_object*); static lean_object* l_Lean_Parser_Term_cdot___elambda__1___closed__6; static lean_object* l_Lean_Parser_Term_borrowed___elambda__1___closed__5; static lean_object* l_Lean_Parser_Term_leading__parser_formatter___closed__1; @@ -6056,6 +6088,7 @@ static lean_object* l_Lean_Parser_Term_structInst___closed__11; static lean_object* l_Lean_Parser_Term_sorry___elambda__1___closed__5; static lean_object* l_Lean_Parser_Term_forall___elambda__1___closed__8; static lean_object* l_Lean_Parser_Term_optExprPrecedence_parenthesizer___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__23; static lean_object* l___regBuiltin_Lean_Parser_Term_stateRefT_declRange___closed__4; static lean_object* l_Lean_Parser_Tactic_quotSeq___closed__2; static lean_object* l_Lean_Parser_Term_have_parenthesizer___closed__5; @@ -6099,7 +6132,6 @@ static lean_object* l_Lean_Parser_Term_structInstLVal_formatter___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Tactic_quot_declRange(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_waitIfTypeMVar_declRange___closed__6; static lean_object* l_Lean_Parser_Term_leading__parser___elambda__1___closed__14; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__46; static lean_object* l_Lean_Parser_Term_inaccessible_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_pipeCompletion_declRange___closed__5; static lean_object* l_Lean_Parser_Term_fun___elambda__1___closed__10; @@ -6114,7 +6146,9 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_haveIdDecl_parenthesizer___c static lean_object* l___regBuiltin_Lean_Parser_Term_ident_declRange___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_dotIdent_formatter___closed__2; static lean_object* l_Lean_Parser_Term_namedPattern_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__38; LEAN_EXPORT lean_object* l_Lean_Parser_Term_macroLastArg_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__21; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_num_declRange(lean_object*); static lean_object* l_Lean_Parser_Term_forall___elambda__1___closed__24; static lean_object* l_Lean_Parser_Term_waitIfTypeMVar___elambda__1___closed__3; @@ -6173,6 +6207,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_subst_declRange___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_forall_formatter___closed__2; static lean_object* l_Lean_Parser_Term_fromTerm___elambda__1___closed__10; static lean_object* l_Lean_Parser_Term_waitIfContainsMVar___elambda__1___closed__9; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__2; static lean_object* l_Lean_Parser_Term_namedPattern_formatter___closed__6; static lean_object* l_Lean_Parser_Term_pipeProj___closed__1; static lean_object* l_Lean_Parser_Term_letIdLhs_formatter___closed__6; @@ -6311,6 +6346,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Level_quot_parenthesizer(lean_object*, le lean_object* l_Lean_Parser_ppDedent_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_doubleQuotedName_formatter___closed__6; static lean_object* l_Lean_Parser_Term_waitIfTypeMVar___closed__2; +LEAN_EXPORT lean_object* l_Lean_Parser_Term_instCoeTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___boxed(lean_object*); static lean_object* l_Lean_Parser_Term_funImplicitBinder___closed__6; static lean_object* l_Lean_Parser_Term_typeOf___closed__6; static lean_object* l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__8; @@ -6333,6 +6369,7 @@ static lean_object* l_Lean_Parser_Term_prop___closed__6; static lean_object* l_Lean_Parser_Term_forInMacro___elambda__1___closed__10; static lean_object* l_Lean_Parser_Term_funBinder_quot___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Term_haveIdLhs_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__47; static lean_object* l_Lean_Parser_Term_structInstFieldAbbrev___closed__4; static lean_object* l_Lean_Parser_Term_waitIfContainsMVar___elambda__1___closed__8; static lean_object* l_Lean_Parser_Term_letRecDecl_parenthesizer___closed__2; @@ -6375,6 +6412,7 @@ static lean_object* l_Lean_Parser_Term_letIdLhs_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_attributes_formatter___closed__3; static lean_object* l_Lean_Parser_Term_letEqnsDecl___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_dotIdent_declRange___closed__6; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__35; LEAN_EXPORT lean_object* l_Lean_Parser_Term_simpleBinderWithoutType_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_whereDecls_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_prop___elambda__1(lean_object*, lean_object*); @@ -6444,6 +6482,7 @@ uint8_t l_Lean_Syntax_isIdent(lean_object*); static lean_object* l_Lean_Parser_Term_proj___elambda__1___closed__3; static lean_object* l_Lean_Parser_Term_letIdDecl___closed__6; lean_object* l_Lean_PrettyPrinter_Formatter_categoryParser_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Term_funImplicitBinder_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Tactic_seq1_formatter___closed__2; static lean_object* l_Lean_Parser_Term_letDecl_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Term_defaultOrOfNonempty_formatter___closed__6; @@ -6470,7 +6509,7 @@ x_5 = l_Lean_Parser_finishCommentBlock(x_4, x_1, x_2); x_6 = lean_ctor_get(x_5, 4); lean_inc(x_6); x_7 = lean_box(0); -x_8 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_6, x_7); +x_8 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_6, x_7); lean_dec(x_6); if (x_8 == 0) { @@ -6858,7 +6897,7 @@ x_22 = l_Lean_Parser_checkPrecFn(x_21, x_1, x_2); x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); x_24 = lean_box(0); -x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_24); +x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_24); lean_dec(x_23); if (x_25 == 0) { @@ -6901,7 +6940,7 @@ goto block_46; lean_object* x_35; uint8_t x_36; x_35 = lean_ctor_get(x_34, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_24); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_24); lean_dec(x_35); if (x_36 == 0) { @@ -6909,7 +6948,7 @@ lean_object* x_37; lean_object* x_38; uint8_t x_39; x_37 = l_Lean_Parser_ParserState_mkNode(x_34, x_7, x_27); x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_24); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_24); lean_dec(x_38); if (x_39 == 0) { @@ -6931,7 +6970,7 @@ x_41 = l_Lean_Parser_Command_commentBody___elambda__1(x_1, x_34); x_42 = l_Lean_Parser_ParserState_mkNode(x_41, x_7, x_27); x_43 = lean_ctor_get(x_42, 4); lean_inc(x_43); -x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_24); +x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_24); lean_dec(x_43); if (x_44 == 0) { @@ -7457,18 +7496,36 @@ return x_3; static lean_object* _init_l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__18() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("group", 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__19() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_groupKind; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__18; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__19; x_2 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__16; x_3 = l_Lean_Parser_nodeInfo(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__19() { +static lean_object* _init_l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_groupKind; +x_1 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__19; x_2 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__17; x_3 = lean_alloc_closure((void*)(l_Lean_Parser_nodeFn), 4, 2); lean_closure_set(x_3, 0, x_1); @@ -7476,7 +7533,7 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__20() { +static lean_object* _init_l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__22() { _start: { lean_object* x_1; @@ -7484,55 +7541,55 @@ x_1 = lean_mk_string_from_bytes("irrelevant", 10); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__21() { +static lean_object* _init_l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__23() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__20; +x_1 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__22; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_checkColGeFn___boxed), 3, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__22() { +static lean_object* _init_l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_commentBody___closed__2; -x_2 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__18; +x_2 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__20; x_3 = l_Lean_Parser_andthenInfo(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__23() { +static lean_object* _init_l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__21; -x_2 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__19; +x_1 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__23; +x_2 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__21; x_3 = lean_alloc_closure((void*)(l_Lean_Parser_andthenFn), 4, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__24() { +static lean_object* _init_l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__22; -x_2 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__23; +x_1 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__24; +x_2 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__25; x_3 = lean_alloc_ctor(0, 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_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__25() { +static lean_object* _init_l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__27() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__24; +x_1 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__26; x_2 = l_Lean_Parser_many1(x_1); return x_2; } @@ -7541,7 +7598,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1(le _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; 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; uint8_t x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint32_t x_22; uint32_t x_23; uint8_t x_24; -x_3 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__25; +x_3 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__27; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___lambda__1), 3, 1); lean_closure_set(x_4, 0, x_3); x_5 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__4; @@ -7591,7 +7648,7 @@ x_26 = l_Lean_Parser_checkPrecFn(x_25, x_1, x_2); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); x_28 = lean_box(0); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_28); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_28); lean_dec(x_27); if (x_29 == 0) { @@ -7630,7 +7687,7 @@ x_36 = lean_apply_2(x_32, x_35, x_26); x_37 = l_Lean_Parser_ParserState_mkNode(x_36, x_5, x_31); x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_28); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_28); lean_dec(x_38); if (x_39 == 0) { @@ -7665,7 +7722,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticSeq1Indented___closed__1() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__25; +x_1 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__27; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__4; @@ -7794,8 +7851,8 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticSeqBracketed___elambda__1__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__18; -x_2 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__19; +x_1 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__20; +x_2 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__21; x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -7982,7 +8039,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -8028,7 +8085,7 @@ x_28 = l_Lean_Parser_Tactic_tacticSeqBracketed___elambda__1___closed__2; x_29 = l_Lean_Parser_ParserState_mkNode(x_27, x_28, x_19); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_16); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_16); lean_dec(x_30); if (x_31 == 0) { @@ -8048,7 +8105,7 @@ return x_32; lean_object* x_35; uint8_t x_36; x_35 = lean_ctor_get(x_34, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_16); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_16); lean_dec(x_35); if (x_36 == 0) { @@ -8064,7 +8121,7 @@ lean_inc(x_1); x_37 = lean_apply_2(x_4, x_1, x_34); x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_16); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_16); lean_dec(x_38); if (x_39 == 0) { @@ -8341,7 +8398,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -8364,7 +8421,7 @@ x_22 = l_Lean_Parser_Tactic_tacticSeq___elambda__1___closed__2; x_23 = l_Lean_Parser_ParserState_mkNode(x_21, x_22, x_17); x_24 = lean_ctor_get(x_23, 4); lean_inc(x_24); -x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_14); +x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_14); lean_dec(x_24); if (x_25 == 0) { @@ -9017,7 +9074,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -9048,7 +9105,7 @@ if (x_31 == 0) lean_object* x_32; uint8_t x_33; x_32 = lean_ctor_get(x_27, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_14); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_14); lean_dec(x_32); if (x_33 == 0) { @@ -9071,7 +9128,7 @@ lean_inc(x_1); x_35 = l_Lean_Parser_tokenAntiquotFn(x_1, x_27); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); lean_dec(x_36); if (x_37 == 0) { @@ -9094,7 +9151,7 @@ x_19 = l_Lean_Parser_Term_byTactic___elambda__1___closed__4; x_20 = l_Lean_Parser_ParserState_mkNode(x_18, x_19, x_17); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); -x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_14); +x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_14); lean_dec(x_21); if (x_22 == 0) { @@ -10514,7 +10571,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -10545,7 +10602,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_20, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -10554,7 +10611,7 @@ x_27 = l_Lean_Parser_Term_byTactic_x27___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_20, x_27, x_17); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -10578,7 +10635,7 @@ x_33 = l_Lean_Parser_Term_byTactic_x27___elambda__1___closed__2; x_34 = l_Lean_Parser_ParserState_mkNode(x_32, x_33, x_17); x_35 = lean_ctor_get(x_34, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_14); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_14); lean_dec(x_35); if (x_36 == 0) { @@ -10601,7 +10658,7 @@ lean_inc(x_1); x_38 = l_Lean_Parser_tokenAntiquotFn(x_1, x_20); x_39 = lean_ctor_get(x_38, 4); lean_inc(x_39); -x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); +x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); lean_dec(x_39); if (x_40 == 0) { @@ -10610,7 +10667,7 @@ x_41 = l_Lean_Parser_Term_byTactic_x27___elambda__1___closed__2; x_42 = l_Lean_Parser_ParserState_mkNode(x_38, x_41, x_17); x_43 = lean_ctor_get(x_42, 4); lean_inc(x_43); -x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_14); +x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_14); lean_dec(x_43); if (x_44 == 0) { @@ -10634,7 +10691,7 @@ x_47 = l_Lean_Parser_Term_byTactic_x27___elambda__1___closed__2; x_48 = l_Lean_Parser_ParserState_mkNode(x_46, x_47, x_17); x_49 = lean_ctor_get(x_48, 4); lean_inc(x_49); -x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_14); +x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_14); lean_dec(x_49); if (x_50 == 0) { @@ -10742,7 +10799,7 @@ x_4 = l_Lean_Parser_semicolonOrLinebreak___elambda__1(x_2, x_3); x_5 = lean_ctor_get(x_4, 4); lean_inc(x_5); x_6 = lean_box(0); -x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); +x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); lean_dec(x_5); if (x_7 == 0) { @@ -10794,7 +10851,7 @@ x_4 = l_Lean_Parser_checkPrecFn(x_3, x_1, x_2); x_5 = lean_ctor_get(x_4, 4); lean_inc(x_5); x_6 = lean_box(0); -x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); +x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); lean_dec(x_5); if (x_7 == 0) { @@ -11045,7 +11102,7 @@ x_4 = l_Lean_Parser_checkPrecFn(x_3, x_1, x_2); x_5 = lean_ctor_get(x_4, 4); lean_inc(x_5); x_6 = lean_box(0); -x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); +x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); lean_dec(x_5); if (x_7 == 0) { @@ -11278,7 +11335,7 @@ x_4 = l_Lean_Parser_checkPrecFn(x_3, x_1, x_2); x_5 = lean_ctor_get(x_4, 4); lean_inc(x_5); x_6 = lean_box(0); -x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); +x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); lean_dec(x_5); if (x_7 == 0) { @@ -11511,7 +11568,7 @@ x_4 = l_Lean_Parser_checkPrecFn(x_3, x_1, x_2); x_5 = lean_ctor_get(x_4, 4); lean_inc(x_5); x_6 = lean_box(0); -x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); +x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); lean_dec(x_5); if (x_7 == 0) { @@ -11744,7 +11801,7 @@ x_4 = l_Lean_Parser_checkPrecFn(x_3, x_1, x_2); x_5 = lean_ctor_get(x_4, 4); lean_inc(x_5); x_6 = lean_box(0); -x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); +x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); lean_dec(x_5); if (x_7 == 0) { @@ -12290,7 +12347,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -12322,7 +12379,7 @@ if (x_26 == 0) lean_object* x_27; uint8_t x_28; x_27 = lean_ctor_get(x_22, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); lean_dec(x_27); if (x_28 == 0) { @@ -12332,7 +12389,7 @@ x_29 = l_Lean_Parser_Term_type___elambda__1___closed__2; x_30 = l_Lean_Parser_ParserState_mkNode(x_22, x_29, x_19); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); lean_dec(x_31); if (x_32 == 0) { @@ -12356,7 +12413,7 @@ x_35 = l_Lean_Parser_Term_type___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_19); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_16); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_16); lean_dec(x_37); if (x_38 == 0) { @@ -12379,7 +12436,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_22); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); lean_dec(x_41); if (x_42 == 0) { @@ -12389,7 +12446,7 @@ x_43 = l_Lean_Parser_Term_type___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_19); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); lean_dec(x_45); if (x_46 == 0) { @@ -12413,7 +12470,7 @@ x_49 = l_Lean_Parser_Term_type___elambda__1___closed__2; x_50 = l_Lean_Parser_ParserState_mkNode(x_48, x_49, x_19); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); lean_dec(x_51); if (x_52 == 0) { @@ -13146,7 +13203,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -13178,7 +13235,7 @@ if (x_26 == 0) lean_object* x_27; uint8_t x_28; x_27 = lean_ctor_get(x_22, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); lean_dec(x_27); if (x_28 == 0) { @@ -13188,7 +13245,7 @@ x_29 = l_Lean_Parser_Term_sort___elambda__1___closed__2; x_30 = l_Lean_Parser_ParserState_mkNode(x_22, x_29, x_19); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); lean_dec(x_31); if (x_32 == 0) { @@ -13212,7 +13269,7 @@ x_35 = l_Lean_Parser_Term_sort___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_19); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_16); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_16); lean_dec(x_37); if (x_38 == 0) { @@ -13235,7 +13292,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_22); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); lean_dec(x_41); if (x_42 == 0) { @@ -13245,7 +13302,7 @@ x_43 = l_Lean_Parser_Term_sort___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_19); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); lean_dec(x_45); if (x_46 == 0) { @@ -13269,7 +13326,7 @@ x_49 = l_Lean_Parser_Term_sort___elambda__1___closed__2; x_50 = l_Lean_Parser_ParserState_mkNode(x_48, x_49, x_19); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); lean_dec(x_51); if (x_52 == 0) { @@ -13833,7 +13890,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -13866,7 +13923,7 @@ x_25 = l_Lean_Parser_Term_prop___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_20, x_25, x_17); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); lean_dec(x_27); if (x_28 == 0) { @@ -13890,7 +13947,7 @@ x_31 = l_Lean_Parser_Term_prop___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_17); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -14416,7 +14473,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -14449,7 +14506,7 @@ x_25 = l_Lean_Parser_Term_hole___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_20, x_25, x_17); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); lean_dec(x_27); if (x_28 == 0) { @@ -14473,7 +14530,7 @@ x_31 = l_Lean_Parser_Term_hole___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_17); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -15023,7 +15080,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -15054,7 +15111,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_20, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -15063,7 +15120,7 @@ x_27 = l_Lean_Parser_Term_syntheticHole___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_20, x_27, x_17); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -15090,7 +15147,7 @@ x_36 = l_Lean_Parser_Term_syntheticHole___elambda__1___closed__2; x_37 = l_Lean_Parser_ParserState_mkNode(x_35, x_36, x_17); x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_14); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_14); lean_dec(x_38); if (x_39 == 0) { @@ -15113,7 +15170,7 @@ lean_inc(x_1); x_41 = l_Lean_Parser_tokenAntiquotFn(x_1, x_20); x_42 = lean_ctor_get(x_41, 4); lean_inc(x_42); -x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_14); +x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_14); lean_dec(x_42); if (x_43 == 0) { @@ -15122,7 +15179,7 @@ x_44 = l_Lean_Parser_Term_syntheticHole___elambda__1___closed__2; x_45 = l_Lean_Parser_ParserState_mkNode(x_41, x_44, x_17); x_46 = lean_ctor_get(x_45, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_14); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_14); lean_dec(x_46); if (x_47 == 0) { @@ -15149,7 +15206,7 @@ x_53 = l_Lean_Parser_Term_syntheticHole___elambda__1___closed__2; x_54 = l_Lean_Parser_ParserState_mkNode(x_52, x_53, x_17); x_55 = lean_ctor_get(x_54, 4); lean_inc(x_55); -x_56 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_55, x_14); +x_56 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_55, x_14); lean_dec(x_55); if (x_56 == 0) { @@ -15756,7 +15813,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -15789,7 +15846,7 @@ x_25 = l_Lean_Parser_Term_sorry___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_20, x_25, x_17); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); lean_dec(x_27); if (x_28 == 0) { @@ -15813,7 +15870,7 @@ x_31 = l_Lean_Parser_Term_sorry___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_17); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -16359,7 +16416,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -16382,7 +16439,7 @@ x_22 = l_Lean_Parser_Term_cdot___elambda__1___closed__2; x_23 = l_Lean_Parser_ParserState_mkNode(x_21, x_22, x_17); x_24 = lean_ctor_get(x_23, 4); lean_inc(x_24); -x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_14); +x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_14); lean_dec(x_24); if (x_25 == 0) { @@ -16993,7 +17050,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -17024,7 +17081,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_20, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -17033,7 +17090,7 @@ x_27 = l_Lean_Parser_Term_typeAscription___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_20, x_27, x_17); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -17059,7 +17116,7 @@ x_35 = l_Lean_Parser_Term_typeAscription___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_17); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); lean_dec(x_37); if (x_38 == 0) { @@ -17082,7 +17139,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_20); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); lean_dec(x_41); if (x_42 == 0) { @@ -17091,7 +17148,7 @@ x_43 = l_Lean_Parser_Term_typeAscription___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_17); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); lean_dec(x_45); if (x_46 == 0) { @@ -17117,7 +17174,7 @@ x_51 = l_Lean_Parser_Term_typeAscription___elambda__1___closed__2; x_52 = l_Lean_Parser_ParserState_mkNode(x_50, x_51, x_17); x_53 = lean_ctor_get(x_52, 4); lean_inc(x_53); -x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); +x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); lean_dec(x_53); if (x_54 == 0) { @@ -17437,7 +17494,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -17469,7 +17526,7 @@ if (x_26 == 0) lean_object* x_27; uint8_t x_28; x_27 = lean_ctor_get(x_22, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); lean_dec(x_27); if (x_28 == 0) { @@ -17479,7 +17536,7 @@ x_29 = l_Lean_Parser_Term_tupleTail___elambda__1___closed__2; x_30 = l_Lean_Parser_ParserState_mkNode(x_22, x_29, x_19); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); lean_dec(x_31); if (x_32 == 0) { @@ -17503,7 +17560,7 @@ x_35 = l_Lean_Parser_Term_tupleTail___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_19); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_16); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_16); lean_dec(x_37); if (x_38 == 0) { @@ -17526,7 +17583,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_22); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); lean_dec(x_41); if (x_42 == 0) { @@ -17536,7 +17593,7 @@ x_43 = l_Lean_Parser_Term_tupleTail___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_19); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); lean_dec(x_45); if (x_46 == 0) { @@ -17560,7 +17617,7 @@ x_49 = l_Lean_Parser_Term_tupleTail___elambda__1___closed__2; x_50 = l_Lean_Parser_ParserState_mkNode(x_48, x_49, x_19); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); lean_dec(x_51); if (x_52 == 0) { @@ -18013,7 +18070,7 @@ x_29 = l_Lean_Parser_checkPrecFn(x_28, x_1, x_2); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); x_31 = lean_box(0); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_31); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_31); lean_dec(x_30); if (x_32 == 0) { @@ -18061,7 +18118,7 @@ goto block_73; lean_object* x_43; uint8_t x_44; x_43 = lean_ctor_get(x_42, 4); lean_inc(x_43); -x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_31); +x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_31); lean_dec(x_43); if (x_44 == 0) { @@ -18075,7 +18132,7 @@ lean_dec(x_17); x_45 = l_Lean_Parser_ParserState_mkNode(x_42, x_9, x_34); x_46 = lean_ctor_get(x_45, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_31); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_31); lean_dec(x_46); if (x_47 == 0) { @@ -18107,7 +18164,7 @@ lean_ctor_set_uint8(x_50, sizeof(void*)*7, x_22); x_51 = lean_apply_2(x_49, x_50, x_42); x_52 = lean_ctor_get(x_51, 4); lean_inc(x_52); -x_53 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_52, x_31); +x_53 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_52, x_31); lean_dec(x_52); if (x_53 == 0) { @@ -18116,7 +18173,7 @@ lean_dec(x_23); x_54 = l_Lean_Parser_ParserState_mkNode(x_51, x_9, x_34); x_55 = lean_ctor_get(x_54, 4); lean_inc(x_55); -x_56 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_55, x_31); +x_56 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_55, x_31); lean_dec(x_55); if (x_56 == 0) { @@ -18150,7 +18207,7 @@ lean_object* x_64; lean_object* x_65; uint8_t x_66; x_64 = l_Lean_Parser_ParserState_mkNode(x_60, x_9, x_34); x_65 = lean_ctor_get(x_64, 4); lean_inc(x_65); -x_66 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_65, x_31); +x_66 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_65, x_31); lean_dec(x_65); if (x_66 == 0) { @@ -18173,7 +18230,7 @@ x_68 = l_Lean_Parser_tokenAntiquotFn(x_1, x_60); x_69 = l_Lean_Parser_ParserState_mkNode(x_68, x_9, x_34); x_70 = lean_ctor_get(x_69, 4); lean_inc(x_70); -x_71 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_70, x_31); +x_71 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_70, x_31); lean_dec(x_70); if (x_71 == 0) { @@ -19483,7 +19540,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -19527,7 +19584,7 @@ goto block_60; lean_object* x_28; uint8_t x_29; x_28 = lean_ctor_get(x_27, 4); lean_inc(x_28); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_16); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_16); lean_dec(x_28); if (x_29 == 0) { @@ -19538,7 +19595,7 @@ x_30 = l_Lean_Parser_Term_anonymousCtor___elambda__1___closed__2; x_31 = l_Lean_Parser_ParserState_mkNode(x_27, x_30, x_19); x_32 = lean_ctor_get(x_31, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_16); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_16); lean_dec(x_32); if (x_33 == 0) { @@ -19560,7 +19617,7 @@ lean_inc(x_1); x_35 = lean_apply_2(x_4, x_1, x_27); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_16); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_16); lean_dec(x_36); if (x_37 == 0) { @@ -19570,7 +19627,7 @@ x_38 = l_Lean_Parser_Term_anonymousCtor___elambda__1___closed__2; x_39 = l_Lean_Parser_ParserState_mkNode(x_35, x_38, x_19); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_16); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_16); lean_dec(x_40); if (x_41 == 0) { @@ -19605,7 +19662,7 @@ x_49 = l_Lean_Parser_Term_anonymousCtor___elambda__1___closed__2; x_50 = l_Lean_Parser_ParserState_mkNode(x_45, x_49, x_19); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); lean_dec(x_51); if (x_52 == 0) { @@ -19629,7 +19686,7 @@ x_55 = l_Lean_Parser_Term_anonymousCtor___elambda__1___closed__2; x_56 = l_Lean_Parser_ParserState_mkNode(x_54, x_55, x_19); x_57 = lean_ctor_get(x_56, 4); lean_inc(x_57); -x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_16); +x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_16); lean_dec(x_57); if (x_58 == 0) { @@ -20367,7 +20424,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -20398,7 +20455,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_20, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -20407,7 +20464,7 @@ x_27 = l_Lean_Parser_Term_fromTerm___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_20, x_27, x_17); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -20433,7 +20490,7 @@ x_35 = l_Lean_Parser_Term_fromTerm___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_17); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); lean_dec(x_37); if (x_38 == 0) { @@ -20456,7 +20513,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_20); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); lean_dec(x_41); if (x_42 == 0) { @@ -20465,7 +20522,7 @@ x_43 = l_Lean_Parser_Term_fromTerm___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_17); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); lean_dec(x_45); if (x_46 == 0) { @@ -20491,7 +20548,7 @@ x_51 = l_Lean_Parser_Term_fromTerm___elambda__1___closed__2; x_52 = l_Lean_Parser_ParserState_mkNode(x_50, x_51, x_17); x_53 = lean_ctor_get(x_52, 4); lean_inc(x_53); -x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); +x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); lean_dec(x_53); if (x_54 == 0) { @@ -20799,7 +20856,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -20818,7 +20875,7 @@ lean_inc(x_1); x_20 = lean_apply_2(x_4, x_1, x_14); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); -x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_16); +x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_16); lean_dec(x_21); if (x_22 == 0) { @@ -20827,7 +20884,7 @@ x_23 = l_Lean_Parser_Term_sufficesDecl___elambda__1___closed__2; x_24 = l_Lean_Parser_ParserState_mkNode(x_20, x_23, x_19); x_25 = lean_ctor_get(x_24, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_16); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_16); lean_dec(x_25); if (x_26 == 0) { @@ -20851,7 +20908,7 @@ lean_inc(x_1); x_30 = l_Lean_Parser_categoryParser___elambda__1(x_28, x_29, x_1, x_20); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); lean_dec(x_31); if (x_32 == 0) { @@ -20860,7 +20917,7 @@ x_33 = l_Lean_Parser_Term_sufficesDecl___elambda__1___closed__2; x_34 = l_Lean_Parser_ParserState_mkNode(x_30, x_33, x_19); x_35 = lean_ctor_get(x_34, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_16); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_16); lean_dec(x_35); if (x_36 == 0) { @@ -20887,7 +20944,7 @@ x_42 = l_Lean_Parser_Term_sufficesDecl___elambda__1___closed__2; x_43 = l_Lean_Parser_ParserState_mkNode(x_41, x_42, x_19); x_44 = lean_ctor_get(x_43, 4); lean_inc(x_44); -x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_16); +x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_16); lean_dec(x_44); if (x_45 == 0) { @@ -21060,7 +21117,7 @@ if (x_17 == 0) lean_object* x_19; uint8_t x_20; x_19 = lean_ctor_get(x_12, 4); lean_inc(x_19); -x_20 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_18); +x_20 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_18); lean_dec(x_19); if (x_20 == 0) { @@ -21081,7 +21138,7 @@ lean_inc(x_2); x_22 = l_Lean_Parser_tokenAntiquotFn(x_2, x_12); x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); -x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_18); +x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_18); lean_dec(x_23); if (x_24 == 0) { @@ -21148,7 +21205,7 @@ if (x_44 == 0) lean_object* x_46; uint8_t x_47; x_46 = lean_ctor_get(x_39, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_45); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_45); lean_dec(x_46); if (x_47 == 0) { @@ -21169,7 +21226,7 @@ lean_inc(x_35); x_49 = l_Lean_Parser_tokenAntiquotFn(x_35, x_39); x_50 = lean_ctor_get(x_49, 4); lean_inc(x_50); -x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_45); +x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_45); lean_dec(x_50); if (x_51 == 0) { @@ -21322,7 +21379,7 @@ x_29 = l_Lean_Parser_checkPrecFn(x_28, x_1, x_2); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); x_31 = lean_box(0); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_31); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_31); lean_dec(x_30); if (x_32 == 0) { @@ -21372,7 +21429,7 @@ if (x_43 == 0) lean_object* x_57; uint8_t x_58; x_57 = lean_ctor_get(x_39, 4); lean_inc(x_57); -x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_31); +x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_31); lean_dec(x_57); if (x_58 == 0) { @@ -21395,7 +21452,7 @@ lean_inc(x_37); x_60 = l_Lean_Parser_tokenAntiquotFn(x_37, x_39); x_61 = lean_ctor_get(x_60, 4); lean_inc(x_61); -x_62 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_61, x_31); +x_62 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_61, x_31); lean_dec(x_61); if (x_62 == 0) { @@ -21416,7 +21473,7 @@ goto block_56; lean_object* x_45; uint8_t x_46; x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_31); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_31); lean_dec(x_45); if (x_46 == 0) { @@ -21425,7 +21482,7 @@ lean_dec(x_6); x_47 = l_Lean_Parser_ParserState_mkNode(x_44, x_8, x_34); x_48 = lean_ctor_get(x_47, 4); lean_inc(x_48); -x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_31); +x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_31); lean_dec(x_48); if (x_49 == 0) { @@ -21448,7 +21505,7 @@ x_51 = lean_apply_2(x_6, x_1, x_44); x_52 = l_Lean_Parser_ParserState_mkNode(x_51, x_8, x_34); x_53 = lean_ctor_get(x_52, 4); lean_inc(x_53); -x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_31); +x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_31); lean_dec(x_53); if (x_54 == 0) { @@ -22913,7 +22970,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -22957,7 +23014,7 @@ goto block_52; lean_object* x_26; uint8_t x_27; x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); lean_dec(x_26); if (x_27 == 0) { @@ -22966,7 +23023,7 @@ x_28 = l_Lean_Parser_Term_show___elambda__1___closed__2; x_29 = l_Lean_Parser_ParserState_mkNode(x_25, x_28, x_17); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); lean_dec(x_30); if (x_31 == 0) { @@ -22990,7 +23047,7 @@ lean_inc(x_1); x_35 = l_Lean_Parser_categoryParser___elambda__1(x_33, x_34, x_1, x_25); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); lean_dec(x_36); if (x_37 == 0) { @@ -22999,7 +23056,7 @@ x_38 = l_Lean_Parser_Term_show___elambda__1___closed__2; x_39 = l_Lean_Parser_ParserState_mkNode(x_35, x_38, x_17); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); lean_dec(x_40); if (x_41 == 0) { @@ -23026,7 +23083,7 @@ x_47 = l_Lean_Parser_Term_show___elambda__1___closed__2; x_48 = l_Lean_Parser_ParserState_mkNode(x_46, x_47, x_17); x_49 = lean_ctor_get(x_48, 4); lean_inc(x_49); -x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_14); +x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_14); lean_dec(x_49); if (x_50 == 0) { @@ -23659,7 +23716,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -23702,7 +23759,7 @@ goto block_60; lean_object* x_26; uint8_t x_27; x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); lean_dec(x_26); if (x_27 == 0) { @@ -23712,7 +23769,7 @@ x_28 = l_Lean_Parser_Term_structInstArrayRef___elambda__1___closed__2; x_29 = l_Lean_Parser_ParserState_mkNode(x_25, x_28, x_17); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); lean_dec(x_30); if (x_31 == 0) { @@ -23736,7 +23793,7 @@ lean_inc(x_1); x_35 = l_Lean_Parser_categoryParser___elambda__1(x_33, x_34, x_1, x_25); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); lean_dec(x_36); if (x_37 == 0) { @@ -23746,7 +23803,7 @@ x_38 = l_Lean_Parser_Term_structInstArrayRef___elambda__1___closed__2; x_39 = l_Lean_Parser_ParserState_mkNode(x_35, x_38, x_17); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); lean_dec(x_40); if (x_41 == 0) { @@ -23781,7 +23838,7 @@ x_49 = l_Lean_Parser_Term_structInstArrayRef___elambda__1___closed__2; x_50 = l_Lean_Parser_ParserState_mkNode(x_45, x_49, x_17); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_14); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_14); lean_dec(x_51); if (x_52 == 0) { @@ -23805,7 +23862,7 @@ x_55 = l_Lean_Parser_Term_structInstArrayRef___elambda__1___closed__2; x_56 = l_Lean_Parser_ParserState_mkNode(x_54, x_55, x_17); x_57 = lean_ctor_get(x_56, 4); lean_inc(x_57); -x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_14); +x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_14); lean_dec(x_57); if (x_58 == 0) { @@ -24123,7 +24180,7 @@ static lean_object* _init_l_Lean_Parser_Term_structInstLVal___elambda__1___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_groupKind; +x_1 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__19; x_2 = l_Lean_Parser_Term_structInstLVal___elambda__1___closed__15; x_3 = l_Lean_Parser_nodeInfo(x_1, x_2); return x_3; @@ -24133,7 +24190,7 @@ static lean_object* _init_l_Lean_Parser_Term_structInstLVal___elambda__1___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_groupKind; +x_1 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__19; x_2 = l_Lean_Parser_Term_structInstLVal___elambda__1___closed__16; x_3 = lean_alloc_closure((void*)(l_Lean_Parser_nodeFn), 4, 2); lean_closure_set(x_3, 0, x_1); @@ -24267,7 +24324,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -24289,7 +24346,7 @@ lean_inc(x_1); x_23 = l_Lean_Parser_orelseFnCore(x_20, x_21, x_22, x_1, x_14); x_24 = lean_ctor_get(x_23, 4); lean_inc(x_24); -x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_16); +x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_16); lean_dec(x_24); if (x_25 == 0) { @@ -24299,7 +24356,7 @@ x_26 = l_Lean_Parser_Term_structInstLVal___elambda__1___closed__2; x_27 = l_Lean_Parser_ParserState_mkNode(x_23, x_26, x_19); x_28 = lean_ctor_get(x_27, 4); lean_inc(x_28); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_16); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_16); lean_dec(x_28); if (x_29 == 0) { @@ -24323,7 +24380,7 @@ x_32 = l_Lean_Parser_Term_structInstLVal___elambda__1___closed__2; x_33 = l_Lean_Parser_ParserState_mkNode(x_31, x_32, x_19); x_34 = lean_ctor_get(x_33, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_16); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_16); lean_dec(x_34); if (x_35 == 0) { @@ -24621,7 +24678,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -24640,7 +24697,7 @@ lean_inc(x_1); x_18 = l_Lean_Parser_Term_structInstLVal___elambda__1(x_1, x_12); x_19 = lean_ctor_get(x_18, 4); lean_inc(x_19); -x_20 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_14); +x_20 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_14); lean_dec(x_19); if (x_20 == 0) { @@ -24667,7 +24724,7 @@ if (x_34 == 0) lean_object* x_35; uint8_t x_36; x_35 = lean_ctor_get(x_30, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_14); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_14); lean_dec(x_35); if (x_36 == 0) { @@ -24692,7 +24749,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_30); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); lean_dec(x_41); if (x_42 == 0) { @@ -24718,7 +24775,7 @@ x_22 = l_Lean_Parser_Term_structInstField___elambda__1___closed__2; x_23 = l_Lean_Parser_ParserState_mkNode(x_21, x_22, x_17); x_24 = lean_ctor_get(x_23, 4); lean_inc(x_24); -x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_14); +x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_14); lean_dec(x_24); if (x_25 == 0) { @@ -25036,7 +25093,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -25057,7 +25114,7 @@ x_20 = l_Lean_Parser_Term_structInstFieldAbbrev___elambda__1___closed__2; x_21 = l_Lean_Parser_ParserState_mkNode(x_19, x_20, x_17); x_22 = lean_ctor_get(x_21, 4); lean_inc(x_22); -x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_22, x_14); +x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_22, x_14); lean_dec(x_22); if (x_23 == 0) { @@ -25321,7 +25378,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -25342,7 +25399,7 @@ x_21 = l_Lean_Parser_Term_optEllipsis___elambda__1___closed__2; x_22 = l_Lean_Parser_ParserState_mkNode(x_20, x_21, x_19); x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); -x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_16); +x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_16); lean_dec(x_23); if (x_24 == 0) { @@ -25710,7 +25767,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_structInst___elambda__1___closed__24; x_2 = lean_ctor_get(x_1, 1); lean_inc(x_2); -x_3 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__21; +x_3 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__23; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_andthenFn), 4, 2); lean_closure_set(x_4, 0, x_3); lean_closure_set(x_4, 1, x_2); @@ -26020,7 +26077,7 @@ x_22 = l_Lean_Parser_checkPrecFn(x_21, x_1, x_2); x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); x_24 = lean_box(0); -x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_24); +x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_24); lean_dec(x_23); if (x_25 == 0) { @@ -26073,7 +26130,7 @@ x_36 = l_Lean_Parser_Term_structInst___elambda__1___closed__2; x_37 = l_Lean_Parser_ParserState_mkNode(x_35, x_36, x_27); x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_24); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_24); lean_dec(x_38); if (x_39 == 0) { @@ -26093,7 +26150,7 @@ return x_40; lean_object* x_43; uint8_t x_44; x_43 = lean_ctor_get(x_42, 4); lean_inc(x_43); -x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_24); +x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_24); lean_dec(x_43); if (x_44 == 0) { @@ -26116,7 +26173,7 @@ lean_inc(x_1); x_45 = lean_apply_2(x_6, x_1, x_42); x_46 = lean_ctor_get(x_45, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_24); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_24); lean_dec(x_46); if (x_47 == 0) { @@ -26153,7 +26210,7 @@ lean_ctor_set_uint8(x_52, sizeof(void*)*7, x_14); x_53 = lean_apply_2(x_49, x_52, x_45); x_54 = lean_ctor_get(x_53, 4); lean_inc(x_54); -x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_24); +x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_24); lean_dec(x_54); if (x_55 == 0) { @@ -26169,7 +26226,7 @@ lean_inc(x_1); x_56 = l_Lean_Parser_Term_optEllipsis___elambda__1(x_1, x_53); x_57 = lean_ctor_get(x_56, 4); lean_inc(x_57); -x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_24); +x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_24); lean_dec(x_57); if (x_58 == 0) { @@ -26185,7 +26242,7 @@ lean_inc(x_1); x_59 = lean_apply_2(x_4, x_1, x_56); x_60 = lean_ctor_get(x_59, 4); lean_inc(x_60); -x_61 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_60, x_24); +x_61 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_60, x_24); lean_dec(x_60); if (x_61 == 0) { @@ -28328,7 +28385,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -28359,7 +28416,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_20, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -28368,7 +28425,7 @@ x_27 = l_Lean_Parser_Term_typeSpec___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_20, x_27, x_17); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -28394,7 +28451,7 @@ x_35 = l_Lean_Parser_Term_typeSpec___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_17); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); lean_dec(x_37); if (x_38 == 0) { @@ -28417,7 +28474,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_20); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); lean_dec(x_41); if (x_42 == 0) { @@ -28426,7 +28483,7 @@ x_43 = l_Lean_Parser_Term_typeSpec___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_17); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); lean_dec(x_45); if (x_46 == 0) { @@ -28452,7 +28509,7 @@ x_51 = l_Lean_Parser_Term_typeSpec___elambda__1___closed__2; x_52 = l_Lean_Parser_ParserState_mkNode(x_50, x_51, x_17); x_53 = lean_ctor_get(x_52, 4); lean_inc(x_53); -x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); +x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); lean_dec(x_53); if (x_54 == 0) { @@ -28732,7 +28789,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -28763,7 +28820,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_20, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -28772,7 +28829,7 @@ x_27 = l_Lean_Parser_Term_explicit___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_20, x_27, x_17); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -28798,7 +28855,7 @@ x_35 = l_Lean_Parser_Term_explicit___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_17); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); lean_dec(x_37); if (x_38 == 0) { @@ -28821,7 +28878,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_20); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); lean_dec(x_41); if (x_42 == 0) { @@ -28830,7 +28887,7 @@ x_43 = l_Lean_Parser_Term_explicit___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_17); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); lean_dec(x_45); if (x_46 == 0) { @@ -28856,7 +28913,7 @@ x_51 = l_Lean_Parser_Term_explicit___elambda__1___closed__2; x_52 = l_Lean_Parser_ParserState_mkNode(x_50, x_51, x_17); x_53 = lean_ctor_get(x_52, 4); lean_inc(x_53); -x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); +x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); lean_dec(x_53); if (x_54 == 0) { @@ -29463,7 +29520,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -29506,7 +29563,7 @@ goto block_60; lean_object* x_26; uint8_t x_27; x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); lean_dec(x_26); if (x_27 == 0) { @@ -29516,7 +29573,7 @@ x_28 = l_Lean_Parser_Term_inaccessible___elambda__1___closed__2; x_29 = l_Lean_Parser_ParserState_mkNode(x_25, x_28, x_17); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); lean_dec(x_30); if (x_31 == 0) { @@ -29540,7 +29597,7 @@ lean_inc(x_1); x_35 = l_Lean_Parser_categoryParser___elambda__1(x_33, x_34, x_1, x_25); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); lean_dec(x_36); if (x_37 == 0) { @@ -29550,7 +29607,7 @@ x_38 = l_Lean_Parser_Term_inaccessible___elambda__1___closed__2; x_39 = l_Lean_Parser_ParserState_mkNode(x_35, x_38, x_17); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); lean_dec(x_40); if (x_41 == 0) { @@ -29585,7 +29642,7 @@ x_49 = l_Lean_Parser_Term_inaccessible___elambda__1___closed__2; x_50 = l_Lean_Parser_ParserState_mkNode(x_45, x_49, x_17); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_14); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_14); lean_dec(x_51); if (x_52 == 0) { @@ -29609,7 +29666,7 @@ x_55 = l_Lean_Parser_Term_inaccessible___elambda__1___closed__2; x_56 = l_Lean_Parser_ParserState_mkNode(x_54, x_55, x_17); x_57 = lean_ctor_get(x_56, 4); lean_inc(x_57); -x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_14); +x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_14); lean_dec(x_57); if (x_58 == 0) { @@ -30107,6 +30164,24 @@ x_1 = l_Lean_Parser_Term_binderIdent___closed__2; return x_1; } } +static lean_object* _init_l_Lean_Parser_Term_binderType___elambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("null", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Term_binderType___elambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_Term_binderType___elambda__1___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Parser_Term_binderType___elambda__1(lean_object* x_1, lean_object* x_2) { _start: { @@ -30137,13 +30212,13 @@ if (x_13 == 0) lean_object* x_15; uint8_t x_16; x_15 = lean_ctor_get(x_7, 4); lean_inc(x_15); -x_16 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_14); +x_16 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_14); lean_dec(x_15); if (x_16 == 0) { lean_object* x_17; lean_object* x_18; lean_dec(x_1); -x_17 = l_Lean_nullKind; +x_17 = l_Lean_Parser_Term_binderType___elambda__1___closed__2; x_18 = l_Lean_Parser_ParserState_mkNode(x_7, x_17, x_4); return x_18; } @@ -30153,7 +30228,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean x_19 = l___regBuiltin_Lean_Parser_Term_byTactic___closed__2; x_20 = lean_unsigned_to_nat(0u); x_21 = l_Lean_Parser_categoryParser___elambda__1(x_19, x_20, x_1, x_7); -x_22 = l_Lean_nullKind; +x_22 = l_Lean_Parser_Term_binderType___elambda__1___closed__2; x_23 = l_Lean_Parser_ParserState_mkNode(x_21, x_22, x_4); return x_23; } @@ -30165,13 +30240,13 @@ lean_inc(x_1); x_24 = l_Lean_Parser_tokenAntiquotFn(x_1, x_7); x_25 = lean_ctor_get(x_24, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { lean_object* x_27; lean_object* x_28; lean_dec(x_1); -x_27 = l_Lean_nullKind; +x_27 = l_Lean_Parser_Term_binderType___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_24, x_27, x_4); return x_28; } @@ -30181,7 +30256,7 @@ lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean x_29 = l___regBuiltin_Lean_Parser_Term_byTactic___closed__2; x_30 = lean_unsigned_to_nat(0u); x_31 = l_Lean_Parser_categoryParser___elambda__1(x_29, x_30, x_1, x_24); -x_32 = l_Lean_nullKind; +x_32 = l_Lean_Parser_Term_binderType___elambda__1___closed__2; x_33 = l_Lean_Parser_ParserState_mkNode(x_31, x_32, x_4); return x_33; } @@ -30213,7 +30288,7 @@ static lean_object* _init_l_Lean_Parser_Term_binderType___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_nullKind; +x_1 = l_Lean_Parser_Term_binderType___elambda__1___closed__2; x_2 = l_Lean_Parser_Term_typeAscription___closed__3; x_3 = l_Lean_Parser_nodeInfo(x_1, x_2); return x_3; @@ -30421,7 +30496,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -30440,7 +30515,7 @@ lean_inc(x_1); x_19 = l_Lean_Parser_atomicFn(x_18, x_1, x_12); x_20 = lean_ctor_get(x_19, 4); lean_inc(x_20); -x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_20, x_14); +x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_20, x_14); lean_dec(x_20); if (x_21 == 0) { @@ -30449,7 +30524,7 @@ x_22 = l_Lean_Parser_Term_binderTactic___elambda__1___closed__2; x_23 = l_Lean_Parser_ParserState_mkNode(x_19, x_22, x_17); x_24 = lean_ctor_get(x_23, 4); lean_inc(x_24); -x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_14); +x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_14); lean_dec(x_24); if (x_25 == 0) { @@ -30473,7 +30548,7 @@ x_28 = l_Lean_Parser_Term_binderTactic___elambda__1___closed__2; x_29 = l_Lean_Parser_ParserState_mkNode(x_27, x_28, x_17); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); lean_dec(x_30); if (x_31 == 0) { @@ -30694,7 +30769,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -30725,7 +30800,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_20, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -30734,7 +30809,7 @@ x_27 = l_Lean_Parser_Term_binderDefault___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_20, x_27, x_17); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -30760,7 +30835,7 @@ x_35 = l_Lean_Parser_Term_binderDefault___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_17); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); lean_dec(x_37); if (x_38 == 0) { @@ -30783,7 +30858,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_20); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); lean_dec(x_41); if (x_42 == 0) { @@ -30792,7 +30867,7 @@ x_43 = l_Lean_Parser_Term_binderDefault___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_17); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); lean_dec(x_45); if (x_46 == 0) { @@ -30818,7 +30893,7 @@ x_51 = l_Lean_Parser_Term_binderDefault___elambda__1___closed__2; x_52 = l_Lean_Parser_ParserState_mkNode(x_50, x_51, x_17); x_53 = lean_ctor_get(x_52, 4); lean_inc(x_53); -x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); +x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); lean_dec(x_53); if (x_54 == 0) { @@ -31282,7 +31357,7 @@ static lean_object* _init_l_Lean_Parser_Term_strictImplicitLeftBracket___elambda _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_groupKind; +x_1 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__19; x_2 = l_Lean_Parser_Term_strictImplicitLeftBracket___elambda__1___closed__1; x_3 = lean_alloc_closure((void*)(l_Lean_Parser_nodeFn), 4, 2); lean_closure_set(x_3, 0, x_1); @@ -31351,7 +31426,7 @@ static lean_object* _init_l_Lean_Parser_Term_strictImplicitLeftBracket___closed_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_groupKind; +x_1 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__19; x_2 = l_Lean_Parser_Term_strictImplicitLeftBracket___closed__1; x_3 = l_Lean_Parser_nodeInfo(x_1, x_2); return x_3; @@ -31419,7 +31494,7 @@ static lean_object* _init_l_Lean_Parser_Term_strictImplicitRightBracket___elambd _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_groupKind; +x_1 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__19; x_2 = l_Lean_Parser_Term_strictImplicitRightBracket___elambda__1___closed__1; x_3 = lean_alloc_closure((void*)(l_Lean_Parser_nodeFn), 4, 2); lean_closure_set(x_3, 0, x_1); @@ -31488,7 +31563,7 @@ static lean_object* _init_l_Lean_Parser_Term_strictImplicitRightBracket___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_groupKind; +x_1 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__19; x_2 = l_Lean_Parser_Term_strictImplicitRightBracket___closed__1; x_3 = l_Lean_Parser_nodeInfo(x_1, x_2); return x_3; @@ -31801,7 +31876,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -31847,7 +31922,7 @@ x_28 = l_Lean_Parser_Term_instBinder___elambda__1___closed__2; x_29 = l_Lean_Parser_ParserState_mkNode(x_27, x_28, x_19); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_16); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_16); lean_dec(x_30); if (x_31 == 0) { @@ -31867,7 +31942,7 @@ return x_32; lean_object* x_35; uint8_t x_36; x_35 = lean_ctor_get(x_34, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_16); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_16); lean_dec(x_35); if (x_36 == 0) { @@ -31883,7 +31958,7 @@ lean_inc(x_1); x_37 = lean_apply_2(x_4, x_1, x_34); x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_16); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_16); lean_dec(x_38); if (x_39 == 0) { @@ -31900,7 +31975,7 @@ lean_inc(x_1); x_42 = l_Lean_Parser_categoryParser___elambda__1(x_40, x_41, x_1, x_37); x_43 = lean_ctor_get(x_42, 4); lean_inc(x_43); -x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_16); +x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_16); lean_dec(x_43); if (x_44 == 0) { @@ -32065,13 +32140,12 @@ return x_3; static lean_object* _init_l_Lean_Parser_Term_bracketedBinder___elambda__1___closed__3() { _start: { -lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; +lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_bracketedBinder___elambda__1___closed__1; x_2 = l_Lean_Parser_Term_bracketedBinder___elambda__1___closed__2; -x_3 = 0; -x_4 = 1; -x_5 = l_Lean_Parser_mkAntiquot(x_1, x_2, x_3, x_4); -return x_5; +x_3 = 1; +x_4 = l_Lean_Parser_mkAntiquot(x_1, x_2, x_3, x_3); +return x_4; } } LEAN_EXPORT lean_object* l_Lean_Parser_Term_bracketedBinder___elambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { @@ -32419,7 +32493,7 @@ x_26 = l_Lean_Parser_checkPrecFn(x_25, x_1, x_2); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); x_28 = lean_box(0); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_28); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_28); lean_dec(x_27); if (x_29 == 0) { @@ -32439,7 +32513,7 @@ lean_inc(x_1); x_32 = lean_apply_2(x_9, x_1, x_26); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_28); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_28); lean_dec(x_33); if (x_34 == 0) { @@ -32465,7 +32539,7 @@ if (x_46 == 0) lean_object* x_47; uint8_t x_48; x_47 = lean_ctor_get(x_42, 4); lean_inc(x_47); -x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_28); +x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_28); lean_dec(x_47); if (x_48 == 0) { @@ -32490,7 +32564,7 @@ lean_inc(x_1); x_52 = l_Lean_Parser_tokenAntiquotFn(x_1, x_42); x_53 = lean_ctor_get(x_52, 4); lean_inc(x_53); -x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_28); +x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_28); lean_dec(x_53); if (x_54 == 0) { @@ -32515,7 +32589,7 @@ lean_object* x_36; lean_object* x_37; uint8_t x_38; x_36 = l_Lean_Parser_ParserState_mkNode(x_35, x_11, x_31); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_28); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_28); lean_dec(x_37); if (x_38 == 0) { @@ -32795,7 +32869,7 @@ return x_8; else { lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_9 = l_Lean_nullKind; +x_9 = l_Lean_Parser_Term_binderType___elambda__1___closed__2; x_10 = l_Lean_Parser_Term_typeAscription_formatter___closed__3; x_11 = l_Lean_PrettyPrinter_Formatter_node_formatter(x_9, x_10, x_2, x_3, x_4, x_5, x_6); return x_11; @@ -33433,19 +33507,18 @@ return x_7; static lean_object* _init_l_Lean_Parser_Term_bracketedBinder_formatter___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Parser_Term_bracketedBinder___elambda__1___closed__1; x_2 = l_Lean_Parser_Term_bracketedBinder___elambda__1___closed__2; -x_3 = 0; -x_4 = 1; +x_3 = 1; +x_4 = lean_box(x_3); x_5 = lean_box(x_3); -x_6 = lean_box(x_4); -x_7 = lean_alloc_closure((void*)(l_Lean_Parser_mkAntiquot_formatter___boxed), 9, 4); -lean_closure_set(x_7, 0, x_1); -lean_closure_set(x_7, 1, x_2); -lean_closure_set(x_7, 2, x_5); -lean_closure_set(x_7, 3, x_6); -return x_7; +x_6 = lean_alloc_closure((void*)(l_Lean_Parser_mkAntiquot_formatter___boxed), 9, 4); +lean_closure_set(x_6, 0, x_1); +lean_closure_set(x_6, 1, x_2); +lean_closure_set(x_6, 2, x_4); +lean_closure_set(x_6, 3, x_5); +return x_6; } } static lean_object* _init_l_Lean_Parser_Term_bracketedBinder_formatter___closed__2() { @@ -33636,7 +33709,7 @@ return x_8; else { lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_9 = l_Lean_nullKind; +x_9 = l_Lean_Parser_Term_binderType___elambda__1___closed__2; x_10 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__3; x_11 = l_Lean_PrettyPrinter_Parenthesizer_node_parenthesizer(x_9, x_10, x_2, x_3, x_4, x_5, x_6); return x_11; @@ -34254,19 +34327,18 @@ return x_8; static lean_object* _init_l_Lean_Parser_Term_bracketedBinder_parenthesizer___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Parser_Term_bracketedBinder___elambda__1___closed__1; x_2 = l_Lean_Parser_Term_bracketedBinder___elambda__1___closed__2; -x_3 = 0; -x_4 = 1; +x_3 = 1; +x_4 = lean_box(x_3); x_5 = lean_box(x_3); -x_6 = lean_box(x_4); -x_7 = lean_alloc_closure((void*)(l_Lean_Parser_mkAntiquot_parenthesizer___boxed), 9, 4); -lean_closure_set(x_7, 0, x_1); -lean_closure_set(x_7, 1, x_2); -lean_closure_set(x_7, 2, x_5); -lean_closure_set(x_7, 3, x_6); -return x_7; +x_6 = lean_alloc_closure((void*)(l_Lean_Parser_mkAntiquot_parenthesizer___boxed), 9, 4); +lean_closure_set(x_6, 0, x_1); +lean_closure_set(x_6, 1, x_2); +lean_closure_set(x_6, 2, x_4); +lean_closure_set(x_6, 3, x_5); +return x_6; } } static lean_object* _init_l_Lean_Parser_Term_bracketedBinder_parenthesizer___closed__2() { @@ -34550,7 +34622,7 @@ x_16 = l_Lean_Parser_checkPrecFn(x_15, x_1, x_2); x_17 = lean_ctor_get(x_16, 4); lean_inc(x_17); x_18 = lean_box(0); -x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); +x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); lean_dec(x_17); if (x_19 == 0) { @@ -34570,7 +34642,7 @@ lean_inc(x_1); x_22 = lean_apply_2(x_4, x_1, x_16); x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); -x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_18); +x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_18); lean_dec(x_23); if (x_24 == 0) { @@ -34580,7 +34652,7 @@ x_25 = l_Lean_Parser_Term_simpleBinder___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_22, x_25, x_21); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_18); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_18); lean_dec(x_27); if (x_28 == 0) { @@ -34604,7 +34676,7 @@ x_31 = l_Lean_Parser_Term_simpleBinder___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_21); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_18); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_18); lean_dec(x_33); if (x_34 == 0) { @@ -35021,7 +35093,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -35068,7 +35140,7 @@ x_29 = l_Lean_Parser_Term_forall___elambda__1___closed__2; x_30 = l_Lean_Parser_ParserState_mkNode(x_28, x_29, x_19); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); lean_dec(x_31); if (x_32 == 0) { @@ -35088,7 +35160,7 @@ return x_33; lean_object* x_36; uint8_t x_37; x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_16); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_16); lean_dec(x_36); if (x_37 == 0) { @@ -35104,7 +35176,7 @@ lean_inc(x_1); x_38 = lean_apply_2(x_4, x_1, x_35); x_39 = lean_ctor_get(x_38, 4); lean_inc(x_39); -x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_16); +x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_16); lean_dec(x_39); if (x_40 == 0) { @@ -35130,7 +35202,7 @@ if (x_46 == 0) lean_object* x_47; uint8_t x_48; x_47 = lean_ctor_get(x_43, 4); lean_inc(x_47); -x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_16); +x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_16); lean_dec(x_47); if (x_48 == 0) { @@ -35155,7 +35227,7 @@ lean_inc(x_1); x_52 = l_Lean_Parser_tokenAntiquotFn(x_1, x_43); x_53 = lean_ctor_get(x_52, 4); lean_inc(x_53); -x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_16); +x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_16); lean_dec(x_53); if (x_54 == 0) { @@ -36356,6 +36428,22 @@ x_1 = l_Lean_Parser_Term_matchAltExpr___closed__1; return x_1; } } +LEAN_EXPORT lean_object* l_Lean_Parser_Term_instCoeTSyntaxConsSyntaxNodeKindMkStrAnonymousNil(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Parser_Term_instCoeTSyntaxConsSyntaxNodeKindMkStrAnonymousNil___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Parser_Term_instCoeTSyntaxConsSyntaxNodeKindMkStrAnonymousNil(x_1); +lean_dec(x_1); +return x_2; +} +} static lean_object* _init_l_Lean_Parser_Term_matchAlts___elambda__1___closed__1() { _start: { @@ -36498,7 +36586,7 @@ lean_closure_set(x_8, 0, x_7); lean_closure_set(x_8, 1, x_6); x_9 = l_Lean_Parser_Command_commentBody___closed__2; x_10 = l_Lean_Parser_andthenInfo(x_9, x_5); -x_11 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__21; +x_11 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__23; x_12 = lean_alloc_closure((void*)(l_Lean_Parser_andthenFn), 4, 2); lean_closure_set(x_12, 0, x_11); lean_closure_set(x_12, 1, x_8); @@ -36648,7 +36736,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -36667,7 +36755,7 @@ lean_inc(x_1); x_20 = lean_apply_2(x_4, x_1, x_14); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); -x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_16); +x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_16); lean_dec(x_21); if (x_22 == 0) { @@ -36676,7 +36764,7 @@ x_23 = l_Lean_Parser_Term_matchDiscr___elambda__1___closed__2; x_24 = l_Lean_Parser_ParserState_mkNode(x_20, x_23, x_19); x_25 = lean_ctor_get(x_24, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_16); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_16); lean_dec(x_25); if (x_26 == 0) { @@ -36702,7 +36790,7 @@ x_31 = l_Lean_Parser_Term_matchDiscr___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_19); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_16); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_16); lean_dec(x_33); if (x_34 == 0) { @@ -36958,7 +37046,7 @@ x_20 = l_Lean_Parser_checkPrecFn(x_19, x_1, x_2); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); x_22 = lean_box(0); -x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_22); +x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_22); lean_dec(x_21); if (x_23 == 0) { @@ -36989,7 +37077,7 @@ lean_object* x_32; lean_object* x_33; uint8_t x_34; x_32 = l_Lean_Parser_ParserState_mkNode(x_27, x_5, x_25); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_22); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_22); lean_dec(x_33); if (x_34 == 0) { @@ -37012,7 +37100,7 @@ x_36 = l_Lean_Parser_tokenAntiquotFn(x_1, x_27); x_37 = l_Lean_Parser_ParserState_mkNode(x_36, x_5, x_25); x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_22); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_22); lean_dec(x_38); if (x_39 == 0) { @@ -37259,7 +37347,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -37292,7 +37380,7 @@ x_25 = l_Lean_Parser_Term_falseVal___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_20, x_25, x_17); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); lean_dec(x_27); if (x_28 == 0) { @@ -37316,7 +37404,7 @@ x_31 = l_Lean_Parser_Term_falseVal___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_17); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -37626,7 +37714,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -37646,7 +37734,7 @@ lean_inc(x_1); x_19 = l_Lean_Parser_atomicFn(x_18, x_1, x_12); x_20 = lean_ctor_get(x_19, 4); lean_inc(x_20); -x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_20, x_14); +x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_20, x_14); lean_dec(x_20); if (x_21 == 0) { @@ -37656,7 +37744,7 @@ x_22 = l_Lean_Parser_Term_generalizingParam___elambda__1___closed__2; x_23 = l_Lean_Parser_ParserState_mkNode(x_19, x_22, x_17); x_24 = lean_ctor_get(x_23, 4); lean_inc(x_24); -x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_14); +x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_14); lean_dec(x_24); if (x_25 == 0) { @@ -37702,7 +37790,7 @@ goto block_70; lean_object* x_35; uint8_t x_36; x_35 = lean_ctor_get(x_34, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_14); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_14); lean_dec(x_35); if (x_36 == 0) { @@ -37712,7 +37800,7 @@ x_37 = l_Lean_Parser_Term_generalizingParam___elambda__1___closed__2; x_38 = l_Lean_Parser_ParserState_mkNode(x_34, x_37, x_17); x_39 = lean_ctor_get(x_38, 4); lean_inc(x_39); -x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); +x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); lean_dec(x_39); if (x_40 == 0) { @@ -37737,7 +37825,7 @@ lean_inc(x_1); x_45 = l_Lean_Parser_orelseFnCore(x_42, x_43, x_44, x_1, x_34); x_46 = lean_ctor_get(x_45, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_14); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_14); lean_dec(x_46); if (x_47 == 0) { @@ -37747,7 +37835,7 @@ x_48 = l_Lean_Parser_Term_generalizingParam___elambda__1___closed__2; x_49 = l_Lean_Parser_ParserState_mkNode(x_45, x_48, x_17); x_50 = lean_ctor_get(x_49, 4); lean_inc(x_50); -x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_14); +x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_14); lean_dec(x_50); if (x_51 == 0) { @@ -37782,7 +37870,7 @@ x_59 = l_Lean_Parser_Term_generalizingParam___elambda__1___closed__2; x_60 = l_Lean_Parser_ParserState_mkNode(x_55, x_59, x_17); x_61 = lean_ctor_get(x_60, 4); lean_inc(x_61); -x_62 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_61, x_14); +x_62 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_61, x_14); lean_dec(x_61); if (x_62 == 0) { @@ -37806,7 +37894,7 @@ x_65 = l_Lean_Parser_Term_generalizingParam___elambda__1___closed__2; x_66 = l_Lean_Parser_ParserState_mkNode(x_64, x_65, x_17); x_67 = lean_ctor_get(x_66, 4); lean_inc(x_67); -x_68 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_67, x_14); +x_68 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_67, x_14); lean_dec(x_67); if (x_68 == 0) { @@ -38152,7 +38240,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -38172,7 +38260,7 @@ lean_inc(x_1); x_19 = l_Lean_Parser_atomicFn(x_18, x_1, x_12); x_20 = lean_ctor_get(x_19, 4); lean_inc(x_20); -x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_20, x_14); +x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_20, x_14); lean_dec(x_20); if (x_21 == 0) { @@ -38182,7 +38270,7 @@ x_22 = l_Lean_Parser_Term_motive___elambda__1___closed__2; x_23 = l_Lean_Parser_ParserState_mkNode(x_19, x_22, x_17); x_24 = lean_ctor_get(x_23, 4); lean_inc(x_24); -x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_14); +x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_14); lean_dec(x_24); if (x_25 == 0) { @@ -38206,7 +38294,7 @@ lean_inc(x_1); x_29 = l_Lean_Parser_categoryParser___elambda__1(x_27, x_28, x_1, x_19); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); lean_dec(x_30); if (x_31 == 0) { @@ -38216,7 +38304,7 @@ x_32 = l_Lean_Parser_Term_motive___elambda__1___closed__2; x_33 = l_Lean_Parser_ParserState_mkNode(x_29, x_32, x_17); x_34 = lean_ctor_get(x_33, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_14); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_14); lean_dec(x_34); if (x_35 == 0) { @@ -38252,7 +38340,7 @@ x_44 = l_Lean_Parser_Term_motive___elambda__1___closed__2; x_45 = l_Lean_Parser_ParserState_mkNode(x_39, x_44, x_17); x_46 = lean_ctor_get(x_45, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_14); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_14); lean_dec(x_46); if (x_47 == 0) { @@ -38276,7 +38364,7 @@ x_50 = l_Lean_Parser_Term_motive___elambda__1___closed__2; x_51 = l_Lean_Parser_ParserState_mkNode(x_49, x_50, x_17); x_52 = lean_ctor_get(x_51, 4); lean_inc(x_52); -x_53 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_52, x_14); +x_53 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_52, x_14); lean_dec(x_52); if (x_53 == 0) { @@ -38707,7 +38795,7 @@ x_20 = l_Lean_Parser_checkPrecFn(x_19, x_1, x_2); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); x_22 = lean_box(0); -x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_22); +x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_22); lean_dec(x_21); if (x_23 == 0) { @@ -38756,7 +38844,7 @@ x_34 = l_Lean_Parser_Term_match___elambda__1___closed__2; x_35 = l_Lean_Parser_ParserState_mkNode(x_33, x_34, x_25); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_22); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_22); lean_dec(x_36); if (x_37 == 0) { @@ -38776,7 +38864,7 @@ return x_38; lean_object* x_41; uint8_t x_42; x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_22); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_22); lean_dec(x_41); if (x_42 == 0) { @@ -38795,7 +38883,7 @@ lean_inc(x_1); x_43 = lean_apply_2(x_10, x_1, x_40); x_44 = lean_ctor_get(x_43, 4); lean_inc(x_44); -x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_22); +x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_22); lean_dec(x_44); if (x_45 == 0) { @@ -38813,7 +38901,7 @@ lean_inc(x_1); x_46 = lean_apply_2(x_8, x_1, x_43); x_47 = lean_ctor_get(x_46, 4); lean_inc(x_47); -x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_22); +x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_22); lean_dec(x_47); if (x_48 == 0) { @@ -38830,7 +38918,7 @@ lean_inc(x_1); x_49 = lean_apply_2(x_6, x_1, x_46); x_50 = lean_ctor_get(x_49, 4); lean_inc(x_50); -x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_22); +x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_22); lean_dec(x_50); if (x_51 == 0) { @@ -38857,7 +38945,7 @@ if (x_57 == 0) lean_object* x_58; uint8_t x_59; x_58 = lean_ctor_get(x_54, 4); lean_inc(x_58); -x_59 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_58, x_22); +x_59 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_58, x_22); lean_dec(x_58); if (x_59 == 0) { @@ -38881,7 +38969,7 @@ lean_inc(x_1); x_61 = l_Lean_Parser_tokenAntiquotFn(x_1, x_54); x_62 = lean_ctor_get(x_61, 4); lean_inc(x_62); -x_63 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_62, x_22); +x_63 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_62, x_22); lean_dec(x_62); if (x_63 == 0) { @@ -39074,7 +39162,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_match_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(157u); +x_1 = lean_unsigned_to_nat(160u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -39086,7 +39174,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_match_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(157u); +x_1 = lean_unsigned_to_nat(160u); x_2 = lean_unsigned_to_nat(177u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -39114,7 +39202,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_match_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(157u); +x_1 = lean_unsigned_to_nat(160u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -39126,7 +39214,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_match_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(157u); +x_1 = lean_unsigned_to_nat(160u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -41180,7 +41268,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -41211,7 +41299,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_20, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -41220,7 +41308,7 @@ x_27 = l_Lean_Parser_Term_nomatch___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_20, x_27, x_17); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -41246,7 +41334,7 @@ x_35 = l_Lean_Parser_Term_nomatch___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_17); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); lean_dec(x_37); if (x_38 == 0) { @@ -41269,7 +41357,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_20); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); lean_dec(x_41); if (x_42 == 0) { @@ -41278,7 +41366,7 @@ x_43 = l_Lean_Parser_Term_nomatch___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_17); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); lean_dec(x_45); if (x_46 == 0) { @@ -41304,7 +41392,7 @@ x_51 = l_Lean_Parser_Term_nomatch___elambda__1___closed__2; x_52 = l_Lean_Parser_ParserState_mkNode(x_50, x_51, x_17); x_53 = lean_ctor_get(x_52, 4); lean_inc(x_53); -x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); +x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); lean_dec(x_53); if (x_54 == 0) { @@ -41441,7 +41529,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_nomatch_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(158u); +x_1 = lean_unsigned_to_nat(161u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -41453,7 +41541,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_nomatch_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(158u); +x_1 = lean_unsigned_to_nat(161u); x_2 = lean_unsigned_to_nat(86u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -41481,7 +41569,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_nomatch_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(158u); +x_1 = lean_unsigned_to_nat(161u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -41493,7 +41581,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_nomatch_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(158u); +x_1 = lean_unsigned_to_nat(161u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -41778,38 +41866,91 @@ return x_2; static lean_object* _init_l_Lean_Parser_Term_funImplicitBinder___elambda__1___closed__5() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_funImplicitBinder___elambda__1___closed__4; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_atomicFn), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Term_funImplicitBinder___elambda__1___closed__6() { +_start: +{ uint8_t x_1; lean_object* x_2; x_1 = 0; x_2 = l_Lean_Parser_Term_implicitBinder(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Term_funImplicitBinder___elambda__1(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_Parser_Term_funImplicitBinder___elambda__1___closed__7() { _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; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Term_funImplicitBinder___elambda__1___closed__6; +x_2 = lean_ctor_get(x_1, 1); +lean_inc(x_2); x_3 = l_Lean_Parser_Term_funImplicitBinder___elambda__1___closed__5; +x_4 = lean_alloc_closure((void*)(l_Lean_Parser_andthenFn), 4, 2); +lean_closure_set(x_4, 0, x_3); +lean_closure_set(x_4, 1, x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Parser_Term_funImplicitBinder___elambda__1(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; uint32_t x_10; uint32_t x_11; uint8_t x_12; +x_3 = l_Lean_Parser_Term_funImplicitBinder___elambda__1___closed__6; x_4 = lean_ctor_get(x_3, 1); lean_inc(x_4); -x_5 = l_Lean_Parser_Term_funImplicitBinder___elambda__1___closed__4; -lean_inc(x_1); -x_6 = l_Lean_Parser_atomicFn(x_5, x_1, x_2); -x_7 = lean_ctor_get(x_6, 4); +x_5 = l_Lean_Parser_Term_implicitBinder___elambda__1___closed__3; +x_6 = lean_ctor_get(x_5, 1); +lean_inc(x_6); +x_7 = lean_ctor_get(x_1, 0); lean_inc(x_7); -x_8 = lean_box(0); -x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); lean_dec(x_7); -if (x_9 == 0) +x_9 = lean_ctor_get(x_2, 2); +lean_inc(x_9); +x_10 = lean_string_utf8_get(x_8, x_9); +lean_dec(x_9); +lean_dec(x_8); +x_11 = 36; +x_12 = lean_uint32_dec_eq(x_10, x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +lean_dec(x_6); +x_13 = l_Lean_Parser_Term_funImplicitBinder___elambda__1___closed__4; +lean_inc(x_1); +x_14 = l_Lean_Parser_atomicFn(x_13, x_1, x_2); +x_15 = lean_ctor_get(x_14, 4); +lean_inc(x_15); +x_16 = lean_box(0); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +lean_dec(x_15); +if (x_17 == 0) { lean_dec(x_4); lean_dec(x_1); -return x_6; +return x_14; } else { -lean_object* x_10; -x_10 = lean_apply_2(x_4, x_1, x_6); -return x_10; +lean_object* x_18; +x_18 = lean_apply_2(x_4, x_1, x_14); +return x_18; +} +} +else +{ +lean_object* x_19; uint8_t x_20; lean_object* x_21; +lean_dec(x_4); +x_19 = l_Lean_Parser_Term_funImplicitBinder___elambda__1___closed__7; +x_20 = 0; +x_21 = l_Lean_Parser_orelseFnCore(x_6, x_19, x_20, x_1, x_2); +return x_21; } } } @@ -41849,7 +41990,7 @@ static lean_object* _init_l_Lean_Parser_Term_funImplicitBinder___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Term_funImplicitBinder___elambda__1___closed__5; +x_1 = l_Lean_Parser_Term_funImplicitBinder___elambda__1___closed__6; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l_Lean_Parser_Term_funImplicitBinder___closed__3; @@ -41860,17 +42001,29 @@ return x_4; static lean_object* _init_l_Lean_Parser_Term_funImplicitBinder___closed__5() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Term_implicitBinder___elambda__1___closed__3; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_3 = l_Lean_Parser_Term_funImplicitBinder___closed__4; +x_4 = l_Lean_Parser_orelseInfo(x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Term_funImplicitBinder___closed__6() { +_start: +{ lean_object* x_1; x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_funImplicitBinder___elambda__1), 2, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_Term_funImplicitBinder___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_funImplicitBinder___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_funImplicitBinder___closed__4; -x_2 = l_Lean_Parser_Term_funImplicitBinder___closed__5; +x_1 = l_Lean_Parser_Term_funImplicitBinder___closed__5; +x_2 = l_Lean_Parser_Term_funImplicitBinder___closed__6; x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -41881,7 +42034,7 @@ static lean_object* _init_l_Lean_Parser_Term_funImplicitBinder() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Term_funImplicitBinder___closed__6; +x_1 = l_Lean_Parser_Term_funImplicitBinder___closed__7; return x_1; } } @@ -41955,7 +42108,7 @@ x_6 = l_Lean_Parser_atomicFn(x_5, x_1, x_2); x_7 = lean_ctor_get(x_6, 4); lean_inc(x_7); x_8 = lean_box(0); -x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); +x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); lean_dec(x_7); if (x_9 == 0) { @@ -42071,28 +42224,78 @@ lean_closure_set(x_2, 0, x_1); return x_2; } } +static lean_object* _init_l_Lean_Parser_Term_funSimpleBinder___elambda__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_funSimpleBinder___elambda__1___closed__2; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_atomicFn), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Term_funSimpleBinder___elambda__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_funSimpleBinder___elambda__1___closed__3; +x_2 = l_Lean_Parser_Term_simpleBinder___closed__6; +x_3 = lean_alloc_closure((void*)(l_Lean_Parser_andthenFn), 4, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Parser_Term_funSimpleBinder___elambda__1(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 = l_Lean_Parser_Term_funSimpleBinder___elambda__1___closed__2; -lean_inc(x_1); -x_4 = l_Lean_Parser_atomicFn(x_3, x_1, x_2); -x_5 = lean_ctor_get(x_4, 4); +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint32_t x_8; uint32_t x_9; uint8_t x_10; +x_3 = l_Lean_Parser_Term_simpleBinder___elambda__1___closed__3; +x_4 = lean_ctor_get(x_3, 1); +lean_inc(x_4); +x_5 = lean_ctor_get(x_1, 0); lean_inc(x_5); -x_6 = lean_box(0); -x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); lean_dec(x_5); -if (x_7 == 0) +x_7 = lean_ctor_get(x_2, 2); +lean_inc(x_7); +x_8 = lean_string_utf8_get(x_6, x_7); +lean_dec(x_7); +lean_dec(x_6); +x_9 = 36; +x_10 = lean_uint32_dec_eq(x_8, x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +lean_dec(x_4); +x_11 = l_Lean_Parser_Term_funSimpleBinder___elambda__1___closed__2; +lean_inc(x_1); +x_12 = l_Lean_Parser_atomicFn(x_11, x_1, x_2); +x_13 = lean_ctor_get(x_12, 4); +lean_inc(x_13); +x_14 = lean_box(0); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +lean_dec(x_13); +if (x_15 == 0) { lean_dec(x_1); -return x_4; +return x_12; } else { -lean_object* x_8; -x_8 = l_Lean_Parser_Term_simpleBinder___elambda__1(x_1, x_4); -return x_8; +lean_object* x_16; +x_16 = l_Lean_Parser_Term_simpleBinder___elambda__1(x_1, x_12); +return x_16; +} +} +else +{ +lean_object* x_17; uint8_t x_18; lean_object* x_19; +x_17 = l_Lean_Parser_Term_funSimpleBinder___elambda__1___closed__4; +x_18 = 0; +x_19 = l_Lean_Parser_orelseFnCore(x_4, x_17, x_18, x_1, x_2); +return x_19; } } } @@ -42123,17 +42326,29 @@ return x_4; static lean_object* _init_l_Lean_Parser_Term_funSimpleBinder___closed__3() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Term_simpleBinder___elambda__1___closed__3; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_3 = l_Lean_Parser_Term_funSimpleBinder___closed__2; +x_4 = l_Lean_Parser_orelseInfo(x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Term_funSimpleBinder___closed__4() { +_start: +{ lean_object* x_1; x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_funSimpleBinder___elambda__1), 2, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_Term_funSimpleBinder___closed__4() { +static lean_object* _init_l_Lean_Parser_Term_funSimpleBinder___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_funSimpleBinder___closed__2; -x_2 = l_Lean_Parser_Term_funSimpleBinder___closed__3; +x_1 = l_Lean_Parser_Term_funSimpleBinder___closed__3; +x_2 = l_Lean_Parser_Term_funSimpleBinder___closed__4; x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -42144,15 +42359,44 @@ static lean_object* _init_l_Lean_Parser_Term_funSimpleBinder() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Term_funSimpleBinder___closed__4; +x_1 = l_Lean_Parser_Term_funSimpleBinder___closed__5; return x_1; } } static lean_object* _init_l_Lean_Parser_Term_funBinder___elambda__1___closed__1() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("funBinder", 9); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Term_funBinder___elambda__1___closed__2() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_funSimpleBinder___closed__3; +x_1 = l_Lean_Parser_Term_byTactic___elambda__1___closed__2; +x_2 = l_Lean_Parser_Term_funBinder___elambda__1___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Term_funBinder___elambda__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Term_funBinder___elambda__1___closed__1; +x_2 = l_Lean_Parser_Term_funBinder___elambda__1___closed__2; +x_3 = 1; +x_4 = l_Lean_Parser_mkAntiquot(x_1, x_2, x_3, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Term_funBinder___elambda__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_funSimpleBinder___closed__4; x_2 = l_Lean_Parser_Term_explicit___elambda__1___closed__7; x_3 = lean_alloc_closure((void*)(l_Lean_Parser_orelseFn), 4, 2); lean_closure_set(x_3, 0, x_1); @@ -42160,24 +42404,36 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_funBinder___elambda__1___closed__2() { +static lean_object* _init_l_Lean_Parser_Term_funBinder___elambda__1___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_instBinder___closed__7; -x_2 = l_Lean_Parser_Term_funBinder___elambda__1___closed__1; +x_2 = l_Lean_Parser_Term_funBinder___elambda__1___closed__4; x_3 = lean_alloc_closure((void*)(l_Lean_Parser_orelseFn), 4, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_funBinder___elambda__1___closed__3() { +static lean_object* _init_l_Lean_Parser_Term_funBinder___elambda__1___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_funImplicitBinder___closed__5; -x_2 = l_Lean_Parser_Term_funBinder___elambda__1___closed__2; +x_1 = l_Lean_Parser_Term_funImplicitBinder___closed__6; +x_2 = l_Lean_Parser_Term_funBinder___elambda__1___closed__5; +x_3 = lean_alloc_closure((void*)(l_Lean_Parser_orelseFn), 4, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Term_funBinder___elambda__1___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_funStrictImplicitBinder___closed__5; +x_2 = l_Lean_Parser_Term_funBinder___elambda__1___closed__6; x_3 = lean_alloc_closure((void*)(l_Lean_Parser_orelseFn), 4, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); @@ -42187,12 +42443,40 @@ return x_3; LEAN_EXPORT lean_object* l_Lean_Parser_Term_funBinder___elambda__1(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; -x_3 = l_Lean_Parser_Term_funStrictImplicitBinder___closed__5; -x_4 = l_Lean_Parser_Term_funBinder___elambda__1___closed__3; -x_5 = 1; -x_6 = l_Lean_Parser_orelseFnCore(x_3, x_4, x_5, x_1, x_2); -return x_6; +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint32_t x_8; uint32_t x_9; uint8_t x_10; +x_3 = l_Lean_Parser_Term_funBinder___elambda__1___closed__3; +x_4 = lean_ctor_get(x_3, 1); +lean_inc(x_4); +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +lean_dec(x_5); +x_7 = lean_ctor_get(x_2, 2); +lean_inc(x_7); +x_8 = lean_string_utf8_get(x_6, x_7); +lean_dec(x_7); +lean_dec(x_6); +x_9 = 36; +x_10 = lean_uint32_dec_eq(x_8, x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; +lean_dec(x_4); +x_11 = l_Lean_Parser_Term_funStrictImplicitBinder___closed__5; +x_12 = l_Lean_Parser_Term_funBinder___elambda__1___closed__6; +x_13 = 1; +x_14 = l_Lean_Parser_orelseFnCore(x_11, x_12, x_13, x_1, x_2); +return x_14; +} +else +{ +lean_object* x_15; uint8_t x_16; lean_object* x_17; +x_15 = l_Lean_Parser_Term_funBinder___elambda__1___closed__7; +x_16 = 0; +x_17 = l_Lean_Parser_orelseFnCore(x_4, x_15, x_16, x_1, x_2); +return x_17; +} } } static lean_object* _init_l_Lean_Parser_Term_funBinder___closed__1() { @@ -42248,17 +42532,29 @@ return x_4; static lean_object* _init_l_Lean_Parser_Term_funBinder___closed__5() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Term_funBinder___elambda__1___closed__3; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_3 = l_Lean_Parser_Term_funBinder___closed__4; +x_4 = l_Lean_Parser_orelseInfo(x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Term_funBinder___closed__6() { +_start: +{ lean_object* x_1; x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_funBinder___elambda__1), 2, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_Term_funBinder___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_funBinder___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_funBinder___closed__4; -x_2 = l_Lean_Parser_Term_funBinder___closed__5; +x_1 = l_Lean_Parser_Term_funBinder___closed__5; +x_2 = l_Lean_Parser_Term_funBinder___closed__6; x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -42269,7 +42565,7 @@ static lean_object* _init_l_Lean_Parser_Term_funBinder() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Term_funBinder___closed__6; +x_1 = l_Lean_Parser_Term_funBinder___closed__7; return x_1; } } @@ -42319,7 +42615,7 @@ static lean_object* _init_l_Lean_Parser_Term_basicFun___elambda__1___closed__5() { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_docComment___elambda__1___closed__12; -x_2 = l_Lean_Parser_Term_funBinder___closed__5; +x_2 = l_Lean_Parser_Term_funBinder___closed__6; x_3 = lean_alloc_closure((void*)(l_Lean_Parser_andthenFn), 4, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); @@ -42498,7 +42794,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -42518,7 +42814,7 @@ lean_inc(x_1); x_20 = lean_apply_2(x_4, x_1, x_14); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); -x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_16); +x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_16); lean_dec(x_21); if (x_22 == 0) { @@ -42559,7 +42855,7 @@ goto block_39; lean_object* x_24; uint8_t x_25; x_24 = lean_ctor_get(x_23, 4); lean_inc(x_24); -x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_16); +x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_16); lean_dec(x_24); if (x_25 == 0) { @@ -42568,7 +42864,7 @@ x_26 = l_Lean_Parser_Term_basicFun___elambda__1___closed__2; x_27 = l_Lean_Parser_ParserState_mkNode(x_23, x_26, x_19); x_28 = lean_ctor_get(x_27, 4); lean_inc(x_28); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_16); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_16); lean_dec(x_28); if (x_29 == 0) { @@ -42594,7 +42890,7 @@ x_34 = l_Lean_Parser_Term_basicFun___elambda__1___closed__2; x_35 = l_Lean_Parser_ParserState_mkNode(x_33, x_34, x_19); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_16); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_16); lean_dec(x_36); if (x_37 == 0) { @@ -42981,7 +43277,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -43014,7 +43310,7 @@ if (x_34 == 0) lean_object* x_35; uint8_t x_36; x_35 = lean_ctor_get(x_30, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_16); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_16); lean_dec(x_35); if (x_36 == 0) { @@ -43040,7 +43336,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_30); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); lean_dec(x_41); if (x_42 == 0) { @@ -43066,7 +43362,7 @@ x_21 = l_Lean_Parser_Term_fun___elambda__1___closed__2; x_22 = l_Lean_Parser_ParserState_mkNode(x_20, x_21, x_19); x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); -x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_16); +x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_16); lean_dec(x_23); if (x_24 == 0) { @@ -43226,7 +43522,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_fun_declRange___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(166u); +x_1 = lean_unsigned_to_nat(169u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -43238,7 +43534,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_fun_declRange___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(166u); +x_1 = lean_unsigned_to_nat(169u); x_2 = lean_unsigned_to_nat(128u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -43266,7 +43562,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_fun_declRange___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(166u); +x_1 = lean_unsigned_to_nat(169u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -43278,7 +43574,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_fun_declRange___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(166u); +x_1 = lean_unsigned_to_nat(169u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -43468,13 +43764,25 @@ lean_closure_set(x_3, 0, x_2); return x_3; } } +static lean_object* _init_l_Lean_Parser_Term_funImplicitBinder_formatter___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_funImplicitBinder_formatter___closed__5; +x_2 = l_Lean_Parser_Term_funImplicitBinder_formatter___closed__6; +x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Parser_Term_funImplicitBinder_formatter(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; -x_6 = l_Lean_Parser_Term_funImplicitBinder_formatter___closed__5; -x_7 = l_Lean_Parser_Term_funImplicitBinder_formatter___closed__6; -x_8 = l_Lean_PrettyPrinter_Formatter_andthen_formatter(x_6, x_7, x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Parser_Term_implicitBinder_formatter___closed__1; +x_7 = l_Lean_Parser_Term_funImplicitBinder_formatter___closed__7; +x_8 = l_Lean_PrettyPrinter_Formatter_orelse_formatter(x_6, x_7, x_1, x_2, x_3, x_4, x_5); return x_8; } } @@ -43510,29 +43818,58 @@ lean_closure_set(x_2, 0, x_1); return x_2; } } +static lean_object* _init_l_Lean_Parser_Term_funSimpleBinder_formatter___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_funSimpleBinder_formatter___closed__3; +x_2 = l___regBuiltin_Lean_Parser_Term_simpleBinder_formatter___closed__2; +x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Parser_Term_funSimpleBinder_formatter(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; -x_6 = l_Lean_Parser_Term_funSimpleBinder_formatter___closed__3; -x_7 = l___regBuiltin_Lean_Parser_Term_simpleBinder_formatter___closed__2; -x_8 = l_Lean_PrettyPrinter_Formatter_andthen_formatter(x_6, x_7, x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Parser_Term_simpleBinder_formatter___closed__1; +x_7 = l_Lean_Parser_Term_funSimpleBinder_formatter___closed__4; +x_8 = l_Lean_PrettyPrinter_Formatter_orelse_formatter(x_6, x_7, x_1, x_2, x_3, x_4, x_5); return x_8; } } static lean_object* _init_l_Lean_Parser_Term_funBinder_formatter___closed__1() { _start: { +lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_Parser_Term_funBinder___elambda__1___closed__1; +x_2 = l_Lean_Parser_Term_funBinder___elambda__1___closed__2; +x_3 = 1; +x_4 = lean_box(x_3); +x_5 = lean_box(x_3); +x_6 = lean_alloc_closure((void*)(l_Lean_Parser_mkAntiquot_formatter___boxed), 9, 4); +lean_closure_set(x_6, 0, x_1); +lean_closure_set(x_6, 1, x_2); +lean_closure_set(x_6, 2, x_4); +lean_closure_set(x_6, 3, x_5); +return x_6; +} +} +static lean_object* _init_l_Lean_Parser_Term_funBinder_formatter___closed__2() { +_start: +{ lean_object* x_1; x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_funSimpleBinder_formatter), 5, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_Term_funBinder_formatter___closed__2() { +static lean_object* _init_l_Lean_Parser_Term_funBinder_formatter___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_funBinder_formatter___closed__1; +x_1 = l_Lean_Parser_Term_funBinder_formatter___closed__2; x_2 = l_Lean_Parser_Term_tupleTail_formatter___closed__3; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_orelse_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -43540,19 +43877,19 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_funBinder_formatter___closed__3() { +static lean_object* _init_l_Lean_Parser_Term_funBinder_formatter___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_bracketedBinder_formatter___closed__2; -x_2 = l_Lean_Parser_Term_funBinder_formatter___closed__2; +x_2 = l_Lean_Parser_Term_funBinder_formatter___closed__3; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_orelse_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_funBinder_formatter___closed__4() { +static lean_object* _init_l_Lean_Parser_Term_funBinder_formatter___closed__5() { _start: { lean_object* x_1; @@ -43560,19 +43897,19 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_funImplicitBinder_formatter) return x_1; } } -static lean_object* _init_l_Lean_Parser_Term_funBinder_formatter___closed__5() { +static lean_object* _init_l_Lean_Parser_Term_funBinder_formatter___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_funBinder_formatter___closed__4; -x_2 = l_Lean_Parser_Term_funBinder_formatter___closed__3; +x_1 = l_Lean_Parser_Term_funBinder_formatter___closed__5; +x_2 = l_Lean_Parser_Term_funBinder_formatter___closed__4; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_orelse_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_funBinder_formatter___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_funBinder_formatter___closed__7() { _start: { lean_object* x_1; @@ -43580,12 +43917,24 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_funStrictImplicitBinder_form return x_1; } } +static lean_object* _init_l_Lean_Parser_Term_funBinder_formatter___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_funBinder_formatter___closed__7; +x_2 = l_Lean_Parser_Term_funBinder_formatter___closed__6; +x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_orelse_formatter), 7, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Parser_Term_funBinder_formatter(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; -x_6 = l_Lean_Parser_Term_funBinder_formatter___closed__6; -x_7 = l_Lean_Parser_Term_funBinder_formatter___closed__5; +x_6 = l_Lean_Parser_Term_funBinder_formatter___closed__1; +x_7 = l_Lean_Parser_Term_funBinder_formatter___closed__8; x_8 = l_Lean_PrettyPrinter_Formatter_orelse_formatter(x_6, x_7, x_1, x_2, x_3, x_4, x_5); return x_8; } @@ -44011,13 +44360,25 @@ lean_closure_set(x_3, 0, x_2); return x_3; } } +static lean_object* _init_l_Lean_Parser_Term_funImplicitBinder_parenthesizer___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_funImplicitBinder_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Term_funImplicitBinder_parenthesizer___closed__6; +x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Parser_Term_funImplicitBinder_parenthesizer(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; -x_6 = l_Lean_Parser_Term_funImplicitBinder_parenthesizer___closed__5; -x_7 = l_Lean_Parser_Term_funImplicitBinder_parenthesizer___closed__6; -x_8 = l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer(x_6, x_7, x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Parser_Term_implicitBinder_parenthesizer___closed__1; +x_7 = l_Lean_Parser_Term_funImplicitBinder_parenthesizer___closed__7; +x_8 = l_Lean_PrettyPrinter_Parenthesizer_withAntiquot_parenthesizer(x_6, x_7, x_1, x_2, x_3, x_4, x_5); return x_8; } } @@ -44053,29 +44414,58 @@ lean_closure_set(x_2, 0, x_1); return x_2; } } +static lean_object* _init_l_Lean_Parser_Term_funSimpleBinder_parenthesizer___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_funSimpleBinder_parenthesizer___closed__3; +x_2 = l___regBuiltin_Lean_Parser_Term_simpleBinder_parenthesizer___closed__2; +x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Parser_Term_funSimpleBinder_parenthesizer(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; -x_6 = l_Lean_Parser_Term_funSimpleBinder_parenthesizer___closed__3; -x_7 = l___regBuiltin_Lean_Parser_Term_simpleBinder_parenthesizer___closed__2; -x_8 = l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer(x_6, x_7, x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Parser_Term_simpleBinder_parenthesizer___closed__1; +x_7 = l_Lean_Parser_Term_funSimpleBinder_parenthesizer___closed__4; +x_8 = l_Lean_PrettyPrinter_Parenthesizer_withAntiquot_parenthesizer(x_6, x_7, x_1, x_2, x_3, x_4, x_5); return x_8; } } static lean_object* _init_l_Lean_Parser_Term_funBinder_parenthesizer___closed__1() { _start: { +lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_Parser_Term_funBinder___elambda__1___closed__1; +x_2 = l_Lean_Parser_Term_funBinder___elambda__1___closed__2; +x_3 = 1; +x_4 = lean_box(x_3); +x_5 = lean_box(x_3); +x_6 = lean_alloc_closure((void*)(l_Lean_Parser_mkAntiquot_parenthesizer___boxed), 9, 4); +lean_closure_set(x_6, 0, x_1); +lean_closure_set(x_6, 1, x_2); +lean_closure_set(x_6, 2, x_4); +lean_closure_set(x_6, 3, x_5); +return x_6; +} +} +static lean_object* _init_l_Lean_Parser_Term_funBinder_parenthesizer___closed__2() { +_start: +{ lean_object* x_1; x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_funSimpleBinder_parenthesizer), 5, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_Term_funBinder_parenthesizer___closed__2() { +static lean_object* _init_l_Lean_Parser_Term_funBinder_parenthesizer___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_funBinder_parenthesizer___closed__1; +x_1 = l_Lean_Parser_Term_funBinder_parenthesizer___closed__2; x_2 = l_Lean_Parser_Term_explicit_parenthesizer___closed__3; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_orelse_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -44083,19 +44473,19 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_funBinder_parenthesizer___closed__3() { +static lean_object* _init_l_Lean_Parser_Term_funBinder_parenthesizer___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_bracketedBinder_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Term_funBinder_parenthesizer___closed__2; +x_2 = l_Lean_Parser_Term_funBinder_parenthesizer___closed__3; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_orelse_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_funBinder_parenthesizer___closed__4() { +static lean_object* _init_l_Lean_Parser_Term_funBinder_parenthesizer___closed__5() { _start: { lean_object* x_1; @@ -44103,19 +44493,19 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_funImplicitBinder_parenthesi return x_1; } } -static lean_object* _init_l_Lean_Parser_Term_funBinder_parenthesizer___closed__5() { +static lean_object* _init_l_Lean_Parser_Term_funBinder_parenthesizer___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_funBinder_parenthesizer___closed__4; -x_2 = l_Lean_Parser_Term_funBinder_parenthesizer___closed__3; +x_1 = l_Lean_Parser_Term_funBinder_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Term_funBinder_parenthesizer___closed__4; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_orelse_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_funBinder_parenthesizer___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_funBinder_parenthesizer___closed__7() { _start: { lean_object* x_1; @@ -44123,13 +44513,25 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_funStrictImplicitBinder_pare return x_1; } } +static lean_object* _init_l_Lean_Parser_Term_funBinder_parenthesizer___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_funBinder_parenthesizer___closed__7; +x_2 = l_Lean_Parser_Term_funBinder_parenthesizer___closed__6; +x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_orelse_parenthesizer), 7, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Parser_Term_funBinder_parenthesizer(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; -x_6 = l_Lean_Parser_Term_funBinder_parenthesizer___closed__6; -x_7 = l_Lean_Parser_Term_funBinder_parenthesizer___closed__5; -x_8 = l_Lean_PrettyPrinter_Parenthesizer_orelse_parenthesizer(x_6, x_7, x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Parser_Term_funBinder_parenthesizer___closed__1; +x_7 = l_Lean_Parser_Term_funBinder_parenthesizer___closed__8; +x_8 = l_Lean_PrettyPrinter_Parenthesizer_withAntiquot_parenthesizer(x_6, x_7, x_1, x_2, x_3, x_4, x_5); return x_8; } } @@ -44667,7 +45069,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -44687,7 +45089,7 @@ lean_inc(x_1); x_19 = l_Lean_Parser_atomicFn(x_18, x_1, x_12); x_20 = lean_ctor_get(x_19, 4); lean_inc(x_20); -x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_20, x_14); +x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_20, x_14); lean_dec(x_20); if (x_21 == 0) { @@ -44697,7 +45099,7 @@ x_22 = l_Lean_Parser_Term_withAnonymousAntiquot___elambda__1___closed__2; x_23 = l_Lean_Parser_ParserState_mkNode(x_19, x_22, x_17); x_24 = lean_ctor_get(x_23, 4); lean_inc(x_24); -x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_14); +x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_14); lean_dec(x_24); if (x_25 == 0) { @@ -44722,7 +45124,7 @@ lean_inc(x_1); x_30 = l_Lean_Parser_orelseFnCore(x_27, x_28, x_29, x_1, x_19); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_14); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_14); lean_dec(x_31); if (x_32 == 0) { @@ -44732,7 +45134,7 @@ x_33 = l_Lean_Parser_Term_withAnonymousAntiquot___elambda__1___closed__2; x_34 = l_Lean_Parser_ParserState_mkNode(x_30, x_33, x_17); x_35 = lean_ctor_get(x_34, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_14); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_14); lean_dec(x_35); if (x_36 == 0) { @@ -44768,7 +45170,7 @@ x_45 = l_Lean_Parser_Term_withAnonymousAntiquot___elambda__1___closed__2; x_46 = l_Lean_Parser_ParserState_mkNode(x_40, x_45, x_17); x_47 = lean_ctor_get(x_46, 4); lean_inc(x_47); -x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_14); +x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_14); lean_dec(x_47); if (x_48 == 0) { @@ -44792,7 +45194,7 @@ x_51 = l_Lean_Parser_Term_withAnonymousAntiquot___elambda__1___closed__2; x_52 = l_Lean_Parser_ParserState_mkNode(x_50, x_51, x_17); x_53 = lean_ctor_get(x_52, 4); lean_inc(x_53); -x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); +x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); lean_dec(x_53); if (x_54 == 0) { @@ -45127,7 +45529,7 @@ x_16 = l_Lean_Parser_checkPrecFn(x_15, x_1, x_2); x_17 = lean_ctor_get(x_16, 4); lean_inc(x_17); x_18 = lean_box(0); -x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); +x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); lean_dec(x_17); if (x_19 == 0) { @@ -45173,7 +45575,7 @@ goto block_61; lean_object* x_30; uint8_t x_31; x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_18); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_18); lean_dec(x_30); if (x_31 == 0) { @@ -45184,7 +45586,7 @@ x_32 = l_Lean_Parser_Term_leading__parser___elambda__1___closed__2; x_33 = l_Lean_Parser_ParserState_mkNode(x_29, x_32, x_21); x_34 = lean_ctor_get(x_33, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_18); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_18); lean_dec(x_34); if (x_35 == 0) { @@ -45206,7 +45608,7 @@ lean_inc(x_1); x_37 = lean_apply_2(x_6, x_1, x_29); x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_18); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_18); lean_dec(x_38); if (x_39 == 0) { @@ -45216,7 +45618,7 @@ x_40 = l_Lean_Parser_Term_leading__parser___elambda__1___closed__2; x_41 = l_Lean_Parser_ParserState_mkNode(x_37, x_40, x_21); x_42 = lean_ctor_get(x_41, 4); lean_inc(x_42); -x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_18); +x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_18); lean_dec(x_42); if (x_43 == 0) { @@ -45238,7 +45640,7 @@ lean_inc(x_1); x_45 = lean_apply_2(x_4, x_1, x_37); x_46 = lean_ctor_get(x_45, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_18); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_18); lean_dec(x_46); if (x_47 == 0) { @@ -45247,7 +45649,7 @@ x_48 = l_Lean_Parser_Term_leading__parser___elambda__1___closed__2; x_49 = l_Lean_Parser_ParserState_mkNode(x_45, x_48, x_21); x_50 = lean_ctor_get(x_49, 4); lean_inc(x_50); -x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_18); +x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_18); lean_dec(x_50); if (x_51 == 0) { @@ -45273,7 +45675,7 @@ x_56 = l_Lean_Parser_Term_leading__parser___elambda__1___closed__2; x_57 = l_Lean_Parser_ParserState_mkNode(x_55, x_56, x_21); x_58 = lean_ctor_get(x_57, 4); lean_inc(x_58); -x_59 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_58, x_18); +x_59 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_58, x_18); lean_dec(x_58); if (x_59 == 0) { @@ -45438,7 +45840,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_leading__parser_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(170u); +x_1 = lean_unsigned_to_nat(173u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -45450,7 +45852,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_leading__parser_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(170u); +x_1 = lean_unsigned_to_nat(173u); x_2 = lean_unsigned_to_nat(156u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -45478,7 +45880,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_leading__parser_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(170u); +x_1 = lean_unsigned_to_nat(173u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -45490,7 +45892,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_leading__parser_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(170u); +x_1 = lean_unsigned_to_nat(173u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -46335,7 +46737,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -46380,7 +46782,7 @@ goto block_59; lean_object* x_28; uint8_t x_29; x_28 = lean_ctor_get(x_27, 4); lean_inc(x_28); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_16); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_16); lean_dec(x_28); if (x_29 == 0) { @@ -46390,7 +46792,7 @@ x_30 = l_Lean_Parser_Term_trailing__parser___elambda__1___closed__2; x_31 = l_Lean_Parser_ParserState_mkNode(x_27, x_30, x_19); x_32 = lean_ctor_get(x_31, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_16); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_16); lean_dec(x_32); if (x_33 == 0) { @@ -46413,7 +46815,7 @@ lean_inc(x_1); x_35 = lean_apply_2(x_4, x_1, x_27); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_16); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_16); lean_dec(x_36); if (x_37 == 0) { @@ -46423,7 +46825,7 @@ x_38 = l_Lean_Parser_Term_trailing__parser___elambda__1___closed__2; x_39 = l_Lean_Parser_ParserState_mkNode(x_35, x_38, x_19); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_16); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_16); lean_dec(x_40); if (x_41 == 0) { @@ -46445,7 +46847,7 @@ lean_inc(x_1); x_43 = lean_apply_2(x_4, x_1, x_35); x_44 = lean_ctor_get(x_43, 4); lean_inc(x_44); -x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_16); +x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_16); lean_dec(x_44); if (x_45 == 0) { @@ -46454,7 +46856,7 @@ x_46 = l_Lean_Parser_Term_trailing__parser___elambda__1___closed__2; x_47 = l_Lean_Parser_ParserState_mkNode(x_43, x_46, x_19); x_48 = lean_ctor_get(x_47, 4); lean_inc(x_48); -x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_16); +x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_16); lean_dec(x_48); if (x_49 == 0) { @@ -46480,7 +46882,7 @@ x_54 = l_Lean_Parser_Term_trailing__parser___elambda__1___closed__2; x_55 = l_Lean_Parser_ParserState_mkNode(x_53, x_54, x_19); x_56 = lean_ctor_get(x_55, 4); lean_inc(x_56); -x_57 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_56, x_16); +x_57 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_56, x_16); lean_dec(x_56); if (x_57 == 0) { @@ -46644,7 +47046,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_trailing__parser_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(171u); +x_1 = lean_unsigned_to_nat(174u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -46656,7 +47058,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_trailing__parser_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(171u); +x_1 = lean_unsigned_to_nat(174u); x_2 = lean_unsigned_to_nat(144u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -46684,7 +47086,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_trailing__parser_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(171u); +x_1 = lean_unsigned_to_nat(174u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -46696,7 +47098,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_trailing__parser_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(171u); +x_1 = lean_unsigned_to_nat(174u); x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -47142,7 +47544,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -47173,7 +47575,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_20, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -47182,7 +47584,7 @@ x_27 = l_Lean_Parser_Term_borrowed___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_20, x_27, x_17); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -47208,7 +47610,7 @@ x_35 = l_Lean_Parser_Term_borrowed___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_17); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); lean_dec(x_37); if (x_38 == 0) { @@ -47231,7 +47633,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_20); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); lean_dec(x_41); if (x_42 == 0) { @@ -47240,7 +47642,7 @@ x_43 = l_Lean_Parser_Term_borrowed___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_17); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); lean_dec(x_45); if (x_46 == 0) { @@ -47266,7 +47668,7 @@ x_51 = l_Lean_Parser_Term_borrowed___elambda__1___closed__2; x_52 = l_Lean_Parser_ParserState_mkNode(x_50, x_51, x_17); x_53 = lean_ctor_get(x_52, 4); lean_inc(x_53); -x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); +x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); lean_dec(x_53); if (x_54 == 0) { @@ -47413,7 +47815,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_borrowed_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(173u); +x_1 = lean_unsigned_to_nat(176u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -47425,7 +47827,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_borrowed_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(173u); +x_1 = lean_unsigned_to_nat(176u); x_2 = lean_unsigned_to_nat(82u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -47453,7 +47855,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_borrowed_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(173u); +x_1 = lean_unsigned_to_nat(176u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -47465,7 +47867,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_borrowed_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(173u); +x_1 = lean_unsigned_to_nat(176u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -47811,7 +48213,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -47831,7 +48233,7 @@ x_19 = l_Lean_Parser_Term_quotedName___elambda__1___closed__2; x_20 = l_Lean_Parser_ParserState_mkNode(x_18, x_19, x_17); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); -x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_14); +x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_14); lean_dec(x_21); if (x_22 == 0) { @@ -47946,7 +48348,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_quotedName_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(174u); +x_1 = lean_unsigned_to_nat(177u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -47958,7 +48360,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_quotedName_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(174u); +x_1 = lean_unsigned_to_nat(177u); x_2 = lean_unsigned_to_nat(61u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -47986,7 +48388,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_quotedName_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(174u); +x_1 = lean_unsigned_to_nat(177u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -47998,7 +48400,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_quotedName_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(174u); +x_1 = lean_unsigned_to_nat(177u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -48421,7 +48823,7 @@ x_13 = l_Lean_Parser_checkPrecFn(x_12, x_1, x_2); x_14 = lean_ctor_get(x_13, 4); lean_inc(x_14); x_15 = lean_box(0); -x_16 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_14, x_15); +x_16 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_14, x_15); lean_dec(x_14); if (x_16 == 0) { @@ -48465,7 +48867,7 @@ goto block_58; lean_object* x_27; uint8_t x_28; x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_15); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_15); lean_dec(x_27); if (x_28 == 0) { @@ -48474,7 +48876,7 @@ x_29 = l_Lean_Parser_Term_doubleQuotedName___elambda__1___closed__2; x_30 = l_Lean_Parser_ParserState_mkNode(x_26, x_29, x_18); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_15); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_15); lean_dec(x_31); if (x_32 == 0) { @@ -48496,7 +48898,7 @@ x_34 = l_Lean_Parser_Term_doubleQuotedName___elambda__1___closed__9; x_35 = l_Lean_Parser_checkNoWsBeforeFn(x_34, x_1, x_26); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_15); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_15); lean_dec(x_36); if (x_37 == 0) { @@ -48505,7 +48907,7 @@ x_38 = l_Lean_Parser_Term_doubleQuotedName___elambda__1___closed__2; x_39 = l_Lean_Parser_ParserState_mkNode(x_35, x_38, x_18); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_15); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_15); lean_dec(x_40); if (x_41 == 0) { @@ -48527,7 +48929,7 @@ x_43 = 0; x_44 = l_Lean_Parser_rawCh___elambda__1(x_3, x_43, x_1, x_35); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_15); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_15); lean_dec(x_45); if (x_46 == 0) { @@ -48536,7 +48938,7 @@ x_47 = l_Lean_Parser_Term_doubleQuotedName___elambda__1___closed__2; x_48 = l_Lean_Parser_ParserState_mkNode(x_44, x_47, x_18); x_49 = lean_ctor_get(x_48, 4); lean_inc(x_49); -x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_15); +x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_15); lean_dec(x_49); if (x_50 == 0) { @@ -48560,7 +48962,7 @@ x_53 = l_Lean_Parser_Term_doubleQuotedName___elambda__1___closed__2; x_54 = l_Lean_Parser_ParserState_mkNode(x_52, x_53, x_18); x_55 = lean_ctor_get(x_54, 4); lean_inc(x_55); -x_56 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_55, x_15); +x_56 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_55, x_15); lean_dec(x_55); if (x_56 == 0) { @@ -48731,7 +49133,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_doubleQuotedName_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(177u); +x_1 = lean_unsigned_to_nat(180u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -48743,7 +49145,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_doubleQuotedName_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(177u); +x_1 = lean_unsigned_to_nat(180u); x_2 = lean_unsigned_to_nat(126u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -48771,7 +49173,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_doubleQuotedName_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(177u); +x_1 = lean_unsigned_to_nat(180u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -48783,7 +49185,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_doubleQuotedName_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(177u); +x_1 = lean_unsigned_to_nat(180u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -49179,6 +49581,150 @@ x_1 = l_Lean_Parser_Term_simpleBinderWithoutType___closed__4; return x_1; } } +static lean_object* _init_l_Lean_Parser_Term_letIdBinder___elambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("letIdBinder", 11); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Term_letIdBinder___elambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_byTactic___elambda__1___closed__2; +x_2 = l_Lean_Parser_Term_letIdBinder___elambda__1___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Term_letIdBinder___elambda__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Term_letIdBinder___elambda__1___closed__1; +x_2 = l_Lean_Parser_Term_letIdBinder___elambda__1___closed__2; +x_3 = 1; +x_4 = l_Lean_Parser_mkAntiquot(x_1, x_2, x_3, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Term_letIdBinder___elambda__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Parser_Term_simpleBinderWithoutType; +x_2 = lean_ctor_get(x_1, 1); +lean_inc(x_2); +x_3 = l_Lean_Parser_Term_forall___elambda__1___closed__8; +x_4 = lean_ctor_get(x_3, 1); +lean_inc(x_4); +x_5 = lean_alloc_closure((void*)(l_Lean_Parser_orelseFn), 4, 2); +lean_closure_set(x_5, 0, x_2); +lean_closure_set(x_5, 1, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Parser_Term_letIdBinder___elambda__1(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; lean_object* x_11; uint32_t x_12; uint32_t x_13; uint8_t x_14; +x_3 = l_Lean_Parser_Term_simpleBinderWithoutType; +x_4 = lean_ctor_get(x_3, 1); +lean_inc(x_4); +x_5 = l_Lean_Parser_Term_forall___elambda__1___closed__8; +x_6 = lean_ctor_get(x_5, 1); +lean_inc(x_6); +x_7 = l_Lean_Parser_Term_letIdBinder___elambda__1___closed__3; +x_8 = lean_ctor_get(x_7, 1); +lean_inc(x_8); +x_9 = lean_ctor_get(x_1, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +lean_dec(x_9); +x_11 = lean_ctor_get(x_2, 2); +lean_inc(x_11); +x_12 = lean_string_utf8_get(x_10, x_11); +lean_dec(x_11); +lean_dec(x_10); +x_13 = 36; +x_14 = lean_uint32_dec_eq(x_12, x_13); +if (x_14 == 0) +{ +uint8_t x_15; lean_object* x_16; +lean_dec(x_8); +x_15 = 1; +x_16 = l_Lean_Parser_orelseFnCore(x_4, x_6, x_15, x_1, x_2); +return x_16; +} +else +{ +lean_object* x_17; uint8_t x_18; lean_object* x_19; +lean_dec(x_6); +lean_dec(x_4); +x_17 = l_Lean_Parser_Term_letIdBinder___elambda__1___closed__4; +x_18 = 0; +x_19 = l_Lean_Parser_orelseFnCore(x_8, x_17, x_18, x_1, x_2); +return x_19; +} +} +} +static lean_object* _init_l_Lean_Parser_Term_letIdBinder___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Parser_Term_simpleBinderWithoutType; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_3 = l_Lean_Parser_Term_forall___elambda__1___closed__8; +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +x_5 = l_Lean_Parser_orelseInfo(x_2, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Parser_Term_letIdBinder___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Term_letIdBinder___elambda__1___closed__3; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_3 = l_Lean_Parser_Term_letIdBinder___closed__1; +x_4 = l_Lean_Parser_orelseInfo(x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Term_letIdBinder___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_letIdBinder___elambda__1), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Term_letIdBinder___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_letIdBinder___closed__2; +x_2 = l_Lean_Parser_Term_letIdBinder___closed__3; +x_3 = lean_alloc_ctor(0, 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_Lean_Parser_Term_letIdBinder() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Parser_Term_letIdBinder___closed__4; +return x_1; +} +} static lean_object* _init_l_Lean_Parser_Term_letIdLhs___elambda__1___closed__1() { _start: { @@ -49204,77 +49750,49 @@ return x_3; static lean_object* _init_l_Lean_Parser_Term_letIdLhs___elambda__1___closed__3() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l_Lean_Parser_Term_simpleBinderWithoutType; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Term_letIdBinder; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l_Lean_Parser_Term_forall___elambda__1___closed__8; -x_4 = lean_ctor_get(x_3, 0); -lean_inc(x_4); -x_5 = l_Lean_Parser_orelseInfo(x_2, x_4); -return x_5; +x_3 = l_Lean_Parser_epsilonInfo; +x_4 = l_Lean_Parser_andthenInfo(x_3, x_2); +return x_4; } } static lean_object* _init_l_Lean_Parser_Term_letIdLhs___elambda__1___closed__4() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l_Lean_Parser_Term_simpleBinderWithoutType; -x_2 = lean_ctor_get(x_1, 1); -lean_inc(x_2); -x_3 = l_Lean_Parser_Term_forall___elambda__1___closed__8; -x_4 = lean_ctor_get(x_3, 1); -lean_inc(x_4); -x_5 = lean_alloc_closure((void*)(l_Lean_Parser_orelseFn), 4, 2); -lean_closure_set(x_5, 0, x_2); -lean_closure_set(x_5, 1, x_4); -return x_5; -} -} -static lean_object* _init_l_Lean_Parser_Term_letIdLhs___elambda__1___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_epsilonInfo; -x_2 = l_Lean_Parser_Term_letIdLhs___elambda__1___closed__3; -x_3 = l_Lean_Parser_andthenInfo(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Term_letIdLhs___elambda__1___closed__6() { -_start: -{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_docComment___elambda__1___closed__12; -x_2 = l_Lean_Parser_Term_letIdLhs___elambda__1___closed__4; +x_2 = l_Lean_Parser_Term_letIdBinder___closed__3; x_3 = lean_alloc_closure((void*)(l_Lean_Parser_andthenFn), 4, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_letIdLhs___elambda__1___closed__7() { +static lean_object* _init_l_Lean_Parser_Term_letIdLhs___elambda__1___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_letIdLhs___elambda__1___closed__5; -x_2 = l_Lean_Parser_Term_letIdLhs___elambda__1___closed__6; +x_1 = l_Lean_Parser_Term_letIdLhs___elambda__1___closed__3; +x_2 = l_Lean_Parser_Term_letIdLhs___elambda__1___closed__4; x_3 = lean_alloc_ctor(0, 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_Lean_Parser_Term_letIdLhs___elambda__1___closed__8() { +static lean_object* _init_l_Lean_Parser_Term_letIdLhs___elambda__1___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_letIdLhs___elambda__1___closed__7; +x_1 = l_Lean_Parser_Term_letIdLhs___elambda__1___closed__5; x_2 = l_Lean_Parser_many(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_letIdLhs___elambda__1___closed__9() { +static lean_object* _init_l_Lean_Parser_Term_letIdLhs___elambda__1___closed__7() { _start: { lean_object* x_1; @@ -49286,7 +49804,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_letIdLhs___elambda__1(lean_object* x _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; uint8_t x_10; -x_3 = l_Lean_Parser_Term_letIdLhs___elambda__1___closed__8; +x_3 = l_Lean_Parser_Term_letIdLhs___elambda__1___closed__6; x_4 = lean_ctor_get(x_3, 1); lean_inc(x_4); x_5 = l_Lean_Parser_Term_optType; @@ -49297,7 +49815,7 @@ x_7 = l_Lean_Parser_Term_ident___elambda__1(x_1, x_2); x_8 = lean_ctor_get(x_7, 4); lean_inc(x_8); x_9 = lean_box(0); -x_10 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_8, x_9); +x_10 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_8, x_9); lean_dec(x_8); if (x_10 == 0) { @@ -49310,12 +49828,12 @@ else { lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; x_11 = l_Lean_Parser_Term_letIdLhs___elambda__1___closed__2; -x_12 = l_Lean_Parser_Term_letIdLhs___elambda__1___closed__9; +x_12 = l_Lean_Parser_Term_letIdLhs___elambda__1___closed__7; lean_inc(x_1); x_13 = l_Lean_Parser_notFollowedByFn(x_11, x_12, x_1, x_7); x_14 = lean_ctor_get(x_13, 4); lean_inc(x_14); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_14, x_9); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_14, x_9); lean_dec(x_14); if (x_15 == 0) { @@ -49331,7 +49849,7 @@ lean_inc(x_1); x_16 = lean_apply_2(x_4, x_1, x_13); x_17 = lean_ctor_get(x_16, 4); lean_inc(x_17); -x_18 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_9); +x_18 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_9); lean_dec(x_17); if (x_18 == 0) { @@ -49353,7 +49871,7 @@ static lean_object* _init_l_Lean_Parser_Term_letIdLhs___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l_Lean_Parser_Term_letIdLhs___elambda__1___closed__8; +x_1 = l_Lean_Parser_Term_letIdLhs___elambda__1___closed__6; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l_Lean_Parser_Term_optType; @@ -49540,7 +50058,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -49559,7 +50077,7 @@ lean_inc(x_1); x_19 = l_Lean_Parser_atomicFn(x_18, x_1, x_12); x_20 = lean_ctor_get(x_19, 4); lean_inc(x_20); -x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_20, x_14); +x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_20, x_14); lean_dec(x_20); if (x_21 == 0) { @@ -49568,7 +50086,7 @@ x_22 = l_Lean_Parser_Term_letIdDecl___elambda__1___closed__2; x_23 = l_Lean_Parser_ParserState_mkNode(x_19, x_22, x_17); x_24 = lean_ctor_get(x_23, 4); lean_inc(x_24); -x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_14); +x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_14); lean_dec(x_24); if (x_25 == 0) { @@ -49594,7 +50112,7 @@ x_30 = l_Lean_Parser_Term_letIdDecl___elambda__1___closed__2; x_31 = l_Lean_Parser_ParserState_mkNode(x_29, x_30, x_17); x_32 = lean_ctor_get(x_31, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_14); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_14); lean_dec(x_32); if (x_33 == 0) { @@ -49868,7 +50386,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -49887,7 +50405,7 @@ lean_inc(x_1); x_19 = l_Lean_Parser_atomicFn(x_18, x_1, x_12); x_20 = lean_ctor_get(x_19, 4); lean_inc(x_20); -x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_20, x_14); +x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_20, x_14); lean_dec(x_20); if (x_21 == 0) { @@ -49896,7 +50414,7 @@ x_22 = l_Lean_Parser_Term_letPatDecl___elambda__1___closed__2; x_23 = l_Lean_Parser_ParserState_mkNode(x_19, x_22, x_17); x_24 = lean_ctor_get(x_23, 4); lean_inc(x_24); -x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_14); +x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_14); lean_dec(x_24); if (x_25 == 0) { @@ -49922,7 +50440,7 @@ x_30 = l_Lean_Parser_Term_letPatDecl___elambda__1___closed__2; x_31 = l_Lean_Parser_ParserState_mkNode(x_29, x_30, x_17); x_32 = lean_ctor_get(x_31, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_14); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_14); lean_dec(x_32); if (x_33 == 0) { @@ -50189,7 +50707,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -50208,7 +50726,7 @@ lean_inc(x_1); x_20 = l_Lean_Parser_Term_letIdLhs___elambda__1(x_1, x_14); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); -x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_16); +x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_16); lean_dec(x_21); if (x_22 == 0) { @@ -50218,7 +50736,7 @@ x_23 = l_Lean_Parser_Term_letEqnsDecl___elambda__1___closed__2; x_24 = l_Lean_Parser_ParserState_mkNode(x_20, x_23, x_19); x_25 = lean_ctor_get(x_24, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_16); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_16); lean_dec(x_25); if (x_26 == 0) { @@ -50244,7 +50762,7 @@ x_31 = l_Lean_Parser_Term_letEqnsDecl___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_19); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_16); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_16); lean_dec(x_33); if (x_34 == 0) { @@ -50534,7 +51052,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -50554,7 +51072,7 @@ lean_inc(x_1); x_20 = l_Lean_Parser_notFollowedByFn(x_18, x_19, x_1, x_12); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); -x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_14); +x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_14); lean_dec(x_21); if (x_22 == 0) { @@ -50563,7 +51081,7 @@ x_23 = l_Lean_Parser_Term_letDecl___elambda__1___closed__2; x_24 = l_Lean_Parser_ParserState_mkNode(x_20, x_23, x_17); x_25 = lean_ctor_get(x_24, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -50590,7 +51108,7 @@ x_32 = l_Lean_Parser_Term_letDecl___elambda__1___closed__2; x_33 = l_Lean_Parser_ParserState_mkNode(x_31, x_32, x_17); x_34 = lean_ctor_get(x_33, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_14); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_14); lean_dec(x_34); if (x_35 == 0) { @@ -50761,7 +51279,7 @@ if (x_17 == 0) lean_object* x_19; uint8_t x_20; x_19 = lean_ctor_get(x_12, 4); lean_inc(x_19); -x_20 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_18); +x_20 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_18); lean_dec(x_19); if (x_20 == 0) { @@ -50782,7 +51300,7 @@ lean_inc(x_2); x_22 = l_Lean_Parser_tokenAntiquotFn(x_2, x_12); x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); -x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_18); +x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_18); lean_dec(x_23); if (x_24 == 0) { @@ -50849,7 +51367,7 @@ if (x_44 == 0) lean_object* x_46; uint8_t x_47; x_46 = lean_ctor_get(x_39, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_45); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_45); lean_dec(x_46); if (x_47 == 0) { @@ -50870,7 +51388,7 @@ lean_inc(x_35); x_49 = l_Lean_Parser_tokenAntiquotFn(x_35, x_39); x_50 = lean_ctor_get(x_49, 4); lean_inc(x_50); -x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_45); +x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_45); lean_dec(x_50); if (x_51 == 0) { @@ -51014,7 +51532,7 @@ x_29 = l_Lean_Parser_checkPrecFn(x_28, x_1, x_2); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); x_31 = lean_box(0); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_31); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_31); lean_dec(x_30); if (x_32 == 0) { @@ -51064,7 +51582,7 @@ if (x_43 == 0) lean_object* x_57; uint8_t x_58; x_57 = lean_ctor_get(x_39, 4); lean_inc(x_57); -x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_31); +x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_31); lean_dec(x_57); if (x_58 == 0) { @@ -51087,7 +51605,7 @@ lean_inc(x_37); x_60 = l_Lean_Parser_tokenAntiquotFn(x_37, x_39); x_61 = lean_ctor_get(x_60, 4); lean_inc(x_61); -x_62 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_61, x_31); +x_62 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_61, x_31); lean_dec(x_61); if (x_62 == 0) { @@ -51108,7 +51626,7 @@ goto block_56; lean_object* x_45; uint8_t x_46; x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_31); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_31); lean_dec(x_45); if (x_46 == 0) { @@ -51117,7 +51635,7 @@ lean_dec(x_6); x_47 = l_Lean_Parser_ParserState_mkNode(x_44, x_8, x_34); x_48 = lean_ctor_get(x_47, 4); lean_inc(x_48); -x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_31); +x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_31); lean_dec(x_48); if (x_49 == 0) { @@ -51140,7 +51658,7 @@ x_51 = lean_apply_2(x_6, x_1, x_44); x_52 = l_Lean_Parser_ParserState_mkNode(x_51, x_8, x_34); x_53 = lean_ctor_get(x_52, 4); lean_inc(x_53); -x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_31); +x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_31); lean_dec(x_53); if (x_54 == 0) { @@ -51304,7 +51822,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_let_declRange___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(202u); +x_1 = lean_unsigned_to_nat(206u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -51316,7 +51834,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_let_declRange___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(202u); +x_1 = lean_unsigned_to_nat(206u); x_2 = lean_unsigned_to_nat(118u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -51344,7 +51862,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_let_declRange___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(202u); +x_1 = lean_unsigned_to_nat(206u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -51356,7 +51874,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_let_declRange___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(202u); +x_1 = lean_unsigned_to_nat(206u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -51474,6 +51992,45 @@ x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); return x_6; } } +static lean_object* _init_l_Lean_Parser_Term_letIdBinder_formatter___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_Parser_Term_letIdBinder___elambda__1___closed__1; +x_2 = l_Lean_Parser_Term_letIdBinder___elambda__1___closed__2; +x_3 = 1; +x_4 = lean_box(x_3); +x_5 = lean_box(x_3); +x_6 = lean_alloc_closure((void*)(l_Lean_Parser_mkAntiquot_formatter___boxed), 9, 4); +lean_closure_set(x_6, 0, x_1); +lean_closure_set(x_6, 1, x_2); +lean_closure_set(x_6, 2, x_4); +lean_closure_set(x_6, 3, x_5); +return x_6; +} +} +static lean_object* _init_l_Lean_Parser_Term_letIdBinder_formatter___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___regBuiltin_Lean_Parser_Term_simpleBinderWithoutType_formatter___closed__4; +x_2 = l_Lean_Parser_Term_forall_formatter___closed__3; +x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_orelse_formatter), 7, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Parser_Term_letIdBinder_formatter(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; +x_6 = l_Lean_Parser_Term_letIdBinder_formatter___closed__1; +x_7 = l_Lean_Parser_Term_letIdBinder_formatter___closed__2; +x_8 = l_Lean_PrettyPrinter_Formatter_orelse_formatter(x_6, x_7, x_1, x_2, x_3, x_4, x_5); +return x_8; +} +} static lean_object* _init_l_Lean_Parser_Term_letIdLhs_formatter___closed__1() { _start: { @@ -51499,13 +52056,9 @@ return x_2; static lean_object* _init_l_Lean_Parser_Term_letIdLhs_formatter___closed__3() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Parser_Term_simpleBinderWithoutType_formatter___closed__4; -x_2 = l_Lean_Parser_Term_forall_formatter___closed__3; -x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_orelse_formatter), 7, 2); -lean_closure_set(x_3, 0, x_1); -lean_closure_set(x_3, 1, x_2); -return x_3; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_letIdBinder_formatter), 5, 0); +return x_1; } } static lean_object* _init_l_Lean_Parser_Term_letIdLhs_formatter___closed__4() { @@ -52201,6 +52754,45 @@ x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); return x_6; } } +static lean_object* _init_l_Lean_Parser_Term_letIdBinder_parenthesizer___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_Parser_Term_letIdBinder___elambda__1___closed__1; +x_2 = l_Lean_Parser_Term_letIdBinder___elambda__1___closed__2; +x_3 = 1; +x_4 = lean_box(x_3); +x_5 = lean_box(x_3); +x_6 = lean_alloc_closure((void*)(l_Lean_Parser_mkAntiquot_parenthesizer___boxed), 9, 4); +lean_closure_set(x_6, 0, x_1); +lean_closure_set(x_6, 1, x_2); +lean_closure_set(x_6, 2, x_4); +lean_closure_set(x_6, 3, x_5); +return x_6; +} +} +static lean_object* _init_l_Lean_Parser_Term_letIdBinder_parenthesizer___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___regBuiltin_Lean_Parser_Term_simpleBinderWithoutType_parenthesizer___closed__2; +x_2 = l_Lean_Parser_Term_forall_parenthesizer___closed__3; +x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_orelse_parenthesizer), 7, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Parser_Term_letIdBinder_parenthesizer(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; +x_6 = l_Lean_Parser_Term_letIdBinder_parenthesizer___closed__1; +x_7 = l_Lean_Parser_Term_letIdBinder_parenthesizer___closed__2; +x_8 = l_Lean_PrettyPrinter_Parenthesizer_withAntiquot_parenthesizer(x_6, x_7, x_1, x_2, x_3, x_4, x_5); +return x_8; +} +} static lean_object* _init_l_Lean_Parser_Term_letIdLhs_parenthesizer___closed__1() { _start: { @@ -52226,13 +52818,9 @@ return x_2; static lean_object* _init_l_Lean_Parser_Term_letIdLhs_parenthesizer___closed__3() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Parser_Term_simpleBinderWithoutType_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Term_forall_parenthesizer___closed__3; -x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_orelse_parenthesizer), 7, 2); -lean_closure_set(x_3, 0, x_1); -lean_closure_set(x_3, 1, x_2); -return x_3; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_letIdBinder_parenthesizer), 5, 0); +return x_1; } } static lean_object* _init_l_Lean_Parser_Term_letIdLhs_parenthesizer___closed__4() { @@ -52895,7 +53483,7 @@ x_10 = l_Lean_Parser_orelseFnCore(x_1, x_2, x_9, x_3, x_4); x_11 = lean_ctor_get(x_10, 4); lean_inc(x_11); x_12 = lean_box(0); -x_13 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_11, x_12); +x_13 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_11, x_12); lean_dec(x_11); if (x_13 == 0) { @@ -52945,7 +53533,7 @@ x_26 = l_Lean_Parser_orelseFnCore(x_1, x_2, x_25, x_24, x_4); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); x_28 = lean_box(0); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_28); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_28); lean_dec(x_27); if (x_29 == 0) { @@ -53108,7 +53696,7 @@ x_30 = l_Lean_Parser_checkPrecFn(x_29, x_1, x_2); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); x_32 = lean_box(0); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_32); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_32); lean_dec(x_31); if (x_33 == 0) { @@ -53147,7 +53735,7 @@ lean_inc(x_38); x_40 = l_Lean_Parser_orelseFnCore(x_3, x_4, x_39, x_38, x_30); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_32); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_32); lean_dec(x_41); if (x_42 == 0) { @@ -53157,7 +53745,7 @@ lean_dec(x_7); x_43 = l_Lean_Parser_ParserState_mkNode(x_40, x_9, x_35); x_44 = lean_ctor_get(x_43, 4); lean_inc(x_44); -x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_32); +x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_32); lean_dec(x_44); if (x_45 == 0) { @@ -53178,7 +53766,7 @@ lean_object* x_47; lean_object* x_48; uint8_t x_49; x_47 = l_Lean_Parser_Term_letDecl___elambda__1(x_38, x_40); x_48 = lean_ctor_get(x_47, 4); lean_inc(x_48); -x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_32); +x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_32); lean_dec(x_48); if (x_49 == 0) { @@ -53187,7 +53775,7 @@ lean_dec(x_7); x_50 = l_Lean_Parser_ParserState_mkNode(x_47, x_9, x_35); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_32); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_32); lean_dec(x_51); if (x_52 == 0) { @@ -53210,7 +53798,7 @@ x_54 = lean_apply_2(x_7, x_1, x_47); x_55 = l_Lean_Parser_ParserState_mkNode(x_54, x_9, x_35); x_56 = lean_ctor_get(x_55, 4); lean_inc(x_56); -x_57 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_56, x_32); +x_57 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_56, x_32); lean_dec(x_56); if (x_57 == 0) { @@ -53383,7 +53971,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_let__fun_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(203u); +x_1 = lean_unsigned_to_nat(207u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -53395,7 +53983,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_let__fun_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(203u); +x_1 = lean_unsigned_to_nat(207u); x_2 = lean_unsigned_to_nat(150u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -53423,7 +54011,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_let__fun_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(203u); +x_1 = lean_unsigned_to_nat(207u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -53435,7 +54023,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_let__fun_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(203u); +x_1 = lean_unsigned_to_nat(207u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -53924,7 +54512,7 @@ x_20 = l_Lean_Parser_checkPrecFn(x_19, x_1, x_2); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); x_22 = lean_box(0); -x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_22); +x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_22); lean_dec(x_21); if (x_23 == 0) { @@ -53975,7 +54563,7 @@ if (x_35 == 0) lean_object* x_51; uint8_t x_52; x_51 = lean_ctor_get(x_31, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_22); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_22); lean_dec(x_51); if (x_52 == 0) { @@ -53998,7 +54586,7 @@ lean_inc(x_28); x_54 = l_Lean_Parser_tokenAntiquotFn(x_28, x_31); x_55 = lean_ctor_get(x_54, 4); lean_inc(x_55); -x_56 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_55, x_22); +x_56 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_55, x_22); lean_dec(x_55); if (x_56 == 0) { @@ -54019,7 +54607,7 @@ goto block_50; lean_object* x_37; uint8_t x_38; x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_22); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_22); lean_dec(x_37); if (x_38 == 0) { @@ -54029,7 +54617,7 @@ x_39 = l_Lean_Parser_Term_let__delayed___elambda__1___closed__2; x_40 = l_Lean_Parser_ParserState_mkNode(x_36, x_39, x_25); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_22); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_22); lean_dec(x_41); if (x_42 == 0) { @@ -54053,7 +54641,7 @@ x_45 = l_Lean_Parser_Term_let__delayed___elambda__1___closed__2; x_46 = l_Lean_Parser_ParserState_mkNode(x_44, x_45, x_25); x_47 = lean_ctor_get(x_46, 4); lean_inc(x_47); -x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_22); +x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_22); lean_dec(x_47); if (x_48 == 0) { @@ -54209,7 +54797,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_let__delayed_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(204u); +x_1 = lean_unsigned_to_nat(208u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -54221,7 +54809,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_let__delayed_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(204u); +x_1 = lean_unsigned_to_nat(208u); x_2 = lean_unsigned_to_nat(133u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -54249,7 +54837,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_let__delayed_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(204u); +x_1 = lean_unsigned_to_nat(208u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -54261,7 +54849,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_let__delayed_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(204u); +x_1 = lean_unsigned_to_nat(208u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -54706,7 +55294,7 @@ x_20 = l_Lean_Parser_checkPrecFn(x_19, x_1, x_2); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); x_22 = lean_box(0); -x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_22); +x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_22); lean_dec(x_21); if (x_23 == 0) { @@ -54757,7 +55345,7 @@ if (x_35 == 0) lean_object* x_51; uint8_t x_52; x_51 = lean_ctor_get(x_31, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_22); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_22); lean_dec(x_51); if (x_52 == 0) { @@ -54780,7 +55368,7 @@ lean_inc(x_28); x_54 = l_Lean_Parser_tokenAntiquotFn(x_28, x_31); x_55 = lean_ctor_get(x_54, 4); lean_inc(x_55); -x_56 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_55, x_22); +x_56 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_55, x_22); lean_dec(x_55); if (x_56 == 0) { @@ -54801,7 +55389,7 @@ goto block_50; lean_object* x_37; uint8_t x_38; x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_22); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_22); lean_dec(x_37); if (x_38 == 0) { @@ -54811,7 +55399,7 @@ x_39 = l_Lean_Parser_Term_let__tmp___elambda__1___closed__2; x_40 = l_Lean_Parser_ParserState_mkNode(x_36, x_39, x_25); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_22); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_22); lean_dec(x_41); if (x_42 == 0) { @@ -54835,7 +55423,7 @@ x_45 = l_Lean_Parser_Term_let__tmp___elambda__1___closed__2; x_46 = l_Lean_Parser_ParserState_mkNode(x_44, x_45, x_25); x_47 = lean_ctor_get(x_46, 4); lean_inc(x_47); -x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_22); +x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_22); lean_dec(x_47); if (x_48 == 0) { @@ -54991,7 +55579,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_let__tmp_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(206u); +x_1 = lean_unsigned_to_nat(210u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -55003,7 +55591,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_let__tmp_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(206u); +x_1 = lean_unsigned_to_nat(210u); x_2 = lean_unsigned_to_nat(125u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -55031,7 +55619,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_let__tmp_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(206u); +x_1 = lean_unsigned_to_nat(210u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -55043,7 +55631,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_let__tmp_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(206u); +x_1 = lean_unsigned_to_nat(210u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -55321,6 +55909,22 @@ x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); return x_6; } } +LEAN_EXPORT lean_object* l_Lean_Parser_Term_instCoeTSyntaxConsSyntaxNodeKindMkStrAnonymousNil__1(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Parser_Term_instCoeTSyntaxConsSyntaxNodeKindMkStrAnonymousNil__1___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Parser_Term_instCoeTSyntaxConsSyntaxNodeKindMkStrAnonymousNil__1(x_1); +lean_dec(x_1); +return x_2; +} +} static lean_object* _init_l_Lean_Parser_Term_haveIdLhs___elambda__1___closed__1() { _start: { @@ -55328,7 +55932,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l_Lean_Parser_Term_ident; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l_Lean_Parser_Term_letIdLhs___elambda__1___closed__8; +x_3 = l_Lean_Parser_Term_letIdLhs___elambda__1___closed__6; x_4 = lean_ctor_get(x_3, 0); lean_inc(x_4); x_5 = l_Lean_Parser_andthenInfo(x_2, x_4); @@ -55339,7 +55943,7 @@ static lean_object* _init_l_Lean_Parser_Term_haveIdLhs___elambda__1___closed__2( _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Term_letIdLhs___elambda__1___closed__8; +x_1 = l_Lean_Parser_Term_letIdLhs___elambda__1___closed__6; x_2 = lean_ctor_get(x_1, 1); lean_inc(x_2); x_3 = l_Lean_Parser_Term_ident___closed__2; @@ -55385,7 +55989,7 @@ x_7 = lean_apply_2(x_4, x_1, x_2); x_8 = lean_ctor_get(x_7, 4); lean_inc(x_8); x_9 = lean_box(0); -x_10 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_8, x_9); +x_10 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_8, x_9); lean_dec(x_8); if (x_10 == 0) { @@ -55570,7 +56174,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -55589,7 +56193,7 @@ lean_inc(x_1); x_19 = l_Lean_Parser_atomicFn(x_18, x_1, x_12); x_20 = lean_ctor_get(x_19, 4); lean_inc(x_20); -x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_20, x_14); +x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_20, x_14); lean_dec(x_20); if (x_21 == 0) { @@ -55598,7 +56202,7 @@ x_22 = l_Lean_Parser_Term_haveIdDecl___elambda__1___closed__2; x_23 = l_Lean_Parser_ParserState_mkNode(x_19, x_22, x_17); x_24 = lean_ctor_get(x_23, 4); lean_inc(x_24); -x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_14); +x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_14); lean_dec(x_24); if (x_25 == 0) { @@ -55624,7 +56228,7 @@ x_30 = l_Lean_Parser_Term_haveIdDecl___elambda__1___closed__2; x_31 = l_Lean_Parser_ParserState_mkNode(x_29, x_30, x_17); x_32 = lean_ctor_get(x_31, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_14); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_14); lean_dec(x_32); if (x_33 == 0) { @@ -55855,7 +56459,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -55874,7 +56478,7 @@ lean_inc(x_1); x_20 = l_Lean_Parser_Term_haveIdLhs___elambda__1(x_1, x_14); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); -x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_16); +x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_16); lean_dec(x_21); if (x_22 == 0) { @@ -55884,7 +56488,7 @@ x_23 = l_Lean_Parser_Term_haveEqnsDecl___elambda__1___closed__2; x_24 = l_Lean_Parser_ParserState_mkNode(x_20, x_23, x_19); x_25 = lean_ctor_get(x_24, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_16); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_16); lean_dec(x_25); if (x_26 == 0) { @@ -55908,7 +56512,7 @@ x_29 = l_Lean_Parser_Term_haveEqnsDecl___elambda__1___closed__2; x_30 = l_Lean_Parser_ParserState_mkNode(x_28, x_29, x_19); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); lean_dec(x_31); if (x_32 == 0) { @@ -56137,7 +56741,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -56160,7 +56764,7 @@ x_22 = l_Lean_Parser_Term_haveDecl___elambda__1___closed__2; x_23 = l_Lean_Parser_ParserState_mkNode(x_21, x_22, x_17); x_24 = lean_ctor_get(x_23, 4); lean_inc(x_24); -x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_14); +x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_14); lean_dec(x_24); if (x_25 == 0) { @@ -56320,7 +56924,7 @@ if (x_17 == 0) lean_object* x_19; uint8_t x_20; x_19 = lean_ctor_get(x_12, 4); lean_inc(x_19); -x_20 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_18); +x_20 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_18); lean_dec(x_19); if (x_20 == 0) { @@ -56341,7 +56945,7 @@ lean_inc(x_2); x_22 = l_Lean_Parser_tokenAntiquotFn(x_2, x_12); x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); -x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_18); +x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_18); lean_dec(x_23); if (x_24 == 0) { @@ -56408,7 +57012,7 @@ if (x_44 == 0) lean_object* x_46; uint8_t x_47; x_46 = lean_ctor_get(x_39, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_45); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_45); lean_dec(x_46); if (x_47 == 0) { @@ -56429,7 +57033,7 @@ lean_inc(x_35); x_49 = l_Lean_Parser_tokenAntiquotFn(x_35, x_39); x_50 = lean_ctor_get(x_49, 4); lean_inc(x_50); -x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_45); +x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_45); lean_dec(x_50); if (x_51 == 0) { @@ -56573,7 +57177,7 @@ x_29 = l_Lean_Parser_checkPrecFn(x_28, x_1, x_2); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); x_31 = lean_box(0); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_31); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_31); lean_dec(x_30); if (x_32 == 0) { @@ -56623,7 +57227,7 @@ if (x_43 == 0) lean_object* x_57; uint8_t x_58; x_57 = lean_ctor_get(x_39, 4); lean_inc(x_57); -x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_31); +x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_31); lean_dec(x_57); if (x_58 == 0) { @@ -56646,7 +57250,7 @@ lean_inc(x_37); x_60 = l_Lean_Parser_tokenAntiquotFn(x_37, x_39); x_61 = lean_ctor_get(x_60, 4); lean_inc(x_61); -x_62 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_61, x_31); +x_62 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_61, x_31); lean_dec(x_61); if (x_62 == 0) { @@ -56667,7 +57271,7 @@ goto block_56; lean_object* x_45; uint8_t x_46; x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_31); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_31); lean_dec(x_45); if (x_46 == 0) { @@ -56676,7 +57280,7 @@ lean_dec(x_6); x_47 = l_Lean_Parser_ParserState_mkNode(x_44, x_8, x_34); x_48 = lean_ctor_get(x_47, 4); lean_inc(x_48); -x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_31); +x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_31); lean_dec(x_48); if (x_49 == 0) { @@ -56699,7 +57303,7 @@ x_51 = lean_apply_2(x_6, x_1, x_44); x_52 = l_Lean_Parser_ParserState_mkNode(x_51, x_8, x_34); x_53 = lean_ctor_get(x_52, 4); lean_inc(x_53); -x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_31); +x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_31); lean_dec(x_53); if (x_54 == 0) { @@ -56863,7 +57467,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_have_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(213u); +x_1 = lean_unsigned_to_nat(220u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -56875,7 +57479,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_have_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(213u); +x_1 = lean_unsigned_to_nat(220u); x_2 = lean_unsigned_to_nat(120u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -56903,7 +57507,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_have_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(213u); +x_1 = lean_unsigned_to_nat(220u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -56915,7 +57519,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_have_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(213u); +x_1 = lean_unsigned_to_nat(220u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -57979,7 +58583,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -58012,7 +58616,7 @@ x_25 = l_Lean_Parser_Term_scoped___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_20, x_25, x_17); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); lean_dec(x_27); if (x_28 == 0) { @@ -58036,7 +58640,7 @@ x_31 = l_Lean_Parser_Term_scoped___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_17); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -58283,7 +58887,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -58316,7 +58920,7 @@ x_25 = l_Lean_Parser_Term_local___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_20, x_25, x_17); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); lean_dec(x_27); if (x_28 == 0) { @@ -58340,7 +58944,7 @@ x_31 = l_Lean_Parser_Term_local___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_17); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -58593,7 +59197,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -58614,7 +59218,7 @@ x_21 = l_Lean_Parser_Term_attrKind___elambda__1___closed__2; x_22 = l_Lean_Parser_ParserState_mkNode(x_20, x_21, x_19); x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); -x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_16); +x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_16); lean_dec(x_23); if (x_24 == 0) { @@ -58849,7 +59453,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -58867,7 +59471,7 @@ lean_inc(x_1); x_18 = l_Lean_Parser_Term_attrKind___elambda__1(x_1, x_12); x_19 = lean_ctor_get(x_18, 4); lean_inc(x_19); -x_20 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_14); +x_20 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_14); lean_dec(x_19); if (x_20 == 0) { @@ -58876,7 +59480,7 @@ x_21 = l_Lean_Parser_Term_attrInstance___elambda__1___closed__2; x_22 = l_Lean_Parser_ParserState_mkNode(x_18, x_21, x_17); x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); -x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_14); +x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_14); lean_dec(x_23); if (x_24 == 0) { @@ -58902,7 +59506,7 @@ x_29 = l_Lean_Parser_Term_attrInstance___elambda__1___closed__2; x_30 = l_Lean_Parser_ParserState_mkNode(x_28, x_29, x_17); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_14); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_14); lean_dec(x_31); if (x_32 == 0) { @@ -59204,7 +59808,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -59248,7 +59852,7 @@ goto block_60; lean_object* x_28; uint8_t x_29; x_28 = lean_ctor_get(x_27, 4); lean_inc(x_28); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_16); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_16); lean_dec(x_28); if (x_29 == 0) { @@ -59259,7 +59863,7 @@ x_30 = l_Lean_Parser_Term_attributes___elambda__1___closed__2; x_31 = l_Lean_Parser_ParserState_mkNode(x_27, x_30, x_19); x_32 = lean_ctor_get(x_31, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_16); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_16); lean_dec(x_32); if (x_33 == 0) { @@ -59281,7 +59885,7 @@ lean_inc(x_1); x_35 = lean_apply_2(x_4, x_1, x_27); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_16); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_16); lean_dec(x_36); if (x_37 == 0) { @@ -59291,7 +59895,7 @@ x_38 = l_Lean_Parser_Term_attributes___elambda__1___closed__2; x_39 = l_Lean_Parser_ParserState_mkNode(x_35, x_38, x_19); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_16); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_16); lean_dec(x_40); if (x_41 == 0) { @@ -59326,7 +59930,7 @@ x_49 = l_Lean_Parser_Term_attributes___elambda__1___closed__2; x_50 = l_Lean_Parser_ParserState_mkNode(x_45, x_49, x_19); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); lean_dec(x_51); if (x_52 == 0) { @@ -59350,7 +59954,7 @@ x_55 = l_Lean_Parser_Term_attributes___elambda__1___closed__2; x_56 = l_Lean_Parser_ParserState_mkNode(x_54, x_55, x_19); x_57 = lean_ctor_get(x_56, 4); lean_inc(x_57); -x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_16); +x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_16); lean_dec(x_57); if (x_58 == 0) { @@ -59629,7 +60233,7 @@ x_16 = l_Lean_Parser_checkPrecFn(x_15, x_1, x_2); x_17 = lean_ctor_get(x_16, 4); lean_inc(x_17); x_18 = lean_box(0); -x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); +x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); lean_dec(x_17); if (x_19 == 0) { @@ -59649,7 +60253,7 @@ lean_inc(x_1); x_22 = lean_apply_2(x_6, x_1, x_16); x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); -x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_18); +x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_18); lean_dec(x_23); if (x_24 == 0) { @@ -59659,7 +60263,7 @@ x_25 = l_Lean_Parser_Term_letRecDecl___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_22, x_25, x_21); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_18); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_18); lean_dec(x_27); if (x_28 == 0) { @@ -59681,7 +60285,7 @@ lean_inc(x_1); x_30 = lean_apply_2(x_4, x_1, x_22); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_18); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_18); lean_dec(x_31); if (x_32 == 0) { @@ -59690,7 +60294,7 @@ x_33 = l_Lean_Parser_Term_letRecDecl___elambda__1___closed__2; x_34 = l_Lean_Parser_ParserState_mkNode(x_30, x_33, x_21); x_35 = lean_ctor_get(x_34, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_18); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_18); lean_dec(x_35); if (x_36 == 0) { @@ -59714,7 +60318,7 @@ x_39 = l_Lean_Parser_Term_letRecDecl___elambda__1___closed__2; x_40 = l_Lean_Parser_ParserState_mkNode(x_38, x_39, x_21); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_18); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_18); lean_dec(x_41); if (x_42 == 0) { @@ -59951,7 +60555,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -59972,7 +60576,7 @@ x_21 = l_Lean_Parser_Term_letRecDecls___elambda__1___closed__2; x_22 = l_Lean_Parser_ParserState_mkNode(x_20, x_21, x_19); x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); -x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_16); +x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_16); lean_dec(x_23); if (x_24 == 0) { @@ -60071,140 +60675,137 @@ x_1 = l_Lean_Parser_Term_letRecDecls___closed__6; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Term_letrec___elambda__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Parser_Term_letrec___elambda__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -uint8_t x_5; -x_5 = !lean_is_exclusive(x_3); -if (x_5 == 0) +uint8_t x_6; +x_6 = !lean_is_exclusive(x_4); +if (x_6 == 0) { -lean_object* x_6; lean_object* x_7; 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; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint32_t x_18; uint32_t x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; -x_6 = lean_ctor_get(x_3, 0); -x_7 = lean_ctor_get(x_3, 5); -lean_dec(x_7); -x_8 = lean_ctor_get(x_4, 2); -lean_inc(x_8); -x_9 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_9, 0, x_8); -lean_inc(x_6); -lean_ctor_set(x_3, 5, x_9); -x_10 = lean_ctor_get(x_4, 0); -lean_inc(x_10); -x_11 = lean_array_get_size(x_10); -lean_dec(x_10); -x_12 = l_Lean_Parser_Command_docComment___elambda__1___lambda__1___closed__1; -x_13 = lean_string_append(x_12, x_1); -x_14 = lean_string_append(x_13, x_12); -lean_inc(x_3); -x_15 = l_Lean_Parser_symbolFnAux(x_1, x_14, x_3, x_4); -x_16 = lean_ctor_get(x_6, 0); -lean_inc(x_16); -lean_dec(x_6); -x_17 = lean_ctor_get(x_15, 2); +lean_object* x_7; 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; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint32_t x_19; uint32_t x_20; uint8_t x_21; lean_object* x_22; lean_object* x_23; +x_7 = lean_ctor_get(x_4, 0); +x_8 = lean_ctor_get(x_4, 5); +lean_dec(x_8); +x_9 = lean_ctor_get(x_5, 2); +lean_inc(x_9); +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_9); +lean_inc(x_7); +lean_ctor_set(x_4, 5, x_10); +x_11 = lean_ctor_get(x_5, 0); +lean_inc(x_11); +x_12 = lean_array_get_size(x_11); +lean_dec(x_11); +x_13 = l_Lean_Parser_Command_docComment___elambda__1___lambda__1___closed__1; +x_14 = lean_string_append(x_13, x_1); +x_15 = lean_string_append(x_14, x_13); +lean_inc(x_4); +x_16 = l_Lean_Parser_symbolFnAux(x_1, x_15, x_4, x_5); +x_17 = lean_ctor_get(x_7, 0); lean_inc(x_17); -x_18 = lean_string_utf8_get(x_16, x_17); -lean_dec(x_17); -x_19 = 37; -x_20 = lean_uint32_dec_eq(x_18, x_19); -x_21 = lean_box(0); -if (x_20 == 0) +lean_dec(x_7); +x_18 = lean_ctor_get(x_16, 2); +lean_inc(x_18); +x_19 = lean_string_utf8_get(x_17, x_18); +lean_dec(x_18); +x_20 = 37; +x_21 = lean_uint32_dec_eq(x_19, x_20); +x_22 = lean_box(0); +if (x_21 == 0) { -x_22 = x_15; -goto block_47; +x_23 = x_16; +goto block_45; } else { -lean_object* x_48; -lean_inc(x_3); -x_48 = l_Lean_Parser_tokenAntiquotFn(x_3, x_15); -x_22 = x_48; -goto block_47; +lean_object* x_46; +lean_inc(x_4); +x_46 = l_Lean_Parser_tokenAntiquotFn(x_4, x_16); +x_23 = x_46; +goto block_45; } -block_47: +block_45: { -lean_object* x_23; uint8_t x_24; -x_23 = lean_ctor_get(x_22, 4); -lean_inc(x_23); -x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_21); -lean_dec(x_23); -if (x_24 == 0) +lean_object* x_24; uint8_t x_25; +x_24 = lean_ctor_get(x_23, 4); +lean_inc(x_24); +x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_22); +lean_dec(x_24); +if (x_25 == 0) { -lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; -lean_dec(x_16); -lean_dec(x_2); -x_25 = l_Lean_groupKind; -x_26 = l_Lean_Parser_ParserState_mkNode(x_22, x_25, x_11); +lean_object* x_26; lean_object* x_27; uint8_t x_28; +lean_dec(x_17); +lean_dec(x_3); +x_26 = l_Lean_Parser_ParserState_mkNode(x_23, x_2, x_12); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_21); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_22); lean_dec(x_27); if (x_28 == 0) { -lean_dec(x_3); +lean_dec(x_4); return x_26; } else { lean_object* x_29; -x_29 = l_Lean_Parser_Term_letRecDecls___elambda__1(x_3, x_26); +x_29 = l_Lean_Parser_Term_letRecDecls___elambda__1(x_4, x_26); return x_29; } } else { lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint32_t x_34; uint8_t x_35; -x_30 = lean_string_append(x_12, x_2); -x_31 = lean_string_append(x_30, x_12); -lean_inc(x_3); -x_32 = l_Lean_Parser_nonReservedSymbolFnAux(x_2, x_31, x_3, x_22); +x_30 = lean_string_append(x_13, x_3); +x_31 = lean_string_append(x_30, x_13); +lean_inc(x_4); +x_32 = l_Lean_Parser_nonReservedSymbolFnAux(x_3, x_31, x_4, x_23); x_33 = lean_ctor_get(x_32, 2); lean_inc(x_33); -x_34 = lean_string_utf8_get(x_16, x_33); +x_34 = lean_string_utf8_get(x_17, x_33); lean_dec(x_33); -lean_dec(x_16); -x_35 = lean_uint32_dec_eq(x_34, x_19); +lean_dec(x_17); +x_35 = lean_uint32_dec_eq(x_34, x_20); if (x_35 == 0) { -lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; -x_36 = l_Lean_groupKind; -x_37 = l_Lean_Parser_ParserState_mkNode(x_32, x_36, x_11); -x_38 = lean_ctor_get(x_37, 4); -lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_21); -lean_dec(x_38); -if (x_39 == 0) +lean_object* x_36; lean_object* x_37; uint8_t x_38; +x_36 = l_Lean_Parser_ParserState_mkNode(x_32, x_2, x_12); +x_37 = lean_ctor_get(x_36, 4); +lean_inc(x_37); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_22); +lean_dec(x_37); +if (x_38 == 0) { -lean_dec(x_3); -return x_37; +lean_dec(x_4); +return x_36; } else { -lean_object* x_40; -x_40 = l_Lean_Parser_Term_letRecDecls___elambda__1(x_3, x_37); -return x_40; +lean_object* x_39; +x_39 = l_Lean_Parser_Term_letRecDecls___elambda__1(x_4, x_36); +return x_39; } } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; -lean_inc(x_3); -x_41 = l_Lean_Parser_tokenAntiquotFn(x_3, x_32); -x_42 = l_Lean_groupKind; -x_43 = l_Lean_Parser_ParserState_mkNode(x_41, x_42, x_11); -x_44 = lean_ctor_get(x_43, 4); -lean_inc(x_44); -x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_21); -lean_dec(x_44); -if (x_45 == 0) +lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; +lean_inc(x_4); +x_40 = l_Lean_Parser_tokenAntiquotFn(x_4, x_32); +x_41 = l_Lean_Parser_ParserState_mkNode(x_40, x_2, x_12); +x_42 = lean_ctor_get(x_41, 4); +lean_inc(x_42); +x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_22); +lean_dec(x_42); +if (x_43 == 0) { -lean_dec(x_3); -return x_43; +lean_dec(x_4); +return x_41; } else { -lean_object* x_46; -x_46 = l_Lean_Parser_Term_letRecDecls___elambda__1(x_3, x_43); -return x_46; +lean_object* x_44; +x_44 = l_Lean_Parser_Term_letRecDecls___elambda__1(x_4, x_41); +return x_44; } } } @@ -60212,152 +60813,149 @@ return x_46; } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; uint32_t x_67; uint32_t x_68; uint8_t x_69; lean_object* x_70; lean_object* x_71; -x_49 = lean_ctor_get(x_3, 0); -x_50 = lean_ctor_get(x_3, 1); -x_51 = lean_ctor_get(x_3, 2); -x_52 = lean_ctor_get(x_3, 3); -x_53 = lean_ctor_get(x_3, 4); -x_54 = lean_ctor_get_uint8(x_3, sizeof(void*)*7); -x_55 = lean_ctor_get(x_3, 6); -lean_inc(x_55); +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; uint32_t x_65; uint32_t x_66; uint8_t x_67; lean_object* x_68; lean_object* x_69; +x_47 = lean_ctor_get(x_4, 0); +x_48 = lean_ctor_get(x_4, 1); +x_49 = lean_ctor_get(x_4, 2); +x_50 = lean_ctor_get(x_4, 3); +x_51 = lean_ctor_get(x_4, 4); +x_52 = lean_ctor_get_uint8(x_4, sizeof(void*)*7); +x_53 = lean_ctor_get(x_4, 6); lean_inc(x_53); -lean_inc(x_52); lean_inc(x_51); lean_inc(x_50); lean_inc(x_49); -lean_dec(x_3); -x_56 = lean_ctor_get(x_4, 2); +lean_inc(x_48); +lean_inc(x_47); +lean_dec(x_4); +x_54 = lean_ctor_get(x_5, 2); +lean_inc(x_54); +x_55 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_55, 0, x_54); +lean_inc(x_47); +x_56 = lean_alloc_ctor(0, 7, 1); +lean_ctor_set(x_56, 0, x_47); +lean_ctor_set(x_56, 1, x_48); +lean_ctor_set(x_56, 2, x_49); +lean_ctor_set(x_56, 3, x_50); +lean_ctor_set(x_56, 4, x_51); +lean_ctor_set(x_56, 5, x_55); +lean_ctor_set(x_56, 6, x_53); +lean_ctor_set_uint8(x_56, sizeof(void*)*7, x_52); +x_57 = lean_ctor_get(x_5, 0); +lean_inc(x_57); +x_58 = lean_array_get_size(x_57); +lean_dec(x_57); +x_59 = l_Lean_Parser_Command_docComment___elambda__1___lambda__1___closed__1; +x_60 = lean_string_append(x_59, x_1); +x_61 = lean_string_append(x_60, x_59); lean_inc(x_56); -x_57 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_57, 0, x_56); -lean_inc(x_49); -x_58 = lean_alloc_ctor(0, 7, 1); -lean_ctor_set(x_58, 0, x_49); -lean_ctor_set(x_58, 1, x_50); -lean_ctor_set(x_58, 2, x_51); -lean_ctor_set(x_58, 3, x_52); -lean_ctor_set(x_58, 4, x_53); -lean_ctor_set(x_58, 5, x_57); -lean_ctor_set(x_58, 6, x_55); -lean_ctor_set_uint8(x_58, sizeof(void*)*7, x_54); -x_59 = lean_ctor_get(x_4, 0); -lean_inc(x_59); -x_60 = lean_array_get_size(x_59); -lean_dec(x_59); -x_61 = l_Lean_Parser_Command_docComment___elambda__1___lambda__1___closed__1; -x_62 = lean_string_append(x_61, x_1); -x_63 = lean_string_append(x_62, x_61); -lean_inc(x_58); -x_64 = l_Lean_Parser_symbolFnAux(x_1, x_63, x_58, x_4); -x_65 = lean_ctor_get(x_49, 0); -lean_inc(x_65); -lean_dec(x_49); -x_66 = lean_ctor_get(x_64, 2); -lean_inc(x_66); -x_67 = lean_string_utf8_get(x_65, x_66); -lean_dec(x_66); -x_68 = 37; -x_69 = lean_uint32_dec_eq(x_67, x_68); -x_70 = lean_box(0); -if (x_69 == 0) +x_62 = l_Lean_Parser_symbolFnAux(x_1, x_61, x_56, x_5); +x_63 = lean_ctor_get(x_47, 0); +lean_inc(x_63); +lean_dec(x_47); +x_64 = lean_ctor_get(x_62, 2); +lean_inc(x_64); +x_65 = lean_string_utf8_get(x_63, x_64); +lean_dec(x_64); +x_66 = 37; +x_67 = lean_uint32_dec_eq(x_65, x_66); +x_68 = lean_box(0); +if (x_67 == 0) { -x_71 = x_64; -goto block_96; +x_69 = x_62; +goto block_91; } else { -lean_object* x_97; -lean_inc(x_58); -x_97 = l_Lean_Parser_tokenAntiquotFn(x_58, x_64); -x_71 = x_97; -goto block_96; +lean_object* x_92; +lean_inc(x_56); +x_92 = l_Lean_Parser_tokenAntiquotFn(x_56, x_62); +x_69 = x_92; +goto block_91; } -block_96: +block_91: { -lean_object* x_72; uint8_t x_73; -x_72 = lean_ctor_get(x_71, 4); -lean_inc(x_72); -x_73 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_72, x_70); -lean_dec(x_72); -if (x_73 == 0) +lean_object* x_70; uint8_t x_71; +x_70 = lean_ctor_get(x_69, 4); +lean_inc(x_70); +x_71 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_70, x_68); +lean_dec(x_70); +if (x_71 == 0) { -lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; -lean_dec(x_65); -lean_dec(x_2); -x_74 = l_Lean_groupKind; -x_75 = l_Lean_Parser_ParserState_mkNode(x_71, x_74, x_60); -x_76 = lean_ctor_get(x_75, 4); -lean_inc(x_76); -x_77 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_76, x_70); -lean_dec(x_76); -if (x_77 == 0) +lean_object* x_72; lean_object* x_73; uint8_t x_74; +lean_dec(x_63); +lean_dec(x_3); +x_72 = l_Lean_Parser_ParserState_mkNode(x_69, x_2, x_58); +x_73 = lean_ctor_get(x_72, 4); +lean_inc(x_73); +x_74 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_73, x_68); +lean_dec(x_73); +if (x_74 == 0) { -lean_dec(x_58); -return x_75; +lean_dec(x_56); +return x_72; } else { -lean_object* x_78; -x_78 = l_Lean_Parser_Term_letRecDecls___elambda__1(x_58, x_75); -return x_78; +lean_object* x_75; +x_75 = l_Lean_Parser_Term_letRecDecls___elambda__1(x_56, x_72); +return x_75; } } else { -lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; uint32_t x_83; uint8_t x_84; -x_79 = lean_string_append(x_61, x_2); -x_80 = lean_string_append(x_79, x_61); -lean_inc(x_58); -x_81 = l_Lean_Parser_nonReservedSymbolFnAux(x_2, x_80, x_58, x_71); -x_82 = lean_ctor_get(x_81, 2); -lean_inc(x_82); -x_83 = lean_string_utf8_get(x_65, x_82); -lean_dec(x_82); -lean_dec(x_65); -x_84 = lean_uint32_dec_eq(x_83, x_68); -if (x_84 == 0) +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; uint32_t x_80; uint8_t x_81; +x_76 = lean_string_append(x_59, x_3); +x_77 = lean_string_append(x_76, x_59); +lean_inc(x_56); +x_78 = l_Lean_Parser_nonReservedSymbolFnAux(x_3, x_77, x_56, x_69); +x_79 = lean_ctor_get(x_78, 2); +lean_inc(x_79); +x_80 = lean_string_utf8_get(x_63, x_79); +lean_dec(x_79); +lean_dec(x_63); +x_81 = lean_uint32_dec_eq(x_80, x_66); +if (x_81 == 0) { -lean_object* x_85; lean_object* x_86; lean_object* x_87; uint8_t x_88; -x_85 = l_Lean_groupKind; -x_86 = l_Lean_Parser_ParserState_mkNode(x_81, x_85, x_60); -x_87 = lean_ctor_get(x_86, 4); -lean_inc(x_87); -x_88 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_87, x_70); -lean_dec(x_87); -if (x_88 == 0) +lean_object* x_82; lean_object* x_83; uint8_t x_84; +x_82 = l_Lean_Parser_ParserState_mkNode(x_78, x_2, x_58); +x_83 = lean_ctor_get(x_82, 4); +lean_inc(x_83); +x_84 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_83, x_68); +lean_dec(x_83); +if (x_84 == 0) { -lean_dec(x_58); -return x_86; +lean_dec(x_56); +return x_82; } else { -lean_object* x_89; -x_89 = l_Lean_Parser_Term_letRecDecls___elambda__1(x_58, x_86); -return x_89; +lean_object* x_85; +x_85 = l_Lean_Parser_Term_letRecDecls___elambda__1(x_56, x_82); +return x_85; } } else { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; uint8_t x_94; -lean_inc(x_58); -x_90 = l_Lean_Parser_tokenAntiquotFn(x_58, x_81); -x_91 = l_Lean_groupKind; -x_92 = l_Lean_Parser_ParserState_mkNode(x_90, x_91, x_60); -x_93 = lean_ctor_get(x_92, 4); -lean_inc(x_93); -x_94 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_93, x_70); -lean_dec(x_93); -if (x_94 == 0) +lean_object* x_86; lean_object* x_87; lean_object* x_88; uint8_t x_89; +lean_inc(x_56); +x_86 = l_Lean_Parser_tokenAntiquotFn(x_56, x_78); +x_87 = l_Lean_Parser_ParserState_mkNode(x_86, x_2, x_58); +x_88 = lean_ctor_get(x_87, 4); +lean_inc(x_88); +x_89 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_88, x_68); +lean_dec(x_88); +if (x_89 == 0) { -lean_dec(x_58); -return x_92; +lean_dec(x_56); +return x_87; } else { -lean_object* x_95; -x_95 = l_Lean_Parser_Term_letRecDecls___elambda__1(x_58, x_92); -return x_95; +lean_object* x_90; +x_90 = l_Lean_Parser_Term_letRecDecls___elambda__1(x_56, x_87); +return x_90; } } } @@ -60435,267 +61033,266 @@ return x_3; LEAN_EXPORT lean_object* l_Lean_Parser_Term_letrec___elambda__1(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; 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; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint32_t x_26; uint32_t x_27; uint8_t x_28; +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; 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; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint32_t x_27; uint32_t x_28; uint8_t x_29; x_3 = l_Lean_Parser_Term_let___elambda__1___closed__5; -x_4 = l_Lean_Parser_Term_letrec___elambda__1___closed__5; -x_5 = lean_alloc_closure((void*)(l_Lean_Parser_Term_letrec___elambda__1___lambda__1___boxed), 4, 2); -lean_closure_set(x_5, 0, x_3); -lean_closure_set(x_5, 1, x_4); -x_6 = l_Lean_Parser_Term_suffices___elambda__1___closed__6; -x_7 = lean_ctor_get(x_6, 1); -lean_inc(x_7); -lean_inc(x_7); -x_8 = lean_alloc_closure((void*)(l_Lean_Parser_andthenFn), 4, 2); -lean_closure_set(x_8, 0, x_5); -lean_closure_set(x_8, 1, x_7); -x_9 = l_Lean_Parser_Term_letrec___elambda__1___closed__2; -x_10 = lean_alloc_closure((void*)(l_Lean_Parser_nodeFn), 4, 2); -lean_closure_set(x_10, 0, x_9); -lean_closure_set(x_10, 1, x_8); -x_11 = l_Lean_Parser_Term_byTactic___elambda__1___closed__13; -x_12 = lean_alloc_closure((void*)(l_Lean_Parser_andthenFn), 4, 2); -lean_closure_set(x_12, 0, x_10); -lean_closure_set(x_12, 1, x_11); -x_13 = l_Lean_Parser_Term_byTactic___elambda__1___closed__11; -x_14 = lean_alloc_closure((void*)(l_Lean_Parser_andthenFn), 4, 2); -lean_closure_set(x_14, 0, x_13); -lean_closure_set(x_14, 1, x_12); -x_15 = l_Lean_Parser_Term_letrec___elambda__1___closed__3; -x_16 = lean_ctor_get(x_15, 1); -lean_inc(x_16); -x_17 = lean_ctor_get(x_1, 0); +x_4 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__19; +x_5 = l_Lean_Parser_Term_letrec___elambda__1___closed__5; +x_6 = lean_alloc_closure((void*)(l_Lean_Parser_Term_letrec___elambda__1___lambda__1___boxed), 5, 3); +lean_closure_set(x_6, 0, x_3); +lean_closure_set(x_6, 1, x_4); +lean_closure_set(x_6, 2, x_5); +x_7 = l_Lean_Parser_Term_suffices___elambda__1___closed__6; +x_8 = lean_ctor_get(x_7, 1); +lean_inc(x_8); +lean_inc(x_8); +x_9 = lean_alloc_closure((void*)(l_Lean_Parser_andthenFn), 4, 2); +lean_closure_set(x_9, 0, x_6); +lean_closure_set(x_9, 1, x_8); +x_10 = l_Lean_Parser_Term_letrec___elambda__1___closed__2; +x_11 = lean_alloc_closure((void*)(l_Lean_Parser_nodeFn), 4, 2); +lean_closure_set(x_11, 0, x_10); +lean_closure_set(x_11, 1, x_9); +x_12 = l_Lean_Parser_Term_byTactic___elambda__1___closed__13; +x_13 = lean_alloc_closure((void*)(l_Lean_Parser_andthenFn), 4, 2); +lean_closure_set(x_13, 0, x_11); +lean_closure_set(x_13, 1, x_12); +x_14 = l_Lean_Parser_Term_byTactic___elambda__1___closed__11; +x_15 = lean_alloc_closure((void*)(l_Lean_Parser_andthenFn), 4, 2); +lean_closure_set(x_15, 0, x_14); +lean_closure_set(x_15, 1, x_13); +x_16 = l_Lean_Parser_Term_letrec___elambda__1___closed__3; +x_17 = lean_ctor_get(x_16, 1); lean_inc(x_17); -x_18 = lean_ctor_get(x_1, 1); +x_18 = lean_ctor_get(x_1, 0); lean_inc(x_18); -x_19 = lean_ctor_get(x_1, 2); +x_19 = lean_ctor_get(x_1, 1); lean_inc(x_19); -x_20 = lean_ctor_get(x_1, 3); +x_20 = lean_ctor_get(x_1, 2); lean_inc(x_20); -x_21 = lean_ctor_get(x_1, 4); +x_21 = lean_ctor_get(x_1, 3); lean_inc(x_21); -x_22 = lean_ctor_get_uint8(x_1, sizeof(void*)*7); -x_23 = lean_ctor_get(x_1, 6); -lean_inc(x_23); -x_24 = lean_ctor_get(x_17, 0); +x_22 = lean_ctor_get(x_1, 4); +lean_inc(x_22); +x_23 = lean_ctor_get_uint8(x_1, sizeof(void*)*7); +x_24 = lean_ctor_get(x_1, 6); lean_inc(x_24); -x_25 = lean_ctor_get(x_2, 2); +x_25 = lean_ctor_get(x_18, 0); lean_inc(x_25); -x_26 = lean_string_utf8_get(x_24, x_25); -lean_dec(x_25); -x_27 = 36; -x_28 = lean_uint32_dec_eq(x_26, x_27); -if (x_28 == 0) +x_26 = lean_ctor_get(x_2, 2); +lean_inc(x_26); +x_27 = lean_string_utf8_get(x_25, x_26); +lean_dec(x_26); +x_28 = 36; +x_29 = lean_uint32_dec_eq(x_27, x_28); +if (x_29 == 0) { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; -lean_dec(x_16); -lean_dec(x_14); -x_29 = l_Lean_Parser_leadPrec; -x_30 = l_Lean_Parser_checkPrecFn(x_29, x_1, x_2); -x_31 = lean_ctor_get(x_30, 4); -lean_inc(x_31); -x_32 = lean_box(0); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_32); -lean_dec(x_31); -if (x_33 == 0) +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; +lean_dec(x_17); +lean_dec(x_15); +x_30 = l_Lean_Parser_leadPrec; +x_31 = l_Lean_Parser_checkPrecFn(x_30, x_1, x_2); +x_32 = lean_ctor_get(x_31, 4); +lean_inc(x_32); +x_33 = lean_box(0); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_33); +lean_dec(x_32); +if (x_34 == 0) { +lean_dec(x_25); lean_dec(x_24); -lean_dec(x_23); +lean_dec(x_22); lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_7); +lean_dec(x_8); lean_dec(x_1); -return x_30; +return x_31; } else { -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; uint32_t x_42; uint32_t x_43; uint8_t x_44; lean_object* x_45; lean_object* x_58; -x_34 = lean_ctor_get(x_30, 0); -lean_inc(x_34); -x_35 = lean_array_get_size(x_34); -lean_dec(x_34); -x_36 = lean_ctor_get(x_30, 2); -lean_inc(x_36); -x_37 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_37, 0, x_36); -x_38 = lean_alloc_ctor(0, 7, 1); -lean_ctor_set(x_38, 0, x_17); -lean_ctor_set(x_38, 1, x_18); -lean_ctor_set(x_38, 2, x_19); -lean_ctor_set(x_38, 3, x_20); -lean_ctor_set(x_38, 4, x_21); -lean_ctor_set(x_38, 5, x_37); -lean_ctor_set(x_38, 6, x_23); -lean_ctor_set_uint8(x_38, sizeof(void*)*7, x_22); -x_39 = l_Lean_Parser_Term_let___elambda__1___closed__7; -lean_inc(x_38); -x_40 = l_Lean_Parser_symbolFnAux(x_3, x_39, x_38, x_30); -x_41 = lean_ctor_get(x_40, 2); -lean_inc(x_41); -x_42 = lean_string_utf8_get(x_24, x_41); -lean_dec(x_41); -x_43 = 37; -x_44 = lean_uint32_dec_eq(x_42, x_43); -if (x_44 == 0) +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; uint32_t x_43; uint32_t x_44; uint8_t x_45; lean_object* x_46; lean_object* x_59; +x_35 = lean_ctor_get(x_31, 0); +lean_inc(x_35); +x_36 = lean_array_get_size(x_35); +lean_dec(x_35); +x_37 = lean_ctor_get(x_31, 2); +lean_inc(x_37); +x_38 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_38, 0, x_37); +x_39 = lean_alloc_ctor(0, 7, 1); +lean_ctor_set(x_39, 0, x_18); +lean_ctor_set(x_39, 1, x_19); +lean_ctor_set(x_39, 2, x_20); +lean_ctor_set(x_39, 3, x_21); +lean_ctor_set(x_39, 4, x_22); +lean_ctor_set(x_39, 5, x_38); +lean_ctor_set(x_39, 6, x_24); +lean_ctor_set_uint8(x_39, sizeof(void*)*7, x_23); +x_40 = l_Lean_Parser_Term_let___elambda__1___closed__7; +lean_inc(x_39); +x_41 = l_Lean_Parser_symbolFnAux(x_3, x_40, x_39, x_31); +x_42 = lean_ctor_get(x_41, 2); +lean_inc(x_42); +x_43 = lean_string_utf8_get(x_25, x_42); +lean_dec(x_42); +x_44 = 37; +x_45 = lean_uint32_dec_eq(x_43, x_44); +if (x_45 == 0) { -x_58 = x_40; -goto block_82; +x_59 = x_41; +goto block_80; } else { -lean_object* x_83; -lean_inc(x_38); -x_83 = l_Lean_Parser_tokenAntiquotFn(x_38, x_40); -x_58 = x_83; -goto block_82; +lean_object* x_81; +lean_inc(x_39); +x_81 = l_Lean_Parser_tokenAntiquotFn(x_39, x_41); +x_59 = x_81; +goto block_80; } -block_57: +block_58: { -lean_object* x_46; uint8_t x_47; -x_46 = lean_ctor_get(x_45, 4); -lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_32); -lean_dec(x_46); -if (x_47 == 0) +lean_object* x_47; uint8_t x_48; +x_47 = lean_ctor_get(x_46, 4); +lean_inc(x_47); +x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_33); +lean_dec(x_47); +if (x_48 == 0) { -lean_object* x_48; lean_object* x_49; uint8_t x_50; -lean_dec(x_7); -x_48 = l_Lean_Parser_ParserState_mkNode(x_45, x_9, x_35); -x_49 = lean_ctor_get(x_48, 4); -lean_inc(x_49); -x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_32); -lean_dec(x_49); -if (x_50 == 0) +lean_object* x_49; lean_object* x_50; uint8_t x_51; +lean_dec(x_8); +x_49 = l_Lean_Parser_ParserState_mkNode(x_46, x_10, x_36); +x_50 = lean_ctor_get(x_49, 4); +lean_inc(x_50); +x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_33); +lean_dec(x_50); +if (x_51 == 0) { lean_dec(x_1); -return x_48; +return x_49; } else { -lean_object* x_51; -x_51 = l_Lean_Parser_setLhsPrecFn(x_29, x_1, x_48); +lean_object* x_52; +x_52 = l_Lean_Parser_setLhsPrecFn(x_30, x_1, x_49); lean_dec(x_1); -return x_51; +return x_52; } } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; +lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; lean_inc(x_1); -x_52 = lean_apply_2(x_7, x_1, x_45); -x_53 = l_Lean_Parser_ParserState_mkNode(x_52, x_9, x_35); -x_54 = lean_ctor_get(x_53, 4); -lean_inc(x_54); -x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_32); -lean_dec(x_54); -if (x_55 == 0) +x_53 = lean_apply_2(x_8, x_1, x_46); +x_54 = l_Lean_Parser_ParserState_mkNode(x_53, x_10, x_36); +x_55 = lean_ctor_get(x_54, 4); +lean_inc(x_55); +x_56 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_55, x_33); +lean_dec(x_55); +if (x_56 == 0) { lean_dec(x_1); -return x_53; +return x_54; } else { -lean_object* x_56; -x_56 = l_Lean_Parser_setLhsPrecFn(x_29, x_1, x_53); +lean_object* x_57; +x_57 = l_Lean_Parser_setLhsPrecFn(x_30, x_1, x_54); lean_dec(x_1); -return x_56; +return x_57; } } } -block_82: +block_80: { -lean_object* x_59; uint8_t x_60; -x_59 = lean_ctor_get(x_58, 4); -lean_inc(x_59); -x_60 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_59, x_32); -lean_dec(x_59); -if (x_60 == 0) +lean_object* x_60; uint8_t x_61; +x_60 = lean_ctor_get(x_59, 4); +lean_inc(x_60); +x_61 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_60, x_33); +lean_dec(x_60); +if (x_61 == 0) { -lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; -lean_dec(x_24); -x_61 = l_Lean_groupKind; -lean_inc(x_35); -x_62 = l_Lean_Parser_ParserState_mkNode(x_58, x_61, x_35); +lean_object* x_62; lean_object* x_63; uint8_t x_64; +lean_dec(x_25); +lean_inc(x_36); +x_62 = l_Lean_Parser_ParserState_mkNode(x_59, x_4, x_36); x_63 = lean_ctor_get(x_62, 4); lean_inc(x_63); -x_64 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_63, x_32); +x_64 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_63, x_33); lean_dec(x_63); if (x_64 == 0) { -lean_dec(x_38); -x_45 = x_62; -goto block_57; +lean_dec(x_39); +x_46 = x_62; +goto block_58; } else { lean_object* x_65; -x_65 = l_Lean_Parser_Term_letRecDecls___elambda__1(x_38, x_62); -x_45 = x_65; -goto block_57; +x_65 = l_Lean_Parser_Term_letRecDecls___elambda__1(x_39, x_62); +x_46 = x_65; +goto block_58; } } else { lean_object* x_66; lean_object* x_67; lean_object* x_68; uint32_t x_69; uint8_t x_70; x_66 = l_Lean_Parser_Term_letrec___elambda__1___closed__7; -lean_inc(x_38); -x_67 = l_Lean_Parser_nonReservedSymbolFnAux(x_4, x_66, x_38, x_58); +lean_inc(x_39); +x_67 = l_Lean_Parser_nonReservedSymbolFnAux(x_5, x_66, x_39, x_59); x_68 = lean_ctor_get(x_67, 2); lean_inc(x_68); -x_69 = lean_string_utf8_get(x_24, x_68); +x_69 = lean_string_utf8_get(x_25, x_68); lean_dec(x_68); -lean_dec(x_24); -x_70 = lean_uint32_dec_eq(x_69, x_43); +lean_dec(x_25); +x_70 = lean_uint32_dec_eq(x_69, x_44); if (x_70 == 0) { -lean_object* x_71; lean_object* x_72; lean_object* x_73; uint8_t x_74; -x_71 = l_Lean_groupKind; -lean_inc(x_35); -x_72 = l_Lean_Parser_ParserState_mkNode(x_67, x_71, x_35); -x_73 = lean_ctor_get(x_72, 4); -lean_inc(x_73); -x_74 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_73, x_32); -lean_dec(x_73); -if (x_74 == 0) +lean_object* x_71; lean_object* x_72; uint8_t x_73; +lean_inc(x_36); +x_71 = l_Lean_Parser_ParserState_mkNode(x_67, x_4, x_36); +x_72 = lean_ctor_get(x_71, 4); +lean_inc(x_72); +x_73 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_72, x_33); +lean_dec(x_72); +if (x_73 == 0) { -lean_dec(x_38); -x_45 = x_72; -goto block_57; +lean_dec(x_39); +x_46 = x_71; +goto block_58; } else { -lean_object* x_75; -x_75 = l_Lean_Parser_Term_letRecDecls___elambda__1(x_38, x_72); -x_45 = x_75; -goto block_57; +lean_object* x_74; +x_74 = l_Lean_Parser_Term_letRecDecls___elambda__1(x_39, x_71); +x_46 = x_74; +goto block_58; } } else { -lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; uint8_t x_80; -lean_inc(x_38); -x_76 = l_Lean_Parser_tokenAntiquotFn(x_38, x_67); -x_77 = l_Lean_groupKind; -lean_inc(x_35); -x_78 = l_Lean_Parser_ParserState_mkNode(x_76, x_77, x_35); -x_79 = lean_ctor_get(x_78, 4); -lean_inc(x_79); -x_80 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_79, x_32); -lean_dec(x_79); -if (x_80 == 0) +lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; +lean_inc(x_39); +x_75 = l_Lean_Parser_tokenAntiquotFn(x_39, x_67); +lean_inc(x_36); +x_76 = l_Lean_Parser_ParserState_mkNode(x_75, x_4, x_36); +x_77 = lean_ctor_get(x_76, 4); +lean_inc(x_77); +x_78 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_77, x_33); +lean_dec(x_77); +if (x_78 == 0) { -lean_dec(x_38); -x_45 = x_78; -goto block_57; +lean_dec(x_39); +x_46 = x_76; +goto block_58; } else { -lean_object* x_81; -x_81 = l_Lean_Parser_Term_letRecDecls___elambda__1(x_38, x_78); -x_45 = x_81; -goto block_57; +lean_object* x_79; +x_79 = l_Lean_Parser_Term_letRecDecls___elambda__1(x_39, x_76); +x_46 = x_79; +goto block_58; } } } @@ -60704,18 +61301,18 @@ goto block_57; } else { -uint8_t x_84; lean_object* x_85; +uint8_t x_82; lean_object* x_83; +lean_dec(x_25); lean_dec(x_24); -lean_dec(x_23); +lean_dec(x_22); lean_dec(x_21); lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_7); -x_84 = 0; -x_85 = l_Lean_Parser_orelseFnCore(x_16, x_14, x_84, x_1, x_2); -return x_85; +lean_dec(x_8); +x_82 = 0; +x_83 = l_Lean_Parser_orelseFnCore(x_17, x_15, x_82, x_1, x_2); +return x_83; } } } @@ -60743,7 +61340,7 @@ static lean_object* _init_l_Lean_Parser_Term_letrec___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_groupKind; +x_1 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__19; x_2 = l_Lean_Parser_Term_letrec___closed__2; x_3 = l_Lean_Parser_nodeInfo(x_1, x_2); return x_3; @@ -60843,13 +61440,13 @@ x_1 = l_Lean_Parser_Term_letrec___closed__11; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Term_letrec___elambda__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Parser_Term_letrec___elambda__1___lambda__1___boxed(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; -x_5 = l_Lean_Parser_Term_letrec___elambda__1___lambda__1(x_1, x_2, x_3, x_4); +lean_object* x_6; +x_6 = l_Lean_Parser_Term_letrec___elambda__1___lambda__1(x_1, x_2, x_3, x_4, x_5); lean_dec(x_1); -return x_5; +return x_6; } } LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_letrec(lean_object* x_1) { @@ -60869,7 +61466,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_letrec_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(224u); +x_1 = lean_unsigned_to_nat(231u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -60881,7 +61478,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_letrec_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(224u); +x_1 = lean_unsigned_to_nat(231u); x_2 = lean_unsigned_to_nat(139u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -60909,7 +61506,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_letrec_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(224u); +x_1 = lean_unsigned_to_nat(231u); x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -60921,7 +61518,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_letrec_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(224u); +x_1 = lean_unsigned_to_nat(231u); x_2 = lean_unsigned_to_nat(12u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -63228,7 +63825,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_whereDecls___elambda__1___closed__7; x_2 = lean_ctor_get(x_1, 1); lean_inc(x_2); -x_3 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__21; +x_3 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__23; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_andthenFn), 4, 2); lean_closure_set(x_4, 0, x_3); lean_closure_set(x_4, 1, x_2); @@ -63408,7 +64005,7 @@ x_18 = l_Lean_Parser_checkPrecFn(x_17, x_1, x_2); x_19 = lean_ctor_get(x_18, 4); lean_inc(x_19); x_20 = lean_box(0); -x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_20); +x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_20); lean_dec(x_19); if (x_21 == 0) { @@ -63458,7 +64055,7 @@ goto block_50; lean_object* x_32; uint8_t x_33; x_32 = lean_ctor_get(x_31, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_20); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_20); lean_dec(x_32); if (x_33 == 0) { @@ -63473,7 +64070,7 @@ x_34 = l_Lean_Parser_Term_whereDecls_formatter___closed__2; x_35 = l_Lean_Parser_ParserState_mkNode(x_31, x_34, x_23); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_20); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_20); lean_dec(x_36); if (x_37 == 0) { @@ -63512,7 +64109,7 @@ x_45 = l_Lean_Parser_Term_whereDecls_formatter___closed__2; x_46 = l_Lean_Parser_ParserState_mkNode(x_44, x_45, x_23); x_47 = lean_ctor_get(x_46, 4); lean_inc(x_47); -x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_20); +x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_20); lean_dec(x_47); if (x_48 == 0) { @@ -63951,7 +64548,7 @@ x_16 = l_Lean_Parser_checkPrecFn(x_15, x_1, x_2); x_17 = lean_ctor_get(x_16, 4); lean_inc(x_17); x_18 = lean_box(0); -x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); +x_19 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_17, x_18); lean_dec(x_17); if (x_19 == 0) { @@ -63971,7 +64568,7 @@ lean_inc(x_1); x_22 = lean_apply_2(x_4, x_1, x_16); x_23 = lean_ctor_get(x_22, 4); lean_inc(x_23); -x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_18); +x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_23, x_18); lean_dec(x_23); if (x_24 == 0) { @@ -63981,7 +64578,7 @@ x_25 = l_Lean_Parser_Term_matchAltsWhereDecls_formatter___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_22, x_25, x_21); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_18); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_18); lean_dec(x_27); if (x_28 == 0) { @@ -64005,7 +64602,7 @@ x_31 = l_Lean_Parser_Term_matchAltsWhereDecls_formatter___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_21); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_18); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_18); lean_dec(x_33); if (x_34 == 0) { @@ -64270,7 +64867,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -64301,7 +64898,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_20, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -64310,7 +64907,7 @@ x_27 = l_Lean_Parser_Term_noindex___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_20, x_27, x_17); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -64336,7 +64933,7 @@ x_35 = l_Lean_Parser_Term_noindex___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_17); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); lean_dec(x_37); if (x_38 == 0) { @@ -64359,7 +64956,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_20); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); lean_dec(x_41); if (x_42 == 0) { @@ -64368,7 +64965,7 @@ x_43 = l_Lean_Parser_Term_noindex___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_17); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); lean_dec(x_45); if (x_46 == 0) { @@ -64394,7 +64991,7 @@ x_51 = l_Lean_Parser_Term_noindex___elambda__1___closed__2; x_52 = l_Lean_Parser_ParserState_mkNode(x_50, x_51, x_17); x_53 = lean_ctor_get(x_52, 4); lean_inc(x_53); -x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); +x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); lean_dec(x_53); if (x_54 == 0) { @@ -64531,7 +65128,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_noindex_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(232u); +x_1 = lean_unsigned_to_nat(239u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -64543,7 +65140,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_noindex_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(232u); +x_1 = lean_unsigned_to_nat(239u); x_2 = lean_unsigned_to_nat(84u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -64571,7 +65168,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_noindex_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(232u); +x_1 = lean_unsigned_to_nat(239u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -64583,7 +65180,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_noindex_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(232u); +x_1 = lean_unsigned_to_nat(239u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -65004,7 +65601,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -65050,7 +65647,7 @@ x_26 = l_Lean_Parser_Term_binrel___elambda__1___closed__2; x_27 = l_Lean_Parser_ParserState_mkNode(x_25, x_26, x_17); x_28 = lean_ctor_get(x_27, 4); lean_inc(x_28); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_14); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_14); lean_dec(x_28); if (x_29 == 0) { @@ -65070,7 +65667,7 @@ return x_30; lean_object* x_33; uint8_t x_34; x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -65084,7 +65681,7 @@ lean_inc(x_1); x_35 = l_Lean_Parser_Term_ident___elambda__1(x_1, x_32); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); lean_dec(x_36); if (x_37 == 0) { @@ -65100,7 +65697,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_categoryParser___elambda__1(x_38, x_39, x_1, x_35); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); lean_dec(x_41); if (x_42 == 0) { @@ -65271,7 +65868,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binrel_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(234u); +x_1 = lean_unsigned_to_nat(241u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -65283,7 +65880,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binrel_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(234u); +x_1 = lean_unsigned_to_nat(241u); x_2 = lean_unsigned_to_nat(124u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -65311,7 +65908,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binrel_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(234u); +x_1 = lean_unsigned_to_nat(241u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -65323,7 +65920,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binrel_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(234u); +x_1 = lean_unsigned_to_nat(241u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -65779,7 +66376,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -65825,7 +66422,7 @@ x_26 = l_Lean_Parser_Term_binrel__no__prop___elambda__1___closed__2; x_27 = l_Lean_Parser_ParserState_mkNode(x_25, x_26, x_17); x_28 = lean_ctor_get(x_27, 4); lean_inc(x_28); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_14); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_14); lean_dec(x_28); if (x_29 == 0) { @@ -65845,7 +66442,7 @@ return x_30; lean_object* x_33; uint8_t x_34; x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -65859,7 +66456,7 @@ lean_inc(x_1); x_35 = l_Lean_Parser_Term_ident___elambda__1(x_1, x_32); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); lean_dec(x_36); if (x_37 == 0) { @@ -65875,7 +66472,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_categoryParser___elambda__1(x_38, x_39, x_1, x_35); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); lean_dec(x_41); if (x_42 == 0) { @@ -66030,7 +66627,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binrel__no__prop_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(236u); +x_1 = lean_unsigned_to_nat(243u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -66042,7 +66639,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binrel__no__prop_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(236u); +x_1 = lean_unsigned_to_nat(243u); x_2 = lean_unsigned_to_nat(140u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -66070,7 +66667,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binrel__no__prop_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(236u); +x_1 = lean_unsigned_to_nat(243u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -66082,7 +66679,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binrel__no__prop_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(236u); +x_1 = lean_unsigned_to_nat(243u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -66468,7 +67065,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -66514,7 +67111,7 @@ x_26 = l_Lean_Parser_Term_binop___elambda__1___closed__2; x_27 = l_Lean_Parser_ParserState_mkNode(x_25, x_26, x_17); x_28 = lean_ctor_get(x_27, 4); lean_inc(x_28); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_14); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_14); lean_dec(x_28); if (x_29 == 0) { @@ -66534,7 +67131,7 @@ return x_30; lean_object* x_33; uint8_t x_34; x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -66548,7 +67145,7 @@ lean_inc(x_1); x_35 = l_Lean_Parser_Term_ident___elambda__1(x_1, x_32); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); lean_dec(x_36); if (x_37 == 0) { @@ -66564,7 +67161,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_categoryParser___elambda__1(x_38, x_39, x_1, x_35); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); lean_dec(x_41); if (x_42 == 0) { @@ -66701,7 +67298,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binop_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(237u); +x_1 = lean_unsigned_to_nat(244u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -66713,7 +67310,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binop_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(237u); +x_1 = lean_unsigned_to_nat(244u); x_2 = lean_unsigned_to_nat(123u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -66741,7 +67338,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binop_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(237u); +x_1 = lean_unsigned_to_nat(244u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -66753,7 +67350,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binop_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(237u); +x_1 = lean_unsigned_to_nat(244u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -67139,7 +67736,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -67185,7 +67782,7 @@ x_26 = l_Lean_Parser_Term_binop__lazy___elambda__1___closed__2; x_27 = l_Lean_Parser_ParserState_mkNode(x_25, x_26, x_17); x_28 = lean_ctor_get(x_27, 4); lean_inc(x_28); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_14); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_14); lean_dec(x_28); if (x_29 == 0) { @@ -67205,7 +67802,7 @@ return x_30; lean_object* x_33; uint8_t x_34; x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -67219,7 +67816,7 @@ lean_inc(x_1); x_35 = l_Lean_Parser_Term_ident___elambda__1(x_1, x_32); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); lean_dec(x_36); if (x_37 == 0) { @@ -67235,7 +67832,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_categoryParser___elambda__1(x_38, x_39, x_1, x_35); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); lean_dec(x_41); if (x_42 == 0) { @@ -67372,7 +67969,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binop__lazy_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(238u); +x_1 = lean_unsigned_to_nat(245u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -67384,7 +67981,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binop__lazy_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(238u); +x_1 = lean_unsigned_to_nat(245u); x_2 = lean_unsigned_to_nat(133u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -67412,7 +68009,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binop__lazy_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(238u); +x_1 = lean_unsigned_to_nat(245u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -67424,7 +68021,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binop__lazy_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(238u); +x_1 = lean_unsigned_to_nat(245u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -67822,7 +68419,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -67866,7 +68463,7 @@ goto block_57; lean_object* x_26; uint8_t x_27; x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); lean_dec(x_26); if (x_27 == 0) { @@ -67875,7 +68472,7 @@ x_28 = l_Lean_Parser_Term_forInMacro___elambda__1___closed__2; x_29 = l_Lean_Parser_ParserState_mkNode(x_25, x_28, x_17); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); lean_dec(x_30); if (x_31 == 0) { @@ -67899,7 +68496,7 @@ lean_inc(x_1); x_35 = l_Lean_Parser_categoryParser___elambda__1(x_33, x_34, x_1, x_25); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); lean_dec(x_36); if (x_37 == 0) { @@ -67908,7 +68505,7 @@ x_38 = l_Lean_Parser_Term_forInMacro___elambda__1___closed__2; x_39 = l_Lean_Parser_ParserState_mkNode(x_35, x_38, x_17); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); lean_dec(x_40); if (x_41 == 0) { @@ -67930,7 +68527,7 @@ lean_inc(x_1); x_43 = l_Lean_Parser_categoryParser___elambda__1(x_33, x_34, x_1, x_35); x_44 = lean_ctor_get(x_43, 4); lean_inc(x_44); -x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_14); +x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_14); lean_dec(x_44); if (x_45 == 0) { @@ -67939,7 +68536,7 @@ x_46 = l_Lean_Parser_Term_forInMacro___elambda__1___closed__2; x_47 = l_Lean_Parser_ParserState_mkNode(x_43, x_46, x_17); x_48 = lean_ctor_get(x_47, 4); lean_inc(x_48); -x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_14); +x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_14); lean_dec(x_48); if (x_49 == 0) { @@ -67963,7 +68560,7 @@ x_52 = l_Lean_Parser_Term_forInMacro___elambda__1___closed__2; x_53 = l_Lean_Parser_ParserState_mkNode(x_51, x_52, x_17); x_54 = lean_ctor_get(x_53, 4); lean_inc(x_54); -x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_14); +x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_14); lean_dec(x_54); if (x_55 == 0) { @@ -68112,7 +68709,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_forInMacro_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(240u); +x_1 = lean_unsigned_to_nat(247u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -68124,7 +68721,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_forInMacro_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(240u); +x_1 = lean_unsigned_to_nat(247u); x_2 = lean_unsigned_to_nat(130u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -68152,7 +68749,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_forInMacro_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(240u); +x_1 = lean_unsigned_to_nat(247u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -68164,7 +68761,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_forInMacro_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(240u); +x_1 = lean_unsigned_to_nat(247u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -68574,7 +69171,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -68618,7 +69215,7 @@ goto block_57; lean_object* x_26; uint8_t x_27; x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); lean_dec(x_26); if (x_27 == 0) { @@ -68627,7 +69224,7 @@ x_28 = l_Lean_Parser_Term_forInMacro_x27___elambda__1___closed__2; x_29 = l_Lean_Parser_ParserState_mkNode(x_25, x_28, x_17); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); lean_dec(x_30); if (x_31 == 0) { @@ -68651,7 +69248,7 @@ lean_inc(x_1); x_35 = l_Lean_Parser_categoryParser___elambda__1(x_33, x_34, x_1, x_25); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); lean_dec(x_36); if (x_37 == 0) { @@ -68660,7 +69257,7 @@ x_38 = l_Lean_Parser_Term_forInMacro_x27___elambda__1___closed__2; x_39 = l_Lean_Parser_ParserState_mkNode(x_35, x_38, x_17); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); lean_dec(x_40); if (x_41 == 0) { @@ -68682,7 +69279,7 @@ lean_inc(x_1); x_43 = l_Lean_Parser_categoryParser___elambda__1(x_33, x_34, x_1, x_35); x_44 = lean_ctor_get(x_43, 4); lean_inc(x_44); -x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_14); +x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_14); lean_dec(x_44); if (x_45 == 0) { @@ -68691,7 +69288,7 @@ x_46 = l_Lean_Parser_Term_forInMacro_x27___elambda__1___closed__2; x_47 = l_Lean_Parser_ParserState_mkNode(x_43, x_46, x_17); x_48 = lean_ctor_get(x_47, 4); lean_inc(x_48); -x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_14); +x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_14); lean_dec(x_48); if (x_49 == 0) { @@ -68715,7 +69312,7 @@ x_52 = l_Lean_Parser_Term_forInMacro_x27___elambda__1___closed__2; x_53 = l_Lean_Parser_ParserState_mkNode(x_51, x_52, x_17); x_54 = lean_ctor_get(x_53, 4); lean_inc(x_54); -x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_14); +x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_14); lean_dec(x_54); if (x_55 == 0) { @@ -68852,7 +69449,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_forInMacro_x27_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(241u); +x_1 = lean_unsigned_to_nat(248u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -68864,7 +69461,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_forInMacro_x27_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(241u); +x_1 = lean_unsigned_to_nat(248u); x_2 = lean_unsigned_to_nat(132u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -68892,7 +69489,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_forInMacro_x27_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(241u); +x_1 = lean_unsigned_to_nat(248u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -68904,7 +69501,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_forInMacro_x27_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(241u); +x_1 = lean_unsigned_to_nat(248u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -69290,7 +69887,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -69321,7 +69918,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_20, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -69330,7 +69927,7 @@ x_27 = l_Lean_Parser_Term_typeOf___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_20, x_27, x_17); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -69356,7 +69953,7 @@ x_35 = l_Lean_Parser_Term_typeOf___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_17); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); lean_dec(x_37); if (x_38 == 0) { @@ -69379,7 +69976,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_20); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); lean_dec(x_41); if (x_42 == 0) { @@ -69388,7 +69985,7 @@ x_43 = l_Lean_Parser_Term_typeOf___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_17); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); lean_dec(x_45); if (x_46 == 0) { @@ -69414,7 +70011,7 @@ x_51 = l_Lean_Parser_Term_typeOf___elambda__1___closed__2; x_52 = l_Lean_Parser_ParserState_mkNode(x_50, x_51, x_17); x_53 = lean_ctor_get(x_52, 4); lean_inc(x_53); -x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); +x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); lean_dec(x_53); if (x_54 == 0) { @@ -69551,7 +70148,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_typeOf_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(243u); +x_1 = lean_unsigned_to_nat(250u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -69563,7 +70160,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_typeOf_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(243u); +x_1 = lean_unsigned_to_nat(250u); x_2 = lean_unsigned_to_nat(95u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -69591,7 +70188,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_typeOf_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(243u); +x_1 = lean_unsigned_to_nat(250u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -69603,7 +70200,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_typeOf_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(243u); +x_1 = lean_unsigned_to_nat(250u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -70021,7 +70618,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -70065,7 +70662,7 @@ goto block_58; lean_object* x_26; uint8_t x_27; x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); lean_dec(x_26); if (x_27 == 0) { @@ -70074,7 +70671,7 @@ x_28 = l_Lean_Parser_Term_ensureTypeOf___elambda__1___closed__2; x_29 = l_Lean_Parser_ParserState_mkNode(x_25, x_28, x_17); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); lean_dec(x_30); if (x_31 == 0) { @@ -70098,7 +70695,7 @@ lean_inc(x_1); x_35 = l_Lean_Parser_categoryParser___elambda__1(x_33, x_34, x_1, x_25); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); lean_dec(x_36); if (x_37 == 0) { @@ -70107,7 +70704,7 @@ x_38 = l_Lean_Parser_Term_ensureTypeOf___elambda__1___closed__2; x_39 = l_Lean_Parser_ParserState_mkNode(x_35, x_38, x_17); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); lean_dec(x_40); if (x_41 == 0) { @@ -70129,7 +70726,7 @@ lean_inc(x_1); x_43 = l_Lean_Parser_strLit___elambda__1(x_1, x_35); x_44 = lean_ctor_get(x_43, 4); lean_inc(x_44); -x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_14); +x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_14); lean_dec(x_44); if (x_45 == 0) { @@ -70138,7 +70735,7 @@ x_46 = l_Lean_Parser_Term_ensureTypeOf___elambda__1___closed__2; x_47 = l_Lean_Parser_ParserState_mkNode(x_43, x_46, x_17); x_48 = lean_ctor_get(x_47, 4); lean_inc(x_48); -x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_14); +x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_14); lean_dec(x_48); if (x_49 == 0) { @@ -70163,7 +70760,7 @@ x_53 = l_Lean_Parser_Term_ensureTypeOf___elambda__1___closed__2; x_54 = l_Lean_Parser_ParserState_mkNode(x_52, x_53, x_17); x_55 = lean_ctor_get(x_54, 4); lean_inc(x_55); -x_56 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_55, x_14); +x_56 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_55, x_14); lean_dec(x_55); if (x_56 == 0) { @@ -70326,7 +70923,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_ensureTypeOf_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(244u); +x_1 = lean_unsigned_to_nat(251u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -70338,7 +70935,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_ensureTypeOf_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(244u); +x_1 = lean_unsigned_to_nat(251u); x_2 = lean_unsigned_to_nat(126u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -70366,7 +70963,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_ensureTypeOf_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(244u); +x_1 = lean_unsigned_to_nat(251u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -70378,7 +70975,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_ensureTypeOf_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(244u); +x_1 = lean_unsigned_to_nat(251u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -70824,7 +71421,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -70868,7 +71465,7 @@ goto block_49; lean_object* x_26; uint8_t x_27; x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); lean_dec(x_26); if (x_27 == 0) { @@ -70877,7 +71474,7 @@ x_28 = l_Lean_Parser_Term_ensureExpectedType___elambda__1___closed__2; x_29 = l_Lean_Parser_ParserState_mkNode(x_25, x_28, x_17); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); lean_dec(x_30); if (x_31 == 0) { @@ -70899,7 +71496,7 @@ lean_inc(x_1); x_33 = l_Lean_Parser_strLit___elambda__1(x_1, x_25); x_34 = lean_ctor_get(x_33, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_14); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_14); lean_dec(x_34); if (x_35 == 0) { @@ -70908,7 +71505,7 @@ x_36 = l_Lean_Parser_Term_ensureExpectedType___elambda__1___closed__2; x_37 = l_Lean_Parser_ParserState_mkNode(x_33, x_36, x_17); x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_14); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_14); lean_dec(x_38); if (x_39 == 0) { @@ -70934,7 +71531,7 @@ x_44 = l_Lean_Parser_Term_ensureExpectedType___elambda__1___closed__2; x_45 = l_Lean_Parser_ParserState_mkNode(x_43, x_44, x_17); x_46 = lean_ctor_get(x_45, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_14); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_14); lean_dec(x_46); if (x_47 == 0) { @@ -71084,7 +71681,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_ensureExpectedType_dec _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(245u); +x_1 = lean_unsigned_to_nat(252u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -71096,7 +71693,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_ensureExpectedType_dec _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(245u); +x_1 = lean_unsigned_to_nat(252u); x_2 = lean_unsigned_to_nat(118u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -71124,7 +71721,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_ensureExpectedType_dec _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(245u); +x_1 = lean_unsigned_to_nat(252u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -71136,7 +71733,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_ensureExpectedType_dec _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(245u); +x_1 = lean_unsigned_to_nat(252u); x_2 = lean_unsigned_to_nat(43u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -71534,7 +72131,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -71565,7 +72162,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_20, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -71574,7 +72171,7 @@ x_27 = l_Lean_Parser_Term_noImplicitLambda___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_20, x_27, x_17); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -71600,7 +72197,7 @@ x_35 = l_Lean_Parser_Term_noImplicitLambda___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_17); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); lean_dec(x_37); if (x_38 == 0) { @@ -71623,7 +72220,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_20); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); lean_dec(x_41); if (x_42 == 0) { @@ -71632,7 +72229,7 @@ x_43 = l_Lean_Parser_Term_noImplicitLambda___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_17); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); lean_dec(x_45); if (x_46 == 0) { @@ -71658,7 +72255,7 @@ x_51 = l_Lean_Parser_Term_noImplicitLambda___elambda__1___closed__2; x_52 = l_Lean_Parser_ParserState_mkNode(x_50, x_51, x_17); x_53 = lean_ctor_get(x_52, 4); lean_inc(x_53); -x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); +x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); lean_dec(x_53); if (x_54 == 0) { @@ -71795,7 +72392,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_noImplicitLambda_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(246u); +x_1 = lean_unsigned_to_nat(253u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -71807,7 +72404,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_noImplicitLambda_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(246u); +x_1 = lean_unsigned_to_nat(253u); x_2 = lean_unsigned_to_nat(106u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -71835,7 +72432,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_noImplicitLambda_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(246u); +x_1 = lean_unsigned_to_nat(253u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -71847,7 +72444,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_noImplicitLambda_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(246u); +x_1 = lean_unsigned_to_nat(253u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -72313,7 +72910,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -72358,7 +72955,7 @@ x_26 = l_Lean_Parser_Term_letMVar___elambda__1___closed__2; x_27 = l_Lean_Parser_ParserState_mkNode(x_25, x_26, x_17); x_28 = lean_ctor_get(x_27, 4); lean_inc(x_28); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_14); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_14); lean_dec(x_28); if (x_29 == 0) { @@ -72378,7 +72975,7 @@ return x_30; lean_object* x_33; uint8_t x_34; x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -72388,7 +72985,7 @@ x_35 = l_Lean_Parser_Term_letMVar___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_32, x_35, x_17); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); lean_dec(x_37); if (x_38 == 0) { @@ -72433,7 +73030,7 @@ goto block_91; lean_object* x_47; uint8_t x_48; x_47 = lean_ctor_get(x_46, 4); lean_inc(x_47); -x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_14); +x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_14); lean_dec(x_47); if (x_48 == 0) { @@ -72443,7 +73040,7 @@ x_49 = l_Lean_Parser_Term_letMVar___elambda__1___closed__2; x_50 = l_Lean_Parser_ParserState_mkNode(x_46, x_49, x_17); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_14); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_14); lean_dec(x_51); if (x_52 == 0) { @@ -72465,7 +73062,7 @@ lean_inc(x_1); x_54 = l_Lean_Parser_Term_ident___elambda__1(x_1, x_46); x_55 = lean_ctor_get(x_54, 4); lean_inc(x_55); -x_56 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_55, x_14); +x_56 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_55, x_14); lean_dec(x_55); if (x_56 == 0) { @@ -72475,7 +73072,7 @@ x_57 = l_Lean_Parser_Term_letMVar___elambda__1___closed__2; x_58 = l_Lean_Parser_ParserState_mkNode(x_54, x_57, x_17); x_59 = lean_ctor_get(x_58, 4); lean_inc(x_59); -x_60 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_59, x_14); +x_60 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_59, x_14); lean_dec(x_59); if (x_60 == 0) { @@ -72520,7 +73117,7 @@ goto block_89; lean_object* x_69; uint8_t x_70; x_69 = lean_ctor_get(x_68, 4); lean_inc(x_69); -x_70 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_69, x_14); +x_70 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_69, x_14); lean_dec(x_69); if (x_70 == 0) { @@ -72537,7 +73134,7 @@ lean_inc(x_1); x_73 = l_Lean_Parser_categoryParser___elambda__1(x_71, x_72, x_1, x_68); x_74 = lean_ctor_get(x_73, 4); lean_inc(x_74); -x_75 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_74, x_14); +x_75 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_74, x_14); lean_dec(x_74); if (x_75 == 0) { @@ -72563,7 +73160,7 @@ if (x_81 == 0) lean_object* x_82; uint8_t x_83; x_82 = lean_ctor_get(x_78, 4); lean_inc(x_82); -x_83 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_82, x_14); +x_83 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_82, x_14); lean_dec(x_82); if (x_83 == 0) { @@ -72586,7 +73183,7 @@ lean_inc(x_1); x_85 = l_Lean_Parser_tokenAntiquotFn(x_1, x_78); x_86 = lean_ctor_get(x_85, 4); lean_inc(x_86); -x_87 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_86, x_14); +x_87 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_86, x_14); lean_dec(x_86); if (x_87 == 0) { @@ -72785,7 +73382,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_letMVar_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(248u); +x_1 = lean_unsigned_to_nat(255u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -72797,7 +73394,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_letMVar_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(248u); +x_1 = lean_unsigned_to_nat(255u); x_2 = lean_unsigned_to_nat(140u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -72825,7 +73422,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_letMVar_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(248u); +x_1 = lean_unsigned_to_nat(255u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -72837,7 +73434,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_letMVar_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(248u); +x_1 = lean_unsigned_to_nat(255u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -73367,7 +73964,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -73412,7 +74009,7 @@ x_26 = l_Lean_Parser_Term_waitIfTypeMVar___elambda__1___closed__2; x_27 = l_Lean_Parser_ParserState_mkNode(x_25, x_26, x_17); x_28 = lean_ctor_get(x_27, 4); lean_inc(x_28); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_14); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_14); lean_dec(x_28); if (x_29 == 0) { @@ -73432,7 +74029,7 @@ return x_30; lean_object* x_33; uint8_t x_34; x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -73442,7 +74039,7 @@ x_35 = l_Lean_Parser_Term_waitIfTypeMVar___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_32, x_35, x_17); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); lean_dec(x_37); if (x_38 == 0) { @@ -73487,7 +74084,7 @@ goto block_69; lean_object* x_47; uint8_t x_48; x_47 = lean_ctor_get(x_46, 4); lean_inc(x_47); -x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_14); +x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_14); lean_dec(x_47); if (x_48 == 0) { @@ -73502,7 +74099,7 @@ lean_inc(x_1); x_49 = l_Lean_Parser_Term_ident___elambda__1(x_1, x_46); x_50 = lean_ctor_get(x_49, 4); lean_inc(x_50); -x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_14); +x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_14); lean_dec(x_50); if (x_51 == 0) { @@ -73528,7 +74125,7 @@ if (x_57 == 0) lean_object* x_58; uint8_t x_59; x_58 = lean_ctor_get(x_54, 4); lean_inc(x_58); -x_59 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_58, x_14); +x_59 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_58, x_14); lean_dec(x_58); if (x_59 == 0) { @@ -73553,7 +74150,7 @@ lean_inc(x_1); x_63 = l_Lean_Parser_tokenAntiquotFn(x_1, x_54); x_64 = lean_ctor_get(x_63, 4); lean_inc(x_64); -x_65 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_64, x_14); +x_65 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_64, x_14); lean_dec(x_64); if (x_65 == 0) { @@ -73729,7 +74326,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_waitIfTypeMVar_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(249u); +x_1 = lean_unsigned_to_nat(256u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -73741,7 +74338,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_waitIfTypeMVar_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(249u); +x_1 = lean_unsigned_to_nat(256u); x_2 = lean_unsigned_to_nat(125u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -73769,7 +74366,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_waitIfTypeMVar_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(249u); +x_1 = lean_unsigned_to_nat(256u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -73781,7 +74378,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_waitIfTypeMVar_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(249u); +x_1 = lean_unsigned_to_nat(256u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -74227,7 +74824,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -74272,7 +74869,7 @@ x_26 = l_Lean_Parser_Term_waitIfTypeContainsMVar___elambda__1___closed__2; x_27 = l_Lean_Parser_ParserState_mkNode(x_25, x_26, x_17); x_28 = lean_ctor_get(x_27, 4); lean_inc(x_28); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_14); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_14); lean_dec(x_28); if (x_29 == 0) { @@ -74292,7 +74889,7 @@ return x_30; lean_object* x_33; uint8_t x_34; x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -74302,7 +74899,7 @@ x_35 = l_Lean_Parser_Term_waitIfTypeContainsMVar___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_32, x_35, x_17); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); lean_dec(x_37); if (x_38 == 0) { @@ -74347,7 +74944,7 @@ goto block_69; lean_object* x_47; uint8_t x_48; x_47 = lean_ctor_get(x_46, 4); lean_inc(x_47); -x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_14); +x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_14); lean_dec(x_47); if (x_48 == 0) { @@ -74362,7 +74959,7 @@ lean_inc(x_1); x_49 = l_Lean_Parser_Term_ident___elambda__1(x_1, x_46); x_50 = lean_ctor_get(x_49, 4); lean_inc(x_50); -x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_14); +x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_14); lean_dec(x_50); if (x_51 == 0) { @@ -74388,7 +74985,7 @@ if (x_57 == 0) lean_object* x_58; uint8_t x_59; x_58 = lean_ctor_get(x_54, 4); lean_inc(x_58); -x_59 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_58, x_14); +x_59 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_58, x_14); lean_dec(x_58); if (x_59 == 0) { @@ -74413,7 +75010,7 @@ lean_inc(x_1); x_63 = l_Lean_Parser_tokenAntiquotFn(x_1, x_54); x_64 = lean_ctor_get(x_63, 4); lean_inc(x_64); -x_65 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_64, x_14); +x_65 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_64, x_14); lean_dec(x_64); if (x_65 == 0) { @@ -74555,7 +75152,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_waitIfTypeContainsMVar _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(250u); +x_1 = lean_unsigned_to_nat(257u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -74567,7 +75164,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_waitIfTypeContainsMVar _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(250u); +x_1 = lean_unsigned_to_nat(257u); x_2 = lean_unsigned_to_nat(134u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -74595,7 +75192,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_waitIfTypeContainsMVar _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(250u); +x_1 = lean_unsigned_to_nat(257u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -74607,7 +75204,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_waitIfTypeContainsMVar _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(250u); +x_1 = lean_unsigned_to_nat(257u); x_2 = lean_unsigned_to_nat(47u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -74993,7 +75590,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -75038,7 +75635,7 @@ x_26 = l_Lean_Parser_Term_waitIfContainsMVar___elambda__1___closed__2; x_27 = l_Lean_Parser_ParserState_mkNode(x_25, x_26, x_17); x_28 = lean_ctor_get(x_27, 4); lean_inc(x_28); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_14); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_14); lean_dec(x_28); if (x_29 == 0) { @@ -75058,7 +75655,7 @@ return x_30; lean_object* x_33; uint8_t x_34; x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -75068,7 +75665,7 @@ x_35 = l_Lean_Parser_Term_waitIfContainsMVar___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_32, x_35, x_17); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); lean_dec(x_37); if (x_38 == 0) { @@ -75113,7 +75710,7 @@ goto block_69; lean_object* x_47; uint8_t x_48; x_47 = lean_ctor_get(x_46, 4); lean_inc(x_47); -x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_14); +x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_14); lean_dec(x_47); if (x_48 == 0) { @@ -75128,7 +75725,7 @@ lean_inc(x_1); x_49 = l_Lean_Parser_Term_ident___elambda__1(x_1, x_46); x_50 = lean_ctor_get(x_49, 4); lean_inc(x_50); -x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_14); +x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_14); lean_dec(x_50); if (x_51 == 0) { @@ -75154,7 +75751,7 @@ if (x_57 == 0) lean_object* x_58; uint8_t x_59; x_58 = lean_ctor_get(x_54, 4); lean_inc(x_58); -x_59 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_58, x_14); +x_59 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_58, x_14); lean_dec(x_58); if (x_59 == 0) { @@ -75179,7 +75776,7 @@ lean_inc(x_1); x_63 = l_Lean_Parser_tokenAntiquotFn(x_1, x_54); x_64 = lean_ctor_get(x_63, 4); lean_inc(x_64); -x_65 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_64, x_14); +x_65 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_64, x_14); lean_dec(x_64); if (x_65 == 0) { @@ -75321,7 +75918,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_waitIfContainsMVar_dec _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(251u); +x_1 = lean_unsigned_to_nat(258u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -75333,7 +75930,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_waitIfContainsMVar_dec _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(251u); +x_1 = lean_unsigned_to_nat(258u); x_2 = lean_unsigned_to_nat(129u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -75361,7 +75958,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_waitIfContainsMVar_dec _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(251u); +x_1 = lean_unsigned_to_nat(258u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -75373,7 +75970,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_waitIfContainsMVar_dec _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(251u); +x_1 = lean_unsigned_to_nat(258u); x_2 = lean_unsigned_to_nat(43u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -75821,7 +76418,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -75853,7 +76450,7 @@ if (x_26 == 0) lean_object* x_27; uint8_t x_28; x_27 = lean_ctor_get(x_22, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_16); lean_dec(x_27); if (x_28 == 0) { @@ -75863,7 +76460,7 @@ x_29 = l_Lean_Parser_Term_defaultOrOfNonempty___elambda__1___closed__2; x_30 = l_Lean_Parser_ParserState_mkNode(x_22, x_29, x_19); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_16); lean_dec(x_31); if (x_32 == 0) { @@ -75887,7 +76484,7 @@ x_35 = l_Lean_Parser_Term_defaultOrOfNonempty___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_19); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_16); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_16); lean_dec(x_37); if (x_38 == 0) { @@ -75910,7 +76507,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_22); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_16); lean_dec(x_41); if (x_42 == 0) { @@ -75920,7 +76517,7 @@ x_43 = l_Lean_Parser_Term_defaultOrOfNonempty___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_19); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_16); lean_dec(x_45); if (x_46 == 0) { @@ -75944,7 +76541,7 @@ x_49 = l_Lean_Parser_Term_defaultOrOfNonempty___elambda__1___closed__2; x_50 = l_Lean_Parser_ParserState_mkNode(x_48, x_49, x_19); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_16); lean_dec(x_51); if (x_52 == 0) { @@ -76082,7 +76679,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_defaultOrOfNonempty_de _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(253u); +x_1 = lean_unsigned_to_nat(260u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -76094,7 +76691,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_defaultOrOfNonempty_de _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(253u); +x_1 = lean_unsigned_to_nat(260u); x_2 = lean_unsigned_to_nat(111u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -76122,7 +76719,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_defaultOrOfNonempty_de _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(253u); +x_1 = lean_unsigned_to_nat(260u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -76134,7 +76731,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_defaultOrOfNonempty_de _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(253u); +x_1 = lean_unsigned_to_nat(260u); x_2 = lean_unsigned_to_nat(44u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -76546,7 +77143,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -76566,7 +77163,7 @@ lean_inc(x_1); x_19 = l_Lean_Parser_atomicFn(x_18, x_1, x_12); x_20 = lean_ctor_get(x_19, 4); lean_inc(x_20); -x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_20, x_14); +x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_20, x_14); lean_dec(x_20); if (x_21 == 0) { @@ -76576,7 +77173,7 @@ x_22 = l_Lean_Parser_Term_namedArgument___elambda__1___closed__2; x_23 = l_Lean_Parser_ParserState_mkNode(x_19, x_22, x_17); x_24 = lean_ctor_get(x_23, 4); lean_inc(x_24); -x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_14); +x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_14); lean_dec(x_24); if (x_25 == 0) { @@ -76600,7 +77197,7 @@ lean_inc(x_1); x_29 = l_Lean_Parser_categoryParser___elambda__1(x_27, x_28, x_1, x_19); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); lean_dec(x_30); if (x_31 == 0) { @@ -76610,7 +77207,7 @@ x_32 = l_Lean_Parser_Term_namedArgument___elambda__1___closed__2; x_33 = l_Lean_Parser_ParserState_mkNode(x_29, x_32, x_17); x_34 = lean_ctor_get(x_33, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_14); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_14); lean_dec(x_34); if (x_35 == 0) { @@ -76646,7 +77243,7 @@ x_44 = l_Lean_Parser_Term_namedArgument___elambda__1___closed__2; x_45 = l_Lean_Parser_ParserState_mkNode(x_39, x_44, x_17); x_46 = lean_ctor_get(x_45, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_14); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_14); lean_dec(x_46); if (x_47 == 0) { @@ -76670,7 +77267,7 @@ x_50 = l_Lean_Parser_Term_namedArgument___elambda__1___closed__2; x_51 = l_Lean_Parser_ParserState_mkNode(x_49, x_50, x_17); x_52 = lean_ctor_get(x_51, 4); lean_inc(x_52); -x_53 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_52, x_14); +x_53 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_52, x_14); lean_dec(x_52); if (x_53 == 0) { @@ -76914,7 +77511,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -76947,7 +77544,7 @@ x_25 = l_Lean_Parser_Term_ellipsis___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_20, x_25, x_17); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); lean_dec(x_27); if (x_28 == 0) { @@ -76971,7 +77568,7 @@ x_31 = l_Lean_Parser_Term_ellipsis___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_17); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -77118,7 +77715,7 @@ x_4 = l_Lean_Parser_checkWsBeforeFn(x_3, x_1, x_2); x_5 = lean_ctor_get(x_4, 4); lean_inc(x_5); x_6 = lean_box(0); -x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); +x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); lean_dec(x_5); if (x_7 == 0) { @@ -77173,7 +77770,7 @@ x_22 = l_Lean_Parser_Term_argument___elambda__1___closed__4; x_23 = l_Lean_Parser_ParserState_mkError(x_4, x_22); x_24 = lean_ctor_get(x_23, 4); lean_inc(x_24); -x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_6); +x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_6); lean_dec(x_24); if (x_25 == 0) { @@ -77326,7 +77923,7 @@ x_6 = l_Lean_Parser_checkPrecFn(x_5, x_1, x_2); x_7 = lean_ctor_get(x_6, 4); lean_inc(x_7); x_8 = lean_box(0); -x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); +x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); lean_dec(x_7); if (x_9 == 0) { @@ -77341,7 +77938,7 @@ x_10 = l_Lean_Parser_maxPrec; x_11 = l_Lean_Parser_checkLhsPrecFn(x_10, x_1, x_6); x_12 = lean_ctor_get(x_11, 4); lean_inc(x_12); -x_13 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_12, x_8); +x_13 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_12, x_8); lean_dec(x_12); if (x_13 == 0) { @@ -77363,7 +77960,7 @@ x_18 = l_Lean_Parser_ParserState_mkTrailingNode(x_16, x_17, x_15); lean_dec(x_15); x_19 = lean_ctor_get(x_18, 4); lean_inc(x_19); -x_20 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_8); +x_20 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_19, x_8); lean_dec(x_19); if (x_20 == 0) { @@ -77468,7 +78065,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_app_declRange___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(264u); +x_1 = lean_unsigned_to_nat(271u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -77480,7 +78077,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_app_declRange___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(264u); +x_1 = lean_unsigned_to_nat(271u); x_2 = lean_unsigned_to_nat(84u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -77508,7 +78105,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_app_declRange___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(264u); +x_1 = lean_unsigned_to_nat(271u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -77520,7 +78117,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_app_declRange___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(264u); +x_1 = lean_unsigned_to_nat(271u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -78219,7 +78816,7 @@ x_4 = l_Lean_Parser_checkPrecFn(x_3, x_1, x_2); x_5 = lean_ctor_get(x_4, 4); lean_inc(x_5); x_6 = lean_box(0); -x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); +x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); lean_dec(x_5); if (x_7 == 0) { @@ -78233,7 +78830,7 @@ x_8 = lean_unsigned_to_nat(0u); x_9 = l_Lean_Parser_checkLhsPrecFn(x_8, x_1, x_4); x_10 = lean_ctor_get(x_9, 4); lean_inc(x_10); -x_11 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_10, x_6); +x_11 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_10, x_6); lean_dec(x_10); if (x_11 == 0) { @@ -78251,7 +78848,7 @@ x_14 = l_Lean_Parser_Term_doubleQuotedName___elambda__1___closed__9; x_15 = l_Lean_Parser_checkNoWsBeforeFn(x_14, x_1, x_9); x_16 = lean_ctor_get(x_15, 4); lean_inc(x_16); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_16, x_6); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_16, x_6); lean_dec(x_16); if (x_17 == 0) { @@ -78282,7 +78879,7 @@ if (x_33 == 0) lean_object* x_34; uint8_t x_35; x_34 = lean_ctor_get(x_27, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_6); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_6); lean_dec(x_34); if (x_35 == 0) { @@ -78295,7 +78892,7 @@ lean_object* x_36; lean_object* x_37; uint8_t x_38; x_36 = l_Lean_Parser_checkNoWsBeforeFn(x_14, x_1, x_27); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_6); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_6); lean_dec(x_37); if (x_38 == 0) { @@ -78322,7 +78919,7 @@ lean_inc(x_1); x_43 = l_Lean_Parser_tokenAntiquotFn(x_1, x_27); x_44 = lean_ctor_get(x_43, 4); lean_inc(x_44); -x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_6); +x_45 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_44, x_6); lean_dec(x_44); if (x_45 == 0) { @@ -78335,7 +78932,7 @@ lean_object* x_46; lean_object* x_47; uint8_t x_48; x_46 = l_Lean_Parser_checkNoWsBeforeFn(x_14, x_1, x_43); x_47 = lean_ctor_get(x_46, 4); lean_inc(x_47); -x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_6); +x_48 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_47, x_6); lean_dec(x_47); if (x_48 == 0) { @@ -78364,7 +78961,7 @@ x_20 = l_Lean_Parser_ParserState_mkTrailingNode(x_18, x_19, x_13); lean_dec(x_13); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); -x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_6); +x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_6); lean_dec(x_21); if (x_22 == 0) { @@ -78510,7 +79107,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_proj_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(266u); +x_1 = lean_unsigned_to_nat(273u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -78522,7 +79119,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_proj_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(266u); +x_1 = lean_unsigned_to_nat(273u); x_2 = lean_unsigned_to_nat(121u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -78550,7 +79147,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_proj_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(266u); +x_1 = lean_unsigned_to_nat(273u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -78562,7 +79159,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_proj_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(266u); +x_1 = lean_unsigned_to_nat(273u); x_2 = lean_unsigned_to_nat(29u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -78831,7 +79428,7 @@ x_4 = l_Lean_Parser_checkPrecFn(x_3, x_1, x_2); x_5 = lean_ctor_get(x_4, 4); lean_inc(x_5); x_6 = lean_box(0); -x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); +x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); lean_dec(x_5); if (x_7 == 0) { @@ -78845,7 +79442,7 @@ x_8 = lean_unsigned_to_nat(0u); x_9 = l_Lean_Parser_checkLhsPrecFn(x_8, x_1, x_4); x_10 = lean_ctor_get(x_9, 4); lean_inc(x_10); -x_11 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_10, x_6); +x_11 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_10, x_6); lean_dec(x_10); if (x_11 == 0) { @@ -78863,7 +79460,7 @@ x_14 = l_Lean_Parser_Term_doubleQuotedName___elambda__1___closed__9; x_15 = l_Lean_Parser_checkNoWsBeforeFn(x_14, x_1, x_9); x_16 = lean_ctor_get(x_15, 4); lean_inc(x_16); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_16, x_6); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_16, x_6); lean_dec(x_16); if (x_17 == 0) { @@ -78873,7 +79470,7 @@ x_19 = l_Lean_Parser_ParserState_mkTrailingNode(x_15, x_18, x_13); lean_dec(x_13); x_20 = lean_ctor_get(x_19, 4); lean_inc(x_20); -x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_20, x_6); +x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_20, x_6); lean_dec(x_20); if (x_21 == 0) { @@ -78915,7 +79512,7 @@ x_33 = l_Lean_Parser_ParserState_mkTrailingNode(x_25, x_32, x_13); lean_dec(x_13); x_34 = lean_ctor_get(x_33, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_6); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_6); lean_dec(x_34); if (x_35 == 0) { @@ -78940,7 +79537,7 @@ x_39 = l_Lean_Parser_ParserState_mkTrailingNode(x_37, x_38, x_13); lean_dec(x_13); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_6); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_6); lean_dec(x_40); if (x_41 == 0) { @@ -79055,7 +79652,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_completion_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(267u); +x_1 = lean_unsigned_to_nat(274u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -79067,7 +79664,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_completion_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(267u); +x_1 = lean_unsigned_to_nat(274u); x_2 = lean_unsigned_to_nat(77u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -79095,7 +79692,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_completion_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(267u); +x_1 = lean_unsigned_to_nat(274u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -79107,7 +79704,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_completion_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(267u); +x_1 = lean_unsigned_to_nat(274u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -79288,7 +79885,7 @@ x_4 = l_Lean_Parser_checkPrecFn(x_3, x_1, x_2); x_5 = lean_ctor_get(x_4, 4); lean_inc(x_5); x_6 = lean_box(0); -x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); +x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); lean_dec(x_5); if (x_7 == 0) { @@ -79302,7 +79899,7 @@ x_8 = lean_unsigned_to_nat(0u); x_9 = l_Lean_Parser_checkLhsPrecFn(x_8, x_1, x_4); x_10 = lean_ctor_get(x_9, 4); lean_inc(x_10); -x_11 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_10, x_6); +x_11 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_10, x_6); lean_dec(x_10); if (x_11 == 0) { @@ -79320,7 +79917,7 @@ x_14 = l_Lean_Parser_Term_doubleQuotedName___elambda__1___closed__9; x_15 = l_Lean_Parser_checkNoWsBeforeFn(x_14, x_1, x_9); x_16 = lean_ctor_get(x_15, 4); lean_inc(x_16); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_16, x_6); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_16, x_6); lean_dec(x_16); if (x_17 == 0) { @@ -79330,7 +79927,7 @@ x_19 = l_Lean_Parser_ParserState_mkTrailingNode(x_15, x_18, x_13); lean_dec(x_13); x_20 = lean_ctor_get(x_19, 4); lean_inc(x_20); -x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_20, x_6); +x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_20, x_6); lean_dec(x_20); if (x_21 == 0) { @@ -79381,7 +79978,7 @@ goto block_66; lean_object* x_33; uint8_t x_34; x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_6); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_6); lean_dec(x_33); if (x_34 == 0) { @@ -79392,7 +79989,7 @@ x_36 = l_Lean_Parser_ParserState_mkTrailingNode(x_32, x_35, x_13); lean_dec(x_13); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_6); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_6); lean_dec(x_37); if (x_38 == 0) { @@ -79415,7 +80012,7 @@ lean_inc(x_1); x_41 = l_Lean_Parser_categoryParser___elambda__1(x_40, x_8, x_1, x_32); x_42 = lean_ctor_get(x_41, 4); lean_inc(x_42); -x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_6); +x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_6); lean_dec(x_42); if (x_43 == 0) { @@ -79426,7 +80023,7 @@ x_45 = l_Lean_Parser_ParserState_mkTrailingNode(x_41, x_44, x_13); lean_dec(x_13); x_46 = lean_ctor_get(x_45, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_6); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_6); lean_dec(x_46); if (x_47 == 0) { @@ -79462,7 +80059,7 @@ x_56 = l_Lean_Parser_ParserState_mkTrailingNode(x_51, x_55, x_13); lean_dec(x_13); x_57 = lean_ctor_get(x_56, 4); lean_inc(x_57); -x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_6); +x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_6); lean_dec(x_57); if (x_58 == 0) { @@ -79487,7 +80084,7 @@ x_62 = l_Lean_Parser_ParserState_mkTrailingNode(x_60, x_61, x_13); lean_dec(x_13); x_63 = lean_ctor_get(x_62, 4); lean_inc(x_63); -x_64 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_63, x_6); +x_64 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_63, x_6); lean_dec(x_63); if (x_64 == 0) { @@ -79605,7 +80202,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_arrayRef_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(268u); +x_1 = lean_unsigned_to_nat(275u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -79617,7 +80214,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_arrayRef_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(268u); +x_1 = lean_unsigned_to_nat(275u); x_2 = lean_unsigned_to_nat(95u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -79645,7 +80242,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_arrayRef_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(268u); +x_1 = lean_unsigned_to_nat(275u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -79657,7 +80254,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_arrayRef_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(268u); +x_1 = lean_unsigned_to_nat(275u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -79838,7 +80435,7 @@ x_4 = l_Lean_Parser_checkPrecFn(x_3, x_1, x_2); x_5 = lean_ctor_get(x_4, 4); lean_inc(x_5); x_6 = lean_box(0); -x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); +x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); lean_dec(x_5); if (x_7 == 0) { @@ -79852,7 +80449,7 @@ x_8 = lean_unsigned_to_nat(0u); x_9 = l_Lean_Parser_checkLhsPrecFn(x_8, x_1, x_4); x_10 = lean_ctor_get(x_9, 4); lean_inc(x_10); -x_11 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_10, x_6); +x_11 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_10, x_6); lean_dec(x_10); if (x_11 == 0) { @@ -79870,7 +80467,7 @@ x_14 = lean_unsigned_to_nat(25u); x_15 = l_Lean_Parser_checkPrecFn(x_14, x_1, x_9); x_16 = lean_ctor_get(x_15, 4); lean_inc(x_16); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_16, x_6); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_16, x_6); lean_dec(x_16); if (x_17 == 0) { @@ -79902,7 +80499,7 @@ if (x_34 == 0) lean_object* x_35; uint8_t x_36; x_35 = lean_ctor_get(x_28, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_6); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_6); lean_dec(x_35); if (x_36 == 0) { @@ -79926,7 +80523,7 @@ lean_inc(x_1); x_39 = l_Lean_Parser_tokenAntiquotFn(x_1, x_28); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_6); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_6); lean_dec(x_40); if (x_41 == 0) { @@ -79952,7 +80549,7 @@ x_20 = l_Lean_Parser_ParserState_mkTrailingNode(x_18, x_19, x_13); lean_dec(x_13); x_21 = lean_ctor_get(x_20, 4); lean_inc(x_21); -x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_6); +x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_6); lean_dec(x_21); if (x_22 == 0) { @@ -80088,7 +80685,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_arrow_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(269u); +x_1 = lean_unsigned_to_nat(276u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -80100,7 +80697,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_arrow_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(269u); +x_1 = lean_unsigned_to_nat(276u); x_2 = lean_unsigned_to_nat(112u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -80128,7 +80725,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_arrow_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(269u); +x_1 = lean_unsigned_to_nat(276u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -80140,7 +80737,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_arrow_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(269u); +x_1 = lean_unsigned_to_nat(276u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -80460,7 +81057,7 @@ x_6 = l_Lean_Parser_checkPrecFn(x_5, x_1, x_2); x_7 = lean_ctor_get(x_6, 4); lean_inc(x_7); x_8 = lean_box(0); -x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); +x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); lean_dec(x_7); if (x_9 == 0) { @@ -80475,7 +81072,7 @@ x_10 = lean_unsigned_to_nat(0u); x_11 = l_Lean_Parser_checkLhsPrecFn(x_10, x_1, x_6); x_12 = lean_ctor_get(x_11, 4); lean_inc(x_12); -x_13 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_12, x_8); +x_13 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_12, x_8); lean_dec(x_12); if (x_13 == 0) { @@ -80513,7 +81110,7 @@ goto block_80; lean_object* x_20; uint8_t x_21; x_20 = lean_ctor_get(x_19, 4); lean_inc(x_20); -x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_20, x_8); +x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_20, x_8); lean_dec(x_20); if (x_21 == 0) { @@ -80524,7 +81121,7 @@ x_23 = l_Lean_Parser_ParserState_mkTrailingNode(x_19, x_22, x_15); lean_dec(x_15); x_24 = lean_ctor_get(x_23, 4); lean_inc(x_24); -x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_8); +x_25 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_24, x_8); lean_dec(x_24); if (x_25 == 0) { @@ -80546,7 +81143,7 @@ x_27 = l_Lean_Parser_Term_explicitUniv___elambda__1___closed__7; x_28 = l_Lean_Parser_checkNoWsBeforeFn(x_27, x_1, x_19); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_8); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_8); lean_dec(x_29); if (x_30 == 0) { @@ -80557,7 +81154,7 @@ x_32 = l_Lean_Parser_ParserState_mkTrailingNode(x_28, x_31, x_15); lean_dec(x_15); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_8); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_8); lean_dec(x_33); if (x_34 == 0) { @@ -80608,7 +81205,7 @@ goto block_78; lean_object* x_46; uint8_t x_47; x_46 = lean_ctor_get(x_45, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_8); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_8); lean_dec(x_46); if (x_47 == 0) { @@ -80620,7 +81217,7 @@ x_49 = l_Lean_Parser_ParserState_mkTrailingNode(x_45, x_48, x_15); lean_dec(x_15); x_50 = lean_ctor_get(x_49, 4); lean_inc(x_50); -x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_8); +x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_8); lean_dec(x_50); if (x_51 == 0) { @@ -80642,7 +81239,7 @@ lean_inc(x_1); x_53 = lean_apply_2(x_4, x_1, x_45); x_54 = lean_ctor_get(x_53, 4); lean_inc(x_54); -x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_8); +x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_8); lean_dec(x_54); if (x_55 == 0) { @@ -80653,7 +81250,7 @@ x_57 = l_Lean_Parser_ParserState_mkTrailingNode(x_53, x_56, x_15); lean_dec(x_15); x_58 = lean_ctor_get(x_57, 4); lean_inc(x_58); -x_59 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_58, x_8); +x_59 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_58, x_8); lean_dec(x_58); if (x_59 == 0) { @@ -80689,7 +81286,7 @@ x_68 = l_Lean_Parser_ParserState_mkTrailingNode(x_63, x_67, x_15); lean_dec(x_15); x_69 = lean_ctor_get(x_68, 4); lean_inc(x_69); -x_70 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_69, x_8); +x_70 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_69, x_8); lean_dec(x_69); if (x_70 == 0) { @@ -80714,7 +81311,7 @@ x_74 = l_Lean_Parser_ParserState_mkTrailingNode(x_72, x_73, x_15); lean_dec(x_15); x_75 = lean_ctor_get(x_74, 4); lean_inc(x_75); -x_76 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_75, x_8); +x_76 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_75, x_8); lean_dec(x_75); if (x_76 == 0) { @@ -80875,7 +81472,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_explicitUniv_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(275u); +x_1 = lean_unsigned_to_nat(282u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -80887,7 +81484,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_explicitUniv_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(275u); +x_1 = lean_unsigned_to_nat(282u); x_2 = lean_unsigned_to_nat(211u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -80915,7 +81512,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_explicitUniv_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(275u); +x_1 = lean_unsigned_to_nat(282u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -80927,7 +81524,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_explicitUniv_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(275u); +x_1 = lean_unsigned_to_nat(282u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -81326,7 +81923,7 @@ x_6 = l_Lean_Parser_checkPrecFn(x_5, x_1, x_2); x_7 = lean_ctor_get(x_6, 4); lean_inc(x_7); x_8 = lean_box(0); -x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); +x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); lean_dec(x_7); if (x_9 == 0) { @@ -81341,7 +81938,7 @@ x_10 = lean_unsigned_to_nat(0u); x_11 = l_Lean_Parser_checkLhsPrecFn(x_10, x_1, x_6); x_12 = lean_ctor_get(x_11, 4); lean_inc(x_12); -x_13 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_12, x_8); +x_13 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_12, x_8); lean_dec(x_12); if (x_13 == 0) { @@ -81382,7 +81979,7 @@ x_21 = l_Lean_Parser_ParserState_mkTrailingNode(x_19, x_20, x_15); lean_dec(x_15); x_22 = lean_ctor_get(x_21, 4); lean_inc(x_22); -x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_22, x_8); +x_23 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_22, x_8); lean_dec(x_22); if (x_23 == 0) { @@ -81402,7 +81999,7 @@ return x_24; lean_object* x_27; uint8_t x_28; x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_8); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_8); lean_dec(x_27); if (x_28 == 0) { @@ -81417,7 +82014,7 @@ x_29 = l_Lean_Parser_Term_namedPattern___elambda__1___closed__8; x_30 = l_Lean_Parser_checkNoWsBeforeFn(x_29, x_1, x_26); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_8); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_8); lean_dec(x_31); if (x_32 == 0) { @@ -81449,7 +82046,7 @@ if (x_41 == 0) lean_object* x_42; uint8_t x_43; x_42 = lean_ctor_get(x_35, 4); lean_inc(x_42); -x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_8); +x_43 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_42, x_8); lean_dec(x_42); if (x_43 == 0) { @@ -81464,7 +82061,7 @@ lean_inc(x_1); x_44 = lean_apply_2(x_4, x_1, x_35); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_8); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_8); lean_dec(x_45); if (x_46 == 0) { @@ -81490,7 +82087,7 @@ lean_inc(x_1); x_50 = l_Lean_Parser_tokenAntiquotFn(x_1, x_35); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_8); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_8); lean_dec(x_51); if (x_52 == 0) { @@ -81505,7 +82102,7 @@ lean_inc(x_1); x_53 = lean_apply_2(x_4, x_1, x_50); x_54 = lean_ctor_get(x_53, 4); lean_inc(x_54); -x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_8); +x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_8); lean_dec(x_54); if (x_55 == 0) { @@ -81660,7 +82257,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_namedPattern_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(276u); +x_1 = lean_unsigned_to_nat(283u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -81672,7 +82269,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_namedPattern_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(276u); +x_1 = lean_unsigned_to_nat(283u); x_2 = lean_unsigned_to_nat(233u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -81700,7 +82297,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_namedPattern_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(276u); +x_1 = lean_unsigned_to_nat(283u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -81712,7 +82309,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_namedPattern_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(276u); +x_1 = lean_unsigned_to_nat(283u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -82078,7 +82675,7 @@ x_6 = l_Lean_Parser_checkPrecFn(x_5, x_1, x_2); x_7 = lean_ctor_get(x_6, 4); lean_inc(x_7); x_8 = lean_box(0); -x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); +x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); lean_dec(x_7); if (x_9 == 0) { @@ -82093,7 +82690,7 @@ x_10 = lean_unsigned_to_nat(0u); x_11 = l_Lean_Parser_checkLhsPrecFn(x_10, x_1, x_6); x_12 = lean_ctor_get(x_11, 4); lean_inc(x_12); -x_13 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_12, x_8); +x_13 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_12, x_8); lean_dec(x_12); if (x_13 == 0) { @@ -82142,7 +82739,7 @@ goto block_59; lean_object* x_26; uint8_t x_27; x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_8); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_8); lean_dec(x_26); if (x_27 == 0) { @@ -82153,7 +82750,7 @@ x_29 = l_Lean_Parser_ParserState_mkTrailingNode(x_25, x_28, x_15); lean_dec(x_15); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_8); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_8); lean_dec(x_30); if (x_31 == 0) { @@ -82175,7 +82772,7 @@ x_33 = l_Lean_Parser_Term_doubleQuotedName___elambda__1___closed__9; x_34 = l_Lean_Parser_checkNoWsBeforeFn(x_33, x_1, x_25); x_35 = lean_ctor_get(x_34, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_8); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_8); lean_dec(x_35); if (x_36 == 0) { @@ -82186,7 +82783,7 @@ x_38 = l_Lean_Parser_ParserState_mkTrailingNode(x_34, x_37, x_15); lean_dec(x_15); x_39 = lean_ctor_get(x_38, 4); lean_inc(x_39); -x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_8); +x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_8); lean_dec(x_39); if (x_40 == 0) { @@ -82211,7 +82808,7 @@ lean_inc(x_1); x_45 = l_Lean_Parser_orelseFnCore(x_42, x_43, x_44, x_1, x_34); x_46 = lean_ctor_get(x_45, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_8); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_8); lean_dec(x_46); if (x_47 == 0) { @@ -82222,7 +82819,7 @@ x_49 = l_Lean_Parser_ParserState_mkTrailingNode(x_45, x_48, x_15); lean_dec(x_15); x_50 = lean_ctor_get(x_49, 4); lean_inc(x_50); -x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_8); +x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_8); lean_dec(x_50); if (x_51 == 0) { @@ -82247,7 +82844,7 @@ x_55 = l_Lean_Parser_ParserState_mkTrailingNode(x_53, x_54, x_15); lean_dec(x_15); x_56 = lean_ctor_get(x_55, 4); lean_inc(x_56); -x_57 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_56, x_8); +x_57 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_56, x_8); lean_dec(x_56); if (x_57 == 0) { @@ -82395,7 +82992,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_pipeProj_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(278u); +x_1 = lean_unsigned_to_nat(285u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -82407,7 +83004,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_pipeProj_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(278u); +x_1 = lean_unsigned_to_nat(285u); x_2 = lean_unsigned_to_nat(132u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -82435,7 +83032,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_pipeProj_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(278u); +x_1 = lean_unsigned_to_nat(285u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -82447,7 +83044,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_pipeProj_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(278u); +x_1 = lean_unsigned_to_nat(285u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -82716,7 +83313,7 @@ x_4 = l_Lean_Parser_checkPrecFn(x_3, x_1, x_2); x_5 = lean_ctor_get(x_4, 4); lean_inc(x_5); x_6 = lean_box(0); -x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); +x_7 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_5, x_6); lean_dec(x_5); if (x_7 == 0) { @@ -82730,7 +83327,7 @@ x_8 = lean_unsigned_to_nat(0u); x_9 = l_Lean_Parser_checkLhsPrecFn(x_8, x_1, x_4); x_10 = lean_ctor_get(x_9, 4); lean_inc(x_10); -x_11 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_10, x_6); +x_11 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_10, x_6); lean_dec(x_10); if (x_11 == 0) { @@ -82768,7 +83365,7 @@ x_24 = l_Lean_Parser_ParserState_mkTrailingNode(x_16, x_23, x_13); lean_dec(x_13); x_25 = lean_ctor_get(x_24, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_6); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_6); lean_dec(x_25); if (x_26 == 0) { @@ -82793,7 +83390,7 @@ x_30 = l_Lean_Parser_ParserState_mkTrailingNode(x_28, x_29, x_13); lean_dec(x_13); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); -x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_6); +x_32 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_6); lean_dec(x_31); if (x_32 == 0) { @@ -82897,7 +83494,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_pipeCompletion_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(279u); +x_1 = lean_unsigned_to_nat(286u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -82909,7 +83506,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_pipeCompletion_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(279u); +x_1 = lean_unsigned_to_nat(286u); x_2 = lean_unsigned_to_nat(73u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -82937,7 +83534,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_pipeCompletion_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(279u); +x_1 = lean_unsigned_to_nat(286u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -82949,7 +83546,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_pipeCompletion_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(279u); +x_1 = lean_unsigned_to_nat(286u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -83199,7 +83796,7 @@ x_6 = l_Lean_Parser_checkPrecFn(x_5, x_1, x_2); x_7 = lean_ctor_get(x_6, 4); lean_inc(x_7); x_8 = lean_box(0); -x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); +x_9 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_7, x_8); lean_dec(x_7); if (x_9 == 0) { @@ -83214,7 +83811,7 @@ x_10 = lean_unsigned_to_nat(0u); x_11 = l_Lean_Parser_checkLhsPrecFn(x_10, x_1, x_6); x_12 = lean_ctor_get(x_11, 4); lean_inc(x_12); -x_13 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_12, x_8); +x_13 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_12, x_8); lean_dec(x_12); if (x_13 == 0) { @@ -83250,7 +83847,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_18, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_8); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_8); lean_dec(x_25); if (x_26 == 0) { @@ -83261,7 +83858,7 @@ x_28 = l_Lean_Parser_ParserState_mkTrailingNode(x_18, x_27, x_15); lean_dec(x_15); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_8); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_8); lean_dec(x_29); if (x_30 == 0) { @@ -83286,7 +83883,7 @@ x_34 = l_Lean_Parser_ParserState_mkTrailingNode(x_32, x_33, x_15); lean_dec(x_15); x_35 = lean_ctor_get(x_34, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_8); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_8); lean_dec(x_35); if (x_36 == 0) { @@ -83309,7 +83906,7 @@ lean_inc(x_1); x_38 = l_Lean_Parser_tokenAntiquotFn(x_1, x_18); x_39 = lean_ctor_get(x_38, 4); lean_inc(x_39); -x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_8); +x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_8); lean_dec(x_39); if (x_40 == 0) { @@ -83320,7 +83917,7 @@ x_42 = l_Lean_Parser_ParserState_mkTrailingNode(x_38, x_41, x_15); lean_dec(x_15); x_43 = lean_ctor_get(x_42, 4); lean_inc(x_43); -x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_8); +x_44 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_43, x_8); lean_dec(x_43); if (x_44 == 0) { @@ -83345,7 +83942,7 @@ x_48 = l_Lean_Parser_ParserState_mkTrailingNode(x_46, x_47, x_15); lean_dec(x_15); x_49 = lean_ctor_get(x_48, 4); lean_inc(x_49); -x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_8); +x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_8); lean_dec(x_49); if (x_50 == 0) { @@ -83462,7 +84059,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_subst_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(281u); +x_1 = lean_unsigned_to_nat(288u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -83474,7 +84071,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_subst_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(281u); +x_1 = lean_unsigned_to_nat(288u); x_2 = lean_unsigned_to_nat(90u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -83502,7 +84099,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_subst_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(281u); +x_1 = lean_unsigned_to_nat(288u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -83514,7 +84111,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_subst_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(281u); +x_1 = lean_unsigned_to_nat(288u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -83736,7 +84333,7 @@ static lean_object* _init_l_Lean_Parser_Term_funBinder_quot___elambda__1___close _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("funBinder", 9); +x_1 = lean_mk_string_from_bytes("quot", 4); return x_1; } } @@ -83744,7 +84341,7 @@ static lean_object* _init_l_Lean_Parser_Term_funBinder_quot___elambda__1___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_byTactic___elambda__1___closed__2; +x_1 = l_Lean_Parser_Term_funBinder___elambda__1___closed__2; x_2 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -83753,34 +84350,16 @@ return x_3; static lean_object* _init_l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__3() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("quot", 4); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__2; -x_2 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__3; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__5() { -_start: -{ lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__3; -x_2 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__4; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__1; +x_2 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__2; x_3 = 1; x_4 = 0; x_5 = l_Lean_Parser_mkAntiquot(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__4() { _start: { lean_object* x_1; @@ -83788,26 +84367,26 @@ x_1 = lean_mk_string_from_bytes("`(funBinder|", 12); return x_1; } } -static lean_object* _init_l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__7() { +static lean_object* _init_l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__6; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__4; x_2 = l_String_trim(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__8() { +static lean_object* _init_l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__7; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__5; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_Command_docComment___elambda__1___lambda__1___boxed), 3, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__9() { +static lean_object* _init_l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__7() { _start: { lean_object* x_1; lean_object* x_2; @@ -83816,11 +84395,11 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__10() { +static lean_object* _init_l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__2; +x_1 = l_Lean_Parser_Term_funBinder___elambda__1___closed__2; x_2 = l_Lean_Parser_Term_funBinder; x_3 = lean_alloc_closure((void*)(l_Lean_Parser_evalInsideQuot___elambda__1), 4, 2); lean_closure_set(x_3, 0, x_1); @@ -83828,23 +84407,23 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__11() { +static lean_object* _init_l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__9; -x_2 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__10; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__7; +x_2 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__8; x_3 = lean_alloc_closure((void*)(l___private_Lean_Parser_Basic_0__Lean_Parser_addQuotDepthFn___boxed), 4, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__12() { +static lean_object* _init_l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__11; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__9; x_2 = l_Lean_Parser_Term_paren___elambda__1___closed__13; x_3 = lean_alloc_closure((void*)(l_Lean_Parser_andthenFn), 4, 2); lean_closure_set(x_3, 0, x_1); @@ -83852,35 +84431,35 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__13() { +static lean_object* _init_l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__8; -x_2 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__12; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__6; +x_2 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__10; x_3 = lean_alloc_closure((void*)(l_Lean_Parser_andthenFn), 4, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__14() { +static lean_object* _init_l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__4; -x_2 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__13; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__2; +x_2 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__11; x_3 = lean_alloc_closure((void*)(l_Lean_Parser_nodeFn), 4, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__15() { +static lean_object* _init_l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__14; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__12; x_2 = l_Lean_Parser_Command_docComment___elambda__1___closed__16; x_3 = lean_alloc_closure((void*)(l_Lean_Parser_andthenFn), 4, 2); lean_closure_set(x_3, 0, x_1); @@ -83888,33 +84467,33 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__16() { +static lean_object* _init_l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_docComment___elambda__1___closed__15; -x_2 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__15; +x_2 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__13; x_3 = lean_alloc_closure((void*)(l_Lean_Parser_andthenFn), 4, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__17() { +static lean_object* _init_l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_docComment___elambda__1___lambda__1___closed__1; -x_2 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__7; +x_2 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__5; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__18() { +static lean_object* _init_l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__17; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__15; x_2 = l_Lean_Parser_Command_docComment___elambda__1___lambda__1___closed__1; x_3 = lean_string_append(x_1, x_2); return x_3; @@ -83924,7 +84503,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_funBinder_quot___elambda__1(lean_obj _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint32_t x_8; uint32_t x_9; uint8_t x_10; -x_3 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__5; +x_3 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__3; x_4 = lean_ctor_get(x_3, 1); lean_inc(x_4); x_5 = lean_ctor_get(x_1, 0); @@ -83947,7 +84526,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -83962,8 +84541,8 @@ x_16 = lean_ctor_get(x_12, 0); lean_inc(x_16); x_17 = lean_array_get_size(x_16); lean_dec(x_16); -x_18 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__7; -x_19 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__18; +x_18 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__5; +x_19 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__16; lean_inc(x_1); x_20 = l_Lean_Parser_symbolFnAux(x_18, x_19, x_1, x_12); x_21 = lean_ctor_get(x_20, 2); @@ -83990,17 +84569,17 @@ goto block_60; lean_object* x_26; uint8_t x_27; x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); lean_dec(x_26); if (x_27 == 0) { lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_dec(x_6); -x_28 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__4; +x_28 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__2; x_29 = l_Lean_Parser_ParserState_mkNode(x_25, x_28, x_17); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); lean_dec(x_30); if (x_31 == 0) { @@ -84018,23 +84597,23 @@ return x_32; else { lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_33 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__9; -x_34 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__10; +x_33 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__7; +x_34 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__8; lean_inc(x_1); x_35 = l___private_Lean_Parser_Basic_0__Lean_Parser_addQuotDepthFn(x_33, x_34, x_1, x_25); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); lean_dec(x_36); if (x_37 == 0) { lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; lean_dec(x_6); -x_38 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__4; +x_38 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__2; x_39 = l_Lean_Parser_ParserState_mkNode(x_35, x_38, x_17); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); lean_dec(x_40); if (x_41 == 0) { @@ -84065,11 +84644,11 @@ x_48 = lean_uint32_dec_eq(x_47, x_23); if (x_48 == 0) { lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; -x_49 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__4; +x_49 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__2; x_50 = l_Lean_Parser_ParserState_mkNode(x_45, x_49, x_17); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_14); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_14); lean_dec(x_51); if (x_52 == 0) { @@ -84089,11 +84668,11 @@ else lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; uint8_t x_58; lean_inc(x_1); x_54 = l_Lean_Parser_tokenAntiquotFn(x_1, x_45); -x_55 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__4; +x_55 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__2; x_56 = l_Lean_Parser_ParserState_mkNode(x_54, x_55, x_17); x_57 = lean_ctor_get(x_56, 4); lean_inc(x_57); -x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_14); +x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_14); lean_dec(x_57); if (x_58 == 0) { @@ -84117,7 +84696,7 @@ else { lean_object* x_62; uint8_t x_63; lean_object* x_64; lean_dec(x_6); -x_62 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__16; +x_62 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__14; x_63 = 0; x_64 = l_Lean_Parser_orelseFnCore(x_4, x_62, x_63, x_1, x_2); return x_64; @@ -84128,7 +84707,7 @@ static lean_object* _init_l_Lean_Parser_Term_funBinder_quot___closed__1() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__7; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__5; x_2 = l_Lean_Parser_symbolInfo(x_1); return x_2; } @@ -84159,7 +84738,7 @@ static lean_object* _init_l_Lean_Parser_Term_funBinder_quot___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__4; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__2; x_2 = l_Lean_Parser_Term_funBinder_quot___closed__3; x_3 = l_Lean_Parser_nodeInfo(x_1, x_2); return x_3; @@ -84189,7 +84768,7 @@ static lean_object* _init_l_Lean_Parser_Term_funBinder_quot___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__5; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__3; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l_Lean_Parser_Term_funBinder_quot___closed__6; @@ -84230,7 +84809,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_funBinder_quot(lean_obj { lean_object* x_2; lean_object* x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; x_2 = l___regBuiltin_Lean_Parser_Term_byTactic___closed__2; -x_3 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__4; +x_3 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__2; x_4 = 1; x_5 = l_Lean_Parser_Term_funBinder_quot; x_6 = lean_unsigned_to_nat(1000u); @@ -84242,7 +84821,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_funBinder_quot_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(284u); +x_1 = lean_unsigned_to_nat(291u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -84254,7 +84833,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_funBinder_quot_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(284u); +x_1 = lean_unsigned_to_nat(291u); x_2 = lean_unsigned_to_nat(144u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -84282,7 +84861,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_funBinder_quot_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(284u); +x_1 = lean_unsigned_to_nat(291u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -84294,7 +84873,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_funBinder_quot_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(284u); +x_1 = lean_unsigned_to_nat(291u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -84334,7 +84913,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_funBinder_quot_declRang _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__4; +x_2 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__2; x_3 = l___regBuiltin_Lean_Parser_Term_funBinder_quot_declRange___closed__7; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; @@ -84344,8 +84923,8 @@ static lean_object* _init_l_Lean_Parser_Term_funBinder_quot_formatter___closed__ _start: { lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__3; -x_2 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__4; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__1; +x_2 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__2; x_3 = 1; x_4 = 0; x_5 = lean_box(x_3); @@ -84362,7 +84941,7 @@ static lean_object* _init_l_Lean_Parser_Term_funBinder_quot_formatter___closed__ _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__6; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__4; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -84416,7 +84995,7 @@ static lean_object* _init_l_Lean_Parser_Term_funBinder_quot_formatter___closed__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__4; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__2; x_2 = lean_unsigned_to_nat(1024u); x_3 = l_Lean_Parser_Term_funBinder_quot_formatter___closed__6; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); @@ -84440,7 +85019,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_funBinder_quot_formatt _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__4; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__2; x_2 = l___regBuiltin_Lean_Parser_Tactic_tacticSeqBracketed_formatter___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -84459,7 +85038,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_funBinder_quot_formatte { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_2 = l___regBuiltin_Lean_Parser_Tactic_tacticSeqBracketed_formatter___closed__3; -x_3 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__4; +x_3 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__2; x_4 = l___regBuiltin_Lean_Parser_Term_funBinder_quot_formatter___closed__1; x_5 = l___regBuiltin_Lean_Parser_Term_funBinder_quot_formatter___closed__2; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); @@ -84470,8 +85049,8 @@ static lean_object* _init_l_Lean_Parser_Term_funBinder_quot_parenthesizer___clos _start: { lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__3; -x_2 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__4; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__1; +x_2 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__2; x_3 = 1; x_4 = 0; x_5 = lean_box(x_3); @@ -84488,7 +85067,7 @@ static lean_object* _init_l_Lean_Parser_Term_funBinder_quot_parenthesizer___clos _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__6; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__4; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -84542,7 +85121,7 @@ static lean_object* _init_l_Lean_Parser_Term_funBinder_quot_parenthesizer___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__4; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__2; x_2 = lean_unsigned_to_nat(1024u); x_3 = l_Lean_Parser_Term_funBinder_quot_parenthesizer___closed__6; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer___boxed), 8, 3); @@ -84566,7 +85145,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_funBinder_quot_parenth _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__4; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__2; x_2 = l___regBuiltin_Lean_Parser_Tactic_tacticSeqBracketed_parenthesizer___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; @@ -84585,7 +85164,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_funBinder_quot_parenthe { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_2 = l___regBuiltin_Lean_Parser_Tactic_tacticSeqBracketed_parenthesizer___closed__3; -x_3 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__4; +x_3 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__2; x_4 = l___regBuiltin_Lean_Parser_Term_funBinder_quot_parenthesizer___closed__1; x_5 = l___regBuiltin_Lean_Parser_Term_funBinder_quot_parenthesizer___closed__2; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); @@ -84605,7 +85184,7 @@ static lean_object* _init_l_Lean_Parser_Term_bracketedBinder_quot___elambda__1__ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_bracketedBinder___elambda__1___closed__2; -x_2 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__3; +x_2 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } @@ -84614,7 +85193,7 @@ static lean_object* _init_l_Lean_Parser_Term_bracketedBinder_quot___elambda__1__ _start: { lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__3; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__1; x_2 = l_Lean_Parser_Term_bracketedBinder_quot___elambda__1___closed__1; x_3 = 1; x_4 = 0; @@ -84683,7 +85262,7 @@ static lean_object* _init_l_Lean_Parser_Term_bracketedBinder_quot___elambda__1__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__9; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__7; x_2 = l_Lean_Parser_Term_bracketedBinder_quot___elambda__1___closed__8; x_3 = lean_alloc_closure((void*)(l___private_Lean_Parser_Basic_0__Lean_Parser_addQuotDepthFn___boxed), 4, 2); lean_closure_set(x_3, 0, x_1); @@ -84798,7 +85377,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -84841,7 +85420,7 @@ goto block_60; lean_object* x_26; uint8_t x_27; x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); lean_dec(x_26); if (x_27 == 0) { @@ -84851,7 +85430,7 @@ x_28 = l_Lean_Parser_Term_bracketedBinder_quot___elambda__1___closed__1; x_29 = l_Lean_Parser_ParserState_mkNode(x_25, x_28, x_17); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); lean_dec(x_30); if (x_31 == 0) { @@ -84869,13 +85448,13 @@ return x_32; else { lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_33 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__9; +x_33 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__7; x_34 = l_Lean_Parser_Term_bracketedBinder_quot___elambda__1___closed__8; lean_inc(x_1); x_35 = l___private_Lean_Parser_Basic_0__Lean_Parser_addQuotDepthFn(x_33, x_34, x_1, x_25); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); lean_dec(x_36); if (x_37 == 0) { @@ -84885,7 +85464,7 @@ x_38 = l_Lean_Parser_Term_bracketedBinder_quot___elambda__1___closed__1; x_39 = l_Lean_Parser_ParserState_mkNode(x_35, x_38, x_17); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); lean_dec(x_40); if (x_41 == 0) { @@ -84920,7 +85499,7 @@ x_49 = l_Lean_Parser_Term_bracketedBinder_quot___elambda__1___closed__1; x_50 = l_Lean_Parser_ParserState_mkNode(x_45, x_49, x_17); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_14); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_14); lean_dec(x_51); if (x_52 == 0) { @@ -84944,7 +85523,7 @@ x_55 = l_Lean_Parser_Term_bracketedBinder_quot___elambda__1___closed__1; x_56 = l_Lean_Parser_ParserState_mkNode(x_54, x_55, x_17); x_57 = lean_ctor_get(x_56, 4); lean_inc(x_57); -x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_14); +x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_14); lean_dec(x_57); if (x_58 == 0) { @@ -85093,7 +85672,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_bracketedBinder_quot_d _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(286u); +x_1 = lean_unsigned_to_nat(293u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -85105,7 +85684,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_bracketedBinder_quot_d _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(286u); +x_1 = lean_unsigned_to_nat(293u); x_2 = lean_unsigned_to_nat(169u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -85133,7 +85712,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_bracketedBinder_quot_d _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(286u); +x_1 = lean_unsigned_to_nat(293u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -85145,7 +85724,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_bracketedBinder_quot_d _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(286u); +x_1 = lean_unsigned_to_nat(293u); x_2 = lean_unsigned_to_nat(45u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -85195,7 +85774,7 @@ static lean_object* _init_l_Lean_Parser_Term_bracketedBinder_quot_formatter___cl _start: { lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__3; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__1; x_2 = l_Lean_Parser_Term_bracketedBinder_quot___elambda__1___closed__1; x_3 = 1; x_4 = 0; @@ -85321,7 +85900,7 @@ static lean_object* _init_l_Lean_Parser_Term_bracketedBinder_quot_parenthesizer_ _start: { lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__3; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__1; x_2 = l_Lean_Parser_Term_bracketedBinder_quot___elambda__1___closed__1; x_3 = 1; x_4 = 0; @@ -85448,7 +86027,7 @@ static lean_object* _init_l_Lean_Parser_Term_matchDiscr_quot___elambda__1___clos { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_matchDiscr___elambda__1___closed__2; -x_2 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__3; +x_2 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } @@ -85457,7 +86036,7 @@ static lean_object* _init_l_Lean_Parser_Term_matchDiscr_quot___elambda__1___clos _start: { lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__3; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__1; x_2 = l_Lean_Parser_Term_matchDiscr_quot___elambda__1___closed__1; x_3 = 1; x_4 = 0; @@ -85508,7 +86087,7 @@ static lean_object* _init_l_Lean_Parser_Term_matchDiscr_quot___elambda__1___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__9; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__7; x_2 = l_Lean_Parser_Term_matchDiscr_quot___elambda__1___closed__6; x_3 = lean_alloc_closure((void*)(l___private_Lean_Parser_Basic_0__Lean_Parser_addQuotDepthFn___boxed), 4, 2); lean_closure_set(x_3, 0, x_1); @@ -85623,7 +86202,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -85666,7 +86245,7 @@ goto block_60; lean_object* x_26; uint8_t x_27; x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); lean_dec(x_26); if (x_27 == 0) { @@ -85676,7 +86255,7 @@ x_28 = l_Lean_Parser_Term_matchDiscr_quot___elambda__1___closed__1; x_29 = l_Lean_Parser_ParserState_mkNode(x_25, x_28, x_17); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); lean_dec(x_30); if (x_31 == 0) { @@ -85694,13 +86273,13 @@ return x_32; else { lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_33 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__9; +x_33 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__7; x_34 = l_Lean_Parser_Term_matchDiscr_quot___elambda__1___closed__6; lean_inc(x_1); x_35 = l___private_Lean_Parser_Basic_0__Lean_Parser_addQuotDepthFn(x_33, x_34, x_1, x_25); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); lean_dec(x_36); if (x_37 == 0) { @@ -85710,7 +86289,7 @@ x_38 = l_Lean_Parser_Term_matchDiscr_quot___elambda__1___closed__1; x_39 = l_Lean_Parser_ParserState_mkNode(x_35, x_38, x_17); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); lean_dec(x_40); if (x_41 == 0) { @@ -85745,7 +86324,7 @@ x_49 = l_Lean_Parser_Term_matchDiscr_quot___elambda__1___closed__1; x_50 = l_Lean_Parser_ParserState_mkNode(x_45, x_49, x_17); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_14); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_14); lean_dec(x_51); if (x_52 == 0) { @@ -85769,7 +86348,7 @@ x_55 = l_Lean_Parser_Term_matchDiscr_quot___elambda__1___closed__1; x_56 = l_Lean_Parser_ParserState_mkNode(x_54, x_55, x_17); x_57 = lean_ctor_get(x_56, 4); lean_inc(x_57); -x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_14); +x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_14); lean_dec(x_57); if (x_58 == 0) { @@ -85918,7 +86497,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_matchDiscr_quot_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(287u); +x_1 = lean_unsigned_to_nat(294u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -85930,7 +86509,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_matchDiscr_quot_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(287u); +x_1 = lean_unsigned_to_nat(294u); x_2 = lean_unsigned_to_nat(148u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -85958,7 +86537,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_matchDiscr_quot_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(287u); +x_1 = lean_unsigned_to_nat(294u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -85970,7 +86549,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_matchDiscr_quot_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(287u); +x_1 = lean_unsigned_to_nat(294u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -86020,7 +86599,7 @@ static lean_object* _init_l_Lean_Parser_Term_matchDiscr_quot_formatter___closed_ _start: { lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__3; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__1; x_2 = l_Lean_Parser_Term_matchDiscr_quot___elambda__1___closed__1; x_3 = 1; x_4 = 0; @@ -86146,7 +86725,7 @@ static lean_object* _init_l_Lean_Parser_Term_matchDiscr_quot_parenthesizer___clo _start: { lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__3; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__1; x_2 = l_Lean_Parser_Term_matchDiscr_quot___elambda__1___closed__1; x_3 = 1; x_4 = 0; @@ -86283,7 +86862,7 @@ static lean_object* _init_l_Lean_Parser_Term_attr_quot___elambda__1___closed__2( { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_attr_quot___elambda__1___closed__1; -x_2 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__3; +x_2 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } @@ -86292,7 +86871,7 @@ static lean_object* _init_l_Lean_Parser_Term_attr_quot___elambda__1___closed__3( _start: { lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__3; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__1; x_2 = l_Lean_Parser_Term_attr_quot___elambda__1___closed__2; x_3 = 1; x_4 = 0; @@ -86331,7 +86910,7 @@ static lean_object* _init_l_Lean_Parser_Term_attr_quot___elambda__1___closed__7( _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__9; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__7; x_2 = l_Lean_Parser_Term_attrInstance___elambda__1___closed__6; x_3 = lean_alloc_closure((void*)(l___private_Lean_Parser_Basic_0__Lean_Parser_addQuotDepthFn___boxed), 4, 2); lean_closure_set(x_3, 0, x_1); @@ -86446,7 +87025,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -86489,7 +87068,7 @@ goto block_60; lean_object* x_26; uint8_t x_27; x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); lean_dec(x_26); if (x_27 == 0) { @@ -86499,7 +87078,7 @@ x_28 = l_Lean_Parser_Term_attr_quot___elambda__1___closed__2; x_29 = l_Lean_Parser_ParserState_mkNode(x_25, x_28, x_17); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); lean_dec(x_30); if (x_31 == 0) { @@ -86517,13 +87096,13 @@ return x_32; else { lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_33 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__9; +x_33 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__7; x_34 = l_Lean_Parser_Term_attrInstance___elambda__1___closed__6; lean_inc(x_1); x_35 = l___private_Lean_Parser_Basic_0__Lean_Parser_addQuotDepthFn(x_33, x_34, x_1, x_25); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); lean_dec(x_36); if (x_37 == 0) { @@ -86533,7 +87112,7 @@ x_38 = l_Lean_Parser_Term_attr_quot___elambda__1___closed__2; x_39 = l_Lean_Parser_ParserState_mkNode(x_35, x_38, x_17); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); lean_dec(x_40); if (x_41 == 0) { @@ -86568,7 +87147,7 @@ x_49 = l_Lean_Parser_Term_attr_quot___elambda__1___closed__2; x_50 = l_Lean_Parser_ParserState_mkNode(x_45, x_49, x_17); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_14); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_14); lean_dec(x_51); if (x_52 == 0) { @@ -86592,7 +87171,7 @@ x_55 = l_Lean_Parser_Term_attr_quot___elambda__1___closed__2; x_56 = l_Lean_Parser_ParserState_mkNode(x_54, x_55, x_17); x_57 = lean_ctor_get(x_56, 4); lean_inc(x_57); -x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_14); +x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_14); lean_dec(x_57); if (x_58 == 0) { @@ -86741,7 +87320,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_attr_quot_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(288u); +x_1 = lean_unsigned_to_nat(295u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -86753,7 +87332,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_attr_quot_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(288u); +x_1 = lean_unsigned_to_nat(295u); x_2 = lean_unsigned_to_nat(106u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -86781,7 +87360,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_attr_quot_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(288u); +x_1 = lean_unsigned_to_nat(295u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -86793,7 +87372,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_attr_quot_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(288u); +x_1 = lean_unsigned_to_nat(295u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -86843,7 +87422,7 @@ static lean_object* _init_l_Lean_Parser_Term_attr_quot_formatter___closed__1() { _start: { lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__3; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__1; x_2 = l_Lean_Parser_Term_attr_quot___elambda__1___closed__2; x_3 = 1; x_4 = 0; @@ -86959,7 +87538,7 @@ static lean_object* _init_l_Lean_Parser_Term_attr_quot_parenthesizer___closed__1 _start: { lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__3; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__1; x_2 = l_Lean_Parser_Term_attr_quot___elambda__1___closed__2; x_3 = 1; x_4 = 0; @@ -87223,7 +87802,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -87254,7 +87833,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_20, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -87263,7 +87842,7 @@ x_27 = l_Lean_Parser_Term_panic___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_20, x_27, x_17); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -87289,7 +87868,7 @@ x_35 = l_Lean_Parser_Term_panic___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_17); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); lean_dec(x_37); if (x_38 == 0) { @@ -87312,7 +87891,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_20); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); lean_dec(x_41); if (x_42 == 0) { @@ -87321,7 +87900,7 @@ x_43 = l_Lean_Parser_Term_panic___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_17); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); lean_dec(x_45); if (x_46 == 0) { @@ -87347,7 +87926,7 @@ x_51 = l_Lean_Parser_Term_panic___elambda__1___closed__2; x_52 = l_Lean_Parser_ParserState_mkNode(x_50, x_51, x_17); x_53 = lean_ctor_get(x_52, 4); lean_inc(x_53); -x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); +x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); lean_dec(x_53); if (x_54 == 0) { @@ -87484,7 +88063,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_panic_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(290u); +x_1 = lean_unsigned_to_nat(297u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -87496,7 +88075,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_panic_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(290u); +x_1 = lean_unsigned_to_nat(297u); x_2 = lean_unsigned_to_nat(87u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -87524,7 +88103,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_panic_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(290u); +x_1 = lean_unsigned_to_nat(297u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -87536,7 +88115,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_panic_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(290u); +x_1 = lean_unsigned_to_nat(297u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -87910,7 +88489,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -87943,7 +88522,7 @@ x_25 = l_Lean_Parser_Term_unreachable___elambda__1___closed__2; x_26 = l_Lean_Parser_ParserState_mkNode(x_20, x_25, x_17); x_27 = lean_ctor_get(x_26, 4); lean_inc(x_27); -x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); +x_28 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_27, x_14); lean_dec(x_27); if (x_28 == 0) { @@ -87967,7 +88546,7 @@ x_31 = l_Lean_Parser_Term_unreachable___elambda__1___closed__2; x_32 = l_Lean_Parser_ParserState_mkNode(x_30, x_31, x_17); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); -x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); +x_34 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_14); lean_dec(x_33); if (x_34 == 0) { @@ -88091,7 +88670,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_unreachable_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(291u); +x_1 = lean_unsigned_to_nat(298u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -88103,7 +88682,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_unreachable_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(291u); +x_1 = lean_unsigned_to_nat(298u); x_2 = lean_unsigned_to_nat(78u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -88131,7 +88710,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_unreachable_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(291u); +x_1 = lean_unsigned_to_nat(298u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -88143,7 +88722,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_unreachable_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(291u); +x_1 = lean_unsigned_to_nat(298u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -88391,7 +88970,7 @@ if (x_19 == 0) lean_object* x_21; uint8_t x_22; x_21 = lean_ctor_get(x_14, 4); lean_inc(x_21); -x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_20); +x_22 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_21, x_20); lean_dec(x_21); if (x_22 == 0) { @@ -88415,7 +88994,7 @@ lean_inc(x_4); x_25 = l_Lean_Parser_tokenAntiquotFn(x_4, x_14); x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_20); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_20); lean_dec(x_26); if (x_27 == 0) { @@ -88485,7 +89064,7 @@ if (x_48 == 0) lean_object* x_50; uint8_t x_51; x_50 = lean_ctor_get(x_43, 4); lean_inc(x_50); -x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_49); +x_51 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_50, x_49); lean_dec(x_50); if (x_51 == 0) { @@ -88509,7 +89088,7 @@ lean_inc(x_39); x_54 = l_Lean_Parser_tokenAntiquotFn(x_39, x_43); x_55 = lean_ctor_get(x_54, 4); lean_inc(x_55); -x_56 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_55, x_49); +x_56 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_55, x_49); lean_dec(x_55); if (x_56 == 0) { @@ -88672,7 +89251,7 @@ x_32 = l_Lean_Parser_checkPrecFn(x_31, x_1, x_2); x_33 = lean_ctor_get(x_32, 4); lean_inc(x_33); x_34 = lean_box(0); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_34); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_33, x_34); lean_dec(x_33); if (x_35 == 0) { @@ -88723,7 +89302,7 @@ if (x_46 == 0) lean_object* x_60; uint8_t x_61; x_60 = lean_ctor_get(x_42, 4); lean_inc(x_60); -x_61 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_60, x_34); +x_61 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_60, x_34); lean_dec(x_60); if (x_61 == 0) { @@ -88748,7 +89327,7 @@ lean_inc(x_40); x_64 = l_Lean_Parser_tokenAntiquotFn(x_40, x_42); x_65 = lean_ctor_get(x_64, 4); lean_inc(x_65); -x_66 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_65, x_34); +x_66 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_65, x_34); lean_dec(x_65); if (x_66 == 0) { @@ -88771,7 +89350,7 @@ goto block_59; lean_object* x_48; uint8_t x_49; x_48 = lean_ctor_get(x_47, 4); lean_inc(x_48); -x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_34); +x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_34); lean_dec(x_48); if (x_49 == 0) { @@ -88780,7 +89359,7 @@ lean_dec(x_9); x_50 = l_Lean_Parser_ParserState_mkNode(x_47, x_11, x_37); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_34); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_34); lean_dec(x_51); if (x_52 == 0) { @@ -88803,7 +89382,7 @@ x_54 = lean_apply_2(x_9, x_1, x_47); x_55 = l_Lean_Parser_ParserState_mkNode(x_54, x_11, x_37); x_56 = lean_ctor_get(x_55, 4); lean_inc(x_56); -x_57 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_56, x_34); +x_57 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_56, x_34); lean_dec(x_56); if (x_57 == 0) { @@ -88980,7 +89559,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_dbgTrace_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(292u); +x_1 = lean_unsigned_to_nat(299u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -88992,7 +89571,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_dbgTrace_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(292u); +x_1 = lean_unsigned_to_nat(299u); x_2 = lean_unsigned_to_nat(166u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -89020,7 +89599,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_dbgTrace_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(292u); +x_1 = lean_unsigned_to_nat(299u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -89032,7 +89611,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_dbgTrace_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(292u); +x_1 = lean_unsigned_to_nat(299u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -89392,7 +89971,7 @@ if (x_18 == 0) lean_object* x_20; uint8_t x_21; x_20 = lean_ctor_get(x_13, 4); lean_inc(x_20); -x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_20, x_19); +x_21 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_20, x_19); lean_dec(x_20); if (x_21 == 0) { @@ -89415,7 +89994,7 @@ lean_inc(x_3); x_24 = l_Lean_Parser_tokenAntiquotFn(x_3, x_13); x_25 = lean_ctor_get(x_24, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_19); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_19); lean_dec(x_25); if (x_26 == 0) { @@ -89484,7 +90063,7 @@ if (x_47 == 0) lean_object* x_49; uint8_t x_50; x_49 = lean_ctor_get(x_42, 4); lean_inc(x_49); -x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_48); +x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_48); lean_dec(x_49); if (x_50 == 0) { @@ -89507,7 +90086,7 @@ lean_inc(x_38); x_53 = l_Lean_Parser_tokenAntiquotFn(x_38, x_42); x_54 = lean_ctor_get(x_53, 4); lean_inc(x_54); -x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_48); +x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_48); lean_dec(x_54); if (x_55 == 0) { @@ -89655,7 +90234,7 @@ x_30 = l_Lean_Parser_checkPrecFn(x_29, x_1, x_2); x_31 = lean_ctor_get(x_30, 4); lean_inc(x_31); x_32 = lean_box(0); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_32); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_31, x_32); lean_dec(x_31); if (x_33 == 0) { @@ -89705,7 +90284,7 @@ if (x_44 == 0) lean_object* x_58; uint8_t x_59; x_58 = lean_ctor_get(x_40, 4); lean_inc(x_58); -x_59 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_58, x_32); +x_59 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_58, x_32); lean_dec(x_58); if (x_59 == 0) { @@ -89729,7 +90308,7 @@ lean_inc(x_38); x_62 = l_Lean_Parser_tokenAntiquotFn(x_38, x_40); x_63 = lean_ctor_get(x_62, 4); lean_inc(x_63); -x_64 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_63, x_32); +x_64 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_63, x_32); lean_dec(x_63); if (x_64 == 0) { @@ -89751,7 +90330,7 @@ goto block_57; lean_object* x_46; uint8_t x_47; x_46 = lean_ctor_get(x_45, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_32); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_32); lean_dec(x_46); if (x_47 == 0) { @@ -89760,7 +90339,7 @@ lean_dec(x_7); x_48 = l_Lean_Parser_ParserState_mkNode(x_45, x_9, x_35); x_49 = lean_ctor_get(x_48, 4); lean_inc(x_49); -x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_32); +x_50 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_49, x_32); lean_dec(x_49); if (x_50 == 0) { @@ -89783,7 +90362,7 @@ x_52 = lean_apply_2(x_7, x_1, x_45); x_53 = l_Lean_Parser_ParserState_mkNode(x_52, x_9, x_35); x_54 = lean_ctor_get(x_53, 4); lean_inc(x_54); -x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_32); +x_55 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_54, x_32); lean_dec(x_54); if (x_55 == 0) { @@ -89947,7 +90526,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_assert_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(293u); +x_1 = lean_unsigned_to_nat(300u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -89959,7 +90538,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_assert_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(293u); +x_1 = lean_unsigned_to_nat(300u); x_2 = lean_unsigned_to_nat(130u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -89987,7 +90566,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_assert_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(293u); +x_1 = lean_unsigned_to_nat(300u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -89999,7 +90578,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_assert_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(293u); +x_1 = lean_unsigned_to_nat(300u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -90449,7 +91028,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -90480,7 +91059,7 @@ if (x_24 == 0) lean_object* x_25; uint8_t x_26; x_25 = lean_ctor_get(x_20, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_14); lean_dec(x_25); if (x_26 == 0) { @@ -90489,7 +91068,7 @@ x_27 = l_Lean_Parser_Term_macroDollarArg___elambda__1___closed__2; x_28 = l_Lean_Parser_ParserState_mkNode(x_20, x_27, x_17); x_29 = lean_ctor_get(x_28, 4); lean_inc(x_29); -x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); +x_30 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_29, x_14); lean_dec(x_29); if (x_30 == 0) { @@ -90515,7 +91094,7 @@ x_35 = l_Lean_Parser_Term_macroDollarArg___elambda__1___closed__2; x_36 = l_Lean_Parser_ParserState_mkNode(x_34, x_35, x_17); x_37 = lean_ctor_get(x_36, 4); lean_inc(x_37); -x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); +x_38 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_37, x_14); lean_dec(x_37); if (x_38 == 0) { @@ -90538,7 +91117,7 @@ lean_inc(x_1); x_40 = l_Lean_Parser_tokenAntiquotFn(x_1, x_20); x_41 = lean_ctor_get(x_40, 4); lean_inc(x_41); -x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); +x_42 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_41, x_14); lean_dec(x_41); if (x_42 == 0) { @@ -90547,7 +91126,7 @@ x_43 = l_Lean_Parser_Term_macroDollarArg___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_40, x_43, x_17); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); lean_dec(x_45); if (x_46 == 0) { @@ -90573,7 +91152,7 @@ x_51 = l_Lean_Parser_Term_macroDollarArg___elambda__1___closed__2; x_52 = l_Lean_Parser_ParserState_mkNode(x_50, x_51, x_17); x_53 = lean_ctor_get(x_52, 4); lean_inc(x_53); -x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); +x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_53, x_14); lean_dec(x_53); if (x_54 == 0) { @@ -90927,7 +91506,7 @@ x_14 = l_Lean_Parser_checkPrecFn(x_13, x_1, x_2); x_15 = lean_ctor_get(x_14, 4); lean_inc(x_15); x_16 = lean_box(0); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_15, x_16); lean_dec(x_15); if (x_17 == 0) { @@ -90972,7 +91551,7 @@ goto block_49; lean_object* x_28; uint8_t x_29; x_28 = lean_ctor_get(x_27, 4); lean_inc(x_28); -x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_16); +x_29 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_28, x_16); lean_dec(x_28); if (x_29 == 0) { @@ -90982,7 +91561,7 @@ x_30 = l_Lean_Parser_Term_stateRefT___elambda__1___closed__2; x_31 = l_Lean_Parser_ParserState_mkNode(x_27, x_30, x_19); x_32 = lean_ctor_get(x_31, 4); lean_inc(x_32); -x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_16); +x_33 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_32, x_16); lean_dec(x_32); if (x_33 == 0) { @@ -91004,7 +91583,7 @@ lean_inc(x_1); x_35 = lean_apply_2(x_4, x_1, x_27); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_16); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_16); lean_dec(x_36); if (x_37 == 0) { @@ -91013,7 +91592,7 @@ x_38 = l_Lean_Parser_Term_stateRefT___elambda__1___closed__2; x_39 = l_Lean_Parser_ParserState_mkNode(x_35, x_38, x_19); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_16); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_16); lean_dec(x_40); if (x_41 == 0) { @@ -91037,7 +91616,7 @@ x_44 = l_Lean_Parser_Term_stateRefT___elambda__1___closed__2; x_45 = l_Lean_Parser_ParserState_mkNode(x_43, x_44, x_19); x_46 = lean_ctor_get(x_45, 4); lean_inc(x_46); -x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_16); +x_47 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_46, x_16); lean_dec(x_46); if (x_47 == 0) { @@ -91188,7 +91767,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_stateRefT_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(301u); +x_1 = lean_unsigned_to_nat(308u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -91200,7 +91779,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_stateRefT_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(301u); +x_1 = lean_unsigned_to_nat(308u); x_2 = lean_unsigned_to_nat(94u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -91228,7 +91807,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_stateRefT_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(301u); +x_1 = lean_unsigned_to_nat(308u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -91240,7 +91819,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_stateRefT_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(301u); +x_1 = lean_unsigned_to_nat(308u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -91838,7 +92417,7 @@ static lean_object* _init_l_Lean_Parser_Term_dynamicQuot___elambda__1___closed__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__9; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__7; x_2 = l_Lean_Parser_Term_dynamicQuot___elambda__1___closed__7; x_3 = lean_alloc_closure((void*)(l___private_Lean_Parser_Basic_0__Lean_Parser_addQuotDepthFn___boxed), 4, 2); lean_closure_set(x_3, 0, x_1); @@ -91997,7 +92576,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -92040,7 +92619,7 @@ goto block_84; lean_object* x_26; uint8_t x_27; x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); lean_dec(x_26); if (x_27 == 0) { @@ -92050,7 +92629,7 @@ x_28 = l_Lean_Parser_Term_dynamicQuot___elambda__1___closed__2; x_29 = l_Lean_Parser_ParserState_mkNode(x_25, x_28, x_17); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); lean_dec(x_30); if (x_31 == 0) { @@ -92072,7 +92651,7 @@ lean_inc(x_1); x_33 = l_Lean_Parser_Term_ident___elambda__1(x_1, x_25); x_34 = lean_ctor_get(x_33, 4); lean_inc(x_34); -x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_14); +x_35 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_34, x_14); lean_dec(x_34); if (x_35 == 0) { @@ -92082,7 +92661,7 @@ x_36 = l_Lean_Parser_Term_dynamicQuot___elambda__1___closed__2; x_37 = l_Lean_Parser_ParserState_mkNode(x_33, x_36, x_17); x_38 = lean_ctor_get(x_37, 4); lean_inc(x_38); -x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_14); +x_39 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_38, x_14); lean_dec(x_38); if (x_39 == 0) { @@ -92127,7 +92706,7 @@ goto block_82; lean_object* x_48; uint8_t x_49; x_48 = lean_ctor_get(x_47, 4); lean_inc(x_48); -x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_14); +x_49 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_48, x_14); lean_dec(x_48); if (x_49 == 0) { @@ -92137,7 +92716,7 @@ x_50 = l_Lean_Parser_Term_dynamicQuot___elambda__1___closed__2; x_51 = l_Lean_Parser_ParserState_mkNode(x_47, x_50, x_17); x_52 = lean_ctor_get(x_51, 4); lean_inc(x_52); -x_53 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_52, x_14); +x_53 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_52, x_14); lean_dec(x_52); if (x_53 == 0) { @@ -92155,13 +92734,13 @@ return x_54; else { lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; -x_55 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__9; +x_55 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__7; x_56 = l_Lean_Parser_Term_dynamicQuot___elambda__1___closed__7; lean_inc(x_1); x_57 = l___private_Lean_Parser_Basic_0__Lean_Parser_addQuotDepthFn(x_55, x_56, x_1, x_47); x_58 = lean_ctor_get(x_57, 4); lean_inc(x_58); -x_59 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_58, x_14); +x_59 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_58, x_14); lean_dec(x_58); if (x_59 == 0) { @@ -92171,7 +92750,7 @@ x_60 = l_Lean_Parser_Term_dynamicQuot___elambda__1___closed__2; x_61 = l_Lean_Parser_ParserState_mkNode(x_57, x_60, x_17); x_62 = lean_ctor_get(x_61, 4); lean_inc(x_62); -x_63 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_62, x_14); +x_63 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_62, x_14); lean_dec(x_62); if (x_63 == 0) { @@ -92206,7 +92785,7 @@ x_71 = l_Lean_Parser_Term_dynamicQuot___elambda__1___closed__2; x_72 = l_Lean_Parser_ParserState_mkNode(x_67, x_71, x_17); x_73 = lean_ctor_get(x_72, 4); lean_inc(x_73); -x_74 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_73, x_14); +x_74 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_73, x_14); lean_dec(x_73); if (x_74 == 0) { @@ -92230,7 +92809,7 @@ x_77 = l_Lean_Parser_Term_dynamicQuot___elambda__1___closed__2; x_78 = l_Lean_Parser_ParserState_mkNode(x_76, x_77, x_17); x_79 = lean_ctor_get(x_78, 4); lean_inc(x_79); -x_80 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_79, x_14); +x_80 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_79, x_14); lean_dec(x_79); if (x_80 == 0) { @@ -92414,7 +92993,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_dynamicQuot_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(303u); +x_1 = lean_unsigned_to_nat(310u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -92426,7 +93005,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_dynamicQuot_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(303u); +x_1 = lean_unsigned_to_nat(310u); x_2 = lean_unsigned_to_nat(116u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -92454,7 +93033,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_dynamicQuot_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(303u); +x_1 = lean_unsigned_to_nat(310u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -92466,7 +93045,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_dynamicQuot_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(303u); +x_1 = lean_unsigned_to_nat(310u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -92933,7 +93512,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -92977,7 +93556,7 @@ goto block_48; lean_object* x_26; uint8_t x_27; x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); lean_dec(x_26); if (x_27 == 0) { @@ -92986,7 +93565,7 @@ x_28 = l_Lean_Parser_Term_dotIdent___elambda__1___closed__2; x_29 = l_Lean_Parser_ParserState_mkNode(x_25, x_28, x_17); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); lean_dec(x_30); if (x_31 == 0) { @@ -93008,7 +93587,7 @@ x_33 = l_Lean_Parser_Term_doubleQuotedName___elambda__1___closed__9; x_34 = l_Lean_Parser_checkNoWsBeforeFn(x_33, x_1, x_25); x_35 = lean_ctor_get(x_34, 4); lean_inc(x_35); -x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_14); +x_36 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_35, x_14); lean_dec(x_35); if (x_36 == 0) { @@ -93017,7 +93596,7 @@ x_37 = l_Lean_Parser_Term_dotIdent___elambda__1___closed__2; x_38 = l_Lean_Parser_ParserState_mkNode(x_34, x_37, x_17); x_39 = lean_ctor_get(x_38, 4); lean_inc(x_39); -x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); +x_40 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_39, x_14); lean_dec(x_39); if (x_40 == 0) { @@ -93041,7 +93620,7 @@ x_43 = l_Lean_Parser_Term_dotIdent___elambda__1___closed__2; x_44 = l_Lean_Parser_ParserState_mkNode(x_42, x_43, x_17); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_14); lean_dec(x_45); if (x_46 == 0) { @@ -93180,7 +93759,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_dotIdent_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(305u); +x_1 = lean_unsigned_to_nat(312u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -93192,7 +93771,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_dotIdent_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(305u); +x_1 = lean_unsigned_to_nat(312u); x_2 = lean_unsigned_to_nat(86u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -93220,7 +93799,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_dotIdent_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(305u); +x_1 = lean_unsigned_to_nat(312u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -93232,7 +93811,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_dotIdent_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(305u); +x_1 = lean_unsigned_to_nat(312u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -93475,7 +94054,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_quot___elambda__1___closed__1() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__2; -x_2 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__3; +x_2 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } @@ -93484,7 +94063,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_quot___elambda__1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__3; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__1; x_2 = l_Lean_Parser_Tactic_quot___elambda__1___closed__1; x_3 = 1; x_4 = 0; @@ -93523,7 +94102,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_quot___elambda__1___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__9; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__7; x_2 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__14; x_3 = lean_alloc_closure((void*)(l___private_Lean_Parser_Basic_0__Lean_Parser_addQuotDepthFn___boxed), 4, 2); lean_closure_set(x_3, 0, x_1); @@ -93638,7 +94217,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -93681,7 +94260,7 @@ goto block_60; lean_object* x_26; uint8_t x_27; x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); lean_dec(x_26); if (x_27 == 0) { @@ -93691,7 +94270,7 @@ x_28 = l_Lean_Parser_Tactic_quot___elambda__1___closed__1; x_29 = l_Lean_Parser_ParserState_mkNode(x_25, x_28, x_17); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); lean_dec(x_30); if (x_31 == 0) { @@ -93709,13 +94288,13 @@ return x_32; else { lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_33 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__9; +x_33 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__7; x_34 = l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__14; lean_inc(x_1); x_35 = l___private_Lean_Parser_Basic_0__Lean_Parser_addQuotDepthFn(x_33, x_34, x_1, x_25); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); lean_dec(x_36); if (x_37 == 0) { @@ -93725,7 +94304,7 @@ x_38 = l_Lean_Parser_Tactic_quot___elambda__1___closed__1; x_39 = l_Lean_Parser_ParserState_mkNode(x_35, x_38, x_17); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); lean_dec(x_40); if (x_41 == 0) { @@ -93760,7 +94339,7 @@ x_49 = l_Lean_Parser_Tactic_quot___elambda__1___closed__1; x_50 = l_Lean_Parser_ParserState_mkNode(x_45, x_49, x_17); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_14); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_14); lean_dec(x_51); if (x_52 == 0) { @@ -93784,7 +94363,7 @@ x_55 = l_Lean_Parser_Tactic_quot___elambda__1___closed__1; x_56 = l_Lean_Parser_ParserState_mkNode(x_54, x_55, x_17); x_57 = lean_ctor_get(x_56, 4); lean_inc(x_57); -x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_14); +x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_14); lean_dec(x_57); if (x_58 == 0) { @@ -93933,7 +94512,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_quot_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(309u); +x_1 = lean_unsigned_to_nat(316u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -93945,7 +94524,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_quot_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(309u); +x_1 = lean_unsigned_to_nat(316u); x_2 = lean_unsigned_to_nat(124u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -93973,7 +94552,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_quot_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(309u); +x_1 = lean_unsigned_to_nat(316u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -93985,7 +94564,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_quot_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(309u); +x_1 = lean_unsigned_to_nat(316u); x_2 = lean_unsigned_to_nat(46u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -94035,7 +94614,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_quot_formatter___closed__1() { _start: { lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__3; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__1; x_2 = l_Lean_Parser_Tactic_quot___elambda__1___closed__1; x_3 = 1; x_4 = 0; @@ -94151,7 +94730,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_quot_parenthesizer___closed__1() _start: { lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__3; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__1; x_2 = l_Lean_Parser_Tactic_quot___elambda__1___closed__1; x_3 = 1; x_4 = 0; @@ -94297,7 +94876,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_quotSeq___elambda__1___closed__4( _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__9; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__7; x_2 = l_Lean_Parser_Tactic_seq1___closed__2; x_3 = lean_alloc_closure((void*)(l___private_Lean_Parser_Basic_0__Lean_Parser_addQuotDepthFn___boxed), 4, 2); lean_closure_set(x_3, 0, x_1); @@ -94392,7 +94971,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -94435,7 +95014,7 @@ goto block_60; lean_object* x_26; uint8_t x_27; x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); lean_dec(x_26); if (x_27 == 0) { @@ -94445,7 +95024,7 @@ x_28 = l_Lean_Parser_Tactic_quotSeq___elambda__1___closed__2; x_29 = l_Lean_Parser_ParserState_mkNode(x_25, x_28, x_17); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); lean_dec(x_30); if (x_31 == 0) { @@ -94463,13 +95042,13 @@ return x_32; else { lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_33 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__9; +x_33 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__7; x_34 = l_Lean_Parser_Tactic_seq1___closed__2; lean_inc(x_1); x_35 = l___private_Lean_Parser_Basic_0__Lean_Parser_addQuotDepthFn(x_33, x_34, x_1, x_25); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); lean_dec(x_36); if (x_37 == 0) { @@ -94479,7 +95058,7 @@ x_38 = l_Lean_Parser_Tactic_quotSeq___elambda__1___closed__2; x_39 = l_Lean_Parser_ParserState_mkNode(x_35, x_38, x_17); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); lean_dec(x_40); if (x_41 == 0) { @@ -94514,7 +95093,7 @@ x_49 = l_Lean_Parser_Tactic_quotSeq___elambda__1___closed__2; x_50 = l_Lean_Parser_ParserState_mkNode(x_45, x_49, x_17); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_14); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_14); lean_dec(x_51); if (x_52 == 0) { @@ -94538,7 +95117,7 @@ x_55 = l_Lean_Parser_Tactic_quotSeq___elambda__1___closed__2; x_56 = l_Lean_Parser_ParserState_mkNode(x_54, x_55, x_17); x_57 = lean_ctor_get(x_56, 4); lean_inc(x_57); -x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_14); +x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_14); lean_dec(x_57); if (x_58 == 0) { @@ -94678,7 +95257,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_quotSeq_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(310u); +x_1 = lean_unsigned_to_nat(317u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -94690,7 +95269,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_quotSeq_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(310u); +x_1 = lean_unsigned_to_nat(317u); x_2 = lean_unsigned_to_nat(113u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -94718,7 +95297,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_quotSeq_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(310u); +x_1 = lean_unsigned_to_nat(317u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -94730,7 +95309,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_quotSeq_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(310u); +x_1 = lean_unsigned_to_nat(317u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -95145,7 +95724,7 @@ static lean_object* _init_l_Lean_Parser_Level_quot___elambda__1___closed__3() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Level_quot___elambda__1___closed__2; -x_2 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__3; +x_2 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } @@ -95154,7 +95733,7 @@ static lean_object* _init_l_Lean_Parser_Level_quot___elambda__1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__3; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__1; x_2 = l_Lean_Parser_Level_quot___elambda__1___closed__3; x_3 = 1; x_4 = 0; @@ -95205,7 +95784,7 @@ static lean_object* _init_l_Lean_Parser_Level_quot___elambda__1___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__9; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__7; x_2 = l_Lean_Parser_Level_quot___elambda__1___closed__8; x_3 = lean_alloc_closure((void*)(l___private_Lean_Parser_Basic_0__Lean_Parser_addQuotDepthFn___boxed), 4, 2); lean_closure_set(x_3, 0, x_1); @@ -95320,7 +95899,7 @@ x_12 = l_Lean_Parser_checkPrecFn(x_11, x_1, x_2); x_13 = lean_ctor_get(x_12, 4); lean_inc(x_13); x_14 = lean_box(0); -x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); +x_15 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -95363,7 +95942,7 @@ goto block_60; lean_object* x_26; uint8_t x_27; x_26 = lean_ctor_get(x_25, 4); lean_inc(x_26); -x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); +x_27 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_26, x_14); lean_dec(x_26); if (x_27 == 0) { @@ -95373,7 +95952,7 @@ x_28 = l_Lean_Parser_Level_quot___elambda__1___closed__3; x_29 = l_Lean_Parser_ParserState_mkNode(x_25, x_28, x_17); x_30 = lean_ctor_get(x_29, 4); lean_inc(x_30); -x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); +x_31 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_30, x_14); lean_dec(x_30); if (x_31 == 0) { @@ -95391,13 +95970,13 @@ return x_32; else { lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_33 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__9; +x_33 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__7; x_34 = l_Lean_Parser_Level_quot___elambda__1___closed__8; lean_inc(x_1); x_35 = l___private_Lean_Parser_Basic_0__Lean_Parser_addQuotDepthFn(x_33, x_34, x_1, x_25); x_36 = lean_ctor_get(x_35, 4); lean_inc(x_36); -x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); +x_37 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_36, x_14); lean_dec(x_36); if (x_37 == 0) { @@ -95407,7 +95986,7 @@ x_38 = l_Lean_Parser_Level_quot___elambda__1___closed__3; x_39 = l_Lean_Parser_ParserState_mkNode(x_35, x_38, x_17); x_40 = lean_ctor_get(x_39, 4); lean_inc(x_40); -x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); +x_41 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_40, x_14); lean_dec(x_40); if (x_41 == 0) { @@ -95442,7 +96021,7 @@ x_49 = l_Lean_Parser_Level_quot___elambda__1___closed__3; x_50 = l_Lean_Parser_ParserState_mkNode(x_45, x_49, x_17); x_51 = lean_ctor_get(x_50, 4); lean_inc(x_51); -x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_14); +x_52 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_51, x_14); lean_dec(x_51); if (x_52 == 0) { @@ -95466,7 +96045,7 @@ x_55 = l_Lean_Parser_Level_quot___elambda__1___closed__3; x_56 = l_Lean_Parser_ParserState_mkNode(x_54, x_55, x_17); x_57 = lean_ctor_get(x_56, 4); lean_inc(x_57); -x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_14); +x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_57, x_14); lean_dec(x_57); if (x_58 == 0) { @@ -95615,7 +96194,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Level_quot_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(312u); +x_1 = lean_unsigned_to_nat(319u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -95627,7 +96206,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Level_quot_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(312u); +x_1 = lean_unsigned_to_nat(319u); x_2 = lean_unsigned_to_nat(109u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -95655,7 +96234,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Level_quot_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(312u); +x_1 = lean_unsigned_to_nat(319u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -95667,7 +96246,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Level_quot_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(312u); +x_1 = lean_unsigned_to_nat(319u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -95717,7 +96296,7 @@ static lean_object* _init_l_Lean_Parser_Level_quot_formatter___closed__1() { _start: { lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__3; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__1; x_2 = l_Lean_Parser_Level_quot___elambda__1___closed__3; x_3 = 1; x_4 = 0; @@ -95833,7 +96412,7 @@ static lean_object* _init_l_Lean_Parser_Level_quot_parenthesizer___closed__1() { _start: { lean_object* x_1; lean_object* x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__3; +x_1 = l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__1; x_2 = l_Lean_Parser_Level_quot___elambda__1___closed__3; x_3 = 1; x_4 = 0; @@ -95945,7 +96524,7 @@ x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); return x_6; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__1() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -95955,7 +96534,7 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__2() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__2() { _start: { lean_object* x_1; lean_object* x_2; @@ -95965,7 +96544,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__3() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__3() { _start: { lean_object* x_1; lean_object* x_2; @@ -95975,7 +96554,29 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__4() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__5() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__4; +x_2 = 1; +x_3 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__6() { _start: { lean_object* x_1; lean_object* x_2; @@ -95985,7 +96586,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__5() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__7() { _start: { lean_object* x_1; @@ -95993,7 +96594,7 @@ x_1 = l_Lean_PrettyPrinter_Formatter_formatterAliasesRef; return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__6() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__8() { _start: { lean_object* x_1; lean_object* x_2; @@ -96003,7 +96604,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__7() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__9() { _start: { lean_object* x_1; @@ -96011,7 +96612,7 @@ x_1 = l_Lean_PrettyPrinter_Parenthesizer_parenthesizerAliasesRef; return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__8() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -96021,7 +96622,7 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__9() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__11() { _start: { lean_object* x_1; lean_object* x_2; @@ -96031,7 +96632,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__10() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__12() { _start: { lean_object* x_1; lean_object* x_2; @@ -96041,7 +96642,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__11() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__13() { _start: { lean_object* x_1; lean_object* x_2; @@ -96051,7 +96652,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__12() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__14() { _start: { lean_object* x_1; lean_object* x_2; @@ -96061,7 +96662,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__13() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -96071,7 +96672,7 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__14() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__16() { _start: { lean_object* x_1; lean_object* x_2; @@ -96081,7 +96682,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__15() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__17() { _start: { lean_object* x_1; lean_object* x_2; @@ -96091,7 +96692,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__16() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__18() { _start: { lean_object* x_1; lean_object* x_2; @@ -96101,7 +96702,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__17() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__19() { _start: { lean_object* x_1; lean_object* x_2; @@ -96111,7 +96712,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__18() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -96121,7 +96722,7 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__19() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__21() { _start: { lean_object* x_1; lean_object* x_2; @@ -96131,7 +96732,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__20() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__22() { _start: { lean_object* x_1; lean_object* x_2; @@ -96141,7 +96742,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__21() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__23() { _start: { lean_object* x_1; lean_object* x_2; @@ -96151,7 +96752,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__22() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__24() { _start: { lean_object* x_1; lean_object* x_2; @@ -96161,7 +96762,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__23() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -96171,7 +96772,7 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__24() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__26() { _start: { lean_object* x_1; lean_object* x_2; @@ -96181,7 +96782,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__25() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__27() { _start: { lean_object* x_1; lean_object* x_2; @@ -96191,7 +96792,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__26() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__28() { _start: { lean_object* x_1; lean_object* x_2; @@ -96201,7 +96802,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__27() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__29() { _start: { lean_object* x_1; lean_object* x_2; @@ -96211,7 +96812,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__28() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__30() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -96221,7 +96822,7 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__29() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__31() { _start: { lean_object* x_1; lean_object* x_2; @@ -96231,7 +96832,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__30() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__32() { _start: { lean_object* x_1; lean_object* x_2; @@ -96241,7 +96842,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__31() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__33() { _start: { lean_object* x_1; lean_object* x_2; @@ -96251,7 +96852,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__32() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__34() { _start: { lean_object* x_1; lean_object* x_2; @@ -96261,7 +96862,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__33() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__35() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -96271,7 +96872,7 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__34() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__36() { _start: { lean_object* x_1; lean_object* x_2; @@ -96281,7 +96882,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__35() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__37() { _start: { lean_object* x_1; lean_object* x_2; @@ -96291,7 +96892,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__36() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__38() { _start: { lean_object* x_1; lean_object* x_2; @@ -96301,7 +96902,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__37() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__39() { _start: { lean_object* x_1; lean_object* x_2; @@ -96311,7 +96912,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__38() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__40() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -96321,7 +96922,7 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__39() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__41() { _start: { lean_object* x_1; lean_object* x_2; @@ -96331,7 +96932,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__40() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__42() { _start: { lean_object* x_1; lean_object* x_2; @@ -96341,7 +96942,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__41() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__43() { _start: { lean_object* x_1; lean_object* x_2; @@ -96351,7 +96952,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__42() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__44() { _start: { lean_object* x_1; lean_object* x_2; @@ -96361,7 +96962,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__43() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__45() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -96371,7 +96972,7 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__44() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__46() { _start: { lean_object* x_1; lean_object* x_2; @@ -96381,7 +96982,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__45() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__47() { _start: { lean_object* x_1; lean_object* x_2; @@ -96391,7 +96992,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__46() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__48() { _start: { lean_object* x_1; lean_object* x_2; @@ -96401,7 +97002,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__47() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__49() { _start: { lean_object* x_1; lean_object* x_2; @@ -96411,837 +97012,838 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883_(lean_object* x_1) { _start: { -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__1; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__2; -x_4 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__3; -x_5 = l_Lean_Parser_registerAlias(x_2, x_3, x_4, x_1); -if (lean_obj_tag(x_5) == 0) -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_6 = lean_ctor_get(x_5, 1); -lean_inc(x_6); -lean_dec(x_5); -x_7 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__5; -x_8 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__4; -x_9 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_2, x_8, x_6); -if (lean_obj_tag(x_9) == 0) +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__1; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__2; +x_4 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__3; +x_5 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__5; +x_6 = l_Lean_Parser_registerAlias(x_2, x_3, x_4, x_5, x_1); +if (lean_obj_tag(x_6) == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_10 = lean_ctor_get(x_9, 1); -lean_inc(x_10); -lean_dec(x_9); -x_11 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__7; -x_12 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__6; -x_13 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_2, x_12, x_10); -if (lean_obj_tag(x_13) == 0) +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_7 = lean_ctor_get(x_6, 1); +lean_inc(x_7); +lean_dec(x_6); +x_8 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__7; +x_9 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__6; +x_10 = l_Lean_Parser_registerAliasCore___rarg(x_8, x_2, x_9, x_7); +if (lean_obj_tag(x_10) == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -lean_dec(x_13); -x_15 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__8; -x_16 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__9; -x_17 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__10; -x_18 = l_Lean_Parser_registerAlias(x_15, x_16, x_17, x_14); -if (lean_obj_tag(x_18) == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -lean_dec(x_18); -x_20 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__11; -x_21 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_15, x_20, x_19); -if (lean_obj_tag(x_21) == 0) +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +lean_dec(x_10); +x_12 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__9; +x_13 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__8; +x_14 = l_Lean_Parser_registerAliasCore___rarg(x_12, x_2, x_13, x_11); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_21, 1); -lean_inc(x_22); -lean_dec(x_21); -x_23 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__12; -x_24 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_15, x_23, x_22); -if (lean_obj_tag(x_24) == 0) +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +lean_dec(x_14); +x_16 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__10; +x_17 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__11; +x_18 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__12; +x_19 = l_Lean_Parser_registerAlias(x_16, x_17, x_18, x_5, x_15); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_19, 1); +lean_inc(x_20); +lean_dec(x_19); +x_21 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__13; +x_22 = l_Lean_Parser_registerAliasCore___rarg(x_8, x_16, x_21, x_20); +if (lean_obj_tag(x_22) == 0) { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_25 = lean_ctor_get(x_24, 1); -lean_inc(x_25); -lean_dec(x_24); -x_26 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__13; -x_27 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__14; -x_28 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__15; -x_29 = l_Lean_Parser_registerAlias(x_26, x_27, x_28, x_25); -if (lean_obj_tag(x_29) == 0) -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_29, 1); -lean_inc(x_30); -lean_dec(x_29); -x_31 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__16; -x_32 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_26, x_31, x_30); -if (lean_obj_tag(x_32) == 0) +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +lean_dec(x_22); +x_24 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__14; +x_25 = l_Lean_Parser_registerAliasCore___rarg(x_12, x_16, x_24, x_23); +if (lean_obj_tag(x_25) == 0) { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_32, 1); -lean_inc(x_33); -lean_dec(x_32); -x_34 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__17; -x_35 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_26, x_34, x_33); -if (lean_obj_tag(x_35) == 0) +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +x_27 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__15; +x_28 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__16; +x_29 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__17; +x_30 = l_Lean_Parser_registerAlias(x_27, x_28, x_29, x_5, x_26); +if (lean_obj_tag(x_30) == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_32 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__18; +x_33 = l_Lean_Parser_registerAliasCore___rarg(x_8, x_27, x_32, x_31); +if (lean_obj_tag(x_33) == 0) { -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_36 = lean_ctor_get(x_35, 1); -lean_inc(x_36); -lean_dec(x_35); -x_37 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__18; -x_38 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__19; -x_39 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__20; -x_40 = l_Lean_Parser_registerAlias(x_37, x_38, x_39, x_36); -if (lean_obj_tag(x_40) == 0) -{ -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_40, 1); -lean_inc(x_41); -lean_dec(x_40); -x_42 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__21; -x_43 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_37, x_42, x_41); -if (lean_obj_tag(x_43) == 0) +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_33, 1); +lean_inc(x_34); +lean_dec(x_33); +x_35 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__19; +x_36 = l_Lean_Parser_registerAliasCore___rarg(x_12, x_27, x_35, x_34); +if (lean_obj_tag(x_36) == 0) { -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_43, 1); -lean_inc(x_44); -lean_dec(x_43); -x_45 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__22; -x_46 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_37, x_45, x_44); -if (lean_obj_tag(x_46) == 0) +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_37 = lean_ctor_get(x_36, 1); +lean_inc(x_37); +lean_dec(x_36); +x_38 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__20; +x_39 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__21; +x_40 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__22; +x_41 = l_Lean_Parser_registerAlias(x_38, x_39, x_40, x_5, x_37); +if (lean_obj_tag(x_41) == 0) +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_41, 1); +lean_inc(x_42); +lean_dec(x_41); +x_43 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__23; +x_44 = l_Lean_Parser_registerAliasCore___rarg(x_8, x_38, x_43, x_42); +if (lean_obj_tag(x_44) == 0) { -lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_47 = lean_ctor_get(x_46, 1); -lean_inc(x_47); -lean_dec(x_46); -x_48 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__23; -x_49 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__24; -x_50 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__25; -x_51 = l_Lean_Parser_registerAlias(x_48, x_49, x_50, x_47); -if (lean_obj_tag(x_51) == 0) -{ -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_51, 1); -lean_inc(x_52); -lean_dec(x_51); -x_53 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__26; -x_54 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_48, x_53, x_52); -if (lean_obj_tag(x_54) == 0) +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_44, 1); +lean_inc(x_45); +lean_dec(x_44); +x_46 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__24; +x_47 = l_Lean_Parser_registerAliasCore___rarg(x_12, x_38, x_46, x_45); +if (lean_obj_tag(x_47) == 0) { -lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_55 = lean_ctor_get(x_54, 1); -lean_inc(x_55); -lean_dec(x_54); -x_56 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__27; -x_57 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_48, x_56, x_55); -if (lean_obj_tag(x_57) == 0) +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_48 = lean_ctor_get(x_47, 1); +lean_inc(x_48); +lean_dec(x_47); +x_49 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__25; +x_50 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__26; +x_51 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__27; +x_52 = l_Lean_Parser_registerAlias(x_49, x_50, x_51, x_5, x_48); +if (lean_obj_tag(x_52) == 0) +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_52, 1); +lean_inc(x_53); +lean_dec(x_52); +x_54 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__28; +x_55 = l_Lean_Parser_registerAliasCore___rarg(x_8, x_49, x_54, x_53); +if (lean_obj_tag(x_55) == 0) { -lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_58 = lean_ctor_get(x_57, 1); -lean_inc(x_58); -lean_dec(x_57); -x_59 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__28; -x_60 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__29; -x_61 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__30; -x_62 = l_Lean_Parser_registerAlias(x_59, x_60, x_61, x_58); -if (lean_obj_tag(x_62) == 0) -{ -lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_63 = lean_ctor_get(x_62, 1); -lean_inc(x_63); -lean_dec(x_62); -x_64 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__31; -x_65 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_59, x_64, x_63); -if (lean_obj_tag(x_65) == 0) +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_55, 1); +lean_inc(x_56); +lean_dec(x_55); +x_57 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__29; +x_58 = l_Lean_Parser_registerAliasCore___rarg(x_12, x_49, x_57, x_56); +if (lean_obj_tag(x_58) == 0) { -lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_66 = lean_ctor_get(x_65, 1); -lean_inc(x_66); -lean_dec(x_65); -x_67 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__32; -x_68 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_59, x_67, x_66); -if (lean_obj_tag(x_68) == 0) +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_59 = lean_ctor_get(x_58, 1); +lean_inc(x_59); +lean_dec(x_58); +x_60 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__30; +x_61 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__31; +x_62 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__32; +x_63 = l_Lean_Parser_registerAlias(x_60, x_61, x_62, x_5, x_59); +if (lean_obj_tag(x_63) == 0) +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_63, 1); +lean_inc(x_64); +lean_dec(x_63); +x_65 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__33; +x_66 = l_Lean_Parser_registerAliasCore___rarg(x_8, x_60, x_65, x_64); +if (lean_obj_tag(x_66) == 0) { -lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; -x_69 = lean_ctor_get(x_68, 1); -lean_inc(x_69); -lean_dec(x_68); -x_70 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__33; -x_71 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__34; -x_72 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__35; -x_73 = l_Lean_Parser_registerAlias(x_70, x_71, x_72, x_69); -if (lean_obj_tag(x_73) == 0) -{ -lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_74 = lean_ctor_get(x_73, 1); -lean_inc(x_74); -lean_dec(x_73); -x_75 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__36; -x_76 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_70, x_75, x_74); -if (lean_obj_tag(x_76) == 0) +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_66, 1); +lean_inc(x_67); +lean_dec(x_66); +x_68 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__34; +x_69 = l_Lean_Parser_registerAliasCore___rarg(x_12, x_60, x_68, x_67); +if (lean_obj_tag(x_69) == 0) { -lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_77 = lean_ctor_get(x_76, 1); -lean_inc(x_77); -lean_dec(x_76); -x_78 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__37; -x_79 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_70, x_78, x_77); -if (lean_obj_tag(x_79) == 0) -{ -lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; -x_80 = lean_ctor_get(x_79, 1); -lean_inc(x_80); -lean_dec(x_79); -x_81 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__38; -x_82 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__39; -x_83 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__40; -x_84 = l_Lean_Parser_registerAlias(x_81, x_82, x_83, x_80); -if (lean_obj_tag(x_84) == 0) -{ -lean_object* x_85; lean_object* x_86; lean_object* x_87; -x_85 = lean_ctor_get(x_84, 1); -lean_inc(x_85); -lean_dec(x_84); -x_86 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__41; -x_87 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_81, x_86, x_85); -if (lean_obj_tag(x_87) == 0) -{ -lean_object* x_88; lean_object* x_89; lean_object* x_90; -x_88 = lean_ctor_get(x_87, 1); -lean_inc(x_88); -lean_dec(x_87); -x_89 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__42; -x_90 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_81, x_89, x_88); -if (lean_obj_tag(x_90) == 0) -{ -lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; -x_91 = lean_ctor_get(x_90, 1); -lean_inc(x_91); -lean_dec(x_90); -x_92 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__43; -x_93 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__44; -x_94 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__45; -x_95 = l_Lean_Parser_registerAlias(x_92, x_93, x_94, x_91); -if (lean_obj_tag(x_95) == 0) -{ -lean_object* x_96; lean_object* x_97; lean_object* x_98; -x_96 = lean_ctor_get(x_95, 1); -lean_inc(x_96); -lean_dec(x_95); -x_97 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__46; -x_98 = l_Lean_Parser_registerAliasCore___rarg(x_7, x_92, x_97, x_96); -if (lean_obj_tag(x_98) == 0) -{ -lean_object* x_99; lean_object* x_100; lean_object* x_101; -x_99 = lean_ctor_get(x_98, 1); -lean_inc(x_99); -lean_dec(x_98); -x_100 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__47; -x_101 = l_Lean_Parser_registerAliasCore___rarg(x_11, x_92, x_100, x_99); -return x_101; -} -else -{ -uint8_t x_102; -x_102 = !lean_is_exclusive(x_98); -if (x_102 == 0) -{ -return x_98; -} -else -{ -lean_object* x_103; lean_object* x_104; lean_object* x_105; -x_103 = lean_ctor_get(x_98, 0); -x_104 = lean_ctor_get(x_98, 1); +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_70 = lean_ctor_get(x_69, 1); +lean_inc(x_70); +lean_dec(x_69); +x_71 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__35; +x_72 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__36; +x_73 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__37; +x_74 = l_Lean_Parser_registerAlias(x_71, x_72, x_73, x_5, x_70); +if (lean_obj_tag(x_74) == 0) +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_75 = lean_ctor_get(x_74, 1); +lean_inc(x_75); +lean_dec(x_74); +x_76 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__38; +x_77 = l_Lean_Parser_registerAliasCore___rarg(x_8, x_71, x_76, x_75); +if (lean_obj_tag(x_77) == 0) +{ +lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_78 = lean_ctor_get(x_77, 1); +lean_inc(x_78); +lean_dec(x_77); +x_79 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__39; +x_80 = l_Lean_Parser_registerAliasCore___rarg(x_12, x_71, x_79, x_78); +if (lean_obj_tag(x_80) == 0) +{ +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_81 = lean_ctor_get(x_80, 1); +lean_inc(x_81); +lean_dec(x_80); +x_82 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__40; +x_83 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__41; +x_84 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__42; +x_85 = l_Lean_Parser_registerAlias(x_82, x_83, x_84, x_5, x_81); +if (lean_obj_tag(x_85) == 0) +{ +lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_86 = lean_ctor_get(x_85, 1); +lean_inc(x_86); +lean_dec(x_85); +x_87 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__43; +x_88 = l_Lean_Parser_registerAliasCore___rarg(x_8, x_82, x_87, x_86); +if (lean_obj_tag(x_88) == 0) +{ +lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_89 = lean_ctor_get(x_88, 1); +lean_inc(x_89); +lean_dec(x_88); +x_90 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__44; +x_91 = l_Lean_Parser_registerAliasCore___rarg(x_12, x_82, x_90, x_89); +if (lean_obj_tag(x_91) == 0) +{ +lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_92 = lean_ctor_get(x_91, 1); +lean_inc(x_92); +lean_dec(x_91); +x_93 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__45; +x_94 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__46; +x_95 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__47; +x_96 = l_Lean_Parser_registerAlias(x_93, x_94, x_95, x_5, x_92); +if (lean_obj_tag(x_96) == 0) +{ +lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_97 = lean_ctor_get(x_96, 1); +lean_inc(x_97); +lean_dec(x_96); +x_98 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__48; +x_99 = l_Lean_Parser_registerAliasCore___rarg(x_8, x_93, x_98, x_97); +if (lean_obj_tag(x_99) == 0) +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_100 = lean_ctor_get(x_99, 1); +lean_inc(x_100); +lean_dec(x_99); +x_101 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__49; +x_102 = l_Lean_Parser_registerAliasCore___rarg(x_12, x_93, x_101, x_100); +return x_102; +} +else +{ +uint8_t x_103; +x_103 = !lean_is_exclusive(x_99); +if (x_103 == 0) +{ +return x_99; +} +else +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_104 = lean_ctor_get(x_99, 0); +x_105 = lean_ctor_get(x_99, 1); +lean_inc(x_105); lean_inc(x_104); -lean_inc(x_103); -lean_dec(x_98); -x_105 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_105, 0, x_103); -lean_ctor_set(x_105, 1, x_104); -return x_105; +lean_dec(x_99); +x_106 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_106, 0, x_104); +lean_ctor_set(x_106, 1, x_105); +return x_106; } } } else { -uint8_t x_106; -x_106 = !lean_is_exclusive(x_95); -if (x_106 == 0) +uint8_t x_107; +x_107 = !lean_is_exclusive(x_96); +if (x_107 == 0) { -return x_95; +return x_96; } else { -lean_object* x_107; lean_object* x_108; lean_object* x_109; -x_107 = lean_ctor_get(x_95, 0); -x_108 = lean_ctor_get(x_95, 1); +lean_object* x_108; lean_object* x_109; lean_object* x_110; +x_108 = lean_ctor_get(x_96, 0); +x_109 = lean_ctor_get(x_96, 1); +lean_inc(x_109); lean_inc(x_108); -lean_inc(x_107); -lean_dec(x_95); -x_109 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_109, 0, x_107); -lean_ctor_set(x_109, 1, x_108); -return x_109; +lean_dec(x_96); +x_110 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_110, 0, x_108); +lean_ctor_set(x_110, 1, x_109); +return x_110; } } } else { -uint8_t x_110; -x_110 = !lean_is_exclusive(x_90); -if (x_110 == 0) +uint8_t x_111; +x_111 = !lean_is_exclusive(x_91); +if (x_111 == 0) { -return x_90; +return x_91; } else { -lean_object* x_111; lean_object* x_112; lean_object* x_113; -x_111 = lean_ctor_get(x_90, 0); -x_112 = lean_ctor_get(x_90, 1); +lean_object* x_112; lean_object* x_113; lean_object* x_114; +x_112 = lean_ctor_get(x_91, 0); +x_113 = lean_ctor_get(x_91, 1); +lean_inc(x_113); lean_inc(x_112); -lean_inc(x_111); -lean_dec(x_90); -x_113 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_113, 0, x_111); -lean_ctor_set(x_113, 1, x_112); -return x_113; +lean_dec(x_91); +x_114 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_114, 0, x_112); +lean_ctor_set(x_114, 1, x_113); +return x_114; } } } else { -uint8_t x_114; -x_114 = !lean_is_exclusive(x_87); -if (x_114 == 0) +uint8_t x_115; +x_115 = !lean_is_exclusive(x_88); +if (x_115 == 0) { -return x_87; +return x_88; } else { -lean_object* x_115; lean_object* x_116; lean_object* x_117; -x_115 = lean_ctor_get(x_87, 0); -x_116 = lean_ctor_get(x_87, 1); +lean_object* x_116; lean_object* x_117; lean_object* x_118; +x_116 = lean_ctor_get(x_88, 0); +x_117 = lean_ctor_get(x_88, 1); +lean_inc(x_117); lean_inc(x_116); -lean_inc(x_115); -lean_dec(x_87); -x_117 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_117, 0, x_115); -lean_ctor_set(x_117, 1, x_116); -return x_117; +lean_dec(x_88); +x_118 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_118, 0, x_116); +lean_ctor_set(x_118, 1, x_117); +return x_118; } } } else { -uint8_t x_118; -x_118 = !lean_is_exclusive(x_84); -if (x_118 == 0) +uint8_t x_119; +x_119 = !lean_is_exclusive(x_85); +if (x_119 == 0) { -return x_84; +return x_85; } else { -lean_object* x_119; lean_object* x_120; lean_object* x_121; -x_119 = lean_ctor_get(x_84, 0); -x_120 = lean_ctor_get(x_84, 1); +lean_object* x_120; lean_object* x_121; lean_object* x_122; +x_120 = lean_ctor_get(x_85, 0); +x_121 = lean_ctor_get(x_85, 1); +lean_inc(x_121); lean_inc(x_120); -lean_inc(x_119); -lean_dec(x_84); -x_121 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_121, 0, x_119); -lean_ctor_set(x_121, 1, x_120); -return x_121; +lean_dec(x_85); +x_122 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_122, 0, x_120); +lean_ctor_set(x_122, 1, x_121); +return x_122; } } } else { -uint8_t x_122; -x_122 = !lean_is_exclusive(x_79); -if (x_122 == 0) +uint8_t x_123; +x_123 = !lean_is_exclusive(x_80); +if (x_123 == 0) { -return x_79; +return x_80; } else { -lean_object* x_123; lean_object* x_124; lean_object* x_125; -x_123 = lean_ctor_get(x_79, 0); -x_124 = lean_ctor_get(x_79, 1); +lean_object* x_124; lean_object* x_125; lean_object* x_126; +x_124 = lean_ctor_get(x_80, 0); +x_125 = lean_ctor_get(x_80, 1); +lean_inc(x_125); lean_inc(x_124); -lean_inc(x_123); -lean_dec(x_79); -x_125 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_125, 0, x_123); -lean_ctor_set(x_125, 1, x_124); -return x_125; +lean_dec(x_80); +x_126 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_126, 0, x_124); +lean_ctor_set(x_126, 1, x_125); +return x_126; } } } else { -uint8_t x_126; -x_126 = !lean_is_exclusive(x_76); -if (x_126 == 0) +uint8_t x_127; +x_127 = !lean_is_exclusive(x_77); +if (x_127 == 0) { -return x_76; +return x_77; } else { -lean_object* x_127; lean_object* x_128; lean_object* x_129; -x_127 = lean_ctor_get(x_76, 0); -x_128 = lean_ctor_get(x_76, 1); +lean_object* x_128; lean_object* x_129; lean_object* x_130; +x_128 = lean_ctor_get(x_77, 0); +x_129 = lean_ctor_get(x_77, 1); +lean_inc(x_129); lean_inc(x_128); -lean_inc(x_127); -lean_dec(x_76); -x_129 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_129, 0, x_127); -lean_ctor_set(x_129, 1, x_128); -return x_129; +lean_dec(x_77); +x_130 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_130, 0, x_128); +lean_ctor_set(x_130, 1, x_129); +return x_130; } } } else { -uint8_t x_130; -x_130 = !lean_is_exclusive(x_73); -if (x_130 == 0) +uint8_t x_131; +x_131 = !lean_is_exclusive(x_74); +if (x_131 == 0) { -return x_73; +return x_74; } else { -lean_object* x_131; lean_object* x_132; lean_object* x_133; -x_131 = lean_ctor_get(x_73, 0); -x_132 = lean_ctor_get(x_73, 1); +lean_object* x_132; lean_object* x_133; lean_object* x_134; +x_132 = lean_ctor_get(x_74, 0); +x_133 = lean_ctor_get(x_74, 1); +lean_inc(x_133); lean_inc(x_132); -lean_inc(x_131); -lean_dec(x_73); -x_133 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_133, 0, x_131); -lean_ctor_set(x_133, 1, x_132); -return x_133; +lean_dec(x_74); +x_134 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_134, 0, x_132); +lean_ctor_set(x_134, 1, x_133); +return x_134; } } } else { -uint8_t x_134; -x_134 = !lean_is_exclusive(x_68); -if (x_134 == 0) +uint8_t x_135; +x_135 = !lean_is_exclusive(x_69); +if (x_135 == 0) { -return x_68; +return x_69; } else { -lean_object* x_135; lean_object* x_136; lean_object* x_137; -x_135 = lean_ctor_get(x_68, 0); -x_136 = lean_ctor_get(x_68, 1); +lean_object* x_136; lean_object* x_137; lean_object* x_138; +x_136 = lean_ctor_get(x_69, 0); +x_137 = lean_ctor_get(x_69, 1); +lean_inc(x_137); lean_inc(x_136); -lean_inc(x_135); -lean_dec(x_68); -x_137 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_137, 0, x_135); -lean_ctor_set(x_137, 1, x_136); -return x_137; +lean_dec(x_69); +x_138 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_138, 0, x_136); +lean_ctor_set(x_138, 1, x_137); +return x_138; } } } else { -uint8_t x_138; -x_138 = !lean_is_exclusive(x_65); -if (x_138 == 0) +uint8_t x_139; +x_139 = !lean_is_exclusive(x_66); +if (x_139 == 0) { -return x_65; +return x_66; } else { -lean_object* x_139; lean_object* x_140; lean_object* x_141; -x_139 = lean_ctor_get(x_65, 0); -x_140 = lean_ctor_get(x_65, 1); +lean_object* x_140; lean_object* x_141; lean_object* x_142; +x_140 = lean_ctor_get(x_66, 0); +x_141 = lean_ctor_get(x_66, 1); +lean_inc(x_141); lean_inc(x_140); -lean_inc(x_139); -lean_dec(x_65); -x_141 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_141, 0, x_139); -lean_ctor_set(x_141, 1, x_140); -return x_141; +lean_dec(x_66); +x_142 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_142, 0, x_140); +lean_ctor_set(x_142, 1, x_141); +return x_142; } } } else { -uint8_t x_142; -x_142 = !lean_is_exclusive(x_62); -if (x_142 == 0) +uint8_t x_143; +x_143 = !lean_is_exclusive(x_63); +if (x_143 == 0) { -return x_62; +return x_63; } else { -lean_object* x_143; lean_object* x_144; lean_object* x_145; -x_143 = lean_ctor_get(x_62, 0); -x_144 = lean_ctor_get(x_62, 1); +lean_object* x_144; lean_object* x_145; lean_object* x_146; +x_144 = lean_ctor_get(x_63, 0); +x_145 = lean_ctor_get(x_63, 1); +lean_inc(x_145); lean_inc(x_144); -lean_inc(x_143); -lean_dec(x_62); -x_145 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_145, 0, x_143); -lean_ctor_set(x_145, 1, x_144); -return x_145; +lean_dec(x_63); +x_146 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_146, 0, x_144); +lean_ctor_set(x_146, 1, x_145); +return x_146; } } } else { -uint8_t x_146; -x_146 = !lean_is_exclusive(x_57); -if (x_146 == 0) +uint8_t x_147; +x_147 = !lean_is_exclusive(x_58); +if (x_147 == 0) { -return x_57; +return x_58; } else { -lean_object* x_147; lean_object* x_148; lean_object* x_149; -x_147 = lean_ctor_get(x_57, 0); -x_148 = lean_ctor_get(x_57, 1); +lean_object* x_148; lean_object* x_149; lean_object* x_150; +x_148 = lean_ctor_get(x_58, 0); +x_149 = lean_ctor_get(x_58, 1); +lean_inc(x_149); lean_inc(x_148); -lean_inc(x_147); -lean_dec(x_57); -x_149 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_149, 0, x_147); -lean_ctor_set(x_149, 1, x_148); -return x_149; +lean_dec(x_58); +x_150 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_150, 0, x_148); +lean_ctor_set(x_150, 1, x_149); +return x_150; } } } else { -uint8_t x_150; -x_150 = !lean_is_exclusive(x_54); -if (x_150 == 0) +uint8_t x_151; +x_151 = !lean_is_exclusive(x_55); +if (x_151 == 0) { -return x_54; +return x_55; } else { -lean_object* x_151; lean_object* x_152; lean_object* x_153; -x_151 = lean_ctor_get(x_54, 0); -x_152 = lean_ctor_get(x_54, 1); +lean_object* x_152; lean_object* x_153; lean_object* x_154; +x_152 = lean_ctor_get(x_55, 0); +x_153 = lean_ctor_get(x_55, 1); +lean_inc(x_153); lean_inc(x_152); -lean_inc(x_151); -lean_dec(x_54); -x_153 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_153, 0, x_151); -lean_ctor_set(x_153, 1, x_152); -return x_153; +lean_dec(x_55); +x_154 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_154, 0, x_152); +lean_ctor_set(x_154, 1, x_153); +return x_154; } } } else { -uint8_t x_154; -x_154 = !lean_is_exclusive(x_51); -if (x_154 == 0) +uint8_t x_155; +x_155 = !lean_is_exclusive(x_52); +if (x_155 == 0) { -return x_51; +return x_52; } else { -lean_object* x_155; lean_object* x_156; lean_object* x_157; -x_155 = lean_ctor_get(x_51, 0); -x_156 = lean_ctor_get(x_51, 1); +lean_object* x_156; lean_object* x_157; lean_object* x_158; +x_156 = lean_ctor_get(x_52, 0); +x_157 = lean_ctor_get(x_52, 1); +lean_inc(x_157); lean_inc(x_156); -lean_inc(x_155); -lean_dec(x_51); -x_157 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_157, 0, x_155); -lean_ctor_set(x_157, 1, x_156); -return x_157; +lean_dec(x_52); +x_158 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_158, 0, x_156); +lean_ctor_set(x_158, 1, x_157); +return x_158; } } } else { -uint8_t x_158; -x_158 = !lean_is_exclusive(x_46); -if (x_158 == 0) +uint8_t x_159; +x_159 = !lean_is_exclusive(x_47); +if (x_159 == 0) { -return x_46; +return x_47; } else { -lean_object* x_159; lean_object* x_160; lean_object* x_161; -x_159 = lean_ctor_get(x_46, 0); -x_160 = lean_ctor_get(x_46, 1); +lean_object* x_160; lean_object* x_161; lean_object* x_162; +x_160 = lean_ctor_get(x_47, 0); +x_161 = lean_ctor_get(x_47, 1); +lean_inc(x_161); lean_inc(x_160); -lean_inc(x_159); -lean_dec(x_46); -x_161 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_161, 0, x_159); -lean_ctor_set(x_161, 1, x_160); -return x_161; +lean_dec(x_47); +x_162 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_162, 0, x_160); +lean_ctor_set(x_162, 1, x_161); +return x_162; } } } else { -uint8_t x_162; -x_162 = !lean_is_exclusive(x_43); -if (x_162 == 0) +uint8_t x_163; +x_163 = !lean_is_exclusive(x_44); +if (x_163 == 0) { -return x_43; +return x_44; } else { -lean_object* x_163; lean_object* x_164; lean_object* x_165; -x_163 = lean_ctor_get(x_43, 0); -x_164 = lean_ctor_get(x_43, 1); +lean_object* x_164; lean_object* x_165; lean_object* x_166; +x_164 = lean_ctor_get(x_44, 0); +x_165 = lean_ctor_get(x_44, 1); +lean_inc(x_165); lean_inc(x_164); -lean_inc(x_163); -lean_dec(x_43); -x_165 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_165, 0, x_163); -lean_ctor_set(x_165, 1, x_164); -return x_165; +lean_dec(x_44); +x_166 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_166, 0, x_164); +lean_ctor_set(x_166, 1, x_165); +return x_166; } } } else { -uint8_t x_166; -x_166 = !lean_is_exclusive(x_40); -if (x_166 == 0) +uint8_t x_167; +x_167 = !lean_is_exclusive(x_41); +if (x_167 == 0) { -return x_40; +return x_41; } else { -lean_object* x_167; lean_object* x_168; lean_object* x_169; -x_167 = lean_ctor_get(x_40, 0); -x_168 = lean_ctor_get(x_40, 1); +lean_object* x_168; lean_object* x_169; lean_object* x_170; +x_168 = lean_ctor_get(x_41, 0); +x_169 = lean_ctor_get(x_41, 1); +lean_inc(x_169); lean_inc(x_168); -lean_inc(x_167); -lean_dec(x_40); -x_169 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_169, 0, x_167); -lean_ctor_set(x_169, 1, x_168); -return x_169; +lean_dec(x_41); +x_170 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_170, 0, x_168); +lean_ctor_set(x_170, 1, x_169); +return x_170; } } } else { -uint8_t x_170; -x_170 = !lean_is_exclusive(x_35); -if (x_170 == 0) +uint8_t x_171; +x_171 = !lean_is_exclusive(x_36); +if (x_171 == 0) { -return x_35; +return x_36; } else { -lean_object* x_171; lean_object* x_172; lean_object* x_173; -x_171 = lean_ctor_get(x_35, 0); -x_172 = lean_ctor_get(x_35, 1); +lean_object* x_172; lean_object* x_173; lean_object* x_174; +x_172 = lean_ctor_get(x_36, 0); +x_173 = lean_ctor_get(x_36, 1); +lean_inc(x_173); lean_inc(x_172); -lean_inc(x_171); -lean_dec(x_35); -x_173 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_173, 0, x_171); -lean_ctor_set(x_173, 1, x_172); -return x_173; +lean_dec(x_36); +x_174 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_174, 0, x_172); +lean_ctor_set(x_174, 1, x_173); +return x_174; } } } else { -uint8_t x_174; -x_174 = !lean_is_exclusive(x_32); -if (x_174 == 0) +uint8_t x_175; +x_175 = !lean_is_exclusive(x_33); +if (x_175 == 0) { -return x_32; +return x_33; } else { -lean_object* x_175; lean_object* x_176; lean_object* x_177; -x_175 = lean_ctor_get(x_32, 0); -x_176 = lean_ctor_get(x_32, 1); +lean_object* x_176; lean_object* x_177; lean_object* x_178; +x_176 = lean_ctor_get(x_33, 0); +x_177 = lean_ctor_get(x_33, 1); +lean_inc(x_177); lean_inc(x_176); -lean_inc(x_175); -lean_dec(x_32); -x_177 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_177, 0, x_175); -lean_ctor_set(x_177, 1, x_176); -return x_177; +lean_dec(x_33); +x_178 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_178, 0, x_176); +lean_ctor_set(x_178, 1, x_177); +return x_178; } } } else { -uint8_t x_178; -x_178 = !lean_is_exclusive(x_29); -if (x_178 == 0) +uint8_t x_179; +x_179 = !lean_is_exclusive(x_30); +if (x_179 == 0) { -return x_29; +return x_30; } else { -lean_object* x_179; lean_object* x_180; lean_object* x_181; -x_179 = lean_ctor_get(x_29, 0); -x_180 = lean_ctor_get(x_29, 1); +lean_object* x_180; lean_object* x_181; lean_object* x_182; +x_180 = lean_ctor_get(x_30, 0); +x_181 = lean_ctor_get(x_30, 1); +lean_inc(x_181); lean_inc(x_180); -lean_inc(x_179); -lean_dec(x_29); -x_181 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_181, 0, x_179); -lean_ctor_set(x_181, 1, x_180); -return x_181; +lean_dec(x_30); +x_182 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_182, 0, x_180); +lean_ctor_set(x_182, 1, x_181); +return x_182; } } } else { -uint8_t x_182; -x_182 = !lean_is_exclusive(x_24); -if (x_182 == 0) +uint8_t x_183; +x_183 = !lean_is_exclusive(x_25); +if (x_183 == 0) { -return x_24; +return x_25; } else { -lean_object* x_183; lean_object* x_184; lean_object* x_185; -x_183 = lean_ctor_get(x_24, 0); -x_184 = lean_ctor_get(x_24, 1); +lean_object* x_184; lean_object* x_185; lean_object* x_186; +x_184 = lean_ctor_get(x_25, 0); +x_185 = lean_ctor_get(x_25, 1); +lean_inc(x_185); lean_inc(x_184); -lean_inc(x_183); -lean_dec(x_24); -x_185 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_185, 0, x_183); -lean_ctor_set(x_185, 1, x_184); -return x_185; +lean_dec(x_25); +x_186 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_186, 0, x_184); +lean_ctor_set(x_186, 1, x_185); +return x_186; } } } else { -uint8_t x_186; -x_186 = !lean_is_exclusive(x_21); -if (x_186 == 0) +uint8_t x_187; +x_187 = !lean_is_exclusive(x_22); +if (x_187 == 0) { -return x_21; +return x_22; } else { -lean_object* x_187; lean_object* x_188; lean_object* x_189; -x_187 = lean_ctor_get(x_21, 0); -x_188 = lean_ctor_get(x_21, 1); +lean_object* x_188; lean_object* x_189; lean_object* x_190; +x_188 = lean_ctor_get(x_22, 0); +x_189 = lean_ctor_get(x_22, 1); +lean_inc(x_189); lean_inc(x_188); -lean_inc(x_187); -lean_dec(x_21); -x_189 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_189, 0, x_187); -lean_ctor_set(x_189, 1, x_188); -return x_189; +lean_dec(x_22); +x_190 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_190, 0, x_188); +lean_ctor_set(x_190, 1, x_189); +return x_190; } } } else { -uint8_t x_190; -x_190 = !lean_is_exclusive(x_18); -if (x_190 == 0) +uint8_t x_191; +x_191 = !lean_is_exclusive(x_19); +if (x_191 == 0) { -return x_18; +return x_19; } else { -lean_object* x_191; lean_object* x_192; lean_object* x_193; -x_191 = lean_ctor_get(x_18, 0); -x_192 = lean_ctor_get(x_18, 1); +lean_object* x_192; lean_object* x_193; lean_object* x_194; +x_192 = lean_ctor_get(x_19, 0); +x_193 = lean_ctor_get(x_19, 1); +lean_inc(x_193); lean_inc(x_192); -lean_inc(x_191); -lean_dec(x_18); -x_193 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_193, 0, x_191); -lean_ctor_set(x_193, 1, x_192); -return x_193; +lean_dec(x_19); +x_194 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_194, 0, x_192); +lean_ctor_set(x_194, 1, x_193); +return x_194; } } } else { -uint8_t x_194; -x_194 = !lean_is_exclusive(x_13); -if (x_194 == 0) +uint8_t x_195; +x_195 = !lean_is_exclusive(x_14); +if (x_195 == 0) { -return x_13; +return x_14; } else { -lean_object* x_195; lean_object* x_196; lean_object* x_197; -x_195 = lean_ctor_get(x_13, 0); -x_196 = lean_ctor_get(x_13, 1); +lean_object* x_196; lean_object* x_197; lean_object* x_198; +x_196 = lean_ctor_get(x_14, 0); +x_197 = lean_ctor_get(x_14, 1); +lean_inc(x_197); lean_inc(x_196); -lean_inc(x_195); -lean_dec(x_13); -x_197 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_197, 0, x_195); -lean_ctor_set(x_197, 1, x_196); -return x_197; +lean_dec(x_14); +x_198 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_198, 0, x_196); +lean_ctor_set(x_198, 1, x_197); +return x_198; } } } else { -uint8_t x_198; -x_198 = !lean_is_exclusive(x_9); -if (x_198 == 0) +uint8_t x_199; +x_199 = !lean_is_exclusive(x_10); +if (x_199 == 0) { -return x_9; +return x_10; } else { -lean_object* x_199; lean_object* x_200; lean_object* x_201; -x_199 = lean_ctor_get(x_9, 0); -x_200 = lean_ctor_get(x_9, 1); +lean_object* x_200; lean_object* x_201; lean_object* x_202; +x_200 = lean_ctor_get(x_10, 0); +x_201 = lean_ctor_get(x_10, 1); +lean_inc(x_201); lean_inc(x_200); -lean_inc(x_199); -lean_dec(x_9); -x_201 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_201, 0, x_199); -lean_ctor_set(x_201, 1, x_200); -return x_201; +lean_dec(x_10); +x_202 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_202, 0, x_200); +lean_ctor_set(x_202, 1, x_201); +return x_202; } } } else { -uint8_t x_202; -x_202 = !lean_is_exclusive(x_5); -if (x_202 == 0) +uint8_t x_203; +x_203 = !lean_is_exclusive(x_6); +if (x_203 == 0) { -return x_5; +return x_6; } else { -lean_object* x_203; lean_object* x_204; lean_object* x_205; -x_203 = lean_ctor_get(x_5, 0); -x_204 = lean_ctor_get(x_5, 1); +lean_object* x_204; lean_object* x_205; lean_object* x_206; +x_204 = lean_ctor_get(x_6, 0); +x_205 = lean_ctor_get(x_6, 1); +lean_inc(x_205); lean_inc(x_204); -lean_inc(x_203); -lean_dec(x_5); -x_205 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_205, 0, x_203); -lean_ctor_set(x_205, 1, x_204); -return x_205; +lean_dec(x_6); +x_206 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_206, 0, x_204); +lean_ctor_set(x_206, 1, x_205); +return x_206; } } } @@ -97402,6 +98004,10 @@ l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__24 = _init_l_Lean lean_mark_persistent(l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__24); l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__25 = _init_l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__25(); lean_mark_persistent(l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__25); +l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__26 = _init_l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__26(); +lean_mark_persistent(l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__26); +l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__27 = _init_l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__27(); +lean_mark_persistent(l_Lean_Parser_Tactic_tacticSeq1Indented___elambda__1___closed__27); l_Lean_Parser_Tactic_tacticSeq1Indented___closed__1 = _init_l_Lean_Parser_Tactic_tacticSeq1Indented___closed__1(); lean_mark_persistent(l_Lean_Parser_Tactic_tacticSeq1Indented___closed__1); l_Lean_Parser_Tactic_tacticSeq1Indented___closed__2 = _init_l_Lean_Parser_Tactic_tacticSeq1Indented___closed__2(); @@ -100504,6 +101110,10 @@ l_Lean_Parser_Term_binderIdent___closed__2 = _init_l_Lean_Parser_Term_binderIden lean_mark_persistent(l_Lean_Parser_Term_binderIdent___closed__2); l_Lean_Parser_Term_binderIdent = _init_l_Lean_Parser_Term_binderIdent(); lean_mark_persistent(l_Lean_Parser_Term_binderIdent); +l_Lean_Parser_Term_binderType___elambda__1___closed__1 = _init_l_Lean_Parser_Term_binderType___elambda__1___closed__1(); +lean_mark_persistent(l_Lean_Parser_Term_binderType___elambda__1___closed__1); +l_Lean_Parser_Term_binderType___elambda__1___closed__2 = _init_l_Lean_Parser_Term_binderType___elambda__1___closed__2(); +lean_mark_persistent(l_Lean_Parser_Term_binderType___elambda__1___closed__2); l_Lean_Parser_Term_binderType___closed__1 = _init_l_Lean_Parser_Term_binderType___closed__1(); lean_mark_persistent(l_Lean_Parser_Term_binderType___closed__1); l_Lean_Parser_Term_binderType___closed__2 = _init_l_Lean_Parser_Term_binderType___closed__2(); @@ -101920,6 +102530,10 @@ l_Lean_Parser_Term_funImplicitBinder___elambda__1___closed__4 = _init_l_Lean_Par lean_mark_persistent(l_Lean_Parser_Term_funImplicitBinder___elambda__1___closed__4); l_Lean_Parser_Term_funImplicitBinder___elambda__1___closed__5 = _init_l_Lean_Parser_Term_funImplicitBinder___elambda__1___closed__5(); lean_mark_persistent(l_Lean_Parser_Term_funImplicitBinder___elambda__1___closed__5); +l_Lean_Parser_Term_funImplicitBinder___elambda__1___closed__6 = _init_l_Lean_Parser_Term_funImplicitBinder___elambda__1___closed__6(); +lean_mark_persistent(l_Lean_Parser_Term_funImplicitBinder___elambda__1___closed__6); +l_Lean_Parser_Term_funImplicitBinder___elambda__1___closed__7 = _init_l_Lean_Parser_Term_funImplicitBinder___elambda__1___closed__7(); +lean_mark_persistent(l_Lean_Parser_Term_funImplicitBinder___elambda__1___closed__7); l_Lean_Parser_Term_funImplicitBinder___closed__1 = _init_l_Lean_Parser_Term_funImplicitBinder___closed__1(); lean_mark_persistent(l_Lean_Parser_Term_funImplicitBinder___closed__1); l_Lean_Parser_Term_funImplicitBinder___closed__2 = _init_l_Lean_Parser_Term_funImplicitBinder___closed__2(); @@ -101932,6 +102546,8 @@ l_Lean_Parser_Term_funImplicitBinder___closed__5 = _init_l_Lean_Parser_Term_funI lean_mark_persistent(l_Lean_Parser_Term_funImplicitBinder___closed__5); l_Lean_Parser_Term_funImplicitBinder___closed__6 = _init_l_Lean_Parser_Term_funImplicitBinder___closed__6(); lean_mark_persistent(l_Lean_Parser_Term_funImplicitBinder___closed__6); +l_Lean_Parser_Term_funImplicitBinder___closed__7 = _init_l_Lean_Parser_Term_funImplicitBinder___closed__7(); +lean_mark_persistent(l_Lean_Parser_Term_funImplicitBinder___closed__7); l_Lean_Parser_Term_funImplicitBinder = _init_l_Lean_Parser_Term_funImplicitBinder(); lean_mark_persistent(l_Lean_Parser_Term_funImplicitBinder); l_Lean_Parser_Term_funStrictImplicitBinder___elambda__1___closed__1 = _init_l_Lean_Parser_Term_funStrictImplicitBinder___elambda__1___closed__1(); @@ -101962,6 +102578,10 @@ l_Lean_Parser_Term_funSimpleBinder___elambda__1___closed__1 = _init_l_Lean_Parse lean_mark_persistent(l_Lean_Parser_Term_funSimpleBinder___elambda__1___closed__1); l_Lean_Parser_Term_funSimpleBinder___elambda__1___closed__2 = _init_l_Lean_Parser_Term_funSimpleBinder___elambda__1___closed__2(); lean_mark_persistent(l_Lean_Parser_Term_funSimpleBinder___elambda__1___closed__2); +l_Lean_Parser_Term_funSimpleBinder___elambda__1___closed__3 = _init_l_Lean_Parser_Term_funSimpleBinder___elambda__1___closed__3(); +lean_mark_persistent(l_Lean_Parser_Term_funSimpleBinder___elambda__1___closed__3); +l_Lean_Parser_Term_funSimpleBinder___elambda__1___closed__4 = _init_l_Lean_Parser_Term_funSimpleBinder___elambda__1___closed__4(); +lean_mark_persistent(l_Lean_Parser_Term_funSimpleBinder___elambda__1___closed__4); l_Lean_Parser_Term_funSimpleBinder___closed__1 = _init_l_Lean_Parser_Term_funSimpleBinder___closed__1(); lean_mark_persistent(l_Lean_Parser_Term_funSimpleBinder___closed__1); l_Lean_Parser_Term_funSimpleBinder___closed__2 = _init_l_Lean_Parser_Term_funSimpleBinder___closed__2(); @@ -101970,6 +102590,8 @@ l_Lean_Parser_Term_funSimpleBinder___closed__3 = _init_l_Lean_Parser_Term_funSim lean_mark_persistent(l_Lean_Parser_Term_funSimpleBinder___closed__3); l_Lean_Parser_Term_funSimpleBinder___closed__4 = _init_l_Lean_Parser_Term_funSimpleBinder___closed__4(); lean_mark_persistent(l_Lean_Parser_Term_funSimpleBinder___closed__4); +l_Lean_Parser_Term_funSimpleBinder___closed__5 = _init_l_Lean_Parser_Term_funSimpleBinder___closed__5(); +lean_mark_persistent(l_Lean_Parser_Term_funSimpleBinder___closed__5); l_Lean_Parser_Term_funSimpleBinder = _init_l_Lean_Parser_Term_funSimpleBinder(); lean_mark_persistent(l_Lean_Parser_Term_funSimpleBinder); l_Lean_Parser_Term_funBinder___elambda__1___closed__1 = _init_l_Lean_Parser_Term_funBinder___elambda__1___closed__1(); @@ -101978,6 +102600,14 @@ l_Lean_Parser_Term_funBinder___elambda__1___closed__2 = _init_l_Lean_Parser_Term lean_mark_persistent(l_Lean_Parser_Term_funBinder___elambda__1___closed__2); l_Lean_Parser_Term_funBinder___elambda__1___closed__3 = _init_l_Lean_Parser_Term_funBinder___elambda__1___closed__3(); lean_mark_persistent(l_Lean_Parser_Term_funBinder___elambda__1___closed__3); +l_Lean_Parser_Term_funBinder___elambda__1___closed__4 = _init_l_Lean_Parser_Term_funBinder___elambda__1___closed__4(); +lean_mark_persistent(l_Lean_Parser_Term_funBinder___elambda__1___closed__4); +l_Lean_Parser_Term_funBinder___elambda__1___closed__5 = _init_l_Lean_Parser_Term_funBinder___elambda__1___closed__5(); +lean_mark_persistent(l_Lean_Parser_Term_funBinder___elambda__1___closed__5); +l_Lean_Parser_Term_funBinder___elambda__1___closed__6 = _init_l_Lean_Parser_Term_funBinder___elambda__1___closed__6(); +lean_mark_persistent(l_Lean_Parser_Term_funBinder___elambda__1___closed__6); +l_Lean_Parser_Term_funBinder___elambda__1___closed__7 = _init_l_Lean_Parser_Term_funBinder___elambda__1___closed__7(); +lean_mark_persistent(l_Lean_Parser_Term_funBinder___elambda__1___closed__7); l_Lean_Parser_Term_funBinder___closed__1 = _init_l_Lean_Parser_Term_funBinder___closed__1(); lean_mark_persistent(l_Lean_Parser_Term_funBinder___closed__1); l_Lean_Parser_Term_funBinder___closed__2 = _init_l_Lean_Parser_Term_funBinder___closed__2(); @@ -101990,6 +102620,8 @@ l_Lean_Parser_Term_funBinder___closed__5 = _init_l_Lean_Parser_Term_funBinder___ lean_mark_persistent(l_Lean_Parser_Term_funBinder___closed__5); l_Lean_Parser_Term_funBinder___closed__6 = _init_l_Lean_Parser_Term_funBinder___closed__6(); lean_mark_persistent(l_Lean_Parser_Term_funBinder___closed__6); +l_Lean_Parser_Term_funBinder___closed__7 = _init_l_Lean_Parser_Term_funBinder___closed__7(); +lean_mark_persistent(l_Lean_Parser_Term_funBinder___closed__7); l_Lean_Parser_Term_funBinder = _init_l_Lean_Parser_Term_funBinder(); lean_mark_persistent(l_Lean_Parser_Term_funBinder); l_Lean_Parser_Term_basicFun___elambda__1___closed__1 = _init_l_Lean_Parser_Term_basicFun___elambda__1___closed__1(); @@ -102156,12 +102788,16 @@ l_Lean_Parser_Term_funImplicitBinder_formatter___closed__5 = _init_l_Lean_Parser lean_mark_persistent(l_Lean_Parser_Term_funImplicitBinder_formatter___closed__5); l_Lean_Parser_Term_funImplicitBinder_formatter___closed__6 = _init_l_Lean_Parser_Term_funImplicitBinder_formatter___closed__6(); lean_mark_persistent(l_Lean_Parser_Term_funImplicitBinder_formatter___closed__6); +l_Lean_Parser_Term_funImplicitBinder_formatter___closed__7 = _init_l_Lean_Parser_Term_funImplicitBinder_formatter___closed__7(); +lean_mark_persistent(l_Lean_Parser_Term_funImplicitBinder_formatter___closed__7); l_Lean_Parser_Term_funSimpleBinder_formatter___closed__1 = _init_l_Lean_Parser_Term_funSimpleBinder_formatter___closed__1(); lean_mark_persistent(l_Lean_Parser_Term_funSimpleBinder_formatter___closed__1); l_Lean_Parser_Term_funSimpleBinder_formatter___closed__2 = _init_l_Lean_Parser_Term_funSimpleBinder_formatter___closed__2(); lean_mark_persistent(l_Lean_Parser_Term_funSimpleBinder_formatter___closed__2); l_Lean_Parser_Term_funSimpleBinder_formatter___closed__3 = _init_l_Lean_Parser_Term_funSimpleBinder_formatter___closed__3(); lean_mark_persistent(l_Lean_Parser_Term_funSimpleBinder_formatter___closed__3); +l_Lean_Parser_Term_funSimpleBinder_formatter___closed__4 = _init_l_Lean_Parser_Term_funSimpleBinder_formatter___closed__4(); +lean_mark_persistent(l_Lean_Parser_Term_funSimpleBinder_formatter___closed__4); l_Lean_Parser_Term_funBinder_formatter___closed__1 = _init_l_Lean_Parser_Term_funBinder_formatter___closed__1(); lean_mark_persistent(l_Lean_Parser_Term_funBinder_formatter___closed__1); l_Lean_Parser_Term_funBinder_formatter___closed__2 = _init_l_Lean_Parser_Term_funBinder_formatter___closed__2(); @@ -102174,6 +102810,10 @@ l_Lean_Parser_Term_funBinder_formatter___closed__5 = _init_l_Lean_Parser_Term_fu lean_mark_persistent(l_Lean_Parser_Term_funBinder_formatter___closed__5); l_Lean_Parser_Term_funBinder_formatter___closed__6 = _init_l_Lean_Parser_Term_funBinder_formatter___closed__6(); lean_mark_persistent(l_Lean_Parser_Term_funBinder_formatter___closed__6); +l_Lean_Parser_Term_funBinder_formatter___closed__7 = _init_l_Lean_Parser_Term_funBinder_formatter___closed__7(); +lean_mark_persistent(l_Lean_Parser_Term_funBinder_formatter___closed__7); +l_Lean_Parser_Term_funBinder_formatter___closed__8 = _init_l_Lean_Parser_Term_funBinder_formatter___closed__8(); +lean_mark_persistent(l_Lean_Parser_Term_funBinder_formatter___closed__8); l_Lean_Parser_Term_basicFun_formatter___closed__1 = _init_l_Lean_Parser_Term_basicFun_formatter___closed__1(); lean_mark_persistent(l_Lean_Parser_Term_basicFun_formatter___closed__1); l_Lean_Parser_Term_basicFun_formatter___closed__2 = _init_l_Lean_Parser_Term_basicFun_formatter___closed__2(); @@ -102244,12 +102884,16 @@ l_Lean_Parser_Term_funImplicitBinder_parenthesizer___closed__5 = _init_l_Lean_Pa lean_mark_persistent(l_Lean_Parser_Term_funImplicitBinder_parenthesizer___closed__5); l_Lean_Parser_Term_funImplicitBinder_parenthesizer___closed__6 = _init_l_Lean_Parser_Term_funImplicitBinder_parenthesizer___closed__6(); lean_mark_persistent(l_Lean_Parser_Term_funImplicitBinder_parenthesizer___closed__6); +l_Lean_Parser_Term_funImplicitBinder_parenthesizer___closed__7 = _init_l_Lean_Parser_Term_funImplicitBinder_parenthesizer___closed__7(); +lean_mark_persistent(l_Lean_Parser_Term_funImplicitBinder_parenthesizer___closed__7); l_Lean_Parser_Term_funSimpleBinder_parenthesizer___closed__1 = _init_l_Lean_Parser_Term_funSimpleBinder_parenthesizer___closed__1(); lean_mark_persistent(l_Lean_Parser_Term_funSimpleBinder_parenthesizer___closed__1); l_Lean_Parser_Term_funSimpleBinder_parenthesizer___closed__2 = _init_l_Lean_Parser_Term_funSimpleBinder_parenthesizer___closed__2(); lean_mark_persistent(l_Lean_Parser_Term_funSimpleBinder_parenthesizer___closed__2); l_Lean_Parser_Term_funSimpleBinder_parenthesizer___closed__3 = _init_l_Lean_Parser_Term_funSimpleBinder_parenthesizer___closed__3(); lean_mark_persistent(l_Lean_Parser_Term_funSimpleBinder_parenthesizer___closed__3); +l_Lean_Parser_Term_funSimpleBinder_parenthesizer___closed__4 = _init_l_Lean_Parser_Term_funSimpleBinder_parenthesizer___closed__4(); +lean_mark_persistent(l_Lean_Parser_Term_funSimpleBinder_parenthesizer___closed__4); l_Lean_Parser_Term_funBinder_parenthesizer___closed__1 = _init_l_Lean_Parser_Term_funBinder_parenthesizer___closed__1(); lean_mark_persistent(l_Lean_Parser_Term_funBinder_parenthesizer___closed__1); l_Lean_Parser_Term_funBinder_parenthesizer___closed__2 = _init_l_Lean_Parser_Term_funBinder_parenthesizer___closed__2(); @@ -102262,6 +102906,10 @@ l_Lean_Parser_Term_funBinder_parenthesizer___closed__5 = _init_l_Lean_Parser_Ter lean_mark_persistent(l_Lean_Parser_Term_funBinder_parenthesizer___closed__5); l_Lean_Parser_Term_funBinder_parenthesizer___closed__6 = _init_l_Lean_Parser_Term_funBinder_parenthesizer___closed__6(); lean_mark_persistent(l_Lean_Parser_Term_funBinder_parenthesizer___closed__6); +l_Lean_Parser_Term_funBinder_parenthesizer___closed__7 = _init_l_Lean_Parser_Term_funBinder_parenthesizer___closed__7(); +lean_mark_persistent(l_Lean_Parser_Term_funBinder_parenthesizer___closed__7); +l_Lean_Parser_Term_funBinder_parenthesizer___closed__8 = _init_l_Lean_Parser_Term_funBinder_parenthesizer___closed__8(); +lean_mark_persistent(l_Lean_Parser_Term_funBinder_parenthesizer___closed__8); l_Lean_Parser_Term_basicFun_parenthesizer___closed__1 = _init_l_Lean_Parser_Term_basicFun_parenthesizer___closed__1(); lean_mark_persistent(l_Lean_Parser_Term_basicFun_parenthesizer___closed__1); l_Lean_Parser_Term_basicFun_parenthesizer___closed__2 = _init_l_Lean_Parser_Term_basicFun_parenthesizer___closed__2(); @@ -102966,6 +103614,24 @@ l_Lean_Parser_Term_simpleBinderWithoutType___closed__4 = _init_l_Lean_Parser_Ter lean_mark_persistent(l_Lean_Parser_Term_simpleBinderWithoutType___closed__4); l_Lean_Parser_Term_simpleBinderWithoutType = _init_l_Lean_Parser_Term_simpleBinderWithoutType(); lean_mark_persistent(l_Lean_Parser_Term_simpleBinderWithoutType); +l_Lean_Parser_Term_letIdBinder___elambda__1___closed__1 = _init_l_Lean_Parser_Term_letIdBinder___elambda__1___closed__1(); +lean_mark_persistent(l_Lean_Parser_Term_letIdBinder___elambda__1___closed__1); +l_Lean_Parser_Term_letIdBinder___elambda__1___closed__2 = _init_l_Lean_Parser_Term_letIdBinder___elambda__1___closed__2(); +lean_mark_persistent(l_Lean_Parser_Term_letIdBinder___elambda__1___closed__2); +l_Lean_Parser_Term_letIdBinder___elambda__1___closed__3 = _init_l_Lean_Parser_Term_letIdBinder___elambda__1___closed__3(); +lean_mark_persistent(l_Lean_Parser_Term_letIdBinder___elambda__1___closed__3); +l_Lean_Parser_Term_letIdBinder___elambda__1___closed__4 = _init_l_Lean_Parser_Term_letIdBinder___elambda__1___closed__4(); +lean_mark_persistent(l_Lean_Parser_Term_letIdBinder___elambda__1___closed__4); +l_Lean_Parser_Term_letIdBinder___closed__1 = _init_l_Lean_Parser_Term_letIdBinder___closed__1(); +lean_mark_persistent(l_Lean_Parser_Term_letIdBinder___closed__1); +l_Lean_Parser_Term_letIdBinder___closed__2 = _init_l_Lean_Parser_Term_letIdBinder___closed__2(); +lean_mark_persistent(l_Lean_Parser_Term_letIdBinder___closed__2); +l_Lean_Parser_Term_letIdBinder___closed__3 = _init_l_Lean_Parser_Term_letIdBinder___closed__3(); +lean_mark_persistent(l_Lean_Parser_Term_letIdBinder___closed__3); +l_Lean_Parser_Term_letIdBinder___closed__4 = _init_l_Lean_Parser_Term_letIdBinder___closed__4(); +lean_mark_persistent(l_Lean_Parser_Term_letIdBinder___closed__4); +l_Lean_Parser_Term_letIdBinder = _init_l_Lean_Parser_Term_letIdBinder(); +lean_mark_persistent(l_Lean_Parser_Term_letIdBinder); l_Lean_Parser_Term_letIdLhs___elambda__1___closed__1 = _init_l_Lean_Parser_Term_letIdLhs___elambda__1___closed__1(); lean_mark_persistent(l_Lean_Parser_Term_letIdLhs___elambda__1___closed__1); l_Lean_Parser_Term_letIdLhs___elambda__1___closed__2 = _init_l_Lean_Parser_Term_letIdLhs___elambda__1___closed__2(); @@ -102980,10 +103646,6 @@ l_Lean_Parser_Term_letIdLhs___elambda__1___closed__6 = _init_l_Lean_Parser_Term_ lean_mark_persistent(l_Lean_Parser_Term_letIdLhs___elambda__1___closed__6); l_Lean_Parser_Term_letIdLhs___elambda__1___closed__7 = _init_l_Lean_Parser_Term_letIdLhs___elambda__1___closed__7(); lean_mark_persistent(l_Lean_Parser_Term_letIdLhs___elambda__1___closed__7); -l_Lean_Parser_Term_letIdLhs___elambda__1___closed__8 = _init_l_Lean_Parser_Term_letIdLhs___elambda__1___closed__8(); -lean_mark_persistent(l_Lean_Parser_Term_letIdLhs___elambda__1___closed__8); -l_Lean_Parser_Term_letIdLhs___elambda__1___closed__9 = _init_l_Lean_Parser_Term_letIdLhs___elambda__1___closed__9(); -lean_mark_persistent(l_Lean_Parser_Term_letIdLhs___elambda__1___closed__9); l_Lean_Parser_Term_letIdLhs___closed__1 = _init_l_Lean_Parser_Term_letIdLhs___closed__1(); lean_mark_persistent(l_Lean_Parser_Term_letIdLhs___closed__1); l_Lean_Parser_Term_letIdLhs___closed__2 = _init_l_Lean_Parser_Term_letIdLhs___closed__2(); @@ -103223,6 +103885,10 @@ lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_simpleBinderWithoutType_for res = l___regBuiltin_Lean_Parser_Term_simpleBinderWithoutType_formatter(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +l_Lean_Parser_Term_letIdBinder_formatter___closed__1 = _init_l_Lean_Parser_Term_letIdBinder_formatter___closed__1(); +lean_mark_persistent(l_Lean_Parser_Term_letIdBinder_formatter___closed__1); +l_Lean_Parser_Term_letIdBinder_formatter___closed__2 = _init_l_Lean_Parser_Term_letIdBinder_formatter___closed__2(); +lean_mark_persistent(l_Lean_Parser_Term_letIdBinder_formatter___closed__2); l_Lean_Parser_Term_letIdLhs_formatter___closed__1 = _init_l_Lean_Parser_Term_letIdLhs_formatter___closed__1(); lean_mark_persistent(l_Lean_Parser_Term_letIdLhs_formatter___closed__1); l_Lean_Parser_Term_letIdLhs_formatter___closed__2 = _init_l_Lean_Parser_Term_letIdLhs_formatter___closed__2(); @@ -103341,6 +104007,10 @@ lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_simpleBinderWithoutType_par res = l___regBuiltin_Lean_Parser_Term_simpleBinderWithoutType_parenthesizer(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +l_Lean_Parser_Term_letIdBinder_parenthesizer___closed__1 = _init_l_Lean_Parser_Term_letIdBinder_parenthesizer___closed__1(); +lean_mark_persistent(l_Lean_Parser_Term_letIdBinder_parenthesizer___closed__1); +l_Lean_Parser_Term_letIdBinder_parenthesizer___closed__2 = _init_l_Lean_Parser_Term_letIdBinder_parenthesizer___closed__2(); +lean_mark_persistent(l_Lean_Parser_Term_letIdBinder_parenthesizer___closed__2); l_Lean_Parser_Term_letIdLhs_parenthesizer___closed__1 = _init_l_Lean_Parser_Term_letIdLhs_parenthesizer___closed__1(); lean_mark_persistent(l_Lean_Parser_Term_letIdLhs_parenthesizer___closed__1); l_Lean_Parser_Term_letIdLhs_parenthesizer___closed__2 = _init_l_Lean_Parser_Term_letIdLhs_parenthesizer___closed__2(); @@ -107447,10 +108117,6 @@ l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__15 = _init_l_Lean_Parse lean_mark_persistent(l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__15); l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__16 = _init_l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__16(); lean_mark_persistent(l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__16); -l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__17 = _init_l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__17(); -lean_mark_persistent(l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__17); -l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__18 = _init_l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__18(); -lean_mark_persistent(l_Lean_Parser_Term_funBinder_quot___elambda__1___closed__18); l_Lean_Parser_Term_funBinder_quot___closed__1 = _init_l_Lean_Parser_Term_funBinder_quot___closed__1(); lean_mark_persistent(l_Lean_Parser_Term_funBinder_quot___closed__1); l_Lean_Parser_Term_funBinder_quot___closed__2 = _init_l_Lean_Parser_Term_funBinder_quot___closed__2(); @@ -108977,101 +109643,105 @@ lean_mark_persistent(l___regBuiltin_Lean_Parser_Level_quot_parenthesizer___close res = l___regBuiltin_Lean_Parser_Level_quot_parenthesizer(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__1(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__1); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__2(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__2); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__3(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__3); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__4 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__4(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__4); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__5 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__5(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__5); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__6 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__6(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__6); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__7 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__7(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__7); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__8 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__8(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__8); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__9 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__9(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__9); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__10 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__10(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__10); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__11 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__11(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__11); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__12 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__12(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__12); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__13 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__13(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__13); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__14 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__14(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__14); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__15 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__15(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__15); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__16 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__16(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__16); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__17 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__17(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__17); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__18 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__18(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__18); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__19 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__19(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__19); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__20 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__20(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__20); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__21 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__21(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__21); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__22 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__22(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__22); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__23 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__23(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__23); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__24 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__24(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__24); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__25 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__25(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__25); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__26 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__26(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__26); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__27 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__27(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__27); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__28 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__28(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__28); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__29 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__29(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__29); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__30 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__30(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__30); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__31 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__31(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__31); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__32 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__32(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__32); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__33 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__33(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__33); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__34 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__34(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__34); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__35 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__35(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__35); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__36 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__36(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__36); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__37 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__37(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__37); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__38 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__38(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__38); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__39 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__39(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__39); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__40 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__40(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__40); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__41 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__41(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__41); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__42 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__42(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__42); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__43 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__43(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__43); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__44 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__44(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__44); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__45 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__45(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__45); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__46 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__46(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__46); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__47 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__47(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812____closed__47); -res = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3812_(lean_io_mk_world()); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__1(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__1); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__2(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__2); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__3(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__3); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__4 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__4(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__4); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__5 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__5(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__5); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__6 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__6(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__6); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__7 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__7(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__7); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__8 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__8(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__8); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__9 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__9(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__9); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__10 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__10(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__10); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__11 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__11(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__11); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__12 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__12(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__12); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__13 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__13(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__13); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__14 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__14(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__14); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__15 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__15(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__15); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__16 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__16(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__16); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__17 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__17(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__17); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__18 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__18(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__18); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__19 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__19(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__19); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__20 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__20(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__20); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__21 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__21(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__21); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__22 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__22(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__22); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__23 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__23(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__23); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__24 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__24(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__24); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__25 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__25(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__25); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__26 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__26(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__26); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__27 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__27(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__27); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__28 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__28(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__28); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__29 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__29(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__29); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__30 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__30(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__30); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__31 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__31(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__31); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__32 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__32(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__32); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__33 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__33(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__33); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__34 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__34(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__34); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__35 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__35(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__35); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__36 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__36(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__36); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__37 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__37(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__37); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__38 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__38(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__38); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__39 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__39(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__39); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__40 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__40(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__40); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__41 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__41(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__41); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__42 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__42(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__42); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__43 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__43(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__43); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__44 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__44(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__44); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__45 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__45(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__45); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__46 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__46(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__46); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__47 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__47(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__47); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__48 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__48(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__48); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__49 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__49(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883____closed__49); +res = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_3883_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/ParserCompiler.c b/stage0/stdlib/Lean/ParserCompiler.c index 54ff0a4d642..fa607a0d611 100644 --- a/stage0/stdlib/Lean/ParserCompiler.c +++ b/stage0/stdlib/Lean/ParserCompiler.c @@ -38,7 +38,6 @@ LEAN_EXPORT lean_object* l_Lean_ParserCompiler_compileParserExpr___rarg___lambda lean_object* lean_mk_empty_array_with_capacity(lean_object*); lean_object* l_Lean_mkSort(lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_ParserCompiler_compileParserExpr___spec__9(lean_object*); -extern lean_object* l_Lean_nullKind; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_ParserCompiler_compileParserExpr___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_ParserCompiler_registerParserCompiler___rarg___lambda__1___closed__6; static lean_object* l_Array_foldrMUnsafe_fold___at_Lean_ParserCompiler_compileParserExpr___spec__1___rarg___closed__1; @@ -84,6 +83,7 @@ lean_object* l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spe LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_ParserCompiler_compileParserExpr___spec__39___at_Lean_ParserCompiler_compileParserExpr___spec__40___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_ParserCompiler_registerParserCompiler___rarg___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_ParserCompiler_parserNodeKind_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__6; lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ParserCompiler_Context_tyName___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_ParserCompiler_parserNodeKind_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -182,6 +182,7 @@ LEAN_EXPORT lean_object* l_Lean_ParserCompiler_compileParserExpr___rarg___lambda LEAN_EXPORT lean_object* l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__16(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_ParserCompiler_registerParserCompiler___rarg___lambda__1___closed__12; LEAN_EXPORT lean_object* l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__33(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__5; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_ParserCompiler_compileParserExpr___spec__18___rarg___boxed(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_ParserCompiler_compileParserExpr___rarg___lambda__16___boxed(lean_object**); lean_object* lean_nat_sub(lean_object*, lean_object*); @@ -15687,6 +15688,24 @@ return x_2; static lean_object* _init_l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__4() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("null", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__4; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__6() { +_start: +{ lean_object* x_1; lean_object* x_2; x_1 = lean_unsigned_to_nat(2u); x_2 = lean_mk_empty_array_with_capacity(x_1); @@ -15759,7 +15778,7 @@ x_37 = lean_mk_syntax_ident(x_32); x_38 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__3; x_39 = lean_array_push(x_38, x_37); x_40 = lean_box(2); -x_41 = l_Lean_nullKind; +x_41 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__5; x_42 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_42, 0, x_40); lean_ctor_set(x_42, 1, x_41); @@ -15775,7 +15794,7 @@ lean_inc(x_44); lean_dec(x_43); lean_inc(x_44); x_45 = lean_mk_syntax_ident(x_44); -x_46 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__4; +x_46 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__6; x_47 = lean_array_push(x_46, x_45); x_48 = lean_array_push(x_47, x_42); x_49 = lean_alloc_ctor(1, 3, 0); @@ -15841,7 +15860,7 @@ lean_inc(x_60); lean_dec(x_59); lean_inc(x_60); x_61 = lean_mk_syntax_ident(x_60); -x_62 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__4; +x_62 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__6; x_63 = lean_array_push(x_62, x_61); x_64 = lean_array_push(x_63, x_42); x_65 = lean_alloc_ctor(1, 3, 0); @@ -16525,7 +16544,7 @@ x_37 = lean_mk_syntax_ident(x_32); x_38 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__3; x_39 = lean_array_push(x_38, x_37); x_40 = lean_box(2); -x_41 = l_Lean_nullKind; +x_41 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__5; x_42 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_42, 0, x_40); lean_ctor_set(x_42, 1, x_41); @@ -16541,7 +16560,7 @@ lean_inc(x_44); lean_dec(x_43); lean_inc(x_44); x_45 = lean_mk_syntax_ident(x_44); -x_46 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__4; +x_46 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__6; x_47 = lean_array_push(x_46, x_45); x_48 = lean_array_push(x_47, x_42); x_49 = lean_alloc_ctor(1, 3, 0); @@ -16607,7 +16626,7 @@ lean_inc(x_60); lean_dec(x_59); lean_inc(x_60); x_61 = lean_mk_syntax_ident(x_60); -x_62 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__4; +x_62 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__6; x_63 = lean_array_push(x_62, x_61); x_64 = lean_array_push(x_63, x_42); x_65 = lean_alloc_ctor(1, 3, 0); @@ -17291,7 +17310,7 @@ x_37 = lean_mk_syntax_ident(x_32); x_38 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__3; x_39 = lean_array_push(x_38, x_37); x_40 = lean_box(2); -x_41 = l_Lean_nullKind; +x_41 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__5; x_42 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_42, 0, x_40); lean_ctor_set(x_42, 1, x_41); @@ -17307,7 +17326,7 @@ lean_inc(x_44); lean_dec(x_43); lean_inc(x_44); x_45 = lean_mk_syntax_ident(x_44); -x_46 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__4; +x_46 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__6; x_47 = lean_array_push(x_46, x_45); x_48 = lean_array_push(x_47, x_42); x_49 = lean_alloc_ctor(1, 3, 0); @@ -17373,7 +17392,7 @@ lean_inc(x_60); lean_dec(x_59); lean_inc(x_60); x_61 = lean_mk_syntax_ident(x_60); -x_62 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__4; +x_62 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__6; x_63 = lean_array_push(x_62, x_61); x_64 = lean_array_push(x_63, x_42); x_65 = lean_alloc_ctor(1, 3, 0); @@ -18057,7 +18076,7 @@ x_37 = lean_mk_syntax_ident(x_32); x_38 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__3; x_39 = lean_array_push(x_38, x_37); x_40 = lean_box(2); -x_41 = l_Lean_nullKind; +x_41 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__5; x_42 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_42, 0, x_40); lean_ctor_set(x_42, 1, x_41); @@ -18073,7 +18092,7 @@ lean_inc(x_44); lean_dec(x_43); lean_inc(x_44); x_45 = lean_mk_syntax_ident(x_44); -x_46 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__4; +x_46 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__6; x_47 = lean_array_push(x_46, x_45); x_48 = lean_array_push(x_47, x_42); x_49 = lean_alloc_ctor(1, 3, 0); @@ -18139,7 +18158,7 @@ lean_inc(x_60); lean_dec(x_59); lean_inc(x_60); x_61 = lean_mk_syntax_ident(x_60); -x_62 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__4; +x_62 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__6; x_63 = lean_array_push(x_62, x_61); x_64 = lean_array_push(x_63, x_42); x_65 = lean_alloc_ctor(1, 3, 0); @@ -18823,7 +18842,7 @@ x_37 = lean_mk_syntax_ident(x_32); x_38 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__3; x_39 = lean_array_push(x_38, x_37); x_40 = lean_box(2); -x_41 = l_Lean_nullKind; +x_41 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__5; x_42 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_42, 0, x_40); lean_ctor_set(x_42, 1, x_41); @@ -18839,7 +18858,7 @@ lean_inc(x_44); lean_dec(x_43); lean_inc(x_44); x_45 = lean_mk_syntax_ident(x_44); -x_46 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__4; +x_46 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__6; x_47 = lean_array_push(x_46, x_45); x_48 = lean_array_push(x_47, x_42); x_49 = lean_alloc_ctor(1, 3, 0); @@ -18905,7 +18924,7 @@ lean_inc(x_60); lean_dec(x_59); lean_inc(x_60); x_61 = lean_mk_syntax_ident(x_60); -x_62 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__4; +x_62 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__6; x_63 = lean_array_push(x_62, x_61); x_64 = lean_array_push(x_63, x_42); x_65 = lean_alloc_ctor(1, 3, 0); @@ -19645,7 +19664,7 @@ x_37 = lean_mk_syntax_ident(x_32); x_38 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__3; x_39 = lean_array_push(x_38, x_37); x_40 = lean_box(2); -x_41 = l_Lean_nullKind; +x_41 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__5; x_42 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_42, 0, x_40); lean_ctor_set(x_42, 1, x_41); @@ -19661,7 +19680,7 @@ lean_inc(x_44); lean_dec(x_43); lean_inc(x_44); x_45 = lean_mk_syntax_ident(x_44); -x_46 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__4; +x_46 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__6; x_47 = lean_array_push(x_46, x_45); x_48 = lean_array_push(x_47, x_42); x_49 = lean_alloc_ctor(1, 3, 0); @@ -19727,7 +19746,7 @@ lean_inc(x_60); lean_dec(x_59); lean_inc(x_60); x_61 = lean_mk_syntax_ident(x_60); -x_62 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__4; +x_62 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__6; x_63 = lean_array_push(x_62, x_61); x_64 = lean_array_push(x_63, x_42); x_65 = lean_alloc_ctor(1, 3, 0); @@ -20411,7 +20430,7 @@ x_37 = lean_mk_syntax_ident(x_32); x_38 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__3; x_39 = lean_array_push(x_38, x_37); x_40 = lean_box(2); -x_41 = l_Lean_nullKind; +x_41 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__5; x_42 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_42, 0, x_40); lean_ctor_set(x_42, 1, x_41); @@ -20427,7 +20446,7 @@ lean_inc(x_44); lean_dec(x_43); lean_inc(x_44); x_45 = lean_mk_syntax_ident(x_44); -x_46 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__4; +x_46 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__6; x_47 = lean_array_push(x_46, x_45); x_48 = lean_array_push(x_47, x_42); x_49 = lean_alloc_ctor(1, 3, 0); @@ -20493,7 +20512,7 @@ lean_inc(x_60); lean_dec(x_59); lean_inc(x_60); x_61 = lean_mk_syntax_ident(x_60); -x_62 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__4; +x_62 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__6; x_63 = lean_array_push(x_62, x_61); x_64 = lean_array_push(x_63, x_42); x_65 = lean_alloc_ctor(1, 3, 0); @@ -21177,7 +21196,7 @@ x_37 = lean_mk_syntax_ident(x_32); x_38 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__3; x_39 = lean_array_push(x_38, x_37); x_40 = lean_box(2); -x_41 = l_Lean_nullKind; +x_41 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__5; x_42 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_42, 0, x_40); lean_ctor_set(x_42, 1, x_41); @@ -21193,7 +21212,7 @@ lean_inc(x_44); lean_dec(x_43); lean_inc(x_44); x_45 = lean_mk_syntax_ident(x_44); -x_46 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__4; +x_46 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__6; x_47 = lean_array_push(x_46, x_45); x_48 = lean_array_push(x_47, x_42); x_49 = lean_alloc_ctor(1, 3, 0); @@ -21259,7 +21278,7 @@ lean_inc(x_60); lean_dec(x_59); lean_inc(x_60); x_61 = lean_mk_syntax_ident(x_60); -x_62 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__4; +x_62 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__6; x_63 = lean_array_push(x_62, x_61); x_64 = lean_array_push(x_63, x_42); x_65 = lean_alloc_ctor(1, 3, 0); @@ -21943,7 +21962,7 @@ x_37 = lean_mk_syntax_ident(x_32); x_38 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__3; x_39 = lean_array_push(x_38, x_37); x_40 = lean_box(2); -x_41 = l_Lean_nullKind; +x_41 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__5; x_42 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_42, 0, x_40); lean_ctor_set(x_42, 1, x_41); @@ -21959,7 +21978,7 @@ lean_inc(x_44); lean_dec(x_43); lean_inc(x_44); x_45 = lean_mk_syntax_ident(x_44); -x_46 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__4; +x_46 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__6; x_47 = lean_array_push(x_46, x_45); x_48 = lean_array_push(x_47, x_42); x_49 = lean_alloc_ctor(1, 3, 0); @@ -22025,7 +22044,7 @@ lean_inc(x_60); lean_dec(x_59); lean_inc(x_60); x_61 = lean_mk_syntax_ident(x_60); -x_62 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__4; +x_62 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__6; x_63 = lean_array_push(x_62, x_61); x_64 = lean_array_push(x_63, x_42); x_65 = lean_alloc_ctor(1, 3, 0); @@ -22709,7 +22728,7 @@ x_37 = lean_mk_syntax_ident(x_32); x_38 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__3; x_39 = lean_array_push(x_38, x_37); x_40 = lean_box(2); -x_41 = l_Lean_nullKind; +x_41 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__5; x_42 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_42, 0, x_40); lean_ctor_set(x_42, 1, x_41); @@ -22725,7 +22744,7 @@ lean_inc(x_44); lean_dec(x_43); lean_inc(x_44); x_45 = lean_mk_syntax_ident(x_44); -x_46 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__4; +x_46 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__6; x_47 = lean_array_push(x_46, x_45); x_48 = lean_array_push(x_47, x_42); x_49 = lean_alloc_ctor(1, 3, 0); @@ -22791,7 +22810,7 @@ lean_inc(x_60); lean_dec(x_59); lean_inc(x_60); x_61 = lean_mk_syntax_ident(x_60); -x_62 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__4; +x_62 = l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__6; x_63 = lean_array_push(x_62, x_61); x_64 = lean_array_push(x_63, x_42); x_65 = lean_alloc_ctor(1, 3, 0); @@ -31046,6 +31065,10 @@ l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__3 = _init_l lean_mark_persistent(l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__3); l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__4 = _init_l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__4(); lean_mark_persistent(l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__4); +l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__5 = _init_l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__5(); +lean_mark_persistent(l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__5); +l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__6 = _init_l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__6(); +lean_mark_persistent(l_Lean_ParserCompiler_compileParserExpr___rarg___lambda__4___closed__6); l_Lean_ParserCompiler_compileParserExpr___rarg___closed__1 = _init_l_Lean_ParserCompiler_compileParserExpr___rarg___closed__1(); lean_mark_persistent(l_Lean_ParserCompiler_compileParserExpr___rarg___closed__1); l_Lean_ParserCompiler_compileParserExpr___rarg___closed__2 = _init_l_Lean_ParserCompiler_compileParserExpr___rarg___closed__2(); diff --git a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Basic.c b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Basic.c index 10e67743d35..3f1739ba3ec 100644 --- a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Basic.c +++ b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Basic.c @@ -102,6 +102,7 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_getExprKind___closed__12; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_addTermInfo_mkTermInfo___boxed(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_resolveGlobalConstNoOverloadCore___at_Lean_PrettyPrinter_Delaborator_mkAppUnexpanderAttribute___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delab___lambda__1___closed__2; +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_PrettyPrinter_delabCore___spec__3___boxed(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_getExprKind___closed__11; lean_object* l_Lean_Expr_FindImpl_findUnsafe_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_mkAppUnexpanderAttribute(lean_object*); @@ -168,6 +169,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_mkDelabAttribute(lean_ lean_object* l_Lean_throwError___at_Lean_registerInitAttrUnsafe___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); lean_object* l_Lean_getPPProofs___boxed(lean_object*); +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_PrettyPrinter_delabCore___spec__3(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delab___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_instOrElseDelabM___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_instMonadWithReaderOfSubExprDelabM(lean_object*); @@ -220,7 +222,6 @@ lean_object* l_Lean_registerInternalExceptionId(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_instMonadQuotationDelabM___closed__11; static lean_object* l_Lean_PrettyPrinter_Delaborator_getUnusedName___closed__5; static lean_object* l_Lean_PrettyPrinter_Delaborator_mkAppUnexpanderAttribute___closed__4; -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_PrettyPrinter_delabCore___spec__3(lean_object*, lean_object*); uint8_t l_Lean_getPPAnalyze(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_mkDelabAttribute___closed__11; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_delab(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -239,7 +240,6 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_instAlternativeDelabM___clo static lean_object* l_Lean_PrettyPrinter_Delaborator_instMonadReaderOfSubExprDelabM___closed__2; lean_object* l_Lean_Expr_bindingName_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_withBindingBodyUnusedName___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_PrettyPrinter_delabCore___spec__3___boxed(lean_object*, lean_object*); extern lean_object* l_Lean_SubExpr_Pos_maxChildren; static lean_object* l_Lean_PrettyPrinter_delabCore___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_delab___lambda__1___closed__21; @@ -313,7 +313,7 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_instMonadQuotationDelabM___ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_PrettyPrinter_Delaborator_delab___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_PrettyPrinter_Delaborator_delab___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instMonadReaderT___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Delaborator_Basic___hyg_2804_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Delaborator_Basic___hyg_2830_(lean_object*); static lean_object* l_Lean_PrettyPrinter_delabCore___closed__5; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_withBindingBodyUnusedName___spec__1(lean_object*); static lean_object* l_Lean_resolveGlobalConstNoOverloadCore___at_Lean_PrettyPrinter_Delaborator_mkAppUnexpanderAttribute___spec__1___closed__1; @@ -8745,7 +8745,7 @@ return x_15; } } } -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_PrettyPrinter_delabCore___spec__3(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_PrettyPrinter_delabCore___spec__3(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -9255,7 +9255,7 @@ lean_inc(x_10); x_11 = l_Lean_PrettyPrinter_delabCore___lambda__4___closed__1; x_12 = l_Lean_Option_get_x3f___at_Lean_PrettyPrinter_delabCore___spec__2(x_10, x_11); x_13 = lean_box(0); -x_14 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_PrettyPrinter_delabCore___spec__3(x_12, x_13); +x_14 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_PrettyPrinter_delabCore___spec__3(x_12, x_13); lean_dec(x_12); if (x_14 == 0) { @@ -9509,11 +9509,11 @@ lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_PrettyPrinter_delabCore___spec__3___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_PrettyPrinter_delabCore___spec__3___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_PrettyPrinter_delabCore___spec__3(x_1, x_2); +x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_PrettyPrinter_delabCore___spec__3(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -9581,7 +9581,7 @@ return x_20; } } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Delaborator_Basic___hyg_2804_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Delaborator_Basic___hyg_2830_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; @@ -9939,7 +9939,7 @@ l_Lean_PrettyPrinter_delabCore___closed__6 = _init_l_Lean_PrettyPrinter_delabCor lean_mark_persistent(l_Lean_PrettyPrinter_delabCore___closed__6); l_Lean_PrettyPrinter_delabCore___closed__7 = _init_l_Lean_PrettyPrinter_delabCore___closed__7(); lean_mark_persistent(l_Lean_PrettyPrinter_delabCore___closed__7); -res = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Delaborator_Basic___hyg_2804_(lean_io_mk_world()); +res = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Delaborator_Basic___hyg_2830_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Builtins.c b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Builtins.c index 5f322c4e717..6e39b691b1c 100644 --- a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Builtins.c +++ b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Builtins.c @@ -13,9 +13,10 @@ #ifdef __cplusplus extern "C" { #endif -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__2___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_PrettyPrinter_Delaborator_getParamKinds(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandCoe___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__1; lean_object* l_Lean_Level_dec(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabFVar___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -26,6 +27,7 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_maybeAddBlockImplicit___clo static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__3___closed__8; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicit___lambda__2___closed__5; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicit___lambda__3___closed__2; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withMDataExpr___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__2(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg(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_PrettyPrinter_Delaborator_isRegularApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -38,7 +40,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_unresolveNameGlobal_un static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabProj___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__3___closed__4; size_t lean_usize_add(size_t, size_t); -extern lean_object* l_Lean_fieldIdxKind; extern lean_object* l_Lean_PrettyPrinter_Delaborator_delabFailureId; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabMData(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getPPLetVarTypes___boxed(lean_object*); @@ -57,11 +58,12 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_L static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabCond___closed__2; lean_object* lean_erase_macro_scopes(lean_object*); lean_object* l_Lean_SubExpr_Pos_push(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__5___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNameMkStr(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withMDataExpr___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__2___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabDIte___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_rootNamespace; +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabAppImplicit___spec__5___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicit___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicit___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__9___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -92,16 +94,15 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabSigmaCore___lambda__1_ lean_object* l_Lean_LocalDecl_userName(lean_object*); lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withType___at_Lean_PrettyPrinter_Delaborator_delab___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_isRegularApp___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_nullKind; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabFVar___closed__4; uint8_t l_Lean_Expr_isMData(lean_object*); static lean_object* l_List_forIn_loop___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__1___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_isRegularApp___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandCoe___closed__1; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__2; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabProj___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConst___closed__2; -LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__11(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicit___lambda__5___closed__4; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabLetE___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDIte___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -110,15 +111,16 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConst___lambda__1 lean_object* l_Array_eraseIdx___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabProj___closed__5; lean_object* l_Lean_Meta_withLocalDecl___at_Lean_PrettyPrinter_Delaborator_withBindingBodyUnusedName___spec__3___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabMData___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabForall___spec__4(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConst___spec__2(size_t, size_t, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__3___closed__2; lean_object* l_Lean_getPPAnalysisNamedArg___boxed(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicit___lambda__4___closed__10; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_bindingDomain_x21(lean_object*); lean_object* l_Array_append___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicit___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -138,6 +140,7 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___closed__2; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabFVar___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNameMkNum(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__3___closed__1; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__15; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); @@ -147,6 +150,7 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabProjectionApp___closed LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetE___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_isRegularApp___lambda__4___closed__4; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__6; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_extract___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabCond___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDIte___lambda__1___closed__8; @@ -161,6 +165,7 @@ lean_object* l_Std_RBNode_find___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyz static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabPSigma___closed__4; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNamedPattern___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabNameMkStr___closed__2; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabAppImplicit___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConst___closed__13; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicit___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicit___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -185,6 +190,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBod lean_object* l_List_tailD___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConst___spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___closed__5; +static lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__12___closed__1; lean_object* l_Lean_SubExpr_Pos_pushNaryArg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_withMDataOptions(lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabFVar___closed__8; @@ -208,6 +214,7 @@ LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_withMDataOp static lean_object* l_Lean_PrettyPrinter_Delaborator_delabCond___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabDo(lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNameMkStr___closed__5; +lean_object* l_Lean_Syntax_mkScientificLit(lean_object*, lean_object*); lean_object* l_Lean_Expr_appFn_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabSort(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_instInhabitedNat; @@ -228,7 +235,6 @@ static lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_u LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabMData___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_reifyName(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicit___lambda__2___closed__4; -static lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__11___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_reifyName___closed__1; static lean_object* l_List_forIn_loop___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__1___closed__2; lean_object* lean_string_append(lean_object*, lean_object*); @@ -251,7 +257,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_withMDataOptions___rar lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_getExprKind___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getOptParamDefault_x3f(lean_object*); lean_object* lean_string_utf8_extract(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__12(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabMVar___closed__4; static lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__7; @@ -262,7 +267,6 @@ lean_object* l_Array_shrink___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_unfoldMDatas___boxed(lean_object*); uint8_t l_Lean_PrettyPrinter_Delaborator_isCoe(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicit___lambda__3(lean_object*, 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_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabAppImplicit___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabNameMkStr___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabNameMkStr___lambda__1___closed__1; lean_object* l_Lean_getPPFullNames___boxed(lean_object*); @@ -292,7 +296,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabForall___lambda__ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__2; uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabMData___closed__2; -extern lean_object* l_Lean_nameLitKind; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___closed__4; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabSigma___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabCond___lambda__1___closed__2; @@ -300,6 +303,7 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_reifyName___closed__4; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabLam___closed__1; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabPSigma___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabNamedPattern___closed__1; +LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__12(lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__12(lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabAppImplicit___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_nextExtraPos___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -315,7 +319,6 @@ static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabProj___cl LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabLam___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabDo___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDIte_delabBranch___closed__1; -lean_object* l_Lean_Syntax_SepArray_ofElems(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_unfoldMDatas(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabAppExplicit___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -337,7 +340,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabOfSc LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_hasIdent___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabMVar___closed__3; static lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_isRegularApp___spec__3___closed__4; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabCond(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfScientific___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -350,7 +352,6 @@ lean_object* l_Lean_getPPAnalysisLetVarType___boxed(lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_isRegularApp___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_mkApp(lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); -lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(size_t, size_t, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__3___closed__6; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfScientific___lambda__1___closed__3; LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -366,11 +367,13 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabFVar(lean_object* LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNamedPattern(lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withMDataExpr___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__2___rarg___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfScientific___lambda__2___closed__1; LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabBinders(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabMVar___closed__3; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabLam___spec__1(lean_object*, lean_object*, size_t, size_t); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConst___spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfScientific___lambda__1___closed__5; LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__9(lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); @@ -383,16 +386,16 @@ lean_object* l_Lean_PrettyPrinter_Delaborator_annotateCurPos(lean_object*, lean_ static lean_object* l_Lean_PrettyPrinter_Delaborator_reifyName___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_isRegularApp___lambda__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandRegularApp___closed__1; +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLit___closed__1; uint8_t l_Lean_getPPFunBinderTypes(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfScientific___lambda__2___closed__2; lean_object* l_Lean_SubExpr_Pos_pushNaryFn(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_isRegularApp___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabSigmaCore___lambda__2(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabMData___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_numLitKind; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabForall___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unresolveNameGlobal_unresolveNameCore___spec__2___lambda__1___boxed(lean_object*, lean_object*, 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_PrettyPrinter_Delaborator_SubExpr_withAppFnArgs___at_Lean_PrettyPrinter_Delaborator_delabAppExplicit___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -466,7 +469,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandCoe___lambda__ lean_object* l_instInhabited___rarg(lean_object*, lean_object*); lean_object* l_Std_RBNode_insert___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_annotateBoolAt___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9(size_t, size_t, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_maybeAddBlockImplicit___closed__4; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -501,6 +504,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_L lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAux___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_LocalDecl_binderInfo(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__4; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabSigmaCore___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_shouldGroupWithNext___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -532,6 +536,7 @@ LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_withMDataOp lean_object* l_Lean_getStructureFields(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabMVar___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__1___closed__2; +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabProj___closed__6; static lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_isRegularApp___spec__3___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___at_Lean_PrettyPrinter_Delaborator_getParamKinds_forallTelescopeArgs___spec__2(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__11; @@ -540,7 +545,6 @@ lean_object* lean_array_to_list(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDIte_delabBranch(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDIte___lambda__1___closed__5; -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabAppImplicit___spec__1(lean_object*, size_t, size_t); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__3; lean_object* l_Lean_KeyedDeclsAttribute_getValues___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__1___closed__4; @@ -641,7 +645,7 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabSigmaCore___lambda__1_ LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_PrettyPrinter_Delaborator_delabLetE___spec__1___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabFVar(lean_object*); lean_object* l_Lean_PrettyPrinter_Delaborator_withBindingBodyUnusedName___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__7; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabLam___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___closed__4; @@ -653,7 +657,6 @@ static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabFVar___cl LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfNat(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_projectionFnInfoExt; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__3; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicit___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getPPCoercions___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabDoElems___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -671,7 +674,7 @@ static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabSort___cl static lean_object* l_Lean_PrettyPrinter_Delaborator_delabProjectionApp___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_shouldGroupWithNext___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__3___closed__6; -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_mkNumLit(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabProj___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabProj___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicit___lambda__4___closed__4; @@ -723,7 +726,7 @@ lean_object* l_Lean_Syntax_getNumArgs(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabForall___spec__5___boxed(lean_object*, lean_object*, 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_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_getParamKinds___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_scientificLitKind; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFnArgs___at_Lean_PrettyPrinter_Delaborator_delabAppImplicit___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabDIte___spec__2(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_PrettyPrinter_Delaborator_maybeAddBlockImplicit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_bindingName_x21(lean_object*); @@ -809,6 +812,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFnArgs_ LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabLetE(lean_object*); lean_object* l_Lean_getPPStructureInstances___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabAppImplicit___spec__5(lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabNameMkNum(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_PrettyPrinter_Delaborator_unresolveNameGlobal___spec__1___boxed(lean_object*, lean_object*, lean_object*, 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_PrettyPrinter_Delaborator_SubExpr_nextExtraPos___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___boxed(lean_object*); @@ -824,6 +828,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda_ static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabMData___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicit___lambda__4___closed__5; static lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withMDataExpr___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__2___rarg___closed__3; +lean_object* l_Lean_Syntax_mkNameLit(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__13; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabNameMkStr___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -841,6 +846,7 @@ static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabPSigma___ static lean_object* l_Array_mapIdxM_map___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__5___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withType___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabAppImplicit___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConst___closed__14; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_isRegularApp___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_instInhabitedName; @@ -901,6 +907,7 @@ static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabFVar___cl lean_object* lean_string_length(lean_object*); static lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__2; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNameMkNum___closed__4; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabAppImplicit___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__9; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_getParamKinds_forallTelescopeArgs___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instMonadReaderT___rarg(lean_object*); @@ -964,7 +971,6 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delabora static lean_object* l_Lean_PrettyPrinter_Delaborator_delabForall___lambda__2___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConst___closed__7; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabForall___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___closed__2; LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabBinders___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Match_Extension_getMatcherInfo_x3f(lean_object*, lean_object*); @@ -1346,7 +1352,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__4 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_PrettyPrinter_Delaborator_delabFVar___closed__1; x_2 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__2; -x_3 = lean_unsigned_to_nat(26u); +x_3 = lean_unsigned_to_nat(27u); x_4 = lean_unsigned_to_nat(31u); x_5 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -1362,7 +1368,7 @@ x_9 = lean_ctor_get(x_8, 0); lean_inc(x_9); if (lean_obj_tag(x_9) == 1) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_24; x_10 = lean_ctor_get(x_8, 1); lean_inc(x_10); lean_dec(x_8); @@ -1371,27 +1377,28 @@ lean_inc(x_11); lean_dec(x_9); lean_inc(x_3); lean_inc(x_11); -x_12 = l_Lean_Meta_getLocalDecl(x_11, x_3, x_4, x_5, x_6, x_10); -if (lean_obj_tag(x_12) == 0) +x_24 = l_Lean_Meta_getLocalDecl(x_11, x_3, x_4, x_5, x_6, x_10); +if (lean_obj_tag(x_24) == 0) { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec(x_12); -x_15 = l_Lean_LocalDecl_userName(x_13); -lean_dec(x_13); -x_16 = lean_mk_syntax_ident(x_15); +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); +x_27 = l_Lean_LocalDecl_userName(x_25); +lean_dec(x_25); +x_28 = lean_mk_syntax_ident(x_27); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_17 = l_Lean_PrettyPrinter_Delaborator_maybeAddBlockImplicit(x_16, x_1, x_2, x_3, x_4, x_5, x_6, x_14); -if (lean_obj_tag(x_17) == 0) +x_29 = l_Lean_PrettyPrinter_Delaborator_maybeAddBlockImplicit(x_28, x_1, x_2, x_3, x_4, x_5, x_6, x_26); +if (lean_obj_tag(x_29) == 0) { +uint8_t x_30; lean_dec(x_11); lean_dec(x_6); lean_dec(x_5); @@ -1399,40 +1406,105 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_17; +x_30 = !lean_is_exclusive(x_29); +if (x_30 == 0) +{ +return x_29; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_17, 1); -lean_inc(x_18); -lean_dec(x_17); -x_19 = lean_mk_syntax_ident(x_11); -x_20 = l_Lean_PrettyPrinter_Delaborator_maybeAddBlockImplicit(x_19, x_1, x_2, x_3, x_4, x_5, x_6, x_18); -return x_20; +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_29, 0); +x_32 = lean_ctor_get(x_29, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_29); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; } } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_12, 1); +lean_object* x_34; +x_34 = lean_ctor_get(x_29, 1); +lean_inc(x_34); +lean_dec(x_29); +x_12 = x_34; +goto block_23; +} +} +else +{ +lean_object* x_35; +x_35 = lean_ctor_get(x_24, 1); +lean_inc(x_35); +lean_dec(x_24); +x_12 = x_35; +goto block_23; +} +block_23: +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_mk_syntax_ident(x_11); +x_14 = l_Lean_PrettyPrinter_Delaborator_maybeAddBlockImplicit(x_13, x_1, x_2, x_3, x_4, x_5, x_6, x_12); +if (lean_obj_tag(x_14) == 0) +{ +uint8_t x_15; +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) +{ +return x_14; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_14, 0); +x_17 = lean_ctor_get(x_14, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_14); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +return x_18; +} +} +else +{ +uint8_t x_19; +x_19 = !lean_is_exclusive(x_14); +if (x_19 == 0) +{ +return x_14; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_14, 0); +x_21 = lean_ctor_get(x_14, 1); lean_inc(x_21); -lean_dec(x_12); -x_22 = lean_mk_syntax_ident(x_11); -x_23 = l_Lean_PrettyPrinter_Delaborator_maybeAddBlockImplicit(x_22, x_1, x_2, x_3, x_4, x_5, x_6, x_21); -return x_23; +lean_inc(x_20); +lean_dec(x_14); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} } } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_dec(x_9); -x_24 = lean_ctor_get(x_8, 1); -lean_inc(x_24); +x_36 = lean_ctor_get(x_8, 1); +lean_inc(x_36); lean_dec(x_8); -x_25 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__4; -x_26 = l_panic___at_Lean_PrettyPrinter_Delaborator_delabFVar___spec__1(x_25, x_1, x_2, x_3, x_4, x_5, x_6, x_24); -return x_26; +x_37 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__4; +x_38 = l_panic___at_Lean_PrettyPrinter_Delaborator_delabFVar___spec__1(x_37, x_1, x_2, x_3, x_4, x_5, x_6, x_36); +return x_38; } } } @@ -1558,7 +1630,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabBVar___closed__3 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_PrettyPrinter_Delaborator_delabFVar___closed__1; x_2 = l_Lean_PrettyPrinter_Delaborator_delabBVar___closed__2; -x_3 = lean_unsigned_to_nat(37u); +x_3 = lean_unsigned_to_nat(38u); x_4 = lean_unsigned_to_nat(34u); x_5 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -1731,7 +1803,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabMVar___closed__5 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_PrettyPrinter_Delaborator_delabFVar___closed__1; x_2 = l_Lean_PrettyPrinter_Delaborator_delabMVar___closed__4; -x_3 = lean_unsigned_to_nat(42u); +x_3 = lean_unsigned_to_nat(43u); x_4 = lean_unsigned_to_nat(32u); x_5 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -1971,7 +2043,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabSort___closed__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; x_1 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__1; x_2 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__1; -x_3 = lean_unsigned_to_nat(52u); +x_3 = lean_unsigned_to_nat(53u); x_4 = lean_unsigned_to_nat(32u); x_5 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -3712,6 +3784,30 @@ goto _start; } } } +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConst___spec__2(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = 1; +x_9 = lean_usize_add(x_2, x_8); +x_10 = lean_array_uset(x_7, x_2, x_5); +x_2 = x_9; +x_3 = x_10; +goto _start; +} +} +} static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConst___lambda__1___closed__1() { _start: { @@ -3923,7 +4019,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConst___closed__ 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_PrettyPrinter_Delaborator_delabFVar___closed__1; x_2 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__1; -x_3 = lean_unsigned_to_nat(89u); +x_3 = lean_unsigned_to_nat(90u); x_4 = lean_unsigned_to_nat(37u); x_5 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -4177,7 +4273,7 @@ x_57 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConst___s x_58 = lean_array_get_size(x_57); x_59 = lean_usize_of_nat(x_58); lean_dec(x_58); -x_60 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_59, x_56, x_57); +x_60 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConst___spec__2(x_59, x_56, x_57); x_61 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__12; x_62 = l_Lean_mkSepArray(x_60, x_61); lean_dec(x_60); @@ -4390,6 +4486,18 @@ x_6 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConst___sp return x_6; } } +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConst___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConst___spec__2(x_4, x_5, x_3); +return x_6; +} +} LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConst___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, lean_object* x_9) { _start: { @@ -5164,7 +5272,7 @@ return x_13; else { lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = l_Lean_instInhabitedSyntax; +x_14 = lean_box(0); x_15 = l_Lean_PrettyPrinter_Delaborator_delabAppFn___closed__1; x_16 = l_Lean_PrettyPrinter_Delaborator_withMDatasOptions___rarg(x_14, x_15, x_1, x_2, x_3, x_4, x_5, x_6, x_10); return x_16; @@ -8249,7 +8357,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_unexpandRegularApp___ 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_PrettyPrinter_Delaborator_delabFVar___closed__1; x_2 = l_Lean_PrettyPrinter_Delaborator_unexpandRegularApp___closed__1; -x_3 = lean_unsigned_to_nat(200u); +x_3 = lean_unsigned_to_nat(201u); x_4 = lean_unsigned_to_nat(63u); x_5 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -8358,89 +8466,88 @@ return x_12; } else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +lean_object* x_13; lean_object* x_14; uint8_t x_15; x_13 = lean_unsigned_to_nat(1u); x_14 = l_Lean_Syntax_getArg(x_1, x_13); lean_dec(x_1); -x_15 = l_Lean_nullKind; lean_inc(x_14); -x_16 = l_Lean_Syntax_isNodeOf(x_14, x_15, x_13); -if (x_16 == 0) +x_15 = l_Lean_Syntax_matchesNull(x_14, x_13); +if (x_15 == 0) { -lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_17 = l_Lean_Syntax_getArgs(x_14); +lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_16 = l_Lean_Syntax_getArgs(x_14); lean_dec(x_14); -x_18 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_PrettyPrinter_Delaborator_delab___spec__3___rarg(x_7, x_8, x_9); -x_19 = !lean_is_exclusive(x_18); -if (x_19 == 0) +x_17 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_PrettyPrinter_Delaborator_delab___spec__3___rarg(x_7, x_8, x_9); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) { -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; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_20 = lean_ctor_get(x_18, 0); -lean_dec(x_20); -x_21 = l_Lean_instInhabitedSyntax; -x_22 = lean_unsigned_to_nat(0u); -x_23 = lean_array_get(x_21, x_17, x_22); -x_24 = l_Array_eraseIdx___rarg(x_17, x_22); -x_25 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__12; -x_26 = l_Array_append___rarg(x_25, x_24); -x_27 = lean_box(2); -x_28 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__11; -x_29 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -lean_ctor_set(x_29, 2, x_26); -x_30 = l_Lean_PrettyPrinter_Delaborator_maybeAddBlockImplicit___closed__11; -x_31 = lean_array_push(x_30, x_23); -x_32 = lean_array_push(x_31, x_29); -x_33 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_33, 0, x_27); -lean_ctor_set(x_33, 1, x_10); -lean_ctor_set(x_33, 2, x_32); -lean_ctor_set(x_18, 0, x_33); -return 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; 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_object* x_32; +x_19 = lean_ctor_get(x_17, 0); +lean_dec(x_19); +x_20 = lean_box(0); +x_21 = lean_unsigned_to_nat(0u); +x_22 = lean_array_get(x_20, x_16, x_21); +x_23 = l_Array_eraseIdx___rarg(x_16, x_21); +x_24 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__12; +x_25 = l_Array_append___rarg(x_24, x_23); +x_26 = lean_box(2); +x_27 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__11; +x_28 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +lean_ctor_set(x_28, 2, x_25); +x_29 = l_Lean_PrettyPrinter_Delaborator_maybeAddBlockImplicit___closed__11; +x_30 = lean_array_push(x_29, x_22); +x_31 = lean_array_push(x_30, x_28); +x_32 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_32, 0, x_26); +lean_ctor_set(x_32, 1, x_10); +lean_ctor_set(x_32, 2, x_31); +lean_ctor_set(x_17, 0, x_32); +return x_17; } else { -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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_34 = lean_ctor_get(x_18, 1); -lean_inc(x_34); -lean_dec(x_18); -x_35 = l_Lean_instInhabitedSyntax; -x_36 = lean_unsigned_to_nat(0u); -x_37 = lean_array_get(x_35, x_17, x_36); -x_38 = l_Array_eraseIdx___rarg(x_17, x_36); -x_39 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__12; -x_40 = l_Array_append___rarg(x_39, x_38); -x_41 = lean_box(2); -x_42 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__11; -x_43 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_43, 0, x_41); -lean_ctor_set(x_43, 1, x_42); -lean_ctor_set(x_43, 2, x_40); -x_44 = l_Lean_PrettyPrinter_Delaborator_maybeAddBlockImplicit___closed__11; -x_45 = lean_array_push(x_44, x_37); -x_46 = lean_array_push(x_45, x_43); -x_47 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_47, 0, x_41); -lean_ctor_set(x_47, 1, x_10); -lean_ctor_set(x_47, 2, x_46); -x_48 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_34); -return x_48; +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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_33 = lean_ctor_get(x_17, 1); +lean_inc(x_33); +lean_dec(x_17); +x_34 = lean_box(0); +x_35 = lean_unsigned_to_nat(0u); +x_36 = lean_array_get(x_34, x_16, x_35); +x_37 = l_Array_eraseIdx___rarg(x_16, x_35); +x_38 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__12; +x_39 = l_Array_append___rarg(x_38, x_37); +x_40 = lean_box(2); +x_41 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__11; +x_42 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_41); +lean_ctor_set(x_42, 2, x_39); +x_43 = l_Lean_PrettyPrinter_Delaborator_maybeAddBlockImplicit___closed__11; +x_44 = lean_array_push(x_43, x_36); +x_45 = lean_array_push(x_44, x_42); +x_46 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_46, 0, x_40); +lean_ctor_set(x_46, 1, x_10); +lean_ctor_set(x_46, 2, x_45); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_33); +return x_47; } } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_dec(x_7); -x_49 = lean_unsigned_to_nat(0u); -x_50 = l_Lean_Syntax_getArg(x_14, x_49); +x_48 = lean_unsigned_to_nat(0u); +x_49 = l_Lean_Syntax_getArg(x_14, x_48); lean_dec(x_14); -x_51 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_9); -return x_51; +x_50 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_50, 0, x_49); +lean_ctor_set(x_50, 1, x_9); +return x_50; } } } @@ -8688,112 +8795,105 @@ x_1 = lean_mk_string_from_bytes(":=", 2); return x_1; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___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_object* x_7, 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, lean_object* x_15, lean_object* x_16, lean_object* x_17) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___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_object* x_7, 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, lean_object* x_15, lean_object* x_16) { _start: { -uint8_t x_18; -x_18 = lean_nat_dec_le(x_8, x_7); -if (x_18 == 0) +uint8_t x_17; +x_17 = lean_nat_dec_le(x_7, x_6); +if (x_17 == 0) { -lean_object* x_19; uint8_t x_20; -x_19 = lean_unsigned_to_nat(0u); -x_20 = lean_nat_dec_eq(x_6, x_19); -if (x_20 == 0) +lean_object* x_18; uint8_t x_19; +x_18 = lean_unsigned_to_nat(0u); +x_19 = lean_nat_dec_eq(x_5, x_18); +if (x_19 == 0) { -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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_21 = lean_unsigned_to_nat(1u); -x_22 = lean_nat_sub(x_6, x_21); -lean_dec(x_6); -x_23 = l_Lean_instInhabitedName; -x_24 = lean_array_get(x_23, x_2, x_7); -lean_inc(x_24); -x_25 = lean_mk_syntax_ident(x_24); -x_26 = l_Lean_PrettyPrinter_Delaborator_SubExpr_nextExtraPos___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___rarg(x_12, x_13, x_14, x_15, x_16, x_17); -x_27 = lean_ctor_get(x_26, 0); +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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_20 = lean_unsigned_to_nat(1u); +x_21 = lean_nat_sub(x_5, x_20); +lean_dec(x_5); +x_22 = l_Lean_instInhabitedName; +x_23 = lean_array_get(x_22, x_2, x_6); +lean_inc(x_23); +x_24 = lean_mk_syntax_ident(x_23); +x_25 = l_Lean_PrettyPrinter_Delaborator_SubExpr_nextExtraPos___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___rarg(x_11, x_12, x_13, x_14, x_15, x_16); +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_26, 1); +lean_dec(x_25); +lean_inc(x_26); +x_28 = l_Lean_PrettyPrinter_Delaborator_annotatePos(x_26, x_24); +lean_inc(x_23); +x_29 = l_Lean_Name_append(x_1, x_23); +x_30 = l_Lean_instInhabitedExpr; +x_31 = lean_array_get(x_30, x_4, x_6); lean_inc(x_28); -lean_dec(x_26); -lean_inc(x_27); -x_29 = l_Lean_PrettyPrinter_Delaborator_annotatePos(x_27, x_25); -lean_inc(x_24); -x_30 = l_Lean_Name_append(x_1, x_24); -x_31 = l_Lean_instInhabitedExpr; -x_32 = lean_array_get(x_31, x_5, x_7); -lean_inc(x_29); -x_33 = l_Lean_PrettyPrinter_Delaborator_addFieldInfo(x_27, x_30, x_24, x_29, x_32, x_11, x_12, x_13, x_14, x_15, x_16, x_28); -x_34 = lean_ctor_get(x_33, 1); -lean_inc(x_34); -lean_dec(x_33); -lean_inc(x_15); -x_35 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_PrettyPrinter_Delaborator_delab___spec__3___rarg(x_15, x_16, x_34); -x_36 = lean_ctor_get(x_35, 0); +x_32 = l_Lean_PrettyPrinter_Delaborator_addFieldInfo(x_26, x_29, x_23, x_28, x_31, x_10, x_11, x_12, x_13, x_14, x_15, x_27); +x_33 = lean_ctor_get(x_32, 1); +lean_inc(x_33); +lean_dec(x_32); +lean_inc(x_14); +x_34 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_PrettyPrinter_Delaborator_delab___spec__3___rarg(x_14, x_15, x_33); +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_34, 1); lean_inc(x_36); -x_37 = lean_ctor_get(x_35, 1); -lean_inc(x_37); -lean_dec(x_35); -x_38 = lean_box(2); -x_39 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__11; -lean_inc(x_3); -x_40 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -lean_ctor_set(x_40, 2, x_3); -x_41 = l_Lean_PrettyPrinter_Delaborator_maybeAddBlockImplicit___closed__11; -x_42 = lean_array_push(x_41, x_29); -x_43 = lean_array_push(x_42, x_40); -x_44 = l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__2___closed__4; -x_45 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_45, 0, x_38); +lean_dec(x_34); +x_37 = l_Lean_PrettyPrinter_Delaborator_maybeAddBlockImplicit___closed__11; +x_38 = lean_array_push(x_37, x_28); +x_39 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__13; +x_40 = lean_array_push(x_38, x_39); +x_41 = lean_box(2); +x_42 = l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__2___closed__4; +x_43 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +lean_ctor_set(x_43, 2, x_40); +x_44 = l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__2___closed__5; +x_45 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_45, 0, x_35); lean_ctor_set(x_45, 1, x_44); -lean_ctor_set(x_45, 2, x_43); -x_46 = l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__2___closed__5; -x_47 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_47, 0, x_36); -lean_ctor_set(x_47, 1, x_46); -x_48 = l_Lean_Syntax_getArg(x_4, x_7); -x_49 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicit___lambda__4___closed__10; -x_50 = lean_array_push(x_49, x_45); -x_51 = lean_array_push(x_50, x_47); -x_52 = lean_array_push(x_51, x_48); -x_53 = l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__2___closed__2; -x_54 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_54, 0, x_38); -lean_ctor_set(x_54, 1, x_53); -lean_ctor_set(x_54, 2, x_52); -x_55 = lean_array_push(x_10, x_54); -x_56 = lean_nat_add(x_7, x_9); -lean_dec(x_7); -x_6 = x_22; -x_7 = x_56; -x_10 = x_55; -x_17 = x_37; -goto _start; +x_46 = l_Lean_Syntax_getArg(x_3, x_6); +x_47 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicit___lambda__4___closed__10; +x_48 = lean_array_push(x_47, x_43); +x_49 = lean_array_push(x_48, x_45); +x_50 = lean_array_push(x_49, x_46); +x_51 = l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__2___closed__2; +x_52 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_52, 0, x_41); +lean_ctor_set(x_52, 1, x_51); +lean_ctor_set(x_52, 2, x_50); +x_53 = lean_array_push(x_9, x_52); +x_54 = lean_nat_add(x_6, x_8); +lean_dec(x_6); +x_5 = x_21; +x_6 = x_54; +x_9 = x_53; +x_16 = x_36; +goto _start; } else { -lean_object* x_58; -lean_dec(x_15); -lean_dec(x_7); +lean_object* x_56; +lean_dec(x_14); lean_dec(x_6); -lean_dec(x_3); -x_58 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_58, 0, x_10); -lean_ctor_set(x_58, 1, x_17); -return x_58; +lean_dec(x_5); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_9); +lean_ctor_set(x_56, 1, x_16); +return x_56; } } else { -lean_object* x_59; -lean_dec(x_15); -lean_dec(x_7); +lean_object* x_57; +lean_dec(x_14); lean_dec(x_6); -lean_dec(x_3); -x_59 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_59, 0, x_10); -lean_ctor_set(x_59, 1, x_17); -return x_59; +lean_dec(x_5); +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_9); +lean_ctor_set(x_57, 1, x_16); +return x_57; } } } @@ -8911,112 +9011,105 @@ return x_20; } } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__5(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_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, lean_object* x_16, lean_object* x_17) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__5(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_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, lean_object* x_16) { _start: { -uint8_t x_18; -x_18 = lean_nat_dec_le(x_8, x_7); -if (x_18 == 0) +uint8_t x_17; +x_17 = lean_nat_dec_le(x_7, x_6); +if (x_17 == 0) { -lean_object* x_19; uint8_t x_20; -x_19 = lean_unsigned_to_nat(0u); -x_20 = lean_nat_dec_eq(x_6, x_19); -if (x_20 == 0) +lean_object* x_18; uint8_t x_19; +x_18 = lean_unsigned_to_nat(0u); +x_19 = lean_nat_dec_eq(x_5, x_18); +if (x_19 == 0) { -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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_21 = lean_unsigned_to_nat(1u); -x_22 = lean_nat_sub(x_6, x_21); -lean_dec(x_6); -x_23 = l_Lean_instInhabitedName; -x_24 = lean_array_get(x_23, x_2, x_7); -lean_inc(x_24); -x_25 = lean_mk_syntax_ident(x_24); -x_26 = l_Lean_PrettyPrinter_Delaborator_SubExpr_nextExtraPos___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___rarg(x_12, x_13, x_14, x_15, x_16, x_17); -x_27 = lean_ctor_get(x_26, 0); +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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_20 = lean_unsigned_to_nat(1u); +x_21 = lean_nat_sub(x_5, x_20); +lean_dec(x_5); +x_22 = l_Lean_instInhabitedName; +x_23 = lean_array_get(x_22, x_2, x_6); +lean_inc(x_23); +x_24 = lean_mk_syntax_ident(x_23); +x_25 = l_Lean_PrettyPrinter_Delaborator_SubExpr_nextExtraPos___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___rarg(x_11, x_12, x_13, x_14, x_15, x_16); +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_26, 1); +lean_dec(x_25); +lean_inc(x_26); +x_28 = l_Lean_PrettyPrinter_Delaborator_annotatePos(x_26, x_24); +lean_inc(x_23); +x_29 = l_Lean_Name_append(x_1, x_23); +x_30 = l_Lean_instInhabitedExpr; +x_31 = lean_array_get(x_30, x_4, x_6); lean_inc(x_28); -lean_dec(x_26); -lean_inc(x_27); -x_29 = l_Lean_PrettyPrinter_Delaborator_annotatePos(x_27, x_25); -lean_inc(x_24); -x_30 = l_Lean_Name_append(x_1, x_24); -x_31 = l_Lean_instInhabitedExpr; -x_32 = lean_array_get(x_31, x_5, x_7); -lean_inc(x_29); -x_33 = l_Lean_PrettyPrinter_Delaborator_addFieldInfo(x_27, x_30, x_24, x_29, x_32, x_11, x_12, x_13, x_14, x_15, x_16, x_28); -x_34 = lean_ctor_get(x_33, 1); -lean_inc(x_34); -lean_dec(x_33); -lean_inc(x_15); -x_35 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_PrettyPrinter_Delaborator_delab___spec__3___rarg(x_15, x_16, x_34); -x_36 = lean_ctor_get(x_35, 0); +x_32 = l_Lean_PrettyPrinter_Delaborator_addFieldInfo(x_26, x_29, x_23, x_28, x_31, x_10, x_11, x_12, x_13, x_14, x_15, x_27); +x_33 = lean_ctor_get(x_32, 1); +lean_inc(x_33); +lean_dec(x_32); +lean_inc(x_14); +x_34 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_PrettyPrinter_Delaborator_delab___spec__3___rarg(x_14, x_15, x_33); +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_34, 1); lean_inc(x_36); -x_37 = lean_ctor_get(x_35, 1); -lean_inc(x_37); -lean_dec(x_35); -x_38 = lean_box(2); -x_39 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__11; -lean_inc(x_3); -x_40 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -lean_ctor_set(x_40, 2, x_3); -x_41 = l_Lean_PrettyPrinter_Delaborator_maybeAddBlockImplicit___closed__11; -x_42 = lean_array_push(x_41, x_29); -x_43 = lean_array_push(x_42, x_40); -x_44 = l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__2___closed__4; -x_45 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_45, 0, x_38); +lean_dec(x_34); +x_37 = l_Lean_PrettyPrinter_Delaborator_maybeAddBlockImplicit___closed__11; +x_38 = lean_array_push(x_37, x_28); +x_39 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__13; +x_40 = lean_array_push(x_38, x_39); +x_41 = lean_box(2); +x_42 = l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__2___closed__4; +x_43 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +lean_ctor_set(x_43, 2, x_40); +x_44 = l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__2___closed__5; +x_45 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_45, 0, x_35); lean_ctor_set(x_45, 1, x_44); -lean_ctor_set(x_45, 2, x_43); -x_46 = l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__2___closed__5; -x_47 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_47, 0, x_36); -lean_ctor_set(x_47, 1, x_46); -x_48 = l_Lean_Syntax_getArg(x_4, x_7); -x_49 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicit___lambda__4___closed__10; -x_50 = lean_array_push(x_49, x_45); -x_51 = lean_array_push(x_50, x_47); -x_52 = lean_array_push(x_51, x_48); -x_53 = l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__2___closed__2; -x_54 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_54, 0, x_38); -lean_ctor_set(x_54, 1, x_53); -lean_ctor_set(x_54, 2, x_52); -x_55 = lean_array_push(x_10, x_54); -x_56 = lean_nat_add(x_7, x_9); -lean_dec(x_7); -x_6 = x_22; -x_7 = x_56; -x_10 = x_55; -x_17 = x_37; +x_46 = l_Lean_Syntax_getArg(x_3, x_6); +x_47 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicit___lambda__4___closed__10; +x_48 = lean_array_push(x_47, x_43); +x_49 = lean_array_push(x_48, x_45); +x_50 = lean_array_push(x_49, x_46); +x_51 = l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__2___closed__2; +x_52 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_52, 0, x_41); +lean_ctor_set(x_52, 1, x_51); +lean_ctor_set(x_52, 2, x_50); +x_53 = lean_array_push(x_9, x_52); +x_54 = lean_nat_add(x_6, x_8); +lean_dec(x_6); +x_5 = x_21; +x_6 = x_54; +x_9 = x_53; +x_16 = x_36; goto _start; } else { -lean_object* x_58; -lean_dec(x_15); -lean_dec(x_7); +lean_object* x_56; +lean_dec(x_14); lean_dec(x_6); -lean_dec(x_3); -x_58 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_58, 0, x_10); -lean_ctor_set(x_58, 1, x_17); -return x_58; +lean_dec(x_5); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_9); +lean_ctor_set(x_56, 1, x_16); +return x_56; } } else { -lean_object* x_59; -lean_dec(x_15); -lean_dec(x_7); +lean_object* x_57; +lean_dec(x_14); lean_dec(x_6); -lean_dec(x_3); -x_59 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_59, 0, x_10); -lean_ctor_set(x_59, 1, x_17); -return x_59; +lean_dec(x_5); +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_9); +lean_ctor_set(x_57, 1, x_16); +return x_57; } } } @@ -9422,7 +9515,7 @@ x_38 = l_Array_extract___rarg(x_35, x_36, x_37); x_39 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__12; lean_inc(x_7); lean_inc(x_24); -x_40 = l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__5(x_16, x_23, x_39, x_26, x_38, x_24, x_30, x_24, x_25, x_39, x_3, x_4, x_5, x_6, x_7, x_8, x_29); +x_40 = l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__5(x_16, x_23, x_26, x_38, x_24, x_30, x_24, x_25, x_39, x_3, x_4, x_5, x_6, x_7, x_8, x_29); lean_dec(x_24); lean_dec(x_38); lean_dec(x_26); @@ -9457,8 +9550,8 @@ lean_inc(x_49); x_51 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_51, 0, x_49); lean_ctor_set(x_51, 1, x_50); -x_52 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__11; -x_53 = l_Lean_Syntax_SepArray_ofElems(x_52, x_41); +x_52 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__12; +x_53 = l_Lean_mkSepArray(x_41, x_52); lean_dec(x_41); x_54 = l_Array_append___rarg(x_39, x_53); x_55 = lean_box(2); @@ -9536,8 +9629,8 @@ lean_inc(x_84); x_87 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_87, 0, x_84); lean_ctor_set(x_87, 1, x_86); -x_88 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__11; -x_89 = l_Lean_Syntax_SepArray_ofElems(x_88, x_41); +x_88 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__12; +x_89 = l_Lean_mkSepArray(x_41, x_88); lean_dec(x_41); x_90 = l_Array_append___rarg(x_39, x_89); x_91 = lean_box(2); @@ -9690,76 +9783,42 @@ lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__2___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__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, lean_object* x_7, 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, lean_object* x_15, lean_object* x_16) { _start: { -lean_object* x_18; -x_18 = l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -lean_dec(x_16); -lean_dec(x_14); +lean_object* x_17; +x_17 = l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_15); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_9); +lean_dec(x_10); lean_dec(x_8); -lean_dec(x_5); +lean_dec(x_7); lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_18; +return x_17; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__5___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__5___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_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, lean_object* x_16) { _start: { -lean_object* x_18; -x_18 = l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -lean_dec(x_16); -lean_dec(x_14); +lean_object* x_17; +x_17 = l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_15); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_9); +lean_dec(x_10); lean_dec(x_8); -lean_dec(x_5); +lean_dec(x_7); lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_18; +return x_17; } } LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___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) { @@ -9795,7 +9854,174 @@ lean_dec(x_1); return x_10; } } -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabAppImplicit___spec__1(lean_object* x_1, size_t x_2, size_t x_3) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabAppImplicit___spec__3(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_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; +x_11 = !lean_is_exclusive(x_4); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_12 = lean_ctor_get(x_4, 4); +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +lean_dec(x_12); +x_14 = l_Lean_SubExpr_Pos_push(x_13, x_2); +lean_dec(x_13); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_1); +lean_ctor_set(x_15, 1, x_14); +lean_ctor_set(x_4, 4, x_15); +x_16 = lean_apply_7(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t 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; +x_17 = lean_ctor_get(x_4, 0); +x_18 = lean_ctor_get(x_4, 1); +x_19 = lean_ctor_get(x_4, 2); +x_20 = lean_ctor_get(x_4, 3); +x_21 = lean_ctor_get_uint8(x_4, sizeof(void*)*5); +x_22 = lean_ctor_get(x_4, 4); +lean_inc(x_22); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_4); +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +lean_dec(x_22); +x_24 = l_Lean_SubExpr_Pos_push(x_23, x_2); +lean_dec(x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_1); +lean_ctor_set(x_25, 1, x_24); +x_26 = lean_alloc_ctor(0, 5, 1); +lean_ctor_set(x_26, 0, x_17); +lean_ctor_set(x_26, 1, x_18); +lean_ctor_set(x_26, 2, x_19); +lean_ctor_set(x_26, 3, x_20); +lean_ctor_set(x_26, 4, x_25); +lean_ctor_set_uint8(x_26, sizeof(void*)*5, x_21); +x_27 = lean_apply_7(x_3, x_26, x_5, x_6, x_7, x_8, x_9, x_10); +return x_27; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabAppImplicit___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_object* x_7, lean_object* x_8) { +_start: +{ +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_9 = l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_getExprKind___spec__1(x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = l_Lean_Expr_appFn_x21(x_10); +lean_dec(x_10); +x_13 = lean_unsigned_to_nat(0u); +x_14 = l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabAppImplicit___spec__3(x_12, x_13, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_11); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabAppImplicit___spec__4(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; lean_object* x_13; lean_object* x_14; +x_9 = l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_getExprKind___spec__1(x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = l_Lean_Expr_appArg_x21(x_10); +lean_dec(x_10); +x_13 = lean_unsigned_to_nat(1u); +x_14 = l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabAppImplicit___spec__3(x_12, x_13, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_11); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFnArgs___at_Lean_PrettyPrinter_Delaborator_delabAppImplicit___spec__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, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_10 = l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_getExprKind___spec__1(x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = l_Lean_Expr_isApp(x_11); +lean_dec(x_11); +if (x_13 == 0) +{ +lean_object* x_14; +lean_dec(x_2); +x_14 = lean_apply_7(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_12); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; +lean_inc(x_2); +x_15 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFnArgs___at_Lean_PrettyPrinter_Delaborator_delabAppImplicit___spec__1), 9, 2); +lean_closure_set(x_15, 0, x_1); +lean_closure_set(x_15, 1, x_2); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_16 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabAppImplicit___spec__2(x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_12); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = lean_apply_1(x_2, x_17); +x_20 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabAppImplicit___spec__4(x_19, x_3, x_4, x_5, x_6, x_7, x_8, x_18); +return x_20; +} +else +{ +uint8_t x_21; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_21 = !lean_is_exclusive(x_16); +if (x_21 == 0) +{ +return x_16; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_16, 0); +x_23 = lean_ctor_get(x_16, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_16); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +return x_24; +} +} +} +} +} +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabAppImplicit___spec__5(lean_object* x_1, size_t x_2, size_t x_3) { _start: { uint8_t x_4; @@ -11046,7 +11272,7 @@ lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_22 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFnArgs___at_Lean_PrettyPrinter_Delaborator_delabAppExplicit___spec__2(x_20, x_21, x_3, x_4, x_5, x_6, x_7, x_8, x_13); +x_22 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFnArgs___at_Lean_PrettyPrinter_Delaborator_delabAppImplicit___spec__1(x_20, x_21, x_3, x_4, x_5, x_6, x_7, x_8, x_13); if (lean_obj_tag(x_22) == 0) { 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; @@ -11973,7 +12199,7 @@ size_t x_27; size_t x_28; uint8_t x_29; x_27 = 0; x_28 = lean_usize_of_nat(x_19); lean_dec(x_19); -x_29 = l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabAppImplicit___spec__1(x_9, x_27, x_28); +x_29 = l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabAppImplicit___spec__5(x_9, x_27, x_28); if (x_29 == 0) { lean_object* x_30; lean_object* x_31; @@ -12046,7 +12272,7 @@ return x_40; } } } -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabAppImplicit___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabAppImplicit___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; size_t x_5; uint8_t x_6; lean_object* x_7; @@ -12054,7 +12280,7 @@ x_4 = lean_unbox_usize(x_2); lean_dec(x_2); x_5 = lean_unbox_usize(x_3); lean_dec(x_3); -x_6 = l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabAppImplicit___spec__1(x_1, x_4, x_5); +x_6 = l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabAppImplicit___spec__5(x_1, x_4, x_5); lean_dec(x_1); x_7 = lean_box(x_6); return x_7; @@ -14006,7 +14232,31 @@ return x_24; } } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___closed__1() { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = 1; +x_9 = lean_usize_add(x_2, x_8); +x_10 = lean_array_uset(x_7, x_2, x_5); +x_2 = x_9; +x_3 = x_10; +goto _start; +} +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__1() { _start: { lean_object* x_1; @@ -14014,7 +14264,7 @@ x_1 = lean_mk_string_from_bytes("matchAlt", 8); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___closed__2() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__2() { _start: { lean_object* x_1; @@ -14022,7 +14272,7 @@ x_1 = lean_mk_string_from_bytes("|", 1); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___closed__3() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__3() { _start: { lean_object* x_1; @@ -14030,43 +14280,44 @@ x_1 = lean_mk_string_from_bytes("=>", 2); return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9(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) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, size_t x_7, lean_object* x_8) { _start: { -uint8_t x_8; -x_8 = lean_usize_dec_lt(x_6, x_5); -if (x_8 == 0) +uint8_t x_9; +x_9 = lean_usize_dec_lt(x_7, x_6); +if (x_9 == 0) { +lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_7; +return x_8; } else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; size_t x_12; size_t 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; 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; 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; -x_9 = lean_array_uget(x_7, x_6); -x_10 = lean_unsigned_to_nat(0u); -x_11 = lean_array_uset(x_7, x_6, x_10); -x_12 = 1; -x_13 = lean_usize_add(x_6, x_12); -x_14 = lean_ctor_get(x_9, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_9, 1); +lean_object* x_10; lean_object* x_11; lean_object* x_12; size_t x_13; size_t 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; 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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_10 = lean_array_uget(x_8, x_7); +x_11 = lean_unsigned_to_nat(0u); +x_12 = lean_array_uset(x_8, x_7, x_11); +x_13 = 1; +x_14 = lean_usize_add(x_7, x_13); +x_15 = lean_ctor_get(x_10, 0); lean_inc(x_15); -lean_dec(x_9); -x_16 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___closed__1; +x_16 = lean_ctor_get(x_10, 1); +lean_inc(x_16); +lean_dec(x_10); +x_17 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__1; lean_inc(x_2); -x_17 = lean_name_mk_string(x_2, x_16); -x_18 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___closed__2; +x_18 = lean_name_mk_string(x_2, x_17); +x_19 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__2; lean_inc(x_1); -x_19 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_19, 0, x_1); -lean_ctor_set(x_19, 1, x_18); -x_20 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__11; -x_21 = l_Lean_Syntax_SepArray_ofElems(x_20, x_14); -lean_dec(x_14); +x_20 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_20, 0, x_1); +lean_ctor_set(x_20, 1, x_19); +lean_inc(x_5); +x_21 = l_Lean_mkSepArray(x_15, x_5); +lean_dec(x_15); lean_inc(x_4); x_22 = l_Array_append___rarg(x_4, x_21); x_23 = lean_box(2); @@ -14082,65 +14333,66 @@ x_27 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_27, 0, x_23); lean_ctor_set(x_27, 1, x_3); lean_ctor_set(x_27, 2, x_26); -x_28 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___closed__3; +x_28 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__3; lean_inc(x_1); x_29 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_29, 0, x_1); lean_ctor_set(x_29, 1, x_28); x_30 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__14; -x_31 = lean_array_push(x_30, x_19); +x_31 = lean_array_push(x_30, x_20); x_32 = lean_array_push(x_31, x_27); x_33 = lean_array_push(x_32, x_29); -x_34 = lean_array_push(x_33, x_15); +x_34 = lean_array_push(x_33, x_16); x_35 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_35, 0, x_23); -lean_ctor_set(x_35, 1, x_17); +lean_ctor_set(x_35, 1, x_18); lean_ctor_set(x_35, 2, x_34); -x_36 = lean_array_uset(x_11, x_6, x_35); -x_6 = x_13; -x_7 = x_36; +x_36 = lean_array_uset(x_12, x_7, x_35); +x_7 = x_14; +x_8 = x_36; goto _start; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, size_t x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, size_t x_7, size_t x_8, lean_object* x_9) { _start: { -uint8_t x_9; -x_9 = lean_usize_dec_lt(x_7, x_6); -if (x_9 == 0) +uint8_t x_10; +x_10 = lean_usize_dec_lt(x_8, x_7); +if (x_10 == 0) { +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_8; +return x_9; } else { -lean_object* x_10; lean_object* x_11; lean_object* x_12; size_t x_13; size_t 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; 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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_10 = lean_array_uget(x_8, x_7); -x_11 = lean_unsigned_to_nat(0u); -x_12 = lean_array_uset(x_8, x_7, x_11); -x_13 = 1; -x_14 = lean_usize_add(x_7, x_13); -x_15 = lean_ctor_get(x_10, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_10, 1); +lean_object* x_11; lean_object* x_12; lean_object* x_13; size_t x_14; size_t 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; 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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_11 = lean_array_uget(x_9, x_8); +x_12 = lean_unsigned_to_nat(0u); +x_13 = lean_array_uset(x_9, x_8, x_12); +x_14 = 1; +x_15 = lean_usize_add(x_8, x_14); +x_16 = lean_ctor_get(x_11, 0); lean_inc(x_16); -lean_dec(x_10); -x_17 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___closed__1; +x_17 = lean_ctor_get(x_11, 1); +lean_inc(x_17); +lean_dec(x_11); +x_18 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__1; lean_inc(x_2); -x_18 = lean_name_mk_string(x_2, x_17); -x_19 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___closed__2; +x_19 = lean_name_mk_string(x_2, x_18); +x_20 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__2; lean_inc(x_1); -x_20 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_20, 0, x_1); -lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__11; -x_22 = l_Lean_Syntax_SepArray_ofElems(x_21, x_15); -lean_dec(x_15); +x_21 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_21, 0, x_1); +lean_ctor_set(x_21, 1, x_20); +lean_inc(x_6); +x_22 = l_Lean_mkSepArray(x_16, x_6); +lean_dec(x_16); lean_inc(x_4); x_23 = l_Array_append___rarg(x_4, x_22); x_24 = lean_box(2); @@ -14156,32 +14408,32 @@ x_27 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_27, 0, x_24); lean_ctor_set(x_27, 1, x_3); lean_ctor_set(x_27, 2, x_26); -x_28 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___closed__3; +x_28 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__3; lean_inc(x_1); x_29 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_29, 0, x_1); lean_ctor_set(x_29, 1, x_28); x_30 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__14; -x_31 = lean_array_push(x_30, x_20); +x_31 = lean_array_push(x_30, x_21); x_32 = lean_array_push(x_31, x_27); x_33 = lean_array_push(x_32, x_29); -x_34 = lean_array_push(x_33, x_16); +x_34 = lean_array_push(x_33, x_17); x_35 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_35, 0, x_24); -lean_ctor_set(x_35, 1, x_18); +lean_ctor_set(x_35, 1, x_19); lean_ctor_set(x_35, 2, x_34); -x_36 = lean_array_uset(x_12, x_7, x_35); -x_7 = x_14; -x_8 = x_36; +x_36 = lean_array_uset(x_13, x_8, x_35); +x_8 = x_15; +x_9 = x_36; goto _start; } } } -static lean_object* _init_l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__11___closed__1() { +static lean_object* _init_l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__12___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_instInhabitedSyntax; +x_1 = lean_box(0); x_2 = l_Lean_instInhabitedExpr; x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -14189,52 +14441,53 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__11(lean_object* x_1) { +LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__12(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__11___closed__1; +x_2 = l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__12___closed__1; x_3 = lean_panic_fn(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__12(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) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__13(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, size_t x_7, lean_object* x_8) { _start: { -uint8_t x_8; -x_8 = lean_usize_dec_lt(x_6, x_5); -if (x_8 == 0) +uint8_t x_9; +x_9 = lean_usize_dec_lt(x_7, x_6); +if (x_9 == 0) { +lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_7; +return x_8; } else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; size_t x_12; size_t 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; 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; 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; -x_9 = lean_array_uget(x_7, x_6); -x_10 = lean_unsigned_to_nat(0u); -x_11 = lean_array_uset(x_7, x_6, x_10); -x_12 = 1; -x_13 = lean_usize_add(x_6, x_12); -x_14 = lean_ctor_get(x_9, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_9, 1); +lean_object* x_10; lean_object* x_11; lean_object* x_12; size_t x_13; size_t 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; 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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_10 = lean_array_uget(x_8, x_7); +x_11 = lean_unsigned_to_nat(0u); +x_12 = lean_array_uset(x_8, x_7, x_11); +x_13 = 1; +x_14 = lean_usize_add(x_7, x_13); +x_15 = lean_ctor_get(x_10, 0); lean_inc(x_15); -lean_dec(x_9); -x_16 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___closed__1; +x_16 = lean_ctor_get(x_10, 1); +lean_inc(x_16); +lean_dec(x_10); +x_17 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__1; lean_inc(x_2); -x_17 = lean_name_mk_string(x_2, x_16); -x_18 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___closed__2; +x_18 = lean_name_mk_string(x_2, x_17); +x_19 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__2; lean_inc(x_1); -x_19 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_19, 0, x_1); -lean_ctor_set(x_19, 1, x_18); -x_20 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__11; -x_21 = l_Lean_Syntax_SepArray_ofElems(x_20, x_14); -lean_dec(x_14); +x_20 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_20, 0, x_1); +lean_ctor_set(x_20, 1, x_19); +lean_inc(x_5); +x_21 = l_Lean_mkSepArray(x_15, x_5); +lean_dec(x_15); lean_inc(x_4); x_22 = l_Array_append___rarg(x_4, x_21); x_23 = lean_box(2); @@ -14250,65 +14503,66 @@ x_27 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_27, 0, x_23); lean_ctor_set(x_27, 1, x_3); lean_ctor_set(x_27, 2, x_26); -x_28 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___closed__3; +x_28 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__3; lean_inc(x_1); x_29 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_29, 0, x_1); lean_ctor_set(x_29, 1, x_28); x_30 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__14; -x_31 = lean_array_push(x_30, x_19); +x_31 = lean_array_push(x_30, x_20); x_32 = lean_array_push(x_31, x_27); x_33 = lean_array_push(x_32, x_29); -x_34 = lean_array_push(x_33, x_15); +x_34 = lean_array_push(x_33, x_16); x_35 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_35, 0, x_23); -lean_ctor_set(x_35, 1, x_17); +lean_ctor_set(x_35, 1, x_18); lean_ctor_set(x_35, 2, x_34); -x_36 = lean_array_uset(x_11, x_6, x_35); -x_6 = x_13; -x_7 = x_36; +x_36 = lean_array_uset(x_12, x_7, x_35); +x_7 = x_14; +x_8 = x_36; goto _start; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__13(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, size_t x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__14(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, size_t x_7, size_t x_8, lean_object* x_9) { _start: { -uint8_t x_9; -x_9 = lean_usize_dec_lt(x_7, x_6); -if (x_9 == 0) +uint8_t x_10; +x_10 = lean_usize_dec_lt(x_8, x_7); +if (x_10 == 0) { +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_8; +return x_9; } else { -lean_object* x_10; lean_object* x_11; lean_object* x_12; size_t x_13; size_t 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; 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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_10 = lean_array_uget(x_8, x_7); -x_11 = lean_unsigned_to_nat(0u); -x_12 = lean_array_uset(x_8, x_7, x_11); -x_13 = 1; -x_14 = lean_usize_add(x_7, x_13); -x_15 = lean_ctor_get(x_10, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_10, 1); +lean_object* x_11; lean_object* x_12; lean_object* x_13; size_t x_14; size_t 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; 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_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_11 = lean_array_uget(x_9, x_8); +x_12 = lean_unsigned_to_nat(0u); +x_13 = lean_array_uset(x_9, x_8, x_12); +x_14 = 1; +x_15 = lean_usize_add(x_8, x_14); +x_16 = lean_ctor_get(x_11, 0); lean_inc(x_16); -lean_dec(x_10); -x_17 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___closed__1; +x_17 = lean_ctor_get(x_11, 1); +lean_inc(x_17); +lean_dec(x_11); +x_18 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__1; lean_inc(x_2); -x_18 = lean_name_mk_string(x_2, x_17); -x_19 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___closed__2; +x_19 = lean_name_mk_string(x_2, x_18); +x_20 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__2; lean_inc(x_1); -x_20 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_20, 0, x_1); -lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__11; -x_22 = l_Lean_Syntax_SepArray_ofElems(x_21, x_15); -lean_dec(x_15); +x_21 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_21, 0, x_1); +lean_ctor_set(x_21, 1, x_20); +lean_inc(x_6); +x_22 = l_Lean_mkSepArray(x_16, x_6); +lean_dec(x_16); lean_inc(x_4); x_23 = l_Array_append___rarg(x_4, x_22); x_24 = lean_box(2); @@ -14324,23 +14578,23 @@ x_27 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_27, 0, x_24); lean_ctor_set(x_27, 1, x_3); lean_ctor_set(x_27, 2, x_26); -x_28 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___closed__3; +x_28 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__3; lean_inc(x_1); x_29 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_29, 0, x_1); lean_ctor_set(x_29, 1, x_28); x_30 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__14; -x_31 = lean_array_push(x_30, x_20); +x_31 = lean_array_push(x_30, x_21); x_32 = lean_array_push(x_31, x_27); x_33 = lean_array_push(x_32, x_29); -x_34 = lean_array_push(x_33, x_16); +x_34 = lean_array_push(x_33, x_17); x_35 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_35, 0, x_24); -lean_ctor_set(x_35, 1, x_18); +lean_ctor_set(x_35, 1, x_19); lean_ctor_set(x_35, 2, x_34); -x_36 = lean_array_uset(x_12, x_7, x_35); -x_7 = x_14; -x_8 = x_36; +x_36 = lean_array_uset(x_13, x_8, x_35); +x_8 = x_15; +x_9 = x_36; goto _start; } } @@ -16172,7 +16426,7 @@ if (lean_obj_tag(x_135) == 0) { lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; x_136 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__10; -x_137 = l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__11(x_136); +x_137 = l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__12(x_136); x_138 = lean_ctor_get(x_137, 0); lean_inc(x_138); x_139 = lean_ctor_get(x_137, 1); @@ -16290,7 +16544,7 @@ x_32 = lean_array_get_size(x_11); x_33 = lean_usize_of_nat(x_32); lean_dec(x_32); x_34 = 0; -x_35 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_33, x_34, x_11); +x_35 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9(x_33, x_34, x_11); x_36 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__12; x_37 = l_Lean_mkSepArray(x_35, x_36); lean_dec(x_35); @@ -16314,7 +16568,7 @@ x_46 = lean_array_get_size(x_45); x_47 = lean_usize_of_nat(x_46); lean_dec(x_46); x_48 = l_Lean_PrettyPrinter_Delaborator_maybeAddBlockImplicit___closed__7; -x_49 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9(x_28, x_48, x_41, x_38, x_47, x_34, x_45); +x_49 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10(x_28, x_48, x_41, x_38, x_36, x_47, x_34, x_45); x_50 = l_Array_append___rarg(x_38, x_49); x_51 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_51, 0, x_40); @@ -16407,7 +16661,7 @@ x_93 = lean_array_get_size(x_11); x_94 = lean_usize_of_nat(x_93); lean_dec(x_93); x_95 = 0; -x_96 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_94, x_95, x_11); +x_96 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9(x_94, x_95, x_11); x_97 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__12; x_98 = l_Lean_mkSepArray(x_96, x_97); lean_dec(x_96); @@ -16429,7 +16683,7 @@ x_105 = lean_array_get_size(x_104); x_106 = lean_usize_of_nat(x_105); lean_dec(x_105); x_107 = l_Lean_PrettyPrinter_Delaborator_maybeAddBlockImplicit___closed__7; -x_108 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10(x_68, x_107, x_91, x_99, x_89, x_106, x_95, x_104); +x_108 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__11(x_68, x_107, x_91, x_99, x_89, x_97, x_106, x_95, x_104); x_109 = l_Array_append___rarg(x_99, x_108); x_110 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_110, 0, x_86); @@ -16546,7 +16800,7 @@ if (lean_obj_tag(x_268) == 0) { lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; x_269 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__10; -x_270 = l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__11(x_269); +x_270 = l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__12(x_269); x_271 = lean_ctor_get(x_270, 0); lean_inc(x_271); x_272 = lean_ctor_get(x_270, 1); @@ -16664,7 +16918,7 @@ x_165 = lean_array_get_size(x_11); x_166 = lean_usize_of_nat(x_165); lean_dec(x_165); x_167 = 0; -x_168 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_166, x_167, x_11); +x_168 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9(x_166, x_167, x_11); x_169 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__12; x_170 = l_Lean_mkSepArray(x_168, x_169); lean_dec(x_168); @@ -16688,7 +16942,7 @@ x_179 = lean_array_get_size(x_178); x_180 = lean_usize_of_nat(x_179); lean_dec(x_179); x_181 = l_Lean_PrettyPrinter_Delaborator_maybeAddBlockImplicit___closed__7; -x_182 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__12(x_161, x_181, x_174, x_171, x_180, x_167, x_178); +x_182 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__13(x_161, x_181, x_174, x_171, x_169, x_180, x_167, x_178); x_183 = l_Array_append___rarg(x_171, x_182); x_184 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_184, 0, x_173); @@ -16781,7 +17035,7 @@ x_226 = lean_array_get_size(x_11); x_227 = lean_usize_of_nat(x_226); lean_dec(x_226); x_228 = 0; -x_229 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_227, x_228, x_11); +x_229 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9(x_227, x_228, x_11); x_230 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__12; x_231 = l_Lean_mkSepArray(x_229, x_230); lean_dec(x_229); @@ -16803,7 +17057,7 @@ x_238 = lean_array_get_size(x_237); x_239 = lean_usize_of_nat(x_238); lean_dec(x_238); x_240 = l_Lean_PrettyPrinter_Delaborator_maybeAddBlockImplicit___closed__7; -x_241 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__13(x_201, x_240, x_224, x_232, x_222, x_239, x_228, x_237); +x_241 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__14(x_201, x_240, x_224, x_232, x_222, x_230, x_239, x_228, x_237); x_242 = l_Array_append___rarg(x_232, x_241); x_243 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_243, 0, x_219); @@ -17152,16 +17406,16 @@ lean_dec(x_2); return x_9; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___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_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -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 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_10 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9(x_1, x_2, x_3, x_4, x_8, x_9, x_7); -return x_10; +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9(x_4, x_5, x_3); +return x_6; } } LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___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) { @@ -17176,16 +17430,16 @@ x_11 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch_ return x_11; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__12___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_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__11___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_object* x_9) { _start: { -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 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_10 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__12(x_1, x_2, x_3, x_4, x_8, x_9, x_7); -return x_10; +size_t x_10; size_t x_11; lean_object* x_12; +x_10 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_11 = lean_unbox_usize(x_8); +lean_dec(x_8); +x_12 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__11(x_1, x_2, x_3, x_4, x_5, x_6, x_10, x_11, x_9); +return x_12; } } LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__13___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) { @@ -17200,6 +17454,18 @@ x_11 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch_ return x_11; } } +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__14___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_object* x_9) { +_start: +{ +size_t x_10; size_t x_11; lean_object* x_12; +x_10 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_11 = lean_unbox_usize(x_8); +lean_dec(x_8); +x_12 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__14(x_1, x_2, x_3, x_4, x_5, x_6, x_10, x_11, x_9); +return x_12; +} +} LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___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: { @@ -17376,7 +17642,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda_ 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_PrettyPrinter_Delaborator_delabFVar___closed__1; x_2 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__1; -x_3 = lean_unsigned_to_nat(437u); +x_3 = lean_unsigned_to_nat(438u); x_4 = lean_unsigned_to_nat(37u); x_5 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -18542,7 +18808,7 @@ lean_dec(x_18); if (lean_obj_tag(x_20) == 0) { lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = l_Lean_instInhabitedSyntax; +x_21 = lean_box(0); x_22 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicit___lambda__4___closed__2; x_23 = l_Lean_PrettyPrinter_Delaborator_withMDataOptions___rarg(x_21, x_22, x_1, x_2, x_3, x_4, x_5, x_6, x_19); return x_23; @@ -18551,7 +18817,7 @@ else { lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_dec(x_20); -x_24 = l_Lean_instInhabitedSyntax; +x_24 = lean_box(0); x_25 = l_Lean_PrettyPrinter_Delaborator_delabMData___closed__2; x_26 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withMDataExpr___at_Lean_PrettyPrinter_Delaborator_delabMData___spec__1(x_24, x_25, x_1, x_2, x_3, x_4, x_5, x_6, x_19); return x_26; @@ -18576,7 +18842,7 @@ lean_dec(x_29); if (lean_obj_tag(x_31) == 0) { lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = l_Lean_instInhabitedSyntax; +x_32 = lean_box(0); x_33 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicit___lambda__4___closed__2; x_34 = l_Lean_PrettyPrinter_Delaborator_withMDataOptions___rarg(x_32, x_33, x_1, x_2, x_3, x_4, x_5, x_6, x_30); return x_34; @@ -18585,7 +18851,7 @@ else { lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_dec(x_31); -x_35 = l_Lean_instInhabitedSyntax; +x_35 = lean_box(0); x_36 = l_Lean_PrettyPrinter_Delaborator_delabMData___closed__2; x_37 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withMDataExpr___at_Lean_PrettyPrinter_Delaborator_delabMData___spec__1(x_35, x_36, x_1, x_2, x_3, x_4, x_5, x_6, x_30); return x_37; @@ -18594,7 +18860,7 @@ return x_37; else { lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = l_Lean_instInhabitedSyntax; +x_38 = lean_box(0); x_39 = l_Lean_PrettyPrinter_Delaborator_delabMData___closed__3; x_40 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withMDataExpr___at_Lean_PrettyPrinter_Delaborator_delabMData___spec__1(x_38, x_39, x_1, x_2, x_3, x_4, x_5, x_6, x_14); return x_40; @@ -18605,7 +18871,7 @@ else { lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_dec(x_11); -x_41 = l_Lean_instInhabitedSyntax; +x_41 = lean_box(0); x_42 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicit___lambda__4___closed__2; lean_inc(x_6); lean_inc(x_5); @@ -20186,7 +20452,7 @@ x_21 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_21, 0, x_19); lean_ctor_set(x_21, 1, x_20); lean_ctor_set(x_21, 2, x_18); -x_22 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___closed__3; +x_22 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__3; x_23 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_23, 0, x_14); lean_ctor_set(x_23, 1, x_22); @@ -20230,7 +20496,7 @@ x_42 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_42, 0, x_40); lean_ctor_set(x_42, 1, x_41); lean_ctor_set(x_42, 2, x_39); -x_43 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___closed__3; +x_43 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__3; x_44 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_44, 0, x_34); lean_ctor_set(x_44, 1, x_43); @@ -20287,7 +20553,7 @@ x_69 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_69, 0, x_67); lean_ctor_set(x_69, 1, x_68); lean_ctor_set(x_69, 2, x_66); -x_70 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___closed__3; +x_70 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__3; x_71 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_71, 0, x_62); lean_ctor_set(x_71, 1, x_70); @@ -20330,7 +20596,7 @@ x_89 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_89, 0, x_87); lean_ctor_set(x_89, 1, x_88); lean_ctor_set(x_89, 2, x_86); -x_90 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___closed__3; +x_90 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__3; x_91 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_91, 0, x_81); lean_ctor_set(x_91, 1, x_90); @@ -20386,7 +20652,7 @@ x_117 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_117, 0, x_115); lean_ctor_set(x_117, 1, x_116); lean_ctor_set(x_117, 2, x_114); -x_118 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___closed__3; +x_118 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__3; x_119 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_119, 0, x_109); lean_ctor_set(x_119, 1, x_118); @@ -20430,7 +20696,7 @@ x_138 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_138, 0, x_136); lean_ctor_set(x_138, 1, x_137); lean_ctor_set(x_138, 2, x_135); -x_139 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___closed__3; +x_139 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__3; x_140 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_140, 0, x_129); lean_ctor_set(x_140, 1, x_139); @@ -23514,7 +23780,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabLetE___closed__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; x_1 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__1; x_2 = l_Lean_PrettyPrinter_Delaborator_delabLetE___closed__1; -x_3 = lean_unsigned_to_nat(594u); +x_3 = lean_unsigned_to_nat(595u); x_4 = lean_unsigned_to_nat(38u); x_5 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -24246,7 +24512,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabLit___closed__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; x_1 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__1; x_2 = l_Lean_PrettyPrinter_Delaborator_delabLit___closed__1; -x_3 = lean_unsigned_to_nat(607u); +x_3 = lean_unsigned_to_nat(608u); x_4 = lean_unsigned_to_nat(31u); x_5 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -24278,85 +24544,83 @@ uint8_t x_11; x_11 = !lean_is_exclusive(x_8); if (x_11 == 0) { -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_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; x_12 = lean_ctor_get(x_8, 0); lean_dec(x_12); x_13 = lean_ctor_get(x_10, 0); lean_inc(x_13); lean_dec(x_10); x_14 = l_Nat_repr(x_13); -x_15 = l_Lean_numLitKind; -x_16 = lean_box(2); -x_17 = l_Lean_Syntax_mkLit(x_15, x_14, x_16); -lean_ctor_set(x_8, 0, x_17); +x_15 = lean_box(2); +x_16 = l_Lean_Syntax_mkNumLit(x_14, x_15); +lean_ctor_set(x_8, 0, x_16); return x_8; } else { -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_18 = lean_ctor_get(x_8, 1); -lean_inc(x_18); +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_17 = lean_ctor_get(x_8, 1); +lean_inc(x_17); lean_dec(x_8); -x_19 = lean_ctor_get(x_10, 0); -lean_inc(x_19); +x_18 = lean_ctor_get(x_10, 0); +lean_inc(x_18); lean_dec(x_10); -x_20 = l_Nat_repr(x_19); -x_21 = l_Lean_numLitKind; -x_22 = lean_box(2); -x_23 = l_Lean_Syntax_mkLit(x_21, x_20, x_22); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_18); -return x_24; +x_19 = l_Nat_repr(x_18); +x_20 = lean_box(2); +x_21 = l_Lean_Syntax_mkNumLit(x_19, x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_17); +return x_22; } } else { -uint8_t x_25; -x_25 = !lean_is_exclusive(x_8); -if (x_25 == 0) +uint8_t x_23; +x_23 = !lean_is_exclusive(x_8); +if (x_23 == 0) { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_26 = lean_ctor_get(x_8, 0); -lean_dec(x_26); -x_27 = lean_ctor_get(x_10, 0); -lean_inc(x_27); +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_24 = lean_ctor_get(x_8, 0); +lean_dec(x_24); +x_25 = lean_ctor_get(x_10, 0); +lean_inc(x_25); lean_dec(x_10); -x_28 = lean_box(2); -x_29 = l_Lean_Syntax_mkStrLit(x_27, x_28); -lean_dec(x_27); -lean_ctor_set(x_8, 0, x_29); +x_26 = lean_box(2); +x_27 = l_Lean_Syntax_mkStrLit(x_25, x_26); +lean_dec(x_25); +lean_ctor_set(x_8, 0, x_27); return x_8; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_30 = lean_ctor_get(x_8, 1); -lean_inc(x_30); +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_28 = lean_ctor_get(x_8, 1); +lean_inc(x_28); lean_dec(x_8); -x_31 = lean_ctor_get(x_10, 0); -lean_inc(x_31); +x_29 = lean_ctor_get(x_10, 0); +lean_inc(x_29); lean_dec(x_10); -x_32 = lean_box(2); -x_33 = l_Lean_Syntax_mkStrLit(x_31, x_32); -lean_dec(x_31); -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_30); -return x_34; +x_30 = lean_box(2); +x_31 = l_Lean_Syntax_mkStrLit(x_29, x_30); +lean_dec(x_29); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_28); +return x_32; } } } else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_dec(x_9); -x_35 = lean_ctor_get(x_8, 1); -lean_inc(x_35); +x_33 = lean_ctor_get(x_8, 1); +lean_inc(x_33); lean_dec(x_8); -x_36 = l_Lean_PrettyPrinter_Delaborator_delabLit___closed__2; -x_37 = l_panic___at_Lean_PrettyPrinter_Delaborator_delabFVar___spec__1(x_36, x_1, x_2, x_3, x_4, x_5, x_6, x_35); -return x_37; +x_34 = l_Lean_PrettyPrinter_Delaborator_delabLit___closed__2; +x_35 = l_panic___at_Lean_PrettyPrinter_Delaborator_delabFVar___spec__1(x_34, x_1, x_2, x_3, x_4, x_5, x_6, x_33); +return x_35; } } } @@ -24439,23 +24703,30 @@ lean_inc(x_11); lean_dec(x_10); if (lean_obj_tag(x_11) == 0) { -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_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; x_12 = lean_ctor_get(x_11, 0); lean_inc(x_12); lean_dec(x_11); x_13 = l_Nat_repr(x_12); -x_14 = l_Lean_numLitKind; -x_15 = lean_box(2); -x_16 = l_Lean_Syntax_mkLit(x_14, x_13, x_15); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_8); +x_14 = lean_box(2); +x_15 = l_Lean_Syntax_mkNumLit(x_13, x_14); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_8); +return x_16; +} +else +{ +lean_object* x_17; +lean_dec(x_11); +x_17 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_8); return x_17; } +} else { lean_object* x_18; -lean_dec(x_11); +lean_dec(x_10); x_18 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_8); return x_18; } @@ -24463,7 +24734,7 @@ return x_18; else { lean_object* x_19; -lean_dec(x_10); +lean_dec(x_9); x_19 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_8); return x_19; } @@ -24471,19 +24742,11 @@ return x_19; else { lean_object* x_20; -lean_dec(x_9); +lean_dec(x_1); x_20 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_8); return x_20; } } -else -{ -lean_object* x_21; -lean_dec(x_1); -x_21 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_8); -return x_21; -} -} } static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabOfNat___closed__1() { _start: @@ -24650,46 +24913,45 @@ lean_object* x_11; lean_object* x_12; x_11 = l_Nat_repr(x_1); if (x_3 == 0) { -lean_object* x_47; -x_47 = lean_box(0); -x_12 = x_47; -goto block_46; +lean_object* x_44; +x_44 = lean_box(0); +x_12 = x_44; +goto block_43; } else { -lean_object* x_48; uint8_t x_49; -x_48 = lean_string_length(x_11); -x_49 = lean_nat_dec_eq(x_2, x_48); -lean_dec(x_48); -if (x_49 == 0) +lean_object* x_45; uint8_t x_46; +x_45 = lean_string_length(x_11); +x_46 = lean_nat_dec_eq(x_2, x_45); +lean_dec(x_45); +if (x_46 == 0) { -lean_object* x_50; -x_50 = lean_box(0); -x_12 = x_50; -goto block_46; +lean_object* x_47; +x_47 = lean_box(0); +x_12 = x_47; +goto block_43; } else { -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_dec(x_2); -x_51 = l_Lean_PrettyPrinter_Delaborator_delabOfScientific___lambda__1___closed__5; -x_52 = lean_string_append(x_51, x_11); +x_48 = l_Lean_PrettyPrinter_Delaborator_delabOfScientific___lambda__1___closed__5; +x_49 = lean_string_append(x_48, x_11); lean_dec(x_11); -x_53 = l_Lean_scientificLitKind; -x_54 = lean_box(2); -x_55 = l_Lean_Syntax_mkLit(x_53, x_52, x_54); -x_56 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_56, 0, x_55); -lean_ctor_set(x_56, 1, x_10); -return x_56; +x_50 = lean_box(2); +x_51 = l_Lean_Syntax_mkScientificLit(x_49, x_50); +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_10); +return x_52; } } -block_46: +block_43: { lean_dec(x_12); if (x_3 == 0) { -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; lean_object* x_20; lean_object* x_21; lean_object* x_22; +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; lean_object* x_20; lean_object* x_21; x_13 = l_Lean_PrettyPrinter_Delaborator_delabOfScientific___lambda__1___closed__1; x_14 = lean_string_append(x_11, x_13); x_15 = l_Nat_repr(x_2); @@ -24697,60 +24959,57 @@ x_16 = l_Lean_PrettyPrinter_Delaborator_delabOfScientific___lambda__1___closed__ x_17 = lean_string_append(x_14, x_16); x_18 = lean_string_append(x_17, x_15); lean_dec(x_15); -x_19 = l_Lean_scientificLitKind; -x_20 = lean_box(2); -x_21 = l_Lean_Syntax_mkLit(x_19, x_18, x_20); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_10); -return x_22; +x_19 = lean_box(2); +x_20 = l_Lean_Syntax_mkScientificLit(x_18, x_19); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_10); +return x_21; } else { -lean_object* x_23; uint8_t x_24; -x_23 = lean_string_length(x_11); -x_24 = lean_nat_dec_lt(x_2, x_23); -if (x_24 == 0) +lean_object* x_22; uint8_t x_23; +x_22 = lean_string_length(x_11); +x_23 = lean_nat_dec_lt(x_2, x_22); +if (x_23 == 0) { -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_object* x_32; lean_object* x_33; lean_object* x_34; -lean_dec(x_23); -x_25 = l_Lean_PrettyPrinter_Delaborator_delabOfScientific___lambda__1___closed__1; -x_26 = lean_string_append(x_11, x_25); -x_27 = l_Nat_repr(x_2); -x_28 = l_Lean_PrettyPrinter_Delaborator_delabOfScientific___lambda__1___closed__3; -x_29 = lean_string_append(x_26, x_28); -x_30 = lean_string_append(x_29, x_27); -lean_dec(x_27); -x_31 = l_Lean_scientificLitKind; -x_32 = lean_box(2); -x_33 = l_Lean_Syntax_mkLit(x_31, x_30, x_32); -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_10); -return x_34; +lean_object* 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; lean_object* x_32; +lean_dec(x_22); +x_24 = l_Lean_PrettyPrinter_Delaborator_delabOfScientific___lambda__1___closed__1; +x_25 = lean_string_append(x_11, x_24); +x_26 = l_Nat_repr(x_2); +x_27 = l_Lean_PrettyPrinter_Delaborator_delabOfScientific___lambda__1___closed__3; +x_28 = lean_string_append(x_25, x_27); +x_29 = lean_string_append(x_28, x_26); +lean_dec(x_26); +x_30 = lean_box(2); +x_31 = l_Lean_Syntax_mkScientificLit(x_29, x_30); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_10); +return x_32; } else { -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; lean_object* x_44; lean_object* x_45; -x_35 = lean_nat_sub(x_23, x_2); +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; +x_33 = lean_nat_sub(x_22, x_2); lean_dec(x_2); -x_36 = lean_unsigned_to_nat(0u); -x_37 = lean_string_utf8_extract(x_11, x_36, x_35); -x_38 = lean_string_utf8_extract(x_11, x_35, x_23); -lean_dec(x_23); -lean_dec(x_35); +x_34 = lean_unsigned_to_nat(0u); +x_35 = lean_string_utf8_extract(x_11, x_34, x_33); +x_36 = lean_string_utf8_extract(x_11, x_33, x_22); +lean_dec(x_22); +lean_dec(x_33); lean_dec(x_11); -x_39 = l_Lean_PrettyPrinter_Delaborator_delabOfScientific___lambda__1___closed__4; -x_40 = lean_string_append(x_37, x_39); -x_41 = lean_string_append(x_40, x_38); -lean_dec(x_38); -x_42 = l_Lean_scientificLitKind; -x_43 = lean_box(2); -x_44 = l_Lean_Syntax_mkLit(x_42, x_41, x_43); -x_45 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_45, 0, x_44); -lean_ctor_set(x_45, 1, x_10); -return x_45; +x_37 = l_Lean_PrettyPrinter_Delaborator_delabOfScientific___lambda__1___closed__4; +x_38 = lean_string_append(x_35, x_37); +x_39 = lean_string_append(x_38, x_36); +lean_dec(x_36); +x_40 = lean_box(2); +x_41 = l_Lean_Syntax_mkScientificLit(x_39, x_40); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_10); +return x_42; } } } @@ -25638,7 +25897,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabProj___closed__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; x_1 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__1; x_2 = l_Lean_PrettyPrinter_Delaborator_delabProj___closed__1; -x_3 = lean_unsigned_to_nat(646u); +x_3 = lean_unsigned_to_nat(647u); x_4 = lean_unsigned_to_nat(38u); x_5 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -25649,7 +25908,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabProj___closed__3 _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("proj", 4); +x_1 = lean_mk_string_from_bytes("fieldIdx", 8); return x_1; } } @@ -25657,12 +25916,30 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabProj___closed__4 _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PrettyPrinter_Delaborator_maybeAddBlockImplicit___closed__7; +x_1 = lean_box(0); x_2 = l_Lean_PrettyPrinter_Delaborator_delabProj___closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabProj___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("proj", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabProj___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_PrettyPrinter_Delaborator_maybeAddBlockImplicit___closed__7; +x_2 = l_Lean_PrettyPrinter_Delaborator_delabProj___closed__5; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabProj(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: { @@ -25679,7 +25956,7 @@ lean_dec(x_8); x_11 = lean_ctor_get(x_9, 1); lean_inc(x_11); lean_dec(x_9); -x_12 = l_Lean_instInhabitedSyntax; +x_12 = lean_box(0); x_13 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicit___lambda__4___closed__2; lean_inc(x_6); lean_inc(x_5); @@ -25696,7 +25973,7 @@ x_17 = lean_unsigned_to_nat(1u); x_18 = lean_nat_add(x_11, x_17); lean_dec(x_11); x_19 = l_Nat_repr(x_18); -x_20 = l_Lean_fieldIdxKind; +x_20 = l_Lean_PrettyPrinter_Delaborator_delabProj___closed__4; x_21 = lean_box(2); x_22 = l_Lean_Syntax_mkLit(x_20, x_19, x_21); x_23 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_PrettyPrinter_Delaborator_delab___spec__3___rarg(x_5, x_6, x_16); @@ -25714,7 +25991,7 @@ x_28 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicit___lambda__4___closed__1 x_29 = lean_array_push(x_28, x_15); x_30 = lean_array_push(x_29, x_27); x_31 = lean_array_push(x_30, x_22); -x_32 = l_Lean_PrettyPrinter_Delaborator_delabProj___closed__4; +x_32 = l_Lean_PrettyPrinter_Delaborator_delabProj___closed__6; x_33 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_33, 0, x_21); lean_ctor_set(x_33, 1, x_32); @@ -25738,7 +26015,7 @@ x_38 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicit___lambda__4___closed__1 x_39 = lean_array_push(x_38, x_15); x_40 = lean_array_push(x_39, x_37); x_41 = lean_array_push(x_40, x_22); -x_42 = l_Lean_PrettyPrinter_Delaborator_delabProj___closed__4; +x_42 = l_Lean_PrettyPrinter_Delaborator_delabProj___closed__6; x_43 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_43, 0, x_21); lean_ctor_set(x_43, 1, x_42); @@ -25793,7 +26070,7 @@ static lean_object* _init_l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabPro { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_PrettyPrinter_Delaborator_delabProj___closed__3; +x_2 = l_Lean_PrettyPrinter_Delaborator_delabProj___closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } @@ -25910,7 +26187,7 @@ lean_dec(x_27); lean_dec(x_24); if (x_28 == 0) { -lean_object* x_123; uint8_t x_124; +lean_object* x_86; uint8_t x_87; lean_dec(x_25); lean_dec(x_12); lean_dec(x_7); @@ -25919,32 +26196,32 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_123 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_15); -x_124 = !lean_is_exclusive(x_123); -if (x_124 == 0) +x_86 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_15); +x_87 = !lean_is_exclusive(x_86); +if (x_87 == 0) { -return x_123; +return x_86; } else { -lean_object* x_125; lean_object* x_126; lean_object* x_127; -x_125 = lean_ctor_get(x_123, 0); -x_126 = lean_ctor_get(x_123, 1); -lean_inc(x_126); -lean_inc(x_125); -lean_dec(x_123); -x_127 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_127, 0, x_125); -lean_ctor_set(x_127, 1, x_126); -return x_127; +lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_88 = lean_ctor_get(x_86, 0); +x_89 = lean_ctor_get(x_86, 1); +lean_inc(x_89); +lean_inc(x_88); +lean_dec(x_86); +x_90 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_90, 0, x_88); +lean_ctor_set(x_90, 1, x_89); +return x_90; } } else { x_29 = x_15; -goto block_122; +goto block_85; } -block_122: +block_85: { lean_object* x_30; lean_object* x_31; x_30 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicit___closed__1; @@ -25957,18 +26234,64 @@ lean_inc(x_2); x_31 = l_Lean_PrettyPrinter_Delaborator_getPPOption(x_30, x_2, x_3, x_4, x_5, x_6, x_7, x_29); if (lean_obj_tag(x_31) == 0) { -lean_object* x_32; uint8_t x_33; +lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_74; x_32 = lean_ctor_get(x_31, 0); lean_inc(x_32); -x_33 = lean_unbox(x_32); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec(x_31); +x_74 = lean_unbox(x_32); lean_dec(x_32); -if (x_33 == 0) +if (x_74 == 0) { -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_dec(x_25); -x_34 = lean_ctor_get(x_31, 1); -lean_inc(x_34); -lean_dec(x_31); +x_34 = x_33; +goto block_73; +} +else +{ +uint8_t x_75; +x_75 = lean_nat_dec_eq(x_25, x_23); +lean_dec(x_25); +if (x_75 == 0) +{ +lean_object* x_76; uint8_t x_77; +lean_dec(x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_76 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_33); +x_77 = !lean_is_exclusive(x_76); +if (x_77 == 0) +{ +return x_76; +} +else +{ +lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_78 = lean_ctor_get(x_76, 0); +x_79 = lean_ctor_get(x_76, 1); +lean_inc(x_79); +lean_inc(x_78); +lean_dec(x_76); +x_80 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_80, 0, x_78); +lean_ctor_set(x_80, 1, x_79); +return x_80; +} +} +else +{ +x_34 = x_33; +goto block_73; +} +} +block_73: +{ +lean_object* x_35; lean_object* x_36; x_35 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicit___lambda__4___closed__2; lean_inc(x_7); lean_inc(x_6); @@ -26000,213 +26323,75 @@ x_48 = lean_array_push(x_47, x_37); x_49 = lean_array_push(x_48, x_43); x_50 = lean_array_push(x_49, x_46); x_51 = lean_box(2); -x_52 = l_Lean_PrettyPrinter_Delaborator_delabProj___closed__4; -x_53 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_53, 0, x_51); -lean_ctor_set(x_53, 1, x_52); -lean_ctor_set(x_53, 2, x_50); -lean_ctor_set(x_39, 0, x_53); -return x_39; -} -else -{ -lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; lean_object* x_67; lean_object* x_68; -x_54 = lean_ctor_get(x_39, 0); -x_55 = lean_ctor_get(x_39, 1); -lean_inc(x_55); -lean_inc(x_54); -lean_dec(x_39); -x_56 = l_Lean_PrettyPrinter_Delaborator_delabOfScientific___lambda__1___closed__4; -x_57 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_57, 0, x_54); -lean_ctor_set(x_57, 1, x_56); -x_58 = lean_box(0); -x_59 = lean_name_mk_string(x_58, x_12); -x_60 = lean_mk_syntax_ident(x_59); -x_61 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicit___lambda__4___closed__10; -x_62 = lean_array_push(x_61, x_37); -x_63 = lean_array_push(x_62, x_57); -x_64 = lean_array_push(x_63, x_60); -x_65 = lean_box(2); -x_66 = l_Lean_PrettyPrinter_Delaborator_delabProj___closed__4; -x_67 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_67, 0, x_65); -lean_ctor_set(x_67, 1, x_66); -lean_ctor_set(x_67, 2, x_64); -x_68 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_68, 0, x_67); -lean_ctor_set(x_68, 1, x_55); -return x_68; -} -} -else -{ -uint8_t x_69; -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -x_69 = !lean_is_exclusive(x_36); -if (x_69 == 0) -{ -return x_36; -} -else -{ -lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_70 = lean_ctor_get(x_36, 0); -x_71 = lean_ctor_get(x_36, 1); -lean_inc(x_71); -lean_inc(x_70); -lean_dec(x_36); -x_72 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_72, 0, x_70); -lean_ctor_set(x_72, 1, x_71); -return x_72; -} -} -} -else -{ -lean_object* x_73; uint8_t x_74; -x_73 = lean_ctor_get(x_31, 1); -lean_inc(x_73); -lean_dec(x_31); -x_74 = lean_nat_dec_eq(x_25, x_23); -lean_dec(x_25); -if (x_74 == 0) -{ -lean_object* x_75; uint8_t x_76; -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_75 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_73); -x_76 = !lean_is_exclusive(x_75); -if (x_76 == 0) -{ -return x_75; -} -else -{ -lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_77 = lean_ctor_get(x_75, 0); -x_78 = lean_ctor_get(x_75, 1); -lean_inc(x_78); -lean_inc(x_77); -lean_dec(x_75); -x_79 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_79, 0, x_77); -lean_ctor_set(x_79, 1, x_78); -return x_79; -} -} -else -{ -lean_object* x_80; lean_object* x_81; -x_80 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicit___lambda__4___closed__2; -lean_inc(x_7); -lean_inc(x_6); -x_81 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__1(x_80, x_2, x_3, x_4, x_5, x_6, x_7, x_73); -if (lean_obj_tag(x_81) == 0) -{ -lean_object* x_82; lean_object* x_83; lean_object* x_84; uint8_t x_85; -x_82 = lean_ctor_get(x_81, 0); -lean_inc(x_82); -x_83 = lean_ctor_get(x_81, 1); -lean_inc(x_83); -lean_dec(x_81); -x_84 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_PrettyPrinter_Delaborator_delab___spec__3___rarg(x_6, x_7, x_83); -lean_dec(x_7); -x_85 = !lean_is_exclusive(x_84); -if (x_85 == 0) -{ -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; -x_86 = lean_ctor_get(x_84, 0); -x_87 = l_Lean_PrettyPrinter_Delaborator_delabOfScientific___lambda__1___closed__4; -x_88 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_88, 0, x_86); -lean_ctor_set(x_88, 1, x_87); -x_89 = lean_box(0); -x_90 = lean_name_mk_string(x_89, x_12); -x_91 = lean_mk_syntax_ident(x_90); -x_92 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicit___lambda__4___closed__10; -x_93 = lean_array_push(x_92, x_82); -x_94 = lean_array_push(x_93, x_88); -x_95 = lean_array_push(x_94, x_91); -x_96 = lean_box(2); -x_97 = l_Lean_PrettyPrinter_Delaborator_delabProj___closed__4; -x_98 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_98, 0, x_96); -lean_ctor_set(x_98, 1, x_97); -lean_ctor_set(x_98, 2, x_95); -lean_ctor_set(x_84, 0, x_98); -return x_84; +x_52 = l_Lean_PrettyPrinter_Delaborator_delabProj___closed__6; +x_53 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_52); +lean_ctor_set(x_53, 2, x_50); +lean_ctor_set(x_39, 0, x_53); +return x_39; } else { -lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; -x_99 = lean_ctor_get(x_84, 0); -x_100 = lean_ctor_get(x_84, 1); -lean_inc(x_100); -lean_inc(x_99); -lean_dec(x_84); -x_101 = l_Lean_PrettyPrinter_Delaborator_delabOfScientific___lambda__1___closed__4; -x_102 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_102, 0, x_99); -lean_ctor_set(x_102, 1, x_101); -x_103 = lean_box(0); -x_104 = lean_name_mk_string(x_103, x_12); -x_105 = lean_mk_syntax_ident(x_104); -x_106 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicit___lambda__4___closed__10; -x_107 = lean_array_push(x_106, x_82); -x_108 = lean_array_push(x_107, x_102); -x_109 = lean_array_push(x_108, x_105); -x_110 = lean_box(2); -x_111 = l_Lean_PrettyPrinter_Delaborator_delabProj___closed__4; -x_112 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_112, 0, x_110); -lean_ctor_set(x_112, 1, x_111); -lean_ctor_set(x_112, 2, x_109); -x_113 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_113, 0, x_112); -lean_ctor_set(x_113, 1, x_100); -return x_113; +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; lean_object* x_67; lean_object* x_68; +x_54 = lean_ctor_get(x_39, 0); +x_55 = lean_ctor_get(x_39, 1); +lean_inc(x_55); +lean_inc(x_54); +lean_dec(x_39); +x_56 = l_Lean_PrettyPrinter_Delaborator_delabOfScientific___lambda__1___closed__4; +x_57 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_57, 0, x_54); +lean_ctor_set(x_57, 1, x_56); +x_58 = lean_box(0); +x_59 = lean_name_mk_string(x_58, x_12); +x_60 = lean_mk_syntax_ident(x_59); +x_61 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicit___lambda__4___closed__10; +x_62 = lean_array_push(x_61, x_37); +x_63 = lean_array_push(x_62, x_57); +x_64 = lean_array_push(x_63, x_60); +x_65 = lean_box(2); +x_66 = l_Lean_PrettyPrinter_Delaborator_delabProj___closed__6; +x_67 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_67, 0, x_65); +lean_ctor_set(x_67, 1, x_66); +lean_ctor_set(x_67, 2, x_64); +x_68 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_68, 0, x_67); +lean_ctor_set(x_68, 1, x_55); +return x_68; } } else { -uint8_t x_114; +uint8_t x_69; lean_dec(x_12); lean_dec(x_7); lean_dec(x_6); -x_114 = !lean_is_exclusive(x_81); -if (x_114 == 0) +x_69 = !lean_is_exclusive(x_36); +if (x_69 == 0) { -return x_81; +return x_36; } else { -lean_object* x_115; lean_object* x_116; lean_object* x_117; -x_115 = lean_ctor_get(x_81, 0); -x_116 = lean_ctor_get(x_81, 1); -lean_inc(x_116); -lean_inc(x_115); -lean_dec(x_81); -x_117 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_117, 0, x_115); -lean_ctor_set(x_117, 1, x_116); -return x_117; -} +lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_70 = lean_ctor_get(x_36, 0); +x_71 = lean_ctor_get(x_36, 1); +lean_inc(x_71); +lean_inc(x_70); +lean_dec(x_36); +x_72 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_72, 0, x_70); +lean_ctor_set(x_72, 1, x_71); +return x_72; } } } } else { -uint8_t x_118; +uint8_t x_81; lean_dec(x_25); lean_dec(x_12); lean_dec(x_7); @@ -26215,30 +26400,30 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_118 = !lean_is_exclusive(x_31); -if (x_118 == 0) +x_81 = !lean_is_exclusive(x_31); +if (x_81 == 0) { return x_31; } else { -lean_object* x_119; lean_object* x_120; lean_object* x_121; -x_119 = lean_ctor_get(x_31, 0); -x_120 = lean_ctor_get(x_31, 1); -lean_inc(x_120); -lean_inc(x_119); +lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_82 = lean_ctor_get(x_31, 0); +x_83 = lean_ctor_get(x_31, 1); +lean_inc(x_83); +lean_inc(x_82); lean_dec(x_31); -x_121 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_121, 0, x_119); -lean_ctor_set(x_121, 1, x_120); -return x_121; +x_84 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_84, 0, x_82); +lean_ctor_set(x_84, 1, x_83); +return x_84; } } } } else { -lean_object* x_128; uint8_t x_129; +lean_object* x_91; uint8_t x_92; lean_dec(x_21); lean_dec(x_12); lean_dec(x_7); @@ -26247,31 +26432,31 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_128 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_15); -x_129 = !lean_is_exclusive(x_128); -if (x_129 == 0) +x_91 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_15); +x_92 = !lean_is_exclusive(x_91); +if (x_92 == 0) { -return x_128; +return x_91; } else { -lean_object* x_130; lean_object* x_131; lean_object* x_132; -x_130 = lean_ctor_get(x_128, 0); -x_131 = lean_ctor_get(x_128, 1); -lean_inc(x_131); -lean_inc(x_130); -lean_dec(x_128); -x_132 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_132, 0, x_130); -lean_ctor_set(x_132, 1, x_131); -return x_132; +lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_93 = lean_ctor_get(x_91, 0); +x_94 = lean_ctor_get(x_91, 1); +lean_inc(x_94); +lean_inc(x_93); +lean_dec(x_91); +x_95 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_95, 0, x_93); +lean_ctor_set(x_95, 1, x_94); +return x_95; } } } } else { -lean_object* x_133; +lean_object* x_96; lean_dec(x_11); lean_dec(x_7); lean_dec(x_6); @@ -26279,13 +26464,13 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_133 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_8); -return x_133; +x_96 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_8); +return x_96; } } else { -lean_object* x_134; +lean_object* x_97; lean_dec(x_10); lean_dec(x_7); lean_dec(x_6); @@ -26293,21 +26478,21 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_134 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_8); -return x_134; +x_97 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_8); +return x_97; } } else { -lean_object* x_135; +lean_object* x_98; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_135 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_8); -return x_135; +x_98 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_8); +return x_98; } } } @@ -26469,7 +26654,7 @@ return x_1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDIte_delabBranch(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; uint8_t x_12; +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; x_9 = l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_getExprKind___spec__1(x_2, x_3, x_4, x_5, x_6, x_7, x_8); x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); @@ -26480,7 +26665,7 @@ x_12 = l_Lean_Expr_isLambda(x_10); lean_dec(x_10); if (x_12 == 0) { -lean_object* x_13; uint8_t x_14; +lean_object* x_31; uint8_t x_32; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); @@ -26488,96 +26673,101 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_13 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_11); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) +x_31 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_11); +x_32 = !lean_is_exclusive(x_31); +if (x_32 == 0) { -return x_13; +return x_31; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_ctor_get(x_13, 1); -lean_inc(x_16); -lean_inc(x_15); -lean_dec(x_13); -x_17 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_17, 0, x_15); -lean_ctor_set(x_17, 1, x_16); -return x_17; +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_31, 0); +x_34 = lean_ctor_get(x_31, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_31); +x_35 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; } } else { +x_13 = x_11; +goto block_30; +} +block_30: +{ if (lean_obj_tag(x_1) == 0) { -lean_object* x_18; lean_object* x_19; -x_18 = l_Lean_PrettyPrinter_Delaborator_delabDIte_delabBranch___closed__1; -x_19 = l_Lean_PrettyPrinter_Delaborator_withBindingBodyUnusedName___rarg(x_18, x_2, x_3, x_4, x_5, x_6, x_7, x_11); -return x_19; +lean_object* x_14; lean_object* x_15; +x_14 = l_Lean_PrettyPrinter_Delaborator_delabDIte_delabBranch___closed__1; +x_15 = l_Lean_PrettyPrinter_Delaborator_withBindingBodyUnusedName___rarg(x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_13); +return x_15; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_1, 0); -lean_inc(x_20); +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_1, 0); +lean_inc(x_16); lean_dec(x_1); -x_21 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicit___lambda__4___closed__2; -lean_inc(x_20); -x_22 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__2(x_20, x_21, x_2, x_3, x_4, x_5, x_6, x_7, x_11); -if (lean_obj_tag(x_22) == 0) +x_17 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicit___lambda__4___closed__2; +lean_inc(x_16); +x_18 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__2(x_16, x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_13); +if (lean_obj_tag(x_18) == 0) { -uint8_t x_23; -x_23 = !lean_is_exclusive(x_22); -if (x_23 == 0) +uint8_t x_19; +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) { -lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_22, 0); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_20); -lean_ctor_set(x_22, 0, x_25); -return x_22; +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_18, 0); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_16); +lean_ctor_set(x_18, 0, x_21); +return x_18; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_26 = lean_ctor_get(x_22, 0); -x_27 = lean_ctor_get(x_22, 1); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_22); -x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_20); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_27); -return x_29; +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_22 = lean_ctor_get(x_18, 0); +x_23 = lean_ctor_get(x_18, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_18); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_16); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_23); +return x_25; } } else { -uint8_t x_30; -lean_dec(x_20); -x_30 = !lean_is_exclusive(x_22); -if (x_30 == 0) +uint8_t x_26; +lean_dec(x_16); +x_26 = !lean_is_exclusive(x_18); +if (x_26 == 0) { -return x_22; +return x_18; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_22, 0); -x_32 = lean_ctor_get(x_22, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_22); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -return x_33; +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_18, 0); +x_28 = lean_ctor_get(x_18, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_18); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; } } } @@ -29601,7 +29791,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabDoElems___closed 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_PrettyPrinter_Delaborator_delabFVar___closed__1; x_2 = l_Lean_PrettyPrinter_Delaborator_delabDoElems___closed__5; -x_3 = lean_unsigned_to_nat(748u); +x_3 = lean_unsigned_to_nat(749u); x_4 = lean_unsigned_to_nat(40u); x_5 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -30986,7 +31176,7 @@ uint8_t x_10; x_10 = !lean_is_exclusive(x_9); if (x_10 == 0) { -lean_object* x_11; uint8_t 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; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +lean_object* x_11; uint8_t 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; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; x_11 = lean_ctor_get(x_9, 0); x_12 = 1; x_13 = l_Lean_Name_toString(x_11, x_12); @@ -30995,70 +31185,68 @@ x_15 = lean_string_append(x_14, x_13); lean_dec(x_13); x_16 = l_Lean_PrettyPrinter_Delaborator_delabOfScientific___lambda__1___closed__2; x_17 = lean_string_append(x_15, x_16); -x_18 = l_Lean_nameLitKind; -x_19 = lean_box(2); -x_20 = l_Lean_Syntax_mkLit(x_18, x_17, x_19); -x_21 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__6; -x_22 = lean_array_push(x_21, x_20); -x_23 = l_Lean_PrettyPrinter_Delaborator_delabNameMkStr___lambda__1___closed__2; -x_24 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_24, 0, x_19); -lean_ctor_set(x_24, 1, x_23); -lean_ctor_set(x_24, 2, x_22); -lean_ctor_set(x_9, 0, x_24); +x_18 = lean_box(2); +x_19 = l_Lean_Syntax_mkNameLit(x_17, x_18); +x_20 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__6; +x_21 = lean_array_push(x_20, x_19); +x_22 = l_Lean_PrettyPrinter_Delaborator_delabNameMkStr___lambda__1___closed__2; +x_23 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_23, 0, x_18); +lean_ctor_set(x_23, 1, x_22); +lean_ctor_set(x_23, 2, x_21); +lean_ctor_set(x_9, 0, x_23); return x_9; } else { -lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; 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; -x_25 = lean_ctor_get(x_9, 0); -x_26 = lean_ctor_get(x_9, 1); -lean_inc(x_26); +lean_object* x_24; lean_object* x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; 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; +x_24 = lean_ctor_get(x_9, 0); +x_25 = lean_ctor_get(x_9, 1); lean_inc(x_25); +lean_inc(x_24); lean_dec(x_9); -x_27 = 1; -x_28 = l_Lean_Name_toString(x_25, x_27); -x_29 = l_Lean_PrettyPrinter_Delaborator_delabNameMkStr___lambda__1___closed__3; -x_30 = lean_string_append(x_29, x_28); -lean_dec(x_28); -x_31 = l_Lean_PrettyPrinter_Delaborator_delabOfScientific___lambda__1___closed__2; -x_32 = lean_string_append(x_30, x_31); -x_33 = l_Lean_nameLitKind; -x_34 = lean_box(2); -x_35 = l_Lean_Syntax_mkLit(x_33, x_32, x_34); -x_36 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__6; -x_37 = lean_array_push(x_36, x_35); -x_38 = l_Lean_PrettyPrinter_Delaborator_delabNameMkStr___lambda__1___closed__2; -x_39 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_39, 0, x_34); -lean_ctor_set(x_39, 1, x_38); -lean_ctor_set(x_39, 2, x_37); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_26); -return x_40; +x_26 = 1; +x_27 = l_Lean_Name_toString(x_24, x_26); +x_28 = l_Lean_PrettyPrinter_Delaborator_delabNameMkStr___lambda__1___closed__3; +x_29 = lean_string_append(x_28, x_27); +lean_dec(x_27); +x_30 = l_Lean_PrettyPrinter_Delaborator_delabOfScientific___lambda__1___closed__2; +x_31 = lean_string_append(x_29, x_30); +x_32 = lean_box(2); +x_33 = l_Lean_Syntax_mkNameLit(x_31, x_32); +x_34 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__6; +x_35 = lean_array_push(x_34, x_33); +x_36 = l_Lean_PrettyPrinter_Delaborator_delabNameMkStr___lambda__1___closed__2; +x_37 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_37, 0, x_32); +lean_ctor_set(x_37, 1, x_36); +lean_ctor_set(x_37, 2, x_35); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_25); +return x_38; } } else { -uint8_t x_41; -x_41 = !lean_is_exclusive(x_9); -if (x_41 == 0) +uint8_t x_39; +x_39 = !lean_is_exclusive(x_9); +if (x_39 == 0) { return x_9; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_9, 0); -x_43 = lean_ctor_get(x_9, 1); -lean_inc(x_43); -lean_inc(x_42); +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_9, 0); +x_41 = lean_ctor_get(x_9, 1); +lean_inc(x_41); +lean_inc(x_40); lean_dec(x_9); -x_44 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_44, 1, x_43); -return x_44; +x_42 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_41); +return x_42; } } } @@ -31654,14 +31842,14 @@ l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2_ lean_mark_persistent(l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__3); l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__4 = _init_l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__4(); lean_mark_persistent(l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__4); -l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___closed__1 = _init_l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___closed__1(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___closed__1); -l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___closed__2 = _init_l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___closed__2(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___closed__2); -l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___closed__3 = _init_l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___closed__3(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___closed__3); -l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__11___closed__1 = _init_l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__11___closed__1(); -lean_mark_persistent(l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__11___closed__1); +l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__1 = _init_l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__1(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__1); +l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__2 = _init_l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__2(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__2); +l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__3 = _init_l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__3(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__3); +l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__12___closed__1 = _init_l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__12___closed__1(); +lean_mark_persistent(l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__12___closed__1); l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___closed__1(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___closed__1); l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___closed__2(); @@ -31964,6 +32152,10 @@ l_Lean_PrettyPrinter_Delaborator_delabProj___closed__3 = _init_l_Lean_PrettyPrin lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabProj___closed__3); l_Lean_PrettyPrinter_Delaborator_delabProj___closed__4 = _init_l_Lean_PrettyPrinter_Delaborator_delabProj___closed__4(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabProj___closed__4); +l_Lean_PrettyPrinter_Delaborator_delabProj___closed__5 = _init_l_Lean_PrettyPrinter_Delaborator_delabProj___closed__5(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabProj___closed__5); +l_Lean_PrettyPrinter_Delaborator_delabProj___closed__6 = _init_l_Lean_PrettyPrinter_Delaborator_delabProj___closed__6(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabProj___closed__6); l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabProj___closed__1 = _init_l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabProj___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabProj___closed__1); l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabProj___closed__2 = _init_l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabProj___closed__2(); diff --git a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/TopDownAnalyze.c b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/TopDownAnalyze.c index c01131d2f6b..17dfdc72029 100644 --- a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/TopDownAnalyze.c +++ b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/TopDownAnalyze.c @@ -300,6 +300,7 @@ static lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownA LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeApp___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_analyzeArg___lambda__1(lean_object*, lean_object*, 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_panic___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeMData___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1(lean_object*, lean_object*); static lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeConst___spec__1___closed__1; LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeLet___spec__18(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_withKnowing___rarg(uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -596,7 +597,6 @@ static lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownA static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isHBinOp___lambda__1___closed__15; static lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_290____closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isTrivialBottomUp___closed__7; -uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_9562____closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_canBottomUp___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_getPPProofs(lean_object*); @@ -3546,7 +3546,7 @@ lean_inc(x_4); lean_dec(x_1); x_5 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_5, 0, x_4); -x_6 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1(x_3, x_5); +x_6 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1(x_3, x_5); lean_dec(x_5); lean_dec(x_3); return x_6; diff --git a/stage0/stdlib/Lean/PrettyPrinter/Formatter.c b/stage0/stdlib/Lean/PrettyPrinter/Formatter.c index 3d648daaa70..657fbc53ca6 100644 --- a/stage0/stdlib/Lean/PrettyPrinter/Formatter.c +++ b/stage0/stdlib/Lean/PrettyPrinter/Formatter.c @@ -14,6 +14,7 @@ extern "C" { #endif lean_object* l_List_reverse___rarg(lean_object*); +static lean_object* l_Lean_PrettyPrinter_Formatter_strLitNoAntiquot_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_rawIdentNoAntiquot_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_State_stack___default; extern lean_object* l_Lean_Parser_builtinTokenTable; @@ -23,7 +24,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_ite___rarg___boxed(lean_ uint8_t l_Lean_Syntax_isAntiquotSuffixSplice(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_visitArgs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); -extern lean_object* l_Lean_fieldIdxKind; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_4333_(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkTailWs_formatter___rarg(lean_object*); uint8_t l_Lean_Syntax_isTokenAntiquot(lean_object*); @@ -45,7 +45,6 @@ static lean_object* l_Lean_PrettyPrinter_mkFormatterAttribute___closed__12; static lean_object* l_Lean_PrettyPrinter_Formatter_symbolNoAntiquot_formatter___closed__4; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_PrettyPrinter_Formatter_unicodeSymbolNoAntiquot_formatter___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Formatter_formatterForKindUnsafe___closed__5; -extern lean_object* l_Lean_nullKind; LEAN_EXPORT lean_object* l_Lean_Syntax_MonadTraverser_goUp___at_Lean_PrettyPrinter_Formatter_visitArgs___spec__4(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkLinebreakBefore_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_withOpenDecl_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -93,6 +92,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_formatterForKindUnsafe(l LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_PrettyPrinter_format___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_pushLine(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_instMonadTraverserFormatterM___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_Formatter_strLitNoAntiquot_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_optionalNoAntiquot_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_instInhabitedNat; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_evalInsideQuot_formatter___boxed(lean_object*); @@ -108,6 +108,7 @@ static lean_object* l_Lean_PrettyPrinter_Formatter_State_leadWord___default___cl LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_identEq_formatter___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkWsBefore_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Formatter_trailingNode_formatter___closed__1; +static lean_object* l_Lean_PrettyPrinter_Formatter_charLitNoAntiquot_formatter___closed__1; lean_object* l_List_range(lean_object*); lean_object* l_Lean_Parser_tokenFn(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_instCoeFormatterFormatterAliasValue(lean_object*); @@ -137,7 +138,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkStackTop_formatter_ lean_object* l_Lean_PrettyPrinter_runForNodeKind___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Formatter_instMonadTraverserFormatterM___closed__4; LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Formatter_0__Lean_PrettyPrinter_Formatter_SourceInfo_getExprPos_x3f(lean_object*); -extern lean_object* l_Lean_nameLitKind; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_withAntiquotSuffixSplice_formatter(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_withAntiquotSuffixSplice_formatter___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_notFollowedBy_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -150,9 +150,11 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_pushToken___lambda__3___ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_fill(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Formatter_categoryFormatterCore___lambda__4___closed__3; static lean_object* l_Lean_PrettyPrinter_mkFormatterAttribute___closed__9; +uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Formatter_categoryFormatterCore___lambda__4___closed__7; lean_object* l_Lean_Parser_registerAliasCore___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_getStack___boxed(lean_object*); +static lean_object* l_Lean_PrettyPrinter_Formatter_scientificLitNoAntiquot_formatter___closed__2; static lean_object* l_Lean_PrettyPrinter_Formatter_checkKind___closed__5; static lean_object* l_Lean_addTrace___at_Lean_PrettyPrinter_Formatter_categoryFormatterCore___spec__2___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkPrec_formatter___rarg(lean_object*); @@ -194,14 +196,13 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_strLitNoAntiquot_formatt LEAN_EXPORT lean_object* l_ReaderT_map___at_Lean_PrettyPrinter_Formatter_instMonadTraverserFormatterM___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_eoi_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); -extern lean_object* l_Lean_numLitKind; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_interpolatedStr_formatter___boxed__const__1; static lean_object* l_Lean_PrettyPrinter_Formatter_formatterForKindUnsafe___closed__4; lean_object* lean_nat_sub(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Formatter_formatterForKindUnsafe___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_visitAtom___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_categoryParserOfStack_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_strLitKind; +static lean_object* l_Lean_PrettyPrinter_Formatter_scientificLitNoAntiquot_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_instMonadTraverserFormatterM___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkPrec_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_format___lambda__1___closed__3; @@ -215,6 +216,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_visitAtom(lean_object*, uint8_t l_Substring_contains(lean_object*, uint32_t); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_withMaybeTag___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_MonadTraverser_goLeft___at_Lean_PrettyPrinter_Formatter_visitArgs___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Formatter_interpolatedStr_formatter___spec__2___closed__1; static lean_object* l_Lean_PrettyPrinter_Formatter_checkKind___closed__6; lean_object* l_Lean_Name_toString(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_notFollowedBy_formatter___rarg(lean_object*); @@ -241,7 +243,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_sepBy1NoAntiquot_formatt LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_categoryFormatterCore___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_skip_formatter(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOptionM___at_Lean_PrettyPrinter_format___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_charLitKind; uint8_t l_String_endsWith(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_withoutForbidden_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkColGt_formatter(lean_object*, lean_object*, lean_object*, lean_object*); @@ -315,6 +316,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_State_leadWord___default static lean_object* l_Lean_PrettyPrinter_Formatter_unicodeSymbolNoAntiquot_formatter___closed__1; static lean_object* l_Lean_PrettyPrinter_mkCombinatorFormatterAttribute___closed__1; static lean_object* l_Lean_PrettyPrinter_mkFormatterAttribute___lambda__1___closed__3; +static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Formatter_interpolatedStr_formatter___spec__2___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_getStack___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at_Lean_PrettyPrinter_Formatter_sepByNoAntiquot_formatter___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_PrettyPrinter_Formatter_interpolatedStr_formatter___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -331,7 +333,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_setStack___boxed(lean_ob static lean_object* l_Lean_addTrace___at_Lean_PrettyPrinter_Formatter_categoryFormatterCore___spec__2___closed__3; static lean_object* l_Lean_PrettyPrinter_Formatter_instMonadTraverserFormatterM___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_parseToken___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_identKind; static lean_object* l_Lean_PrettyPrinter_Formatter_parseToken___closed__1; static lean_object* l_Lean_PrettyPrinter_mkFormatterAttribute___closed__11; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_parserOfStack_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -345,7 +346,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_unicodeSymbolNoAntiquot_ lean_object* l_Lean_Syntax_Traverser_down(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_formatterAliasesRef; static lean_object* l_Lean_PrettyPrinter_mkFormatterAttribute___lambda__1___closed__2; -extern lean_object* l_Lean_scientificLitKind; +static lean_object* l_Lean_PrettyPrinter_Formatter_many1Unbox_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkNoImmediateColon_formatter(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_MonadTraverser_goUp___at_Lean_PrettyPrinter_Formatter_visitArgs___spec__4___boxed(lean_object*); static lean_object* l_Lean_PrettyPrinter_Formatter_categoryFormatterCore___closed__3; @@ -357,6 +358,7 @@ uint8_t lean_nat_dec_le(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_PrettyPrinter_format___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkLhsPrec_formatter___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_error_formatter___rarg(lean_object*); +static lean_object* l_Lean_PrettyPrinter_Formatter_many1Unbox_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_pushToken(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArgs(lean_object*); lean_object* l_Lean_Name_append(lean_object*, lean_object*); @@ -367,9 +369,11 @@ extern lean_object* l_Lean_Core_instMonadCoreM; uint8_t l_Lean_Parser_isValidSyntaxNodeKind(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_instMonadTraverserFormatterM; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_incQuotDepth_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_Formatter_fieldIdx_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_instMonadTraverserFormatterM___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_evalInsideQuot_formatter___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_getStackSize___boxed(lean_object*); +static lean_object* l_Lean_PrettyPrinter_Formatter_identNoAntiquot_formatter___closed__4; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_errorAtSavedPos_formatter(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_MonadTraverser_getCur___at_Lean_PrettyPrinter_Formatter_visitArgs___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_withForbidden_formatter___boxed(lean_object*); @@ -379,6 +383,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkNoWsBefore_formatte static lean_object* l_Lean_PrettyPrinter_Formatter_categoryFormatterCore___closed__2; static lean_object* l_Lean_PrettyPrinter_Formatter_categoryFormatterCore___lambda__4___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_setLhsPrec_formatter(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_Formatter_fieldIdx_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_pushToken___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_categoryFormatterCore___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_withOpen_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -411,13 +416,16 @@ lean_object* lean_panic_fn(lean_object*, lean_object*); lean_object* l_String_trim(lean_object*); static lean_object* l_Lean_PrettyPrinter_Formatter_instMonadTraverserFormatterM___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOptionM___at_Lean_PrettyPrinter_format___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_Formatter_numLitNoAntiquot_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_fold(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Formatter_instMonadTraverserFormatterM___closed__2; LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Formatter_0__Lean_PrettyPrinter_Formatter_SourceInfo_getExprPos_x3f___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_withoutInfo_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Std_instInhabitedFormat; +static lean_object* l_Lean_PrettyPrinter_Formatter_numLitNoAntiquot_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_concat___lambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_parserOfStack_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_Formatter_nameLitNoAntiquot_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_4641_(lean_object*); static lean_object* l_Lean_addTrace___at_Lean_PrettyPrinter_Formatter_categoryFormatterCore___spec__2___closed__7; lean_object* lean_int_sub(lean_object*, lean_object*); @@ -425,6 +433,7 @@ static lean_object* l_Lean_PrettyPrinter_Formatter_symbolNoAntiquot_formatter___ static lean_object* l_Lean_PrettyPrinter_format___lambda__1___closed__1; lean_object* l_Lean_indentD(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_charLitNoAntiquot_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_Formatter_identNoAntiquot_formatter___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_concat___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkStackTop_formatter___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_interpretParserDescr_x27___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -442,6 +451,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_pushToken___lambda__3(le static lean_object* l_Lean_PrettyPrinter_mkFormatterAttribute___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkColGe_formatter___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_FormatterM_orElse(lean_object*); +static lean_object* l_Lean_PrettyPrinter_Formatter_nameLitNoAntiquot_formatter___closed__1; lean_object* l_Lean_KeyedDeclsAttribute_init___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkKind(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_unicodeSymbolNoAntiquot_formatter___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -452,7 +462,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_suppressInsideQuot_forma LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkLinebreakBefore_formatter___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_rawCh_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_Traverser_left(lean_object*); -uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_map___at_Lean_PrettyPrinter_Formatter_instMonadTraverserFormatterM___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_format___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_throwBacktrack___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -474,11 +483,11 @@ static lean_object* l_Lean_PrettyPrinter_formatTerm___closed__1; extern lean_object* l_instInhabitedPUnit; lean_object* l_Lean_Syntax_formatStxAux(lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PrettyPrinter_Formatter_State_isUngrouped___default; -extern lean_object* l_Lean_interpolatedStrLitKind; static lean_object* l_Lean_PrettyPrinter_Formatter_pushToken___lambda__4___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_andthen_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Formatter_checkKind___closed__1; uint8_t l_Std_Format_isNil(lean_object*); +static lean_object* l_Lean_PrettyPrinter_Formatter_charLitNoAntiquot_formatter___closed__2; static lean_object* l_Lean_PrettyPrinter_formatTactic___closed__1; static lean_object* l_Lean_PrettyPrinter_mkCombinatorFormatterAttribute___closed__2; static lean_object* l_Lean_PrettyPrinter_Formatter_State_stack___default___closed__1; @@ -6515,7 +6524,7 @@ lean_dec(x_1); x_24 = l_Lean_Parser_whitespace(x_22, x_23); x_25 = lean_ctor_get(x_24, 4); lean_inc(x_25); -x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_19); +x_26 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_25, x_19); lean_dec(x_25); if (x_26 == 0) { @@ -6577,7 +6586,7 @@ lean_dec(x_1); x_44 = l_Lean_Parser_whitespace(x_42, x_43); x_45 = lean_ctor_get(x_44, 4); lean_inc(x_45); -x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_39); +x_46 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Parser_ParserState_hasError___spec__1(x_45, x_39); lean_dec(x_45); if (x_46 == 0) { @@ -8371,15 +8380,33 @@ static lean_object* _init_l_Lean_PrettyPrinter_Formatter_identNoAntiquot_formatt _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("not an ident: ", 14); +x_1 = lean_mk_string_from_bytes("ident", 5); return x_1; } } static lean_object* _init_l_Lean_PrettyPrinter_Formatter_identNoAntiquot_formatter___closed__2() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_PrettyPrinter_Formatter_identNoAntiquot_formatter___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_Formatter_identNoAntiquot_formatter___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("not an ident: ", 14); +return x_1; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_Formatter_identNoAntiquot_formatter___closed__4() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_PrettyPrinter_Formatter_identNoAntiquot_formatter___closed__1; +x_1 = l_Lean_PrettyPrinter_Formatter_identNoAntiquot_formatter___closed__3; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } @@ -8388,7 +8415,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_identNoAntiquot_formatte _start: { lean_object* x_6; lean_object* x_7; -x_6 = l_Lean_identKind; +x_6 = l_Lean_PrettyPrinter_Formatter_identNoAntiquot_formatter___closed__2; lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); @@ -8480,7 +8507,7 @@ lean_inc(x_29); lean_dec(x_12); x_30 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_30, 0, x_10); -x_31 = l_Lean_PrettyPrinter_Formatter_identNoAntiquot_formatter___closed__2; +x_31 = l_Lean_PrettyPrinter_Formatter_identNoAntiquot_formatter___closed__4; x_32 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_32, 0, x_31); lean_ctor_set(x_32, 1, x_30); @@ -8528,7 +8555,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_rawIdentNoAntiquot_forma _start: { lean_object* x_6; lean_object* x_7; -x_6 = l_Lean_identKind; +x_6 = l_Lean_PrettyPrinter_Formatter_identNoAntiquot_formatter___closed__2; lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); @@ -8614,7 +8641,7 @@ lean_inc(x_26); lean_dec(x_12); x_27 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_27, 0, x_10); -x_28 = l_Lean_PrettyPrinter_Formatter_identNoAntiquot_formatter___closed__2; +x_28 = l_Lean_PrettyPrinter_Formatter_identNoAntiquot_formatter___closed__4; x_29 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_29, 0, x_28); lean_ctor_set(x_29, 1, x_27); @@ -8940,56 +8967,164 @@ lean_dec(x_2); return x_8; } } +static lean_object* _init_l_Lean_PrettyPrinter_Formatter_charLitNoAntiquot_formatter___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("char", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_Formatter_charLitNoAntiquot_formatter___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_PrettyPrinter_Formatter_charLitNoAntiquot_formatter___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_charLitNoAntiquot_formatter(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; -x_6 = l_Lean_charLitKind; +x_6 = l_Lean_PrettyPrinter_Formatter_charLitNoAntiquot_formatter___closed__2; x_7 = l_Lean_PrettyPrinter_Formatter_visitAtom(x_6, x_1, x_2, x_3, x_4, x_5); return x_7; } } +static lean_object* _init_l_Lean_PrettyPrinter_Formatter_strLitNoAntiquot_formatter___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("str", 3); +return x_1; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_Formatter_strLitNoAntiquot_formatter___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_PrettyPrinter_Formatter_strLitNoAntiquot_formatter___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_strLitNoAntiquot_formatter(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; -x_6 = l_Lean_strLitKind; +x_6 = l_Lean_PrettyPrinter_Formatter_strLitNoAntiquot_formatter___closed__2; x_7 = l_Lean_PrettyPrinter_Formatter_visitAtom(x_6, x_1, x_2, x_3, x_4, x_5); return x_7; } } +static lean_object* _init_l_Lean_PrettyPrinter_Formatter_nameLitNoAntiquot_formatter___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("name", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_Formatter_nameLitNoAntiquot_formatter___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_PrettyPrinter_Formatter_nameLitNoAntiquot_formatter___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_nameLitNoAntiquot_formatter(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; -x_6 = l_Lean_nameLitKind; +x_6 = l_Lean_PrettyPrinter_Formatter_nameLitNoAntiquot_formatter___closed__2; x_7 = l_Lean_PrettyPrinter_Formatter_visitAtom(x_6, x_1, x_2, x_3, x_4, x_5); return x_7; } } +static lean_object* _init_l_Lean_PrettyPrinter_Formatter_numLitNoAntiquot_formatter___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("num", 3); +return x_1; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_Formatter_numLitNoAntiquot_formatter___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_PrettyPrinter_Formatter_numLitNoAntiquot_formatter___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_numLitNoAntiquot_formatter(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; -x_6 = l_Lean_numLitKind; +x_6 = l_Lean_PrettyPrinter_Formatter_numLitNoAntiquot_formatter___closed__2; x_7 = l_Lean_PrettyPrinter_Formatter_visitAtom(x_6, x_1, x_2, x_3, x_4, x_5); return x_7; } } +static lean_object* _init_l_Lean_PrettyPrinter_Formatter_scientificLitNoAntiquot_formatter___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("scientific", 10); +return x_1; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_Formatter_scientificLitNoAntiquot_formatter___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_PrettyPrinter_Formatter_scientificLitNoAntiquot_formatter___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_scientificLitNoAntiquot_formatter(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; -x_6 = l_Lean_scientificLitKind; +x_6 = l_Lean_PrettyPrinter_Formatter_scientificLitNoAntiquot_formatter___closed__2; x_7 = l_Lean_PrettyPrinter_Formatter_visitAtom(x_6, x_1, x_2, x_3, x_4, x_5); return x_7; } } +static lean_object* _init_l_Lean_PrettyPrinter_Formatter_fieldIdx_formatter___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("fieldIdx", 8); +return x_1; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_Formatter_fieldIdx_formatter___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_PrettyPrinter_Formatter_fieldIdx_formatter___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_fieldIdx_formatter(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; -x_6 = l_Lean_fieldIdxKind; +x_6 = l_Lean_PrettyPrinter_Formatter_fieldIdx_formatter___closed__2; x_7 = l_Lean_PrettyPrinter_Formatter_visitAtom(x_6, x_1, x_2, x_3, x_4, x_5); return x_7; } @@ -9116,6 +9251,24 @@ x_7 = l_Lean_PrettyPrinter_Formatter_visitArgs(x_1, x_2, x_3, x_4, x_5, x_6); return x_7; } } +static lean_object* _init_l_Lean_PrettyPrinter_Formatter_many1Unbox_formatter___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("null", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_Formatter_many1Unbox_formatter___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_PrettyPrinter_Formatter_many1Unbox_formatter___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_many1Unbox_formatter(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: { @@ -9127,7 +9280,7 @@ x_9 = lean_ctor_get(x_7, 1); lean_inc(x_9); lean_dec(x_7); x_10 = l_Lean_Syntax_getKind(x_8); -x_11 = l_Lean_nullKind; +x_11 = l_Lean_PrettyPrinter_Formatter_many1Unbox_formatter___closed__2; x_12 = lean_name_eq(x_10, x_11); lean_dec(x_10); if (x_12 == 0) @@ -10089,6 +10242,24 @@ x_2 = lean_alloc_closure((void*)(l_ReaderT_pure___at_Lean_PrettyPrinter_Formatte return x_2; } } +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Formatter_interpolatedStr_formatter___spec__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("interpolatedStrLitKind", 22); +return x_1; +} +} +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Formatter_interpolatedStr_formatter___spec__2___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Formatter_interpolatedStr_formatter___spec__2___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Formatter_interpolatedStr_formatter___spec__2(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_object* x_9, lean_object* x_10) { _start: { @@ -10099,7 +10270,7 @@ if (x_11 == 0) lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_dec(x_5); x_12 = lean_array_uget(x_2, x_3); -x_13 = l_Lean_interpolatedStrLitKind; +x_13 = l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Formatter_interpolatedStr_formatter___spec__2___closed__2; x_14 = l_Lean_Syntax_isLit_x3f(x_13, x_12); lean_dec(x_12); if (lean_obj_tag(x_14) == 0) @@ -11396,6 +11567,42 @@ l_Lean_PrettyPrinter_Formatter_identNoAntiquot_formatter___closed__1 = _init_l_L lean_mark_persistent(l_Lean_PrettyPrinter_Formatter_identNoAntiquot_formatter___closed__1); l_Lean_PrettyPrinter_Formatter_identNoAntiquot_formatter___closed__2 = _init_l_Lean_PrettyPrinter_Formatter_identNoAntiquot_formatter___closed__2(); lean_mark_persistent(l_Lean_PrettyPrinter_Formatter_identNoAntiquot_formatter___closed__2); +l_Lean_PrettyPrinter_Formatter_identNoAntiquot_formatter___closed__3 = _init_l_Lean_PrettyPrinter_Formatter_identNoAntiquot_formatter___closed__3(); +lean_mark_persistent(l_Lean_PrettyPrinter_Formatter_identNoAntiquot_formatter___closed__3); +l_Lean_PrettyPrinter_Formatter_identNoAntiquot_formatter___closed__4 = _init_l_Lean_PrettyPrinter_Formatter_identNoAntiquot_formatter___closed__4(); +lean_mark_persistent(l_Lean_PrettyPrinter_Formatter_identNoAntiquot_formatter___closed__4); +l_Lean_PrettyPrinter_Formatter_charLitNoAntiquot_formatter___closed__1 = _init_l_Lean_PrettyPrinter_Formatter_charLitNoAntiquot_formatter___closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_Formatter_charLitNoAntiquot_formatter___closed__1); +l_Lean_PrettyPrinter_Formatter_charLitNoAntiquot_formatter___closed__2 = _init_l_Lean_PrettyPrinter_Formatter_charLitNoAntiquot_formatter___closed__2(); +lean_mark_persistent(l_Lean_PrettyPrinter_Formatter_charLitNoAntiquot_formatter___closed__2); +l_Lean_PrettyPrinter_Formatter_strLitNoAntiquot_formatter___closed__1 = _init_l_Lean_PrettyPrinter_Formatter_strLitNoAntiquot_formatter___closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_Formatter_strLitNoAntiquot_formatter___closed__1); +l_Lean_PrettyPrinter_Formatter_strLitNoAntiquot_formatter___closed__2 = _init_l_Lean_PrettyPrinter_Formatter_strLitNoAntiquot_formatter___closed__2(); +lean_mark_persistent(l_Lean_PrettyPrinter_Formatter_strLitNoAntiquot_formatter___closed__2); +l_Lean_PrettyPrinter_Formatter_nameLitNoAntiquot_formatter___closed__1 = _init_l_Lean_PrettyPrinter_Formatter_nameLitNoAntiquot_formatter___closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_Formatter_nameLitNoAntiquot_formatter___closed__1); +l_Lean_PrettyPrinter_Formatter_nameLitNoAntiquot_formatter___closed__2 = _init_l_Lean_PrettyPrinter_Formatter_nameLitNoAntiquot_formatter___closed__2(); +lean_mark_persistent(l_Lean_PrettyPrinter_Formatter_nameLitNoAntiquot_formatter___closed__2); +l_Lean_PrettyPrinter_Formatter_numLitNoAntiquot_formatter___closed__1 = _init_l_Lean_PrettyPrinter_Formatter_numLitNoAntiquot_formatter___closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_Formatter_numLitNoAntiquot_formatter___closed__1); +l_Lean_PrettyPrinter_Formatter_numLitNoAntiquot_formatter___closed__2 = _init_l_Lean_PrettyPrinter_Formatter_numLitNoAntiquot_formatter___closed__2(); +lean_mark_persistent(l_Lean_PrettyPrinter_Formatter_numLitNoAntiquot_formatter___closed__2); +l_Lean_PrettyPrinter_Formatter_scientificLitNoAntiquot_formatter___closed__1 = _init_l_Lean_PrettyPrinter_Formatter_scientificLitNoAntiquot_formatter___closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_Formatter_scientificLitNoAntiquot_formatter___closed__1); +l_Lean_PrettyPrinter_Formatter_scientificLitNoAntiquot_formatter___closed__2 = _init_l_Lean_PrettyPrinter_Formatter_scientificLitNoAntiquot_formatter___closed__2(); +lean_mark_persistent(l_Lean_PrettyPrinter_Formatter_scientificLitNoAntiquot_formatter___closed__2); +l_Lean_PrettyPrinter_Formatter_fieldIdx_formatter___closed__1 = _init_l_Lean_PrettyPrinter_Formatter_fieldIdx_formatter___closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_Formatter_fieldIdx_formatter___closed__1); +l_Lean_PrettyPrinter_Formatter_fieldIdx_formatter___closed__2 = _init_l_Lean_PrettyPrinter_Formatter_fieldIdx_formatter___closed__2(); +lean_mark_persistent(l_Lean_PrettyPrinter_Formatter_fieldIdx_formatter___closed__2); +l_Lean_PrettyPrinter_Formatter_many1Unbox_formatter___closed__1 = _init_l_Lean_PrettyPrinter_Formatter_many1Unbox_formatter___closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_Formatter_many1Unbox_formatter___closed__1); +l_Lean_PrettyPrinter_Formatter_many1Unbox_formatter___closed__2 = _init_l_Lean_PrettyPrinter_Formatter_many1Unbox_formatter___closed__2(); +lean_mark_persistent(l_Lean_PrettyPrinter_Formatter_many1Unbox_formatter___closed__2); +l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Formatter_interpolatedStr_formatter___spec__2___closed__1 = _init_l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Formatter_interpolatedStr_formatter___spec__2___closed__1(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Formatter_interpolatedStr_formatter___spec__2___closed__1); +l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Formatter_interpolatedStr_formatter___spec__2___closed__2 = _init_l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Formatter_interpolatedStr_formatter___spec__2___closed__2(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Formatter_interpolatedStr_formatter___spec__2___closed__2); l_Lean_PrettyPrinter_Formatter_interpolatedStr_formatter___boxed__const__1 = _init_l_Lean_PrettyPrinter_Formatter_interpolatedStr_formatter___boxed__const__1(); lean_mark_persistent(l_Lean_PrettyPrinter_Formatter_interpolatedStr_formatter___boxed__const__1); if (builtin) {res = l_Lean_PrettyPrinter_Formatter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_4333_(lean_io_mk_world()); diff --git a/stage0/stdlib/Lean/PrettyPrinter/Parenthesizer.c b/stage0/stdlib/Lean/PrettyPrinter/Parenthesizer.c index 595a95485db..cbd20af948a 100644 --- a/stage0/stdlib/Lean/PrettyPrinter/Parenthesizer.c +++ b/stage0/stdlib/Lean/PrettyPrinter/Parenthesizer.c @@ -49,7 +49,6 @@ static lean_object* l_Lean_PrettyPrinter_mkCategoryParenthesizerAttribute___lamb LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_withAntiquot_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Parenthesizer_instMonadTraverserParenthesizerM___closed__3; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_PrettyPrinter_parenthesize___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_nullKind; static lean_object* l_Lean_PrettyPrinter_Parenthesizer_term_parenthesizer___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_nonReservedSymbolNoAntiquot_parenthesizer(lean_object*, uint8_t, lean_object*); lean_object* lean_name_mk_string(lean_object*, lean_object*); @@ -89,7 +88,7 @@ lean_object* lean_environment_find(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_MonadTraverser_goUp___at_Lean_PrettyPrinter_Parenthesizer_visitArgs___spec__4___boxed(lean_object*); uint8_t l_Lean_checkTraceOption(lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_4159_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_4183_(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___lambda__9___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_Syntax_MonadTraverser_setCur___at_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Parenthesizer_term_parenthesizer___lambda__1___closed__6; @@ -204,7 +203,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_dbgTraceState_parent static lean_object* l_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___lambda__10___closed__13; static lean_object* l_Lean_PrettyPrinter_Parenthesizer_parenthesizerForKindUnsafe___closed__1; static lean_object* l_Lean_PrettyPrinter_Parenthesizer_node_parenthesizer___closed__2; -static lean_object* l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___lambda__1___closed__6; static lean_object* l_Lean_PrettyPrinter_parenthesizeCommand___closed__2; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Parenthesizer_level_parenthesizer___lambda__1___closed__1; @@ -257,7 +255,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_node_parenthesizer__ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_categoryParserOfStack_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_instMonadTraverserParenthesizerM___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_manyNoAntiquot_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instMonadFunctorReaderT___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_interpolatedStr_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -385,7 +382,7 @@ lean_object* l_Lean_Syntax_Traverser_down(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_level_parenthesizer(lean_object*); static lean_object* l_Lean_PrettyPrinter_Parenthesizer_level_parenthesizer___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_mkCategoryParenthesizerAttribute___lambda__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_3878_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_3902_(lean_object*); static lean_object* l_Lean_addTrace___at_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___spec__2___closed__8; static lean_object* l_Option_format___at_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___spec__6___closed__4; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_setExpected_parenthesizer(lean_object*); @@ -430,11 +427,13 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_identNoAntiquot_pare static lean_object* l_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___closed__1; static lean_object* l_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___lambda__10___closed__12; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_unicodeSymbolNoAntiquot_parenthesizer___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Parenthesizer_interpolatedStr_parenthesizer___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addTrace___at_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___spec__2___closed__2; static lean_object* l_Lean_PrettyPrinter_parenthesizeCommand___closed__1; static lean_object* l_Lean_PrettyPrinter_mkParenthesizerAttribute___closed__9; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_withAntiquotSuffixSplice_parenthesizer___boxed(lean_object*); +static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Parenthesizer_interpolatedStr_parenthesizer___spec__1___closed__2; LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_PrettyPrinter_Parenthesizer_instMonadQuotationParenthesizerM___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Parenthesizer_instMonadTraverserParenthesizerM___closed__1; LEAN_EXPORT lean_object* l_Option_format___at_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___spec__6(lean_object*); @@ -538,7 +537,6 @@ static lean_object* l_Lean_PrettyPrinter_Parenthesizer_term_parenthesizer___lamb LEAN_EXPORT lean_object* l_Lean_Syntax_MonadTraverser_getIdx___at_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___closed__1; LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_PrettyPrinter_Parenthesizer_trailingNode_parenthesizer___spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___lambda__1___closed__4; lean_object* l_Lean_Syntax_Traverser_left(lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Parenthesizer_term_parenthesizer___lambda__2___closed__1; @@ -565,7 +563,6 @@ static lean_object* l_Lean_PrettyPrinter_parenthesize___closed__2; static lean_object* l_Lean_PrettyPrinter_Parenthesizer_level_parenthesizer___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_addPrecCheck___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___lambda__10___closed__9; -extern lean_object* l_Lean_interpolatedStrLitKind; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_identNoAntiquot_parenthesizer___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_notFollowedBy_parenthesizer___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_MonadTraverser_getCur___at_Lean_PrettyPrinter_Parenthesizer_visitArgs___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -7510,67 +7507,33 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___lambda__1___closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("tacticSeq", 9); -return x_1; -} -} -static lean_object* _init_l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___lambda__1___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___lambda__1___closed__2; -x_2 = l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___lambda__1___closed__4; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___lambda__1___closed__6() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(1u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; -} -} LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___lambda__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; 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; lean_object* x_20; +lean_object* x_5; lean_object* x_6; lean_object* x_7; 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; lean_object* x_15; lean_object* x_16; x_5 = l_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___lambda__10___closed__13; lean_inc(x_2); x_6 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_6, 0, x_2); lean_ctor_set(x_6, 1, x_5); -x_7 = l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___lambda__1___closed__6; -x_8 = lean_array_push(x_7, x_1); -x_9 = lean_box(2); -x_10 = l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___lambda__1___closed__5; -x_11 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_11, 0, x_9); -lean_ctor_set(x_11, 1, x_10); -lean_ctor_set(x_11, 2, x_8); -x_12 = l_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___lambda__10___closed__17; -x_13 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_13, 0, x_2); -lean_ctor_set(x_13, 1, x_12); -x_14 = l_Lean_PrettyPrinter_Parenthesizer_term_parenthesizer___lambda__1___closed__12; -x_15 = lean_array_push(x_14, x_6); -x_16 = lean_array_push(x_15, x_11); -x_17 = lean_array_push(x_16, x_13); -x_18 = l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___lambda__1___closed__3; -x_19 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_19, 0, x_9); -lean_ctor_set(x_19, 1, x_18); -lean_ctor_set(x_19, 2, x_17); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_4); -return x_20; +x_7 = l_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___lambda__10___closed__17; +x_8 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_8, 0, x_2); +lean_ctor_set(x_8, 1, x_7); +x_9 = l_Lean_PrettyPrinter_Parenthesizer_term_parenthesizer___lambda__1___closed__12; +x_10 = lean_array_push(x_9, x_6); +x_11 = lean_array_push(x_10, x_1); +x_12 = lean_array_push(x_11, x_8); +x_13 = lean_box(2); +x_14 = l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___lambda__1___closed__3; +x_15 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_14); +lean_ctor_set(x_15, 2, x_12); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_4); +return x_16; } } LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___lambda__2(lean_object* x_1) { @@ -9425,7 +9388,7 @@ x_9 = lean_ctor_get(x_7, 1); lean_inc(x_9); lean_dec(x_7); x_10 = l_Lean_Syntax_getKind(x_8); -x_11 = l_Lean_nullKind; +x_11 = l_Lean_PrettyPrinter_Parenthesizer_term_parenthesizer___lambda__1___closed__8; x_12 = lean_name_eq(x_10, x_11); lean_dec(x_10); if (x_12 == 0) @@ -10204,6 +10167,24 @@ x_7 = lean_apply_5(x_1, x_2, x_3, x_4, x_5, x_6); return x_7; } } +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Parenthesizer_interpolatedStr_parenthesizer___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("interpolatedStrLitKind", 22); +return x_1; +} +} +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Parenthesizer_interpolatedStr_parenthesizer___spec__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Parenthesizer_interpolatedStr_parenthesizer___spec__1___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Parenthesizer_interpolatedStr_parenthesizer___spec__1(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_object* x_9, lean_object* x_10) { _start: { @@ -10214,7 +10195,7 @@ if (x_11 == 0) lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_dec(x_5); x_12 = lean_array_uget(x_2, x_3); -x_13 = l_Lean_interpolatedStrLitKind; +x_13 = l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Parenthesizer_interpolatedStr_parenthesizer___spec__1___closed__2; x_14 = l_Lean_Syntax_isOfKind(x_12, x_13); if (x_14 == 0) { @@ -10454,7 +10435,7 @@ x_10 = l_Lean_PrettyPrinter_Parenthesizer_ite___rarg(x_9, x_2, x_3, x_4, x_5, x_ return x_10; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_3878_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_3902_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; @@ -11209,7 +11190,7 @@ x_6 = l_Lean_PrettyPrinter_parenthesize(x_5, x_1, x_2, x_3, x_4); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_4159_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_4183_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; @@ -11530,12 +11511,6 @@ l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___lambda__1___closed__2 lean_mark_persistent(l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___lambda__1___closed__2); l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___lambda__1___closed__3 = _init_l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___lambda__1___closed__3(); lean_mark_persistent(l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___lambda__1___closed__3); -l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___lambda__1___closed__4 = _init_l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___lambda__1___closed__4(); -lean_mark_persistent(l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___lambda__1___closed__4); -l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___lambda__1___closed__5 = _init_l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___lambda__1___closed__5(); -lean_mark_persistent(l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___lambda__1___closed__5); -l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___lambda__1___closed__6 = _init_l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___lambda__1___closed__6(); -lean_mark_persistent(l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___lambda__1___closed__6); l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___closed__1 = _init_l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___closed__1(); lean_mark_persistent(l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___closed__1); l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___closed__2 = _init_l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___closed__2(); @@ -11599,11 +11574,15 @@ l_Lean_PrettyPrinter_Parenthesizer_node_parenthesizer___closed__6 = _init_l_Lean lean_mark_persistent(l_Lean_PrettyPrinter_Parenthesizer_node_parenthesizer___closed__6); l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer___closed__1 = _init_l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer___closed__1(); lean_mark_persistent(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer___closed__1); +l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Parenthesizer_interpolatedStr_parenthesizer___spec__1___closed__1 = _init_l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Parenthesizer_interpolatedStr_parenthesizer___spec__1___closed__1(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Parenthesizer_interpolatedStr_parenthesizer___spec__1___closed__1); +l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Parenthesizer_interpolatedStr_parenthesizer___spec__1___closed__2 = _init_l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Parenthesizer_interpolatedStr_parenthesizer___spec__1___closed__2(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Parenthesizer_interpolatedStr_parenthesizer___spec__1___closed__2); l_Lean_PrettyPrinter_Parenthesizer_interpolatedStr_parenthesizer___closed__1 = _init_l_Lean_PrettyPrinter_Parenthesizer_interpolatedStr_parenthesizer___closed__1(); lean_mark_persistent(l_Lean_PrettyPrinter_Parenthesizer_interpolatedStr_parenthesizer___closed__1); l_Lean_PrettyPrinter_Parenthesizer_interpolatedStr_parenthesizer___boxed__const__1 = _init_l_Lean_PrettyPrinter_Parenthesizer_interpolatedStr_parenthesizer___boxed__const__1(); lean_mark_persistent(l_Lean_PrettyPrinter_Parenthesizer_interpolatedStr_parenthesizer___boxed__const__1); -if (builtin) {res = l_Lean_PrettyPrinter_Parenthesizer_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_3878_(lean_io_mk_world()); +if (builtin) {res = l_Lean_PrettyPrinter_Parenthesizer_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_3902_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_PrettyPrinter_Parenthesizer_parenthesizerAliasesRef = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_PrettyPrinter_Parenthesizer_parenthesizerAliasesRef); @@ -11628,7 +11607,7 @@ l_Lean_PrettyPrinter_parenthesizeCommand___closed__2 = _init_l_Lean_PrettyPrinte lean_mark_persistent(l_Lean_PrettyPrinter_parenthesizeCommand___closed__2); l_Lean_PrettyPrinter_parenthesizeCommand___closed__3 = _init_l_Lean_PrettyPrinter_parenthesizeCommand___closed__3(); lean_mark_persistent(l_Lean_PrettyPrinter_parenthesizeCommand___closed__3); -res = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_4159_(lean_io_mk_world()); +res = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_4183_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/ResolveName.c b/stage0/stdlib/Lean/ResolveName.c index 0079ae2da74..5b3eb2a68ec 100644 --- a/stage0/stdlib/Lean/ResolveName.c +++ b/stage0/stdlib/Lean/ResolveName.c @@ -27,6 +27,7 @@ lean_object* l___private_Std_Data_HashMap_0__Std_numBucketsForCapacity(lean_obje static lean_object* l_List_toStringAux___at_Lean_resolveGlobalConstNoOverloadCore___spec__3___closed__1; uint8_t l_Lean_Environment_isNamespace(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_getRevAliases___spec__9(lean_object*); +LEAN_EXPORT lean_object* l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_resolveNamespace___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_toStringAux___at_Lean_resolveGlobalConstNoOverloadCore___spec__3___boxed(lean_object*, lean_object*); lean_object* lean_name_mk_string(lean_object*, lean_object*); @@ -57,6 +58,7 @@ LEAN_EXPORT lean_object* l_Lean_resolveGlobalName(lean_object*); static lean_object* l_Lean_ResolveName_resolveNamespaceUsingScope_x3f___closed__2; lean_object* l_Lean_throwUnknownConstant___rarg(lean_object*, lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); +static lean_object* l_List_filterTRAux___at_Lean_getAliases___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_ResolveName_resolveNamespace___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_resolveGlobalConst___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_toString___at_Lean_resolveGlobalConstNoOverloadCore___spec__2___closed__3; @@ -71,7 +73,7 @@ LEAN_EXPORT lean_object* l_Std_HashMapImp_moveEntries___at_Lean_addAliasEntry___ uint8_t l_Lean_Name_isPrefixOf(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_getRevAliases___spec__5(lean_object*); LEAN_EXPORT lean_object* lean_add_alias(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_getAliases___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_getAliases___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_aliasExtension; size_t lean_usize_shift_right(size_t, size_t); LEAN_EXPORT lean_object* l_Std_HashMapImp_find_x3f___at_Lean_addAliasEntry___spec__5(lean_object*, lean_object*); @@ -160,7 +162,8 @@ size_t lean_usize_mul(size_t, size_t); static lean_object* l_Lean_initFn____x40_Lean_ResolveName___hyg_66____closed__5; lean_object* l_Lean_throwErrorAt___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ResolveName_resolveGlobalName(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_getAliases(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_filterTRAux___at_Lean_getAliases___spec__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_getAliases(lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_addAlias___closed__1; static lean_object* l_Lean_ResolveName_resolveNamespaceUsingScope_x3f___closed__1; static lean_object* l_Lean_initFn____x40_Lean_ResolveName___hyg_66____closed__3; @@ -174,7 +177,6 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_getRevAliases___spe LEAN_EXPORT lean_object* l_Std_PersistentHashMap_findAtAux___at_Lean_addAliasEntry___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_resolveGlobalName___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_foldlMAux___at_Lean_getRevAliases___spec__4(lean_object*); -LEAN_EXPORT lean_object* l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object*, lean_object*); uint8_t l_List_elem___at_Lean_NameHashSet_insert___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstNoOverloadCore___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_ResolveName___hyg_66____spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -197,7 +199,6 @@ LEAN_EXPORT lean_object* l_Lean_getRevAliases___lambda__1(lean_object*, lean_obj LEAN_EXPORT lean_object* l_Lean_getAliasState___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_resolveGlobalConst(lean_object*); LEAN_EXPORT lean_object* l_Lean_resolveNamespace___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_ResolveName_0__Lean_ResolveName_resolveQualifiedName___closed__1; static lean_object* l_Lean_ResolveName_resolveNamespaceUsingScope_x3f___closed__4; LEAN_EXPORT lean_object* l___private_Lean_ResolveName_0__Lean_ResolveName_resolveQualifiedName___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_resolveUniqueNamespace___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); @@ -1914,37 +1915,135 @@ lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_getAliases(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_List_filterTRAux___at_Lean_getAliases___spec__1___closed__1() { _start: { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_3 = l_Lean_getAliasState___closed__1; -x_4 = l_Lean_addAlias___closed__1; -x_5 = l_Lean_SimplePersistentEnvExtension_getState___rarg(x_3, x_4, x_1); -x_6 = l_Lean_SMap_find_x3f___at_Lean_addAliasEntry___spec__1(x_5, x_2); -if (lean_obj_tag(x_6) == 0) +lean_object* x_1; +x_1 = l_Lean_protectedExt; +return x_1; +} +} +LEAN_EXPORT lean_object* l_List_filterTRAux___at_Lean_getAliases___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_7; -x_7 = lean_box(0); -return x_7; +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_4; +lean_dec(x_1); +x_4 = l_List_reverse___rarg(x_3); +return x_4; } else { -lean_object* x_8; -x_8 = lean_ctor_get(x_6, 0); -lean_inc(x_8); +uint8_t x_5; +x_5 = !lean_is_exclusive(x_2); +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_2, 0); +x_7 = lean_ctor_get(x_2, 1); +x_8 = l_List_filterTRAux___at_Lean_getAliases___spec__1___closed__1; +lean_inc(x_6); +lean_inc(x_1); +x_9 = l_Lean_TagDeclarationExtension_isTagged(x_8, x_1, x_6); +if (x_9 == 0) +{ +lean_ctor_set(x_2, 1, x_3); +{ +lean_object* _tmp_1 = x_7; +lean_object* _tmp_2 = x_2; +x_2 = _tmp_1; +x_3 = _tmp_2; +} +goto _start; +} +else +{ +lean_free_object(x_2); lean_dec(x_6); -return x_8; +x_2 = x_7; +goto _start; } } +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = lean_ctor_get(x_2, 0); +x_13 = lean_ctor_get(x_2, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_2); +x_14 = l_List_filterTRAux___at_Lean_getAliases___spec__1___closed__1; +lean_inc(x_12); +lean_inc(x_1); +x_15 = l_Lean_TagDeclarationExtension_isTagged(x_14, x_1, x_12); +if (x_15 == 0) +{ +lean_object* x_16; +x_16 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_16, 0, x_12); +lean_ctor_set(x_16, 1, x_3); +x_2 = x_13; +x_3 = x_16; +goto _start; +} +else +{ +lean_dec(x_12); +x_2 = x_13; +goto _start; +} } -LEAN_EXPORT lean_object* l_Lean_getAliases___boxed(lean_object* x_1, lean_object* x_2) { +} +} +} +LEAN_EXPORT lean_object* l_Lean_getAliases(lean_object* x_1, lean_object* x_2, uint8_t x_3) { _start: { -lean_object* x_3; -x_3 = l_Lean_getAliases(x_1, x_2); +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_4 = l_Lean_getAliasState___closed__1; +x_5 = l_Lean_addAlias___closed__1; +x_6 = l_Lean_SimplePersistentEnvExtension_getState___rarg(x_4, x_5, x_1); +x_7 = l_Lean_SMap_find_x3f___at_Lean_addAliasEntry___spec__1(x_6, x_2); +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_8; lean_dec(x_1); -return x_3; +x_8 = lean_box(0); +return x_8; +} +else +{ +if (x_3 == 0) +{ +lean_object* x_9; +lean_dec(x_1); +x_9 = lean_ctor_get(x_7, 0); +lean_inc(x_9); +lean_dec(x_7); +return x_9; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_ctor_get(x_7, 0); +lean_inc(x_10); +lean_dec(x_7); +x_11 = lean_box(0); +x_12 = l_List_filterTRAux___at_Lean_getAliases___spec__1(x_1, x_10, x_11); +return x_12; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_getAliases___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = lean_unbox(x_3); +lean_dec(x_3); +x_5 = l_Lean_getAliases(x_1, x_2, x_4); +return x_5; } } LEAN_EXPORT lean_object* l_Std_AssocList_foldlM___at_Lean_getRevAliases___spec__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { @@ -2603,65 +2702,56 @@ lean_dec(x_1); return x_5; } } -static lean_object* _init_l___private_Lean_ResolveName_0__Lean_ResolveName_resolveQualifiedName___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_protectedExt; -return x_1; -} -} LEAN_EXPORT lean_object* l___private_Lean_ResolveName_0__Lean_ResolveName_resolveQualifiedName(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_4; lean_object* x_5; uint8_t x_6; +lean_object* x_4; uint8_t x_5; lean_object* x_6; uint8_t x_7; lean_inc(x_3); x_4 = l_Lean_Name_append(x_2, x_3); +x_5 = l_Lean_Name_isAtomic(x_3); +lean_dec(x_3); lean_inc(x_4); -x_5 = l_Lean_getAliases(x_1, x_4); +lean_inc(x_1); +x_6 = l_Lean_getAliases(x_1, x_4, x_5); lean_inc(x_4); lean_inc(x_1); -x_6 = l_Lean_Environment_contains(x_1, x_4); -if (x_6 == 0) +x_7 = l_Lean_Environment_contains(x_1, x_4); +if (x_7 == 0) { -lean_object* x_7; uint8_t x_8; -lean_dec(x_3); +lean_object* x_8; uint8_t x_9; lean_inc(x_1); -x_7 = l_Lean_mkPrivateName(x_1, x_4); -lean_inc(x_7); -x_8 = l_Lean_Environment_contains(x_1, x_7); -if (x_8 == 0) +x_8 = l_Lean_mkPrivateName(x_1, x_4); +lean_inc(x_8); +x_9 = l_Lean_Environment_contains(x_1, x_8); +if (x_9 == 0) { -lean_dec(x_7); -return x_5; +lean_dec(x_8); +return x_6; } else { -lean_object* x_9; -x_9 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_9, 0, x_7); -lean_ctor_set(x_9, 1, x_5); -return x_9; +lean_object* x_10; +x_10 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_10, 0, x_8); +lean_ctor_set(x_10, 1, x_6); +return x_10; } } else { -uint8_t x_10; -x_10 = l_Lean_Name_isAtomic(x_3); -lean_dec(x_3); -if (x_10 == 0) +if (x_5 == 0) { lean_object* x_11; lean_dec(x_1); x_11 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_11, 0, x_4); -lean_ctor_set(x_11, 1, x_5); +lean_ctor_set(x_11, 1, x_6); return x_11; } else { lean_object* x_12; uint8_t x_13; -x_12 = l___private_Lean_ResolveName_0__Lean_ResolveName_resolveQualifiedName___closed__1; +x_12 = l_List_filterTRAux___at_Lean_getAliases___spec__1___closed__1; lean_inc(x_4); lean_inc(x_1); x_13 = l_Lean_TagDeclarationExtension_isTagged(x_12, x_1, x_4); @@ -2671,7 +2761,7 @@ lean_object* x_14; lean_dec(x_1); x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_4); -lean_ctor_set(x_14, 1, x_5); +lean_ctor_set(x_14, 1, x_6); return x_14; } else @@ -2684,14 +2774,14 @@ x_16 = l_Lean_Environment_contains(x_1, x_15); if (x_16 == 0) { lean_dec(x_15); -return x_5; +return x_6; } else { lean_object* x_17; x_17 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_17, 0, x_15); -lean_ctor_set(x_17, 1, x_5); +lean_ctor_set(x_17, 1, x_6); return x_17; } } @@ -3374,7 +3464,7 @@ lean_inc(x_1); x_15 = l___private_Lean_ResolveName_0__Lean_ResolveName_resolveExact(x_1, x_13); if (lean_obj_tag(x_15) == 0) { -uint8_t x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; +uint8_t x_16; lean_object* x_17; uint8_t x_18; uint8_t x_19; lean_object* x_20; lean_inc(x_13); lean_inc(x_1); x_16 = l_Lean_Environment_contains(x_1, x_13); @@ -3384,200 +3474,202 @@ x_17 = l_Lean_mkPrivateName(x_1, x_13); lean_inc(x_17); lean_inc(x_1); x_18 = l_Lean_Environment_contains(x_1, x_17); +x_19 = l_Lean_Name_isAtomic(x_13); lean_inc(x_13); -x_19 = l_Lean_getAliases(x_1, x_13); +lean_inc(x_1); +x_20 = l_Lean_getAliases(x_1, x_13, x_19); if (x_16 == 0) { -lean_object* x_20; -x_20 = lean_box(0); +lean_object* x_21; +x_21 = lean_box(0); if (x_18 == 0) { -lean_object* x_21; lean_object* x_22; +lean_object* x_22; lean_object* x_23; lean_dec(x_17); lean_inc(x_3); lean_inc(x_1); -x_21 = l___private_Lean_ResolveName_0__Lean_ResolveName_resolveOpenDecls(x_1, x_13, x_3, x_20); -x_22 = l_List_appendTR___rarg(x_19, x_21); -if (lean_obj_tag(x_22) == 0) -{ -lean_object* x_23; -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_8); -lean_ctor_set(x_23, 1, x_6); +x_22 = l___private_Lean_ResolveName_0__Lean_ResolveName_resolveOpenDecls(x_1, x_13, x_3, x_21); +x_23 = l_List_appendTR___rarg(x_20, x_22); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_8); +lean_ctor_set(x_24, 1, x_6); x_5 = x_7; -x_6 = x_23; +x_6 = x_24; goto _start; } else { -lean_object* x_25; lean_object* x_26; +lean_object* x_26; lean_object* x_27; lean_dec(x_8); lean_dec(x_7); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_25 = l_List_eraseDups___at_Lean_ResolveName_resolveGlobalName_loop___spec__1(x_22); -x_26 = l_List_mapTRAux___at_Lean_ResolveName_resolveGlobalName_loop___spec__3(x_6, x_25, x_20); -return x_26; +x_26 = l_List_eraseDups___at_Lean_ResolveName_resolveGlobalName_loop___spec__1(x_23); +x_27 = l_List_mapTRAux___at_Lean_ResolveName_resolveGlobalName_loop___spec__3(x_6, x_26, x_21); +return x_27; } } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_27 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_27, 0, x_17); -lean_ctor_set(x_27, 1, x_20); -x_28 = l_List_appendTR___rarg(x_27, x_20); +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_17); +lean_ctor_set(x_28, 1, x_21); +x_29 = l_List_appendTR___rarg(x_28, x_21); lean_inc(x_3); lean_inc(x_1); -x_29 = l___private_Lean_ResolveName_0__Lean_ResolveName_resolveOpenDecls(x_1, x_13, x_3, x_28); -x_30 = l_List_appendTR___rarg(x_19, x_29); -if (lean_obj_tag(x_30) == 0) -{ -lean_object* x_31; -x_31 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_31, 0, x_8); -lean_ctor_set(x_31, 1, x_6); +x_30 = l___private_Lean_ResolveName_0__Lean_ResolveName_resolveOpenDecls(x_1, x_13, x_3, x_29); +x_31 = l_List_appendTR___rarg(x_20, x_30); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_32; +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_8); +lean_ctor_set(x_32, 1, x_6); x_5 = x_7; -x_6 = x_31; +x_6 = x_32; goto _start; } else { -lean_object* x_33; lean_object* x_34; +lean_object* x_34; lean_object* x_35; lean_dec(x_8); lean_dec(x_7); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_33 = l_List_eraseDups___at_Lean_ResolveName_resolveGlobalName_loop___spec__1(x_30); -x_34 = l_List_mapTRAux___at_Lean_ResolveName_resolveGlobalName_loop___spec__4(x_6, x_33, x_20); -return x_34; +x_34 = l_List_eraseDups___at_Lean_ResolveName_resolveGlobalName_loop___spec__1(x_31); +x_35 = l_List_mapTRAux___at_Lean_ResolveName_resolveGlobalName_loop___spec__4(x_6, x_34, x_21); +return x_35; } } } else { -lean_object* x_35; lean_object* x_36; -x_35 = lean_box(0); +lean_object* x_36; lean_object* x_37; +x_36 = lean_box(0); lean_inc(x_13); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_13); -lean_ctor_set(x_36, 1, x_35); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_13); +lean_ctor_set(x_37, 1, x_36); if (x_18 == 0) { -lean_object* x_37; lean_object* x_38; +lean_object* x_38; lean_object* x_39; lean_dec(x_17); lean_inc(x_3); lean_inc(x_1); -x_37 = l___private_Lean_ResolveName_0__Lean_ResolveName_resolveOpenDecls(x_1, x_13, x_3, x_36); -x_38 = l_List_appendTR___rarg(x_19, x_37); -if (lean_obj_tag(x_38) == 0) -{ -lean_object* x_39; -x_39 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_39, 0, x_8); -lean_ctor_set(x_39, 1, x_6); +x_38 = l___private_Lean_ResolveName_0__Lean_ResolveName_resolveOpenDecls(x_1, x_13, x_3, x_37); +x_39 = l_List_appendTR___rarg(x_20, x_38); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; +x_40 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_40, 0, x_8); +lean_ctor_set(x_40, 1, x_6); x_5 = x_7; -x_6 = x_39; +x_6 = x_40; goto _start; } else { -lean_object* x_41; lean_object* x_42; +lean_object* x_42; lean_object* x_43; lean_dec(x_8); lean_dec(x_7); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_41 = l_List_eraseDups___at_Lean_ResolveName_resolveGlobalName_loop___spec__1(x_38); -x_42 = l_List_mapTRAux___at_Lean_ResolveName_resolveGlobalName_loop___spec__5(x_6, x_41, x_35); -return x_42; +x_42 = l_List_eraseDups___at_Lean_ResolveName_resolveGlobalName_loop___spec__1(x_39); +x_43 = l_List_mapTRAux___at_Lean_ResolveName_resolveGlobalName_loop___spec__5(x_6, x_42, x_36); +return x_43; } } else { -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_43 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_43, 0, x_17); -lean_ctor_set(x_43, 1, x_35); -x_44 = l_List_appendTR___rarg(x_43, x_36); +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_44 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_44, 0, x_17); +lean_ctor_set(x_44, 1, x_36); +x_45 = l_List_appendTR___rarg(x_44, x_37); lean_inc(x_3); lean_inc(x_1); -x_45 = l___private_Lean_ResolveName_0__Lean_ResolveName_resolveOpenDecls(x_1, x_13, x_3, x_44); -x_46 = l_List_appendTR___rarg(x_19, x_45); -if (lean_obj_tag(x_46) == 0) -{ -lean_object* x_47; -x_47 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_47, 0, x_8); -lean_ctor_set(x_47, 1, x_6); +x_46 = l___private_Lean_ResolveName_0__Lean_ResolveName_resolveOpenDecls(x_1, x_13, x_3, x_45); +x_47 = l_List_appendTR___rarg(x_20, x_46); +if (lean_obj_tag(x_47) == 0) +{ +lean_object* x_48; +x_48 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_48, 0, x_8); +lean_ctor_set(x_48, 1, x_6); x_5 = x_7; -x_6 = x_47; +x_6 = x_48; goto _start; } else { -lean_object* x_49; lean_object* x_50; +lean_object* x_50; lean_object* x_51; lean_dec(x_8); lean_dec(x_7); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_49 = l_List_eraseDups___at_Lean_ResolveName_resolveGlobalName_loop___spec__1(x_46); -x_50 = l_List_mapTRAux___at_Lean_ResolveName_resolveGlobalName_loop___spec__6(x_6, x_49, x_35); -return x_50; +x_50 = l_List_eraseDups___at_Lean_ResolveName_resolveGlobalName_loop___spec__1(x_47); +x_51 = l_List_mapTRAux___at_Lean_ResolveName_resolveGlobalName_loop___spec__6(x_6, x_50, x_36); +return x_51; } } } } else { -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_dec(x_13); lean_dec(x_8); lean_dec(x_7); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_51 = lean_ctor_get(x_15, 0); -lean_inc(x_51); +x_52 = lean_ctor_get(x_15, 0); +lean_inc(x_52); lean_dec(x_15); -x_52 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_52, 0, x_51); -lean_ctor_set(x_52, 1, x_6); -x_53 = lean_box(0); -x_54 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_54, 0, x_52); -lean_ctor_set(x_54, 1, x_53); -return x_54; +x_53 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_6); +x_54 = lean_box(0); +x_55 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +return x_55; } } else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_dec(x_13); lean_dec(x_8); lean_dec(x_7); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_55 = l_List_eraseDups___at_Lean_ResolveName_resolveGlobalName_loop___spec__1(x_14); -x_56 = lean_box(0); -x_57 = l_List_mapTRAux___at_Lean_ResolveName_resolveGlobalName_loop___spec__7(x_6, x_55, x_56); -return x_57; +x_56 = l_List_eraseDups___at_Lean_ResolveName_resolveGlobalName_loop___spec__1(x_14); +x_57 = lean_box(0); +x_58 = l_List_mapTRAux___at_Lean_ResolveName_resolveGlobalName_loop___spec__7(x_6, x_56, x_57); +return x_58; } } else { -lean_object* x_58; +lean_object* x_59; lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_58 = lean_box(0); -return x_58; +x_59 = lean_box(0); +return x_59; } } } @@ -3642,7 +3734,7 @@ static lean_object* _init_l_Lean_ResolveName_resolveNamespaceUsingScope_x3f___cl 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_ResolveName_resolveNamespaceUsingScope_x3f___closed__1; x_2 = l_Lean_ResolveName_resolveNamespaceUsingScope_x3f___closed__2; -x_3 = lean_unsigned_to_nat(137u); +x_3 = lean_unsigned_to_nat(146u); x_4 = lean_unsigned_to_nat(27u); x_5 = l_Lean_ResolveName_resolveNamespaceUsingScope_x3f___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -4219,7 +4311,7 @@ x_2 = lean_alloc_closure((void*)(l_Lean_resolveUniqueNamespace___rarg), 5, 0); return x_2; } } -LEAN_EXPORT lean_object* l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -4364,7 +4456,7 @@ LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstCore___rarg___lambda__2(lean_o { lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; x_6 = lean_box(0); -x_7 = l_List_filterAux___at_Lean_resolveGlobalConstCore___spec__1(x_5, x_6); +x_7 = l_List_filterTRAux___at_Lean_resolveGlobalConstCore___spec__1(x_5, x_6); lean_inc(x_7); lean_inc(x_1); x_8 = lean_alloc_closure((void*)(l_Lean_resolveGlobalConstCore___rarg___lambda__1___boxed), 4, 3); @@ -5170,8 +5262,8 @@ lean_dec_ref(res); lean_mark_persistent(l_Lean_addAlias___closed__1); l_Lean_getAliasState___closed__1 = _init_l_Lean_getAliasState___closed__1(); lean_mark_persistent(l_Lean_getAliasState___closed__1); -l___private_Lean_ResolveName_0__Lean_ResolveName_resolveQualifiedName___closed__1 = _init_l___private_Lean_ResolveName_0__Lean_ResolveName_resolveQualifiedName___closed__1(); -lean_mark_persistent(l___private_Lean_ResolveName_0__Lean_ResolveName_resolveQualifiedName___closed__1); +l_List_filterTRAux___at_Lean_getAliases___spec__1___closed__1 = _init_l_List_filterTRAux___at_Lean_getAliases___spec__1___closed__1(); +lean_mark_persistent(l_List_filterTRAux___at_Lean_getAliases___spec__1___closed__1); l_Lean_ResolveName_resolveNamespaceUsingScope_x3f___closed__1 = _init_l_Lean_ResolveName_resolveNamespaceUsingScope_x3f___closed__1(); lean_mark_persistent(l_Lean_ResolveName_resolveNamespaceUsingScope_x3f___closed__1); l_Lean_ResolveName_resolveNamespaceUsingScope_x3f___closed__2 = _init_l_Lean_ResolveName_resolveNamespaceUsingScope_x3f___closed__2(); diff --git a/stage0/stdlib/Lean/Server/Completion.c b/stage0/stdlib/Lean/Server/Completion.c index 9b99c767bc0..115973148ad 100644 --- a/stage0/stdlib/Lean/Server/Completion.c +++ b/stage0/stdlib/Lean/Server/Completion.c @@ -26,7 +26,6 @@ lean_object* lean_erase_macro_scopes(lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces___lambda__1(uint8_t, lean_object*, lean_object*, lean_object*, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static double l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__12; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_foldlMAux_traverse___at_Lean_Server_Completion_completeNamespaces___spec__8___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); lean_object* lean_nat_div(lean_object*, lean_object*); @@ -52,12 +51,11 @@ lean_object* l_Array_append___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_tacticCompletion___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_noConfusionExt; static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__5; -lean_object* lean_nat_log2(lean_object*); +double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); extern lean_object* l_Std_Format_defWidth; lean_object* l_Lean_Elab_Info_occursBefore_x3f(lean_object*, lean_object*); uint8_t l_Lean_Elab_Info_isSmaller(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_tacticCompletion___boxed(lean_object*, lean_object*); -double lean_float_of_nat(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -104,7 +102,6 @@ static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentArray_forIn___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__36(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__1; -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__9; extern lean_object* l_Lean_Lsp_instInhabitedCompletionItem; uint8_t l_Lean_Name_isPrefixOf(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__2___boxed(lean_object*, lean_object*, lean_object*); @@ -113,7 +110,6 @@ lean_object* lean_string_utf8_extract(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__39(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Completion_completionBlackListExt; lean_object* l_Lean_Meta_unfoldDefinition_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__4; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_foldlMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_qpartition_loop___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -133,7 +129,6 @@ uint8_t lean_usize_dec_lt(size_t, size_t); uint8_t l_Lean_isPrivateName(lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_foldlMAux___at_Lean_Server_Completion_completeNamespaces___spec__6___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__2___closed__8; -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__11; LEAN_EXPORT lean_object* l_Std_AssocList_forM___at_Lean_Server_Completion_completeNamespaces___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addCompletionItem___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); @@ -142,7 +137,6 @@ extern lean_object* l_Lean_Parser_instInhabitedParserCategory; static lean_object* l_Std_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__3; lean_object* l_Lean_FileMap_utf8PosToLspPos(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__40(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -double l_Float_ofBinaryScientific(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_runM___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_Completion_matchNamespace___closed__1; static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_runM___closed__1; @@ -168,7 +162,6 @@ LEAN_EXPORT lean_object* l_Std_PersistentArray_forInAux___at___private_Lean_Serv lean_object* lean_st_ref_take(lean_object*, lean_object*); static lean_object* l_Std_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_runM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_int_mul(lean_object*, lean_object*); extern lean_object* l_instInhabitedFloat; lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__33___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -185,8 +178,7 @@ extern lean_object* l_Lean_NameSSet_instInhabitedNameSSet; static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__31(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__8; -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__1; +static double l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__1; LEAN_EXPORT lean_object* l_ReaderT_bind___at___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addCompletionItem___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -229,7 +221,6 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Com LEAN_EXPORT lean_object* l_Std_PersistentHashMap_foldlMAux_traverse___at_Lean_Server_Completion_completeNamespaces___spec__8___rarg(lean_object*, lean_object*, 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_Std_PersistentHashMap_foldlMAux___at_Lean_Server_Completion_completeNamespaces___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__10(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfo___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Name_isAtomic(lean_object*); @@ -275,7 +266,6 @@ static lean_object* l_Lean_Server_Completion_addToBlackList___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_isProjectionFn___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__1___closed__1; static lean_object* l_Lean_getConstInfo___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__2; -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__5; uint8_t l_Lean_Expr_Data_binderInfo(uint64_t); LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_SavedState_restore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -300,12 +290,10 @@ lean_object* l_Lean_SimplePersistentEnvExtension_getState___rarg(lean_object*, l LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addKeywordCompletionItem___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__28___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_InfoTree_foldInfo___rarg(lean_object*, lean_object*, lean_object*); -lean_object* lean_int_neg(lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__22(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getAllParentStructures(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_foldlM___at_Lean_Server_Completion_completeNamespaces___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Completion_State_itemsMain___default; -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__6; LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_tacticCompletion(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -368,7 +356,6 @@ static lean_object* l_Lean_Server_Completion_initFn____x40_Lean_Server_Completio static lean_object* l_Lean_Server_Completion_initFn____x40_Lean_Server_Completion___hyg_8____closed__1; LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__21(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_nat_shiftl(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_getPrefix(lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); @@ -406,7 +393,6 @@ static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion static lean_object* l_List_forIn_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__2; static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_tacticCompletion___lambda__1___closed__6; LEAN_EXPORT lean_object* l_Std_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_int_sub(lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___closed__3; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_string_length(lean_object*); @@ -420,7 +406,6 @@ uint8_t l_Lean_NameSet_contains(lean_object*, lean_object*); extern lean_object* l_Lean_namespacesExt; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__39___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__7; LEAN_EXPORT lean_object* l_Lean_SMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__2(lean_object*, 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_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -440,7 +425,6 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_c LEAN_EXPORT lean_object* l_Std_PersistentHashMap_foldlMAux___at_Lean_Server_Completion_completeNamespaces___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addNamespaceCompletionItem___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__3; lean_object* l_Lean_TagDeclarationExtension_tag(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore_visitNamespaces(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_mkCompletionItem___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -459,13 +443,11 @@ lean_object* l_Lean_mkConst(lean_object*, lean_object*); static lean_object* l_panic___at___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___spec__1___closed__1; lean_object* l_Lean_Name_replacePrefix(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_runM(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_nat_to_int(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Completion_initFn____x40_Lean_Server_Completion___hyg_8_(lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_is_class(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_mkCompletionItem(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__34___closed__2; -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__10; LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isTypeApplicable___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_isStructure(lean_object*, lean_object*); uint8_t lean_string_dec_lt(lean_object*, lean_object*); @@ -2308,7 +2290,6 @@ return x_28; else { uint8_t x_29; -lean_dec(x_1); x_29 = !lean_is_exclusive(x_11); if (x_29 == 0) { @@ -2449,7 +2430,6 @@ else { uint8_t x_71; lean_dec(x_3); -lean_dec(x_1); x_71 = !lean_is_exclusive(x_11); if (x_71 == 0) { @@ -2491,6 +2471,7 @@ uint8_t x_10; lean_object* x_11; x_10 = lean_unbox(x_4); lean_dec(x_4); x_11 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_mkCompletionItem(x_1, x_2, x_3, x_10, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_1); return x_11; } } @@ -2814,6 +2795,7 @@ x_14 = lean_unbox_float(x_5); lean_dec(x_5); x_15 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addCompletionItem___lambda__1(x_1, x_2, x_13, x_4, x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_7); +lean_dec(x_1); return x_15; } } @@ -2827,6 +2809,7 @@ x_14 = lean_unbox_float(x_6); lean_dec(x_6); x_15 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addCompletionItem(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_7); +lean_dec(x_1); return x_15; } } @@ -4340,7 +4323,6 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); x_17 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addCompletionItem___lambda__1___closed__1; lean_ctor_set(x_11, 0, x_17); return x_11; @@ -4372,7 +4354,6 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); x_21 = !lean_is_exclusive(x_19); if (x_21 == 0) { @@ -4442,7 +4423,6 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); x_38 = !lean_is_exclusive(x_19); if (x_38 == 0) { @@ -4486,7 +4466,6 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); x_46 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addCompletionItem___lambda__1___closed__1; x_47 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_47, 0, x_46); @@ -4519,7 +4498,6 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); x_51 = lean_ctor_get(x_49, 1); lean_inc(x_51); if (lean_is_exclusive(x_49)) { @@ -4579,7 +4557,6 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); x_62 = lean_ctor_get(x_49, 0); lean_inc(x_62); x_63 = lean_ctor_get(x_49, 1); @@ -4613,6 +4590,7 @@ x_11 = lean_unbox_float(x_4); lean_dec(x_4); x_12 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addCompletionItemForDecl(x_1, x_2, x_3, x_11, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_5); +lean_dec(x_1); return x_12; } } @@ -4648,10 +4626,12 @@ return x_3; static double _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addKeywordCompletionItem___closed__4() { _start: { -lean_object* x_1; double x_2; +lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; x_1 = lean_unsigned_to_nat(1u); -x_2 = lean_float_of_nat(x_1); -return x_2; +x_2 = 0; +x_3 = lean_unsigned_to_nat(0u); +x_4 = l_Float_ofScientific(x_1, x_2, x_3); +return x_4; } } static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addKeywordCompletionItem___boxed__const__1() { @@ -5130,119 +5110,15 @@ lean_dec(x_1); return x_5; } } -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(2u); -x_2 = lean_nat_log2(x_1); -return x_2; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(64u); -x_2 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__1; -x_3 = lean_nat_sub(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__3() { +static double _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(3u); -x_2 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__2; -x_3 = lean_nat_add(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; +lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; x_1 = lean_unsigned_to_nat(2u); -x_2 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__3; -x_3 = lean_nat_shiftl(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__4; -x_2 = lean_unsigned_to_nat(5u); -x_3 = lean_nat_div(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__6() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(4u); -x_2 = lean_nat_to_int(x_1); -return x_2; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__7() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__6; -x_2 = lean_int_neg(x_1); -return x_2; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(1u); -x_2 = lean_nat_to_int(x_1); -return x_2; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__9() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__7; -x_2 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__8; -x_3 = lean_int_mul(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__10() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__2; -x_2 = lean_nat_to_int(x_1); -return x_2; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__11() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__9; -x_2 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__10; -x_3 = lean_int_sub(x_1, x_2); -return x_3; -} -} -static double _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__12() { -_start: -{ -lean_object* x_1; lean_object* x_2; double x_3; -x_1 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__5; -x_2 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__11; -x_3 = l_Float_ofBinaryScientific(x_1, x_2); -return x_3; +x_2 = 1; +x_3 = lean_unsigned_to_nat(1u); +x_4 = l_Float_ofScientific(x_1, x_2, x_3); +return x_4; } } LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic(lean_object* x_1, lean_object* x_2) { @@ -5269,7 +5145,7 @@ lean_dec(x_1); x_6 = lean_ctor_get(x_2, 1); lean_inc(x_6); lean_dec(x_2); -x_7 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__12; +x_7 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__1; x_8 = l_Lean_FuzzyMatching_fuzzyMatchScoreWithThreshold_x3f(x_5, x_6, x_7); return x_8; } @@ -5573,7 +5449,7 @@ return x_26; else { double x_27; lean_object* x_28; -x_27 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__12; +x_27 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__1; lean_inc(x_23); x_28 = l_Lean_FuzzyMatching_fuzzyMatchScoreWithThreshold_x3f(x_21, x_23, x_27); if (lean_obj_tag(x_28) == 0) @@ -6057,7 +5933,7 @@ return x_9; else { double x_10; lean_object* x_11; -x_10 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__12; +x_10 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__1; x_11 = l_Lean_FuzzyMatching_fuzzyMatchScoreWithThreshold_x3f(x_7, x_5, x_10); return x_11; } @@ -8649,6 +8525,7 @@ lean_inc(x_11); lean_inc(x_10); lean_inc(x_1); x_71 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addCompletionItem(x_59, x_68, x_1, x_62, x_69, x_70, x_9, x_10, x_11, x_12, x_13, x_65); +lean_dec(x_59); if (lean_obj_tag(x_71) == 0) { lean_object* x_72; @@ -8729,6 +8606,7 @@ lean_inc(x_11); lean_inc(x_10); lean_inc(x_1); x_86 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addCompletionItem(x_59, x_82, x_1, x_83, x_84, x_85, x_9, x_10, x_11, x_12, x_13, x_65); +lean_dec(x_59); if (lean_obj_tag(x_86) == 0) { lean_object* x_87; @@ -10418,7 +10296,6 @@ lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_2); lean_dec(x_1); x_11 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addCompletionItem___lambda__1___closed__1; x_12 = lean_alloc_ctor(0, 2, 0); @@ -10451,7 +10328,6 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_1); -lean_inc(x_2); x_19 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addCompletionItemForDecl(x_2, x_13, x_1, x_3, x_5, x_6, x_7, x_8, x_9, x_18); if (lean_obj_tag(x_19) == 0) { @@ -10466,7 +10342,6 @@ lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_2); lean_dec(x_1); x_21 = !lean_is_exclusive(x_19); if (x_21 == 0) @@ -10511,7 +10386,6 @@ lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_2); lean_dec(x_1); x_29 = !lean_is_exclusive(x_19); if (x_29 == 0) @@ -10557,7 +10431,6 @@ lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_2); lean_dec(x_1); x_11 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addCompletionItem___lambda__1___closed__1; x_12 = lean_alloc_ctor(0, 2, 0); @@ -10590,7 +10463,6 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_1); -lean_inc(x_2); x_19 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addCompletionItemForDecl(x_2, x_13, x_1, x_3, x_5, x_6, x_7, x_8, x_9, x_18); if (lean_obj_tag(x_19) == 0) { @@ -10605,7 +10477,6 @@ lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_2); lean_dec(x_1); x_21 = !lean_is_exclusive(x_19); if (x_21 == 0) @@ -10650,7 +10521,6 @@ lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_2); lean_dec(x_1); x_29 = !lean_is_exclusive(x_19); if (x_29 == 0) @@ -12125,6 +11995,7 @@ x_21 = lean_unbox_float(x_20); lean_dec(x_20); x_22 = l_List_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__18(x_3, x_4, x_21, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_6); +lean_dec(x_4); return x_22; } } @@ -12305,6 +12176,7 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_1); x_47 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addCompletionItemForDecl(x_37, x_38, x_1, x_46, x_6, x_7, x_8, x_9, x_10, x_42); +lean_dec(x_37); if (lean_obj_tag(x_47) == 0) { lean_object* x_48; @@ -13004,6 +12876,7 @@ lean_inc(x_10); lean_inc(x_9); lean_inc(x_1); x_77 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addCompletionItem(x_69, x_73, x_1, x_74, x_75, x_76, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_69); if (lean_obj_tag(x_77) == 0) { lean_object* x_78; @@ -13763,6 +13636,7 @@ lean_inc(x_10); lean_inc(x_9); lean_inc(x_1); x_59 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addCompletionItem(x_51, x_55, x_1, x_56, x_57, x_58, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_51); if (lean_obj_tag(x_59) == 0) { lean_object* x_60; @@ -14599,6 +14473,7 @@ lean_dec(x_35); x_38 = lean_unbox_float(x_24); lean_dec(x_24); x_39 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addCompletionItem(x_23, x_36, x_5, x_26, x_37, x_38, x_7, x_8, x_9, x_10, x_11, x_33); +lean_dec(x_23); if (lean_obj_tag(x_39) == 0) { lean_object* x_40; @@ -14695,6 +14570,7 @@ lean_dec(x_57); x_61 = lean_unbox_float(x_24); lean_dec(x_24); x_62 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addCompletionItem(x_23, x_58, x_5, x_59, x_60, x_61, x_7, x_8, x_9, x_10, x_11, x_33); +lean_dec(x_23); if (lean_obj_tag(x_62) == 0) { lean_object* x_63; @@ -15219,6 +15095,7 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_3); x_50 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addCompletionItem(x_38, x_47, x_3, x_41, x_48, x_49, x_7, x_8, x_9, x_10, x_11, x_44); +lean_dec(x_38); if (lean_obj_tag(x_50) == 0) { lean_object* x_51; @@ -15303,6 +15180,7 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_3); x_65 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addCompletionItem(x_38, x_61, x_3, x_62, x_63, x_64, x_7, x_8, x_9, x_10, x_11, x_44); +lean_dec(x_38); if (lean_obj_tag(x_65) == 0) { lean_object* x_66; @@ -16124,6 +16002,7 @@ x_11 = lean_unbox_float(x_3); lean_dec(x_3); x_12 = l_List_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__17(x_1, x_2, x_11, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_5); +lean_dec(x_2); return x_12; } } @@ -16135,6 +16014,7 @@ x_11 = lean_unbox_float(x_3); lean_dec(x_3); x_12 = l_List_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__18(x_1, x_2, x_11, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_5); +lean_dec(x_2); return x_12; } } @@ -16350,6 +16230,7 @@ x_11 = lean_unbox_float(x_4); lean_dec(x_4); x_12 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__6(x_1, x_2, x_3, x_11, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_5); +lean_dec(x_2); return x_12; } } @@ -17258,7 +17139,6 @@ if (x_42 == 0) { lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; double x_48; uint8_t x_49; lean_object* x_50; x_43 = lean_ctor_get(x_34, 0); -lean_inc(x_18); x_44 = l_Lean_Name_getString_x21(x_18); x_45 = lean_box(0); x_46 = lean_name_mk_string(x_45, x_44); @@ -17268,6 +17148,7 @@ x_48 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addKeywordCo x_49 = lean_unbox(x_43); lean_dec(x_43); x_50 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addCompletionItem(x_46, x_47, x_2, x_34, x_49, x_48, x_5, x_6, x_7, x_8, x_9, x_41); +lean_dec(x_46); return x_50; } else @@ -17276,7 +17157,6 @@ lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean x_51 = lean_ctor_get(x_34, 0); lean_inc(x_51); lean_dec(x_34); -lean_inc(x_18); x_52 = l_Lean_Name_getString_x21(x_18); x_53 = lean_box(0); x_54 = lean_name_mk_string(x_53, x_52); @@ -17287,6 +17167,7 @@ x_57 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addKeywordCo x_58 = lean_unbox(x_51); lean_dec(x_51); x_59 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addCompletionItem(x_54, x_55, x_2, x_56, x_58, x_57, x_5, x_6, x_7, x_8, x_9, x_41); +lean_dec(x_54); return x_59; } } @@ -17528,7 +17409,6 @@ if (lean_is_exclusive(x_94)) { lean_dec_ref(x_94); x_101 = lean_box(0); } -lean_inc(x_80); x_102 = l_Lean_Name_getString_x21(x_80); x_103 = lean_box(0); x_104 = lean_name_mk_string(x_103, x_102); @@ -17543,6 +17423,7 @@ x_107 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addKeywordC x_108 = lean_unbox(x_100); lean_dec(x_100); x_109 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addCompletionItem(x_104, x_105, x_2, x_106, x_108, x_107, x_5, x_6, x_7, x_8, x_9, x_99); +lean_dec(x_104); return x_109; } } @@ -18126,7 +18007,7 @@ if (lean_is_exclusive(x_21)) { x_32 = 1; lean_inc(x_17); x_33 = l_Lean_Name_toString(x_17, x_32); -x_34 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__12; +x_34 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__1; lean_inc(x_33); lean_inc(x_4); x_35 = l_Lean_FuzzyMatching_fuzzyMatchScoreWithThreshold_x3f(x_4, x_33, x_34); @@ -20848,28 +20729,6 @@ lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completio l___private_Lean_Server_Completion_0__Lean_Server_Completion_runM___closed__1 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_runM___closed__1(); lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_runM___closed__1); l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__1 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__1(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__1); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__2 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__2(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__2); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__3 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__3(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__3); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__4 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__4(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__4); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__5 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__5(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__5); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__6 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__6(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__6); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__7 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__7(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__7); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__8 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__8(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__8); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__9 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__9(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__9); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__10 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__10(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__10); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__11 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__11(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__11); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__12 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__12(); l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___boxed__const__1 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___boxed__const__1(); lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___boxed__const__1); l_panic___at___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___spec__1___closed__1 = _init_l_panic___at___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___spec__1___closed__1(); diff --git a/stage0/stdlib/Lean/Server/FileWorker.c b/stage0/stdlib/Lean/Server/FileWorker.c index a682475452f..17dbb8a4e6e 100644 --- a/stage0/stdlib/Lean/Server/FileWorker.c +++ b/stage0/stdlib/Lean/Server/FileWorker.c @@ -16,7 +16,6 @@ extern "C" { lean_object* l_List_reverse___rarg(lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__40; static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__67; -lean_object* l_IO_setRandSeed(lean_object*, lean_object*); lean_object* lean_string_push(lean_object*, uint32_t); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__8; lean_object* l_Lean_Server_FileWorker_RpcSession_keptAlive(lean_object*, lean_object*); @@ -32,11 +31,11 @@ static lean_object* l_Lean_Server_FileWorker_compileHeader___lambda__1___closed_ static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__45; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_parseParams___at_Lean_Server_FileWorker_handleNotification___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__2; +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleRequest___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_lakeSetupSearchPath___closed__4; LEAN_EXPORT lean_object* l_Std_RBNode_foldM___at_Lean_Server_FileWorker_mainLoop___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapM___at_Lean_Server_FileWorker_lakeSetupSearchPath___spec__2(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_updateDocument___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_AsyncElabState_lastSnap___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_updateDocument___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Stream_readLspRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__65; @@ -57,12 +56,13 @@ lean_object* lean_name_mk_string(lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); lean_object* l_Lean_Parser_parseHeader(lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_parseParams___at_Lean_Server_FileWorker_handleRequest___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_io_error_to_string(lean_object*); +LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_Server_FileWorker_handleRequest___spec__2(lean_object*, lean_object*); lean_object* l_Array_append___rarg(lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__36; static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__70; static lean_object* l_Lean_Server_FileWorker_compileHeader___lambda__1___closed__8; -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_updateDocument___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcConnectParams____x40_Lean_Data_Lsp_Extra___hyg_1088_(lean_object*); static lean_object* l_Lean_Server_FileWorker_lakeSetupSearchPath___closed__5; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_parseParams___at_Lean_Server_FileWorker_handleNotification___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -80,6 +80,7 @@ lean_object* l_Lean_JsonNumber_toString(lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__35; uint8_t lean_uint32_to_uint8(uint32_t); static lean_object* l_Std_RBNode_foldM___at_Lean_Server_FileWorker_mainLoop___spec__1___closed__3; +LEAN_EXPORT lean_object* l_IO_FS_Stream_writeLspResponse___at_Lean_Server_FileWorker_handleRequest___spec__4(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_compileHeader___lambda__1___closed__11; static lean_object* l_Lean_Server_FileWorker_handleNotification___closed__1; lean_object* l_Lean_Server_publishProgressDone(lean_object*, lean_object*, lean_object*); @@ -116,9 +117,11 @@ LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleRpcConnect___rarg___boxe lean_object* l_IO_throwServerError___rarg(lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__19; lean_object* l_Array_toPersistentArray___rarg(lean_object*); +lean_object* l_List_head_x21___rarg(lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__27; lean_object* lean_io_process_spawn(lean_object*, lean_object*); lean_object* l_Lean_Server_DocumentMeta_mkInputContext(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initializeWorker___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l_Lean_Server_FileWorker_handleNotification___closed__4; LEAN_EXPORT lean_object* l_Std_RBNode_insert___at_Lean_Server_FileWorker_handleRpcConnect___spec__1(lean_object*, uint64_t, lean_object*); @@ -136,6 +139,7 @@ static lean_object* l_Lean_Server_FileWorker_parseParams___rarg___closed__1; static lean_object* l_Lean_Server_FileWorker_initAndRunWorker___closed__2; LEAN_EXPORT lean_object* l_IO_FS_Stream_writeLspNotification___at___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishIleanInfo___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_compileHeader___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleRequest___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleRpcConnect___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_parseParams___at_Lean_Server_FileWorker_handleNotification___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_get_set_stderr(lean_object*, lean_object*); @@ -158,14 +162,12 @@ static lean_object* l_Lean_Server_FileWorker_compileHeader___closed__2; lean_object* l_Std_RBNode_setBlack___rarg(lean_object*); lean_object* l_IO_appDir(lean_object*); lean_object* l_Lean_Widget_msgToInteractiveDiagnostic(lean_object*, lean_object*, uint8_t, lean_object*); -lean_object* lean_array_fget(lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_lakeSetupSearchPath_processStderr___closed__2; -uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_mainLoop___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCancelRequest___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_instDecidableNot___rarg(uint8_t); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__17; -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleRequest___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleRequest___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Json_compress(lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); lean_object* l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(lean_object*, lean_object*); @@ -180,7 +182,6 @@ static lean_object* l_Lean_Server_FileWorker_handleNotification___closed__3; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_parseParams___at_Lean_Server_FileWorker_handleNotification___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__43; static lean_object* l_Lean_Server_FileWorker_lakeSetupSearchPath___closed__2; -lean_object* l_Lean_Server_Snapshots_reparseHeader(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_RBNode_forIn_visit___at_Lean_Server_FileWorker_mainLoop___spec__4___at_Lean_Server_FileWorker_mainLoop___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_lakeSetupSearchPath___spec__1(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Std_RBNode_ins___at_Lean_Server_FileWorker_handleRpcConnect___spec__2(lean_object*, uint64_t, lean_object*); @@ -194,14 +195,13 @@ LEAN_EXPORT lean_object* l_Std_RBNode_insert___at_Lean_Server_FileWorker_queueRe static lean_object* l_Lean_Server_FileWorker_handleDidChange___closed__5; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_compileHeader___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_toString(lean_object*, uint8_t); -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_parseParams___at_Lean_Server_FileWorker_handleRequest___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__28; +lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__14; static lean_object* l_List_mapTRAux___at_Lean_Server_FileWorker_compileHeader___spec__1___closed__1; LEAN_EXPORT lean_object* l_Std_RBNode_ins___at_Lean_Server_FileWorker_handleRpcConnect___spec__2___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_io_as_task(lean_object*, lean_object*, lean_object*); static lean_object* l_List_mapTRAux___at_Lean_Server_FileWorker_compileHeader___spec__1___closed__2; -uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_SearchPath_findAllWithExt___spec__1(lean_object*, lean_object*); lean_object* l_Std_RBNode_appendTrees___rarg(lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_mkState(lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcKeepAliveParams____x40_Lean_Data_Lsp_Extra___hyg_1605_(lean_object*); @@ -227,13 +227,13 @@ lean_object* lean_io_prim_handle_get_line(lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__61; static lean_object* l_Lean_Server_FileWorker_lakeSetupSearchPath___closed__1; static lean_object* l_Std_RBNode_foldM___at_Lean_Server_FileWorker_mainLoop___spec__1___closed__6; -static lean_object* l_Lean_Server_FileWorker_updateDocument___lambda__3___closed__1; lean_object* l_Std_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_queueRequest___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_compileHeader___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleRpcKeepAlive(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__12; -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleRequest___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Array_back___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleRequest___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_Lsp_instFromJsonInitializeParams___spec__2(lean_object*, lean_object*); lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_parseParams(lean_object*); @@ -262,7 +262,7 @@ static lean_object* l_Lean_Server_FileWorker_initAndRunWorker___closed__6; lean_object* l_Lean_realPathNormalized(lean_object*, lean_object*); uint8_t l_Array_isEmpty___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_parseParams___at_Lean_Server_FileWorker_handleNotification___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_updateDocument___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_updateDocument___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_lakeSetupSearchPath_processStderr___closed__4; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_updatePendingRequests___boxed(lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_firstFrontendMacroScope; @@ -285,7 +285,6 @@ lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_Lsp_instFromJsonInitializePar extern lean_object* l_System_FilePath_exeExtension; lean_object* l_IO_AsyncList_ofList___rarg(lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentArray_mapM___at_Lean_Server_FileWorker_compileHeader___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_parseParams___at_Lean_Server_FileWorker_handleRequest___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__42; lean_object* l_Lean_Server_Snapshots_Snapshot_diagnostics(lean_object*); lean_object* l_Lean_initSearchPath(lean_object*, lean_object*, lean_object*); @@ -298,11 +297,11 @@ static lean_object* l_Lean_Server_FileWorker_compileHeader___lambda__1___closed_ static lean_object* l_Std_RBNode_foldM___at_Lean_Server_FileWorker_mainLoop___spec__1___closed__2; LEAN_EXPORT lean_object* l_Std_PersistentArray_mapMAux___at_Lean_Server_FileWorker_compileHeader___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initializeWorker(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_IO_FS_Stream_writeLspResponse___at_Lean_Server_FileWorker_handleRequest___spec__3(lean_object*, lean_object*, lean_object*); static lean_object* l_List_mapTRAux___at_Lean_Server_FileWorker_compileHeader___spec__1___closed__3; LEAN_EXPORT lean_object* l_Std_RBNode_foldM___at_Lean_Server_FileWorker_mainLoop___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTRAux___at_Lean_Server_FileWorker_compileHeader___spec__1(lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__37; +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_parseParams___at_Lean_Server_FileWorker_handleRequest___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_environment(uint32_t, lean_object*); static lean_object* l_Lean_Server_FileWorker_compileHeader___lambda__1___closed__5; lean_object* l_IO_AsyncList_updateFinishedPrefix___rarg(lean_object*, lean_object*); @@ -319,6 +318,7 @@ lean_object* l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(lean_ LEAN_EXPORT lean_object* l_Std_PersistentArray_mapM___at_Lean_Server_FileWorker_compileHeader___spec__7(uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleRpcConnect___boxed(lean_object*, lean_object*); extern lean_object* l_Lean_MessageLog_empty; +lean_object* lean_io_bind_task(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_HashMap_ofList___at_Lean_Server_ModuleRefs_instCoeModuleRefsModuleRefs___spec__5(lean_object*); lean_object* lean_io_getenv(lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_server_worker_main(lean_object*, lean_object*); @@ -382,7 +382,6 @@ LEAN_EXPORT lean_object* l_Std_RBNode_erase___at_Lean_Server_FileWorker_mainLoop LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_nextCmdSnap___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_nextCmdSnap___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_IO_eprint___at_IO_eprintln___spec__1(lean_object*, lean_object*); -lean_object* lean_io_get_random_bytes(size_t, lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__34; lean_object* l_String_trim(lean_object*); lean_object* l_List_dropLast___rarg(lean_object*); @@ -393,7 +392,6 @@ lean_object* l_Lean_Server_foldDocumentChanges(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_RBNode_forIn_visit___at_Lean_Server_FileWorker_mainLoop___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_nextCmdSnap___lambda__1___closed__1; static lean_object* l_Lean_Server_FileWorker_handleDidChange___closed__4; -lean_object* lean_uint64_to_nat(uint64_t); lean_object* l_IO_FS_Stream_withPrefix(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_parseParams___at_Lean_Server_FileWorker_handleNotification___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_Lsp_instFromJsonInitializeParams___spec__4(lean_object*, lean_object*); @@ -407,13 +405,14 @@ static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_ini LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_updatePendingRequests(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_getLast_x21___rarg(lean_object*, lean_object*); lean_object* lean_get_stderr(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_unfoldCmdSnaps(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_unfoldCmdSnaps(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__38; lean_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_compileHeader___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__21; lean_object* l_Lean_Server_FileWorker_CancelToken_set(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initAndRunWorker(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_SearchPath_findAllWithExt___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_compileHeader___spec__4(uint8_t, lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_lakeSetupSearchPath___spec__1___boxed(lean_object*, lean_object*, lean_object*); @@ -431,10 +430,8 @@ static lean_object* l_Lean_Json_toStructured_x3f___at___private_Lean_Server_File static lean_object* l_Lean_Server_FileWorker_compileHeader___lambda__3___closed__1; lean_object* l_IO_FS_Handle_readToEnd_loop(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_workerMain___closed__1; -lean_object* l_ByteArray_toUInt64LE_x21(lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__11; lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidChangeTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_439_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_AsyncElabState_lastSnap(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishIleanInfoUpdate(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Stream_writeLspResponse___at_Lean_Server_FileWorker_handleRequest___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_workerMain___boxed__const__1; @@ -448,10 +445,10 @@ lean_object* l_IO_AsyncList_unfoldAsync___rarg(lean_object*, lean_object*, lean_ static lean_object* l_Lean_Server_FileWorker_lakeSetupSearchPath___closed__10; lean_object* l_List_mapTRAux___at_Lean_Server_FileWorker_handleDocumentHighlight___spec__2(size_t, lean_object*, lean_object*); lean_object* lean_io_map_task(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_task_pure(lean_object*); lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_Lsp_instFromJsonInitializeParams___spec__3(lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__62; LEAN_EXPORT lean_object* l_Std_RBNode_forIn_visit___at_Lean_Server_FileWorker_mainLoop___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Array_insertAt___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__33; lean_object* lean_nat_to_int(lean_object*); lean_object* l_List_appendTR___rarg(lean_object*, lean_object*); @@ -474,41 +471,6 @@ static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_ini uint8_t lean_string_dec_eq(lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__54; -LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_AsyncElabState_lastSnap(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; uint8_t x_7; -x_2 = lean_ctor_get(x_1, 1); -x_3 = lean_array_get_size(x_2); -x_4 = lean_unsigned_to_nat(1u); -x_5 = lean_nat_sub(x_3, x_4); -x_6 = lean_ctor_get(x_1, 0); -x_7 = lean_nat_dec_lt(x_5, x_3); -lean_dec(x_3); -if (x_7 == 0) -{ -lean_dec(x_5); -lean_inc(x_6); -return x_6; -} -else -{ -lean_object* x_8; -x_8 = lean_array_fget(x_2, x_5); -lean_dec(x_5); -return x_8; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_AsyncElabState_lastSnap___boxed(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_AsyncElabState_lastSnap(x_1); -lean_dec(x_1); -return x_2; -} -} static lean_object* _init_l_Lean_Json_toStructured_x3f___at___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishIleanInfo___spec__2___closed__1() { _start: { @@ -733,524 +695,347 @@ lean_dec(x_2); x_16 = l_Lean_Server_Snapshots_compileNextCmd(x_14, x_1, x_15, x_13); if (lean_obj_tag(x_16) == 0) { -lean_object* x_17; lean_object* x_18; uint8_t x_19; +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); x_18 = lean_ctor_get(x_16, 1); lean_inc(x_18); lean_dec(x_16); -x_19 = !lean_is_exclusive(x_4); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_4, 1); lean_inc(x_17); -x_21 = lean_array_push(x_20, x_17); -lean_ctor_set(x_4, 1, x_21); -x_22 = l_Lean_Server_FileWorker_CancelToken_check___at___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_nextCmdSnap___spec__1(x_5, x_4, x_18); -if (lean_obj_tag(x_22) == 0) +x_19 = lean_array_push(x_4, x_17); +x_20 = l_Lean_Server_FileWorker_CancelToken_check___at___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_nextCmdSnap___spec__1(x_5, x_19, x_18); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_22, 1); -lean_inc(x_24); -lean_dec(x_22); -x_25 = !lean_is_exclusive(x_23); -if (x_25 == 0) +lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = !lean_is_exclusive(x_21); +if (x_23 == 0) { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_26 = lean_ctor_get(x_23, 1); -x_27 = lean_ctor_get(x_23, 0); -lean_dec(x_27); +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_24 = lean_ctor_get(x_21, 1); +x_25 = lean_ctor_get(x_21, 0); +lean_dec(x_25); lean_inc(x_17); -x_28 = l_Lean_Server_Snapshots_Snapshot_diagnostics(x_17); -x_29 = l_Std_PersistentArray_toArray___rarg(x_28); +x_26 = l_Lean_Server_Snapshots_Snapshot_diagnostics(x_17); +x_27 = l_Std_PersistentArray_toArray___rarg(x_26); lean_inc(x_10); lean_inc(x_3); -x_30 = l_Lean_Server_publishDiagnostics(x_3, x_29, x_10, x_24); -if (lean_obj_tag(x_30) == 0) +x_28 = l_Lean_Server_publishDiagnostics(x_3, x_27, x_10, x_22); +if (lean_obj_tag(x_28) == 0) { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_31 = lean_ctor_get(x_30, 1); -lean_inc(x_31); -lean_dec(x_30); -x_32 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_nextCmdSnap___lambda__1___closed__1; +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_29 = lean_ctor_get(x_28, 1); +lean_inc(x_29); +lean_dec(x_28); +x_30 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_nextCmdSnap___lambda__1___closed__1; lean_inc(x_17); -x_33 = lean_array_push(x_32, x_17); -x_34 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishIleanInfoUpdate___closed__1; -x_35 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishIleanInfo(x_34, x_3, x_10, x_33, x_31); -if (lean_obj_tag(x_35) == 0) +x_31 = lean_array_push(x_30, x_17); +x_32 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishIleanInfoUpdate___closed__1; +x_33 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishIleanInfo(x_32, x_3, x_10, x_31, x_29); +if (lean_obj_tag(x_33) == 0) { -uint8_t x_36; -x_36 = !lean_is_exclusive(x_35); -if (x_36 == 0) +uint8_t x_34; +x_34 = !lean_is_exclusive(x_33); +if (x_34 == 0) { -lean_object* x_37; lean_object* x_38; -x_37 = lean_ctor_get(x_35, 0); -lean_dec(x_37); -x_38 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_38, 0, x_17); -lean_ctor_set(x_23, 0, x_38); -lean_ctor_set(x_35, 0, x_23); -return x_35; +lean_object* x_35; lean_object* x_36; +x_35 = lean_ctor_get(x_33, 0); +lean_dec(x_35); +x_36 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_36, 0, x_17); +lean_ctor_set(x_21, 0, x_36); +lean_ctor_set(x_33, 0, x_21); +return x_33; } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_39 = lean_ctor_get(x_35, 1); -lean_inc(x_39); -lean_dec(x_35); -x_40 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_40, 0, x_17); -lean_ctor_set(x_23, 0, x_40); -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_23); -lean_ctor_set(x_41, 1, x_39); -return x_41; +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_33, 1); +lean_inc(x_37); +lean_dec(x_33); +x_38 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_38, 0, x_17); +lean_ctor_set(x_21, 0, x_38); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_21); +lean_ctor_set(x_39, 1, x_37); +return x_39; } } else { -uint8_t x_42; -lean_free_object(x_23); -lean_dec(x_26); +uint8_t x_40; +lean_free_object(x_21); +lean_dec(x_24); lean_dec(x_17); -x_42 = !lean_is_exclusive(x_35); -if (x_42 == 0) +x_40 = !lean_is_exclusive(x_33); +if (x_40 == 0) { -lean_object* x_43; lean_object* x_44; -x_43 = lean_ctor_get(x_35, 0); -x_44 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_35, 0, x_44); -return x_35; +lean_object* x_41; lean_object* x_42; +x_41 = lean_ctor_get(x_33, 0); +x_42 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_33, 0, x_42); +return x_33; } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_45 = lean_ctor_get(x_35, 0); -x_46 = lean_ctor_get(x_35, 1); -lean_inc(x_46); -lean_inc(x_45); -lean_dec(x_35); -x_47 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_47, 0, x_45); -x_48 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_46); -return x_48; +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_43 = lean_ctor_get(x_33, 0); +x_44 = lean_ctor_get(x_33, 1); +lean_inc(x_44); +lean_inc(x_43); +lean_dec(x_33); +x_45 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_45, 0, x_43); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_44); +return x_46; } } } else { -uint8_t x_49; -lean_free_object(x_23); -lean_dec(x_26); +uint8_t x_47; +lean_free_object(x_21); +lean_dec(x_24); lean_dec(x_17); lean_dec(x_10); lean_dec(x_3); -x_49 = !lean_is_exclusive(x_30); -if (x_49 == 0) +x_47 = !lean_is_exclusive(x_28); +if (x_47 == 0) { -lean_object* x_50; lean_object* x_51; -x_50 = lean_ctor_get(x_30, 0); -x_51 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_30, 0, x_51); -return x_30; +lean_object* x_48; lean_object* x_49; +x_48 = lean_ctor_get(x_28, 0); +x_49 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_28, 0, x_49); +return x_28; } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_52 = lean_ctor_get(x_30, 0); -x_53 = lean_ctor_get(x_30, 1); -lean_inc(x_53); -lean_inc(x_52); -lean_dec(x_30); -x_54 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_54, 0, x_52); -x_55 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_55, 0, x_54); -lean_ctor_set(x_55, 1, x_53); -return x_55; +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_50 = lean_ctor_get(x_28, 0); +x_51 = lean_ctor_get(x_28, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_28); +x_52 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_52, 0, x_50); +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_51); +return x_53; } } } else { -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_56 = lean_ctor_get(x_23, 1); -lean_inc(x_56); -lean_dec(x_23); +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_54 = lean_ctor_get(x_21, 1); +lean_inc(x_54); +lean_dec(x_21); lean_inc(x_17); -x_57 = l_Lean_Server_Snapshots_Snapshot_diagnostics(x_17); -x_58 = l_Std_PersistentArray_toArray___rarg(x_57); +x_55 = l_Lean_Server_Snapshots_Snapshot_diagnostics(x_17); +x_56 = l_Std_PersistentArray_toArray___rarg(x_55); lean_inc(x_10); lean_inc(x_3); -x_59 = l_Lean_Server_publishDiagnostics(x_3, x_58, x_10, x_24); -if (lean_obj_tag(x_59) == 0) +x_57 = l_Lean_Server_publishDiagnostics(x_3, x_56, x_10, x_22); +if (lean_obj_tag(x_57) == 0) { -lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_60 = lean_ctor_get(x_59, 1); -lean_inc(x_60); -lean_dec(x_59); -x_61 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_nextCmdSnap___lambda__1___closed__1; +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_58 = lean_ctor_get(x_57, 1); +lean_inc(x_58); +lean_dec(x_57); +x_59 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_nextCmdSnap___lambda__1___closed__1; lean_inc(x_17); -x_62 = lean_array_push(x_61, x_17); -x_63 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishIleanInfoUpdate___closed__1; -x_64 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishIleanInfo(x_63, x_3, x_10, x_62, x_60); -if (lean_obj_tag(x_64) == 0) +x_60 = lean_array_push(x_59, x_17); +x_61 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishIleanInfoUpdate___closed__1; +x_62 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishIleanInfo(x_61, x_3, x_10, x_60, x_58); +if (lean_obj_tag(x_62) == 0) { -lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_65 = lean_ctor_get(x_64, 1); -lean_inc(x_65); -if (lean_is_exclusive(x_64)) { - lean_ctor_release(x_64, 0); - lean_ctor_release(x_64, 1); - x_66 = x_64; +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_63 = lean_ctor_get(x_62, 1); +lean_inc(x_63); +if (lean_is_exclusive(x_62)) { + lean_ctor_release(x_62, 0); + lean_ctor_release(x_62, 1); + x_64 = x_62; } else { - lean_dec_ref(x_64); - x_66 = lean_box(0); + lean_dec_ref(x_62); + x_64 = lean_box(0); } -x_67 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_67, 0, x_17); -x_68 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_68, 0, x_67); -lean_ctor_set(x_68, 1, x_56); -if (lean_is_scalar(x_66)) { - x_69 = lean_alloc_ctor(0, 2, 0); +x_65 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_65, 0, x_17); +x_66 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_66, 0, x_65); +lean_ctor_set(x_66, 1, x_54); +if (lean_is_scalar(x_64)) { + x_67 = lean_alloc_ctor(0, 2, 0); } else { - x_69 = x_66; + x_67 = x_64; } -lean_ctor_set(x_69, 0, x_68); -lean_ctor_set(x_69, 1, x_65); -return x_69; +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_67, 1, x_63); +return x_67; } else { -lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; -lean_dec(x_56); +lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +lean_dec(x_54); lean_dec(x_17); -x_70 = lean_ctor_get(x_64, 0); -lean_inc(x_70); -x_71 = lean_ctor_get(x_64, 1); -lean_inc(x_71); -if (lean_is_exclusive(x_64)) { - lean_ctor_release(x_64, 0); - lean_ctor_release(x_64, 1); - x_72 = x_64; +x_68 = lean_ctor_get(x_62, 0); +lean_inc(x_68); +x_69 = lean_ctor_get(x_62, 1); +lean_inc(x_69); +if (lean_is_exclusive(x_62)) { + lean_ctor_release(x_62, 0); + lean_ctor_release(x_62, 1); + x_70 = x_62; } else { - lean_dec_ref(x_64); - x_72 = lean_box(0); + lean_dec_ref(x_62); + x_70 = lean_box(0); } -x_73 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_73, 0, x_70); -if (lean_is_scalar(x_72)) { - x_74 = lean_alloc_ctor(1, 2, 0); +x_71 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_71, 0, x_68); +if (lean_is_scalar(x_70)) { + x_72 = lean_alloc_ctor(1, 2, 0); } else { - x_74 = x_72; + x_72 = x_70; } -lean_ctor_set(x_74, 0, x_73); -lean_ctor_set(x_74, 1, x_71); -return x_74; +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set(x_72, 1, x_69); +return x_72; } } else { -lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; -lean_dec(x_56); +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +lean_dec(x_54); lean_dec(x_17); lean_dec(x_10); lean_dec(x_3); -x_75 = lean_ctor_get(x_59, 0); -lean_inc(x_75); -x_76 = lean_ctor_get(x_59, 1); -lean_inc(x_76); -if (lean_is_exclusive(x_59)) { - lean_ctor_release(x_59, 0); - lean_ctor_release(x_59, 1); - x_77 = x_59; +x_73 = lean_ctor_get(x_57, 0); +lean_inc(x_73); +x_74 = lean_ctor_get(x_57, 1); +lean_inc(x_74); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + lean_ctor_release(x_57, 1); + x_75 = x_57; } else { - lean_dec_ref(x_59); - x_77 = lean_box(0); + lean_dec_ref(x_57); + x_75 = lean_box(0); } -x_78 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_78, 0, x_75); -if (lean_is_scalar(x_77)) { - x_79 = lean_alloc_ctor(1, 2, 0); +x_76 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_76, 0, x_73); +if (lean_is_scalar(x_75)) { + x_77 = lean_alloc_ctor(1, 2, 0); } else { - x_79 = x_77; + x_77 = x_75; } -lean_ctor_set(x_79, 0, x_78); -lean_ctor_set(x_79, 1, x_76); -return x_79; +lean_ctor_set(x_77, 0, x_76); +lean_ctor_set(x_77, 1, x_74); +return x_77; } } } else { -uint8_t x_80; +uint8_t x_78; lean_dec(x_17); lean_dec(x_10); lean_dec(x_3); -x_80 = !lean_is_exclusive(x_22); -if (x_80 == 0) +x_78 = !lean_is_exclusive(x_20); +if (x_78 == 0) { -return x_22; +return x_20; } else { -lean_object* x_81; lean_object* x_82; lean_object* x_83; -x_81 = lean_ctor_get(x_22, 0); -x_82 = lean_ctor_get(x_22, 1); -lean_inc(x_82); -lean_inc(x_81); -lean_dec(x_22); -x_83 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_83, 0, x_81); -lean_ctor_set(x_83, 1, x_82); -return x_83; +lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_79 = lean_ctor_get(x_20, 0); +x_80 = lean_ctor_get(x_20, 1); +lean_inc(x_80); +lean_inc(x_79); +lean_dec(x_20); +x_81 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_81, 0, x_79); +lean_ctor_set(x_81, 1, x_80); +return x_81; } } } else { -lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_84 = lean_ctor_get(x_4, 0); -x_85 = lean_ctor_get(x_4, 1); -lean_inc(x_85); -lean_inc(x_84); +uint8_t x_82; +lean_dec(x_10); lean_dec(x_4); -lean_inc(x_17); -x_86 = lean_array_push(x_85, x_17); -x_87 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_87, 0, x_84); -lean_ctor_set(x_87, 1, x_86); -x_88 = l_Lean_Server_FileWorker_CancelToken_check___at___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_nextCmdSnap___spec__1(x_5, x_87, x_18); -if (lean_obj_tag(x_88) == 0) -{ -lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; -x_89 = lean_ctor_get(x_88, 0); -lean_inc(x_89); -x_90 = lean_ctor_get(x_88, 1); -lean_inc(x_90); -lean_dec(x_88); -x_91 = lean_ctor_get(x_89, 1); -lean_inc(x_91); -if (lean_is_exclusive(x_89)) { - lean_ctor_release(x_89, 0); - lean_ctor_release(x_89, 1); - x_92 = x_89; -} else { - lean_dec_ref(x_89); - x_92 = lean_box(0); -} -lean_inc(x_17); -x_93 = l_Lean_Server_Snapshots_Snapshot_diagnostics(x_17); -x_94 = l_Std_PersistentArray_toArray___rarg(x_93); -lean_inc(x_10); -lean_inc(x_3); -x_95 = l_Lean_Server_publishDiagnostics(x_3, x_94, x_10, x_90); -if (lean_obj_tag(x_95) == 0) +lean_dec(x_3); +x_82 = !lean_is_exclusive(x_16); +if (x_82 == 0) { -lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_96 = lean_ctor_get(x_95, 1); -lean_inc(x_96); -lean_dec(x_95); -x_97 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_nextCmdSnap___lambda__1___closed__1; -lean_inc(x_17); -x_98 = lean_array_push(x_97, x_17); -x_99 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishIleanInfoUpdate___closed__1; -x_100 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishIleanInfo(x_99, x_3, x_10, x_98, x_96); -if (lean_obj_tag(x_100) == 0) -{ -lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; -x_101 = lean_ctor_get(x_100, 1); -lean_inc(x_101); -if (lean_is_exclusive(x_100)) { - lean_ctor_release(x_100, 0); - lean_ctor_release(x_100, 1); - x_102 = x_100; -} else { - lean_dec_ref(x_100); - x_102 = lean_box(0); -} -x_103 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_103, 0, x_17); -if (lean_is_scalar(x_92)) { - x_104 = lean_alloc_ctor(0, 2, 0); -} else { - x_104 = x_92; -} -lean_ctor_set(x_104, 0, x_103); -lean_ctor_set(x_104, 1, x_91); -if (lean_is_scalar(x_102)) { - x_105 = lean_alloc_ctor(0, 2, 0); -} else { - x_105 = x_102; -} -lean_ctor_set(x_105, 0, x_104); -lean_ctor_set(x_105, 1, x_101); -return x_105; +lean_object* x_83; lean_object* x_84; +x_83 = lean_ctor_get(x_16, 0); +x_84 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_84, 0, x_83); +lean_ctor_set(x_16, 0, x_84); +return x_16; } else { -lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; -lean_dec(x_92); -lean_dec(x_91); -lean_dec(x_17); -x_106 = lean_ctor_get(x_100, 0); -lean_inc(x_106); -x_107 = lean_ctor_get(x_100, 1); -lean_inc(x_107); -if (lean_is_exclusive(x_100)) { - lean_ctor_release(x_100, 0); - lean_ctor_release(x_100, 1); - x_108 = x_100; -} else { - lean_dec_ref(x_100); - x_108 = lean_box(0); -} -x_109 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_109, 0, x_106); -if (lean_is_scalar(x_108)) { - x_110 = lean_alloc_ctor(1, 2, 0); -} else { - x_110 = x_108; +lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_85 = lean_ctor_get(x_16, 0); +x_86 = lean_ctor_get(x_16, 1); +lean_inc(x_86); +lean_inc(x_85); +lean_dec(x_16); +x_87 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_87, 0, x_85); +x_88 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_88, 0, x_87); +lean_ctor_set(x_88, 1, x_86); +return x_88; } -lean_ctor_set(x_110, 0, x_109); -lean_ctor_set(x_110, 1, x_107); -return x_110; } } else { -lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; -lean_dec(x_92); -lean_dec(x_91); -lean_dec(x_17); +uint8_t x_89; lean_dec(x_10); +lean_dec(x_4); lean_dec(x_3); -x_111 = lean_ctor_get(x_95, 0); -lean_inc(x_111); -x_112 = lean_ctor_get(x_95, 1); -lean_inc(x_112); -if (lean_is_exclusive(x_95)) { - lean_ctor_release(x_95, 0); - lean_ctor_release(x_95, 1); - x_113 = x_95; -} else { - lean_dec_ref(x_95); - x_113 = lean_box(0); -} -x_114 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_114, 0, x_111); -if (lean_is_scalar(x_113)) { - x_115 = lean_alloc_ctor(1, 2, 0); -} else { - x_115 = x_113; -} -lean_ctor_set(x_115, 0, x_114); -lean_ctor_set(x_115, 1, x_112); -return x_115; -} +lean_dec(x_2); +lean_dec(x_1); +x_89 = !lean_is_exclusive(x_12); +if (x_89 == 0) +{ +lean_object* x_90; lean_object* x_91; +x_90 = lean_ctor_get(x_12, 0); +x_91 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_91, 0, x_90); +lean_ctor_set(x_12, 0, x_91); +return x_12; } else { -lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; -lean_dec(x_17); -lean_dec(x_10); -lean_dec(x_3); -x_116 = lean_ctor_get(x_88, 0); -lean_inc(x_116); -x_117 = lean_ctor_get(x_88, 1); -lean_inc(x_117); -if (lean_is_exclusive(x_88)) { - lean_ctor_release(x_88, 0); - lean_ctor_release(x_88, 1); - x_118 = x_88; -} else { - lean_dec_ref(x_88); - x_118 = lean_box(0); -} -if (lean_is_scalar(x_118)) { - x_119 = lean_alloc_ctor(1, 2, 0); -} else { - x_119 = x_118; -} -lean_ctor_set(x_119, 0, x_116); -lean_ctor_set(x_119, 1, x_117); -return x_119; -} -} -} -else -{ -uint8_t x_120; -lean_dec(x_10); -lean_dec(x_4); -lean_dec(x_3); -x_120 = !lean_is_exclusive(x_16); -if (x_120 == 0) -{ -lean_object* x_121; lean_object* x_122; -x_121 = lean_ctor_get(x_16, 0); -x_122 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_122, 0, x_121); -lean_ctor_set(x_16, 0, x_122); -return x_16; -} -else -{ -lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; -x_123 = lean_ctor_get(x_16, 0); -x_124 = lean_ctor_get(x_16, 1); -lean_inc(x_124); -lean_inc(x_123); -lean_dec(x_16); -x_125 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_125, 0, x_123); -x_126 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_126, 0, x_125); -lean_ctor_set(x_126, 1, x_124); -return x_126; -} -} -} -else -{ -uint8_t x_127; -lean_dec(x_10); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_127 = !lean_is_exclusive(x_12); -if (x_127 == 0) -{ -lean_object* x_128; lean_object* x_129; -x_128 = lean_ctor_get(x_12, 0); -x_129 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_129, 0, x_128); -lean_ctor_set(x_12, 0, x_129); -return x_12; -} -else -{ -lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; -x_130 = lean_ctor_get(x_12, 0); -x_131 = lean_ctor_get(x_12, 1); -lean_inc(x_131); -lean_inc(x_130); +lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_92 = lean_ctor_get(x_12, 0); +x_93 = lean_ctor_get(x_12, 1); +lean_inc(x_93); +lean_inc(x_92); lean_dec(x_12); -x_132 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_132, 0, x_130); -x_133 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_133, 0, x_132); -lean_ctor_set(x_133, 1, x_131); -return x_133; +x_94 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_94, 0, x_92); +x_95 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_95, 0, x_94); +lean_ctor_set(x_95, 1, x_93); +return x_95; } } } @@ -1271,381 +1056,373 @@ lean_dec(x_6); x_9 = !lean_is_exclusive(x_7); if (x_9 == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; x_10 = lean_ctor_get(x_7, 1); x_11 = lean_ctor_get(x_7, 0); lean_dec(x_11); -x_12 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_AsyncElabState_lastSnap(x_10); -lean_inc(x_12); -x_13 = l_Lean_Server_Snapshots_Snapshot_isAtEnd(x_12); -if (x_13 == 0) +x_12 = l_Lean_Server_Snapshots_instInhabitedSnapshot; +x_13 = l_Array_back___rarg(x_12, x_10); +lean_inc(x_13); +x_14 = l_Lean_Server_Snapshots_Snapshot_isAtEnd(x_13); +if (x_14 == 0) { -lean_object* x_14; lean_object* x_15; +lean_object* x_15; lean_object* x_16; lean_free_object(x_7); -x_14 = lean_box(0); +x_15 = lean_box(0); lean_inc(x_10); -x_15 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_nextCmdSnap___lambda__1(x_12, x_1, x_2, x_10, x_3, x_14, x_10, x_8); +x_16 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_nextCmdSnap___lambda__1(x_13, x_1, x_2, x_10, x_3, x_15, x_10, x_8); lean_dec(x_3); lean_dec(x_10); -return x_15; +return x_16; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_dec(x_3); -x_16 = l_Lean_Server_Snapshots_Snapshot_diagnostics(x_12); -x_17 = l_Std_PersistentArray_toArray___rarg(x_16); -x_18 = lean_ctor_get(x_1, 1); -lean_inc(x_18); +x_17 = l_Lean_Server_Snapshots_Snapshot_diagnostics(x_13); +x_18 = l_Std_PersistentArray_toArray___rarg(x_17); +x_19 = lean_ctor_get(x_1, 1); +lean_inc(x_19); lean_dec(x_1); -lean_inc(x_18); +lean_inc(x_19); lean_inc(x_2); -x_19 = l_Lean_Server_publishDiagnostics(x_2, x_17, x_18, x_8); -if (lean_obj_tag(x_19) == 0) +x_20 = l_Lean_Server_publishDiagnostics(x_2, x_18, x_19, x_8); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_19, 1); -lean_inc(x_20); -lean_dec(x_19); -lean_inc(x_18); +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_20, 1); +lean_inc(x_21); +lean_dec(x_20); +lean_inc(x_19); lean_inc(x_2); -x_21 = l_Lean_Server_publishProgressDone(x_2, x_18, x_20); -if (lean_obj_tag(x_21) == 0) +x_22 = l_Lean_Server_publishProgressDone(x_2, x_19, x_21); +if (lean_obj_tag(x_22) == 0) { -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; -x_22 = lean_ctor_get(x_21, 1); -lean_inc(x_22); -lean_dec(x_21); -x_23 = lean_ctor_get(x_10, 0); +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_22, 1); lean_inc(x_23); -x_24 = lean_ctor_get(x_10, 1); -lean_inc(x_24); -x_25 = lean_unsigned_to_nat(0u); -x_26 = l_Array_insertAt___rarg(x_24, x_25, x_23); -x_27 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishIleanInfoFinal___closed__1; -x_28 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishIleanInfo(x_27, x_2, x_18, x_26, x_22); -if (lean_obj_tag(x_28) == 0) +lean_dec(x_22); +x_24 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishIleanInfoFinal___closed__1; +lean_inc(x_10); +x_25 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishIleanInfo(x_24, x_2, x_19, x_10, x_23); +if (lean_obj_tag(x_25) == 0) { -uint8_t x_29; -x_29 = !lean_is_exclusive(x_28); -if (x_29 == 0) +uint8_t x_26; +x_26 = !lean_is_exclusive(x_25); +if (x_26 == 0) { -lean_object* x_30; lean_object* x_31; -x_30 = lean_ctor_get(x_28, 0); -lean_dec(x_30); -x_31 = lean_box(0); -lean_ctor_set(x_7, 0, x_31); -lean_ctor_set(x_28, 0, x_7); -return x_28; +lean_object* x_27; lean_object* x_28; +x_27 = lean_ctor_get(x_25, 0); +lean_dec(x_27); +x_28 = lean_box(0); +lean_ctor_set(x_7, 0, x_28); +lean_ctor_set(x_25, 0, x_7); +return x_25; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_28, 1); -lean_inc(x_32); -lean_dec(x_28); -x_33 = lean_box(0); -lean_ctor_set(x_7, 0, x_33); -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_7); -lean_ctor_set(x_34, 1, x_32); -return x_34; +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_25, 1); +lean_inc(x_29); +lean_dec(x_25); +x_30 = lean_box(0); +lean_ctor_set(x_7, 0, x_30); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_7); +lean_ctor_set(x_31, 1, x_29); +return x_31; } } else { -uint8_t x_35; +uint8_t x_32; lean_free_object(x_7); lean_dec(x_10); -x_35 = !lean_is_exclusive(x_28); -if (x_35 == 0) +x_32 = !lean_is_exclusive(x_25); +if (x_32 == 0) { -lean_object* x_36; lean_object* x_37; -x_36 = lean_ctor_get(x_28, 0); -x_37 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_28, 0, x_37); -return x_28; +lean_object* x_33; lean_object* x_34; +x_33 = lean_ctor_get(x_25, 0); +x_34 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_25, 0, x_34); +return x_25; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_38 = lean_ctor_get(x_28, 0); -x_39 = lean_ctor_get(x_28, 1); -lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_28); -x_40 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_40, 0, x_38); -x_41 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set(x_41, 1, x_39); -return x_41; +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_35 = lean_ctor_get(x_25, 0); +x_36 = lean_ctor_get(x_25, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_25); +x_37 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_37, 0, x_35); +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_36); +return x_38; } } } else { -uint8_t x_42; -lean_dec(x_18); +uint8_t x_39; +lean_dec(x_19); lean_free_object(x_7); lean_dec(x_10); lean_dec(x_2); -x_42 = !lean_is_exclusive(x_21); -if (x_42 == 0) +x_39 = !lean_is_exclusive(x_22); +if (x_39 == 0) { -lean_object* x_43; lean_object* x_44; -x_43 = lean_ctor_get(x_21, 0); -x_44 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_21, 0, x_44); -return x_21; +lean_object* x_40; lean_object* x_41; +x_40 = lean_ctor_get(x_22, 0); +x_41 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_22, 0, x_41); +return x_22; } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_45 = lean_ctor_get(x_21, 0); -x_46 = lean_ctor_get(x_21, 1); -lean_inc(x_46); -lean_inc(x_45); -lean_dec(x_21); -x_47 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_47, 0, x_45); -x_48 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_46); -return x_48; +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_42 = lean_ctor_get(x_22, 0); +x_43 = lean_ctor_get(x_22, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_22); +x_44 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_44, 0, x_42); +x_45 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_43); +return x_45; } } } else { -uint8_t x_49; -lean_dec(x_18); +uint8_t x_46; +lean_dec(x_19); lean_free_object(x_7); lean_dec(x_10); lean_dec(x_2); -x_49 = !lean_is_exclusive(x_19); -if (x_49 == 0) +x_46 = !lean_is_exclusive(x_20); +if (x_46 == 0) { -lean_object* x_50; lean_object* x_51; -x_50 = lean_ctor_get(x_19, 0); -x_51 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_19, 0, x_51); -return x_19; +lean_object* x_47; lean_object* x_48; +x_47 = lean_ctor_get(x_20, 0); +x_48 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_20, 0, x_48); +return x_20; } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_52 = lean_ctor_get(x_19, 0); -x_53 = lean_ctor_get(x_19, 1); -lean_inc(x_53); -lean_inc(x_52); -lean_dec(x_19); -x_54 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_54, 0, x_52); -x_55 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_55, 0, x_54); -lean_ctor_set(x_55, 1, x_53); -return x_55; +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_49 = lean_ctor_get(x_20, 0); +x_50 = lean_ctor_get(x_20, 1); +lean_inc(x_50); +lean_inc(x_49); +lean_dec(x_20); +x_51 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_51, 0, x_49); +x_52 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_50); +return x_52; } } } } else { -lean_object* x_56; lean_object* x_57; uint8_t x_58; -x_56 = lean_ctor_get(x_7, 1); -lean_inc(x_56); +lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; +x_53 = lean_ctor_get(x_7, 1); +lean_inc(x_53); lean_dec(x_7); -x_57 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_AsyncElabState_lastSnap(x_56); -lean_inc(x_57); -x_58 = l_Lean_Server_Snapshots_Snapshot_isAtEnd(x_57); -if (x_58 == 0) +x_54 = l_Lean_Server_Snapshots_instInhabitedSnapshot; +x_55 = l_Array_back___rarg(x_54, x_53); +lean_inc(x_55); +x_56 = l_Lean_Server_Snapshots_Snapshot_isAtEnd(x_55); +if (x_56 == 0) { -lean_object* x_59; lean_object* x_60; -x_59 = lean_box(0); -lean_inc(x_56); -x_60 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_nextCmdSnap___lambda__1(x_57, x_1, x_2, x_56, x_3, x_59, x_56, x_8); +lean_object* x_57; lean_object* x_58; +x_57 = lean_box(0); +lean_inc(x_53); +x_58 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_nextCmdSnap___lambda__1(x_55, x_1, x_2, x_53, x_3, x_57, x_53, x_8); lean_dec(x_3); -lean_dec(x_56); -return x_60; +lean_dec(x_53); +return x_58; } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_dec(x_3); -x_61 = l_Lean_Server_Snapshots_Snapshot_diagnostics(x_57); -x_62 = l_Std_PersistentArray_toArray___rarg(x_61); -x_63 = lean_ctor_get(x_1, 1); -lean_inc(x_63); +x_59 = l_Lean_Server_Snapshots_Snapshot_diagnostics(x_55); +x_60 = l_Std_PersistentArray_toArray___rarg(x_59); +x_61 = lean_ctor_get(x_1, 1); +lean_inc(x_61); lean_dec(x_1); +lean_inc(x_61); +lean_inc(x_2); +x_62 = l_Lean_Server_publishDiagnostics(x_2, x_60, x_61, x_8); +if (lean_obj_tag(x_62) == 0) +{ +lean_object* x_63; lean_object* x_64; +x_63 = lean_ctor_get(x_62, 1); lean_inc(x_63); +lean_dec(x_62); +lean_inc(x_61); lean_inc(x_2); -x_64 = l_Lean_Server_publishDiagnostics(x_2, x_62, x_63, x_8); +x_64 = l_Lean_Server_publishProgressDone(x_2, x_61, x_63); if (lean_obj_tag(x_64) == 0) { -lean_object* x_65; lean_object* x_66; +lean_object* x_65; lean_object* x_66; lean_object* x_67; x_65 = lean_ctor_get(x_64, 1); lean_inc(x_65); lean_dec(x_64); -lean_inc(x_63); -lean_inc(x_2); -x_66 = l_Lean_Server_publishProgressDone(x_2, x_63, x_65); -if (lean_obj_tag(x_66) == 0) +x_66 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishIleanInfoFinal___closed__1; +lean_inc(x_53); +x_67 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishIleanInfo(x_66, x_2, x_61, x_53, x_65); +if (lean_obj_tag(x_67) == 0) { -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; -x_67 = lean_ctor_get(x_66, 1); -lean_inc(x_67); -lean_dec(x_66); -x_68 = lean_ctor_get(x_56, 0); +lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_68 = lean_ctor_get(x_67, 1); lean_inc(x_68); -x_69 = lean_ctor_get(x_56, 1); -lean_inc(x_69); -x_70 = lean_unsigned_to_nat(0u); -x_71 = l_Array_insertAt___rarg(x_69, x_70, x_68); -x_72 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishIleanInfoFinal___closed__1; -x_73 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishIleanInfo(x_72, x_2, x_63, x_71, x_67); -if (lean_obj_tag(x_73) == 0) +if (lean_is_exclusive(x_67)) { + lean_ctor_release(x_67, 0); + lean_ctor_release(x_67, 1); + x_69 = x_67; +} else { + lean_dec_ref(x_67); + x_69 = lean_box(0); +} +x_70 = lean_box(0); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set(x_71, 1, x_53); +if (lean_is_scalar(x_69)) { + x_72 = lean_alloc_ctor(0, 2, 0); +} else { + x_72 = x_69; +} +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set(x_72, 1, x_68); +return x_72; +} +else { -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_74 = lean_ctor_get(x_73, 1); +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +lean_dec(x_53); +x_73 = lean_ctor_get(x_67, 0); +lean_inc(x_73); +x_74 = lean_ctor_get(x_67, 1); lean_inc(x_74); -if (lean_is_exclusive(x_73)) { - lean_ctor_release(x_73, 0); - lean_ctor_release(x_73, 1); - x_75 = x_73; +if (lean_is_exclusive(x_67)) { + lean_ctor_release(x_67, 0); + lean_ctor_release(x_67, 1); + x_75 = x_67; } else { - lean_dec_ref(x_73); + lean_dec_ref(x_67); x_75 = lean_box(0); } -x_76 = lean_box(0); -x_77 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_77, 0, x_76); -lean_ctor_set(x_77, 1, x_56); +x_76 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_76, 0, x_73); if (lean_is_scalar(x_75)) { - x_78 = lean_alloc_ctor(0, 2, 0); + x_77 = lean_alloc_ctor(1, 2, 0); } else { - x_78 = x_75; + x_77 = x_75; +} +lean_ctor_set(x_77, 0, x_76); +lean_ctor_set(x_77, 1, x_74); +return x_77; } -lean_ctor_set(x_78, 0, x_77); -lean_ctor_set(x_78, 1, x_74); -return x_78; } else { -lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; -lean_dec(x_56); -x_79 = lean_ctor_get(x_73, 0); +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; +lean_dec(x_61); +lean_dec(x_53); +lean_dec(x_2); +x_78 = lean_ctor_get(x_64, 0); +lean_inc(x_78); +x_79 = lean_ctor_get(x_64, 1); lean_inc(x_79); -x_80 = lean_ctor_get(x_73, 1); -lean_inc(x_80); -if (lean_is_exclusive(x_73)) { - lean_ctor_release(x_73, 0); - lean_ctor_release(x_73, 1); - x_81 = x_73; +if (lean_is_exclusive(x_64)) { + lean_ctor_release(x_64, 0); + lean_ctor_release(x_64, 1); + x_80 = x_64; } else { - lean_dec_ref(x_73); - x_81 = lean_box(0); + lean_dec_ref(x_64); + x_80 = lean_box(0); } -x_82 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_82, 0, x_79); -if (lean_is_scalar(x_81)) { - x_83 = lean_alloc_ctor(1, 2, 0); +x_81 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_81, 0, x_78); +if (lean_is_scalar(x_80)) { + x_82 = lean_alloc_ctor(1, 2, 0); } else { - x_83 = x_81; + x_82 = x_80; } -lean_ctor_set(x_83, 0, x_82); -lean_ctor_set(x_83, 1, x_80); -return x_83; +lean_ctor_set(x_82, 0, x_81); +lean_ctor_set(x_82, 1, x_79); +return x_82; } } else { -lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; -lean_dec(x_63); -lean_dec(x_56); +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +lean_dec(x_61); +lean_dec(x_53); lean_dec(x_2); -x_84 = lean_ctor_get(x_66, 0); +x_83 = lean_ctor_get(x_62, 0); +lean_inc(x_83); +x_84 = lean_ctor_get(x_62, 1); lean_inc(x_84); -x_85 = lean_ctor_get(x_66, 1); -lean_inc(x_85); -if (lean_is_exclusive(x_66)) { - lean_ctor_release(x_66, 0); - lean_ctor_release(x_66, 1); - x_86 = x_66; +if (lean_is_exclusive(x_62)) { + lean_ctor_release(x_62, 0); + lean_ctor_release(x_62, 1); + x_85 = x_62; } else { - lean_dec_ref(x_66); - x_86 = lean_box(0); + lean_dec_ref(x_62); + x_85 = lean_box(0); } -x_87 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_87, 0, x_84); -if (lean_is_scalar(x_86)) { - x_88 = lean_alloc_ctor(1, 2, 0); +x_86 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_86, 0, x_83); +if (lean_is_scalar(x_85)) { + x_87 = lean_alloc_ctor(1, 2, 0); } else { - x_88 = x_86; + x_87 = x_85; } -lean_ctor_set(x_88, 0, x_87); -lean_ctor_set(x_88, 1, x_85); -return x_88; -} -} -else -{ -lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; -lean_dec(x_63); -lean_dec(x_56); -lean_dec(x_2); -x_89 = lean_ctor_get(x_64, 0); -lean_inc(x_89); -x_90 = lean_ctor_get(x_64, 1); -lean_inc(x_90); -if (lean_is_exclusive(x_64)) { - lean_ctor_release(x_64, 0); - lean_ctor_release(x_64, 1); - x_91 = x_64; -} else { - lean_dec_ref(x_64); - x_91 = lean_box(0); -} -x_92 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_92, 0, x_89); -if (lean_is_scalar(x_91)) { - x_93 = lean_alloc_ctor(1, 2, 0); -} else { - x_93 = x_91; -} -lean_ctor_set(x_93, 0, x_92); -lean_ctor_set(x_93, 1, x_90); -return x_93; +lean_ctor_set(x_87, 0, x_86); +lean_ctor_set(x_87, 1, x_84); +return x_87; } } } } else { -uint8_t x_94; +uint8_t x_88; lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_94 = !lean_is_exclusive(x_6); -if (x_94 == 0) +x_88 = !lean_is_exclusive(x_6); +if (x_88 == 0) { return x_6; } else { -lean_object* x_95; lean_object* x_96; lean_object* x_97; -x_95 = lean_ctor_get(x_6, 0); -x_96 = lean_ctor_get(x_6, 1); -lean_inc(x_96); -lean_inc(x_95); +lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_89 = lean_ctor_get(x_6, 0); +x_90 = lean_ctor_get(x_6, 1); +lean_inc(x_90); +lean_inc(x_89); lean_dec(x_6); -x_97 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_97, 0, x_95); -lean_ctor_set(x_97, 1, x_96); -return x_97; +x_91 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_91, 0, x_89); +lean_ctor_set(x_91, 1, x_90); +return x_91; } } } @@ -1670,260 +1447,195 @@ lean_dec(x_5); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_unfoldCmdSnaps(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_Lean_Server_FileWorker_unfoldCmdSnaps(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -uint8_t x_7; -x_7 = l_Array_isEmpty___rarg(x_3); -if (x_7 == 0) +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_6 = l_Lean_Server_Snapshots_instInhabitedSnapshot; +x_7 = lean_unsigned_to_nat(0u); +x_8 = lean_array_get(x_6, x_2, x_7); +x_9 = l_Lean_Server_Snapshots_Snapshot_msgLog(x_8); +x_10 = l_Std_PersistentArray_anyM___at_Lean_MessageLog_hasErrors___spec__1(x_9); +if (x_10 == 0) { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_8 = lean_ctor_get(x_5, 1); -lean_inc(x_8); -x_9 = lean_unsigned_to_nat(0u); +lean_object* x_11; lean_object* x_12; lean_object* x_13; +lean_dec(x_8); +x_11 = lean_ctor_get(x_4, 1); +lean_inc(x_11); +x_12 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishIleanInfoUpdate___closed__1; lean_inc(x_2); -lean_inc(x_3); -x_10 = l_Array_insertAt___rarg(x_3, x_9, x_2); -x_11 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishIleanInfoUpdate___closed__1; lean_inc(x_1); -x_12 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishIleanInfo(x_11, x_1, x_8, x_10, x_6); -if (lean_obj_tag(x_12) == 0) +x_13 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishIleanInfo(x_12, x_1, x_11, x_2, x_5); +if (lean_obj_tag(x_13) == 0) { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_13 = lean_ctor_get(x_12, 1); -lean_inc(x_13); -lean_dec(x_12); -x_14 = lean_alloc_closure((void*)(l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_nextCmdSnap), 5, 3); -lean_closure_set(x_14, 0, x_5); -lean_closure_set(x_14, 1, x_1); -lean_closure_set(x_14, 2, x_4); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_2); -lean_ctor_set(x_15, 1, x_3); -x_16 = l_IO_AsyncList_unfoldAsync___rarg(x_14, x_15, x_13); +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_14 = lean_ctor_get(x_13, 1); +lean_inc(x_14); +lean_dec(x_13); +x_15 = lean_alloc_closure((void*)(l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_nextCmdSnap), 5, 3); +lean_closure_set(x_15, 0, x_4); +lean_closure_set(x_15, 1, x_1); +lean_closure_set(x_15, 2, x_3); +lean_inc(x_2); +x_16 = l_IO_AsyncList_unfoldAsync___rarg(x_15, x_2, x_14); x_17 = !lean_is_exclusive(x_16); if (x_17 == 0) { +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_18 = lean_ctor_get(x_16, 0); +x_19 = lean_array_to_list(lean_box(0), x_2); +x_20 = l_IO_AsyncList_ofList___rarg(x_19); +x_21 = l_IO_AsyncList_append___rarg(x_20, x_18); +lean_ctor_set(x_16, 0, x_21); return x_16; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_16, 0); -x_19 = lean_ctor_get(x_16, 1); -lean_inc(x_19); -lean_inc(x_18); +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_22 = lean_ctor_get(x_16, 0); +x_23 = lean_ctor_get(x_16, 1); +lean_inc(x_23); +lean_inc(x_22); lean_dec(x_16); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_18); -lean_ctor_set(x_20, 1, x_19); -return x_20; +x_24 = lean_array_to_list(lean_box(0), x_2); +x_25 = l_IO_AsyncList_ofList___rarg(x_24); +x_26 = l_IO_AsyncList_append___rarg(x_25, x_22); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_23); +return x_27; } } else { -uint8_t x_21; -lean_dec(x_5); +uint8_t x_28; lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_21 = !lean_is_exclusive(x_12); -if (x_21 == 0) -{ -return x_12; -} -else +x_28 = !lean_is_exclusive(x_13); +if (x_28 == 0) { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_12, 0); -x_23 = lean_ctor_get(x_12, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_12); -x_24 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_23); -return x_24; -} -} +return x_13; } else { -lean_object* x_25; uint8_t x_26; -x_25 = l_Lean_Server_Snapshots_Snapshot_msgLog(x_2); -x_26 = l_Std_PersistentArray_anyM___at_Lean_MessageLog_hasErrors___spec__1(x_25); -if (x_26 == 0) -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_27 = lean_ctor_get(x_5, 1); -lean_inc(x_27); -x_28 = lean_unsigned_to_nat(0u); -lean_inc(x_2); -lean_inc(x_3); -x_29 = l_Array_insertAt___rarg(x_3, x_28, x_2); -x_30 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishIleanInfoUpdate___closed__1; -lean_inc(x_1); -x_31 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishIleanInfo(x_30, x_1, x_27, x_29, x_6); -if (lean_obj_tag(x_31) == 0) -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; -x_32 = lean_ctor_get(x_31, 1); -lean_inc(x_32); -lean_dec(x_31); -x_33 = lean_alloc_closure((void*)(l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_nextCmdSnap), 5, 3); -lean_closure_set(x_33, 0, x_5); -lean_closure_set(x_33, 1, x_1); -lean_closure_set(x_33, 2, x_4); -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_2); -lean_ctor_set(x_34, 1, x_3); -x_35 = l_IO_AsyncList_unfoldAsync___rarg(x_33, x_34, x_32); -x_36 = !lean_is_exclusive(x_35); -if (x_36 == 0) -{ -return x_35; +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_13, 0); +x_30 = lean_ctor_get(x_13, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_13); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +return x_31; } -else -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_35, 0); -x_38 = lean_ctor_get(x_35, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_35); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_37); -lean_ctor_set(x_39, 1, x_38); -return x_39; } } else { -uint8_t x_40; -lean_dec(x_5); -lean_dec(x_4); +lean_object* x_32; lean_object* x_33; uint8_t x_34; lean_object* x_35; lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_40 = !lean_is_exclusive(x_31); -if (x_40 == 0) -{ -return x_31; -} -else -{ -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_31, 0); -x_42 = lean_ctor_get(x_31, 1); -lean_inc(x_42); -lean_inc(x_41); -lean_dec(x_31); -x_43 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_43, 0, x_41); -lean_ctor_set(x_43, 1, x_42); -return x_43; -} -} -} -else -{ -lean_object* x_44; lean_object* x_45; uint8_t x_46; lean_object* x_47; +x_32 = lean_ctor_get(x_8, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_4, 1); +lean_inc(x_33); lean_dec(x_4); -lean_dec(x_3); -x_44 = lean_ctor_get(x_2, 0); -lean_inc(x_44); -x_45 = lean_ctor_get(x_5, 1); -lean_inc(x_45); -lean_dec(x_5); -x_46 = 1; -lean_inc(x_45); +x_34 = 1; +lean_inc(x_33); lean_inc(x_1); -x_47 = l_Lean_Server_publishProgressAtPos(x_1, x_44, x_45, x_46, x_6); -lean_dec(x_44); -if (lean_obj_tag(x_47) == 0) +x_35 = l_Lean_Server_publishProgressAtPos(x_1, x_32, x_33, x_34, x_5); +lean_dec(x_32); +if (lean_obj_tag(x_35) == 0) { -lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_48 = lean_ctor_get(x_47, 1); -lean_inc(x_48); -lean_dec(x_47); -x_49 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_nextCmdSnap___lambda__1___closed__1; -x_50 = lean_array_push(x_49, x_2); -x_51 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishIleanInfoFinal___closed__1; -x_52 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishIleanInfo(x_51, x_1, x_45, x_50, x_48); -if (lean_obj_tag(x_52) == 0) +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; +x_36 = lean_ctor_get(x_35, 1); +lean_inc(x_36); +lean_dec(x_35); +x_37 = lean_box(0); +lean_inc(x_8); +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_8); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_nextCmdSnap___lambda__1___closed__1; +x_40 = lean_array_push(x_39, x_8); +x_41 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishIleanInfoFinal___closed__1; +x_42 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishIleanInfo(x_41, x_1, x_33, x_40, x_36); +if (lean_obj_tag(x_42) == 0) { -uint8_t x_53; -x_53 = !lean_is_exclusive(x_52); -if (x_53 == 0) +uint8_t x_43; +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) { -lean_object* x_54; lean_object* x_55; -x_54 = lean_ctor_get(x_52, 0); -lean_dec(x_54); -x_55 = lean_box(2); -lean_ctor_set(x_52, 0, x_55); -return x_52; +lean_object* x_44; lean_object* x_45; +x_44 = lean_ctor_get(x_42, 0); +lean_dec(x_44); +x_45 = l_IO_AsyncList_ofList___rarg(x_38); +lean_ctor_set(x_42, 0, x_45); +return x_42; } else { -lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_56 = lean_ctor_get(x_52, 1); -lean_inc(x_56); -lean_dec(x_52); -x_57 = lean_box(2); -x_58 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_58, 0, x_57); -lean_ctor_set(x_58, 1, x_56); -return x_58; +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_42, 1); +lean_inc(x_46); +lean_dec(x_42); +x_47 = l_IO_AsyncList_ofList___rarg(x_38); +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_46); +return x_48; } } else { -uint8_t x_59; -x_59 = !lean_is_exclusive(x_52); -if (x_59 == 0) +uint8_t x_49; +lean_dec(x_38); +x_49 = !lean_is_exclusive(x_42); +if (x_49 == 0) { -return x_52; +return x_42; } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_52, 0); -x_61 = lean_ctor_get(x_52, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_52); -x_62 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_62, 0, x_60); -lean_ctor_set(x_62, 1, x_61); -return x_62; +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_42, 0); +x_51 = lean_ctor_get(x_42, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_42); +x_52 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_50); +lean_ctor_set(x_52, 1, x_51); +return x_52; } } } else { -uint8_t x_63; -lean_dec(x_45); -lean_dec(x_2); +uint8_t x_53; +lean_dec(x_33); +lean_dec(x_8); lean_dec(x_1); -x_63 = !lean_is_exclusive(x_47); -if (x_63 == 0) +x_53 = !lean_is_exclusive(x_35); +if (x_53 == 0) { -return x_47; +return x_35; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_ctor_get(x_47, 0); -x_65 = lean_ctor_get(x_47, 1); -lean_inc(x_65); -lean_inc(x_64); -lean_dec(x_47); -x_66 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_66, 0, x_64); -lean_ctor_set(x_66, 1, x_65); -return x_66; -} +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_ctor_get(x_35, 0); +x_55 = lean_ctor_get(x_35, 1); +lean_inc(x_55); +lean_inc(x_54); +lean_dec(x_35); +x_56 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_56, 0, x_54); +lean_ctor_set(x_56, 1, x_55); +return x_56; } } } @@ -5123,7 +4835,7 @@ lean_inc(x_51); lean_dec(x_49); x_52 = l_System_FilePath_fileName(x_48); x_53 = l_Lean_Server_FileWorker_compileHeader___lambda__3___closed__4; -x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_SearchPath_findAllWithExt___spec__1(x_52, x_53); +x_54 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_SearchPath_findAllWithExt___spec__1(x_52, x_53); lean_dec(x_52); if (x_54 == 0) { @@ -5835,284 +5547,323 @@ x_7 = l_Lean_Server_FileWorker_compileHeader(x_1, x_2, x_3, x_6, x_5); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initializeWorker(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_EXPORT lean_object* l_Lean_Server_FileWorker_initializeWorker___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -uint8_t x_8; lean_object* x_61; -x_61 = lean_ctor_get(x_5, 3); -lean_inc(x_61); -if (lean_obj_tag(x_61) == 0) -{ -uint8_t x_62; -x_62 = 0; -x_8 = x_62; -goto block_60; -} -else +if (lean_obj_tag(x_4) == 0) { -lean_object* x_63; lean_object* x_64; -x_63 = lean_ctor_get(x_61, 0); -lean_inc(x_63); -lean_dec(x_61); -x_64 = lean_ctor_get(x_63, 1); -lean_inc(x_64); -lean_dec(x_63); -if (lean_obj_tag(x_64) == 0) +uint8_t x_6; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_6 = !lean_is_exclusive(x_4); +if (x_6 == 0) { -uint8_t x_65; -x_65 = 0; -x_8 = x_65; -goto block_60; +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_ctor_get(x_4, 0); +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_4, 0, x_8); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_4); +lean_ctor_set(x_9, 1, x_5); +return x_9; } else { -lean_object* x_66; uint8_t x_67; -x_66 = lean_ctor_get(x_64, 0); -lean_inc(x_66); -lean_dec(x_64); -x_67 = lean_unbox(x_66); -lean_dec(x_66); -x_8 = x_67; -goto block_60; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_ctor_get(x_4, 0); +lean_inc(x_10); +lean_dec(x_4); +x_11 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_11, 0, x_10); +x_12 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_12, 0, x_11); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_5); +return x_13; } } -block_60: -{ -lean_object* x_9; -lean_inc(x_3); -lean_inc(x_1); -x_9 = l_Lean_Server_FileWorker_compileHeader(x_1, x_3, x_6, x_8, x_7); -if (lean_obj_tag(x_9) == 0) +else { -lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_9, 1); -lean_inc(x_11); -lean_dec(x_9); -x_12 = !lean_is_exclusive(x_10); -if (x_12 == 0) +uint8_t x_14; +x_14 = !lean_is_exclusive(x_4); +if (x_14 == 0) { -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; lean_object* x_20; -x_13 = lean_ctor_get(x_10, 0); -x_14 = lean_ctor_get(x_10, 1); -x_15 = l_Lean_Server_FileWorker_CancelToken_new(x_11); +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_15 = lean_ctor_get(x_4, 0); x_16 = lean_ctor_get(x_15, 0); lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); -lean_inc(x_17); lean_dec(x_15); -x_18 = lean_alloc_ctor(0, 5, 1); -lean_ctor_set(x_18, 0, x_2); -lean_ctor_set(x_18, 1, x_3); -lean_ctor_set(x_18, 2, x_4); -lean_ctor_set(x_18, 3, x_14); -lean_ctor_set(x_18, 4, x_5); -lean_ctor_set_uint8(x_18, sizeof(void*)*5, x_8); -x_19 = l_Lean_Server_FileWorker_lakeSetupSearchPath___closed__5; -lean_inc(x_18); -lean_inc(x_16); -lean_inc(x_13); -lean_inc(x_1); -x_20 = l_Lean_Server_FileWorker_unfoldCmdSnaps(x_1, x_13, x_19, x_16, x_18, x_17); -if (lean_obj_tag(x_20) == 0) +x_17 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_nextCmdSnap___lambda__1___closed__1; +x_18 = lean_array_push(x_17, x_16); +x_19 = l_Lean_Server_FileWorker_unfoldCmdSnaps(x_1, x_18, x_2, x_3, x_5); +if (lean_obj_tag(x_19) == 0) { -uint8_t x_21; -x_21 = !lean_is_exclusive(x_20); -if (x_21 == 0) +uint8_t x_20; +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_22 = lean_ctor_get(x_20, 0); -x_23 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_23, 0, x_1); -lean_ctor_set(x_23, 1, x_13); -lean_ctor_set(x_23, 2, x_22); -lean_ctor_set(x_23, 3, x_16); -x_24 = lean_box(0); -x_25 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_25, 0, x_23); -lean_ctor_set(x_25, 1, x_24); -lean_ctor_set(x_25, 2, x_24); -lean_ctor_set(x_10, 1, x_25); -lean_ctor_set(x_10, 0, x_18); -lean_ctor_set(x_20, 0, x_10); -return x_20; +lean_object* x_21; +x_21 = lean_ctor_get(x_19, 0); +lean_ctor_set(x_4, 0, x_21); +lean_ctor_set(x_19, 0, x_4); +return x_19; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_26 = lean_ctor_get(x_20, 0); -x_27 = lean_ctor_get(x_20, 1); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_20); -x_28 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_28, 0, x_1); -lean_ctor_set(x_28, 1, x_13); -lean_ctor_set(x_28, 2, x_26); -lean_ctor_set(x_28, 3, x_16); -x_29 = lean_box(0); -x_30 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -lean_ctor_set(x_30, 2, x_29); -lean_ctor_set(x_10, 1, x_30); -lean_ctor_set(x_10, 0, x_18); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_10); -lean_ctor_set(x_31, 1, x_27); -return x_31; +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_19, 0); +x_23 = lean_ctor_get(x_19, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_19); +lean_ctor_set(x_4, 0, x_22); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_4); +lean_ctor_set(x_24, 1, x_23); +return x_24; } } else { -uint8_t x_32; -lean_dec(x_18); -lean_dec(x_16); -lean_free_object(x_10); -lean_dec(x_13); -lean_dec(x_1); -x_32 = !lean_is_exclusive(x_20); -if (x_32 == 0) +uint8_t x_25; +x_25 = !lean_is_exclusive(x_19); +if (x_25 == 0) { -return x_20; +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_19, 0); +x_27 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set_tag(x_4, 0); +lean_ctor_set(x_4, 0, x_27); +lean_ctor_set_tag(x_19, 0); +lean_ctor_set(x_19, 0, x_4); +return x_19; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_20, 0); -x_34 = lean_ctor_get(x_20, 1); -lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_20); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -return x_35; +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_28 = lean_ctor_get(x_19, 0); +x_29 = lean_ctor_get(x_19, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_19); +x_30 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set_tag(x_4, 0); +lean_ctor_set(x_4, 0, x_30); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_4); +lean_ctor_set(x_31, 1, x_29); +return x_31; } } } else { -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_36 = lean_ctor_get(x_10, 0); -x_37 = lean_ctor_get(x_10, 1); -lean_inc(x_37); -lean_inc(x_36); -lean_dec(x_10); -x_38 = l_Lean_Server_FileWorker_CancelToken_new(x_11); -x_39 = lean_ctor_get(x_38, 0); -lean_inc(x_39); -x_40 = lean_ctor_get(x_38, 1); -lean_inc(x_40); -lean_dec(x_38); -x_41 = lean_alloc_ctor(0, 5, 1); -lean_ctor_set(x_41, 0, x_2); -lean_ctor_set(x_41, 1, x_3); -lean_ctor_set(x_41, 2, x_4); -lean_ctor_set(x_41, 3, x_37); -lean_ctor_set(x_41, 4, x_5); -lean_ctor_set_uint8(x_41, sizeof(void*)*5, x_8); -x_42 = l_Lean_Server_FileWorker_lakeSetupSearchPath___closed__5; -lean_inc(x_41); -lean_inc(x_39); -lean_inc(x_36); -lean_inc(x_1); -x_43 = l_Lean_Server_FileWorker_unfoldCmdSnaps(x_1, x_36, x_42, x_39, x_41, x_40); -if (lean_obj_tag(x_43) == 0) +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_32 = lean_ctor_get(x_4, 0); +lean_inc(x_32); +lean_dec(x_4); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +lean_dec(x_32); +x_34 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_nextCmdSnap___lambda__1___closed__1; +x_35 = lean_array_push(x_34, x_33); +x_36 = l_Lean_Server_FileWorker_unfoldCmdSnaps(x_1, x_35, x_2, x_3, x_5); +if (lean_obj_tag(x_36) == 0) { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_44 = lean_ctor_get(x_43, 0); -lean_inc(x_44); -x_45 = lean_ctor_get(x_43, 1); -lean_inc(x_45); -if (lean_is_exclusive(x_43)) { - lean_ctor_release(x_43, 0); - lean_ctor_release(x_43, 1); - x_46 = x_43; +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +if (lean_is_exclusive(x_36)) { + lean_ctor_release(x_36, 0); + lean_ctor_release(x_36, 1); + x_39 = x_36; } else { - lean_dec_ref(x_43); - x_46 = lean_box(0); -} -x_47 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_47, 0, x_1); -lean_ctor_set(x_47, 1, x_36); -lean_ctor_set(x_47, 2, x_44); -lean_ctor_set(x_47, 3, x_39); -x_48 = lean_box(0); -x_49 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_49, 0, x_47); -lean_ctor_set(x_49, 1, x_48); -lean_ctor_set(x_49, 2, x_48); -x_50 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_50, 0, x_41); -lean_ctor_set(x_50, 1, x_49); -if (lean_is_scalar(x_46)) { - x_51 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_36); + x_39 = lean_box(0); +} +x_40 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_40, 0, x_37); +if (lean_is_scalar(x_39)) { + x_41 = lean_alloc_ctor(0, 2, 0); } else { - x_51 = x_46; + x_41 = x_39; } -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_45); -return x_51; +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_38); +return x_41; } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -lean_dec(x_41); -lean_dec(x_39); -lean_dec(x_36); -lean_dec(x_1); -x_52 = lean_ctor_get(x_43, 0); -lean_inc(x_52); -x_53 = lean_ctor_get(x_43, 1); -lean_inc(x_53); -if (lean_is_exclusive(x_43)) { - lean_ctor_release(x_43, 0); - lean_ctor_release(x_43, 1); - x_54 = x_43; +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_42 = lean_ctor_get(x_36, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_36, 1); +lean_inc(x_43); +if (lean_is_exclusive(x_36)) { + lean_ctor_release(x_36, 0); + lean_ctor_release(x_36, 1); + x_44 = x_36; } else { - lean_dec_ref(x_43); - x_54 = lean_box(0); + lean_dec_ref(x_36); + x_44 = lean_box(0); } -if (lean_is_scalar(x_54)) { - x_55 = lean_alloc_ctor(1, 2, 0); +x_45 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_45, 0, x_42); +x_46 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_46, 0, x_45); +if (lean_is_scalar(x_44)) { + x_47 = lean_alloc_ctor(0, 2, 0); } else { - x_55 = x_54; + x_47 = x_44; + lean_ctor_set_tag(x_47, 0); +} +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_43); +return x_47; +} +} } -lean_ctor_set(x_55, 0, x_52); -lean_ctor_set(x_55, 1, x_53); -return x_55; } } +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initializeWorker(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: +{ +uint8_t x_8; lean_object* x_38; +x_38 = lean_ctor_get(x_5, 3); +lean_inc(x_38); +if (lean_obj_tag(x_38) == 0) +{ +uint8_t x_39; +x_39 = 0; +x_8 = x_39; +goto block_37; } else { -uint8_t x_56; -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_56 = !lean_is_exclusive(x_9); -if (x_56 == 0) +lean_object* x_40; lean_object* x_41; +x_40 = lean_ctor_get(x_38, 0); +lean_inc(x_40); +lean_dec(x_38); +x_41 = lean_ctor_get(x_40, 1); +lean_inc(x_41); +lean_dec(x_40); +if (lean_obj_tag(x_41) == 0) { -return x_9; +uint8_t x_42; +x_42 = 0; +x_8 = x_42; +goto block_37; } else { -lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_57 = lean_ctor_get(x_9, 0); -x_58 = lean_ctor_get(x_9, 1); -lean_inc(x_58); -lean_inc(x_57); -lean_dec(x_9); -x_59 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_59, 0, x_57); -lean_ctor_set(x_59, 1, x_58); -return x_59; +lean_object* x_43; uint8_t x_44; +x_43 = lean_ctor_get(x_41, 0); +lean_inc(x_43); +lean_dec(x_41); +x_44 = lean_unbox(x_43); +lean_dec(x_43); +x_8 = x_44; +goto block_37; } } +block_37: +{ +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; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_9 = lean_box(x_8); +lean_inc(x_3); +lean_inc(x_1); +x_10 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_compileHeader___boxed), 5, 4); +lean_closure_set(x_10, 0, x_1); +lean_closure_set(x_10, 1, x_3); +lean_closure_set(x_10, 2, x_6); +lean_closure_set(x_10, 3, x_9); +x_11 = lean_alloc_closure((void*)(l_EIO_toBaseIO___rarg), 2, 1); +lean_closure_set(x_11, 0, x_10); +x_12 = l_Task_Priority_default; +x_13 = lean_io_as_task(x_11, x_12, x_7); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_Server_FileWorker_CancelToken_new(x_15); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +lean_inc(x_14); +x_19 = lean_alloc_ctor(0, 5, 1); +lean_ctor_set(x_19, 0, x_2); +lean_ctor_set(x_19, 1, x_3); +lean_ctor_set(x_19, 2, x_4); +lean_ctor_set(x_19, 3, x_14); +lean_ctor_set(x_19, 4, x_5); +lean_ctor_set_uint8(x_19, sizeof(void*)*5, x_8); +lean_inc(x_19); +lean_inc(x_17); +lean_inc(x_1); +x_20 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_initializeWorker___lambda__1), 5, 3); +lean_closure_set(x_20, 0, x_1); +lean_closure_set(x_20, 1, x_17); +lean_closure_set(x_20, 2, x_19); +x_21 = lean_io_map_task(x_20, x_14, x_12, x_18); +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_23 = lean_ctor_get(x_21, 0); +x_24 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_24, 0, x_23); +x_25 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_25, 0, x_1); +lean_ctor_set(x_25, 1, x_24); +lean_ctor_set(x_25, 2, x_17); +x_26 = lean_box(0); +x_27 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +lean_ctor_set(x_27, 2, x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_19); +lean_ctor_set(x_28, 1, x_27); +lean_ctor_set(x_21, 0, x_28); +return x_21; +} +else +{ +lean_object* x_29; 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; +x_29 = lean_ctor_get(x_21, 0); +x_30 = lean_ctor_get(x_21, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_21); +x_31 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_31, 0, x_29); +x_32 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_32, 0, x_1); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set(x_32, 2, x_17); +x_33 = lean_box(0); +x_34 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_34, 2, x_33); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_19); +lean_ctor_set(x_35, 1, x_34); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_30); +return x_36; +} } } } @@ -6215,140 +5966,133 @@ lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_updateDocument___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_Server_FileWorker_updateDocument___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) { _start: { -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 = l_Lean_Server_FileWorker_CancelToken_new(x_8); -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_9, 1); -lean_inc(x_11); -lean_dec(x_9); -x_12 = l_List_redLength___rarg(x_4); -x_13 = lean_mk_empty_array_with_capacity(x_12); -lean_dec(x_12); -lean_inc(x_4); -x_14 = l_List_toArrayAux___rarg(x_4, x_13); +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 = l_Lean_Server_FileWorker_CancelToken_new(x_7); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); lean_inc(x_10); -lean_inc(x_2); +lean_dec(x_8); +x_11 = l_List_redLength___rarg(x_3); +x_12 = lean_mk_empty_array_with_capacity(x_11); +lean_dec(x_11); +x_13 = l_List_toArrayAux___rarg(x_3, x_12); +lean_inc(x_9); lean_inc(x_1); -x_15 = l_Lean_Server_FileWorker_unfoldCmdSnaps(x_1, x_2, x_14, x_10, x_3, x_11); -if (lean_obj_tag(x_15) == 0) +x_14 = l_Lean_Server_FileWorker_unfoldCmdSnaps(x_1, x_13, x_9, x_2, x_10); +if (lean_obj_tag(x_14) == 0) { -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; uint8_t x_24; -x_16 = lean_ctor_get(x_15, 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 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); -lean_inc(x_17); -lean_dec(x_15); -x_18 = l_IO_AsyncList_ofList___rarg(x_4); -x_19 = l_IO_AsyncList_append___rarg(x_18, x_16); -x_20 = lean_st_ref_take(x_7, x_17); -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_20, 1); -lean_inc(x_22); -lean_dec(x_20); -x_23 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_23, 0, x_1); -lean_ctor_set(x_23, 1, x_2); -lean_ctor_set(x_23, 2, x_19); -lean_ctor_set(x_23, 3, x_10); -x_24 = !lean_is_exclusive(x_21); +lean_dec(x_14); +x_17 = lean_st_ref_take(x_6, x_16); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_20, 0, x_1); +lean_ctor_set(x_20, 1, x_15); +lean_ctor_set(x_20, 2, x_9); +x_21 = !lean_is_exclusive(x_18); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_22 = lean_ctor_get(x_18, 0); +lean_dec(x_22); +lean_ctor_set(x_18, 0, x_20); +x_23 = lean_st_ref_set(x_6, x_18, x_19); +x_24 = !lean_is_exclusive(x_23); if (x_24 == 0) { -lean_object* x_25; lean_object* x_26; uint8_t x_27; -x_25 = lean_ctor_get(x_21, 0); +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_23, 0); lean_dec(x_25); -lean_ctor_set(x_21, 0, x_23); -x_26 = lean_st_ref_set(x_7, x_21, x_22); -x_27 = !lean_is_exclusive(x_26); -if (x_27 == 0) -{ -lean_object* x_28; lean_object* x_29; -x_28 = lean_ctor_get(x_26, 0); -lean_dec(x_28); -x_29 = lean_box(0); -lean_ctor_set(x_26, 0, x_29); -return x_26; +x_26 = lean_box(0); +lean_ctor_set(x_23, 0, x_26); +return x_23; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_26, 1); -lean_inc(x_30); -lean_dec(x_26); -x_31 = lean_box(0); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_30); -return x_32; +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_23, 1); +lean_inc(x_27); +lean_dec(x_23); +x_28 = lean_box(0); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_27); +return x_29; } } else { -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; -x_33 = lean_ctor_get(x_21, 1); -x_34 = lean_ctor_get(x_21, 2); +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; +x_30 = lean_ctor_get(x_18, 1); +x_31 = lean_ctor_get(x_18, 2); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_18); +x_32 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_32, 0, x_20); +lean_ctor_set(x_32, 1, x_30); +lean_ctor_set(x_32, 2, x_31); +x_33 = lean_st_ref_set(x_6, x_32, x_19); +x_34 = lean_ctor_get(x_33, 1); lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_21); -x_35 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_35, 0, x_23); -lean_ctor_set(x_35, 1, x_33); -lean_ctor_set(x_35, 2, x_34); -x_36 = lean_st_ref_set(x_7, x_35, x_22); -x_37 = lean_ctor_get(x_36, 1); -lean_inc(x_37); -if (lean_is_exclusive(x_36)) { - lean_ctor_release(x_36, 0); - lean_ctor_release(x_36, 1); - x_38 = x_36; +if (lean_is_exclusive(x_33)) { + lean_ctor_release(x_33, 0); + lean_ctor_release(x_33, 1); + x_35 = x_33; } else { - lean_dec_ref(x_36); - x_38 = lean_box(0); + lean_dec_ref(x_33); + x_35 = lean_box(0); } -x_39 = lean_box(0); -if (lean_is_scalar(x_38)) { - x_40 = lean_alloc_ctor(0, 2, 0); +x_36 = lean_box(0); +if (lean_is_scalar(x_35)) { + x_37 = lean_alloc_ctor(0, 2, 0); } else { - x_40 = x_38; + x_37 = x_35; } -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_37); -return x_40; +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_34); +return x_37; } } else { -uint8_t x_41; -lean_dec(x_10); -lean_dec(x_4); -lean_dec(x_2); +uint8_t x_38; +lean_dec(x_9); lean_dec(x_1); -x_41 = !lean_is_exclusive(x_15); -if (x_41 == 0) +x_38 = !lean_is_exclusive(x_14); +if (x_38 == 0) { -return x_15; +return x_14; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_15, 0); -x_43 = lean_ctor_get(x_15, 1); -lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_15); -x_44 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_44, 1, x_43); -return x_44; +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_14, 0); +x_40 = lean_ctor_get(x_14, 1); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_14); +x_41 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +return x_41; } } } } -static lean_object* _init_l_Lean_Server_FileWorker_updateDocument___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_FileWorker_updateDocument___closed__1() { _start: { lean_object* x_1; @@ -6356,436 +6100,433 @@ x_1 = lean_mk_string_from_bytes("Internal server error: elab task was aborted wh return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_updateDocument___lambda__3(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_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_updateDocument(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_5 = lean_st_ref_get(x_3, x_4); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_5, 1); +lean_inc(x_7); +lean_dec(x_5); +x_8 = lean_ctor_get(x_6, 0); +lean_inc(x_8); lean_dec(x_6); -x_10 = lean_ctor_get(x_1, 2); -lean_inc(x_10); -x_11 = l_IO_AsyncList_updateFinishedPrefix___rarg(x_10, x_9); -x_12 = lean_ctor_get(x_11, 0); +x_9 = lean_ctor_get(x_8, 1); +lean_inc(x_9); +x_10 = l_IO_AsyncList_updateFinishedPrefix___rarg(x_9, x_7); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_11, 1); lean_inc(x_12); -x_13 = lean_ctor_get(x_12, 1); -lean_inc(x_13); -if (lean_obj_tag(x_13) == 0) +if (lean_obj_tag(x_12) == 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; 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_14 = lean_ctor_get(x_11, 1); +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); +lean_dec(x_10); +x_14 = lean_ctor_get(x_11, 0); lean_inc(x_14); lean_dec(x_11); -x_15 = lean_ctor_get(x_12, 0); +lean_inc(x_1); +x_15 = l_Lean_Server_DocumentMeta_mkInputContext(x_1); lean_inc(x_15); -lean_dec(x_12); -x_16 = lean_ctor_get(x_1, 3); -lean_inc(x_16); -x_17 = l_Lean_Server_FileWorker_CancelToken_set(x_16, x_14); -lean_dec(x_16); +x_16 = l_Lean_Parser_parseHeader(x_15, x_13); +if (lean_obj_tag(x_16) == 0) +{ +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; uint8_t x_25; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); x_18 = lean_ctor_get(x_17, 1); lean_inc(x_18); -lean_dec(x_17); -x_19 = lean_ctor_get(x_1, 0); +x_19 = lean_ctor_get(x_16, 1); lean_inc(x_19); -lean_dec(x_1); -x_20 = lean_ctor_get(x_19, 2); +lean_dec(x_16); +x_20 = lean_ctor_get(x_17, 0); lean_inc(x_20); -lean_dec(x_19); -x_21 = lean_ctor_get(x_20, 0); +lean_dec(x_17); +x_21 = lean_ctor_get(x_18, 0); lean_inc(x_21); -lean_dec(x_20); -x_22 = lean_ctor_get(x_2, 2); -lean_inc(x_22); -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); -lean_dec(x_22); -x_24 = l_String_firstDiffPos(x_21, x_23); -lean_dec(x_23); -lean_dec(x_21); -x_25 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_updateDocument___lambda__1___boxed), 2, 1); -lean_closure_set(x_25, 0, x_24); -x_26 = l_IO_AsyncList_finishedPrefix___rarg(x_15); -lean_dec(x_15); -x_27 = l_List_takeWhile___rarg(x_25, x_26); -x_28 = lean_unsigned_to_nat(0u); -x_29 = l_List_lengthTRAux___rarg(x_27, x_28); -x_30 = lean_nat_dec_eq(x_29, x_28); -if (x_30 == 0) -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; -x_31 = l_Lean_Server_Snapshots_instInhabitedSnapshot; -lean_inc(x_27); -x_32 = l_List_getLast_x21___rarg(x_31, x_27); -x_33 = lean_unsigned_to_nat(2u); -x_34 = lean_nat_dec_le(x_33, x_29); -if (x_34 == 0) +lean_dec(x_18); +x_22 = l_IO_AsyncList_finishedPrefix___rarg(x_14); +lean_dec(x_14); +x_23 = l_Lean_Server_Snapshots_instInhabitedSnapshot; +x_24 = l_List_head_x21___rarg(x_23, x_22); +x_25 = !lean_is_exclusive(x_24); +if (x_25 == 0) { -lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; +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_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; uint8_t x_42; +x_26 = lean_ctor_get(x_24, 2); +lean_dec(x_26); +x_27 = lean_ctor_get(x_24, 1); +lean_dec(x_27); +lean_ctor_set(x_24, 2, x_21); +lean_ctor_set(x_24, 1, x_20); +x_28 = lean_ctor_get(x_8, 2); +lean_inc(x_28); +x_29 = l_Lean_Server_FileWorker_CancelToken_set(x_28, x_19); +lean_dec(x_28); +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); lean_dec(x_29); -lean_inc(x_4); -x_35 = l_Lean_Server_Snapshots_parseNextCmd(x_3, x_4, x_18); -x_36 = lean_ctor_get(x_35, 0); -lean_inc(x_36); -x_37 = lean_ctor_get(x_35, 1); -lean_inc(x_37); -lean_dec(x_35); -x_38 = lean_ctor_get(x_32, 1); -lean_inc(x_38); -lean_dec(x_32); -x_39 = l_Lean_Syntax_structEq(x_36, x_38); -if (x_39 == 0) -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = l_List_dropLast___rarg(x_27); -x_41 = lean_box(0); -x_42 = l_Lean_Server_FileWorker_updateDocument___lambda__2(x_2, x_4, x_5, x_40, x_41, x_7, x_8, x_37); +x_31 = lean_ctor_get(x_8, 0); +lean_inc(x_31); lean_dec(x_8); -lean_dec(x_7); -return x_42; -} -else +x_32 = lean_ctor_get(x_31, 2); +lean_inc(x_32); +lean_dec(x_31); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +lean_dec(x_32); +x_34 = lean_ctor_get(x_1, 2); +lean_inc(x_34); +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +lean_dec(x_34); +x_36 = l_String_firstDiffPos(x_33, x_35); +lean_dec(x_35); +lean_dec(x_33); +x_37 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_updateDocument___lambda__1___boxed), 2, 1); +lean_closure_set(x_37, 0, x_36); +x_38 = l_List_takeWhile___rarg(x_37, x_22); +x_39 = lean_unsigned_to_nat(0u); +x_40 = l_List_lengthTRAux___rarg(x_38, x_39); +x_41 = lean_unsigned_to_nat(1u); +x_42 = lean_nat_dec_le(x_40, x_41); +if (x_42 == 0) { -lean_object* x_43; lean_object* x_44; -x_43 = lean_box(0); -x_44 = l_Lean_Server_FileWorker_updateDocument___lambda__2(x_2, x_4, x_5, x_27, x_43, x_7, x_8, x_37); -lean_dec(x_8); -lean_dec(x_7); -return x_44; -} -} -else +lean_object* x_43; lean_object* x_44; uint8_t x_45; +lean_inc(x_38); +x_43 = l_List_getLast_x21___rarg(x_23, x_38); +x_44 = lean_unsigned_to_nat(2u); +x_45 = lean_nat_dec_le(x_44, x_40); +if (x_45 == 0) { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; -x_45 = lean_nat_sub(x_29, x_33); -lean_dec(x_29); -x_46 = l_List_get_x21___rarg(x_31, x_27, x_45); -x_47 = l_Lean_Server_Snapshots_parseNextCmd(x_3, x_46, x_18); -x_48 = lean_ctor_get(x_47, 0); +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; +lean_dec(x_40); +x_46 = l_Lean_Server_Snapshots_parseNextCmd(x_15, x_24, x_30); +x_47 = lean_ctor_get(x_46, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_46, 1); lean_inc(x_48); -x_49 = lean_ctor_get(x_47, 1); +lean_dec(x_46); +x_49 = lean_ctor_get(x_43, 1); lean_inc(x_49); -lean_dec(x_47); -x_50 = lean_ctor_get(x_32, 1); -lean_inc(x_50); -lean_dec(x_32); -x_51 = l_Lean_Syntax_structEq(x_48, x_50); -if (x_51 == 0) +lean_dec(x_43); +x_50 = l_Lean_Syntax_structEq(x_47, x_49); +if (x_50 == 0) { -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = l_List_dropLast___rarg(x_27); -x_53 = lean_box(0); -x_54 = l_Lean_Server_FileWorker_updateDocument___lambda__2(x_2, x_4, x_5, x_52, x_53, x_7, x_8, x_49); -lean_dec(x_8); -lean_dec(x_7); -return x_54; +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = l_List_dropLast___rarg(x_38); +x_52 = lean_box(0); +lean_inc(x_2); +x_53 = l_Lean_Server_FileWorker_updateDocument___lambda__2(x_1, x_2, x_51, x_52, x_2, x_3, x_48); +lean_dec(x_3); +lean_dec(x_2); +return x_53; } else { -lean_object* x_55; lean_object* x_56; -x_55 = lean_box(0); -x_56 = l_Lean_Server_FileWorker_updateDocument___lambda__2(x_2, x_4, x_5, x_27, x_55, x_7, x_8, x_49); -lean_dec(x_8); -lean_dec(x_7); -return x_56; -} +lean_object* x_54; lean_object* x_55; +x_54 = lean_box(0); +lean_inc(x_2); +x_55 = l_Lean_Server_FileWorker_updateDocument___lambda__2(x_1, x_2, x_38, x_54, x_2, x_3, x_48); +lean_dec(x_3); +lean_dec(x_2); +return x_55; } } else { -lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; -lean_dec(x_29); -lean_dec(x_27); -lean_dec(x_7); -lean_dec(x_3); -x_57 = l_Lean_Server_FileWorker_CancelToken_new(x_18); -x_58 = lean_ctor_get(x_57, 0); -lean_inc(x_58); -x_59 = lean_ctor_get(x_57, 1); +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; +lean_dec(x_24); +x_56 = lean_nat_sub(x_40, x_44); +lean_dec(x_40); +x_57 = l_List_get_x21___rarg(x_23, x_38, x_56); +x_58 = l_Lean_Server_Snapshots_parseNextCmd(x_15, x_57, x_30); +x_59 = lean_ctor_get(x_58, 0); lean_inc(x_59); -lean_dec(x_57); -x_60 = l_Lean_Server_FileWorker_lakeSetupSearchPath___closed__5; -lean_inc(x_58); -lean_inc(x_4); -lean_inc(x_2); -x_61 = l_Lean_Server_FileWorker_unfoldCmdSnaps(x_2, x_4, x_60, x_58, x_5, x_59); -if (lean_obj_tag(x_61) == 0) -{ -lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; -x_62 = lean_ctor_get(x_61, 0); -lean_inc(x_62); -x_63 = lean_ctor_get(x_61, 1); -lean_inc(x_63); -lean_dec(x_61); -x_64 = lean_st_ref_take(x_8, x_63); -x_65 = lean_ctor_get(x_64, 0); -lean_inc(x_65); -x_66 = lean_ctor_get(x_64, 1); -lean_inc(x_66); -lean_dec(x_64); -x_67 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_67, 0, x_2); -lean_ctor_set(x_67, 1, x_4); -lean_ctor_set(x_67, 2, x_62); -lean_ctor_set(x_67, 3, x_58); -x_68 = !lean_is_exclusive(x_65); -if (x_68 == 0) -{ -lean_object* x_69; lean_object* x_70; uint8_t x_71; -x_69 = lean_ctor_get(x_65, 0); -lean_dec(x_69); -lean_ctor_set(x_65, 0, x_67); -x_70 = lean_st_ref_set(x_8, x_65, x_66); -lean_dec(x_8); -x_71 = !lean_is_exclusive(x_70); -if (x_71 == 0) +x_60 = lean_ctor_get(x_58, 1); +lean_inc(x_60); +lean_dec(x_58); +x_61 = lean_ctor_get(x_43, 1); +lean_inc(x_61); +lean_dec(x_43); +x_62 = l_Lean_Syntax_structEq(x_59, x_61); +if (x_62 == 0) { -lean_object* x_72; lean_object* x_73; -x_72 = lean_ctor_get(x_70, 0); -lean_dec(x_72); -x_73 = lean_box(0); -lean_ctor_set(x_70, 0, x_73); -return x_70; +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = l_List_dropLast___rarg(x_38); +x_64 = lean_box(0); +lean_inc(x_2); +x_65 = l_Lean_Server_FileWorker_updateDocument___lambda__2(x_1, x_2, x_63, x_64, x_2, x_3, x_60); +lean_dec(x_3); +lean_dec(x_2); +return x_65; } else { -lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_74 = lean_ctor_get(x_70, 1); -lean_inc(x_74); -lean_dec(x_70); -x_75 = lean_box(0); -x_76 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_76, 0, x_75); -lean_ctor_set(x_76, 1, x_74); -return x_76; +lean_object* x_66; lean_object* x_67; +x_66 = lean_box(0); +lean_inc(x_2); +x_67 = l_Lean_Server_FileWorker_updateDocument___lambda__2(x_1, x_2, x_38, x_66, x_2, x_3, x_60); +lean_dec(x_3); +lean_dec(x_2); +return x_67; +} } } else { -lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; -x_77 = lean_ctor_get(x_65, 1); -x_78 = lean_ctor_get(x_65, 2); -lean_inc(x_78); -lean_inc(x_77); -lean_dec(x_65); -x_79 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_79, 0, x_67); -lean_ctor_set(x_79, 1, x_77); -lean_ctor_set(x_79, 2, x_78); -x_80 = lean_st_ref_set(x_8, x_79, x_66); -lean_dec(x_8); -x_81 = lean_ctor_get(x_80, 1); -lean_inc(x_81); -if (lean_is_exclusive(x_80)) { - lean_ctor_release(x_80, 0); - lean_ctor_release(x_80, 1); - x_82 = x_80; -} else { - lean_dec_ref(x_80); - x_82 = lean_box(0); -} -x_83 = lean_box(0); -if (lean_is_scalar(x_82)) { - x_84 = lean_alloc_ctor(0, 2, 0); -} else { - x_84 = x_82; -} -lean_ctor_set(x_84, 0, x_83); -lean_ctor_set(x_84, 1, x_81); -return x_84; +lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; +lean_dec(x_40); +lean_dec(x_38); +lean_dec(x_15); +x_68 = lean_box(0); +x_69 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_69, 0, x_24); +lean_ctor_set(x_69, 1, x_68); +x_70 = lean_box(0); +lean_inc(x_2); +x_71 = l_Lean_Server_FileWorker_updateDocument___lambda__2(x_1, x_2, x_69, x_70, x_2, x_3, x_30); +lean_dec(x_3); +lean_dec(x_2); +return x_71; } } else { -uint8_t x_85; -lean_dec(x_58); +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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; uint8_t x_91; +x_72 = lean_ctor_get(x_24, 0); +x_73 = lean_ctor_get(x_24, 3); +x_74 = lean_ctor_get(x_24, 4); +x_75 = lean_ctor_get(x_24, 5); +lean_inc(x_75); +lean_inc(x_74); +lean_inc(x_73); +lean_inc(x_72); +lean_dec(x_24); +x_76 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_76, 0, x_72); +lean_ctor_set(x_76, 1, x_20); +lean_ctor_set(x_76, 2, x_21); +lean_ctor_set(x_76, 3, x_73); +lean_ctor_set(x_76, 4, x_74); +lean_ctor_set(x_76, 5, x_75); +x_77 = lean_ctor_get(x_8, 2); +lean_inc(x_77); +x_78 = l_Lean_Server_FileWorker_CancelToken_set(x_77, x_19); +lean_dec(x_77); +x_79 = lean_ctor_get(x_78, 1); +lean_inc(x_79); +lean_dec(x_78); +x_80 = lean_ctor_get(x_8, 0); +lean_inc(x_80); lean_dec(x_8); -lean_dec(x_4); -lean_dec(x_2); -x_85 = !lean_is_exclusive(x_61); -if (x_85 == 0) +x_81 = lean_ctor_get(x_80, 2); +lean_inc(x_81); +lean_dec(x_80); +x_82 = lean_ctor_get(x_81, 0); +lean_inc(x_82); +lean_dec(x_81); +x_83 = lean_ctor_get(x_1, 2); +lean_inc(x_83); +x_84 = lean_ctor_get(x_83, 0); +lean_inc(x_84); +lean_dec(x_83); +x_85 = l_String_firstDiffPos(x_82, x_84); +lean_dec(x_84); +lean_dec(x_82); +x_86 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_updateDocument___lambda__1___boxed), 2, 1); +lean_closure_set(x_86, 0, x_85); +x_87 = l_List_takeWhile___rarg(x_86, x_22); +x_88 = lean_unsigned_to_nat(0u); +x_89 = l_List_lengthTRAux___rarg(x_87, x_88); +x_90 = lean_unsigned_to_nat(1u); +x_91 = lean_nat_dec_le(x_89, x_90); +if (x_91 == 0) +{ +lean_object* x_92; lean_object* x_93; uint8_t x_94; +lean_inc(x_87); +x_92 = l_List_getLast_x21___rarg(x_23, x_87); +x_93 = lean_unsigned_to_nat(2u); +x_94 = lean_nat_dec_le(x_93, x_89); +if (x_94 == 0) { -return x_61; -} -else +lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; uint8_t x_99; +lean_dec(x_89); +x_95 = l_Lean_Server_Snapshots_parseNextCmd(x_15, x_76, x_79); +x_96 = lean_ctor_get(x_95, 0); +lean_inc(x_96); +x_97 = lean_ctor_get(x_95, 1); +lean_inc(x_97); +lean_dec(x_95); +x_98 = lean_ctor_get(x_92, 1); +lean_inc(x_98); +lean_dec(x_92); +x_99 = l_Lean_Syntax_structEq(x_96, x_98); +if (x_99 == 0) { -lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_86 = lean_ctor_get(x_61, 0); -x_87 = lean_ctor_get(x_61, 1); -lean_inc(x_87); -lean_inc(x_86); -lean_dec(x_61); -x_88 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_88, 0, x_86); -lean_ctor_set(x_88, 1, x_87); -return x_88; -} -} -} +lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_100 = l_List_dropLast___rarg(x_87); +x_101 = lean_box(0); +lean_inc(x_2); +x_102 = l_Lean_Server_FileWorker_updateDocument___lambda__2(x_1, x_2, x_100, x_101, x_2, x_3, x_97); +lean_dec(x_3); +lean_dec(x_2); +return x_102; } else { -lean_object* x_89; -lean_dec(x_12); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_4); +lean_object* x_103; lean_object* x_104; +x_103 = lean_box(0); +lean_inc(x_2); +x_104 = l_Lean_Server_FileWorker_updateDocument___lambda__2(x_1, x_2, x_87, x_103, x_2, x_3, x_97); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_89 = lean_ctor_get(x_13, 0); -lean_inc(x_89); -lean_dec(x_13); -if (lean_obj_tag(x_89) == 0) -{ -lean_object* x_90; lean_object* x_91; lean_object* x_92; -x_90 = lean_ctor_get(x_11, 1); -lean_inc(x_90); -lean_dec(x_11); -x_91 = l_Lean_Server_FileWorker_updateDocument___lambda__3___closed__1; -x_92 = l_IO_throwServerError___rarg(x_91, x_90); -return x_92; +return x_104; +} } else { -uint8_t x_93; -x_93 = !lean_is_exclusive(x_11); -if (x_93 == 0) -{ -lean_object* x_94; lean_object* x_95; -x_94 = lean_ctor_get(x_11, 0); -lean_dec(x_94); -x_95 = lean_ctor_get(x_89, 0); -lean_inc(x_95); +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; uint8_t x_111; +lean_dec(x_76); +x_105 = lean_nat_sub(x_89, x_93); lean_dec(x_89); -lean_ctor_set_tag(x_11, 1); -lean_ctor_set(x_11, 0, x_95); -return x_11; +x_106 = l_List_get_x21___rarg(x_23, x_87, x_105); +x_107 = l_Lean_Server_Snapshots_parseNextCmd(x_15, x_106, x_79); +x_108 = lean_ctor_get(x_107, 0); +lean_inc(x_108); +x_109 = lean_ctor_get(x_107, 1); +lean_inc(x_109); +lean_dec(x_107); +x_110 = lean_ctor_get(x_92, 1); +lean_inc(x_110); +lean_dec(x_92); +x_111 = l_Lean_Syntax_structEq(x_108, x_110); +if (x_111 == 0) +{ +lean_object* x_112; lean_object* x_113; lean_object* x_114; +x_112 = l_List_dropLast___rarg(x_87); +x_113 = lean_box(0); +lean_inc(x_2); +x_114 = l_Lean_Server_FileWorker_updateDocument___lambda__2(x_1, x_2, x_112, x_113, x_2, x_3, x_109); +lean_dec(x_3); +lean_dec(x_2); +return x_114; } else { -lean_object* x_96; lean_object* x_97; lean_object* x_98; -x_96 = lean_ctor_get(x_11, 1); -lean_inc(x_96); -lean_dec(x_11); -x_97 = lean_ctor_get(x_89, 0); -lean_inc(x_97); -lean_dec(x_89); -x_98 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_98, 0, x_97); -lean_ctor_set(x_98, 1, x_96); -return x_98; -} -} +lean_object* x_115; lean_object* x_116; +x_115 = lean_box(0); +lean_inc(x_2); +x_116 = l_Lean_Server_FileWorker_updateDocument___lambda__2(x_1, x_2, x_87, x_115, x_2, x_3, x_109); +lean_dec(x_3); +lean_dec(x_2); +return x_116; } } } -static lean_object* _init_l_Lean_Server_FileWorker_updateDocument___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Internal server error: header changed but worker wasn't restarted.", 66); -return x_1; +lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; +lean_dec(x_89); +lean_dec(x_87); +lean_dec(x_15); +x_117 = lean_box(0); +x_118 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_118, 0, x_76); +lean_ctor_set(x_118, 1, x_117); +x_119 = lean_box(0); +lean_inc(x_2); +x_120 = l_Lean_Server_FileWorker_updateDocument___lambda__2(x_1, x_2, x_118, x_119, x_2, x_3, x_79); +lean_dec(x_3); +lean_dec(x_2); +return x_120; } } -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_updateDocument(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; lean_object* x_11; -x_5 = lean_st_ref_get(x_3, x_4); -x_6 = lean_ctor_get(x_5, 0); -lean_inc(x_6); -x_7 = lean_ctor_get(x_5, 1); -lean_inc(x_7); -lean_dec(x_5); -x_8 = lean_ctor_get(x_6, 0); -lean_inc(x_8); -lean_dec(x_6); -lean_inc(x_1); -x_9 = l_Lean_Server_DocumentMeta_mkInputContext(x_1); -x_10 = lean_ctor_get(x_8, 1); -lean_inc(x_10); -lean_inc(x_10); -lean_inc(x_9); -x_11 = l_Lean_Server_Snapshots_reparseHeader(x_9, x_10, x_7); -if (lean_obj_tag(x_11) == 0) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); -x_15 = lean_ctor_get(x_10, 1); -lean_inc(x_15); -lean_dec(x_10); -x_16 = l_Lean_Syntax_structEq(x_14, x_15); -if (x_16 == 0) +} +else { -lean_object* x_17; lean_object* x_18; uint8_t x_19; -lean_dec(x_12); -lean_dec(x_9); +uint8_t x_121; +lean_dec(x_15); +lean_dec(x_14); lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_17 = l_Lean_Server_FileWorker_updateDocument___closed__1; -x_18 = l_IO_throwServerError___rarg(x_17, x_13); -x_19 = !lean_is_exclusive(x_18); -if (x_19 == 0) +x_121 = !lean_is_exclusive(x_16); +if (x_121 == 0) { -return x_18; +return x_16; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_18, 0); -x_21 = lean_ctor_get(x_18, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_18); -x_22 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_22, 0, x_20); -lean_ctor_set(x_22, 1, x_21); -return x_22; -} +lean_object* x_122; lean_object* x_123; lean_object* x_124; +x_122 = lean_ctor_get(x_16, 0); +x_123 = lean_ctor_get(x_16, 1); +lean_inc(x_123); +lean_inc(x_122); +lean_dec(x_16); +x_124 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_124, 0, x_122); +lean_ctor_set(x_124, 1, x_123); +return x_124; } -else -{ -lean_object* x_23; lean_object* x_24; -x_23 = lean_box(0); -lean_inc(x_2); -x_24 = l_Lean_Server_FileWorker_updateDocument___lambda__3(x_8, x_1, x_9, x_12, x_2, x_23, x_2, x_3, x_13); -return x_24; } } else { -uint8_t x_25; -lean_dec(x_10); -lean_dec(x_9); +lean_object* x_125; +lean_dec(x_11); lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_25 = !lean_is_exclusive(x_11); -if (x_25 == 0) +x_125 = lean_ctor_get(x_12, 0); +lean_inc(x_125); +lean_dec(x_12); +if (lean_obj_tag(x_125) == 0) { -return x_11; +lean_object* x_126; lean_object* x_127; lean_object* x_128; +x_126 = lean_ctor_get(x_10, 1); +lean_inc(x_126); +lean_dec(x_10); +x_127 = l_Lean_Server_FileWorker_updateDocument___closed__1; +x_128 = l_IO_throwServerError___rarg(x_127, x_126); +return x_128; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_11, 0); -x_27 = lean_ctor_get(x_11, 1); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_11); -x_28 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -return x_28; +uint8_t x_129; +x_129 = !lean_is_exclusive(x_10); +if (x_129 == 0) +{ +lean_object* x_130; lean_object* x_131; +x_130 = lean_ctor_get(x_10, 0); +lean_dec(x_130); +x_131 = lean_ctor_get(x_125, 0); +lean_inc(x_131); +lean_dec(x_125); +lean_ctor_set_tag(x_10, 1); +lean_ctor_set(x_10, 0, x_131); +return x_10; +} +else +{ +lean_object* x_132; lean_object* x_133; lean_object* x_134; +x_132 = lean_ctor_get(x_10, 1); +lean_inc(x_132); +lean_dec(x_10); +x_133 = lean_ctor_get(x_125, 0); +lean_inc(x_133); +lean_dec(x_125); +x_134 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_134, 0, x_133); +lean_ctor_set(x_134, 1, x_132); +return x_134; +} } } } @@ -6801,15 +6542,15 @@ x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_updateDocument___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, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_updateDocument___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, lean_object* x_7) { _start: { -lean_object* x_9; -x_9 = l_Lean_Server_FileWorker_updateDocument___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); +lean_object* x_8; +x_8 = l_Lean_Server_FileWorker_updateDocument___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_6); lean_dec(x_5); -return x_9; +lean_dec(x_4); +return x_8; } } static lean_object* _init_l_Lean_Server_FileWorker_handleDidChange___closed__1() { @@ -13369,7 +13110,37 @@ lean_dec(x_6); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_parseParams___at_Lean_Server_FileWorker_handleRequest___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_Server_FileWorker_handleRequest___spec__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +lean_dec(x_1); +x_4 = lean_io_error_to_string(x_3); +x_5 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_5, 0, x_4); +x_6 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_2); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_1, 0); +lean_inc(x_7); +lean_dec(x_1); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_2); +return x_8; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_parseParams___at_Lean_Server_FileWorker_handleRequest___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; @@ -13407,7 +13178,7 @@ return x_17; } } } -LEAN_EXPORT lean_object* l_IO_FS_Stream_writeLspResponse___at_Lean_Server_FileWorker_handleRequest___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_IO_FS_Stream_writeLspResponse___at_Lean_Server_FileWorker_handleRequest___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; uint64_t x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; @@ -13716,107 +13487,198 @@ return x_62; } } } -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleRequest___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_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleRequest___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) { _start: { -lean_object* x_10; lean_object* x_11; 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; -x_30 = lean_ctor_get(x_3, 2); -x_31 = lean_ctor_get(x_1, 2); +lean_object* x_8; lean_object* x_9; lean_object* x_30; +x_30 = l_IO_ofExcept___at_Lean_Server_FileWorker_handleRequest___spec__2(x_6, x_7); +if (lean_obj_tag(x_30) == 0) +{ +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; +x_31 = lean_ctor_get(x_30, 0); lean_inc(x_31); -x_32 = lean_ctor_get(x_1, 3); +x_32 = lean_ctor_get(x_30, 1); lean_inc(x_32); -x_33 = lean_ctor_get(x_1, 4); +lean_dec(x_30); +x_33 = lean_ctor_get(x_31, 1); lean_inc(x_33); -x_34 = lean_ctor_get(x_3, 0); +lean_dec(x_31); +x_34 = lean_ctor_get(x_1, 2); +x_35 = lean_ctor_get(x_1, 0); +x_36 = lean_ctor_get(x_2, 2); +lean_inc(x_36); +x_37 = lean_ctor_get(x_2, 4); +lean_inc(x_37); +lean_inc(x_35); lean_inc(x_34); -lean_inc(x_30); -x_35 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_35, 0, x_30); -lean_ctor_set(x_35, 1, x_32); -lean_ctor_set(x_35, 2, x_34); -lean_ctor_set(x_35, 3, x_31); -lean_ctor_set(x_35, 4, x_33); -x_36 = l_Lean_Server_handleLspRequest(x_4, x_5, x_35, x_9); -if (lean_obj_tag(x_36) == 0) +x_38 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_38, 0, x_34); +lean_ctor_set(x_38, 1, x_33); +lean_ctor_set(x_38, 2, x_35); +lean_ctor_set(x_38, 3, x_36); +lean_ctor_set(x_38, 4, x_37); +x_39 = l_Lean_Server_handleLspRequest(x_3, x_4, x_38, x_32); +if (lean_obj_tag(x_39) == 0) { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_36, 0); -lean_inc(x_37); -x_38 = lean_ctor_get(x_36, 1); -lean_inc(x_38); -lean_dec(x_36); -x_39 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_39, 0, x_37); -x_10 = x_39; -x_11 = x_38; +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +x_41 = lean_ctor_get(x_39, 1); +lean_inc(x_41); +lean_dec(x_39); +x_42 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_42, 0, x_40); +x_8 = x_42; +x_9 = x_41; +goto block_29; +} +else +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_39, 0); +lean_inc(x_43); +x_44 = lean_ctor_get(x_39, 1); +lean_inc(x_44); +lean_dec(x_39); +x_45 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_45, 0, x_43); +x_8 = x_45; +x_9 = x_44; goto block_29; } +} else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_36, 0); -lean_inc(x_40); -x_41 = lean_ctor_get(x_36, 1); -lean_inc(x_41); -lean_dec(x_36); -x_42 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_42, 0, x_40); -x_10 = x_42; -x_11 = x_41; -goto block_29; +uint8_t x_46; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_46 = !lean_is_exclusive(x_30); +if (x_46 == 0) +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_30, 0); +x_48 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_48, 0, x_47); +x_49 = lean_task_pure(x_48); +lean_ctor_set_tag(x_30, 0); +lean_ctor_set(x_30, 0, x_49); +return x_30; +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_50 = lean_ctor_get(x_30, 0); +x_51 = lean_ctor_get(x_30, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_30); +x_52 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_52, 0, x_50); +x_53 = lean_task_pure(x_52); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_53); +lean_ctor_set(x_54, 1, x_51); +return x_54; +} } block_29: { -if (lean_obj_tag(x_10) == 0) +if (lean_obj_tag(x_8) == 0) { -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; lean_object* x_20; lean_object* x_21; -x_12 = lean_ctor_get(x_10, 0); -lean_inc(x_12); +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; uint8_t x_17; +x_10 = lean_ctor_get(x_8, 0); +lean_inc(x_10); +lean_dec(x_8); +x_11 = lean_ctor_get(x_2, 1); +lean_inc(x_11); +lean_dec(x_2); +x_12 = l_Lean_Server_RequestError_toLspResponseError(x_5, x_10); lean_dec(x_10); -x_13 = lean_ctor_get(x_1, 1); -lean_inc(x_13); -lean_dec(x_1); -lean_inc(x_2); -x_14 = l_Lean_Server_RequestError_toLspResponseError(x_2, x_12); -lean_dec(x_12); -x_15 = lean_alloc_closure((void*)(l_IO_FS_Stream_writeLspResponseError___boxed), 3, 2); -lean_closure_set(x_15, 0, x_13); -lean_closure_set(x_15, 1, x_14); -x_16 = lean_alloc_closure((void*)(l_EIO_toBaseIO___rarg), 2, 1); -lean_closure_set(x_16, 0, x_15); -x_17 = l_Task_Priority_default; -x_18 = lean_io_as_task(x_16, x_17, x_11); -x_19 = lean_ctor_get(x_18, 0); +x_13 = lean_alloc_closure((void*)(l_IO_FS_Stream_writeLspResponseError___boxed), 3, 2); +lean_closure_set(x_13, 0, x_11); +lean_closure_set(x_13, 1, x_12); +x_14 = lean_alloc_closure((void*)(l_EIO_toBaseIO___rarg), 2, 1); +lean_closure_set(x_14, 0, x_13); +x_15 = l_Task_Priority_default; +x_16 = lean_io_as_task(x_14, x_15, x_9); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +return x_16; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_16, 0); +x_19 = lean_ctor_get(x_16, 1); lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); -lean_inc(x_20); -lean_dec(x_18); -x_21 = l_Lean_Server_FileWorker_queueRequest(x_2, x_19, x_7, x_8, x_20); -return x_21; +lean_inc(x_18); +lean_dec(x_16); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +return x_20; +} } else { -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; -x_22 = lean_ctor_get(x_10, 0); -lean_inc(x_22); -lean_dec(x_10); -lean_inc(x_2); -x_23 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleRequest___lambda__1), 4, 2); -lean_closure_set(x_23, 0, x_1); -lean_closure_set(x_23, 1, x_2); -x_24 = l_Task_Priority_default; -x_25 = lean_io_map_task(x_23, x_22, x_24, x_11); -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 1); +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_21 = lean_ctor_get(x_8, 0); +lean_inc(x_21); +lean_dec(x_8); +x_22 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleRequest___lambda__1), 4, 2); +lean_closure_set(x_22, 0, x_2); +lean_closure_set(x_22, 1, x_5); +x_23 = l_Task_Priority_default; +x_24 = lean_io_map_task(x_22, x_21, x_23, x_9); +x_25 = !lean_is_exclusive(x_24); +if (x_25 == 0) +{ +return x_24; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_24, 0); +x_27 = lean_ctor_get(x_24, 1); lean_inc(x_27); -lean_dec(x_25); -x_28 = l_Lean_Server_FileWorker_queueRequest(x_2, x_26, x_7, x_8, x_27); +lean_inc(x_26); +lean_dec(x_24); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); return x_28; } } } } +} +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleRequest___lambda__3(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_object* x_9) { +_start: +{ +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; +x_10 = lean_ctor_get(x_1, 3); +lean_inc(x_10); +lean_inc(x_5); +x_11 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleRequest___lambda__2___boxed), 7, 5); +lean_closure_set(x_11, 0, x_2); +lean_closure_set(x_11, 1, x_1); +lean_closure_set(x_11, 2, x_3); +lean_closure_set(x_11, 3, x_4); +lean_closure_set(x_11, 4, x_5); +x_12 = l_Task_Priority_default; +x_13 = lean_io_bind_task(x_10, x_11, x_12, x_9); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_Server_FileWorker_queueRequest(x_5, x_14, x_7, x_8, x_15); +return x_16; +} +} static lean_object* _init_l_Lean_Server_FileWorker_handleRequest___closed__1() { _start: { @@ -13842,9 +13704,8 @@ if (x_11 == 0) lean_object* x_12; lean_object* x_13; x_12 = lean_box(0); lean_inc(x_4); -x_13 = l_Lean_Server_FileWorker_handleRequest___lambda__2(x_4, x_1, x_8, x_2, x_3, x_12, x_4, x_5, x_9); +x_13 = l_Lean_Server_FileWorker_handleRequest___lambda__3(x_4, x_8, x_2, x_3, x_1, x_12, x_4, x_5, x_9); lean_dec(x_5); -lean_dec(x_8); lean_dec(x_4); return x_13; } @@ -13853,7 +13714,7 @@ else lean_object* x_14; lean_dec(x_8); lean_dec(x_2); -x_14 = l_Lean_Server_FileWorker_parseParams___at_Lean_Server_FileWorker_handleRequest___spec__2(x_3, x_4, x_5, x_9); +x_14 = l_Lean_Server_FileWorker_parseParams___at_Lean_Server_FileWorker_handleRequest___spec__3(x_3, x_4, x_5, x_9); if (lean_obj_tag(x_14) == 0) { lean_object* x_15; lean_object* x_16; @@ -13878,7 +13739,7 @@ x_20 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_20, 0, x_1); lean_ctor_set(x_20, 1, x_17); lean_inc(x_19); -x_21 = l_IO_FS_Stream_writeLspResponse___at_Lean_Server_FileWorker_handleRequest___spec__3(x_19, x_20, x_18); +x_21 = l_IO_FS_Stream_writeLspResponse___at_Lean_Server_FileWorker_handleRequest___spec__4(x_19, x_20, x_18); if (lean_obj_tag(x_21) == 0) { uint8_t x_22; @@ -14129,25 +13990,33 @@ lean_dec(x_2); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_parseParams___at_Lean_Server_FileWorker_handleRequest___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_parseParams___at_Lean_Server_FileWorker_handleRequest___spec__3___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_Server_FileWorker_parseParams___at_Lean_Server_FileWorker_handleRequest___spec__2(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_FileWorker_parseParams___at_Lean_Server_FileWorker_handleRequest___spec__3(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_Server_FileWorker_handleRequest___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, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleRequest___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, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_Server_FileWorker_handleRequest___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_1); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleRequest___lambda__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_object* x_9) { _start: { lean_object* x_10; -x_10 = l_Lean_Server_FileWorker_handleRequest___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_10 = l_Lean_Server_FileWorker_handleRequest___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_3); return x_10; } } @@ -15252,7 +15121,7 @@ lean_dec(x_63); x_66 = lean_ctor_get(x_64, 0); lean_inc(x_66); lean_dec(x_64); -x_67 = lean_ctor_get(x_66, 3); +x_67 = lean_ctor_get(x_66, 2); lean_inc(x_67); lean_dec(x_66); x_68 = l_Lean_Server_FileWorker_CancelToken_set(x_67, x_65); @@ -19204,7 +19073,7 @@ lean_inc(x_9); x_21 = l_IO_FS_Stream_readLspNotificationAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__3(x_9, x_20, x_19); if (lean_obj_tag(x_21) == 0) { -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; lean_object* x_30; lean_object* x_31; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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_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; lean_object* x_30; lean_object* x_31; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; x_22 = lean_ctor_get(x_21, 0); lean_inc(x_22); x_23 = lean_ctor_get(x_22, 1); @@ -19243,9 +19112,6 @@ lean_dec(x_18); lean_inc(x_14); lean_inc(x_29); x_65 = l_Lean_Server_FileWorker_initializeWorker(x_29, x_9, x_14, x_61, x_64, x_4, x_63); -if (lean_obj_tag(x_65) == 0) -{ -lean_object* x_66; 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; x_66 = lean_ctor_get(x_65, 0); lean_inc(x_66); x_67 = lean_ctor_get(x_65, 1); @@ -19310,19 +19176,6 @@ x_30 = x_82; x_31 = x_83; goto block_56; } -} -else -{ -lean_object* x_84; lean_object* x_85; -x_84 = lean_ctor_get(x_65, 0); -lean_inc(x_84); -x_85 = lean_ctor_get(x_65, 1); -lean_inc(x_85); -lean_dec(x_65); -x_30 = x_84; -x_31 = x_85; -goto block_56; -} block_56: { lean_object* x_32; @@ -19428,108 +19281,108 @@ return x_55; } else { -uint8_t x_86; +uint8_t x_84; lean_dec(x_18); lean_dec(x_14); lean_dec(x_9); lean_dec(x_4); lean_dec(x_3); -x_86 = !lean_is_exclusive(x_21); -if (x_86 == 0) +x_84 = !lean_is_exclusive(x_21); +if (x_84 == 0) { return x_21; } else { -lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_87 = lean_ctor_get(x_21, 0); -x_88 = lean_ctor_get(x_21, 1); -lean_inc(x_88); -lean_inc(x_87); +lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_85 = lean_ctor_get(x_21, 0); +x_86 = lean_ctor_get(x_21, 1); +lean_inc(x_86); +lean_inc(x_85); lean_dec(x_21); -x_89 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_89, 0, x_87); -lean_ctor_set(x_89, 1, x_88); -return x_89; +x_87 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_87, 0, x_85); +lean_ctor_set(x_87, 1, x_86); +return x_87; } } } else { -uint8_t x_90; +uint8_t x_88; lean_dec(x_14); lean_dec(x_9); lean_dec(x_4); lean_dec(x_3); -x_90 = !lean_is_exclusive(x_17); -if (x_90 == 0) +x_88 = !lean_is_exclusive(x_17); +if (x_88 == 0) { return x_17; } else { -lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_91 = lean_ctor_get(x_17, 0); -x_92 = lean_ctor_get(x_17, 1); -lean_inc(x_92); -lean_inc(x_91); +lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_89 = lean_ctor_get(x_17, 0); +x_90 = lean_ctor_get(x_17, 1); +lean_inc(x_90); +lean_inc(x_89); lean_dec(x_17); -x_93 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_93, 0, x_91); -lean_ctor_set(x_93, 1, x_92); -return x_93; +x_91 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_91, 0, x_89); +lean_ctor_set(x_91, 1, x_90); +return x_91; } } } else { -uint8_t x_94; +uint8_t x_92; lean_dec(x_9); lean_dec(x_4); lean_dec(x_3); -x_94 = !lean_is_exclusive(x_13); -if (x_94 == 0) +x_92 = !lean_is_exclusive(x_13); +if (x_92 == 0) { return x_13; } else { -lean_object* x_95; lean_object* x_96; lean_object* x_97; -x_95 = lean_ctor_get(x_13, 0); -x_96 = lean_ctor_get(x_13, 1); -lean_inc(x_96); -lean_inc(x_95); +lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_93 = lean_ctor_get(x_13, 0); +x_94 = lean_ctor_get(x_13, 1); +lean_inc(x_94); +lean_inc(x_93); lean_dec(x_13); -x_97 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_97, 0, x_95); -lean_ctor_set(x_97, 1, x_96); -return x_97; +x_95 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_95, 0, x_93); +lean_ctor_set(x_95, 1, x_94); +return x_95; } } } else { -uint8_t x_98; +uint8_t x_96; lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_98 = !lean_is_exclusive(x_8); -if (x_98 == 0) +x_96 = !lean_is_exclusive(x_8); +if (x_96 == 0) { return x_8; } else { -lean_object* x_99; lean_object* x_100; lean_object* x_101; -x_99 = lean_ctor_get(x_8, 0); -x_100 = lean_ctor_get(x_8, 1); -lean_inc(x_100); -lean_inc(x_99); +lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_97 = lean_ctor_get(x_8, 0); +x_98 = lean_ctor_get(x_8, 1); +lean_inc(x_98); +lean_inc(x_97); lean_dec(x_8); -x_101 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_101, 0, x_99); -lean_ctor_set(x_101, 1, x_100); -return x_101; +x_99 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_99, 0, x_97); +lean_ctor_set(x_99, 1, x_98); +return x_99; } } } @@ -19554,7 +19407,7 @@ return x_2; LEAN_EXPORT lean_object* lean_server_worker_main(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; lean_object* x_11; lean_object* x_12; lean_object* x_13; size_t x_31; lean_object* x_32; +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; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_31; x_3 = lean_get_stdin(x_2); x_4 = lean_ctor_get(x_3, 0); lean_inc(x_4); @@ -19573,150 +19426,115 @@ lean_inc(x_10); x_11 = lean_ctor_get(x_9, 1); lean_inc(x_11); lean_dec(x_9); -x_31 = 8; -x_32 = lean_io_get_random_bytes(x_31, x_11); -if (lean_obj_tag(x_32) == 0) +lean_inc(x_10); +lean_inc(x_7); +x_31 = l_Lean_Server_FileWorker_initAndRunWorker(x_4, x_7, x_10, x_1, x_11); +if (lean_obj_tag(x_31) == 0) { -lean_object* x_33; lean_object* x_34; lean_object* x_35; uint64_t x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_33 = lean_ctor_get(x_32, 0); +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_31, 1); lean_inc(x_33); -x_34 = lean_ctor_get(x_32, 1); +lean_dec(x_31); +x_34 = lean_ctor_get(x_7, 1); lean_inc(x_34); -lean_dec(x_32); -x_35 = l_ByteArray_toUInt64LE_x21(x_33); -x_36 = lean_unbox_uint64(x_35); +lean_dec(x_7); +x_35 = lean_apply_1(x_34, x_33); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_35, 1); +lean_inc(x_36); lean_dec(x_35); -x_37 = lean_uint64_to_nat(x_36); -x_38 = l_IO_setRandSeed(x_37, x_34); -lean_dec(x_37); +x_37 = lean_ctor_get(x_10, 1); +lean_inc(x_37); +x_38 = lean_apply_1(x_37, x_36); +if (lean_obj_tag(x_38) == 0) +{ +lean_object* x_39; uint32_t x_40; uint8_t x_41; lean_object* x_42; x_39 = lean_ctor_get(x_38, 1); lean_inc(x_39); lean_dec(x_38); -lean_inc(x_10); -lean_inc(x_7); -x_40 = l_Lean_Server_FileWorker_initAndRunWorker(x_4, x_7, x_10, x_1, x_39); -if (lean_obj_tag(x_40) == 0) -{ -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_41 = lean_ctor_get(x_40, 0); -lean_inc(x_41); -x_42 = lean_ctor_get(x_40, 1); -lean_inc(x_42); -lean_dec(x_40); -x_43 = lean_ctor_get(x_7, 1); -lean_inc(x_43); -lean_dec(x_7); -x_44 = lean_apply_1(x_43, x_42); -if (lean_obj_tag(x_44) == 0) -{ -lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_44, 1); -lean_inc(x_45); -lean_dec(x_44); -x_46 = lean_ctor_get(x_10, 1); -lean_inc(x_46); -x_47 = lean_apply_1(x_46, x_45); -if (lean_obj_tag(x_47) == 0) -{ -lean_object* x_48; uint32_t x_49; uint8_t x_50; lean_object* x_51; -x_48 = lean_ctor_get(x_47, 1); -lean_inc(x_48); -lean_dec(x_47); -x_49 = lean_unbox_uint32(x_41); -lean_dec(x_41); -x_50 = lean_uint32_to_uint8(x_49); -x_51 = lean_io_exit(x_50, x_48); -if (lean_obj_tag(x_51) == 0) +x_40 = lean_unbox_uint32(x_32); +lean_dec(x_32); +x_41 = lean_uint32_to_uint8(x_40); +x_42 = lean_io_exit(x_41, x_39); +if (lean_obj_tag(x_42) == 0) { -uint8_t x_52; +uint8_t x_43; lean_dec(x_10); -x_52 = !lean_is_exclusive(x_51); -if (x_52 == 0) -{ -return x_51; -} -else +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) { -lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_53 = lean_ctor_get(x_51, 0); -x_54 = lean_ctor_get(x_51, 1); -lean_inc(x_54); -lean_inc(x_53); -lean_dec(x_51); -x_55 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_55, 0, x_53); -lean_ctor_set(x_55, 1, x_54); -return x_55; -} +return x_42; } else { -lean_object* x_56; lean_object* x_57; -x_56 = lean_ctor_get(x_51, 0); -lean_inc(x_56); -x_57 = lean_ctor_get(x_51, 1); -lean_inc(x_57); -lean_dec(x_51); -x_12 = x_56; -x_13 = x_57; -goto block_30; +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_42, 0); +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_42); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; } } else { -lean_object* x_58; lean_object* x_59; -lean_dec(x_41); -x_58 = lean_ctor_get(x_47, 0); -lean_inc(x_58); -x_59 = lean_ctor_get(x_47, 1); -lean_inc(x_59); -lean_dec(x_47); -x_12 = x_58; -x_13 = x_59; +lean_object* x_47; lean_object* x_48; +x_47 = lean_ctor_get(x_42, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_42, 1); +lean_inc(x_48); +lean_dec(x_42); +x_12 = x_47; +x_13 = x_48; goto block_30; } } else { -lean_object* x_60; lean_object* x_61; -lean_dec(x_41); -x_60 = lean_ctor_get(x_44, 0); -lean_inc(x_60); -x_61 = lean_ctor_get(x_44, 1); -lean_inc(x_61); -lean_dec(x_44); -x_12 = x_60; -x_13 = x_61; +lean_object* x_49; lean_object* x_50; +lean_dec(x_32); +x_49 = lean_ctor_get(x_38, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_38, 1); +lean_inc(x_50); +lean_dec(x_38); +x_12 = x_49; +x_13 = x_50; goto block_30; } } else { -lean_object* x_62; lean_object* x_63; -lean_dec(x_7); -x_62 = lean_ctor_get(x_40, 0); -lean_inc(x_62); -x_63 = lean_ctor_get(x_40, 1); -lean_inc(x_63); -lean_dec(x_40); -x_12 = x_62; -x_13 = x_63; +lean_object* x_51; lean_object* x_52; +lean_dec(x_32); +x_51 = lean_ctor_get(x_35, 0); +lean_inc(x_51); +x_52 = lean_ctor_get(x_35, 1); +lean_inc(x_52); +lean_dec(x_35); +x_12 = x_51; +x_13 = x_52; goto block_30; } } else { -lean_object* x_64; lean_object* x_65; +lean_object* x_53; lean_object* x_54; lean_dec(x_7); -lean_dec(x_4); -lean_dec(x_1); -x_64 = lean_ctor_get(x_32, 0); -lean_inc(x_64); -x_65 = lean_ctor_get(x_32, 1); -lean_inc(x_65); -lean_dec(x_32); -x_12 = x_64; -x_13 = x_65; +x_53 = lean_ctor_get(x_31, 0); +lean_inc(x_53); +x_54 = lean_ctor_get(x_31, 1); +lean_inc(x_54); +lean_dec(x_31); +x_12 = x_53; +x_13 = x_54; goto block_30; } block_30: @@ -19941,8 +19759,6 @@ l_Lean_Server_FileWorker_compileHeader___closed__3 = _init_l_Lean_Server_FileWor lean_mark_persistent(l_Lean_Server_FileWorker_compileHeader___closed__3); l_Lean_Server_FileWorker_compileHeader___closed__4 = _init_l_Lean_Server_FileWorker_compileHeader___closed__4(); lean_mark_persistent(l_Lean_Server_FileWorker_compileHeader___closed__4); -l_Lean_Server_FileWorker_updateDocument___lambda__3___closed__1 = _init_l_Lean_Server_FileWorker_updateDocument___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_FileWorker_updateDocument___lambda__3___closed__1); l_Lean_Server_FileWorker_updateDocument___closed__1 = _init_l_Lean_Server_FileWorker_updateDocument___closed__1(); lean_mark_persistent(l_Lean_Server_FileWorker_updateDocument___closed__1); l_Lean_Server_FileWorker_handleDidChange___closed__1 = _init_l_Lean_Server_FileWorker_handleDidChange___closed__1(); diff --git a/stage0/stdlib/Lean/Server/FileWorker/RequestHandling.c b/stage0/stdlib/Lean/Server/FileWorker/RequestHandling.c index 2608ec12c57..a4690b13aa6 100644 --- a/stage0/stdlib/Lean/Server/FileWorker/RequestHandling.c +++ b/stage0/stdlib/Lean/Server/FileWorker/RequestHandling.c @@ -14,34 +14,39 @@ extern "C" { #endif lean_object* l_List_reverse___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Server_FileWorker_handleCompletion___spec__1___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__6; lean_object* l_String_csize(uint32_t); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__1(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__24___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handleDocumentHighlight___spec__1(size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__13; static lean_object* l_Lean_Server_FileWorker_handleFoldingRange_isImport___closed__4; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleWaitForDiagnostics___lambda__1___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__15(lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__3___closed__1; LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Server_FileWorker_handleSemanticTokens___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); lean_object* l_Lean_Expr_mvarId_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokensRange(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__2(lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__4; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_FileWorker_handleDocumentHighlight___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__21; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___closed__6; lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokens___lambda__1___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__2; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__3___closed__1; +static lean_object* l_Lean_Server_FileWorker_handleSemanticTokens_go___closed__3; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_HashMapImp_find_x3f___at_Lean_Server_References_updateWorkerRefs___spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleHover___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_1974_(lean_object*); -extern lean_object* l_Lean_nullKind; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Server_FileWorker_getInteractiveGoals___lambda__1(lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__4___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__22(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__19(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_sectionLikeToDocumentSymbols___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__3___closed__1; static lean_object* l_Lean_Server_FileWorker_handleFoldingRange_isImport___closed__1; static lean_object* l_Lean_Server_FileWorker_handleFoldingRange_isImport___closed__2; uint8_t l_Lean_LocalDecl_isAuxDecl(lean_object*); @@ -52,212 +57,216 @@ LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight___lamb lean_object* lean_array_uget(lean_object*, size_t); lean_object* lean_io_error_to_string(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_getInteractiveTermGoal___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1761_(lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1760_(lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__7; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__3___closed__2; lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoalParams____x40_Lean_Data_Lsp_Extra___hyg_490_(lean_object*); lean_object* l_Array_append___rarg(lean_object*, lean_object*); static lean_object* l_List_mapM___at_Lean_Server_FileWorker_getInteractiveGoals___spec__3___closed__7; -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handlePlainTermGoal(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__3___closed__1; +static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_sectionLikeToDocumentSymbols___closed__2; extern lean_object* l_Std_Format_defWidth; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_isImport___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__7; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__24___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleHover___closed__1; static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__3; static lean_object* l_Lean_Server_FileWorker_handleFoldingRange___rarg___closed__1; -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__20(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__23; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); static lean_object* l_Lean_Server_FileWorker_handlePlainGoal___closed__2; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__8; -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__17; -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__8; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__9___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__10; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__2___closed__1; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__24___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__11; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handleDocumentHighlight___spec__4___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__4___closed__1; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__30___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__2; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__3___closed__2; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__24___lambda__3___closed__1; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__3___closed__2; static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__1; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Server_Completion_find_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__3___closed__1; static lean_object* l_Lean_Server_FileWorker_handleFoldingRange_isImport___closed__3; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__12; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addRange___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__18; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addCommandRange___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__2; uint8_t lean_name_eq(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__24___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__14; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__24(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__1; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__3___closed__1; -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__12(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__22; +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__5; static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__4; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handlePlainGoal___closed__1; static lean_object* l_Lean_Server_FileWorker_handlePlainTermGoal___closed__2; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__3___closed__2; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__4___closed__1; lean_object* lean_array_push(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__24___lambda__3___closed__1; lean_object* l_Lean_Server_Snapshots_Snapshot_infoTree(lean_object*); lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainTermGoalParams____x40_Lean_Data_Lsp_Extra___hyg_705_(lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__4___closed__1; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__4___closed__1; lean_object* lean_string_append(lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__6; static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__1; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_sectionLikeToDocumentSymbols___closed__1; LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Server_FileWorker_handleDocumentSymbol___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokens_addToken___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__5___closed__1; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_withPPInaccessibleNames___at_Lean_Server_FileWorker_getInteractiveGoals___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__8; lean_object* l_Lean_FileMap_lspPosToUtf8Pos(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__4; static lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__11; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__2___closed__2; lean_object* l_IO_sleep(uint32_t, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___closed__2; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__2___closed__1; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__3___closed__2; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleWaitForDiagnostics___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__36___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__9; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletion___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__1(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__2(lean_object*); lean_object* l_Lean_Server_DocumentMeta_mkInputContext(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__1(lean_object*); static lean_object* l_Lean_Server_FileWorker_handleFoldingRange_isImport___closed__5; LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_handleWaitForDiagnostics___spec__1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__13; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_join___rarg(lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__2___boxed(lean_object*); LEAN_EXPORT lean_object* l_List_mapM___at_Lean_Server_FileWorker_getInteractiveGoals___spec__3___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__3; -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__23; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__4___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addRangeFromSyntax(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_mapM___at_Lean_Server_FileWorker_getInteractiveGoals___spec__3___closed__1; +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__25; static lean_object* l_Lean_Server_FileWorker_handleWaitForDiagnostics___lambda__1___closed__1; +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__24; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol___boxed(lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handlePlainTermGoal___closed__3; -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__26(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__1(lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__3___closed__2; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__3___closed__1; -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__14; lean_object* l_Lean_FileMap_utf8PosToLspPos(lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__5; -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__25; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__15; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleWaitForDiagnostics___lambda__1(lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__18; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__3___closed__1; lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonPlainTermGoal____x40_Lean_Data_Lsp_Extra___hyg_868_(lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__4___closed__1; lean_object* l_Lean_Elab_InfoTree_deepestNodes___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__1(lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__9(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__3___closed__2; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Server_FileWorker_handleSemanticTokens_go___spec__2(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight___lambda__2___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__1(lean_object*); lean_object* l_Lean_Widget_InteractiveGoal_pretty(lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___closed__5; static lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addCommandRange___closed__2; -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__22; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__7; +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__29___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleWaitForDiagnostics_waitLoop(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__3___closed__2; static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__2___closed__1; static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight___lambda__2(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Server_FileWorker_handleSemanticTokens_highlightId___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Except_map___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__2___closed__2; -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1665_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__2(lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1664_(lean_object*); lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonWaitForDiagnosticsParams____x40_Lean_Data_Lsp_Extra___hyg_24_(lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___closed__3; static lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addCommandRange___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDefinition___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__11___closed__1; +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__33(lean_object*, lean_object*, lean_object*); static lean_object* l_List_mapM___at_Lean_Server_FileWorker_getInteractiveGoals___spec__3___closed__6; lean_object* l_Lean_Elab_InfoTree_goalsAt_x3f(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_InfoTree_findInfo_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addCommandRange___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Server_Snapshots_parseAhead(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handlePlainTermGoal___closed__1; -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__28(lean_object*); +uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handlePlainGoal___spec__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_String_Range_contains(lean_object*, lean_object*, uint8_t); lean_object* l_Lean_Json_compress(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__24___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__3___closed__1; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__4___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__3; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__24___lambda__1(lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___closed__1; static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__2; static lean_object* l_Lean_Server_FileWorker_handleWaitForDiagnostics___closed__1; lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1019_(lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__3___closed__2; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__1(lean_object*); lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonDocumentSymbol_go___spec__3(size_t, size_t, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Lsp_SemanticTokenType_toNat(uint8_t); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__4___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleFoldingRange_isImport___closed__6; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_getInteractiveTermGoal___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__12___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__7; -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__3___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___closed__2; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__2___closed__2; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__13(size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10(lean_object*, lean_object*, lean_object*); lean_object* l_String_Range_toLspRange(lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletion___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addRanges(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___closed__10; -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__6___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__2(lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__24; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27(lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3698_(lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__4___closed__1; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addCommandRange___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_toString(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handlePlainGoal___spec__2(size_t, size_t, lean_object*); +static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_sectionLikeToDocumentSymbols___closed__3; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Server_FileWorker_handleFoldingRange___rarg___lambda__1(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__2___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__10; lean_object* l_Std_PersistentHashMap_insert___at_Lean_Server_registerLspRequestHandler___spec__2(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Widget_InteractiveTermGoal_toInteractiveGoal(lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3739_(lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__20; +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738_(lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__26___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_IO_AsyncList_waitAll___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__2; lean_object* l_Lean_Server_documentUriFromModule(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokens_highlightKeyword(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDefinition___closed__1; @@ -265,48 +274,45 @@ LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokens_highlight LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleSemanticTokensFull___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__3(lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3560_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__1(lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3559_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokensFull(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokens_highlightId___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__2(lean_object*); lean_object* lean_task_map(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handlePlainGoal___spec__2___closed__2; static lean_object* l_List_mapM___at_Lean_Server_FileWorker_getInteractiveGoals___spec__3___closed__3; -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__23(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_handleWaitForDiagnostics___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__3___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__1(lean_object*); lean_object* l_Lean_Syntax_getId(lean_object*); LEAN_EXPORT lean_object* l_List_mapTRAux___at_Lean_Server_FileWorker_handleDocumentHighlight___spec__2___boxed(lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Server_requestHandlers; lean_object* lean_format_pretty(lean_object*, lean_object*); lean_object* l_Lean_Server_RequestM_bindTask___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_choiceKind; static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__16___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Std_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__16; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_getInteractiveGoals(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__15; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokens_highlightId(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Std_PersistentHashMap_contains___at_Lean_Server_registerLspRequestHandler___spec__6(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__26___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__1(lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__12; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__2___closed__2; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handlePlainGoal___spec__2___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__23___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3477_(lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3476_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handlePlainGoal(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__3; lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_RefInfo_instCoeRefInfoRefInfo___spec__1(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDefinition___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__1(lean_object*); lean_object* l_Lean_Server_RequestM_mapTask___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__36(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_getInteractiveGoals___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint32_t lean_string_utf8_get(lean_object*, lean_object*); @@ -315,191 +321,190 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_hand LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__3___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handlePlainGoal___spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__30(size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletion(lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonPlainGoal____x40_Lean_Data_Lsp_Extra___hyg_654_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2648_(lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2647_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_getInteractiveTermGoal___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2600_(lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2599_(lean_object*); lean_object* lean_local_ctx_pop(lean_object*); static lean_object* l_Lean_Server_FileWorker_handleCompletion___closed__1; +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__12; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol___rarg___lambda__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__30(size_t, size_t, lean_object*); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletion___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__2___closed__1; uint8_t l_Array_isEmpty___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokensFull___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange___boxed(lean_object*); extern lean_object* l_Lean_instInhabitedSyntax; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__2___closed__1; lean_object* l_Lean_Elab_Info_stx(lean_object*); lean_object* l_Lean_Elab_Info_toElabInfo_x3f(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__1(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18(lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_Server_Snapshots_instInhabitedSnapshot; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__19(lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__22(lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__16(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__29(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__20; +static lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__1; static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__9; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__4___closed__1; -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552_(lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__4___closed__1; +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551_(lean_object*); extern lean_object* l_Task_Priority_default; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__2(lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__3___closed__2; lean_object* l_Lean_Elab_Info_range_x3f(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1(lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleCompletion___closed__2; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__24___lambda__1(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__15; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__13___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__3; +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__6___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletion___lambda__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__6(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Server_FileWorker_handleCompletion___lambda__1(lean_object*, lean_object*); uint8_t l_Char_isAlpha(uint32_t); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_getInteractiveTermGoal___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__5; -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1501_(lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1500_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__11(lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__3___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__4; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addCommandRange(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__7; -uint8_t l_Lean_Syntax_isNodeOf(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_FileWorker_handleSemanticTokens_go___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__16; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__5(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapM___at_Lean_Server_FileWorker_getInteractiveGoals___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_panic___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__1(lean_object*); lean_object* l_List_drop___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__36___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokens_highlightId___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletion___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_isIdOrAtom_x3f(lean_object*); static lean_object* l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__6___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__32(lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__14; static lean_object* l_Lean_Server_FileWorker_handlePlainGoal___closed__3; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__1(lean_object*); lean_object* l_Lean_Syntax_getPos_x3f(lean_object*, uint8_t); lean_object* l_IO_AsyncList_updateFinishedPrefix___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleHover(lean_object*, lean_object*, lean_object*); uint8_t lean_uint32_dec_eq(uint32_t, uint32_t); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__9___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__5(lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__17(size_t, size_t, lean_object*); lean_object* l_String_intercalate(lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__3___closed__2; lean_object* l_List_redLength___rarg(lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__12; +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDefinition___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__20(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__10; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__2(lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__11; lean_object* l_Std_HashMap_ofList___at_Lean_Server_ModuleRefs_instCoeModuleRefsModuleRefs___spec__5(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__25(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Server_FileWorker_handleSemanticTokens_highlightId___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Server_findModuleRefs(lean_object*, lean_object*, uint8_t); -lean_object* l_Lean_Server_FileWorker_EditableDocument_allSnaps(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokensFull___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__11(lean_object*); lean_object* l_IO_AsyncList_finishedPrefix___rarg(lean_object*); static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__2___closed__4; lean_object* l_Lean_Lsp_ModuleRefs_findAt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDefinition(uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleHover___lambda__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__1(lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___closed__9; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokens___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Info_tailPos_x3f(lean_object*); lean_object* l_Lean_Syntax_getArgs(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__24(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__5(lean_object*); lean_object* l_Lean_Syntax_getRange_x3f(lean_object*, uint8_t); lean_object* l_Lean_Syntax_getKind(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__25(lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___closed__4; -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__29___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_sectionLikeToDocumentSymbols(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__16; -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__10; +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__3___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_withPPInaccessibleNamesImp___rarg(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__26; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Server_RequestError_fileChanged; lean_object* l_Lean_Elab_Info_fmtHover_x3f(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Json_mkObj(lean_object*); lean_object* l_Lean_Elab_InfoTree_hoverableInfoAt_x3f(lean_object*, lean_object*, uint8_t, uint8_t); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__4___closed__1; lean_object* lean_nat_shiftl(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__8(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__1(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokens_highlightKeyword___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__1(lean_object*); LEAN_EXPORT uint8_t l_Array_contains___at_Lean_Server_FileWorker_handleSemanticTokens_go___spec__1(lean_object*, lean_object*); static lean_object* l_List_mapM___at_Lean_Server_FileWorker_getInteractiveGoals___spec__3___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addRanges___closed__1; +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__23___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Server_Snapshots_Snapshot_endPos(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokens_addToken(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Server_FileWorker_handleDocumentSymbol___spec__2(lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleSemanticTokens_go___closed__2; -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__11; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__1(lean_object*); static lean_object* l_Lean_Server_FileWorker_handlePlainGoal___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addRanges___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__2; static lean_object* l_Lean_Server_FileWorker_handleHover___lambda__2___closed__2; lean_object* l_Lean_Expr_constName_x3f(lean_object*); lean_object* l_List_toArrayAux___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isNone(lean_object*); lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol(lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarImpl(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__9(lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addRanges___closed__2; lean_object* l_Option_map___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l___private_Lean_Server_GoTo_0__Lean_Server_beqGoToKind____x40_Lean_Server_GoTo___hyg_8_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange___rarg___lambda__1___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__1(lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_getInteractiveGoals___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Server_FileWorker_handleSemanticTokens_highlightId___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___closed__7; lean_object* l_Std_HashMap_toList___at_Lean_Server_ModuleRefs_instCoeModuleRefsModuleRefs___spec__1(lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__33___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__9; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addRange(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_noHighlightKinds; lean_object* l_Lean_Server_RequestM_withWaitFindSnap___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__17___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_List_mapM___at_Lean_Server_FileWorker_getInteractiveGoals___spec__3___closed__4; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleHover___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__13___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__2(lean_object*); lean_object* l_Lean_Meta_instantiateMVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__2(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__24___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__3___closed__2; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__3___closed__1; +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__9___boxed(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__15(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__13; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getTailPos_x3f(lean_object*, uint8_t); static lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__6; static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__6; -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__33(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__12(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapM___at_Lean_Server_FileWorker_getInteractiveGoals___spec__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_contains___at_Lean_Server_FileWorker_handleSemanticTokens_go___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleHover___lambda__2___closed__1; @@ -507,124 +512,121 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Server_FileWorker_hand lean_object* lean_io_initializing(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleWaitForDiagnostics___lambda__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__9; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_getInteractiveGoals___lambda__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__20___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Server_locationLinksFromDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__5; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokens(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleWaitForDiagnostics(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__3___closed__1; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__3___closed__1; lean_object* lean_string_length(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handlePlainTermGoal___lambda__1___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__32(lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__20___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_getLast_x21___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Server_FileWorker_handleCompletion___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__35(lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__33___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Server_FileWorker_handleHover___lambda__1(lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handlePlainGoal___spec__1(size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__2(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Server_FileWorker_handleSemanticTokens_highlightId___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__26; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__6; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handlePlainTermGoal___lambda__1(lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__19; LEAN_EXPORT uint8_t l_Lean_Server_FileWorker_handleFoldingRange_isImport(lean_object*); lean_object* l_Lean_Server_RequestM_asTask___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleSemanticTokens___lambda__2___closed__1; lean_object* l_Lean_Elab_Info_pos_x3f(lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__3(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_FileWorker_handleDocumentHighlight___spec__3(lean_object*, lean_object*, size_t, size_t, lean_object*); -static lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__1; LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Server_FileWorker_handleCompletion___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__2(lean_object*); LEAN_EXPORT lean_object* l_List_spanAux___at_Lean_Server_FileWorker_handleFoldingRange_addRanges___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__26(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleHover___lambda__2___closed__4; -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__9(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Server_FileWorker_handleSemanticTokens___lambda__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__17(size_t, size_t, lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3857_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3856_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__12___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__8; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__9; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__4___closed__1; -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__6(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleWaitForDiagnostics_waitLoop___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_local_ctx_find(lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__4; +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__21; static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___closed__8; static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___closed__2; lean_object* l_Lean_Elab_Info_lctx(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handlePlainGoal___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Server_FileWorker_handleSemanticTokens___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_handleSemanticTokens_go___closed__4; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___closed__11; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_FileWorker_handleSemanticTokens_go___spec__3(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handlePlainGoal___lambda__1___closed__3; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__2(lean_object*); lean_object* l_Lean_Elab_ContextInfo_runMetaM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokens_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__8(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__28(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_getInteractiveTermGoal(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__2; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__4(uint8_t, 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_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__35(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__30___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__3___closed__2; static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__5; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__2; +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__29(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleSemanticTokens_go___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_List_mapTRAux___at_Lean_Server_FileWorker_handleDocumentHighlight___spec__2(size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTRAux___at_Lean_Server_FileWorker_handleDocumentSymbol___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___closed__1; static lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__1; static lean_object* l_List_forIn_loop___at_Lean_Server_FileWorker_handleSemanticTokens___spec__1___lambda__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__2(lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__23(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleHover___lambda__2___closed__3; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__3___closed__1; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addRangeFromSyntax___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withPPInaccessibleNames___at_Lean_Server_FileWorker_getInteractiveGoals___spec__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__2(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Server_FileWorker_handleSemanticTokens___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__16(lean_object*, lean_object*, lean_object*); lean_object* l_List_appendTR___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__1; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_FileWorker_handleDocumentHighlight___spec__3___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleWaitForDiagnostics___lambda__2(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__17___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_reprint(lean_object*); static lean_object* l_List_mapM___at_Lean_Server_FileWorker_getInteractiveGoals___spec__3___closed__5; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__13(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Server_RequestM_readDoc(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_getLastD___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__4; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__2; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__24___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__19; -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__16___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__36(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__8; lean_object* l_Lean_Widget_goalToInteractive(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__17; uint8_t lean_string_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___lambda__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__24___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_panic___at_Lean_Elab_InfoTree_smallestInfo_x3f___spec__1(lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_FileWorker_handleDocumentHighlight___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2462_(lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2461_(lean_object*); static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__2___closed__3; static lean_object* l_Lean_Server_FileWorker_handlePlainGoal___lambda__1___closed__1; LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Server_FileWorker_handleCompletion___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { @@ -839,7 +841,6 @@ lean_closure_set(x_16, 1, x_10); lean_closure_set(x_16, 2, x_12); lean_closure_set(x_16, 3, x_14); x_17 = l_Lean_Server_RequestM_withWaitFindSnap___rarg(x_5, x_13, x_15, x_16, x_2, x_6); -lean_dec(x_5); return x_17; } } @@ -1903,7 +1904,6 @@ lean_closure_set(x_13, 1, x_10); lean_closure_set(x_13, 2, x_8); x_14 = l_Lean_Server_FileWorker_handleHover___closed__1; x_15 = l_Lean_Server_RequestM_withWaitFindSnap___rarg(x_5, x_11, x_14, x_13, x_2, x_6); -lean_dec(x_5); return x_15; } } @@ -4528,7 +4528,6 @@ lean_closure_set(x_15, 1, x_13); lean_closure_set(x_15, 2, x_14); x_16 = l_Lean_Server_FileWorker_handleDefinition___closed__1; x_17 = l_Lean_Server_RequestM_withWaitFindSnap___rarg(x_6, x_12, x_16, x_15, x_3, x_7); -lean_dec(x_6); return x_17; } } @@ -5884,7 +5883,6 @@ lean_closure_set(x_13, 1, x_10); lean_closure_set(x_13, 2, x_12); x_14 = l_Lean_Server_FileWorker_handleHover___closed__1; x_15 = l_Lean_Server_RequestM_withWaitFindSnap___rarg(x_5, x_11, x_14, x_13, x_2, x_6); -lean_dec(x_5); return x_15; } } @@ -7094,7 +7092,6 @@ lean_closure_set(x_13, 2, x_9); lean_closure_set(x_13, 3, x_8); x_14 = l_Lean_Server_FileWorker_handleHover___closed__1; x_15 = l_Lean_Server_RequestM_withWaitFindSnap___rarg(x_5, x_11, x_14, x_13, x_2, x_6); -lean_dec(x_5); return x_15; } } @@ -7558,187 +7555,186 @@ goto _start; } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; x_39 = lean_unsigned_to_nat(0u); x_40 = l_Lean_Syntax_getArg(x_4, x_39); x_41 = lean_unsigned_to_nat(1u); x_42 = l_Lean_Syntax_getArg(x_4, x_41); -x_43 = l_Lean_nullKind; lean_inc(x_42); -x_44 = l_Lean_Syntax_isNodeOf(x_42, x_43, x_41); -if (x_44 == 0) +x_43 = l_Lean_Syntax_matchesNull(x_42, x_41); +if (x_43 == 0) { -lean_object* x_45; lean_object* x_46; lean_object* x_47; size_t x_48; size_t x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +lean_object* x_44; lean_object* x_45; lean_object* x_46; size_t x_47; size_t x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_dec(x_42); lean_dec(x_40); -x_45 = l_Lean_Syntax_getArgs(x_4); +x_44 = l_Lean_Syntax_getArgs(x_4); lean_dec(x_4); -x_46 = lean_box(0); -x_47 = lean_array_get_size(x_45); -x_48 = lean_usize_of_nat(x_47); -lean_dec(x_47); -x_49 = 0; -x_50 = l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__11; -x_51 = l_Array_forInUnsafe_loop___at_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___spec__2(x_1, x_2, x_3, x_50, x_45, x_48, x_49, x_50); -lean_dec(x_45); -x_52 = lean_ctor_get(x_51, 0); -lean_inc(x_52); -lean_dec(x_51); -if (lean_obj_tag(x_52) == 0) +x_45 = lean_box(0); +x_46 = lean_array_get_size(x_44); +x_47 = lean_usize_of_nat(x_46); +lean_dec(x_46); +x_48 = 0; +x_49 = l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__11; +x_50 = l_Array_forInUnsafe_loop___at_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___spec__2(x_1, x_2, x_3, x_49, x_44, x_47, x_48, x_49); +lean_dec(x_44); +x_51 = lean_ctor_get(x_50, 0); +lean_inc(x_51); +lean_dec(x_50); +if (lean_obj_tag(x_51) == 0) { -return x_46; +return x_45; } else { -uint8_t x_53; -x_53 = !lean_is_exclusive(x_52); -if (x_53 == 0) +uint8_t x_52; +x_52 = !lean_is_exclusive(x_51); +if (x_52 == 0) { -return x_52; +return x_51; } else { -lean_object* x_54; lean_object* x_55; -x_54 = lean_ctor_get(x_52, 0); -lean_inc(x_54); -lean_dec(x_52); -x_55 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_55, 0, x_54); -return x_55; +lean_object* x_53; lean_object* x_54; +x_53 = lean_ctor_get(x_51, 0); +lean_inc(x_53); +lean_dec(x_51); +x_54 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_54, 0, x_53); +return x_54; } } } else { -lean_object* x_56; uint8_t x_57; lean_object* x_58; +lean_object* x_55; uint8_t x_56; lean_object* x_57; lean_dec(x_4); -x_56 = l_Lean_Syntax_getArg(x_42, x_39); +x_55 = l_Lean_Syntax_getArg(x_42, x_39); lean_dec(x_42); -x_57 = 0; -x_58 = l_Lean_Syntax_getRange_x3f(x_40, x_57); -if (lean_obj_tag(x_58) == 0) +x_56 = 0; +x_57 = l_Lean_Syntax_getRange_x3f(x_40, x_56); +if (lean_obj_tag(x_57) == 0) { -x_4 = x_56; +x_4 = x_55; goto _start; } else { -uint8_t x_60; -x_60 = !lean_is_exclusive(x_58); -if (x_60 == 0) +uint8_t x_59; +x_59 = !lean_is_exclusive(x_57); +if (x_59 == 0) { -lean_object* x_61; uint8_t x_62; -x_61 = lean_ctor_get(x_58, 0); -x_62 = l_String_Range_contains(x_61, x_2, x_57); -if (x_62 == 0) +lean_object* x_60; uint8_t x_61; +x_60 = lean_ctor_get(x_57, 0); +x_61 = l_String_Range_contains(x_60, x_2, x_56); +if (x_61 == 0) { -lean_free_object(x_58); -lean_dec(x_61); -x_4 = x_56; +lean_free_object(x_57); +lean_dec(x_60); +x_4 = x_55; goto _start; } else { -lean_dec(x_56); +lean_dec(x_55); if (lean_obj_tag(x_3) == 0) { -lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = l_String_Range_toLspRange(x_1, x_61); -lean_dec(x_61); -x_65 = l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__12; -x_66 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_66, 0, x_64); -lean_ctor_set(x_66, 1, x_65); -lean_ctor_set(x_58, 0, x_66); -return x_58; +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = l_String_Range_toLspRange(x_1, x_60); +lean_dec(x_60); +x_64 = l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__12; +x_65 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_64); +lean_ctor_set(x_57, 0, x_65); +return x_57; } else { -uint8_t x_67; -lean_free_object(x_58); -lean_dec(x_61); -x_67 = !lean_is_exclusive(x_3); -if (x_67 == 0) +uint8_t x_66; +lean_free_object(x_57); +lean_dec(x_60); +x_66 = !lean_is_exclusive(x_3); +if (x_66 == 0) { -lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_68 = lean_ctor_get(x_3, 0); -x_69 = l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__12; -x_70 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_69); -lean_ctor_set(x_3, 0, x_70); +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_3, 0); +x_68 = l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__12; +x_69 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_69, 0, x_67); +lean_ctor_set(x_69, 1, x_68); +lean_ctor_set(x_3, 0, x_69); return x_3; } else { -lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_71 = lean_ctor_get(x_3, 0); -lean_inc(x_71); +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_70 = lean_ctor_get(x_3, 0); +lean_inc(x_70); lean_dec(x_3); -x_72 = l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__12; -x_73 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_73, 0, x_71); -lean_ctor_set(x_73, 1, x_72); -x_74 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_74, 0, x_73); -return x_74; +x_71 = l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__12; +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_70); +lean_ctor_set(x_72, 1, x_71); +x_73 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_73, 0, x_72); +return x_73; } } } } else { -lean_object* x_75; uint8_t x_76; -x_75 = lean_ctor_get(x_58, 0); -lean_inc(x_75); -lean_dec(x_58); -x_76 = l_String_Range_contains(x_75, x_2, x_57); -if (x_76 == 0) +lean_object* x_74; uint8_t x_75; +x_74 = lean_ctor_get(x_57, 0); +lean_inc(x_74); +lean_dec(x_57); +x_75 = l_String_Range_contains(x_74, x_2, x_56); +if (x_75 == 0) { -lean_dec(x_75); -x_4 = x_56; +lean_dec(x_74); +x_4 = x_55; goto _start; } else { -lean_dec(x_56); +lean_dec(x_55); if (lean_obj_tag(x_3) == 0) { -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_78 = l_String_Range_toLspRange(x_1, x_75); -lean_dec(x_75); -x_79 = l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__12; -x_80 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_80, 0, x_78); -lean_ctor_set(x_80, 1, x_79); -x_81 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_81, 0, x_80); -return x_81; +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_77 = l_String_Range_toLspRange(x_1, x_74); +lean_dec(x_74); +x_78 = l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__12; +x_79 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_79, 0, x_77); +lean_ctor_set(x_79, 1, x_78); +x_80 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_80, 0, x_79); +return x_80; } else { -lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; -lean_dec(x_75); -x_82 = lean_ctor_get(x_3, 0); -lean_inc(x_82); +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +lean_dec(x_74); +x_81 = lean_ctor_get(x_3, 0); +lean_inc(x_81); if (lean_is_exclusive(x_3)) { lean_ctor_release(x_3, 0); - x_83 = x_3; + x_82 = x_3; } else { lean_dec_ref(x_3); - x_83 = lean_box(0); + x_82 = lean_box(0); } -x_84 = l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__12; -x_85 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_85, 0, x_82); -lean_ctor_set(x_85, 1, x_84); -if (lean_is_scalar(x_83)) { - x_86 = lean_alloc_ctor(1, 1, 0); +x_83 = l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__12; +x_84 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_84, 0, x_81); +lean_ctor_set(x_84, 1, x_83); +if (lean_is_scalar(x_82)) { + x_85 = lean_alloc_ctor(1, 1, 0); } else { - x_86 = x_83; + x_85 = x_82; } -lean_ctor_set(x_86, 0, x_85); -return x_86; +lean_ctor_set(x_85, 0, x_84); +return x_85; } } } @@ -8265,7 +8261,8 @@ LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight___lamb _start: { lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_9 = l_Lean_Server_FileWorker_EditableDocument_allSnaps(x_1); +x_9 = lean_ctor_get(x_1, 1); +lean_inc(x_9); lean_dec(x_1); x_10 = l_IO_AsyncList_updateFinishedPrefix___rarg(x_9, x_8); x_11 = !lean_is_exclusive(x_10); @@ -8434,7 +8431,6 @@ lean_closure_set(x_13, 3, x_10); lean_closure_set(x_13, 4, x_9); x_14 = l_Lean_Server_FileWorker_handleDefinition___closed__1; x_15 = l_Lean_Server_RequestM_withWaitFindSnap___rarg(x_5, x_11, x_14, x_13, x_2, x_6); -lean_dec(x_5); return x_15; } } @@ -8727,375 +8723,373 @@ lean_dec(x_14); x_36 = l_Lean_Syntax_isNone(x_35); if (x_36 == 0) { -lean_object* x_37; lean_object* x_38; uint8_t x_39; -x_37 = l_Lean_nullKind; -x_38 = lean_unsigned_to_nat(3u); +lean_object* x_37; uint8_t x_38; +x_37 = lean_unsigned_to_nat(3u); lean_inc(x_35); -x_39 = l_Lean_Syntax_isNodeOf(x_35, x_37, x_38); -if (x_39 == 0) +x_38 = l_Lean_Syntax_matchesNull(x_35, x_37); +if (x_38 == 0) { -lean_object* x_40; lean_object* x_41; +lean_object* x_39; lean_object* x_40; lean_dec(x_35); lean_dec(x_26); -x_40 = l_Lean_Syntax_getArg(x_13, x_25); +x_39 = l_Lean_Syntax_getArg(x_13, x_25); lean_dec(x_13); -x_41 = l_Lean_Syntax_isIdOrAtom_x3f(x_40); -if (lean_obj_tag(x_41) == 0) +x_40 = l_Lean_Syntax_isIdOrAtom_x3f(x_39); +if (lean_obj_tag(x_40) == 0) { -lean_object* x_42; lean_object* x_43; -x_42 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_43, 1, x_40); -return x_43; +lean_object* x_41; lean_object* x_42; +x_41 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_39); +return x_42; } else { -lean_object* x_44; lean_object* x_45; -x_44 = lean_ctor_get(x_41, 0); -lean_inc(x_44); -lean_dec(x_41); -x_45 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_45, 0, x_44); -lean_ctor_set(x_45, 1, x_40); -return x_45; +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_40, 0); +lean_inc(x_43); +lean_dec(x_40); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_39); +return x_44; } } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_dec(x_13); -x_46 = l_Lean_Syntax_getArg(x_35, x_12); +x_45 = l_Lean_Syntax_getArg(x_35, x_12); lean_dec(x_35); -x_47 = l_Lean_Syntax_getArgs(x_46); -lean_dec(x_46); -x_48 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_48, 0, x_47); -x_49 = lean_box(0); -x_50 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_26, x_49, x_48); -lean_dec(x_48); -return x_50; +x_46 = l_Lean_Syntax_getArgs(x_45); +lean_dec(x_45); +x_47 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_47, 0, x_46); +x_48 = lean_box(0); +x_49 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_26, x_48, x_47); +lean_dec(x_47); +return x_49; } } else { -lean_object* x_51; lean_object* x_52; lean_object* x_53; +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_dec(x_35); lean_dec(x_13); +x_50 = lean_box(0); x_51 = lean_box(0); -x_52 = lean_box(0); -x_53 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_26, x_52, x_51); -return x_53; +x_52 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_26, x_51, x_50); +return x_52; } } } } else { -lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; uint8_t x_58; -x_54 = lean_unsigned_to_nat(6u); -x_55 = l_Lean_Syntax_getArg(x_1, x_54); -x_56 = l_Lean_nullKind; -x_57 = lean_unsigned_to_nat(0u); -x_58 = l_Lean_Syntax_isNodeOf(x_55, x_56, x_57); -if (x_58 == 0) +lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; +x_53 = lean_unsigned_to_nat(6u); +x_54 = l_Lean_Syntax_getArg(x_1, x_53); +x_55 = lean_unsigned_to_nat(0u); +x_56 = l_Lean_Syntax_matchesNull(x_54, x_55); +if (x_56 == 0) { -lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; lean_dec(x_8); lean_dec(x_1); -x_59 = lean_unsigned_to_nat(1u); -x_60 = l_Lean_Syntax_getArg(x_3, x_59); +x_57 = lean_unsigned_to_nat(1u); +x_58 = l_Lean_Syntax_getArg(x_3, x_57); lean_dec(x_3); -x_61 = l_Lean_Syntax_getArg(x_60, x_59); -x_62 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__2; -x_63 = lean_name_mk_string(x_2, x_62); -lean_inc(x_61); -x_64 = l_Lean_Syntax_isOfKind(x_61, x_63); -lean_dec(x_63); -if (x_64 == 0) +x_59 = l_Lean_Syntax_getArg(x_58, x_57); +x_60 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__2; +x_61 = lean_name_mk_string(x_2, x_60); +lean_inc(x_59); +x_62 = l_Lean_Syntax_isOfKind(x_59, x_61); +lean_dec(x_61); +if (x_62 == 0) +{ +lean_object* x_63; lean_object* x_64; +lean_dec(x_59); +x_63 = l_Lean_Syntax_getArg(x_58, x_55); +lean_dec(x_58); +x_64 = l_Lean_Syntax_isIdOrAtom_x3f(x_63); +if (lean_obj_tag(x_64) == 0) { lean_object* x_65; lean_object* x_66; -lean_dec(x_61); -x_65 = l_Lean_Syntax_getArg(x_60, x_57); -lean_dec(x_60); -x_66 = l_Lean_Syntax_isIdOrAtom_x3f(x_65); -if (lean_obj_tag(x_66) == 0) +x_65 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_66 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_66, 0, x_65); +lean_ctor_set(x_66, 1, x_63); +return x_66; +} +else { lean_object* x_67; lean_object* x_68; -x_67 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_67 = lean_ctor_get(x_64, 0); +lean_inc(x_67); +lean_dec(x_64); x_68 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_68, 0, x_67); -lean_ctor_set(x_68, 1, x_65); +lean_ctor_set(x_68, 1, x_63); return x_68; } +} else { -lean_object* x_69; lean_object* x_70; -x_69 = lean_ctor_get(x_66, 0); +lean_object* x_69; lean_object* x_70; uint8_t x_71; +x_69 = l_Lean_Syntax_getArg(x_59, x_55); +x_70 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; lean_inc(x_69); -lean_dec(x_66); -x_70 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_65); -return x_70; -} -} -else +x_71 = l_Lean_Syntax_isOfKind(x_69, x_70); +if (x_71 == 0) { -lean_object* x_71; lean_object* x_72; uint8_t x_73; -x_71 = l_Lean_Syntax_getArg(x_61, x_57); -x_72 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; -lean_inc(x_71); -x_73 = l_Lean_Syntax_isOfKind(x_71, x_72); -if (x_73 == 0) +lean_object* x_72; lean_object* x_73; +lean_dec(x_69); +lean_dec(x_59); +x_72 = l_Lean_Syntax_getArg(x_58, x_55); +lean_dec(x_58); +x_73 = l_Lean_Syntax_isIdOrAtom_x3f(x_72); +if (lean_obj_tag(x_73) == 0) { lean_object* x_74; lean_object* x_75; -lean_dec(x_71); -lean_dec(x_61); -x_74 = l_Lean_Syntax_getArg(x_60, x_57); -lean_dec(x_60); -x_75 = l_Lean_Syntax_isIdOrAtom_x3f(x_74); -if (lean_obj_tag(x_75) == 0) +x_74 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_75 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_75, 0, x_74); +lean_ctor_set(x_75, 1, x_72); +return x_75; +} +else { lean_object* x_76; lean_object* x_77; -x_76 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_76 = lean_ctor_get(x_73, 0); +lean_inc(x_76); +lean_dec(x_73); x_77 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_77, 0, x_76); -lean_ctor_set(x_77, 1, x_74); +lean_ctor_set(x_77, 1, x_72); return x_77; } -else -{ -lean_object* x_78; lean_object* x_79; -x_78 = lean_ctor_get(x_75, 0); -lean_inc(x_78); -lean_dec(x_75); -x_79 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_79, 0, x_78); -lean_ctor_set(x_79, 1, x_74); -return x_79; -} } else { +lean_object* x_78; uint8_t x_79; +x_78 = l_Lean_Syntax_getArg(x_59, x_57); +lean_dec(x_59); +x_79 = l_Lean_Syntax_isNone(x_78); +if (x_79 == 0) +{ lean_object* x_80; uint8_t x_81; -x_80 = l_Lean_Syntax_getArg(x_61, x_59); -lean_dec(x_61); -x_81 = l_Lean_Syntax_isNone(x_80); +x_80 = lean_unsigned_to_nat(3u); +lean_inc(x_78); +x_81 = l_Lean_Syntax_matchesNull(x_78, x_80); if (x_81 == 0) { -lean_object* x_82; uint8_t x_83; -x_82 = lean_unsigned_to_nat(3u); -lean_inc(x_80); -x_83 = l_Lean_Syntax_isNodeOf(x_80, x_56, x_82); -if (x_83 == 0) +lean_object* x_82; lean_object* x_83; +lean_dec(x_78); +lean_dec(x_69); +x_82 = l_Lean_Syntax_getArg(x_58, x_55); +lean_dec(x_58); +x_83 = l_Lean_Syntax_isIdOrAtom_x3f(x_82); +if (lean_obj_tag(x_83) == 0) { lean_object* x_84; lean_object* x_85; -lean_dec(x_80); -lean_dec(x_71); -x_84 = l_Lean_Syntax_getArg(x_60, x_57); -lean_dec(x_60); -x_85 = l_Lean_Syntax_isIdOrAtom_x3f(x_84); -if (lean_obj_tag(x_85) == 0) +x_84 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_85 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_85, 0, x_84); +lean_ctor_set(x_85, 1, x_82); +return x_85; +} +else { lean_object* x_86; lean_object* x_87; -x_86 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_86 = lean_ctor_get(x_83, 0); +lean_inc(x_86); +lean_dec(x_83); x_87 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_87, 0, x_86); -lean_ctor_set(x_87, 1, x_84); +lean_ctor_set(x_87, 1, x_82); return x_87; } -else -{ -lean_object* x_88; lean_object* x_89; -x_88 = lean_ctor_get(x_85, 0); -lean_inc(x_88); -lean_dec(x_85); -x_89 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_89, 0, x_88); -lean_ctor_set(x_89, 1, x_84); -return x_89; -} } else { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; -lean_dec(x_60); -x_90 = l_Lean_Syntax_getArg(x_80, x_59); -lean_dec(x_80); -x_91 = l_Lean_Syntax_getArgs(x_90); +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; +lean_dec(x_58); +x_88 = l_Lean_Syntax_getArg(x_78, x_57); +lean_dec(x_78); +x_89 = l_Lean_Syntax_getArgs(x_88); +lean_dec(x_88); +x_90 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_90, 0, x_89); +x_91 = lean_box(0); +x_92 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_69, x_91, x_90); lean_dec(x_90); -x_92 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_92, 0, x_91); -x_93 = lean_box(0); -x_94 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_71, x_93, x_92); -lean_dec(x_92); -return x_94; +return x_92; } } else { -lean_object* x_95; lean_object* x_96; lean_object* x_97; -lean_dec(x_80); -lean_dec(x_60); -x_95 = lean_box(0); -x_96 = lean_box(0); -x_97 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_71, x_96, x_95); -return x_97; +lean_object* x_93; lean_object* x_94; lean_object* x_95; +lean_dec(x_78); +lean_dec(x_58); +x_93 = lean_box(0); +x_94 = lean_box(0); +x_95 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_69, x_94, x_93); +return x_95; } } } } else { -lean_object* x_98; lean_object* x_99; uint8_t x_100; -x_98 = lean_unsigned_to_nat(7u); -x_99 = l_Lean_Syntax_getArg(x_1, x_98); +lean_object* x_96; lean_object* x_97; uint8_t x_98; +x_96 = lean_unsigned_to_nat(7u); +x_97 = l_Lean_Syntax_getArg(x_1, x_96); lean_dec(x_1); -x_100 = l_Lean_Syntax_isNodeOf(x_99, x_56, x_57); -if (x_100 == 0) +x_98 = l_Lean_Syntax_matchesNull(x_97, x_55); +if (x_98 == 0) { -lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; uint8_t x_106; +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; uint8_t x_104; lean_dec(x_8); -x_101 = lean_unsigned_to_nat(1u); -x_102 = l_Lean_Syntax_getArg(x_3, x_101); +x_99 = lean_unsigned_to_nat(1u); +x_100 = l_Lean_Syntax_getArg(x_3, x_99); lean_dec(x_3); -x_103 = l_Lean_Syntax_getArg(x_102, x_101); -x_104 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__2; -x_105 = lean_name_mk_string(x_2, x_104); -lean_inc(x_103); -x_106 = l_Lean_Syntax_isOfKind(x_103, x_105); -lean_dec(x_105); -if (x_106 == 0) +x_101 = l_Lean_Syntax_getArg(x_100, x_99); +x_102 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__2; +x_103 = lean_name_mk_string(x_2, x_102); +lean_inc(x_101); +x_104 = l_Lean_Syntax_isOfKind(x_101, x_103); +lean_dec(x_103); +if (x_104 == 0) +{ +lean_object* x_105; lean_object* x_106; +lean_dec(x_101); +x_105 = l_Lean_Syntax_getArg(x_100, x_55); +lean_dec(x_100); +x_106 = l_Lean_Syntax_isIdOrAtom_x3f(x_105); +if (lean_obj_tag(x_106) == 0) { lean_object* x_107; lean_object* x_108; -lean_dec(x_103); -x_107 = l_Lean_Syntax_getArg(x_102, x_57); -lean_dec(x_102); -x_108 = l_Lean_Syntax_isIdOrAtom_x3f(x_107); -if (lean_obj_tag(x_108) == 0) +x_107 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_108 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_108, 0, x_107); +lean_ctor_set(x_108, 1, x_105); +return x_108; +} +else { lean_object* x_109; lean_object* x_110; -x_109 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_109 = lean_ctor_get(x_106, 0); +lean_inc(x_109); +lean_dec(x_106); x_110 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_110, 0, x_109); -lean_ctor_set(x_110, 1, x_107); +lean_ctor_set(x_110, 1, x_105); return x_110; } +} else { -lean_object* x_111; lean_object* x_112; -x_111 = lean_ctor_get(x_108, 0); +lean_object* x_111; lean_object* x_112; uint8_t x_113; +x_111 = l_Lean_Syntax_getArg(x_101, x_55); +x_112 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; lean_inc(x_111); -lean_dec(x_108); -x_112 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_112, 0, x_111); -lean_ctor_set(x_112, 1, x_107); -return x_112; -} -} -else +x_113 = l_Lean_Syntax_isOfKind(x_111, x_112); +if (x_113 == 0) { -lean_object* x_113; lean_object* x_114; uint8_t x_115; -x_113 = l_Lean_Syntax_getArg(x_103, x_57); -x_114 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; -lean_inc(x_113); -x_115 = l_Lean_Syntax_isOfKind(x_113, x_114); -if (x_115 == 0) +lean_object* x_114; lean_object* x_115; +lean_dec(x_111); +lean_dec(x_101); +x_114 = l_Lean_Syntax_getArg(x_100, x_55); +lean_dec(x_100); +x_115 = l_Lean_Syntax_isIdOrAtom_x3f(x_114); +if (lean_obj_tag(x_115) == 0) { lean_object* x_116; lean_object* x_117; -lean_dec(x_113); -lean_dec(x_103); -x_116 = l_Lean_Syntax_getArg(x_102, x_57); -lean_dec(x_102); -x_117 = l_Lean_Syntax_isIdOrAtom_x3f(x_116); -if (lean_obj_tag(x_117) == 0) +x_116 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_117 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_117, 0, x_116); +lean_ctor_set(x_117, 1, x_114); +return x_117; +} +else { lean_object* x_118; lean_object* x_119; -x_118 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_118 = lean_ctor_get(x_115, 0); +lean_inc(x_118); +lean_dec(x_115); x_119 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_119, 0, x_118); -lean_ctor_set(x_119, 1, x_116); +lean_ctor_set(x_119, 1, x_114); return x_119; } -else -{ -lean_object* x_120; lean_object* x_121; -x_120 = lean_ctor_get(x_117, 0); -lean_inc(x_120); -lean_dec(x_117); -x_121 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_121, 0, x_120); -lean_ctor_set(x_121, 1, x_116); -return x_121; -} } else { +lean_object* x_120; uint8_t x_121; +x_120 = l_Lean_Syntax_getArg(x_101, x_99); +lean_dec(x_101); +x_121 = l_Lean_Syntax_isNone(x_120); +if (x_121 == 0) +{ lean_object* x_122; uint8_t x_123; -x_122 = l_Lean_Syntax_getArg(x_103, x_101); -lean_dec(x_103); -x_123 = l_Lean_Syntax_isNone(x_122); +x_122 = lean_unsigned_to_nat(3u); +lean_inc(x_120); +x_123 = l_Lean_Syntax_matchesNull(x_120, x_122); if (x_123 == 0) { -lean_object* x_124; uint8_t x_125; -x_124 = lean_unsigned_to_nat(3u); -lean_inc(x_122); -x_125 = l_Lean_Syntax_isNodeOf(x_122, x_56, x_124); -if (x_125 == 0) +lean_object* x_124; lean_object* x_125; +lean_dec(x_120); +lean_dec(x_111); +x_124 = l_Lean_Syntax_getArg(x_100, x_55); +lean_dec(x_100); +x_125 = l_Lean_Syntax_isIdOrAtom_x3f(x_124); +if (lean_obj_tag(x_125) == 0) { lean_object* x_126; lean_object* x_127; -lean_dec(x_122); -lean_dec(x_113); -x_126 = l_Lean_Syntax_getArg(x_102, x_57); -lean_dec(x_102); -x_127 = l_Lean_Syntax_isIdOrAtom_x3f(x_126); -if (lean_obj_tag(x_127) == 0) +x_126 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_127 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_127, 0, x_126); +lean_ctor_set(x_127, 1, x_124); +return x_127; +} +else { lean_object* x_128; lean_object* x_129; -x_128 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_128 = lean_ctor_get(x_125, 0); +lean_inc(x_128); +lean_dec(x_125); x_129 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_129, 0, x_128); -lean_ctor_set(x_129, 1, x_126); +lean_ctor_set(x_129, 1, x_124); return x_129; } -else -{ -lean_object* x_130; lean_object* x_131; -x_130 = lean_ctor_get(x_127, 0); -lean_inc(x_130); -lean_dec(x_127); -x_131 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_131, 0, x_130); -lean_ctor_set(x_131, 1, x_126); -return x_131; -} } else { -lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; -lean_dec(x_102); -x_132 = l_Lean_Syntax_getArg(x_122, x_101); -lean_dec(x_122); -x_133 = l_Lean_Syntax_getArgs(x_132); +lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; +lean_dec(x_100); +x_130 = l_Lean_Syntax_getArg(x_120, x_99); +lean_dec(x_120); +x_131 = l_Lean_Syntax_getArgs(x_130); +lean_dec(x_130); +x_132 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_132, 0, x_131); +x_133 = lean_box(0); +x_134 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_111, x_133, x_132); lean_dec(x_132); -x_134 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_134, 0, x_133); -x_135 = lean_box(0); -x_136 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_113, x_135, x_134); -lean_dec(x_134); -return x_136; +return x_134; } } else { -lean_object* x_137; lean_object* x_138; lean_object* x_139; -lean_dec(x_122); -lean_dec(x_102); -x_137 = lean_box(0); -x_138 = lean_box(0); -x_139 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_113, x_138, x_137); -return x_139; +lean_object* x_135; lean_object* x_136; lean_object* x_137; +lean_dec(x_120); +lean_dec(x_100); +x_135 = lean_box(0); +x_136 = lean_box(0); +x_137 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_111, x_136, x_135); +return x_137; } } } @@ -9106,48 +9100,48 @@ lean_dec(x_3); lean_dec(x_2); if (lean_obj_tag(x_5) == 0) { -lean_object* x_140; +lean_object* x_138; lean_inc(x_8); -x_140 = l_Lean_Syntax_reprint(x_8); -if (lean_obj_tag(x_140) == 0) +x_138 = l_Lean_Syntax_reprint(x_8); +if (lean_obj_tag(x_138) == 0) { -lean_object* x_141; lean_object* x_142; -x_141 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__9; -x_142 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_142, 0, x_141); -lean_ctor_set(x_142, 1, x_8); -return x_142; +lean_object* x_139; lean_object* x_140; +x_139 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__9; +x_140 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_140, 0, x_139); +lean_ctor_set(x_140, 1, x_8); +return x_140; } else { -lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; -x_143 = lean_ctor_get(x_140, 0); -lean_inc(x_143); -lean_dec(x_140); -x_144 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__6; -x_145 = lean_string_append(x_144, x_143); -lean_dec(x_143); -x_146 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__7; -x_147 = lean_string_append(x_145, x_146); -x_148 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_148, 0, x_147); -lean_ctor_set(x_148, 1, x_8); -return x_148; +lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; +x_141 = lean_ctor_get(x_138, 0); +lean_inc(x_141); +lean_dec(x_138); +x_142 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__6; +x_143 = lean_string_append(x_142, x_141); +lean_dec(x_141); +x_144 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__7; +x_145 = lean_string_append(x_143, x_144); +x_146 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_146, 0, x_145); +lean_ctor_set(x_146, 1, x_8); +return x_146; } } else { -lean_object* x_149; lean_object* x_150; uint8_t x_151; lean_object* x_152; lean_object* x_153; +lean_object* x_147; lean_object* x_148; uint8_t x_149; lean_object* x_150; lean_object* x_151; lean_dec(x_8); -x_149 = lean_ctor_get(x_5, 0); -x_150 = l_Lean_Syntax_getId(x_149); -x_151 = 1; -x_152 = l_Lean_Name_toString(x_150, x_151); -lean_inc(x_149); -x_153 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_153, 0, x_152); -lean_ctor_set(x_153, 1, x_149); -return x_153; +x_147 = lean_ctor_get(x_5, 0); +x_148 = l_Lean_Syntax_getId(x_147); +x_149 = 1; +x_150 = l_Lean_Name_toString(x_148, x_149); +lean_inc(x_147); +x_151 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_151, 0, x_150); +lean_ctor_set(x_151, 1, x_147); +return x_151; } } } @@ -9185,646 +9179,645 @@ lean_closure_set(x_8, 2, x_3); x_9 = l_Lean_Syntax_isNone(x_7); if (x_9 == 0) { -lean_object* x_10; lean_object* x_11; uint8_t x_12; +lean_object* x_10; uint8_t x_11; lean_dec(x_1); -x_10 = l_Lean_nullKind; -x_11 = lean_unsigned_to_nat(1u); +x_10 = lean_unsigned_to_nat(1u); lean_inc(x_7); -x_12 = l_Lean_Syntax_isNodeOf(x_7, x_10, x_11); -if (x_12 == 0) +x_11 = l_Lean_Syntax_matchesNull(x_7, x_10); +if (x_11 == 0) { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_dec(x_8); lean_dec(x_7); -x_13 = l_Lean_Syntax_getArg(x_3, x_11); +x_12 = l_Lean_Syntax_getArg(x_3, x_10); lean_dec(x_3); -x_14 = l_Lean_Syntax_getArg(x_13, x_11); -x_15 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__2; -x_16 = lean_name_mk_string(x_2, x_15); -lean_inc(x_14); -x_17 = l_Lean_Syntax_isOfKind(x_14, x_16); -lean_dec(x_16); -if (x_17 == 0) +x_13 = l_Lean_Syntax_getArg(x_12, x_10); +x_14 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__2; +x_15 = lean_name_mk_string(x_2, x_14); +lean_inc(x_13); +x_16 = l_Lean_Syntax_isOfKind(x_13, x_15); +lean_dec(x_15); +if (x_16 == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -lean_dec(x_14); -x_18 = lean_unsigned_to_nat(0u); -x_19 = l_Lean_Syntax_getArg(x_13, x_18); +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_dec(x_13); -x_20 = l_Lean_Syntax_isIdOrAtom_x3f(x_19); -if (lean_obj_tag(x_20) == 0) +x_17 = lean_unsigned_to_nat(0u); +x_18 = l_Lean_Syntax_getArg(x_12, x_17); +lean_dec(x_12); +x_19 = l_Lean_Syntax_isIdOrAtom_x3f(x_18); +if (lean_obj_tag(x_19) == 0) { -lean_object* x_21; lean_object* x_22; -x_21 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_19); -return x_22; +lean_object* x_20; lean_object* x_21; +x_20 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_18); +return x_21; } else { -lean_object* x_23; lean_object* x_24; -x_23 = lean_ctor_get(x_20, 0); -lean_inc(x_23); -lean_dec(x_20); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_19); -return x_24; +lean_object* x_22; lean_object* x_23; +x_22 = lean_ctor_get(x_19, 0); +lean_inc(x_22); +lean_dec(x_19); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_18); +return x_23; } } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; -x_25 = lean_unsigned_to_nat(0u); -x_26 = l_Lean_Syntax_getArg(x_14, x_25); -x_27 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; -lean_inc(x_26); -x_28 = l_Lean_Syntax_isOfKind(x_26, x_27); -if (x_28 == 0) +lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_24 = lean_unsigned_to_nat(0u); +x_25 = l_Lean_Syntax_getArg(x_13, x_24); +x_26 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; +lean_inc(x_25); +x_27 = l_Lean_Syntax_isOfKind(x_25, x_26); +if (x_27 == 0) { -lean_object* x_29; lean_object* x_30; -lean_dec(x_26); -lean_dec(x_14); -x_29 = l_Lean_Syntax_getArg(x_13, x_25); +lean_object* x_28; lean_object* x_29; +lean_dec(x_25); lean_dec(x_13); -x_30 = l_Lean_Syntax_isIdOrAtom_x3f(x_29); -if (lean_obj_tag(x_30) == 0) +x_28 = l_Lean_Syntax_getArg(x_12, x_24); +lean_dec(x_12); +x_29 = l_Lean_Syntax_isIdOrAtom_x3f(x_28); +if (lean_obj_tag(x_29) == 0) { -lean_object* x_31; lean_object* x_32; -x_31 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_29); -return x_32; +lean_object* x_30; lean_object* x_31; +x_30 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_28); +return x_31; } else { -lean_object* x_33; lean_object* x_34; -x_33 = lean_ctor_get(x_30, 0); -lean_inc(x_33); -lean_dec(x_30); -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_29); -return x_34; +lean_object* x_32; lean_object* x_33; +x_32 = lean_ctor_get(x_29, 0); +lean_inc(x_32); +lean_dec(x_29); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_28); +return x_33; } } else { -lean_object* x_35; uint8_t x_36; -x_35 = l_Lean_Syntax_getArg(x_14, x_11); -lean_dec(x_14); -x_36 = l_Lean_Syntax_isNone(x_35); -if (x_36 == 0) +lean_object* x_34; uint8_t x_35; +x_34 = l_Lean_Syntax_getArg(x_13, x_10); +lean_dec(x_13); +x_35 = l_Lean_Syntax_isNone(x_34); +if (x_35 == 0) { -uint8_t x_37; -lean_inc(x_35); -x_37 = l_Lean_Syntax_isNodeOf(x_35, x_10, x_6); -if (x_37 == 0) +uint8_t x_36; +lean_inc(x_34); +x_36 = l_Lean_Syntax_matchesNull(x_34, x_6); +if (x_36 == 0) { -lean_object* x_38; lean_object* x_39; -lean_dec(x_35); -lean_dec(x_26); -x_38 = l_Lean_Syntax_getArg(x_13, x_25); -lean_dec(x_13); -x_39 = l_Lean_Syntax_isIdOrAtom_x3f(x_38); -if (lean_obj_tag(x_39) == 0) +lean_object* x_37; lean_object* x_38; +lean_dec(x_34); +lean_dec(x_25); +x_37 = l_Lean_Syntax_getArg(x_12, x_24); +lean_dec(x_12); +x_38 = l_Lean_Syntax_isIdOrAtom_x3f(x_37); +if (lean_obj_tag(x_38) == 0) { -lean_object* x_40; lean_object* x_41; -x_40 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set(x_41, 1, x_38); -return x_41; +lean_object* x_39; lean_object* x_40; +x_39 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_37); +return x_40; } else { -lean_object* x_42; lean_object* x_43; -x_42 = lean_ctor_get(x_39, 0); -lean_inc(x_42); -lean_dec(x_39); -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_43, 1, x_38); -return x_43; +lean_object* x_41; lean_object* x_42; +x_41 = lean_ctor_get(x_38, 0); +lean_inc(x_41); +lean_dec(x_38); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_37); +return x_42; } } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -lean_dec(x_13); -x_44 = l_Lean_Syntax_getArg(x_35, x_11); -lean_dec(x_35); -x_45 = l_Lean_Syntax_getArgs(x_44); -lean_dec(x_44); -x_46 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_46, 0, x_45); -x_47 = lean_box(0); -x_48 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_26, x_47, x_46); -lean_dec(x_46); -return x_48; +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +lean_dec(x_12); +x_43 = l_Lean_Syntax_getArg(x_34, x_10); +lean_dec(x_34); +x_44 = l_Lean_Syntax_getArgs(x_43); +lean_dec(x_43); +x_45 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_45, 0, x_44); +x_46 = lean_box(0); +x_47 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_25, x_46, x_45); +lean_dec(x_45); +return x_47; } } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; -lean_dec(x_35); -lean_dec(x_13); +lean_object* x_48; lean_object* x_49; lean_object* x_50; +lean_dec(x_34); +lean_dec(x_12); +x_48 = lean_box(0); x_49 = lean_box(0); -x_50 = lean_box(0); -x_51 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_26, x_50, x_49); -return x_51; +x_50 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_25, x_49, x_48); +return x_50; } } } } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; -x_52 = lean_unsigned_to_nat(0u); -x_53 = l_Lean_Syntax_getArg(x_7, x_52); +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_51 = lean_unsigned_to_nat(0u); +x_52 = l_Lean_Syntax_getArg(x_7, x_51); lean_dec(x_7); -x_54 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__2; -x_55 = lean_name_mk_string(x_2, x_54); -lean_inc(x_53); -x_56 = l_Lean_Syntax_isOfKind(x_53, x_55); -if (x_56 == 0) +x_53 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__2; +x_54 = lean_name_mk_string(x_2, x_53); +lean_inc(x_52); +x_55 = l_Lean_Syntax_isOfKind(x_52, x_54); +if (x_55 == 0) { -lean_object* x_57; lean_object* x_58; uint8_t x_59; -lean_dec(x_53); +lean_object* x_56; lean_object* x_57; uint8_t x_58; +lean_dec(x_52); lean_dec(x_8); -x_57 = l_Lean_Syntax_getArg(x_3, x_11); +x_56 = l_Lean_Syntax_getArg(x_3, x_10); lean_dec(x_3); -x_58 = l_Lean_Syntax_getArg(x_57, x_11); -lean_inc(x_58); -x_59 = l_Lean_Syntax_isOfKind(x_58, x_55); -lean_dec(x_55); -if (x_59 == 0) +x_57 = l_Lean_Syntax_getArg(x_56, x_10); +lean_inc(x_57); +x_58 = l_Lean_Syntax_isOfKind(x_57, x_54); +lean_dec(x_54); +if (x_58 == 0) { -lean_object* x_60; lean_object* x_61; -lean_dec(x_58); -x_60 = l_Lean_Syntax_getArg(x_57, x_52); +lean_object* x_59; lean_object* x_60; lean_dec(x_57); -x_61 = l_Lean_Syntax_isIdOrAtom_x3f(x_60); -if (lean_obj_tag(x_61) == 0) +x_59 = l_Lean_Syntax_getArg(x_56, x_51); +lean_dec(x_56); +x_60 = l_Lean_Syntax_isIdOrAtom_x3f(x_59); +if (lean_obj_tag(x_60) == 0) { -lean_object* x_62; lean_object* x_63; -x_62 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_63 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_63, 0, x_62); -lean_ctor_set(x_63, 1, x_60); -return x_63; +lean_object* x_61; lean_object* x_62; +x_61 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_61); +lean_ctor_set(x_62, 1, x_59); +return x_62; } else { -lean_object* x_64; lean_object* x_65; -x_64 = lean_ctor_get(x_61, 0); -lean_inc(x_64); -lean_dec(x_61); -x_65 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_65, 0, x_64); -lean_ctor_set(x_65, 1, x_60); -return x_65; +lean_object* x_63; lean_object* x_64; +x_63 = lean_ctor_get(x_60, 0); +lean_inc(x_63); +lean_dec(x_60); +x_64 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set(x_64, 1, x_59); +return x_64; } } else { -lean_object* x_66; lean_object* x_67; uint8_t x_68; -x_66 = l_Lean_Syntax_getArg(x_58, x_52); -x_67 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; -lean_inc(x_66); -x_68 = l_Lean_Syntax_isOfKind(x_66, x_67); -if (x_68 == 0) +lean_object* x_65; lean_object* x_66; uint8_t x_67; +x_65 = l_Lean_Syntax_getArg(x_57, x_51); +x_66 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; +lean_inc(x_65); +x_67 = l_Lean_Syntax_isOfKind(x_65, x_66); +if (x_67 == 0) { -lean_object* x_69; lean_object* x_70; -lean_dec(x_66); -lean_dec(x_58); -x_69 = l_Lean_Syntax_getArg(x_57, x_52); +lean_object* x_68; lean_object* x_69; +lean_dec(x_65); lean_dec(x_57); -x_70 = l_Lean_Syntax_isIdOrAtom_x3f(x_69); -if (lean_obj_tag(x_70) == 0) +x_68 = l_Lean_Syntax_getArg(x_56, x_51); +lean_dec(x_56); +x_69 = l_Lean_Syntax_isIdOrAtom_x3f(x_68); +if (lean_obj_tag(x_69) == 0) { -lean_object* x_71; lean_object* x_72; -x_71 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_72 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_72, 0, x_71); -lean_ctor_set(x_72, 1, x_69); -return x_72; +lean_object* x_70; lean_object* x_71; +x_70 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set(x_71, 1, x_68); +return x_71; } else { -lean_object* x_73; lean_object* x_74; -x_73 = lean_ctor_get(x_70, 0); -lean_inc(x_73); -lean_dec(x_70); -x_74 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_74, 0, x_73); -lean_ctor_set(x_74, 1, x_69); -return x_74; +lean_object* x_72; lean_object* x_73; +x_72 = lean_ctor_get(x_69, 0); +lean_inc(x_72); +lean_dec(x_69); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_72); +lean_ctor_set(x_73, 1, x_68); +return x_73; } } else { -lean_object* x_75; uint8_t x_76; -x_75 = l_Lean_Syntax_getArg(x_58, x_11); -lean_dec(x_58); -x_76 = l_Lean_Syntax_isNone(x_75); -if (x_76 == 0) +lean_object* x_74; uint8_t x_75; +x_74 = l_Lean_Syntax_getArg(x_57, x_10); +lean_dec(x_57); +x_75 = l_Lean_Syntax_isNone(x_74); +if (x_75 == 0) { -uint8_t x_77; -lean_inc(x_75); -x_77 = l_Lean_Syntax_isNodeOf(x_75, x_10, x_6); -if (x_77 == 0) +uint8_t x_76; +lean_inc(x_74); +x_76 = l_Lean_Syntax_matchesNull(x_74, x_6); +if (x_76 == 0) { -lean_object* x_78; lean_object* x_79; -lean_dec(x_75); -lean_dec(x_66); -x_78 = l_Lean_Syntax_getArg(x_57, x_52); -lean_dec(x_57); -x_79 = l_Lean_Syntax_isIdOrAtom_x3f(x_78); -if (lean_obj_tag(x_79) == 0) +lean_object* x_77; lean_object* x_78; +lean_dec(x_74); +lean_dec(x_65); +x_77 = l_Lean_Syntax_getArg(x_56, x_51); +lean_dec(x_56); +x_78 = l_Lean_Syntax_isIdOrAtom_x3f(x_77); +if (lean_obj_tag(x_78) == 0) { -lean_object* x_80; lean_object* x_81; -x_80 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_81 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_81, 0, x_80); -lean_ctor_set(x_81, 1, x_78); -return x_81; +lean_object* x_79; lean_object* x_80; +x_79 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_80 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_80, 0, x_79); +lean_ctor_set(x_80, 1, x_77); +return x_80; } else { -lean_object* x_82; lean_object* x_83; -x_82 = lean_ctor_get(x_79, 0); -lean_inc(x_82); -lean_dec(x_79); -x_83 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_83, 0, x_82); -lean_ctor_set(x_83, 1, x_78); -return x_83; +lean_object* x_81; lean_object* x_82; +x_81 = lean_ctor_get(x_78, 0); +lean_inc(x_81); +lean_dec(x_78); +x_82 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_82, 0, x_81); +lean_ctor_set(x_82, 1, x_77); +return x_82; } } else { -lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; -lean_dec(x_57); -x_84 = l_Lean_Syntax_getArg(x_75, x_11); -lean_dec(x_75); -x_85 = l_Lean_Syntax_getArgs(x_84); -lean_dec(x_84); -x_86 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_86, 0, x_85); -x_87 = lean_box(0); -x_88 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_66, x_87, x_86); -lean_dec(x_86); -return x_88; +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +lean_dec(x_56); +x_83 = l_Lean_Syntax_getArg(x_74, x_10); +lean_dec(x_74); +x_84 = l_Lean_Syntax_getArgs(x_83); +lean_dec(x_83); +x_85 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_85, 0, x_84); +x_86 = lean_box(0); +x_87 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_65, x_86, x_85); +lean_dec(x_85); +return x_87; } } else { -lean_object* x_89; lean_object* x_90; lean_object* x_91; -lean_dec(x_75); -lean_dec(x_57); +lean_object* x_88; lean_object* x_89; lean_object* x_90; +lean_dec(x_74); +lean_dec(x_56); +x_88 = lean_box(0); x_89 = lean_box(0); -x_90 = lean_box(0); -x_91 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_66, x_90, x_89); -return x_91; +x_90 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_65, x_89, x_88); +return x_90; } } } } else { -lean_object* x_92; lean_object* x_93; uint8_t x_94; -x_92 = l_Lean_Syntax_getArg(x_53, x_52); -x_93 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; -lean_inc(x_92); -x_94 = l_Lean_Syntax_isOfKind(x_92, x_93); -if (x_94 == 0) +lean_object* x_91; lean_object* x_92; uint8_t x_93; +x_91 = l_Lean_Syntax_getArg(x_52, x_51); +x_92 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; +lean_inc(x_91); +x_93 = l_Lean_Syntax_isOfKind(x_91, x_92); +if (x_93 == 0) { -lean_object* x_95; lean_object* x_96; uint8_t x_97; -lean_dec(x_92); -lean_dec(x_53); +lean_object* x_94; lean_object* x_95; uint8_t x_96; +lean_dec(x_91); +lean_dec(x_52); lean_dec(x_8); -x_95 = l_Lean_Syntax_getArg(x_3, x_11); +x_94 = l_Lean_Syntax_getArg(x_3, x_10); lean_dec(x_3); -x_96 = l_Lean_Syntax_getArg(x_95, x_11); -lean_inc(x_96); -x_97 = l_Lean_Syntax_isOfKind(x_96, x_55); -lean_dec(x_55); -if (x_97 == 0) +x_95 = l_Lean_Syntax_getArg(x_94, x_10); +lean_inc(x_95); +x_96 = l_Lean_Syntax_isOfKind(x_95, x_54); +lean_dec(x_54); +if (x_96 == 0) { -lean_object* x_98; lean_object* x_99; -lean_dec(x_96); -x_98 = l_Lean_Syntax_getArg(x_95, x_52); +lean_object* x_97; lean_object* x_98; lean_dec(x_95); -x_99 = l_Lean_Syntax_isIdOrAtom_x3f(x_98); -if (lean_obj_tag(x_99) == 0) +x_97 = l_Lean_Syntax_getArg(x_94, x_51); +lean_dec(x_94); +x_98 = l_Lean_Syntax_isIdOrAtom_x3f(x_97); +if (lean_obj_tag(x_98) == 0) { -lean_object* x_100; lean_object* x_101; -x_100 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_101 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_101, 0, x_100); -lean_ctor_set(x_101, 1, x_98); -return x_101; +lean_object* x_99; lean_object* x_100; +x_99 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_100 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_100, 0, x_99); +lean_ctor_set(x_100, 1, x_97); +return x_100; } else { -lean_object* x_102; lean_object* x_103; -x_102 = lean_ctor_get(x_99, 0); -lean_inc(x_102); -lean_dec(x_99); -x_103 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_103, 0, x_102); -lean_ctor_set(x_103, 1, x_98); -return x_103; +lean_object* x_101; lean_object* x_102; +x_101 = lean_ctor_get(x_98, 0); +lean_inc(x_101); +lean_dec(x_98); +x_102 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_102, 0, x_101); +lean_ctor_set(x_102, 1, x_97); +return x_102; } } else { -lean_object* x_104; uint8_t x_105; -x_104 = l_Lean_Syntax_getArg(x_96, x_52); -lean_inc(x_104); -x_105 = l_Lean_Syntax_isOfKind(x_104, x_93); -if (x_105 == 0) +lean_object* x_103; uint8_t x_104; +x_103 = l_Lean_Syntax_getArg(x_95, x_51); +lean_inc(x_103); +x_104 = l_Lean_Syntax_isOfKind(x_103, x_92); +if (x_104 == 0) { -lean_object* x_106; lean_object* x_107; -lean_dec(x_104); -lean_dec(x_96); -x_106 = l_Lean_Syntax_getArg(x_95, x_52); +lean_object* x_105; lean_object* x_106; +lean_dec(x_103); lean_dec(x_95); -x_107 = l_Lean_Syntax_isIdOrAtom_x3f(x_106); -if (lean_obj_tag(x_107) == 0) +x_105 = l_Lean_Syntax_getArg(x_94, x_51); +lean_dec(x_94); +x_106 = l_Lean_Syntax_isIdOrAtom_x3f(x_105); +if (lean_obj_tag(x_106) == 0) { -lean_object* x_108; lean_object* x_109; -x_108 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_109 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_109, 0, x_108); -lean_ctor_set(x_109, 1, x_106); -return x_109; +lean_object* x_107; lean_object* x_108; +x_107 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_108 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_108, 0, x_107); +lean_ctor_set(x_108, 1, x_105); +return x_108; } else { -lean_object* x_110; lean_object* x_111; -x_110 = lean_ctor_get(x_107, 0); -lean_inc(x_110); -lean_dec(x_107); -x_111 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_111, 0, x_110); -lean_ctor_set(x_111, 1, x_106); -return x_111; +lean_object* x_109; lean_object* x_110; +x_109 = lean_ctor_get(x_106, 0); +lean_inc(x_109); +lean_dec(x_106); +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_109); +lean_ctor_set(x_110, 1, x_105); +return x_110; } } else { -lean_object* x_112; uint8_t x_113; -x_112 = l_Lean_Syntax_getArg(x_96, x_11); -lean_dec(x_96); -x_113 = l_Lean_Syntax_isNone(x_112); -if (x_113 == 0) +lean_object* x_111; uint8_t x_112; +x_111 = l_Lean_Syntax_getArg(x_95, x_10); +lean_dec(x_95); +x_112 = l_Lean_Syntax_isNone(x_111); +if (x_112 == 0) { -uint8_t x_114; -lean_inc(x_112); -x_114 = l_Lean_Syntax_isNodeOf(x_112, x_10, x_6); -if (x_114 == 0) +uint8_t x_113; +lean_inc(x_111); +x_113 = l_Lean_Syntax_matchesNull(x_111, x_6); +if (x_113 == 0) { -lean_object* x_115; lean_object* x_116; -lean_dec(x_112); -lean_dec(x_104); -x_115 = l_Lean_Syntax_getArg(x_95, x_52); -lean_dec(x_95); -x_116 = l_Lean_Syntax_isIdOrAtom_x3f(x_115); -if (lean_obj_tag(x_116) == 0) +lean_object* x_114; lean_object* x_115; +lean_dec(x_111); +lean_dec(x_103); +x_114 = l_Lean_Syntax_getArg(x_94, x_51); +lean_dec(x_94); +x_115 = l_Lean_Syntax_isIdOrAtom_x3f(x_114); +if (lean_obj_tag(x_115) == 0) { -lean_object* x_117; lean_object* x_118; -x_117 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_118 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_118, 0, x_117); -lean_ctor_set(x_118, 1, x_115); -return x_118; +lean_object* x_116; lean_object* x_117; +x_116 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_117 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_117, 0, x_116); +lean_ctor_set(x_117, 1, x_114); +return x_117; } else { -lean_object* x_119; lean_object* x_120; -x_119 = lean_ctor_get(x_116, 0); -lean_inc(x_119); -lean_dec(x_116); -x_120 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_120, 0, x_119); -lean_ctor_set(x_120, 1, x_115); -return x_120; +lean_object* x_118; lean_object* x_119; +x_118 = lean_ctor_get(x_115, 0); +lean_inc(x_118); +lean_dec(x_115); +x_119 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_119, 0, x_118); +lean_ctor_set(x_119, 1, x_114); +return x_119; } } else { -lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; -lean_dec(x_95); -x_121 = l_Lean_Syntax_getArg(x_112, x_11); -lean_dec(x_112); -x_122 = l_Lean_Syntax_getArgs(x_121); -lean_dec(x_121); -x_123 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_123, 0, x_122); -x_124 = lean_box(0); -x_125 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_104, x_124, x_123); -lean_dec(x_123); -return x_125; +lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; +lean_dec(x_94); +x_120 = l_Lean_Syntax_getArg(x_111, x_10); +lean_dec(x_111); +x_121 = l_Lean_Syntax_getArgs(x_120); +lean_dec(x_120); +x_122 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_122, 0, x_121); +x_123 = lean_box(0); +x_124 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_103, x_123, x_122); +lean_dec(x_122); +return x_124; } } else { -lean_object* x_126; lean_object* x_127; lean_object* x_128; -lean_dec(x_112); -lean_dec(x_95); +lean_object* x_125; lean_object* x_126; lean_object* x_127; +lean_dec(x_111); +lean_dec(x_94); +x_125 = lean_box(0); x_126 = lean_box(0); -x_127 = lean_box(0); -x_128 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_104, x_127, x_126); -return x_128; +x_127 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_103, x_126, x_125); +return x_127; } } } } else { -lean_object* x_129; uint8_t x_130; -x_129 = l_Lean_Syntax_getArg(x_53, x_11); -lean_dec(x_53); -x_130 = l_Lean_Syntax_isNone(x_129); -if (x_130 == 0) +lean_object* x_128; uint8_t x_129; +x_128 = l_Lean_Syntax_getArg(x_52, x_10); +lean_dec(x_52); +x_129 = l_Lean_Syntax_isNone(x_128); +if (x_129 == 0) { -uint8_t x_131; -lean_inc(x_129); -x_131 = l_Lean_Syntax_isNodeOf(x_129, x_10, x_6); -if (x_131 == 0) +uint8_t x_130; +lean_inc(x_128); +x_130 = l_Lean_Syntax_matchesNull(x_128, x_6); +if (x_130 == 0) { -lean_object* x_132; lean_object* x_133; uint8_t x_134; -lean_dec(x_129); -lean_dec(x_92); +lean_object* x_131; lean_object* x_132; uint8_t x_133; +lean_dec(x_128); +lean_dec(x_91); lean_dec(x_8); -x_132 = l_Lean_Syntax_getArg(x_3, x_11); +x_131 = l_Lean_Syntax_getArg(x_3, x_10); lean_dec(x_3); -x_133 = l_Lean_Syntax_getArg(x_132, x_11); -lean_inc(x_133); -x_134 = l_Lean_Syntax_isOfKind(x_133, x_55); -lean_dec(x_55); -if (x_134 == 0) +x_132 = l_Lean_Syntax_getArg(x_131, x_10); +lean_inc(x_132); +x_133 = l_Lean_Syntax_isOfKind(x_132, x_54); +lean_dec(x_54); +if (x_133 == 0) { -lean_object* x_135; lean_object* x_136; -lean_dec(x_133); -x_135 = l_Lean_Syntax_getArg(x_132, x_52); +lean_object* x_134; lean_object* x_135; lean_dec(x_132); -x_136 = l_Lean_Syntax_isIdOrAtom_x3f(x_135); -if (lean_obj_tag(x_136) == 0) +x_134 = l_Lean_Syntax_getArg(x_131, x_51); +lean_dec(x_131); +x_135 = l_Lean_Syntax_isIdOrAtom_x3f(x_134); +if (lean_obj_tag(x_135) == 0) { -lean_object* x_137; lean_object* x_138; -x_137 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_138 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_138, 0, x_137); -lean_ctor_set(x_138, 1, x_135); -return x_138; +lean_object* x_136; lean_object* x_137; +x_136 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_137 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_137, 0, x_136); +lean_ctor_set(x_137, 1, x_134); +return x_137; } else { -lean_object* x_139; lean_object* x_140; -x_139 = lean_ctor_get(x_136, 0); -lean_inc(x_139); -lean_dec(x_136); -x_140 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_140, 0, x_139); -lean_ctor_set(x_140, 1, x_135); -return x_140; +lean_object* x_138; lean_object* x_139; +x_138 = lean_ctor_get(x_135, 0); +lean_inc(x_138); +lean_dec(x_135); +x_139 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_139, 0, x_138); +lean_ctor_set(x_139, 1, x_134); +return x_139; } } else { -lean_object* x_141; uint8_t x_142; -x_141 = l_Lean_Syntax_getArg(x_133, x_52); -lean_inc(x_141); -x_142 = l_Lean_Syntax_isOfKind(x_141, x_93); -if (x_142 == 0) +lean_object* x_140; uint8_t x_141; +x_140 = l_Lean_Syntax_getArg(x_132, x_51); +lean_inc(x_140); +x_141 = l_Lean_Syntax_isOfKind(x_140, x_92); +if (x_141 == 0) { -lean_object* x_143; lean_object* x_144; -lean_dec(x_141); -lean_dec(x_133); -x_143 = l_Lean_Syntax_getArg(x_132, x_52); +lean_object* x_142; lean_object* x_143; +lean_dec(x_140); lean_dec(x_132); -x_144 = l_Lean_Syntax_isIdOrAtom_x3f(x_143); -if (lean_obj_tag(x_144) == 0) +x_142 = l_Lean_Syntax_getArg(x_131, x_51); +lean_dec(x_131); +x_143 = l_Lean_Syntax_isIdOrAtom_x3f(x_142); +if (lean_obj_tag(x_143) == 0) { -lean_object* x_145; lean_object* x_146; -x_145 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_146 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_146, 0, x_145); -lean_ctor_set(x_146, 1, x_143); -return x_146; +lean_object* x_144; lean_object* x_145; +x_144 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_145 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_145, 0, x_144); +lean_ctor_set(x_145, 1, x_142); +return x_145; } else { -lean_object* x_147; lean_object* x_148; -x_147 = lean_ctor_get(x_144, 0); -lean_inc(x_147); -lean_dec(x_144); -x_148 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_148, 0, x_147); -lean_ctor_set(x_148, 1, x_143); -return x_148; +lean_object* x_146; lean_object* x_147; +x_146 = lean_ctor_get(x_143, 0); +lean_inc(x_146); +lean_dec(x_143); +x_147 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_147, 0, x_146); +lean_ctor_set(x_147, 1, x_142); +return x_147; } } else { -lean_object* x_149; uint8_t x_150; -x_149 = l_Lean_Syntax_getArg(x_133, x_11); -lean_dec(x_133); -x_150 = l_Lean_Syntax_isNone(x_149); -if (x_150 == 0) +lean_object* x_148; uint8_t x_149; +x_148 = l_Lean_Syntax_getArg(x_132, x_10); +lean_dec(x_132); +x_149 = l_Lean_Syntax_isNone(x_148); +if (x_149 == 0) { -uint8_t x_151; -lean_inc(x_149); -x_151 = l_Lean_Syntax_isNodeOf(x_149, x_10, x_6); -if (x_151 == 0) +uint8_t x_150; +lean_inc(x_148); +x_150 = l_Lean_Syntax_matchesNull(x_148, x_6); +if (x_150 == 0) { -lean_object* x_152; lean_object* x_153; -lean_dec(x_149); -lean_dec(x_141); -x_152 = l_Lean_Syntax_getArg(x_132, x_52); -lean_dec(x_132); -x_153 = l_Lean_Syntax_isIdOrAtom_x3f(x_152); -if (lean_obj_tag(x_153) == 0) +lean_object* x_151; lean_object* x_152; +lean_dec(x_148); +lean_dec(x_140); +x_151 = l_Lean_Syntax_getArg(x_131, x_51); +lean_dec(x_131); +x_152 = l_Lean_Syntax_isIdOrAtom_x3f(x_151); +if (lean_obj_tag(x_152) == 0) { -lean_object* x_154; lean_object* x_155; -x_154 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_155 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_155, 0, x_154); -lean_ctor_set(x_155, 1, x_152); -return x_155; +lean_object* x_153; lean_object* x_154; +x_153 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_154 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_154, 0, x_153); +lean_ctor_set(x_154, 1, x_151); +return x_154; } else { -lean_object* x_156; lean_object* x_157; -x_156 = lean_ctor_get(x_153, 0); -lean_inc(x_156); -lean_dec(x_153); -x_157 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_157, 0, x_156); -lean_ctor_set(x_157, 1, x_152); -return x_157; +lean_object* x_155; lean_object* x_156; +x_155 = lean_ctor_get(x_152, 0); +lean_inc(x_155); +lean_dec(x_152); +x_156 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_156, 0, x_155); +lean_ctor_set(x_156, 1, x_151); +return x_156; } } else { -lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; -lean_dec(x_132); -x_158 = l_Lean_Syntax_getArg(x_149, x_11); -lean_dec(x_149); -x_159 = l_Lean_Syntax_getArgs(x_158); -lean_dec(x_158); -x_160 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_160, 0, x_159); -x_161 = lean_box(0); -x_162 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_141, x_161, x_160); -lean_dec(x_160); -return x_162; +lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; +lean_dec(x_131); +x_157 = l_Lean_Syntax_getArg(x_148, x_10); +lean_dec(x_148); +x_158 = l_Lean_Syntax_getArgs(x_157); +lean_dec(x_157); +x_159 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_159, 0, x_158); +x_160 = lean_box(0); +x_161 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_140, x_160, x_159); +lean_dec(x_159); +return x_161; } } else { -lean_object* x_163; lean_object* x_164; lean_object* x_165; -lean_dec(x_149); -lean_dec(x_132); +lean_object* x_162; lean_object* x_163; lean_object* x_164; +lean_dec(x_148); +lean_dec(x_131); +x_162 = lean_box(0); x_163 = lean_box(0); -x_164 = lean_box(0); -x_165 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_141, x_164, x_163); -return x_165; +x_164 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_140, x_163, x_162); +return x_164; } } } } else { -lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; -lean_dec(x_55); +lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; +lean_dec(x_54); lean_dec(x_3); -x_166 = l_Lean_Syntax_getArg(x_129, x_11); -lean_dec(x_129); -x_167 = l_Lean_Syntax_getArgs(x_166); -lean_dec(x_166); -x_168 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_168, 0, x_167); -x_169 = lean_box(0); -x_170 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__4(x_92, x_8, x_169, x_168); -return x_170; +x_165 = l_Lean_Syntax_getArg(x_128, x_10); +lean_dec(x_128); +x_166 = l_Lean_Syntax_getArgs(x_165); +lean_dec(x_165); +x_167 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_167, 0, x_166); +x_168 = lean_box(0); +x_169 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__4(x_91, x_8, x_168, x_167); +return x_169; } } else { -lean_object* x_171; lean_object* x_172; lean_object* x_173; -lean_dec(x_129); -lean_dec(x_55); +lean_object* x_170; lean_object* x_171; lean_object* x_172; +lean_dec(x_128); +lean_dec(x_54); lean_dec(x_3); +x_170 = lean_box(0); x_171 = lean_box(0); -x_172 = lean_box(0); -x_173 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__4(x_92, x_8, x_172, x_171); -return x_173; +x_172 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__4(x_91, x_8, x_171, x_170); +return x_172; } } } @@ -9832,13 +9825,13 @@ return x_173; } else { -lean_object* x_174; lean_object* x_175; lean_object* x_176; +lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_dec(x_8); lean_dec(x_7); +x_173 = lean_box(0); x_174 = lean_box(0); -x_175 = lean_box(0); -x_176 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3(x_1, x_2, x_3, x_175, x_174, x_174); -return x_176; +x_175 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3(x_1, x_2, x_3, x_174, x_173, x_173); +return x_175; } } } @@ -10001,77 +9994,76 @@ lean_dec(x_33); x_51 = l_Lean_Syntax_isNone(x_50); if (x_51 == 0) { -lean_object* x_52; lean_object* x_53; uint8_t x_54; -x_52 = l_Lean_nullKind; -x_53 = lean_unsigned_to_nat(3u); +lean_object* x_52; uint8_t x_53; +x_52 = lean_unsigned_to_nat(3u); lean_inc(x_50); -x_54 = l_Lean_Syntax_isNodeOf(x_50, x_52, x_53); -if (x_54 == 0) +x_53 = l_Lean_Syntax_matchesNull(x_50, x_52); +if (x_53 == 0) { -lean_object* x_55; lean_object* x_56; +lean_object* x_54; lean_object* x_55; lean_dec(x_50); lean_dec(x_43); lean_free_object(x_10); -x_55 = l_Lean_Syntax_getArg(x_32, x_42); +x_54 = l_Lean_Syntax_getArg(x_32, x_42); lean_dec(x_32); -x_56 = l_Lean_Syntax_isIdOrAtom_x3f(x_55); -if (lean_obj_tag(x_56) == 0) +x_55 = l_Lean_Syntax_isIdOrAtom_x3f(x_54); +if (lean_obj_tag(x_55) == 0) { -lean_object* x_57; -x_57 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_14 = x_57; -x_15 = x_55; +lean_object* x_56; +x_56 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_14 = x_56; +x_15 = x_54; goto block_27; } else { -lean_object* x_58; -x_58 = lean_ctor_get(x_56, 0); -lean_inc(x_58); -lean_dec(x_56); -x_14 = x_58; -x_15 = x_55; +lean_object* x_57; +x_57 = lean_ctor_get(x_55, 0); +lean_inc(x_57); +lean_dec(x_55); +x_14 = x_57; +x_15 = x_54; goto block_27; } } else { -lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_dec(x_32); -x_59 = l_Lean_Syntax_getArg(x_50, x_31); +x_58 = l_Lean_Syntax_getArg(x_50, x_31); lean_dec(x_50); -x_60 = l_Lean_Syntax_getArgs(x_59); -lean_dec(x_59); -lean_ctor_set(x_10, 0, x_60); -x_61 = lean_box(0); -x_62 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_43, x_61, x_10); +x_59 = l_Lean_Syntax_getArgs(x_58); +lean_dec(x_58); +lean_ctor_set(x_10, 0, x_59); +x_60 = lean_box(0); +x_61 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_43, x_60, x_10); lean_dec(x_10); -x_63 = lean_ctor_get(x_62, 0); +x_62 = lean_ctor_get(x_61, 0); +lean_inc(x_62); +x_63 = lean_ctor_get(x_61, 1); lean_inc(x_63); -x_64 = lean_ctor_get(x_62, 1); -lean_inc(x_64); -lean_dec(x_62); -x_14 = x_63; -x_15 = x_64; +lean_dec(x_61); +x_14 = x_62; +x_15 = x_63; goto block_27; } } else { -lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_dec(x_50); lean_dec(x_32); lean_free_object(x_10); +x_64 = lean_box(0); x_65 = lean_box(0); -x_66 = lean_box(0); -x_67 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_43, x_66, x_65); -x_68 = lean_ctor_get(x_67, 0); +x_66 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_43, x_65, x_64); +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_66, 1); lean_inc(x_68); -x_69 = lean_ctor_get(x_67, 1); -lean_inc(x_69); -lean_dec(x_67); -x_14 = x_68; -x_15 = x_69; +lean_dec(x_66); +x_14 = x_67; +x_15 = x_68; goto block_27; } } @@ -10079,169 +10071,168 @@ goto block_27; } else { -lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; uint8_t x_74; -x_70 = lean_unsigned_to_nat(0u); -x_71 = l_Lean_Syntax_getArg(x_6, x_70); -x_72 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__2; +lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; uint8_t x_73; +x_69 = lean_unsigned_to_nat(0u); +x_70 = l_Lean_Syntax_getArg(x_6, x_69); +x_71 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__2; lean_inc(x_5); -x_73 = lean_name_mk_string(x_5, x_72); -x_74 = l_Lean_Syntax_isOfKind(x_71, x_73); -lean_dec(x_73); -if (x_74 == 0) +x_72 = lean_name_mk_string(x_5, x_71); +x_73 = l_Lean_Syntax_isOfKind(x_70, x_72); +lean_dec(x_72); +if (x_73 == 0) { -lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; uint8_t x_80; +lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; lean_dec(x_7); lean_dec(x_6); -x_75 = lean_unsigned_to_nat(1u); -x_76 = l_Lean_Syntax_getArg(x_3, x_75); +x_74 = lean_unsigned_to_nat(1u); +x_75 = l_Lean_Syntax_getArg(x_3, x_74); lean_dec(x_3); -x_77 = l_Lean_Syntax_getArg(x_76, x_75); -x_78 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__2; -x_79 = lean_name_mk_string(x_5, x_78); -lean_inc(x_77); -x_80 = l_Lean_Syntax_isOfKind(x_77, x_79); -lean_dec(x_79); -if (x_80 == 0) +x_76 = l_Lean_Syntax_getArg(x_75, x_74); +x_77 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__2; +x_78 = lean_name_mk_string(x_5, x_77); +lean_inc(x_76); +x_79 = l_Lean_Syntax_isOfKind(x_76, x_78); +lean_dec(x_78); +if (x_79 == 0) { -lean_object* x_81; lean_object* x_82; -lean_dec(x_77); -lean_free_object(x_10); -x_81 = l_Lean_Syntax_getArg(x_76, x_70); +lean_object* x_80; lean_object* x_81; lean_dec(x_76); -x_82 = l_Lean_Syntax_isIdOrAtom_x3f(x_81); -if (lean_obj_tag(x_82) == 0) +lean_free_object(x_10); +x_80 = l_Lean_Syntax_getArg(x_75, x_69); +lean_dec(x_75); +x_81 = l_Lean_Syntax_isIdOrAtom_x3f(x_80); +if (lean_obj_tag(x_81) == 0) { -lean_object* x_83; -x_83 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_14 = x_83; -x_15 = x_81; +lean_object* x_82; +x_82 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_14 = x_82; +x_15 = x_80; goto block_27; } else { -lean_object* x_84; -x_84 = lean_ctor_get(x_82, 0); -lean_inc(x_84); -lean_dec(x_82); -x_14 = x_84; -x_15 = x_81; +lean_object* x_83; +x_83 = lean_ctor_get(x_81, 0); +lean_inc(x_83); +lean_dec(x_81); +x_14 = x_83; +x_15 = x_80; goto block_27; } } else { -lean_object* x_85; lean_object* x_86; uint8_t x_87; -x_85 = l_Lean_Syntax_getArg(x_77, x_70); -x_86 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; -lean_inc(x_85); -x_87 = l_Lean_Syntax_isOfKind(x_85, x_86); -if (x_87 == 0) +lean_object* x_84; lean_object* x_85; uint8_t x_86; +x_84 = l_Lean_Syntax_getArg(x_76, x_69); +x_85 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; +lean_inc(x_84); +x_86 = l_Lean_Syntax_isOfKind(x_84, x_85); +if (x_86 == 0) { -lean_object* x_88; lean_object* x_89; -lean_dec(x_85); -lean_dec(x_77); -lean_free_object(x_10); -x_88 = l_Lean_Syntax_getArg(x_76, x_70); +lean_object* x_87; lean_object* x_88; +lean_dec(x_84); lean_dec(x_76); -x_89 = l_Lean_Syntax_isIdOrAtom_x3f(x_88); -if (lean_obj_tag(x_89) == 0) +lean_free_object(x_10); +x_87 = l_Lean_Syntax_getArg(x_75, x_69); +lean_dec(x_75); +x_88 = l_Lean_Syntax_isIdOrAtom_x3f(x_87); +if (lean_obj_tag(x_88) == 0) { -lean_object* x_90; -x_90 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_14 = x_90; -x_15 = x_88; +lean_object* x_89; +x_89 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_14 = x_89; +x_15 = x_87; goto block_27; } else { -lean_object* x_91; -x_91 = lean_ctor_get(x_89, 0); -lean_inc(x_91); -lean_dec(x_89); -x_14 = x_91; -x_15 = x_88; +lean_object* x_90; +x_90 = lean_ctor_get(x_88, 0); +lean_inc(x_90); +lean_dec(x_88); +x_14 = x_90; +x_15 = x_87; goto block_27; } } else { -lean_object* x_92; uint8_t x_93; -x_92 = l_Lean_Syntax_getArg(x_77, x_75); -lean_dec(x_77); -x_93 = l_Lean_Syntax_isNone(x_92); -if (x_93 == 0) +lean_object* x_91; uint8_t x_92; +x_91 = l_Lean_Syntax_getArg(x_76, x_74); +lean_dec(x_76); +x_92 = l_Lean_Syntax_isNone(x_91); +if (x_92 == 0) { -lean_object* x_94; lean_object* x_95; uint8_t x_96; -x_94 = l_Lean_nullKind; -x_95 = lean_unsigned_to_nat(3u); -lean_inc(x_92); -x_96 = l_Lean_Syntax_isNodeOf(x_92, x_94, x_95); -if (x_96 == 0) +lean_object* x_93; uint8_t x_94; +x_93 = lean_unsigned_to_nat(3u); +lean_inc(x_91); +x_94 = l_Lean_Syntax_matchesNull(x_91, x_93); +if (x_94 == 0) { -lean_object* x_97; lean_object* x_98; -lean_dec(x_92); -lean_dec(x_85); +lean_object* x_95; lean_object* x_96; +lean_dec(x_91); +lean_dec(x_84); lean_free_object(x_10); -x_97 = l_Lean_Syntax_getArg(x_76, x_70); -lean_dec(x_76); -x_98 = l_Lean_Syntax_isIdOrAtom_x3f(x_97); -if (lean_obj_tag(x_98) == 0) +x_95 = l_Lean_Syntax_getArg(x_75, x_69); +lean_dec(x_75); +x_96 = l_Lean_Syntax_isIdOrAtom_x3f(x_95); +if (lean_obj_tag(x_96) == 0) { -lean_object* x_99; -x_99 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_14 = x_99; -x_15 = x_97; +lean_object* x_97; +x_97 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_14 = x_97; +x_15 = x_95; goto block_27; } else { -lean_object* x_100; -x_100 = lean_ctor_get(x_98, 0); -lean_inc(x_100); -lean_dec(x_98); -x_14 = x_100; -x_15 = x_97; +lean_object* x_98; +x_98 = lean_ctor_get(x_96, 0); +lean_inc(x_98); +lean_dec(x_96); +x_14 = x_98; +x_15 = x_95; goto block_27; } } else { -lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; -lean_dec(x_76); -x_101 = l_Lean_Syntax_getArg(x_92, x_75); -lean_dec(x_92); -x_102 = l_Lean_Syntax_getArgs(x_101); -lean_dec(x_101); -lean_ctor_set(x_10, 0, x_102); -x_103 = lean_box(0); -x_104 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_85, x_103, x_10); +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +lean_dec(x_75); +x_99 = l_Lean_Syntax_getArg(x_91, x_74); +lean_dec(x_91); +x_100 = l_Lean_Syntax_getArgs(x_99); +lean_dec(x_99); +lean_ctor_set(x_10, 0, x_100); +x_101 = lean_box(0); +x_102 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_84, x_101, x_10); lean_dec(x_10); -x_105 = lean_ctor_get(x_104, 0); -lean_inc(x_105); -x_106 = lean_ctor_get(x_104, 1); -lean_inc(x_106); -lean_dec(x_104); -x_14 = x_105; -x_15 = x_106; +x_103 = lean_ctor_get(x_102, 0); +lean_inc(x_103); +x_104 = lean_ctor_get(x_102, 1); +lean_inc(x_104); +lean_dec(x_102); +x_14 = x_103; +x_15 = x_104; goto block_27; } } else { -lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; -lean_dec(x_92); -lean_dec(x_76); +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; +lean_dec(x_91); +lean_dec(x_75); lean_free_object(x_10); -x_107 = lean_box(0); -x_108 = lean_box(0); -x_109 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_85, x_108, x_107); -x_110 = lean_ctor_get(x_109, 0); -lean_inc(x_110); -x_111 = lean_ctor_get(x_109, 1); -lean_inc(x_111); -lean_dec(x_109); -x_14 = x_110; -x_15 = x_111; +x_105 = lean_box(0); +x_106 = lean_box(0); +x_107 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_84, x_106, x_105); +x_108 = lean_ctor_get(x_107, 0); +lean_inc(x_108); +x_109 = lean_ctor_get(x_107, 1); +lean_inc(x_109); +lean_dec(x_107); +x_14 = x_108; +x_15 = x_109; goto block_27; } } @@ -10249,170 +10240,169 @@ goto block_27; } else { -lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; uint8_t x_116; -x_112 = lean_unsigned_to_nat(1u); -x_113 = l_Lean_Syntax_getArg(x_6, x_112); +lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; uint8_t x_114; +x_110 = lean_unsigned_to_nat(1u); +x_111 = l_Lean_Syntax_getArg(x_6, x_110); lean_dec(x_6); -x_114 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__3; +x_112 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__3; lean_inc(x_5); -x_115 = lean_name_mk_string(x_5, x_114); -lean_inc(x_113); -x_116 = l_Lean_Syntax_isOfKind(x_113, x_115); -lean_dec(x_115); -if (x_116 == 0) -{ -lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; uint8_t x_121; +x_113 = lean_name_mk_string(x_5, x_112); +lean_inc(x_111); +x_114 = l_Lean_Syntax_isOfKind(x_111, x_113); lean_dec(x_113); +if (x_114 == 0) +{ +lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; uint8_t x_119; +lean_dec(x_111); lean_dec(x_7); -x_117 = l_Lean_Syntax_getArg(x_3, x_112); +x_115 = l_Lean_Syntax_getArg(x_3, x_110); lean_dec(x_3); -x_118 = l_Lean_Syntax_getArg(x_117, x_112); -x_119 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__2; -x_120 = lean_name_mk_string(x_5, x_119); -lean_inc(x_118); -x_121 = l_Lean_Syntax_isOfKind(x_118, x_120); -lean_dec(x_120); -if (x_121 == 0) -{ -lean_object* x_122; lean_object* x_123; +x_116 = l_Lean_Syntax_getArg(x_115, x_110); +x_117 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__2; +x_118 = lean_name_mk_string(x_5, x_117); +lean_inc(x_116); +x_119 = l_Lean_Syntax_isOfKind(x_116, x_118); lean_dec(x_118); +if (x_119 == 0) +{ +lean_object* x_120; lean_object* x_121; +lean_dec(x_116); lean_free_object(x_10); -x_122 = l_Lean_Syntax_getArg(x_117, x_70); -lean_dec(x_117); -x_123 = l_Lean_Syntax_isIdOrAtom_x3f(x_122); -if (lean_obj_tag(x_123) == 0) -{ -lean_object* x_124; -x_124 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_14 = x_124; -x_15 = x_122; +x_120 = l_Lean_Syntax_getArg(x_115, x_69); +lean_dec(x_115); +x_121 = l_Lean_Syntax_isIdOrAtom_x3f(x_120); +if (lean_obj_tag(x_121) == 0) +{ +lean_object* x_122; +x_122 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_14 = x_122; +x_15 = x_120; goto block_27; } else { -lean_object* x_125; -x_125 = lean_ctor_get(x_123, 0); -lean_inc(x_125); -lean_dec(x_123); -x_14 = x_125; -x_15 = x_122; +lean_object* x_123; +x_123 = lean_ctor_get(x_121, 0); +lean_inc(x_123); +lean_dec(x_121); +x_14 = x_123; +x_15 = x_120; goto block_27; } } else { -lean_object* x_126; lean_object* x_127; uint8_t x_128; -x_126 = l_Lean_Syntax_getArg(x_118, x_70); -x_127 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; -lean_inc(x_126); -x_128 = l_Lean_Syntax_isOfKind(x_126, x_127); -if (x_128 == 0) +lean_object* x_124; lean_object* x_125; uint8_t x_126; +x_124 = l_Lean_Syntax_getArg(x_116, x_69); +x_125 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; +lean_inc(x_124); +x_126 = l_Lean_Syntax_isOfKind(x_124, x_125); +if (x_126 == 0) { -lean_object* x_129; lean_object* x_130; -lean_dec(x_126); -lean_dec(x_118); +lean_object* x_127; lean_object* x_128; +lean_dec(x_124); +lean_dec(x_116); lean_free_object(x_10); -x_129 = l_Lean_Syntax_getArg(x_117, x_70); -lean_dec(x_117); -x_130 = l_Lean_Syntax_isIdOrAtom_x3f(x_129); -if (lean_obj_tag(x_130) == 0) +x_127 = l_Lean_Syntax_getArg(x_115, x_69); +lean_dec(x_115); +x_128 = l_Lean_Syntax_isIdOrAtom_x3f(x_127); +if (lean_obj_tag(x_128) == 0) { -lean_object* x_131; -x_131 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_14 = x_131; -x_15 = x_129; +lean_object* x_129; +x_129 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_14 = x_129; +x_15 = x_127; goto block_27; } else { -lean_object* x_132; -x_132 = lean_ctor_get(x_130, 0); -lean_inc(x_132); -lean_dec(x_130); -x_14 = x_132; -x_15 = x_129; +lean_object* x_130; +x_130 = lean_ctor_get(x_128, 0); +lean_inc(x_130); +lean_dec(x_128); +x_14 = x_130; +x_15 = x_127; goto block_27; } } else { +lean_object* x_131; uint8_t x_132; +x_131 = l_Lean_Syntax_getArg(x_116, x_110); +lean_dec(x_116); +x_132 = l_Lean_Syntax_isNone(x_131); +if (x_132 == 0) +{ lean_object* x_133; uint8_t x_134; -x_133 = l_Lean_Syntax_getArg(x_118, x_112); -lean_dec(x_118); -x_134 = l_Lean_Syntax_isNone(x_133); +x_133 = lean_unsigned_to_nat(3u); +lean_inc(x_131); +x_134 = l_Lean_Syntax_matchesNull(x_131, x_133); if (x_134 == 0) { -lean_object* x_135; lean_object* x_136; uint8_t x_137; -x_135 = l_Lean_nullKind; -x_136 = lean_unsigned_to_nat(3u); -lean_inc(x_133); -x_137 = l_Lean_Syntax_isNodeOf(x_133, x_135, x_136); -if (x_137 == 0) -{ -lean_object* x_138; lean_object* x_139; -lean_dec(x_133); -lean_dec(x_126); +lean_object* x_135; lean_object* x_136; +lean_dec(x_131); +lean_dec(x_124); lean_free_object(x_10); -x_138 = l_Lean_Syntax_getArg(x_117, x_70); -lean_dec(x_117); -x_139 = l_Lean_Syntax_isIdOrAtom_x3f(x_138); -if (lean_obj_tag(x_139) == 0) +x_135 = l_Lean_Syntax_getArg(x_115, x_69); +lean_dec(x_115); +x_136 = l_Lean_Syntax_isIdOrAtom_x3f(x_135); +if (lean_obj_tag(x_136) == 0) { -lean_object* x_140; -x_140 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_14 = x_140; -x_15 = x_138; +lean_object* x_137; +x_137 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_14 = x_137; +x_15 = x_135; goto block_27; } else { -lean_object* x_141; -x_141 = lean_ctor_get(x_139, 0); -lean_inc(x_141); -lean_dec(x_139); -x_14 = x_141; -x_15 = x_138; +lean_object* x_138; +x_138 = lean_ctor_get(x_136, 0); +lean_inc(x_138); +lean_dec(x_136); +x_14 = x_138; +x_15 = x_135; goto block_27; } } else { -lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; -lean_dec(x_117); -x_142 = l_Lean_Syntax_getArg(x_133, x_112); -lean_dec(x_133); -x_143 = l_Lean_Syntax_getArgs(x_142); -lean_dec(x_142); -lean_ctor_set(x_10, 0, x_143); -x_144 = lean_box(0); -x_145 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_126, x_144, x_10); +lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; +lean_dec(x_115); +x_139 = l_Lean_Syntax_getArg(x_131, x_110); +lean_dec(x_131); +x_140 = l_Lean_Syntax_getArgs(x_139); +lean_dec(x_139); +lean_ctor_set(x_10, 0, x_140); +x_141 = lean_box(0); +x_142 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_124, x_141, x_10); lean_dec(x_10); -x_146 = lean_ctor_get(x_145, 0); -lean_inc(x_146); -x_147 = lean_ctor_get(x_145, 1); -lean_inc(x_147); -lean_dec(x_145); -x_14 = x_146; -x_15 = x_147; +x_143 = lean_ctor_get(x_142, 0); +lean_inc(x_143); +x_144 = lean_ctor_get(x_142, 1); +lean_inc(x_144); +lean_dec(x_142); +x_14 = x_143; +x_15 = x_144; goto block_27; } } else { -lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; -lean_dec(x_133); -lean_dec(x_117); +lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; +lean_dec(x_131); +lean_dec(x_115); lean_free_object(x_10); -x_148 = lean_box(0); -x_149 = lean_box(0); -x_150 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_126, x_149, x_148); -x_151 = lean_ctor_get(x_150, 0); -lean_inc(x_151); -x_152 = lean_ctor_get(x_150, 1); -lean_inc(x_152); -lean_dec(x_150); -x_14 = x_151; -x_15 = x_152; +x_145 = lean_box(0); +x_146 = lean_box(0); +x_147 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_124, x_146, x_145); +x_148 = lean_ctor_get(x_147, 0); +lean_inc(x_148); +x_149 = lean_ctor_get(x_147, 1); +lean_inc(x_149); +lean_dec(x_147); +x_14 = x_148; +x_15 = x_149; goto block_27; } } @@ -10420,167 +10410,166 @@ goto block_27; } else { -lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; uint8_t x_158; -x_153 = l_Lean_Syntax_getArg(x_113, x_70); -x_154 = l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__5; -x_155 = lean_name_mk_string(x_7, x_154); -x_156 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__4; -x_157 = lean_name_mk_string(x_155, x_156); -x_158 = l_Lean_Syntax_isOfKind(x_153, x_157); -lean_dec(x_157); -if (x_158 == 0) +lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; uint8_t x_155; +x_150 = l_Lean_Syntax_getArg(x_111, x_69); +x_151 = l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__5; +x_152 = lean_name_mk_string(x_7, x_151); +x_153 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__4; +x_154 = lean_name_mk_string(x_152, x_153); +x_155 = l_Lean_Syntax_isOfKind(x_150, x_154); +lean_dec(x_154); +if (x_155 == 0) { -lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; uint8_t x_163; -lean_dec(x_113); -x_159 = l_Lean_Syntax_getArg(x_3, x_112); +lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; uint8_t x_160; +lean_dec(x_111); +x_156 = l_Lean_Syntax_getArg(x_3, x_110); lean_dec(x_3); -x_160 = l_Lean_Syntax_getArg(x_159, x_112); -x_161 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__2; -x_162 = lean_name_mk_string(x_5, x_161); -lean_inc(x_160); -x_163 = l_Lean_Syntax_isOfKind(x_160, x_162); -lean_dec(x_162); -if (x_163 == 0) -{ -lean_object* x_164; lean_object* x_165; -lean_dec(x_160); -lean_free_object(x_10); -x_164 = l_Lean_Syntax_getArg(x_159, x_70); +x_157 = l_Lean_Syntax_getArg(x_156, x_110); +x_158 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__2; +x_159 = lean_name_mk_string(x_5, x_158); +lean_inc(x_157); +x_160 = l_Lean_Syntax_isOfKind(x_157, x_159); lean_dec(x_159); -x_165 = l_Lean_Syntax_isIdOrAtom_x3f(x_164); -if (lean_obj_tag(x_165) == 0) +if (x_160 == 0) { -lean_object* x_166; -x_166 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_14 = x_166; -x_15 = x_164; +lean_object* x_161; lean_object* x_162; +lean_dec(x_157); +lean_free_object(x_10); +x_161 = l_Lean_Syntax_getArg(x_156, x_69); +lean_dec(x_156); +x_162 = l_Lean_Syntax_isIdOrAtom_x3f(x_161); +if (lean_obj_tag(x_162) == 0) +{ +lean_object* x_163; +x_163 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_14 = x_163; +x_15 = x_161; goto block_27; } else { -lean_object* x_167; -x_167 = lean_ctor_get(x_165, 0); -lean_inc(x_167); -lean_dec(x_165); -x_14 = x_167; -x_15 = x_164; +lean_object* x_164; +x_164 = lean_ctor_get(x_162, 0); +lean_inc(x_164); +lean_dec(x_162); +x_14 = x_164; +x_15 = x_161; goto block_27; } } else { -lean_object* x_168; lean_object* x_169; uint8_t x_170; -x_168 = l_Lean_Syntax_getArg(x_160, x_70); -x_169 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; -lean_inc(x_168); -x_170 = l_Lean_Syntax_isOfKind(x_168, x_169); -if (x_170 == 0) +lean_object* x_165; lean_object* x_166; uint8_t x_167; +x_165 = l_Lean_Syntax_getArg(x_157, x_69); +x_166 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; +lean_inc(x_165); +x_167 = l_Lean_Syntax_isOfKind(x_165, x_166); +if (x_167 == 0) { -lean_object* x_171; lean_object* x_172; -lean_dec(x_168); -lean_dec(x_160); +lean_object* x_168; lean_object* x_169; +lean_dec(x_165); +lean_dec(x_157); lean_free_object(x_10); -x_171 = l_Lean_Syntax_getArg(x_159, x_70); -lean_dec(x_159); -x_172 = l_Lean_Syntax_isIdOrAtom_x3f(x_171); -if (lean_obj_tag(x_172) == 0) -{ -lean_object* x_173; -x_173 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_14 = x_173; -x_15 = x_171; +x_168 = l_Lean_Syntax_getArg(x_156, x_69); +lean_dec(x_156); +x_169 = l_Lean_Syntax_isIdOrAtom_x3f(x_168); +if (lean_obj_tag(x_169) == 0) +{ +lean_object* x_170; +x_170 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_14 = x_170; +x_15 = x_168; goto block_27; } else { -lean_object* x_174; -x_174 = lean_ctor_get(x_172, 0); -lean_inc(x_174); -lean_dec(x_172); -x_14 = x_174; -x_15 = x_171; +lean_object* x_171; +x_171 = lean_ctor_get(x_169, 0); +lean_inc(x_171); +lean_dec(x_169); +x_14 = x_171; +x_15 = x_168; goto block_27; } } else { -lean_object* x_175; uint8_t x_176; -x_175 = l_Lean_Syntax_getArg(x_160, x_112); -lean_dec(x_160); -x_176 = l_Lean_Syntax_isNone(x_175); -if (x_176 == 0) +lean_object* x_172; uint8_t x_173; +x_172 = l_Lean_Syntax_getArg(x_157, x_110); +lean_dec(x_157); +x_173 = l_Lean_Syntax_isNone(x_172); +if (x_173 == 0) { -lean_object* x_177; lean_object* x_178; uint8_t x_179; -x_177 = l_Lean_nullKind; -x_178 = lean_unsigned_to_nat(3u); -lean_inc(x_175); -x_179 = l_Lean_Syntax_isNodeOf(x_175, x_177, x_178); -if (x_179 == 0) +lean_object* x_174; uint8_t x_175; +x_174 = lean_unsigned_to_nat(3u); +lean_inc(x_172); +x_175 = l_Lean_Syntax_matchesNull(x_172, x_174); +if (x_175 == 0) { -lean_object* x_180; lean_object* x_181; -lean_dec(x_175); -lean_dec(x_168); +lean_object* x_176; lean_object* x_177; +lean_dec(x_172); +lean_dec(x_165); lean_free_object(x_10); -x_180 = l_Lean_Syntax_getArg(x_159, x_70); -lean_dec(x_159); -x_181 = l_Lean_Syntax_isIdOrAtom_x3f(x_180); -if (lean_obj_tag(x_181) == 0) -{ -lean_object* x_182; -x_182 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_14 = x_182; -x_15 = x_180; +x_176 = l_Lean_Syntax_getArg(x_156, x_69); +lean_dec(x_156); +x_177 = l_Lean_Syntax_isIdOrAtom_x3f(x_176); +if (lean_obj_tag(x_177) == 0) +{ +lean_object* x_178; +x_178 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_14 = x_178; +x_15 = x_176; goto block_27; } else { -lean_object* x_183; -x_183 = lean_ctor_get(x_181, 0); -lean_inc(x_183); -lean_dec(x_181); -x_14 = x_183; -x_15 = x_180; +lean_object* x_179; +x_179 = lean_ctor_get(x_177, 0); +lean_inc(x_179); +lean_dec(x_177); +x_14 = x_179; +x_15 = x_176; goto block_27; } } else { -lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; -lean_dec(x_159); -x_184 = l_Lean_Syntax_getArg(x_175, x_112); -lean_dec(x_175); -x_185 = l_Lean_Syntax_getArgs(x_184); -lean_dec(x_184); -lean_ctor_set(x_10, 0, x_185); -x_186 = lean_box(0); -x_187 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_168, x_186, x_10); +lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; +lean_dec(x_156); +x_180 = l_Lean_Syntax_getArg(x_172, x_110); +lean_dec(x_172); +x_181 = l_Lean_Syntax_getArgs(x_180); +lean_dec(x_180); +lean_ctor_set(x_10, 0, x_181); +x_182 = lean_box(0); +x_183 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_165, x_182, x_10); lean_dec(x_10); -x_188 = lean_ctor_get(x_187, 0); -lean_inc(x_188); -x_189 = lean_ctor_get(x_187, 1); -lean_inc(x_189); -lean_dec(x_187); -x_14 = x_188; -x_15 = x_189; +x_184 = lean_ctor_get(x_183, 0); +lean_inc(x_184); +x_185 = lean_ctor_get(x_183, 1); +lean_inc(x_185); +lean_dec(x_183); +x_14 = x_184; +x_15 = x_185; goto block_27; } } else { -lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; -lean_dec(x_175); -lean_dec(x_159); +lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; +lean_dec(x_172); +lean_dec(x_156); lean_free_object(x_10); -x_190 = lean_box(0); -x_191 = lean_box(0); -x_192 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_168, x_191, x_190); -x_193 = lean_ctor_get(x_192, 0); -lean_inc(x_193); -x_194 = lean_ctor_get(x_192, 1); -lean_inc(x_194); -lean_dec(x_192); -x_14 = x_193; -x_15 = x_194; +x_186 = lean_box(0); +x_187 = lean_box(0); +x_188 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_165, x_187, x_186); +x_189 = lean_ctor_get(x_188, 0); +lean_inc(x_189); +x_190 = lean_ctor_get(x_188, 1); +lean_inc(x_190); +lean_dec(x_188); +x_14 = x_189; +x_15 = x_190; goto block_27; } } @@ -10588,143 +10577,161 @@ goto block_27; } else { -lean_object* x_195; lean_object* x_196; uint8_t x_197; -x_195 = lean_unsigned_to_nat(2u); -x_196 = l_Lean_Syntax_getArg(x_113, x_195); -x_197 = l_Lean_Syntax_isNone(x_196); -if (x_197 == 0) +lean_object* x_191; lean_object* x_192; uint8_t x_193; +x_191 = lean_unsigned_to_nat(2u); +x_192 = l_Lean_Syntax_getArg(x_111, x_191); +x_193 = l_Lean_Syntax_isNone(x_192); +if (x_193 == 0) +{ +uint8_t x_194; +lean_inc(x_192); +x_194 = l_Lean_Syntax_matchesNull(x_192, x_110); +if (x_194 == 0) { -lean_object* x_198; uint8_t x_199; -x_198 = l_Lean_nullKind; +lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; uint8_t x_199; +lean_dec(x_192); +lean_dec(x_111); +x_195 = l_Lean_Syntax_getArg(x_3, x_110); +lean_dec(x_3); +x_196 = l_Lean_Syntax_getArg(x_195, x_110); +x_197 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__2; +x_198 = lean_name_mk_string(x_5, x_197); lean_inc(x_196); -x_199 = l_Lean_Syntax_isNodeOf(x_196, x_198, x_112); +x_199 = l_Lean_Syntax_isOfKind(x_196, x_198); +lean_dec(x_198); if (x_199 == 0) { -lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; uint8_t x_204; +lean_object* x_200; lean_object* x_201; lean_dec(x_196); -lean_dec(x_113); -x_200 = l_Lean_Syntax_getArg(x_3, x_112); -lean_dec(x_3); -x_201 = l_Lean_Syntax_getArg(x_200, x_112); -x_202 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__2; -x_203 = lean_name_mk_string(x_5, x_202); -lean_inc(x_201); -x_204 = l_Lean_Syntax_isOfKind(x_201, x_203); -lean_dec(x_203); -if (x_204 == 0) -{ -lean_object* x_205; lean_object* x_206; -lean_dec(x_201); lean_free_object(x_10); -x_205 = l_Lean_Syntax_getArg(x_200, x_70); -lean_dec(x_200); -x_206 = l_Lean_Syntax_isIdOrAtom_x3f(x_205); -if (lean_obj_tag(x_206) == 0) -{ -lean_object* x_207; -x_207 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_14 = x_207; -x_15 = x_205; +x_200 = l_Lean_Syntax_getArg(x_195, x_69); +lean_dec(x_195); +x_201 = l_Lean_Syntax_isIdOrAtom_x3f(x_200); +if (lean_obj_tag(x_201) == 0) +{ +lean_object* x_202; +x_202 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_14 = x_202; +x_15 = x_200; goto block_27; } else { -lean_object* x_208; -x_208 = lean_ctor_get(x_206, 0); -lean_inc(x_208); -lean_dec(x_206); -x_14 = x_208; -x_15 = x_205; +lean_object* x_203; +x_203 = lean_ctor_get(x_201, 0); +lean_inc(x_203); +lean_dec(x_201); +x_14 = x_203; +x_15 = x_200; goto block_27; } } else { -lean_object* x_209; lean_object* x_210; uint8_t x_211; -x_209 = l_Lean_Syntax_getArg(x_201, x_70); -x_210 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; -lean_inc(x_209); -x_211 = l_Lean_Syntax_isOfKind(x_209, x_210); -if (x_211 == 0) +lean_object* x_204; lean_object* x_205; uint8_t x_206; +x_204 = l_Lean_Syntax_getArg(x_196, x_69); +x_205 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; +lean_inc(x_204); +x_206 = l_Lean_Syntax_isOfKind(x_204, x_205); +if (x_206 == 0) { -lean_object* x_212; lean_object* x_213; -lean_dec(x_209); -lean_dec(x_201); +lean_object* x_207; lean_object* x_208; +lean_dec(x_204); +lean_dec(x_196); lean_free_object(x_10); -x_212 = l_Lean_Syntax_getArg(x_200, x_70); -lean_dec(x_200); -x_213 = l_Lean_Syntax_isIdOrAtom_x3f(x_212); -if (lean_obj_tag(x_213) == 0) -{ -lean_object* x_214; -x_214 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_14 = x_214; -x_15 = x_212; +x_207 = l_Lean_Syntax_getArg(x_195, x_69); +lean_dec(x_195); +x_208 = l_Lean_Syntax_isIdOrAtom_x3f(x_207); +if (lean_obj_tag(x_208) == 0) +{ +lean_object* x_209; +x_209 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_14 = x_209; +x_15 = x_207; goto block_27; } else { -lean_object* x_215; -x_215 = lean_ctor_get(x_213, 0); -lean_inc(x_215); -lean_dec(x_213); -x_14 = x_215; -x_15 = x_212; +lean_object* x_210; +x_210 = lean_ctor_get(x_208, 0); +lean_inc(x_210); +lean_dec(x_208); +x_14 = x_210; +x_15 = x_207; goto block_27; } } else { -lean_object* x_216; uint8_t x_217; -x_216 = l_Lean_Syntax_getArg(x_201, x_112); -lean_dec(x_201); -x_217 = l_Lean_Syntax_isNone(x_216); -if (x_217 == 0) +lean_object* x_211; uint8_t x_212; +x_211 = l_Lean_Syntax_getArg(x_196, x_110); +lean_dec(x_196); +x_212 = l_Lean_Syntax_isNone(x_211); +if (x_212 == 0) { -lean_object* x_218; uint8_t x_219; -x_218 = lean_unsigned_to_nat(3u); -lean_inc(x_216); -x_219 = l_Lean_Syntax_isNodeOf(x_216, x_198, x_218); -if (x_219 == 0) +lean_object* x_213; uint8_t x_214; +x_213 = lean_unsigned_to_nat(3u); +lean_inc(x_211); +x_214 = l_Lean_Syntax_matchesNull(x_211, x_213); +if (x_214 == 0) { -lean_object* x_220; lean_object* x_221; -lean_dec(x_216); -lean_dec(x_209); +lean_object* x_215; lean_object* x_216; +lean_dec(x_211); +lean_dec(x_204); lean_free_object(x_10); -x_220 = l_Lean_Syntax_getArg(x_200, x_70); -lean_dec(x_200); -x_221 = l_Lean_Syntax_isIdOrAtom_x3f(x_220); -if (lean_obj_tag(x_221) == 0) -{ -lean_object* x_222; -x_222 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_14 = x_222; -x_15 = x_220; +x_215 = l_Lean_Syntax_getArg(x_195, x_69); +lean_dec(x_195); +x_216 = l_Lean_Syntax_isIdOrAtom_x3f(x_215); +if (lean_obj_tag(x_216) == 0) +{ +lean_object* x_217; +x_217 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_14 = x_217; +x_15 = x_215; +goto block_27; +} +else +{ +lean_object* x_218; +x_218 = lean_ctor_get(x_216, 0); +lean_inc(x_218); +lean_dec(x_216); +x_14 = x_218; +x_15 = x_215; goto block_27; } +} else { -lean_object* x_223; -x_223 = lean_ctor_get(x_221, 0); +lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; +lean_dec(x_195); +x_219 = l_Lean_Syntax_getArg(x_211, x_110); +lean_dec(x_211); +x_220 = l_Lean_Syntax_getArgs(x_219); +lean_dec(x_219); +lean_ctor_set(x_10, 0, x_220); +x_221 = lean_box(0); +x_222 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_204, x_221, x_10); +lean_dec(x_10); +x_223 = lean_ctor_get(x_222, 0); lean_inc(x_223); -lean_dec(x_221); +x_224 = lean_ctor_get(x_222, 1); +lean_inc(x_224); +lean_dec(x_222); x_14 = x_223; -x_15 = x_220; +x_15 = x_224; goto block_27; } } else { -lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; -lean_dec(x_200); -x_224 = l_Lean_Syntax_getArg(x_216, x_112); -lean_dec(x_216); -x_225 = l_Lean_Syntax_getArgs(x_224); -lean_dec(x_224); -lean_ctor_set(x_10, 0, x_225); +lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; +lean_dec(x_211); +lean_dec(x_195); +lean_free_object(x_10); +x_225 = lean_box(0); x_226 = lean_box(0); -x_227 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_209, x_226, x_10); -lean_dec(x_10); +x_227 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_204, x_226, x_225); x_228 = lean_ctor_get(x_227, 0); lean_inc(x_228); x_229 = lean_ctor_get(x_227, 1); @@ -10735,165 +10742,165 @@ x_15 = x_229; goto block_27; } } -else -{ -lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; -lean_dec(x_216); -lean_dec(x_200); -lean_free_object(x_10); -x_230 = lean_box(0); -x_231 = lean_box(0); -x_232 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_209, x_231, x_230); -x_233 = lean_ctor_get(x_232, 0); -lean_inc(x_233); -x_234 = lean_ctor_get(x_232, 1); -lean_inc(x_234); -lean_dec(x_232); -x_14 = x_233; -x_15 = x_234; -goto block_27; -} -} } } else { -lean_object* x_235; lean_object* x_236; lean_object* x_237; uint8_t x_238; -x_235 = l_Lean_Syntax_getArg(x_196, x_70); -lean_dec(x_196); -x_236 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__5; +lean_object* x_230; lean_object* x_231; lean_object* x_232; uint8_t x_233; +x_230 = l_Lean_Syntax_getArg(x_192, x_69); +lean_dec(x_192); +x_231 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__5; lean_inc(x_5); +x_232 = lean_name_mk_string(x_5, x_231); +lean_inc(x_230); +x_233 = l_Lean_Syntax_isOfKind(x_230, x_232); +lean_dec(x_232); +if (x_233 == 0) +{ +lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; uint8_t x_238; +lean_dec(x_230); +lean_dec(x_111); +x_234 = l_Lean_Syntax_getArg(x_3, x_110); +lean_dec(x_3); +x_235 = l_Lean_Syntax_getArg(x_234, x_110); +x_236 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__2; x_237 = lean_name_mk_string(x_5, x_236); lean_inc(x_235); x_238 = l_Lean_Syntax_isOfKind(x_235, x_237); lean_dec(x_237); if (x_238 == 0) { -lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; uint8_t x_243; +lean_object* x_239; lean_object* x_240; lean_dec(x_235); -lean_dec(x_113); -x_239 = l_Lean_Syntax_getArg(x_3, x_112); -lean_dec(x_3); -x_240 = l_Lean_Syntax_getArg(x_239, x_112); -x_241 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__2; -x_242 = lean_name_mk_string(x_5, x_241); -lean_inc(x_240); -x_243 = l_Lean_Syntax_isOfKind(x_240, x_242); -lean_dec(x_242); -if (x_243 == 0) -{ -lean_object* x_244; lean_object* x_245; -lean_dec(x_240); lean_free_object(x_10); -x_244 = l_Lean_Syntax_getArg(x_239, x_70); -lean_dec(x_239); -x_245 = l_Lean_Syntax_isIdOrAtom_x3f(x_244); -if (lean_obj_tag(x_245) == 0) -{ -lean_object* x_246; -x_246 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_14 = x_246; -x_15 = x_244; +x_239 = l_Lean_Syntax_getArg(x_234, x_69); +lean_dec(x_234); +x_240 = l_Lean_Syntax_isIdOrAtom_x3f(x_239); +if (lean_obj_tag(x_240) == 0) +{ +lean_object* x_241; +x_241 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_14 = x_241; +x_15 = x_239; goto block_27; } else { -lean_object* x_247; -x_247 = lean_ctor_get(x_245, 0); -lean_inc(x_247); -lean_dec(x_245); -x_14 = x_247; -x_15 = x_244; +lean_object* x_242; +x_242 = lean_ctor_get(x_240, 0); +lean_inc(x_242); +lean_dec(x_240); +x_14 = x_242; +x_15 = x_239; goto block_27; } } else { -lean_object* x_248; lean_object* x_249; uint8_t x_250; -x_248 = l_Lean_Syntax_getArg(x_240, x_70); -x_249 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; -lean_inc(x_248); -x_250 = l_Lean_Syntax_isOfKind(x_248, x_249); -if (x_250 == 0) +lean_object* x_243; lean_object* x_244; uint8_t x_245; +x_243 = l_Lean_Syntax_getArg(x_235, x_69); +x_244 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; +lean_inc(x_243); +x_245 = l_Lean_Syntax_isOfKind(x_243, x_244); +if (x_245 == 0) { -lean_object* x_251; lean_object* x_252; -lean_dec(x_248); -lean_dec(x_240); +lean_object* x_246; lean_object* x_247; +lean_dec(x_243); +lean_dec(x_235); lean_free_object(x_10); -x_251 = l_Lean_Syntax_getArg(x_239, x_70); -lean_dec(x_239); -x_252 = l_Lean_Syntax_isIdOrAtom_x3f(x_251); -if (lean_obj_tag(x_252) == 0) -{ -lean_object* x_253; -x_253 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_14 = x_253; -x_15 = x_251; +x_246 = l_Lean_Syntax_getArg(x_234, x_69); +lean_dec(x_234); +x_247 = l_Lean_Syntax_isIdOrAtom_x3f(x_246); +if (lean_obj_tag(x_247) == 0) +{ +lean_object* x_248; +x_248 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_14 = x_248; +x_15 = x_246; goto block_27; } else { -lean_object* x_254; -x_254 = lean_ctor_get(x_252, 0); -lean_inc(x_254); -lean_dec(x_252); -x_14 = x_254; -x_15 = x_251; +lean_object* x_249; +x_249 = lean_ctor_get(x_247, 0); +lean_inc(x_249); +lean_dec(x_247); +x_14 = x_249; +x_15 = x_246; goto block_27; } } else { -lean_object* x_255; uint8_t x_256; -x_255 = l_Lean_Syntax_getArg(x_240, x_112); -lean_dec(x_240); -x_256 = l_Lean_Syntax_isNone(x_255); -if (x_256 == 0) -{ -lean_object* x_257; uint8_t x_258; -x_257 = lean_unsigned_to_nat(3u); -lean_inc(x_255); -x_258 = l_Lean_Syntax_isNodeOf(x_255, x_198, x_257); -if (x_258 == 0) +lean_object* x_250; uint8_t x_251; +x_250 = l_Lean_Syntax_getArg(x_235, x_110); +lean_dec(x_235); +x_251 = l_Lean_Syntax_isNone(x_250); +if (x_251 == 0) +{ +lean_object* x_252; uint8_t x_253; +x_252 = lean_unsigned_to_nat(3u); +lean_inc(x_250); +x_253 = l_Lean_Syntax_matchesNull(x_250, x_252); +if (x_253 == 0) +{ +lean_object* x_254; lean_object* x_255; +lean_dec(x_250); +lean_dec(x_243); +lean_free_object(x_10); +x_254 = l_Lean_Syntax_getArg(x_234, x_69); +lean_dec(x_234); +x_255 = l_Lean_Syntax_isIdOrAtom_x3f(x_254); +if (lean_obj_tag(x_255) == 0) +{ +lean_object* x_256; +x_256 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_14 = x_256; +x_15 = x_254; +goto block_27; +} +else { -lean_object* x_259; lean_object* x_260; +lean_object* x_257; +x_257 = lean_ctor_get(x_255, 0); +lean_inc(x_257); lean_dec(x_255); -lean_dec(x_248); -lean_free_object(x_10); -x_259 = l_Lean_Syntax_getArg(x_239, x_70); -lean_dec(x_239); -x_260 = l_Lean_Syntax_isIdOrAtom_x3f(x_259); -if (lean_obj_tag(x_260) == 0) -{ -lean_object* x_261; -x_261 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_14 = x_261; -x_15 = x_259; +x_14 = x_257; +x_15 = x_254; goto block_27; } +} else { -lean_object* x_262; -x_262 = lean_ctor_get(x_260, 0); +lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; +lean_dec(x_234); +x_258 = l_Lean_Syntax_getArg(x_250, x_110); +lean_dec(x_250); +x_259 = l_Lean_Syntax_getArgs(x_258); +lean_dec(x_258); +lean_ctor_set(x_10, 0, x_259); +x_260 = lean_box(0); +x_261 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_243, x_260, x_10); +lean_dec(x_10); +x_262 = lean_ctor_get(x_261, 0); lean_inc(x_262); -lean_dec(x_260); +x_263 = lean_ctor_get(x_261, 1); +lean_inc(x_263); +lean_dec(x_261); x_14 = x_262; -x_15 = x_259; +x_15 = x_263; goto block_27; } } else { -lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; -lean_dec(x_239); -x_263 = l_Lean_Syntax_getArg(x_255, x_112); -lean_dec(x_255); -x_264 = l_Lean_Syntax_getArgs(x_263); -lean_dec(x_263); -lean_ctor_set(x_10, 0, x_264); +lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; +lean_dec(x_250); +lean_dec(x_234); +lean_free_object(x_10); +x_264 = lean_box(0); x_265 = lean_box(0); -x_266 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_248, x_265, x_10); -lean_dec(x_10); +x_266 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_243, x_265, x_264); x_267 = lean_ctor_get(x_266, 0); lean_inc(x_267); x_268 = lean_ctor_get(x_266, 1); @@ -10904,33 +10911,33 @@ x_15 = x_268; goto block_27; } } +} +} else { -lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; -lean_dec(x_255); -lean_dec(x_239); -lean_free_object(x_10); +lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; +lean_ctor_set(x_10, 0, x_230); x_269 = lean_box(0); -x_270 = lean_box(0); -x_271 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_248, x_270, x_269); -x_272 = lean_ctor_get(x_271, 0); +x_270 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__5(x_111, x_5, x_3, x_269, x_10); +x_271 = lean_ctor_get(x_270, 0); +lean_inc(x_271); +x_272 = lean_ctor_get(x_270, 1); lean_inc(x_272); -x_273 = lean_ctor_get(x_271, 1); -lean_inc(x_273); -lean_dec(x_271); -x_14 = x_272; -x_15 = x_273; +lean_dec(x_270); +x_14 = x_271; +x_15 = x_272; goto block_27; } } } -} else { -lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; -lean_ctor_set(x_10, 0, x_235); +lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; +lean_dec(x_192); +lean_free_object(x_10); +x_273 = lean_box(0); x_274 = lean_box(0); -x_275 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__5(x_113, x_5, x_3, x_274, x_10); +x_275 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__5(x_111, x_5, x_3, x_274, x_273); x_276 = lean_ctor_get(x_275, 0); lean_inc(x_276); x_277 = lean_ctor_get(x_275, 1); @@ -10942,25 +10949,6 @@ goto block_27; } } } -else -{ -lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; -lean_dec(x_196); -lean_free_object(x_10); -x_278 = lean_box(0); -x_279 = lean_box(0); -x_280 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__5(x_113, x_5, x_3, x_279, x_278); -x_281 = lean_ctor_get(x_280, 0); -lean_inc(x_281); -x_282 = lean_ctor_get(x_280, 1); -lean_inc(x_282); -lean_dec(x_280); -x_14 = x_281; -x_15 = x_282; -goto block_27; -} -} -} } } block_27: @@ -11010,381 +10998,413 @@ return x_26; } else { -lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_298; lean_object* x_299; uint8_t x_300; -x_283 = lean_ctor_get(x_10, 0); -lean_inc(x_283); +lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_293; lean_object* x_294; uint8_t x_295; +x_278 = lean_ctor_get(x_10, 0); +lean_inc(x_278); lean_dec(x_10); -x_298 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__1; +x_293 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__1; lean_inc(x_5); -x_299 = lean_name_mk_string(x_5, x_298); +x_294 = lean_name_mk_string(x_5, x_293); lean_inc(x_6); -x_300 = l_Lean_Syntax_isOfKind(x_6, x_299); -lean_dec(x_299); -if (x_300 == 0) +x_295 = l_Lean_Syntax_isOfKind(x_6, x_294); +lean_dec(x_294); +if (x_295 == 0) { -lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; uint8_t x_306; +lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; uint8_t x_301; lean_dec(x_7); lean_dec(x_6); -x_301 = lean_unsigned_to_nat(1u); -x_302 = l_Lean_Syntax_getArg(x_3, x_301); +x_296 = lean_unsigned_to_nat(1u); +x_297 = l_Lean_Syntax_getArg(x_3, x_296); lean_dec(x_3); -x_303 = l_Lean_Syntax_getArg(x_302, x_301); -x_304 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__2; -x_305 = lean_name_mk_string(x_5, x_304); -lean_inc(x_303); -x_306 = l_Lean_Syntax_isOfKind(x_303, x_305); -lean_dec(x_305); -if (x_306 == 0) -{ -lean_object* x_307; lean_object* x_308; lean_object* x_309; -lean_dec(x_303); -x_307 = lean_unsigned_to_nat(0u); -x_308 = l_Lean_Syntax_getArg(x_302, x_307); -lean_dec(x_302); -x_309 = l_Lean_Syntax_isIdOrAtom_x3f(x_308); -if (lean_obj_tag(x_309) == 0) -{ -lean_object* x_310; -x_310 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_284 = x_310; -x_285 = x_308; -goto block_297; -} -else +x_298 = l_Lean_Syntax_getArg(x_297, x_296); +x_299 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__2; +x_300 = lean_name_mk_string(x_5, x_299); +lean_inc(x_298); +x_301 = l_Lean_Syntax_isOfKind(x_298, x_300); +lean_dec(x_300); +if (x_301 == 0) { -lean_object* x_311; -x_311 = lean_ctor_get(x_309, 0); -lean_inc(x_311); -lean_dec(x_309); -x_284 = x_311; -x_285 = x_308; -goto block_297; -} -} -else -{ -lean_object* x_312; lean_object* x_313; lean_object* x_314; uint8_t x_315; -x_312 = lean_unsigned_to_nat(0u); -x_313 = l_Lean_Syntax_getArg(x_303, x_312); -x_314 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; -lean_inc(x_313); -x_315 = l_Lean_Syntax_isOfKind(x_313, x_314); -if (x_315 == 0) +lean_object* x_302; lean_object* x_303; lean_object* x_304; +lean_dec(x_298); +x_302 = lean_unsigned_to_nat(0u); +x_303 = l_Lean_Syntax_getArg(x_297, x_302); +lean_dec(x_297); +x_304 = l_Lean_Syntax_isIdOrAtom_x3f(x_303); +if (lean_obj_tag(x_304) == 0) { -lean_object* x_316; lean_object* x_317; -lean_dec(x_313); -lean_dec(x_303); -x_316 = l_Lean_Syntax_getArg(x_302, x_312); -lean_dec(x_302); -x_317 = l_Lean_Syntax_isIdOrAtom_x3f(x_316); -if (lean_obj_tag(x_317) == 0) -{ -lean_object* x_318; -x_318 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_284 = x_318; -x_285 = x_316; -goto block_297; +lean_object* x_305; +x_305 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_279 = x_305; +x_280 = x_303; +goto block_292; } else { -lean_object* x_319; -x_319 = lean_ctor_get(x_317, 0); -lean_inc(x_319); -lean_dec(x_317); -x_284 = x_319; -x_285 = x_316; -goto block_297; +lean_object* x_306; +x_306 = lean_ctor_get(x_304, 0); +lean_inc(x_306); +lean_dec(x_304); +x_279 = x_306; +x_280 = x_303; +goto block_292; } } else { -lean_object* x_320; uint8_t x_321; -x_320 = l_Lean_Syntax_getArg(x_303, x_301); -lean_dec(x_303); -x_321 = l_Lean_Syntax_isNone(x_320); -if (x_321 == 0) -{ -lean_object* x_322; lean_object* x_323; uint8_t x_324; -x_322 = l_Lean_nullKind; -x_323 = lean_unsigned_to_nat(3u); -lean_inc(x_320); -x_324 = l_Lean_Syntax_isNodeOf(x_320, x_322, x_323); -if (x_324 == 0) -{ -lean_object* x_325; lean_object* x_326; +lean_object* x_307; lean_object* x_308; lean_object* x_309; uint8_t x_310; +x_307 = lean_unsigned_to_nat(0u); +x_308 = l_Lean_Syntax_getArg(x_298, x_307); +x_309 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; +lean_inc(x_308); +x_310 = l_Lean_Syntax_isOfKind(x_308, x_309); +if (x_310 == 0) +{ +lean_object* x_311; lean_object* x_312; +lean_dec(x_308); +lean_dec(x_298); +x_311 = l_Lean_Syntax_getArg(x_297, x_307); +lean_dec(x_297); +x_312 = l_Lean_Syntax_isIdOrAtom_x3f(x_311); +if (lean_obj_tag(x_312) == 0) +{ +lean_object* x_313; +x_313 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_279 = x_313; +x_280 = x_311; +goto block_292; +} +else +{ +lean_object* x_314; +x_314 = lean_ctor_get(x_312, 0); +lean_inc(x_314); +lean_dec(x_312); +x_279 = x_314; +x_280 = x_311; +goto block_292; +} +} +else +{ +lean_object* x_315; uint8_t x_316; +x_315 = l_Lean_Syntax_getArg(x_298, x_296); +lean_dec(x_298); +x_316 = l_Lean_Syntax_isNone(x_315); +if (x_316 == 0) +{ +lean_object* x_317; uint8_t x_318; +x_317 = lean_unsigned_to_nat(3u); +lean_inc(x_315); +x_318 = l_Lean_Syntax_matchesNull(x_315, x_317); +if (x_318 == 0) +{ +lean_object* x_319; lean_object* x_320; +lean_dec(x_315); +lean_dec(x_308); +x_319 = l_Lean_Syntax_getArg(x_297, x_307); +lean_dec(x_297); +x_320 = l_Lean_Syntax_isIdOrAtom_x3f(x_319); +if (lean_obj_tag(x_320) == 0) +{ +lean_object* x_321; +x_321 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_279 = x_321; +x_280 = x_319; +goto block_292; +} +else +{ +lean_object* x_322; +x_322 = lean_ctor_get(x_320, 0); +lean_inc(x_322); lean_dec(x_320); -lean_dec(x_313); -x_325 = l_Lean_Syntax_getArg(x_302, x_312); -lean_dec(x_302); -x_326 = l_Lean_Syntax_isIdOrAtom_x3f(x_325); -if (lean_obj_tag(x_326) == 0) -{ -lean_object* x_327; -x_327 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_284 = x_327; -x_285 = x_325; -goto block_297; +x_279 = x_322; +x_280 = x_319; +goto block_292; +} } else { -lean_object* x_328; -x_328 = lean_ctor_get(x_326, 0); +lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; +lean_dec(x_297); +x_323 = l_Lean_Syntax_getArg(x_315, x_296); +lean_dec(x_315); +x_324 = l_Lean_Syntax_getArgs(x_323); +lean_dec(x_323); +x_325 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_325, 0, x_324); +x_326 = lean_box(0); +x_327 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_308, x_326, x_325); +lean_dec(x_325); +x_328 = lean_ctor_get(x_327, 0); lean_inc(x_328); -lean_dec(x_326); -x_284 = x_328; -x_285 = x_325; -goto block_297; +x_329 = lean_ctor_get(x_327, 1); +lean_inc(x_329); +lean_dec(x_327); +x_279 = x_328; +x_280 = x_329; +goto block_292; } } else { -lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; -lean_dec(x_302); -x_329 = l_Lean_Syntax_getArg(x_320, x_301); -lean_dec(x_320); -x_330 = l_Lean_Syntax_getArgs(x_329); -lean_dec(x_329); -x_331 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_331, 0, x_330); -x_332 = lean_box(0); -x_333 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_313, x_332, x_331); -lean_dec(x_331); -x_334 = lean_ctor_get(x_333, 0); +lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; +lean_dec(x_315); +lean_dec(x_297); +x_330 = lean_box(0); +x_331 = lean_box(0); +x_332 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_308, x_331, x_330); +x_333 = lean_ctor_get(x_332, 0); +lean_inc(x_333); +x_334 = lean_ctor_get(x_332, 1); lean_inc(x_334); -x_335 = lean_ctor_get(x_333, 1); -lean_inc(x_335); -lean_dec(x_333); -x_284 = x_334; -x_285 = x_335; -goto block_297; -} -} -else -{ -lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; -lean_dec(x_320); -lean_dec(x_302); -x_336 = lean_box(0); -x_337 = lean_box(0); -x_338 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_313, x_337, x_336); -x_339 = lean_ctor_get(x_338, 0); -lean_inc(x_339); -x_340 = lean_ctor_get(x_338, 1); -lean_inc(x_340); -lean_dec(x_338); -x_284 = x_339; -x_285 = x_340; -goto block_297; +lean_dec(x_332); +x_279 = x_333; +x_280 = x_334; +goto block_292; } } } } else { -lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; uint8_t x_345; -x_341 = lean_unsigned_to_nat(0u); -x_342 = l_Lean_Syntax_getArg(x_6, x_341); -x_343 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__2; +lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; uint8_t x_339; +x_335 = lean_unsigned_to_nat(0u); +x_336 = l_Lean_Syntax_getArg(x_6, x_335); +x_337 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__2; lean_inc(x_5); -x_344 = lean_name_mk_string(x_5, x_343); -x_345 = l_Lean_Syntax_isOfKind(x_342, x_344); -lean_dec(x_344); -if (x_345 == 0) +x_338 = lean_name_mk_string(x_5, x_337); +x_339 = l_Lean_Syntax_isOfKind(x_336, x_338); +lean_dec(x_338); +if (x_339 == 0) { -lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; uint8_t x_351; +lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; uint8_t x_345; lean_dec(x_7); lean_dec(x_6); -x_346 = lean_unsigned_to_nat(1u); -x_347 = l_Lean_Syntax_getArg(x_3, x_346); +x_340 = lean_unsigned_to_nat(1u); +x_341 = l_Lean_Syntax_getArg(x_3, x_340); lean_dec(x_3); -x_348 = l_Lean_Syntax_getArg(x_347, x_346); -x_349 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__2; -x_350 = lean_name_mk_string(x_5, x_349); -lean_inc(x_348); -x_351 = l_Lean_Syntax_isOfKind(x_348, x_350); -lean_dec(x_350); -if (x_351 == 0) +x_342 = l_Lean_Syntax_getArg(x_341, x_340); +x_343 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__2; +x_344 = lean_name_mk_string(x_5, x_343); +lean_inc(x_342); +x_345 = l_Lean_Syntax_isOfKind(x_342, x_344); +lean_dec(x_344); +if (x_345 == 0) { -lean_object* x_352; lean_object* x_353; -lean_dec(x_348); -x_352 = l_Lean_Syntax_getArg(x_347, x_341); -lean_dec(x_347); -x_353 = l_Lean_Syntax_isIdOrAtom_x3f(x_352); -if (lean_obj_tag(x_353) == 0) +lean_object* x_346; lean_object* x_347; +lean_dec(x_342); +x_346 = l_Lean_Syntax_getArg(x_341, x_335); +lean_dec(x_341); +x_347 = l_Lean_Syntax_isIdOrAtom_x3f(x_346); +if (lean_obj_tag(x_347) == 0) { -lean_object* x_354; -x_354 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_284 = x_354; -x_285 = x_352; -goto block_297; +lean_object* x_348; +x_348 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_279 = x_348; +x_280 = x_346; +goto block_292; } else { -lean_object* x_355; -x_355 = lean_ctor_get(x_353, 0); -lean_inc(x_355); -lean_dec(x_353); -x_284 = x_355; -x_285 = x_352; -goto block_297; +lean_object* x_349; +x_349 = lean_ctor_get(x_347, 0); +lean_inc(x_349); +lean_dec(x_347); +x_279 = x_349; +x_280 = x_346; +goto block_292; } } else { -lean_object* x_356; lean_object* x_357; uint8_t x_358; -x_356 = l_Lean_Syntax_getArg(x_348, x_341); -x_357 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; -lean_inc(x_356); -x_358 = l_Lean_Syntax_isOfKind(x_356, x_357); -if (x_358 == 0) +lean_object* x_350; lean_object* x_351; uint8_t x_352; +x_350 = l_Lean_Syntax_getArg(x_342, x_335); +x_351 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; +lean_inc(x_350); +x_352 = l_Lean_Syntax_isOfKind(x_350, x_351); +if (x_352 == 0) { -lean_object* x_359; lean_object* x_360; -lean_dec(x_356); -lean_dec(x_348); -x_359 = l_Lean_Syntax_getArg(x_347, x_341); -lean_dec(x_347); -x_360 = l_Lean_Syntax_isIdOrAtom_x3f(x_359); -if (lean_obj_tag(x_360) == 0) +lean_object* x_353; lean_object* x_354; +lean_dec(x_350); +lean_dec(x_342); +x_353 = l_Lean_Syntax_getArg(x_341, x_335); +lean_dec(x_341); +x_354 = l_Lean_Syntax_isIdOrAtom_x3f(x_353); +if (lean_obj_tag(x_354) == 0) { -lean_object* x_361; -x_361 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_284 = x_361; -x_285 = x_359; -goto block_297; +lean_object* x_355; +x_355 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_279 = x_355; +x_280 = x_353; +goto block_292; } else { -lean_object* x_362; -x_362 = lean_ctor_get(x_360, 0); -lean_inc(x_362); -lean_dec(x_360); -x_284 = x_362; -x_285 = x_359; -goto block_297; +lean_object* x_356; +x_356 = lean_ctor_get(x_354, 0); +lean_inc(x_356); +lean_dec(x_354); +x_279 = x_356; +x_280 = x_353; +goto block_292; } } else { -lean_object* x_363; uint8_t x_364; -x_363 = l_Lean_Syntax_getArg(x_348, x_346); -lean_dec(x_348); -x_364 = l_Lean_Syntax_isNone(x_363); -if (x_364 == 0) -{ -lean_object* x_365; lean_object* x_366; uint8_t x_367; -x_365 = l_Lean_nullKind; -x_366 = lean_unsigned_to_nat(3u); -lean_inc(x_363); -x_367 = l_Lean_Syntax_isNodeOf(x_363, x_365, x_366); -if (x_367 == 0) -{ -lean_object* x_368; lean_object* x_369; -lean_dec(x_363); -lean_dec(x_356); -x_368 = l_Lean_Syntax_getArg(x_347, x_341); -lean_dec(x_347); -x_369 = l_Lean_Syntax_isIdOrAtom_x3f(x_368); -if (lean_obj_tag(x_369) == 0) +lean_object* x_357; uint8_t x_358; +x_357 = l_Lean_Syntax_getArg(x_342, x_340); +lean_dec(x_342); +x_358 = l_Lean_Syntax_isNone(x_357); +if (x_358 == 0) { -lean_object* x_370; -x_370 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_284 = x_370; -x_285 = x_368; -goto block_297; -} -else +lean_object* x_359; uint8_t x_360; +x_359 = lean_unsigned_to_nat(3u); +lean_inc(x_357); +x_360 = l_Lean_Syntax_matchesNull(x_357, x_359); +if (x_360 == 0) { -lean_object* x_371; -x_371 = lean_ctor_get(x_369, 0); +lean_object* x_361; lean_object* x_362; +lean_dec(x_357); +lean_dec(x_350); +x_361 = l_Lean_Syntax_getArg(x_341, x_335); +lean_dec(x_341); +x_362 = l_Lean_Syntax_isIdOrAtom_x3f(x_361); +if (lean_obj_tag(x_362) == 0) +{ +lean_object* x_363; +x_363 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_279 = x_363; +x_280 = x_361; +goto block_292; +} +else +{ +lean_object* x_364; +x_364 = lean_ctor_get(x_362, 0); +lean_inc(x_364); +lean_dec(x_362); +x_279 = x_364; +x_280 = x_361; +goto block_292; +} +} +else +{ +lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; +lean_dec(x_341); +x_365 = l_Lean_Syntax_getArg(x_357, x_340); +lean_dec(x_357); +x_366 = l_Lean_Syntax_getArgs(x_365); +lean_dec(x_365); +x_367 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_367, 0, x_366); +x_368 = lean_box(0); +x_369 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_350, x_368, x_367); +lean_dec(x_367); +x_370 = lean_ctor_get(x_369, 0); +lean_inc(x_370); +x_371 = lean_ctor_get(x_369, 1); lean_inc(x_371); lean_dec(x_369); -x_284 = x_371; -x_285 = x_368; -goto block_297; +x_279 = x_370; +x_280 = x_371; +goto block_292; } } else { -lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; -lean_dec(x_347); -x_372 = l_Lean_Syntax_getArg(x_363, x_346); -lean_dec(x_363); -x_373 = l_Lean_Syntax_getArgs(x_372); -lean_dec(x_372); -x_374 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_374, 0, x_373); -x_375 = lean_box(0); -x_376 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_356, x_375, x_374); +lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; +lean_dec(x_357); +lean_dec(x_341); +x_372 = lean_box(0); +x_373 = lean_box(0); +x_374 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_350, x_373, x_372); +x_375 = lean_ctor_get(x_374, 0); +lean_inc(x_375); +x_376 = lean_ctor_get(x_374, 1); +lean_inc(x_376); lean_dec(x_374); -x_377 = lean_ctor_get(x_376, 0); -lean_inc(x_377); -x_378 = lean_ctor_get(x_376, 1); -lean_inc(x_378); -lean_dec(x_376); -x_284 = x_377; -x_285 = x_378; -goto block_297; -} -} -else -{ -lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; -lean_dec(x_363); -lean_dec(x_347); -x_379 = lean_box(0); -x_380 = lean_box(0); -x_381 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_356, x_380, x_379); -x_382 = lean_ctor_get(x_381, 0); -lean_inc(x_382); -x_383 = lean_ctor_get(x_381, 1); -lean_inc(x_383); -lean_dec(x_381); -x_284 = x_382; -x_285 = x_383; -goto block_297; +x_279 = x_375; +x_280 = x_376; +goto block_292; } } } } else { -lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; uint8_t x_388; -x_384 = lean_unsigned_to_nat(1u); -x_385 = l_Lean_Syntax_getArg(x_6, x_384); +lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; uint8_t x_381; +x_377 = lean_unsigned_to_nat(1u); +x_378 = l_Lean_Syntax_getArg(x_6, x_377); lean_dec(x_6); -x_386 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__3; +x_379 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__3; lean_inc(x_5); -x_387 = lean_name_mk_string(x_5, x_386); -lean_inc(x_385); -x_388 = l_Lean_Syntax_isOfKind(x_385, x_387); -lean_dec(x_387); -if (x_388 == 0) +x_380 = lean_name_mk_string(x_5, x_379); +lean_inc(x_378); +x_381 = l_Lean_Syntax_isOfKind(x_378, x_380); +lean_dec(x_380); +if (x_381 == 0) { -lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; uint8_t x_393; -lean_dec(x_385); +lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; uint8_t x_386; +lean_dec(x_378); lean_dec(x_7); -x_389 = l_Lean_Syntax_getArg(x_3, x_384); +x_382 = l_Lean_Syntax_getArg(x_3, x_377); lean_dec(x_3); -x_390 = l_Lean_Syntax_getArg(x_389, x_384); -x_391 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__2; -x_392 = lean_name_mk_string(x_5, x_391); +x_383 = l_Lean_Syntax_getArg(x_382, x_377); +x_384 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__2; +x_385 = lean_name_mk_string(x_5, x_384); +lean_inc(x_383); +x_386 = l_Lean_Syntax_isOfKind(x_383, x_385); +lean_dec(x_385); +if (x_386 == 0) +{ +lean_object* x_387; lean_object* x_388; +lean_dec(x_383); +x_387 = l_Lean_Syntax_getArg(x_382, x_335); +lean_dec(x_382); +x_388 = l_Lean_Syntax_isIdOrAtom_x3f(x_387); +if (lean_obj_tag(x_388) == 0) +{ +lean_object* x_389; +x_389 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_279 = x_389; +x_280 = x_387; +goto block_292; +} +else +{ +lean_object* x_390; +x_390 = lean_ctor_get(x_388, 0); lean_inc(x_390); -x_393 = l_Lean_Syntax_isOfKind(x_390, x_392); -lean_dec(x_392); +lean_dec(x_388); +x_279 = x_390; +x_280 = x_387; +goto block_292; +} +} +else +{ +lean_object* x_391; lean_object* x_392; uint8_t x_393; +x_391 = l_Lean_Syntax_getArg(x_383, x_335); +x_392 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; +lean_inc(x_391); +x_393 = l_Lean_Syntax_isOfKind(x_391, x_392); if (x_393 == 0) { lean_object* x_394; lean_object* x_395; -lean_dec(x_390); -x_394 = l_Lean_Syntax_getArg(x_389, x_341); -lean_dec(x_389); +lean_dec(x_391); +lean_dec(x_383); +x_394 = l_Lean_Syntax_getArg(x_382, x_335); +lean_dec(x_382); x_395 = l_Lean_Syntax_isIdOrAtom_x3f(x_394); if (lean_obj_tag(x_395) == 0) { lean_object* x_396; x_396 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_284 = x_396; -x_285 = x_394; -goto block_297; +x_279 = x_396; +x_280 = x_394; +goto block_292; } else { @@ -11392,198 +11412,202 @@ lean_object* x_397; x_397 = lean_ctor_get(x_395, 0); lean_inc(x_397); lean_dec(x_395); -x_284 = x_397; -x_285 = x_394; -goto block_297; +x_279 = x_397; +x_280 = x_394; +goto block_292; } } else { -lean_object* x_398; lean_object* x_399; uint8_t x_400; -x_398 = l_Lean_Syntax_getArg(x_390, x_341); -x_399 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; +lean_object* x_398; uint8_t x_399; +x_398 = l_Lean_Syntax_getArg(x_383, x_377); +lean_dec(x_383); +x_399 = l_Lean_Syntax_isNone(x_398); +if (x_399 == 0) +{ +lean_object* x_400; uint8_t x_401; +x_400 = lean_unsigned_to_nat(3u); lean_inc(x_398); -x_400 = l_Lean_Syntax_isOfKind(x_398, x_399); -if (x_400 == 0) +x_401 = l_Lean_Syntax_matchesNull(x_398, x_400); +if (x_401 == 0) { -lean_object* x_401; lean_object* x_402; +lean_object* x_402; lean_object* x_403; lean_dec(x_398); -lean_dec(x_390); -x_401 = l_Lean_Syntax_getArg(x_389, x_341); -lean_dec(x_389); -x_402 = l_Lean_Syntax_isIdOrAtom_x3f(x_401); -if (lean_obj_tag(x_402) == 0) +lean_dec(x_391); +x_402 = l_Lean_Syntax_getArg(x_382, x_335); +lean_dec(x_382); +x_403 = l_Lean_Syntax_isIdOrAtom_x3f(x_402); +if (lean_obj_tag(x_403) == 0) { -lean_object* x_403; -x_403 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_284 = x_403; -x_285 = x_401; -goto block_297; +lean_object* x_404; +x_404 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_279 = x_404; +x_280 = x_402; +goto block_292; } else { -lean_object* x_404; -x_404 = lean_ctor_get(x_402, 0); -lean_inc(x_404); -lean_dec(x_402); -x_284 = x_404; -x_285 = x_401; -goto block_297; +lean_object* x_405; +x_405 = lean_ctor_get(x_403, 0); +lean_inc(x_405); +lean_dec(x_403); +x_279 = x_405; +x_280 = x_402; +goto block_292; } } else { -lean_object* x_405; uint8_t x_406; -x_405 = l_Lean_Syntax_getArg(x_390, x_384); -lean_dec(x_390); -x_406 = l_Lean_Syntax_isNone(x_405); -if (x_406 == 0) -{ -lean_object* x_407; lean_object* x_408; uint8_t x_409; -x_407 = l_Lean_nullKind; -x_408 = lean_unsigned_to_nat(3u); -lean_inc(x_405); -x_409 = l_Lean_Syntax_isNodeOf(x_405, x_407, x_408); -if (x_409 == 0) -{ -lean_object* x_410; lean_object* x_411; -lean_dec(x_405); +lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; +lean_dec(x_382); +x_406 = l_Lean_Syntax_getArg(x_398, x_377); lean_dec(x_398); -x_410 = l_Lean_Syntax_getArg(x_389, x_341); -lean_dec(x_389); -x_411 = l_Lean_Syntax_isIdOrAtom_x3f(x_410); -if (lean_obj_tag(x_411) == 0) -{ -lean_object* x_412; -x_412 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_284 = x_412; -x_285 = x_410; -goto block_297; -} -else -{ -lean_object* x_413; -x_413 = lean_ctor_get(x_411, 0); -lean_inc(x_413); -lean_dec(x_411); -x_284 = x_413; -x_285 = x_410; -goto block_297; -} -} -else -{ -lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; -lean_dec(x_389); -x_414 = l_Lean_Syntax_getArg(x_405, x_384); -lean_dec(x_405); -x_415 = l_Lean_Syntax_getArgs(x_414); -lean_dec(x_414); -x_416 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_416, 0, x_415); -x_417 = lean_box(0); -x_418 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_398, x_417, x_416); -lean_dec(x_416); -x_419 = lean_ctor_get(x_418, 0); -lean_inc(x_419); -x_420 = lean_ctor_get(x_418, 1); -lean_inc(x_420); -lean_dec(x_418); -x_284 = x_419; -x_285 = x_420; -goto block_297; -} -} -else -{ -lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; -lean_dec(x_405); -lean_dec(x_389); -x_421 = lean_box(0); -x_422 = lean_box(0); -x_423 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_398, x_422, x_421); -x_424 = lean_ctor_get(x_423, 0); -lean_inc(x_424); -x_425 = lean_ctor_get(x_423, 1); -lean_inc(x_425); -lean_dec(x_423); -x_284 = x_424; -x_285 = x_425; -goto block_297; +x_407 = l_Lean_Syntax_getArgs(x_406); +lean_dec(x_406); +x_408 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_408, 0, x_407); +x_409 = lean_box(0); +x_410 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_391, x_409, x_408); +lean_dec(x_408); +x_411 = lean_ctor_get(x_410, 0); +lean_inc(x_411); +x_412 = lean_ctor_get(x_410, 1); +lean_inc(x_412); +lean_dec(x_410); +x_279 = x_411; +x_280 = x_412; +goto block_292; +} +} +else +{ +lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; +lean_dec(x_398); +lean_dec(x_382); +x_413 = lean_box(0); +x_414 = lean_box(0); +x_415 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_391, x_414, x_413); +x_416 = lean_ctor_get(x_415, 0); +lean_inc(x_416); +x_417 = lean_ctor_get(x_415, 1); +lean_inc(x_417); +lean_dec(x_415); +x_279 = x_416; +x_280 = x_417; +goto block_292; } } } } else { -lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; uint8_t x_431; -x_426 = l_Lean_Syntax_getArg(x_385, x_341); -x_427 = l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__5; -x_428 = lean_name_mk_string(x_7, x_427); -x_429 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__4; -x_430 = lean_name_mk_string(x_428, x_429); -x_431 = l_Lean_Syntax_isOfKind(x_426, x_430); -lean_dec(x_430); -if (x_431 == 0) +lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; uint8_t x_423; +x_418 = l_Lean_Syntax_getArg(x_378, x_335); +x_419 = l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__5; +x_420 = lean_name_mk_string(x_7, x_419); +x_421 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__4; +x_422 = lean_name_mk_string(x_420, x_421); +x_423 = l_Lean_Syntax_isOfKind(x_418, x_422); +lean_dec(x_422); +if (x_423 == 0) { -lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; uint8_t x_436; -lean_dec(x_385); -x_432 = l_Lean_Syntax_getArg(x_3, x_384); +lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; uint8_t x_428; +lean_dec(x_378); +x_424 = l_Lean_Syntax_getArg(x_3, x_377); lean_dec(x_3); -x_433 = l_Lean_Syntax_getArg(x_432, x_384); -x_434 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__2; -x_435 = lean_name_mk_string(x_5, x_434); +x_425 = l_Lean_Syntax_getArg(x_424, x_377); +x_426 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__2; +x_427 = lean_name_mk_string(x_5, x_426); +lean_inc(x_425); +x_428 = l_Lean_Syntax_isOfKind(x_425, x_427); +lean_dec(x_427); +if (x_428 == 0) +{ +lean_object* x_429; lean_object* x_430; +lean_dec(x_425); +x_429 = l_Lean_Syntax_getArg(x_424, x_335); +lean_dec(x_424); +x_430 = l_Lean_Syntax_isIdOrAtom_x3f(x_429); +if (lean_obj_tag(x_430) == 0) +{ +lean_object* x_431; +x_431 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_279 = x_431; +x_280 = x_429; +goto block_292; +} +else +{ +lean_object* x_432; +x_432 = lean_ctor_get(x_430, 0); +lean_inc(x_432); +lean_dec(x_430); +x_279 = x_432; +x_280 = x_429; +goto block_292; +} +} +else +{ +lean_object* x_433; lean_object* x_434; uint8_t x_435; +x_433 = l_Lean_Syntax_getArg(x_425, x_335); +x_434 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; lean_inc(x_433); -x_436 = l_Lean_Syntax_isOfKind(x_433, x_435); -lean_dec(x_435); -if (x_436 == 0) +x_435 = l_Lean_Syntax_isOfKind(x_433, x_434); +if (x_435 == 0) { -lean_object* x_437; lean_object* x_438; +lean_object* x_436; lean_object* x_437; lean_dec(x_433); -x_437 = l_Lean_Syntax_getArg(x_432, x_341); -lean_dec(x_432); -x_438 = l_Lean_Syntax_isIdOrAtom_x3f(x_437); -if (lean_obj_tag(x_438) == 0) +lean_dec(x_425); +x_436 = l_Lean_Syntax_getArg(x_424, x_335); +lean_dec(x_424); +x_437 = l_Lean_Syntax_isIdOrAtom_x3f(x_436); +if (lean_obj_tag(x_437) == 0) { -lean_object* x_439; -x_439 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_284 = x_439; -x_285 = x_437; -goto block_297; +lean_object* x_438; +x_438 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_279 = x_438; +x_280 = x_436; +goto block_292; } else { -lean_object* x_440; -x_440 = lean_ctor_get(x_438, 0); -lean_inc(x_440); -lean_dec(x_438); -x_284 = x_440; -x_285 = x_437; -goto block_297; +lean_object* x_439; +x_439 = lean_ctor_get(x_437, 0); +lean_inc(x_439); +lean_dec(x_437); +x_279 = x_439; +x_280 = x_436; +goto block_292; } } else { -lean_object* x_441; lean_object* x_442; uint8_t x_443; -x_441 = l_Lean_Syntax_getArg(x_433, x_341); -x_442 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; -lean_inc(x_441); -x_443 = l_Lean_Syntax_isOfKind(x_441, x_442); +lean_object* x_440; uint8_t x_441; +x_440 = l_Lean_Syntax_getArg(x_425, x_377); +lean_dec(x_425); +x_441 = l_Lean_Syntax_isNone(x_440); +if (x_441 == 0) +{ +lean_object* x_442; uint8_t x_443; +x_442 = lean_unsigned_to_nat(3u); +lean_inc(x_440); +x_443 = l_Lean_Syntax_matchesNull(x_440, x_442); if (x_443 == 0) { lean_object* x_444; lean_object* x_445; -lean_dec(x_441); +lean_dec(x_440); lean_dec(x_433); -x_444 = l_Lean_Syntax_getArg(x_432, x_341); -lean_dec(x_432); +x_444 = l_Lean_Syntax_getArg(x_424, x_335); +lean_dec(x_424); x_445 = l_Lean_Syntax_isIdOrAtom_x3f(x_444); if (lean_obj_tag(x_445) == 0) { lean_object* x_446; x_446 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_284 = x_446; -x_285 = x_444; -goto block_297; +x_279 = x_446; +x_280 = x_444; +goto block_292; } else { @@ -11591,507 +11615,466 @@ lean_object* x_447; x_447 = lean_ctor_get(x_445, 0); lean_inc(x_447); lean_dec(x_445); -x_284 = x_447; -x_285 = x_444; -goto block_297; -} -} -else -{ -lean_object* x_448; uint8_t x_449; -x_448 = l_Lean_Syntax_getArg(x_433, x_384); -lean_dec(x_433); -x_449 = l_Lean_Syntax_isNone(x_448); -if (x_449 == 0) -{ -lean_object* x_450; lean_object* x_451; uint8_t x_452; -x_450 = l_Lean_nullKind; -x_451 = lean_unsigned_to_nat(3u); -lean_inc(x_448); -x_452 = l_Lean_Syntax_isNodeOf(x_448, x_450, x_451); -if (x_452 == 0) -{ -lean_object* x_453; lean_object* x_454; -lean_dec(x_448); -lean_dec(x_441); -x_453 = l_Lean_Syntax_getArg(x_432, x_341); -lean_dec(x_432); -x_454 = l_Lean_Syntax_isIdOrAtom_x3f(x_453); -if (lean_obj_tag(x_454) == 0) -{ -lean_object* x_455; -x_455 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_284 = x_455; -x_285 = x_453; -goto block_297; -} -else -{ -lean_object* x_456; -x_456 = lean_ctor_get(x_454, 0); -lean_inc(x_456); -lean_dec(x_454); -x_284 = x_456; -x_285 = x_453; -goto block_297; +x_279 = x_447; +x_280 = x_444; +goto block_292; } } else { -lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; -lean_dec(x_432); -x_457 = l_Lean_Syntax_getArg(x_448, x_384); +lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; +lean_dec(x_424); +x_448 = l_Lean_Syntax_getArg(x_440, x_377); +lean_dec(x_440); +x_449 = l_Lean_Syntax_getArgs(x_448); lean_dec(x_448); -x_458 = l_Lean_Syntax_getArgs(x_457); +x_450 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_450, 0, x_449); +x_451 = lean_box(0); +x_452 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_433, x_451, x_450); +lean_dec(x_450); +x_453 = lean_ctor_get(x_452, 0); +lean_inc(x_453); +x_454 = lean_ctor_get(x_452, 1); +lean_inc(x_454); +lean_dec(x_452); +x_279 = x_453; +x_280 = x_454; +goto block_292; +} +} +else +{ +lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; +lean_dec(x_440); +lean_dec(x_424); +x_455 = lean_box(0); +x_456 = lean_box(0); +x_457 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_433, x_456, x_455); +x_458 = lean_ctor_get(x_457, 0); +lean_inc(x_458); +x_459 = lean_ctor_get(x_457, 1); +lean_inc(x_459); lean_dec(x_457); -x_459 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_459, 0, x_458); -x_460 = lean_box(0); -x_461 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_441, x_460, x_459); -lean_dec(x_459); -x_462 = lean_ctor_get(x_461, 0); -lean_inc(x_462); -x_463 = lean_ctor_get(x_461, 1); -lean_inc(x_463); -lean_dec(x_461); -x_284 = x_462; -x_285 = x_463; -goto block_297; -} -} -else -{ -lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; -lean_dec(x_448); -lean_dec(x_432); -x_464 = lean_box(0); -x_465 = lean_box(0); -x_466 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_441, x_465, x_464); -x_467 = lean_ctor_get(x_466, 0); -lean_inc(x_467); -x_468 = lean_ctor_get(x_466, 1); -lean_inc(x_468); -lean_dec(x_466); -x_284 = x_467; -x_285 = x_468; -goto block_297; +x_279 = x_458; +x_280 = x_459; +goto block_292; } } } } else { -lean_object* x_469; lean_object* x_470; uint8_t x_471; -x_469 = lean_unsigned_to_nat(2u); -x_470 = l_Lean_Syntax_getArg(x_385, x_469); -x_471 = l_Lean_Syntax_isNone(x_470); -if (x_471 == 0) +lean_object* x_460; lean_object* x_461; uint8_t x_462; +x_460 = lean_unsigned_to_nat(2u); +x_461 = l_Lean_Syntax_getArg(x_378, x_460); +x_462 = l_Lean_Syntax_isNone(x_461); +if (x_462 == 0) { -lean_object* x_472; uint8_t x_473; -x_472 = l_Lean_nullKind; -lean_inc(x_470); -x_473 = l_Lean_Syntax_isNodeOf(x_470, x_472, x_384); -if (x_473 == 0) +uint8_t x_463; +lean_inc(x_461); +x_463 = l_Lean_Syntax_matchesNull(x_461, x_377); +if (x_463 == 0) { -lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; uint8_t x_478; -lean_dec(x_470); -lean_dec(x_385); -x_474 = l_Lean_Syntax_getArg(x_3, x_384); +lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; uint8_t x_468; +lean_dec(x_461); +lean_dec(x_378); +x_464 = l_Lean_Syntax_getArg(x_3, x_377); lean_dec(x_3); -x_475 = l_Lean_Syntax_getArg(x_474, x_384); -x_476 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__2; -x_477 = lean_name_mk_string(x_5, x_476); -lean_inc(x_475); -x_478 = l_Lean_Syntax_isOfKind(x_475, x_477); -lean_dec(x_477); -if (x_478 == 0) -{ -lean_object* x_479; lean_object* x_480; -lean_dec(x_475); -x_479 = l_Lean_Syntax_getArg(x_474, x_341); -lean_dec(x_474); -x_480 = l_Lean_Syntax_isIdOrAtom_x3f(x_479); -if (lean_obj_tag(x_480) == 0) -{ -lean_object* x_481; -x_481 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_284 = x_481; -x_285 = x_479; -goto block_297; -} -else -{ -lean_object* x_482; -x_482 = lean_ctor_get(x_480, 0); -lean_inc(x_482); -lean_dec(x_480); -x_284 = x_482; -x_285 = x_479; -goto block_297; +x_465 = l_Lean_Syntax_getArg(x_464, x_377); +x_466 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__2; +x_467 = lean_name_mk_string(x_5, x_466); +lean_inc(x_465); +x_468 = l_Lean_Syntax_isOfKind(x_465, x_467); +lean_dec(x_467); +if (x_468 == 0) +{ +lean_object* x_469; lean_object* x_470; +lean_dec(x_465); +x_469 = l_Lean_Syntax_getArg(x_464, x_335); +lean_dec(x_464); +x_470 = l_Lean_Syntax_isIdOrAtom_x3f(x_469); +if (lean_obj_tag(x_470) == 0) +{ +lean_object* x_471; +x_471 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_279 = x_471; +x_280 = x_469; +goto block_292; +} +else +{ +lean_object* x_472; +x_472 = lean_ctor_get(x_470, 0); +lean_inc(x_472); +lean_dec(x_470); +x_279 = x_472; +x_280 = x_469; +goto block_292; } } else { -lean_object* x_483; lean_object* x_484; uint8_t x_485; -x_483 = l_Lean_Syntax_getArg(x_475, x_341); -x_484 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; -lean_inc(x_483); -x_485 = l_Lean_Syntax_isOfKind(x_483, x_484); -if (x_485 == 0) +lean_object* x_473; lean_object* x_474; uint8_t x_475; +x_473 = l_Lean_Syntax_getArg(x_465, x_335); +x_474 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; +lean_inc(x_473); +x_475 = l_Lean_Syntax_isOfKind(x_473, x_474); +if (x_475 == 0) { -lean_object* x_486; lean_object* x_487; -lean_dec(x_483); -lean_dec(x_475); -x_486 = l_Lean_Syntax_getArg(x_474, x_341); -lean_dec(x_474); -x_487 = l_Lean_Syntax_isIdOrAtom_x3f(x_486); -if (lean_obj_tag(x_487) == 0) +lean_object* x_476; lean_object* x_477; +lean_dec(x_473); +lean_dec(x_465); +x_476 = l_Lean_Syntax_getArg(x_464, x_335); +lean_dec(x_464); +x_477 = l_Lean_Syntax_isIdOrAtom_x3f(x_476); +if (lean_obj_tag(x_477) == 0) { -lean_object* x_488; -x_488 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_284 = x_488; -x_285 = x_486; -goto block_297; +lean_object* x_478; +x_478 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_279 = x_478; +x_280 = x_476; +goto block_292; } else { -lean_object* x_489; -x_489 = lean_ctor_get(x_487, 0); -lean_inc(x_489); -lean_dec(x_487); -x_284 = x_489; -x_285 = x_486; -goto block_297; +lean_object* x_479; +x_479 = lean_ctor_get(x_477, 0); +lean_inc(x_479); +lean_dec(x_477); +x_279 = x_479; +x_280 = x_476; +goto block_292; } } else { -lean_object* x_490; uint8_t x_491; -x_490 = l_Lean_Syntax_getArg(x_475, x_384); -lean_dec(x_475); -x_491 = l_Lean_Syntax_isNone(x_490); -if (x_491 == 0) +lean_object* x_480; uint8_t x_481; +x_480 = l_Lean_Syntax_getArg(x_465, x_377); +lean_dec(x_465); +x_481 = l_Lean_Syntax_isNone(x_480); +if (x_481 == 0) { -lean_object* x_492; uint8_t x_493; -x_492 = lean_unsigned_to_nat(3u); -lean_inc(x_490); -x_493 = l_Lean_Syntax_isNodeOf(x_490, x_472, x_492); -if (x_493 == 0) +lean_object* x_482; uint8_t x_483; +x_482 = lean_unsigned_to_nat(3u); +lean_inc(x_480); +x_483 = l_Lean_Syntax_matchesNull(x_480, x_482); +if (x_483 == 0) { -lean_object* x_494; lean_object* x_495; -lean_dec(x_490); -lean_dec(x_483); -x_494 = l_Lean_Syntax_getArg(x_474, x_341); -lean_dec(x_474); -x_495 = l_Lean_Syntax_isIdOrAtom_x3f(x_494); -if (lean_obj_tag(x_495) == 0) +lean_object* x_484; lean_object* x_485; +lean_dec(x_480); +lean_dec(x_473); +x_484 = l_Lean_Syntax_getArg(x_464, x_335); +lean_dec(x_464); +x_485 = l_Lean_Syntax_isIdOrAtom_x3f(x_484); +if (lean_obj_tag(x_485) == 0) { -lean_object* x_496; -x_496 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_284 = x_496; -x_285 = x_494; -goto block_297; +lean_object* x_486; +x_486 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_279 = x_486; +x_280 = x_484; +goto block_292; } else { -lean_object* x_497; -x_497 = lean_ctor_get(x_495, 0); -lean_inc(x_497); -lean_dec(x_495); -x_284 = x_497; -x_285 = x_494; -goto block_297; +lean_object* x_487; +x_487 = lean_ctor_get(x_485, 0); +lean_inc(x_487); +lean_dec(x_485); +x_279 = x_487; +x_280 = x_484; +goto block_292; } } else { -lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; -lean_dec(x_474); -x_498 = l_Lean_Syntax_getArg(x_490, x_384); +lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; +lean_dec(x_464); +x_488 = l_Lean_Syntax_getArg(x_480, x_377); +lean_dec(x_480); +x_489 = l_Lean_Syntax_getArgs(x_488); +lean_dec(x_488); +x_490 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_490, 0, x_489); +x_491 = lean_box(0); +x_492 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_473, x_491, x_490); lean_dec(x_490); -x_499 = l_Lean_Syntax_getArgs(x_498); -lean_dec(x_498); -x_500 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_500, 0, x_499); -x_501 = lean_box(0); -x_502 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_483, x_501, x_500); -lean_dec(x_500); -x_503 = lean_ctor_get(x_502, 0); -lean_inc(x_503); -x_504 = lean_ctor_get(x_502, 1); -lean_inc(x_504); -lean_dec(x_502); -x_284 = x_503; -x_285 = x_504; -goto block_297; +x_493 = lean_ctor_get(x_492, 0); +lean_inc(x_493); +x_494 = lean_ctor_get(x_492, 1); +lean_inc(x_494); +lean_dec(x_492); +x_279 = x_493; +x_280 = x_494; +goto block_292; } } else { -lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; -lean_dec(x_490); -lean_dec(x_474); -x_505 = lean_box(0); -x_506 = lean_box(0); -x_507 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_483, x_506, x_505); -x_508 = lean_ctor_get(x_507, 0); -lean_inc(x_508); -x_509 = lean_ctor_get(x_507, 1); -lean_inc(x_509); -lean_dec(x_507); -x_284 = x_508; -x_285 = x_509; -goto block_297; +lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; +lean_dec(x_480); +lean_dec(x_464); +x_495 = lean_box(0); +x_496 = lean_box(0); +x_497 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_473, x_496, x_495); +x_498 = lean_ctor_get(x_497, 0); +lean_inc(x_498); +x_499 = lean_ctor_get(x_497, 1); +lean_inc(x_499); +lean_dec(x_497); +x_279 = x_498; +x_280 = x_499; +goto block_292; } } } } else { -lean_object* x_510; lean_object* x_511; lean_object* x_512; uint8_t x_513; -x_510 = l_Lean_Syntax_getArg(x_470, x_341); -lean_dec(x_470); -x_511 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__5; +lean_object* x_500; lean_object* x_501; lean_object* x_502; uint8_t x_503; +x_500 = l_Lean_Syntax_getArg(x_461, x_335); +lean_dec(x_461); +x_501 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__5; lean_inc(x_5); -x_512 = lean_name_mk_string(x_5, x_511); -lean_inc(x_510); -x_513 = l_Lean_Syntax_isOfKind(x_510, x_512); -lean_dec(x_512); -if (x_513 == 0) +x_502 = lean_name_mk_string(x_5, x_501); +lean_inc(x_500); +x_503 = l_Lean_Syntax_isOfKind(x_500, x_502); +lean_dec(x_502); +if (x_503 == 0) { -lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; uint8_t x_518; -lean_dec(x_510); -lean_dec(x_385); -x_514 = l_Lean_Syntax_getArg(x_3, x_384); +lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; uint8_t x_508; +lean_dec(x_500); +lean_dec(x_378); +x_504 = l_Lean_Syntax_getArg(x_3, x_377); lean_dec(x_3); -x_515 = l_Lean_Syntax_getArg(x_514, x_384); -x_516 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__2; -x_517 = lean_name_mk_string(x_5, x_516); -lean_inc(x_515); -x_518 = l_Lean_Syntax_isOfKind(x_515, x_517); -lean_dec(x_517); -if (x_518 == 0) +x_505 = l_Lean_Syntax_getArg(x_504, x_377); +x_506 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__2; +x_507 = lean_name_mk_string(x_5, x_506); +lean_inc(x_505); +x_508 = l_Lean_Syntax_isOfKind(x_505, x_507); +lean_dec(x_507); +if (x_508 == 0) +{ +lean_object* x_509; lean_object* x_510; +lean_dec(x_505); +x_509 = l_Lean_Syntax_getArg(x_504, x_335); +lean_dec(x_504); +x_510 = l_Lean_Syntax_isIdOrAtom_x3f(x_509); +if (lean_obj_tag(x_510) == 0) +{ +lean_object* x_511; +x_511 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_279 = x_511; +x_280 = x_509; +goto block_292; +} +else +{ +lean_object* x_512; +x_512 = lean_ctor_get(x_510, 0); +lean_inc(x_512); +lean_dec(x_510); +x_279 = x_512; +x_280 = x_509; +goto block_292; +} +} +else { -lean_object* x_519; lean_object* x_520; -lean_dec(x_515); -x_519 = l_Lean_Syntax_getArg(x_514, x_341); -lean_dec(x_514); -x_520 = l_Lean_Syntax_isIdOrAtom_x3f(x_519); -if (lean_obj_tag(x_520) == 0) +lean_object* x_513; lean_object* x_514; uint8_t x_515; +x_513 = l_Lean_Syntax_getArg(x_505, x_335); +x_514 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; +lean_inc(x_513); +x_515 = l_Lean_Syntax_isOfKind(x_513, x_514); +if (x_515 == 0) { -lean_object* x_521; -x_521 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_284 = x_521; -x_285 = x_519; -goto block_297; +lean_object* x_516; lean_object* x_517; +lean_dec(x_513); +lean_dec(x_505); +x_516 = l_Lean_Syntax_getArg(x_504, x_335); +lean_dec(x_504); +x_517 = l_Lean_Syntax_isIdOrAtom_x3f(x_516); +if (lean_obj_tag(x_517) == 0) +{ +lean_object* x_518; +x_518 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_279 = x_518; +x_280 = x_516; +goto block_292; } else { -lean_object* x_522; -x_522 = lean_ctor_get(x_520, 0); -lean_inc(x_522); -lean_dec(x_520); -x_284 = x_522; -x_285 = x_519; -goto block_297; +lean_object* x_519; +x_519 = lean_ctor_get(x_517, 0); +lean_inc(x_519); +lean_dec(x_517); +x_279 = x_519; +x_280 = x_516; +goto block_292; } } else { -lean_object* x_523; lean_object* x_524; uint8_t x_525; -x_523 = l_Lean_Syntax_getArg(x_515, x_341); -x_524 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; -lean_inc(x_523); -x_525 = l_Lean_Syntax_isOfKind(x_523, x_524); -if (x_525 == 0) +lean_object* x_520; uint8_t x_521; +x_520 = l_Lean_Syntax_getArg(x_505, x_377); +lean_dec(x_505); +x_521 = l_Lean_Syntax_isNone(x_520); +if (x_521 == 0) +{ +lean_object* x_522; uint8_t x_523; +x_522 = lean_unsigned_to_nat(3u); +lean_inc(x_520); +x_523 = l_Lean_Syntax_matchesNull(x_520, x_522); +if (x_523 == 0) { -lean_object* x_526; lean_object* x_527; -lean_dec(x_523); -lean_dec(x_515); -x_526 = l_Lean_Syntax_getArg(x_514, x_341); -lean_dec(x_514); -x_527 = l_Lean_Syntax_isIdOrAtom_x3f(x_526); -if (lean_obj_tag(x_527) == 0) +lean_object* x_524; lean_object* x_525; +lean_dec(x_520); +lean_dec(x_513); +x_524 = l_Lean_Syntax_getArg(x_504, x_335); +lean_dec(x_504); +x_525 = l_Lean_Syntax_isIdOrAtom_x3f(x_524); +if (lean_obj_tag(x_525) == 0) { -lean_object* x_528; -x_528 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_284 = x_528; -x_285 = x_526; -goto block_297; +lean_object* x_526; +x_526 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +x_279 = x_526; +x_280 = x_524; +goto block_292; } else { -lean_object* x_529; -x_529 = lean_ctor_get(x_527, 0); -lean_inc(x_529); -lean_dec(x_527); -x_284 = x_529; -x_285 = x_526; -goto block_297; +lean_object* x_527; +x_527 = lean_ctor_get(x_525, 0); +lean_inc(x_527); +lean_dec(x_525); +x_279 = x_527; +x_280 = x_524; +goto block_292; } } else { -lean_object* x_530; uint8_t x_531; -x_530 = l_Lean_Syntax_getArg(x_515, x_384); -lean_dec(x_515); -x_531 = l_Lean_Syntax_isNone(x_530); -if (x_531 == 0) -{ -lean_object* x_532; uint8_t x_533; -x_532 = lean_unsigned_to_nat(3u); -lean_inc(x_530); -x_533 = l_Lean_Syntax_isNodeOf(x_530, x_472, x_532); -if (x_533 == 0) -{ -lean_object* x_534; lean_object* x_535; +lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; +lean_dec(x_504); +x_528 = l_Lean_Syntax_getArg(x_520, x_377); +lean_dec(x_520); +x_529 = l_Lean_Syntax_getArgs(x_528); +lean_dec(x_528); +x_530 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_530, 0, x_529); +x_531 = lean_box(0); +x_532 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_513, x_531, x_530); lean_dec(x_530); -lean_dec(x_523); -x_534 = l_Lean_Syntax_getArg(x_514, x_341); -lean_dec(x_514); -x_535 = l_Lean_Syntax_isIdOrAtom_x3f(x_534); -if (lean_obj_tag(x_535) == 0) -{ -lean_object* x_536; -x_536 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; -x_284 = x_536; -x_285 = x_534; -goto block_297; +x_533 = lean_ctor_get(x_532, 0); +lean_inc(x_533); +x_534 = lean_ctor_get(x_532, 1); +lean_inc(x_534); +lean_dec(x_532); +x_279 = x_533; +x_280 = x_534; +goto block_292; +} } else { -lean_object* x_537; -x_537 = lean_ctor_get(x_535, 0); -lean_inc(x_537); -lean_dec(x_535); -x_284 = x_537; -x_285 = x_534; -goto block_297; +lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; +lean_dec(x_520); +lean_dec(x_504); +x_535 = lean_box(0); +x_536 = lean_box(0); +x_537 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_513, x_536, x_535); +x_538 = lean_ctor_get(x_537, 0); +lean_inc(x_538); +x_539 = lean_ctor_get(x_537, 1); +lean_inc(x_539); +lean_dec(x_537); +x_279 = x_538; +x_280 = x_539; +goto block_292; +} +} } } else { -lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; -lean_dec(x_514); -x_538 = l_Lean_Syntax_getArg(x_530, x_384); -lean_dec(x_530); -x_539 = l_Lean_Syntax_getArgs(x_538); -lean_dec(x_538); +lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; x_540 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_540, 0, x_539); +lean_ctor_set(x_540, 0, x_500); x_541 = lean_box(0); -x_542 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_523, x_541, x_540); -lean_dec(x_540); +x_542 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__5(x_378, x_5, x_3, x_541, x_540); x_543 = lean_ctor_get(x_542, 0); lean_inc(x_543); x_544 = lean_ctor_get(x_542, 1); lean_inc(x_544); lean_dec(x_542); -x_284 = x_543; -x_285 = x_544; -goto block_297; +x_279 = x_543; +x_280 = x_544; +goto block_292; +} } } else { lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; -lean_dec(x_530); -lean_dec(x_514); +lean_dec(x_461); x_545 = lean_box(0); x_546 = lean_box(0); -x_547 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(x_523, x_546, x_545); +x_547 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__5(x_378, x_5, x_3, x_546, x_545); x_548 = lean_ctor_get(x_547, 0); lean_inc(x_548); x_549 = lean_ctor_get(x_547, 1); lean_inc(x_549); lean_dec(x_547); -x_284 = x_548; -x_285 = x_549; -goto block_297; -} -} -} -} -else -{ -lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; -x_550 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_550, 0, x_510); -x_551 = lean_box(0); -x_552 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__5(x_385, x_5, x_3, x_551, x_550); -x_553 = lean_ctor_get(x_552, 0); -lean_inc(x_553); -x_554 = lean_ctor_get(x_552, 1); -lean_inc(x_554); -lean_dec(x_552); -x_284 = x_553; -x_285 = x_554; -goto block_297; -} -} -} -else -{ -lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; -lean_dec(x_470); -x_555 = lean_box(0); -x_556 = lean_box(0); -x_557 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__5(x_385, x_5, x_3, x_556, x_555); -x_558 = lean_ctor_get(x_557, 0); -lean_inc(x_558); -x_559 = lean_ctor_get(x_557, 1); -lean_inc(x_559); -lean_dec(x_557); -x_284 = x_558; -x_285 = x_559; -goto block_297; +x_279 = x_548; +x_280 = x_549; +goto block_292; } } } } } -block_297: +block_292: { -lean_object* x_286; -x_286 = l_Lean_Syntax_getRange_x3f(x_285, x_9); -if (lean_obj_tag(x_286) == 0) +lean_object* x_281; +x_281 = l_Lean_Syntax_getRange_x3f(x_280, x_9); +if (lean_obj_tag(x_281) == 0) { -lean_object* x_287; -lean_dec(x_284); -lean_dec(x_283); -x_287 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_287, 0, x_1); -lean_ctor_set(x_287, 1, x_2); -return x_287; +lean_object* x_282; +lean_dec(x_279); +lean_dec(x_278); +x_282 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_282, 0, x_1); +lean_ctor_set(x_282, 1, x_2); +return x_282; } else { -lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; uint8_t x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; -x_288 = lean_ctor_get(x_286, 0); -lean_inc(x_288); -lean_dec(x_286); -x_289 = lean_box(0); -x_290 = l_String_Range_toLspRange(x_4, x_283); +lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; uint8_t x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; +x_283 = lean_ctor_get(x_281, 0); +lean_inc(x_283); +lean_dec(x_281); +x_284 = lean_box(0); +x_285 = l_String_Range_toLspRange(x_4, x_278); +lean_dec(x_278); +x_286 = l_String_Range_toLspRange(x_4, x_283); lean_dec(x_283); -x_291 = l_String_Range_toLspRange(x_4, x_288); -lean_dec(x_288); -x_292 = 5; -x_293 = lean_alloc_ctor(0, 5, 1); -lean_ctor_set(x_293, 0, x_284); -lean_ctor_set(x_293, 1, x_289); -lean_ctor_set(x_293, 2, x_290); -lean_ctor_set(x_293, 3, x_291); -lean_ctor_set(x_293, 4, x_289); -lean_ctor_set_uint8(x_293, sizeof(void*)*5, x_292); -x_294 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_294, 0, x_293); -x_295 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_295, 0, x_294); -lean_ctor_set(x_295, 1, x_1); -x_296 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_296, 0, x_295); -lean_ctor_set(x_296, 1, x_2); -return x_296; +x_287 = 5; +x_288 = lean_alloc_ctor(0, 5, 1); +lean_ctor_set(x_288, 0, x_279); +lean_ctor_set(x_288, 1, x_284); +lean_ctor_set(x_288, 2, x_285); +lean_ctor_set(x_288, 3, x_286); +lean_ctor_set(x_288, 4, x_284); +lean_ctor_set_uint8(x_288, sizeof(void*)*5, x_287); +x_289 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_289, 0, x_288); +x_290 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_290, 0, x_289); +lean_ctor_set(x_290, 1, x_1); +x_291 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_291, 0, x_290); +lean_ctor_set(x_291, 1, x_2); +return x_291; } } } @@ -12498,6 +12481,24 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } +static lean_object* _init_l_Lean_Server_FileWorker_handleDocumentSymbol_sectionLikeToDocumentSymbols___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("null", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Server_FileWorker_handleDocumentSymbol_sectionLikeToDocumentSymbols___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Server_FileWorker_handleDocumentSymbol_sectionLikeToDocumentSymbols___closed__2; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_sectionLikeToDocumentSymbols(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6) { _start: { @@ -12563,7 +12564,7 @@ goto block_56; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; x_26 = lean_array_push(x_17, x_25); x_27 = lean_box(2); -x_28 = l_Lean_nullKind; +x_28 = l_Lean_Server_FileWorker_handleDocumentSymbol_sectionLikeToDocumentSymbols___closed__3; x_29 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_29, 0, x_27); lean_ctor_set(x_29, 1, x_28); @@ -12892,13 +12893,11 @@ LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol___rarg___ _start: { lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_7 = lean_ctor_get(x_1, 1); -lean_inc(x_7); -x_8 = l_List_getLastD___rarg(x_2, x_7); -lean_dec(x_7); -x_9 = lean_ctor_get(x_1, 0); +x_7 = l_Lean_Server_Snapshots_instInhabitedSnapshot; +x_8 = l_List_getLast_x21___rarg(x_7, x_1); +x_9 = lean_ctor_get(x_2, 0); lean_inc(x_9); -lean_dec(x_1); +lean_dec(x_2); lean_inc(x_9); x_10 = l_Lean_Server_DocumentMeta_mkInputContext(x_9); x_11 = l_Lean_Server_Snapshots_parseAhead(x_10, x_8, x_6); @@ -12969,7 +12968,7 @@ x_8 = lean_box(0); lean_inc(x_7); x_9 = l_List_mapTRAux___at_Lean_Server_FileWorker_handleDocumentSymbol___spec__1(x_7, x_8); x_10 = lean_box(0); -x_11 = l_Lean_Server_FileWorker_handleDocumentSymbol___rarg___lambda__2(x_1, x_7, x_9, x_10, x_3, x_4); +x_11 = l_Lean_Server_FileWorker_handleDocumentSymbol___rarg___lambda__2(x_7, x_1, x_9, x_10, x_3, x_4); lean_dec(x_3); return x_11; } @@ -13020,7 +13019,7 @@ lean_inc(x_4); x_5 = lean_ctor_get(x_3, 1); lean_inc(x_5); lean_dec(x_3); -x_6 = lean_ctor_get(x_4, 2); +x_6 = lean_ctor_get(x_4, 1); lean_inc(x_6); x_7 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleDocumentSymbol___rarg___lambda__1___boxed), 3, 1); lean_closure_set(x_7, 0, x_6); @@ -14271,6 +14270,24 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } +static lean_object* _init_l_Lean_Server_FileWorker_handleSemanticTokens_go___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("choice", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Server_FileWorker_handleSemanticTokens_go___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Server_FileWorker_handleSemanticTokens_go___closed__3; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokens_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { @@ -14310,7 +14327,7 @@ x_17 = lean_ctor_get(x_13, 1); x_18 = lean_ctor_get(x_15, 1); x_19 = lean_ctor_get(x_15, 0); lean_dec(x_19); -x_20 = l_Lean_choiceKind; +x_20 = l_Lean_Server_FileWorker_handleSemanticTokens_go___closed__4; lean_inc(x_1); x_21 = l_Lean_Syntax_isOfKind(x_1, x_20); if (x_21 == 0) @@ -14383,7 +14400,7 @@ x_36 = lean_ctor_get(x_13, 1); x_37 = lean_ctor_get(x_15, 1); lean_inc(x_37); lean_dec(x_15); -x_38 = l_Lean_choiceKind; +x_38 = l_Lean_Server_FileWorker_handleSemanticTokens_go___closed__4; lean_inc(x_1); x_39 = l_Lean_Syntax_isOfKind(x_1, x_38); if (x_39 == 0) @@ -14472,7 +14489,7 @@ if (lean_is_exclusive(x_56)) { lean_dec_ref(x_56); x_59 = lean_box(0); } -x_60 = l_Lean_choiceKind; +x_60 = l_Lean_Server_FileWorker_handleSemanticTokens_go___closed__4; lean_inc(x_1); x_61 = l_Lean_Syntax_isOfKind(x_1, x_60); if (x_61 == 0) @@ -14617,7 +14634,7 @@ x_97 = lean_ctor_get(x_93, 1); x_98 = lean_ctor_get(x_95, 1); x_99 = lean_ctor_get(x_95, 0); lean_dec(x_99); -x_100 = l_Lean_choiceKind; +x_100 = l_Lean_Server_FileWorker_handleSemanticTokens_go___closed__4; lean_inc(x_1); x_101 = l_Lean_Syntax_isOfKind(x_1, x_100); if (x_101 == 0) @@ -14688,7 +14705,7 @@ x_114 = lean_ctor_get(x_93, 1); x_115 = lean_ctor_get(x_95, 1); lean_inc(x_115); lean_dec(x_95); -x_116 = l_Lean_choiceKind; +x_116 = l_Lean_Server_FileWorker_handleSemanticTokens_go___closed__4; lean_inc(x_1); x_117 = l_Lean_Syntax_isOfKind(x_1, x_116); if (x_117 == 0) @@ -14775,7 +14792,7 @@ if (lean_is_exclusive(x_132)) { lean_dec_ref(x_132); x_135 = lean_box(0); } -x_136 = l_Lean_choiceKind; +x_136 = l_Lean_Server_FileWorker_handleSemanticTokens_go___closed__4; lean_inc(x_1); x_137 = l_Lean_Syntax_isOfKind(x_1, x_136); if (x_137 == 0) @@ -15171,7 +15188,7 @@ lean_dec(x_8); lean_inc(x_2); x_10 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleSemanticTokens___lambda__1___boxed), 2, 1); lean_closure_set(x_10, 0, x_2); -x_11 = lean_ctor_get(x_6, 2); +x_11 = lean_ctor_get(x_6, 1); lean_inc(x_11); lean_dec(x_6); x_12 = l_IO_AsyncList_waitAll___rarg(x_10, x_11, x_7); @@ -16028,159 +16045,158 @@ goto _start; } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; x_41 = lean_unsigned_to_nat(1u); x_42 = l_Lean_Syntax_getArg(x_11, x_41); x_43 = lean_unsigned_to_nat(3u); x_44 = l_Lean_Syntax_getArg(x_11, x_43); -x_45 = l_Lean_nullKind; -x_46 = lean_unsigned_to_nat(0u); -x_47 = l_Lean_Syntax_isNodeOf(x_44, x_45, x_46); -if (x_47 == 0) +x_45 = lean_unsigned_to_nat(0u); +x_46 = l_Lean_Syntax_matchesNull(x_44, x_45); +if (x_46 == 0) { -uint8_t x_48; +uint8_t x_47; lean_dec(x_42); lean_inc(x_11); -x_48 = l_Lean_Server_FileWorker_handleFoldingRange_isImport(x_11); -if (x_48 == 0) +x_47 = l_Lean_Server_FileWorker_handleFoldingRange_isImport(x_11); +if (x_47 == 0) { -lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_49 = l_Lean_Server_FileWorker_handleFoldingRange_addCommandRange(x_1, x_11, x_4, x_5, x_6); -x_50 = lean_ctor_get(x_49, 0); +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_48 = l_Lean_Server_FileWorker_handleFoldingRange_addCommandRange(x_1, x_11, x_4, x_5, x_6); +x_49 = lean_ctor_get(x_48, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_48, 1); lean_inc(x_50); +lean_dec(x_48); x_51 = lean_ctor_get(x_49, 1); lean_inc(x_51); lean_dec(x_49); -x_52 = lean_ctor_get(x_50, 1); -lean_inc(x_52); -lean_dec(x_50); x_3 = x_12; -x_4 = x_52; -x_6 = x_51; +x_4 = x_51; +x_6 = x_50; goto _start; } else { -lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_54 = lean_box(0); -x_55 = l_List_spanAux___at_Lean_Server_FileWorker_handleFoldingRange_addRanges___spec__1(x_12, x_54); -x_56 = lean_ctor_get(x_55, 0); +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; uint8_t x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_53 = lean_box(0); +x_54 = l_List_spanAux___at_Lean_Server_FileWorker_handleFoldingRange_addRanges___spec__1(x_12, x_53); +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_54, 1); lean_inc(x_56); -x_57 = lean_ctor_get(x_55, 1); -lean_inc(x_57); -lean_dec(x_55); -x_58 = l_List_getLastD___rarg(x_56, x_11); -x_59 = 0; -x_60 = l_Lean_Syntax_getPos_x3f(x_11, x_59); +lean_dec(x_54); +x_57 = l_List_getLastD___rarg(x_55, x_11); +x_58 = 0; +x_59 = l_Lean_Syntax_getPos_x3f(x_11, x_58); lean_dec(x_11); -x_61 = l_Lean_Syntax_getTailPos_x3f(x_58, x_59); -x_62 = 1; -x_63 = l_Lean_Server_FileWorker_handleFoldingRange_addRange(x_1, x_62, x_60, x_61, x_4, x_5, x_6); -lean_dec(x_60); -x_64 = lean_ctor_get(x_63, 0); +x_60 = l_Lean_Syntax_getTailPos_x3f(x_57, x_58); +x_61 = 1; +x_62 = l_Lean_Server_FileWorker_handleFoldingRange_addRange(x_1, x_61, x_59, x_60, x_4, x_5, x_6); +lean_dec(x_59); +x_63 = lean_ctor_get(x_62, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_62, 1); lean_inc(x_64); +lean_dec(x_62); x_65 = lean_ctor_get(x_63, 1); lean_inc(x_65); lean_dec(x_63); -x_66 = lean_ctor_get(x_64, 1); -lean_inc(x_66); -lean_dec(x_64); -x_3 = x_57; -x_4 = x_66; -x_6 = x_65; +x_3 = x_56; +x_4 = x_65; +x_6 = x_64; goto _start; } } else { -lean_object* x_68; lean_object* x_69; uint8_t x_70; -x_68 = lean_unsigned_to_nat(4u); -x_69 = l_Lean_Syntax_getArg(x_11, x_68); -x_70 = l_Lean_Syntax_isNodeOf(x_69, x_45, x_46); -if (x_70 == 0) +lean_object* x_67; lean_object* x_68; uint8_t x_69; +x_67 = lean_unsigned_to_nat(4u); +x_68 = l_Lean_Syntax_getArg(x_11, x_67); +x_69 = l_Lean_Syntax_matchesNull(x_68, x_45); +if (x_69 == 0) { -uint8_t x_71; +uint8_t x_70; lean_dec(x_42); lean_inc(x_11); -x_71 = l_Lean_Server_FileWorker_handleFoldingRange_isImport(x_11); -if (x_71 == 0) +x_70 = l_Lean_Server_FileWorker_handleFoldingRange_isImport(x_11); +if (x_70 == 0) { -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_72 = l_Lean_Server_FileWorker_handleFoldingRange_addCommandRange(x_1, x_11, x_4, x_5, x_6); -x_73 = lean_ctor_get(x_72, 0); +lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_71 = l_Lean_Server_FileWorker_handleFoldingRange_addCommandRange(x_1, x_11, x_4, x_5, x_6); +x_72 = lean_ctor_get(x_71, 0); +lean_inc(x_72); +x_73 = lean_ctor_get(x_71, 1); lean_inc(x_73); +lean_dec(x_71); x_74 = lean_ctor_get(x_72, 1); lean_inc(x_74); lean_dec(x_72); -x_75 = lean_ctor_get(x_73, 1); -lean_inc(x_75); -lean_dec(x_73); x_3 = x_12; -x_4 = x_75; -x_6 = x_74; +x_4 = x_74; +x_6 = x_73; goto _start; } else { -lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; uint8_t x_82; lean_object* x_83; lean_object* x_84; uint8_t x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_77 = lean_box(0); -x_78 = l_List_spanAux___at_Lean_Server_FileWorker_handleFoldingRange_addRanges___spec__1(x_12, x_77); -x_79 = lean_ctor_get(x_78, 0); +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; uint8_t x_81; lean_object* x_82; lean_object* x_83; uint8_t x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_76 = lean_box(0); +x_77 = l_List_spanAux___at_Lean_Server_FileWorker_handleFoldingRange_addRanges___spec__1(x_12, x_76); +x_78 = lean_ctor_get(x_77, 0); +lean_inc(x_78); +x_79 = lean_ctor_get(x_77, 1); lean_inc(x_79); -x_80 = lean_ctor_get(x_78, 1); -lean_inc(x_80); -lean_dec(x_78); -x_81 = l_List_getLastD___rarg(x_79, x_11); -x_82 = 0; -x_83 = l_Lean_Syntax_getPos_x3f(x_11, x_82); +lean_dec(x_77); +x_80 = l_List_getLastD___rarg(x_78, x_11); +x_81 = 0; +x_82 = l_Lean_Syntax_getPos_x3f(x_11, x_81); lean_dec(x_11); -x_84 = l_Lean_Syntax_getTailPos_x3f(x_81, x_82); -x_85 = 1; -x_86 = l_Lean_Server_FileWorker_handleFoldingRange_addRange(x_1, x_85, x_83, x_84, x_4, x_5, x_6); -lean_dec(x_83); -x_87 = lean_ctor_get(x_86, 0); +x_83 = l_Lean_Syntax_getTailPos_x3f(x_80, x_81); +x_84 = 1; +x_85 = l_Lean_Server_FileWorker_handleFoldingRange_addRange(x_1, x_84, x_82, x_83, x_4, x_5, x_6); +lean_dec(x_82); +x_86 = lean_ctor_get(x_85, 0); +lean_inc(x_86); +x_87 = lean_ctor_get(x_85, 1); lean_inc(x_87); +lean_dec(x_85); x_88 = lean_ctor_get(x_86, 1); lean_inc(x_88); lean_dec(x_86); -x_89 = lean_ctor_get(x_87, 1); -lean_inc(x_89); -lean_dec(x_87); -x_3 = x_80; -x_4 = x_89; -x_6 = x_88; +x_3 = x_79; +x_4 = x_88; +x_6 = x_87; goto _start; } } else { -lean_object* x_91; uint8_t x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; -x_91 = l_Lean_Syntax_getArgs(x_42); +lean_object* x_90; uint8_t x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_90 = l_Lean_Syntax_getArgs(x_42); lean_dec(x_42); -x_92 = 2; -x_93 = l_Lean_Server_FileWorker_handleFoldingRange_addRangeFromSyntax(x_1, x_92, x_11, x_4, x_5, x_6); -x_94 = lean_ctor_get(x_93, 0); +x_91 = 2; +x_92 = l_Lean_Server_FileWorker_handleFoldingRange_addRangeFromSyntax(x_1, x_91, x_11, x_4, x_5, x_6); +x_93 = lean_ctor_get(x_92, 0); +lean_inc(x_93); +x_94 = lean_ctor_get(x_92, 1); lean_inc(x_94); +lean_dec(x_92); x_95 = lean_ctor_get(x_93, 1); lean_inc(x_95); lean_dec(x_93); -x_96 = lean_ctor_get(x_94, 1); -lean_inc(x_96); -lean_dec(x_94); -x_97 = lean_box(0); -x_98 = lean_array_to_list(lean_box(0), x_91); -x_99 = l_Lean_Server_FileWorker_handleFoldingRange_addRanges(x_1, x_97, x_98, x_96, x_5, x_95); -x_100 = lean_ctor_get(x_99, 0); +x_96 = lean_box(0); +x_97 = lean_array_to_list(lean_box(0), x_90); +x_98 = l_Lean_Server_FileWorker_handleFoldingRange_addRanges(x_1, x_96, x_97, x_95, x_5, x_94); +x_99 = lean_ctor_get(x_98, 0); +lean_inc(x_99); +x_100 = lean_ctor_get(x_98, 1); lean_inc(x_100); +lean_dec(x_98); x_101 = lean_ctor_get(x_99, 1); lean_inc(x_101); lean_dec(x_99); -x_102 = lean_ctor_get(x_100, 1); -lean_inc(x_102); -lean_dec(x_100); x_3 = x_12; -x_4 = x_102; -x_6 = x_101; +x_4 = x_101; +x_6 = x_100; goto _start; } } @@ -16196,41 +16212,41 @@ goto _start; } else { -lean_object* x_105; lean_object* x_106; uint8_t x_107; lean_object* x_108; uint8_t x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; -x_105 = lean_ctor_get(x_2, 0); +lean_object* x_104; lean_object* x_105; uint8_t x_106; lean_object* x_107; uint8_t x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_104 = lean_ctor_get(x_2, 0); +lean_inc(x_104); +x_105 = lean_ctor_get(x_2, 1); lean_inc(x_105); -x_106 = lean_ctor_get(x_2, 1); -lean_inc(x_106); lean_dec(x_2); -x_107 = 0; -x_108 = l_Lean_Syntax_getTailPos_x3f(x_11, x_107); -x_109 = 2; -x_110 = l_Lean_Server_FileWorker_handleFoldingRange_addRange(x_1, x_109, x_105, x_108, x_4, x_5, x_6); -lean_dec(x_105); -x_111 = lean_ctor_get(x_110, 0); +x_106 = 0; +x_107 = l_Lean_Syntax_getTailPos_x3f(x_11, x_106); +x_108 = 2; +x_109 = l_Lean_Server_FileWorker_handleFoldingRange_addRange(x_1, x_108, x_104, x_107, x_4, x_5, x_6); +lean_dec(x_104); +x_110 = lean_ctor_get(x_109, 0); +lean_inc(x_110); +x_111 = lean_ctor_get(x_109, 1); lean_inc(x_111); +lean_dec(x_109); x_112 = lean_ctor_get(x_110, 1); lean_inc(x_112); lean_dec(x_110); -x_113 = lean_ctor_get(x_111, 1); -lean_inc(x_113); -lean_dec(x_111); -x_2 = x_106; +x_2 = x_105; x_3 = x_12; -x_4 = x_113; -x_6 = x_112; +x_4 = x_112; +x_6 = x_111; goto _start; } } } else { -uint8_t x_115; lean_object* x_116; -x_115 = 0; -x_116 = l_Lean_Syntax_getPos_x3f(x_11, x_115); +uint8_t x_114; lean_object* x_115; +x_114 = 0; +x_115 = l_Lean_Syntax_getPos_x3f(x_11, x_114); lean_dec(x_11); lean_ctor_set(x_3, 1, x_2); -lean_ctor_set(x_3, 0, x_116); +lean_ctor_set(x_3, 0, x_115); x_2 = x_3; x_3 = x_12; goto _start; @@ -16238,12 +16254,12 @@ goto _start; } else { -uint8_t x_118; lean_object* x_119; -x_118 = 0; -x_119 = l_Lean_Syntax_getPos_x3f(x_11, x_118); +uint8_t x_117; lean_object* x_118; +x_117 = 0; +x_118 = l_Lean_Syntax_getPos_x3f(x_11, x_117); lean_dec(x_11); lean_ctor_set(x_3, 1, x_2); -lean_ctor_set(x_3, 0, x_119); +lean_ctor_set(x_3, 0, x_118); x_2 = x_3; x_3 = x_12; goto _start; @@ -16251,242 +16267,241 @@ goto _start; } else { -lean_object* x_121; lean_object* x_122; lean_object* x_123; uint8_t x_124; -x_121 = lean_ctor_get(x_3, 0); -x_122 = lean_ctor_get(x_3, 1); -lean_inc(x_122); +lean_object* x_120; lean_object* x_121; lean_object* x_122; uint8_t x_123; +x_120 = lean_ctor_get(x_3, 0); +x_121 = lean_ctor_get(x_3, 1); lean_inc(x_121); +lean_inc(x_120); lean_dec(x_3); -x_123 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___closed__5; -lean_inc(x_121); -x_124 = l_Lean_Syntax_isOfKind(x_121, x_123); -if (x_124 == 0) +x_122 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___closed__5; +lean_inc(x_120); +x_123 = l_Lean_Syntax_isOfKind(x_120, x_122); +if (x_123 == 0) { -lean_object* x_125; uint8_t x_126; -x_125 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___closed__7; -lean_inc(x_121); -x_126 = l_Lean_Syntax_isOfKind(x_121, x_125); -if (x_126 == 0) +lean_object* x_124; uint8_t x_125; +x_124 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___closed__7; +lean_inc(x_120); +x_125 = l_Lean_Syntax_isOfKind(x_120, x_124); +if (x_125 == 0) { -lean_object* x_127; uint8_t x_128; -x_127 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___closed__9; -lean_inc(x_121); -x_128 = l_Lean_Syntax_isOfKind(x_121, x_127); -if (x_128 == 0) +lean_object* x_126; uint8_t x_127; +x_126 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___closed__9; +lean_inc(x_120); +x_127 = l_Lean_Syntax_isOfKind(x_120, x_126); +if (x_127 == 0) { -lean_object* x_129; uint8_t x_130; -x_129 = l_Lean_Server_FileWorker_handleFoldingRange_addRanges___closed__2; -lean_inc(x_121); -x_130 = l_Lean_Syntax_isOfKind(x_121, x_129); -if (x_130 == 0) +lean_object* x_128; uint8_t x_129; +x_128 = l_Lean_Server_FileWorker_handleFoldingRange_addRanges___closed__2; +lean_inc(x_120); +x_129 = l_Lean_Syntax_isOfKind(x_120, x_128); +if (x_129 == 0) { -uint8_t x_131; -lean_inc(x_121); -x_131 = l_Lean_Server_FileWorker_handleFoldingRange_isImport(x_121); -if (x_131 == 0) +uint8_t x_130; +lean_inc(x_120); +x_130 = l_Lean_Server_FileWorker_handleFoldingRange_isImport(x_120); +if (x_130 == 0) { -lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; -x_132 = l_Lean_Server_FileWorker_handleFoldingRange_addCommandRange(x_1, x_121, x_4, x_5, x_6); -x_133 = lean_ctor_get(x_132, 0); +lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; +x_131 = l_Lean_Server_FileWorker_handleFoldingRange_addCommandRange(x_1, x_120, x_4, x_5, x_6); +x_132 = lean_ctor_get(x_131, 0); +lean_inc(x_132); +x_133 = lean_ctor_get(x_131, 1); lean_inc(x_133); +lean_dec(x_131); x_134 = lean_ctor_get(x_132, 1); lean_inc(x_134); lean_dec(x_132); -x_135 = lean_ctor_get(x_133, 1); -lean_inc(x_135); -lean_dec(x_133); -x_3 = x_122; -x_4 = x_135; -x_6 = x_134; +x_3 = x_121; +x_4 = x_134; +x_6 = x_133; goto _start; } else { -lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; uint8_t x_142; lean_object* x_143; lean_object* x_144; uint8_t x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; -x_137 = lean_box(0); -x_138 = l_List_spanAux___at_Lean_Server_FileWorker_handleFoldingRange_addRanges___spec__1(x_122, x_137); -x_139 = lean_ctor_get(x_138, 0); +lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; uint8_t x_141; lean_object* x_142; lean_object* x_143; uint8_t x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; +x_136 = lean_box(0); +x_137 = l_List_spanAux___at_Lean_Server_FileWorker_handleFoldingRange_addRanges___spec__1(x_121, x_136); +x_138 = lean_ctor_get(x_137, 0); +lean_inc(x_138); +x_139 = lean_ctor_get(x_137, 1); lean_inc(x_139); -x_140 = lean_ctor_get(x_138, 1); -lean_inc(x_140); -lean_dec(x_138); -x_141 = l_List_getLastD___rarg(x_139, x_121); -x_142 = 0; -x_143 = l_Lean_Syntax_getPos_x3f(x_121, x_142); -lean_dec(x_121); -x_144 = l_Lean_Syntax_getTailPos_x3f(x_141, x_142); -x_145 = 1; -x_146 = l_Lean_Server_FileWorker_handleFoldingRange_addRange(x_1, x_145, x_143, x_144, x_4, x_5, x_6); -lean_dec(x_143); -x_147 = lean_ctor_get(x_146, 0); +lean_dec(x_137); +x_140 = l_List_getLastD___rarg(x_138, x_120); +x_141 = 0; +x_142 = l_Lean_Syntax_getPos_x3f(x_120, x_141); +lean_dec(x_120); +x_143 = l_Lean_Syntax_getTailPos_x3f(x_140, x_141); +x_144 = 1; +x_145 = l_Lean_Server_FileWorker_handleFoldingRange_addRange(x_1, x_144, x_142, x_143, x_4, x_5, x_6); +lean_dec(x_142); +x_146 = lean_ctor_get(x_145, 0); +lean_inc(x_146); +x_147 = lean_ctor_get(x_145, 1); lean_inc(x_147); +lean_dec(x_145); x_148 = lean_ctor_get(x_146, 1); lean_inc(x_148); lean_dec(x_146); -x_149 = lean_ctor_get(x_147, 1); -lean_inc(x_149); -lean_dec(x_147); -x_3 = x_140; -x_4 = x_149; -x_6 = x_148; +x_3 = x_139; +x_4 = x_148; +x_6 = x_147; goto _start; } } else { -lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; uint8_t x_157; -x_151 = lean_unsigned_to_nat(1u); -x_152 = l_Lean_Syntax_getArg(x_121, x_151); -x_153 = lean_unsigned_to_nat(3u); -x_154 = l_Lean_Syntax_getArg(x_121, x_153); -x_155 = l_Lean_nullKind; -x_156 = lean_unsigned_to_nat(0u); -x_157 = l_Lean_Syntax_isNodeOf(x_154, x_155, x_156); -if (x_157 == 0) +lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; uint8_t x_155; +x_150 = lean_unsigned_to_nat(1u); +x_151 = l_Lean_Syntax_getArg(x_120, x_150); +x_152 = lean_unsigned_to_nat(3u); +x_153 = l_Lean_Syntax_getArg(x_120, x_152); +x_154 = lean_unsigned_to_nat(0u); +x_155 = l_Lean_Syntax_matchesNull(x_153, x_154); +if (x_155 == 0) { -uint8_t x_158; -lean_dec(x_152); -lean_inc(x_121); -x_158 = l_Lean_Server_FileWorker_handleFoldingRange_isImport(x_121); -if (x_158 == 0) +uint8_t x_156; +lean_dec(x_151); +lean_inc(x_120); +x_156 = l_Lean_Server_FileWorker_handleFoldingRange_isImport(x_120); +if (x_156 == 0) { -lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; -x_159 = l_Lean_Server_FileWorker_handleFoldingRange_addCommandRange(x_1, x_121, x_4, x_5, x_6); -x_160 = lean_ctor_get(x_159, 0); +lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; +x_157 = l_Lean_Server_FileWorker_handleFoldingRange_addCommandRange(x_1, x_120, x_4, x_5, x_6); +x_158 = lean_ctor_get(x_157, 0); +lean_inc(x_158); +x_159 = lean_ctor_get(x_157, 1); +lean_inc(x_159); +lean_dec(x_157); +x_160 = lean_ctor_get(x_158, 1); lean_inc(x_160); -x_161 = lean_ctor_get(x_159, 1); -lean_inc(x_161); -lean_dec(x_159); -x_162 = lean_ctor_get(x_160, 1); -lean_inc(x_162); -lean_dec(x_160); -x_3 = x_122; -x_4 = x_162; -x_6 = x_161; +lean_dec(x_158); +x_3 = x_121; +x_4 = x_160; +x_6 = x_159; goto _start; } else { -lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; uint8_t x_169; lean_object* x_170; lean_object* x_171; uint8_t x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; -x_164 = lean_box(0); -x_165 = l_List_spanAux___at_Lean_Server_FileWorker_handleFoldingRange_addRanges___spec__1(x_122, x_164); -x_166 = lean_ctor_get(x_165, 0); -lean_inc(x_166); -x_167 = lean_ctor_get(x_165, 1); -lean_inc(x_167); -lean_dec(x_165); -x_168 = l_List_getLastD___rarg(x_166, x_121); -x_169 = 0; -x_170 = l_Lean_Syntax_getPos_x3f(x_121, x_169); -lean_dec(x_121); -x_171 = l_Lean_Syntax_getTailPos_x3f(x_168, x_169); -x_172 = 1; -x_173 = l_Lean_Server_FileWorker_handleFoldingRange_addRange(x_1, x_172, x_170, x_171, x_4, x_5, x_6); -lean_dec(x_170); -x_174 = lean_ctor_get(x_173, 0); +lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; uint8_t x_167; lean_object* x_168; lean_object* x_169; uint8_t x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; +x_162 = lean_box(0); +x_163 = l_List_spanAux___at_Lean_Server_FileWorker_handleFoldingRange_addRanges___spec__1(x_121, x_162); +x_164 = lean_ctor_get(x_163, 0); +lean_inc(x_164); +x_165 = lean_ctor_get(x_163, 1); +lean_inc(x_165); +lean_dec(x_163); +x_166 = l_List_getLastD___rarg(x_164, x_120); +x_167 = 0; +x_168 = l_Lean_Syntax_getPos_x3f(x_120, x_167); +lean_dec(x_120); +x_169 = l_Lean_Syntax_getTailPos_x3f(x_166, x_167); +x_170 = 1; +x_171 = l_Lean_Server_FileWorker_handleFoldingRange_addRange(x_1, x_170, x_168, x_169, x_4, x_5, x_6); +lean_dec(x_168); +x_172 = lean_ctor_get(x_171, 0); +lean_inc(x_172); +x_173 = lean_ctor_get(x_171, 1); +lean_inc(x_173); +lean_dec(x_171); +x_174 = lean_ctor_get(x_172, 1); lean_inc(x_174); -x_175 = lean_ctor_get(x_173, 1); -lean_inc(x_175); -lean_dec(x_173); -x_176 = lean_ctor_get(x_174, 1); -lean_inc(x_176); -lean_dec(x_174); -x_3 = x_167; -x_4 = x_176; -x_6 = x_175; +lean_dec(x_172); +x_3 = x_165; +x_4 = x_174; +x_6 = x_173; goto _start; } } else { -lean_object* x_178; lean_object* x_179; uint8_t x_180; -x_178 = lean_unsigned_to_nat(4u); -x_179 = l_Lean_Syntax_getArg(x_121, x_178); -x_180 = l_Lean_Syntax_isNodeOf(x_179, x_155, x_156); -if (x_180 == 0) +lean_object* x_176; lean_object* x_177; uint8_t x_178; +x_176 = lean_unsigned_to_nat(4u); +x_177 = l_Lean_Syntax_getArg(x_120, x_176); +x_178 = l_Lean_Syntax_matchesNull(x_177, x_154); +if (x_178 == 0) { -uint8_t x_181; -lean_dec(x_152); -lean_inc(x_121); -x_181 = l_Lean_Server_FileWorker_handleFoldingRange_isImport(x_121); -if (x_181 == 0) +uint8_t x_179; +lean_dec(x_151); +lean_inc(x_120); +x_179 = l_Lean_Server_FileWorker_handleFoldingRange_isImport(x_120); +if (x_179 == 0) { -lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; -x_182 = l_Lean_Server_FileWorker_handleFoldingRange_addCommandRange(x_1, x_121, x_4, x_5, x_6); -x_183 = lean_ctor_get(x_182, 0); +lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; +x_180 = l_Lean_Server_FileWorker_handleFoldingRange_addCommandRange(x_1, x_120, x_4, x_5, x_6); +x_181 = lean_ctor_get(x_180, 0); +lean_inc(x_181); +x_182 = lean_ctor_get(x_180, 1); +lean_inc(x_182); +lean_dec(x_180); +x_183 = lean_ctor_get(x_181, 1); lean_inc(x_183); -x_184 = lean_ctor_get(x_182, 1); -lean_inc(x_184); -lean_dec(x_182); -x_185 = lean_ctor_get(x_183, 1); -lean_inc(x_185); -lean_dec(x_183); -x_3 = x_122; -x_4 = x_185; -x_6 = x_184; +lean_dec(x_181); +x_3 = x_121; +x_4 = x_183; +x_6 = x_182; goto _start; } else { -lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; uint8_t x_192; lean_object* x_193; lean_object* x_194; uint8_t x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; -x_187 = lean_box(0); -x_188 = l_List_spanAux___at_Lean_Server_FileWorker_handleFoldingRange_addRanges___spec__1(x_122, x_187); -x_189 = lean_ctor_get(x_188, 0); -lean_inc(x_189); -x_190 = lean_ctor_get(x_188, 1); -lean_inc(x_190); -lean_dec(x_188); -x_191 = l_List_getLastD___rarg(x_189, x_121); -x_192 = 0; -x_193 = l_Lean_Syntax_getPos_x3f(x_121, x_192); -lean_dec(x_121); -x_194 = l_Lean_Syntax_getTailPos_x3f(x_191, x_192); -x_195 = 1; -x_196 = l_Lean_Server_FileWorker_handleFoldingRange_addRange(x_1, x_195, x_193, x_194, x_4, x_5, x_6); -lean_dec(x_193); -x_197 = lean_ctor_get(x_196, 0); +lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; uint8_t x_190; lean_object* x_191; lean_object* x_192; uint8_t x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; +x_185 = lean_box(0); +x_186 = l_List_spanAux___at_Lean_Server_FileWorker_handleFoldingRange_addRanges___spec__1(x_121, x_185); +x_187 = lean_ctor_get(x_186, 0); +lean_inc(x_187); +x_188 = lean_ctor_get(x_186, 1); +lean_inc(x_188); +lean_dec(x_186); +x_189 = l_List_getLastD___rarg(x_187, x_120); +x_190 = 0; +x_191 = l_Lean_Syntax_getPos_x3f(x_120, x_190); +lean_dec(x_120); +x_192 = l_Lean_Syntax_getTailPos_x3f(x_189, x_190); +x_193 = 1; +x_194 = l_Lean_Server_FileWorker_handleFoldingRange_addRange(x_1, x_193, x_191, x_192, x_4, x_5, x_6); +lean_dec(x_191); +x_195 = lean_ctor_get(x_194, 0); +lean_inc(x_195); +x_196 = lean_ctor_get(x_194, 1); +lean_inc(x_196); +lean_dec(x_194); +x_197 = lean_ctor_get(x_195, 1); lean_inc(x_197); -x_198 = lean_ctor_get(x_196, 1); -lean_inc(x_198); -lean_dec(x_196); -x_199 = lean_ctor_get(x_197, 1); -lean_inc(x_199); -lean_dec(x_197); -x_3 = x_190; -x_4 = x_199; -x_6 = x_198; +lean_dec(x_195); +x_3 = x_188; +x_4 = x_197; +x_6 = x_196; goto _start; } } else { -lean_object* x_201; uint8_t x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; -x_201 = l_Lean_Syntax_getArgs(x_152); -lean_dec(x_152); -x_202 = 2; -x_203 = l_Lean_Server_FileWorker_handleFoldingRange_addRangeFromSyntax(x_1, x_202, x_121, x_4, x_5, x_6); -x_204 = lean_ctor_get(x_203, 0); +lean_object* x_199; uint8_t x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; +x_199 = l_Lean_Syntax_getArgs(x_151); +lean_dec(x_151); +x_200 = 2; +x_201 = l_Lean_Server_FileWorker_handleFoldingRange_addRangeFromSyntax(x_1, x_200, x_120, x_4, x_5, x_6); +x_202 = lean_ctor_get(x_201, 0); +lean_inc(x_202); +x_203 = lean_ctor_get(x_201, 1); +lean_inc(x_203); +lean_dec(x_201); +x_204 = lean_ctor_get(x_202, 1); lean_inc(x_204); -x_205 = lean_ctor_get(x_203, 1); -lean_inc(x_205); -lean_dec(x_203); -x_206 = lean_ctor_get(x_204, 1); -lean_inc(x_206); -lean_dec(x_204); -x_207 = lean_box(0); -x_208 = lean_array_to_list(lean_box(0), x_201); -x_209 = l_Lean_Server_FileWorker_handleFoldingRange_addRanges(x_1, x_207, x_208, x_206, x_5, x_205); -x_210 = lean_ctor_get(x_209, 0); +lean_dec(x_202); +x_205 = lean_box(0); +x_206 = lean_array_to_list(lean_box(0), x_199); +x_207 = l_Lean_Server_FileWorker_handleFoldingRange_addRanges(x_1, x_205, x_206, x_204, x_5, x_203); +x_208 = lean_ctor_get(x_207, 0); +lean_inc(x_208); +x_209 = lean_ctor_get(x_207, 1); +lean_inc(x_209); +lean_dec(x_207); +x_210 = lean_ctor_get(x_208, 1); lean_inc(x_210); -x_211 = lean_ctor_get(x_209, 1); -lean_inc(x_211); -lean_dec(x_209); -x_212 = lean_ctor_get(x_210, 1); -lean_inc(x_212); -lean_dec(x_210); -x_3 = x_122; -x_4 = x_212; -x_6 = x_211; +lean_dec(x_208); +x_3 = x_121; +x_4 = x_210; +x_6 = x_209; goto _start; } } @@ -16496,64 +16511,64 @@ else { if (lean_obj_tag(x_2) == 0) { -lean_dec(x_121); -x_3 = x_122; +lean_dec(x_120); +x_3 = x_121; goto _start; } else { -lean_object* x_215; lean_object* x_216; uint8_t x_217; lean_object* x_218; uint8_t x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; -x_215 = lean_ctor_get(x_2, 0); -lean_inc(x_215); -x_216 = lean_ctor_get(x_2, 1); -lean_inc(x_216); +lean_object* x_213; lean_object* x_214; uint8_t x_215; lean_object* x_216; uint8_t x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; +x_213 = lean_ctor_get(x_2, 0); +lean_inc(x_213); +x_214 = lean_ctor_get(x_2, 1); +lean_inc(x_214); lean_dec(x_2); -x_217 = 0; -x_218 = l_Lean_Syntax_getTailPos_x3f(x_121, x_217); -x_219 = 2; -x_220 = l_Lean_Server_FileWorker_handleFoldingRange_addRange(x_1, x_219, x_215, x_218, x_4, x_5, x_6); -lean_dec(x_215); -x_221 = lean_ctor_get(x_220, 0); +x_215 = 0; +x_216 = l_Lean_Syntax_getTailPos_x3f(x_120, x_215); +x_217 = 2; +x_218 = l_Lean_Server_FileWorker_handleFoldingRange_addRange(x_1, x_217, x_213, x_216, x_4, x_5, x_6); +lean_dec(x_213); +x_219 = lean_ctor_get(x_218, 0); +lean_inc(x_219); +x_220 = lean_ctor_get(x_218, 1); +lean_inc(x_220); +lean_dec(x_218); +x_221 = lean_ctor_get(x_219, 1); lean_inc(x_221); -x_222 = lean_ctor_get(x_220, 1); -lean_inc(x_222); -lean_dec(x_220); -x_223 = lean_ctor_get(x_221, 1); -lean_inc(x_223); -lean_dec(x_221); -x_2 = x_216; -x_3 = x_122; -x_4 = x_223; -x_6 = x_222; +lean_dec(x_219); +x_2 = x_214; +x_3 = x_121; +x_4 = x_221; +x_6 = x_220; goto _start; } } } else { -uint8_t x_225; lean_object* x_226; lean_object* x_227; -x_225 = 0; -x_226 = l_Lean_Syntax_getPos_x3f(x_121, x_225); -lean_dec(x_121); -x_227 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_227, 0, x_226); -lean_ctor_set(x_227, 1, x_2); -x_2 = x_227; -x_3 = x_122; +uint8_t x_223; lean_object* x_224; lean_object* x_225; +x_223 = 0; +x_224 = l_Lean_Syntax_getPos_x3f(x_120, x_223); +lean_dec(x_120); +x_225 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_225, 0, x_224); +lean_ctor_set(x_225, 1, x_2); +x_2 = x_225; +x_3 = x_121; goto _start; } } else { -uint8_t x_229; lean_object* x_230; lean_object* x_231; -x_229 = 0; -x_230 = l_Lean_Syntax_getPos_x3f(x_121, x_229); -lean_dec(x_121); -x_231 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_231, 0, x_230); -lean_ctor_set(x_231, 1, x_2); -x_2 = x_231; -x_3 = x_122; +uint8_t x_227; lean_object* x_228; lean_object* x_229; +x_227 = 0; +x_228 = l_Lean_Syntax_getPos_x3f(x_120, x_227); +lean_dec(x_120); +x_229 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_229, 0, x_228); +lean_ctor_set(x_229, 1, x_2); +x_2 = x_229; +x_3 = x_121; goto _start; } } @@ -16638,7 +16653,8 @@ lean_inc(x_4); x_5 = lean_ctor_get(x_3, 1); lean_inc(x_5); lean_dec(x_3); -x_6 = l_Lean_Server_FileWorker_EditableDocument_allSnaps(x_4); +x_6 = lean_ctor_get(x_4, 1); +lean_inc(x_6); x_7 = l_Lean_Server_FileWorker_handleFoldingRange___rarg___closed__1; x_8 = l_IO_AsyncList_waitAll___rarg(x_7, x_6, x_5); x_9 = lean_ctor_get(x_8, 0); @@ -16911,7 +16927,7 @@ lean_inc(x_5); x_6 = lean_ctor_get(x_4, 1); lean_inc(x_6); lean_dec(x_4); -x_7 = lean_ctor_get(x_5, 2); +x_7 = lean_ctor_get(x_5, 1); lean_inc(x_7); lean_dec(x_5); x_8 = l_Lean_Server_FileWorker_handleFoldingRange___rarg___closed__1; @@ -17023,7 +17039,7 @@ lean_dec(x_1); return x_4; } } -static lean_object* _init_l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__1() { +static lean_object* _init_l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__1() { _start: { lean_object* x_1; @@ -17031,7 +17047,7 @@ x_1 = lean_mk_string_from_bytes("Cannot parse request params: ", 29); return x_1; } } -static lean_object* _init_l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__2() { +static lean_object* _init_l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__2() { _start: { lean_object* x_1; @@ -17039,7 +17055,7 @@ x_1 = lean_mk_string_from_bytes("\n", 1); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2(lean_object* x_1) { _start: { lean_object* x_2; @@ -17053,10 +17069,10 @@ if (x_3 == 0) 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; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -17076,10 +17092,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -17116,7 +17132,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -17141,11 +17157,11 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -17195,7 +17211,7 @@ return x_11; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__2___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__2___closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -17204,38 +17220,38 @@ x_2 = l_Lean_Json_mkObj(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__2(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__2___closed__1; +x_2 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__2___closed__1; return x_2; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__2___boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__2___boxed), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__3___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__3___closed__1; +x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__3___closed__1; x_2 = lean_alloc_closure((void*)(l_Except_map___rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__3(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; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__3(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__3(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -17254,7 +17270,7 @@ if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; x_11 = lean_ctor_get(x_9, 0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__3___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__3___closed__2; x_13 = l_Task_Priority_default; x_14 = lean_task_map(x_12, x_11, x_13); lean_ctor_set(x_9, 0, x_14); @@ -17268,7 +17284,7 @@ x_16 = lean_ctor_get(x_9, 1); lean_inc(x_16); lean_inc(x_15); lean_dec(x_9); -x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__3___closed__2; +x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__3___closed__2; x_18 = l_Task_Priority_default; x_19 = lean_task_map(x_17, x_15, x_18); x_20 = lean_alloc_ctor(0, 2, 0); @@ -17326,7 +17342,7 @@ return x_28; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___closed__1() { _start: { lean_object* x_1; @@ -17334,28 +17350,28 @@ x_1 = l_Lean_Server_requestHandlers; return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___closed__2() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4(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; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__3), 4, 1); +x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__3), 4, 1); lean_closure_set(x_5, 0, x_1); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_take(x_6, x_4); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); x_9 = lean_ctor_get(x_7, 1); lean_inc(x_9); lean_dec(x_7); -x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___closed__2; +x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___closed__2; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_5); @@ -17381,7 +17397,7 @@ return x_17; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1() { _start: { lean_object* x_1; @@ -17389,7 +17405,7 @@ x_1 = lean_mk_string_from_bytes("Failed to register LSP request handler for '", return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__2() { _start: { lean_object* x_1; @@ -17397,12 +17413,12 @@ x_1 = lean_mk_string_from_bytes("': already registered", 21); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5(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; uint8_t x_7; lean_dec(x_3); -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___closed__1; +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___closed__1; x_6 = lean_st_ref_get(x_5, x_4); x_7 = !lean_is_exclusive(x_6); if (x_7 == 0) @@ -17417,17 +17433,17 @@ if (x_10 == 0) lean_object* x_11; lean_object* x_12; lean_free_object(x_6); x_11 = lean_box(0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4(x_1, x_2, x_11, x_9); +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4(x_1, x_2, x_11, x_9); return x_12; } else { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_dec(x_1); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_14 = lean_string_append(x_13, x_2); lean_dec(x_2); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__2; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__2; x_16 = lean_string_append(x_14, x_15); x_17 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_17, 0, x_16); @@ -17450,17 +17466,17 @@ if (x_20 == 0) { lean_object* x_21; lean_object* x_22; x_21 = lean_box(0); -x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4(x_1, x_2, x_21, x_19); +x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4(x_1, x_2, x_21, x_19); return x_22; } else { 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_dec(x_1); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_24 = lean_string_append(x_23, x_2); lean_dec(x_2); -x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__2; +x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__2; x_26 = lean_string_append(x_24, x_25); x_27 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_27, 0, x_26); @@ -17472,7 +17488,7 @@ return x_28; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___closed__1() { _start: { lean_object* x_1; @@ -17480,7 +17496,7 @@ x_1 = lean_mk_string_from_bytes("': only possible during initialization", 38); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; uint8_t x_6; @@ -17499,10 +17515,10 @@ if (x_7 == 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; x_8 = lean_ctor_get(x_4, 0); lean_dec(x_8); -x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_10 = lean_string_append(x_9, x_1); lean_dec(x_1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___closed__1; x_12 = lean_string_append(x_10, x_11); x_13 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_13, 0, x_12); @@ -17516,10 +17532,10 @@ lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean x_14 = lean_ctor_get(x_4, 1); lean_inc(x_14); lean_dec(x_4); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_16 = lean_string_append(x_15, x_1); lean_dec(x_1); -x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___closed__1; +x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___closed__1; x_18 = lean_string_append(x_16, x_17); x_19 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_19, 0, x_18); @@ -17536,16 +17552,16 @@ x_21 = lean_ctor_get(x_4, 1); lean_inc(x_21); lean_dec(x_4); x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5(x_2, x_1, x_22, x_21); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5(x_2, x_1, x_22, x_21); return x_23; } } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__5(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__5(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1552_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1551_(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -17555,10 +17571,10 @@ if (x_3 == 0) 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; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -17578,10 +17594,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -17618,7 +17634,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -17643,11 +17659,11 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__5(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__5(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -17697,30 +17713,30 @@ return x_11; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__2___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__2___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1501_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1500_), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__2___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__2___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__2___closed__1; +x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__2___closed__1; x_2 = lean_alloc_closure((void*)(l_Except_map___rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__2(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; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__5(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__6(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__5(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__6(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -17739,7 +17755,7 @@ if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; x_11 = lean_ctor_get(x_9, 0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__2___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__2___closed__2; x_13 = l_Task_Priority_default; x_14 = lean_task_map(x_12, x_11, x_13); lean_ctor_set(x_9, 0, x_14); @@ -17753,7 +17769,7 @@ x_16 = lean_ctor_get(x_9, 1); lean_inc(x_16); lean_inc(x_15); lean_dec(x_9); -x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__2___closed__2; +x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__2___closed__2; x_18 = l_Task_Priority_default; x_19 = lean_task_map(x_17, x_15, x_18); x_20 = lean_alloc_ctor(0, 2, 0); @@ -17811,28 +17827,28 @@ return x_28; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__3(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; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__2), 4, 1); +x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__2), 4, 1); lean_closure_set(x_5, 0, x_1); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_take(x_6, x_4); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); x_9 = lean_ctor_get(x_7, 1); lean_inc(x_9); lean_dec(x_7); -x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__3___closed__1; +x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__3___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_5); @@ -17858,12 +17874,12 @@ return x_17; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__4(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; uint8_t x_7; lean_dec(x_3); -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___closed__1; +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___closed__1; x_6 = lean_st_ref_get(x_5, x_4); x_7 = !lean_is_exclusive(x_6); if (x_7 == 0) @@ -17878,17 +17894,17 @@ if (x_10 == 0) lean_object* x_11; lean_object* x_12; lean_free_object(x_6); x_11 = lean_box(0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__3(x_1, x_2, x_11, x_9); +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__3(x_1, x_2, x_11, x_9); return x_12; } else { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_dec(x_1); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_14 = lean_string_append(x_13, x_2); lean_dec(x_2); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__2; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__2; x_16 = lean_string_append(x_14, x_15); x_17 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_17, 0, x_16); @@ -17911,17 +17927,17 @@ if (x_20 == 0) { lean_object* x_21; lean_object* x_22; x_21 = lean_box(0); -x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__3(x_1, x_2, x_21, x_19); +x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__3(x_1, x_2, x_21, x_19); return x_22; } else { 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_dec(x_1); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_24 = lean_string_append(x_23, x_2); lean_dec(x_2); -x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__2; +x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__2; x_26 = lean_string_append(x_24, x_25); x_27 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_27, 0, x_26); @@ -17933,7 +17949,7 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; uint8_t x_6; @@ -17952,10 +17968,10 @@ if (x_7 == 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; x_8 = lean_ctor_get(x_4, 0); lean_dec(x_8); -x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_10 = lean_string_append(x_9, x_1); lean_dec(x_1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___closed__1; x_12 = lean_string_append(x_10, x_11); x_13 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_13, 0, x_12); @@ -17969,10 +17985,10 @@ lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean x_14 = lean_ctor_get(x_4, 1); lean_inc(x_14); lean_dec(x_4); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_16 = lean_string_append(x_15, x_1); lean_dec(x_1); -x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___closed__1; +x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___closed__1; x_18 = lean_string_append(x_16, x_17); x_19 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_19, 0, x_18); @@ -17989,16 +18005,16 @@ x_21 = lean_ctor_get(x_4, 1); lean_inc(x_21); lean_dec(x_4); x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__4(x_2, x_1, x_22, x_21); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__4(x_2, x_1, x_22, x_21); return x_23; } } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__8(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__8(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1761_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1760_(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -18008,10 +18024,10 @@ if (x_3 == 0) 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; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -18031,10 +18047,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -18071,7 +18087,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -18096,11 +18112,11 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__8(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__8(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -18150,7 +18166,7 @@ return x_11; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__2(lean_object* x_1) { _start: { if (lean_obj_tag(x_1) == 0) @@ -18165,35 +18181,35 @@ lean_object* x_3; lean_object* x_4; x_3 = lean_ctor_get(x_1, 0); lean_inc(x_3); lean_dec(x_1); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1665_(x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1664_(x_3); return x_4; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__2), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__2), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__3___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__3___closed__1; +x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__3___closed__1; x_2 = lean_alloc_closure((void*)(l_Except_map___rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__3(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; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__8(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__9(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__8(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__9(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -18212,7 +18228,7 @@ if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; x_11 = lean_ctor_get(x_9, 0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__3___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__3___closed__2; x_13 = l_Task_Priority_default; x_14 = lean_task_map(x_12, x_11, x_13); lean_ctor_set(x_9, 0, x_14); @@ -18226,7 +18242,7 @@ x_16 = lean_ctor_get(x_9, 1); lean_inc(x_16); lean_inc(x_15); lean_dec(x_9); -x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__3___closed__2; +x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__3___closed__2; x_18 = l_Task_Priority_default; x_19 = lean_task_map(x_17, x_15, x_18); x_20 = lean_alloc_ctor(0, 2, 0); @@ -18284,28 +18300,28 @@ return x_28; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__4___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__4___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__4(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; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__3), 4, 1); +x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__3), 4, 1); lean_closure_set(x_5, 0, x_1); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_take(x_6, x_4); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); x_9 = lean_ctor_get(x_7, 1); lean_inc(x_9); lean_dec(x_7); -x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__4___closed__1; +x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__4___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_5); @@ -18331,12 +18347,12 @@ return x_17; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__5(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; uint8_t x_7; lean_dec(x_3); -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___closed__1; +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___closed__1; x_6 = lean_st_ref_get(x_5, x_4); x_7 = !lean_is_exclusive(x_6); if (x_7 == 0) @@ -18351,17 +18367,17 @@ if (x_10 == 0) lean_object* x_11; lean_object* x_12; lean_free_object(x_6); x_11 = lean_box(0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__4(x_1, x_2, x_11, x_9); +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__4(x_1, x_2, x_11, x_9); return x_12; } else { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_dec(x_1); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_14 = lean_string_append(x_13, x_2); lean_dec(x_2); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__2; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__2; x_16 = lean_string_append(x_14, x_15); x_17 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_17, 0, x_16); @@ -18384,17 +18400,17 @@ if (x_20 == 0) { lean_object* x_21; lean_object* x_22; x_21 = lean_box(0); -x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__4(x_1, x_2, x_21, x_19); +x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__4(x_1, x_2, x_21, x_19); return x_22; } else { 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_dec(x_1); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_24 = lean_string_append(x_23, x_2); lean_dec(x_2); -x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__2; +x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__2; x_26 = lean_string_append(x_24, x_25); x_27 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_27, 0, x_26); @@ -18406,7 +18422,7 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; uint8_t x_6; @@ -18425,10 +18441,10 @@ if (x_7 == 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; x_8 = lean_ctor_get(x_4, 0); lean_dec(x_8); -x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_10 = lean_string_append(x_9, x_1); lean_dec(x_1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___closed__1; x_12 = lean_string_append(x_10, x_11); x_13 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_13, 0, x_12); @@ -18442,10 +18458,10 @@ lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean x_14 = lean_ctor_get(x_4, 1); lean_inc(x_14); lean_dec(x_4); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_16 = lean_string_append(x_15, x_1); lean_dec(x_1); -x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___closed__1; +x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___closed__1; x_18 = lean_string_append(x_16, x_17); x_19 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_19, 0, x_18); @@ -18462,12 +18478,12 @@ x_21 = lean_ctor_get(x_4, 1); lean_inc(x_21); lean_dec(x_4); x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__5(x_2, x_1, x_22, x_21); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__5(x_2, x_1, x_22, x_21); return x_23; } } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__11(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__11(lean_object* x_1) { _start: { lean_object* x_2; @@ -18481,10 +18497,10 @@ if (x_3 == 0) 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; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -18504,10 +18520,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -18544,7 +18560,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -18569,7 +18585,7 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__13(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__13(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -18594,11 +18610,11 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__11(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__11(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -18648,7 +18664,7 @@ return x_11; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__2(lean_object* x_1) { _start: { lean_object* x_2; size_t x_3; size_t x_4; lean_object* x_5; lean_object* x_6; @@ -18656,36 +18672,36 @@ x_2 = lean_array_get_size(x_1); x_3 = lean_usize_of_nat(x_2); lean_dec(x_2); x_4 = 0; -x_5 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__13(x_3, x_4, x_1); +x_5 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__13(x_3, x_4, x_1); x_6 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_6, 0, x_5); return x_6; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__2), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__2), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__3___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__3___closed__1; +x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__3___closed__1; x_2 = lean_alloc_closure((void*)(l_Except_map___rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__3(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; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__11(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__12(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__11(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__12(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -18704,7 +18720,7 @@ if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; x_11 = lean_ctor_get(x_9, 0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__3___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__3___closed__2; x_13 = l_Task_Priority_default; x_14 = lean_task_map(x_12, x_11, x_13); lean_ctor_set(x_9, 0, x_14); @@ -18718,7 +18734,7 @@ x_16 = lean_ctor_get(x_9, 1); lean_inc(x_16); lean_inc(x_15); lean_dec(x_9); -x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__3___closed__2; +x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__3___closed__2; x_18 = l_Task_Priority_default; x_19 = lean_task_map(x_17, x_15, x_18); x_20 = lean_alloc_ctor(0, 2, 0); @@ -18776,28 +18792,28 @@ return x_28; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__4___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__4___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__4(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; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__3), 4, 1); +x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__3), 4, 1); lean_closure_set(x_5, 0, x_1); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_take(x_6, x_4); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); x_9 = lean_ctor_get(x_7, 1); lean_inc(x_9); lean_dec(x_7); -x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__4___closed__1; +x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__4___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_5); @@ -18823,12 +18839,12 @@ return x_17; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__5(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; uint8_t x_7; lean_dec(x_3); -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___closed__1; +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___closed__1; x_6 = lean_st_ref_get(x_5, x_4); x_7 = !lean_is_exclusive(x_6); if (x_7 == 0) @@ -18843,17 +18859,17 @@ if (x_10 == 0) lean_object* x_11; lean_object* x_12; lean_free_object(x_6); x_11 = lean_box(0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__4(x_1, x_2, x_11, x_9); +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__4(x_1, x_2, x_11, x_9); return x_12; } else { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_dec(x_1); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_14 = lean_string_append(x_13, x_2); lean_dec(x_2); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__2; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__2; x_16 = lean_string_append(x_14, x_15); x_17 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_17, 0, x_16); @@ -18876,17 +18892,17 @@ if (x_20 == 0) { lean_object* x_21; lean_object* x_22; x_21 = lean_box(0); -x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__4(x_1, x_2, x_21, x_19); +x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__4(x_1, x_2, x_21, x_19); return x_22; } else { 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_dec(x_1); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_24 = lean_string_append(x_23, x_2); lean_dec(x_2); -x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__2; +x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__2; x_26 = lean_string_append(x_24, x_25); x_27 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_27, 0, x_26); @@ -18898,7 +18914,7 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; uint8_t x_6; @@ -18917,10 +18933,10 @@ if (x_7 == 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; x_8 = lean_ctor_get(x_4, 0); lean_dec(x_8); -x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_10 = lean_string_append(x_9, x_1); lean_dec(x_1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___closed__1; x_12 = lean_string_append(x_10, x_11); x_13 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_13, 0, x_12); @@ -18934,10 +18950,10 @@ lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean x_14 = lean_ctor_get(x_4, 1); lean_inc(x_14); lean_dec(x_4); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_16 = lean_string_append(x_15, x_1); lean_dec(x_1); -x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___closed__1; +x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___closed__1; x_18 = lean_string_append(x_16, x_17); x_19 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_19, 0, x_18); @@ -18954,16 +18970,16 @@ x_21 = lean_ctor_get(x_4, 1); lean_inc(x_21); lean_dec(x_4); x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__5(x_2, x_1, x_22, x_21); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__5(x_2, x_1, x_22, x_21); return x_23; } } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__15(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__15(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2462_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2461_(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -18973,10 +18989,10 @@ if (x_3 == 0) 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; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -18996,10 +19012,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -19036,7 +19052,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__16(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__16(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -19061,7 +19077,7 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__17(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__17(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -19076,7 +19092,7 @@ lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x x_5 = lean_array_uget(x_3, x_2); x_6 = lean_unsigned_to_nat(0u); x_7 = lean_array_uset(x_3, x_2, x_6); -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2600_(x_5); +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2599_(x_5); x_9 = 1; x_10 = lean_usize_add(x_2, x_9); x_11 = lean_array_uset(x_7, x_2, x_8); @@ -19086,11 +19102,11 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__15(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__15(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -19140,7 +19156,7 @@ return x_11; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__2(lean_object* x_1) { _start: { lean_object* x_2; size_t x_3; size_t x_4; lean_object* x_5; lean_object* x_6; @@ -19148,36 +19164,36 @@ x_2 = lean_array_get_size(x_1); x_3 = lean_usize_of_nat(x_2); lean_dec(x_2); x_4 = 0; -x_5 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__17(x_3, x_4, x_1); +x_5 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__17(x_3, x_4, x_1); x_6 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_6, 0, x_5); return x_6; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__2), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__2), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__3___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__3___closed__1; +x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__3___closed__1; x_2 = lean_alloc_closure((void*)(l_Except_map___rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__3(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; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__15(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__16(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__15(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__16(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -19196,7 +19212,7 @@ if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; x_11 = lean_ctor_get(x_9, 0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__3___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__3___closed__2; x_13 = l_Task_Priority_default; x_14 = lean_task_map(x_12, x_11, x_13); lean_ctor_set(x_9, 0, x_14); @@ -19210,7 +19226,7 @@ x_16 = lean_ctor_get(x_9, 1); lean_inc(x_16); lean_inc(x_15); lean_dec(x_9); -x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__3___closed__2; +x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__3___closed__2; x_18 = l_Task_Priority_default; x_19 = lean_task_map(x_17, x_15, x_18); x_20 = lean_alloc_ctor(0, 2, 0); @@ -19268,28 +19284,28 @@ return x_28; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__4___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__4___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__4(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; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__3), 4, 1); +x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__3), 4, 1); lean_closure_set(x_5, 0, x_1); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_take(x_6, x_4); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); x_9 = lean_ctor_get(x_7, 1); lean_inc(x_9); lean_dec(x_7); -x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__4___closed__1; +x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__4___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_5); @@ -19315,12 +19331,12 @@ return x_17; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__5(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; uint8_t x_7; lean_dec(x_3); -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___closed__1; +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___closed__1; x_6 = lean_st_ref_get(x_5, x_4); x_7 = !lean_is_exclusive(x_6); if (x_7 == 0) @@ -19335,17 +19351,17 @@ if (x_10 == 0) lean_object* x_11; lean_object* x_12; lean_free_object(x_6); x_11 = lean_box(0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__4(x_1, x_2, x_11, x_9); +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__4(x_1, x_2, x_11, x_9); return x_12; } else { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_dec(x_1); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_14 = lean_string_append(x_13, x_2); lean_dec(x_2); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__2; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__2; x_16 = lean_string_append(x_14, x_15); x_17 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_17, 0, x_16); @@ -19368,17 +19384,17 @@ if (x_20 == 0) { lean_object* x_21; lean_object* x_22; x_21 = lean_box(0); -x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__4(x_1, x_2, x_21, x_19); +x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__4(x_1, x_2, x_21, x_19); return x_22; } else { 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_dec(x_1); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_24 = lean_string_append(x_23, x_2); lean_dec(x_2); -x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__2; +x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__2; x_26 = lean_string_append(x_24, x_25); x_27 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_27, 0, x_26); @@ -19390,7 +19406,7 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; uint8_t x_6; @@ -19409,10 +19425,10 @@ if (x_7 == 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; x_8 = lean_ctor_get(x_4, 0); lean_dec(x_8); -x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_10 = lean_string_append(x_9, x_1); lean_dec(x_1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___closed__1; x_12 = lean_string_append(x_10, x_11); x_13 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_13, 0, x_12); @@ -19426,10 +19442,10 @@ lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean x_14 = lean_ctor_get(x_4, 1); lean_inc(x_14); lean_dec(x_4); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_16 = lean_string_append(x_15, x_1); lean_dec(x_1); -x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___closed__1; +x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___closed__1; x_18 = lean_string_append(x_16, x_17); x_19 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_19, 0, x_18); @@ -19446,16 +19462,16 @@ x_21 = lean_ctor_get(x_4, 1); lean_inc(x_21); lean_dec(x_4); x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__5(x_2, x_1, x_22, x_21); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__5(x_2, x_1, x_22, x_21); return x_23; } } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__19(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__19(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2648_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2647_(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -19465,10 +19481,10 @@ if (x_3 == 0) 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; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -19488,10 +19504,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -19528,7 +19544,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__20(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__20(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -19553,11 +19569,11 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__19(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__19(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -19598,7 +19614,7 @@ return x_8; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__2(lean_object* x_1) { _start: { lean_object* x_2; size_t x_3; size_t x_4; lean_object* x_5; lean_object* x_6; @@ -19612,30 +19628,30 @@ lean_ctor_set(x_6, 0, x_5); return x_6; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__2), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__2), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__3___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__3___closed__1; +x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__3___closed__1; x_2 = lean_alloc_closure((void*)(l_Except_map___rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__3(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; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__19(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__20(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__19(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__20(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -19654,7 +19670,7 @@ if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; x_11 = lean_ctor_get(x_9, 0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__3___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__3___closed__2; x_13 = l_Task_Priority_default; x_14 = lean_task_map(x_12, x_11, x_13); lean_ctor_set(x_9, 0, x_14); @@ -19668,7 +19684,7 @@ x_16 = lean_ctor_get(x_9, 1); lean_inc(x_16); lean_inc(x_15); lean_dec(x_9); -x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__3___closed__2; +x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__3___closed__2; x_18 = l_Task_Priority_default; x_19 = lean_task_map(x_17, x_15, x_18); x_20 = lean_alloc_ctor(0, 2, 0); @@ -19726,28 +19742,28 @@ return x_28; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__4___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__4___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__4(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; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__3), 4, 1); +x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__3), 4, 1); lean_closure_set(x_5, 0, x_1); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_take(x_6, x_4); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); x_9 = lean_ctor_get(x_7, 1); lean_inc(x_9); lean_dec(x_7); -x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__4___closed__1; +x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__4___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_5); @@ -19773,12 +19789,12 @@ return x_17; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__5(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; uint8_t x_7; lean_dec(x_3); -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___closed__1; +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___closed__1; x_6 = lean_st_ref_get(x_5, x_4); x_7 = !lean_is_exclusive(x_6); if (x_7 == 0) @@ -19793,17 +19809,17 @@ if (x_10 == 0) lean_object* x_11; lean_object* x_12; lean_free_object(x_6); x_11 = lean_box(0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__4(x_1, x_2, x_11, x_9); +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__4(x_1, x_2, x_11, x_9); return x_12; } else { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_dec(x_1); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_14 = lean_string_append(x_13, x_2); lean_dec(x_2); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__2; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__2; x_16 = lean_string_append(x_14, x_15); x_17 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_17, 0, x_16); @@ -19826,17 +19842,17 @@ if (x_20 == 0) { lean_object* x_21; lean_object* x_22; x_21 = lean_box(0); -x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__4(x_1, x_2, x_21, x_19); +x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__4(x_1, x_2, x_21, x_19); return x_22; } else { 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_dec(x_1); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_24 = lean_string_append(x_23, x_2); lean_dec(x_2); -x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__2; +x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__2; x_26 = lean_string_append(x_24, x_25); x_27 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_27, 0, x_26); @@ -19848,7 +19864,7 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; uint8_t x_6; @@ -19867,10 +19883,10 @@ if (x_7 == 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; x_8 = lean_ctor_get(x_4, 0); lean_dec(x_8); -x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_10 = lean_string_append(x_9, x_1); lean_dec(x_1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___closed__1; x_12 = lean_string_append(x_10, x_11); x_13 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_13, 0, x_12); @@ -19884,10 +19900,10 @@ lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean x_14 = lean_ctor_get(x_4, 1); lean_inc(x_14); lean_dec(x_4); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_16 = lean_string_append(x_15, x_1); lean_dec(x_1); -x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___closed__1; +x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___closed__1; x_18 = lean_string_append(x_16, x_17); x_19 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_19, 0, x_18); @@ -19904,16 +19920,16 @@ x_21 = lean_ctor_get(x_4, 1); lean_inc(x_21); lean_dec(x_4); x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__5(x_2, x_1, x_22, x_21); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__5(x_2, x_1, x_22, x_21); return x_23; } } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__22(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__22(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3477_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3476_(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -19923,10 +19939,10 @@ if (x_3 == 0) 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; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -19946,10 +19962,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -19986,7 +20002,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__23(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__23(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -20011,11 +20027,11 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__22(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__22(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -20056,30 +20072,30 @@ return x_8; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__2___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__2___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3698_), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__2___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__2___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__2___closed__1; +x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__2___closed__1; x_2 = lean_alloc_closure((void*)(l_Except_map___rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__2(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; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__22(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__23(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__22(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__23(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -20098,7 +20114,7 @@ if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; x_11 = lean_ctor_get(x_9, 0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__2___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__2___closed__2; x_13 = l_Task_Priority_default; x_14 = lean_task_map(x_12, x_11, x_13); lean_ctor_set(x_9, 0, x_14); @@ -20112,7 +20128,7 @@ x_16 = lean_ctor_get(x_9, 1); lean_inc(x_16); lean_inc(x_15); lean_dec(x_9); -x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__2___closed__2; +x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__2___closed__2; x_18 = l_Task_Priority_default; x_19 = lean_task_map(x_17, x_15, x_18); x_20 = lean_alloc_ctor(0, 2, 0); @@ -20170,28 +20186,28 @@ return x_28; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__3(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; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__2), 4, 1); +x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__2), 4, 1); lean_closure_set(x_5, 0, x_1); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_take(x_6, x_4); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); x_9 = lean_ctor_get(x_7, 1); lean_inc(x_9); lean_dec(x_7); -x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__3___closed__1; +x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__3___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_5); @@ -20217,12 +20233,12 @@ return x_17; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__4(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; uint8_t x_7; lean_dec(x_3); -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___closed__1; +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___closed__1; x_6 = lean_st_ref_get(x_5, x_4); x_7 = !lean_is_exclusive(x_6); if (x_7 == 0) @@ -20237,17 +20253,17 @@ if (x_10 == 0) lean_object* x_11; lean_object* x_12; lean_free_object(x_6); x_11 = lean_box(0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__3(x_1, x_2, x_11, x_9); +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__3(x_1, x_2, x_11, x_9); return x_12; } else { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_dec(x_1); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_14 = lean_string_append(x_13, x_2); lean_dec(x_2); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__2; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__2; x_16 = lean_string_append(x_14, x_15); x_17 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_17, 0, x_16); @@ -20270,17 +20286,17 @@ if (x_20 == 0) { lean_object* x_21; lean_object* x_22; x_21 = lean_box(0); -x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__3(x_1, x_2, x_21, x_19); +x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__3(x_1, x_2, x_21, x_19); return x_22; } else { 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_dec(x_1); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_24 = lean_string_append(x_23, x_2); lean_dec(x_2); -x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__2; +x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__2; x_26 = lean_string_append(x_24, x_25); x_27 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_27, 0, x_26); @@ -20292,7 +20308,7 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; uint8_t x_6; @@ -20311,10 +20327,10 @@ if (x_7 == 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; x_8 = lean_ctor_get(x_4, 0); lean_dec(x_8); -x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_10 = lean_string_append(x_9, x_1); lean_dec(x_1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___closed__1; x_12 = lean_string_append(x_10, x_11); x_13 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_13, 0, x_12); @@ -20328,10 +20344,10 @@ lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean x_14 = lean_ctor_get(x_4, 1); lean_inc(x_14); lean_dec(x_4); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_16 = lean_string_append(x_15, x_1); lean_dec(x_1); -x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___closed__1; +x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___closed__1; x_18 = lean_string_append(x_16, x_17); x_19 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_19, 0, x_18); @@ -20348,16 +20364,16 @@ x_21 = lean_ctor_get(x_4, 1); lean_inc(x_21); lean_dec(x_4); x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__4(x_2, x_1, x_22, x_21); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__4(x_2, x_1, x_22, x_21); return x_23; } } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__25(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__25(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3560_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3559_(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -20367,10 +20383,10 @@ if (x_3 == 0) 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; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -20390,10 +20406,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -20430,7 +20446,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__26(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__26(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -20455,11 +20471,11 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__24___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__24___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__25(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__25(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -20509,12 +20525,12 @@ return x_11; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__24___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__24___lambda__2(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; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__25(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__26(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__25(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__26(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -20533,7 +20549,7 @@ if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; x_11 = lean_ctor_get(x_9, 0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__2___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__2___closed__2; x_13 = l_Task_Priority_default; x_14 = lean_task_map(x_12, x_11, x_13); lean_ctor_set(x_9, 0, x_14); @@ -20547,7 +20563,7 @@ x_16 = lean_ctor_get(x_9, 1); lean_inc(x_16); lean_inc(x_15); lean_dec(x_9); -x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__2___closed__2; +x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__2___closed__2; x_18 = l_Task_Priority_default; x_19 = lean_task_map(x_17, x_15, x_18); x_20 = lean_alloc_ctor(0, 2, 0); @@ -20605,28 +20621,28 @@ return x_28; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__24___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__24___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__24___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__24___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__24___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__24___lambda__3(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; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__24___lambda__2), 4, 1); +x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__24___lambda__2), 4, 1); lean_closure_set(x_5, 0, x_1); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_take(x_6, x_4); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); x_9 = lean_ctor_get(x_7, 1); lean_inc(x_9); lean_dec(x_7); -x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__24___lambda__3___closed__1; +x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__24___lambda__3___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_5); @@ -20652,12 +20668,12 @@ return x_17; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__24___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__24___lambda__4(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; uint8_t x_7; lean_dec(x_3); -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___closed__1; +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___closed__1; x_6 = lean_st_ref_get(x_5, x_4); x_7 = !lean_is_exclusive(x_6); if (x_7 == 0) @@ -20672,17 +20688,17 @@ if (x_10 == 0) lean_object* x_11; lean_object* x_12; lean_free_object(x_6); x_11 = lean_box(0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__24___lambda__3(x_1, x_2, x_11, x_9); +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__24___lambda__3(x_1, x_2, x_11, x_9); return x_12; } else { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_dec(x_1); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_14 = lean_string_append(x_13, x_2); lean_dec(x_2); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__2; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__2; x_16 = lean_string_append(x_14, x_15); x_17 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_17, 0, x_16); @@ -20705,17 +20721,17 @@ if (x_20 == 0) { lean_object* x_21; lean_object* x_22; x_21 = lean_box(0); -x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__24___lambda__3(x_1, x_2, x_21, x_19); +x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__24___lambda__3(x_1, x_2, x_21, x_19); return x_22; } else { 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_dec(x_1); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_24 = lean_string_append(x_23, x_2); lean_dec(x_2); -x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__2; +x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__2; x_26 = lean_string_append(x_24, x_25); x_27 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_27, 0, x_26); @@ -20727,7 +20743,7 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__24(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__24(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; uint8_t x_6; @@ -20746,10 +20762,10 @@ if (x_7 == 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; x_8 = lean_ctor_get(x_4, 0); lean_dec(x_8); -x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_10 = lean_string_append(x_9, x_1); lean_dec(x_1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___closed__1; x_12 = lean_string_append(x_10, x_11); x_13 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_13, 0, x_12); @@ -20763,10 +20779,10 @@ lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean x_14 = lean_ctor_get(x_4, 1); lean_inc(x_14); lean_dec(x_4); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_16 = lean_string_append(x_15, x_1); lean_dec(x_1); -x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___closed__1; +x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___closed__1; x_18 = lean_string_append(x_16, x_17); x_19 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_19, 0, x_18); @@ -20783,16 +20799,16 @@ x_21 = lean_ctor_get(x_4, 1); lean_inc(x_21); lean_dec(x_4); x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__24___lambda__4(x_2, x_1, x_22, x_21); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__24___lambda__4(x_2, x_1, x_22, x_21); return x_23; } } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__28(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__28(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3739_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738_(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -20802,10 +20818,10 @@ if (x_3 == 0) 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; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -20825,10 +20841,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -20865,7 +20881,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__29(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__29(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -20890,7 +20906,7 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__30(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__30(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -20905,7 +20921,7 @@ lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x x_5 = lean_array_uget(x_3, x_2); x_6 = lean_unsigned_to_nat(0u); x_7 = lean_array_uset(x_3, x_2, x_6); -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3857_(x_5); +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3856_(x_5); x_9 = 1; x_10 = lean_usize_add(x_2, x_9); x_11 = lean_array_uset(x_7, x_2, x_8); @@ -20915,11 +20931,11 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__28(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__28(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -20960,7 +20976,7 @@ return x_8; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__2(lean_object* x_1) { _start: { lean_object* x_2; size_t x_3; size_t x_4; lean_object* x_5; lean_object* x_6; @@ -20968,36 +20984,36 @@ x_2 = lean_array_get_size(x_1); x_3 = lean_usize_of_nat(x_2); lean_dec(x_2); x_4 = 0; -x_5 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__30(x_3, x_4, x_1); +x_5 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__30(x_3, x_4, x_1); x_6 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_6, 0, x_5); return x_6; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__2), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__2), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__3___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__3___closed__1; +x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__3___closed__1; x_2 = lean_alloc_closure((void*)(l_Except_map___rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__3(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; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__28(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__29(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__28(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__29(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -21016,7 +21032,7 @@ if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; x_11 = lean_ctor_get(x_9, 0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__3___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__3___closed__2; x_13 = l_Task_Priority_default; x_14 = lean_task_map(x_12, x_11, x_13); lean_ctor_set(x_9, 0, x_14); @@ -21030,7 +21046,7 @@ x_16 = lean_ctor_get(x_9, 1); lean_inc(x_16); lean_inc(x_15); lean_dec(x_9); -x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__3___closed__2; +x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__3___closed__2; x_18 = l_Task_Priority_default; x_19 = lean_task_map(x_17, x_15, x_18); x_20 = lean_alloc_ctor(0, 2, 0); @@ -21088,28 +21104,28 @@ return x_28; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__4___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__4___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__4(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; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__3), 4, 1); +x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__3), 4, 1); lean_closure_set(x_5, 0, x_1); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_take(x_6, x_4); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); x_9 = lean_ctor_get(x_7, 1); lean_inc(x_9); lean_dec(x_7); -x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__4___closed__1; +x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__4___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_5); @@ -21135,12 +21151,12 @@ return x_17; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__5(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; uint8_t x_7; lean_dec(x_3); -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___closed__1; +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___closed__1; x_6 = lean_st_ref_get(x_5, x_4); x_7 = !lean_is_exclusive(x_6); if (x_7 == 0) @@ -21155,17 +21171,17 @@ if (x_10 == 0) lean_object* x_11; lean_object* x_12; lean_free_object(x_6); x_11 = lean_box(0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__4(x_1, x_2, x_11, x_9); +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__4(x_1, x_2, x_11, x_9); return x_12; } else { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_dec(x_1); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_14 = lean_string_append(x_13, x_2); lean_dec(x_2); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__2; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__2; x_16 = lean_string_append(x_14, x_15); x_17 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_17, 0, x_16); @@ -21188,17 +21204,17 @@ if (x_20 == 0) { lean_object* x_21; lean_object* x_22; x_21 = lean_box(0); -x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__4(x_1, x_2, x_21, x_19); +x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__4(x_1, x_2, x_21, x_19); return x_22; } else { 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_dec(x_1); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_24 = lean_string_append(x_23, x_2); lean_dec(x_2); -x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__2; +x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__2; x_26 = lean_string_append(x_24, x_25); x_27 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_27, 0, x_26); @@ -21210,7 +21226,7 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; uint8_t x_6; @@ -21229,10 +21245,10 @@ if (x_7 == 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; x_8 = lean_ctor_get(x_4, 0); lean_dec(x_8); -x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_10 = lean_string_append(x_9, x_1); lean_dec(x_1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___closed__1; x_12 = lean_string_append(x_10, x_11); x_13 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_13, 0, x_12); @@ -21246,10 +21262,10 @@ lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean x_14 = lean_ctor_get(x_4, 1); lean_inc(x_14); lean_dec(x_4); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_16 = lean_string_append(x_15, x_1); lean_dec(x_1); -x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___closed__1; +x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___closed__1; x_18 = lean_string_append(x_16, x_17); x_19 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_19, 0, x_18); @@ -21266,12 +21282,12 @@ x_21 = lean_ctor_get(x_4, 1); lean_inc(x_21); lean_dec(x_4); x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__5(x_2, x_1, x_22, x_21); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__5(x_2, x_1, x_22, x_21); return x_23; } } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__32(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__32(lean_object* x_1) { _start: { lean_object* x_2; @@ -21285,10 +21301,10 @@ if (x_3 == 0) 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; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -21308,10 +21324,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -21348,7 +21364,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__33(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__33(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -21373,11 +21389,11 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__32(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__32(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -21427,7 +21443,7 @@ return x_11; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__2(lean_object* x_1) { _start: { if (lean_obj_tag(x_1) == 0) @@ -21447,30 +21463,30 @@ return x_4; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__2), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__2), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__3___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__3___closed__1; +x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__3___closed__1; x_2 = lean_alloc_closure((void*)(l_Except_map___rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__3(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; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__32(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__33(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__32(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__33(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -21489,7 +21505,7 @@ if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; x_11 = lean_ctor_get(x_9, 0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__3___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__3___closed__2; x_13 = l_Task_Priority_default; x_14 = lean_task_map(x_12, x_11, x_13); lean_ctor_set(x_9, 0, x_14); @@ -21503,7 +21519,7 @@ x_16 = lean_ctor_get(x_9, 1); lean_inc(x_16); lean_inc(x_15); lean_dec(x_9); -x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__3___closed__2; +x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__3___closed__2; x_18 = l_Task_Priority_default; x_19 = lean_task_map(x_17, x_15, x_18); x_20 = lean_alloc_ctor(0, 2, 0); @@ -21561,28 +21577,28 @@ return x_28; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__4___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__4___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__4(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; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__3), 4, 1); +x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__3), 4, 1); lean_closure_set(x_5, 0, x_1); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_take(x_6, x_4); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); x_9 = lean_ctor_get(x_7, 1); lean_inc(x_9); lean_dec(x_7); -x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__4___closed__1; +x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__4___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_5); @@ -21608,12 +21624,12 @@ return x_17; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__5(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; uint8_t x_7; lean_dec(x_3); -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___closed__1; +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___closed__1; x_6 = lean_st_ref_get(x_5, x_4); x_7 = !lean_is_exclusive(x_6); if (x_7 == 0) @@ -21628,17 +21644,17 @@ if (x_10 == 0) lean_object* x_11; lean_object* x_12; lean_free_object(x_6); x_11 = lean_box(0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__4(x_1, x_2, x_11, x_9); +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__4(x_1, x_2, x_11, x_9); return x_12; } else { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_dec(x_1); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_14 = lean_string_append(x_13, x_2); lean_dec(x_2); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__2; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__2; x_16 = lean_string_append(x_14, x_15); x_17 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_17, 0, x_16); @@ -21661,17 +21677,17 @@ if (x_20 == 0) { lean_object* x_21; lean_object* x_22; x_21 = lean_box(0); -x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__4(x_1, x_2, x_21, x_19); +x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__4(x_1, x_2, x_21, x_19); return x_22; } else { 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_dec(x_1); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_24 = lean_string_append(x_23, x_2); lean_dec(x_2); -x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__2; +x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__2; x_26 = lean_string_append(x_24, x_25); x_27 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_27, 0, x_26); @@ -21683,7 +21699,7 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; uint8_t x_6; @@ -21702,10 +21718,10 @@ if (x_7 == 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; x_8 = lean_ctor_get(x_4, 0); lean_dec(x_8); -x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_10 = lean_string_append(x_9, x_1); lean_dec(x_1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___closed__1; x_12 = lean_string_append(x_10, x_11); x_13 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_13, 0, x_12); @@ -21719,10 +21735,10 @@ lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean x_14 = lean_ctor_get(x_4, 1); lean_inc(x_14); lean_dec(x_4); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_16 = lean_string_append(x_15, x_1); lean_dec(x_1); -x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___closed__1; +x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___closed__1; x_18 = lean_string_append(x_16, x_17); x_19 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_19, 0, x_18); @@ -21739,12 +21755,12 @@ x_21 = lean_ctor_get(x_4, 1); lean_inc(x_21); lean_dec(x_4); x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__5(x_2, x_1, x_22, x_21); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__5(x_2, x_1, x_22, x_21); return x_23; } } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__35(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__35(lean_object* x_1) { _start: { lean_object* x_2; @@ -21758,10 +21774,10 @@ if (x_3 == 0) 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; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -21781,10 +21797,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -21821,7 +21837,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__36(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__36(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -21846,11 +21862,11 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__35(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__35(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -21900,7 +21916,7 @@ return x_11; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__2(lean_object* x_1) { _start: { if (lean_obj_tag(x_1) == 0) @@ -21920,30 +21936,30 @@ return x_4; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__2), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__2), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__3___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__3___closed__1; +x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__3___closed__1; x_2 = lean_alloc_closure((void*)(l_Except_map___rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__3(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; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__35(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__36(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__35(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__36(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -21962,7 +21978,7 @@ if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; x_11 = lean_ctor_get(x_9, 0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__3___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__3___closed__2; x_13 = l_Task_Priority_default; x_14 = lean_task_map(x_12, x_11, x_13); lean_ctor_set(x_9, 0, x_14); @@ -21976,7 +21992,7 @@ x_16 = lean_ctor_get(x_9, 1); lean_inc(x_16); lean_inc(x_15); lean_dec(x_9); -x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__3___closed__2; +x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__3___closed__2; x_18 = l_Task_Priority_default; x_19 = lean_task_map(x_17, x_15, x_18); x_20 = lean_alloc_ctor(0, 2, 0); @@ -22034,28 +22050,28 @@ return x_28; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__4___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__4___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__4(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; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__3), 4, 1); +x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__3), 4, 1); lean_closure_set(x_5, 0, x_1); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_take(x_6, x_4); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); x_9 = lean_ctor_get(x_7, 1); lean_inc(x_9); lean_dec(x_7); -x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__4___closed__1; +x_10 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__4___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_5); @@ -22081,12 +22097,12 @@ return x_17; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__5(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; uint8_t x_7; lean_dec(x_3); -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___closed__1; +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___closed__1; x_6 = lean_st_ref_get(x_5, x_4); x_7 = !lean_is_exclusive(x_6); if (x_7 == 0) @@ -22101,17 +22117,17 @@ if (x_10 == 0) lean_object* x_11; lean_object* x_12; lean_free_object(x_6); x_11 = lean_box(0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__4(x_1, x_2, x_11, x_9); +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__4(x_1, x_2, x_11, x_9); return x_12; } else { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_dec(x_1); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_14 = lean_string_append(x_13, x_2); lean_dec(x_2); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__2; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__2; x_16 = lean_string_append(x_14, x_15); x_17 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_17, 0, x_16); @@ -22134,17 +22150,17 @@ if (x_20 == 0) { lean_object* x_21; lean_object* x_22; x_21 = lean_box(0); -x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__4(x_1, x_2, x_21, x_19); +x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__4(x_1, x_2, x_21, x_19); return x_22; } else { 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_dec(x_1); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_24 = lean_string_append(x_23, x_2); lean_dec(x_2); -x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__2; +x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__2; x_26 = lean_string_append(x_24, x_25); x_27 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_27, 0, x_26); @@ -22156,7 +22172,7 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; uint8_t x_6; @@ -22175,10 +22191,10 @@ if (x_7 == 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; x_8 = lean_ctor_get(x_4, 0); lean_dec(x_8); -x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_10 = lean_string_append(x_9, x_1); lean_dec(x_1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___closed__1; x_12 = lean_string_append(x_10, x_11); x_13 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_13, 0, x_12); @@ -22192,10 +22208,10 @@ lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean x_14 = lean_ctor_get(x_4, 1); lean_inc(x_14); lean_dec(x_4); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1; x_16 = lean_string_append(x_15, x_1); lean_dec(x_1); -x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___closed__1; +x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___closed__1; x_18 = lean_string_append(x_16, x_17); x_19 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_19, 0, x_18); @@ -22212,12 +22228,12 @@ x_21 = lean_ctor_get(x_4, 1); lean_inc(x_21); lean_dec(x_4); x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__5(x_2, x_1, x_22, x_21); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__5(x_2, x_1, x_22, x_21); return x_23; } } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__1() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__1() { _start: { lean_object* x_1; @@ -22225,7 +22241,7 @@ x_1 = lean_mk_string_from_bytes("textDocument/waitForDiagnostics", 31); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__2() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__2() { _start: { lean_object* x_1; @@ -22233,7 +22249,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleWaitForDiagnosti return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__3() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__3() { _start: { lean_object* x_1; @@ -22241,7 +22257,7 @@ x_1 = lean_mk_string_from_bytes("textDocument/completion", 23); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__4() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__4() { _start: { lean_object* x_1; @@ -22249,7 +22265,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleCompletion), 3, return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__5() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__5() { _start: { lean_object* x_1; @@ -22257,7 +22273,7 @@ x_1 = lean_mk_string_from_bytes("textDocument/hover", 18); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__6() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__6() { _start: { lean_object* x_1; @@ -22265,7 +22281,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleHover), 3, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__7() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__7() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -22276,7 +22292,7 @@ lean_closure_set(x_3, 0, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__8() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__8() { _start: { lean_object* x_1; @@ -22284,7 +22300,7 @@ x_1 = lean_mk_string_from_bytes("textDocument/declaration", 24); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__9() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__9() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -22295,7 +22311,7 @@ lean_closure_set(x_3, 0, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__10() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__10() { _start: { lean_object* x_1; @@ -22303,7 +22319,7 @@ x_1 = lean_mk_string_from_bytes("textDocument/definition", 23); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__11() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__11() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -22314,7 +22330,7 @@ lean_closure_set(x_3, 0, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__12() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__12() { _start: { lean_object* x_1; @@ -22322,7 +22338,7 @@ x_1 = lean_mk_string_from_bytes("textDocument/typeDefinition", 27); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__13() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__13() { _start: { lean_object* x_1; @@ -22330,7 +22346,7 @@ x_1 = lean_mk_string_from_bytes("textDocument/documentHighlight", 30); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__14() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__14() { _start: { lean_object* x_1; @@ -22338,7 +22354,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleDocumentHighligh return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__15() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__15() { _start: { lean_object* x_1; @@ -22346,7 +22362,7 @@ x_1 = lean_mk_string_from_bytes("textDocument/documentSymbol", 27); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__16() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__16() { _start: { lean_object* x_1; @@ -22354,7 +22370,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleDocumentSymbol__ return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__17() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__17() { _start: { lean_object* x_1; @@ -22362,7 +22378,7 @@ x_1 = lean_mk_string_from_bytes("textDocument/semanticTokens/full", 32); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__18() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__18() { _start: { lean_object* x_1; @@ -22370,7 +22386,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleSemanticTokensFu return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__19() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__19() { _start: { lean_object* x_1; @@ -22378,7 +22394,7 @@ x_1 = lean_mk_string_from_bytes("textDocument/semanticTokens/range", 33); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__20() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__20() { _start: { lean_object* x_1; @@ -22386,7 +22402,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleSemanticTokensRa return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__21() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__21() { _start: { lean_object* x_1; @@ -22394,7 +22410,7 @@ x_1 = lean_mk_string_from_bytes("textDocument/foldingRange", 25); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__22() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__22() { _start: { lean_object* x_1; @@ -22402,7 +22418,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleFoldingRange___b return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__23() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__23() { _start: { lean_object* x_1; @@ -22410,7 +22426,7 @@ x_1 = lean_mk_string_from_bytes("$/lean/plainGoal", 16); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__24() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__24() { _start: { lean_object* x_1; @@ -22418,7 +22434,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handlePlainGoal), 3, 0 return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__25() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__25() { _start: { lean_object* x_1; @@ -22426,7 +22442,7 @@ x_1 = lean_mk_string_from_bytes("$/lean/plainTermGoal", 20); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__26() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__26() { _start: { lean_object* x_1; @@ -22434,121 +22450,121 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handlePlainTermGoal), return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__1; -x_3 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__2; -x_4 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1(x_2, x_3, x_1); +x_2 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__1; +x_3 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__2; +x_4 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1(x_2, x_3, x_1); if (lean_obj_tag(x_4) == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; x_5 = lean_ctor_get(x_4, 1); lean_inc(x_5); lean_dec(x_4); -x_6 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__3; -x_7 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__4; -x_8 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4(x_6, x_7, x_5); +x_6 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__3; +x_7 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__4; +x_8 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4(x_6, x_7, x_5); if (lean_obj_tag(x_8) == 0) { lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; x_9 = lean_ctor_get(x_8, 1); lean_inc(x_9); lean_dec(x_8); -x_10 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__5; -x_11 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__6; -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7(x_10, x_11, x_9); +x_10 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__5; +x_11 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__6; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7(x_10, x_11, x_9); if (lean_obj_tag(x_12) == 0) { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; x_13 = lean_ctor_get(x_12, 1); lean_inc(x_13); lean_dec(x_12); -x_14 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__8; -x_15 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__7; -x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10(x_14, x_15, x_13); +x_14 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__8; +x_15 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__7; +x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10(x_14, x_15, x_13); if (lean_obj_tag(x_16) == 0) { lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; x_17 = lean_ctor_get(x_16, 1); lean_inc(x_17); lean_dec(x_16); -x_18 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__10; -x_19 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__9; -x_20 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10(x_18, x_19, x_17); +x_18 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__10; +x_19 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__9; +x_20 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10(x_18, x_19, x_17); if (lean_obj_tag(x_20) == 0) { lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; x_21 = lean_ctor_get(x_20, 1); lean_inc(x_21); lean_dec(x_20); -x_22 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__12; -x_23 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__11; -x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10(x_22, x_23, x_21); +x_22 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__12; +x_23 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__11; +x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10(x_22, x_23, x_21); if (lean_obj_tag(x_24) == 0) { lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; x_25 = lean_ctor_get(x_24, 1); lean_inc(x_25); lean_dec(x_24); -x_26 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__13; -x_27 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__14; -x_28 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14(x_26, x_27, x_25); +x_26 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__13; +x_27 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__14; +x_28 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14(x_26, x_27, x_25); if (lean_obj_tag(x_28) == 0) { lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; x_29 = lean_ctor_get(x_28, 1); lean_inc(x_29); lean_dec(x_28); -x_30 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__15; -x_31 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__16; -x_32 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18(x_30, x_31, x_29); +x_30 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__15; +x_31 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__16; +x_32 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18(x_30, x_31, x_29); if (lean_obj_tag(x_32) == 0) { lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; x_33 = lean_ctor_get(x_32, 1); lean_inc(x_33); lean_dec(x_32); -x_34 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__17; -x_35 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__18; -x_36 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21(x_34, x_35, x_33); +x_34 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__17; +x_35 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__18; +x_36 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21(x_34, x_35, x_33); if (lean_obj_tag(x_36) == 0) { lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; x_37 = lean_ctor_get(x_36, 1); lean_inc(x_37); lean_dec(x_36); -x_38 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__19; -x_39 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__20; -x_40 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__24(x_38, x_39, x_37); +x_38 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__19; +x_39 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__20; +x_40 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__24(x_38, x_39, x_37); if (lean_obj_tag(x_40) == 0) { lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; x_41 = lean_ctor_get(x_40, 1); lean_inc(x_41); lean_dec(x_40); -x_42 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__21; -x_43 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__22; -x_44 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27(x_42, x_43, x_41); +x_42 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__21; +x_43 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__22; +x_44 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27(x_42, x_43, x_41); if (lean_obj_tag(x_44) == 0) { lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; x_45 = lean_ctor_get(x_44, 1); lean_inc(x_45); lean_dec(x_44); -x_46 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__23; -x_47 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__24; -x_48 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31(x_46, x_47, x_45); +x_46 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__23; +x_47 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__24; +x_48 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31(x_46, x_47, x_45); if (lean_obj_tag(x_48) == 0) { lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; x_49 = lean_ctor_get(x_48, 1); lean_inc(x_49); lean_dec(x_48); -x_50 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__25; -x_51 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__26; -x_52 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34(x_50, x_51, x_49); +x_50 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__25; +x_51 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__26; +x_52 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34(x_50, x_51, x_49); return x_52; } else @@ -22828,83 +22844,83 @@ return x_100; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__3(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__3(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__2___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__2___boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__2(x_1); +x_2 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__2(x_1); lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__6(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__6(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__3___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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__3(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__3(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__9(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__9(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__4___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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__4(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__4(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__12(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__12(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; size_t x_5; lean_object* x_6; @@ -22912,30 +22928,30 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__13(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__13(x_4, x_5, x_3); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__4___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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__4(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__4(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__16___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__16___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__16(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__16(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__17___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__17___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; size_t x_5; lean_object* x_6; @@ -22943,87 +22959,87 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__17(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__17(x_4, x_5, x_3); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__4___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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__4(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__4(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__20___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__20___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__20(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__20(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__4___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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__4(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__4(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__23___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__23___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__23(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__23(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__3___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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__3(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__3(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__26___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__26___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__26(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__26(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__24___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__24___lambda__3___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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__24___lambda__3(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__24___lambda__3(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__29___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__29___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__29(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__29(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__30___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__30___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; size_t x_5; lean_object* x_6; @@ -23031,53 +23047,53 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__30(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__30(x_4, x_5, x_3); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__4___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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__4(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__4(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__33___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__33___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__33(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__33(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__4___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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__4(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__4(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__36___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__36___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__36(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__36(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__4___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_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__4(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__4(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } @@ -23271,6 +23287,10 @@ l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___closed__11 = _ lean_mark_persistent(l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___closed__11); l_Lean_Server_FileWorker_handleDocumentSymbol_sectionLikeToDocumentSymbols___closed__1 = _init_l_Lean_Server_FileWorker_handleDocumentSymbol_sectionLikeToDocumentSymbols___closed__1(); lean_mark_persistent(l_Lean_Server_FileWorker_handleDocumentSymbol_sectionLikeToDocumentSymbols___closed__1); +l_Lean_Server_FileWorker_handleDocumentSymbol_sectionLikeToDocumentSymbols___closed__2 = _init_l_Lean_Server_FileWorker_handleDocumentSymbol_sectionLikeToDocumentSymbols___closed__2(); +lean_mark_persistent(l_Lean_Server_FileWorker_handleDocumentSymbol_sectionLikeToDocumentSymbols___closed__2); +l_Lean_Server_FileWorker_handleDocumentSymbol_sectionLikeToDocumentSymbols___closed__3 = _init_l_Lean_Server_FileWorker_handleDocumentSymbol_sectionLikeToDocumentSymbols___closed__3(); +lean_mark_persistent(l_Lean_Server_FileWorker_handleDocumentSymbol_sectionLikeToDocumentSymbols___closed__3); l_Lean_Server_FileWorker_noHighlightKinds___closed__1 = _init_l_Lean_Server_FileWorker_noHighlightKinds___closed__1(); lean_mark_persistent(l_Lean_Server_FileWorker_noHighlightKinds___closed__1); l_Lean_Server_FileWorker_noHighlightKinds___closed__2 = _init_l_Lean_Server_FileWorker_noHighlightKinds___closed__2(); @@ -23309,6 +23329,10 @@ l_Lean_Server_FileWorker_handleSemanticTokens_go___closed__1 = _init_l_Lean_Serv lean_mark_persistent(l_Lean_Server_FileWorker_handleSemanticTokens_go___closed__1); l_Lean_Server_FileWorker_handleSemanticTokens_go___closed__2 = _init_l_Lean_Server_FileWorker_handleSemanticTokens_go___closed__2(); lean_mark_persistent(l_Lean_Server_FileWorker_handleSemanticTokens_go___closed__2); +l_Lean_Server_FileWorker_handleSemanticTokens_go___closed__3 = _init_l_Lean_Server_FileWorker_handleSemanticTokens_go___closed__3(); +lean_mark_persistent(l_Lean_Server_FileWorker_handleSemanticTokens_go___closed__3); +l_Lean_Server_FileWorker_handleSemanticTokens_go___closed__4 = _init_l_Lean_Server_FileWorker_handleSemanticTokens_go___closed__4(); +lean_mark_persistent(l_Lean_Server_FileWorker_handleSemanticTokens_go___closed__4); l_List_forIn_loop___at_Lean_Server_FileWorker_handleSemanticTokens___spec__1___lambda__1___closed__1 = _init_l_List_forIn_loop___at_Lean_Server_FileWorker_handleSemanticTokens___spec__1___lambda__1___closed__1(); lean_mark_persistent(l_List_forIn_loop___at_Lean_Server_FileWorker_handleSemanticTokens___spec__1___lambda__1___closed__1); l_Lean_Server_FileWorker_handleSemanticTokens___lambda__2___closed__1 = _init_l_Lean_Server_FileWorker_handleSemanticTokens___lambda__2___closed__1(); @@ -23343,135 +23367,135 @@ l_Lean_Server_FileWorker_handleWaitForDiagnostics___lambda__2___closed__1 = _ini lean_mark_persistent(l_Lean_Server_FileWorker_handleWaitForDiagnostics___lambda__2___closed__1); l_Lean_Server_FileWorker_handleWaitForDiagnostics___closed__1 = _init_l_Lean_Server_FileWorker_handleWaitForDiagnostics___closed__1(); lean_mark_persistent(l_Lean_Server_FileWorker_handleWaitForDiagnostics___closed__1); -l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__1 = _init_l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__1(); -lean_mark_persistent(l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__1); -l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__2 = _init_l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__2(); -lean_mark_persistent(l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__2___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__2___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__2___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__3___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__3___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__4___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___lambda__5___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__1___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__2___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__2___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__2___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__2___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__2___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__4___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__3___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__3___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__7___lambda__4___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__3___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__3___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__10___lambda__4___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__3___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__3___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__14___lambda__4___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__3___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__3___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__18___lambda__4___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__2___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__2___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__2___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__2___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__2___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__21___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__24___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__24___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__24___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__3___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__3___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__27___lambda__4___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__3___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__3___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__31___lambda__4___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__3___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__3___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____spec__34___lambda__4___closed__1); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__1 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__1(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__1); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__2 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__2(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__2); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__3 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__3(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__3); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__4 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__4(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__4); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__5 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__5(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__5); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__6 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__6(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__6); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__7 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__7(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__7); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__8 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__8(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__8); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__9 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__9(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__9); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__10 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__10(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__10); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__11 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__11(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__11); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__12 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__12(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__12); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__13 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__13(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__13); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__14 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__14(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__14); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__15 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__15(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__15); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__16 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__16(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__16); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__17 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__17(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__17); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__18 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__18(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__18); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__19 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__19(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__19); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__20 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__20(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__20); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__21 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__21(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__21); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__22 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__22(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__22); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__23 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__23(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__23); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__24 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__24(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__24); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__25 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__25(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__25); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__26 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__26(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373____closed__26); -res = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9373_(lean_io_mk_world()); +l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__1 = _init_l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__1(); +lean_mark_persistent(l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__1); +l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__2 = _init_l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__2(); +lean_mark_persistent(l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__2___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__2___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__2___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__3___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__4___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___lambda__5___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__1___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__2___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__2___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__2___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__2___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__4___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__3___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__7___lambda__4___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__3___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__10___lambda__4___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__3___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__14___lambda__4___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__3___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__18___lambda__4___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__2___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__2___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__2___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__2___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__21___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__24___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__24___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__24___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__3___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__27___lambda__4___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__3___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__31___lambda__4___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__3___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____spec__34___lambda__4___closed__1); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__1 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__1(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__1); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__2 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__2(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__2); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__3 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__3(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__3); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__4 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__4(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__4); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__5 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__5(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__5); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__6 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__6(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__6); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__7 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__7(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__7); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__8 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__8(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__8); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__9 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__9(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__9); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__10 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__10(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__10); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__11 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__11(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__11); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__12 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__12(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__12); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__13 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__13(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__13); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__14 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__14(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__14); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__15 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__15(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__15); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__16 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__16(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__16); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__17 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__17(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__17); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__18 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__18(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__18); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__19 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__19(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__19); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__20 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__20(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__20); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__21 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__21(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__21); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__22 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__22(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__22); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__23 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__23(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__23); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__24 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__24(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__24); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__25 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__25(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__25); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__26 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__26(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631____closed__26); +res = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_9631_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Server/FileWorker/Utils.c b/stage0/stdlib/Lean/Server/FileWorker/Utils.c index 307ec5a0eb5..0fa3d7e35a9 100644 --- a/stage0/stdlib/Lean/Server/FileWorker/Utils.c +++ b/stage0/stdlib/Lean/Server/FileWorker/Utils.c @@ -44,7 +44,6 @@ LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_instMonadRpcSession___rarg___l static lean_object* l_Lean_Server_FileWorker_logSnapContent___closed__4; size_t lean_usize_shift_right(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_instMonadLiftIOEIOElabTaskError___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_EditableDocument_allSnaps___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_CancelToken_new(lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_CancelToken_check___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -91,7 +90,6 @@ LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_CancelToken_check___rarg___lam LEAN_EXPORT lean_object* l_Std_PersistentHashMap_findAtAux___at_Lean_Server_FileWorker_instMonadRpcSession___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_instMonadRpcSession___rarg___lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Std_PersistentHashMap_empty___at_Lean_Server_FileWorker_RpcSession_new___spec__1___closed__2; -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_EditableDocument_allSnaps(lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); uint8_t lean_usize_dec_le(size_t, size_t); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAux___at_Lean_Server_FileWorker_RpcSession_store___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -404,29 +402,6 @@ lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_EditableDocument_allSnaps(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = lean_ctor_get(x_1, 1); -x_3 = lean_ctor_get(x_1, 2); -lean_inc(x_3); -lean_inc(x_2); -x_4 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_4, 0, x_2); -lean_ctor_set(x_4, 1, x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_EditableDocument_allSnaps___boxed(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = l_Lean_Server_FileWorker_EditableDocument_allSnaps(x_1); -lean_dec(x_1); -return x_2; -} -} static lean_object* _init_l_Lean_Server_FileWorker_RpcSession_keepAliveTimeMs() { _start: { diff --git a/stage0/stdlib/Lean/Server/FileWorker/WidgetRequests.c b/stage0/stdlib/Lean/Server/FileWorker/WidgetRequests.c index 487ecd0c802..a4c0946fac9 100644 --- a/stage0/stdlib/Lean/Server/FileWorker/WidgetRequests.c +++ b/stage0/stdlib/Lean/Server/FileWorker/WidgetRequests.c @@ -4175,7 +4175,7 @@ if (lean_is_exclusive(x_4)) { if (lean_obj_tag(x_1) == 0) { lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_15 = lean_ctor_get(x_5, 2); +x_15 = lean_ctor_get(x_5, 1); lean_inc(x_15); lean_dec(x_5); x_16 = l_Lean_Widget_getInteractiveDiagnostics___closed__1; @@ -4192,7 +4192,7 @@ goto block_14; else { 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; lean_object* x_30; lean_object* x_31; -x_20 = lean_ctor_get(x_5, 2); +x_20 = lean_ctor_get(x_5, 1); lean_inc(x_20); x_21 = lean_ctor_get(x_1, 0); lean_inc(x_21); diff --git a/stage0/stdlib/Lean/Server/InfoUtils.c b/stage0/stdlib/Lean/Server/InfoUtils.c index 9000e67162a..55520185b16 100644 --- a/stage0/stdlib/Lean/Server/InfoUtils.c +++ b/stage0/stdlib/Lean/Server/InfoUtils.c @@ -73,6 +73,7 @@ uint8_t l_Lean_Syntax_structEq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Info_fmtHover_x3f_isAtomicFormat___boxed(lean_object*); size_t lean_usize_sub(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM___at_Lean_Elab_InfoTree_collectNodesBottomUp___spec__2___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_filterTRAux___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__8___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_ppExpr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfo_go___spec__4(lean_object*); static lean_object* l_List_mapTRAux___at_Lean_Elab_InfoTree_smallestInfo_x3f___spec__2___closed__2; @@ -87,9 +88,7 @@ static lean_object* l_String_instInhabitedRange___closed__1; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__24___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_InfoTree_visitM_go___spec__1(lean_object*, lean_object*); -lean_object* l_Nat_decLt___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_maximum_x3f___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__9(lean_object*); -LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__10(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_instHashableRange; LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_collectNodesBottomUp(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__4___boxed(lean_object*); @@ -97,7 +96,6 @@ static lean_object* l___private_Lean_Server_InfoUtils_0__String_reprRange____x40 LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_getRange_x3f___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___closed__1; -LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__10___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfo_go___spec__11___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Info_fmtHover_x3f_fmtTerm_x3f___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__4(lean_object*); @@ -109,11 +107,9 @@ LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM_go___rarg(lean_object*, lea LEAN_EXPORT lean_object* l_String_Range_contains___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_instBEqRange; LEAN_EXPORT lean_object* l_Lean_Elab_Info_fmtHover_x3f_fmtTerm_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_prodHasDecidableLt___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfo_go___spec__7___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_Format_joinSep___at_instReprProd___spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___closed__4; size_t lean_usize_shift_right(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_string_utf8_byte_size(lean_object*); @@ -139,19 +135,16 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry LEAN_EXPORT lean_object* l_Lean_Elab_Info_fmtHover_x3f___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Info_fmtHover_x3f_fmtTerm_x3f___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Info_isTerm___boxed(lean_object*); -static lean_object* l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___closed__3; LEAN_EXPORT lean_object* l_Std_PersistentArray_anyM___at_Lean_Elab_InfoTree_hasSorry_go___spec__6(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__15(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentArray_anyM___at_Lean_Elab_InfoTree_hasSorry_go___spec__21(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Info_fmtHover_x3f___lambda__2___closed__1; -lean_object* l_Int_decLt___boxed(lean_object*, lean_object*); extern lean_object* l_Lean_LocalContext_empty; LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_hoverableInfoAt_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfo_go___spec__5___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Info_docString_x3f___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentArray_anyM___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_deepestNodes___rarg(lean_object*, lean_object*); -lean_object* l_Int_decEq___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Std_PersistentArray_anyM___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Std_Data_PersistentArray_0__Std_PersistentArray_foldlFromMAux___at_Lean_Elab_InfoTree_foldInfo_go___spec__15___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Info_fmtHover_x3f_fmtTerm_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -185,15 +178,13 @@ static lean_object* l_Lean_Elab_Info_fmtHover_x3f_fmtTerm_x3f___lambda__1___clos lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Info_docString_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Std_Data_PersistentArray_0__Std_PersistentArray_foldlFromMAux___at_Lean_Elab_InfoTree_foldInfo_go___spec__2(lean_object*); +uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_collectNodesBottomUp___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentArray_anyM___at_Lean_Elab_InfoTree_hasSorry_go___spec__26(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentArray_anyMAux___at_Lean_Elab_InfoTree_hasSorry_go___spec__22(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__35(lean_object*, lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_List_filterAux___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__11(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___closed__5; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__35___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10___rarg___closed__3; LEAN_EXPORT uint8_t l_Lean_Elab_Info_contains(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__30___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapM___at_Lean_Elab_InfoTree_visitM_go___spec__2___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -204,7 +195,6 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldI lean_object* l_Lean_Name_toString(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_hoverableInfoAt_x3f___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_max___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__9(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_instReprRange; LEAN_EXPORT lean_object* l_Std_PersistentArray_anyMAux___at_Lean_Elab_InfoTree_hasSorry_go___spec__27(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Info_docString_x3f___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -212,7 +202,6 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldI static lean_object* l_Lean_Elab_InfoTree_getCompletionInfos___closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfo_go___spec__24___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_goalsAt_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM_go___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_termGoalAt_x3f(lean_object*, lean_object*); lean_object* l_instInhabited___rarg(lean_object*, lean_object*); @@ -222,7 +211,6 @@ LEAN_EXPORT lean_object* l___private_Std_Data_PersistentArray_0__Std_PersistentA LEAN_EXPORT lean_object* l_Lean_Elab_Info_fmtHover_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Std_Data_PersistentArray_0__Std_PersistentArray_foldlFromMAux___at_Lean_Elab_InfoTree_foldInfo_go___spec__15___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Info_docString_x3f___lambda__2___closed__1; -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__12(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__29(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__30(lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* l_Nat_repr(lean_object*); @@ -253,18 +241,19 @@ static lean_object* l_List_mapTRAux___at_Lean_Elab_InfoTree_smallestInfo_x3f___s LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM___at_Lean_Elab_InfoTree_collectNodesBottomUp___spec__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_smallestInfo_x3f___spec__4(lean_object*, size_t, size_t, lean_object*); extern lean_object* l_instInhabitedPos; -LEAN_EXPORT lean_object* l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10___boxed(lean_object*); +LEAN_EXPORT lean_object* l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_smallestInfo_x3f___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Info_stx(lean_object*); lean_object* l_Lean_Elab_Info_toElabInfo_x3f(lean_object*); static lean_object* l___private_Lean_Server_InfoUtils_0__String_reprRange____x40_Lean_Server_InfoUtils___hyg_34____closed__7; +LEAN_EXPORT lean_object* l_List_filterTRAux___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_collectNodesBottomUp___rarg___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__4___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfo_go___spec__11(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Info_occursInside_x3f(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10(lean_object*); +LEAN_EXPORT lean_object* l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfo_go___spec__17(lean_object*); -LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__11___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__11___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Info_range_x3f(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_goalsAt_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Info_docString_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -281,7 +270,6 @@ size_t lean_usize_land(size_t, size_t); static lean_object* l_Lean_Elab_InfoTree_getCompletionInfos___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_foldInfo___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentArray_anyM___at_Lean_Elab_InfoTree_hasSorry_go___spec__31(lean_object*, lean_object*, lean_object*); -lean_object* lean_int_neg(lean_object*); LEAN_EXPORT lean_object* l_List_mapTRAux___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Std_Data_PersistentArray_0__Std_PersistentArray_foldlFromMAux___at_Lean_Elab_InfoTree_foldInfo_go___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_PersistentArray_toList___rarg(lean_object*); @@ -297,9 +285,9 @@ static lean_object* l___private_Lean_Server_InfoUtils_0__String_reprRange____x40 static lean_object* l___private_Lean_Server_InfoUtils_0__String_reprRange____x40_Lean_Server_InfoUtils___hyg_34____closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_deepestNodes___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_redLength___rarg(lean_object*); +LEAN_EXPORT lean_object* l_List_filterTRAux___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__8(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfo_go___spec__9(lean_object*); LEAN_EXPORT lean_object* l_List_mapM___at_Lean_Elab_InfoTree_collectNodesBottomUp___spec__5(lean_object*); -static lean_object* l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10___rarg___closed__2; LEAN_EXPORT lean_object* l_List_mapM___at_Lean_Elab_InfoTree_visitM_go___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfo_go___spec__26___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfo_go___spec__26___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -344,7 +332,6 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_hoverableInfoAt_x3f(lean_object*, lean_object*, uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfo_go___spec__8(lean_object*); static lean_object* l_Lean_Elab_InfoTree_hoverableInfoAt_x3f___lambda__2___closed__5; -uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_InfoTree_collectNodesBottomUp___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_smallestInfo_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint64_t l___private_Lean_Server_InfoUtils_0__String_hashRange____x40_Lean_Server_InfoUtils___hyg_146_(lean_object*); @@ -367,7 +354,6 @@ lean_object* l_List_toArrayAux___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfo_go___spec__12___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Std_Data_PersistentArray_0__Std_PersistentArray_foldlMAux___at_Lean_Elab_InfoTree_foldInfo_go___spec__10___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10___rarg___closed__4; LEAN_EXPORT lean_object* l___private_Std_Data_PersistentArray_0__Std_PersistentArray_foldlMAux___at_Lean_Elab_InfoTree_foldInfo_go___spec__23(lean_object*); lean_object* l_Lean_Meta_instantiateMVars___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Std_PersistentArray_anyMAux___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); @@ -382,11 +368,9 @@ static lean_object* l___private_Lean_Server_InfoUtils_0__String_reprRange____x40 LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfo_go___spec__11___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_Meta_instantiateMVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__34(lean_object*, lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__12___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentArray_anyMAux___at_Lean_Elab_InfoTree_hasSorry_go___spec__32(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Info_stx___boxed(lean_object*); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); -static lean_object* l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10___rarg___closed__1; uint64_t lean_uint64_mix_hash(uint64_t, uint64_t); LEAN_EXPORT lean_object* l_Std_PersistentArray_anyM___at_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Info_fmtHover_x3f___lambda__1___closed__1; @@ -402,7 +386,6 @@ LEAN_EXPORT lean_object* l_Std_PersistentArray_anyM___at_Lean_Elab_InfoTree_hasS lean_object* lean_string_length(lean_object*); static lean_object* l_String_instBEqRange___closed__1; LEAN_EXPORT uint8_t l_Lean_Elab_Info_isTerm(lean_object*); -LEAN_EXPORT lean_object* l_List_maximum_x3f___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__8(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_collectNodesBottomUp___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentArray_anyM___at_Lean_Elab_InfoTree_hasSorry_go___spec__11(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); @@ -413,7 +396,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Info_pos_x3f(lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentArray_foldlM___at_Lean_Elab_InfoTree_foldInfo_go___spec__14(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__8(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_List_mapM___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_filterAux___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__11___boxed(lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_pp_universes; static lean_object* l_Lean_Elab_InfoTree_hoverableInfoAt_x3f___lambda__2___closed__2; uint8_t l_List_isEmpty___rarg(lean_object*); @@ -424,7 +406,6 @@ LEAN_EXPORT lean_object* l_Std_PersistentArray_anyM___at_Lean_Elab_InfoTree_goal LEAN_EXPORT lean_object* l_Lean_Elab_Info_lctx(lean_object*); static lean_object* l_panic___at_Lean_Elab_InfoTree_collectNodesBottomUp___spec__4___rarg___closed__1; lean_object* lean_usize_to_nat(size_t); -uint8_t lean_int_dec_eq(lean_object*, lean_object*); lean_object* l_Std_instInhabitedPersistentArrayNode(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__24(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_foldInfo(lean_object*); @@ -432,7 +413,6 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry lean_object* l_Lean_Elab_ContextInfo_runMetaM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__3(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentArray_anyMAux___at_Lean_Elab_InfoTree_hasSorry_go___spec__7(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_max___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__9___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfo_go___spec__24___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t); @@ -442,7 +422,7 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldI static lean_object* l___private_Lean_Server_InfoUtils_0__String_reprRange____x40_Lean_Server_InfoUtils___hyg_34____closed__12; LEAN_EXPORT lean_object* l_Std_PersistentArray_foldlM___at_Lean_Elab_InfoTree_foldInfo_go___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__34___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__11(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__11(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTRAux___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__7(lean_object*, lean_object*); lean_object* lean_nat_to_int(lean_object*); LEAN_EXPORT lean_object* l_List_mapM___at_Lean_Elab_InfoTree_collectNodesBottomUp___spec__5___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -452,7 +432,6 @@ LEAN_EXPORT lean_object* l___private_Std_Data_PersistentArray_0__Std_PersistentA LEAN_EXPORT lean_object* l_Std_PersistentArray_foldlM___at_Lean_Elab_InfoTree_foldInfo_go___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_hoverableInfoAt_x3f___lambda__4(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_filterAux___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapM___at_Lean_Elab_InfoTree_visitM_go___spec__2___rarg___lambda__1(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_CompletionInfo_stx(lean_object*); @@ -464,8 +443,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Info_occursInside_x3f___boxed(lean_object*, LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__18(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_getCompletionInfos___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM_go___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___closed__6; -lean_object* l_Nat_decEq___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentArray_anyMAux___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__29___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentArray_anyMAux___at_Lean_Elab_InfoTree_hasSorry_go___spec__17(lean_object*, lean_object*, lean_object*); @@ -5051,7 +5028,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_List_filterAux___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_filterTRAux___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_2) == 0) @@ -5405,89 +5382,39 @@ goto _start; } } } -static lean_object* _init_l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10___rarg___closed__1() { +LEAN_EXPORT lean_object* l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Int_decEq___boxed), 2, 0); -return x_1; -} -} -static lean_object* _init_l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10___rarg___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Nat_decEq___boxed), 2, 0); -return x_1; -} -} -static lean_object* _init_l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10___rarg___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Int_decLt___boxed), 2, 0); -return x_1; -} -} -static lean_object* _init_l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10___rarg___closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Nat_decLt___boxed), 2, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10___rarg(lean_object* x_1, lean_object* x_2) { -_start: +uint8_t x_3; +x_3 = lean_nat_dec_lt(x_2, x_1); +if (x_3 == 0) { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_3 = l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10___rarg___closed__1; -x_4 = l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10___rarg___closed__2; -x_5 = l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10___rarg___closed__3; -x_6 = l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10___rarg___closed__4; -lean_inc(x_1); lean_inc(x_2); -x_7 = l_prodHasDecidableLt___rarg(x_3, x_4, x_5, x_6, x_2, x_1); -x_8 = lean_unbox(x_7); -lean_dec(x_7); -if (x_8 == 0) -{ -lean_dec(x_1); return x_2; } else { -lean_dec(x_2); +lean_inc(x_1); return x_1; } } } -LEAN_EXPORT lean_object* l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10___rarg), 2, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__11(lean_object* x_1, lean_object* x_2) { _start: { -if (lean_obj_tag(x_3) == 0) +if (lean_obj_tag(x_2) == 0) { -return x_2; +return x_1; } else { -lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_4 = lean_ctor_get(x_3, 0); -lean_inc(x_4); -x_5 = lean_ctor_get(x_3, 1); -lean_inc(x_5); -lean_dec(x_3); -x_6 = l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10___rarg(x_2, x_4); -x_2 = x_6; -x_3 = x_5; +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_ctor_get(x_2, 0); +x_4 = lean_ctor_get(x_2, 1); +x_5 = l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10(x_1, x_3); +lean_dec(x_1); +x_1 = x_5; +x_2 = x_4; goto _start; } } @@ -5503,78 +5430,26 @@ return x_2; } else { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_3 = lean_ctor_get(x_1, 0); lean_inc(x_3); x_4 = lean_ctor_get(x_1, 1); lean_inc(x_4); lean_dec(x_1); -x_5 = lean_box(0); -x_6 = l_List_foldl___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__11(x_5, x_3, x_4); -x_7 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_7, 0, x_6); -return x_7; -} -} -} -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__12(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -if (lean_obj_tag(x_2) == 0) -{ -uint8_t x_3; -x_3 = 1; -return x_3; -} -else -{ -uint8_t x_4; -x_4 = 0; -return x_4; -} -} -else -{ -if (lean_obj_tag(x_2) == 0) -{ -uint8_t x_5; -x_5 = 0; -return x_5; -} -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; uint8_t x_12; -x_6 = lean_ctor_get(x_1, 0); -x_7 = lean_ctor_get(x_2, 0); -x_8 = lean_ctor_get(x_6, 0); -x_9 = lean_ctor_get(x_6, 1); -x_10 = lean_ctor_get(x_7, 0); -x_11 = lean_ctor_get(x_7, 1); -x_12 = lean_int_dec_eq(x_8, x_10); -if (x_12 == 0) -{ -uint8_t x_13; -x_13 = 0; -return x_13; -} -else -{ -uint8_t x_14; -x_14 = lean_nat_dec_eq(x_9, x_11); -return x_14; -} -} +x_5 = l_List_foldl___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__11(x_3, x_4); +lean_dec(x_4); +x_6 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_6, 0, x_5); +return x_6; } } } LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_hoverableInfoAt_x3f___lambda__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_8; lean_object* x_41; uint8_t x_46; -x_46 = l_List_isEmpty___rarg(x_6); -if (x_46 == 0) +lean_object* x_8; lean_object* x_40; uint8_t x_45; +x_45 = l_List_isEmpty___rarg(x_6); +if (x_45 == 0) { lean_dec(x_5); lean_dec(x_4); @@ -5586,9 +5461,9 @@ else { if (lean_obj_tag(x_1) == 4) { -uint8_t x_47; -x_47 = l_Lean_Elab_Info_contains(x_1, x_2, x_3); -if (x_47 == 0) +uint8_t x_46; +x_46 = l_Lean_Elab_Info_contains(x_1, x_2, x_3); +if (x_46 == 0) { lean_dec(x_5); lean_dec(x_4); @@ -5598,21 +5473,21 @@ return x_6; } else { -lean_object* x_48; -x_48 = lean_box(0); -x_8 = x_48; -goto block_40; +lean_object* x_47; +x_47 = lean_box(0); +x_8 = x_47; +goto block_39; } } else { -lean_object* x_49; -x_49 = lean_box(0); -x_41 = x_49; -goto block_45; +lean_object* x_48; +x_48 = lean_box(0); +x_40 = x_48; +goto block_44; } } -block_40: +block_39: { lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_dec(x_8); @@ -5623,105 +5498,111 @@ lean_ctor_set(x_10, 0, x_4); lean_ctor_set(x_10, 1, x_1); if (lean_obj_tag(x_9) == 0) { -lean_object* x_37; lean_object* x_38; -x_37 = l_List_mapTRAux___at_Lean_Elab_InfoTree_smallestInfo_x3f___spec__2___closed__4; -x_38 = l_panic___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__1(x_37); -x_11 = x_38; -goto block_36; +lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_24 = l_List_mapTRAux___at_Lean_Elab_InfoTree_smallestInfo_x3f___spec__2___closed__4; +x_25 = l_panic___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__1(x_24); +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +x_27 = lean_nat_dec_eq(x_26, x_2); +lean_dec(x_26); +if (x_27 == 0) +{ +lean_object* x_28; +x_28 = lean_box(0); +x_11 = x_28; +goto block_23; } else { -lean_object* x_39; -x_39 = lean_ctor_get(x_9, 0); -lean_inc(x_39); +lean_object* x_29; lean_object* x_30; lean_object* x_31; +lean_dec(x_1); +x_29 = lean_unsigned_to_nat(0u); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_10); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_5); +return x_31; +} +} +else +{ +lean_object* x_32; lean_object* x_33; uint8_t x_34; +x_32 = lean_ctor_get(x_9, 0); +lean_inc(x_32); lean_dec(x_9); -x_11 = x_39; -goto block_36; +x_33 = lean_ctor_get(x_32, 1); +lean_inc(x_33); +lean_dec(x_32); +x_34 = lean_nat_dec_eq(x_33, x_2); +lean_dec(x_33); +if (x_34 == 0) +{ +lean_object* x_35; +x_35 = lean_box(0); +x_11 = x_35; +goto block_23; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +lean_dec(x_1); +x_36 = lean_unsigned_to_nat(0u); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_10); +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_5); +return x_38; +} } -block_36: +block_23: { -lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_12 = lean_ctor_get(x_11, 1); -lean_inc(x_12); -x_13 = lean_nat_dec_eq(x_12, x_2); -x_14 = lean_ctor_get(x_11, 0); -lean_inc(x_14); lean_dec(x_11); -x_15 = lean_nat_sub(x_12, x_14); -lean_dec(x_14); -lean_dec(x_12); -x_16 = lean_nat_to_int(x_15); -x_17 = lean_int_neg(x_16); -lean_dec(x_16); -if (x_13 == 0) -{ if (lean_obj_tag(x_1) == 1) { -lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_1, 0); -lean_inc(x_24); +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_1, 0); +lean_inc(x_12); lean_dec(x_1); -x_25 = lean_ctor_get(x_24, 3); -lean_inc(x_25); -lean_dec(x_24); -if (lean_obj_tag(x_25) == 1) -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -lean_dec(x_25); -x_26 = lean_unsigned_to_nat(0u); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_17); -lean_ctor_set(x_27, 1, x_26); -x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_27); -lean_ctor_set(x_28, 1, x_10); -x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_5); -return x_29; -} -else +x_13 = lean_ctor_get(x_12, 3); +lean_inc(x_13); +lean_dec(x_12); +if (lean_obj_tag(x_13) == 1) { -lean_object* x_30; -lean_dec(x_25); -x_30 = lean_box(0); -x_18 = x_30; -goto block_23; -} +lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_dec(x_13); +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_10); +x_16 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_5); +return x_16; } else { -lean_object* x_31; -lean_dec(x_1); -x_31 = lean_box(0); -x_18 = x_31; -goto block_23; +lean_object* x_17; lean_object* x_18; lean_object* x_19; +lean_dec(x_13); +x_17 = lean_unsigned_to_nat(1u); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_10); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_5); +return x_19; } } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_dec(x_1); -x_32 = lean_unsigned_to_nat(0u); -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_17); -lean_ctor_set(x_33, 1, x_32); -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_10); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_5); -return x_35; -} -block_23: -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -lean_dec(x_18); -x_19 = lean_unsigned_to_nat(1u); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_17); -lean_ctor_set(x_20, 1, x_19); +x_20 = lean_unsigned_to_nat(1u); x_21 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_21, 0, x_20); lean_ctor_set(x_21, 1, x_10); @@ -5732,12 +5613,12 @@ return x_22; } } } -block_45: +block_44: { -lean_object* x_42; -lean_dec(x_41); -x_42 = l_Lean_Elab_Info_toElabInfo_x3f(x_1); -if (lean_obj_tag(x_42) == 0) +lean_object* x_41; +lean_dec(x_40); +x_41 = l_Lean_Elab_Info_toElabInfo_x3f(x_1); +if (lean_obj_tag(x_41) == 0) { lean_dec(x_5); lean_dec(x_4); @@ -5747,10 +5628,10 @@ return x_6; } else { -uint8_t x_43; -lean_dec(x_42); -x_43 = l_Lean_Elab_Info_contains(x_1, x_2, x_3); -if (x_43 == 0) +uint8_t x_42; +lean_dec(x_41); +x_42 = l_Lean_Elab_Info_contains(x_1, x_2, x_3); +if (x_42 == 0) { lean_dec(x_5); lean_dec(x_4); @@ -5760,10 +5641,10 @@ return x_6; } else { -lean_object* x_44; -x_44 = lean_box(0); -x_8 = x_44; -goto block_40; +lean_object* x_43; +x_43 = lean_box(0); +x_8 = x_43; +goto block_39; } } } @@ -5846,16 +5727,15 @@ LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_hoverableInfoAt_x3f___lambda__2(le { lean_object* x_9; lean_object* x_10; lean_dec(x_7); -lean_inc(x_2); -x_9 = l_List_mapTRAux___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__2(x_8, x_2); +lean_inc(x_1); +x_9 = l_List_mapTRAux___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__2(x_8, x_1); x_10 = l_List_join___rarg(x_9); if (x_4 == 0) { lean_object* x_11; lean_object* x_12; x_11 = lean_box(0); -x_12 = l_Lean_Elab_InfoTree_hoverableInfoAt_x3f___lambda__1(x_6, x_1, x_3, x_5, x_2, x_10, x_11); +x_12 = l_Lean_Elab_InfoTree_hoverableInfoAt_x3f___lambda__1(x_6, x_2, x_3, x_5, x_1, x_10, x_11); lean_dec(x_10); -lean_dec(x_1); return x_12; } else @@ -5870,9 +5750,8 @@ if (x_15 == 0) lean_object* x_16; lean_object* x_17; lean_dec(x_13); x_16 = lean_box(0); -x_17 = l_Lean_Elab_InfoTree_hoverableInfoAt_x3f___lambda__1(x_6, x_1, x_3, x_5, x_2, x_10, x_16); +x_17 = l_Lean_Elab_InfoTree_hoverableInfoAt_x3f___lambda__1(x_6, x_2, x_3, x_5, x_1, x_10, x_16); lean_dec(x_10); -lean_dec(x_1); return x_17; } else @@ -5887,20 +5766,18 @@ if (x_20 == 0) lean_object* x_21; lean_object* x_22; lean_dec(x_19); x_21 = lean_box(0); -x_22 = l_Lean_Elab_InfoTree_hoverableInfoAt_x3f___lambda__1(x_6, x_1, x_3, x_5, x_2, x_10, x_21); +x_22 = l_Lean_Elab_InfoTree_hoverableInfoAt_x3f___lambda__1(x_6, x_2, x_3, x_5, x_1, x_10, x_21); lean_dec(x_10); -lean_dec(x_1); return x_22; } else { lean_object* x_23; lean_object* x_24; lean_object* x_25; -lean_inc(x_2); -x_23 = l_List_filterAux___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__3(x_19, x_10, x_2); +lean_inc(x_1); +x_23 = l_List_filterTRAux___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__3(x_19, x_10, x_1); x_24 = lean_box(0); -x_25 = l_Lean_Elab_InfoTree_hoverableInfoAt_x3f___lambda__1(x_6, x_1, x_3, x_5, x_2, x_23, x_24); +x_25 = l_Lean_Elab_InfoTree_hoverableInfoAt_x3f___lambda__1(x_6, x_2, x_3, x_5, x_1, x_23, x_24); lean_dec(x_23); -lean_dec(x_1); return x_25; } } @@ -5913,11 +5790,9 @@ LEAN_EXPORT uint8_t l_Lean_Elab_InfoTree_hoverableInfoAt_x3f___lambda__3(lean_ob lean_object* x_3; lean_object* x_4; uint8_t x_5; x_3 = lean_ctor_get(x_2, 0); lean_inc(x_3); -lean_dec(x_2); x_4 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_4, 0, x_3); -x_5 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__12(x_4, x_1); -lean_dec(x_1); +x_5 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1(x_4, x_1); lean_dec(x_4); return x_5; } @@ -5937,8 +5812,8 @@ x_5 = lean_box(0); x_6 = lean_box(x_3); x_7 = lean_box(x_4); x_8 = lean_alloc_closure((void*)(l_Lean_Elab_InfoTree_hoverableInfoAt_x3f___lambda__2___boxed), 8, 4); -lean_closure_set(x_8, 0, x_2); -lean_closure_set(x_8, 1, x_5); +lean_closure_set(x_8, 0, x_5); +lean_closure_set(x_8, 1, x_2); lean_closure_set(x_8, 2, x_6); lean_closure_set(x_8, 3, x_7); x_9 = l_Lean_Elab_InfoTree_collectNodesBottomUp___rarg___closed__1; @@ -6074,33 +5949,23 @@ lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10___boxed(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10(x_1); -lean_dec(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_4; -x_4 = l_List_foldl___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__11(x_1, x_2, x_3); +lean_object* x_3; +x_3 = l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10(x_1, x_2); +lean_dec(x_2); lean_dec(x_1); -return x_4; +return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__12___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__11___boxed(lean_object* x_1, lean_object* x_2) { _start: { -uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__12(x_1, x_2); +lean_object* x_3; +x_3 = l_List_foldl___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__11(x_1, x_2); lean_dec(x_2); -lean_dec(x_1); -x_4 = lean_box(x_3); -return x_4; +return x_3; } } LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_hoverableInfoAt_x3f___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) { @@ -6125,6 +5990,7 @@ lean_dec(x_3); x_10 = lean_unbox(x_4); lean_dec(x_4); x_11 = l_Lean_Elab_InfoTree_hoverableInfoAt_x3f___lambda__2(x_1, x_2, x_9, x_10, x_5, x_6, x_7, x_8); +lean_dec(x_2); return x_11; } } @@ -6133,6 +5999,8 @@ LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_hoverableInfoAt_x3f___lambda__3___ { uint8_t x_3; lean_object* x_4; x_3 = l_Lean_Elab_InfoTree_hoverableInfoAt_x3f___lambda__3(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); x_4 = lean_box(x_3); return x_4; } @@ -8255,42 +8123,6 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Tactic", 6); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_InfoTree_hoverableInfoAt_x3f___lambda__2___closed__4; -x_2 = l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___closed__3; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___closed__5() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("tacticSeq", 9); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___closed__6() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___closed__4; -x_2 = l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___closed__5; -x_3 = lean_name_mk_string(x_1, x_2); -return x_3; -} -} LEAN_EXPORT uint8_t l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { @@ -8308,12 +8140,10 @@ lean_inc(x_6); lean_dec(x_4); x_7 = l_Lean_Elab_Info_stx(x_5); x_8 = l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___closed__2; -lean_inc(x_7); x_9 = l_Lean_Syntax_isOfKind(x_7, x_8); if (x_9 == 0) { lean_object* x_10; uint8_t x_11; -lean_dec(x_7); x_10 = lean_box(0); x_11 = l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___lambda__2(x_1, x_2, x_3, x_6, x_5, x_10); lean_dec(x_5); @@ -8321,56 +8151,39 @@ return x_11; } else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_12 = lean_unsigned_to_nat(1u); -x_13 = l_Lean_Syntax_getArg(x_7, x_12); -lean_dec(x_7); -x_14 = l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___closed__6; -x_15 = l_Lean_Syntax_isOfKind(x_13, x_14); -if (x_15 == 0) -{ -lean_object* x_16; uint8_t x_17; -x_16 = lean_box(0); -x_17 = l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___lambda__2(x_1, x_2, x_3, x_6, x_5, x_16); -lean_dec(x_5); -return x_17; -} -else -{ -uint8_t x_18; +uint8_t x_12; lean_dec(x_6); lean_dec(x_5); -x_18 = 0; -return x_18; -} +x_12 = 0; +return x_12; } } case 3: { -lean_object* x_19; uint8_t x_20; +lean_object* x_13; uint8_t x_14; lean_dec(x_5); -x_19 = lean_ctor_get(x_4, 1); -lean_inc(x_19); +x_13 = lean_ctor_get(x_4, 1); +lean_inc(x_13); lean_dec(x_4); -x_20 = l_Std_PersistentArray_anyM___at_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___spec__6(x_1, x_2, x_3, x_19); -return x_20; +x_14 = l_Std_PersistentArray_anyM___at_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___spec__6(x_1, x_2, x_3, x_13); +return x_14; } default: { -uint8_t x_21; +uint8_t x_15; lean_dec(x_5); lean_dec(x_4); -x_21 = 0; -return x_21; +x_15 = 0; +return x_15; } } } else { -uint8_t x_22; +uint8_t x_16; lean_dec(x_4); -x_22 = 0; -return x_22; +x_16 = 0; +return x_16; } } } @@ -8884,69 +8697,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_max___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__9(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint8_t x_3; -x_3 = lean_nat_dec_lt(x_2, x_1); -if (x_3 == 0) -{ -lean_inc(x_2); -return x_2; -} -else -{ -lean_inc(x_1); -return x_1; -} -} -} -LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__10(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_2) == 0) -{ -return x_1; -} -else -{ -lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_3 = lean_ctor_get(x_2, 0); -x_4 = lean_ctor_get(x_2, 1); -x_5 = l_max___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__9(x_1, x_3); -lean_dec(x_1); -x_1 = x_5; -x_2 = x_4; -goto _start; -} -} -} -LEAN_EXPORT lean_object* l_List_maximum_x3f___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__8(lean_object* x_1) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_2; -x_2 = lean_box(0); -return x_2; -} -else -{ -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_3 = lean_ctor_get(x_1, 0); -lean_inc(x_3); -x_4 = lean_ctor_get(x_1, 1); -lean_inc(x_4); -lean_dec(x_1); -x_5 = l_List_foldl___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__10(x_3, x_4); -lean_dec(x_4); -x_6 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_6, 0, x_5); -return x_6; -} -} -} -LEAN_EXPORT lean_object* l_List_filterAux___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_filterTRAux___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_2) == 0) @@ -8968,7 +8719,7 @@ x_8 = lean_ctor_get(x_6, 2); lean_inc(x_8); x_9 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_9, 0, x_8); -x_10 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1(x_9, x_1); +x_10 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1(x_9, x_1); lean_dec(x_9); if (x_10 == 0) { @@ -9001,7 +8752,7 @@ x_15 = lean_ctor_get(x_13, 2); lean_inc(x_15); x_16 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_16, 0, x_15); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1(x_16, x_1); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_341____spec__1(x_16, x_1); lean_dec(x_16); if (x_17 == 0) { @@ -9367,8 +9118,8 @@ x_5 = l_Lean_Elab_InfoTree_collectNodesBottomUp___rarg(x_4, x_2); x_6 = lean_box(0); lean_inc(x_5); x_7 = l_List_mapTRAux___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__7(x_5, x_6); -x_8 = l_List_maximum_x3f___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__8(x_7); -x_9 = l_List_filterAux___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__11(x_8, x_5, x_6); +x_8 = l_List_maximum_x3f___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__9(x_7); +x_9 = l_List_filterTRAux___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__8(x_8, x_5, x_6); lean_dec(x_8); return x_9; } @@ -9460,30 +9211,11 @@ x_5 = lean_box(x_4); return x_5; } } -LEAN_EXPORT lean_object* l_max___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__9___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_max___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__9(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__10___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_List_foldl___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__10(x_1, x_2); -lean_dec(x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_List_filterAux___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_filterTRAux___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_List_filterAux___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__11(x_1, x_2, x_3); +x_4 = l_List_filterTRAux___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__8(x_1, x_2, x_3); lean_dec(x_1); return x_4; } @@ -14103,14 +13835,6 @@ l_List_mapTRAux___at_Lean_Elab_InfoTree_smallestInfo_x3f___spec__2___closed__3 = lean_mark_persistent(l_List_mapTRAux___at_Lean_Elab_InfoTree_smallestInfo_x3f___spec__2___closed__3); l_List_mapTRAux___at_Lean_Elab_InfoTree_smallestInfo_x3f___spec__2___closed__4 = _init_l_List_mapTRAux___at_Lean_Elab_InfoTree_smallestInfo_x3f___spec__2___closed__4(); lean_mark_persistent(l_List_mapTRAux___at_Lean_Elab_InfoTree_smallestInfo_x3f___spec__2___closed__4); -l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10___rarg___closed__1 = _init_l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10___rarg___closed__1(); -lean_mark_persistent(l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10___rarg___closed__1); -l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10___rarg___closed__2 = _init_l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10___rarg___closed__2(); -lean_mark_persistent(l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10___rarg___closed__2); -l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10___rarg___closed__3 = _init_l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10___rarg___closed__3(); -lean_mark_persistent(l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10___rarg___closed__3); -l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10___rarg___closed__4 = _init_l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10___rarg___closed__4(); -lean_mark_persistent(l_max___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10___rarg___closed__4); l_Lean_Elab_InfoTree_hoverableInfoAt_x3f___lambda__2___closed__1 = _init_l_Lean_Elab_InfoTree_hoverableInfoAt_x3f___lambda__2___closed__1(); lean_mark_persistent(l_Lean_Elab_InfoTree_hoverableInfoAt_x3f___lambda__2___closed__1); l_Lean_Elab_InfoTree_hoverableInfoAt_x3f___lambda__2___closed__2 = _init_l_Lean_Elab_InfoTree_hoverableInfoAt_x3f___lambda__2___closed__2(); @@ -14161,14 +13885,6 @@ l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___closed__1 = _init_l_Lean_Elab lean_mark_persistent(l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___closed__1); l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___closed__2 = _init_l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___closed__2(); lean_mark_persistent(l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___closed__2); -l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___closed__3 = _init_l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___closed__3(); -lean_mark_persistent(l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___closed__3); -l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___closed__4 = _init_l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___closed__4(); -lean_mark_persistent(l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___closed__4); -l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___closed__5 = _init_l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___closed__5(); -lean_mark_persistent(l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___closed__5); -l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___closed__6 = _init_l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___closed__6(); -lean_mark_persistent(l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___closed__6); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Server/References.c b/stage0/stdlib/Lean/Server/References.c index b864fb930bc..a5817041c4b 100644 --- a/stage0/stdlib/Lean/Server/References.c +++ b/stage0/stdlib/Lean/Server/References.c @@ -74,7 +74,6 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_combineFvars_ LEAN_EXPORT lean_object* l___private_Lean_Server_References_0__Lean_Server_fromJsonIlean____x40_Lean_Server_References___hyg_793____boxed(lean_object*); static lean_object* l_Lean_Server_instInhabitedReference___closed__3; LEAN_EXPORT lean_object* l_Std_AssocList_foldlM___at_Lean_Server_References_allRefs___spec__10(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_filterAux___at_Lean_Server_References_removeIlean___spec__4___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Std_AssocList_contains___at_Lean_Server_References_allRefs___spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Server_References_0__Lean_Server_fromJsonIlean____x40_Lean_Server_References___hyg_793____closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_findReferences___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -237,6 +236,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_RefInfo_merge(lean_object*, lean_object*); lean_object* l_instBEqProd___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_AssocList_contains___at_Lean_Server_combineFvars___spec__13___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Server_References_allRefs___spec__12(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_filterTRAux___at_Lean_Server_References_removeIlean___spec__4___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_identOf___boxed(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_combineFvars___spec__20___closed__1; LEAN_EXPORT lean_object* l_Std_HashMap_insert___at_Lean_Server_combineFvars___spec__2(lean_object*, lean_object*, lean_object*); @@ -303,12 +303,12 @@ uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_RefInfo_addRef(lean_object*, lean_object*); static lean_object* l_Lean_Server_instToJsonIlean___closed__1; LEAN_EXPORT lean_object* l_Std_mkHashMap___at_Lean_Server_dedupReferences___spec__1___boxed(lean_object*); -LEAN_EXPORT lean_object* l_List_filterAux___at_Lean_Server_References_removeIlean___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashMapImp_expand___at_Lean_Server_combineFvars___spec__14(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashMap_toList___at_Lean_Server_ModuleRefs_instCoeModuleRefsModuleRefs___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Std_AssocList_foldlM___at_Lean_Server_dedupReferences___spec__8(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_References_definitionsMatching___rarg___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Server_References_0__Lean_Server_fromJsonIlean____x40_Lean_Server_References___hyg_793_(lean_object*); +LEAN_EXPORT lean_object* l_List_filterTRAux___at_Lean_Server_References_removeIlean___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_findModuleRefs___spec__2(lean_object*, size_t, size_t, lean_object*); lean_object* l_instHashableProd___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Std_AssocList_contains___at_Lean_Server_dedupReferences___spec__5(lean_object*, lean_object*); @@ -6049,7 +6049,7 @@ return x_10; } } } -LEAN_EXPORT lean_object* l_List_filterAux___at_Lean_Server_References_removeIlean___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_filterTRAux___at_Lean_Server_References_removeIlean___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_2) == 0) @@ -6355,7 +6355,7 @@ x_3 = lean_ctor_get(x_1, 0); lean_inc(x_3); x_4 = l_Std_HashMap_toList___at_Lean_Server_References_removeIlean___spec__1(x_3); x_5 = lean_box(0); -x_6 = l_List_filterAux___at_Lean_Server_References_removeIlean___spec__4(x_2, x_4, x_5); +x_6 = l_List_filterTRAux___at_Lean_Server_References_removeIlean___spec__4(x_2, x_4, x_5); lean_dec(x_2); x_7 = l_List_mapTRAux___at_Lean_Server_References_removeIlean___spec__5(x_6, x_5); x_8 = l_List_foldl___at_Lean_Server_References_removeIlean___spec__8(x_1, x_7); @@ -6385,11 +6385,11 @@ lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l_List_filterAux___at_Lean_Server_References_removeIlean___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_filterTRAux___at_Lean_Server_References_removeIlean___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_List_filterAux___at_Lean_Server_References_removeIlean___spec__4(x_1, x_2, x_3); +x_4 = l_List_filterTRAux___at_Lean_Server_References_removeIlean___spec__4(x_1, x_2, x_3); lean_dec(x_1); return x_4; } diff --git a/stage0/stdlib/Lean/Server/Requests.c b/stage0/stdlib/Lean/Server/Requests.c index 48d636e50c3..b59a92902e2 100644 --- a/stage0/stdlib/Lean/Server/Requests.c +++ b/stage0/stdlib/Lean/Server/Requests.c @@ -96,7 +96,6 @@ LEAN_EXPORT lean_object* l_Lean_Server_RequestM_mapTask___rarg(lean_object*, lea LEAN_EXPORT lean_object* l_Std_PersistentHashMap_findAux___at_Lean_Server_lookupLspRequestHandler___spec__2(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_instMonadLiftIORequestM___rarg(lean_object*, lean_object*, lean_object*); size_t lean_usize_mul(size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_Server_RequestM_withWaitFindSnap___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_containsAux___at_Lean_Server_registerLspRequestHandler___spec__7___boxed(lean_object*, lean_object*, lean_object*); extern lean_object* l_Task_Priority_default; static lean_object* l_Lean_Server_initFn____x40_Lean_Server_Requests___hyg_575____closed__5; @@ -110,7 +109,6 @@ LEAN_EXPORT lean_object* l_Lean_Server_instMonadLiftIORequestM(lean_object*); lean_object* lean_io_bind_task(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_initFn____x40_Lean_Server_Requests___hyg_575____closed__3; LEAN_EXPORT lean_object* l_Lean_Server_RequestM_asTask(lean_object*); -lean_object* l_Lean_Server_FileWorker_EditableDocument_allSnaps(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_RequestM_bindTask___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_RequestM_waitFindSnapAux___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___lambda__1(lean_object*, lean_object*, lean_object*); @@ -151,7 +149,6 @@ static lean_object* l_Lean_Server_registerLspRequestHandler___closed__1; lean_object* lean_io_map_task(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_task_pure(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_RequestM_bindWaitFindSnap___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_RequestM_mapTask___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_instMonadLiftIORequestM___rarg___boxed(lean_object*, lean_object*, lean_object*); @@ -975,7 +972,9 @@ LEAN_EXPORT lean_object* l_Lean_Server_RequestM_withWaitFindSnap___rarg(lean_obj _start: { lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_7 = l_Lean_Server_FileWorker_EditableDocument_allSnaps(x_1); +x_7 = lean_ctor_get(x_1, 1); +lean_inc(x_7); +lean_dec(x_1); x_8 = l_IO_AsyncList_waitFind_x3f___rarg(x_2, x_7, x_6); x_9 = lean_ctor_get(x_8, 0); lean_inc(x_9); @@ -993,24 +992,17 @@ LEAN_EXPORT lean_object* l_Lean_Server_RequestM_withWaitFindSnap(lean_object* x_ _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Server_RequestM_withWaitFindSnap___rarg___boxed), 6, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Server_RequestM_withWaitFindSnap___rarg), 6, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_RequestM_withWaitFindSnap___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) { -_start: -{ -lean_object* x_7; -x_7 = l_Lean_Server_RequestM_withWaitFindSnap___rarg(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_1); -return x_7; -} -} LEAN_EXPORT lean_object* l_Lean_Server_RequestM_bindWaitFindSnap___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; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_7 = l_Lean_Server_FileWorker_EditableDocument_allSnaps(x_1); +x_7 = lean_ctor_get(x_1, 1); +lean_inc(x_7); +lean_dec(x_1); x_8 = l_IO_AsyncList_waitFind_x3f___rarg(x_2, x_7, x_6); x_9 = lean_ctor_get(x_8, 0); lean_inc(x_9); @@ -1028,19 +1020,10 @@ LEAN_EXPORT lean_object* l_Lean_Server_RequestM_bindWaitFindSnap(lean_object* x_ _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Server_RequestM_bindWaitFindSnap___rarg___boxed), 6, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Server_RequestM_bindWaitFindSnap___rarg), 6, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_RequestM_bindWaitFindSnap___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) { -_start: -{ -lean_object* x_7; -x_7 = l_Lean_Server_RequestM_bindWaitFindSnap___rarg(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_1); -return x_7; -} -} static lean_object* _init_l_Lean_Server_initFn____x40_Lean_Server_Requests___hyg_575____closed__1() { _start: { diff --git a/stage0/stdlib/Lean/Server/Rpc/Deriving.c b/stage0/stdlib/Lean/Server/Rpc/Deriving.c index db064ab2d5a..0600607bffb 100644 --- a/stage0/stdlib/Lean/Server/Rpc/Deriving.c +++ b/stage0/stdlib/Lean/Server/Rpc/Deriving.c @@ -25,7 +25,6 @@ static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncod LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__21(lean_object*, lean_object*, size_t, size_t); size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Server_RpcEncoding_foldWithConstructors___spec__1(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___lambda__2(lean_object*, lean_object*, 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_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__20(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__53; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_RpcEncoding_withFieldsFlattened___spec__2___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -51,12 +50,12 @@ lean_object* lean_nat_div(lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__9; LEAN_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__34___rarg(lean_object*, lean_object*, 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_Array_anyMUnsafe_any___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__14(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__20; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__238; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__228; lean_object* l_Lean_LocalDecl_userName(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_RpcEncoding_withFieldsAux___spec__1___rarg___lambda__3(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__40(lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInstance___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_RpcEncoding_withFields___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__36(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -70,9 +69,10 @@ LEAN_EXPORT lean_object* l___private_Init_Data_Array_BinSearch_0__Array_binInser static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__169; uint8_t lean_usize_dec_eq(size_t, size_t); lean_object* lean_array_uget(lean_object*, size_t); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__19(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__19(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_insert___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAux_traverse___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__8(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__22(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__20; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__2___closed__6; LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -95,15 +95,17 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Server_RpcEncoding_ static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__244; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_MetavarContext_DependsOn_dep_visitApp___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__18(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_findAtAux___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__138; +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__1; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__39; -LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___boxed(lean_object*, lean_object*, lean_object*, 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___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Server_RpcEncoding_withFields___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__109; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__65; LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_withFieldsFlattened___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_RpcEncoding_withFields___spec__2___rarg___lambda__1(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*); @@ -118,7 +120,6 @@ lean_object* l_Lean_MetavarContext_getExprAssignment_x3f(lean_object*, lean_obje LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_DiscrTree_insertCore___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__2___closed__4; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__206; -LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_withFieldsAux___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__23___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_structEq(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___lambda__2___closed__1; lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -144,13 +145,13 @@ lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__90; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAtCollisionNodeAux___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__10(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__66; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__39(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__39(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__128; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__30___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__38(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__38(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Server_RpcEncoding_withFieldsAux___spec__2___rarg___lambda__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__116; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__38___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__38___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__231; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__192; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__71; @@ -162,6 +163,7 @@ static lean_object* l_Lean_getConstInfoInduct___at___private_Lean_Server_Rpc_Der static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__5; LEAN_EXPORT uint8_t l_Std_PersistentArray_anyMAux___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__20(lean_object*, lean_object*); lean_object* l_Lean_Meta_DiscrTree_instInhabitedDiscrTree(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_findAux___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__4___boxed(lean_object*, lean_object*, lean_object*); static size_t l_Std_PersistentHashMap_findAux___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__4___closed__1; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_matchF___closed__2; @@ -179,7 +181,6 @@ LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Rpc_Derivi static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__7; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__43; lean_object* l_Lean_MessageData_ofList(lean_object*); -static lean_object* l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___closed__1; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__120; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__32; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__10; @@ -193,7 +194,6 @@ static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncod static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__2___closed__5; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__29(lean_object*, lean_object*, size_t, size_t); lean_object* l_Std_mkHashSetImp___rarg(lean_object*); -static lean_object* l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInstance___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_right(size_t, size_t); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_dispatchDeriveInstanceUnsafe___lambda__3___closed__5; @@ -209,7 +209,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_instInhabitedCtorState; lean_object* l_Lean_Elab_Term_instInhabitedTermElabM(lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__136; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_RpcEncoding_withFieldsAux___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__19___boxed(lean_object*, lean_object*, lean_object*, 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_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__196; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__37___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__75; @@ -218,7 +218,6 @@ uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__58; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__21___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__49; -extern lean_object* l_Lean_nameLitKind; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__26; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__114; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__12; @@ -230,7 +229,6 @@ static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncod LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_MetavarContext_DependsOn_dep_visit___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__41; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__32(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41(lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__79; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__34; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insert___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__6(lean_object*, lean_object*, lean_object*); @@ -266,7 +264,6 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Rpc_D static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__2___closed__1; lean_object* l_Lean_Syntax_mkApp(lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__54; -lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(size_t, size_t, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__32___lambda__1___closed__2; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__132; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__103; @@ -280,7 +277,6 @@ static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncod LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, 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_Server_RpcEncoding_withFields___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__40___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__2___closed__7; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__32___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -291,6 +287,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_Rpc static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__55; LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_withFieldsAux___at_Lean_Server_RpcEncoding_withFields___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Array_contains___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__12___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_withFields___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__23(lean_object*); uint8_t l_instDecidableNot___rarg(uint8_t); LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_withFields(lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__30; @@ -301,13 +298,16 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deri static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__208; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__32___lambda__1(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_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___lambda__1___boxed(lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__230; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__92; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_MetavarContext_DependsOn_dep_visitApp___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__42(size_t, size_t, lean_object*); lean_object* l_Lean_Core_mkFreshUserName___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_withFields___rarg___lambda__2(lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__134; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__160; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_RpcEncoding_withFieldsFlattened___spec__2___rarg___lambda__1(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*); @@ -323,6 +323,7 @@ lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_createNodes_ LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Server_RpcEncoding_withFieldsAux___spec__2___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfoInduct___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInstance___spec__1___closed__4; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__214; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__25___boxed(lean_object*, lean_object*, 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_Meta_DiscrTree_insert___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__13; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__29; @@ -331,7 +332,7 @@ static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncod lean_object* l_Lean_getStructureFieldsFlattened(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_dispatchDeriveInstanceUnsafe(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__234; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInstance___spec__4(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInstance___spec__4(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__2___closed__8; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__16; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__52; @@ -340,7 +341,7 @@ lean_object* l_Lean_Elab_Term_elabTerm(lean_object*, lean_object*, uint8_t, uint lean_object* l_Lean_Elab_Term_mkFreshBinderName___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___spec__2___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_withFieldsAux(lean_object*); lean_object* l_instInhabitedForAll__1___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__3(lean_object*, size_t, lean_object*, lean_object*, lean_object*, 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___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__3(lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__177; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_dispatchDeriveInstanceUnsafe___lambda__3___closed__4; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__24; @@ -353,10 +354,12 @@ static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncod static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__106; lean_object* l_Lean_Name_toString(lean_object*, uint8_t); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__5___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__5(size_t, lean_object*, lean_object*, lean_object*, 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___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__5(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__21(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__78; +static lean_object* l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___closed__3; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__40(lean_object*, size_t, size_t, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__91; LEAN_EXPORT lean_object* l_panic___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInstance___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__49; @@ -374,6 +377,7 @@ lean_object* l_instInhabited___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__22; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__110; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__3___closed__5; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__22___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_RpcEncoding_withFieldsAux___rarg___closed__2; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__99; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInstance___lambda__1___closed__4; @@ -395,6 +399,7 @@ static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncod static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__7; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__71; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__63; +static lean_object* l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___closed__2; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__32___closed__6; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__35___closed__2; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInstance___spec__4___lambda__1___closed__2; @@ -402,9 +407,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_Rpc static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__62; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAux_traverse___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_insertCore___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__3(lean_object*, lean_object*, lean_object*); -lean_object* l___private_Init_Meta_0__Lean_quoteNameMk(lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__166; -static lean_object* l_Lean_Server_RpcEncoding_initFn____x40_Lean_Server_Rpc_Deriving___hyg_8525____closed__1; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__16; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Server_RpcEncoding_withFields___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_withFields___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -421,9 +424,9 @@ static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncod static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__117; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__224; LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_foldWithConstructors___rarg___lambda__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__24___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getStructureFields(lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__66; +static lean_object* l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___closed__1; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__56; lean_object* l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__25; @@ -434,7 +437,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_withFieldsFlattened___rarg___ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__32___closed__4; lean_object* lean_array_to_list(lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__105; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___boxed(lean_object*, lean_object*, lean_object*, 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_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___boxed(lean_object*, lean_object*, 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_Meta_withLocalDecl___at_Lean_Server_RpcEncoding_withFieldsAux___spec__2(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__1; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__221; @@ -468,7 +471,6 @@ lean_object* l_Array_reverse___rarg(lean_object*); uint8_t l_Array_isEmpty___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__37(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__76; -extern lean_object* l_Lean_instInhabitedSyntax; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__176; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Server_RpcEncoding_withFieldsAux___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__18; @@ -478,6 +480,7 @@ static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncod static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__35; lean_object* l_Lean_Elab_Command_elabCommand(lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_mul(size_t, size_t); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInstance___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__32___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__3___closed__6; @@ -489,7 +492,6 @@ static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncod LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_withFieldsAux___at_Lean_Server_RpcEncoding_withFields___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__171; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__89; -LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_withFieldsAux___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__23___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__3___closed__3; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__157; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__4; @@ -497,11 +499,10 @@ static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncod static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInstance___lambda__3___closed__2; lean_object* l_Lean_Elab_Command_liftCoreM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInstance___lambda__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_withFields___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__22(lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__2; size_t lean_usize_of_nat(lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__50; -LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, 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___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_List_forIn_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__34(lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__162; lean_object* l_Lean_addTrace___at___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTerm___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -526,7 +527,6 @@ static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncod LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_RpcEncoding_withFieldsAux___spec__1___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__7; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__141; -static lean_object* l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___lambda__1___closed__1; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__226; lean_object* l_Lean_Meta_DiscrTree_mkPath(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___closed__5; @@ -539,24 +539,23 @@ lean_object* l_Lean_Meta_mkProjection(lean_object*, lean_object*, lean_object*, LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_dispatchDeriveInstanceUnsafe___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__53; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__3; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_delab(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__142; -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg(lean_object*, 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_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44(lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_find_x3f___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__4(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_RpcEncoding_withFields___spec__2___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_withFields___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__32___closed__5; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__213; -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___lambda__1___boxed(lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__5___closed__2; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_dispatchDeriveInstanceUnsafe___lambda__3___closed__1; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__184; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__229; extern uint8_t l_Lean_instInhabitedBinderInfo; -LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_withFieldsAux___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__23(lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__119; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_MetavarContext_DependsOn_dep_visitMain___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__17(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg(lean_object*, 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_Server_RpcEncoding_withFieldsAux___at_Lean_Server_RpcEncoding_withFieldsFlattened___spec__1(lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__29; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -564,13 +563,14 @@ static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncod static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__17; uint8_t l_Array_contains___at___private_Lean_Meta_Match_Value_0__Lean_Meta_isUIntTypeName___spec__1(lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__5; +LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_withFieldsAux___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__24___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_getString_x21(lean_object*); lean_object* l_String_intercalate(lean_object*, lean_object*); uint8_t l_Lean_Expr_isType(lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__44; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__25; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__57; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___boxed(lean_object*, lean_object*, 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_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__46; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAux_traverse___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__9(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_DiscrTree_insertCore___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__2___closed__2; @@ -591,7 +591,6 @@ lean_object* lean_environment_main_module(lean_object*); uint8_t lean_expr_eqv(lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__173; static lean_object* l_Lean_Meta_DiscrTree_insertCore___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__2___closed__3; -lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkInductiveApp___spec__1(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_MetavarContext_DependsOn_dep_visitMain___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__25___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isMVar(lean_object*); lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -606,6 +605,7 @@ lean_object* l_Lean_Parser_categoryParser(lean_object*, lean_object*); uint8_t l_Lean_Expr_hasMVar(lean_object*); static lean_object* l_Lean_Server_RpcEncoding_withFields___rarg___closed__1; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__74; +LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_withFieldsAux___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__24___rarg___lambda__1(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_Array_binInsertM___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___closed__2; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__47; @@ -615,6 +615,7 @@ static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncod LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInstance(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__197; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__64; +lean_object* l_Lean_quoteNameMk(lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__235; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__144; LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_CtorState_uniqEncArgTypes___default; @@ -646,7 +647,6 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_RpcEncoding_w uint8_t l_Lean_Meta_DiscrTree_Key_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__35___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_getMainModule___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___closed__3; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__198; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_find_x3f___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_RpcEncoding_withFieldsFlattened___spec__2(lean_object*); @@ -656,13 +656,15 @@ LEAN_EXPORT lean_object* l_Array_binInsertM___at___private_Lean_Server_Rpc_Deriv uint8_t l_Array_contains___at___private_Lean_Class_0__Lean_checkOutParam___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__8; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_MetavarContext_DependsOn_dep_visit___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__24(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4(size_t, lean_object*, lean_object*, 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___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___closed__8; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__21___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__220; lean_object* l_Lean_Elab_Command_getRef(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__195; +lean_object* l_Lean_Syntax_mkNameLit(lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInstance___lambda__1___closed__2; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__40___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__96; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__216; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__72; @@ -672,6 +674,7 @@ uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_BinSearch_0__Array_binInsertAux___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkForall(lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_foldWithConstructors(lean_object*); +static lean_object* l_Lean_Server_RpcEncoding_initFn____x40_Lean_Server_Rpc_Deriving___hyg_8911____closed__1; static size_t l_Std_PersistentHashMap_findAux___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__4___closed__2; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_findAux___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__4(lean_object*, size_t, lean_object*); extern lean_object* l_Lean_instInhabitedName; @@ -699,8 +702,11 @@ static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncod static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__59; static lean_object* l_panic___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInstance___spec__3___closed__1; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insert___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__7(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__43___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___closed__6; +LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_withFields___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__23___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__209; +static lean_object* l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___closed__4; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___lambda__1___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_CtorState_ctors___default; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Server_RpcEncoding_withFieldsFlattened___spec__3(lean_object*); @@ -718,6 +724,7 @@ LEAN_EXPORT lean_object* l___private_Init_Data_Array_BinSearch_0__Array_binInser static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__11; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__104; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__2___closed__2; +LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_withFieldsAux___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__24(lean_object*); lean_object* l_Lean_Meta_forallTelescopeReducing___at_Lean_Elab_Deriving_mkInductArgNames___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__183; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__59; @@ -729,8 +736,8 @@ static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncod LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_withFieldsAux___at_Lean_Server_RpcEncoding_withFieldsFlattened___spec__1___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__17; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__8; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__94; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__24(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_indentD(lean_object*); static lean_object* l_Lean_Server_RpcEncoding_isOptField___closed__1; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__2___closed__9; @@ -745,7 +752,6 @@ lean_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__245; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Server_RpcEncoding_withFields___spec__3(lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___closed__3; -static lean_object* l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___closed__4; lean_object* lean_mk_syntax_ident(lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__233; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__34; @@ -756,7 +762,7 @@ static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncod static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Array_binInsertM___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__14; -LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_withFields___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__139; @@ -789,6 +795,7 @@ static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncod LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_MetavarContext_DependsOn_dep_visitMain___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__25(lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_levelOne; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__143; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__42___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___closed__11; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInstance___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__87; @@ -797,17 +804,17 @@ uint8_t l_Lean_Expr_hasFVar(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_MetavarContext_DependsOn_dep_visit___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__16(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__2___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__43(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Server_RpcEncoding_withFieldsFlattened___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__27; lean_object* l_Lean_mkConst(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__10___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_dispatchDeriveInstanceUnsafe___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__11; -lean_object* l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__10(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_insertAt___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__25(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___lambda__1___closed__2; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__187; -LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_withFields___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__22___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__159; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_dispatchDeriveInstanceUnsafe___lambda__3___closed__6; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__45; @@ -832,23 +839,21 @@ lean_object* l_Lean_Meta_DiscrTree_instInhabitedTrie(lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__194; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__40; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__67; -LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_initFn____x40_Lean_Server_Rpc_Deriving___hyg_8525_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_initFn____x40_Lean_Server_Rpc_Deriving___hyg_8911_(lean_object*); lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__2___closed__4; uint8_t l_Lean_isStructure(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInstance___spec__4___boxed(lean_object*, lean_object*, 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_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInstance___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__122; LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__44; -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___lambda__1(lean_object*); lean_object* l_Std_PersistentHashMap_mkCollisionNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_withFieldsFlattened(lean_object*); lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); -lean_object* l_Lean_Syntax_mkLit(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInstance___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__9; lean_object* l___private_Lean_MetavarContext_0__Lean_MetavarContext_DependsOn_shouldVisit(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__39___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__39___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); lean_object* l_Lean_Meta_getFVarLocalDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_RpcEncoding_withFieldsAux___spec__1___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -4256,13 +4261,13 @@ x_470 = lean_array_push(x_469, x_70); if (lean_obj_tag(x_183) == 0) { lean_object* x_524; -x_524 = l___private_Init_Meta_0__Lean_quoteNameMk(x_1); +x_524 = l_Lean_quoteNameMk(x_1); x_471 = x_524; goto block_523; } else { -lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; +lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_dec(x_1); x_525 = lean_ctor_get(x_183, 0); lean_inc(x_525); @@ -4272,15 +4277,14 @@ x_527 = l_String_intercalate(x_526, x_525); x_528 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__246; x_529 = lean_string_append(x_528, x_527); lean_dec(x_527); -x_530 = l_Lean_nameLitKind; -x_531 = l_Lean_Syntax_mkLit(x_530, x_529, x_21); -x_532 = lean_array_push(x_33, x_531); -x_533 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__244; -x_534 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_534, 0, x_21); -lean_ctor_set(x_534, 1, x_533); -lean_ctor_set(x_534, 2, x_532); -x_471 = x_534; +x_530 = l_Lean_Syntax_mkNameLit(x_529, x_21); +x_531 = lean_array_push(x_33, x_530); +x_532 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__244; +x_533 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_533, 0, x_21); +lean_ctor_set(x_533, 1, x_532); +lean_ctor_set(x_533, 2, x_531); +x_471 = x_533; goto block_523; } block_523: @@ -7272,300 +7276,282 @@ x_3 = lean_array_push(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___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_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, lean_object* x_16) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___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_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) { _start: { -lean_object* x_17; lean_object* x_18; lean_object* x_19; -lean_dec(x_9); +lean_object* x_16; lean_object* x_17; lean_object* x_18; +lean_dec(x_8); lean_inc(x_1); -x_17 = lean_mk_syntax_ident(x_1); -x_18 = lean_array_push(x_6, x_17); -lean_inc(x_15); +x_16 = lean_mk_syntax_ident(x_1); +x_17 = lean_array_push(x_5, x_16); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -lean_inc(x_7); -lean_inc(x_5); -x_19 = l_Lean_Meta_DiscrTree_getMatch___rarg(x_5, x_7, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_19) == 0) +lean_inc(x_11); +lean_inc(x_6); +lean_inc(x_4); +x_18 = l_Lean_Meta_DiscrTree_getMatch___rarg(x_4, x_6, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_19, 0); +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); lean_inc(x_20); -x_21 = lean_ctor_get(x_19, 1); -lean_inc(x_21); +lean_dec(x_18); +x_21 = l_Array_back_x3f___rarg(x_19); lean_dec(x_19); -x_22 = l_Array_back_x3f___rarg(x_20); -lean_dec(x_20); -if (lean_obj_tag(x_22) == 0) +if (lean_obj_tag(x_21) == 0) { -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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_inc(x_1); -x_23 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_1, x_14, x_15, x_21); -x_24 = lean_ctor_get(x_23, 0); +x_22 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_1, x_13, x_14, x_20); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); -lean_inc(x_25); -lean_dec(x_23); -x_26 = lean_mk_syntax_ident(x_24); -lean_inc(x_14); -x_27 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_14, x_15, x_25); -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); -lean_dec(x_27); -x_30 = lean_st_ref_get(x_15, x_29); -x_31 = lean_ctor_get(x_30, 1); -lean_inc(x_31); -lean_dec(x_30); -x_32 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__80; +lean_dec(x_22); +x_25 = lean_mk_syntax_ident(x_23); +lean_inc(x_13); +x_26 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_13, x_14, x_24); +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); lean_inc(x_28); -x_33 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_33, 0, x_28); -lean_ctor_set(x_33, 1, x_32); -x_34 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__25; -lean_inc(x_26); -x_35 = lean_array_push(x_34, x_26); -x_36 = lean_box(2); -x_37 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__2; -x_38 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -lean_ctor_set(x_38, 2, x_35); -x_39 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_37); -lean_ctor_set(x_39, 2, x_2); -x_40 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__92; -x_41 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_41, 0, x_28); -lean_ctor_set(x_41, 1, x_40); -x_42 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__93; -x_43 = lean_array_push(x_42, x_33); -x_44 = lean_array_push(x_43, x_38); -lean_inc(x_39); -x_45 = lean_array_push(x_44, x_39); -lean_inc(x_39); +lean_dec(x_26); +x_29 = lean_st_ref_get(x_14, x_28); +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +lean_dec(x_29); +x_31 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__80; +lean_inc(x_27); +x_32 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_32, 0, x_27); +lean_ctor_set(x_32, 1, x_31); +x_33 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__25; +lean_inc(x_25); +x_34 = lean_array_push(x_33, x_25); +x_35 = lean_box(2); +x_36 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__2; +x_37 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +lean_ctor_set(x_37, 2, x_34); +x_38 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__92; +x_39 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_39, 0, x_27); +lean_ctor_set(x_39, 1, x_38); +x_40 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__93; +x_41 = lean_array_push(x_40, x_32); +x_42 = lean_array_push(x_41, x_37); +x_43 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__12; +x_44 = lean_array_push(x_42, x_43); +x_45 = lean_array_push(x_44, x_43); x_46 = lean_array_push(x_45, x_39); -x_47 = lean_array_push(x_46, x_41); -x_48 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__79; -x_49 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_49, 0, x_36); -lean_ctor_set(x_49, 1, x_48); -lean_ctor_set(x_49, 2, x_47); -x_50 = lean_array_push(x_3, x_49); -x_51 = lean_box(0); -lean_inc(x_15); +x_47 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__79; +x_48 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_48, 0, x_35); +lean_ctor_set(x_48, 1, x_47); +lean_ctor_set(x_48, 2, x_46); +x_49 = lean_array_push(x_2, x_48); +x_50 = lean_box(0); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -lean_inc(x_7); -x_52 = l_Lean_PrettyPrinter_delab(x_7, x_51, x_12, x_13, x_14, x_15, x_31); -if (lean_obj_tag(x_52) == 0) +lean_inc(x_11); +lean_inc(x_6); +x_51 = l_Lean_PrettyPrinter_delab(x_6, x_50, x_11, x_12, x_13, x_14, x_30); +if (lean_obj_tag(x_51) == 0) { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; -x_53 = lean_ctor_get(x_52, 0); +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_51, 1); lean_inc(x_53); -x_54 = lean_ctor_get(x_52, 1); -lean_inc(x_54); -lean_dec(x_52); -lean_inc(x_14); -x_55 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_14, x_15, x_54); -x_56 = lean_ctor_get(x_55, 0); -lean_inc(x_56); -x_57 = lean_ctor_get(x_55, 1); -lean_inc(x_57); -lean_dec(x_55); -x_58 = lean_st_ref_get(x_15, x_57); -x_59 = lean_ctor_get(x_58, 1); -lean_inc(x_59); -lean_dec(x_58); -x_60 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__59; +lean_dec(x_51); +lean_inc(x_13); +x_54 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_13, x_14, x_53); +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_54, 1); lean_inc(x_56); -x_61 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_61, 0, x_56); -lean_ctor_set(x_61, 1, x_60); -x_62 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__13; -x_63 = lean_array_push(x_62, x_53); -lean_inc(x_26); -x_64 = lean_array_push(x_63, x_26); -x_65 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_65, 0, x_36); -lean_ctor_set(x_65, 1, x_37); -lean_ctor_set(x_65, 2, x_64); -x_66 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___lambda__2___closed__2; -x_67 = lean_array_push(x_66, x_65); -x_68 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__61; -x_69 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_69, 0, x_36); -lean_ctor_set(x_69, 1, x_68); -lean_ctor_set(x_69, 2, x_67); -x_70 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__68; -x_71 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_71, 0, x_56); -lean_ctor_set(x_71, 1, x_70); -x_72 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__35; -x_73 = lean_array_push(x_72, x_61); -x_74 = lean_array_push(x_73, x_39); -x_75 = lean_array_push(x_74, x_69); -x_76 = lean_array_push(x_75, x_71); -x_77 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__58; -x_78 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_78, 0, x_36); -lean_ctor_set(x_78, 1, x_77); -lean_ctor_set(x_78, 2, x_76); -x_79 = lean_array_push(x_50, x_78); -lean_inc(x_15); +lean_dec(x_54); +x_57 = lean_st_ref_get(x_14, x_56); +x_58 = lean_ctor_get(x_57, 1); +lean_inc(x_58); +lean_dec(x_57); +x_59 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__59; +lean_inc(x_55); +x_60 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_60, 0, x_55); +lean_ctor_set(x_60, 1, x_59); +x_61 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__13; +x_62 = lean_array_push(x_61, x_52); +lean_inc(x_25); +x_63 = lean_array_push(x_62, x_25); +x_64 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_64, 0, x_35); +lean_ctor_set(x_64, 1, x_36); +lean_ctor_set(x_64, 2, x_63); +x_65 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___lambda__2___closed__2; +x_66 = lean_array_push(x_65, x_64); +x_67 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__61; +x_68 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_68, 0, x_35); +lean_ctor_set(x_68, 1, x_67); +lean_ctor_set(x_68, 2, x_66); +x_69 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__68; +x_70 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_70, 0, x_55); +lean_ctor_set(x_70, 1, x_69); +x_71 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__35; +x_72 = lean_array_push(x_71, x_60); +x_73 = lean_array_push(x_72, x_43); +x_74 = lean_array_push(x_73, x_68); +x_75 = lean_array_push(x_74, x_70); +x_76 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__58; +x_77 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_77, 0, x_35); +lean_ctor_set(x_77, 1, x_76); +lean_ctor_set(x_77, 2, x_75); +x_78 = lean_array_push(x_49, x_77); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -lean_inc(x_26); -x_80 = l_Lean_Meta_DiscrTree_insert___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__1(x_5, x_7, x_26, x_12, x_13, x_14, x_15, x_59); -if (lean_obj_tag(x_80) == 0) +lean_inc(x_11); +lean_inc(x_25); +x_79 = l_Lean_Meta_DiscrTree_insert___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__1(x_4, x_6, x_25, x_11, x_12, x_13, x_14, x_58); +if (lean_obj_tag(x_79) == 0) { -lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_81 = lean_ctor_get(x_80, 0); +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_80 = lean_ctor_get(x_79, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_79, 1); lean_inc(x_81); -x_82 = lean_ctor_get(x_80, 1); -lean_inc(x_82); -lean_dec(x_80); -lean_inc(x_26); -x_83 = lean_array_push(x_8, x_26); -x_84 = lean_box(0); -x_85 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___lambda__1(x_1, x_18, x_79, x_26, x_4, x_81, x_83, x_84, x_10, x_11, x_12, x_13, x_14, x_15, x_82); -lean_dec(x_15); -lean_dec(x_13); +lean_dec(x_79); +lean_inc(x_25); +x_82 = lean_array_push(x_7, x_25); +x_83 = lean_box(0); +x_84 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___lambda__1(x_1, x_17, x_78, x_25, x_3, x_80, x_82, x_83, x_9, x_10, x_11, x_12, x_13, x_14, x_81); +lean_dec(x_14); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); -return x_85; +return x_84; } else { -uint8_t x_86; -lean_dec(x_79); -lean_dec(x_26); -lean_dec(x_18); -lean_dec(x_15); +uint8_t x_85; +lean_dec(x_78); +lean_dec(x_25); +lean_dec(x_17); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_8); -lean_dec(x_4); +lean_dec(x_7); +lean_dec(x_3); lean_dec(x_1); -x_86 = !lean_is_exclusive(x_80); -if (x_86 == 0) +x_85 = !lean_is_exclusive(x_79); +if (x_85 == 0) { -return x_80; +return x_79; } else { -lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_87 = lean_ctor_get(x_80, 0); -x_88 = lean_ctor_get(x_80, 1); -lean_inc(x_88); +lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_86 = lean_ctor_get(x_79, 0); +x_87 = lean_ctor_get(x_79, 1); lean_inc(x_87); -lean_dec(x_80); -x_89 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_89, 0, x_87); -lean_ctor_set(x_89, 1, x_88); -return x_89; +lean_inc(x_86); +lean_dec(x_79); +x_88 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_88, 0, x_86); +lean_ctor_set(x_88, 1, x_87); +return x_88; } } } else { -uint8_t x_90; -lean_dec(x_50); -lean_dec(x_39); -lean_dec(x_26); -lean_dec(x_18); -lean_dec(x_15); +uint8_t x_89; +lean_dec(x_49); +lean_dec(x_25); +lean_dec(x_17); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_8); lean_dec(x_7); -lean_dec(x_5); +lean_dec(x_6); lean_dec(x_4); +lean_dec(x_3); lean_dec(x_1); -x_90 = !lean_is_exclusive(x_52); -if (x_90 == 0) +x_89 = !lean_is_exclusive(x_51); +if (x_89 == 0) { -return x_52; +return x_51; } else { -lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_91 = lean_ctor_get(x_52, 0); -x_92 = lean_ctor_get(x_52, 1); -lean_inc(x_92); +lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_90 = lean_ctor_get(x_51, 0); +x_91 = lean_ctor_get(x_51, 1); lean_inc(x_91); -lean_dec(x_52); -x_93 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_93, 0, x_91); -lean_ctor_set(x_93, 1, x_92); -return x_93; +lean_inc(x_90); +lean_dec(x_51); +x_92 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_91); +return x_92; } } } else { -lean_object* x_94; lean_object* x_95; lean_object* x_96; -lean_dec(x_7); -lean_dec(x_2); -x_94 = lean_ctor_get(x_22, 0); -lean_inc(x_94); -lean_dec(x_22); -x_95 = lean_box(0); -x_96 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___lambda__1(x_1, x_18, x_3, x_94, x_4, x_5, x_8, x_95, x_10, x_11, x_12, x_13, x_14, x_15, x_21); -lean_dec(x_15); -lean_dec(x_13); +lean_object* x_93; lean_object* x_94; lean_object* x_95; +lean_dec(x_6); +x_93 = lean_ctor_get(x_21, 0); +lean_inc(x_93); +lean_dec(x_21); +x_94 = lean_box(0); +x_95 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___lambda__1(x_1, x_17, x_2, x_93, x_3, x_4, x_7, x_94, x_9, x_10, x_11, x_12, x_13, x_14, x_20); +lean_dec(x_14); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); -return x_96; +return x_95; } } else { -uint8_t x_97; -lean_dec(x_18); -lean_dec(x_15); +uint8_t x_96; +lean_dec(x_17); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_8); lean_dec(x_7); -lean_dec(x_5); +lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_97 = !lean_is_exclusive(x_19); -if (x_97 == 0) +x_96 = !lean_is_exclusive(x_18); +if (x_96 == 0) { -return x_19; +return x_18; } else { -lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_98 = lean_ctor_get(x_19, 0); -x_99 = lean_ctor_get(x_19, 1); -lean_inc(x_99); +lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_97 = lean_ctor_get(x_18, 0); +x_98 = lean_ctor_get(x_18, 1); lean_inc(x_98); -lean_dec(x_19); -x_100 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_100, 0, x_98); -lean_ctor_set(x_100, 1, x_99); -return x_100; +lean_inc(x_97); +lean_dec(x_18); +x_99 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_99, 0, x_97); +lean_ctor_set(x_99, 1, x_98); +return x_99; } } } @@ -7692,336 +7678,325 @@ x_2 = l_Lean_indentD(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17(lean_object* x_1, size_t x_2, size_t 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, lean_object* x_11) { _start: { -uint8_t x_13; -x_13 = lean_usize_dec_lt(x_4, x_3); -if (x_13 == 0) +uint8_t x_12; +x_12 = lean_usize_dec_lt(x_3, x_2); +if (x_12 == 0) { -lean_object* x_14; -lean_dec(x_11); +lean_object* x_13; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_1); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_5); -lean_ctor_set(x_14, 1, x_12); -return x_14; +lean_dec(x_5); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_4); +lean_ctor_set(x_13, 1, x_11); +return x_13; } else { -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; lean_object* x_25; lean_object* x_26; uint8_t x_27; -x_15 = lean_array_uget(x_2, x_4); -x_16 = lean_ctor_get(x_15, 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; lean_object* x_25; uint8_t x_26; +x_14 = lean_array_uget(x_1, x_3); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +lean_dec(x_14); +x_17 = lean_ctor_get(x_4, 0); lean_inc(x_17); -lean_dec(x_15); -x_18 = lean_ctor_get(x_5, 0); +x_18 = lean_ctor_get(x_4, 1); lean_inc(x_18); -x_19 = lean_ctor_get(x_5, 1); +lean_dec(x_4); +x_19 = lean_ctor_get(x_18, 0); lean_inc(x_19); -lean_dec(x_5); -x_20 = lean_ctor_get(x_19, 0); +x_20 = lean_ctor_get(x_18, 1); lean_inc(x_20); -x_21 = lean_ctor_get(x_19, 1); +lean_dec(x_18); +x_21 = lean_ctor_get(x_20, 0); lean_inc(x_21); -lean_dec(x_19); -x_22 = lean_ctor_get(x_21, 0); +x_22 = lean_ctor_get(x_20, 1); lean_inc(x_22); -x_23 = lean_ctor_get(x_21, 1); +lean_dec(x_20); +x_23 = lean_ctor_get(x_22, 0); lean_inc(x_23); -lean_dec(x_21); -x_24 = lean_ctor_get(x_23, 0); +x_24 = lean_ctor_get(x_22, 1); lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); -lean_inc(x_25); -lean_dec(x_23); -lean_inc(x_1); -lean_inc(x_16); -x_26 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___lambda__2), 16, 2); -lean_closure_set(x_26, 0, x_16); -lean_closure_set(x_26, 1, x_1); -lean_inc(x_16); -x_27 = l_Lean_Server_RpcEncoding_isOptField(x_16); -if (x_27 == 0) +lean_dec(x_22); +lean_inc(x_15); +x_25 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___lambda__2___boxed), 15, 1); +lean_closure_set(x_25, 0, x_15); +lean_inc(x_15); +x_26 = l_Lean_Server_RpcEncoding_isOptField(x_15); +if (x_26 == 0) { -lean_object* x_28; lean_object* x_29; -lean_dec(x_26); -x_28 = lean_box(0); -lean_inc(x_11); +lean_object* x_27; lean_object* x_28; +lean_dec(x_25); +x_27 = lean_box(0); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_1); -x_29 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___lambda__2(x_16, x_1, x_18, x_20, x_22, x_24, x_17, x_25, x_28, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_29) == 0) +x_28 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___lambda__2(x_15, x_17, x_19, x_21, x_23, x_16, x_24, x_27, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_28) == 0) { -lean_object* x_30; -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -if (lean_obj_tag(x_30) == 0) +lean_object* x_29; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +if (lean_obj_tag(x_29) == 0) { -uint8_t x_31; -lean_dec(x_11); +uint8_t x_30; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_1); -x_31 = !lean_is_exclusive(x_29); -if (x_31 == 0) +lean_dec(x_5); +x_30 = !lean_is_exclusive(x_28); +if (x_30 == 0) { -lean_object* x_32; lean_object* x_33; +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_28, 0); +lean_dec(x_31); x_32 = lean_ctor_get(x_29, 0); -lean_dec(x_32); -x_33 = lean_ctor_get(x_30, 0); -lean_inc(x_33); -lean_dec(x_30); -lean_ctor_set(x_29, 0, x_33); -return x_29; +lean_inc(x_32); +lean_dec(x_29); +lean_ctor_set(x_28, 0, x_32); +return x_28; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_29, 1); +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_28, 1); +lean_inc(x_33); +lean_dec(x_28); +x_34 = lean_ctor_get(x_29, 0); lean_inc(x_34); lean_dec(x_29); -x_35 = lean_ctor_get(x_30, 0); -lean_inc(x_35); -lean_dec(x_30); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_34); -return x_36; +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_33); +return x_35; } } else { -lean_object* x_37; lean_object* x_38; size_t x_39; size_t x_40; -x_37 = lean_ctor_get(x_29, 1); +lean_object* x_36; lean_object* x_37; size_t x_38; size_t x_39; +x_36 = lean_ctor_get(x_28, 1); +lean_inc(x_36); +lean_dec(x_28); +x_37 = lean_ctor_get(x_29, 0); lean_inc(x_37); lean_dec(x_29); -x_38 = lean_ctor_get(x_30, 0); -lean_inc(x_38); -lean_dec(x_30); -x_39 = 1; -x_40 = lean_usize_add(x_4, x_39); -x_4 = x_40; -x_5 = x_38; -x_12 = x_37; +x_38 = 1; +x_39 = lean_usize_add(x_3, x_38); +x_3 = x_39; +x_4 = x_37; +x_11 = x_36; goto _start; } } else { -uint8_t x_42; -lean_dec(x_11); +uint8_t x_41; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_1); -x_42 = !lean_is_exclusive(x_29); -if (x_42 == 0) +lean_dec(x_5); +x_41 = !lean_is_exclusive(x_28); +if (x_41 == 0) { -return x_29; +return x_28; } else { -lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_43 = lean_ctor_get(x_29, 0); -x_44 = lean_ctor_get(x_29, 1); -lean_inc(x_44); +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_28, 0); +x_43 = lean_ctor_get(x_28, 1); lean_inc(x_43); -lean_dec(x_29); -x_45 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_45, 0, x_43); -lean_ctor_set(x_45, 1, x_44); -return x_45; +lean_inc(x_42); +lean_dec(x_28); +x_44 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_44, 0, x_42); +lean_ctor_set(x_44, 1, x_43); +return x_44; } } } else { -lean_object* x_46; uint8_t x_47; -x_46 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___lambda__1___closed__4; -x_47 = l_Lean_Expr_isAppOf(x_17, x_46); -if (x_47 == 0) +lean_object* x_45; uint8_t x_46; +x_45 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___lambda__1___closed__4; +x_46 = l_Lean_Expr_isAppOf(x_16, x_45); +if (x_46 == 0) { -lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; uint8_t x_65; -lean_dec(x_26); +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; lean_dec(x_25); lean_dec(x_24); -lean_dec(x_22); -lean_dec(x_20); -lean_dec(x_18); -lean_dec(x_1); -x_48 = lean_alloc_ctor(4, 1, 0); -lean_ctor_set(x_48, 0, x_16); -x_49 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___closed__2; -x_50 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_50, 0, x_49); -lean_ctor_set(x_50, 1, x_48); -x_51 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___closed__4; -x_52 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_52, 0, x_50); -lean_ctor_set(x_52, 1, x_51); -x_53 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_53, 0, x_17); -x_54 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___closed__6; +lean_dec(x_23); +lean_dec(x_21); +lean_dec(x_19); +lean_dec(x_17); +x_47 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_47, 0, x_15); +x_48 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___closed__2; +x_49 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_47); +x_50 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___closed__4; +x_51 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +x_52 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_52, 0, x_16); +x_53 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___closed__6; +x_54 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_54, 0, x_53); +lean_ctor_set(x_54, 1, x_52); x_55 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_55, 0, x_54); lean_ctor_set(x_55, 1, x_53); -x_56 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_56, 0, x_55); -lean_ctor_set(x_56, 1, x_54); -x_57 = l_Lean_indentD(x_56); -x_58 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_58, 0, x_52); -lean_ctor_set(x_58, 1, x_57); -x_59 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___closed__8; -x_60 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_60, 0, x_58); -lean_ctor_set(x_60, 1, x_59); -x_61 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___closed__12; +x_56 = l_Lean_indentD(x_55); +x_57 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_57, 0, x_51); +lean_ctor_set(x_57, 1, x_56); +x_58 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___closed__8; +x_59 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_58); +x_60 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___closed__12; +x_61 = lean_alloc_ctor(10, 2, 0); +lean_ctor_set(x_61, 0, x_59); +lean_ctor_set(x_61, 1, x_60); x_62 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_62, 0, x_60); -lean_ctor_set(x_62, 1, x_61); -x_63 = lean_alloc_ctor(10, 2, 0); -lean_ctor_set(x_63, 0, x_62); -lean_ctor_set(x_63, 1, x_54); -x_64 = l_Lean_throwError___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__1(x_63, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); +lean_ctor_set(x_62, 0, x_61); +lean_ctor_set(x_62, 1, x_53); +x_63 = l_Lean_throwError___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__1(x_62, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -x_65 = !lean_is_exclusive(x_64); -if (x_65 == 0) +lean_dec(x_6); +x_64 = !lean_is_exclusive(x_63); +if (x_64 == 0) { -return x_64; +return x_63; } else { -lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_66 = lean_ctor_get(x_64, 0); -x_67 = lean_ctor_get(x_64, 1); -lean_inc(x_67); +lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = lean_ctor_get(x_63, 0); +x_66 = lean_ctor_get(x_63, 1); lean_inc(x_66); -lean_dec(x_64); -x_68 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_68, 0, x_66); -lean_ctor_set(x_68, 1, x_67); -return x_68; +lean_inc(x_65); +lean_dec(x_63); +x_67 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_67, 0, x_65); +lean_ctor_set(x_67, 1, x_66); +return x_67; } } else { -lean_object* x_69; lean_object* x_70; -lean_dec(x_16); -x_69 = lean_box(0); -lean_inc(x_11); +lean_object* x_68; lean_object* x_69; +lean_dec(x_15); +x_68 = lean_box(0); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_70 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___lambda__3(x_26, x_18, x_20, x_22, x_24, x_25, x_17, x_69, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_70) == 0) +lean_inc(x_5); +x_69 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___lambda__3(x_25, x_17, x_19, x_21, x_23, x_24, x_16, x_68, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_69) == 0) { -lean_object* x_71; -x_71 = lean_ctor_get(x_70, 0); -lean_inc(x_71); -if (lean_obj_tag(x_71) == 0) +lean_object* x_70; +x_70 = lean_ctor_get(x_69, 0); +lean_inc(x_70); +if (lean_obj_tag(x_70) == 0) { -uint8_t x_72; -lean_dec(x_11); +uint8_t x_71; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_1); -x_72 = !lean_is_exclusive(x_70); -if (x_72 == 0) +lean_dec(x_5); +x_71 = !lean_is_exclusive(x_69); +if (x_71 == 0) { -lean_object* x_73; lean_object* x_74; +lean_object* x_72; lean_object* x_73; +x_72 = lean_ctor_get(x_69, 0); +lean_dec(x_72); x_73 = lean_ctor_get(x_70, 0); -lean_dec(x_73); -x_74 = lean_ctor_get(x_71, 0); -lean_inc(x_74); -lean_dec(x_71); -lean_ctor_set(x_70, 0, x_74); -return x_70; +lean_inc(x_73); +lean_dec(x_70); +lean_ctor_set(x_69, 0, x_73); +return x_69; } else { -lean_object* x_75; lean_object* x_76; lean_object* x_77; -x_75 = lean_ctor_get(x_70, 1); +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_69, 1); +lean_inc(x_74); +lean_dec(x_69); +x_75 = lean_ctor_get(x_70, 0); lean_inc(x_75); lean_dec(x_70); -x_76 = lean_ctor_get(x_71, 0); -lean_inc(x_76); -lean_dec(x_71); -x_77 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_77, 0, x_76); -lean_ctor_set(x_77, 1, x_75); -return x_77; +x_76 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_76, 0, x_75); +lean_ctor_set(x_76, 1, x_74); +return x_76; } } else { -lean_object* x_78; lean_object* x_79; size_t x_80; size_t x_81; -x_78 = lean_ctor_get(x_70, 1); +lean_object* x_77; lean_object* x_78; size_t x_79; size_t x_80; +x_77 = lean_ctor_get(x_69, 1); +lean_inc(x_77); +lean_dec(x_69); +x_78 = lean_ctor_get(x_70, 0); lean_inc(x_78); lean_dec(x_70); -x_79 = lean_ctor_get(x_71, 0); -lean_inc(x_79); -lean_dec(x_71); -x_80 = 1; -x_81 = lean_usize_add(x_4, x_80); -x_4 = x_81; -x_5 = x_79; -x_12 = x_78; +x_79 = 1; +x_80 = lean_usize_add(x_3, x_79); +x_3 = x_80; +x_4 = x_78; +x_11 = x_77; goto _start; } } else { -uint8_t x_83; -lean_dec(x_11); +uint8_t x_82; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_1); -x_83 = !lean_is_exclusive(x_70); -if (x_83 == 0) +lean_dec(x_5); +x_82 = !lean_is_exclusive(x_69); +if (x_82 == 0) { -return x_70; +return x_69; } else { -lean_object* x_84; lean_object* x_85; lean_object* x_86; -x_84 = lean_ctor_get(x_70, 0); -x_85 = lean_ctor_get(x_70, 1); -lean_inc(x_85); +lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_83 = lean_ctor_get(x_69, 0); +x_84 = lean_ctor_get(x_69, 1); lean_inc(x_84); -lean_dec(x_70); -x_86 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_86, 0, x_84); -lean_ctor_set(x_86, 1, x_85); -return x_86; +lean_inc(x_83); +lean_dec(x_69); +x_85 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_85, 0, x_83); +lean_ctor_set(x_85, 1, x_84); +return x_85; } } } @@ -8094,292 +8069,282 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18(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_object* x_7, 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_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -uint8_t x_14; -x_14 = lean_usize_dec_lt(x_5, x_4); -if (x_14 == 0) +uint8_t x_13; +x_13 = lean_usize_dec_lt(x_4, x_3); +if (x_13 == 0) { -lean_object* x_15; -lean_dec(x_11); -lean_dec(x_3); +lean_object* x_14; +lean_dec(x_10); lean_dec(x_2); lean_dec(x_1); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_6); -lean_ctor_set(x_15, 1, x_13); -return x_15; +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_5); +lean_ctor_set(x_14, 1, x_12); +return x_14; } else { -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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; size_t x_79; size_t x_80; lean_object* x_81; -x_16 = lean_array_uget(x_6, x_5); -x_17 = lean_unsigned_to_nat(0u); -x_18 = lean_array_uset(x_6, x_5, x_17); -lean_inc(x_11); -x_19 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_11, x_12, x_13); -x_20 = lean_ctor_get(x_19, 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; lean_object* 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; size_t x_78; size_t x_79; lean_object* x_80; +x_15 = lean_array_uget(x_5, x_4); +x_16 = lean_unsigned_to_nat(0u); +x_17 = lean_array_uset(x_5, x_4, x_16); +lean_inc(x_10); +x_18 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_10, x_11, x_12); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); lean_inc(x_20); -x_21 = lean_ctor_get(x_19, 1); +lean_dec(x_18); +x_21 = lean_ctor_get(x_10, 10); lean_inc(x_21); -lean_dec(x_19); -x_22 = lean_ctor_get(x_11, 10); -lean_inc(x_22); -x_23 = lean_st_ref_get(x_12, x_21); -x_24 = lean_ctor_get(x_23, 0); +x_22 = lean_st_ref_get(x_11, x_20); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); +lean_dec(x_22); +x_25 = lean_ctor_get(x_23, 0); lean_inc(x_25); lean_dec(x_23); -x_26 = lean_ctor_get(x_24, 0); -lean_inc(x_26); -lean_dec(x_24); -x_27 = lean_environment_main_module(x_26); -x_28 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__5; -lean_inc(x_2); -x_29 = lean_name_mk_string(x_2, x_28); -x_30 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__16; -x_31 = lean_name_mk_string(x_29, x_30); -x_32 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__217; -lean_inc(x_31); -x_33 = lean_name_mk_string(x_31, x_32); -x_34 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__219; -lean_inc(x_31); -x_35 = lean_name_mk_string(x_31, x_34); -x_36 = lean_box(2); -x_37 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__2; +x_26 = lean_environment_main_module(x_25); +x_27 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__5; lean_inc(x_1); -x_38 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -lean_ctor_set(x_38, 2, x_1); -x_39 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__13; -lean_inc(x_16); -x_40 = lean_array_push(x_39, x_16); -x_41 = lean_array_push(x_40, x_38); -x_42 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_42, 0, x_36); -lean_ctor_set(x_42, 1, x_35); -lean_ctor_set(x_42, 2, x_41); -x_43 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__109; -lean_inc(x_20); -x_44 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_44, 0, x_20); -lean_ctor_set(x_44, 1, x_43); -x_45 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__1; -lean_inc(x_31); -x_46 = lean_name_mk_string(x_31, x_45); -x_47 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__2; -lean_inc(x_20); -x_48 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_48, 0, x_20); -lean_ctor_set(x_48, 1, x_47); -x_49 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__60; -lean_inc(x_31); -x_50 = lean_name_mk_string(x_31, x_49); -lean_inc(x_3); -x_51 = lean_mk_syntax_ident(x_3); -x_52 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__3; -x_53 = lean_name_mk_string(x_31, x_52); -x_54 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__7; -x_55 = l_Lean_addMacroScope(x_27, x_54, x_22); -x_56 = lean_box(0); -x_57 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__6; -lean_inc(x_20); -x_58 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_58, 0, x_20); +x_28 = lean_name_mk_string(x_1, x_27); +x_29 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__16; +x_30 = lean_name_mk_string(x_28, x_29); +x_31 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__217; +lean_inc(x_30); +x_32 = lean_name_mk_string(x_30, x_31); +x_33 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__219; +lean_inc(x_30); +x_34 = lean_name_mk_string(x_30, x_33); +x_35 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__13; +lean_inc(x_15); +x_36 = lean_array_push(x_35, x_15); +x_37 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__12; +x_38 = lean_array_push(x_36, x_37); +x_39 = lean_box(2); +x_40 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_34); +lean_ctor_set(x_40, 2, x_38); +x_41 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__109; +lean_inc(x_19); +x_42 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_42, 0, x_19); +lean_ctor_set(x_42, 1, x_41); +x_43 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__1; +lean_inc(x_30); +x_44 = lean_name_mk_string(x_30, x_43); +x_45 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__2; +lean_inc(x_19); +x_46 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_46, 0, x_19); +lean_ctor_set(x_46, 1, x_45); +x_47 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__60; +lean_inc(x_30); +x_48 = lean_name_mk_string(x_30, x_47); +lean_inc(x_2); +x_49 = lean_mk_syntax_ident(x_2); +x_50 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__3; +x_51 = lean_name_mk_string(x_30, x_50); +x_52 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__7; +x_53 = l_Lean_addMacroScope(x_26, x_52, x_21); +x_54 = lean_box(0); +x_55 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__6; +lean_inc(x_19); +x_56 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_56, 0, x_19); +lean_ctor_set(x_56, 1, x_55); +lean_ctor_set(x_56, 2, x_53); +lean_ctor_set(x_56, 3, x_54); +x_57 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__245; +x_58 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_58, 0, x_19); lean_ctor_set(x_58, 1, x_57); -lean_ctor_set(x_58, 2, x_55); -lean_ctor_set(x_58, 3, x_56); -x_59 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__245; -x_60 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_60, 0, x_20); -lean_ctor_set(x_60, 1, x_59); -x_61 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__33; -x_62 = lean_array_push(x_61, x_58); -x_63 = lean_array_push(x_62, x_60); -x_64 = lean_array_push(x_63, x_16); -x_65 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_65, 0, x_36); -lean_ctor_set(x_65, 1, x_53); -lean_ctor_set(x_65, 2, x_64); -x_66 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__25; -x_67 = lean_array_push(x_66, x_65); -x_68 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_68, 0, x_36); -lean_ctor_set(x_68, 1, x_37); -lean_ctor_set(x_68, 2, x_67); -x_69 = lean_array_push(x_39, x_51); -x_70 = lean_array_push(x_69, x_68); -x_71 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_71, 0, x_36); -lean_ctor_set(x_71, 1, x_50); -lean_ctor_set(x_71, 2, x_70); -x_72 = lean_array_push(x_39, x_48); -x_73 = lean_array_push(x_72, x_71); -x_74 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_74, 0, x_36); -lean_ctor_set(x_74, 1, x_46); -lean_ctor_set(x_74, 2, x_73); -x_75 = lean_array_push(x_61, x_42); -x_76 = lean_array_push(x_75, x_44); -x_77 = lean_array_push(x_76, x_74); -x_78 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_78, 0, x_36); -lean_ctor_set(x_78, 1, x_33); -lean_ctor_set(x_78, 2, x_77); -x_79 = 1; -x_80 = lean_usize_add(x_5, x_79); -x_81 = lean_array_uset(x_18, x_5, x_78); +x_59 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__33; +x_60 = lean_array_push(x_59, x_56); +x_61 = lean_array_push(x_60, x_58); +x_62 = lean_array_push(x_61, x_15); +x_63 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_63, 0, x_39); +lean_ctor_set(x_63, 1, x_51); +lean_ctor_set(x_63, 2, x_62); +x_64 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__25; +x_65 = lean_array_push(x_64, x_63); +x_66 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__2; +x_67 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_67, 0, x_39); +lean_ctor_set(x_67, 1, x_66); +lean_ctor_set(x_67, 2, x_65); +x_68 = lean_array_push(x_35, x_49); +x_69 = lean_array_push(x_68, x_67); +x_70 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_70, 0, x_39); +lean_ctor_set(x_70, 1, x_48); +lean_ctor_set(x_70, 2, x_69); +x_71 = lean_array_push(x_35, x_46); +x_72 = lean_array_push(x_71, x_70); +x_73 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_73, 0, x_39); +lean_ctor_set(x_73, 1, x_44); +lean_ctor_set(x_73, 2, x_72); +x_74 = lean_array_push(x_59, x_40); +x_75 = lean_array_push(x_74, x_42); +x_76 = lean_array_push(x_75, x_73); +x_77 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_77, 0, x_39); +lean_ctor_set(x_77, 1, x_32); +lean_ctor_set(x_77, 2, x_76); +x_78 = 1; +x_79 = lean_usize_add(x_4, x_78); +x_80 = lean_array_uset(x_17, x_4, x_77); +x_4 = x_79; x_5 = x_80; -x_6 = x_81; -x_13 = x_25; +x_12 = x_24; goto _start; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__19(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_object* x_7, 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_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__19(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -uint8_t x_14; -x_14 = lean_usize_dec_lt(x_5, x_4); -if (x_14 == 0) +uint8_t x_13; +x_13 = lean_usize_dec_lt(x_4, x_3); +if (x_13 == 0) { -lean_object* x_15; -lean_dec(x_11); -lean_dec(x_3); +lean_object* x_14; +lean_dec(x_10); lean_dec(x_2); lean_dec(x_1); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_6); -lean_ctor_set(x_15, 1, x_13); -return x_15; +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_5); +lean_ctor_set(x_14, 1, x_12); +return x_14; } else { -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; 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; size_t x_79; size_t x_80; lean_object* x_81; -x_16 = lean_array_uget(x_6, x_5); -x_17 = lean_unsigned_to_nat(0u); -x_18 = lean_array_uset(x_6, x_5, x_17); -lean_inc(x_11); -x_19 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_11, x_12, x_13); -x_20 = lean_ctor_get(x_19, 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; lean_object* 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; size_t x_78; size_t x_79; lean_object* x_80; +x_15 = lean_array_uget(x_5, x_4); +x_16 = lean_unsigned_to_nat(0u); +x_17 = lean_array_uset(x_5, x_4, x_16); +lean_inc(x_10); +x_18 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_10, x_11, x_12); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); lean_inc(x_20); -x_21 = lean_ctor_get(x_19, 1); +lean_dec(x_18); +x_21 = lean_ctor_get(x_10, 10); lean_inc(x_21); -lean_dec(x_19); -x_22 = lean_ctor_get(x_11, 10); -lean_inc(x_22); -x_23 = lean_st_ref_get(x_12, x_21); -x_24 = lean_ctor_get(x_23, 0); +x_22 = lean_st_ref_get(x_11, x_20); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); +lean_dec(x_22); +x_25 = lean_ctor_get(x_23, 0); lean_inc(x_25); lean_dec(x_23); -x_26 = lean_ctor_get(x_24, 0); -lean_inc(x_26); -lean_dec(x_24); -x_27 = lean_environment_main_module(x_26); -x_28 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__5; -lean_inc(x_2); -x_29 = lean_name_mk_string(x_2, x_28); -x_30 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__16; -x_31 = lean_name_mk_string(x_29, x_30); -x_32 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__217; -lean_inc(x_31); -x_33 = lean_name_mk_string(x_31, x_32); -x_34 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__219; -lean_inc(x_31); -x_35 = lean_name_mk_string(x_31, x_34); -x_36 = lean_box(2); -x_37 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__2; +x_26 = lean_environment_main_module(x_25); +x_27 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__5; lean_inc(x_1); -x_38 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -lean_ctor_set(x_38, 2, x_1); -x_39 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__13; -lean_inc(x_16); -x_40 = lean_array_push(x_39, x_16); -x_41 = lean_array_push(x_40, x_38); -x_42 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_42, 0, x_36); -lean_ctor_set(x_42, 1, x_35); -lean_ctor_set(x_42, 2, x_41); -x_43 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__109; -lean_inc(x_20); -x_44 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_44, 0, x_20); -lean_ctor_set(x_44, 1, x_43); -x_45 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__1; -lean_inc(x_31); -x_46 = lean_name_mk_string(x_31, x_45); -x_47 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__2; -lean_inc(x_20); -x_48 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_48, 0, x_20); -lean_ctor_set(x_48, 1, x_47); -x_49 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__60; -lean_inc(x_31); -x_50 = lean_name_mk_string(x_31, x_49); -lean_inc(x_3); -x_51 = lean_mk_syntax_ident(x_3); -x_52 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__3; -x_53 = lean_name_mk_string(x_31, x_52); -x_54 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__7; -x_55 = l_Lean_addMacroScope(x_27, x_54, x_22); -x_56 = lean_box(0); -x_57 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__6; -lean_inc(x_20); -x_58 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_58, 0, x_20); +x_28 = lean_name_mk_string(x_1, x_27); +x_29 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__16; +x_30 = lean_name_mk_string(x_28, x_29); +x_31 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__217; +lean_inc(x_30); +x_32 = lean_name_mk_string(x_30, x_31); +x_33 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__219; +lean_inc(x_30); +x_34 = lean_name_mk_string(x_30, x_33); +x_35 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__13; +lean_inc(x_15); +x_36 = lean_array_push(x_35, x_15); +x_37 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__12; +x_38 = lean_array_push(x_36, x_37); +x_39 = lean_box(2); +x_40 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_34); +lean_ctor_set(x_40, 2, x_38); +x_41 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__109; +lean_inc(x_19); +x_42 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_42, 0, x_19); +lean_ctor_set(x_42, 1, x_41); +x_43 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__1; +lean_inc(x_30); +x_44 = lean_name_mk_string(x_30, x_43); +x_45 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__2; +lean_inc(x_19); +x_46 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_46, 0, x_19); +lean_ctor_set(x_46, 1, x_45); +x_47 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__60; +lean_inc(x_30); +x_48 = lean_name_mk_string(x_30, x_47); +lean_inc(x_2); +x_49 = lean_mk_syntax_ident(x_2); +x_50 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__3; +x_51 = lean_name_mk_string(x_30, x_50); +x_52 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__7; +x_53 = l_Lean_addMacroScope(x_26, x_52, x_21); +x_54 = lean_box(0); +x_55 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__6; +lean_inc(x_19); +x_56 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_56, 0, x_19); +lean_ctor_set(x_56, 1, x_55); +lean_ctor_set(x_56, 2, x_53); +lean_ctor_set(x_56, 3, x_54); +x_57 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__245; +x_58 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_58, 0, x_19); lean_ctor_set(x_58, 1, x_57); -lean_ctor_set(x_58, 2, x_55); -lean_ctor_set(x_58, 3, x_56); -x_59 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__245; -x_60 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_60, 0, x_20); -lean_ctor_set(x_60, 1, x_59); -x_61 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__33; -x_62 = lean_array_push(x_61, x_58); -x_63 = lean_array_push(x_62, x_60); -x_64 = lean_array_push(x_63, x_16); -x_65 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_65, 0, x_36); -lean_ctor_set(x_65, 1, x_53); -lean_ctor_set(x_65, 2, x_64); -x_66 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__25; -x_67 = lean_array_push(x_66, x_65); -x_68 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_68, 0, x_36); -lean_ctor_set(x_68, 1, x_37); -lean_ctor_set(x_68, 2, x_67); -x_69 = lean_array_push(x_39, x_51); -x_70 = lean_array_push(x_69, x_68); -x_71 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_71, 0, x_36); -lean_ctor_set(x_71, 1, x_50); -lean_ctor_set(x_71, 2, x_70); -x_72 = lean_array_push(x_39, x_48); -x_73 = lean_array_push(x_72, x_71); -x_74 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_74, 0, x_36); -lean_ctor_set(x_74, 1, x_46); -lean_ctor_set(x_74, 2, x_73); -x_75 = lean_array_push(x_61, x_42); -x_76 = lean_array_push(x_75, x_44); -x_77 = lean_array_push(x_76, x_74); -x_78 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_78, 0, x_36); -lean_ctor_set(x_78, 1, x_33); -lean_ctor_set(x_78, 2, x_77); -x_79 = 1; -x_80 = lean_usize_add(x_5, x_79); -x_81 = lean_array_uset(x_18, x_5, x_78); +x_59 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__33; +x_60 = lean_array_push(x_59, x_56); +x_61 = lean_array_push(x_60, x_58); +x_62 = lean_array_push(x_61, x_15); +x_63 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_63, 0, x_39); +lean_ctor_set(x_63, 1, x_51); +lean_ctor_set(x_63, 2, x_62); +x_64 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__25; +x_65 = lean_array_push(x_64, x_63); +x_66 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__2; +x_67 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_67, 0, x_39); +lean_ctor_set(x_67, 1, x_66); +lean_ctor_set(x_67, 2, x_65); +x_68 = lean_array_push(x_35, x_49); +x_69 = lean_array_push(x_68, x_67); +x_70 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_70, 0, x_39); +lean_ctor_set(x_70, 1, x_48); +lean_ctor_set(x_70, 2, x_69); +x_71 = lean_array_push(x_35, x_46); +x_72 = lean_array_push(x_71, x_70); +x_73 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_73, 0, x_39); +lean_ctor_set(x_73, 1, x_44); +lean_ctor_set(x_73, 2, x_72); +x_74 = lean_array_push(x_59, x_40); +x_75 = lean_array_push(x_74, x_42); +x_76 = lean_array_push(x_75, x_73); +x_77 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_77, 0, x_39); +lean_ctor_set(x_77, 1, x_32); +lean_ctor_set(x_77, 2, x_76); +x_78 = 1; +x_79 = lean_usize_add(x_4, x_78); +x_80 = lean_array_uset(x_17, x_4, x_77); +x_4 = x_79; x_5 = x_80; -x_6 = x_81; -x_13 = x_25; +x_12 = x_24; goto _start; } } @@ -8565,7 +8530,31 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__24(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__22(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = 1; +x_9 = lean_usize_add(x_2, x_8); +x_10 = lean_array_uset(x_7, x_2, x_5); +x_2 = x_9; +x_3 = x_10; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__25(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { uint8_t x_13; @@ -8688,7 +8677,7 @@ return x_34; } } } -LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_withFieldsAux___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__23___rarg___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, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_withFieldsAux___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__24___rarg___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, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; lean_object* x_12; lean_object* x_13; size_t x_14; size_t x_15; lean_object* x_16; lean_object* x_17; @@ -8703,7 +8692,7 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_17 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__24(x_3, x_12, x_14, x_15, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_17 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__25(x_3, x_12, x_14, x_15, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_12); if (lean_obj_tag(x_17) == 0) { @@ -8747,7 +8736,7 @@ return x_24; } } } -LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_withFieldsAux___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__23___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, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_withFieldsAux___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__24___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, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { 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; lean_object* x_20; uint8_t x_21; lean_object* x_22; @@ -8763,7 +8752,7 @@ lean_inc(x_16); lean_dec(x_12); x_17 = l_Lean_mkConst(x_16, x_15); x_18 = l_Lean_mkAppN(x_17, x_2); -x_19 = lean_alloc_closure((void*)(l_Lean_Server_RpcEncoding_withFieldsAux___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__23___rarg___lambda__1), 10, 2); +x_19 = lean_alloc_closure((void*)(l_Lean_Server_RpcEncoding_withFieldsAux___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__24___rarg___lambda__1), 10, 2); lean_closure_set(x_19, 0, x_4); lean_closure_set(x_19, 1, x_3); x_20 = l_Lean_Server_RpcEncoding_withFieldsAux___rarg___closed__2; @@ -8772,15 +8761,15 @@ x_22 = l_Lean_Meta_withLocalDecl___at___private_Lean_Elab_Term_0__Lean_Elab_Term return x_22; } } -LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_withFieldsAux___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__23(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_withFieldsAux___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__24(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Server_RpcEncoding_withFieldsAux___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__23___rarg), 11, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Server_RpcEncoding_withFieldsAux___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__24___rarg), 11, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_withFields___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__22___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, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_withFields___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__23___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, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; @@ -8797,15 +8786,15 @@ lean_inc(x_1); x_15 = lean_alloc_closure((void*)(l_Lean_Server_RpcEncoding_withFields___rarg___lambda__2___boxed), 3, 2); lean_closure_set(x_15, 0, x_1); lean_closure_set(x_15, 1, x_14); -x_16 = l_Lean_Server_RpcEncoding_withFieldsAux___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__23___rarg(x_1, x_2, x_3, x_15, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +x_16 = l_Lean_Server_RpcEncoding_withFieldsAux___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__24___rarg(x_1, x_2, x_3, x_15, x_4, x_5, x_6, x_7, x_8, x_9, x_13); return x_16; } } -LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_withFields___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__22(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_withFields___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__23(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Server_RpcEncoding_withFields___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__22___rarg), 10, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Server_RpcEncoding_withFields___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__23___rarg), 10, 0); return x_2; } } @@ -9521,192 +9510,192 @@ return x_3; LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___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, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_12; size_t x_13; size_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +lean_object* x_12; size_t x_13; size_t x_14; lean_object* x_15; lean_object* x_16; x_12 = lean_array_get_size(x_1); x_13 = lean_usize_of_nat(x_12); lean_dec(x_12); x_14 = 0; -x_15 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__11; -x_16 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__7; +x_15 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__7; lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_17 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17(x_15, x_1, x_13, x_14, x_16, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_17) == 0) +x_16 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17(x_1, x_13, x_14, x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_16) == 0) { -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; lean_object* x_25; lean_object* x_26; lean_object* x_27; size_t x_28; lean_object* x_29; 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; size_t x_39; lean_object* x_40; -x_18 = lean_ctor_get(x_17, 0); +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; lean_object* x_25; lean_object* x_26; size_t x_27; lean_object* x_28; lean_object* x_29; 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; size_t x_38; lean_object* x_39; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_17, 1); lean_inc(x_18); x_19 = lean_ctor_get(x_18, 1); lean_inc(x_19); x_20 = lean_ctor_get(x_19, 1); lean_inc(x_20); -x_21 = lean_ctor_get(x_20, 1); +lean_dec(x_19); +x_21 = lean_ctor_get(x_16, 1); lean_inc(x_21); -lean_dec(x_20); -x_22 = lean_ctor_get(x_17, 1); +lean_dec(x_16); +x_22 = lean_ctor_get(x_17, 0); lean_inc(x_22); lean_dec(x_17); x_23 = lean_ctor_get(x_18, 0); lean_inc(x_23); lean_dec(x_18); -x_24 = lean_ctor_get(x_19, 0); +x_24 = lean_ctor_get(x_20, 0); lean_inc(x_24); -lean_dec(x_19); -x_25 = lean_ctor_get(x_21, 0); +x_25 = lean_ctor_get(x_20, 1); lean_inc(x_25); -x_26 = lean_ctor_get(x_21, 1); -lean_inc(x_26); -lean_dec(x_21); -x_27 = lean_array_get_size(x_25); -x_28 = lean_usize_of_nat(x_27); -lean_dec(x_27); -x_29 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__4; -x_30 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__225; +lean_dec(x_20); +x_26 = lean_array_get_size(x_24); +x_27 = lean_usize_of_nat(x_26); +lean_dec(x_26); +x_28 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__4; +x_29 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__225; lean_inc(x_9); -lean_inc(x_25); -x_31 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18(x_15, x_29, x_30, x_28, x_14, x_25, x_5, x_6, x_7, x_8, x_9, x_10, x_22); -x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_24); +x_30 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18(x_28, x_29, x_27, x_14, x_24, x_5, x_6, x_7, x_8, x_9, x_10, x_21); +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); lean_inc(x_32); -x_33 = lean_ctor_get(x_31, 1); -lean_inc(x_33); -lean_dec(x_31); -x_34 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__232; +lean_dec(x_30); +x_33 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__232; lean_inc(x_9); -lean_inc(x_25); -x_35 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__19(x_15, x_29, x_34, x_28, x_14, x_25, x_5, x_6, x_7, x_8, x_9, x_10, x_33); -x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_24); +x_34 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__19(x_28, x_33, x_27, x_14, x_24, x_5, x_6, x_7, x_8, x_9, x_10, x_32); +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_34, 1); lean_inc(x_36); -x_37 = lean_ctor_get(x_35, 1); -lean_inc(x_37); -lean_dec(x_35); -x_38 = lean_array_get_size(x_2); -x_39 = lean_usize_of_nat(x_38); -lean_dec(x_38); -x_40 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__20(x_39, x_14, x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_37); +lean_dec(x_34); +x_37 = lean_array_get_size(x_2); +x_38 = lean_usize_of_nat(x_37); +lean_dec(x_37); +x_39 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__20(x_38, x_14, x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_36); lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); -if (lean_obj_tag(x_40) == 0) +if (lean_obj_tag(x_39) == 0) { -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; -x_41 = lean_ctor_get(x_40, 0); +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; lean_object* x_67; lean_object* x_68; lean_object* x_69; uint8_t x_70; +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +x_41 = lean_ctor_get(x_39, 1); lean_inc(x_41); -x_42 = lean_ctor_get(x_40, 1); -lean_inc(x_42); -lean_dec(x_40); +lean_dec(x_39); lean_inc(x_9); -x_43 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_9, x_10, x_42); -x_44 = lean_ctor_get(x_43, 0); +x_42 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_9, x_10, x_41); +x_43 = lean_ctor_get(x_42, 0); +lean_inc(x_43); +x_44 = lean_ctor_get(x_42, 1); lean_inc(x_44); -x_45 = lean_ctor_get(x_43, 1); -lean_inc(x_45); -lean_dec(x_43); -x_46 = lean_st_ref_get(x_10, x_45); -x_47 = lean_ctor_get(x_46, 1); -lean_inc(x_47); -lean_dec(x_46); -x_48 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__10; -x_49 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_49, 0, x_44); -lean_ctor_set(x_49, 1, x_48); -x_50 = lean_ctor_get(x_3, 0); -lean_inc(x_50); +lean_dec(x_42); +x_45 = lean_st_ref_get(x_10, x_44); +x_46 = lean_ctor_get(x_45, 1); +lean_inc(x_46); +lean_dec(x_45); +x_47 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__10; +x_48 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_48, 0, x_43); +lean_ctor_set(x_48, 1, x_47); +x_49 = lean_ctor_get(x_3, 0); +lean_inc(x_49); lean_dec(x_3); -x_51 = lean_ctor_get(x_50, 0); -lean_inc(x_51); -lean_dec(x_50); -x_52 = lean_mk_syntax_ident(x_51); -x_53 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__13; -x_54 = lean_array_push(x_53, x_49); -x_55 = lean_array_push(x_54, x_52); -x_56 = lean_box(2); -x_57 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__9; -x_58 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_58, 0, x_56); -lean_ctor_set(x_58, 1, x_57); -lean_ctor_set(x_58, 2, x_55); -x_59 = l_Lean_Syntax_mkApp(x_58, x_41); -x_60 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__12; -x_61 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_60, x_9, x_10, x_47); -x_62 = lean_ctor_get(x_61, 0); +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +lean_dec(x_49); +x_51 = lean_mk_syntax_ident(x_50); +x_52 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__13; +x_53 = lean_array_push(x_52, x_48); +x_54 = lean_array_push(x_53, x_51); +x_55 = lean_box(2); +x_56 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__9; +x_57 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_56); +lean_ctor_set(x_57, 2, x_54); +x_58 = l_Lean_Syntax_mkApp(x_57, x_40); +x_59 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__12; +x_60 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_59, x_9, x_10, x_46); +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_60, 1); lean_inc(x_62); -x_63 = lean_ctor_get(x_61, 1); +lean_dec(x_60); +x_63 = lean_mk_syntax_ident(x_61); lean_inc(x_63); -lean_dec(x_61); -x_64 = lean_mk_syntax_ident(x_62); -lean_inc(x_64); -x_65 = l_Lean_Syntax_mkApp(x_64, x_26); +x_64 = l_Lean_Syntax_mkApp(x_63, x_25); lean_inc(x_9); -x_66 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_9, x_10, x_63); -x_67 = lean_ctor_get(x_66, 0); +x_65 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_9, x_10, x_62); +x_66 = lean_ctor_get(x_65, 0); +lean_inc(x_66); +x_67 = lean_ctor_get(x_65, 1); lean_inc(x_67); -x_68 = lean_ctor_get(x_66, 1); +lean_dec(x_65); +x_68 = lean_ctor_get(x_9, 10); lean_inc(x_68); -lean_dec(x_66); -x_69 = lean_ctor_get(x_9, 10); -lean_inc(x_69); lean_dec(x_9); -x_70 = lean_st_ref_get(x_10, x_68); +x_69 = lean_st_ref_get(x_10, x_67); lean_dec(x_10); -x_71 = !lean_is_exclusive(x_70); -if (x_71 == 0) +x_70 = !lean_is_exclusive(x_69); +if (x_70 == 0) { -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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; size_t x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; size_t x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; size_t x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; -x_72 = lean_ctor_get(x_70, 0); -x_73 = lean_ctor_get(x_72, 0); -lean_inc(x_73); -lean_dec(x_72); -x_74 = lean_environment_main_module(x_73); -x_75 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__14; -lean_inc(x_67); -x_76 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_76, 0, x_67); -lean_ctor_set(x_76, 1, x_75); -x_77 = l_Array_append___rarg(x_15, x_23); +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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; size_t x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; size_t x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; size_t x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; +x_71 = lean_ctor_get(x_69, 0); +x_72 = lean_ctor_get(x_71, 0); +lean_inc(x_72); +lean_dec(x_71); +x_73 = lean_environment_main_module(x_72); +x_74 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__14; +lean_inc(x_66); +x_75 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_75, 0, x_66); +lean_ctor_set(x_75, 1, x_74); +x_76 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__11; +x_77 = l_Array_append___rarg(x_76, x_22); x_78 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__2; x_79 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_79, 0, x_56); +lean_ctor_set(x_79, 0, x_55); lean_ctor_set(x_79, 1, x_78); lean_ctor_set(x_79, 2, x_77); -x_80 = lean_array_push(x_53, x_76); +x_80 = lean_array_push(x_52, x_75); x_81 = lean_array_push(x_80, x_79); x_82 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__15; x_83 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_83, 0, x_56); +lean_ctor_set(x_83, 0, x_55); lean_ctor_set(x_83, 1, x_82); lean_ctor_set(x_83, 2, x_81); x_84 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__21; -lean_inc(x_67); +lean_inc(x_66); x_85 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_85, 0, x_67); +lean_ctor_set(x_85, 0, x_66); lean_ctor_set(x_85, 1, x_84); x_86 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__25; x_87 = lean_array_push(x_86, x_85); x_88 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__24; x_89 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_89, 0, x_56); +lean_ctor_set(x_89, 0, x_55); lean_ctor_set(x_89, 1, x_88); lean_ctor_set(x_89, 2, x_87); -x_90 = lean_array_push(x_53, x_64); +x_90 = lean_array_push(x_52, x_63); x_91 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__13; x_92 = lean_array_push(x_90, x_91); x_93 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__50; x_94 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_94, 0, x_56); +lean_ctor_set(x_94, 0, x_55); lean_ctor_set(x_94, 1, x_93); lean_ctor_set(x_94, 2, x_92); x_95 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__25; -lean_inc(x_67); +lean_inc(x_66); x_96 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_96, 0, x_67); +lean_ctor_set(x_96, 0, x_66); lean_ctor_set(x_96, 1, x_95); -x_97 = l_Array_zip___rarg(x_25, x_24); +x_97 = l_Array_zip___rarg(x_24, x_23); +lean_dec(x_23); lean_dec(x_24); -lean_dec(x_25); x_98 = lean_array_get_size(x_97); x_99 = lean_usize_of_nat(x_98); lean_dec(x_98); @@ -9714,17 +9703,17 @@ x_100 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWi x_101 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__8; x_102 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__42; x_103 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__20; -lean_inc(x_67); -x_104 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__21(x_100, x_53, x_67, x_78, x_101, x_91, x_102, x_103, x_86, x_99, x_14, x_97); -x_105 = l_Array_append___rarg(x_15, x_104); +lean_inc(x_66); +x_104 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__21(x_100, x_52, x_66, x_78, x_101, x_91, x_102, x_103, x_86, x_99, x_14, x_97); +x_105 = l_Array_append___rarg(x_76, x_104); x_106 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_106, 0, x_56); +lean_ctor_set(x_106, 0, x_55); lean_ctor_set(x_106, 1, x_78); lean_ctor_set(x_106, 2, x_105); x_107 = lean_array_push(x_86, x_106); x_108 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__27; x_109 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_109, 0, x_56); +lean_ctor_set(x_109, 0, x_55); lean_ctor_set(x_109, 1, x_108); lean_ctor_set(x_109, 2, x_107); x_110 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__33; @@ -9733,37 +9722,37 @@ lean_inc(x_111); x_112 = lean_array_push(x_111, x_91); x_113 = lean_array_push(x_112, x_109); x_114 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_114, 0, x_56); +lean_ctor_set(x_114, 0, x_55); lean_ctor_set(x_114, 1, x_78); lean_ctor_set(x_114, 2, x_113); x_115 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__30; -lean_inc(x_67); +lean_inc(x_66); x_116 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_116, 0, x_67); +lean_ctor_set(x_116, 0, x_66); lean_ctor_set(x_116, 1, x_115); x_117 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__39; -lean_inc(x_67); +lean_inc(x_66); x_118 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_118, 0, x_67); +lean_ctor_set(x_118, 0, x_66); lean_ctor_set(x_118, 1, x_117); x_119 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__46; x_120 = lean_array_push(x_119, x_118); x_121 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__45; x_122 = lean_array_push(x_120, x_121); x_123 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_123, 0, x_56); +lean_ctor_set(x_123, 0, x_55); lean_ctor_set(x_123, 1, x_78); lean_ctor_set(x_123, 2, x_122); -x_124 = lean_array_push(x_53, x_116); +x_124 = lean_array_push(x_52, x_116); x_125 = lean_array_push(x_124, x_123); x_126 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_126, 0, x_56); +lean_ctor_set(x_126, 0, x_55); lean_ctor_set(x_126, 1, x_78); lean_ctor_set(x_126, 2, x_125); x_127 = lean_array_push(x_86, x_126); x_128 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__29; x_129 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_129, 0, x_56); +lean_ctor_set(x_129, 0, x_55); lean_ctor_set(x_129, 1, x_128); lean_ctor_set(x_129, 2, x_127); x_130 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__117; @@ -9776,126 +9765,126 @@ x_136 = lean_array_push(x_135, x_114); x_137 = lean_array_push(x_136, x_129); x_138 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__22; x_139 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_139, 0, x_56); +lean_ctor_set(x_139, 0, x_55); lean_ctor_set(x_139, 1, x_138); lean_ctor_set(x_139, 2, x_137); x_140 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__47; x_141 = lean_array_push(x_140, x_139); x_142 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__37; x_143 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_143, 0, x_56); +lean_ctor_set(x_143, 0, x_55); lean_ctor_set(x_143, 1, x_142); lean_ctor_set(x_143, 2, x_141); x_144 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__205; -lean_inc(x_67); +lean_inc(x_66); x_145 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_145, 0, x_67); +lean_ctor_set(x_145, 0, x_66); lean_ctor_set(x_145, 1, x_144); x_146 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__26; -lean_inc(x_67); +lean_inc(x_66); x_147 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_147, 0, x_67); +lean_ctor_set(x_147, 0, x_66); lean_ctor_set(x_147, 1, x_146); -x_148 = lean_array_push(x_53, x_59); -x_149 = lean_array_push(x_148, x_65); +x_148 = lean_array_push(x_52, x_58); +x_149 = lean_array_push(x_148, x_64); x_150 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_150, 0, x_56); +lean_ctor_set(x_150, 0, x_55); lean_ctor_set(x_150, 1, x_78); lean_ctor_set(x_150, 2, x_149); x_151 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__50; x_152 = lean_array_push(x_151, x_150); x_153 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__61; x_154 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_154, 0, x_56); +lean_ctor_set(x_154, 0, x_55); lean_ctor_set(x_154, 1, x_153); lean_ctor_set(x_154, 2, x_152); -x_155 = lean_array_push(x_53, x_147); +x_155 = lean_array_push(x_52, x_147); x_156 = lean_array_push(x_155, x_154); x_157 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__95; x_158 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_158, 0, x_56); +lean_ctor_set(x_158, 0, x_55); lean_ctor_set(x_158, 1, x_157); lean_ctor_set(x_158, 2, x_156); x_159 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__51; x_160 = lean_array_push(x_159, x_158); x_161 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__143; x_162 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_162, 0, x_56); +lean_ctor_set(x_162, 0, x_55); lean_ctor_set(x_162, 1, x_161); lean_ctor_set(x_162, 2, x_160); x_163 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__224; -lean_inc(x_69); -lean_inc(x_74); -x_164 = l_Lean_addMacroScope(x_74, x_163, x_69); +lean_inc(x_68); +lean_inc(x_73); +x_164 = l_Lean_addMacroScope(x_73, x_163, x_68); x_165 = lean_box(0); x_166 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__223; x_167 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__61; -lean_inc(x_67); +lean_inc(x_66); x_168 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_168, 0, x_67); +lean_ctor_set(x_168, 0, x_66); lean_ctor_set(x_168, 1, x_166); lean_ctor_set(x_168, 2, x_164); lean_ctor_set(x_168, 3, x_167); x_169 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__7; -lean_inc(x_69); -lean_inc(x_74); -x_170 = l_Lean_addMacroScope(x_74, x_169, x_69); +lean_inc(x_68); +lean_inc(x_73); +x_170 = l_Lean_addMacroScope(x_73, x_169, x_68); x_171 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__6; -lean_inc(x_67); +lean_inc(x_66); x_172 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_172, 0, x_67); +lean_ctor_set(x_172, 0, x_66); lean_ctor_set(x_172, 1, x_171); lean_ctor_set(x_172, 2, x_170); lean_ctor_set(x_172, 3, x_165); x_173 = lean_array_push(x_86, x_172); x_174 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_174, 0, x_56); +lean_ctor_set(x_174, 0, x_55); lean_ctor_set(x_174, 1, x_78); lean_ctor_set(x_174, 2, x_173); -x_175 = lean_array_push(x_53, x_174); +x_175 = lean_array_push(x_52, x_174); x_176 = lean_array_push(x_175, x_91); x_177 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__63; x_178 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_178, 0, x_56); +lean_ctor_set(x_178, 0, x_55); lean_ctor_set(x_178, 1, x_177); lean_ctor_set(x_178, 2, x_176); x_179 = lean_array_push(x_86, x_178); x_180 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_180, 0, x_56); +lean_ctor_set(x_180, 0, x_55); lean_ctor_set(x_180, 1, x_78); lean_ctor_set(x_180, 2, x_179); x_181 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__109; -lean_inc(x_67); +lean_inc(x_66); x_182 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_182, 0, x_67); +lean_ctor_set(x_182, 0, x_66); lean_ctor_set(x_182, 1, x_181); x_183 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__66; -lean_inc(x_67); +lean_inc(x_66); x_184 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_184, 0, x_67); +lean_ctor_set(x_184, 0, x_66); lean_ctor_set(x_184, 1, x_183); x_185 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__20; -lean_inc(x_67); +lean_inc(x_66); x_186 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_186, 0, x_67); +lean_ctor_set(x_186, 0, x_66); lean_ctor_set(x_186, 1, x_185); -x_187 = lean_array_get_size(x_32); +x_187 = lean_array_get_size(x_31); x_188 = lean_usize_of_nat(x_187); lean_dec(x_187); -x_189 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_188, x_14, x_32); +x_189 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__22(x_188, x_14, x_31); x_190 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__67; x_191 = l_Lean_mkSepArray(x_189, x_190); lean_dec(x_189); -x_192 = l_Array_append___rarg(x_15, x_191); +x_192 = l_Array_append___rarg(x_76, x_191); x_193 = lean_array_push(x_192, x_91); x_194 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_194, 0, x_56); +lean_ctor_set(x_194, 0, x_55); lean_ctor_set(x_194, 1, x_78); lean_ctor_set(x_194, 2, x_193); x_195 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__34; -lean_inc(x_67); +lean_inc(x_66); x_196 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_196, 0, x_67); +lean_ctor_set(x_196, 0, x_66); lean_ctor_set(x_196, 1, x_195); x_197 = lean_array_push(x_102, x_186); x_198 = lean_array_push(x_197, x_91); @@ -9908,20 +9897,20 @@ lean_inc(x_196); x_203 = lean_array_push(x_202, x_196); x_204 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__216; x_205 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_205, 0, x_56); +lean_ctor_set(x_205, 0, x_55); lean_ctor_set(x_205, 1, x_204); lean_ctor_set(x_205, 2, x_203); x_206 = lean_array_push(x_86, x_205); x_207 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_207, 0, x_56); +lean_ctor_set(x_207, 0, x_55); lean_ctor_set(x_207, 1, x_78); lean_ctor_set(x_207, 2, x_206); -x_208 = lean_array_push(x_53, x_184); +x_208 = lean_array_push(x_52, x_184); lean_inc(x_208); x_209 = lean_array_push(x_208, x_207); x_210 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__65; x_211 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_211, 0, x_56); +lean_ctor_set(x_211, 0, x_55); lean_ctor_set(x_211, 1, x_210); lean_ctor_set(x_211, 2, x_209); x_212 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__93; @@ -9934,40 +9923,40 @@ x_216 = lean_array_push(x_215, x_182); x_217 = lean_array_push(x_216, x_211); x_218 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__59; x_219 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_219, 0, x_56); +lean_ctor_set(x_219, 0, x_55); lean_ctor_set(x_219, 1, x_218); lean_ctor_set(x_219, 2, x_217); x_220 = lean_array_push(x_86, x_219); x_221 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__57; x_222 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_222, 0, x_56); +lean_ctor_set(x_222, 0, x_55); lean_ctor_set(x_222, 1, x_221); lean_ctor_set(x_222, 2, x_220); x_223 = lean_array_push(x_86, x_222); x_224 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__55; x_225 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_225, 0, x_56); +lean_ctor_set(x_225, 0, x_55); lean_ctor_set(x_225, 1, x_224); lean_ctor_set(x_225, 2, x_223); x_226 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__231; -x_227 = l_Lean_addMacroScope(x_74, x_226, x_69); +x_227 = l_Lean_addMacroScope(x_73, x_226, x_68); x_228 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__230; x_229 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__70; x_230 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_230, 0, x_67); +lean_ctor_set(x_230, 0, x_66); lean_ctor_set(x_230, 1, x_228); lean_ctor_set(x_230, 2, x_227); lean_ctor_set(x_230, 3, x_229); -x_231 = lean_array_get_size(x_36); +x_231 = lean_array_get_size(x_35); x_232 = lean_usize_of_nat(x_231); lean_dec(x_231); -x_233 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_232, x_14, x_36); +x_233 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__22(x_232, x_14, x_35); x_234 = l_Lean_mkSepArray(x_233, x_190); lean_dec(x_233); -x_235 = l_Array_append___rarg(x_15, x_234); +x_235 = l_Array_append___rarg(x_76, x_234); x_236 = lean_array_push(x_235, x_91); x_237 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_237, 0, x_56); +lean_ctor_set(x_237, 0, x_55); lean_ctor_set(x_237, 1, x_78); lean_ctor_set(x_237, 2, x_236); x_238 = lean_array_push(x_198, x_237); @@ -9975,17 +9964,17 @@ x_239 = lean_array_push(x_238, x_200); x_240 = lean_array_push(x_239, x_91); x_241 = lean_array_push(x_240, x_196); x_242 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_242, 0, x_56); +lean_ctor_set(x_242, 0, x_55); lean_ctor_set(x_242, 1, x_204); lean_ctor_set(x_242, 2, x_241); x_243 = lean_array_push(x_86, x_242); x_244 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_244, 0, x_56); +lean_ctor_set(x_244, 0, x_55); lean_ctor_set(x_244, 1, x_78); lean_ctor_set(x_244, 2, x_243); x_245 = lean_array_push(x_208, x_244); x_246 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_246, 0, x_56); +lean_ctor_set(x_246, 0, x_55); lean_ctor_set(x_246, 1, x_210); lean_ctor_set(x_246, 2, x_245); x_247 = lean_array_push(x_212, x_230); @@ -9994,17 +9983,17 @@ x_249 = lean_array_push(x_248, x_91); x_250 = lean_array_push(x_249, x_182); x_251 = lean_array_push(x_250, x_246); x_252 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_252, 0, x_56); +lean_ctor_set(x_252, 0, x_55); lean_ctor_set(x_252, 1, x_218); lean_ctor_set(x_252, 2, x_251); x_253 = lean_array_push(x_86, x_252); x_254 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_254, 0, x_56); +lean_ctor_set(x_254, 0, x_55); lean_ctor_set(x_254, 1, x_221); lean_ctor_set(x_254, 2, x_253); x_255 = lean_array_push(x_86, x_254); x_256 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_256, 0, x_56); +lean_ctor_set(x_256, 0, x_55); lean_ctor_set(x_256, 1, x_224); lean_ctor_set(x_256, 2, x_255); x_257 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__35; @@ -10013,14 +10002,14 @@ x_259 = lean_array_push(x_258, x_91); x_260 = lean_array_push(x_259, x_256); x_261 = lean_array_push(x_260, x_91); x_262 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_262, 0, x_56); +lean_ctor_set(x_262, 0, x_55); lean_ctor_set(x_262, 1, x_78); lean_ctor_set(x_262, 2, x_261); x_263 = lean_array_push(x_111, x_262); x_264 = lean_array_push(x_263, x_91); x_265 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__53; x_266 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_266, 0, x_56); +lean_ctor_set(x_266, 0, x_55); lean_ctor_set(x_266, 1, x_265); lean_ctor_set(x_266, 2, x_264); x_267 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__71; @@ -10033,464 +10022,465 @@ x_273 = lean_array_push(x_272, x_91); x_274 = lean_array_push(x_273, x_91); x_275 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__206; x_276 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_276, 0, x_56); +lean_ctor_set(x_276, 0, x_55); lean_ctor_set(x_276, 1, x_275); lean_ctor_set(x_276, 2, x_274); x_277 = lean_array_push(x_140, x_276); x_278 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_278, 0, x_56); +lean_ctor_set(x_278, 0, x_55); lean_ctor_set(x_278, 1, x_142); lean_ctor_set(x_278, 2, x_277); x_279 = lean_array_push(x_110, x_83); x_280 = lean_array_push(x_279, x_143); x_281 = lean_array_push(x_280, x_278); x_282 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_282, 0, x_56); +lean_ctor_set(x_282, 0, x_55); lean_ctor_set(x_282, 1, x_78); lean_ctor_set(x_282, 2, x_281); -lean_ctor_set(x_70, 0, x_282); -return x_70; +lean_ctor_set(x_69, 0, x_282); +return x_69; } else { -lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; size_t x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; size_t x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; size_t x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; -x_283 = lean_ctor_get(x_70, 0); -x_284 = lean_ctor_get(x_70, 1); +lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; size_t x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; size_t x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; size_t x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; +x_283 = lean_ctor_get(x_69, 0); +x_284 = lean_ctor_get(x_69, 1); lean_inc(x_284); lean_inc(x_283); -lean_dec(x_70); +lean_dec(x_69); x_285 = lean_ctor_get(x_283, 0); lean_inc(x_285); lean_dec(x_283); x_286 = lean_environment_main_module(x_285); x_287 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__14; -lean_inc(x_67); +lean_inc(x_66); x_288 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_288, 0, x_67); +lean_ctor_set(x_288, 0, x_66); lean_ctor_set(x_288, 1, x_287); -x_289 = l_Array_append___rarg(x_15, x_23); -x_290 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__2; -x_291 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_291, 0, x_56); -lean_ctor_set(x_291, 1, x_290); -lean_ctor_set(x_291, 2, x_289); -x_292 = lean_array_push(x_53, x_288); -x_293 = lean_array_push(x_292, x_291); -x_294 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__15; -x_295 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_295, 0, x_56); -lean_ctor_set(x_295, 1, x_294); -lean_ctor_set(x_295, 2, x_293); -x_296 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__21; -lean_inc(x_67); -x_297 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_297, 0, x_67); -lean_ctor_set(x_297, 1, x_296); -x_298 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__25; -x_299 = lean_array_push(x_298, x_297); -x_300 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__24; -x_301 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_301, 0, x_56); -lean_ctor_set(x_301, 1, x_300); -lean_ctor_set(x_301, 2, x_299); -x_302 = lean_array_push(x_53, x_64); -x_303 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__13; -x_304 = lean_array_push(x_302, x_303); -x_305 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__50; -x_306 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_306, 0, x_56); -lean_ctor_set(x_306, 1, x_305); -lean_ctor_set(x_306, 2, x_304); -x_307 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__25; -lean_inc(x_67); -x_308 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_308, 0, x_67); -lean_ctor_set(x_308, 1, x_307); -x_309 = l_Array_zip___rarg(x_25, x_24); +x_289 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__11; +x_290 = l_Array_append___rarg(x_289, x_22); +x_291 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__2; +x_292 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_292, 0, x_55); +lean_ctor_set(x_292, 1, x_291); +lean_ctor_set(x_292, 2, x_290); +x_293 = lean_array_push(x_52, x_288); +x_294 = lean_array_push(x_293, x_292); +x_295 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__15; +x_296 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_296, 0, x_55); +lean_ctor_set(x_296, 1, x_295); +lean_ctor_set(x_296, 2, x_294); +x_297 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__21; +lean_inc(x_66); +x_298 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_298, 0, x_66); +lean_ctor_set(x_298, 1, x_297); +x_299 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__25; +x_300 = lean_array_push(x_299, x_298); +x_301 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__24; +x_302 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_302, 0, x_55); +lean_ctor_set(x_302, 1, x_301); +lean_ctor_set(x_302, 2, x_300); +x_303 = lean_array_push(x_52, x_63); +x_304 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__13; +x_305 = lean_array_push(x_303, x_304); +x_306 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__50; +x_307 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_307, 0, x_55); +lean_ctor_set(x_307, 1, x_306); +lean_ctor_set(x_307, 2, x_305); +x_308 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__25; +lean_inc(x_66); +x_309 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_309, 0, x_66); +lean_ctor_set(x_309, 1, x_308); +x_310 = l_Array_zip___rarg(x_24, x_23); +lean_dec(x_23); lean_dec(x_24); -lean_dec(x_25); -x_310 = lean_array_get_size(x_309); -x_311 = lean_usize_of_nat(x_310); -lean_dec(x_310); -x_312 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__17; -x_313 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__8; -x_314 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__42; -x_315 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__20; -lean_inc(x_67); -x_316 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__21(x_312, x_53, x_67, x_290, x_313, x_303, x_314, x_315, x_298, x_311, x_14, x_309); -x_317 = l_Array_append___rarg(x_15, x_316); -x_318 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_318, 0, x_56); -lean_ctor_set(x_318, 1, x_290); -lean_ctor_set(x_318, 2, x_317); -x_319 = lean_array_push(x_298, x_318); -x_320 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__27; -x_321 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_321, 0, x_56); -lean_ctor_set(x_321, 1, x_320); -lean_ctor_set(x_321, 2, x_319); -x_322 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__33; -x_323 = lean_array_push(x_322, x_308); -lean_inc(x_323); -x_324 = lean_array_push(x_323, x_303); -x_325 = lean_array_push(x_324, x_321); -x_326 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_326, 0, x_56); -lean_ctor_set(x_326, 1, x_290); -lean_ctor_set(x_326, 2, x_325); -x_327 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__30; -lean_inc(x_67); -x_328 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_328, 0, x_67); -lean_ctor_set(x_328, 1, x_327); -x_329 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__39; -lean_inc(x_67); -x_330 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_330, 0, x_67); -lean_ctor_set(x_330, 1, x_329); -x_331 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__46; -x_332 = lean_array_push(x_331, x_330); -x_333 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__45; -x_334 = lean_array_push(x_332, x_333); -x_335 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_335, 0, x_56); -lean_ctor_set(x_335, 1, x_290); -lean_ctor_set(x_335, 2, x_334); -x_336 = lean_array_push(x_53, x_328); -x_337 = lean_array_push(x_336, x_335); -x_338 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_338, 0, x_56); -lean_ctor_set(x_338, 1, x_290); -lean_ctor_set(x_338, 2, x_337); -x_339 = lean_array_push(x_298, x_338); -x_340 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__29; -x_341 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_341, 0, x_56); -lean_ctor_set(x_341, 1, x_340); -lean_ctor_set(x_341, 2, x_339); -x_342 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__117; -x_343 = lean_array_push(x_342, x_301); -x_344 = lean_array_push(x_343, x_306); -x_345 = lean_array_push(x_344, x_303); -x_346 = lean_array_push(x_345, x_303); -x_347 = lean_array_push(x_346, x_303); -x_348 = lean_array_push(x_347, x_326); -x_349 = lean_array_push(x_348, x_341); -x_350 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__22; -x_351 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_351, 0, x_56); -lean_ctor_set(x_351, 1, x_350); -lean_ctor_set(x_351, 2, x_349); -x_352 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__47; -x_353 = lean_array_push(x_352, x_351); -x_354 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__37; -x_355 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_355, 0, x_56); -lean_ctor_set(x_355, 1, x_354); -lean_ctor_set(x_355, 2, x_353); -x_356 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__205; -lean_inc(x_67); -x_357 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_357, 0, x_67); -lean_ctor_set(x_357, 1, x_356); -x_358 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__26; -lean_inc(x_67); -x_359 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_359, 0, x_67); -lean_ctor_set(x_359, 1, x_358); -x_360 = lean_array_push(x_53, x_59); -x_361 = lean_array_push(x_360, x_65); -x_362 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_362, 0, x_56); -lean_ctor_set(x_362, 1, x_290); -lean_ctor_set(x_362, 2, x_361); -x_363 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__50; -x_364 = lean_array_push(x_363, x_362); -x_365 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__61; -x_366 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_366, 0, x_56); -lean_ctor_set(x_366, 1, x_365); -lean_ctor_set(x_366, 2, x_364); -x_367 = lean_array_push(x_53, x_359); -x_368 = lean_array_push(x_367, x_366); -x_369 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__95; -x_370 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_370, 0, x_56); -lean_ctor_set(x_370, 1, x_369); -lean_ctor_set(x_370, 2, x_368); -x_371 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__51; -x_372 = lean_array_push(x_371, x_370); -x_373 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__143; -x_374 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_374, 0, x_56); -lean_ctor_set(x_374, 1, x_373); -lean_ctor_set(x_374, 2, x_372); -x_375 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__224; -lean_inc(x_69); +x_311 = lean_array_get_size(x_310); +x_312 = lean_usize_of_nat(x_311); +lean_dec(x_311); +x_313 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__17; +x_314 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__8; +x_315 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__42; +x_316 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__20; +lean_inc(x_66); +x_317 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__21(x_313, x_52, x_66, x_291, x_314, x_304, x_315, x_316, x_299, x_312, x_14, x_310); +x_318 = l_Array_append___rarg(x_289, x_317); +x_319 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_319, 0, x_55); +lean_ctor_set(x_319, 1, x_291); +lean_ctor_set(x_319, 2, x_318); +x_320 = lean_array_push(x_299, x_319); +x_321 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__27; +x_322 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_322, 0, x_55); +lean_ctor_set(x_322, 1, x_321); +lean_ctor_set(x_322, 2, x_320); +x_323 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__33; +x_324 = lean_array_push(x_323, x_309); +lean_inc(x_324); +x_325 = lean_array_push(x_324, x_304); +x_326 = lean_array_push(x_325, x_322); +x_327 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_327, 0, x_55); +lean_ctor_set(x_327, 1, x_291); +lean_ctor_set(x_327, 2, x_326); +x_328 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__30; +lean_inc(x_66); +x_329 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_329, 0, x_66); +lean_ctor_set(x_329, 1, x_328); +x_330 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__39; +lean_inc(x_66); +x_331 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_331, 0, x_66); +lean_ctor_set(x_331, 1, x_330); +x_332 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__46; +x_333 = lean_array_push(x_332, x_331); +x_334 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__45; +x_335 = lean_array_push(x_333, x_334); +x_336 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_336, 0, x_55); +lean_ctor_set(x_336, 1, x_291); +lean_ctor_set(x_336, 2, x_335); +x_337 = lean_array_push(x_52, x_329); +x_338 = lean_array_push(x_337, x_336); +x_339 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_339, 0, x_55); +lean_ctor_set(x_339, 1, x_291); +lean_ctor_set(x_339, 2, x_338); +x_340 = lean_array_push(x_299, x_339); +x_341 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__29; +x_342 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_342, 0, x_55); +lean_ctor_set(x_342, 1, x_341); +lean_ctor_set(x_342, 2, x_340); +x_343 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__117; +x_344 = lean_array_push(x_343, x_302); +x_345 = lean_array_push(x_344, x_307); +x_346 = lean_array_push(x_345, x_304); +x_347 = lean_array_push(x_346, x_304); +x_348 = lean_array_push(x_347, x_304); +x_349 = lean_array_push(x_348, x_327); +x_350 = lean_array_push(x_349, x_342); +x_351 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__22; +x_352 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_352, 0, x_55); +lean_ctor_set(x_352, 1, x_351); +lean_ctor_set(x_352, 2, x_350); +x_353 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__47; +x_354 = lean_array_push(x_353, x_352); +x_355 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__37; +x_356 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_356, 0, x_55); +lean_ctor_set(x_356, 1, x_355); +lean_ctor_set(x_356, 2, x_354); +x_357 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__205; +lean_inc(x_66); +x_358 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_358, 0, x_66); +lean_ctor_set(x_358, 1, x_357); +x_359 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__26; +lean_inc(x_66); +x_360 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_360, 0, x_66); +lean_ctor_set(x_360, 1, x_359); +x_361 = lean_array_push(x_52, x_58); +x_362 = lean_array_push(x_361, x_64); +x_363 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_363, 0, x_55); +lean_ctor_set(x_363, 1, x_291); +lean_ctor_set(x_363, 2, x_362); +x_364 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__50; +x_365 = lean_array_push(x_364, x_363); +x_366 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__61; +x_367 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_367, 0, x_55); +lean_ctor_set(x_367, 1, x_366); +lean_ctor_set(x_367, 2, x_365); +x_368 = lean_array_push(x_52, x_360); +x_369 = lean_array_push(x_368, x_367); +x_370 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__95; +x_371 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_371, 0, x_55); +lean_ctor_set(x_371, 1, x_370); +lean_ctor_set(x_371, 2, x_369); +x_372 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__51; +x_373 = lean_array_push(x_372, x_371); +x_374 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__143; +x_375 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_375, 0, x_55); +lean_ctor_set(x_375, 1, x_374); +lean_ctor_set(x_375, 2, x_373); +x_376 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__224; +lean_inc(x_68); lean_inc(x_286); -x_376 = l_Lean_addMacroScope(x_286, x_375, x_69); -x_377 = lean_box(0); -x_378 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__223; -x_379 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__61; -lean_inc(x_67); -x_380 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_380, 0, x_67); -lean_ctor_set(x_380, 1, x_378); -lean_ctor_set(x_380, 2, x_376); -lean_ctor_set(x_380, 3, x_379); -x_381 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__7; -lean_inc(x_69); +x_377 = l_Lean_addMacroScope(x_286, x_376, x_68); +x_378 = lean_box(0); +x_379 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__223; +x_380 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__61; +lean_inc(x_66); +x_381 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_381, 0, x_66); +lean_ctor_set(x_381, 1, x_379); +lean_ctor_set(x_381, 2, x_377); +lean_ctor_set(x_381, 3, x_380); +x_382 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__7; +lean_inc(x_68); lean_inc(x_286); -x_382 = l_Lean_addMacroScope(x_286, x_381, x_69); -x_383 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__6; -lean_inc(x_67); -x_384 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_384, 0, x_67); -lean_ctor_set(x_384, 1, x_383); -lean_ctor_set(x_384, 2, x_382); -lean_ctor_set(x_384, 3, x_377); -x_385 = lean_array_push(x_298, x_384); -x_386 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_386, 0, x_56); -lean_ctor_set(x_386, 1, x_290); -lean_ctor_set(x_386, 2, x_385); -x_387 = lean_array_push(x_53, x_386); -x_388 = lean_array_push(x_387, x_303); -x_389 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__63; -x_390 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_390, 0, x_56); -lean_ctor_set(x_390, 1, x_389); -lean_ctor_set(x_390, 2, x_388); -x_391 = lean_array_push(x_298, x_390); -x_392 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_392, 0, x_56); -lean_ctor_set(x_392, 1, x_290); -lean_ctor_set(x_392, 2, x_391); -x_393 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__109; -lean_inc(x_67); -x_394 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_394, 0, x_67); -lean_ctor_set(x_394, 1, x_393); -x_395 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__66; -lean_inc(x_67); -x_396 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_396, 0, x_67); -lean_ctor_set(x_396, 1, x_395); -x_397 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__20; -lean_inc(x_67); -x_398 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_398, 0, x_67); -lean_ctor_set(x_398, 1, x_397); -x_399 = lean_array_get_size(x_32); -x_400 = lean_usize_of_nat(x_399); -lean_dec(x_399); -x_401 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_400, x_14, x_32); -x_402 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__67; -x_403 = l_Lean_mkSepArray(x_401, x_402); -lean_dec(x_401); -x_404 = l_Array_append___rarg(x_15, x_403); -x_405 = lean_array_push(x_404, x_303); -x_406 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_406, 0, x_56); -lean_ctor_set(x_406, 1, x_290); -lean_ctor_set(x_406, 2, x_405); -x_407 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__34; -lean_inc(x_67); -x_408 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_408, 0, x_67); -lean_ctor_set(x_408, 1, x_407); -x_409 = lean_array_push(x_314, x_398); -x_410 = lean_array_push(x_409, x_303); -lean_inc(x_410); -x_411 = lean_array_push(x_410, x_406); -x_412 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__68; -x_413 = lean_array_push(x_411, x_412); -x_414 = lean_array_push(x_413, x_303); -lean_inc(x_408); -x_415 = lean_array_push(x_414, x_408); -x_416 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__216; -x_417 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_417, 0, x_56); -lean_ctor_set(x_417, 1, x_416); -lean_ctor_set(x_417, 2, x_415); -x_418 = lean_array_push(x_298, x_417); -x_419 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_419, 0, x_56); -lean_ctor_set(x_419, 1, x_290); -lean_ctor_set(x_419, 2, x_418); -x_420 = lean_array_push(x_53, x_396); -lean_inc(x_420); -x_421 = lean_array_push(x_420, x_419); -x_422 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__65; -x_423 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_423, 0, x_56); -lean_ctor_set(x_423, 1, x_422); -lean_ctor_set(x_423, 2, x_421); -x_424 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__93; -x_425 = lean_array_push(x_424, x_380); -lean_inc(x_392); -x_426 = lean_array_push(x_425, x_392); -x_427 = lean_array_push(x_426, x_303); -lean_inc(x_394); -x_428 = lean_array_push(x_427, x_394); -x_429 = lean_array_push(x_428, x_423); -x_430 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__59; -x_431 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_431, 0, x_56); -lean_ctor_set(x_431, 1, x_430); -lean_ctor_set(x_431, 2, x_429); -x_432 = lean_array_push(x_298, x_431); -x_433 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__57; -x_434 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_434, 0, x_56); -lean_ctor_set(x_434, 1, x_433); -lean_ctor_set(x_434, 2, x_432); -x_435 = lean_array_push(x_298, x_434); -x_436 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__55; -x_437 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_437, 0, x_56); -lean_ctor_set(x_437, 1, x_436); -lean_ctor_set(x_437, 2, x_435); -x_438 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__231; -x_439 = l_Lean_addMacroScope(x_286, x_438, x_69); -x_440 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__230; -x_441 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__70; -x_442 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_442, 0, x_67); -lean_ctor_set(x_442, 1, x_440); -lean_ctor_set(x_442, 2, x_439); -lean_ctor_set(x_442, 3, x_441); -x_443 = lean_array_get_size(x_36); -x_444 = lean_usize_of_nat(x_443); -lean_dec(x_443); -x_445 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_444, x_14, x_36); -x_446 = l_Lean_mkSepArray(x_445, x_402); -lean_dec(x_445); -x_447 = l_Array_append___rarg(x_15, x_446); -x_448 = lean_array_push(x_447, x_303); -x_449 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_449, 0, x_56); -lean_ctor_set(x_449, 1, x_290); -lean_ctor_set(x_449, 2, x_448); -x_450 = lean_array_push(x_410, x_449); -x_451 = lean_array_push(x_450, x_412); -x_452 = lean_array_push(x_451, x_303); -x_453 = lean_array_push(x_452, x_408); -x_454 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_454, 0, x_56); -lean_ctor_set(x_454, 1, x_416); -lean_ctor_set(x_454, 2, x_453); -x_455 = lean_array_push(x_298, x_454); -x_456 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_456, 0, x_56); -lean_ctor_set(x_456, 1, x_290); -lean_ctor_set(x_456, 2, x_455); -x_457 = lean_array_push(x_420, x_456); -x_458 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_458, 0, x_56); -lean_ctor_set(x_458, 1, x_422); -lean_ctor_set(x_458, 2, x_457); -x_459 = lean_array_push(x_424, x_442); -x_460 = lean_array_push(x_459, x_392); -x_461 = lean_array_push(x_460, x_303); -x_462 = lean_array_push(x_461, x_394); -x_463 = lean_array_push(x_462, x_458); -x_464 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_464, 0, x_56); -lean_ctor_set(x_464, 1, x_430); -lean_ctor_set(x_464, 2, x_463); -x_465 = lean_array_push(x_298, x_464); -x_466 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_466, 0, x_56); -lean_ctor_set(x_466, 1, x_433); -lean_ctor_set(x_466, 2, x_465); -x_467 = lean_array_push(x_298, x_466); -x_468 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_468, 0, x_56); -lean_ctor_set(x_468, 1, x_436); -lean_ctor_set(x_468, 2, x_467); -x_469 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__35; -x_470 = lean_array_push(x_469, x_437); -x_471 = lean_array_push(x_470, x_303); -x_472 = lean_array_push(x_471, x_468); -x_473 = lean_array_push(x_472, x_303); -x_474 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_474, 0, x_56); -lean_ctor_set(x_474, 1, x_290); -lean_ctor_set(x_474, 2, x_473); -x_475 = lean_array_push(x_323, x_474); -x_476 = lean_array_push(x_475, x_303); -x_477 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__53; -x_478 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_478, 0, x_56); -lean_ctor_set(x_478, 1, x_477); -lean_ctor_set(x_478, 2, x_476); -x_479 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__71; -x_480 = lean_array_push(x_479, x_357); -x_481 = lean_array_push(x_480, x_303); -x_482 = lean_array_push(x_481, x_303); -x_483 = lean_array_push(x_482, x_374); -x_484 = lean_array_push(x_483, x_478); -x_485 = lean_array_push(x_484, x_303); -x_486 = lean_array_push(x_485, x_303); -x_487 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__206; -x_488 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_488, 0, x_56); -lean_ctor_set(x_488, 1, x_487); -lean_ctor_set(x_488, 2, x_486); -x_489 = lean_array_push(x_352, x_488); -x_490 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_490, 0, x_56); -lean_ctor_set(x_490, 1, x_354); -lean_ctor_set(x_490, 2, x_489); -x_491 = lean_array_push(x_322, x_295); -x_492 = lean_array_push(x_491, x_355); -x_493 = lean_array_push(x_492, x_490); -x_494 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_494, 0, x_56); -lean_ctor_set(x_494, 1, x_290); -lean_ctor_set(x_494, 2, x_493); -x_495 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_495, 0, x_494); -lean_ctor_set(x_495, 1, x_284); -return x_495; +x_383 = l_Lean_addMacroScope(x_286, x_382, x_68); +x_384 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__6; +lean_inc(x_66); +x_385 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_385, 0, x_66); +lean_ctor_set(x_385, 1, x_384); +lean_ctor_set(x_385, 2, x_383); +lean_ctor_set(x_385, 3, x_378); +x_386 = lean_array_push(x_299, x_385); +x_387 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_387, 0, x_55); +lean_ctor_set(x_387, 1, x_291); +lean_ctor_set(x_387, 2, x_386); +x_388 = lean_array_push(x_52, x_387); +x_389 = lean_array_push(x_388, x_304); +x_390 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__63; +x_391 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_391, 0, x_55); +lean_ctor_set(x_391, 1, x_390); +lean_ctor_set(x_391, 2, x_389); +x_392 = lean_array_push(x_299, x_391); +x_393 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_393, 0, x_55); +lean_ctor_set(x_393, 1, x_291); +lean_ctor_set(x_393, 2, x_392); +x_394 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__109; +lean_inc(x_66); +x_395 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_395, 0, x_66); +lean_ctor_set(x_395, 1, x_394); +x_396 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__66; +lean_inc(x_66); +x_397 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_397, 0, x_66); +lean_ctor_set(x_397, 1, x_396); +x_398 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__20; +lean_inc(x_66); +x_399 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_399, 0, x_66); +lean_ctor_set(x_399, 1, x_398); +x_400 = lean_array_get_size(x_31); +x_401 = lean_usize_of_nat(x_400); +lean_dec(x_400); +x_402 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__22(x_401, x_14, x_31); +x_403 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__67; +x_404 = l_Lean_mkSepArray(x_402, x_403); +lean_dec(x_402); +x_405 = l_Array_append___rarg(x_289, x_404); +x_406 = lean_array_push(x_405, x_304); +x_407 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_407, 0, x_55); +lean_ctor_set(x_407, 1, x_291); +lean_ctor_set(x_407, 2, x_406); +x_408 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__34; +lean_inc(x_66); +x_409 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_409, 0, x_66); +lean_ctor_set(x_409, 1, x_408); +x_410 = lean_array_push(x_315, x_399); +x_411 = lean_array_push(x_410, x_304); +lean_inc(x_411); +x_412 = lean_array_push(x_411, x_407); +x_413 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__68; +x_414 = lean_array_push(x_412, x_413); +x_415 = lean_array_push(x_414, x_304); +lean_inc(x_409); +x_416 = lean_array_push(x_415, x_409); +x_417 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__216; +x_418 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_418, 0, x_55); +lean_ctor_set(x_418, 1, x_417); +lean_ctor_set(x_418, 2, x_416); +x_419 = lean_array_push(x_299, x_418); +x_420 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_420, 0, x_55); +lean_ctor_set(x_420, 1, x_291); +lean_ctor_set(x_420, 2, x_419); +x_421 = lean_array_push(x_52, x_397); +lean_inc(x_421); +x_422 = lean_array_push(x_421, x_420); +x_423 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__65; +x_424 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_424, 0, x_55); +lean_ctor_set(x_424, 1, x_423); +lean_ctor_set(x_424, 2, x_422); +x_425 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__93; +x_426 = lean_array_push(x_425, x_381); +lean_inc(x_393); +x_427 = lean_array_push(x_426, x_393); +x_428 = lean_array_push(x_427, x_304); +lean_inc(x_395); +x_429 = lean_array_push(x_428, x_395); +x_430 = lean_array_push(x_429, x_424); +x_431 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__59; +x_432 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_432, 0, x_55); +lean_ctor_set(x_432, 1, x_431); +lean_ctor_set(x_432, 2, x_430); +x_433 = lean_array_push(x_299, x_432); +x_434 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__57; +x_435 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_435, 0, x_55); +lean_ctor_set(x_435, 1, x_434); +lean_ctor_set(x_435, 2, x_433); +x_436 = lean_array_push(x_299, x_435); +x_437 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__55; +x_438 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_438, 0, x_55); +lean_ctor_set(x_438, 1, x_437); +lean_ctor_set(x_438, 2, x_436); +x_439 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__231; +x_440 = l_Lean_addMacroScope(x_286, x_439, x_68); +x_441 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__230; +x_442 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__70; +x_443 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_443, 0, x_66); +lean_ctor_set(x_443, 1, x_441); +lean_ctor_set(x_443, 2, x_440); +lean_ctor_set(x_443, 3, x_442); +x_444 = lean_array_get_size(x_35); +x_445 = lean_usize_of_nat(x_444); +lean_dec(x_444); +x_446 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__22(x_445, x_14, x_35); +x_447 = l_Lean_mkSepArray(x_446, x_403); +lean_dec(x_446); +x_448 = l_Array_append___rarg(x_289, x_447); +x_449 = lean_array_push(x_448, x_304); +x_450 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_450, 0, x_55); +lean_ctor_set(x_450, 1, x_291); +lean_ctor_set(x_450, 2, x_449); +x_451 = lean_array_push(x_411, x_450); +x_452 = lean_array_push(x_451, x_413); +x_453 = lean_array_push(x_452, x_304); +x_454 = lean_array_push(x_453, x_409); +x_455 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_455, 0, x_55); +lean_ctor_set(x_455, 1, x_417); +lean_ctor_set(x_455, 2, x_454); +x_456 = lean_array_push(x_299, x_455); +x_457 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_457, 0, x_55); +lean_ctor_set(x_457, 1, x_291); +lean_ctor_set(x_457, 2, x_456); +x_458 = lean_array_push(x_421, x_457); +x_459 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_459, 0, x_55); +lean_ctor_set(x_459, 1, x_423); +lean_ctor_set(x_459, 2, x_458); +x_460 = lean_array_push(x_425, x_443); +x_461 = lean_array_push(x_460, x_393); +x_462 = lean_array_push(x_461, x_304); +x_463 = lean_array_push(x_462, x_395); +x_464 = lean_array_push(x_463, x_459); +x_465 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_465, 0, x_55); +lean_ctor_set(x_465, 1, x_431); +lean_ctor_set(x_465, 2, x_464); +x_466 = lean_array_push(x_299, x_465); +x_467 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_467, 0, x_55); +lean_ctor_set(x_467, 1, x_434); +lean_ctor_set(x_467, 2, x_466); +x_468 = lean_array_push(x_299, x_467); +x_469 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_469, 0, x_55); +lean_ctor_set(x_469, 1, x_437); +lean_ctor_set(x_469, 2, x_468); +x_470 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__35; +x_471 = lean_array_push(x_470, x_438); +x_472 = lean_array_push(x_471, x_304); +x_473 = lean_array_push(x_472, x_469); +x_474 = lean_array_push(x_473, x_304); +x_475 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_475, 0, x_55); +lean_ctor_set(x_475, 1, x_291); +lean_ctor_set(x_475, 2, x_474); +x_476 = lean_array_push(x_324, x_475); +x_477 = lean_array_push(x_476, x_304); +x_478 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__53; +x_479 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_479, 0, x_55); +lean_ctor_set(x_479, 1, x_478); +lean_ctor_set(x_479, 2, x_477); +x_480 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__71; +x_481 = lean_array_push(x_480, x_358); +x_482 = lean_array_push(x_481, x_304); +x_483 = lean_array_push(x_482, x_304); +x_484 = lean_array_push(x_483, x_375); +x_485 = lean_array_push(x_484, x_479); +x_486 = lean_array_push(x_485, x_304); +x_487 = lean_array_push(x_486, x_304); +x_488 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__206; +x_489 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_489, 0, x_55); +lean_ctor_set(x_489, 1, x_488); +lean_ctor_set(x_489, 2, x_487); +x_490 = lean_array_push(x_353, x_489); +x_491 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_491, 0, x_55); +lean_ctor_set(x_491, 1, x_355); +lean_ctor_set(x_491, 2, x_490); +x_492 = lean_array_push(x_323, x_296); +x_493 = lean_array_push(x_492, x_356); +x_494 = lean_array_push(x_493, x_491); +x_495 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_495, 0, x_55); +lean_ctor_set(x_495, 1, x_291); +lean_ctor_set(x_495, 2, x_494); +x_496 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_496, 0, x_495); +lean_ctor_set(x_496, 1, x_284); +return x_496; } } else { -uint8_t x_496; -lean_dec(x_36); -lean_dec(x_32); -lean_dec(x_26); +uint8_t x_497; +lean_dec(x_35); +lean_dec(x_31); lean_dec(x_25); lean_dec(x_24); lean_dec(x_23); +lean_dec(x_22); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); -x_496 = !lean_is_exclusive(x_40); -if (x_496 == 0) +x_497 = !lean_is_exclusive(x_39); +if (x_497 == 0) { -return x_40; +return x_39; } else { -lean_object* x_497; lean_object* x_498; lean_object* x_499; -x_497 = lean_ctor_get(x_40, 0); -x_498 = lean_ctor_get(x_40, 1); +lean_object* x_498; lean_object* x_499; lean_object* x_500; +x_498 = lean_ctor_get(x_39, 0); +x_499 = lean_ctor_get(x_39, 1); +lean_inc(x_499); lean_inc(x_498); -lean_inc(x_497); -lean_dec(x_40); -x_499 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_499, 0, x_497); -lean_ctor_set(x_499, 1, x_498); -return x_499; +lean_dec(x_39); +x_500 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_500, 0, x_498); +lean_ctor_set(x_500, 1, x_499); +return x_500; } } } else { -uint8_t x_500; +uint8_t x_501; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -10499,23 +10489,23 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); -x_500 = !lean_is_exclusive(x_17); -if (x_500 == 0) +x_501 = !lean_is_exclusive(x_16); +if (x_501 == 0) { -return x_17; +return x_16; } else { -lean_object* x_501; lean_object* x_502; lean_object* x_503; -x_501 = lean_ctor_get(x_17, 0); -x_502 = lean_ctor_get(x_17, 1); +lean_object* x_502; lean_object* x_503; lean_object* x_504; +x_502 = lean_ctor_get(x_16, 0); +x_503 = lean_ctor_get(x_16, 1); +lean_inc(x_503); lean_inc(x_502); -lean_inc(x_501); -lean_dec(x_17); -x_503 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_503, 0, x_501); -lean_ctor_set(x_503, 1, x_502); -return x_503; +lean_dec(x_16); +x_504 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_504, 0, x_502); +lean_ctor_set(x_504, 1, x_503); +return x_504; } } } @@ -10706,7 +10696,7 @@ lean_inc(x_2); x_10 = lean_alloc_closure((void*)(l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__2), 10, 2); lean_closure_set(x_10, 0, x_2); lean_closure_set(x_10, 1, x_1); -x_11 = l_Lean_Server_RpcEncoding_withFields___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__22___rarg(x_1, x_2, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l_Lean_Server_RpcEncoding_withFields___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__23___rarg(x_1, x_2, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } @@ -10825,6 +10815,16 @@ lean_dec(x_8); return x_16; } } +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___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, lean_object* x_7, 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, lean_object* x_15) { +_start: +{ +lean_object* x_16; +x_16 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_10); +lean_dec(x_9); +return x_16; +} +} LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___lambda__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_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) { _start: { @@ -10834,51 +10834,51 @@ lean_dec(x_8); return x_16; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17___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_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -size_t x_13; size_t x_14; lean_object* x_15; +size_t x_12; size_t x_13; lean_object* x_14; +x_12 = lean_unbox_usize(x_2); +lean_dec(x_2); x_13 = lean_unbox_usize(x_3); lean_dec(x_3); -x_14 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_15 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_2); -return x_15; +x_14 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__17(x_1, x_12, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_1); +return x_14; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -size_t x_14; size_t x_15; lean_object* x_16; +size_t x_13; size_t x_14; lean_object* x_15; +x_13 = lean_unbox_usize(x_3); +lean_dec(x_3); x_14 = lean_unbox_usize(x_4); lean_dec(x_4); -x_15 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_16 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18(x_1, x_2, x_3, x_14, x_15, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_12); -lean_dec(x_10); +x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -return x_16; +lean_dec(x_6); +return x_15; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__19___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__19___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -size_t x_14; size_t x_15; lean_object* x_16; +size_t x_13; size_t x_14; lean_object* x_15; +x_13 = lean_unbox_usize(x_3); +lean_dec(x_3); x_14 = lean_unbox_usize(x_4); lean_dec(x_4); -x_15 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_16 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__19(x_1, x_2, x_3, x_14, x_15, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_12); -lean_dec(x_10); +x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__19(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -return x_16; +lean_dec(x_6); +return x_15; } } LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__20___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_object* x_9, lean_object* x_10) { @@ -10910,7 +10910,19 @@ x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Se return x_15; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__24___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__22___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__22(x_4, x_5, x_3); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__25___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { size_t x_13; size_t x_14; lean_object* x_15; @@ -10918,7 +10930,7 @@ x_13 = lean_unbox_usize(x_3); lean_dec(x_3); x_14 = lean_unbox_usize(x_4); lean_dec(x_4); -x_15 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__24(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_15 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__25(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_7); lean_dec(x_6); lean_dec(x_2); @@ -15592,167 +15604,240 @@ return x_26; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__38(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, size_t x_7, 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, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__38(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, size_t x_7, size_t 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, lean_object* x_15, lean_object* x_16) { _start: { -uint8_t x_16; -x_16 = lean_usize_dec_lt(x_7, x_6); -if (x_16 == 0) +uint8_t x_17; +x_17 = lean_usize_dec_lt(x_8, x_7); +if (x_17 == 0) { -lean_object* x_17; -lean_dec(x_13); +lean_object* x_18; +lean_dec(x_14); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_8); -lean_ctor_set(x_17, 1, x_15); -return x_17; +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_9); +lean_ctor_set(x_18, 1, x_16); +return x_18; } else { -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; 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_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; size_t x_42; size_t x_43; lean_object* x_44; -x_18 = lean_array_uget(x_8, x_7); -x_19 = lean_unsigned_to_nat(0u); -x_20 = lean_array_uset(x_8, x_7, x_19); -lean_inc(x_13); -x_21 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_13, x_14, x_15); -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_21, 1); +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; 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; size_t x_43; size_t x_44; lean_object* x_45; +x_19 = lean_array_uget(x_9, x_8); +x_20 = lean_unsigned_to_nat(0u); +x_21 = lean_array_uset(x_9, x_8, x_20); +lean_inc(x_14); +x_22 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_14, x_15, x_16); +x_23 = lean_ctor_get(x_22, 0); lean_inc(x_23); -lean_dec(x_21); -x_24 = lean_st_ref_get(x_14, x_23); -x_25 = lean_ctor_get(x_24, 1); -lean_inc(x_25); -lean_dec(x_24); -x_26 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__1; -lean_inc(x_2); -x_27 = lean_name_mk_string(x_2, x_26); -x_28 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__2; -x_29 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_29, 0, x_22); -lean_ctor_set(x_29, 1, x_28); -x_30 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__60; -lean_inc(x_2); -x_31 = lean_name_mk_string(x_2, x_30); -lean_inc(x_5); -x_32 = lean_mk_syntax_ident(x_5); -lean_inc(x_4); -x_33 = lean_array_push(x_4, x_18); -x_34 = lean_box(2); -lean_inc(x_1); -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_1); -lean_ctor_set(x_35, 2, x_33); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_st_ref_get(x_15, x_24); +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +x_27 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__1; lean_inc(x_3); -x_36 = lean_array_push(x_3, x_32); -x_37 = lean_array_push(x_36, x_35); -x_38 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_38, 0, x_34); -lean_ctor_set(x_38, 1, x_31); -lean_ctor_set(x_38, 2, x_37); +x_28 = lean_name_mk_string(x_3, x_27); +x_29 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__2; +x_30 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_30, 0, x_23); +lean_ctor_set(x_30, 1, x_29); +x_31 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__60; lean_inc(x_3); -x_39 = lean_array_push(x_3, x_29); -x_40 = lean_array_push(x_39, x_38); -x_41 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_41, 0, x_34); -lean_ctor_set(x_41, 1, x_27); -lean_ctor_set(x_41, 2, x_40); -x_42 = 1; -x_43 = lean_usize_add(x_7, x_42); -x_44 = lean_array_uset(x_20, x_7, x_41); -x_7 = x_43; +x_32 = lean_name_mk_string(x_3, x_31); +lean_inc(x_6); +x_33 = lean_mk_syntax_ident(x_6); +lean_inc(x_5); +x_34 = lean_array_push(x_5, x_19); +x_35 = lean_box(2); +lean_inc(x_2); +x_36 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_2); +lean_ctor_set(x_36, 2, x_34); +lean_inc(x_4); +x_37 = lean_array_push(x_4, x_33); +x_38 = lean_array_push(x_37, x_36); +x_39 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_39, 0, x_35); +lean_ctor_set(x_39, 1, x_32); +lean_ctor_set(x_39, 2, x_38); +lean_inc(x_4); +x_40 = lean_array_push(x_4, x_30); +x_41 = lean_array_push(x_40, x_39); +x_42 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_42, 0, x_35); +lean_ctor_set(x_42, 1, x_28); +lean_ctor_set(x_42, 2, x_41); +x_43 = 1; +x_44 = lean_usize_add(x_8, x_43); +x_45 = lean_array_uset(x_21, x_8, x_42); x_8 = x_44; -x_15 = x_25; +x_9 = x_45; +x_16 = x_26; goto _start; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__39(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, size_t x_7, 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, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__39(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, size_t x_7, size_t 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, lean_object* x_15, lean_object* x_16) { _start: { -uint8_t x_16; -x_16 = lean_usize_dec_lt(x_7, x_6); -if (x_16 == 0) +uint8_t x_17; +x_17 = lean_usize_dec_lt(x_8, x_7); +if (x_17 == 0) { -lean_object* x_17; -lean_dec(x_13); -lean_dec(x_5); +lean_object* x_18; +lean_dec(x_14); +lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_8); -lean_ctor_set(x_17, 1, x_15); -return x_17; +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_9); +lean_ctor_set(x_18, 1, x_16); +return x_18; } else { -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; 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_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; size_t x_42; size_t x_43; lean_object* x_44; -x_18 = lean_array_uget(x_8, x_7); -x_19 = lean_unsigned_to_nat(0u); -x_20 = lean_array_uset(x_8, x_7, x_19); -lean_inc(x_13); -x_21 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_13, x_14, x_15); -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_21, 1); +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; 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; size_t x_43; size_t x_44; lean_object* x_45; +x_19 = lean_array_uget(x_9, x_8); +x_20 = lean_unsigned_to_nat(0u); +x_21 = lean_array_uset(x_9, x_8, x_20); +lean_inc(x_14); +x_22 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_14, x_15, x_16); +x_23 = lean_ctor_get(x_22, 0); lean_inc(x_23); -lean_dec(x_21); -x_24 = lean_st_ref_get(x_14, x_23); -x_25 = lean_ctor_get(x_24, 1); -lean_inc(x_25); -lean_dec(x_24); -x_26 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__1; +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_st_ref_get(x_15, x_24); +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +x_27 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__1; lean_inc(x_2); -x_27 = lean_name_mk_string(x_2, x_26); -x_28 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__2; -x_29 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_29, 0, x_22); -lean_ctor_set(x_29, 1, x_28); -x_30 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__60; +x_28 = lean_name_mk_string(x_2, x_27); +x_29 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__18___closed__2; +x_30 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_30, 0, x_23); +lean_ctor_set(x_30, 1, x_29); +x_31 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__60; lean_inc(x_2); -x_31 = lean_name_mk_string(x_2, x_30); -lean_inc(x_5); -x_32 = lean_mk_syntax_ident(x_5); +x_32 = lean_name_mk_string(x_2, x_31); +lean_inc(x_6); +x_33 = lean_mk_syntax_ident(x_6); lean_inc(x_4); -x_33 = lean_array_push(x_4, x_18); -x_34 = lean_box(2); +x_34 = lean_array_push(x_4, x_19); +x_35 = lean_box(2); lean_inc(x_1); -x_35 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_1); -lean_ctor_set(x_35, 2, x_33); +x_36 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_1); +lean_ctor_set(x_36, 2, x_34); lean_inc(x_3); -x_36 = lean_array_push(x_3, x_32); -x_37 = lean_array_push(x_36, x_35); -x_38 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_38, 0, x_34); -lean_ctor_set(x_38, 1, x_31); -lean_ctor_set(x_38, 2, x_37); +x_37 = lean_array_push(x_3, x_33); +x_38 = lean_array_push(x_37, x_36); +x_39 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_39, 0, x_35); +lean_ctor_set(x_39, 1, x_32); +lean_ctor_set(x_39, 2, x_38); lean_inc(x_3); -x_39 = lean_array_push(x_3, x_29); -x_40 = lean_array_push(x_39, x_38); -x_41 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_41, 0, x_34); -lean_ctor_set(x_41, 1, x_27); -lean_ctor_set(x_41, 2, x_40); -x_42 = 1; -x_43 = lean_usize_add(x_7, x_42); -x_44 = lean_array_uset(x_20, x_7, x_41); -x_7 = x_43; +x_40 = lean_array_push(x_3, x_30); +x_41 = lean_array_push(x_40, x_39); +x_42 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_42, 0, x_35); +lean_ctor_set(x_42, 1, x_28); +lean_ctor_set(x_42, 2, x_41); +x_43 = 1; +x_44 = lean_usize_add(x_8, x_43); +x_45 = lean_array_uset(x_21, x_8, x_42); x_8 = x_44; -x_15 = x_25; +x_9 = x_45; +x_16 = x_26; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__40(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; +x_5 = lean_usize_dec_lt(x_3, x_2); +if (x_5 == 0) +{ +return x_4; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; size_t x_10; size_t x_11; lean_object* x_12; +x_6 = lean_array_uget(x_4, x_3); +x_7 = lean_unsigned_to_nat(0u); +x_8 = lean_array_uset(x_4, x_3, x_7); +x_9 = lean_mk_syntax_ident(x_6); +x_10 = 1; +x_11 = lean_usize_add(x_3, x_10); +x_12 = lean_array_uset(x_8, x_3, x_9); +x_3 = x_11; +x_4 = x_12; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = 1; +x_9 = lean_usize_add(x_2, x_8); +x_10 = lean_array_uset(x_7, x_2, x_5); +x_2 = x_9; +x_3 = x_10; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__42(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = 1; +x_9 = lean_usize_add(x_2, x_8); +x_10 = lean_array_uset(x_7, x_2, x_5); +x_2 = x_9; +x_3 = x_10; goto _start; } } } -static lean_object* _init_l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___lambda__1___closed__1() { +static lean_object* _init_l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___lambda__1___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -15762,47 +15847,47 @@ x_3 = l_instInhabited___rarg(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___lambda__1___closed__1; +x_2 = l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___lambda__1___closed__1; return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { lean_object* x_13; lean_object* x_14; x_13 = lean_array_push(x_1, x_5); -x_14 = l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg(x_2, x_3, x_4, x_13, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_14 = l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg(x_2, x_3, x_4, x_13, x_6, x_7, x_8, x_9, x_10, x_11, x_12); return x_14; } } -static lean_object* _init_l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___closed__1() { +static lean_object* _init_l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___lambda__1___boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___lambda__1___boxed), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___closed__2() { +static lean_object* _init_l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___closed__1; +x_1 = l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___closed__1; x_2 = lean_alloc_closure((void*)(l_instInhabitedForAll__1___rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___closed__3() { +static lean_object* _init_l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___closed__3() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_instInhabitedBinderInfo; -x_2 = l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___closed__2; +x_2 = l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___closed__2; x_3 = lean_box(x_1); x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); @@ -15810,19 +15895,19 @@ lean_ctor_set(x_4, 1, x_2); return x_4; } } -static lean_object* _init_l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___closed__4() { +static lean_object* _init_l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_instInhabitedName; -x_2 = l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___closed__3; +x_2 = l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___closed__3; x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___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, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___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, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; lean_object* x_13; uint8_t x_14; @@ -15842,7 +15927,7 @@ return x_15; else { 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; -x_16 = l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___closed__4; +x_16 = l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___closed__4; x_17 = lean_array_get(x_16, x_1, x_12); lean_dec(x_12); x_18 = lean_ctor_get(x_17, 1); @@ -15871,7 +15956,7 @@ lean_inc(x_23); x_24 = lean_ctor_get(x_22, 1); lean_inc(x_24); lean_dec(x_22); -x_25 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___lambda__2), 12, 4); +x_25 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___lambda__2), 12, 4); lean_closure_set(x_25, 0, x_4); lean_closure_set(x_25, 1, x_1); lean_closure_set(x_25, 2, x_2); @@ -15918,28 +16003,28 @@ return x_31; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg), 11, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg), 11, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__40___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, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__43___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, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; lean_object* x_12; x_11 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__11; -x_12 = l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg(x_2, x_3, x_1, x_11, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_12 = l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg(x_2, x_3, x_1, x_11, x_4, x_5, x_6, x_7, x_8, x_9, x_10); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__40(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__43(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecls___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__40___rarg), 10, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecls___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__43___rarg), 10, 0); return x_2; } } @@ -16217,26 +16302,26 @@ x_1 = lean_mk_string_from_bytes("=>", 2); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__3(lean_object* x_1, size_t 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, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__3(lean_object* x_1, lean_object* x_2, size_t 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, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; size_t x_20; lean_object* x_21; -x_16 = lean_box(0); +lean_object* x_17; lean_object* x_18; lean_object* x_19; size_t x_20; lean_object* x_21; +lean_inc(x_2); lean_inc(x_1); -x_17 = l_Lean_mkConst(x_1, x_16); -lean_inc(x_7); -x_18 = l_Array_reverse___rarg(x_7); +x_17 = l_Lean_mkConst(x_1, x_2); +lean_inc(x_8); +x_18 = l_Array_reverse___rarg(x_8); x_19 = lean_array_get_size(x_18); x_20 = lean_usize_of_nat(x_19); lean_dec(x_19); +lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_5); -x_21 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__36(x_5, x_18, x_20, x_2, x_17, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_inc(x_6); +x_21 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__36(x_6, x_18, x_20, x_3, x_17, x_10, x_11, x_12, x_13, x_14, x_15, x_16); lean_dec(x_18); if (lean_obj_tag(x_21) == 0) { @@ -16247,11 +16332,11 @@ x_23 = lean_ctor_get(x_21, 1); lean_inc(x_23); lean_dec(x_21); x_24 = lean_box(0); +lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -lean_inc(x_11); -x_25 = l_Lean_PrettyPrinter_delab(x_22, x_24, x_11, x_12, x_13, x_14, x_23); +x_25 = l_Lean_PrettyPrinter_delab(x_22, x_24, x_12, x_13, x_14, x_15, x_23); if (lean_obj_tag(x_25) == 0) { 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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; size_t x_73; lean_object* x_74; @@ -16260,14 +16345,14 @@ lean_inc(x_26); x_27 = lean_ctor_get(x_25, 1); lean_inc(x_27); lean_dec(x_25); -lean_inc(x_13); -x_28 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_13, x_14, x_27); +lean_inc(x_14); +x_28 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_14, x_15, x_27); x_29 = lean_ctor_get(x_28, 0); lean_inc(x_29); x_30 = lean_ctor_get(x_28, 1); lean_inc(x_30); lean_dec(x_28); -x_31 = lean_st_ref_get(x_14, x_30); +x_31 = lean_st_ref_get(x_15, x_30); x_32 = lean_ctor_get(x_31, 1); lean_inc(x_32); lean_dec(x_31); @@ -16281,7 +16366,7 @@ x_36 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWit x_37 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_37, 0, x_35); lean_ctor_set(x_37, 1, x_36); -lean_ctor_set(x_37, 2, x_3); +lean_ctor_set(x_37, 2, x_4); x_38 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__42; lean_inc(x_37); x_39 = lean_array_push(x_38, x_37); @@ -16300,7 +16385,7 @@ x_46 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_46, 0, x_35); lean_ctor_set(x_46, 1, x_45); lean_ctor_set(x_46, 2, x_44); -x_47 = l_Lean_Name_getString_x21(x_6); +x_47 = l_Lean_Name_getString_x21(x_7); x_48 = lean_box(0); lean_inc(x_47); x_49 = lean_name_mk_string(x_48, x_47); @@ -16340,416 +16425,424 @@ x_71 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_71, 0, x_35); lean_ctor_set(x_71, 1, x_70); lean_ctor_set(x_71, 2, x_69); -x_72 = lean_array_get_size(x_7); +x_72 = lean_array_get_size(x_8); x_73 = lean_usize_of_nat(x_72); lean_dec(x_72); +lean_inc(x_15); lean_inc(x_14); -lean_inc(x_13); -x_74 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__37(x_73, x_2, x_7, x_9, x_10, x_11, x_12, x_13, x_14, x_32); +x_74 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__37(x_73, x_3, x_8, x_10, x_11, x_12, x_13, x_14, x_15, x_32); if (lean_obj_tag(x_74) == 0) { -lean_object* x_75; lean_object* x_76; lean_object* x_77; size_t x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; uint8_t x_143; +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; size_t x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; uint8_t x_145; x_75 = lean_ctor_get(x_74, 0); lean_inc(x_75); x_76 = lean_ctor_get(x_74, 1); lean_inc(x_76); lean_dec(x_74); -x_77 = lean_array_get_size(x_75); -x_78 = lean_usize_of_nat(x_77); -lean_dec(x_77); -x_79 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__17; -x_80 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__225; -lean_inc(x_13); +x_77 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_matchF___closed__2; +lean_inc(x_2); +x_78 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_78, 0, x_77); +lean_ctor_set(x_78, 1, x_2); +x_79 = lean_array_get_size(x_75); +x_80 = lean_usize_of_nat(x_79); +lean_dec(x_79); +x_81 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__17; +x_82 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__225; +lean_inc(x_14); lean_inc(x_75); -x_81 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__38(x_36, x_79, x_53, x_58, x_80, x_78, x_2, x_75, x_9, x_10, x_11, x_12, x_13, x_14, x_76); -x_82 = lean_ctor_get(x_81, 0); -lean_inc(x_82); -x_83 = lean_ctor_get(x_81, 1); -lean_inc(x_83); -lean_dec(x_81); -lean_inc(x_47); -x_84 = lean_name_mk_string(x_1, x_47); -x_85 = lean_mk_syntax_ident(x_84); +x_83 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__38(x_2, x_36, x_81, x_53, x_58, x_82, x_80, x_3, x_75, x_10, x_11, x_12, x_13, x_14, x_15, x_76); +lean_dec(x_2); +x_84 = lean_ctor_get(x_83, 0); +lean_inc(x_84); +x_85 = lean_ctor_get(x_83, 1); lean_inc(x_85); -x_86 = l_Lean_Syntax_mkApp(x_85, x_82); -lean_inc(x_13); -x_87 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_13, x_14, x_83); -x_88 = lean_ctor_get(x_87, 0); -lean_inc(x_88); -x_89 = lean_ctor_get(x_87, 1); -lean_inc(x_89); -lean_dec(x_87); -x_90 = lean_st_ref_get(x_14, x_89); -x_91 = lean_ctor_get(x_90, 1); +lean_dec(x_83); +lean_inc(x_47); +x_86 = lean_name_mk_string(x_1, x_47); +x_87 = lean_mk_syntax_ident(x_86); +lean_inc(x_87); +x_88 = l_Lean_Syntax_mkApp(x_87, x_84); +lean_inc(x_14); +x_89 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_14, x_15, x_85); +x_90 = lean_ctor_get(x_89, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_89, 1); lean_inc(x_91); -lean_dec(x_90); -x_92 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__66; -x_93 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_93, 0, x_88); -lean_ctor_set(x_93, 1, x_92); -x_94 = lean_array_push(x_58, x_86); -x_95 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_95, 0, x_35); -lean_ctor_set(x_95, 1, x_36); -lean_ctor_set(x_95, 2, x_94); -x_96 = lean_array_push(x_53, x_93); -x_97 = lean_array_push(x_96, x_95); -x_98 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__65; -x_99 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_99, 0, x_35); -lean_ctor_set(x_99, 1, x_98); -lean_ctor_set(x_99, 2, x_97); -lean_inc(x_13); -x_100 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_13, x_14, x_91); -x_101 = lean_ctor_get(x_100, 0); -lean_inc(x_101); -x_102 = lean_ctor_get(x_100, 1); -lean_inc(x_102); -lean_dec(x_100); -x_103 = lean_st_ref_get(x_14, x_102); -x_104 = lean_ctor_get(x_103, 1); +lean_dec(x_89); +x_92 = lean_st_ref_get(x_15, x_91); +x_93 = lean_ctor_get(x_92, 1); +lean_inc(x_93); +lean_dec(x_92); +x_94 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__66; +x_95 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_95, 0, x_90); +lean_ctor_set(x_95, 1, x_94); +x_96 = lean_array_push(x_58, x_88); +x_97 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_97, 0, x_35); +lean_ctor_set(x_97, 1, x_36); +lean_ctor_set(x_97, 2, x_96); +x_98 = lean_array_push(x_53, x_95); +x_99 = lean_array_push(x_98, x_97); +x_100 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__65; +x_101 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_101, 0, x_35); +lean_ctor_set(x_101, 1, x_100); +lean_ctor_set(x_101, 2, x_99); +lean_inc(x_14); +x_102 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_14, x_15, x_93); +x_103 = lean_ctor_get(x_102, 0); +lean_inc(x_103); +x_104 = lean_ctor_get(x_102, 1); lean_inc(x_104); -lean_dec(x_103); -lean_inc(x_101); -x_105 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_105, 0, x_101); -lean_ctor_set(x_105, 1, x_33); -x_106 = lean_ctor_get(x_4, 0); +lean_dec(x_102); +x_105 = lean_st_ref_get(x_15, x_104); +x_106 = lean_ctor_get(x_105, 1); lean_inc(x_106); -lean_dec(x_4); -x_107 = lean_ctor_get(x_106, 0); -lean_inc(x_107); -lean_dec(x_106); -x_108 = lean_name_mk_string(x_107, x_47); -x_109 = lean_mk_syntax_ident(x_108); -lean_inc(x_75); +lean_dec(x_105); +lean_inc(x_103); +x_107 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_107, 0, x_103); +lean_ctor_set(x_107, 1, x_33); +x_108 = lean_ctor_get(x_5, 0); +lean_inc(x_108); +lean_dec(x_5); +x_109 = lean_ctor_get(x_108, 0); lean_inc(x_109); -x_110 = l_Lean_Syntax_mkApp(x_109, x_75); -x_111 = lean_array_push(x_58, x_110); -x_112 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_112, 0, x_35); -lean_ctor_set(x_112, 1, x_36); -lean_ctor_set(x_112, 2, x_111); +lean_dec(x_108); +x_110 = lean_name_mk_string(x_109, x_47); +x_111 = lean_mk_syntax_ident(x_110); +lean_inc(x_75); +lean_inc(x_111); +x_112 = l_Lean_Syntax_mkApp(x_111, x_75); x_113 = lean_array_push(x_58, x_112); x_114 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_114, 0, x_35); lean_ctor_set(x_114, 1, x_36); lean_ctor_set(x_114, 2, x_113); -x_115 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__3___closed__6; -x_116 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_116, 0, x_101); -lean_ctor_set(x_116, 1, x_115); -x_117 = lean_array_push(x_65, x_105); -x_118 = lean_array_push(x_117, x_114); -x_119 = lean_array_push(x_118, x_116); -x_120 = lean_array_push(x_119, x_99); -x_121 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__3___closed__5; -x_122 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_122, 0, x_35); -lean_ctor_set(x_122, 1, x_121); -lean_ctor_set(x_122, 2, x_120); -x_123 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__232; -lean_inc(x_13); +x_115 = lean_array_push(x_58, x_114); +x_116 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_116, 0, x_35); +lean_ctor_set(x_116, 1, x_36); +lean_ctor_set(x_116, 2, x_115); +x_117 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__3___closed__6; +x_118 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_118, 0, x_103); +lean_ctor_set(x_118, 1, x_117); +x_119 = lean_array_push(x_65, x_107); +x_120 = lean_array_push(x_119, x_116); +x_121 = lean_array_push(x_120, x_118); +x_122 = lean_array_push(x_121, x_101); +x_123 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__3___closed__5; +x_124 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_124, 0, x_35); +lean_ctor_set(x_124, 1, x_123); +lean_ctor_set(x_124, 2, x_122); +x_125 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__232; +lean_inc(x_14); lean_inc(x_75); -x_124 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__39(x_36, x_79, x_53, x_58, x_123, x_78, x_2, x_75, x_9, x_10, x_11, x_12, x_13, x_14, x_104); +x_126 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__39(x_36, x_81, x_53, x_58, x_78, x_125, x_80, x_3, x_75, x_10, x_11, x_12, x_13, x_14, x_15, x_106); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); -x_125 = lean_ctor_get(x_124, 0); -lean_inc(x_125); -x_126 = lean_ctor_get(x_124, 1); -lean_inc(x_126); -lean_dec(x_124); -x_127 = l_Lean_Syntax_mkApp(x_109, x_125); -lean_inc(x_13); -x_128 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_13, x_14, x_126); -x_129 = lean_ctor_get(x_128, 0); -lean_inc(x_129); -x_130 = lean_ctor_get(x_128, 1); -lean_inc(x_130); -lean_dec(x_128); -x_131 = lean_st_ref_get(x_14, x_130); -x_132 = lean_ctor_get(x_131, 1); +lean_dec(x_78); +x_127 = lean_ctor_get(x_126, 0); +lean_inc(x_127); +x_128 = lean_ctor_get(x_126, 1); +lean_inc(x_128); +lean_dec(x_126); +x_129 = l_Lean_Syntax_mkApp(x_111, x_127); +lean_inc(x_14); +x_130 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_14, x_15, x_128); +x_131 = lean_ctor_get(x_130, 0); +lean_inc(x_131); +x_132 = lean_ctor_get(x_130, 1); lean_inc(x_132); -lean_dec(x_131); -x_133 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_133, 0, x_129); -lean_ctor_set(x_133, 1, x_92); -x_134 = lean_array_push(x_58, x_127); -x_135 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_135, 0, x_35); -lean_ctor_set(x_135, 1, x_36); -lean_ctor_set(x_135, 2, x_134); -x_136 = lean_array_push(x_53, x_133); -x_137 = lean_array_push(x_136, x_135); -x_138 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_138, 0, x_35); -lean_ctor_set(x_138, 1, x_98); -lean_ctor_set(x_138, 2, x_137); -x_139 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_13, x_14, x_132); -x_140 = lean_ctor_get(x_139, 0); -lean_inc(x_140); -x_141 = lean_ctor_get(x_139, 1); -lean_inc(x_141); -lean_dec(x_139); -x_142 = lean_st_ref_get(x_14, x_141); -lean_dec(x_14); -x_143 = !lean_is_exclusive(x_142); -if (x_143 == 0) +lean_dec(x_130); +x_133 = lean_st_ref_get(x_15, x_132); +x_134 = lean_ctor_get(x_133, 1); +lean_inc(x_134); +lean_dec(x_133); +x_135 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_135, 0, x_131); +lean_ctor_set(x_135, 1, x_94); +x_136 = lean_array_push(x_58, x_129); +x_137 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_137, 0, x_35); +lean_ctor_set(x_137, 1, x_36); +lean_ctor_set(x_137, 2, x_136); +x_138 = lean_array_push(x_53, x_135); +x_139 = lean_array_push(x_138, x_137); +x_140 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_140, 0, x_35); +lean_ctor_set(x_140, 1, x_100); +lean_ctor_set(x_140, 2, x_139); +x_141 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_14, x_15, x_134); +x_142 = lean_ctor_get(x_141, 0); +lean_inc(x_142); +x_143 = lean_ctor_get(x_141, 1); +lean_inc(x_143); +lean_dec(x_141); +x_144 = lean_st_ref_get(x_15, x_143); +lean_dec(x_15); +x_145 = !lean_is_exclusive(x_144); +if (x_145 == 0) { -lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; uint8_t x_157; -x_144 = lean_ctor_get(x_142, 0); -lean_dec(x_144); -lean_inc(x_140); -x_145 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_145, 0, x_140); -lean_ctor_set(x_145, 1, x_33); -x_146 = l_Lean_Syntax_mkApp(x_85, x_75); -x_147 = lean_array_push(x_58, x_146); -x_148 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_148, 0, x_35); -lean_ctor_set(x_148, 1, x_36); -lean_ctor_set(x_148, 2, x_147); +lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; uint8_t x_159; +x_146 = lean_ctor_get(x_144, 0); +lean_dec(x_146); +lean_inc(x_142); +x_147 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_147, 0, x_142); +lean_ctor_set(x_147, 1, x_33); +x_148 = l_Lean_Syntax_mkApp(x_87, x_75); x_149 = lean_array_push(x_58, x_148); x_150 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_150, 0, x_35); lean_ctor_set(x_150, 1, x_36); lean_ctor_set(x_150, 2, x_149); -x_151 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_151, 0, x_140); -lean_ctor_set(x_151, 1, x_115); -x_152 = lean_array_push(x_65, x_145); -x_153 = lean_array_push(x_152, x_150); -x_154 = lean_array_push(x_153, x_151); -x_155 = lean_array_push(x_154, x_138); -x_156 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_156, 0, x_35); -lean_ctor_set(x_156, 1, x_121); -lean_ctor_set(x_156, 2, x_155); -x_157 = !lean_is_exclusive(x_5); -if (x_157 == 0) -{ -lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; -x_158 = lean_ctor_get(x_5, 3); -x_159 = lean_ctor_get(x_5, 4); -x_160 = lean_ctor_get(x_5, 5); -x_161 = lean_array_push(x_158, x_71); -x_162 = lean_array_push(x_159, x_122); -x_163 = lean_array_push(x_160, x_156); -lean_ctor_set(x_5, 5, x_163); -lean_ctor_set(x_5, 4, x_162); -lean_ctor_set(x_5, 3, x_161); -lean_ctor_set(x_142, 0, x_5); -return x_142; -} -else -{ -lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; -x_164 = lean_ctor_get(x_5, 0); -x_165 = lean_ctor_get(x_5, 1); -x_166 = lean_ctor_get(x_5, 2); -x_167 = lean_ctor_get(x_5, 3); -x_168 = lean_ctor_get(x_5, 4); -x_169 = lean_ctor_get(x_5, 5); +x_151 = lean_array_push(x_58, x_150); +x_152 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_152, 0, x_35); +lean_ctor_set(x_152, 1, x_36); +lean_ctor_set(x_152, 2, x_151); +x_153 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_153, 0, x_142); +lean_ctor_set(x_153, 1, x_117); +x_154 = lean_array_push(x_65, x_147); +x_155 = lean_array_push(x_154, x_152); +x_156 = lean_array_push(x_155, x_153); +x_157 = lean_array_push(x_156, x_140); +x_158 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_158, 0, x_35); +lean_ctor_set(x_158, 1, x_123); +lean_ctor_set(x_158, 2, x_157); +x_159 = !lean_is_exclusive(x_6); +if (x_159 == 0) +{ +lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; +x_160 = lean_ctor_get(x_6, 3); +x_161 = lean_ctor_get(x_6, 4); +x_162 = lean_ctor_get(x_6, 5); +x_163 = lean_array_push(x_160, x_71); +x_164 = lean_array_push(x_161, x_124); +x_165 = lean_array_push(x_162, x_158); +lean_ctor_set(x_6, 5, x_165); +lean_ctor_set(x_6, 4, x_164); +lean_ctor_set(x_6, 3, x_163); +lean_ctor_set(x_144, 0, x_6); +return x_144; +} +else +{ +lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; +x_166 = lean_ctor_get(x_6, 0); +x_167 = lean_ctor_get(x_6, 1); +x_168 = lean_ctor_get(x_6, 2); +x_169 = lean_ctor_get(x_6, 3); +x_170 = lean_ctor_get(x_6, 4); +x_171 = lean_ctor_get(x_6, 5); +lean_inc(x_171); +lean_inc(x_170); lean_inc(x_169); lean_inc(x_168); lean_inc(x_167); lean_inc(x_166); -lean_inc(x_165); -lean_inc(x_164); -lean_dec(x_5); -x_170 = lean_array_push(x_167, x_71); -x_171 = lean_array_push(x_168, x_122); -x_172 = lean_array_push(x_169, x_156); -x_173 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_173, 0, x_164); -lean_ctor_set(x_173, 1, x_165); -lean_ctor_set(x_173, 2, x_166); -lean_ctor_set(x_173, 3, x_170); -lean_ctor_set(x_173, 4, x_171); -lean_ctor_set(x_173, 5, x_172); -lean_ctor_set(x_142, 0, x_173); -return x_142; -} -} -else -{ -lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; -x_174 = lean_ctor_get(x_142, 1); -lean_inc(x_174); -lean_dec(x_142); -lean_inc(x_140); -x_175 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_175, 0, x_140); -lean_ctor_set(x_175, 1, x_33); -x_176 = l_Lean_Syntax_mkApp(x_85, x_75); -x_177 = lean_array_push(x_58, x_176); -x_178 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_178, 0, x_35); -lean_ctor_set(x_178, 1, x_36); -lean_ctor_set(x_178, 2, x_177); +lean_dec(x_6); +x_172 = lean_array_push(x_169, x_71); +x_173 = lean_array_push(x_170, x_124); +x_174 = lean_array_push(x_171, x_158); +x_175 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_175, 0, x_166); +lean_ctor_set(x_175, 1, x_167); +lean_ctor_set(x_175, 2, x_168); +lean_ctor_set(x_175, 3, x_172); +lean_ctor_set(x_175, 4, x_173); +lean_ctor_set(x_175, 5, x_174); +lean_ctor_set(x_144, 0, x_175); +return x_144; +} +} +else +{ +lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; +x_176 = lean_ctor_get(x_144, 1); +lean_inc(x_176); +lean_dec(x_144); +lean_inc(x_142); +x_177 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_177, 0, x_142); +lean_ctor_set(x_177, 1, x_33); +x_178 = l_Lean_Syntax_mkApp(x_87, x_75); x_179 = lean_array_push(x_58, x_178); x_180 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_180, 0, x_35); lean_ctor_set(x_180, 1, x_36); lean_ctor_set(x_180, 2, x_179); -x_181 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_181, 0, x_140); -lean_ctor_set(x_181, 1, x_115); -x_182 = lean_array_push(x_65, x_175); -x_183 = lean_array_push(x_182, x_180); -x_184 = lean_array_push(x_183, x_181); -x_185 = lean_array_push(x_184, x_138); -x_186 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_186, 0, x_35); -lean_ctor_set(x_186, 1, x_121); -lean_ctor_set(x_186, 2, x_185); -x_187 = lean_ctor_get(x_5, 0); -lean_inc(x_187); -x_188 = lean_ctor_get(x_5, 1); -lean_inc(x_188); -x_189 = lean_ctor_get(x_5, 2); +x_181 = lean_array_push(x_58, x_180); +x_182 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_182, 0, x_35); +lean_ctor_set(x_182, 1, x_36); +lean_ctor_set(x_182, 2, x_181); +x_183 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_183, 0, x_142); +lean_ctor_set(x_183, 1, x_117); +x_184 = lean_array_push(x_65, x_177); +x_185 = lean_array_push(x_184, x_182); +x_186 = lean_array_push(x_185, x_183); +x_187 = lean_array_push(x_186, x_140); +x_188 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_188, 0, x_35); +lean_ctor_set(x_188, 1, x_123); +lean_ctor_set(x_188, 2, x_187); +x_189 = lean_ctor_get(x_6, 0); lean_inc(x_189); -x_190 = lean_ctor_get(x_5, 3); +x_190 = lean_ctor_get(x_6, 1); lean_inc(x_190); -x_191 = lean_ctor_get(x_5, 4); +x_191 = lean_ctor_get(x_6, 2); lean_inc(x_191); -x_192 = lean_ctor_get(x_5, 5); +x_192 = lean_ctor_get(x_6, 3); lean_inc(x_192); -if (lean_is_exclusive(x_5)) { - lean_ctor_release(x_5, 0); - lean_ctor_release(x_5, 1); - lean_ctor_release(x_5, 2); - lean_ctor_release(x_5, 3); - lean_ctor_release(x_5, 4); - lean_ctor_release(x_5, 5); - x_193 = x_5; +x_193 = lean_ctor_get(x_6, 4); +lean_inc(x_193); +x_194 = lean_ctor_get(x_6, 5); +lean_inc(x_194); +if (lean_is_exclusive(x_6)) { + lean_ctor_release(x_6, 0); + lean_ctor_release(x_6, 1); + lean_ctor_release(x_6, 2); + lean_ctor_release(x_6, 3); + lean_ctor_release(x_6, 4); + lean_ctor_release(x_6, 5); + x_195 = x_6; } else { - lean_dec_ref(x_5); - x_193 = lean_box(0); -} -x_194 = lean_array_push(x_190, x_71); -x_195 = lean_array_push(x_191, x_122); -x_196 = lean_array_push(x_192, x_186); -if (lean_is_scalar(x_193)) { - x_197 = lean_alloc_ctor(0, 6, 0); + lean_dec_ref(x_6); + x_195 = lean_box(0); +} +x_196 = lean_array_push(x_192, x_71); +x_197 = lean_array_push(x_193, x_124); +x_198 = lean_array_push(x_194, x_188); +if (lean_is_scalar(x_195)) { + x_199 = lean_alloc_ctor(0, 6, 0); } else { - x_197 = x_193; + x_199 = x_195; } -lean_ctor_set(x_197, 0, x_187); -lean_ctor_set(x_197, 1, x_188); -lean_ctor_set(x_197, 2, x_189); -lean_ctor_set(x_197, 3, x_194); -lean_ctor_set(x_197, 4, x_195); -lean_ctor_set(x_197, 5, x_196); -x_198 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_198, 0, x_197); -lean_ctor_set(x_198, 1, x_174); -return x_198; +lean_ctor_set(x_199, 0, x_189); +lean_ctor_set(x_199, 1, x_190); +lean_ctor_set(x_199, 2, x_191); +lean_ctor_set(x_199, 3, x_196); +lean_ctor_set(x_199, 4, x_197); +lean_ctor_set(x_199, 5, x_198); +x_200 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_200, 0, x_199); +lean_ctor_set(x_200, 1, x_176); +return x_200; } } else { -uint8_t x_199; +uint8_t x_201; lean_dec(x_71); lean_dec(x_47); +lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); +lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); +lean_dec(x_2); lean_dec(x_1); -x_199 = !lean_is_exclusive(x_74); -if (x_199 == 0) +x_201 = !lean_is_exclusive(x_74); +if (x_201 == 0) { return x_74; } else { -lean_object* x_200; lean_object* x_201; lean_object* x_202; -x_200 = lean_ctor_get(x_74, 0); -x_201 = lean_ctor_get(x_74, 1); -lean_inc(x_201); -lean_inc(x_200); +lean_object* x_202; lean_object* x_203; lean_object* x_204; +x_202 = lean_ctor_get(x_74, 0); +x_203 = lean_ctor_get(x_74, 1); +lean_inc(x_203); +lean_inc(x_202); lean_dec(x_74); -x_202 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_202, 0, x_200); -lean_ctor_set(x_202, 1, x_201); -return x_202; +x_204 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_204, 0, x_202); +lean_ctor_set(x_204, 1, x_203); +return x_204; } } } else { -uint8_t x_203; +uint8_t x_205; +lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_7); +lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -x_203 = !lean_is_exclusive(x_25); -if (x_203 == 0) +x_205 = !lean_is_exclusive(x_25); +if (x_205 == 0) { return x_25; } else { -lean_object* x_204; lean_object* x_205; lean_object* x_206; -x_204 = lean_ctor_get(x_25, 0); -x_205 = lean_ctor_get(x_25, 1); -lean_inc(x_205); -lean_inc(x_204); +lean_object* x_206; lean_object* x_207; lean_object* x_208; +x_206 = lean_ctor_get(x_25, 0); +x_207 = lean_ctor_get(x_25, 1); +lean_inc(x_207); +lean_inc(x_206); lean_dec(x_25); -x_206 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_206, 0, x_204); -lean_ctor_set(x_206, 1, x_205); -return x_206; +x_208 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_208, 0, x_206); +lean_ctor_set(x_208, 1, x_207); +return x_208; } } } else { -uint8_t x_207; +uint8_t x_209; +lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_7); +lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -x_207 = !lean_is_exclusive(x_21); -if (x_207 == 0) +x_209 = !lean_is_exclusive(x_21); +if (x_209 == 0) { return x_21; } else { -lean_object* x_208; lean_object* x_209; lean_object* x_210; -x_208 = lean_ctor_get(x_21, 0); -x_209 = lean_ctor_get(x_21, 1); -lean_inc(x_209); -lean_inc(x_208); +lean_object* x_210; lean_object* x_211; lean_object* x_212; +x_210 = lean_ctor_get(x_21, 0); +x_211 = lean_ctor_get(x_21, 1); +lean_inc(x_211); +lean_inc(x_210); lean_dec(x_21); -x_210 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_210, 0, x_208); -lean_ctor_set(x_210, 1, x_209); -return x_210; +x_212 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_212, 0, x_210); +lean_ctor_set(x_212, 1, x_211); +return x_212; } } } @@ -16911,1086 +17004,1142 @@ x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4(size_t 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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4(lean_object* x_1, size_t 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, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -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; -x_14 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__12; -x_15 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_14, x_11, x_12, x_13); -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +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; +x_15 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__12; +x_16 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_15, x_12, x_13, x_14); +x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); -lean_dec(x_15); -x_18 = lean_box_usize(x_1); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = lean_box_usize(x_2); +lean_inc(x_4); lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_16); -x_19 = lean_alloc_closure((void*)(l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__3___boxed), 15, 4); -lean_closure_set(x_19, 0, x_16); -lean_closure_set(x_19, 1, x_18); -lean_closure_set(x_19, 2, x_2); -lean_closure_set(x_19, 3, x_3); +lean_inc(x_1); +lean_inc(x_17); +x_20 = lean_alloc_closure((void*)(l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__3___boxed), 16, 5); +lean_closure_set(x_20, 0, x_17); +lean_closure_set(x_20, 1, x_1); +lean_closure_set(x_20, 2, x_19); +lean_closure_set(x_20, 3, x_3); +lean_closure_set(x_20, 4, x_4); +lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_7); +lean_inc(x_5); lean_inc(x_4); -lean_inc(x_3); -x_20 = l_Lean_Server_RpcEncoding_foldWithConstructors___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__33___rarg(x_3, x_4, x_19, x_5, x_7, x_8, x_9, x_10, x_11, x_12, x_17); -if (lean_obj_tag(x_20) == 0) +x_21 = l_Lean_Server_RpcEncoding_foldWithConstructors___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__33___rarg(x_4, x_5, x_20, x_6, x_8, x_9, x_10, x_11, x_12, x_13, x_18); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; size_t x_24; lean_object* x_25; -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_20, 1); +lean_object* x_22; lean_object* x_23; lean_object* x_24; size_t x_25; lean_object* x_26; +x_22 = lean_ctor_get(x_21, 0); lean_inc(x_22); -lean_dec(x_20); -x_23 = lean_array_get_size(x_4); -x_24 = lean_usize_of_nat(x_23); -lean_dec(x_23); -x_25 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__20(x_24, x_1, x_4, x_7, x_8, x_9, x_10, x_11, x_12, x_22); -lean_dec(x_10); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = lean_array_get_size(x_5); +x_25 = lean_usize_of_nat(x_24); +lean_dec(x_24); +x_26 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___spec__20(x_25, x_2, x_5, x_8, x_9, x_10, x_11, x_12, x_13, x_23); +lean_dec(x_11); +lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); -if (lean_obj_tag(x_25) == 0) +if (lean_obj_tag(x_26) == 0) { -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_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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; size_t x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 1); +lean_object* x_27; lean_object* x_28; lean_object* x_29; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; size_t x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; +x_27 = lean_ctor_get(x_26, 0); lean_inc(x_27); -lean_dec(x_25); -lean_inc(x_11); -x_28 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_11, x_12, x_27); -x_29 = lean_ctor_get(x_28, 0); -lean_inc(x_29); -x_30 = lean_ctor_get(x_28, 1); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +lean_inc(x_12); +x_29 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_12, x_13, x_28); +x_30 = lean_ctor_get(x_29, 0); lean_inc(x_30); -lean_dec(x_28); -x_31 = lean_st_ref_get(x_12, x_30); -x_32 = lean_ctor_get(x_31, 1); -lean_inc(x_32); -lean_dec(x_31); -x_33 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__10; -x_34 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_34, 0, x_29); -lean_ctor_set(x_34, 1, x_33); -x_35 = lean_ctor_get(x_3, 0); -lean_inc(x_35); -lean_dec(x_3); -x_36 = lean_ctor_get(x_35, 0); -lean_inc(x_36); -lean_dec(x_35); -x_37 = lean_mk_syntax_ident(x_36); -x_38 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__13; -x_39 = lean_array_push(x_38, x_34); -x_40 = lean_array_push(x_39, x_37); -x_41 = lean_box(2); -x_42 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__9; -x_43 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_43, 0, x_41); -lean_ctor_set(x_43, 1, x_42); -lean_ctor_set(x_43, 2, x_40); -x_44 = l_Lean_Syntax_mkApp(x_43, x_26); -x_45 = lean_mk_syntax_ident(x_16); -x_46 = lean_ctor_get(x_21, 1); -lean_inc(x_46); -x_47 = lean_array_get_size(x_46); -x_48 = lean_usize_of_nat(x_47); -lean_dec(x_47); -x_49 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkInductiveApp___spec__1(x_48, x_1, x_46); -lean_inc(x_45); -x_50 = l_Lean_Syntax_mkApp(x_45, x_49); -lean_inc(x_11); -x_51 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_11, x_12, x_32); -x_52 = lean_ctor_get(x_51, 0); -lean_inc(x_52); -x_53 = lean_ctor_get(x_51, 1); -lean_inc(x_53); -lean_dec(x_51); -x_54 = lean_ctor_get(x_11, 10); -lean_inc(x_54); -lean_dec(x_11); -x_55 = lean_st_ref_get(x_12, x_53); +x_31 = lean_ctor_get(x_29, 1); +lean_inc(x_31); +lean_dec(x_29); +x_32 = lean_st_ref_get(x_13, x_31); +x_33 = lean_ctor_get(x_32, 1); +lean_inc(x_33); +lean_dec(x_32); +x_34 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_matchF___closed__2; +lean_inc(x_1); +x_35 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_1); +x_36 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__10; +x_37 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_37, 0, x_30); +lean_ctor_set(x_37, 1, x_36); +x_38 = lean_ctor_get(x_4, 0); +lean_inc(x_38); +lean_dec(x_4); +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +lean_dec(x_38); +x_40 = lean_mk_syntax_ident(x_39); +x_41 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__13; +x_42 = lean_array_push(x_41, x_37); +x_43 = lean_array_push(x_42, x_40); +x_44 = lean_box(2); +x_45 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__9; +x_46 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +lean_ctor_set(x_46, 2, x_43); +x_47 = l_Lean_Syntax_mkApp(x_46, x_27); +x_48 = lean_mk_syntax_ident(x_17); +x_49 = lean_ctor_get(x_22, 1); +lean_inc(x_49); +x_50 = lean_array_get_size(x_49); +x_51 = lean_usize_of_nat(x_50); +lean_dec(x_50); +x_52 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__40(x_35, x_51, x_2, x_49); +lean_dec(x_35); +lean_inc(x_48); +x_53 = l_Lean_Syntax_mkApp(x_48, x_52); +lean_inc(x_12); +x_54 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_12, x_13, x_33); +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_54, 1); +lean_inc(x_56); +lean_dec(x_54); +x_57 = lean_ctor_get(x_12, 10); +lean_inc(x_57); lean_dec(x_12); -x_56 = !lean_is_exclusive(x_55); -if (x_56 == 0) +x_58 = lean_st_ref_get(x_13, x_56); +lean_dec(x_13); +x_59 = !lean_is_exclusive(x_58); +if (x_59 == 0) { -lean_object* x_57; lean_object* x_58; lean_object* x_59; 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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; size_t x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; size_t x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; size_t x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; -x_57 = lean_ctor_get(x_55, 0); -x_58 = lean_ctor_get(x_57, 0); -lean_inc(x_58); -lean_dec(x_57); -x_59 = lean_environment_main_module(x_58); -x_60 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__14; -lean_inc(x_52); -x_61 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_61, 0, x_52); -lean_ctor_set(x_61, 1, x_60); -x_62 = lean_ctor_get(x_21, 2); -lean_inc(x_62); -lean_inc(x_2); -x_63 = l_Array_append___rarg(x_2, x_62); -x_64 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__2; -x_65 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_65, 0, x_41); -lean_ctor_set(x_65, 1, x_64); -lean_ctor_set(x_65, 2, x_63); -x_66 = lean_array_push(x_38, x_61); -x_67 = lean_array_push(x_66, x_65); -x_68 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__15; -x_69 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_69, 0, x_41); -lean_ctor_set(x_69, 1, x_68); -lean_ctor_set(x_69, 2, x_67); -lean_inc(x_2); -x_70 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_70, 0, x_41); -lean_ctor_set(x_70, 1, x_64); -lean_ctor_set(x_70, 2, x_2); -x_71 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__42; -lean_inc(x_70); -x_72 = lean_array_push(x_71, x_70); -lean_inc(x_70); -x_73 = lean_array_push(x_72, x_70); -lean_inc(x_70); -x_74 = lean_array_push(x_73, x_70); -lean_inc(x_70); -x_75 = lean_array_push(x_74, x_70); -lean_inc(x_70); -x_76 = lean_array_push(x_75, x_70); -lean_inc(x_70); -x_77 = lean_array_push(x_76, x_70); -x_78 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__39; -x_79 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_79, 0, x_41); -lean_ctor_set(x_79, 1, x_78); -lean_ctor_set(x_79, 2, x_77); -x_80 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__1; -lean_inc(x_52); -x_81 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_81, 0, x_52); -lean_ctor_set(x_81, 1, x_80); -lean_inc(x_70); -x_82 = lean_array_push(x_38, x_70); -lean_inc(x_70); -lean_inc(x_82); -x_83 = lean_array_push(x_82, x_70); -x_84 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__56; -x_85 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_85, 0, x_41); -lean_ctor_set(x_85, 1, x_84); -lean_ctor_set(x_85, 2, x_83); -x_86 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__25; -lean_inc(x_52); -x_87 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_87, 0, x_52); -lean_ctor_set(x_87, 1, x_86); -x_88 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__25; -lean_inc(x_87); -x_89 = lean_array_push(x_88, x_87); -x_90 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_90, 0, x_41); -lean_ctor_set(x_90, 1, x_64); -lean_ctor_set(x_90, 2, x_89); -x_91 = lean_ctor_get(x_21, 3); -lean_inc(x_91); -x_92 = lean_array_get_size(x_91); -x_93 = lean_usize_of_nat(x_92); -lean_dec(x_92); -x_94 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_93, x_1, x_91); -lean_inc(x_2); -x_95 = l_Array_append___rarg(x_2, x_94); -x_96 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_96, 0, x_41); -lean_ctor_set(x_96, 1, x_64); -lean_ctor_set(x_96, 2, x_95); -x_97 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__30; -lean_inc(x_52); -x_98 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_98, 0, x_52); -lean_ctor_set(x_98, 1, x_97); -x_99 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__36; -lean_inc(x_70); -x_100 = lean_array_push(x_99, x_70); -x_101 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__32; -x_102 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_102, 0, x_41); -lean_ctor_set(x_102, 1, x_101); -lean_ctor_set(x_102, 2, x_100); -x_103 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__39; -lean_inc(x_52); -x_104 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_104, 0, x_52); -lean_ctor_set(x_104, 1, x_103); -x_105 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__43; -lean_inc(x_70); -x_106 = lean_array_push(x_105, x_70); -x_107 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_107, 0, x_41); -lean_ctor_set(x_107, 1, x_101); -lean_ctor_set(x_107, 2, x_106); -x_108 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__33; -x_109 = lean_array_push(x_108, x_102); -x_110 = lean_array_push(x_109, x_104); -x_111 = lean_array_push(x_110, x_107); -x_112 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_112, 0, x_41); -lean_ctor_set(x_112, 1, x_64); -lean_ctor_set(x_112, 2, x_111); -x_113 = lean_array_push(x_38, x_98); -x_114 = lean_array_push(x_113, x_112); -x_115 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_115, 0, x_41); -lean_ctor_set(x_115, 1, x_64); -lean_ctor_set(x_115, 2, x_114); -x_116 = lean_array_push(x_88, x_115); -x_117 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__29; -x_118 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_118, 0, x_41); -lean_ctor_set(x_118, 1, x_117); -lean_ctor_set(x_118, 2, x_116); -x_119 = lean_array_push(x_71, x_81); -x_120 = lean_array_push(x_119, x_45); -x_121 = lean_array_push(x_120, x_85); -x_122 = lean_array_push(x_121, x_90); -x_123 = lean_array_push(x_122, x_96); -x_124 = lean_array_push(x_123, x_118); -x_125 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__2; -x_126 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_126, 0, x_41); -lean_ctor_set(x_126, 1, x_125); -lean_ctor_set(x_126, 2, x_124); -x_127 = lean_array_push(x_38, x_79); -lean_inc(x_127); -x_128 = lean_array_push(x_127, x_126); -x_129 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__37; -x_130 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_130, 0, x_41); -lean_ctor_set(x_130, 1, x_129); -lean_ctor_set(x_130, 2, x_128); -lean_inc(x_70); -x_131 = lean_array_push(x_88, x_70); -x_132 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__124; +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; 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; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; size_t x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; size_t x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; size_t x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; +x_60 = lean_ctor_get(x_58, 0); +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +lean_dec(x_60); +x_62 = lean_environment_main_module(x_61); +x_63 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__14; +lean_inc(x_55); +x_64 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_64, 0, x_55); +lean_ctor_set(x_64, 1, x_63); +x_65 = lean_ctor_get(x_22, 2); +lean_inc(x_65); +lean_inc(x_3); +x_66 = l_Array_append___rarg(x_3, x_65); +x_67 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__2; +x_68 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_68, 0, x_44); +lean_ctor_set(x_68, 1, x_67); +lean_ctor_set(x_68, 2, x_66); +x_69 = lean_array_push(x_41, x_64); +x_70 = lean_array_push(x_69, x_68); +x_71 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__15; +x_72 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_72, 0, x_44); +lean_ctor_set(x_72, 1, x_71); +lean_ctor_set(x_72, 2, x_70); +lean_inc(x_3); +x_73 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_73, 0, x_44); +lean_ctor_set(x_73, 1, x_67); +lean_ctor_set(x_73, 2, x_3); +x_74 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__42; +lean_inc(x_73); +x_75 = lean_array_push(x_74, x_73); +lean_inc(x_73); +x_76 = lean_array_push(x_75, x_73); +lean_inc(x_73); +x_77 = lean_array_push(x_76, x_73); +lean_inc(x_73); +x_78 = lean_array_push(x_77, x_73); +lean_inc(x_73); +x_79 = lean_array_push(x_78, x_73); +lean_inc(x_73); +x_80 = lean_array_push(x_79, x_73); +x_81 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__39; +x_82 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_82, 0, x_44); +lean_ctor_set(x_82, 1, x_81); +lean_ctor_set(x_82, 2, x_80); +x_83 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__1; +lean_inc(x_55); +x_84 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_84, 0, x_55); +lean_ctor_set(x_84, 1, x_83); +x_85 = lean_array_push(x_41, x_48); +lean_inc(x_73); +x_86 = lean_array_push(x_85, x_73); +x_87 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__50; +x_88 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_88, 0, x_44); +lean_ctor_set(x_88, 1, x_87); +lean_ctor_set(x_88, 2, x_86); +lean_inc(x_73); +x_89 = lean_array_push(x_41, x_73); +lean_inc(x_73); +lean_inc(x_89); +x_90 = lean_array_push(x_89, x_73); +x_91 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__56; +x_92 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_92, 0, x_44); +lean_ctor_set(x_92, 1, x_91); +lean_ctor_set(x_92, 2, x_90); +x_93 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__25; +lean_inc(x_55); +x_94 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_94, 0, x_55); +lean_ctor_set(x_94, 1, x_93); +x_95 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__25; +lean_inc(x_94); +x_96 = lean_array_push(x_95, x_94); +x_97 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_97, 0, x_44); +lean_ctor_set(x_97, 1, x_67); +lean_ctor_set(x_97, 2, x_96); +x_98 = lean_ctor_get(x_22, 3); +lean_inc(x_98); +x_99 = lean_array_get_size(x_98); +x_100 = lean_usize_of_nat(x_99); +lean_dec(x_99); +x_101 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41(x_100, x_2, x_98); +lean_inc(x_3); +x_102 = l_Array_append___rarg(x_3, x_101); +x_103 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_103, 0, x_44); +lean_ctor_set(x_103, 1, x_67); +lean_ctor_set(x_103, 2, x_102); +x_104 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__30; +lean_inc(x_55); +x_105 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_105, 0, x_55); +lean_ctor_set(x_105, 1, x_104); +x_106 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__36; +lean_inc(x_73); +x_107 = lean_array_push(x_106, x_73); +x_108 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__32; +x_109 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_109, 0, x_44); +lean_ctor_set(x_109, 1, x_108); +lean_ctor_set(x_109, 2, x_107); +x_110 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__39; +lean_inc(x_55); +x_111 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_111, 0, x_55); +lean_ctor_set(x_111, 1, x_110); +x_112 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__43; +lean_inc(x_73); +x_113 = lean_array_push(x_112, x_73); +x_114 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_114, 0, x_44); +lean_ctor_set(x_114, 1, x_108); +lean_ctor_set(x_114, 2, x_113); +x_115 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__33; +x_116 = lean_array_push(x_115, x_109); +x_117 = lean_array_push(x_116, x_111); +x_118 = lean_array_push(x_117, x_114); +x_119 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_119, 0, x_44); +lean_ctor_set(x_119, 1, x_67); +lean_ctor_set(x_119, 2, x_118); +x_120 = lean_array_push(x_41, x_105); +x_121 = lean_array_push(x_120, x_119); +x_122 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_122, 0, x_44); +lean_ctor_set(x_122, 1, x_67); +lean_ctor_set(x_122, 2, x_121); +x_123 = lean_array_push(x_95, x_122); +x_124 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__29; +x_125 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_125, 0, x_44); +lean_ctor_set(x_125, 1, x_124); +lean_ctor_set(x_125, 2, x_123); +x_126 = lean_array_push(x_74, x_84); +x_127 = lean_array_push(x_126, x_88); +x_128 = lean_array_push(x_127, x_92); +x_129 = lean_array_push(x_128, x_97); +x_130 = lean_array_push(x_129, x_103); +x_131 = lean_array_push(x_130, x_125); +x_132 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__2; x_133 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_133, 0, x_41); +lean_ctor_set(x_133, 0, x_44); lean_ctor_set(x_133, 1, x_132); lean_ctor_set(x_133, 2, x_131); -x_134 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__205; -lean_inc(x_52); -x_135 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_135, 0, x_52); -lean_ctor_set(x_135, 1, x_134); -x_136 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__26; -lean_inc(x_52); -x_137 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_137, 0, x_52); +x_134 = lean_array_push(x_41, x_82); +lean_inc(x_134); +x_135 = lean_array_push(x_134, x_133); +x_136 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__37; +x_137 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_137, 0, x_44); lean_ctor_set(x_137, 1, x_136); -x_138 = lean_array_push(x_38, x_44); -x_139 = lean_array_push(x_138, x_50); +lean_ctor_set(x_137, 2, x_135); +lean_inc(x_73); +x_138 = lean_array_push(x_95, x_73); +x_139 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__124; x_140 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_140, 0, x_41); -lean_ctor_set(x_140, 1, x_64); -lean_ctor_set(x_140, 2, x_139); -x_141 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__50; -x_142 = lean_array_push(x_141, x_140); -x_143 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__61; -x_144 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_144, 0, x_41); +lean_ctor_set(x_140, 0, x_44); +lean_ctor_set(x_140, 1, x_139); +lean_ctor_set(x_140, 2, x_138); +x_141 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__205; +lean_inc(x_55); +x_142 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_142, 0, x_55); +lean_ctor_set(x_142, 1, x_141); +x_143 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__26; +lean_inc(x_55); +x_144 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_144, 0, x_55); lean_ctor_set(x_144, 1, x_143); -lean_ctor_set(x_144, 2, x_142); -x_145 = lean_array_push(x_38, x_137); -x_146 = lean_array_push(x_145, x_144); -x_147 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__95; -x_148 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_148, 0, x_41); -lean_ctor_set(x_148, 1, x_147); -lean_ctor_set(x_148, 2, x_146); -lean_inc(x_82); -x_149 = lean_array_push(x_82, x_148); -x_150 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__143; +x_145 = lean_array_push(x_41, x_47); +x_146 = lean_array_push(x_145, x_53); +x_147 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_147, 0, x_44); +lean_ctor_set(x_147, 1, x_67); +lean_ctor_set(x_147, 2, x_146); +x_148 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__50; +x_149 = lean_array_push(x_148, x_147); +x_150 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__61; x_151 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_151, 0, x_41); +lean_ctor_set(x_151, 0, x_44); lean_ctor_set(x_151, 1, x_150); lean_ctor_set(x_151, 2, x_149); -x_152 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__224; -lean_inc(x_54); -lean_inc(x_59); -x_153 = l_Lean_addMacroScope(x_59, x_152, x_54); -x_154 = lean_box(0); -x_155 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__223; -x_156 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__61; -lean_inc(x_52); -x_157 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_157, 0, x_52); -lean_ctor_set(x_157, 1, x_155); -lean_ctor_set(x_157, 2, x_153); -lean_ctor_set(x_157, 3, x_156); -x_158 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__109; -lean_inc(x_52); -x_159 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_159, 0, x_52); -lean_ctor_set(x_159, 1, x_158); -x_160 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__3; -lean_inc(x_52); -x_161 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_161, 0, x_52); -lean_ctor_set(x_161, 1, x_160); -x_162 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__10; -lean_inc(x_54); -lean_inc(x_59); -x_163 = l_Lean_addMacroScope(x_59, x_162, x_54); -x_164 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__9; -lean_inc(x_52); +x_152 = lean_array_push(x_41, x_144); +x_153 = lean_array_push(x_152, x_151); +x_154 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__95; +x_155 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_155, 0, x_44); +lean_ctor_set(x_155, 1, x_154); +lean_ctor_set(x_155, 2, x_153); +lean_inc(x_89); +x_156 = lean_array_push(x_89, x_155); +x_157 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__143; +x_158 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_158, 0, x_44); +lean_ctor_set(x_158, 1, x_157); +lean_ctor_set(x_158, 2, x_156); +x_159 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__224; +lean_inc(x_57); +lean_inc(x_62); +x_160 = l_Lean_addMacroScope(x_62, x_159, x_57); +x_161 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__225; +lean_inc(x_1); +x_162 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_162, 0, x_161); +lean_ctor_set(x_162, 1, x_1); +lean_inc(x_1); +x_163 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_163, 0, x_162); +lean_ctor_set(x_163, 1, x_1); +x_164 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__223; +lean_inc(x_55); x_165 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_165, 0, x_52); +lean_ctor_set(x_165, 0, x_55); lean_ctor_set(x_165, 1, x_164); -lean_ctor_set(x_165, 2, x_163); -lean_ctor_set(x_165, 3, x_154); -lean_inc(x_165); -x_166 = lean_array_push(x_88, x_165); -x_167 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_167, 0, x_41); -lean_ctor_set(x_167, 1, x_64); -lean_ctor_set(x_167, 2, x_166); -x_168 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__3___closed__6; -lean_inc(x_52); +lean_ctor_set(x_165, 2, x_160); +lean_ctor_set(x_165, 3, x_163); +x_166 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__109; +lean_inc(x_55); +x_167 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_167, 0, x_55); +lean_ctor_set(x_167, 1, x_166); +x_168 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__3; +lean_inc(x_55); x_169 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_169, 0, x_52); +lean_ctor_set(x_169, 0, x_55); lean_ctor_set(x_169, 1, x_168); -x_170 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__11; -lean_inc(x_52); -x_171 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_171, 0, x_52); -lean_ctor_set(x_171, 1, x_170); -x_172 = lean_array_push(x_82, x_165); -x_173 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__14; -x_174 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_174, 0, x_41); -lean_ctor_set(x_174, 1, x_173); -lean_ctor_set(x_174, 2, x_172); -x_175 = lean_array_push(x_88, x_174); -x_176 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_176, 0, x_41); -lean_ctor_set(x_176, 1, x_64); -lean_ctor_set(x_176, 2, x_175); -x_177 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__15; -lean_inc(x_52); -x_178 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_178, 0, x_52); -lean_ctor_set(x_178, 1, x_177); -x_179 = lean_ctor_get(x_21, 4); -lean_inc(x_179); -x_180 = lean_array_get_size(x_179); -x_181 = lean_usize_of_nat(x_180); -lean_dec(x_180); -x_182 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_181, x_1, x_179); -lean_inc(x_2); -x_183 = l_Array_append___rarg(x_2, x_182); +x_170 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__10; +lean_inc(x_57); +lean_inc(x_62); +x_171 = l_Lean_addMacroScope(x_62, x_170, x_57); +x_172 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__9; +lean_inc(x_1); +lean_inc(x_55); +x_173 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_173, 0, x_55); +lean_ctor_set(x_173, 1, x_172); +lean_ctor_set(x_173, 2, x_171); +lean_ctor_set(x_173, 3, x_1); +lean_inc(x_173); +x_174 = lean_array_push(x_95, x_173); +x_175 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_175, 0, x_44); +lean_ctor_set(x_175, 1, x_67); +lean_ctor_set(x_175, 2, x_174); +x_176 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__3___closed__6; +lean_inc(x_55); +x_177 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_177, 0, x_55); +lean_ctor_set(x_177, 1, x_176); +x_178 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__11; +lean_inc(x_55); +x_179 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_179, 0, x_55); +lean_ctor_set(x_179, 1, x_178); +x_180 = lean_array_push(x_89, x_173); +x_181 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__14; +x_182 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_182, 0, x_44); +lean_ctor_set(x_182, 1, x_181); +lean_ctor_set(x_182, 2, x_180); +x_183 = lean_array_push(x_95, x_182); x_184 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_184, 0, x_41); -lean_ctor_set(x_184, 1, x_64); +lean_ctor_set(x_184, 0, x_44); +lean_ctor_set(x_184, 1, x_67); lean_ctor_set(x_184, 2, x_183); -x_185 = lean_array_push(x_88, x_184); -x_186 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__17; -x_187 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_187, 0, x_41); -lean_ctor_set(x_187, 1, x_186); -lean_ctor_set(x_187, 2, x_185); -x_188 = lean_array_push(x_71, x_171); -lean_inc(x_70); -x_189 = lean_array_push(x_188, x_70); -lean_inc(x_70); -x_190 = lean_array_push(x_189, x_70); -x_191 = lean_array_push(x_190, x_176); -x_192 = lean_array_push(x_191, x_178); -lean_inc(x_192); -x_193 = lean_array_push(x_192, x_187); -x_194 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__12; +x_185 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__15; +lean_inc(x_55); +x_186 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_186, 0, x_55); +lean_ctor_set(x_186, 1, x_185); +x_187 = lean_ctor_get(x_22, 4); +lean_inc(x_187); +x_188 = lean_array_get_size(x_187); +x_189 = lean_usize_of_nat(x_188); +lean_dec(x_188); +x_190 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__42(x_189, x_2, x_187); +lean_inc(x_3); +x_191 = l_Array_append___rarg(x_3, x_190); +x_192 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_192, 0, x_44); +lean_ctor_set(x_192, 1, x_67); +lean_ctor_set(x_192, 2, x_191); +x_193 = lean_array_push(x_95, x_192); +x_194 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__17; x_195 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_195, 0, x_41); +lean_ctor_set(x_195, 0, x_44); lean_ctor_set(x_195, 1, x_194); lean_ctor_set(x_195, 2, x_193); -x_196 = lean_array_push(x_108, x_167); -x_197 = lean_array_push(x_196, x_169); -lean_inc(x_197); -x_198 = lean_array_push(x_197, x_195); -x_199 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__6; -x_200 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_200, 0, x_41); -lean_ctor_set(x_200, 1, x_199); -lean_ctor_set(x_200, 2, x_198); -x_201 = lean_array_push(x_38, x_161); -lean_inc(x_201); -x_202 = lean_array_push(x_201, x_200); -x_203 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__4; -x_204 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_204, 0, x_41); -lean_ctor_set(x_204, 1, x_203); -lean_ctor_set(x_204, 2, x_202); -x_205 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__93; -x_206 = lean_array_push(x_205, x_157); -lean_inc(x_70); -x_207 = lean_array_push(x_206, x_70); -lean_inc(x_70); -x_208 = lean_array_push(x_207, x_70); -lean_inc(x_159); -x_209 = lean_array_push(x_208, x_159); -x_210 = lean_array_push(x_209, x_204); -x_211 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__59; +x_196 = lean_array_push(x_74, x_179); +lean_inc(x_73); +x_197 = lean_array_push(x_196, x_73); +lean_inc(x_73); +x_198 = lean_array_push(x_197, x_73); +x_199 = lean_array_push(x_198, x_184); +x_200 = lean_array_push(x_199, x_186); +lean_inc(x_200); +x_201 = lean_array_push(x_200, x_195); +x_202 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__12; +x_203 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_203, 0, x_44); +lean_ctor_set(x_203, 1, x_202); +lean_ctor_set(x_203, 2, x_201); +x_204 = lean_array_push(x_115, x_175); +x_205 = lean_array_push(x_204, x_177); +lean_inc(x_205); +x_206 = lean_array_push(x_205, x_203); +x_207 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__6; +x_208 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_208, 0, x_44); +lean_ctor_set(x_208, 1, x_207); +lean_ctor_set(x_208, 2, x_206); +x_209 = lean_array_push(x_41, x_169); +lean_inc(x_209); +x_210 = lean_array_push(x_209, x_208); +x_211 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__4; x_212 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_212, 0, x_41); +lean_ctor_set(x_212, 0, x_44); lean_ctor_set(x_212, 1, x_211); lean_ctor_set(x_212, 2, x_210); -x_213 = lean_array_push(x_88, x_212); -x_214 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__57; -x_215 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_215, 0, x_41); -lean_ctor_set(x_215, 1, x_214); -lean_ctor_set(x_215, 2, x_213); -x_216 = lean_array_push(x_88, x_215); -x_217 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__55; -x_218 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_218, 0, x_41); -lean_ctor_set(x_218, 1, x_217); -lean_ctor_set(x_218, 2, x_216); -x_219 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__231; -x_220 = l_Lean_addMacroScope(x_59, x_219, x_54); -x_221 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__230; -x_222 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__234; -x_223 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_223, 0, x_52); -lean_ctor_set(x_223, 1, x_221); -lean_ctor_set(x_223, 2, x_220); -lean_ctor_set(x_223, 3, x_222); -x_224 = lean_ctor_get(x_21, 5); -lean_inc(x_224); -lean_dec(x_21); -x_225 = lean_array_get_size(x_224); -x_226 = lean_usize_of_nat(x_225); -lean_dec(x_225); -x_227 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_226, x_1, x_224); -x_228 = l_Array_append___rarg(x_2, x_227); -x_229 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_229, 0, x_41); -lean_ctor_set(x_229, 1, x_64); -lean_ctor_set(x_229, 2, x_228); -x_230 = lean_array_push(x_88, x_229); -x_231 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_231, 0, x_41); -lean_ctor_set(x_231, 1, x_186); -lean_ctor_set(x_231, 2, x_230); -x_232 = lean_array_push(x_192, x_231); -x_233 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_233, 0, x_41); -lean_ctor_set(x_233, 1, x_194); -lean_ctor_set(x_233, 2, x_232); -x_234 = lean_array_push(x_197, x_233); -x_235 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_235, 0, x_41); -lean_ctor_set(x_235, 1, x_199); -lean_ctor_set(x_235, 2, x_234); -x_236 = lean_array_push(x_201, x_235); -x_237 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_237, 0, x_41); -lean_ctor_set(x_237, 1, x_203); -lean_ctor_set(x_237, 2, x_236); -x_238 = lean_array_push(x_205, x_223); -lean_inc(x_70); -x_239 = lean_array_push(x_238, x_70); -lean_inc(x_70); -x_240 = lean_array_push(x_239, x_70); -x_241 = lean_array_push(x_240, x_159); -x_242 = lean_array_push(x_241, x_237); +x_213 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__93; +x_214 = lean_array_push(x_213, x_165); +lean_inc(x_73); +x_215 = lean_array_push(x_214, x_73); +lean_inc(x_73); +x_216 = lean_array_push(x_215, x_73); +lean_inc(x_167); +x_217 = lean_array_push(x_216, x_167); +x_218 = lean_array_push(x_217, x_212); +x_219 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__59; +x_220 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_220, 0, x_44); +lean_ctor_set(x_220, 1, x_219); +lean_ctor_set(x_220, 2, x_218); +x_221 = lean_array_push(x_95, x_220); +x_222 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__57; +x_223 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_223, 0, x_44); +lean_ctor_set(x_223, 1, x_222); +lean_ctor_set(x_223, 2, x_221); +x_224 = lean_array_push(x_95, x_223); +x_225 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__55; +x_226 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_226, 0, x_44); +lean_ctor_set(x_226, 1, x_225); +lean_ctor_set(x_226, 2, x_224); +x_227 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__231; +x_228 = l_Lean_addMacroScope(x_62, x_227, x_57); +x_229 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__232; +lean_inc(x_1); +x_230 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_230, 0, x_229); +lean_ctor_set(x_230, 1, x_1); +x_231 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_231, 0, x_230); +lean_ctor_set(x_231, 1, x_1); +x_232 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__230; +x_233 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_233, 0, x_55); +lean_ctor_set(x_233, 1, x_232); +lean_ctor_set(x_233, 2, x_228); +lean_ctor_set(x_233, 3, x_231); +x_234 = lean_ctor_get(x_22, 5); +lean_inc(x_234); +lean_dec(x_22); +x_235 = lean_array_get_size(x_234); +x_236 = lean_usize_of_nat(x_235); +lean_dec(x_235); +x_237 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__42(x_236, x_2, x_234); +x_238 = l_Array_append___rarg(x_3, x_237); +x_239 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_239, 0, x_44); +lean_ctor_set(x_239, 1, x_67); +lean_ctor_set(x_239, 2, x_238); +x_240 = lean_array_push(x_95, x_239); +x_241 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_241, 0, x_44); +lean_ctor_set(x_241, 1, x_194); +lean_ctor_set(x_241, 2, x_240); +x_242 = lean_array_push(x_200, x_241); x_243 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_243, 0, x_41); -lean_ctor_set(x_243, 1, x_211); +lean_ctor_set(x_243, 0, x_44); +lean_ctor_set(x_243, 1, x_202); lean_ctor_set(x_243, 2, x_242); -x_244 = lean_array_push(x_88, x_243); +x_244 = lean_array_push(x_205, x_243); x_245 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_245, 0, x_41); -lean_ctor_set(x_245, 1, x_214); +lean_ctor_set(x_245, 0, x_44); +lean_ctor_set(x_245, 1, x_207); lean_ctor_set(x_245, 2, x_244); -x_246 = lean_array_push(x_88, x_245); +x_246 = lean_array_push(x_209, x_245); x_247 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_247, 0, x_41); -lean_ctor_set(x_247, 1, x_217); +lean_ctor_set(x_247, 0, x_44); +lean_ctor_set(x_247, 1, x_211); lean_ctor_set(x_247, 2, x_246); -x_248 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__35; -x_249 = lean_array_push(x_248, x_218); -lean_inc(x_70); -x_250 = lean_array_push(x_249, x_70); -x_251 = lean_array_push(x_250, x_247); -lean_inc(x_70); -x_252 = lean_array_push(x_251, x_70); +x_248 = lean_array_push(x_213, x_233); +lean_inc(x_73); +x_249 = lean_array_push(x_248, x_73); +lean_inc(x_73); +x_250 = lean_array_push(x_249, x_73); +x_251 = lean_array_push(x_250, x_167); +x_252 = lean_array_push(x_251, x_247); x_253 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_253, 0, x_41); -lean_ctor_set(x_253, 1, x_64); +lean_ctor_set(x_253, 0, x_44); +lean_ctor_set(x_253, 1, x_219); lean_ctor_set(x_253, 2, x_252); -x_254 = lean_array_push(x_108, x_87); -x_255 = lean_array_push(x_254, x_253); -lean_inc(x_70); -x_256 = lean_array_push(x_255, x_70); -x_257 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__53; -x_258 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_258, 0, x_41); -lean_ctor_set(x_258, 1, x_257); -lean_ctor_set(x_258, 2, x_256); -x_259 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__238; -x_260 = lean_array_push(x_259, x_133); -x_261 = lean_array_push(x_260, x_135); -lean_inc(x_70); -x_262 = lean_array_push(x_261, x_70); -lean_inc(x_70); -x_263 = lean_array_push(x_262, x_70); -x_264 = lean_array_push(x_263, x_151); -x_265 = lean_array_push(x_264, x_258); -lean_inc(x_70); -x_266 = lean_array_push(x_265, x_70); -x_267 = lean_array_push(x_266, x_70); -x_268 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__206; -x_269 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_269, 0, x_41); -lean_ctor_set(x_269, 1, x_268); -lean_ctor_set(x_269, 2, x_267); -x_270 = lean_array_push(x_127, x_269); -x_271 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_271, 0, x_41); -lean_ctor_set(x_271, 1, x_129); -lean_ctor_set(x_271, 2, x_270); -x_272 = lean_array_push(x_108, x_69); -x_273 = lean_array_push(x_272, x_130); -x_274 = lean_array_push(x_273, x_271); -x_275 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_275, 0, x_41); -lean_ctor_set(x_275, 1, x_64); -lean_ctor_set(x_275, 2, x_274); -lean_ctor_set(x_55, 0, x_275); -return x_55; +x_254 = lean_array_push(x_95, x_253); +x_255 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_255, 0, x_44); +lean_ctor_set(x_255, 1, x_222); +lean_ctor_set(x_255, 2, x_254); +x_256 = lean_array_push(x_95, x_255); +x_257 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_257, 0, x_44); +lean_ctor_set(x_257, 1, x_225); +lean_ctor_set(x_257, 2, x_256); +x_258 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__35; +x_259 = lean_array_push(x_258, x_226); +lean_inc(x_73); +x_260 = lean_array_push(x_259, x_73); +x_261 = lean_array_push(x_260, x_257); +lean_inc(x_73); +x_262 = lean_array_push(x_261, x_73); +x_263 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_263, 0, x_44); +lean_ctor_set(x_263, 1, x_67); +lean_ctor_set(x_263, 2, x_262); +x_264 = lean_array_push(x_115, x_94); +x_265 = lean_array_push(x_264, x_263); +lean_inc(x_73); +x_266 = lean_array_push(x_265, x_73); +x_267 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__53; +x_268 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_268, 0, x_44); +lean_ctor_set(x_268, 1, x_267); +lean_ctor_set(x_268, 2, x_266); +x_269 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__238; +x_270 = lean_array_push(x_269, x_140); +x_271 = lean_array_push(x_270, x_142); +lean_inc(x_73); +x_272 = lean_array_push(x_271, x_73); +lean_inc(x_73); +x_273 = lean_array_push(x_272, x_73); +x_274 = lean_array_push(x_273, x_158); +x_275 = lean_array_push(x_274, x_268); +lean_inc(x_73); +x_276 = lean_array_push(x_275, x_73); +x_277 = lean_array_push(x_276, x_73); +x_278 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__206; +x_279 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_279, 0, x_44); +lean_ctor_set(x_279, 1, x_278); +lean_ctor_set(x_279, 2, x_277); +x_280 = lean_array_push(x_134, x_279); +x_281 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_281, 0, x_44); +lean_ctor_set(x_281, 1, x_136); +lean_ctor_set(x_281, 2, x_280); +x_282 = lean_array_push(x_115, x_72); +x_283 = lean_array_push(x_282, x_137); +x_284 = lean_array_push(x_283, x_281); +x_285 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_285, 0, x_44); +lean_ctor_set(x_285, 1, x_67); +lean_ctor_set(x_285, 2, x_284); +lean_ctor_set(x_58, 0, x_285); +return x_58; } else { -lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; size_t x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; size_t x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; size_t x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; -x_276 = lean_ctor_get(x_55, 0); -x_277 = lean_ctor_get(x_55, 1); -lean_inc(x_277); -lean_inc(x_276); -lean_dec(x_55); -x_278 = lean_ctor_get(x_276, 0); -lean_inc(x_278); -lean_dec(x_276); -x_279 = lean_environment_main_module(x_278); -x_280 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__14; -lean_inc(x_52); -x_281 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_281, 0, x_52); -lean_ctor_set(x_281, 1, x_280); -x_282 = lean_ctor_get(x_21, 2); -lean_inc(x_282); -lean_inc(x_2); -x_283 = l_Array_append___rarg(x_2, x_282); -x_284 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__2; -x_285 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_285, 0, x_41); -lean_ctor_set(x_285, 1, x_284); -lean_ctor_set(x_285, 2, x_283); -x_286 = lean_array_push(x_38, x_281); -x_287 = lean_array_push(x_286, x_285); -x_288 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__15; -x_289 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_289, 0, x_41); -lean_ctor_set(x_289, 1, x_288); -lean_ctor_set(x_289, 2, x_287); -lean_inc(x_2); -x_290 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_290, 0, x_41); -lean_ctor_set(x_290, 1, x_284); -lean_ctor_set(x_290, 2, x_2); -x_291 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__42; -lean_inc(x_290); -x_292 = lean_array_push(x_291, x_290); -lean_inc(x_290); -x_293 = lean_array_push(x_292, x_290); -lean_inc(x_290); -x_294 = lean_array_push(x_293, x_290); -lean_inc(x_290); -x_295 = lean_array_push(x_294, x_290); -lean_inc(x_290); -x_296 = lean_array_push(x_295, x_290); -lean_inc(x_290); -x_297 = lean_array_push(x_296, x_290); -x_298 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__39; +lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; size_t x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; size_t x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; size_t x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; +x_286 = lean_ctor_get(x_58, 0); +x_287 = lean_ctor_get(x_58, 1); +lean_inc(x_287); +lean_inc(x_286); +lean_dec(x_58); +x_288 = lean_ctor_get(x_286, 0); +lean_inc(x_288); +lean_dec(x_286); +x_289 = lean_environment_main_module(x_288); +x_290 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__14; +lean_inc(x_55); +x_291 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_291, 0, x_55); +lean_ctor_set(x_291, 1, x_290); +x_292 = lean_ctor_get(x_22, 2); +lean_inc(x_292); +lean_inc(x_3); +x_293 = l_Array_append___rarg(x_3, x_292); +x_294 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__2; +x_295 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_295, 0, x_44); +lean_ctor_set(x_295, 1, x_294); +lean_ctor_set(x_295, 2, x_293); +x_296 = lean_array_push(x_41, x_291); +x_297 = lean_array_push(x_296, x_295); +x_298 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__15; x_299 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_299, 0, x_41); +lean_ctor_set(x_299, 0, x_44); lean_ctor_set(x_299, 1, x_298); lean_ctor_set(x_299, 2, x_297); -x_300 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__1; -lean_inc(x_52); -x_301 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_301, 0, x_52); -lean_ctor_set(x_301, 1, x_300); -lean_inc(x_290); -x_302 = lean_array_push(x_38, x_290); -lean_inc(x_290); -lean_inc(x_302); -x_303 = lean_array_push(x_302, x_290); -x_304 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__56; -x_305 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_305, 0, x_41); -lean_ctor_set(x_305, 1, x_304); -lean_ctor_set(x_305, 2, x_303); -x_306 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__25; -lean_inc(x_52); -x_307 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_307, 0, x_52); -lean_ctor_set(x_307, 1, x_306); -x_308 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__25; -lean_inc(x_307); -x_309 = lean_array_push(x_308, x_307); -x_310 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_310, 0, x_41); -lean_ctor_set(x_310, 1, x_284); -lean_ctor_set(x_310, 2, x_309); -x_311 = lean_ctor_get(x_21, 3); -lean_inc(x_311); -x_312 = lean_array_get_size(x_311); -x_313 = lean_usize_of_nat(x_312); -lean_dec(x_312); -x_314 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_313, x_1, x_311); -lean_inc(x_2); -x_315 = l_Array_append___rarg(x_2, x_314); -x_316 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_316, 0, x_41); -lean_ctor_set(x_316, 1, x_284); -lean_ctor_set(x_316, 2, x_315); -x_317 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__30; -lean_inc(x_52); -x_318 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_318, 0, x_52); -lean_ctor_set(x_318, 1, x_317); -x_319 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__36; -lean_inc(x_290); -x_320 = lean_array_push(x_319, x_290); -x_321 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__32; -x_322 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_322, 0, x_41); -lean_ctor_set(x_322, 1, x_321); -lean_ctor_set(x_322, 2, x_320); -x_323 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__39; -lean_inc(x_52); -x_324 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_324, 0, x_52); -lean_ctor_set(x_324, 1, x_323); -x_325 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__43; -lean_inc(x_290); -x_326 = lean_array_push(x_325, x_290); -x_327 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_327, 0, x_41); -lean_ctor_set(x_327, 1, x_321); -lean_ctor_set(x_327, 2, x_326); -x_328 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__33; -x_329 = lean_array_push(x_328, x_322); -x_330 = lean_array_push(x_329, x_324); -x_331 = lean_array_push(x_330, x_327); -x_332 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_332, 0, x_41); -lean_ctor_set(x_332, 1, x_284); -lean_ctor_set(x_332, 2, x_331); -x_333 = lean_array_push(x_38, x_318); -x_334 = lean_array_push(x_333, x_332); -x_335 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_335, 0, x_41); -lean_ctor_set(x_335, 1, x_284); -lean_ctor_set(x_335, 2, x_334); -x_336 = lean_array_push(x_308, x_335); -x_337 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__29; -x_338 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_338, 0, x_41); +lean_inc(x_3); +x_300 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_300, 0, x_44); +lean_ctor_set(x_300, 1, x_294); +lean_ctor_set(x_300, 2, x_3); +x_301 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__42; +lean_inc(x_300); +x_302 = lean_array_push(x_301, x_300); +lean_inc(x_300); +x_303 = lean_array_push(x_302, x_300); +lean_inc(x_300); +x_304 = lean_array_push(x_303, x_300); +lean_inc(x_300); +x_305 = lean_array_push(x_304, x_300); +lean_inc(x_300); +x_306 = lean_array_push(x_305, x_300); +lean_inc(x_300); +x_307 = lean_array_push(x_306, x_300); +x_308 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__39; +x_309 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_309, 0, x_44); +lean_ctor_set(x_309, 1, x_308); +lean_ctor_set(x_309, 2, x_307); +x_310 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__1; +lean_inc(x_55); +x_311 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_311, 0, x_55); +lean_ctor_set(x_311, 1, x_310); +x_312 = lean_array_push(x_41, x_48); +lean_inc(x_300); +x_313 = lean_array_push(x_312, x_300); +x_314 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__50; +x_315 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_315, 0, x_44); +lean_ctor_set(x_315, 1, x_314); +lean_ctor_set(x_315, 2, x_313); +lean_inc(x_300); +x_316 = lean_array_push(x_41, x_300); +lean_inc(x_300); +lean_inc(x_316); +x_317 = lean_array_push(x_316, x_300); +x_318 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__56; +x_319 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_319, 0, x_44); +lean_ctor_set(x_319, 1, x_318); +lean_ctor_set(x_319, 2, x_317); +x_320 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__25; +lean_inc(x_55); +x_321 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_321, 0, x_55); +lean_ctor_set(x_321, 1, x_320); +x_322 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__25; +lean_inc(x_321); +x_323 = lean_array_push(x_322, x_321); +x_324 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_324, 0, x_44); +lean_ctor_set(x_324, 1, x_294); +lean_ctor_set(x_324, 2, x_323); +x_325 = lean_ctor_get(x_22, 3); +lean_inc(x_325); +x_326 = lean_array_get_size(x_325); +x_327 = lean_usize_of_nat(x_326); +lean_dec(x_326); +x_328 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41(x_327, x_2, x_325); +lean_inc(x_3); +x_329 = l_Array_append___rarg(x_3, x_328); +x_330 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_330, 0, x_44); +lean_ctor_set(x_330, 1, x_294); +lean_ctor_set(x_330, 2, x_329); +x_331 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__30; +lean_inc(x_55); +x_332 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_332, 0, x_55); +lean_ctor_set(x_332, 1, x_331); +x_333 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__36; +lean_inc(x_300); +x_334 = lean_array_push(x_333, x_300); +x_335 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__32; +x_336 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_336, 0, x_44); +lean_ctor_set(x_336, 1, x_335); +lean_ctor_set(x_336, 2, x_334); +x_337 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__39; +lean_inc(x_55); +x_338 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_338, 0, x_55); lean_ctor_set(x_338, 1, x_337); -lean_ctor_set(x_338, 2, x_336); -x_339 = lean_array_push(x_291, x_301); -x_340 = lean_array_push(x_339, x_45); -x_341 = lean_array_push(x_340, x_305); -x_342 = lean_array_push(x_341, x_310); -x_343 = lean_array_push(x_342, x_316); +x_339 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__43; +lean_inc(x_300); +x_340 = lean_array_push(x_339, x_300); +x_341 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_341, 0, x_44); +lean_ctor_set(x_341, 1, x_335); +lean_ctor_set(x_341, 2, x_340); +x_342 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__33; +x_343 = lean_array_push(x_342, x_336); x_344 = lean_array_push(x_343, x_338); -x_345 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__2; +x_345 = lean_array_push(x_344, x_341); x_346 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_346, 0, x_41); -lean_ctor_set(x_346, 1, x_345); -lean_ctor_set(x_346, 2, x_344); -x_347 = lean_array_push(x_38, x_299); -lean_inc(x_347); +lean_ctor_set(x_346, 0, x_44); +lean_ctor_set(x_346, 1, x_294); +lean_ctor_set(x_346, 2, x_345); +x_347 = lean_array_push(x_41, x_332); x_348 = lean_array_push(x_347, x_346); -x_349 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__37; -x_350 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_350, 0, x_41); -lean_ctor_set(x_350, 1, x_349); -lean_ctor_set(x_350, 2, x_348); -lean_inc(x_290); -x_351 = lean_array_push(x_308, x_290); -x_352 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__124; -x_353 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_353, 0, x_41); -lean_ctor_set(x_353, 1, x_352); -lean_ctor_set(x_353, 2, x_351); -x_354 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__205; -lean_inc(x_52); -x_355 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_355, 0, x_52); -lean_ctor_set(x_355, 1, x_354); -x_356 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__26; -lean_inc(x_52); -x_357 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_357, 0, x_52); -lean_ctor_set(x_357, 1, x_356); -x_358 = lean_array_push(x_38, x_44); -x_359 = lean_array_push(x_358, x_50); +x_349 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_349, 0, x_44); +lean_ctor_set(x_349, 1, x_294); +lean_ctor_set(x_349, 2, x_348); +x_350 = lean_array_push(x_322, x_349); +x_351 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__29; +x_352 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_352, 0, x_44); +lean_ctor_set(x_352, 1, x_351); +lean_ctor_set(x_352, 2, x_350); +x_353 = lean_array_push(x_301, x_311); +x_354 = lean_array_push(x_353, x_315); +x_355 = lean_array_push(x_354, x_319); +x_356 = lean_array_push(x_355, x_324); +x_357 = lean_array_push(x_356, x_330); +x_358 = lean_array_push(x_357, x_352); +x_359 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__2; x_360 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_360, 0, x_41); -lean_ctor_set(x_360, 1, x_284); -lean_ctor_set(x_360, 2, x_359); -x_361 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__50; +lean_ctor_set(x_360, 0, x_44); +lean_ctor_set(x_360, 1, x_359); +lean_ctor_set(x_360, 2, x_358); +x_361 = lean_array_push(x_41, x_309); +lean_inc(x_361); x_362 = lean_array_push(x_361, x_360); -x_363 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__61; +x_363 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__37; x_364 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_364, 0, x_41); +lean_ctor_set(x_364, 0, x_44); lean_ctor_set(x_364, 1, x_363); lean_ctor_set(x_364, 2, x_362); -x_365 = lean_array_push(x_38, x_357); -x_366 = lean_array_push(x_365, x_364); -x_367 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__95; -x_368 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_368, 0, x_41); -lean_ctor_set(x_368, 1, x_367); -lean_ctor_set(x_368, 2, x_366); -lean_inc(x_302); -x_369 = lean_array_push(x_302, x_368); -x_370 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__143; -x_371 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_371, 0, x_41); +lean_inc(x_300); +x_365 = lean_array_push(x_322, x_300); +x_366 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__124; +x_367 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_367, 0, x_44); +lean_ctor_set(x_367, 1, x_366); +lean_ctor_set(x_367, 2, x_365); +x_368 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__205; +lean_inc(x_55); +x_369 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_369, 0, x_55); +lean_ctor_set(x_369, 1, x_368); +x_370 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__26; +lean_inc(x_55); +x_371 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_371, 0, x_55); lean_ctor_set(x_371, 1, x_370); -lean_ctor_set(x_371, 2, x_369); -x_372 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__224; -lean_inc(x_54); -lean_inc(x_279); -x_373 = l_Lean_addMacroScope(x_279, x_372, x_54); -x_374 = lean_box(0); -x_375 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__223; -x_376 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__61; -lean_inc(x_52); -x_377 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_377, 0, x_52); -lean_ctor_set(x_377, 1, x_375); -lean_ctor_set(x_377, 2, x_373); -lean_ctor_set(x_377, 3, x_376); -x_378 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__109; -lean_inc(x_52); -x_379 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_379, 0, x_52); -lean_ctor_set(x_379, 1, x_378); -x_380 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__3; -lean_inc(x_52); -x_381 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_381, 0, x_52); -lean_ctor_set(x_381, 1, x_380); -x_382 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__10; -lean_inc(x_54); -lean_inc(x_279); -x_383 = l_Lean_addMacroScope(x_279, x_382, x_54); -x_384 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__9; -lean_inc(x_52); -x_385 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_385, 0, x_52); +x_372 = lean_array_push(x_41, x_47); +x_373 = lean_array_push(x_372, x_53); +x_374 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_374, 0, x_44); +lean_ctor_set(x_374, 1, x_294); +lean_ctor_set(x_374, 2, x_373); +x_375 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__50; +x_376 = lean_array_push(x_375, x_374); +x_377 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__61; +x_378 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_378, 0, x_44); +lean_ctor_set(x_378, 1, x_377); +lean_ctor_set(x_378, 2, x_376); +x_379 = lean_array_push(x_41, x_371); +x_380 = lean_array_push(x_379, x_378); +x_381 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__95; +x_382 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_382, 0, x_44); +lean_ctor_set(x_382, 1, x_381); +lean_ctor_set(x_382, 2, x_380); +lean_inc(x_316); +x_383 = lean_array_push(x_316, x_382); +x_384 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__143; +x_385 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_385, 0, x_44); lean_ctor_set(x_385, 1, x_384); lean_ctor_set(x_385, 2, x_383); -lean_ctor_set(x_385, 3, x_374); -lean_inc(x_385); -x_386 = lean_array_push(x_308, x_385); -x_387 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_387, 0, x_41); -lean_ctor_set(x_387, 1, x_284); -lean_ctor_set(x_387, 2, x_386); -x_388 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__3___closed__6; -lean_inc(x_52); -x_389 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_389, 0, x_52); -lean_ctor_set(x_389, 1, x_388); -x_390 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__11; -lean_inc(x_52); -x_391 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_391, 0, x_52); -lean_ctor_set(x_391, 1, x_390); -x_392 = lean_array_push(x_302, x_385); -x_393 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__14; -x_394 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_394, 0, x_41); +x_386 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__224; +lean_inc(x_57); +lean_inc(x_289); +x_387 = l_Lean_addMacroScope(x_289, x_386, x_57); +x_388 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__225; +lean_inc(x_1); +x_389 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_389, 0, x_388); +lean_ctor_set(x_389, 1, x_1); +lean_inc(x_1); +x_390 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_390, 0, x_389); +lean_ctor_set(x_390, 1, x_1); +x_391 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__223; +lean_inc(x_55); +x_392 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_392, 0, x_55); +lean_ctor_set(x_392, 1, x_391); +lean_ctor_set(x_392, 2, x_387); +lean_ctor_set(x_392, 3, x_390); +x_393 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__109; +lean_inc(x_55); +x_394 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_394, 0, x_55); lean_ctor_set(x_394, 1, x_393); -lean_ctor_set(x_394, 2, x_392); -x_395 = lean_array_push(x_308, x_394); -x_396 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_396, 0, x_41); -lean_ctor_set(x_396, 1, x_284); -lean_ctor_set(x_396, 2, x_395); -x_397 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__15; -lean_inc(x_52); -x_398 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_398, 0, x_52); -lean_ctor_set(x_398, 1, x_397); -x_399 = lean_ctor_get(x_21, 4); -lean_inc(x_399); -x_400 = lean_array_get_size(x_399); -x_401 = lean_usize_of_nat(x_400); -lean_dec(x_400); -x_402 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_401, x_1, x_399); -lean_inc(x_2); -x_403 = l_Array_append___rarg(x_2, x_402); -x_404 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_404, 0, x_41); -lean_ctor_set(x_404, 1, x_284); -lean_ctor_set(x_404, 2, x_403); -x_405 = lean_array_push(x_308, x_404); -x_406 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__17; -x_407 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_407, 0, x_41); -lean_ctor_set(x_407, 1, x_406); -lean_ctor_set(x_407, 2, x_405); -x_408 = lean_array_push(x_291, x_391); -lean_inc(x_290); -x_409 = lean_array_push(x_408, x_290); -lean_inc(x_290); -x_410 = lean_array_push(x_409, x_290); -x_411 = lean_array_push(x_410, x_396); -x_412 = lean_array_push(x_411, x_398); -lean_inc(x_412); -x_413 = lean_array_push(x_412, x_407); -x_414 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__12; -x_415 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_415, 0, x_41); -lean_ctor_set(x_415, 1, x_414); -lean_ctor_set(x_415, 2, x_413); -x_416 = lean_array_push(x_328, x_387); -x_417 = lean_array_push(x_416, x_389); -lean_inc(x_417); -x_418 = lean_array_push(x_417, x_415); -x_419 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__6; -x_420 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_420, 0, x_41); -lean_ctor_set(x_420, 1, x_419); -lean_ctor_set(x_420, 2, x_418); -x_421 = lean_array_push(x_38, x_381); -lean_inc(x_421); -x_422 = lean_array_push(x_421, x_420); -x_423 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__4; -x_424 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_424, 0, x_41); -lean_ctor_set(x_424, 1, x_423); -lean_ctor_set(x_424, 2, x_422); -x_425 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__93; -x_426 = lean_array_push(x_425, x_377); -lean_inc(x_290); -x_427 = lean_array_push(x_426, x_290); -lean_inc(x_290); -x_428 = lean_array_push(x_427, x_290); -lean_inc(x_379); -x_429 = lean_array_push(x_428, x_379); -x_430 = lean_array_push(x_429, x_424); -x_431 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__59; -x_432 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_432, 0, x_41); -lean_ctor_set(x_432, 1, x_431); -lean_ctor_set(x_432, 2, x_430); -x_433 = lean_array_push(x_308, x_432); -x_434 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__57; +x_395 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__3; +lean_inc(x_55); +x_396 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_396, 0, x_55); +lean_ctor_set(x_396, 1, x_395); +x_397 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__10; +lean_inc(x_57); +lean_inc(x_289); +x_398 = l_Lean_addMacroScope(x_289, x_397, x_57); +x_399 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__9; +lean_inc(x_1); +lean_inc(x_55); +x_400 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_400, 0, x_55); +lean_ctor_set(x_400, 1, x_399); +lean_ctor_set(x_400, 2, x_398); +lean_ctor_set(x_400, 3, x_1); +lean_inc(x_400); +x_401 = lean_array_push(x_322, x_400); +x_402 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_402, 0, x_44); +lean_ctor_set(x_402, 1, x_294); +lean_ctor_set(x_402, 2, x_401); +x_403 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__3___closed__6; +lean_inc(x_55); +x_404 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_404, 0, x_55); +lean_ctor_set(x_404, 1, x_403); +x_405 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__11; +lean_inc(x_55); +x_406 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_406, 0, x_55); +lean_ctor_set(x_406, 1, x_405); +x_407 = lean_array_push(x_316, x_400); +x_408 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__14; +x_409 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_409, 0, x_44); +lean_ctor_set(x_409, 1, x_408); +lean_ctor_set(x_409, 2, x_407); +x_410 = lean_array_push(x_322, x_409); +x_411 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_411, 0, x_44); +lean_ctor_set(x_411, 1, x_294); +lean_ctor_set(x_411, 2, x_410); +x_412 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__15; +lean_inc(x_55); +x_413 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_413, 0, x_55); +lean_ctor_set(x_413, 1, x_412); +x_414 = lean_ctor_get(x_22, 4); +lean_inc(x_414); +x_415 = lean_array_get_size(x_414); +x_416 = lean_usize_of_nat(x_415); +lean_dec(x_415); +x_417 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__42(x_416, x_2, x_414); +lean_inc(x_3); +x_418 = l_Array_append___rarg(x_3, x_417); +x_419 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_419, 0, x_44); +lean_ctor_set(x_419, 1, x_294); +lean_ctor_set(x_419, 2, x_418); +x_420 = lean_array_push(x_322, x_419); +x_421 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__17; +x_422 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_422, 0, x_44); +lean_ctor_set(x_422, 1, x_421); +lean_ctor_set(x_422, 2, x_420); +x_423 = lean_array_push(x_301, x_406); +lean_inc(x_300); +x_424 = lean_array_push(x_423, x_300); +lean_inc(x_300); +x_425 = lean_array_push(x_424, x_300); +x_426 = lean_array_push(x_425, x_411); +x_427 = lean_array_push(x_426, x_413); +lean_inc(x_427); +x_428 = lean_array_push(x_427, x_422); +x_429 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__12; +x_430 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_430, 0, x_44); +lean_ctor_set(x_430, 1, x_429); +lean_ctor_set(x_430, 2, x_428); +x_431 = lean_array_push(x_342, x_402); +x_432 = lean_array_push(x_431, x_404); +lean_inc(x_432); +x_433 = lean_array_push(x_432, x_430); +x_434 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__6; x_435 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_435, 0, x_41); +lean_ctor_set(x_435, 0, x_44); lean_ctor_set(x_435, 1, x_434); lean_ctor_set(x_435, 2, x_433); -x_436 = lean_array_push(x_308, x_435); -x_437 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__55; -x_438 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_438, 0, x_41); -lean_ctor_set(x_438, 1, x_437); -lean_ctor_set(x_438, 2, x_436); -x_439 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__231; -x_440 = l_Lean_addMacroScope(x_279, x_439, x_54); -x_441 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__230; -x_442 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__234; -x_443 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_443, 0, x_52); -lean_ctor_set(x_443, 1, x_441); -lean_ctor_set(x_443, 2, x_440); -lean_ctor_set(x_443, 3, x_442); -x_444 = lean_ctor_get(x_21, 5); -lean_inc(x_444); -lean_dec(x_21); -x_445 = lean_array_get_size(x_444); -x_446 = lean_usize_of_nat(x_445); -lean_dec(x_445); -x_447 = l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(x_446, x_1, x_444); -x_448 = l_Array_append___rarg(x_2, x_447); -x_449 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_449, 0, x_41); -lean_ctor_set(x_449, 1, x_284); -lean_ctor_set(x_449, 2, x_448); -x_450 = lean_array_push(x_308, x_449); -x_451 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_451, 0, x_41); -lean_ctor_set(x_451, 1, x_406); -lean_ctor_set(x_451, 2, x_450); -x_452 = lean_array_push(x_412, x_451); +x_436 = lean_array_push(x_41, x_396); +lean_inc(x_436); +x_437 = lean_array_push(x_436, x_435); +x_438 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___closed__4; +x_439 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_439, 0, x_44); +lean_ctor_set(x_439, 1, x_438); +lean_ctor_set(x_439, 2, x_437); +x_440 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__93; +x_441 = lean_array_push(x_440, x_392); +lean_inc(x_300); +x_442 = lean_array_push(x_441, x_300); +lean_inc(x_300); +x_443 = lean_array_push(x_442, x_300); +lean_inc(x_394); +x_444 = lean_array_push(x_443, x_394); +x_445 = lean_array_push(x_444, x_439); +x_446 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__59; +x_447 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_447, 0, x_44); +lean_ctor_set(x_447, 1, x_446); +lean_ctor_set(x_447, 2, x_445); +x_448 = lean_array_push(x_322, x_447); +x_449 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__57; +x_450 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_450, 0, x_44); +lean_ctor_set(x_450, 1, x_449); +lean_ctor_set(x_450, 2, x_448); +x_451 = lean_array_push(x_322, x_450); +x_452 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__55; x_453 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_453, 0, x_41); -lean_ctor_set(x_453, 1, x_414); -lean_ctor_set(x_453, 2, x_452); -x_454 = lean_array_push(x_417, x_453); -x_455 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_455, 0, x_41); -lean_ctor_set(x_455, 1, x_419); -lean_ctor_set(x_455, 2, x_454); -x_456 = lean_array_push(x_421, x_455); -x_457 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_457, 0, x_41); -lean_ctor_set(x_457, 1, x_423); -lean_ctor_set(x_457, 2, x_456); -x_458 = lean_array_push(x_425, x_443); -lean_inc(x_290); -x_459 = lean_array_push(x_458, x_290); -lean_inc(x_290); -x_460 = lean_array_push(x_459, x_290); -x_461 = lean_array_push(x_460, x_379); -x_462 = lean_array_push(x_461, x_457); -x_463 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_463, 0, x_41); -lean_ctor_set(x_463, 1, x_431); -lean_ctor_set(x_463, 2, x_462); -x_464 = lean_array_push(x_308, x_463); -x_465 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_465, 0, x_41); -lean_ctor_set(x_465, 1, x_434); -lean_ctor_set(x_465, 2, x_464); -x_466 = lean_array_push(x_308, x_465); -x_467 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_467, 0, x_41); -lean_ctor_set(x_467, 1, x_437); -lean_ctor_set(x_467, 2, x_466); -x_468 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__35; -x_469 = lean_array_push(x_468, x_438); -lean_inc(x_290); -x_470 = lean_array_push(x_469, x_290); -x_471 = lean_array_push(x_470, x_467); -lean_inc(x_290); -x_472 = lean_array_push(x_471, x_290); -x_473 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_473, 0, x_41); -lean_ctor_set(x_473, 1, x_284); -lean_ctor_set(x_473, 2, x_472); -x_474 = lean_array_push(x_328, x_307); -x_475 = lean_array_push(x_474, x_473); -lean_inc(x_290); -x_476 = lean_array_push(x_475, x_290); -x_477 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__53; -x_478 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_478, 0, x_41); -lean_ctor_set(x_478, 1, x_477); -lean_ctor_set(x_478, 2, x_476); -x_479 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__238; -x_480 = lean_array_push(x_479, x_353); -x_481 = lean_array_push(x_480, x_355); -lean_inc(x_290); -x_482 = lean_array_push(x_481, x_290); -lean_inc(x_290); -x_483 = lean_array_push(x_482, x_290); -x_484 = lean_array_push(x_483, x_371); -x_485 = lean_array_push(x_484, x_478); -lean_inc(x_290); -x_486 = lean_array_push(x_485, x_290); -x_487 = lean_array_push(x_486, x_290); -x_488 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__206; -x_489 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_489, 0, x_41); -lean_ctor_set(x_489, 1, x_488); -lean_ctor_set(x_489, 2, x_487); -x_490 = lean_array_push(x_347, x_489); -x_491 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_491, 0, x_41); -lean_ctor_set(x_491, 1, x_349); -lean_ctor_set(x_491, 2, x_490); -x_492 = lean_array_push(x_328, x_289); -x_493 = lean_array_push(x_492, x_350); -x_494 = lean_array_push(x_493, x_491); +lean_ctor_set(x_453, 0, x_44); +lean_ctor_set(x_453, 1, x_452); +lean_ctor_set(x_453, 2, x_451); +x_454 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__231; +x_455 = l_Lean_addMacroScope(x_289, x_454, x_57); +x_456 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__232; +lean_inc(x_1); +x_457 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_457, 0, x_456); +lean_ctor_set(x_457, 1, x_1); +x_458 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_458, 0, x_457); +lean_ctor_set(x_458, 1, x_1); +x_459 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__230; +x_460 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_460, 0, x_55); +lean_ctor_set(x_460, 1, x_459); +lean_ctor_set(x_460, 2, x_455); +lean_ctor_set(x_460, 3, x_458); +x_461 = lean_ctor_get(x_22, 5); +lean_inc(x_461); +lean_dec(x_22); +x_462 = lean_array_get_size(x_461); +x_463 = lean_usize_of_nat(x_462); +lean_dec(x_462); +x_464 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__42(x_463, x_2, x_461); +x_465 = l_Array_append___rarg(x_3, x_464); +x_466 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_466, 0, x_44); +lean_ctor_set(x_466, 1, x_294); +lean_ctor_set(x_466, 2, x_465); +x_467 = lean_array_push(x_322, x_466); +x_468 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_468, 0, x_44); +lean_ctor_set(x_468, 1, x_421); +lean_ctor_set(x_468, 2, x_467); +x_469 = lean_array_push(x_427, x_468); +x_470 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_470, 0, x_44); +lean_ctor_set(x_470, 1, x_429); +lean_ctor_set(x_470, 2, x_469); +x_471 = lean_array_push(x_432, x_470); +x_472 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_472, 0, x_44); +lean_ctor_set(x_472, 1, x_434); +lean_ctor_set(x_472, 2, x_471); +x_473 = lean_array_push(x_436, x_472); +x_474 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_474, 0, x_44); +lean_ctor_set(x_474, 1, x_438); +lean_ctor_set(x_474, 2, x_473); +x_475 = lean_array_push(x_440, x_460); +lean_inc(x_300); +x_476 = lean_array_push(x_475, x_300); +lean_inc(x_300); +x_477 = lean_array_push(x_476, x_300); +x_478 = lean_array_push(x_477, x_394); +x_479 = lean_array_push(x_478, x_474); +x_480 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_480, 0, x_44); +lean_ctor_set(x_480, 1, x_446); +lean_ctor_set(x_480, 2, x_479); +x_481 = lean_array_push(x_322, x_480); +x_482 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_482, 0, x_44); +lean_ctor_set(x_482, 1, x_449); +lean_ctor_set(x_482, 2, x_481); +x_483 = lean_array_push(x_322, x_482); +x_484 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_484, 0, x_44); +lean_ctor_set(x_484, 1, x_452); +lean_ctor_set(x_484, 2, x_483); +x_485 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__35; +x_486 = lean_array_push(x_485, x_453); +lean_inc(x_300); +x_487 = lean_array_push(x_486, x_300); +x_488 = lean_array_push(x_487, x_484); +lean_inc(x_300); +x_489 = lean_array_push(x_488, x_300); +x_490 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_490, 0, x_44); +lean_ctor_set(x_490, 1, x_294); +lean_ctor_set(x_490, 2, x_489); +x_491 = lean_array_push(x_342, x_321); +x_492 = lean_array_push(x_491, x_490); +lean_inc(x_300); +x_493 = lean_array_push(x_492, x_300); +x_494 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveStructureInstance___lambda__1___closed__53; x_495 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_495, 0, x_41); -lean_ctor_set(x_495, 1, x_284); -lean_ctor_set(x_495, 2, x_494); -x_496 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_496, 0, x_495); -lean_ctor_set(x_496, 1, x_277); -return x_496; -} -} -else -{ -uint8_t x_497; -lean_dec(x_21); -lean_dec(x_16); +lean_ctor_set(x_495, 0, x_44); +lean_ctor_set(x_495, 1, x_494); +lean_ctor_set(x_495, 2, x_493); +x_496 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__238; +x_497 = lean_array_push(x_496, x_367); +x_498 = lean_array_push(x_497, x_369); +lean_inc(x_300); +x_499 = lean_array_push(x_498, x_300); +lean_inc(x_300); +x_500 = lean_array_push(x_499, x_300); +x_501 = lean_array_push(x_500, x_385); +x_502 = lean_array_push(x_501, x_495); +lean_inc(x_300); +x_503 = lean_array_push(x_502, x_300); +x_504 = lean_array_push(x_503, x_300); +x_505 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__206; +x_506 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_506, 0, x_44); +lean_ctor_set(x_506, 1, x_505); +lean_ctor_set(x_506, 2, x_504); +x_507 = lean_array_push(x_361, x_506); +x_508 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_508, 0, x_44); +lean_ctor_set(x_508, 1, x_363); +lean_ctor_set(x_508, 2, x_507); +x_509 = lean_array_push(x_342, x_299); +x_510 = lean_array_push(x_509, x_364); +x_511 = lean_array_push(x_510, x_508); +x_512 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_512, 0, x_44); +lean_ctor_set(x_512, 1, x_294); +lean_ctor_set(x_512, 2, x_511); +x_513 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_513, 0, x_512); +lean_ctor_set(x_513, 1, x_287); +return x_513; +} +} +else +{ +uint8_t x_514; +lean_dec(x_22); +lean_dec(x_17); +lean_dec(x_13); lean_dec(x_12); -lean_dec(x_11); +lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -x_497 = !lean_is_exclusive(x_25); -if (x_497 == 0) +lean_dec(x_1); +x_514 = !lean_is_exclusive(x_26); +if (x_514 == 0) { -return x_25; +return x_26; } else { -lean_object* x_498; lean_object* x_499; lean_object* x_500; -x_498 = lean_ctor_get(x_25, 0); -x_499 = lean_ctor_get(x_25, 1); -lean_inc(x_499); -lean_inc(x_498); -lean_dec(x_25); -x_500 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_500, 0, x_498); -lean_ctor_set(x_500, 1, x_499); -return x_500; +lean_object* x_515; lean_object* x_516; lean_object* x_517; +x_515 = lean_ctor_get(x_26, 0); +x_516 = lean_ctor_get(x_26, 1); +lean_inc(x_516); +lean_inc(x_515); +lean_dec(x_26); +x_517 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_517, 0, x_515); +lean_ctor_set(x_517, 1, x_516); +return x_517; } } } else { -uint8_t x_501; -lean_dec(x_16); +uint8_t x_518; +lean_dec(x_17); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -x_501 = !lean_is_exclusive(x_20); -if (x_501 == 0) +lean_dec(x_1); +x_518 = !lean_is_exclusive(x_21); +if (x_518 == 0) { -return x_20; +return x_21; } else { -lean_object* x_502; lean_object* x_503; lean_object* x_504; -x_502 = lean_ctor_get(x_20, 0); -x_503 = lean_ctor_get(x_20, 1); -lean_inc(x_503); -lean_inc(x_502); -lean_dec(x_20); -x_504 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_504, 0, x_502); -lean_ctor_set(x_504, 1, x_503); -return x_504; +lean_object* x_519; lean_object* x_520; lean_object* x_521; +x_519 = lean_ctor_get(x_21, 0); +x_520 = lean_ctor_get(x_21, 1); +lean_inc(x_520); +lean_inc(x_519); +lean_dec(x_21); +x_521 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_521, 0, x_519); +lean_ctor_set(x_521, 1, x_520); +return x_521; } } } @@ -18012,11 +18161,11 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__5(size_t 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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__5(lean_object* x_1, size_t 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, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { -uint8_t x_15; lean_object* x_16; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; -x_32 = lean_st_ref_get(x_13, x_14); +uint8_t x_16; lean_object* x_17; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_32 = lean_st_ref_get(x_14, x_15); x_33 = lean_ctor_get(x_32, 0); lean_inc(x_33); x_34 = lean_ctor_get(x_33, 3); @@ -18031,8 +18180,8 @@ x_36 = lean_ctor_get(x_32, 1); lean_inc(x_36); lean_dec(x_32); x_37 = 0; -x_15 = x_37; -x_16 = x_36; +x_16 = x_37; +x_17 = x_36; goto block_31; } else @@ -18041,8 +18190,8 @@ lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint x_38 = lean_ctor_get(x_32, 1); lean_inc(x_38); lean_dec(x_32); -lean_inc(x_6); -x_39 = l___private_Lean_Util_Trace_0__Lean_checkTraceOptionM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTerm___spec__2(x_6, x_8, x_9, x_10, x_11, x_12, x_13, x_38); +lean_inc(x_7); +x_39 = l___private_Lean_Util_Trace_0__Lean_checkTraceOptionM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTerm___spec__2(x_7, x_9, x_10, x_11, x_12, x_13, x_14, x_38); x_40 = lean_ctor_get(x_39, 0); lean_inc(x_40); x_41 = lean_ctor_get(x_39, 1); @@ -18050,27 +18199,27 @@ lean_inc(x_41); lean_dec(x_39); x_42 = lean_unbox(x_40); lean_dec(x_40); -x_15 = x_42; -x_16 = x_41; +x_16 = x_42; +x_17 = x_41; goto block_31; } block_31: { -if (x_15 == 0) +if (x_16 == 0) { -lean_object* x_17; lean_object* x_18; +lean_object* x_18; lean_object* x_19; +lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); -x_17 = lean_box(0); -x_18 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4(x_1, x_2, x_3, x_4, x_5, x_17, x_8, x_9, x_10, x_11, x_12, x_13, x_16); -return x_18; +x_18 = lean_box(0); +x_19 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_18, x_9, x_10, x_11, x_12, x_13, x_14, x_17); +return x_19; } 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; lean_object* x_30; -x_19 = lean_array_to_list(lean_box(0), x_7); -x_20 = lean_box(0); -x_21 = l_List_mapTRAux___at_Lean_MessageData_instCoeListExprMessageData___spec__1(x_19, x_20); +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; lean_object* x_30; +x_20 = lean_array_to_list(lean_box(0), x_8); +lean_inc(x_1); +x_21 = l_List_mapTRAux___at_Lean_MessageData_instCoeListExprMessageData___spec__1(x_20, x_1); x_22 = l_Lean_MessageData_ofList(x_21); lean_dec(x_21); x_23 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__5___closed__2; @@ -18081,13 +18230,13 @@ x_25 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_ x_26 = lean_alloc_ctor(10, 2, 0); lean_ctor_set(x_26, 0, x_24); lean_ctor_set(x_26, 1, x_25); -x_27 = l_Lean_addTrace___at___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTerm___spec__1(x_6, x_26, x_8, x_9, x_10, x_11, x_12, x_13, x_16); +x_27 = l_Lean_addTrace___at___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTerm___spec__1(x_7, x_26, x_9, x_10, x_11, x_12, x_13, x_14, x_17); x_28 = lean_ctor_get(x_27, 0); lean_inc(x_28); x_29 = lean_ctor_get(x_27, 1); lean_inc(x_29); lean_dec(x_27); -x_30 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4(x_1, x_2, x_3, x_4, x_5, x_28, x_8, x_9, x_10, x_11, x_12, x_13, x_29); +x_30 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_28, x_9, x_10, x_11, x_12, x_13, x_14, x_29); lean_dec(x_28); return x_30; } @@ -18125,35 +18274,37 @@ lean_inc(x_4); x_16 = l_Lean_Server_RpcEncoding_foldWithConstructors___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__33___rarg(x_4, x_5, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); if (lean_obj_tag(x_16) == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; size_t x_21; size_t 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_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; size_t x_22; size_t 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_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); x_18 = lean_ctor_get(x_16, 1); lean_inc(x_18); lean_dec(x_16); -x_19 = lean_ctor_get(x_17, 1); -lean_inc(x_19); -x_20 = lean_array_get_size(x_19); -x_21 = lean_usize_of_nat(x_20); -lean_dec(x_20); -x_22 = 0; -x_23 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__35(x_21, x_22, x_19); -x_24 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__11; -x_25 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__6___boxed__const__1; -x_26 = lean_alloc_closure((void*)(l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__5___boxed), 14, 6); -lean_closure_set(x_26, 0, x_25); -lean_closure_set(x_26, 1, x_24); -lean_closure_set(x_26, 2, x_4); -lean_closure_set(x_26, 3, x_5); -lean_closure_set(x_26, 4, x_17); -lean_closure_set(x_26, 5, x_3); -x_27 = l_Lean_instInhabitedSyntax; -x_28 = l_Lean_Meta_withLocalDecls___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__40___rarg(x_27, x_23, x_26, x_7, x_8, x_9, x_10, x_11, x_12, x_18); -return x_28; +x_19 = lean_box(0); +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +x_21 = lean_array_get_size(x_20); +x_22 = lean_usize_of_nat(x_21); +lean_dec(x_21); +x_23 = 0; +x_24 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__35(x_22, x_23, x_20); +x_25 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__11; +x_26 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__6___boxed__const__1; +x_27 = lean_alloc_closure((void*)(l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__5___boxed), 15, 7); +lean_closure_set(x_27, 0, x_19); +lean_closure_set(x_27, 1, x_26); +lean_closure_set(x_27, 2, x_25); +lean_closure_set(x_27, 3, x_4); +lean_closure_set(x_27, 4, x_5); +lean_closure_set(x_27, 5, x_17); +lean_closure_set(x_27, 6, x_3); +x_28 = lean_box(0); +x_29 = l_Lean_Meta_withLocalDecls___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__43___rarg(x_28, x_24, x_27, x_7, x_8, x_9, x_10, x_11, x_12, x_18); +return x_29; } else { -uint8_t x_29; +uint8_t x_30; lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -18163,23 +18314,23 @@ lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_29 = !lean_is_exclusive(x_16); -if (x_29 == 0) +x_30 = !lean_is_exclusive(x_16); +if (x_30 == 0) { return x_16; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_16, 0); -x_31 = lean_ctor_get(x_16, 1); +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_16, 0); +x_32 = lean_ctor_get(x_16, 1); +lean_inc(x_32); lean_inc(x_31); -lean_inc(x_30); lean_dec(x_16); -x_32 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -return x_32; +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; } } } @@ -18685,45 +18836,84 @@ lean_dec(x_4); return x_13; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__38___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_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) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__38___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_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, lean_object* x_16) { _start: { -size_t x_16; size_t x_17; lean_object* x_18; -x_16 = lean_unbox_usize(x_6); -lean_dec(x_6); +size_t x_17; size_t x_18; lean_object* x_19; x_17 = lean_unbox_usize(x_7); lean_dec(x_7); -x_18 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__38(x_1, x_2, x_3, x_4, x_5, x_16, x_17, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -lean_dec(x_14); +x_18 = lean_unbox_usize(x_8); +lean_dec(x_8); +x_19 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__38(x_1, x_2, x_3, x_4, x_5, x_6, x_17, x_18, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_15); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); -return x_18; +lean_dec(x_1); +return x_19; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__39___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_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) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__39___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_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, lean_object* x_16) { _start: { -size_t x_16; size_t x_17; lean_object* x_18; -x_16 = lean_unbox_usize(x_6); -lean_dec(x_6); +size_t x_17; size_t x_18; lean_object* x_19; x_17 = lean_unbox_usize(x_7); lean_dec(x_7); -x_18 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__39(x_1, x_2, x_3, x_4, x_5, x_16, x_17, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -lean_dec(x_14); +x_18 = lean_unbox_usize(x_8); +lean_dec(x_8); +x_19 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__39(x_1, x_2, x_3, x_4, x_5, x_6, x_17, x_18, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_15); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); -return x_18; +lean_dec(x_5); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__40___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; size_t x_6; lean_object* x_7; +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__40(x_1, x_5, x_6, x_4); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41(x_4, x_5, x_3); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__42___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__42(x_4, x_5, x_3); +return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___lambda__1___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___lambda__1___boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___lambda__1(x_1); +x_2 = l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___lambda__1(x_1); lean_dec(x_1); return x_2; } @@ -18740,36 +18930,37 @@ lean_dec(x_2); return x_13; } } -LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__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_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) { +LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__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_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, lean_object* x_16) { _start: { -size_t x_16; lean_object* x_17; -x_16 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_17 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__3(x_1, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -lean_dec(x_8); -return x_17; +size_t x_17; lean_object* x_18; +x_17 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_18 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__3(x_1, x_2, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_9); +lean_dec(x_7); +return x_18; } } -LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -size_t x_14; lean_object* x_15; -x_14 = lean_unbox_usize(x_1); -lean_dec(x_1); -x_15 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4(x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_6); -return x_15; +size_t x_15; lean_object* x_16; +x_15 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_16 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__4(x_1, x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_7); +return x_16; } } -LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__5___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__5___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_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) { _start: { -size_t x_15; lean_object* x_16; -x_15 = lean_unbox_usize(x_1); -lean_dec(x_1); -x_16 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__5(x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -return x_16; +size_t x_16; lean_object* x_17; +x_16 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_17 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__5(x_1, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_17; } } LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInstance___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { @@ -18866,7 +19057,7 @@ LEAN_EXPORT lean_object* l_Lean_getConstInfoInduct___at___private_Lean_Server_Rp lean_object* x_5; lean_inc(x_2); lean_inc(x_1); -x_5 = l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__10(x_1, x_2, x_3, x_4); +x_5 = l_Lean_getConstInfo___at_Lean_Elab_elabDeriving___spec__14(x_1, x_2, x_3, x_4); if (lean_obj_tag(x_5) == 0) { lean_object* x_6; @@ -19136,205 +19327,194 @@ return x_93; } } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInstance___spec__4(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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInstance___spec__4(lean_object* x_1, size_t x_2, size_t 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, lean_object* x_11) { _start: { -uint8_t x_13; -x_13 = lean_usize_dec_lt(x_4, x_3); -if (x_13 == 0) +uint8_t x_12; +x_12 = lean_usize_dec_lt(x_3, x_2); +if (x_12 == 0) { -lean_object* x_14; -lean_dec(x_11); +lean_object* x_13; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_1); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_5); -lean_ctor_set(x_14, 1, x_12); -return x_14; +lean_dec(x_7); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_4); +lean_ctor_set(x_13, 1, x_11); +return x_13; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_23; -x_15 = lean_array_uget(x_2, x_4); -lean_inc(x_8); -lean_inc(x_15); -x_23 = l_Lean_Meta_getFVarLocalDecl(x_15, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_23) == 0) +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_22; +x_14 = lean_array_uget(x_1, x_3); +lean_inc(x_7); +lean_inc(x_14); +x_22 = l_Lean_Meta_getFVarLocalDecl(x_14, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_22) == 0) { -lean_object* 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_24 = lean_ctor_get(x_23, 0); +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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); -lean_inc(x_25); +lean_dec(x_22); +x_25 = l_Lean_LocalDecl_userName(x_23); lean_dec(x_23); -x_26 = l_Lean_LocalDecl_userName(x_24); -lean_dec(x_24); -lean_inc(x_10); -x_27 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_10, x_11, x_25); -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); -lean_dec(x_27); -x_30 = lean_st_ref_get(x_11, x_29); -x_31 = lean_ctor_get(x_30, 1); -lean_inc(x_31); -lean_dec(x_30); -x_32 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__80; +lean_inc(x_9); +x_26 = l_Lean_MonadRef_mkInfoFromRefPos___at_Lean_Elab_Deriving_mkInductiveApp___spec__2___rarg(x_9, x_10, x_24); +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); lean_inc(x_28); -x_33 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_33, 0, x_28); -lean_ctor_set(x_33, 1, x_32); -x_34 = lean_mk_syntax_ident(x_26); -x_35 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__25; -lean_inc(x_34); -x_36 = lean_array_push(x_35, x_34); -x_37 = lean_box(2); -x_38 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__2; -x_39 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_39, 0, x_37); -lean_ctor_set(x_39, 1, x_38); -lean_ctor_set(x_39, 2, x_36); -lean_inc(x_1); -x_40 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_40, 0, x_37); -lean_ctor_set(x_40, 1, x_38); -lean_ctor_set(x_40, 2, x_1); -x_41 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__92; -x_42 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_42, 0, x_28); -lean_ctor_set(x_42, 1, x_41); -x_43 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__93; -x_44 = lean_array_push(x_43, x_33); -x_45 = lean_array_push(x_44, x_39); -lean_inc(x_40); -x_46 = lean_array_push(x_45, x_40); -lean_inc(x_40); +lean_dec(x_26); +x_29 = lean_st_ref_get(x_10, x_28); +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +lean_dec(x_29); +x_31 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__80; +lean_inc(x_27); +x_32 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_32, 0, x_27); +lean_ctor_set(x_32, 1, x_31); +x_33 = lean_mk_syntax_ident(x_25); +x_34 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__25; +lean_inc(x_33); +x_35 = lean_array_push(x_34, x_33); +x_36 = lean_box(2); +x_37 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__2; +x_38 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +lean_ctor_set(x_38, 2, x_35); +x_39 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__92; +x_40 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_40, 0, x_27); +lean_ctor_set(x_40, 1, x_39); +x_41 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__93; +x_42 = lean_array_push(x_41, x_32); +x_43 = lean_array_push(x_42, x_38); +x_44 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__12; +x_45 = lean_array_push(x_43, x_44); +x_46 = lean_array_push(x_45, x_44); x_47 = lean_array_push(x_46, x_40); -x_48 = lean_array_push(x_47, x_42); -x_49 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__79; -x_50 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_50, 0, x_37); -lean_ctor_set(x_50, 1, x_49); -lean_ctor_set(x_50, 2, x_48); -x_51 = lean_array_push(x_5, x_50); -lean_inc(x_11); +x_48 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__79; +x_49 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_49, 0, x_36); +lean_ctor_set(x_49, 1, x_48); +lean_ctor_set(x_49, 2, x_47); +x_50 = lean_array_push(x_4, x_49); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -x_52 = lean_infer_type(x_15, x_8, x_9, x_10, x_11, x_31); -if (lean_obj_tag(x_52) == 0) +lean_inc(x_7); +x_51 = lean_infer_type(x_14, x_7, x_8, x_9, x_10, x_30); +if (lean_obj_tag(x_51) == 0) { -lean_object* x_53; lean_object* x_54; uint8_t x_55; -x_53 = lean_ctor_get(x_52, 0); +lean_object* x_52; lean_object* x_53; uint8_t x_54; +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_51, 1); lean_inc(x_53); -x_54 = lean_ctor_get(x_52, 1); -lean_inc(x_54); +lean_dec(x_51); +x_54 = l_Lean_Expr_isType(x_52); lean_dec(x_52); -x_55 = l_Lean_Expr_isType(x_53); -lean_dec(x_53); -if (x_55 == 0) +if (x_54 == 0) { -lean_object* x_56; -lean_dec(x_40); -lean_dec(x_34); -x_56 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_56, 0, x_51); -x_16 = x_56; -x_17 = x_54; -goto block_22; +lean_object* x_55; +lean_dec(x_33); +x_55 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_55, 0, x_50); +x_15 = x_55; +x_16 = x_53; +goto block_21; } else { -lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_57 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__17; -x_58 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__4; -x_59 = lean_box(0); -lean_inc(x_10); -x_60 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInstance___spec__4___lambda__1(x_57, x_58, x_35, x_34, x_38, x_40, x_51, x_59, x_6, x_7, x_8, x_9, x_10, x_11, x_54); -x_61 = lean_ctor_get(x_60, 0); +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_56 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__17; +x_57 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__4; +x_58 = lean_box(0); +lean_inc(x_9); +x_59 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInstance___spec__4___lambda__1(x_56, x_57, x_34, x_33, x_37, x_44, x_50, x_58, x_5, x_6, x_7, x_8, x_9, x_10, x_53); +x_60 = lean_ctor_get(x_59, 0); +lean_inc(x_60); +x_61 = lean_ctor_get(x_59, 1); lean_inc(x_61); -x_62 = lean_ctor_get(x_60, 1); -lean_inc(x_62); -lean_dec(x_60); +lean_dec(x_59); +x_15 = x_60; x_16 = x_61; -x_17 = x_62; -goto block_22; +goto block_21; } } else { -uint8_t x_63; -lean_dec(x_51); -lean_dec(x_40); -lean_dec(x_34); -lean_dec(x_11); +uint8_t x_62; +lean_dec(x_50); +lean_dec(x_33); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_1); -x_63 = !lean_is_exclusive(x_52); -if (x_63 == 0) +lean_dec(x_7); +x_62 = !lean_is_exclusive(x_51); +if (x_62 == 0) { -return x_52; +return x_51; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_ctor_get(x_52, 0); -x_65 = lean_ctor_get(x_52, 1); -lean_inc(x_65); +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_51, 0); +x_64 = lean_ctor_get(x_51, 1); lean_inc(x_64); -lean_dec(x_52); -x_66 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_66, 0, x_64); -lean_ctor_set(x_66, 1, x_65); -return x_66; +lean_inc(x_63); +lean_dec(x_51); +x_65 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_64); +return x_65; } } } else { -uint8_t x_67; -lean_dec(x_15); -lean_dec(x_11); +uint8_t x_66; +lean_dec(x_14); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_5); -lean_dec(x_1); -x_67 = !lean_is_exclusive(x_23); -if (x_67 == 0) +lean_dec(x_7); +lean_dec(x_4); +x_66 = !lean_is_exclusive(x_22); +if (x_66 == 0) { -return x_23; +return x_22; } else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_68 = lean_ctor_get(x_23, 0); -x_69 = lean_ctor_get(x_23, 1); -lean_inc(x_69); +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_22, 0); +x_68 = lean_ctor_get(x_22, 1); lean_inc(x_68); -lean_dec(x_23); -x_70 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_69); -return x_70; +lean_inc(x_67); +lean_dec(x_22); +x_69 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_69, 0, x_67); +lean_ctor_set(x_69, 1, x_68); +return x_69; } } -block_22: +block_21: { -lean_object* x_18; size_t x_19; size_t x_20; -x_18 = lean_ctor_get(x_16, 0); -lean_inc(x_18); -lean_dec(x_16); -x_19 = 1; -x_20 = lean_usize_add(x_4, x_19); -x_4 = x_20; -x_5 = x_18; -x_12 = x_17; +lean_object* x_17; size_t x_18; size_t x_19; +x_17 = lean_ctor_get(x_15, 0); +lean_inc(x_17); +lean_dec(x_15); +x_18 = 1; +x_19 = lean_usize_add(x_3, x_18); +x_3 = x_19; +x_4 = x_17; +x_11 = x_16; goto _start; } } @@ -19491,7 +19671,7 @@ lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -x_20 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInstance___spec__4(x_19, x_3, x_17, x_18, x_19, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_20 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInstance___spec__4(x_3, x_17, x_18, x_19, x_5, x_6, x_7, x_8, x_9, x_10, x_11); if (lean_obj_tag(x_20) == 0) { 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; 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; lean_object* x_56; lean_object* x_57; @@ -20061,19 +20241,19 @@ lean_dec(x_8); return x_16; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInstance___spec__4___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInstance___spec__4___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_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -size_t x_13; size_t x_14; lean_object* x_15; +size_t x_12; size_t x_13; lean_object* x_14; +x_12 = lean_unbox_usize(x_2); +lean_dec(x_2); x_13 = lean_unbox_usize(x_3); lean_dec(x_3); -x_14 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_15 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInstance___spec__4(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_7); +x_14 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInstance___spec__4(x_1, x_12, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_6); -lean_dec(x_2); -return x_15; +lean_dec(x_5); +lean_dec(x_1); +return x_14; } } LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInstance___spec__5___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) { @@ -20511,7 +20691,7 @@ lean_dec(x_1); return x_6; } } -static lean_object* _init_l_Lean_Server_RpcEncoding_initFn____x40_Lean_Server_Rpc_Deriving___hyg_8525____closed__1() { +static lean_object* _init_l_Lean_Server_RpcEncoding_initFn____x40_Lean_Server_Rpc_Deriving___hyg_8911____closed__1() { _start: { lean_object* x_1; @@ -20519,12 +20699,12 @@ x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Rpc_Deriving_0__Lean_Se return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_initFn____x40_Lean_Server_Rpc_Deriving___hyg_8525_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_RpcEncoding_initFn____x40_Lean_Server_Rpc_Deriving___hyg_8911_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; x_2 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveWithRefInstance___closed__211; -x_3 = l_Lean_Server_RpcEncoding_initFn____x40_Lean_Server_Rpc_Deriving___hyg_8525____closed__1; +x_3 = l_Lean_Server_RpcEncoding_initFn____x40_Lean_Server_Rpc_Deriving___hyg_8911____closed__1; x_4 = l_Lean_Elab_registerBuiltinDerivingHandlerWithArgs(x_2, x_3, x_1); if (lean_obj_tag(x_4) == 0) { @@ -21377,16 +21557,16 @@ l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_Rp lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__35___closed__1); l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__35___closed__2 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__35___closed__2(); lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__35___closed__2); -l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___lambda__1___closed__1 = _init_l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___lambda__1___closed__1); -l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___closed__1 = _init_l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___closed__1(); -lean_mark_persistent(l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___closed__1); -l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___closed__2 = _init_l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___closed__2(); -lean_mark_persistent(l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___closed__2); -l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___closed__3 = _init_l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___closed__3(); -lean_mark_persistent(l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___closed__3); -l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___closed__4 = _init_l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___closed__4(); -lean_mark_persistent(l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__41___rarg___closed__4); +l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___lambda__1___closed__1 = _init_l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___lambda__1___closed__1); +l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___closed__1 = _init_l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___closed__1(); +lean_mark_persistent(l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___closed__1); +l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___closed__2 = _init_l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___closed__2(); +lean_mark_persistent(l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___closed__2); +l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___closed__3 = _init_l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___closed__3(); +lean_mark_persistent(l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___closed__3); +l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___closed__4 = _init_l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___closed__4(); +lean_mark_persistent(l_Lean_Meta_withLocalDecls_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___spec__44___rarg___closed__4); l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__2___closed__1 = _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__2___closed__1(); lean_mark_persistent(l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__2___closed__1); l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__2___closed__2 = _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_deriveInductiveInstance___lambda__2___closed__2(); @@ -21499,9 +21679,9 @@ l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_dispatchDeriveIn lean_mark_persistent(l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_dispatchDeriveInstanceUnsafe___lambda__3___closed__6); l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_dispatchDeriveInstanceUnsafe___lambda__3___closed__7 = _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_dispatchDeriveInstanceUnsafe___lambda__3___closed__7(); lean_mark_persistent(l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncoding_dispatchDeriveInstanceUnsafe___lambda__3___closed__7); -l_Lean_Server_RpcEncoding_initFn____x40_Lean_Server_Rpc_Deriving___hyg_8525____closed__1 = _init_l_Lean_Server_RpcEncoding_initFn____x40_Lean_Server_Rpc_Deriving___hyg_8525____closed__1(); -lean_mark_persistent(l_Lean_Server_RpcEncoding_initFn____x40_Lean_Server_Rpc_Deriving___hyg_8525____closed__1); -res = l_Lean_Server_RpcEncoding_initFn____x40_Lean_Server_Rpc_Deriving___hyg_8525_(lean_io_mk_world()); +l_Lean_Server_RpcEncoding_initFn____x40_Lean_Server_Rpc_Deriving___hyg_8911____closed__1 = _init_l_Lean_Server_RpcEncoding_initFn____x40_Lean_Server_Rpc_Deriving___hyg_8911____closed__1(); +lean_mark_persistent(l_Lean_Server_RpcEncoding_initFn____x40_Lean_Server_Rpc_Deriving___hyg_8911____closed__1); +res = l_Lean_Server_RpcEncoding_initFn____x40_Lean_Server_Rpc_Deriving___hyg_8911_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Server/Rpc/RequestHandling.c b/stage0/stdlib/Lean/Server/Rpc/RequestHandling.c index 83648606b7a..7d4b76f7507 100644 --- a/stage0/stdlib/Lean/Server/Rpc/RequestHandling.c +++ b/stage0/stdlib/Lean/Server/Rpc/RequestHandling.c @@ -14,6 +14,7 @@ extern "C" { #endif static lean_object* l_Lean_Server_registerRpcProcedure___lambda__1___closed__7; +static lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__6; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_find_x3f___at___private_Lean_Server_Rpc_RequestHandling_0__Lean_Server_handleRpcCallUnsafe___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_322____spec__3(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_KeyedDeclsAttribute_ExtensionState_erase___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); @@ -27,7 +28,6 @@ static lean_object* l___private_Lean_Server_Rpc_RequestHandling_0__Lean_Server_h size_t lean_usize_add(size_t, size_t); static lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_69____closed__2; static lean_object* l_Lean_Server_wrapRpcProcedure___elambda__1___closed__2; -static lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__6; LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_RequestHandling_0__Lean_Server_handleRpcCallUnsafe___lambda__3___closed__11; LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_RequestHandling_0__Lean_Server_handleRpcCallUnsafe(lean_object*, lean_object*, lean_object*); @@ -40,6 +40,7 @@ uint8_t lean_uint64_dec_eq(uint64_t, uint64_t); static lean_object* l_Lean_Server_registerBuiltinRpcProcedure___closed__1; static lean_object* l_Lean_Server_registerRpcProcedure___lambda__1___closed__24; static lean_object* l_Lean_Server_registerRpcProcedure___lambda__3___closed__22; +LEAN_EXPORT lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_registerRpcProcedure___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_name_mk_string(lean_object*, lean_object*); static lean_object* l_Lean_Server_registerBuiltinRpcProcedure___lambda__2___closed__1; @@ -49,7 +50,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Serve LEAN_EXPORT lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_32_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_69_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_322_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256_(lean_object*); static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_322____spec__1___lambda__3___closed__1; static lean_object* l_Lean_Server_registerRpcProcedure___lambda__3___closed__18; lean_object* l_Lean_Server_FileWorker_instMonadRpcSession___rarg(lean_object*, lean_object*); @@ -66,7 +67,6 @@ static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_ini static lean_object* l_Lean_Server_wrapRpcProcedure___elambda__1___lambda__1___closed__1; static lean_object* l_Lean_Server_wrapRpcProcedure___elambda__1___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Server_wrapRpcProcedure___elambda__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____lambda__2___closed__2; static lean_object* l_Lean_Server_registerRpcProcedure___lambda__3___closed__3; static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_322____spec__1___lambda__3___closed__2; size_t lean_usize_sub(size_t, size_t); @@ -80,6 +80,7 @@ LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Server_registerRpcProcedure___s uint8_t lean_name_eq(lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_RequestHandling_0__Lean_Server_handleRpcCallUnsafe___lambda__3___closed__12; LEAN_EXPORT lean_object* l_Lean_Server_builtinRpcProcedures; +static lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__4; LEAN_EXPORT lean_object* l_Lean_Server_wrapRpcProcedure___elambda__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_wrapRpcProcedure___elambda__1___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_322____spec__2(lean_object*); @@ -102,11 +103,9 @@ LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Server_registerRpcProcedure___s LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_wrapRpcProcedure___elambda__1___spec__2(lean_object*); lean_object* lean_string_utf8_byte_size(lean_object*); lean_object* l_List_head_x21___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_registerRpcProcedure___lambda__3___closed__25; uint8_t l___private_Lean_Message_0__Lean_beqMessageSeverity____x40_Lean_Message___hyg_101_(uint8_t, uint8_t); -static lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__4; -extern lean_object* l_Lean_nameLitKind; +static lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__7; lean_object* lean_nat_add(lean_object*, lean_object*); static lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_32____closed__3; static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_322____spec__1___lambda__4___closed__1; @@ -117,7 +116,6 @@ LEAN_EXPORT lean_object* l_Lean_log___at_Lean_Server_registerRpcProcedure___spec static lean_object* l___private_Lean_Server_Rpc_RequestHandling_0__Lean_Server_handleRpcCallUnsafe___lambda__3___closed__21; size_t lean_uint64_to_usize(uint64_t); lean_object* l_Lean_Expr_FindImpl_findUnsafe_x3f(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MapDeclarationExtension_find_x3f___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_compileDecl___at_Lean_Server_registerRpcProcedure___spec__6(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_RequestHandling_0__Lean_Server_handleRpcCallUnsafe___lambda__3___closed__2; @@ -128,9 +126,9 @@ lean_object* l_Except_map___rarg(lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); lean_object* l_Lean_MapDeclarationExtension_insert___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_registerRpcProcedure(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__7; static lean_object* l___private_Lean_Server_Rpc_RequestHandling_0__Lean_Server_handleRpcCallUnsafe___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_322____spec__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__1; static lean_object* l_Lean_Server_registerRpcProcedure___lambda__1___closed__13; LEAN_EXPORT lean_object* l_Lean_Server_userRpcProcedures; lean_object* l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(lean_object*, lean_object*); @@ -140,11 +138,9 @@ lean_object* l_Lean_Json_compress(lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_findAux___at___private_Lean_Server_Rpc_RequestHandling_0__Lean_Server_handleRpcCallUnsafe___spec__2___boxed(lean_object*, lean_object*, lean_object*); static size_t l_Std_PersistentHashMap_findAux___at___private_Lean_Server_Rpc_RequestHandling_0__Lean_Server_handleRpcCallUnsafe___spec__2___closed__1; -static lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__1; LEAN_EXPORT uint8_t l_Lean_Server_registerRpcProcedure___lambda__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_322____spec__1___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_containsAux___at_Lean_Server_registerBuiltinRpcProcedure___spec__6___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__2; LEAN_EXPORT lean_object* l_Lean_addDecl___at_Lean_Server_registerRpcProcedure___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_instInhabitedRpcProcedure___rarg___closed__2; lean_object* l_Lean_KernelException_toMessageData(lean_object*, lean_object*); @@ -159,11 +155,11 @@ static lean_object* l_Lean_Server_registerRpcProcedure___closed__2; lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); lean_object* l_Std_PersistentHashMap_insert___at_Lean_Server_registerLspRequestHandler___spec__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_registerRpcProcedure___lambda__1___closed__21; +static lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____lambda__2___closed__2; LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Server_wrapRpcProcedure___elambda__1___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwKernelException___at_Lean_Server_registerRpcProcedure___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_322____closed__1; lean_object* l_Lean_mkMapDeclarationExtension___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_containsAtAux___at_Lean_Server_registerBuiltinRpcProcedure___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_RequestHandling_0__Lean_Server_handleRpcCallUnsafe___lambda__2(lean_object*, lean_object*, lean_object*); uint64_t l_Lean_Name_hash(lean_object*); @@ -174,8 +170,8 @@ static lean_object* l_Lean_Server_registerRpcProcedure___closed__4; lean_object* lean_st_mk_ref(lean_object*, lean_object*); static lean_object* l_Lean_Server_registerRpcProcedure___lambda__1___closed__1; uint8_t lean_uint64_dec_lt(uint64_t, uint64_t); +static lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__2; extern lean_object* l_Lean_Server_requestHandlers; -lean_object* l___private_Init_Meta_0__Lean_quoteNameMk(lean_object*); extern lean_object* l_Lean_Expr_instHashableExpr; LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_Server_registerRpcProcedure___spec__5(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_registerRpcProcedure___lambda__1___closed__17; @@ -183,6 +179,7 @@ static lean_object* l_Lean_Server_registerRpcProcedure___lambda__3___closed__24; lean_object* l_Lean_Server_RequestM_bindTask___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); uint8_t l_Std_PersistentHashMap_contains___at_Lean_Server_registerLspRequestHandler___spec__6(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_left(size_t, size_t); lean_object* l_Lean_Meta_InfoCacheKey_instHashableInfoCacheKey___boxed(lean_object*); @@ -198,6 +195,7 @@ static lean_object* l_Lean_Server_instInhabitedRpcProcedure___rarg___closed__4; static lean_object* l___private_Lean_Server_Rpc_RequestHandling_0__Lean_Server_handleRpcCallUnsafe___lambda__3___closed__5; lean_object* l_Lean_registerBuiltinAttribute(lean_object*, lean_object*); lean_object* l_Lean_FileMap_toPosition(lean_object*, lean_object*); +static lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__3; lean_object* l_Lean_throwError___at_Lean_AttributeImpl_erase___default___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_RequestHandling_0__Lean_Server_handleRpcCallUnsafe___lambda__3___closed__14; static lean_object* l_Lean_Server_registerRpcProcedure___lambda__1___closed__4; @@ -207,6 +205,7 @@ size_t lean_usize_mul(size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_RequestHandling_0__Lean_Server_handleRpcCallUnsafe___lambda__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_registerRpcProcedure___lambda__4___closed__5; LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAux_traverse___at_Lean_Server_registerBuiltinRpcProcedure___spec__3(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_PersistentHashMap_insertAux___at_Lean_Server_registerBuiltinRpcProcedure___spec__2___closed__1; static lean_object* l_Lean_Server_registerRpcProcedure___lambda__1___closed__9; static lean_object* l_Lean_Server_registerRpcProcedure___lambda__1___closed__8; @@ -220,7 +219,6 @@ static lean_object* l_Lean_Server_registerRpcProcedure___lambda__4___closed__3; static lean_object* l___private_Lean_Server_Rpc_RequestHandling_0__Lean_Server_handleRpcCallUnsafe___lambda__3___closed__18; LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_322____spec__1___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_RequestHandling_0__Lean_Server_handleRpcCallUnsafe___lambda__3___closed__15; -static lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__3; LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_wrapRpcProcedure___elambda__1___spec__2___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_RequestHandling_0__Lean_Server_handleRpcCallUnsafe___lambda__3___closed__22; lean_object* l_instBEqProd___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -235,6 +233,7 @@ LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAux___at_Lean_Server_regi static lean_object* l_Lean_Server_registerRpcProcedure___lambda__3___closed__19; static lean_object* l___private_Lean_Server_Rpc_RequestHandling_0__Lean_Server_handleRpcCallUnsafe___lambda__3___closed__17; static lean_object* l___private_Lean_Server_Rpc_RequestHandling_0__Lean_Server_handleRpcCallUnsafe___lambda__3___closed__8; +static lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____lambda__2___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_RequestHandling_0__Lean_Server_handleRpcCallUnsafe___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getPos_x3f(lean_object*, uint8_t); uint8_t l_Array_contains___at___private_Lean_Meta_Match_Value_0__Lean_Meta_isUIntTypeName___spec__1(lean_object*, lean_object*); @@ -251,7 +250,6 @@ LEAN_EXPORT lean_object* l___private_Lean_MonadEnv_0__Lean_checkUnsupported___at static lean_object* l_Lean_Server_registerRpcProcedure___lambda__1___closed__14; static lean_object* l_Lean_Server_registerRpcProcedure___lambda__3___closed__26; lean_object* lean_environment_main_module(lean_object*); -static lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__5; static lean_object* l_Lean_Server_registerRpcProcedure___lambda__1___closed__15; static lean_object* l_Lean_Server_registerRpcProcedure___lambda__1___closed__18; LEAN_EXPORT lean_object* l_Lean_Server_instInhabitedRpcProcedure(uint64_t, lean_object*, lean_object*); @@ -265,6 +263,7 @@ static lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_initFn____ static lean_object* l_Lean_Server_wrapRpcProcedure___elambda__1___lambda__1___closed__5; lean_object* l_Lean_Name_append(lean_object*, lean_object*); static lean_object* l_Lean_Server_registerRpcProcedure___lambda__3___closed__11; +lean_object* l_Lean_quoteNameMk(lean_object*); static lean_object* l_Lean_Server_wrapRpcProcedure___elambda__1___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Server_wrapRpcProcedure(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Elab_Command_instInhabitedScope; @@ -275,10 +274,13 @@ static lean_object* l___private_Lean_Server_Rpc_RequestHandling_0__Lean_Server_h static lean_object* l_Lean_Server_registerRpcProcedure___lambda__3___closed__12; LEAN_EXPORT lean_object* l_Std_RBNode_find___at_Lean_Server_wrapRpcProcedure___elambda__1___spec__1(lean_object*, uint64_t); LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Server_registerRpcProcedure___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__5; static lean_object* l___private_Lean_Server_Rpc_RequestHandling_0__Lean_Server_handleRpcCallUnsafe___lambda__3___closed__19; LEAN_EXPORT uint8_t l_Std_PersistentHashMap_contains___at_Lean_Server_registerBuiltinRpcProcedure___spec__5(lean_object*, lean_object*); static lean_object* l_Lean_Server_registerRpcProcedure___lambda__1___closed__6; +lean_object* l_Lean_Syntax_mkNameLit(lean_object*, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Server_Snapshots_Snapshot_endPos(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_registerRpcProcedure___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_wrapRpcProcedure___elambda__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -338,20 +340,17 @@ lean_object* l_Lean_Server_Snapshots_Snapshot_env(lean_object*); LEAN_EXPORT lean_object* l_Lean_addAndCompile___at_Lean_Server_registerRpcProcedure___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l___private_Lean_MonadEnv_0__Lean_supportedRecursors; lean_object* l_Std_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); -static lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____lambda__2___closed__1; static lean_object* l_Lean_Server_registerRpcProcedure___lambda__3___closed__9; static lean_object* l_Lean_Server_registerRpcProcedure___lambda__1___closed__16; static lean_object* l___private_Lean_Server_Rpc_RequestHandling_0__Lean_Server_handleRpcCallUnsafe___closed__3; lean_object* l_Lean_Server_RequestM_readDoc(lean_object*, lean_object*); static lean_object* l_Lean_Server_registerRpcProcedure___lambda__3___closed__4; static lean_object* l___private_Lean_Server_Rpc_RequestHandling_0__Lean_Server_handleRpcCallUnsafe___lambda__3___closed__7; -LEAN_EXPORT lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Server_registerBuiltinRpcProcedure___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_registerRpcProcedure___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_PersistentHashMap_mkCollisionNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_RequestHandling_0__Lean_Server_handleRpcCallUnsafe___lambda__3___closed__13; static lean_object* l_Lean_Server_registerRpcProcedure___lambda__3___closed__17; -lean_object* l_Lean_Syntax_mkLit(lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); lean_object* lean_add_decl(lean_object*, lean_object*); static lean_object* l_Lean_Server_registerRpcProcedure___lambda__4___closed__1; @@ -1210,7 +1209,6 @@ lean_closure_set(x_32, 0, x_31); x_33 = lean_alloc_closure((void*)(l___private_Lean_Server_Rpc_RequestHandling_0__Lean_Server_handleRpcCallUnsafe___lambda__3___boxed), 4, 1); lean_closure_set(x_33, 0, x_1); x_34 = l_Lean_Server_RequestM_bindWaitFindSnap___rarg(x_5, x_12, x_32, x_33, x_2, x_6); -lean_dec(x_5); return x_34; } } @@ -5794,7 +5792,7 @@ x_37 = lean_array_push(x_36, x_26); if (lean_obj_tag(x_27) == 0) { lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; lean_object* x_51; -x_38 = l___private_Init_Meta_0__Lean_quoteNameMk(x_2); +x_38 = l_Lean_quoteNameMk(x_2); x_39 = l_Lean_Server_registerRpcProcedure___lambda__1___closed__20; x_40 = lean_array_push(x_39, x_38); lean_inc(x_34); @@ -5868,171 +5866,169 @@ lean_dec(x_2); x_59 = !lean_is_exclusive(x_27); if (x_59 == 0) { -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; 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; lean_object* x_79; uint8_t x_80; lean_object* x_81; +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; 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; uint8_t x_79; lean_object* x_80; x_60 = lean_ctor_get(x_27, 0); x_61 = l_Lean_Server_registerRpcProcedure___lambda__1___closed__24; x_62 = l_String_intercalate(x_61, x_60); x_63 = l_Lean_Server_registerRpcProcedure___lambda__1___closed__25; x_64 = lean_string_append(x_63, x_62); lean_dec(x_62); -x_65 = l_Lean_nameLitKind; -x_66 = l_Lean_Syntax_mkLit(x_65, x_64, x_32); -x_67 = lean_array_push(x_30, x_66); -x_68 = l_Lean_Server_registerRpcProcedure___lambda__1___closed__23; -x_69 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_69, 0, x_32); -lean_ctor_set(x_69, 1, x_68); -lean_ctor_set(x_69, 2, x_67); -x_70 = l_Lean_Server_registerRpcProcedure___lambda__1___closed__20; -x_71 = lean_array_push(x_70, x_69); +x_65 = l_Lean_Syntax_mkNameLit(x_64, x_32); +x_66 = lean_array_push(x_30, x_65); +x_67 = l_Lean_Server_registerRpcProcedure___lambda__1___closed__23; +x_68 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_68, 0, x_32); +lean_ctor_set(x_68, 1, x_67); +lean_ctor_set(x_68, 2, x_66); +x_69 = l_Lean_Server_registerRpcProcedure___lambda__1___closed__20; +x_70 = lean_array_push(x_69, x_68); lean_inc(x_34); +x_71 = lean_array_push(x_70, x_34); x_72 = lean_array_push(x_71, x_34); -x_73 = lean_array_push(x_72, x_34); -x_74 = lean_array_push(x_73, x_35); -x_75 = l_Lean_Server_registerRpcProcedure___lambda__1___closed__15; -x_76 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_76, 0, x_32); -lean_ctor_set(x_76, 1, x_75); -lean_ctor_set(x_76, 2, x_74); -x_77 = lean_array_push(x_37, x_76); -x_78 = l_Lean_Server_registerRpcProcedure___lambda__1___closed__7; -x_79 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_79, 0, x_32); -lean_ctor_set(x_79, 1, x_78); -lean_ctor_set(x_79, 2, x_77); +x_73 = lean_array_push(x_72, x_35); +x_74 = l_Lean_Server_registerRpcProcedure___lambda__1___closed__15; +x_75 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_75, 0, x_32); +lean_ctor_set(x_75, 1, x_74); +lean_ctor_set(x_75, 2, x_73); +x_76 = lean_array_push(x_37, x_75); +x_77 = l_Lean_Server_registerRpcProcedure___lambda__1___closed__7; +x_78 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_78, 0, x_32); +lean_ctor_set(x_78, 1, x_77); +lean_ctor_set(x_78, 2, x_76); lean_ctor_set(x_27, 0, x_3); -x_80 = 1; +x_79 = 1; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_81 = l_Lean_Elab_Term_elabTerm(x_79, x_27, x_80, x_80, x_4, x_5, x_6, x_7, x_8, x_9, x_17); -if (lean_obj_tag(x_81) == 0) +x_80 = l_Lean_Elab_Term_elabTerm(x_78, x_27, x_79, x_79, x_4, x_5, x_6, x_7, x_8, x_9, x_17); +if (lean_obj_tag(x_80) == 0) { -lean_object* x_82; lean_object* x_83; lean_object* x_84; -x_82 = lean_ctor_get(x_81, 0); +lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_81 = lean_ctor_get(x_80, 0); +lean_inc(x_81); +x_82 = lean_ctor_get(x_80, 1); lean_inc(x_82); -x_83 = lean_ctor_get(x_81, 1); -lean_inc(x_83); -lean_dec(x_81); -x_84 = l_Lean_Meta_instantiateMVars(x_82, x_6, x_7, x_8, x_9, x_83); +lean_dec(x_80); +x_83 = l_Lean_Meta_instantiateMVars(x_81, x_6, x_7, x_8, x_9, x_82); lean_dec(x_9); lean_dec(x_8); lean_dec(x_6); -return x_84; +return x_83; } else { -uint8_t x_85; +uint8_t x_84; lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -x_85 = !lean_is_exclusive(x_81); -if (x_85 == 0) +x_84 = !lean_is_exclusive(x_80); +if (x_84 == 0) { -return x_81; +return x_80; } else { -lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_86 = lean_ctor_get(x_81, 0); -x_87 = lean_ctor_get(x_81, 1); -lean_inc(x_87); +lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_85 = lean_ctor_get(x_80, 0); +x_86 = lean_ctor_get(x_80, 1); lean_inc(x_86); -lean_dec(x_81); -x_88 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_88, 0, x_86); -lean_ctor_set(x_88, 1, x_87); -return x_88; +lean_inc(x_85); +lean_dec(x_80); +x_87 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_87, 0, x_85); +lean_ctor_set(x_87, 1, x_86); +return x_87; } } } else { -lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; uint8_t x_110; lean_object* x_111; -x_89 = lean_ctor_get(x_27, 0); -lean_inc(x_89); +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; uint8_t x_108; lean_object* x_109; +x_88 = lean_ctor_get(x_27, 0); +lean_inc(x_88); lean_dec(x_27); -x_90 = l_Lean_Server_registerRpcProcedure___lambda__1___closed__24; -x_91 = l_String_intercalate(x_90, x_89); -x_92 = l_Lean_Server_registerRpcProcedure___lambda__1___closed__25; -x_93 = lean_string_append(x_92, x_91); -lean_dec(x_91); -x_94 = l_Lean_nameLitKind; -x_95 = l_Lean_Syntax_mkLit(x_94, x_93, x_32); -x_96 = lean_array_push(x_30, x_95); -x_97 = l_Lean_Server_registerRpcProcedure___lambda__1___closed__23; -x_98 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_98, 0, x_32); -lean_ctor_set(x_98, 1, x_97); -lean_ctor_set(x_98, 2, x_96); -x_99 = l_Lean_Server_registerRpcProcedure___lambda__1___closed__20; -x_100 = lean_array_push(x_99, x_98); +x_89 = l_Lean_Server_registerRpcProcedure___lambda__1___closed__24; +x_90 = l_String_intercalate(x_89, x_88); +x_91 = l_Lean_Server_registerRpcProcedure___lambda__1___closed__25; +x_92 = lean_string_append(x_91, x_90); +lean_dec(x_90); +x_93 = l_Lean_Syntax_mkNameLit(x_92, x_32); +x_94 = lean_array_push(x_30, x_93); +x_95 = l_Lean_Server_registerRpcProcedure___lambda__1___closed__23; +x_96 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_96, 0, x_32); +lean_ctor_set(x_96, 1, x_95); +lean_ctor_set(x_96, 2, x_94); +x_97 = l_Lean_Server_registerRpcProcedure___lambda__1___closed__20; +x_98 = lean_array_push(x_97, x_96); lean_inc(x_34); -x_101 = lean_array_push(x_100, x_34); -x_102 = lean_array_push(x_101, x_34); -x_103 = lean_array_push(x_102, x_35); -x_104 = l_Lean_Server_registerRpcProcedure___lambda__1___closed__15; -x_105 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_105, 0, x_32); -lean_ctor_set(x_105, 1, x_104); -lean_ctor_set(x_105, 2, x_103); -x_106 = lean_array_push(x_37, x_105); -x_107 = l_Lean_Server_registerRpcProcedure___lambda__1___closed__7; -x_108 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_108, 0, x_32); -lean_ctor_set(x_108, 1, x_107); -lean_ctor_set(x_108, 2, x_106); -x_109 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_109, 0, x_3); -x_110 = 1; +x_99 = lean_array_push(x_98, x_34); +x_100 = lean_array_push(x_99, x_34); +x_101 = lean_array_push(x_100, x_35); +x_102 = l_Lean_Server_registerRpcProcedure___lambda__1___closed__15; +x_103 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_103, 0, x_32); +lean_ctor_set(x_103, 1, x_102); +lean_ctor_set(x_103, 2, x_101); +x_104 = lean_array_push(x_37, x_103); +x_105 = l_Lean_Server_registerRpcProcedure___lambda__1___closed__7; +x_106 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_106, 0, x_32); +lean_ctor_set(x_106, 1, x_105); +lean_ctor_set(x_106, 2, x_104); +x_107 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_107, 0, x_3); +x_108 = 1; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_111 = l_Lean_Elab_Term_elabTerm(x_108, x_109, x_110, x_110, x_4, x_5, x_6, x_7, x_8, x_9, x_17); -if (lean_obj_tag(x_111) == 0) +x_109 = l_Lean_Elab_Term_elabTerm(x_106, x_107, x_108, x_108, x_4, x_5, x_6, x_7, x_8, x_9, x_17); +if (lean_obj_tag(x_109) == 0) { -lean_object* x_112; lean_object* x_113; lean_object* x_114; -x_112 = lean_ctor_get(x_111, 0); -lean_inc(x_112); -x_113 = lean_ctor_get(x_111, 1); -lean_inc(x_113); -lean_dec(x_111); -x_114 = l_Lean_Meta_instantiateMVars(x_112, x_6, x_7, x_8, x_9, x_113); +lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_110 = lean_ctor_get(x_109, 0); +lean_inc(x_110); +x_111 = lean_ctor_get(x_109, 1); +lean_inc(x_111); +lean_dec(x_109); +x_112 = l_Lean_Meta_instantiateMVars(x_110, x_6, x_7, x_8, x_9, x_111); lean_dec(x_9); lean_dec(x_8); lean_dec(x_6); -return x_114; +return x_112; } else { -lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -x_115 = lean_ctor_get(x_111, 0); -lean_inc(x_115); -x_116 = lean_ctor_get(x_111, 1); -lean_inc(x_116); -if (lean_is_exclusive(x_111)) { - lean_ctor_release(x_111, 0); - lean_ctor_release(x_111, 1); - x_117 = x_111; +x_113 = lean_ctor_get(x_109, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_109, 1); +lean_inc(x_114); +if (lean_is_exclusive(x_109)) { + lean_ctor_release(x_109, 0); + lean_ctor_release(x_109, 1); + x_115 = x_109; } else { - lean_dec_ref(x_111); - x_117 = lean_box(0); + lean_dec_ref(x_109); + x_115 = lean_box(0); } -if (lean_is_scalar(x_117)) { - x_118 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_115)) { + x_116 = lean_alloc_ctor(1, 2, 0); } else { - x_118 = x_117; + x_116 = x_115; } -lean_ctor_set(x_118, 0, x_115); -lean_ctor_set(x_118, 1, x_116); -return x_118; +lean_ctor_set(x_116, 0, x_113); +lean_ctor_set(x_116, 1, x_114); +return x_116; } } } @@ -6847,7 +6843,7 @@ lean_dec(x_2); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____lambda__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____lambda__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_7; @@ -6855,7 +6851,7 @@ x_7 = l_Lean_Server_registerRpcProcedure(x_1, x_4, x_5, x_6); return x_7; } } -static lean_object* _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____lambda__2___closed__1() { +static lean_object* _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____lambda__2___closed__1() { _start: { lean_object* x_1; @@ -6863,25 +6859,25 @@ x_1 = lean_mk_string_from_bytes("attribute cannot be erased", 26); return x_1; } } -static lean_object* _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____lambda__2___closed__2() { +static lean_object* _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____lambda__2___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____lambda__2___closed__1; +x_1 = l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____lambda__2___closed__1; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____lambda__2(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; -x_5 = l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____lambda__2___closed__2; +x_5 = l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____lambda__2___closed__2; x_6 = l_Lean_throwError___at_Lean_AttributeImpl_erase___default___spec__1(x_5, x_2, x_3, x_4); return x_6; } } -static lean_object* _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__1() { +static lean_object* _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__1() { _start: { lean_object* x_1; @@ -6889,17 +6885,17 @@ x_1 = lean_mk_string_from_bytes("serverRpcMethod", 15); return x_1; } } -static lean_object* _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__2() { +static lean_object* _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__1; +x_2 = l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__3() { +static lean_object* _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__3() { _start: { lean_object* x_1; @@ -6907,12 +6903,12 @@ x_1 = lean_mk_string_from_bytes("Marks a function as a Lean server RPC method.\n return x_1; } } -static lean_object* _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__4() { +static lean_object* _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__4() { _start: { lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; -x_1 = l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__2; -x_2 = l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__3; +x_1 = l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__2; +x_2 = l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__3; x_3 = 1; x_4 = lean_alloc_ctor(0, 2, 1); lean_ctor_set(x_4, 0, x_1); @@ -6921,29 +6917,29 @@ lean_ctor_set_uint8(x_4, sizeof(void*)*2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__5() { +static lean_object* _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__5() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____lambda__1___boxed), 6, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____lambda__1___boxed), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__6() { +static lean_object* _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__6() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____lambda__2___boxed), 4, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____lambda__2___boxed), 4, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__7() { +static lean_object* _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__4; -x_2 = l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__5; -x_3 = l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__6; +x_1 = l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__4; +x_2 = l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__5; +x_3 = l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__6; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -6951,31 +6947,31 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__7; +x_2 = l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__7; x_3 = l_Lean_registerBuiltinAttribute(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____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_EXPORT lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____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) { _start: { uint8_t x_7; lean_object* x_8; x_7 = lean_unbox(x_3); lean_dec(x_3); -x_8 = l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____lambda__1(x_1, x_2, x_7, x_4, x_5, x_6); +x_8 = l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____lambda__1(x_1, x_2, x_7, x_4, x_5, x_6); lean_dec(x_2); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____lambda__2___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_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____lambda__2(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____lambda__2(x_1, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); @@ -7268,25 +7264,25 @@ l_Lean_Server_registerRpcProcedure___closed__3 = _init_l_Lean_Server_registerRpc lean_mark_persistent(l_Lean_Server_registerRpcProcedure___closed__3); l_Lean_Server_registerRpcProcedure___closed__4 = _init_l_Lean_Server_registerRpcProcedure___closed__4(); lean_mark_persistent(l_Lean_Server_registerRpcProcedure___closed__4); -l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____lambda__2___closed__1 = _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____lambda__2___closed__1); -l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____lambda__2___closed__2 = _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____lambda__2___closed__2(); -lean_mark_persistent(l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____lambda__2___closed__2); -l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__1 = _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__1(); -lean_mark_persistent(l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__1); -l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__2 = _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__2(); -lean_mark_persistent(l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__2); -l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__3 = _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__3(); -lean_mark_persistent(l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__3); -l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__4 = _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__4(); -lean_mark_persistent(l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__4); -l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__5 = _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__5(); -lean_mark_persistent(l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__5); -l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__6 = _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__6(); -lean_mark_persistent(l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__6); -l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__7 = _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__7(); -lean_mark_persistent(l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244____closed__7); -res = l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1244_(lean_io_mk_world()); +l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____lambda__2___closed__1 = _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____lambda__2___closed__1); +l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____lambda__2___closed__2 = _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____lambda__2___closed__2); +l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__1 = _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__1(); +lean_mark_persistent(l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__1); +l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__2 = _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__2(); +lean_mark_persistent(l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__2); +l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__3 = _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__3(); +lean_mark_persistent(l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__3); +l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__4 = _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__4(); +lean_mark_persistent(l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__4); +l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__5 = _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__5(); +lean_mark_persistent(l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__5); +l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__6 = _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__6(); +lean_mark_persistent(l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__6); +l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__7 = _init_l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__7(); +lean_mark_persistent(l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256____closed__7); +res = l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_1256_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Server/Snapshots.c b/stage0/stdlib/Lean/Server/Snapshots.c index 8ce9ac01633..debe8a999a4 100644 --- a/stage0/stdlib/Lean/Server/Snapshots.c +++ b/stage0/stdlib/Lean/Server/Snapshots.c @@ -13,32 +13,29 @@ #ifdef __cplusplus extern "C" { #endif -static lean_object* l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__1; size_t lean_usize_add(size_t, size_t); -static lean_object* l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__2; LEAN_EXPORT lean_object* l_Lean_Server_Snapshots_Snapshot_msgLog___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_Snapshots_Snapshot_diagnostics___spec__3___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); +static lean_object* l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__3; LEAN_EXPORT lean_object* l_Lean_Server_Snapshots_compileNextCmd___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_name_mk_string(lean_object*, lean_object*); -lean_object* l_Lean_Parser_parseHeader(lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Server_Snapshots_compileNextCmd_withNewInteractiveDiags___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static uint32_t l_Lean_Server_Snapshots_instInhabitedSnapshot___closed__6; -static lean_object* l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__4; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Snapshots_server_stderrAsMessages; static lean_object* l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_6____closed__3; -static lean_object* l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__5; static lean_object* l_Lean_Server_Snapshots_instInhabitedSnapshot___closed__3; lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Snapshots_parseAhead_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__1; lean_object* lean_array_push(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Snapshots_Snapshot_infoTree(lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Snapshots_parseAhead_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__3; +static lean_object* l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__2; static lean_object* l_Lean_Server_Snapshots_Snapshot_infoTree___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_Snapshots_Snapshot_diagnostics___spec__4(size_t, size_t, lean_object*); static lean_object* l_Lean_Server_Snapshots_instInhabitedSnapshot___closed__14; @@ -68,7 +65,6 @@ static lean_object* l_Lean_Server_Snapshots_Snapshot_infoTree___closed__6; lean_object* lean_st_ref_take(lean_object*, lean_object*); static lean_object* l_Lean_Server_Snapshots_instInhabitedSnapshot___closed__7; lean_object* lean_nat_sub(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_Snapshots_reparseHeader(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentArray_mapMAux___at_Lean_Server_Snapshots_Snapshot_diagnostics___spec__2(lean_object*); lean_object* l_Std_PersistentArray_get_x21___rarg(lean_object*, lean_object*, lean_object*); lean_object* lean_string_from_utf8_unchecked(lean_object*); @@ -80,13 +76,14 @@ extern lean_object* l_Lean_instInhabitedMessage; lean_object* lean_st_mk_ref(lean_object*, lean_object*); lean_object* l_Std_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); static lean_object* l_Lean_Server_Snapshots_compileNextCmd___closed__6; +static lean_object* l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__5; lean_object* l_Lean_Elab_Command_withLogging(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_FileMap_toPosition(lean_object*, lean_object*); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_firstFrontendMacroScope; LEAN_EXPORT lean_object* l_Lean_Server_Snapshots_instInhabitedSnapshot; LEAN_EXPORT lean_object* l_Lean_Server_Snapshots_Snapshot_endPos___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_6_(lean_object*); size_t lean_usize_of_nat(lean_object*); lean_object* l_Lean_Widget_InteractiveDiagnostic_toDiagnostic(lean_object*); @@ -115,6 +112,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_Snapshots_Snapshot_endPos(lean_object*); static lean_object* l_Lean_Server_Snapshots_instInhabitedSnapshot___closed__8; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_Server_Snapshots_compileNextCmd___closed__1; +static lean_object* l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__4; LEAN_EXPORT lean_object* l_panic___at_Lean_Server_Snapshots_Snapshot_infoTree___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Snapshots_compileNextCmd___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); @@ -128,8 +126,8 @@ uint8_t l_Lean_Parser_isEOI(lean_object*); static lean_object* l_Lean_Server_Snapshots_compileNextCmd___closed__2; static lean_object* l_Lean_Server_Snapshots_instInhabitedSnapshot___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_Snapshots_Snapshot_diagnostics___spec__4___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__6; LEAN_EXPORT lean_object* l_Lean_Server_Snapshots_compileNextCmd_withNewInteractiveDiags(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); -static lean_object* l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__6; static lean_object* l_Lean_Server_Snapshots_Snapshot_infoTree___closed__3; static lean_object* l_Lean_Server_Snapshots_instInhabitedSnapshot___closed__13; uint32_t lean_uint32_of_nat(lean_object*); @@ -853,140 +851,6 @@ x_3 = lean_box(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Server_Snapshots_reparseHeader(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = l_Lean_Parser_parseHeader(x_1, x_3); -if (lean_obj_tag(x_4) == 0) -{ -lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_5 = lean_ctor_get(x_4, 0); -lean_inc(x_5); -x_6 = lean_ctor_get(x_5, 1); -lean_inc(x_6); -x_7 = !lean_is_exclusive(x_4); -if (x_7 == 0) -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_ctor_get(x_4, 0); -lean_dec(x_8); -x_9 = lean_ctor_get(x_5, 0); -lean_inc(x_9); -lean_dec(x_5); -x_10 = lean_ctor_get(x_6, 0); -lean_inc(x_10); -lean_dec(x_6); -x_11 = !lean_is_exclusive(x_2); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_2, 2); -lean_dec(x_12); -x_13 = lean_ctor_get(x_2, 1); -lean_dec(x_13); -lean_ctor_set(x_2, 2, x_10); -lean_ctor_set(x_2, 1, x_9); -lean_ctor_set(x_4, 0, x_2); -return x_4; -} -else -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_14 = lean_ctor_get(x_2, 0); -x_15 = lean_ctor_get(x_2, 3); -x_16 = lean_ctor_get(x_2, 4); -x_17 = lean_ctor_get(x_2, 5); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_14); -lean_dec(x_2); -x_18 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_18, 0, x_14); -lean_ctor_set(x_18, 1, x_9); -lean_ctor_set(x_18, 2, x_10); -lean_ctor_set(x_18, 3, x_15); -lean_ctor_set(x_18, 4, x_16); -lean_ctor_set(x_18, 5, x_17); -lean_ctor_set(x_4, 0, x_18); -return x_4; -} -} -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; -x_19 = lean_ctor_get(x_4, 1); -lean_inc(x_19); -lean_dec(x_4); -x_20 = lean_ctor_get(x_5, 0); -lean_inc(x_20); -lean_dec(x_5); -x_21 = lean_ctor_get(x_6, 0); -lean_inc(x_21); -lean_dec(x_6); -x_22 = lean_ctor_get(x_2, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_2, 3); -lean_inc(x_23); -x_24 = lean_ctor_get(x_2, 4); -lean_inc(x_24); -x_25 = lean_ctor_get(x_2, 5); -lean_inc(x_25); -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); - lean_ctor_release(x_2, 5); - x_26 = x_2; -} else { - lean_dec_ref(x_2); - x_26 = lean_box(0); -} -if (lean_is_scalar(x_26)) { - x_27 = lean_alloc_ctor(0, 6, 0); -} else { - x_27 = x_26; -} -lean_ctor_set(x_27, 0, x_22); -lean_ctor_set(x_27, 1, x_20); -lean_ctor_set(x_27, 2, x_21); -lean_ctor_set(x_27, 3, x_23); -lean_ctor_set(x_27, 4, x_24); -lean_ctor_set(x_27, 5, x_25); -x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_27); -lean_ctor_set(x_28, 1, x_19); -return x_28; -} -} -else -{ -uint8_t x_29; -lean_dec(x_2); -x_29 = !lean_is_exclusive(x_4); -if (x_29 == 0) -{ -return x_4; -} -else -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_4, 0); -x_31 = lean_ctor_get(x_4, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_4); -x_32 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -return x_32; -} -} -} -} LEAN_EXPORT lean_object* l_Lean_Server_Snapshots_parseNextCmd(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { @@ -1128,7 +992,7 @@ lean_dec(x_2); return x_15; } } -static lean_object* _init_l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__1() { +static lean_object* _init_l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__1() { _start: { lean_object* x_1; @@ -1136,17 +1000,17 @@ x_1 = lean_mk_string_from_bytes("server", 6); return x_1; } } -static lean_object* _init_l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__2() { +static lean_object* _init_l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__1; +x_2 = l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__3() { +static lean_object* _init_l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__3() { _start: { lean_object* x_1; @@ -1154,17 +1018,17 @@ x_1 = lean_mk_string_from_bytes("stderrAsMessages", 16); return x_1; } } -static lean_object* _init_l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__4() { +static lean_object* _init_l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__2; -x_2 = l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__3; +x_1 = l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__2; +x_2 = l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__5() { +static lean_object* _init_l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__5() { _start: { lean_object* x_1; @@ -1172,13 +1036,13 @@ x_1 = lean_mk_string_from_bytes("(server) capture output to the Lean stderr chan return x_1; } } -static lean_object* _init_l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__6() { +static lean_object* _init_l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__6() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = 1; -x_2 = l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__1; -x_3 = l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__5; +x_2 = l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__1; +x_3 = l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__5; x_4 = lean_box(x_1); x_5 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_5, 0, x_4); @@ -1187,12 +1051,12 @@ lean_ctor_set(x_5, 2, x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__4; -x_3 = l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__6; +x_2 = l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__4; +x_3 = l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__6; x_4 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_Options___hyg_6____spec__1(x_2, x_3, x_1); return x_4; } @@ -2765,19 +2629,19 @@ l_Lean_Server_Snapshots_Snapshot_infoTree___closed__5 = _init_l_Lean_Server_Snap lean_mark_persistent(l_Lean_Server_Snapshots_Snapshot_infoTree___closed__5); l_Lean_Server_Snapshots_Snapshot_infoTree___closed__6 = _init_l_Lean_Server_Snapshots_Snapshot_infoTree___closed__6(); lean_mark_persistent(l_Lean_Server_Snapshots_Snapshot_infoTree___closed__6); -l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__1 = _init_l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__1(); -lean_mark_persistent(l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__1); -l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__2 = _init_l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__2(); -lean_mark_persistent(l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__2); -l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__3 = _init_l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__3(); -lean_mark_persistent(l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__3); -l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__4 = _init_l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__4(); -lean_mark_persistent(l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__4); -l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__5 = _init_l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__5(); -lean_mark_persistent(l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__5); -l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__6 = _init_l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__6(); -lean_mark_persistent(l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422____closed__6); -if (builtin) {res = l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_422_(lean_io_mk_world()); +l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__1 = _init_l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__1(); +lean_mark_persistent(l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__1); +l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__2 = _init_l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__2(); +lean_mark_persistent(l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__2); +l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__3 = _init_l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__3(); +lean_mark_persistent(l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__3); +l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__4 = _init_l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__4(); +lean_mark_persistent(l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__4); +l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__5 = _init_l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__5(); +lean_mark_persistent(l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__5); +l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__6 = _init_l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__6(); +lean_mark_persistent(l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358____closed__6); +if (builtin) {res = l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_358_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Server_Snapshots_server_stderrAsMessages = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Server_Snapshots_server_stderrAsMessages); diff --git a/stage0/stdlib/Lean/Server/Watchdog.c b/stage0/stdlib/Lean/Server/Watchdog.c index 7ed84238c73..9a973ea0c1d 100644 --- a/stage0/stdlib/Lean/Server/Watchdog.c +++ b/stage0/stdlib/Lean/Server/Watchdog.c @@ -30,7 +30,6 @@ lean_object* lean_get_stdout(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_Watchdog_tryWriteMessage___closed__52; extern lean_object* l_String_instInhabitedString; -static lean_object* l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__10; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_Watchdog_loadReferences___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*); static lean_object* l_Lean_Server_Watchdog_tryWriteMessage___closed__29; static lean_object* l_Lean_Server_Watchdog_tryWriteMessage___closed__28; @@ -71,7 +70,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_forwardNotification___at_Lean_Se LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_Watchdog_mainLoop___spec__4(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_Watchdog_initAndRunWatchdog___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_terminateFileWorker___boxed(lean_object*, lean_object*, lean_object*); -lean_object* lean_nat_log2(lean_object*); +double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); static lean_object* l_Lean_Server_Watchdog_startFileWorker___closed__3; static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_Watchdog_initAndRunWatchdog___spec__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_parseParams___at_Lean_Server_Watchdog_forwardRequestToWorker___spec__4(lean_object*, lean_object*, lean_object*); @@ -131,7 +130,6 @@ static lean_object* l_Lean_Server_Watchdog_mkLeanServerCapabilities___closed__4; static lean_object* l___private_Lean_Server_Watchdog_0__Lean_Server_Watchdog_forwardMessages_loop___closed__3; LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_loadReferences(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_findFileWorker_x3f___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__7; lean_object* l___private_Lean_Data_Lsp_Workspace_0__Lean_Lsp_toJsonDidChangeWatchedFilesRegistrationOptions____x40_Lean_Data_Lsp_Workspace___hyg_287_(lean_object*); lean_object* l_Lean_Server_References_definitionsMatching___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_initAndRunWatchdogAux(lean_object*, lean_object*); @@ -170,7 +168,6 @@ LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_handleEdits(lean_object*, lean_o LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_Watchdog_handleReference___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_Watchdog_mainLoop___closed__4; static lean_object* l_Lean_Json_toStructured_x3f___at_Lean_Server_Watchdog_initAndRunWatchdog___spec__5___closed__3; -double l_Float_ofBinaryScientific(lean_object*, lean_object*); lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_Lsp_instFromJsonInitializeParams___spec__6(lean_object*, lean_object*); static lean_object* l_Lean_Server_Watchdog_startFileWorker___closed__10; static lean_object* l_Lean_Server_Watchdog_initAndRunWatchdog___closed__6; @@ -202,7 +199,6 @@ uint8_t l_instDecidableNot___rarg(uint8_t); static lean_object* l___private_Lean_Server_Watchdog_0__Lean_Server_Watchdog_forwardMessages_loop___closed__9; static lean_object* l_Lean_Server_Watchdog_runClientTask___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_handleDidClose(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__11; static lean_object* l_Lean_Server_Watchdog_findWorkerPath___closed__3; lean_object* l_Lean_Json_compress(lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); @@ -212,7 +208,6 @@ lean_object* l_EStateM_bind___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_findWorkerPath___lambda__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_Lsp_instFromJsonInitializeParams___spec__1(lean_object*, lean_object*); extern lean_object* l_Task_Priority_dedicated; -lean_object* lean_int_mul(lean_object*, lean_object*); lean_object* l_Lean_Server_routeLspRequest(lean_object*, lean_object*, lean_object*); extern lean_object* l_instInhabitedFloat; lean_object* lean_nat_sub(lean_object*, lean_object*); @@ -254,7 +249,6 @@ LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_terminateFileWorker___lambda__1( static lean_object* l_Lean_Server_Watchdog_initAndRunWatchdog___closed__19; lean_object* l_Std_RBNode_appendTrees___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Server_Watchdog_tryWriteMessage___closed__46; -static double l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__12; lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcKeepAliveParams____x40_Lean_Data_Lsp_Extra___hyg_1605_(lean_object*); LEAN_EXPORT lean_object* l_Std_RBNode_find___at_Lean_Server_Watchdog_handleCancelRequest___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Server_Watchdog_parseParams___rarg___closed__1; @@ -312,15 +306,13 @@ LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_forwardNotification___at_Lean_Se LEAN_EXPORT lean_object* l_Lean_Json_toStructured_x3f___at_Lean_Server_Watchdog_handleCancelRequest___spec__2(lean_object*); static lean_object* l_Lean_Server_Watchdog_tryWriteMessage___closed__64; static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_Watchdog_initAndRunWatchdog___spec__2___closed__3; -static lean_object* l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__8; lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_110_(lean_object*); uint8_t l_Array_isEmpty___rarg(lean_object*); static lean_object* l_Lean_Server_Watchdog_mkLeanServerCapabilities___closed__6; -static lean_object* l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__2; lean_object* lean_get_prefix(lean_object*); static lean_object* l_Lean_Server_Watchdog_tryWriteMessage___closed__11; lean_object* l_Lean_FileMap_ofString(lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3088_(lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3087_(lean_object*); lean_object* l_Std_RBNode_balLeft___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_FileWorker_errorPendingRequests___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_mainLoop(lean_object*, lean_object*, lean_object*); @@ -366,7 +358,6 @@ static lean_object* l_IO_FS_Stream_readLspNotificationAs___at_Lean_Server_Watchd static lean_object* l_Std_RBNode_forIn_visit___at_Lean_Server_Watchdog_handleCancelRequest___spec__3___closed__1; lean_object* lean_int_neg(lean_object*); static lean_object* l_Lean_Server_Watchdog_tryWriteMessage___closed__44; -static lean_object* l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__3; static lean_object* l_Lean_Server_Watchdog_initAndRunWatchdog___closed__10; static lean_object* l_Lean_Server_Watchdog_mainLoop___closed__6; static lean_object* l___private_Lean_Server_Watchdog_0__Lean_Server_Watchdog_forwardMessages_loop___closed__8; @@ -377,7 +368,6 @@ static lean_object* l___private_Lean_Server_Watchdog_0__Lean_Server_Watchdog_for static lean_object* l_Lean_Server_Watchdog_findWorkerPath___lambda__2___closed__2; static lean_object* l_Lean_Server_Watchdog_tryWriteMessage___closed__1; static lean_object* l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__2___closed__2; -static lean_object* l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__9; LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_shutdown(lean_object*, lean_object*); static lean_object* l_Lean_Server_Watchdog_tryWriteMessage___closed__4; LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_findWorkerPath___lambda__2(lean_object*, lean_object*, lean_object*); @@ -410,7 +400,6 @@ static lean_object* l_Lean_Server_Watchdog_initAndRunWatchdog___closed__13; static lean_object* l_Lean_Server_Watchdog_tryWriteMessage___closed__45; LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_handleDidChangeWatchedFiles___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_Watchdog_tryWriteMessage___closed__40; -static lean_object* l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__4; static lean_object* l_Lean_Server_Watchdog_handleNotification___closed__5; static lean_object* l_Lean_Server_Watchdog_tryWriteMessage___closed__9; uint8_t l_String_isPrefixOf(lean_object*, lean_object*); @@ -430,7 +419,7 @@ LEAN_EXPORT lean_object* l_Lean_Json_toStructured_x3f___at_Lean_Server_Watchdog_ LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_findDefinitions(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_Watchdog_tryWriteMessage___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_handleRequest(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__1; +static double l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Std_RBNode_forIn_visit___at_Lean_Server_Watchdog_shutdown___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_Watchdog_startFileWorker___closed__7; static lean_object* l_Lean_Server_Watchdog_tryWriteMessage___closed__58; @@ -451,7 +440,6 @@ static lean_object* l_Lean_Server_Watchdog_FileWorker_runEditsSignalTask_loopAct lean_object* l_Lean_Json_mkObj(lean_object*); LEAN_EXPORT lean_object* l_Std_RBNode_find___at_Lean_Server_Watchdog_findFileWorker_x3f___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_FileWorker_stdin(lean_object*); -lean_object* lean_nat_shiftl(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Stream_writeLspResponse___at_Lean_Server_Watchdog_initAndRunWatchdog___spec__3(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_Watchdog_handleEdits___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_findFileWorker_x3f(lean_object*, lean_object*, lean_object*); @@ -474,7 +462,6 @@ static lean_object* l_Lean_Server_Watchdog_mkLeanServerCapabilities___closed__5; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_Watchdog_mainLoop___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Server_foldDocumentChanges(lean_object*, lean_object*); lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidOpenTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_102_(lean_object*); -static lean_object* l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__5; lean_object* l___private_Lean_Data_Lsp_Client_0__Lean_Lsp_toJsonRegistrationParams____x40_Lean_Data_Lsp_Client___hyg_165_(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_Watchdog_mainLoop___spec__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_handleIleanInfoUpdate___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -492,7 +479,6 @@ LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_handleDidChangeWatchedFiles(lean static lean_object* l_Lean_Server_Watchdog_tryWriteMessage___closed__19; LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_findFileWorker_x21___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_Watchdog_tryWriteMessage___closed__51; -lean_object* lean_int_sub(lean_object*, lean_object*); static lean_object* l_Lean_Server_Watchdog_mkLeanServerCapabilities___closed__3; LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_forwardNotification___at_Lean_Server_Watchdog_handleNotification___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_get_stderr(lean_object*); @@ -545,7 +531,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_handleWorkspaceSymbol(lean_objec LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_Watchdog_handleWorkspaceSymbol___spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readLspRequestAs___at_Lean_Server_Watchdog_initAndRunWatchdog___spec__1___closed__1; static lean_object* l_Lean_Server_Watchdog_findFileWorker_x21___closed__2; -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2256_(lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2255_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_parseParams___at_Lean_Server_Watchdog_handleNotification___spec__1(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Server_References_definitionOf_x3f(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_to_int(lean_object*); @@ -553,10 +539,9 @@ static lean_object* l_Lean_Server_Watchdog_startFileWorker___closed__6; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_Watchdog_handleRequest___spec__6(size_t, size_t, lean_object*); lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_778_(lean_object*); lean_object* l_List_appendTR___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__6; static lean_object* l_Lean_Server_Watchdog_forwardRequestToWorker___closed__1; static uint8_t l_Lean_Server_Watchdog_tryWriteMessage___lambda__1___closed__1; -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2387_(lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2386_(lean_object*); LEAN_EXPORT lean_object* l_Std_RBNode_ins___at_Lean_Server_Watchdog_forwardRequestToWorker___spec__3(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_Watchdog_handleNotification___closed__3; lean_object* l___private_Lean_Data_Lsp_InitShutdown_0__Lean_Lsp_toJsonInitializeResult____x40_Lean_Data_Lsp_InitShutdown___hyg_666_(lean_object*); @@ -9023,119 +9008,15 @@ goto _start; } } } -static lean_object* _init_l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__1() { +static double _init_l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(2u); -x_2 = lean_nat_log2(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(64u); -x_2 = l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__1; -x_3 = lean_nat_sub(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(3u); -x_2 = l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__2; -x_3 = lean_nat_add(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; +lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; x_1 = lean_unsigned_to_nat(2u); -x_2 = l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__3; -x_3 = lean_nat_shiftl(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__4; -x_2 = lean_unsigned_to_nat(5u); -x_3 = lean_nat_div(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__6() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(4u); -x_2 = lean_nat_to_int(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__7() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__6; -x_2 = lean_int_neg(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(1u); -x_2 = lean_nat_to_int(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__9() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__7; -x_2 = l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__8; -x_3 = lean_int_mul(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__10() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__2; -x_2 = lean_nat_to_int(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__11() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__9; -x_2 = l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__10; -x_3 = lean_int_sub(x_1, x_2); -return x_3; -} -} -static double _init_l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__12() { -_start: -{ -lean_object* x_1; lean_object* x_2; double x_3; -x_1 = l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__5; -x_2 = l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__11; -x_3 = l_Float_ofBinaryScientific(x_1, x_2); -return x_3; +x_2 = 1; +x_3 = lean_unsigned_to_nat(1u); +x_4 = l_Float_ofScientific(x_1, x_2, x_3); +return x_4; } } LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1(lean_object* x_1, lean_object* x_2) { @@ -9149,7 +9030,7 @@ if (lean_obj_tag(x_3) == 0) uint8_t x_4; lean_object* x_5; double x_6; lean_object* x_7; x_4 = 1; x_5 = l_Lean_Name_toString(x_2, x_4); -x_6 = l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__12; +x_6 = l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__1; lean_inc(x_5); x_7 = l_Lean_FuzzyMatching_fuzzyMatchScoreWithThreshold_x3f(x_1, x_5, x_6); if (lean_obj_tag(x_7) == 0) @@ -9197,7 +9078,7 @@ lean_inc(x_15); lean_dec(x_3); x_16 = 1; x_17 = l_Lean_Name_toString(x_15, x_16); -x_18 = l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__12; +x_18 = l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__1; lean_inc(x_17); x_19 = l_Lean_FuzzyMatching_fuzzyMatchScoreWithThreshold_x3f(x_1, x_17, x_18); if (lean_obj_tag(x_19) == 0) @@ -14296,7 +14177,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_parseParams___at_Lean_Server_Wat _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2387_(x_1); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2386_(x_1); if (lean_obj_tag(x_4) == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; 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; @@ -14345,7 +14226,7 @@ lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x x_5 = lean_array_uget(x_3, x_2); x_6 = lean_unsigned_to_nat(0u); x_7 = lean_array_uset(x_3, x_2, x_6); -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3088_(x_5); +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3087_(x_5); x_9 = 1; x_10 = lean_usize_add(x_2, x_9); x_11 = lean_array_uset(x_7, x_2, x_8); @@ -14383,7 +14264,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_parseParams___at_Lean_Server_Wat _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2256_(x_1); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2255_(x_1); if (lean_obj_tag(x_4) == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; 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; @@ -24259,28 +24140,6 @@ lean_mark_persistent(l_Lean_Server_Watchdog_findDefinitions___closed__1); l_Array_qsort_sort___at_Lean_Server_Watchdog_handleWorkspaceSymbol___spec__1___closed__1 = _init_l_Array_qsort_sort___at_Lean_Server_Watchdog_handleWorkspaceSymbol___spec__1___closed__1(); lean_mark_persistent(l_Array_qsort_sort___at_Lean_Server_Watchdog_handleWorkspaceSymbol___spec__1___closed__1); l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__1 = _init_l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__1); -l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__2 = _init_l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__2(); -lean_mark_persistent(l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__2); -l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__3 = _init_l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__3(); -lean_mark_persistent(l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__3); -l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__4 = _init_l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__4(); -lean_mark_persistent(l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__4); -l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__5 = _init_l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__5(); -lean_mark_persistent(l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__5); -l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__6 = _init_l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__6(); -lean_mark_persistent(l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__6); -l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__7 = _init_l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__7(); -lean_mark_persistent(l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__7); -l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__8 = _init_l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__8(); -lean_mark_persistent(l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__8); -l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__9 = _init_l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__9(); -lean_mark_persistent(l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__9); -l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__10 = _init_l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__10(); -lean_mark_persistent(l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__10); -l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__11 = _init_l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__11(); -lean_mark_persistent(l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__11); -l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__12 = _init_l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__1___closed__12(); l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__2___closed__1 = _init_l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__2___closed__1(); lean_mark_persistent(l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__2___closed__1); l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__2___closed__2 = _init_l_Lean_Server_Watchdog_handleWorkspaceSymbol___lambda__2___closed__2(); diff --git a/stage0/stdlib/Lean/SubExpr.c b/stage0/stdlib/Lean/SubExpr.c index 3dc4d731f5a..d9320748e13 100644 --- a/stage0/stdlib/Lean/SubExpr.c +++ b/stage0/stdlib/Lean/SubExpr.c @@ -89,7 +89,6 @@ static lean_object* l___private_Lean_SubExpr_0__Lean_SubExpr_Pos_ofStringCoord__ lean_object* l_Nat_repr(lean_object*); LEAN_EXPORT uint8_t l_Lean_SubExpr_Pos_all(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SubExpr_Pos_foldlM___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); -lean_object* l_panic___at_String_toNat_x21___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_SubExpr_Pos_instDecidableEqPos___boxed(lean_object*, lean_object*); lean_object* l_List_toString___at_Lean_MetavarContext_MkBinding_instToStringException___spec__2(lean_object*); lean_object* l_Function_comp___rarg(lean_object*, lean_object*, lean_object*); @@ -137,6 +136,7 @@ lean_object* lean_nat_mul(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_SubExpr_Pos_fromString_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SubExpr_Pos_pushAppArg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_SubExpr_Pos_foldl___rarg___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_panic___at_Lean_TSyntax_getNat___spec__1(lean_object*); lean_object* l_List_toArrayAux___rarg(lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedSubExpr___closed__3; static lean_object* l___private_Lean_SubExpr_0__Lean_SubExpr_Pos_ofStringCoord___closed__7; @@ -295,7 +295,7 @@ else { lean_object* x_5; lean_object* x_6; x_5 = l_Lean_SubExpr_Pos_head___closed__4; -x_6 = l_panic___at_String_toNat_x21___spec__1(x_5); +x_6 = l_panic___at_Lean_TSyntax_getNat___spec__1(x_5); return x_6; } } diff --git a/stage0/stdlib/Lean/Syntax.c b/stage0/stdlib/Lean/Syntax.c index 64f7b92940f..481a9b0131d 100644 --- a/stage0/stdlib/Lean/Syntax.c +++ b/stage0/stdlib/Lean/Syntax.c @@ -35,7 +35,6 @@ lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_instForInTopDownSyntax_loop___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_getAntiquotSpliceContents___boxed(lean_object*); static lean_object* l_Lean_Syntax_mkAntiquotNode___closed__12; -extern lean_object* l_Lean_nullKind; LEAN_EXPORT lean_object* l_Lean_Syntax_getAntiquotSuffixSpliceInner(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_rewriteBottomUpM___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_isLitKind(lean_object*); @@ -44,6 +43,7 @@ lean_object* lean_name_mk_string(lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Syntax_replaceM___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); +static lean_object* l_Lean_Syntax_mkAntiquotNode___closed__24; LEAN_EXPORT lean_object* l_Lean_Syntax_ifNodeKind___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Syntax_reprint___spec__3___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Syntax_mkAntiquotSuffixSpliceNode___closed__1; @@ -66,6 +66,7 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_isEscapedAntiquot___boxed(lean_object*); static lean_object* l_Lean_Syntax_getAtomVal_x21___closed__2; LEAN_EXPORT lean_object* l_Lean_Syntax_MonadTraverser_getCur(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_MonadTraverser_goRight(lean_object*); +static lean_object* l_Lean_isLitKind___closed__7; static lean_object* l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Syntax_hasMissing___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_SyntaxNode_withArgs(lean_object*); static lean_object* l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Syntax_hasMissing___spec__1___closed__4; @@ -81,6 +82,7 @@ extern lean_object* l_instInhabitedNat; LEAN_EXPORT lean_object* l_List_mapTRAux___at_Lean_Syntax_identComponents___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Syntax_mkAntiquotNode___closed__18; static lean_object* l_Lean_Syntax_getAtomVal_x21___closed__1; +static lean_object* l_Lean_isLitKind___closed__10; lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Syntax_reprint___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_ifNodeKind(lean_object*); @@ -111,10 +113,10 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_getAntiquotTerm___boxed(lean_object*); static lean_object* l_Lean_Syntax_mkAntiquotSpliceNode___closed__2; LEAN_EXPORT lean_object* l_Lean_Syntax_instForInTopDownSyntax_loop___rarg___lambda__6(lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); -extern lean_object* l_Lean_nameLitKind; LEAN_EXPORT lean_object* l_Lean_Syntax_MonadTraverser_goDown(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_reprint_reprintLeaf___boxed(lean_object*, lean_object*); lean_object* l_Array_back_x3f___rarg(lean_object*); +static lean_object* l_Lean_isLitKind___closed__5; LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Syntax_identComponents_nameComps___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_mkAntiquotNode(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Syntax_hasMissing(lean_object*); @@ -143,13 +145,13 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_replaceM___rarg___lambda__1(lean_object*, lean_object* l_Array_forInUnsafe_loop___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Syntax_identComponents___spec__3___boxed(lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); +static lean_object* l_Lean_isLitKind___closed__3; LEAN_EXPORT lean_object* l_Lean_SyntaxNode_getKind___boxed(lean_object*); LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at_Lean_Syntax_reprint___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_Syntax_isQuot___closed__3; LEAN_EXPORT lean_object* l_Lean_Syntax_updateTrailing(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Syntax_reprint___spec__4(uint8_t, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -extern lean_object* l_Lean_numLitKind; lean_object* lean_nat_sub(lean_object*, lean_object*); static lean_object* l_Lean_Syntax_identComponents___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_mkAntiquotSpliceNode___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -157,11 +159,11 @@ LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Syntax_identComponents___spec__3 lean_object* l_List_zip___rarg___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_instForInTopDownSyntax_loop___rarg___lambda__4(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Syntax_isQuot___closed__5; -extern lean_object* l_Lean_strLitKind; LEAN_EXPORT lean_object* l_Lean_Syntax_MonadTraverser_getIdx(lean_object*); static lean_object* l_Lean_Syntax_identComponents___closed__8; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Syntax_updateLeading___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Syntax_mkAntiquotSpliceNode___closed__1; +static lean_object* l_Lean_isLitKind___closed__6; LEAN_EXPORT lean_object* l_Lean_Syntax_MonadTraverser_goLeft___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_MonadTraverser_goDown___rarg(lean_object*, lean_object*); lean_object* l_Lean_Name_toString(lean_object*, uint8_t); @@ -170,6 +172,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Syntax_updateLeading__ LEAN_EXPORT lean_object* l_Lean_unreachIsNodeMissing(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_replaceM(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_getIdAt___boxed(lean_object*, lean_object*); +lean_object* l_panic___at_Lean_TSyntax_getString___spec__1(lean_object*); static lean_object* l_Lean_Syntax_mkAntiquotSpliceNode___closed__8; lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_updateLeading(lean_object*); @@ -181,9 +184,7 @@ static lean_object* l_Lean_Syntax_mkAntiquotNode___closed__14; LEAN_EXPORT lean_object* l_Lean_Syntax_reprint_reprintLeaf(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Syntax_0__Lean_Syntax_updateInfo(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getId(lean_object*); -extern lean_object* l_Lean_choiceKind; static lean_object* l_Lean_Syntax_mkAntiquotNode___closed__15; -extern lean_object* l_Lean_charLitKind; LEAN_EXPORT lean_object* l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Syntax_hasMissing___spec__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); lean_object* l_Array_back___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_MonadTraverser_setCur___rarg(lean_object*, lean_object*); @@ -201,6 +202,7 @@ LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at_Lean_Syntax_reprint___ LEAN_EXPORT lean_object* l_Lean_SyntaxNode_modifyArgs(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_hasMissing___lambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_SyntaxNode_getArgs___boxed(lean_object*); +static lean_object* l_Lean_isLitKind___closed__2; lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Array_isEmpty___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_unreachIsNodeIdent___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -226,11 +228,14 @@ static lean_object* l_Lean_Syntax_mkAntiquotSpliceNode___closed__5; LEAN_EXPORT lean_object* l_Lean_SyntaxNode_getNumArgs___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Syntax_hasMissing___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isAtom(lean_object*); +static lean_object* l_Lean_isLitKind___closed__4; LEAN_EXPORT lean_object* l_Lean_Syntax_replaceM___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Array_mapMUnsafe_map___rarg(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_Traverser_fromSyntax(lean_object*); LEAN_EXPORT lean_object* l_Lean_unreachIsNodeAtom(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Syntax_asNode___closed__4; LEAN_EXPORT lean_object* l_Lean_Syntax_mkAntiquotNode___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_isLitKind___closed__9; lean_object* l_List_drop___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Syntax_mkAntiquotNode___closed__1; uint8_t l_Lean_Syntax_isMissing(lean_object*); @@ -244,7 +249,6 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_instForInTopDownSyntax___rarg(lean_object static lean_object* l_Lean_Syntax_mkAntiquotNode___closed__8; lean_object* l_Lean_Syntax_getNumArgs(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_Traverser_down(lean_object*, lean_object*); -extern lean_object* l_Lean_scientificLitKind; LEAN_EXPORT lean_object* l_Lean_Syntax_MonadTraverser_getIdx___rarg___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_antiquotKinds___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_setArg(lean_object*, lean_object*, lean_object*); @@ -252,6 +256,7 @@ LEAN_EXPORT lean_object* l_Lean_SourceInfo_updateTrailing(lean_object*, lean_obj LEAN_EXPORT lean_object* l_Lean_Syntax_rewriteBottomUpM___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Syntax_getQuotContent___closed__3; uint8_t lean_nat_dec_le(lean_object*, lean_object*); +static lean_object* l_Lean_isLitKind___closed__1; static lean_object* l_Lean_Syntax_asNode___closed__1; lean_object* l_Lean_Syntax_getArgs(lean_object*); lean_object* l_Lean_Name_append(lean_object*, lean_object*); @@ -280,6 +285,7 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_instForInTopDownSyntax(lean_object*, lean static lean_object* l_Lean_Syntax_isTokenAntiquot___closed__2; LEAN_EXPORT lean_object* l_Lean_Syntax_MonadTraverser_goDown___rarg___lambda__1(lean_object*, lean_object*); static lean_object* l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Syntax_hasMissing___spec__1___closed__2; +static lean_object* l_Lean_isLitKind___closed__8; LEAN_EXPORT lean_object* l_Lean_isLitKind___boxed(lean_object*); static lean_object* l_Lean_Syntax_mkAntiquotNode___closed__13; lean_object* lean_panic_fn(lean_object*, lean_object*); @@ -298,6 +304,7 @@ LEAN_EXPORT lean_object* l_panic___at_Lean_Syntax_identComponents___spec__1(lean LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Syntax_isAntiquots___spec__1(lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Syntax_instForInTopDownSyntax_loop___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Syntax_hasMissing___closed__2; +static lean_object* l_Lean_Syntax_asNode___closed__3; static lean_object* l_Lean_Syntax_mkAntiquotNode___closed__6; LEAN_EXPORT lean_object* l_Array_findSomeRevM_x3f_find___at_Lean_Syntax_getTailWithPos___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Syntax_isEscapedAntiquot(lean_object*); @@ -316,11 +323,12 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_asNode(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_antiquotSpliceKind_x3f(lean_object*); LEAN_EXPORT lean_object* l_Lean_unreachIsNodeIdent(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_instForInTopDownSyntax_loop___rarg___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Syntax_instForInTopDownSyntax_loop___rarg___lambda__7___closed__2; LEAN_EXPORT lean_object* l_Lean_Syntax_antiquotKinds(lean_object*); -lean_object* l_panic___at_Lean_Name_getString_x21___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_getAtomVal_x21(lean_object*); lean_object* l_Nat_min(lean_object*, lean_object*); static lean_object* l_Lean_Syntax_isQuot___closed__2; +static lean_object* l_Lean_Syntax_instForInTopDownSyntax_loop___rarg___lambda__7___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_ifNodeKind___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Syntax_mkAntiquotNode___closed__21; static lean_object* l_Lean_Syntax_identComponents___closed__6; @@ -426,31 +434,121 @@ lean_dec(x_2); return x_7; } } +static lean_object* _init_l_Lean_isLitKind___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("str", 3); +return x_1; +} +} +static lean_object* _init_l_Lean_isLitKind___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_isLitKind___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_isLitKind___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("num", 3); +return x_1; +} +} +static lean_object* _init_l_Lean_isLitKind___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_isLitKind___closed__3; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_isLitKind___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("char", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_isLitKind___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_isLitKind___closed__5; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_isLitKind___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("name", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_isLitKind___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_isLitKind___closed__7; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_isLitKind___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("scientific", 10); +return x_1; +} +} +static lean_object* _init_l_Lean_isLitKind___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_isLitKind___closed__9; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT uint8_t l_Lean_isLitKind(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; -x_2 = l_Lean_strLitKind; +x_2 = l_Lean_isLitKind___closed__2; x_3 = lean_name_eq(x_1, x_2); if (x_3 == 0) { lean_object* x_4; uint8_t x_5; -x_4 = l_Lean_numLitKind; +x_4 = l_Lean_isLitKind___closed__4; x_5 = lean_name_eq(x_1, x_4); if (x_5 == 0) { lean_object* x_6; uint8_t x_7; -x_6 = l_Lean_charLitKind; +x_6 = l_Lean_isLitKind___closed__6; x_7 = lean_name_eq(x_1, x_6); if (x_7 == 0) { lean_object* x_8; uint8_t x_9; -x_8 = l_Lean_nameLitKind; +x_8 = l_Lean_isLitKind___closed__8; x_9 = lean_name_eq(x_1, x_8); if (x_9 == 0) { lean_object* x_10; uint8_t x_11; -x_10 = l_Lean_scientificLitKind; +x_10 = l_Lean_isLitKind___closed__10; x_11 = lean_name_eq(x_1, x_10); return x_11; } @@ -669,7 +767,7 @@ else { lean_object* x_3; lean_object* x_4; x_3 = l_Lean_Syntax_getAtomVal_x21___closed__4; -x_4 = l_panic___at_Lean_Name_getString_x21___spec__1(x_3); +x_4 = l_panic___at_Lean_TSyntax_getString___spec__1(x_3); return x_4; } } @@ -877,9 +975,27 @@ return x_2; static lean_object* _init_l_Lean_Syntax_asNode___closed__2() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("null", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Syntax_asNode___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Syntax_asNode___closed__2; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Syntax_asNode___closed__4() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean_nullKind; +x_2 = l_Lean_Syntax_asNode___closed__3; x_3 = l_Lean_Syntax_asNode___closed__1; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -920,7 +1036,7 @@ else { lean_object* x_7; lean_dec(x_1); -x_7 = l_Lean_Syntax_asNode___closed__2; +x_7 = l_Lean_Syntax_asNode___closed__4; return x_7; } } @@ -3323,6 +3439,24 @@ x_5 = lean_apply_2(x_4, lean_box(0), x_2); return x_5; } } +static lean_object* _init_l_Lean_Syntax_instForInTopDownSyntax_loop___rarg___lambda__7___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("choice", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Syntax_instForInTopDownSyntax_loop___rarg___lambda__7___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Syntax_instForInTopDownSyntax_loop___rarg___lambda__7___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Syntax_instForInTopDownSyntax_loop___rarg___lambda__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7) { _start: { @@ -3391,7 +3525,7 @@ goto block_32; else { lean_object* x_34; uint8_t x_35; -x_34 = l_Lean_choiceKind; +x_34 = l_Lean_Syntax_instForInTopDownSyntax_loop___rarg___lambda__7___closed__2; x_35 = lean_name_eq(x_18, x_34); lean_dec(x_18); if (x_35 == 0) @@ -4018,7 +4152,7 @@ x_43 = lean_ctor_get(x_2, 1); lean_inc(x_43); x_44 = lean_ctor_get(x_2, 2); lean_inc(x_44); -x_45 = l_Lean_choiceKind; +x_45 = l_Lean_Syntax_instForInTopDownSyntax_loop___rarg___lambda__7___closed__2; x_46 = lean_name_eq(x_43, x_45); lean_dec(x_43); if (x_46 == 0) @@ -4159,7 +4293,7 @@ goto block_26; else { lean_object* x_28; uint8_t x_29; -x_28 = l_Lean_choiceKind; +x_28 = l_Lean_Syntax_instForInTopDownSyntax_loop___rarg___lambda__7___closed__2; x_29 = lean_name_eq(x_8, x_28); lean_dec(x_8); if (x_29 == 0) @@ -4588,7 +4722,7 @@ goto block_22; else { lean_object* x_24; uint8_t x_25; -x_24 = l_Lean_choiceKind; +x_24 = l_Lean_Syntax_instForInTopDownSyntax_loop___rarg___lambda__7___closed__2; x_25 = lean_name_eq(x_8, x_24); lean_dec(x_8); if (x_25 == 0) @@ -5328,7 +5462,7 @@ LEAN_EXPORT lean_object* l_Lean_mkListNode(lean_object* x_1) { { lean_object* x_2; lean_object* x_3; lean_object* x_4; x_2 = lean_box(2); -x_3 = l_Lean_nullKind; +x_3 = l_Lean_Syntax_asNode___closed__3; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_2); lean_ctor_set(x_4, 1, x_3); @@ -5709,7 +5843,7 @@ x_2 = l_Lean_Syntax_isAntiquot(x_1); if (x_2 == 0) { lean_object* x_3; uint8_t x_4; -x_3 = l_Lean_choiceKind; +x_3 = l_Lean_Syntax_instForInTopDownSyntax_loop___rarg___lambda__7___closed__2; lean_inc(x_1); x_4 = l_Lean_Syntax_isOfKind(x_1, x_3); if (x_4 == 0) @@ -5821,7 +5955,7 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_getCanonicalAntiquot(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; -x_2 = l_Lean_choiceKind; +x_2 = l_Lean_Syntax_instForInTopDownSyntax_loop___rarg___lambda__7___closed__2; lean_inc(x_1); x_3 = l_Lean_Syntax_isOfKind(x_1, x_2); if (x_3 == 0) @@ -5890,22 +6024,36 @@ return x_3; static lean_object* _init_l_Lean_Syntax_mkAntiquotNode___closed__6() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(2); +x_2 = l_Lean_Syntax_asNode___closed__3; +x_3 = l_Lean_Syntax_asNode___closed__1; +x_4 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Syntax_mkAntiquotNode___closed__7() { +_start: +{ lean_object* x_1; x_1 = lean_mk_string_from_bytes("pseudo", 6); return x_1; } } -static lean_object* _init_l_Lean_Syntax_mkAntiquotNode___closed__7() { +static lean_object* _init_l_Lean_Syntax_mkAntiquotNode___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Syntax_mkAntiquotNode___closed__6; +x_2 = l_Lean_Syntax_mkAntiquotNode___closed__7; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Syntax_mkAntiquotNode___closed__8() { +static lean_object* _init_l_Lean_Syntax_mkAntiquotNode___closed__9() { _start: { lean_object* x_1; @@ -5913,17 +6061,17 @@ x_1 = lean_mk_string_from_bytes("antiquotName", 12); return x_1; } } -static lean_object* _init_l_Lean_Syntax_mkAntiquotNode___closed__9() { +static lean_object* _init_l_Lean_Syntax_mkAntiquotNode___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Syntax_mkAntiquotNode___closed__8; +x_2 = l_Lean_Syntax_mkAntiquotNode___closed__9; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Syntax_mkAntiquotNode___closed__10() { +static lean_object* _init_l_Lean_Syntax_mkAntiquotNode___closed__11() { _start: { lean_object* x_1; @@ -5931,19 +6079,19 @@ x_1 = lean_mk_string_from_bytes(":", 1); return x_1; } } -static lean_object* _init_l_Lean_Syntax_mkAntiquotNode___closed__11() { +static lean_object* _init_l_Lean_Syntax_mkAntiquotNode___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(2); -x_2 = l_Lean_Syntax_mkAntiquotNode___closed__10; +x_2 = l_Lean_Syntax_mkAntiquotNode___closed__11; x_3 = lean_alloc_ctor(2, 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_Lean_Syntax_mkAntiquotNode___closed__12() { +static lean_object* _init_l_Lean_Syntax_mkAntiquotNode___closed__13() { _start: { lean_object* x_1; lean_object* x_2; @@ -5952,17 +6100,17 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Syntax_mkAntiquotNode___closed__13() { +static lean_object* _init_l_Lean_Syntax_mkAntiquotNode___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Syntax_mkAntiquotNode___closed__12; -x_2 = l_Lean_Syntax_mkAntiquotNode___closed__11; +x_1 = l_Lean_Syntax_mkAntiquotNode___closed__13; +x_2 = l_Lean_Syntax_mkAntiquotNode___closed__12; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Syntax_mkAntiquotNode___closed__14() { +static lean_object* _init_l_Lean_Syntax_mkAntiquotNode___closed__15() { _start: { lean_object* x_1; @@ -5970,17 +6118,17 @@ x_1 = lean_mk_string_from_bytes("hole", 4); return x_1; } } -static lean_object* _init_l_Lean_Syntax_mkAntiquotNode___closed__15() { +static lean_object* _init_l_Lean_Syntax_mkAntiquotNode___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Syntax_getQuotContent___closed__3; -x_2 = l_Lean_Syntax_mkAntiquotNode___closed__14; +x_2 = l_Lean_Syntax_mkAntiquotNode___closed__15; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Syntax_mkAntiquotNode___closed__16() { +static lean_object* _init_l_Lean_Syntax_mkAntiquotNode___closed__17() { _start: { lean_object* x_1; @@ -5988,17 +6136,17 @@ x_1 = lean_mk_string_from_bytes("antiquotNestedExpr", 18); return x_1; } } -static lean_object* _init_l_Lean_Syntax_mkAntiquotNode___closed__17() { +static lean_object* _init_l_Lean_Syntax_mkAntiquotNode___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Syntax_mkAntiquotNode___closed__16; +x_2 = l_Lean_Syntax_mkAntiquotNode___closed__17; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Syntax_mkAntiquotNode___closed__18() { +static lean_object* _init_l_Lean_Syntax_mkAntiquotNode___closed__19() { _start: { lean_object* x_1; @@ -6006,19 +6154,19 @@ x_1 = lean_mk_string_from_bytes("(", 1); return x_1; } } -static lean_object* _init_l_Lean_Syntax_mkAntiquotNode___closed__19() { +static lean_object* _init_l_Lean_Syntax_mkAntiquotNode___closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(2); -x_2 = l_Lean_Syntax_mkAntiquotNode___closed__18; +x_2 = l_Lean_Syntax_mkAntiquotNode___closed__19; x_3 = lean_alloc_ctor(2, 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_Lean_Syntax_mkAntiquotNode___closed__20() { +static lean_object* _init_l_Lean_Syntax_mkAntiquotNode___closed__21() { _start: { lean_object* x_1; @@ -6026,19 +6174,19 @@ x_1 = lean_mk_string_from_bytes(")", 1); return x_1; } } -static lean_object* _init_l_Lean_Syntax_mkAntiquotNode___closed__21() { +static lean_object* _init_l_Lean_Syntax_mkAntiquotNode___closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(2); -x_2 = l_Lean_Syntax_mkAntiquotNode___closed__20; +x_2 = l_Lean_Syntax_mkAntiquotNode___closed__21; x_3 = lean_alloc_ctor(2, 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_Lean_Syntax_mkAntiquotNode___closed__22() { +static lean_object* _init_l_Lean_Syntax_mkAntiquotNode___closed__23() { _start: { lean_object* x_1; lean_object* x_2; @@ -6047,12 +6195,12 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Syntax_mkAntiquotNode___closed__23() { +static lean_object* _init_l_Lean_Syntax_mkAntiquotNode___closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Syntax_mkAntiquotNode___closed__22; -x_2 = l_Lean_Syntax_mkAntiquotNode___closed__19; +x_1 = l_Lean_Syntax_mkAntiquotNode___closed__23; +x_2 = l_Lean_Syntax_mkAntiquotNode___closed__20; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -6064,7 +6212,7 @@ lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_obj x_6 = l_Lean_Syntax_mkAntiquotNode___closed__2; x_7 = lean_mk_array(x_3, x_6); x_8 = lean_box(2); -x_9 = l_Lean_nullKind; +x_9 = l_Lean_Syntax_asNode___closed__3; x_10 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_10, 0, x_8); lean_ctor_set(x_10, 1, x_9); @@ -6075,17 +6223,17 @@ x_13 = lean_array_push(x_12, x_10); if (x_11 == 0) { lean_object* x_46; uint8_t x_47; -x_46 = l_Lean_Syntax_mkAntiquotNode___closed__15; +x_46 = l_Lean_Syntax_mkAntiquotNode___closed__16; lean_inc(x_2); x_47 = l_Lean_Syntax_isOfKind(x_2, x_46); if (x_47 == 0) { lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; -x_48 = l_Lean_Syntax_mkAntiquotNode___closed__23; +x_48 = l_Lean_Syntax_mkAntiquotNode___closed__24; x_49 = lean_array_push(x_48, x_2); -x_50 = l_Lean_Syntax_mkAntiquotNode___closed__21; +x_50 = l_Lean_Syntax_mkAntiquotNode___closed__22; x_51 = lean_array_push(x_49, x_50); -x_52 = l_Lean_Syntax_mkAntiquotNode___closed__17; +x_52 = l_Lean_Syntax_mkAntiquotNode___closed__18; x_53 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_53, 0, x_8); lean_ctor_set(x_53, 1, x_52); @@ -6115,7 +6263,7 @@ x_15 = lean_array_push(x_13, x_14); if (lean_obj_tag(x_4) == 0) { lean_object* x_16; lean_object* x_17; -x_16 = l_Lean_Syntax_asNode___closed__2; +x_16 = l_Lean_Syntax_mkAntiquotNode___closed__6; x_17 = lean_array_push(x_15, x_16); if (x_5 == 0) { @@ -6134,7 +6282,7 @@ return x_22; else { lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_23 = l_Lean_Syntax_mkAntiquotNode___closed__7; +x_23 = l_Lean_Syntax_mkAntiquotNode___closed__8; x_24 = l_Lean_Name_append(x_1, x_23); x_25 = l_Lean_Syntax_mkAntiquotNode___closed__3; x_26 = l_Lean_Name_append(x_24, x_25); @@ -6154,9 +6302,9 @@ lean_inc(x_28); x_29 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_29, 0, x_8); lean_ctor_set(x_29, 1, x_28); -x_30 = l_Lean_Syntax_mkAntiquotNode___closed__13; +x_30 = l_Lean_Syntax_mkAntiquotNode___closed__14; x_31 = lean_array_push(x_30, x_29); -x_32 = l_Lean_Syntax_mkAntiquotNode___closed__9; +x_32 = l_Lean_Syntax_mkAntiquotNode___closed__10; x_33 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_33, 0, x_8); lean_ctor_set(x_33, 1, x_32); @@ -6179,7 +6327,7 @@ return x_39; else { lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_40 = l_Lean_Syntax_mkAntiquotNode___closed__7; +x_40 = l_Lean_Syntax_mkAntiquotNode___closed__8; x_41 = l_Lean_Name_append(x_1, x_40); x_42 = l_Lean_Syntax_mkAntiquotNode___closed__3; x_43 = l_Lean_Name_append(x_41, x_42); @@ -6258,7 +6406,7 @@ x_5 = l_Lean_Syntax_getArgs(x_4); lean_dec(x_4); x_6 = lean_array_pop(x_5); x_7 = lean_box(2); -x_8 = l_Lean_nullKind; +x_8 = l_Lean_Syntax_asNode___closed__3; x_9 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_9, 0, x_7); lean_ctor_set(x_9, 1, x_8); @@ -6306,7 +6454,7 @@ lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_ x_9 = l_Lean_Syntax_getAntiquotTerm___closed__1; x_10 = lean_array_push(x_9, x_4); x_11 = lean_box(2); -x_12 = l_Lean_Syntax_mkAntiquotNode___closed__15; +x_12 = l_Lean_Syntax_mkAntiquotNode___closed__16; x_13 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_13, 0, x_11); lean_ctor_set(x_13, 1, x_12); @@ -6343,7 +6491,7 @@ lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean x_20 = l_Lean_Syntax_getAntiquotTerm___closed__1; x_21 = lean_array_push(x_20, x_15); x_22 = lean_box(2); -x_23 = l_Lean_Syntax_mkAntiquotNode___closed__15; +x_23 = l_Lean_Syntax_mkAntiquotNode___closed__16; x_24 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_24, 0, x_22); lean_ctor_set(x_24, 1, x_23); @@ -6384,7 +6532,7 @@ if (lean_obj_tag(x_3) == 1) lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_ctor_get(x_3, 0); x_15 = lean_ctor_get(x_3, 1); -x_16 = l_Lean_Syntax_mkAntiquotNode___closed__6; +x_16 = l_Lean_Syntax_mkAntiquotNode___closed__7; x_17 = lean_string_dec_eq(x_15, x_16); if (x_17 == 0) { @@ -6558,7 +6706,7 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_antiquotKinds(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; -x_2 = l_Lean_choiceKind; +x_2 = l_Lean_Syntax_instForInTopDownSyntax_loop___rarg___lambda__7___closed__2; lean_inc(x_1); x_3 = l_Lean_Syntax_isOfKind(x_1, x_2); if (x_3 == 0) @@ -6846,7 +6994,7 @@ lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_obj x_5 = l_Lean_Syntax_mkAntiquotNode___closed__2; x_6 = lean_mk_array(x_4, x_5); x_7 = lean_box(2); -x_8 = l_Lean_nullKind; +x_8 = l_Lean_Syntax_asNode___closed__3; x_9 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_9, 0, x_7); lean_ctor_set(x_9, 1, x_8); @@ -7013,7 +7161,7 @@ x_6 = lean_box(2); x_7 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_3); -x_8 = l_Lean_Syntax_mkAntiquotNode___closed__12; +x_8 = l_Lean_Syntax_mkAntiquotNode___closed__13; x_9 = lean_array_push(x_8, x_2); x_10 = lean_array_push(x_9, x_7); x_11 = lean_alloc_ctor(1, 3, 0); @@ -7138,6 +7286,26 @@ lean_dec_ref(res); res = initialize_Lean_Data_Format(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +l_Lean_isLitKind___closed__1 = _init_l_Lean_isLitKind___closed__1(); +lean_mark_persistent(l_Lean_isLitKind___closed__1); +l_Lean_isLitKind___closed__2 = _init_l_Lean_isLitKind___closed__2(); +lean_mark_persistent(l_Lean_isLitKind___closed__2); +l_Lean_isLitKind___closed__3 = _init_l_Lean_isLitKind___closed__3(); +lean_mark_persistent(l_Lean_isLitKind___closed__3); +l_Lean_isLitKind___closed__4 = _init_l_Lean_isLitKind___closed__4(); +lean_mark_persistent(l_Lean_isLitKind___closed__4); +l_Lean_isLitKind___closed__5 = _init_l_Lean_isLitKind___closed__5(); +lean_mark_persistent(l_Lean_isLitKind___closed__5); +l_Lean_isLitKind___closed__6 = _init_l_Lean_isLitKind___closed__6(); +lean_mark_persistent(l_Lean_isLitKind___closed__6); +l_Lean_isLitKind___closed__7 = _init_l_Lean_isLitKind___closed__7(); +lean_mark_persistent(l_Lean_isLitKind___closed__7); +l_Lean_isLitKind___closed__8 = _init_l_Lean_isLitKind___closed__8(); +lean_mark_persistent(l_Lean_isLitKind___closed__8); +l_Lean_isLitKind___closed__9 = _init_l_Lean_isLitKind___closed__9(); +lean_mark_persistent(l_Lean_isLitKind___closed__9); +l_Lean_isLitKind___closed__10 = _init_l_Lean_isLitKind___closed__10(); +lean_mark_persistent(l_Lean_isLitKind___closed__10); l_Lean_Syntax_getAtomVal_x21___closed__1 = _init_l_Lean_Syntax_getAtomVal_x21___closed__1(); lean_mark_persistent(l_Lean_Syntax_getAtomVal_x21___closed__1); l_Lean_Syntax_getAtomVal_x21___closed__2 = _init_l_Lean_Syntax_getAtomVal_x21___closed__2(); @@ -7150,6 +7318,10 @@ l_Lean_Syntax_asNode___closed__1 = _init_l_Lean_Syntax_asNode___closed__1(); lean_mark_persistent(l_Lean_Syntax_asNode___closed__1); l_Lean_Syntax_asNode___closed__2 = _init_l_Lean_Syntax_asNode___closed__2(); lean_mark_persistent(l_Lean_Syntax_asNode___closed__2); +l_Lean_Syntax_asNode___closed__3 = _init_l_Lean_Syntax_asNode___closed__3(); +lean_mark_persistent(l_Lean_Syntax_asNode___closed__3); +l_Lean_Syntax_asNode___closed__4 = _init_l_Lean_Syntax_asNode___closed__4(); +lean_mark_persistent(l_Lean_Syntax_asNode___closed__4); l_List_mapTRAux___at_Lean_Syntax_identComponents___spec__2___closed__1 = _init_l_List_mapTRAux___at_Lean_Syntax_identComponents___spec__2___closed__1(); lean_mark_persistent(l_List_mapTRAux___at_Lean_Syntax_identComponents___spec__2___closed__1); l_List_mapTRAux___at_Lean_Syntax_identComponents___spec__2___closed__2 = _init_l_List_mapTRAux___at_Lean_Syntax_identComponents___spec__2___closed__2(); @@ -7174,6 +7346,10 @@ l_Lean_Syntax_identComponents___closed__8 = _init_l_Lean_Syntax_identComponents_ lean_mark_persistent(l_Lean_Syntax_identComponents___closed__8); l_Lean_Syntax_identComponents___closed__9 = _init_l_Lean_Syntax_identComponents___closed__9(); lean_mark_persistent(l_Lean_Syntax_identComponents___closed__9); +l_Lean_Syntax_instForInTopDownSyntax_loop___rarg___lambda__7___closed__1 = _init_l_Lean_Syntax_instForInTopDownSyntax_loop___rarg___lambda__7___closed__1(); +lean_mark_persistent(l_Lean_Syntax_instForInTopDownSyntax_loop___rarg___lambda__7___closed__1); +l_Lean_Syntax_instForInTopDownSyntax_loop___rarg___lambda__7___closed__2 = _init_l_Lean_Syntax_instForInTopDownSyntax_loop___rarg___lambda__7___closed__2(); +lean_mark_persistent(l_Lean_Syntax_instForInTopDownSyntax_loop___rarg___lambda__7___closed__2); l_Lean_Syntax_reprint_reprintLeaf___closed__1 = _init_l_Lean_Syntax_reprint_reprintLeaf___closed__1(); lean_mark_persistent(l_Lean_Syntax_reprint_reprintLeaf___closed__1); l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Syntax_reprint___spec__3___closed__1 = _init_l_Lean_Syntax_instForInTopDownSyntax_loop___at_Lean_Syntax_reprint___spec__3___closed__1(); @@ -7266,6 +7442,8 @@ l_Lean_Syntax_mkAntiquotNode___closed__22 = _init_l_Lean_Syntax_mkAntiquotNode__ lean_mark_persistent(l_Lean_Syntax_mkAntiquotNode___closed__22); l_Lean_Syntax_mkAntiquotNode___closed__23 = _init_l_Lean_Syntax_mkAntiquotNode___closed__23(); lean_mark_persistent(l_Lean_Syntax_mkAntiquotNode___closed__23); +l_Lean_Syntax_mkAntiquotNode___closed__24 = _init_l_Lean_Syntax_mkAntiquotNode___closed__24(); +lean_mark_persistent(l_Lean_Syntax_mkAntiquotNode___closed__24); l_Lean_Syntax_getAntiquotTerm___closed__1 = _init_l_Lean_Syntax_getAntiquotTerm___closed__1(); lean_mark_persistent(l_Lean_Syntax_getAntiquotTerm___closed__1); l_Lean_Syntax_antiquotSpliceKind_x3f___closed__1 = _init_l_Lean_Syntax_antiquotSpliceKind_x3f___closed__1(); diff --git a/stage0/stdlib/Lean/Util/Path.c b/stage0/stdlib/Lean/Util/Path.c index 1fd039f4d75..ca6670906a2 100644 --- a/stage0/stdlib/Lean/Util/Path.c +++ b/stage0/stdlib/Lean/Util/Path.c @@ -77,9 +77,7 @@ extern lean_object* l_System_instInhabitedFilePath; LEAN_EXPORT lean_object* l_Lean_SearchPath_findModuleWithExt___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_addSearchPathFromEnv___closed__1; static lean_object* l_Lean_modToFilePath_go___closed__3; -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_SearchPath_findAllWithExt___spec__1___boxed(lean_object*, lean_object*); uint8_t l_Lean_Internal_isStage0(lean_object*); -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_SearchPath_findAllWithExt___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at_Lean_moduleNameOfFileName___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_findOLean___lambda__1(lean_object*, lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); @@ -119,6 +117,7 @@ static lean_object* l_Lean_findOLean_maybeThisOne___closed__1; lean_object* l_System_FilePath_isDir(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_modToFilePath(lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_SearchPath_findAllWithExt___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_SearchPath_findAllWithExt___closed__1; lean_object* lean_panic_fn(lean_object*, lean_object*); lean_object* l_String_trim(lean_object*); @@ -131,6 +130,7 @@ LEAN_EXPORT lean_object* l_Lean_SearchPath_findModuleWithExt___boxed(lean_object lean_object* l_System_FilePath_extension(lean_object*); lean_object* lean_string_length(lean_object*); lean_object* l_System_FilePath_parent(lean_object*); +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_SearchPath_findAllWithExt___spec__1(lean_object*, lean_object*); lean_object* l_System_FilePath_withExtension(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_findOLean_maybeThisOne(lean_object*, lean_object*, lean_object*); lean_object* l_String_drop(lean_object*, lean_object*); @@ -689,7 +689,7 @@ lean_dec(x_2); return x_5; } } -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_SearchPath_findAllWithExt___spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_SearchPath_findAllWithExt___spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -740,7 +740,7 @@ x_8 = l_System_FilePath_extension(x_7); lean_inc(x_1); x_9 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_9, 0, x_1); -x_10 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_SearchPath_findAllWithExt___spec__1(x_8, x_9); +x_10 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_SearchPath_findAllWithExt___spec__1(x_8, x_9); lean_dec(x_9); lean_dec(x_8); x_11 = 1; @@ -980,11 +980,11 @@ return x_13; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_SearchPath_findAllWithExt___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_SearchPath_findAllWithExt___spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_SearchPath_findAllWithExt___spec__1(x_1, x_2); +x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_SearchPath_findAllWithExt___spec__1(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); diff --git a/stage0/stdlib/Lean/Util/SCC.c b/stage0/stdlib/Lean/Util/SCC.c index 844a5e6a06a..74b1784e32e 100644 --- a/stage0/stdlib/Lean/Util/SCC.c +++ b/stage0/stdlib/Lean/Util/SCC.c @@ -18,7 +18,6 @@ LEAN_EXPORT lean_object* l_Lean_SCC_State_data___default___boxed(lean_object*, l static lean_object* l_Lean_SCC_State_data___default___closed__1; LEAN_EXPORT lean_object* l_Lean_SCC_scc(lean_object*); LEAN_EXPORT lean_object* l_List_forM___at_Lean_SCC_scc___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Util_SCC_0__Lean_SCC_sccAux___spec__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_SCC_0__Lean_SCC_getDataOf(lean_object*); LEAN_EXPORT lean_object* l_Lean_SCC_Data_index_x3f___default; LEAN_EXPORT lean_object* l___private_Lean_Util_SCC_0__Lean_SCC_push___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -53,7 +52,8 @@ uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at_Lean_SCC_scc___spec__1(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_SCC_0__Lean_SCC_resetOnStack___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_SCC_0__Lean_SCC_updateLowLinkOf(lean_object*); -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Util_SCC_0__Lean_SCC_sccAux___spec__2(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Util_SCC_0__Lean_SCC_sccAux___spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Util_SCC_0__Lean_SCC_sccAux___spec__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_SCC_0__Lean_SCC_addSCC_add___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Util_SCC_0__Lean_SCC_getDataOf___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_SCC_State_sccs___default(lean_object*); @@ -1273,7 +1273,7 @@ x_2 = lean_alloc_closure((void*)(l_List_forM___at___private_Lean_Util_SCC_0__Lea return x_2; } } -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Util_SCC_0__Lean_SCC_sccAux___spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Util_SCC_0__Lean_SCC_sccAux___spec__2(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -1346,7 +1346,7 @@ lean_inc(x_15); x_16 = lean_ctor_get(x_13, 0); lean_inc(x_16); lean_dec(x_13); -x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Util_SCC_0__Lean_SCC_sccAux___spec__2(x_15, x_16); +x_17 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Util_SCC_0__Lean_SCC_sccAux___spec__2(x_15, x_16); lean_dec(x_16); lean_dec(x_15); if (x_17 == 0) @@ -1380,7 +1380,7 @@ lean_inc(x_22); x_23 = lean_ctor_get(x_20, 0); lean_inc(x_23); lean_dec(x_20); -x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Util_SCC_0__Lean_SCC_sccAux___spec__2(x_22, x_23); +x_24 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Util_SCC_0__Lean_SCC_sccAux___spec__2(x_22, x_23); lean_dec(x_23); lean_dec(x_22); if (x_24 == 0) @@ -1412,11 +1412,11 @@ x_2 = lean_alloc_closure((void*)(l___private_Lean_Util_SCC_0__Lean_SCC_sccAux___ return x_2; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Util_SCC_0__Lean_SCC_sccAux___spec__2___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Util_SCC_0__Lean_SCC_sccAux___spec__2___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at___private_Lean_Util_SCC_0__Lean_SCC_sccAux___spec__2(x_1, x_2); +x_3 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at___private_Lean_Util_SCC_0__Lean_SCC_sccAux___spec__2(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); diff --git a/stage0/stdlib/Lean/Util/Trace.c b/stage0/stdlib/Lean/Util/Trace.c index ba7e5e0e134..38c1fff00fd 100644 --- a/stage0/stdlib/Lean/Util/Trace.c +++ b/stage0/stdlib/Lean/Util/Trace.c @@ -13,82 +13,85 @@ #ifdef __cplusplus extern "C" { #endif -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__42; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__37; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__14; lean_object* l_Lean_KVMap_setBool(lean_object*, lean_object*, uint8_t); +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__35; size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_registerTraceClass(lean_object*, lean_object*); lean_object* lean_erase_macro_scopes(lean_object*); +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__15; lean_object* l_Lean_stringToMessageData(lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__2; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__36; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_printTraces___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addRawTrace(lean_object*); -static lean_object* l_Lean_doElemTrace_x5b_____x5d_______closed__24; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__18; uint8_t l_Lean_MessageData_isNest(lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); -static lean_object* l_Lean_doElemTrace_x5b_____x5d_______closed__7; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__12; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__11; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__5; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__19; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Util_Trace_0__Lean_withNestedTracesFinalizer___spec__4(lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__9; LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__4; LEAN_EXPORT lean_object* l_Lean_traceCtx___rarg___lambda__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); +static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__13; lean_object* lean_name_mk_string(lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); lean_object* lean_array_uget(lean_object*, size_t); -static lean_object* l_Lean_doElemTrace_x5b_____x5d_______closed__16; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__21; +static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__23; LEAN_EXPORT lean_object* l_Std_PersistentArray_forM___at_Lean_printTraces___spec__1___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__38; lean_object* l_Lean_MessageData_format(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_printTraces___spec__4___rarg___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_printTraces___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_registerTraceClass___closed__1; static lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceOptions___closed__3; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__44; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_withNestedTracesFinalizer___rarg___lambda__1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__17; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__8; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOptionM___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOptionM___rarg___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_traceCtx___rarg___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); size_t lean_usize_sub(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_enableTracing___rarg___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_printTraces___spec__4___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_trace___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__17; LEAN_EXPORT uint8_t l_Lean_checkTraceOption(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_printTraces___spec__5___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__26; lean_object* l_Std_PersistentArray_append___rarg(lean_object*, lean_object*); static lean_object* l_Lean_TraceState_traces___default___closed__1; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__7; +static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__11; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__8; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__4; uint8_t lean_name_eq(lean_object*, lean_object*); static lean_object* l_Lean_resetTraceState___rarg___closed__1; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__11; -static lean_object* l_Lean_doElemTrace_x5b_____x5d_______closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addNode___spec__1(size_t, size_t, lean_object*); extern lean_object* l_Lean_MessageData_nil; static lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceOptions___closed__4; static lean_object* l_Lean_registerTraceClass___closed__3; LEAN_EXPORT lean_object* l_Lean_enableTracing___rarg(lean_object*, lean_object*, uint8_t); +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__9; lean_object* l_Lean_MonadRef_mkInfoFromRefPos___at___aux__Init__Notation______macroRules__precMax__1___spec__1(lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__30; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Lean_traceCtx___rarg___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); -static lean_object* l_Lean_doElemTrace_x5b_____x5d_______closed__21; lean_object* lean_string_append(lean_object*, lean_object*); -extern lean_object* l_Lean_interpolatedStrKind; +static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__8; LEAN_EXPORT lean_object* l_Lean_getTraces(lean_object*); -static lean_object* l_Lean_doElemTrace_x5b_____x5d_______closed__13; static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Util_Trace_0__Lean_withNestedTracesFinalizer___spec__5___closed__1; -static lean_object* l_Lean_doElemTrace_x5b_____x5d_______closed__23; LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_printTraces___spec__5___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_printTraces___spec__5___rarg(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__24; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__12; +static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__17; LEAN_EXPORT lean_object* l_Lean_addRawTrace___rarg___lambda__2(lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_right(size_t, size_t); static lean_object* l_Lean_instInhabitedTraceState___closed__5; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_printTraces___spec__3___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*); +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__39; static lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___rarg___lambda__3___closed__1; lean_object* lean_string_utf8_byte_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_withNestedTraces___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -96,80 +99,79 @@ LEAN_EXPORT lean_object* l_Lean_traceM___rarg___lambda__1___boxed(lean_object*, LEAN_EXPORT lean_object* l_Std_PersistentArray_forM___at_Lean_printTraces___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOptionM___rarg(lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_nameLitKind; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Util_Trace_0__Lean_withNestedTracesFinalizer___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_withNestedTracesFinalizer___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__33; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__3; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__34; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__13; lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentArray_foldlM___at___private_Lean_Util_Trace_0__Lean_withNestedTracesFinalizer___spec__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__23; static lean_object* l_Lean_addTrace___rarg___lambda__1___closed__4; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__41; LEAN_EXPORT lean_object* l_Lean_instInhabitedTraceState; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__13; LEAN_EXPORT lean_object* l_Lean_traceCtx___rarg___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOptionM(lean_object*); -LEAN_EXPORT lean_object* l_Lean_doElemTrace_x5b_____x5d____; LEAN_EXPORT lean_object* l_Lean_trace(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Util_Trace_0__Lean_withNestedTracesFinalizer___spec__5(lean_object*, size_t, size_t, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__9; +static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__12; static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Util_Trace_0__Lean_withNestedTracesFinalizer___spec__5___closed__2; LEAN_EXPORT lean_object* l_Lean_setTraceState(lean_object*); LEAN_EXPORT lean_object* l_Lean_withNestedTraces(lean_object*, lean_object*); -static lean_object* l_Lean_doElemTrace_x5b_____x5d_______closed__12; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__10; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__31; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addNode(lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__11; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__33; LEAN_EXPORT lean_object* l_Nat_foldM_loop___at___private_Lean_Util_Trace_0__Lean_TraceState_toFormat___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_traceM___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_printTraces___spec__3(lean_object*); LEAN_EXPORT lean_object* l_Lean_setTraceState___rarg___lambda__1(lean_object*, lean_object*); -static lean_object* l_Lean_doElemTrace_x5b_____x5d_______closed__26; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__22; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__2; LEAN_EXPORT lean_object* l_Lean_enableTracing___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_modifyTraces(lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__19; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_withNestedTracesFinalizer(lean_object*, lean_object*); +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__21; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__38; lean_object* l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__26; LEAN_EXPORT lean_object* l_Lean_addTrace___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__4; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__44; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__1; +LEAN_EXPORT lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_KVMap_getBool(lean_object*, lean_object*, uint8_t); -static lean_object* l_Lean_doElemTrace_x5b_____x5d_______closed__22; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__4; lean_object* l_Std_PersistentArray_get_x21___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__6; +LEAN_EXPORT lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__20; static lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceOptions___closed__1; LEAN_EXPORT lean_object* l_Lean_enableTracing___rarg___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__1; static lean_object* l_Lean_registerTraceClass___closed__2; -LEAN_EXPORT lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_doElemTrace_x5b_____x5d_______closed__5; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_withNestedTracesFinalizer___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_doElemTrace_x5b_____x5d_______closed__14; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__45; uint8_t l_Lean_MessageData_isNil(lean_object*); -static lean_object* l_Lean_doElemTrace_x5b_____x5d_______closed__15; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__40; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__17; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__8; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__12; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addNode___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_traceCtx___rarg___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_doElemTrace_x5b_____x5d_______closed__18; LEAN_EXPORT lean_object* l_Lean_traceM(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_printTraces___spec__4___rarg(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__46; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__34; static lean_object* l_Lean_addTrace___rarg___lambda__1___closed__1; -static lean_object* l_Lean_doElemTrace_x5b_____x5d_______closed__3; +static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__26; static lean_object* l_Lean_addTrace___rarg___lambda__1___closed__8; static lean_object* l_Lean_addTrace___rarg___lambda__1___closed__2; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__41; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__13; static lean_object* l_Lean_TraceState_traces___default___closed__2; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__9; lean_object* l_Lean_Syntax_getId(lean_object*); static lean_object* l_Lean_traceCtx___rarg___lambda__2___closed__1; -lean_object* l___private_Init_Meta_0__Lean_quoteNameMk(lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor(lean_object*); LEAN_EXPORT lean_object* l_Lean_traceCtx___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_printTraces___spec__4(lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__8; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__27; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces(lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withNestedTraces___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -180,161 +182,159 @@ LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___ra LEAN_EXPORT lean_object* l_Lean_traceCtx___rarg___lambda__4(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_KVMap_contains(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withNestedTraces___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__23; LEAN_EXPORT lean_object* l_Lean_withNestedTraces___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__1; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__39; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__35; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__13; +static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__10; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__3; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__37; LEAN_EXPORT lean_object* l_Lean_enableTracing(lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentArray_forMAux___at_Lean_printTraces___spec__2(lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__43; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_printTraces___spec__4___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__6; uint8_t l_Std_PersistentArray_isEmpty___rarg(lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__20; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__10; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__32; +static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__24; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__28; +static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__7; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_printTraces___spec__5(lean_object*); +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__14; LEAN_EXPORT lean_object* l_Lean_addTrace___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_modifyTraces___rarg___lambda__1(lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__9; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_MonadTracer_trace___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_doElemTrace_x5b_____x5d_______closed__9; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_printTraces___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__27; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__11; static lean_object* l_Lean_addTrace___rarg___lambda__1___closed__6; -static lean_object* l_Lean_doElemTrace_x5b_____x5d_______closed__17; +static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__16; LEAN_EXPORT lean_object* l_Lean_addTrace___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Lean_traceCtx___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_doElemTrace_x5b_____x5d_______closed__4; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__2; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__10; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__32; lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___rarg(lean_object*, lean_object*); +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__20; static lean_object* l_Lean_instInhabitedTraceElem___closed__1; LEAN_EXPORT lean_object* l_Lean_modifyTraces___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_doElemTrace_x5b_____x5d_______closed__8; size_t lean_usize_land(size_t, size_t); +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__18; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceOptions(lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__16; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__36; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__30; static lean_object* l_Lean_instInhabitedTraceState___closed__1; LEAN_EXPORT lean_object* l_Lean_traceCtx___rarg___lambda__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceOptions___closed__2; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__25; static lean_object* l_Lean_instInhabitedTraceState___closed__2; -static lean_object* l_Lean_doElemTrace_x5b_____x5d_______closed__20; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__14; lean_object* l_String_intercalate(lean_object*, lean_object*); static lean_object* l_Lean_addTrace___rarg___lambda__1___closed__3; -static lean_object* l_Lean_doElemTrace_x5b_____x5d_______closed__10; lean_object* l_Std_PersistentArray_push___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___rarg___lambda__2(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__14; LEAN_EXPORT lean_object* l_Lean_resetTraceState(lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__3; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_printTraces___spec__3___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MonadTracer_trace(lean_object*); -LEAN_EXPORT lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_setTraceState___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedTraceElem; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__15; LEAN_EXPORT lean_object* l_Lean_enableTracing___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_traceCtx(lean_object*, lean_object*); +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__6; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__10; LEAN_EXPORT lean_object* l_Lean_TraceState_traces___default; static lean_object* l_Lean_addTrace___rarg___lambda__1___closed__5; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__46; static lean_object* l_Lean_resetTraceState___rarg___lambda__1___closed__1; static lean_object* l___private_Lean_Util_Trace_0__Lean_addNode___rarg___lambda__1___closed__1; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__5; static lean_object* l_Lean_TraceState_traces___default___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addNode___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedTraceState___closed__4; LEAN_EXPORT lean_object* l_Lean_instMonadTrace___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__5; uint8_t lean_nat_dec_le(lean_object*, lean_object*); +static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__15; +static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__18; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__29; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__7; lean_object* l_Std_PersistentArray_toArray___rarg(lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__40; lean_object* l_Lean_Name_append(lean_object*, lean_object*); lean_object* l_Lean_Syntax_getKind(lean_object*); LEAN_EXPORT lean_object* l___private_Std_Data_PersistentArray_0__Std_PersistentArray_foldlMAux___at___private_Lean_Util_Trace_0__Lean_withNestedTracesFinalizer___spec__3(lean_object*, lean_object*); +lean_object* l_Lean_quoteNameMk(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__45; LEAN_EXPORT lean_object* l_Lean_traceCtx___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_enableTracing___rarg___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedTraceElem___closed__2; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__19; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__42; LEAN_EXPORT lean_object* l_Lean_enableTracing___rarg___lambda__2(lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Std_PersistentArray_forMAux___at_Lean_printTraces___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_register_option(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__22; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Util_Trace_0__Lean_withNestedTracesFinalizer___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addRawTrace___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__15; LEAN_EXPORT lean_object* l_Lean_printTraces(lean_object*); +lean_object* l_Lean_Syntax_mkNameLit(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_checkTraceOption___boxed(lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_checkTraceOption___closed__1; LEAN_EXPORT lean_object* l_Lean_traceM___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__7; LEAN_EXPORT lean_object* l_Lean_addTrace(lean_object*); LEAN_EXPORT lean_object* l___private_Std_Data_PersistentArray_0__Std_PersistentArray_foldlFromMAux___at___private_Lean_Util_Trace_0__Lean_withNestedTracesFinalizer___spec__2(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_addRawTrace___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__16; LEAN_EXPORT lean_object* l_Lean_resetTraceState___rarg___lambda__1___boxed(lean_object*); lean_object* l_IO_println___at_Lean_instEval__1___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_addTrace___rarg___lambda__1___closed__7; LEAN_EXPORT lean_object* l___private_Std_Data_PersistentArray_0__Std_PersistentArray_foldlFromMAux___at___private_Lean_Util_Trace_0__Lean_withNestedTracesFinalizer___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_printTraces___spec__4___rarg___lambda__2(size_t, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*); -static lean_object* l_Lean_doElemTrace_x5b_____x5d_______closed__6; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__22; +static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__6; +static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__25; static lean_object* l___private_Std_Data_PersistentArray_0__Std_PersistentArray_foldlFromMAux___at___private_Lean_Util_Trace_0__Lean_withNestedTracesFinalizer___spec__2___closed__1; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__43; static size_t l_Lean_instInhabitedTraceState___closed__3; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__12; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addNode___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_printTraces___spec__5___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_resetTraceState___rarg___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_enableTracing___rarg___lambda__1(uint8_t, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__14; uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); static lean_object* l___private_Lean_Util_Trace_0__Lean_addNode___rarg___lambda__1___closed__2; -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__18; -static lean_object* l_Lean_doElemTrace_x5b_____x5d_______closed__25; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__31; +static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__3; LEAN_EXPORT lean_object* l_Lean_trace___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__28; LEAN_EXPORT lean_object* l_Lean_traceCtx___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__29; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__25; LEAN_EXPORT lean_object* l_Lean_traceCtx___rarg___lambda__5___boxed(lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__15; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__3; LEAN_EXPORT lean_object* l_Lean_printTraces___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_TraceState_enabled___default; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instMonadTrace___rarg___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_setTraceState___rarg___lambda__1___boxed(lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__16; +LEAN_EXPORT lean_object* l_Lean_doElemTrace_x5b___x5d____; LEAN_EXPORT lean_object* l_Lean_getTraces___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_doElemTrace_x5b_____x5d_______closed__1; +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__16; LEAN_EXPORT lean_object* l_Lean_trace___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_foldM_loop___at___private_Lean_Util_Trace_0__Lean_TraceState_toFormat___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_traceCtx___rarg___lambda__1(lean_object*); uint8_t l_List_isEmpty___rarg(lean_object*); +static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__2; lean_object* lean_usize_to_nat(size_t); LEAN_EXPORT lean_object* l_Lean_instMonadTrace(lean_object*, lean_object*); lean_object* l_Std_instInhabitedPersistentArrayNode(lean_object*); +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__7; LEAN_EXPORT lean_object* l_Lean_traceCtx___rarg___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__24; static lean_object* l_Lean_traceCtx___rarg___lambda__6___closed__1; LEAN_EXPORT lean_object* l_Lean_getTraces___rarg___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_traceCtx___rarg___lambda__5(lean_object*); LEAN_EXPORT lean_object* l_Lean_withNestedTraces___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__21; LEAN_EXPORT lean_object* l_Lean_traceCtx___rarg___lambda__3(lean_object*, uint8_t, lean_object*); +static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__6; static lean_object* l_Lean_checkTraceOption___closed__2; LEAN_EXPORT lean_object* l_Lean_addRawTrace___rarg___lambda__1(lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Util_Trace_0__Lean_checkTraceOptionAux(lean_object*, lean_object*); -static lean_object* l_Lean_doElemTrace_x5b_____x5d_______closed__11; -static lean_object* l_Lean_doElemTrace_x5b_____x5d_______closed__19; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_TraceState_toFormat(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__1; -lean_object* l_Lean_Syntax_mkLit(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__1; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOptionAux___boxed(lean_object*, lean_object*); static lean_object* _init_l_Lean_instInhabitedTraceElem___closed__1() { @@ -2973,7 +2973,7 @@ x_6 = lean_register_option(x_4, x_5, x_2); return x_6; } } -static lean_object* _init_l_Lean_doElemTrace_x5b_____x5d_______closed__1() { +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__1() { _start: { lean_object* x_1; @@ -2981,35 +2981,35 @@ x_1 = lean_mk_string_from_bytes("Lean", 4); return x_1; } } -static lean_object* _init_l_Lean_doElemTrace_x5b_____x5d_______closed__2() { +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_doElemTrace_x5b_____x5d_______closed__1; +x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_doElemTrace_x5b_____x5d_______closed__3() { +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("doElemTrace[__]__", 17); +x_1 = lean_mk_string_from_bytes("doElemTrace[_]__", 16); return x_1; } } -static lean_object* _init_l_Lean_doElemTrace_x5b_____x5d_______closed__4() { +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_doElemTrace_x5b_____x5d_______closed__2; -x_2 = l_Lean_doElemTrace_x5b_____x5d_______closed__3; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__2; +x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__3; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_doElemTrace_x5b_____x5d_______closed__5() { +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__5() { _start: { lean_object* x_1; @@ -3017,17 +3017,17 @@ x_1 = lean_mk_string_from_bytes("andthen", 7); return x_1; } } -static lean_object* _init_l_Lean_doElemTrace_x5b_____x5d_______closed__6() { +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_doElemTrace_x5b_____x5d_______closed__5; +x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__5; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_doElemTrace_x5b_____x5d_______closed__7() { +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__7() { _start: { lean_object* x_1; @@ -3035,17 +3035,17 @@ x_1 = lean_mk_string_from_bytes("trace[", 6); return x_1; } } -static lean_object* _init_l_Lean_doElemTrace_x5b_____x5d_______closed__8() { +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_doElemTrace_x5b_____x5d_______closed__7; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__7; x_2 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_doElemTrace_x5b_____x5d_______closed__9() { +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__9() { _start: { lean_object* x_1; @@ -3053,33 +3053,33 @@ x_1 = lean_mk_string_from_bytes("ident", 5); return x_1; } } -static lean_object* _init_l_Lean_doElemTrace_x5b_____x5d_______closed__10() { +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_doElemTrace_x5b_____x5d_______closed__9; +x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__9; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_doElemTrace_x5b_____x5d_______closed__11() { +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_doElemTrace_x5b_____x5d_______closed__10; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__10; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_doElemTrace_x5b_____x5d_______closed__12() { +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_doElemTrace_x5b_____x5d_______closed__6; -x_2 = l_Lean_doElemTrace_x5b_____x5d_______closed__8; -x_3 = l_Lean_doElemTrace_x5b_____x5d_______closed__11; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__6; +x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__8; +x_3 = l_Lean_doElemTrace_x5b___x5d_______closed__11; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3087,7 +3087,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_doElemTrace_x5b_____x5d_______closed__13() { +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__13() { _start: { lean_object* x_1; @@ -3095,23 +3095,23 @@ x_1 = lean_mk_string_from_bytes("]", 1); return x_1; } } -static lean_object* _init_l_Lean_doElemTrace_x5b_____x5d_______closed__14() { +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__14() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_doElemTrace_x5b_____x5d_______closed__13; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__13; x_2 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_doElemTrace_x5b_____x5d_______closed__15() { +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_doElemTrace_x5b_____x5d_______closed__6; -x_2 = l_Lean_doElemTrace_x5b_____x5d_______closed__12; -x_3 = l_Lean_doElemTrace_x5b_____x5d_______closed__14; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__6; +x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__12; +x_3 = l_Lean_doElemTrace_x5b___x5d_______closed__14; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3119,7 +3119,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_doElemTrace_x5b_____x5d_______closed__16() { +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__16() { _start: { lean_object* x_1; @@ -3127,17 +3127,17 @@ x_1 = lean_mk_string_from_bytes("orelse", 6); return x_1; } } -static lean_object* _init_l_Lean_doElemTrace_x5b_____x5d_______closed__17() { +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_doElemTrace_x5b_____x5d_______closed__16; +x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__16; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_doElemTrace_x5b_____x5d_______closed__18() { +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__18() { _start: { lean_object* x_1; @@ -3145,17 +3145,17 @@ x_1 = lean_mk_string_from_bytes("interpolatedStr", 15); return x_1; } } -static lean_object* _init_l_Lean_doElemTrace_x5b_____x5d_______closed__19() { +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_doElemTrace_x5b_____x5d_______closed__18; +x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__18; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_doElemTrace_x5b_____x5d_______closed__20() { +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__20() { _start: { lean_object* x_1; @@ -3163,21 +3163,21 @@ x_1 = lean_mk_string_from_bytes("term", 4); return x_1; } } -static lean_object* _init_l_Lean_doElemTrace_x5b_____x5d_______closed__21() { +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_doElemTrace_x5b_____x5d_______closed__20; +x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__20; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_doElemTrace_x5b_____x5d_______closed__22() { +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_doElemTrace_x5b_____x5d_______closed__21; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__21; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -3185,25 +3185,25 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_doElemTrace_x5b_____x5d_______closed__23() { +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_doElemTrace_x5b_____x5d_______closed__19; -x_2 = l_Lean_doElemTrace_x5b_____x5d_______closed__22; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__19; +x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__22; x_3 = lean_alloc_ctor(1, 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_Lean_doElemTrace_x5b_____x5d_______closed__24() { +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_doElemTrace_x5b_____x5d_______closed__17; -x_2 = l_Lean_doElemTrace_x5b_____x5d_______closed__23; -x_3 = l_Lean_doElemTrace_x5b_____x5d_______closed__22; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__17; +x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__23; +x_3 = l_Lean_doElemTrace_x5b___x5d_______closed__22; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3211,13 +3211,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_doElemTrace_x5b_____x5d_______closed__25() { +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_doElemTrace_x5b_____x5d_______closed__6; -x_2 = l_Lean_doElemTrace_x5b_____x5d_______closed__15; -x_3 = l_Lean_doElemTrace_x5b_____x5d_______closed__24; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__6; +x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__15; +x_3 = l_Lean_doElemTrace_x5b___x5d_______closed__24; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3225,13 +3225,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_doElemTrace_x5b_____x5d_______closed__26() { +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d_______closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_doElemTrace_x5b_____x5d_______closed__4; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__4; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_doElemTrace_x5b_____x5d_______closed__25; +x_3 = l_Lean_doElemTrace_x5b___x5d_______closed__25; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3239,15 +3239,15 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_doElemTrace_x5b_____x5d____() { +static lean_object* _init_l_Lean_doElemTrace_x5b___x5d____() { _start: { lean_object* x_1; -x_1 = l_Lean_doElemTrace_x5b_____x5d_______closed__26; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__26; return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__1() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__1() { _start: { lean_object* x_1; @@ -3255,7 +3255,7 @@ x_1 = lean_mk_string_from_bytes("Parser", 6); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__2() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__2() { _start: { lean_object* x_1; @@ -3263,7 +3263,7 @@ x_1 = lean_mk_string_from_bytes("Term", 4); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__3() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__3() { _start: { lean_object* x_1; @@ -3271,7 +3271,7 @@ x_1 = lean_mk_string_from_bytes("doNested", 8); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__4() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__4() { _start: { lean_object* x_1; @@ -3279,7 +3279,7 @@ x_1 = lean_mk_string_from_bytes("do", 2); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__5() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__5() { _start: { lean_object* x_1; @@ -3287,7 +3287,7 @@ x_1 = lean_mk_string_from_bytes("doSeqIndent", 11); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__6() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__6() { _start: { lean_object* x_1; @@ -3295,17 +3295,17 @@ x_1 = lean_mk_string_from_bytes("null", 4); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__7() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__6; +x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__6; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__8() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__8() { _start: { lean_object* x_1; @@ -3313,7 +3313,7 @@ x_1 = lean_mk_string_from_bytes("doSeqItem", 9); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__9() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__9() { _start: { lean_object* x_1; @@ -3321,7 +3321,7 @@ x_1 = lean_mk_string_from_bytes("doLet", 5); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__10() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__10() { _start: { lean_object* x_1; @@ -3329,12 +3329,12 @@ x_1 = lean_mk_string_from_bytes("let", 3); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__11() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__7; +x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__7; x_3 = l_Lean_instInhabitedTraceState___closed__1; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -3343,7 +3343,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__12() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__12() { _start: { lean_object* x_1; @@ -3351,7 +3351,7 @@ x_1 = lean_mk_string_from_bytes("letDecl", 7); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__13() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__13() { _start: { lean_object* x_1; @@ -3359,7 +3359,7 @@ x_1 = lean_mk_string_from_bytes("letIdDecl", 9); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__14() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__14() { _start: { lean_object* x_1; @@ -3367,22 +3367,22 @@ x_1 = lean_mk_string_from_bytes("cls", 3); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__15() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__15() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__14; +x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__14; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__16() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__14; +x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__14; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__15; +x_3 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__15; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3390,17 +3390,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__17() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__14; +x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__14; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__18() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__18() { _start: { lean_object* x_1; @@ -3408,7 +3408,7 @@ x_1 = lean_mk_string_from_bytes(":=", 2); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__19() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__19() { _start: { lean_object* x_1; lean_object* x_2; @@ -3417,7 +3417,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__20() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__20() { _start: { lean_object* x_1; lean_object* x_2; @@ -3426,7 +3426,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__21() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__21() { _start: { lean_object* x_1; lean_object* x_2; @@ -3435,7 +3435,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__22() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__22() { _start: { lean_object* x_1; lean_object* x_2; @@ -3444,7 +3444,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__23() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__23() { _start: { lean_object* x_1; @@ -3452,7 +3452,7 @@ x_1 = lean_mk_string_from_bytes("doIf", 4); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__24() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__24() { _start: { lean_object* x_1; @@ -3460,7 +3460,7 @@ x_1 = lean_mk_string_from_bytes("if", 2); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__25() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__25() { _start: { lean_object* x_1; @@ -3468,7 +3468,7 @@ x_1 = lean_mk_string_from_bytes("doIfProp", 8); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__26() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__26() { _start: { lean_object* x_1; @@ -3476,7 +3476,7 @@ x_1 = lean_mk_string_from_bytes("paren", 5); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__27() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__27() { _start: { lean_object* x_1; @@ -3484,7 +3484,7 @@ x_1 = lean_mk_string_from_bytes("(", 1); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__28() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__28() { _start: { lean_object* x_1; @@ -3492,7 +3492,7 @@ x_1 = lean_mk_string_from_bytes("liftMethod", 10); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__29() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__29() { _start: { lean_object* x_1; @@ -3500,7 +3500,7 @@ x_1 = lean_mk_string_from_bytes("←", 3); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__30() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__30() { _start: { lean_object* x_1; @@ -3508,7 +3508,7 @@ x_1 = lean_mk_string_from_bytes("app", 3); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__31() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__31() { _start: { lean_object* x_1; @@ -3516,22 +3516,22 @@ x_1 = lean_mk_string_from_bytes("Lean.isTracingEnabledFor", 24); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__32() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__32() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__31; +x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__31; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__33() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__33() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__31; +x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__31; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__32; +x_3 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__32; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3539,7 +3539,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__34() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__34() { _start: { lean_object* x_1; @@ -3547,7 +3547,7 @@ x_1 = lean_mk_string_from_bytes("isTracingEnabledFor", 19); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__35() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__35() { _start: { lean_object* x_1; @@ -3555,17 +3555,17 @@ x_1 = lean_mk_string_from_bytes(")", 1); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__36() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__36() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__22; -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__11; +x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__22; +x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__11; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__37() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__37() { _start: { lean_object* x_1; @@ -3573,7 +3573,7 @@ x_1 = lean_mk_string_from_bytes("then", 4); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__38() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__38() { _start: { lean_object* x_1; @@ -3581,7 +3581,7 @@ x_1 = lean_mk_string_from_bytes("doExpr", 6); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__39() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__39() { _start: { lean_object* x_1; @@ -3589,22 +3589,22 @@ x_1 = lean_mk_string_from_bytes("Lean.addTrace", 13); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__40() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__40() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__39; +x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__39; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__41() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__41() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__39; +x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__39; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__40; +x_3 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__40; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3612,7 +3612,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__42() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__42() { _start: { lean_object* x_1; @@ -3620,7 +3620,7 @@ x_1 = lean_mk_string_from_bytes("addTrace", 8); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__43() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__43() { _start: { lean_object* x_1; lean_object* x_2; @@ -3629,7 +3629,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__44() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__44() { _start: { lean_object* x_1; @@ -3637,7 +3637,7 @@ x_1 = lean_mk_string_from_bytes("quotedName", 10); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__45() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__45() { _start: { lean_object* x_1; @@ -3645,7 +3645,7 @@ x_1 = lean_mk_string_from_bytes(".", 1); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__46() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__46() { _start: { lean_object* x_1; @@ -3653,7 +3653,7 @@ x_1 = lean_mk_string_from_bytes("`", 1); return x_1; } } -LEAN_EXPORT lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1(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; uint8_t x_7; @@ -3669,52 +3669,52 @@ lean_inc(x_9); x_10 = lean_ctor_get(x_4, 1); lean_inc(x_10); lean_dec(x_4); -x_11 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__1; +x_11 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__1; lean_inc(x_1); x_12 = lean_name_mk_string(x_1, x_11); -x_13 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__2; +x_13 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__2; x_14 = lean_name_mk_string(x_12, x_13); -x_15 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__3; +x_15 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__3; lean_inc(x_14); x_16 = lean_name_mk_string(x_14, x_15); -x_17 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__4; +x_17 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__4; lean_inc(x_8); x_18 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_18, 0, x_8); lean_ctor_set(x_18, 1, x_17); -x_19 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__5; +x_19 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__5; lean_inc(x_14); x_20 = lean_name_mk_string(x_14, x_19); -x_21 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__8; +x_21 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__8; lean_inc(x_14); x_22 = lean_name_mk_string(x_14, x_21); -x_23 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__9; +x_23 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__9; lean_inc(x_14); x_24 = lean_name_mk_string(x_14, x_23); -x_25 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__10; +x_25 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__10; lean_inc(x_8); x_26 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_26, 0, x_8); lean_ctor_set(x_26, 1, x_25); -x_27 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__12; +x_27 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__12; lean_inc(x_14); x_28 = lean_name_mk_string(x_14, x_27); -x_29 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__13; +x_29 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__13; lean_inc(x_14); x_30 = lean_name_mk_string(x_14, x_29); -x_31 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__17; +x_31 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__17; lean_inc(x_9); lean_inc(x_10); x_32 = l_Lean_addMacroScope(x_10, x_31, x_9); x_33 = lean_box(0); -x_34 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__16; +x_34 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__16; lean_inc(x_8); x_35 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_35, 0, x_8); lean_ctor_set(x_35, 1, x_34); lean_ctor_set(x_35, 2, x_32); lean_ctor_set(x_35, 3, x_33); -x_36 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__18; +x_36 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__18; lean_inc(x_8); x_37 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_37, 0, x_8); @@ -3723,47 +3723,47 @@ x_38 = l_Lean_Syntax_getId(x_2); x_39 = lean_erase_macro_scopes(x_38); lean_inc(x_39); x_40 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_33, x_39); -x_41 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__19; +x_41 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__19; lean_inc(x_35); x_42 = lean_array_push(x_41, x_35); -x_43 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__11; +x_43 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__11; x_44 = lean_array_push(x_42, x_43); x_45 = lean_array_push(x_44, x_43); x_46 = lean_array_push(x_45, x_37); -x_47 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__21; +x_47 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__21; x_48 = lean_array_push(x_47, x_26); x_49 = lean_array_push(x_48, x_43); -x_50 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__23; +x_50 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__23; lean_inc(x_14); x_51 = lean_name_mk_string(x_14, x_50); -x_52 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__24; +x_52 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__24; lean_inc(x_8); x_53 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_53, 0, x_8); lean_ctor_set(x_53, 1, x_52); -x_54 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__25; +x_54 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__25; lean_inc(x_14); x_55 = lean_name_mk_string(x_14, x_54); -x_56 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__26; +x_56 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__26; lean_inc(x_14); x_57 = lean_name_mk_string(x_14, x_56); -x_58 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__27; +x_58 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__27; lean_inc(x_8); x_59 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_59, 0, x_8); lean_ctor_set(x_59, 1, x_58); -x_60 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__28; +x_60 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__28; lean_inc(x_14); x_61 = lean_name_mk_string(x_14, x_60); -x_62 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__29; +x_62 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__29; lean_inc(x_8); x_63 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_63, 0, x_8); lean_ctor_set(x_63, 1, x_62); -x_64 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__30; +x_64 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__30; lean_inc(x_14); x_65 = lean_name_mk_string(x_14, x_64); -x_66 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__34; +x_66 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__34; lean_inc(x_1); x_67 = lean_name_mk_string(x_1, x_66); lean_inc(x_9); @@ -3776,23 +3776,23 @@ lean_ctor_set(x_69, 1, x_33); x_70 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_70, 0, x_69); lean_ctor_set(x_70, 1, x_33); -x_71 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__33; +x_71 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__33; lean_inc(x_8); x_72 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_72, 0, x_8); lean_ctor_set(x_72, 1, x_71); lean_ctor_set(x_72, 2, x_68); lean_ctor_set(x_72, 3, x_70); -x_73 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__20; +x_73 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__20; lean_inc(x_35); x_74 = lean_array_push(x_73, x_35); x_75 = lean_box(2); -x_76 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__7; +x_76 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__7; x_77 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_77, 0, x_75); lean_ctor_set(x_77, 1, x_76); lean_ctor_set(x_77, 2, x_74); -x_78 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__22; +x_78 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__22; x_79 = lean_array_push(x_78, x_72); x_80 = lean_array_push(x_79, x_77); lean_inc(x_65); @@ -3812,7 +3812,7 @@ x_87 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_87, 0, x_75); lean_ctor_set(x_87, 1, x_76); lean_ctor_set(x_87, 2, x_86); -x_88 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__35; +x_88 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__35; lean_inc(x_8); x_89 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_89, 0, x_8); @@ -3824,21 +3824,21 @@ x_93 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_93, 0, x_75); lean_ctor_set(x_93, 1, x_57); lean_ctor_set(x_93, 2, x_92); -x_94 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__36; +x_94 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__36; x_95 = lean_array_push(x_94, x_93); x_96 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_96, 0, x_75); lean_ctor_set(x_96, 1, x_55); lean_ctor_set(x_96, 2, x_95); -x_97 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__37; +x_97 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__37; lean_inc(x_8); x_98 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_98, 0, x_8); lean_ctor_set(x_98, 1, x_97); -x_99 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__38; +x_99 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__38; lean_inc(x_14); x_100 = lean_name_mk_string(x_14, x_99); -x_101 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__42; +x_101 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__42; x_102 = lean_name_mk_string(x_1, x_101); lean_inc(x_102); x_103 = l_Lean_addMacroScope(x_10, x_102, x_9); @@ -3848,7 +3848,7 @@ lean_ctor_set(x_104, 1, x_33); x_105 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_105, 0, x_104); lean_ctor_set(x_105, 1, x_33); -x_106 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__41; +x_106 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__41; x_107 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_107, 0, x_8); lean_ctor_set(x_107, 1, x_106); @@ -3889,7 +3889,7 @@ x_122 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_122, 0, x_75); lean_ctor_set(x_122, 1, x_20); lean_ctor_set(x_122, 2, x_121); -x_123 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__43; +x_123 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__43; x_124 = lean_array_push(x_123, x_53); x_125 = lean_array_push(x_124, x_96); x_126 = lean_array_push(x_125, x_98); @@ -3912,7 +3912,7 @@ if (lean_obj_tag(x_40) == 0) { lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_dec(x_14); -x_135 = l___private_Init_Meta_0__Lean_quoteNameMk(x_39); +x_135 = l_Lean_quoteNameMk(x_39); x_136 = lean_array_push(x_46, x_135); x_137 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_137, 0, x_75); @@ -3955,462 +3955,478 @@ return x_6; } else { -lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; +lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_dec(x_39); x_152 = lean_ctor_get(x_40, 0); lean_inc(x_152); lean_dec(x_40); -x_153 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__44; +x_153 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__44; x_154 = lean_name_mk_string(x_14, x_153); -x_155 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__45; +x_155 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__45; x_156 = l_String_intercalate(x_155, x_152); -x_157 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__46; +x_157 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__46; x_158 = lean_string_append(x_157, x_156); lean_dec(x_156); -x_159 = l_Lean_nameLitKind; -x_160 = l_Lean_Syntax_mkLit(x_159, x_158, x_75); -x_161 = lean_array_push(x_73, x_160); -x_162 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_162, 0, x_75); -lean_ctor_set(x_162, 1, x_154); -lean_ctor_set(x_162, 2, x_161); -x_163 = lean_array_push(x_46, x_162); -x_164 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_164, 0, x_75); -lean_ctor_set(x_164, 1, x_30); -lean_ctor_set(x_164, 2, x_163); -x_165 = lean_array_push(x_73, x_164); -x_166 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_166, 0, x_75); -lean_ctor_set(x_166, 1, x_28); -lean_ctor_set(x_166, 2, x_165); -x_167 = lean_array_push(x_49, x_166); -x_168 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_168, 0, x_75); -lean_ctor_set(x_168, 1, x_24); -lean_ctor_set(x_168, 2, x_167); -x_169 = lean_array_push(x_78, x_168); -x_170 = lean_array_push(x_169, x_43); -x_171 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_171, 0, x_75); -lean_ctor_set(x_171, 1, x_22); -lean_ctor_set(x_171, 2, x_170); -x_172 = lean_array_push(x_78, x_171); -x_173 = lean_array_push(x_172, x_133); -x_174 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_174, 0, x_75); -lean_ctor_set(x_174, 1, x_76); -lean_ctor_set(x_174, 2, x_173); -x_175 = lean_array_push(x_73, x_174); -x_176 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_176, 0, x_75); -lean_ctor_set(x_176, 1, x_20); -lean_ctor_set(x_176, 2, x_175); -x_177 = lean_array_push(x_134, x_176); -x_178 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_178, 0, x_75); -lean_ctor_set(x_178, 1, x_16); -lean_ctor_set(x_178, 2, x_177); -lean_ctor_set(x_6, 0, x_178); +x_159 = l_Lean_Syntax_mkNameLit(x_158, x_75); +x_160 = lean_array_push(x_73, x_159); +x_161 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_161, 0, x_75); +lean_ctor_set(x_161, 1, x_154); +lean_ctor_set(x_161, 2, x_160); +x_162 = lean_array_push(x_46, x_161); +x_163 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_163, 0, x_75); +lean_ctor_set(x_163, 1, x_30); +lean_ctor_set(x_163, 2, x_162); +x_164 = lean_array_push(x_73, x_163); +x_165 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_165, 0, x_75); +lean_ctor_set(x_165, 1, x_28); +lean_ctor_set(x_165, 2, x_164); +x_166 = lean_array_push(x_49, x_165); +x_167 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_167, 0, x_75); +lean_ctor_set(x_167, 1, x_24); +lean_ctor_set(x_167, 2, x_166); +x_168 = lean_array_push(x_78, x_167); +x_169 = lean_array_push(x_168, x_43); +x_170 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_170, 0, x_75); +lean_ctor_set(x_170, 1, x_22); +lean_ctor_set(x_170, 2, x_169); +x_171 = lean_array_push(x_78, x_170); +x_172 = lean_array_push(x_171, x_133); +x_173 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_173, 0, x_75); +lean_ctor_set(x_173, 1, x_76); +lean_ctor_set(x_173, 2, x_172); +x_174 = lean_array_push(x_73, x_173); +x_175 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_175, 0, x_75); +lean_ctor_set(x_175, 1, x_20); +lean_ctor_set(x_175, 2, x_174); +x_176 = lean_array_push(x_134, x_175); +x_177 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_177, 0, x_75); +lean_ctor_set(x_177, 1, x_16); +lean_ctor_set(x_177, 2, x_176); +lean_ctor_set(x_6, 0, x_177); return x_6; } } else { -lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; -x_179 = lean_ctor_get(x_6, 0); -x_180 = lean_ctor_get(x_6, 1); -lean_inc(x_180); +lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; +x_178 = lean_ctor_get(x_6, 0); +x_179 = lean_ctor_get(x_6, 1); lean_inc(x_179); +lean_inc(x_178); lean_dec(x_6); -x_181 = lean_ctor_get(x_4, 2); +x_180 = lean_ctor_get(x_4, 2); +lean_inc(x_180); +x_181 = lean_ctor_get(x_4, 1); lean_inc(x_181); -x_182 = lean_ctor_get(x_4, 1); -lean_inc(x_182); lean_dec(x_4); -x_183 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__1; +x_182 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__1; lean_inc(x_1); -x_184 = lean_name_mk_string(x_1, x_183); -x_185 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__2; -x_186 = lean_name_mk_string(x_184, x_185); -x_187 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__3; -lean_inc(x_186); -x_188 = lean_name_mk_string(x_186, x_187); -x_189 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__4; -lean_inc(x_179); -x_190 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_190, 0, x_179); -lean_ctor_set(x_190, 1, x_189); -x_191 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__5; -lean_inc(x_186); -x_192 = lean_name_mk_string(x_186, x_191); -x_193 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__8; -lean_inc(x_186); -x_194 = lean_name_mk_string(x_186, x_193); -x_195 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__9; -lean_inc(x_186); -x_196 = lean_name_mk_string(x_186, x_195); -x_197 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__10; -lean_inc(x_179); -x_198 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_198, 0, x_179); -lean_ctor_set(x_198, 1, x_197); -x_199 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__12; -lean_inc(x_186); -x_200 = lean_name_mk_string(x_186, x_199); -x_201 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__13; -lean_inc(x_186); -x_202 = lean_name_mk_string(x_186, x_201); -x_203 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__17; +x_183 = lean_name_mk_string(x_1, x_182); +x_184 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__2; +x_185 = lean_name_mk_string(x_183, x_184); +x_186 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__3; +lean_inc(x_185); +x_187 = lean_name_mk_string(x_185, x_186); +x_188 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__4; +lean_inc(x_178); +x_189 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_189, 0, x_178); +lean_ctor_set(x_189, 1, x_188); +x_190 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__5; +lean_inc(x_185); +x_191 = lean_name_mk_string(x_185, x_190); +x_192 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__8; +lean_inc(x_185); +x_193 = lean_name_mk_string(x_185, x_192); +x_194 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__9; +lean_inc(x_185); +x_195 = lean_name_mk_string(x_185, x_194); +x_196 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__10; +lean_inc(x_178); +x_197 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_197, 0, x_178); +lean_ctor_set(x_197, 1, x_196); +x_198 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__12; +lean_inc(x_185); +x_199 = lean_name_mk_string(x_185, x_198); +x_200 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__13; +lean_inc(x_185); +x_201 = lean_name_mk_string(x_185, x_200); +x_202 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__17; +lean_inc(x_180); lean_inc(x_181); -lean_inc(x_182); -x_204 = l_Lean_addMacroScope(x_182, x_203, x_181); -x_205 = lean_box(0); -x_206 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__16; -lean_inc(x_179); -x_207 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_207, 0, x_179); -lean_ctor_set(x_207, 1, x_206); -lean_ctor_set(x_207, 2, x_204); -lean_ctor_set(x_207, 3, x_205); -x_208 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__18; -lean_inc(x_179); -x_209 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_209, 0, x_179); -lean_ctor_set(x_209, 1, x_208); -x_210 = l_Lean_Syntax_getId(x_2); -x_211 = lean_erase_macro_scopes(x_210); -lean_inc(x_211); -x_212 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_205, x_211); -x_213 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__19; -lean_inc(x_207); -x_214 = lean_array_push(x_213, x_207); -x_215 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__11; -x_216 = lean_array_push(x_214, x_215); -x_217 = lean_array_push(x_216, x_215); -x_218 = lean_array_push(x_217, x_209); -x_219 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__21; -x_220 = lean_array_push(x_219, x_198); -x_221 = lean_array_push(x_220, x_215); -x_222 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__23; -lean_inc(x_186); -x_223 = lean_name_mk_string(x_186, x_222); -x_224 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__24; -lean_inc(x_179); -x_225 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_225, 0, x_179); -lean_ctor_set(x_225, 1, x_224); -x_226 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__25; -lean_inc(x_186); -x_227 = lean_name_mk_string(x_186, x_226); -x_228 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__26; -lean_inc(x_186); -x_229 = lean_name_mk_string(x_186, x_228); -x_230 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__27; -lean_inc(x_179); -x_231 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_231, 0, x_179); -lean_ctor_set(x_231, 1, x_230); -x_232 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__28; -lean_inc(x_186); -x_233 = lean_name_mk_string(x_186, x_232); -x_234 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__29; -lean_inc(x_179); -x_235 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_235, 0, x_179); -lean_ctor_set(x_235, 1, x_234); -x_236 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__30; -lean_inc(x_186); -x_237 = lean_name_mk_string(x_186, x_236); -x_238 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__34; +x_203 = l_Lean_addMacroScope(x_181, x_202, x_180); +x_204 = lean_box(0); +x_205 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__16; +lean_inc(x_178); +x_206 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_206, 0, x_178); +lean_ctor_set(x_206, 1, x_205); +lean_ctor_set(x_206, 2, x_203); +lean_ctor_set(x_206, 3, x_204); +x_207 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__18; +lean_inc(x_178); +x_208 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_208, 0, x_178); +lean_ctor_set(x_208, 1, x_207); +x_209 = l_Lean_Syntax_getId(x_2); +x_210 = lean_erase_macro_scopes(x_209); +lean_inc(x_210); +x_211 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_204, x_210); +x_212 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__19; +lean_inc(x_206); +x_213 = lean_array_push(x_212, x_206); +x_214 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__11; +x_215 = lean_array_push(x_213, x_214); +x_216 = lean_array_push(x_215, x_214); +x_217 = lean_array_push(x_216, x_208); +x_218 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__21; +x_219 = lean_array_push(x_218, x_197); +x_220 = lean_array_push(x_219, x_214); +x_221 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__23; +lean_inc(x_185); +x_222 = lean_name_mk_string(x_185, x_221); +x_223 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__24; +lean_inc(x_178); +x_224 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_224, 0, x_178); +lean_ctor_set(x_224, 1, x_223); +x_225 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__25; +lean_inc(x_185); +x_226 = lean_name_mk_string(x_185, x_225); +x_227 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__26; +lean_inc(x_185); +x_228 = lean_name_mk_string(x_185, x_227); +x_229 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__27; +lean_inc(x_178); +x_230 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_230, 0, x_178); +lean_ctor_set(x_230, 1, x_229); +x_231 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__28; +lean_inc(x_185); +x_232 = lean_name_mk_string(x_185, x_231); +x_233 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__29; +lean_inc(x_178); +x_234 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_234, 0, x_178); +lean_ctor_set(x_234, 1, x_233); +x_235 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__30; +lean_inc(x_185); +x_236 = lean_name_mk_string(x_185, x_235); +x_237 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__34; lean_inc(x_1); -x_239 = lean_name_mk_string(x_1, x_238); +x_238 = lean_name_mk_string(x_1, x_237); +lean_inc(x_180); +lean_inc(x_238); lean_inc(x_181); -lean_inc(x_239); -lean_inc(x_182); -x_240 = l_Lean_addMacroScope(x_182, x_239, x_181); -x_241 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_241, 0, x_239); -lean_ctor_set(x_241, 1, x_205); -x_242 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_242, 0, x_241); -lean_ctor_set(x_242, 1, x_205); -x_243 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__33; -lean_inc(x_179); -x_244 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_244, 0, x_179); -lean_ctor_set(x_244, 1, x_243); -lean_ctor_set(x_244, 2, x_240); -lean_ctor_set(x_244, 3, x_242); -x_245 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__20; -lean_inc(x_207); -x_246 = lean_array_push(x_245, x_207); -x_247 = lean_box(2); -x_248 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__7; -x_249 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_249, 0, x_247); -lean_ctor_set(x_249, 1, x_248); -lean_ctor_set(x_249, 2, x_246); -x_250 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__22; -x_251 = lean_array_push(x_250, x_244); -x_252 = lean_array_push(x_251, x_249); -lean_inc(x_237); -x_253 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_253, 0, x_247); -lean_ctor_set(x_253, 1, x_237); -lean_ctor_set(x_253, 2, x_252); -x_254 = lean_array_push(x_250, x_235); -x_255 = lean_array_push(x_254, x_253); -x_256 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_256, 0, x_247); -lean_ctor_set(x_256, 1, x_233); -lean_ctor_set(x_256, 2, x_255); -x_257 = lean_array_push(x_250, x_256); -x_258 = lean_array_push(x_257, x_215); -x_259 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_259, 0, x_247); -lean_ctor_set(x_259, 1, x_248); -lean_ctor_set(x_259, 2, x_258); -x_260 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__35; -lean_inc(x_179); -x_261 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_261, 0, x_179); -lean_ctor_set(x_261, 1, x_260); -x_262 = lean_array_push(x_219, x_231); -x_263 = lean_array_push(x_262, x_259); -x_264 = lean_array_push(x_263, x_261); -x_265 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_265, 0, x_247); -lean_ctor_set(x_265, 1, x_229); -lean_ctor_set(x_265, 2, x_264); -x_266 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__36; -x_267 = lean_array_push(x_266, x_265); -x_268 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_268, 0, x_247); -lean_ctor_set(x_268, 1, x_227); -lean_ctor_set(x_268, 2, x_267); -x_269 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__37; -lean_inc(x_179); -x_270 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_270, 0, x_179); -lean_ctor_set(x_270, 1, x_269); -x_271 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__38; -lean_inc(x_186); -x_272 = lean_name_mk_string(x_186, x_271); -x_273 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__42; -x_274 = lean_name_mk_string(x_1, x_273); -lean_inc(x_274); -x_275 = l_Lean_addMacroScope(x_182, x_274, x_181); -x_276 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_276, 0, x_274); -lean_ctor_set(x_276, 1, x_205); -x_277 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_277, 0, x_276); -lean_ctor_set(x_277, 1, x_205); -x_278 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__41; -x_279 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_279, 0, x_179); -lean_ctor_set(x_279, 1, x_278); -lean_ctor_set(x_279, 2, x_275); -lean_ctor_set(x_279, 3, x_277); -x_280 = lean_array_push(x_250, x_207); -x_281 = lean_array_push(x_280, x_3); -x_282 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_282, 0, x_247); -lean_ctor_set(x_282, 1, x_248); -lean_ctor_set(x_282, 2, x_281); -x_283 = lean_array_push(x_250, x_279); -x_284 = lean_array_push(x_283, x_282); -x_285 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_285, 0, x_247); -lean_ctor_set(x_285, 1, x_237); -lean_ctor_set(x_285, 2, x_284); -x_286 = lean_array_push(x_245, x_285); -x_287 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_287, 0, x_247); -lean_ctor_set(x_287, 1, x_272); -lean_ctor_set(x_287, 2, x_286); -x_288 = lean_array_push(x_250, x_287); -x_289 = lean_array_push(x_288, x_215); -lean_inc(x_194); -x_290 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_290, 0, x_247); -lean_ctor_set(x_290, 1, x_194); -lean_ctor_set(x_290, 2, x_289); -x_291 = lean_array_push(x_245, x_290); -x_292 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_292, 0, x_247); -lean_ctor_set(x_292, 1, x_248); -lean_ctor_set(x_292, 2, x_291); -x_293 = lean_array_push(x_245, x_292); -lean_inc(x_192); -x_294 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_294, 0, x_247); -lean_ctor_set(x_294, 1, x_192); -lean_ctor_set(x_294, 2, x_293); -x_295 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__43; -x_296 = lean_array_push(x_295, x_225); -x_297 = lean_array_push(x_296, x_268); -x_298 = lean_array_push(x_297, x_270); -x_299 = lean_array_push(x_298, x_294); -x_300 = lean_array_push(x_299, x_215); -x_301 = lean_array_push(x_300, x_215); -x_302 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_302, 0, x_247); -lean_ctor_set(x_302, 1, x_223); -lean_ctor_set(x_302, 2, x_301); -x_303 = lean_array_push(x_250, x_302); -x_304 = lean_array_push(x_303, x_215); -lean_inc(x_194); -x_305 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_305, 0, x_247); -lean_ctor_set(x_305, 1, x_194); -lean_ctor_set(x_305, 2, x_304); -x_306 = lean_array_push(x_250, x_190); -if (lean_obj_tag(x_212) == 0) -{ -lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; -lean_dec(x_186); -x_307 = l___private_Init_Meta_0__Lean_quoteNameMk(x_211); -x_308 = lean_array_push(x_218, x_307); -x_309 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_309, 0, x_247); -lean_ctor_set(x_309, 1, x_202); -lean_ctor_set(x_309, 2, x_308); -x_310 = lean_array_push(x_245, x_309); -x_311 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_311, 0, x_247); -lean_ctor_set(x_311, 1, x_200); -lean_ctor_set(x_311, 2, x_310); -x_312 = lean_array_push(x_221, x_311); -x_313 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_313, 0, x_247); -lean_ctor_set(x_313, 1, x_196); -lean_ctor_set(x_313, 2, x_312); -x_314 = lean_array_push(x_250, x_313); -x_315 = lean_array_push(x_314, x_215); -x_316 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_316, 0, x_247); -lean_ctor_set(x_316, 1, x_194); -lean_ctor_set(x_316, 2, x_315); -x_317 = lean_array_push(x_250, x_316); -x_318 = lean_array_push(x_317, x_305); -x_319 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_319, 0, x_247); -lean_ctor_set(x_319, 1, x_248); -lean_ctor_set(x_319, 2, x_318); -x_320 = lean_array_push(x_245, x_319); -x_321 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_321, 0, x_247); -lean_ctor_set(x_321, 1, x_192); -lean_ctor_set(x_321, 2, x_320); -x_322 = lean_array_push(x_306, x_321); -x_323 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_323, 0, x_247); -lean_ctor_set(x_323, 1, x_188); -lean_ctor_set(x_323, 2, x_322); -x_324 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_324, 0, x_323); -lean_ctor_set(x_324, 1, x_180); -return x_324; +x_239 = l_Lean_addMacroScope(x_181, x_238, x_180); +x_240 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_240, 0, x_238); +lean_ctor_set(x_240, 1, x_204); +x_241 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_241, 0, x_240); +lean_ctor_set(x_241, 1, x_204); +x_242 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__33; +lean_inc(x_178); +x_243 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_243, 0, x_178); +lean_ctor_set(x_243, 1, x_242); +lean_ctor_set(x_243, 2, x_239); +lean_ctor_set(x_243, 3, x_241); +x_244 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__20; +lean_inc(x_206); +x_245 = lean_array_push(x_244, x_206); +x_246 = lean_box(2); +x_247 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__7; +x_248 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_248, 0, x_246); +lean_ctor_set(x_248, 1, x_247); +lean_ctor_set(x_248, 2, x_245); +x_249 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__22; +x_250 = lean_array_push(x_249, x_243); +x_251 = lean_array_push(x_250, x_248); +lean_inc(x_236); +x_252 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_252, 0, x_246); +lean_ctor_set(x_252, 1, x_236); +lean_ctor_set(x_252, 2, x_251); +x_253 = lean_array_push(x_249, x_234); +x_254 = lean_array_push(x_253, x_252); +x_255 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_255, 0, x_246); +lean_ctor_set(x_255, 1, x_232); +lean_ctor_set(x_255, 2, x_254); +x_256 = lean_array_push(x_249, x_255); +x_257 = lean_array_push(x_256, x_214); +x_258 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_258, 0, x_246); +lean_ctor_set(x_258, 1, x_247); +lean_ctor_set(x_258, 2, x_257); +x_259 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__35; +lean_inc(x_178); +x_260 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_260, 0, x_178); +lean_ctor_set(x_260, 1, x_259); +x_261 = lean_array_push(x_218, x_230); +x_262 = lean_array_push(x_261, x_258); +x_263 = lean_array_push(x_262, x_260); +x_264 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_264, 0, x_246); +lean_ctor_set(x_264, 1, x_228); +lean_ctor_set(x_264, 2, x_263); +x_265 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__36; +x_266 = lean_array_push(x_265, x_264); +x_267 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_267, 0, x_246); +lean_ctor_set(x_267, 1, x_226); +lean_ctor_set(x_267, 2, x_266); +x_268 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__37; +lean_inc(x_178); +x_269 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_269, 0, x_178); +lean_ctor_set(x_269, 1, x_268); +x_270 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__38; +lean_inc(x_185); +x_271 = lean_name_mk_string(x_185, x_270); +x_272 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__42; +x_273 = lean_name_mk_string(x_1, x_272); +lean_inc(x_273); +x_274 = l_Lean_addMacroScope(x_181, x_273, x_180); +x_275 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_275, 0, x_273); +lean_ctor_set(x_275, 1, x_204); +x_276 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_276, 0, x_275); +lean_ctor_set(x_276, 1, x_204); +x_277 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__41; +x_278 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_278, 0, x_178); +lean_ctor_set(x_278, 1, x_277); +lean_ctor_set(x_278, 2, x_274); +lean_ctor_set(x_278, 3, x_276); +x_279 = lean_array_push(x_249, x_206); +x_280 = lean_array_push(x_279, x_3); +x_281 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_281, 0, x_246); +lean_ctor_set(x_281, 1, x_247); +lean_ctor_set(x_281, 2, x_280); +x_282 = lean_array_push(x_249, x_278); +x_283 = lean_array_push(x_282, x_281); +x_284 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_284, 0, x_246); +lean_ctor_set(x_284, 1, x_236); +lean_ctor_set(x_284, 2, x_283); +x_285 = lean_array_push(x_244, x_284); +x_286 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_286, 0, x_246); +lean_ctor_set(x_286, 1, x_271); +lean_ctor_set(x_286, 2, x_285); +x_287 = lean_array_push(x_249, x_286); +x_288 = lean_array_push(x_287, x_214); +lean_inc(x_193); +x_289 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_289, 0, x_246); +lean_ctor_set(x_289, 1, x_193); +lean_ctor_set(x_289, 2, x_288); +x_290 = lean_array_push(x_244, x_289); +x_291 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_291, 0, x_246); +lean_ctor_set(x_291, 1, x_247); +lean_ctor_set(x_291, 2, x_290); +x_292 = lean_array_push(x_244, x_291); +lean_inc(x_191); +x_293 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_293, 0, x_246); +lean_ctor_set(x_293, 1, x_191); +lean_ctor_set(x_293, 2, x_292); +x_294 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__43; +x_295 = lean_array_push(x_294, x_224); +x_296 = lean_array_push(x_295, x_267); +x_297 = lean_array_push(x_296, x_269); +x_298 = lean_array_push(x_297, x_293); +x_299 = lean_array_push(x_298, x_214); +x_300 = lean_array_push(x_299, x_214); +x_301 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_301, 0, x_246); +lean_ctor_set(x_301, 1, x_222); +lean_ctor_set(x_301, 2, x_300); +x_302 = lean_array_push(x_249, x_301); +x_303 = lean_array_push(x_302, x_214); +lean_inc(x_193); +x_304 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_304, 0, x_246); +lean_ctor_set(x_304, 1, x_193); +lean_ctor_set(x_304, 2, x_303); +x_305 = lean_array_push(x_249, x_189); +if (lean_obj_tag(x_211) == 0) +{ +lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; +lean_dec(x_185); +x_306 = l_Lean_quoteNameMk(x_210); +x_307 = lean_array_push(x_217, x_306); +x_308 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_308, 0, x_246); +lean_ctor_set(x_308, 1, x_201); +lean_ctor_set(x_308, 2, x_307); +x_309 = lean_array_push(x_244, x_308); +x_310 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_310, 0, x_246); +lean_ctor_set(x_310, 1, x_199); +lean_ctor_set(x_310, 2, x_309); +x_311 = lean_array_push(x_220, x_310); +x_312 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_312, 0, x_246); +lean_ctor_set(x_312, 1, x_195); +lean_ctor_set(x_312, 2, x_311); +x_313 = lean_array_push(x_249, x_312); +x_314 = lean_array_push(x_313, x_214); +x_315 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_315, 0, x_246); +lean_ctor_set(x_315, 1, x_193); +lean_ctor_set(x_315, 2, x_314); +x_316 = lean_array_push(x_249, x_315); +x_317 = lean_array_push(x_316, x_304); +x_318 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_318, 0, x_246); +lean_ctor_set(x_318, 1, x_247); +lean_ctor_set(x_318, 2, x_317); +x_319 = lean_array_push(x_244, x_318); +x_320 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_320, 0, x_246); +lean_ctor_set(x_320, 1, x_191); +lean_ctor_set(x_320, 2, x_319); +x_321 = lean_array_push(x_305, x_320); +x_322 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_322, 0, x_246); +lean_ctor_set(x_322, 1, x_187); +lean_ctor_set(x_322, 2, x_321); +x_323 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_323, 0, x_322); +lean_ctor_set(x_323, 1, x_179); +return x_323; } else { -lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; +lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; +lean_dec(x_210); +x_324 = lean_ctor_get(x_211, 0); +lean_inc(x_324); lean_dec(x_211); -x_325 = lean_ctor_get(x_212, 0); -lean_inc(x_325); -lean_dec(x_212); -x_326 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__44; -x_327 = lean_name_mk_string(x_186, x_326); -x_328 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__45; -x_329 = l_String_intercalate(x_328, x_325); -x_330 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__46; -x_331 = lean_string_append(x_330, x_329); -lean_dec(x_329); -x_332 = l_Lean_nameLitKind; -x_333 = l_Lean_Syntax_mkLit(x_332, x_331, x_247); -x_334 = lean_array_push(x_245, x_333); +x_325 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__44; +x_326 = lean_name_mk_string(x_185, x_325); +x_327 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__45; +x_328 = l_String_intercalate(x_327, x_324); +x_329 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__46; +x_330 = lean_string_append(x_329, x_328); +lean_dec(x_328); +x_331 = l_Lean_Syntax_mkNameLit(x_330, x_246); +x_332 = lean_array_push(x_244, x_331); +x_333 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_333, 0, x_246); +lean_ctor_set(x_333, 1, x_326); +lean_ctor_set(x_333, 2, x_332); +x_334 = lean_array_push(x_217, x_333); x_335 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_335, 0, x_247); -lean_ctor_set(x_335, 1, x_327); +lean_ctor_set(x_335, 0, x_246); +lean_ctor_set(x_335, 1, x_201); lean_ctor_set(x_335, 2, x_334); -x_336 = lean_array_push(x_218, x_335); +x_336 = lean_array_push(x_244, x_335); x_337 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_337, 0, x_247); -lean_ctor_set(x_337, 1, x_202); +lean_ctor_set(x_337, 0, x_246); +lean_ctor_set(x_337, 1, x_199); lean_ctor_set(x_337, 2, x_336); -x_338 = lean_array_push(x_245, x_337); +x_338 = lean_array_push(x_220, x_337); x_339 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_339, 0, x_247); -lean_ctor_set(x_339, 1, x_200); +lean_ctor_set(x_339, 0, x_246); +lean_ctor_set(x_339, 1, x_195); lean_ctor_set(x_339, 2, x_338); -x_340 = lean_array_push(x_221, x_339); -x_341 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_341, 0, x_247); -lean_ctor_set(x_341, 1, x_196); -lean_ctor_set(x_341, 2, x_340); -x_342 = lean_array_push(x_250, x_341); -x_343 = lean_array_push(x_342, x_215); -x_344 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_344, 0, x_247); -lean_ctor_set(x_344, 1, x_194); -lean_ctor_set(x_344, 2, x_343); -x_345 = lean_array_push(x_250, x_344); -x_346 = lean_array_push(x_345, x_305); +x_340 = lean_array_push(x_249, x_339); +x_341 = lean_array_push(x_340, x_214); +x_342 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_342, 0, x_246); +lean_ctor_set(x_342, 1, x_193); +lean_ctor_set(x_342, 2, x_341); +x_343 = lean_array_push(x_249, x_342); +x_344 = lean_array_push(x_343, x_304); +x_345 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_345, 0, x_246); +lean_ctor_set(x_345, 1, x_247); +lean_ctor_set(x_345, 2, x_344); +x_346 = lean_array_push(x_244, x_345); x_347 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_347, 0, x_247); -lean_ctor_set(x_347, 1, x_248); +lean_ctor_set(x_347, 0, x_246); +lean_ctor_set(x_347, 1, x_191); lean_ctor_set(x_347, 2, x_346); -x_348 = lean_array_push(x_245, x_347); +x_348 = lean_array_push(x_305, x_347); x_349 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_349, 0, x_247); -lean_ctor_set(x_349, 1, x_192); +lean_ctor_set(x_349, 0, x_246); +lean_ctor_set(x_349, 1, x_187); lean_ctor_set(x_349, 2, x_348); -x_350 = lean_array_push(x_306, x_349); -x_351 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_351, 0, x_247); -lean_ctor_set(x_351, 1, x_188); -lean_ctor_set(x_351, 2, x_350); -x_352 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_352, 0, x_351); -lean_ctor_set(x_352, 1, x_180); -return x_352; +x_350 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_350, 0, x_349); +lean_ctor_set(x_350, 1, x_179); +return x_350; } } } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__1() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("interpolatedStrKind", 19); +return x_1; +} +} +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__1; +x_3 = lean_name_mk_string(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_doElemTrace_x5b_____x5d_______closed__2; -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__1; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__2; +x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__1; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__2() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__1; -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__2; +x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__3; +x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__2; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__3() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__2; -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__26; +x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__4; +x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__26; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__4() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__6() { _start: { lean_object* x_1; @@ -4418,17 +4434,17 @@ x_1 = lean_mk_string_from_bytes("typeAscription", 14); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__5() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__2; -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__4; +x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__4; +x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__6; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__6() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__8() { _start: { lean_object* x_1; @@ -4436,7 +4452,7 @@ x_1 = lean_mk_string_from_bytes(":", 1); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__7() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__9() { _start: { lean_object* x_1; @@ -4444,22 +4460,22 @@ x_1 = lean_mk_string_from_bytes("MessageData", 11); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__8() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__7; +x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__9; x_2 = lean_string_utf8_byte_size(x_1); return x_2; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__9() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__7; +x_1 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__9; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__8; +x_3 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__10; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -4467,51 +4483,51 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__10() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__7; +x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__9; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__11() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_doElemTrace_x5b_____x5d_______closed__2; -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__7; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__2; +x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__9; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__12() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__11; +x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__13; 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___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__13() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__12; +x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__14; x_3 = lean_alloc_ctor(1, 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___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__14() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__16() { _start: { lean_object* x_1; @@ -4519,17 +4535,17 @@ x_1 = lean_mk_string_from_bytes("termM!_", 7); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__15() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_doElemTrace_x5b_____x5d_______closed__2; -x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__14; +x_1 = l_Lean_doElemTrace_x5b___x5d_______closed__2; +x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__16; x_3 = lean_name_mk_string(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__16() { +static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__18() { _start: { lean_object* x_1; @@ -4537,11 +4553,11 @@ x_1 = lean_mk_string_from_bytes("m!", 2); return x_1; } } -LEAN_EXPORT lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; uint8_t x_5; -x_4 = l_Lean_doElemTrace_x5b_____x5d_______closed__4; +x_4 = l_Lean_doElemTrace_x5b___x5d_______closed__4; lean_inc(x_1); x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); if (x_5 == 0) @@ -4565,7 +4581,7 @@ x_11 = l_Lean_Syntax_getArg(x_1, x_10); lean_dec(x_1); lean_inc(x_11); x_12 = l_Lean_Syntax_getKind(x_11); -x_13 = l_Lean_interpolatedStrKind; +x_13 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__2; x_14 = lean_name_eq(x_12, x_13); lean_dec(x_12); if (x_14 == 0) @@ -4582,38 +4598,38 @@ x_18 = lean_ctor_get(x_2, 2); lean_inc(x_18); x_19 = lean_ctor_get(x_2, 1); lean_inc(x_19); -x_20 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__27; +x_20 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__27; lean_inc(x_16); x_21 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_21, 0, x_16); lean_ctor_set(x_21, 1, x_20); -x_22 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__6; +x_22 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__8; lean_inc(x_16); x_23 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_23, 0, x_16); lean_ctor_set(x_23, 1, x_22); -x_24 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__10; +x_24 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__12; x_25 = l_Lean_addMacroScope(x_19, x_24, x_18); -x_26 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__9; -x_27 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__13; +x_26 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__11; +x_27 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__15; lean_inc(x_16); x_28 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_28, 0, x_16); lean_ctor_set(x_28, 1, x_26); lean_ctor_set(x_28, 2, x_25); lean_ctor_set(x_28, 3, x_27); -x_29 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__22; +x_29 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__22; x_30 = lean_array_push(x_29, x_23); x_31 = lean_array_push(x_30, x_28); x_32 = lean_box(2); -x_33 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__5; +x_33 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__7; x_34 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_34, 0, x_32); lean_ctor_set(x_34, 1, x_33); lean_ctor_set(x_34, 2, x_31); -x_35 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__20; +x_35 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__20; x_36 = lean_array_push(x_35, x_34); -x_37 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__7; +x_37 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__7; x_38 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_38, 0, x_32); lean_ctor_set(x_38, 1, x_37); @@ -4624,21 +4640,21 @@ x_41 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_41, 0, x_32); lean_ctor_set(x_41, 1, x_37); lean_ctor_set(x_41, 2, x_40); -x_42 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__35; +x_42 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__35; x_43 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_43, 0, x_16); lean_ctor_set(x_43, 1, x_42); -x_44 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__21; +x_44 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__21; x_45 = lean_array_push(x_44, x_21); x_46 = lean_array_push(x_45, x_41); x_47 = lean_array_push(x_46, x_43); -x_48 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__3; +x_48 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__5; x_49 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_49, 0, x_32); lean_ctor_set(x_49, 1, x_48); lean_ctor_set(x_49, 2, x_47); -x_50 = l_Lean_doElemTrace_x5b_____x5d_______closed__2; -x_51 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1(x_50, x_9, x_49, x_2, x_17); +x_50 = l_Lean_doElemTrace_x5b___x5d_______closed__2; +x_51 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1(x_50, x_9, x_49, x_2, x_17); lean_dec(x_9); return x_51; } @@ -4652,32 +4668,32 @@ lean_inc(x_53); x_54 = lean_ctor_get(x_52, 1); lean_inc(x_54); lean_dec(x_52); -x_55 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__16; +x_55 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__18; x_56 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_56, 0, x_53); lean_ctor_set(x_56, 1, x_55); -x_57 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__22; +x_57 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__22; x_58 = lean_array_push(x_57, x_56); x_59 = lean_array_push(x_58, x_11); x_60 = lean_box(2); -x_61 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__15; +x_61 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__17; x_62 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_62, 0, x_60); lean_ctor_set(x_62, 1, x_61); lean_ctor_set(x_62, 2, x_59); -x_63 = l_Lean_doElemTrace_x5b_____x5d_______closed__2; -x_64 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1(x_63, x_9, x_62, x_2, x_54); +x_63 = l_Lean_doElemTrace_x5b___x5d_______closed__2; +x_64 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1(x_63, x_9, x_62, x_2, x_54); lean_dec(x_9); return x_64; } } } } -LEAN_EXPORT lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___boxed(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; -x_6 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1(x_1, x_2, x_3, x_4, x_5); lean_dec(x_2); return x_6; } @@ -5526,184 +5542,188 @@ l_Lean_registerTraceClass___closed__2 = _init_l_Lean_registerTraceClass___closed lean_mark_persistent(l_Lean_registerTraceClass___closed__2); l_Lean_registerTraceClass___closed__3 = _init_l_Lean_registerTraceClass___closed__3(); lean_mark_persistent(l_Lean_registerTraceClass___closed__3); -l_Lean_doElemTrace_x5b_____x5d_______closed__1 = _init_l_Lean_doElemTrace_x5b_____x5d_______closed__1(); -lean_mark_persistent(l_Lean_doElemTrace_x5b_____x5d_______closed__1); -l_Lean_doElemTrace_x5b_____x5d_______closed__2 = _init_l_Lean_doElemTrace_x5b_____x5d_______closed__2(); -lean_mark_persistent(l_Lean_doElemTrace_x5b_____x5d_______closed__2); -l_Lean_doElemTrace_x5b_____x5d_______closed__3 = _init_l_Lean_doElemTrace_x5b_____x5d_______closed__3(); -lean_mark_persistent(l_Lean_doElemTrace_x5b_____x5d_______closed__3); -l_Lean_doElemTrace_x5b_____x5d_______closed__4 = _init_l_Lean_doElemTrace_x5b_____x5d_______closed__4(); -lean_mark_persistent(l_Lean_doElemTrace_x5b_____x5d_______closed__4); -l_Lean_doElemTrace_x5b_____x5d_______closed__5 = _init_l_Lean_doElemTrace_x5b_____x5d_______closed__5(); -lean_mark_persistent(l_Lean_doElemTrace_x5b_____x5d_______closed__5); -l_Lean_doElemTrace_x5b_____x5d_______closed__6 = _init_l_Lean_doElemTrace_x5b_____x5d_______closed__6(); -lean_mark_persistent(l_Lean_doElemTrace_x5b_____x5d_______closed__6); -l_Lean_doElemTrace_x5b_____x5d_______closed__7 = _init_l_Lean_doElemTrace_x5b_____x5d_______closed__7(); -lean_mark_persistent(l_Lean_doElemTrace_x5b_____x5d_______closed__7); -l_Lean_doElemTrace_x5b_____x5d_______closed__8 = _init_l_Lean_doElemTrace_x5b_____x5d_______closed__8(); -lean_mark_persistent(l_Lean_doElemTrace_x5b_____x5d_______closed__8); -l_Lean_doElemTrace_x5b_____x5d_______closed__9 = _init_l_Lean_doElemTrace_x5b_____x5d_______closed__9(); -lean_mark_persistent(l_Lean_doElemTrace_x5b_____x5d_______closed__9); -l_Lean_doElemTrace_x5b_____x5d_______closed__10 = _init_l_Lean_doElemTrace_x5b_____x5d_______closed__10(); -lean_mark_persistent(l_Lean_doElemTrace_x5b_____x5d_______closed__10); -l_Lean_doElemTrace_x5b_____x5d_______closed__11 = _init_l_Lean_doElemTrace_x5b_____x5d_______closed__11(); -lean_mark_persistent(l_Lean_doElemTrace_x5b_____x5d_______closed__11); -l_Lean_doElemTrace_x5b_____x5d_______closed__12 = _init_l_Lean_doElemTrace_x5b_____x5d_______closed__12(); -lean_mark_persistent(l_Lean_doElemTrace_x5b_____x5d_______closed__12); -l_Lean_doElemTrace_x5b_____x5d_______closed__13 = _init_l_Lean_doElemTrace_x5b_____x5d_______closed__13(); -lean_mark_persistent(l_Lean_doElemTrace_x5b_____x5d_______closed__13); -l_Lean_doElemTrace_x5b_____x5d_______closed__14 = _init_l_Lean_doElemTrace_x5b_____x5d_______closed__14(); -lean_mark_persistent(l_Lean_doElemTrace_x5b_____x5d_______closed__14); -l_Lean_doElemTrace_x5b_____x5d_______closed__15 = _init_l_Lean_doElemTrace_x5b_____x5d_______closed__15(); -lean_mark_persistent(l_Lean_doElemTrace_x5b_____x5d_______closed__15); -l_Lean_doElemTrace_x5b_____x5d_______closed__16 = _init_l_Lean_doElemTrace_x5b_____x5d_______closed__16(); -lean_mark_persistent(l_Lean_doElemTrace_x5b_____x5d_______closed__16); -l_Lean_doElemTrace_x5b_____x5d_______closed__17 = _init_l_Lean_doElemTrace_x5b_____x5d_______closed__17(); -lean_mark_persistent(l_Lean_doElemTrace_x5b_____x5d_______closed__17); -l_Lean_doElemTrace_x5b_____x5d_______closed__18 = _init_l_Lean_doElemTrace_x5b_____x5d_______closed__18(); -lean_mark_persistent(l_Lean_doElemTrace_x5b_____x5d_______closed__18); -l_Lean_doElemTrace_x5b_____x5d_______closed__19 = _init_l_Lean_doElemTrace_x5b_____x5d_______closed__19(); -lean_mark_persistent(l_Lean_doElemTrace_x5b_____x5d_______closed__19); -l_Lean_doElemTrace_x5b_____x5d_______closed__20 = _init_l_Lean_doElemTrace_x5b_____x5d_______closed__20(); -lean_mark_persistent(l_Lean_doElemTrace_x5b_____x5d_______closed__20); -l_Lean_doElemTrace_x5b_____x5d_______closed__21 = _init_l_Lean_doElemTrace_x5b_____x5d_______closed__21(); -lean_mark_persistent(l_Lean_doElemTrace_x5b_____x5d_______closed__21); -l_Lean_doElemTrace_x5b_____x5d_______closed__22 = _init_l_Lean_doElemTrace_x5b_____x5d_______closed__22(); -lean_mark_persistent(l_Lean_doElemTrace_x5b_____x5d_______closed__22); -l_Lean_doElemTrace_x5b_____x5d_______closed__23 = _init_l_Lean_doElemTrace_x5b_____x5d_______closed__23(); -lean_mark_persistent(l_Lean_doElemTrace_x5b_____x5d_______closed__23); -l_Lean_doElemTrace_x5b_____x5d_______closed__24 = _init_l_Lean_doElemTrace_x5b_____x5d_______closed__24(); -lean_mark_persistent(l_Lean_doElemTrace_x5b_____x5d_______closed__24); -l_Lean_doElemTrace_x5b_____x5d_______closed__25 = _init_l_Lean_doElemTrace_x5b_____x5d_______closed__25(); -lean_mark_persistent(l_Lean_doElemTrace_x5b_____x5d_______closed__25); -l_Lean_doElemTrace_x5b_____x5d_______closed__26 = _init_l_Lean_doElemTrace_x5b_____x5d_______closed__26(); -lean_mark_persistent(l_Lean_doElemTrace_x5b_____x5d_______closed__26); -l_Lean_doElemTrace_x5b_____x5d____ = _init_l_Lean_doElemTrace_x5b_____x5d____(); -lean_mark_persistent(l_Lean_doElemTrace_x5b_____x5d____); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__1 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__1(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__1); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__2 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__2(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__2); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__3 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__3(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__3); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__4 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__4(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__4); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__5 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__5(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__5); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__6 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__6(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__6); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__7 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__7(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__7); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__8 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__8(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__8); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__9 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__9(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__9); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__10 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__10(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__10); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__11 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__11(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__11); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__12 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__12(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__12); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__13 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__13(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__13); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__14 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__14(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__14); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__15 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__15(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__15); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__16 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__16(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__16); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__17 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__17(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__17); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__18 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__18(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__18); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__19 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__19(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__19); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__20 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__20(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__20); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__21 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__21(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__21); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__22 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__22(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__22); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__23 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__23(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__23); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__24 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__24(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__24); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__25 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__25(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__25); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__26 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__26(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__26); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__27 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__27(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__27); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__28 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__28(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__28); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__29 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__29(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__29); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__30 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__30(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__30); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__31 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__31(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__31); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__32 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__32(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__32); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__33 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__33(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__33); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__34 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__34(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__34); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__35 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__35(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__35); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__36 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__36(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__36); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__37 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__37(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__37); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__38 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__38(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__38); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__39 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__39(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__39); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__40 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__40(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__40); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__41 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__41(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__41); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__42 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__42(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__42); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__43 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__43(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__43); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__44 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__44(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__44); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__45 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__45(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__45); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__46 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__46(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___lambda__1___closed__46); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__1 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__1(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__1); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__2 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__2(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__2); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__3 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__3(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__3); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__4 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__4(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__4); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__5 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__5(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__5); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__6 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__6(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__6); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__7 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__7(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__7); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__8 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__8(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__8); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__9 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__9(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__9); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__10 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__10(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__10); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__11 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__11(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__11); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__12 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__12(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__12); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__13 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__13(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__13); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__14 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__14(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__14); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__15 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__15(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__15); -l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__16 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__16(); -lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b_____x5d______1___closed__16); +l_Lean_doElemTrace_x5b___x5d_______closed__1 = _init_l_Lean_doElemTrace_x5b___x5d_______closed__1(); +lean_mark_persistent(l_Lean_doElemTrace_x5b___x5d_______closed__1); +l_Lean_doElemTrace_x5b___x5d_______closed__2 = _init_l_Lean_doElemTrace_x5b___x5d_______closed__2(); +lean_mark_persistent(l_Lean_doElemTrace_x5b___x5d_______closed__2); +l_Lean_doElemTrace_x5b___x5d_______closed__3 = _init_l_Lean_doElemTrace_x5b___x5d_______closed__3(); +lean_mark_persistent(l_Lean_doElemTrace_x5b___x5d_______closed__3); +l_Lean_doElemTrace_x5b___x5d_______closed__4 = _init_l_Lean_doElemTrace_x5b___x5d_______closed__4(); +lean_mark_persistent(l_Lean_doElemTrace_x5b___x5d_______closed__4); +l_Lean_doElemTrace_x5b___x5d_______closed__5 = _init_l_Lean_doElemTrace_x5b___x5d_______closed__5(); +lean_mark_persistent(l_Lean_doElemTrace_x5b___x5d_______closed__5); +l_Lean_doElemTrace_x5b___x5d_______closed__6 = _init_l_Lean_doElemTrace_x5b___x5d_______closed__6(); +lean_mark_persistent(l_Lean_doElemTrace_x5b___x5d_______closed__6); +l_Lean_doElemTrace_x5b___x5d_______closed__7 = _init_l_Lean_doElemTrace_x5b___x5d_______closed__7(); +lean_mark_persistent(l_Lean_doElemTrace_x5b___x5d_______closed__7); +l_Lean_doElemTrace_x5b___x5d_______closed__8 = _init_l_Lean_doElemTrace_x5b___x5d_______closed__8(); +lean_mark_persistent(l_Lean_doElemTrace_x5b___x5d_______closed__8); +l_Lean_doElemTrace_x5b___x5d_______closed__9 = _init_l_Lean_doElemTrace_x5b___x5d_______closed__9(); +lean_mark_persistent(l_Lean_doElemTrace_x5b___x5d_______closed__9); +l_Lean_doElemTrace_x5b___x5d_______closed__10 = _init_l_Lean_doElemTrace_x5b___x5d_______closed__10(); +lean_mark_persistent(l_Lean_doElemTrace_x5b___x5d_______closed__10); +l_Lean_doElemTrace_x5b___x5d_______closed__11 = _init_l_Lean_doElemTrace_x5b___x5d_______closed__11(); +lean_mark_persistent(l_Lean_doElemTrace_x5b___x5d_______closed__11); +l_Lean_doElemTrace_x5b___x5d_______closed__12 = _init_l_Lean_doElemTrace_x5b___x5d_______closed__12(); +lean_mark_persistent(l_Lean_doElemTrace_x5b___x5d_______closed__12); +l_Lean_doElemTrace_x5b___x5d_______closed__13 = _init_l_Lean_doElemTrace_x5b___x5d_______closed__13(); +lean_mark_persistent(l_Lean_doElemTrace_x5b___x5d_______closed__13); +l_Lean_doElemTrace_x5b___x5d_______closed__14 = _init_l_Lean_doElemTrace_x5b___x5d_______closed__14(); +lean_mark_persistent(l_Lean_doElemTrace_x5b___x5d_______closed__14); +l_Lean_doElemTrace_x5b___x5d_______closed__15 = _init_l_Lean_doElemTrace_x5b___x5d_______closed__15(); +lean_mark_persistent(l_Lean_doElemTrace_x5b___x5d_______closed__15); +l_Lean_doElemTrace_x5b___x5d_______closed__16 = _init_l_Lean_doElemTrace_x5b___x5d_______closed__16(); +lean_mark_persistent(l_Lean_doElemTrace_x5b___x5d_______closed__16); +l_Lean_doElemTrace_x5b___x5d_______closed__17 = _init_l_Lean_doElemTrace_x5b___x5d_______closed__17(); +lean_mark_persistent(l_Lean_doElemTrace_x5b___x5d_______closed__17); +l_Lean_doElemTrace_x5b___x5d_______closed__18 = _init_l_Lean_doElemTrace_x5b___x5d_______closed__18(); +lean_mark_persistent(l_Lean_doElemTrace_x5b___x5d_______closed__18); +l_Lean_doElemTrace_x5b___x5d_______closed__19 = _init_l_Lean_doElemTrace_x5b___x5d_______closed__19(); +lean_mark_persistent(l_Lean_doElemTrace_x5b___x5d_______closed__19); +l_Lean_doElemTrace_x5b___x5d_______closed__20 = _init_l_Lean_doElemTrace_x5b___x5d_______closed__20(); +lean_mark_persistent(l_Lean_doElemTrace_x5b___x5d_______closed__20); +l_Lean_doElemTrace_x5b___x5d_______closed__21 = _init_l_Lean_doElemTrace_x5b___x5d_______closed__21(); +lean_mark_persistent(l_Lean_doElemTrace_x5b___x5d_______closed__21); +l_Lean_doElemTrace_x5b___x5d_______closed__22 = _init_l_Lean_doElemTrace_x5b___x5d_______closed__22(); +lean_mark_persistent(l_Lean_doElemTrace_x5b___x5d_______closed__22); +l_Lean_doElemTrace_x5b___x5d_______closed__23 = _init_l_Lean_doElemTrace_x5b___x5d_______closed__23(); +lean_mark_persistent(l_Lean_doElemTrace_x5b___x5d_______closed__23); +l_Lean_doElemTrace_x5b___x5d_______closed__24 = _init_l_Lean_doElemTrace_x5b___x5d_______closed__24(); +lean_mark_persistent(l_Lean_doElemTrace_x5b___x5d_______closed__24); +l_Lean_doElemTrace_x5b___x5d_______closed__25 = _init_l_Lean_doElemTrace_x5b___x5d_______closed__25(); +lean_mark_persistent(l_Lean_doElemTrace_x5b___x5d_______closed__25); +l_Lean_doElemTrace_x5b___x5d_______closed__26 = _init_l_Lean_doElemTrace_x5b___x5d_______closed__26(); +lean_mark_persistent(l_Lean_doElemTrace_x5b___x5d_______closed__26); +l_Lean_doElemTrace_x5b___x5d____ = _init_l_Lean_doElemTrace_x5b___x5d____(); +lean_mark_persistent(l_Lean_doElemTrace_x5b___x5d____); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__1 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__1(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__1); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__2 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__2(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__2); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__3 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__3(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__3); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__4 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__4(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__4); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__5 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__5(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__5); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__6 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__6(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__6); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__7 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__7(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__7); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__8 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__8(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__8); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__9 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__9(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__9); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__10 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__10(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__10); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__11 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__11(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__11); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__12 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__12(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__12); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__13 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__13(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__13); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__14 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__14(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__14); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__15 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__15(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__15); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__16 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__16(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__16); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__17 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__17(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__17); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__18 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__18(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__18); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__19 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__19(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__19); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__20 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__20(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__20); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__21 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__21(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__21); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__22 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__22(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__22); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__23 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__23(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__23); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__24 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__24(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__24); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__25 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__25(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__25); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__26 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__26(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__26); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__27 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__27(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__27); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__28 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__28(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__28); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__29 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__29(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__29); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__30 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__30(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__30); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__31 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__31(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__31); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__32 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__32(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__32); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__33 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__33(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__33); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__34 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__34(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__34); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__35 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__35(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__35); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__36 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__36(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__36); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__37 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__37(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__37); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__38 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__38(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__38); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__39 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__39(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__39); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__40 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__40(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__40); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__41 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__41(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__41); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__42 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__42(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__42); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__43 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__43(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__43); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__44 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__44(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__44); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__45 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__45(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__45); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__46 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__46(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__46); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__1 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__1(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__1); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__2 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__2(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__2); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__3 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__3(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__3); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__4 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__4(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__4); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__5 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__5(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__5); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__6 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__6(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__6); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__7 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__7(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__7); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__8 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__8(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__8); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__9 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__9(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__9); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__10 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__10(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__10); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__11 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__11(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__11); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__12 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__12(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__12); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__13 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__13(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__13); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__14 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__14(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__14); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__15 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__15(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__15); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__16 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__16(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__16); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__17 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__17(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__17); +l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__18 = _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__18(); +lean_mark_persistent(l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__18); l_Array_foldlMUnsafe_fold___at___private_Lean_Util_Trace_0__Lean_withNestedTracesFinalizer___spec__5___closed__1 = _init_l_Array_foldlMUnsafe_fold___at___private_Lean_Util_Trace_0__Lean_withNestedTracesFinalizer___spec__5___closed__1(); lean_mark_persistent(l_Array_foldlMUnsafe_fold___at___private_Lean_Util_Trace_0__Lean_withNestedTracesFinalizer___spec__5___closed__1); l_Array_foldlMUnsafe_fold___at___private_Lean_Util_Trace_0__Lean_withNestedTracesFinalizer___spec__5___closed__2 = _init_l_Array_foldlMUnsafe_fold___at___private_Lean_Util_Trace_0__Lean_withNestedTracesFinalizer___spec__5___closed__2(); diff --git a/stage0/stdlib/Lean/Widget/InteractiveGoal.c b/stage0/stdlib/Lean/Widget/InteractiveGoal.c index 949e08e71cc..db57ec8ae15 100644 --- a/stage0/stdlib/Lean/Widget/InteractiveGoal.c +++ b/stage0/stdlib/Lean/Widget/InteractiveGoal.c @@ -47,6 +47,7 @@ LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodingInteractiveHypothesisBundl LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodingInteractiveGoalsRpcEncodingPacket(lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* l_Lean_throwError___at___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Meta_ppGoal_ppVars___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Widget_instRpcEncodingInteractiveHypothesisBundleRpcEncodingPacket___rarg___lambda__5___closed__1; extern lean_object* l_Lean_Meta_pp_auxDecls; LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodingPacket____x40_Lean_Widget_InteractiveGoal___hyg_1032____rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -91,6 +92,7 @@ LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodingInteractiveTermGoalRpcEnco static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodingPacket____x40_Lean_Widget_InteractiveGoal___hyg_2578____rarg___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodingPacket____x40_Lean_Widget_InteractiveGoal___hyg_199____rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_toJsonRpcEncodingPacket____x40_Lean_Widget_InteractiveGoal___hyg_1225____rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_filterTRAux___at_Lean_Widget_InteractiveGoal_pretty___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodingPacket____x40_Lean_Widget_InteractiveGoal___hyg_199____rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodingInteractiveHypothesisBundleRpcEncodingPacket___rarg___lambda__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodingInteractiveHypothesisBundleRpcEncodingPacket___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -165,7 +167,6 @@ LEAN_EXPORT lean_object* l_List_mapTRAux___at_Lean_Widget_InteractiveGoal_pretty uint8_t l_Array_isEmpty___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodingPacket____x40_Lean_Widget_InteractiveGoal___hyg_1032____spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Widget_goalToInteractive___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_List_filterAux___at_Lean_Widget_InteractiveGoal_pretty___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodingInteractiveTermGoalRpcEncodingPacket___rarg___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Widget_goalToInteractive___spec__6(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instFromJsonRpcEncodingPacket__4(lean_object*); @@ -197,7 +198,6 @@ LEAN_EXPORT lean_object* l_Std_PersistentArray_forInAux___at_Lean_Widget_goalToI LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodingInteractiveGoalRpcEncodingPacket(lean_object*); static lean_object* l_Lean_Widget_goalToInteractive___closed__2; LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodingInteractiveGoalRpcEncodingPacket___rarg___lambda__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Meta_ppGoal_ppVars___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_InteractiveTermGoal_toInteractiveGoal___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodingInteractiveHypothesisBundleRpcEncodingPacket(lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodingPacket____x40_Lean_Widget_InteractiveGoal___hyg_2097____rarg___closed__1; @@ -3140,7 +3140,7 @@ return x_1; } } } -LEAN_EXPORT lean_object* l_List_filterAux___at_Lean_Widget_InteractiveGoal_pretty___spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_List_filterTRAux___at_Lean_Widget_InteractiveGoal_pretty___spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -3366,7 +3366,7 @@ x_9 = lean_ctor_get(x_7, 0); lean_inc(x_9); x_10 = lean_array_to_list(lean_box(0), x_9); x_11 = lean_box(0); -x_12 = l_List_filterAux___at_Lean_Widget_InteractiveGoal_pretty___spec__1(x_10, x_11); +x_12 = l_List_filterTRAux___at_Lean_Widget_InteractiveGoal_pretty___spec__1(x_10, x_11); x_13 = l_List_mapTRAux___at_Lean_Widget_InteractiveGoal_pretty___spec__2(x_12, x_11); x_14 = l_Array_forInUnsafe_loop___at_Lean_Widget_InteractiveGoal_pretty___spec__3___closed__1; x_15 = l_String_intercalate(x_14, x_13); @@ -5650,7 +5650,7 @@ lean_inc(x_56); x_57 = lean_ctor_get(x_55, 1); lean_inc(x_57); lean_dec(x_55); -x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Meta_ppGoal_ppVars___spec__1(x_45, x_4); +x_58 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Meta_ppGoal_ppVars___spec__1(x_45, x_4); if (x_58 == 0) { lean_object* x_59; uint8_t x_60; @@ -5661,7 +5661,7 @@ if (lean_is_scalar(x_43)) { x_59 = x_43; } lean_ctor_set(x_59, 0, x_56); -x_60 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Meta_ppGoal_ppVars___spec__1(x_45, x_59); +x_60 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Meta_ppGoal_ppVars___spec__1(x_45, x_59); lean_dec(x_59); if (x_60 == 0) { @@ -6860,7 +6860,7 @@ lean_inc(x_57); x_58 = lean_ctor_get(x_56, 1); lean_inc(x_58); lean_dec(x_56); -x_59 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Meta_ppGoal_ppVars___spec__1(x_46, x_4); +x_59 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Meta_ppGoal_ppVars___spec__1(x_46, x_4); if (x_59 == 0) { lean_object* x_60; uint8_t x_61; @@ -6871,7 +6871,7 @@ if (lean_is_scalar(x_44)) { x_60 = x_44; } lean_ctor_set(x_60, 0, x_57); -x_61 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_822____at_Lean_Meta_ppGoal_ppVars___spec__1(x_46, x_60); +x_61 = l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_800____at_Lean_Meta_ppGoal_ppVars___spec__1(x_46, x_60); lean_dec(x_60); if (x_61 == 0) { From 9bdb7033cd104f3a2207f5452b8157cc47cbf3df Mon Sep 17 00:00:00 2001 From: Sebastian Ullrich Date: Sat, 18 Jun 2022 11:16:35 +0200 Subject: [PATCH 3/4] =?UTF-8?q?refactor:=20remove=20`group(=C2=B7)`s?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- src/Init/Conv.lean | 8 ++++---- src/Init/NotationExtra.lean | 4 ++-- src/Init/Tactics.lean | 6 +++--- 3 files changed, 9 insertions(+), 9 deletions(-) diff --git a/src/Init/Conv.lean b/src/Init/Conv.lean index 8636c6c97cb..ad15eea9680 100644 --- a/src/Init/Conv.lean +++ b/src/Init/Conv.lean @@ -12,8 +12,8 @@ namespace Lean.Parser.Tactic.Conv declare_syntax_cat conv (behavior := both) -syntax convSeq1Indented := withPosition((group(colGe conv ";"?))+) -syntax convSeqBracketed := "{" (group(conv ";"?))+ "}" +syntax convSeq1Indented := withPosition((colGe conv ";"?)+) +syntax convSeqBracketed := "{" (conv ";"?)+ "}" -- Order is important: a missing `conv` proof should not be parsed as `{ }`, -- automatically closing goals syntax convSeq := convSeqBracketed <|> convSeq1Indented @@ -58,7 +58,7 @@ syntax "intro " (colGt ident)* : conv macro_rules | `(conv| intro $[$xs:ident]*) => `(conv| ext $xs*) -syntax enterArg := ident <|> group("@"? num) +syntax enterArg := ident <|> ("@"? num) syntax "enter " "[" (colGt enterArg),+ "]": conv macro_rules | `(conv| enter [$i:num]) => `(conv| arg $i) @@ -72,7 +72,7 @@ macro "trace_state" : conv => `(tactic' => trace_state) macro "apply " e:term : conv => `(tactic => apply $e) /-- `first | conv | ...` runs each `conv` until one succeeds, or else fails. -/ -syntax (name := first) "first " withPosition((group(colGe "|" convSeq))+) : conv +syntax (name := first) "first " withPosition((colGe "|" convSeq)+) : conv syntax "repeat " convSeq : conv macro_rules diff --git a/src/Init/NotationExtra.lean b/src/Init/NotationExtra.lean index 6b978e8012d..1dd7027cca3 100644 --- a/src/Init/NotationExtra.lean +++ b/src/Init/NotationExtra.lean @@ -212,14 +212,14 @@ macro_rules attribute [instance] $ctor) /-- `· tac` focuses on the main goal and tries to solve it using `tac`, or else fails. -/ -syntax ("·" <|> ".") ppHardSpace many1Indent(group(tactic ";"? ppLine)) : tactic +syntax ("·" <|> ".") ppHardSpace many1Indent(tactic ";"? ppLine) : tactic macro_rules | `(tactic| ·%$dot $[$tacs:tactic $[;%$sc]?]*) => `(tactic| {%$dot $[$tacs:tactic $[;%$sc]?]*}) /-- Similar to `first`, but succeeds only if one the given tactics solves the current goal. -/ -syntax (name := solve) "solve " withPosition((group(colGe "|" tacticSeq))+) : tactic +syntax (name := solve) "solve " withPosition((colGe "|" tacticSeq)+) : tactic macro_rules | `(tactic| solve $[| $ts]* ) => `(tactic| focus first $[| ($ts); done]*) diff --git a/src/Init/Tactics.lean b/src/Init/Tactics.lean index 074f1988f88..2eb420f85e5 100644 --- a/src/Init/Tactics.lean +++ b/src/Init/Tactics.lean @@ -141,7 +141,7 @@ syntax (name := withReducible) "with_reducible " tacticSeq : tactic syntax (name := withReducibleAndInstances) "with_reducible_and_instances " tacticSeq : tactic syntax (name := withUnfoldingAll) "with_unfolding_all " tacticSeq : tactic /-- `first | tac | ...` runs each `tac` until one succeeds, or else fails. -/ -syntax (name := first) "first " withPosition((group(colGe "|" tacticSeq))+) : tactic +syntax (name := first) "first " withPosition((colGe "|" tacticSeq)+) : tactic syntax (name := rotateLeft) "rotate_left" (num)? : tactic syntax (name := rotateRight) "rotate_right" (num)? : tactic /-- `try tac` runs `tac` and succeeds even if `tac` failed. -/ @@ -302,7 +302,7 @@ macro "let " d:letDecl : tactic => `(refine_lift let $d:letDecl; ?_) performs the unification, and replaces the target with the unified version of `t`. -/ macro "show " e:term : tactic => `(refine_lift show $e:term from ?_) -- TODO: fix, see comment -syntax (name := letrec) withPosition(atomic(group("let " &"rec ")) letRecDecls) : tactic +syntax (name := letrec) withPosition(atomic("let " &"rec ") letRecDecls) : tactic macro_rules | `(tactic| let rec $d:letRecDecls) => `(tactic| refine_lift let rec $d:letRecDecls; ?_) @@ -312,7 +312,7 @@ macro "have' " d:haveDecl : tactic => `(refine_lift' have $d:haveDecl; ?_) macro (priority := high) "have'" x:ident " := " p:term : tactic => `(have' $x:ident : _ := $p) macro "let' " d:letDecl : tactic => `(refine_lift' let $d:letDecl; ?_) -syntax inductionAltLHS := "| " (group("@"? ident) <|> "_") (ident <|> "_")* +syntax inductionAltLHS := "| " (("@"? ident) <|> "_") (ident <|> "_")* syntax inductionAlt := ppDedent(ppLine) inductionAltLHS+ " => " (hole <|> syntheticHole <|> tacticSeq) syntax inductionAlts := "with " (tactic)? withPosition( (colGe inductionAlt)+) /-- From 035d17f6d3a3472cc98ebee1a32cddd942e47492 Mon Sep 17 00:00:00 2001 From: Sebastian Ullrich Date: Sat, 18 Jun 2022 14:40:52 +0200 Subject: [PATCH 4/4] refactor: revert `toParserDescr` signature change --- src/Lean/Elab/Syntax.lean | 11 ++++++----- 1 file changed, 6 insertions(+), 5 deletions(-) diff --git a/src/Lean/Elab/Syntax.lean b/src/Lean/Elab/Syntax.lean index b668e8b9c6f..4540d4265d8 100644 --- a/src/Lean/Elab/Syntax.lean +++ b/src/Lean/Elab/Syntax.lean @@ -89,13 +89,14 @@ def resolveParserName [Monad m] [MonadInfoTree m] [MonadResolveName m] [MonadEnv open TSyntax.Compat in /-- - Given a `stx` of category `syntax`, return a pair `(newStx, lhsPrec?)`, + Given a `stx` of category `syntax`, return a `(newStx, lhsPrec?)`, where `newStx` is of category `term`. After elaboration, `newStx` should have type `TrailingParserDescr` if `lhsPrec?.isSome`, and `ParserDescr` otherwise. -/ -partial def toParserDescr (stx : Syntax) (catName : Name) : TermElabM ((Term × Nat) × Option Nat) := do +partial def toParserDescr (stx : Syntax) (catName : Name) : TermElabM (Term × Option Nat) := do let env ← getEnv let behavior := Parser.leadingIdentBehavior env catName - (process stx { catName := catName, first := true, leftRec := true, behavior := behavior }).run none + let ((newStx, _), lhsPrec?) ← (process stx { catName := catName, first := true, leftRec := true, behavior := behavior }).run none + return (newStx, lhsPrec?) where process (stx : Syntax) : ToParserDescr := withRef stx do let kind := stx.getKind @@ -341,7 +342,7 @@ def resolveSyntaxKind (k : Name) : CommandElabM Name := do let prio ← liftMacroM <| evalOptPrio prio? let stxNodeKind := (← getCurrNamespace) ++ name let catParserId := mkIdentFrom stx (cat.appendAfter "Parser") - let ((val, _), lhsPrec?) ← runTermElabM none fun _ => Term.toParserDescr syntaxParser cat + let (val, lhsPrec?) ← runTermElabM none fun _ => Term.toParserDescr syntaxParser cat let declName := mkIdentFrom stx name let d ← if let some lhsPrec := lhsPrec? then `($[$doc?:docComment]? @[$attrKind:attrKind $catParserId:ident $(quote prio):num] def $declName:ident : Lean.TrailingParserDescr := @@ -355,7 +356,7 @@ def resolveSyntaxKind (k : Name) : CommandElabM Name := do @[builtinCommandElab «syntaxAbbrev»] def elabSyntaxAbbrev : CommandElab := fun stx => do let `($[$doc?:docComment]? syntax $declName:ident := $[$ps:stx]*) ← pure stx | throwUnsupportedSyntax -- TODO: nonatomic names - let ((val, _), _) ← runTermElabM none fun _ => Term.toParserDescr (mkNullNode ps) Name.anonymous + let (val, _) ← runTermElabM none fun _ => Term.toParserDescr (mkNullNode ps) Name.anonymous let stxNodeKind := (← getCurrNamespace) ++ declName.getId let stx' ← `($[$doc?:docComment]? def $declName:ident : Lean.ParserDescr := ParserDescr.nodeWithAntiquot $(quote (toString declName.getId)) $(quote stxNodeKind) $val) withMacroExpansion stx stx' <| elabCommand stx'